--- a/LFex.thy Fri Dec 04 09:01:13 2009 +0100
+++ b/LFex.thy Fri Dec 04 09:08:51 2009 +0100
@@ -298,10 +298,7 @@
\<And>trm1 trm2. \<lbrakk>P3 trm1; P3 trm2\<rbrakk> \<Longrightarrow> P3 (APP trm1 trm2);
\<And>ty name trm. \<lbrakk>P2 ty; P3 trm\<rbrakk> \<Longrightarrow> P3 (LAM ty name trm)\<rbrakk>
\<Longrightarrow> P1 mkind \<and> P2 mty \<and> P3 mtrm"
-apply(tactic {* procedure_tac @{context} @{thm kind_ty_trm.induct} 1 *})
-apply(tactic {* regularize_tac @{context} @{thms alpha_EQUIVs} 1 *})
-apply(tactic {* all_inj_repabs_tac @{context} quot rel_refl trans2 1 *})
-apply(tactic {* clean_tac @{context} quot 1 *})
+apply(tactic {* lift_tac @{context} @{thm kind_ty_trm.induct} @{thms alpha_EQUIVs} quot 1 *})
done
print_quotients