RunST.fundamentalNN

From RunST Require Import typing IC IC_lifting rules regions
     logrel_shared logrelNN lang ST_Lang_reduction reduction.
From iris.proofmode Require Import tactics.
From iris.algebra Require Import gmap.

Set Bullet Behavior "Strict Subproofs".

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

  Definition bin_log_relatedNN (Γ : list type) (e e' : expr) (τ : type) :=
     Δ vvs, env_PersistentP Δ reg_inv
       Γ NN* Δ vvs τ NN Δ (e.[env_subst (vvs.*1)], e'.[env_subst (vvs.*2)]).

End bin_log_def.

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

Section fundamental.
  Context `{heapG Σ, LogRelG Σ}.
  Notation D := (prodC valC valC -n> iProp Σ).
  Notation Cx := (prodC ((prodC valC valC) -n> iProp Σ) (leibnizC positive)).
  Implicit Types e : expr.
  Implicit Types Δ : listC Cx.
  Hint Resolve to_of_val.

  Lemma logrel_bind K K' τ' {e e' τ} Δ Δ' :
     τ NN Δ' (e, e')
         ( vv, τ NN Δ' vv -★
           ( τ' NN Δ (fill K (of_val (vv.1)), fill K' (of_val (vv.2)))))
      τ' NN Δ (fill K e, fill K' e').
  Proof.
    iIntros "[Hee' Hcont]".
    iIntros (γh γh' σ1') "Hσ1'" => /=.
    iApply (ic_ectx_bind with "[-]").
    iApply (ic_wand_l with "[-]"); iSplitR "Hee' Hσ1'"; last by iApply "Hee'".
    simpl. iIntros (v n) "Hc". iDestruct "Hc" as (σ2' v') "(Hred & Ho & Hvv')".
    iDestruct "Hred" as %Hred.
    iSpecialize ("Hcont" $! (v, v') with "[Hvv']"); first done; simpl.
    iApply (ic_wand_l with "[-]"); iSplitL "";
      last by iApply "Hcont".
    simpl. iIntros (w m) "Hc". iDestruct "Hc" as (σ3' w') "(Hred & Ho & Hww')".
    iDestruct "Hred" as %Hred'.
    iExists σ3', w'; iFrame; iPureIntro.
    eapply nsteps_trans => //. eapply (@nsteps_ctx det_lang) => //; first by
                             eapply (@ectx_lang_ctx _ _ _ _ det_ectx_lang).
  Qed.

  Lemma logrel_pure_det_head_step e1 e1' e2 e2' τ Δ :
    ( σ, head_step e1 σ e2 σ [])
    ( σ e3 σ3, head_step e1 σ e3 σ3 [] e3 = e2 σ3 = σ)
    ( σ, @ectx_language.head_step _ _ _ _ det_ectx_lang e1' σ e2' σ [])
     τ NN Δ (e2, e2') τ NN Δ (e1, e1').
  Proof.
    iIntros (He1 He2 He3) "Hee". iIntros (γh γh' σ1') "Hσ1' /=".
    unfold interp_exprNN. rewrite ic_eq /ic_def.
    iIntros (σ1 σ2 v n) "[Hred Hs2]"; iDestruct "Hred" as %Hred.
    inversion Hred as [|k Hs2 [ew σw] Hs4 Hs5 Hs6]; subst.
    { specialize (He1 σ1').
      by apply val_stuck in He1; rewrite to_of_val in He1. }
    apply head_reducible_prim_step in Hs5; last by eexists _, _, _; eapply He1.
    simpl in *.
    apply He2 in Hs5; destruct Hs5 as [??]; subst.
    rewrite future_S. iModIntro. iNext.
    iSpecialize ("Hee" $! γh γh' σ1' with "Hσ1'").
    iSpecialize ("Hee" $! _ _ _ _ with "[Hs2]"); first by iFrame.
    iMod "Hee" as "[Hee Ho]". iDestruct "Hee" as (σ3 v3) "(Hred' & Ho' & Hrel)".
    iDestruct "Hred'" as %Hred'.
    iModIntro. iFrame. iExists _, _; iFrame.
    iPureIntro. econstructor; eauto; simpl. apply head_prim_step; simpl; eauto.
  Qed.

  Local Ltac value_case := iApply ic_value; [cbn; rewrite ?to_of_val; trivial|].

  Lemma bin_log_related_var Γ x τ :
    Γ !! x = Some τ Γ Var x logNN Var x : τ.
  Proof.
    iIntros (? Δ vvs ?) "#(Hinv & HΓ)"; iIntros (γh γh' σ1') "Hσ1' /=".
    iDestruct (interp_env_Some_l with "HΓ") as ([v v']) "[% ?]"; first done.
    rewrite /env_subst !list_lookup_fmap; simplify_option_eq. value_case.
    iExists _, _; iFrame; iFrame "#"; iPureIntro; econstructor.
  Qed.

  Lemma bin_log_related_unit Γ : Γ (Lit LitUnit) logNN (Lit LitUnit) : TUnit.
  Proof.
    iIntros (Δ vvs ?) "#(Hinv & HΓ)"; iIntros (γh γh' σ1') "Hσ1' /=".
    value_case. iExists _, (LitV LitUnit); eauto.
    iSplitR; last by eauto. iPureIntro. constructor 1.
  Qed.

  Lemma bin_log_related_int Γ z :
    Γ (Lit $ LitInt z) logNN (Lit $ LitInt z) : TInt.
  Proof.
    iIntros (Δ vvs ?) "#(Hinv & HΓ)"; iIntros (γh γh' σ1') "Hσ1' /=".
    value_case. iExists _, (LitV $ LitInt _); simpl; eauto.
    iSplitR; last by eauto. iPureIntro. constructor 1.
  Qed.

  Lemma bin_log_related_bool Γ b :
    Γ (Lit $ LitBool b) logNN (Lit $ LitBool b) : TBool.
  Proof.
    iIntros (Δ vvs ?) "#(Hinv & HΓ)"; iIntros (γh γh' σ1') "Hσ1' /=".
    value_case. iExists _, (LitV $ LitBool _); simpl; eauto.
    iSplitR; last by eauto. iPureIntro. constructor 1.
  Qed.

  Lemma bin_log_related_pair Γ e1 e2 e1' e2' τ1 τ2
      (IHHtyped1 : Γ e1 logNN e1' : τ1)
      (IHHtyped2 : Γ e2 logNN e2' : τ2) :
    Γ Pair e1 e2 logNN Pair e1' e2' : TProd τ1 τ2.
  Proof.
    iIntros (Δ vvs ?) "#(Hinv & HΓ)". simpl.
    iApply ((logrel_bind [PairLCtx _] [PairLCtx _] (TProd _ _)) with "[]").
    iSplitL ""; first by iApply IHHtyped1; iFrame "#".
    iIntros (vv) "#Hvv". simpl.
    iApply ((logrel_bind [PairRCtx _] [PairRCtx _] (TProd _ _)) with "[]").
    iSplitL ""; first by iApply IHHtyped2; iFrame "#".
    iIntros (ww) "#Hww". simpl.
    iIntros (? ? ?) "?"; value_case.
    iExists _, _; iFrame; iSplit; last first; simpl; eauto.
    iPureIntro. by constructor 1.
  Qed.

  Lemma bin_log_related_fst Γ e e' τ1 τ2
      (IHHtyped : Γ e logNN e' : TProd τ1 τ2) :
    Γ Fst e logNN Fst e' : τ1.
  Proof.
    iIntros (Δ vvs ?) "#(Hinv & HΓ)". simpl.
    iApply ((logrel_bind [FstCtx] [FstCtx]) with "[]").
    iSplitL ""; first by iApply IHHtyped; iFrame "#".
    iIntros (vv) "#Hvv". simpl.
    iDestruct "Hvv" as ([w1 w1'] [w2 w2']) "#[% [Hw1 Hw2]]"; simplify_eq; simpl.
    iIntros (? ? ?) "? /=".
    iApply ic_fst; auto.
    iNext; iModIntro. iExists _, _; iFrame; iFrame "#".
    iPureIntro; econstructor; last econstructor; simpl.
    apply head_prim_step; econstructor; eauto.
  Qed.

  Lemma bin_log_related_snd Γ e e' τ1 τ2
      (IHHtyped : Γ e logNN e' : TProd τ1 τ2) :
    Γ Snd e logNN Snd e' : τ2.
  Proof.
    iIntros (Δ vvs ?) "#(Hinv & HΓ)". simpl.
    iApply ((logrel_bind [SndCtx] [SndCtx]) with "[]").
    iSplitL ""; first by iApply IHHtyped; iFrame "#".
    iIntros (vv) "#Hvv". simpl.
    iDestruct "Hvv" as ([w1 w1'] [w2 w2']) "#[% [Hw1 Hw2]]"; simplify_eq; simpl.
    iIntros (? ? ?) "? /=".
    iApply ic_snd; auto.
    iNext; iModIntro. iExists _, _; iFrame; iFrame "#".
    iPureIntro; econstructor; last econstructor; simpl.
    apply head_prim_step; econstructor; eauto.
  Qed.

  Lemma bin_log_related_injl Γ e e' τ1 τ2
      (IHHtyped : Γ e logNN e' : τ1) :
    Γ InjL e logNN InjL e' : (TSum τ1 τ2).
  Proof.
    iIntros (Δ vvs ?) "#(Hinv & HΓ)". simpl.
    iApply ((logrel_bind [InjLCtx] [InjLCtx] (TSum _ _)) with "[]").
    iSplitL ""; first by iApply IHHtyped; iFrame "#".
    iIntros (vv) "#Hvv". simpl.
    iIntros (? ? ?) "? /=".
    value_case. iExists _, _; iFrame; iFrame "#".
    iSplit; eauto.
    iPureIntro; econstructor.
  Qed.

  Lemma bin_log_related_injr Γ e e' τ1 τ2
      (IHHtyped : Γ e logNN e' : τ2) :
    Γ InjR e logNN InjR e' : TSum τ1 τ2.
  Proof.
    iIntros (Δ vvs ?) "#(Hinv & HΓ)". simpl.
    iApply ((logrel_bind [InjRCtx] [InjRCtx] (TSum _ _)) with "[]").
    iSplitL ""; first by iApply IHHtyped; iFrame "#".
    iIntros (vv) "#Hvv". simpl.
    iIntros (? ? ?) "? /=".
    value_case. iExists _, _; iFrame; iFrame "#".
    iSplit; eauto.
    iPureIntro; econstructor.
  Qed.

  Lemma bin_log_related_case Γ e0 e1 e2 e0' e1' e2' τ1 τ2 τ3
      (Hclosed2 : f, e1.[upn (S (length Γ)) f] = e1)
      (Hclosed3 : f, e2.[upn (S (length Γ)) f] = e2)
      (Hclosed2' : f, e1'.[upn (S (length Γ)) f] = e1')
      (Hclosed3' : f, e2'.[upn (S (length Γ)) f] = e2')
      (IHHtyped1 : Γ e0 logNN e0' : TSum τ1 τ2)
      (IHHtyped2 : τ1 :: Γ e1 logNN e1' : τ3)
      (IHHtyped3 : τ2 :: Γ e2 logNN e2' : τ3) :
    Γ Case e0 e1 e2 logNN Case e0' e1' e2' : τ3.
  Proof.
    iIntros (Δ vvs ?) "#(Hinv & HΓ)". simpl.
    iDestruct (interp_env_length with "HΓ") as %?.
    iApply ((logrel_bind [CaseCtx _ _] [CaseCtx _ _]) with "[]").
    iSplitL ""; first by iApply IHHtyped1; iFrame "#".
    iIntros (vv) "#Hvv". simpl.
    iDestruct "Hvv" as "[Hvv|Hvv]".
    - iDestruct "Hvv" as ([w w']) "[% Hw]"; simplify_eq; simpl.
      iApply logrel_pure_det_head_step.
      { intros σ; econstructor; eauto. }
      { intros ? ? ? Hrd. inversion Hrd; subst; auto. }
      { intros σ. econstructor; eauto. }
      asimpl. erewrite !n_closed_subst_head_simpl; (rewrite ?fmap_length; eauto).
      iNext. iApply (IHHtyped2 _ ((w,w') :: vvs)); iFrame "#".
      iApply interp_env_cons; auto.
    - iDestruct "Hvv" as ([w w']) "[% Hw]"; simplify_eq; simpl.
      iApply logrel_pure_det_head_step.
      { intros σ; econstructor; eauto. }
      { intros ? ? ? Hrd. inversion Hrd; subst; auto. }
      { intros σ. econstructor; eauto. }
      asimpl. erewrite !n_closed_subst_head_simpl; (rewrite ?fmap_length; eauto).
      iNext. iApply (IHHtyped3 _ ((w,w') :: vvs)); iFrame "#".
      iApply interp_env_cons; auto.
  Qed.

  Lemma bin_log_related_if Γ e0 e1 e2 e0' e1' e2' τ
      (IHHtyped1 : Γ e0 logNN e0' : TBool)
      (IHHtyped2 : Γ e1 logNN e1' : τ)
      (IHHtyped3 : Γ e2 logNN e2' : τ) :
    Γ If e0 e1 e2 logNN If e0' e1' e2' : τ.
  Proof.
    iIntros (Δ vvs ?) "#(Hinv & HΓ)". simpl.
    iApply ((logrel_bind [IfCtx _ _] [IfCtx _ _]) with "[]").
    iSplitL ""; first by iApply IHHtyped1; iFrame "#".
    iIntros ([v v']) "#Hvv". simpl.
    iDestruct "Hvv" as ([]) "[% %]"; simplify_eq; simpl in *.
    - iApply logrel_pure_det_head_step.
      { intros σ; econstructor; eauto. }
      { intros ? ? ? Hrd. inversion Hrd; subst; auto. }
      { intros σ. econstructor; eauto. }
      iNext. iApply IHHtyped2; iFrame "#".
    - iApply logrel_pure_det_head_step.
      { intros σ; econstructor; eauto. }
      { intros ? ? ? Hrd. inversion Hrd; subst; auto. }
      { intros σ. econstructor; eauto. }
      iNext. iApply IHHtyped3; iFrame "#".
  Qed.

  Lemma bin_log_related_int_binop Γ op e1 e2 e1' e2'
      (IHHtyped1 : Γ e1 logNN e1' : TInt)
      (IHHtyped2 : Γ e2 logNN e2' : TInt) :
    Γ BinOp op e1 e2 logNN BinOp op e1' e2' : binop_res_type op.
  Proof.
    iIntros (Δ vvs ?) "#(Hinv & HΓ)". simpl.
    iApply ((logrel_bind [BinOpLCtx _ _] [BinOpLCtx _ _]) with "[]").
    iSplitL ""; first by iApply IHHtyped1; iFrame "#".
    iIntros ([v v']) "#Hvv". simpl.
    iApply ((logrel_bind [BinOpRCtx _ _] [BinOpRCtx _ _]) with "[]").
    iSplitL ""; first by iApply IHHtyped2; iFrame "#".
    iIntros ([w w']) "#Hww". simpl.
    iDestruct "Hvv" as (z) "[% %]"; simplify_eq; simpl in *.
    iDestruct "Hww" as (z') "[% %]"; simplify_eq; simpl in *.
    iIntros (? ? ?) "? /=".
    iApply ic_bin_op; iNext; iModIntro; iExists _, _; simpl; iFrame; iSplit.
    - iPureIntro. eapply nsteps_once.
      eapply head_prim_step; simpl; econstructor; eauto.
    - destruct op; iExists _; iSplit; simpl; eauto.
  Qed.

  Lemma bin_log_related_rec Γ (e e' : expr) τ1 τ2
      (Hclosed : f, e.[upn (S (S (length Γ))) f] = e)
      (Hclosed' : f, e'.[upn (S (S (length Γ))) f] = e')
      (IHHtyped : TArrow τ1 τ2 :: τ1 :: Γ e logNN e' : τ2) :
    Γ Rec e logNN Rec e' : TArrow τ1 τ2.
  Proof.
    iIntros (Δ vvs ?) "#(Hinv & HΓ)". simpl.
    iDestruct (interp_env_length with "HΓ") as %?.
    iIntros (? ? ?) "? /=". value_case.
    iExists _, (RecV _); iFrame; iSplit; first by iPureIntro; econstructor.
    iAlways. iLöb as "IH". iIntros ([v v']) "#Hvv".
    iApply logrel_pure_det_head_step.
    { intros σ; econstructor; eauto. }
    { intros ? ? ? Hrd. inversion Hrd; subst; auto. }
    { intros σ. econstructor; eauto. }
    iNext. asimpl. change (Rec ?e) with (of_val (RecV e)).
    erewrite !n_closed_subst_head_simpl_2; (rewrite ?fmap_length; eauto).
    iApply (IHHtyped _ ((_,_) :: (v,v') :: vvs)); repeat iSplit; eauto.
    iApply interp_env_cons; iSplit; [|iApply interp_env_cons]; auto.
  Qed.

  Lemma bin_log_related_app Γ e1 e2 e1' e2' τ1 τ2
      (IHHtyped1 : Γ e1 logNN e1' : TArrow τ1 τ2)
      (IHHtyped2 : Γ e2 logNN e2' : τ1) :
    Γ App e1 e2 logNN App e1' e2' : τ2.
  Proof.
    iIntros (Δ vvs ?) "#(Hinv & HΓ)". simpl.
    iApply ((logrel_bind [AppLCtx _] [AppLCtx _]) with "[]").
    iSplitL ""; first by iApply IHHtyped1; iFrame "#".
    iIntros ([f f']) "#Hff". simpl.
    iApply ((logrel_bind [AppRCtx _] [AppRCtx _]) with "[]").
    iSplitL ""; first by iApply IHHtyped2; iFrame "#".
    iIntros ([v v']) "#Hvv". simpl.
    by iApply ("Hff" $! (v, v') with "").
  Qed.

  Lemma bin_log_related_tlam Γ e e' τ
      (IHHtyped : (subst (ren (+1)) <$> Γ) e logNN e' : τ) :
    Γ TLam e logNN TLam e' : TForall τ.
  Proof.
    iIntros (Δ vvs ?) "#(Hinv & HΓ)". simpl.
    iIntros (? ? ?) "? /=". value_case.
    iExists _, (TLamV _); iFrame.
    iSplit; first by iPureIntro; econstructor.
    iAlways. iIntrosi rn Hps).
    iApply logrel_pure_det_head_step.
    { intros σ; econstructor; eauto. }
    { intros ? ? ? Hrd. inversion Hrd; subst; auto. }
    { intros σ. econstructor; eauto. }
    iNext. iApply IHHtyped; iFrame "#". by iApply interpNN_env_ren.
  Qed.

  Lemma bin_log_related_tapp Γ e e' τ τ'
      (IHHtyped : Γ e logNN e' : TForall τ) :
    Γ TApp e logNN TApp e' : τ.[τ'/].
  Proof.
    iIntros (Δ vvs ?) "#(Hinv & HΓ)". simpl.
    iApply ((logrel_bind [TAppCtx] [TAppCtx]) with "[]").
    iSplitL ""; first by iApply IHHtyped; iFrame "#".
    iIntros ([v v']) "#Hvv". simpl.
    iIntros (? ? ?) "? /=".
    iApply ic_wand_r; iSplitL.
    { iSpecialize ("Hvv" $! (interpNN τ' Δ) ((type_to_reg τ') Δ) with "[#]");
        [iPureIntro; apply _|].
      iApply "Hvv"; eauto. }
    simpl.
    iIntros (w n). iDestruct 1 as (σ2' w') "(Hw & Ho & #Hww)".
    iExists _, _; iFrame. rewrite -interpNN_subst; eauto.
  Qed.

  Lemma bin_log_related_fold Γ e e' τ
      (IHHtyped : Γ e logNN e' : τ.[(TRec τ)/]) :
    Γ Fold e logNN Fold e' : TRec τ.
  Proof.
    iIntros (Δ vvs ?) "#(Hinv & HΓ)". simpl.
    iApply ((logrel_bind [FoldCtx] [FoldCtx] (TRec _)) with "[]").
    iSplitL ""; first by iApply IHHtyped; iFrame "#".
    iIntros ([v v']) "#Hvv". simpl.
    iIntros (? ? ?) "? /=". value_case.
    iExists _, (FoldV _); iFrame; iSplit; first by iPureIntro; econstructor.
    rewrite fixpoint_unfold /= -interpNN_subst.
    iAlways. iExists (_, _); iSplit; eauto.
  Qed.

  Lemma bin_log_related_unfold Γ e e' τ
      (IHHtyped : Γ e logNN e' : TRec τ) :
    Γ Unfold e logNN Unfold e' : τ.[(TRec τ)/].
  Proof.
    iIntros (Δ vvs ?) "#(Hinv & HΓ)". simpl.
    iApply ((logrel_bind [UnfoldCtx] [UnfoldCtx]) with "[]").
    iSplitL ""; first by iApply IHHtyped; iFrame "#".
    iIntros ([v v']) "#Hvv". simpl.
    rewrite /= fixpoint_unfold /=.
    change (fixpoint _) with (interpNN (TRec τ) Δ).
    iDestruct "Hvv" as ([w w']) "#[% Hww]"; simplify_eq/=.
    iApply logrel_pure_det_head_step.
    { intros σ; econstructor; eauto. }
    { intros ? ? ? Hrd. inversion Hrd; subst; auto. }
    { intros σ. econstructor; eauto. }
    iNext. iIntros (? ? ?) "? /=". value_case.
    iExists _, _; iFrame; iSplit; first by iPureIntro; econstructor.
    by rewrite -interpNN_subst.
  Qed.

  Lemma stupid e s e' s': det_head_step e s e' s' [] ->
                          det_head_step ((e,s).1) ((e,s).2)
                                        ((e',s').1) ((e',s').2) [].
  Proof. done. Qed.

  Lemma bin_log_related_alloc Γ e e' ρ τ
      (IHHtyped : Γ e logNN e' : τ) :
    Γ Alloc e logNN Alloc e' : TST ρ (TSTref ρ τ).
  Proof.
    iIntros (Δ vvs ?) "#(Hinv & HΓ)". simpl.
    iApply ((logrel_bind [AllocCtx] [AllocCtx]) (TST _ (TSTref _ _)) with "[]").
    iSplitL ""; first by iApply IHHtyped; iFrame "#".
    iIntros ([v v']) "#Hvv". simpl.
    iIntros (? ? ?) "? /=". value_case.
    iExists _, (AllocV _); iFrame; iSplit; first by iPureIntro; econstructor.
    iAlways. iIntros (γh γh' σ1') "[Ho Hreg]".
    iApply ic_runst. iApply ic_value'. simpl.
    iApply ic_alloc; auto.
    iNext. iIntros (l) "Hl".
    iMod ((mapsto_alloc _ _ (fresh (dom (gset loc) σ1')) v') with "Ho")
      as "[Ho Hl']".
    { rewrite -(@not_elem_of_dom _ _ (gset loc)). apply is_fresh. }
    iApply fupd_ic.
    iMod (Region_extend with "[$Hreg $Hl $Hl']") as "[Hreg Hrefrel]";
      trivial; iFrame "#".
    iApply ic_return. iNext.
    iExists _, (LitV $ LitLoc (fresh (dom (gset loc) σ1'))).
    iFrame. iPureIntro. eapply nsteps_r. eapply nsteps_once.
    eapply head_prim_step; simpl. eapply stupid.
    eapply (det_EffS σ1' _ (Alloc (of_val v'))).
    eapply head_prim_step; simpl; econstructor; eauto.
    eapply head_prim_step; simpl. eapply det_RunRet; eauto.
  Qed.

  Lemma bin_log_related_read Γ e e' ρ τ
      (IHHtyped : Γ e logNN e' : (TSTref ρ τ)) :
    Γ Read e logNN Read e' : TST ρ τ.
  Proof.
    iIntros (Δ vvs ?) "#(Hinv & HΓ)". simpl.
    iApply ((logrel_bind [ReadCtx] [ReadCtx]) (TST _ _) with "[]").
    iSplitL ""; first by iApply IHHtyped; iFrame "#".
    iIntros ([v v']) "#Hvv". simpl.
    iIntros (? ? ?) "? /=". value_case.
    iExists _, (ReadV _); iFrame; iSplit; first by iPureIntro; econstructor.
    iAlways. iIntros (γh γh' σ1') "[Ho Hreg]".
    iApply ic_runst. iApply ic_value'. simpl.
    iDestruct (STRefREL_locs with "[]") as %(l & l' & ? & ?);
      first iFrame "#"; subst.
    iDestruct (region_open with "[$Hreg]")
      as (w w') "(Horeg & Hl & Hl' & #Hrel)"; iFrame "#".
    iApply (ic_read with "[-]"); auto.
    iNext; iFrame. iIntros "Hl".
    iDestruct (mapsto_allocated with "[$Ho $Hl']") as %His.
    iDestruct (region_close with "[$Horeg $Hl $Hl']") as "Hreg"; iFrame "#".
    iApply ic_return. iNext. iExists _, w'.
    iFrame; iFrame "#". iPureIntro.
    eapply nsteps_r. eapply nsteps_once.
    eapply head_prim_step; simpl. eapply stupid. econstructor.
    eapply head_prim_step; simpl; econstructor; eauto.
    eapply head_prim_step; simpl. eapply det_RunRet; eauto.
  Qed.

  Lemma bin_log_related_write Γ e1 e2 e1' e2' ρ τ
      (IHHtyped1 : Γ e1 logNN e1' : (TSTref ρ τ))
      (IHHtyped2 : Γ e2 logNN e2' : τ) :
    Γ Write e1 e2 logNN Write e1' e2' : TST ρ TUnit.
  Proof.
    iIntros (Δ vvs ?) "#(Hinv & HΓ)". simpl.
    iApply ((logrel_bind [WriteLCtx _] [WriteLCtx _]) (TST _ TUnit) with "[]").
    iSplitL ""; first by iApply IHHtyped1; iFrame "#".
    iIntros ([v v']) "#Hvv". simpl.
    iApply ((logrel_bind [WriteRCtx _] [WriteRCtx _]) (TST _ TUnit) with "[]").
    iSplitL ""; first by iApply IHHtyped2; iFrame "#".
    iIntros ([w w']) "#Hww". simpl.
    iIntros (? ? ?) "? /=". value_case.
    iExists _, (WriteV _ _); iFrame; iSplit; first by iPureIntro; econstructor.
    iAlways. iIntros (γh γh' σ1') "[Ho Hreg]".
    iApply ic_runst. iApply ic_value'. simpl.
    iDestruct (STRefREL_locs with "[]")
      as %(l & l' & ? & ?); first iFrame "#"; subst. simpl.
    iDestruct (region_open with "[$Hreg]")
      as (u u') "(Horeg & Hl & Hl' & #Hrel)"; iFrame "#".
    iApply (ic_write with "[-]"); auto.
    iNext; iFrame. iIntros "Hl".
    iDestruct (mapsto_allocated with "[$Ho $Hl']") as %His.
    iMod ((mapsto_update _ _ _ _ w') with "[$Ho $Hl']") as "[Ho Hl']".
    iDestruct (region_close with "[$Horeg $Hl $Hl']") as "Hreg"; iFrame "#".
    iApply ic_return. iNext. iExists _, (LitV LitUnit).
    iFrame; iFrame "#". iSplit; eauto. iPureIntro.
    eapply nsteps_r. eapply nsteps_once.
    eapply head_prim_step; simpl. eapply stupid. econstructor.
    eapply head_prim_step; simpl; econstructor; eauto.
    eapply head_prim_step; simpl. eapply det_RunRet; eauto.
  Qed.

  Lemma bin_log_related_bind Γ e1 e2 e1' e2' ρ τ τ'
      (IHHtyped1 : Γ e1 logNN e1' : (TST ρ τ))
      (IHHtyped2 : Γ e2 logNN e2' : TArrow τ (TST ρ τ')) :
    Γ Bind e1 e2 logNN Bind e1' e2' : TST ρ τ'.
  Proof.
    iIntros (Δ vvs ?) "#(Hinv & HΓ)". simpl.
    iApply ((logrel_bind [BindLCtx _] [BindLCtx _]) (TST _ _) with "[]").
    iSplitL ""; first by iApply IHHtyped1; iFrame "#".
    iIntros ([v v']) "#Hvv". simpl.
    iApply ((logrel_bind [BindRCtx _] [BindRCtx _]) (TST _ _) with "[]").
    iSplitL ""; first by iApply IHHtyped2; iFrame "#".
    iIntros ([w w']) "#Hww". simpl.
    iIntros (? ? ?) "? /=". value_case.
    iExists _, (BindV _ _); iFrame; iSplit; first by iPureIntro; econstructor.
    iAlways. iIntros (γh γh' σ1') "[Ho Hreg]".
    iApply (ic_runst_to_bind with "[-]").
    iApply (ic_wand_r with "[-]"); iSplitL. iApply "Hvv"; iFrame. simpl.
    iIntros (u n). iDestruct 1 as (σ2' u') "(Hred & Ho & #Huu & Hreg)".
    iDestruct "Hred" as %Hred.
    iApply (@ic_bind _ _ _ _ (fill_item RunSTCtx)); simpl.
    iApply (ic_wand_r with "[-]"); iSplitL "Ho".
    iSpecialize ("Hww" $! _ with "[$Huu]").
    iApply ("Hww" with "[$Ho]"). simpl.
    iIntros (u2 m). iDestruct 1 as (σ2'' u3) "(Hred & Ho & #Huu3)".
    iDestruct "Hred" as %Hred'.
    iApply (ic_wand_r with "[-]"); iSplitL.
    iApply ("Huu3" $! _ _ _ with "[-]"); iFrame.
    iIntros (u4 nn). iDestruct 1 as (σ3 u5) "(Hred & Ho & #Huu5 & Hreg)".
    iDestruct "Hred" as %Hred''.
    iExists _, _; iFrame; iFrame "#".
    iPureIntro.
    eapply red_RunST_Bind'; eauto.
    eapply nsteps_trans; eauto.
    eapply (@nsteps_ctx det_lang (fill_item RunSTCtx));
      eauto; typeclasses eauto.
  Qed.

  Lemma bin_log_related_return Γ e e' ρ τ
        (IHHtyped : Γ e logNN e' : τ) :
    Γ Return e logNN Return e' : TST ρ τ.
  Proof.
    iIntros (Δ vvs ?) "#(Hinv & HΓ)". simpl.
    iApply ((logrel_bind [ReturnCtx] [ReturnCtx]) (TST _ _) with "[]").
    iSplitL ""; first by iApply IHHtyped; iFrame "#".
    iIntros ([v v']) "#Hvv". simpl.
    iIntros (? ? ?) "? /=". value_case.
    iExists _, (ReturnV _); iFrame; iSplit; first by iPureIntro; econstructor.
    iAlways. iIntros (γh γh' σ1') "[Ho Hreg]".
    iApply ic_runst. iApply ic_value'. simpl.
    (* change (Return  ReturnV v). *)
    iApply ((@ic_value eff_lang) with "[-]").
    rewrite /= to_of_val; eauto. simpl.
    iApply ic_return. iNext. iExists _, _.
    iFrame; iFrame "#". iPureIntro.
    eapply nsteps_once.
    eapply head_prim_step; simpl. econstructor; eauto.
  Qed.

  Lemma bin_log_related_runST Γ e e' τ
        (IHtyped : (subst (ren (+1)) <$> Γ) e logNN e'
         : TST (TVar 0) τ.[ren (+1)]):
    Γ RunST e logNN RunST e' : τ.
  Proof.
    iIntros (Δ vvs ?) "#(Hinv & HΓ)". simpl.
    iIntros (γh γh' σ1') "Hσ1' /=".
    iApply fupd_ic.
    iMod ((region_alloc _ γh γh') with "[]")
      as (rn) "Hreg"; first done; iFrame "#".
    iModIntro.
    change (RunST e.[env_subst (vvs.*1)]) with
      (fill [RunSTCtx] e.[env_subst (vvs.*1)]).
    iApply (ic_ectx_bind with "[-]").
    iApply (ic_wand_r with "[-]"); iSplitL "Hσ1'".
    iPoseProof (IHtyped ((interpNN (TUnit) Δ, rn) :: Δ) (vvs) with "[]") as "Hiht".
    { rewrite interpNN_env_ren; eauto. }
    iApply "Hiht"; iFrame. simpl.
    iIntros (v n) "Hc". iDestruct "Hc" as (σ2' v') "(Hred & Ho & Hvv')".
    iDestruct "Hred" as %Hred.
    iApply (ic_wand_r with "[-]"); iSplitR "".
    iApply ("Hvv'" with "[-]"); rewrite type_to_reg_var /=; iFrame.
    simpl.
    iIntros (w m) "Hc". iDestruct "Hc" as (σ3' w') "(Hred & Ho & Hww' & Hreg)".
    iDestruct "Hred" as %Hred'.
    iExists σ3', w'; iFrame.
    iSplit.
    - iPureIntro.
      eapply (@nsteps_bind' det_lang); eauto.
      eapply (@ectx_lang_ctx _ _ _ _ det_ectx_lang [RunSTCtx]).
    - by rewrite (interpNN_weaken [] [_]).
  Qed.

  Lemma bin_log_related_compare Γ e1 e2 e1' e2' ρ τ
        (IHtyped1 : Γ e1 logNN e1' : TSTref ρ τ)
        (IHtyped2 : Γ e2 logNN e2' : TSTref ρ τ) :
    Γ Compare e1 e2 logNN Compare e1' e2' : TBool.
  Proof.
    iIntros (Δ vvs ?) "#(Hinv & HΓ)". simpl.
    iApply ((logrel_bind [CompLCtx _] [CompLCtx _]) TBool with "[]").
    iSplitL ""; first by iApply IHtyped1; iFrame "#".
    iIntros ([v v']) "#Hvv". simpl.
    iApply ((logrel_bind [CompRCtx _] [CompRCtx _]) TBool with "[]").
    iSplitL ""; first by iApply IHtyped2; iFrame "#".
    iIntros ([w w']) "#Hww". simpl.
    iDestruct (STRefREL_locs with "[]")
      as %(l1 & l1' & ? & ?); first iFrame "Hvv"; subst.
    iDestruct (STRefREL_locs with "[]")
      as %(l2 & l2' & ? & ?); first iFrame "Hww"; subst.
    iIntros (γh γh' σ1') "Hσ1' /=".
    iApply fupd_ic.
    iMod (related_pairs_eq_iff with "[]") as %Heq; first auto;
      first iFrame "Hvv Hww Hinv".
    iModIntro. iRevert (γh γh' σ1') "Hσ1'". simpl.
    destruct (decide (l1 = l2)); subst.
    - rewrite (proj1 Heq) //.
      iApply (logrel_pure_det_head_step _ _ _ _ TBool Δ).
      { intros σ; econstructor; eauto. }
      { intros ? ? ? Hrd. by inversion Hrd; subst; auto. }
      { intros σ. econstructor; eauto. }
      iNext. iIntros (? ? ?) "? /=". value_case.
      iExists _, _; iFrame; iSplit; eauto.
      iPureIntro. econstructor.
    - assert (Hneq : l1' l2') by by intros Heq'%Heq.
      iApply (logrel_pure_det_head_step _ _ _ _ TBool Δ).
      { intros σ; econstructor; eauto. }
      { intros ? ? ? Hrd. by inversion Hrd; subst; auto. }
      { intros σ. econstructor; eauto. }
      iNext. iIntros (? ? ?) "? /=". value_case.
      iExists _, _; iFrame; iSplit; eauto.
      iPureIntro. econstructor.
  Qed.

  Theorem binary_fundamental Γ e τ :
    Γ ⊢ₜ e : τ Γ e logNN e : τ.
  Proof.
    induction 1.
    - by apply bin_log_related_var.
    - by apply bin_log_related_unit.
    - by apply bin_log_related_bool.
    - by apply bin_log_related_int.
    - apply bin_log_related_int_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;
        match goal with H : _ |- _ => eapply (typed_n_closed _ _ _ H) end.
    - eapply bin_log_related_if; eauto.
    - eapply bin_log_related_rec; eauto;
        match goal with H : _ |- _ => eapply (typed_n_closed _ _ _ H) end.
    - eapply bin_log_related_app; eauto.
    - eapply bin_log_related_tlam; eauto with typeclass_instances.
    - eapply bin_log_related_tapp; eauto.
    - eapply bin_log_related_fold; eauto.
    - eapply bin_log_related_unfold; eauto.
    - eapply bin_log_related_alloc; eauto.
    - eapply bin_log_related_read; eauto.
    - eapply bin_log_related_write; eauto.
    - eapply bin_log_related_bind; eauto.
    - eapply bin_log_related_return; eauto.
    - eapply bin_log_related_runST; eauto.
    - eapply bin_log_related_compare; eauto.
  Qed.
End fundamental.