Skip to content
Snippets Groups Projects
Commit 620ac27c authored by David Swasey's avatar David Swasey
Browse files

Rename stuckness lemmas.

parent dd0a95de
No related branches found
No related tags found
No related merge requests found
...@@ -57,9 +57,9 @@ Proof. solve_proper. Qed. ...@@ -57,9 +57,9 @@ Proof. solve_proper. Qed.
Lemma ht_mono s E P P' Φ Φ' e : Lemma ht_mono s E P P' Φ Φ' e :
(P P') ( v, Φ' v Φ v) {{ P' }} e @ s; E {{ Φ' }} {{ P }} e @ s; E {{ Φ }}. (P P') ( v, Φ' v Φ v) {{ P' }} e @ s; E {{ Φ' }} {{ P }} e @ s; E {{ Φ }}.
Proof. by intros; apply persistently_mono, wand_mono, wp_mono. Qed. Proof. by intros; apply persistently_mono, wand_mono, wp_mono. Qed.
Lemma ht_stuckness_mono s1 s2 E P Φ e : Lemma ht_stuck_mono s1 s2 E P Φ e :
(s1 s2)%stuckness {{ P }} e @ s1; E {{ Φ }} {{ P }} e @ s2; E {{ Φ }}. (s1 s2)%stuckness {{ P }} e @ s1; E {{ Φ }} {{ P }} e @ s2; E {{ Φ }}.
Proof. by intros; apply persistently_mono, wand_mono, wp_stuckness_mono. Qed. Proof. by intros; apply persistently_mono, wand_mono, wp_stuck_mono. Qed.
Global Instance ht_mono' s E : Global Instance ht_mono' s E :
Proper (flip () ==> eq ==> pointwise_relation _ () ==> ()) (ht s E). Proper (flip () ==> eq ==> pointwise_relation _ () ==> ()) (ht s E).
Proof. solve_proper. Qed. Proof. solve_proper. Qed.
...@@ -98,10 +98,10 @@ Proof. ...@@ -98,10 +98,10 @@ Proof.
iIntros (v) "Hv". by iApply "HwpK". iIntros (v) "Hv". by iApply "HwpK".
Qed. Qed.
Lemma ht_forget_not_stuck s E P Φ e : Lemma ht_stuck_weaken s E P Φ e :
{{ P }} e @ s; E {{ Φ }} {{ P }} e @ E ?{{ Φ }}. {{ P }} e @ s; E {{ Φ }} {{ P }} e @ E ?{{ Φ }}.
Proof. Proof.
by iIntros "#Hwp !# ?"; iApply wp_forget_not_stuck; iApply "Hwp". by iIntros "#Hwp !# ?"; iApply wp_stuck_weaken; iApply "Hwp".
Qed. Qed.
Lemma ht_mask_weaken s E1 E2 P Φ e : Lemma ht_mask_weaken s E1 E2 P Φ e :
......
...@@ -210,7 +210,7 @@ Proof. ...@@ -210,7 +210,7 @@ Proof.
iMod "Hclose" as "_". by iApply ("IH" with "HΦ"). iMod "Hclose" as "_". by iApply ("IH" with "HΦ").
Qed. Qed.
Lemma wp_forget_not_stuck s E e Φ : Lemma wp_stuck_weaken s E e Φ :
WP e @ s; E {{ Φ }} WP e @ E ?{{ Φ }}. WP e @ s; E {{ Φ }} WP e @ E ?{{ Φ }}.
Proof. Proof.
iIntros "H". iLöb as "IH" forall (E e Φ). rewrite !wp_unfold /wp_pre. iIntros "H". iLöb as "IH" forall (E e Φ). rewrite !wp_unfold /wp_pre.
...@@ -311,9 +311,9 @@ Proof. ...@@ -311,9 +311,9 @@ Proof.
iIntros () "H"; iApply (wp_strong_mono s E E); auto. iIntros () "H"; iApply (wp_strong_mono s E E); auto.
iIntros "{$H}" (v) "?". by iApply . iIntros "{$H}" (v) "?". by iApply .
Qed. Qed.
Lemma wp_stuckness_mono s1 s2 E e Φ : Lemma wp_stuck_mono s1 s2 E e Φ :
(s1 s2)%stuckness WP e @ s1; E {{ Φ }} WP e @ s2; E {{ Φ }}. (s1 s2)%stuckness WP e @ s1; E {{ Φ }} WP e @ s2; E {{ Φ }}.
Proof. case: s1; case: s2 => // _. exact: wp_forget_not_stuck. Qed. Proof. case: s1; case: s2 => // _. exact: wp_stuck_weaken. Qed.
Lemma wp_mask_mono s E1 E2 e Φ : E1 E2 WP e @ s; E1 {{ Φ }} WP e @ s; E2 {{ Φ }}. Lemma wp_mask_mono s E1 E2 e Φ : E1 E2 WP e @ s; E1 {{ Φ }} WP e @ s; E2 {{ Φ }}.
Proof. iIntros (?) "H"; iApply (wp_strong_mono s E1 E2); auto. iFrame; eauto. Qed. Proof. iIntros (?) "H"; iApply (wp_strong_mono s E1 E2); auto. iFrame; eauto. Qed.
Global Instance wp_mono' s E e : Global Instance wp_mono' s E e :
......
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