Skip to content
Snippets Groups Projects
Commit c0e8dc23 authored by Robbert Krebbers's avatar Robbert Krebbers
Browse files

Strengthen `bi_mono_pred` to ensure that functions are non-expansive.

parent ce8bd0c6
No related branches found
No related tags found
No related merge requests found
......@@ -73,7 +73,7 @@ Section definition.
Local Instance atomic_update_pre_mono : BiMonoPred atomic_update_pre.
Proof.
constructor.
- iIntros (P1 P2) "#HP12". iIntros ([]) "AU".
- iIntros (P1 P2 ??) "#HP12". iIntros ([]) "AU".
iApply (make_laterable_intuitionistic_wand with "[] AU").
iIntros "!> AA". iApply (atomic_acc_wand with "[HP12] AA").
iSplit; last by eauto. iApply "HP12".
......
......@@ -6,7 +6,10 @@ Import bi.
(** Least and greatest fixpoint of a monotone function, defined entirely inside
the logic. *)
Class BiMonoPred {PROP : bi} {A : ofe} (F : (A PROP) (A PROP)) := {
bi_mono_pred Φ Ψ : ( x, Φ x -∗ Ψ x) -∗ x, F Φ x -∗ F Ψ x;
bi_mono_pred Φ Ψ :
NonExpansive Φ
NonExpansive Ψ
( x, Φ x -∗ Ψ x) -∗ x, F Φ x -∗ F Ψ x;
bi_mono_pred_ne Φ : NonExpansive Φ NonExpansive (F Φ)
}.
Global Arguments bi_mono_pred {_ _ _ _} _ _.
......@@ -37,7 +40,7 @@ Section least.
Lemma least_fixpoint_unfold_2 x : F (bi_least_fixpoint F) x bi_least_fixpoint F x.
Proof using Type*.
rewrite /bi_least_fixpoint /=. iIntros "HF" (Φ) "#Hincl".
iApply "Hincl". iApply (bi_mono_pred _ Φ with "[#]"); last done.
iApply "Hincl". iApply (bi_mono_pred _ Φ with "[#] HF"); [solve_proper|].
iIntros "!>" (y) "Hy". iApply ("Hy" with "[# //]").
Qed.
......@@ -45,7 +48,7 @@ Section least.
bi_least_fixpoint F x F (bi_least_fixpoint F) x.
Proof using Type*.
iIntros "HF". iApply ("HF" $! (OfeMor (F (bi_least_fixpoint F))) with "[#]").
iIntros "!>" (y) "Hy /=". iApply (bi_mono_pred with "[#]"); last done.
iIntros "!>" (y) "Hy /=". iApply (bi_mono_pred with "[#] Hy").
iIntros "!>" (z) "?". by iApply least_fixpoint_unfold_2.
Qed.
......@@ -68,7 +71,8 @@ Section least.
trans ( x, bi_least_fixpoint F x -∗ Φ x bi_least_fixpoint F x)%I.
{ iIntros "#HΦ". iApply (least_fixpoint_ind with "[]"); first solve_proper.
iIntros "!>" (y) "H". iSplit; first by iApply "HΦ".
iApply least_fixpoint_unfold_2. iApply (bi_mono_pred with "[#] H").
iApply least_fixpoint_unfold_2.
iApply (bi_mono_pred with "[#] H"); [solve_proper|].
by iIntros "!> * [_ ?]". }
by setoid_rewrite and_elim_l.
Qed.
......
......@@ -17,7 +17,7 @@ Global Instance bi_rtc_pre_mono `{!BiInternalEq PROP}
BiMonoPred (bi_rtc_pre R x).
Proof.
constructor; [|solve_proper].
iIntros (rec1 rec2) "#H". iIntros (x1) "[Hrec | Hrec]".
iIntros (rec1 rec2 ??) "#H". iIntros (x1) "[Hrec | Hrec]".
{ by iLeft. }
iRight.
iDestruct "Hrec" as (x') "[HP Hrec]".
......
......@@ -27,7 +27,7 @@ Qed.
Local Instance twptp_pre_mono' : BiMonoPred twptp_pre.
Proof.
constructor; first apply twptp_pre_mono.
constructor; first (intros ????; apply twptp_pre_mono).
intros wp Hwp n t1 t2 ?%(discrete_iff _ _)%leibniz_equiv; solve_proper.
Qed.
......
......@@ -50,7 +50,7 @@ Definition twp_pre' `{!irisGS Λ Σ} (s : stuckness) :
Local Instance twp_pre_mono' `{!irisGS Λ Σ} s : BiMonoPred (twp_pre' s).
Proof.
constructor.
- iIntros (wp1 wp2) "#H"; iIntros ([[E e1] Φ]); iRevert (E e1 Φ).
- iIntros (wp1 wp2 ??) "#H"; iIntros ([[E e1] Φ]); iRevert (E e1 Φ).
iApply twp_pre_mono. iIntros "!>" (E e Φ). iApply ("H" $! (E,e,Φ)).
- intros wp Hwp n [[E1 e1] Φ1] [[E2 e2] Φ2]
[[?%leibniz_equiv ?%leibniz_equiv] ?]; simplify_eq/=.
......
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