--- a/LFex.thy Fri Dec 04 15:20:06 2009 +0100
+++ b/LFex.thy Fri Dec 04 15:25:26 2009 +0100
@@ -67,18 +67,18 @@
apply(auto intro: akind_aty_atrm.intros)
done
-lemma alpha_EQUIVs:
- shows "EQUIV akind"
- and "EQUIV aty"
- and "EQUIV atrm"
+lemma alpha_equivps:
+ shows "equivp akind"
+ and "equivp aty"
+ and "equivp atrm"
sorry
quotient KIND = kind / akind
- by (rule alpha_EQUIVs)
+ by (rule alpha_equivps)
quotient TY = ty / aty
and TRM = trm / atrm
- by (auto intro: alpha_EQUIVs)
+ by (auto intro: alpha_equivps)
print_quotients
@@ -217,10 +217,10 @@
thm kind_ty_trm.induct
ML {*
- val quot = @{thms QUOTIENT_KIND QUOTIENT_TY QUOTIENT_TRM}
- val rel_refl = map (fn x => @{thm EQUIV_REFL} OF [x]) @{thms alpha_EQUIVs}
- (*val trans2 = map (fn x => @{thm equiv_trans2} OF [x]) @{thms alpha_EQUIVs}*)
- val reps_same = map (fn x => @{thm QUOTIENT_REL_REP} OF [x]) quot
+ val quot = @{thms Quotient_KIND Quotient_TY Quotient_TRM}
+ val rel_refl = map (fn x => @{thm equivp_reflp} OF [x]) @{thms alpha_equivps}
+ (*val trans2 = map (fn x => @{thm EQUALS_PRS} OF [x]) quot*)
+ val reps_same = map (fn x => @{thm Quotient_REL_REP} OF [x]) quot
val meta_reps_same = map (fn x => @{thm eq_reflection} OF [x]) reps_same
*}
@@ -258,8 +258,8 @@
((x5 :: TRM) = x6 \<longrightarrow> 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_EQUIVs} 1 *})
-(*apply(tactic {* regularize_tac @{context} @{thms alpha_EQUIVs} 1 *})
+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 {* all_inj_repabs_tac @{context} rel_refl trans2 1 *})
apply(tactic {* clean_tac @{context} 1 *})
*)
@@ -268,7 +268,7 @@
Profiling:
ML_prf {* fun ith i = (#concl (fst (Subgoal.focus @{context} i (#goal (Isar.goal ()))))) *}
ML_prf {* profile 2 Seq.list_of ((clean_tac @{context} quot defs 1) (ith 3)) *}
-ML_prf {* profile 2 Seq.list_of ((regularize_tac @{context} @{thms alpha_EQUIVs} 1) (ith 1)) *}
+ML_prf {* profile 2 Seq.list_of ((regularize_tac @{context} @{thms alpha_equivps} 1) (ith 1)) *}
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 ()))) *}
*)
@@ -298,7 +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 {* lift_tac @{context} @{thm kind_ty_trm.induct} @{thms alpha_EQUIVs} 1 *})
+apply(tactic {* lift_tac @{context} @{thm kind_ty_trm.induct} @{thms alpha_equivps} 1 *})
done
print_quotients