Skip to content
Snippets Groups Projects
Commit 52f90871 authored by Ralf Jung's avatar Ralf Jung
Browse files

remove unnecessary side-conditions from ownP lemmas

parent bb37a795
No related branches found
No related tags found
No related merge requests found
...@@ -36,17 +36,19 @@ Qed. ...@@ -36,17 +36,19 @@ Qed.
(** Derived lifting lemmas. *) (** Derived lifting lemmas. *)
Lemma wp_lift_pure_step `{Inhabited (state Λ)} s E E' Φ e1 : Lemma wp_lift_pure_step `{Inhabited (state Λ)} s E E' Φ e1 :
to_val e1 = None ( σ1, if s is not_stuck then reducible e1 σ1 else to_val e1 = None)
( σ1, if s is not_stuck then reducible e1 σ1 else True)
( σ1 e2 σ2 efs, prim_step e1 σ1 e2 σ2 efs σ1 = σ2) ( σ1 e2 σ2 efs, prim_step e1 σ1 e2 σ2 efs σ1 = σ2)
(|={E,E'}▷=> e2 efs σ, prim_step e1 σ e2 σ efs (|={E,E'}▷=> e2 efs σ, prim_step e1 σ e2 σ efs
WP e2 @ s; E {{ Φ }} [ list] ef efs, WP ef @ s; {{ _, True }}) WP e2 @ s; E {{ Φ }} [ list] ef efs, WP ef @ s; {{ _, True }})
WP e1 @ s; E {{ Φ }}. WP e1 @ s; E {{ Φ }}.
Proof. Proof.
iIntros (? Hsafe Hstep) "H". iApply wp_lift_step; first done. iIntros (Hsafe Hstep) "H". iApply wp_lift_step.
{ specialize (Hsafe inhabitant). destruct s; last done.
by eapply reducible_not_val. }
iIntros (σ1) "Hσ". iMod "H". iIntros (σ1) "Hσ". iMod "H".
iMod fupd_intro_mask' as "Hclose"; last iModIntro; first by set_solver. iMod fupd_intro_mask' as "Hclose"; last iModIntro; first by set_solver. iSplit.
iSplit; first by iPureIntro; apply Hsafe. iNext. iIntros (e2 σ2 efs ?). { iPureIntro. destruct s; done. }
iNext. iIntros (e2 σ2 efs ?).
destruct (Hstep σ1 e2 σ2 efs); auto; subst. destruct (Hstep σ1 e2 σ2 efs); auto; subst.
iMod "Hclose" as "_". iFrame "Hσ". iMod "H". iApply "H"; auto. iMod "Hclose" as "_". iFrame "Hσ". iMod "H". iApply "H"; auto.
Qed. Qed.
...@@ -83,13 +85,12 @@ Proof. ...@@ -83,13 +85,12 @@ Proof.
Qed. Qed.
Lemma wp_lift_pure_det_step `{Inhabited (state Λ)} {s E E' Φ} e1 e2 efs : Lemma wp_lift_pure_det_step `{Inhabited (state Λ)} {s E E' Φ} e1 e2 efs :
to_val e1 = None ( σ1, if s is not_stuck then reducible e1 σ1 else to_val e1 = None)
( σ1, if s is not_stuck then reducible e1 σ1 else true)
( σ1 e2' σ2 efs', prim_step e1 σ1 e2' σ2 efs' σ1 = σ2 e2 = e2' efs = efs') ( σ1 e2' σ2 efs', prim_step e1 σ1 e2' σ2 efs' σ1 = σ2 e2 = e2' efs = efs')
(|={E,E'}▷=> WP e2 @ s; E {{ Φ }} [ list] ef efs, WP ef @ s; {{ _, True }}) (|={E,E'}▷=> WP e2 @ s; E {{ Φ }} [ list] ef efs, WP ef @ s; {{ _, True }})
WP e1 @ s; E {{ Φ }}. WP e1 @ s; E {{ Φ }}.
Proof. Proof.
iIntros (?? Hpuredet) "H". iApply (wp_lift_pure_step s E E'); try done. iIntros (? Hpuredet) "H". iApply (wp_lift_pure_step s E E'); try done.
{ by intros; eapply Hpuredet. } { by intros; eapply Hpuredet. }
iApply (step_fupd_wand with "H"); iIntros "H". iApply (step_fupd_wand with "H"); iIntros "H".
by iIntros (e' efs' σ (_&->&->)%Hpuredet). by iIntros (e' efs' σ (_&->&->)%Hpuredet).
...@@ -102,9 +103,8 @@ Lemma wp_pure_step_fupd `{Inhabited (state Λ)} s E E' e1 e2 φ Φ : ...@@ -102,9 +103,8 @@ Lemma wp_pure_step_fupd `{Inhabited (state Λ)} s E E' e1 e2 φ Φ :
Proof. Proof.
iIntros ([??] ) "HWP". iIntros ([??] ) "HWP".
iApply (wp_lift_pure_det_step with "[HWP]"). iApply (wp_lift_pure_det_step with "[HWP]").
- apply (reducible_not_val _ inhabitant). by auto. - intros σ. specialize (pure_exec_safe σ). destruct s; eauto using reducible_not_val.
- destruct s; naive_solver. - destruct s; naive_solver.
- naive_solver.
- by rewrite big_sepL_nil right_id. - by rewrite big_sepL_nil right_id.
Qed. Qed.
......
...@@ -86,20 +86,23 @@ Section lifting. ...@@ -86,20 +86,23 @@ Section lifting.
Proof. rewrite /ownP; apply _. Qed. Proof. rewrite /ownP; apply _. Qed.
Lemma ownP_lift_step s E Φ e1 : Lemma ownP_lift_step s E Φ e1 :
to_val e1 = None (|={E,}=> σ1, if s is not_stuck then reducible e1 σ1 else to_val e1 = None ownP σ1
(|={E,}=> σ1, if s is not_stuck then reducible e1 σ1 else True ownP σ1
e2 σ2 efs, prim_step e1 σ1 e2 σ2 efs -∗ ownP σ2 e2 σ2 efs, prim_step e1 σ1 e2 σ2 efs -∗ ownP σ2
={,E}=∗ WP e2 @ s; E {{ Φ }} [ list] ef efs, WP ef @ s; {{ _, True }}) ={,E}=∗ WP e2 @ s; E {{ Φ }} [ list] ef efs, WP ef @ s; {{ _, True }})
WP e1 @ s; E {{ Φ }}. WP e1 @ s; E {{ Φ }}.
Proof. Proof.
iIntros (?) "H". iApply wp_lift_step; first done. iIntros "H". destruct (to_val e1) as [v|] eqn:EQe1.
iIntros (σ1) "Hσ"; iMod "H" as (σ1') "(% & >Hσf & H)". - apply of_to_val in EQe1 as <-. iApply fupd_wp.
iDestruct (ownP_eq with "Hσ Hσf") as %->. iMod "H" as (σ1) "[Hred _]"; iDestruct "Hred" as %Hred.
iModIntro. iSplit; first done. iNext. iIntros (e2 σ2 efs Hstep). destruct s; last done. apply reducible_not_val in Hred.
rewrite /ownP; iMod (own_update_2 with "Hσ Hσf") as "[Hσ Hσf]". move: Hred; by rewrite to_of_val.
{ by apply auth_update, option_local_update, - iApply wp_lift_step; [done|]; iIntros (σ1) "Hσ".
(exclusive_local_update _ (Excl σ2)). } iMod "H" as (σ1' ?) "[>Hσf H]". iDestruct (ownP_eq with "Hσ Hσf") as %->.
iFrame "Hσ". by iApply ("H" with "[]"); eauto. iModIntro; iSplit; [by destruct s|]; iNext; iIntros (e2 σ2 efs Hstep).
rewrite /ownP; iMod (own_update_2 with "Hσ Hσf") as "[Hσ Hσf]".
{ by apply auth_update, option_local_update,
(exclusive_local_update _ (Excl σ2)). }
iFrame "Hσ". iApply ("H" with "[]"); eauto.
Qed. Qed.
Lemma ownP_lift_stuck E Φ e : Lemma ownP_lift_stuck E Φ e :
...@@ -115,15 +118,16 @@ Section lifting. ...@@ -115,15 +118,16 @@ Section lifting.
by iDestruct (ownP_eq with "Hσ Hσf") as %->. by iDestruct (ownP_eq with "Hσ Hσf") as %->.
Qed. Qed.
Lemma ownP_lift_pure_step s E Φ e1 : Lemma ownP_lift_pure_step `{Inhabited (state Λ)} s E Φ e1 :
to_val e1 = None ( σ1, if s is not_stuck then reducible e1 σ1 else to_val e1 = None)
( σ1, if s is not_stuck then reducible e1 σ1 else True)
( σ1 e2 σ2 efs, prim_step e1 σ1 e2 σ2 efs σ1 = σ2) ( σ1 e2 σ2 efs, prim_step e1 σ1 e2 σ2 efs σ1 = σ2)
( e2 efs σ, prim_step e1 σ e2 σ efs ( e2 efs σ, prim_step e1 σ e2 σ efs
WP e2 @ s; E {{ Φ }} [ list] ef efs, WP ef @ s; {{ _, True }}) WP e2 @ s; E {{ Φ }} [ list] ef efs, WP ef @ s; {{ _, True }})
WP e1 @ s; E {{ Φ }}. WP e1 @ s; E {{ Φ }}.
Proof. Proof.
iIntros (? Hsafe Hstep) "H"; iApply wp_lift_step; first done. iIntros (Hsafe Hstep) "H"; iApply wp_lift_step.
{ specialize (Hsafe inhabitant). destruct s; last done.
by eapply reducible_not_val. }
iIntros (σ1) "Hσ". iMod (fupd_intro_mask' E ) as "Hclose"; first set_solver. iIntros (σ1) "Hσ". iMod (fupd_intro_mask' E ) as "Hclose"; first set_solver.
iModIntro; iSplit; [by destruct s|]; iNext; iIntros (e2 σ2 efs ?). iModIntro; iSplit; [by destruct s|]; iNext; iIntros (e2 σ2 efs ?).
destruct (Hstep σ1 e2 σ2 efs); auto; subst. destruct (Hstep σ1 e2 σ2 efs); auto; subst.
...@@ -132,13 +136,12 @@ Section lifting. ...@@ -132,13 +136,12 @@ Section lifting.
(** Derived lifting lemmas. *) (** Derived lifting lemmas. *)
Lemma ownP_lift_atomic_step {s E Φ} e1 σ1 : Lemma ownP_lift_atomic_step {s E Φ} e1 σ1 :
to_val e1 = None (if s is not_stuck then reducible e1 σ1 else to_val e1 = None)
(if s is not_stuck then reducible e1 σ1 else True)
( ownP σ1 e2 σ2 efs, prim_step e1 σ1 e2 σ2 efs -∗ ownP σ2 -∗ ( ownP σ1 e2 σ2 efs, prim_step e1 σ1 e2 σ2 efs -∗ ownP σ2 -∗
default False (to_val e2) Φ [ list] ef efs, WP ef @ s; {{ _, True }}) default False (to_val e2) Φ [ list] ef efs, WP ef @ s; {{ _, True }})
WP e1 @ s; E {{ Φ }}. WP e1 @ s; E {{ Φ }}.
Proof. Proof.
iIntros (??) "[Hσ H]"; iApply ownP_lift_step; first done. iIntros (?) "[Hσ H]"; iApply ownP_lift_step.
iMod (fupd_intro_mask' E ) as "Hclose"; first set_solver. iMod (fupd_intro_mask' E ) as "Hclose"; first set_solver.
iModIntro; iExists σ1; iFrame; iSplit; first by destruct s. iModIntro; iExists σ1; iFrame; iSplit; first by destruct s.
iNext; iIntros (e2 σ2 efs) "% Hσ". iNext; iIntros (e2 σ2 efs) "% Hσ".
...@@ -148,22 +151,20 @@ Section lifting. ...@@ -148,22 +151,20 @@ Section lifting.
Qed. Qed.
Lemma ownP_lift_atomic_det_step {s E Φ e1} σ1 v2 σ2 efs : Lemma ownP_lift_atomic_det_step {s E Φ e1} σ1 v2 σ2 efs :
to_val e1 = None (if s is not_stuck then reducible e1 σ1 else to_val e1 = None)
(if s is not_stuck then reducible e1 σ1 else True)
( e2' σ2' efs', prim_step e1 σ1 e2' σ2' efs' ( e2' σ2' efs', prim_step e1 σ1 e2' σ2' efs'
σ2 = σ2' to_val e2' = Some v2 efs = efs') σ2 = σ2' to_val e2' = Some v2 efs = efs')
ownP σ1 (ownP σ2 -∗ ownP σ1 (ownP σ2 -∗
Φ v2 [ list] ef efs, WP ef @ s; {{ _, True }}) Φ v2 [ list] ef efs, WP ef @ s; {{ _, True }})
WP e1 @ s; E {{ Φ }}. WP e1 @ s; E {{ Φ }}.
Proof. Proof.
iIntros (?? Hdet) "[Hσ1 Hσ2]"; iApply ownP_lift_atomic_step; try done. iIntros (? Hdet) "[Hσ1 Hσ2]"; iApply ownP_lift_atomic_step; try done.
iFrame; iNext; iIntros (e2' σ2' efs') "% Hσ2'". iFrame; iNext; iIntros (e2' σ2' efs') "% Hσ2'".
edestruct Hdet as (->&Hval&->). done. by rewrite Hval; iApply "Hσ2". edestruct Hdet as (->&Hval&->). done. by rewrite Hval; iApply "Hσ2".
Qed. Qed.
Lemma ownP_lift_atomic_det_step_no_fork {s E e1} σ1 v2 σ2 : Lemma ownP_lift_atomic_det_step_no_fork {s E e1} σ1 v2 σ2 :
to_val e1 = None (if s is not_stuck then reducible e1 σ1 else to_val e1 = None)
(if s is not_stuck then reducible e1 σ1 else True)
( e2' σ2' efs', prim_step e1 σ1 e2' σ2' efs' ( e2' σ2' efs', prim_step e1 σ1 e2' σ2' efs'
σ2 = σ2' to_val e2' = Some v2 [] = efs') σ2 = σ2' to_val e2' = Some v2 [] = efs')
{{{ ownP σ1 }}} e1 @ s; E {{{ RET v2; ownP σ2 }}}. {{{ ownP σ1 }}} e1 @ s; E {{{ RET v2; ownP σ2 }}}.
...@@ -172,20 +173,18 @@ Section lifting. ...@@ -172,20 +173,18 @@ Section lifting.
rewrite big_sepL_nil right_id. by apply uPred.wand_intro_r. rewrite big_sepL_nil right_id. by apply uPred.wand_intro_r.
Qed. Qed.
Lemma ownP_lift_pure_det_step {s E Φ} e1 e2 efs : Lemma ownP_lift_pure_det_step `{Inhabited (state Λ)} {s E Φ} e1 e2 efs :
to_val e1 = None ( σ1, if s is not_stuck then reducible e1 σ1 else to_val e1 = None)
( σ1, if s is not_stuck then reducible e1 σ1 else True)
( σ1 e2' σ2 efs', prim_step e1 σ1 e2' σ2 efs' σ1 = σ2 e2 = e2' efs = efs') ( σ1 e2' σ2 efs', prim_step e1 σ1 e2' σ2 efs' σ1 = σ2 e2 = e2' efs = efs')
(WP e2 @ s; E {{ Φ }} [ list] ef efs, WP ef @ s; {{ _, True }}) (WP e2 @ s; E {{ Φ }} [ list] ef efs, WP ef @ s; {{ _, True }})
WP e1 @ s; E {{ Φ }}. WP e1 @ s; E {{ Φ }}.
Proof. Proof.
iIntros (?? Hpuredet) "?"; iApply ownP_lift_pure_step=>//. iIntros (? Hpuredet) "?"; iApply ownP_lift_pure_step=>//.
by apply Hpuredet. by iNext; iIntros (e' efs' σ (_&->&->)%Hpuredet). by apply Hpuredet. by iNext; iIntros (e' efs' σ (_&->&->)%Hpuredet).
Qed. Qed.
Lemma ownP_lift_pure_det_step_no_fork `{Inhabited (state Λ)} {s E Φ} e1 e2 : Lemma ownP_lift_pure_det_step_no_fork `{Inhabited (state Λ)} {s E Φ} e1 e2 :
to_val e1 = None ( σ1, if s is not_stuck then reducible e1 σ1 else to_val e1 = None)
( σ1, if s is not_stuck then reducible e1 σ1 else True)
( σ1 e2' σ2 efs', prim_step e1 σ1 e2' σ2 efs' σ1 = σ2 e2 = e2' [] = efs') ( σ1 e2' σ2 efs', prim_step e1 σ1 e2' σ2 efs' σ1 = σ2 e2 = e2' [] = efs')
WP e2 @ s; E {{ Φ }} WP e1 @ s; E {{ Φ }}. WP e2 @ s; E {{ Φ }} WP e1 @ s; E {{ Φ }}.
Proof. Proof.
...@@ -204,15 +203,15 @@ Section ectx_lifting. ...@@ -204,15 +203,15 @@ Section ectx_lifting.
Hint Resolve head_stuck_stuck. Hint Resolve head_stuck_stuck.
Lemma ownP_lift_head_step s E Φ e1 : Lemma ownP_lift_head_step s E Φ e1 :
to_val e1 = None
(|={E,}=> σ1, head_reducible e1 σ1 ownP σ1 (|={E,}=> σ1, head_reducible e1 σ1 ownP σ1
e2 σ2 efs, head_step e1 σ1 e2 σ2 efs -∗ ownP σ2 e2 σ2 efs, head_step e1 σ1 e2 σ2 efs -∗ ownP σ2
={,E}=∗ WP e2 @ s; E {{ Φ }} [ list] ef efs, WP ef @ s; {{ _, True }}) ={,E}=∗ WP e2 @ s; E {{ Φ }} [ list] ef efs, WP ef @ s; {{ _, True }})
WP e1 @ s; E {{ Φ }}. WP e1 @ s; E {{ Φ }}.
Proof. Proof.
iIntros (?) "H". iApply ownP_lift_step; first done. iIntros "H". iApply ownP_lift_step.
iMod "H" as (σ1 ?) "[Hσ1 Hwp]". iModIntro. iExists σ1. iMod "H" as (σ1 ?) "[Hσ1 Hwp]". iModIntro. iExists σ1. iSplit.
iSplit; first by destruct s; eauto. iFrame. iNext. iIntros (e2 σ2 efs) "% ?". { destruct s; try by eauto using reducible_not_val. }
iFrame. iNext. iIntros (e2 σ2 efs) "% ?".
iApply ("Hwp" with "[]"); eauto. iApply ("Hwp" with "[]"); eauto.
Qed. Qed.
...@@ -226,71 +225,65 @@ Section ectx_lifting. ...@@ -226,71 +225,65 @@ Section ectx_lifting.
Qed. Qed.
Lemma ownP_lift_pure_head_step s E Φ e1 : Lemma ownP_lift_pure_head_step s E Φ e1 :
to_val e1 = None
( σ1, head_reducible e1 σ1) ( σ1, head_reducible e1 σ1)
( σ1 e2 σ2 efs, head_step e1 σ1 e2 σ2 efs σ1 = σ2) ( σ1 e2 σ2 efs, head_step e1 σ1 e2 σ2 efs σ1 = σ2)
( e2 efs σ, head_step e1 σ e2 σ efs ( e2 efs σ, head_step e1 σ e2 σ efs
WP e2 @ s; E {{ Φ }} [ list] ef efs, WP ef @ s; {{ _, True }}) WP e2 @ s; E {{ Φ }} [ list] ef efs, WP ef @ s; {{ _, True }})
WP e1 @ s; E {{ Φ }}. WP e1 @ s; E {{ Φ }}.
Proof using Hinh. Proof using Hinh.
iIntros (???) "H". iApply ownP_lift_pure_step; eauto. iIntros (??) "H". iApply ownP_lift_pure_step; eauto.
{ by destruct s; auto. } { by destruct s; auto. }
iNext. iIntros (????). iApply "H"; eauto. iNext. iIntros (????). iApply "H"; eauto.
Qed. Qed.
Lemma ownP_lift_atomic_head_step {s E Φ} e1 σ1 : Lemma ownP_lift_atomic_head_step {s E Φ} e1 σ1 :
to_val e1 = None
head_reducible e1 σ1 head_reducible e1 σ1
ownP σ1 ( e2 σ2 efs, ownP σ1 ( e2 σ2 efs,
head_step e1 σ1 e2 σ2 efs -∗ ownP σ2 -∗ head_step e1 σ1 e2 σ2 efs -∗ ownP σ2 -∗
default False (to_val e2) Φ [ list] ef efs, WP ef @ s; {{ _, True }}) default False (to_val e2) Φ [ list] ef efs, WP ef @ s; {{ _, True }})
WP e1 @ s; E {{ Φ }}. WP e1 @ s; E {{ Φ }}.
Proof. Proof.
iIntros (??) "[? H]". iApply ownP_lift_atomic_step; eauto. iIntros (?) "[? H]". iApply ownP_lift_atomic_step; eauto.
{ by destruct s; eauto. } { by destruct s; eauto using reducible_not_val. }
iFrame. iNext. iIntros (???) "% ?". iApply ("H" with "[]"); eauto. iFrame. iNext. iIntros (???) "% ?". iApply ("H" with "[]"); eauto.
Qed. Qed.
Lemma ownP_lift_atomic_det_head_step {s E Φ e1} σ1 v2 σ2 efs : Lemma ownP_lift_atomic_det_head_step {s E Φ e1} σ1 v2 σ2 efs :
to_val e1 = None
head_reducible e1 σ1 head_reducible e1 σ1
( e2' σ2' efs', head_step e1 σ1 e2' σ2' efs' ( e2' σ2' efs', head_step e1 σ1 e2' σ2' efs'
σ2 = σ2' to_val e2' = Some v2 efs = efs') σ2 = σ2' to_val e2' = Some v2 efs = efs')
ownP σ1 (ownP σ2 -∗ Φ v2 [ list] ef efs, WP ef @ s; {{ _, True }}) ownP σ1 (ownP σ2 -∗ Φ v2 [ list] ef efs, WP ef @ s; {{ _, True }})
WP e1 @ s; E {{ Φ }}. WP e1 @ s; E {{ Φ }}.
Proof. Proof.
by destruct s; eauto 10 using ownP_lift_atomic_det_step. by destruct s; eauto 10 using ownP_lift_atomic_det_step, reducible_not_val.
Qed. Qed.
Lemma ownP_lift_atomic_det_head_step_no_fork {s E e1} σ1 v2 σ2 : Lemma ownP_lift_atomic_det_head_step_no_fork {s E e1} σ1 v2 σ2 :
to_val e1 = None
head_reducible e1 σ1 head_reducible e1 σ1
( e2' σ2' efs', head_step e1 σ1 e2' σ2' efs' ( e2' σ2' efs', head_step e1 σ1 e2' σ2' efs'
σ2 = σ2' to_val e2' = Some v2 [] = efs') σ2 = σ2' to_val e2' = Some v2 [] = efs')
{{{ ownP σ1 }}} e1 @ s; E {{{ RET v2; ownP σ2 }}}. {{{ ownP σ1 }}} e1 @ s; E {{{ RET v2; ownP σ2 }}}.
Proof. Proof.
intros ???; apply ownP_lift_atomic_det_step_no_fork; eauto. intros ???; apply ownP_lift_atomic_det_step_no_fork; eauto.
by destruct s; eauto. by destruct s; eauto using reducible_not_val.
Qed. Qed.
Lemma ownP_lift_pure_det_head_step {s E Φ} e1 e2 efs : Lemma ownP_lift_pure_det_head_step {s E Φ} e1 e2 efs :
to_val e1 = None
( σ1, head_reducible e1 σ1) ( σ1, head_reducible e1 σ1)
( σ1 e2' σ2 efs', head_step e1 σ1 e2' σ2 efs' σ1 = σ2 e2 = e2' efs = efs') ( σ1 e2' σ2 efs', head_step e1 σ1 e2' σ2 efs' σ1 = σ2 e2 = e2' efs = efs')
(WP e2 @ s; E {{ Φ }} [ list] ef efs, WP ef @ s; {{ _, True }}) (WP e2 @ s; E {{ Φ }} [ list] ef efs, WP ef @ s; {{ _, True }})
WP e1 @ s; E {{ Φ }}. WP e1 @ s; E {{ Φ }}.
Proof using Hinh. Proof using Hinh.
iIntros (???) "H"; iApply wp_lift_pure_det_step; eauto. iIntros (??) "H"; iApply wp_lift_pure_det_step; eauto.
by destruct s; eauto. by destruct s; eauto using reducible_not_val.
Qed. Qed.
Lemma ownP_lift_pure_det_head_step_no_fork {s E Φ} e1 e2 : Lemma ownP_lift_pure_det_head_step_no_fork {s E Φ} e1 e2 :
to_val e1 = None
( σ1, head_reducible e1 σ1) ( σ1, head_reducible e1 σ1)
( σ1 e2' σ2 efs', head_step e1 σ1 e2' σ2 efs' σ1 = σ2 e2 = e2' [] = efs') ( σ1 e2' σ2 efs', head_step e1 σ1 e2' σ2 efs' σ1 = σ2 e2 = e2' [] = efs')
WP e2 @ s; E {{ Φ }} WP e1 @ s; E {{ Φ }}. WP e2 @ s; E {{ Φ }} WP e1 @ s; E {{ Φ }}.
Proof using Hinh. Proof using Hinh.
iIntros (???) "H". iApply ownP_lift_pure_det_step_no_fork; eauto. iIntros (??) "H". iApply ownP_lift_pure_det_step_no_fork; eauto.
by destruct s; eauto. by destruct s; eauto using reducible_not_val.
Qed. Qed.
End ectx_lifting. End ectx_lifting.
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment