# HG changeset patch # User urbanc # Date 1174328907 -3600 # Node ID bfd9c0fd70b1e88900b1a606a3f88a97a4b8f40f # Parent 7c51f1a799f35aca62270dfee30d74e0d2cb75fd tuned the proof diff -r 7c51f1a799f3 -r bfd9c0fd70b1 src/HOL/Nominal/Examples/SOS.thy --- a/src/HOL/Nominal/Examples/SOS.thy Mon Mar 19 15:58:02 2007 +0100 +++ b/src/HOL/Nominal/Examples/SOS.thy Mon Mar 19 19:28:27 2007 +0100 @@ -1,10 +1,12 @@ (* "$Id$" *) (* *) -(* Formalisation of some typical SOS-proofs from a *) -(* challenge suggested by Adam Chlipala. *) +(* Formalisation of some typical SOS-proofs *) (* *) -(* We thank Nick Benton for hellping us with the *) -(* termination-proof for evaluation . *) +(* This work arose from challenge suggested by Adam *) +(* Chlipala suggested on the POPLmark mailing list. *) +(* *) +(* We thank Nick Benton for helping us with the *) +(* termination-proof for evaluation. *) theory SOS imports "Nominal" @@ -87,6 +89,8 @@ by (induct rule: lookup.induct) (auto simp add: fresh_list_cons fresh_prod fresh_atm) +text {* Parallel Substitution *} + consts psubst :: "(name\trm) list \ trm \ trm" ("_<_>" [95,95] 105) @@ -101,20 +105,19 @@ "\<(InL e)> = InL (\)" "\<(InR e)> = InR (\)" "\y\x; x\(e,e\<^isub>2,\); y\(e,e\<^isub>1,\)\ - \ \<(Case e of inl x \ e\<^isub>1 | inr y \ e\<^isub>2)> = - (Case (\) of inl x \ (\1>) | inr y \ (\2>))" - apply(finite_guess add: lookup_eqvt)+ - apply(rule TrueI)+ - apply(simp add: abs_fresh)+ - apply(fresh_guess add: fs_name1 lookup_eqvt)+ - done + \ \ e\<^isub>1 | inr y \ e\<^isub>2> = (Case (\) of inl x \ (\1>) | inr y \ (\2>))" +apply(finite_guess add: lookup_eqvt)+ +apply(rule TrueI)+ +apply(simp add: abs_fresh)+ +apply(fresh_guess add: fs_name1 lookup_eqvt)+ +done lemma psubst_eqvt[eqvt]: fixes pi::"name prm" and t::"trm" shows "pi\(\) = (pi\\)<(pi\t)>" - by (nominal_induct t avoiding: \ rule: trm.induct) - (perm_simp add: fresh_bij lookup_eqvt)+ +by (nominal_induct t avoiding: \ rule: trm.induct) + (perm_simp add: fresh_bij lookup_eqvt)+ lemma fresh_psubst: fixes z::"name" @@ -142,7 +145,7 @@ and "\z\x; x\(y,e,e\<^isub>2,t'); z\(y,e,e\<^isub>1,t')\ \ (Case e of inl x \ e\<^isub>1 | inr z \ e\<^isub>2)[y::=t'] = (Case (e[y::=t']) of inl x \ (e\<^isub>1[y::=t']) | inr z \ (e\<^isub>2[y::=t']))" - by (simp_all add: fresh_list_cons fresh_list_nil) +by (simp_all add: fresh_list_cons fresh_list_nil) lemma subst_eqvt[eqvt]: fixes pi::"name prm" @@ -151,60 +154,6 @@ by (nominal_induct t avoiding: x t' rule: trm.induct) (perm_simp add: fresh_bij)+ - -lemma subst_rename: - fixes c::"name" - and t\<^isub>1::"trm" - assumes "c\t\<^isub>1" - shows "t\<^isub>1[a::=t\<^isub>2] = ([(c,a)]\t\<^isub>1)[c::=t\<^isub>2]" - using assms - apply(nominal_induct t\<^isub>1 avoiding: a c t\<^isub>2 rule: trm.induct) - apply(simp add: trm.inject calc_atm fresh_atm abs_fresh perm_nat_def) - apply(simp add: trm.inject calc_atm fresh_atm abs_fresh perm_nat_def) - apply(simp add: trm.inject calc_atm fresh_atm abs_fresh perm_nat_def) - apply(simp add: trm.inject calc_atm fresh_atm abs_fresh perm_nat_def) - apply(simp add: trm.inject calc_atm fresh_atm abs_fresh perm_nat_def) - apply(simp add: trm.inject calc_atm fresh_atm abs_fresh perm_nat_def) - apply(simp add: trm.inject calc_atm fresh_atm abs_fresh perm_nat_def) - apply(simp add: trm.inject calc_atm fresh_atm abs_fresh perm_nat_def) - apply(simp add: trm.inject calc_atm fresh_atm abs_fresh perm_nat_def) - apply(simp (no_asm_use)) - apply(rule sym) - apply(rule trans) - apply(rule subst) - apply(simp add: perm_bij) - apply(simp add: fresh_prod) - apply(simp add: fresh_bij) - apply(simp add: calc_atm fresh_atm) - apply(simp add: fresh_prod) - apply(simp add: fresh_bij) - apply(simp add: calc_atm fresh_atm) - apply(rule sym) - apply(rule trans) - apply(rule subst) - apply(simp add: fresh_atm) - apply(simp) - apply(simp) - apply(simp (no_asm_use) add: trm.inject) - apply(rule conjI) - apply(blast) - apply(rule conjI) - apply(rotate_tac 12) - apply(drule_tac x="a" in meta_spec) - apply(rotate_tac 14) - apply(drule_tac x="c" in meta_spec) - apply(rotate_tac 14) - apply(drule_tac x="t\<^isub>2" in meta_spec) - apply(simp add: calc_atm fresh_atm alpha abs_fresh) - apply(rotate_tac 13) - apply(drule_tac x="a" in meta_spec) - apply(rotate_tac 14) - apply(drule_tac x="c" in meta_spec) - apply(rotate_tac 14) - apply(drule_tac x="t\<^isub>2" in meta_spec) - apply(simp add: calc_atm fresh_atm alpha abs_fresh) - done - lemma fresh_subst: fixes z::"name" and t\<^isub>1::"trm" @@ -234,24 +183,6 @@ by (nominal_induct L avoiding: x P rule: trm.induct) (auto simp add: fresh_atm abs_fresh) -lemma subst_fun_eq: - fixes u::trm - assumes "[x].t\<^isub>1 = [y].t\<^isub>2" - shows "t\<^isub>1[x::=u] = t\<^isub>2[y::=u]" -proof - - { - assume "x=y" and "t\<^isub>1=t\<^isub>2" - then have ?thesis using assms by simp - } - moreover - { - assume h1:"x \ y" and h2:"t\<^isub>1=[(x,y)]\t\<^isub>2" and h3:"x \ t\<^isub>2" - then have "([(x,y)]\t\<^isub>2)[x::=u] = t\<^isub>2[y::=u]" by (simp add: subst_rename) - then have ?thesis using h2 by simp - } - ultimately show ?thesis using alpha assms by blast -qed - lemma psubst_empty[simp]: shows "[] = t" by (nominal_induct t rule: trm.induct, auto simp add:fresh_list_nil) @@ -271,7 +202,7 @@ by (nominal_induct t avoiding: a e rule: trm.induct) (auto simp add: fresh_atm abs_fresh fresh_nat) -text {* Typing *} +text {* Typing-Judgements *} inductive2 valid :: "(name \ 'a::pt_name) list \ bool" @@ -335,7 +266,7 @@ (x\<^isub>1,Data(S\<^isub>1))#\ \ e\<^isub>1 : T; (x\<^isub>2,Data(S\<^isub>2))#\ \ e\<^isub>2 : T\ \ \ \ (Case e of inl x\<^isub>1 \ e\<^isub>1 | inr x\<^isub>2 \ e\<^isub>2) : T" -lemma typing_valid: +lemma typing_implies_valid: assumes "\ \ t : T" shows "valid \" using assms @@ -535,7 +466,7 @@ have "[(x\<^isub>1',c)]\((x\<^isub>1',Data \\<^isub>1)# \) \ [(x\<^isub>1',c)]\e\<^isub>1' : T" using h1 typing_eqvt by blast then have x:"(c,Data \\<^isub>1)#( [(x\<^isub>1',c)]\\) \ [(x\<^isub>1',c)]\e\<^isub>1': T" using f' by (auto simp add: fresh_atm calc_atm) - have "x\<^isub>1' \ \" using h1 typing_valid by auto + have "x\<^isub>1' \ \" using h1 typing_implies_valid by auto then have "(c,Data \\<^isub>1)#\ \ [(x\<^isub>1 ,c)]\e\<^isub>1 : T" using f' x e1' by (auto simp add: perm_fresh_fresh) then have "[(x\<^isub>1,c)]\((c,Data \\<^isub>1)#\) \ [(x\<^isub>1,c)]\[(x\<^isub>1 ,c)]\e\<^isub>1 : T" using typing_eqvt by blast then have "([(x\<^isub>1,c)]\(c,Data \\<^isub>1)) #\ \ [(x\<^isub>1,c)]\[(x\<^isub>1 ,c)]\e\<^isub>1 : T" using f f' @@ -548,7 +479,7 @@ have "[(x\<^isub>2',c)]\((x\<^isub>2',Data \\<^isub>2)# \) \ [(x\<^isub>2',c)]\e\<^isub>2' : T" using h2 typing_eqvt by blast then have x:"(c,Data \\<^isub>2)#([(x\<^isub>2',c)]\\) \ [(x\<^isub>2',c)]\e\<^isub>2': T" using f' by (auto simp add: fresh_atm calc_atm) - have "x\<^isub>2' \ \" using h2 typing_valid by auto + have "x\<^isub>2' \ \" using h2 typing_implies_valid by auto then have "(c,Data \\<^isub>2)#\ \ [(x\<^isub>2 ,c)]\e\<^isub>2 : T" using f' x e2' by (auto simp add: perm_fresh_fresh) then have "[(x\<^isub>2,c)]\((c,Data \\<^isub>2)#\) \ [(x\<^isub>2,c)]\[(x\<^isub>2 ,c)]\e\<^isub>2 : T" using typing_eqvt by blast then have "([(x\<^isub>2,c)]\(c,Data \\<^isub>2))#\ \ [(x\<^isub>2,c)]\[(x\<^isub>2 ,c)]\e\<^isub>2 : T" using f f' @@ -589,7 +520,7 @@ assumes a: "(x\<^isub>1,T\<^isub>1)#(x\<^isub>2,T\<^isub>2)#\ \ e : T" shows "(x\<^isub>2,T\<^isub>2)#(x\<^isub>1,T\<^isub>1)#\ \ e : T" proof - - from a have "valid ((x\<^isub>1,T\<^isub>1)#(x\<^isub>2,T\<^isub>2)#\)" by (simp add: typing_valid) + from a have "valid ((x\<^isub>1,T\<^isub>1)#(x\<^isub>2,T\<^isub>2)#\)" by (simp add: typing_implies_valid) then have "x\<^isub>1\x\<^isub>2" "x\<^isub>1\\" "x\<^isub>2\\" "valid \" by (auto simp: fresh_list_cons fresh_atm[symmetric]) then have "valid ((x\<^isub>2,T\<^isub>2)#(x\<^isub>1,T\<^isub>1)#\)" @@ -607,7 +538,6 @@ thus "t\<^isub>1=t\<^isub>2" by (simp only: type_unicity_in_context) qed - lemma typing_substitution: fixes \::"(name \ ty) list" assumes "(x,T')#\ \ e : T" @@ -630,7 +560,7 @@ have "(y,T) \ set ((x,T')#\)" using h1 by auto then have "(y,T) \ set \" using as by auto moreover - have "valid \" using h2 by (simp only: typing_valid) + have "valid \" using h2 by (simp only: typing_implies_valid) ultimately show "\ \ (Var y)[x::=e'] : T" using as by (simp add: t_Var) qed next @@ -642,7 +572,7 @@ using vc by (auto simp add: fresh_list_cons) then have pr2'':"(x,T')#(y,T\<^isub>1)#\ \ t : T\<^isub>2" by (simp add: context_exchange) have ih: "\(x,T')#(y,T\<^isub>1)#\ \ t : T\<^isub>2; (y,T\<^isub>1)#\ \ e' : T'\ \ (y,T\<^isub>1)#\ \ t[x::=e'] : T\<^isub>2" by fact - have "valid \" using pr1 by (simp add: typing_valid) + have "valid \" using pr1 by (simp add: typing_implies_valid) then have "valid ((y,T\<^isub>1)#\)" using vc by auto then have "(y,T\<^isub>1)#\ \ e' : T'" using pr1 by (auto intro: weakening) then have "(y,T\<^isub>1)#\ \ t[x::=e'] : T\<^isub>2" using ih pr2'' by simp @@ -667,11 +597,11 @@ from h3 have h3': "(x,T')#(x\<^isub>2,Data S\<^isub>2)#\ \ t3 : T" by (rule context_exchange) have "\ \ t\<^isub>1[x::=e'] : Data (DSum S\<^isub>1 S\<^isub>2)" using h1 ih1 as1 by simp moreover - have "valid ((x\<^isub>1,Data S\<^isub>1)#\)" using h2' by (auto dest: typing_valid) + have "valid ((x\<^isub>1,Data S\<^isub>1)#\)" using h2' by (auto dest: typing_implies_valid) then have "(x\<^isub>1,Data S\<^isub>1)#\ \ e' : T'" using as1 by (auto simp add: weakening) then have "(x\<^isub>1,Data S\<^isub>1)#\ \ t\<^isub>2[x::=e'] : T" using ih2 h2' by simp moreover - have "valid ((x\<^isub>2,Data S\<^isub>2)#\)" using h3' by (auto dest: typing_valid) + have "valid ((x\<^isub>2,Data S\<^isub>2)#\)" using h3' by (auto dest: typing_implies_valid) then have "(x\<^isub>2,Data S\<^isub>2)#\ \ e' : T'" using as1 by (auto simp add: weakening) then have "(x\<^isub>2,Data S\<^isub>2)#\ \ t3[x::=e'] : T" using ih3 h3' by simp ultimately have "\ \ Case (t\<^isub>1[x::=e']) of inl x\<^isub>1 \ (t\<^isub>2[x::=e']) | inr x\<^isub>2 \ (t3[x::=e']) : T" @@ -699,6 +629,27 @@ nominal_inductive big +lemma big_eqvt[eqvt]: + fixes pi::"name prm" + assumes a: "t \ t'" + shows "(pi\t) \ (pi\t')" + using a + apply(induct) + apply(auto simp add: eqvt) + apply(rule_tac x="pi\x" in b_App) + apply(auto simp add: eqvt fresh_bij fresh_prod) + done + +lemma big_eqvt': + fixes pi::"name prm" + assumes a: "(pi\t) \ (pi\t')" + shows "t \ t'" +using a +apply - +apply(drule_tac pi="rev pi" in big_eqvt) +apply(perm_simp) +done + lemma fresh_preserved: fixes x::name fixes t::trm @@ -732,24 +683,20 @@ assumes a: "t \ t'" and a1: "\x e c. P c (Lam [x].e) (Lam [x].e)" and a2: "\x e\<^isub>1 e\<^isub>2 e\<^isub>2' e' e\<^isub>1' c. - \x\(e\<^isub>1,e\<^isub>2,e',c); e\<^isub>1\Lam [x].e\<^isub>1'; (\c. P c e\<^isub>1 (Lam [x].e\<^isub>1')); - e\<^isub>2\e\<^isub>2'; (\c. P c e\<^isub>2 e\<^isub>2'); e\<^isub>1'[x::=e\<^isub>2']\e'; (\c. P c (e\<^isub>1'[x::=e\<^isub>2']) e')\ + \x\(e\<^isub>1,e\<^isub>2,e',c); \c. P c e\<^isub>1 (Lam [x].e\<^isub>1'); \c. P c e\<^isub>2 e\<^isub>2'; \c. P c (e\<^isub>1'[x::=e\<^isub>2']) e'\ \ P c (App e\<^isub>1 e\<^isub>2) e'" and a3: "\n c. P c (Const n) (Const n)" - and a4: "\e\<^isub>1 e\<^isub>1' e\<^isub>2 e\<^isub>2' c. - \e\<^isub>1 \ e\<^isub>1'; (\c. P c e\<^isub>1 e\<^isub>1'); e\<^isub>2 \ e\<^isub>2'; (\c. P c e\<^isub>2 e\<^isub>2')\ + and a4: "\e\<^isub>1 e\<^isub>1' e\<^isub>2 e\<^isub>2' c. \\c. P c e\<^isub>1 e\<^isub>1'; \c. P c e\<^isub>2 e\<^isub>2'\ \ P c (Pr e\<^isub>1 e\<^isub>2) (Pr e\<^isub>1' e\<^isub>2')" - and a5: "\e e\<^isub>1 e\<^isub>2 c. \e \ Pr e\<^isub>1 e\<^isub>2; (\c. P c e (Pr e\<^isub>1 e\<^isub>2))\ \ P c (Fst e) e\<^isub>1" - and a6: "\e e\<^isub>1 e\<^isub>2 c. \e \ Pr e\<^isub>1 e\<^isub>2; (\c. P c e (Pr e\<^isub>1 e\<^isub>2))\ \ P c (Snd e) e\<^isub>2" - and a7: "\e e' c. \e \ e'; (\c. P c e e')\ \ P c (InL e) (InL e')" - and a8: "\e e' c. \e \ e'; (\c. P c e e')\ \ P c (InR e) (InR e')" + and a5: "\e e\<^isub>1 e\<^isub>2 c. \\c. P c e (Pr e\<^isub>1 e\<^isub>2)\ \ P c (Fst e) e\<^isub>1" + and a6: "\e e\<^isub>1 e\<^isub>2 c. \\c. P c e (Pr e\<^isub>1 e\<^isub>2)\ \ P c (Snd e) e\<^isub>2" + and a7: "\e e' c. \\c. P c e e'\ \ P c (InL e) (InL e')" + and a8: "\e e' c. \\c. P c e e'\ \ P c (InR e) (InR e')" and a9: "\x\<^isub>1 e e\<^isub>2 e'' x\<^isub>2 e\<^isub>1 e' c. - \x\<^isub>1\(e,e\<^isub>2,e'',x\<^isub>2,c); x\<^isub>2\(e,e\<^isub>1,e'',x\<^isub>1,c); e \ InL e'; (\c. P c e (InL e')); - e\<^isub>1[x\<^isub>1::=e'] \ e''; (\c. P c (e\<^isub>1[x\<^isub>1::=e']) e'')\ + \x\<^isub>1\(e,e\<^isub>2,e'',x\<^isub>2,c); x\<^isub>2\(e,e\<^isub>1,e'',x\<^isub>1,c); \c. P c e (InL e'); \c. P c (e\<^isub>1[x\<^isub>1::=e']) e''\ \ P c (Case e of inl x\<^isub>1 \ e\<^isub>1 | inr x\<^isub>2 \ e\<^isub>2) e''" and a10:"\x\<^isub>1 e e\<^isub>2 e'' x\<^isub>2 e\<^isub>1 e' c. - \x\<^isub>1\(e,e\<^isub>2,e'',x\<^isub>2,c); x\<^isub>2\(e,e\<^isub>1,e'',x\<^isub>1,c); e \ InR e'; (\c. P c e (InR e')); - e\<^isub>2[x\<^isub>2::=e'] \ e''; (\c. P c (e\<^isub>2[x\<^isub>2::=e']) e'')\ + \x\<^isub>1\(e,e\<^isub>2,e'',x\<^isub>2,c); x\<^isub>2\(e,e\<^isub>1,e'',x\<^isub>1,c); \c. P c e (InR e'); \c. P c (e\<^isub>2[x\<^isub>2::=e']) e''\ \ P c (Case e of inl x\<^isub>1 \ e\<^isub>1 | inr x\<^isub>2 \ e\<^isub>2) e''" shows "P c t t'" proof - @@ -767,16 +714,6 @@ have f1: "(pi\x)\(pi\e\<^isub>1,pi\e\<^isub>2,pi\e')" using f0 by (simp add: fresh_bij) have f2: "y\(?pi'\e\<^isub>1,?pi'\e\<^isub>2,?pi'\e')" using f0 by (auto simp add: pt_name2 fresh_left calc_atm perm_pi_simp fresh_prod) - have p1: "e\<^isub>1 \ Lam [x].e\<^isub>1'" by fact - then have "(?pi'\e\<^isub>1)\(?pi'\Lam [x].e\<^isub>1')" by (simp only: big_eqvt) - moreover - have p2: "e\<^isub>2 \ e\<^isub>2'" by fact - then have "(?pi'\e\<^isub>2) \ (?pi'\e\<^isub>2')" by (simp only: big_eqvt) - moreover - have p3: "e\<^isub>1'[x::=e\<^isub>2'] \ e'" by fact - then have "(?pi'\(e\<^isub>1'[x::=e\<^isub>2'])) \ (?pi'\e')" by (simp only: big_eqvt) - then have "(?pi'\e\<^isub>1')[y::=(?pi'\e\<^isub>2')] \ (?pi'\e')" by (simp add: subst_eqvt calc_atm) - moreover have ih1: "\c. P c (?pi'\e\<^isub>1) (?pi'\(Lam [x].e\<^isub>1'))" by fact then have "\c. P c (?pi'\e\<^isub>1) (Lam [y].(?pi'\e\<^isub>1'))" by (simp add: calc_atm) moreover @@ -799,36 +736,25 @@ show "P c (pi\(Const n)) (pi\(Const n))" using a3 by simp next case (b_Pr e\<^isub>1 e\<^isub>1' e\<^isub>2 e\<^isub>2' pi c) - then show "P c (pi\(Pr e\<^isub>1 e\<^isub>2)) (pi\(Pr e\<^isub>1' e\<^isub>2'))" using a4 - by (simp, blast intro: big_eqvt) + then show "P c (pi\(Pr e\<^isub>1 e\<^isub>2)) (pi\(Pr e\<^isub>1' e\<^isub>2'))" using a4 by simp next case (b_Fst e e\<^isub>1 e\<^isub>2 pi c) - have p1: "e \ Pr e\<^isub>1 e\<^isub>2" by fact - then have "(pi\e)\(pi\(Pr e\<^isub>1 e\<^isub>2))" by (simp only: big_eqvt) - moreover - have ih1: "\c. P c (pi\e) (pi\(Pr e\<^isub>1 e\<^isub>2))" by fact - ultimately show "P c (pi\(Fst e)) (pi\e\<^isub>1)" using a5 by simp + then show "P c (pi\(Fst e)) (pi\e\<^isub>1)" using a5 by (simp, blast) next case (b_Snd e e\<^isub>1 e\<^isub>2 pi c) - have p1: "e \ Pr e\<^isub>1 e\<^isub>2" by fact - then have "(pi\e)\(pi\(Pr e\<^isub>1 e\<^isub>2))" by (simp only: big_eqvt) - moreover - have ih1: "\c. P c (pi\e) (pi\(Pr e\<^isub>1 e\<^isub>2))" by fact - ultimately show "P c (pi\(Snd e)) (pi\e\<^isub>2)" using a6 by simp + then show "P c (pi\(Snd e)) (pi\e\<^isub>2)" using a6 by (simp, blast) next case (b_InL e e' pi c) - then show "P c (pi\(InL e)) (pi\(InL e'))" using a7 - by (simp, blast intro: big_eqvt) + then show "P c (pi\(InL e)) (pi\(InL e'))" using a7 by (simp) next case (b_InR e e' pi c) - then show "P c (pi\(InR e)) (pi\(InR e'))" using a8 - by (simp, blast intro: big_eqvt) + then show "P c (pi\(InR e)) (pi\(InR e'))" using a8 by (simp) next case (b_CaseL x\<^isub>1 e e\<^isub>2 e'' x\<^isub>2 e\<^isub>1 e' pi c) obtain y\<^isub>1::"name" where fs1: "y\<^isub>1\(pi\x\<^isub>1,pi\e,pi\e\<^isub>1,pi\e\<^isub>2,pi\e'',pi\x\<^isub>2,c)" - by (rule exists_fresh[OF fs_name1]) + by (rule exists_fresh[OF fin_supp]) obtain y\<^isub>2::"name" where fs2: "y\<^isub>2\(pi\x\<^isub>2,pi\e,pi\e\<^isub>1,pi\e\<^isub>2,pi\e'',pi\x\<^isub>1,c,y\<^isub>1)" - by (rule exists_fresh[OF fs_name1]) + by (rule exists_fresh[OF fin_supp]) let ?sw1 = "[(pi\x\<^isub>1,y\<^isub>1)]" let ?sw2 = "[(pi\x\<^isub>2,y\<^isub>2)]" let ?pi' = "?sw2@?sw1@pi" @@ -840,14 +766,6 @@ have f12: "(pi\x\<^isub>2)\(pi\e,pi\e\<^isub>1,pi\e'',pi\x\<^isub>1)" using f02 by (simp add: fresh_bij) have f22: "y\<^isub>2\(?pi'\e,?pi'\e\<^isub>1,?pi'\e'')" using f02 fs1 fs2 by (auto simp add: fresh_atm fresh_prod fresh_left calc_atm pt_name2 perm_pi_simp) - have p1: "e \ InL e'" by fact - then have "(?pi'\e) \ (?pi'\(InL e'))" by (simp only: big_eqvt) - moreover - have p2: "e\<^isub>1[x\<^isub>1::=e'] \ e''" by fact - then have "(?pi'\(e\<^isub>1[x\<^isub>1::=e'])) \ (?pi'\e'')" by (simp only: big_eqvt) - then have "(?pi'\e\<^isub>1)[y\<^isub>1::=(?pi'\e')] \ (?pi'\e'')" using fs1 fs2 - by (auto simp add: calc_atm subst_eqvt fresh_prod fresh_atm del: append_Cons) - moreover have ih1: "\c. P c (?pi'\e) (?pi'\(InL e'))" by fact moreover have ih2: "\c. P c (?pi'\(e\<^isub>1[x\<^isub>1::=e'])) (?pi'\e'')" by fact @@ -874,9 +792,9 @@ next case (b_CaseR x\<^isub>1 e e\<^isub>2 e'' x\<^isub>2 e\<^isub>1 e' pi c) obtain y\<^isub>1::"name" where fs1: "y\<^isub>1\(pi\x\<^isub>1,pi\e,pi\e\<^isub>1,pi\e\<^isub>2,pi\e'',pi\x\<^isub>2,c)" - by (rule exists_fresh[OF fs_name1]) + by (rule exists_fresh[OF fin_supp]) obtain y\<^isub>2::"name" where fs2: "y\<^isub>2\(pi\x\<^isub>2,pi\e,pi\e\<^isub>1,pi\e\<^isub>2,pi\e'',pi\x\<^isub>1,c,y\<^isub>1)" - by (rule exists_fresh[OF fs_name1]) + by (rule exists_fresh[OF fin_supp]) let ?sw1 = "[(pi\x\<^isub>1,y\<^isub>1)]" let ?sw2 = "[(pi\x\<^isub>2,y\<^isub>2)]" let ?pi' = "?sw2@?sw1@pi" @@ -888,14 +806,6 @@ have f12: "(pi\x\<^isub>2)\(pi\e,pi\e\<^isub>1,pi\e'',pi\x\<^isub>1)" using f02 by (simp add: fresh_bij) have f22: "y\<^isub>2\(?pi'\e,?pi'\e\<^isub>1,?pi'\e'')" using f02 fs1 fs2 by (auto simp add: fresh_atm fresh_prod fresh_left calc_atm pt_name2 perm_pi_simp) - have p1: "e \ InR e'" by fact - then have "(?pi'\e) \ (?pi'\(InR e'))" by (simp only: big_eqvt) - moreover - have p2: "e\<^isub>2[x\<^isub>2::=e'] \ e''" by fact - then have "(?pi'\(e\<^isub>2[x\<^isub>2::=e'])) \ (?pi'\e'')" by (simp only: big_eqvt) - then have "(?pi'\e\<^isub>2)[y\<^isub>2::=(?pi'\e')] \ (?pi'\e'')" using fs1 fs2 f11 f12 - by (auto simp add: calc_atm subst_eqvt fresh_prod fresh_atm del: append_Cons) - moreover have ih1: "\c. P c (?pi'\e) (?pi'\(InR e'))" by fact moreover have ih2: "\c. P c (?pi'\(e\<^isub>2[x\<^isub>2::=e'])) (?pi'\e'')" by fact @@ -937,24 +847,62 @@ done lemma b_CaseL_elim[elim]: - assumes "Case e of inl x\<^isub>1 \ e\<^isub>1 | inr x\<^isub>2 \ e\<^isub>2 \ e''" and "(\ t. \ e \ InR t)" + assumes "Case e of inl x\<^isub>1 \ e\<^isub>1 | inr x\<^isub>2 \ e\<^isub>2 \ e''" + and "\ t. \ e \ InR t" + and "x\<^isub>1\e''" "x\<^isub>1\e" "x\<^isub>2\e''" "x\<^isub>1\e" obtains e' where "e \ InL e'" and "e\<^isub>1[x\<^isub>1::=e'] \ e''" using assms apply - - apply (rule b_Case_inv_auto, auto) - apply(drule_tac u="e'" in subst_fun_eq) - apply(simp) - done + apply(rule b_Case_inv_auto) + apply(auto) + apply(simp add: alpha) + apply(auto) + apply(drule_tac x="[(x\<^isub>1,x\<^isub>1')]\e'" in meta_spec) + apply(drule meta_mp) + apply(rule_tac pi="[(x\<^isub>1,x\<^isub>1')]" in big_eqvt') + apply(perm_simp add: fresh_prod) + apply(drule meta_mp) + apply(rule_tac pi="[(x\<^isub>1,x\<^isub>1')]" in big_eqvt') + apply(perm_simp add: eqvt calc_atm) + apply(assumption) + apply(drule_tac x="[(x\<^isub>1,x\<^isub>1')]\e'" in meta_spec) + apply(drule meta_mp) + apply(rule_tac pi="[(x\<^isub>1,x\<^isub>1')]" in big_eqvt') + apply(perm_simp add: fresh_prod) + apply(drule meta_mp) + apply(rule_tac pi="[(x\<^isub>1,x\<^isub>1')]" in big_eqvt') + apply(perm_simp add: eqvt calc_atm) + apply(assumption) +done lemma b_CaseR_elim[elim]: - assumes "Case e of inl x\<^isub>1 \ e\<^isub>1 | inr x\<^isub>2 \ e\<^isub>2 \ e''" and "\ t. \ e \ InL t" + assumes "Case e of inl x\<^isub>1 \ e\<^isub>1 | inr x\<^isub>2 \ e\<^isub>2 \ e''" + and "\ t. \ e \ InL t" + and "x\<^isub>1\e''" "x\<^isub>1\e" "x\<^isub>2\e''" "x\<^isub>2\e" obtains e' where "e \ InR e'" and "e\<^isub>2[x\<^isub>2::=e'] \ e''" - using assms + using assms apply - - apply (rule b_Case_inv_auto, auto) - apply(drule_tac u="e'" in subst_fun_eq)+ - apply(simp) - done + apply(rule b_Case_inv_auto) + apply(auto) + apply(simp add: alpha) + apply(auto) + apply(drule_tac x="[(x\<^isub>2,x\<^isub>2')]\e'" in meta_spec) + apply(drule meta_mp) + apply(rule_tac pi="[(x\<^isub>2,x\<^isub>2')]" in big_eqvt') + apply(perm_simp add: fresh_prod) + apply(drule meta_mp) + apply(rule_tac pi="[(x\<^isub>2,x\<^isub>2')]" in big_eqvt') + apply(perm_simp add: eqvt calc_atm) + apply(assumption) + apply(drule_tac x="[(x\<^isub>2,x\<^isub>2')]\e'" in meta_spec) + apply(drule meta_mp) + apply(rule_tac pi="[(x\<^isub>2,x\<^isub>2')]" in big_eqvt') + apply(perm_simp add: fresh_prod) + apply(drule meta_mp) + apply(rule_tac pi="[(x\<^isub>2,x\<^isub>2')]" in big_eqvt') + apply(perm_simp add: eqvt calc_atm) + apply(assumption) +done inductive2 val :: "trm\bool" @@ -965,7 +913,6 @@ | v_InL[intro]: "val e \ val (InL e)" | v_InR[intro]: "val e \ val (InR e)" - declare trm.inject [simp add] declare ty.inject [simp add] declare data.inject [simp add] @@ -986,17 +933,18 @@ declare data.inject [simp del] lemma subject_reduction: - assumes "e \ e'" and "\ \ e : T" + assumes a: "e \ e'" + and b: "\ \ e : T" shows "\ \ e' : T" - using assms + using a b proof (nominal_induct avoiding: \ arbitrary: T rule: big_induct_strong) case (b_App x e\<^isub>1 e\<^isub>2 e\<^isub>2' e' e \ T) have vc: "x\\" by fact have "\ \ App e\<^isub>1 e\<^isub>2 : T" by fact - then obtain T' where + then obtain T' where a1: "\ \ e\<^isub>1 : T'\T" and a2: "\ \ e\<^isub>2 : T'" by auto - have ih1: "\ \ e\<^isub>1 : T'\T \ \ \ Lam [x].e : T'\T" by fact + have ih1: "\ \ e\<^isub>1 : T' \ T \ \ \ Lam [x].e : T' \ T" by fact have ih2: "\ \ e\<^isub>2 : T' \ \ \ e\<^isub>2' : T'" by fact have ih3: "\ \ e[x::=e\<^isub>2'] : T \ \ \ e' : T" by fact have "\ \ Lam [x].e : T'\T" using ih1 a1 by simp @@ -1023,76 +971,57 @@ then show "\ \ e'' : T" by (blast intro: typing_substitution) qed (blast)+ -lemma challenge_5: - assumes "x\y" - shows "App (App (Lam [x].(Lam [y].Var y)) (Const n\<^isub>1)) (Const n\<^isub>2) \ (Const n\<^isub>2)" - using assms - by (auto intro!: big.intros simp add: forget abs_fresh fresh_atm fresh_nat) - -lemma challenge_6: - shows "Fst (App (Lam [x].Pr (Var x) (Var x)) (Const n)) \ Const n" - by (auto intro!: big.intros) (simp add: fresh_nat abs_fresh) - -lemma challenge_4_unicity: - assumes "e \ e\<^isub>1" and "e \ e\<^isub>2" +lemma unicity_of_evaluation: + assumes a: "e \ e\<^isub>1" + and b: "e \ e\<^isub>2" shows "e\<^isub>1 = e\<^isub>2" - using assms -proof (induct arbitrary: e\<^isub>2) + using a b +proof (nominal_induct e e\<^isub>1 avoiding: e\<^isub>2 rule: big_induct_strong) case (b_Lam x e t\<^isub>2) have "Lam [x].e \ t\<^isub>2" by fact thus "Lam [x].e = t\<^isub>2" by (cases, simp_all add: trm.inject) next - case (b_Fst e e\<^isub>1 e\<^isub>2 t\<^isub>2) - have "Fst e \ t\<^isub>2" by fact - then obtain e\<^isub>1' e\<^isub>2' where "e \ Pr e\<^isub>1' e\<^isub>2'" and eq: "t\<^isub>2 = e\<^isub>1'" by auto - then have "Pr e\<^isub>1 e\<^isub>2 = Pr e\<^isub>1' e\<^isub>2'" by auto - thus "e\<^isub>1 = t\<^isub>2" using eq by (simp add: trm.inject) -next - case (b_Snd e e\<^isub>1 e\<^isub>2 t\<^isub>2) - thus ?case by (force simp add: trm.inject) -next - case (b_App x e\<^isub>1 e\<^isub>2 e' e\<^isub>1' e\<^isub>2' t\<^isub>2) - have "e\<^isub>1 \ Lam [x].e\<^isub>1'" by fact + case (b_App x e\<^isub>1 e\<^isub>2 e\<^isub>2' e' e\<^isub>1' t\<^isub>2) have ih1: "\t. e\<^isub>1 \ t \ Lam [x].e\<^isub>1' = t" by fact - have "e\<^isub>2 \ e\<^isub>2'" by fact have ih2:"\t. e\<^isub>2 \ t \ e\<^isub>2' = t" by fact - have "e\<^isub>1'[x::=e\<^isub>2'] \ e'" by fact have ih3: "\t. e\<^isub>1'[x::=e\<^isub>2'] \ t \ e' = t" by fact - have f:"x\(e\<^isub>1,e\<^isub>2,e')" by fact + have app: "App e\<^isub>1 e\<^isub>2 \ t\<^isub>2" by fact + have vc: "x\e\<^isub>1" "x\e\<^isub>2" by fact then have "x \ App e\<^isub>1 e\<^isub>2" by auto - moreover - have app:"App e\<^isub>1 e\<^isub>2 \ t\<^isub>2" by fact - ultimately have "x\t\<^isub>2" using fresh_preserved by blast - then have "x\(e\<^isub>1,e\<^isub>2,t\<^isub>2)" using f by auto - then obtain f\<^isub>1'' f\<^isub>2'' where x1:"e\<^isub>1 \ Lam [x]. f\<^isub>1''" and x2:"e\<^isub>2 \ f\<^isub>2''" and x3:"f\<^isub>1''[x::=f\<^isub>2''] \ t\<^isub>2" - using app by auto - then have "Lam [x]. f\<^isub>1'' = Lam [x]. e\<^isub>1'" using ih1 by simp - then have "f\<^isub>1'' = e\<^isub>1'" by (auto simp add: trm.inject alpha) - moreover have "f\<^isub>2''=e\<^isub>2'" using x2 ih2 by simp + then have vc': "x\t\<^isub>2" using fresh_preserved app by blast + from vc vc' obtain f\<^isub>1 f\<^isub>2 where x1: "e\<^isub>1 \ Lam [x]. f\<^isub>1" and x2: "e\<^isub>2 \ f\<^isub>2" and x3: "f\<^isub>1[x::=f\<^isub>2] \ t\<^isub>2" + using app by (auto simp add: fresh_prod) + then have "Lam [x]. f\<^isub>1 = Lam [x]. e\<^isub>1'" using ih1 by simp + then + have "f\<^isub>1 = e\<^isub>1'" by (auto simp add: trm.inject alpha) + moreover + have "f\<^isub>2 = e\<^isub>2'" using x2 ih2 by simp ultimately have "e\<^isub>1'[x::=e\<^isub>2'] \ t\<^isub>2" using x3 by simp thus ?case using ih3 by simp next - case (b_CaseL x\<^isub>1 e e\<^isub>2 e'' x\<^isub>2 e\<^isub>1 e' t\<^isub>2) - have ih1:"\ t. e \ t \ InL e' = t" by fact + case (b_CaseL x\<^isub>1 e e\<^isub>2 e'' x\<^isub>2 e\<^isub>1 e' t\<^isub>2) + have fs: "x\<^isub>1\e" "x\<^isub>1\t\<^isub>2" "x\<^isub>2\e" "x\<^isub>2\t\<^isub>2" by fact + have ih1:"\t. e \ t \ InL e' = t" by fact have ih2:"\t. e\<^isub>1[x\<^isub>1::=e'] \ t \ e'' = t" by fact - have ha:"\ t. (e \ InR t) \ False" using ih1 by force + have ha: "\(\t. e \ InR t)" using ih1 by force have "Case e of inl x\<^isub>1 \ e\<^isub>1 | inr x\<^isub>2 \ e\<^isub>2 \ t\<^isub>2" by fact - then obtain xe' where "e\InL xe'" and h:"e\<^isub>1[x\<^isub>1::=xe']\t\<^isub>2" using ha by auto - then have "InL xe'=InL e'" using ih1 by simp - then have "xe'=e'" by (simp add: trm.inject) + then obtain f' where "e \ InL f'" and h: "e\<^isub>1[x\<^isub>1::=f']\t\<^isub>2" using ha fs by auto + then have "InL f' = InL e'" using ih1 by simp + then have "f' = e'" by (simp add: trm.inject) then have "e\<^isub>1[x\<^isub>1::=e'] \ t\<^isub>2" using h by simp - then show "e''=t\<^isub>2" using ih2 by simp + then show "e'' = t\<^isub>2" using ih2 by simp next case (b_CaseR x\<^isub>1 e e\<^isub>2 e'' x\<^isub>2 e\<^isub>1 e' t\<^isub>2 ) - have ih1:"\ t. e \ t \ InR e' = t" by fact - have ih2:"\t. e\<^isub>2[x\<^isub>2::=e'] \ t \ e'' = t" by fact - have a:"\ t. (e \ InL t \ False)" using ih1 by force + have fs: "x\<^isub>1\e" "x\<^isub>1\t\<^isub>2" "x\<^isub>2\e" "x\<^isub>2\t\<^isub>2" by fact + have ih1: "\t. e \ t \ InR e' = t" by fact + have ih2: "\t. e\<^isub>2[x\<^isub>2::=e'] \ t \ e'' = t" by fact + have ha: "\(\t. e \ InL t)" using ih1 by force have "Case e of inl x\<^isub>1 \ e\<^isub>1 | inr x\<^isub>2 \ e\<^isub>2 \ t\<^isub>2" by fact - then obtain xe' where "e\InR xe'" and h:"e\<^isub>2[x\<^isub>2::=xe']\t\<^isub>2" using a by auto - then have "InR xe'=InR e'" using ih1 by simp + then obtain f' where "e \ InR f'" and h: "e\<^isub>2[x\<^isub>2::=f']\t\<^isub>2" using ha fs by auto + then have "InR f' = InR e'" using ih1 by simp then have "e\<^isub>2[x\<^isub>2::=e'] \ t\<^isub>2" using h by (simp add: trm.inject) - thus "e''=t\<^isub>2" using ih2 by simp -qed (fast)+ + thus "e'' = t\<^isub>2" using ih2 by simp +qed (force simp add: trm.inject)+ lemma not_val_App[simp]: shows @@ -1103,51 +1032,52 @@ "\ val (Case e of inl x\<^isub>1 \ e\<^isub>1 | inr x\<^isub>2 \ e\<^isub>2)" by auto -lemma reduces_to_value: +lemma reduces_evaluates_to_values: assumes h:"t \ t'" shows "val t'" - using h by (induct, auto) + using h by (induct) (auto) -lemma type_prod_down_pair: - assumes "\ \ t : Data (DProd S\<^isub>1 S\<^isub>2)" and "t \ t'" +lemma type_prod_evaluates_to_pairs: + assumes a: "\ \ t : Data (DProd S\<^isub>1 S\<^isub>2)" + and b: "t \ t'" obtains t\<^isub>1 t\<^isub>2 where "t' = Pr t\<^isub>1 t\<^isub>2" proof - have "\ \ t' : Data (DProd S\<^isub>1 S\<^isub>2)" using assms subject_reduction by simp moreover - have "val t'" using reduces_to_value assms by simp + have "val t'" using reduces_evaluates_to_values assms by simp ultimately obtain t\<^isub>1 t\<^isub>2 where "t' = Pr t\<^isub>1 t\<^isub>2" by (cases, auto simp add:ty.inject data.inject) thus ?thesis using prems by auto qed -lemma type_sum_down_or: +lemma type_sum_evaluates_to_ins: assumes "\ \ t : Data (DSum \\<^isub>1 \\<^isub>2)" and "t \ t'" - shows "(\ t''. t' = InL t'') \ (\ t''. t' = InR t'')" + shows "(\t''. t' = InL t'') \ (\t''. t' = InR t'')" proof - have "\ \ t' : Data (DSum \\<^isub>1 \\<^isub>2)" using assms subject_reduction by simp moreover - have "val t'" using reduces_to_value assms by simp + have "val t'" using reduces_evaluates_to_values assms by simp ultimately obtain t'' where "t' = InL t'' \ t' = InR t''" by (cases, auto simp add:ty.inject data.inject) thus ?thesis by auto qed -lemma type_arrow_down_lam: +lemma type_arrow_evaluates_to_lams: assumes "\ \ t : \ \ \" and "t \ t'" obtains x t'' where "t' = Lam [x]. t''" proof - have "\ \ t' : \ \ \" using assms subject_reduction by simp moreover - have "val t'" using reduces_to_value assms by simp + have "val t'" using reduces_evaluates_to_values assms by simp ultimately obtain x t'' where "t' = Lam [x]. t''" by (cases, auto simp add:ty.inject data.inject) thus ?thesis using prems by auto qed -lemma type_nat_down_const: +lemma type_nat_evaluates_to_consts: assumes "\ \ t : Data DNat" and "t \ t'" obtains n where "t' = Const n" proof - have "\ \ t' : Data DNat " using assms subject_reduction by simp - moreover have "val t'" using reduces_to_value assms by simp + moreover have "val t'" using reduces_evaluates_to_values assms by simp ultimately obtain n where "t' = Const n" by (cases, auto simp add:ty.inject data.inject) thus ?thesis using prems by auto qed @@ -1170,6 +1100,14 @@ by (nominal_induct S arbitrary: e rule:data.induct) (auto) +lemma V'_eqvt: + fixes pi::"name prm" + assumes a: "v \ V' S" + shows "(pi\v) \ V' S" +using a +by (nominal_induct S arbitrary: v rule: data.induct) + (auto simp add: trm.inject) + consts V :: "ty \ trm set" @@ -1179,6 +1117,27 @@ apply(rule TrueI)+ done +lemma V_eqvt: + fixes pi::"name prm" + assumes a: "x\V T" + shows "(pi\x)\V T" +using a +apply(nominal_induct T arbitrary: pi x rule: ty.induct) +apply(auto simp add: trm.inject perm_set_def) +apply(perm_simp add: V'_eqvt) +apply(rule_tac x="pi\xa" in exI) +apply(rule_tac x="pi\e" in exI) +apply(simp) +apply(auto) +apply(drule_tac x="(rev pi)\v" in bspec) +apply(force) +apply(auto) +apply(rule_tac x="pi\v'" in exI) +apply(auto) +apply(drule_tac pi="pi" in big_eqvt) +apply(perm_simp add: eqvt) +done + lemma V_arrow_elim_weak[elim] : assumes h:"u \ (V (T\<^isub>1 \ T\<^isub>2))" obtains a t where "u = Lam[a].t" and "\ v \ (V T\<^isub>1). \ v'. t[a::=v] \ v' \ v' \ V T\<^isub>2" @@ -1186,27 +1145,34 @@ lemma V_arrow_elim_strong[elim]: fixes c::"'a::fs_name" - assumes h: "u \ (V (T\<^isub>1 \ T\<^isub>2))" - obtains a t where "a\c" "u = Lam[a].t" "\ v \ (V T\<^isub>1). \ v'. t[a::=v] \ v' \ v' \ V T\<^isub>2" + assumes h: "u \ V (T\<^isub>1 \ T\<^isub>2)" + obtains a t where "a\c" "u = Lam[a].t" "\v \ (V T\<^isub>1). \ v'. t[a::=v] \ v' \ v' \ V T\<^isub>2" using h apply - apply(erule V_arrow_elim_weak) -apply(subgoal_tac "\a'::name. a'\(a,t,c)") +apply(subgoal_tac "\a'::name. a'\(a,t,c)") (*A*) apply(erule exE) apply(drule_tac x="a'" in meta_spec) +apply(drule_tac x="[(a,a')]\t" in meta_spec) +apply(drule meta_mp) apply(simp) -apply(drule_tac x="[(a,a')]\t" in meta_spec) -apply(simp add: trm.inject alpha fresh_prod fresh_atm) +apply(drule meta_mp) +apply(simp add: trm.inject alpha fresh_left fresh_prod calc_atm fresh_atm) apply(perm_simp) -apply(simp add: fresh_left calc_atm) +apply(force) +apply(drule meta_mp) +apply(rule ballI) +apply(drule_tac x="[(a,a')]\v" in bspec) +apply(simp add: V_eqvt) apply(auto) -apply(simp add: subst_rename) -apply(subgoal_tac "[(a',a)]\t = [(a,a')]\t") -apply(simp) -apply(rule pt_name3) -apply(rule at_ds5[OF at_name_inst]) +apply(rule_tac x="[(a,a')]\v'" in exI) +apply(auto) +apply(drule_tac pi="[(a,a')]" in big_eqvt) +apply(perm_simp add: eqvt calc_atm) +apply(simp add: V_eqvt) +(*A*) apply(rule exists_fresh') -apply(simp add: fs_name1) +apply(simp add: fin_supp) done lemma V_are_values : @@ -1306,7 +1272,7 @@ then obtain v\<^isub>1 v\<^isub>2 where "\1> \ v\<^isub>1 \ v\<^isub>1 \ V (Data T\<^isub>a)" "\2> \ v\<^isub>2 \ v\<^isub>2 \ V (Data T\<^isub>b)" using prems by blast thus "\v. \1 t\<^isub>2> \ v \ v \ V T" using eq by auto -next +next case (Lam x e \ \ T) have ih:"\\ \ T. \\ Vcloses \; \ \ e : T\ \ \v. \ \ v \ v \ V T" by fact have as\<^isub>1: "\ Vcloses \" by fact @@ -1314,15 +1280,14 @@ have fs: "x\\" "x\\" by fact from as\<^isub>2 fs obtain T\<^isub>1 T\<^isub>2 where "(i)": "(x,T\<^isub>1)#\ \ e:T\<^isub>2" and "(ii)": "T = T\<^isub>1 \ T\<^isub>2" by auto - from "(i)" have "(iii)": "valid ((x,T\<^isub>1)#\)" by (simp add: typing_valid) + from "(i)" have "(iii)": "valid ((x,T\<^isub>1)#\)" by (simp add: typing_implies_valid) have "\v \ (V T\<^isub>1). \v'. (\)[x::=v] \ v' \ v' \ V T\<^isub>2" proof fix v assume "v \ (V T\<^isub>1)" with "(iii)" as\<^isub>1 have "(x,v)#\ Vcloses (x,T\<^isub>1)#\" using monotonicity by auto with ih "(i)" obtain v' where "((x,v)#\) \ v' \ v' \ V T\<^isub>2" by blast - then have "\[x::=v] \ v' \ v' \ V T\<^isub>2" using fs - by (simp add: psubst_subst_psubst) + then have "\[x::=v] \ v' \ v' \ V T\<^isub>2" using fs by (simp add: psubst_subst_psubst) then show "\v'. \[x::=v] \ v' \ v' \ V T\<^isub>2" by auto qed then have "Lam[x].\ \ V (T\<^isub>1 \ T\<^isub>2)" by auto @@ -1336,21 +1301,20 @@ have th:"\ \ Case t' of inl n\<^isub>1 \ t\<^isub>1 | inr n\<^isub>2 \ t\<^isub>2 : T" by fact then obtain S\<^isub>1 S\<^isub>2 where hm:"\ \ t' : Data (DSum S\<^isub>1 S\<^isub>2)" and - hl:"(n\<^isub>1,Data S\<^isub>1)# \ \ t\<^isub>1 : T" and - hr:"(n\<^isub>2,Data S\<^isub>2)# \ \ t\<^isub>2 : T" using f by auto + hl:"(n\<^isub>1,Data S\<^isub>1)#\ \ t\<^isub>1 : T" and + hr:"(n\<^isub>2,Data S\<^isub>2)#\ \ t\<^isub>2 : T" using f by auto then obtain v\<^isub>0 where ht':"\ \ v\<^isub>0" and hS:"v\<^isub>0 \ V (Data (DSum S\<^isub>1 S\<^isub>2))" using prems h by blast (* We distinguish between the cases InL and InR *) - { - fix v\<^isub>0' + { fix v\<^isub>0' assume eqc:"v\<^isub>0 = InL v\<^isub>0'" and "v\<^isub>0' \ V' S\<^isub>1" then have inc: "v\<^isub>0' \ V (Data S\<^isub>1)" by auto - have "valid \" using th typing_valid by auto + have "valid \" using th typing_implies_valid by auto then moreover have "valid ((n\<^isub>1,Data S\<^isub>1)#\)" using f by auto then moreover have "(n\<^isub>1,v\<^isub>0')#\ Vcloses (n\<^isub>1,Data S\<^isub>1)#\" using inc h monotonicity by blast - moreover have ih:"\\ \ T. \\ Vcloses \; \ \ t\<^isub>1 : T\ \ - \v. \1> \ v \ v \ V T" by fact - ultimately obtain v\<^isub>1 where ho:"((n\<^isub>1,v\<^isub>0')#\)1> \ v\<^isub>1 \ v\<^isub>1 \ V T" using hl by blast + moreover + have ih:"\\ \ T. \\ Vcloses \; \ \ t\<^isub>1 : T\ \ \v. \1> \ v \ v \ V T" by fact + ultimately obtain v\<^isub>1 where ho: "((n\<^isub>1,v\<^isub>0')#\)1> \ v\<^isub>1 \ v\<^isub>1 \ V T" using hl by blast then have r:"\1>[n\<^isub>1::=v\<^isub>0'] \ v\<^isub>1 \ v\<^isub>1 \ V T" using psubst_subst_psubst f by simp then moreover have "n\<^isub>1\(\,\2>,v\<^isub>1,n\<^isub>2)" proof - @@ -1375,11 +1339,10 @@ ultimately have "\v. \1 \ t\<^isub>1 | inr n\<^isub>2 \ t\<^isub>2> \ v \ v \ V T" by auto } moreover - { - fix v\<^isub>0' + { fix v\<^isub>0' assume eqc:"v\<^isub>0 = InR v\<^isub>0'" and "v\<^isub>0' \ V' S\<^isub>2" then have inc:"v\<^isub>0' \ V (Data S\<^isub>2)" by auto - have "valid \" using th typing_valid by auto + have "valid \" using th typing_implies_valid by auto then moreover have "valid ((n\<^isub>2,Data S\<^isub>2)#\)" using f by auto then moreover have "(n\<^isub>2,v\<^isub>0')#\ Vcloses (n\<^isub>2,Data S\<^isub>2)#\" using inc h monotonicity by blast @@ -1407,7 +1370,7 @@ qed ultimately have "Case \ of inl n\<^isub>1 \ \1> | inr n\<^isub>2 \ \2> \ v\<^isub>2 \ v\<^isub>2 \ V T" using ht' eqc by auto then have "\v. \1 \ t\<^isub>1 | inr n\<^isub>2 \ t\<^isub>2> \ v \ v \ V T" using f by auto -} + } ultimately show "\v. \1 \ t\<^isub>1 | inr n\<^isub>2 \ t\<^isub>2> \ v \ v \ V T" using hS V_sum by blast qed (force)+