NormInTypeAlternate Version of Normalization, for Extraction
(* $Date: 2012-04-16 20:14:27 -0400 (Mon, 16 Apr 2012) $ *)
(* Chapter maintained by Andrew Tolmach and Benjamin Pierce *)
Note
Normalization
Require Import Stlc.
Exercise: 1 star
Where do we fail if we attempt to prove normalization by a straightforward induction on the size of a well-typed term?(* FILL IN HERE *)
☐
Language
Inductive ty : Type :=
| TBool : ty
| TArrow : ty → ty → ty
.
Tactic Notation "T_cases" tactic(first) ident(c) :=
first;
[ Case_aux c "TBool" | Case_aux c "TArrow" | Case_aux c "TProd" ].
Inductive tm : Type :=
(* pure STLC *)
| tvar : id → tm
| tapp : tm → tm → tm
| tabs : id → ty → tm → tm
(* booleans *)
| ttrue : tm
| tfalse : tm
| tif : tm → tm → tm → tm.
(* i.e., if t0 then t1 else t2 *)
Tactic Notation "t_cases" tactic(first) ident(c) :=
first;
[ Case_aux c "tvar" | Case_aux c "tapp" | Case_aux c "tabs"
| Case_aux c "ttrue" | Case_aux c "tfalse" | Case_aux c "tif" ].
Fixpoint subst (x:id) (s:tm) (t:tm) : tm :=
match t with
| tvar y => if beq_id x y then s else t
| tabs y T t1 => tabs y T (if beq_id x y then t1 else (subst x s t1))
| tapp t1 t2 => tapp (subst x s t1) (subst x s t2)
| ttrue => ttrue
| tfalse => tfalse
| tif t0 t1 t2 => tif (subst x s t0) (subst x s t1) (subst x s t2)
end.
Notation "'[' x ':=' s ']' t" := (subst x s t) (at level 20).
Inductive value : tm → Type :=
| v_abs : ∀x T11 t12,
value (tabs x T11 t12)
| v_true : value ttrue
| v_false : value tfalse
.
Hint Constructors value.
Reserved Notation "t1 '⇒' t2" (at level 40).
Inductive step : tm → tm → Type :=
| ST_AppAbs : ∀x T11 t12 v2,
value v2 →
(tapp (tabs x T11 t12) v2) ⇒ [x:=v2]t12
| ST_App1 : ∀t1 t1' t2,
t1 ⇒ t1' →
(tapp t1 t2) ⇒ (tapp t1' t2)
| ST_App2 : ∀v1 t2 t2',
value v1 →
t2 ⇒ t2' →
(tapp v1 t2) ⇒ (tapp v1 t2')
(* booleans *)
| ST_IfTrue : ∀t1 t2,
(tif ttrue t1 t2) ⇒ t1
| ST_IfFalse : ∀t1 t2,
(tif tfalse t1 t2) ⇒ t2
| ST_If : ∀t0 t0' t1 t2,
t0 ⇒ t0' →
(tif t0 t1 t2) ⇒ (tif t0' t1 t2)
where "t1 '⇒' t2" := (step t1 t2).
Tactic Notation "step_cases" tactic(first) ident(c) :=
first;
[ Case_aux c "ST_AppAbs" | Case_aux c "ST_App1" | Case_aux c "ST_App2"
| Case_aux c "ST_IfTrue" | Case_aux c "ST_IfFalse" | Case_aux c "ST_If" ].
Definition relation (X:Type) := X → X → Type.
Inductive multi (X:Type) (R: relation X)
: X → X → Type :=
| multi_refl : ∀(x : X),
multi X R x x
| multi_step : ∀(x y z : X),
R x y →
multi X R y z →
multi X R x z.
Implicit Arguments multi [[X]].
Tactic Notation "multi_cases" tactic(first) ident(c) :=
first;
[ Case_aux c "multi_refl" | Case_aux c "multi_step" ].
Notation multistep := (multi step).
Notation "t1 '⇒*' t2" := (multistep t1 t2) (at level 40).
Hint Constructors step.
Inductive ex (X:Type) (P : X→Type) : Type :=
ex_intro : ∀(witness:X), P witness → ex X P.
Notation "'exists' x , p" := (ex _ (fun x => p))
(at level 200, x ident, right associativity) : type_scope.
Notation "'exists' x : X , p" := (ex _ (fun x:X => p))
(at level 200, x ident, right associativity) : type_scope.
Definition normal_form {X:Type} (R:relation X) (t:X) : Type :=
(∃t', R t t') → False.
Notation step_normal_form := (normal_form step).
Definition deterministic {X: Type} (R: relation X) :=
∀x y1 y2 : X, R x y1 → R x y2 → y1 = y2.
Hint Constructors ex.
Lemma value__normal : ∀t, value t → step_normal_form t.
Proof with eauto.
intros t H; induction H; intros [t' ST]; inversion ST...
Qed.
Definition context := partial_map ty.
Inductive has_type : context → tm → ty → Type :=
(* Typing rules for proper terms *)
| T_Var : ∀Γ x T,
Γ x = Some T →
has_type Γ (tvar x) T
| T_Abs : ∀Γ x T11 T12 t12,
has_type (extend Γ x T11) t12 T12 →
has_type Γ (tabs x T11 t12) (TArrow T11 T12)
| T_App : ∀T1 T2 Γ t1 t2,
has_type Γ t1 (TArrow T1 T2) →
has_type Γ t2 T1 →
has_type Γ (tapp t1 t2) T2
(* booleans *)
| T_True : ∀Γ,
has_type Γ ttrue TBool
| T_False : ∀Γ,
has_type Γ tfalse TBool
| T_If : ∀Γ t0 t1 t2 T,
has_type Γ t0 TBool →
has_type Γ t1 T →
has_type Γ t2 T →
has_type Γ (tif t0 t1 t2) T
.
Hint Constructors has_type.
Tactic Notation "has_type_cases" tactic(first) ident(c) :=
first;
[ Case_aux c "T_Var" | Case_aux c "T_Abs" | Case_aux c "T_App"
| Case_aux c "T_True" | Case_aux c "T_False" | Case_aux c "T_If" ].
Hint Extern 2 (has_type _ (tapp _ _) _) => eapply T_App; auto.
Hint Extern 2 (_ = _) => compute; reflexivity.
Inductive appears_free_in : id → tm → Type :=
| afi_var : ∀x,
appears_free_in x (tvar x)
| afi_app1 : ∀x t1 t2,
appears_free_in x t1 → appears_free_in x (tapp t1 t2)
| afi_app2 : ∀x t1 t2,
appears_free_in x t2 → appears_free_in x (tapp t1 t2)
| afi_abs : ∀x y T11 t12,
y <> x →
appears_free_in x t12 →
appears_free_in x (tabs y T11 t12)
(* booleans *)
| afi_if0 : ∀x t0 t1 t2,
appears_free_in x t0 →
appears_free_in x (tif t0 t1 t2)
| afi_if1 : ∀x t0 t1 t2,
appears_free_in x t1 →
appears_free_in x (tif t0 t1 t2)
| afi_if2 : ∀x t0 t1 t2,
appears_free_in x t2 →
appears_free_in x (tif t0 t1 t2)
.
Hint Constructors appears_free_in.
Definition closed (t:tm) :=
∀x, appears_free_in x t → False.
Lemma context_invariance : ∀Γ Γ' t S,
has_type Γ t S →
(∀x, appears_free_in x t → Γ x = Γ' x) →
has_type Γ' t S.
Proof with eauto.
intros Γ Γ' t S H H0. generalize dependent Γ'.
has_type_cases (induction H) Case;
intros Γ' Heqv...
Case "T_Var".
apply T_Var... rewrite ← Heqv...
Case "T_Abs".
apply T_Abs... apply IHhas_type. intros y Hafi.
unfold extend. remember (beq_id x y) as e.
destruct e...
Case "T_If".
eapply T_If...
Qed.
Lemma free_in_context : ∀x t T Γ,
appears_free_in x t →
has_type Γ t T →
∃T', Γ x = Some T'.
Proof with eauto.
intros x t T Γ Hafi Htyp.
has_type_cases (induction Htyp) Case; inversion Hafi; subst...
Case "T_Abs".
destruct IHHtyp as [T' Hctx]... ∃T'.
unfold extend in Hctx.
apply not_eq_beq_id_false in H2. rewrite H2 in Hctx...
Qed.
Corollary typable_empty__closed : ∀t T,
has_type empty t T →
closed t.
Proof.
intros t T H. unfold closed. intros x H1.
destruct (free_in_context _ _ _ _ H1 H) as [T' C].
inversion C. Qed.
Lemma substitution_preserves_typing : ∀Γ x U v t S,
has_type (extend Γ x U) t S →
has_type empty v U →
has_type Γ ([x:=v]t) S.
Proof with eauto.
(* Theorem: If Gamma,x:U |- t : S and empty |- v : U, then
Gamma |- (x:=vt) S. *)
intros Γ x U v t S Htypt Htypv.
generalize dependent Γ. generalize dependent S.
(* Proof: By induction on the term t. Most cases follow directly
from the IH, with the exception of tvar and tabs.
The former aren't automatic because we must reason about how the
variables interact. *)
t_cases (induction t) Case;
intros S Γ Htypt; simpl; inversion Htypt; subst...
Case "tvar".
simpl. rename i into y.
(* If t = y, we know that
empty ⊢ v : U and
Γ,x:U ⊢ y : S
and, by inversion, extend Γ x U y = Some S. We want to
show that Γ ⊢ [x:=v]y : S.
There are two cases to consider: either x=y or x<>y. *)
remember (beq_id x y) as e. destruct e.
SCase "x=y".
(* If x = y, then we know that U = S, and that [x:=v]y = v.
So what we really must show is that if empty ⊢ v : U then
Γ ⊢ v : U. We have already proven a more general version
of this theorem, called context invariance. *)
apply beq_id_eq in Heqe. subst.
unfold extend in H1. rewrite ← beq_id_refl in H1.
inversion H1; subst. clear H1.
eapply context_invariance...
intros x Hcontra.
destruct (free_in_context _ _ S empty Hcontra) as [T' HT']...
inversion HT'.
SCase "x<>y".
(* If x <> y, then Γ y = Some S and the substitution has no
effect. We can show that Γ ⊢ y : S by T_Var. *)
apply T_Var... unfold extend in H1. rewrite ← Heqe in H1...
Case "tabs".
rename i into y. rename t into T11.
(* If t = tabs y T11 t0, then we know that
Γ,x:U ⊢ tabs y T11 t0 : T11→T12
Γ,x:U,y:T11 ⊢ t0 : T12
empty ⊢ v : U
As our IH, we know that forall S Gamma,
Γ,x:U ⊢ t0 : S → Γ ⊢ [x:=v]t0 S.
We can calculate that
x:=vt = tabs y T11 (if beq_id x y then t0 else x:=vt0)
And we must show that Γ ⊢ [x:=v]t : T11→T12. We know
we will do so using T_Abs, so it remains to be shown that:
Γ,y:T11 ⊢ if beq_id x y then t0 else [x:=v]t0 : T12
We consider two cases: x = y and x <> y.
*)
apply T_Abs...
remember (beq_id x y) as e. destruct e.
SCase "x=y".
(* If x = y, then the substitution has no effect. Context
invariance shows that Γ,y:U,y:T11 and Γ,y:T11 are
equivalent. Since the former context shows that t0 : T12, so
does the latter. *)
eapply context_invariance...
apply beq_id_eq in Heqe. subst.
intros x Hafi. unfold extend.
destruct (beq_id y x)...
SCase "x<>y".
(* If x <> y, then the IH and context invariance allow us to show that
Γ,x:U,y:T11 ⊢ t0 : T12 =>
Γ,y:T11,x:U ⊢ t0 : T12 =>
Γ,y:T11 ⊢ [x:=v]t0 : T12 *)
apply IHt. eapply context_invariance...
intros z Hafi. unfold extend.
remember (beq_id y z) as e0. destruct e0...
apply beq_id_eq in Heqe0. subst.
rewrite ← Heqe...
Qed.
Theorem preservation : ∀t t' T,
has_type empty t T →
t ⇒ t' →
has_type empty t' T.
Proof with eauto.
intros t t' T HT.
(* Theorem: If empty ⊢ t : T and t ⇒ t', then empty ⊢ t' : T. *)
remember (@empty ty) as Γ. generalize dependent HeqGamma.
generalize dependent t'.
(* Proof: By induction on the given typing derivation. Many cases are
contradictory (T_Var, T_Abs). We show just the interesting ones. *)
has_type_cases (induction HT) Case;
intros t' HeqGamma HE; subst; inversion HE; subst...
Case "T_App".
(* If the last rule used was T_App, then t = t1 t2, and three rules
could have been used to show t ⇒ t': ST_App1, ST_App2, and
ST_AppAbs. In the first two cases, the result follows directly from
the IH. *)
inversion HE; subst...
SCase "ST_AppAbs".
(* For the third case, suppose
t1 = tabs x T11 t12
and
t2 = v2.
We must show that empty ⊢ [x:=v2]t12 : T2.
We know by assumption that
empty ⊢ tabs x T11 t12 : T1→T2
and by inversion
x:T1 ⊢ t12 : T2
We have already proven that substitution_preserves_typing and
empty ⊢ v2 : T1
by assumption, so we are done. *)
apply substitution_preserves_typing with T1...
inversion HT1...
Qed.
☐
Lemma step_deterministic :
deterministic step.
Proof with eauto.
unfold deterministic.
intros t t' t'' E1 E2.
generalize dependent t''.
step_cases (induction E1) Case; intros t'' E2; inversion E2; subst; clear E2...
Case "ST_AppAbs"...
inversion H2.
apply ex_falso_quodlibet; apply value__normal in v...
Case "ST_App1".
inversion E1.
f_equal...
apply ex_falso_quodlibet; apply value__normal in H1...
Case "ST_App2".
apply ex_falso_quodlibet; apply value__normal in H2...
apply ex_falso_quodlibet; apply value__normal in v...
f_equal...
Case "ST_IfTrue".
inversion H3.
Case "ST_IfFalse".
inversion H3.
Case "ST_If".
inversion E1.
inversion E1.
f_equal...
Qed.
Normalization
Definition halts (t:tm) : Type := (∃t' : tm, (t ⇒* t') * value t') %type.
A trivial fact:
Lemma value_halts : ∀v, value v → halts v.
Proof.
intros v H. unfold halts.
∃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
evaluation 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 predicates.) 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
predicates, indexed by types. For the base type A, the predicate is
just P. For functional types, it says that the function should map
values satisfying the predicate at the input type to values satisfying
the predicate 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:
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 (tapp 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
predicates 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, has_type empty t TBool →
halts t →
R TBool t
| R_arrow : ∀ T1 T2 t, has_type empty t (TArrow T1 T2) →
halts t →
(∀ s, R T1 s → R T2 (tapp t s)) →
R (TArrow T1 T2) t.
| R_bool : ∀ b t, has_type empty t TBool →
halts t →
R TBool t
| R_arrow : ∀ T1 T2 t, has_type empty t (TArrow T1 T2) →
halts t →
(∀ s, R T1 s → R T2 (tapp t s)) →
R (TArrow T1 T2) t.
Fixpoint R (T:ty) (t:tm) {struct T} : Type :=
(has_type empty t T * (halts t *
(match T with
| TBool => True
| TArrow T1 T2 => (∀s, R T1 s → R T2 (tapp t s))
end))) % type.
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.
Proof.
intros. destruct T; unfold R in X; inversion X; inversion X1; subst; assumption.
Qed.
Lemma R_typable_empty : ∀{T} {t}, R T t → has_type empty t T.
Proof.
intros. destruct T; unfold R in X; inversion X; inversion X1; assumption.
Qed.
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 evaluation
Definition iff (P Q : Type) := ((P → Q) * (Q → P)) % type.
Notation "P ↔ Q" := (iff P Q)
(at level 95, no associativity) : type_scope.
Theorem ex_falso_quodlibet : ∀(P:Type),
False → P.
Proof.
intros P contra.
inversion contra. Qed.
Lemma step_preserves_halting : ∀t t', (t ⇒ t') → (halts t ↔ halts t').
Proof.
intros t t' ST. unfold halts.
split.
Case "→".
intros [t'' [STM V]].
inversion STM; subst.
apply ex_falso_quodlibet. apply value__normal in V. unfold normal_form in V. apply V. ∃t'. auto.
rewrite (step_deterministic _ _ _ ST H). ∃t''. split; assumption.
Case "←".
intros [t'0 [STM V]].
∃t'0. split; eauto.
eapply multi_step. eassumption. assumption.
Qed.
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
finiteness 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.
Lemma step_preserves_R : ∀T t t', (t ⇒ t') → R T t → R T t'.
Proof.
induction T; intros t t' E Rt; unfold R; fold R; unfold R in Rt; fold R in Rt;
destruct Rt as [typable_empty_t [halts_t RRt]].
(* TBool *)
split. eapply preservation; eauto.
split. apply (step_preserves_halting _ _ E); eauto.
auto.
(* TArrow *)
split. eapply preservation; eauto.
split. apply (step_preserves_halting _ _ E); eauto.
intros.
eapply IHT2.
apply ST_App1. apply E.
apply RRt; auto.
Qed.
The generalization to multiple steps is trivial:
Lemma multistep_preserves_R : ∀T t t',
(t ⇒* t') → R T t → R T t'.
Proof.
intros T t t' STM; induction STM; intros.
assumption.
apply IHSTM. eapply step_preserves_R. apply r. assumption.
Qed.
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',
has_type empty t T → (t ⇒ t') → R T t' → R T t.
Proof.
induction T; intros t t' typable_empty_t E Rt'; unfold R; fold R;
unfold R in Rt'; fold R in Rt';
destruct Rt' as [typable_empty_t' [halts_t' RRt']].
(* TBool *)
split. assumption.
split. apply (step_preserves_halting _ _ E); eauto.
auto.
(* TArrow *)
split. assumption.
split. apply (step_preserves_halting _ _ E); eauto.
intros.
eapply IHT2. eapply T_App. apply typable_empty_t.
eapply R_typable_empty. assumption.
apply ST_App1. apply E.
apply RRt'; auto.
Qed.
Lemma multistep_preserves_R' : ∀T t t',
has_type empty t T → (t ⇒* t') → R T t' → R T t.
Proof.
intros T t t' HT STM.
induction STM; intros.
assumption.
eapply step_preserves_R'. assumption. apply r. apply IHSTM.
eapply preservation; eauto. auto.
Qed.
Closed instances of terms of type T belong to R_T
Multisubstitutions, multi-extensions, and instantiations
Definition env := list (id * tm).
Fixpoint msubst (ss:env) (t:tm) {struct ss} : 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 (id * ty).
Fixpoint mextend (Γ : context) (xts : tass) :=
match xts with
| nil => Γ
| ((x,v)::xts') => extend (mextend Γ xts') x v
end.
We will need some simple operations that work uniformly on
environments and type assigments
Fixpoint lookup {X:Set} (k : id) (l : list (id * X)) {struct l} : option X :=
match l with
| nil => None
| (j,x) :: l' =>
if beq_id j k then Some x else lookup k l'
end.
Fixpoint drop {X:Set} (n:id) (nxs:list (id * X)) {struct nxs} : list (id * X) :=
match nxs with
| nil => nil
| ((n',x)::nxs') => if beq_id 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 → Type :=
| 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 → False) →
∀t', [x:=t']t = t.
Proof with eauto.
t_cases (induction t) Case; intros; simpl.
Case "tvar".
remember (beq_id x i) as e. destruct e...
destruct H. apply beq_id_eq in Heqe. subst...
Case "tapp".
rewrite IHt1... rewrite IHt2...
Case "tabs".
remember (beq_id x i) as e. destruct e...
symmetry in Heqe; apply beq_id_false_not_eq in Heqe.
rewrite IHt...
Case "ttrue"...
Case "tfalse"...
Case "tif".
rewrite IHt1... rewrite IHt2... rewrite IHt3...
Qed.
Lemma subst_closed: ∀t,
closed t →
∀x t', [x:=t']t = t.
Proof.
intros. apply vacuous_substitution. apply H. Qed.
Lemma subst_not_afi : ∀t x v, closed v → (appears_free_in x ([x:=v]t) → False).
Proof with eauto. (* rather slow this way *)
unfold closed, not.
t_cases (induction t) Case; intros x v P A; simpl in A.
Case "tvar".
remember (beq_id x i) as e; destruct e...
inversion A; subst. rewrite ← beq_id_refl in Heqe; inversion Heqe.
Case "tapp".
inversion A; subst...
Case "tabs".
remember (beq_id x i) as e; destruct e...
apply beq_id_eq in Heqe; subst. inversion A; subst...
inversion A; subst...
Case "ttrue".
inversion A.
Case "tfalse".
inversion A.
Case "tif".
inversion A; subst...
Qed.
Lemma duplicate_subst : ∀t' x t v,
closed v → [x:=t]([x:=v]t') = [x:=v]t'.
Proof.
intros. eapply vacuous_substitution. apply subst_not_afi. auto.
Qed.
Lemma swap_subst : ∀t x x1 v v1, x <> x1 → closed v → closed v1 →
[x1:=v1]([x:=v]t) = [x:=v]([x1:=v1]t).
Proof with eauto.
t_cases (induction t) Case; intros; simpl.
Case "tvar".
remember (beq_id x i) as e; destruct e; remember (beq_id x1 i) as e; destruct e.
apply beq_id_eq in Heqe. apply beq_id_eq in Heqe0. subst.
apply ex_falso_quodlibet...
apply beq_id_eq in Heqe; subst. simpl.
rewrite ← beq_id_refl. apply subst_closed...
apply beq_id_eq in Heqe0; subst. simpl.
rewrite ← beq_id_refl. rewrite subst_closed...
simpl. rewrite ← Heqe. rewrite ← Heqe0...
Case "tapp".
rewrite IHt1... rewrite IHt2...
Case "tabs".
destruct (beq_id x i); destruct (beq_id x1 i)...
rewrite IHt...
Case "ttrue"...
Case "tfalse"...
Case "tif".
f_equal...
Qed.
Lemma msubst_closed: ∀t, closed t → ∀ss, msubst ss t = t.
Proof.
induction ss.
reflexivity.
destruct a. simpl. rewrite subst_closed; assumption.
Qed.
Closed environments are those that contain only closed terms.
Fixpoint closed_env (env:env) {struct env} :=
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).
Proof.
induction env0; intros.
auto.
destruct a. simpl.
inversion H0. fold closed_env in H2.
remember (beq_id i x) as e; destruct e.
apply beq_id_eq in Heqe; subst.
rewrite duplicate_subst; auto.
symmetry in Heqe. apply beq_id_false_not_eq in Heqe.
simpl. rewrite swap_subst; eauto.
Qed.
Lemma msubst_var: ∀ss x, closed_env ss →
msubst ss (tvar x) =
match lookup x ss with
| Some t => t
| None => tvar x
end.
Proof.
induction ss; intros.
reflexivity.
destruct a.
simpl. destruct (beq_id i x).
apply msubst_closed. inversion H; auto.
apply IHss. inversion H; auto.
Qed.
Lemma msubst_abs: ∀ss x T t,
msubst ss (tabs x T t) = tabs x T (msubst (drop x ss) t).
Proof.
induction ss; intros.
reflexivity.
destruct a.
simpl. destruct (beq_id i x); simpl; auto.
Qed.
Lemma msubst_app : ∀ss t1 t2, msubst ss (tapp t1 t2) = tapp (msubst ss t1) (msubst ss t2).
Proof.
induction ss; intros.
reflexivity.
destruct a.
simpl. rewrite ← IHss. auto.
Qed.
You'll need similar functions for the other term constructors.
(* FILL IN HERE *)
Lemma msubst_true : ∀ss, msubst ss ttrue = ttrue.
Proof.
induction ss; intros.
auto.
destruct a.
simpl. auto.
Qed.
Lemma msubst_false : ∀ss, msubst ss tfalse = tfalse.
Proof.
induction ss; intros.
auto.
destruct a.
simpl. auto.
Qed.
Lemma msubst_if : ∀ss t0 t1 t2,
msubst ss (tif t0 t1 t2) = tif (msubst ss t0) (msubst ss t1) (msubst ss t2).
Proof.
induction ss; intros.
auto.
destruct a. simpl. auto.
Qed.
Properties of multi-extensions
Lemma mextend_lookup : ∀(c : tass) (x:id), lookup x c = (mextend empty c) x.
Proof.
induction c; intros.
auto.
destruct a. unfold lookup, mextend, extend. destruct (beq_id i x); auto.
Qed.
Lemma mextend_drop : ∀(c: tass) Γ x x',
mextend Γ (drop x c) x' = if beq_id x x' then Γ x' else mextend Γ c x'.
induction c; intros.
destruct (beq_id x x'); auto.
destruct a. simpl.
remember (beq_id i x) as e; destruct e.
apply beq_id_eq in Heqe; subst. rewrite IHc.
remember (beq_id x x') as e; destruct e. auto. unfold extend. rewrite ← Heqe. auto.
simpl. unfold extend. remember (beq_id i x') as e; destruct e.
apply beq_id_eq in Heqe0; subst.
remember (beq_id x x') as e; destruct e.
apply beq_id_eq in Heqe0; subst. rewrite ← beq_id_refl in Heqe. inversion Heqe.
auto.
auto.
Qed.
Lemma instantiation_domains_match: ∀{c} {e},
instantiation c e → ∀{x} {T}, lookup x c = Some T → ∃t, lookup x e = Some t.
Proof.
intros c e V. induction V; intros x0 T0 C.
solve by inversion .
simpl in *.
destruct (beq_id x x0); eauto.
Qed.
Lemma instantiation_env_closed : ∀c e, instantiation c e → closed_env e.
Proof.
intros c e V; induction V; intros.
econstructor.
unfold closed_env. fold closed_env.
split. eapply typable_empty__closed. eapply R_typable_empty. eauto.
auto.
Qed.
Lemma instantiation_R : ∀c e, instantiation c e →
∀x t T, lookup x c = Some T →
lookup x e = Some t → R T t.
Proof.
intros c e V. induction V; intros x' t' T' G E.
solve by inversion.
unfold lookup in *. destruct (beq_id x x').
inversion G; inversion E; subst. auto.
eauto.
Qed.
Lemma instantiation_drop : ∀c env,
instantiation c env → ∀x, instantiation (drop x c) (drop x env).
Proof.
intros c e V. induction V.
intros. simpl. constructor.
intros. unfold drop. destruct (beq_id x x0); auto. constructor; eauto.
Qed.
Lemma multistep_App2 : ∀v t t',
value v → (t ⇒* t') → (tapp v t) ⇒* (tapp v t').
Proof.
intros v t t' V STM. induction STM.
apply multi_refl.
eapply multi_step.
apply ST_App2; eauto. auto.
Qed.
Lemma multistep_If : ∀t1 t1' t2 t3,
(t1 ⇒* t1') → (tif t1 t2 t3) ⇒* (tif t1' t2 t3).
Proof.
intros t1 t1' t2 t3 STM. induction STM.
apply multi_refl.
eapply multi_step.
apply ST_If; eauto. auto.
Qed.
(* FILL IN HERE *)
The R Lemma.
Lemma msubst_preserves_typing : ∀c e,
instantiation c e →
∀Γ t S, has_type (mextend Γ c) t S →
has_type Γ (msubst e t) S.
Proof.
induction 1; intros.
simpl in X. simpl. auto.
simpl in X0. simpl.
apply IHX.
eapply substitution_preserves_typing; eauto.
apply (R_typable_empty r).
Qed.
Theorem multi_trans :
∀(X:Type) (R: relation X) (x y z : X),
multi R x y →
multi R y z →
multi R x z.
Proof.
intros X R x y z G H.
multi_cases (induction G) Case.
Case "multi_refl". assumption.
Case "multi_step".
apply multi_step with y. assumption.
apply IHG. assumption. Qed.
Theorem multi_R : ∀(X:Type) (R:relation X) (x y : X),
R x y → multi R x y.
Proof.
intros X R x y r.
apply multi_step with y. apply r. apply multi_refl. Qed.
And at long last, the main lemma.
Lemma msubst_R : ∀c env t T,
has_type (mextend empty c) t T → instantiation c env → R T (msubst env t).
Proof.
intros c env0 t T HT V.
generalize dependent env0.
(* We need to generalize the hypothesis a bit before setting up the induction. *)
remember (mextend empty c) as Γ.
assert (∀x, Γ x = lookup x c).
intros. rewrite HeqGamma. rewrite mextend_lookup. auto.
clear HeqGamma.
generalize dependent c.
has_type_cases (induction HT) Case; intros.
Case "T_Var".
rewrite H in e. destruct (instantiation_domains_match V e) as [t P].
eapply instantiation_R; eauto.
rewrite msubst_var. rewrite P. auto. eapply instantiation_env_closed; eauto.
Case "T_Abs".
rewrite msubst_abs.
(* We'll need variants of the following fact several times, so its simplest to
establish it just once. *)
assert (WT: has_type empty (tabs x T11 (msubst (drop x env0) t12)) (TArrow T11 T12)).
eapply T_Abs. eapply msubst_preserves_typing. eapply instantiation_drop; eauto.
eapply context_invariance. apply HT.
intros.
unfold extend. rewrite mextend_drop. remember (beq_id x x0) as e; destruct e. auto.
rewrite H.
clear - c Heqe. induction c.
simpl. rewrite ← Heqe. auto.
simpl. destruct a. unfold extend. destruct (beq_id i x0); auto.
unfold R. fold R. split.
auto.
split. apply value_halts. apply v_abs.
intros.
destruct (R_halts X) as [v [P Q]].
pose proof (multistep_preserves_R _ _ _ P X).
apply multistep_preserves_R' with (msubst ((x,v)::env0) t12).
eapply T_App. eauto.
apply R_typable_empty; auto.
eapply multi_trans. eapply multistep_App2; eauto.
eapply multi_R.
simpl. rewrite subst_msubst.
eapply ST_AppAbs; eauto.
eapply typable_empty__closed.
apply (R_typable_empty X0).
eapply instantiation_env_closed; eauto.
eapply (IHHT ((x,T11)::c)).
intros. unfold extend, lookup. destruct (beq_id x x0); auto.
constructor; auto.
Case "T_App".
rewrite msubst_app.
destruct (IHHT1 c H env0 V) as [_ [_ P1]].
pose proof (IHHT2 c H env0 V) as P2. fold R in P1. auto.
Case "T_True".
rewrite msubst_true.
unfold R. split.
apply T_True.
split. unfold halts. ∃ttrue. split. apply multi_refl. apply v_true. auto.
Case "T_False".
rewrite msubst_false.
unfold R. split.
apply T_False.
split. unfold halts. ∃tfalse. split. apply multi_refl. apply v_false. auto.
Case "T_If".
rewrite msubst_if.
assert (WT: has_type empty (tif (msubst env0 t0) (msubst env0 t1) (msubst env0 t2)) T).
apply T_If; eapply msubst_preserves_typing; eauto;
eapply context_invariance; eauto; intros; rewrite ← mextend_lookup; auto.
pose proof (IHHT1 c H env0 V) as IH1.
destruct (R_halts IH1) as [v [P Q]].
assert (R TBool v).
eapply multistep_preserves_R. apply P. apply IH1.
(* The following is a somewhat easier approach than that taken in Pierce's TAPL
sample solutions; in essence we apply canonical forms, on the fly *)
pose proof (R_typable_empty X).
inversion Q; subst.
(* abs : impossible *)
inversion X0.
(* pair : impossible *)
inversion X0.
(* true *)
eapply multistep_preserves_R' with (msubst env0 t1).
assumption.
eapply multi_trans. apply multistep_If. eapply P.
eapply multi_R. apply ST_IfTrue.
apply (IHHT2 c H env0 V).
(* false *)
eapply multistep_preserves_R' with (msubst env0 t2).
assumption.
eapply multi_trans. apply multistep_If. eapply P.
eapply multi_R. apply ST_IfFalse.
apply (IHHT3 c H env0 V).
Qed.
Theorem normalization : ∀t T, has_type empty t T → halts t.
Proof.
intros.
replace t with (msubst nil t).
eapply R_halts.
eapply msubst_R; eauto. instantiate (2:= nil). eauto.
eapply V_nil.
auto.
Qed.
(* $Date: 2012-04-16 14:02:51 -0400 (Mon, 16 Apr 2012) $ *)
The has_type relation of the STLC defines what it means for a
term to belong to a type (in some context). But it doesn't, by
itself, tell us how to check whether or not a term is well
typed.
Fortunately, the rules defining has_type are syntax directed
— they exactly follow the shape of the term. This makes it
straightforward to translate the typing rules into clauses of a
typechecking function that takes a term and a context and either
returns the term's type or else signals that the term is not
typable.
Fixpoint beq_ty (T1 T2:ty) : bool :=
match T1,T2 with
| TBool, TBool =>
true
| TArrow T11 T12, TArrow T21 T22 =>
andb (beq_ty T11 T21) (beq_ty T12 T22)
| _,_ =>
false
end.
... and we need to establish the usual two-way connection between
the boolean result returned by beq_ty and the logical
proposition that its inputs are equal.
Lemma beq_ty_refl : ∀T1,
beq_ty T1 T1 = true.
Proof.
intros T1. induction T1; simpl.
reflexivity.
rewrite IHT1_1. rewrite IHT1_2. reflexivity.
Qed.
Lemma beq_ty__eq : ∀T1 T2,
beq_ty T1 T2 = true → T1 = T2.
Proof with auto.
intros T1. induction T1; intros T2 Hbeq; destruct T2; inversion Hbeq.
Case "T1=TBool".
reflexivity.
Case "T1=TArrow T1_1 T1_2".
apply andb_true in H0. inversion H0 as [Hbeq1 Hbeq2].
apply IHT1_1 in Hbeq1. apply IHT1_2 in Hbeq2. subst...
Qed.
The Typechecker
Fixpoint type_check (Γ:context) (t:tm) : option ty :=
match t with
| tvar x => Γ x
| tabs x T11 t12 => match type_check (extend Γ x T11) t12 with
| Some T12 => Some (TArrow T11 T12)
| _ => None
end
| tapp t1 t2 => match type_check Γ t1, type_check Γ t2 with
| Some (TArrow T11 T12),Some T2 =>
if beq_ty T11 T2 then Some T12 else None
| _,_ => None
end
| ttrue => Some TBool
| tfalse => Some TBool
| tif x t f => match type_check Γ x with
| Some TBool =>
match type_check Γ t, type_check Γ f with
| Some T1, Some T2 =>
if beq_ty T1 T2 then Some T1 else None
| _,_ => None
end
| _ => None
end
end.
Properties
Theorem type_checking_sound : ∀Γ t T,
type_check Γ t = Some T → has_type Γ t T.
Proof with eauto.
intros Γ t. generalize dependent Γ.
t_cases (induction t) Case; intros Γ T Htc; inversion Htc; clear Htc.
Case "tvar"...
Case "tapp".
remember (type_check Γ t1) as TO1.
remember (type_check Γ t2) as TO2.
destruct TO1 as [T1|]; try solve by inversion;
destruct T1 as [|T11 T12]; try solve by inversion.
destruct TO2 as [T2|]; try solve by inversion.
remember (beq_ty T11 T2) as b.
destruct b; try solve by inversion.
symmetry in Heqb. apply beq_ty__eq in Heqb.
inversion H0; subst...
Case "tabs".
rename i into y. rename t into T1.
remember (extend Γ y T1) as G'.
remember (type_check G' t0) as TO2.
destruct TO2; try solve by inversion.
inversion H0; subst...
Case "ttrue"...
Case "tfalse"...
Case "tif".
remember (type_check Γ t1) as TOc.
remember (type_check Γ t2) as TO1.
remember (type_check Γ t3) as TO2.
destruct TOc as [Tc|]; try solve by inversion.
destruct Tc; try solve by inversion.
destruct TO1 as [T1|]; try solve by inversion.
destruct TO2 as [T2|]; try solve by inversion.
remember (beq_ty T1 T2) as b.
destruct b; try solve by inversion.
symmetry in Heqb. apply beq_ty__eq in Heqb.
inversion H0. subst. subst...
Qed.
Theorem type_checking_complete : ∀Γ t T,
has_type Γ t T → type_check Γ t = Some T.
Proof with auto.
intros Γ t T Hty.
has_type_cases (induction Hty) Case; simpl.
Case "T_Var"...
Case "T_Abs". rewrite IHHty...
Case "T_App".
rewrite IHHty1. rewrite IHHty2.
rewrite (beq_ty_refl T1)...
Case "T_True"...
Case "T_False"...
Case "T_If". rewrite IHHty1. rewrite IHHty2.
rewrite IHHty3. rewrite (beq_ty_refl T)...
Qed.