RecordSub: レコードのサブタイプ


In this chapter, we combine two significant extensions of the pure STLC -- records (from chapter Records) and subtyping (from chapter Sub) -- and explore their interactions. Most of the concepts have already been discussed in those chapters, so the presentation here is somewhat terse. We just comment where things are nonstandard.

Set Warnings "-notation-overridden,-parsing".
From Coq Require Import Strings.String.
From PLF Require Import Maps.
From PLF Require Import Smallstep.
From PLF Require Import MoreStlc.

中核部の定義


構文


Inductive ty : Type :=
  
  | Top : ty
  | Base : string -> ty
  | Arrow : ty -> ty -> ty
  
  | RNil : ty
  | RCons : string -> ty -> ty -> ty.

Inductive tm : Type :=
  
  | var : string -> tm
  | app : tm -> tm -> tm
  | abs : string -> ty -> tm -> tm
  | rproj : tm -> string -> tm
  
  | rnil : tm
  | rcons : string -> tm -> tm -> tm.

Well-Formedness

The syntax of terms and types is a bit too loose, in the sense that it admits things like a record type whose final "tail" is Top or some arrow type rather than Nil. To avoid such cases, it is useful to assume that all the record types and terms that we see will obey some simple well-formedness conditions.
An interesting technical question is whether the basic properties of the system -- progress and preservation -- remain true if we drop these conditions. I believe they do, and I would encourage motivated readers to try to check this by dropping the conditions from the definitions of typing and subtyping and adjusting the proofs in the rest of the chapter accordingly. This is not a trivial exercise (or I'd have done it!), but it should not involve changing the basic structure of the proofs. If someone does do it, please let me know. --BCP 5/16.

Inductive record_ty : ty -> Prop :=
  | RTnil :
        record_ty RNil
  | RTcons : forall i T1 T2,
        record_ty (RCons i T1 T2).

Inductive record_tm : tm -> Prop :=
  | rtnil :
        record_tm rnil
  | rtcons : forall i t1 t2,
        record_tm (rcons i t1 t2).

Inductive well_formed_ty : ty -> Prop :=
  | wfTop :
        well_formed_ty Top
  | wfBase : forall i,
        well_formed_ty (Base i)
  | wfArrow : forall T1 T2,
        well_formed_ty T1 ->
        well_formed_ty T2 ->
        well_formed_ty (Arrow T1 T2)
  | wfRNil :
        well_formed_ty RNil
  | wfRCons : forall i T1 T2,
        well_formed_ty T1 ->
        well_formed_ty T2 ->
        record_ty T2 ->
        well_formed_ty (RCons i T1 T2).

Hint Constructors record_ty record_tm well_formed_ty.

置換

Substitution and reduction are as before.

Fixpoint subst (x:string) (s:tm) (t:tm) : tm :=
  match t with
  | var y => if eqb_string x y then s else t
  | abs y T t1 => abs y T (if eqb_string x y then t1
                             else (subst x s t1))
  | app t1 t2 => app (subst x s t1) (subst x s t2)
  | rproj t1 i => rproj (subst x s t1) i
  | rnil => rnil
  | rcons i t1 tr2 => rcons i (subst x s t1) (subst x s tr2)
  end.

Notation "'[' x ':=' s ']' t" := (subst x s t) (at level 20).

簡約


Inductive value : tm -> Prop :=
  | v_abs : forall x T t,
      value (abs x T t)
  | v_rnil : value rnil
  | v_rcons : forall i v vr,
      value v ->
      value vr ->
      value (rcons i v vr).

Hint Constructors value.

Fixpoint Tlookup (i:string) (Tr:ty) : option ty :=
  match Tr with
  | RCons i' T Tr' =>
      if eqb_string i i' then Some T else Tlookup i Tr'
  | _ => None
  end.

Fixpoint tlookup (i:string) (tr:tm) : option tm :=
  match tr with
  | rcons i' t tr' =>
      if eqb_string i i' then Some t else tlookup i tr'
  | _ => None
  end.

Reserved Notation "t1 '-->' t2" (at level 40).

Inductive step : tm -> tm -> Prop :=
  | ST_AppAbs : forall x T t12 v2,
         value v2 ->
         (app (abs x T t12) v2) --> [x:=v2]t12
  | ST_App1 : forall t1 t1' t2,
         t1 --> t1' ->
         (app t1 t2) --> (app t1' t2)
  | ST_App2 : forall v1 t2 t2',
         value v1 ->
         t2 --> t2' ->
         (app v1 t2) --> (app v1 t2')
  | ST_Proj1 : forall tr tr' i,
        tr --> tr' ->
        (rproj tr i) --> (rproj tr' i)
  | ST_ProjRcd : forall tr i vi,
        value tr ->
        tlookup i tr = Some vi ->
       (rproj tr i) --> vi
  | ST_Rcd_Head : forall i t1 t1' tr2,
        t1 --> t1' ->
        (rcons i t1 tr2) --> (rcons i t1' tr2)
  | ST_Rcd_Tail : forall i v1 tr2 tr2',
        value v1 ->
        tr2 --> tr2' ->
        (rcons i v1 tr2) --> (rcons i v1 tr2')

where "t1 '-->' t2" := (step t1 t2).

Hint Constructors step.

サブタイプ


追加した要素が関係してくる、おもしろい部分に来ました。 サブタイプ関係を定義し、その重要な技術的性質のいくつかを調べることから始めます。

定義


サブタイプの定義は、本質的には、Sub章で議論したレコードのサブタイプの通りです。 ただ、いくつかの規則に付加条件として well-formedness を追加する必要があります。 また、「n引数」の幅、深さ、順列などの規則を一つ目のフィールドを対象にした2引数のものに置き換えます。

Reserved Notation "T '<:' U" (at level 40).

Inductive subtype : ty -> ty -> Prop :=
  
  | S_Refl : forall T,
    well_formed_ty T ->
    T <: T
  | S_Trans : forall S U T,
    S <: U ->
    U <: T ->
    S <: T
  | S_Top : forall S,
    well_formed_ty S ->
    S <: Top
  | S_Arrow : forall S1 S2 T1 T2,
    T1 <: S1 ->
    S2 <: T2 ->
    Arrow S1 S2 <: Arrow T1 T2
  
  | S_RcdWidth : forall i T1 T2,
    well_formed_ty (RCons i T1 T2) ->
    RCons i T1 T2 <: RNil
  | S_RcdDepth : forall i S1 T1 Sr2 Tr2,
    S1 <: T1 ->
    Sr2 <: Tr2 ->
    record_ty Sr2 ->
    record_ty Tr2 ->
    RCons i S1 Sr2 <: RCons i T1 Tr2
  | S_RcdPerm : forall i1 i2 T1 T2 Tr3,
    well_formed_ty (RCons i1 T1 (RCons i2 T2 Tr3)) ->
    i1 <> i2 ->
       RCons i1 T1 (RCons i2 T2 Tr3)
    <: RCons i2 T2 (RCons i1 T1 Tr3)

where "T '<:' U" := (subtype T U).

Hint Constructors subtype.


Module Examples.
Open Scope string_scope.

Notation x := "x".
Notation y := "y".
Notation z := "z".
Notation j := "j".
Notation k := "k".
Notation i := "i".
Notation A := (Base "A").
Notation B := (Base "B").
Notation C := (Base "C").

Definition TRcd_j :=
  (RCons j (Arrow B B) RNil). Definition TRcd_kj :=
  RCons k (Arrow A A) TRcd_j.
Example subtyping_example_0 :
  subtype (Arrow C TRcd_kj)
          (Arrow C RNil).
Proof.
  apply S_Arrow.
    apply S_Refl. auto.
    unfold TRcd_kj, TRcd_j. apply S_RcdWidth; auto.
Qed.

以下の事実のほとんどはCoqで証明することは簡単です。 最大限理解するために、どのように証明するかを理解していることを紙の上でも確認しなさい!

練習問題: ★★, standard (subtyping_example_1)

Example subtyping_example_1 :
  subtype TRcd_kj TRcd_j.
Proof with eauto.
Admitted.

練習問題: ★, standard (subtyping_example_2)

Example subtyping_example_2 :
  subtype (Arrow Top TRcd_kj)
          (Arrow (Arrow C C) TRcd_j).
Proof with eauto.
Admitted.

練習問題: ★, standard (subtyping_example_3)

Example subtyping_example_3 :
  subtype (Arrow RNil (RCons j A RNil))
          (Arrow (RCons k B RNil) RNil).
Proof with eauto.
Admitted.

練習問題: ★★, standard (subtyping_example_4)

Example subtyping_example_4 :
  subtype (RCons x A (RCons y B (RCons z C RNil)))
          (RCons z C (RCons y B (RCons x A RNil))).
Proof with eauto.
Admitted.

End Examples.

サブタイプの性質


Well-Formedness

To get started proving things about subtyping, we need a couple of technical lemmas that intuitively (1) allow us to extract the well-formedness assumptions embedded in subtyping derivations and (2) record the fact that fields of well-formed record types are themselves well-formed types.

Lemma subtype__wf : forall S T,
  subtype S T ->
  well_formed_ty T /\ well_formed_ty S.
Proof with eauto.
  intros S T Hsub.
  induction Hsub;
    intros; try (destruct IHHsub1; destruct IHHsub2)...
  -
    split... inversion H. subst. inversion H5... Qed.

Lemma wf_rcd_lookup : forall i T Ti,
  well_formed_ty T ->
  Tlookup i T = Some Ti ->
  well_formed_ty Ti.
Proof with eauto.
  intros i T.
  induction T; intros; try solve_by_invert.
  -
    inversion H. subst. unfold Tlookup in H0.
    destruct (eqb_string i s)... inversion H0; subst... Qed.

フィールド参照


サブタイプがあることで、レコードマッチング補題はいくらか複雑になります。 それには2つの理由があります。 1つはレコード型が対応する項の正確な構造を記述する必要がなくなることです。 2つ目は、has_typeの導出に関する帰納法を使う推論が一般には難しくなることです。 なぜなら、has_typeが構文指向ではなくなるからです。

Lemma rcd_types_match : forall S T i Ti,
  subtype S T ->
  Tlookup i T = Some Ti ->
  exists Si, Tlookup i S = Some Si /\ subtype Si Ti.
Proof with (eauto using wf_rcd_lookup).
  intros S T i Ti Hsub Hget. generalize dependent Ti.
  induction Hsub; intros Ti Hget;
    try solve_by_invert.
  -
    exists Ti...
  -
    destruct (IHHsub2 Ti) as [Ui Hui]... destruct Hui.
    destruct (IHHsub1 Ui) as [Si Hsi]... destruct Hsi.
    exists Si...
  -
    rename i0 into k.
    unfold Tlookup. unfold Tlookup in Hget.
    destruct (eqb_string i k)...
    +
      inversion Hget. subst. exists S1...
  -
    exists Ti. split.
    +
      unfold Tlookup. unfold Tlookup in Hget.
      destruct (eqb_stringP i i1)...
      *
        destruct (eqb_stringP i i2)...
        destruct H0.
        subst...
    +
      inversion H. subst. inversion H5. subst... Qed.

練習問題: ★★★, standard (rcd_types_match_informal)

rcd_types_match補題の非形式的証明を注意深く記述しなさい。

反転補題


練習問題: ★★★, standard, optional (sub_inversion_arrow)

Lemma sub_inversion_arrow : forall U V1 V2,
     subtype U (Arrow V1 V2) ->
     exists U1 U2,
       (U=(Arrow U1 U2)) /\ (subtype V1 U1) /\ (subtype U2 V2).
Proof with eauto.
  intros U V1 V2 Hs.
  remember (Arrow V1 V2) as V.
  generalize dependent V2. generalize dependent V1.
Admitted.

型付け


Definition context := partial_map ty.

Reserved Notation "Gamma '|-' t '\in' T" (at level 40).

Inductive has_type : context -> tm -> ty -> Prop :=
  | T_Var : forall Gamma x T,
      Gamma x = Some T ->
      well_formed_ty T ->
      Gamma |- var x \in T
  | T_Abs : forall Gamma x T11 T12 t12,
      well_formed_ty T11 ->
      update Gamma x T11 |- t12 \in T12 ->
      Gamma |- abs x T11 t12 \in Arrow T11 T12
  | T_App : forall T1 T2 Gamma t1 t2,
      Gamma |- t1 \in Arrow T1 T2 ->
      Gamma |- t2 \in T1 ->
      Gamma |- app t1 t2 \in T2
  | T_Proj : forall Gamma i t T Ti,
      Gamma |- t \in T ->
      Tlookup i T = Some Ti ->
      Gamma |- rproj t i \in Ti
  
  | T_Sub : forall Gamma t S T,
      Gamma |- t \in S ->
      subtype S T ->
      Gamma |- t \in T
  
  | T_RNil : forall Gamma,
      Gamma |- rnil \in RNil
  | T_RCons : forall Gamma i t T tr Tr,
      Gamma |- t \in T ->
      Gamma |- tr \in Tr ->
      record_ty Tr ->
      record_tm tr ->
      Gamma |- rcons i t tr \in RCons i T Tr

where "Gamma '|-' t '\in' T" := (has_type Gamma t T).

Hint Constructors has_type.

型付けの例


Module Examples2.
Import Examples.

練習問題: ★, standard (typing_example_0)

Definition trcd_kj :=
  (rcons k (abs z A (var z))
           (rcons j (abs z B (var z))
                      rnil)).

Example typing_example_0 :
  has_type empty
           (rcons k (abs z A (var z))
                     (rcons j (abs z B (var z))
                               rnil))
           TRcd_kj.
Proof.
Admitted.

練習問題: ★★, standard (typing_example_1)

Example typing_example_1 :
  has_type empty
           (app (abs x TRcd_j (rproj (var x) j))
                   (trcd_kj))
           (Arrow B B).
Proof with eauto.
Admitted.

練習問題: ★★, standard, optional (typing_example_2)

Example typing_example_2 :
  has_type empty
           (app (abs z (Arrow (Arrow C C) TRcd_j)
                           (rproj (app (var z)
                                            (abs x C (var x)))
                                    j))
                   (abs z (Arrow C C) trcd_kj))
           (Arrow B B).
Proof with eauto.
Admitted.

End Examples2.

型付けの性質


Well-Formedness


Lemma has_type__wf : forall Gamma t T,
  has_type Gamma t T -> well_formed_ty T.
Proof with eauto.
  intros Gamma t T Htyp.
  induction Htyp...
  -
    inversion IHHtyp1...
  -
    eapply wf_rcd_lookup...
  -
    apply subtype__wf in H.
    destruct H...
Qed.

Lemma step_preserves_record_tm : forall tr tr',
  record_tm tr ->
  tr --> tr' ->
  record_tm tr'.
Proof.
  intros tr tr' Hrt Hstp.
  inversion Hrt; subst; inversion Hstp; subst; eauto.
Qed.

フィールド参照


Lemma lookup_field_in_value : forall v T i Ti,
  value v ->
  has_type empty v T ->
  Tlookup i T = Some Ti ->
  exists vi, tlookup i v = Some vi /\ has_type empty vi Ti.
Proof with eauto.
  remember empty as Gamma.
  intros t T i Ti Hval Htyp. revert Ti HeqGamma Hval.
  induction Htyp; intros; subst; try solve_by_invert.
  -
    apply (rcd_types_match S) in H0...
    destruct H0 as [Si [HgetSi Hsub]].
    destruct (IHHtyp Si) as [vi [Hget Htyvi]]...
  -
    simpl in H0. simpl. simpl in H1.
    destruct (eqb_string i i0).
    +
      inversion H1. subst. exists t...
    +
      destruct (IHHtyp2 Ti) as [vi [get Htyvi]]...
      inversion Hval... Qed.

進行


練習問題: ★★★, standard (canonical_forms_of_arrow_types)

Lemma canonical_forms_of_arrow_types : forall Gamma s T1 T2,
     has_type Gamma s (Arrow T1 T2) ->
     value s ->
     exists x S1 s2,
        s = abs x S1 s2.
Proof with eauto.
Admitted.

Theorem progress : forall t T,
     has_type empty t T ->
     value t \/ exists t', t --> t'.
Proof with eauto.
  intros t T Ht.
  remember empty as Gamma.
  revert HeqGamma.
  induction Ht;
    intros HeqGamma; subst...
  -
    inversion H.
  -
    right.
    destruct IHHt1; subst...
    +
      destruct IHHt2; subst...
      *
        destruct (canonical_forms_of_arrow_types empty t1 T1 T2)
          as [x [S1 [t12 Heqt1]]]...
        subst. exists ([x:=t2]t12)...
      *
        destruct H0 as [t2' Hstp]. exists (app t1 t2')...
    +
      destruct H as [t1' Hstp]. exists (app t1' t2)...
  -
    right. destruct IHHt...
    +
      destruct (lookup_field_in_value t T i Ti)
        as [t' [Hget Ht']]...
    +
      destruct H0 as [t' Hstp]. exists (rproj t' i)...
  -
    destruct IHHt1...
    +
      destruct IHHt2...
      *
        right. destruct H2 as [tr' Hstp].
        exists (rcons i t tr')...
    +
      right. destruct H1 as [t' Hstp].
      exists (rcons i t' tr)... Qed.

定理: 任意の項tと型Tについて、empty |- t : Tならば、tは値であるか、ある項t'について t --> t' である。
証明 : tTempty |- t : T を満たすとする。 型付け導出についての帰納法を使う。
  • T_AbsT_RNilの場合は自明である。関数抽象と{}は常に値だからである。 T_Varの場合は考えなくて良い。なぜなら変数は空コンテキストで型付けできないからである。
  • 型付け導出の最後のステップがT_Appの適用だった場合、 項 t1 t2 と型 T1 T2 があって t = t1 t2T = T2empty |- t1 : T1 -> T2empty |- t2 : T1 となる。
    これらの型付け導出の帰納法の仮定より、t1は値であるかステップを進めることができ、 またt2は値であるかステップを進めることができる。
    • ある項t1'について t1 --> t1' とする。 するとST_App1より t1 t2 --> t1' t2 である。
    • そうでないならばt1は値である。
      • ある項t2'について t2 --> t2' とする。 するとt1が値であることから規則ST_App2により t1 t2 --> t1 t2' となる。
      • そうでなければt2は値である。補題canonical_forms_for_arrow_typesにより、 あるxS1s2について t1 = \x:S1.s2 である。 するとt2が値であることから、ST_AppAbsにより (\x:S1.s2) t2 --> [x:=t2]s2 となる。
  • 導出の最後のステップがT_Projの適用だった場合、 項tr、型Tr、ラベルiが存在して t = tr.iempty |- tr : TrTlookup i Tr = Some T となる。
    帰納仮定より、trは値であるかステップを進むことができる。 もしある項tr'について tr --> tr' ならば、規則ST_Proj1より tr.i --> tr'.i となる。
    もしtrが値であれば、補題lookup_field_in_value により項tiが存在して tlookup i tr = Some ti となる。 これから、規則ST_ProjRcdより tr.i ==> ti となる。
  • 導出の最後のステップがT_Subの適用だった場合、 型Sが存在して S <: T かつ empty |- t : S となる。 求める結果は帰納法の仮定そのものである。
  • 導出の最後のステップがT_RConsの適用だった場合、 項t1 tr、型 T1 Tr、ラベルiが存在して t = {i=t1, tr}T = {i:T1, Tr}record_tm trrecord_ty Trempty |- t1 : T1empty |- tr : Tr となる。
    これらの型付け導出についての帰納法の仮定より、t1は値であるか、ステップを進めることができ、 trは値であるかステップを進めることができることが言える。 それそれの場合を考える:
    • ある項t1'について t1 --> t1' とする。 すると規則ST_Rcd_Headから {i=t1, tr} --> {i=t1', tr} となる。
    • そうではないとき、t1は値である。
      • ある項tr'について tr --> tr' とする。 するとt1が値であることから、規則ST_Rcd_Tailより {i=t1, tr} --> {i=t1, tr'} となる。
      • そうではないとき、trも値である。すると、v_rconsから {i=t1, tr} は値である。

反転補題


Lemma typing_inversion_var : forall Gamma x T,
  has_type Gamma (var x) T ->
  exists S,
    Gamma x = Some S /\ subtype S T.
Proof with eauto.
  intros Gamma x T Hty.
  remember (var x) as t.
  induction Hty; intros;
    inversion Heqt; subst; try solve_by_invert.
  -
    exists T...
  -
    destruct IHHty as [U [Hctx HsubU]]... Qed.

Lemma typing_inversion_app : forall Gamma t1 t2 T2,
  has_type Gamma (app t1 t2) T2 ->
  exists T1,
    has_type Gamma t1 (Arrow T1 T2) /\
    has_type Gamma t2 T1.
Proof with eauto.
  intros Gamma t1 t2 T2 Hty.
  remember (app t1 t2) as t.
  induction Hty; intros;
    inversion Heqt; subst; try solve_by_invert.
  -
    exists T1...
  -
    destruct IHHty as [U1 [Hty1 Hty2]]...
    assert (Hwf := has_type__wf _ _ _ Hty2).
    exists U1... Qed.

Lemma typing_inversion_abs : forall Gamma x S1 t2 T,
     has_type Gamma (abs x S1 t2) T ->
     (exists S2, subtype (Arrow S1 S2) T
              /\ has_type (update Gamma x S1) t2 S2).
Proof with eauto.
  intros Gamma x S1 t2 T H.
  remember (abs x S1 t2) as t.
  induction H;
    inversion Heqt; subst; intros; try solve_by_invert.
  -
    assert (Hwf := has_type__wf _ _ _ H0).
    exists T12...
  -
    destruct IHhas_type as [S2 [Hsub Hty]]...
    Qed.

Lemma typing_inversion_proj : forall Gamma i t1 Ti,
  has_type Gamma (rproj t1 i) Ti ->
  exists T Si,
    Tlookup i T = Some Si /\ subtype Si Ti /\ has_type Gamma t1 T.
Proof with eauto.
  intros Gamma i t1 Ti H.
  remember (rproj t1 i) as t.
  induction H;
    inversion Heqt; subst; intros; try solve_by_invert.
  -
    assert (well_formed_ty Ti) as Hwf.
    {
      apply (wf_rcd_lookup i T Ti)...
      apply has_type__wf in H... }
    exists T, Ti...
  -
    destruct IHhas_type as [U [Ui [Hget [Hsub Hty]]]]...
    exists U, Ui... Qed.

Lemma typing_inversion_rcons : forall Gamma i ti tr T,
  has_type Gamma (rcons i ti tr) T ->
  exists Si Sr,
    subtype (RCons i Si Sr) T /\ has_type Gamma ti Si /\
    record_tm tr /\ has_type Gamma tr Sr.
Proof with eauto.
  intros Gamma i ti tr T Hty.
  remember (rcons i ti tr) as t.
  induction Hty;
    inversion Heqt; subst...
  -
    apply IHHty in H0.
    destruct H0 as [Ri [Rr [HsubRS [HtypRi HtypRr]]]].
    exists Ri, Rr...
  -
    assert (well_formed_ty (RCons i T Tr)) as Hwf.
    {
      apply has_type__wf in Hty1.
      apply has_type__wf in Hty2... }
    exists T, Tr... Qed.

Lemma abs_arrow : forall x S1 s2 T1 T2,
  has_type empty (abs x S1 s2) (Arrow T1 T2) ->
     subtype T1 S1
  /\ has_type (update empty x S1) s2 T2.
Proof with eauto.
  intros x S1 s2 T1 T2 Hty.
  apply typing_inversion_abs in Hty.
  destruct Hty as [S2 [Hsub Hty]].
  apply sub_inversion_arrow in Hsub.
  destruct Hsub as [U1 [U2 [Heq [Hsub1 Hsub2]]]].
  inversion Heq; subst... Qed.

コンテキスト不変性


Inductive appears_free_in : string -> tm -> Prop :=
  | afi_var : forall x,
      appears_free_in x (var x)
  | afi_app1 : forall x t1 t2,
      appears_free_in x t1 -> appears_free_in x (app t1 t2)
  | afi_app2 : forall x t1 t2,
      appears_free_in x t2 -> appears_free_in x (app t1 t2)
  | afi_abs : forall x y T11 t12,
        y <> x ->
        appears_free_in x t12 ->
        appears_free_in x (abs y T11 t12)
  | afi_proj : forall x t i,
      appears_free_in x t ->
      appears_free_in x (rproj t i)
  | afi_rhead : forall x i t tr,
      appears_free_in x t ->
      appears_free_in x (rcons i t tr)
  | afi_rtail : forall x i t tr,
      appears_free_in x tr ->
      appears_free_in x (rcons i t tr).

Hint Constructors appears_free_in.

Lemma context_invariance : forall Gamma Gamma' t S,
     has_type Gamma t S ->
     (forall x, appears_free_in x t -> Gamma x = Gamma' x) ->
     has_type Gamma' t S.
Proof with eauto.
  intros. generalize dependent Gamma'.
  induction H;
    intros Gamma' Heqv...
  -
    apply T_Var... rewrite <- Heqv...
  -
    apply T_Abs... apply IHhas_type. intros x0 Hafi.
    unfold update, t_update. destruct (eqb_stringP x x0)...
  -
    apply T_App with T1...
  -
    apply T_RCons... Qed.

Lemma free_in_context : forall x t T Gamma,
   appears_free_in x t ->
   has_type Gamma t T ->
   exists T', Gamma x = Some T'.
Proof with eauto.
  intros x t T Gamma Hafi Htyp.
  induction Htyp; subst; inversion Hafi; subst...
  -
    destruct (IHHtyp H5) as [T Hctx]. exists T.
    unfold update, t_update in Hctx.
    rewrite false_eqb_string in Hctx... Qed.

保存


Lemma substitution_preserves_typing : forall Gamma x U v t S,
     has_type (update Gamma x U) t S ->
     has_type empty v U ->
     has_type Gamma ([x:=v]t) S.
Proof with eauto.
  intros Gamma x U v t S Htypt Htypv.
  generalize dependent S. generalize dependent Gamma.
  induction t; intros; simpl.
  -
    rename s into y.
    destruct (typing_inversion_var _ _ _ Htypt) as [T [Hctx Hsub]].
    unfold update, t_update in Hctx.
    destruct (eqb_stringP x y)...
    +
      subst.
      inversion Hctx; subst. clear Hctx.
      apply context_invariance with empty...
      intros x Hcontra.
      destruct (free_in_context _ _ S empty Hcontra) as [T' HT']...
      inversion HT'.
    +
      destruct (subtype__wf _ _ Hsub)...
  -
    destruct (typing_inversion_app _ _ _ _ Htypt)
      as [T1 [Htypt1 Htypt2]].
    eapply T_App...
  -
    rename s into y. rename t into T1.
    destruct (typing_inversion_abs _ _ _ _ _ Htypt)
      as [T2 [Hsub Htypt2]].
    destruct (subtype__wf _ _ Hsub) as [Hwf1 Hwf2].
    inversion Hwf2. subst.
    apply T_Sub with (Arrow T1 T2)... apply T_Abs...
    destruct (eqb_stringP x y).
    +
      eapply context_invariance...
      subst.
      intros x Hafi. unfold update, t_update.
      destruct (eqb_string y x)...
    +
      apply IHt. eapply context_invariance...
      intros z Hafi. unfold update, t_update.
      destruct (eqb_stringP y z)...
      subst. rewrite false_eqb_string...
  -
    destruct (typing_inversion_proj _ _ _ _ Htypt)
      as [T [Ti [Hget [Hsub Htypt1]]]]...
  -
    eapply context_invariance...
    intros y Hcontra. inversion Hcontra.
  -
    destruct (typing_inversion_rcons _ _ _ _ _ Htypt) as
      [Ti [Tr [Hsub [HtypTi [Hrcdt2 HtypTr]]]]].
    apply T_Sub with (RCons s Ti Tr)...
    apply T_RCons...
    +
      apply subtype__wf in Hsub. destruct Hsub. inversion H0...
    +
      inversion Hrcdt2; subst; simpl... Qed.

Theorem preservation : forall t t' T,
     has_type empty t T ->
     t --> t' ->
     has_type empty t' T.
Proof with eauto.
  intros t t' T HT.
  remember empty as Gamma. generalize dependent HeqGamma.
  generalize dependent t'.
  induction HT;
    intros t' HeqGamma HE; subst; inversion HE; subst...
  -
    inversion HE; subst...
    +
      destruct (abs_arrow _ _ _ _ _ HT1) as [HA1 HA2].
      apply substitution_preserves_typing with T...
  -
    destruct (lookup_field_in_value _ _ _ _ H2 HT H)
      as [vi [Hget Hty]].
    rewrite H4 in Hget. inversion Hget. subst...
  -
    eauto using step_preserves_record_tm. Qed.

定理: tt'が項でTが型であり empty |- t : T かつ t --> t' ならば、 empty |- t' : T である。
証明: tTempty |- t : T であるとする。 t'を特化しないまま、型付け導出の構造についての帰納法で証明を進める。 T_AbsT_RNilの場合は考える必要はない。関数抽象と{}は進まないからである。 T_Varの場合も同様に考える必要はない。コンテキストが空だからである。
  • もし導出の最後ステップで使った規則がT_Appならば、 項t1 t2と型T1 T2があり、t = t1 t2T = T2empty |- t1 : T1 -> T2empty |- t2 : T1 である。
    ステップ関係の定義を見ることにより、 t1 t2 がステップを進める方法は3通りであることがわかる。 ST_App1ST_App2の場合は型付け導出の帰納法の仮定とT_Appを使うことで、 すぐに証明が完了する。
    そうではなくST_AppAbsにより t1 t2 がステップを進めるとする。 すると、ある型Sと項t12について t1 = \x:S.t12 かつ t' = [x:=t2]t12 となる。
    補題abs_arrowより、T1 <: S かつ x:S1 |- s2 : T2 であるから、 補題substitution_preserves_typingより empty |- [x:=t2] t12 : T2 となり、これが求める結果である。
  • もし導出の最後ステップで使った規則がT_Projならば、 項tr、型Tr、ラベルiが存在して t = tr.i かつ empty |- tr : Tr かつ Tlookup i Tr = Some T となる。
    型付け導出の帰納仮定より、任意の項tr'について、tr --> tr' ならば empty |- tr' Tr である。 ステップ関係の定義より、射影がステップを進める方法は2つである。 ST_Proj1の場合は帰納仮定からすぐに証明される。
    そうではなくtr.iのステップがST_ProjRcdによる場合、 trは値であり、ある項viがあって tlookup i tr = Some vi かつ t' = vi となる。しかし補題lookup_field_in_valueより empty |- vi : Ti となるが、これが求める結果である。
  • もし導出の最後ステップで使った規則がT_Subならば、型Sが存在して S <: T かつ empty |- t : S となる。 型付け導出の帰納法の仮定とT_Subの適用により結果がすぐに得られる。
  • もし導出の最後ステップで使った規則がT_RConsならば、 項 t1 tr、型 T1 Tr、ラベルiが存在して t = {i=t1, tr}T = {i:T1, Tr}record_tm trrecord_ty Trempty |- t1 : T1empty |- tr : Tr となる。
    ステップ関係の定義より、tST_Rcd_HeadまたはST_Rcd_Tail によってステップを進めたはずである。ST_Rcd_Headの場合、 t1の型付け導出についての帰納仮定とT_RConsより求める結果が得られる。 ST_Rcd_Tailの場合、trの型付け導出についての帰納仮定、T_RCons、 およびstep_preserves_record_tm補題の使用から求める結果が得られる。