ReferencesTyping Mutable References
Set Warnings "-notation-overridden,-parsing".
Require Import Coq.Arith.Arith.
Require Import Coq.omega.Omega.
From PLF Require Import Maps.
From PLF Require Import Smallstep.
Require Import Coq.Lists.List.
Import ListNotations.
Require Import Coq.Arith.Arith.
Require Import Coq.omega.Omega.
From PLF Require Import Maps.
From PLF Require Import Smallstep.
Require Import Coq.Lists.List.
Import ListNotations.
Definitions
Syntax
Module STLCRef.
The basic operations on references are allocation,
    dereferencing, and assignment.
 
-  To allocate a reference, we use the ref operator, providing
         an initial value for the new cell.  For example, ref 5
         creates a new cell containing the value 5, and reduces to
         a reference to that cell.
-  To read the current value of this cell, we use the
         dereferencing operator !; for example, !(ref 5) reduces
         to 5.
- To change the value stored in a cell, we use the assignment operator. If r is a reference, r := 7 will store the value 7 in the cell referenced by r.
Types
      T ::= Nat
| Unit
| T → T
| Ref T 
| Unit
| T → T
| Ref T
Inductive ty : Type :=
| Nat : ty
| Unit : ty
| Arrow : ty → ty → ty
| Ref : ty → ty.
| Nat : ty
| Unit : ty
| Arrow : ty → ty → ty
| Ref : ty → ty.
Terms
      t ::= ...              Terms
          | ref t              allocation
          | !t                 dereference
          | t := t             assignment
          | l                  location
Inductive tm  : Type :=
(* STLC with numbers: *)
| var : string → tm
| app : tm → tm → tm
| abs : string → ty → tm → tm
| const : nat → tm
| scc : tm → tm
| prd : tm → tm
| mlt : tm → tm → tm
| test0 : tm → tm → tm → tm
(* New terms: *)
| unit : tm
| ref : tm → tm
| deref : tm → tm
| assign : tm → tm → tm
| loc : nat → tm.
(* STLC with numbers: *)
| var : string → tm
| app : tm → tm → tm
| abs : string → ty → tm → tm
| const : nat → tm
| scc : tm → tm
| prd : tm → tm
| mlt : tm → tm → tm
| test0 : tm → tm → tm → tm
(* New terms: *)
| unit : tm
| ref : tm → tm
| deref : tm → tm
| assign : tm → tm → tm
| loc : nat → tm.
Intuitively:
 
 
 In informal examples, we'll also freely use the extensions
    of the STLC developed in the MoreStlc chapter; however, to keep
    the proofs small, we won't bother formalizing them again here.  (It
    would be easy to do so, since there are no very interesting
    interactions between those features and references.) 
-  ref t (formally, ref t) allocates a new reference cell
      with the value t and reduces to the location of the newly
      allocated cell;
-  !t (formally, deref t) reduces to the contents of the
      cell referenced by t;
-  t1 := t2 (formally, assign t1 t2) assigns t2 to the
      cell referenced by t1; and
- l (formally, loc l) is a reference to the cell at location l. We'll discuss locations later.
Typing (Preview)
| Gamma ⊢ t1 : T1 | (T_Ref) | 
| Gamma ⊢ ref t1 : Ref T1 | 
| Gamma ⊢ t1 : Ref T11 | (T_Deref) | 
| Gamma ⊢ !t1 : T11 | 
| Gamma ⊢ t1 : Ref T11 | |
| Gamma ⊢ t2 : T11 | (T_Assign) | 
| Gamma ⊢ t1 := t2 : Unit | 
Values and Substitution
Inductive value : tm → Prop :=
| v_abs : ∀x T t,
value (abs x T t)
| v_nat : ∀n,
value (const n)
| v_unit :
value unit
| v_loc : ∀l,
value (loc l).
Hint Constructors value.
| v_abs : ∀x T t,
value (abs x T t)
| v_nat : ∀n,
value (const n)
| v_unit :
value unit
| v_loc : ∀l,
value (loc l).
Hint Constructors value.
Extending substitution to handle the new syntax of terms is
    straightforward.  
Fixpoint subst (x:string) (s:tm) (t:tm) : tm :=
match t with
| var x' ⇒
if eqb_string x x' then s else t
| app t1 t2 ⇒
app (subst x s t1) (subst x s t2)
| abs x' T t1 ⇒
if eqb_string x x' then t else abs x' T (subst x s t1)
| const n ⇒
t
| scc t1 ⇒
scc (subst x s t1)
| prd t1 ⇒
prd (subst x s t1)
| mlt t1 t2 ⇒
mlt (subst x s t1) (subst x s t2)
| test0 t1 t2 t3 ⇒
test0 (subst x s t1) (subst x s t2) (subst x s t3)
| unit ⇒
t
| ref t1 ⇒
ref (subst x s t1)
| deref t1 ⇒
deref (subst x s t1)
| assign t1 t2 ⇒
assign (subst x s t1) (subst x s t2)
| loc _ ⇒
t
end.
Notation "'[' x ':=' s ']' t" := (subst x s t) (at level 20).
match t with
| var x' ⇒
if eqb_string x x' then s else t
| app t1 t2 ⇒
app (subst x s t1) (subst x s t2)
| abs x' T t1 ⇒
if eqb_string x x' then t else abs x' T (subst x s t1)
| const n ⇒
t
| scc t1 ⇒
scc (subst x s t1)
| prd t1 ⇒
prd (subst x s t1)
| mlt t1 t2 ⇒
mlt (subst x s t1) (subst x s t2)
| test0 t1 t2 t3 ⇒
test0 (subst x s t1) (subst x s t2) (subst x s t3)
| unit ⇒
t
| ref t1 ⇒
ref (subst x s t1)
| deref t1 ⇒
deref (subst x s t1)
| assign t1 t2 ⇒
assign (subst x s t1) (subst x s t2)
| loc _ ⇒
t
end.
Notation "'[' x ':=' s ']' t" := (subst x s t) (at level 20).
Side Effects and Sequencing
       r:=succ(!r); !r
    as an abbreviation for
       (\x:Unit. !r) (r:=succ(!r)).
    This has the effect of reducing two expressions in order and
    returning the value of the second.  Restricting the type of the
    first expression to Unit helps the typechecker to catch some
    silly errors by permitting us to throw away the first value only
    if it is really guaranteed to be trivial.
       r:=succ(!r); r:=succ(!r); r:=succ(!r); r:=succ(!r); !r
 Formally, we introduce sequencing as a derived form
    tseq that expands into an abstraction and an application. 
Definition tseq t1 t2 :=
app (abs "x" Unit t2) t1.
app (abs "x" Unit t2) t1.
References and Aliasing
      let r = ref 5 in
      let s = r in
      s := 82;
      (!r)+1
    the cell referenced by r will contain the value 82, while the
    result of the whole expression will be 83.  The references r
    and s are said to be aliases for the same cell.
      r := 5; r := !s
    assigns 5 to r and then immediately overwrites it with s's
    current value; this has exactly the same effect as the single
    assignment
      r := !s
    unless we happen to do it in a context where r and s are
    aliases for the same cell! 
Shared State
      let c = ref 0 in
      let incc = \_:Unit. (c := succ (!c); !c) in
      let decc = \_:Unit. (c := pred (!c); !c) in
      ...
Objects
      newcounter =
          \_:Unit.
             let c = ref 0 in
             let incc = \_:Unit. (c := succ (!c); !c) in
             let decc = \_:Unit. (c := pred (!c); !c) in
             {i=incc, d=decc}
      let c1 = newcounter unit in
      let c2 = newcounter unit in
      // Note that we've allocated two separate storage cells now!
      let r1 = c1.i unit in
      let r2 = c2.i unit in
      r2  // yields 1, not 2!
 
Exercise: 1 star, standard, optional (store_draw)
Draw (on paper) the contents of the store at the point in execution where the first two lets have finished and the third one is about to begin.
(* FILL IN HERE *)
☐ 
References to Compound Types
      equal =
        fix
          (\eq:Nat->Nat->Bool.
             \m:Nat. \n:Nat.
               if m=0 then iszero n
               else if n=0 then false
               else eq (pred m) (pred n))
    To build a new array, we allocate a reference cell and fill
    it with a function that, when given an index, always returns 0.
      newarray = \_:Unit. ref (\n:Nat.0)
    To look up an element of an array, we simply apply
    the function to the desired index.
      lookup = \a:NatArray. \n:Nat. (!a) n
    The interesting part of the encoding is the update function.  It
    takes an array, an index, and a new value to be stored at that index, and
    does its job by creating (and storing in the reference) a new function
    that, when it is asked for the value at this very index, returns the new
    value that was given to update, while on all other indices it passes the
    lookup to the function that was previously stored in the reference.
      update = \a:NatArray. \m:Nat. \v:Nat.
                   let oldf = !a in
                   a := (\n:Nat. if equal m n then v else oldf n);
    References to values containing other references can also be very
    useful, allowing us to define data structures such as mutable
    lists and trees. 
Exercise: 2 stars, standard, recommended (compact_update)
If we defined update more compactly like this
      update = \a:NatArray. \m:Nat. \v:Nat.
                  a := (\n:Nat. if equal m n then v else (!a) n)
would it behave the same? 
(* FILL IN HERE *)
(* Do not modify the following line: *)
Definition manual_grade_for_compact_update : option (nat*string) := None.
☐ 
(* Do not modify the following line: *)
Definition manual_grade_for_compact_update : option (nat*string) := None.
Null References
Garbage Collection
Exercise: 2 stars, standard (type_safety_violation)
Show how this can lead to a violation of type safety.
(* FILL IN HERE *)
(* Do not modify the following line: *)
Definition manual_grade_for_type_safety_violation : option (nat*string) := None.
☐ 
(* Do not modify the following line: *)
Definition manual_grade_for_type_safety_violation : option (nat*string) := None.
Locations
Stores
Definition store := list tm.
We use store_lookup n st to retrieve the value of the reference
    cell at location n in the store st.  Note that we must give a
    default value to nth in case we try looking up an index which is
    too large. (In fact, we will never actually do this, but proving
    that we don't will require a bit of work.) 
Definition store_lookup (n:nat) (st:store) :=
nth n st unit.
nth n st unit.
To update the store, we use the replace function, which replaces
    the contents of a cell at a particular index. 
Fixpoint replace {A:Type} (n:nat) (x:A) (l:list A) : list A :=
match l with
| nil ⇒ nil
| h :: t ⇒
match n with
| O ⇒ x :: t
| S n' ⇒ h :: replace n' x t
end
end.
match l with
| nil ⇒ nil
| h :: t ⇒
match n with
| O ⇒ x :: t
| S n' ⇒ h :: replace n' x t
end
end.
As might be expected, we will also need some technical
    lemmas about replace; they are straightforward to prove. 
Lemma replace_nil : ∀A n (x:A),
replace n x nil = nil.
Lemma length_replace : ∀A n x (l:list A),
length (replace n x l) = length l.
Lemma lookup_replace_eq : ∀l t st,
l < length st →
store_lookup l (replace l t st) = t.
Lemma lookup_replace_neq : ∀l1 l2 t st,
l1 ≠ l2 →
store_lookup l1 (replace l2 t st) = store_lookup l1 st.
replace n x nil = nil.
Proof.
destruct n; auto.
Qed.
destruct n; auto.
Qed.
Lemma length_replace : ∀A n x (l:list A),
length (replace n x l) = length l.
Proof with auto.
intros A n x l. generalize dependent n.
induction l; intros n.
destruct n...
destruct n...
simpl. rewrite IHl...
Qed.
intros A n x l. generalize dependent n.
induction l; intros n.
destruct n...
destruct n...
simpl. rewrite IHl...
Qed.
Lemma lookup_replace_eq : ∀l t st,
l < length st →
store_lookup l (replace l t st) = t.
Proof with auto.
intros l t st.
unfold store_lookup.
generalize dependent l.
induction st as [|t' st']; intros l Hlen.
- (* st = *)
inversion Hlen.
- (* st = t' :: st' *)
destruct l; simpl...
apply IHst'. simpl in Hlen. omega.
Qed.
intros l t st.
unfold store_lookup.
generalize dependent l.
induction st as [|t' st']; intros l Hlen.
- (* st = *)
inversion Hlen.
- (* st = t' :: st' *)
destruct l; simpl...
apply IHst'. simpl in Hlen. omega.
Qed.
Lemma lookup_replace_neq : ∀l1 l2 t st,
l1 ≠ l2 →
store_lookup l1 (replace l2 t st) = store_lookup l1 st.
Proof with auto.
unfold store_lookup.
induction l1 as [|l1']; intros l2 t st Hneq.
- (* l1 = 0 *)
destruct st.
+ (* st = *) rewrite replace_nil...
+ (* st = _ :: _ *) destruct l2... contradict Hneq...
- (* l1 = S l1' *)
destruct st as [|t2 st2].
+ (* st = *) destruct l2...
+ (* st = t2 :: st2 *)
destruct l2...
simpl; apply IHl1'...
Qed.
unfold store_lookup.
induction l1 as [|l1']; intros l2 t st Hneq.
- (* l1 = 0 *)
destruct st.
+ (* st = *) rewrite replace_nil...
+ (* st = _ :: _ *) destruct l2... contradict Hneq...
- (* l1 = S l1' *)
destruct st as [|t2 st2].
+ (* st = *) destruct l2...
+ (* st = t2 :: st2 *)
destruct l2...
simpl; apply IHl1'...
Qed.
Reduction
| value v2 | (ST_AppAbs) | 
| (\x:T.t12) v2 / st --> [x:=v2]t12 / st | 
| t1 / st --> t1' / st' | (ST_App1) | 
| t1 t2 / st --> t1' t2 / st' | 
| value v1 t2 / st --> t2' / st' | (ST_App2) | 
| v1 t2 / st --> v1 t2' / st' | 
| t1 / st --> t1' / st' | (ST_Deref) | 
| !t1 / st --> !t1' / st' | 
| l < |st| | (ST_DerefLoc) | 
| !(loc l) / st --> lookup l st / st | 
| t1 / st --> t1' / st' | (ST_Assign1) | 
| t1 := t2 / st --> t1' := t2 / st' | 
| t2 / st --> t2' / st' | (ST_Assign2) | 
| v1 := t2 / st --> v1 := t2' / st' | 
| l < |st| | (ST_Assign) | 
| loc l := v2 / st --> unit / [l:=v2]st | 
| t1 / st --> t1' / st' | (ST_Ref) | 
| ref t1 / st --> ref t1' / st' | 
| (ST_RefValue) | |
| ref v1 / st --> loc |st| / st,v1 | 
Reserved Notation "t1 '/' st1 '-->' t2 '/' st2"
(at level 40, st1 at level 39, t2 at level 39).
Import ListNotations.
Inductive step : tm * store → tm * store → Prop :=
| ST_AppAbs : ∀x T t12 v2 st,
value v2 →
app (abs x T t12) v2 / st --> [x:=v2]t12 / st
| ST_App1 : ∀t1 t1' t2 st st',
t1 / st --> t1' / st' →
app t1 t2 / st --> app t1' t2 / st'
| ST_App2 : ∀v1 t2 t2' st st',
value v1 →
t2 / st --> t2' / st' →
app v1 t2 / st --> app v1 t2'/ st'
| ST_SuccNat : ∀n st,
scc (const n) / st --> const (S n) / st
| ST_Succ : ∀t1 t1' st st',
t1 / st --> t1' / st' →
scc t1 / st --> scc t1' / st'
| ST_PredNat : ∀n st,
prd (const n) / st --> const (pred n) / st
| ST_Pred : ∀t1 t1' st st',
t1 / st --> t1' / st' →
prd t1 / st --> prd t1' / st'
| ST_MultNats : ∀n1 n2 st,
mlt (const n1) (const n2) / st --> const (mult n1 n2) / st
| ST_Mult1 : ∀t1 t2 t1' st st',
t1 / st --> t1' / st' →
mlt t1 t2 / st --> mlt t1' t2 / st'
| ST_Mult2 : ∀v1 t2 t2' st st',
value v1 →
t2 / st --> t2' / st' →
mlt v1 t2 / st --> mlt v1 t2' / st'
| ST_If0 : ∀t1 t1' t2 t3 st st',
t1 / st --> t1' / st' →
test0 t1 t2 t3 / st --> test0 t1' t2 t3 / st'
| ST_If0_Zero : ∀t2 t3 st,
test0 (const 0) t2 t3 / st --> t2 / st
| ST_If0_Nonzero : ∀n t2 t3 st,
test0 (const (S n)) t2 t3 / st --> t3 / st
| ST_RefValue : ∀v1 st,
value v1 →
ref v1 / st --> loc (length st) / (st ++ v1::nil)
| ST_Ref : ∀t1 t1' st st',
t1 / st --> t1' / st' →
ref t1 / st --> ref t1' / st'
| ST_DerefLoc : ∀st l,
l < length st →
deref (loc l) / st --> store_lookup l st / st
| ST_Deref : ∀t1 t1' st st',
t1 / st --> t1' / st' →
deref t1 / st --> deref t1' / st'
| ST_Assign : ∀v2 l st,
value v2 →
l < length st →
assign (loc l) v2 / st --> unit / replace l v2 st
| ST_Assign1 : ∀t1 t1' t2 st st',
t1 / st --> t1' / st' →
assign t1 t2 / st --> assign t1' t2 / st'
| ST_Assign2 : ∀v1 t2 t2' st st',
value v1 →
t2 / st --> t2' / st' →
assign v1 t2 / st --> assign v1 t2' / st'
where "t1 '/' st1 '-->' t2 '/' st2" := (step (t1,st1) (t2,st2)).
(at level 40, st1 at level 39, t2 at level 39).
Import ListNotations.
Inductive step : tm * store → tm * store → Prop :=
| ST_AppAbs : ∀x T t12 v2 st,
value v2 →
app (abs x T t12) v2 / st --> [x:=v2]t12 / st
| ST_App1 : ∀t1 t1' t2 st st',
t1 / st --> t1' / st' →
app t1 t2 / st --> app t1' t2 / st'
| ST_App2 : ∀v1 t2 t2' st st',
value v1 →
t2 / st --> t2' / st' →
app v1 t2 / st --> app v1 t2'/ st'
| ST_SuccNat : ∀n st,
scc (const n) / st --> const (S n) / st
| ST_Succ : ∀t1 t1' st st',
t1 / st --> t1' / st' →
scc t1 / st --> scc t1' / st'
| ST_PredNat : ∀n st,
prd (const n) / st --> const (pred n) / st
| ST_Pred : ∀t1 t1' st st',
t1 / st --> t1' / st' →
prd t1 / st --> prd t1' / st'
| ST_MultNats : ∀n1 n2 st,
mlt (const n1) (const n2) / st --> const (mult n1 n2) / st
| ST_Mult1 : ∀t1 t2 t1' st st',
t1 / st --> t1' / st' →
mlt t1 t2 / st --> mlt t1' t2 / st'
| ST_Mult2 : ∀v1 t2 t2' st st',
value v1 →
t2 / st --> t2' / st' →
mlt v1 t2 / st --> mlt v1 t2' / st'
| ST_If0 : ∀t1 t1' t2 t3 st st',
t1 / st --> t1' / st' →
test0 t1 t2 t3 / st --> test0 t1' t2 t3 / st'
| ST_If0_Zero : ∀t2 t3 st,
test0 (const 0) t2 t3 / st --> t2 / st
| ST_If0_Nonzero : ∀n t2 t3 st,
test0 (const (S n)) t2 t3 / st --> t3 / st
| ST_RefValue : ∀v1 st,
value v1 →
ref v1 / st --> loc (length st) / (st ++ v1::nil)
| ST_Ref : ∀t1 t1' st st',
t1 / st --> t1' / st' →
ref t1 / st --> ref t1' / st'
| ST_DerefLoc : ∀st l,
l < length st →
deref (loc l) / st --> store_lookup l st / st
| ST_Deref : ∀t1 t1' st st',
t1 / st --> t1' / st' →
deref t1 / st --> deref t1' / st'
| ST_Assign : ∀v2 l st,
value v2 →
l < length st →
assign (loc l) v2 / st --> unit / replace l v2 st
| ST_Assign1 : ∀t1 t1' t2 st st',
t1 / st --> t1' / st' →
assign t1 t2 / st --> assign t1' t2 / st'
| ST_Assign2 : ∀v1 t2 t2' st st',
value v1 →
t2 / st --> t2' / st' →
assign v1 t2 / st --> assign v1 t2' / st'
where "t1 '/' st1 '-->' t2 '/' st2" := (step (t1,st1) (t2,st2)).
One slightly ugly point should be noted here: In the ST_RefValue
    rule, we extend the state by writing st ++ v1::nil rather than
    the more natural st ++ [v1].  The reason for this is that the
    notation we've defined for substitution uses square brackets,
    which clash with the standard library's notation for lists. 
Hint Constructors step.
Definition multistep := (multi step).
Notation "t1 '/' st '-->*' t2 '/' st'" :=
(multistep (t1,st) (t2,st'))
(at level 40, st at level 39, t2 at level 39).
Definition multistep := (multi step).
Notation "t1 '/' st '-->*' t2 '/' st'" :=
(multistep (t1,st) (t2,st'))
(at level 40, st at level 39, t2 at level 39).
Typing
Definition context := partial_map ty.
Store typings
| Gamma ⊢ lookup l st : T1 | |
| Gamma ⊢ loc l : Ref T1 | 
| Gamma; st ⊢ lookup l st : T1 | |
| Gamma; st ⊢ loc l : Ref T1 | 
[\x:Nat. (!(loc 1)) x, \x:Nat. (!(loc 0)) x]
Exercise: 2 stars, standard (cyclic_store)
Can you find a term whose reduction will create this particular cyclic store?
(* Do not modify the following line: *)
Definition manual_grade_for_cyclic_store : option (nat*string) := None.
☐ 
Definition manual_grade_for_cyclic_store : option (nat*string) := None.
Definition store_ty := list ty.
The store_Tlookup function retrieves the type at a particular
    index. 
Definition store_Tlookup (n:nat) (ST:store_ty) :=
nth n ST Unit.
nth n ST Unit.
Suppose we are given a store typing ST describing the store
    st in which some term t will be reduced.  Then we can use
    ST to calculate the type of the result of t without ever
    looking directly at st.  For example, if ST is [Unit,
    Unit→Unit], then we can immediately infer that !(loc 1) has
    type Unit→Unit.  More generally, the typing rule for locations
    can be reformulated in terms of store typings like this:
| l < |ST| | |
| Gamma; ST ⊢ loc l : Ref (lookup l ST) | 
The Typing Relation
| l < |ST| | (T_Loc) | 
| Gamma; ST ⊢ loc l : Ref (lookup l ST) | 
| Gamma; ST ⊢ t1 : T1 | (T_Ref) | 
| Gamma; ST ⊢ ref t1 : Ref T1 | 
| Gamma; ST ⊢ t1 : Ref T11 | (T_Deref) | 
| Gamma; ST ⊢ !t1 : T11 | 
| Gamma; ST ⊢ t1 : Ref T11 | |
| Gamma; ST ⊢ t2 : T11 | (T_Assign) | 
| Gamma; ST ⊢ t1 := t2 : Unit | 
Reserved Notation "Gamma ';' ST '⊢' t '∈' T" (at level 40).
Inductive has_type : context → store_ty → tm → ty → Prop :=
| T_Var : ∀Gamma ST x T,
Gamma x = Some T →
Gamma; ST ⊢ (var x) ∈ T
| T_Abs : ∀Gamma ST x T11 T12 t12,
(update Gamma x T11); ST ⊢ t12 ∈ T12 →
Gamma; ST ⊢ (abs x T11 t12) ∈ (Arrow T11 T12)
| T_App : ∀T1 T2 Gamma ST t1 t2,
Gamma; ST ⊢ t1 ∈ (Arrow T1 T2) →
Gamma; ST ⊢ t2 ∈ T1 →
Gamma; ST ⊢ (app t1 t2) ∈ T2
| T_Nat : ∀Gamma ST n,
Gamma; ST ⊢ (const n) ∈ Nat
| T_Succ : ∀Gamma ST t1,
Gamma; ST ⊢ t1 ∈ Nat →
Gamma; ST ⊢ (scc t1) ∈ Nat
| T_Pred : ∀Gamma ST t1,
Gamma; ST ⊢ t1 ∈ Nat →
Gamma; ST ⊢ (prd t1) ∈ Nat
| T_Mult : ∀Gamma ST t1 t2,
Gamma; ST ⊢ t1 ∈ Nat →
Gamma; ST ⊢ t2 ∈ Nat →
Gamma; ST ⊢ (mlt t1 t2) ∈ Nat
| T_If0 : ∀Gamma ST t1 t2 t3 T,
Gamma; ST ⊢ t1 ∈ Nat →
Gamma; ST ⊢ t2 ∈ T →
Gamma; ST ⊢ t3 ∈ T →
Gamma; ST ⊢ (test0 t1 t2 t3) ∈ T
| T_Unit : ∀Gamma ST,
Gamma; ST ⊢ unit ∈ Unit
| T_Loc : ∀Gamma ST l,
l < length ST →
Gamma; ST ⊢ (loc l) ∈ (Ref (store_Tlookup l ST))
| T_Ref : ∀Gamma ST t1 T1,
Gamma; ST ⊢ t1 ∈ T1 →
Gamma; ST ⊢ (ref t1) ∈ (Ref T1)
| T_Deref : ∀Gamma ST t1 T11,
Gamma; ST ⊢ t1 ∈ (Ref T11) →
Gamma; ST ⊢ (deref t1) ∈ T11
| T_Assign : ∀Gamma ST t1 t2 T11,
Gamma; ST ⊢ t1 ∈ (Ref T11) →
Gamma; ST ⊢ t2 ∈ T11 →
Gamma; ST ⊢ (assign t1 t2) ∈ Unit
where "Gamma ';' ST '⊢' t '∈' T" := (has_type Gamma ST t T).
Hint Constructors has_type.
Inductive has_type : context → store_ty → tm → ty → Prop :=
| T_Var : ∀Gamma ST x T,
Gamma x = Some T →
Gamma; ST ⊢ (var x) ∈ T
| T_Abs : ∀Gamma ST x T11 T12 t12,
(update Gamma x T11); ST ⊢ t12 ∈ T12 →
Gamma; ST ⊢ (abs x T11 t12) ∈ (Arrow T11 T12)
| T_App : ∀T1 T2 Gamma ST t1 t2,
Gamma; ST ⊢ t1 ∈ (Arrow T1 T2) →
Gamma; ST ⊢ t2 ∈ T1 →
Gamma; ST ⊢ (app t1 t2) ∈ T2
| T_Nat : ∀Gamma ST n,
Gamma; ST ⊢ (const n) ∈ Nat
| T_Succ : ∀Gamma ST t1,
Gamma; ST ⊢ t1 ∈ Nat →
Gamma; ST ⊢ (scc t1) ∈ Nat
| T_Pred : ∀Gamma ST t1,
Gamma; ST ⊢ t1 ∈ Nat →
Gamma; ST ⊢ (prd t1) ∈ Nat
| T_Mult : ∀Gamma ST t1 t2,
Gamma; ST ⊢ t1 ∈ Nat →
Gamma; ST ⊢ t2 ∈ Nat →
Gamma; ST ⊢ (mlt t1 t2) ∈ Nat
| T_If0 : ∀Gamma ST t1 t2 t3 T,
Gamma; ST ⊢ t1 ∈ Nat →
Gamma; ST ⊢ t2 ∈ T →
Gamma; ST ⊢ t3 ∈ T →
Gamma; ST ⊢ (test0 t1 t2 t3) ∈ T
| T_Unit : ∀Gamma ST,
Gamma; ST ⊢ unit ∈ Unit
| T_Loc : ∀Gamma ST l,
l < length ST →
Gamma; ST ⊢ (loc l) ∈ (Ref (store_Tlookup l ST))
| T_Ref : ∀Gamma ST t1 T1,
Gamma; ST ⊢ t1 ∈ T1 →
Gamma; ST ⊢ (ref t1) ∈ (Ref T1)
| T_Deref : ∀Gamma ST t1 T11,
Gamma; ST ⊢ t1 ∈ (Ref T11) →
Gamma; ST ⊢ (deref t1) ∈ T11
| T_Assign : ∀Gamma ST t1 t2 T11,
Gamma; ST ⊢ t1 ∈ (Ref T11) →
Gamma; ST ⊢ t2 ∈ T11 →
Gamma; ST ⊢ (assign t1 t2) ∈ Unit
where "Gamma ';' ST '⊢' t '∈' T" := (has_type Gamma ST t T).
Hint Constructors has_type.
Of course, these typing rules will accurately predict the results
    of reduction only if the concrete store used during reduction
    actually conforms to the store typing that we assume for purposes
    of typechecking.  This proviso exactly parallels the situation
    with free variables in the basic STLC: the substitution lemma
    promises that, if Gamma ⊢ t : T, then we can replace the free
    variables in t with values of the types listed in Gamma to
    obtain a closed term of type T, which, by the type preservation
    theorem will reduce to a final result of type T if it yields
    any result at all.  We will see below how to formalize an
    analogous intuition for stores and store typings.
 
    However, for purposes of typechecking the terms that programmers
    actually write, we do not need to do anything tricky to guess what
    store typing we should use.  Concrete locations arise only in
    terms that are the intermediate results of reduction; they are
    not in the language that programmers write.  Thus, we can simply
    typecheck the programmer's terms with respect to the empty store
    typing.  As reduction proceeds and new locations are created, we
    will always be able to see how to extend the store typing by
    looking at the type of the initial values being placed in newly
    allocated cells; this intuition is formalized in the statement of
    the type preservation theorem below.  
Properties
Well-Typed Stores
Theorem preservation_wrong1 : ∀ST T t st t' st',
empty; ST ⊢ t ∈ T →
t / st --> t' / st' →
empty; ST ⊢ t' ∈ T.
Abort.
empty; ST ⊢ t ∈ T →
t / st --> t' / st' →
empty; ST ⊢ t' ∈ T.
Abort.
If we typecheck with respect to some set of assumptions about the
    types of the values in the store and then reduce with respect to
    a store that violates these assumptions, the result will be
    disaster.  We say that a store st is well typed with respect a
    store typing ST if the term at each location l in st has the
    type at location l in ST.  Since only closed terms ever get
    stored in locations (why?), it suffices to type them in the empty
    context. The following definition of store_well_typed formalizes
    this.  
Definition store_well_typed (ST:store_ty) (st:store) :=
length ST = length st ∧
(∀l, l < length st →
empty; ST ⊢ (store_lookup l st) ∈ (store_Tlookup l ST)).
length ST = length st ∧
(∀l, l < length st →
empty; ST ⊢ (store_lookup l st) ∈ (store_Tlookup l ST)).
Informally, we will write ST ⊢ st for store_well_typed ST st. 
 
 Intuitively, a store st is consistent with a store typing
    ST if every value in the store has the type predicted by the
    store typing.  The only subtle point is the fact that, when
    typing the values in the store, we supply the very same store
    typing to the typing relation.  This allows us to type circular
    stores like the one we saw above. 
 
Exercise: 2 stars, standard (store_not_unique)
Can you find a store st, and two different store typings ST1 and ST2 such that both ST1 ⊢ st and ST2 ⊢ st?
(* FILL IN HERE *)
(* Do not modify the following line: *)
Definition manual_grade_for_store_not_unique : option (nat*string) := None.
☐ 
(* Do not modify the following line: *)
Definition manual_grade_for_store_not_unique : option (nat*string) := None.
Theorem preservation_wrong2 : ∀ST T t st t' st',
empty; ST ⊢ t ∈ T →
t / st --> t' / st' →
store_well_typed ST st →
empty; ST ⊢ t' ∈ T.
Abort.
empty; ST ⊢ t ∈ T →
t / st --> t' / st' →
store_well_typed ST st →
empty; ST ⊢ t' ∈ T.
Abort.
This statement is fine for all of the reduction rules except
    the allocation rule ST_RefValue.  The problem is that this rule
    yields a store with a larger domain than the initial store, which
    falsifies the conclusion of the above statement: if st' includes
    a binding for a fresh location l, then l cannot be in the
    domain of ST, and it will not be the case that t' (which
    definitely mentions l) is typable under ST. 
Extending Store Typings
Inductive extends : store_ty → store_ty → Prop :=
| extends_nil : ∀ST',
extends ST' nil
| extends_cons : ∀x ST' ST,
extends ST' ST →
extends (x::ST') (x::ST).
Hint Constructors extends.
| extends_nil : ∀ST',
extends ST' nil
| extends_cons : ∀x ST' ST,
extends ST' ST →
extends (x::ST') (x::ST).
Hint Constructors extends.
We'll need a few technical lemmas about extended contexts.
 
    First, looking up a type in an extended store typing yields the
    same result as in the original: 
Lemma extends_lookup : ∀l ST ST',
l < length ST →
extends ST' ST →
store_Tlookup l ST' = store_Tlookup l ST.
l < length ST →
extends ST' ST →
store_Tlookup l ST' = store_Tlookup l ST.
Proof with auto.
intros l ST ST' Hlen H.
generalize dependent ST'. generalize dependent l.
induction ST as [|a ST2]; intros l Hlen ST' HST'.
- (* nil *) inversion Hlen.
- (* cons *) unfold store_Tlookup in *.
destruct ST'.
+ (* ST' = nil *) inversion HST'.
+ (* ST' = a' :: ST'2 *)
inversion HST'; subst.
destruct l as [|l'].
* (* l = 0 *) auto.
* (* l = S l' *) simpl. apply IHST2...
simpl in Hlen; omega.
Qed.
intros l ST ST' Hlen H.
generalize dependent ST'. generalize dependent l.
induction ST as [|a ST2]; intros l Hlen ST' HST'.
- (* nil *) inversion Hlen.
- (* cons *) unfold store_Tlookup in *.
destruct ST'.
+ (* ST' = nil *) inversion HST'.
+ (* ST' = a' :: ST'2 *)
inversion HST'; subst.
destruct l as [|l'].
* (* l = 0 *) auto.
* (* l = S l' *) simpl. apply IHST2...
simpl in Hlen; omega.
Qed.
Next, if ST' extends ST, the length of ST' is at least that
    of ST. 
Lemma length_extends : ∀l ST ST',
l < length ST →
extends ST' ST →
l < length ST'.
l < length ST →
extends ST' ST →
l < length ST'.
Proof with eauto.
intros. generalize dependent l. induction H0; intros l Hlen.
inversion Hlen.
simpl in *.
destruct l; try omega.
apply lt_n_S. apply IHextends. omega.
Qed.
intros. generalize dependent l. induction H0; intros l Hlen.
inversion Hlen.
simpl in *.
destruct l; try omega.
apply lt_n_S. apply IHextends. omega.
Qed.
Finally, ST ++ T extends ST, and extends is reflexive. 
Lemma extends_app : ∀ST T,
extends (ST ++ T) ST.
Lemma extends_refl : ∀ST,
extends ST ST.
extends (ST ++ T) ST.
Proof with auto.
induction ST; intros T...
simpl...
Qed.
induction ST; intros T...
simpl...
Qed.
Lemma extends_refl : ∀ST,
extends ST ST.
Proof.
induction ST; auto.
Qed.
induction ST; auto.
Qed.
Preservation, Finally
Definition preservation_theorem := ∀ST t t' T st st',
empty; ST ⊢ t ∈ T →
store_well_typed ST st →
t / st --> t' / st' →
∃ST',
(extends ST' ST ∧
empty; ST' ⊢ t' ∈ T ∧
store_well_typed ST' st').
empty; ST ⊢ t ∈ T →
store_well_typed ST st →
t / st --> t' / st' →
∃ST',
(extends ST' ST ∧
empty; ST' ⊢ t' ∈ T ∧
store_well_typed ST' st').
Note that the preservation theorem merely asserts that there is
    some store typing ST' extending ST (i.e., agreeing with ST
    on the values of all the old locations) such that the new term
    t' is well typed with respect to ST'; it does not tell us
    exactly what ST' is.  It is intuitively clear, of course, that
    ST' is either ST or else exactly ST ++ T1::nil, where
    T1 is the type of the value v1 in the extended store st ++
    v1::nil, but stating this explicitly would complicate the statement of
    the theorem without actually making it any more useful: the weaker
    version above is already in the right form (because its conclusion
    implies its hypothesis) to "turn the crank" repeatedly and
    conclude that every sequence of reduction steps preserves
    well-typedness.  Combining this with the progress property, we
    obtain the usual guarantee that "well-typed programs never go
    wrong."
 
    In order to prove this, we'll need a few lemmas, as usual. 
Substitution Lemma
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_succ : ∀x t1,
appears_free_in x t1 →
appears_free_in x (scc t1)
| afi_pred : ∀x t1,
appears_free_in x t1 →
appears_free_in x (prd t1)
| afi_mult1 : ∀x t1 t2,
appears_free_in x t1 →
appears_free_in x (mlt t1 t2)
| afi_mult2 : ∀x t1 t2,
appears_free_in x t2 →
appears_free_in x (mlt t1 t2)
| afi_if0_1 : ∀x t1 t2 t3,
appears_free_in x t1 →
appears_free_in x (test0 t1 t2 t3)
| afi_if0_2 : ∀x t1 t2 t3,
appears_free_in x t2 →
appears_free_in x (test0 t1 t2 t3)
| afi_if0_3 : ∀x t1 t2 t3,
appears_free_in x t3 →
appears_free_in x (test0 t1 t2 t3)
| afi_ref : ∀x t1,
appears_free_in x t1 → appears_free_in x (ref t1)
| afi_deref : ∀x t1,
appears_free_in x t1 → appears_free_in x (deref t1)
| afi_assign1 : ∀x t1 t2,
appears_free_in x t1 → appears_free_in x (assign t1 t2)
| afi_assign2 : ∀x t1 t2,
appears_free_in x t2 → appears_free_in x (assign t1 t2).
Hint Constructors appears_free_in.
Lemma free_in_context : ∀x t T Gamma ST,
appears_free_in x t →
Gamma; ST ⊢ t ∈ T →
∃T', Gamma x = Some T'.
Lemma context_invariance : ∀Gamma Gamma' ST t T,
Gamma; ST ⊢ t ∈ T →
(∀x, appears_free_in x t → Gamma x = Gamma' x) →
Gamma'; ST ⊢ t ∈ T.
Lemma substitution_preserves_typing : ∀Gamma ST x s S t T,
empty; ST ⊢ s ∈ S →
(update Gamma x S); ST ⊢ t ∈ T →
Gamma; ST ⊢ ([x:=s]t) ∈ T.
| 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_succ : ∀x t1,
appears_free_in x t1 →
appears_free_in x (scc t1)
| afi_pred : ∀x t1,
appears_free_in x t1 →
appears_free_in x (prd t1)
| afi_mult1 : ∀x t1 t2,
appears_free_in x t1 →
appears_free_in x (mlt t1 t2)
| afi_mult2 : ∀x t1 t2,
appears_free_in x t2 →
appears_free_in x (mlt t1 t2)
| afi_if0_1 : ∀x t1 t2 t3,
appears_free_in x t1 →
appears_free_in x (test0 t1 t2 t3)
| afi_if0_2 : ∀x t1 t2 t3,
appears_free_in x t2 →
appears_free_in x (test0 t1 t2 t3)
| afi_if0_3 : ∀x t1 t2 t3,
appears_free_in x t3 →
appears_free_in x (test0 t1 t2 t3)
| afi_ref : ∀x t1,
appears_free_in x t1 → appears_free_in x (ref t1)
| afi_deref : ∀x t1,
appears_free_in x t1 → appears_free_in x (deref t1)
| afi_assign1 : ∀x t1 t2,
appears_free_in x t1 → appears_free_in x (assign t1 t2)
| afi_assign2 : ∀x t1 t2,
appears_free_in x t2 → appears_free_in x (assign t1 t2).
Hint Constructors appears_free_in.
Lemma free_in_context : ∀x t T Gamma ST,
appears_free_in x t →
Gamma; ST ⊢ t ∈ T →
∃T', Gamma x = Some T'.
Proof with eauto.
intros. generalize dependent Gamma. generalize dependent T.
induction H;
intros; (try solve [ inversion H0; subst; eauto ]).
- (* afi_abs *)
inversion H1; subst.
apply IHappears_free_in in H8.
rewrite update_neq in H8; assumption.
Qed.
intros. generalize dependent Gamma. generalize dependent T.
induction H;
intros; (try solve [ inversion H0; subst; eauto ]).
- (* afi_abs *)
inversion H1; subst.
apply IHappears_free_in in H8.
rewrite update_neq in H8; assumption.
Qed.
Lemma context_invariance : ∀Gamma Gamma' ST t T,
Gamma; ST ⊢ t ∈ T →
(∀x, appears_free_in x t → Gamma x = Gamma' x) →
Gamma'; ST ⊢ t ∈ T.
Proof with eauto.
intros.
generalize dependent Gamma'.
induction H; intros...
- (* T_Var *)
apply T_Var. symmetry. rewrite <- H...
- (* T_Abs *)
apply T_Abs. apply IHhas_type; intros.
unfold update, t_update.
destruct (eqb_stringP x x0)...
- (* T_App *)
eapply T_App.
apply IHhas_type1...
apply IHhas_type2...
- (* T_Mult *)
eapply T_Mult.
apply IHhas_type1...
apply IHhas_type2...
- (* T_If0 *)
eapply T_If0.
apply IHhas_type1...
apply IHhas_type2...
apply IHhas_type3...
- (* T_Assign *)
eapply T_Assign.
apply IHhas_type1...
apply IHhas_type2...
Qed.
intros.
generalize dependent Gamma'.
induction H; intros...
- (* T_Var *)
apply T_Var. symmetry. rewrite <- H...
- (* T_Abs *)
apply T_Abs. apply IHhas_type; intros.
unfold update, t_update.
destruct (eqb_stringP x x0)...
- (* T_App *)
eapply T_App.
apply IHhas_type1...
apply IHhas_type2...
- (* T_Mult *)
eapply T_Mult.
apply IHhas_type1...
apply IHhas_type2...
- (* T_If0 *)
eapply T_If0.
apply IHhas_type1...
apply IHhas_type2...
apply IHhas_type3...
- (* T_Assign *)
eapply T_Assign.
apply IHhas_type1...
apply IHhas_type2...
Qed.
Lemma substitution_preserves_typing : ∀Gamma ST x s S t T,
empty; ST ⊢ s ∈ S →
(update Gamma x S); ST ⊢ t ∈ T →
Gamma; ST ⊢ ([x:=s]t) ∈ T.
Proof with eauto.
intros Gamma ST x s S t T Hs Ht.
generalize dependent Gamma. generalize dependent T.
induction t; intros T Gamma H;
inversion H; subst; simpl...
- (* var *)
rename s0 into y.
destruct (eqb_stringP x y).
+ (* x = y *)
subst.
rewrite update_eq in H3.
inversion H3; subst.
eapply context_invariance...
intros x Hcontra.
destruct (free_in_context _ _ _ _ _ Hcontra Hs)
as [T' HT'].
inversion HT'.
+ (* x <> y *)
apply T_Var.
rewrite update_neq in H3...
- (* abs *) subst.
rename s0 into y.
destruct (eqb_stringP x y).
+ (* x = y *)
subst.
apply T_Abs. eapply context_invariance...
intros. rewrite update_shadow. reflexivity.
+ (* x <> x0 *)
apply T_Abs. apply IHt.
eapply context_invariance...
intros. unfold update, t_update.
destruct (eqb_stringP y x0)...
subst.
rewrite false_eqb_string...
Qed.
intros Gamma ST x s S t T Hs Ht.
generalize dependent Gamma. generalize dependent T.
induction t; intros T Gamma H;
inversion H; subst; simpl...
- (* var *)
rename s0 into y.
destruct (eqb_stringP x y).
+ (* x = y *)
subst.
rewrite update_eq in H3.
inversion H3; subst.
eapply context_invariance...
intros x Hcontra.
destruct (free_in_context _ _ _ _ _ Hcontra Hs)
as [T' HT'].
inversion HT'.
+ (* x <> y *)
apply T_Var.
rewrite update_neq in H3...
- (* abs *) subst.
rename s0 into y.
destruct (eqb_stringP x y).
+ (* x = y *)
subst.
apply T_Abs. eapply context_invariance...
intros. rewrite update_shadow. reflexivity.
+ (* x <> x0 *)
apply T_Abs. apply IHt.
eapply context_invariance...
intros. unfold update, t_update.
destruct (eqb_stringP y x0)...
subst.
rewrite false_eqb_string...
Qed.
Assignment Preserves Store Typing
Lemma assign_pres_store_typing : ∀ST st l t,
l < length st →
store_well_typed ST st →
empty; ST ⊢ t ∈ (store_Tlookup l ST) →
store_well_typed ST (replace l t st).
l < length st →
store_well_typed ST st →
empty; ST ⊢ t ∈ (store_Tlookup l ST) →
store_well_typed ST (replace l t st).
Proof with auto.
intros ST st l t Hlen HST Ht.
inversion HST; subst.
split. rewrite length_replace...
intros l' Hl'.
destruct (l' =? l) eqn: Heqll'.
- (* l' = l *)
apply Nat.eqb_eq in Heqll'; subst.
rewrite lookup_replace_eq...
- (* l' <> l *)
apply Nat.eqb_neq in Heqll'.
rewrite lookup_replace_neq...
rewrite length_replace in Hl'.
apply H0...
Qed.
intros ST st l t Hlen HST Ht.
inversion HST; subst.
split. rewrite length_replace...
intros l' Hl'.
destruct (l' =? l) eqn: Heqll'.
- (* l' = l *)
apply Nat.eqb_eq in Heqll'; subst.
rewrite lookup_replace_eq...
- (* l' <> l *)
apply Nat.eqb_neq in Heqll'.
rewrite lookup_replace_neq...
rewrite length_replace in Hl'.
apply H0...
Qed.
Weakening for Stores
Lemma store_weakening : ∀Gamma ST ST' t T,
extends ST' ST →
Gamma; ST ⊢ t ∈ T →
Gamma; ST' ⊢ t ∈ T.
extends ST' ST →
Gamma; ST ⊢ t ∈ T →
Gamma; ST' ⊢ t ∈ T.
Proof with eauto.
intros. induction H0; eauto.
- (* T_Loc *)
erewrite <- extends_lookup...
apply T_Loc.
eapply length_extends...
Qed.
intros. induction H0; eauto.
- (* T_Loc *)
erewrite <- extends_lookup...
apply T_Loc.
eapply length_extends...
Qed.
We can use the store_weakening lemma to prove that if a store is
    well typed with respect to a store typing, then the store extended
    with a new term t will still be well typed with respect to the
    store typing extended with t's type. 
Lemma store_well_typed_app : ∀ST st t1 T1,
store_well_typed ST st →
empty; ST ⊢ t1 ∈ T1 →
store_well_typed (ST ++ T1::nil) (st ++ t1::nil).
store_well_typed ST st →
empty; ST ⊢ t1 ∈ T1 →
store_well_typed (ST ++ T1::nil) (st ++ t1::nil).
Proof with auto.
intros.
unfold store_well_typed in *.
inversion H as [Hlen Hmatch]; clear H.
rewrite app_length, plus_comm. simpl.
rewrite app_length, plus_comm. simpl.
split...
- (* types match. *)
intros l Hl.
unfold store_lookup, store_Tlookup.
apply le_lt_eq_dec in Hl; inversion Hl as [Hlt | Heq].
+ (* l < length st *)
apply lt_S_n in Hlt.
rewrite !app_nth1...
* apply store_weakening with ST. apply extends_app.
apply Hmatch...
* rewrite Hlen...
+ (* l = length st *)
inversion Heq.
rewrite app_nth2; try omega.
rewrite <- Hlen.
rewrite minus_diag. simpl.
apply store_weakening with ST...
{ apply extends_app. }
rewrite app_nth2; try omega.
rewrite minus_diag. simpl. trivial.
Qed.
intros.
unfold store_well_typed in *.
inversion H as [Hlen Hmatch]; clear H.
rewrite app_length, plus_comm. simpl.
rewrite app_length, plus_comm. simpl.
split...
- (* types match. *)
intros l Hl.
unfold store_lookup, store_Tlookup.
apply le_lt_eq_dec in Hl; inversion Hl as [Hlt | Heq].
+ (* l < length st *)
apply lt_S_n in Hlt.
rewrite !app_nth1...
* apply store_weakening with ST. apply extends_app.
apply Hmatch...
* rewrite Hlen...
+ (* l = length st *)
inversion Heq.
rewrite app_nth2; try omega.
rewrite <- Hlen.
rewrite minus_diag. simpl.
apply store_weakening with ST...
{ apply extends_app. }
rewrite app_nth2; try omega.
rewrite minus_diag. simpl. trivial.
Qed.
Preservation!
Lemma nth_eq_last : ∀A (l:list A) x d,
nth (length l) (l ++ x::nil) d = x.
nth (length l) (l ++ x::nil) d = x.
Proof.
induction l; intros; [ auto | simpl; rewrite IHl; auto ].
Qed.
induction l; intros; [ auto | simpl; rewrite IHl; auto ].
Qed.
And here, at last, is the preservation theorem and proof: 
Theorem preservation : ∀ST t t' T st st',
empty; ST ⊢ t ∈ T →
store_well_typed ST st →
t / st --> t' / st' →
∃ST',
(extends ST' ST ∧
empty; ST' ⊢ t' ∈ T ∧
store_well_typed ST' st').
empty; ST ⊢ t ∈ T →
store_well_typed ST st →
t / st --> t' / st' →
∃ST',
(extends ST' ST ∧
empty; ST' ⊢ t' ∈ T ∧
store_well_typed ST' st').
Proof with eauto using store_weakening, extends_refl.
remember (@empty ty) as Gamma.
intros ST t t' T st st' Ht.
generalize dependent t'.
induction Ht; intros t' HST Hstep;
subst; try solve_by_invert; inversion Hstep; subst;
try (eauto using store_weakening, extends_refl).
(* T_App *)
- (* ST_AppAbs *) ∃ST.
inversion Ht1; subst.
split; try split... eapply substitution_preserves_typing...
- (* ST_App1 *)
eapply IHHt1 in H0...
inversion H0 as [ST' [Hext [Hty Hsty]]].
∃ST'...
- (* ST_App2 *)
eapply IHHt2 in H5...
inversion H5 as [ST' [Hext [Hty Hsty]]].
∃ST'...
- (* T_Succ *)
+ (* ST_Succ *)
eapply IHHt in H0...
inversion H0 as [ST' [Hext [Hty Hsty]]].
∃ST'...
- (* T_Pred *)
+ (* ST_Pred *)
eapply IHHt in H0...
inversion H0 as [ST' [Hext [Hty Hsty]]].
∃ST'...
(* T_Mult *)
- (* ST_Mult1 *)
eapply IHHt1 in H0...
inversion H0 as [ST' [Hext [Hty Hsty]]].
∃ST'...
- (* ST_Mult2 *)
eapply IHHt2 in H5...
inversion H5 as [ST' [Hext [Hty Hsty]]].
∃ST'...
- (* T_If0 *)
+ (* ST_If0_1 *)
eapply IHHt1 in H0...
inversion H0 as [ST' [Hext [Hty Hsty]]].
∃ST'... split...
(* T_Ref *)
- (* ST_RefValue *)
∃(ST ++ T1::nil).
inversion HST; subst.
split.
apply extends_app.
split.
replace (Ref T1)
with (Ref (store_Tlookup (length st) (ST ++ T1::nil))).
apply T_Loc.
rewrite <- H. rewrite app_length, plus_comm. simpl. omega.
unfold store_Tlookup. rewrite <- H. rewrite nth_eq_last.
reflexivity.
apply store_well_typed_app; assumption.
- (* ST_Ref *)
eapply IHHt in H0...
inversion H0 as [ST' [Hext [Hty Hsty]]].
∃ST'...
(* T_Deref *)
- (* ST_DerefLoc *)
∃ST. split; try split...
inversion HST as [_ Hsty].
replace T11 with (store_Tlookup l ST).
apply Hsty...
inversion Ht; subst...
- (* ST_Deref *)
eapply IHHt in H0...
inversion H0 as [ST' [Hext [Hty Hsty]]].
∃ST'...
(* T_Assign *)
- (* ST_Assign *)
∃ST. split; try split...
eapply assign_pres_store_typing...
inversion Ht1; subst...
- (* ST_Assign1 *)
eapply IHHt1 in H0...
inversion H0 as [ST' [Hext [Hty Hsty]]].
∃ST'...
- (* ST_Assign2 *)
eapply IHHt2 in H5...
inversion H5 as [ST' [Hext [Hty Hsty]]].
∃ST'...
Qed.
remember (@empty ty) as Gamma.
intros ST t t' T st st' Ht.
generalize dependent t'.
induction Ht; intros t' HST Hstep;
subst; try solve_by_invert; inversion Hstep; subst;
try (eauto using store_weakening, extends_refl).
(* T_App *)
- (* ST_AppAbs *) ∃ST.
inversion Ht1; subst.
split; try split... eapply substitution_preserves_typing...
- (* ST_App1 *)
eapply IHHt1 in H0...
inversion H0 as [ST' [Hext [Hty Hsty]]].
∃ST'...
- (* ST_App2 *)
eapply IHHt2 in H5...
inversion H5 as [ST' [Hext [Hty Hsty]]].
∃ST'...
- (* T_Succ *)
+ (* ST_Succ *)
eapply IHHt in H0...
inversion H0 as [ST' [Hext [Hty Hsty]]].
∃ST'...
- (* T_Pred *)
+ (* ST_Pred *)
eapply IHHt in H0...
inversion H0 as [ST' [Hext [Hty Hsty]]].
∃ST'...
(* T_Mult *)
- (* ST_Mult1 *)
eapply IHHt1 in H0...
inversion H0 as [ST' [Hext [Hty Hsty]]].
∃ST'...
- (* ST_Mult2 *)
eapply IHHt2 in H5...
inversion H5 as [ST' [Hext [Hty Hsty]]].
∃ST'...
- (* T_If0 *)
+ (* ST_If0_1 *)
eapply IHHt1 in H0...
inversion H0 as [ST' [Hext [Hty Hsty]]].
∃ST'... split...
(* T_Ref *)
- (* ST_RefValue *)
∃(ST ++ T1::nil).
inversion HST; subst.
split.
apply extends_app.
split.
replace (Ref T1)
with (Ref (store_Tlookup (length st) (ST ++ T1::nil))).
apply T_Loc.
rewrite <- H. rewrite app_length, plus_comm. simpl. omega.
unfold store_Tlookup. rewrite <- H. rewrite nth_eq_last.
reflexivity.
apply store_well_typed_app; assumption.
- (* ST_Ref *)
eapply IHHt in H0...
inversion H0 as [ST' [Hext [Hty Hsty]]].
∃ST'...
(* T_Deref *)
- (* ST_DerefLoc *)
∃ST. split; try split...
inversion HST as [_ Hsty].
replace T11 with (store_Tlookup l ST).
apply Hsty...
inversion Ht; subst...
- (* ST_Deref *)
eapply IHHt in H0...
inversion H0 as [ST' [Hext [Hty Hsty]]].
∃ST'...
(* T_Assign *)
- (* ST_Assign *)
∃ST. split; try split...
eapply assign_pres_store_typing...
inversion Ht1; subst...
- (* ST_Assign1 *)
eapply IHHt1 in H0...
inversion H0 as [ST' [Hext [Hty Hsty]]].
∃ST'...
- (* ST_Assign2 *)
eapply IHHt2 in H5...
inversion H5 as [ST' [Hext [Hty Hsty]]].
∃ST'...
Qed.
Exercise: 3 stars, standard (preservation_informal)
Write a careful informal proof of the preservation theorem, concentrating on the T_App, T_Deref, T_Assign, and T_Ref cases.
(* Do not modify the following line: *)
Definition manual_grade_for_preservation_informal : option (nat*string) := None.
☐ 
Definition manual_grade_for_preservation_informal : option (nat*string) := None.
Progress
Theorem progress : ∀ST t T st,
empty; ST ⊢ t ∈ T →
store_well_typed ST st →
(value t ∨ ∃t' st', t / st --> t' / st').
empty; ST ⊢ t ∈ T →
store_well_typed ST st →
(value t ∨ ∃t' st', t / st --> t' / st').
Proof with eauto.
intros ST t T st Ht HST. remember (@empty ty) as Gamma.
induction Ht; subst; try solve_by_invert...
- (* T_App *)
right. destruct IHHt1 as [Ht1p | Ht1p]...
+ (* t1 is a value *)
inversion Ht1p; subst; try solve_by_invert.
destruct IHHt2 as [Ht2p | Ht2p]...
* (* t2 steps *)
inversion Ht2p as [t2' [st' Hstep]].
∃(app (abs x T t) t2'), st'...
+ (* t1 steps *)
inversion Ht1p as [t1' [st' Hstep]].
∃(app t1' t2), st'...
- (* T_Succ *)
right. destruct IHHt as [Ht1p | Ht1p]...
+ (* t1 is a value *)
inversion Ht1p; subst; try solve [ inversion Ht ].
* (* t1 is a const *)
∃(const (S n)), st...
+ (* t1 steps *)
inversion Ht1p as [t1' [st' Hstep]].
∃(scc t1'), st'...
- (* T_Pred *)
right. destruct IHHt as [Ht1p | Ht1p]...
+ (* t1 is a value *)
inversion Ht1p; subst; try solve [inversion Ht ].
* (* t1 is a const *)
∃(const (pred n)), st...
+ (* t1 steps *)
inversion Ht1p as [t1' [st' Hstep]].
∃(prd t1'), st'...
- (* T_Mult *)
right. destruct IHHt1 as [Ht1p | Ht1p]...
+ (* t1 is a value *)
inversion Ht1p; subst; try solve [inversion Ht1].
destruct IHHt2 as [Ht2p | Ht2p]...
* (* t2 is a value *)
inversion Ht2p; subst; try solve [inversion Ht2].
∃(const (mult n n0)), st...
* (* t2 steps *)
inversion Ht2p as [t2' [st' Hstep]].
∃(mlt (const n) t2'), st'...
+ (* t1 steps *)
inversion Ht1p as [t1' [st' Hstep]].
∃(mlt t1' t2), st'...
- (* T_If0 *)
right. destruct IHHt1 as [Ht1p | Ht1p]...
+ (* t1 is a value *)
inversion Ht1p; subst; try solve [inversion Ht1].
destruct n.
* (* n = 0 *) ∃t2, st...
* (* n = S n' *) ∃t3, st...
+ (* t1 steps *)
inversion Ht1p as [t1' [st' Hstep]].
∃(test0 t1' t2 t3), st'...
- (* T_Ref *)
right. destruct IHHt as [Ht1p | Ht1p]...
+ (* t1 steps *)
inversion Ht1p as [t1' [st' Hstep]].
∃(ref t1'), st'...
- (* T_Deref *)
right. destruct IHHt as [Ht1p | Ht1p]...
+ (* t1 is a value *)
inversion Ht1p; subst; try solve_by_invert.
eexists. eexists. apply ST_DerefLoc...
inversion Ht; subst. inversion HST; subst.
rewrite <- H...
+ (* t1 steps *)
inversion Ht1p as [t1' [st' Hstep]].
∃(deref t1'), st'...
- (* T_Assign *)
right. destruct IHHt1 as [Ht1p|Ht1p]...
+ (* t1 is a value *)
destruct IHHt2 as [Ht2p|Ht2p]...
* (* t2 is a value *)
inversion Ht1p; subst; try solve_by_invert.
eexists. eexists. apply ST_Assign...
inversion HST; subst. inversion Ht1; subst.
rewrite H in H5...
* (* t2 steps *)
inversion Ht2p as [t2' [st' Hstep]].
∃(assign t1 t2'), st'...
+ (* t1 steps *)
inversion Ht1p as [t1' [st' Hstep]].
∃(assign t1' t2), st'...
Qed.
intros ST t T st Ht HST. remember (@empty ty) as Gamma.
induction Ht; subst; try solve_by_invert...
- (* T_App *)
right. destruct IHHt1 as [Ht1p | Ht1p]...
+ (* t1 is a value *)
inversion Ht1p; subst; try solve_by_invert.
destruct IHHt2 as [Ht2p | Ht2p]...
* (* t2 steps *)
inversion Ht2p as [t2' [st' Hstep]].
∃(app (abs x T t) t2'), st'...
+ (* t1 steps *)
inversion Ht1p as [t1' [st' Hstep]].
∃(app t1' t2), st'...
- (* T_Succ *)
right. destruct IHHt as [Ht1p | Ht1p]...
+ (* t1 is a value *)
inversion Ht1p; subst; try solve [ inversion Ht ].
* (* t1 is a const *)
∃(const (S n)), st...
+ (* t1 steps *)
inversion Ht1p as [t1' [st' Hstep]].
∃(scc t1'), st'...
- (* T_Pred *)
right. destruct IHHt as [Ht1p | Ht1p]...
+ (* t1 is a value *)
inversion Ht1p; subst; try solve [inversion Ht ].
* (* t1 is a const *)
∃(const (pred n)), st...
+ (* t1 steps *)
inversion Ht1p as [t1' [st' Hstep]].
∃(prd t1'), st'...
- (* T_Mult *)
right. destruct IHHt1 as [Ht1p | Ht1p]...
+ (* t1 is a value *)
inversion Ht1p; subst; try solve [inversion Ht1].
destruct IHHt2 as [Ht2p | Ht2p]...
* (* t2 is a value *)
inversion Ht2p; subst; try solve [inversion Ht2].
∃(const (mult n n0)), st...
* (* t2 steps *)
inversion Ht2p as [t2' [st' Hstep]].
∃(mlt (const n) t2'), st'...
+ (* t1 steps *)
inversion Ht1p as [t1' [st' Hstep]].
∃(mlt t1' t2), st'...
- (* T_If0 *)
right. destruct IHHt1 as [Ht1p | Ht1p]...
+ (* t1 is a value *)
inversion Ht1p; subst; try solve [inversion Ht1].
destruct n.
* (* n = 0 *) ∃t2, st...
* (* n = S n' *) ∃t3, st...
+ (* t1 steps *)
inversion Ht1p as [t1' [st' Hstep]].
∃(test0 t1' t2 t3), st'...
- (* T_Ref *)
right. destruct IHHt as [Ht1p | Ht1p]...
+ (* t1 steps *)
inversion Ht1p as [t1' [st' Hstep]].
∃(ref t1'), st'...
- (* T_Deref *)
right. destruct IHHt as [Ht1p | Ht1p]...
+ (* t1 is a value *)
inversion Ht1p; subst; try solve_by_invert.
eexists. eexists. apply ST_DerefLoc...
inversion Ht; subst. inversion HST; subst.
rewrite <- H...
+ (* t1 steps *)
inversion Ht1p as [t1' [st' Hstep]].
∃(deref t1'), st'...
- (* T_Assign *)
right. destruct IHHt1 as [Ht1p|Ht1p]...
+ (* t1 is a value *)
destruct IHHt2 as [Ht2p|Ht2p]...
* (* t2 is a value *)
inversion Ht1p; subst; try solve_by_invert.
eexists. eexists. apply ST_Assign...
inversion HST; subst. inversion Ht1; subst.
rewrite H in H5...
* (* t2 steps *)
inversion Ht2p as [t2' [st' Hstep]].
∃(assign t1 t2'), st'...
+ (* t1 steps *)
inversion Ht1p as [t1' [st' Hstep]].
∃(assign t1' t2), st'...
Qed.
References and Nontermination
   (\r:Ref (Unit -> Unit).
        r := (\x:Unit.(!r) unit); (!r) unit)
   (ref (\x:Unit.unit))
   First, ref (\x:Unit.unit) creates a reference to a cell of type
   Unit → Unit.  We then pass this reference as the argument to a
   function which binds it to the name r, and assigns to it the
   function \x:Unit.(!r) unit — that is, the function which ignores
   its argument and calls the function stored in r on the argument
   unit; but of course, that function is itself!  To start the
   divergent loop, we execute the function stored in the cell by
   evaluating (!r) unit.
Module ExampleVariables.
Open Scope string_scope.
Definition x := "x".
Definition y := "y".
Definition r := "r".
Definition s := "s".
End ExampleVariables.
Module RefsAndNontermination.
Import ExampleVariables.
Definition loop_fun :=
abs x Unit (app (deref (var r)) unit).
Definition loop :=
app
(abs r (Ref (Arrow Unit Unit))
(tseq (assign (var r) loop_fun)
(app (deref (var r)) unit)))
(ref (abs x Unit unit)).
Open Scope string_scope.
Definition x := "x".
Definition y := "y".
Definition r := "r".
Definition s := "s".
End ExampleVariables.
Module RefsAndNontermination.
Import ExampleVariables.
Definition loop_fun :=
abs x Unit (app (deref (var r)) unit).
Definition loop :=
app
(abs r (Ref (Arrow Unit Unit))
(tseq (assign (var r) loop_fun)
(app (deref (var r)) unit)))
(ref (abs x Unit unit)).
This term is well typed: 
Lemma loop_typeable : ∃T, empty; nil ⊢ loop ∈ T.
Proof with eauto.
eexists. unfold loop. unfold loop_fun.
eapply T_App...
eapply T_Abs...
eapply T_App...
eapply T_Abs. eapply T_App. eapply T_Deref. eapply T_Var.
unfold update, t_update. simpl. reflexivity. auto.
eapply T_Assign.
eapply T_Var. unfold update, t_update. simpl. reflexivity.
eapply T_Abs.
eapply T_App...
eapply T_Deref. eapply T_Var. reflexivity.
Qed.
eexists. unfold loop. unfold loop_fun.
eapply T_App...
eapply T_Abs...
eapply T_App...
eapply T_Abs. eapply T_App. eapply T_Deref. eapply T_Var.
unfold update, t_update. simpl. reflexivity. auto.
eapply T_Assign.
eapply T_Var. unfold update, t_update. simpl. reflexivity.
eapply T_Abs.
eapply T_App...
eapply T_Deref. eapply T_Var. reflexivity.
Qed.
To show formally that the term diverges, we first define the
    step_closure of the single-step reduction relation, written
    -->+.  This is just like the reflexive step closure of
    single-step reduction (which we're been writing -->*), except
    that it is not reflexive: t -->+ t' means that t can reach
    t' by one or more steps of reduction. 
Inductive step_closure {X:Type} (R: relation X) : X → X → Prop :=
| sc_one : ∀(x y : X),
R x y → step_closure R x y
| sc_step : ∀(x y z : X),
R x y →
step_closure R y z →
step_closure R x z.
Definition multistep1 := (step_closure step).
Notation "t1 '/' st '-->+' t2 '/' st'" :=
(multistep1 (t1,st) (t2,st'))
(at level 40, st at level 39, t2 at level 39).
| sc_one : ∀(x y : X),
R x y → step_closure R x y
| sc_step : ∀(x y z : X),
R x y →
step_closure R y z →
step_closure R x z.
Definition multistep1 := (step_closure step).
Notation "t1 '/' st '-->+' t2 '/' st'" :=
(multistep1 (t1,st) (t2,st'))
(at level 40, st at level 39, t2 at level 39).
Now, we can show that the expression loop reduces to the
    expression !(loc 0) unit and the size-one store
    [r:=(loc 0)]loop_fun. 
 
 As a convenience, we introduce a slight variant of the normalize
    tactic, called reduce, which tries solving the goal with
    multi_refl at each step, instead of waiting until the goal can't
    be reduced any more. Of course, the whole point is that loop
    doesn't normalize, so the old normalize tactic would just go
    into an infinite loop reducing it forever! 
Ltac print_goal := match goal with ⊢ ?x ⇒ idtac x end.
Ltac reduce :=
repeat (print_goal; eapply multi_step ;
[ (eauto 10; fail) | (instantiate; compute)];
try solve [apply multi_refl]).
Ltac reduce :=
repeat (print_goal; eapply multi_step ;
[ (eauto 10; fail) | (instantiate; compute)];
try solve [apply multi_refl]).
Next, we use reduce to show that loop steps to
    !(loc 0) unit, starting from the empty store. 
Lemma loop_steps_to_loop_fun :
loop / nil -->*
app (deref (loc 0)) unit / cons ([r:=loc 0]loop_fun) nil.
Proof.
unfold loop.
reduce.
Qed.
loop / nil -->*
app (deref (loc 0)) unit / cons ([r:=loc 0]loop_fun) nil.
Proof.
unfold loop.
reduce.
Qed.
Finally, we show that the latter expression reduces in
    two steps to itself! 
Lemma loop_fun_step_self :
app (deref (loc 0)) unit / cons ([r:=loc 0]loop_fun) nil -->+
app (deref (loc 0)) unit / cons ([r:=loc 0]loop_fun) nil.
app (deref (loc 0)) unit / cons ([r:=loc 0]loop_fun) nil -->+
app (deref (loc 0)) unit / cons ([r:=loc 0]loop_fun) nil.
Proof with eauto.
unfold loop_fun; simpl.
eapply sc_step. apply ST_App1...
eapply sc_one. compute. apply ST_AppAbs...
Qed.
unfold loop_fun; simpl.
eapply sc_step. apply ST_App1...
eapply sc_one. compute. apply ST_AppAbs...
Qed.
Exercise: 4 stars, standard (factorial_ref)
Use the above ideas to implement a factorial function in STLC with references. (There is no need to prove formally that it really behaves like the factorial. Just uncomment the example below to make sure it gives the correct result when applied to the argument 4.)
Definition factorial : tm
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
Lemma factorial_type : empty; nil ⊢ factorial ∈ (Arrow Nat Nat).
Proof with eauto.
(* FILL IN HERE *) Admitted.
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
Lemma factorial_type : empty; nil ⊢ factorial ∈ (Arrow Nat Nat).
Proof with eauto.
(* FILL IN HERE *) Admitted.
If your definition is correct, you should be able to just
    uncomment the example below; the proof should be fully
    automatic using the reduce tactic. 
(* 
Lemma factorial_4 : exists st,
app factorial (const 4) / nil -->* const 24 / st.
Proof.
eexists. unfold factorial. reduce.
Qed.
*)
☐ 
Lemma factorial_4 : exists st,
app factorial (const 4) / nil -->* const 24 / st.
Proof.
eexists. unfold factorial. reduce.
Qed.
*)
Additional Exercises
Exercise: 5 stars, standard, optional (garabage_collector)
Challenge problem: modify our formalization to include an account of garbage collection, and prove that it satisfies whatever nice properties you can think to prove about it.
End RefsAndNontermination.
End STLCRef.
(* Sat 22 Dec 2018 10:17:49 EST *)
End STLCRef.
(* Sat 22 Dec 2018 10:17:49 EST *)
