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"
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)
nominal_primrec
psubst :: "(name×trm) list => trm => trm" ("_<_>" [95,95] 105)
where
"ϑ<(Var x)> = (lookup ϑ x)"
| "ϑ<(App e\<^isub>1 e\<^isub>2)> = App (ϑ<e\<^isub>1>) (ϑ<e\<^isub>2>)"
| "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 t\<^isub>1 t\<^isub>2)[y::=t'] = App (t\<^isub>1[y::=t']) (t\<^isub>2[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"
shows "[|z\<sharp>s; (z=y ∨ z\<sharp>t)|] ==> z\<sharp>t[y::=s]"
by (nominal_induct t avoiding: z y s rule: trm.strong_induct)
(auto simp add: abs_fresh fresh_prod fresh_atm)
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_elim[elim]: "valid ((x,T)#Γ)"
lemma valid_insert:
assumes a: "valid (Δ@[(x,T)]@Γ)"
shows "valid (Δ @ Γ)"
using a
by (induct Δ)
(auto simp add: fresh_list_append fresh_list_cons elim!: valid_elim)
lemma fresh_set:
shows "y\<sharp>xs = (∀x∈set xs. y\<sharp>x)"
by (induct xs) (simp_all add: fresh_list_nil fresh_list_cons)
lemma context_unique:
assumes a1: "valid Γ"
and a2: "(x,T) ∈ set Γ"
and a3: "(x,U) ∈ set Γ"
shows "T = U"
using a1 a2 a3
by (induct) (auto simp add: fresh_set fresh_prod fresh_atm)
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> e\<^isub>1 : T\<^isub>1->T\<^isub>2; Γ \<turnstile> e\<^isub>2 : T\<^isub>1|] ==> Γ \<turnstile> App e\<^isub>1 e\<^isub>2 : T\<^isub>2"
| t_Lam[intro]: "[|x\<sharp>Γ; (x,T\<^isub>1)#Γ \<turnstile> e : T\<^isub>2|] ==> Γ \<turnstile> Lam [x].e : T\<^isub>1->T\<^isub>2"
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_App_elim:
assumes a: "Γ \<turnstile> App t1 t2 : T"
obtains T' where "Γ \<turnstile> t1 : T' -> T" and "Γ \<turnstile> t2 : T'"
using a
by (cases) (auto simp add: trm.inject)
lemma t_Lam_elim:
assumes a: "Γ \<turnstile> Lam [x].t : T" "x\<sharp>Γ"
obtains T\<^isub>1 and T\<^isub>2 where "(x,T\<^isub>1)#Γ \<turnstile> t : T\<^isub>2" and "T=T\<^isub>1->T\<^isub>2"
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
"Γ\<^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"
assumes "Γ\<^isub>1 \<turnstile> e: T" and "valid Γ\<^isub>2" and "Γ\<^isub>1 ⊆ Γ\<^isub>2"
shows "Γ\<^isub>2 \<turnstile> 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\<sharp>Γ\<^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 \<turnstile> t : T\<^isub>2" by fact
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 \<turnstile> t : T\<^isub>2" using ih by simp
with vc show "Γ\<^isub>2 \<turnstile> Lam [x].t : T\<^isub>1->T\<^isub>2" by auto
qed (auto)
lemma type_substitutivity_aux:
assumes a: "(Δ@[(x,T')]@Γ) \<turnstile> e : T"
and b: "Γ \<turnstile> e' : T'"
shows "(Δ@Γ) \<turnstile> e[x::=e'] : T"
using a b
proof (nominal_induct Γ≡"Δ@[(x,T')]@Γ" e T avoiding: e' Δ rule: typing.strong_induct)
case (t_Var y T e' Δ)
then have a1: "valid (Δ@[(x,T')]@Γ)"
and a2: "(y,T) ∈ set (Δ@[(x,T')]@Γ)"
and a3: "Γ \<turnstile> e' : T'" .
from a1 have a4: "valid (Δ@Γ)" by (rule valid_insert)
{ assume eq: "x=y"
from a1 a2 have "T=T'" using eq by (auto intro: context_unique)
with a3 have "Δ@Γ \<turnstile> Var y[x::=e'] : T" using eq a4 by (auto intro: weakening)
}
moreover
{ assume ineq: "x≠y"
from a2 have "(y,T) ∈ set (Δ@Γ)" using ineq by simp
then have "Δ@Γ \<turnstile> Var y[x::=e'] : T" using ineq a4 by auto
}
ultimately show "Δ@Γ \<turnstile> Var y[x::=e'] : T" by blast
qed (force simp add: fresh_list_append fresh_list_cons)+
corollary type_substitutivity:
assumes a: "(x,T')#Γ \<turnstile> e : T"
and b: "Γ \<turnstile> e' : T'"
shows "Γ \<turnstile> e[x::=e'] : T"
using a b type_substitutivity_aux[where Δ="[]"]
by (auto)
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 *}
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>(e\<^isub>1,e\<^isub>2,e'); e\<^isub>1\<Down>Lam [x].e; e\<^isub>2\<Down>e\<^isub>2'; e[x::=e\<^isub>2']\<Down>e'|] ==> App e\<^isub>1 e\<^isub>2 \<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 e\<^isub>1 e\<^isub>2 \<Down> e'" "x\<sharp>(e\<^isub>1,e\<^isub>2,e')"
obtains f\<^isub>1 and f\<^isub>2 where "e\<^isub>1 \<Down> Lam [x]. f\<^isub>1" "e\<^isub>2 \<Down> f\<^isub>2" "f\<^isub>1[x::=f\<^isub>2] \<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 e\<^isub>1 e\<^isub>2 e' e e\<^isub>2' Γ T)
have vc: "x\<sharp>Γ" by fact
have "Γ \<turnstile> App e\<^isub>1 e\<^isub>2 : T" by fact
then obtain T' where a1: "Γ \<turnstile> e\<^isub>1 : T'->T" and a2: "Γ \<turnstile> e\<^isub>2 : T'"
by (cases) (auto simp add: trm.inject)
have ih1: "Γ \<turnstile> e\<^isub>1 : T' -> T ==> Γ \<turnstile> Lam [x].e : T' -> T" by fact
have ih2: "Γ \<turnstile> e\<^isub>2 : T' ==> Γ \<turnstile> e\<^isub>2' : T'" by fact
have ih3: "Γ \<turnstile> e[x::=e\<^isub>2'] : 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> e\<^isub>2': T'" using ih2 a2 by simp
ultimately have "Γ \<turnstile> e[x::=e\<^isub>2'] : T" by (simp add: type_substitutivity)
thus "Γ \<turnstile> e' : T" using ih3 by simp
qed (blast)
lemma subject_reduction2:
assumes a: "e \<Down> e'" and b: "Γ \<turnstile> e : T"
shows "Γ \<turnstile> e' : T"
using a b
by (nominal_induct avoiding: Γ T rule: big.strong_induct)
(force elim: t_App_elim t_Lam_elim simp add: ty.inject type_substitutivity)+
lemma unicity_of_evaluation:
assumes a: "e \<Down> e\<^isub>1"
and b: "e \<Down> e\<^isub>2"
shows "e\<^isub>1 = e\<^isub>2"
using a b
proof (nominal_induct e e\<^isub>1 avoiding: e\<^isub>2 rule: big.strong_induct)
case (b_Lam x e t\<^isub>2)
have "Lam [x].e \<Down> t\<^isub>2" by fact
thus "Lam [x].e = t\<^isub>2" by (cases, simp_all 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 ih1: "!!t. e\<^isub>1 \<Down> t ==> Lam [x].e\<^isub>1' = t" by fact
have ih2:"!!t. e\<^isub>2 \<Down> t ==> e\<^isub>2' = t" by fact
have ih3: "!!t. e\<^isub>1'[x::=e\<^isub>2'] \<Down> t ==> e' = t" by fact
have app: "App e\<^isub>1 e\<^isub>2 \<Down> t\<^isub>2" by fact
have vc: "x\<sharp>e\<^isub>1" "x\<sharp>e\<^isub>2" "x\<sharp>t\<^isub>2" by fact+
then have "x\<sharp>App e\<^isub>1 e\<^isub>2" by auto
from app vc obtain f\<^isub>1 f\<^isub>2 where x1: "e\<^isub>1 \<Down> Lam [x]. f\<^isub>1" and x2: "e\<^isub>2 \<Down> f\<^isub>2" and x3: "f\<^isub>1[x::=f\<^isub>2] \<Down> t\<^isub>2"
by (auto elim!: b_App_elim)
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'] \<Down> t\<^isub>2" using x3 by simp
thus "e' = t\<^isub>2" using ih3 by simp
qed
lemma reduces_evaluates_to_values:
assumes h: "t \<Down> t'"
shows "val t'"
using h by (induct) (auto)
nominal_primrec
V :: "ty => trm set"
where
"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] \<Down> v' ∧ v' ∈ V T\<^isub>2}"
by (rule TrueI)+
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 (T\<^isub>1 -> T\<^isub>2)"
obtains a t where "u = Lam [a].t" and "∀ v ∈ (V T\<^isub>1). ∃ v'. t[a::=v] \<Down> v' ∧ v' ∈ V T\<^isub>2"
using h by (auto)
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\<sharp>c" "u = Lam [a].t" "∀v ∈ (V T\<^isub>1). ∃ v'. t[a::=v] \<Down> v' ∧ v' ∈ V T\<^isub>2"
using h
apply -
apply(erule V_arrow_elim_weak)
apply(subgoal_tac "∃a'::name. a'\<sharp>(a,t,c)")
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)
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 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
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
{
assume "(x',T') = (x,T)"
then have "∃e'. ((x,e)#ϑ) maps x to e' ∧ e' ∈ V T'" using h2 by auto
}
moreover
{
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 e\<^isub>1 e\<^isub>2 Γ ϑ T)
have ih\<^isub>1: "!!ϑ Γ T. [|ϑ Vcloses Γ; Γ \<turnstile> e\<^isub>1 : T|] ==> ∃v. ϑ<e\<^isub>1> \<Down> v ∧ v ∈ V T" by fact
have ih\<^isub>2: "!!ϑ Γ T. [|ϑ Vcloses Γ; Γ \<turnstile> e\<^isub>2 : T|] ==> ∃v. ϑ<e\<^isub>2> \<Down> v ∧ v ∈ V T" by fact
have as\<^isub>1: "ϑ Vcloses Γ" by fact
have as\<^isub>2: "Γ \<turnstile> App e\<^isub>1 e\<^isub>2 : T" by fact
then obtain T' where "Γ \<turnstile> e\<^isub>1 : T' -> T" and "Γ \<turnstile> e\<^isub>2 : T'" by (auto elim: t_App_elim)
then obtain v\<^isub>1 v\<^isub>2 where "(i)": "ϑ<e\<^isub>1> \<Down> v\<^isub>1" "v\<^isub>1 ∈ V (T' -> T)"
and "(ii)": "ϑ<e\<^isub>2> \<Down> 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] \<Down> v' ∧ v' ∈ V T)"
and "(iv)": "ϑ<e\<^isub>1> \<Down> Lam [x].e'"
and fr: "x\<sharp>(ϑ,e\<^isub>1,e\<^isub>2)" by (blast elim: V_arrow_elim_strong)
from fr have fr\<^isub>1: "x\<sharp>ϑ<e\<^isub>1>" and fr\<^isub>2: "x\<sharp>ϑ<e\<^isub>2>" by (simp_all add: fresh_psubst)
from "(ii)" "(iii)" obtain v\<^isub>3 where "(v)": "e'[x::=v\<^isub>2] \<Down> v\<^isub>3" "v\<^isub>3 ∈ V T" by auto
from fr\<^isub>2 "(ii)" have "x\<sharp>v\<^isub>2" by (simp add: big_preserves_fresh)
then have "x\<sharp>e'[x::=v\<^isub>2]" by (simp add: fresh_subst)
then have fr\<^isub>3: "x\<sharp>v\<^isub>3" using "(v)" by (simp add: big_preserves_fresh)
from fr\<^isub>1 fr\<^isub>2 fr\<^isub>3 have "x\<sharp>(ϑ<e\<^isub>1>,ϑ<e\<^isub>2>,v\<^isub>3)" by simp
with "(iv)" "(ii)" "(v)" have "App (ϑ<e\<^isub>1>) (ϑ<e\<^isub>2>) \<Down> v\<^isub>3" by auto
then show "∃v. ϑ<App e\<^isub>1 e\<^isub>2> \<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 as\<^isub>1: "ϑ Vcloses Γ" by fact
have as\<^isub>2: "Γ \<turnstile> Lam [x].e : T" by fact
have fs: "x\<sharp>Γ" "x\<sharp>ϑ" by fact+
from as\<^isub>2 fs obtain T\<^isub>1 T\<^isub>2
where "(i)": "(x,T\<^isub>1)#Γ \<turnstile> e:T\<^isub>2" and "(ii)": "T = T\<^isub>1 -> T\<^isub>2" using fs
by (auto elim: t_Lam_elim)
from "(i)" have "(iii)": "valid ((x,T\<^isub>1)#Γ)" by (simp add: typing_implies_valid)
have "∀v ∈ (V T\<^isub>1). ∃v'. (ϑ<e>)[x::=v] \<Down> 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)#ϑ)<e> \<Down> v' ∧ v' ∈ V T\<^isub>2" by blast
then have "ϑ<e>[x::=v] \<Down> v' ∧ v' ∈ V T\<^isub>2" using fs by (simp add: psubst_subst_psubst)
then show "∃v'. ϑ<e>[x::=v] \<Down> v' ∧ v' ∈ V T\<^isub>2" by auto
qed
then have "Lam[x].ϑ<e> ∈ V (T\<^isub>1 -> T\<^isub>2)" by auto
then have "ϑ<Lam [x].e> \<Down> Lam [x].ϑ<e> ∧ Lam [x].ϑ<e> ∈ V (T\<^isub>1->T\<^isub>2)" 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