StlcPropProperties of STLC
In this chapter, we develop the fundamental theory of the Simply
Typed Lambda Calculus -- in particular, the type safety
theorem.
Canonical Forms
Lemma canonical_forms_bool : ∀ t,
empty |-- t \in Bool →
value t →
(t = <{true}>) ∨ (t = <{false}>).
Lemma canonical_forms_fun : ∀ t T1 T2,
empty |-- t \in (T1 → T2) →
value t →
∃ x u, t = <{\x:T1, u}>.
empty |-- t \in Bool →
value t →
(t = <{true}>) ∨ (t = <{false}>).
Proof.
intros t HT HVal.
destruct HVal; auto.
inversion HT.
Qed.
intros t HT HVal.
destruct HVal; auto.
inversion HT.
Qed.
Lemma canonical_forms_fun : ∀ t T1 T2,
empty |-- t \in (T1 → T2) →
value t →
∃ x u, t = <{\x:T1, u}>.
Proof.
intros t T1 T2 HT HVal.
destruct HVal as [x ? t1| |] ; inversion HT; subst.
∃ x, t1. reflexivity.
Qed.
intros t T1 T2 HT HVal.
destruct HVal as [x ? t1| |] ; inversion HT; subst.
∃ x, t1. reflexivity.
Qed.
Theorem progress : ∀ t T,
empty |-- t \in T →
value t ∨ ∃ t', t --> t'.
empty |-- t \in T →
value t ∨ ∃ t', t --> t'.
Proof with eauto.
intros t T Ht.
remember empty as Gamma.
induction Ht; subst Gamma; auto.
(* auto solves all three cases in which t is a value *)
- (* T_Var *)
(* contradictory: variables cannot be typed in an
empty context *)
discriminate H.
- (* T_App *)
(* t = t1 t2. Proceed by cases on whether t1 is a
value or steps... *)
right. destruct IHHt1...
+ (* t1 is a value *)
destruct IHHt2...
× (* t2 is also a value *)
eapply canonical_forms_fun in Ht1; [|assumption].
destruct Ht1 as [x [t0 H1]]. subst.
∃ (<{ [x:=t2]t0 }>)...
× (* t2 steps *)
destruct H0 as [t2' Hstp]. ∃ (<{t1 t2'}>)...
+ (* t1 steps *)
destruct H as [t1' Hstp]. ∃ (<{t1' t2}>)...
- (* T_If *)
right. destruct IHHt1...
+ (* t1 is a value *)
destruct (canonical_forms_bool t1); subst; eauto.
+ (* t1 also steps *)
destruct H as [t1' Hstp]. ∃ <{if t1' then t2 else t3}>...
Qed.
intros t T Ht.
remember empty as Gamma.
induction Ht; subst Gamma; auto.
(* auto solves all three cases in which t is a value *)
- (* T_Var *)
(* contradictory: variables cannot be typed in an
empty context *)
discriminate H.
- (* T_App *)
(* t = t1 t2. Proceed by cases on whether t1 is a
value or steps... *)
right. destruct IHHt1...
+ (* t1 is a value *)
destruct IHHt2...
× (* t2 is also a value *)
eapply canonical_forms_fun in Ht1; [|assumption].
destruct Ht1 as [x [t0 H1]]. subst.
∃ (<{ [x:=t2]t0 }>)...
× (* t2 steps *)
destruct H0 as [t2' Hstp]. ∃ (<{t1 t2'}>)...
+ (* t1 steps *)
destruct H as [t1' Hstp]. ∃ (<{t1' t2}>)...
- (* T_If *)
right. destruct IHHt1...
+ (* t1 is a value *)
destruct (canonical_forms_bool t1); subst; eauto.
+ (* t1 also steps *)
destruct H as [t1' Hstp]. ∃ <{if t1' then t2 else t3}>...
Qed.
Preservation
- The preservation theorem is proved by induction on a typing
derivation, pretty much as we did in the Types chapter.
- substitution lemma, stating that substituting a (closed,
well-typed) term s for a variable x in a term t
preserves the type of t.
- weakening lemma, showing that typing is preserved under "extensions" to the context Gamma.
To make Coq happy, we need to formalize all this in the opposite order...
The Weakening Lemma
Lemma weakening : ∀ Gamma Gamma' t T,
includedin Gamma Gamma' →
Gamma |-- t \in T →
Gamma' |-- t \in T.
includedin Gamma Gamma' →
Gamma |-- t \in T →
Gamma' |-- t \in T.
Proof.
intros Gamma Gamma' t T H Ht.
generalize dependent Gamma'.
induction Ht; eauto using includedin_update.
Qed.
intros Gamma Gamma' t T H Ht.
generalize dependent Gamma'.
induction Ht; eauto using includedin_update.
Qed.
Lemma weakening_empty : ∀ Gamma t T,
empty |-- t \in T →
Gamma |-- t \in T.
empty |-- t \in T →
Gamma |-- t \in T.
Proof.
intros Gamma t T.
eapply weakening.
discriminate.
Qed.
intros Gamma t T.
eapply weakening.
discriminate.
Qed.
The Substitution Lemma
- Suppose we have a term t with a free variable x, and
suppose we've been able to assign a type T to t under the
assumption that x has some type U.
- Also, suppose that we have some other term v and that we've
shown that v has type U.
- Then we can substitute v for each of the occurrences of x in t and obtain a new term that still has type T.
Lemma substitution_preserves_typing : ∀ Gamma x U t v T,
x ⊢> U ; Gamma |-- t \in T →
empty |-- v \in U →
Gamma |-- [x:=v]t \in T.
x ⊢> U ; Gamma |-- t \in T →
empty |-- v \in U →
Gamma |-- [x:=v]t \in T.
Main Theorem
Theorem preservation : ∀ t t' T,
empty |-- t \in T →
t --> t' →
empty |-- t' \in T.
empty |-- t \in T →
t --> t' →
empty |-- t' \in T.
Proof with eauto.
intros t t' T HT. generalize dependent t'.
remember empty as Gamma.
induction HT;
intros t' HE; subst;
try solve [inversion HE; subst; auto].
- (* T_App *)
inversion HE; subst...
(* Most of the cases are immediate by induction,
and eauto takes care of them *)
+ (* ST_AppAbs *)
apply substitution_preserves_typing with T2...
inversion HT1...
Qed.
intros t t' T HT. generalize dependent t'.
remember empty as Gamma.
induction HT;
intros t' HE; subst;
try solve [inversion HE; subst; auto].
- (* T_App *)
inversion HE; subst...
(* Most of the cases are immediate by induction,
and eauto takes care of them *)
+ (* ST_AppAbs *)
apply substitution_preserves_typing with T2...
inversion HT1...
Qed.