(* "$Id: SOS.thy,v 1.21 2008/05/22 14:34:49 urbanc Exp $" *) (* *) (* Formalisation of some typical SOS-proofs. *) (* *) (* This work was inspired by 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" begin atom_decl name text {* types and terms *} nominal_datatype ty = TVar "nat" | Arrow "ty" "ty" ("_->_" [100,100] 100) nominal_datatype trm = Var "name" | Lam "«name»trm" ("Lam [_]._" [100,100] 100) | App "trm" "trm" lemma fresh_ty: fixes x::"name" and T::"ty" shows "x\<sharp>T" by (induct T rule: ty.induct) (auto simp add: fresh_nat) 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[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: eqvts) lemma lookup_fresh: fixes z::"name" assumes a: "z\<sharp>ϑ" and b: "z\<sharp>x" shows "z \<sharp>lookup ϑ x" using a b by (induct rule: lookup.induct) (auto simp add: fresh_list_cons) lemma lookup_fresh': assumes "z\<sharp>ϑ" shows "lookup ϑ z = Var z" using assms by (induct rule: lookup.induct) (auto simp add: fresh_list_cons fresh_prod fresh_atm) (* parallel substitution *) consts psubst :: "(name×trm) list => trm => trm" ("_<_>" [95,95] 105) nominal_primrec "ϑ<(Var x)> = (lookup ϑ x)" "ϑ<(App e1 e2)> = App (ϑ<e1>) (ϑ<e2>)" "x\<sharp>ϑ ==> ϑ<(Lam [x].e)> = Lam [x].(ϑ<e>)" apply(finite_guess)+ apply(rule TrueI)+ apply(simp add: abs_fresh)+ apply(fresh_guess)+ done lemma psubst_eqvt[eqvt]: fixes pi::"name prm" and t::"trm" shows "pi•(ϑ<t>) = (pi•ϑ)<(pi•t)>" by (nominal_induct t avoiding: ϑ rule: trm.strong_induct) (perm_simp add: fresh_bij lookup_eqvt)+ lemma fresh_psubst: fixes z::"name" and t::"trm" assumes "z\<sharp>t" and "z\<sharp>ϑ" shows "z\<sharp>(ϑ<t>)" using assms by (nominal_induct t avoiding: z ϑ t rule: trm.strong_induct) (auto simp add: abs_fresh lookup_fresh) lemma psubst_empty[simp]: shows "[]<t> = t" by (nominal_induct t rule: trm.strong_induct) (auto simp add: fresh_list_nil) text {* Single substitution *} abbreviation subst :: "trm => name => trm => trm" ("_[_::=_]" [100,100,100] 100) where "t[x::=t'] ≡ ([(x,t')])<t>" lemma subst[simp]: shows "(Var x)[y::=t'] = (if x=y then t' else (Var x))" and "(App t1 t2)[y::=t'] = App (t1[y::=t']) (t2[y::=t'])" and "x\<sharp>(y,t') ==> (Lam [x].t)[y::=t'] = Lam [x].(t[y::=t'])" by (simp_all add: fresh_list_cons fresh_list_nil) lemma fresh_subst: fixes z::"name" and t1::"trm" and t2::"trm" assumes a: "z\<sharp>t1" "z\<sharp>t2" shows "z\<sharp>t1[y::=t2]" using a by (nominal_induct t1 avoiding: z y t2 rule: trm.strong_induct) (auto simp add: abs_fresh fresh_atm) lemma fresh_subst': assumes "x\<sharp>e'" shows "x\<sharp>e[x::=e']" using assms by (nominal_induct e avoiding: x e' rule: trm.strong_induct) (auto simp add: fresh_atm abs_fresh fresh_nat) lemma forget: assumes a: "x\<sharp>e" shows "e[x::=e'] = e" using a by (nominal_induct e avoiding: x e' rule: trm.strong_induct) (auto simp add: fresh_atm abs_fresh) lemma psubst_subst_psubst: assumes h: "x\<sharp>ϑ" shows "ϑ<e>[x::=e'] = ((x,e')#ϑ)<e>" using h by (nominal_induct e avoiding: ϑ x e' rule: trm.strong_induct) (auto simp add: fresh_list_cons fresh_atm forget lookup_fresh lookup_fresh') text {* Typing Judgements *} inductive valid :: "(name×ty) list => bool" where v_nil[intro]: "valid []" | v_cons[intro]: "[|valid Γ;x\<sharp>Γ|] ==> valid ((x,T)#Γ)" equivariance valid inductive_cases valid_cons_inv_auto[elim]: "valid ((x,T)#Γ)" lemma type_unicity_in_context: assumes asm1: "(x,t2) ∈ set ((x,t1)#Γ)" and asm2: "valid ((x,t1)#Γ)" shows "t1=t2" proof - from asm2 have "x\<sharp>Γ" by (cases, auto) then have "(x,t2) ∉ set Γ" by (induct Γ) (auto simp add: fresh_list_cons fresh_prod fresh_atm) then have "(x,t2) = (x,t1)" using asm1 by auto then show "t1 = t2" 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\<sharp>Γ" by (induct Γ) (auto simp add: fresh_list_cons fresh_prod fresh_atm) moreover have "m\<sharp>Γ" using asm1 by auto ultimately have False using eq by auto } ultimately show ?thesis by auto qed text {* Typing Relation *} inductive typing :: "(name×ty) list=>trm=>ty=>bool" ("_ \<turnstile> _ : _" [60,60,60] 60) where t_Var[intro]: "[|valid Γ; (x,T)∈set Γ|] ==> Γ \<turnstile> Var x : T" | t_App[intro]: "[|Γ \<turnstile> e1 : T1->T2; Γ \<turnstile> e2 : T1|] ==> Γ \<turnstile> App e1 e2 : T2" | t_Lam[intro]: "[|x\<sharp>Γ; (x,T1)#Γ \<turnstile> e : T2|] ==> Γ \<turnstile> Lam [x].e : T1->T2" equivariance typing nominal_inductive typing by (simp_all add: abs_fresh fresh_ty) lemma typing_implies_valid: assumes a: "Γ \<turnstile> t : T" shows "valid Γ" using a by (induct) (auto) lemma t_Lam_elim: assumes a: "Γ \<turnstile> Lam [x].t : T" "x\<sharp>Γ" obtains T1 and T2 where "(x,T1)#Γ \<turnstile> t : T2" and "T=T1->T2" using a by (cases rule: typing.strong_cases [where x="x"]) (auto simp add: abs_fresh fresh_ty alpha trm.inject) abbreviation "sub_context" :: "(name×ty) list => (name×ty) list => bool" ("_ ⊆ _" [55,55] 55) where "Γ1 ⊆ Γ2 ≡ ∀x T. (x,T)∈set Γ1 --> (x,T)∈set Γ2" lemma weakening: fixes Γ1 Γ2::"(name×ty) list" assumes "Γ1 \<turnstile> e: T" and "valid Γ2" and "Γ1 ⊆ Γ2" shows "Γ2 \<turnstile> e: T" using assms proof(nominal_induct Γ1 e T avoiding: Γ2 rule: typing.strong_induct) case (t_Lam x Γ1 T1 t T2 Γ2) have vc: "x\<sharp>Γ2" by fact have ih: "[|valid ((x,T1)#Γ2); (x,T1)#Γ1 ⊆ (x,T1)#Γ2|] ==> (x,T1)#Γ2 \<turnstile> t : T2" by fact have "valid Γ2" by fact then have "valid ((x,T1)#Γ2)" using vc by auto moreover have "Γ1 ⊆ Γ2" by fact then have "(x,T1)#Γ1 ⊆ (x,T1)#Γ2" by simp ultimately have "(x,T1)#Γ2 \<turnstile> t : T2" using ih by simp with vc show "Γ2 \<turnstile> Lam [x].t : T1->T2" by auto qed (auto) lemma context_exchange: assumes a: "(x1,T1)#(x2,T2)#Γ \<turnstile> e : T" shows "(x2,T2)#(x1,T1)#Γ \<turnstile> e : T" proof - from a have "valid ((x1,T1)#(x2,T2)#Γ)" by (simp add: typing_implies_valid) then have "x1≠x2" "x1\<sharp>Γ" "x2\<sharp>Γ" "valid Γ" by (auto simp: fresh_list_cons fresh_atm[symmetric]) then have "valid ((x2,T2)#(x1,T1)#Γ)" by (auto simp: fresh_list_cons fresh_atm fresh_ty) moreover have "(x1,T1)#(x2,T2)#Γ ⊆ (x2,T2)#(x1,T1)#Γ" by auto ultimately show "(x2,T2)#(x1,T1)#Γ \<turnstile> e : T" using a by (auto intro: weakening) qed lemma typing_var_unicity: assumes a: "(x,T1)#Γ \<turnstile> Var x : T2" shows "T1=T2" proof - have "(x,T2) ∈ set ((x,T1)#Γ)" using a by (cases) (auto simp add: trm.inject) moreover have "valid ((x,T1)#Γ)" using a by (simp add: typing_implies_valid) ultimately show "T1=T2" by (simp only: type_unicity_in_context) qed lemma typing_substitution: fixes Γ::"(name × ty) list" assumes "(x,T')#Γ \<turnstile> e : T" and "Γ \<turnstile> e': T'" shows "Γ \<turnstile> e[x::=e'] : T" using assms proof (nominal_induct e avoiding: Γ e' x arbitrary: T rule: trm.strong_induct) case (Var y Γ e' x T) have h1: "(x,T')#Γ \<turnstile> Var y : T" by fact have h2: "Γ \<turnstile> e' : T'" by fact show "Γ \<turnstile> (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 "Γ \<turnstile> (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 (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) ultimately show "Γ \<turnstile> (Var y)[x::=e'] : T" using as by (simp add: t_Var) qed next case (Lam y t Γ e' x T) have vc: "y\<sharp>Γ" "y\<sharp>x" "y\<sharp>e'" by fact have pr1: "Γ \<turnstile> e' : T'" by fact have pr2: "(x,T')#Γ \<turnstile> Lam [y].t : T" by fact then obtain T1 T2 where pr2': "(y,T1)#(x,T')#Γ \<turnstile> t : T2" and eq: "T = T1->T2" using vc by (auto elim: t_Lam_elim simp add: fresh_list_cons fresh_ty) then have pr2'':"(x,T')#(y,T1)#Γ \<turnstile> t : T2" by (simp add: context_exchange) have ih: "[|(x,T')#(y,T1)#Γ \<turnstile> t : T2; (y,T1)#Γ \<turnstile> e' : T'|] ==> (y,T1)#Γ \<turnstile> t[x::=e'] : T2" by fact have "valid Γ" using pr1 by (simp add: typing_implies_valid) then have "valid ((y,T1)#Γ)" using vc by auto then have "(y,T1)#Γ \<turnstile> e' : T'" using pr1 by (simp add: weakening) then have "(y,T1)#Γ \<turnstile> t[x::=e'] : T2" using ih pr2'' by simp then have "Γ \<turnstile> Lam [y].(t[x::=e']) : T1->T2" using vc by auto then show "Γ \<turnstile> (Lam [y].t)[x::=e'] : T" using vc eq by simp next case (App e1 e2 Γ e' x T) have "(x,T')#Γ \<turnstile> App e1 e2 : T" by fact then obtain Tn where a1: "(x,T')#Γ \<turnstile> e1 : Tn -> T" and a2: "(x,T')#Γ \<turnstile> e2 : Tn" by (cases) (auto simp add: trm.inject) have a3: "Γ \<turnstile> e' : T'" by fact have ih1: "[|(x,T')#Γ \<turnstile> e1 : Tn->T; Γ \<turnstile> e' : T'|] ==> Γ \<turnstile> e1[x::=e'] : Tn->T" by fact have ih2: "[|(x,T')#Γ \<turnstile> e2 : Tn; Γ \<turnstile> e' : T'|] ==> Γ \<turnstile> e2[x::=e'] : Tn" by fact then show "Γ \<turnstile> (App e1 e2)[x::=e'] : T" 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 e1 e2)" "¬ val (Var x)" by (auto elim: val.cases) text {* Big-Step Evaluation *} inductive big :: "trm=>trm=>bool" ("_ \<Down> _" [80,80] 80) where b_Lam[intro]: "Lam [x].e \<Down> Lam [x].e" | b_App[intro]: "[|x\<sharp>(e1,e2,e'); e1\<Down>Lam [x].e; e2\<Down>e2'; e[x::=e2']\<Down>e'|] ==> App e1 e2 \<Down> e'" equivariance big nominal_inductive big by (simp_all add: abs_fresh) lemma big_preserves_fresh: fixes x::"name" assumes a: "e \<Down> e'" "x\<sharp>e" shows "x\<sharp>e'" using a by (induct) (auto simp add: abs_fresh fresh_subst) lemma b_App_elim: assumes a: "App e1 e2 \<Down> e'" "x\<sharp>(e1,e2,e')" obtains f1 and f2 where "e1 \<Down> Lam [x]. f1" "e2 \<Down> f2" "f1[x::=f2] \<Down> 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 \<Down> e'" and b: "Γ \<turnstile> e : T" shows "Γ \<turnstile> e' : T" using a b proof (nominal_induct avoiding: Γ arbitrary: T rule: big.strong_induct) case (b_App x e1 e2 e' e e2' Γ T) have vc: "x\<sharp>Γ" by fact have "Γ \<turnstile> App e1 e2 : T" by fact then obtain T' where a1: "Γ \<turnstile> e1 : T'->T" and a2: "Γ \<turnstile> e2 : T'" by (cases) (auto simp add: trm.inject) have ih1: "Γ \<turnstile> e1 : T' -> T ==> Γ \<turnstile> Lam [x].e : T' -> T" by fact have ih2: "Γ \<turnstile> e2 : T' ==> Γ \<turnstile> e2' : T'" by fact have ih3: "Γ \<turnstile> e[x::=e2'] : T ==> Γ \<turnstile> e' : T" by fact have "Γ \<turnstile> Lam [x].e : T'->T" using ih1 a1 by simp then have "((x,T')#Γ) \<turnstile> e : T" using vc by (auto elim: t_Lam_elim simp add: ty.inject) moreover have "Γ \<turnstile> e2': T'" using ih2 a2 by simp ultimately have "Γ \<turnstile> e[x::=e2'] : T" by (simp add: typing_substitution) thus "Γ \<turnstile> e' : T" using ih3 by simp qed (blast)+ lemma unicity_of_evaluation: assumes a: "e \<Down> e1" and b: "e \<Down> e2" shows "e1 = e2" using a b proof (nominal_induct e e1 avoiding: e2 rule: big.strong_induct) case (b_Lam x e t2) have "Lam [x].e \<Down> t2" by fact thus "Lam [x].e = t2" by (cases, simp_all add: trm.inject) next case (b_App x e1 e2 e' e1' e2' t2) have ih1: "!!t. e1 \<Down> t ==> Lam [x].e1' = t" by fact have ih2:"!!t. e2 \<Down> t ==> e2' = t" by fact have ih3: "!!t. e1'[x::=e2'] \<Down> t ==> e' = t" by fact have app: "App e1 e2 \<Down> t2" by fact have vc: "x\<sharp>e1" "x\<sharp>e2" "x\<sharp>t2" by fact then have "x\<sharp>App e1 e2" by auto from app vc obtain f1 f2 where x1: "e1 \<Down> Lam [x]. f1" and x2: "e2 \<Down> f2" and x3: "f1[x::=f2] \<Down> t2" by (cases rule: big.strong_cases[where x="x" and xa="x"]) (auto simp add: trm.inject) then have "Lam [x]. f1 = Lam [x]. e1'" using ih1 by simp then have "f1 = e1'" by (auto simp add: trm.inject alpha) moreover have "f2 = e2'" using x2 ih2 by simp ultimately have "e1'[x::=e2'] \<Down> t2" using x3 by simp thus "e' = t2" using ih3 by simp qed lemma reduces_evaluates_to_values: assumes h:"t \<Down> t'" shows "val t'" using h by (induct) (auto) (* Valuation *) consts V :: "ty => trm set" nominal_primrec "V (TVar x) = {e. val e}" "V (T1 -> T2) = {Lam [x].e | x e. ∀ v ∈ (V T1). ∃ v'. e[x::=v] \<Down> v' ∧ v' ∈ V T2}" by (rule TrueI)+ (* can go with strong inversion rules *) lemma V_eqvt: fixes pi::"name prm" assumes a: "x∈V T" shows "(pi•x)∈V T" using a apply(nominal_induct T arbitrary: pi x rule: ty.strong_induct) apply(auto simp add: trm.inject) apply(simp add: eqvts) apply(rule_tac x="pi•xa" in exI) apply(rule_tac x="pi•e" in exI) apply(simp) apply(auto) apply(drule_tac x="(rev pi)•v" in bspec) apply(force) apply(auto) apply(rule_tac x="pi•v'" in exI) apply(auto) apply(drule_tac pi="pi" in big.eqvt) apply(perm_simp add: eqvts) done lemma V_arrow_elim_weak: assumes h:"u ∈ V (T1 -> T2)" obtains a t where "u = Lam[a].t" and "∀ v ∈ (V T1). ∃ v'. t[a::=v] \<Down> v' ∧ v' ∈ V T2" using h by (auto) lemma V_arrow_elim_strong: fixes c::"'a::fs_name" assumes h: "u ∈ V (T1 -> T2)" obtains a t where "a\<sharp>c" "u = Lam[a].t" "∀v ∈ (V T1). ∃ v'. t[a::=v] \<Down> v' ∧ v' ∈ V T2" using h apply - apply(erule V_arrow_elim_weak) apply(subgoal_tac "∃a'::name. a'\<sharp>(a,t,c)") (*A*) apply(erule exE) apply(drule_tac x="a'" in meta_spec) apply(drule_tac x="[(a,a')]•t" in meta_spec) apply(drule meta_mp) apply(simp) apply(drule meta_mp) apply(simp add: trm.inject alpha fresh_left fresh_prod calc_atm fresh_atm) apply(perm_simp) apply(force) apply(drule meta_mp) apply(rule ballI) apply(drule_tac x="[(a,a')]•v" in bspec) apply(simp add: V_eqvt) apply(auto) apply(rule_tac x="[(a,a')]•v'" in exI) apply(auto) apply(drule_tac pi="[(a,a')]" in big.eqvt) apply(perm_simp add: eqvts calc_atm) apply(simp add: V_eqvt) (*A*) apply(rule exists_fresh') apply(simp add: fin_supp) done lemma Vs_are_values: assumes a: "e ∈ V T" shows "val e" using a by (nominal_induct T arbitrary: e rule: ty.strong_induct) (auto) lemma values_reduce_to_themselves: assumes a: "val v" shows "v \<Down> v" using a by (induct) (auto) lemma Vs_reduce_to_themselves: assumes a: "v ∈ V T" shows "v \<Down> 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) 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: assumes h1: "Γ \<turnstile> e : T" and h2: "ϑ Vcloses Γ" shows "∃v. ϑ<e> \<Down> v ∧ v ∈ V T" using h2 h1 proof(nominal_induct e avoiding: Γ ϑ arbitrary: T rule: trm.strong_induct) case (App e1 e2 Γ ϑ T) have ih1:"!!ϑ Γ T. [|ϑ Vcloses Γ; Γ \<turnstile> e1 : T|] ==> ∃v. ϑ<e1> \<Down> v ∧ v ∈ V T" by fact have ih2:"!!ϑ Γ T. [|ϑ Vcloses Γ; Γ \<turnstile> e2 : T|] ==> ∃v. ϑ<e2> \<Down> v ∧ v ∈ V T" by fact have as1:"ϑ Vcloses Γ" by fact have as2: "Γ \<turnstile> App e1 e2 : T" by fact then obtain T' where "Γ \<turnstile> e1 : T' -> T" and "Γ \<turnstile> e2 : T'" by (cases) (auto simp add: trm.inject) then obtain v1 v2 where "(i)": "ϑ<e1> \<Down> v1" "v1 ∈ V (T' -> T)" and "(ii)":"ϑ<e2> \<Down> v2" "v2 ∈ V T'" using ih1 ih2 as1 by blast from "(i)" obtain x e' where "v1 = Lam[x].e'" and "(iii)": "(∀v ∈ (V T').∃ v'. e'[x::=v] \<Down> v' ∧ v' ∈ V T)" and "(iv)": "ϑ<e1> \<Down> Lam [x].e'" and fr: "x\<sharp>(ϑ,e1,e2)" by (blast elim: V_arrow_elim_strong) from fr have fr1: "x\<sharp>ϑ<e1>" and fr2: "x\<sharp>ϑ<e2>" by (simp_all add: fresh_psubst) from "(ii)" "(iii)" obtain v3 where "(v)": "e'[x::=v2] \<Down> v3" "v3 ∈ V T" by auto from fr2 "(ii)" have "x\<sharp>v2" by (simp add: big_preserves_fresh) then have "x\<sharp>e'[x::=v2]" by (simp add: fresh_subst') then have fr3: "x\<sharp>v3" using "(v)" by (simp add: big_preserves_fresh) from fr1 fr2 fr3 have "x\<sharp>(ϑ<e1>,ϑ<e2>,v3)" by simp with "(iv)" "(ii)" "(v)" have "App (ϑ<e1>) (ϑ<e2>) \<Down> v3" by auto then show "∃v. ϑ<App e1 e2> \<Down> v ∧ v ∈ V T" using "(v)" by auto next case (Lam x e Γ ϑ T) have ih:"!!ϑ Γ T. [|ϑ Vcloses Γ; Γ \<turnstile> e : T|] ==> ∃v. ϑ<e> \<Down> v ∧ v ∈ V T" by fact have as1: "ϑ Vcloses Γ" by fact have as2: "Γ \<turnstile> Lam [x].e : T" by fact have fs: "x\<sharp>Γ" "x\<sharp>ϑ" by fact from as2 fs obtain T1 T2 where "(i)": "(x,T1)#Γ \<turnstile> e:T2" and "(ii)": "T = T1 -> T2" 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,T1)#Γ)" by (simp add: typing_implies_valid) have "∀v ∈ (V T1). ∃v'. (ϑ<e>)[x::=v] \<Down> v' ∧ v' ∈ V T2" proof fix v assume "v ∈ (V T1)" with "(iii)" as1 have "(x,v)#ϑ Vcloses (x,T1)#Γ" using monotonicity by auto with ih "(i)" obtain v' where "((x,v)#ϑ)<e> \<Down> v' ∧ v' ∈ V T2" by blast then have "ϑ<e>[x::=v] \<Down> v' ∧ v' ∈ V T2" using fs by (simp add: psubst_subst_psubst) then show "∃v'. ϑ<e>[x::=v] \<Down> v' ∧ v' ∈ V T2" by auto qed then have "Lam[x].ϑ<e> ∈ V (T1 -> T2)" by auto then have "ϑ<Lam [x].e> \<Down> Lam[x].ϑ<e> ∧ Lam[x].ϑ<e> ∈ V (T1->T2)" using fs by auto thus "∃v. ϑ<Lam [x].e> \<Down> v ∧ v ∈ V T" using "(ii)" by auto next case (Var x Γ ϑ T) have "Γ \<turnstile> (Var x) : T" by fact then have "(x,T)∈set Γ" by (cases) (auto simp add: trm.inject) with prems have "ϑ<Var x> \<Down> ϑ<Var x> ∧ ϑ<Var x>∈ V T" by (auto intro!: Vs_reduce_to_themselves) then show "∃v. ϑ<Var x> \<Down> v ∧ v ∈ V T" by auto qed theorem termination_of_evaluation: assumes a: "[] \<turnstile> e : T" shows "∃v. e \<Down> v ∧ val v" proof - from a have "∃v. []<e> \<Down> v ∧ v ∈ V T" by (rule termination_aux) (auto) thus "∃v. e \<Down> v ∧ val v" using Vs_are_values by auto qed end
lemma fresh_ty:
x \<sharp> T
lemma lookup_eqvt:
pi • lookup ϑ X = lookup (pi • ϑ) (pi • X)
lemma lookup_fresh:
[| z \<sharp> ϑ; z \<sharp> x |] ==> z \<sharp> lookup ϑ x
lemma lookup_fresh':
z \<sharp> ϑ ==> lookup ϑ z = Var z
lemma psubst_eqvt:
pi • ϑ<t> = (pi • ϑ)<(pi • t)>
lemma fresh_psubst:
[| z \<sharp> t; z \<sharp> ϑ |] ==> z \<sharp> ϑ<t>
lemma psubst_empty:
[]<t> = t
lemma subst(1):
Var x[y::=t'] = (if x = y then t' else Var x)
and subst(2):
App t1 t2[y::=t'] = App (t1[y::=t']) (t2[y::=t'])
and subst(3):
x \<sharp> (y, t') ==> Lam [x].t[y::=t'] = Lam [x].t[y::=t']
lemma fresh_subst:
[| z \<sharp> t1; z \<sharp> t2 |] ==> z \<sharp> t1[y::=t2]
lemma fresh_subst':
x \<sharp> e' ==> x \<sharp> e[x::=e']
lemma forget:
x \<sharp> e ==> e[x::=e'] = e
lemma psubst_subst_psubst:
x \<sharp> ϑ ==> ϑ<e>[x::=e'] = ((x, e') # ϑ)<e>
lemma type_unicity_in_context:
[| (x, t2) ∈ set ((x, t1) # Γ); valid ((x, t1) # Γ) |] ==> t1 = t2
lemma case_distinction_on_context:
[| valid ((m, t) # Γ); (n, U) ∈ set ((m, T) # Γ) |]
==> (n, U) = (m, T) ∨ (n, U) ∈ set Γ ∧ n ≠ m
lemma typing_implies_valid:
Γ \<turnstile> t : T ==> valid Γ
lemma t_Lam_elim:
[| Γ \<turnstile> Lam [x].t : T; x \<sharp> Γ;
!!T1 T2. [| (x, T1) # Γ \<turnstile> t : T2; T = T1->T2 |] ==> thesis |]
==> thesis
lemma weakening:
[| Γ1 \<turnstile> e : T; valid Γ2; Γ1 ⊆ Γ2 |] ==> Γ2 \<turnstile> e : T
lemma context_exchange:
(x1, T1) # (x2, T2) # Γ \<turnstile> e : T
==> (x2, T2) # (x1, T1) # Γ \<turnstile> e : T
lemma typing_var_unicity:
(x, T1) # Γ \<turnstile> Var x : T2 ==> T1 = T2
lemma typing_substitution:
[| (x, T') # Γ \<turnstile> e : T; Γ \<turnstile> e' : T' |]
==> Γ \<turnstile> e[x::=e'] : T
lemma not_val_App:
¬ val (App e1 e2)
¬ val (Var x)
lemma big_preserves_fresh:
[| e \<Down> e'; x \<sharp> e |] ==> x \<sharp> e'
lemma b_App_elim:
[| App e1 e2 \<Down> e'; x \<sharp> (e1, e2, e');
!!f1 f2.
[| e1 \<Down> Lam [x].f1; e2 \<Down> f2; f1[x::=f2] \<Down> e' |]
==> thesis |]
==> thesis
lemma subject_reduction:
[| e \<Down> e'; Γ \<turnstile> e : T |] ==> Γ \<turnstile> e' : T
lemma unicity_of_evaluation:
[| e \<Down> e1; e \<Down> e2 |] ==> e1 = e2
lemma reduces_evaluates_to_values:
t \<Down> t' ==> val t'
lemma V_eqvt:
x ∈ V T ==> pi • x ∈ V T
lemma V_arrow_elim_weak:
[| u ∈ V (T1->T2);
!!a t. [| u = Lam [a].t; ∀v∈V T1. ∃v'. t[a::=v] \<Down> v' ∧ v' ∈ V T2 |]
==> thesis |]
==> thesis
lemma V_arrow_elim_strong:
[| u ∈ V (T1->T2);
!!a t. [| a \<sharp> c; u = Lam [a].t;
∀v∈V T1. ∃v'. t[a::=v] \<Down> v' ∧ v' ∈ V T2 |]
==> thesis |]
==> thesis
lemma Vs_are_values:
e ∈ V T ==> val e
lemma values_reduce_to_themselves:
val v ==> v \<Down> v
lemma Vs_reduce_to_themselves:
v ∈ V T ==> v \<Down> v
lemma monotonicity:
[| ϑ Vcloses Γ; e ∈ V T; valid ((x, T) # Γ) |] ==> (x, e) # ϑ Vcloses (x, T) # Γ
lemma termination_aux:
[| Γ \<turnstile> e : T; ϑ Vcloses Γ |] ==> ∃v. ϑ<e> \<Down> v ∧ v ∈ V T
theorem termination_of_evaluation:
[] \<turnstile> e : T ==> ∃v. e \<Down> v ∧ val v