Nominal-General/nominal_eqvt.ML
changeset 2477 2f289c1f6cf1
parent 2311 4da5c5c29009
--- a/Nominal-General/nominal_eqvt.ML	Sat Sep 11 05:56:49 2010 +0800
+++ b/Nominal-General/nominal_eqvt.ML	Sun Sep 12 22:46:40 2010 +0800
@@ -34,97 +34,101 @@
 val perm_cancel = @{thms permute_minus_cancel(2)}
 
 fun eqvt_rel_single_case_tac ctxt pred_names pi intro  = 
-let
-  val thy = ProofContext.theory_of ctxt
-  val cpi = Thm.cterm_of thy (mk_minus pi)
-  val pi_intro_rule = Drule.instantiate' [] [SOME cpi] perm_boolE
-  val simps1 = HOL_basic_ss addsimps @{thms permute_fun_def minus_minus split_paired_all}
-  val simps2 = HOL_basic_ss addsimps @{thms permute_bool_def}
-in
-  eqvt_strict_tac ctxt [] pred_names THEN'
-  SUBPROOF (fn {prems, context as ctxt, ...} =>
-    let
-      val prems' = map (transform_prem2 ctxt pred_names) prems
-      val tac1 = resolve_tac prems'
-      val tac2 = EVERY' [ rtac pi_intro_rule, 
-            eqvt_strict_tac ctxt perm_cancel pred_names, resolve_tac prems' ]
-      val tac3 = EVERY' [ rtac pi_intro_rule, 
-            eqvt_strict_tac ctxt perm_cancel pred_names, simp_tac simps1, 
-            simp_tac simps2, resolve_tac prems']
-    in
-      (rtac intro THEN_ALL_NEW FIRST' [tac1, tac2, tac3]) 1 
-    end) ctxt
-end
+  let
+    val thy = ProofContext.theory_of ctxt
+    val cpi = Thm.cterm_of thy (mk_minus pi)
+    val pi_intro_rule = Drule.instantiate' [] [SOME cpi] perm_boolE
+    val simps1 = HOL_basic_ss addsimps @{thms permute_fun_def minus_minus split_paired_all}
+    val simps2 = HOL_basic_ss addsimps @{thms permute_bool_def}
+  in
+    eqvt_strict_tac ctxt [] pred_names THEN'
+    SUBPROOF (fn {prems, context as ctxt, ...} =>
+      let
+        val prems' = map (transform_prem2 ctxt pred_names) prems
+        val tac1 = resolve_tac prems'
+        val tac2 = EVERY' [ rtac pi_intro_rule, 
+          eqvt_strict_tac ctxt perm_cancel pred_names, resolve_tac prems' ]
+        val tac3 = EVERY' [ rtac pi_intro_rule, 
+          eqvt_strict_tac ctxt perm_cancel pred_names, simp_tac simps1, 
+          simp_tac simps2, resolve_tac prems']
+      in
+        (rtac intro THEN_ALL_NEW FIRST' [tac1, tac2, tac3]) 1 
+      end) ctxt
+  end
 
 fun eqvt_rel_tac ctxt pred_names pi induct intros =
-let
-  val cases = map (eqvt_rel_single_case_tac ctxt pred_names pi) intros
-in
-  EVERY' (rtac induct :: cases)
-end
+  let
+    val cases = map (eqvt_rel_single_case_tac ctxt pred_names pi) intros
+  in
+    EVERY' (rtac induct :: cases)
+  end
 
 
 (** equivariance procedure *)
 
-(* sets up goal and makes sure parameters
-   are untouched PROBLEM: this violates the 
-   form of eqvt lemmas *)
 fun prepare_goal pi pred =
-let
-  val (c, xs) = strip_comb pred;
-in
-  HOLogic.mk_imp (pred, list_comb (c, map (mk_perm pi) xs))
-end
+  let
+    val (c, xs) = strip_comb pred;
+  in
+    HOLogic.mk_imp (pred, list_comb (c, map (mk_perm pi) xs))
+  end
 
 (* stores thm under name.eqvt and adds [eqvt]-attribute *)
+
 fun note_named_thm (name, thm) ctxt = 
-let
-  val thm_name = Binding.qualified_name 
-    (Long_Name.qualify (Long_Name.base_name name) "eqvt")
-  val attr = Attrib.internal (K eqvt_add)
-  val ((_, [thm']), ctxt') =  Local_Theory.note ((thm_name, [attr]), [thm]) ctxt
-in
-  (thm', ctxt')
-end
+  let
+    val thm_name = Binding.qualified_name 
+      (Long_Name.qualify (Long_Name.base_name name) "eqvt")
+    val attr = Attrib.internal (K eqvt_add)
+    val ((_, [thm']), ctxt') =  Local_Theory.note ((thm_name, [attr]), [thm]) ctxt
+  in
+    (thm', ctxt')
+  end
 
 fun equivariance note_flag pred_trms raw_induct intrs ctxt = 
-let
-  val is_already_eqvt = 
-    filter (is_eqvt ctxt) pred_trms
-    |> map (Syntax.string_of_term ctxt)
-  val _ = if null is_already_eqvt then ()
-    else error ("Already equivariant: " ^ commas is_already_eqvt)
+  let
+    val is_already_eqvt = 
+      filter (is_eqvt ctxt) pred_trms
+      |> map (Syntax.string_of_term ctxt)
+    val _ = if null is_already_eqvt then ()
+      else error ("Already equivariant: " ^ commas is_already_eqvt)
 
-  val pred_names = map (fst o dest_Const) pred_trms
-  val raw_induct' = atomize_induct ctxt raw_induct
-  val intrs' = map atomize_intr intrs
-  val (([raw_concl], [raw_pi]), ctxt') = 
-    ctxt 
-    |> Variable.import_terms false [concl_of raw_induct'] 
-    ||>> Variable.variant_fixes ["p"]
-  val pi = Free (raw_pi, @{typ perm})
-  val preds = map (fst o HOLogic.dest_imp)
-    (HOLogic.dest_conj (HOLogic.dest_Trueprop raw_concl));
-  val goal = HOLogic.mk_Trueprop 
-    (foldr1 HOLogic.mk_conj (map (prepare_goal pi) preds))
-  val thms = Datatype_Aux.split_conj_thm (Goal.prove ctxt' [] [] goal 
-    (fn {context,...} => eqvt_rel_tac context pred_names pi raw_induct' intrs' 1)
-    |> singleton (ProofContext.export ctxt' ctxt))
-  val thms' = map (fn th => zero_var_indexes (th RS mp)) thms
-in
-  if note_flag
-  then ctxt |> fold_map note_named_thm (pred_names ~~ thms')  
-  else (thms', ctxt) 
-end
+    val pred_names = map (fst o dest_Const) pred_trms
+    val raw_induct' = atomize_induct ctxt raw_induct
+    val intrs' = map atomize_intr intrs
+  
+    val (([raw_concl], [raw_pi]), ctxt') = 
+      ctxt 
+      |> Variable.import_terms false [concl_of raw_induct'] 
+      ||>> Variable.variant_fixes ["p"]
+    val pi = Free (raw_pi, @{typ perm})
+  
+    val preds = map (fst o HOLogic.dest_imp)
+      (HOLogic.dest_conj (HOLogic.dest_Trueprop raw_concl));
+  
+    val goal = HOLogic.mk_Trueprop 
+      (foldr1 HOLogic.mk_conj (map (prepare_goal pi) preds))
+  
+    val thms = Goal.prove ctxt' [] [] goal 
+      (fn {context,...} => eqvt_rel_tac context pred_names pi raw_induct' intrs' 1)
+      |> Datatype_Aux.split_conj_thm 
+      |> ProofContext.export ctxt' ctxt
+      |> map (fn th => th RS mp)
+      |> map zero_var_indexes
+  in
+    if note_flag
+    then fold_map note_named_thm (pred_names ~~ thms) ctxt 
+    else (thms, ctxt) 
+  end
 
 fun equivariance_cmd pred_name ctxt =
-let
-  val thy = ProofContext.theory_of ctxt
-  val (_, {preds, raw_induct, intrs, ...}) =
-    Inductive.the_inductive ctxt (Sign.intern_const thy pred_name)
-in
-  equivariance true preds raw_induct intrs ctxt |> snd
-end
+  let
+    val thy = ProofContext.theory_of ctxt
+    val (_, {preds, raw_induct, intrs, ...}) =
+      Inductive.the_inductive ctxt (Sign.intern_const thy pred_name)
+  in
+    equivariance true preds raw_induct intrs ctxt |> snd
+  end
 
 local structure P = Parse and K = Keyword in
 
@@ -132,6 +136,7 @@
   Outer_Syntax.local_theory "equivariance"
     "Proves equivariance for inductive predicate involving nominal datatypes." 
       K.thy_decl (P.xname >> equivariance_cmd);
+
 end;
 
 end (* structure *)