RunST.soundness

From RunST Require Import IC lang rules typing regions logrel fundamental
     reduction contextual_refinement IC_adequacy saved_pred ST_Lang_reduction
     logrel_shared logrel_ctx_closure.
From iris.algebra Require Import gmap excl dec_agree.
From iris.base_logic Require Import big_op auth.
From iris.proofmode Require Import tactics.

Notation ICG_ST Σ := (@ICG' state ICState_state Σ).

Global Instance ICG_heapG `{H: ICG_ST Σ} : heapG Σ :=
  {| heap_invG := IC_invG; heap_inG := ICstate_inG |}.

Lemma basic_soundness Σ `{ICPreG ST_lang Σ} `{PreLogRelG Σ}
      e σ1 e' σ1' τ v1 σ2 :
  ( `{ICG_ST Σ} `{LogRelG Σ}, [] e log e' : τ)
  rtc pstep (e, σ1) (of_val v1, σ2)
  ( σ2' v2', rtc pstep (e', σ1') (of_val v2', σ2')).
Proof.
  intros Hlog Hsteps.
  apply rtc_nsteps in Hsteps. destruct Hsteps as [n Hsteps].
  cut ( v n σ', nsteps pstep n (e, σ1) (of_val v, σ')
          ((λ _ _, σ2' v2',
               rtc pstep (e', σ1') (of_val v2', σ2')) v n)).
  { eauto. } simpl. clear Hsteps n v1 σ2.
  intros v b σ' Hsteps.
  eapply ic_adequacy; eauto.
  - intros. simpl.
    iIntros. iMod (allocate_full_state σ1) as (γh) "[H1 H2]".
    iModIntro. iExists γh; rewrite /ownP_full; iFrame.
  - iIntros (HICG γh) "Hown".
    iApply fupd_ic.
    iMod (@reg_inv_alloc _ ICG_heapG _) as (HLogRelG) "#Hinv".
    specialize (Hlog HICG HLogRelG).
    iPoseProof (Hlog [] [] with "[]") as "He".
    { iFrame "#". iApply interp_env_nil. }
    iMod (allocate_full_state σ1') as (γh') "[Hfl' Hown']".
    iSpecialize ("He" $! γh γh' _ with "[$Hfl']").
    simpl in *. rewrite !empty_env_subst.
    iModIntro.
    iApply ic_mono; last iApply "He".
    iIntros (v' n'); simpl. iDestruct 1 as (σ2' v2') "(Hrd & _ & _)".
    iDestruct "Hrd" as %Hrd%rtc_nsteps. destruct Hrd as [n Hrd].
    iPureIntro. eauto using nsteps_rtc, nsteps_det.
Qed.

Lemma binary_soundness Γ e e' τ :
  typed Γ e τ typed Γ e' τ
  ( Σ `{ICG_ST Σ} `{LogRelG Σ}, Γ e log e' : τ)
  Γ e ctx e' : τ.
Proof.
  set (Σ := #[ICΣ; PreLogRelΣ]).
  intros He He' Hlog. repeat split; auto.
  intros K σ1 σ2 v σ1' Hctx. eapply (basic_soundness Σ)=> ? ?.
  eapply (bin_log_related_under_typed_ctx _ _ _ _ []);
    eauto using typed_n_closed.
Qed.