Use raw_induct instead of induct
authorCezary Kaliszyk <kaliszyk@in.tum.de>
Wed, 12 May 2010 16:08:32 +0200
changeset 2108 c5b7be27f105
parent 2106 409ecb7284dd
child 2109 287aa0d3d23a
Use raw_induct instead of induct
Nominal/Equivp.thy
Nominal/NewAlpha.thy
Nominal/Rsp.thy
Nominal/Tacs.thy
--- a/Nominal/Equivp.thy	Wed May 12 13:43:48 2010 +0100
+++ b/Nominal/Equivp.thy	Wed May 12 16:08:32 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/NewAlpha.thy	Wed May 12 13:43:48 2010 +0100
+++ b/Nominal/NewAlpha.thy	Wed May 12 16:08:32 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/Rsp.thy	Wed May 12 13:43:48 2010 +0100
+++ b/Nominal/Rsp.thy	Wed May 12 16:08:32 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_basic_ss addsimps @{thms alphas2}) THEN_ALL_NEW
   asm_full_simp_tac (HOL_ss addsimps (@{thms alphas prod_rel.simps prod_fv.simps} @ fvbv_simps)) THEN_ALL_NEW
@@ -96,7 +96,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 13:43:48 2010 +0100
+++ b/Nominal/Tacs.thy	Wed May 12 16:08:32 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