--- a/Quot/Examples/LamEx.thy Mon Dec 07 18:49:14 2009 +0100
+++ b/Quot/Examples/LamEx.thy Mon Dec 07 21:53:50 2009 +0100
@@ -169,54 +169,48 @@
apply (simp_all add: rlam.inject alpha_refl)
done
-ML {* val qty = @{typ "lam"} *}
-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 (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "lam" *}
-ML {* fun lift_tac_lam lthy t = lift_tac lthy t *}
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 *})
+apply (tactic {* lift_tac @{context} @{thm pi_var_com} 1 *})
done
lemma pi_app: "(pi\<Colon>('x \<times> 'x) list) \<bullet> App (x\<Colon>lam) (xa\<Colon>lam) = App (pi \<bullet> x) (pi \<bullet> xa)"
-apply (tactic {* lift_tac_lam @{context} @{thm pi_app_com} 1 *})
+apply (tactic {* lift_tac @{context} @{thm pi_app_com} 1 *})
done
lemma pi_lam: "(pi\<Colon>('x \<times> 'x) list) \<bullet> Lam (a\<Colon>name) (x\<Colon>lam) = Lam (pi \<bullet> a) (pi \<bullet> x)"
-apply (tactic {* lift_tac_lam @{context} @{thm pi_lam_com} 1 *})
+apply (tactic {* lift_tac @{context} @{thm pi_lam_com} 1 *})
done
lemma fv_var: "fv (Var (a\<Colon>name)) = {a}"
-apply (tactic {* lift_tac_lam @{context} @{thm rfv_var} 1 *})
+apply (tactic {* lift_tac @{context} @{thm rfv_var} 1 *})
done
lemma fv_app: "fv (App (x\<Colon>lam) (xa\<Colon>lam)) = fv x \<union> fv xa"
-apply (tactic {* lift_tac_lam @{context} @{thm rfv_app} 1 *})
+apply (tactic {* lift_tac @{context} @{thm rfv_app} 1 *})
done
lemma fv_lam: "fv (Lam (a\<Colon>name) (x\<Colon>lam)) = fv x - {a}"
-apply (tactic {* lift_tac_lam @{context} @{thm rfv_lam} 1 *})
+apply (tactic {* lift_tac @{context} @{thm rfv_lam} 1 *})
done
lemma a1: "(a\<Colon>name) = (b\<Colon>name) \<Longrightarrow> Var a = Var b"
-apply (tactic {* lift_tac_lam @{context} @{thm a1} 1 *})
+apply (tactic {* lift_tac @{context} @{thm a1} 1 *})
done
lemma a2: "\<lbrakk>(x\<Colon>lam) = (xa\<Colon>lam); (xb\<Colon>lam) = (xc\<Colon>lam)\<rbrakk> \<Longrightarrow> App x xb = App xa xc"
-apply (tactic {* lift_tac_lam @{context} @{thm a2} 1 *})
+apply (tactic {* lift_tac @{context} @{thm a2} 1 *})
done
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 (tactic {* lift_tac @{context} @{thm a3} 1 *})
done
lemma alpha_cases: "\<lbrakk>a1 = a2; \<And>a b. \<lbrakk>a1 = Var a; a2 = Var b; a = b\<rbrakk> \<Longrightarrow> P;
\<And>x xa xb xc. \<lbrakk>a1 = App x xb; a2 = App xa xc; x = xa; xb = xc\<rbrakk> \<Longrightarrow> P;
\<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 (tactic {* lift_tac @{context} @{thm alpha.cases} 1 *})
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);
@@ -224,16 +218,16 @@
\<And>(x\<Colon>lam) (a\<Colon>name) (b\<Colon>name) xa\<Colon>lam.
\<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 (tactic {* lift_tac @{context} @{thm alpha.induct} 1 *})
done
lemma var_inject: "(Var a = Var b) = (a = b)"
-apply (tactic {* lift_tac_lam @{context} @{thm rvar_inject} 1 *})
+apply (tactic {* lift_tac @{context} @{thm rvar_inject} 1 *})
done
lemma lam_induct:" \<lbrakk>\<And>name. P (Var name); \<And>lam1 lam2. \<lbrakk>P lam1; P lam2\<rbrakk> \<Longrightarrow> P (App lam1 lam2);
\<And>name lam. P lam \<Longrightarrow> P (Lam name lam)\<rbrakk> \<Longrightarrow> P lam"
-apply (tactic {* lift_tac_lam @{context} @{thm rlam.induct} 1 *})
+apply (tactic {* lift_tac @{context} @{thm rlam.induct} 1 *})
done
lemma var_supp: