Skip to content
Snippets Groups Projects

Compare revisions

Changes are shown as if the source revision was being merged into the target revision. Learn more about comparing revisions.

Source

Select target project
No results found

Target

Select target project
  • iris/iris
  • jeehoon.kang/iris-coq
  • amintimany/iris-coq
  • dfrumin/iris-coq
  • Villetaneuse/iris
  • gares/iris
  • shiatsumat/iris
  • Blaisorblade/iris
  • jihgfee/iris-coq
  • mrhaandi/iris
  • tlsomers/iris
  • Quarkbeast/iris-coq
  • janno/iris
  • amaurremi/iris-coq
  • proux/iris
  • tchajed/iris
  • herbelin/iris-coq
  • msammler/iris-coq
  • maximedenes/iris-coq
  • bpeters/iris
  • haidang/iris
  • lepigre/iris
  • lczch/iris
  • simonspies/iris
  • gpirlea/iris
  • dkhalanskiyjb/iris
  • gmalecha/iris
  • germanD/iris
  • aa755/iris
  • jules/iris
  • abeln/iris
  • simonfv/iris
  • atrieu/iris
  • arthuraa/iris
  • simonh/iris
  • jung/iris
  • mattam82/iris
  • Armael/iris
  • adamAndMath/iris
  • gmevel/iris
  • snyke7/iris
  • johannes/iris
  • NiklasM/iris
  • simonspies/iris-parametric-index
  • svancollem/iris
  • proux1/iris
  • wmansky/iris
  • LukeXuan/iris
  • ivanbakel/iris
  • SkySkimmer/iris
  • tjhance/iris
  • yiyunliu/iris
  • Lee-Janggun/iris
  • thomas-lamiaux/iris
  • dongjae/iris
  • dnezam/iris
  • Tragicus/iris
  • clef-men/iris
  • ffengyu/iris
59 results
Show changes
Showing
with 6178 additions and 0 deletions
From iris.proofmode Require Import proofmode.
From iris.bi Require Import internal_eq.
From iris.algebra Require Import cmra csum excl agree.
From iris.prelude Require Import options.
(** Derived [≼] connective on [cmra] elements. This can be defined on
any [bi] that has internal equality [≡]. It corresponds to the
step-indexed [≼{n}] connective in the [uPred] model. *)
Definition internal_included `{!BiInternalEq PROP} {A : cmra} (a b : A) : PROP :=
c, b a c.
Global Arguments internal_included {_ _ _} _ _ : simpl never.
Global Instance: Params (@internal_included) 3 := {}.
Global Typeclasses Opaque internal_included.
Infix "≼" := internal_included : bi_scope.
Section internal_included_laws.
Context `{!BiInternalEq PROP}.
Implicit Type A B : cmra.
(* Force implicit argument PROP *)
Notation "P ⊢ Q" := (P ⊢@{PROP} Q).
Notation "P ⊣⊢ Q" := (P ⊣⊢@{PROP} Q).
(** Propers *)
Global Instance internal_included_nonexpansive A :
NonExpansive2 (internal_included (PROP := PROP) (A := A)).
Proof. solve_proper. Qed.
Global Instance internal_included_proper A :
Proper (() ==> () ==> (⊣⊢)) (internal_included (PROP := PROP) (A := A)).
Proof. solve_proper. Qed.
(** Proofmode support *)
Global Instance into_pure_internal_included {A} (a b : A) `{!Discrete b} :
@IntoPure PROP (a b) (a b).
Proof. rewrite /IntoPure /internal_included. eauto. Qed.
Global Instance from_pure_internal_included {A} (a b : A) :
@FromPure PROP false (a b) (a b).
Proof. rewrite /FromPure /= /internal_included. eauto. Qed.
Global Instance into_exist_internal_included {A} (a b : A) :
IntoExist (PROP := PROP) (a b) (λ c, b a c)%I (λ x, x).
Proof. by rewrite /IntoExist. Qed.
Global Instance from_exist_internal_included {A} (a b : A) :
FromExist (PROP := PROP) (a b) (λ c, b a c)%I.
Proof. by rewrite /FromExist. Qed.
Global Instance internal_included_persistent {A} (a b : A) :
Persistent (PROP := PROP) (a b).
Proof. rewrite /internal_included. apply _. Qed.
Global Instance internal_included_absorbing {A} (a b : A) :
Absorbing (PROP := PROP) (a b).
Proof. rewrite /internal_included. apply _. Qed.
Lemma internal_included_refl `{!CmraTotal A} (x : A) : ⊢@{PROP} x x.
Proof. iExists (core x). by rewrite cmra_core_r. Qed.
Lemma internal_included_trans {A} (x y z : A) :
⊢@{PROP} x y -∗ y z -∗ x z.
Proof.
iIntros "#[%x' Hx'] #[%y' Hy']". iExists (x' y').
rewrite assoc. by iRewrite -"Hx'".
Qed.
(** Simplification lemmas *)
Lemma f_homom_includedI {A B} (x y : A) (f : A B) `{!NonExpansive f} :
(* This is a slightly weaker condition than being a [CmraMorphism] *)
( c, f x f c f (x c))
(x y f x f y).
Proof.
intros f_homom. iDestruct 1 as (z) "Hz".
iExists (f z). rewrite f_homom.
by iApply f_equivI.
Qed.
Lemma prod_includedI {A B} (x y : A * B) :
x y ⊣⊢ (x.1 y.1) (x.2 y.2).
Proof.
destruct x as [x1 x2], y as [y1 y2]; simpl; iSplit.
- iIntros "#[%z H]". rewrite prod_equivI /=. iDestruct "H" as "[??]".
iSplit; by iExists _.
- iIntros "#[[%z1 Hz1] [%z2 Hz2]]". iExists (z1, z2).
rewrite prod_equivI /=; auto.
Qed.
Lemma option_includedI {A} (mx my : option A) :
mx my ⊣⊢ match mx, my with
| Some x, Some y => (x y) (x y)
| None, _ => True
| Some x, None => False
end.
Proof.
iSplit.
- iIntros "[%mz H]". rewrite option_equivI.
destruct mx as [x|], my as [y|], mz as [z|]; simpl; auto; [|].
+ iLeft. by iExists z.
+ iRight. by iRewrite "H".
- destruct mx as [x|], my as [y|]; simpl; auto; [|].
+ iDestruct 1 as "[[%z H]|H]"; iRewrite "H".
* by iExists (Some z).
* by iExists None.
+ iIntros "_". by iExists (Some y).
Qed.
Lemma option_included_totalI `{!CmraTotal A} (mx my : option A) :
mx my ⊣⊢ match mx, my with
| Some x, Some y => x y
| None, _ => True
| Some x, None => False
end.
Proof.
rewrite option_includedI. destruct mx as [x|], my as [y|]; [|done..].
iSplit; [|by auto].
iIntros "[Hx|Hx] //". iRewrite "Hx". iApply (internal_included_refl y).
Qed.
Lemma Some_included_totalI `{!CmraTotal A} (x y : A) :
Some x Some y ⊣⊢ x y.
Proof. by rewrite option_included_totalI. Qed.
Lemma csum_includedI {A B} (sx sy : csum A B) :
sx sy ⊣⊢ match sx, sy with
| Cinl x, Cinl y => x y
| Cinr x, Cinr y => x y
| _, CsumInvalid => True
| _, _ => False
end.
Proof.
iSplit.
- iDestruct 1 as (sz) "H". rewrite csum_equivI.
destruct sx, sy, sz; rewrite /internal_included /=; auto.
- destruct sx as [x|x|], sy as [y|y|]; eauto; [|].
+ iIntros "#[%z H]". iExists (Cinl z). by rewrite csum_equivI.
+ iIntros "#[%z H]". iExists (Cinr z). by rewrite csum_equivI.
Qed.
Lemma excl_includedI {O : ofe} (x y : excl O) :
x y ⊣⊢ y = ExclInvalid ⌝.
Proof.
iSplit.
- iIntros "[%z Hz]". rewrite excl_equivI. destruct y, x, z; auto.
- iIntros (->). by iExists ExclInvalid.
Qed.
Lemma agree_includedI {O : ofe} (x y : agree O) : x y ⊣⊢ y x y.
Proof.
iSplit.
+ iIntros "[%z Hz]". iRewrite "Hz". by rewrite assoc agree_idemp.
+ iIntros "H". by iExists _.
Qed.
End internal_included_laws.
From iris.bi Require Export bi plainly.
From iris.proofmode Require Import proofmode.
From iris.prelude Require Import options.
Import bi.
(** The "core" of an assertion is its maximal persistent part,
i.e. the conjunction of all persistent assertions that are weaker
than P (as in, implied by P). *)
Definition coreP `{!BiPlainly PROP} (P : PROP) : PROP :=
(* TODO: Looks like we want notation for affinely-plainly; that lets us avoid
using conjunction/implication here. *)
Q : PROP, <affine> (Q -∗ <pers> Q) -∗ <affine> (P -∗ Q) -∗ Q.
Global Instance: Params (@coreP) 1 := {}.
Global Typeclasses Opaque coreP.
Section core.
Context {PROP : bi} `{!BiPlainly PROP}.
Implicit Types P Q : PROP.
Lemma coreP_intro P : P -∗ coreP P.
Proof.
rewrite /coreP. iIntros "HP" (Q) "_ HPQ".
(* FIXME: Cannot apply HPQ directly. This works if we move it to the
persistent context, but why should we? *)
iDestruct (affinely_plainly_elim with "HPQ") as "HPQ".
by iApply "HPQ".
Qed.
Global Instance coreP_persistent
`{!BiPersistentlyForall PROP, !BiPersistentlyImplPlainly PROP} P :
Persistent (coreP P).
Proof.
rewrite /coreP /Persistent. iIntros "HC" (Q).
iApply persistently_wand_affinely_plainly. iIntros "#HQ".
iApply persistently_wand_affinely_plainly. iIntros "#HPQ".
iApply "HQ". iApply "HC"; auto.
Qed.
Global Instance coreP_affine P : Affine P Affine (coreP P).
Proof. intros ?. rewrite /coreP /Affine. iIntros "HC". iApply "HC"; eauto. Qed.
Global Instance coreP_ne : NonExpansive (coreP (PROP:=PROP)).
Proof. solve_proper. Qed.
Global Instance coreP_proper : Proper ((⊣⊢) ==> (⊣⊢)) (coreP (PROP:=PROP)).
Proof. solve_proper. Qed.
Global Instance coreP_mono : Proper (() ==> ()) (coreP (PROP:=PROP)).
Proof. solve_proper. Qed.
Global Instance coreP_flip_mono :
Proper (flip () ==> flip ()) (coreP (PROP:=PROP)).
Proof. solve_proper. Qed.
Lemma coreP_wand P Q : <affine> (P -∗ Q) -∗ coreP P -∗ coreP Q.
Proof.
rewrite /coreP. iIntros "#HPQ HP" (R) "#HR #HQR". iApply ("HP" with "HR").
iIntros "!> !> HP". iApply "HQR". by iApply "HPQ".
Qed.
Lemma coreP_elim P : Persistent P coreP P -∗ P.
Proof. rewrite /coreP. iIntros (?) "HCP". iApply "HCP"; auto. Qed.
(** The [<affine>] modality is needed for general BIs:
- The right-to-left direction corresponds to elimination of [<pers>], which
cannot be done without [<affine>].
- The left-to-right direction corresponds the introduction of [<pers>]. The
[<affine>] modality makes it stronger since it appears in the LHS of the
[⊢] in the premise. As a user, you have to prove [<affine> coreP P ⊢ Q],
which is weaker than [coreP P ⊢ Q]. *)
Lemma coreP_entails `{!BiPersistentlyForall PROP, !BiPersistentlyImplPlainly PROP} P Q :
(<affine> coreP P Q) (P <pers> Q).
Proof.
split.
- iIntros (HP) "HP". iDestruct (coreP_intro with "HP") as "#HcP {HP}".
iIntros "!>". by iApply HP.
- iIntros (->) "HcQ". by iDestruct (coreP_elim with "HcQ") as "#HQ".
Qed.
(** A more convenient variant of the above lemma for affine [P]. *)
Lemma coreP_entails' `{!BiPersistentlyForall PROP, !BiPersistentlyImplPlainly PROP}
P Q `{!Affine P} :
(coreP P Q) (P Q).
Proof.
rewrite -(affine_affinely (coreP P)) coreP_entails. split.
- rewrite -{2}(affine_affinely P). by intros ->.
- intros ->. apply affinely_elim.
Qed.
End core.
From iris.bi Require Export bi.
From iris.proofmode Require Import proofmode.
From iris.prelude Require Import options.
(* The sections add extra BI assumptions, which is only picked up with "Type"*. *)
Set Default Proof Using "Type*".
(** This proves that in an affine BI (i.e., a BI that enjoys [P ∗ Q ⊢ P]), the
classical excluded middle ([P ∨ ¬P]) axiom makes the separation conjunction
trivial, i.e., it gives [P -∗ P ∗ P] and [P ∧ Q -∗ P ∗ Q].
Our proof essentially follows the structure of the proof of Theorem 3 in
https://scholar.princeton.edu/sites/default/files/qinxiang/files/putting_order_to_the_separation_logic_jungle_revised_version.pdf *)
Module affine_em. Section affine_em.
Context {PROP : bi} `{!BiAffine PROP}.
Context (em : P : PROP, P ¬P).
Implicit Types P Q : PROP.
Lemma sep_dup P : P -∗ P P.
Proof.
iIntros "HP". iDestruct (em P) as "[HP'|HnotP]".
- iFrame "HP HP'".
- iExFalso. by iApply "HnotP".
Qed.
Lemma and_sep P Q : P Q -∗ P Q.
Proof.
iIntros "HPQ". iDestruct (sep_dup with "HPQ") as "[HPQ HPQ']".
iSplitL "HPQ".
- by iDestruct "HPQ" as "[HP _]".
- by iDestruct "HPQ'" as "[_ HQ]".
Qed.
End affine_em. End affine_em.
(** This proves that the combination of Löb induction [(▷ P → P) ⊢ P] and the
classical excluded-middle [P ∨ ¬P] axiom makes the later operator trivial, i.e.,
it gives [▷ P] for any [P], or equivalently [▷ P ≡ True]. *)
Module löb_em. Section löb_em.
Context {PROP : bi} `{!BiLöb PROP}.
Context (em : P : PROP, P ¬P).
Implicit Types P : PROP.
Lemma later_anything P : ⊢@{PROP} P.
Proof.
iDestruct (em ( False)) as "#[HP|HnotP]".
- iNext. done.
- iExFalso. iLöb as "IH". iSpecialize ("HnotP" with "IH"). done.
Qed.
End löb_em. End löb_em.
(** This proves that we need the ▷ in a "Saved Proposition" construction with
name-dependent allocation. *)
Module savedprop. Section savedprop.
Context {PROP : bi} `{!BiAffine PROP}.
Implicit Types P : PROP.
Context (bupd : PROP PROP).
Notation "|==> Q" := (bupd Q) : bi_scope.
Hypothesis bupd_intro : P, P |==> P.
Hypothesis bupd_mono : P Q, (P Q) (|==> P) |==> Q.
Hypothesis bupd_trans : P, (|==> |==> P) |==> P.
Hypothesis bupd_frame_r : P R, (|==> P) R |==> (P R).
Context (ident : Type) (saved : ident PROP PROP).
Hypothesis sprop_persistent : i P, Persistent (saved i P).
Hypothesis sprop_alloc_dep :
(P : ident PROP), (|==> i, saved i (P i)).
Hypothesis sprop_agree : i P Q, saved i P saved i Q (P Q).
(** We assume that we cannot update to false. *)
Hypothesis consistency : ¬( |==> False).
Global Instance bupd_mono' : Proper (() ==> ()) bupd.
Proof. intros P Q ?. by apply bupd_mono. Qed.
Global Instance elim_modal_bupd p P Q : ElimModal True p false (|==> P) P (|==> Q) (|==> Q).
Proof.
by rewrite /ElimModal bi.intuitionistically_if_elim
bupd_frame_r bi.wand_elim_r bupd_trans.
Qed.
(** A bad recursive reference: "Assertion with name [i] does not hold" *)
Definition A (i : ident) : PROP := P, ¬ P saved i P.
Lemma A_alloc : |==> i, saved i (A i).
Proof. by apply sprop_alloc_dep. Qed.
Lemma saved_NA i : saved i (A i) ¬ A i.
Proof.
iIntros "#Hs #HA". iPoseProof "HA" as "HA'".
iDestruct "HA'" as (P) "[HNP HsP]". iApply "HNP".
iDestruct (sprop_agree i P (A i) with "[]") as "#[_ HP]".
{ eauto. }
iApply "HP". done.
Qed.
Lemma saved_A i : saved i (A i) A i.
Proof.
iIntros "#Hs". iExists (A i). iFrame "Hs".
iIntros "!>". by iApply saved_NA.
Qed.
Lemma contradiction : False.
Proof using All.
apply consistency.
iMod A_alloc as (i) "#H".
iPoseProof (saved_NA with "H") as "HN".
iApply bupd_intro. iApply "HN". iApply saved_A. done.
Qed.
End savedprop. End savedprop.
(** This proves that we need the ▷ when opening invariants. We have two
paradoxes in this section, but they share the general axiomatization of
invariants. *)
Module inv. Section inv.
Context {PROP : bi} `{!BiAffine PROP}.
Implicit Types P : PROP.
(** Assumptions *)
(** We have the update modality (two classes: empty/full mask) *)
Inductive mask := M0 | M1.
Context (fupd : mask PROP PROP).
Hypothesis fupd_intro : E P, P fupd E P.
Hypothesis fupd_mono : E P Q, (P Q) fupd E P fupd E Q.
Hypothesis fupd_fupd : E P, fupd E (fupd E P) fupd E P.
Hypothesis fupd_frame_l : E P Q, P fupd E Q fupd E (P Q).
Hypothesis fupd_mask_mono : P, fupd M0 P fupd M1 P.
(** We have invariants *)
Context (name : Type) (inv : name PROP PROP).
Global Arguments inv _ _%_I.
Hypothesis inv_persistent : i P, Persistent (inv i P).
Hypothesis inv_alloc : P, P fupd M1 ( i, inv i P).
Hypothesis inv_fupd :
i P Q R, (P Q fupd M0 (P R)) (inv i P Q fupd M1 R).
(** We assume that we cannot update to false. *)
Hypothesis consistency : ¬ ( fupd M1 False).
(** This completes the general assumptions shared by both paradoxes. We set up
some general lemmas and proof mode compatibility before proceeding with
the paradoxes. *)
Lemma inv_fupd' i P R : inv i P (P -∗ fupd M0 (P fupd M1 R)) fupd M1 R.
Proof.
iIntros "(#HiP & HP)". iApply fupd_fupd. iApply inv_fupd; last first.
{ iSplit; first done. iExact "HP". }
iIntros "(HP & HPw)". by iApply "HPw".
Qed.
Global Instance fupd_mono' E : Proper (() ==> ()) (fupd E).
Proof. intros P Q ?. by apply fupd_mono. Qed.
Global Instance fupd_proper E : Proper ((⊣⊢) ==> (⊣⊢)) (fupd E).
Proof.
intros P Q; rewrite !bi.equiv_entails=> -[??]; split; by apply fupd_mono.
Qed.
Lemma fupd_frame_r E P Q : fupd E P Q fupd E (P Q).
Proof. by rewrite comm fupd_frame_l comm. Qed.
Global Instance elim_fupd_fupd p E P Q :
ElimModal True p false (fupd E P) P (fupd E Q) (fupd E Q).
Proof.
by rewrite /ElimModal bi.intuitionistically_if_elim
fupd_frame_r bi.wand_elim_r fupd_fupd.
Qed.
Global Instance elim_fupd0_fupd1 p P Q :
ElimModal True p false (fupd M0 P) P (fupd M1 Q) (fupd M1 Q).
Proof.
by rewrite /ElimModal bi.intuitionistically_if_elim
fupd_frame_r bi.wand_elim_r fupd_mask_mono fupd_fupd.
Qed.
Global Instance exists_split_fupd0 {A} E P (Φ : A PROP) :
FromExist P Φ FromExist (fupd E P) (λ a, fupd E (Φ a)).
Proof.
rewrite /FromExist=>HP. apply bi.exist_elim=> a.
apply fupd_mono. by rewrite -HP -(bi.exist_intro a).
Qed.
(** The original paradox, as found in the "Iris from the Ground Up" paper. *)
Section inv1.
(** On top of invariants themselves, we need a particular kind of ghost state:
we have tokens for a little "two-state STS": [start] -> [finish].
[start] also asserts the exact state; it is only ever owned by the
invariant. [finished] is duplicable. *)
(** Posssible implementations of these axioms:
- Using the STS monoid of a two-state STS, where [start] is the
authoritative saying the state is exactly [start], and [finish]
is the "we are at least in state [finish]" typically owned by threads.
- Ex () +_## ()
*)
Context (gname : Type).
Context (start finished : gname PROP).
Hypothesis sts_alloc : fupd M0 ( γ, start γ).
Hypotheses start_finish : γ, start γ fupd M0 (finished γ).
Hypothesis finished_not_start : γ, start γ finished γ False.
Hypothesis finished_dup : γ, finished γ finished γ finished γ.
(** Now to the actual counterexample. We start with a weird form of saved propositions. *)
Definition saved (γ : gname) (P : PROP) : PROP :=
i, inv i (start γ (finished γ P)).
Global Instance saved_persistent γ P : Persistent (saved γ P) := _.
Lemma saved_alloc (P : gname PROP) : fupd M1 ( γ, saved γ (P γ)).
Proof.
iIntros "". iMod (sts_alloc) as (γ) "Hs".
iMod (inv_alloc (start γ (finished γ (P γ))) with "[Hs]") as (i) "#Hi".
{ auto. }
iApply fupd_intro. by iExists γ, i.
Qed.
Lemma saved_cast γ P Q : saved γ P saved γ Q P fupd M1 ( Q).
Proof.
iIntros "(#HsP & #HsQ & #HP)". iDestruct "HsP" as (i) "HiP".
iApply (inv_fupd' i). iSplit; first done.
iIntros "HaP". iAssert (fupd M0 (finished γ)) with "[HaP]" as "> Hf".
{ iDestruct "HaP" as "[Hs | [Hf _]]".
- by iApply start_finish.
- by iApply fupd_intro. }
iDestruct (finished_dup with "Hf") as "[Hf Hf']".
iApply fupd_intro. iSplitL "Hf'"; first by eauto.
(* Step 2: Open the Q-invariant. *)
iClear (i) "HiP ". iDestruct "HsQ" as (i) "HiQ".
iApply (inv_fupd' i). iSplit; first done.
iIntros "[HaQ | [_ #HQ]]".
{ iExFalso. iApply finished_not_start. by iFrame. }
iApply fupd_intro. iSplitL "Hf".
{ iRight. by iFrame. }
by iApply fupd_intro.
Qed.
(** And now we tie a bad knot. *)
Notation not_fupd P := ( (P -∗ fupd M1 False))%I.
Definition A i : PROP := P, not_fupd P saved i P.
Global Instance A_persistent i : Persistent (A i) := _.
Lemma A_alloc : fupd M1 ( i, saved i (A i)).
Proof. by apply saved_alloc. Qed.
Lemma saved_NA i : saved i (A i) not_fupd (A i).
Proof.
iIntros "#Hi !> #HA". iPoseProof "HA" as "HA'".
iDestruct "HA'" as (P) "#[HNP Hi']".
iMod (saved_cast i (A i) P with "[]") as "HP".
{ eauto. }
by iApply "HNP".
Qed.
Lemma saved_A i : saved i (A i) A i.
Proof.
iIntros "#Hi". iExists (A i). iFrame "#".
by iApply saved_NA.
Qed.
Lemma contradiction : False.
Proof using All.
apply consistency. iIntros "".
iMod A_alloc as (i) "#H".
iPoseProof (saved_NA with "H") as "HN".
iApply "HN". iApply saved_A. done.
Qed.
End inv1.
(** This is another proof showing that we need the ▷ when opening invariants.
Unlike the two paradoxes above, this proof does not rely on impredicative
quantification -- at least, not directly. Instead it exploits the impredicative
quantification that is implicit in [fupd]. Unlike the previous paradox,
the [finish] token needs to be persistent for this paradox to work.
This paradox is due to Yusuke Matsushita. *)
Section inv2.
(** On top of invariants themselves, we need a particular kind of ghost state:
we have tokens for a little "two-state STS": [start] -> [finish].
[start] also asserts the exact state; it is only ever owned by the
invariant. [finished] is persistent. *)
(** Posssible implementations of these axioms:
- Using the STS monoid of a two-state STS, where [start] is the
authoritative saying the state is exactly [start], and [finish]
is the "we are at least in state [finish]" typically owned by threads.
- Ex () +_## ()
*)
Context (gname : Type).
Context (start finished : gname PROP).
Hypothesis sts_alloc : fupd M0 ( γ, start γ).
Hypotheses start_finish : γ, start γ fupd M0 (finished γ).
Hypothesis finished_not_start : γ, start γ finished γ False.
Hypothesis finished_persistent : γ, Persistent (finished γ).
(** Now to the actual counterexample. *)
(** A version of ⊥ behind a persistent update. *)
Definition B : PROP := fupd M1 False.
(** A delayed-initialization invariant storing [B]. *)
Definition P (γ : gname) : PROP := start γ B.
Definition I (i : name) (γ : gname) : PROP := inv i (P γ).
(** If we can ever finish initializing the invariant, we have a
contradiction. *)
Lemma finished_contradiction γ i :
finished γ I i γ -∗ B.
Proof.
iIntros "[#Hfin #HI] !>".
iApply (inv_fupd' i). iSplit; first done.
iIntros "[Hstart|#Hfalse]".
{ iDestruct (finished_not_start with "[$Hfin $Hstart]") as %[]. }
iApply fupd_intro. iSplitR; last done.
by iRight.
Qed.
(** If we can even just create the invariant, we can finish initializing it
using the above lemma, and then get the contradiction. *)
Lemma invariant_contradiction γ i :
I i γ -∗ B.
Proof.
iIntros "#HI !>".
iApply (inv_fupd' i). iSplit; first done.
iIntros "HP".
iAssert (fupd M0 B) with "[HP]" as ">#Hfalse".
{ iDestruct "HP" as "[Hstart|#Hfalse]"; last by iApply fupd_intro.
iMod (start_finish with "Hstart"). iApply fupd_intro.
(** There's a magic moment here where we have the invariant open,
but inside [finished_contradiction] we will be proving
a [fupd M1] and so we can open the invariant *again*.
Really we are just building up a thunk that can be used
later when the invariant is closed again. But to build up that
thunk we can use resources that we just got out of the invariant,
before closing it again. *)
iApply finished_contradiction. eauto. }
iApply fupd_intro. iSplitR; last done.
by iRight.
Qed.
(** Of course, creating the invariant is trivial. *)
Lemma contradiction' : False.
Proof using All.
apply consistency.
iMod sts_alloc as (γ) "Hstart".
iMod (inv_alloc (P γ) with "[Hstart]") as (i) "HI".
{ by iLeft. }
iDestruct (invariant_contradiction with "HI") as "#>[]".
Qed.
End inv2.
End inv. End inv.
(** This proves that if we have linear impredicative invariants, we can still
drop arbitrary resources (i.e., we can "defeat" linearity).
We assume [cinv_alloc] without any bells or whistles.
Moreover, we also have an accessor that gives back the invariant token
immediately, not just after the invariant got closed again.
The assumptions here match the proof rules in Iron, save for the side-condition
that the invariant must be "uniform". In particular, [cinv_alloc] delays
handing out the [cinv_own] token until after the invariant has been created so
that this can match Iron by picking [cinv_own γ := fcinv_own γ 1 ∗
fcinv_cancel_own γ 1]. This means [cinv_own] is not "uniform" in Iron terms,
which is why Iron does not suffer from this contradiction.
This also loosely matches VST's locks with stored resource invariants.
There, the stronger variant of the "unlock" rule (see Aquinas Hobor's PhD thesis
"Oracle Semantics", §4.7, p. 88) also permits putting the token of a lock
entirely into that lock.
*)
Module linear. Section linear.
Context {PROP: bi}.
Implicit Types P : PROP.
(** Assumptions. *)
(** We have the mask-changing update modality (two classes: empty/full mask) *)
Inductive mask := M0 | M1.
Context (fupd : mask mask PROP PROP).
Hypothesis fupd_intro : E P, P fupd E E P.
Hypothesis fupd_mono : E1 E2 P Q, (P Q) fupd E1 E2 P fupd E1 E2 Q.
Hypothesis fupd_fupd : E1 E2 E3 P, fupd E1 E2 (fupd E2 E3 P) fupd E1 E3 P.
Hypothesis fupd_frame_l : E1 E2 P Q, P fupd E1 E2 Q fupd E1 E2 (P Q).
(** We have cancelable invariants. (Really they would have fractions at
[cinv_own] but we do not need that. They would also have a name matching
the [mask] type, but we do not need that either.) *)
Context (gname : Type) (cinv : gname PROP PROP) (cinv_own : gname PROP).
Hypothesis cinv_alloc : E P,
P -∗ fupd E E ( γ, cinv γ P cinv_own γ).
Hypothesis cinv_acc : P γ,
cinv γ P -∗ cinv_own γ -∗ fupd M1 M0 ( P cinv_own γ ( P -∗ fupd M0 M1 emp)).
(** Some general lemmas and proof mode compatibility. *)
Global Instance fupd_mono' E1 E2 : Proper (() ==> ()) (fupd E1 E2).
Proof. intros P Q ?. by apply fupd_mono. Qed.
Global Instance fupd_proper E1 E2 : Proper ((⊣⊢) ==> (⊣⊢)) (fupd E1 E2).
Proof.
intros P Q; rewrite !bi.equiv_entails=> -[??]; split; by apply fupd_mono.
Qed.
Lemma fupd_frame_r E1 E2 P Q : fupd E1 E2 P Q fupd E1 E2 (P Q).
Proof. by rewrite comm fupd_frame_l comm. Qed.
Global Instance elim_fupd_fupd p E1 E2 E3 P Q :
ElimModal True p false (fupd E1 E2 P) P (fupd E1 E3 Q) (fupd E2 E3 Q).
Proof.
by rewrite /ElimModal bi.intuitionistically_if_elim
fupd_frame_r bi.wand_elim_r fupd_fupd.
Qed.
(** Counterexample: now we can make any resource disappear. *)
Lemma leak P : P -∗ fupd M1 M1 emp.
Proof.
iIntros "HP".
iMod (cinv_alloc _ True with "[//]") as (γ) "[Hinv Htok]".
iMod (cinv_acc with "Hinv Htok") as "(Htrue & Htok & Hclose)".
iApply "Hclose". done.
Qed.
End linear. End linear.
(** Various lemmas showing that [fixpoint] is closed under the BI properties.
We only prove these lemmas for the unary case; if you need them for an
n-ary function, it is probably easier to copy these proofs than to try and apply
these lemmas via (un)currying. *)
From iris.bi Require Export bi.
From iris.proofmode Require Import proofmode.
From iris.prelude Require Import options.
Import bi.
Section fixpoint_laws.
Context {PROP: bi}.
Implicit Types P Q : PROP.
Lemma fixpoint_plain `{!BiPlainly PROP} {A}
(F : (A -d> PROP) A -d> PROP) `{!Contractive F} :
( Φ, ( x, Plain (Φ x)) ( x, Plain (F Φ x)))
x, Plain (fixpoint F x).
Proof.
intros ?. apply fixpoint_ind.
- intros Φ1 Φ2 ??. by rewrite -( _).
- exists (λ _, emp%I); apply _.
- done.
- apply limit_preserving_forall=> x.
apply limit_preserving_Plain; solve_proper.
Qed.
Lemma fixpoint_persistent {A} (F : (A -d> PROP) A -d> PROP) `{!Contractive F} :
( Φ, ( x, Persistent (Φ x)) ( x, Persistent (F Φ x)))
x, Persistent (fixpoint F x).
Proof.
intros ?. apply fixpoint_ind.
- intros Φ1 Φ2 ??. by rewrite -( _).
- exists (λ _, emp%I); apply _.
- done.
- apply limit_preserving_forall=> x.
apply limit_preserving_Persistent; solve_proper.
Qed.
Lemma fixpoint_absorbing {A} (F : (A -d> PROP) A -d> PROP) `{!Contractive F} :
( Φ, ( x, Absorbing (Φ x)) ( x, Absorbing (F Φ x)))
x, Absorbing (fixpoint F x).
Proof.
intros ?. apply fixpoint_ind.
- intros Φ1 Φ2 ??. by rewrite -( _).
- exists (λ _, True%I); apply _.
- done.
- apply limit_preserving_forall=> x.
apply limit_preserving_entails; solve_proper.
Qed.
Lemma fixpoint_affine {A} (F : (A -d> PROP) A -d> PROP) `{!Contractive F} :
( Φ, ( x, Affine (Φ x)) ( x, Affine (F Φ x)))
x, Affine (fixpoint F x).
Proof.
intros ?. apply fixpoint_ind.
- intros Φ1 Φ2 ??. by rewrite -( _).
- exists (λ _, emp%I); apply _.
- done.
- apply limit_preserving_forall=> x.
apply limit_preserving_entails; solve_proper.
Qed.
Lemma fixpoint_persistent_absoring {A}
(F : (A -d> PROP) A -d> PROP) `{!Contractive F} :
( Φ, ( x, Persistent (Φ x)) ( x, Absorbing (Φ x))
( x, Persistent (F Φ x) Absorbing (F Φ x)))
x, Persistent (fixpoint F x) Absorbing (fixpoint F x).
Proof.
intros ?. apply fixpoint_ind.
- intros Φ1 Φ2 ??. by rewrite -( _).
- exists (λ _, True%I); split; apply _.
- naive_solver.
- apply limit_preserving_forall=> x.
apply limit_preserving_and; apply limit_preserving_entails; solve_proper.
Qed.
Lemma fixpoint_persistent_affine {A}
(F : (A -d> PROP) A -d> PROP) `{!Contractive F} :
( Φ, ( x, Persistent (Φ x)) ( x, Affine (Φ x))
( x, Persistent (F Φ x) Affine (F Φ x)))
x, Persistent (fixpoint F x) Affine (fixpoint F x).
Proof.
intros ?. apply fixpoint_ind.
- intros Φ1 Φ2 ??. by rewrite -( _).
- exists (λ _, emp%I); split; apply _.
- naive_solver.
- apply limit_preserving_forall=> x.
apply limit_preserving_and; apply limit_preserving_entails; solve_proper.
Qed.
Lemma fixpoint_plain_absoring `{!BiPlainly PROP} {A}
(F : (A -d> PROP) A -d> PROP) `{!Contractive F} :
( Φ, ( x, Plain (Φ x)) ( x, Absorbing (Φ x))
( x, Plain (F Φ x) Absorbing (F Φ x)))
x, Plain (fixpoint F x) Absorbing (fixpoint F x).
Proof.
intros ?. apply fixpoint_ind.
- intros Φ1 Φ2 ??. by rewrite -( _).
- exists (λ _, True%I); split; apply _.
- naive_solver.
- apply limit_preserving_forall=> x.
apply limit_preserving_and; apply limit_preserving_entails; solve_proper.
Qed.
Lemma fixpoint_plain_affine `{!BiPlainly PROP} {A}
(F : (A -d> PROP) A -d> PROP) `{!Contractive F} :
( Φ, ( x, Plain (Φ x)) ( x, Affine (Φ x))
( x, Plain (F Φ x) Affine (F Φ x)))
x, Plain (fixpoint F x) Affine (fixpoint F x).
Proof.
intros ?. apply fixpoint_ind.
- intros Φ1 Φ2 ??. by rewrite -( _).
- exists (λ _, emp%I); split; apply _.
- naive_solver.
- apply limit_preserving_forall=> x.
apply limit_preserving_and; apply limit_preserving_entails; solve_proper.
Qed.
End fixpoint_laws.
From iris.bi Require Export bi.
From iris.proofmode Require Import proofmode.
From iris.prelude Require Import options.
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 Φ Ψ :
NonExpansive Φ
NonExpansive Ψ
( x, Φ x -∗ Ψ x) -∗ x, F Φ x -∗ F Ψ x;
bi_mono_pred_ne Φ : NonExpansive Φ NonExpansive (F Φ)
}.
Global Arguments bi_mono_pred {_ _ _ _} _ _.
Local Existing Instance bi_mono_pred_ne.
Definition bi_least_fixpoint {PROP : bi} {A : ofe}
(F : (A PROP) (A PROP)) (x : A) : PROP :=
tc_opaque ( Φ : A -n> PROP, ( x, F Φ x -∗ Φ x) -∗ Φ x)%I.
Global Arguments bi_least_fixpoint : simpl never.
Definition bi_greatest_fixpoint {PROP : bi} {A : ofe}
(F : (A PROP) (A PROP)) (x : A) : PROP :=
tc_opaque ( Φ : A -n> PROP, ( x, Φ x -∗ F Φ x) Φ x)%I.
Global Arguments bi_greatest_fixpoint : simpl never.
(* Both non-expansiveness lemmas do not seem to be interderivable.
FIXME: is there some lemma that subsumes both? *)
Lemma least_fixpoint_ne' {PROP : bi} {A : ofe} (F : (A PROP) (A PROP)):
( Φ, NonExpansive Φ NonExpansive (F Φ)) NonExpansive (bi_least_fixpoint F).
Proof. solve_proper. Qed.
Global Instance least_fixpoint_ne {PROP : bi} {A : ofe} n :
Proper (pointwise_relation (A PROP) (pointwise_relation A (dist n)) ==>
dist n ==> dist n) bi_least_fixpoint.
Proof. solve_proper. Qed.
Global Instance least_fixpoint_proper {PROP : bi} {A : ofe} :
Proper (pointwise_relation (A PROP) (pointwise_relation A ()) ==>
() ==> ()) bi_least_fixpoint.
Proof. solve_proper. Qed.
Section least.
Context {PROP : bi} {A : ofe} (F : (A PROP) (A PROP)) `{!BiMonoPred F}.
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 "[#] HF"); [solve_proper|].
iIntros "!>" (y) "Hy". iApply ("Hy" with "[# //]").
Qed.
Lemma least_fixpoint_unfold_1 x :
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 "[#] Hy").
iIntros "!>" (z) "?". by iApply least_fixpoint_unfold_2.
Qed.
Corollary least_fixpoint_unfold x :
bi_least_fixpoint F x F (bi_least_fixpoint F) x.
Proof using Type*.
apply (anti_symm _); auto using least_fixpoint_unfold_1, least_fixpoint_unfold_2.
Qed.
(**
The basic induction principle for least fixpoints: as inductive hypothesis,
it allows to assume the statement to prove below exactly one application of [F].
*)
Lemma least_fixpoint_iter (Φ : A PROP) `{!NonExpansive Φ} :
( y, F Φ y -∗ Φ y) -∗ x, bi_least_fixpoint F x -∗ Φ x.
Proof.
iIntros "#HΦ" (x) "HF". by iApply ("HF" $! (OfeMor Φ) with "[#]").
Qed.
Lemma least_fixpoint_affine :
( x, Affine (F (λ _, emp%I) x))
x, Affine (bi_least_fixpoint F x).
Proof.
intros ?. rewrite /Affine. iApply least_fixpoint_iter.
by iIntros "!> %y HF".
Qed.
Lemma least_fixpoint_absorbing :
( Φ, ( x, Absorbing (Φ x)) ( x, Absorbing (F Φ x)))
x, Absorbing (bi_least_fixpoint F x).
Proof using Type*.
intros ? x. rewrite /Absorbing /bi_absorbingly.
apply wand_elim_r'. revert x.
iApply least_fixpoint_iter; first solve_proper.
iIntros "!> %y HF Htrue". iApply least_fixpoint_unfold.
iAssert (F (λ x : A, True -∗ bi_least_fixpoint F x) y)%I with "[-]" as "HF".
{ by iClear "Htrue". }
iApply (bi_mono_pred with "[] HF"); first solve_proper.
iIntros "!> %x HF". by iApply "HF".
Qed.
Lemma least_fixpoint_persistent_affine :
( Φ, ( x, Affine (Φ x)) ( x, Affine (F Φ x)))
( Φ, ( x, Persistent (Φ x)) ( x, Persistent (F Φ x)))
x, Persistent (bi_least_fixpoint F x).
Proof using Type*.
intros ?? x. rewrite /Persistent -intuitionistically_into_persistently_1.
revert x. iApply least_fixpoint_iter; first solve_proper.
iIntros "!> %y #HF !>". iApply least_fixpoint_unfold.
iApply (bi_mono_pred with "[] HF"); first solve_proper.
by iIntros "!> %x #?".
Qed.
Lemma least_fixpoint_persistent_absorbing :
( Φ, ( x, Absorbing (Φ x)) ( x, Absorbing (F Φ x)))
( Φ, ( x, Persistent (Φ x)) ( x, Persistent (F Φ x)))
x, Persistent (bi_least_fixpoint F x).
Proof using Type*.
intros ??. pose proof (least_fixpoint_absorbing _). unfold Persistent.
iApply least_fixpoint_iter; first solve_proper.
iIntros "!> %y #HF !>". rewrite least_fixpoint_unfold.
iApply (bi_mono_pred with "[] HF"); first solve_proper.
by iIntros "!> %x #?".
Qed.
End least.
Lemma least_fixpoint_strong_mono
{PROP : bi} {A : ofe} (F : (A PROP) (A PROP)) `{!BiMonoPred F}
(G : (A PROP) (A PROP)) `{!BiMonoPred G} :
( Φ x, F Φ x -∗ G Φ x) -∗
x, bi_least_fixpoint F x -∗ bi_least_fixpoint G x.
Proof.
iIntros "#Hmon". iApply least_fixpoint_iter.
iIntros "!>" (y) "IH". iApply least_fixpoint_unfold.
by iApply "Hmon".
Qed.
(** In addition to [least_fixpoint_iter], we provide two derived, stronger
induction principles:
- [least_fixpoint_ind] allows to assume [F (λ x, Φ x ∧ bi_least_fixpoint F x) y]
when proving the inductive step.
Intuitively, it does not only offer the induction hypothesis ([Φ] under an
application of [F]), but also the induction predicate [bi_least_fixpoint F]
itself (under an application of [F]).
- [least_fixpoint_ind_wf] intuitively corresponds to a kind of well-founded
induction. It provides the hypothesis [F (bi_least_fixpoint (λ Ψ a, Φ a ∧ F Ψ a)) y]
and thus allows to assume the induction hypothesis not just below one
application of [F], but below any positive number (by unfolding the least
fixpoint). The unfolding lemma [least_fixpoint_unfold] as well as
[least_fixpoint_strong_mono] can be useful to work with the hypothesis. *)
Section least_ind.
Context {PROP : bi} {A : ofe} (F : (A PROP) (A PROP)) `{!BiMonoPred F}.
Local Lemma Private_wf_pred_mono `{!NonExpansive Φ} :
BiMonoPred (λ (Ψ : A PROP) (a : A), Φ a F Ψ a)%I.
Proof using Type*.
split; last solve_proper.
intros Ψ Ψ' Hne Hne'. iIntros "#Mon" (x) "Ha". iSplit; first by iDestruct "Ha" as "[$ _]".
iDestruct "Ha" as "[_ Hr]". iApply (bi_mono_pred with "[] Hr"). by iModIntro.
Qed.
Local Existing Instance Private_wf_pred_mono.
Lemma least_fixpoint_ind_wf (Φ : A PROP) `{!NonExpansive Φ} :
( y, F (bi_least_fixpoint (λ Ψ a, Φ a F Ψ a)) y -∗ Φ y) -∗
x, bi_least_fixpoint F x -∗ Φ x.
Proof using Type*.
iIntros "#Hmon" (x). rewrite least_fixpoint_unfold. iIntros "Hx".
iApply "Hmon". iApply (bi_mono_pred with "[] Hx").
iModIntro. iApply least_fixpoint_iter.
iIntros "!> %y Hy". rewrite least_fixpoint_unfold.
iSplit; last done. by iApply "Hmon".
Qed.
Lemma least_fixpoint_ind (Φ : A PROP) `{!NonExpansive Φ} :
( y, F (λ x, Φ x bi_least_fixpoint F x) y -∗ Φ y) -∗
x, bi_least_fixpoint F x -∗ Φ x.
Proof using Type*.
iIntros "#Hmon". iApply least_fixpoint_ind_wf. iIntros "!> %y Hy".
iApply "Hmon". iApply (bi_mono_pred with "[] Hy"). { solve_proper. }
iIntros "!> %x Hx". iSplit.
- rewrite least_fixpoint_unfold. iDestruct "Hx" as "[$ _]".
- iApply (least_fixpoint_strong_mono with "[] Hx"). iIntros "!>" (??) "[_ $]".
Qed.
End least_ind.
Lemma greatest_fixpoint_ne_outer {PROP : bi} {A : ofe}
(F1 : (A PROP) (A PROP)) (F2 : (A PROP) (A PROP)):
( Φ x n, F1 Φ x {n} F2 Φ x) x1 x2 n,
x1 {n} x2 bi_greatest_fixpoint F1 x1 {n} bi_greatest_fixpoint F2 x2.
Proof.
intros HF x1 x2 n Hx. rewrite /bi_greatest_fixpoint /=.
do 3 f_equiv; last solve_proper. repeat f_equiv. apply HF.
Qed.
(* Both non-expansiveness lemmas do not seem to be interderivable.
FIXME: is there some lemma that subsumes both? *)
Lemma greatest_fixpoint_ne' {PROP : bi} {A : ofe} (F : (A PROP) (A PROP)):
( Φ, NonExpansive Φ NonExpansive (F Φ)) NonExpansive (bi_greatest_fixpoint F).
Proof. solve_proper. Qed.
Global Instance greatest_fixpoint_ne {PROP : bi} {A : ofe} n :
Proper (pointwise_relation (A PROP) (pointwise_relation A (dist n)) ==>
dist n ==> dist n) bi_greatest_fixpoint.
Proof. solve_proper. Qed.
Global Instance greatest_fixpoint_proper {PROP : bi} {A : ofe} :
Proper (pointwise_relation (A PROP) (pointwise_relation A ()) ==>
() ==> ()) bi_greatest_fixpoint.
Proof. solve_proper. Qed.
Section greatest.
Context {PROP : bi} {A : ofe} (F : (A PROP) (A PROP)) `{!BiMonoPred F}.
Lemma greatest_fixpoint_unfold_1 x :
bi_greatest_fixpoint F x F (bi_greatest_fixpoint F) x.
Proof using Type*.
iDestruct 1 as (Φ) "[#Hincl HΦ]".
iApply (bi_mono_pred Φ (bi_greatest_fixpoint F) with "[#]").
- iIntros "!>" (y) "Hy". iExists Φ. auto.
- by iApply "Hincl".
Qed.
Lemma greatest_fixpoint_unfold_2 x :
F (bi_greatest_fixpoint F) x bi_greatest_fixpoint F x.
Proof using Type*.
iIntros "HF". iExists (OfeMor (F (bi_greatest_fixpoint F))).
iSplit; last done. iIntros "!>" (y) "Hy /=". iApply (bi_mono_pred with "[#] Hy").
iIntros "!>" (z) "?". by iApply greatest_fixpoint_unfold_1.
Qed.
Corollary greatest_fixpoint_unfold x :
bi_greatest_fixpoint F x F (bi_greatest_fixpoint F) x.
Proof using Type*.
apply (anti_symm _); auto using greatest_fixpoint_unfold_1, greatest_fixpoint_unfold_2.
Qed.
(**
The following lemma provides basic coinduction capabilities,
by requiring to reestablish the coinduction hypothesis after exactly one step.
*)
Lemma greatest_fixpoint_coiter (Φ : A PROP) `{!NonExpansive Φ} :
( y, Φ y -∗ F Φ y) -∗ x, Φ x -∗ bi_greatest_fixpoint F x.
Proof. iIntros "#HΦ" (x) "Hx". iExists (OfeMor Φ). auto. Qed.
Lemma greatest_fixpoint_absorbing :
( Φ, ( x, Absorbing (Φ x)) ( x, Absorbing (F Φ x)))
x, Absorbing (bi_greatest_fixpoint F x).
Proof using Type*.
intros ?. rewrite /Absorbing.
iApply greatest_fixpoint_coiter; first solve_proper.
iIntros "!> %y >HF". iDestruct (greatest_fixpoint_unfold with "HF") as "HF".
iApply (bi_mono_pred with "[] HF"); first solve_proper.
by iIntros "!> %x HF !>".
Qed.
End greatest.
Lemma greatest_fixpoint_strong_mono {PROP : bi} {A : ofe}
(F : (A PROP) (A PROP)) `{!BiMonoPred F}
(G : (A PROP) (A PROP)) `{!BiMonoPred G} :
( Φ x, F Φ x -∗ G Φ x) -∗
x, bi_greatest_fixpoint F x -∗ bi_greatest_fixpoint G x.
Proof using Type*.
iIntros "#Hmon". iApply greatest_fixpoint_coiter.
iIntros "!>" (y) "IH". rewrite greatest_fixpoint_unfold.
by iApply "Hmon".
Qed.
(** In addition to [greatest_fixpoint_coiter], we provide two derived, stronger
coinduction principles:
- [greatest_fixpoint_coind] requires to prove
[F (λ x, Φ x ∨ bi_greatest_fixpoint F x) y] in the coinductive step instead of
[F Φ y] and thus instead allows to prove the original fixpoint again, after
taking one step.
- [greatest_fixpoint_paco] allows for so-called parameterized coinduction, a
stronger coinduction principle, where [F (bi_greatest_fixpoint (λ Ψ a, Φ a ∨ F Ψ a)) y]
needs to be established in the coinductive step. It allows to prove the
hypothesis [Φ] not just after one step, but after any positive number of
unfoldings of the greatest fixpoint. This encodes a way of accumulating
"knowledge" in the coinduction hypothesis: if you return to the "initial
point" [Φ] of the coinduction after some number of unfoldings (not just one),
the proof is done. (Interestingly, this is the dual to [least_fixpoint_ind_wf]).
The unfolding lemma [greatest_fixpoint_unfold] and
[greatest_fixpoint_strong_mono] may be useful when using this lemma.
*Example use case:*
Suppose that [F] defines a coinductive simulation relation, e.g.,
[F rec '(e_t, e_s) :=
(is_val e_s ∧ is_val e_t ∧ post e_t e_s) ∨
(safe e_t ∧ ∀ e_t', step e_t e_t' → ∃ e_s', step e_s e_s' ∧ rec e_t' e_s')],
and [sim e_t e_s := bi_greatest_fixpoint F].
Suppose you want to show a simulation of two loops,
[sim (while ...) (while ...)],
i.e., [Φ '(e_t, e_s) := e_t = while ... ∧ e_s = while ...].
Then the standard coinduction principle [greatest_fixpoint_iter] requires to
establish the coinduction hypothesis [Φ] after precisely one unfolding of [sim],
which is clearly not strong enough if the loop takes multiple steps of
computation per iteration. But [greatest_fixpoint_paco] allows to establish a
fixpoint to which [Φ] has been added as a disjunct. This fixpoint can be
unfolded arbitrarily many times, allowing to establish the coinduction
hypothesis after any number of steps. This enables to take multiple simulation
steps, before closing the coinduction by establishing the hypothesis [Φ]
again. *)
Section greatest_coind.
Context {PROP : bi} {A : ofe} (F : (A PROP) (A PROP)) `{!BiMonoPred F}.
Local Lemma Private_paco_mono `{!NonExpansive Φ} :
BiMonoPred (λ (Ψ : A PROP) (a : A), Φ a F Ψ a)%I.
Proof using Type*.
split; last solve_proper.
intros Ψ Ψ' Hne Hne'. iIntros "#Mon" (x) "[H1|H2]"; first by iLeft.
iRight. iApply (bi_mono_pred with "[] H2"). by iModIntro.
Qed.
Local Existing Instance Private_paco_mono.
Lemma greatest_fixpoint_paco (Φ : A PROP) `{!NonExpansive Φ} :
( y, Φ y -∗ F (bi_greatest_fixpoint (λ Ψ a, Φ a F Ψ a)) y) -∗
x, Φ x -∗ bi_greatest_fixpoint F x.
Proof using Type*.
iIntros "#Hmon" (x) "HΦ". iDestruct ("Hmon" with "HΦ") as "HF".
rewrite greatest_fixpoint_unfold. iApply (bi_mono_pred with "[] HF").
iIntros "!>" (y) "HG". iApply (greatest_fixpoint_coiter with "[] HG").
iIntros "!>" (z) "Hf". rewrite greatest_fixpoint_unfold.
iDestruct "Hf" as "[HΦ|$]". by iApply "Hmon".
Qed.
Lemma greatest_fixpoint_coind (Φ : A PROP) `{!NonExpansive Φ} :
( y, Φ y -∗ F (λ x, Φ x bi_greatest_fixpoint F x) y) -∗
x, Φ x -∗ bi_greatest_fixpoint F x.
Proof using Type*.
iIntros "#Ha". iApply greatest_fixpoint_paco. iModIntro.
iIntros (y) "Hy". iSpecialize ("Ha" with "Hy").
iApply (bi_mono_pred with "[] Ha"). { solve_proper. }
iIntros "!> %x [Hphi | Hgfp]".
- iApply greatest_fixpoint_unfold. eauto.
- iApply (greatest_fixpoint_strong_mono with "[] Hgfp"); eauto.
Qed.
End greatest_coind.
From iris.bi Require Export bi.
From iris.proofmode Require Import classes classes_make proofmode.
From iris.prelude Require Import options.
Class Fractional {PROP : bi} (Φ : Qp PROP) :=
fractional p q : Φ (p + q)%Qp ⊣⊢ Φ p Φ q.
Global Arguments Fractional {_} _%_I : simpl never.
Global Arguments fractional {_ _ _} _ _.
(** The [AsFractional] typeclass eta-expands a proposition [P] into [Φ q] such
that [Φ] is a fractional predicate. This is needed because higher-order
unification cannot be relied upon to guess the right [Φ].
[AsFractional] should generally be used in APIs that work with fractional
predicates (instead of [Fractional]): when the user provides the original
resource [P], have a premise [AsFractional P Φ 1] to convert that into some
fractional predicate.
The equivalence in [as_fractional] should hold definitionally; various typeclass
instances assume that [Φ q] will un-do the eta-expansion performed by
[AsFractional]. *)
Class AsFractional {PROP : bi} (P : PROP) (Φ : Qp PROP) (q : Qp) := {
as_fractional : P ⊣⊢ Φ q;
as_fractional_fractional : Fractional Φ
}.
Global Arguments AsFractional {_} _%_I _%_I _%_Qp.
Global Hint Mode AsFractional - ! - - : typeclass_instances.
(** The class [FrameFractionalQp] is used for fractional framing, it subtracts
the fractional of the hypothesis from the goal: it computes [r := qP - qR].
See [frame_fractional] for how it is used. *)
Class FrameFractionalQp (qR qP r : Qp) :=
frame_fractional_qp : qP = (qR + r)%Qp.
Global Hint Mode FrameFractionalQp ! ! - : typeclass_instances.
Section fractional.
Context {PROP : bi}.
Implicit Types P Q : PROP.
Implicit Types Φ : Qp PROP.
Implicit Types q : Qp.
Global Instance Fractional_proper :
Proper (pointwise_relation _ () ==> iff) (@Fractional PROP).
Proof.
rewrite /Fractional.
intros Φ1 Φ2 Hequiv.
by setoid_rewrite Hequiv.
Qed.
(* Every [Fractional] predicate admits an obvious [AsFractional] instance.
Ideally, this instance would mean that a user never has to define a manual
[AsFractional] instance for a [Fractional] predicate (even if it's of the
form [λ q, Φ a1 ‥ q ‥ an] for some n-ary predicate [Φ].) However, Coq's
lack of guarantees for higher-order unification mean this instance wouldn't
guarantee to apply for every [AsFractional] goal.
Therefore, this instance is not global to avoid conflicts with existing instances
defined by our users, since we can't ask users to universally delete their
manually-defined [AsFractional] instances for their own [Fractional] predicates.
(We could just support this instance for predicates with the fractional argument
in the final position, but that was felt to be a bit of a foot-gun - users would
have to remember to *not* define an [AsFractional] some of the time.) *)
Local Instance fractional_as_fractional Φ q :
Fractional Φ AsFractional (Φ q) Φ q.
Proof. done. Qed.
(** This lemma is meant to be used when [P] is known. But really you should be
using [iDestruct "H" as "[H1 H2]"], which supports splitting at fractions. *)
Lemma fractional_split P Φ q1 q2 :
AsFractional P Φ (q1 + q2)
P ⊣⊢ Φ q1 Φ q2.
Proof. by move=>-[-> ->]. Qed.
(** This lemma is meant to be used when [P] is known. But really you should be
using [iDestruct "H" as "[H1 H2]"], which supports halving fractions. *)
Lemma fractional_half P Φ q :
AsFractional P Φ q
P ⊣⊢ Φ (q/2)%Qp Φ (q/2)%Qp.
Proof. by rewrite -{1}(Qp.div_2 q)=>-[->->]. Qed.
(** This lemma is meant to be used when [P1], [P2] are known. But really you
should be using [iCombine "H1 H2" as "H"] (for forwards reasoning) or
[iSplitL]/[iSplitR] (for goal-oriented reasoning), which support merging
fractions. *)
Lemma fractional_merge P1 P2 Φ q1 q2 `{!Fractional Φ} :
AsFractional P1 Φ q1 AsFractional P2 Φ q2
P1 P2 ⊣⊢ Φ (q1 + q2)%Qp.
Proof. move=>-[-> _] [-> _]. rewrite fractional //. Qed.
(** Fractional and logical connectives *)
Global Instance persistent_fractional (P : PROP) :
Persistent P TCOr (Affine P) (Absorbing P) Fractional (λ _, P).
Proof. intros ?? q q'. apply: bi.persistent_sep_dup. Qed.
(** We do not have [AsFractional] instances for [∗] and the big operators
because the [iDestruct] tactic already turns [P ∗ Q] into [P] and [Q],
[[∗ list] k↦x ∈ y :: l, Φ k x] into [Φ 0 i] and [[∗ list] k↦x ∈ l, Φ (S k) x],
etc. Hence, an [AsFractional] instance would cause ambiguity because for
example [l ↦{1} v ∗ l' ↦{1} v'] could be turned into [l ↦{1} v] and
[l' ↦{1} v'], or into two times [l ↦{1/2} v ∗ l' ↦{1/2} v'].
We do provide the [Fractional] instances so that when one defines a derived
connection in terms of [∗] or a big operator (and makes that opaque in some
way), one could choose to split along the [∗] or along the fraction. *)
Global Instance fractional_sep Φ Ψ :
Fractional Φ Fractional Ψ Fractional (λ q, Φ q Ψ q)%I.
Proof.
intros ?? q q'. rewrite !fractional -!assoc. f_equiv.
rewrite !assoc. f_equiv. by rewrite comm.
Qed.
Global Instance fractional_embed `{!BiEmbed PROP PROP'} Φ :
Fractional Φ Fractional (λ q, Φ q : PROP')%I.
Proof. intros ???. by rewrite fractional embed_sep. Qed.
Global Instance as_fractional_embed `{!BiEmbed PROP PROP'} P Φ q :
AsFractional P Φ q AsFractional ( P ) (λ q, Φ q )%I q.
Proof. intros [??]; split; [by f_equiv|apply _]. Qed.
Global Instance fractional_big_sepL {A} (l : list A) Ψ :
( k x, Fractional (Ψ k x))
Fractional (PROP:=PROP) (λ q, [ list] kx l, Ψ k x q)%I.
Proof. intros ? q q'. rewrite -big_sepL_sep. by setoid_rewrite fractional. Qed.
Global Instance fractional_big_sepL2 {A B} (l1 : list A) (l2 : list B) Ψ :
( k x1 x2, Fractional (Ψ k x1 x2))
Fractional (PROP:=PROP) (λ q, [ list] kx1; x2 l1; l2, Ψ k x1 x2 q)%I.
Proof. intros ? q q'. rewrite -big_sepL2_sep. by setoid_rewrite fractional. Qed.
Global Instance fractional_big_sepM `{Countable K} {A} (m : gmap K A) Ψ :
( k x, Fractional (Ψ k x))
Fractional (PROP:=PROP) (λ q, [ map] kx m, Ψ k x q)%I.
Proof. intros ? q q'. rewrite -big_sepM_sep. by setoid_rewrite fractional. Qed.
Global Instance fractional_big_sepS `{Countable A} (X : gset A) Ψ :
( x, Fractional (Ψ x))
Fractional (PROP:=PROP) (λ q, [ set] x X, Ψ x q)%I.
Proof. intros ? q q'. rewrite -big_sepS_sep. by setoid_rewrite fractional. Qed.
Global Instance fractional_big_sepMS `{Countable A} (X : gmultiset A) Ψ :
( x, Fractional (Ψ x))
Fractional (PROP:=PROP) (λ q, [ mset] x X, Ψ x q)%I.
Proof. intros ? q q'. rewrite -big_sepMS_sep. by setoid_rewrite fractional. Qed.
(** Proof mode instances *)
Global Instance from_sep_fractional P Φ q1 q2 :
AsFractional P Φ (q1 + q2) FromSep P (Φ q1) (Φ q2).
Proof. rewrite /FromSep=>-[-> ->] //. Qed.
Global Instance combine_sep_as_fractional P1 P2 Φ q1 q2 :
AsFractional P1 Φ q1 AsFractional P2 Φ q2
CombineSepAs P1 P2 (Φ (q1 + q2)%Qp) | 50.
(* Explicit cost, to make it easier to provide instances with higher or
lower cost. Higher-cost instances exist to combine (for example)
[l ↦{q1} v1] with [l ↦{q2} v2] where [v1] and [v2] are not unifiable. *)
Proof. rewrite /CombineSepAs =>-[-> _] [-> <-] //. Qed.
Global Instance from_sep_fractional_half P Φ q :
AsFractional P Φ q FromSep P (Φ (q / 2)%Qp) (Φ (q / 2)%Qp) | 10.
Proof. rewrite /FromSep -{1}(Qp.div_2 q). intros [-> <-]. rewrite Qp.div_2 //. Qed.
Global Instance combine_sep_as_fractional_half P Φ q :
AsFractional P Φ (q/2) CombineSepAs P P (Φ q) | 50.
(* Explicit cost, to make it easier to provide instances with higher or
lower cost. Higher-cost instances exist to combine (for example)
[l ↦{q1} v1] with [l ↦{q2} v2] where [v1] and [v2] are not unifiable. *)
Proof. rewrite /CombineSepAs=>-[-> <-]. by rewrite Qp.div_2. Qed.
Global Instance into_sep_fractional P Φ q1 q2 :
AsFractional P Φ (q1 + q2) IntoSep P (Φ q1) (Φ q2).
Proof. intros [??]. rewrite /IntoSep [P]fractional_split //. Qed.
Global Instance into_sep_fractional_half P Φ q :
AsFractional P Φ q IntoSep P (Φ (q / 2)%Qp) (Φ (q / 2)%Qp) | 100.
Proof. intros [??]. rewrite /IntoSep [P]fractional_half //. Qed.
Global Instance frame_fractional_qp_add_l q q' : FrameFractionalQp q (q + q') q'.
Proof. by rewrite /FrameFractionalQp. Qed.
Global Instance frame_fractional_qp_add_r q q' : FrameFractionalQp q' (q + q') q.
Proof. by rewrite /FrameFractionalQp Qp.add_comm. Qed.
Global Instance frame_fractional_qp_half q : FrameFractionalQp (q/2) q (q/2).
Proof. by rewrite /FrameFractionalQp Qp.div_2. Qed.
(* Not an instance because of performance; you can locally add it if you are
willing to pay the cost. We have concrete instances for certain fractional
assertions such as ↦.
Coq is sometimes unable to infer the [Φ], hence it might be useful to write
[apply: (frame_fractional (λ q, ...))] when using the lemma to prove your
custom instance. See also https://github.com/coq/coq/issues/17688 *)
Lemma frame_fractional Φ p R P qR qP r :
AsFractional R Φ qR
AsFractional P Φ qP
FrameFractionalQp qR qP r
Frame p R P (Φ r).
Proof.
rewrite /Frame /FrameFractionalQp=> -[-> _] [-> ?] ->.
by rewrite bi.intuitionistically_if_elim fractional.
Qed.
End fractional.
(** Marked [tc_opaque] instead [Typeclasses Opaque] so that you can use
[iDestruct] to eliminate and [iModIntro] to introduce [internal_fractional],
while still preventing [iFrame] and [iNext] from unfolding. *)
Definition internal_fractional {PROP : bi} (Φ : Qp PROP) : PROP :=
tc_opaque ( p q, Φ (p + q)%Qp ∗-∗ Φ p Φ q)%I.
Global Instance: Params (@internal_fractional) 1 := {}.
Section internal_fractional.
Context {PROP : bi}.
Implicit Types Φ Ψ : Qp PROP.
Implicit Types q : Qp.
Global Instance internal_fractional_ne n :
Proper (pointwise_relation _ (dist n) ==> dist n) (@internal_fractional PROP).
Proof. solve_proper. Qed.
Global Instance internal_fractional_proper :
Proper (pointwise_relation _ () ==> ()) (@internal_fractional PROP).
Proof. solve_proper. Qed.
Global Instance internal_fractional_affine Φ : Affine (internal_fractional Φ).
Proof. rewrite /internal_fractional /=. apply _. Qed.
Global Instance internal_fractional_persistent Φ :
Persistent (internal_fractional Φ).
Proof. rewrite /internal_fractional /=. apply _. Qed.
Lemma fractional_internal_fractional Φ :
Fractional Φ internal_fractional Φ.
Proof.
intros. iIntros "!>" (q1 q2).
rewrite [Φ (q1 + q2)%Qp]fractional //=; auto.
Qed.
Lemma internal_fractional_iff Φ Ψ:
( q, Φ q ∗-∗ Ψ q) internal_fractional Φ -∗ internal_fractional Ψ.
Proof.
iIntros "#Hiff #HΦdup !>" (p q).
iSplit.
- iIntros "H".
iDestruct ("Hiff" with "H") as "HΦ".
iDestruct ("HΦdup" with "HΦ") as "[H1 ?]".
iSplitL "H1"; iApply "Hiff"; iFrame.
- iIntros "[H1 H2]".
iDestruct ("Hiff" with "H1") as "HΦ1".
iDestruct ("Hiff" with "H2") as "HΦ2".
iApply "Hiff".
iApply "HΦdup".
iFrame.
Qed.
End internal_fractional.
From iris.bi Require Export bi.
From iris.proofmode Require Import proofmode.
From iris.prelude Require Import options.
(** The class of laterable assertions *)
Class Laterable {PROP : bi} (P : PROP) := laterable :
P Q, Q ( Q -∗ P).
Global Arguments Laterable {_} _%_I : simpl never.
Global Arguments laterable {_} _%_I {_}.
Global Hint Mode Laterable + ! : typeclass_instances.
(** Proofmode class for turning [P] into a laterable [Q].
Will be the identity if [P] already is laterable, and add
[▷] otherwise. *)
Class IntoLaterable {PROP : bi} (P Q : PROP) : Prop := {
(** This is non-standard; usually we would just demand
[P ⊢ make_laterable Q]. However, we need these stronger properties for
the [make_laterable_id] hack in [atomic.v]. *)
into_laterable : P Q;
into_laterable_result_laterable : Laterable Q;
}.
Global Arguments IntoLaterable {_} P%_I Q%_I.
Global Arguments into_laterable {_} P%_I Q%_I {_}.
Global Arguments into_laterable_result_laterable {_} P%_I Q%_I {_}.
Global Hint Mode IntoLaterable + ! - : typeclass_instances.
Section instances.
Context {PROP : bi}.
Implicit Types P : PROP.
Implicit Types Ps : list PROP.
Global Instance laterable_proper : Proper ((⊣⊢) ==> ()) (@Laterable PROP).
Proof. solve_proper. Qed.
Global Instance later_laterable P : Laterable ( P).
Proof.
rewrite /Laterable. iIntros "HP". iExists P. iFrame.
iIntros "!> HP !>". done.
Qed.
Global Instance timeless_laterable P :
Timeless P Laterable P.
Proof.
rewrite /Laterable. iIntros (?) "HP". iExists P%I. iFrame.
iSplitR; first by iNext. iIntros "!> >HP !>". done.
Qed.
(** This lemma is not very useful: It needs a strange assumption about
emp, and most of the time intuitionistic propositions can be just kept
around anyway and don't need to be "latered". The lemma exists
because the fact that it needs the side-condition is interesting;
it is not an instance because it won't usually get used. *)
Lemma intuitionistic_laterable P :
Timeless (PROP:=PROP) emp Affine P Persistent P Laterable P.
Proof.
rewrite /Laterable. iIntros (???) "#HP".
iExists emp%I. iSplitL; first by iNext.
iIntros "!> >_". done.
Qed.
(** This instance, together with the one below, can lead to massive
backtracking, but only when searching for [Laterable]. In the future, it
should be rewritten using [Hint Immediate] or [Hint Cut], where the latter
is preferred once we figure out how to use it. *)
Global Instance persistent_laterable `{!BiAffine PROP} P :
Persistent P Laterable P.
Proof.
intros ?. apply intuitionistic_laterable; apply _.
Qed.
Global Instance sep_laterable P Q :
Laterable P Laterable Q Laterable (P Q).
Proof.
rewrite /Laterable. iIntros (LP LQ) "[HP HQ]".
iDestruct (LP with "HP") as (P') "[HP' #HP]".
iDestruct (LQ with "HQ") as (Q') "[HQ' #HQ]".
iExists (P' Q')%I. iSplitL; first by iFrame.
iIntros "!> [HP' HQ']". iSplitL "HP'".
- iApply "HP". done.
- iApply "HQ". done.
Qed.
Global Instance exist_laterable {A} (Φ : A PROP) :
( x, Laterable (Φ x)) Laterable ( x, Φ x).
Proof.
rewrite /Laterable. iIntros (). iDestruct 1 as (x) "H".
iDestruct ( with "H") as (Q) "[HQ #HΦ]".
iExists Q. iIntros "{$HQ} !> HQ". iExists x. by iApply "HΦ".
Qed.
Lemma big_sep_sepL_laterable Q Ps :
Laterable Q
TCForall Laterable Ps
Laterable (Q [] Ps).
Proof.
intros HQ HPs. revert Q HQ. induction HPs as [|P Ps ?? IH]; intros Q HQ.
{ simpl. rewrite right_id. done. }
simpl. rewrite assoc. apply IH; by apply _.
Qed.
Global Instance big_sepL_laterable Ps :
Laterable (PROP:=PROP) emp
TCForall Laterable Ps
Laterable ([] Ps).
Proof.
intros. assert (Laterable (emp [] Ps)) as Hlater.
{ apply big_sep_sepL_laterable; done. }
rewrite ->left_id in Hlater; last by apply _. done.
Qed.
(** A wrapper to obtain a weaker, laterable form of any assertion.
Alternatively: the modality corresponding to [Laterable].
The ◇ is required to make [make_laterable_intro'] hold.
TODO: Define [Laterable] in terms of this (see [laterable_alt] below). *)
Definition make_laterable (Q : PROP) : PROP :=
P, P ( P -∗ Q).
Global Instance make_laterable_ne : NonExpansive make_laterable.
Proof. solve_proper. Qed.
Global Instance make_laterable_proper : Proper (() ==> ()) make_laterable := ne_proper _.
Global Instance make_laterable_mono' : Proper (() ==> ()) make_laterable.
Proof. solve_proper. Qed.
Global Instance make_laterable_flip_mono' :
Proper (flip () ==> flip ()) make_laterable.
Proof. solve_proper. Qed.
Lemma make_laterable_mono Q1 Q2 :
(Q1 Q2) (make_laterable Q1 make_laterable Q2).
Proof. by intros ->. Qed.
Lemma make_laterable_except_0 Q :
make_laterable ( Q) make_laterable Q.
Proof.
iIntros "(%P & HP & #HPQ)". iExists P. iFrame.
iIntros "!# HP". iMod ("HPQ" with "HP"). done.
Qed.
Lemma make_laterable_sep Q1 Q2 :
make_laterable Q1 make_laterable Q2 make_laterable (Q1 Q2).
Proof.
iIntros "[HQ1 HQ2]".
iDestruct "HQ1" as (P1) "[HP1 #HQ1]".
iDestruct "HQ2" as (P2) "[HP2 #HQ2]".
iExists (P1 P2)%I. iFrame. iIntros "!# [HP1 HP2]".
iDestruct ("HQ1" with "HP1") as ">$".
iDestruct ("HQ2" with "HP2") as ">$".
done.
Qed.
(** A stronger version of [make_laterable_mono] that lets us keep laterable
resources. We cannot keep arbitrary resources since that would let us "frame
in" non-laterable things. *)
Lemma make_laterable_wand Q1 Q2 :
make_laterable (Q1 -∗ Q2) (make_laterable Q1 -∗ make_laterable Q2).
Proof.
iIntros "HQ HQ1".
iDestruct (make_laterable_sep with "[$HQ $HQ1 //]") as "HQ".
iApply make_laterable_mono; last done.
by rewrite bi.wand_elim_l.
Qed.
(** A variant of the above for keeping arbitrary intuitionistic resources.
Sadly, this is not implied by the above for non-affine BIs. *)
Lemma make_laterable_intuitionistic_wand Q1 Q2 :
(Q1 -∗ Q2) (make_laterable Q1 -∗ make_laterable Q2).
Proof.
iIntros "#HQ HQ1". iDestruct "HQ1" as (P) "[HP #HQ1]".
iExists P. iFrame. iIntros "!> HP".
iDestruct ("HQ1" with "HP") as "{HQ1} >HQ1".
iModIntro. iApply "HQ". done.
Qed.
Global Instance make_laterable_laterable Q :
Laterable (make_laterable Q).
Proof.
rewrite /Laterable. iIntros "HQ". iDestruct "HQ" as (P) "[HP #HQ]".
iExists P. iFrame. iIntros "!> HP !>". iExists P. by iFrame.
Qed.
Lemma make_laterable_elim Q :
make_laterable Q Q.
Proof.
iIntros "HQ". iDestruct "HQ" as (P) "[HP #HQ]". by iApply "HQ".
Qed.
(** Written internally (as an entailment of wands) to reflect
that persistent assertions can be kept unchanged. *)
Lemma make_laterable_intro P Q :
Laterable P
(P -∗ Q) -∗ P -∗ make_laterable Q.
Proof.
iIntros (?) "#HQ HP".
iDestruct (laterable with "HP") as (P') "[HP' #HPi]". iExists P'.
iFrame. iIntros "!> HP'".
iDestruct ("HPi" with "HP'") as ">HP". iModIntro.
iApply "HQ". done.
Qed.
Lemma make_laterable_intro' Q :
Laterable Q
Q make_laterable Q.
Proof.
intros ?. iApply make_laterable_intro. iIntros "!# $".
Qed.
Lemma make_laterable_idemp Q :
make_laterable (make_laterable Q) ⊣⊢ make_laterable Q.
Proof.
apply (anti_symm ()).
- trans (make_laterable ( Q)).
+ apply make_laterable_mono, make_laterable_elim.
+ apply make_laterable_except_0.
- apply make_laterable_intro', _.
Qed.
Lemma laterable_alt Q :
Laterable Q (Q make_laterable Q).
Proof.
split.
- intros ?. apply make_laterable_intro'. done.
- intros ?. done.
Qed.
(** * Proofmode integration
We integrate [make_laterable] with [iModIntro]. All non-laterable
hypotheses have a ▷ added in front of them to ensure a laterable context.
*)
Global Instance into_laterable_laterable P :
Laterable P
IntoLaterable P P.
Proof. intros ?. constructor; done. Qed.
Global Instance into_laterable_fallback P :
IntoLaterable P ( P) | 100.
Proof. constructor; last by apply _. apply bi.later_intro. Qed.
Lemma modality_make_laterable_mixin `{!Timeless (PROP:=PROP) emp} :
modality_mixin make_laterable MIEnvId (MIEnvTransform IntoLaterable).
Proof.
split; simpl; eauto using make_laterable_intro', make_laterable_mono,
make_laterable_sep, intuitionistic_laterable with typeclass_instances; [].
intros P Q ?. rewrite (into_laterable P).
apply make_laterable_intro'. eapply (into_laterable_result_laterable P), _.
Qed.
Definition modality_make_laterable `{!Timeless (PROP:=PROP) emp} :=
Modality _ modality_make_laterable_mixin.
Global Instance from_modal_make_laterable `{!Timeless (PROP:=PROP) emp} P :
FromModal True modality_make_laterable (make_laterable P) (make_laterable P) P.
Proof. by rewrite /FromModal. Qed.
End instances.
Global Typeclasses Opaque make_laterable.
(** This file provides constructions to lift
a PROP-level binary relation to various closures. *)
From iris.bi.lib Require Export fixpoint_mono.
From iris.proofmode Require Import proofmode.
From iris.prelude Require Import options.
(* The sections add extra BI assumptions, which is only picked up with "Type"*. *)
Set Default Proof Using "Type*".
(** * Definitions *)
Section definitions.
Context {PROP : bi} `{!BiInternalEq PROP}.
Context {A : ofe}.
Local Definition bi_rtc_pre (R : A A PROP)
(x2 : A) (rec : A PROP) (x1 : A) : PROP :=
<affine> (x1 x2) x', R x1 x' rec x'.
(** The reflexive transitive closure. *)
Definition bi_rtc (R : A A PROP) (x1 x2 : A) : PROP :=
bi_least_fixpoint (bi_rtc_pre R x2) x1.
Global Instance: Params (@bi_rtc) 4 := {}.
Local Definition bi_tc_pre (R : A A PROP)
(x2 : A) (rec : A PROP) (x1 : A) : PROP :=
R x1 x2 x', R x1 x' rec x'.
(** The transitive closure. *)
Definition bi_tc (R : A A PROP) (x1 x2 : A) : PROP :=
bi_least_fixpoint (bi_tc_pre R x2) x1.
Global Instance: Params (@bi_tc) 4 := {}.
(** Reductions of exactly [n] steps. *)
Fixpoint bi_nsteps (R : A A PROP) (n : nat) (x1 x2 : A) : PROP :=
match n with
| 0 => <affine> (x1 x2)
| S n' => x', R x1 x' bi_nsteps R n' x' x2
end.
Global Instance: Params (@bi_nsteps) 5 := {}.
End definitions.
Local Instance bi_rtc_pre_mono {PROP : bi} `{!BiInternalEq PROP}
{A : ofe} (R : A A PROP) `{!NonExpansive2 R} (x : A) :
BiMonoPred (bi_rtc_pre R x).
Proof.
constructor; [|solve_proper].
iIntros (rec1 rec2 ??) "#H". iIntros (x1) "[Hrec | Hrec]".
{ by iLeft. }
iRight.
iDestruct "Hrec" as (x') "[HP Hrec]".
iDestruct ("H" with "Hrec") as "Hrec". eauto with iFrame.
Qed.
Global Instance bi_rtc_ne {PROP : bi} `{!BiInternalEq PROP} {A : ofe} (R : A A PROP) :
NonExpansive2 (bi_rtc R).
Proof.
intros n x1 x2 Hx y1 y2 Hy. rewrite /bi_rtc Hx. f_equiv=> rec z.
solve_proper.
Qed.
Global Instance bi_rtc_proper {PROP : bi} `{!BiInternalEq PROP} {A : ofe} (R : A A PROP)
: Proper (() ==> () ==> (⊣⊢)) (bi_rtc R).
Proof. apply ne_proper_2. apply _. Qed.
Local Instance bi_tc_pre_mono `{!BiInternalEq PROP}
{A : ofe} (R : A A PROP) `{NonExpansive2 R} (x : A) :
BiMonoPred (bi_tc_pre R x).
Proof.
constructor; [|solve_proper].
iIntros (rec1 rec2 ??) "#H". iIntros (x1) "Hrec".
iDestruct "Hrec" as "[Hrec | Hrec]".
{ by iLeft. }
iDestruct "Hrec" as (x') "[HR Hrec]".
iRight. iExists x'. iFrame "HR". by iApply "H".
Qed.
Global Instance bi_tc_ne `{!BiInternalEq PROP} {A : ofe}
(R : A A PROP) `{NonExpansive2 R} :
NonExpansive2 (bi_tc R).
Proof.
intros n x1 x2 Hx y1 y2 Hy. rewrite /bi_tc Hx. f_equiv=> rec z.
solve_proper.
Qed.
Global Instance bi_tc_proper `{!BiInternalEq PROP} {A : ofe}
(R : A A PROP) `{NonExpansive2 R}
: Proper (() ==> () ==> (⊣⊢)) (bi_tc R).
Proof. apply ne_proper_2. apply _. Qed.
Global Instance bi_nsteps_ne {PROP : bi} `{!BiInternalEq PROP}
{A : ofe} (R : A A PROP) `{NonExpansive2 R} (n : nat) :
NonExpansive2 (bi_nsteps R n).
Proof. induction n; solve_proper. Qed.
Global Instance bi_nsteps_proper {PROP : bi} `{!BiInternalEq PROP}
{A : ofe} (R : A A PROP) `{NonExpansive2 R} (n : nat)
: Proper (() ==> () ==> (⊣⊢)) (bi_nsteps R n).
Proof. apply ne_proper_2. apply _. Qed.
(** * General theorems *)
Section general.
Context {PROP : bi} `{!BiInternalEq PROP}.
Context {A : ofe}.
Context (R : A A PROP) `{!NonExpansive2 R}.
(** ** Results about the reflexive-transitive closure [bi_rtc] *)
Local Lemma bi_rtc_unfold (x1 x2 : A) :
bi_rtc R x1 x2 bi_rtc_pre R x2 (λ x1, bi_rtc R x1 x2) x1.
Proof. by rewrite /bi_rtc; rewrite -least_fixpoint_unfold. Qed.
Lemma bi_rtc_strong_ind_l x2 Φ :
NonExpansive Φ
( x1, <affine> (x1 x2) ( x', R x1 x' (Φ x' bi_rtc R x' x2)) -∗ Φ x1) -∗
x1, bi_rtc R x1 x2 -∗ Φ x1.
Proof.
iIntros (?) "#IH". rewrite /bi_rtc.
by iApply (least_fixpoint_ind (bi_rtc_pre R x2) with "IH").
Qed.
Lemma bi_rtc_ind_l x2 Φ :
NonExpansive Φ
( x1, <affine> (x1 x2) ( x', R x1 x' Φ x') -∗ Φ x1) -∗
x1, bi_rtc R x1 x2 -∗ Φ x1.
Proof.
iIntros (?) "#IH". rewrite /bi_rtc.
by iApply (least_fixpoint_iter (bi_rtc_pre R x2) with "IH").
Qed.
Lemma bi_rtc_refl x : bi_rtc R x x.
Proof. rewrite bi_rtc_unfold. by iLeft. Qed.
Lemma bi_rtc_l x1 x2 x3 : R x1 x2 -∗ bi_rtc R x2 x3 -∗ bi_rtc R x1 x3.
Proof.
iIntros "H1 H2".
iEval (rewrite bi_rtc_unfold /bi_rtc_pre). iRight.
iExists x2. iFrame.
Qed.
Lemma bi_rtc_once x1 x2 : R x1 x2 -∗ bi_rtc R x1 x2.
Proof. iIntros "H". iApply (bi_rtc_l with "H"). iApply bi_rtc_refl. Qed.
Lemma bi_rtc_trans x1 x2 x3 : bi_rtc R x1 x2 -∗ bi_rtc R x2 x3 -∗ bi_rtc R x1 x3.
Proof.
iRevert (x1).
iApply bi_rtc_ind_l.
{ solve_proper. }
iIntros "!>" (x1) "[H | H] H2".
{ by iRewrite "H". }
iDestruct "H" as (x') "[H IH]".
iApply (bi_rtc_l with "H").
by iApply "IH".
Qed.
Lemma bi_rtc_r x y z : bi_rtc R x y -∗ R y z -∗ bi_rtc R x z.
Proof.
iIntros "H H'".
iApply (bi_rtc_trans with "H").
by iApply bi_rtc_once.
Qed.
Lemma bi_rtc_inv x z :
bi_rtc R x z -∗ <affine> (x z) y, R x y bi_rtc R y z.
Proof. rewrite bi_rtc_unfold. iIntros "[H | H]"; eauto. Qed.
Global Instance bi_rtc_affine :
( x y, Affine (R x y))
x y, Affine (bi_rtc R x y).
Proof. intros. apply least_fixpoint_affine; apply _. Qed.
(* FIXME: It would be nicer to use the least_fixpoint_persistent lemmas,
but they seem to weak. *)
Global Instance bi_rtc_persistent :
( x y, Persistent (R x y))
x y, Persistent (bi_rtc R x y).
Proof.
intros ? x y. rewrite /Persistent.
iRevert (x). iApply bi_rtc_ind_l; first solve_proper.
iIntros "!>" (x) "[#Heq | (%x' & #Hxx' & #Hx'y)]".
{ iRewrite "Heq". iApply bi_rtc_refl. }
iApply (bi_rtc_l with "Hxx' Hx'y").
Qed.
(** ** Results about the transitive closure [bi_tc] *)
Local Lemma bi_tc_unfold (x1 x2 : A) :
bi_tc R x1 x2 bi_tc_pre R x2 (λ x1, bi_tc R x1 x2) x1.
Proof. by rewrite /bi_tc; rewrite -least_fixpoint_unfold. Qed.
Lemma bi_tc_strong_ind_l x2 Φ :
NonExpansive Φ
( x1, (R x1 x2 ( x', R x1 x' (Φ x' bi_tc R x' x2))) -∗ Φ x1) -∗
x1, bi_tc R x1 x2 -∗ Φ x1.
Proof.
iIntros (?) "#IH". rewrite /bi_tc.
iApply (least_fixpoint_ind (bi_tc_pre R x2) with "IH").
Qed.
Lemma bi_tc_ind_l x2 Φ :
NonExpansive Φ
( x1, (R x1 x2 ( x', R x1 x' Φ x')) -∗ Φ x1) -∗
x1, bi_tc R x1 x2 -∗ Φ x1.
Proof.
iIntros (?) "#IH". rewrite /bi_tc.
iApply (least_fixpoint_iter (bi_tc_pre R x2) with "IH").
Qed.
Lemma bi_tc_l x1 x2 x3 : R x1 x2 -∗ bi_tc R x2 x3 -∗ bi_tc R x1 x3.
Proof.
iIntros "H1 H2".
iEval (rewrite bi_tc_unfold /bi_tc_pre).
iRight. iExists x2. iFrame.
Qed.
Lemma bi_tc_once x1 x2 : R x1 x2 -∗ bi_tc R x1 x2.
Proof.
iIntros "H".
iEval (rewrite bi_tc_unfold /bi_tc_pre).
by iLeft.
Qed.
Lemma bi_tc_trans x1 x2 x3 : bi_tc R x1 x2 -∗ bi_tc R x2 x3 -∗ bi_tc R x1 x3.
Proof.
iRevert (x1).
iApply bi_tc_ind_l.
{ solve_proper. }
iIntros "!>" (x1) "H H2".
iDestruct "H" as "[H | H]".
{ iApply (bi_tc_l with "H H2"). }
iDestruct "H" as (x') "[H IH]".
iApply (bi_tc_l with "H").
by iApply "IH".
Qed.
Lemma bi_tc_r x y z : bi_tc R x y -∗ R y z -∗ bi_tc R x z.
Proof.
iIntros "H H'".
iApply (bi_tc_trans with "H").
by iApply bi_tc_once.
Qed.
Lemma bi_tc_rtc_l x y z : bi_rtc R x y -∗ bi_tc R y z -∗ bi_tc R x z.
Proof.
iRevert (x).
iApply bi_rtc_ind_l. { solve_proper. }
iIntros "!>" (x) "[Heq | H] Hyz".
{ by iRewrite "Heq". }
iDestruct "H" as (x') "[H IH]".
iApply (bi_tc_l with "H").
by iApply "IH".
Qed.
Lemma bi_tc_rtc_r x y z : bi_tc R x y -∗ bi_rtc R y z -∗ bi_tc R x z.
Proof.
iIntros "Hxy Hyz".
iRevert (x) "Hxy". iRevert (y) "Hyz".
iApply bi_rtc_ind_l. { solve_proper. }
iIntros "!>" (y) "[Heq | H] %x Hxy".
{ by iRewrite -"Heq". }
iDestruct "H" as (y') "[H IH]".
iApply "IH". iApply (bi_tc_r with "Hxy H").
Qed.
Lemma bi_tc_rtc x y : bi_tc R x y -∗ bi_rtc R x y.
Proof.
iRevert (x). iApply bi_tc_ind_l. { solve_proper. }
iIntros "!>" (x) "[Hxy | H]".
{ by iApply bi_rtc_once. }
iDestruct "H" as (x') "[H H']".
iApply (bi_rtc_l with "H H'").
Qed.
Global Instance bi_tc_affine :
( x y, Affine (R x y))
x y, Affine (bi_tc R x y).
Proof. intros. apply least_fixpoint_affine; apply _. Qed.
Global Instance bi_tc_absorbing :
( x y, Absorbing (R x y))
x y, Absorbing (bi_tc R x y).
Proof. intros. apply least_fixpoint_absorbing; apply _. Qed.
(* FIXME: It would be nicer to use the least_fixpoint_persistent lemmas,
but they seem to weak. *)
Global Instance bi_tc_persistent :
( x y, Persistent (R x y))
x y, Persistent (bi_tc R x y).
Proof.
intros ? x y. rewrite /Persistent.
iRevert (x). iApply bi_tc_ind_l; first solve_proper.
iIntros "!# %x [#H|(%x'&#?&#?)] !>"; first by iApply bi_tc_once.
by iApply bi_tc_l.
Qed.
(** ** Results about [bi_nsteps] *)
Lemma bi_nsteps_O x : bi_nsteps R 0 x x.
Proof. auto. Qed.
Lemma bi_nsteps_once x y : R x y -∗ bi_nsteps R 1 x y.
Proof. simpl. eauto. Qed.
Lemma bi_nsteps_once_inv x y : bi_nsteps R 1 x y -∗ R x y.
Proof. iDestruct 1 as (x') "[Hxx' Heq]". by iRewrite -"Heq". Qed.
Lemma bi_nsteps_l n x y z : R x y -∗ bi_nsteps R n y z -∗ bi_nsteps R (S n) x z.
Proof. iIntros "? ?". iExists y. iFrame. Qed.
Lemma bi_nsteps_trans n m x y z :
bi_nsteps R n x y -∗ bi_nsteps R m y z -∗ bi_nsteps R (n + m) x z.
Proof.
iInduction n as [|n IH] forall (x); simpl.
- iIntros "Heq". iRewrite "Heq". auto.
- iDestruct 1 as (x') "[Hxx' Hx'y]". iIntros "Hyz".
iExists x'. iFrame "Hxx'". iApply ("IH" with "Hx'y Hyz").
Qed.
Lemma bi_nsteps_r n x y z :
bi_nsteps R n x y -∗ R y z -∗ bi_nsteps R (S n) x z.
Proof.
iIntros "Hxy Hyx". rewrite -Nat.add_1_r.
iApply (bi_nsteps_trans with "Hxy").
by iApply bi_nsteps_once.
Qed.
Lemma bi_nsteps_add_inv n m x z :
bi_nsteps R (n + m) x z y, bi_nsteps R n x y bi_nsteps R m y z.
Proof.
iInduction n as [|n IH] forall (x).
- iIntros "Hxz". iExists x. auto.
- iDestruct 1 as (y) "[Hxy Hyz]".
iDestruct ("IH" with "Hyz") as (y') "[Hyy' Hy'z]".
iExists y'. iFrame "Hy'z".
iApply (bi_nsteps_l with "Hxy Hyy'").
Qed.
Lemma bi_nsteps_inv_r n x z :
bi_nsteps R (S n) x z y, bi_nsteps R n x y R y z.
Proof.
rewrite -Nat.add_1_r bi_nsteps_add_inv /=.
iDestruct 1 as (y) "[Hxy (%x' & Hxx' & Heq)]".
iExists y. iRewrite -"Heq". iFrame.
Qed.
(** ** Equivalences between closure operators *)
Lemma bi_rtc_tc x y : bi_rtc R x y ⊣⊢ <affine> (x y) bi_tc R x y.
Proof.
iSplit.
- iRevert (x). iApply bi_rtc_ind_l. { solve_proper. }
iIntros "!>" (x) "[Heq | H]".
{ by iLeft. }
iRight.
iDestruct "H" as (x') "[H [Heq | IH]]".
{ iRewrite -"Heq". by iApply bi_tc_once. }
iApply (bi_tc_l with "H IH").
- iIntros "[Heq | Hxy]".
{ iRewrite "Heq". iApply bi_rtc_refl. }
by iApply bi_tc_rtc.
Qed.
Lemma bi_tc_nsteps x y :
bi_tc R x y ⊣⊢ n, <affine> 0 < n bi_nsteps R n x y.
Proof.
iSplit.
- iRevert (x). iApply bi_tc_ind_l. { solve_proper. }
iIntros "!>" (x) "[Hxy | H]".
{ iExists 1. iSplitR; first auto with lia.
iApply (bi_nsteps_l with "Hxy"). iApply bi_nsteps_O. }
iDestruct "H" as (x') "[Hxx' IH]".
iDestruct "IH" as (n ?) "Hx'y".
iExists (S n). iSplitR; first auto with lia.
iApply (bi_nsteps_l with "Hxx' Hx'y").
- iDestruct 1 as (n ?) "Hxy".
iInduction n as [|n IH] forall (y). { lia. }
rewrite bi_nsteps_inv_r.
iDestruct "Hxy" as (x') "[Hxx' Hx'y]".
destruct n.
{ simpl. iRewrite "Hxx'". by iApply bi_tc_once. }
iApply (bi_tc_r with "[Hxx'] Hx'y").
iApply ("IH" with "[%] Hxx'"). lia.
Qed.
Lemma bi_rtc_nsteps x y : bi_rtc R x y ⊣⊢ n, bi_nsteps R n x y.
Proof.
iSplit.
- iRevert (x). iApply bi_rtc_ind_l. { solve_proper. }
iIntros "!>" (x) "[Heq | H]".
{ iExists 0. iRewrite "Heq". iApply bi_nsteps_O. }
iDestruct "H" as (x') "[Hxx' IH]".
iDestruct "IH" as (n) "Hx'y".
iExists (S n). iApply (bi_nsteps_l with "Hxx' Hx'y").
- iDestruct 1 as (n) "Hxy".
iInduction n as [|n IH] forall (y).
{ simpl. iRewrite "Hxy". iApply bi_rtc_refl. }
iDestruct (bi_nsteps_inv_r with "Hxy") as (x') "[Hxx' Hx'y]".
iApply (bi_rtc_r with "[Hxx'] Hx'y").
by iApply "IH".
Qed.
End general.
Section timeless.
Context {PROP : bi} `{!BiInternalEq PROP, !BiAffine PROP}.
Context `{!OfeDiscrete A}.
Context (R : A A PROP) `{!NonExpansive2 R}.
Global Instance bi_nsteps_timeless n :
( x y, Timeless (R x y))
x y, Timeless (bi_nsteps R n x y).
Proof. induction n; apply _. Qed.
Global Instance bi_rtc_timeless :
( x y, Timeless (R x y))
x y, Timeless (bi_rtc R x y).
Proof. intros ? x y. rewrite bi_rtc_nsteps. apply _. Qed.
Global Instance bi_tc_timeless :
( x y, Timeless (R x y))
x y, Timeless (bi_tc R x y).
Proof. intros ? x y. rewrite bi_tc_nsteps. apply _. Qed.
End timeless.
Global Typeclasses Opaque bi_rtc.
Global Typeclasses Opaque bi_tc.
Global Typeclasses Opaque bi_nsteps.
From stdpp Require Import coPset.
From iris.bi Require Import bi.
From iris.prelude Require Import options.
(** Definitions. *)
Structure biIndex :=
BiIndex
{ bi_index_type :> Type;
bi_index_inhabited : Inhabited bi_index_type;
bi_index_rel : SqSubsetEq bi_index_type;
bi_index_rel_preorder : PreOrder (⊑@{bi_index_type}) }.
Global Existing Instances bi_index_inhabited bi_index_rel bi_index_rel_preorder.
(* We may want to instantiate monPred with the reflexivity relation in
the case where there is no relevant order. In that case, there is
no bottom element, so that we do not want to force any BI index to
have one. *)
Class BiIndexBottom {I : biIndex} (bot : I) :=
bi_index_bot i : bot i.
Section cofe.
Context {I : biIndex} {PROP : bi}.
Implicit Types i : I.
Record monPred :=
MonPred { monPred_at :> I PROP;
monPred_mono : Proper (() ==> ()) monPred_at }.
Local Existing Instance monPred_mono.
Bind Scope bi_scope with monPred.
Implicit Types P Q : monPred.
(** Ofe + Cofe instances *)
Section cofe_def.
Inductive monPred_equiv' P Q : Prop :=
{ monPred_in_equiv i : P i Q i } .
Local Instance monPred_equiv : Equiv monPred := monPred_equiv'.
Inductive monPred_dist' (n : nat) (P Q : monPred) : Prop :=
{ monPred_in_dist i : P i {n} Q i }.
Local Instance monPred_dist : Dist monPred := monPred_dist'.
Definition monPred_sig P : { f : I -d> PROP | Proper (() ==> ()) f } :=
exist _ (monPred_at P) (monPred_mono P).
Definition sig_monPred (P' : { f : I -d> PROP | Proper (() ==> ()) f })
: monPred :=
MonPred (proj1_sig P') (proj2_sig P').
(* These two lemma use the wrong Equiv and Dist instance for
monPred. so we make sure they are not accessible outside of the
section by using Let. *)
Let monPred_sig_equiv:
P Q, P Q monPred_sig P monPred_sig Q.
Proof. by split; [intros []|]. Defined.
Let monPred_sig_dist:
n, P Q : monPred, P {n} Q monPred_sig P {n} monPred_sig Q.
Proof. by split; [intros []|]. Defined.
Definition monPred_ofe_mixin : OfeMixin monPred.
Proof.
by apply (iso_ofe_mixin monPred_sig monPred_sig_equiv monPred_sig_dist).
Qed.
Canonical Structure monPredO := Ofe monPred monPred_ofe_mixin.
Global Instance monPred_cofe `{!Cofe PROP} : Cofe monPredO.
Proof.
refine ((iso_cofe_subtype' (A:=I-d>PROP) _ MonPred monPred_at _ _ _) _);
[done..|].
apply limit_preserving_forall=> y. apply limit_preserving_forall=> x.
apply limit_preserving_impl; [solve_proper|].
apply bi.limit_preserving_entails; solve_proper.
Qed.
End cofe_def.
Lemma monPred_sig_monPred (P' : { f : I -d> PROP | Proper (() ==> ()) f }) :
monPred_sig (sig_monPred P') P'.
Proof. by change (P' P'). Qed.
Lemma sig_monPred_sig P : sig_monPred (monPred_sig P) P.
Proof. done. Qed.
Global Instance monPred_sig_ne : NonExpansive monPred_sig.
Proof. move=> ??? [?] ? //=. Qed.
Global Instance monPred_sig_proper : Proper (() ==> ()) monPred_sig.
Proof. eapply (ne_proper _). Qed.
Global Instance sig_monPred_ne : NonExpansive (@sig_monPred).
Proof. split=>? //=. Qed.
Global Instance sig_monPred_proper : Proper (() ==> ()) sig_monPred.
Proof. eapply (ne_proper _). Qed.
(* We generalize over the relation R which is morally the equivalence
relation over B. That way, the BI index can use equality as an
equivalence relation (and Coq is able to infer the Proper and
Reflexive instances properly), or any other equivalence relation,
provided it is compatible with (⊑). *)
Global Instance monPred_at_ne (R : relation I) :
Proper (R ==> R ==> iff) () Reflexive R
n, Proper (dist n ==> R ==> dist n) monPred_at.
Proof.
intros ????? [Hd] ?? HR. rewrite Hd.
apply equiv_dist, bi.equiv_entails; split; f_equiv; rewrite ->HR; done.
Qed.
Global Instance monPred_at_proper (R : relation I) :
Proper (R ==> R ==> iff) () Reflexive R
Proper (() ==> R ==> ()) monPred_at.
Proof. repeat intro. apply equiv_dist=>?. f_equiv=>//. by apply equiv_dist. Qed.
End cofe.
Global Arguments monPred _ _ : clear implicits.
Global Arguments monPred_at {_ _} _%_I _.
Local Existing Instance monPred_mono.
Global Arguments monPredO _ _ : clear implicits.
(** BI canonical structure and type class instances *)
Module Export monPred_defs.
Section monPred_defs.
Context {I : biIndex} {PROP : bi}.
Implicit Types i : I.
Notation monPred := (monPred I PROP).
Implicit Types P Q : monPred.
Inductive monPred_entails (P1 P2 : monPred) : Prop :=
{ monPred_in_entails i : P1 i P2 i }.
Local Hint Immediate monPred_in_entails : core.
Program Definition monPred_upclosed (Φ : I PROP) : monPred :=
MonPred (λ i, ( j, i j Φ j)%I) _.
Next Obligation. solve_proper. Qed.
Local Definition monPred_embed_def : Embed PROP monPred := λ (P : PROP),
MonPred (λ _, P) _.
Local Definition monPred_embed_aux : seal (@monPred_embed_def).
Proof. by eexists. Qed.
Definition monPred_embed := monPred_embed_aux.(unseal).
Local Definition monPred_embed_unseal :
@embed _ _ monPred_embed = _ := monPred_embed_aux.(seal_eq).
Local Definition monPred_emp_def : monPred := MonPred (λ _, emp)%I _.
Local Definition monPred_emp_aux : seal (@monPred_emp_def). Proof. by eexists. Qed.
Definition monPred_emp := monPred_emp_aux.(unseal).
Local Definition monPred_emp_unseal :
@monPred_emp = _ := monPred_emp_aux.(seal_eq).
Local Definition monPred_pure_def (φ : Prop) : monPred := MonPred (λ _, φ)%I _.
Local Definition monPred_pure_aux : seal (@monPred_pure_def).
Proof. by eexists. Qed.
Definition monPred_pure := monPred_pure_aux.(unseal).
Local Definition monPred_pure_unseal :
@monPred_pure = _ := monPred_pure_aux.(seal_eq).
Local Definition monPred_objectively_def P : monPred :=
MonPred (λ _, i, P i)%I _.
Local Definition monPred_objectively_aux : seal (@monPred_objectively_def).
Proof. by eexists. Qed.
Definition monPred_objectively := monPred_objectively_aux.(unseal).
Local Definition monPred_objectively_unseal :
@monPred_objectively = _ := monPred_objectively_aux.(seal_eq).
Local Definition monPred_subjectively_def P : monPred := MonPred (λ _, i, P i)%I _.
Local Definition monPred_subjectively_aux : seal (@monPred_subjectively_def).
Proof. by eexists. Qed.
Definition monPred_subjectively := monPred_subjectively_aux.(unseal).
Local Definition monPred_subjectively_unseal :
@monPred_subjectively = _ := monPred_subjectively_aux.(seal_eq).
Local Program Definition monPred_and_def P Q : monPred :=
MonPred (λ i, P i Q i)%I _.
Next Obligation. solve_proper. Qed.
Local Definition monPred_and_aux : seal (@monPred_and_def).
Proof. by eexists. Qed.
Definition monPred_and := monPred_and_aux.(unseal).
Local Definition monPred_and_unseal :
@monPred_and = _ := monPred_and_aux.(seal_eq).
Local Program Definition monPred_or_def P Q : monPred :=
MonPred (λ i, P i Q i)%I _.
Next Obligation. solve_proper. Qed.
Local Definition monPred_or_aux : seal (@monPred_or_def).
Proof. by eexists. Qed.
Definition monPred_or := monPred_or_aux.(unseal).
Local Definition monPred_or_unseal :
@monPred_or = _ := monPred_or_aux.(seal_eq).
Local Definition monPred_impl_def P Q : monPred :=
monPred_upclosed (λ i, P i Q i)%I.
Local Definition monPred_impl_aux : seal (@monPred_impl_def).
Proof. by eexists. Qed.
Definition monPred_impl := monPred_impl_aux.(unseal).
Local Definition monPred_impl_unseal :
@monPred_impl = _ := monPred_impl_aux.(seal_eq).
Local Program Definition monPred_forall_def A (Φ : A monPred) : monPred :=
MonPred (λ i, x : A, Φ x i)%I _.
Next Obligation. solve_proper. Qed.
Local Definition monPred_forall_aux : seal (@monPred_forall_def).
Proof. by eexists. Qed.
Definition monPred_forall := monPred_forall_aux.(unseal).
Local Definition monPred_forall_unseal :
@monPred_forall = _ := monPred_forall_aux.(seal_eq).
Local Program Definition monPred_exist_def A (Φ : A monPred) : monPred :=
MonPred (λ i, x : A, Φ x i)%I _.
Next Obligation. solve_proper. Qed.
Local Definition monPred_exist_aux : seal (@monPred_exist_def).
Proof. by eexists. Qed.
Definition monPred_exist := monPred_exist_aux.(unseal).
Local Definition monPred_exist_unseal :
@monPred_exist = _ := monPred_exist_aux.(seal_eq).
Local Program Definition monPred_sep_def P Q : monPred :=
MonPred (λ i, P i Q i)%I _.
Next Obligation. solve_proper. Qed.
Local Definition monPred_sep_aux : seal (@monPred_sep_def).
Proof. by eexists. Qed.
Definition monPred_sep := monPred_sep_aux.(unseal).
Local Definition monPred_sep_unseal :
@monPred_sep = _ := monPred_sep_aux.(seal_eq).
Local Definition monPred_wand_def P Q : monPred :=
monPred_upclosed (λ i, P i -∗ Q i)%I.
Local Definition monPred_wand_aux : seal (@monPred_wand_def).
Proof. by eexists. Qed.
Definition monPred_wand := monPred_wand_aux.(unseal).
Local Definition monPred_wand_unseal :
@monPred_wand = _ := monPred_wand_aux.(seal_eq).
Local Program Definition monPred_persistently_def P : monPred :=
MonPred (λ i, <pers> (P i))%I _.
Next Obligation. solve_proper. Qed.
Local Definition monPred_persistently_aux : seal (@monPred_persistently_def).
Proof. by eexists. Qed.
Definition monPred_persistently := monPred_persistently_aux.(unseal).
Local Definition monPred_persistently_unseal :
@monPred_persistently = _ := monPred_persistently_aux.(seal_eq).
Local Program Definition monPred_in_def (i0 : I) : monPred :=
MonPred (λ i : I, i0 i⌝%I) _.
Next Obligation. solve_proper. Qed.
Local Definition monPred_in_aux : seal (@monPred_in_def). Proof. by eexists. Qed.
Definition monPred_in := monPred_in_aux.(unseal).
Local Definition monPred_in_unseal :
@monPred_in = _ := monPred_in_aux.(seal_eq).
Local Program Definition monPred_later_def P : monPred := MonPred (λ i, (P i))%I _.
Next Obligation. solve_proper. Qed.
Local Definition monPred_later_aux : seal monPred_later_def.
Proof. by eexists. Qed.
Definition monPred_later := monPred_later_aux.(unseal).
Local Definition monPred_later_unseal :
monPred_later = _ := monPred_later_aux.(seal_eq).
Local Definition monPred_internal_eq_def `{!BiInternalEq PROP}
(A : ofe) (a b : A) : monPred := MonPred (λ _, a b)%I _.
Local Definition monPred_internal_eq_aux : seal (@monPred_internal_eq_def).
Proof. by eexists. Qed.
Definition monPred_internal_eq := monPred_internal_eq_aux.(unseal).
Global Arguments monPred_internal_eq {_}.
Local Definition monPred_internal_eq_unseal `{!BiInternalEq PROP} :
@internal_eq _ monPred_internal_eq = monPred_internal_eq_def.
Proof. by rewrite -monPred_internal_eq_aux.(seal_eq). Qed.
Local Program Definition monPred_bupd_def `{BiBUpd PROP}
(P : monPred) : monPred := MonPred (λ i, |==> P i)%I _.
Next Obligation. solve_proper. Qed.
Local Definition monPred_bupd_aux : seal (@monPred_bupd_def).
Proof. by eexists. Qed.
Definition monPred_bupd := monPred_bupd_aux.(unseal).
Global Arguments monPred_bupd {_}.
Local Definition monPred_bupd_unseal `{BiBUpd PROP} :
@bupd _ monPred_bupd = monPred_bupd_def.
Proof. by rewrite -monPred_bupd_aux.(seal_eq). Qed.
Local Program Definition monPred_fupd_def `{BiFUpd PROP} (E1 E2 : coPset)
(P : monPred) : monPred := MonPred (λ i, |={E1,E2}=> P i)%I _.
Next Obligation. solve_proper. Qed.
Local Definition monPred_fupd_aux : seal (@monPred_fupd_def).
Proof. by eexists. Qed.
Definition monPred_fupd := monPred_fupd_aux.(unseal).
Global Arguments monPred_fupd {_}.
Local Definition monPred_fupd_unseal `{BiFUpd PROP} :
@fupd _ monPred_fupd = monPred_fupd_def.
Proof. by rewrite -monPred_fupd_aux.(seal_eq). Qed.
Local Definition monPred_plainly_def `{BiPlainly PROP} P : monPred :=
MonPred (λ _, i, (P i))%I _.
Local Definition monPred_plainly_aux : seal (@monPred_plainly_def).
Proof. by eexists. Qed.
Definition monPred_plainly := monPred_plainly_aux.(unseal).
Global Arguments monPred_plainly {_}.
Local Definition monPred_plainly_unseal `{BiPlainly PROP} :
@plainly _ monPred_plainly = monPred_plainly_def.
Proof. by rewrite -monPred_plainly_aux.(seal_eq). Qed.
End monPred_defs.
(** This is not the final collection of unsealing lemmas, below we redefine
[monPred_unseal] to also unfold the BI layer (i.e., the projections of the BI
structures/classes). *)
Local Definition monPred_unseal :=
(@monPred_embed_unseal, @monPred_emp_unseal, @monPred_pure_unseal,
@monPred_objectively_unseal, @monPred_subjectively_unseal,
@monPred_and_unseal, @monPred_or_unseal, @monPred_impl_unseal,
@monPred_forall_unseal, @monPred_exist_unseal, @monPred_sep_unseal,
@monPred_wand_unseal, @monPred_persistently_unseal,
@monPred_in_unseal, @monPred_later_unseal).
End monPred_defs.
Global Arguments monPred_objectively {_ _} _%_I.
Global Arguments monPred_subjectively {_ _} _%_I.
Notation "'<obj>' P" := (monPred_objectively P) : bi_scope.
Notation "'<subj>' P" := (monPred_subjectively P) : bi_scope.
Section instances.
Context (I : biIndex) (PROP : bi).
Lemma monPred_bi_mixin : BiMixin (PROP:=monPred I PROP)
monPred_entails monPred_emp monPred_pure monPred_and monPred_or
monPred_impl monPred_forall monPred_exist monPred_sep monPred_wand.
Proof.
split; rewrite ?monPred_defs.monPred_unseal;
try by (split=> ? /=; repeat f_equiv).
- split.
+ intros P. by split.
+ intros P Q R [H1] [H2]. split => ?. by rewrite H1 H2.
- split.
+ intros [HPQ]. split; split => i; move: (HPQ i); by apply bi.equiv_entails.
+ intros [[] []]. split=>i. by apply bi.equiv_entails.
- intros P φ ?. split=> i. by apply bi.pure_intro.
- intros φ P HP. split=> i. apply bi.pure_elim'=> ?. by apply HP.
- intros P Q. split=> i. by apply bi.and_elim_l.
- intros P Q. split=> i. by apply bi.and_elim_r.
- intros P Q R [?] [?]. split=> i. by apply bi.and_intro.
- intros P Q. split=> i. by apply bi.or_intro_l.
- intros P Q. split=> i. by apply bi.or_intro_r.
- intros P Q R [?] [?]. split=> i. by apply bi.or_elim.
- intros P Q R [HR]. split=> i /=. setoid_rewrite bi.pure_impl_forall.
apply bi.forall_intro=> j. apply bi.forall_intro=> Hij.
apply bi.impl_intro_r. by rewrite -HR /= !Hij.
- intros P Q R [HR]. split=> i /=.
rewrite HR /= bi.forall_elim bi.pure_impl_forall bi.forall_elim //.
apply bi.impl_elim_l.
- intros A P Ψ . split=> i. apply bi.forall_intro => ?. by apply .
- intros A Ψ. split=> i. by apply: bi.forall_elim.
- intros A Ψ a. split=> i. by rewrite /= -bi.exist_intro.
- intros A Ψ Q . split=> i. apply bi.exist_elim => a. by apply .
- intros P P' Q Q' [?] [?]. split=> i. by apply bi.sep_mono.
- intros P. split=> i. by apply bi.emp_sep_1.
- intros P. split=> i. by apply bi.emp_sep_2.
- intros P Q. split=> i. by apply bi.sep_comm'.
- intros P Q R. split=> i. by apply bi.sep_assoc'.
- intros P Q R [HR]. split=> i /=. setoid_rewrite bi.pure_impl_forall.
apply bi.forall_intro=> j. apply bi.forall_intro=> Hij.
apply bi.wand_intro_r. by rewrite -HR /= !Hij.
- intros P Q R [HP]. split=> i. apply bi.wand_elim_l'.
rewrite HP /= bi.forall_elim bi.pure_impl_forall bi.forall_elim //.
Qed.
Lemma monPred_bi_persistently_mixin :
BiPersistentlyMixin (PROP:=monPred I PROP)
monPred_entails monPred_emp monPred_and
monPred_exist monPred_sep monPred_persistently.
Proof.
split; rewrite ?monPred_defs.monPred_unseal;
try by (split=> ? /=; repeat f_equiv).
- intros P Q [?]. split=> i /=. by f_equiv.
- intros P. split=> i. by apply bi.persistently_idemp_2.
- split=> i. by apply bi.persistently_emp_intro.
- intros A Ψ. split=> i. by apply bi.persistently_and_2.
- intros A Ψ. split=> i. by apply bi.persistently_exist_1.
- intros P Q. split=> i. apply bi.sep_elim_l, _.
- intros P Q. split=> i. by apply bi.persistently_and_sep_elim.
Qed.
Lemma monPred_bi_later_mixin :
BiLaterMixin (PROP:=monPred I PROP)
monPred_entails monPred_pure
monPred_or monPred_impl monPred_forall monPred_exist
monPred_sep monPred_persistently monPred_later.
Proof.
split; rewrite ?monPred_defs.monPred_unseal.
- by split=> ? /=; repeat f_equiv.
- intros P Q [?]. split=> i. by apply bi.later_mono.
- intros P. split=> i /=. by apply bi.later_intro.
- intros A Ψ. split=> i. by apply bi.later_forall_2.
- intros A Ψ. split=> i. by apply bi.later_exist_false.
- intros P Q. split=> i. by apply bi.later_sep_1.
- intros P Q. split=> i. by apply bi.later_sep_2.
- intros P. split=> i. by apply bi.later_persistently_1.
- intros P. split=> i. by apply bi.later_persistently_2.
- intros P. split=> i /=. rewrite -bi.forall_intro.
+ apply bi.later_false_em.
+ intros j. rewrite bi.pure_impl_forall. apply bi.forall_intro=> Hij.
by rewrite Hij.
Qed.
Canonical Structure monPredI : bi :=
{| bi_ofe_mixin := monPred_ofe_mixin;
bi_bi_mixin := monPred_bi_mixin;
bi_bi_persistently_mixin := monPred_bi_persistently_mixin;
bi_bi_later_mixin := monPred_bi_later_mixin |}.
(** We restate the unsealing lemmas so that they also unfold the BI layer. The
sealing lemmas are partially applied so that they also work under binders. *)
Local Lemma monPred_emp_unseal :
bi_emp = @monPred_defs.monPred_emp_def I PROP.
Proof. by rewrite -monPred_defs.monPred_emp_unseal. Qed.
Local Lemma monPred_pure_unseal :
bi_pure = @monPred_defs.monPred_pure_def I PROP.
Proof. by rewrite -monPred_defs.monPred_pure_unseal. Qed.
Local Lemma monPred_and_unseal :
bi_and = @monPred_defs.monPred_and_def I PROP.
Proof. by rewrite -monPred_defs.monPred_and_unseal. Qed.
Local Lemma monPred_or_unseal :
bi_or = @monPred_defs.monPred_or_def I PROP.
Proof. by rewrite -monPred_defs.monPred_or_unseal. Qed.
Local Lemma monPred_impl_unseal :
bi_impl = @monPred_defs.monPred_impl_def I PROP.
Proof. by rewrite -monPred_defs.monPred_impl_unseal. Qed.
Local Lemma monPred_forall_unseal :
@bi_forall _ = @monPred_defs.monPred_forall_def I PROP.
Proof. by rewrite -monPred_defs.monPred_forall_unseal. Qed.
Local Lemma monPred_exist_unseal :
@bi_exist _ = @monPred_defs.monPred_exist_def I PROP.
Proof. by rewrite -monPred_defs.monPred_exist_unseal. Qed.
Local Lemma monPred_sep_unseal :
bi_sep = @monPred_defs.monPred_sep_def I PROP.
Proof. by rewrite -monPred_defs.monPred_sep_unseal. Qed.
Local Lemma monPred_wand_unseal :
bi_wand = @monPred_defs.monPred_wand_def I PROP.
Proof. by rewrite -monPred_defs.monPred_wand_unseal. Qed.
Local Lemma monPred_persistently_unseal :
bi_persistently = @monPred_defs.monPred_persistently_def I PROP.
Proof. by rewrite -monPred_defs.monPred_persistently_unseal. Qed.
Local Lemma monPred_later_unseal :
bi_later = @monPred_defs.monPred_later_def I PROP.
Proof. by rewrite -monPred_defs.monPred_later_unseal. Qed.
(** This definition only includes the unseal lemmas for the [bi] connectives.
After we have defined the right class instances, we define [monPred_unseal],
which also includes [embed], [internal_eq], [bupd], [fupd], [plainly],
[monPred_objectively], [monPred_subjectively] and [monPred_in]. *)
Local Definition monPred_unseal_bi :=
(monPred_emp_unseal, monPred_pure_unseal, monPred_and_unseal,
monPred_or_unseal, monPred_impl_unseal, monPred_forall_unseal,
monPred_exist_unseal, monPred_sep_unseal, monPred_wand_unseal,
monPred_persistently_unseal, monPred_later_unseal).
Definition monPred_embedding_mixin : BiEmbedMixin PROP monPredI monPred_embed.
Proof.
split; try apply _; rewrite /bi_emp_valid
!(monPred_defs.monPred_embed_unseal, monPred_unseal_bi); try done.
- move=> P /= [/(_ inhabitant) ?] //.
- intros PROP' ? P Q.
set (f P := @monPred_at I PROP P inhabitant).
assert (NonExpansive f) by solve_proper.
apply (f_equivI f).
- intros P Q. split=> i /=.
by rewrite bi.forall_elim bi.pure_impl_forall bi.forall_elim.
- intros P Q. split=> i /=.
by rewrite bi.forall_elim bi.pure_impl_forall bi.forall_elim.
Qed.
Global Instance monPred_bi_embed : BiEmbed PROP monPredI :=
{| bi_embed_mixin := monPred_embedding_mixin |}.
Lemma monPred_internal_eq_mixin `{!BiInternalEq PROP} :
BiInternalEqMixin monPredI monPred_internal_eq.
Proof.
split;
rewrite !(monPred_defs.monPred_internal_eq_unseal, monPred_unseal_bi).
- split=> i /=. solve_proper.
- intros A P a. split=> i /=. apply internal_eq_refl.
- intros A a b Ψ ?. split=> i /=.
setoid_rewrite bi.pure_impl_forall. do 2 apply bi.forall_intro => ?.
erewrite (internal_eq_rewrite _ _ (flip Ψ _)) => //=. solve_proper.
- intros A1 A2 f g. split=> i /=. by apply fun_extI.
- intros A P x y. split=> i /=. by apply sig_equivI_1.
- intros A a b ?. split=> i /=. by apply discrete_eq_1.
- intros A x y. split=> i /=. by apply later_equivI_1.
- intros A x y. split=> i /=. by apply later_equivI_2.
Qed.
Global Instance monPred_bi_internal_eq `{BiInternalEq PROP} :
BiInternalEq monPredI :=
{| bi_internal_eq_mixin := monPred_internal_eq_mixin |}.
Lemma monPred_bupd_mixin `{BiBUpd PROP} : BiBUpdMixin monPredI monPred_bupd.
Proof.
split; rewrite !(monPred_defs.monPred_bupd_unseal, monPred_unseal_bi).
- split=>/= i. solve_proper.
- intros P. split=>/= i. apply bupd_intro.
- intros P Q [HPQ]. split=>/= i. by rewrite HPQ.
- intros P. split=>/= i. apply bupd_trans.
- intros P Q. split=>/= i. apply bupd_frame_r.
Qed.
Global Instance monPred_bi_bupd `{BiBUpd PROP} : BiBUpd monPredI :=
{| bi_bupd_mixin := monPred_bupd_mixin |}.
Lemma monPred_fupd_mixin `{BiFUpd PROP} : BiFUpdMixin monPredI monPred_fupd.
Proof.
split; rewrite /bi_emp_valid /bi_except_0
!(monPred_defs.monPred_fupd_unseal, monPred_unseal_bi).
- split=>/= i. solve_proper.
- intros E1 E2 HE12. split=>/= i. by apply fupd_mask_intro_subseteq.
- intros E1 E2 P. split=>/= i. apply except_0_fupd.
- intros E1 E2 P Q [HPQ]. split=>/= i. by rewrite HPQ.
- intros E1 E2 E3 P. split=>/= i. apply fupd_trans.
- intros E1 E2 Ef P HE1f. split=>/= i.
by rewrite (bi.forall_elim i) bi.pure_True // left_id fupd_mask_frame_r'.
- intros E1 E2 P Q. split=>/= i. apply fupd_frame_r.
Qed.
Global Instance monPred_bi_fupd `{BiFUpd PROP} : BiFUpd monPredI :=
{| bi_fupd_mixin := monPred_fupd_mixin |}.
Lemma monPred_plainly_mixin `{BiPlainly PROP} :
BiPlainlyMixin monPredI monPred_plainly.
Proof.
split; rewrite !(monPred_defs.monPred_plainly_unseal, monPred_unseal_bi).
- by (split=> ? /=; repeat f_equiv).
- intros P Q [?]. split=> i /=. by do 3 f_equiv.
- intros P. split=> i /=. by rewrite bi.forall_elim plainly_elim_persistently.
- intros P. split=> i /=. do 3 setoid_rewrite <-plainly_forall.
rewrite -plainly_idemp_2. f_equiv. by apply bi.forall_intro=>_.
- intros A Ψ. split=> i /=. apply bi.forall_intro=> j.
rewrite plainly_forall. apply bi.forall_intro=> a. by rewrite !bi.forall_elim.
- intros P Q. split=> i /=.
setoid_rewrite bi.pure_impl_forall. rewrite 2!bi.forall_elim //.
do 2 setoid_rewrite <-plainly_forall.
setoid_rewrite plainly_impl_plainly. f_equiv.
do 3 apply bi.forall_intro => ?. f_equiv. rewrite bi.forall_elim //.
- intros P. split=> i /=. apply bi.forall_intro=>_. by apply plainly_emp_intro.
- intros P Q. split=> i. apply bi.sep_elim_l, _.
- intros P. split=> i /=.
rewrite bi.later_forall. f_equiv=> j. by rewrite -later_plainly_1.
- intros P. split=> i /=.
rewrite bi.later_forall. f_equiv=> j. by rewrite -later_plainly_2.
Qed.
Global Instance monPred_bi_plainly `{BiPlainly PROP} : BiPlainly monPredI :=
{| bi_plainly_mixin := monPred_plainly_mixin |}.
Local Lemma monPred_embed_unseal :
embed = @monPred_defs.monPred_embed_def I PROP.
Proof. by rewrite -monPred_defs.monPred_embed_unseal. Qed.
Local Lemma monPred_internal_eq_unseal `{!BiInternalEq PROP} :
@internal_eq _ _ = @monPred_defs.monPred_internal_eq_def I PROP _.
Proof. by rewrite -monPred_defs.monPred_internal_eq_unseal. Qed.
Local Lemma monPred_bupd_unseal `{BiBUpd PROP} :
bupd = @monPred_defs.monPred_bupd_def I PROP _.
Proof. by rewrite -monPred_defs.monPred_bupd_unseal. Qed.
Local Lemma monPred_fupd_unseal `{BiFUpd PROP} :
fupd = @monPred_defs.monPred_fupd_def I PROP _.
Proof. by rewrite -monPred_defs.monPred_fupd_unseal. Qed.
Local Lemma monPred_plainly_unseal `{BiPlainly PROP} :
plainly = @monPred_defs.monPred_plainly_def I PROP _.
Proof. by rewrite -monPred_defs.monPred_plainly_unseal. Qed.
(** And finally the proper [unseal] tactic (which we also redefine outside
of the section since Ltac definitions do not outlive a section). *)
Local Definition monPred_unseal :=
(monPred_unseal_bi,
@monPred_defs.monPred_objectively_unseal,
@monPred_defs.monPred_subjectively_unseal,
@monPred_embed_unseal, @monPred_internal_eq_unseal,
@monPred_bupd_unseal, @monPred_fupd_unseal, @monPred_plainly_unseal,
@monPred_defs.monPred_in_unseal).
Ltac unseal := rewrite !monPred_unseal /=.
Global Instance monPred_bi_löb : BiLöb PROP BiLöb monPredI.
Proof. rewrite {2}/BiLöb; unseal=> ? P HP; split=> i /=. apply löb_weak, HP. Qed.
Global Instance monPred_bi_positive : BiPositive PROP BiPositive monPredI.
Proof. split => ?. rewrite /bi_affinely. unseal. apply bi_positive. Qed.
Global Instance monPred_bi_affine : BiAffine PROP BiAffine monPredI.
Proof. split => ?. unseal. by apply affine. Qed.
Global Instance monPred_bi_persistently_forall :
BiPersistentlyForall PROP BiPersistentlyForall monPredI.
Proof. intros ? A φ. split=> /= i. unseal. by apply persistently_forall_2. Qed.
Global Instance monPred_bi_pure_forall :
BiPureForall PROP BiPureForall monPredI.
Proof. intros ? A φ. split=> /= i. unseal. by apply pure_forall_2. Qed.
Global Instance monPred_bi_later_contractive :
BiLaterContractive PROP BiLaterContractive monPredI.
Proof. intros ? n. unseal=> P Q HPQ. split=> i /=. f_contractive. apply HPQ. Qed.
Global Instance monPred_bi_embed_emp : BiEmbedEmp PROP monPredI.
Proof. split. by unseal. Qed.
Global Instance monPred_bi_embed_later : BiEmbedLater PROP monPredI.
Proof. split; by unseal. Qed.
Global Instance monPred_bi_embed_internal_eq `{BiInternalEq PROP} :
BiEmbedInternalEq PROP monPredI.
Proof. split. by unseal. Qed.
Global Instance monPred_bi_bupd_fupd `{BiBUpdFUpd PROP} : BiBUpdFUpd monPredI.
Proof. intros E P. split=> i. unseal. apply bupd_fupd. Qed.
Global Instance monPred_bi_embed_bupd `{!BiBUpd PROP} :
BiEmbedBUpd PROP monPredI.
Proof. split. by unseal. Qed.
Global Instance monPred_bi_embed_fupd `{BiFUpd PROP} : BiEmbedFUpd PROP monPredI.
Proof. split. by unseal. Qed.
Global Instance monPred_bi_persistently_impl_plainly
`{!BiPlainly PROP, !BiPersistentlyForall PROP, !BiPersistentlyImplPlainly PROP} :
BiPersistentlyImplPlainly monPredI.
Proof.
intros P Q. split=> i. unseal. setoid_rewrite bi.pure_impl_forall.
setoid_rewrite <-plainly_forall.
do 2 setoid_rewrite bi.persistently_forall.
by setoid_rewrite persistently_impl_plainly.
Qed.
Global Instance monPred_bi_prop_ext
`{!BiPlainly PROP, !BiInternalEq PROP, !BiPropExt PROP} : BiPropExt monPredI.
Proof.
intros P Q. split=> i /=. rewrite /bi_wand_iff. unseal.
rewrite -{3}(sig_monPred_sig P) -{3}(sig_monPred_sig Q)
-f_equivI -sig_equivI !discrete_fun_equivI /=.
f_equiv=> j. rewrite prop_ext.
by rewrite !(bi.forall_elim j) !bi.pure_True // !bi.True_impl.
Qed.
Global Instance monPred_bi_plainly_exist `{!BiPlainly PROP, @BiIndexBottom I bot} :
BiPlainlyExist PROP BiPlainlyExist monPredI.
Proof.
split=> ? /=. unseal. rewrite (bi.forall_elim bot) plainly_exist_1.
do 2 f_equiv. apply bi.forall_intro=> ?. by do 2 f_equiv.
Qed.
Global Instance monPred_bi_embed_plainly `{BiPlainly PROP} :
BiEmbedPlainly PROP monPredI.
Proof.
split=> i. unseal. apply (anti_symm _).
- by apply bi.forall_intro.
- by rewrite (bi.forall_elim inhabitant).
Qed.
Global Instance monPred_bi_bupd_plainly `{BiBUpdPlainly PROP} :
BiBUpdPlainly monPredI.
Proof.
intros P. split=> /= i. unseal. by rewrite bi.forall_elim bupd_plainly.
Qed.
Global Instance monPred_bi_fupd_plainly `{BiFUpdPlainly PROP} :
BiFUpdPlainly monPredI.
Proof.
split; rewrite /bi_except_0; unseal.
- intros E E' P R. split=>/= i.
rewrite (bi.forall_elim i) bi.pure_True // bi.True_impl.
by rewrite (bi.forall_elim i) fupd_plainly_keep_l.
- intros E P. split=>/= i.
by rewrite (bi.forall_elim i) fupd_plainly_later.
- intros E A Φ. split=>/= i.
rewrite -fupd_plainly_forall_2. apply bi.forall_mono=> x.
by rewrite (bi.forall_elim i).
Qed.
End instances.
(** The final unseal tactic that also unfolds the BI layer. *)
Module Import monPred.
Ltac unseal := rewrite !monPred_unseal /=.
End monPred.
Class Objective {I : biIndex} {PROP : bi} (P : monPred I PROP) :=
objective_at i j : P i P j.
Global Arguments Objective {_ _} _%_I.
Global Arguments objective_at {_ _} _%_I {_}.
Global Hint Mode Objective + + ! : typeclass_instances.
Global Instance: Params (@Objective) 2 := {}.
(** Primitive facts that cannot be deduced from the BI structure. *)
Section bi_facts.
Context {I : biIndex} {PROP : bi}.
Local Notation monPred := (monPred I PROP).
Local Notation monPredI := (monPredI I PROP).
Local Notation monPred_at := (@monPred_at I PROP).
Local Notation BiIndexBottom := (@BiIndexBottom I).
Implicit Types i : I.
Implicit Types P Q : monPred.
(** monPred_at unfolding laws *)
Lemma monPred_at_pure i (φ : Prop) : monPred_at φ i ⊣⊢ φ⌝.
Proof. by unseal. Qed.
Lemma monPred_at_emp i : monPred_at emp i ⊣⊢ emp.
Proof. by unseal. Qed.
Lemma monPred_at_and i P Q : (P Q) i ⊣⊢ P i Q i.
Proof. by unseal. Qed.
Lemma monPred_at_or i P Q : (P Q) i ⊣⊢ P i Q i.
Proof. by unseal. Qed.
Lemma monPred_at_impl i P Q : (P Q) i ⊣⊢ j, i j P j Q j.
Proof. by unseal. Qed.
Lemma monPred_at_forall {A} i (Φ : A monPred) : ( x, Φ x) i ⊣⊢ x, Φ x i.
Proof. by unseal. Qed.
Lemma monPred_at_exist {A} i (Φ : A monPred) : ( x, Φ x) i ⊣⊢ x, Φ x i.
Proof. by unseal. Qed.
Lemma monPred_at_sep i P Q : (P Q) i ⊣⊢ P i Q i.
Proof. by unseal. Qed.
Lemma monPred_at_wand i P Q : (P -∗ Q) i ⊣⊢ j, i j P j -∗ Q j.
Proof. by unseal. Qed.
Lemma monPred_at_persistently i P : (<pers> P) i ⊣⊢ <pers> (P i).
Proof. by unseal. Qed.
Lemma monPred_at_in i j : monPred_at (monPred_in j) i ⊣⊢ j i⌝.
Proof. by unseal. Qed.
Lemma monPred_at_objectively i P : (<obj> P) i ⊣⊢ j, P j.
Proof. by unseal. Qed.
Lemma monPred_at_subjectively i P : (<subj> P) i ⊣⊢ j, P j.
Proof. by unseal. Qed.
Lemma monPred_at_persistently_if i p P : (<pers>?p P) i ⊣⊢ <pers>?p (P i).
Proof. destruct p=>//=. apply monPred_at_persistently. Qed.
Lemma monPred_at_affinely i P : (<affine> P) i ⊣⊢ <affine> (P i).
Proof. by rewrite /bi_affinely monPred_at_and monPred_at_emp. Qed.
Lemma monPred_at_affinely_if i p P : (<affine>?p P) i ⊣⊢ <affine>?p (P i).
Proof. destruct p=>//=. apply monPred_at_affinely. Qed.
Lemma monPred_at_intuitionistically i P : ( P) i ⊣⊢ (P i).
Proof.
by rewrite /bi_intuitionistically monPred_at_affinely monPred_at_persistently.
Qed.
Lemma monPred_at_intuitionistically_if i p P : (?p P) i ⊣⊢ ?p (P i).
Proof. destruct p=>//=. apply monPred_at_intuitionistically. Qed.
Lemma monPred_at_absorbingly i P : (<absorb> P) i ⊣⊢ <absorb> (P i).
Proof. by rewrite /bi_absorbingly monPred_at_sep monPred_at_pure. Qed.
Lemma monPred_at_absorbingly_if i p P : (<absorb>?p P) i ⊣⊢ <absorb>?p (P i).
Proof. destruct p=>//=. apply monPred_at_absorbingly. Qed.
Lemma monPred_wand_force i P Q : (P -∗ Q) i (P i -∗ Q i).
Proof. unseal. rewrite bi.forall_elim bi.pure_impl_forall bi.forall_elim //. Qed.
Lemma monPred_impl_force i P Q : (P Q) i (P i Q i).
Proof. unseal. rewrite bi.forall_elim bi.pure_impl_forall bi.forall_elim //. Qed.
(** Instances *)
Global Instance monPred_at_mono :
Proper (() ==> () ==> ()) monPred_at.
Proof. by move=> ?? [?] ?? ->. Qed.
Global Instance monPred_at_flip_mono :
Proper (flip () ==> flip () ==> flip ()) monPred_at.
Proof. solve_proper. Qed.
Global Instance monPred_in_proper (R : relation I) :
Proper (R ==> R ==> iff) () Reflexive R
Proper (R ==> ()) (@monPred_in I PROP).
Proof. unseal. split. solve_proper. Qed.
Global Instance monPred_in_mono : Proper (flip () ==> ()) (@monPred_in I PROP).
Proof. unseal. split. solve_proper. Qed.
Global Instance monPred_in_flip_mono : Proper (() ==> flip ()) (@monPred_in I PROP).
Proof. solve_proper. Qed.
Lemma monPred_persistent P : ( i, Persistent (P i)) Persistent P.
Proof. intros HP. constructor=> i. unseal. apply HP. Qed.
Lemma monPred_absorbing P : ( i, Absorbing (P i)) Absorbing P.
Proof. intros HP. constructor=> i. rewrite /bi_absorbingly. unseal. apply HP. Qed.
Lemma monPred_affine P : ( i, Affine (P i)) Affine P.
Proof. intros HP. constructor=> i. unseal. apply HP. Qed.
Global Instance monPred_at_persistent P i : Persistent P Persistent (P i).
Proof. move => [] /(_ i). by unseal. Qed.
Global Instance monPred_at_absorbing P i : Absorbing P Absorbing (P i).
Proof. move => [] /(_ i). rewrite /Absorbing /bi_absorbingly. by unseal. Qed.
Global Instance monPred_at_affine P i : Affine P Affine (P i).
Proof. move => [] /(_ i). unfold Affine. by unseal. Qed.
(** Note that [monPred_in] is *not* [Plain], because it depends on the index. *)
Global Instance monPred_in_persistent i : Persistent (@monPred_in I PROP i).
Proof. apply monPred_persistent=> j. rewrite monPred_at_in. apply _. Qed.
Global Instance monPred_in_absorbing i : Absorbing (@monPred_in I PROP i).
Proof. apply monPred_absorbing=> j. rewrite monPred_at_in. apply _. Qed.
Lemma monPred_at_embed i (P : PROP) : monPred_at P i ⊣⊢ P.
Proof. by unseal. Qed.
Lemma monPred_emp_unfold : emp%I =@{monPred} emp : PROP⎤%I.
Proof. by unseal. Qed.
Lemma monPred_pure_unfold : bi_pure =@{_ monPred} λ φ, φ : PROP⎤%I.
Proof. by unseal. Qed.
Lemma monPred_objectively_unfold : monPred_objectively = λ P, ⎡∀ i, P i⎤%I.
Proof. by unseal. Qed.
Lemma monPred_subjectively_unfold : monPred_subjectively = λ P, ⎡∃ i, P i⎤%I.
Proof. by unseal. Qed.
Global Instance monPred_objectively_ne : NonExpansive (@monPred_objectively I PROP).
Proof. rewrite monPred_objectively_unfold. solve_proper. Qed.
Global Instance monPred_objectively_proper :
Proper (() ==> ()) (@monPred_objectively I PROP).
Proof. apply (ne_proper _). Qed.
Lemma monPred_objectively_mono P Q : (P Q) (<obj> P <obj> Q).
Proof. rewrite monPred_objectively_unfold. solve_proper. Qed.
Global Instance monPred_objectively_mono' :
Proper (() ==> ()) (@monPred_objectively I PROP).
Proof. intros ???. by apply monPred_objectively_mono. Qed.
Global Instance monPred_objectively_flip_mono' :
Proper (flip () ==> flip ()) (@monPred_objectively I PROP).
Proof. intros ???. by apply monPred_objectively_mono. Qed.
Global Instance monPred_objectively_persistent `{!BiPersistentlyForall PROP} P :
Persistent P Persistent (<obj> P).
Proof. rewrite monPred_objectively_unfold. apply _. Qed.
Global Instance monPred_objectively_absorbing P : Absorbing P Absorbing (<obj> P).
Proof. rewrite monPred_objectively_unfold. apply _. Qed.
Global Instance monPred_objectively_affine P : Affine P Affine (<obj> P).
Proof. rewrite monPred_objectively_unfold. apply _. Qed.
Global Instance monPred_subjectively_ne : NonExpansive (@monPred_subjectively I PROP).
Proof. rewrite monPred_subjectively_unfold. solve_proper. Qed.
Global Instance monPred_subjectively_proper :
Proper (() ==> ()) (@monPred_subjectively I PROP).
Proof. apply (ne_proper _). Qed.
Lemma monPred_subjectively_mono P Q : (P Q) <subj> P <subj> Q.
Proof. rewrite monPred_subjectively_unfold. solve_proper. Qed.
Global Instance monPred_subjectively_mono' :
Proper (() ==> ()) (@monPred_subjectively I PROP).
Proof. intros ???. by apply monPred_subjectively_mono. Qed.
Global Instance monPred_subjectively_flip_mono' :
Proper (flip () ==> flip ()) (@monPred_subjectively I PROP).
Proof. intros ???. by apply monPred_subjectively_mono. Qed.
Global Instance monPred_subjectively_persistent P :
Persistent P Persistent (<subj> P).
Proof. rewrite monPred_subjectively_unfold. apply _. Qed.
Global Instance monPred_subjectively_absorbing P :
Absorbing P Absorbing (<subj> P).
Proof. rewrite monPred_subjectively_unfold. apply _. Qed.
Global Instance monPred_subjectively_affine P : Affine P Affine (<subj> P).
Proof. rewrite monPred_subjectively_unfold. apply _. Qed.
(* Laws for monPred_objectively and of Objective. *)
Lemma monPred_objectively_elim P : <obj> P P.
Proof. rewrite monPred_objectively_unfold. unseal. split=>?. apply bi.forall_elim. Qed.
Lemma monPred_objectively_idemp P : <obj> <obj> P ⊣⊢ <obj> P.
Proof.
apply bi.equiv_entails; split; [by apply monPred_objectively_elim|].
unseal. split=>i /=. by apply bi.forall_intro=>_.
Qed.
Lemma monPred_objectively_forall {A} (Φ : A monPred) :
<obj> ( x, Φ x) ⊣⊢ x, <obj> (Φ x).
Proof.
unseal. split=>i. apply bi.equiv_entails; split=>/=;
do 2 apply bi.forall_intro=>?; by do 2 rewrite bi.forall_elim.
Qed.
Lemma monPred_objectively_and P Q : <obj> (P Q) ⊣⊢ <obj> P <obj> Q.
Proof.
unseal. split=>i. apply bi.equiv_entails; split=>/=.
- apply bi.and_intro; do 2 f_equiv.
+ apply bi.and_elim_l.
+ apply bi.and_elim_r.
- apply bi.forall_intro=>?. by rewrite !bi.forall_elim.
Qed.
Lemma monPred_objectively_exist {A} (Φ : A monPred) :
( x, <obj> (Φ x)) <obj> ( x, (Φ x)).
Proof. apply bi.exist_elim=>?. f_equiv. apply bi.exist_intro. Qed.
Lemma monPred_objectively_or P Q : <obj> P <obj> Q <obj> (P Q).
Proof.
apply bi.or_elim; f_equiv.
- apply bi.or_intro_l.
- apply bi.or_intro_r.
Qed.
Lemma monPred_objectively_sep_2 P Q : <obj> P <obj> Q <obj> (P Q).
Proof.
unseal. split=>i /=. apply bi.forall_intro=>?. by rewrite !bi.forall_elim.
Qed.
Lemma monPred_objectively_sep `{BiIndexBottom bot} P Q :
<obj> (P Q) ⊣⊢ <obj> P <obj> Q.
Proof.
apply bi.equiv_entails, conj, monPred_objectively_sep_2. unseal. split=>i /=.
rewrite (bi.forall_elim bot). by f_equiv; apply bi.forall_intro=>j; f_equiv.
Qed.
Lemma monPred_objectively_embed (P : PROP) : <obj> P ⊣⊢@{monPredI} P⎤.
Proof.
apply bi.equiv_entails; split; unseal; split=>i /=.
- by rewrite (bi.forall_elim inhabitant).
- by apply bi.forall_intro.
Qed.
Lemma monPred_objectively_emp : <obj> (emp : monPred) ⊣⊢ emp.
Proof. rewrite monPred_emp_unfold. apply monPred_objectively_embed. Qed.
Lemma monPred_objectively_pure φ : <obj> ( φ : monPred) ⊣⊢ φ ⌝.
Proof. rewrite monPred_pure_unfold. apply monPred_objectively_embed. Qed.
Lemma monPred_subjectively_intro P : P <subj> P.
Proof. unseal. split=>?. apply bi.exist_intro. Qed.
Lemma monPred_subjectively_forall {A} (Φ : A monPred) :
(<subj> ( x, Φ x)) x, <subj> (Φ x).
Proof. apply bi.forall_intro=>?. f_equiv. apply bi.forall_elim. Qed.
Lemma monPred_subjectively_and P Q : <subj> (P Q) <subj> P <subj> Q.
Proof.
apply bi.and_intro; f_equiv.
- apply bi.and_elim_l.
- apply bi.and_elim_r.
Qed.
Lemma monPred_subjectively_exist {A} (Φ : A monPred) :
<subj> ( x, Φ x) ⊣⊢ x, <subj> (Φ x).
Proof.
unseal. split=>i. apply bi.equiv_entails; split=>/=;
do 2 apply bi.exist_elim=>?; by do 2 rewrite -bi.exist_intro.
Qed.
Lemma monPred_subjectively_or P Q : <subj> (P Q) ⊣⊢ <subj> P <subj> Q.
Proof. split=>i. unseal. apply bi.or_exist. Qed.
Lemma monPred_subjectively_sep P Q : <subj> (P Q) <subj> P <subj> Q.
Proof.
unseal. split=>i /=. apply bi.exist_elim=>?. by rewrite -!bi.exist_intro.
Qed.
Lemma monPred_subjectively_idemp P : <subj> <subj> P ⊣⊢ <subj> P.
Proof.
apply bi.equiv_entails; split; [|by apply monPred_subjectively_intro].
unseal. split=>i /=. by apply bi.exist_elim=>_.
Qed.
Lemma objective_objectively P `{!Objective P} : P <obj> P.
Proof.
rewrite monPred_objectively_unfold /= embed_forall. apply bi.forall_intro=>?.
split=>?. unseal. apply objective_at, _.
Qed.
Lemma objective_subjectively P `{!Objective P} : <subj> P P.
Proof.
rewrite monPred_subjectively_unfold /= embed_exist. apply bi.exist_elim=>?.
split=>?. unseal. apply objective_at, _.
Qed.
Global Instance embed_objective (P : PROP) : @Objective I PROP P⎤.
Proof. intros ??. by unseal. Qed.
Global Instance pure_objective φ : @Objective I PROP φ⌝.
Proof. intros ??. by unseal. Qed.
Global Instance emp_objective : @Objective I PROP emp.
Proof. intros ??. by unseal. Qed.
Global Instance objectively_objective P : Objective (<obj> P).
Proof. intros ??. by unseal. Qed.
Global Instance subjectively_objective P : Objective (<subj> P).
Proof. intros ??. by unseal. Qed.
Global Instance and_objective P Q `{!Objective P, !Objective Q} :
Objective (P Q).
Proof. intros i j. unseal. by rewrite !(objective_at _ i j). Qed.
Global Instance or_objective P Q `{!Objective P, !Objective Q} :
Objective (P Q).
Proof. intros i j. by rewrite !monPred_at_or !(objective_at _ i j). Qed.
Global Instance impl_objective P Q `{!Objective P, !Objective Q} :
Objective (P Q).
Proof.
intros i j. unseal. rewrite (bi.forall_elim i) bi.pure_impl_forall.
rewrite bi.forall_elim //. apply bi.forall_intro=> k.
rewrite bi.pure_impl_forall. apply bi.forall_intro=>_.
rewrite (objective_at Q i). by rewrite (objective_at P k).
Qed.
Global Instance forall_objective {A} Φ {H : x : A, Objective (Φ x)} :
@Objective I PROP ( x, Φ x)%I.
Proof. intros i j. unseal. do 2 f_equiv. by apply objective_at. Qed.
Global Instance exists_objective {A} Φ {H : x : A, Objective (Φ x)} :
@Objective I PROP ( x, Φ x)%I.
Proof. intros i j. unseal. do 2 f_equiv. by apply objective_at. Qed.
Global Instance sep_objective P Q `{!Objective P, !Objective Q} :
Objective (P Q).
Proof. intros i j. unseal. by rewrite !(objective_at _ i j). Qed.
Global Instance wand_objective P Q `{!Objective P, !Objective Q} :
Objective (P -∗ Q).
Proof.
intros i j. unseal. rewrite (bi.forall_elim i) bi.pure_impl_forall.
rewrite bi.forall_elim //. apply bi.forall_intro=> k.
rewrite bi.pure_impl_forall. apply bi.forall_intro=>_.
rewrite (objective_at Q i). by rewrite (objective_at P k).
Qed.
Global Instance persistently_objective P `{!Objective P} : Objective (<pers> P).
Proof. intros i j. unseal. by rewrite objective_at. Qed.
Global Instance affinely_objective P `{!Objective P} : Objective (<affine> P).
Proof. rewrite /bi_affinely. apply _. Qed.
Global Instance intuitionistically_objective P `{!Objective P} : Objective ( P).
Proof. rewrite /bi_intuitionistically. apply _. Qed.
Global Instance absorbingly_objective P `{!Objective P} : Objective (<absorb> P).
Proof. rewrite /bi_absorbingly. apply _. Qed.
Global Instance persistently_if_objective P p `{!Objective P} :
Objective (<pers>?p P).
Proof. rewrite /bi_persistently_if. destruct p; apply _. Qed.
Global Instance affinely_if_objective P p `{!Objective P} :
Objective (<affine>?p P).
Proof. rewrite /bi_affinely_if. destruct p; apply _. Qed.
Global Instance absorbingly_if_objective P p `{!Objective P} :
Objective (<absorb>?p P).
Proof. rewrite /bi_absorbingly_if. destruct p; apply _. Qed.
Global Instance intuitionistically_if_objective P p `{!Objective P} :
Objective (?p P).
Proof. rewrite /bi_intuitionistically_if. destruct p; apply _. Qed.
(** monPred_in *)
Lemma monPred_in_intro P : P i, monPred_in i P i⎤.
Proof.
unseal. split=>i /=.
rewrite /= -(bi.exist_intro i). apply bi.and_intro=>//. by apply bi.pure_intro.
Qed.
Lemma monPred_in_elim P i : monPred_in i P i P .
Proof.
apply bi.impl_intro_r. unseal. split=>i' /=.
eapply bi.pure_elim; [apply bi.and_elim_l|]=>?. rewrite bi.and_elim_r. by f_equiv.
Qed.
(** Big op *)
Global Instance monPred_at_monoid_and_homomorphism i :
MonoidHomomorphism bi_and bi_and () (flip monPred_at i).
Proof.
split; [split|]; try apply _; [apply monPred_at_and | apply monPred_at_pure].
Qed.
Global Instance monPred_at_monoid_or_homomorphism i :
MonoidHomomorphism bi_or bi_or () (flip monPred_at i).
Proof.
split; [split|]; try apply _; [apply monPred_at_or | apply monPred_at_pure].
Qed.
Global Instance monPred_at_monoid_sep_homomorphism i :
MonoidHomomorphism bi_sep bi_sep () (flip monPred_at i).
Proof.
split; [split|]; try apply _; [apply monPred_at_sep | apply monPred_at_emp].
Qed.
Lemma monPred_at_big_sepL {A} i (Φ : nat A monPred) l :
([ list] kx l, Φ k x) i ⊣⊢ [ list] kx l, Φ k x i.
Proof. apply (big_opL_commute (flip monPred_at i)). Qed.
Lemma monPred_at_big_sepM `{Countable K} {A} i (Φ : K A monPred) (m : gmap K A) :
([ map] kx m, Φ k x) i ⊣⊢ [ map] kx m, Φ k x i.
Proof. apply (big_opM_commute (flip monPred_at i)). Qed.
Lemma monPred_at_big_sepS `{Countable A} i (Φ : A monPred) (X : gset A) :
([ set] y X, Φ y) i ⊣⊢ [ set] y X, Φ y i.
Proof. apply (big_opS_commute (flip monPred_at i)). Qed.
Lemma monPred_at_big_sepMS `{Countable A} i (Φ : A monPred) (X : gmultiset A) :
([ mset] y X, Φ y) i ⊣⊢ ([ mset] y X, Φ y i).
Proof. apply (big_opMS_commute (flip monPred_at i)). Qed.
Global Instance monPred_objectively_monoid_and_homomorphism :
MonoidHomomorphism bi_and bi_and () (@monPred_objectively I PROP).
Proof.
split; [split|]; try apply _.
- apply monPred_objectively_and.
- apply monPred_objectively_pure.
Qed.
Global Instance monPred_objectively_monoid_sep_entails_homomorphism :
MonoidHomomorphism bi_sep bi_sep (flip ()) (@monPred_objectively I PROP).
Proof.
split; [split|]; try apply _.
- apply monPred_objectively_sep_2.
- by rewrite monPred_objectively_emp.
Qed.
Global Instance monPred_objectively_monoid_sep_homomorphism `{BiIndexBottom bot} :
MonoidHomomorphism bi_sep bi_sep () (@monPred_objectively I PROP).
Proof.
split; [split|]; try apply _.
- apply monPred_objectively_sep.
- by rewrite monPred_objectively_emp.
Qed.
Lemma monPred_objectively_big_sepL_entails {A} (Φ : nat A monPred) l :
([ list] kx l, <obj> (Φ k x)) <obj> ([ list] kx l, Φ k x).
Proof. apply (big_opL_commute monPred_objectively (R:=flip ())). Qed.
Lemma monPred_objectively_big_sepM_entails
`{Countable K} {A} (Φ : K A monPred) (m : gmap K A) :
([ map] kx m, <obj> (Φ k x)) <obj> ([ map] kx m, Φ k x).
Proof. apply (big_opM_commute monPred_objectively (R:=flip ())). Qed.
Lemma monPred_objectively_big_sepS_entails `{Countable A}
(Φ : A monPred) (X : gset A) :
([ set] y X, <obj> (Φ y)) <obj> ([ set] y X, Φ y).
Proof. apply (big_opS_commute monPred_objectively (R:=flip ())). Qed.
Lemma monPred_objectively_big_sepMS_entails `{Countable A}
(Φ : A monPred) (X : gmultiset A) :
([ mset] y X, <obj> (Φ y)) <obj> ([ mset] y X, Φ y).
Proof. apply (big_opMS_commute monPred_objectively (R:=flip ())). Qed.
Lemma monPred_objectively_big_sepL `{BiIndexBottom bot} {A}
(Φ : nat A monPred) l :
<obj> ([ list] kx l, Φ k x) ⊣⊢ ([ list] kx l, <obj> (Φ k x)).
Proof. apply (big_opL_commute _). Qed.
Lemma monPred_objectively_big_sepM `{BiIndexBottom bot} `{Countable K} {A}
(Φ : K A monPred) (m : gmap K A) :
<obj> ([ map] kx m, Φ k x) ⊣⊢ ([ map] kx m, <obj> (Φ k x)).
Proof. apply (big_opM_commute _). Qed.
Lemma monPred_objectively_big_sepS `{BiIndexBottom bot} `{Countable A}
(Φ : A monPred) (X : gset A) :
<obj> ([ set] y X, Φ y) ⊣⊢ ([ set] y X, <obj> (Φ y)).
Proof. apply (big_opS_commute _). Qed.
Lemma monPred_objectively_big_sepMS `{BiIndexBottom bot} `{Countable A}
(Φ : A monPred) (X : gmultiset A) :
<obj> ([ mset] y X, Φ y) ⊣⊢ ([ mset] y X, <obj> (Φ y)).
Proof. apply (big_opMS_commute _). Qed.
Global Instance big_sepL_objective {A} (l : list A) Φ `{ n x, Objective (Φ n x)} :
@Objective I PROP ([ list] nx l, Φ n x).
Proof. generalize dependent Φ. induction l=>/=; apply _. Qed.
Global Instance big_sepM_objective `{Countable K} {A}
(Φ : K A monPred) (m : gmap K A) `{ k x, Objective (Φ k x)} :
Objective ([ map] kx m, Φ k x).
Proof.
intros ??. rewrite !monPred_at_big_sepM. do 3 f_equiv. by apply objective_at.
Qed.
Global Instance big_sepS_objective `{Countable A} (Φ : A monPred)
(X : gset A) `{ y, Objective (Φ y)} :
Objective ([ set] y X, Φ y).
Proof.
intros ??. rewrite !monPred_at_big_sepS. do 2 f_equiv. by apply objective_at.
Qed.
Global Instance big_sepMS_objective `{Countable A} (Φ : A monPred)
(X : gmultiset A) `{ y, Objective (Φ y)} :
Objective ([ mset] y X, Φ y).
Proof.
intros ??. rewrite !monPred_at_big_sepMS. do 2 f_equiv. by apply objective_at.
Qed.
(** BUpd *)
Lemma monPred_at_bupd `{!BiBUpd PROP} i P : (|==> P) i ⊣⊢ |==> P i.
Proof. by rewrite monPred_bupd_unseal. Qed.
Global Instance bupd_objective `{!BiBUpd PROP} P `{!Objective P} :
Objective (|==> P).
Proof. intros ??. by rewrite !monPred_at_bupd objective_at. Qed.
(** Later *)
Global Instance monPred_at_timeless P i : Timeless P Timeless (P i).
Proof. move => [] /(_ i). rewrite /Timeless /bi_except_0. by unseal. Qed.
Global Instance monPred_in_timeless i0 : Timeless (@monPred_in I PROP i0).
Proof. split => ? /=. rewrite /bi_except_0. unseal. apply timeless, _. Qed.
Global Instance monPred_objectively_timeless P : Timeless P Timeless (<obj> P).
Proof.
move=>[]. rewrite /Timeless /bi_except_0. unseal=>Hti. split=> ? /=.
by apply timeless, bi.forall_timeless.
Qed.
Global Instance monPred_subjectively_timeless P : Timeless P Timeless (<subj> P).
Proof.
move=>[]. rewrite /Timeless /bi_except_0. unseal=>Hti. split=> ? /=.
by apply timeless, bi.exist_timeless.
Qed.
Lemma monPred_at_later i P : ( P) i ⊣⊢ P i.
Proof. by unseal. Qed.
Lemma monPred_at_laterN n i P : (▷^n P) i ⊣⊢ ▷^n P i.
Proof. induction n as [|? IHn]; first done. rewrite /= monPred_at_later IHn //. Qed.
Lemma monPred_at_except_0 i P : ( P) i ⊣⊢ P i.
Proof. rewrite /bi_except_0. by unseal. Qed.
Global Instance later_objective P `{!Objective P} : Objective ( P).
Proof. intros ??. unseal. by rewrite objective_at. Qed.
Global Instance laterN_objective P `{!Objective P} n : Objective (▷^n P).
Proof. induction n; apply _. Qed.
Global Instance except0_objective P `{!Objective P} : Objective ( P).
Proof. rewrite /bi_except_0. apply _. Qed.
(** Internal equality *)
Lemma monPred_internal_eq_unfold `{!BiInternalEq PROP} :
@internal_eq monPredI _ = λ A x y, x y ⎤%I.
Proof. rewrite monPred_internal_eq_unseal. by unseal. Qed.
Lemma monPred_at_internal_eq `{!BiInternalEq PROP} {A : ofe} i (a b : A) :
@monPred_at (a b) i ⊣⊢ a b.
Proof. rewrite monPred_internal_eq_unfold. by apply monPred_at_embed. Qed.
Lemma monPred_equivI `{!BiInternalEq PROP'} P Q :
P Q ⊣⊢@{PROP'} i, P i Q i.
Proof.
apply bi.equiv_entails. split.
- apply bi.forall_intro=> ?. apply (f_equivI (flip monPred_at _)).
- by rewrite -{2}(sig_monPred_sig P) -{2}(sig_monPred_sig Q)
-f_equivI -sig_equivI !discrete_fun_equivI.
Qed.
Global Instance internal_eq_objective `{!BiInternalEq PROP} {A : ofe} (x y : A) :
@Objective I PROP (x y).
Proof. intros ??. rewrite monPred_internal_eq_unfold. by unseal. Qed.
(** FUpd *)
Lemma monPred_at_fupd `{!BiFUpd PROP} i E1 E2 P :
(|={E1,E2}=> P) i ⊣⊢ |={E1,E2}=> P i.
Proof. by rewrite monPred_fupd_unseal. Qed.
Global Instance fupd_objective E1 E2 P `{!Objective P} `{!BiFUpd PROP} :
Objective (|={E1,E2}=> P).
Proof. intros ??. by rewrite !monPred_at_fupd objective_at. Qed.
(** Plainly *)
Lemma monPred_plainly_unfold `{!BiPlainly PROP} : plainly = λ P, i, (P i) ⎤%I.
Proof. by rewrite monPred_plainly_unseal monPred_embed_unseal. Qed.
Lemma monPred_at_plainly `{!BiPlainly PROP} i P : ( P) i ⊣⊢ j, (P j).
Proof. by rewrite monPred_plainly_unseal. Qed.
Global Instance monPred_at_plain `{!BiPlainly PROP} P i : Plain P Plain (P i).
Proof. move => [] /(_ i). rewrite /Plain monPred_at_plainly bi.forall_elim //. Qed.
Global Instance plainly_objective `{!BiPlainly PROP} P : Objective ( P).
Proof. rewrite monPred_plainly_unfold. apply _. Qed.
Global Instance plainly_if_objective `{!BiPlainly PROP} P p `{!Objective P} :
Objective (?p P).
Proof. rewrite /plainly_if. destruct p; apply _. Qed.
Global Instance monPred_objectively_plain `{!BiPlainly PROP} P :
Plain P Plain (<obj> P).
Proof. rewrite monPred_objectively_unfold. apply _. Qed.
Global Instance monPred_subjectively_plain `{!BiPlainly PROP} P :
Plain P Plain (<subj> P).
Proof. rewrite monPred_subjectively_unfold. apply _. Qed.
End bi_facts.
From iris.prelude Require Import options.
(** Just reserve the notation. *)
(** * Turnstiles *)
Reserved Notation "P ⊢ Q" (at level 99, Q at level 200, right associativity).
Reserved Notation "P '⊢@{' PROP } Q" (at level 99, Q at level 200, right associativity).
Reserved Notation "(⊢)".
Reserved Notation "'(⊢@{' PROP } )".
Reserved Notation "( P ⊣⊢.)".
Reserved Notation "(.⊣⊢ Q )".
Reserved Notation "P ⊣⊢ Q" (at level 95, no associativity).
Reserved Notation "P '⊣⊢@{' PROP } Q" (at level 95, no associativity).
Reserved Notation "(⊣⊢)".
Reserved Notation "'(⊣⊢@{' PROP } )".
Reserved Notation "(.⊢ Q )".
Reserved Notation "( P ⊢.)".
Reserved Notation "⊢ Q" (at level 20, Q at level 200).
Reserved Notation "'⊢@{' PROP } Q" (at level 20, Q at level 200).
(** The definition must coincide with "'⊢@{' PROP } Q". *)
Reserved Notation "'(⊢@{' PROP } Q )".
(**
Rationale:
Notation [( '⊢@{' PROP } )] prevents parsing [(⊢@{PROP} Q)] using the
[⊢@{PROP} Q] notation; since the latter parse arises from composing two
notations, it is missed by the automatic left-factorization.
To fix that, we force left-factorization by explicitly composing parentheses with
['⊢@{' PROP } Q] into the new notation [( '⊢@{' PROP } Q )],
which successfully undergoes automatic left-factoring. *)
(** * BI connectives *)
Reserved Notation "'emp'".
Reserved Notation "'⌜' φ '⌝'" (at level 0, φ at level 200, format "⌜ φ ⌝").
Reserved Notation "P ∗ Q" (at level 80, right associativity, format "P ∗ '/' Q").
Reserved Notation "P -∗ Q"
(at level 99, Q at level 200, right associativity,
format "'[' P -∗ '/' '[' Q ']' ']'").
Reserved Notation "⎡ P ⎤".
(** Modalities *)
Reserved Notation "'<pers>' P" (at level 20, right associativity).
Reserved Notation "'<pers>?' p P" (at level 20, p at level 9, P at level 20,
right associativity, format "'<pers>?' p P").
Reserved Notation "▷ P" (at level 20, right associativity).
Reserved Notation "▷? p P" (at level 20, p at level 9, P at level 20,
format "▷? p P").
Reserved Notation "▷^ n P" (at level 20, n at level 9, P at level 20,
format "▷^ n P").
Reserved Infix "∗-∗" (at level 95, no associativity).
Reserved Notation "'<affine>' P" (at level 20, right associativity).
Reserved Notation "'<affine>?' p P" (at level 20, p at level 9, P at level 20,
right associativity, format "'<affine>?' p P").
Reserved Notation "'<absorb>' P" (at level 20, right associativity).
Reserved Notation "'<absorb>?' p P" (at level 20, p at level 9, P at level 20,
right associativity, format "'<absorb>?' p P").
Reserved Notation "□ P" (at level 20, right associativity).
Reserved Notation "'□?' p P" (at level 20, p at level 9, P at level 20,
right associativity, format "'□?' p P").
Reserved Notation "◇ P" (at level 20, right associativity).
Reserved Notation "■ P" (at level 20, right associativity).
Reserved Notation "■? p P" (at level 20, p at level 9, P at level 20,
right associativity, format "■? p P").
Reserved Notation "'<obj>' P" (at level 20, right associativity).
Reserved Notation "'<subj>' P" (at level 20, right associativity).
(** * Update modalities *)
Reserved Notation "|==> Q" (at level 99, Q at level 200, format "'[ ' |==> '/' Q ']'").
Reserved Notation "P ==∗ Q"
(at level 99, Q at level 200, format "'[' P ==∗ '/' Q ']'").
Reserved Notation "|={ E1 , E2 }=> Q"
(at level 99, E1, E2 at level 50, Q at level 200,
format "'[ ' |={ E1 , E2 }=> '/' Q ']'").
Reserved Notation "P ={ E1 , E2 }=∗ Q"
(at level 99, E1,E2 at level 50, Q at level 200,
format "'[' P ={ E1 , E2 }=∗ '/' '[' Q ']' ']'").
Reserved Notation "|={ E }=> Q"
(at level 99, E at level 50, Q at level 200,
format "'[ ' |={ E }=> '/' Q ']'").
Reserved Notation "P ={ E }=∗ Q"
(at level 99, E at level 50, Q at level 200,
format "'[' P ={ E }=∗ '/' '[' Q ']' ']'").
(** Step-taking fancy updates *)
Reserved Notation "|={ E1 } [ E2 ]▷=> Q"
(at level 99, E1, E2 at level 50, Q at level 200,
format "'[ ' |={ E1 } [ E2 ]▷=> '/' Q ']'").
Reserved Notation "P ={ E1 } [ E2 ]▷=∗ Q"
(at level 99, E1, E2 at level 50, Q at level 200,
format "'[' P ={ E1 } [ E2 ]▷=∗ '/' '[' Q ']' ']'").
Reserved Notation "|={ E }▷=> Q"
(at level 99, E at level 50, Q at level 200,
format "'[ ' |={ E }▷=> '/' Q ']'").
Reserved Notation "P ={ E }▷=∗ Q"
(at level 99, E at level 50, Q at level 200,
format "'[' P ={ E }▷=∗ '/' '[' Q ']' ']'").
(** Multi-step-taking fancy updates *)
Reserved Notation "|={ E1 } [ E2 ]▷=>^ n Q"
(at level 99, E1, E2 at level 50, n at level 9, Q at level 200,
format "'[ ' |={ E1 } [ E2 ]▷=>^ n '/' Q ']'").
Reserved Notation "P ={ E1 } [ E2 ]▷=∗^ n Q"
(at level 99, E1, E2 at level 50, n at level 9, Q at level 200,
format "'[' P ={ E1 } [ E2 ]▷=∗^ n '/' '[' Q ']' ']'").
Reserved Notation "|={ E }▷=>^ n Q"
(at level 99, E at level 50, n at level 9, Q at level 200,
format "'[ ' |={ E }▷=>^ n '/' Q ']'").
Reserved Notation "P ={ E }▷=∗^ n Q"
(at level 99, E at level 50, n at level 9, Q at level 200,
format "'[' P ={ E }▷=∗^ n '/' '[' Q ']' ']'").
(** * Big Ops *)
Reserved Notation "'[∗' 'list]' k ↦ x ∈ l , P"
(at level 200, l at level 10, k binder, x binder, right associativity,
format "[∗ list] k ↦ x ∈ l , P").
Reserved Notation "'[∗' 'list]' x ∈ l , P"
(at level 200, l at level 10, x binder, right associativity,
format "[∗ list] x ∈ l , P").
Reserved Notation "'[∗' 'list]' k ↦ x1 ; x2 ∈ l1 ; l2 , P"
(at level 200, l1, l2 at level 10, k binder, x1 binder, x2 binder,
right associativity,
format "[∗ list] k ↦ x1 ; x2 ∈ l1 ; l2 , P").
Reserved Notation "'[∗' 'list]' x1 ; x2 ∈ l1 ; l2 , P"
(at level 200, l1, l2 at level 10, x1 binder, x2 binder, right associativity,
format "[∗ list] x1 ; x2 ∈ l1 ; l2 , P").
Reserved Notation "'[∗]' Ps" (at level 20).
Reserved Notation "'[∧' 'list]' k ↦ x ∈ l , P"
(at level 200, l at level 10, k binder, x binder, right associativity,
format "[∧ list] k ↦ x ∈ l , P").
Reserved Notation "'[∧' 'list]' x ∈ l , P"
(at level 200, l at level 10, x binder, right associativity,
format "[∧ list] x ∈ l , P").
Reserved Notation "'[∧]' Ps" (at level 20).
Reserved Notation "'[∨' 'list]' k ↦ x ∈ l , P"
(at level 200, l at level 10, k binder, x binder, right associativity,
format "[∨ list] k ↦ x ∈ l , P").
Reserved Notation "'[∨' 'list]' x ∈ l , P"
(at level 200, l at level 10, x binder, right associativity,
format "[∨ list] x ∈ l , P").
Reserved Notation "'[∨]' Ps" (at level 20).
Reserved Notation "'[∗' 'map]' k ↦ x ∈ m , P"
(at level 200, m at level 10, k binder, x binder, right associativity,
format "[∗ map] k ↦ x ∈ m , P").
Reserved Notation "'[∗' 'map]' x ∈ m , P"
(at level 200, m at level 10, x binder, right associativity,
format "[∗ map] x ∈ m , P").
Reserved Notation "'[∗' 'map]' k ↦ x1 ; x2 ∈ m1 ; m2 , P"
(at level 200, m1, m2 at level 10,
k binder, x1 binder, x2 binder, right associativity,
format "[∗ map] k ↦ x1 ; x2 ∈ m1 ; m2 , P").
Reserved Notation "'[∗' 'map]' x1 ; x2 ∈ m1 ; m2 , P"
(at level 200, m1, m2 at level 10, x1 binder, x2 binder, right associativity,
format "[∗ map] x1 ; x2 ∈ m1 ; m2 , P").
Reserved Notation "'[∧' 'map]' k ↦ x ∈ m , P"
(at level 200, m at level 10, k binder, x binder, right associativity,
format "[∧ map] k ↦ x ∈ m , P").
Reserved Notation "'[∧' 'map]' x ∈ m , P"
(at level 200, m at level 10, x binder, right associativity,
format "[∧ map] x ∈ m , P").
Reserved Notation "'[∗' 'set]' x ∈ X , P"
(at level 200, X at level 10, x binder, right associativity,
format "[∗ set] x ∈ X , P").
Reserved Notation "'[∗' 'mset]' x ∈ X , P"
(at level 200, X at level 10, x binder, right associativity,
format "[∗ mset] x ∈ X , P").
(** Define the scope *)
Declare Scope bi_scope.
Delimit Scope bi_scope with I.
From iris.algebra Require Import monoid.
From iris.bi Require Import derived_laws_later big_op internal_eq.
From iris.prelude Require Import options.
Import interface.bi derived_laws.bi derived_laws_later.bi.
(* We enable primitive projections in this file to improve the performance of the Iris proofmode:
primitive projections for the bi-records makes the proofmode faster.
*)
Local Set Primitive Projections.
(* The sections add [BiAffine] and the like, which is only picked up with "Type"*. *)
Set Default Proof Using "Type*".
Class Plainly (PROP : Type) := plainly : PROP PROP.
Global Arguments plainly {PROP}%_type_scope {_} _%_I.
Global Hint Mode Plainly ! : typeclass_instances.
Global Instance: Params (@plainly) 2 := {}.
Global Typeclasses Opaque plainly.
Notation "■ P" := (plainly P) : bi_scope.
(* Mixins allow us to create instances easily without having to use Program *)
Record BiPlainlyMixin (PROP : bi) `(Plainly PROP) := {
bi_plainly_mixin_plainly_ne : NonExpansive (plainly (PROP:=PROP));
bi_plainly_mixin_plainly_mono (P Q : PROP) : (P Q) P Q;
bi_plainly_mixin_plainly_elim_persistently (P : PROP) : P <pers> P;
bi_plainly_mixin_plainly_idemp_2 (P : PROP) : P P;
bi_plainly_mixin_plainly_forall_2 {A} (Ψ : A PROP) :
( a, (Ψ a)) ( a, Ψ a);
(* The following law and [persistently_impl_plainly] below are very similar,
and indeed they hold not just for persistently and plainly, but for any
modality defined as [M P n x := ∀ y, R x y → P n y]. *)
bi_plainly_mixin_plainly_impl_plainly (P Q : PROP) :
( P Q) ( P Q);
bi_plainly_mixin_plainly_emp_intro (P : PROP) : P emp;
bi_plainly_mixin_plainly_absorb (P Q : PROP) : P Q P;
bi_plainly_mixin_later_plainly_1 (P : PROP) : P P;
bi_plainly_mixin_later_plainly_2 (P : PROP) : P P;
}.
Class BiPlainly (PROP : bi) := {
#[global] bi_plainly_plainly :: Plainly PROP;
bi_plainly_mixin : BiPlainlyMixin PROP bi_plainly_plainly;
}.
Global Hint Mode BiPlainly ! : typeclass_instances.
Global Arguments bi_plainly_plainly : simpl never.
Class BiPersistentlyImplPlainly `{!BiPlainly PROP} :=
persistently_impl_plainly (P Q : PROP) :
( P <pers> Q) <pers> ( P Q).
Global Arguments BiPersistentlyImplPlainly : clear implicits.
Global Arguments BiPersistentlyImplPlainly _ {_}.
Global Arguments persistently_impl_plainly _ {_ _} _.
Global Hint Mode BiPersistentlyImplPlainly ! - : typeclass_instances.
Class BiPlainlyExist {PROP: bi} `{!BiPlainly PROP} :=
plainly_exist_1 A (Ψ : A PROP) :
( a, Ψ a) a, (Ψ a).
Global Arguments BiPlainlyExist : clear implicits.
Global Arguments BiPlainlyExist _ {_}.
Global Arguments plainly_exist_1 _ {_ _} _.
Global Hint Mode BiPlainlyExist ! - : typeclass_instances.
Class BiPropExt {PROP: bi} `{!BiPlainly PROP, !BiInternalEq PROP} :=
prop_ext_2 (P Q : PROP) : (P ∗-∗ Q) P Q.
Global Arguments BiPropExt : clear implicits.
Global Arguments BiPropExt _ {_ _}.
Global Arguments prop_ext_2 _ {_ _ _} _.
Global Hint Mode BiPropExt ! - - : typeclass_instances.
Section plainly_laws.
Context {PROP: bi} `{!BiPlainly PROP}.
Implicit Types P Q : PROP.
Global Instance plainly_ne : NonExpansive (@plainly PROP _).
Proof. eapply bi_plainly_mixin_plainly_ne, bi_plainly_mixin. Qed.
Lemma plainly_mono P Q : (P Q) P Q.
Proof. eapply bi_plainly_mixin_plainly_mono, bi_plainly_mixin. Qed.
Lemma plainly_elim_persistently P : P <pers> P.
Proof. eapply bi_plainly_mixin_plainly_elim_persistently, bi_plainly_mixin. Qed.
Lemma plainly_idemp_2 P : P P.
Proof. eapply bi_plainly_mixin_plainly_idemp_2, bi_plainly_mixin. Qed.
Lemma plainly_forall_2 {A} (Ψ : A PROP) : ( a, (Ψ a)) ( a, Ψ a).
Proof. eapply bi_plainly_mixin_plainly_forall_2, bi_plainly_mixin. Qed.
Lemma plainly_impl_plainly P Q : ( P Q) ( P Q).
Proof. eapply bi_plainly_mixin_plainly_impl_plainly, bi_plainly_mixin. Qed.
Lemma plainly_absorb P Q : P Q P.
Proof. eapply bi_plainly_mixin_plainly_absorb, bi_plainly_mixin. Qed.
Lemma plainly_emp_intro P : P emp.
Proof. eapply bi_plainly_mixin_plainly_emp_intro, bi_plainly_mixin. Qed.
Lemma later_plainly_1 P : P ( P).
Proof. eapply bi_plainly_mixin_later_plainly_1, bi_plainly_mixin. Qed.
Lemma later_plainly_2 P : P P.
Proof. eapply bi_plainly_mixin_later_plainly_2, bi_plainly_mixin. Qed.
End plainly_laws.
(* Derived properties and connectives *)
Class Plain {PROP: bi} `{!BiPlainly PROP} (P : PROP) := plain : P P.
Global Arguments Plain {_ _} _%_I : simpl never.
Global Arguments plain {_ _} _%_I {_}.
Global Hint Mode Plain + - ! : typeclass_instances.
Global Instance: Params (@Plain) 1 := {}.
Global Hint Extern 100 (Plain (match ?x with _ => _ end)) =>
destruct x : typeclass_instances.
Definition plainly_if {PROP: bi} `{!BiPlainly PROP} (p : bool) (P : PROP) : PROP :=
(if p then P else P)%I.
Global Arguments plainly_if {_ _} !_ _%_I /.
Global Instance: Params (@plainly_if) 2 := {}.
Global Typeclasses Opaque plainly_if.
Notation "■? p P" := (plainly_if p P) : bi_scope.
(* Derived laws *)
Section plainly_derived.
Context {PROP: bi} `{!BiPlainly PROP}.
Implicit Types P : PROP.
Local Hint Resolve pure_intro forall_intro : core.
Local Hint Resolve or_elim or_intro_l' or_intro_r' : core.
Local Hint Resolve and_intro and_elim_l' and_elim_r' : core.
Global Instance plainly_proper :
Proper ((⊣⊢) ==> (⊣⊢)) (@plainly PROP _) := ne_proper _.
Global Instance plainly_mono' : Proper (() ==> ()) (@plainly PROP _).
Proof. intros P Q; apply plainly_mono. Qed.
Global Instance plainly_flip_mono' :
Proper (flip () ==> flip ()) (@plainly PROP _).
Proof. intros P Q; apply plainly_mono. Qed.
Lemma affinely_plainly_elim P : <affine> P P.
Proof. by rewrite plainly_elim_persistently /bi_affinely persistently_and_emp_elim. Qed.
Lemma persistently_elim_plainly P : <pers> P ⊣⊢ P.
Proof.
apply (anti_symm _).
- by rewrite persistently_into_absorbingly /bi_absorbingly comm plainly_absorb.
- by rewrite {1}plainly_idemp_2 plainly_elim_persistently.
Qed.
Lemma persistently_if_elim_plainly P p : <pers>?p P ⊣⊢ P.
Proof. destruct p; last done. exact: persistently_elim_plainly. Qed.
Lemma plainly_persistently_elim P : <pers> P ⊣⊢ P.
Proof.
apply (anti_symm _).
- rewrite -{1}(left_id True%I bi_and ( _)%I) (plainly_emp_intro True).
rewrite -{2}(persistently_and_emp_elim P).
rewrite !and_alt -plainly_forall_2. by apply forall_mono=> -[].
- by rewrite {1}plainly_idemp_2 (plainly_elim_persistently P).
Qed.
Lemma absorbingly_elim_plainly P : <absorb> P ⊣⊢ P.
Proof. by rewrite -(persistently_elim_plainly P) absorbingly_elim_persistently. Qed.
Lemma plainly_and_sep_elim P Q : P Q (emp P) Q.
Proof. by rewrite plainly_elim_persistently persistently_and_sep_elim_emp. Qed.
Lemma plainly_and_sep_assoc P Q R : P (Q R) ⊣⊢ ( P Q) R.
Proof. by rewrite -(persistently_elim_plainly P) persistently_and_sep_assoc. Qed.
Lemma plainly_and_emp_elim P : emp P P.
Proof. by rewrite plainly_elim_persistently persistently_and_emp_elim. Qed.
Lemma plainly_into_absorbingly P : P <absorb> P.
Proof. by rewrite plainly_elim_persistently persistently_into_absorbingly. Qed.
Lemma plainly_elim P `{!Absorbing P} : P P.
Proof. by rewrite plainly_elim_persistently persistently_elim. Qed.
Lemma plainly_idemp_1 P : P P.
Proof. by rewrite plainly_into_absorbingly absorbingly_elim_plainly. Qed.
Lemma plainly_idemp P : P ⊣⊢ P.
Proof. apply (anti_symm _); auto using plainly_idemp_1, plainly_idemp_2. Qed.
Lemma plainly_intro' P Q : ( P Q) P Q.
Proof. intros <-. apply plainly_idemp_2. Qed.
Lemma plainly_pure φ : φ ⊣⊢@{PROP} φ⌝.
Proof.
apply (anti_symm _); auto.
- by rewrite plainly_elim_persistently persistently_pure.
- apply pure_elim'=> .
trans ( x : False, True : PROP)%I; [by apply forall_intro|].
rewrite plainly_forall_2. by rewrite -(pure_intro φ).
Qed.
Lemma plainly_forall {A} (Ψ : A PROP) : ( a, Ψ a) ⊣⊢ a, (Ψ a).
Proof.
apply (anti_symm _); auto using plainly_forall_2.
apply forall_intro=> x. by rewrite (forall_elim x).
Qed.
Lemma plainly_exist_2 {A} (Ψ : A PROP) : ( a, (Ψ a)) ( a, Ψ a).
Proof. apply exist_elim=> x. by rewrite (exist_intro x). Qed.
Lemma plainly_exist `{!BiPlainlyExist PROP} {A} (Ψ : A PROP) :
( a, Ψ a) ⊣⊢ a, (Ψ a).
Proof. apply (anti_symm _); auto using plainly_exist_1, plainly_exist_2. Qed.
Lemma plainly_and P Q : (P Q) ⊣⊢ P Q.
Proof. rewrite !and_alt plainly_forall. by apply forall_proper=> -[]. Qed.
Lemma plainly_or_2 P Q : P Q (P Q).
Proof. rewrite !or_alt -plainly_exist_2. by apply exist_mono=> -[]. Qed.
Lemma plainly_or `{!BiPlainlyExist PROP} P Q : (P Q) ⊣⊢ P Q.
Proof. rewrite !or_alt plainly_exist. by apply exist_proper=> -[]. Qed.
Lemma plainly_impl P Q : (P Q) P Q.
Proof.
apply impl_intro_l; rewrite -plainly_and.
apply plainly_mono, impl_elim with P; auto.
Qed.
Lemma plainly_emp_2 : emp ⊢@{PROP} emp.
Proof. apply plainly_emp_intro. Qed.
Lemma plainly_sep_dup P : P ⊣⊢ P P.
Proof.
apply (anti_symm _).
- rewrite -{1}(idemp bi_and ( _)%I).
by rewrite -{2}(emp_sep ( _)%I) plainly_and_sep_assoc and_elim_l.
- by rewrite plainly_absorb.
Qed.
Lemma plainly_and_sep_l_1 P Q : P Q P Q.
Proof. by rewrite -{1}(emp_sep Q) plainly_and_sep_assoc and_elim_l. Qed.
Lemma plainly_and_sep_r_1 P Q : P Q P Q.
Proof. by rewrite !(comm _ P) plainly_and_sep_l_1. Qed.
Lemma plainly_True_emp : True ⊣⊢@{PROP} emp.
Proof. apply (anti_symm _); eauto using plainly_mono, plainly_emp_intro. Qed.
Lemma plainly_and_sep P Q : (P Q) (P Q).
Proof.
rewrite plainly_and.
rewrite -{1}plainly_idemp -plainly_and -{1}(emp_sep Q).
by rewrite plainly_and_sep_assoc (comm bi_and) plainly_and_emp_elim.
Qed.
Lemma plainly_affinely_elim P : <affine> P ⊣⊢ P.
Proof. by rewrite /bi_affinely plainly_and -plainly_True_emp plainly_pure left_id. Qed.
Lemma intuitionistically_plainly_elim P : P P.
Proof. rewrite intuitionistically_affinely plainly_elim_persistently //. Qed.
Lemma intuitionistically_plainly P : P P.
Proof.
rewrite /bi_intuitionistically plainly_affinely_elim affinely_elim.
rewrite persistently_elim_plainly plainly_persistently_elim. done.
Qed.
Lemma and_sep_plainly P Q : P Q ⊣⊢ P Q.
Proof.
apply (anti_symm _); auto using plainly_and_sep_l_1.
apply and_intro.
- by rewrite plainly_absorb.
- by rewrite comm plainly_absorb.
Qed.
Lemma plainly_sep_2 P Q : P Q (P Q).
Proof. by rewrite -plainly_and_sep plainly_and -and_sep_plainly. Qed.
Lemma plainly_sep `{!BiPositive PROP} P Q : (P Q) ⊣⊢ P Q.
Proof.
apply (anti_symm _); auto using plainly_sep_2.
rewrite -(plainly_affinely_elim (_ _)) affinely_sep -and_sep_plainly. apply and_intro.
- by rewrite (affinely_elim_emp Q) right_id affinely_elim.
- by rewrite (affinely_elim_emp P) left_id affinely_elim.
Qed.
Lemma plainly_wand P Q : (P -∗ Q) P -∗ Q.
Proof. apply wand_intro_r. by rewrite plainly_sep_2 wand_elim_l. Qed.
Lemma plainly_entails_l P Q : (P Q) P Q P.
Proof. intros; rewrite -plainly_and_sep_l_1; auto. Qed.
Lemma plainly_entails_r P Q : (P Q) P P Q.
Proof. intros; rewrite -plainly_and_sep_r_1; auto. Qed.
Lemma plainly_impl_wand_2 P Q : (P -∗ Q) (P Q).
Proof.
apply plainly_intro', impl_intro_r.
rewrite -{2}(emp_sep P) plainly_and_sep_assoc.
by rewrite (comm bi_and) plainly_and_emp_elim wand_elim_l.
Qed.
Lemma impl_wand_plainly_2 P Q : ( P -∗ Q) ( P Q).
Proof. apply impl_intro_l. by rewrite plainly_and_sep_l_1 wand_elim_r. Qed.
Lemma impl_wand_affinely_plainly P Q : ( P Q) ⊣⊢ (<affine> P -∗ Q).
Proof. by rewrite -(persistently_elim_plainly P) impl_wand_intuitionistically. Qed.
Lemma persistently_wand_affinely_plainly `{!BiPersistentlyImplPlainly PROP} P Q :
(<affine> P -∗ <pers> Q) <pers> (<affine> P -∗ Q).
Proof. rewrite -!impl_wand_affinely_plainly. apply: persistently_impl_plainly. Qed.
Lemma plainly_wand_affinely_plainly P Q :
(<affine> P -∗ Q) (<affine> P -∗ Q).
Proof. rewrite -!impl_wand_affinely_plainly. apply plainly_impl_plainly. Qed.
Section plainly_affine_bi.
Context `{!BiAffine PROP}.
Lemma plainly_emp : emp ⊣⊢@{PROP} emp.
Proof. by rewrite -!True_emp plainly_pure. Qed.
Lemma plainly_and_sep_l P Q : P Q ⊣⊢ P Q.
Proof.
apply (anti_symm ());
eauto using plainly_and_sep_l_1, sep_and with typeclass_instances.
Qed.
Lemma plainly_and_sep_r P Q : P Q ⊣⊢ P Q.
Proof. by rewrite !(comm _ P) plainly_and_sep_l. Qed.
Lemma plainly_impl_wand P Q : (P Q) ⊣⊢ (P -∗ Q).
Proof.
apply (anti_symm ()); auto using plainly_impl_wand_2.
apply plainly_intro', wand_intro_l.
by rewrite -plainly_and_sep_r plainly_elim impl_elim_r.
Qed.
Lemma impl_wand_plainly P Q : ( P Q) ⊣⊢ ( P -∗ Q).
Proof.
apply (anti_symm ()).
- by rewrite -impl_wand_1.
- by rewrite impl_wand_plainly_2.
Qed.
End plainly_affine_bi.
(* Conditional plainly *)
Global Instance plainly_if_ne p : NonExpansive (@plainly_if PROP _ p).
Proof. solve_proper. Qed.
Global Instance plainly_if_proper p : Proper ((⊣⊢) ==> (⊣⊢)) (@plainly_if PROP _ p).
Proof. solve_proper. Qed.
Global Instance plainly_if_mono' p : Proper (() ==> ()) (@plainly_if PROP _ p).
Proof. solve_proper. Qed.
Global Instance plainly_if_flip_mono' p :
Proper (flip () ==> flip ()) (@plainly_if PROP _ p).
Proof. solve_proper. Qed.
Lemma plainly_if_mono p P Q : (P Q) ?p P ?p Q.
Proof. by intros ->. Qed.
Lemma plainly_if_pure p φ : ?p φ ⊣⊢@{PROP} φ⌝.
Proof. destruct p; simpl; auto using plainly_pure. Qed.
Lemma plainly_if_and p P Q : ?p (P Q) ⊣⊢ ?p P ?p Q.
Proof. destruct p; simpl; auto using plainly_and. Qed.
Lemma plainly_if_or_2 p P Q : ?p P ?p Q ?p (P Q).
Proof. destruct p; simpl; auto using plainly_or_2. Qed.
Lemma plainly_if_or `{!BiPlainlyExist PROP} p P Q : ?p (P Q) ⊣⊢ ?p P ?p Q.
Proof. destruct p; simpl; auto using plainly_or. Qed.
Lemma plainly_if_exist_2 {A} p (Ψ : A PROP) : ( a, ?p (Ψ a)) ?p ( a, Ψ a).
Proof. destruct p; simpl; auto using plainly_exist_2. Qed.
Lemma plainly_if_exist `{!BiPlainlyExist PROP} {A} p (Ψ : A PROP) :
?p ( a, Ψ a) ⊣⊢ a, ?p (Ψ a).
Proof. destruct p; simpl; auto using plainly_exist. Qed.
Lemma plainly_if_sep_2 `{!BiPositive PROP} p P Q : ?p P ?p Q ?p (P Q).
Proof. destruct p; simpl; auto using plainly_sep_2. Qed.
Lemma plainly_if_idemp p P : ?p ?p P ⊣⊢ ?p P.
Proof. destruct p; simpl; auto using plainly_idemp. Qed.
(* Properties of plain propositions *)
Global Instance Plain_proper : Proper (() ==> iff) (@Plain PROP _).
Proof. solve_proper. Qed.
Lemma plain_plainly_2 P `{!Plain P} : P P.
Proof. done. Qed.
Lemma plain_plainly P `{!Plain P, !Absorbing P} : P ⊣⊢ P.
Proof. apply (anti_symm _), plain_plainly_2, _. by apply plainly_elim. Qed.
Lemma plainly_intro P Q `{!Plain P} : (P Q) P Q.
Proof. by intros <-. Qed.
(* Typeclass instances *)
Global Instance plainly_absorbing P : Absorbing ( P).
Proof. by rewrite /Absorbing /bi_absorbingly comm plainly_absorb. Qed.
Global Instance plainly_if_absorbing P p :
Absorbing P Absorbing (plainly_if p P).
Proof. intros; destruct p; simpl; apply _. Qed.
(* Not an instance, see the bottom of this file *)
Lemma plain_persistent P : Plain P Persistent P.
Proof. intros. by rewrite /Persistent -plainly_elim_persistently. Qed.
Global Instance impl_persistent `{!BiPersistentlyImplPlainly PROP} P Q :
Absorbing P Plain P Persistent Q Persistent (P Q).
Proof.
intros. by rewrite /Persistent {2}(plain P) -persistently_impl_plainly
-(persistent Q) (plainly_into_absorbingly P) absorbing.
Qed.
Global Instance plainly_persistent P : Persistent ( P).
Proof. by rewrite /Persistent persistently_elim_plainly. Qed.
Global Instance wand_persistent `{!BiPersistentlyImplPlainly PROP} P Q :
Plain P Persistent Q Absorbing Q Persistent (P -∗ Q).
Proof.
intros. rewrite /Persistent {2}(plain P). trans (<pers> ( P Q))%I.
- rewrite -persistently_impl_plainly impl_wand_affinely_plainly -(persistent Q).
by rewrite affinely_plainly_elim.
- apply persistently_mono, wand_intro_l. by rewrite sep_and impl_elim_r.
Qed.
Global Instance limit_preserving_Plain {A : ofe} `{!Cofe A} (Φ : A PROP) :
NonExpansive Φ LimitPreserving (λ x, Plain (Φ x)).
Proof. intros. apply limit_preserving_entails; solve_proper. Qed.
(* Instances for big operators *)
Global Instance plainly_sep_weak_homomorphism `{!BiPositive PROP, !BiAffine PROP} :
WeakMonoidHomomorphism bi_sep bi_sep () (@plainly PROP _).
Proof. split; try apply _. apply plainly_sep. Qed.
Global Instance plainly_sep_entails_weak_homomorphism :
WeakMonoidHomomorphism bi_sep bi_sep (flip ()) (@plainly PROP _).
Proof. split; try apply _. intros P Q; by rewrite plainly_sep_2. Qed.
Global Instance plainly_sep_entails_homomorphism `{!BiAffine PROP} :
MonoidHomomorphism bi_sep bi_sep (flip ()) (@plainly PROP _).
Proof. split; try apply _. simpl. rewrite plainly_emp. done. Qed.
Global Instance plainly_sep_homomorphism `{!BiAffine PROP} :
MonoidHomomorphism bi_sep bi_sep () (@plainly PROP _).
Proof. split; try apply _. apply plainly_emp. Qed.
Global Instance plainly_and_homomorphism :
MonoidHomomorphism bi_and bi_and () (@plainly PROP _).
Proof. split; [split|]; try apply _; [apply plainly_and | apply plainly_pure]. Qed.
Global Instance plainly_or_homomorphism `{!BiPlainlyExist PROP} :
MonoidHomomorphism bi_or bi_or () (@plainly PROP _).
Proof. split; [split|]; try apply _; [apply plainly_or | apply plainly_pure]. Qed.
Lemma big_sepL_plainly `{!BiAffine PROP} {A} (Φ : nat A PROP) l :
([ list] kx l, Φ k x) ⊣⊢ [ list] kx l, (Φ k x).
Proof. apply (big_opL_commute _). Qed.
Lemma big_andL_plainly {A} (Φ : nat A PROP) l :
([ list] kx l, Φ k x) ⊣⊢ [ list] kx l, (Φ k x).
Proof. apply (big_opL_commute _). Qed.
Lemma big_orL_plainly `{!BiPlainlyExist PROP} {A} (Φ : nat A PROP) l :
([ list] kx l, Φ k x) ⊣⊢ [ list] kx l, (Φ k x).
Proof. apply (big_opL_commute _). Qed.
Lemma big_sepL2_plainly `{!BiAffine PROP} {A B} (Φ : nat A B PROP) l1 l2 :
([ list] ky1;y2 l1;l2, Φ k y1 y2)
⊣⊢ [ list] ky1;y2 l1;l2, (Φ k y1 y2).
Proof. by rewrite !big_sepL2_alt plainly_and plainly_pure big_sepL_plainly. Qed.
Lemma big_sepM_plainly `{!BiAffine PROP, Countable K} {A} (Φ : K A PROP) m :
([ map] kx m, Φ k x) ⊣⊢ [ map] kx m, (Φ k x).
Proof. apply (big_opM_commute _). Qed.
Lemma big_sepM2_plainly `{!BiAffine PROP, Countable K} {A B} (Φ : K A B PROP) m1 m2 :
([ map] kx1;x2 m1;m2, Φ k x1 x2) ⊣⊢ [ map] kx1;x2 m1;m2, (Φ k x1 x2).
Proof. by rewrite !big_sepM2_alt plainly_and plainly_pure big_sepM_plainly. Qed.
Lemma big_sepS_plainly `{!BiAffine PROP, Countable A} (Φ : A PROP) X :
([ set] y X, Φ y) ⊣⊢ [ set] y X, (Φ y).
Proof. apply (big_opS_commute _). Qed.
Lemma big_sepMS_plainly `{!BiAffine PROP, Countable A} (Φ : A PROP) X :
([ mset] y X, Φ y) ⊣⊢ [ mset] y X, (Φ y).
Proof. apply (big_opMS_commute _). Qed.
(* Plainness instances *)
Global Instance pure_plain φ : Plain (PROP:=PROP) φ⌝.
Proof. by rewrite /Plain plainly_pure. Qed.
Global Instance emp_plain : Plain (PROP:=PROP) emp.
Proof. apply plainly_emp_intro. Qed.
Global Instance and_plain P Q : Plain P Plain Q Plain (P Q).
Proof. intros. by rewrite /Plain plainly_and -!plain. Qed.
Global Instance or_plain P Q : Plain P Plain Q Plain (P Q).
Proof. intros. by rewrite /Plain -plainly_or_2 -!plain. Qed.
Global Instance forall_plain {A} (Ψ : A PROP) :
( x, Plain (Ψ x)) Plain ( x, Ψ x).
Proof.
intros. rewrite /Plain plainly_forall. apply forall_mono=> x. by rewrite -plain.
Qed.
Global Instance exist_plain {A} (Ψ : A PROP) :
( x, Plain (Ψ x)) Plain ( x, Ψ x).
Proof.
intros. rewrite /Plain -plainly_exist_2. apply exist_mono=> x. by rewrite -plain.
Qed.
Global Instance impl_plain P Q : Absorbing P Plain P Plain Q Plain (P Q).
Proof.
intros. by rewrite /Plain {2}(plain P) -plainly_impl_plainly -(plain Q)
(plainly_into_absorbingly P) absorbing.
Qed.
Global Instance wand_plain P Q :
Plain P Plain Q Absorbing Q Plain (P -∗ Q).
Proof.
intros. rewrite /Plain {2}(plain P). trans ( ( P Q))%I.
- rewrite -plainly_impl_plainly impl_wand_affinely_plainly -(plain Q).
by rewrite affinely_plainly_elim.
- apply plainly_mono, wand_intro_l. by rewrite sep_and impl_elim_r.
Qed.
Global Instance sep_plain P Q : Plain P Plain Q Plain (P Q).
Proof. intros. by rewrite /Plain -plainly_sep_2 -!plain. Qed.
Global Instance plainly_plain P : Plain ( P).
Proof. by rewrite /Plain plainly_idemp. Qed.
Global Instance persistently_plain P : Plain P Plain (<pers> P).
Proof.
rewrite /Plain=> HP. rewrite {1}HP plainly_persistently_elim persistently_elim_plainly //.
Qed.
Global Instance affinely_plain P : Plain P Plain (<affine> P).
Proof. rewrite /bi_affinely. apply _. Qed.
Global Instance intuitionistically_plain P : Plain P Plain ( P).
Proof. rewrite /bi_intuitionistically. apply _. Qed.
Global Instance absorbingly_plain P : Plain P Plain (<absorb> P).
Proof. rewrite /bi_absorbingly. apply _. Qed.
Global Instance from_option_plain {A} P (Ψ : A PROP) (mx : option A) :
( x, Plain (Ψ x)) Plain P Plain (from_option Ψ P mx).
Proof. destruct mx; apply _. Qed.
Global Instance big_sepL_nil_plain {A} (Φ : nat A PROP) :
Plain ([ list] kx [], Φ k x).
Proof. simpl; apply _. Qed.
Global Instance big_sepL_plain {A} (Φ : nat A PROP) l :
( k x, Plain (Φ k x)) Plain ([ list] kx l, Φ k x).
Proof. revert Φ. induction l as [|x l IH]=> Φ ? /=; apply _. Qed.
Global Instance big_andL_nil_plain {A} (Φ : nat A PROP) :
Plain ([ list] kx [], Φ k x).
Proof. simpl; apply _. Qed.
Global Instance big_andL_plain {A} (Φ : nat A PROP) l :
( k x, Plain (Φ k x)) Plain ([ list] kx l, Φ k x).
Proof. revert Φ. induction l as [|x l IH]=> Φ ? /=; apply _. Qed.
Global Instance big_orL_nil_plain {A} (Φ : nat A PROP) :
Plain ([ list] kx [], Φ k x).
Proof. simpl; apply _. Qed.
Global Instance big_orL_plain {A} (Φ : nat A PROP) l :
( k x, Plain (Φ k x)) Plain ([ list] kx l, Φ k x).
Proof. revert Φ. induction l as [|x l IH]=> Φ ? /=; apply _. Qed.
Global Instance big_sepL2_nil_plain {A B}
(Φ : nat A B PROP) :
Plain ([ list] ky1;y2 []; [], Φ k y1 y2).
Proof. simpl; apply _. Qed.
Global Instance big_sepL2_plain {A B} (Φ : nat A B PROP) l1 l2 :
( k x1 x2, Plain (Φ k x1 x2))
Plain ([ list] ky1;y2 l1;l2, Φ k y1 y2).
Proof. rewrite big_sepL2_alt. apply _. Qed.
Global Instance big_sepM_empty_plain `{Countable K} {A} (Φ : K A PROP) :
Plain ([ map] kx , Φ k x).
Proof. rewrite big_opM_empty. apply _. Qed.
Global Instance big_sepM_plain `{Countable K} {A} (Φ : K A PROP) m :
( k x, Plain (Φ k x)) Plain ([ map] kx m, Φ k x).
Proof.
induction m using map_ind;
[rewrite big_opM_empty|rewrite big_opM_insert //]; apply _.
Qed.
Global Instance big_sepM2_empty_plain `{Countable K}
{A B} (Φ : K A B PROP) :
Plain ([ map] kx1;x2 ;, Φ k x1 x2).
Proof. rewrite big_sepM2_empty. apply _. Qed.
Global Instance big_sepM2_plain `{Countable K}
{A B} (Φ : K A B PROP) m1 m2 :
( k x1 x2, Plain (Φ k x1 x2))
Plain ([ map] kx1;x2 m1;m2, Φ k x1 x2).
Proof. intros. rewrite big_sepM2_alt. apply _. Qed.
Global Instance big_sepS_empty_plain `{Countable A} (Φ : A PROP) :
Plain ([ set] x , Φ x).
Proof. rewrite big_opS_empty. apply _. Qed.
Global Instance big_sepS_plain `{Countable A} (Φ : A PROP) X :
( x, Plain (Φ x)) Plain ([ set] x X, Φ x).
Proof.
induction X using set_ind_L;
[rewrite big_opS_empty|rewrite big_opS_insert //]; apply _.
Qed.
Global Instance big_sepMS_empty_plain `{Countable A} (Φ : A PROP) :
Plain ([ mset] x , Φ x).
Proof. rewrite big_opMS_empty. apply _. Qed.
Global Instance big_sepMS_plain `{Countable A} (Φ : A PROP) X :
( x, Plain (Φ x)) Plain ([ mset] x X, Φ x).
Proof.
induction X using gmultiset_ind;
[rewrite big_opMS_empty|rewrite big_opMS_insert]; apply _.
Qed.
Global Instance plainly_timeless P `{!BiPlainlyExist PROP} :
Timeless P Timeless ( P).
Proof.
intros. rewrite /Timeless /bi_except_0 later_plainly_1.
by rewrite (timeless P) /bi_except_0 plainly_or {1}plainly_elim.
Qed.
(* Interaction with equality *)
Section internal_eq.
Context `{!BiInternalEq PROP}.
Lemma plainly_internal_eq {A:ofe} (a b : A) : (a b) ⊣⊢@{PROP} a b.
Proof.
apply (anti_symm ()).
{ by rewrite plainly_elim. }
apply (internal_eq_rewrite' a b (λ b, (a b))%I); [solve_proper|done|].
rewrite -(internal_eq_refl True%I a) plainly_pure; auto.
Qed.
Global Instance internal_eq_plain {A : ofe} (a b : A) :
Plain (PROP:=PROP) (a b).
Proof. by intros; rewrite /Plain plainly_internal_eq. Qed.
End internal_eq.
Section prop_ext.
Context `{!BiInternalEq PROP, !BiPropExt PROP}.
Lemma prop_ext P Q : P Q ⊣⊢ (P ∗-∗ Q).
Proof.
apply (anti_symm ()); last exact: prop_ext_2.
apply (internal_eq_rewrite' P Q (λ Q, (P ∗-∗ Q))%I);
[ solve_proper | done | ].
rewrite (plainly_emp_intro (P Q)).
apply plainly_mono, wand_iff_refl.
Qed.
Lemma plainly_alt P : P ⊣⊢ <affine> P emp.
Proof.
rewrite -plainly_affinely_elim. apply (anti_symm ()).
- rewrite prop_ext. apply plainly_mono, and_intro; apply wand_intro_l.
+ by rewrite affinely_elim_emp left_id.
+ by rewrite left_id.
- rewrite internal_eq_sym (internal_eq_rewrite _ _ plainly).
by rewrite -plainly_True_emp plainly_pure True_impl.
Qed.
Lemma plainly_alt_absorbing P `{!Absorbing P} : P ⊣⊢ P True.
Proof.
apply (anti_symm ()).
- rewrite prop_ext. apply plainly_mono, and_intro; apply wand_intro_l; auto.
- rewrite internal_eq_sym (internal_eq_rewrite _ _ plainly).
by rewrite plainly_pure True_impl.
Qed.
Lemma plainly_True_alt P : (True -∗ P) ⊣⊢ P True.
Proof.
apply (anti_symm ()).
- rewrite prop_ext. apply plainly_mono, and_intro; apply wand_intro_l; auto.
by rewrite wand_elim_r.
- rewrite internal_eq_sym (internal_eq_rewrite _ _
(λ Q, (True -∗ Q))%I ltac:(shelve)); last solve_proper.
by rewrite -entails_wand // -(plainly_emp_intro True) True_impl.
Qed.
(* This proof uses [BiPlainlyExist] and [BiLöb] via [plainly_timeless] and
[wand_timeless]. *)
Global Instance internal_eq_timeless `{!BiPlainlyExist PROP, !BiLöb PROP}
`{!Timeless P} `{!Timeless Q} :
Timeless (PROP := PROP) (P Q).
Proof. rewrite prop_ext. apply _. Qed.
End prop_ext.
(* Interaction with ▷ *)
Lemma later_plainly P : P ⊣⊢ P.
Proof. apply (anti_symm _); auto using later_plainly_1, later_plainly_2. Qed.
Lemma laterN_plainly n P : ▷^n P ⊣⊢ ▷^n P.
Proof. induction n as [|n IH]; simpl; auto. by rewrite IH later_plainly. Qed.
Lemma later_plainly_if p P : ?p P ⊣⊢ ?p P.
Proof. destruct p; simpl; auto using later_plainly. Qed.
Lemma laterN_plainly_if n p P : ▷^n ?p P ⊣⊢ ?p (▷^n P).
Proof. destruct p; simpl; auto using laterN_plainly. Qed.
Lemma except_0_plainly_1 P : P P.
Proof. by rewrite /bi_except_0 -plainly_or_2 -later_plainly plainly_pure. Qed.
Lemma except_0_plainly `{!BiPlainlyExist PROP} P : P ⊣⊢ P.
Proof. by rewrite /bi_except_0 plainly_or -later_plainly plainly_pure. Qed.
Global Instance later_plain P : Plain P Plain ( P).
Proof. intros. by rewrite /Plain -later_plainly {1}(plain P). Qed.
Global Instance laterN_plain n P : Plain P Plain (▷^n P).
Proof. induction n; apply _. Qed.
Global Instance except_0_plain P : Plain P Plain ( P).
Proof. rewrite /bi_except_0; apply _. Qed.
End plainly_derived.
(* When declared as an actual instance, [plain_persistent] will cause
failing proof searches to take exponential time, as Coq will try to
apply it the instance at any node in the proof search tree.
To avoid that, we declare it using a [Hint Immediate], so that it will
only be used at the leaves of the proof search tree, i.e. when the
premise of the hint can be derived from just the current context. *)
Global Hint Immediate plain_persistent : typeclass_instances.
From stdpp Require Export telescopes.
From iris.bi Require Export bi.
From iris.prelude Require Import options.
Import bi.
(* This cannot import the proofmode because it is imported by the proofmode! *)
(** Telescopic quantifiers *)
Definition bi_texist {PROP : bi} {TT : tele@{Quant}} (Ψ : TT PROP) : PROP :=
tele_fold (@bi_exist PROP) (λ x, x) (tele_bind Ψ).
Global Arguments bi_texist {_ !_} _ /.
Definition bi_tforall {PROP : bi} {TT : tele@{Quant}} (Ψ : TT PROP) : PROP :=
tele_fold (@bi_forall PROP) (λ x, x) (tele_bind Ψ).
Global Arguments bi_tforall {_ !_} _ /.
Notation "'∃..' x .. y , P" := (bi_texist (λ x, .. (bi_texist (λ y, P)) .. )%I)
(at level 200, x binder, y binder, right associativity,
format "∃.. x .. y , P") : bi_scope.
Notation "'∀..' x .. y , P" := (bi_tforall (λ x, .. (bi_tforall (λ y, P)) .. )%I)
(at level 200, x binder, y binder, right associativity,
format "∀.. x .. y , P") : bi_scope.
Section telescopes.
Context {PROP : bi} {TT : tele@{Quant}}.
Implicit Types Ψ : TT PROP.
Lemma bi_tforall_forall Ψ : bi_tforall Ψ ⊣⊢ bi_forall Ψ.
Proof.
symmetry. unfold bi_tforall. induction TT as [|X ft IH].
- simpl. apply (anti_symm _).
+ by rewrite (forall_elim TargO).
+ rewrite -forall_intro; first done.
intros p. rewrite (tele_arg_O_inv p) /= //.
- simpl. apply (anti_symm _); apply forall_intro; intros a.
+ rewrite /= -IH. apply forall_intro; intros p.
by rewrite (forall_elim (TargS a p)).
+ destruct a=> /=.
setoid_rewrite <- IH.
rewrite 2!forall_elim. done.
Qed.
Lemma bi_texist_exist Ψ : bi_texist Ψ ⊣⊢ bi_exist Ψ.
Proof.
symmetry. unfold bi_texist. induction TT as [|X ft IH].
- simpl. apply (anti_symm _).
+ apply exist_elim; intros p.
rewrite (tele_arg_O_inv p) //.
+ by rewrite -(exist_intro TargO).
- simpl. apply (anti_symm _); apply exist_elim.
+ intros p. destruct p => /=.
by rewrite -exist_intro -IH -exist_intro.
+ intros x.
rewrite /= -IH. apply exist_elim; intros p.
by rewrite -(exist_intro (TargS x p)).
Qed.
Global Instance bi_tforall_ne n :
Proper (pointwise_relation _ (dist n) ==> dist n) (@bi_tforall PROP TT).
Proof. intros ?? EQ. rewrite !bi_tforall_forall. rewrite EQ //. Qed.
Global Instance bi_tforall_proper :
Proper (pointwise_relation _ (⊣⊢) ==> (⊣⊢)) (@bi_tforall PROP TT).
Proof. intros ?? EQ. rewrite !bi_tforall_forall. rewrite EQ //. Qed.
Global Instance bi_texist_ne n :
Proper (pointwise_relation _ (dist n) ==> dist n) (@bi_texist PROP TT).
Proof. intros ?? EQ. rewrite !bi_texist_exist. rewrite EQ //. Qed.
Global Instance bi_texist_proper :
Proper (pointwise_relation _ (⊣⊢) ==> (⊣⊢)) (@bi_texist PROP TT).
Proof. intros ?? EQ. rewrite !bi_texist_exist. rewrite EQ //. Qed.
Global Instance bi_tforall_absorbing Ψ :
( x, Absorbing (Ψ x)) Absorbing (.. x, Ψ x).
Proof. rewrite bi_tforall_forall. apply _. Qed.
Global Instance bi_tforall_persistent `{!BiPersistentlyForall PROP} Ψ :
( x, Persistent (Ψ x)) Persistent (.. x, Ψ x).
Proof. rewrite bi_tforall_forall. apply _. Qed.
Global Instance bi_texist_affine Ψ :
( x, Affine (Ψ x)) Affine (.. x, Ψ x).
Proof. rewrite bi_texist_exist. apply _. Qed.
Global Instance bi_texist_absorbing Ψ :
( x, Absorbing (Ψ x)) Absorbing (.. x, Ψ x).
Proof. rewrite bi_texist_exist. apply _. Qed.
Global Instance bi_texist_persistent Ψ :
( x, Persistent (Ψ x)) Persistent (.. x, Ψ x).
Proof. rewrite bi_texist_exist. apply _. Qed.
Global Instance bi_tforall_timeless Ψ :
( x, Timeless (Ψ x)) Timeless (.. x, Ψ x).
Proof. rewrite bi_tforall_forall. apply _. Qed.
Global Instance bi_texist_timeless Ψ :
( x, Timeless (Ψ x)) Timeless (.. x, Ψ x).
Proof. rewrite bi_texist_exist. apply _. Qed.
End telescopes.
From stdpp Require Import coPset.
From iris.bi Require Import interface derived_laws_later big_op plainly.
From iris.prelude Require Import options.
Import interface.bi derived_laws.bi derived_laws_later.bi.
(* We enable primitive projections in this file to improve the performance of the Iris proofmode:
primitive projections for the bi-records makes the proofmode faster.
*)
Local Set Primitive Projections.
(* The sections add extra BI assumptions, which is only picked up with "Type"*. *)
Set Default Proof Using "Type*".
(* We first define operational type classes for the notations, and then later
bundle these operational type classes with the laws. *)
Class BUpd (PROP : Type) : Type := bupd : PROP PROP.
Global Instance : Params (@bupd) 2 := {}.
Global Hint Mode BUpd ! : typeclass_instances.
Global Arguments bupd {_}%_type_scope {_} _%_bi_scope.
Global Typeclasses Opaque bupd.
Notation "|==> Q" := (bupd Q) : bi_scope.
Notation "P ==∗ Q" := (P -∗ |==> Q)%I : bi_scope.
Notation "P ==∗ Q" := (P -∗ |==> Q) : stdpp_scope.
Class FUpd (PROP : Type) : Type := fupd : coPset coPset PROP PROP.
Global Instance: Params (@fupd) 4 := {}.
Global Hint Mode FUpd ! : typeclass_instances.
Global Arguments fupd {_}%_type_scope {_} _ _ _%_bi_scope.
Global Typeclasses Opaque fupd.
Notation "|={ E1 , E2 }=> Q" := (fupd E1 E2 Q) : bi_scope.
Notation "P ={ E1 , E2 }=∗ Q" := (P -∗ |={E1,E2}=> Q)%I : bi_scope.
Notation "P ={ E1 , E2 }=∗ Q" := (P -∗ |={E1,E2}=> Q) : stdpp_scope.
Notation "|={ E }=> Q" := (fupd E E Q) : bi_scope.
Notation "P ={ E }=∗ Q" := (P -∗ |={E}=> Q)%I : bi_scope.
Notation "P ={ E }=∗ Q" := (P -∗ |={E}=> Q) : stdpp_scope.
(** * Step-taking fancy updates. *)
(** These have two masks, but they are different than the two masks of a
mask-changing update: in [|={Eo}[Ei]▷=> Q], the first mask [Eo] ("outer
mask") holds at the beginning and the end; the second mask [Ei] ("inner
mask") holds around each ▷. This is also why we use a different notation
than for the two masks of a mask-changing updates. *)
Notation "|={ Eo } [ Ei ]▷=> Q" := (|={Eo,Ei}=> |={Ei,Eo}=> Q)%I : bi_scope.
Notation "P ={ Eo } [ Ei ]▷=∗ Q" := (P -∗ |={Eo}[Ei]▷=> Q)%I : bi_scope.
Notation "P ={ Eo } [ Ei ]▷=∗ Q" := (P -∗ |={Eo}[Ei]▷=> Q) : stdpp_scope.
Notation "|={ E }▷=> Q" := (|={E}[E]▷=> Q)%I : bi_scope.
Notation "P ={ E }▷=∗ Q" := (P ={E}[E]▷=∗ Q)%I : bi_scope.
Notation "P ={ E }▷=∗ Q" := (P ={E}[E]▷=∗ Q) : stdpp_scope.
(** For the iterated version, in principle there are 4 masks: "outer" and
"inner" of [|={Eo}[Ei]▷=>], as well as "begin" and "end" masks [E1] and [E2]
that could potentially differ from [Eo]. The latter can be obtained from
this notation by adding normal mask-changing update modalities: [
|={E1,Eo}=> |={Eo}[Ei]▷=>^n |={Eo,E2}=> Q] *)
Notation "|={ Eo } [ Ei ]▷=>^ n Q" := (Nat.iter n (λ P, |={Eo}[Ei]▷=> P) Q)%I : bi_scope.
Notation "P ={ Eo } [ Ei ]▷=∗^ n Q" := (P -∗ |={Eo}[Ei]▷=>^n Q)%I : bi_scope.
Notation "P ={ Eo } [ Ei ]▷=∗^ n Q" := (P -∗ |={Eo}[Ei]▷=>^n Q) : stdpp_scope.
Notation "|={ E }▷=>^ n Q" := (|={E}[E]▷=>^n Q)%I : bi_scope.
Notation "P ={ E }▷=∗^ n Q" := (P ={E}[E]▷=∗^n Q)%I : bi_scope.
Notation "P ={ E }▷=∗^ n Q" := (P ={E}[E]▷=∗^n Q) : stdpp_scope.
(** Bundled versions *)
(* Mixins allow us to create instances easily without having to use Program *)
Record BiBUpdMixin (PROP : bi) `(BUpd PROP) := {
bi_bupd_mixin_bupd_ne : NonExpansive (bupd (PROP:=PROP));
bi_bupd_mixin_bupd_intro (P : PROP) : P |==> P;
bi_bupd_mixin_bupd_mono (P Q : PROP) : (P Q) (|==> P) |==> Q;
bi_bupd_mixin_bupd_trans (P : PROP) : (|==> |==> P) |==> P;
bi_bupd_mixin_bupd_frame_r (P R : PROP) : (|==> P) R |==> P R;
}.
Record BiFUpdMixin (PROP : bi) `(FUpd PROP) := {
bi_fupd_mixin_fupd_ne E1 E2 :
NonExpansive (fupd (PROP:=PROP) E1 E2);
bi_fupd_mixin_fupd_mask_subseteq E1 E2 :
E2 E1 ⊢@{PROP} |={E1,E2}=> |={E2,E1}=> emp;
bi_fupd_mixin_except_0_fupd E1 E2 (P : PROP) :
(|={E1,E2}=> P) |={E1,E2}=> P;
bi_fupd_mixin_fupd_mono E1 E2 (P Q : PROP) :
(P Q) (|={E1,E2}=> P) |={E1,E2}=> Q;
bi_fupd_mixin_fupd_trans E1 E2 E3 (P : PROP) :
(|={E1,E2}=> |={E2,E3}=> P) |={E1,E3}=> P;
bi_fupd_mixin_fupd_mask_frame_r' E1 E2 Ef (P : PROP) :
E1 ## Ef (|={E1,E2}=> E2 ## Ef P) |={E1 Ef,E2 Ef}=> P;
bi_fupd_mixin_fupd_frame_r E1 E2 (P R : PROP) :
(|={E1,E2}=> P) R |={E1,E2}=> P R;
}.
Class BiBUpd (PROP : bi) := {
#[global] bi_bupd_bupd :: BUpd PROP;
bi_bupd_mixin : BiBUpdMixin PROP bi_bupd_bupd;
}.
Global Hint Mode BiBUpd ! : typeclass_instances.
Global Arguments bi_bupd_bupd : simpl never.
Class BiFUpd (PROP : bi) := {
#[global] bi_fupd_fupd :: FUpd PROP;
bi_fupd_mixin : BiFUpdMixin PROP bi_fupd_fupd;
}.
Global Hint Mode BiFUpd ! : typeclass_instances.
Global Arguments bi_fupd_fupd : simpl never.
Class BiBUpdFUpd (PROP : bi) `{BiBUpd PROP, BiFUpd PROP} :=
bupd_fupd E (P : PROP) : (|==> P) |={E}=> P.
Global Hint Mode BiBUpdFUpd ! - - : typeclass_instances.
Class BiBUpdPlainly (PROP : bi) `{!BiBUpd PROP, !BiPlainly PROP} :=
bupd_plainly (P : PROP) : (|==> P) P.
Global Hint Mode BiBUpdPlainly ! - - : typeclass_instances.
(** These rules for the interaction between the [■] and [|={E1,E2=>] modalities
only make sense for affine logics. From the axioms below, one could derive
[■ P ={E}=∗ P] (see the lemma [fupd_plainly_elim]), which in turn gives
[True ={E}=∗ emp]. *)
Class BiFUpdPlainly (PROP : bi) `{!BiFUpd PROP, !BiPlainly PROP} := {
(** Given a mask-changing non-persistent view shift ending in a plain
proposition, and given the precondition of the view shift, we can eliminate
the view shift without consuming the precondition and without changing the
mask. *)
fupd_plainly_keep_l E E' (P R : PROP) :
(R ={E,E'}=∗ P) R |={E}=> P R;
(** Later "almost" commutes with fancy updates over plain propositions. It
commutes "almost" because of the ◇ modality, which is needed in the definition
of fancy updates so one can remove laters of timeless propositions. *)
fupd_plainly_later E (P : PROP) :
( |={E}=> P) |={E}=> P;
(** Forall quantifiers commute with fancy updates over plain propositions. *)
fupd_plainly_forall_2 E {A} (Φ : A PROP) :
( x, |={E}=> Φ x) |={E}=> x, Φ x
}.
Global Hint Mode BiBUpdFUpd ! - - : typeclass_instances.
Section bupd_laws.
Context {PROP : bi} `{!BiBUpd PROP}.
Implicit Types P : PROP.
Global Instance bupd_ne : NonExpansive (@bupd PROP _).
Proof. eapply bi_bupd_mixin_bupd_ne, bi_bupd_mixin. Qed.
Lemma bupd_intro P : P |==> P.
Proof. eapply bi_bupd_mixin_bupd_intro, bi_bupd_mixin. Qed.
Lemma bupd_mono (P Q : PROP) : (P Q) (|==> P) |==> Q.
Proof. eapply bi_bupd_mixin_bupd_mono, bi_bupd_mixin. Qed.
Lemma bupd_trans (P : PROP) : (|==> |==> P) |==> P.
Proof. eapply bi_bupd_mixin_bupd_trans, bi_bupd_mixin. Qed.
Lemma bupd_frame_r (P R : PROP) : (|==> P) R |==> P R.
Proof. eapply bi_bupd_mixin_bupd_frame_r, bi_bupd_mixin. Qed.
End bupd_laws.
Section fupd_laws.
Context {PROP : bi} `{!BiFUpd PROP}.
Implicit Types P : PROP.
Global Instance fupd_ne E1 E2 : NonExpansive (@fupd PROP _ E1 E2).
Proof. eapply bi_fupd_mixin_fupd_ne, bi_fupd_mixin. Qed.
(** [iMod] with this lemma is useful to change the current mask to a subset,
and obtain a fupd for changing it back. For the case where you want to get rid
of a mask-changing fupd in the goal, [iApply fupd_mask_intro] avoids having to
specify the mask. *)
Lemma fupd_mask_subseteq {E1} E2 : E2 E1 ⊢@{PROP} |={E1,E2}=> |={E2,E1}=> emp.
Proof. eapply bi_fupd_mixin_fupd_mask_subseteq, bi_fupd_mixin. Qed.
Lemma except_0_fupd E1 E2 (P : PROP) : (|={E1,E2}=> P) |={E1,E2}=> P.
Proof. eapply bi_fupd_mixin_except_0_fupd, bi_fupd_mixin. Qed.
Lemma fupd_mono E1 E2 (P Q : PROP) : (P Q) (|={E1,E2}=> P) |={E1,E2}=> Q.
Proof. eapply bi_fupd_mixin_fupd_mono, bi_fupd_mixin. Qed.
Lemma fupd_trans E1 E2 E3 (P : PROP) : (|={E1,E2}=> |={E2,E3}=> P) |={E1,E3}=> P.
Proof. eapply bi_fupd_mixin_fupd_trans, bi_fupd_mixin. Qed.
Lemma fupd_mask_frame_r' E1 E2 Ef (P : PROP) :
E1 ## Ef (|={E1,E2}=> E2 ## Ef P) |={E1 Ef,E2 Ef}=> P.
Proof. eapply bi_fupd_mixin_fupd_mask_frame_r', bi_fupd_mixin. Qed.
Lemma fupd_frame_r E1 E2 (P R : PROP) : (|={E1,E2}=> P) R |={E1,E2}=> P R.
Proof. eapply bi_fupd_mixin_fupd_frame_r, bi_fupd_mixin. Qed.
End fupd_laws.
Section bupd_derived.
Context {PROP : bi} `{!BiBUpd PROP}.
Implicit Types P Q R : PROP.
Global Instance bupd_proper :
Proper (() ==> ()) (bupd (PROP:=PROP)) := ne_proper _.
(** BUpd derived rules *)
Global Instance bupd_mono' : Proper (() ==> ()) (bupd (PROP:=PROP)).
Proof. intros P Q; apply bupd_mono. Qed.
Global Instance bupd_flip_mono' : Proper (flip () ==> flip ()) (bupd (PROP:=PROP)).
Proof. intros P Q; apply bupd_mono. Qed.
Lemma bupd_frame_l R Q : (R |==> Q) |==> R Q.
Proof. rewrite !(comm _ R); apply bupd_frame_r. Qed.
Lemma bupd_wand_l P Q : (P -∗ Q) (|==> P) |==> Q.
Proof. by rewrite bupd_frame_l wand_elim_l. Qed.
Lemma bupd_wand_r P Q : (|==> P) (P -∗ Q) |==> Q.
Proof. by rewrite bupd_frame_r wand_elim_r. Qed.
Lemma bupd_sep P Q : (|==> P) (|==> Q) |==> P Q.
Proof. by rewrite bupd_frame_r bupd_frame_l bupd_trans. Qed.
Lemma bupd_idemp P : (|==> |==> P) ⊣⊢ |==> P.
Proof.
apply: anti_symm.
- apply bupd_trans.
- apply bupd_intro.
Qed.
Global Instance bupd_sep_homomorphism :
MonoidHomomorphism bi_sep bi_sep (flip ()) (bupd (PROP:=PROP)).
Proof. split; [split|]; try apply _; [apply bupd_sep | apply bupd_intro]. Qed.
Lemma bupd_or P Q : (|==> P) (|==> Q) |==> (P Q).
Proof. apply or_elim; apply bupd_mono; [ apply or_intro_l | apply or_intro_r ]. Qed.
Global Instance bupd_or_homomorphism :
MonoidHomomorphism bi_or bi_or (flip ()) (bupd (PROP:=PROP)).
Proof. split; [split|]; try apply _; [apply bupd_or | apply bupd_intro]. Qed.
Lemma bupd_and P Q : (|==> (P Q)) (|==> P) (|==> Q).
Proof. apply and_intro; apply bupd_mono; [apply and_elim_l | apply and_elim_r]. Qed.
Lemma bupd_exist A (Φ : A PROP) : ( x : A, |==> Φ x) |==> x : A, Φ x.
Proof. apply exist_elim=> a. by rewrite -(exist_intro a). Qed.
Lemma bupd_forall A (Φ : A PROP) : (|==> x : A, Φ x) x : A, |==> Φ x.
Proof. apply forall_intro=> a. by rewrite -(forall_elim a). Qed.
Lemma big_sepL_bupd {A} (Φ : nat A PROP) l :
([ list] kx l, |==> Φ k x) |==> [ list] kx l, Φ k x.
Proof. by rewrite (big_opL_commute _). Qed.
Lemma big_sepM_bupd {A} `{Countable K} (Φ : K A PROP) l :
([ map] kx l, |==> Φ k x) |==> [ map] kx l, Φ k x.
Proof. by rewrite (big_opM_commute _). Qed.
Lemma big_sepS_bupd `{Countable A} (Φ : A PROP) l :
([ set] x l, |==> Φ x) |==> [ set] x l, Φ x.
Proof. by rewrite (big_opS_commute _). Qed.
Lemma big_sepMS_bupd `{Countable A} (Φ : A PROP) l :
([ mset] x l, |==> Φ x) |==> [ mset] x l, Φ x.
Proof. by rewrite (big_opMS_commute _). Qed.
Lemma except_0_bupd P : (|==> P) (|==> P).
Proof.
rewrite /bi_except_0. apply or_elim; eauto using bupd_mono, or_intro_r.
by rewrite -bupd_intro -or_intro_l.
Qed.
Global Instance bupd_absorbing P :
Absorbing P Absorbing (|==> P).
Proof. rewrite /Absorbing /bi_absorbingly bupd_frame_l =>-> //. Qed.
Section bupd_plainly.
Context `{!BiPlainly PROP, !BiBUpdPlainly PROP}.
Lemma bupd_elim P `{!Plain P} : (|==> P) P.
Proof. by rewrite {1}(plain P) bupd_plainly. Qed.
Lemma bupd_plain_forall {A} (Φ : A PROP) `{ x, Plain (Φ x)} :
(|==> x, Φ x) ⊣⊢ ( x, |==> Φ x).
Proof.
apply (anti_symm _).
- apply bupd_forall.
- rewrite -bupd_intro. apply forall_intro=> x.
by rewrite (forall_elim x) bupd_elim.
Qed.
Global Instance bupd_plain P : Plain P Plain (|==> P).
Proof.
intros. rewrite /Plain. rewrite {1}(plain P) {1}bupd_elim.
by rewrite -bupd_intro.
Qed.
End bupd_plainly.
End bupd_derived.
Section fupd_derived.
Context {PROP : bi} `{!BiFUpd PROP}.
Implicit Types P Q R : PROP.
Global Instance fupd_proper E1 E2 :
Proper (() ==> ()) (fupd (PROP:=PROP) E1 E2) := ne_proper _.
(** FUpd derived rules *)
Global Instance fupd_mono' E1 E2 : Proper (() ==> ()) (fupd (PROP:=PROP) E1 E2).
Proof. intros P Q; apply fupd_mono. Qed.
Global Instance fupd_flip_mono' E1 E2 :
Proper (flip () ==> flip ()) (fupd (PROP:=PROP) E1 E2).
Proof. intros P Q; apply fupd_mono. Qed.
Lemma fupd_mask_intro_subseteq E1 E2 P :
E2 E1 P |={E1,E2}=> |={E2,E1}=> P.
Proof.
intros HE.
apply wand_entails', wand_intro_r.
rewrite fupd_mask_subseteq; last exact: HE.
rewrite !fupd_frame_r. rewrite left_id. done.
Qed.
Lemma fupd_intro E P : P |={E}=> P.
Proof. by rewrite {1}(fupd_mask_intro_subseteq E E P) // fupd_trans. Qed.
Lemma fupd_except_0 E1 E2 P : (|={E1,E2}=> P) |={E1,E2}=> P.
Proof. by rewrite {1}(fupd_intro E2 P) except_0_fupd fupd_trans. Qed.
Lemma fupd_idemp E P : (|={E}=> |={E}=> P) ⊣⊢ |={E}=> P.
Proof.
apply: anti_symm.
- apply fupd_trans.
- apply fupd_intro.
Qed.
(** Weaken the first mask of the goal from [E1] to [E2].
This lemma is intended to be [iApply]ed.
However, usually you can [iMod (fupd_mask_subseteq E2)] instead and that
will be slightly more convenient. *)
Lemma fupd_mask_weaken {E1} E2 {E3 P} :
E2 E1
((|={E2,E1}=> emp) ={E2,E3}=∗ P) |={E1,E3}=> P.
Proof.
intros HE.
apply wand_entails', wand_intro_r.
rewrite {1}(fupd_mask_subseteq E2) //.
rewrite fupd_frame_r. by rewrite wand_elim_r fupd_trans.
Qed.
(** Introduction lemma for a mask-changing fupd.
This lemma is intended to be [iApply]ed. *)
Lemma fupd_mask_intro E1 E2 P :
E2 E1
((|={E2,E1}=> emp) -∗ P) |={E1,E2}=> P.
Proof.
intros. etrans; [|by apply fupd_mask_weaken]. by rewrite -fupd_intro.
Qed.
Lemma fupd_mask_intro_discard E1 E2 P `{!Absorbing P} :
E2 E1 P |={E1,E2}=> P.
Proof.
intros. etrans; [|by apply fupd_mask_intro].
apply wand_intro_r. rewrite sep_elim_l. done.
Qed.
Lemma fupd_frame_l E1 E2 R Q : (R |={E1,E2}=> Q) |={E1,E2}=> R Q.
Proof. rewrite !(comm _ R); apply fupd_frame_r. Qed.
Lemma fupd_wand_l E1 E2 P Q : (P -∗ Q) (|={E1,E2}=> P) |={E1,E2}=> Q.
Proof. by rewrite fupd_frame_l wand_elim_l. Qed.
Lemma fupd_wand_r E1 E2 P Q : (|={E1,E2}=> P) (P -∗ Q) |={E1,E2}=> Q.
Proof. by rewrite fupd_frame_r wand_elim_r. Qed.
Global Instance fupd_absorbing E1 E2 P :
Absorbing P Absorbing (|={E1,E2}=> P).
Proof. rewrite /Absorbing /bi_absorbingly fupd_frame_l =>-> //. Qed.
Lemma fupd_trans_frame E1 E2 E3 P Q :
((Q ={E2,E3}=∗ emp) |={E1,E2}=> (Q P)) |={E1,E3}=> P.
Proof.
rewrite fupd_frame_l assoc -(comm _ Q) wand_elim_r.
by rewrite fupd_frame_r left_id fupd_trans.
Qed.
Lemma fupd_elim E1 E2 E3 P Q :
(Q (|={E2,E3}=> P)) (|={E1,E2}=> Q) (|={E1,E3}=> P).
Proof. intros ->. rewrite fupd_trans //. Qed.
Lemma fupd_mask_frame_r E1 E2 Ef P :
E1 ## Ef (|={E1,E2}=> P) |={E1 Ef,E2 Ef}=> P.
Proof.
intros ?. rewrite -fupd_mask_frame_r' //. f_equiv.
apply impl_intro_l, and_elim_r.
Qed.
Lemma fupd_mask_mono E1 E2 P : E1 E2 (|={E1}=> P) |={E2}=> P.
Proof.
intros (Ef&->&?)%subseteq_disjoint_union_L. by apply fupd_mask_frame_r.
Qed.
(** How to apply an arbitrary mask-changing view shift when having
an arbitrary mask. *)
Lemma fupd_mask_frame E E' E1 E2 P :
E1 E
(|={E1,E2}=> |={E2 (E E1),E'}=> P) (|={E,E'}=> P).
Proof.
intros ?. rewrite (fupd_mask_frame_r _ _ (E E1)); last set_solver.
rewrite fupd_trans.
by replace (E1 E E1) with E by (by apply union_difference_L).
Qed.
(* A variant of [fupd_mask_frame] that works well for accessors: Tailored to
eliminate updates of the form [|={E1,E1∖E2}=> Q] and provides a way to
transform the closing view shift instead of letting you prove the same
side-conditions twice. *)
Lemma fupd_mask_frame_acc E E' E1(*Eo*) E2(*Em*) P Q :
E1 E
(|={E1,E1E2}=> Q) -∗
(Q -∗ |={EE2,E'}=> ( R, (|={E1E2,E1}=> R) -∗ |={EE2,E}=> R) -∗ P) -∗
(|={E,E'}=> P).
Proof.
intros HE. apply entails_wand, wand_intro_r. rewrite fupd_frame_r.
rewrite wand_elim_r. clear Q.
rewrite -(fupd_mask_frame E E'); first apply fupd_mono; last done.
(* The most horrible way to apply fupd_intro_mask *)
rewrite -[X in (X _)](right_id emp%I).
rewrite (fupd_mask_intro_subseteq (E1 E2 E E1) (E E2) emp); last first.
{ rewrite {1}(union_difference_L _ _ HE). set_solver. }
rewrite fupd_frame_l fupd_frame_r. apply fupd_elim.
apply fupd_mono.
eapply wand_apply;
last (apply sep_mono; first reflexivity); first reflexivity.
apply forall_intro=>R. apply wand_intro_r.
rewrite fupd_frame_r. apply fupd_elim. rewrite left_id.
rewrite (fupd_mask_frame_r _ _ (E E1)); last set_solver+.
rewrite {4}(union_difference_L _ _ HE). done.
Qed.
Lemma fupd_mask_subseteq_emptyset_difference E1 E2 :
E2 E1
⊢@{PROP} |={E1, E2}=> |={, E1E2}=> emp.
Proof.
intros ?. rewrite [in fupd E1](union_difference_L E2 E1); [|done].
rewrite (comm_L ())
-[X in fupd _ X](left_id_L () E2) -fupd_mask_frame_r; [|set_solver+].
apply fupd_mask_intro_subseteq; set_solver.
Qed.
Lemma fupd_or E1 E2 P Q :
(|={E1,E2}=> P) (|={E1,E2}=> Q) ⊢@{PROP}
(|={E1,E2}=> (P Q)).
Proof. apply or_elim; apply fupd_mono; [ apply or_intro_l | apply or_intro_r ]. Qed.
Global Instance fupd_or_homomorphism E :
MonoidHomomorphism bi_or bi_or (flip ()) (fupd (PROP:=PROP) E E).
Proof. split; [split|]; try apply _; [apply fupd_or | apply fupd_intro]. Qed.
Lemma fupd_and E1 E2 P Q :
(|={E1,E2}=> (P Q)) ⊢@{PROP} (|={E1,E2}=> P) (|={E1,E2}=> Q).
Proof. apply and_intro; apply fupd_mono; [apply and_elim_l | apply and_elim_r]. Qed.
Lemma fupd_exist E1 E2 A (Φ : A PROP) : ( x : A, |={E1, E2}=> Φ x) |={E1, E2}=> x : A, Φ x.
Proof. apply exist_elim=> a. by rewrite -(exist_intro a). Qed.
Lemma fupd_forall E1 E2 A (Φ : A PROP) : (|={E1, E2}=> x : A, Φ x) x : A, |={E1, E2}=> Φ x.
Proof. apply forall_intro=> a. by rewrite -(forall_elim a). Qed.
Lemma fupd_sep E P Q : (|={E}=> P) (|={E}=> Q) |={E}=> P Q.
Proof. by rewrite fupd_frame_r fupd_frame_l fupd_trans. Qed.
Global Instance fupd_sep_homomorphism E :
MonoidHomomorphism bi_sep bi_sep (flip ()) (fupd (PROP:=PROP) E E).
Proof. split; [split|]; try apply _; [apply fupd_sep | apply fupd_intro]. Qed.
Lemma big_sepL_fupd {A} E (Φ : nat A PROP) l :
([ list] kx l, |={E}=> Φ k x) |={E}=> [ list] kx l, Φ k x.
Proof. by rewrite (big_opL_commute _). Qed.
Lemma big_sepL2_fupd {A B} E (Φ : nat A B PROP) l1 l2 :
([ list] kx;y l1;l2, |={E}=> Φ k x y) |={E}=> [ list] kx;y l1;l2, Φ k x y.
Proof.
rewrite !big_sepL2_alt !persistent_and_affinely_sep_l.
etrans; [| by apply fupd_frame_l]. apply sep_mono_r. apply big_sepL_fupd.
Qed.
Lemma big_sepM_fupd `{Countable K} {A} E (Φ : K A PROP) m :
([ map] kx m, |={E}=> Φ k x) |={E}=> [ map] kx m, Φ k x.
Proof. by rewrite (big_opM_commute _). Qed.
Lemma big_sepS_fupd `{Countable A} E (Φ : A PROP) X :
([ set] x X, |={E}=> Φ x) |={E}=> [ set] x X, Φ x.
Proof. by rewrite (big_opS_commute _). Qed.
Lemma big_sepMS_fupd `{Countable A} E (Φ : A PROP) l :
([ mset] x l, |={E}=> Φ x) |={E}=> [ mset] x l, Φ x.
Proof. by rewrite (big_opMS_commute _). Qed.
(** Fancy updates that take a step derived rules. *)
Lemma step_fupd_wand Eo Ei P Q : (|={Eo}[Ei]▷=> P) -∗ (P -∗ Q) -∗ |={Eo}[Ei]▷=> Q.
Proof.
apply entails_wand, wand_intro_l.
by rewrite (later_intro (P -∗ Q)) fupd_frame_l -later_sep fupd_frame_l
wand_elim_l.
Qed.
Lemma step_fupd_mask_frame_r Eo Ei Ef P :
Eo ## Ef Ei ## Ef (|={Eo}[Ei]▷=> P) |={Eo Ef}[Ei Ef]▷=> P.
Proof.
intros. rewrite -fupd_mask_frame_r //. do 2 f_equiv. by apply fupd_mask_frame_r.
Qed.
Lemma step_fupd_mask_mono Eo1 Eo2 Ei1 Ei2 P :
Ei2 Ei1 Eo1 Eo2 (|={Eo1}[Ei1]▷=> P) |={Eo2}[Ei2]▷=> P.
Proof.
intros ??. rewrite -(emp_sep (|={Eo1}[Ei1]▷=> P)%I).
rewrite (fupd_mask_intro_subseteq Eo2 Eo1 emp) //.
rewrite fupd_frame_r -(fupd_trans Eo2 Eo1 Ei2). f_equiv.
rewrite fupd_frame_l -(fupd_trans Eo1 Ei1 Ei2). f_equiv.
rewrite (fupd_mask_intro_subseteq Ei1 Ei2 (|={_,_}=> emp)) //.
rewrite fupd_frame_r. f_equiv.
rewrite [X in (X _)%I]later_intro -later_sep. f_equiv.
rewrite fupd_frame_r -(fupd_trans Ei2 Ei1 Eo2). f_equiv.
rewrite fupd_frame_l -(fupd_trans Ei1 Eo1 Eo2). f_equiv.
by rewrite fupd_frame_r left_id.
Qed.
Lemma step_fupd_intro Ei Eo P : Ei Eo P |={Eo}[Ei]▷=> P.
Proof. intros. by rewrite -(step_fupd_mask_mono Ei _ Ei _) // -!fupd_intro. Qed.
Lemma step_fupd_frame_l Eo Ei R Q :
(R |={Eo}[Ei]▷=> Q) |={Eo}[Ei]▷=> (R Q).
Proof.
rewrite fupd_frame_l.
apply fupd_mono.
rewrite [P in P _ _](later_intro R) -later_sep fupd_frame_l.
by apply later_mono, fupd_mono.
Qed.
Lemma step_fupd_fupd Eo Ei P : (|={Eo}[Ei]▷=> P) ⊣⊢ (|={Eo}[Ei]▷=> |={Eo}=> P).
Proof.
apply (anti_symm ()).
- by rewrite -fupd_intro.
- by rewrite fupd_trans.
Qed.
Lemma step_fupdN_mono Eo Ei n P Q :
(P Q) (|={Eo}[Ei]▷=>^n P) (|={Eo}[Ei]▷=>^n Q).
Proof.
intros HPQ. induction n as [|n IH]=> //=. rewrite IH //.
Qed.
Lemma step_fupdN_wand Eo Ei n P Q :
(|={Eo}[Ei]▷=>^n P) -∗ (P -∗ Q) -∗ (|={Eo}[Ei]▷=>^n Q).
Proof.
apply entails_wand, wand_intro_l. induction n as [|n IH]=> /=.
{ by rewrite wand_elim_l. }
rewrite -IH -fupd_frame_l later_sep -fupd_frame_l.
by apply sep_mono; first apply later_intro.
Qed.
Lemma step_fupdN_intro Ei Eo n P : Ei Eo ▷^n P |={Eo}[Ei]▷=>^n P.
Proof.
induction n as [|n IH]=> ?; [done|].
rewrite /= -step_fupd_intro; [|done]. by rewrite IH.
Qed.
Lemma step_fupdN_S_fupd n E P :
(|={E}[]▷=>^(S n) P) ⊣⊢ (|={E}[]▷=>^(S n) |={E}=> P).
Proof.
apply (anti_symm ()); rewrite !Nat.iter_succ_r; apply step_fupdN_mono;
rewrite -step_fupd_fupd //.
Qed.
Lemma step_fupdN_frame_l Eo Ei n R Q :
(R |={Eo}[Ei]▷=>^n Q) |={Eo}[Ei]▷=>^n (R Q).
Proof.
induction n as [|n IH]; simpl; [done|].
rewrite step_fupd_frame_l IH //=.
Qed.
Lemma step_fupdN_add n m Eo Ei P :
(|={Eo}[Ei]▷=>^(n+m) P) ⊣⊢ (|={Eo}[Ei]▷=>^n |={Eo}[Ei]▷=>^m P).
Proof.
induction n as [ | n IH]; simpl; [done | by rewrite IH].
Qed.
(** The sidecondition [Ei ⊆ Eo] is needed because for [n = 0],
this lemma introduces updates in the same way as [step_fupdN_intro]
(in fact, for [n = 0] it is essentially [step_fupdN_intro], modulo laters). *)
Lemma step_fupdN_le n m Eo Ei P :
n m Ei Eo (|={Eo}[Ei]▷=>^n P) (|={Eo}[Ei]▷=>^m P).
Proof.
intros ??. replace m with ((m - n) + n) by lia.
rewrite step_fupdN_add.
rewrite -(step_fupdN_intro _ _ (m - n)); last done.
by rewrite -laterN_intro.
Qed.
Section fupd_plainly_derived.
Context `{!BiPlainly PROP, !BiFUpdPlainly PROP}.
Lemma fupd_plainly_mask E E' P : (|={E,E'}=> P) |={E}=> P.
Proof.
trans (|={E}=> P emp)%I; last by rewrite right_id.
rewrite -fupd_plainly_keep_l.
rewrite right_id left_id. done.
Qed.
Lemma fupd_plainly_elim E P : P |={E}=> P.
Proof. by rewrite (fupd_intro E ( P)) fupd_plainly_mask. Qed.
Lemma fupd_plainly_keep_r E E' P R : R (R ={E,E'}=∗ P) |={E}=> R P.
Proof. by rewrite !(comm _ R) fupd_plainly_keep_l. Qed.
Lemma fupd_plainly_laterN E n P : (▷^n |={E}=> P) |={E}=> ▷^n P.
Proof.
revert P. induction n as [|n IH]=> P /=.
{ by rewrite -except_0_intro (fupd_plainly_elim E) fupd_trans. }
rewrite -!later_laterN !laterN_later.
rewrite -plainly_idemp fupd_plainly_later.
by rewrite except_0_plainly_1 later_plainly_1 IH except_0_later.
Qed.
(** Laws for general plain propositions. *)
Lemma fupd_plain_mask E E' P `{!Plain P} : (|={E,E'}=> P) |={E}=> P.
Proof. by rewrite {1}(plain P) fupd_plainly_mask. Qed.
Lemma fupd_plain_keep_l E E' P R `{!Plain P} : (R ={E,E'}=∗ P) R |={E}=> P R.
Proof. rewrite {1}(plain P) fupd_plainly_keep_l //. Qed.
Lemma fupd_plain_keep_r E E' P R `{!Plain P} : R (R ={E,E'}=∗ P) |={E}=> R P.
Proof. by rewrite !(comm _ R) fupd_plain_keep_l. Qed.
Lemma fupd_plain_keep E E' P R `{!Plain P} : (R ={E,E'}=∗ P) -∗ R -∗ |={E}=> P R.
Proof. apply entails_wand, wand_intro_r, fupd_plain_keep_l. done. Qed.
Lemma fupd_plain_later E P `{!Plain P} : ( |={E}=> P) |={E}=> P.
Proof. by rewrite {1}(plain P) fupd_plainly_later. Qed.
Lemma fupd_plain_laterN E n P `{!Plain P} : (▷^n |={E}=> P) |={E}=> ▷^n P.
Proof. by rewrite {1}(plain P) fupd_plainly_laterN. Qed.
Lemma fupd_plain_forall_2 E {A} (Φ : A PROP) `{!∀ x, Plain (Φ x)} :
( x, |={E}=> Φ x) |={E}=> x, Φ x.
Proof.
rewrite -fupd_plainly_forall_2. apply forall_mono=> x.
by rewrite {1}(plain (Φ _)).
Qed.
Lemma fupd_plain_forall E1 E2 {A} (Φ : A PROP) `{!∀ x, Plain (Φ x)} :
E2 E1
(|={E1,E2}=> x, Φ x) ⊣⊢ ( x, |={E1,E2}=> Φ x).
Proof.
intros. apply (anti_symm _); first apply fupd_forall.
trans ( x, |={E1}=> Φ x)%I.
{ apply forall_mono=> x. by rewrite fupd_plain_mask. }
rewrite fupd_plain_forall_2. apply fupd_elim.
rewrite {1}(plain ( x, Φ x)) (fupd_mask_intro_discard E1 E2 ( _)) //.
apply fupd_elim. by rewrite fupd_plainly_elim.
Qed.
Lemma fupd_plain_forall' E {A} (Φ : A PROP) `{!∀ x, Plain (Φ x)} :
(|={E}=> x, Φ x) ⊣⊢ ( x, |={E}=> Φ x).
Proof. by apply fupd_plain_forall. Qed.
Lemma step_fupd_plain Eo Ei P `{!Plain P} : (|={Eo}[Ei]▷=> P) |={Eo}=> P.
Proof.
rewrite -(fupd_plain_mask _ Ei ( P)).
apply fupd_elim. by rewrite fupd_plain_mask -fupd_plain_later.
Qed.
Lemma step_fupdN_plain Eo Ei n P `{!Plain P} : (|={Eo}[Ei]▷=>^n P) |={Eo}=> ▷^n P.
Proof.
induction n as [|n IH].
- by rewrite -fupd_intro -except_0_intro.
- rewrite Nat.iter_succ step_fupd_fupd IH !fupd_trans step_fupd_plain.
apply fupd_mono. destruct n as [|n]; simpl.
* by rewrite except_0_idemp.
* by rewrite except_0_later.
Qed.
Lemma step_fupd_plain_forall Eo Ei {A} (Φ : A PROP) `{!∀ x, Plain (Φ x)} :
Ei Eo
(|={Eo}[Ei]▷=> x, Φ x) ⊣⊢ ( x, |={Eo}[Ei]▷=> Φ x).
Proof.
intros. apply (anti_symm _).
{ apply forall_intro=> x. by rewrite (forall_elim x). }
trans ( x, |={Eo}=> Φ x)%I.
{ apply forall_mono=> x. by rewrite step_fupd_plain. }
rewrite -fupd_plain_forall'. apply fupd_elim.
rewrite -(fupd_except_0 Ei Eo) -step_fupd_intro //.
by rewrite -later_forall -except_0_forall.
Qed.
End fupd_plainly_derived.
End fupd_derived.
(** Shared notation file for WP connectives. *)
From stdpp Require Export coPset.
From iris.bi Require Import interface derived_connectives.
From iris.prelude Require Import options.
Declare Scope expr_scope.
Delimit Scope expr_scope with E.
Declare Scope val_scope.
Delimit Scope val_scope with V.
Inductive stuckness := NotStuck | MaybeStuck.
Global Instance stuckness_le : SqSubsetEq stuckness := λ s1 s2,
match s1, s2 with
| MaybeStuck, NotStuck => False
| _, _ => True
end.
Global Instance stuckness_le_po : PreOrder (⊑@{stuckness}).
Proof. split; by repeat intros []. Qed.
(** Weakest preconditions [WP e @ s ; E {{ Φ }}] have an additional argument [s]
of arbitrary type [A], that can be chosen by the one instantiating the [Wp] type
class. This argument can be used for e.g. the stuckness bit (as in Iris) or
thread IDs (as in iGPS).
For the case of stuckness bits, there are two specific notations
[WP e @ E {{ Φ }}] and [WP e @ E ?{{ Φ }}], which forces [A] to be [stuckness],
and [s] to be [NotStuck] or [MaybeStuck]. This will fail to typecheck if [A] is
not [stuckness]. If we ever want to use the notation [WP e @ E {{ Φ }}] with a
different [A], the plan is to generalize the notation to use [Inhabited] instead
to pick a default value depending on [A]. *)
Class Wp (PROP EXPR VAL A : Type) :=
wp : A coPset EXPR (VAL PROP) PROP.
Global Arguments wp {_ _ _ _ _} _ _ _%_E _%_I.
Global Instance: Params (@wp) 8 := {}.
Class Twp (PROP EXPR VAL A : Type) :=
twp : A coPset EXPR (VAL PROP) PROP.
Global Arguments twp {_ _ _ _ _} _ _ _%_E _%_I.
Global Instance: Params (@twp) 8 := {}.
(** Notations for partial weakest preconditions *)
(** Notations without binder -- only parsing because they overlap with the
notations with binder. *)
Notation "'WP' e @ s ; E {{ Φ } }" := (wp s E e%E Φ)
(at level 20, e, Φ at level 200, only parsing) : bi_scope.
Notation "'WP' e @ E {{ Φ } }" := (wp NotStuck E e%E Φ)
(at level 20, e, Φ at level 200, only parsing) : bi_scope.
Notation "'WP' e @ E ? {{ Φ } }" := (wp MaybeStuck E e%E Φ)
(at level 20, e, Φ at level 200, only parsing) : bi_scope.
Notation "'WP' e {{ Φ } }" := (wp NotStuck e%E Φ)
(at level 20, e, Φ at level 200, only parsing) : bi_scope.
Notation "'WP' e ? {{ Φ } }" := (wp MaybeStuck e%E Φ)
(at level 20, e, Φ at level 200, only parsing) : bi_scope.
(** Notations with binder. *)
(** The general approach we use for all these complex notations: an outer '[hv'
to switch bwteen "horizontal mode" where it all fits on one line, and "vertical
mode" where each '/' becomes a line break. Then, as appropriate, nested boxes
('['...']') for things like preconditions and postconditions such that they are
maximally horizontal and suitably indented inside the parentheses that surround
them. *)
Notation "'WP' e @ s ; E {{ v , Q } }" := (wp s E e%E (λ v, Q))
(at level 20, e, Q at level 200, v at level 200 as pattern,
format "'[hv' 'WP' e '/' @ '[' s ; '/' E ']' '/' {{ '[' v , '/' Q ']' } } ']'") : bi_scope.
Notation "'WP' e @ E {{ v , Q } }" := (wp NotStuck E e%E (λ v, Q))
(at level 20, e, Q at level 200, v at level 200 as pattern,
format "'[hv' 'WP' e '/' @ E '/' {{ '[' v , '/' Q ']' } } ']'") : bi_scope.
Notation "'WP' e @ E ? {{ v , Q } }" := (wp MaybeStuck E e%E (λ v, Q))
(at level 20, e, Q at level 200, v at level 200 as pattern,
format "'[hv' 'WP' e '/' @ E '/' ? {{ '[' v , '/' Q ']' } } ']'") : bi_scope.
Notation "'WP' e {{ v , Q } }" := (wp NotStuck e%E (λ v, Q))
(at level 20, e, Q at level 200, v at level 200 as pattern,
format "'[hv' 'WP' e '/' {{ '[' v , '/' Q ']' } } ']'") : bi_scope.
Notation "'WP' e ? {{ v , Q } }" := (wp MaybeStuck e%E (λ v, Q))
(at level 20, e, Q at level 200, v at level 200 as pattern,
format "'[hv' 'WP' e '/' ? {{ '[' v , '/' Q ']' } } ']'") : bi_scope.
(* Texan triples *)
Notation "'{{{' P } } } e @ s ; E {{{ x .. y , 'RET' pat ; Q } } }" :=
( Φ,
P -∗ ( x, .. (∀ y, Q -∗ Φ pat%V) .. ) -∗ WP e @ s; E {{ Φ }})%I
(at level 20, x closed binder, y closed binder,
format "'[hv' {{{ '[' P ']' } } } '/ ' e '/' @ '[' s ; '/' E ']' '/' {{{ '[' x .. y , RET pat ; '/' Q ']' } } } ']'") : bi_scope.
Notation "'{{{' P } } } e @ E {{{ x .. y , 'RET' pat ; Q } } }" :=
( Φ,
P -∗ ( x, .. (∀ y, Q -∗ Φ pat%V) .. ) -∗ WP e @ E {{ Φ }})%I
(at level 20, x closed binder, y closed binder,
format "'[hv' {{{ '[' P ']' } } } '/ ' e '/' @ E '/' {{{ '[' x .. y , RET pat ; '/' Q ']' } } } ']'") : bi_scope.
Notation "'{{{' P } } } e @ E ? {{{ x .. y , 'RET' pat ; Q } } }" :=
( Φ,
P -∗ ( x, .. (∀ y, Q -∗ Φ pat%V) .. ) -∗ WP e @ E ?{{ Φ }})%I
(at level 20, x closed binder, y closed binder,
format "'[hv' {{{ '[' P ']' } } } '/ ' e '/' @ E '/' ? {{{ '[' x .. y , RET pat ; '/' Q ']' } } } ']'") : bi_scope.
Notation "'{{{' P } } } e {{{ x .. y , 'RET' pat ; Q } } }" :=
( Φ,
P -∗ ( x, .. (∀ y, Q -∗ Φ pat%V) .. ) -∗ WP e {{ Φ }})%I
(at level 20, x closed binder, y closed binder,
format "'[hv' {{{ '[' P ']' } } } '/ ' e '/' {{{ '[' x .. y , RET pat ; '/' Q ']' } } } ']'") : bi_scope.
Notation "'{{{' P } } } e ? {{{ x .. y , 'RET' pat ; Q } } }" :=
( Φ,
P -∗ ( x, .. (∀ y, Q -∗ Φ pat%V) .. ) -∗ WP e ?{{ Φ }})%I
(at level 20, x closed binder, y closed binder,
format "'[hv' {{{ '[' P ']' } } } '/ ' e '/' ? {{{ '[' x .. y , RET pat ; '/' Q ']' } } } ']'") : bi_scope.
Notation "'{{{' P } } } e @ s ; E {{{ 'RET' pat ; Q } } }" :=
( Φ, P -∗ (Q -∗ Φ pat%V) -∗ WP e @ s; E {{ Φ }})%I
(at level 20,
format "'[hv' {{{ '[' P ']' } } } '/ ' e '/' @ '[' s ; '/' E ']' '/' {{{ '[' RET pat ; '/' Q ']' } } } ']'") : bi_scope.
Notation "'{{{' P } } } e @ E {{{ 'RET' pat ; Q } } }" :=
( Φ, P -∗ (Q -∗ Φ pat%V) -∗ WP e @ E {{ Φ }})%I
(at level 20,
format "'[hv' {{{ '[' P ']' } } } '/ ' e '/' @ E '/' {{{ '[' RET pat ; '/' Q ']' } } } ']'") : bi_scope.
Notation "'{{{' P } } } e @ E ? {{{ 'RET' pat ; Q } } }" :=
( Φ, P -∗ (Q -∗ Φ pat%V) -∗ WP e @ E ?{{ Φ }})%I
(at level 20,
format "'[hv' {{{ '[' P ']' } } } '/ ' e '/' @ E '/' ? {{{ '[' RET pat ; '/' Q ']' } } } ']'") : bi_scope.
Notation "'{{{' P } } } e {{{ 'RET' pat ; Q } } }" :=
( Φ, P -∗ (Q -∗ Φ pat%V) -∗ WP e {{ Φ }})%I
(at level 20,
format "'[hv' {{{ '[' P ']' } } } '/ ' e '/' {{{ '[' RET pat ; '/' Q ']' } } } ']'") : bi_scope.
Notation "'{{{' P } } } e ? {{{ 'RET' pat ; Q } } }" :=
( Φ, P -∗ (Q -∗ Φ pat%V) -∗ WP e ?{{ Φ }})%I
(at level 20,
format "'[hv' {{{ '[' P ']' } } } '/ ' e '/' ? {{{ '[' RET pat ; '/' Q ']' } } } ']'") : bi_scope.
(** Aliases for stdpp scope -- they inherit the levels and format from above. *)
Notation "'{{{' P } } } e @ s ; E {{{ x .. y , 'RET' pat ; Q } } }" :=
( Φ, P -∗ ( x, .. (∀ y, Q -∗ Φ pat%V) .. ) -∗ WP e @ s; E {{ Φ }}) : stdpp_scope.
Notation "'{{{' P } } } e @ E {{{ x .. y , 'RET' pat ; Q } } }" :=
( Φ, P -∗ ( x, .. (∀ y, Q -∗ Φ pat%V) .. ) -∗ WP e @ E {{ Φ }}) : stdpp_scope.
Notation "'{{{' P } } } e @ E ? {{{ x .. y , 'RET' pat ; Q } } }" :=
( Φ, P -∗ ( x, .. (∀ y, Q -∗ Φ pat%V) .. ) -∗ WP e @ E ?{{ Φ }}) : stdpp_scope.
Notation "'{{{' P } } } e {{{ x .. y , 'RET' pat ; Q } } }" :=
( Φ, P -∗ ( x, .. (∀ y, Q -∗ Φ pat%V) .. ) -∗ WP e {{ Φ }}) : stdpp_scope.
Notation "'{{{' P } } } e ? {{{ x .. y , 'RET' pat ; Q } } }" :=
( Φ, P -∗ ( x, .. (∀ y, Q -∗ Φ pat%V) .. ) -∗ WP e ?{{ Φ }}) : stdpp_scope.
Notation "'{{{' P } } } e @ s ; E {{{ 'RET' pat ; Q } } }" :=
( Φ, P -∗ (Q -∗ Φ pat%V) -∗ WP e @ s; E {{ Φ }}) : stdpp_scope.
Notation "'{{{' P } } } e @ E {{{ 'RET' pat ; Q } } }" :=
( Φ, P -∗ (Q -∗ Φ pat%V) -∗ WP e @ E {{ Φ }}) : stdpp_scope.
Notation "'{{{' P } } } e @ E ? {{{ 'RET' pat ; Q } } }" :=
( Φ, P -∗ (Q -∗ Φ pat%V) -∗ WP e @ E ?{{ Φ }}) : stdpp_scope.
Notation "'{{{' P } } } e {{{ 'RET' pat ; Q } } }" :=
( Φ, P -∗ (Q -∗ Φ pat%V) -∗ WP e {{ Φ }}) : stdpp_scope.
Notation "'{{{' P } } } e ? {{{ 'RET' pat ; Q } } }" :=
( Φ, P -∗ (Q -∗ Φ pat%V) -∗ WP e ?{{ Φ }}) : stdpp_scope.
(** Notations for total weakest preconditions *)
(** Notations without binder -- only parsing because they overlap with the
notations with binder. *)
Notation "'WP' e @ s ; E [{ Φ } ]" := (twp s E e%E Φ)
(at level 20, e, Φ at level 200, only parsing) : bi_scope.
Notation "'WP' e @ E [{ Φ } ]" := (twp NotStuck E e%E Φ)
(at level 20, e, Φ at level 200, only parsing) : bi_scope.
Notation "'WP' e @ E ? [{ Φ } ]" := (twp MaybeStuck E e%E Φ)
(at level 20, e, Φ at level 200, only parsing) : bi_scope.
Notation "'WP' e [{ Φ } ]" := (twp NotStuck e%E Φ)
(at level 20, e, Φ at level 200, only parsing) : bi_scope.
Notation "'WP' e ? [{ Φ } ]" := (twp MaybeStuck e%E Φ)
(at level 20, e, Φ at level 200, only parsing) : bi_scope.
(** Notations with binder. *)
Notation "'WP' e @ s ; E [{ v , Q } ]" := (twp s E e%E (λ v, Q))
(at level 20, e, Q at level 200, v at level 200 as pattern,
format "'[hv' 'WP' e '/' @ '[' s ; '/' E ']' '/' [{ '[' v , '/' Q ']' } ] ']'") : bi_scope.
Notation "'WP' e @ E [{ v , Q } ]" := (twp NotStuck E e%E (λ v, Q))
(at level 20, e, Q at level 200, v at level 200 as pattern,
format "'[hv' 'WP' e '/' @ E '/' [{ '[' v , '/' Q ']' } ] ']'") : bi_scope.
Notation "'WP' e @ E ? [{ v , Q } ]" := (twp MaybeStuck E e%E (λ v, Q))
(at level 20, e, Q at level 200, v at level 200 as pattern,
format "'[hv' 'WP' e '/' @ E '/' ? [{ '[' v , '/' Q ']' } ] ']'") : bi_scope.
Notation "'WP' e [{ v , Q } ]" := (twp NotStuck e%E (λ v, Q))
(at level 20, e, Q at level 200, v at level 200 as pattern,
format "'[hv' 'WP' e '/' [{ '[' v , '/' Q ']' } ] ']'") : bi_scope.
Notation "'WP' e ? [{ v , Q } ]" := (twp MaybeStuck e%E (λ v, Q))
(at level 20, e, Q at level 200, v at level 200 as pattern,
format "'[hv' 'WP' e '/' ? [{ '[' v , '/' Q ']' } ] ']'") : bi_scope.
(* Texan triples *)
Notation "'[[{' P } ] ] e @ s ; E [[{ x .. y , 'RET' pat ; Q } ] ]" :=
( Φ,
P -∗ ( x, .. (∀ y, Q -∗ Φ pat%V) .. ) -∗ WP e @ s; E [{ Φ }])%I
(at level 20, x closed binder, y closed binder,
format "'[hv' [[{ '[' P ']' } ] ] '/ ' e '/' @ '[' s ; '/' E ']' '/' [[{ '[hv' x .. y , RET pat ; '/' Q ']' } ] ] ']'") : bi_scope.
Notation "'[[{' P } ] ] e @ E [[{ x .. y , 'RET' pat ; Q } ] ]" :=
( Φ,
P -∗ ( x, .. (∀ y, Q -∗ Φ pat%V) .. ) -∗ WP e @ E [{ Φ }])%I
(at level 20, x closed binder, y closed binder,
format "'[hv' [[{ '[' P ']' } ] ] '/ ' e '/' @ E '/' [[{ '[hv' x .. y , RET pat ; '/' Q ']' } ] ] ']'") : bi_scope.
Notation "'[[{' P } ] ] e @ E ? [[{ x .. y , 'RET' pat ; Q } ] ]" :=
( Φ,
P -∗ ( x, .. (∀ y, Q -∗ Φ pat%V) .. ) -∗ WP e @ E ?[{ Φ }])%I
(at level 20, x closed binder, y closed binder,
format "'[hv' [[{ '[' P ']' } ] ] '/ ' e '/' @ E '/' ? [[{ '[hv' x .. y , RET pat ; '/' Q ']' } ] ] ']'") : bi_scope.
Notation "'[[{' P } ] ] e [[{ x .. y , 'RET' pat ; Q } ] ]" :=
( Φ,
P -∗ ( x, .. (∀ y, Q -∗ Φ pat%V) .. ) -∗ WP e [{ Φ }])%I
(at level 20, x closed binder, y closed binder,
format "'[hv' [[{ '[' P ']' } ] ] '/ ' e '/' [[{ '[hv' x .. y , RET pat ; '/' Q ']' } ] ] ']'") : bi_scope.
Notation "'[[{' P } ] ] e ? [[{ x .. y , 'RET' pat ; Q } ] ]" :=
( Φ,
P -∗ ( x, .. (∀ y, Q -∗ Φ pat%V) .. ) -∗ WP e ?[{ Φ }])%I
(at level 20, x closed binder, y closed binder,
format "'[hv' [[{ '[' P ']' } ] ] '/ ' e '/' ? [[{ '[hv' x .. y , RET pat ; '/' Q ']' } ] ] ']'") : bi_scope.
Notation "'[[{' P } ] ] e @ s ; E [[{ 'RET' pat ; Q } ] ]" :=
( Φ, P -∗ (Q -∗ Φ pat%V) -∗ WP e @ s; E [{ Φ }])%I
(at level 20,
format "'[hv' [[{ '[' P ']' } ] ] '/ ' e '/' @ '[' s ; '/' E ']' '/' [[{ '[hv' RET pat ; '/' Q ']' } ] ] ']'") : bi_scope.
Notation "'[[{' P } ] ] e @ E [[{ 'RET' pat ; Q } ] ]" :=
( Φ, P -∗ (Q -∗ Φ pat%V) -∗ WP e @ E [{ Φ }])%I
(at level 20,
format "'[hv' [[{ '[' P ']' } ] ] '/ ' e '/' @ E '/' [[{ '[hv' RET pat ; '/' Q ']' } ] ] ']'") : bi_scope.
Notation "'[[{' P } ] ] e @ E ? [[{ 'RET' pat ; Q } ] ]" :=
( Φ, P -∗ (Q -∗ Φ pat%V) -∗ WP e @ E ?[{ Φ }])%I
(at level 20,
format "'[hv' [[{ '[' P ']' } ] ] '/ ' e '/' @ E '/' ? [[{ '[hv' RET pat ; '/' Q ']' } ] ] ']'") : bi_scope.
Notation "'[[{' P } ] ] e [[{ 'RET' pat ; Q } ] ]" :=
( Φ, P -∗ (Q -∗ Φ pat%V) -∗ WP e [{ Φ }])%I
(at level 20,
format "'[hv' [[{ '[' P ']' } ] ] '/ ' e '/' [[{ '[hv' RET pat ; '/' Q ']' } ] ] ']'") : bi_scope.
Notation "'[[{' P } ] ] e ? [[{ 'RET' pat ; Q } ] ]" :=
( Φ, P -∗ (Q -∗ Φ pat%V) -∗ WP e ?[{ Φ }])%I
(at level 20,
format "'[hv' [[{ '[' P ']' } ] ] '/ ' e '/' ? [[{ '[hv' RET pat ; '/' Q ']' } ] ] ']'") : bi_scope.
(** Aliases for stdpp scope -- they inherit the levels and format from above. *)
Notation "'[[{' P } ] ] e @ s ; E [[{ x .. y , 'RET' pat ; Q } ] ]" :=
( Φ, P -∗ ( x, .. (∀ y, Q -∗ Φ pat%V) .. ) -∗ WP e @ s; E [{ Φ }]) : stdpp_scope.
Notation "'[[{' P } ] ] e @ E [[{ x .. y , 'RET' pat ; Q } ] ]" :=
( Φ, P -∗ ( x, .. (∀ y, Q -∗ Φ pat%V) .. ) -∗ WP e @ E [{ Φ }]) : stdpp_scope.
Notation "'[[{' P } ] ] e @ E ? [[{ x .. y , 'RET' pat ; Q } ] ]" :=
( Φ, P -∗ ( x, .. (∀ y, Q -∗ Φ pat%V) .. ) -∗ WP e @ E ?[{ Φ }]) : stdpp_scope.
Notation "'[[{' P } ] ] e [[{ x .. y , 'RET' pat ; Q } ] ]" :=
( Φ, P -∗ ( x, .. (∀ y, Q -∗ Φ pat%V) .. ) -∗ WP e [{ Φ }]) : stdpp_scope.
Notation "'[[{' P } ] ] e ? [[{ x .. y , 'RET' pat ; Q } ] ]" :=
( Φ, P -∗ ( x, .. (∀ y, Q -∗ Φ pat%V) .. ) -∗ WP e ?[{ Φ }]) : stdpp_scope.
Notation "'[[{' P } ] ] e @ s ; E [[{ 'RET' pat ; Q } ] ]" :=
( Φ, P -∗ (Q -∗ Φ pat%V) -∗ WP e @ s; E [{ Φ }]) : stdpp_scope.
Notation "'[[{' P } ] ] e @ E [[{ 'RET' pat ; Q } ] ]" :=
( Φ, P -∗ (Q -∗ Φ pat%V) -∗ WP e @ E [{ Φ }]) : stdpp_scope.
Notation "'[[{' P } ] ] e @ E ? [[{ 'RET' pat ; Q } ] ]" :=
( Φ, P -∗ (Q -∗ Φ pat%V) -∗ WP e @ E ?[{ Φ }]) : stdpp_scope.
Notation "'[[{' P } ] ] e [[{ 'RET' pat ; Q } ] ]" :=
( Φ, P -∗ (Q -∗ Φ pat%V) -∗ WP e [{ Φ }]) : stdpp_scope.
Notation "'[[{' P } ] ] e ? [[{ 'RET' pat ; Q } ] ]" :=
( Φ, P -∗ (Q -∗ Φ pat%V) -∗ WP e ?[{ Φ }]) : stdpp_scope.
(include_subdirs qualified)
(coq.theory
(name iris)
(package coq-iris)
(theories Ltac2 stdpp))
(** Coq configuration for Iris (not meant to leak to clients).
If you are a user of Iris, note that importing this file means
you are implicitly opting-in to every new option we will add here
in the future. We are *not* guaranteeing any kind of stability here.
Instead our advice is for you to have your own options file; then
you can re-export the Iris file there but if we ever add an option
you disagree with you can easily overwrite it in one central location. *)
(* Everything here should be [Export Set], which means when this
file is *imported*, the option will only apply on the import site
but not transitively. *)
From stdpp Require Export options.
#[export] Set Suggest Proof Using. (* also warns about forgotten [Proof.] *)
(* "Fake" import to whitelist this file for the check that ensures we import
this file everywhere.
From iris.prelude Require Import options.
*)
From stdpp Require Export ssreflect.
From iris.prelude Require Import options.
(** Iris itself and many dependencies still rely on this coercion. *)
Coercion Z.of_nat : nat >-> Z.
From iris.algebra Require Import gmap auth agree gset coPset.
From iris.proofmode Require Import proofmode.
From iris.base_logic.lib Require Import wsat.
From iris.program_logic Require Export weakestpre.
From iris.prelude Require Import options.
Import uPred.
(** This file contains the adequacy statements of the Iris program logic. First
we prove a number of auxiliary results. *)
Section adequacy.
Context `{!irisGS_gen hlc Λ Σ}.
Implicit Types e : expr Λ.
Implicit Types P Q : iProp Σ.
Implicit Types Φ : val Λ iProp Σ.
Implicit Types Φs : list (val Λ iProp Σ).
Notation wptp s t Φs := ([ list] e;Φ t;Φs, WP e @ s; {{ Φ }})%I.
Local Lemma wp_step s e1 σ1 ns κ κs e2 σ2 efs nt Φ :
prim_step e1 σ1 κ e2 σ2 efs
state_interp σ1 ns (κ ++ κs) nt -∗
£ (S (num_laters_per_step ns)) -∗
WP e1 @ s; {{ Φ }}
={,}=∗ |={}▷=>^(S $ num_laters_per_step ns) |={,}=>
state_interp σ2 (S ns) κs (nt + length efs) WP e2 @ s; {{ Φ }}
wptp s efs (replicate (length efs) fork_post).
Proof.
rewrite {1}wp_unfold /wp_pre. iIntros (?) "Hσ Hcred H".
rewrite (val_stuck e1 σ1 κ e2 σ2 efs) //.
iMod ("H" $! σ1 ns with "Hσ") as "(_ & H)". iModIntro.
iApply (step_fupdN_wand with "(H [//] Hcred)"). iIntros ">H".
by rewrite Nat.add_comm big_sepL2_replicate_r.
Qed.
Local Lemma wptp_step s es1 es2 κ κs σ1 ns σ2 Φs nt :
step (es1,σ1) κ (es2, σ2)
state_interp σ1 ns (κ ++ κs) nt -∗
£ (S (num_laters_per_step ns)) -∗
wptp s es1 Φs -∗
nt', |={,}=> |={}▷=>^(S $ num_laters_per_step$ ns) |={,}=>
state_interp σ2 (S ns) κs (nt + nt')
wptp s es2 (Φs ++ replicate nt' fork_post).
Proof.
iIntros (Hstep) "Hσ Hcred Ht".
destruct Hstep as [e1' σ1' e2' σ2' efs t2' t3 Hstep]; simplify_eq/=.
iDestruct (big_sepL2_app_inv_l with "Ht") as (Φs1 Φs2 ->) "[? Ht]".
iDestruct (big_sepL2_cons_inv_l with "Ht") as (Φ Φs3 ->) "[Ht ?]".
iExists _. iMod (wp_step with "Hσ Hcred Ht") as "H"; first done. iModIntro.
iApply (step_fupdN_wand with "H"). iIntros ">($ & He2 & Hefs) !>".
rewrite -(assoc_L app) -app_comm_cons. iFrame.
Qed.
(* The total number of laters used between the physical steps number
[start] (included) to [start+ns] (excluded). *)
Local Fixpoint steps_sum (num_laters_per_step : nat nat) (start ns : nat) : nat :=
match ns with
| O => 0
| S ns =>
S $ num_laters_per_step start + steps_sum num_laters_per_step (S start) ns
end.
Local Lemma wptp_preservation s n es1 es2 κs κs' σ1 ns σ2 Φs nt :
nsteps n (es1, σ1) κs (es2, σ2)
state_interp σ1 ns (κs ++ κs') nt -∗
£ (steps_sum num_laters_per_step ns n) -∗
wptp s es1 Φs
={,}=∗ |={}▷=>^(steps_sum num_laters_per_step ns n) |={,}=> nt',
state_interp σ2 (n + ns) κs' (nt + nt')
wptp s es2 (Φs ++ replicate nt' fork_post).
Proof.
revert nt es1 es2 κs κs' σ1 ns σ2 Φs.
induction n as [|n IH]=> nt es1 es2 κs κs' σ1 ns σ2 Φs /=.
{ inversion_clear 1; iIntros "? ? ?"; iExists 0=> /=.
rewrite Nat.add_0_r right_id_L. iFrame. by iApply fupd_mask_subseteq. }
iIntros (Hsteps) "Hσ Hcred He". inversion_clear Hsteps as [|?? [t1' σ1']].
rewrite -(assoc_L (++)) Nat.iter_add -{1}plus_Sn_m plus_n_Sm.
rewrite lc_split. iDestruct "Hcred" as "[Hc1 Hc2]".
iDestruct (wptp_step with "Hσ Hc1 He") as (nt') ">H"; first eauto; simplify_eq.
iModIntro. iApply step_fupdN_S_fupd. iApply (step_fupdN_wand with "H").
iIntros ">(Hσ & He)". iMod (IH with "Hσ Hc2 He") as "IH"; first done. iModIntro.
iApply (step_fupdN_wand with "IH"). iIntros ">IH".
iDestruct "IH" as (nt'') "[??]".
rewrite -Nat.add_assoc -(assoc_L app) -replicate_add. by eauto with iFrame.
Qed.
Local Lemma wp_not_stuck κs nt e σ ns Φ :
state_interp σ ns κs nt -∗ WP e {{ Φ }} ={, }=∗ not_stuck e σ⌝.
Proof.
rewrite wp_unfold /wp_pre /not_stuck. iIntros "Hσ H".
destruct (to_val e) as [v|] eqn:?.
{ iMod (fupd_mask_subseteq ); first set_solver. iModIntro. eauto. }
iSpecialize ("H" $! σ ns [] κs with "Hσ"). rewrite sep_elim_l.
iMod "H" as "%". iModIntro. eauto.
Qed.
(** The adequacy statement of Iris consists of two parts:
(1) the postcondition for all threads that have terminated in values
and (2) progress (i.e., after n steps the program is not stuck).
For an n-step execution of a thread pool, the two parts are given by
[wptp_strong_adequacy] and [wptp_progress] below.
For the final adequacy theorem of Iris, [wp_strong_adequacy_gen], we would
like to instantiate the Iris proof (i.e., instantiate the
[∀ {Hinv : !invGS_gen hlc Σ} κs, ...]) and then use both lemmas to get
progress and the postconditions. Unfortunately, since the addition of later
credits, this is no longer possible, because the original proof relied on an
interaction of the update modality and plain propositions. So instead, we
employ a trick: we duplicate the instantiation of the Iris proof, such
that we can "run the WP proof twice". That is, we instantiate the
[∀ {Hinv : !invGS_gen hlc Σ} κs, ...] both in [wp_progress_gen] and
[wp_strong_adequacy_gen]. In doing so, we can avoid the interactions with
the plain modality. In [wp_strong_adequacy_gen], we can then make use of
[wp_progress_gen] to prove the progress component of the main adequacy theorem.
*)
Local Lemma wptp_postconditions Φs κs' s n es1 es2 κs σ1 ns σ2 nt:
nsteps n (es1, σ1) κs (es2, σ2)
state_interp σ1 ns (κs ++ κs') nt -∗
£ (steps_sum num_laters_per_step ns n) -∗
wptp s es1 Φs
={,}=∗ |={}▷=>^(steps_sum num_laters_per_step ns n) |={,}=> nt',
state_interp σ2 (n + ns) κs' (nt + nt')
[ list] e;Φ es2;Φs ++ replicate nt' fork_post, from_option Φ True (to_val e).
Proof.
iIntros (Hstep) "Hσ Hcred He". iMod (wptp_preservation with "Hσ Hcred He") as "Hwp"; first done.
iModIntro. iApply (step_fupdN_wand with "Hwp").
iMod 1 as (nt') "(Hσ & Ht)"; simplify_eq/=.
iExists _. iFrame "Hσ".
iApply big_sepL2_fupd.
iApply (big_sepL2_impl with "Ht").
iIntros "!#" (? e Φ ??) "Hwp".
destruct (to_val e) as [v2|] eqn:He2'; last done.
apply of_to_val in He2' as <-. simpl. iApply wp_value_fupd'. done.
Qed.
Local Lemma wptp_progress Φs κs' n es1 es2 κs σ1 ns σ2 nt e2 :
nsteps n (es1, σ1) κs (es2, σ2)
e2 es2
state_interp σ1 ns (κs ++ κs') nt -∗
£ (steps_sum num_laters_per_step ns n) -∗
wptp NotStuck es1 Φs
={,}=∗ |={}▷=>^(steps_sum num_laters_per_step ns n) |={}=> not_stuck e2 σ2⌝.
Proof.
iIntros (Hstep Hel) "Hσ Hcred He". iMod (wptp_preservation with "Hσ Hcred He") as "Hwp"; first done.
iModIntro. iApply (step_fupdN_wand with "Hwp").
iMod 1 as (nt') "(Hσ & Ht)"; simplify_eq/=.
eapply elem_of_list_lookup in Hel as [i Hlook].
destruct ((Φs ++ replicate nt' fork_post) !! i) as [Φ|] eqn: Hlook2; last first.
{ rewrite big_sepL2_alt. iDestruct "Ht" as "[%Hlen _]". exfalso.
eapply lookup_lt_Some in Hlook. rewrite Hlen in Hlook.
eapply lookup_lt_is_Some_2 in Hlook. rewrite Hlook2 in Hlook.
destruct Hlook as [? ?]. naive_solver. }
iDestruct (big_sepL2_lookup with "Ht") as "Ht"; [done..|].
by iApply (wp_not_stuck with "Hσ").
Qed.
End adequacy.
Local Lemma wp_progress_gen (hlc : has_lc) Σ Λ `{!invGpreS Σ} es σ1 n κs t2 σ2 e2
(num_laters_per_step : nat nat) :
( `{Hinv : !invGS_gen hlc Σ},
|={}=>
(stateI : state Λ nat list (observation Λ) nat iProp Σ)
(Φs : list (val Λ iProp Σ))
(fork_post : val Λ iProp Σ)
state_interp_mono,
let _ : irisGS_gen hlc Λ Σ := IrisG Hinv stateI fork_post num_laters_per_step
state_interp_mono
in
stateI σ1 0 κs 0
([ list] e;Φ es;Φs, WP e @ {{ Φ }}))
nsteps n (es, σ1) κs (t2, σ2)
e2 t2
not_stuck e2 σ2.
Proof.
intros Hwp ??.
eapply pure_soundness.
eapply (step_fupdN_soundness_gen _ hlc (steps_sum num_laters_per_step 0 n)
(steps_sum num_laters_per_step 0 n)).
iIntros (Hinv) "Hcred".
iMod Hwp as (stateI Φ fork_post state_interp_mono) "(Hσ & Hwp)".
iDestruct (big_sepL2_length with "Hwp") as %Hlen1.
iMod (@wptp_progress _ _ _
(IrisG Hinv stateI fork_post num_laters_per_step state_interp_mono) _ []
with "[Hσ] Hcred Hwp") as "H"; [done| done |by rewrite right_id_L|].
iAssert (|={}▷=>^(steps_sum num_laters_per_step 0 n) |={}=> not_stuck e2 σ2)%I
with "[-]" as "H"; last first.
{ destruct steps_sum; [done|]. by iApply step_fupdN_S_fupd. }
iApply (step_fupdN_wand with "H"). iIntros "$".
Qed.
(** Iris's generic adequacy result *)
(** The lemma is parameterized by [use_credits] over whether to make later credits available or not.
Below, a concrete instances is provided with later credits (see [wp_strong_adequacy]). *)
Lemma wp_strong_adequacy_gen (hlc : has_lc) Σ Λ `{!invGpreS Σ} s es σ1 n κs t2 σ2 φ
(num_laters_per_step : nat nat) :
(* WP *)
( `{Hinv : !invGS_gen hlc Σ},
|={}=>
(stateI : state Λ nat list (observation Λ) nat iProp Σ)
(Φs : list (val Λ iProp Σ))
(fork_post : val Λ iProp Σ)
(* Note: existentially quantifying over Iris goal! [iExists _] should
usually work. *)
state_interp_mono,
let _ : irisGS_gen hlc Λ Σ := IrisG Hinv stateI fork_post num_laters_per_step
state_interp_mono
in
stateI σ1 0 κs 0
([ list] e;Φ es;Φs, WP e @ s; {{ Φ }})
( es' t2',
(* es' is the final state of the initial threads, t2' the rest *)
t2 = es' ++ t2' -∗
(* es' corresponds to the initial threads *)
length es' = length es -∗
(* If this is a stuck-free triple (i.e. [s = NotStuck]), then all
threads in [t2] are not stuck *)
e2, s = NotStuck e2 t2 not_stuck e2 σ2 -∗
(* The state interpretation holds for [σ2] *)
stateI σ2 n [] (length t2') -∗
(* If the initial threads are done, their post-condition [Φ] holds *)
([ list] e;Φ es';Φs, from_option Φ True (to_val e)) -∗
(* For all forked-off threads that are done, their postcondition
[fork_post] holds. *)
([ list] v omap to_val t2', fork_post v) -∗
(* Under all these assumptions, and while opening all invariants, we
can conclude [φ] in the logic. After opening all required invariants,
one can use [fupd_mask_subseteq] to introduce the fancy update. *)
|={,}=> φ ))
nsteps n (es, σ1) κs (t2, σ2)
(* Then we can conclude [φ] at the meta-level. *)
φ.
Proof.
intros Hwp ?.
eapply pure_soundness.
eapply (step_fupdN_soundness_gen _ hlc (steps_sum num_laters_per_step 0 n)
(steps_sum num_laters_per_step 0 n)).
iIntros (Hinv) "Hcred".
iMod Hwp as (stateI Φ fork_post state_interp_mono) "(Hσ & Hwp & Hφ)".
iDestruct (big_sepL2_length with "Hwp") as %Hlen1.
iMod (@wptp_postconditions _ _ _
(IrisG Hinv stateI fork_post num_laters_per_step state_interp_mono) _ []
with "[Hσ] Hcred Hwp") as "H"; [done|by rewrite right_id_L|].
iAssert (|={}▷=>^(steps_sum num_laters_per_step 0 n) |={}=> φ)%I
with "[-]" as "H"; last first.
{ destruct steps_sum; [done|]. by iApply step_fupdN_S_fupd. }
iApply (step_fupdN_wand with "H").
iMod 1 as (nt') "(Hσ & Hval) /=".
iDestruct (big_sepL2_app_inv_r with "Hval") as (es' t2' ->) "[Hes' Ht2']".
iDestruct (big_sepL2_length with "Ht2'") as %Hlen2.
rewrite length_replicate in Hlen2; subst.
iDestruct (big_sepL2_length with "Hes'") as %Hlen3.
rewrite -plus_n_O.
iApply ("Hφ" with "[//] [%] [ ] Hσ Hes'");
(* FIXME: Different implicit types for [length] are inferred, so [lia] and
[congruence] do not work due to https://github.com/coq/coq/issues/16634 *)
[by rewrite Hlen1 Hlen3| |]; last first.
{ by rewrite big_sepL2_replicate_r // big_sepL_omap. }
(* At this point in the adequacy proof, we use a trick: we effectively run the
user-provided WP proof again (i.e., instantiate the `invGS_gen` and execute the
program) by using the lemma [wp_progress_gen]. In doing so, we can obtain
the progress part of the adequacy theorem.
*)
iPureIntro. intros e2 -> Hel.
eapply (wp_progress_gen hlc);
[ done | clear stateI Φ fork_post state_interp_mono Hlen1 Hlen3 | done|done].
iIntros (?).
iMod Hwp as (stateI Φ fork_post state_interp_mono) "(Hσ & Hwp & Hφ)".
iModIntro. iExists _, _, _, _. iFrame.
Qed.
(** Adequacy when using later credits (the default) *)
Definition wp_strong_adequacy := wp_strong_adequacy_gen HasLc.
Global Arguments wp_strong_adequacy _ _ {_}.
(** Since the full adequacy statement is quite a mouthful, we prove some more
intuitive and simpler corollaries. These lemmas are morover stated in terms of
[rtc erased_step] so one does not have to provide the trace. *)
Record adequate {Λ} (s : stuckness) (e1 : expr Λ) (σ1 : state Λ)
(φ : val Λ state Λ Prop) := {
adequate_result t2 σ2 v2 :
rtc erased_step ([e1], σ1) (of_val v2 :: t2, σ2) φ v2 σ2;
adequate_not_stuck t2 σ2 e2 :
s = NotStuck
rtc erased_step ([e1], σ1) (t2, σ2)
e2 t2 not_stuck e2 σ2
}.
Lemma adequate_alt {Λ} s e1 σ1 (φ : val Λ state Λ Prop) :
adequate s e1 σ1 φ t2 σ2,
rtc erased_step ([e1], σ1) (t2, σ2)
( v2 t2', t2 = of_val v2 :: t2' φ v2 σ2)
( e2, s = NotStuck e2 t2 not_stuck e2 σ2).
Proof.
split.
- intros []; naive_solver.
- constructor; naive_solver.
Qed.
Theorem adequate_tp_safe {Λ} (e1 : expr Λ) t2 σ1 σ2 φ :
adequate NotStuck e1 σ1 φ
rtc erased_step ([e1], σ1) (t2, σ2)
Forall (λ e, is_Some (to_val e)) t2 t3 σ3, erased_step (t2, σ2) (t3, σ3).
Proof.
intros Had ?.
destruct (decide (Forall (λ e, is_Some (to_val e)) t2)) as [|Ht2]; [by left|].
apply (not_Forall_Exists _), Exists_exists in Ht2; destruct Ht2 as (e2&?&He2).
destruct (adequate_not_stuck NotStuck e1 σ1 φ Had t2 σ2 e2) as [?|(κ&e3&σ3&efs&?)];
rewrite ?eq_None_not_Some; auto.
{ exfalso. eauto. }
destruct (elem_of_list_split t2 e2) as (t2'&t2''&->); auto.
right; exists (t2' ++ e3 :: t2'' ++ efs), σ3, κ; econstructor; eauto.
Qed.
(** This simpler form of adequacy requires the [irisGS] instance that you use
everywhere to syntactically be of the form
{|
iris_invGS := ...;
state_interp σ _ κs _ := ...;
fork_post v := ...;
num_laters_per_step _ := 0;
state_interp_mono _ _ _ _ := fupd_intro _ _;
|}
In other words, the state interpretation must ignore [ns] and [nt], the number
of laters per step must be 0, and the proof of [state_interp_mono] must have
this specific proof term.
*)
(** Again, we first prove a lemma generic over the usage of credits. *)
Lemma wp_adequacy_gen (hlc : has_lc) Σ Λ `{!invGpreS Σ} s e σ φ :
( `{Hinv : !invGS_gen hlc Σ} κs,
|={}=>
(stateI : state Λ list (observation Λ) iProp Σ)
(fork_post : val Λ iProp Σ),
let _ : irisGS_gen hlc Λ Σ :=
IrisG Hinv (λ σ _ κs _, stateI σ κs) fork_post (λ _, 0)
(λ _ _ _ _, fupd_intro _ _)
in
stateI σ κs WP e @ s; {{ v, φ v }})
adequate s e σ (λ v _, φ v).
Proof.
intros Hwp. apply adequate_alt; intros t2 σ2 [n [κs ?]]%erased_steps_nsteps.
eapply (wp_strong_adequacy_gen hlc Σ _); [ | done]=> ?.
iMod Hwp as (stateI fork_post) "[Hσ Hwp]".
iExists (λ σ _ κs _, stateI σ κs), [(λ v, φ v⌝%I)], fork_post, _ => /=.
iIntros "{$Hσ $Hwp} !>" (e2 t2' -> ? ?) "_ H _".
iApply fupd_mask_intro_discard; [done|]. iSplit; [|done].
iDestruct (big_sepL2_cons_inv_r with "H") as (e' ? ->) "[Hwp H]".
iDestruct (big_sepL2_nil_inv_r with "H") as %->.
iIntros (v2 t2'' [= -> <-]). by rewrite to_of_val.
Qed.
(** Instance for using credits *)
Definition wp_adequacy := wp_adequacy_gen HasLc.
Global Arguments wp_adequacy _ _ {_}.
Lemma wp_invariance_gen (hlc : has_lc) Σ Λ `{!invGpreS Σ} s e1 σ1 t2 σ2 φ :
( `{Hinv : !invGS_gen hlc Σ} κs,
|={}=>
(stateI : state Λ list (observation Λ) nat iProp Σ)
(fork_post : val Λ iProp Σ),
let _ : irisGS_gen hlc Λ Σ := IrisG Hinv (λ σ _, stateI σ) fork_post
(λ _, 0) (λ _ _ _ _, fupd_intro _ _) in
stateI σ1 κs 0 WP e1 @ s; {{ _, True }}
(stateI σ2 [] (pred (length t2)) -∗ E, |={,E}=> φ))
rtc erased_step ([e1], σ1) (t2, σ2)
φ.
Proof.
intros Hwp [n [κs ?]]%erased_steps_nsteps.
eapply (wp_strong_adequacy_gen hlc Σ); [done| |done]=> ?.
iMod (Hwp _ κs) as (stateI fork_post) "(Hσ & Hwp & Hφ)".
iExists (λ σ _, stateI σ), [(λ _, True)%I], fork_post, _ => /=.
iIntros "{$Hσ $Hwp} !>" (e2 t2' -> _ _) "Hσ H _ /=".
iDestruct (big_sepL2_cons_inv_r with "H") as (? ? ->) "[_ H]".
iDestruct (big_sepL2_nil_inv_r with "H") as %->.
iDestruct ("Hφ" with "Hσ") as (E) ">Hφ".
by iApply fupd_mask_intro_discard; first set_solver.
Qed.
Definition wp_invariance := wp_invariance_gen HasLc.
Global Arguments wp_invariance _ _ {_}.
(** This file declares notation for logically atomic Hoare triples, and some
generic lemmas about them. To enable the definition of a shared theory applying
to triples with any number of binders, the triples themselves are defined via telescopes, but as a user
you need not be concerned with that. You can just use the following notation:
<<{ ∀∀ x, atomic_precondition }>>
code @ E
<<{ ∃∃ y, atomic_postcondition | z, RET return_value; private_postcondition }>>
Here, [x] (which can be any number of binders, including 0) is bound in all of
the atomic pre- and postcondition and the private (non-atomic) postcondition and
the return value, [y] (which can be any number of binders, including 0) is bound
in both postconditions and the return value, and [z] (which can be any number of
binders, including 0) is bound in the return value and the private
postcondition.
Note that atomic triples are *not* implicitly persistent, unlike Texan triples.
If you need a private (non-atomic) precondition, you can use a magic wand:
private_precondition -∗
<<{ ∀∀ x, atomic_precondition }>>
code @ E
<<{ ∃∃ y, atomic_postcondition
| z, RET return_value; private_postcondition }>>
If you don't need a private postcondition, you can leave it away, e.g.:
<<{ ∀∀ x, atomic_precondition }>>
code @ E
<<{ ∃∃ y, atomic_postcondition | RET return_value }>>
Note that due to combinatorial explosion and because Coq does not have a
facility to declare such notation in a compositional way, not *all* variants of
this notation exist: if you have binders before the [RET] (which is very
uncommon), you must have a private postcondition (it can be [True]), and you
must have [∀∀] and [∃∃] binders (they can be [_: ()]).
For an example for how to prove and use logically atomic specifications, see
[iris_heap_lang/lib/increment.v].
*)
From stdpp Require Import namespaces.
From iris.bi Require Import telescopes.
From iris.bi.lib Require Export atomic.
From iris.proofmode Require Import proofmode classes.
From iris.program_logic Require Export weakestpre.
From iris.base_logic Require Import invariants.
From iris.prelude Require Import options.
(* This hard-codes the inner mask to be empty, because we have yet to find an
example where we want it to be anything else.
For the non-atomic post-condition, we use an [option PROP], combined with a
[-∗?]. This is to avoid introducing spurious [True -∗] into proofs that do not
need a non-atomic post-condition (which is most of them). *)
Definition atomic_wp `{!irisGS_gen hlc Λ Σ} {TA TB TP : tele}
(e: expr Λ) (* expression *)
(E : coPset) (* *implementation* mask *)
(α: TA iProp Σ) (* atomic pre-condition *)
(β: TA TB iProp Σ) (* atomic post-condition *)
(POST: TA TB TP option (iProp Σ)) (* non-atomic post-condition *)
(f: TA TB TP val Λ) (* Turn the return data into the return value *)
: iProp Σ :=
(Φ : val Λ iProp Σ),
(* The (outer) user mask is what is left after the implementation
opened its things. *)
atomic_update (⊤∖E) α β (λ.. x y, .. z, POST x y z -∗? Φ (f x y z)) -∗
WP e {{ Φ }}.
(** We avoid '<<{'/'}>>' since those can also reasonably be infix operators
(and in fact Autosubst uses the latter). *)
Notation "'<<{' ∀∀ x1 .. xn , α '}>>' e @ E '<<{' ∃∃ y1 .. yn , β '|' z1 .. zn , 'RET' v ; POST '}>>'" :=
(* The way to read the [tele_app foo] here is that they convert the n-ary
function [foo] into a unary function taking a telescope as the argument. *)
(atomic_wp (TA:=TeleS (λ x1, .. (TeleS (λ xn, TeleO)) .. ))
(TB:=TeleS (λ y1, .. (TeleS (λ yn, TeleO)) .. ))
(TP:=TeleS (λ z1, .. (TeleS (λ zn, TeleO)) .. ))
e%E
E
(tele_app $ λ x1, .. (λ xn, α%I) ..)
(tele_app $ λ x1, .. (λ xn,
tele_app $ λ y1, .. (λ yn, β%I) ..
) .. )
(tele_app $ λ x1, .. (λ xn,
tele_app $ λ y1, .. (λ yn,
tele_app $ λ z1, .. (λ zn, Some POST%I) ..
) ..
) .. )
(tele_app $ λ x1, .. (λ xn,
tele_app $ λ y1, .. (λ yn,
tele_app $ λ z1, .. (λ zn, v%V) ..
) ..
) .. )
)
(at level 20, E, β, α, v, POST at level 200, x1 binder, xn binder, y1 binder, yn binder, z1 binder, zn binder,
format "'[hv' '<<{' '[' ∀∀ x1 .. xn , '/' α ']' '}>>' '/ ' e @ E '/' '<<{' '[' ∃∃ y1 .. yn , '/' β '|' '/' z1 .. zn , RET v ; '/' POST ']' '}>>' ']'")
: bi_scope.
(* There are overall 16 of possible variants of this notation:
- with and without ∀∀ binders
- with and without ∃∃ binders
- with and without RET binders
- with and without POST
However we don't support the case where RET binders are present but anything
else is missing. Below we declare the 8 notations that involve no RET binders.
*)
(* No RET binders *)
Notation "'<<{' ∀∀ x1 .. xn , α '}>>' e @ E '<<{' ∃∃ y1 .. yn , β '|' 'RET' v ; POST '}>>'" :=
(atomic_wp (TA:=TeleS (λ x1, .. (TeleS (λ xn, TeleO)) .. ))
(TB:=TeleS (λ y1, .. (TeleS (λ yn, TeleO)) .. ))
(TP:=TeleO)
e%E
E
(tele_app $ λ x1, .. (λ xn, α%I) ..)
(tele_app $ λ x1, .. (λ xn,
tele_app $ λ y1, .. (λ yn, β%I) ..
) .. )
(tele_app $ λ x1, .. (λ xn,
tele_app $ λ y1, .. (λ yn, tele_app $ Some POST%I) ..
) .. )
(tele_app $ λ x1, .. (λ xn,
tele_app $ λ y1, .. (λ yn, tele_app $ v%V) ..
) .. )
)
(at level 20, E, β, α, v, POST at level 200, x1 binder, xn binder, y1 binder, yn binder,
format "'[hv' '<<{' '[' ∀∀ x1 .. xn , '/' α ']' '}>>' '/ ' e @ E '/' '<<{' '[' ∃∃ y1 .. yn , '/' β '|' '/' RET v ; '/' POST ']' '}>>' ']'")
: bi_scope.
(* No ∃∃ binders, no RET binders *)
Notation "'<<{' ∀∀ x1 .. xn , α '}>>' e @ E '<<{' β '|' 'RET' v ; POST '}>>'" :=
(atomic_wp (TA:=TeleS (λ x1, .. (TeleS (λ xn, TeleO)) .. ))
(TB:=TeleO)
(TP:=TeleO)
e%E
E
(tele_app $ λ x1, .. (λ xn, α%I) ..)
(tele_app $ λ x1, .. (λ xn, tele_app β%I) .. )
(tele_app $ λ x1, .. (λ xn,
tele_app $ tele_app $ Some POST%I
) .. )
(tele_app $ λ x1, .. (λ xn,
tele_app $ tele_app $ v%V
) .. )
)
(at level 20, E, β, α, v, POST at level 200, x1 binder, xn binder,
format "'[hv' '<<{' '[' ∀∀ x1 .. xn , '/' α ']' '}>>' '/ ' e @ E '/' '<<{' '[' β '|' '/' RET v ; '/' POST ']' '}>>' ']'")
: bi_scope.
(* No ∀∀ binders, no RET binders *)
Notation "'<<{' α '}>>' e @ E '<<{' ∃∃ y1 .. yn , β '|' 'RET' v ; POST '}>>'" :=
(atomic_wp (TA:=TeleO)
(TB:=TeleS (λ y1, .. (TeleS (λ yn, TeleO)) .. ))
(TP:=TeleO)
e%E
E
(tele_app $ α%I)
(tele_app $ tele_app $ λ y1, .. (λ yn, β%I) .. )
(tele_app $ tele_app $ λ y1, .. (λ yn, tele_app $ Some POST%I) .. )
(tele_app $ tele_app $ λ y1, .. (λ yn, tele_app $ v%V) .. )
)
(at level 20, E, β, α, v, POST at level 200, y1 binder, yn binder,
format "'[hv' '<<{' '[' α ']' '}>>' '/ ' e @ E '/' '<<{' '[' ∃∃ y1 .. yn , '/' β '|' '/' RET v ; '/' POST ']' '}>>' ']'")
: bi_scope.
(* No ∀∀ binders, no ∃∃ binders, no RET binders *)
Notation "'<<{' α '}>>' e @ E '<<{' β '|' 'RET' v ; POST '}>>'" :=
(atomic_wp (TA:=TeleO)
(TB:=TeleO)
(TP:=TeleO)
e%E
E
(tele_app $ α%I)
(tele_app $ tele_app β%I)
(tele_app $ tele_app $ tele_app $ Some POST%I)
(tele_app $ tele_app $ tele_app $ v%V)
)
(at level 20, E, β, α, v, POST at level 200,
format "'[hv' '<<{' '[' α ']' '}>>' '/ ' e @ E '/' '<<{' '[' β '|' '/' RET v ; '/' POST ']' '}>>' ']'")
: bi_scope.
(* No RET binders, no POST *)
Notation "'<<{' ∀∀ x1 .. xn , α '}>>' e @ E '<<{' ∃∃ y1 .. yn , β '|' 'RET' v '}>>'" :=
(atomic_wp (TA:=TeleS (λ x1, .. (TeleS (λ xn, TeleO)) .. ))
(TB:=TeleS (λ y1, .. (TeleS (λ yn, TeleO)) .. ))
(TP:=TeleO)
e%E
E
(tele_app $ λ x1, .. (λ xn, α%I) ..)
(tele_app $ λ x1, .. (λ xn,
tele_app $ λ y1, .. (λ yn, β%I) ..
) .. )
(tele_app $ λ x1, .. (λ xn,
tele_app $ λ y1, .. (λ yn, tele_app None) ..
) .. )
(tele_app $ λ x1, .. (λ xn,
tele_app $ λ y1, .. (λ yn, tele_app v%V) ..
) .. )
)
(at level 20, E, β, α, v at level 200, x1 binder, xn binder, y1 binder, yn binder,
format "'[hv' '<<{' '[' ∀∀ x1 .. xn , '/' α ']' '}>>' '/ ' e @ E '/' '<<{' '[' ∃∃ y1 .. yn , '/' β '|' '/' RET v ']' '}>>' ']'")
: bi_scope.
(* No ∃∃ binders, no RET binders, no POST *)
Notation "'<<{' ∀∀ x1 .. xn , α '}>>' e @ E '<<{' β '|' 'RET' v '}>>'" :=
(atomic_wp (TA:=TeleS (λ x1, .. (TeleS (λ xn, TeleO)) .. ))
(TB:=TeleO)
(TP:=TeleO)
e%E
E
(tele_app $ λ x1, .. (λ xn, α%I) ..)
(tele_app $ λ x1, .. (λ xn, tele_app β%I) .. )
(tele_app $ λ x1, .. (λ xn, tele_app $ tele_app None) .. )
(tele_app $ λ x1, .. (λ xn, tele_app $ tele_app v%V) .. )
)
(at level 20, E, β, α, v at level 200, x1 binder, xn binder,
format "'[hv' '<<{' '[' ∀∀ x1 .. xn , '/' α ']' '}>>' '/ ' e @ E '/' '<<{' '[' β '|' '/' RET v ']' '}>>' ']'")
: bi_scope.
(* No ∀∀ binders, no RET binders, no POST *)
Notation "'<<{' α '}>>' e @ E '<<{' ∃∃ y1 .. yn , β '|' 'RET' v '}>>'" :=
(atomic_wp (TA:=TeleO)
(TB:=TeleS (λ y1, .. (TeleS (λ yn, TeleO)) .. ))
(TP:=TeleO)
e%E
E
(tele_app α%I)
(tele_app $ tele_app $ λ y1, .. (λ yn, β%I) .. )
(tele_app $ tele_app $ λ y1, .. (λ yn, tele_app None) .. )
(tele_app $ tele_app $ λ y1, .. (λ yn, tele_app v%V) .. )
)
(at level 20, E, β, α, v at level 200, y1 binder, yn binder,
format "'[hv' '<<{' '[' α ']' '}>>' '/ ' e @ E '/' '<<{' '[' ∃∃ y1 .. yn , '/' β '|' '/' RET v ']' '}>>' ']'")
: bi_scope.
(* No ∀∀ binders, no ∃∃ binders, no RET binders, no POST *)
Notation "'<<{' α '}>>' e @ E '<<{' β '|' 'RET' v '}>>'" :=
(atomic_wp (TA:=TeleO)
(TB:=TeleO)
(TP:=TeleO)
e%E
E
(tele_app α%I)
(tele_app $ tele_app β%I)
(tele_app $ tele_app $ tele_app None)
(tele_app $ tele_app $ tele_app v%V)
)
(at level 20, E, β, α, v at level 200,
format "'[hv' '<<{' '[' α ']' '}>>' '/ ' e @ E '/' '<<{' '[' β '|' '/' RET v ']' '}>>' ']'")
: bi_scope.
(** Theory *)
Section lemmas.
Context `{!irisGS_gen hlc Λ Σ} {TA TB TP : tele}.
Notation iProp := (iProp Σ).
Implicit Types (α : TA iProp) (β : TA TB iProp) (POST : TA TB TP option iProp) (f : TA TB TP val Λ).
(* Atomic triples imply sequential triples. *)
Lemma atomic_wp_seq e E α β POST f :
atomic_wp e E α β POST f -∗
Φ, .. x, α x -∗ (.. y, β x y -∗ .. z, POST x y z -∗? Φ (f x y z)) -∗ WP e {{ Φ }}.
Proof.
iIntros "Hwp" (Φ x) "Hα HΦ".
iApply (wp_frame_wand with "HΦ"). iApply "Hwp".
iAuIntro. iAaccIntro with "Hα"; first by eauto. iIntros (y) "Hβ !>".
(* FIXME: Using ssreflect rewrite does not work, see Coq bug #7773. *)
rewrite ->!tele_app_bind. iIntros (z) "Hpost HΦ". iApply ("HΦ" with "Hβ Hpost").
Qed.
(** This version matches the Texan triple, i.e., with a later in front of the
[(∀.. y, β x y -∗ Φ (f x y))]. *)
Lemma atomic_wp_seq_step e E α β POST f :
TCEq (to_val e) None
atomic_wp e E α β POST f -∗
Φ, .. x, α x -∗ (.. y, β x y -∗ .. z, POST x y z -∗? Φ (f x y z)) -∗ WP e {{ Φ }}.
Proof.
iIntros (?) "H"; iIntros (Φ x) "Hα HΦ".
iApply (wp_step_fupd _ _ _ (.. y : TB, _)
with "[$HΦ //]"); first done.
iApply (atomic_wp_seq with "H Hα").
iIntros "%y Hβ %z Hpost HΦ". iApply ("HΦ" with "Hβ Hpost").
Qed.
(* Sequential triples with the empty mask for a physically atomic [e] are atomic. *)
Lemma atomic_seq_wp_atomic e E α β POST f `{!Atomic WeaklyAtomic e} :
( Φ, .. x, α x -∗ (.. y, β x y -∗ .. z, POST x y z -∗? Φ (f x y z)) -∗ WP e @ {{ Φ }}) -∗
atomic_wp e E α β POST f.
Proof.
iIntros "Hwp" (Φ) "AU". iMod "AU" as (x) "[Hα [_ Hclose]]".
iApply ("Hwp" with "Hα"). iIntros "%y Hβ %z Hpost".
iMod ("Hclose" with "Hβ") as "HΦ".
rewrite ->!tele_app_bind. iApply "HΦ". done.
Qed.
(** Sequential triples with a persistent precondition and no initial quantifier
are atomic. *)
Lemma persistent_seq_wp_atomic e E (α : [tele] iProp) (β : [tele] TB iProp)
(POST : [tele] TB TP option iProp) (f : [tele] TB TP val Λ)
{HP : Persistent (α [tele_arg])} :
( Φ, α [tele_arg] -∗ (.. y, β [tele_arg] y -∗ .. z, POST [tele_arg] y z -∗? Φ (f [tele_arg] y z)) -∗ WP e {{ Φ }}) -∗
atomic_wp e E α β POST f.
Proof.
simpl in HP. iIntros "Hwp" (Φ) "HΦ". iApply fupd_wp.
iMod ("HΦ") as "[#Hα [Hclose _]]". iMod ("Hclose" with "Hα") as "HΦ".
iApply wp_fupd. iApply ("Hwp" with "Hα"). iIntros "!> %y Hβ %z Hpost".
iMod ("HΦ") as "[_ [_ Hclose]]". iMod ("Hclose" with "Hβ") as "HΦ".
(* FIXME: Using ssreflect rewrite does not work, see Coq bug #7773. *)
rewrite ->!tele_app_bind. iApply "HΦ". done.
Qed.
Lemma atomic_wp_mask_weaken e E1 E2 α β POST f :
E1 E2 atomic_wp e E1 α β POST f -∗ atomic_wp e E2 α β POST f.
Proof.
iIntros (HE) "Hwp". iIntros (Φ) "AU". iApply "Hwp".
iApply atomic_update_mask_weaken; last done. set_solver.
Qed.
(** We can open invariants around atomic triples.
(Just for demonstration purposes; we always use [iInv] in proofs.) *)
Lemma atomic_wp_inv e E α β POST f N I :
N E
atomic_wp e (E N) (λ.. x, I α x) (λ.. x y, I β x y) POST f -∗
inv N I -∗ atomic_wp e E α β POST f.
Proof.
intros ?. iIntros "Hwp #Hinv" (Φ) "AU". iApply "Hwp". iAuIntro.
iInv N as "HI". iApply (aacc_aupd with "AU"); first solve_ndisj.
iIntros (x) "Hα". iAaccIntro with "[HI Hα]"; rewrite ->!tele_app_bind; first by iFrame.
- (* abort *)
iIntros "[HI $]". by eauto with iFrame.
- (* commit *)
iIntros (y). rewrite ->!tele_app_bind. iIntros "[HI Hβ]". iRight.
iExists y. rewrite ->!tele_app_bind. by eauto with iFrame.
Qed.
End lemmas.
(** An axiomatization of evaluation-context based languages, including a proof
that this gives rise to a "language" in the Iris sense. *)
From iris.prelude Require Export prelude.
From iris.program_logic Require Import language.
From iris.prelude Require Import options.
(** TAKE CARE: When you define an [ectxLanguage] canonical structure for your
language, you need to also define a corresponding [language] canonical
structure. Use the coercion [LanguageOfEctx] as defined in the bottom of this
file for doing that. *)
Section ectx_language_mixin.
Context {expr val ectx state observation : Type}.
Context (of_val : val expr).
Context (to_val : expr option val).
Context (empty_ectx : ectx).
Context (comp_ectx : ectx ectx ectx).
Context (fill : ectx expr expr).
Context (base_step : expr state list observation expr state list expr Prop).
Record EctxLanguageMixin := {
mixin_to_of_val v : to_val (of_val v) = Some v;
mixin_of_to_val e v : to_val e = Some v of_val v = e;
mixin_val_base_stuck e1 σ1 κ e2 σ2 efs :
base_step e1 σ1 κ e2 σ2 efs to_val e1 = None;
mixin_fill_empty e : fill empty_ectx e = e;
mixin_fill_comp K1 K2 e : fill K1 (fill K2 e) = fill (comp_ectx K1 K2) e;
mixin_fill_inj K : Inj (=) (=) (fill K);
mixin_fill_val K e : is_Some (to_val (fill K e)) is_Some (to_val e);
(** Given a base redex [e1_redex] somewhere in a term, and another
decomposition of the same term into [fill K' e1'] such that [e1'] is not
a value, then the base redex context is [e1']'s context [K'] filled with
another context [K'']. In particular, this implies [e1 = fill K''
e1_redex] by [fill_inj], i.e., [e1]' contains the base redex.)
This implies there can be only one base redex, see
[base_redex_unique]. *)
mixin_step_by_val K' K_redex e1' e1_redex σ1 κ e2 σ2 efs :
fill K' e1' = fill K_redex e1_redex
to_val e1' = None
base_step e1_redex σ1 κ e2 σ2 efs
K'', K_redex = comp_ectx K' K'';
(** If [fill K e] takes a base step, then either [e] is a value or [K] is
the empty evaluation context. In other words, if [e] is not a value
wrapping it in a context does not add new base redex positions. *)
mixin_base_ctx_step_val K e σ1 κ e2 σ2 efs :
base_step (fill K e) σ1 κ e2 σ2 efs is_Some (to_val e) K = empty_ectx;
}.
End ectx_language_mixin.
Structure ectxLanguage := EctxLanguage {
expr : Type;
val : Type;
ectx : Type;
state : Type;
observation : Type;
of_val : val expr;
to_val : expr option val;
empty_ectx : ectx;
comp_ectx : ectx ectx ectx;
fill : ectx expr expr;
base_step : expr state list observation expr state list expr Prop;
ectx_language_mixin :
EctxLanguageMixin of_val to_val empty_ectx comp_ectx fill base_step
}.
Bind Scope expr_scope with expr.
Bind Scope val_scope with val.
Global Arguments EctxLanguage {_ _ _ _ _ _ _ _ _ _ _} _.
Global Arguments of_val {_} _.
Global Arguments to_val {_} _.
Global Arguments empty_ectx {_}.
Global Arguments comp_ectx {_} _ _.
Global Arguments fill {_} _ _.
Global Arguments base_step {_} _ _ _ _ _ _.
(* From an ectx_language, we can construct a language. *)
Section ectx_language.
Context {Λ : ectxLanguage}.
Implicit Types v : val Λ.
Implicit Types e : expr Λ.
Implicit Types K : ectx Λ.
(* Only project stuff out of the mixin that is not also in language *)
Lemma val_base_stuck e1 σ1 κ e2 σ2 efs : base_step e1 σ1 κ e2 σ2 efs to_val e1 = None.
Proof. apply ectx_language_mixin. Qed.
Lemma fill_empty e : fill empty_ectx e = e.
Proof. apply ectx_language_mixin. Qed.
Lemma fill_comp K1 K2 e : fill K1 (fill K2 e) = fill (comp_ectx K1 K2) e.
Proof. apply ectx_language_mixin. Qed.
Global Instance fill_inj K : Inj (=) (=) (fill K).
Proof. apply ectx_language_mixin. Qed.
Lemma fill_val K e : is_Some (to_val (fill K e)) is_Some (to_val e).
Proof. apply ectx_language_mixin. Qed.
Lemma step_by_val K' K_redex e1' e1_redex σ1 κ e2 σ2 efs :
fill K' e1' = fill K_redex e1_redex
to_val e1' = None
base_step e1_redex σ1 κ e2 σ2 efs
K'', K_redex = comp_ectx K' K''.
Proof. apply ectx_language_mixin. Qed.
Lemma base_ctx_step_val K e σ1 κ e2 σ2 efs :
base_step (fill K e) σ1 κ e2 σ2 efs is_Some (to_val e) K = empty_ectx.
Proof. apply ectx_language_mixin. Qed.
Definition base_reducible (e : expr Λ) (σ : state Λ) :=
κ e' σ' efs, base_step e σ κ e' σ' efs.
Definition base_reducible_no_obs (e : expr Λ) (σ : state Λ) :=
e' σ' efs, base_step e σ [] e' σ' efs.
Definition base_irreducible (e : expr Λ) (σ : state Λ) :=
κ e' σ' efs, ¬base_step e σ κ e' σ' efs.
Definition base_stuck (e : expr Λ) (σ : state Λ) :=
to_val e = None base_irreducible e σ.
(* All non-value redexes are at the root. In other words, all sub-redexes are
values. *)
Definition sub_redexes_are_values (e : expr Λ) :=
K e', e = fill K e' to_val e' = None K = empty_ectx.
Inductive prim_step (e1 : expr Λ) (σ1 : state Λ) (κ : list (observation Λ))
(e2 : expr Λ) (σ2 : state Λ) (efs : list (expr Λ)) : Prop :=
Ectx_step K e1' e2' :
e1 = fill K e1' e2 = fill K e2'
base_step e1' σ1 κ e2' σ2 efs prim_step e1 σ1 κ e2 σ2 efs.
Lemma Ectx_step' K e1 σ1 κ e2 σ2 efs :
base_step e1 σ1 κ e2 σ2 efs prim_step (fill K e1) σ1 κ (fill K e2) σ2 efs.
Proof. econstructor; eauto. Qed.
Definition ectx_lang_mixin : LanguageMixin of_val to_val prim_step.
Proof.
split.
- apply ectx_language_mixin.
- apply ectx_language_mixin.
- intros ?????? [??? -> -> ?%val_base_stuck].
apply eq_None_not_Some. by intros ?%fill_val%eq_None_not_Some.
Qed.
Canonical Structure ectx_lang : language := Language ectx_lang_mixin.
Definition base_atomic (a : atomicity) (e : expr Λ) : Prop :=
σ κ e' σ' efs,
base_step e σ κ e' σ' efs
if a is WeaklyAtomic then irreducible e' σ' else is_Some (to_val e').
(** * Some lemmas about this language *)
Lemma fill_not_val K e : to_val e = None to_val (fill K e) = None.
Proof. rewrite !eq_None_not_Some. eauto using fill_val. Qed.
Lemma base_reducible_no_obs_reducible e σ :
base_reducible_no_obs e σ base_reducible e σ.
Proof. intros (?&?&?&?). eexists. eauto. Qed.
Lemma not_base_reducible e σ : ¬base_reducible e σ base_irreducible e σ.
Proof. unfold base_reducible, base_irreducible. naive_solver. Qed.
(** The decomposition into base redex and context is unique.
In all sensible instances, [comp_ectx K' empty_ectx] will be the same as
[K'], so the conclusion is [K = K' ∧ e = e'], but we do not require a law
to actually prove that so we cannot use that fact here. *)
Lemma base_redex_unique K K' e e' σ :
fill K e = fill K' e'
base_reducible e σ
base_reducible e' σ
K = comp_ectx K' empty_ectx e = e'.
Proof.
intros Heq (κ & e2 & σ2 & efs & Hred) (κ' & e2' & σ2' & efs' & Hred').
edestruct (step_by_val K' K e' e) as [K'' HK];
[by eauto using val_base_stuck..|].
subst K. move: Heq. rewrite -fill_comp. intros <-%(inj (fill _)).
destruct (base_ctx_step_val _ _ _ _ _ _ _ Hred') as [[]%not_eq_None_Some|HK''].
{ by eapply val_base_stuck. }
subst K''. rewrite fill_empty. done.
Qed.
Lemma base_prim_step e1 σ1 κ e2 σ2 efs :
base_step e1 σ1 κ e2 σ2 efs prim_step e1 σ1 κ e2 σ2 efs.
Proof. apply Ectx_step with empty_ectx; by rewrite ?fill_empty. Qed.
Lemma base_step_not_stuck e σ κ e' σ' efs : base_step e σ κ e' σ' efs not_stuck e σ.
Proof. rewrite /not_stuck /reducible /=. eauto 10 using base_prim_step. Qed.
Lemma fill_prim_step K e1 σ1 κ e2 σ2 efs :
prim_step e1 σ1 κ e2 σ2 efs prim_step (fill K e1) σ1 κ (fill K e2) σ2 efs.
Proof.
destruct 1 as [K' e1' e2' -> ->].
rewrite !fill_comp. by econstructor.
Qed.
Lemma fill_reducible K e σ : reducible e σ reducible (fill K e) σ.
Proof.
intros (κ&e'&σ'&efs&?). exists κ, (fill K e'), σ', efs.
by apply fill_prim_step.
Qed.
Lemma fill_reducible_no_obs K e σ : reducible_no_obs e σ reducible_no_obs (fill K e) σ.
Proof.
intros (e'&σ'&efs&?). exists (fill K e'), σ', efs.
by apply fill_prim_step.
Qed.
Lemma base_prim_reducible e σ : base_reducible e σ reducible e σ.
Proof. intros (κ&e'&σ'&efs&?). eexists κ, e', σ', efs. by apply base_prim_step. Qed.
Lemma base_prim_fill_reducible e K σ :
base_reducible e σ reducible (fill K e) σ.
Proof. intro. by apply fill_reducible, base_prim_reducible. Qed.
Lemma base_prim_reducible_no_obs e σ : base_reducible_no_obs e σ reducible_no_obs e σ.
Proof. intros (e'&σ'&efs&?). eexists e', σ', efs. by apply base_prim_step. Qed.
Lemma base_prim_irreducible e σ : irreducible e σ base_irreducible e σ.
Proof.
rewrite -not_reducible -not_base_reducible. eauto using base_prim_reducible.
Qed.
Lemma base_prim_fill_reducible_no_obs e K σ :
base_reducible_no_obs e σ reducible_no_obs (fill K e) σ.
Proof. intro. by apply fill_reducible_no_obs, base_prim_reducible_no_obs. Qed.
Lemma prim_base_reducible e σ :
reducible e σ sub_redexes_are_values e base_reducible e σ.
Proof.
intros (κ&e'&σ'&efs&[K e1' e2' -> -> Hstep]) ?.
assert (K = empty_ectx) as -> by eauto 10 using val_base_stuck.
rewrite fill_empty /base_reducible; eauto.
Qed.
Lemma prim_base_irreducible e σ :
base_irreducible e σ sub_redexes_are_values e irreducible e σ.
Proof.
rewrite -not_reducible -not_base_reducible. eauto using prim_base_reducible.
Qed.
Lemma base_stuck_stuck e σ :
base_stuck e σ sub_redexes_are_values e stuck e σ.
Proof.
intros [] ?. split; first done.
by apply prim_base_irreducible.
Qed.
Lemma ectx_language_atomic a e :
base_atomic a e sub_redexes_are_values e Atomic a e.
Proof.
intros Hatomic_step Hatomic_fill σ κ e' σ' efs [K e1' e2' -> -> Hstep].
assert (K = empty_ectx) as -> by eauto 10 using val_base_stuck.
rewrite fill_empty. eapply Hatomic_step. by rewrite fill_empty.
Qed.
Lemma base_reducible_prim_step_ctx K e1 σ1 κ e2 σ2 efs :
base_reducible e1 σ1
prim_step (fill K e1) σ1 κ e2 σ2 efs
e2', e2 = fill K e2' base_step e1 σ1 κ e2' σ2 efs.
Proof.
intros (κ'&e2''&σ2''&efs''&HhstepK) [K' e1' e2' HKe1 -> Hstep].
edestruct (step_by_val K) as [K'' ?];
eauto using val_base_stuck; simplify_eq/=.
rewrite -fill_comp in HKe1; simplify_eq.
exists (fill K'' e2'). rewrite fill_comp; split; first done.
apply base_ctx_step_val in HhstepK as [[v ?]|?]; simplify_eq.
{ apply val_base_stuck in Hstep; simplify_eq. }
by rewrite !fill_empty.
Qed.
Lemma base_reducible_prim_step e1 σ1 κ e2 σ2 efs :
base_reducible e1 σ1
prim_step e1 σ1 κ e2 σ2 efs
base_step e1 σ1 κ e2 σ2 efs.
Proof.
intros.
edestruct (base_reducible_prim_step_ctx empty_ectx) as (?&?&?);
rewrite ?fill_empty; eauto.
by simplify_eq; rewrite fill_empty.
Qed.
(* Every evaluation context is a context. *)
Global Instance ectx_lang_ctx K : LanguageCtx (fill K).
Proof.
split; simpl.
- eauto using fill_not_val.
- intros ?????? [K' e1' e2' Heq1 Heq2 Hstep].
by exists (comp_ectx K K') e1' e2'; rewrite ?Heq1 ?Heq2 ?fill_comp.
- intros e1 σ1 κ e2 σ2 efs Hnval [K'' e1'' e2'' Heq1 -> Hstep].
destruct (step_by_val K K'' e1 e1'' σ1 κ e2'' σ2 efs) as [K' ->]; eauto.
rewrite -fill_comp in Heq1; apply (inj (fill _)) in Heq1.
exists (fill K' e2''); rewrite -fill_comp; split; auto.
econstructor; eauto.
Qed.
Record pure_base_step (e1 e2 : expr Λ) := {
pure_base_step_safe σ1 : base_reducible_no_obs e1 σ1;
pure_base_step_det σ1 κ e2' σ2 efs :
base_step e1 σ1 κ e2' σ2 efs κ = [] σ2 = σ1 e2' = e2 efs = []
}.
Lemma pure_base_step_pure_step e1 e2 : pure_base_step e1 e2 pure_step e1 e2.
Proof.
intros [Hp1 Hp2]. split.
- intros σ. destruct (Hp1 σ) as (e2' & σ2 & efs & ?).
eexists e2', σ2, efs. by apply base_prim_step.
- intros σ1 κ e2' σ2 efs ?%base_reducible_prim_step; eauto using base_reducible_no_obs_reducible.
Qed.
(** This is not an instance because HeapLang's [wp_pure] tactic already takes
care of handling the evaluation context. So the instance is redundant.
If you are defining your own language and your [wp_pure] works
differently, you might want to specialize this lemma to your language and
register that as an instance. *)
Lemma pure_exec_fill K φ n e1 e2 :
PureExec φ n e1 e2
PureExec φ n (fill K e1) (fill K e2).
Proof. apply: pure_exec_ctx. Qed.
End ectx_language.
Global Arguments ectx_lang : clear implicits.
Global Arguments Ectx_step {Λ e1 σ1 κ e2 σ2 efs}.
Coercion ectx_lang : ectxLanguage >-> language.
(* This definition makes sure that the fields of the [language] record do not
refer to the projections of the [ectxLanguage] record but to the actual fields
of the [ectxLanguage] record. This is crucial for canonical structure search to
work.
Note that this trick no longer works when we switch to canonical projections
because then the pattern match [let '...] will be desugared into projections. *)
Definition LanguageOfEctx (Λ : ectxLanguage) : language :=
let '@EctxLanguage E V C St K of_val to_val empty comp fill base mix := Λ in
@Language E V St K of_val to_val _
(@ectx_lang_mixin (@EctxLanguage E V C St K of_val to_val empty comp fill base mix)).
Global Arguments LanguageOfEctx : simpl never.