--- a/LamEx.thy Wed Nov 25 03:47:07 2009 +0100
+++ b/LamEx.thy Wed Nov 25 10:34:03 2009 +0100
@@ -119,18 +119,6 @@
shows "(pi\<bullet>rLam a t) \<approx> rLam (pi\<bullet>a) (pi\<bullet>t)"
sorry
-lemma fv_var:
- shows "fv (Var a) = {a}"
-sorry
-
-lemma fv_app:
- shows "fv (App t1 t2) = (fv t1) \<union> (fv t2)"
-sorry
-
-lemma fv_lam:
- shows "fv (Lam a t) = (fv t) - {a}"
-sorry
-
lemma real_alpha:
assumes "t = [(a,b)]\<bullet>s" "a\<sharp>[b].s"
shows "Lam a t = Lam b s"
@@ -194,51 +182,48 @@
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 rel_refl rty quot trans2 rsp_thms reps_same defs *}
-lemma "(pi\<Colon>('x \<times> 'x) list) \<bullet> Var a = Var (pi \<bullet> a)"
+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 *})
done
-ML {* val pi_app = lift_thm_lam @{context} @{thm pi_app_com} *}
-lemma "(pi\<Colon>('x \<times> 'x) list) \<bullet> App (x\<Colon>lam) (xa\<Colon>lam) = App (pi \<bullet> x) (pi \<bullet> xa)"
+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 *})
done
-ML {* val pi_lam = lift_thm_lam @{context} @{thm pi_lam_com} *}
-lemma "(pi\<Colon>('x \<times> 'x) list) \<bullet> Lam (a\<Colon>name) (x\<Colon>lam) = Lam (pi \<bullet> a) (pi \<bullet> x)"
+
+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 *})
done
-ML {* val fv_var = lift_thm_lam @{context} @{thm rfv_var} *}
-lemma "\<forall>a. fv (Var (a\<Colon>name)) = {a}"
+lemma fv_var: "fv (Var (a\<Colon>name)) = {a}"
apply (tactic {* lift_tac_lam @{context} @{thm rfv_var} 1 *})
done
-ML {* val fv_app = lift_thm_lam @{context} @{thm rfv_app} *}
-lemma "fv (App (x\<Colon>lam) (xa\<Colon>lam)) = fv x \<union> fv xa"
-(*apply (tactic {* lift_tac_lam @{context} @{thm rfv_app} 1 *})*)
-sorry
-ML {* val fv_lam = lift_thm_lam @{context} @{thm rfv_lam} *}
-lemma "fv (Lam (a\<Colon>name) (x\<Colon>lam)) = fv x - {a}"
-(*apply (tactic {* lift_tac_lam @{context} @{thm rfv_lam} 1 *})*)
-sorry
+
+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 *})
+done
-ML {* val a1 = lift_thm_lam @{context} @{thm a1} *}
-lemma "(a\<Colon>name) = (b\<Colon>name) \<Longrightarrow> Var a = Var b"
+lemma fv_lam: "fv (Lam (a\<Colon>name) (x\<Colon>lam)) = fv x - {a}"
+apply (tactic {* lift_tac_lam @{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 *})
done
-ML {* val a2 = lift_thm_lam @{context} @{thm a2} *}
-lemma "\<lbrakk>(x\<Colon>lam) = (xa\<Colon>lam); (xb\<Colon>lam) = (xc\<Colon>lam)\<rbrakk> \<Longrightarrow> App x xb = App xa xc"
+
+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 *})
done
-ML {* val a3 = lift_thm_lam @{context} @{thm a3} *}
-lemma "\<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 *})*)
-sorry
+
+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 *})
+done
ML {* val alpha_cases = lift_thm_lam @{context} @{thm alpha.cases} *}
lemma "\<lbrakk>(x\<Colon>lam) = (xa\<Colon>lam); \<And>(a\<Colon>name) b\<Colon>name. \<lbrakk>x = Var a; xa = Var b; a = b\<rbrakk> \<Longrightarrow> P\<Colon>bool;
\<And>(x\<Colon>lam) (xa\<Colon>lam) (xb\<Colon>lam) xc\<Colon>lam. \<lbrakk>x = App x xb; xa = App xa xc; x = xa; xb = xc\<rbrakk> \<Longrightarrow> P;
\<And>(x\<Colon>lam) (a\<Colon>name) (b\<Colon>name) xa\<Colon>lam. \<lbrakk>x = Lam a x; xa = 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 {* procedure_tac @{thm alpha.cases} @{context} 1 *})
sorry
ML {* val alpha_induct = lift_thm_lam @{context} @{thm alpha.induct} *}
lemma "\<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);
@@ -249,27 +234,10 @@
(* apply (tactic {* lift_tac_lam @{context} @{thm alpha.induct} 1 *}) *)
sorry
-lemma "(Var a = Var b) = (a = b)"
+lemma var_inject: "(Var a = Var b) = (a = b)"
apply (tactic {* lift_tac_lam @{context} @{thm rvar_inject} 1 *})
done
-local_setup {*
- Quotient.note (@{binding "pi_var"}, pi_var) #> snd #>
- Quotient.note (@{binding "pi_app"}, pi_app) #> snd #>
- Quotient.note (@{binding "pi_lam"}, pi_lam) #> snd #>
- Quotient.note (@{binding "a1"}, a1) #> snd #>
- Quotient.note (@{binding "a2"}, a2) #> snd #>
- Quotient.note (@{binding "a3"}, a3) #> snd #>
- Quotient.note (@{binding "alpha_cases"}, alpha_cases) #> snd #>
- Quotient.note (@{binding "alpha_induct"}, alpha_induct) #> snd #>
- Quotient.note (@{binding "var_inject"}, var_inject) #> snd
-*}
-
-thm alpha.cases
-thm alpha_cases
-thm alpha.induct
-thm alpha_induct
-
lemma var_supp:
shows "supp (Var a) = ((supp a)::name set)"
apply(simp add: supp_def)
@@ -354,24 +322,6 @@
ML_prf {*
fun r_mk_comb_tac_lam lthy = r_mk_comb_tac lthy rty quot rel_refl trans2 rsp_thms
*}
-apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
-apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
-apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
-apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
-apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
-apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
-apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
-apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
-apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
-apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
-apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
-apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
-apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
-apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
-apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
-apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
-apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
-apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})
apply (tactic {* (r_mk_comb_tac_lam @{context}) 1 *})