src/HOL/Nominal/Examples/Weakening.thy
changeset 21107 e69c0e82955a
parent 21052 ec5531061ed6
child 21364 3baf57a27269
--- 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