Updated the examples
authorCezary Kaliszyk <kaliszyk@in.tum.de>
Thu, 03 Dec 2009 13:45:52 +0100
changeset 500 184d74813679
parent 499 f122816d7729
child 501 375e28eedee7
Updated the examples
IntEx.thy
LFex.thy
LamEx.thy
--- a/IntEx.thy	Thu Dec 03 12:33:05 2009 +0100
+++ b/IntEx.thy	Thu Dec 03 13:45:52 2009 +0100
@@ -135,12 +135,10 @@
 ML {* val qty = @{typ "my_int"} *}
 ML {* val ty_name = "my_int" *}
 ML {* val rsp_thms = @{thms ho_plus_rsp} *}
-ML {* val defs = @{thms PLUS_def} *}
 ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
 ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "my_int"; *}
-ML {* val consts = lookup_quot_consts defs *}
 
-ML {* fun lift_tac_intex lthy t = lift_tac lthy t [rel_eqv] [quot] defs *}
+ML {* fun lift_tac_intex lthy t = lift_tac lthy t [rel_eqv] [quot] *}
 
 ML {* fun inj_repabs_tac_intex lthy = inj_repabs_tac lthy [quot] [rel_refl] [trans2] *}
 ML {* fun all_inj_repabs_tac_intex lthy = all_inj_repabs_tac lthy [quot] [rel_refl] [trans2] *}
@@ -150,7 +148,7 @@
 apply(tactic {* procedure_tac @{context} @{thm plus_sym_pre} 1 *})
 apply(tactic {* regularize_tac @{context} [rel_eqv] 1 *})
 prefer 2
-apply(tactic {* clean_tac @{context} [quot] defs 1 *})
+apply(tactic {* clean_tac @{context} [quot] 1 *})
 apply(tactic {* inj_repabs_tac_intex @{context} 1*})
 apply(tactic {* inj_repabs_tac_intex @{context} 1*})
 apply(tactic {* inj_repabs_tac_intex @{context} 1*})
@@ -188,7 +186,7 @@
 apply(tactic {* procedure_tac @{context} @{thm plus_assoc_pre} 1 *})
 apply(tactic {* regularize_tac @{context} [rel_eqv] 1 *})
 apply(tactic {* all_inj_repabs_tac_intex @{context} 1*})
-apply(tactic {* clean_tac @{context} [quot] defs 1 *})
+apply(tactic {* clean_tac @{context} [quot] 1 *})
 done
 
 lemma ho_tst: "foldl my_plus x [] = x"
@@ -285,7 +283,7 @@
 apply(tactic {* all_inj_repabs_tac_intex @{context} 1*})
 apply(simp only: foldl_prs[OF QUOTIENT_my_int QUOTIENT_my_int])
 apply(simp only: nil_prs)
-apply(tactic {* clean_tac @{context} [quot] defs 1 *})
+apply(tactic {* clean_tac @{context} [quot] 1 *})
 done
 
 lemma ho_tst2: "foldl my_plus x (h # t) \<approx> my_plus h (foldl my_plus x t)"
@@ -297,6 +295,6 @@
 apply(tactic {* all_inj_repabs_tac_intex @{context} 1*})
 apply(simp only: foldl_prs[OF QUOTIENT_my_int QUOTIENT_my_int])
 apply(simp only: cons_prs[OF QUOTIENT_my_int])
-apply(tactic {* clean_tac @{context} [quot] defs 1 *})
+apply(tactic {* clean_tac @{context} [quot] 1 *})
 done
 
--- a/LFex.thy	Thu Dec 03 12:33:05 2009 +0100
+++ b/LFex.thy	Thu Dec 03 13:45:52 2009 +0100
@@ -266,17 +266,16 @@
 apply(tactic {* procedure_tac @{context} @{thm akind_aty_atrm.induct} 1 *})
 apply(tactic {* regularize_tac @{context} @{thms alpha_EQUIVs} 1 *})
 prefer 2
-apply(tactic {* clean_tac @{context} quot defs 1 *})
-
+apply(tactic {* simp_tac ((Simplifier.context @{context} empty_ss) addsimps @{thms perm_kind_def perm_ty_def perm_trm_def}) 1 *})
+apply(tactic {* clean_tac @{context} quot 1 *})
 (*
 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 {* PolyML.profiling 1 *}
 ML_prf {* profile 2 Seq.list_of ((all_inj_repabs_tac @{context} quot rel_refl trans2 1) (#goal (Isar.goal ()))) *}
-
+*)
 apply(tactic {* all_inj_repabs_tac @{context} quot rel_refl trans2 1 *})
 done
 
@@ -307,7 +306,7 @@
 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 defs 1 *})
+apply(tactic {* clean_tac @{context} quot 1 *})
 done
 
 print_quotients
--- a/LamEx.thy	Thu Dec 03 12:33:05 2009 +0100
+++ b/LamEx.thy	Thu Dec 03 13:45:52 2009 +0100
@@ -170,13 +170,11 @@
 done
 
 ML {* val qty = @{typ "lam"} *}
-ML {* val defs = @{thms Var_def App_def Lam_def perm_lam_def fv_def} *}
 ML {* val rsp_thms = @{thms perm_rsp fresh_rsp rVar_rsp rApp_rsp rLam_rsp rfv_rsp} *}
 
 ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
-ML {* val consts = lookup_quot_consts defs *}
 ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "lam" *}
-ML {* fun lift_tac_lam lthy t = lift_tac lthy t [rel_eqv] rty [quot] defs *}
+ML {* fun lift_tac_lam lthy t = lift_tac lthy t [rel_eqv] [quot] *}
 
 lemma pi_var: "(pi\<Colon>('x \<times> 'x) list) \<bullet> Var a = Var (pi \<bullet> a)"
 apply (tactic {* lift_tac_lam @{context} @{thm pi_var_com} 1 *})
@@ -212,6 +210,7 @@
 
 lemma a3: "\<lbrakk>(x\<Colon>lam) = [(a\<Colon>name, b\<Colon>name)] \<bullet> (xa\<Colon>lam); a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> Lam a x = Lam b xa"
 apply (tactic {* lift_tac_lam @{context} @{thm a3} 1 *})
+apply (simp add:perm_lam_def)
 done
 
 lemma alpha_cases: "\<lbrakk>a1 = a2; \<And>a b. \<lbrakk>a1 = Var a; a2 = Var b; a = b\<rbrakk> \<Longrightarrow> P;
@@ -219,6 +218,7 @@
      \<And>x a b xa. \<lbrakk>a1 = Lam a x; a2 = Lam b xa; x = [(a, b)] \<bullet> xa; a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> P\<rbrakk>
     \<Longrightarrow> P"
 apply (tactic {* lift_tac_lam @{context} @{thm alpha.cases} 1 *})
+apply (simp add:perm_lam_def)
 done
 
 lemma alpha_induct: "\<lbrakk>(qx\<Colon>lam) = (qxa\<Colon>lam); \<And>(a\<Colon>name) b\<Colon>name. a = b \<Longrightarrow> (qxb\<Colon>lam \<Rightarrow> lam \<Rightarrow> bool) (Var a) (Var b);
@@ -227,6 +227,7 @@
         \<lbrakk>x = [(a, b)] \<bullet> xa; qxb x ([(a, b)] \<bullet> xa); a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> qxb (Lam a x) (Lam b xa)\<rbrakk>
     \<Longrightarrow> qxb qx qxa"
 apply (tactic {* lift_tac_lam @{context} @{thm alpha.induct} 1 *})
+apply (simp add:perm_lam_def)
 done
 
 lemma var_inject: "(Var a = Var b) = (a = b)"
@@ -337,7 +338,7 @@
 prefer 2
 apply (tactic {* all_inj_repabs_tac @{context} rty [quot] [rel_refl] [trans2] 1 *})
 prefer 3
-apply (tactic {* clean_tac @{context} [quot] defs 1 *})
+apply (tactic {* clean_tac @{context} [quot] 1 *})
 
 thm all_prs
 thm REP_ABS_RSP