diff -r 0384e039b7f2 -r e121ac0028f8 LFex.thy --- a/LFex.thy Sun Dec 06 00:13:35 2009 +0100 +++ b/LFex.thy Sun Dec 06 00:19:45 2009 +0100 @@ -220,6 +220,7 @@ val quot = @{thms Quotient_KIND Quotient_TY Quotient_TRM} val rel_refl = map (fn x => @{thm equivp_reflp} OF [x]) @{thms alpha_equivps} val reps_same = map (fn x => @{thm Quotient_rel_rep} OF [x]) quot + val trans2 = map (fn x => @{thm equals_rsp} OF [x]) quot *} lemma @@ -256,12 +257,11 @@ ((x5 :: TRM) = x6 \ P3 x5 x6)" using a0 a1 a2 a3 a4 a5 a6 a7 a8 a9 a10 a11 apply - -apply(tactic {* lift_tac @{context} @{thm akind_aty_atrm.induct} @{thms alpha_equivps} 1 *}) -(*apply(tactic {* regularize_tac @{context} @{thms alpha_equivps} 1 *}) +apply(tactic {* procedure_tac @{context} @{thm akind_aty_atrm.induct} 1 *}) +apply(tactic {* regularize_tac @{context} @{thms alpha_equivps} 1 *}) apply(tactic {* all_inj_repabs_tac @{context} rel_refl trans2 1 *}) +apply(fold perm_kind_def perm_ty_def perm_trm_def FV_ty_def[simplified id_simps] FV_kind_def[simplified id_simps] FV_trm_def[simplified id_simps]) apply(tactic {* clean_tac @{context} 1 *}) -*) -(*apply(tactic {* all_inj_repabs_tac @{context} rel_refl trans2 1 *})*) (* Profiling: ML_prf {* fun ith i = (#concl (fst (Subgoal.focus @{context} i (#goal (Isar.goal ()))))) *} @@ -270,8 +270,6 @@ ML_prf {* PolyML.profiling 1 *} ML_prf {* profile 2 Seq.list_of ((all_inj_repabs_tac @{context} quot rel_refl trans2 1) (#goal (Isar.goal ()))) *} *) -apply(unfold perm_kind_def perm_ty_def perm_trm_def) -apply(tactic {* clean_tac @{context} 1 *}) done (* Does not work: