merge
authorCezary Kaliszyk <kaliszyk@in.tum.de>
Wed, 12 May 2010 16:11:23 +0200
changeset 2112 7c9746795767
parent 2109 287aa0d3d23a (diff)
parent 2111 502b1f3b282a (current diff)
child 2113 af11e9fbc45a
merge
Nominal/Rsp.thy
--- a/Nominal-General/nominal_eqvt.ML	Wed May 12 16:11:03 2010 +0200
+++ b/Nominal-General/nominal_eqvt.ML	Wed May 12 16:11:23 2010 +0200
@@ -7,9 +7,11 @@
 
 signature NOMINAL_EQVT =
 sig
-  val equivariance: string -> Proof.context -> local_theory
   val eqvt_rel_tac: Proof.context -> string list -> term -> thm -> thm list -> int -> tactic
   val eqvt_rel_single_case_tac: Proof.context -> string list -> term -> thm -> int -> tactic
+  
+  val equivariance: term list -> thm -> thm list -> Proof.context -> (string * thm list) list * local_theory
+  val equivariance_cmd: string -> Proof.context -> local_theory
 end
 
 structure Nominal_Eqvt : NOMINAL_EQVT =
@@ -127,12 +129,11 @@
 (* sets up goal and makes sure parameters
    are untouched PROBLEM: this violates the 
    form of eqvt lemmas *)
-fun prepare_goal params_no pi pred =
+fun prepare_goal pi pred =
 let
   val (c, xs) = strip_comb pred;
-  val (xs1, xs2) = chop params_no xs
 in
-  HOLogic.mk_imp (pred, list_comb (c, xs1 @ map (mk_perm pi) xs2))
+  HOLogic.mk_imp (pred, list_comb (c, map (mk_perm pi) xs))
 end
 
 (* stores thm under name.eqvt and adds [eqvt]-attribute *)
@@ -145,38 +146,43 @@
   Local_Theory.note ((thm_name, [attr]), [thm]) ctxt
 end
 
-fun equivariance pred_name ctxt = 
+fun equivariance pred_trms raw_induct intrs ctxt = 
 let
-  val thy = ProofContext.theory_of ctxt
-  val ({names, ...}, {raw_induct, intrs, ...}) =
-    Inductive.the_inductive ctxt (Sign.intern_const thy pred_name)
-  val raw_induct = atomize_induct ctxt raw_induct
-  val intros = map atomize_intr intrs
-  val params_no = length (Inductive.params_of raw_induct)
+  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.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 params_no pi) preds))
+    (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 names pi raw_induct intros 1)
+    (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
-  ctxt |> fold_map note_named_thm (names ~~ thms') |> snd  
+  ctxt |> fold_map note_named_thm (pred_names ~~ thms')   
 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 preds raw_induct intrs ctxt |> snd
+end
 
 local structure P = OuterParse and K = OuterKeyword in
 
 val _ =
   OuterSyntax.local_theory "equivariance"
     "Proves equivariance for inductive predicate involving nominal datatypes." 
-      K.thy_decl (P.xname >> equivariance);
+      K.thy_decl (P.xname >> equivariance_cmd);
 end;
 
 end (* structure *)
--- a/Nominal/Equivp.thy	Wed May 12 16:11:03 2010 +0200
+++ b/Nominal/Equivp.thy	Wed May 12 16:11:23 2010 +0200
@@ -76,7 +76,7 @@
 
 ML {*
 fun symp_tac induct inj eqvt ctxt =
-  rel_indtac induct THEN_ALL_NEW
+  rtac induct THEN_ALL_NEW
   simp_tac (HOL_basic_ss addsimps inj) THEN_ALL_NEW split_conj_tac
   THEN_ALL_NEW
   REPEAT o etac @{thm exi_neg}
@@ -111,7 +111,7 @@
 
 ML {*
 fun transp_tac ctxt induct alpha_inj term_inj distinct cases eqvt =
-  rel_indtac induct THEN_ALL_NEW
+  rtac induct THEN_ALL_NEW
   (TRY o rtac allI THEN' imp_elim_tac cases ctxt) THEN_ALL_NEW
   asm_full_simp_tac (HOL_basic_ss addsimps alpha_inj) THEN_ALL_NEW
   split_conj_tac THEN_ALL_NEW REPEAT o (eetac @{thm exi_sum} ctxt) THEN_ALL_NEW split_conj_tac
@@ -222,7 +222,7 @@
 *}
 
 ML {*
-fun fs_tac induct supports = rel_indtac induct THEN_ALL_NEW (
+fun fs_tac induct supports = rtac induct THEN_ALL_NEW (
   rtac @{thm supports_finite} THEN' resolve_tac supports) THEN_ALL_NEW
   asm_full_simp_tac (HOL_ss addsimps @{thms supp_atom supp_atom_image supp_fset_to_set
     supp_fmap_atom finite_insert finite.emptyI finite_Un finite_supp})
@@ -317,7 +317,7 @@
 
 ML {*
 fun supp_eq_tac ind fv perm eqiff ctxt =
-  rel_indtac ind THEN_ALL_NEW
+  rtac ind THEN_ALL_NEW
   asm_full_simp_tac (HOL_basic_ss addsimps fv) THEN_ALL_NEW
   asm_full_simp_tac (HOL_basic_ss addsimps @{thms supp_abs_atom[symmetric]}) THEN_ALL_NEW
   asm_full_simp_tac (HOL_basic_ss addsimps [choose_alpha_abs eqiff]) THEN_ALL_NEW
--- a/Nominal/Ex/SingleLet.thy	Wed May 12 16:11:03 2010 +0200
+++ b/Nominal/Ex/SingleLet.thy	Wed May 12 16:11:23 2010 +0200
@@ -16,7 +16,7 @@
 where
   "bn (As x t) = {atom x}"
 
-
+ML {* Inductive.the_inductive *}
 thm trm_assg.fv
 thm trm_assg.supp
 thm trm_assg.eq_iff
@@ -31,6 +31,7 @@
 equivariance alpha_trm_raw
 
 
+
 end
 
 
--- a/Nominal/NewAlpha.thy	Wed May 12 16:11:03 2010 +0200
+++ b/Nominal/NewAlpha.thy	Wed May 12 16:11:23 2010 +0200
@@ -242,7 +242,7 @@
      all_alpha_names [] all_alpha_eqs [] lthy
 
   val alpha_ts_loc = #preds alphas;
-  val alpha_induct_loc = #induct alphas;
+  val alpha_induct_loc = #raw_induct alphas;
   val alpha_intros_loc = #intrs alphas;
   val alpha_cases_loc = #elims alphas;
   val morphism = ProofContext.export_morphism lthy' lthy;
--- a/Nominal/NewParser.thy	Wed May 12 16:11:03 2010 +0200
+++ b/Nominal/NewParser.thy	Wed May 12 16:11:23 2010 +0200
@@ -367,11 +367,6 @@
     if STEPS > 1 then raw_nominal_decls dts bn_funs bn_eqs bclauses lthy
     else raise TEST lthy
 
-  val _ = tracing ("raw_bn_eqs\n" ^ cat_lines (map (@{make_string} o prop_of) raw_bn_eqs))
-  
-  val _ = tracing ("raw_bns\n" ^ @{make_string} raw_bns)
-  val _ = tracing ("raw_bclauses\n" ^ @{make_string} raw_bclauses)
-
   val dtinfo = Datatype.the_info (ProofContext.theory_of lthy1) (hd raw_dt_names);
   val {descr, sorts, ...} = dtinfo;
   val raw_tys = map (fn (i, _) => nth_dtyp descr sorts i) descr;
@@ -406,12 +401,11 @@
   val lthy3 = Theory_Target.init NONE thy;
   val raw_bn_funs = map (fn (f, _, _) => f) bn_funs_decls;
 
-  (*
   val _ = tracing ("raw_bns\n" ^ @{make_string} raw_bns)
   val _ = tracing ("raw_bns\n" ^ @{make_string} raw_bns_exp)
   val _ = tracing ("bn_funs\n" ^ @{make_string} bn_funs_decls)
   val _ = tracing ("raw_bclauses\n" ^ @{make_string} raw_bclauses)
-  *)
+  
   val ((fv, fvbn), fv_def, lthy3a) = 
     if STEPS > 3 then define_raw_fv descr sorts bn_funs_decls raw_bclauses lthy3
     else raise TEST lthy3
--- a/Nominal/Rsp.thy	Wed May 12 16:11:03 2010 +0200
+++ b/Nominal/Rsp.thy	Wed May 12 16:11:23 2010 +0200
@@ -60,7 +60,7 @@
 
 ML {*
 fun fvbv_rsp_tac induct fvbv_simps ctxt =
-  rel_indtac induct THEN_ALL_NEW
+  rtac induct THEN_ALL_NEW
   (TRY o rtac @{thm TrueI}) THEN_ALL_NEW
   asm_full_simp_tac (HOL_ss addsimps (@{thms prod_fv.simps prod_rel.simps alphas} @ fvbv_simps)) THEN_ALL_NEW
   REPEAT o eresolve_tac [conjE, exE] THEN_ALL_NEW
@@ -95,7 +95,7 @@
 
 ML {*
 fun alpha_eqvt_tac induct simps ctxt =
-  rel_indtac induct THEN_ALL_NEW
+  rtac induct THEN_ALL_NEW
   simp_tac (HOL_basic_ss addsimps simps) THEN_ALL_NEW split_conj_tac THEN_ALL_NEW
   REPEAT o etac @{thm exi[of _ _ "p"]} THEN' split_conj_tac THEN_ALL_NEW
   asm_full_simp_tac (HOL_ss addsimps (all_eqvts ctxt @ simps)) THEN_ALL_NEW
--- a/Nominal/Tacs.thy	Wed May 12 16:11:03 2010 +0200
+++ b/Nominal/Tacs.thy	Wed May 12 16:11:23 2010 +0200
@@ -56,12 +56,6 @@
 end
 *}
 
-(* An induction for a single relation is "R x y \<Longrightarrow> P x y"
-   but for multiple relations is "(R1 x y \<longrightarrow> P x y) \<and> (R2 a b \<longrightarrow> P2 a b)" *)
-ML {*
-fun rel_indtac induct = (rtac impI THEN' etac induct) ORELSE' rtac induct
-*}
-
 ML {*
 fun prove_by_rel_induct alphas build_goal ind utac inputs ctxt =
 let
@@ -80,7 +74,7 @@
     (fn ((alpha, gl), (l, r)) => HOLogic.mk_imp (alpha $ l $ r, gl)) 
     ((alphas ~~ trm_gls) ~~ (freesl ~~ freesr))
   val gl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj pgls);
-  fun tac {context,...} = (rel_indtac ind THEN_ALL_NEW split_conj_tac THEN_ALL_NEW
+  fun tac {context,...} = (rtac ind THEN_ALL_NEW split_conj_tac THEN_ALL_NEW
     TRY o rtac @{thm TrueI} THEN_ALL_NEW utac context) 1
   val th_loc = Goal.prove ctxt'' [] [] gl tac
   val ths_loc = HOLogic.conj_elims th_loc