LogrelCC.program_logic.CC_ectx_lifting
Some derived lemmas for ectx-based languages with continuations
From LogrelCC.program_logic Require Export weakestpre lifting.
From LogrelCC.program_logic Require Export CC_ectx_language.
From iris.proofmode Require Import tactics.
Set Default Proof Using "Type".
Section wp.
Context {expr val ectx state} {Λ : CCEctxLanguage expr val ectx state}.
Context `{irisG (CC_ectx_lang expr) Σ} {Hinh : Inhabited state}.
Implicit Types P : iProp Σ.
Implicit Types Φ : val → iProp Σ.
Implicit Types v : val.
Implicit Types e : expr.
Hint Resolve head_prim_reducible
head_nonthrow_prim_reducible
head_normal_prim_reducible
head_capture_prim_reducible
head_throw_prim_reducible
head_normal_reducible_prim_step
head_capture_reducible_prim_step
head_nonthrow_reducible_prim_step
head_throw_reducible_prim_step.
Lemma wp_lift_nonthrow_head_step {s E Φ} K e1 :
to_val e1 = None →
(∀ σ1, state_interp σ1 ={E,∅}=∗
⌜head_nonthrow_reducible K e1 σ1⌝ ∗
▷ ∀ rm e2 σ2 efs, ⌜head_step K e1 σ1 e2 σ2 efs rm⌝ ={∅,E}=∗
state_interp σ2 ∗ WP fill K e2 @ s; E {{ Φ }} ∗
[∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ _, True }})
⊢ WP fill K e1 @ s; E {{ Φ }}.
Proof.
iIntros (?) "H". iApply wp_lift_step=>//; eauto using CC_fill_not_val.
iIntros (σ1) "Hσ".
iMod ("H" $! σ1 with "Hσ") as "[% H]"; iModIntro.
iSplit; first by destruct s; eauto. iNext. iIntros (e2 σ2 efs) "H1".
iDestruct "H1" as %Hps.
eapply head_nonthrow_reducible_prim_step in Hps; eauto.
destruct Hps as (?&e2'&?&?&Hps); subst.
iApply "H". by eauto.
Qed.
Lemma wp_lift_throw_head_step {s E Φ} K e1 :
to_val e1 = None →
(∀ σ1, state_interp σ1 ={E,∅}=∗
⌜head_throw_reducible K e1 σ1⌝ ∗
▷ ∀ e2 σ2 efs, ⌜head_step K e1 σ1 e2 σ2 efs ThrowMode⌝ ={∅,E}=∗
state_interp σ2 ∗ WP e2 @ s; E {{ Φ }} ∗
[∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ _, True }})
⊢ WP fill K e1 @ s; E {{ Φ }}.
Proof.
iIntros (?) "H". iApply wp_lift_step=>//; eauto using CC_fill_not_val.
iIntros (σ1) "Hσ".
iMod ("H" $! σ1 with "Hσ") as "[% H]"; iModIntro.
iSplit; first by destruct s; eauto. iNext. iIntros (e2 σ2 efs) "H1".
iDestruct "H1" as %Hps.
eapply head_throw_reducible_prim_step in Hps; eauto.
iApply "H". by eauto.
Qed.
Lemma wp_lift_nonthrow_pure_head_step {s E E' Φ} K e1 :
(∀ σ1, head_nonthrow_reducible K e1 σ1) →
(∀ rm σ1 e2 σ2 efs, head_step K e1 σ1 e2 σ2 efs rm → σ1 = σ2) →
(|={E,E'}▷=> ∀ rm e2 efs σ, ⌜head_step K e1 σ e2 σ efs rm⌝ →
WP fill K e2 @ s; E {{ Φ }} ∗ [∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ _, True }})
⊢ WP fill K e1 @ s; E {{ Φ }}.
Proof using Hinh.
iIntros (Hnthsp Hpure) "H".
iApply wp_lift_pure_step;
first destruct s; eauto.
- intros σ; specialize (Hnthsp σ).
eauto using reducible_not_val,
head_prim_reducible, head_nonthrow_prim_reducible.
- intros ? ? ? ? Hps.
eapply head_nonthrow_reducible_prim_step in Hps; eauto.
destruct Hps as (?&e2'&?&?&Hps); subst; eauto.
- iApply (step_fupd_wand with "H"); iIntros "H".
iIntros (??? Hps).
eapply head_nonthrow_reducible_prim_step in Hps; eauto.
destruct Hps as (?&e2'&?&?&Hps); subst; eauto.
iApply "H"; eauto.
Qed.
Lemma wp_lift_throw_pure_head_step {s E E' Φ} K e1 :
(∀ σ1, head_throw_reducible K e1 σ1) →
(∀ rm σ1 e2 σ2 efs, head_step K e1 σ1 e2 σ2 efs rm → σ1 = σ2) →
(|={E,E'}▷=> ∀ rm e2 efs σ, ⌜head_step K e1 σ e2 σ efs rm⌝ →
WP e2 @ s; E {{ Φ }} ∗ [∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ _, True }})
⊢ WP fill K e1 @s; E {{ Φ }}.
Proof using Hinh.
iIntros (Hthsp Hpure) "H". iApply wp_lift_pure_step; first destruct s; eauto.
- intros σ; specialize (Hthsp σ).
eauto using reducible_not_val,
head_prim_reducible, head_throw_prim_reducible.
- iApply (step_fupd_wand with "H"); iIntros "H".
iIntros (??? Hps).
eapply head_throw_reducible_prim_step in Hps; eauto.
iApply "H"; eauto.
Qed.
(* Atomic steps don't need any mask-changing business here, one can
use the generic lemmas here. These we only show for non-throw
operational steps as atomic throw steps don't happen in practice! *)
Lemma wp_lift_nonthrow_atomic_head_step {s E Φ} K e1 :
to_val e1 = None →
(∀ σ1, state_interp σ1 ={E}=∗
⌜head_nonthrow_reducible K e1 σ1⌝ ∗
▷ ∀ rm e2 σ2 efs, ⌜head_step K e1 σ1 e2 σ2 efs rm⌝ ={E}=∗
state_interp σ2 ∗
from_option (λ v, WP fill K (of_val v) @ s; E {{Φ}}) False (to_val e2) ∗
[∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ _, True }})
⊢ WP fill K e1 @ s; E {{ Φ }}.
Proof.
iIntros (?) "H".
iApply (wp_lift_nonthrow_head_step)=>//; eauto using CC_fill_not_val.
iIntros (σ1) "Hσ1".
iMod ("H" $! σ1 with "Hσ1") as "[$ H]".
iMod (fupd_intro_mask' E ∅) as "Hclose"; first set_solver.
iModIntro; iNext; iIntros (rm e2 σ2 efs) "%". iMod "Hclose" as "_".
iMod ("H" $! rm e2 σ2 efs with "[#]") as "($ & HΦ & $)"; auto.
iModIntro.
destruct (to_val e2) eqn:Heq; last by iExFalso.
apply of_to_val in Heq; subst; auto.
Qed.
Lemma wp_lift_nonthrow_atomic_head_step_no_fork {s E Φ} K e1 :
to_val e1 = None →
(∀ σ1, state_interp σ1 ={E}=∗
⌜head_nonthrow_reducible K e1 σ1⌝ ∗
▷ ∀ rm e2 σ2 efs, ⌜head_step K e1 σ1 e2 σ2 efs rm⌝ ={E}=∗
⌜efs = []⌝ ∗ state_interp σ2 ∗
from_option (λ v, WP fill K (of_val v) @ s; E {{Φ}}) False (to_val e2))
⊢ WP fill K e1 @ s; E {{ Φ }}.
Proof.
iIntros (?) "H". iApply wp_lift_nonthrow_atomic_head_step; eauto.
iIntros (σ1) "Hσ1". iMod ("H" $! σ1 with "Hσ1") as "[$ H]"; iModIntro.
iNext; iIntros (rm v2 σ2 efs) "%".
iMod ("H" $! rm v2 σ2 efs with "[# //]") as "(% & $ & $)"; subst; auto.
Qed.
Lemma wp_lift_nonthrow_pure_det_head_step {s E E' Φ} K e1 e2 efs :
(∀ σ1, head_nonthrow_reducible K e1 σ1) →
(∀ σ1 e2' σ2 efs' rm,
head_step K e1 σ1 e2' σ2 efs' rm → σ1 = σ2 ∧ e2 = e2' ∧ efs = efs') →
(|={E,E'}▷=> WP fill K e2 @ s; E {{ Φ }} ∗ [∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ _, True }})
⊢ WP fill K e1 @ s; E {{ Φ }}.
Proof using Hinh.
intros Hnthsp Hdet.
apply wp_lift_pure_det_step; first destruct s; eauto; simpl.
- intros σ; specialize (Hnthsp σ).
eauto using reducible_not_val,
head_prim_reducible, head_nonthrow_prim_reducible.
- intros ???? Hps.
apply head_nonthrow_reducible_prim_step in Hps; auto.
destruct Hps as (?&?&?&?&Hhs); subst.
apply Hdet in Hhs. intuition congruence.
Qed.
Lemma wp_lift_throw_pure_det_head_step {s E E' Φ} K e1 e2 efs :
(∀ σ1, head_throw_reducible K e1 σ1) →
(∀ σ1 e2' σ2 efs' rm,
head_step K e1 σ1 e2' σ2 efs' rm → σ1 = σ2 ∧ e2 = e2' ∧ efs = efs') →
(|={E,E'}▷=> WP e2 @ s; E {{ Φ }} ∗ [∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ _, True }})
⊢ WP fill K e1 @ s; E {{ Φ }}.
Proof using Hinh.
intros Hthsp Hdet. apply wp_lift_pure_det_step; destruct s; eauto.
intros σ; specialize (Hthsp σ).
eauto using reducible_not_val,
head_prim_reducible, head_throw_prim_reducible.
Qed.
Lemma wp_lift_nonthrow_pure_det_head_step_no_fork {s E E' Φ} K e1 e2 :
to_val e1 = None →
(∀ σ1, head_nonthrow_reducible K e1 σ1) →
(∀ rm σ1 e2' σ2 efs',
head_step K e1 σ1 e2' σ2 efs' rm → σ1 = σ2 ∧ e2 = e2' ∧ [] = efs') →
(|={E,E'}▷=> WP fill K e2 @ s; E {{ Φ }}) ⊢ WP fill K e1 @ s; E {{ Φ }}.
Proof using Hinh.
intros ? Hnthsp Hdet.
rewrite -(wp_lift_pure_det_step (fill K e1) (fill K e2) []) /= ?right_id;
eauto.
- destruct s; eauto.
intros σ; specialize (Hnthsp σ).
eauto using reducible_not_val,
head_prim_reducible, head_nonthrow_prim_reducible.
- intros ???? Hps.
apply head_nonthrow_reducible_prim_step in Hps; auto.
destruct Hps as (?&?&?&?&Hhs); subst.
apply Hdet in Hhs. intuition congruence.
Qed.
Lemma wp_lift_throw_pure_det_head_step_no_fork {s E E' Φ} K e1 e2 :
to_val e1 = None →
(∀ σ1, head_throw_reducible K e1 σ1) →
(∀ rm σ1 e2' σ2 efs',
head_step K e1 σ1 e2' σ2 efs' rm → σ1 = σ2 ∧ e2 = e2' ∧ [] = efs') →
(|={E,E'}▷=> WP e2 @ s; E {{ Φ }}) ⊢ WP fill K e1 @ s; E {{ Φ }}.
Proof using Hinh.
intros ? Hthsp Hdet.
rewrite -(wp_lift_pure_det_step (fill K e1) e2 []) /= ?right_id;
eauto.
destruct s; eauto.
intros σ; specialize (Hthsp σ).
eauto using reducible_not_val,
head_prim_reducible, head_throw_prim_reducible.
Qed.
Lemma wp_lift_nonthrow_pure_det_head_step_no_fork' {s E Φ} K e1 e2 :
to_val e1 = None →
(∀ σ1, head_nonthrow_reducible K e1 σ1) →
(∀ rm σ1 e2' σ2 efs',
head_step K e1 σ1 e2' σ2 efs' rm → σ1 = σ2 ∧ e2 = e2' ∧ [] = efs') →
(▷ WP fill K e2 @ s; E {{ Φ }}) ⊢ WP fill K e1 @ s; E {{ Φ }}.
Proof using Hinh.
intros.
rewrite -[(WP (fill K e1) @ _; _ {{ _ }})%I]
wp_lift_nonthrow_pure_det_head_step_no_fork //.
rewrite -step_fupd_intro //.
Qed.
Lemma wp_lift_throw_pure_det_head_step_no_fork' {s E Φ} K e1 e2 :
to_val e1 = None →
(∀ σ1, head_throw_reducible K e1 σ1) →
(∀ rm σ1 e2' σ2 efs',
head_step K e1 σ1 e2' σ2 efs' rm → σ1 = σ2 ∧ e2 = e2' ∧ [] = efs') →
(▷ WP e2 @ s; E {{ Φ }}) ⊢ WP fill K e1 @ s; E {{ Φ }}.
Proof using Hinh.
intros.
rewrite -[(WP (fill K e1) @ _; _ {{ _ }})%I]
wp_lift_throw_pure_det_head_step_no_fork //.
rewrite -step_fupd_intro //.
Qed.
Lemma wp_atomic_under_ectx E1 E2 K e Φ :
Atomic StronglyAtomic e → sub_values e → is_normal e →
(|={E1,E2}=> WP e @ E2 {{ v, |={E2,E1}=> WP fill K (of_val v) @ E1 {{ Φ }} }})
⊢ WP fill K e @ E1 {{ Φ }}.
Proof.
iIntros (Hatomic Hsv Hin) "H". rewrite !wp_unfold /wp_pre /=.
destruct (to_val (fill K e)) as [v|] eqn:HKe.
- iMod "H".
destruct (to_val e) as [e'|] eqn:Heqe'; last
by eapply (CC_fill_not_val K) in Heqe'; rewrite Heqe' in HKe.
repeat iMod "H".
rewrite (of_to_val _ _ Heqe').
iApply wp_value_inv; simpl; eauto.
- iIntros (σ1) "Hσ". iMod "H".
destruct (to_val e) as [v|] eqn:He.
+ do 2 iMod "H". rewrite (of_to_val _ _ He).
rewrite !wp_unfold /wp_pre /= HKe.
by iApply ("H" with "Hσ").
+ iMod ("H" with "Hσ") as "[Hr H]".
iDestruct "Hr" as %Hr.
iModIntro; iSplit; first by iPureIntro;
auto using reducible_under_ectx.
iNext. iIntros (e2 σ2 efs2) "Hps". iDestruct "Hps" as %Hps.
destruct (reducible_prim_step _ _ _ _ _ _ Hr Hsv Hin Hps) as
(e2' & He2 & Hps'); simpl in *; subst.
iMod ("H" with "[]") as "[Hσ [H $]]"; eauto.
apply Hatomic in Hps'.
rewrite {1}wp_unfold /wp_pre /=.
destruct (to_val e2') as [z|] eqn:Heqz.
++ rewrite (of_to_val _ _ Heqz). by repeat iMod "H"; iFrame.
++ iMod ("H" with "[$]") as "[H _]".
iDestruct "H" as %(? & ? & ? & Heq%val_prim_stuck); simpl in *.
rewrite Heq in Hps'. by inversion Hps'.
Qed.
End wp.
From LogrelCC.program_logic Require Export CC_ectx_language.
From iris.proofmode Require Import tactics.
Set Default Proof Using "Type".
Section wp.
Context {expr val ectx state} {Λ : CCEctxLanguage expr val ectx state}.
Context `{irisG (CC_ectx_lang expr) Σ} {Hinh : Inhabited state}.
Implicit Types P : iProp Σ.
Implicit Types Φ : val → iProp Σ.
Implicit Types v : val.
Implicit Types e : expr.
Hint Resolve head_prim_reducible
head_nonthrow_prim_reducible
head_normal_prim_reducible
head_capture_prim_reducible
head_throw_prim_reducible
head_normal_reducible_prim_step
head_capture_reducible_prim_step
head_nonthrow_reducible_prim_step
head_throw_reducible_prim_step.
Lemma wp_lift_nonthrow_head_step {s E Φ} K e1 :
to_val e1 = None →
(∀ σ1, state_interp σ1 ={E,∅}=∗
⌜head_nonthrow_reducible K e1 σ1⌝ ∗
▷ ∀ rm e2 σ2 efs, ⌜head_step K e1 σ1 e2 σ2 efs rm⌝ ={∅,E}=∗
state_interp σ2 ∗ WP fill K e2 @ s; E {{ Φ }} ∗
[∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ _, True }})
⊢ WP fill K e1 @ s; E {{ Φ }}.
Proof.
iIntros (?) "H". iApply wp_lift_step=>//; eauto using CC_fill_not_val.
iIntros (σ1) "Hσ".
iMod ("H" $! σ1 with "Hσ") as "[% H]"; iModIntro.
iSplit; first by destruct s; eauto. iNext. iIntros (e2 σ2 efs) "H1".
iDestruct "H1" as %Hps.
eapply head_nonthrow_reducible_prim_step in Hps; eauto.
destruct Hps as (?&e2'&?&?&Hps); subst.
iApply "H". by eauto.
Qed.
Lemma wp_lift_throw_head_step {s E Φ} K e1 :
to_val e1 = None →
(∀ σ1, state_interp σ1 ={E,∅}=∗
⌜head_throw_reducible K e1 σ1⌝ ∗
▷ ∀ e2 σ2 efs, ⌜head_step K e1 σ1 e2 σ2 efs ThrowMode⌝ ={∅,E}=∗
state_interp σ2 ∗ WP e2 @ s; E {{ Φ }} ∗
[∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ _, True }})
⊢ WP fill K e1 @ s; E {{ Φ }}.
Proof.
iIntros (?) "H". iApply wp_lift_step=>//; eauto using CC_fill_not_val.
iIntros (σ1) "Hσ".
iMod ("H" $! σ1 with "Hσ") as "[% H]"; iModIntro.
iSplit; first by destruct s; eauto. iNext. iIntros (e2 σ2 efs) "H1".
iDestruct "H1" as %Hps.
eapply head_throw_reducible_prim_step in Hps; eauto.
iApply "H". by eauto.
Qed.
Lemma wp_lift_nonthrow_pure_head_step {s E E' Φ} K e1 :
(∀ σ1, head_nonthrow_reducible K e1 σ1) →
(∀ rm σ1 e2 σ2 efs, head_step K e1 σ1 e2 σ2 efs rm → σ1 = σ2) →
(|={E,E'}▷=> ∀ rm e2 efs σ, ⌜head_step K e1 σ e2 σ efs rm⌝ →
WP fill K e2 @ s; E {{ Φ }} ∗ [∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ _, True }})
⊢ WP fill K e1 @ s; E {{ Φ }}.
Proof using Hinh.
iIntros (Hnthsp Hpure) "H".
iApply wp_lift_pure_step;
first destruct s; eauto.
- intros σ; specialize (Hnthsp σ).
eauto using reducible_not_val,
head_prim_reducible, head_nonthrow_prim_reducible.
- intros ? ? ? ? Hps.
eapply head_nonthrow_reducible_prim_step in Hps; eauto.
destruct Hps as (?&e2'&?&?&Hps); subst; eauto.
- iApply (step_fupd_wand with "H"); iIntros "H".
iIntros (??? Hps).
eapply head_nonthrow_reducible_prim_step in Hps; eauto.
destruct Hps as (?&e2'&?&?&Hps); subst; eauto.
iApply "H"; eauto.
Qed.
Lemma wp_lift_throw_pure_head_step {s E E' Φ} K e1 :
(∀ σ1, head_throw_reducible K e1 σ1) →
(∀ rm σ1 e2 σ2 efs, head_step K e1 σ1 e2 σ2 efs rm → σ1 = σ2) →
(|={E,E'}▷=> ∀ rm e2 efs σ, ⌜head_step K e1 σ e2 σ efs rm⌝ →
WP e2 @ s; E {{ Φ }} ∗ [∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ _, True }})
⊢ WP fill K e1 @s; E {{ Φ }}.
Proof using Hinh.
iIntros (Hthsp Hpure) "H". iApply wp_lift_pure_step; first destruct s; eauto.
- intros σ; specialize (Hthsp σ).
eauto using reducible_not_val,
head_prim_reducible, head_throw_prim_reducible.
- iApply (step_fupd_wand with "H"); iIntros "H".
iIntros (??? Hps).
eapply head_throw_reducible_prim_step in Hps; eauto.
iApply "H"; eauto.
Qed.
(* Atomic steps don't need any mask-changing business here, one can
use the generic lemmas here. These we only show for non-throw
operational steps as atomic throw steps don't happen in practice! *)
Lemma wp_lift_nonthrow_atomic_head_step {s E Φ} K e1 :
to_val e1 = None →
(∀ σ1, state_interp σ1 ={E}=∗
⌜head_nonthrow_reducible K e1 σ1⌝ ∗
▷ ∀ rm e2 σ2 efs, ⌜head_step K e1 σ1 e2 σ2 efs rm⌝ ={E}=∗
state_interp σ2 ∗
from_option (λ v, WP fill K (of_val v) @ s; E {{Φ}}) False (to_val e2) ∗
[∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ _, True }})
⊢ WP fill K e1 @ s; E {{ Φ }}.
Proof.
iIntros (?) "H".
iApply (wp_lift_nonthrow_head_step)=>//; eauto using CC_fill_not_val.
iIntros (σ1) "Hσ1".
iMod ("H" $! σ1 with "Hσ1") as "[$ H]".
iMod (fupd_intro_mask' E ∅) as "Hclose"; first set_solver.
iModIntro; iNext; iIntros (rm e2 σ2 efs) "%". iMod "Hclose" as "_".
iMod ("H" $! rm e2 σ2 efs with "[#]") as "($ & HΦ & $)"; auto.
iModIntro.
destruct (to_val e2) eqn:Heq; last by iExFalso.
apply of_to_val in Heq; subst; auto.
Qed.
Lemma wp_lift_nonthrow_atomic_head_step_no_fork {s E Φ} K e1 :
to_val e1 = None →
(∀ σ1, state_interp σ1 ={E}=∗
⌜head_nonthrow_reducible K e1 σ1⌝ ∗
▷ ∀ rm e2 σ2 efs, ⌜head_step K e1 σ1 e2 σ2 efs rm⌝ ={E}=∗
⌜efs = []⌝ ∗ state_interp σ2 ∗
from_option (λ v, WP fill K (of_val v) @ s; E {{Φ}}) False (to_val e2))
⊢ WP fill K e1 @ s; E {{ Φ }}.
Proof.
iIntros (?) "H". iApply wp_lift_nonthrow_atomic_head_step; eauto.
iIntros (σ1) "Hσ1". iMod ("H" $! σ1 with "Hσ1") as "[$ H]"; iModIntro.
iNext; iIntros (rm v2 σ2 efs) "%".
iMod ("H" $! rm v2 σ2 efs with "[# //]") as "(% & $ & $)"; subst; auto.
Qed.
Lemma wp_lift_nonthrow_pure_det_head_step {s E E' Φ} K e1 e2 efs :
(∀ σ1, head_nonthrow_reducible K e1 σ1) →
(∀ σ1 e2' σ2 efs' rm,
head_step K e1 σ1 e2' σ2 efs' rm → σ1 = σ2 ∧ e2 = e2' ∧ efs = efs') →
(|={E,E'}▷=> WP fill K e2 @ s; E {{ Φ }} ∗ [∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ _, True }})
⊢ WP fill K e1 @ s; E {{ Φ }}.
Proof using Hinh.
intros Hnthsp Hdet.
apply wp_lift_pure_det_step; first destruct s; eauto; simpl.
- intros σ; specialize (Hnthsp σ).
eauto using reducible_not_val,
head_prim_reducible, head_nonthrow_prim_reducible.
- intros ???? Hps.
apply head_nonthrow_reducible_prim_step in Hps; auto.
destruct Hps as (?&?&?&?&Hhs); subst.
apply Hdet in Hhs. intuition congruence.
Qed.
Lemma wp_lift_throw_pure_det_head_step {s E E' Φ} K e1 e2 efs :
(∀ σ1, head_throw_reducible K e1 σ1) →
(∀ σ1 e2' σ2 efs' rm,
head_step K e1 σ1 e2' σ2 efs' rm → σ1 = σ2 ∧ e2 = e2' ∧ efs = efs') →
(|={E,E'}▷=> WP e2 @ s; E {{ Φ }} ∗ [∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ _, True }})
⊢ WP fill K e1 @ s; E {{ Φ }}.
Proof using Hinh.
intros Hthsp Hdet. apply wp_lift_pure_det_step; destruct s; eauto.
intros σ; specialize (Hthsp σ).
eauto using reducible_not_val,
head_prim_reducible, head_throw_prim_reducible.
Qed.
Lemma wp_lift_nonthrow_pure_det_head_step_no_fork {s E E' Φ} K e1 e2 :
to_val e1 = None →
(∀ σ1, head_nonthrow_reducible K e1 σ1) →
(∀ rm σ1 e2' σ2 efs',
head_step K e1 σ1 e2' σ2 efs' rm → σ1 = σ2 ∧ e2 = e2' ∧ [] = efs') →
(|={E,E'}▷=> WP fill K e2 @ s; E {{ Φ }}) ⊢ WP fill K e1 @ s; E {{ Φ }}.
Proof using Hinh.
intros ? Hnthsp Hdet.
rewrite -(wp_lift_pure_det_step (fill K e1) (fill K e2) []) /= ?right_id;
eauto.
- destruct s; eauto.
intros σ; specialize (Hnthsp σ).
eauto using reducible_not_val,
head_prim_reducible, head_nonthrow_prim_reducible.
- intros ???? Hps.
apply head_nonthrow_reducible_prim_step in Hps; auto.
destruct Hps as (?&?&?&?&Hhs); subst.
apply Hdet in Hhs. intuition congruence.
Qed.
Lemma wp_lift_throw_pure_det_head_step_no_fork {s E E' Φ} K e1 e2 :
to_val e1 = None →
(∀ σ1, head_throw_reducible K e1 σ1) →
(∀ rm σ1 e2' σ2 efs',
head_step K e1 σ1 e2' σ2 efs' rm → σ1 = σ2 ∧ e2 = e2' ∧ [] = efs') →
(|={E,E'}▷=> WP e2 @ s; E {{ Φ }}) ⊢ WP fill K e1 @ s; E {{ Φ }}.
Proof using Hinh.
intros ? Hthsp Hdet.
rewrite -(wp_lift_pure_det_step (fill K e1) e2 []) /= ?right_id;
eauto.
destruct s; eauto.
intros σ; specialize (Hthsp σ).
eauto using reducible_not_val,
head_prim_reducible, head_throw_prim_reducible.
Qed.
Lemma wp_lift_nonthrow_pure_det_head_step_no_fork' {s E Φ} K e1 e2 :
to_val e1 = None →
(∀ σ1, head_nonthrow_reducible K e1 σ1) →
(∀ rm σ1 e2' σ2 efs',
head_step K e1 σ1 e2' σ2 efs' rm → σ1 = σ2 ∧ e2 = e2' ∧ [] = efs') →
(▷ WP fill K e2 @ s; E {{ Φ }}) ⊢ WP fill K e1 @ s; E {{ Φ }}.
Proof using Hinh.
intros.
rewrite -[(WP (fill K e1) @ _; _ {{ _ }})%I]
wp_lift_nonthrow_pure_det_head_step_no_fork //.
rewrite -step_fupd_intro //.
Qed.
Lemma wp_lift_throw_pure_det_head_step_no_fork' {s E Φ} K e1 e2 :
to_val e1 = None →
(∀ σ1, head_throw_reducible K e1 σ1) →
(∀ rm σ1 e2' σ2 efs',
head_step K e1 σ1 e2' σ2 efs' rm → σ1 = σ2 ∧ e2 = e2' ∧ [] = efs') →
(▷ WP e2 @ s; E {{ Φ }}) ⊢ WP fill K e1 @ s; E {{ Φ }}.
Proof using Hinh.
intros.
rewrite -[(WP (fill K e1) @ _; _ {{ _ }})%I]
wp_lift_throw_pure_det_head_step_no_fork //.
rewrite -step_fupd_intro //.
Qed.
Lemma wp_atomic_under_ectx E1 E2 K e Φ :
Atomic StronglyAtomic e → sub_values e → is_normal e →
(|={E1,E2}=> WP e @ E2 {{ v, |={E2,E1}=> WP fill K (of_val v) @ E1 {{ Φ }} }})
⊢ WP fill K e @ E1 {{ Φ }}.
Proof.
iIntros (Hatomic Hsv Hin) "H". rewrite !wp_unfold /wp_pre /=.
destruct (to_val (fill K e)) as [v|] eqn:HKe.
- iMod "H".
destruct (to_val e) as [e'|] eqn:Heqe'; last
by eapply (CC_fill_not_val K) in Heqe'; rewrite Heqe' in HKe.
repeat iMod "H".
rewrite (of_to_val _ _ Heqe').
iApply wp_value_inv; simpl; eauto.
- iIntros (σ1) "Hσ". iMod "H".
destruct (to_val e) as [v|] eqn:He.
+ do 2 iMod "H". rewrite (of_to_val _ _ He).
rewrite !wp_unfold /wp_pre /= HKe.
by iApply ("H" with "Hσ").
+ iMod ("H" with "Hσ") as "[Hr H]".
iDestruct "Hr" as %Hr.
iModIntro; iSplit; first by iPureIntro;
auto using reducible_under_ectx.
iNext. iIntros (e2 σ2 efs2) "Hps". iDestruct "Hps" as %Hps.
destruct (reducible_prim_step _ _ _ _ _ _ Hr Hsv Hin Hps) as
(e2' & He2 & Hps'); simpl in *; subst.
iMod ("H" with "[]") as "[Hσ [H $]]"; eauto.
apply Hatomic in Hps'.
rewrite {1}wp_unfold /wp_pre /=.
destruct (to_val e2') as [z|] eqn:Heqz.
++ rewrite (of_to_val _ _ Heqz). by repeat iMod "H"; iFrame.
++ iMod ("H" with "[$]") as "[H _]".
iDestruct "H" as %(? & ? & ? & Heq%val_prim_stuck); simpl in *.
rewrite Heq in Hps'. by inversion Hps'.
Qed.
End wp.