LogrelCC.cooperative.cooplang.typing

From stdpp Require Import prelude.
From LogrelCC.cooperative.cooplang Require Export lang.
From LogrelCC Require Export types.

Inductive typed (Γ : list type) : coopexpr type Prop :=
  | Var_typed x τ : Γ !! x = Some τ Γ ⊢ₜ Var x : τ
  | Unit_typed : Γ ⊢ₜ Unit : TUnit
  | Nat_typed n : Γ ⊢ₜ #n n : TNat
  | Bool_typed b : Γ ⊢ₜ #♭ b : TBool
  | BinOp_typed op e1 e2 :
     Γ ⊢ₜ e1 : TNat Γ ⊢ₜ e2 : TNat Γ ⊢ₜ BinOp op e1 e2 : binop_res_type op
  | Pair_typed e1 e2 τ1 τ2 : Γ ⊢ₜ e1 : τ1 Γ ⊢ₜ e2 : τ2
     Γ ⊢ₜ Pair e1 e2 : TProd τ1 τ2
  | Fst_typed e τ1 τ2 : Γ ⊢ₜ e : TProd τ1 τ2 Γ ⊢ₜ Fst e : τ1
  | Snd_typed e τ1 τ2 : Γ ⊢ₜ e : TProd τ1 τ2 Γ ⊢ₜ Snd e : τ2
  | InjL_typed e τ1 τ2 : Γ ⊢ₜ e : τ1 Γ ⊢ₜ InjL e : TSum τ1 τ2
  | InjR_typed e τ1 τ2 : Γ ⊢ₜ e : τ2 Γ ⊢ₜ InjR e : TSum τ1 τ2
  | Case_typed e0 e1 e2 τ1 τ2 τ3 :
     Γ ⊢ₜ e0 : TSum τ1 τ2 τ1 :: Γ ⊢ₜ e1 : τ3 τ2 :: Γ ⊢ₜ e2 : τ3
     Γ ⊢ₜ Case e0 e1 e2 : τ3
  | If_typed e0 e1 e2 τ :
     Γ ⊢ₜ e0 : TBool Γ ⊢ₜ e1 : τ Γ ⊢ₜ e2 : τ Γ ⊢ₜ If e0 e1 e2 : τ
  | Rec_typed e τ1 τ2 :
     TArrow τ1 τ2 :: τ1 :: Γ ⊢ₜ e : τ2 Γ ⊢ₜ Rec e : TArrow τ1 τ2
  | Lam_typed e τ τ' : τ :: Γ ⊢ₜ e : τ' Γ ⊢ₜ Lam e : TArrow τ τ'
  | LetIn_typed e e' τ τ' :
      Γ ⊢ₜ e : τ τ :: Γ ⊢ₜ e' : τ' Γ ⊢ₜ LetIn e e' : τ'
  | Seq_typed e e' τ τ' :
      Γ ⊢ₜ e : τ Γ ⊢ₜ e' : τ' Γ ⊢ₜ Seq e e' : τ'
  | App_typed e1 e2 τ1 τ2 :
     Γ ⊢ₜ e1 : TArrow τ1 τ2 Γ ⊢ₜ e2 : τ1 Γ ⊢ₜ App e1 e2 : τ2
  | TLam_typed e τ :
     subst (ren (+1)) <$> Γ ⊢ₜ e : τ Γ ⊢ₜ TLam e : TForall τ
  | TApp_typed e τ τ' : Γ ⊢ₜ e : TForall τ Γ ⊢ₜ TApp e : τ.[τ'/]
  | Fold_typed e τ : Γ ⊢ₜ e : τ.[TRec τ/] Γ ⊢ₜ Fold e : TRec τ
  | Unfold_typed e τ : Γ ⊢ₜ e : TRec τ Γ ⊢ₜ Unfold e : τ.[TRec τ/]
  | CoFork_typed e : Γ ⊢ₜ e : TUnit Γ ⊢ₜ CoFork e : TUnit
  | CoYield_typed : Γ ⊢ₜ CoYield : TUnit
  | Alloc_typed e τ : Γ ⊢ₜ e : τ Γ ⊢ₜ Alloc e : Tref τ
  | Load_typed e τ : Γ ⊢ₜ e : Tref τ Γ ⊢ₜ Load e : τ
  | Store_typed e e' τ : Γ ⊢ₜ e : Tref τ Γ ⊢ₜ e' : τ Γ ⊢ₜ Store e e' : TUnit
  | Callcc_typed e τ : TCont τ :: Γ ⊢ₜ e : τ Γ ⊢ₜ Callcc e : τ
  | Throw_typed e1 e2 τ τ' : Γ ⊢ₜ e1 : τ Γ ⊢ₜ e2 : TCont τ
     Γ ⊢ₜ Throw e1 e2 : τ'
where "Γ ⊢ₜ e : τ" := (typed Γ e τ).

Lemma n_closed_invariant n (e : coopexpr) s1 s2 :
  ( f, e.[upn n f] = e) ( x, x < n s1 x = s2 x) e.[s1] = e.[s2].
Proof.
  intros Hnc. specialize (Hnc (ren (+1))).
  revert n Hnc s1 s2.
  apply (coopexpr_rect'
            (λ e, n : nat, e.[upn n (ren (+1))] = e
                            s1 s2 : nat coopexpr,
                             ( x : nat, x < n s1 x = s2 x) e.[s1] = e.[s2])
            (λ v, n : nat, coopval_subst (upn n (ren (+1))) v = v
                            s1 s2 : nat coopexpr,
                             ( x : nat, x < n s1 x = s2 x)
                             coopval_subst s1 v = coopval_subst s2 v));
     intros; simpl in *; asimpl in *;
       try (f_equal;
            match goal with H : _ |- _ => eapply H end; eauto;
            try
              match goal with
                H' : _ = _ |- _ => inversion H'
              end; try match goal with H : _ |- _ => by rewrite H end;
            fail);
       try match goal with
  n : nat,
  H1 : Cont (_ :: ?A) = Cont (_ :: ?B),
  H2 : n : nat, Cont _ = Cont _ _,
  H3 : x : nat, x < _ _ x = _ x |- _ =>
  let Hf := fresh in
  assert (Cont A = Cont B) as Hf by
      (by f_equal; inversion H1;
          match goal with
            H' : A = _ |- _ => rewrite !H' end);
    specialize (H2 n Hf _ _ H3);
    let Hf := fresh in
    inversion H as [[Hf]]; rewrite Hf;
      inversion H1;
      repeat match goal with
             H : _ |- _ => erewrite H; eauto
             end; eauto
  end;
    try match goal with
  n : nat,
  H1 : ContV (_ :: ?A) = ContV (_ :: ?B),
  H2 : n : nat, ContV _ = ContV _ _,
  H3 : x : nat, x < _ _ x = _ x |- _ =>
  let Hf := fresh in
  assert (ContV A = ContV B) as Hf by
      (by f_equal; inversion H1;
          match goal with
            H' : A = _ |- _ => rewrite !H' end);
    specialize (H2 n Hf _ _ H3);
    let Hf := fresh in
    inversion H as [[Hf]]; rewrite Hf;
      inversion H1;
      repeat match goal with
             H : _ |- _ => erewrite H; eauto
             end; eauto
  end;
    try (repeat match goal with
             H : _ |- _ => erewrite H; eauto; clear H
           end; fail).
  - apply H0. rewrite iter_up in H. destruct lt_dec; try omega.
    asimpl in *. cbv in x. replace (n + (x - n)) with x in n0 by omega.
    inversion H; omega.
  - inversion H0 as [[H']]; rewrite H'. erewrite (H (S (S n))); eauto.
    intros [|[|x]] H2; [by cbv|by cbv |].
      asimpl; rewrite H1; auto with omega.
  - inversion H0 as [[H']]; rewrite H'. erewrite (H (S n)); eauto.
    intros [|x] H2; [by cbv|].
      asimpl; rewrite H1; auto with omega.
  - inversion H1 as [[H'1 H'2]]. rewrite H'1 H'2.
    erewrite H; eauto.
    erewrite (H0 (S n)); eauto.
    intros [|x] ?; [by cbv|].
    asimpl; rewrite H2; auto with omega.
  - inversion H2 as [[H'1 H'2 H'3]]; rewrite H'1 H'2 H'3.
    erewrite H; eauto.
    erewrite H0; [|eauto|].
    + erewrite H1; [|eauto|]; eauto.
     { intros [|x] ? ; [by cbv |].
       asimpl; rewrite H3; auto with omega. }
    + intros [|x] ? ; [by cbv |].
       asimpl; rewrite H3; auto with omega.
  - intros []; asimpl; auto.
    intros; rewrite H2; auto with omega.
  - inversion H2 as [[H'1 H'2 H'3]]; rewrite H'1 H'2 H'3.
    erewrite H0; [|eauto|].
    + erewrite H1; [|eauto|]; eauto.
      { intros [|x] ? ; [by cbv |].
       asimpl; rewrite H3; auto with omega. }
    + intros [|x] ? ; [by cbv |].
       asimpl; rewrite H3; auto with omega.
  - inversion H0 as [[H']]; rewrite H'. erewrite (H (S (S n))); eauto.
    intros [|[|x]] H2; [by cbv|by cbv |].
      asimpl; rewrite H1; auto with omega.
  - inversion H0 as [[H']]; rewrite H'. erewrite (H (S n)); eauto.
    intros [|x] H2; [by cbv |].
      asimpl; rewrite H1; auto with omega.
  - inversion H0 as [[H']]; clear H0; rewrite H'.
    match type of H' with
      ?A = ?B => assert (Cont A = Cont B) as H'' by by rewrite H'
    end.
    specialize (H n H'' _ _ H1).
    by inversion H as [[H''']]; rewrite H'''.
  - intros []; asimpl; auto.
    intros; rewrite H2; auto with omega.
  - inversion H2 as [[H'1 H'2 H'3]]; rewrite H'1 H'2 H'3.
    erewrite H0; [|eauto|].
    + erewrite H1; [|eauto|]; eauto.
      { intros [|x] ? ; [by cbv |].
       asimpl; rewrite H3; auto with omega. }
    + intros [|x] ? ; [by cbv |].
       asimpl; rewrite H3; auto with omega.
  - inversion H0 as [[H']]; rewrite H'.
    erewrite (H (S n)); eauto.
    intros [|x] ? ; [by cbv |].
       asimpl; rewrite H1; auto with omega.
Qed.

Definition env_subst (vs : list coopval) (x : var) : coopexpr :=
  from_option id (Var x) (of_val <$> vs !! x).

Lemma typed_n_closed Γ τ e : Γ ⊢ₜ e : τ ( f, e.[upn (length Γ) f] = e).
Proof.
  intros H. induction H => f; asimpl; simpl in *; auto with f_equal.
  - apply lookup_lt_Some in H. rewrite iter_up. destruct lt_dec; auto with omega.
  - f_equal. apply IHtyped.
  - by f_equal; rewrite map_length in IHtyped.
Qed.

Lemma n_closed_subst_head_simpl n e w ws :
  ( f, e.[upn n f] = e)
  S (length ws) = n
  e.[of_val w .: env_subst ws] = e.[env_subst (w :: ws)].
Proof.
  intros H1 H2.
  rewrite /env_subst. eapply n_closed_invariant; eauto=> /= -[|x] ? //=.
  destruct (lookup_lt_is_Some_2 ws x) as [v' Hv]; first omega; simpl.
  by rewrite Hv.
Qed.

Lemma typed_subst_head_simpl Δ τ e w ws :
  Δ ⊢ₜ e : τ length Δ = S (length ws)
  e.[of_val w .: env_subst ws] = e.[env_subst (w :: ws)].
Proof. eauto using n_closed_subst_head_simpl, typed_n_closed. Qed.

Lemma n_closed_subst_head_simpl_2 n e w w' ws :
  ( f, e.[upn n f] = e) (S (S (length ws))) = n
  e.[of_val w .: of_val w' .: env_subst ws] = e.[env_subst (w :: w' :: ws)].
Proof.
  intros H1 H2.
  rewrite /env_subst. eapply n_closed_invariant; eauto => /= -[|[|x]] H3 //=.
  destruct (lookup_lt_is_Some_2 ws x) as [v' Hv]; first omega; simpl.
  by rewrite Hv.
Qed.

Lemma typed_subst_head_simpl_2 Δ τ e w w' ws :
  Δ ⊢ₜ e : τ length Δ = 2 + length ws
  e.[of_val w .: of_val w' .: env_subst ws] = e.[env_subst (w :: w' :: ws)].
Proof. eauto using n_closed_subst_head_simpl_2, typed_n_closed. Qed.

Lemma empty_env_subst e : e.[env_subst []] = e.
Proof. change (env_subst []) with (@ids coopexpr _). by asimpl. Qed.

Weakening
Lemma context_gen_weakening ξ Γ' Γ e τ :
  Γ' ++ Γ ⊢ₜ e : τ
  Γ' ++ ξ ++ Γ ⊢ₜ e.[upn (length Γ') (ren (+ (length ξ)))] : τ.
Proof.
  intros H1.
  remember (Γ' ++ Γ) as Ξ. revert Γ' Γ ξ HeqΞ.
  induction H1 => Γ1 Γ2 ξ HeqΞ; subst; asimpl in *; eauto using typed.
  - rewrite iter_up; destruct lt_dec as [Hl | Hl].
    + constructor. rewrite lookup_app_l; trivial. by rewrite lookup_app_l in H.
    + asimpl. constructor. rewrite lookup_app_r; auto with omega.
      rewrite lookup_app_r; auto with omega.
      rewrite lookup_app_r in H; auto with omega.
      match goal with
        |- _ !! ?A = _ => by replace A with (x - length Γ1) by omega
      end.
  - econstructor; eauto. by apply (IHtyped2 (_::_)). by apply (IHtyped3 (_::_)).
  - constructor. by apply (IHtyped (_ :: _ :: _)).
  - constructor. by apply (IHtyped (_ :: _)).
  - econstructor; eauto. by apply (IHtyped2 (_ :: _)).
  - constructor.
    specialize (IHtyped
      (subst (ren (+1)) <$> Γ1) (subst (ren (+1)) <$> Γ2) (subst (ren (+1)) <$> ξ)).
    asimpl in *. rewrite ?map_length in IHtyped.
    repeat rewrite fmap_app. apply IHtyped.
    by repeat rewrite fmap_app.
  - constructor. by apply (IHtyped (_ :: _)).
Qed.

Lemma context_weakening ξ Γ e τ :
  Γ ⊢ₜ e : τ ξ ++ Γ ⊢ₜ e.[(ren (+ (length ξ)))] : τ.
Proof. eapply (context_gen_weakening _ []). Qed.

Lemma closed_context_weakening ξ Γ e τ :
  ( f, e.[f] = e) Γ ⊢ₜ e : τ ξ ++ Γ ⊢ₜ e : τ.
Proof. intros H1 H2. erewrite <- H1. by eapply context_weakening. Qed.

Module coop_type_notations.

  Notation coop_typed := typed.
  Notation coop_Var_typed := Var_typed.
  Notation coop_Unit_typed := Unit_typed.
  Notation coop_Nat_typed := Nat_typed.
  Notation coop_Bool_typed := Bool_typed.
  Notation coop_BinOp_typed := BinOp_typed.
  Notation coop_Pair_typed := Pair_typed.
  Notation coop_Fst_typed := Fst_typed.
  Notation coop_Snd_typed := Snd_typed.
  Notation coop_InjL_typed := InjL_typed.
  Notation coop_InjR_typed := InjR_typed.
  Notation coop_Case_typed := Case_typed.
  Notation coop_If_typed := If_typed.
  Notation coop_Rec_typed := Rec_typed.
  Notation coop_Lam_typed := Lam_typed.
  Notation coop_LetIn_typed := LetIn_typed.
  Notation coop_Seq_typed := Seq_typed.
  Notation coop_App_typed := App_typed.
  Notation coop_TLam_typed := TLam_typed.
  Notation coop_TApp_typed := TApp_typed.
  Notation coop_Fold_typed := Fold_typed.
  Notation coop_Unfold_typed := Unfold_typed.
  Notation coop_CoFork_typed := CoFork_typed.
  Notation coop_CoYield_typed := CoYield_typed.
  Notation coop_Alloc_typed := Alloc_typed.
  Notation coop_Load_typed := Load_typed.
  Notation coop_Store_typed := Store_typed.
  Notation coop_Callcc_typed := Callcc_typed.
  Notation coop_Throw_typed := Throw_typed.

  Notation coop_n_closed_invariant := n_closed_invariant.
  Notation coop_env_subst := env_subst.

  Notation coop_typed_n_closed := typed_n_closed.
  Notation coop_n_closed_subst_head_simpl := n_closed_subst_head_simpl.

  Notation coop_typed_subst_head_simpl := typed_subst_head_simpl.

  Notation coop_n_closed_subst_head_simpl_2 := n_closed_subst_head_simpl_2.

  Notation coop_typed_subst_head_simpl_2 := typed_subst_head_simpl_2.

  Notation coop_empty_env_subst := empty_env_subst.

  Notation coop_context_gen_weakening := context_gen_weakening.
  Notation coop_context_weakening := context_weakening.

  Notation coop_closed_context_weakening := closed_context_weakening.

End coop_type_notations.