new file for defining functions in the lambda-calculus
authorurbanc
Mon, 30 Oct 2006 13:07:51 +0100
changeset 21107 e69c0e82955a
parent 21106 51599a81b308
child 21108 04d8e6eb9a2e
new file for defining functions in the lambda-calculus
src/HOL/Nominal/Examples/Lam_Funs.thy
src/HOL/Nominal/Examples/SN.thy
src/HOL/Nominal/Examples/Weakening.thy
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Nominal/Examples/Lam_Funs.thy	Mon Oct 30 13:07:51 2006 +0100
@@ -0,0 +1,244 @@
+(* $Id: *)
+
+theory Lam_Funs
+imports Nominal
+begin
+
+atom_decl name
+
+nominal_datatype lam = Var "name"
+                     | App "lam" "lam"
+                     | Lam "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [100,100] 100)
+
+
+constdefs 
+  depth_Var :: "name \<Rightarrow> nat"
+  "depth_Var \<equiv> \<lambda>(a::name). 1"
+  
+  depth_App :: "lam \<Rightarrow> lam \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
+  "depth_App \<equiv> \<lambda>_ _ n1 n2. (max n1 n2)+1"
+
+  depth_Lam :: "name \<Rightarrow> lam \<Rightarrow> nat \<Rightarrow> nat"
+  "depth_Lam \<equiv> \<lambda>_ _ n. n+1"
+
+  depth_lam :: "lam \<Rightarrow> nat"
+  "depth_lam t \<equiv> (lam_rec depth_Var depth_App depth_Lam) t"
+
+lemma fin_supp_depth_lam:
+  shows "finite ((supp depth_Var)::name set)"
+  and   "finite ((supp depth_Lam)::name set)"
+  and   "finite ((supp depth_App)::name set)"
+  by (finite_guess add: depth_Var_def depth_Lam_def depth_App_def perm_nat_def)+
+
+lemma fcb_depth_lam:
+  fixes a::"name"
+  shows "a\<sharp>(depth_Lam a t n)"
+  by (simp add: fresh_nat)
+
+lemma depth_lam:
+  shows "depth_lam (Var a)     = 1"
+  and   "depth_lam (App t1 t2) = (max (depth_lam t1) (depth_lam t2))+1"
+  and   "depth_lam (Lam [a].t) = (depth_lam t)+1"
+apply -
+apply(unfold depth_lam_def)
+apply(rule trans)
+apply(rule lam.recs[where P="\<lambda>_. True" and z="()", simplified])
+apply(simp_all add: fin_supp_depth_lam supp_unit)
+apply(simp add: fcb_depth_lam)
+apply(simp add: depth_Var_def)
+apply(rule trans)
+apply(rule lam.recs[where P="\<lambda>_. True" and z="()", simplified])
+apply(simp_all add: fin_supp_depth_lam supp_unit)
+apply(simp add: fcb_depth_lam)
+apply(simp add: depth_App_def)
+apply(rule trans)
+apply(rule lam.recs[where P="\<lambda>_. True" and z="()", simplified])
+apply(simp_all add: fin_supp_depth_lam supp_unit)
+apply(simp add: fcb_depth_lam)
+apply(simp add: depth_Var_def, fresh_guess add: perm_nat_def)
+apply(simp add: depth_App_def, fresh_guess add: perm_nat_def)
+apply(simp add: depth_Lam_def, fresh_guess add: perm_nat_def)
+apply(simp add: depth_Lam_def)
+done
+
+text {* capture-avoiding substitution *}
+constdefs 
+  subst_Var :: "name \<Rightarrow>lam \<Rightarrow> name \<Rightarrow> lam"
+  "subst_Var b t \<equiv> \<lambda>a. (if a=b then t else (Var a))"
+  
+  subst_App :: "name \<Rightarrow> lam \<Rightarrow> lam \<Rightarrow> lam \<Rightarrow> lam \<Rightarrow> lam \<Rightarrow> lam"
+  "subst_App b t \<equiv> \<lambda>_ _ r1 r2. App r1 r2"
+
+  subst_Lam :: "name \<Rightarrow> lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam \<Rightarrow> lam"
+  "subst_Lam b t \<equiv> \<lambda>a _ r. Lam [a].r"
+
+abbreviation
+  subst_syn  :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_[_::=_]" [100,100,100] 900) 
+  "t'[b::=t] \<equiv> (lam_rec (subst_Var b t) (subst_App b t) (subst_Lam b t)) t'"
+
+(* FIXME: this lemma needs to be in Nominal.thy *)
+lemma eq_eqvt:
+  fixes pi::"name prm"
+  and   x::"'a::pt_name"
+  shows "pi\<bullet>(x=y) = ((pi\<bullet>x)=(pi\<bullet>y))"
+  apply(simp add: perm_bool perm_bij)
+  done
+
+lemma fin_supp_subst:
+  shows "finite ((supp (subst_Var b t))::name set)"
+  and   "finite ((supp (subst_Lam b t))::name set)"
+  and   "finite ((supp (subst_App b t))::name set)"
+apply -
+apply(finite_guess add: subst_Var_def perm_if eq_eqvt fs_name1)
+apply(finite_guess add: subst_Lam_def fs_name1)
+apply(finite_guess add: subst_App_def fs_name1)
+done
+
+lemma fcb_subst:
+  fixes a::"name"
+  shows "a\<sharp>(subst_Lam b t) a t' r"
+  by (simp add: subst_Lam_def abs_fresh)
+
+lemma subst[simp]:
+  shows "(Var a)[b::=t] = (if a=b then t else (Var a))"
+  and   "(App t1 t2)[b::=t] = App (t1[b::=t]) (t2[b::=t])"
+  and   "a\<sharp>(b,t)    \<Longrightarrow> (Lam [a].t1)[b::=t] = Lam [a].(t1[b::=t])"
+  and   "\<lbrakk>a\<noteq>b; a\<sharp>t\<rbrakk> \<Longrightarrow> (Lam [a].t1)[b::=t] = Lam [a].(t1[b::=t])"
+  and   "\<lbrakk>a\<sharp>b; a\<sharp>t\<rbrakk> \<Longrightarrow> (Lam [a].t1)[b::=t] = Lam [a].(t1[b::=t])"
+apply -
+apply(rule trans)
+apply(rule lam.recs[where P="\<lambda>_. True" and z="()", simplified])
+apply(simp add: fs_name1 fin_supp_subst)+
+apply(simp add: fcb_subst)
+apply(simp add: subst_Var_def)
+apply(rule trans)
+apply(rule lam.recs[where P="\<lambda>_. True" and z="()", simplified])
+apply(simp add: fs_name1 fin_supp_subst)+
+apply(simp add: fcb_subst)
+apply(simp add: subst_App_def)
+apply(rule trans)
+apply(rule lam.recs[where P="\<lambda>_. True" and z="()", simplified])
+apply(simp add: fs_name1 fin_supp_subst)+
+apply(simp add: fcb_subst)
+apply(fresh_guess add: subst_Var_def perm_if eq_eqvt fs_name1)
+apply(fresh_guess add: subst_App_def fs_name1)
+apply(fresh_guess add: subst_Lam_def fs_name1)
+apply(simp add: subst_Lam_def)
+apply(rule trans)
+apply(rule lam.recs[where P="\<lambda>_. True" and z="()", simplified])
+apply(simp add: fs_name1 fin_supp_subst)+
+apply(simp add: fcb_subst)
+apply(fresh_guess add: subst_Var_def perm_if eq_eqvt fs_name1 fresh_atm)
+apply(fresh_guess add: subst_App_def fs_name1 fresh_atm)
+apply(fresh_guess add: subst_Lam_def fs_name1 fresh_atm)
+apply(simp add: subst_Lam_def)
+apply(rule trans)
+apply(rule lam.recs[where P="\<lambda>_. True" and z="()", simplified])
+apply(simp add: fs_name1 fin_supp_subst)+
+apply(simp add: fcb_subst)
+apply(fresh_guess add: subst_Var_def perm_if eq_eqvt fs_name1)
+apply(fresh_guess add: subst_App_def fs_name1)
+apply(fresh_guess add: subst_Lam_def fs_name1)
+apply(simp add: subst_Lam_def)
+done
+
+lemma subst_eqvt[simp]:
+  fixes pi:: "name prm"
+  and   t1:: "lam"
+  and   t2:: "lam"
+  and   a :: "name"
+  shows "pi\<bullet>(t1[b::=t2]) = (pi\<bullet>t1)[(pi\<bullet>b)::=(pi\<bullet>t2)]"
+apply(nominal_induct t1 avoiding: b t2 rule: lam.induct)
+apply(auto simp add: perm_bij fresh_prod fresh_atm fresh_bij)
+done
+
+lemma subst_supp: 
+  shows "supp(t1[a::=t2]) \<subseteq> (((supp(t1)-{a})\<union>supp(t2))::name set)"
+apply(nominal_induct t1 avoiding: a t2 rule: lam.induct)
+apply(auto simp add: lam.supp supp_atm fresh_prod abs_supp)
+apply(blast)+
+done
+
+(* parallel substitution *)
+
+(* simultaneous substitution *)
+consts
+  "domain" :: "(name\<times>lam) list \<Rightarrow> (name list)"
+primrec
+  "domain []    = []"
+  "domain (x#\<theta>) = (fst x)#(domain \<theta>)" 
+
+consts
+  "apply_sss"  :: "(name\<times>lam) list \<Rightarrow> name \<Rightarrow> lam" (" _ < _ >" [80,80] 80)
+primrec  
+"(x#\<theta>)<a>  = (if (a = fst x) then (snd x) else \<theta><a>)" 
+
+lemma apply_sss_eqvt:
+  fixes pi::"name prm"
+  assumes a: "a\<in>set (domain \<theta>)"
+  shows "pi\<bullet>(\<theta><a>) = (pi\<bullet>\<theta>)<pi\<bullet>a>"
+using a
+by (induct \<theta>)
+   (auto simp add: perm_bij)
+
+lemma domain_eqvt:
+  fixes pi::"name prm"
+  shows "((pi\<bullet>a)\<in>set (domain \<theta>)) = (a\<in>set (domain ((rev pi)\<bullet>\<theta>)))"
+apply(induct \<theta>)
+apply(auto)
+apply(perm_simp)+
+done
+
+constdefs 
+  psubst_Var :: "(name\<times>lam) list \<Rightarrow> name \<Rightarrow> lam"
+  "psubst_Var \<theta> \<equiv> \<lambda>a. (if a\<in>set (domain \<theta>) then \<theta><a> else (Var a))"
+  
+  psubst_App :: "(name\<times>lam) list \<Rightarrow> lam \<Rightarrow> lam \<Rightarrow> lam \<Rightarrow> lam \<Rightarrow> lam"
+  "psubst_App \<theta> \<equiv> \<lambda>_ _ r1 r2. App r1 r2"
+
+  psubst_Lam :: "(name\<times>lam) list \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam \<Rightarrow> lam"
+  "psubst_Lam \<theta> \<equiv> \<lambda>a _ r. Lam [a].r"
+
+abbreviation
+  psubst_lam :: "lam \<Rightarrow> (name\<times>lam) list \<Rightarrow> lam" ("_[<_>]" [100,100] 900)
+  "t[<\<theta>>] \<equiv> (lam_rec (psubst_Var \<theta>) (psubst_App \<theta>) (psubst_Lam \<theta>)) t"
+
+lemma fin_supp_psubst:
+  shows "finite ((supp (psubst_Var \<theta>))::name set)"
+  and   "finite ((supp (psubst_Lam \<theta>))::name set)"
+  and   "finite ((supp (psubst_App \<theta>))::name set)"
+  apply -
+  apply(finite_guess add: fs_name1 psubst_Var_def domain_eqvt apply_sss_eqvt)
+  apply(finite_guess add: fs_name1 psubst_Lam_def)
+  apply(finite_guess add: fs_name1 psubst_App_def)
+done
+
+lemma fcb_psubst_Lam:
+  shows "x\<sharp>(psubst_Lam \<theta>) x t r"
+  by (simp add: psubst_Lam_def abs_fresh)
+
+lemma psubst[simp]:
+  shows "(Var a)[<\<theta>>] = (if a\<in>set (domain \<theta>) then \<theta><a> else (Var a))"
+  and   "(App t1 t2)[<\<theta>>] = App (t1[<\<theta>>]) (t2[<\<theta>>])"
+  and   "a\<sharp>\<theta> \<Longrightarrow> (Lam [a].t1)[<\<theta>>] = Lam [a].(t1[<\<theta>>])"
+  apply(rule trans)
+  apply(rule lam.recs[where P="\<lambda>_. True" and z="()", simplified])
+  apply(simp add: fin_supp_psubst fs_name1)+
+  apply(simp add: fcb_psubst_Lam)
+  apply(simp add: psubst_Var_def)
+  apply(rule trans)
+  apply(rule lam.recs[where P="\<lambda>_. True" and z="()", simplified])
+  apply(simp add: fin_supp_psubst fs_name1)+
+  apply(simp add: fcb_psubst_Lam)
+  apply(simp add: psubst_App_def)
+  apply(rule trans)
+  apply(rule lam.recs[where P="\<lambda>_. True" and z="()", simplified])
+  apply(simp add: fin_supp_psubst fs_name1)+
+  apply(simp add: fcb_psubst_Lam)
+  apply(fresh_guess add: psubst_Var_def domain_eqvt apply_sss_eqvt fs_name1)
+  apply(fresh_guess add: psubst_App_def fs_name1)
+  apply(fresh_guess add: psubst_Lam_def fs_name1)
+  apply(simp add: psubst_Lam_def)
+done
+
+end
\ No newline at end of file
--- a/src/HOL/Nominal/Examples/SN.thy	Thu Oct 26 16:08:40 2006 +0200
+++ b/src/HOL/Nominal/Examples/SN.thy	Mon Oct 30 13:07:51 2006 +0100
@@ -1,7 +1,7 @@
 (* $Id$ *)
 
 theory SN
-imports Lam_substs
+imports Lam_Funs
 begin
 
 text {* Strong Normalisation proof from the Proofs and Types book *}
@@ -343,7 +343,7 @@
   thus "P x \<Gamma> t \<tau>" by simp
 qed
 
-constdefs
+abbreviation
   "sub" :: "(name\<times>ty) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" (" _ \<lless> _ " [80,80] 80)
   "\<Gamma>1 \<lless> \<Gamma>2 \<equiv> \<forall>a \<sigma>. (a,\<sigma>)\<in>set \<Gamma>1 \<longrightarrow>  (a,\<sigma>)\<in>set \<Gamma>2"
 
@@ -354,11 +354,9 @@
   shows "\<Gamma>2 \<turnstile> t:\<sigma>"
 using a b c
 apply(nominal_induct \<Gamma>1 t \<sigma> avoiding: \<Gamma>2 rule: typing_induct)
-apply(auto simp add: sub_def)
+apply(auto | atomize)+
 (* FIXME: before using meta-connectives and the new induction *)
 (* method, this was completely automatic *)
-apply(atomize)
-apply(auto)
 done
 
 lemma in_ctxt: 
@@ -456,7 +454,7 @@
     by (auto dest: valid_elim simp add: fresh_list_cons) 
   from c12 have c14: "((c,\<sigma>)#(a,\<tau>1)#\<Gamma>) \<turnstile> s : \<tau>2"
   proof -
-    have c2: "((a,\<tau>1)#(c,\<sigma>)#\<Gamma>) \<lless> ((c,\<sigma>)#(a,\<tau>1)#\<Gamma>)" by (force simp add: sub_def)
+    have c2: "((a,\<tau>1)#(c,\<sigma>)#\<Gamma>) \<lless> ((c,\<sigma>)#(a,\<tau>1)#\<Gamma>)" by force
     have c3: "valid ((c,\<sigma>)#(a,\<tau>1)#\<Gamma>)"
       by (rule v2, rule v2, auto simp add: fresh_list_cons fresh_prod ca cb cc f2' fresh_ty)
     from c12 c2 c3 show ?thesis by (force intro: weakening)
@@ -464,7 +462,7 @@
   assume c8: "\<Gamma> \<turnstile> t2 : \<sigma>"
   have c81: "((a,\<tau>1)#\<Gamma>)\<turnstile>t2 :\<sigma>"
   proof -
-    have c82: "\<Gamma> \<lless> ((a,\<tau>1)#\<Gamma>)" by (force simp add: sub_def)
+    have c82: "\<Gamma> \<lless> ((a,\<tau>1)#\<Gamma>)" by force
     have c83: "valid ((a,\<tau>1)#\<Gamma>)" using f1 ca by force
     with c8 c82 c83 show ?thesis by (force intro: weakening)
   qed
@@ -857,6 +855,7 @@
 (*A*)
 apply(simp add: fresh_list_cons fresh_prod)
 done
+
 lemma all_RED: 
   assumes a: "\<Gamma>\<turnstile>t:\<tau>"
   and     b: "\<forall>a \<sigma>. (a,\<sigma>)\<in>set(\<Gamma>) \<longrightarrow> (a\<in>set(domain \<theta>)\<and>\<theta><a>\<in>RED \<sigma>)" 
--- a/src/HOL/Nominal/Examples/Weakening.thy	Thu Oct 26 16:08:40 2006 +0200
+++ b/src/HOL/Nominal/Examples/Weakening.thy	Mon Oct 30 13:07:51 2006 +0100
@@ -124,6 +124,72 @@
   thus "P x \<Gamma> t \<tau>" by simp
 qed
 
+lemma typing_induct_test[consumes 1, case_names t_Var t_App t_Lam]:
+  fixes  P :: "'a::fs_name\<Rightarrow>(name\<times>ty) list \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow>bool"
+  and    \<Gamma> :: "(name\<times>ty) list"
+  and    t :: "lam"
+  and    \<tau> :: "ty"
+  and    x :: "'a::fs_name"
+  assumes a: "\<Gamma> \<turnstile> t : \<tau>"
+  and a1:    "\<And>\<Gamma> a \<tau> x. \<lbrakk>valid \<Gamma>; (a,\<tau>) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> P x \<Gamma> (Var a) \<tau>"
+  and a2:    "\<And>\<Gamma> \<tau> \<sigma> t1 t2 x. 
+              \<lbrakk>\<Gamma> \<turnstile> t1 : \<tau>\<rightarrow>\<sigma>; \<And>z. P z \<Gamma> t1 (\<tau>\<rightarrow>\<sigma>); \<Gamma> \<turnstile> t2 : \<tau>; \<And>z. P z \<Gamma> t2 \<tau>\<rbrakk>
+              \<Longrightarrow> P x \<Gamma> (App t1 t2) \<sigma>"
+  and a3:    "\<And>a \<Gamma> \<tau> \<sigma> t x. \<lbrakk>a\<sharp>x; a\<sharp>\<Gamma>; ((a,\<tau>)#\<Gamma>) \<turnstile> t : \<sigma>; \<And>z. P z ((a,\<tau>)#\<Gamma>) t \<sigma>\<rbrakk>
+              \<Longrightarrow> P x \<Gamma> (Lam [a].t) (\<tau>\<rightarrow>\<sigma>)"
+  shows "P x \<Gamma> t \<tau>"
+proof -
+  from a have "\<And>(pi::name prm) x. P x (pi\<bullet>\<Gamma>) (pi\<bullet>t) \<tau>"
+  proof (induct)
+    case (t_Var \<Gamma> a \<tau>)
+    have "valid \<Gamma>" by fact
+    then have "valid (pi\<bullet>\<Gamma>)" by (rule eqvt_valid)
+    moreover
+    have "(a,\<tau>)\<in>set \<Gamma>" by fact
+    then have "pi\<bullet>(a,\<tau>)\<in>pi\<bullet>(set \<Gamma>)" by (simp only: pt_set_bij[OF pt_name_inst, OF at_name_inst])  
+    then have "(pi\<bullet>a,\<tau>)\<in>set (pi\<bullet>\<Gamma>)" by (simp add: pt_list_set_pi[OF pt_name_inst])
+    ultimately show "P x (pi\<bullet>\<Gamma>) (pi\<bullet>(Var a)) \<tau>" using a1 by simp
+  next
+    case (t_App \<Gamma> t1 \<tau> \<sigma> t2)
+    thus "P x (pi\<bullet>\<Gamma>) (pi\<bullet>(App t1 t2)) \<sigma>" using a2 by (simp, blast intro: eqvt_typing)
+  next
+    case (t_Lam a \<Gamma> \<tau> t \<sigma> pi x)
+    have p1: "((a,\<tau>)#\<Gamma>)\<turnstile>t:\<sigma>" by fact
+    have ih1: "\<And>(pi::name prm) x. P x (pi\<bullet>((a,\<tau>)#\<Gamma>)) (pi\<bullet>t) \<sigma>" by fact
+    have f: "a\<sharp>\<Gamma>" by fact
+    then have f': "(pi\<bullet>a)\<sharp>(pi\<bullet>\<Gamma>)"  by (simp add: fresh_bij)
+    have "\<exists>c::name. c\<sharp>(pi\<bullet>a,pi\<bullet>t,pi\<bullet>\<Gamma>,x)"
+      by (rule exists_fresh, simp add: fs_name1)
+    then obtain c::"name" 
+      where fs: "c\<noteq>(pi\<bullet>a)" "c\<sharp>x" "c\<sharp>(pi\<bullet>t)" "c\<sharp>(pi\<bullet>\<Gamma>)"
+      by (force simp add: fresh_prod fresh_atm)    
+    let ?pi'="[(pi\<bullet>a,c)]@pi"
+    have eq: "((pi\<bullet>a,c)#pi)\<bullet>a = c" by (simp add: calc_atm)
+    have p1': "(?pi'\<bullet>((a,\<tau>)#\<Gamma>))\<turnstile>(?pi'\<bullet>t):\<sigma>" using p1 by (simp only: eqvt_typing)
+    have ih1': "\<And>x. P x (?pi'\<bullet>((a,\<tau>)#\<Gamma>)) (?pi'\<bullet>t) \<sigma>" using ih1 by simp
+    have "P x (?pi'\<bullet>\<Gamma>) (?pi'\<bullet>(Lam [a].t)) (\<tau>\<rightarrow>\<sigma>)" using f f' fs p1' ih1' eq
+      apply -
+      apply(simp del: append_Cons)
+      apply(rule a3)
+      apply(simp_all add: fresh_left calc_atm pt_name2)
+      done
+    then have "P x ([(pi\<bullet>a,c)]\<bullet>(pi\<bullet>\<Gamma>)) ([(pi\<bullet>a,c)]\<bullet>(Lam [(pi\<bullet>a)].(pi\<bullet>t))) (\<tau>\<rightarrow>\<sigma>)" 
+      by (simp del: append_Cons add: pt_name2)
+    then show "P x (pi\<bullet>\<Gamma>) (pi\<bullet>(Lam [a].t)) (\<tau> \<rightarrow> \<sigma>)" using f f' fs      
+      apply -
+      apply(subgoal_tac "c\<sharp>Lam [(pi\<bullet>a)].(pi\<bullet>t)")
+      apply(subgoal_tac "(pi\<bullet>a)\<sharp>Lam [(pi\<bullet>a)].(pi\<bullet>t)")
+      apply(simp only: perm_fresh_fresh)
+      apply(simp)
+      apply(simp add: abs_fresh)
+      apply(simp add: abs_fresh)
+      done 
+  qed
+  hence "P x (([]::name prm)\<bullet>\<Gamma>) (([]::name prm)\<bullet>t) \<tau>" by blast
+  thus "P x \<Gamma> t \<tau>" by simp
+qed
+
+
 text {* definition of a subcontext *}
 
 abbreviation