Library MoreStlc
Simple Extensions to STLC
Numbers
let-bindings
t ::= Terms | ... (other terms same as before) | let x=t in t let-binding
(ST_Let1) let x=t1 in t2 ==> let x=t1' in t2
(ST_LetValue) let x=v1 in t2 ==> x:=v1t2 Typing: Gamma |- t1 : T1 Gamma , x:T1 |- t2 : T2
(T_Let) Gamma |- let x=t1 in t2 : T2
Pairs
\x:Nat*Nat. let sum = x.fst + x.snd in let diff = x.fst - x.snd in (sum,diff)
t ::= Terms | ... | (t,t) pair | t.fst first projection | t.snd second projection v ::= Values | ... | (v,v) pair value T ::= Types | ... | T * T product type
(ST_Pair1) (t1,t2) ==> (t1',t2)
(ST_Pair2) (v1,t2) ==> (v1,t2')
(ST_Fst1) t1.fst ==> t1'.fst
(ST_FstPair) (v1,v2).fst ==> v1
(ST_Snd1) t1.snd ==> t1'.snd
(ST_SndPair) (v1,v2).snd ==> v2
(T_Pair) Gamma |- (t1,t2) : T1*T2
(T_Fst) Gamma |- t1.fst : T11
(T_Snd) Gamma |- t1.snd : T12
Unit
t ::= Terms | ... | unit unit value v ::= Values | ... | unit unit T ::= Types | ... | Unit Unit typeTyping:
(T_Unit) Gamma |- unit : Unit
Sums
Nat + Bool
inl : Nat -> Nat + Bool inr : Bool -> Nat + Bool
div : Nat -> Nat -> (Nat + Unit) = div = \x:Nat. \y:Nat. if iszero y then inr unit else inl ...The type Nat + Unit above is in fact isomorphic to option nat in Coq, and we've already seen how to signal errors with options.
getNat = \x:Nat+Bool. case x of inl n => n | inr b => if b then 1 else 0
t ::= Terms | ... | inl T t tagging (left) | inr T t tagging (right) | case t of case inl x => t | inr x => t v ::= Values | ... | inl T v tagged value (left) | inr T v tagged value (right) T ::= Types | ... | T + T sum type
(ST_Inl) inl T t1 ==> inl T t1'
(ST_Inr) inr T t1 ==> inr T t1'
(ST_Case) case t0 of inl x1 => t1 | inr x2 => t2 ==> case t0' of inl x1 => t1 | inr x2 => t2
(ST_CaseInl) case (inl T v0) of inl x1 => t1 | inr x2 => t2 ==> x1:=v0t1
(ST_CaseInr) case (inr T v0) of inl x1 => t1 | inr x2 => t2 ==> x2:=v0t2
(T_Inl) Gamma |- inl T2 t1 : T1 + T2
(T_Inr) Gamma |- inr T1 t1 : T1 + T2
(T_Case) Gamma |- case t0 of inl x1 => t1 | inr x2 => t2 : T
Lists
\x:List Nat. lcase x of nil -> 0 | a::x' -> lcase x' of nil -> a | b::x'' -> a+b
t ::= Terms | ... | nil T | cons t t | lcase t of nil -> t | x::x -> t v ::= Values | ... | nil T nil value | cons v v cons value T ::= Types | ... | List T list of Ts
(ST_Cons1) cons t1 t2 ==> cons t1' t2
(ST_Cons2) cons v1 t2 ==> cons v1 t2'
(ST_Lcase1) (lcase t1 of nil -> t2 | xh::xt -> t3) ==> (lcase t1' of nil -> t2 | xh::xt -> t3)
(ST_LcaseNil) (lcase nil T of nil -> t2 | xh::xt -> t3) ==> t2
(ST_LcaseCons) (lcase (cons vh vt) of nil -> t2 | xh::xt -> t3) ==> xh:=vh,xt:=vtt3
(T_Nil) Gamma |- nil T : List T
(T_Cons) Gamma |- cons t1 t2: List T
(T_Lcase) Gamma |- (lcase t1 of nil -> t2 | h::t -> t3) : T
General Recursion
fact = \x:Nat. if x=0 then 1 else x * (fact (pred x)))But this would require quite a bit of work to formalize: we'd have to introduce a notion of "function definitions" and carry around an "environment" of such definitions in the definition of the step relation.
fact = fix (\f:Nat->Nat. \x:Nat. if x=0 then 1 else x * (f (pred x)))
t ::= Terms | ... | fix t fixed-point operatorReduction: t1 ==> t1'
(ST_Fix1) fix t1 ==> fix t1'
(ST_FixAbs) fix F ==> xf:=fix Ft2 Typing: Gamma |- t1 : T1->T1
(T_Fix) Gamma |- fix t1 : T1
fix F 3==> ST_FixAbs
(\x. if x=0 then 1 else x * (fix F (pred x))) 3==> ST_AppAbs
if 3=0 then 1 else 3 * (fix F (pred 3))==> ST_If0_Nonzero
3 * (fix F (pred 3))==> ST_FixAbs + ST_Mult2
3 * ((\x. if x=0 then 1 else x * (fix F (pred x))) (pred 3))==> ST_PredNat + ST_Mult2 + ST_App2
3 * ((\x. if x=0 then 1 else x * (fix F (pred x))) 2)==> ST_AppAbs + ST_Mult2
3 * (if 2=0 then 1 else 2 * (fix F (pred 2)))==> ST_If0_Nonzero + ST_Mult2
3 * (2 * (fix F (pred 2)))==> ST_FixAbs + 2 x ST_Mult2
3 * (2 * ((\x. if x=0 then 1 else x * (fix F (pred x))) (pred 2)))==> ST_PredNat + 2 x ST_Mult2 + ST_App2
3 * (2 * ((\x. if x=0 then 1 else x * (fix F (pred x))) 1))==> ST_AppAbs + 2 x ST_Mult2
3 * (2 * (if 1=0 then 1 else 1 * (fix F (pred 1))))==> ST_If0_Nonzero + 2 x ST_Mult2
3 * (2 * (1 * (fix F (pred 1))))==> ST_FixAbs + 3 x ST_Mult2
3 * (2 * (1 * ((\x. if x=0 then 1 else x * (fix F (pred x))) (pred 1))))==> ST_PredNat + 3 x ST_Mult2 + ST_App2
3 * (2 * (1 * ((\x. if x=0 then 1 else x * (fix F (pred x))) 0)))==> ST_AppAbs + 3 x ST_Mult2
3 * (2 * (1 * (if 0=0 then 1 else 0 * (fix F (pred 0)))))==> ST_If0Zero + 3 x ST_Mult2
3 * (2 * (1 * 1))==> ST_MultNats + 2 x ST_Mult2
3 * (2 * 1)==> ST_MultNats + ST_Mult2
3 * 2==> ST_MultNats
6
Exercise: 1 star (halve_fix)
Translate this informal recursive definition into one using fix:halve = \x:Nat. if x=0 then 0 else if (pred x)=0 then 0 else 1 + (halve (pred (pred x))))☐
Exercise: 1 star (fact_steps)
Write down the sequence of steps that the term fact 1 goes through to reduce to a normal form (assuming the usual reduction rules for arithmetic operations).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))
evenodd = fix (\eo: (Nat->Bool * Nat->Bool). let e = \n:Nat. if n=0 then true else eo.snd (pred n) in let o = \n:Nat. if n=0 then false else eo.fst (pred n) in (e,o)) even = evenodd.fst odd = evenodd.snd
Records
t ::= Terms | ... | {i1=t1, ..., in=tn} record | t.i projection v ::= Values | ... | {i1=v1, ..., in=vn} record value T ::= Types | ... | {i1:T1, ..., in:Tn} record typeIntuitively, the generalization is pretty obvious. But it's worth noticing that what we've actually written is rather informal: in particular, we've written "..." in several places to mean "any number of these," and we've omitted explicit mention of the usual side-condition that the labels of a record should not contain repetitions.
Reduction:
ti ==> ti'
(ST_Rcd) {i1=v1, ..., im=vm, in=ti, ...} ==> {i1=v1, ..., im=vm, in=ti', ...}
t1 ==> t1'
(ST_Proj1) t1.i ==> t1'.i
(ST_ProjRcd) {..., i=vi, ...}.i ==> vi Again, these rules are a bit informal. For example, the first rule is intended to be read "if ti is the leftmost field that is not a value and if ti steps to ti', then the whole record steps..." In the last rule, the intention is that there should only be one field called i, and that all the other fields must contain values.
Typing:
Gamma |- t1 : T1 ... Gamma |- tn : Tn
(T_Rcd) Gamma |- {i1=t1, ..., in=tn} : {i1:T1, ..., in:Tn}
Gamma |- t : {..., i:Ti, ...}
(T_Proj) Gamma |- t.i : Ti
(ST_Rcd) {i1=v1, ..., im=vm, in=ti, ...} ==> {i1=v1, ..., im=vm, in=ti', ...}
(ST_Proj1) t1.i ==> t1'.i
(ST_ProjRcd) {..., i=vi, ...}.i ==> vi Again, these rules are a bit informal. For example, the first rule is intended to be read "if ti is the leftmost field that is not a value and if ti steps to ti', then the whole record steps..." In the last rule, the intention is that there should only be one field called i, and that all the other fields must contain values.
(T_Rcd) Gamma |- {i1=t1, ..., in=tn} : {i1:T1, ..., in:Tn}
(T_Proj) Gamma |- t.i : Ti
Encoding Records (Optional)
- We can directly formalize the syntactic forms and inference
rules, staying as close as possible to the form we've given
them above. This is conceptually straightforward, and it's
probably what we'd want to do if we were building a real
compiler -- in particular, it will allow is to print error
messages in the form that programmers will find easy to
understand. But the formal versions of the rules will not be
pretty at all!
- We could look for a smoother way of presenting records -- for
example, a binary presentation with one constructor for the
empty record and another constructor for adding a single field
to an existing record, instead of a single monolithic
constructor that builds a whole record at once. This is the
right way to go if we are primarily interested in studying the
metatheory of the calculi with records, since it leads to
clean and elegant definitions and proofs. Chapter Records
shows how this can be done.
- Alternatively, if we like, we can avoid formalizing records altogether, by stipulating that record notations are just informal shorthands for more complex expressions involving pairs and product types. We sketch this approach here.
{} ----> unit {t1, t2, ..., tn} ----> (t1, trest) where {t2, ..., tn} ----> trestSimilarly, we can encode tuple types using nested product types:
{} ----> Unit {T1, T2, ..., Tn} ----> T1 * TRest where {T2, ..., Tn} ----> TRestThe operation of projecting a field from a tuple can be encoded using a sequence of second projections followed by a first projection:
t.0 ----> t.fst t.(n+1) ----> (t.snd).n
LABEL POSITION a 0 b 1 c 2 ... ... foo 1004 ... ... bar 10562 ... ...
{a=5, b=6} ----> {5,6} {a=5, c=7} ----> {5,unit,7} {c=7, a=5} ----> {5,unit,7} {c=5, b=3} ----> {unit,3,5} {f=8,c=5,a=7} ----> {7,unit,5,unit,unit,8} {f=8,c=5} ----> {unit,unit,5,unit,unit,8}Note that each field appears in the position associated with its label, that the size of the tuple is determined by the label with the highest position, and that we fill in unused positions with unit.
{a:Nat, b:Nat} ----> {Nat,Nat} {c:Nat, a:Nat} ----> {Nat,Unit,Nat} {f:Nat,c:Nat} ----> {Unit,Unit,Nat,Unit,Unit,Nat}
t.l ----> t.(position of l)
Variants (Optional Reading)
Exercise: Formalizing the Extensions
Exercise: 4 stars, advanced (STLC_extensions)
In this problem you will formalize a couple of the extensions described above. We've provided the necessary additions to the syntax of terms and types, and we've included a few examples that you can test your definitions with to make sure they are working as expected. You'll fill in the rest of the definitions and extend all the proofs accordingly.- numbers
- pairs and units
- sums
- lists
- let (which involves binding)
- fix
Inductive ty : Type :=
| TArrow : ty → ty → ty
| TNat : ty
| TUnit : ty
| TProd : ty → ty → ty
| TSum : ty → ty → ty
| TList : ty → ty.
Tactic Notation "T_cases" tactic(first) ident(c) :=
first;
[ Case_aux c "TArrow" | Case_aux c "TNat"
| Case_aux c "TProd" | Case_aux c "TUnit"
| Case_aux c "TSum" | Case_aux c "TList" ].
Inductive tm : Type :=
| tvar : id → tm
| tapp : tm → tm → tm
| tabs : id → ty → tm → tm
| tnat : nat → tm
| tsucc : tm → tm
| tpred : tm → tm
| tmult : tm → tm → tm
| tif0 : tm → tm → tm → tm
| tpair : tm → tm → tm
| tfst : tm → tm
| tsnd : tm → tm
| tunit : tm
| tlet : id → tm → tm → tm
| tinl : ty → tm → tm
| tinr : ty → tm → tm
| tcase : tm → id → tm → id → tm → tm
| tnil : ty → tm
| tcons : tm → tm → tm
| tlcase : tm → tm → id → id → tm → tm
| tfix : tm → tm.
Note that, for brevity, we've omitted booleans and instead
provided a single if0 form combining a zero test and a
conditional. That is, instead of writing
if x = 0 then ... else ...we'll write this:
if0 x then ... else ...
Tactic Notation "t_cases" tactic(first) ident(c) :=
first;
[ Case_aux c "tvar" | Case_aux c "tapp" | Case_aux c "tabs"
| Case_aux c "tnat" | Case_aux c "tsucc" | Case_aux c "tpred"
| Case_aux c "tmult" | Case_aux c "tif0"
| Case_aux c "tpair" | Case_aux c "tfst" | Case_aux c "tsnd"
| Case_aux c "tunit" | Case_aux c "tlet"
| Case_aux c "tinl" | Case_aux c "tinr" | Case_aux c "tcase"
| Case_aux c "tnil" | Case_aux c "tcons" | Case_aux c "tlcase"
| Case_aux c "tfix" ].
Fixpoint subst (x:id) (s:tm) (t:tm) : tm :=
match t with
| tvar y ⇒
if eq_id_dec x y then s else t
| tabs y T t1 ⇒
tabs y T (if eq_id_dec x y then t1 else (subst x s t1))
| tapp t1 t2 ⇒
tapp (subst x s t1) (subst x s t2)
| tnat n ⇒
tnat n
| tsucc t1 ⇒
tsucc (subst x s t1)
| tpred t1 ⇒
tpred (subst x s t1)
| tmult t1 t2 ⇒
tmult (subst x s t1) (subst x s t2)
| tif0 t1 t2 t3 ⇒
tif0 (subst x s t1) (subst x s t2) (subst x s t3)
| tpair t1 t2 ⇒
tpair (subst x s t1) (subst x s t2)
| tfst t1 ⇒
tfst (subst x s t1)
| tsnd t1 ⇒
tsnd (subst x s t1)
| tunit ⇒ tunit
| tinl T t1 ⇒
tinl T (subst x s t1)
| tinr T t1 ⇒
tinr T (subst x s t1)
| tcase t0 y1 t1 y2 t2 ⇒
tcase (subst x s t0)
y1 (if eq_id_dec x y1 then t1 else (subst x s t1))
y2 (if eq_id_dec x y2 then t2 else (subst x s t2))
| tnil T ⇒
tnil T
| tcons t1 t2 ⇒
tcons (subst x s t1) (subst x s t2)
| tlcase t1 t2 y1 y2 t3 ⇒
tlcase (subst x s t1) (subst x s t2) y1 y2
(if eq_id_dec x y1 then
t3
else if eq_id_dec x y2 then t3
else (subst x s t3))
| _ ⇒ t
end.
Notation "'[' x ':=' s ']' t" := (subst x s t) (at level 20).
Inductive value : tm → Prop :=
| v_abs : ∀ x T11 t12,
value (tabs x T11 t12)
| v_nat : ∀ n1,
value (tnat n1)
| v_pair : ∀ v1 v2,
value v1 →
value v2 →
value (tpair v1 v2)
| v_unit : value tunit
| v_inl : ∀ v T,
value v →
value (tinl T v)
| v_inr : ∀ v T,
value v →
value (tinr T v)
| v_lnil : ∀ T, value (tnil T)
| v_lcons : ∀ v1 vl,
value v1 →
value vl →
value (tcons v1 vl)
.
Hint Constructors value.
Reserved Notation "t1 '==>' t2" (at level 40).
Inductive step : tm → tm → Prop :=
| 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')
| ST_Succ1 : ∀ t1 t1',
t1 ==> t1' →
(tsucc t1) ==> (tsucc t1')
| ST_SuccNat : ∀ n1,
(tsucc (tnat n1)) ==> (tnat (S n1))
| ST_Pred : ∀ t1 t1',
t1 ==> t1' →
(tpred t1) ==> (tpred t1')
| ST_PredNat : ∀ n1,
(tpred (tnat n1)) ==> (tnat (pred n1))
| ST_Mult1 : ∀ t1 t1' t2,
t1 ==> t1' →
(tmult t1 t2) ==> (tmult t1' t2)
| ST_Mult2 : ∀ v1 t2 t2',
value v1 →
t2 ==> t2' →
(tmult v1 t2) ==> (tmult v1 t2')
| ST_MultNats : ∀ n1 n2,
(tmult (tnat n1) (tnat n2)) ==> (tnat (mult n1 n2))
| ST_If01 : ∀ t1 t1' t2 t3,
t1 ==> t1' →
(tif0 t1 t2 t3) ==> (tif0 t1' t2 t3)
| ST_If0Zero : ∀ t2 t3,
(tif0 (tnat 0) t2 t3) ==> t2
| ST_If0Nonzero : ∀ n t2 t3,
(tif0 (tnat (S n)) t2 t3) ==> t3
| ST_Pair1 : ∀ t1 t1' t2,
t1 ==> t1' →
(tpair t1 t2) ==> (tpair t1' t2)
| ST_Pair2 : ∀ v1 t2 t2',
value v1 →
t2 ==> t2' →
(tpair v1 t2) ==> (tpair v1 t2')
| ST_Fst1 : ∀ t1 t1',
t1 ==> t1' →
(tfst t1) ==> (tfst t1')
| ST_FstPair : ∀ v1 v2,
value v1 →
value v2 →
(tfst (tpair v1 v2)) ==> v1
| ST_Snd1 : ∀ t1 t1',
t1 ==> t1' →
(tsnd t1) ==> (tsnd t1')
| ST_SndPair : ∀ v1 v2,
value v1 →
value v2 →
(tsnd (tpair v1 v2)) ==> v2
| ST_Inl : ∀ t1 t1' T,
t1 ==> t1' →
(tinl T t1) ==> (tinl T t1')
| ST_Inr : ∀ t1 t1' T,
t1 ==> t1' →
(tinr T t1) ==> (tinr T t1')
| ST_Case : ∀ t0 t0' x1 t1 x2 t2,
t0 ==> t0' →
(tcase t0 x1 t1 x2 t2) ==> (tcase t0' x1 t1 x2 t2)
| ST_CaseInl : ∀ v0 x1 t1 x2 t2 T,
value v0 →
(tcase (tinl T v0) x1 t1 x2 t2) ==> [x1:=v0]t1
| ST_CaseInr : ∀ v0 x1 t1 x2 t2 T,
value v0 →
(tcase (tinr T v0) x1 t1 x2 t2) ==> [x2:=v0]t2
| ST_Cons1 : ∀ t1 t1' t2,
t1 ==> t1' →
(tcons t1 t2) ==> (tcons t1' t2)
| ST_Cons2 : ∀ v1 t2 t2',
value v1 →
t2 ==> t2' →
(tcons v1 t2) ==> (tcons v1 t2')
| ST_Lcase1 : ∀ t1 t1' t2 x1 x2 t3,
t1 ==> t1' →
(tlcase t1 t2 x1 x2 t3) ==> (tlcase t1' t2 x1 x2 t3)
| ST_LcaseNil : ∀ T t2 x1 x2 t3,
(tlcase (tnil T) t2 x1 x2 t3) ==> t2
| ST_LcaseCons : ∀ v1 vl t2 x1 x2 t3,
value v1 →
value vl →
(tlcase (tcons v1 vl) t2 x1 x2 t3) ==> (subst x2 vl (subst x1 v1 t3))
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_Succ1" | Case_aux c "ST_SuccNat"
| Case_aux c "ST_Pred1" | Case_aux c "ST_PredNat"
| Case_aux c "ST_Mult1" | Case_aux c "ST_Mult2"
| Case_aux c "ST_MultNats" | Case_aux c "ST_If01"
| Case_aux c "ST_If0Zero" | Case_aux c "ST_If0Nonzero"
| Case_aux c "ST_Pair1" | Case_aux c "ST_Pair2"
| Case_aux c "ST_Fst1" | Case_aux c "ST_FstPair"
| Case_aux c "ST_Snd1" | Case_aux c "ST_SndPair"
| Case_aux c "ST_Inl" | Case_aux c "ST_Inr" | Case_aux c "ST_Case"
| Case_aux c "ST_CaseInl" | Case_aux c "ST_CaseInr"
| Case_aux c "ST_Cons1" | Case_aux c "ST_Cons2" | Case_aux c "ST_Lcase1"
| Case_aux c "ST_LcaseNil" | Case_aux c "ST_LcaseCons"
].
Notation multistep := (multi step).
Notation "t1 '==>*' t2" := (multistep t1 t2) (at level 40).
Hint Constructors step.
Next we define the typing rules. These are nearly direct
transcriptions of the inference rules shown above.
Reserved Notation "Gamma '|-' t '\in' T" (at level 40).
Inductive has_type : context → tm → ty → Prop :=
| T_Var : ∀ Gamma x T,
Gamma x = Some T →
Gamma |- (tvar x) \in T
| T_Abs : ∀ Gamma x T11 T12 t12,
(extend Gamma x T11) |- t12 \in T12 →
Gamma |- (tabs x T11 t12) \in (TArrow T11 T12)
| T_App : ∀ T1 T2 Gamma t1 t2,
Gamma |- t1 \in (TArrow T1 T2) →
Gamma |- t2 \in T1 →
Gamma |- (tapp t1 t2) \in T2
| T_Nat : ∀ Gamma n1,
Gamma |- (tnat n1) \in TNat
| T_Succ : ∀ Gamma t1,
Gamma |- t1 \in TNat →
Gamma |- (tsucc t1) \in TNat
| T_Pred : ∀ Gamma t1,
Gamma |- t1 \in TNat →
Gamma |- (tpred t1) \in TNat
| T_Mult : ∀ Gamma t1 t2,
Gamma |- t1 \in TNat →
Gamma |- t2 \in TNat →
Gamma |- (tmult t1 t2) \in TNat
| T_If0 : ∀ Gamma t1 t2 t3 T1,
Gamma |- t1 \in TNat →
Gamma |- t2 \in T1 →
Gamma |- t3 \in T1 →
Gamma |- (tif0 t1 t2 t3) \in T1
| T_Pair : ∀ Gamma t1 t2 T1 T2,
Gamma |- t1 \in T1 →
Gamma |- t2 \in T2 →
Gamma |- (tpair t1 t2) \in (TProd T1 T2)
| T_Fst : ∀ Gamma t T1 T2,
Gamma |- t \in (TProd T1 T2) →
Gamma |- (tfst t) \in T1
| T_Snd : ∀ Gamma t T1 T2,
Gamma |- t \in (TProd T1 T2) →
Gamma |- (tsnd t) \in T2
| T_Unit : ∀ Gamma,
Gamma |- tunit \in TUnit
| T_Inl : ∀ Gamma t1 T1 T2,
Gamma |- t1 \in T1 →
Gamma |- (tinl T2 t1) \in (TSum T1 T2)
| T_Inr : ∀ Gamma t2 T1 T2,
Gamma |- t2 \in T2 →
Gamma |- (tinr T1 t2) \in (TSum T1 T2)
| T_Case : ∀ Gamma t0 x1 T1 t1 x2 T2 t2 T,
Gamma |- t0 \in (TSum T1 T2) →
(extend Gamma x1 T1) |- t1 \in T →
(extend Gamma x2 T2) |- t2 \in T →
Gamma |- (tcase t0 x1 t1 x2 t2) \in T
| T_Nil : ∀ Gamma T,
Gamma |- (tnil T) \in (TList T)
| T_Cons : ∀ Gamma t1 t2 T1,
Gamma |- t1 \in T1 →
Gamma |- t2 \in (TList T1) →
Gamma |- (tcons t1 t2) \in (TList T1)
| T_Lcase : ∀ Gamma t1 T1 t2 x1 x2 t3 T2,
Gamma |- t1 \in (TList T1) →
Gamma |- t2 \in T2 →
(extend (extend Gamma x2 (TList T1)) x1 T1) |- t3 \in T2 →
Gamma |- (tlcase t1 t2 x1 x2 t3) \in T2
where "Gamma '|-' t '\in' T" := (has_type Gamma t 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_Nat" | Case_aux c "T_Succ" | Case_aux c "T_Pred"
| Case_aux c "T_Mult" | Case_aux c "T_If0"
| Case_aux c "T_Pair" | Case_aux c "T_Fst" | Case_aux c "T_Snd"
| Case_aux c "T_Unit"
| Case_aux c "T_Inl" | Case_aux c "T_Inr" | Case_aux c "T_Case"
| Case_aux c "T_Nil" | Case_aux c "T_Cons" | Case_aux c "T_Lcase"
].
Examples
Notation a := (Id 0).
Notation f := (Id 1).
Notation g := (Id 2).
Notation l := (Id 3).
Notation k := (Id 6).
Notation i1 := (Id 7).
Notation i2 := (Id 8).
Notation x := (Id 9).
Notation y := (Id 10).
Notation processSum := (Id 11).
Notation n := (Id 12).
Notation eq := (Id 13).
Notation m := (Id 14).
Notation evenodd := (Id 15).
Notation even := (Id 16).
Notation odd := (Id 17).
Notation eo := (Id 18).
Next, a bit of Coq hackery to automate searching for typing
derivations. You don't need to understand this bit in detail --
just have a look over it so that you'll know what to look for if
you ever find yourself needing to make custom extensions to
auto.
The following Hint declarations say that, whenever auto
arrives at a goal of the form (Gamma |- (tapp e1 e1) \in T), it
should consider eapply T_App, leaving an existential variable
for the middle type T1, and similar for lcase. That variable
will then be filled in during the search for type derivations for
e1 and e2. We also include a hint to "try harder" when
solving equality goals; this is useful to automate uses of
T_Var (which includes an equality as a precondition).
Hint Extern 2 (has_type _ (tapp _ _) _) ⇒
eapply T_App; auto.
Hint Extern 2 (_ = _) ⇒ compute; reflexivity.
Module Numtest.
Definition test :=
tif0
(tpred
(tsucc
(tpred
(tmult
(tnat 2)
(tnat 0)))))
(tnat 5)
(tnat 6).
Remove the comment braces once you've implemented enough of the
definitions that you think this should work.
Module Prodtest.
Definition test :=
tsnd
(tfst
(tpair
(tpair
(tnat 5)
(tnat 6))
(tnat 7))).
End Prodtest.
Module Sumtest1.
Definition test :=
tcase (tinl TNat (tnat 5))
x (tvar x)
y (tvar y).
End Sumtest1.
Module Sumtest2.
Definition test :=
tlet
processSum
(tabs x (TSum TNat TNat)
(tcase (tvar x)
n (tvar n)
n (tif0 (tvar n) (tnat 1) (tnat 0))))
(tpair
(tapp (tvar processSum) (tinl TNat (tnat 5)))
(tapp (tvar processSum) (tinr TNat (tnat 5)))).
End Sumtest2.
Module ListTest.
Definition test :=
tlet l
(tcons (tnat 5) (tcons (tnat 6) (tnil TNat)))
(tlcase (tvar l)
(tnat 0)
x y (tmult (tvar x) (tvar x))).
End ListTest.
Module FixTest1.
Definition fact :=
tfix
(tabs f (TArrow TNat TNat)
(tabs a TNat
(tif0
(tvar a)
(tnat 1)
(tmult
(tvar a)
(tapp (tvar f) (tpred (tvar a))))))).
(Warning: you may be able to typecheck fact but still have some
rules wrong!)
End FixTest1.
Module FixTest2.
Definition map :=
tabs g (TArrow TNat TNat)
(tfix
(tabs f (TArrow (TList TNat) (TList TNat))
(tabs l (TList TNat)
(tlcase (tvar l)
(tnil TNat)
a l (tcons (tapp (tvar g) (tvar a))
(tapp (tvar f) (tvar l))))))).
End FixTest2.
Module FixTest3.
Definition equal :=
tfix
(tabs eq (TArrow TNat (TArrow TNat TNat))
(tabs m TNat
(tabs n TNat
(tif0 (tvar m)
(tif0 (tvar n) (tnat 1) (tnat 0))
(tif0 (tvar n)
(tnat 0)
(tapp (tapp (tvar eq)
(tpred (tvar m)))
(tpred (tvar n)))))))).
End FixTest3.
Module FixTest4.
Definition eotest :=
tlet evenodd
(tfix
(tabs eo (TProd (TArrow TNat TNat) (TArrow TNat TNat))
(tpair
(tabs n TNat
(tif0 (tvar n)
(tnat 1)
(tapp (tsnd (tvar eo)) (tpred (tvar n)))))
(tabs n TNat
(tif0 (tvar n)
(tnat 0)
(tapp (tfst (tvar eo)) (tpred (tvar n))))))))
(tlet even (tfst (tvar evenodd))
(tlet odd (tsnd (tvar evenodd))
(tpair
(tapp (tvar even) (tnat 3))
(tapp (tvar even) (tnat 4))))).
End FixTest4.
End Examples.
Properties of Typing
Theorem progress : ∀ t T,
empty |- t \in T →
value t ∨ ∃ t', t ==> t'.
Proof with eauto.
intros t T Ht.
remember (@empty ty) as Gamma.
generalize dependent HeqGamma.
has_type_cases (induction Ht) Case; intros HeqGamma; subst.
Case "T_Var".
inversion H.
Case "T_Abs".
left...
Case "T_App".
right.
destruct IHHt1; subst...
SCase "t1 is a value".
destruct IHHt2; subst...
SSCase "t2 is a value".
inversion H; subst; try (solve by inversion).
∃ (subst x t2 t12)...
SSCase "t2 steps".
inversion H0 as [t2' Hstp]. ∃ (tapp t1 t2')...
SCase "t1 steps".
inversion H as [t1' Hstp]. ∃ (tapp t1' t2)...
Case "T_Nat".
left...
Case "T_Succ".
right.
destruct IHHt...
SCase "t1 is a value".
inversion H; subst; try solve by inversion.
∃ (tnat (S n1))...
SCase "t1 steps".
inversion H as [t1' Hstp].
∃ (tsucc t1')...
Case "T_Pred".
right.
destruct IHHt...
SCase "t1 is a value".
inversion H; subst; try solve by inversion.
∃ (tnat (pred n1))...
SCase "t1 steps".
inversion H as [t1' Hstp].
∃ (tpred t1')...
Case "T_Mult".
right.
destruct IHHt1...
SCase "t1 is a value".
destruct IHHt2...
SSCase "t2 is a value".
inversion H; subst; try solve by inversion.
inversion H0; subst; try solve by inversion.
∃ (tnat (mult n1 n0))...
SSCase "t2 steps".
inversion H0 as [t2' Hstp].
∃ (tmult t1 t2')...
SCase "t1 steps".
inversion H as [t1' Hstp].
∃ (tmult t1' t2)...
Case "T_If0".
right.
destruct IHHt1...
SCase "t1 is a value".
inversion H; subst; try solve by inversion.
destruct n1 as [|n1'].
SSCase "n1=0".
∃ t2...
SSCase "n1<>0".
∃ t3...
SCase "t1 steps".
inversion H as [t1' H0].
∃ (tif0 t1' t2 t3)...
Case "T_Pair".
destruct IHHt1...
SCase "t1 is a value".
destruct IHHt2...
SSCase "t2 steps".
right. inversion H0 as [t2' Hstp].
∃ (tpair t1 t2')...
SCase "t1 steps".
right. inversion H as [t1' Hstp].
∃ (tpair t1' t2)...
Case "T_Fst".
right.
destruct IHHt...
SCase "t1 is a value".
inversion H; subst; try solve by inversion.
∃ v1...
SCase "t1 steps".
inversion H as [t1' Hstp].
∃ (tfst t1')...
Case "T_Snd".
right.
destruct IHHt...
SCase "t1 is a value".
inversion H; subst; try solve by inversion.
∃ v2...
SCase "t1 steps".
inversion H as [t1' Hstp].
∃ (tsnd t1')...
Case "T_Unit".
left...
Case "T_Inl".
destruct IHHt...
SCase "t1 steps".
right. inversion H as [t1' Hstp]...
Case "T_Inr".
destruct IHHt...
SCase "t1 steps".
right. inversion H as [t1' Hstp]...
Case "T_Case".
right.
destruct IHHt1...
SCase "t0 is a value".
inversion H; subst; try solve by inversion.
SSCase "t0 is inl".
∃ ([x1:=v]t1)...
SSCase "t0 is inr".
∃ ([x2:=v]t2)...
SCase "t0 steps".
inversion H as [t0' Hstp].
∃ (tcase t0' x1 t1 x2 t2)...
Case "T_Nil".
left...
Case "T_Cons".
destruct IHHt1...
SCase "head is a value".
destruct IHHt2...
SSCase "tail steps".
right. inversion H0 as [t2' Hstp].
∃ (tcons t1 t2')...
SCase "head steps".
right. inversion H as [t1' Hstp].
∃ (tcons t1' t2)...
Case "T_Lcase".
right.
destruct IHHt1...
SCase "t1 is a value".
inversion H; subst; try solve by inversion.
SSCase "t1=tnil".
∃ t2...
SSCase "t1=tcons v1 vl".
∃ ([x2:=vl]([x1:=v1]t3))...
SCase "t1 steps".
inversion H as [t1' Hstp].
∃ (tlcase t1' t2 x1 x2 t3)...
Qed.
Inductive appears_free_in : id → tm → Prop :=
| 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)
| afi_succ : ∀ x t,
appears_free_in x t →
appears_free_in x (tsucc t)
| afi_pred : ∀ x t,
appears_free_in x t →
appears_free_in x (tpred t)
| afi_mult1 : ∀ x t1 t2,
appears_free_in x t1 →
appears_free_in x (tmult t1 t2)
| afi_mult2 : ∀ x t1 t2,
appears_free_in x t2 →
appears_free_in x (tmult t1 t2)
| afi_if01 : ∀ x t1 t2 t3,
appears_free_in x t1 →
appears_free_in x (tif0 t1 t2 t3)
| afi_if02 : ∀ x t1 t2 t3,
appears_free_in x t2 →
appears_free_in x (tif0 t1 t2 t3)
| afi_if03 : ∀ x t1 t2 t3,
appears_free_in x t3 →
appears_free_in x (tif0 t1 t2 t3)
| afi_pair1 : ∀ x t1 t2,
appears_free_in x t1 →
appears_free_in x (tpair t1 t2)
| afi_pair2 : ∀ x t1 t2,
appears_free_in x t2 →
appears_free_in x (tpair t1 t2)
| afi_fst : ∀ x t,
appears_free_in x t →
appears_free_in x (tfst t)
| afi_snd : ∀ x t,
appears_free_in x t →
appears_free_in x (tsnd t)
| afi_inl : ∀ x t T,
appears_free_in x t →
appears_free_in x (tinl T t)
| afi_inr : ∀ x t T,
appears_free_in x t →
appears_free_in x (tinr T t)
| afi_case0 : ∀ x t0 x1 t1 x2 t2,
appears_free_in x t0 →
appears_free_in x (tcase t0 x1 t1 x2 t2)
| afi_case1 : ∀ x t0 x1 t1 x2 t2,
x1 ≠ x →
appears_free_in x t1 →
appears_free_in x (tcase t0 x1 t1 x2 t2)
| afi_case2 : ∀ x t0 x1 t1 x2 t2,
x2 ≠ x →
appears_free_in x t2 →
appears_free_in x (tcase t0 x1 t1 x2 t2)
| afi_cons1 : ∀ x t1 t2,
appears_free_in x t1 →
appears_free_in x (tcons t1 t2)
| afi_cons2 : ∀ x t1 t2,
appears_free_in x t2 →
appears_free_in x (tcons t1 t2)
| afi_lcase1 : ∀ x t1 t2 y1 y2 t3,
appears_free_in x t1 →
appears_free_in x (tlcase t1 t2 y1 y2 t3)
| afi_lcase2 : ∀ x t1 t2 y1 y2 t3,
appears_free_in x t2 →
appears_free_in x (tlcase t1 t2 y1 y2 t3)
| afi_lcase3 : ∀ x t1 t2 y1 y2 t3,
y1 ≠ x →
y2 ≠ x →
appears_free_in x t3 →
appears_free_in x (tlcase t1 t2 y1 y2 t3)
.
Hint Constructors appears_free_in.
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 with eauto.
intros. generalize dependent Gamma'.
has_type_cases (induction H) Case;
intros Gamma' Heqv...
Case "T_Var".
apply T_Var... rewrite <- Heqv...
Case "T_Abs".
apply T_Abs... apply IHhas_type. intros y Hafi.
unfold extend.
destruct (eq_id_dec x y)...
Case "T_Mult".
apply T_Mult...
Case "T_If0".
apply T_If0...
Case "T_Pair".
apply T_Pair...
Case "T_Case".
eapply T_Case...
apply IHhas_type2. intros y Hafi.
unfold extend.
destruct (eq_id_dec x1 y)...
apply IHhas_type3. intros y Hafi.
unfold extend.
destruct (eq_id_dec x2 y)...
Case "T_Cons".
apply T_Cons...
Case "T_Lcase".
eapply T_Lcase... apply IHhas_type3. intros y Hafi.
unfold extend.
destruct (eq_id_dec x1 y)...
destruct (eq_id_dec x2 y)...
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.
has_type_cases (induction Htyp) Case; inversion Hafi; subst...
Case "T_Abs".
destruct IHHtyp as [T' Hctx]... ∃ T'.
unfold extend in Hctx.
rewrite neq_id in Hctx...
Case "T_Case".
SCase "left".
destruct IHHtyp2 as [T' Hctx]... ∃ T'.
unfold extend in Hctx.
rewrite neq_id in Hctx...
SCase "right".
destruct IHHtyp3 as [T' Hctx]... ∃ T'.
unfold extend in Hctx.
rewrite neq_id in Hctx...
Case "T_Lcase".
clear Htyp1 IHHtyp1 Htyp2 IHHtyp2.
destruct IHHtyp3 as [T' Hctx]... ∃ T'.
unfold extend in Hctx.
rewrite neq_id in Hctx... rewrite neq_id in Hctx...
Qed.
Lemma substitution_preserves_typing : ∀ Gamma x U v t S,
(extend Gamma x U) |- t \in S →
empty |- v \in U →
Gamma |- ([x:=v]t) \in S.
Proof with eauto.
intros Gamma x U v t S Htypt Htypv.
generalize dependent Gamma. generalize dependent S.
t_cases (induction t) Case;
intros S Gamma Htypt; simpl; inversion Htypt; subst...
Case "tvar".
simpl. rename i into y.
destruct (eq_id_dec x y).
SCase "x=y".
subst.
unfold extend in H1. rewrite eq_id 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".
apply T_Var... unfold extend in H1. rewrite neq_id in H1...
Case "tabs".
rename i into y. rename t into T11.
apply T_Abs...
destruct (eq_id_dec x y).
SCase "x=y".
eapply context_invariance...
subst.
intros x Hafi. unfold extend.
destruct (eq_id_dec y x)...
SCase "x<>y".
apply IHt. eapply context_invariance...
intros z Hafi. unfold extend.
destruct (eq_id_dec y z)...
subst. rewrite neq_id...
Case "tcase".
rename i into x1. rename i0 into x2.
eapply T_Case...
SCase "left arm".
destruct (eq_id_dec x x1).
SSCase "x = x1".
eapply context_invariance...
subst.
intros z Hafi. unfold extend.
destruct (eq_id_dec x1 z)...
SSCase "x <> x1".
apply IHt2. eapply context_invariance...
intros z Hafi. unfold extend.
destruct (eq_id_dec x1 z)...
subst. rewrite neq_id...
SCase "right arm".
destruct (eq_id_dec x x2).
SSCase "x = x2".
eapply context_invariance...
subst.
intros z Hafi. unfold extend.
destruct (eq_id_dec x2 z)...
SSCase "x <> x2".
apply IHt3. eapply context_invariance...
intros z Hafi. unfold extend.
destruct (eq_id_dec x2 z)...
subst. rewrite neq_id...
Case "tlcase".
rename i into y1. rename i0 into y2.
eapply T_Lcase...
destruct (eq_id_dec x y1).
SCase "x=y1".
simpl.
eapply context_invariance...
subst.
intros z Hafi. unfold extend.
destruct (eq_id_dec y1 z)...
SCase "x<>y1".
destruct (eq_id_dec x y2).
SSCase "x=y2".
eapply context_invariance...
subst.
intros z Hafi. unfold extend.
destruct (eq_id_dec y2 z)...
SSCase "x<>y2".
apply IHt3. eapply context_invariance...
intros z Hafi. unfold extend.
destruct (eq_id_dec y1 z)...
subst. rewrite neq_id...
destruct (eq_id_dec y2 z)...
subst. rewrite neq_id...
Qed.
Theorem preservation : ∀ t t' T,
empty |- t \in T →
t ==> t' →
empty |- t' \in T.
Proof with eauto.
intros t t' T HT.
remember (@empty ty) as Gamma. generalize dependent HeqGamma.
generalize dependent t'.
has_type_cases (induction HT) Case;
intros t' HeqGamma HE; subst; inversion HE; subst...
Case "T_App".
inversion HE; subst...
SCase "ST_AppAbs".
apply substitution_preserves_typing with T1...
inversion HT1...
Case "T_Fst".
inversion HT...
Case "T_Snd".
inversion HT...
Case "T_Case".
SCase "ST_CaseInl".
inversion HT1; subst.
eapply substitution_preserves_typing...
SCase "ST_CaseInr".
inversion HT1; subst.
eapply substitution_preserves_typing...
Case "T_Lcase".
SCase "ST_LcaseCons".
inversion HT1; subst.
apply substitution_preserves_typing with (TList T1)...
apply substitution_preserves_typing with T1...
Qed.
☐