Theory Type_Preservation

Up to index of Isabelle/HOL/HOL-Nominal/Examples

theory Type_Preservation
imports Nominal

theory Type_Preservation
imports Nominal
begin


text {*

This theory shows how to prove the type preservation
property for the simply-typed lambda-calculus and
beta-reduction.

*}


atom_decl name

nominal_datatype lam =
Var "name"
| App "lam" "lam"
| Lam "«name»lam" ("Lam [_]._")


text {* Capture-Avoiding Substitution *}

nominal_primrec
subst :: "lam => name => lam => lam" ("_[_::=_]")
where
"(Var x)[y::=s] = (if x=y then s else (Var x))"
| "(App t1 t2)[y::=s] = App (t1[y::=s]) (t2[y::=s])"
| "x\<sharp>(y,s) ==> (Lam [x].t)[y::=s] = Lam [x].(t[y::=s])"

apply(finite_guess)+
apply(rule TrueI)+
apply(simp add: abs_fresh)
apply(fresh_guess)+
done

lemma subst_eqvt[eqvt]:
fixes pi::"name prm"
shows "pi•(t1[x::=t2]) = (pi•t1)[(pi•x)::=(pi•t2)]"

by (nominal_induct t1 avoiding: x t2 rule: lam.strong_induct)
(auto simp add: perm_bij fresh_atm fresh_bij)


lemma fresh_fact:
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: lam.strong_induct)
(auto simp add: abs_fresh fresh_prod fresh_atm)


text {* Types *}

nominal_datatype ty =
TVar "string"
| TArr "ty" "ty" ("_ -> _")


lemma ty_fresh:
fixes x::"name"
and T::"ty"
shows "x\<sharp>T"

by (induct T rule: ty.induct)
(auto simp add: fresh_string)


text {* Typing Contexts *}

types ctx = "(name×ty) list"

abbreviation
"sub_ctx" :: "ctx => ctx => bool" ("_ ⊆ _")
where
1 ⊆ Γ2 ≡ ∀x. x ∈ set Γ1 --> x ∈ set Γ2"


text {* Validity of Typing Contexts *}

inductive
valid :: "(name×ty) list => bool"
where
v1[intro]: "valid []"
| v2[intro]: "[|valid Γ; x\<sharp>Γ|]==> valid ((x,T)#Γ)"


equivariance valid

lemma valid_elim[dest]:
assumes a: "valid ((x,T)#Γ)"
shows "x\<sharp>Γ ∧ valid Γ"

using a by (cases) (auto)

lemma valid_insert:
assumes a: "valid (Δ@[(x,T)]@Γ)"
shows "valid (Δ @ Γ)"

using a
by (induct Δ)
(auto simp add: fresh_list_append fresh_list_cons dest!: 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 :: "ctx => lam => ty => bool" ("_ \<turnstile> _ : _")
where
t_Var[intro]: "[|valid Γ; (x,T)∈set Γ|] ==> Γ \<turnstile> Var x : T"
| t_App[intro]: "[|Γ \<turnstile> t1 : T1->T2; Γ \<turnstile> t2 : T1|] ==> Γ \<turnstile> App t1 t2 : T2"
| t_Lam[intro]: "[|x\<sharp>Γ; (x,T1)#Γ \<turnstile> t : T2|] ==> Γ \<turnstile> Lam [x].t : T1->T2"


equivariance typing
nominal_inductive typing
by (simp_all add: abs_fresh ty_fresh)

lemma t_Lam_inversion[dest]:
assumes ty: "Γ \<turnstile> Lam [x].t : T"
and fc: "x\<sharp>Γ"
shows "∃T1 T2. T = T1 -> T2 ∧ (x,T1)#Γ \<turnstile> t : T2"

using ty fc
by (cases rule: typing.strong_cases)
(auto simp add: alpha lam.inject abs_fresh ty_fresh)


lemma t_App_inversion[dest]:
assumes ty: "Γ \<turnstile> App t1 t2 : T"
shows "∃T'. Γ \<turnstile> t1 : T' -> T ∧ Γ \<turnstile> t2 : T'"

using ty
by (cases) (auto simp add: lam.inject)

lemma weakening:
fixes Γ1 Γ2::"ctx"
assumes a: "Γ1 \<turnstile> t : T"
and b: "valid Γ2"
and c: "Γ1 ⊆ Γ2"
shows "Γ2 \<turnstile> t : T"

using a b c
by (nominal_induct Γ1 t T avoiding: Γ2 rule: typing.strong_induct)
(auto | atomize)+


lemma type_substitution_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: x e' Δ rule: typing.strong_induct)
case (t_Var y T x 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_substitution:
assumes a: "(x,T')#Γ \<turnstile> e : T"
and b: "Γ \<turnstile> e' : T'"
shows "Γ \<turnstile> e[x::=e'] : T"

using a b
by (auto intro: type_substitution_aux[where Δ="[]",simplified])

text {* Beta Reduction *}

inductive
"beta" :: "lam=>lam=>bool" (" _ -->β _")
where
b1[intro]: "t1 -->β t2 ==> App t1 s -->β App t2 s"
| b2[intro]: "s1 -->β s2 ==> App t s1 -->β App t s2"
| b3[intro]: "t1 -->β t2 ==> Lam [x].t1 -->β Lam [x].t2"
| b4[intro]: "x\<sharp>s ==> App (Lam [x].t) s -->β t[x::=s]"


equivariance beta
nominal_inductive beta
by (auto simp add: abs_fresh fresh_fact)


theorem type_preservation:
assumes a: "t -->β t'"
and b: "Γ \<turnstile> t : T"
shows "Γ \<turnstile> t' : T"

using a b
proof (nominal_induct avoiding: Γ T rule: beta.strong_induct)
case (b1 t1 t2 s Γ T)
have "Γ \<turnstile> App t1 s : T" by fact
then obtain T' where a1: "Γ \<turnstile> t1 : T' -> T" and a2: "Γ \<turnstile> s : T'" by auto
have ih: "Γ \<turnstile> t1 : T' -> T ==> Γ \<turnstile> t2 : T' -> T" by fact
with a1 a2 show "Γ \<turnstile> App t2 s : T" by auto
next
case (b2 s1 s2 t Γ T)
have "Γ \<turnstile> App t s1 : T" by fact
then obtain T' where a1: "Γ \<turnstile> t : T' -> T" and a2: "Γ \<turnstile> s1 : T'" by auto
have ih: "Γ \<turnstile> s1 : T' ==> Γ \<turnstile> s2 : T'" by fact
with a1 a2 show "Γ \<turnstile> App t s2 : T" by auto
next
case (b3 t1 t2 x Γ T)
have vc: "x\<sharp>Γ" by fact
have "Γ \<turnstile> Lam [x].t1 : T" by fact
then obtain T1 T2 where a1: "(x,T1)#Γ \<turnstile> t1 : T2" and a2: "T = T1 -> T2" using vc by auto
have ih: "(x,T1)#Γ \<turnstile> t1 : T2 ==> (x,T1)#Γ \<turnstile> t2 : T2" by fact
with a1 a2 show "Γ \<turnstile> Lam [x].t2 : T" using vc by auto
next
case (b4 x s t Γ T)
have vc: "x\<sharp>Γ" by fact
have "Γ \<turnstile> App (Lam [x].t) s : T" by fact
then obtain T' where a1: "Γ \<turnstile> Lam [x].t : T' -> T" and a2: "Γ \<turnstile> s : T'" by auto
from a1 obtain T1 T2 where a3: "(x,T')#Γ \<turnstile> t : T" using vc by (auto simp add: ty.inject)
from a3 a2 show "Γ \<turnstile> t[x::=s] : T" by (simp add: type_substitution)
qed


theorem type_preservation_automatic:
assumes a: "t -->β t'"
and b: "Γ \<turnstile> t : T"
shows "Γ \<turnstile> t' : T"

using a b
by (nominal_induct avoiding: Γ T rule: beta.strong_induct)
(auto dest!: t_Lam_inversion t_App_inversion simp add: ty.inject type_substitution)


end