RunST.tlam_hoisting

From RunST Require Import IC base lang rules typing reduction regions
     future contextual_refinement saved_pred ST_Lang_reduction
     logrel_shared soundness soundnessNN logrelNN fundamentalNN
     logrel fundamental.
From iris.proofmode Require Import tactics.

Hint Resolve to_of_val.

Lemma TLam_Hoisting_LR1 `{ICG_ST Σ} `{LogRelG Σ} Γ e1 e1' τ e2 e2' τ'' :
  Γ ⊢ₜ e1 : τ subst (ren (+1)) <$> (τ :: Γ) ⊢ₜ e2 : τ''
  Γ ⊢ₜ e1' : τ subst (ren (+1)) <$> (τ :: Γ) ⊢ₜ e2' : τ''
  Γ e1 log e1' : τ subst (ren (+1)) <$> (τ :: Γ) e2 log e2' : τ''
  Γ App (Rec (TLam e2).[ren (+1)]) e1 log
      App (Rec (TLam (App (Rec e2'.[ren (+2)]) e1'.[ren (+1)])).[ren (+1)]) e1'
  : TForall τ''.
Proof.
  intros Ht1 Ht2 Ht1' Ht2' HLR1 HLR2.
  iIntros (Δ vvs ) "[#Hinv #Hvvs]".
  iIntros (γh γh' σ1') "Hσ1' /=". asimpl.
  rewrite ic_eq /ic_def /=. iIntros (σ1 σ2 v n) "[Hrd1 How]".
  iDestruct "Hrd1" as %Hrd1.
  apply (nsteps_bind (fill_item (AppRCtx (RecV _)))) in Hrd1.
  destruct Hrd1 as (k & σ3 & w & Hk & Hrd11 & Hrd12). simpl in Hrd11, Hrd12.
  set (Q := (|={}=> Nat.iter k (λ Q, Q) ( ( σ1', (σ2' : state) (v' : val),
           (rtc (pstep' det_lang)
          (e1'.[env_subst (vvs.*2)], σ1') (of_val v', σ2')
          to_heap σ1' to_heap σ2' to_heap σ2'))))%I : iProp Σ).
  iAssert Q as "HQ".
  { unfold Q. rewrite later_n_except_0_forall.
    iApply (fupd_plain_forall_comm (E2 := )).
    iIntros (σ3').
    iMod (allocate_full_state σ3') as (γh3) "[Hnfσ3 Hnoσ3]".
    iMod (allocate_full_state σ1) as (γh4) "[Hnfσ1 Hnoσ1]".
    iPoseProof (HLR1 with "[]") as "He1"; iFrame "#".
    iSpecialize ("He1" $! _ _ _ with "[$Hnfσ3]"); simpl.
    rewrite ic_eq /ic_def /=.
    iSpecialize ("He1" $! _ _ _ _ with "[Hnfσ1]"); iFrame; auto.
    iApply (future_plain with "[-]").
    iMod "He1" as "[He1 _]". iModIntro.
    iDestruct "He1" as (σ4' ?) "(Hrd2&Hnfσ3&?)".
    iDestruct "Hrd2" as %Hrd2.
    iDestruct (heap_included with "[$Hnoσ3 $Hnfσ3]") as %[??].
    iExists _, _; iPureIntro; repeat split; eauto. }
  iMod "HQ" as "HQ". iModIntro. clear Q.
  iPoseProof (HLR1 with "[]") as "He1"; iFrame "#".
  iSpecialize ("He1" $! _ _ _ with "[$Hσ1']"); simpl.
  rewrite ic_eq /ic_def /=.
  iSpecialize ("He1" $! _ _ _ _ with "[How]"); iFrame; auto.
  iDestruct (later_n_except_0_future (E := ) with "[$HQ]") as "HQ".
  iCombine "HQ" "He1" as "He1".
  iMod "He1" as "(#HQ & [He1 HPh1])".
  iDestruct "He1" as (σ2'' v'' ) "(Hrd21 & Ho1 & #Hrl1)".
  iDestruct "Hrd21" as %Hrd21.
  apply rtc_nsteps in Hrd21. destruct Hrd21 as [k1 Hrd21].
  apply head_red_app_TLam in Hrd12. destruct Hrd12 as [? [? ?]]; subst.
  iModIntro; iFrame. asimpl.
  iExists σ2'', (TLamV _); iSplit; iFrame; eauto.
  { iPureIntro. etrans.
    - eapply (nsteps_rtc _).
      eapply (@nsteps_ctx det_lang (fill_item (AppRCtx (RecV _))));
        eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
    - eapply (rtc_l _ (_, _) (_, _)); simpl.
      { apply head_prim_step. econstructor; eauto. }
      asimpl. econstructor. }
  iAlways. iIntrosi rn) "Hvv". iDestruct "Hvv" as %Hvv. asimpl.
  iIntros (γh2 γh3 σ3') "Hσ3'".
  iApply ic_tlam. iNext.
  rewrite ic_eq /ic_def /=.
  iIntros (σ4 σ5 v k') "(Hrd3 & Hσ4)". iDestruct "Hrd3" as %Hrd3.
  iDestruct (interp_env_length with "Hvvs") as %Hvvsl.
  iPoseProof (HLR2 ((τi, rn) :: Δ) ((w, v'') :: vvs) with "[]") as "He2".
  { rewrite interp_env_ren interp_env_cons; iFrame "#". }
  iDestruct ("HQ" $! σ3') as (σ3'' u'') "HQ'".
  iDestruct "HQ'" as %(Hrd4 & Hincl & Hvl).
  iMod (heap_catch_up with "Hσ3'") as "Hσ3''"; eauto.
  iSpecialize ("He2" $! _ _ _ with "[$Hσ3'']"); simpl.
  rewrite ic_eq /ic_def /=.
  iSpecialize ("He2" $! _ _ _ _ with "[Hσ4]"); iFrame; auto.
  erewrite <- typed_subst_head_simpl; eauto;
    last by simpl; rewrite ?map_length Hvvsl.
  iMod "He2" as "[He2 Hσ5]"; iModIntro; iFrame.
  iDestruct "He2" as (σ5' v') "(Hrd5 & Hσ5' & #Hvv)".
  iDestruct "Hrd5" as %Hrd5.
  erewrite <- typed_subst_head_simpl in Hrd5; eauto;
    last by simpl; rewrite ?map_length Hvvsl.
  iExists _, _; iFrame; iFrame "#".
  iPureIntro.
  etrans.
  - eapply (rtc_l _ _ (_, _) (_, _)); simpl.
    eapply head_prim_step; econstructor.
    econstructor.
  - etrans.
    eapply (@rtc_ctx det_lang (fill_item (AppRCtx (RecV _))));
      eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
    etrans.
    eapply (rtc_l _ _ (_, _) (_, _)); simpl.
    eapply head_prim_step; econstructor; eauto.
    econstructor.
    by asimpl.
Qed.

Lemma TLam_Hoisting_LR2 `{ICG_ST Σ} `{LogRelG Σ} Γ e1' τ e2' τ'' :
  Γ ⊢ₜ e1' : τ subst (ren (+1)) <$> (τ :: Γ) ⊢ₜ e2' : τ''
  Γ App (Rec (TLam (App (Rec e2'.[ren (+2)]) e1'.[ren (+1)])).[ren (+1)]) e1'
      logNN
      App (Rec (TLam (App (Rec e2'.[ren (+1)])
                    e1')).[ren (+2)]) e1'
  : TForall τ''.
Proof.
  iIntros (Ht1 Ht2 Δ vvs ) "#(Hinv & Hvvs)".
  iIntros (γh γh' σ1') "Hσ1'". asimpl.
  rewrite ic_eq /ic_def /=. iIntros (σ1 σ2 v n) "[Hrd1 How]".
  iDestruct "Hrd1" as %Hrd1.
  apply (nsteps_bind (fill_item (AppRCtx (RecV _)))) in Hrd1.
  destruct Hrd1 as (k & σ3 & w & Hk & Hrd11 & Hrd12).
  iPoseProof (fundamentalNN.binary_fundamental _ _ _ Ht1 with "[]") as "He1";
    iFrame "#".
  iSpecialize ("He1" $! _ _ _ with "[$Hσ1']"); simpl.
  rewrite ic_eq /ic_def /=.
  iSpecialize ("He1" $! _ _ _ _ with "[How]"); iFrame; auto.
  iMod "He1" as "[Hrd How]". iDestruct "Hrd" as (σ2' w') "(Hrd2 & Hσ2' & #Hww)".
  iDestruct "Hrd2" as %Hrd2.
  apply head_red_app_TLam in Hrd12; destruct Hrd12 as [? [? ?]]; subst.
  iModIntro; iFrame. iExists _, (TLamV _); iFrame; iSplit.
  { iPureIntro. asimpl.
    replace n with (k + 1) by omega; eapply nsteps_trans.
    - eapply (@nsteps_ctx det_lang (fill_item (AppRCtx (RecV _))));
        eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
    - eapply (nsteps_l _ _ _ (_, _) (_, _)); simpl.
      { apply head_prim_step. econstructor; eauto. }
      asimpl. econstructor. }
  asimpl.
  iAlways. iIntrosi rn) "Hvv". iDestruct "Hvv" as %Hvv. asimpl.
  iIntros (γh2 γh3 σ3') "Hσ3'".
  iApply ic_tlam. iNext.
  rewrite ic_eq /ic_def /=.
  iIntros (σ4 σ5 v k') "(Hrd3 & Hσ4)". iDestruct "Hrd3" as %Hrd3.
  iDestruct (interp_env_length with "Hvvs") as %Hvvsl.
  apply (nsteps_bind (fill_item (AppRCtx (RecV _)))) in Hrd3.
  destruct Hrd3 as (m & σ6 & w3 & Hm & Hrd31 & Hrd32).
  iMod (allocate_full_state σ4) as (γ3) "[H341 _]".
  iMod (allocate_full_state σ1') as (γ4) "[H411 _]".
  iPoseProof (fundamentalNN.binary_fundamental _ _ _ Ht1 with "[]") as "He1";
    iFrame "#".
  iSpecialize ("He1" $! _ _ _ with "[$H411]"); simpl.
  rewrite ic_eq /ic_def /=.
  iSpecialize ("He1" $! _ _ _ _ with "[$H341]"); auto.
  iMod (allocate_full_state σ1) as (γ5) "[H511 _]".
  iMod (allocate_full_state σ3') as (γ6) "[H611 _]".
  iPoseProof (fundamentalNN.binary_fundamental _ _ _ Ht1 with "[]") as "He1'";
    iFrame "#".
  iSpecialize ("He1'" $! _ _ _ with "[$H611]"); simpl.
  rewrite ic_eq /ic_def /=.
  iSpecialize ("He1'" $! _ _ _ _ with "[$H511]"); auto.
  iPoseProof (fundamentalNN.binary_fundamental _ _ _ Ht1 with "[]") as "He1''";
    iFrame "#".
  iSpecialize ("He1''" $! _ _ _ with "[$Hσ3']"); simpl.
  rewrite ic_eq /ic_def /=.
  iSpecialize ("He1''" $! _ _ _ _ with "[$Hσ4]"); auto.
  iCombine "He1" "He1''" as "He1".
  iPoseProof (future_cancel_2 with "[He1 He1']") as "He"; first trivial.
  { iSplitL "He1'"; iFrame. }
  iMod "He" as "[He1 [[He2 Hσ26] [He3 Hσ36]]]".
  iDestruct "He2" as (σ7' v7') "(Hrd7 & Hσ7' & #Hvv7)".
  iDestruct "Hrd7" as %Hrd7.
  destruct (nsteps_deterministic' Hrd7 Hrd2) as [?[??]]; subst.
  iDestruct "He3" as (σ8' v8') "(Hrd8 & Hσ8' & #Hvv8)".
  iDestruct "Hrd8" as %Hrd8.
  replace (k- k) with 0 by omega. rewrite future_unfold_O.
  iMod "He1" as "[He1 Hσ5]".
  iDestruct "He1" as (σ9' v9') "(Hrd9 & Hσ9' & #Hvv9)".
  iDestruct "Hrd9" as %Hrd9.
  destruct (nsteps_deterministic' Hrd8 Hrd9) as [?[??]]; subst.
  apply rec_red_step in Hrd32. destruct Hrd32 as [m [Hkm Hrd32]].
  asimpl in Hrd32.
  erewrite typed_subst_head_simpl in Hrd32; eauto;
    last by rewrite /= ?map_length -Hvvsl.
  iPoseProof (fundamentalNN.binary_fundamental _ _ _
                Ht2 ((τi, rn) :: Δ) (((w, v9') :: vvs)) with "[]") as "He2".
  { rewrite interpNN_env_ren interp_env_cons; iFrame "#". }
  iSpecialize ("He2" $! _ _ _ with "[$Hσ8']"); simpl.
  rewrite ic_eq /ic_def /=.
  iSpecialize ("He2" $! _ _ _ _ with "[$Hσ36]"); auto.
  iModIntro. iMod "He2" as "[He2 Hσ8']".
  iDestruct "He2" as (σ10 v10) "(Hrd10 & Hσ10 & #Hvv10)".
  iDestruct "Hrd10" as %Hrd10.
  iModIntro; iFrame. iExists _, _; iFrame; iFrame "#".
  iPureIntro.
  replace k' with (k + S m) by omega.
  eapply (nsteps_l _ _ (_, _) (_, _)); simpl.
  { apply head_prim_step; econstructor. }
  eapply nsteps_trans.
  + eapply (@nsteps_ctx det_lang (fill_item (AppRCtx (RecV _))));
      eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
  + eapply (nsteps_l _ _ (_, _) (_, _)); simpl.
    { apply head_prim_step. econstructor; eauto. }
    asimpl.
    erewrite <- typed_subst_head_simpl in Hrd10; eauto;
        last rewrite /= map_length fmap_length -Hvvsl //.
Qed.

Lemma TLam_Hoisting_LR3 `{ICG_ST Σ} `{LogRelG Σ} Γ e1' τ e2' τ'' :
  Γ ⊢ₜ e1' : τ subst (ren (+1)) <$> (τ :: Γ) ⊢ₜ e2' : τ''
  Γ App (Rec (TLam (App (Rec e2'.[ren (+1)]) e1')).[ren (+2)]) e1'
      log
      (TLam (App (Rec e2'.[ren (+1)]) e1')) : TForall τ''.
Proof.
  iIntros (Ht1 Ht2 Δ vvs ) "#(Hinv & Hvvs)".
  iIntros (γh γh' σ1') "Hσ1'". asimpl.
  rewrite ic_eq /ic_def /=. iIntros (σ1 σ2 v n) "[Hrd1 How]".
  iDestruct "Hrd1" as %Hrd1.
  apply (nsteps_bind (fill_item (AppRCtx (RecV _)))) in Hrd1.
  destruct Hrd1 as (k & σ3 & w & Hk & Hrd11 & Hrd12).
  iPoseProof (binary_fundamental _ _ _ Ht1 with "[]") as "He1";
    iFrame "#".
  iMod (allocate_full_state σ1') as (γh3) "[Hnfσ1' Hnoσ1']".
  iSpecialize ("He1" $! _ _ _ with "[$Hnfσ1']"); simpl.
  rewrite ic_eq /ic_def /=.
  iSpecialize ("He1" $! _ _ _ _ with "[How]"); iFrame; auto.
  iMod "He1" as "[Hrd How]". iDestruct "Hrd" as (σ2' w') "(Hrd2 & Hσ2' & #Hww)".
  iDestruct "Hrd2" as %Hrd2.
  apply rtc_nsteps in Hrd2. destruct Hrd2 as [k' Hrd2].
  apply head_red_app_TLam in Hrd12; destruct Hrd12 as [? [? ?]]; subst.
  iModIntro; iFrame. iExists _, (TLamV _); iFrame; iSplit.
  { iPureIntro. asimpl. econstructor 1. }
  asimpl.
  replace
    (TLam (App (Rec
       e2'.[up (env_subst (vvs.*1)) >> ren (+1)])
         e1'.[env_subst (vvs.*1)]))
  with
  (TLam (App (Rec e2'.[ren (+1)])
            e1')).[env_subst (vvs.*1)]
    by by asimpl.
  replace
    (TLam (App (Rec
       e2'.[up (env_subst (vvs.*2)) >> ren (+1)])
         e1'.[env_subst (vvs.*2)]))
  with
  (TLam (App (Rec e2'.[ren (+1)])
            e1')).[env_subst (vvs.*2)]
    by by asimpl.
  assert (Γ ⊢ₜ (TLam (App (Rec e2'.[ren (+1)]) e1')) : TForall τ'') as Ht2'.
  { constructor. econstructor; last by apply context_rename.
    econstructor. eapply (context_weakening [_]); auto. }
  iClear "Hnoσ1' Hσ2'". iAlways.
  iIntrosi rn Hτi).
  iPoseProof (binary_fundamental _ _ _ Ht2' with "[]") as "Hff";
    first by iFrame "#".
  iApply (logrel_bind [TAppCtx] [TAppCtx]).
  iSplitL; first iExact "Hff".
  by iIntros (ff) "Hff"; iApply "Hff".
Qed.

Lemma TLam_Hoisting_CR1 Γ e1 e1' τ e2 e2' τ'' :
  Γ ⊢ₜ e1 : τ subst (ren (+1)) <$> (τ :: Γ) ⊢ₜ e2 : τ''
  Γ ⊢ₜ e1' : τ subst (ren (+1)) <$> (τ :: Γ) ⊢ₜ e2' : τ''
  ( `{ICG_ST Σ} `{LogRelG Σ}, Γ e1 log e1' : τ)
  ( `{ICG_ST Σ} `{LogRelG Σ},
      subst (ren (+1)) <$> (τ :: Γ) e2 log e2' : τ'')
  Γ App (Rec (TLam e2).[ren (+1)]) e1 ctx
      App (Rec (TLam (App (Rec e2'.[ren (+2)]) e1'.[ren (+1)])).[ren (+1)]) e1'
  : TForall τ''.
Proof.
  intros Ht1 Ht2 Ht1' Ht2' HLR1 HLR2.
  apply binary_soundness; eauto.
  - econstructor; last eauto. econstructor.
    eapply (context_weakening [_]). econstructor; eauto.
  - econstructor; eauto. econstructor.
    eapply (context_weakening [_]).
    econstructor.
    econstructor;
      last by asimpl; eapply (context_weakening [_]); apply context_rename.
    econstructor; eapply (context_weakening [_; _]); eauto.
  - intros. eapply TLam_Hoisting_LR1; eauto.
Qed.

Lemma TLam_Hoisting_CR2 Γ e1' τ e2' τ'' :
  Γ ⊢ₜ e1' : τ subst (ren (+1)) <$> (τ :: Γ) ⊢ₜ e2' : τ''
  Γ App (Rec (TLam (App (Rec e2'.[ren (+2)]) e1'.[ren (+1)])).[ren (+1)]) e1'
      ctx
      App (Rec (TLam (App (Rec e2'.[ren (+1)])
                    e1')).[ren (+2)]) e1'
  : TForall τ''.
Proof.
  intros Ht1 Ht2.
  apply binary_soundnessNN; eauto.
  - econstructor; eauto. econstructor.
    eapply (context_weakening [_]).
    econstructor.
    econstructor;
      last by asimpl; eapply (context_weakening [_]); apply context_rename.
    econstructor; eapply (context_weakening [_; _]); eauto.
  - econstructor; eauto. econstructor.
    eapply (context_weakening [_; _]).
    econstructor; eauto.
    econstructor; last by eapply context_rename.
    econstructor. eapply (context_weakening [_]); auto.
  - intros. eapply TLam_Hoisting_LR2; eauto.
Qed.

Lemma TLam_Hoisting_CR3 Γ e1' τ e2' τ'' :
  Γ ⊢ₜ e1' : τ subst (ren (+1)) <$> (τ :: Γ) ⊢ₜ e2' : τ''
  Γ App (Rec (TLam (App (Rec e2'.[ren (+1)]) e1')).[ren (+2)]) e1'
      ctx
      (TLam (App (Rec e2'.[ren (+1)]) e1'))
  : TForall τ''.
Proof.
  intros Ht1 Ht2.
  apply binary_soundness; eauto.
  - econstructor; eauto. econstructor.
    eapply (context_weakening [_; _]).
    econstructor; eauto.
    econstructor; last by eapply context_rename.
    econstructor. eapply (context_weakening [_]); auto.
  - econstructor. econstructor; last by eapply context_rename.
    econstructor. eapply (context_weakening [_]); auto.
  - intros. eapply TLam_Hoisting_LR3; eauto.
Qed.

Lemma TLam_Hoisting_basic Γ e1 e1' τ e2 e2' τ'' :
  Γ ⊢ₜ e1 : τ subst (ren (+1)) <$> (τ :: Γ) ⊢ₜ e2 : τ''
  Γ ⊢ₜ e1' : τ subst (ren (+1)) <$> (τ :: Γ) ⊢ₜ e2' : τ''
  ( `{ICG_ST Σ} `{LogRelG Σ}, Γ e1 log e1' : τ)
  ( `{ICG_ST Σ} `{LogRelG Σ},
      subst (ren (+1)) <$> (τ :: Γ) e2 log e2' : τ'')
  Γ App (Rec (TLam e2).[ren (+1)]) e1 ctx
      (TLam (App (Rec e2'.[ren (+1)]) e1'))
  : TForall τ''.
Proof.
  intros Ht1 Ht2 Ht1' Ht2' HLR1 HLR2.
  eapply ctx_refines_trans.
  eapply TLam_Hoisting_CR1; eauto.
  eapply ctx_refines_trans.
  eapply TLam_Hoisting_CR2; eauto.
  eapply TLam_Hoisting_CR3; eauto.
Qed.

Lemma TLam_Hoisting Γ e1 τ e2 τ'' :
  Γ ⊢ₜ e1 : τ subst (ren (+1)) <$> (τ :: Γ) ⊢ₜ e2 : τ''
  Γ App (Rec (TLam e2).[ren (+1)]) e1 ctx
      (TLam (App (Rec e2.[ren (+1)]) e1))
  : TForall τ''.
Proof.
  intros.
  eapply TLam_Hoisting_basic; eauto.
  - intros; apply binary_fundamental; auto.
  - intros; apply binary_fundamental; auto.
Qed.