
NormNormalization of STLC
x
Set Warnings "-notation-overridden,-parsing,-deprecated-hint-without-locality".
From Coq Require Import Lists.List.
From Coq Require Import Strings.String.
From PLF Require Import Maps.
From PLF Require Import Smallstep.
Hint Constructors multi : core.
(* Chapter written and maintained by Andrew Tolmach *)
This optional chapter is based on chapter 12 of Types and
Programming Languages (Pierce). It may be useful to look at the
two together, as that chapter includes explanations and informal
proofs that are not repeated here.
In this chapter, we consider another fundamental theoretical
property of the simply typed lambda-calculus: the fact that the
evaluation of a well-typed program is guaranteed to halt in a
finite number of steps---i.e., every well-typed term is
normalizable.
Unlike the type-safety properties we have considered so far, the
normalization property does not extend to full-blown programming
languages, because these languages nearly always extend the simply
typed lambda-calculus with constructs, such as general
recursion (see the MoreStlc chapter) or recursive types, that
can be used to write nonterminating programs. However, the issue
of normalization reappears at the level of types when we
consider the metatheory of polymorphic versions of the lambda
calculus such as System F-omega: in this system, the language of
types effectively contains a copy of the simply typed
lambda-calculus, and the termination of the typechecking algorithm
will hinge on the fact that a "normalization" operation on type
expressions is guaranteed to terminate.
Another reason for studying normalization proofs is that they are
some of the most beautiful---and mind-blowing---mathematics to be
found in the type theory literature, often (as here) involving the
fundamental proof technique of logical relations.
The calculus we shall consider here is the simply typed
lambda-calculus over a single base type bool and with
pairs. We'll give most details of the development for the basic
lambda-calculus terms treating bool as an uninterpreted base
type, and leave the extension to the boolean operators and pairs
to the reader. Even for the base calculus, normalization is not
entirely trivial to prove, since each reduction of a term can
duplicate redexes in subterms.
Exercise: 2 stars, standard (norm_fail)
Where do we fail if we attempt to prove normalization by a straightforward induction on the size of a well-typed term?xxxxxxxxxx
10
(* FILL IN HERE *)
(* Do not modify the following line: *)
Definition manual_grade_for_norm_fail : option (*string) := None.
Exercise: 5 stars, standard, especially useful (norm)
The best ways to understand an intricate proof like this is are (1) to help fill it in and (2) to extend it. We've left out some parts of the following development, including some proofs of lemmas and the all the cases involving products and conditionals. Fill them in.xxxxxxxxxx
12
(* Do not modify the following line: *)
Definition manual_grade_for_norm : option (*string) := None.
Language
xxxxxxxxxx
49
Inductive ty : Type :=
| Ty_Bool : ty
| Ty_Arrow : ty ty ty
| Ty_Prod : ty ty ty
.
Inductive tm : Type :=
(* pure STLC *)
| tm_var : string tm
| tm_app : tm tm tm
| tm_abs : string ty tm tm
(* booleans *)
| tm_true : tm
| tm_false : tm
| tm_if : tm tm tm tm
(* pairs *)
| tm_pair : tm tm tm
| tm_fst : tm tm
| tm_snd : tm tm.
Declare Custom Entry stlc.
Notation "<{ e }>" := e (e custom stlc at level 99).
Notation "( x )" := x (in custom stlc, x at level 99).
Notation "x" := x (in custom stlc at level 0, x constr at level 0).
Notation "S -> T" := (Ty_Arrow S T) (in custom stlc at level 50, right associativity).
Notation "x y" := (tm_app x y) (in custom stlc at level 1, left associativity).
Notation "\ x : t , y" :=
(tm_abs x t y) (in custom stlc at level 90, x at level 99,
t custom stlc at level 99,
y custom stlc at level 99,
left associativity).
Coercion tm_var : string tm.
Notation "{ x }" := x (in custom stlc at level 1, x constr).
Notation "'Bool'" := Ty_Bool (in custom stlc at level 0).
Notation "'if' x 'then' y 'else' z" :=
(tm_if x y z) (in custom stlc at level 89,
x custom stlc at level 99,
y custom stlc at level 99,
z custom stlc at level 99,
left associativity).
Notation "'true'" := true (at level 1).
Notation "'true'" := tm_true (in custom stlc at level 0).
Notation "'false'" := false (at level 1).
Notation "'false'" := tm_false (in custom stlc at level 0).
Notation "X * Y" :=
(Ty_Prod X Y) (in custom stlc at level 2, X custom stlc, Y custom stlc at level 0).
Notation "( x ',' y )" := (tm_pair x y) (in custom stlc at level 0,
x custom stlc at level 99,
y custom stlc at level 99).
Notation "t '.fst'" := (tm_fst t) (in custom stlc at level 0).
Notation "t '.snd'" := (tm_snd t) (in custom stlc at level 0).
xxxxxxxxxx
24
Reserved Notation "'[' x ':=' s ']' t" (in custom stlc at level 20, x constr).
Fixpoint subst (x : string) (s : tm) (t : tm) : tm :=
match t with
| tm_var y
if String.eqb x y then s else t
| <{ \ y : T, t1 }>
if String.eqb x y then t else <{ \y:T, [x:=s] t1 }>
| <{t1 t2}>
<{ ([x:=s]t1) ([x:=s]t2)}>
| <{true}>
<{true}>
| <{false}>
<{false}>
| <{if t1 then t2 else t3}>
<{if ([x:=s] t1) then ([x:=s] t2) else ([x:=s] t3)}>
| <{(t1, t2)}>
<{( ([x:=s] t1), ([x:=s] t2) )}>
| <{t0.fst}>
<{ ([x:=s] t0).fst}>
| <{t0.snd}>
<{ ([x:=s] t0).snd}>
end
where "'[' x ':=' s ']' t" := (subst x s t) (in custom stlc).
xxxxxxxxxx
65
150
Inductive value : tm :=
| v_abs : ∀ x T2 t1,
value <{\x:T2, t1}>
| v_true :
value <{true}>
| v_false :
value <{false}>
| v_pair : ∀ v1 v2,
value v1
value v2
value <{(v1, v2)}>.
Hint Constructors value : core.
Reserved Notation "t '-->' t'" (at level 40).
Inductive step : tm tm :=
| ST_AppAbs : ∀ x T2 t1 v2,
value v2
<{(\x:T2, t1) v2}> --> <{ [x:=v2]t1 }>
| ST_App1 : ∀ t1 t1' t2,
t1 --> t1'
<{t1 t2}> --> <{t1' t2}>
| ST_App2 : ∀ v1 t2 t2',
value v1
t2 --> t2'
<{v1 t2}> --> <{v1 t2'}>
| ST_IfTrue : ∀ t1 t2,
<{if true then t1 else t2}> --> t1
| ST_IfFalse : ∀ t1 t2,
<{if false then t1 else t2}> --> t2
| ST_If : ∀ t1 t1' t2 t3,
t1 --> t1'
<{if t1 then t2 else t3}> --> <{if t1' then t2 else t3}>
| ST_Pair1 : ∀ t1 t1' t2,
t1 --> t1'
<{ (t1,t2) }> --> <{ (t1' , t2) }>
| ST_Pair2 : ∀ v1 t2 t2',
value v1
t2 --> t2'
<{ (v1, t2) }> --> <{ (v1, t2') }>
| ST_Fst1 : ∀ t0 t0',
t0 --> t0'
<{ t0.fst }> --> <{ t0'.fst }>
| ST_FstPair : ∀ v1 v2,
value v1
value v2
<{ (v1,v2).fst }> --> v1
| ST_Snd1 : ∀ t0 t0',
t0 --> t0'
<{ t0.snd }> --> <{ t0'.snd }>
| ST_SndPair : ∀ v1 v2,
value v1
value v2
<{ (v1,v2).snd }> --> v2
where "t '-->' t'" := (step t t').
Hint Constructors step : core.
Notation multistep := (multi step).
Notation "t1 '-->*' t2" := (multistep t1 t2) (at level 40).
Notation step_normal_form := (normal_form step).
Lemma valuenormal : ∀ t, value t step_normal_form t.
Proof with eauto.
intros t H; induction H; intros [t' ST]; inversion ST...
Qed.
xxxxxxxxxx
40
190
Definition context := partial_map ty.
Reserved Notation "Gamma '|-' t '\in' T" (at level 40,
t custom stlc, T custom stlc at level 0).
Inductive has_type : context tm ty :=
(* Same as before: *)
(* pure STLC *)
| T_Var : ∀ Gamma x T1,
Gamma x = Some T1
Gamma |- x \in T1
| T_Abs : ∀ Gamma x T1 T2 t1,
(x | T2 ; Gamma) |- t1 \in T1
Gamma |- \x:T2, t1 \in (T2 T1)
| T_App : ∀ T1 T2 Gamma t1 t2,
Gamma |- t1 \in (T2 T1)
Gamma |- t2 \in T2
Gamma |- t1 t2 \in T1
| T_True : ∀ Gamma,
Gamma |- true \in Bool
| T_False : ∀ Gamma,
Gamma |- false \in Bool
| T_If : ∀ t1 t2 t3 T1 Gamma,
Gamma |- t1 \in Bool
Gamma |- t2 \in T1
Gamma |- t3 \in T1
Gamma |- if t1 then t2 else t3 \in T1
| T_Pair : ∀ Gamma t1 t2 T1 T2,
Gamma |- t1 \in T1
Gamma |- t2 \in T2
Gamma |- (t1, t2) \in (T1 * T2)
| T_Fst : ∀ Gamma t0 T1 T2,
Gamma |- t0 \in (T1 * T2)
Gamma |- t0.fst \in T1
| T_Snd : ∀ Gamma t0 T1 T2,
Gamma |- t0 \in (T1 * T2)
Gamma |- t0.snd \in T2
where "Gamma '|-' t '\in' T" := (has_type Gamma t T).
Hint Constructors has_type : core.
Hint Extern 2 (has_type _ (app _ _) _) eapply T_App; auto : core.
Hint Extern 2 (_ = _) compute; reflexivity : core.
xxxxxxxxxx
17
207
Lemma weakening : ∀ Gamma Gamma' t 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.
Lemma weakening_empty : ∀ Gamma t T,
empty |- t \in T
Gamma |- t \in T.
Proof.
intros Gamma t T.
eapply weakening.
discriminate.
Qed.
xxxxxxxxxx
27
234
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.
Proof with eauto.
intros Gamma x U t v T Ht Hv.
generalize dependent Gamma. generalize dependent T.
induction t; intros T Gamma H;
(* in each case, we'll want to get at the derivation of H *)
inversion H; clear H; subst; simpl; eauto.
- (* var *)
rename s into y. destruct (eqb_spec x y); subst.
+ (* x=y *)
rewrite update_eq in H2.
injection H2 as H2; subst.
apply weakening_empty. assumption.
+ (* x<>y *)
apply T_Var. rewrite update_neq in H2; auto.
- (* abs *)
rename s into y, t into S.
destruct (eqb_spec x y); subst; apply T_Abs.
+ (* x=y *)
rewrite update_shadow in H5. assumption.
+ (* x<>y *)
apply IHt.
rewrite update_permute; auto.
Qed.
xxxxxxxxxx
20
254
Theorem preservation : ∀ t t' 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; inversion HE; subst...
- (* T_App *)
inversion HE; subst...
+ (* ST_AppAbs *)
apply substitution_preserves_typing with T2...
inversion HT1...
- (* T_Fst *)
inversion HT...
- (* T_Snd *)
inversion HT...
Qed.
xxxxxxxxxx
98
352
Inductive appears_free_in : string tm :=
| afi_var : ∀ (x : string),
appears_free_in x <{ x }>
| afi_app1 : ∀ x t1 t2,
appears_free_in x t1 appears_free_in x <{ t1 t2 }>
| afi_app2 : ∀ x t1 t2,
appears_free_in x t2 appears_free_in x <{ t1 t2 }>
| afi_abs : ∀ x y T11 t12,
y x
appears_free_in x t12
appears_free_in x <{ \y : T11, t12 }>
(* booleans *)
| afi_test0 : ∀ x t0 t1 t2,
appears_free_in x t0
appears_free_in x <{ if t0 then t1 else t2 }>
| afi_test1 : ∀ x t0 t1 t2,
appears_free_in x t1
appears_free_in x <{ if t0 then t1 else t2 }>
| afi_test2 : ∀ x t0 t1 t2,
appears_free_in x t2
appears_free_in x <{ if t0 then t1 else t2 }>
(* pairs *)
| afi_pair1 : ∀ x t1 t2,
appears_free_in x t1
appears_free_in x <{ (t1, t2) }>
| afi_pair2 : ∀ x t1 t2,
appears_free_in x t2
appears_free_in x <{ (t1 , t2) }>
| afi_fst : ∀ x t,
appears_free_in x t
appears_free_in x <{ t.fst }>
| afi_snd : ∀ x t,
appears_free_in x t
appears_free_in x <{ t.snd }>
.
Hint Constructors appears_free_in : core.
Definition closed (t:tm) :=
∀ x, ~ appears_free_in x t.
Lemma context_invariance : ∀ Gamma Gamma' t S,
Gamma |- t \in S
(∀ x, appears_free_in x t Gamma x = Gamma' x)
Gamma' |- t \in S.
Proof.
intros.
generalize dependent Gamma'.
induction H; intros; eauto 12.
- (* T_Var *)
apply T_Var. rewrite H0; auto.
- (* T_Abs *)
apply T_Abs.
apply IHhas_type. intros x1 Hafi.
(* the only tricky step... *)
destruct (eqb_spec x x1); subst.
+ rewrite update_eq.
rewrite update_eq.
reflexivity.
+ rewrite update_neq; [| assumption].
rewrite update_neq; [| assumption].
auto.
Qed.
(* A handy consequence of eqb_neq *)
Theorem false_eqb_string : ∀ x y : string,
x y String.eqb x y = false.
Proof.
intros x y. rewrite String.eqb_neq.
intros H. apply H. Qed.
Lemma free_in_context : ∀ x t T Gamma,
appears_free_in x t
Gamma |- t \in T
T', Gamma x = Some T'.
Proof with eauto.
intros x t T Gamma Hafi Htyp.
induction Htyp; inversion Hafi; subst...
- (* T_Abs *)
destruct IHHtyp as [T' Hctx]... exists T'.
unfold update, t_update in Hctx.
rewrite false_eqb_string in Hctx...
Qed.
Corollary typable_emptyclosed : ∀ t T,
empty |- t \in T
closed t.
Proof.
intros. unfold closed. intros x H1.
destruct (free_in_context _ _ _ _ H1 H) as [T' C].
discriminate C. Qed.
xxxxxxxxxx
58
410
Lemma step_deterministic :
deterministic step.
Proof with eauto.
unfold deterministic.
intros t t' t'' E1 E2.
generalize dependent t''.
induction E1; intros t'' E2; inversion E2; subst; clear E2...
(* ST_AppAbs *)
- inversion H3.
- exfalso; apply valuenormal in H...
(* ST_App1 *)
- inversion E1.
- f_equal...
- exfalso; apply valuenormal in H1...
(* ST_App2 *)
- exfalso; apply valuenormal in H3...
- exfalso; apply valuenormal in H...
- f_equal...
- (* ST_IfTrue *)
inversion H3.
- (* ST_IfFalse *)
inversion H3.
(* ST_If *)
- inversion E1.
- inversion E1.
- f_equal...
(* ST_Pair1 *)
- f_equal...
- exfalso; apply valuenormal in H1...
(* ST_Pair2 *)
- exfalso; apply valuenormal in H...
- f_equal...
(* ST_Fst1 *)
- f_equal...
- exfalso.
inversion E1; subst.
+ apply valuenormal in H0...
+ apply valuenormal in H1...
(* ST_FstPair *)
- exfalso.
inversion H2; subst.
+ apply valuenormal in H...
+ apply valuenormal in H0...
(* ST_Snd1 *)
- f_equal...
- exfalso.
inversion E1; subst.
+ apply valuenormal in H0...
+ apply valuenormal in H1...
(* ST_SndPair *)
- exfalso.
inversion H2; subst.
+ apply valuenormal in H...
+ apply valuenormal in H0...
Qed.
Normalization
xxxxxxxxxx
411
Definition halts (t:tm) : := t', t -->* t' value t'.
A trivial fact:
xxxxxxxxxx
10
421
Lemma value_halts : ∀ v, value v halts v.
Proof.
intros v H. unfold halts.
exists v. split.
- apply multi_refl.
- assumption.
Qed.
The key issue in the normalization proof (as in many proofs by
induction) is finding a strong enough induction hypothesis. To
this end, we begin by defining, for each type T, a set R_T of
closed terms of type T. We will specify these sets using a
relation R and write R T t when t is in R_T. (The sets
R_T are sometimes called saturated sets or reducibility
candidates.)
Here is the definition of R for the base language:
This definition gives us the strengthened induction hypothesis that we
need. Our primary goal is to show that all programs ---i.e., all
closed terms of base type---halt. But closed terms of base type can
contain subterms of functional type, so we need to know something
about these as well. Moreover, it is not enough to know that these
subterms halt, because the application of a normalized function to a
normalized argument involves a substitution, which may enable more
reduction steps. So we need a stronger condition for terms of
functional type: not only should they halt themselves, but, when
applied to halting arguments, they should yield halting results.
The form of R is characteristic of the logical relations proof
technique. (Since we are just dealing with unary relations here, we
could perhaps more properly say logical properties.) If we want to
prove some property P of all closed terms of type A, we proceed by
proving, by induction on types, that all terms of type A possess
property P, all terms of type A→A preserve property P, all
terms of type (A→A)->(A→A) preserve the property of preserving
property P, and so on. We do this by defining a family of
properties, indexed by types. For the base type A, the property is
just P. For functional types, it says that the function should map
values satisfying the property at the input type to values satisfying
the property at the output type.
When we come to formalize the definition of R in Coq, we hit a
problem. The most obvious formulation would be as a parameterized
Inductive proposition like this:
Inductive R : ty → tm → Prop :=
| R_bool : ∀ b t, empty ⊢ t \in Bool →
halts t →
R Bool t
| R_arrow : ∀ T1 T2 t, empty ⊢ t \in (Arrow T1 T2) →
halts t →
(∀ s, R T1 s → R T2 (app t s)) →
R (Arrow T1 T2) t. Unfortunately, Coq rejects this definition because it violates the strict positivity requirement for inductive definitions, which says that the type being defined must not occur to the left of an arrow in the type of a constructor argument. Here, it is the third argument to R_arrow, namely (∀ s, R T1 s → R TS (app t s)), and specifically the R T1 s part, that violates this rule. (The outermost arrows separating the constructor arguments don't count when applying this rule; otherwise we could never have genuinely inductive properties at all!) The reason for the rule is that types defined with non-positive recursion can be used to build non-terminating functions, which as we know would be a disaster for Coq's logical soundness. Even though the relation we want in this case might be perfectly innocent, Coq still rejects it because it fails the positivity test.
Fortunately, it turns out that we can define R using a
Fixpoint:
- R bool t iff t is a closed term of type bool and t halts
in a value
- R (T1 → T2) t iff t is a closed term of type T1 → T2 and t halts in a value and for any term s such that R T1 s, we have R T2 (t s).
Inductive R : ty → tm → Prop :=
| R_bool : ∀ b t, empty ⊢ t \in Bool →
halts t →
R Bool t
| R_arrow : ∀ T1 T2 t, empty ⊢ t \in (Arrow T1 T2) →
halts t →
(∀ s, R T1 s → R T2 (app t s)) →
R (Arrow T1 T2) t. Unfortunately, Coq rejects this definition because it violates the strict positivity requirement for inductive definitions, which says that the type being defined must not occur to the left of an arrow in the type of a constructor argument. Here, it is the third argument to R_arrow, namely (∀ s, R T1 s → R TS (app t s)), and specifically the R T1 s part, that violates this rule. (The outermost arrows separating the constructor arguments don't count when applying this rule; otherwise we could never have genuinely inductive properties at all!) The reason for the rule is that types defined with non-positive recursion can be used to build non-terminating functions, which as we know would be a disaster for Coq's logical soundness. Even though the relation we want in this case might be perfectly innocent, Coq still rejects it because it fails the positivity test.
xxxxxxxxxx
430
Fixpoint R (T:ty) (t:tm) : :=
empty |- t \in T halts t
(match T with
| <{ Bool }> True
| <{ T1 T2 }> (∀ s, R T1 s R T2 <{t s}> )
(* ... edit the next line when dealing with products *)
| <{ T1 * T2 }> False (* FILL IN HERE *)
end).
As immediate consequences of this definition, we have that every
element of every set R_T halts in a value and is closed with type
T :
Lemma R_halts : ∀ {T} {t}, R T t → halts t.
Lemma R_typable_empty : ∀ {T} {t}, R T t → empty ⊢ t \in T.
Lemma R_typable_empty : ∀ {T} {t}, R T t → empty ⊢ t \in T.
Now we proceed to show the main result, which is that every
well-typed term of type T is an element of R_T. Together with
R_halts, that will show that every well-typed term halts in a
value.
Membership in R_T Is Invariant Under Reduction
Now the main lemma, which comes in two parts, one for each
direction. Each proceeds by induction on the structure of the type
T. In fact, this is where we make fundamental use of the
structure of types.
One requirement for staying in R_T is to stay in type T. In the
forward direction, we get this from ordinary type Preservation.
The generalization to multiple steps is trivial:
In the reverse direction, we must add the fact that t has type
T before stepping as an additional hypothesis.
Lemma step_preserves_R' : ∀ T t t',
empty ⊢ t \in T → (t --> t') → R T t' → R T t.
Lemma multistep_preserves_R' : ∀ T t t',
empty ⊢ t \in T → (t -->* t') → R T t' → R T t.
empty ⊢ t \in T → (t --> t') → R T t' → R T t.
Lemma multistep_preserves_R' : ∀ T t t',
empty ⊢ t \in T → (t -->* t') → R T t' → R T t.
Closed Instances of Terms of Type t Belong to R_T
Multisubstitutions, Multi-Extensions, and Instantiations
Definition env := list (string × tm).
Fixpoint msubst (ss:env) (t:tm) : tm :=
match ss with
| nil ⇒ t
| ((x,s)::ss') ⇒ msubst ss' <{ [x:=s]t }>
end.
Fixpoint msubst (ss:env) (t:tm) : tm :=
match ss with
| nil ⇒ t
| ((x,s)::ss') ⇒ msubst ss' <{ [x:=s]t }>
end.
We need similar machinery to talk about repeated extension of a
typing context using a list of (identifier, type) pairs, which we
call a type assignment.
Definition tass := list (string × ty).
Fixpoint mupdate (Gamma : context) (xts : tass) :=
match xts with
| nil ⇒ Gamma
| ((x,v)::xts') ⇒ update (mupdate Gamma xts') x v
end.
Fixpoint mupdate (Gamma : context) (xts : tass) :=
match xts with
| nil ⇒ Gamma
| ((x,v)::xts') ⇒ update (mupdate Gamma xts') x v
end.
We will need some simple operations that work uniformly on
environments and type assigments
Fixpoint lookup {X:Set} (k : string) (l : list (string × X))
: option X :=
match l with
| nil ⇒ None
| (j,x) :: l' ⇒
if String.eqb j k then Some x else lookup k l'
end.
Fixpoint drop {X:Set} (n:string) (nxs:list (string × X))
: list (string × X) :=
match nxs with
| nil ⇒ nil
| ((n',x)::nxs') ⇒
if String.eqb n' n then drop n nxs'
else (n',x)::(drop n nxs')
end.
: option X :=
match l with
| nil ⇒ None
| (j,x) :: l' ⇒
if String.eqb j k then Some x else lookup k l'
end.
Fixpoint drop {X:Set} (n:string) (nxs:list (string × X))
: list (string × X) :=
match nxs with
| nil ⇒ nil
| ((n',x)::nxs') ⇒
if String.eqb n' n then drop n nxs'
else (n',x)::(drop n nxs')
end.
An instantiation combines a type assignment and a value
environment with the same domains, where corresponding elements are
in R.
Inductive instantiation : tass → env → Prop :=
| V_nil :
instantiation nil nil
| V_cons : ∀ x T v c e,
value v → R T v →
instantiation c e →
instantiation ((x,T)::c) ((x,v)::e).
| V_nil :
instantiation nil nil
| V_cons : ∀ x T v c e,
value v → R T v →
instantiation c e →
instantiation ((x,T)::c) ((x,v)::e).
We now proceed to prove various properties of these definitions.
Lemma vacuous_substitution : ∀ t x,
¬ appears_free_in x t →
∀ t', <{ [x:=t']t }> = t.
Lemma subst_closed: ∀ t,
closed t →
∀ x t', <{ [x:=t']t }> = t.
Lemma subst_not_afi : ∀ t x v,
closed v → ¬ appears_free_in x <{ [x:=v]t }>.
Lemma duplicate_subst : ∀ t' x t v,
closed v → <{ [x:=t]([x:=v]t') }> = <{ [x:=v]t' }>.
Lemma swap_subst : ∀ t x x1 v v1,
x ≠ x1 →
closed v → closed v1 →
<{ [x1:=v1]([x:=v]t) }> = <{ [x:=v]([x1:=v1]t) }>.
¬ appears_free_in x t →
∀ t', <{ [x:=t']t }> = t.
Lemma subst_closed: ∀ t,
closed t →
∀ x t', <{ [x:=t']t }> = t.
Lemma subst_not_afi : ∀ t x v,
closed v → ¬ appears_free_in x <{ [x:=v]t }>.
Lemma duplicate_subst : ∀ t' x t v,
closed v → <{ [x:=t]([x:=v]t') }> = <{ [x:=v]t' }>.
Lemma swap_subst : ∀ t x x1 v v1,
x ≠ x1 →
closed v → closed v1 →
<{ [x1:=v1]([x:=v]t) }> = <{ [x:=v]([x1:=v1]t) }>.
Closed environments are those that contain only closed terms.
Fixpoint closed_env (env:env) :=
match env with
| nil ⇒ True
| (x,t)::env' ⇒ closed t ∧ closed_env env'
end.
match env with
| nil ⇒ True
| (x,t)::env' ⇒ closed t ∧ closed_env env'
end.
Next come a series of lemmas charcterizing how msubst of closed terms
distributes over subst and over each term form
Lemma subst_msubst: ∀ env x v t, closed v → closed_env env →
msubst env <{ [x:=v]t }> = <{ [x:=v] { msubst (drop x env) t } }> .
Lemma msubst_var: ∀ ss x, closed_env ss →
msubst ss (tm_var x) =
match lookup x ss with
| Some t ⇒ t
| None ⇒ tm_var x
end.
Lemma msubst_abs: ∀ ss x T t,
msubst ss <{ \ x : T, t }> = <{ \x : T, {msubst (drop x ss) t} }>.
Lemma msubst_app : ∀ ss t1 t2,
msubst ss <{ t1 t2 }> = <{ {msubst ss t1} ({msubst ss t2}) }>.
msubst env <{ [x:=v]t }> = <{ [x:=v] { msubst (drop x env) t } }> .
Lemma msubst_var: ∀ ss x, closed_env ss →
msubst ss (tm_var x) =
match lookup x ss with
| Some t ⇒ t
| None ⇒ tm_var x
end.
Lemma msubst_abs: ∀ ss x T t,
msubst ss <{ \ x : T, t }> = <{ \x : T, {msubst (drop x ss) t} }>.
Lemma msubst_app : ∀ ss t1 t2,
msubst ss <{ t1 t2 }> = <{ {msubst ss t1} ({msubst ss t2}) }>.
You'll need similar functions for the other term constructors.
(* FILL IN HERE *)
Properties of Multi-Extensions
Lemma mupdate_lookup : ∀ (c : tass) (x:string),
lookup x c = (mupdate empty c) x.
Lemma mupdate_drop : ∀ (c: tass) Gamma x x',
mupdate Gamma (drop x c) x'
= if String.eqb x x' then Gamma x' else mupdate Gamma c x'.
lookup x c = (mupdate empty c) x.
Lemma mupdate_drop : ∀ (c: tass) Gamma x x',
mupdate Gamma (drop x c) x'
= if String.eqb x x' then Gamma x' else mupdate Gamma c x'.
Lemma instantiation_domains_match: ∀ {c} {e},
instantiation c e →
∀ {x} {T},
lookup x c = Some T → ∃ t, lookup x e = Some t.
Lemma instantiation_env_closed : ∀ c e,
instantiation c e → closed_env e.
Lemma instantiation_R : ∀ c e,
instantiation c e →
∀ x t T,
lookup x c = Some T →
lookup x e = Some t → R T t.
Lemma instantiation_drop : ∀ c env,
instantiation c env →
∀ x, instantiation (drop x c) (drop x env).
instantiation c e →
∀ {x} {T},
lookup x c = Some T → ∃ t, lookup x e = Some t.
Lemma instantiation_env_closed : ∀ c e,
instantiation c e → closed_env e.
Lemma instantiation_R : ∀ c e,
instantiation c e →
∀ x t T,
lookup x c = Some T →
lookup x e = Some t → R T t.
Lemma instantiation_drop : ∀ c env,
instantiation c env →
∀ x, instantiation (drop x c) (drop x env).
Lemma multistep_App2 : ∀ v t t',
value v → (t -->* t') → <{ v t }> -->* <{ v t' }>.
(* FILL IN HERE *)
value v → (t -->* t') → <{ v t }> -->* <{ v t' }>.
(* FILL IN HERE *)
The R Lemma
Lemma msubst_preserves_typing : ∀ c e,
instantiation c e →
∀ Gamma t S, (mupdate Gamma c) ⊢ t \in S →
Gamma ⊢ { (msubst e t) } \in S.
instantiation c e →
∀ Gamma t S, (mupdate Gamma c) ⊢ t \in S →
Gamma ⊢ { (msubst e t) } \in S.
And at long last, the main lemma.