RunST.IC_adequacy

From RunST Require Import IC reduction future.
From iris.algebra Require Import gmap excl auth agree gset coPset.
From iris.base_logic Require Import auth big_op soundness derived.
From iris.base_logic.lib Require Import wsat.
From iris.proofmode Require Import tactics.
Import uPred.

(* Global functor setup *)
Definition invΣ : gFunctors :=
  #[GFunctor (authRF (gmapURF positive (agreeRF (laterCF idCF))));
    GFunctor (constRF coPset_disjUR);
    GFunctor (constRF (gset_disjUR positive))].

Class invPreG (Σ : gFunctors) : Set := WsatPreG {
  inv_inPreG :> inG Σ (authR (gmapUR positive (agreeR (laterC (iPreProp Σ)))));
  enabled_inPreG :> inG Σ coPset_disjR;
  disabled_inPreG :> inG Σ (gset_disjR positive);
}.

Instance subG_invΣ Σ : subG invΣ Σ invPreG Σ.
Proof.
  intros [?%subG_inG [?%subG_inG ?%subG_inG]%subG_inv]%subG_inv; by constructor.
Qed.

Definition ICΣ `{ICState' Λstate} : gFunctors :=
  #[invΣ; GFunctor (constRF (authR IC_stateR))].

Class ICPreG' {Λstate : Type} `{ICState' Λstate} (Σ : gFunctors) :
  Type := Build_ICPreG {
  inv_inG :> invPreG Σ;
  state_inG :> inG Σ (authR IC_stateR);
}.

Instance subG_ICΣ `{ICState' Λstate} Σ : subG ICΣ Σ ICPreG' Σ.
Proof. intros [??%subG_inG]%subG_inv. constructor; apply _. Qed.

(* Allocation *)
Lemma wsat_alloc `{invPreG Σ} : True ==★ _ : invG Σ, wsat ownE .
Proof.
  iIntros.
  iMod (own_alloc ( ( : gmap _ _))) as (γI) "HI"; first done.
  iMod (own_alloc (CoPset )) as (γE) "HE"; first done.
  iMod (own_alloc (GSet )) as (γD) "HD"; first done.
  iModIntro; iExists (WsatG _ _ _ _ γI γE γD).
  rewrite /wsat /ownE; iFrame.
  iExists . rewrite fmap_empty big_sepM_empty. by iFrame.
Qed.

Lemma IC_alloc `{Hicpreg : ICPreG' Λstate Σ} :
    True ==★ (_ : ICG' Λstate Σ), wsat ownE .
Proof.
  iIntros.
  iMod wsat_alloc as (?) "[Hws HE]". iModIntro.
  iExists (Build_ICG _ _ _ _ (AuthG _ _ _ IC_stateR_discrete)).
  by iFrame.
Qed.

Section adequacy.
Context `{ICState Λ, ICG Λ Σ}.
Implicit Types e : expr Λ.
Implicit Types P Q : iProp Σ.
Implicit Types Φ : val Λ nat iProp Σ.
Implicit Types Ψ : val Λ nat Prop.

Notation world γ σ := (wsat ownE ownP_full γ σ)%I.

Class Reds e σ :=
{
  nstps : nat;
  end_val : val Λ;
  end_state : state Λ;
  reds : nsteps pstep nstps (e, σ) (of_val end_val, end_state)
}.

Lemma ic_adequacy_basic γ e σ Ψ :
  ownP_full γ σ IC e {{v; n, Ψ v n}}[γ] ={}=★
         (rd : Reds e σ),
          Nat.iter nstps (λ P, P) ( ( Ψ end_val nstps)).
Proof.
  iIntros "[HPh Hic]".
  iApply (@fupd_plain_forall_comm _ _ _ ).
  iIntros ([n v σ' Hrd]); simpl.
  rewrite ic_eq /ic_def.
  iSpecialize ("Hic" $! _ _ _ _ with "[HPh]"); iFrame; eauto.
  iApply future_plain. iMod "Hic" as "[Hr _]". by iModIntro.
Qed.

Lemma ic_adequacy_basic' γ e σ Ψ :
  world γ σ IC e {{v; n, Ψ v n}}[γ] ==★
         (rd : Reds e σ),
          Nat.iter (S (S nstps)) (λ P, P) ( Ψ end_val nstps).
Proof.
  iIntros "[(Hsat & How & HPh) Hic]".
  iPoseProof (ic_adequacy_basic with "[$HPh $Hic]") as "Hbase"; auto.
  rewrite fupd_eq /fupd_def.
  iMod ("Hbase" with "[$Hsat $How]") as "H".
  rewrite ?except_0_sep.
  iDestruct "H" as "(_&_&H)".
  iModIntro. iIntros (rd).
  rewrite Nat_iter_S Nat_iter_S_r.
  rewrite /uPred_except_0.
  iDestruct "H" as "[H|H]"; first by iNext; iExFalso.
  iNext. iSpecialize ("H" $! rd).
  iApply later_n_mono; last eauto.
  iIntros "H". iDestruct "H" as "[H|H]"; first by iNext; iExFalso.
  by iNext.
Qed.

Lemma ic_adequacy_basic'' γ e σ Ψ :
  world γ σ IC e {{v; n, Ψ v n}}[γ]
         (rd : Reds e σ),
          Nat.iter (S (S nstps)) (λ P, |==> P) ( Ψ end_val nstps).
Proof.
  iIntros "[Hw Hic]".
  iPoseProof (ic_adequacy_basic' with "[$Hw $Hic]") as "H"; auto.
  iIntros (rd). simpl.
  iMod "H". iModIntro. iSpecialize ("H" $! rd).
  iNext; iModIntro; iNext.
  generalize nstps at 1 3. intros n; induction n; simpl; auto.
Qed.

End adequacy.

Instance bupd_iter_mono {M : ucmraT} n :
    Proper ((⊢) ==> (@uPred_entails M)) (Nat.iter n (λ P, |==> P)%I).
Proof. intros P Q HP. induction n; simpl; do 2?f_equiv; auto. Qed.

Notation ICPreG Λ Σ := (@ICPreG' (state Λ) _ Σ).

Lemma ic_adequacy `{HICst : ICState Λ, Hipg : ICPreG Λ Σ} e σ Ψ :
  (forall `{@ICG' (state Λ) HICst Σ},
      True ==★ γ, ownP_full γ σ auth_own γ (IC_to_stateR σ))
  ( `{@ICG' (state Λ) HICst Σ} γ,
      auth_own γ (IC_to_stateR σ) IC e {{v; n, Ψ v n}}[γ])
  ( v n σ', nsteps pstep n (e, σ) (of_val v, σ') (Ψ v n)).
Proof.
  intros Halloc Hic v n σ' Hrd.
  eapply (@soundness (iResUR Σ) _ (S (S nstps))).
  iIntros "/=".
  iMod (IC_alloc) as (HICG) "[Hwsat How]".
  iMod Halloc as (γ) "[HPh Ho]".
  iPoseProof (ic_adequacy_basic'' with "[-]") as "H"; iFrame.
  { by iApply Hic. }
  iSpecialize ("H" $! (Build_Reds _ _ _ _ _ Hrd)).
  rewrite -?(Nat_iter_S _ (λ P : uPred (iResUR Σ), |==> P)%I).
  iApply (bupd_iter_mono with "H").
  trivial.
Qed.