LogrelCC.coop_logrel.logical_correctness

From LogrelCC.program_logic Require Import cl_weakestpre.
From LogrelCC.cooperative Require Import rules_binary.
From LogrelCC.F_mu_ref_cc Require Import rules_unary queue.
From LogrelCC.coop_logrel Require Import logrel translation light_weight_threads.
From iris.proofmode Require Import tactics.

Import coop_notations.
Import coop_type_notations.

Class coop_lib_info :=
  {
    queue_loc : loc;
  }.

Section bin_log_def.
  Context `{heapG Σ, cfgSG Σ, logrel_na_invs Σ, coop_lib_info}.
  Notation D := (prodC valC coopvalC -n> iProp Σ).

  Definition coop_inv : iProp Σ :=
    na_inv logrel_nais (nroot .@ "coop_inv")
           ( ths, queue_loc ↦ᵢ to_list ths
           [∗ list] th ths,
               K j' e',
                th = ContV K j' e'
            (obs_refine (fill K Unit, e')))%I.

  Definition bin_log_related (Γ : list type) (e : expr) (e' : coopexpr) (τ : type) :=
     Δ vvs ρ,
    env_Persistent Δ
    spec_ctx ρ
    coop_inv
     Γ ⟧* Δ vvs
              τ ⟧ₑ Δ (e.[env_subst
                            (vvs.*1 ++
                                [sim_fork queue_loc; sim_yield queue_loc])],
                       e'.[coop_env_subst (vvs.*2)]).
End bin_log_def.

Notation "Γ ⊨ e '≤log≤' e' : τ" :=
  (bin_log_related Γ e e' τ) (at level 74, e, e', τ at next level).

Section fundamental.
  Context `{heapG Σ, cfgSG Σ, logrel_na_invs Σ, coop_lib_info}.
  Notation D := (prodC valC coopvalC -n> iProp Σ).
  Implicit Types Δ : listC D.
  Hint Resolve to_of_val.

  Lemma bin_log_related_alt {Γ e e' τ} : Γ e log e' : τ
     Δ vvs ρ j KK,
    env_Persistent Δ
    spec_ctx ρ Γ ⟧* Δ vvs interp_ectx (interp τ) KK Δ
    coop_inv na_own logrel_nais CurTh j
    j coop_fill (KK.2) (e'.[coop_env_subst (vvs.*2)])
     WP fill (KK.1) e.[env_subst
                            (vvs.*1 ++
                                [sim_fork queue_loc; sim_yield queue_loc])]
                         {{ v, v' j',
                              na_own logrel_nais
                                     CurTh j' j' (coop_of_val v') }}.
  Proof.
    iIntros (Hlog Δ vvs ρ j K ?) "(#Hs & HΓ & HKK & #HCI & Hna & HCT & Hj)".
    iApply (Hlog with "[HΓ]"); iFrame; eauto.
  Qed.

  Notation "`' H" := (bin_log_related_alt H) (at level 8).

  Local Tactic Notation "smart_bind"
        constr(HP) uconstr(ctx) ident(vv) ident(j) constr(Hv) :=
    iApply (`'HP _ _ _ _ ctx); iFrame; iFrame "#"; iAlways; iIntros (vv j) Hv.

  Local Ltac DoLogrelIntros :=
    iIntros (Δ vvs ρ HenvPer) "#(Hspc & HCI & HΓ) /=";
    iIntros ([K1 K2] j) "(Hna & HCT & Hj & #HKK) /=";
    iDestruct (interp_env_length with "HΓ") as %?.

  Lemma bin_log_related_var Γ x τ :
    Γ !! x = Some τ
    Γ Var x log coop_Var x : τ.
  Proof.
    intros Htp. DoLogrelIntros.
    assert (x < length Γ) by eauto using lookup_lt_is_Some_1.
    rewrite /env_subst lookup_app_l; eauto;
      last by rewrite fmap_length; eauto with omega.
    iDestruct (interp_env_Some_l with "HΓ") as ([v v']) "[% ?]"; first done.
    rewrite /coop_env_subst !list_lookup_fmap; simplify_option_eq.
    by iApply ("HKK" $! (v,v') j with "[-]"); iFrame.
  Qed.

  Lemma bin_log_related_unit Γ : Γ Unit log coop_Unit : TUnit.
  Proof.
    DoLogrelIntros.
    by iApply ("HKK" $! (UnitV, coop_UnitV)); iFrame.
  Qed.

  Lemma bin_log_related_nat Γ n : Γ #n n log coop_Nat n : TNat.
  Proof.
    DoLogrelIntros.
    iApply ("HKK" $! (#nv _, coop_NatV _)); iFrame; eauto.
  Qed.

  Lemma bin_log_related_bool Γ b : Γ #♭ b log coop_Bool b : TBool.
  Proof.
    DoLogrelIntros.
    iApply ("HKK" $! (#♭v _, coop_BoolV _)); iFrame; eauto.
  Qed.

  Lemma bin_log_related_pair Γ e1 e2 τ1 τ2
      (IHHtyped1 : Γ (Translate_coop e1 (length Γ)) log e1 : τ1)
      (IHHtyped2 : Γ (Translate_coop e2 (length Γ)) log e2 : τ2) :
    Γ (Translate_coop (coop_Pair e1 e2) (length Γ)) log coop_Pair e1 e2 :
      TProd τ1 τ2.
  Proof.
    DoLogrelIntros.
    smart_bind IHHtyped1 (PairLCtx _ :: K1, coop_PairLCtx _ :: K2)
               vv j1 "(Hna & HCT & Hj & #Hvv)"; simpl.
    smart_bind IHHtyped2 (PairRCtx _ :: K1, coop_PairRCtx _ :: K2)
               ww j2 "(Hna & HCT & Hj & #Hww)"; simpl.
    iApply ("HKK" $! (PairV _ _, coop_PairV _ _)); iFrame; eauto.
  Qed.

  Lemma bin_log_related_fst Γ e τ1 τ2
      (IHHtyped : Γ (Translate_coop e (length Γ)) log e : TProd τ1 τ2) :
    Γ Translate_coop (coop_Fst e) (length Γ) log coop_Fst e : τ1.
  Proof.
    DoLogrelIntros.
    smart_bind IHHtyped (FstCtx :: K1, coop_FstCtx :: K2)
               vv j1 "(Hna & HCT & Hj & #Hvv)"; simpl.
    iDestruct "Hvv" as ([w1 w1'] [w2 w2']) "#[% [Hw1 Hw2]]"; simplify_eq.
    simpl.
    iApply wp_fst; eauto. iNext.
    iMod (step_fst with "[$HCT $Hj]") as "[HCT Hj]"; eauto.
    iApply ("HKK" $! (_, _)); iFrame; eauto.
  Qed.

  Lemma bin_log_related_snd Γ e τ1 τ2
      (IHHtyped : Γ Translate_coop e (length Γ) log e : TProd τ1 τ2) :
    Γ Translate_coop (coop_Snd e) (length Γ) log coop_Snd e : τ2.
  Proof.
    DoLogrelIntros.
    smart_bind IHHtyped (SndCtx :: K1, coop_SndCtx :: K2)
               vv j1 "(Hna & HCT & Hj & #Hvv)"; simpl.
    iDestruct "Hvv" as ([w1 w1'] [w2 w2']) "#[% [Hw1 Hw2]]"; simplify_eq.
    simpl.
    iApply wp_snd; eauto. iNext.
    iMod (step_snd with "[$HCT $Hj]") as "[HCT Hj]"; eauto.
    iApply ("HKK" $! (_, _)); iFrame; eauto.
  Qed.

  Lemma bin_log_related_injl Γ e τ1 τ2
      (IHHtyped : Γ Translate_coop e (length Γ) log e : τ1) :
    Γ Translate_coop (coop_InjL e) (length Γ) log coop_InjL e : (TSum τ1 τ2).
  Proof.
    DoLogrelIntros.
    smart_bind IHHtyped (InjLCtx :: K1, coop_InjLCtx :: K2)
               vv j1 "(Hna & HCT & Hj & #Hvv)"; simpl.
    iApply ("HKK" $! (InjLV _, coop_InjLV _)); iFrame; eauto.
  Qed.

  Lemma bin_log_related_injr Γ e τ1 τ2
      (IHHtyped : Γ Translate_coop e (length Γ) log e : τ2) :
    Γ Translate_coop (coop_InjR e) (length Γ) log coop_InjR e : TSum τ1 τ2.
  Proof.
    DoLogrelIntros.
    smart_bind IHHtyped (InjRCtx :: K1, coop_InjRCtx :: K2)
               vv j1 "(Hna & HCT & Hj & #Hvv)"; simpl.
    iApply ("HKK" $! (InjRV _, coop_InjRV _)); iFrame; eauto.
  Qed.

  Lemma bin_log_related_case Γ e0 e1 e2 τ1 τ2 τ3
      (Htyped2 : coop_typed (τ1 :: Γ) e1 τ3)
      (Htyped3 : coop_typed (τ2 :: Γ) e2 τ3)
      (IHHtyped1 : Γ Translate_coop e0 (length Γ) log e0 : TSum τ1 τ2)
      (IHHtyped2 : τ1 :: Γ Translate_coop e1 (S (length Γ)) log e1 : τ3)
      (IHHtyped3 : τ2 :: Γ Translate_coop e2 (S (length Γ)) log e2 : τ3) :
    Γ Translate_coop (coop_Case e0 e1 e2) (length Γ) log coop_Case e0 e1 e2
    : τ3.
  Proof.
    DoLogrelIntros.
    smart_bind IHHtyped1 (CaseCtx _ _ :: K1, coop_CaseCtx _ _ :: K2)
               vv j1 "(Hna & HCT & Hj & #Hvv)"; simpl.
    iDestruct "Hvv" as "[Hvv|Hvv]";
      iDestruct "Hvv" as ([w w']) "[% Hw]"; simplify_eq; simpl.
    + iMod (step_case_inl with "[$HCT $Hj]") as "[HCT Hj]"; eauto.
      iApply wp_case_injl; eauto using to_of_val. iNext.
      asimpl.
      erewrite !n_closed_subst_head_simpl;
          [|eapply (typed_Translate_n_closed (_ :: _))|]; eauto;
            last rewrite !app_length fmap_length /=; eauto.
      erewrite !coop_n_closed_subst_head_simpl;
        eauto using coop_typed_n_closed;
        last by rewrite /= ?fmap_length; eauto.
      iApply (`'IHHtyped2 _ ((w,w') :: vvs) _ _ (_, _) with "[-]"); iFrame.
      repeat iSplit; eauto.
      iApply interp_env_cons; auto.
    + iMod (step_case_inr with "[$HCT $Hj]") as "[HCT Hj]"; eauto.
      iApply wp_case_injr; eauto using to_of_val. iNext.
      asimpl.
      erewrite !n_closed_subst_head_simpl;
        [|eapply (typed_Translate_n_closed (_ :: _))|]; eauto;
          last rewrite !app_length fmap_length /=; eauto.
      erewrite !coop_n_closed_subst_head_simpl;
          eauto using coop_typed_n_closed;
          last by rewrite /= ?fmap_length; eauto.
      iApply (`'IHHtyped3 _ ((w,w') :: vvs) _ _ (_, _) with "[-]"); iFrame.
      repeat iSplit; eauto.
      iApply interp_env_cons; auto.
  Qed.

  Lemma bin_log_related_if Γ e0 e1 e2 τ
      (IHHtyped1 : Γ Translate_coop e0 (length Γ) log e0 : TBool)
      (IHHtyped2 : Γ Translate_coop e1 (length Γ) log e1 : τ)
      (IHHtyped3 : Γ Translate_coop e2 (length Γ) log e2 : τ) :
    Γ Translate_coop (coop_If e0 e1 e2) (length Γ) log coop_If e0 e1 e2 : τ.
  Proof.
    DoLogrelIntros.
    smart_bind IHHtyped1 (IfCtx _ _ :: K1, coop_IfCtx _ _ :: K2)
               vv j1 "(Hna & HCT & Hj & #Hvv)"; simpl.
    destruct vv as [v1 v2].
    iDestruct "Hvv" as ([]) "[% %]"; simplify_eq/=.
    + iMod (step_if_true with "[$HCT $Hj]") as "[HCT Hj]"; eauto.
      iApply wp_if_true. iNext.
      iApply (`'IHHtyped2 _ _ _ _ (_, _)); iFrame; eauto.
    + iMod (step_if_false with "[$HCT $Hj]") as "[HCT Hj]"; eauto.
      iApply wp_if_false. iNext.
      iApply (`'IHHtyped3 _ _ _ _ (_, _)); iFrame; eauto.
  Qed.

  Lemma bin_log_related_nat_binop Γ op e1 e2
      (IHHtyped1 : Γ Translate_coop e1 (length Γ) log e1 : TNat)
      (IHHtyped2 : Γ Translate_coop e2 (length Γ) log e2 : TNat) :
    Γ Translate_coop (coop_BinOp op e1 e2) (length Γ) log
      coop_BinOp op e1 e2 : binop_res_type op.
  Proof.
    DoLogrelIntros.
    smart_bind IHHtyped1 (BinOpLCtx _ _ :: K1, coop_BinOpLCtx _ _ :: K2)
               vv j1 "(Hna & HCT & Hj & #Hvv)"; simpl.
    smart_bind IHHtyped2 (BinOpRCtx _ _ :: K1, coop_BinOpRCtx _ _ :: K2)
               ww j2 "(Hna & HCT & Hj & #Hww)"; simpl.
    destruct vv as [v1 v2]; destruct ww as [w1 w2].
    iDestruct "Hvv" as (n) "[% %]"; simplify_eq/=.
    iDestruct "Hww" as (n') "[% %]"; simplify_eq/=.
    destruct op;
    iMod (step_nat_binop _ _ _ _ _ _ _ (coop_NatV _) (coop_NatV _)
            with "[$HCT $Hj]") as "[HCT Hj]";
      eauto; simpl; eauto;
        iApply wp_bin_op; eauto; iNext;
          try (iApply ("HKK" $! (#nv _, coop_NatV _) with "[-]"); iFrame; eauto);
          try (iApply ("HKK" $! (#♭v _, coop_BoolV _) with "[-]"); iFrame; eauto).
  Qed.

  Lemma bin_log_related_rec Γ e τ1 τ2
      (Htyped : coop_typed (TArrow τ1 τ2 :: τ1 :: Γ) e τ2)
      (IHHtyped : TArrow τ1 τ2 :: τ1 :: Γ Translate_coop e (S (S (length Γ)))
                         log e : τ2) :
    Γ Translate_coop (coop_Rec e) (length Γ) log coop_Rec e : TArrow τ1 τ2.
  Proof.
    DoLogrelIntros.
    iApply ("HKK" $! (RecV _, coop_RecV _)); iFrame; iAlways.
    iClear (K1 K2 j) "HKK".
    iLöb as "IH".
    iIntros ([v v']) "#Hvv". iIntros ([K1 K2] j) "(Hna & HCT & Hj & HKK) /=".
    iDestruct (interp_env_length with "HΓ") as %?.
    iApply wp_rec; eauto using to_of_val. iNext.
    iMod (step_rec with "[$HCT $Hj]") as "[HCT Hj]"; eauto.
    asimpl. change (Rec ?e) with (of_val (RecV e)).
    change (coop_Rec ?e) with (coop_of_val (coop_RecV e)).
    erewrite !n_closed_subst_head_simpl_2;
        [|eapply (typed_Translate_n_closed (_ :: _ :: _))|]; eauto;
          last rewrite !app_length fmap_length /=; eauto.
      erewrite !coop_n_closed_subst_head_simpl_2;
          eauto using coop_typed_n_closed;
          last by rewrite /= ?fmap_length; eauto.
    iApply (`'IHHtyped _ ((_,_) :: (v,v') :: vvs) _ _ (_, _) with "[-]");
      repeat iSplit; iFrame; eauto.
    rewrite !interp_env_cons; iSplit; auto.
  Qed.

  Lemma bin_log_related_Lam Γ e τ1 τ2
      (Htyped : coop_typed (τ1 :: Γ) e τ2)
      (IHHtyped : τ1 :: Γ Translate_coop e (S (length Γ)) log e : τ2) :
    Γ Translate_coop (coop_Lam e) (length Γ) log coop_Lam e : TArrow τ1 τ2.
  Proof.
    DoLogrelIntros.
    iApply ("HKK" $! (LamV _, coop_LamV _)); iFrame; iAlways.
    iClear (K1 K2 j) "HKK".
    iLöb as "IH".
    iIntros ([v v']) "#Hvv". iIntros ([K1 K2] j) "(Hna & HCT & Hj & HKK) /=".
    iDestruct (interp_env_length with "HΓ") as %?.
    iApply wp_Lam; eauto. iNext.
    iMod (step_Lam with "[$HCT $Hj]") as "[HCT Hj]"; eauto.
    asimpl.
    erewrite !n_closed_subst_head_simpl;
        [|eapply (typed_Translate_n_closed (_ :: _))|]; eauto;
          last rewrite !app_length fmap_length /=; eauto.
      erewrite !coop_n_closed_subst_head_simpl;
          eauto using coop_typed_n_closed;
          last by rewrite /= ?fmap_length; eauto.
    iApply (`'IHHtyped _ ((v,v') :: vvs) _ _ (_, _) with "[-]"); repeat iSplit;
      iFrame; eauto.
    rewrite !interp_env_cons; iSplit; auto.
  Qed.

  Lemma bin_log_related_LetIn Γ e1 e2 τ1 τ2
      (Hclosed2 : coop_typed (τ1 :: Γ) e2 τ2)
      (IHHtyped1 : Γ Translate_coop e1 (length Γ) log e1 : τ1)
      (IHHtyped2 : τ1 :: Γ Translate_coop e2 (S (length Γ)) log e2 : τ2) :
    Γ Translate_coop (coop_LetIn e1 e2) (length Γ) log
      coop_LetIn e1 e2 : τ2.
  Proof.
    DoLogrelIntros.
    smart_bind IHHtyped1 (LetInCtx _ :: K1, coop_LetInCtx _ :: K2)
               vv j1 "(Hna & HCT & Hj & #Hvv)"; simpl.
    iApply wp_LetIn; eauto. iNext.
    iMod (step_LetIn with "[$HCT $Hj]") as "[HCT Hj]"; eauto.
    asimpl.
    erewrite !n_closed_subst_head_simpl;
        [|eapply (typed_Translate_n_closed (_ :: _))|]; eauto;
          last rewrite !app_length fmap_length /=; eauto.
      erewrite !coop_n_closed_subst_head_simpl;
          eauto using coop_typed_n_closed;
          last by rewrite /= ?fmap_length; eauto.
    iApply (`'IHHtyped2 _ (vv :: vvs) _ _ (_, _)); repeat iSplit; iFrame; eauto.
    rewrite !interp_env_cons; iSplit; auto.
  Qed.

  Lemma bin_log_related_Seq Γ e1 e2 τ1 τ2
      (IHHtyped1 : Γ Translate_coop e1 (length Γ) log e1 : τ1)
      (IHHtyped2 : Γ Translate_coop e2 (length Γ) log e2 : τ2) :
    Γ Translate_coop (coop_Seq e1 e2) (length Γ) log coop_Seq e1 e2 : τ2.
  Proof.
    DoLogrelIntros.
    smart_bind IHHtyped1 (SeqCtx _ :: K1, coop_SeqCtx _ :: K2)
               vv j1 "(Hna & HCT & Hj & #Hvv)"; simpl.
    iApply wp_Seq; eauto. iNext.
    iMod (step_Seq with "[$HCT $Hj]") as "[HCT Hj]"; eauto.
    asimpl.
    iApply (`'IHHtyped2 _ vvs _ _ (_, _)); iFrame; eauto.
  Qed.

  Lemma bin_log_related_app Γ e1 e2 τ1 τ2
      (IHHtyped1 : Γ Translate_coop e1 (length Γ) log e1 : TArrow τ1 τ2)
      (IHHtyped2 : Γ Translate_coop e2 (length Γ) log e2 : τ1) :
    Γ Translate_coop (coop_App e1 e2) (length Γ) log coop_App e1 e2 : τ2.
  Proof.
    DoLogrelIntros.
    smart_bind IHHtyped1 (AppLCtx _ :: K1, coop_AppLCtx _ :: K2)
               vv j1 "(Hna & HCT & Hj & #Hvv)"; rewrite -/interp /=.
    smart_bind IHHtyped2 (AppRCtx _ :: K1, coop_AppRCtx _ :: K2)
               ww j2 "(Hna & HCT & Hj & #Hww)"; simpl.
    iSpecialize ("Hvv" with "[]"); auto.
    iApply ("Hvv" $! (_, _)); iFrame; eauto.
  Qed.

  Lemma bin_log_related_tlam Γ e τ
      (IHHtyped : (subst (ren (+1)) <$> Γ) Translate_coop e (length Γ) log
                                           e : τ) :
    Γ Translate_coop (coop_TLam e) (length Γ) log coop_TLam e : TForall τ.
  Proof.
    DoLogrelIntros.
    iApply ("HKK" $! (TLamV _, coop_TLamV _)); iFrame; iAlways.
    iClear (K1 K2 j) "HKK".
    iIntrosi ? KK' j') "(Hna & HCT & Hj & HKK) /=".
    iApply wp_tapp; iNext.
    iMod (step_tlam with "[$HCT $Hj]") as "[HCT Hj]"; eauto.
    iApply (`'IHHtypedi :: Δ)); iFrame "#"; iFrame.
    by rewrite interp_env_ren.
  Qed.

  Lemma bin_log_related_tapp Γ e τ τ'
      (IHHtyped : Γ Translate_coop e (length Γ) log e : TForall τ) :
    Γ Translate_coop (coop_TApp e) (length Γ) log coop_TApp e : τ.[τ'/].
  Proof.
    DoLogrelIntros.
    smart_bind IHHtyped (TAppCtx :: K1, coop_TAppCtx :: K2)
               vv j1 "(Hna & HCT & Hj & #Hvv)"; rewrite -/interp /=.
    iSpecialize ("Hvv" $! (interp τ' Δ) with "[#]"); [iPureIntro; apply _|].
    iApply ("Hvv" $! (_, _)); iFrame; eauto.
    iAlways. iIntros (??) "?". iApply "HKK".
    erewrite <- (interp_subst _ _ _ _); eauto.
  Qed.

  Lemma bin_log_related_fold Γ e τ
      (IHHtyped : Γ Translate_coop e (length Γ) log e : τ.[(TRec τ)/]) :
    Γ Translate_coop (coop_Fold e) (length Γ) log coop_Fold e : TRec τ.
  Proof.
    DoLogrelIntros.
    smart_bind IHHtyped (FoldCtx :: K1, coop_FoldCtx :: K2)
               vv j1 "(Hna & HCT & Hj & #Hvv)"; simpl.
    destruct vv as [v1 v2].
    iApply ("HKK" $! (FoldV _, coop_FoldV _)); iFrame; simpl.
    rewrite -> (fixpoint_unfold (interp_rec1 _ _) _).
    rewrite /= -(interp_subst _ _ _ _).
    iAlways; iExists (_, _); eauto.
  Qed.

  Lemma bin_log_related_unfold Γ e τ
      (IHHtyped : Γ Translate_coop e (length Γ) log e : TRec τ) :
    Γ Translate_coop (coop_Unfold e) (length Γ) log
      coop_Unfold e : τ.[(TRec τ)/].
  Proof.
    DoLogrelIntros.
    smart_bind IHHtyped (UnfoldCtx :: K1, coop_UnfoldCtx :: K2)
               vv j1 "(Hna & HCT & Hj & #Hvv)"; simpl.
    destruct vv as [v1 v2].
    rewrite /= (fixpoint_unfold (interp_rec1 _ _) _) /=.
    change (fixpoint _) with (interp (TRec τ) Δ).
    iDestruct "Hvv" as ([w w']) "#[% Hiz]"; simplify_eq/=.
    iMod (step_Fold with "[$HCT $Hj]") as "[HCT Hj]"; eauto.
    iApply wp_unfold; cbn; auto.
    iNext.
    iApply ("HKK" $! (w, w')); iFrame.
    by rewrite <- (interp_subst _ _ _ _).
  Qed.

  Lemma bin_log_related_fork Γ e
      (IHHtyped : Γ Translate_coop e (length Γ) log e : TUnit) :
    Γ Translate_coop (coop_CoFork e) (length Γ) log coop_CoFork e : TUnit.
  Proof.
    DoLogrelIntros.
    iMod (step_fork with "[$HCT $Hj]") as (j') "(HCT & Hj & Hj')"; eauto.
    rewrite {1}/env_subst lookup_app_r fmap_length; eauto with omega.
    replace (length Γ - length vvs) with 0 by lia; simpl.
    iApply wp_Lam; eauto.
    iNext. asimpl.
    iApply wp_callcc.
    iNext. asimpl.
    iMod (na_inv_open _ _ _ (nroot.@"coop_inv") with "HCI Hna") as
        "(Hinv & Hna & Hcl)"; auto.
    remember ( nroot.@"coop_inv") as mask.
    iDestruct "Hinv" as (ths) "(Hq & Hths)".
    iApply (wp_load (LetInCtx _ :: StoreRCtx (LocV _) :: SeqCtx _ :: K1)
              with "[-]"); iFrame "Hq".
    iNext; iIntros "Hq /=".
    iApply (wp_LetIn (StoreRCtx (LocV _) :: SeqCtx _ :: K1)); eauto.
    iNext; simpl. asimpl.
    iApply (clwp_cl (StoreRCtx (LocV _) :: SeqCtx _ :: K1));
      first by iApply (wp_EnQueue _ (ContV _)).
    iIntros (ths') "%"; subst ths'; simpl.
    iApply (wp_store (SeqCtx _ :: K1) with "[-]"); eauto; iFrame "Hq".
    iNext; simpl. iIntros "Hq".
    iMod ("Hcl" with "[Hths Hj $Hna Hq]") as "Hna".
    { iNext. iExists _; iFrame.
      rewrite big_sepL_singleton.
      iExists _, _, _; iFrame; iSplit; eauto.
      iAlways.
      iIntros (j'') "(Hna & HCT & Hj'') /=".
      iApply ("HKK" $! (UnitV, coop_UnitV) with "[-]"); iFrame; eauto. }
    iApply wp_Seq; eauto.
    iNext; simpl.
    iApply wp_throw; eauto.
    iNext; simpl.
    iApply (wp_Lam []); eauto.
    iNext; simpl. asimpl.
    iPoseProof (IHHtyped with "[]") as "He"; iFrame "#".
    iApply ("He" $! ([], []) j' with "[$HCT $Hj' $Hna]"); eauto.
    iIntros ((v, v') j''); iAlways.
    iIntros "(Hna & HCT & Hj'' & [% %])"; simpl in *; subst.
    iApply wp_value; eauto. iExists _, _; iFrame.
  Qed.

  Lemma bin_log_related_yield Γ :
    Γ Translate_coop coop_CoYield (length Γ) log coop_CoYield : TUnit.
  Proof.
    DoLogrelIntros.
    rewrite {1}/env_subst lookup_app_r fmap_length; eauto with omega.
    replace (S (length Γ) - length vvs) with 1 by lia; simpl.
    iApply wp_Lam; eauto.
    iNext. asimpl.
    iApply wp_callcc.
    iNext. asimpl.
    iMod (na_inv_open _ _ _ (nroot.@"coop_inv") with "HCI Hna") as
        "(Hinv & Hna & Hcl)"; auto.
    remember ( nroot.@"coop_inv") as mask.
    iDestruct "Hinv" as (ths) "(Hq & Hths)".
    iApply (wp_load (LetInCtx _ :: K1)
              with "[-]"); iFrame "Hq".
    iNext; iIntros "Hq /=".
    iApply wp_LetIn; eauto.
    iNext; simpl. asimpl.
    iApply (clwp_cl (LetInCtx _ :: K1));
      first by iApply (wp_EnQueue _ (ContV _)).
    iIntros (ths') "%"; subst ths'; simpl.
    iApply wp_LetIn; eauto.
    iNext; simpl. asimpl.
    iApply (clwp_cl (CaseCtx _ _ :: K1));
      first iApply wp_DeQueue.
    iIntros (v) "Hv".
    destruct ths as [|th ths]; simpl; iDestruct "Hv" as %?; subst v; simpl.
    - iApply wp_case_injl; eauto; simpl.
      iNext.
      iApply (wp_snd (StoreRCtx (LocV _) :: SeqCtx _ :: K1)); eauto.
      iNext; simpl.
      iApply (wp_store (SeqCtx _ :: K1)); eauto; iFrame "Hq".
      iNext. iIntros "Hq /=".
      iApply wp_Seq; eauto. iNext; simpl.
      iApply (wp_fst (ThrowRCtx UnitV :: K1)); eauto.
      iNext; simpl.
      iApply wp_throw; eauto; simpl.
      iNext.
      iMod ("Hcl" with "[Hths $Hna Hq]") as "Hna".
      { iNext. iExists []; iFrame.
        by rewrite big_sepL_nil. }
      iMod (step_yield_dont_change with "[$HCT $Hj]") as "[HCT Hj]"; eauto.
      iApply ("HKK" $! (UnitV, coop_UnitV)); simpl; iFrame; eauto.
    - iApply wp_case_injl; eauto; simpl.
      iNext.
      iApply (wp_snd (StoreRCtx (LocV _) :: SeqCtx _ :: K1)); eauto.
      iNext; simpl.
      iApply (wp_store (SeqCtx _ :: K1)); eauto; iFrame "Hq".
      iNext. iIntros "Hq /=".
      iApply wp_Seq; eauto. iNext; simpl.
      iApply (wp_fst (ThrowRCtx UnitV :: K1)); eauto.
      iNext; simpl.
      iDestruct "Hths" as "[Hth Hths]".
      iDestruct "Hth" as (K j' e') "(% & Hj' & #Hobs)"; subst th.
      iApply wp_throw; eauto; simpl.
      iNext.
      iMod (step_yield_change with "[$HCT $Hj $Hj']") as "(HCT & Hj & Hj')";
        eauto.
      iMod ("Hcl" with "[Hths Hj $Hna Hq]") as "Hna".
    { iNext. iExists _; iFrame.
      rewrite big_sepL_singleton.
      iExists _, _, _; iFrame; iSplit; eauto.
      iAlways.
      iIntros (j'') "(Hna & HCT & Hj'')".
      iApply ("HKK" $! (UnitV, coop_UnitV) with "[-]"); iFrame; eauto. }
      iApply "Hobs"; iFrame.
  Qed.

  Lemma bin_log_related_alloc Γ e τ
      (IHHtyped : Γ Translate_coop e (length Γ) log e : τ) :
    Γ Translate_coop (coop_Alloc e) (length Γ) log coop_Alloc e : Tref τ.
  Proof.
    DoLogrelIntros.
    smart_bind IHHtyped (AllocCtx :: K1, coop_AllocCtx :: K2)
               vv j1 "(Hna & HCT & Hj & #Hvv)"; simpl.
    iMod (step_alloc with "[$HCT $Hj]") as "[HCT Hlj]"; eauto.
    iDestruct "Hlj" as (l') "[Hj Hl']".
    iApply wp_alloc; eauto.
    iNext. iIntros (l) "Hl".
    iMod (inv_alloc (logN .@ (l,l')) _ ( w : val * coopval,
      l ↦ᵢ w.1 l' ↦ₛ w.2 interp τ Δ w)%I with "[Hl Hl']") as "HN"; eauto.
    { iNext. iExists vv; iFrame; iFrame "#". }
    iApply ("HKK" $! (LocV _, coop_LocV _)); iFrame.
    iExists (_, _); eauto.
  Qed.

  Lemma bin_log_related_load Γ e τ
      (IHHtyped : Γ Translate_coop e (length Γ) log e : (Tref τ)) :
    Γ Translate_coop (coop_Load e) (length Γ) log coop_Load e : τ.
  Proof.
    DoLogrelIntros.
    smart_bind IHHtyped (LoadCtx :: K1, coop_LoadCtx :: K2)
               vv j1 "(Hna & HCT & Hj & #Hvv)"; simpl.
    iDestruct "Hvv" as ([l l']) "[% Hll]"; simplify_eq/=.
    iApply wp_atomic_under_ectx; eauto.
    iInv (logN .@ (l,l')) as ([w w']) "[Hw1 [>Hw2 #Hw]]" "Hclose"; simpl.
    iModIntro.
    iApply (wp_load' with "[-]"); iFrame "Hw1".
    iNext. iIntros "Hw1".
    iMod (step_load with "[$HCT $Hj $Hw2]") as "(HCT & Hj & Hw2)";
      [solve_ndisj|by iFrame|].
    iApply wp_value; eauto.
    iMod ("Hclose" with "[Hw1 Hw2]") as "_".
    { iNext. iExists (w,w'); by iFrame. }
    iModIntro. iApply ("HKK" $! (w, w') with "[-]"); iFrame; eauto.
  Qed.

  Lemma bin_log_related_store Γ e1 e2 τ
      (IHHtyped1 : Γ Translate_coop e1 (length Γ) log e1 : (Tref τ))
      (IHHtyped2 : Γ Translate_coop e2 (length Γ) log e2 : τ) :
    Γ Translate_coop (coop_Store e1 e2) (length Γ) log
      coop_Store e1 e2 : TUnit.
  Proof.
    DoLogrelIntros.
    smart_bind IHHtyped1 (StoreLCtx _ :: K1, coop_StoreLCtx _:: K2)
               vv j1 "(Hna & HCT & Hj & #Hvv)"; simpl.
    smart_bind IHHtyped2 (StoreRCtx _ :: K1, coop_StoreRCtx _:: K2)
               ww j2 "(Hna & HCT & Hj & #Hww)"; simpl.
    iDestruct "Hvv" as ([l l']) "[% Hll]"; simplify_eq/=.
    iApply wp_atomic_under_ectx; eauto.
    iInv (logN .@ (l,l')) as ([v v']) "[Hv1 [>Hv2 #Hv]]" "Hclose".
    iModIntro.
    iApply (wp_store' with "[-]"); auto using to_of_val.
    iFrame "Hv1". iNext. iIntros "Hv1".
    iMod (step_store with "[$HCT $Hj $Hv2]") as "(HCT & Hj & Hv2)"; eauto;
    first solve_ndisj.
    iApply wp_value; eauto.
    iMod ("Hclose" with "[Hv1 Hv2]") as "_".
    { iNext; iExists _; simpl; iFrame; iFrame "#". }
    iModIntro. iApply ("HKK" $! (UnitV, coop_UnitV) with "[-]"); iFrame; eauto.
  Qed.

  Lemma bin_log_related_callcc Γ e τ
      (Hclosed : coop_typed (TCont τ :: Γ) e τ)
      (IHHtyped : TCont τ :: Γ Translate_coop e (S (length Γ)) log e : τ) :
    Γ Translate_coop (coop_Callcc e) (length Γ) log coop_Callcc e : τ.
  Proof.
    DoLogrelIntros.
    iMod (step_callcc with "[$HCT $Hj]") as "[HCT Hj]"; eauto.
    iApply wp_callcc; iNext.
    asimpl.
    change (coop_Cont K2) with (coop_of_val (coop_ContV K2)).
    change (Cont (K1)) with (of_val (ContV (K1))).
    erewrite !n_closed_subst_head_simpl;
        [|eapply (typed_Translate_n_closed (_ :: _))|]; eauto;
          last rewrite !app_length fmap_length /=; eauto.
    erewrite !coop_n_closed_subst_head_simpl;
      eauto using coop_typed_n_closed;
      last by rewrite /= ?fmap_length; eauto.
    iApply (`'IHHtyped _ ((ContV K1, coop_ContV K2) :: vvs) _ _ (_, _) with "[-]");
      iFrame; iFrame "#".
    iApply interp_env_cons; iFrame "#"; auto.
    simpl; iExists _, _; eauto.
  Qed.

  Lemma bin_log_related_throw Γ e1 e2 τ τ'
      (IHHtyped1 : Γ Translate_coop e1 (length Γ) log e1 : τ)
      (IHHtyped2 : Γ Translate_coop e2 (length Γ) log e2 : TCont τ) :
    Γ Translate_coop (coop_Throw e1 e2) (length Γ) log coop_Throw e1 e2 : τ'.
  Proof.
    DoLogrelIntros.
    smart_bind IHHtyped1 (ThrowLCtx _ :: K1, coop_ThrowLCtx _ :: K2)
               vv j1 "(Hna & HCT & Hj & #Hvv)"; simpl.
    smart_bind IHHtyped2 (ThrowRCtx _ :: K1, coop_ThrowRCtx _ :: K2)
               ww j2 "(Hna & HCT & Hj & #Hww)"; simpl.
    destruct ww as [w1 w2].
    iDestruct "Hww" as (K K') "[% #HKK2]"; simplify_eq; simpl.
    iMod (step_throw with "[$HCT $Hj]") as "[HCT Hj]"; eauto.
    iApply wp_throw; eauto.
    iNext.
    iApply "HKK2"; iFrame; eauto.
  Qed.

  Theorem binary_fundamental Γ e τ :
    coop_typed Γ e τ Γ Translate_coop e (length Γ) log e : τ.
  Proof.
    induction 1.
    - by apply bin_log_related_var.
    - by apply bin_log_related_unit.
    - by apply bin_log_related_nat.
    - by apply bin_log_related_bool.
    - apply bin_log_related_nat_binop; eauto.
    - apply bin_log_related_pair; eauto.
    - eapply bin_log_related_fst; eauto.
    - eapply bin_log_related_snd; eauto.
    - eapply bin_log_related_injl; eauto.
    - eapply bin_log_related_injr; eauto.
    - eapply bin_log_related_case; eauto.
    - eapply bin_log_related_if; eauto.
    - eapply bin_log_related_rec; eauto.
    - eapply bin_log_related_Lam; eauto.
    - eapply bin_log_related_LetIn; eauto.
    - eapply bin_log_related_Seq; eauto.
    - eapply (bin_log_related_app _ _ _ τ1); auto.
    - eapply bin_log_related_tlam; eauto.
      by rewrite fmap_length in IHtyped.
    - eapply bin_log_related_tapp; eauto.
    - eapply bin_log_related_fold; eauto.
    - eapply bin_log_related_unfold; eauto.
    - eapply bin_log_related_fork; eauto.
    - eapply bin_log_related_yield; eauto.
    - eapply bin_log_related_alloc; eauto.
    - eapply bin_log_related_load; eauto.
    - eapply (bin_log_related_store _ _ _ τ) ; eauto.
    - eapply bin_log_related_callcc; eauto.
    - eapply bin_log_related_throw; eauto.
  Qed.
End fundamental.