RunST.IC_lifting

From iris.program_logic Require Export ectx_language lifting.
From RunST Require Export IC future reduction.
From iris.base_logic Require Export big_op.
From iris.proofmode Require Import tactics.

Section lifting.
Context `{ICState Λ, ICG Λ Σ}.
Implicit Types v : val Λ.
Implicit Types e : expr Λ.
Implicit Types σ : state Λ.
Implicit Types P Q : iProp Σ.
Implicit Types Φ : val Λ nat iProp Σ.

Lemma ic_lift_step γ E1 E2 Φ e1 :
  ( σ1, ownP_full γ σ1 ={E1, E2}=★
     ( nf_reducible e1 σ1)
      e2 σ2, prim_step e1 σ1 e2 σ2 [] -★
       (|={E2, E1}=> ownP_full γ σ2 IC e2 @ E1 {{w; n, Φ w (S n) }}[γ]))
     IC e1 @ E1 {{ Φ }}[γ].
Proof.
  iIntros "H". rewrite ic_eq /ic_def.
  iIntros (σ1 σ2 v n) "[Hr Ho]"; iDestruct "Hr" as %Hr.
  iMod ("H" $! σ1 with "[Ho]") as "(Hr1 & Hcont)"; first done.
  iDestruct "Hr1" as %Hr1.
  inversion Hr as [|k Hs2 [e' σ'] Hs4 Hs5 Hs6]; subst; simpl in *.
  { apply nf_red_red, reducible_not_val in Hr1; by rewrite to_of_val in Hr1. }
  iDestruct ("Hcont" $! e' σ' with "[]") as "Hrest"; first done.
  iMod ("Hrest") as "[Ho Hrest]".
  iModIntro; rewrite future_S; iModIntro; iNext.
  iApply "Hrest"; by iFrame.
Qed.

Lemma ic_lift_pure_step γ E Φ e1 :
  ( σ1, nf_reducible e1 σ1)
  ( σ1 e2 σ2, prim_step e1 σ1 e2 σ2 [] σ1 = σ2)
  ( σ1 e2 σ2,
         prim_step e1 σ1 e2 σ2 []
        IC e2 @ E {{w; n, Φ w (S n) }}[γ]) IC e1 @ E {{ Φ }}[γ].
Proof.
  iIntros (Hnfr Hpr) "H". rewrite ic_eq /ic_def.
  iIntros (σ1 σ2 v n) "[Hr Ho]"; iDestruct "Hr" as %Hr.
  specialize (Hnfr σ1); specialize (Hpr σ1).
  inversion Hr as [|k Hs2 [e' σ'] Hs4 Hs5 Hs6]; subst; simpl in *.
  { apply nf_red_red, reducible_not_val in Hnfr; by rewrite to_of_val in Hnfr. }
  rewrite future_S; iModIntro; iNext.
  iSpecialize ("H" $! σ1 e' σ' with "[]"); first done.
  rewrite (Hpr _ _ Hs5).
  iSpecialize ("H" $! σ' σ2 v k with "[Ho]"); by iFrame.
Qed.

Derived lifting lemmas.
Lemma ic_lift_atomic_step {γ E1 E2 Φ} e1 :
  atomic e1
  ( σ1, ownP_full γ σ1 ={E1, E2}=★ ( nf_reducible e1 σ1)
      v2 σ2, prim_step e1 σ1 (of_val v2) σ2 [] -★
     (|={E2, E1}=> ownP_full γ σ2 Φ v2 1))
     IC e1 @ E1 {{ Φ }}[γ].
Proof.
  iIntros (Hatomic) "H". iApply (ic_lift_step _ E1 _ _ e1).
  iIntros (σ1) "Ho"; iMod ("H" $! _ with "Ho") as "[Hr H]"; iModIntro; iFrame.
  iIntros (e2 σ2 Hstp).
  destruct (Hatomic _ _ _ _ Hstp) as [v2 Hv2].
  iMod ("H" $! _ _ with "[]") as "[H1 H2]";
    first by erewrite of_to_val; eauto; iFrame.
  iModIntro; iFrame; iNext.
  rewrite -(of_to_val e2 v2); trivial.
  by iApply ic_value'.
Qed.

Lemma ic_lift_atomic_det_step {γ E1 E2 Φ e1} :
  atomic e1
  ( σ1, ownP_full γ σ1 ={E1, E2}=★ ( nf_reducible e1 σ1)
      v2 σ2,
       ( e2' σ2', prim_step e1 σ1 e2' σ2' []
                   σ2 = σ2' to_val e2' = Some v2)
       (|={E2, E1}=> ownP_full γ σ2 Φ v2 1))
     IC e1 @ E1 {{ Φ }}[γ].
Proof.
  iIntros (Hatomic) "H". iApply (ic_lift_atomic_step); first done.
  iIntros (σ1) "Ho"; iMod ("H" $! _ with "Ho") as "[Hr H]"; iModIntro; iFrame.
  iDestruct "H" as (v2 σ2) "[Hdet H2]"; iDestruct "Hdet" as %Hdet.
  iIntros (v2' σ2' Hred).
  edestruct Hdet as (->&->%of_to_val%(inj of_val)); done.
Qed.

Lemma ic_lift_pure_det_step `{Inhabited (state Λ)} {γ E Φ} e1 e2 :
  ( σ1, nf_reducible e1 σ1)
  ( σ1 e2' σ2, prim_step e1 σ1 e2' σ2 [] σ1 = σ2 e2 = e2')→
   IC e2 @ E {{ v; n, Φ v (S n) }}[γ] IC e1 @ E {{ Φ }}[γ].
Proof.
  iIntros (? Hpuredet) "?". iApply (ic_lift_pure_step _ E); try done.
  by intros; eapply Hpuredet. iNext. by iIntros (e' σ σ' (->&->)%Hpuredet).
Qed.
End lifting.

Some derived lemmas for ectx-based languages
Section ic_ectx_lifting.
Context {expr val ectx state} {Λ : EctxLanguage expr val ectx state}.
Context `{ICState (ectx_lang expr), ICG (ectx_lang expr) Σ} `{Inhabited state}.
Implicit Types P : iProp Σ.
Implicit Types Φ : val nat iProp Σ.
Implicit Types v : val.
Implicit Types e : expr.
Hint Resolve head_prim_reducible head_reducible_prim_step
     nf_head_red_head_red nf_head_reducible_nf_reducible.

Lemma ic_ectx_bind {γ E e} K Φ :
  IC e @ E {{ v ; n, IC fill K (of_val v) @ E {{w; m, Φ w (n + m)}}[γ] }}[γ]
      IC fill K e @ E {{ Φ }}[γ].
Proof. apply: ic_bind. Qed.

Lemma ic_lift_head_step γ E1 E2 Φ e1 :
( σ1, ownP_full γ σ1 ={E1, E2}=★ ( nf_head_reducible e1 σ1)
      e2 σ2, head_step e1 σ1 e2 σ2 [] -★
       (|={E2, E1}=> ownP_full γ σ2 IC e2 @ E1 {{w; n, Φ w (S n) }}[γ]))
   IC e1 @ E1 {{ Φ }}[γ].
Proof.
  iIntros "H". iApply (ic_lift_step _ E1 E2); try done.
  iIntros (σ1) "Ho"; iMod ("H" $! _ with "Ho") as "[Hr H]"; iModIntro.
  iDestruct "Hr" as %Hr.
  iSplit; first iPureIntro; auto using head_reducible_prim_step.
  iIntros (e2 σ2) "%".
  iSpecialize ("H" $! e2 σ2 with "[]");
    auto using head_reducible_prim_step, nf_head_red_head_red.
Qed.

Lemma ic_lift_pure_head_step γ E Φ e1 :
  ( σ1, nf_head_reducible e1 σ1)
  ( σ1 e2 σ2, head_step e1 σ1 e2 σ2 [] σ1 = σ2)
  ( σ1 e2 σ2,
         head_step e1 σ1 e2 σ2 []
        IC e2 @ E {{w; n, Φ w (S n) }}[γ]) IC e1 @ E {{ Φ }}[γ].
Proof.
  iIntros (??) "H".
  iApply ic_lift_pure_step; eauto.
  iNext. iIntros (????). iApply "H"; eauto.
Qed.

Lemma ic_lift_atomic_head_step {γ E1 E2 Φ} e1 :
  atomic e1
  ( σ1, ownP_full γ σ1 ={E1, E2}=★ ( nf_head_reducible e1 σ1)
      v2 σ2, head_step e1 σ1 (of_val v2) σ2 [] -★
     (|={E2, E1}=> ownP_full γ σ2 Φ v2 1))
     IC e1 @ E1 {{ Φ }}[γ].
Proof.
  iIntros (?) "H". iApply ic_lift_atomic_step; first done.
  iIntros (σ1) "Ho"; iMod ("H" $! _ with "Ho") as "[Hr H]"; iModIntro.
  iDestruct "Hr" as %Hr.
  iSplit; first iPureIntro; auto using head_reducible_prim_step.
  iIntros (e2 σ2) "%".
  iSpecialize ("H" $! e2 σ2 with "[]");
    auto using head_reducible_prim_step, nf_head_red_head_red.
Qed.

Lemma ic_lift_pure_det_head_step {γ E Φ} e1 e2 :
  ( σ1, nf_head_reducible e1 σ1)
  ( σ1 e2' σ2, head_step e1 σ1 e2' σ2 [] σ1 = σ2 e2 = e2')→
   IC e2 @ E {{ v; n, Φ v (S n) }}[γ] IC e1 @ E {{ Φ }}[γ].
Proof. intros. apply ic_lift_pure_det_step; eauto. Qed.

End ic_ectx_lifting.