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 ∈ Bool →
value t →
(t = tru) ∨ (t = fls).
Lemma canonical_forms_fun : ∀t T1 T2,
empty ⊢ t ∈ (Arrow T1 T2) →
value t →
∃x u, t = abs x T1 u.
empty ⊢ t ∈ Bool →
value t →
(t = tru) ∨ (t = fls).
Lemma canonical_forms_fun : ∀t T1 T2,
empty ⊢ t ∈ (Arrow T1 T2) →
value t →
∃x u, t = abs x T1 u.
Theorem progress : ∀t T,
empty ⊢ t ∈ T →
value t ∨ ∃t', t --> t'.
empty ⊢ t ∈ T →
value t ∨ ∃t', t --> t'.
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)
term s for a variable x in a term t preserves the type
of t.
- context invariance lemma, showing that typing is preserved
under "inessential changes" to the context Gamma — in
particular, changes that do not affect any of the free
variables of the term.
- the free variables of a term — i.e., those variables mentioned in a term and not in the scope of an enclosing function abstraction binding a variable of the same name.
Free Occurrences
- y appears free, but x does not, in \x:T→U. x y
- both x and y appear free in (\x:T→U. x y) x
- no variables appear free in \x:T→U. \y:T. x y
Inductive appears_free_in : string → tm → Prop :=
| afi_var : ∀x,
appears_free_in x (var x)
| afi_app1 : ∀x t1 t2,
appears_free_in x t1 →
appears_free_in x (app t1 t2)
| afi_app2 : ∀x t1 t2,
appears_free_in x t2 →
appears_free_in x (app t1 t2)
| afi_abs : ∀x y T11 t12,
y ≠ x →
appears_free_in x t12 →
appears_free_in x (abs y T11 t12)
| afi_test1 : ∀x t1 t2 t3,
appears_free_in x t1 →
appears_free_in x (test t1 t2 t3)
| afi_test2 : ∀x t1 t2 t3,
appears_free_in x t2 →
appears_free_in x (test t1 t2 t3)
| afi_test3 : ∀x t1 t2 t3,
appears_free_in x t3 →
appears_free_in x (test t1 t2 t3).
Hint Constructors appears_free_in.
| afi_var : ∀x,
appears_free_in x (var x)
| afi_app1 : ∀x t1 t2,
appears_free_in x t1 →
appears_free_in x (app t1 t2)
| afi_app2 : ∀x t1 t2,
appears_free_in x t2 →
appears_free_in x (app t1 t2)
| afi_abs : ∀x y T11 t12,
y ≠ x →
appears_free_in x t12 →
appears_free_in x (abs y T11 t12)
| afi_test1 : ∀x t1 t2 t3,
appears_free_in x t1 →
appears_free_in x (test t1 t2 t3)
| afi_test2 : ∀x t1 t2 t3,
appears_free_in x t2 →
appears_free_in x (test t1 t2 t3)
| afi_test3 : ∀x t1 t2 t3,
appears_free_in x t3 →
appears_free_in x (test t1 t2 t3).
Hint Constructors appears_free_in.
The free variables of a term are just the variables that appear free in it. A term with no free variables is said to be closed.
Definition closed (t:tm) :=
∀x, ¬appears_free_in x t.
∀x, ¬appears_free_in x t.
Substitution
Lemma free_in_context : ∀x t T Gamma,
appears_free_in x t →
Gamma ⊢ t ∈ T →
∃T', Gamma x = Some T'.
appears_free_in x t →
Gamma ⊢ t ∈ T →
∃T', Gamma x = Some T'.
From the free_in_context lemma, it immediately follows that any term t that is well typed in the empty context is closed (it has no free variables).
Corollary typable_empty__closed : ∀t T,
empty ⊢ t ∈ T →
closed t.
empty ⊢ t ∈ T →
closed t.
Sometimes, when we have a proof of some typing relation Gamma ⊢ t ∈ T, we will need to replace Gamma by a different context Gamma'.
Lemma context_invariance : ∀Gamma Gamma' t T,
Gamma ⊢ t ∈ T →
(∀x, appears_free_in x t → Gamma x = Gamma' x) →
Gamma' ⊢ t ∈ T.
Gamma ⊢ t ∈ T →
(∀x, appears_free_in x t → Gamma x = Gamma' x) →
Gamma' ⊢ t ∈ T.
Now we come to the conceptual heart of the proof that reduction preserves types — namely, the observation that substitution preserves types.
- 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 ∈ T →
empty ⊢ v ∈ U →
Gamma ⊢ [x:=v]t ∈ T.
(x ⊢> U ; Gamma) ⊢ t ∈ T →
empty ⊢ v ∈ U →
Gamma ⊢ [x:=v]t ∈ T.
Main Theorem
Theorem preservation : ∀t t' T,
empty ⊢ t ∈ T →
t --> t' →
empty ⊢ t' ∈ T.
empty ⊢ t ∈ T →
t --> t' →
empty ⊢ t' ∈ T.