# HG changeset patch # User urbanc # Date 1174061572 -3600 # Node ID 013dbd8234f0a6a1676990016f554f896841a115 # Parent 91951d4177d3e27e726722a00dbf4173f705ca68 added formalisations of typical SOS-proofs diff -r 91951d4177d3 -r 013dbd8234f0 src/HOL/Nominal/Examples/SOS.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Nominal/Examples/SOS.thy Fri Mar 16 17:12:52 2007 +0100 @@ -0,0 +1,1423 @@ +(* "$Id$" *) +(* *) +(* Formalisation of some typical SOS-proofs from a *) +(* challenge suggested by Adam Chlipala. *) +(* *) +(* We thank Nick Benton for hellping us with the *) +(* termination-proof for evaluation . *) + +theory SOS + imports "Nominal" +begin + +atom_decl name + +nominal_datatype data = + DNat + | DProd "data" "data" + | DSum "data" "data" + +nominal_datatype ty = + Data "data" + | 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.induct_weak) (simp_all) + +lemma perm_ty[simp]: + fixes T::"ty" + and pi::"name prm" + shows "pi\T = T" + by (induct T rule: ty.induct_weak) (simp_all) + +lemma fresh_ty[simp]: + fixes x::"name" + and T::"ty" + shows "x\T" + by (simp add: fresh_def supp_def) + +text {* 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: + 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) + +lemma lookup_fresh: + fixes z::"name" + assumes "z\\" and "z\x" + shows "z \lookup \ x" +using assms +by (induct rule: lookup.induct) (auto simp add: fresh_list_cons) + +lemma lookup_fresh': + assumes "z\\" + shows "lookup \ z = Var z" +using assms +by (induct rule: lookup.induct) + (auto simp add: fresh_list_cons fresh_prod fresh_atm) + +consts + psubst :: "(name\trm) list \ trm \ trm" ("_<_>" [95,95] 105) + +nominal_primrec + "\<(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,\)\ + \ \<(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 + +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)+ + +lemma fresh_psubst: + fixes z::"name" + and t::"trm" + assumes "z\t" and "z\\" + shows "z\(\)" +using assms +by (nominal_induct t avoiding: z \ t rule: trm.induct) + (auto simp add: abs_fresh lookup_fresh) + +abbreviation + 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 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" + and t2::"trm" + assumes "z\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: z y t\<^isub>2 rule: trm.induct) + (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) + +lemma forget: + fixes x::"name" + and L::"trm" + assumes "x\L" + shows "L[x::=P] = L" + using assms + 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) + +lemma psubst_subst_psubst: +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) +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 *} + +inductive2 + valid :: "(name \ 'a::pt_name) list \ bool" +where + v_nil[intro]: "valid []" + | v_cons[intro]: "\valid \;x\\\ \ valid ((x,T)#\)" + +nominal_inductive valid + +inductive_cases2 + 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" + +lemma type_unicity_in_context: + assumes asm1: "(x,t\<^isub>2) \ set ((x,t\<^isub>1)#\)" + and asm2: "valid ((x,t\<^isub>1)#\)" + shows "t\<^isub>1=t\<^isub>2" +proof - + from asm2 have "x\\" by (cases, auto) + then have "(x,t\<^isub>2) \ set \" + by (induct \) (auto simp add: fresh_list_cons fresh_prod fresh_atm) + then have "(x,t\<^isub>2) = (x,t\<^isub>1)" using asm1 by auto + then show "t\<^isub>1 = t\<^isub>2" by auto +qed + +lemma case_distinction_on_context: + 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 +qed + +inductive2 + 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_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" + +lemma typing_valid: + assumes "\ \ t : T" + shows "valid \" + using assms + by (induct) (auto) + +lemma typing_eqvt: + fixes pi::"name prm" + assumes a: "\ \ t : T" + shows "(pi\\) \ (pi\t) : T" +using a +apply(induct) +apply(auto simp add: fresh_bij set_eqvt valid_eqvt) +apply(rule t_Var) +apply(drule valid_eqvt) +apply(assumption) +apply(drule in_eqvt) +apply(simp add: set_eqvt) +done + +declare trm.inject [simp add] +declare ty.inject [simp add] +declare data.inject [simp add] + +inductive_cases2 t_Lam_inv_auto[elim]: "\ \ Lam [x].t : T" +inductive_cases2 t_Var_inv_auto[elim]: "\ \ Var x : T" +inductive_cases2 t_App_inv_auto[elim]: "\ \ App x y : T" +inductive_cases2 t_Const_inv_auto[elim]: "\ \ Const n : T" +inductive_cases2 t_Fst_inv_auto[elim]: "\ \ Fst x : T" +inductive_cases2 t_Snd_inv_auto[elim]: "\ \ Snd x : T" +inductive_cases2 t_InL_inv_auto[elim]: "\ \ InL x : T" +inductive_cases2 t_InL_inv_auto'[elim]: "\ \ InL x : Data (DSum T\<^isub>1 T2)" +inductive_cases2 t_InR_inv_auto[elim]: "\ \ InR x : T" +inductive_cases2 t_InR_inv_auto'[elim]: "\ \ InR x : Data (DSum T\<^isub>1 T2)" +inductive_cases2 t_Pr_inv_auto[elim]: "\ \ Pr x y : T" +inductive_cases2 t_Pr_inv_auto'[elim]: "\ \ Pr e\<^isub>1 e\<^isub>2 : Data (DProd \1 \\<^isub>2)" +inductive_cases2 t_Case_inv_auto[elim]: "\ \ 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 typing_induct_strong + [consumes 1, case_names t_Var t_App t_Lam t_Const t_Pr t_Fst t_Snd t_InL t_InR t_Case]: + fixes P::"'a::fs_name \ (name\ty) list \ trm \ ty \bool" + and x :: "'a::fs_name" + assumes a: "\ \ e : T" + and a1: "\\ x T c. \valid \; (x,T)\set \\ \ P c \ (Var x) T" + and a2: "\\ e\<^isub>1 T\<^isub>1 T\<^isub>2 e\<^isub>2 c. \\c. P c \ e\<^isub>1 (T\<^isub>1\T\<^isub>2); \c. P c \ e\<^isub>2 T\<^isub>1\ \ P c \ (App e\<^isub>1 e\<^isub>2) T\<^isub>2" + and a3: "\x \ T\<^isub>1 t T\<^isub>2 c. \x\(\,c); \c. P c ((x,T\<^isub>1)#\) t T\<^isub>2\ \ P c \ (Lam [x].t) (T\<^isub>1\T\<^isub>2)" + and a4: "\\ n c. valid \ \ P c \ (Const n) (Data DNat)" + and a5: "\\ e\<^isub>1 S\<^isub>1 e\<^isub>2 S\<^isub>2 c. \\c. P c \ e\<^isub>1 (Data S\<^isub>1); \c. P c \ e\<^isub>2 (Data S\<^isub>2)\ + \ P c \ (Pr e\<^isub>1 e\<^isub>2) (Data (DProd S\<^isub>1 S\<^isub>2))" + and a6: "\\ e S\<^isub>1 S\<^isub>2 c. \\c. P c \ e (Data (DProd S\<^isub>1 S\<^isub>2))\ \ P c \ (Fst e) (Data S\<^isub>1)" + and a7: "\\ e S\<^isub>1 S\<^isub>2 c. \\c. P c \ e (Data (DProd S\<^isub>1 S\<^isub>2))\ \ P c \ (Snd e) (Data S\<^isub>2)" + and a8: "\\ e S\<^isub>1 S\<^isub>2 c. \\c. P c \ e (Data S\<^isub>1)\ \ P c \ (InL e) (Data (DSum S\<^isub>1 S\<^isub>2))" + and a9: "\\ e S\<^isub>2 S\<^isub>1 c. \\c. P c \ e (Data S\<^isub>2)\ \ P c \ (InR e) (Data (DSum S\<^isub>1 S\<^isub>2))" + and a10:"\x\<^isub>1 \ e e\<^isub>2 x\<^isub>2 e\<^isub>1 S\<^isub>1 S\<^isub>2 T c. + \x\<^isub>1\(\,e,e\<^isub>2,x\<^isub>2,c); x\<^isub>2\(\,e,e\<^isub>1,x\<^isub>1,c); + \c. P c \ e (Data (DSum S\<^isub>1 S\<^isub>2)); + \c. P c ((x\<^isub>1,Data S\<^isub>1)#\) e\<^isub>1 T; + \c. P c ((x\<^isub>2,Data S\<^isub>2)#\) e\<^isub>2 T\ + \ P c \ (Case e of inl x\<^isub>1 \ e\<^isub>1 | inr x\<^isub>2 \ e\<^isub>2) T" + shows "P c \ e T" +proof - + from a have "\(pi::name prm) c. P c (pi\\) (pi\e) T" + proof (induct) + case (t_Var \ x T pi c) + have "valid \" by fact + then have "valid (pi\\)" by (simp only: eqvt) + moreover + have "(x,T)\set \" by fact + then have "pi\(x,T)\pi\(set \)" by (simp only: pt_set_bij[OF pt_name_inst, OF at_name_inst]) + then have "(pi\x,T)\set (pi\\)" by (simp add: eqvt) + ultimately show "P c (pi\\) (pi\(Var x)) T" using a1 by simp + next + case (t_App \ e\<^isub>1 T\<^isub>1 T\<^isub>2 e\<^isub>2 pi c) + thus "P c (pi\\) (pi\(App e\<^isub>1 e\<^isub>2)) T\<^isub>2" using a2 by (simp, blast) + next + case (t_Lam x \ T\<^isub>1 t T\<^isub>2 pi c) + obtain y::"name" where fs: "y\(pi\x,pi\\,pi\t,c)" by (erule exists_fresh[OF fs_name1]) + let ?sw = "[(pi\x,y)]" + let ?pi' = "?sw@pi" + have f0: "x\\" by fact + have f1: "(pi\x)\(pi\\)" using f0 by (simp add: fresh_bij) + have f2: "y\?pi'\\" by (auto simp add: pt_name2 fresh_left calc_atm perm_pi_simp) + have ih1: "\c. P c (?pi'\((x,T\<^isub>1)#\)) (?pi'\t) T\<^isub>2" by fact + then have "P c (?pi'\\) (Lam [y].(?pi'\t)) (T\<^isub>1\T\<^isub>2)" using fs f2 a3 by (simp add: calc_atm) + then have "P c (?sw\pi\\) (?sw\(Lam [(pi\x)].(pi\t))) (T\<^isub>1\T\<^isub>2)" + by (simp del: append_Cons add: calc_atm pt_name2) + moreover have "(?sw\pi\\) = (pi\\)" + by (rule perm_fresh_fresh) (simp_all add: fs f1) + moreover have "(?sw\(Lam [(pi\x)].(pi\t))) = Lam [(pi\x)].(pi\t)" + by (rule perm_fresh_fresh) (simp_all add: fs f1 fresh_atm abs_fresh) + ultimately show "P c (pi\\) (pi\(Lam [x].t)) (T\<^isub>1\T\<^isub>2)" + by simp + next + case (t_Const \ n pi c) + thus "P c (pi\\) (pi\(Const n)) (Data DNat)" using a4 by (simp, blast intro: eqvt) + next + case (t_Pr \ e\<^isub>1 S\<^isub>1 e\<^isub>2 S\<^isub>2 pi c) + thus "P c (pi\\) (pi\(Pr e\<^isub>1 e\<^isub>2)) (Data (DProd S\<^isub>1 S\<^isub>2))" using a5 + by (simp) + next + case (t_Fst \ e S\<^isub>1 S\<^isub>2 pi c) + thus "P c (pi\\) (pi\(Fst e)) (Data S\<^isub>1)" using a6 by (simp, blast) + next + case (t_Snd \ e S\<^isub>1 S\<^isub>2 pi c) + thus "P c (pi\\) (pi\(Snd e)) (Data S\<^isub>2)" using a7 by (simp, blast) + next + case (t_InL \ e S\<^isub>1 S\<^isub>2 pi c) + thus "P c (pi\\) (pi\(InL e)) (Data (DSum S\<^isub>1 S\<^isub>2))" using a8 by (simp) + next + case (t_InR \ e S\<^isub>2 S\<^isub>1 pi c) + thus "P c (pi\\) (pi\(InR e)) (Data (DSum S\<^isub>1 S\<^isub>2))" using a9 by (simp) + next + case (t_Case x\<^isub>1 \ e e\<^isub>2 x\<^isub>2 e\<^isub>1 S\<^isub>1 S\<^isub>2 T pi c) + obtain y\<^isub>1::"name" where fs1: "y\<^isub>1\(pi\x\<^isub>1,pi\x\<^isub>2,pi\e,pi\e\<^isub>1,pi\e\<^isub>2,pi\\,c)" + by (erule exists_fresh[OF fs_name1]) + obtain y\<^isub>2::"name" where fs2: "y\<^isub>2\(pi\x\<^isub>1,pi\x\<^isub>2,pi\e,pi\e\<^isub>1,pi\e\<^isub>2,pi\\,c,y\<^isub>1)" + by (erule exists_fresh[OF fs_name1]) + let ?sw1 = "[(pi\x\<^isub>1,y\<^isub>1)]" + let ?sw2 = "[(pi\x\<^isub>2,y\<^isub>2)]" + let ?pi' = "?sw2@?sw1@pi" + have f01: "x\<^isub>1\(\,e,e\<^isub>2,x\<^isub>2)" by fact + have f11: "(pi\x\<^isub>1)\(pi\\,pi\e,pi\e\<^isub>2,pi\x\<^isub>2)" using f01 by (simp add: fresh_bij) + have f21: "y\<^isub>1\(?pi'\\,?pi'\e,?pi'\e\<^isub>2)" using f01 fs1 fs2 + by (simp add: fresh_atm fresh_prod fresh_left calc_atm pt_name2 perm_pi_simp) + have f02: "x\<^isub>2\(\,e,e\<^isub>1,x\<^isub>1)" by fact + have f12: "(pi\x\<^isub>2)\(pi\\,pi\e,pi\e\<^isub>1,pi\x\<^isub>1)" using f02 by (simp add: fresh_bij) + have f22: "y\<^isub>2\(?pi'\\,?pi'\e,?pi'\e\<^isub>1)" using f02 fs1 fs2 + by (auto simp add: fresh_atm fresh_prod fresh_left calc_atm pt_name2 perm_pi_simp) + have ih1: "\c. P c (?pi'\\) (?pi'\e) (Data (DSum S\<^isub>1 S\<^isub>2))" by fact + moreover + have ih2: "\c. P c (?pi'\((x\<^isub>1,Data S\<^isub>1)#\)) (?pi'\e\<^isub>1) T" by fact + then have "\c. P c ((y\<^isub>1,Data S\<^isub>1)#(?pi'\\)) (?pi'\e\<^isub>1) T" using fs1 fs2 + by (auto simp add: calc_atm fresh_prod fresh_atm) + moreover + have ih3: "\c. P c (?pi'\((x\<^isub>2,Data S\<^isub>2)#\)) (?pi'\e\<^isub>2) T" by fact + then have "\c. P c ((y\<^isub>2,Data S\<^isub>2)#(?pi'\\)) (?pi'\e\<^isub>2) T" using fs1 fs2 f11 f12 + by (simp add: calc_atm fresh_prod fresh_atm) + ultimately have "P c (?pi'\\) (Case (?pi'\e) of inl y\<^isub>1 \ (?pi'\e\<^isub>1) | inr y\<^isub>2 \ (?pi'\e\<^isub>2)) T" + using f21 f22 fs1 fs2 a10 by (force simp add: fresh_atm fresh_prod) + then have "P c (?sw2\?sw1\pi\\) + (?sw2\?sw1\(Case (pi\e) of inl (pi\x\<^isub>1) \ (pi\e\<^isub>1) | inr (pi\x\<^isub>2) \ (pi\e\<^isub>2))) T" + using fs1 fs2 f01 f02 f11 f12 + by (auto simp del: append_Cons simp add: pt_name2 fresh_atm fresh_prod calc_atm) + moreover have "(?sw1\pi\\) = (pi\\)" + by (rule perm_fresh_fresh) (simp_all add: fs1 f11) + moreover have "(?sw2\pi\\) = (pi\\)" + by (rule perm_fresh_fresh) (simp_all add: fs2 f12) + moreover have "?sw1\(Case (pi\e) of inl (pi\x\<^isub>1) \ (pi\e\<^isub>1) | inr (pi\x\<^isub>2) \ (pi\e\<^isub>2)) + = (Case (pi\e) of inl (pi\x\<^isub>1) \ (pi\e\<^isub>1) | inr (pi\x\<^isub>2) \ (pi\e\<^isub>2))" + by (rule perm_fresh_fresh) (simp_all add: fs1 fs2 f11 f12 abs_fresh) + moreover have "?sw2\(Case (pi\e) of inl (pi\x\<^isub>1) \ (pi\e\<^isub>1) | inr (pi\x\<^isub>2) \ (pi\e\<^isub>2)) + = (Case (pi\e) of inl (pi\x\<^isub>1) \ (pi\e\<^isub>1) | inr (pi\x\<^isub>2) \ (pi\e\<^isub>2))" + by (rule perm_fresh_fresh) (simp_all add: fs1 fs2 f11 f12 abs_fresh) + ultimately show "P c (pi\\) (pi\(Case e of inl x\<^isub>1 \ e\<^isub>1 | inr x\<^isub>2 \ e\<^isub>2)) T" + by (simp only:, simp) + qed + then have "P c (([]::name prm)\\) (([]::name prm)\e) T" by blast + then show "P c \ e T" by simp +qed + +lemma t_Lam_elim [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') + have "([(x,c)]\[(x',c)]\((x',T\<^isub>1)#\)) \ ([(x,c)]\[(x',c)]\t') : T\<^isub>2" using b2 + by (simp only: typing_eqvt[simplified perm_ty]) + 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 + +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_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_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_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_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 + +lemma weakening: + 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_induct_strong) + case (t_Lam x \\<^isub>1 T\<^isub>1 t T\<^isub>2 \\<^isub>2) + 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 + 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 +qed (auto) + +lemma context_exchange: + 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) + 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) + 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" +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) +qed + + +lemma typing_substitution: + fixes \::"(name \ ty) list" + assumes "(x,T')#\ \ e : T" + and "\ \ e': T'" + shows "\ \ e[x::=e'] : T" + using assms +proof (nominal_induct e avoiding: \ e' x arbitrary: T rule: trm.induct) + case (Var y \ e' x T) + have h1: "(x,T')#\ \ Var y : T" by fact + have h2: "\ \ e' : T'" by fact + show "\ \ (Var y)[x::=e'] : T" + proof (cases "x=y") + case True + assume as: "x=y" + then have "T=T'" using h1 typing_var_unicity by auto + then show "\ \ (Var y)[x::=e'] : T" using as h2 by simp + next + case False + assume as: "x\y" + 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) + ultimately show "\ \ (Var y)[x::=e'] : T" using as by (simp add: t_Var) + qed +next + case (Lam y t \ e' x T) + 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) + 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) + 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 + 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 +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 : T; \ \ e' : T'\ \ \ \ t\<^isub>1[x::=e'] : T" + 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_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) + 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)+ + +text {* Big-Step Evaluation *} + +inductive2 + big :: "trm\trm\bool" ("_ \ _" [80,80] 80) +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''" + +nominal_inductive big + +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_cases2 b_App_inv_auto[elim]: "App e\<^isub>1 e\<^isub>2 \ t" +inductive_cases2 b_Case_inv_auto[elim]: "Case e of inl x\<^isub>1 \ e\<^isub>1 | inr x\<^isub>2 \ e\<^isub>2 \ t" +inductive_cases2 b_Lam_inv_auto[elim]: "Lam[x].t \ t" +inductive_cases2 b_Const_inv_auto[elim]: "Const n \ t" +inductive_cases2 b_Fst_inv_auto[elim]: "Fst e \ t" +inductive_cases2 b_Snd_inv_auto[elim]: "Snd e \ t" +inductive_cases2 b_InL_inv_auto[elim]: "InL e \ t" +inductive_cases2 b_InR_inv_auto[elim]: "InR e \ t" +inductive_cases2 b_Pr_inv_auto[elim]: "Pr e\<^isub>1 e\<^isub>2 \ t" + +declare trm.inject [simp del] +declare ty.inject [simp del] +declare data.inject [simp del] + +lemma big_induct_strong + [consumes 1, case_names b_Lam b_App b_Const b_Pr b_Fst b_Snd b_InL b_InR b_CaseL b_CaseR]: + fixes P::"'a::fs_name \ trm \ trm \bool" + and x :: "'a::fs_name" + 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')\ + \ 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')\ + \ 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 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'')\ + \ 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'')\ + \ 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 - + from a have "\(pi::name prm) c. P c (pi\t) (pi\t')" + proof (induct) + case (b_Lam x e pi c) + show "P c (pi\(Lam [x].e)) (pi\(Lam [x].e))" using a1 by simp + next + case (b_App x e\<^isub>1 e\<^isub>2 e' e\<^isub>1' e\<^isub>2' pi c) + obtain y::"name" where fs: "y\(pi\x,pi\e\<^isub>1,pi\e\<^isub>2,pi\e\<^isub>2',pi\e',pi\e\<^isub>1,c)" + by (erule exists_fresh[OF fs_name1]) + let ?sw = "[(pi\x,y)]" + let ?pi' = "?sw@pi" + have f0: "x\(e\<^isub>1,e\<^isub>2,e')" by fact + 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 + have ih2: "\c. P c (?pi'\e\<^isub>2) (?pi'\e\<^isub>2')" by fact + moreover + have ih3: "\c. P c (?pi'\(e\<^isub>1'[x::=e\<^isub>2'])) (?pi'\e')" by fact + then have "\c. P c ((?pi'\e\<^isub>1')[y::=(?pi'\e\<^isub>2')]) (?pi'\e')" by (simp add: calc_atm subst_eqvt) + ultimately have "P c (App (?pi'\e\<^isub>1) (?pi'\e\<^isub>2)) (?pi'\e')" using fs f2 + by (auto intro!: a2 simp add: calc_atm) + then have "P c (?sw\(App (pi\e\<^isub>1) (pi\e\<^isub>2))) (?sw\(pi\e'))" + by (simp del: append_Cons add: pt_name2) + moreover have "(?sw\(App (pi\e\<^isub>1) (pi\e\<^isub>2))) = App (pi\e\<^isub>1) (pi\e\<^isub>2)" + by (rule perm_fresh_fresh) (simp_all add: fs f1) + moreover have "(?sw\(pi\e')) = pi\e'" + by (rule perm_fresh_fresh) (simp_all add: fs f1) + ultimately show "P c (pi\(App e\<^isub>1 e\<^isub>2)) (pi\e')" + by simp + next + case (b_Const n pi c) + 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) + 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 + 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 + 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) + 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) + 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]) + 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]) + let ?sw1 = "[(pi\x\<^isub>1,y\<^isub>1)]" + let ?sw2 = "[(pi\x\<^isub>2,y\<^isub>2)]" + let ?pi' = "?sw2@?sw1@pi" + have f01: "x\<^isub>1\(e,e\<^isub>2,e'',x\<^isub>2)" by fact + have f11: "(pi\x\<^isub>1)\(pi\e,pi\e\<^isub>2,pi\e'',pi\x\<^isub>2)" using f01 by (simp add: fresh_bij) + have f21: "y\<^isub>1\(?pi'\e,?pi'\e\<^isub>2,?pi'\e'')" using f01 fs1 fs2 + by (simp add: fresh_atm fresh_prod fresh_left calc_atm pt_name2 perm_pi_simp) + have f02: "x\<^isub>2\(e,e\<^isub>1,e'',x\<^isub>1)" by fact + 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 + then have "\c. P c ((?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) + ultimately have "P c (Case (?pi'\e) of inl y\<^isub>1 \ (?pi'\e\<^isub>1) | inr y\<^isub>2 \ (?pi'\e\<^isub>2)) (?pi'\e'')" + using f21 f22 fs1 fs2 by (auto intro!: a9 simp add: fresh_atm fresh_prod) + then have "P c (?sw2\?sw1\(Case (pi\e) of inl (pi\x\<^isub>1) \ (pi\e\<^isub>1) | inr (pi\x\<^isub>2) \ (pi\e\<^isub>2))) + (?sw2\?sw1\(pi\e''))" + using fs1 fs2 f01 f02 f11 f12 + by (auto simp del: append_Cons simp add: pt_name2 fresh_atm fresh_prod calc_atm) + moreover have "?sw1\(Case (pi\e) of inl (pi\x\<^isub>1) \ (pi\e\<^isub>1) | inr (pi\x\<^isub>2) \ (pi\e\<^isub>2)) + = (Case (pi\e) of inl (pi\x\<^isub>1) \ (pi\e\<^isub>1) | inr (pi\x\<^isub>2) \ (pi\e\<^isub>2))" + by (rule perm_fresh_fresh) (simp_all add: fs1 fs2 f11 f12 abs_fresh) + moreover have "?sw2\(Case (pi\e) of inl (pi\x\<^isub>1) \ (pi\e\<^isub>1) | inr (pi\x\<^isub>2) \ (pi\e\<^isub>2)) + = (Case (pi\e) of inl (pi\x\<^isub>1) \ (pi\e\<^isub>1) | inr (pi\x\<^isub>2) \ (pi\e\<^isub>2))" + by (rule perm_fresh_fresh) (simp_all add: fs1 fs2 f11 f12 abs_fresh) + moreover have "(?sw1\(pi\e'')) = (pi\e'')" + by (rule perm_fresh_fresh) (simp_all add: fs1 fs2 f11 f12) + moreover have "(?sw2\(pi\e'')) = (pi\e'')" + by (rule perm_fresh_fresh) (simp_all add: fs1 fs2 f11 f12) + ultimately show "P c (pi\(Case e of inl x\<^isub>1 \ e\<^isub>1 | inr x\<^isub>2 \ e\<^isub>2)) (pi\e'')" + by (simp only:, simp) + 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]) + 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]) + let ?sw1 = "[(pi\x\<^isub>1,y\<^isub>1)]" + let ?sw2 = "[(pi\x\<^isub>2,y\<^isub>2)]" + let ?pi' = "?sw2@?sw1@pi" + have f01: "x\<^isub>1\(e,e\<^isub>2,e'',x\<^isub>2)" by fact + have f11: "(pi\x\<^isub>1)\(pi\e,pi\e\<^isub>2,pi\e'',pi\x\<^isub>2)" using f01 by (simp add: fresh_bij) + have f21: "y\<^isub>1\(?pi'\e,?pi'\e\<^isub>2,?pi'\e'')" using f01 fs1 fs2 + by (simp add: fresh_atm fresh_prod fresh_left calc_atm pt_name2 perm_pi_simp) + have f02: "x\<^isub>2\(e,e\<^isub>1,e'',x\<^isub>1)" by fact + 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 + then have "\c. P c ((?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) + ultimately have "P c (Case (?pi'\e) of inl y\<^isub>1 \ (?pi'\e\<^isub>1) | inr y\<^isub>2 \ (?pi'\e\<^isub>2)) (?pi'\e'')" + using f21 f22 fs1 fs2 by (auto intro!: a10 simp add: fresh_atm fresh_prod) + then have "P c (?sw2\?sw1\(Case (pi\e) of inl (pi\x\<^isub>1) \ (pi\e\<^isub>1) | inr (pi\x\<^isub>2) \ (pi\e\<^isub>2))) + (?sw2\?sw1\(pi\e''))" + using fs1 fs2 f01 f02 f11 f12 + by (auto simp del: append_Cons simp add: pt_name2 fresh_atm fresh_prod calc_atm) + moreover have "?sw1\(Case (pi\e) of inl (pi\x\<^isub>1) \ (pi\e\<^isub>1) | inr (pi\x\<^isub>2) \ (pi\e\<^isub>2)) + = (Case (pi\e) of inl (pi\x\<^isub>1) \ (pi\e\<^isub>1) | inr (pi\x\<^isub>2) \ (pi\e\<^isub>2))" + by (rule perm_fresh_fresh) (simp_all add: fs1 fs2 f11 f12 abs_fresh) + moreover have "?sw2\(Case (pi\e) of inl (pi\x\<^isub>1) \ (pi\e\<^isub>1) | inr (pi\x\<^isub>2) \ (pi\e\<^isub>2)) + = (Case (pi\e) of inl (pi\x\<^isub>1) \ (pi\e\<^isub>1) | inr (pi\x\<^isub>2) \ (pi\e\<^isub>2))" + by (rule perm_fresh_fresh) (simp_all add: fs1 fs2 f11 f12 abs_fresh) + moreover have "(?sw1\(pi\e'')) = (pi\e'')" + by (rule perm_fresh_fresh) (simp_all add: fs1 fs2 f11 f12) + moreover have "(?sw2\(pi\e'')) = (pi\e'')" + by (rule perm_fresh_fresh) (simp_all add: fs1 fs2 f11 f12) + ultimately show "P c (pi\(Case e of inl x\<^isub>1 \ e\<^isub>1 | inr x\<^isub>2 \ e\<^isub>2)) (pi\e'')" + by (simp only:, simp) + qed + then have "P c (([]::name prm)\t) (([]::name prm)\t')" by blast + then show "P c t t'" by simp +qed + +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_App_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 eqvt) + 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)" + 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 + +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" + obtains e' where "e \ InR e'" and "e\<^isub>2[x\<^isub>2::=e'] \ e''" + using assms + apply - + apply (rule b_Case_inv_auto, auto) + apply(drule_tac u="e'" in subst_fun_eq)+ + apply(simp) + done + +inductive2 + 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_cases2 v_Const_inv_auto[elim]: "val (Const n)" +inductive_cases2 v_Pr_inv_auto[elim]: "val (Pr e\<^isub>1 e\<^isub>2)" +inductive_cases2 v_InL_inv_auto[elim]: "val (InL e)" +inductive_cases2 v_InR_inv_auto[elim]: "val (InR e)" +inductive_cases2 v_Fst_inv_auto[elim]: "val (Fst e)" +inductive_cases2 v_Snd_inv_auto[elim]: "val (Snd e)" +inductive_cases2 v_Case_inv_auto[elim]: "val (Case e of inl x\<^isub>1 \ e\<^isub>1 | inr x\<^isub>2 \ e\<^isub>2)" +inductive_cases2 v_Var_inv_auto[elim]: "val (Var x)" +inductive_cases2 v_Lam_inv_auto[elim]: "val (Lam [x].e)" +inductive_cases2 v_App_inv_auto[elim]: "val (App e\<^isub>1 e\<^isub>2)" + +declare trm.inject [simp del] +declare ty.inject [simp del] +declare data.inject [simp del] + +lemma subject_reduction: + assumes "e \ e'" and "\ \ e : T" + shows "\ \ e' : T" + using assms +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 + 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 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) + 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 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" + shows "e\<^isub>1 = e\<^isub>2" + using assms +proof (induct arbitrary: e\<^isub>2) + 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 + 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 + 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 + 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 + 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 "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 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 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 "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 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)+ + +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 + +lemma reduces_to_value: + assumes h:"t \ t'" + shows "val t'" + using h by (induct, auto) + +lemma type_prod_down_pair: + assumes "\ \ t : Data (DProd S\<^isub>1 S\<^isub>2)" and "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 + 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: + 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_to_value 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: + 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 + 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: + 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 + ultimately obtain n where "t' = Const n" by (cases, auto simp add:ty.inject data.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) + +consts + V :: "ty \ trm set" + +nominal_primrec + "V (Data S) = V' S" + "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 + +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" +using h by (auto) + +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" +using h +apply - +apply(erule V_arrow_elim_weak) +apply(subgoal_tac "\a'::name. a'\(a,t,c)") +apply(erule exE) +apply(drule_tac x="a'" in meta_spec) +apply(simp) +apply(drule_tac x="[(a,a')]\t" in meta_spec) +apply(simp add: trm.inject alpha fresh_prod fresh_atm) +apply(perm_simp) +apply(simp add: fresh_left calc_atm) +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 exists_fresh') +apply(simp add: fs_name1) +done + +lemma V_are_values : + fixes T::"ty" + assumes h:"e \ V T" + shows "val e" +using h by (nominal_induct T arbitrary: e rule:ty.induct, auto simp add: Vprimes_are_values) + +lemma values_reduce_to_themselves: + assumes h:"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) + +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 + +abbreviation + mapsto :: "(name\trm) list \ name \ trm \ bool" ("_ maps _ to _" [55,55,55] 55) +where + "\ 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)))" + +lemma monotonicity: + fixes m::"name" + fixes \::"(name \ trm) list" + assumes h1: "\ Vcloses \" + and h2: "e \ V T" + and h3: "valid ((x,T)#\)" + shows "(x,e)#\ Vcloses (x,T)#\" +proof(intro strip) + fix x' T' + assume "(x',T') \ set ((x,T)#\)" + then have "((x',T')=(x,T)) \ ((x',T')\set \ \ x'\x)" using h3 + by (rule_tac case_distinction_on_context) + moreover + { (* first case *) + assume "(x',T') = (x,T)" + then have "\e'. ((x,e)#\) maps x to e' \ e' \ V T'" using h2 by auto + } + moreover + { (* second case *) + assume "(x',T') \ set \" and neq:"x' \ x" + then have "\e'. \ maps x' to e' \ e' \ V T'" using h1 by auto + then have "\e'. ((x,e)#\) maps x' to e' \ e' \ V T'" using neq by auto + } + ultimately show "\e'. ((x,e)#\) maps x' to e' \ e' \ V T'" by blast +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" +using h2 h1 +proof(nominal_induct e avoiding: \ \ arbitrary: T rule: trm.induct) + case (App e\<^isub>1 e\<^isub>2 \ \ T) + have ih\<^isub>1:"\\ \ T. \\ Vcloses \; \ \ e\<^isub>1 : T\ \ \v. \1> \ v \ v \ V T" by fact + 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 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 + 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>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 + 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) + 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 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 + 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_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 fresh_list_cons fresh_atm by force + 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_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 fresh_list_cons fresh_atm by force + 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)+ + +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 +qed + +end