RunST.rec_unfolding

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

Hint Resolve to_of_val.

Lemma rec_unfolding_LR `{ICG_ST Σ} `{LogRelG Σ} Γ e e' τ τ' :
  TArrow τ τ' :: τ :: Γ ⊢ₜ e : τ'
  TArrow τ τ' :: τ :: Γ ⊢ₜ e' : τ'
  TArrow τ τ' :: τ :: Γ e log e' : τ'
  Γ Rec e log Rec (e'.[Rec e'.[upn 2 (ren (+1))]/]).[ren (+1)]
  : TArrow τ τ'.
Proof.
  iIntros (Ht Ht' HLR Δ vvs ) "[#Hinv #Hvvs]".
  iIntros (γh γh' σ1') "Hσ1' /=".
  iApply ic_value; [cbn; rewrite ?to_of_val; trivial|].
  iExists _, (RecV _); iSplit; iFrame.
  { iPureIntro. econstructor. }
  clear γh γh' σ1'.
  iAlways. iIntros (vv) "#Hvv". asimpl.
  iApply logrel_pure_det_head_step.
  { intros σ; econstructor; eauto. }
  { intros ? ? ? Hrd. inversion Hrd; subst; auto. }
  { intros σ. econstructor; eauto. }
  iNext. asimpl.
  iDestruct (interp_env_length with "[]") as %Hvvsl; first iFrame "#".
  replace (e.[Rec e.[upn 2 (env_subst (vvs.*1))]
                      .: of_val (vv.1) .: env_subst (vvs.*1)])
  with
  (e.[env_subst (RecV e.[upn 2 (env_subst (vvs.*1))] :: vv.1:: (vvs.*1))]);
    last by erewrite <- typed_subst_head_simpl_2; eauto;
      last by rewrite /= Hvvsl map_length.
  replace (e'.[Rec e'.[upn 2 (env_subst (vvs.*2))]
                      .: of_val (vv.2) .: env_subst (vvs.*2)])
  with
  (e'.[env_subst (RecV e'.[upn 2 (env_subst (vvs.*2))] :: vv.2:: (vvs.*2))]);
    last by erewrite <- typed_subst_head_simpl_2; eauto;
      last by rewrite /= Hvvsl map_length.
  iIntros (γh γh' σ1') "Hσ1' /=".
  unshelve iPoseProof (bin_log_related_rec _ _ _ _ _ _ _ HLR _ vvs with "[]")
    as "Hrl"; try iFrame "#".
  { eapply (typed_n_closed (_ :: _ :: _)); eauto. }
  { eapply (typed_n_closed (_ :: _ :: _)); eauto. }
  iSpecialize ("Hrl" $! _ _ _ with "[$Hσ1']"); simpl.
  change (Rec e.[upn 2 (env_subst (vvs.*1))]) with
    (of_val (RecV e.[upn 2 (env_subst (vvs.*1))])).
  iMod "Hrl".
  iDestruct "Hrl" as (σ2' v') "(Hrd1 & Hσ2' & #Hrl)".
  iDestruct "Hrd1" as %Hrd1.
  apply rtc_nsteps in Hrd1. destruct Hrd1 as [n Hrd1].
  eapply (nsteps_val det_lang _ (RecV _)) in Hrd1;
    destruct Hrd1 as (?&?%of_val_inj&?); subst.
  iRevert (γh γh' σ2') "Hσ2'".
  iPoseProof
    (HLR _ ((RecV e.[upn 2 (env_subst (vvs.*1))],
             RecV e'.[upn 2 (env_subst (vvs.*2))]) :: vv :: vvs)
     with "[]") as "He"; iFrame "#"; auto.
  rewrite ?interp_env_cons; iFrame "#"; auto.
Qed.

Lemma rec_unfolding_basic Γ e e' τ τ' :
  TArrow τ τ' :: τ :: Γ ⊢ₜ e : τ'
  TArrow τ τ' :: τ :: Γ ⊢ₜ e' : τ'
  ( `{ICG_ST Σ} `{LogRelG Σ}, TArrow τ τ' :: τ :: Γ e log e' : τ')
  Γ Rec e ctx Rec (e'.[Rec e'.[upn 2 (ren (+1))]/]).[ren (+1)]
  : TArrow τ τ'.
Proof.
  intros Ht1 Ht2 HLR.
  apply binary_soundness.
  - by econstructor.
  - econstructor. apply (context_weakening [_]).
    eapply typed_subst; eauto.
    econstructor. by eapply (context_gen_weakening [_] [_; _]).
  - intros; by apply rec_unfolding_LR.
Qed.

Lemma rec_unfolding Γ e τ τ' :
  TArrow τ τ' :: τ :: Γ ⊢ₜ e : τ'
  Γ Rec e ctx Rec (e.[Rec e.[upn 2 (ren (+1))]/]).[ren (+1)] : TArrow τ τ'.
Proof.
  intros.
  eapply rec_unfolding_basic; eauto.
  intros; apply binary_fundamental; auto.
Qed.