# HG changeset patch # User urbanc # Date 1199460922 -3600 # Node ID 41a014cc44c031e68733af98009b4139a0eb3563 # Parent 7711d60a5293d5312080b349302e7512d5490de7 partially adapted to new inversion rules diff -r 7711d60a5293 -r 41a014cc44c0 src/HOL/Nominal/Examples/SOS.thy --- a/src/HOL/Nominal/Examples/SOS.thy Fri Jan 04 09:34:11 2008 +0100 +++ b/src/HOL/Nominal/Examples/SOS.thy Fri Jan 04 16:35:22 2008 +0100 @@ -1,15 +1,15 @@ (* "$Id$" *) -(* *) -(* Formalisation of some typical SOS-proofs *) -(* *) -(* 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. *) -(* *) -(* The formalisation was done by Julien Narboux and *) -(* Christian Urban. *) +(* *) +(* Formalisation of some typical SOS-proofs. *) +(* *) +(* This work arose from a challenge suggested by Adam *) +(* Chlipala on the POPLmark mailing list. *) +(* *) +(* We thank Nick Benton for helping us with the *) +(* termination-proof for evaluation. *) +(* *) +(* The formalisation was done by Julien Narboux and *) +(* Christian Urban. *) theory SOS imports "../Nominal" @@ -17,72 +17,42 @@ atom_decl name -nominal_datatype data = - DNat - | DProd "data" "data" - | DSum "data" "data" - +text {* types and terms *} nominal_datatype ty = - Data "data" + TVar "nat" | Arrow "ty" "ty" ("_\_" [100,100] 100) nominal_datatype trm = Var "name" | Lam "\name\trm" ("Lam [_]._" [100,100] 100) | App "trm" "trm" - | Const "nat" - | Pr "trm" "trm" - | Fst "trm" - | Snd "trm" - | InL "trm" - | InR "trm" - | Case "trm" "\name\trm" "\name\trm" ("Case _ of inl _ \ _ | inr _ \ _" [100,100,100,100,100] 100) -lemma in_eqvt[eqvt]: - fixes pi::"name prm" - and x::"'a::pt_name" - assumes "x\X" - shows "pi\x \ pi\X" - using assms by (perm_simp add: pt_set_bij1a[OF pt_name_inst, OF at_name_inst]) - -lemma perm_data[simp]: - fixes D::"data" - and pi::"name prm" - shows "pi\D = D" - by (induct D rule: data.weak_induct) (simp_all) - -lemma perm_ty[simp]: - fixes T::"ty" - and pi::"name prm" - shows "pi\T = T" - by (induct T rule: ty.weak_induct) (simp_all) - -lemma fresh_ty[simp]: +lemma fresh_ty: fixes x::"name" and T::"ty" shows "x\T" - by (simp add: fresh_def supp_def) +by (induct T rule: ty.weak_induct) + (auto simp add: fresh_nat) -text {* substitution *} - +text {* Parallel and single substitution. *} fun lookup :: "(name\trm) list \ name \ trm" where "lookup [] x = Var x" | "lookup ((y,e)#\) x = (if x=y then e else lookup \ x)" -lemma lookup_eqvt: +lemma lookup_eqvt[eqvt]: fixes pi::"name prm" and \::"(name\trm) list" and X::"name" shows "pi\(lookup \ X) = lookup (pi\\) (pi\X)" -by (induct \, auto simp add: perm_bij) +by (induct \) (auto simp add: eqvts) lemma lookup_fresh: fixes z::"name" - assumes "z\\" and "z\x" + assumes a: "z\\" and b: "z\x" shows "z \lookup \ x" -using assms +using a b by (induct rule: lookup.induct) (auto simp add: fresh_list_cons) lemma lookup_fresh': @@ -92,8 +62,7 @@ by (induct rule: lookup.induct) (auto simp add: fresh_list_cons fresh_prod fresh_atm) -text {* Parallel Substitution *} - +(* parallel substitution *) consts psubst :: "(name\trm) list \ trm \ trm" ("_<_>" [95,95] 105) @@ -101,18 +70,10 @@ "\<(Var x)> = (lookup \ x)" "\<(App e\<^isub>1 e\<^isub>2)> = App (\1>) (\2>)" "x\\ \ \<(Lam [x].e)> = Lam [x].(\)" - "\<(Const n)> = Const n" - "\<(Pr e\<^isub>1 e\<^isub>2)> = Pr (\1>) (\2>)" - "\<(Fst e)> = Fst (\)" - "\<(Snd e)> = Snd (\)" - "\<(InL e)> = InL (\)" - "\<(InR e)> = InR (\)" - "\y\x; x\(e,e\<^isub>2,\); y\(e,e\<^isub>1,\)\ - \ \ e\<^isub>1 | inr y \ e\<^isub>2> = (Case (\) of inl x \ (\1>) | inr y \ (\2>))" -apply(finite_guess add: lookup_eqvt)+ +apply(finite_guess)+ apply(rule TrueI)+ apply(simp add: abs_fresh)+ -apply(fresh_guess add: fs_name1 lookup_eqvt)+ +apply(fresh_guess)+ done lemma psubst_eqvt[eqvt]: @@ -131,32 +92,23 @@ by (nominal_induct t avoiding: z \ t rule: trm.induct) (auto simp add: abs_fresh lookup_fresh) +lemma psubst_empty[simp]: + shows "[] = t" + by (nominal_induct t rule: trm.induct) + (auto simp add: fresh_list_nil) + +(* Single substitution *) abbreviation - subst :: "trm \ name \ trm \ trm" ("_[_::=_]" [100,100,100] 100) - where "t[x::=t'] \ ([(x,t')])" + subst :: "trm \ name \ trm \ trm" ("_[_::=_]" [100,100,100] 100) +where + "t[x::=t'] \ ([(x,t')])" lemma subst[simp]: shows "(Var x)[y::=t'] = (if x=y then t' else (Var x))" and "(App t\<^isub>1 t\<^isub>2)[y::=t'] = App (t\<^isub>1[y::=t']) (t\<^isub>2[y::=t'])" and "x\(y,t') \ (Lam [x].t)[y::=t'] = Lam [x].(t[y::=t'])" - and "(Const n)[y::=t'] = Const n" - and "(Pr e\<^isub>1 e\<^isub>2)[y::=t'] = Pr (e\<^isub>1[y::=t']) (e\<^isub>2[y::=t'])" - and "(Fst e)[y::=t'] = Fst (e[y::=t'])" - and "(Snd e)[y::=t'] = Snd (e[y::=t'])" - and "(InL e)[y::=t'] = InL (e[y::=t'])" - and "(InR e)[y::=t'] = InR (e[y::=t'])" - 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) -lemma subst_eqvt[eqvt]: - fixes pi::"name prm" - and t::"trm" - shows "pi\(t[x::=t']) = (pi\t)[(pi\x)::=(pi\t')]" - by (nominal_induct t avoiding: x t' rule: trm.induct) - (perm_simp add: fresh_bij)+ - lemma fresh_subst: fixes z::"name" and t\<^isub>1::"trm" @@ -168,60 +120,41 @@ (auto simp add: abs_fresh fresh_atm) lemma fresh_subst': - fixes z::"name" - and t\<^isub>1::"trm" - and t2::"trm" - assumes "z\[y].t\<^isub>1" and "z\t\<^isub>2" - shows "z\t\<^isub>1[y::=t\<^isub>2]" -using assms -by (nominal_induct t\<^isub>1 avoiding: y t\<^isub>2 z rule: trm.induct) - (auto simp add: abs_fresh fresh_nat fresh_atm) + assumes "x\e'" + shows "x\e[x::=e']" + using assms +by (nominal_induct e avoiding: x e' rule: trm.induct) + (auto simp add: fresh_atm abs_fresh fresh_nat) lemma forget: fixes x::"name" and L::"trm" - assumes "x\L" - shows "L[x::=P] = L" + assumes "x\e" + shows "e[x::=e'] = e" using assms - by (nominal_induct L avoiding: x P rule: trm.induct) + by (nominal_induct e avoiding: x e' rule: trm.induct) (auto simp add: fresh_atm abs_fresh) -lemma psubst_empty[simp]: - shows "[] = t" - by (nominal_induct t rule: trm.induct, auto simp add:fresh_list_nil) - lemma psubst_subst_psubst: -assumes h:"x\\" -shows "\[x::=e'] = ((x,e')#\)" -using h + assumes h:"x\\" + shows "\[x::=e'] = ((x,e')#\)" + using h apply(nominal_induct e avoiding: \ x e' rule: trm.induct) -apply(auto simp add: fresh_list_cons fresh_atm forget lookup_fresh lookup_fresh' fresh_psubst) +apply(auto simp add: fresh_list_cons fresh_atm forget lookup_fresh lookup_fresh') done -lemma fresh_subst_fresh: - assumes "a\e" - shows "a\t[a::=e]" -using assms -by (nominal_induct t avoiding: a e rule: trm.induct) - (auto simp add: fresh_atm abs_fresh fresh_nat) - -text {* Typing-Judgements *} +text {* Typing Judgements *} inductive - valid :: "(name \ 'a::pt_name) list \ bool" + valid :: "(name\ty) list \ bool" where - v_nil[intro]: "valid []" - | v_cons[intro]: "\valid \;x\\\ \ valid ((x,T)#\)" + v_nil[intro]: "valid []" +| v_cons[intro]: "\valid \;x\\\ \ valid ((x,T)#\)" equivariance valid -inductive_cases - valid_cons_inv_auto[elim]:"valid ((x,T)#\)" - -abbreviation - "sub" :: "(name\ty) list \ (name\ty) list \ bool" ("_ \ _" [55,55] 55) -where - "\\<^isub>1 \ \\<^isub>2 \ \x T. (x,T)\set \\<^isub>1 \ (x,T)\set \\<^isub>2" +inductive_cases + valid_cons_inv_auto[elim]: "valid ((x,T)#\)" lemma type_unicity_in_context: assumes asm1: "(x,t\<^isub>2) \ set ((x,t\<^isub>1)#\)" @@ -236,161 +169,70 @@ qed lemma case_distinction_on_context: - fixes \::"(name \ ty) list" + fixes \::"(name\ty) list" assumes asm1: "valid ((m,t)#\)" and asm2: "(n,U) \ set ((m,T)#\)" shows "(n,U) = (m,T) \ ((n,U) \ set \ \ n \ m)" proof - -from asm2 have "(n,U) \ set [(m,T)] \ (n,U) \ set \" by auto -moreover -{ assume eq: "m=n" - assume "(n,U) \ set \" - then have "\ n\\" - by (induct \) (auto simp add: fresh_list_cons fresh_prod fresh_atm) - moreover have "m\\" using asm1 by auto - ultimately have False using eq by auto -} -ultimately show ?thesis by auto + from asm2 have "(n,U) \ set [(m,T)] \ (n,U) \ set \" by auto + moreover + { assume eq: "m=n" + assume "(n,U) \ set \" + then have "\ n\\" + by (induct \) (auto simp add: fresh_list_cons fresh_prod fresh_atm) + moreover have "m\\" using asm1 by auto + ultimately have False using eq by auto + } + ultimately show ?thesis by auto qed inductive typing :: "(name\ty) list\trm\ty\bool" ("_ \ _ : _" [60,60,60] 60) where - t_Var[intro]: "\valid \; (x,T)\set \\\ \ \ Var x : T" -| t_App[intro]: "\\ \ e\<^isub>1 : T\<^isub>1\T\<^isub>2; \ \ e\<^isub>2 : T\<^isub>1\\ \ \ App e\<^isub>1 e\<^isub>2 : T\<^isub>2" + t_Var[intro]: "\valid \; (x,T)\set \\ \ \ \ Var x : T" +| t_App[intro]: "\\ \ e\<^isub>1 : T\<^isub>1\T\<^isub>2; \ \ e\<^isub>2 : T\<^isub>1\ \ \ \ App e\<^isub>1 e\<^isub>2 : T\<^isub>2" | t_Lam[intro]: "\x\\; (x,T\<^isub>1)#\ \ e : T\<^isub>2\ \ \ \ Lam [x].e : T\<^isub>1\T\<^isub>2" -| t_Const[intro]: "valid \ \ \ \ Const n : Data(DNat)" -| t_Pr[intro]: "\\ \ e\<^isub>1 : Data(S\<^isub>1); \ \ e\<^isub>2 : Data(S\<^isub>2)\ \ \ \ Pr e\<^isub>1 e\<^isub>2 : Data (DProd S\<^isub>1 S\<^isub>2)" -| t_Fst[intro]: "\\ \ e : Data(DProd S\<^isub>1 S\<^isub>2)\ \ \ \ Fst e : Data(S\<^isub>1)" -| t_Snd[intro]: "\\ \ e : Data(DProd S\<^isub>1 S\<^isub>2)\ \ \ \ Snd e : Data(S\<^isub>2)" -| t_InL[intro]: "\\ \ e : Data(S\<^isub>1)\ \ \ \ InL e : Data(DSum S\<^isub>1 S\<^isub>2)" -| t_InR[intro]: "\\ \ e : Data(S\<^isub>2)\ \ \ \ InR e : Data(DSum S\<^isub>1 S\<^isub>2)" -| t_Case[intro]: "\x\<^isub>1\(\,e,e\<^isub>2,x\<^isub>2); x\<^isub>2\(\,e,e\<^isub>1,x\<^isub>1); \ \ e: Data(DSum S\<^isub>1 S\<^isub>2); - (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" equivariance typing nominal_inductive typing - by (simp_all add: abs_fresh fresh_prod fresh_atm) - -lemmas typing_eqvt' = typing.eqvt[simplified] + by (simp_all add: abs_fresh fresh_ty) lemma typing_implies_valid: - assumes "\ \ t : T" + assumes a: "\ \ t : T" shows "valid \" - using assms + using a by (induct) (auto) -declare trm.inject [simp add] -declare ty.inject [simp add] -declare data.inject [simp add] - - -inductive_cases typing_inv_auto[elim]: - "\ \ Lam [x].t : T" - "\ \ Var x : T" - "\ \ App x y : T" - "\ \ Const n : T" - "\ \ Fst x : T" - "\ \ Snd x : T" - "\ \ InL x : T" - "\ \ InL x : Data (DSum T\<^isub>1 T2)" - "\ \ InR x : T" - "\ \ InR x : Data (DSum T\<^isub>1 T2)" - "\ \ Pr x y : T" - "\ \ Pr e\<^isub>1 e\<^isub>2 : Data (DProd \1 \\<^isub>2)" - "\ \ Case e of inl x\<^isub>1 \ e\<^isub>1 | inr x\<^isub>2 \ e\<^isub>2 : T" - -declare trm.inject [simp del] -declare ty.inject [simp del] -declare data.inject [simp del] - -lemma t_Lam_elim[elim]: +lemma t_Lam_elim: assumes a1:"\ \ Lam [x].t : T" and a2: "x\\" obtains T\<^isub>1 and T\<^isub>2 where "(x,T\<^isub>1)#\ \ t : T\<^isub>2" and "T=T\<^isub>1\T\<^isub>2" -proof - - from a1 obtain x' t' T\<^isub>1 T\<^isub>2 - where b1: "x'\\" and b2: "(x',T\<^isub>1)#\ \ t' : T\<^isub>2" and b3: "[x'].t' = [x].t" and b4: "T=T\<^isub>1\T\<^isub>2" - by auto - obtain c::"name" where "c\(\,x,x',t,t')" by (erule exists_fresh[OF fs_name1]) - then have fs: "c\\" "c\x" "c\x'" "c\t" "c\t'" by (simp_all add: fresh_atm[symmetric]) - then have b5: "[(x',c)]\t'=[(x,c)]\t" using b3 fs by (simp add: alpha_fresh) - have "([(x,c)]\[(x',c)]\((x',T\<^isub>1)#\)) \ ([(x,c)]\[(x',c)]\t') : T\<^isub>2" using b2 - by (simp only: typing_eqvt') - then have "(x,T\<^isub>1)#\ \ t : T\<^isub>2" using fs b1 a2 b5 by (perm_simp add: calc_atm) - then show ?thesis using prems b4 by simp -qed +using a1 a2 +by (cases rule: typing.strong_cases [where x="x"]) + (auto simp add: abs_fresh fresh_ty alpha trm.inject) -lemma t_Case_elim[elim]: - assumes "\ \ Case e of inl x\<^isub>1 \ e\<^isub>1 | inr x\<^isub>2 \ e\<^isub>2 : T" and "x\<^isub>1\\" and "x\<^isub>2\\" - obtains \\<^isub>1 \\<^isub>2 where "\ \ e : Data (DSum \\<^isub>1 \\<^isub>2)" - and "(x\<^isub>1, Data \\<^isub>1)#\ \ e\<^isub>1 : T" - and "(x\<^isub>2, Data \\<^isub>2)#\ \ e\<^isub>2 : T" -proof - - have f:"x\<^isub>1\\" "x\<^isub>2\\" by fact+ - have "\ \ Case e of inl x\<^isub>1 \ e\<^isub>1 | inr x\<^isub>2 \ e\<^isub>2 : T" by fact - then obtain \\<^isub>1 \\<^isub>2 x\<^isub>1' x\<^isub>2' e\<^isub>1' e\<^isub>2' where - h:"\ \ e : Data (DSum \\<^isub>1 \\<^isub>2)" and - h1:"(x\<^isub>1',Data \\<^isub>1)#\ \ e\<^isub>1' : T" and - h2:"(x\<^isub>2',Data \\<^isub>2)#\ \ e\<^isub>2' : T" and - e1:"[x\<^isub>1].e\<^isub>1=[x\<^isub>1'].e\<^isub>1'" and e2:"[x\<^isub>2].e\<^isub>2=[x\<^isub>2'].e\<^isub>2'" - by auto - obtain c::name where f':"c \ (x\<^isub>1,x\<^isub>1',e\<^isub>1,e\<^isub>1',\)" by (erule exists_fresh[OF fs_name1]) - have e1':"[(x\<^isub>1,c)]\e\<^isub>1 = [(x\<^isub>1',c)]\e\<^isub>1'" using e1 f' by (auto simp add: alpha_fresh fresh_prod fresh_atm) - 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_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' - by (auto simp add: perm_fresh_fresh) - then have "([(x\<^isub>1,c)]\(c,Data \\<^isub>1)) #\ \ e\<^isub>1 : T" by perm_simp - then have g1:"(x\<^isub>1, Data \\<^isub>1)#\ \ e\<^isub>1 : T" using f' by (auto simp add: fresh_atm calc_atm fresh_prod) - (* The second part of the proof is the same *) - obtain c::name where f':"c \ (x\<^isub>2,x\<^isub>2',e\<^isub>2,e\<^isub>2',\)" by (erule exists_fresh[OF fs_name1]) - have e2':"[(x\<^isub>2,c)]\e\<^isub>2 = [(x\<^isub>2',c)]\e\<^isub>2'" using e2 f' by (auto simp add: alpha_fresh fresh_prod fresh_atm) - 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_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' - by (auto simp add: perm_fresh_fresh) - then have "([(x\<^isub>2,c)]\(c,Data \\<^isub>2)) #\ \ e\<^isub>2 : T" by perm_simp - then have g2:"(x\<^isub>2,Data \\<^isub>2)#\ \ e\<^isub>2 : T" using f' by (auto simp add: fresh_atm calc_atm fresh_prod) - show ?thesis using g1 g2 prems by auto -qed +abbreviation + "sub_context" :: "(name\ty) list \ (name\ty) list \ bool" ("_ \ _" [55,55] 55) +where + "\\<^isub>1 \ \\<^isub>2 \ \x T. (x,T)\set \\<^isub>1 \ (x,T)\set \\<^isub>2" lemma weakening: - fixes \\<^isub>1 \\<^isub>2 :: "(name\ty) list" + fixes \\<^isub>1 \\<^isub>2::"(name\ty) list" assumes "\\<^isub>1 \ e: T" and "valid \\<^isub>2" and "\\<^isub>1 \ \\<^isub>2" shows "\\<^isub>2 \ e: T" using assms proof(nominal_induct \\<^isub>1 e T avoiding: \\<^isub>2 rule: typing.strong_induct) case (t_Lam x \\<^isub>1 T\<^isub>1 t T\<^isub>2 \\<^isub>2) + have vc: "x\\\<^isub>2" by fact have ih: "\valid ((x,T\<^isub>1)#\\<^isub>2); (x,T\<^isub>1)#\\<^isub>1 \ (x,T\<^isub>1)#\\<^isub>2\ \ (x,T\<^isub>1)#\\<^isub>2 \ t : T\<^isub>2" by fact - have H1: "valid \\<^isub>2" by fact - have H2: "\\<^isub>1 \ \\<^isub>2" by fact - have fs: "x\\\<^isub>2" by fact - then have "valid ((x,T\<^isub>1)#\\<^isub>2)" using H1 by auto - moreover have "(x,T\<^isub>1)#\\<^isub>1 \ (x,T\<^isub>1)#\\<^isub>2" using H2 by auto + have "valid \\<^isub>2" by fact + then have "valid ((x,T\<^isub>1)#\\<^isub>2)" using vc by auto + moreover + have "\\<^isub>1 \ \\<^isub>2" by fact + then have "(x,T\<^isub>1)#\\<^isub>1 \ (x,T\<^isub>1)#\\<^isub>2" by simp ultimately have "(x,T\<^isub>1)#\\<^isub>2 \ t : T\<^isub>2" using ih by simp - thus "\\<^isub>2 \ Lam [x].t : T\<^isub>1\T\<^isub>2" using fs by auto -next - case (t_Case x\<^isub>1 \\<^isub>1 e e\<^isub>2 x\<^isub>2 e\<^isub>1 S\<^isub>1 S\<^isub>2 T \\<^isub>2) - then have ih\<^isub>1: "valid ((x\<^isub>1,Data S\<^isub>1)#\\<^isub>2) \ (x\<^isub>1,Data S\<^isub>1)#\\<^isub>2 \ e\<^isub>1 : T" - and ih\<^isub>2: "valid ((x\<^isub>2,Data S\<^isub>2)#\\<^isub>2) \ (x\<^isub>2,Data S\<^isub>2)#\\<^isub>2 \ e\<^isub>2 : T" - and ih\<^isub>3: "\\<^isub>2 \ e : Data (DSum S\<^isub>1 S\<^isub>2)" by auto - have fs\<^isub>1: "x\<^isub>1\\\<^isub>2" "x\<^isub>1\e" "x\<^isub>1\e\<^isub>2" "x\<^isub>1\x\<^isub>2" by fact+ - have fs\<^isub>2: "x\<^isub>2\\\<^isub>2" "x\<^isub>2\e" "x\<^isub>2\e\<^isub>1" "x\<^isub>2\x\<^isub>1" by fact+ - have "valid \\<^isub>2" by fact - then have "valid ((x\<^isub>1,Data S\<^isub>1)#\\<^isub>2)" and "valid ((x\<^isub>2,Data S\<^isub>2)#\\<^isub>2)" using fs\<^isub>1 fs\<^isub>2 by auto - then have "(x\<^isub>1, Data S\<^isub>1)#\\<^isub>2 \ e\<^isub>1 : T" and "(x\<^isub>2, Data S\<^isub>2)#\\<^isub>2 \ e\<^isub>2 : T" using ih\<^isub>1 ih\<^isub>2 by simp_all - with ih\<^isub>3 show "\\<^isub>2 \ Case e of inl x\<^isub>1 \ e\<^isub>1 | inr x\<^isub>2 \ e\<^isub>2 : T" using fs\<^isub>1 fs\<^isub>2 by auto + with vc show "\\<^isub>2 \ Lam [x].t : T\<^isub>1\T\<^isub>2" by auto qed (auto) lemma context_exchange: @@ -401,18 +243,21 @@ 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)#\)" - by (auto simp: fresh_list_cons fresh_atm) + by (auto simp: fresh_list_cons fresh_atm fresh_ty) moreover have "(x\<^isub>1,T\<^isub>1)#(x\<^isub>2,T\<^isub>2)#\ \ (x\<^isub>2,T\<^isub>2)#(x\<^isub>1,T\<^isub>1)#\" by auto ultimately show "(x\<^isub>2,T\<^isub>2)#(x\<^isub>1,T\<^isub>1)#\ \ e : T" using a by (auto intro: weakening) qed lemma typing_var_unicity: - assumes "(x,t\<^isub>1)#\ \ Var x : t\<^isub>2" - shows "t\<^isub>1=t\<^isub>2" + assumes a: "(x,T\<^isub>1)#\ \ Var x : T\<^isub>2" + shows "T\<^isub>1=T\<^isub>2" proof - - have "(x,t\<^isub>2) \ set ((x,t\<^isub>1)#\)" and "valid ((x,t\<^isub>1)#\)" using assms by auto - thus "t\<^isub>1=t\<^isub>2" by (simp only: type_unicity_in_context) + have "(x,T\<^isub>2) \ set ((x,T\<^isub>1)#\)" using a + by (cases) (auto simp add: trm.inject) + moreover + have "valid ((x,T\<^isub>1)#\)" using a by (simp add: typing_implies_valid) + ultimately show "T\<^isub>1=T\<^isub>2" by (simp only: type_unicity_in_context) qed lemma typing_substitution: @@ -434,7 +279,7 @@ next case False assume as: "x\y" - have "(y,T) \ set ((x,T')#\)" using h1 by auto + have "(y,T) \ set ((x,T')#\)" using h1 by (cases) (auto simp add: trm.inject) then have "(y,T) \ set \" using as by auto moreover have "valid \" using h2 by (simp only: typing_implies_valid) @@ -442,49 +287,44 @@ qed next case (Lam y t \ e' x T) - have vc: "y\\" "y\x" "y\e'" by fact+ + have vc: "y\\" "y\x" "y\e'" by fact have pr1: "\ \ e' : T'" by fact have pr2: "(x,T')#\ \ Lam [y].t : T" by fact then obtain T\<^isub>1 T\<^isub>2 where pr2': "(y,T\<^isub>1)#(x,T')#\ \ t : T\<^isub>2" and eq: "T = T\<^isub>1\T\<^isub>2" - using vc by (auto simp add: fresh_list_cons) + using vc by (auto elim: t_Lam_elim simp add: fresh_list_cons fresh_ty) 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_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)#\ \ e' : T'" using pr1 by (simp add: weakening) then have "(y,T\<^isub>1)#\ \ t[x::=e'] : T\<^isub>2" using ih pr2'' by simp - then have "\ \ Lam [y].(t[x::=e']) : T\<^isub>1\T\<^isub>2" using vc by (auto intro: t_Lam) - thus "\ \ (Lam [y].t)[x::=e'] : T" using vc eq by simp + then have "\ \ Lam [y].(t[x::=e']) : T\<^isub>1\T\<^isub>2" using vc by auto + then show "\ \ (Lam [y].t)[x::=e'] : T" using vc eq by simp next - case (Case t\<^isub>1 x\<^isub>1 t\<^isub>2 x\<^isub>2 t3 \ e' x T) - have vc: "x\<^isub>1\\" "x\<^isub>1\e'" "x\<^isub>1\x""x\<^isub>1\t\<^isub>1" "x\<^isub>1\t3" "x\<^isub>2\\" - "x\<^isub>2\e'" "x\<^isub>2\x" "x\<^isub>2\t\<^isub>1" "x\<^isub>2\t\<^isub>2" "x\<^isub>2\x\<^isub>1" by fact+ - have as1: "\ \ e' : T'" by fact - have as2: "(x,T')#\ \ Case t\<^isub>1 of inl x\<^isub>1 \ t\<^isub>2 | inr x\<^isub>2 \ t3 : T" by fact - then obtain S\<^isub>1 S\<^isub>2 where - h1:"(x,T')#\ \ t\<^isub>1 : Data (DSum S\<^isub>1 S\<^isub>2)" and - h2:"(x\<^isub>1,Data S\<^isub>1)#(x,T')#\ \ t\<^isub>2 : T" and - h3:"(x\<^isub>2,Data S\<^isub>2)#(x,T')#\ \ t3 : T" - using vc by (auto simp add: fresh_list_cons) - have ih1: "\(x,T')#\ \ t\<^isub>1 : Data (DSum S\<^isub>1 S\<^isub>2); \ \ e' : T'\ \ \ \ t\<^isub>1[x::=e'] : Data (DSum S\<^isub>1 S\<^isub>2)" - and ih2: "\(x,T')#(x\<^isub>1,Data S\<^isub>1)#\ \ t\<^isub>2:T; (x\<^isub>1,Data S\<^isub>1)#\ \ e':T'\ \ (x\<^isub>1,Data S\<^isub>1)#\ \ t\<^isub>2[x::=e']:T" - and ih3: "\(x,T')#(x\<^isub>2,Data S\<^isub>2)#\ \ t3:T; (x\<^isub>2,Data S\<^isub>2)#\ \ e':T'\ \ (x\<^isub>2,Data S\<^isub>2)#\ \ t3[x::=e']:T" - by fact+ - from h2 have h2': "(x,T')#(x\<^isub>1,Data S\<^isub>1)#\ \ t\<^isub>2 : T" by (rule context_exchange) - 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_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_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" - using vc by (auto simp add: fresh_atm fresh_subst) - thus "\ \ (Case t\<^isub>1 of inl x\<^isub>1 \ t\<^isub>2 | inr x\<^isub>2 \ t3)[x::=e'] : T" using vc by simp -qed (simp, fast)+ + case (App e1 e2 \ e' x T) + have "(x,T')#\ \ App e1 e2 : T" by fact + then obtain Tn where a1: "(x,T')#\ \ e1 : Tn \ T" + and a2: "(x,T')#\ \ e2 : Tn" + by (cases) (auto simp add: trm.inject) + have a3: "\ \ e' : T'" by fact + have ih1: "\(x,T')#\ \ e1 : Tn\T; \ \ e' : T'\ \ \ \ e1[x::=e'] : Tn\T" by fact + have ih2: "\(x,T')#\ \ e2 : Tn; \ \ e' : T'\ \ \ \ e2[x::=e'] : Tn" by fact + then show ?case using a1 a2 a3 ih1 ih2 by auto +qed + +text {* Values *} +inductive + val :: "trm\bool" +where + v_Lam[intro]: "val (Lam [x].e)" + +equivariance val + +lemma not_val_App[simp]: + shows + "\ val (App e\<^isub>1 e\<^isub>2)" + "\ val (Var x)" + by (auto elim: val.cases) text {* Big-Step Evaluation *} @@ -493,195 +333,45 @@ where b_Lam[intro]: "Lam [x].e \ Lam [x].e" | b_App[intro]: "\x\(e\<^isub>1,e\<^isub>2,e'); e\<^isub>1\Lam [x].e; e\<^isub>2\e\<^isub>2'; e[x::=e\<^isub>2']\e'\ \ App e\<^isub>1 e\<^isub>2 \ e'" -| b_Const[intro]: "Const n \ Const n" -| b_Pr[intro]: "\e\<^isub>1\e\<^isub>1'; e\<^isub>2\e\<^isub>2'\ \ Pr e\<^isub>1 e\<^isub>2 \ Pr e\<^isub>1' e\<^isub>2'" -| b_Fst[intro]: "e\Pr e\<^isub>1 e\<^isub>2 \ Fst e\e\<^isub>1" -| b_Snd[intro]: "e\Pr e\<^isub>1 e\<^isub>2 \ Snd e\e\<^isub>2" -| b_InL[intro]: "e\e' \ InL e \ InL e'" -| b_InR[intro]: "e\e' \ InR e \ InR e'" -| b_CaseL[intro]: "\x\<^isub>1\(e,e\<^isub>2,e'',x\<^isub>2); x\<^isub>2\(e,e\<^isub>1,e'',x\<^isub>1) ; e\InL e'; e\<^isub>1[x\<^isub>1::=e']\e''\ - \ Case e of inl x\<^isub>1 \ e\<^isub>1 | inr x\<^isub>2 \ e\<^isub>2 \ e''" -| b_CaseR[intro]: "\x\<^isub>1\(e,e\<^isub>2,e'',x\<^isub>2); x\<^isub>2\(e,e\<^isub>1,e'',x\<^isub>1) ; e\InR e'; e\<^isub>2[x\<^isub>2::=e']\e''\ - \ Case e of inl x\<^isub>1 \ e\<^isub>1 | inr x\<^isub>2 \ e\<^isub>2 \ e''" equivariance big nominal_inductive big - by (simp_all add: abs_fresh fresh_prod fresh_atm) - -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 - fixes t'::trm - assumes "e \ e'" and "x\e" - shows "x\e'" - using assms by (induct) (auto simp add:fresh_subst') - -declare trm.inject [simp add] -declare ty.inject [simp add] -declare data.inject [simp add] - -inductive_cases b_inv_auto[elim]: - "App e\<^isub>1 e\<^isub>2 \ t" - "Case e of inl x\<^isub>1 \ e\<^isub>1 | inr x\<^isub>2 \ e\<^isub>2 \ t" - "Lam[x].t \ t" - "Const n \ t" - "Fst e \ t" - "Snd e \ t" - "InL e \ t" - "InR e \ t" - "Pr e\<^isub>1 e\<^isub>2 \ t" - -declare trm.inject [simp del] -declare ty.inject [simp del] -declare data.inject [simp del] - -lemma b_App_elim[elim]: - assumes "App e\<^isub>1 e\<^isub>2 \ e'" and "x\(e\<^isub>1,e\<^isub>2,e')" - obtains f\<^isub>1 and f\<^isub>2 where "e\<^isub>1 \ Lam [x]. f\<^isub>1" "e\<^isub>2 \ f\<^isub>2" "f\<^isub>1[x::=f\<^isub>2] \ e'" - using assms - apply - - apply(erule b_inv_auto) - apply(drule_tac pi="[(xa,x)]" in big.eqvt) - apply(drule_tac pi="[(xa,x)]" in big.eqvt) - apply(drule_tac pi="[(xa,x)]" in big.eqvt) - apply(perm_simp add: calc_atm eqvts) - done + by (simp_all add: abs_fresh) -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" - 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_inv_auto(2)) - 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: eqvts 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: eqvts calc_atm) - apply(assumption) -done +lemma big_preserves_fresh: + fixes x::"name" + assumes a: "e \ e'" "x\e" + shows "x\e'" + using a by (induct) (auto simp add: abs_fresh fresh_subst) -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" - 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 - apply - - apply(rule b_inv_auto(2)) - 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: eqvts 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: eqvts calc_atm) - apply(assumption) -done - -inductive - val :: "trm\bool" -where - v_Lam[intro]: "val (Lam [x].e)" -| v_Const[intro]: "val (Const n)" -| v_Pr[intro]: "\val e\<^isub>1; val e\<^isub>2\ \ val (Pr e\<^isub>1 e\<^isub>2)" -| 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] - -inductive_cases v_inv_auto[elim]: - "val (Const n)" - "val (Pr e\<^isub>1 e\<^isub>2)" - "val (InL e)" - "val (InR e)" - "val (Fst e)" - "val (Snd e)" - "val (Case e of inl x\<^isub>1 \ e\<^isub>1 | inr x\<^isub>2 \ e\<^isub>2)" - "val (Var x)" - "val (Lam [x].e)" - "val (App e\<^isub>1 e\<^isub>2)" - -declare trm.inject [simp del] -declare ty.inject [simp del] -declare data.inject [simp del] +lemma b_App_elim: + assumes a: "App e\<^isub>1 e\<^isub>2 \ e'" "x\(e\<^isub>1,e\<^isub>2,e')" + obtains f\<^isub>1 and f\<^isub>2 where "e\<^isub>1 \ Lam [x]. f\<^isub>1" "e\<^isub>2 \ f\<^isub>2" "f\<^isub>1[x::=f\<^isub>2] \ e'" + using a +by (cases rule: big.strong_cases[where x="x" and xa="x"]) + (auto simp add: trm.inject) lemma subject_reduction: - assumes a: "e \ e'" - and b: "\ \ e : T" + assumes a: "e \ e'" and b: "\ \ e : T" shows "\ \ e' : T" using a b proof (nominal_induct avoiding: \ arbitrary: T rule: big.strong_induct) case (b_App x e\<^isub>1 e\<^isub>2 e' e e\<^isub>2' \ T) have vc: "x\\" by fact have "\ \ App e\<^isub>1 e\<^isub>2 : T" by fact - then obtain T' where - a1: "\ \ e\<^isub>1 : T'\T" and - a2: "\ \ e\<^isub>2 : T'" by auto + then obtain T' where a1: "\ \ e\<^isub>1 : T'\T" and a2: "\ \ e\<^isub>2 : T'" + by (cases) (auto simp add: trm.inject) 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 - then have "((x,T')#\) \ e : T" using vc by (auto simp add: ty.inject) + then have "((x,T')#\) \ e : T" using vc + by (auto elim: t_Lam_elim simp add: ty.inject) moreover have "\ \ e\<^isub>2': T'" using ih2 a2 by simp ultimately have "\ \ e[x::=e\<^isub>2'] : T" by (simp add: typing_substitution) thus "\ \ e' : T" using ih3 by simp -next - case (b_CaseL x\<^isub>1 e e\<^isub>2 e'' x\<^isub>2 e\<^isub>1 e' \) - have vc: "x\<^isub>1\\" "x\<^isub>2\\" by fact+ - have "\ \ Case e of inl x\<^isub>1 \ e\<^isub>1 | inr x\<^isub>2 \ e\<^isub>2 : T" by fact - then obtain S\<^isub>1 S\<^isub>2 e\<^isub>1' e\<^isub>2' where - a1: "\ \ e : Data (DSum S\<^isub>1 S\<^isub>2)" and - a2: "((x\<^isub>1,Data S\<^isub>1)#\) \ e\<^isub>1 : T" using vc by auto - have ih1:"\ \ e : Data (DSum S\<^isub>1 S\<^isub>2) \ \ \ InL e' : Data (DSum S\<^isub>1 S\<^isub>2)" by fact - have ih2:"\ \ e\<^isub>1[x\<^isub>1::=e'] : T \ \ \ e'' : T " by fact - have "\ \ InL e' : Data (DSum S\<^isub>1 S\<^isub>2)" using ih1 a1 by simp - then have "\ \ e' : Data S\<^isub>1" by auto - then have "\ \ e\<^isub>1[x\<^isub>1::=e'] : T" using a2 by (simp add: typing_substitution) - then show "\ \ e'' : T" using ih2 by simp -next - case (b_CaseR x\<^isub>1 e e\<^isub>2 e'' x\<^isub>2 e\<^isub>1 e' \ T) - then show "\ \ e'' : T" by (blast intro: typing_substitution) qed (blast)+ lemma unicity_of_evaluation: @@ -699,93 +389,25 @@ have ih2:"\t. e\<^isub>2 \ t \ e\<^isub>2' = t" by fact have ih3: "\t. e\<^isub>1'[x::=e\<^isub>2'] \ t \ e' = t" 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 - 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) + have vc: "x\e\<^isub>1" "x\e\<^isub>2" "x\t\<^isub>2" by fact + then have "x\App e\<^isub>1 e\<^isub>2" by auto + from app 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" + by (cases rule: big.strong_cases[where x="x" and xa="x"]) + (auto simp add: trm.inject) 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 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)" 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 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 -next - case (b_CaseR 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 \ 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 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 -next - case (b_Fst e e\<^isub>1 e\<^isub>2 e\<^isub>2') - have "e \ Pr e\<^isub>1 e\<^isub>2" by fact - have "\ b. e \ b \ Pr e\<^isub>1 e\<^isub>2 = b" by fact - have "Fst e \ e\<^isub>2'" by fact - show "e\<^isub>1 = e\<^isub>2'" using prems by (force simp add: trm.inject) -next - case (b_Snd e e\<^isub>1 e\<^isub>2 e\<^isub>2') - have "e \ Pr e\<^isub>1 e\<^isub>2" by fact - have "\ b. e \ b \ Pr e\<^isub>1 e\<^isub>2 = b" by fact - have "Snd e \ e\<^isub>2'" by fact - show "e\<^isub>2 = e\<^isub>2'" using prems by (force simp add: trm.inject) -qed (blast)+ - -lemma not_val_App[simp]: - shows - "\ val (App e\<^isub>1 e\<^isub>2)" - "\ val (Fst e)" - "\ val (Snd e)" - "\ val (Var x)" - "\ val (Case e of inl x\<^isub>1 \ e\<^isub>1 | inr x\<^isub>2 \ e\<^isub>2)" -by auto + thus "e' = t\<^isub>2" using ih3 by simp +qed lemma reduces_evaluates_to_values: assumes h:"t \ t'" shows "val t'" using h by (induct) (auto) -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_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_evaluates_to_ins: - assumes "\ \ t : Data (DSum \\<^isub>1 \\<^isub>2)" and "t \ 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_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_evaluates_to_lams: assumes "\ \ t : \ \ \" and "t \ t'" obtains x t'' where "t' = Lam [x]. t''" @@ -793,55 +415,20 @@ have "\ \ t' : \ \ \" using assms subject_reduction by simp moreover 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_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_evaluates_to_values assms by simp - ultimately obtain n where "t' = Const n" by (cases, auto simp add:ty.inject data.inject) + ultimately obtain x t'' where "t' = Lam [x]. t''" by (cases, auto simp add:ty.inject) thus ?thesis using prems by auto qed -consts - V' :: "data \ trm set" - -nominal_primrec - "V' (DNat) = {Const n | n. n \ (UNIV::nat set)}" - "V' (DProd S\<^isub>1 S\<^isub>2) = {Pr x y | x y. x \ V' S\<^isub>1 \ y \ V' S\<^isub>2}" - "V' (DSum S\<^isub>1 S\<^isub>2) = {InL x | x. x \ V' S\<^isub>1} \ {InR y | y. y \ V' S\<^isub>2}" -apply(rule TrueI)+ -done - -lemma Vprimes_are_values : - fixes S::"data" - assumes h: "e \ V' S" - shows "val e" -using h -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) - +(* Valuation *) consts V :: "ty \ trm set" nominal_primrec - "V (Data S) = V' S" + "V (TVar x) = {e. val e}" "V (T\<^isub>1 \ T\<^isub>2) = {Lam [x].e | x e. \ v \ (V T\<^isub>1). \ v'. e[x::=v] \ v' \ v' \ V T\<^isub>2}" -apply(rule TrueI)+ -done + by (rule TrueI)+ +(* can go with strong inversion rules *) lemma V_eqvt: fixes pi::"name prm" assumes a: "x\V T" @@ -849,7 +436,7 @@ 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(simp add: eqvts) apply(rule_tac x="pi\xa" in exI) apply(rule_tac x="pi\e" in exI) apply(simp) @@ -863,12 +450,12 @@ apply(perm_simp add: eqvts) done -lemma V_arrow_elim_weak[elim] : +lemma V_arrow_elim_weak: 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" using h by (auto) -lemma V_arrow_elim_strong[elim]: +lemma V_arrow_elim_strong: 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" @@ -900,36 +487,31 @@ apply(simp add: fin_supp) done -lemma V_are_values : - fixes T::"ty" - assumes h:"e \ V T" +lemma Vs_are_values: + assumes a: "e \ V T" shows "val e" -using h by (nominal_induct T arbitrary: e rule:ty.induct, auto simp add: Vprimes_are_values) +using a by (nominal_induct T arbitrary: e rule: ty.induct) (auto) lemma values_reduce_to_themselves: - assumes h:"val v" + assumes a: "val v" shows "v \ v" -using h by (induct,auto) - -lemma Vs_reduce_to_themselves[simp]: - assumes h:"v \ V T" - shows "v \ v" -using h by (simp add: values_reduce_to_themselves V_are_values) +using a by (induct) (auto) -lemma V_sum: - assumes h:"x \ V (Data (DSum S\<^isub>1 S\<^isub>2))" - shows "(\ y. x= InL y \ y \ V' S\<^isub>1) \ (\ y. x= InR y \ y \ V' S\<^isub>2)" -using h by simp +lemma Vs_reduce_to_themselves: + assumes a: "v \ V T" + shows "v \ v" +using a by (simp add: values_reduce_to_themselves Vs_are_values) +text {* '\ maps x to e' asserts that \ substitutes x with e *} abbreviation - mapsto :: "(name\trm) list \ name \ trm \ bool" ("_ maps _ to _" [55,55,55] 55) + mapsto :: "(name\trm) list \ name \ trm \ bool" ("_ maps _ to _" [55,55,55] 55) where - "\ maps x to e\ (lookup \ x) = e" + "\ maps x to e \ (lookup \ x) = e" abbreviation v_closes :: "(name\trm) list \ (name\ty) list \ bool" ("_ Vcloses _" [55,55] 55) where - "\ Vcloses \ \ \x T. ((x,T) \ set \ \ (\v. \ maps x to v \ v \ (V T)))" + "\ Vcloses \ \ \x T. (x,T) \ set \ \ (\v. \ maps x to v \ v \ V T)" lemma monotonicity: fixes m::"name" @@ -958,10 +540,6 @@ qed lemma termination_aux: - fixes T :: "ty" - fixes \ :: "(name \ ty) list" - fixes \ :: "(name \ trm) list" - fixes e :: "trm" assumes h1: "\ \ e : T" and h2: "\ Vcloses \" shows "\v. \ \ v \ v \ V T" @@ -972,39 +550,33 @@ have ih\<^isub>2:"\\ \ T. \\ Vcloses \; \ \ e\<^isub>2 : T\ \ \v. \2> \ v \ v \ V T" by fact have as\<^isub>1:"\ Vcloses \" by fact have as\<^isub>2: "\ \ App e\<^isub>1 e\<^isub>2 : T" by fact - from as\<^isub>2 obtain T' where "\ \ e\<^isub>1 : T' \ T" and "\ \ e\<^isub>2 : T'" by auto + then obtain T' where "\ \ e\<^isub>1 : T' \ T" and "\ \ e\<^isub>2 : T'" + by (cases) (auto simp add: trm.inject) then obtain v\<^isub>1 v\<^isub>2 where "(i)": "\1> \ v\<^isub>1" "v\<^isub>1 \ V (T' \ T)" and "(ii)":"\2> \ v\<^isub>2" "v\<^isub>2 \ V T'" using ih\<^isub>1 ih\<^isub>2 as\<^isub>1 by blast from "(i)" obtain x e' where "v\<^isub>1 = Lam[x].e'" and "(iii)": "(\v \ (V T').\ v'. e'[x::=v] \ v' \ v' \ V T)" and "(iv)": "\1> \ Lam [x].e'" - and fr: "x\(\,e\<^isub>1,e\<^isub>2)" by blast + and fr: "x\(\,e\<^isub>1,e\<^isub>2)" by (blast elim: V_arrow_elim_strong) from fr have fr\<^isub>1: "x\\1>" and fr\<^isub>2: "x\\2>" by (simp_all add: fresh_psubst) from "(ii)" "(iii)" obtain v\<^isub>3 where "(v)": "e'[x::=v\<^isub>2] \ v\<^isub>3" "v\<^isub>3 \ V T" by auto - from fr\<^isub>2 "(ii)" have "x\v\<^isub>2" by (simp add: fresh_preserved) - then have "x\e'[x::=v\<^isub>2]" by (simp add: fresh_subst_fresh) - then have fr\<^isub>3: "x\v\<^isub>3" using "(v)" by (simp add: fresh_preserved) + from fr\<^isub>2 "(ii)" have "x\v\<^isub>2" by (simp add: big_preserves_fresh) + then have "x\e'[x::=v\<^isub>2]" by (simp add: fresh_subst') + then have fr\<^isub>3: "x\v\<^isub>3" using "(v)" by (simp add: big_preserves_fresh) from fr\<^isub>1 fr\<^isub>2 fr\<^isub>3 have "x\(\1>,\2>,v\<^isub>3)" by simp with "(iv)" "(ii)" "(v)" have "App (\1>) (\2>) \ v\<^isub>3" by auto then show "\v. \1 e\<^isub>2> \ v \ v \ V T" using "(v)" by auto next - case (Pr t\<^isub>1 t\<^isub>2 \ \ T) - have "\ \ Pr t\<^isub>1 t\<^isub>2 : T" by fact - then obtain T\<^isub>a T\<^isub>b where ta:"\ \ t\<^isub>1 : Data T\<^isub>a" and "\ \ t\<^isub>2 : Data T\<^isub>b" - and eq:"T=Data (DProd T\<^isub>a T\<^isub>b)" by auto - have h:"\ Vcloses \" by fact - 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 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 have as\<^isub>2: "\ \ Lam [x].e : T" by fact - have fs: "x\\" "x\\" by fact+ + 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 + where "(i)": "(x,T\<^isub>1)#\ \ e:T\<^isub>2" and "(ii)": "T = T\<^isub>1 \ T\<^isub>2" using fs + by (cases rule: typing.strong_cases[where x="x"]) + (auto simp add: trm.inject alpha abs_fresh fresh_ty) 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 @@ -1019,93 +591,20 @@ then have "\ \ Lam[x].\ \ Lam[x].\ \ V (T\<^isub>1\T\<^isub>2)" using fs by auto thus "\v. \ \ v \ v \ V T" using "(ii)" by auto next - case (Case t' n\<^isub>1 t\<^isub>1 n\<^isub>2 t\<^isub>2 \ \ T) - have f: "n\<^isub>1\\" "n\<^isub>1\\" "n\<^isub>2\\" "n\<^isub>2\\" "n\<^isub>2\n\<^isub>1" "n\<^isub>1\t'" - "n\<^isub>1\t\<^isub>2" "n\<^isub>2\t'" "n\<^isub>2\t\<^isub>1" by fact+ - have h:"\ Vcloses \" by fact - 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 - 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' - 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_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 - 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 - - have "n\<^isub>1\v\<^isub>0" using ht' fresh_preserved fresh_psubst f by auto - then have "n\<^isub>1\v\<^isub>0'" using eqc by auto - then have "n\<^isub>1\v\<^isub>1" using f r fresh_preserved fresh_subst_fresh by blast - thus "n\<^isub>1\(\,\2>,v\<^isub>1,n\<^isub>2)" using f by (simp add: fresh_atm fresh_psubst) - qed - moreover have "n\<^isub>2\(\,\1>,v\<^isub>1,n\<^isub>1)" - proof - - have "n\<^isub>2\v\<^isub>0" using ht' fresh_preserved fresh_psubst f by auto - then have "n\<^isub>2\v\<^isub>0'" using eqc by auto - then have "n\<^isub>2\((n\<^isub>1,v\<^isub>0')#\)" using f by (simp add: fresh_list_cons fresh_prod fresh_atm) - then have "n\<^isub>2\((n\<^isub>1,v\<^isub>0')#\)1>" using f fresh_psubst by auto - moreover then have "n\<^isub>2 \ v\<^isub>1" using fresh_preserved ho by auto - ultimately show "n\<^isub>2\(\,\1>,v\<^isub>1,n\<^isub>1)" using f by (simp add: fresh_psubst fresh_atm) - qed - ultimately have "Case \ of inl n\<^isub>1 \ \1> | inr n\<^isub>2 \ \2> \ v\<^isub>1 \ v\<^isub>1 \ V T" using ht' eqc by auto - moreover - have "Case \ of inl n\<^isub>1 \ \1> | inr n\<^isub>2 \ \2> = \1 \ t\<^isub>1 | inr n\<^isub>2 \ t\<^isub>2>" - using f by auto - 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' - 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_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 - moreover have ih:"\\ \ T. \\ Vcloses \; \ \ t\<^isub>2 : T\ \ \v. \2> \ v \ v \ V T" by fact - ultimately obtain v\<^isub>2 where ho:"((n\<^isub>2,v\<^isub>0')#\)2> \ v\<^isub>2 \ v\<^isub>2 \ V T" using hr by blast - then have r:"\2>[n\<^isub>2::=v\<^isub>0'] \ v\<^isub>2 \ v\<^isub>2 \ V T" using psubst_subst_psubst f by simp - moreover have "n\<^isub>1\(\,\2>,v\<^isub>2,n\<^isub>2)" - proof - - have "n\<^isub>1\\" using fresh_psubst f by simp - then have "n\<^isub>1\v\<^isub>0" using ht' fresh_preserved by auto - then have "n\<^isub>1\v\<^isub>0'" using eqc by auto - then have "n\<^isub>1\((n\<^isub>2,v\<^isub>0')#\)" using f by (simp add: fresh_list_cons fresh_atm fresh_prod) - then have "n\<^isub>1\((n\<^isub>2,v\<^isub>0')#\)2>" using f fresh_psubst by auto - moreover then have "n\<^isub>1\v\<^isub>2" using fresh_preserved ho by auto - ultimately show "n\<^isub>1 \ (\,\2>,v\<^isub>2,n\<^isub>2)" using f by (simp add: fresh_psubst fresh_atm) - qed - moreover have "n\<^isub>2 \ (\,\1>,v\<^isub>2,n\<^isub>1)" - proof - - have "n\<^isub>2\\" using fresh_psubst f by simp - then have "n\<^isub>2\v\<^isub>0" using ht' fresh_preserved by auto - then have "n\<^isub>2\v\<^isub>0'" using eqc by auto - then have "n\<^isub>2\\2>[n\<^isub>2::=v\<^isub>0']" using f fresh_subst_fresh by auto - then have "n\<^isub>2\v\<^isub>2" using f fresh_preserved r by blast - then show "n\<^isub>2\(\,\1>,v\<^isub>2,n\<^isub>1)" using f by (simp add: fresh_atm fresh_psubst) - 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)+ + case (Var x \ \ T) + have "\ \ (Var x) : T" by fact + then have "(x,T)\set \" by (cases) (auto simp add: trm.inject) + with prems have "\ \ \ \ \\ V T" + by (auto intro!: Vs_reduce_to_themselves) + then show "\v. \ \ v \ v \ V T" by auto +qed theorem termination_of_evaluation: assumes a: "[] \ e : T" shows "\v. e \ v \ val v" proof - - from a have "\v. (([]::(name \ trm) list)) \ v \ v \ V T" - by (rule termination_aux) (auto) - thus "\v. e \ v \ val v" using V_are_values by auto + from a have "\v. [] \ v \ v \ V T" by (rule termination_aux) (auto) + thus "\v. e \ v \ val v" using Vs_are_values by auto qed end