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 9827 additions and 0 deletions
(** This library provides assertions that represent "unique tokens".
The [token γ] assertion provides ownership of the token named [γ],
and the key lemma [token_exclusive] proves only one token exists. *)
From iris.algebra Require Import excl.
From iris.proofmode Require Import proofmode.
From iris.base_logic.lib Require Export own.
From iris.prelude Require Import options.
(** The CMRA we need. *)
Class tokenG Σ := TokenG {
#[local] token_inG :: inG Σ (exclR unitO);
}.
Global Hint Mode tokenG - : typeclass_instances.
Definition tokenΣ : gFunctors :=
#[ GFunctor (exclR unitO) ].
Global Instance subG_tokenΣ Σ : subG tokenΣ Σ tokenG Σ.
Proof. solve_inG. Qed.
Local Definition token_def `{!tokenG Σ} (γ : gname) : iProp Σ :=
own γ (Excl ()).
Local Definition token_aux : seal (@token_def). Proof. by eexists. Qed.
Definition token := token_aux.(unseal).
Local Definition token_unseal :
@token = @token_def := token_aux.(seal_eq).
Global Arguments token {Σ _} γ.
Local Ltac unseal := rewrite ?token_unseal /token_def.
Section lemmas.
Context `{!tokenG Σ}.
Global Instance token_timeless γ : Timeless (token γ).
Proof. unseal. apply _. Qed.
Lemma token_alloc_strong (P : gname Prop) :
pred_infinite P
|==> γ, P γ token γ.
Proof. unseal. intros. iApply own_alloc_strong; done. Qed.
Lemma token_alloc :
|==> γ, token γ.
Proof. unseal. iApply own_alloc. done. Qed.
Lemma token_exclusive γ :
token γ -∗ token γ -∗ False.
Proof.
unseal. iIntros "Htok1 Htok2".
iCombine "Htok1 Htok2" gives %[].
Qed.
Global Instance token_combine_gives γ :
CombineSepGives (token γ) (token γ) False⌝.
Proof.
rewrite /CombineSepGives. iIntros "[H1 H2]".
iDestruct (token_exclusive with "H1 H2") as %[].
Qed.
End lemmas.
From stdpp Require Export coPset.
From iris.algebra Require Import gmap_view gset coPset.
From iris.bi Require Import lib.cmra.
From iris.proofmode Require Import proofmode.
From iris.base_logic.lib Require Export own.
From iris.prelude Require Import options.
(** All definitions in this file are internal to [fancy_updates] with the
exception of what's in the [wsatGS] module. The module [wsatGS] is thus exported in
[fancy_updates], where [wsat] is only imported. *)
Module wsatGS.
Class wsatGpreS (Σ : gFunctors) : Set := WsatGpreS {
wsatGpreS_inv : inG Σ (gmap_viewR positive (agreeR $ laterO (iPropO Σ)));
wsatGpreS_enabled : inG Σ coPset_disjR;
wsatGpreS_disabled : inG Σ (gset_disjR positive);
}.
Class wsatGS (Σ : gFunctors) : Set := WsatG {
wsat_inG : wsatGpreS Σ;
invariant_name : gname;
enabled_name : gname;
disabled_name : gname;
}.
Definition wsatΣ : gFunctors :=
#[GFunctor (gmap_viewRF positive (agreeRF $ laterOF idOF));
GFunctor coPset_disjR;
GFunctor (gset_disjR positive)].
Global Instance subG_wsatΣ {Σ} : subG wsatΣ Σ wsatGpreS Σ.
Proof. solve_inG. Qed.
End wsatGS.
Import wsatGS.
(* Make the instances local to this file. We cannot use #[local] in the code
above, as that would make the instances local to the module. *)
Local Existing Instances wsat_inG wsatGpreS_inv wsatGpreS_enabled wsatGpreS_disabled.
Definition invariant_unfold {Σ} (P : iProp Σ) : later (iProp Σ) :=
Next P.
Definition ownI `{!wsatGS Σ} (i : positive) (P : iProp Σ) : iProp Σ :=
own invariant_name
(gmap_view_frag i DfracDiscarded (to_agree $ invariant_unfold P)).
Global Typeclasses Opaque ownI.
Global Instance: Params (@invariant_unfold) 1 := {}.
Global Instance: Params (@ownI) 3 := {}.
Definition ownE `{!wsatGS Σ} (E : coPset) : iProp Σ :=
own enabled_name (CoPset E).
Global Typeclasses Opaque ownE.
Global Instance: Params (@ownE) 3 := {}.
Definition ownD `{!wsatGS Σ} (E : gset positive) : iProp Σ :=
own disabled_name (GSet E).
Global Typeclasses Opaque ownD.
Global Instance: Params (@ownD) 3 := {}.
Definition wsat `{!wsatGS Σ} : iProp Σ :=
locked ( I : gmap positive (iProp Σ),
own invariant_name
(gmap_view_auth (DfracOwn 1) (to_agree <$> (invariant_unfold <$> I)))
[ map] i Q I, Q ownD {[i]} ownE {[i]})%I.
Section wsat.
Context `{!wsatGS Σ}.
Implicit Types P : iProp Σ.
(* Invariants *)
Local Instance invariant_unfold_contractive : Contractive (@invariant_unfold Σ).
Proof. solve_contractive. Qed.
Global Instance ownI_contractive i : Contractive (@ownI Σ _ i).
Proof. solve_contractive. Qed.
Global Instance ownI_persistent i P : Persistent (ownI i P).
Proof. rewrite /ownI. apply _. Qed.
Lemma ownE_empty : |==> ownE ∅.
Proof.
rewrite /bi_emp_valid.
by rewrite (own_unit (coPset_disjUR) enabled_name).
Qed.
Lemma ownE_op E1 E2 : E1 ## E2 ownE (E1 E2) ⊣⊢ ownE E1 ownE E2.
Proof. intros. by rewrite /ownE -own_op coPset_disj_union. Qed.
Lemma ownE_disjoint E1 E2 : ownE E1 ownE E2 E1 ## E2⌝.
Proof. rewrite /ownE -own_op own_valid. by iIntros (?%coPset_disj_valid_op). Qed.
Lemma ownE_op' E1 E2 : E1 ## E2 ownE (E1 E2) ⊣⊢ ownE E1 ownE E2.
Proof.
iSplit; [iIntros "[% ?]"; by iApply ownE_op|].
iIntros "HE". iDestruct (ownE_disjoint with "HE") as %?.
iSplit; first done. iApply ownE_op; by try iFrame.
Qed.
Lemma ownE_singleton_twice i : ownE {[i]} ownE {[i]} False.
Proof. rewrite ownE_disjoint. iIntros (?); set_solver. Qed.
Lemma ownD_empty : |==> ownD ∅.
Proof.
rewrite /bi_emp_valid.
by rewrite (own_unit (gset_disjUR positive) disabled_name).
Qed.
Lemma ownD_op E1 E2 : E1 ## E2 ownD (E1 E2) ⊣⊢ ownD E1 ownD E2.
Proof. intros. by rewrite /ownD -own_op gset_disj_union. Qed.
Lemma ownD_disjoint E1 E2 : ownD E1 ownD E2 E1 ## E2⌝.
Proof. rewrite /ownD -own_op own_valid. by iIntros (?%gset_disj_valid_op). Qed.
Lemma ownD_op' E1 E2 : E1 ## E2 ownD (E1 E2) ⊣⊢ ownD E1 ownD E2.
Proof.
iSplit; [iIntros "[% ?]"; by iApply ownD_op|].
iIntros "HE". iDestruct (ownD_disjoint with "HE") as %?.
iSplit; first done. iApply ownD_op; by try iFrame.
Qed.
Lemma ownD_singleton_twice i : ownD {[i]} ownD {[i]} False.
Proof. rewrite ownD_disjoint. iIntros (?); set_solver. Qed.
Lemma invariant_lookup (I : gmap positive (iProp Σ)) i P :
own invariant_name (gmap_view_auth (DfracOwn 1) (to_agree <$> (invariant_unfold <$> I)))
own invariant_name (gmap_view_frag i DfracDiscarded (to_agree $ invariant_unfold P))
Q, I !! i = Some Q (Q P).
Proof.
rewrite -own_op own_valid gmap_view_both_validI_total.
iIntros "[%Q' (_& _ & HQ' & Hval & Hincl)]". rewrite !lookup_fmap.
case: (I !! i)=> [Q|] /=; last first.
{ iDestruct "HQ'" as %?. done. }
iDestruct "HQ'" as %[= <-]. iExists Q; iSplit; first done.
rewrite to_agree_includedI internal_eq_sym -later_equivI. done.
Qed.
Lemma ownI_open i P : wsat ownI i P ownE {[i]} wsat P ownD {[i]}.
Proof.
rewrite /ownI /wsat -!lock.
iIntros "(Hw & Hi & HiE)". iDestruct "Hw" as (I) "[Hw HI]".
iDestruct (invariant_lookup I i P with "[$]") as (Q ?) "#HPQ".
iDestruct (big_sepM_delete _ _ i with "HI") as "[[[HQ $]|HiE'] HI]"; eauto.
- iSplitR "HQ"; last by iNext; iRewrite -"HPQ".
iExists I. iFrame "Hw". iApply (big_sepM_delete _ _ i); eauto.
iFrame "HI"; eauto.
- iDestruct (ownE_singleton_twice with "[$HiE $HiE']") as %[].
Qed.
Lemma ownI_close i P : wsat ownI i P P ownD {[i]} wsat ownE {[i]}.
Proof.
rewrite /ownI /wsat -!lock.
iIntros "(Hw & Hi & HP & HiD)". iDestruct "Hw" as (I) "[Hw HI]".
iDestruct (invariant_lookup with "[$]") as (Q ?) "#HPQ".
iDestruct (big_sepM_delete _ _ i with "HI") as "[[[HQ ?]|$] HI]"; eauto.
- iDestruct (ownD_singleton_twice with "[$]") as %[].
- iExists I. iFrame "Hw". iApply (big_sepM_delete _ _ i); eauto.
iFrame "HI". iLeft. iFrame "HiD". by iNext; iRewrite "HPQ".
Qed.
Lemma ownI_alloc φ P :
( E : gset positive, i, i E φ i)
wsat P ==∗ i, φ i wsat ownI i P.
Proof.
iIntros (Hfresh) "[Hw HP]". rewrite /wsat -!lock.
iDestruct "Hw" as (I) "[Hw HI]".
iMod (own_unit (gset_disjUR positive) disabled_name) as "HE".
iMod (own_updateP with "[$]") as "HE".
{ apply (gset_disj_alloc_empty_updateP_strong' (λ i, I !! i = None φ i)).
intros E. destruct (Hfresh (E dom I))
as (i & [? HIi%not_elem_of_dom]%not_elem_of_union & ?); eauto. }
iDestruct "HE" as (X) "[Hi HE]"; iDestruct "Hi" as %(i & -> & HIi & ?).
iMod (own_update with "Hw") as "[Hw HiP]".
{ eapply (gmap_view_alloc _ i DfracDiscarded (to_agree _)); [|done..].
by rewrite /= !lookup_fmap HIi. }
iModIntro; iExists i; iSplit; [done|]. rewrite /ownI; iFrame "HiP".
iExists (<[i:=P]>I); iSplitL "Hw".
{ by rewrite !fmap_insert. }
iApply (big_sepM_insert _ I); first done.
iFrame "HI". iLeft. by rewrite /ownD; iFrame.
Qed.
Lemma ownI_alloc_open φ P :
( E : gset positive, i, i E φ i)
wsat ==∗ i, φ i (ownE {[i]} -∗ wsat) ownI i P ownD {[i]}.
Proof.
iIntros (Hfresh) "Hw". rewrite /wsat -!lock. iDestruct "Hw" as (I) "[Hw HI]".
iMod (own_unit (gset_disjUR positive) disabled_name) as "HD".
iMod (own_updateP with "[$]") as "HD".
{ apply (gset_disj_alloc_empty_updateP_strong' (λ i, I !! i = None φ i)).
intros E. destruct (Hfresh (E dom I))
as (i & [? HIi%not_elem_of_dom]%not_elem_of_union & ?); eauto. }
iDestruct "HD" as (X) "[Hi HD]"; iDestruct "Hi" as %(i & -> & HIi & ?).
iMod (own_update with "Hw") as "[Hw HiP]".
{ eapply (gmap_view_alloc _ i DfracDiscarded (to_agree _)); [|done..].
by rewrite /= !lookup_fmap HIi. }
iModIntro; iExists i; iSplit; [done|]. rewrite /ownI; iFrame "HiP".
rewrite -/(ownD _). iFrame "HD".
iIntros "HE". iExists (<[i:=P]>I); iSplitL "Hw".
{ by rewrite !fmap_insert. }
iApply (big_sepM_insert _ I); first done.
iFrame "HI". by iRight.
Qed.
End wsat.
(* Allocation of an initial world *)
Lemma wsat_alloc `{!wsatGpreS Σ} : |==> _ : wsatGS Σ, wsat ownE .
Proof.
iIntros.
iMod (own_alloc (gmap_view_auth (DfracOwn 1) )) as (γI) "HI";
first by apply gmap_view_auth_valid.
iMod (own_alloc (CoPset )) as (γE) "HE"; first done.
iMod (own_alloc (GSet )) as (γD) "HD"; first done.
iModIntro; iExists (WsatG _ _ γI γE γD).
rewrite /wsat /ownE -lock; iFrame.
iExists ∅. rewrite fmap_empty big_opM_empty. by iFrame.
Qed.
From iris.algebra Require Import proofmode_classes.
From iris.proofmode Require Import classes.
From iris.base_logic Require Export derived.
From iris.prelude Require Import options.
Import base_logic.bi.uPred.
(* Setup of the proof mode *)
Section class_instances.
Context {M : ucmra}.
Implicit Types P Q R : uPred M.
Global Instance into_pure_cmra_valid `{!CmraDiscrete A} (a : A) :
@IntoPure (uPredI M) ( a) ( a).
Proof. rewrite /IntoPure. by rewrite uPred.discrete_valid. Qed.
Global Instance from_pure_cmra_valid {A : cmra} (a : A) :
@FromPure (uPredI M) false ( a) ( a).
Proof.
rewrite /FromPure /=. eapply bi.pure_elim=> // ?.
rewrite -uPred.cmra_valid_intro //.
Qed.
Global Instance from_sep_ownM (a b1 b2 : M) :
IsOp a b1 b2
FromSep (uPred_ownM a) (uPred_ownM b1) (uPred_ownM b2).
Proof. intros. by rewrite /FromSep -ownM_op -is_op. Qed.
(* TODO: Improve this instance with generic own simplification machinery
once https://gitlab.mpi-sws.org/iris/iris/-/issues/460 is fixed *)
(* Cost > 50 to give priority to [combine_sep_as_fractional]. *)
Global Instance combine_sep_as_ownM (a b1 b2 : M) :
IsOp a b1 b2
CombineSepAs (uPred_ownM b1) (uPred_ownM b2) (uPred_ownM a) | 60.
Proof. intros. by rewrite /CombineSepAs -ownM_op -is_op. Qed.
(* TODO: Improve this instance with generic own validity simplification
machinery once https://gitlab.mpi-sws.org/iris/iris/-/issues/460 is fixed *)
Global Instance combine_sep_gives_ownM (b1 b2 : M) :
CombineSepGives (uPred_ownM b1) (uPred_ownM b2) ( (b1 b2)).
Proof.
intros. rewrite /CombineSepGives -ownM_op ownM_valid.
by apply: bi.persistently_intro.
Qed.
Global Instance from_sep_ownM_core_id (a b1 b2 : M) :
IsOp a b1 b2 TCOr (CoreId b1) (CoreId b2)
FromAnd (uPred_ownM a) (uPred_ownM b1) (uPred_ownM b2).
Proof.
intros ? H. rewrite /FromAnd (is_op a) ownM_op.
destruct H; by rewrite bi.persistent_and_sep.
Qed.
Global Instance into_and_ownM p (a b1 b2 : M) :
IsOp a b1 b2 IntoAnd p (uPred_ownM a) (uPred_ownM b1) (uPred_ownM b2).
Proof.
intros. apply bi.intuitionistically_if_mono. by rewrite (is_op a) ownM_op bi.sep_and.
Qed.
Global Instance into_sep_ownM (a b1 b2 : M) :
IsOp a b1 b2 IntoSep (uPred_ownM a) (uPred_ownM b1) (uPred_ownM b2).
Proof. intros. by rewrite /IntoSep (is_op a) ownM_op. Qed.
End class_instances.
From iris.algebra Require Export cmra updates.
From iris.algebra Require Export stepindex_finite.
From iris.bi Require Import notation.
From iris.prelude Require Import options.
Local Hint Extern 1 (_ _) => etrans; [eassumption|] : core.
Local Hint Extern 1 (_ _) => etrans; [|eassumption] : core.
Local Hint Extern 10 (_ _) => lia : core.
(** The basic definition of the uPred type, its metric and functor laws.
You probably do not want to import this file. Instead, import
base_logic.base_logic; that will also give you all the primitive
and many derived laws for the logic. *)
(** A good way of understanding this definition of the uPred OFE is to
consider the OFE uPred0 of monotonous SProp predicates. That is,
uPred0 is the OFE of non-expansive functions from M to SProp that
are monotonous with respect to CMRA inclusion. This notion of
monotonicity has to be stated in the SProp logic. Together with the
usual closedness property of SProp, this gives exactly uPred_mono.
Then, we quotient uPred0 *in the siProp logic* with respect to
equivalence on valid elements of M. That is, we quotient with
respect to the following *siProp* equivalence relation:
P1 ≡ P2 := ∀ x, ✓ x → (P1(x) ↔ P2(x)) (1)
When seen from the ambiant logic, obtaining this quotient requires
definig both a custom Equiv and Dist.
It is worth noting that this equivalence relation admits canonical
representatives. More precisely, one can show that every
equivalence class contains exactly one element P0 such that:
∀ x, (✓ x → P0(x)) → P0(x) (2)
(Again, this assertion has to be understood in siProp). Intuitively,
this says that P0 trivially holds whenever the resource is invalid.
Starting from any element P, one can find this canonical
representative by choosing:
P0(x) := ✓ x → P(x) (3)
Hence, as an alternative definition of uPred, we could use the set
of canonical representatives (i.e., the subtype of monotonous
siProp predicates that verify (2)). This alternative definition would
save us from using a quotient. However, the definitions of the various
connectives would get more complicated, because we have to make sure
they all verify (2), which sometimes requires some adjustments. We
would moreover need to prove one more property for every logical
connective.
*)
(** Note that, somewhat curiously, [uPred M] is *not* in general a Camera,
at least not if all propositions are considered "valid" Camera elements.
It fails to satisfy the extension axiom. Here is the counterexample:
We use [M := (option Ex {A,B})^2] -- so we have pairs
whose components are ε, A or B.
Let
[[
P r n := (ownM (A,A) ∧ ▷ False) ∨ ownM (A,B) ∨ ownM (B,A) ∨ ownM (B,B)
↔ r = (A,A) ∧ n = 0 ∨
r = (A,B) ∨
r = (B,A) ∨
r = (B,B)
Q1 r n := ownM (A, ε) ∨ ownM (B, ε)
↔ (A, ε) ≼ r ∨ (B, ε) ≼ r
("Left component is not ε")
Q2 r n := ownM (ε, A) ∨ ownM (ε, B)
↔ (ε, A) ≼ r ∨ (ε, B) ≼ r
("Right component is not ε")
]]
These are all sufficiently closed and non-expansive and whatnot.
We have [P ≡{0}≡ Q1 * Q2]. So assume extension holds, then we get Q1', Q2'
such that
[[
P ≡ Q1' ∗ Q2'
Q1 ≡{0}≡ Q1'
Q2 ≡{0}≡ Q2'
]]
Now comes the contradiction:
We know that [P (A,A) 1] does *not* hold, but I am going to show that
[(Q1' ∗ Q2') (A,A) 1] holds, which would be a contraction.
To this end, I will show (a) [Q1' (A,ε) 1] and (b) [Q2' (ε,A) 1].
The result [(Q1' ∗ Q2') (A,A)] follows from [(A,ε) ⋅ (ε,A) = (A,A)].
(a) Proof of [Q1' (A,ε) 1].
We have [P (A,B) 1], and thus [Q1' r1 1] and [Q2' r2 1] for some
[r1 ⋅ r2 = (A,B)]. There are four possible decompositions [r1 ⋅ r2]:
- [(ε,ε) ⋅ (A,B)]: This would give us [Q1' (ε,ε) 1], from which we
obtain (through down-closure and the equality [Q1 ≡{0}≡ Q1'] above) that
[Q1 (ε,ε) 0]. However, we know that's false.
- [(A,B) ⋅ (ε,ε)]: Can be excluded for similar reasons
(the second resource must not be ε in the 2nd component).
- [(ε,B) ⋅ (A,ε)]: Can be excluded for similar reasons
(the first resource must not be ε in the 1st component).
- [(A,ε) ⋅ (ε,B)]: This gives us the desired [Q1' (A,ε) 1].
(b) Proof of [Q2' (ε,A) 1].
We have [P (B,A) 1], and thus [Q1' r1 1] and [Q2' r2 1] for some
[r1 ⋅ r2 = (B,A)]. There are again four possible decompositions,
and like above we can exclude three of them. This leaves us with
[(B,ε) ⋅ (ε,A)] and thus [Q2' (ε,A) 1].
This completes the proof.
*)
Record uPred (M : ucmra) : Type := UPred {
uPred_holds : nat M Prop;
uPred_mono n1 n2 x1 x2 :
uPred_holds n1 x1 x1 {n2} x2 n2 n1 uPred_holds n2 x2
}.
(** When working in the model, it is convenient to be able to treat [uPred] as
[nat → M → Prop]. But we only want to locally break the [uPred] abstraction
this way. *)
Local Coercion uPred_holds : uPred >-> Funclass.
Bind Scope bi_scope with uPred.
Global Arguments uPred_holds {_} _ _ _ : simpl never.
Add Printing Constructor uPred.
Global Instance: Params (@uPred_holds) 3 := {}.
Section cofe.
Context {M : ucmra}.
Inductive uPred_equiv' (P Q : uPred M) : Prop :=
{ uPred_in_equiv : n x, {n} x P n x Q n x }.
Local Instance uPred_equiv : Equiv (uPred M) := uPred_equiv'.
Inductive uPred_dist' (n : nat) (P Q : uPred M) : Prop :=
{ uPred_in_dist : n' x, n' n {n'} x P n' x Q n' x }.
Local Instance uPred_dist : Dist (uPred M) := uPred_dist'.
Definition uPred_ofe_mixin : OfeMixin (uPred M).
Proof.
apply ofe_mixin_finite.
- intros P Q; split.
+ by intros HPQ n; split=> i x ??; apply HPQ.
+ intros HPQ; split=> n x ?; apply HPQ with n; auto.
- intros n; split.
+ by intros P; split=> x i.
+ by intros P Q HPQ; split=> x i ??; symmetry; apply HPQ.
+ intros P Q Q' HP HQ; split=> i x ??.
by trans (Q i x);[apply HP|apply HQ].
- intros n P Q HPQ. split=> i x ??; apply HPQ; eauto with lia.
Qed.
Canonical Structure uPredO : ofe := Ofe (uPred M) uPred_ofe_mixin.
Program Definition uPred_compl : Compl uPredO := λ c,
{| uPred_holds n x := n', n' n {n'} x c n' n' x |}.
Next Obligation.
move=> /= c n1 n2 x1 x2 HP Hx12 Hn12 n3 Hn23 Hv. eapply uPred_mono.
- eapply HP, cmra_validN_includedN, cmra_includedN_le=>//; lia.
- eapply cmra_includedN_le=>//; lia.
- done.
Qed.
Global Program Instance uPred_cofe : Cofe uPredO := cofe_finite uPred_compl _.
Next Obligation.
intros n c; split=>i x Hin Hv.
etrans; [|by symmetry; apply (chain_cauchy c i n)].
split=>H; [by apply H|].
repeat intro. apply (chain_cauchy c _ i)=>//. by eapply uPred_mono.
Qed.
End cofe.
Global Arguments uPredO : clear implicits.
Global Instance uPred_ne {M} (P : uPred M) n : Proper (dist n ==> iff) (P n).
Proof.
intros x1 x2 Hx; split=> ?; eapply uPred_mono; eauto; by rewrite Hx.
Qed.
Global Instance uPred_proper {M} (P : uPred M) n : Proper (() ==> iff) (P n).
Proof. by intros x1 x2 Hx; apply uPred_ne, equiv_dist. Qed.
Lemma uPred_holds_ne {M} (P Q : uPred M) n1 n2 x :
P {n2} Q n2 n1 {n2} x Q n1 x P n2 x.
Proof.
intros [Hne] ???. eapply Hne; try done. eauto using uPred_mono, cmra_validN_le.
Qed.
(* Equivalence to the definition of uPred in the appendix. *)
Lemma uPred_alt {M : ucmra} (P: nat M Prop) :
( n1 n2 x1 x2, P n1 x1 x1 {n1} x2 n2 n1 P n2 x2)
( ( x n1 n2, n2 n1 P n1 x P n2 x) (* Pointwise down-closed *)
( n x1 x2, x1 {n} x2 m, m n P m x1 P m x2) (* Non-expansive *)
( n x1 x2, x1 {n} x2 m, m n P m x1 P m x2) (* Monotonicity *)
).
Proof.
(* Provide this lemma to eauto. *)
assert ( n1 n2 (x1 x2 : M), n2 n1 x1 {n1} x2 x1 {n2} x2).
{ intros ????? H. eapply cmra_includedN_le; last done. by rewrite H. }
(* Now go ahead. *)
split.
- intros Hupred. repeat split; eauto using cmra_includedN_le.
- intros (Hdown & _ & Hmono) **. eapply Hmono; [done..|]. eapply Hdown; done.
Qed.
(** functor *)
Program Definition uPred_map {M1 M2 : ucmra} (f : M2 -n> M1)
`{!CmraMorphism f} (P : uPred M1) :
uPred M2 := {| uPred_holds n x := P n (f x) |}.
Next Obligation. naive_solver eauto using uPred_mono, cmra_morphism_monotoneN. Qed.
Global Instance uPred_map_ne {M1 M2 : ucmra} (f : M2 -n> M1)
`{!CmraMorphism f} n : Proper (dist n ==> dist n) (uPred_map f).
Proof.
intros x1 x2 Hx; split=> n' y ??.
split; apply Hx; auto using cmra_morphism_validN.
Qed.
Lemma uPred_map_id {M : ucmra} (P : uPred M): uPred_map cid P P.
Proof. by split=> n x ?. Qed.
Lemma uPred_map_compose {M1 M2 M3 : ucmra} (f : M1 -n> M2) (g : M2 -n> M3)
`{!CmraMorphism f, !CmraMorphism g} (P : uPred M3):
uPred_map (g f) P uPred_map f (uPred_map g P).
Proof. by split=> n x Hx. Qed.
Lemma uPred_map_ext {M1 M2 : ucmra} (f g : M1 -n> M2)
`{!CmraMorphism f} `{!CmraMorphism g}:
( x, f x g x) x, uPred_map f x uPred_map g x.
Proof. intros Hf P; split=> n x Hx /=; by rewrite /uPred_holds /= Hf. Qed.
Definition uPredO_map {M1 M2 : ucmra} (f : M2 -n> M1) `{!CmraMorphism f} :
uPredO M1 -n> uPredO M2 := OfeMor (uPred_map f : uPredO M1 uPredO M2).
Lemma uPredO_map_ne {M1 M2 : ucmra} (f g : M2 -n> M1)
`{!CmraMorphism f, !CmraMorphism g} n :
f {n} g uPredO_map f {n} uPredO_map g.
Proof.
by intros Hfg P; split=> n' y ??;
rewrite /uPred_holds /= (dist_le _ _ _ _(Hfg y)).
Qed.
Program Definition uPredOF (F : urFunctor) : oFunctor := {|
oFunctor_car A _ B _ := uPredO (urFunctor_car F B A);
oFunctor_map A1 _ A2 _ B1 _ B2 _ fg := uPredO_map (urFunctor_map F (fg.2, fg.1))
|}.
Next Obligation.
intros F A1 ? A2 ? B1 ? B2 ? n P Q HPQ.
apply uPredO_map_ne, urFunctor_map_ne; split; by apply HPQ.
Qed.
Next Obligation.
intros F A ? B ? P; simpl. rewrite -{2}(uPred_map_id P).
apply uPred_map_ext=>y. by rewrite urFunctor_map_id.
Qed.
Next Obligation.
intros F A1 ? A2 ? A3 ? B1 ? B2 ? B3 ? f g f' g' P; simpl. rewrite -uPred_map_compose.
apply uPred_map_ext=>y; apply urFunctor_map_compose.
Qed.
Global Instance uPredOF_contractive F :
urFunctorContractive F oFunctorContractive (uPredOF F).
Proof.
intros ? A1 ? A2 ? B1 ? B2 ? n P Q HPQ.
apply uPredO_map_ne, urFunctor_map_contractive.
destruct HPQ as [HPQ]; split; intros ??; split; simpl; by eapply HPQ.
Qed.
(** logical entailement *)
Inductive uPred_entails {M} (P Q : uPred M) : Prop :=
{ uPred_in_entails : n x, {n} x P n x Q n x }.
Global Hint Resolve uPred_mono : uPred_def.
(** logical connectives *)
Local Program Definition uPred_pure_def {M} (φ : Prop) : uPred M :=
{| uPred_holds n x := φ |}.
Solve Obligations with done.
Local Definition uPred_pure_aux : seal (@uPred_pure_def). Proof. by eexists. Qed.
Definition uPred_pure := uPred_pure_aux.(unseal).
Global Arguments uPred_pure {M}.
Local Definition uPred_pure_unseal :
@uPred_pure = @uPred_pure_def := uPred_pure_aux.(seal_eq).
Local Program Definition uPred_and_def {M} (P Q : uPred M) : uPred M :=
{| uPred_holds n x := P n x Q n x |}.
Solve Obligations with naive_solver eauto 2 with uPred_def.
Local Definition uPred_and_aux : seal (@uPred_and_def). Proof. by eexists. Qed.
Definition uPred_and := uPred_and_aux.(unseal).
Global Arguments uPred_and {M}.
Local Definition uPred_and_unseal :
@uPred_and = @uPred_and_def := uPred_and_aux.(seal_eq).
Local Program Definition uPred_or_def {M} (P Q : uPred M) : uPred M :=
{| uPred_holds n x := P n x Q n x |}.
Solve Obligations with naive_solver eauto 2 with uPred_def.
Local Definition uPred_or_aux : seal (@uPred_or_def). Proof. by eexists. Qed.
Definition uPred_or := uPred_or_aux.(unseal).
Global Arguments uPred_or {M}.
Local Definition uPred_or_unseal :
@uPred_or = @uPred_or_def := uPred_or_aux.(seal_eq).
Local Program Definition uPred_impl_def {M} (P Q : uPred M) : uPred M :=
{| uPred_holds n x := n' x',
x x' n' n {n'} x' P n' x' Q n' x' |}.
Next Obligation.
intros M P Q n1 n1' x1 x1' HPQ [x2 Hx1'] Hn1 n2 x3 [x4 Hx3] ?; simpl in *.
rewrite Hx3 (dist_le _ _ _ _ Hx1'); auto. intros ??.
eapply HPQ; auto. exists (x2 x4); by rewrite assoc.
Qed.
Local Definition uPred_impl_aux : seal (@uPred_impl_def). Proof. by eexists. Qed.
Definition uPred_impl := uPred_impl_aux.(unseal).
Global Arguments uPred_impl {M}.
Local Definition uPred_impl_unseal :
@uPred_impl = @uPred_impl_def := uPred_impl_aux.(seal_eq).
Local Program Definition uPred_forall_def {M A} (Ψ : A uPred M) : uPred M :=
{| uPred_holds n x := a, Ψ a n x |}.
Solve Obligations with naive_solver eauto 2 with uPred_def.
Local Definition uPred_forall_aux : seal (@uPred_forall_def). Proof. by eexists. Qed.
Definition uPred_forall := uPred_forall_aux.(unseal).
Global Arguments uPred_forall {M A}.
Local Definition uPred_forall_unseal :
@uPred_forall = @uPred_forall_def := uPred_forall_aux.(seal_eq).
Local Program Definition uPred_exist_def {M A} (Ψ : A uPred M) : uPred M :=
{| uPred_holds n x := a, Ψ a n x |}.
Solve Obligations with naive_solver eauto 2 with uPred_def.
Local Definition uPred_exist_aux : seal (@uPred_exist_def). Proof. by eexists. Qed.
Definition uPred_exist := uPred_exist_aux.(unseal).
Global Arguments uPred_exist {M A}.
Local Definition uPred_exist_unseal :
@uPred_exist = @uPred_exist_def := uPred_exist_aux.(seal_eq).
Local Program Definition uPred_internal_eq_def {M} {A : ofe} (a1 a2 : A) : uPred M :=
{| uPred_holds n x := a1 {n} a2 |}.
Solve Obligations with naive_solver eauto 2 using dist_le.
Local Definition uPred_internal_eq_aux : seal (@uPred_internal_eq_def).
Proof. by eexists. Qed.
Definition uPred_internal_eq := uPred_internal_eq_aux.(unseal).
Global Arguments uPred_internal_eq {M A}.
Local Definition uPred_internal_eq_unseal :
@uPred_internal_eq = @uPred_internal_eq_def := uPred_internal_eq_aux.(seal_eq).
Local Program Definition uPred_sep_def {M} (P Q : uPred M) : uPred M :=
{| uPred_holds n x := x1 x2, x {n} x1 x2 P n x1 Q n x2 |}.
Next Obligation.
intros M P Q n1 n2 x y (x1&x2&Hx&?&?) [z Hy] Hn.
exists x1, (x2 z); split_and?; eauto using uPred_mono, cmra_includedN_l.
rewrite Hy. eapply dist_le, Hn. by rewrite Hx assoc.
Qed.
Local Definition uPred_sep_aux : seal (@uPred_sep_def). Proof. by eexists. Qed.
Definition uPred_sep := uPred_sep_aux.(unseal).
Global Arguments uPred_sep {M}.
Local Definition uPred_sep_unseal :
@uPred_sep = @uPred_sep_def := uPred_sep_aux.(seal_eq).
Local Program Definition uPred_wand_def {M} (P Q : uPred M) : uPred M :=
{| uPred_holds n x := n' x',
n' n {n'} (x x') P n' x' Q n' (x x') |}.
Next Obligation.
intros M P Q n1 n1' x1 x1' HPQ ? Hn n3 x3 ???; simpl in *.
eapply uPred_mono with n3 (x1 x3);
eauto using cmra_validN_includedN, cmra_monoN_r, cmra_includedN_le.
Qed.
Local Definition uPred_wand_aux : seal (@uPred_wand_def). Proof. by eexists. Qed.
Definition uPred_wand := uPred_wand_aux.(unseal).
Global Arguments uPred_wand {M}.
Local Definition uPred_wand_unseal :
@uPred_wand = @uPred_wand_def := uPred_wand_aux.(seal_eq).
(* Equivalently, this could be `∀ y, P n y`. That's closer to the intuition
of "embedding the step-indexed logic in Iris", but the two are equivalent
because Iris is afine. The following is easier to work with. *)
Local Program Definition uPred_plainly_def {M} (P : uPred M) : uPred M :=
{| uPred_holds n x := P n ε |}.
Solve Obligations with naive_solver eauto using uPred_mono, ucmra_unit_validN.
Local Definition uPred_plainly_aux : seal (@uPred_plainly_def). Proof. by eexists. Qed.
Definition uPred_plainly := uPred_plainly_aux.(unseal).
Global Arguments uPred_plainly {M}.
Local Definition uPred_plainly_unseal :
@uPred_plainly = @uPred_plainly_def := uPred_plainly_aux.(seal_eq).
Local Program Definition uPred_persistently_def {M} (P : uPred M) : uPred M :=
{| uPred_holds n x := P n (core x) |}.
Solve Obligations with naive_solver eauto using uPred_mono, cmra_core_monoN.
Local Definition uPred_persistently_aux : seal (@uPred_persistently_def).
Proof. by eexists. Qed.
Definition uPred_persistently := uPred_persistently_aux.(unseal).
Global Arguments uPred_persistently {M}.
Local Definition uPred_persistently_unseal :
@uPred_persistently = @uPred_persistently_def := uPred_persistently_aux.(seal_eq).
Local Program Definition uPred_later_def {M} (P : uPred M) : uPred M :=
{| uPred_holds n x := match n return _ with 0 => True | S n' => P n' x end |}.
Next Obligation.
intros M P [|n1] [|n2] x1 x2; eauto using uPred_mono, cmra_includedN_S with lia.
Qed.
Local Definition uPred_later_aux : seal (@uPred_later_def). Proof. by eexists. Qed.
Definition uPred_later := uPred_later_aux.(unseal).
Global Arguments uPred_later {M}.
Local Definition uPred_later_unseal :
@uPred_later = @uPred_later_def := uPred_later_aux.(seal_eq).
Local Program Definition uPred_ownM_def {M : ucmra} (a : M) : uPred M :=
{| uPred_holds n x := a {n} x |}.
Next Obligation.
intros M a n1 n2 x1 x [a' Hx1] [x2 Hx] Hn.
exists (a' x2). rewrite Hx. eapply dist_le, Hn. rewrite (assoc op) -Hx1 //.
Qed.
Local Definition uPred_ownM_aux : seal (@uPred_ownM_def). Proof. by eexists. Qed.
Definition uPred_ownM := uPred_ownM_aux.(unseal).
Global Arguments uPred_ownM {M}.
Local Definition uPred_ownM_unseal :
@uPred_ownM = @uPred_ownM_def := uPred_ownM_aux.(seal_eq).
Local Program Definition uPred_cmra_valid_def {M} {A : cmra} (a : A) : uPred M :=
{| uPred_holds n x := {n} a |}.
Solve Obligations with naive_solver eauto 2 using cmra_validN_le.
Local Definition uPred_cmra_valid_aux : seal (@uPred_cmra_valid_def).
Proof. by eexists. Qed.
Definition uPred_cmra_valid := uPred_cmra_valid_aux.(unseal).
Global Arguments uPred_cmra_valid {M A}.
Local Definition uPred_cmra_valid_unseal :
@uPred_cmra_valid = @uPred_cmra_valid_def := uPred_cmra_valid_aux.(seal_eq).
Local Program Definition uPred_bupd_def {M} (Q : uPred M) : uPred M :=
{| uPred_holds n x := k yf,
k n {k} (x yf) x', {k} (x' yf) Q k x' |}.
Next Obligation.
intros M Q n1 n2 x1 x2 HQ [x3 Hx] Hn k yf Hk.
rewrite (dist_le _ _ _ _ Hx Hk). intros Hxy.
destruct (HQ k (x3 yf)) as (x'&?&?); [auto|by rewrite assoc|].
exists (x' x3); split; first by rewrite -assoc.
eauto using uPred_mono, cmra_includedN_l.
Qed.
Local Definition uPred_bupd_aux : seal (@uPred_bupd_def). Proof. by eexists. Qed.
Definition uPred_bupd := uPred_bupd_aux.(unseal).
Global Arguments uPred_bupd {M}.
Local Definition uPred_bupd_unseal :
@uPred_bupd = @uPred_bupd_def := uPred_bupd_aux.(seal_eq).
(** Global uPred-specific Notation *)
Notation "✓ x" := (uPred_cmra_valid x) (at level 20) : bi_scope.
(** Primitive logical rules.
These are not directly usable later because they do not refer to the BI
connectives. *)
Module uPred_primitive.
Local Definition uPred_unseal :=
(uPred_pure_unseal, uPred_and_unseal, uPred_or_unseal, uPred_impl_unseal,
uPred_forall_unseal, uPred_exist_unseal, uPred_internal_eq_unseal,
uPred_sep_unseal, uPred_wand_unseal, uPred_plainly_unseal,
uPred_persistently_unseal, uPred_later_unseal, uPred_ownM_unseal,
uPred_cmra_valid_unseal, @uPred_bupd_unseal).
Ltac unseal :=
rewrite !uPred_unseal /=.
Section primitive.
Context {M : ucmra}.
Implicit Types φ : Prop.
Implicit Types P Q : uPred M.
Implicit Types A : Type.
Local Arguments uPred_holds {_} !_ _ _ /.
Local Hint Immediate uPred_in_entails : core.
(** The notations below are implicitly local due to the section, so we do not
mind the overlap with the general BI notations. *)
Notation "P ⊢ Q" := (@uPred_entails M P%I Q%I) : stdpp_scope.
Notation "(⊢)" := (@uPred_entails M) (only parsing) : stdpp_scope.
Notation "P ⊣⊢ Q" := (@uPred_equiv M P%I Q%I) : stdpp_scope.
Notation "(⊣⊢)" := (@uPred_equiv M) (only parsing) : stdpp_scope.
Notation "'True'" := (uPred_pure True) : bi_scope.
Notation "'False'" := (uPred_pure False) : bi_scope.
Notation "'⌜' φ '⌝'" := (uPred_pure φ%type%stdpp) : bi_scope.
Infix "∧" := uPred_and : bi_scope.
Infix "∨" := uPred_or : bi_scope.
Infix "→" := uPred_impl : bi_scope.
Notation "∀ x .. y , P" :=
(uPred_forall (λ x, .. (uPred_forall (λ y, P)) ..)) : bi_scope.
Notation "∃ x .. y , P" :=
(uPred_exist (λ x, .. (uPred_exist (λ y, P)) ..)) : bi_scope.
Infix "∗" := uPred_sep : bi_scope.
Infix "-∗" := uPred_wand : bi_scope.
Notation "□ P" := (uPred_persistently P) : bi_scope.
Notation "■ P" := (uPred_plainly P) : bi_scope.
Notation "x ≡ y" := (uPred_internal_eq x y) : bi_scope.
Notation "▷ P" := (uPred_later P) : bi_scope.
Notation "|==> P" := (uPred_bupd P) : bi_scope.
(** Entailment *)
Lemma entails_po : PreOrder ().
Proof.
split.
- by intros P; split=> x i.
- by intros P Q Q' HP HQ; split=> x i ??; apply HQ, HP.
Qed.
Lemma entails_anti_sym : AntiSymm (⊣⊢) ().
Proof. intros P Q HPQ HQP; split=> x n; by split; [apply HPQ|apply HQP]. Qed.
Lemma equiv_entails P Q : (P ⊣⊢ Q) (P Q) (Q P).
Proof.
split.
- intros HPQ; split; split=> x i; apply HPQ.
- intros [??]. exact: entails_anti_sym.
Qed.
Lemma entails_lim (cP cQ : chain (uPredO M)) :
( n, cP n cQ n) compl cP compl cQ.
Proof.
intros Hlim; split=> n m ? HP.
eapply uPred_holds_ne, Hlim; [..|apply: HP]; rewrite ?conv_compl; eauto.
Qed.
(** Non-expansiveness and setoid morphisms *)
Lemma pure_ne n : Proper (iff ==> dist n) (@uPred_pure M).
Proof. intros φ1 φ2 . by unseal; split=> -[|m] ?; try apply . Qed.
Lemma and_ne : NonExpansive2 (@uPred_and M).
Proof.
intros n P P' HP Q Q' HQ; unseal; split=> x n' ??.
split; (intros [??]; split; [by apply HP|by apply HQ]).
Qed.
Lemma or_ne : NonExpansive2 (@uPred_or M).
Proof.
intros n P P' HP Q Q' HQ; split=> x n' ??.
unseal; split; (intros [?|?]; [left; by apply HP|right; by apply HQ]).
Qed.
Lemma impl_ne :
NonExpansive2 (@uPred_impl M).
Proof.
intros n P P' HP Q Q' HQ; split=> x n' ??.
unseal; split; intros HPQ x' n'' ????; apply HQ, HPQ, HP; auto.
Qed.
Lemma sep_ne : NonExpansive2 (@uPred_sep M).
Proof.
intros n P P' HP Q Q' HQ; split=> n' x ??.
unseal; split; intros (x1&x2&?&?&?); ofe_subst x;
exists x1, x2; split_and!; try (apply HP || apply HQ);
eauto using cmra_validN_op_l, cmra_validN_op_r.
Qed.
Lemma wand_ne :
NonExpansive2 (@uPred_wand M).
Proof.
intros n P P' HP Q Q' HQ; split=> n' x ??; unseal; split; intros HPQ x' n'' ???;
apply HQ, HPQ, HP; eauto using cmra_validN_op_r.
Qed.
Lemma internal_eq_ne (A : ofe) :
NonExpansive2 (@uPred_internal_eq M A).
Proof.
intros n x x' Hx y y' Hy; split=> n' z; unseal; split; intros; simpl in *.
- by rewrite -(dist_le _ _ _ _ Hx) -?(dist_le _ _ _ _ Hy); auto.
- by rewrite (dist_le _ _ _ _ Hx) ?(dist_le _ _ _ _ Hy); auto.
Qed.
Lemma forall_ne A n :
Proper (pointwise_relation _ (dist n) ==> dist n) (@uPred_forall M A).
Proof.
by intros Ψ1 Ψ2 ; unseal; split=> n' x; split; intros HP a; apply .
Qed.
Lemma exist_ne A n :
Proper (pointwise_relation _ (dist n) ==> dist n) (@uPred_exist M A).
Proof.
intros Ψ1 Ψ2 .
unseal; split=> n' x ??; split; intros [a ?]; exists a; by apply .
Qed.
Lemma later_contractive : Contractive (@uPred_later M).
Proof.
unseal; intros [|n] P Q HPQ; split=> -[|n'] x ?? //=; try lia.
eapply HPQ; eauto using cmra_validN_S.
Qed.
Lemma plainly_ne : NonExpansive (@uPred_plainly M).
Proof.
intros n P1 P2 HP.
unseal; split=> n' x; split; apply HP; eauto using ucmra_unit_validN.
Qed.
Lemma persistently_ne : NonExpansive (@uPred_persistently M).
Proof.
intros n P1 P2 HP.
unseal; split=> n' x; split; apply HP; eauto using cmra_core_validN.
Qed.
Lemma ownM_ne : NonExpansive (@uPred_ownM M).
Proof.
intros n a b Ha.
unseal; split=> n' x ? /=. by rewrite (dist_le _ _ _ _ Ha).
Qed.
Lemma cmra_valid_ne {A : cmra} :
NonExpansive (@uPred_cmra_valid M A).
Proof.
intros n a b Ha; unseal; split=> n' x ? /=.
by rewrite (dist_le _ _ _ _ Ha).
Qed.
Lemma bupd_ne : NonExpansive (@uPred_bupd M).
Proof.
intros n P Q HPQ.
unseal; split=> n' x; split; intros HP k yf ??;
destruct (HP k yf) as (x'&?&?); auto;
exists x'; split; auto; apply HPQ; eauto using cmra_validN_op_l.
Qed.
(** Introduction and elimination rules *)
Lemma pure_intro φ P : φ P φ⌝.
Proof. by intros ?; unseal; split. Qed.
Lemma pure_elim' φ P : (φ True P) φ P.
Proof. unseal; intros HP; split=> n x ??. by apply HP. Qed.
Lemma pure_forall_2 {A} (φ : A Prop) : ( x : A, φ x) ⌜∀ x : A, φ x⌝.
Proof. by unseal. Qed.
Lemma and_elim_l P Q : P Q P.
Proof. by unseal; split=> n x ? [??]. Qed.
Lemma and_elim_r P Q : P Q Q.
Proof. by unseal; split=> n x ? [??]. Qed.
Lemma and_intro P Q R : (P Q) (P R) P Q R.
Proof. intros HQ HR; unseal; split=> n x ??; by split; [apply HQ|apply HR]. Qed.
Lemma or_intro_l P Q : P P Q.
Proof. unseal; split=> n x ??; left; auto. Qed.
Lemma or_intro_r P Q : Q P Q.
Proof. unseal; split=> n x ??; right; auto. Qed.
Lemma or_elim P Q R : (P R) (Q R) P Q R.
Proof.
intros HP HQ; unseal; split=> n x ? [?|?].
- by apply HP.
- by apply HQ.
Qed.
Lemma impl_intro_r P Q R : (P Q R) P Q R.
Proof.
unseal; intros HQ; split=> n x ?? n' x' ????. apply HQ;
naive_solver eauto using uPred_mono, cmra_included_includedN.
Qed.
Lemma impl_elim_l' P Q R : (P Q R) P Q R.
Proof. unseal; intros HP ; split=> n x ? [??]; apply HP with n x; auto. Qed.
Lemma forall_intro {A} P (Ψ : A uPred M): ( a, P Ψ a) P a, Ψ a.
Proof. unseal; intros HPΨ; split=> n x ?? a; by apply HPΨ. Qed.
Lemma forall_elim {A} {Ψ : A uPred M} a : ( a, Ψ a) Ψ a.
Proof. unseal; split=> n x ? HP; apply HP. Qed.
Lemma exist_intro {A} {Ψ : A uPred M} a : Ψ a a, Ψ a.
Proof. unseal; split=> n x ??; by exists a. Qed.
Lemma exist_elim {A} (Φ : A uPred M) Q : ( a, Φ a Q) ( a, Φ a) Q.
Proof. unseal; intros HΦΨ; split=> n x ? [a ?]; by apply HΦΨ with a. Qed.
(** BI connectives *)
Lemma sep_mono P P' Q Q' : (P Q) (P' Q') P P' Q Q'.
Proof.
intros HQ HQ'; unseal.
split; intros n' x ? (x1&x2&?&?&?); exists x1,x2; ofe_subst x;
eauto 7 using cmra_validN_op_l, cmra_validN_op_r, uPred_in_entails.
Qed.
Lemma True_sep_1 P : P True P.
Proof.
unseal; split; intros n x ??. exists (core x), x. by rewrite cmra_core_l.
Qed.
Lemma True_sep_2 P : True P P.
Proof.
unseal; split; intros n x ? (x1&x2&?&_&?); ofe_subst;
eauto using uPred_mono, cmra_includedN_r.
Qed.
Lemma sep_comm' P Q : P Q Q P.
Proof.
unseal; split; intros n x ? (x1&x2&?&?&?); exists x2, x1; by rewrite (comm op).
Qed.
Lemma sep_assoc' P Q R : (P Q) R P (Q R).
Proof.
unseal; split; intros n x ? (x1&x2&Hx&(y1&y2&Hy&?&?)&?).
exists y1, (y2 x2); split_and?; auto.
+ by rewrite (assoc op) -Hy -Hx.
+ by exists y2, x2.
Qed.
Lemma wand_intro_r P Q R : (P Q R) P Q -∗ R.
Proof.
unseal=> HPQR; split=> n x ?? n' x' ???; apply HPQR; auto.
exists x, x'; split_and?; auto.
eapply uPred_mono with n x; eauto using cmra_validN_op_l.
Qed.
Lemma wand_elim_l' P Q R : (P Q -∗ R) P Q R.
Proof.
unseal =>HPQR. split; intros n x ? (?&?&?&?&?). ofe_subst.
eapply HPQR; eauto using cmra_validN_op_l.
Qed.
(** Persistently *)
Lemma persistently_mono P Q : (P Q) P Q.
Proof. intros HP; unseal; split=> n x ? /=. by apply HP, cmra_core_validN. Qed.
Lemma persistently_elim P : P P.
Proof.
unseal; split=> n x ? /=.
eauto using uPred_mono, cmra_included_core, cmra_included_includedN.
Qed.
Lemma persistently_idemp_2 P : P P.
Proof. unseal; split=> n x ?? /=. by rewrite cmra_core_idemp. Qed.
Lemma persistently_forall_2 {A} (Ψ : A uPred M) : ( a, Ψ a) ( a, Ψ a).
Proof. by unseal. Qed.
Lemma persistently_exist_1 {A} (Ψ : A uPred M) : ( a, Ψ a) ( a, Ψ a).
Proof. by unseal. Qed.
Lemma persistently_and_sep_l_1 P Q : P Q P Q.
Proof.
unseal; split=> n x ? [??]; exists (core x), x; simpl in *.
by rewrite cmra_core_l.
Qed.
(** Plainly *)
Lemma plainly_mono P Q : (P Q) P Q.
Proof. intros HP; unseal; split=> n x ? /=. apply HP, ucmra_unit_validN. Qed.
Lemma plainly_elim_persistently P : P P.
Proof. unseal; split; simpl; eauto using uPred_mono, ucmra_unit_leastN. Qed.
Lemma plainly_idemp_2 P : P P.
Proof. unseal; split=> n x ?? //. Qed.
Lemma plainly_forall_2 {A} (Ψ : A uPred M) : ( a, Ψ a) ( a, Ψ a).
Proof. by unseal. Qed.
Lemma plainly_exist_1 {A} (Ψ : A uPred M) : ( a, Ψ a) ( a, Ψ a).
Proof. by unseal. Qed.
Lemma prop_ext_2 P Q : ((P -∗ Q) (Q -∗ P)) P Q.
Proof.
unseal; split=> n x ? /=. setoid_rewrite (left_id ε op). split; naive_solver.
Qed.
(* The following two laws are very similar, and indeed they hold not just for □
and ■, but for any modality defined as `M P n x := ∀ y, R x y → P n y`. *)
Lemma persistently_impl_plainly P Q : ( P Q) ( P Q).
Proof.
unseal; split=> /= n x ? HPQ n' x' ????.
eapply uPred_mono with n' (core x)=>//; [|by apply cmra_included_includedN].
apply (HPQ n' x); eauto using cmra_validN_le.
Qed.
Lemma plainly_impl_plainly P Q : ( P Q) ( P Q).
Proof.
unseal; split=> /= n x ? HPQ n' x' ????.
eapply uPred_mono with n' ε=>//; [|by apply cmra_included_includedN].
apply (HPQ n' x); eauto using cmra_validN_le.
Qed.
(** Later *)
Lemma later_mono P Q : (P Q) P Q.
Proof.
unseal=> HP; split=>-[|n] x ??; [done|apply HP; eauto using cmra_validN_S].
Qed.
Lemma later_intro P : P P.
Proof.
unseal; split=> -[|n] /= x ? HP; first done.
apply uPred_mono with (S n) x; eauto using cmra_validN_S.
Qed.
Lemma later_forall_2 {A} (Φ : A uPred M) : ( a, Φ a) a, Φ a.
Proof. unseal; by split=> -[|n] x. Qed.
Lemma later_exist_false {A} (Φ : A uPred M) :
( a, Φ a) False ( a, Φ a).
Proof. unseal; split=> -[|[|n]] x /=; eauto. Qed.
Lemma later_sep_1 P Q : (P Q) P Q.
Proof.
unseal; split=> n x ?.
destruct n as [|n]; simpl.
{ by exists x, (core x); rewrite cmra_core_r. }
intros (x1&x2&Hx&?&?); destruct (cmra_extend n x x1 x2)
as (y1&y2&Hx'&Hy1&Hy2); eauto using cmra_validN_S; simpl in *.
exists y1, y2; split; [by rewrite Hx'|by rewrite Hy1 Hy2].
Qed.
Lemma later_sep_2 P Q : P Q (P Q).
Proof.
unseal; split=> n x ?.
destruct n as [|n]; simpl; [done|intros (x1&x2&Hx&?&?)].
exists x1, x2; eauto using dist_S.
Qed.
Lemma later_false_em P : P False ( False P).
Proof.
unseal; split=> -[|n] x ? /= HP; [by left|right].
intros [|n'] x' ????; eauto using uPred_mono, cmra_included_includedN.
Qed.
Lemma later_persistently_1 P : P P.
Proof. by unseal. Qed.
Lemma later_persistently_2 P : P P.
Proof. by unseal. Qed.
Lemma later_plainly_1 P : P P.
Proof. by unseal. Qed.
Lemma later_plainly_2 P : P P.
Proof. by unseal. Qed.
(** Internal equality *)
Lemma internal_eq_refl {A : ofe} P (a : A) : P (a a).
Proof. unseal; by split=> n x ??; simpl. Qed.
Lemma internal_eq_rewrite {A : ofe} a b (Ψ : A uPred M) :
NonExpansive Ψ a b Ψ a Ψ b.
Proof. intros . unseal; split=> n x ?? n' x' ??? Ha. by apply with n a. Qed.
Lemma fun_ext {A} {B : A ofe} (g1 g2 : discrete_fun B) :
( i, g1 i g2 i) g1 g2.
Proof. by unseal. Qed.
Lemma sig_eq {A : ofe} (P : A Prop) (x y : sigO P) :
proj1_sig x proj1_sig y x y.
Proof. by unseal. Qed.
Lemma later_eq_1 {A : ofe} (x y : A) : Next x Next y (x y).
Proof.
unseal. split. intros [|n]; simpl; [done|].
intros ?? Heq; apply Heq, SIdx.lt_succ_diag_r.
Qed.
Lemma later_eq_2 {A : ofe} (x y : A) : (x y) Next x Next y.
Proof.
unseal. split. intros n ? ? Hn; split; intros m Hlt; simpl in *.
destruct n as [|n]; first lia.
by eapply dist_le, SIdx.lt_succ_r.
Qed.
Lemma discrete_eq_1 {A : ofe} (a b : A) : Discrete a a b a b⌝.
Proof.
unseal=> ?. split=> n x ?. by apply (discrete_iff n).
Qed.
(** This is really just a special case of an entailment
between two [siProp], but we do not have the infrastructure
to express the more general case. This temporary proof rule will
be replaced by the proper one eventually. *)
Lemma internal_eq_entails {A B : ofe} (a1 a2 : A) (b1 b2 : B) :
(a1 a2 b1 b2) ( n, a1 {n} a2 b1 {n} b2).
Proof.
split.
- unseal=> -[Hsi] n. apply (Hsi _ ε), ucmra_unit_validN.
- unseal=> Hsi. split=>n x ?. apply Hsi.
Qed.
(** Basic update modality *)
Lemma bupd_intro P : P |==> P.
Proof.
unseal. split=> n x ? HP k yf ?; exists x; split; first done.
apply uPred_mono with n x; eauto using cmra_validN_op_l.
Qed.
Lemma bupd_mono P Q : (P Q) (|==> P) |==> Q.
Proof.
unseal. intros HPQ; split=> n x ? HP k yf ??.
destruct (HP k yf) as (x'&?&?); eauto.
exists x'; split; eauto using uPred_in_entails, cmra_validN_op_l.
Qed.
Lemma bupd_trans P : (|==> |==> P) |==> P.
Proof. unseal; split; naive_solver. Qed.
Lemma bupd_frame_r P R : (|==> P) R |==> P R.
Proof.
unseal; split; intros n x ? (x1&x2&Hx&HP&?) k yf ??.
destruct (HP k (x2 yf)) as (x'&?&?); eauto.
{ by rewrite assoc -(dist_le _ _ _ _ Hx). }
exists (x' x2); split; first by rewrite -assoc.
exists x', x2. eauto using uPred_mono, cmra_validN_op_l, cmra_validN_op_r.
Qed.
Lemma bupd_plainly P : (|==> P) P.
Proof.
unseal; split => n x Hnx /= Hng.
destruct (Hng n ε) as [? [_ Hng']]; try rewrite right_id; auto.
eapply uPred_mono; eauto using ucmra_unit_leastN.
Qed.
(** Own *)
Lemma ownM_op (a1 a2 : M) :
uPred_ownM (a1 a2) ⊣⊢ uPred_ownM a1 uPred_ownM a2.
Proof.
unseal; split=> n x ?; split.
- intros [z ?]; exists a1, (a2 z); split; [by rewrite (assoc op)|].
split.
+ by exists (core a1); rewrite cmra_core_r.
+ by exists z.
- intros (y1&y2&Hx&[z1 Hy1]&[z2 Hy2]); exists (z1 z2).
by rewrite (assoc op _ z1) -(comm op z1) (assoc op z1)
-(assoc op _ a2) (comm op z1) -Hy1 -Hy2.
Qed.
Lemma persistently_ownM_core (a : M) : uPred_ownM a uPred_ownM (core a).
Proof.
split=> n x /=; unseal; intros Hx. simpl. by apply cmra_core_monoN.
Qed.
Lemma ownM_unit P : P (uPred_ownM ε).
Proof. unseal; split=> n x ??; by exists x; rewrite left_id. Qed.
Lemma later_ownM a : uPred_ownM a b, uPred_ownM b (a b).
Proof.
unseal; split=> -[|n] x /= ? Hax; first by eauto using ucmra_unit_leastN.
destruct Hax as [y ?].
destruct (cmra_extend n x a y) as (a'&y'&Hx&?&?); auto using cmra_validN_S.
exists a'. rewrite Hx. eauto using cmra_includedN_l.
Qed.
Lemma bupd_ownM_updateP x (Φ : M Prop) :
x ~~>: Φ uPred_ownM x |==> y, Φ y uPred_ownM y.
Proof.
unseal=> Hup; split=> n x2 ? [x3 Hx] k yf ??.
destruct (Hup k (Some (x3 yf))) as (y&?&?); simpl in *.
{ rewrite /= assoc -(dist_le _ _ _ _ Hx); auto. }
exists (y x3); split; first by rewrite -assoc.
exists y; eauto using cmra_includedN_l.
Qed.
Lemma ownM_forall {A} (f : A M) :
( a, uPred_ownM (f a)) z, uPred_ownM z ( a, xf, z f a xf).
Proof.
unseal. split=> n y /= Hval Hf. exists y. split; [done|]. intros a.
destruct (Hf a) as [xf ?]; eauto.
Qed.
(** Valid *)
Lemma ownM_valid (a : M) : uPred_ownM a a.
Proof.
unseal; split=> n x Hv [a' ?]; ofe_subst; eauto using cmra_validN_op_l.
Qed.
Lemma cmra_valid_intro {A : cmra} P (a : A) : a P ( a).
Proof. unseal=> ?; split=> n x ? _ /=; by apply cmra_valid_validN. Qed.
Lemma cmra_valid_elim {A : cmra} (a : A) : a {0} a ⌝.
Proof. unseal; split=> n x ??; by apply cmra_validN_le with n, SIdx.le_0_l. Qed.
Lemma plainly_cmra_valid_1 {A : cmra} (a : A) : a a.
Proof. by unseal. Qed.
Lemma cmra_valid_weaken {A : cmra} (a b : A) : (a b) a.
Proof. unseal; split=> n x _; apply cmra_validN_op_l. Qed.
(** This is really just a special case of an entailment
between two [siProp], but we do not have the infrastructure
to express the more general case. This temporary proof rule will
be replaced by the proper one eventually. *)
Lemma valid_entails {A B : cmra} (a : A) (b : B) :
( n, {n} a {n} b) a b.
Proof. unseal=> Hval. split=>n x ?. apply Hval. Qed.
(** Consistency/soundness statement *)
(** The lemmas [pure_soundness] and [internal_eq_soundness] should become an
instance of [siProp] soundness in the future. *)
Lemma pure_soundness φ : (True φ ) φ.
Proof. unseal=> -[H]. by apply (H 0 ε); eauto using ucmra_unit_validN. Qed.
Lemma internal_eq_soundness {A : ofe} (x y : A) : (True x y) x y.
Proof.
unseal=> -[H]. apply equiv_dist=> n.
by apply (H n ε); eauto using ucmra_unit_validN.
Qed.
Lemma later_soundness P : (True P) (True P).
Proof.
unseal=> -[HP]; split=> n x Hx _.
apply uPred_mono with n ε; eauto using ucmra_unit_leastN.
by apply (HP (S n)); eauto using ucmra_unit_validN.
Qed.
End primitive.
End uPred_primitive.
From iris.bi Require Import interface derived_connectives updates.
From iris.prelude Require Import options.
Notation "P |- Q" := (P Q)
(at level 99, Q at level 200, right associativity, only parsing) : stdpp_scope.
Notation "P '|-@{' PROP } Q" := (P ⊢@{PROP} Q)
(at level 99, Q at level 200, right associativity, only parsing)
: stdpp_scope.
Notation "(|-)" := () (only parsing) : stdpp_scope.
Notation "'(|-@{' PROP } )" := (⊢@{PROP}) (only parsing) : stdpp_scope.
Notation "|- Q" := ( Q%I) (at level 20, Q at level 200, only parsing) : stdpp_scope.
Notation "'|-@{' PROP } Q" := (⊢@{PROP} Q) (at level 20, Q at level 200, only parsing) : stdpp_scope.
(** Work around parsing issues: see [notation.v] for details. *)
Notation "'(|-@{' PROP } Q )" := (⊢@{PROP} Q) (only parsing) : stdpp_scope.
Notation "P -|- Q" := (P ⊣⊢ Q) (at level 95, no associativity, only parsing) : stdpp_scope.
Notation "P '-|-@{' PROP } Q" := (P ⊣⊢@{PROP} Q)
(at level 95, no associativity, only parsing) : stdpp_scope.
Notation "(-|-)" := (⊣⊢) (only parsing) : stdpp_scope.
Notation "'(-|-@{' PROP } )" := (⊣⊢@{PROP}) (only parsing) : stdpp_scope.
Notation "P -* Q" := (P Q)%stdpp
(at level 99, Q at level 200, right associativity, only parsing)
: stdpp_scope.
(* FIXME: Notation "'⌜' φ '⌝'" := (bi_pure φ%type%stdpp) : bi_scope. *)
Notation "P /\ Q" := (P Q)%I (only parsing) : bi_scope.
Notation "(/\)" := bi_and (only parsing) : bi_scope.
Notation "P \/ Q" := (P Q)%I (only parsing) : bi_scope.
Notation "(\/)" := bi_or (only parsing) : bi_scope.
Notation "P -> Q" := (P Q)%I (only parsing) : bi_scope.
Notation "~ P" := (P False)%I (only parsing) : bi_scope.
Notation "P ** Q" := (P Q)%I
(at level 80, right associativity, only parsing) : bi_scope.
Notation "(**)" := bi_sep (only parsing) : bi_scope.
Notation "P -* Q" := (P -∗ Q)%I
(at level 99, Q at level 200, right associativity, only parsing) : bi_scope.
(* ∀ x .. y , P *)
Notation "'forall' x .. y , P" :=
(bi_forall (fun x => .. (bi_forall (fun y => P%I)) ..))
(at level 200, x binder, right associativity, only parsing) : bi_scope.
(* ∃ x .. y , P *)
Notation "'exists' x .. y , P" :=
(bi_exist (fun x => .. (bi_exist (fun y => P%I)) ..))
(at level 200, x binder, right associativity, only parsing) : bi_scope.
Notation "|> P" := ( P)%I
(at level 20, right associativity, only parsing) : bi_scope.
Notation "|>? p P" := (?p P)%I (at level 20, p at level 9, P at level 20,
only parsing) : bi_scope.
Notation "|>^ n P" := (▷^n P)%I (at level 20, n at level 9, P at level 20,
only parsing) : bi_scope.
Notation "P <-> Q" := (P Q)%I
(at level 95, no associativity, only parsing) : bi_scope.
Notation "P *-* Q" := (P ∗-∗ Q)%I
(at level 95, no associativity, only parsing) : bi_scope.
Notation "'<#>' P" := ( P)%I
(at level 20, right associativity, only parsing) : bi_scope.
Notation "'<#>?' p P" := (?p P)%I (at level 20, p at level 9, P at level 20,
right associativity, only parsing) : bi_scope.
Notation "'<except_0>' P" := ( P)%I
(at level 20, right associativity, only parsing) : bi_scope.
Notation "mP -*? Q" := (mP -∗? Q)%I
(at level 99, Q at level 200, right associativity, only parsing) : bi_scope.
Notation "P ==* Q" := (P ==∗ Q)%stdpp
(at level 99, Q at level 200, only parsing) : stdpp_scope.
Notation "P ==* Q" := (P ==∗ Q)%I
(at level 99, Q at level 200, only parsing) : bi_scope.
Notation "P ={ E1 , E2 }=* Q" := (P ={E1,E2}=∗ Q)%I
(at level 99, E1,E2 at level 50, Q at level 200, only parsing) : bi_scope.
Notation "P ={ E1 , E2 }=* Q" := (P ={E1,E2}=∗ Q)
(at level 99, E1,E2 at level 50, Q at level 200, only parsing) : stdpp_scope.
Notation "P ={ E }=* Q" := (P ={E}=∗ Q)%I
(at level 99, E at level 50, Q at level 200, only parsing) : bi_scope.
Notation "P ={ E }=* Q" := (P ={E}=∗ Q)
(at level 99, E at level 50, Q at level 200, only parsing) : stdpp_scope
.
Notation "P ={ E1 } [ E2 ]|>=* Q" := (P ={E1}[E2]▷=∗ Q)
(at level 99, E1,E2 at level 50, Q at level 200, only parsing) : stdpp_scope.
Notation "P ={ E1 } [ E2 ]|>=* Q" := (P ={E1}[E2]▷=∗ Q)%I
(at level 99, E1,E2 at level 50, Q at level 200, only parsing) : bi_scope.
Notation "|={ E }|>=> Q" := (|={E}▷=> Q)%I
(at level 99, E at level 50, Q at level 200, only parsing) : bi_scope.
Notation "P ={ E }|>=* Q" := (P ={E}▷=∗ Q)%I
(at level 99, E at level 50, Q at level 200, only parsing) : bi_scope.
Notation "|={ E1 } [ E2 ]|>=>^ n Q" := (|={E1}[E2]▷=>^n Q)%I
(at level 99, E1,E2 at level 50, n at level 9, Q at level 200, only parsing)
: bi_scope.
Notation "P ={ E1 } [ E2 ]|>=*^ n Q" := (P ={E1}[E2]▷=∗^n Q)%I
(at level 99, E1,E2 at level 50, n at level 9, Q at level 200, only parsing)
: stdpp_scope.
Notation "P ={ E1 } [ E2 ]|>=*^ n Q" := (P ={E1}[E2]▷=∗^n Q)%I
(at level 99, E1,E2 at level 50, n at level 9, Q at level 200, only parsing)
: bi_scope.
From iris.bi Require Export derived_laws derived_laws_later big_op.
From iris.bi Require Export updates internal_eq plainly embedding.
From iris.prelude Require Import options.
Module Import bi.
(** Get universes into the desired scope *)
Universe Logic.
Constraint Logic = bi.interface.universes.Logic.
Universe Quant.
Constraint Quant = bi.interface.universes.Quant.
(** Get other symbols into the desired scope *)
Export bi.interface.bi.
Export bi.derived_laws.bi.
Export bi.derived_laws_later.bi.
End bi.
From stdpp Require Import countable fin_sets functions.
From iris.algebra Require Export big_op.
From iris.algebra Require Import list gmap.
From iris.bi Require Import derived_laws_later.
From iris.prelude Require Import options.
Import interface.bi derived_laws.bi derived_laws_later.bi.
(** Notations for unary variants *)
Notation "'[∗' 'list]' k ↦ x ∈ l , P" :=
(big_opL bi_sep (λ k x, P%I) l) : bi_scope.
Notation "'[∗' 'list]' x ∈ l , P" :=
(big_opL bi_sep (λ _ x, P%I) l) : bi_scope.
Notation "'[∗]' Ps" := (big_opL bi_sep (λ _ x, x) Ps%I) : bi_scope.
Notation "'[∧' 'list]' k ↦ x ∈ l , P" :=
(big_opL bi_and (λ k x, P%I) l) : bi_scope.
Notation "'[∧' 'list]' x ∈ l , P" :=
(big_opL bi_and (λ _ x, P%I) l) : bi_scope.
Notation "'[∧]' Ps" := (big_opL bi_and (λ _ x, x) Ps%I) : bi_scope.
Notation "'[∨' 'list]' k ↦ x ∈ l , P" :=
(big_opL bi_or (λ k x, P%I) l) : bi_scope.
Notation "'[∨' 'list]' x ∈ l , P" :=
(big_opL bi_or (λ _ x, P%I) l) : bi_scope.
Notation "'[∨]' Ps" := (big_opL bi_or (λ _ x, x) Ps%I) : bi_scope.
Notation "'[∗' 'map]' k ↦ x ∈ m , P" := (big_opM bi_sep (λ k x, P%I) m) : bi_scope.
Notation "'[∗' 'map]' x ∈ m , P" := (big_opM bi_sep (λ _ x, P%I) m) : bi_scope.
Notation "'[∧' 'map]' k ↦ x ∈ m , P" := (big_opM bi_and (λ k x, P) m) : bi_scope.
Notation "'[∧' 'map]' x ∈ m , P" := (big_opM bi_and (λ _ x, P) m) : bi_scope.
Notation "'[∗' 'set]' x ∈ X , P" := (big_opS bi_sep (λ x, P%I) X) : bi_scope.
Notation "'[∗' 'mset]' x ∈ X , P" := (big_opMS bi_sep (λ x, P%I) X) : bi_scope.
(** Definitions and notations for binary variants *)
(** A version of the separating big operator that ranges over two lists. This
version also ensures that both lists have the same length. Although this version
can be defined in terms of the unary using a [zip] (see [big_sepL2_alt]), we do
not define it that way to get better computational behavior (for [simpl]). *)
Fixpoint big_sepL2 {PROP : bi} {A B}
(Φ : nat A B PROP) (l1 : list A) (l2 : list B) : PROP :=
match l1, l2 with
| [], [] => emp
| x1 :: l1, x2 :: l2 => Φ 0 x1 x2 big_sepL2 (λ n, Φ (S n)) l1 l2
| _, _ => False
end%I.
Global Instance: Params (@big_sepL2) 3 := {}.
Global Arguments big_sepL2 {PROP A B} _ !_ !_ /.
Global Typeclasses Opaque big_sepL2.
Notation "'[∗' 'list]' k ↦ x1 ; x2 ∈ l1 ; l2 , P" :=
(big_sepL2 (λ k x1 x2, P%I) l1 l2) : bi_scope.
Notation "'[∗' 'list]' x1 ; x2 ∈ l1 ; l2 , P" :=
(big_sepL2 (λ _ x1 x2, P%I) l1 l2) : bi_scope.
Local Definition big_sepM2_def {PROP : bi} `{Countable K} {A B}
(Φ : K A B PROP) (m1 : gmap K A) (m2 : gmap K B) : PROP :=
dom m1 = dom m2 [ map] k xy map_zip m1 m2, Φ k xy.1 xy.2.
Local Definition big_sepM2_aux : seal (@big_sepM2_def). Proof. by eexists. Qed.
Definition big_sepM2 := big_sepM2_aux.(unseal).
Global Arguments big_sepM2 {PROP K _ _ A B} _ _ _.
Local Definition big_sepM2_unseal : @big_sepM2 = _ := big_sepM2_aux.(seal_eq).
Global Instance: Params (@big_sepM2) 6 := {}.
Notation "'[∗' 'map]' k ↦ x1 ; x2 ∈ m1 ; m2 , P" :=
(big_sepM2 (λ k x1 x2, P%I) m1 m2) : bi_scope.
Notation "'[∗' 'map]' x1 ; x2 ∈ m1 ; m2 , P" :=
(big_sepM2 (λ _ x1 x2, P%I) m1 m2) : bi_scope.
(** * Properties *)
Section big_op.
Context {PROP : bi}.
Implicit Types P Q : PROP.
Implicit Types Ps Qs : list PROP.
Implicit Types A : Type.
(** ** Big ops over lists *)
Section sep_list.
Context {A : Type}.
Implicit Types l : list A.
Implicit Types Φ Ψ : nat A PROP.
Lemma big_sepL_nil Φ : ([ list] ky nil, Φ k y) ⊣⊢ emp.
Proof. done. Qed.
Lemma big_sepL_nil' P `{!Affine P} Φ : P [ list] ky nil, Φ k y.
Proof. apply: affine. Qed.
Lemma big_sepL_cons Φ x l :
([ list] ky x :: l, Φ k y) ⊣⊢ Φ 0 x [ list] ky l, Φ (S k) y.
Proof. by rewrite big_opL_cons. Qed.
Lemma big_sepL_singleton Φ x : ([ list] ky [x], Φ k y) ⊣⊢ Φ 0 x.
Proof. by rewrite big_opL_singleton. Qed.
Lemma big_sepL_app Φ l1 l2 :
([ list] ky l1 ++ l2, Φ k y)
⊣⊢ ([ list] ky l1, Φ k y) ([ list] ky l2, Φ (length l1 + k) y).
Proof. by rewrite big_opL_app. Qed.
Lemma big_sepL_snoc Φ l x :
([ list] ky l ++ [x], Φ k y) ⊣⊢ ([ list] ky l, Φ k y) Φ (length l) x.
Proof. by rewrite big_opL_snoc. Qed.
Lemma big_sepL_take_drop Φ l n :
([ list] k x l, Φ k x) ⊣⊢
([ list] k x take n l, Φ k x) ([ list] k x drop n l, Φ (n + k) x).
Proof. by rewrite big_opL_take_drop. Qed.
Lemma big_sepL_submseteq (Φ : A PROP) `{!∀ x, Affine (Φ x)} l1 l2 :
l1 ⊆+ l2 ([ list] y l2, Φ y) [ list] y l1, Φ y.
Proof.
intros [l ->]%submseteq_Permutation. rewrite big_sepL_app.
induction l as [|x l IH]; simpl; [by rewrite right_id|by rewrite sep_elim_r].
Qed.
(** The lemmas [big_sepL_mono], [big_sepL_ne] and [big_sepL_proper] are more
generic than the instances as they also give [l !! k = Some y] in the premise. *)
Lemma big_sepL_mono Φ Ψ l :
( k y, l !! k = Some y Φ k y Ψ k y)
([ list] k y l, Φ k y) [ list] k y l, Ψ k y.
Proof. apply big_opL_gen_proper; apply _. Qed.
Lemma big_sepL_ne Φ Ψ l n :
( k y, l !! k = Some y Φ k y {n} Ψ k y)
([ list] k y l, Φ k y)%I {n} ([ list] k y l, Ψ k y)%I.
Proof. apply big_opL_ne. Qed.
Lemma big_sepL_proper Φ Ψ l :
( k y, l !! k = Some y Φ k y ⊣⊢ Ψ k y)
([ list] k y l, Φ k y) ⊣⊢ ([ list] k y l, Ψ k y).
Proof. apply big_opL_proper. Qed.
(** No need to declare instances for non-expansiveness and properness, we
get both from the generic [big_opL] instances. *)
Global Instance big_sepL_mono' :
Proper (pointwise_relation _ (pointwise_relation _ ()) ==> (=) ==> ())
(big_opL (@bi_sep PROP) (A:=A)).
Proof. intros f g Hf m ? <-. apply big_sepL_mono; intros; apply Hf. Qed.
Global Instance big_sepL_flip_mono' :
Proper (pointwise_relation _ (pointwise_relation _ (flip ())) ==> (=) ==> flip ())
(big_opL (@bi_sep PROP) (A:=A)).
Proof. solve_proper. Qed.
Global Instance big_sepL_id_mono' :
Proper (Forall2 () ==> ()) (big_opL (@bi_sep PROP) (λ _ P, P)).
Proof. by induction 1 as [|P Q Ps Qs HPQ ? IH]; rewrite /= ?HPQ ?IH. Qed.
Global Instance big_sepL_nil_persistent Φ :
Persistent ([ list] kx [], Φ k x).
Proof. simpl; apply _. Qed.
Lemma big_sepL_persistent Φ l :
( k x, l !! k = Some x Persistent (Φ k x))
Persistent ([ list] kx l, Φ k x).
Proof. apply big_opL_closed; apply _. Qed.
Global Instance big_sepL_persistent' Φ l :
( k x, Persistent (Φ k x)) Persistent ([ list] kx l, Φ k x).
Proof. intros; apply big_sepL_persistent, _. Qed.
Global Instance big_sepL_persistent_id Ps :
TCForall Persistent Ps Persistent ([] Ps).
Proof. induction 1; simpl; apply _. Qed.
Global Instance big_sepL_nil_affine Φ :
Affine ([ list] kx [], Φ k x).
Proof. simpl; apply _. Qed.
Lemma big_sepL_affine Φ l :
( k x, l !! k = Some x Affine (Φ k x))
Affine ([ list] kx l, Φ k x).
Proof. apply big_opL_closed; apply _. Qed.
Global Instance big_sepL_affine' Φ l :
( k x, Affine (Φ k x)) Affine ([ list] kx l, Φ k x).
Proof. intros. apply big_sepL_affine, _. Qed.
Global Instance big_sepL_affine_id Ps : TCForall Affine Ps Affine ([] Ps).
Proof. induction 1; simpl; apply _. Qed.
Global Instance big_sepL_nil_timeless `{!Timeless (emp%I : PROP)} Φ :
Timeless ([ list] kx [], Φ k x).
Proof. simpl; apply _. Qed.
Lemma big_sepL_timeless `{!Timeless (emp%I : PROP)} Φ l :
( k x, l !! k = Some x Timeless (Φ k x))
Timeless ([ list] kx l, Φ k x).
Proof. apply big_opL_closed; apply _. Qed.
Global Instance big_sepL_timeless' `{!Timeless (emp%I : PROP)} Φ l :
( k x, Timeless (Φ k x))
Timeless ([ list] kx l, Φ k x).
Proof. intros. apply big_sepL_timeless, _. Qed.
Global Instance big_sepL_timeless_id `{!Timeless (emp%I : PROP)} Ps :
TCForall Timeless Ps Timeless ([] Ps).
Proof. induction 1; simpl; apply _. Qed.
Lemma big_sepL_emp l : ([ list] ky l, emp) ⊣⊢@{PROP} emp.
Proof. by rewrite big_opL_unit. Qed.
Lemma big_sepL_insert_acc Φ l i x :
l !! i = Some x
([ list] ky l, Φ k y) Φ i x ( y, Φ i y -∗ ([ list] ky <[i:=y]>l, Φ k y)).
Proof.
intros Hli. assert (i length l) by eauto using lookup_lt_Some, Nat.lt_le_incl.
rewrite -(take_drop_middle l i x) // big_sepL_app /=.
rewrite Nat.add_0_r length_take_le //.
rewrite assoc -!(comm _ (Φ _ _)) -assoc. apply sep_mono_r, forall_intro=> y.
rewrite insert_app_r_alt ?length_take_le //.
rewrite Nat.sub_diag /=. apply wand_intro_l.
rewrite assoc !(comm _ (Φ _ _)) -assoc big_sepL_app /=.
by rewrite Nat.add_0_r length_take_le.
Qed.
Lemma big_sepL_lookup_acc Φ l i x :
l !! i = Some x
([ list] ky l, Φ k y) Φ i x (Φ i x -∗ ([ list] ky l, Φ k y)).
Proof. intros. by rewrite {1}big_sepL_insert_acc // (forall_elim x) list_insert_id. Qed.
Lemma big_sepL_lookup Φ l i x
`{!TCOr ( j y, Affine (Φ j y)) (Absorbing (Φ i x))} :
l !! i = Some x ([ list] ky l, Φ k y) Φ i x.
Proof.
intros Hi. destruct select (TCOr _ _).
- rewrite -(take_drop_middle l i x) // big_sepL_app /= length_take.
apply lookup_lt_Some in Hi. rewrite (_ : _ + 0 = i); last lia.
rewrite sep_elim_r sep_elim_l //.
- rewrite big_sepL_lookup_acc // sep_elim_l //.
Qed.
Lemma big_sepL_elem_of (Φ : A PROP) l x
`{!TCOr ( y, Affine (Φ y)) (Absorbing (Φ x))} :
x l ([ list] y l, Φ y) Φ x.
Proof.
intros [i ?]%elem_of_list_lookup.
destruct select (TCOr _ _); by apply: big_sepL_lookup.
Qed.
Lemma big_sepL_fmap {B} (f : A B) (Φ : nat B PROP) l :
([ list] ky f <$> l, Φ k y) ⊣⊢ ([ list] ky l, Φ k (f y)).
Proof. by rewrite big_opL_fmap. Qed.
Lemma big_sepL_omap {B} (f : A option B) (Φ : B PROP) l :
([ list] y omap f l, Φ y) ⊣⊢ ([ list] y l, from_option Φ emp (f y)).
Proof. by rewrite big_opL_omap. Qed.
Lemma big_sepL_bind {B} (f : A list B) (Φ : B PROP) l :
([ list] y l ≫= f, Φ y) ⊣⊢ ([ list] x l, [ list] y f x, Φ y).
Proof. by rewrite big_opL_bind. Qed.
Lemma big_sepL_sep Φ Ψ l :
([ list] kx l, Φ k x Ψ k x)
⊣⊢ ([ list] kx l, Φ k x) ([ list] kx l, Ψ k x).
Proof. by rewrite big_opL_op. Qed.
Lemma big_sepL_sep_2 Φ Ψ l :
([ list] kx l, Φ k x) -∗
([ list] kx l, Ψ k x) -∗
([ list] kx l, Φ k x Ψ k x).
Proof. apply entails_wand, wand_intro_r. rewrite big_sepL_sep //. Qed.
Lemma big_sepL_and Φ Ψ l :
([ list] kx l, Φ k x Ψ k x)
([ list] kx l, Φ k x) ([ list] kx l, Ψ k x).
Proof. auto using and_intro, big_sepL_mono, and_elim_l, and_elim_r. Qed.
Lemma big_sepL_pure_1 (φ : nat A Prop) l :
([ list] kx l, φ k x) ⊢@{PROP} ⌜∀ k x, l !! k = Some x φ k x⌝.
Proof.
induction l as [|x l IH] using rev_ind.
{ apply pure_intro=>??. rewrite lookup_nil. done. }
rewrite big_sepL_snoc // IH sep_and -pure_and.
f_equiv=>-[Hl Hx] k y /lookup_app_Some =>-[Hy|[Hlen Hy]].
- by apply Hl.
- apply list_lookup_singleton_Some in Hy as [Hk ->].
replace k with (length l) by lia. done.
Qed.
Lemma big_sepL_affinely_pure_2 (φ : nat A Prop) l :
<affine> ⌜∀ k x, l !! k = Some x φ k x ⊢@{PROP} ([ list] kx l, <affine> φ k x).
Proof.
induction l as [|x l IH] using rev_ind.
{ rewrite big_sepL_nil. apply affinely_elim_emp. }
rewrite big_sepL_snoc // -IH.
rewrite -persistent_and_sep_1 -affinely_and -pure_and.
f_equiv. f_equiv=>- Hlx. split.
- intros k y Hy. apply Hlx. rewrite lookup_app Hy //.
- apply Hlx. rewrite lookup_app lookup_ge_None_2 //.
rewrite Nat.sub_diag //.
Qed.
(** The general backwards direction requires [BiAffine] to cover the empty case. *)
Lemma big_sepL_pure `{!BiAffine PROP} (φ : nat A Prop) l :
([ list] kx l, φ k x) ⊣⊢@{PROP} ⌜∀ k x, l !! k = Some x φ k x⌝.
Proof.
apply (anti_symm ()); first by apply big_sepL_pure_1.
rewrite -(affine_affinely ⌜_⌝).
rewrite big_sepL_affinely_pure_2. by setoid_rewrite affinely_elim.
Qed.
Lemma big_sepL_persistently `{!BiAffine PROP} Φ l :
<pers> ([ list] kx l, Φ k x) ⊣⊢ [ list] kx l, <pers> (Φ k x).
Proof. apply (big_opL_commute _). Qed.
Lemma big_sepL_intro Φ l :
( k x, l !! k = Some x Φ k x) [ list] kx l, Φ k x.
Proof.
revert Φ. induction l as [|x l IH]=> Φ /=; [by apply (affine _)|].
rewrite intuitionistically_sep_dup. f_equiv.
- rewrite (forall_elim 0) (forall_elim x) pure_True // True_impl.
by rewrite intuitionistically_elim.
- rewrite -IH. f_equiv.
apply forall_intro=> k; by rewrite (forall_elim (S k)).
Qed.
Lemma big_sepL_forall `{!BiAffine PROP} Φ l :
( k x, Persistent (Φ k x))
([ list] kx l, Φ k x) ⊣⊢ ( k x, l !! k = Some x Φ k x).
Proof.
intros . apply (anti_symm _).
{ apply forall_intro=> k; apply forall_intro=> x.
apply impl_intro_l, pure_elim_l=> ?; by apply: big_sepL_lookup. }
revert Φ . induction l as [|x l IH]=> Φ /=.
{ apply: affine. }
rewrite -persistent_and_sep_1. apply and_intro.
- rewrite (forall_elim 0) (forall_elim x) pure_True // True_impl. done.
- rewrite -IH. apply forall_intro => k. by rewrite (forall_elim (S k)).
Qed.
Lemma big_sepL_impl Φ Ψ l :
([ list] kx l, Φ k x) -∗
( k x, l !! k = Some x Φ k x -∗ Ψ k x) -∗
[ list] kx l, Ψ k x.
Proof.
apply entails_wand, wand_intro_l. rewrite big_sepL_intro -big_sepL_sep.
by setoid_rewrite wand_elim_l.
Qed.
Lemma big_sepL_wand Φ Ψ l :
([ list] kx l, Φ k x) -∗
([ list] kx l, Φ k x -∗ Ψ k x) -∗
[ list] kx l, Ψ k x.
Proof.
apply entails_wand, wand_intro_r. rewrite -big_sepL_sep.
setoid_rewrite wand_elim_r. done.
Qed.
Lemma big_sepL_dup P `{!Affine P} l :
(P -∗ P P) -∗ P -∗ [ list] kx l, P.
Proof.
apply entails_wand, wand_intro_l.
induction l as [|x l IH]=> /=; first by apply: affine.
rewrite intuitionistically_sep_dup {1}intuitionistically_elim.
rewrite assoc wand_elim_r -assoc. apply sep_mono; done.
Qed.
Lemma big_sepL_delete Φ l i x :
l !! i = Some x
([ list] ky l, Φ k y) ⊣⊢
Φ i x [ list] ky l, if decide (k = i) then emp else Φ k y.
Proof.
intros. rewrite -(take_drop_middle l i x) // !big_sepL_app /= Nat.add_0_r.
rewrite length_take_le; last eauto using lookup_lt_Some, Nat.lt_le_incl.
rewrite decide_True // left_id.
rewrite assoc -!(comm _ (Φ _ _)) -assoc. do 2 f_equiv.
- apply big_sepL_proper=> k y Hk. apply lookup_lt_Some in Hk.
rewrite length_take in Hk. by rewrite decide_False; last lia.
- apply big_sepL_proper=> k y _. by rewrite decide_False; last lia.
Qed.
Lemma big_sepL_delete' `{!BiAffine PROP} Φ l i x :
l !! i = Some x
([ list] ky l, Φ k y) ⊣⊢ Φ i x [ list] ky l, k i Φ k y.
Proof.
intros. rewrite big_sepL_delete //. (do 2 f_equiv)=> k y.
rewrite -decide_emp. by repeat case_decide.
Qed.
Lemma big_sepL_lookup_acc_impl {Φ l} i x :
l !! i = Some x
([ list] ky l, Φ k y) -∗
(* We obtain [Φ] for [x] *)
Φ i x
(* We reobtain the bigop for a predicate [Ψ] selected by the user *)
Ψ,
( k y, l !! k = Some y k i Φ k y -∗ Ψ k y) -∗
Ψ i x -∗
[ list] ky l, Ψ k y.
Proof.
intros. apply entails_wand.
rewrite big_sepL_delete //. apply sep_mono_r, forall_intro=> Ψ.
apply wand_intro_r, wand_intro_l.
rewrite (big_sepL_delete Ψ l i x) //. apply sep_mono_r.
eapply wand_apply; [apply wand_entails, big_sepL_impl|apply sep_mono_r].
apply intuitionistically_intro', forall_intro=> k; apply forall_intro=> y.
apply impl_intro_l, pure_elim_l=> ?; apply wand_intro_r.
rewrite (forall_elim ) (forall_elim y) pure_True // left_id.
destruct (decide _) as [->|]; [by apply: affine|].
by rewrite pure_True //left_id intuitionistically_elim wand_elim_l.
Qed.
Lemma big_sepL_replicate l P :
[] replicate (length l) P ⊣⊢ [ list] y l, P.
Proof. induction l as [|x l]=> //=; by f_equiv. Qed.
Lemma big_sepL_later `{!BiAffine PROP} Φ l :
([ list] kx l, Φ k x) ⊣⊢ ([ list] kx l, Φ k x).
Proof. apply (big_opL_commute _). Qed.
Lemma big_sepL_later_2 Φ l :
([ list] kx l, Φ k x) [ list] kx l, Φ k x.
Proof. by rewrite (big_opL_commute _). Qed.
Lemma big_sepL_laterN `{!BiAffine PROP} Φ n l :
▷^n ([ list] kx l, Φ k x) ⊣⊢ ([ list] kx l, ▷^n Φ k x).
Proof. apply (big_opL_commute _). Qed.
Lemma big_sepL_laterN_2 Φ n l :
([ list] kx l, ▷^n Φ k x) ▷^n [ list] kx l, Φ k x.
Proof. by rewrite (big_opL_commute _). Qed.
End sep_list.
(* Some lemmas depend on the generalized versions of the above ones. *)
Lemma big_sepL_sep_zip_with {A B C} (f : A B C) (g1 : C A) (g2 : C B)
(Φ1 : nat A PROP) (Φ2 : nat B PROP) l1 l2 :
( x y, g1 (f x y) = x)
( x y, g2 (f x y) = y)
length l1 = length l2
([ list] kxy zip_with f l1 l2, Φ1 k (g1 xy) Φ2 k (g2 xy)) ⊣⊢
([ list] kx l1, Φ1 k x) ([ list] ky l2, Φ2 k y).
Proof. apply big_opL_sep_zip_with. Qed.
Lemma big_sepL_sep_zip {A B} (Φ1 : nat A PROP) (Φ2 : nat B PROP) l1 l2 :
length l1 = length l2
([ list] kxy zip l1 l2, Φ1 k xy.1 Φ2 k xy.2) ⊣⊢
([ list] kx l1, Φ1 k x) ([ list] ky l2, Φ2 k y).
Proof. apply big_opL_sep_zip. Qed.
Lemma big_sepL_zip_with {A B C} (Φ : nat A PROP) f (l1 : list B) (l2 : list C) :
([ list] kx zip_with f l1 l2, Φ k x) ⊣⊢
([ list] kx l1, if l2 !! k is Some y then Φ k (f x y) else emp).
Proof.
revert Φ l2; induction l1 as [|x l1 IH]=> Φ [|y l2] //=.
- by rewrite big_sepL_emp left_id.
- by rewrite IH.
Qed.
(** ** Big ops over two lists *)
Lemma big_sepL2_alt {A B} (Φ : nat A B PROP) l1 l2 :
([ list] ky1;y2 l1; l2, Φ k y1 y2) ⊣⊢
length l1 = length l2 [ list] k xy zip l1 l2, Φ k (xy.1) (xy.2).
Proof.
apply (anti_symm _).
- apply and_intro.
+ revert Φ l2. induction l1 as [|x1 l1 IH]=> Φ -[|x2 l2] /=;
auto using pure_intro, False_elim.
rewrite IH sep_elim_r. apply pure_mono; auto.
+ revert Φ l2. induction l1 as [|x1 l1 IH]=> Φ -[|x2 l2] /=;
auto using pure_intro, False_elim.
by rewrite IH.
- apply pure_elim_l=> /Forall2_same_length Hl. revert Φ.
induction Hl as [|x1 l1 x2 l2 _ _ IH]=> Φ //=. by rewrite -IH.
Qed.
Section sep_list2.
Context {A B : Type}.
Implicit Types Φ Ψ : nat A B PROP.
Lemma big_sepL2_nil Φ : ([ list] ky1;y2 []; [], Φ k y1 y2) ⊣⊢ emp.
Proof. done. Qed.
Lemma big_sepL2_nil' P `{!Affine P} Φ : P [ list] ky1;y2 [];[], Φ k y1 y2.
Proof. apply: affine. Qed.
Lemma big_sepL2_nil_inv_l Φ l2 :
([ list] ky1;y2 []; l2, Φ k y1 y2) l2 = []⌝.
Proof. destruct l2; simpl; auto using False_elim, pure_intro. Qed.
Lemma big_sepL2_nil_inv_r Φ l1 :
([ list] ky1;y2 l1; [], Φ k y1 y2) l1 = []⌝.
Proof. destruct l1; simpl; auto using False_elim, pure_intro. Qed.
Lemma big_sepL2_cons Φ x1 x2 l1 l2 :
([ list] ky1;y2 x1 :: l1; x2 :: l2, Φ k y1 y2)
⊣⊢ Φ 0 x1 x2 [ list] ky1;y2 l1;l2, Φ (S k) y1 y2.
Proof. done. Qed.
Lemma big_sepL2_cons_inv_l Φ x1 l1 l2 :
([ list] ky1;y2 x1 :: l1; l2, Φ k y1 y2)
x2 l2', l2 = x2 :: l2'
Φ 0 x1 x2 [ list] ky1;y2 l1;l2', Φ (S k) y1 y2.
Proof.
destruct l2 as [|x2 l2]; simpl; auto using False_elim.
by rewrite -(exist_intro x2) -(exist_intro l2) pure_True // left_id.
Qed.
Lemma big_sepL2_cons_inv_r Φ x2 l1 l2 :
([ list] ky1;y2 l1; x2 :: l2, Φ k y1 y2)
x1 l1', l1 = x1 :: l1'
Φ 0 x1 x2 [ list] ky1;y2 l1';l2, Φ (S k) y1 y2.
Proof.
destruct l1 as [|x1 l1]; simpl; auto using False_elim.
by rewrite -(exist_intro x1) -(exist_intro l1) pure_True // left_id.
Qed.
Lemma big_sepL2_singleton Φ x1 x2 :
([ list] ky1;y2 [x1];[x2], Φ k y1 y2) ⊣⊢ Φ 0 x1 x2.
Proof. by rewrite /= right_id. Qed.
Lemma big_sepL2_length Φ l1 l2 :
([ list] ky1;y2 l1; l2, Φ k y1 y2) length l1 = length l2 ⌝.
Proof. by rewrite big_sepL2_alt and_elim_l. Qed.
Lemma big_sepL2_fst_snd Φ l :
([ list] ky1;y2 l.*1; l.*2, Φ k y1 y2) ⊣⊢
[ list] k xy l, Φ k (xy.1) (xy.2).
Proof.
rewrite big_sepL2_alt !length_fmap.
by rewrite pure_True // True_and zip_fst_snd.
Qed.
Lemma big_sepL2_app Φ l1 l2 l1' l2' :
([ list] ky1;y2 l1; l1', Φ k y1 y2)
([ list] ky1;y2 l2; l2', Φ (length l1 + k) y1 y2) -∗
([ list] ky1;y2 l1 ++ l2; l1' ++ l2', Φ k y1 y2).
Proof.
apply wand_intro_r. revert Φ l1'. induction l1 as [|x1 l1 IH]=> Φ -[|x1' l1'] /=.
- by rewrite left_id.
- rewrite left_absorb. apply False_elim.
- rewrite left_absorb. apply False_elim.
- by rewrite -assoc IH.
Qed.
Lemma big_sepL2_app_inv_l Φ l1' l1'' l2 :
([ list] ky1;y2 l1' ++ l1''; l2, Φ k y1 y2)
l2' l2'', l2 = l2' ++ l2''
([ list] ky1;y2 l1';l2', Φ k y1 y2)
([ list] ky1;y2 l1'';l2'', Φ (length l1' + k) y1 y2).
Proof.
rewrite -(exist_intro (take (length l1') l2))
-(exist_intro (drop (length l1') l2)) take_drop pure_True // left_id.
revert Φ l2. induction l1' as [|x1 l1' IH]=> Φ -[|x2 l2] /=;
[by rewrite left_id|by rewrite left_id|apply False_elim|].
by rewrite IH -assoc.
Qed.
Lemma big_sepL2_app_inv_r Φ l1 l2' l2'' :
([ list] ky1;y2 l1; l2' ++ l2'', Φ k y1 y2)
l1' l1'', l1 = l1' ++ l1''
([ list] ky1;y2 l1';l2', Φ k y1 y2)
([ list] ky1;y2 l1'';l2'', Φ (length l2' + k) y1 y2).
Proof.
rewrite -(exist_intro (take (length l2') l1))
-(exist_intro (drop (length l2') l1)) take_drop pure_True // left_id.
revert Φ l1. induction l2' as [|x2 l2' IH]=> Φ -[|x1 l1] /=;
[by rewrite left_id|by rewrite left_id|apply False_elim|].
by rewrite IH -assoc.
Qed.
Lemma big_sepL2_app_inv Φ l1 l2 l1' l2' :
length l1 = length l1' length l2 = length l2'
([ list] ky1;y2 l1 ++ l2; l1' ++ l2', Φ k y1 y2)
([ list] ky1;y2 l1; l1', Φ k y1 y2)
([ list] ky1;y2 l2; l2', Φ (length l1 + k)%nat y1 y2).
Proof.
revert Φ l1'. induction l1 as [|x1 l1 IH]=> Φ -[|x1' l1'] /= Hlen.
- by rewrite left_id.
- destruct Hlen as [[=]|Hlen]. rewrite big_sepL2_length Hlen /= length_app.
apply pure_elim'; lia.
- destruct Hlen as [[=]|Hlen]. rewrite big_sepL2_length -Hlen /= length_app.
apply pure_elim'; lia.
- by rewrite -assoc IH; last lia.
Qed.
Lemma big_sepL2_app_same_length Φ l1 l2 l1' l2' :
length l1 = length l1' length l2 = length l2'
([ list] ky1;y2 l1 ++ l2; l1' ++ l2', Φ k y1 y2) ⊣⊢
([ list] ky1;y2 l1; l1', Φ k y1 y2)
([ list] ky1;y2 l2; l2', Φ (length l1 + k)%nat y1 y2).
Proof.
intros. apply (anti_symm _).
- by apply big_sepL2_app_inv.
- apply wand_elim_l', big_sepL2_app.
Qed.
Lemma big_sepL2_snoc Φ x1 x2 l1 l2 :
([ list] ky1;y2 l1 ++ [x1]; l2 ++ [x2], Φ k y1 y2) ⊣⊢
([ list] ky1;y2 l1; l2, Φ k y1 y2) Φ (length l1) x1 x2.
Proof.
rewrite big_sepL2_app_same_length; last by auto.
by rewrite big_sepL2_singleton Nat.add_0_r.
Qed.
(** The lemmas [big_sepL2_mono], [big_sepL2_ne] and [big_sepL2_proper] are more
generic than the instances as they also give [li !! k = Some yi] in the premise. *)
Lemma big_sepL2_mono Φ Ψ l1 l2 :
( k y1 y2, l1 !! k = Some y1 l2 !! k = Some y2 Φ k y1 y2 Ψ k y1 y2)
([ list] k y1;y2 l1;l2, Φ k y1 y2) [ list] k y1;y2 l1;l2, Ψ k y1 y2.
Proof.
intros H. rewrite !big_sepL2_alt. f_equiv. apply big_sepL_mono=> k [y1 y2].
rewrite lookup_zip_with=> ?; simplify_option_eq; auto.
Qed.
Lemma big_sepL2_ne Φ Ψ l1 l2 n :
( k y1 y2, l1 !! k = Some y1 l2 !! k = Some y2 Φ k y1 y2 {n} Ψ k y1 y2)
([ list] k y1;y2 l1;l2, Φ k y1 y2)%I {n} ([ list] k y1;y2 l1;l2, Ψ k y1 y2)%I.
Proof.
intros H. rewrite !big_sepL2_alt. f_equiv. apply big_sepL_ne=> k [y1 y2].
rewrite lookup_zip_with=> ?; simplify_option_eq; auto.
Qed.
Lemma big_sepL2_proper Φ Ψ l1 l2 :
( k y1 y2, l1 !! k = Some y1 l2 !! k = Some y2 Φ k y1 y2 ⊣⊢ Ψ k y1 y2)
([ list] k y1;y2 l1;l2, Φ k y1 y2) ⊣⊢ [ list] k y1;y2 l1;l2, Ψ k y1 y2.
Proof.
intros; apply (anti_symm _);
apply big_sepL2_mono; auto using equiv_entails_1_1, equiv_entails_1_2.
Qed.
Lemma big_sepL2_proper_2 `{!Equiv A, !Equiv B} Φ Ψ l1 l2 l1' l2' :
l1 l1' l2 l2'
( k y1 y1' y2 y2',
l1 !! k = Some y1 l1' !! k = Some y1' y1 y1'
l2 !! k = Some y2 l2' !! k = Some y2' y2 y2'
Φ k y1 y2 ⊣⊢ Ψ k y1' y2')
([ list] k y1;y2 l1;l2, Φ k y1 y2) ⊣⊢ [ list] k y1;y2 l1';l2', Ψ k y1 y2.
Proof.
intros Hl1 Hl2 Hf. rewrite !big_sepL2_alt. f_equiv.
{ do 2 f_equiv; by apply length_proper. }
apply big_opL_proper_2; [by f_equiv|].
intros k [x1 y1] [x2 y2] (?&?&[=<- <-]&?&?)%lookup_zip_with_Some
(?&?&[=<- <-]&?&?)%lookup_zip_with_Some [??]; naive_solver.
Qed.
Global Instance big_sepL2_ne' n :
Proper (pointwise_relation _ (pointwise_relation _ (pointwise_relation _ (dist n)))
==> (=) ==> (=) ==> (dist n))
(big_sepL2 (PROP:=PROP) (A:=A) (B:=B)).
Proof. intros f g Hf l1 ? <- l2 ? <-. apply big_sepL2_ne; intros; apply Hf. Qed.
Global Instance big_sepL2_mono' :
Proper (pointwise_relation _ (pointwise_relation _ (pointwise_relation _ ()))
==> (=) ==> (=) ==> ())
(big_sepL2 (PROP:=PROP) (A:=A) (B:=B)).
Proof. intros f g Hf l1 ? <- l2 ? <-. apply big_sepL2_mono; intros; apply Hf. Qed.
Global Instance big_sepL2_flip_mono' :
Proper (pointwise_relation _ (pointwise_relation _ (pointwise_relation _ (flip ())))
==> (=) ==> (=) ==> flip ())
(big_sepL2 (PROP:=PROP) (A:=A) (B:=B)).
Proof. solve_proper. Qed.
Global Instance big_sepL2_proper' :
Proper (pointwise_relation _ (pointwise_relation _ (pointwise_relation _ (⊣⊢)))
==> (=) ==> (=) ==> (⊣⊢))
(big_sepL2 (PROP:=PROP) (A:=A) (B:=B)).
Proof. intros f g Hf l1 ? <- l2 ? <-. apply big_sepL2_proper; intros; apply Hf. Qed.
(** Shows that some property [P] is closed under [big_sepL2]. Examples of [P]
are [Persistent], [Affine], [Timeless]. *)
Lemma big_sepL2_closed (P : PROP Prop) Φ l1 l2 :
P emp%I P False%I
( Q1 Q2, P Q1 P Q2 P (Q1 Q2)%I)
( k x1 x2, l1 !! k = Some x1 l2 !! k = Some x2 P (Φ k x1 x2))
P ([ list] kx1;x2 l1; l2, Φ k x1 x2)%I.
Proof.
intros ?? Hsep. revert l2 Φ. induction l1 as [|x1 l1 IH]=> -[|x2 l2] Φ //=.
apply Hsep; first by auto. apply IH=> k. apply ( (S k)).
Qed.
Global Instance big_sepL2_nil_persistent Φ :
Persistent ([ list] ky1;y2 []; [], Φ k y1 y2).
Proof. simpl; apply _. Qed.
Lemma big_sepL2_persistent Φ l1 l2 :
( k x1 x2, l1 !! k = Some x1 l2 !! k = Some x2 Persistent (Φ k x1 x2))
Persistent ([ list] ky1;y2 l1;l2, Φ k y1 y2).
Proof. apply big_sepL2_closed; apply _. Qed.
Global Instance big_sepL2_persistent' Φ l1 l2 :
( k x1 x2, Persistent (Φ k x1 x2))
Persistent ([ list] ky1;y2 l1;l2, Φ k y1 y2).
Proof. intros; apply big_sepL2_persistent, _. Qed.
Global Instance big_sepL2_nil_affine Φ :
Affine ([ list] ky1;y2 []; [], Φ k y1 y2).
Proof. simpl; apply _. Qed.
Lemma big_sepL2_affine Φ l1 l2 :
( k x1 x2, l1 !! k = Some x1 l2 !! k = Some x2 Affine (Φ k x1 x2))
Affine ([ list] ky1;y2 l1;l2, Φ k y1 y2).
Proof. apply big_sepL2_closed; apply _. Qed.
Global Instance big_sepL2_affine' Φ l1 l2 :
( k x1 x2, Affine (Φ k x1 x2))
Affine ([ list] ky1;y2 l1;l2, Φ k y1 y2).
Proof. intros; apply big_sepL2_affine, _. Qed.
Global Instance big_sepL2_nil_timeless `{!Timeless (emp%I : PROP)} Φ :
Timeless ([ list] ky1;y2 []; [], Φ k y1 y2).
Proof. simpl; apply _. Qed.
Lemma big_sepL2_timeless `{!Timeless (emp%I : PROP)} Φ l1 l2 :
( k x1 x2, l1 !! k = Some x1 l2 !! k = Some x2 Timeless (Φ k x1 x2))
Timeless ([ list] ky1;y2 l1;l2, Φ k y1 y2).
Proof. apply big_sepL2_closed; apply _. Qed.
Global Instance big_sepL2_timeless' `{!Timeless (emp%I : PROP)} Φ l1 l2 :
( k x1 x2, Timeless (Φ k x1 x2))
Timeless ([ list] ky1;y2 l1;l2, Φ k y1 y2).
Proof. intros; apply big_sepL2_timeless, _. Qed.
Lemma big_sepL2_insert_acc Φ l1 l2 i x1 x2 :
l1 !! i = Some x1 l2 !! i = Some x2
([ list] ky1;y2 l1;l2, Φ k y1 y2)
Φ i x1 x2 ( y1 y2, Φ i y1 y2 -∗ ([ list] ky1;y2 <[i:=y1]>l1;<[i:=y2]>l2, Φ k y1 y2)).
Proof.
intros Hl1 Hl2. rewrite big_sepL2_alt. apply pure_elim_l=> Hl.
rewrite {1}big_sepL_insert_acc; last by rewrite lookup_zip_with; simplify_option_eq.
apply sep_mono_r. apply forall_intro => y1. apply forall_intro => y2.
rewrite big_sepL2_alt !length_insert pure_True // left_id -insert_zip_with.
by rewrite (forall_elim (y1, y2)).
Qed.
Lemma big_sepL2_lookup_acc Φ l1 l2 i x1 x2 :
l1 !! i = Some x1 l2 !! i = Some x2
([ list] ky1;y2 l1;l2, Φ k y1 y2)
Φ i x1 x2 (Φ i x1 x2 -∗ ([ list] ky1;y2 l1;l2, Φ k y1 y2)).
Proof.
intros. rewrite {1}big_sepL2_insert_acc // (forall_elim x1) (forall_elim x2).
by rewrite !list_insert_id.
Qed.
Lemma big_sepL2_lookup Φ l1 l2 i x1 x2
`{!TCOr ( j y1 y2, Affine (Φ j y1 y2)) (Absorbing (Φ i x1 x2))} :
l1 !! i = Some x1 l2 !! i = Some x2
([ list] ky1;y2 l1;l2, Φ k y1 y2) Φ i x1 x2.
Proof.
intros Hx1 Hx2. destruct select (TCOr _ _).
- rewrite -(take_drop_middle l1 i x1) // -(take_drop_middle l2 i x2) //.
apply lookup_lt_Some in Hx1. apply lookup_lt_Some in Hx2.
rewrite big_sepL2_app_same_length /= 2?length_take; last lia.
rewrite (_ : _ + 0 = i); last lia.
rewrite sep_elim_r sep_elim_l //.
- rewrite big_sepL2_lookup_acc // sep_elim_l //.
Qed.
Lemma big_sepL2_fmap_l {A'} (f : A A') (Φ : nat A' B PROP) l1 l2 :
([ list] ky1;y2 f <$> l1; l2, Φ k y1 y2)
⊣⊢ ([ list] ky1;y2 l1;l2, Φ k (f y1) y2).
Proof.
rewrite !big_sepL2_alt length_fmap zip_with_fmap_l zip_with_zip big_sepL_fmap.
by f_equiv; f_equiv=> k [??].
Qed.
Lemma big_sepL2_fmap_r {B'} (g : B B') (Φ : nat A B' PROP) l1 l2 :
([ list] ky1;y2 l1; g <$> l2, Φ k y1 y2)
⊣⊢ ([ list] ky1;y2 l1;l2, Φ k y1 (g y2)).
Proof.
rewrite !big_sepL2_alt length_fmap zip_with_fmap_r zip_with_zip big_sepL_fmap.
by f_equiv; f_equiv=> k [??].
Qed.
Lemma big_sepL2_reverse_2 (Φ : A B PROP) l1 l2 :
([ list] y1;y2 l1;l2, Φ y1 y2) ([ list] y1;y2 reverse l1;reverse l2, Φ y1 y2).
Proof.
revert l2. induction l1 as [|x1 l1 IH]; intros [|x2 l2]; simpl; auto using False_elim.
rewrite !reverse_cons (comm bi_sep) IH.
by rewrite (big_sepL2_app _ _ [x1] _ [x2]) big_sepL2_singleton wand_elim_l.
Qed.
Lemma big_sepL2_reverse (Φ : A B PROP) l1 l2 :
([ list] y1;y2 reverse l1;reverse l2, Φ y1 y2) ⊣⊢ ([ list] y1;y2 l1;l2, Φ y1 y2).
Proof. apply (anti_symm _); by rewrite big_sepL2_reverse_2 ?reverse_involutive. Qed.
Lemma big_sepL2_replicate_l l x Φ n :
length l = n
([ list] kx1;x2 replicate n x; l, Φ k x1 x2) ⊣⊢ [ list] kx2 l, Φ k x x2.
Proof. intros <-. revert Φ. induction l as [|y l IH]=> //= Φ. by rewrite IH. Qed.
Lemma big_sepL2_replicate_r l x Φ n :
length l = n
([ list] kx1;x2 l;replicate n x, Φ k x1 x2) ⊣⊢ [ list] kx1 l, Φ k x1 x.
Proof. intros <-. revert Φ. induction l as [|y l IH]=> //= Φ. by rewrite IH. Qed.
Lemma big_sepL2_sep Φ Ψ l1 l2 :
([ list] ky1;y2 l1;l2, Φ k y1 y2 Ψ k y1 y2)
⊣⊢ ([ list] ky1;y2 l1;l2, Φ k y1 y2) ([ list] ky1;y2 l1;l2, Ψ k y1 y2).
Proof.
rewrite !big_sepL2_alt big_sepL_sep !persistent_and_affinely_sep_l.
rewrite -assoc (assoc _ _ (<affine> _)%I). rewrite -(comm bi_sep (<affine> _)%I).
rewrite -assoc (assoc _ _ (<affine> _)%I) -!persistent_and_affinely_sep_l.
by rewrite affinely_and_r persistent_and_affinely_sep_l idemp.
Qed.
Lemma big_sepL2_sep_2 Φ Ψ l1 l2 :
([ list] ky1;y2 l1;l2, Φ k y1 y2) -∗
([ list] ky1;y2 l1;l2, Ψ k y1 y2) -∗
([ list] ky1;y2 l1;l2, Φ k y1 y2 Ψ k y1 y2).
Proof. apply entails_wand, wand_intro_r. rewrite big_sepL2_sep //. Qed.
Lemma big_sepL2_and Φ Ψ l1 l2 :
([ list] ky1;y2 l1;l2, Φ k y1 y2 Ψ k y1 y2)
([ list] ky1;y2 l1;l2, Φ k y1 y2) ([ list] ky1;y2 l1;l2, Ψ k y1 y2).
Proof. auto using and_intro, big_sepL2_mono, and_elim_l, and_elim_r. Qed.
Lemma big_sepL2_pure_1 (φ : nat A B Prop) l1 l2 :
([ list] ky1;y2 l1;l2, φ k y1 y2) ⊢@{PROP}
⌜∀ k y1 y2, l1 !! k = Some y1 l2 !! k = Some y2 φ k y1 y2⌝.
Proof.
rewrite big_sepL2_alt big_sepL_pure_1.
rewrite -pure_and. f_equiv=>-[Hlen Hlookup] k y1 y2 Hy1 Hy2.
eapply (Hlookup k (y1, y2)).
rewrite lookup_zip_with Hy1 /= Hy2 /= //.
Qed.
Lemma big_sepL2_affinely_pure_2 (φ : nat A B Prop) l1 l2 :
length l1 = length l2
<affine> ⌜∀ k y1 y2, l1 !! k = Some y1 l2 !! k = Some y2 φ k y1 y2 ⊢@{PROP}
([ list] ky1;y2 l1;l2, <affine> φ k y1 y2).
Proof.
intros Hdom. rewrite big_sepL2_alt.
rewrite -big_sepL_affinely_pure_2.
rewrite affinely_and_r -pure_and. f_equiv. f_equiv=>-Hforall.
split; first done.
intros k [y1 y2] (? & ? & [= <- <-] & Hy1 & Hy2)%lookup_zip_with_Some.
by eapply Hforall.
Qed.
(** The general backwards direction requires [BiAffine] to cover the empty case. *)
Lemma big_sepL2_pure `{!BiAffine PROP} (φ : nat A B Prop) l1 l2 :
([ list] ky1;y2 l1;l2, φ k y1 y2) ⊣⊢@{PROP}
length l1 = length l2
k y1 y2, l1 !! k = Some y1 l2 !! k = Some y2 φ k y1 y2⌝.
Proof.
apply (anti_symm ()).
{ rewrite pure_and. apply and_intro.
- apply big_sepL2_length.
- apply big_sepL2_pure_1. }
rewrite -(affine_affinely ⌜_⌝%I).
rewrite pure_and -affinely_and_r.
apply pure_elim_l=>Hdom.
rewrite big_sepL2_affinely_pure_2 //. by setoid_rewrite affinely_elim.
Qed.
Lemma big_sepL2_persistently `{!BiAffine PROP} Φ l1 l2 :
<pers> ([ list] ky1;y2 l1;l2, Φ k y1 y2)
⊣⊢ [ list] ky1;y2 l1;l2, <pers> (Φ k y1 y2).
Proof.
by rewrite !big_sepL2_alt persistently_and persistently_pure big_sepL_persistently.
Qed.
Lemma big_sepL2_intro Φ l1 l2 :
length l1 = length l2
( k x1 x2, l1 !! k = Some x1 l2 !! k = Some x2 Φ k x1 x2)
[ list] kx1;x2 l1;l2, Φ k x1 x2.
Proof.
revert l2 Φ. induction l1 as [|x1 l1 IH]=> -[|x2 l2] Φ ?; simplify_eq/=.
{ by apply (affine _). }
rewrite intuitionistically_sep_dup. f_equiv.
- rewrite (forall_elim 0) (forall_elim x1) (forall_elim x2).
by rewrite !pure_True // !True_impl intuitionistically_elim.
- rewrite -IH //. f_equiv.
by apply forall_intro=> k; by rewrite (forall_elim (S k)).
Qed.
Lemma big_sepL2_forall `{!BiAffine PROP} Φ l1 l2 :
( k x1 x2, Persistent (Φ k x1 x2))
([ list] kx1;x2 l1;l2, Φ k x1 x2) ⊣⊢
length l1 = length l2
( k x1 x2, l1 !! k = Some x1 l2 !! k = Some x2 Φ k x1 x2).
Proof.
intros . apply (anti_symm _).
{ apply and_intro; [apply big_sepL2_length|].
apply forall_intro=> k. apply forall_intro=> x1. apply forall_intro=> x2.
do 2 (apply impl_intro_l; apply pure_elim_l=> ?). by apply: big_sepL2_lookup. }
apply pure_elim_l=> Hlen.
revert l2 Φ Hlen. induction l1 as [|x1 l1 IH]=> -[|x2 l2] Φ Hlen; simplify_eq/=.
{ by apply (affine _). }
rewrite -persistent_and_sep_1. apply and_intro.
- rewrite (forall_elim 0) (forall_elim x1) (forall_elim x2).
by rewrite !pure_True // !True_impl.
- rewrite -IH //.
by apply forall_intro=> k; by rewrite (forall_elim (S k)).
Qed.
Lemma big_sepL2_impl Φ Ψ l1 l2 :
([ list] ky1;y2 l1;l2, Φ k y1 y2) -∗
( k x1 x2,
l1 !! k = Some x1 l2 !! k = Some x2 Φ k x1 x2 -∗ Ψ k x1 x2) -∗
[ list] ky1;y2 l1;l2, Ψ k y1 y2.
Proof.
apply entails_wand. rewrite -(idemp bi_and (big_sepL2 _ _ _)) {1}big_sepL2_length.
apply pure_elim_l=> ?. rewrite big_sepL2_intro //.
apply bi.wand_intro_l. rewrite -big_sepL2_sep. by setoid_rewrite wand_elim_l.
Qed.
Lemma big_sepL2_wand Φ Ψ l1 l2 :
([ list] ky1;y2 l1;l2, Φ k y1 y2) -∗
([ list] ky1;y2 l1;l2, Φ k y1 y2 -∗ Ψ k y1 y2) -∗
[ list] ky1;y2 l1;l2, Ψ k y1 y2.
Proof.
apply entails_wand, wand_intro_r. rewrite -big_sepL2_sep.
setoid_rewrite wand_elim_r. done.
Qed.
Lemma big_sepL2_delete Φ l1 l2 i x1 x2 :
l1 !! i = Some x1 l2 !! i = Some x2
([ list] ky1;y2 l1;l2, Φ k y1 y2) ⊣⊢
Φ i x1 x2 [ list] ky1;y2 l1;l2, if decide (k = i) then emp else Φ k y1 y2.
Proof.
intros. rewrite -(take_drop_middle l1 i x1) // -(take_drop_middle l2 i x2) //.
assert (i < length l1 i < length l2) as [??] by eauto using lookup_lt_Some.
rewrite !big_sepL2_app_same_length /=; [|rewrite ?length_take; lia..].
rewrite Nat.add_0_r length_take_le; [|lia].
rewrite decide_True // left_id.
rewrite assoc -!(comm _ (Φ _ _ _)) -assoc. do 2 f_equiv.
- apply big_sepL2_proper=> k y1 y2 Hk. apply lookup_lt_Some in Hk.
rewrite length_take in Hk. by rewrite decide_False; last lia.
- apply big_sepL2_proper=> k y1 y2 _. by rewrite decide_False; last lia.
Qed.
Lemma big_sepL2_delete' `{!BiAffine PROP} Φ l1 l2 i x1 x2 :
l1 !! i = Some x1 l2 !! i = Some x2
([ list] ky1;y2 l1;l2, Φ k y1 y2) ⊣⊢
Φ i x1 x2 [ list] ky1;y2 l1;l2, k i Φ k y1 y2.
Proof.
intros. rewrite big_sepL2_delete //. (do 2 f_equiv)=> k y1 y2.
rewrite -decide_emp. by repeat case_decide.
Qed.
Lemma big_sepL2_lookup_acc_impl {Φ l1 l2} i x1 x2 :
l1 !! i = Some x1
l2 !! i = Some x2
([ list] ky1;y2 l1;l2, Φ k y1 y2) -∗
(* We obtain [Φ] for the [x1] and [x2] *)
Φ i x1 x2
(* We reobtain the bigop for a predicate [Ψ] selected by the user *)
Ψ,
( k y1 y2,
l1 !! k = Some y1 l2 !! k = Some y2 k i
Φ k y1 y2 -∗ Ψ k y1 y2) -∗
Ψ i x1 x2 -∗
[ list] ky1;y2 l1;l2, Ψ k y1 y2.
Proof.
intros. rewrite big_sepL2_delete //. apply entails_wand, sep_mono_r, forall_intro=> Ψ.
apply wand_intro_r, wand_intro_l.
rewrite (big_sepL2_delete Ψ l1 l2 i) //. apply sep_mono_r.
eapply wand_apply; [apply wand_entails, big_sepL2_impl|apply sep_mono_r].
apply intuitionistically_intro', forall_intro=> k;
apply forall_intro=> y1; apply forall_intro=> y2.
do 2 (apply impl_intro_l, pure_elim_l=> ?); apply wand_intro_r.
rewrite (forall_elim k) (forall_elim y1) (forall_elim y2).
rewrite !(pure_True (_ = Some _)) // !left_id.
destruct (decide _) as [->|]; [by apply: affine|].
by rewrite pure_True //left_id intuitionistically_elim wand_elim_l.
Qed.
Lemma big_sepL2_later_1 `{!BiAffine PROP} Φ l1 l2 :
( [ list] ky1;y2 l1;l2, Φ k y1 y2) [ list] ky1;y2 l1;l2, Φ k y1 y2.
Proof.
rewrite !big_sepL2_alt later_and big_sepL_later (timeless _ ).
rewrite except_0_and. auto using and_mono, except_0_intro.
Qed.
Lemma big_sepL2_later_2 Φ l1 l2 :
([ list] ky1;y2 l1;l2, Φ k y1 y2) [ list] ky1;y2 l1;l2, Φ k y1 y2.
Proof.
rewrite !big_sepL2_alt later_and big_sepL_later_2.
auto using and_mono, later_intro.
Qed.
Lemma big_sepL2_laterN_2 Φ n l1 l2 :
([ list] ky1;y2 l1;l2, ▷^n Φ k y1 y2) ▷^n [ list] ky1;y2 l1;l2, Φ k y1 y2.
Proof.
rewrite !big_sepL2_alt laterN_and big_sepL_laterN_2.
auto using and_mono, laterN_intro.
Qed.
Lemma big_sepL2_flip Φ l1 l2 :
([ list] ky1;y2 l2; l1, Φ k y2 y1) ⊣⊢ ([ list] ky1;y2 l1; l2, Φ k y1 y2).
Proof.
revert Φ l2. induction l1 as [|x1 l1 IH]=> Φ -[|x2 l2]//=; simplify_eq.
by rewrite IH.
Qed.
Lemma big_sepL2_sepL (Φ1 : nat A PROP) (Φ2 : nat B PROP) l1 l2 :
length l1 = length l2
([ list] ky1;y2 l1;l2, Φ1 k y1 Φ2 k y2) ⊣⊢
([ list] ky1 l1, Φ1 k y1) ([ list] ky2 l2, Φ2 k y2).
Proof.
intros. rewrite -big_sepL_sep_zip // big_sepL2_alt pure_True // left_id //.
Qed.
Lemma big_sepL2_sepL_2 (Φ1 : nat A PROP) (Φ2 : nat B PROP) l1 l2 :
length l1 = length l2
([ list] ky1 l1, Φ1 k y1) -∗
([ list] ky2 l2, Φ2 k y2) -∗
[ list] ky1;y2 l1;l2, Φ1 k y1 Φ2 k y2.
Proof. intros. apply entails_wand, wand_intro_r. by rewrite big_sepL2_sepL. Qed.
End sep_list2.
Lemma big_sepL2_const_sepL_l {A B} (Φ : nat A PROP) (l1 : list A) (l2 : list B) :
([ list] ky1;y2 l1;l2, Φ k y1)
⊣⊢ length l1 = length l2 ([ list] ky1 l1, Φ k y1).
Proof.
rewrite big_sepL2_alt.
trans (length l1 = length l2 [ list] ky1 (zip l1 l2).*1, Φ k y1)%I.
{ rewrite big_sepL_fmap //. }
apply (anti_symm ()); apply pure_elim_l=> Hl; rewrite fst_zip;
rewrite ?Hl //;
(apply and_intro; [by apply pure_intro|done]).
Qed.
Lemma big_sepL2_const_sepL_r {A B} (Φ : nat B PROP) (l1 : list A) (l2 : list B) :
([ list] ky1;y2 l1;l2, Φ k y2)
⊣⊢ length l1 = length l2 ([ list] ky2 l2, Φ k y2).
Proof. by rewrite big_sepL2_flip big_sepL2_const_sepL_l (symmetry_iff (=)). Qed.
Lemma big_sepL2_sep_sepL_l {A B} (Φ : nat A PROP)
(Ψ : nat A B PROP) l1 l2 :
([ list] ky1;y2 l1;l2, Φ k y1 Ψ k y1 y2)
⊣⊢ ([ list] ky1 l1, Φ k y1) ([ list] ky1;y2 l1;l2, Ψ k y1 y2).
Proof.
rewrite big_sepL2_sep big_sepL2_const_sepL_l. apply (anti_symm _).
{ rewrite and_elim_r. done. }
rewrite !big_sepL2_alt [(_ _)%I]comm -!persistent_and_sep_assoc.
apply pure_elim_l=>Hl. apply and_intro.
{ apply pure_intro. done. }
rewrite [(_ _)%I]comm. apply sep_mono; first done.
apply and_intro; last done.
apply pure_intro. done.
Qed.
Lemma big_sepL2_sep_sepL_r {A B} (Φ : nat A B PROP)
(Ψ : nat B PROP) l1 l2 :
([ list] ky1;y2 l1;l2, Φ k y1 y2 Ψ k y2)
⊣⊢ ([ list] ky1;y2 l1;l2, Φ k y1 y2) ([ list] ky2 l2, Ψ k y2).
Proof.
rewrite !(big_sepL2_flip _ _ l1). setoid_rewrite (comm bi_sep).
by rewrite big_sepL2_sep_sepL_l.
Qed.
Lemma big_sepL_sepL2_diag {A} (Φ : nat A A PROP) (l : list A) :
([ list] ky l, Φ k y y)
([ list] ky1;y2 l;l, Φ k y1 y2).
Proof.
rewrite big_sepL2_alt. rewrite pure_True // left_id.
rewrite zip_diag big_sepL_fmap /=. done.
Qed.
Lemma big_sepL2_ne_2 {A B : ofe}
(Φ Ψ : nat A B PROP) l1 l2 l1' l2' n :
l1 {n} l1' l2 {n} l2'
( k y1 y1' y2 y2',
l1 !! k = Some y1 l1' !! k = Some y1' y1 {n} y1'
l2 !! k = Some y2 l2' !! k = Some y2' y2 {n} y2'
Φ k y1 y2 {n} Ψ k y1' y2')
([ list] k y1;y2 l1;l2, Φ k y1 y2)%I {n} ([ list] k y1;y2 l1';l2', Ψ k y1 y2)%I.
Proof.
intros Hl1 Hl2 Hf. rewrite !big_sepL2_alt. f_equiv.
{ do 2 f_equiv; by apply: length_ne. }
apply big_opL_ne_2; [by f_equiv|].
intros k [x1 y1] [x2 y2] (?&?&[=<- <-]&?&?)%lookup_zip_with_Some
(?&?&[=<- <-]&?&?)%lookup_zip_with_Some [??]; naive_solver.
Qed.
Section and_list.
Context {A : Type}.
Implicit Types l : list A.
Implicit Types Φ Ψ : nat A PROP.
Lemma big_andL_nil Φ : ([ list] ky nil, Φ k y) ⊣⊢ True.
Proof. done. Qed.
Lemma big_andL_nil' P Φ : P [ list] ky nil, Φ k y.
Proof. by apply pure_intro. Qed.
Lemma big_andL_cons Φ x l :
([ list] ky x :: l, Φ k y) ⊣⊢ Φ 0 x [ list] ky l, Φ (S k) y.
Proof. by rewrite big_opL_cons. Qed.
Lemma big_andL_singleton Φ x : ([ list] ky [x], Φ k y) ⊣⊢ Φ 0 x.
Proof. by rewrite big_opL_singleton. Qed.
Lemma big_andL_app Φ l1 l2 :
([ list] ky l1 ++ l2, Φ k y)
⊣⊢ ([ list] ky l1, Φ k y) ([ list] ky l2, Φ (length l1 + k) y).
Proof. by rewrite big_opL_app. Qed.
Lemma big_andL_snoc Φ l x :
([ list] ky l ++ [x], Φ k y) ⊣⊢ ([ list] ky l, Φ k y) Φ (length l) x.
Proof. by rewrite big_opL_snoc. Qed.
Lemma big_andL_submseteq (Φ : A PROP) l1 l2 :
l1 ⊆+ l2 ([ list] y l2, Φ y) [ list] y l1, Φ y.
Proof.
intros [l ->]%submseteq_Permutation. by rewrite big_andL_app and_elim_l.
Qed.
(** The lemmas [big_andL_mono], [big_andL_ne] and [big_andL_proper] are more
generic than the instances as they also give [l !! k = Some y] in the premise. *)
Lemma big_andL_mono Φ Ψ l :
( k y, l !! k = Some y Φ k y Ψ k y)
([ list] k y l, Φ k y) [ list] k y l, Ψ k y.
Proof. apply big_opL_gen_proper; apply _. Qed.
Lemma big_andL_ne Φ Ψ l n :
( k y, l !! k = Some y Φ k y {n} Ψ k y)
([ list] k y l, Φ k y)%I {n} ([ list] k y l, Ψ k y)%I.
Proof. apply big_opL_ne. Qed.
Lemma big_andL_proper Φ Ψ l :
( k y, l !! k = Some y Φ k y ⊣⊢ Ψ k y)
([ list] k y l, Φ k y) ⊣⊢ ([ list] k y l, Ψ k y).
Proof. apply big_opL_proper. Qed.
(** No need to declare instances for non-expansiveness and properness, we
get both from the generic [big_opL] instances. *)
Global Instance big_andL_mono' :
Proper (pointwise_relation _ (pointwise_relation _ ()) ==> (=) ==> ())
(big_opL (@bi_and PROP) (A:=A)).
Proof. intros f g Hf m ? <-. apply big_andL_mono; intros; apply Hf. Qed.
Global Instance big_andL_id_mono' :
Proper (Forall2 () ==> ()) (big_opL (@bi_and PROP) (λ _ P, P)).
Proof. by induction 1 as [|P Q Ps Qs HPQ ? IH]; rewrite /= ?HPQ ?IH. Qed.
Global Instance big_andL_nil_absorbing Φ :
Absorbing ([ list] kx [], Φ k x).
Proof. simpl; apply _. Qed.
Lemma big_andL_absorbing Φ l :
( k x, l !! k = Some x Absorbing (Φ k x))
Absorbing ([ list] kx l, Φ k x).
Proof. apply big_opL_closed; apply _. Qed.
Global Instance big_andL_absorbing' Φ l :
( k x, Absorbing (Φ k x)) Absorbing ([ list] kx l, Φ k x).
Proof. intros; apply big_andL_absorbing, _. Qed.
Global Instance big_andL_nil_persistent Φ :
Persistent ([ list] kx [], Φ k x).
Proof. simpl; apply _. Qed.
Lemma big_andL_persistent Φ l :
( k x, l !! k = Some x Persistent (Φ k x)) Persistent ([ list] kx l, Φ k x).
Proof. apply big_opL_closed; apply _. Qed.
Global Instance big_andL_persistent' Φ l :
( k x, Persistent (Φ k x)) Persistent ([ list] kx l, Φ k x).
Proof. intros; apply big_andL_persistent, _. Qed.
Global Instance big_andL_nil_timeless Φ :
Timeless ([ list] kx [], Φ k x).
Proof. simpl; apply _. Qed.
Lemma big_andL_timeless Φ l :
( k x, l !! k = Some x Timeless (Φ k x))
Timeless ([ list] kx l, Φ k x).
Proof. apply big_opL_closed; apply _. Qed.
Global Instance big_andL_timeless' Φ l :
( k x, Timeless (Φ k x)) Timeless ([ list] kx l, Φ k x).
Proof. intros; apply big_andL_timeless, _. Qed.
Lemma big_andL_lookup Φ l i x :
l !! i = Some x ([ list] ky l, Φ k y) Φ i x.
Proof.
intros. rewrite -(take_drop_middle l i x) // big_andL_app /=.
rewrite Nat.add_0_r length_take_le;
eauto using lookup_lt_Some, Nat.lt_le_incl, and_elim_l', and_elim_r'.
Qed.
Lemma big_andL_elem_of (Φ : A PROP) l x :
x l ([ list] y l, Φ y) Φ x.
Proof.
intros [i ?]%elem_of_list_lookup. by eapply (big_andL_lookup (λ _, Φ)).
Qed.
Lemma big_andL_fmap {B} (f : A B) (Φ : nat B PROP) l :
([ list] ky f <$> l, Φ k y) ⊣⊢ ([ list] ky l, Φ k (f y)).
Proof. by rewrite big_opL_fmap. Qed.
Lemma big_andL_bind {B} (f : A list B) (Φ : B PROP) l :
([ list] y l ≫= f, Φ y) ⊣⊢ ([ list] x l, [ list] y f x, Φ y).
Proof. by rewrite big_opL_bind. Qed.
Lemma big_andL_and Φ Ψ l :
([ list] kx l, Φ k x Ψ k x)
⊣⊢ ([ list] kx l, Φ k x) ([ list] kx l, Ψ k x).
Proof. by rewrite big_opL_op. Qed.
Lemma big_andL_persistently Φ l :
<pers> ([ list] kx l, Φ k x) ⊣⊢ [ list] kx l, <pers> (Φ k x).
Proof. apply (big_opL_commute _). Qed.
Lemma big_andL_forall Φ l :
([ list] kx l, Φ k x) ⊣⊢ ( k x, l !! k = Some x Φ k x).
Proof.
apply (anti_symm _).
{ apply forall_intro=> k; apply forall_intro=> x.
apply impl_intro_l, pure_elim_l=> ?; by apply: big_andL_lookup. }
revert Φ. induction l as [|x l IH]=> Φ; [by auto using big_andL_nil'|].
rewrite big_andL_cons. apply and_intro.
- by rewrite (forall_elim 0) (forall_elim x) pure_True // True_impl.
- rewrite -IH. apply forall_intro=> k; by rewrite (forall_elim (S k)).
Qed.
Lemma big_andL_intro Φ l :
( k x, l !! k = Some x Φ k x) [ list] kx l, Φ k x.
Proof. rewrite big_andL_forall //. Qed.
Lemma big_andL_impl Φ Ψ m :
([ list] kx m, Φ k x)
( k x, m !! k = Some x Φ k x Ψ k x)
[ list] kx m, Ψ k x.
Proof.
rewrite -big_andL_forall -big_andL_and.
by setoid_rewrite bi.impl_elim_r.
Qed.
Lemma big_andL_pure_1 (φ : nat A Prop) l :
([ list] kx l, φ k x) ⊢@{PROP} ⌜∀ k x, l !! k = Some x φ k x⌝.
Proof.
induction l as [|x l IH] using rev_ind.
{ apply pure_intro=>??. rewrite lookup_nil. done. }
rewrite big_andL_snoc // IH -pure_and.
f_equiv=>-[Hl Hx] k y /lookup_app_Some =>-[Hy|[Hlen Hy]].
- by apply Hl.
- apply list_lookup_singleton_Some in Hy as [Hk ->].
replace k with (length l) by lia. done.
Qed.
Lemma big_andL_pure_2 (φ : nat A Prop) l :
⌜∀ k x, l !! k = Some x φ k x ⊢@{PROP} ([ list] kx l, φ k x).
Proof.
rewrite big_andL_forall pure_forall_1. f_equiv=>k.
rewrite pure_forall_1. f_equiv=>x. apply pure_impl_1.
Qed.
Lemma big_andL_pure (φ : nat A Prop) l :
([ list] kx l, φ k x) ⊣⊢@{PROP} ⌜∀ k x, l !! k = Some x φ k x⌝.
Proof.
apply (anti_symm ()); first by apply big_andL_pure_1.
apply big_andL_pure_2.
Qed.
Lemma big_andL_later Φ l :
([ list] kx l, Φ k x) ⊣⊢ ([ list] kx l, Φ k x).
Proof. apply (big_opL_commute _). Qed.
Lemma big_andL_laterN Φ n l :
▷^n ([ list] kx l, Φ k x) ⊣⊢ ([ list] kx l, ▷^n Φ k x).
Proof. apply (big_opL_commute _). Qed.
End and_list.
Section or_list.
Context {A : Type}.
Implicit Types l : list A.
Implicit Types Φ Ψ : nat A PROP.
Lemma big_orL_nil Φ : ([ list] ky nil, Φ k y) ⊣⊢ False.
Proof. done. Qed.
Lemma big_orL_cons Φ x l :
([ list] ky x :: l, Φ k y) ⊣⊢ Φ 0 x [ list] ky l, Φ (S k) y.
Proof. by rewrite big_opL_cons. Qed.
Lemma big_orL_singleton Φ x : ([ list] ky [x], Φ k y) ⊣⊢ Φ 0 x.
Proof. by rewrite big_opL_singleton. Qed.
Lemma big_orL_app Φ l1 l2 :
([ list] ky l1 ++ l2, Φ k y)
⊣⊢ ([ list] ky l1, Φ k y) ([ list] ky l2, Φ (length l1 + k) y).
Proof. by rewrite big_opL_app. Qed.
Lemma big_orL_snoc Φ l x :
([ list] ky l ++ [x], Φ k y) ⊣⊢ ([ list] ky l, Φ k y) Φ (length l) x.
Proof. by rewrite big_opL_snoc. Qed.
Lemma big_orL_submseteq (Φ : A PROP) l1 l2 :
l1 ⊆+ l2 ([ list] y l1, Φ y) [ list] y l2, Φ y.
Proof.
intros [l ->]%submseteq_Permutation. by rewrite big_orL_app -or_intro_l.
Qed.
(** The lemmas [big_orL_mono], [big_orL_ne] and [big_orL_proper] are more
generic than the instances as they also give [l !! k = Some y] in the premise. *)
Lemma big_orL_mono Φ Ψ l :
( k y, l !! k = Some y Φ k y Ψ k y)
([ list] k y l, Φ k y) [ list] k y l, Ψ k y.
Proof. apply big_opL_gen_proper; apply _. Qed.
Lemma big_orL_ne Φ Ψ l n :
( k y, l !! k = Some y Φ k y {n} Ψ k y)
([ list] k y l, Φ k y)%I {n} ([ list] k y l, Ψ k y)%I.
Proof. apply big_opL_ne. Qed.
Lemma big_orL_proper Φ Ψ l :
( k y, l !! k = Some y Φ k y ⊣⊢ Ψ k y)
([ list] k y l, Φ k y) ⊣⊢ ([ list] k y l, Ψ k y).
Proof. apply big_opL_proper. Qed.
(** No need to declare instances for non-expansiveness and properness, we
get both from the generic [big_opL] instances. *)
Global Instance big_orL_mono' :
Proper (pointwise_relation _ (pointwise_relation _ ()) ==> (=) ==> ())
(big_opL (@bi_or PROP) (A:=A)).
Proof. intros f g Hf m ? <-. apply big_orL_mono; intros; apply Hf. Qed.
Global Instance big_orL_id_mono' :
Proper (Forall2 () ==> ()) (big_opL (@bi_or PROP) (λ _ P, P)).
Proof. by induction 1 as [|P Q Ps Qs HPQ ? IH]; rewrite /= ?HPQ ?IH. Qed.
Global Instance big_orL_nil_persistent Φ :
Persistent ([ list] kx [], Φ k x).
Proof. simpl; apply _. Qed.
Global Instance big_orL_persistent Φ l :
( k x, l !! k = Some x Persistent (Φ k x))
Persistent ([ list] kx l, Φ k x).
Proof. apply big_opL_closed; apply _. Qed.
Global Instance big_orL_persistent' Φ l :
( k x, Persistent (Φ k x)) Persistent ([ list] kx l, Φ k x).
Proof. intros; apply big_orL_persistent, _. Qed.
Global Instance big_orL_nil_timeless Φ :
Timeless ([ list] kx [], Φ k x).
Proof. simpl; apply _. Qed.
Global Instance big_orL_timeless Φ l :
( k x, l !! k = Some x Timeless (Φ k x))
Timeless ([ list] kx l, Φ k x).
Proof. apply big_opL_closed; apply _. Qed.
Global Instance big_orL_timeless' Φ l :
( k x, Timeless (Φ k x)) Timeless ([ list] kx l, Φ k x).
Proof. intros; apply big_orL_timeless, _. Qed.
Lemma big_orL_intro Φ l i x :
l !! i = Some x Φ i x ([ list] ky l, Φ k y).
Proof.
intros. rewrite -(take_drop_middle l i x) // big_orL_app /=.
rewrite Nat.add_0_r length_take_le;
eauto using lookup_lt_Some, Nat.lt_le_incl, or_intro_l', or_intro_r'.
Qed.
Lemma big_orL_elem_of (Φ : A PROP) l x :
x l Φ x ([ list] y l, Φ y).
Proof.
intros [i ?]%elem_of_list_lookup; by eapply (big_orL_intro (λ _, Φ)).
Qed.
Lemma big_orL_fmap {B} (f : A B) (Φ : nat B PROP) l :
([ list] ky f <$> l, Φ k y) ⊣⊢ ([ list] ky l, Φ k (f y)).
Proof. by rewrite big_opL_fmap. Qed.
Lemma big_orL_bind {B} (f : A list B) (Φ : B PROP) l :
([ list] y l ≫= f, Φ y) ⊣⊢ ([ list] x l, [ list] y f x, Φ y).
Proof. by rewrite big_opL_bind. Qed.
Lemma big_orL_or Φ Ψ l :
([ list] kx l, Φ k x Ψ k x)
⊣⊢ ([ list] kx l, Φ k x) ([ list] kx l, Ψ k x).
Proof. by rewrite big_opL_op. Qed.
Lemma big_orL_persistently Φ l :
<pers> ([ list] kx l, Φ k x) ⊣⊢ [ list] kx l, <pers> (Φ k x).
Proof. apply (big_opL_commute _). Qed.
Lemma big_orL_exist Φ l :
([ list] kx l, Φ k x) ⊣⊢ ( k x, l !! k = Some x Φ k x).
Proof.
apply (anti_symm _).
{ revert Φ. induction l as [|x l IH]=> Φ.
{ rewrite big_orL_nil. apply False_elim. }
rewrite big_orL_cons. apply or_elim.
- by rewrite -(exist_intro 0) -(exist_intro x) pure_True // left_id.
- rewrite IH. apply exist_elim=> k. by rewrite -(exist_intro (S k)). }
apply exist_elim=> k; apply exist_elim=> x. apply pure_elim_l=> ?.
by apply: big_orL_intro.
Qed.
Lemma big_orL_pure (φ : nat A Prop) l :
([ list] kx l, φ k x) ⊣⊢@{PROP} ⌜∃ k x, l !! k = Some x φ k x⌝.
Proof.
rewrite big_orL_exist.
rewrite pure_exist. apply exist_proper=>k.
rewrite pure_exist. apply exist_proper=>x.
rewrite -pure_and. done.
Qed.
Lemma big_orL_sep_l P Φ l :
P ([ list] kx l, Φ k x) ⊣⊢ ([ list] kx l, P Φ k x).
Proof.
rewrite !big_orL_exist sep_exist_l.
f_equiv=> k. rewrite sep_exist_l. f_equiv=> x.
by rewrite !persistent_and_affinely_sep_l !assoc (comm _ P).
Qed.
Lemma big_orL_sep_r Q Φ l :
([ list] kx l, Φ k x) Q ⊣⊢ ([ list] kx l, Φ k x Q).
Proof. setoid_rewrite (comm bi_sep). apply big_orL_sep_l. Qed.
Lemma big_orL_later Φ l :
l []
([ list] kx l, Φ k x) ⊣⊢ ([ list] kx l, Φ k x).
Proof. apply (big_opL_commute1 _). Qed.
Lemma big_orL_laterN Φ n l :
l []
▷^n ([ list] kx l, Φ k x) ⊣⊢ ([ list] kx l, ▷^n Φ k x).
Proof. apply (big_opL_commute1 _). Qed.
End or_list.
(** ** Big ops over finite maps *)
Section sep_map.
Context `{Countable K} {A : Type}.
Implicit Types m : gmap K A.
Implicit Types Φ Ψ : K A PROP.
(** The lemmas [big_sepM_mono], [big_sepM_ne] and [big_sepM_proper] are more
generic than the instances as they also give [l !! k = Some y] in the premise. *)
Lemma big_sepM_mono Φ Ψ m :
( k x, m !! k = Some x Φ k x Ψ k x)
([ map] k x m, Φ k x) [ map] k x m, Ψ k x.
Proof. apply big_opM_gen_proper; apply _ || auto. Qed.
Lemma big_sepM_ne Φ Ψ m n :
( k x, m !! k = Some x Φ k x {n} Ψ k x)
([ map] k x m, Φ k x)%I {n} ([ map] k x m, Ψ k x)%I.
Proof. apply big_opM_ne. Qed.
Lemma big_sepM_proper Φ Ψ m :
( k x, m !! k = Some x Φ k x ⊣⊢ Ψ k x)
([ map] k x m, Φ k x) ⊣⊢ ([ map] k x m, Ψ k x).
Proof. apply big_opM_proper. Qed.
(** No need to declare instances for non-expansiveness and properness, we
get both from the generic [big_opM] instances. *)
Global Instance big_sepM_mono' :
Proper (pointwise_relation _ (pointwise_relation _ ()) ==> (=) ==> ())
(big_opM (@bi_sep PROP) (K:=K) (A:=A)).
Proof. intros f g Hf m ? <-. apply big_sepM_mono; intros; apply Hf. Qed.
Global Instance big_sepM_flip_mono' :
Proper (pointwise_relation _ (pointwise_relation _ (flip ())) ==> (=) ==> flip ())
(big_opM (@bi_sep PROP) (K:=K) (A:=A)).
Proof. solve_proper. Qed.
Global Instance big_sepM_empty_persistent Φ :
Persistent ([ map] kx , Φ k x).
Proof. rewrite big_opM_empty. apply _. Qed.
Lemma big_sepM_persistent Φ m :
( k x, m !! k = Some x Persistent (Φ k x))
Persistent ([ map] kx m, Φ k x).
Proof. apply big_opM_closed; apply _. Qed.
Global Instance big_sepM_persistent' Φ m :
( k x, Persistent (Φ k x)) Persistent ([ map] kx m, Φ k x).
Proof. intros; apply big_sepM_persistent, _. Qed.
Global Instance big_sepM_empty_affine Φ :
Affine ([ map] kx , Φ k x).
Proof. rewrite big_opM_empty. apply _. Qed.
Lemma big_sepM_affine Φ m :
( k x, m !! k = Some x Affine (Φ k x))
Affine ([ map] kx m, Φ k x).
Proof. apply big_opM_closed; apply _. Qed.
Global Instance big_sepM_affine' Φ m :
( k x, Affine (Φ k x)) Affine ([ map] kx m, Φ k x).
Proof. intros; apply big_sepM_affine, _. Qed.
Global Instance big_sepM_empty_timeless `{!Timeless (emp%I : PROP)} Φ :
Timeless ([ map] kx , Φ k x).
Proof. rewrite big_opM_empty. apply _. Qed.
Lemma big_sepM_timeless `{!Timeless (emp%I : PROP)} Φ m :
( k x, m !! k = Some x Timeless (Φ k x))
Timeless ([ map] kx m, Φ k x).
Proof. apply big_opM_closed; apply _. Qed.
Global Instance big_sepM_timeless' `{!Timeless (emp%I : PROP)} Φ m :
( k x, Timeless (Φ k x)) Timeless ([ map] kx m, Φ k x).
Proof. intros; apply big_sepM_timeless, _. Qed.
(* We place the [Affine] instance after [m1] and [m2], so that
manually instantiating [m1] or [m2] in [iApply] does not also
implicitly instantiate the [Affine] instance. If it gets
instantiated too early, [Φ] is not known, and typeclass inference
fails. *)
Lemma big_sepM_subseteq Φ m1 m2 `{!∀ k x, Affine (Φ k x)} :
m2 m1 ([ map] k x m1, Φ k x) [ map] k x m2, Φ k x.
Proof.
intros ?. rewrite -(map_difference_union m2 m1) //.
rewrite big_opM_union; last by apply map_disjoint_difference_r.
assert ( kx, Affine (uncurry Φ kx)) by (intros []; apply _).
by rewrite sep_elim_l.
Qed.
Lemma big_sepM_empty Φ : ([ map] kx , Φ k x) ⊣⊢ emp.
Proof. by rewrite big_opM_empty. Qed.
Lemma big_sepM_empty' P `{!Affine P} Φ : P [ map] kx , Φ k x.
Proof. rewrite big_sepM_empty. apply: affine. Qed.
Lemma big_sepM_insert Φ m i x :
m !! i = None
([ map] ky <[i:=x]> m, Φ k y) ⊣⊢ Φ i x [ map] ky m, Φ k y.
Proof. apply big_opM_insert. Qed.
Lemma big_sepM_delete Φ m i x :
m !! i = Some x
([ map] ky m, Φ k y) ⊣⊢ Φ i x [ map] ky delete i m, Φ k y.
Proof. apply big_opM_delete. Qed.
Lemma big_sepM_insert_delete Φ m i x :
([ map] ky <[i:=x]> m, Φ k y) ⊣⊢ Φ i x [ map] ky delete i m, Φ k y.
Proof. apply big_opM_insert_delete. Qed.
Lemma big_sepM_insert_2 Φ m i x
`{!TCOr ( y, Affine (Φ i y)) (Absorbing (Φ i x))} :
Φ i x ([ map] ky m, Φ k y) -∗ [ map] ky <[i:=x]> m, Φ k y.
Proof.
apply wand_intro_r. destruct (m !! i) as [y|] eqn:Hi; last first.
{ by rewrite -big_sepM_insert. }
assert (TCOr (Affine (Φ i y)) (Absorbing (Φ i x))).
{ destruct select (TCOr _ _); apply _. }
rewrite big_sepM_delete // assoc.
rewrite (sep_elim_l (Φ i x)) -big_sepM_insert ?lookup_delete //.
by rewrite insert_delete_insert.
Qed.
Lemma big_sepM_lookup_acc Φ m i x :
m !! i = Some x
([ map] ky m, Φ k y) Φ i x (Φ i x -∗ ([ map] ky m, Φ k y)).
Proof.
intros. rewrite big_sepM_delete //. by apply sep_mono_r, wand_intro_l.
Qed.
Lemma big_sepM_lookup Φ m i x
`{!TCOr ( j y, Affine (Φ j y)) (Absorbing (Φ i x))} :
m !! i = Some x ([ map] ky m, Φ k y) Φ i x.
Proof.
intros Hi. destruct select (TCOr _ _).
- rewrite big_sepM_delete // sep_elim_l //.
- rewrite big_sepM_lookup_acc // sep_elim_l //.
Qed.
Lemma big_sepM_lookup_dom (Φ : K PROP) m i
`{!TCOr ( j, Affine (Φ j)) (Absorbing (Φ i))} :
is_Some (m !! i) ([ map] k↦_ m, Φ k) Φ i.
Proof. intros [x ?]. destruct select (TCOr _ _); by apply: big_sepM_lookup. Qed.
Lemma big_sepM_singleton Φ i x : ([ map] ky {[i:=x]}, Φ k y) ⊣⊢ Φ i x.
Proof. by rewrite big_opM_singleton. Qed.
Lemma big_sepM_fmap {B} (f : A B) (Φ : K B PROP) m :
([ map] ky f <$> m, Φ k y) ⊣⊢ ([ map] ky m, Φ k (f y)).
Proof. by rewrite big_opM_fmap. Qed.
Lemma big_sepM_omap {B} (f : A option B) (Φ : K B PROP) m :
([ map] ky omap f m, Φ k y) ⊣⊢ ([ map] ky m, from_option (Φ k) emp (f y)).
Proof. by rewrite big_opM_omap. Qed.
Lemma big_sepM_insert_override Φ m i x x' :
m !! i = Some x (Φ i x ⊣⊢ Φ i x')
([ map] ky <[i:=x']> m, Φ k y) ⊣⊢ ([ map] ky m, Φ k y).
Proof. apply big_opM_insert_override. Qed.
Lemma big_sepM_insert_override_1 Φ m i x x' :
m !! i = Some x
([ map] ky <[i:=x']> m, Φ k y)
(Φ i x' -∗ Φ i x) -∗ ([ map] ky m, Φ k y).
Proof.
intros ?. apply wand_intro_l.
rewrite -insert_delete_insert big_sepM_insert ?lookup_delete //.
by rewrite assoc wand_elim_l -big_sepM_delete.
Qed.
Lemma big_sepM_insert_override_2 Φ m i x x' :
m !! i = Some x
([ map] ky m, Φ k y)
(Φ i x -∗ Φ i x') -∗ ([ map] ky <[i:=x']> m, Φ k y).
Proof.
intros ?. apply wand_intro_l.
rewrite {1}big_sepM_delete //; rewrite assoc wand_elim_l.
rewrite -insert_delete_insert big_sepM_insert ?lookup_delete //.
Qed.
Lemma big_sepM_insert_acc Φ m i x :
m !! i = Some x
([ map] ky m, Φ k y)
Φ i x ( x', Φ i x' -∗ ([ map] ky <[i:=x']> m, Φ k y)).
Proof.
intros ?. rewrite {1}big_sepM_delete //. apply sep_mono; [done|].
apply forall_intro=> x'.
rewrite -insert_delete_insert big_sepM_insert ?lookup_delete //.
by apply wand_intro_l.
Qed.
Lemma big_sepM_fn_insert {B} (Ψ : K A B PROP) (f : K B) m i x b :
m !! i = None
([ map] ky <[i:=x]> m, Ψ k y (<[i:=b]> f k))
⊣⊢ (Ψ i x b [ map] ky m, Ψ k y (f k)).
Proof. apply big_opM_fn_insert. Qed.
Lemma big_sepM_fn_insert' (Φ : K PROP) m i x P :
m !! i = None
([ map] ky <[i:=x]> m, <[i:=P]> Φ k) ⊣⊢ (P [ map] ky m, Φ k).
Proof. apply big_opM_fn_insert'. Qed.
Lemma big_sepM_filter' (φ : K * A Prop) `{ kx, Decision (φ kx)} Φ m :
([ map] k x filter φ m, Φ k x) ⊣⊢
([ map] k x m, if decide (φ (k, x)) then Φ k x else emp).
Proof. apply: big_opM_filter'. Qed.
Lemma big_sepM_filter `{!BiAffine PROP}
(φ : K * A Prop) `{ kx, Decision (φ kx)} Φ m :
([ map] k x filter φ m, Φ k x) ⊣⊢
([ map] k x m, φ (k, x) Φ k x).
Proof. setoid_rewrite <-decide_emp. apply big_sepM_filter'. Qed.
Lemma big_sepM_union Φ m1 m2 :
m1 ## m2
([ map] ky m1 m2, Φ k y)
⊣⊢ ([ map] ky m1, Φ k y) ([ map] ky m2, Φ k y).
Proof. apply big_opM_union. Qed.
Lemma big_sepM_sep Φ Ψ m :
([ map] kx m, Φ k x Ψ k x)
⊣⊢ ([ map] kx m, Φ k x) ([ map] kx m, Ψ k x).
Proof. apply big_opM_op. Qed.
Lemma big_sepM_sep_2 Φ Ψ m :
([ map] kx m, Φ k x) -∗
([ map] kx m, Ψ k x) -∗
([ map] kx m, Φ k x Ψ k x).
Proof. apply entails_wand, wand_intro_r. rewrite big_sepM_sep //. Qed.
Lemma big_sepM_and Φ Ψ m :
([ map] kx m, Φ k x Ψ k x)
([ map] kx m, Φ k x) ([ map] kx m, Ψ k x).
Proof. auto using and_intro, big_sepM_mono, and_elim_l, and_elim_r. Qed.
Lemma big_sepM_pure_1 (φ : K A Prop) m :
([ map] kx m, φ k x) ⊢@{PROP} map_Forall φ m⌝.
Proof.
induction m as [|k x m ? IH] using map_ind.
{ apply pure_intro, map_Forall_empty. }
rewrite big_sepM_insert // IH sep_and -pure_and.
by rewrite -map_Forall_insert.
Qed.
Lemma big_sepM_affinely_pure_2 (φ : K A Prop) m :
<affine> map_Forall φ m ⊢@{PROP} ([ map] kx m, <affine> φ k x).
Proof.
induction m as [|k x m ? IH] using map_ind.
{ rewrite big_sepM_empty. apply affinely_elim_emp. }
rewrite big_sepM_insert // -IH.
by rewrite -persistent_and_sep_1 -affinely_and -pure_and map_Forall_insert.
Qed.
(** The general backwards direction requires [BiAffine] to cover the empty case. *)
Lemma big_sepM_pure `{!BiAffine PROP} (φ : K A Prop) m :
([ map] kx m, φ k x) ⊣⊢@{PROP} map_Forall φ m⌝.
Proof.
apply (anti_symm ()); first by apply big_sepM_pure_1.
rewrite -(affine_affinely ⌜_⌝).
rewrite big_sepM_affinely_pure_2. by setoid_rewrite affinely_elim.
Qed.
Lemma big_sepM_persistently `{!BiAffine PROP} Φ m :
(<pers> ([ map] kx m, Φ k x)) ⊣⊢ ([ map] kx m, <pers> (Φ k x)).
Proof. apply (big_opM_commute _). Qed.
Lemma big_sepM_intro Φ m :
( k x, m !! k = Some x Φ k x) [ map] kx m, Φ k x.
Proof.
revert Φ. induction m as [|i x m ? IH] using map_ind=> Φ.
{ by rewrite (affine ( _)) big_sepM_empty. }
rewrite big_sepM_insert // intuitionistically_sep_dup. f_equiv.
- rewrite (forall_elim i) (forall_elim x) lookup_insert.
by rewrite pure_True // True_impl intuitionistically_elim.
- rewrite -IH. f_equiv. apply forall_mono=> k; apply forall_mono=> y.
apply impl_intro_l, pure_elim_l=> ?.
rewrite lookup_insert_ne; last by intros ?; simplify_map_eq.
by rewrite pure_True // True_impl.
Qed.
Lemma big_sepM_forall `{!BiAffine PROP} Φ m :
( k x, Persistent (Φ k x))
([ map] kx m, Φ k x) ⊣⊢ ( k x, m !! k = Some x Φ k x).
Proof.
intros . apply (anti_symm _).
{ apply forall_intro=> k; apply forall_intro=> x.
apply impl_intro_l, pure_elim_l=> ?; by apply: big_sepM_lookup. }
revert Φ . induction m as [|i x m ? IH] using map_ind=> Φ .
{ rewrite big_sepM_empty. apply: affine. }
rewrite big_sepM_insert // -persistent_and_sep_1. apply and_intro.
- rewrite (forall_elim i) (forall_elim x) lookup_insert.
by rewrite pure_True // True_impl.
- rewrite -IH. apply forall_mono=> k; apply forall_mono=> y.
apply impl_intro_l, pure_elim_l=> ?.
rewrite lookup_insert_ne; last by intros ?; simplify_map_eq.
by rewrite pure_True // True_impl.
Qed.
Lemma big_sepM_impl Φ Ψ m :
([ map] kx m, Φ k x) -∗
( k x, m !! k = Some x Φ k x -∗ Ψ k x) -∗
[ map] kx m, Ψ k x.
Proof.
apply entails_wand, wand_intro_l. rewrite big_sepM_intro -big_sepM_sep.
by setoid_rewrite wand_elim_l.
Qed.
Lemma big_sepM_wand Φ Ψ m :
([ map] kx m, Φ k x) -∗
([ map] kx m, Φ k x -∗ Ψ k x) -∗
[ map] kx m, Ψ k x.
Proof.
apply entails_wand, wand_intro_r. rewrite -big_sepM_sep.
setoid_rewrite wand_elim_r. done.
Qed.
Lemma big_sepM_dup P `{!Affine P} m :
(P -∗ P P) -∗ P -∗ [ map] kx m, P.
Proof.
apply entails_wand, wand_intro_l. induction m as [|i x m ? IH] using map_ind.
{ apply: big_sepM_empty'. }
rewrite !big_sepM_insert //.
rewrite intuitionistically_sep_dup {1}intuitionistically_elim.
rewrite assoc wand_elim_r -assoc. apply sep_mono; done.
Qed.
Lemma big_sepM_lookup_acc_impl {Φ m} i x :
m !! i = Some x
([ map] ky m, Φ k y) -∗
(* We obtain [Φ] for [x] *)
Φ i x
(* We reobtain the bigop for a predicate [Ψ] selected by the user *)
Ψ,
( k y, m !! k = Some y k i Φ k y -∗ Ψ k y) -∗
Ψ i x -∗
[ map] ky m, Ψ k y.
Proof.
intros. rewrite big_sepM_delete //. apply entails_wand, sep_mono_r, forall_intro=> Ψ.
apply wand_intro_r, wand_intro_l.
rewrite (big_sepM_delete Ψ m i x) //. apply sep_mono_r.
eapply wand_apply; [apply wand_entails, big_sepM_impl|apply sep_mono_r].
f_equiv; f_equiv=> k; f_equiv=> y.
rewrite impl_curry -pure_and lookup_delete_Some.
do 2 f_equiv. intros ?; naive_solver.
Qed.
Lemma big_sepM_later `{!BiAffine PROP} Φ m :
([ map] kx m, Φ k x) ⊣⊢ ([ map] kx m, Φ k x).
Proof. apply (big_opM_commute _). Qed.
Lemma big_sepM_later_2 Φ m :
([ map] kx m, Φ k x) [ map] kx m, Φ k x.
Proof. by rewrite big_opM_commute. Qed.
Lemma big_sepM_laterN `{!BiAffine PROP} Φ n m :
▷^n ([ map] kx m, Φ k x) ⊣⊢ ([ map] kx m, ▷^n Φ k x).
Proof. apply (big_opM_commute _). Qed.
Lemma big_sepM_laterN_2 Φ n m :
([ map] kx m, ▷^n Φ k x) ▷^n [ map] kx m, Φ k x.
Proof. by rewrite big_opM_commute. Qed.
Lemma big_sepM_map_to_list Φ m :
([ map] kx m, Φ k x) ⊣⊢ [ list] xk map_to_list m, Φ (xk.1) (xk.2).
Proof. apply big_opM_map_to_list. Qed.
Lemma big_sepM_list_to_map Φ l :
NoDup l.*1
([ map] kx list_to_map l, Φ k x) ⊣⊢ [ list] xk l, Φ (xk.1) (xk.2).
Proof. apply big_opM_list_to_map. Qed.
End sep_map.
(* Some lemmas depend on the generalized versions of the above ones. *)
Lemma big_sepM_sep_zip_with `{Countable K} {A B C}
(f : A B C) (g1 : C A) (g2 : C B)
(Φ1 : K A PROP) (Φ2 : K B PROP) m1 m2 :
( x y, g1 (f x y) = x)
( x y, g2 (f x y) = y)
( k, is_Some (m1 !! k) is_Some (m2 !! k))
([ map] kxy map_zip_with f m1 m2, Φ1 k (g1 xy) Φ2 k (g2 xy)) ⊣⊢
([ map] kx m1, Φ1 k x) ([ map] ky m2, Φ2 k y).
Proof. apply big_opM_sep_zip_with. Qed.
Lemma big_sepM_sep_zip `{Countable K} {A B}
(Φ1 : K A PROP) (Φ2 : K B PROP) m1 m2 :
( k, is_Some (m1 !! k) is_Some (m2 !! k))
([ map] kxy map_zip m1 m2, Φ1 k xy.1 Φ2 k xy.2) ⊣⊢
([ map] kx m1, Φ1 k x) ([ map] ky m2, Φ2 k y).
Proof. apply big_opM_sep_zip. Qed.
Lemma big_sepM_impl_strong `{Countable K} {A B}
(Φ : K A PROP) (Ψ : K B PROP) (m1 : gmap K A) (m2 : gmap K B) :
([ map] kx m1, Φ k x)
( (k : K) (y : B),
(if m1 !! k is Some x then Φ k x else emp) -∗
m2 !! k = Some y
Ψ k y) -∗
([ map] ky m2, Ψ k y)
([ map] kx filter (λ '(k, _), m2 !! k = None) m1, Φ k x).
Proof.
apply wand_intro_r.
revert m1. induction m2 as [|i y m ? IH] using map_ind=> m1.
{ rewrite big_sepM_empty left_id sep_elim_l.
rewrite map_filter_id //. }
rewrite big_sepM_insert; last done.
rewrite intuitionistically_sep_dup.
rewrite assoc. rewrite (comm _ _ ( _))%I.
rewrite {1}intuitionistically_elim {1}(forall_elim i) {1}(forall_elim y).
rewrite lookup_insert pure_True // left_id.
destruct (m1 !! i) as [x|] eqn:Hx.
- rewrite big_sepM_delete; last done.
rewrite assoc assoc wand_elim_l -2!assoc. apply sep_mono_r.
assert (filter (λ '(k, _), <[i:=y]> m !! k = None) m1
= filter (λ '(k, _), m !! k = None) (delete i m1)) as ->.
{ apply map_filter_strong_ext_1=> k z.
rewrite lookup_insert_None lookup_delete_Some. naive_solver. }
rewrite -IH. do 2 f_equiv. f_equiv=> k. f_equiv=> z.
apply wand_intro_r. apply impl_intro_l, pure_elim_l=> ?.
assert (i k) by congruence.
rewrite lookup_insert_ne // pure_True // left_id.
rewrite lookup_delete_ne // wand_elim_l //.
- rewrite left_id -assoc. apply sep_mono_r.
assert (filter (λ '(k, _), <[i:=y]> m !! k = None) m1
= filter (λ '(k, _), m !! k = None) m1) as ->.
{ apply map_filter_strong_ext_1=> k z.
rewrite lookup_insert_None. naive_solver. }
rewrite -IH. do 2 f_equiv. f_equiv=> k. f_equiv=> z.
apply wand_intro_r. apply impl_intro_l, pure_elim_l=> ?.
rewrite lookup_insert_ne; last congruence.
rewrite pure_True // left_id // wand_elim_l //.
Qed.
Lemma big_sepM_impl_dom_subseteq `{Countable K} {A B}
(Φ : K A PROP) (Ψ : K B PROP) (m1 : gmap K A) (m2 : gmap K B) :
dom m2 dom m1
([ map] kx m1, Φ k x) -∗
( (k : K) (x : A) (y : B),
m1 !! k = Some x m2 !! k = Some y
Φ k x -∗ Ψ k y) -∗
([ map] ky m2, Ψ k y)
([ map] kx filter (λ '(k, _), m2 !! k = None) m1, Φ k x).
Proof.
intros. apply entails_wand. rewrite big_sepM_impl_strong.
apply wand_mono; last done. f_equiv. f_equiv=> k. apply forall_intro=> y.
apply wand_intro_r, impl_intro_l, pure_elim_l=> Hlm2.
destruct (m1 !! k) as [x|] eqn:Hlm1.
- rewrite (forall_elim x) (forall_elim y).
do 2 rewrite pure_True // left_id //. apply wand_elim_l.
- apply elem_of_dom_2 in Hlm2. apply not_elem_of_dom in Hlm1.
set_solver.
Qed.
Section and_map.
Context `{Countable K} {A : Type}.
Implicit Types m : gmap K A.
Implicit Types Φ Ψ : K A PROP.
(** The lemmas [big_andM_mono], [big_andM_ne] and [big_andM_proper] are more
generic than the instances as they also give [l !! k = Some y] in the premise. *)
Lemma big_andM_mono Φ Ψ m :
( k x, m !! k = Some x Φ k x Ψ k x)
([ map] k x m, Φ k x) [ map] k x m, Ψ k x.
Proof. apply big_opM_gen_proper; apply _ || auto. Qed.
Lemma big_andM_ne Φ Ψ m n :
( k x, m !! k = Some x Φ k x {n} Ψ k x)
([ map] k x m, Φ k x)%I {n} ([ map] k x m, Ψ k x)%I.
Proof. apply big_opM_ne. Qed.
Lemma big_andM_proper Φ Ψ m :
( k x, m !! k = Some x Φ k x ⊣⊢ Ψ k x)
([ map] k x m, Φ k x) ⊣⊢ ([ map] k x m, Ψ k x).
Proof. apply big_opM_proper. Qed.
(** No need to declare instances for non-expansiveness and properness, we
get both from the generic [big_opM] instances. *)
Global Instance big_andM_mono' :
Proper (pointwise_relation _ (pointwise_relation _ ()) ==> (=) ==> ())
(big_opM (@bi_and PROP) (K:=K) (A:=A)).
Proof. intros f g Hf m ? <-. apply big_andM_mono; intros; apply Hf. Qed.
Global Instance big_andM_empty_persistent Φ :
Persistent ([ map] kx , Φ k x).
Proof. rewrite big_opM_empty. apply _. Qed.
Lemma big_andM_persistent Φ m :
( k x, m !! k = Some x Persistent (Φ k x))
Persistent ([ map] kx m, Φ k x).
Proof. apply big_opM_closed; apply _. Qed.
Global Instance big_andM_persistent' Φ m :
( k x, Persistent (Φ k x)) Persistent ([ map] kx m, Φ k x).
Proof. intros; apply big_andM_persistent, _. Qed.
Global Instance big_andM_empty_timeless Φ :
Timeless ([ map] kx , Φ k x).
Proof. rewrite big_opM_empty. apply _. Qed.
Lemma big_andM_timeless Φ m :
( k x, m !! k = Some x Timeless (Φ k x))
Timeless ([ map] kx m, Φ k x).
Proof. apply big_opM_closed; apply _. Qed.
Global Instance big_andM_timeless' Φ m :
( k x, Timeless (Φ k x)) Timeless ([ map] kx m, Φ k x).
Proof. intros; apply big_andM_timeless, _. Qed.
Lemma big_andM_subseteq Φ m1 m2 :
m2 m1 ([ map] k x m1, Φ k x) [ map] k x m2, Φ k x.
Proof.
intros ?. rewrite -(map_difference_union m2 m1) //.
rewrite big_opM_union; last by apply map_disjoint_difference_r.
by rewrite and_elim_l.
Qed.
Lemma big_andM_empty Φ : ([ map] kx , Φ k x) ⊣⊢ True.
Proof. by rewrite big_opM_empty. Qed.
Lemma big_andM_empty' P Φ : P [ map] kx , Φ k x.
Proof. rewrite big_andM_empty. apply: True_intro. Qed.
Lemma big_andM_insert Φ m i x :
m !! i = None
([ map] ky <[i:=x]> m, Φ k y) ⊣⊢ Φ i x [ map] ky m, Φ k y.
Proof. apply big_opM_insert. Qed.
Lemma big_andM_delete Φ m i x :
m !! i = Some x
([ map] ky m, Φ k y) ⊣⊢ Φ i x [ map] ky delete i m, Φ k y.
Proof. apply big_opM_delete. Qed.
Lemma big_andM_insert_delete Φ m i x :
([ map] ky <[i:=x]> m, Φ k y) ⊣⊢ Φ i x [ map] ky delete i m, Φ k y.
Proof. apply big_opM_insert_delete. Qed.
Lemma big_andM_insert_2 Φ m i x :
Φ i x ([ map] ky m, Φ k y) [ map] ky <[i:=x]> m, Φ k y.
Proof.
rewrite big_andM_insert_delete.
destruct (m !! i) eqn:Hi; [ | by rewrite delete_notin ].
rewrite big_andM_delete //. apply and_mono_r, and_elim_r.
Qed.
Lemma big_andM_lookup Φ m i x :
m !! i = Some x ([ map] ky m, Φ k y) Φ i x.
Proof.
intros. rewrite -(insert_id m i x) // big_andM_insert_delete. apply and_elim_l.
Qed.
Lemma big_andM_lookup_dom (Φ : K PROP) m i :
is_Some (m !! i) ([ map] k↦_ m, Φ k) Φ i.
Proof. intros [x ?]. by eapply (big_andM_lookup (λ i x, Φ i)). Qed.
Lemma big_andM_singleton Φ i x : ([ map] ky {[i:=x]}, Φ k y) ⊣⊢ Φ i x.
Proof. by rewrite big_opM_singleton. Qed.
Lemma big_andM_fmap {B} (f : A B) (Φ : K B PROP) m :
([ map] ky f <$> m, Φ k y) ⊣⊢ ([ map] ky m, Φ k (f y)).
Proof. by rewrite big_opM_fmap. Qed.
Lemma big_andM_omap {B} (f : A option B) (Φ : K B PROP) m :
([ map] ky omap f m, Φ k y) ⊣⊢ ([ map] ky m, from_option (Φ k) True (f y)).
Proof. by rewrite big_opM_omap. Qed.
Lemma big_andM_fn_insert {B} (Ψ : K A B PROP) (f : K B) m i x b :
m !! i = None
([ map] ky <[i:=x]> m, Ψ k y (<[i:=b]> f k))
⊣⊢ (Ψ i x b [ map] ky m, Ψ k y (f k)).
Proof. apply big_opM_fn_insert. Qed.
Lemma big_andM_fn_insert' (Φ : K PROP) m i x P :
m !! i = None
([ map] ky <[i:=x]> m, <[i:=P]> Φ k) ⊣⊢ (P [ map] ky m, Φ k).
Proof. apply big_opM_fn_insert'. Qed.
Lemma big_andM_filter' (φ : K * A Prop) `{ kx, Decision (φ kx)} Φ m :
([ map] k x filter φ m, Φ k x) ⊣⊢
([ map] k x m, if decide (φ (k, x)) then Φ k x else True).
Proof. apply: big_opM_filter'. Qed.
Lemma big_andM_filter (φ : K * A Prop) `{ kx, Decision (φ kx)} Φ m :
([ map] k x filter φ m, Φ k x) ⊣⊢
([ map] k x m, φ (k, x) Φ k x).
Proof. setoid_rewrite <-decide_bi_True. apply big_andM_filter'. Qed.
Lemma big_andM_union Φ m1 m2 :
m1 ## m2
([ map] ky m1 m2, Φ k y)
⊣⊢ ([ map] ky m1, Φ k y) ([ map] ky m2, Φ k y).
Proof. apply big_opM_union. Qed.
Lemma big_andM_and Φ Ψ m :
([ map] kx m, Φ k x Ψ k x)
⊣⊢ ([ map] kx m, Φ k x) ([ map] kx m, Ψ k x).
Proof. apply big_opM_op. Qed.
Lemma big_andM_persistently Φ m :
<pers> ([ map] kx m, Φ k x) ⊣⊢ ([ map] kx m, <pers> (Φ k x)).
Proof. apply (big_opM_commute _). Qed.
Lemma big_andM_intro Φ m :
( k x, m !! k = Some x Φ k x) [ map] kx m, Φ k x.
Proof.
revert Φ. induction m as [|i x m ? IH] using map_ind=> Φ.
{ rewrite big_andM_empty. apply: True_intro. }
rewrite big_andM_insert //. apply and_intro.
- rewrite (forall_elim i) (forall_elim x) lookup_insert.
by rewrite pure_True // True_impl.
- rewrite -IH. apply forall_intro=> k. apply forall_intro=> x'.
rewrite (forall_elim k) (forall_elim x').
apply impl_intro_l, pure_elim_l=> ?.
rewrite lookup_insert_ne; last by intros ?; simplify_map_eq.
by rewrite pure_True // True_impl.
Qed.
Lemma big_andM_forall Φ m :
([ map] kx m, Φ k x) ⊣⊢ ( k x, m !! k = Some x Φ k x).
Proof.
intros. apply (anti_symm _); [| by rewrite -big_andM_intro].
apply forall_intro=> k; apply forall_intro=> x.
apply impl_intro_l, pure_elim_l=> ?; by apply: big_andM_lookup.
Qed.
Lemma big_andM_impl Φ Ψ m :
([ map] kx m, Φ k x)
( k x, m !! k = Some x Φ k x Ψ k x)
[ map] kx m, Ψ k x.
Proof.
rewrite -big_andM_forall -big_andM_and.
by setoid_rewrite bi.impl_elim_r.
Qed.
Lemma big_andM_pure_1 (φ : K A Prop) m :
([ map] kx m, φ k x) ⊢@{PROP} map_Forall φ m⌝.
Proof.
induction m as [|k x m ? IH] using map_ind.
{ apply pure_intro, map_Forall_empty. }
rewrite big_andM_insert // IH -pure_and.
by rewrite -map_Forall_insert.
Qed.
Lemma big_andM_pure_2 (φ : K A Prop) m :
map_Forall φ m ⊢@{PROP} ([ map] kx m, φ k x).
Proof.
rewrite big_andM_forall pure_forall_1. f_equiv=>k.
rewrite pure_forall_1. f_equiv=>x. apply pure_impl_1.
Qed.
Lemma big_andM_pure (φ : K A Prop) m :
([ map] kx m, φ k x) ⊣⊢@{PROP} map_Forall φ m⌝.
Proof.
apply (anti_symm ()); [ by apply big_andM_pure_1 | by apply big_andM_pure_2].
Qed.
Lemma big_andM_later Φ m :
([ map] kx m, Φ k x) ⊣⊢ ([ map] kx m, Φ k x).
Proof. apply (big_opM_commute _). Qed.
Lemma big_andM_laterN Φ n m :
▷^n ([ map] kx m, Φ k x) ⊣⊢ ([ map] kx m, ▷^n Φ k x).
Proof. apply (big_opM_commute _). Qed.
End and_map.
(** ** Big ops over two maps *)
Lemma big_sepM2_alt `{Countable K} {A B} (Φ : K A B PROP) m1 m2 :
([ map] ky1;y2 m1; m2, Φ k y1 y2) ⊣⊢
dom m1 = dom m2 [ map] k xy map_zip m1 m2, Φ k xy.1 xy.2.
Proof. by rewrite big_sepM2_unseal. Qed.
Section map2.
Context `{Countable K} {A B : Type}.
Implicit Types Φ Ψ : K A B PROP.
Lemma big_sepM2_alt_lookup (Φ : K A B PROP) m1 m2 :
([ map] ky1;y2 m1; m2, Φ k y1 y2) ⊣⊢
k, is_Some (m1 !! k) is_Some (m2 !! k)
[ map] k xy map_zip m1 m2, Φ k xy.1 xy.2.
Proof. rewrite big_sepM2_alt set_eq. by setoid_rewrite elem_of_dom. Qed.
Lemma big_sepM2_lookup_iff Φ m1 m2 :
([ map] ky1;y2 m1; m2, Φ k y1 y2)
k, is_Some (m1 !! k) is_Some (m2 !! k) ⌝.
Proof. by rewrite big_sepM2_alt_lookup and_elim_l. Qed.
Lemma big_sepM2_dom Φ m1 m2 :
([ map] ky1;y2 m1; m2, Φ k y1 y2)
dom m1 = dom m2 ⌝.
Proof. by rewrite big_sepM2_alt and_elim_l. Qed.
Lemma big_sepM2_flip Φ m1 m2 :
([ map] ky1;y2 m2; m1, Φ k y2 y1) ⊣⊢ ([ map] ky1;y2 m1; m2, Φ k y1 y2).
Proof.
rewrite !big_sepM2_alt. apply and_proper; [apply pure_proper; naive_solver |].
rewrite -map_zip_with_flip map_zip_with_map_zip big_sepM_fmap.
apply big_sepM_proper. by intros k [b a].
Qed.
Lemma big_sepM2_empty Φ : ([ map] ky1;y2 ; , Φ k y1 y2) ⊣⊢ emp.
Proof.
rewrite big_sepM2_alt map_zip_with_empty big_sepM_empty pure_True ?left_id //.
Qed.
Lemma big_sepM2_empty' P `{!Affine P} Φ : P [ map] ky1;y2 ;, Φ k y1 y2.
Proof. rewrite big_sepM2_empty. apply: affine. Qed.
Lemma big_sepM2_empty_l m1 Φ : ([ map] ky1;y2 m1; , Φ k y1 y2) m1 = ∅⌝.
Proof.
rewrite big_sepM2_dom dom_empty_L.
apply pure_mono, dom_empty_iff_L.
Qed.
Lemma big_sepM2_empty_r m2 Φ : ([ map] ky1;y2 ; m2, Φ k y1 y2) m2 = ∅⌝.
Proof.
rewrite big_sepM2_dom dom_empty_L.
apply pure_mono=>?. eapply dom_empty_iff_L. eauto.
Qed.
Lemma big_sepM2_insert Φ m1 m2 i x1 x2 :
m1 !! i = None m2 !! i = None
([ map] ky1;y2 <[i:=x1]>m1; <[i:=x2]>m2, Φ k y1 y2)
⊣⊢ Φ i x1 x2 [ map] ky1;y2 m1;m2, Φ k y1 y2.
Proof.
intros Hm1 Hm2. rewrite !big_sepM2_alt -map_insert_zip_with.
rewrite big_sepM_insert;
last by rewrite map_lookup_zip_with Hm1.
rewrite !persistent_and_affinely_sep_l /=.
rewrite sep_assoc (sep_comm _ (Φ _ _ _)) -sep_assoc.
repeat apply sep_proper=>//.
apply affinely_proper, pure_proper. rewrite !dom_insert_L.
apply not_elem_of_dom in Hm1. apply not_elem_of_dom in Hm2. set_solver.
Qed.
(** The lemmas [big_sepM2_mono], [big_sepM2_ne] and [big_sepM2_proper] are more
generic than the instances as they also give [mi !! k = Some yi] in the premise. *)
Lemma big_sepM2_mono Φ Ψ m1 m2 :
( k y1 y2, m1 !! k = Some y1 m2 !! k = Some y2 Φ k y1 y2 Ψ k y1 y2)
([ map] k y1;y2 m1;m2, Φ k y1 y2) [ map] k y1;y2 m1;m2, Ψ k y1 y2.
Proof.
intros Hm1m2. rewrite !big_sepM2_alt.
apply and_mono_r, big_sepM_mono.
intros k [x1 x2]. rewrite map_lookup_zip_with.
specialize (Hm1m2 k x1 x2).
destruct (m1 !! k) as [y1|]; last done.
destruct (m2 !! k) as [y2|]; simpl; last done.
intros ?; simplify_eq. by apply Hm1m2.
Qed.
Lemma big_sepM2_ne Φ Ψ m1 m2 n :
( k y1 y2, m1 !! k = Some y1 m2 !! k = Some y2 Φ k y1 y2 {n} Ψ k y1 y2)
([ map] k y1;y2 m1;m2, Φ k y1 y2)%I {n} ([ map] k y1;y2 m1;m2, Ψ k y1 y2)%I.
Proof.
intros Hm1m2. rewrite !big_sepM2_alt.
f_equiv. apply big_sepM_ne=> k [x1 x2]. rewrite map_lookup_zip_with.
specialize (Hm1m2 k x1 x2).
destruct (m1 !! k) as [y1|]; last done.
destruct (m2 !! k) as [y2|]; simpl; last done.
intros ?; simplify_eq. by apply Hm1m2.
Qed.
Lemma big_sepM2_proper Φ Ψ m1 m2 :
( k y1 y2, m1 !! k = Some y1 m2 !! k = Some y2 Φ k y1 y2 ⊣⊢ Ψ k y1 y2)
([ map] k y1;y2 m1;m2, Φ k y1 y2) ⊣⊢ [ map] k y1;y2 m1;m2, Ψ k y1 y2.
Proof.
intros; apply (anti_symm _);
apply big_sepM2_mono; auto using equiv_entails_1_1, equiv_entails_1_2.
Qed.
Lemma big_sepM2_proper_2 `{!Equiv A, !Equiv B} Φ Ψ m1 m2 m1' m2' :
m1 m1' m2 m2'
( k y1 y1' y2 y2',
m1 !! k = Some y1 m1' !! k = Some y1' y1 y1'
m2 !! k = Some y2 m2' !! k = Some y2' y2 y2'
Φ k y1 y2 ⊣⊢ Ψ k y1' y2')
([ map] k y1;y2 m1;m2, Φ k y1 y2) ⊣⊢ [ map] k y1;y2 m1';m2', Ψ k y1 y2.
Proof.
intros Hm1 Hm2 Hf. rewrite !big_sepM2_alt. f_equiv.
{ by rewrite Hm1 Hm2. }
apply big_opM_proper_2; [by f_equiv|].
intros k [x1 y1] [x2 y2] (?&?&[=<- <-]&?&?)%map_lookup_zip_with_Some
(?&?&[=<- <-]&?&?)%map_lookup_zip_with_Some [??]; naive_solver.
Qed.
Global Instance big_sepM2_ne' n :
Proper (pointwise_relation _ (pointwise_relation _ (pointwise_relation _ (dist n)))
==> (=) ==> (=) ==> (dist n))
(big_sepM2 (PROP:=PROP) (K:=K) (A:=A) (B:=B)).
Proof. intros f g Hf m1 ? <- m2 ? <-. apply big_sepM2_ne; intros; apply Hf. Qed.
Global Instance big_sepM2_mono' :
Proper (pointwise_relation _ (pointwise_relation _ (pointwise_relation _ ()))
==> (=) ==> (=) ==> ()) (big_sepM2 (PROP:=PROP) (K:=K) (A:=A) (B:=B)).
Proof. intros f g Hf m1 ? <- m2 ? <-. apply big_sepM2_mono; intros; apply Hf. Qed.
Global Instance big_sepM2_flip_mono' :
Proper (pointwise_relation _ (pointwise_relation _ (pointwise_relation _ (flip ())))
==> (=) ==> (=) ==> flip ())
(big_sepM2 (PROP:=PROP) (K:=K) (A:=A) (B:=B)).
Proof. solve_proper. Qed.
Global Instance big_sepM2_proper' :
Proper (pointwise_relation _ (pointwise_relation _ (pointwise_relation _ (⊣⊢)))
==> (=) ==> (=) ==> (⊣⊢))
(big_sepM2 (PROP:=PROP) (K:=K) (A:=A) (B:=B)).
Proof. intros f g Hf m1 ? <- m2 ? <-. apply big_sepM2_proper; intros; apply Hf. Qed.
(** Shows that some property [P] is closed under [big_sepM2]. Examples of [P]
are [Persistent], [Affine], [Timeless]. *)
Lemma big_sepM2_closed (P : PROP Prop) Φ m1 m2 :
Proper (() ==> iff) P
P emp%I P False%I
( Q1 Q2, P Q1 P Q2 P (Q1 Q2)%I)
( k x1 x2, m1 !! k = Some x1 m2 !! k = Some x2 P (Φ k x1 x2))
P ([ map] kx1;x2 m1; m2, Φ k x1 x2)%I.
Proof.
intros ??? Hsep .
rewrite big_sepM2_alt. destruct (decide (dom m1 = dom m2)).
- rewrite pure_True // left_id. apply big_opM_closed; [done..|].
intros k [x1 x2] Hk. rewrite map_lookup_zip_with in Hk.
simplify_option_eq; auto.
- by rewrite pure_False // left_absorb.
Qed.
Global Instance big_sepM2_empty_persistent Φ :
Persistent ([ map] ky1;y2 ; , Φ k y1 y2).
Proof. rewrite big_sepM2_empty. apply _. Qed.
Lemma big_sepM2_persistent Φ m1 m2 :
( k x1 x2, m1 !! k = Some x1 m2 !! k = Some x2 Persistent (Φ k x1 x2))
Persistent ([ map] ky1;y2 m1;m2, Φ k y1 y2).
Proof. apply big_sepM2_closed; apply _. Qed.
Global Instance big_sepM2_persistent' Φ m1 m2 :
( k x1 x2, Persistent (Φ k x1 x2))
Persistent ([ map] ky1;y2 m1;m2, Φ k y1 y2).
Proof. intros; apply big_sepM2_persistent, _. Qed.
Global Instance big_sepM2_empty_affine Φ :
Affine ([ map] ky1;y2 ; , Φ k y1 y2).
Proof. rewrite big_sepM2_empty. apply _. Qed.
Lemma big_sepM2_affine Φ m1 m2 :
( k x1 x2, m1 !! k = Some x1 m2 !! k = Some x2 Affine (Φ k x1 x2))
Affine ([ map] ky1;y2 m1;m2, Φ k y1 y2).
Proof. apply big_sepM2_closed; apply _. Qed.
Global Instance big_sepM2_affine' Φ m1 m2 :
( k x1 x2, Affine (Φ k x1 x2))
Affine ([ map] ky1;y2 m1;m2, Φ k y1 y2).
Proof. intros; apply big_sepM2_affine, _. Qed.
Global Instance big_sepM2_empty_timeless `{!Timeless (emp%I : PROP)} Φ :
Timeless ([ map] kx1;x2 ;, Φ k x1 x2).
Proof. rewrite big_sepM2_alt map_zip_with_empty. apply _. Qed.
Lemma big_sepM2_timeless `{!Timeless (emp%I : PROP)} Φ m1 m2 :
( k x1 x2, m1 !! k = Some x1 m2 !! k = Some x2 Timeless (Φ k x1 x2))
Timeless ([ map] kx1;x2 m1;m2, Φ k x1 x2).
Proof. apply big_sepM2_closed; apply _. Qed.
Global Instance big_sepM2_timeless' `{!Timeless (emp%I : PROP)} Φ m1 m2 :
( k x1 x2, Timeless (Φ k x1 x2))
Timeless ([ map] kx1;x2 m1;m2, Φ k x1 x2).
Proof. intros; apply big_sepM2_timeless, _. Qed.
Lemma big_sepM2_delete Φ m1 m2 i x1 x2 :
m1 !! i = Some x1 m2 !! i = Some x2
([ map] kx;y m1;m2, Φ k x y) ⊣⊢
Φ i x1 x2 [ map] kx;y delete i m1;delete i m2, Φ k x y.
Proof.
rewrite !big_sepM2_alt=> Hx1 Hx2.
rewrite !persistent_and_affinely_sep_l /=.
rewrite sep_assoc (sep_comm (Φ _ _ _)) -sep_assoc.
apply sep_proper.
- apply affinely_proper, pure_proper. rewrite !dom_delete_L.
apply elem_of_dom_2 in Hx1; apply elem_of_dom_2 in Hx2. set_unfold.
apply base.forall_proper=> k. destruct (decide (k = i)); naive_solver.
- rewrite -map_delete_zip_with.
apply (big_sepM_delete (λ i xx, Φ i xx.1 xx.2) (map_zip m1 m2) i (x1,x2)).
by rewrite map_lookup_zip_with Hx1 Hx2.
Qed.
Lemma big_sepM2_delete_l Φ m1 m2 i x1 :
m1 !! i = Some x1
([ map] ky1;y2 m1;m2, Φ k y1 y2)
⊣⊢ x2, m2 !! i = Some x2
(Φ i x1 x2 [ map] ky1;y2 delete i m1;delete i m2, Φ k y1 y2).
Proof.
intros Hm1. apply (anti_symm _).
- rewrite big_sepM2_alt_lookup. apply pure_elim_l=> Hm.
assert (is_Some (m2 !! i)) as [x2 Hm2].
{ apply Hm. by econstructor. }
rewrite -(exist_intro x2).
rewrite big_sepM2_alt_lookup (big_sepM_delete _ _ i (x1,x2)) /=;
last by rewrite map_lookup_zip_with Hm1 Hm2.
rewrite pure_True // left_id.
f_equiv. apply and_intro.
+ apply pure_intro=> k. by rewrite !lookup_delete_is_Some Hm.
+ by rewrite -map_delete_zip_with.
- apply exist_elim=> x2. apply pure_elim_l=> ?.
by rewrite -big_sepM2_delete.
Qed.
Lemma big_sepM2_delete_r Φ m1 m2 i x2 :
m2 !! i = Some x2
([ map] ky1;y2 m1;m2, Φ k y1 y2)
⊣⊢ x1, m1 !! i = Some x1
(Φ i x1 x2 [ map] ky1;y2 delete i m1;delete i m2, Φ k y1 y2).
Proof.
intros Hm2. apply (anti_symm _).
- rewrite big_sepM2_alt_lookup.
apply pure_elim_l=> Hm.
assert (is_Some (m1 !! i)) as [x1 Hm1].
{ apply Hm. by econstructor. }
rewrite -(exist_intro x1).
rewrite big_sepM2_alt_lookup (big_sepM_delete _ _ i (x1,x2)) /=;
last by rewrite map_lookup_zip_with Hm1 Hm2.
rewrite pure_True // left_id.
f_equiv. apply and_intro.
+ apply pure_intro=> k. by rewrite !lookup_delete_is_Some Hm.
+ by rewrite -map_delete_zip_with.
- apply exist_elim=> x1. apply pure_elim_l=> ?.
by rewrite -big_sepM2_delete.
Qed.
Lemma big_sepM2_insert_delete Φ m1 m2 i x1 x2 :
([ map] ky1;y2 <[i:=x1]>m1; <[i:=x2]>m2, Φ k y1 y2)
⊣⊢ Φ i x1 x2 [ map] ky1;y2 delete i m1;delete i m2, Φ k y1 y2.
Proof.
rewrite -(insert_delete_insert m1) -(insert_delete_insert m2).
apply big_sepM2_insert; by rewrite lookup_delete.
Qed.
Lemma big_sepM2_insert_acc Φ m1 m2 i x1 x2 :
m1 !! i = Some x1 m2 !! i = Some x2
([ map] ky1;y2 m1;m2, Φ k y1 y2)
Φ i x1 x2 ( x1' x2', Φ i x1' x2' -∗
([ map] ky1;y2 <[i:=x1']>m1;<[i:=x2']>m2, Φ k y1 y2)).
Proof.
intros ??. rewrite {1}big_sepM2_delete //. apply sep_mono; [done|].
apply forall_intro=> x1'. apply forall_intro=> x2'.
rewrite -(insert_delete_insert m1) -(insert_delete_insert m2)
big_sepM2_insert ?lookup_delete //.
by apply wand_intro_l.
Qed.
Lemma big_sepM2_insert_2 Φ m1 m2 i x1 x2
`{!TCOr ( x y, Affine (Φ i x y)) (Absorbing (Φ i x1 x2))} :
Φ i x1 x2 -∗
([ map] ky1;y2 m1;m2, Φ k y1 y2) -∗
([ map] ky1;y2 <[i:=x1]>m1; <[i:=x2]>m2, Φ k y1 y2).
Proof.
rewrite !big_sepM2_alt.
assert (TCOr ( x, Affine (Φ i x.1 x.2)) (Absorbing (Φ i x1 x2))).
{ destruct select (TCOr _ _); apply _. }
apply entails_wand, wand_intro_r.
rewrite !persistent_and_affinely_sep_l /=.
rewrite (sep_comm (Φ _ _ _)) -sep_assoc. apply sep_mono.
{ apply affinely_mono, pure_mono. rewrite !dom_insert_L. set_solver. }
rewrite (big_sepM_insert_2 (λ k xx, Φ k xx.1 xx.2) (map_zip m1 m2) i (x1, x2)).
rewrite map_insert_zip_with. apply wand_elim_r.
Qed.
Lemma big_sepM2_lookup_acc Φ m1 m2 i x1 x2 :
m1 !! i = Some x1 m2 !! i = Some x2
([ map] ky1;y2 m1;m2, Φ k y1 y2)
Φ i x1 x2 (Φ i x1 x2 -∗ ([ map] ky1;y2 m1;m2, Φ k y1 y2)).
Proof.
intros Hm1 Hm2. etrans; first apply big_sepM2_insert_acc=>//.
apply sep_mono_r. rewrite (forall_elim x1) (forall_elim x2).
rewrite !insert_id //.
Qed.
Lemma big_sepM2_lookup Φ m1 m2 i x1 x2
`{!TCOr ( j y1 y2, Affine (Φ j y1 y2)) (Absorbing (Φ i x1 x2))} :
m1 !! i = Some x1 m2 !! i = Some x2
([ map] ky1;y2 m1;m2, Φ k y1 y2) Φ i x1 x2.
Proof.
intros Hx1 Hx2. destruct select (TCOr _ _).
- rewrite big_sepM2_delete // sep_elim_l //.
- rewrite big_sepM2_lookup_acc // sep_elim_l //.
Qed.
Lemma big_sepM2_lookup_l Φ m1 m2 i x1
`{!TCOr ( j y1 y2, Affine (Φ j y1 y2)) ( x2, Absorbing (Φ i x1 x2))} :
m1 !! i = Some x1
([ map] ky1;y2 m1;m2, Φ k y1 y2)
x2, m2 !! i = Some x2 Φ i x1 x2.
Proof.
intros Hm1. rewrite big_sepM2_delete_l //.
f_equiv=> x2. destruct select (TCOr _ _); by rewrite sep_elim_l.
Qed.
Lemma big_sepM2_lookup_r Φ m1 m2 i x2
`{!TCOr ( j y1 y2, Affine (Φ j y1 y2)) ( x1, Absorbing (Φ i x1 x2))} :
m2 !! i = Some x2
([ map] ky1;y2 m1;m2, Φ k y1 y2)
x1, m1 !! i = Some x1 Φ i x1 x2.
Proof.
intros Hm2. rewrite big_sepM2_delete_r //.
f_equiv=> x1. destruct select (TCOr _ _); by rewrite sep_elim_l.
Qed.
Lemma big_sepM2_singleton Φ i x1 x2 :
([ map] ky1;y2 {[ i := x1 ]}; {[ i := x2 ]}, Φ k y1 y2) ⊣⊢ Φ i x1 x2.
Proof.
rewrite big_sepM2_alt.
rewrite map_zip_with_singleton big_sepM_singleton.
apply (anti_symm _).
- apply and_elim_r.
- rewrite <- (left_id True%I ()%I (Φ i x1 x2)).
apply and_mono=> //. apply pure_mono=> _. set_solver.
Qed.
Lemma big_sepM2_fst_snd Φ m :
([ map] ky1;y2 fst <$> m; snd <$> m, Φ k y1 y2) ⊣⊢
[ map] k xy m, Φ k (xy.1) (xy.2).
Proof.
rewrite big_sepM2_alt. rewrite !dom_fmap_L.
by rewrite pure_True // True_and map_zip_fst_snd.
Qed.
Lemma big_sepM2_fmap {A' B'} (f : A A') (g : B B') (Φ : K A' B' PROP) m1 m2 :
([ map] ky1;y2 f <$> m1; g <$> m2, Φ k y1 y2)
⊣⊢ ([ map] ky1;y2 m1;m2, Φ k (f y1) (g y2)).
Proof.
rewrite !big_sepM2_alt. by rewrite map_fmap_zip !dom_fmap_L big_sepM_fmap.
Qed.
Lemma big_sepM2_fmap_l {A'} (f : A A') (Φ : K A' B PROP) m1 m2 :
([ map] ky1;y2 f <$> m1; m2, Φ k y1 y2)
⊣⊢ ([ map] ky1;y2 m1;m2, Φ k (f y1) y2).
Proof. rewrite -{1}(map_fmap_id m2). apply big_sepM2_fmap. Qed.
Lemma big_sepM2_fmap_r {B'} (g : B B') (Φ : K A B' PROP) m1 m2 :
([ map] ky1;y2 m1; g <$> m2, Φ k y1 y2)
⊣⊢ ([ map] ky1;y2 m1;m2, Φ k y1 (g y2)).
Proof. rewrite -{1}(map_fmap_id m1). apply big_sepM2_fmap. Qed.
Lemma big_sepM2_sep Φ Ψ m1 m2 :
([ map] ky1;y2 m1;m2, Φ k y1 y2 Ψ k y1 y2)
⊣⊢ ([ map] ky1;y2 m1;m2, Φ k y1 y2) ([ map] ky1;y2 m1;m2, Ψ k y1 y2).
Proof.
rewrite !big_sepM2_alt.
rewrite -{1}(idemp bi_and dom m1 = dom m2 ⌝%I).
rewrite -assoc.
rewrite !persistent_and_affinely_sep_l /=.
rewrite -assoc. apply sep_proper=>//.
rewrite assoc (comm _ _ (<affine> _)%I) -assoc.
apply sep_proper=>//. apply big_sepM_sep.
Qed.
Lemma big_sepM2_sep_2 Φ Ψ m1 m2 :
([ map] ky1;y2 m1;m2, Φ k y1 y2) -∗
([ map] ky1;y2 m1;m2, Ψ k y1 y2) -∗
([ map] ky1;y2 m1;m2, Φ k y1 y2 Ψ k y1 y2).
Proof. apply entails_wand, wand_intro_r. rewrite big_sepM2_sep //. Qed.
Lemma big_sepM2_and Φ Ψ m1 m2 :
([ map] ky1;y2 m1;m2, Φ k y1 y2 Ψ k y1 y2)
([ map] ky1;y2 m1;m2, Φ k y1 y2) ([ map] ky1;y2 m1;m2, Ψ k y1 y2).
Proof. auto using and_intro, big_sepM2_mono, and_elim_l, and_elim_r. Qed.
Lemma big_sepM2_pure_1 (φ : K A B Prop) m1 m2 :
([ map] ky1;y2 m1;m2, φ k y1 y2) ⊢@{PROP}
⌜∀ k y1 y2, m1 !! k = Some y1 m2 !! k = Some y2 φ k y1 y2⌝.
Proof.
rewrite big_sepM2_alt.
rewrite big_sepM_pure_1 -pure_and.
f_equiv=>-[Hdom Hforall] k y1 y2 Hy1 Hy2.
eapply (Hforall k (y1, y2)). clear Hforall.
apply map_lookup_zip_with_Some. naive_solver.
Qed.
Lemma big_sepM2_affinely_pure_2 (φ : K A B Prop) m1 m2 :
( k : K, is_Some (m1 !! k) is_Some (m2 !! k))
<affine> ⌜∀ k y1 y2, m1 !! k = Some y1 m2 !! k = Some y2 φ k y1 y2 ⊢@{PROP}
([ map] ky1;y2 m1;m2, <affine> φ k y1 y2).
Proof.
intros Hdom.
rewrite big_sepM2_alt_lookup.
rewrite -big_sepM_affinely_pure_2.
rewrite affinely_and_r -pure_and. f_equiv. f_equiv=>-Hforall.
split; first done.
intros k [y1 y2] (? & ? & [= <- <-] & Hy1 & Hy2)%map_lookup_zip_with_Some; simpl.
by eapply Hforall.
Qed.
(** The general backwards direction requires [BiAffine] to cover the empty case. *)
Lemma big_sepM2_pure `{!BiAffine PROP} (φ : K A B Prop) m1 m2 :
([ map] ky1;y2 m1;m2, φ k y1 y2) ⊣⊢@{PROP}
( k : K, is_Some (m1 !! k) is_Some (m2 !! k))
( k y1 y2, m1 !! k = Some y1 m2 !! k = Some y2 φ k y1 y2)⌝.
Proof.
apply (anti_symm ()).
{ rewrite pure_and. apply and_intro.
- apply big_sepM2_lookup_iff.
- apply big_sepM2_pure_1. }
rewrite -(affine_affinely ⌜_⌝%I).
rewrite pure_and -affinely_and_r.
apply pure_elim_l=>Hdom.
rewrite big_sepM2_affinely_pure_2 //. by setoid_rewrite affinely_elim.
Qed.
Lemma big_sepM2_persistently `{!BiAffine PROP} Φ m1 m2 :
<pers> ([ map] ky1;y2 m1;m2, Φ k y1 y2)
⊣⊢ [ map] ky1;y2 m1;m2, <pers> (Φ k y1 y2).
Proof.
by rewrite !big_sepM2_alt persistently_and
persistently_pure big_sepM_persistently.
Qed.
Lemma big_sepM2_intro Φ m1 m2 :
( k : K, is_Some (m1 !! k) is_Some (m2 !! k))
( k x1 x2, m1 !! k = Some x1 m2 !! k = Some x2 Φ k x1 x2)
[ map] kx1;x2 m1;m2, Φ k x1 x2.
Proof.
intros. rewrite big_sepM2_alt_lookup.
apply and_intro; [by apply pure_intro|].
rewrite -big_sepM_intro. f_equiv; f_equiv=> k.
apply forall_intro=> -[x1 x2]. rewrite (forall_elim x1) (forall_elim x2).
apply impl_intro_l, pure_elim_l.
intros (?&?&[= <- <-]&?&?)%map_lookup_zip_with_Some.
by rewrite !pure_True // !True_impl.
Qed.
Lemma big_sepM2_forall `{!BiAffine PROP} Φ m1 m2 :
( k x1 x2, Persistent (Φ k x1 x2))
([ map] kx1;x2 m1;m2, Φ k x1 x2) ⊣⊢
⌜∀ k : K, is_Some (m1 !! k) is_Some (m2 !! k)
( k x1 x2, m1 !! k = Some x1 m2 !! k = Some x2 Φ k x1 x2).
Proof.
intros. apply (anti_symm _).
{ apply and_intro; [apply big_sepM2_lookup_iff|].
apply forall_intro=> k. apply forall_intro=> x1. apply forall_intro=> x2.
do 2 (apply impl_intro_l; apply pure_elim_l=> ?). by apply: big_sepM2_lookup. }
apply pure_elim_l=> Hdom. rewrite big_sepM2_alt_lookup.
apply and_intro; [by apply pure_intro|].
rewrite big_sepM_forall. f_equiv=> k.
apply forall_intro=> -[x1 x2]. rewrite (forall_elim x1) (forall_elim x2).
apply impl_intro_l, pure_elim_l.
intros (?&?&[= <- <-]&?&?)%map_lookup_zip_with_Some.
by rewrite !pure_True // !True_impl.
Qed.
Lemma big_sepM2_impl Φ Ψ m1 m2 :
([ map] ky1;y2 m1;m2, Φ k y1 y2) -∗
( k x1 x2,
m1 !! k = Some x1 m2 !! k = Some x2 Φ k x1 x2 -∗ Ψ k x1 x2) -∗
[ map] ky1;y2 m1;m2, Ψ k y1 y2.
Proof.
apply entails_wand.
rewrite -(idemp bi_and (big_sepM2 _ _ _)) {1}big_sepM2_lookup_iff.
apply pure_elim_l=> ?. rewrite big_sepM2_intro //.
apply bi.wand_intro_l. rewrite -big_sepM2_sep. by setoid_rewrite wand_elim_l.
Qed.
Lemma big_sepM2_wand Φ Ψ m1 m2 :
([ map] ky1;y2 m1;m2, Φ k y1 y2) -∗
([ map] ky1;y2 m1;m2, Φ k y1 y2 -∗ Ψ k y1 y2) -∗
[ map] ky1;y2 m1;m2, Ψ k y1 y2.
Proof.
apply entails_wand, wand_intro_r. rewrite -big_sepM2_sep.
setoid_rewrite wand_elim_r. done.
Qed.
Lemma big_sepM2_lookup_acc_impl {Φ m1 m2} i x1 x2 :
m1 !! i = Some x1
m2 !! i = Some x2
([ map] ky1;y2 m1;m2, Φ k y1 y2) -∗
(* We obtain [Φ] for [x1] and [x2] *)
Φ i x1 x2
(* We reobtain the bigop for a predicate [Ψ] selected by the user *)
Ψ,
( k y1 y2,
m1 !! k = Some y1 m2 !! k = Some y2 k i
Φ k y1 y2 -∗ Ψ k y1 y2) -∗
Ψ i x1 x2 -∗
[ map] ky1;y2 m1;m2, Ψ k y1 y2.
Proof.
intros. rewrite big_sepM2_delete //. apply entails_wand, sep_mono_r, forall_intro=> Ψ.
apply wand_intro_r, wand_intro_l.
rewrite (big_sepM2_delete Ψ m1 m2 i) //. apply sep_mono_r.
eapply wand_apply; [apply wand_entails, big_sepM2_impl|apply sep_mono_r].
f_equiv; f_equiv=> k; f_equiv=> y1; f_equiv=> y2.
rewrite !impl_curry -!pure_and !lookup_delete_Some.
do 2 f_equiv. intros ?; naive_solver.
Qed.
Lemma big_sepM2_later_1 `{!BiAffine PROP} Φ m1 m2 :
( [ map] kx1;x2 m1;m2, Φ k x1 x2)
([ map] kx1;x2 m1;m2, Φ k x1 x2).
Proof.
rewrite !big_sepM2_alt later_and (timeless ⌜_⌝).
rewrite big_sepM_later except_0_and.
auto using and_mono_r, except_0_intro.
Qed.
Lemma big_sepM2_later_2 Φ m1 m2 :
([ map] kx1;x2 m1;m2, Φ k x1 x2)
[ map] kx1;x2 m1;m2, Φ k x1 x2.
Proof.
rewrite !big_sepM2_alt later_and -(later_intro ⌜_⌝).
apply and_mono_r. by rewrite big_opM_commute.
Qed.
Lemma big_sepM2_laterN_2 Φ n m1 m2 :
([ map] kx1;x2 m1;m2, ▷^n Φ k x1 x2)
▷^n [ map] kx1;x2 m1;m2, Φ k x1 x2.
Proof.
induction n as [|n IHn]; first done.
rewrite later_laterN -IHn -big_sepM2_later_2.
apply big_sepM2_mono. eauto.
Qed.
Lemma big_sepM2_sepM (Φ1 : K A PROP) (Φ2 : K B PROP) m1 m2 :
( k, is_Some (m1 !! k) is_Some (m2 !! k))
([ map] ky1;y2 m1;m2, Φ1 k y1 Φ2 k y2) ⊣⊢
([ map] ky1 m1, Φ1 k y1) ([ map] ky2 m2, Φ2 k y2).
Proof.
intros.
rewrite -big_sepM_sep_zip // big_sepM2_alt_lookup pure_True // left_id //.
Qed.
Lemma big_sepM2_sepM_2 (Φ1 : K A PROP) (Φ2 : K B PROP) m1 m2 :
( k, is_Some (m1 !! k) is_Some (m2 !! k))
([ map] ky1 m1, Φ1 k y1) -∗
([ map] ky2 m2, Φ2 k y2) -∗
[ map] ky1;y2 m1;m2, Φ1 k y1 Φ2 k y2.
Proof. intros. apply entails_wand, wand_intro_r. by rewrite big_sepM2_sepM. Qed.
Lemma big_sepM2_union_inv_l (Φ : K A B PROP) m1 m2 m' :
m1 ## m2
([ map] kx;y (m1 m2);m', Φ k x y)
m1' m2', m' = m1' m2' m1' ## m2'
([ map] kx;y m1;m1', Φ k x y)
([ map] kx;y m2;m2', Φ k x y).
Proof.
revert m'. induction m1 as [|i x m1 ? IH] using map_ind;
intros m' ?; decompose_map_disjoint.
{ rewrite -(exist_intro ) -(exist_intro m') !left_id_L.
rewrite !pure_True //; last by apply map_disjoint_empty_l.
rewrite big_sepM2_empty !left_id //. }
rewrite -insert_union_l big_sepM2_delete_l; last by apply lookup_insert.
apply exist_elim=> y. apply pure_elim_l=> ?.
rewrite delete_insert; last by apply lookup_union_None.
rewrite IH //.
rewrite sep_exist_l. eapply exist_elim=> m1'.
rewrite sep_exist_l. eapply exist_elim=> m2'.
rewrite comm. apply wand_elim_l', pure_elim_l=> Hm'. apply pure_elim_l=> ?.
assert ((m1' m2') !! i = None) as [??]%lookup_union_None.
{ by rewrite -Hm' lookup_delete. }
apply wand_intro_l.
rewrite -(exist_intro (<[i:=y]> m1')) -(exist_intro m2'). apply and_intro.
{ apply pure_intro. by rewrite -insert_union_l -Hm' insert_delete. }
apply and_intro.
{ apply pure_intro. by apply map_disjoint_insert_l. }
by rewrite big_sepM2_insert // -assoc.
Qed.
End map2.
Lemma big_sepM2_union_inv_r `{Countable K} {A B}
(Φ : K A B PROP) (m1 m2 : gmap K B) (m' : gmap K A) :
m1 ## m2
([ map] kx;y m';(m1 m2), Φ k x y)
m1' m2', m' = m1' m2' m1' ## m2'
([ map] kx;y m1';m1, Φ k x y)
([ map] kx;y m2';m2, Φ k x y).
Proof.
intros Hm. rewrite -(big_sepM2_flip Φ).
rewrite (big_sepM2_union_inv_l (λ k (x : B) y, Φ k y x)) //.
f_equiv=>n1. f_equiv=>n2. f_equiv.
by rewrite -!(big_sepM2_flip Φ).
Qed.
Lemma big_sepM_sepM2_diag `{Countable K} {A} (Φ : K A A PROP) (m : gmap K A) :
([ map] ky m, Φ k y y)
([ map] ky1;y2 m;m, Φ k y1 y2).
Proof.
rewrite big_sepM2_alt. rewrite pure_True; last naive_solver. rewrite left_id.
rewrite map_zip_diag big_sepM_fmap. done.
Qed.
Lemma big_sepM2_ne_2 `{Countable K} (A B : ofe)
(Φ Ψ : K A B PROP) m1 m2 m1' m2' n :
m1 {n} m1' m2 {n} m2'
( k y1 y1' y2 y2',
m1 !! k = Some y1 m1' !! k = Some y1' y1 {n} y1'
m2 !! k = Some y2 m2' !! k = Some y2' y2 {n} y2'
Φ k y1 y2 {n} Ψ k y1' y2')
([ map] k y1;y2 m1;m2, Φ k y1 y2)%I {n} ([ map] k y1;y2 m1';m2', Ψ k y1 y2)%I.
Proof.
intros Hm1 Hm2 Hf. rewrite !big_sepM2_alt. f_equiv.
{ by rewrite Hm1 Hm2. }
apply big_opM_ne_2; [by f_equiv|].
intros k [x1 y1] [x2 y2] (?&?&[=<- <-]&?&?)%map_lookup_zip_with_Some
(?&?&[=<- <-]&?&?)%map_lookup_zip_with_Some [??]; naive_solver.
Qed.
(** ** Big ops over finite sets *)
Section gset.
Context `{Countable A}.
Implicit Types X : gset A.
Implicit Types Φ : A PROP.
(** The lemmas [big_sepS_mono], [big_sepS_ne] and [big_sepS_proper] are more
generic than the instances as they also give [x ∈ X] in the premise. *)
Lemma big_sepS_mono Φ Ψ X :
( x, x X Φ x Ψ x)
([ set] x X, Φ x) [ set] x X, Ψ x.
Proof. intros. apply big_opS_gen_proper; apply _ || auto. Qed.
Lemma big_sepS_ne Φ Ψ X n :
( x, x X Φ x {n} Ψ x)
([ set] x X, Φ x)%I {n} ([ set] x X, Ψ x)%I.
Proof. apply big_opS_ne. Qed.
Lemma big_sepS_proper Φ Ψ X :
( x, x X Φ x ⊣⊢ Ψ x)
([ set] x X, Φ x) ⊣⊢ ([ set] x X, Ψ x).
Proof. apply big_opS_proper. Qed.
(** No need to declare instances for non-expansiveness and properness, we
get both from the generic [big_opS] instances. *)
Global Instance big_sepS_mono' :
Proper (pointwise_relation _ () ==> (=) ==> ()) (big_opS (@bi_sep PROP) (A:=A)).
Proof. intros f g Hf m ? <-. by apply big_sepS_mono. Qed.
Global Instance big_sepS_flip_mono' :
Proper (pointwise_relation _ (flip ()) ==> (=) ==> flip ())
(big_opS (@bi_sep PROP) (A:=A)).
Proof. solve_proper. Qed.
Global Instance big_sepS_empty_persistent Φ :
Persistent ([ set] x , Φ x).
Proof. rewrite big_opS_empty. apply _. Qed.
Lemma big_sepS_persistent Φ X :
( x, x X Persistent (Φ x)) Persistent ([ set] x X, Φ x).
Proof. apply big_opS_closed; apply _. Qed.
Global Instance big_sepS_persistent' Φ X :
( x, Persistent (Φ x)) Persistent ([ set] x X, Φ x).
Proof. intros; apply big_sepS_persistent, _. Qed.
Global Instance big_sepS_empty_affine Φ : Affine ([ set] x , Φ x).
Proof. rewrite big_opS_empty. apply _. Qed.
Lemma big_sepS_affine Φ X :
( x, x X Affine (Φ x)) Affine ([ set] x X, Φ x).
Proof. apply big_opS_closed; apply _. Qed.
Global Instance big_sepS_affine' Φ X :
( x, Affine (Φ x)) Affine ([ set] x X, Φ x).
Proof. intros; apply big_sepS_affine, _. Qed.
Global Instance big_sepS_empty_timeless `{!Timeless (emp%I : PROP)} Φ :
Timeless ([ set] x , Φ x).
Proof. rewrite big_opS_empty. apply _. Qed.
Lemma big_sepS_timeless `{!Timeless (emp%I : PROP)} Φ X :
( x, x X Timeless (Φ x)) Timeless ([ set] x X, Φ x).
Proof. apply big_opS_closed; apply _. Qed.
Global Instance big_sepS_timeless' `{!Timeless (emp%I : PROP)} Φ X :
( x, Timeless (Φ x)) Timeless ([ set] x X, Φ x).
Proof. intros; apply big_sepS_timeless, _. Qed.
(* See comment above [big_sepM_subseteq] as for why the [Affine]
instance is placed late. *)
Lemma big_sepS_subseteq Φ X Y `{!∀ x, Affine (Φ x)} :
Y X ([ set] x X, Φ x) [ set] x Y, Φ x.
Proof.
intros ->%union_difference_L. rewrite big_opS_union; last set_solver.
by rewrite sep_elim_l.
Qed.
Lemma big_sepS_elements Φ X :
([ set] x X, Φ x) ⊣⊢ ([ list] x elements X, Φ x).
Proof. by rewrite big_opS_elements. Qed.
Lemma big_sepS_empty Φ : ([ set] x , Φ x) ⊣⊢ emp.
Proof. by rewrite big_opS_empty. Qed.
Lemma big_sepS_empty' P `{!Affine P} Φ : P [ set] x , Φ x.
Proof. rewrite big_sepS_empty. apply: affine. Qed.
Lemma big_sepS_emp X : ([ set] x X, emp) ⊣⊢@{PROP} emp.
Proof. by rewrite big_opS_unit. Qed.
Lemma big_sepS_insert Φ X x :
x X ([ set] y {[ x ]} X, Φ y) ⊣⊢ (Φ x [ set] y X, Φ y).
Proof. apply big_opS_insert. Qed.
Lemma big_sepS_fn_insert {B} (Ψ : A B PROP) f X x b :
x X
([ set] y {[ x ]} X, Ψ y (<[x:=b]> f y))
⊣⊢ (Ψ x b [ set] y X, Ψ y (f y)).
Proof. apply big_opS_fn_insert. Qed.
Lemma big_sepS_fn_insert' Φ X x P :
x X ([ set] y {[ x ]} X, <[x:=P]> Φ y) ⊣⊢ (P [ set] y X, Φ y).
Proof. apply big_opS_fn_insert'. Qed.
Lemma big_sepS_union Φ X Y :
X ## Y
([ set] y X Y, Φ y) ⊣⊢ ([ set] y X, Φ y) ([ set] y Y, Φ y).
Proof. apply big_opS_union. Qed.
Lemma big_sepS_delete Φ X x :
x X ([ set] y X, Φ y) ⊣⊢ Φ x [ set] y X {[ x ]}, Φ y.
Proof. apply big_opS_delete. Qed.
Lemma big_sepS_insert_2 {Φ X} x
`{!TCOr (Affine (Φ x)) (Absorbing (Φ x))} :
Φ x ([ set] y X, Φ y) -∗ ([ set] y {[ x ]} X, Φ y).
Proof.
apply wand_intro_r. destruct (decide (x X)); last first.
{ rewrite -big_sepS_insert //. }
rewrite big_sepS_delete // assoc.
rewrite (sep_elim_l (Φ x)) -big_sepS_insert; last set_solver.
rewrite -union_difference_singleton_L //.
replace ({[x]} X) with X by set_solver.
auto.
Qed.
Lemma big_sepS_insert_2' {Φ X} x
`{!TCOr (Affine (Φ x)) (Absorbing (Φ x))} :
Φ x -∗ ([ set] y X, Φ y) -∗ ([ set] y X {[ x ]}, Φ y).
Proof. apply entails_wand. rewrite comm_L. by apply big_sepS_insert_2. Qed.
Lemma big_sepS_union_2 {Φ} X Y
`{!∀ x, TCOr (Affine (Φ x)) (Absorbing (Φ x))} :
([ set] y X, Φ y) -∗ ([ set] y Y, Φ y) -∗ ([ set] y X Y, Φ y).
Proof.
apply entails_wand, wand_intro_r. induction X as [|x X ? IH] using set_ind_L.
{ by rewrite left_id_L big_sepS_empty left_id. }
rewrite big_sepS_insert // -assoc IH -assoc_L.
destruct (decide (x Y)).
{ replace ({[x]} (X Y)) with (X Y) by set_solver.
rewrite (big_sepS_delete _ _ x); last set_solver.
by rewrite assoc sep_elim_r. }
by rewrite big_sepS_insert; last set_solver.
Qed.
Lemma big_sepS_delete_2 {Φ X} x :
Affine (Φ x)
Φ x ([ set] y X {[ x ]}, Φ y) -∗ [ set] y X, Φ y.
Proof.
intros Haff. apply wand_intro_r. destruct (decide (x X)).
{ rewrite -big_sepS_delete //. }
rewrite sep_elim_r.
replace (X {[x]}) with X by set_solver.
auto.
Qed.
Lemma big_sepS_elem_of Φ X x
`{!TCOr ( y, Affine (Φ y)) (Absorbing (Φ x))} :
x X ([ set] y X, Φ y) Φ x.
Proof.
intros ?. rewrite big_sepS_delete //.
destruct select (TCOr _ _); by rewrite sep_elim_l.
Qed.
Lemma big_sepS_elem_of_acc Φ X x :
x X
([ set] y X, Φ y) Φ x (Φ x -∗ ([ set] y X, Φ y)).
Proof.
intros. rewrite big_sepS_delete //. by apply sep_mono_r, wand_intro_l.
Qed.
Lemma big_sepS_singleton Φ x : ([ set] y {[ x ]}, Φ y) ⊣⊢ Φ x.
Proof. apply big_opS_singleton. Qed.
Lemma big_sepS_filter' (φ : A Prop) `{ x, Decision (φ x)} Φ X :
([ set] y filter φ X, Φ y)
⊣⊢ ([ set] y X, if decide (φ y) then Φ y else emp).
Proof. apply: big_opS_filter'. Qed.
Lemma big_sepS_filter `{!BiAffine PROP}
(φ : A Prop) `{ x, Decision (φ x)} Φ X :
([ set] y filter φ X, Φ y) ⊣⊢ ([ set] y X, φ y Φ y).
Proof. setoid_rewrite <-decide_emp. apply big_sepS_filter'. Qed.
Lemma big_sepS_filter_acc' (φ : A Prop) `{ y, Decision (φ y)} Φ X Y :
( y, y Y φ y y X)
([ set] y X, Φ y) -∗
([ set] y Y, if decide (φ y) then Φ y else emp)
(([ set] y Y, if decide (φ y) then Φ y else emp) -∗ [ set] y X, Φ y).
Proof.
intros ?. destruct (proj1 (subseteq_disjoint_union_L (filter φ Y) X))
as (Z&->&?); first set_solver.
rewrite big_sepS_union // big_sepS_filter'.
by apply entails_wand, sep_mono_r, wand_intro_l.
Qed.
Lemma big_sepS_filter_acc `{!BiAffine PROP}
(φ : A Prop) `{ y, Decision (φ y)} Φ X Y :
( y, y Y φ y y X)
([ set] y X, Φ y) -∗
([ set] y Y, φ y Φ y)
(([ set] y Y, φ y Φ y) -∗ [ set] y X, Φ y).
Proof. intros. setoid_rewrite <-decide_emp. by apply big_sepS_filter_acc'. Qed.
Lemma big_sepS_list_to_set Φ (l : list A) :
NoDup l
([ set] x list_to_set l, Φ x) ⊣⊢ [ list] x l, Φ x.
Proof. apply big_opS_list_to_set. Qed.
Lemma big_sepS_sep Φ Ψ X :
([ set] y X, Φ y Ψ y) ⊣⊢ ([ set] y X, Φ y) ([ set] y X, Ψ y).
Proof. apply big_opS_op. Qed.
Lemma big_sepS_sep_2 Φ Ψ X :
([ set] y X, Φ y) -∗
([ set] y X, Ψ y) -∗
([ set] y X, Φ y Ψ y).
Proof. apply entails_wand, wand_intro_r. rewrite big_sepS_sep //. Qed.
Lemma big_sepS_and Φ Ψ X :
([ set] y X, Φ y Ψ y) ([ set] y X, Φ y) ([ set] y X, Ψ y).
Proof. auto using and_intro, big_sepS_mono, and_elim_l, and_elim_r. Qed.
Lemma big_sepS_pure_1 (φ : A Prop) X :
([ set] y X, φ y) ⊢@{PROP} set_Forall φ X⌝.
Proof.
induction X as [|x X ? IH] using set_ind_L.
{ apply pure_intro, set_Forall_empty. }
rewrite big_sepS_insert // IH sep_and -pure_and.
f_equiv=>-[Hx HX]. apply set_Forall_union.
- apply set_Forall_singleton. done.
- done.
Qed.
Lemma big_sepS_affinely_pure_2 (φ : A Prop) X :
<affine> set_Forall φ X ⊢@{PROP} ([ set] y X, <affine> φ y).
Proof.
induction X as [|x X ? IH] using set_ind_L.
{ rewrite big_sepS_empty. apply affinely_elim_emp. }
rewrite big_sepS_insert // -IH.
rewrite -persistent_and_sep_1 -affinely_and -pure_and.
f_equiv. f_equiv=>HX. split.
- apply HX. set_solver+.
- apply set_Forall_union_inv_2 in HX. done.
Qed.
(** The general backwards direction requires [BiAffine] to cover the empty case. *)
Lemma big_sepS_pure `{!BiAffine PROP} (φ : A Prop) X :
([ set] y X, φ y) ⊣⊢@{PROP} set_Forall φ X⌝.
Proof.
apply (anti_symm ()); first by apply big_sepS_pure_1.
rewrite -(affine_affinely ⌜_⌝%I).
rewrite big_sepS_affinely_pure_2. by setoid_rewrite affinely_elim.
Qed.
Lemma big_sepS_persistently `{!BiAffine PROP} Φ X :
<pers> ([ set] y X, Φ y) ⊣⊢ [ set] y X, <pers> (Φ y).
Proof. apply (big_opS_commute _). Qed.
Lemma big_sepS_intro Φ X :
( x, x X Φ x) [ set] x X, Φ x.
Proof.
revert Φ. induction X as [|x X ? IH] using set_ind_L=> Φ.
{ by rewrite (affine ( _)) big_sepS_empty. }
rewrite intuitionistically_sep_dup big_sepS_insert //. f_equiv.
- rewrite (forall_elim x) pure_True ?True_impl; last set_solver.
by rewrite intuitionistically_elim.
- rewrite -IH. f_equiv. apply forall_mono=> y.
apply impl_intro_l, pure_elim_l=> ?.
by rewrite pure_True ?True_impl; last set_solver.
Qed.
Lemma big_sepS_forall `{!BiAffine PROP} Φ X :
( x, Persistent (Φ x))
([ set] x X, Φ x) ⊣⊢ ( x, x X Φ x).
Proof.
intros . apply (anti_symm _).
{ apply forall_intro=> x.
apply impl_intro_l, pure_elim_l=> ?; by apply: big_sepS_elem_of. }
revert Φ . induction X as [|x X ? IH] using set_ind_L=> Φ .
{ rewrite big_sepS_empty. apply: affine. }
rewrite big_sepS_insert // -persistent_and_sep_1. apply and_intro.
- rewrite (forall_elim x) pure_True ?True_impl; last set_solver. done.
- rewrite -IH. apply forall_mono=> y.
apply impl_intro_l, pure_elim_l=> ?.
by rewrite pure_True ?True_impl; last set_solver.
Qed.
Lemma big_sepS_impl Φ Ψ X :
([ set] x X, Φ x) -∗
( x, x X Φ x -∗ Ψ x) -∗
[ set] x X, Ψ x.
Proof.
apply entails_wand, wand_intro_l. rewrite big_sepS_intro -big_sepS_sep.
by setoid_rewrite wand_elim_l.
Qed.
Lemma big_sepS_wand Φ Ψ X :
([ set] x X, Φ x) -∗
([ set] x X, Φ x -∗ Ψ x) -∗
[ set] x X, Ψ x.
Proof.
apply entails_wand, wand_intro_r. rewrite -big_sepS_sep.
setoid_rewrite wand_elim_r. done.
Qed.
Lemma big_sepS_elem_of_acc_impl {Φ X} x :
x X
([ set] y X, Φ y) -∗
(* we get [Φ] for the element [x] *)
Φ x
(* we reobtain the bigop for a predicate [Ψ] selected by the user *)
Ψ,
( y, y X x y Φ y -∗ Ψ y) -∗
Ψ x -∗
[ set] y X, Ψ y.
Proof.
intros. rewrite big_sepS_delete //. apply entails_wand, sep_mono_r, forall_intro=> Ψ.
apply wand_intro_r, wand_intro_l.
rewrite (big_sepS_delete Ψ X x) //. apply sep_mono_r.
eapply wand_apply; [apply wand_entails, big_sepS_impl|apply sep_mono_r].
f_equiv; f_equiv=> y. rewrite impl_curry -pure_and.
do 2 f_equiv. intros ?; set_solver.
Qed.
Lemma big_sepS_dup P `{!Affine P} X :
(P -∗ P P) -∗ P -∗ [ set] x X, P.
Proof.
apply entails_wand, wand_intro_l. induction X as [|x X ? IH] using set_ind_L.
{ apply: big_sepS_empty'. }
rewrite !big_sepS_insert //.
rewrite intuitionistically_sep_dup {1}intuitionistically_elim.
rewrite assoc wand_elim_r -assoc. apply sep_mono; done.
Qed.
Lemma big_sepS_later `{!BiAffine PROP} Φ X :
([ set] y X, Φ y) ⊣⊢ ([ set] y X, Φ y).
Proof. apply (big_opS_commute _). Qed.
Lemma big_sepS_later_2 Φ X :
([ set] y X, Φ y) ([ set] y X, Φ y).
Proof. by rewrite big_opS_commute. Qed.
Lemma big_sepS_laterN `{!BiAffine PROP} Φ n X :
▷^n ([ set] y X, Φ y) ⊣⊢ ([ set] y X, ▷^n Φ y).
Proof. apply (big_opS_commute _). Qed.
Lemma big_sepS_laterN_2 Φ n X :
([ set] y X, ▷^n Φ y) ▷^n ([ set] y X, Φ y).
Proof. by rewrite big_opS_commute. Qed.
End gset.
Lemma big_sepM_dom `{Countable K} {A} (Φ : K PROP) (m : gmap K A) :
([ map] k↦_ m, Φ k) ⊣⊢ ([ set] k dom m, Φ k).
Proof. apply big_opM_dom. Qed.
Lemma big_sepM_gset_to_gmap `{Countable K} {A} (Φ : K A PROP) (X : gset K) c :
([ map] ka gset_to_gmap c X, Φ k a) ⊣⊢ ([ set] k X, Φ k c).
Proof. apply big_opM_gset_to_gmap. Qed.
(** ** Big ops over finite multisets *)
Section gmultiset.
Context `{Countable A}.
Implicit Types X : gmultiset A.
Implicit Types Φ : A PROP.
(** The lemmas [big_sepMS_mono], [big_sepMS_ne] and [big_sepMS_proper] are more
generic than the instances as they also give [l !! k = Some y] in the premise. *)
Lemma big_sepMS_mono Φ Ψ X :
( x, x X Φ x Ψ x)
([ mset] x X, Φ x) [ mset] x X, Ψ x.
Proof. intros. apply big_opMS_gen_proper; apply _ || auto. Qed.
Lemma big_sepMS_ne Φ Ψ X n :
( x, x X Φ x {n} Ψ x)
([ mset] x X, Φ x)%I {n} ([ mset] x X, Ψ x)%I.
Proof. apply big_opMS_ne. Qed.
Lemma big_sepMS_proper Φ Ψ X :
( x, x X Φ x ⊣⊢ Ψ x)
([ mset] x X, Φ x) ⊣⊢ ([ mset] x X, Ψ x).
Proof. apply big_opMS_proper. Qed.
(** No need to declare instances for non-expansiveness and properness, we
get both from the generic [big_opMS] instances. *)
Global Instance big_sepMS_mono' :
Proper (pointwise_relation _ () ==> (=) ==> ()) (big_opMS (@bi_sep PROP) (A:=A)).
Proof. intros f g Hf m ? <-. by apply big_sepMS_mono. Qed.
Global Instance big_sepMS_flip_mono' :
Proper (pointwise_relation _ (flip ()) ==> (=) ==> flip ())
(big_opMS (@bi_sep PROP) (A:=A)).
Proof. solve_proper. Qed.
Global Instance big_sepMS_empty_persistent Φ :
Persistent ([ mset] x , Φ x).
Proof. rewrite big_opMS_empty. apply _. Qed.
Lemma big_sepMS_persistent Φ X :
( x, x X Persistent (Φ x)) Persistent ([ mset] x X, Φ x).
Proof. apply big_opMS_closed; apply _. Qed.
Global Instance big_sepMS_persistent' Φ X :
( x, Persistent (Φ x)) Persistent ([ mset] x X, Φ x).
Proof. intros; apply big_sepMS_persistent, _. Qed.
Global Instance big_sepMS_empty_affine Φ : Affine ([ mset] x , Φ x).
Proof. rewrite big_opMS_empty. apply _. Qed.
Lemma big_sepMS_affine Φ X :
( x, x X Affine (Φ x)) Affine ([ mset] x X, Φ x).
Proof. apply big_opMS_closed; apply _. Qed.
Global Instance big_sepMS_affine' Φ X :
( x, Affine (Φ x)) Affine ([ mset] x X, Φ x).
Proof. intros; apply big_sepMS_affine, _. Qed.
Global Instance big_sepMS_empty_timeless `{!Timeless (emp%I : PROP)} Φ :
Timeless ([ mset] x , Φ x).
Proof. rewrite big_opMS_empty. apply _. Qed.
Lemma big_sepMS_timeless `{!Timeless (emp%I : PROP)} Φ X :
( x, x X Timeless (Φ x)) Timeless ([ mset] x X, Φ x).
Proof. apply big_opMS_closed; apply _. Qed.
Global Instance big_sepMS_timeless' `{!Timeless (emp%I : PROP)} Φ X :
( x, Timeless (Φ x)) Timeless ([ mset] x X, Φ x).
Proof. intros; apply big_sepMS_timeless, _. Qed.
(* See comment above [big_sepM_subseteq] as for why the [Affine]
instance is placed late. *)
Lemma big_sepMS_subseteq Φ X Y `{!∀ x, Affine (Φ x)} :
Y X ([ mset] x X, Φ x) [ mset] x Y, Φ x.
Proof.
intros ->%gmultiset_disj_union_difference.
by rewrite big_opMS_disj_union sep_elim_l.
Qed.
Lemma big_sepMS_empty Φ : ([ mset] x , Φ x) ⊣⊢ emp.
Proof. by rewrite big_opMS_empty. Qed.
Lemma big_sepMS_empty' P `{!Affine P} Φ : P [ mset] x , Φ x.
Proof. rewrite big_sepMS_empty. apply: affine. Qed.
Lemma big_sepMS_disj_union Φ X Y :
([ mset] y X Y, Φ y) ⊣⊢ ([ mset] y X, Φ y) [ mset] y Y, Φ y.
Proof. apply big_opMS_disj_union. Qed.
Lemma big_sepMS_delete Φ X x :
x X ([ mset] y X, Φ y) ⊣⊢ Φ x [ mset] y X {[+ x +]}, Φ y.
Proof. apply big_opMS_delete. Qed.
Lemma big_sepMS_elem_of Φ X x
`{!TCOr ( y, Affine (Φ y)) (Absorbing (Φ x))} :
x X ([ mset] y X, Φ y) Φ x.
Proof.
intros ?. rewrite big_sepMS_delete //.
destruct select (TCOr _ _); by rewrite sep_elim_l.
Qed.
Lemma big_sepMS_elem_of_acc Φ X x :
x X
([ mset] y X, Φ y) Φ x (Φ x -∗ ([ mset] y X, Φ y)).
Proof.
intros. rewrite big_sepMS_delete //. by apply sep_mono_r, wand_intro_l.
Qed.
Lemma big_sepMS_singleton Φ x : ([ mset] y {[+ x +]}, Φ y) ⊣⊢ Φ x.
Proof. apply big_opMS_singleton. Qed.
Lemma big_sepMS_insert Φ X x :
([ mset] y {[+ x +]} X, Φ y) ⊣⊢ (Φ x [ mset] y X, Φ y).
Proof. apply big_opMS_insert. Qed.
Lemma big_sepMS_sep Φ Ψ X :
([ mset] y X, Φ y Ψ y) ⊣⊢ ([ mset] y X, Φ y) ([ mset] y X, Ψ y).
Proof. apply big_opMS_op. Qed.
Lemma big_sepMS_sep_2 Φ Ψ X :
([ mset] y X, Φ y) -∗
([ mset] y X, Ψ y) -∗
([ mset] y X, Φ y Ψ y).
Proof. apply entails_wand, wand_intro_r. rewrite big_sepMS_sep //. Qed.
Lemma big_sepMS_and Φ Ψ X :
([ mset] y X, Φ y Ψ y) ([ mset] y X, Φ y) ([ mset] y X, Ψ y).
Proof. auto using and_intro, big_sepMS_mono, and_elim_l, and_elim_r. Qed.
Lemma big_sepMS_later `{!BiAffine PROP} Φ X :
([ mset] y X, Φ y) ⊣⊢ ([ mset] y X, Φ y).
Proof. apply (big_opMS_commute _). Qed.
Lemma big_sepMS_later_2 Φ X :
([ mset] y X, Φ y) [ mset] y X, Φ y.
Proof. by rewrite big_opMS_commute. Qed.
Lemma big_sepMS_laterN `{!BiAffine PROP} Φ n X :
▷^n ([ mset] y X, Φ y) ⊣⊢ ([ mset] y X, ▷^n Φ y).
Proof. apply (big_opMS_commute _). Qed.
Lemma big_sepMS_laterN_2 Φ n X :
([ mset] y X, ▷^n Φ y) ▷^n [ mset] y X, Φ y.
Proof. by rewrite big_opMS_commute. Qed.
Lemma big_sepMS_pure_1 (φ : A Prop) X :
([ mset] y X, φ y) ⊢@{PROP} ⌜∀ y : A, y X φ y⌝.
Proof.
induction X as [|x X IH] using gmultiset_ind.
{ apply pure_intro=>??. exfalso. multiset_solver. }
rewrite big_sepMS_insert // IH sep_and -pure_and.
f_equiv=>-[Hx HX] y /gmultiset_elem_of_disj_union [|].
- move=>/gmultiset_elem_of_singleton =>->. done.
- eauto.
Qed.
Lemma big_sepMS_affinely_pure_2 (φ : A Prop) X :
<affine> ⌜∀ y : A, y X φ y ⊢@{PROP} ([ mset] y X, <affine> φ y).
Proof.
induction X as [|x X IH] using gmultiset_ind.
{ rewrite big_sepMS_empty. apply affinely_elim_emp. }
rewrite big_sepMS_insert // -IH.
rewrite -persistent_and_sep_1 -affinely_and -pure_and.
f_equiv. f_equiv=>HX. split.
- apply HX. set_solver+.
- intros y Hy. apply HX. multiset_solver.
Qed.
(** The general backwards direction requires [BiAffine] to cover the empty case. *)
Lemma big_sepMS_pure `{!BiAffine PROP} (φ : A Prop) X :
([ mset] y X, φ y) ⊣⊢@{PROP} ⌜∀ y : A, y X φ y⌝.
Proof.
apply (anti_symm ()); first by apply big_sepMS_pure_1.
rewrite -(affine_affinely ⌜_⌝%I).
rewrite big_sepMS_affinely_pure_2. by setoid_rewrite affinely_elim.
Qed.
Lemma big_sepMS_persistently `{!BiAffine PROP} Φ X :
<pers> ([ mset] y X, Φ y) ⊣⊢ [ mset] y X, <pers> (Φ y).
Proof. apply (big_opMS_commute _). Qed.
Lemma big_sepMS_intro Φ X :
( x, x X Φ x) [ mset] x X, Φ x.
Proof.
revert Φ. induction X as [|x X IH] using gmultiset_ind=> Φ.
{ by rewrite (affine ( _)) big_sepMS_empty. }
rewrite intuitionistically_sep_dup big_sepMS_disj_union.
rewrite big_sepMS_singleton. f_equiv.
- rewrite (forall_elim x) pure_True ?True_impl; last multiset_solver.
by rewrite intuitionistically_elim.
- rewrite -IH. f_equiv. apply forall_mono=> y.
apply impl_intro_l, pure_elim_l=> ?.
by rewrite pure_True ?True_impl; last multiset_solver.
Qed.
Lemma big_sepMS_forall `{!BiAffine PROP} Φ X :
( x, Persistent (Φ x))
([ mset] x X, Φ x) ⊣⊢ ( x, x X Φ x).
Proof.
intros . apply (anti_symm _).
{ apply forall_intro=> x.
apply impl_intro_l, pure_elim_l=> ?; by apply: big_sepMS_elem_of. }
revert Φ . induction X as [|x X IH] using gmultiset_ind=> Φ .
{ rewrite big_sepMS_empty. apply: affine. }
rewrite big_sepMS_disj_union.
rewrite big_sepMS_singleton -persistent_and_sep_1. apply and_intro.
- rewrite (forall_elim x) pure_True ?True_impl; last multiset_solver. done.
- rewrite -IH. apply forall_mono=> y.
apply impl_intro_l, pure_elim_l=> ?.
by rewrite pure_True ?True_impl; last multiset_solver.
Qed.
Lemma big_sepMS_impl Φ Ψ X :
([ mset] x X, Φ x) -∗
( x, x X Φ x -∗ Ψ x) -∗
[ mset] x X, Ψ x.
Proof.
apply entails_wand, wand_intro_l. rewrite big_sepMS_intro -big_sepMS_sep.
by setoid_rewrite wand_elim_l.
Qed.
Lemma big_sepMS_wand Φ Ψ X :
([ mset] x X, Φ x) -∗
([ mset] x X, Φ x -∗ Ψ x) -∗
[ mset] x X, Ψ x.
Proof.
apply entails_wand, wand_intro_r. rewrite -big_sepMS_sep.
setoid_rewrite wand_elim_r. done.
Qed.
Lemma big_sepMS_dup P `{!Affine P} X :
(P -∗ P P) -∗ P -∗ [ mset] x X, P.
Proof.
apply entails_wand, wand_intro_l. induction X as [|x X IH] using gmultiset_ind.
{ apply: big_sepMS_empty'. }
rewrite !big_sepMS_disj_union big_sepMS_singleton.
rewrite intuitionistically_sep_dup {1}intuitionistically_elim.
rewrite assoc wand_elim_r -assoc. apply sep_mono; done.
Qed.
Lemma big_sepMS_elem_of_acc_impl {Φ X} x :
x X
([ mset] y X, Φ y) -∗
(* we get [Φ] for [x] *)
Φ x
(* we reobtain the bigop for a predicate [Ψ] selected by the user *)
Ψ,
( y, y X {[+ x +]} Φ y -∗ Ψ y) -∗
Ψ x -∗
[ mset] y X, Ψ y.
Proof.
intros. rewrite big_sepMS_delete //. apply entails_wand, sep_mono_r, forall_intro=> Ψ.
apply wand_intro_r, wand_intro_l.
rewrite (big_sepMS_delete Ψ X x) //. apply sep_mono_r.
apply wand_elim_l', wand_entails, big_sepMS_impl.
Qed.
End gmultiset.
(** Commuting lemmas *)
Lemma big_sepL_sepL {A B} (Φ : nat A nat B PROP) (l1 : list A) (l2 : list B) :
([ list] k1x1 l1, [ list] k2x2 l2, Φ k1 x1 k2 x2) ⊣⊢
([ list] k2x2 l2, [ list] k1x1 l1, Φ k1 x1 k2 x2).
Proof. apply big_opL_opL. Qed.
Lemma big_sepL_sepM {A} `{Countable K} {B}
(Φ : nat A K B PROP) (l1 : list A) (m2 : gmap K B) :
([ list] k1x1 l1, [ map] k2x2 m2, Φ k1 x1 k2 x2) ⊣⊢
([ map] k2x2 m2, [ list] k1x1 l1, Φ k1 x1 k2 x2).
Proof. apply big_opL_opM. Qed.
Lemma big_sepL_sepS {A} `{Countable B}
(Φ : nat A B PROP) (l1 : list A) (X2 : gset B) :
([ list] k1x1 l1, [ set] x2 X2, Φ k1 x1 x2) ⊣⊢
([ set] x2 X2, [ list] k1x1 l1, Φ k1 x1 x2).
Proof. apply big_opL_opS. Qed.
Lemma big_sepL_sepMS {A} `{Countable B}
(Φ : nat A B PROP) (l1 : list A) (X2 : gmultiset B) :
([ list] k1x1 l1, [ mset] x2 X2, Φ k1 x1 x2) ⊣⊢
([ mset] x2 X2, [ list] k1x1 l1, Φ k1 x1 x2).
Proof. apply big_opL_opMS. Qed.
Lemma big_sepM_sepL `{Countable K} {A} {B}
(Φ : K A nat B PROP) (m1 : gmap K A) (l2 : list B) :
([ map] k1x1 m1, [ list] k2x2 l2, Φ k1 x1 k2 x2) ⊣⊢
([ list] k2x2 l2, [ map] k1x1 m1, Φ k1 x1 k2 x2).
Proof. apply big_opM_opL. Qed.
Lemma big_sepM_sepM `{Countable K1} {A} `{Countable K2} {B}
(Φ : K1 A K2 B PROP) (m1 : gmap K1 A) (m2 : gmap K2 B) :
([ map] k1x1 m1, [ map] k2x2 m2, Φ k1 x1 k2 x2) ⊣⊢
([ map] k2x2 m2, [ map] k1x1 m1, Φ k1 x1 k2 x2).
Proof. apply big_opM_opM. Qed.
Lemma big_sepM_sepS `{Countable K} {A} `{Countable B}
(Φ : K A B PROP) (m1 : gmap K A) (X2 : gset B) :
([ map] k1x1 m1, [ set] x2 X2, Φ k1 x1 x2) ⊣⊢
([ set] x2 X2, [ map] k1x1 m1, Φ k1 x1 x2).
Proof. apply big_opM_opS. Qed.
Lemma big_sepM_sepMS `{Countable K} {A} `{Countable B} (Φ : K A B PROP)
(m1 : gmap K A) (X2 : gmultiset B) :
([ map] k1x1 m1, [ mset] x2 X2, Φ k1 x1 x2) ⊣⊢
([ mset] x2 X2, [ map] k1x1 m1, Φ k1 x1 x2).
Proof. apply big_opM_opMS. Qed.
Lemma big_sepS_sepL `{Countable A} {B}
(f : A nat B PROP) (X1 : gset A) (l2 : list B) :
([ set] x1 X1, [ list] k2x2 l2, f x1 k2 x2) ⊣⊢
([ list] k2x2 l2, [ set] x1 X1, f x1 k2 x2).
Proof. apply big_opS_opL. Qed.
Lemma big_sepS_sepM `{Countable A} `{Countable K} {B}
(f : A K B PROP) (X1 : gset A) (m2 : gmap K B) :
([ set] x1 X1, [ map] k2x2 m2, f x1 k2 x2) ⊣⊢
([ map] k2x2 m2, [ set] x1 X1, f x1 k2 x2).
Proof. apply big_opS_opM. Qed.
Lemma big_sepS_sepS `{Countable A, Countable B}
(X : gset A) (Y : gset B) (Φ : A B PROP) :
([ set] x X, [ set] y Y, Φ x y) ⊣⊢ ([ set] y Y, [ set] x X, Φ x y).
Proof. apply big_opS_opS. Qed.
Lemma big_sepS_sepMS `{Countable A, Countable B}
(X : gset A) (Y : gmultiset B) (Φ : A B PROP) :
([ set] x X, [ mset] y Y, Φ x y) ⊣⊢ ([ mset] y Y, [ set] x X, Φ x y).
Proof. apply big_opS_opMS. Qed.
Lemma big_sepMS_sepL `{Countable A} {B}
(f : A nat B PROP) (X1 : gmultiset A) (l2 : list B) :
([ mset] x1 X1, [ list] k2x2 l2, f x1 k2 x2) ⊣⊢
([ list] k2x2 l2, [ mset] x1 X1, f x1 k2 x2).
Proof. apply big_opMS_opL. Qed.
Lemma big_sepMS_sepM `{Countable A} `{Countable K} {B} (f : A K B PROP)
(X1 : gmultiset A) (m2 : gmap K B) :
([ mset] x1 X1, [ map] k2x2 m2, f x1 k2 x2) ⊣⊢
([ map] k2x2 m2, [ mset] x1 X1, f x1 k2 x2).
Proof. apply big_opMS_opM. Qed.
Lemma big_sepMS_sepS `{Countable A, Countable B}
(X : gmultiset A) (Y : gset B) (f : A B PROP) :
([ mset] x X, [ set] y Y, f x y) ⊣⊢ ([ set] y Y, [ mset] x X, f x y).
Proof. apply big_opMS_opS. Qed.
Lemma big_sepMS_sepMS `{Countable A, Countable B}
(X : gmultiset A) (Y : gmultiset B) (Φ : A B PROP) :
([ mset] x X, [ mset] y Y, Φ x y) ⊣⊢ ([ mset] y Y, [ mset] x X, Φ x y).
Proof. apply big_opMS_opMS. Qed.
End big_op.
From iris.algebra Require Import monoid.
From iris.bi Require Export interface.
From iris.prelude Require Import options.
Definition bi_iff {PROP : bi} (P Q : PROP) : PROP := (P Q) (Q P).
Global Arguments bi_iff {_} _%_I _%_I : simpl never.
Global Instance: Params (@bi_iff) 1 := {}.
Infix "↔" := bi_iff : bi_scope.
Definition bi_wand_iff {PROP : bi} (P Q : PROP) : PROP :=
(P -∗ Q) (Q -∗ P).
Global Arguments bi_wand_iff {_} _%_I _%_I : simpl never.
Global Instance: Params (@bi_wand_iff) 1 := {}.
Infix "∗-∗" := bi_wand_iff : bi_scope.
Notation "P ∗-∗ Q" := ( P ∗-∗ Q) : stdpp_scope.
Class Persistent {PROP : bi} (P : PROP) := persistent : P <pers> P.
Global Arguments Persistent {_} _%_I : simpl never.
Global Arguments persistent {_} _%_I {_}.
Global Hint Mode Persistent + ! : typeclass_instances.
Global Instance: Params (@Persistent) 1 := {}.
Global Hint Extern 100 (Persistent (match ?x with _ => _ end)) =>
destruct x : typeclass_instances.
Definition bi_affinely {PROP : bi} (P : PROP) : PROP := emp P.
Global Arguments bi_affinely {_} _%_I : simpl never.
Global Instance: Params (@bi_affinely) 1 := {}.
Global Typeclasses Opaque bi_affinely.
Notation "'<affine>' P" := (bi_affinely P) : bi_scope.
Class Affine {PROP : bi} (Q : PROP) := affine : Q emp.
Global Arguments Affine {_} _%_I : simpl never.
Global Arguments affine {_} _%_I {_}.
Global Hint Mode Affine + ! : typeclass_instances.
Global Hint Extern 100 (Affine (match ?x with _ => _ end)) =>
destruct x : typeclass_instances.
Definition bi_absorbingly {PROP : bi} (P : PROP) : PROP := True P.
Global Arguments bi_absorbingly {_} _%_I : simpl never.
Global Instance: Params (@bi_absorbingly) 1 := {}.
Global Typeclasses Opaque bi_absorbingly.
Notation "'<absorb>' P" := (bi_absorbingly P) : bi_scope.
Class Absorbing {PROP : bi} (P : PROP) := absorbing : <absorb> P P.
Global Arguments Absorbing {_} _%_I : simpl never.
Global Arguments absorbing {_} _%_I.
Global Hint Mode Absorbing + ! : typeclass_instances.
Global Hint Extern 100 (Absorbing (match ?x with _ => _ end)) =>
destruct x : typeclass_instances.
Definition bi_persistently_if {PROP : bi} (p : bool) (P : PROP) : PROP :=
(if p then <pers> P else P)%I.
Global Arguments bi_persistently_if {_} !_ _%_I /.
Global Instance: Params (@bi_persistently_if) 2 := {}.
Global Typeclasses Opaque bi_persistently_if.
Notation "'<pers>?' p P" := (bi_persistently_if p P) : bi_scope.
Definition bi_affinely_if {PROP : bi} (p : bool) (P : PROP) : PROP :=
(if p then <affine> P else P)%I.
Global Arguments bi_affinely_if {_} !_ _%_I /.
Global Instance: Params (@bi_affinely_if) 2 := {}.
Global Typeclasses Opaque bi_affinely_if.
Notation "'<affine>?' p P" := (bi_affinely_if p P) : bi_scope.
Definition bi_absorbingly_if {PROP : bi} (p : bool) (P : PROP) : PROP :=
(if p then <absorb> P else P)%I.
Global Arguments bi_absorbingly_if {_} !_ _%_I /.
Global Instance: Params (@bi_absorbingly_if) 2 := {}.
Global Typeclasses Opaque bi_absorbingly_if.
Notation "'<absorb>?' p P" := (bi_absorbingly_if p P) : bi_scope.
Definition bi_intuitionistically {PROP : bi} (P : PROP) : PROP :=
(<affine> <pers> P)%I.
Global Arguments bi_intuitionistically {_} _%_I : simpl never.
Global Instance: Params (@bi_intuitionistically) 1 := {}.
Global Typeclasses Opaque bi_intuitionistically.
Notation "□ P" := (bi_intuitionistically P) : bi_scope.
Definition bi_intuitionistically_if {PROP : bi} (p : bool) (P : PROP) : PROP :=
(if p then P else P)%I.
Global Arguments bi_intuitionistically_if {_} !_ _%_I /.
Global Instance: Params (@bi_intuitionistically_if) 2 := {}.
Global Typeclasses Opaque bi_intuitionistically_if.
Notation "'□?' p P" := (bi_intuitionistically_if p P) : bi_scope.
Fixpoint bi_laterN {PROP : bi} (n : nat) (P : PROP) : PROP :=
match n with
| O => P
| S n' => ▷^n' P
end
where "▷^ n P" := (bi_laterN n P) : bi_scope.
Global Arguments bi_laterN {_} !_%_nat_scope _%_I.
Global Instance: Params (@bi_laterN) 2 := {}.
Notation "▷? p P" := (bi_laterN (Nat.b2n p) P) : bi_scope.
Definition bi_except_0 {PROP : bi} (P : PROP) : PROP := False P.
Global Arguments bi_except_0 {_} _%_I : simpl never.
Notation "◇ P" := (bi_except_0 P) : bi_scope.
Global Instance: Params (@bi_except_0) 1 := {}.
Global Typeclasses Opaque bi_except_0.
(* Timeless propositions are propositions that do not depend on the step-index.
There are two equivalent ways of expressing that a step-indexed proposition
[P : nat → Prop] is timeless:
* Option one, used here, says that [∀ n, P n → P (S n)]. In the logic, this
is stated as [▷ P ⊢ ◇ P] (which actually reads [∀ n > 0, P (n-1) → P n],
but this is trivially equivalent).
* Option two says that [∀ n, P 0 → P n]. In the logic, this is stated as a
meta-entailment [∀ P, (▷ False ∧ P ⊢ Q) → (P ⊢ Q)]. The assumption
[▷ False] expresses that the step-index is 0.
Both formulations are equivalent. In the logic, this can be shown using Löb
induction and [later_false_em]. In the model, this follows from regular
natural induction.
The law [timeless_alt] in [derived_laws_later.v] provides option two, by
proving that both versions are equivalent in the logic. *)
Class Timeless {PROP : bi} (P : PROP) := timeless : P P.
Global Arguments Timeless {_} _%_I : simpl never.
Global Arguments timeless {_} _%_I {_}.
Global Hint Mode Timeless + ! : typeclass_instances.
Global Instance: Params (@Timeless) 1 := {}.
Global Hint Extern 100 (Timeless (match ?x with _ => _ end)) =>
destruct x : typeclass_instances.
(** An optional precondition [mP] to [Q].
TODO: We may actually consider generalizing this to a list of preconditions,
and e.g. also using it for texan triples. *)
Definition bi_wandM {PROP : bi} (mP : option PROP) (Q : PROP) : PROP :=
match mP with
| None => Q
| Some P => P -∗ Q
end.
Global Arguments bi_wandM {_} !_%_I _%_I /.
Notation "mP -∗? Q" := (bi_wandM mP Q)
(at level 99, Q at level 200, right associativity) : bi_scope.
From iris.algebra Require Import monoid.
From iris.bi Require Export extensions.
From iris.prelude Require Import options.
(* The sections add [BiAffine] and the like, which is only picked up with [Type*]. *)
Set Default Proof Using "Type*".
(** Naming schema for lemmas about modalities:
M1_into_M2: M1 P ⊢ M2 P
M1_M2_elim: M1 (M2 P) ⊣⊢ M1 P
M1_elim_M2: M1 (M2 P) ⊣⊢ M2 P
M1_M2: M1 (M2 P) ⊣⊢ M2 (M1 P)
*)
Module bi.
Import interface.bi.
Section derived.
Context {PROP : bi}.
Implicit Types φ : Prop.
Implicit Types P Q R : PROP.
Implicit Types Ps : list PROP.
Implicit Types A : Type.
Local Hint Extern 100 (NonExpansive _) => solve_proper : core.
(* Force implicit argument PROP *)
Notation "P ⊢ Q" := (P ⊢@{PROP} Q).
Notation "P ⊣⊢ Q" := (P ⊣⊢@{PROP} Q).
(* Derived stuff about the entailment *)
Global Instance entails_anti_sym : AntiSymm (⊣⊢) (@bi_entails PROP).
Proof. intros P Q ??. by apply equiv_entails. Qed.
Lemma equiv_entails_1_1 P Q : (P ⊣⊢ Q) (P Q).
Proof. apply equiv_entails. Qed.
Lemma equiv_entails_1_2 P Q : (P ⊣⊢ Q) (Q P).
Proof. apply equiv_entails. Qed.
Lemma equiv_entails_2 P Q : (P Q) (Q P) (P ⊣⊢ Q).
Proof. intros. by apply equiv_entails. Qed.
Global Instance entails_proper :
Proper ((⊣⊢) ==> (⊣⊢) ==> iff) (() : relation PROP).
Proof.
move => P1 P2 /equiv_entails [HP1 HP2] Q1 Q2 /equiv_entails [HQ1 HQ2]; split=>?.
- by trans P1; [|trans Q1].
- by trans P2; [|trans Q2].
Qed.
Lemma entails_equiv_l P Q R : (P ⊣⊢ Q) (Q R) (P R).
Proof. by intros ->. Qed.
Lemma entails_equiv_r P Q R : (P Q) (Q ⊣⊢ R) (P R).
Proof. by intros ? <-. Qed.
Global Instance bi_emp_valid_proper : Proper ((⊣⊢) ==> iff) (@bi_emp_valid PROP).
Proof. solve_proper. Qed.
Global Instance bi_emp_valid_mono : Proper (() ==> impl) (@bi_emp_valid PROP).
Proof. solve_proper. Qed.
Global Instance bi_emp_valid_flip_mono :
Proper (flip () ==> flip impl) (@bi_emp_valid PROP).
Proof. solve_proper. Qed.
(* Propers *)
Global Instance pure_proper : Proper (iff ==> (⊣⊢)) (@bi_pure PROP) | 0.
Proof. intros φ1 φ2 . apply equiv_dist=> n. by apply pure_ne. Qed.
Global Instance and_proper :
Proper ((⊣⊢) ==> (⊣⊢) ==> (⊣⊢)) (@bi_and PROP) := ne_proper_2 _.
Global Instance or_proper :
Proper ((⊣⊢) ==> (⊣⊢) ==> (⊣⊢)) (@bi_or PROP) := ne_proper_2 _.
Global Instance impl_proper :
Proper ((⊣⊢) ==> (⊣⊢) ==> (⊣⊢)) (@bi_impl PROP) := ne_proper_2 _.
Global Instance sep_proper :
Proper ((⊣⊢) ==> (⊣⊢) ==> (⊣⊢)) (@bi_sep PROP) := ne_proper_2 _.
Global Instance wand_proper :
Proper ((⊣⊢) ==> (⊣⊢) ==> (⊣⊢)) (@bi_wand PROP) := ne_proper_2 _.
Global Instance forall_proper A :
Proper (pointwise_relation _ (⊣⊢) ==> (⊣⊢)) (@bi_forall PROP A).
Proof.
intros Φ1 Φ2 . apply equiv_dist=> n.
apply forall_ne=> x. apply equiv_dist, .
Qed.
Global Instance exist_proper A :
Proper (pointwise_relation _ (⊣⊢) ==> (⊣⊢)) (@bi_exist PROP A).
Proof.
intros Φ1 Φ2 . apply equiv_dist=> n.
apply exist_ne=> x. apply equiv_dist, .
Qed.
Global Instance persistently_proper :
Proper ((⊣⊢) ==> (⊣⊢)) (@bi_persistently PROP) := ne_proper _.
(* Derived logical stuff *)
Lemma and_elim_l' P Q R : (P R) P Q R.
Proof. by rewrite and_elim_l. Qed.
Lemma and_elim_r' P Q R : (Q R) P Q R.
Proof. by rewrite and_elim_r. Qed.
Lemma or_intro_l' P Q R : (P Q) P Q R.
Proof. intros ->; apply or_intro_l. Qed.
Lemma or_intro_r' P Q R : (P R) P Q R.
Proof. intros ->; apply or_intro_r. Qed.
Lemma exist_intro' {A} P (Ψ : A PROP) a : (P Ψ a) P a, Ψ a.
Proof. intros ->; apply exist_intro. Qed.
Lemma forall_elim' {A} P (Ψ : A PROP) : (P a, Ψ a) a, P Ψ a.
Proof. move=> HP a. by rewrite HP forall_elim. Qed.
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.
Lemma impl_intro_l P Q R : (Q P R) P Q R.
Proof. intros HR; apply impl_intro_r; rewrite -HR; auto. Qed.
Lemma impl_elim P Q R : (P Q R) (P Q) P R.
Proof. intros. rewrite -(impl_elim_l' P Q R); auto. Qed.
Lemma impl_elim_r' P Q R : (Q P R) P Q R.
Proof. intros; apply impl_elim with P; auto. Qed.
Lemma impl_elim_l P Q : (P Q) P Q.
Proof. by apply impl_elim_l'. Qed.
Lemma impl_elim_r P Q : P (P Q) Q.
Proof. by apply impl_elim_r'. Qed.
Lemma False_elim P : False P.
Proof. by apply (pure_elim' False). Qed.
Lemma True_intro P : P True.
Proof. by apply pure_intro. Qed.
Local Hint Immediate False_elim : core.
Lemma entails_eq_True P Q : (P Q) ((P Q)%I True%I).
Proof.
split=>EQ.
- apply bi.equiv_entails; split; [by apply True_intro|].
apply impl_intro_r. rewrite and_elim_r //.
- trans (P True)%I.
+ apply and_intro; first done. by apply pure_intro.
+ rewrite -EQ impl_elim_r. done.
Qed.
Lemma entails_impl_True P Q : (P Q) (True (P Q)).
Proof. rewrite entails_eq_True equiv_entails; naive_solver. Qed.
Lemma and_mono P P' Q Q' : (P Q) (P' Q') P P' Q Q'.
Proof. auto. Qed.
Lemma and_mono_l P P' Q : (P Q) P P' Q P'.
Proof. by intros; apply and_mono. Qed.
Lemma and_mono_r P P' Q' : (P' Q') P P' P Q'.
Proof. by apply and_mono. Qed.
Lemma or_mono P P' Q Q' : (P Q) (P' Q') P P' Q Q'.
Proof. auto. Qed.
Lemma or_mono_l P P' Q : (P Q) P P' Q P'.
Proof. by intros; apply or_mono. Qed.
Lemma or_mono_r P P' Q' : (P' Q') P P' P Q'.
Proof. by apply or_mono. Qed.
Lemma impl_mono P P' Q Q' : (Q P) (P' Q') (P P') Q Q'.
Proof.
intros HP HQ'; apply impl_intro_l; rewrite -HQ'.
apply impl_elim with P; eauto.
Qed.
Lemma forall_mono {A} (Φ Ψ : A PROP) :
( a, Φ a Ψ a) ( a, Φ a) a, Ψ a.
Proof.
intros HP. apply forall_intro=> a; rewrite -(HP a); apply forall_elim.
Qed.
Lemma exist_mono {A} (Φ Ψ : A PROP) :
( a, Φ a Ψ a) ( a, Φ a) a, Ψ a.
Proof. intros . apply exist_elim=> a; rewrite ( a); apply exist_intro. Qed.
Global Instance and_mono' : Proper (() ==> () ==> ()) (@bi_and PROP).
Proof. by intros P P' HP Q Q' HQ; apply and_mono. Qed.
Global Instance and_flip_mono' :
Proper (flip () ==> flip () ==> flip ()) (@bi_and PROP).
Proof. by intros P P' HP Q Q' HQ; apply and_mono. Qed.
Global Instance or_mono' : Proper (() ==> () ==> ()) (@bi_or PROP).
Proof. by intros P P' HP Q Q' HQ; apply or_mono. Qed.
Global Instance or_flip_mono' :
Proper (flip () ==> flip () ==> flip ()) (@bi_or PROP).
Proof. by intros P P' HP Q Q' HQ; apply or_mono. Qed.
Global Instance impl_mono' :
Proper (flip () ==> () ==> ()) (@bi_impl PROP).
Proof. by intros P P' HP Q Q' HQ; apply impl_mono. Qed.
Global Instance impl_flip_mono' :
Proper (() ==> flip () ==> flip ()) (@bi_impl PROP).
Proof. by intros P P' HP Q Q' HQ; apply impl_mono. Qed.
Global Instance forall_mono' A :
Proper (pointwise_relation _ () ==> ()) (@bi_forall PROP A).
Proof. intros P1 P2; apply forall_mono. Qed.
Global Instance forall_flip_mono' A :
Proper (pointwise_relation _ (flip ()) ==> flip ()) (@bi_forall PROP A).
Proof. intros P1 P2; apply forall_mono. Qed.
Global Instance exist_mono' A :
Proper (pointwise_relation _ (()) ==> ()) (@bi_exist PROP A).
Proof. intros P1 P2; apply exist_mono. Qed.
Global Instance exist_flip_mono' A :
Proper (pointwise_relation _ (flip ()) ==> flip ()) (@bi_exist PROP A).
Proof. intros P1 P2; apply exist_mono. Qed.
Global Instance and_idem : IdemP (⊣⊢) (@bi_and PROP).
Proof. intros P; apply (anti_symm ()); auto. Qed.
Global Instance or_idem : IdemP (⊣⊢) (@bi_or PROP).
Proof. intros P; apply (anti_symm ()); auto. Qed.
Global Instance and_comm : Comm (⊣⊢) (@bi_and PROP).
Proof. intros P Q; apply (anti_symm ()); auto. Qed.
Global Instance True_and : LeftId (⊣⊢) True%I (@bi_and PROP).
Proof. intros P; apply (anti_symm ()); auto. Qed.
Global Instance and_True : RightId (⊣⊢) True%I (@bi_and PROP).
Proof. intros P; apply (anti_symm ()); auto. Qed.
Global Instance False_and : LeftAbsorb (⊣⊢) False%I (@bi_and PROP).
Proof. intros P; apply (anti_symm ()); auto. Qed.
Global Instance and_False : RightAbsorb (⊣⊢) False%I (@bi_and PROP).
Proof. intros P; apply (anti_symm ()); auto. Qed.
Global Instance True_or : LeftAbsorb (⊣⊢) True%I (@bi_or PROP).
Proof. intros P; apply (anti_symm ()); auto. Qed.
Global Instance or_True : RightAbsorb (⊣⊢) True%I (@bi_or PROP).
Proof. intros P; apply (anti_symm ()); auto. Qed.
Global Instance False_or : LeftId (⊣⊢) False%I (@bi_or PROP).
Proof. intros P; apply (anti_symm ()); auto. Qed.
Global Instance or_False : RightId (⊣⊢) False%I (@bi_or PROP).
Proof. intros P; apply (anti_symm ()); auto. Qed.
Global Instance and_assoc : Assoc (⊣⊢) (@bi_and PROP).
Proof. intros P Q R; apply (anti_symm ()); auto. Qed.
Global Instance or_comm : Comm (⊣⊢) (@bi_or PROP).
Proof. intros P Q; apply (anti_symm ()); auto. Qed.
Global Instance or_assoc : Assoc (⊣⊢) (@bi_or PROP).
Proof. intros P Q R; apply (anti_symm ()); auto. Qed.
Global Instance True_impl : LeftId (⊣⊢) True%I (@bi_impl PROP).
Proof.
intros P; apply (anti_symm ()).
- by rewrite -(left_id True%I ()%I (_ _)%I) impl_elim_r.
- by apply impl_intro_l; rewrite left_id.
Qed.
Lemma impl_refl P Q : Q P P.
Proof. apply impl_intro_l, and_elim_l. Qed.
Lemma impl_trans P Q R : (P Q) (Q R) (P R).
Proof. apply impl_intro_l. by rewrite assoc !impl_elim_r. Qed.
Lemma False_impl P : (False P) ⊣⊢ True.
Proof.
apply (anti_symm ()); [by auto|].
apply impl_intro_l. rewrite left_absorb. auto.
Qed.
Lemma exist_impl_forall {A} P (Ψ : A PROP) :
(( x : A, Ψ x) P) ⊣⊢ x : A, Ψ x P.
Proof.
apply equiv_entails; split.
- apply forall_intro=>x. by rewrite -exist_intro.
- apply impl_intro_r, impl_elim_r', exist_elim=>x.
apply impl_intro_r. by rewrite (forall_elim x) impl_elim_r.
Qed.
Lemma forall_unit (Ψ : unit PROP) :
( x, Ψ x) ⊣⊢ Ψ ().
Proof.
apply (anti_symm ()).
- rewrite (forall_elim ()) //.
- apply forall_intro=>[[]]. done.
Qed.
Lemma exist_unit (Ψ : unit PROP) :
( x, Ψ x) ⊣⊢ Ψ ().
Proof.
apply (anti_symm ()).
- apply exist_elim=>[[]]. done.
- rewrite -(exist_intro ()). done.
Qed.
Lemma exist_exist {A B} (Ψ : A B PROP) :
( x y, Ψ x y) ⊣⊢ ( y x, Ψ x y).
Proof.
apply (anti_symm ());
do 2 (apply exist_elim=>?); rewrite -2!exist_intro; eauto.
Qed.
Lemma forall_forall {A B} (Ψ : A B PROP) :
( x y, Ψ x y) ⊣⊢ ( y x, Ψ x y).
Proof.
apply (anti_symm ());
do 2 (apply forall_intro=>?); rewrite 2!forall_elim; eauto.
Qed.
Lemma exist_forall {A B} (Ψ : A B PROP) :
( x, y, Ψ x y) ( y, x, Ψ x y).
Proof.
apply forall_intro=>?. apply exist_elim=>?.
rewrite -exist_intro forall_elim ; eauto.
Qed.
Lemma impl_curry P Q R : (P Q R) ⊣⊢ (P Q R).
Proof.
apply (anti_symm _).
- apply impl_intro_l. by rewrite (comm _ P) -and_assoc !impl_elim_r.
- do 2 apply impl_intro_l. by rewrite assoc (comm _ Q) impl_elim_r.
Qed.
Lemma or_and_l P Q R : P Q R ⊣⊢ (P Q) (P R).
Proof.
apply (anti_symm ()); first auto.
do 2 (apply impl_elim_l', or_elim; apply impl_intro_l); auto.
Qed.
Lemma or_and_r P Q R : P Q R ⊣⊢ (P R) (Q R).
Proof. by rewrite -!(comm _ R) or_and_l. Qed.
Lemma and_or_l P Q R : P (Q R) ⊣⊢ P Q P R.
Proof.
apply (anti_symm ()); last auto.
apply impl_elim_r', or_elim; apply impl_intro_l; auto.
Qed.
Lemma and_or_r P Q R : (P Q) R ⊣⊢ P R Q R.
Proof. by rewrite -!(comm _ R) and_or_l. Qed.
Lemma and_exist_l {A} P (Ψ : A PROP) : P ( a, Ψ a) ⊣⊢ a, P Ψ a.
Proof.
apply (anti_symm ()).
- apply impl_elim_r'. apply exist_elim=>a. apply impl_intro_l.
by rewrite -(exist_intro a).
- apply exist_elim=>a. apply and_intro; first by rewrite and_elim_l.
by rewrite -(exist_intro a) and_elim_r.
Qed.
Lemma and_exist_r {A} P (Φ: A PROP) : ( a, Φ a) P ⊣⊢ a, Φ a P.
Proof.
rewrite -(comm _ P) and_exist_l. apply exist_proper=>a. by rewrite comm.
Qed.
Lemma or_exist {A} (Φ Ψ : A PROP) :
( a, Φ a Ψ a) ⊣⊢ ( a, Φ a) ( a, Ψ a).
Proof.
apply (anti_symm ()).
- apply exist_elim=> a. by rewrite -!(exist_intro a).
- apply or_elim; apply exist_elim=> a; rewrite -(exist_intro a); auto.
Qed.
Lemma and_alt P Q : P Q ⊣⊢ b : bool, if b then P else Q.
Proof.
apply (anti_symm _); first apply forall_intro=> -[]; auto.
by apply and_intro; [rewrite (forall_elim true)|rewrite (forall_elim false)].
Qed.
Lemma or_alt P Q : P Q ⊣⊢ b : bool, if b then P else Q.
Proof.
apply (anti_symm _); last apply exist_elim=> -[]; auto.
by apply or_elim; [rewrite -(exist_intro true)|rewrite -(exist_intro false)].
Qed.
Lemma entails_equiv_and P Q : (P ⊣⊢ Q P) (P Q).
Proof.
split.
- intros ->; auto.
- intros; apply (anti_symm _); auto.
Qed.
Global Instance iff_ne : NonExpansive2 (@bi_iff PROP).
Proof. unfold bi_iff; solve_proper. Qed.
Global Instance iff_proper :
Proper ((⊣⊢) ==> (⊣⊢) ==> (⊣⊢)) (@bi_iff PROP) := ne_proper_2 _.
Lemma iff_refl Q P : Q P P.
Proof. rewrite /bi_iff. apply and_intro; apply impl_refl. Qed.
Lemma iff_sym P Q : (P Q) ⊣⊢ (Q P).
Proof.
apply equiv_entails. split; apply and_intro;
try apply and_elim_r; apply and_elim_l.
Qed.
Lemma iff_trans P Q R : (P Q) (Q R) (P R).
Proof.
apply and_intro.
- rewrite /bi_iff (and_elim_l _ (_ _)) (and_elim_l _ (_ _)).
apply impl_trans.
- rewrite /bi_iff (and_elim_r _ (_ _)) (and_elim_r _ (_ _)) comm.
apply impl_trans.
Qed.
(* BI Stuff *)
Local Hint Resolve sep_mono : core.
Lemma sep_mono_l P P' Q : (P Q) P P' Q P'.
Proof. by intros; apply sep_mono. Qed.
Lemma sep_mono_r P P' Q' : (P' Q') P P' P Q'.
Proof. by apply sep_mono. Qed.
Global Instance sep_mono' : Proper (() ==> () ==> ()) (@bi_sep PROP).
Proof. by intros P P' HP Q Q' HQ; apply sep_mono. Qed.
Global Instance sep_flip_mono' :
Proper (flip () ==> flip () ==> flip ()) (@bi_sep PROP).
Proof. by intros P P' HP Q Q' HQ; apply sep_mono. Qed.
Lemma wand_mono P P' Q Q' : (Q P) (P' Q') (P -∗ P') Q -∗ Q'.
Proof.
intros HP HQ; apply wand_intro_r. rewrite HP -HQ. by apply wand_elim_l'.
Qed.
Global Instance wand_mono' : Proper (flip () ==> () ==> ()) (@bi_wand PROP).
Proof. by intros P P' HP Q Q' HQ; apply wand_mono. Qed.
Global Instance wand_flip_mono' :
Proper (() ==> flip () ==> flip ()) (@bi_wand PROP).
Proof. by intros P P' HP Q Q' HQ; apply wand_mono. Qed.
Global Instance sep_comm : Comm (⊣⊢) (@bi_sep PROP).
Proof. intros P Q; apply (anti_symm _); auto using sep_comm'. Qed.
Global Instance sep_assoc : Assoc (⊣⊢) (@bi_sep PROP).
Proof.
intros P Q R; apply (anti_symm _); auto using sep_assoc'.
by rewrite !(comm _ P) !(comm _ _ R) sep_assoc'.
Qed.
Global Instance emp_sep : LeftId (⊣⊢) emp%I (@bi_sep PROP).
Proof. intros P; apply (anti_symm _); auto using emp_sep_1, emp_sep_2. Qed.
Global Instance sep_emp : RightId (⊣⊢) emp%I (@bi_sep PROP).
Proof. by intros P; rewrite comm left_id. Qed.
Global Instance sep_False : LeftAbsorb (⊣⊢) False%I (@bi_sep PROP).
Proof. intros P; apply (anti_symm _); auto using wand_elim_l'. Qed.
Global Instance False_sep : RightAbsorb (⊣⊢) False%I (@bi_sep PROP).
Proof. intros P. by rewrite comm left_absorb. Qed.
Lemma True_sep_2 P : P True P.
Proof. rewrite -{1}[P](left_id emp%I bi_sep). auto using sep_mono. Qed.
Lemma sep_True_2 P : P P True.
Proof. by rewrite comm -True_sep_2. Qed.
Lemma sep_intro_emp_valid_l P Q R : ( P) (R Q) R P Q.
Proof. intros ? ->. rewrite -{1}(left_id emp%I _ Q). by apply sep_mono. Qed.
Lemma sep_intro_emp_valid_r P Q R : (R P) ( Q) R P Q.
Proof. intros -> ?. rewrite comm. by apply sep_intro_emp_valid_l. Qed.
Lemma sep_elim_emp_valid_l P Q R : ( P) (P R Q) R Q.
Proof. intros <- <-. by rewrite left_id. Qed.
Lemma sep_elim_emp_valid_r P Q R : (P) (R P Q) R Q.
Proof. intros <- <-. by rewrite right_id. Qed.
Lemma wand_intro_l P Q R : (Q P R) P Q -∗ R.
Proof. rewrite comm; apply wand_intro_r. Qed.
Lemma wand_elim_l P Q : (P -∗ Q) P Q.
Proof. by apply wand_elim_l'. Qed.
Lemma wand_elim_r P Q : P (P -∗ Q) Q.
Proof. rewrite (comm _ P); apply wand_elim_l. Qed.
Lemma wand_elim_r' P Q R : (Q P -∗ R) P Q R.
Proof. intros ->; apply wand_elim_r. Qed.
Lemma wand_apply P Q R S : (P Q -∗ R) (S P Q) S R.
Proof. intros HR%wand_elim_l' HQ. by rewrite HQ. Qed.
Lemma wand_frame_l P Q R : (Q -∗ R) P Q -∗ P R.
Proof. apply wand_intro_l. rewrite -assoc. apply sep_mono_r, wand_elim_r. Qed.
Lemma wand_frame_r P Q R : (Q -∗ R) Q P -∗ R P.
Proof.
apply wand_intro_l. rewrite ![(_ P)%I]comm -assoc.
apply sep_mono_r, wand_elim_r.
Qed.
Global Instance emp_wand : LeftId (⊣⊢) emp%I (@bi_wand PROP).
Proof.
intros P. apply (anti_symm _).
- by rewrite -[(emp -∗ P)%I]left_id wand_elim_r.
- apply wand_intro_l. by rewrite left_id.
Qed.
Lemma False_wand P : (False -∗ P) ⊣⊢ True.
Proof.
apply (anti_symm ()); [by auto|].
apply wand_intro_l. rewrite left_absorb. auto.
Qed.
Lemma wand_refl P : P -∗ P.
Proof. apply wand_intro_l. by rewrite right_id. Qed.
Lemma wand_trans P Q R : (P -∗ Q) (Q -∗ R) (P -∗ R).
Proof. apply wand_intro_l. by rewrite assoc !wand_elim_r. Qed.
Lemma wand_curry P Q R : (P -∗ Q -∗ R) ⊣⊢ (P Q -∗ R).
Proof.
apply (anti_symm _).
- apply wand_intro_l. by rewrite (comm _ P) -assoc !wand_elim_r.
- do 2 apply wand_intro_l. by rewrite assoc (comm _ Q) wand_elim_r.
Qed.
Lemma sep_and_l P Q R : P (Q R) (P Q) (P R).
Proof. auto. Qed.
Lemma sep_and_r P Q R : (P Q) R (P R) (Q R).
Proof. auto. Qed.
Lemma sep_or_l P Q R : P (Q R) ⊣⊢ (P Q) (P R).
Proof.
apply (anti_symm ()); last by eauto 8.
apply wand_elim_r', or_elim; apply wand_intro_l; auto.
Qed.
Lemma sep_or_r P Q R : (P Q) R ⊣⊢ (P R) (Q R).
Proof. by rewrite -!(comm _ R) sep_or_l. Qed.
Lemma sep_exist_l {A} P (Ψ : A PROP) : P ( a, Ψ a) ⊣⊢ a, P Ψ a.
Proof.
intros; apply (anti_symm ()).
- apply wand_elim_r', exist_elim=>a. apply wand_intro_l.
by rewrite -(exist_intro a).
- apply exist_elim=> a; apply sep_mono; auto using exist_intro.
Qed.
Lemma sep_exist_r {A} (Φ: A PROP) Q: ( a, Φ a) Q ⊣⊢ a, Φ a Q.
Proof. setoid_rewrite (comm _ _ Q); apply sep_exist_l. Qed.
Lemma sep_forall_l {A} P (Ψ : A PROP) : P ( a, Ψ a) a, P Ψ a.
Proof. by apply forall_intro=> a; rewrite forall_elim. Qed.
Lemma sep_forall_r {A} (Φ : A PROP) Q : ( a, Φ a) Q a, Φ a Q.
Proof. by apply forall_intro=> a; rewrite forall_elim. Qed.
Lemma exist_wand_forall {A} P (Ψ : A PROP) :
(( x : A, Ψ x) -∗ P) ⊣⊢ x : A, Ψ x -∗ P.
Proof.
apply equiv_entails; split.
- apply forall_intro=>x. by rewrite -exist_intro.
- apply wand_intro_r, wand_elim_r', exist_elim=>x.
apply wand_intro_r. by rewrite (forall_elim x) wand_elim_r.
Qed.
Global Instance wand_iff_ne : NonExpansive2 (@bi_wand_iff PROP).
Proof. solve_proper. Qed.
Global Instance wand_iff_proper :
Proper ((⊣⊢) ==> (⊣⊢) ==> (⊣⊢)) (@bi_wand_iff PROP) := ne_proper_2 _.
Lemma wand_iff_refl P : P ∗-∗ P.
Proof. apply and_intro; apply wand_intro_l; by rewrite right_id. Qed.
Lemma wand_iff_sym P Q : (P ∗-∗ Q) ⊣⊢ (Q ∗-∗ P).
Proof. apply equiv_entails; split; rewrite /bi_wand_iff; eauto. Qed.
Lemma wand_iff_trans P Q R : (P ∗-∗ Q) (Q ∗-∗ R) (P ∗-∗ R).
Proof.
apply and_intro.
- rewrite /bi_wand_iff !and_elim_l. apply wand_trans.
- rewrite /bi_wand_iff !and_elim_r comm. apply wand_trans.
Qed.
Lemma wand_entails P Q : ( P -∗ Q) P Q.
Proof. intros. rewrite -[P]emp_sep. by apply wand_elim_l'. Qed.
Lemma entails_wand P Q : (P Q) P -∗ Q.
Proof. intros ->. apply wand_intro_r. by rewrite left_id. Qed.
(* A version that works with rewrite, in which bi_emp_valid is unfolded. *)
Lemma entails_wand' P Q : (P Q) emp (P -∗ Q).
Proof. apply entails_wand. Qed.
Lemma wand_entails' P Q : (emp (P -∗ Q)) P Q.
Proof. apply wand_entails. Qed.
Lemma equiv_wand_iff P Q : (P ⊣⊢ Q) P ∗-∗ Q.
Proof. intros ->; apply wand_iff_refl. Qed.
Lemma wand_iff_equiv P Q : ( P ∗-∗ Q) (P ⊣⊢ Q).
Proof.
intros HPQ; apply (anti_symm ());
apply wand_entails; rewrite /bi_emp_valid HPQ /bi_wand_iff; auto.
Qed.
Lemma entails_impl P Q : (P Q) ( P Q).
Proof. intros ->. apply impl_intro_l. auto. Qed.
Lemma impl_entails P Q `{!Affine P} : ( P Q) P Q.
Proof. intros HPQ. apply impl_elim with P=>//. by rewrite {1}(affine P). Qed.
Lemma equiv_iff P Q : (P ⊣⊢ Q) ( P Q).
Proof. intros ->; apply iff_refl. Qed.
Lemma iff_equiv P Q `{!Affine P, !Affine Q} : ( P Q)%I (P ⊣⊢ Q).
Proof.
intros HPQ; apply (anti_symm ());
apply: impl_entails; rewrite /bi_emp_valid HPQ /bi_iff; auto.
Qed.
Lemma and_parallel P1 P2 Q1 Q2 :
(P1 P2) -∗ ((P1 -∗ Q1) (P2 -∗ Q2)) -∗ Q1 Q2.
Proof.
apply entails_wand, wand_intro_r, and_intro.
- rewrite !and_elim_l wand_elim_r. done.
- rewrite !and_elim_r wand_elim_r. done.
Qed.
Lemma wandM_sound (mP : option PROP) Q :
(mP -∗? Q) ⊣⊢ (default emp mP -∗ Q).
Proof. destruct mP; simpl; first done. rewrite emp_wand //. Qed.
(* Properties of the affinely modality *)
Global Instance affinely_ne : NonExpansive (@bi_affinely PROP).
Proof. solve_proper. Qed.
Global Instance affinely_proper : Proper ((⊣⊢) ==> (⊣⊢)) (@bi_affinely PROP).
Proof. solve_proper. Qed.
Global Instance affinely_mono' : Proper (() ==> ()) (@bi_affinely PROP).
Proof. solve_proper. Qed.
Global Instance affinely_flip_mono' :
Proper (flip () ==> flip ()) (@bi_affinely PROP).
Proof. solve_proper. Qed.
Lemma affinely_elim_emp P : <affine> P emp.
Proof. rewrite /bi_affinely; auto. Qed.
Global Instance affinely_affine P : Affine (<affine> P).
Proof. by rewrite /Affine affinely_elim_emp. Qed.
Lemma affinely_elim P : <affine> P P.
Proof. rewrite /bi_affinely; auto. Qed.
Lemma affinely_mono P Q : (P Q) <affine> P <affine> Q.
Proof. by intros ->. Qed.
Lemma affinely_idemp P : <affine> <affine> P ⊣⊢ <affine> P.
Proof. by rewrite /bi_affinely assoc idemp. Qed.
Lemma affinely_intro' P Q : (<affine> P Q) <affine> P <affine> Q.
Proof. intros <-. by rewrite affinely_idemp. Qed.
Lemma affinely_False : <affine> False ⊣⊢ False.
Proof. by rewrite /bi_affinely right_absorb. Qed.
Lemma affinely_emp : <affine> emp ⊣⊢ emp.
Proof. by rewrite /bi_affinely (idemp bi_and). Qed.
Lemma affinely_or P Q : <affine> (P Q) ⊣⊢ <affine> P <affine> Q.
Proof. by rewrite /bi_affinely and_or_l. Qed.
Lemma affinely_and P Q : <affine> (P Q) ⊣⊢ <affine> P <affine> Q.
Proof.
rewrite /bi_affinely -(comm _ P) (assoc _ (_ _)%I) -!(assoc _ P).
by rewrite idemp !assoc (comm _ P).
Qed.
Lemma affinely_sep_2 P Q : <affine> P <affine> Q <affine> (P Q).
Proof.
rewrite /bi_affinely. apply and_intro.
- by rewrite !and_elim_l right_id.
- by rewrite !and_elim_r.
Qed.
Lemma affinely_sep `{!BiPositive PROP} P Q :
<affine> (P Q) ⊣⊢ <affine> P <affine> Q.
Proof.
apply (anti_symm _), affinely_sep_2.
by rewrite -{1}affinely_idemp bi_positive !(comm _ (<affine> P)%I) bi_positive.
Qed.
Lemma affinely_forall {A} (Φ : A PROP) :
<affine> ( a, Φ a) a, <affine> (Φ a).
Proof. apply forall_intro=> a. by rewrite (forall_elim a). Qed.
Lemma affinely_exist {A} (Φ : A PROP) :
<affine> ( a, Φ a) ⊣⊢ a, <affine> (Φ a).
Proof. by rewrite /bi_affinely and_exist_l. Qed.
Lemma affinely_True_emp : <affine> True ⊣⊢ emp.
Proof. apply (anti_symm _); rewrite /bi_affinely; auto. Qed.
Lemma affinely_and_l P Q : <affine> P Q ⊣⊢ <affine> (P Q).
Proof. by rewrite /bi_affinely assoc. Qed.
Lemma affinely_and_r P Q : P <affine> Q ⊣⊢ <affine> (P Q).
Proof. by rewrite /bi_affinely !assoc (comm _ P). Qed.
Lemma affinely_and_lr P Q : <affine> P Q ⊣⊢ P <affine> Q.
Proof. by rewrite affinely_and_l affinely_and_r. Qed.
(* Affine instances *)
Global Instance emp_affine : Affine (PROP:=PROP) emp.
Proof. by rewrite /Affine. Qed.
Global Instance False_affine : Affine (PROP:=PROP) False.
Proof. by rewrite /Affine False_elim. Qed.
Global Instance and_affine_l P Q : Affine P Affine (P Q).
Proof. rewrite /Affine=> ->; auto. Qed.
Global Instance and_affine_r P Q : Affine Q Affine (P Q).
Proof. rewrite /Affine=> ->; auto. Qed.
Global Instance or_affine P Q : Affine P Affine Q Affine (P Q).
Proof. rewrite /Affine=> -> ->; auto. Qed.
Global Instance forall_affine `{Inhabited A} (Φ : A PROP) :
( x, Affine (Φ x)) Affine ( x, Φ x).
Proof. intros. rewrite /Affine (forall_elim inhabitant). apply: affine. Qed.
Global Instance exist_affine {A} (Φ : A PROP) :
( x, Affine (Φ x)) Affine ( x, Φ x).
Proof. rewrite /Affine=> H. apply exist_elim=> a. by rewrite H. Qed.
Global Instance sep_affine P Q : Affine P Affine Q Affine (P Q).
Proof. rewrite /Affine=>-> ->. by rewrite left_id. Qed.
(* Properties of the absorbingly modality *)
Global Instance absorbingly_ne : NonExpansive (@bi_absorbingly PROP).
Proof. solve_proper. Qed.
Global Instance absorbingly_proper : Proper ((⊣⊢) ==> (⊣⊢)) (@bi_absorbingly PROP).
Proof. solve_proper. Qed.
Global Instance absorbingly_mono' : Proper (() ==> ()) (@bi_absorbingly PROP).
Proof. solve_proper. Qed.
Global Instance absorbingly_flip_mono' :
Proper (flip () ==> flip ()) (@bi_absorbingly PROP).
Proof. solve_proper. Qed.
Lemma absorbingly_intro P : P <absorb> P.
Proof. by rewrite /bi_absorbingly -True_sep_2. Qed.
Lemma absorbingly_mono P Q : (P Q) <absorb> P <absorb> Q.
Proof. by intros ->. Qed.
Lemma absorbingly_idemp P : <absorb> <absorb> P ⊣⊢ <absorb> P.
Proof.
apply (anti_symm _), absorbingly_intro.
rewrite /bi_absorbingly assoc. apply sep_mono; auto.
Qed.
Global Instance absorbingly_absorbing P : Absorbing (<absorb> P).
Proof. by rewrite /Absorbing absorbingly_idemp. Qed.
Lemma absorbingly_pure φ : <absorb> φ ⊣⊢ φ ⌝.
Proof.
apply (anti_symm _), absorbingly_intro.
apply wand_elim_r', pure_elim'=> ?. apply wand_intro_l; auto.
Qed.
Lemma absorbingly_True : <absorb> True ⊣⊢ True.
Proof. apply absorbingly_pure. Qed.
Lemma absorbingly_or P Q : <absorb> (P Q) ⊣⊢ <absorb> P <absorb> Q.
Proof. by rewrite /bi_absorbingly sep_or_l. Qed.
Lemma absorbingly_and_1 P Q : <absorb> (P Q) <absorb> P <absorb> Q.
Proof. apply and_intro; apply absorbingly_mono; auto. Qed.
Lemma absorbingly_forall {A} (Φ : A PROP) :
<absorb> ( a, Φ a) a, <absorb> (Φ a).
Proof. apply forall_intro=> a. by rewrite (forall_elim a). Qed.
Lemma absorbingly_exist {A} (Φ : A PROP) :
<absorb> ( a, Φ a) ⊣⊢ a, <absorb> (Φ a).
Proof. by rewrite /bi_absorbingly sep_exist_l. Qed.
Lemma absorbingly_sep P Q : <absorb> (P Q) ⊣⊢ <absorb> P <absorb> Q.
Proof.
by rewrite -{1}absorbingly_idemp /bi_absorbingly !assoc -!(comm _ P) !assoc.
Qed.
Lemma absorbingly_emp_True : <absorb> emp ⊣⊢ True.
Proof. rewrite /bi_absorbingly right_id //. Qed.
Lemma absorbingly_wand P Q : <absorb> (P -∗ Q) <absorb> P -∗ <absorb> Q.
Proof. apply wand_intro_l. by rewrite -absorbingly_sep wand_elim_r. Qed.
Lemma absorbingly_sep_l P Q : <absorb> P Q ⊣⊢ <absorb> (P Q).
Proof. by rewrite /bi_absorbingly assoc. Qed.
Lemma absorbingly_sep_r P Q : P <absorb> Q ⊣⊢ <absorb> (P Q).
Proof. by rewrite /bi_absorbingly !assoc (comm _ P). Qed.
Lemma absorbingly_sep_lr P Q : <absorb> P Q ⊣⊢ P <absorb> Q.
Proof. by rewrite absorbingly_sep_l absorbingly_sep_r. Qed.
Lemma affinely_absorbingly_elim `{!BiPositive PROP} P :
<affine> <absorb> P ⊣⊢ <affine> P.
Proof.
apply (anti_symm _), affinely_mono, absorbingly_intro.
by rewrite /bi_absorbingly affinely_sep affinely_True_emp left_id.
Qed.
(* Absorbing instances *)
Global Instance pure_absorbing φ : Absorbing (PROP:=PROP) φ⌝.
Proof. by rewrite /Absorbing absorbingly_pure. Qed.
Global Instance and_absorbing P Q : Absorbing P Absorbing Q Absorbing (P Q).
Proof. intros. by rewrite /Absorbing absorbingly_and_1 !absorbing. Qed.
Global Instance or_absorbing P Q : Absorbing P Absorbing Q Absorbing (P Q).
Proof. intros. by rewrite /Absorbing absorbingly_or !absorbing. Qed.
Global Instance forall_absorbing {A} (Φ : A PROP) :
( x, Absorbing (Φ x)) Absorbing ( x, Φ x).
Proof.
rewrite /Absorbing=> ?. rewrite absorbingly_forall. auto using forall_mono.
Qed.
Global Instance exist_absorbing {A} (Φ : A PROP) :
( x, Absorbing (Φ x)) Absorbing ( x, Φ x).
Proof.
rewrite /Absorbing=> ?. rewrite absorbingly_exist. auto using exist_mono.
Qed.
(* The instance for [Absorbing (P → Q)] is in the persistence section *)
Global Instance sep_absorbing_l P Q : Absorbing P Absorbing (P Q).
Proof. intros. by rewrite /Absorbing -absorbingly_sep_l absorbing. Qed.
Global Instance sep_absorbing_r P Q : Absorbing Q Absorbing (P Q).
Proof. intros. by rewrite /Absorbing -absorbingly_sep_r absorbing. Qed.
Global Instance wand_absorbing_l P Q : Absorbing P Absorbing (P -∗ Q).
Proof.
intros. rewrite /Absorbing. apply wand_intro_l.
by rewrite absorbingly_sep_r -absorbingly_sep_l absorbing wand_elim_r.
Qed.
Global Instance wand_absorbing_r P Q : Absorbing Q Absorbing (P -∗ Q).
Proof.
intros. by rewrite /Absorbing absorbingly_wand !absorbing -absorbingly_intro.
Qed.
(* Affine and absorbing propositions *)
Global Instance Affine_proper : Proper ((⊣⊢) ==> iff) (@Affine PROP).
Proof. solve_proper. Qed.
Global Instance Absorbing_proper : Proper ((⊣⊢) ==> iff) (@Absorbing PROP).
Proof. solve_proper. Qed.
Lemma affine_affinely P `{!Affine P} : <affine> P ⊣⊢ P.
Proof. rewrite /bi_affinely. apply (anti_symm _); auto. Qed.
Lemma absorbing_absorbingly P `{!Absorbing P} : <absorb> P ⊣⊢ P.
Proof. by apply (anti_symm _), absorbingly_intro. Qed.
Lemma True_affine_all_affine P : Affine (PROP:=PROP) True Affine P.
Proof. rewrite /Affine=> <-; auto. Qed.
Lemma emp_absorbing_all_absorbing P : Absorbing (PROP:=PROP) emp Absorbing P.
Proof.
intros. rewrite /Absorbing -{2}(emp_sep P).
rewrite -(absorbing emp) absorbingly_sep_l left_id //.
Qed.
Lemma sep_elim_l P Q `{HQP : TCOr (Affine Q) (Absorbing P)} : P Q P.
Proof.
destruct HQP.
- by rewrite (affine Q) right_id.
- by rewrite (True_intro Q) comm.
Qed.
Lemma sep_elim_r P Q `{TCOr (Affine P) (Absorbing Q)} : P Q Q.
Proof. by rewrite comm sep_elim_l. Qed.
Lemma sep_and P Q :
TCOr (Affine P) (Absorbing Q) TCOr (Affine Q) (Absorbing P)
P Q P Q.
Proof.
intros [?|?] [?|?];
apply and_intro; apply: sep_elim_l || apply: sep_elim_r.
Qed.
Lemma affinely_intro P Q `{!Affine P} : (P Q) P <affine> Q.
Proof. intros <-. by rewrite affine_affinely. Qed.
Lemma emp_and P `{!Affine P} : emp P ⊣⊢ P.
Proof. apply (anti_symm _); auto. Qed.
Lemma and_emp P `{!Affine P} : P emp ⊣⊢ P.
Proof. apply (anti_symm _); auto. Qed.
Lemma emp_or P `{!Affine P} : emp P ⊣⊢ emp.
Proof. apply (anti_symm _); auto. Qed.
Lemma or_emp P `{!Affine P} : P emp ⊣⊢ emp.
Proof. apply (anti_symm _); auto. Qed.
Lemma True_sep P `{!Absorbing P} : True P ⊣⊢ P.
Proof. apply (anti_symm _); auto using True_sep_2. Qed.
Lemma sep_True P `{!Absorbing P} : P True ⊣⊢ P.
Proof. by rewrite comm True_sep. Qed.
Lemma True_emp_iff_BiAffine :
BiAffine PROP (True emp).
Proof.
split.
- intros ?. exact: affine.
- rewrite /BiAffine /Affine=>Hemp ?. rewrite -Hemp.
exact: True_intro.
Qed.
Section bi_affine.
Context `{!BiAffine PROP}.
Global Instance bi_affine_absorbing P : Absorbing P | 0.
Proof. by rewrite /Absorbing /bi_absorbingly (affine True) left_id. Qed.
Global Instance bi_affine_positive : BiPositive PROP.
Proof. intros P Q. by rewrite !affine_affinely. Qed.
Lemma True_emp : True ⊣⊢ emp.
Proof. apply (anti_symm _); auto using affine. Qed.
Global Instance emp_and' : LeftId (⊣⊢) emp%I (@bi_and PROP).
Proof. intros P. by rewrite -True_emp left_id. Qed.
Global Instance and_emp' : RightId (⊣⊢) emp%I (@bi_and PROP).
Proof. intros P. by rewrite -True_emp right_id. Qed.
Global Instance True_sep' : LeftId (⊣⊢) True%I (@bi_sep PROP).
Proof. intros P. by rewrite True_emp left_id. Qed.
Global Instance sep_True' : RightId (⊣⊢) True%I (@bi_sep PROP).
Proof. intros P. by rewrite True_emp right_id. Qed.
Lemma impl_wand_1 P Q : (P Q) P -∗ Q.
Proof. apply wand_intro_l. by rewrite sep_and impl_elim_r. Qed.
End bi_affine.
(* Pure stuff *)
Lemma pure_elim φ Q R : (Q φ) (φ Q R) Q R.
Proof.
intros HQ HQR. rewrite -(idemp ()%I Q) {1}HQ.
apply impl_elim_l', pure_elim'=> ?. apply impl_intro_l.
rewrite and_elim_l; auto.
Qed.
Lemma pure_mono φ1 φ2 : (φ1 φ2) φ1 φ2⌝.
Proof. auto using pure_elim', pure_intro. Qed.
Global Instance pure_mono' : Proper (impl ==> ()) (@bi_pure PROP).
Proof. intros φ1 φ2; apply pure_mono. Qed.
Global Instance pure_flip_mono : Proper (flip impl ==> flip ()) (@bi_pure PROP).
Proof. intros φ1 φ2; apply pure_mono. Qed.
Lemma pure_iff φ1 φ2 : (φ1 φ2) φ1 ⊣⊢ φ2⌝.
Proof. intros [??]; apply (anti_symm _); auto using pure_mono. Qed.
Lemma pure_elim_l φ Q R : (φ Q R) φ Q R.
Proof. intros; apply pure_elim with φ; eauto. Qed.
Lemma pure_elim_r φ Q R : (φ Q R) Q φ R.
Proof. intros; apply pure_elim with φ; eauto. Qed.
Lemma pure_True (φ : Prop) : φ φ ⊣⊢ True.
Proof. intros; apply (anti_symm _); auto. Qed.
Lemma pure_False (φ : Prop) : ¬φ φ ⊣⊢ False.
Proof. intros; apply (anti_symm _); eauto using pure_mono. Qed.
Lemma pure_and φ1 φ2 : φ1 φ2 ⊣⊢ φ1 φ2⌝.
Proof.
apply (anti_symm _).
- apply and_intro; apply pure_mono; tauto.
- eapply (pure_elim φ1); [auto|]=> ?. rewrite and_elim_r. auto using pure_mono.
Qed.
Lemma pure_or φ1 φ2 : φ1 φ2 ⊣⊢ φ1 φ2⌝.
Proof.
apply (anti_symm _).
- eapply pure_elim=> // -[?|?]; auto using pure_mono.
- apply or_elim; eauto using pure_mono.
Qed.
Lemma pure_impl_1 φ1 φ2 : φ1 φ2 (φ1 φ2).
Proof. apply impl_intro_l. rewrite -pure_and. apply pure_mono. naive_solver. Qed.
Lemma pure_impl_2 `{!BiPureForall PROP} φ1 φ2 : (φ1 φ2) φ1 φ2⌝.
Proof.
rewrite -pure_forall_2. apply forall_intro=> ?.
by rewrite -(left_id True bi_and (__))%I (pure_True φ1) // impl_elim_r.
Qed.
Lemma pure_impl `{!BiPureForall PROP} φ1 φ2 : φ1 φ2 ⊣⊢ (φ1 φ2).
Proof. apply (anti_symm _); auto using pure_impl_1, pure_impl_2. Qed.
Lemma pure_forall_1 {A} (φ : A Prop) : ⌜∀ x, φ x x, φ x⌝.
Proof. apply forall_intro=> x. eauto using pure_mono. Qed.
Lemma pure_forall `{!BiPureForall PROP} {A} (φ : A Prop) :
⌜∀ x, φ x ⊣⊢ x, φ x⌝.
Proof. apply (anti_symm _); auto using pure_forall_1, pure_forall_2. Qed.
Lemma pure_exist {A} (φ : A Prop) : ⌜∃ x, φ x ⊣⊢ x, φ x⌝.
Proof.
apply (anti_symm _).
- eapply pure_elim=> // -[x ?]. rewrite -(exist_intro x); auto using pure_mono.
- apply exist_elim=> x. eauto using pure_mono.
Qed.
Lemma bi_pure_forall_em : ( φ : Prop, φ ¬φ) BiPureForall PROP.
Proof.
intros Hem A φ. destruct (Hem ( a, ¬φ a)) as [[a ]|].
{ rewrite (forall_elim a). by apply pure_elim'. }
apply pure_intro=> a. destruct (Hem (φ a)); naive_solver.
Qed.
Lemma pure_impl_forall φ P : (φ P) ⊣⊢ ( _ : φ, P).
Proof.
apply (anti_symm _).
- apply forall_intro=> ?. by rewrite pure_True // left_id.
- apply impl_intro_l, pure_elim_l=> . by rewrite (forall_elim ).
Qed.
Lemma pure_alt φ : φ ⊣⊢ _ : φ, True.
Proof.
apply (anti_symm _).
- eapply pure_elim; eauto=> H. rewrite -(exist_intro H); auto.
- by apply exist_elim, pure_intro.
Qed.
Lemma pure_wand_forall φ P `{!Absorbing P} : (φ -∗ P) ⊣⊢ ( _ : φ, P).
Proof.
apply (anti_symm _).
- apply forall_intro=> .
rewrite -(pure_intro φ emp) // emp_wand //.
- apply wand_intro_l, wand_elim_l', pure_elim'=> .
apply wand_intro_l. rewrite (forall_elim ) comm. by apply absorbing.
Qed.
Lemma decide_bi_True φ `{!Decision φ} (P : PROP) :
(if decide φ then P else True) ⊣⊢ (φ P).
Proof.
destruct (decide _).
- by rewrite pure_True // True_impl.
- by rewrite (pure_False φ) // False_impl.
Qed.
Lemma decide_emp `{!BiAffine PROP} φ `{!Decision φ} (P : PROP) :
(if decide φ then P else emp) ⊣⊢ (φ P).
Proof.
rewrite -decide_bi_True. destruct (decide _); [done|]. by rewrite True_emp.
Qed.
(* Properties of the persistence modality *)
Local Hint Resolve persistently_mono : core.
Global Instance persistently_mono' : Proper (() ==> ()) (@bi_persistently PROP).
Proof. intros P Q; apply persistently_mono. Qed.
Global Instance persistently_flip_mono' :
Proper (flip () ==> flip ()) (@bi_persistently PROP).
Proof. intros P Q; apply persistently_mono. Qed.
Global Instance persistently_persistent P : Persistent (<pers> P).
Proof. by rewrite /Persistent -persistently_idemp_2. Qed.
Lemma absorbingly_elim_persistently P : <absorb> <pers> P ⊣⊢ <pers> P.
Proof.
apply (anti_symm _), absorbingly_intro.
by rewrite /bi_absorbingly comm persistently_absorbing.
Qed.
Global Instance persistently_absorbing P : Absorbing (<pers> P).
Proof. by rewrite /Absorbing absorbingly_elim_persistently. Qed.
Lemma persistently_forall_1 {A} (Ψ : A PROP) :
<pers> ( a, Ψ a) a, <pers> (Ψ a).
Proof. apply forall_intro=> x. by rewrite (forall_elim x). Qed.
Lemma persistently_forall `{!BiPersistentlyForall PROP} {A} (Ψ : A PROP) :
<pers> ( a, Ψ a) ⊣⊢ a, <pers> (Ψ a).
Proof.
apply (anti_symm _); auto using persistently_forall_1, persistently_forall_2.
Qed.
Lemma persistently_exist {A} (Ψ : A PROP) :
<pers> ( a, Ψ a) ⊣⊢ a, <pers> (Ψ a).
Proof.
apply (anti_symm _); first by auto using persistently_exist_1.
apply exist_elim=> x. by rewrite (exist_intro x).
Qed.
Lemma persistently_and P Q : <pers> (P Q) ⊣⊢ <pers> P <pers> Q.
Proof. apply (anti_symm _); by auto using persistently_and_2. Qed.
Lemma persistently_or P Q : <pers> (P Q) ⊣⊢ <pers> P <pers> Q.
Proof. rewrite !or_alt persistently_exist. by apply exist_proper=> -[]. Qed.
Lemma persistently_impl P Q : <pers> (P Q) <pers> P <pers> Q.
Proof.
apply impl_intro_l; rewrite -persistently_and.
apply persistently_mono, impl_elim with P; auto.
Qed.
Lemma persistently_emp_intro P : P <pers> emp.
Proof.
rewrite -(left_id emp%I bi_sep P).
by rewrite {1}persistently_emp_2 sep_elim_l.
Qed.
Lemma persistently_True_emp : <pers> True ⊣⊢ <pers> emp.
Proof. apply (anti_symm _); auto using persistently_emp_intro. Qed.
Lemma persistently_True : <pers> True ⊣⊢ True.
Proof.
apply (anti_symm _); auto.
rewrite persistently_True_emp. apply persistently_emp_intro.
Qed.
Lemma persistently_and_emp P : <pers> P ⊣⊢ <pers> (emp P).
Proof.
apply (anti_symm ()); last by rewrite and_elim_r.
rewrite persistently_and. apply and_intro; last done.
apply persistently_emp_intro.
Qed.
Lemma persistently_and_sep_elim_emp P Q : <pers> P Q (emp P) Q.
Proof.
rewrite persistently_and_emp.
apply persistently_and_sep_elim.
Qed.
Lemma persistently_and_sep_assoc P Q R : <pers> P (Q R) ⊣⊢ (<pers> P Q) R.
Proof.
apply (anti_symm ()).
- rewrite {1}persistently_idemp_2 persistently_and_sep_elim_emp assoc.
apply sep_mono_l, and_intro.
+ by rewrite and_elim_r sep_elim_l.
+ by rewrite and_elim_l left_id.
- apply and_intro.
+ by rewrite and_elim_l sep_elim_l.
+ by rewrite and_elim_r.
Qed.
Lemma persistently_and_emp_elim P : emp <pers> P P.
Proof. by rewrite comm persistently_and_sep_elim_emp right_id and_elim_r. Qed.
Lemma persistently_into_absorbingly P : <pers> P <absorb> P.
Proof.
rewrite -(right_id True%I _ (<pers> _)%I) -{1}(emp_sep True%I).
rewrite persistently_and_sep_assoc.
rewrite (comm bi_and) persistently_and_emp_elim comm //.
Qed.
Lemma persistently_elim P `{!Absorbing P} : <pers> P P.
Proof. by rewrite persistently_into_absorbingly absorbing_absorbingly. Qed.
Lemma persistently_idemp_1 P : <pers> <pers> P <pers> P.
Proof.
by rewrite persistently_into_absorbingly absorbingly_elim_persistently.
Qed.
Lemma persistently_idemp P : <pers> <pers> P ⊣⊢ <pers> P.
Proof.
apply (anti_symm _); auto using persistently_idemp_1, persistently_idemp_2.
Qed.
Lemma persistently_intro' P Q : (<pers> P Q) <pers> P <pers> Q.
Proof. intros <-. apply persistently_idemp_2. Qed.
Lemma persistently_pure φ : <pers> φ ⊣⊢ φ⌝.
Proof.
apply (anti_symm _).
{ by rewrite persistently_into_absorbingly absorbingly_pure. }
apply pure_elim'=> . rewrite -persistently_True.
auto using persistently_mono, pure_intro.
Qed.
Lemma persistently_sep_dup P : <pers> P ⊣⊢ <pers> P <pers> P.
Proof.
apply (anti_symm _).
- rewrite -{1}(idemp bi_and (<pers> _)%I).
by rewrite -{2}(emp_sep (<pers> _)%I)
persistently_and_sep_assoc and_elim_l.
- by rewrite sep_elim_l.
Qed.
Lemma persistently_and_sep_l_1 P Q : <pers> P Q <pers> P Q.
Proof.
by rewrite -{1}(emp_sep Q) persistently_and_sep_assoc and_elim_l.
Qed.
Lemma persistently_and_sep_r_1 P Q : P <pers> Q P <pers> Q.
Proof. by rewrite !(comm _ P) persistently_and_sep_l_1. Qed.
Lemma persistently_and_sep P Q : <pers> (P Q) <pers> (P Q).
Proof.
rewrite persistently_and.
rewrite -{1}persistently_idemp -persistently_and -{1}(emp_sep Q).
by rewrite persistently_and_sep_assoc (comm bi_and) persistently_and_emp_elim.
Qed.
Lemma persistently_affinely_elim P : <pers> <affine> P ⊣⊢ <pers> P.
Proof.
by rewrite /bi_affinely persistently_and -persistently_True_emp
persistently_pure left_id.
Qed.
Lemma and_sep_persistently P Q : <pers> P <pers> Q ⊣⊢ <pers> P <pers> Q.
Proof.
apply (anti_symm _); auto using persistently_and_sep_l_1.
apply and_intro.
- by rewrite sep_elim_l.
- by rewrite sep_elim_r.
Qed.
Lemma persistently_sep_2 P Q : <pers> P <pers> Q <pers> (P Q).
Proof.
by rewrite -persistently_and_sep persistently_and -and_sep_persistently.
Qed.
Lemma persistently_sep `{!BiPositive PROP} P Q :
<pers> (P Q) ⊣⊢ <pers> P <pers> Q.
Proof.
apply (anti_symm _); auto using persistently_sep_2.
rewrite -persistently_affinely_elim affinely_sep -and_sep_persistently.
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 persistently_alt_fixpoint P :
<pers> P ⊣⊢ P <pers> P.
Proof.
apply (anti_symm _).
- rewrite -persistently_and_sep_elim. apply and_intro; done.
- by rewrite sep_elim_r.
Qed.
Lemma persistently_alt_fixpoint' P :
<pers> P ⊣⊢ <affine> P <pers> P.
Proof.
rewrite -{1}persistently_affinely_elim {1}persistently_alt_fixpoint
persistently_affinely_elim //.
Qed.
Lemma persistently_wand P Q : <pers> (P -∗ Q) <pers> P -∗ <pers> Q.
Proof. apply wand_intro_r. by rewrite persistently_sep_2 wand_elim_l. Qed.
Lemma persistently_entails_l P Q : (P <pers> Q) P <pers> Q P.
Proof. intros; rewrite -persistently_and_sep_l_1; auto. Qed.
Lemma persistently_entails_r P Q : (P <pers> Q) P P <pers> Q.
Proof. intros; rewrite -persistently_and_sep_r_1; auto. Qed.
Lemma persistently_impl_wand_2 P Q : <pers> (P -∗ Q) <pers> (P Q).
Proof.
apply persistently_intro', impl_intro_r.
rewrite -{2}(emp_sep P) persistently_and_sep_assoc.
by rewrite (comm bi_and) persistently_and_emp_elim wand_elim_l.
Qed.
Lemma impl_wand_persistently_2 P Q : (<pers> P -∗ Q) (<pers> P Q).
Proof. apply impl_intro_l. by rewrite persistently_and_sep_l_1 wand_elim_r. Qed.
Section persistently_affine_bi.
Context `{!BiAffine PROP}.
Lemma persistently_emp : <pers> emp ⊣⊢ emp.
Proof. by rewrite -!True_emp persistently_pure. Qed.
Lemma persistently_and_sep_l P Q : <pers> P Q ⊣⊢ <pers> P Q.
Proof.
apply (anti_symm ());
eauto using persistently_and_sep_l_1, sep_and with typeclass_instances.
Qed.
Lemma persistently_and_sep_r P Q : P <pers> Q ⊣⊢ P <pers> Q.
Proof. by rewrite !(comm _ P) persistently_and_sep_l. Qed.
Lemma persistently_impl_wand P Q : <pers> (P Q) ⊣⊢ <pers> (P -∗ Q).
Proof.
apply (anti_symm ()); auto using persistently_impl_wand_2.
apply persistently_intro', wand_intro_l.
by rewrite -persistently_and_sep_r persistently_elim impl_elim_r.
Qed.
Lemma impl_wand_persistently P Q : (<pers> P Q) ⊣⊢ (<pers> P -∗ Q).
Proof.
apply (anti_symm ()).
- by rewrite -impl_wand_1.
- apply impl_wand_persistently_2.
Qed.
Lemma wand_alt P Q : (P -∗ Q) ⊣⊢ R, R <pers> (P R Q).
Proof.
apply (anti_symm ()).
- rewrite -(right_id True%I bi_sep (P -∗ Q)%I) -(exist_intro (P -∗ Q)%I).
apply sep_mono_r. rewrite -persistently_pure.
apply persistently_intro', impl_intro_l.
by rewrite wand_elim_r persistently_pure right_id.
- apply exist_elim=> R. apply wand_intro_l.
rewrite assoc -persistently_and_sep_r.
by rewrite persistently_elim impl_elim_r.
Qed.
End persistently_affine_bi.
(* Persistence instances *)
Global Instance pure_persistent φ : Persistent (PROP:=PROP) φ⌝.
Proof. by rewrite /Persistent persistently_pure. Qed.
Global Instance emp_persistent : Persistent (PROP:=PROP) emp.
Proof. rewrite /Persistent. apply persistently_emp_intro. Qed.
Global Instance and_persistent P Q :
Persistent P Persistent Q Persistent (P Q).
Proof. intros. by rewrite /Persistent persistently_and -!persistent. Qed.
Global Instance or_persistent P Q :
Persistent P Persistent Q Persistent (P Q).
Proof. intros. by rewrite /Persistent persistently_or -!persistent. Qed.
Global Instance forall_persistent `{!BiPersistentlyForall PROP} {A} (Ψ : A PROP) :
( x, Persistent (Ψ x)) Persistent ( x, Ψ x).
Proof.
intros. rewrite /Persistent persistently_forall.
apply forall_mono=> x. by rewrite -!persistent.
Qed.
Global Instance exist_persistent {A} (Ψ : A PROP) :
( x, Persistent (Ψ x)) Persistent ( x, Ψ x).
Proof.
intros. rewrite /Persistent persistently_exist.
apply exist_mono=> x. by rewrite -!persistent.
Qed.
Global Instance sep_persistent P Q :
Persistent P Persistent Q Persistent (P Q).
Proof. intros. by rewrite /Persistent -persistently_sep_2 -!persistent. Qed.
Global Instance affinely_persistent P : Persistent P Persistent (<affine> P).
Proof. rewrite /bi_affinely. apply _. Qed.
Global Instance absorbingly_persistent P : Persistent P Persistent (<absorb> P).
Proof. rewrite /bi_absorbingly. apply _. Qed.
Global Instance from_option_persistent {A} P (Ψ : A PROP) (mx : option A) :
( x, Persistent (Ψ x)) Persistent P Persistent (from_option Ψ P mx).
Proof. destruct mx; apply _. Qed.
(* The intuitionistic modality *)
Global Instance intuitionistically_ne : NonExpansive (@bi_intuitionistically PROP).
Proof. solve_proper. Qed.
Global Instance intuitionistically_proper :
Proper ((⊣⊢) ==> (⊣⊢)) (@bi_intuitionistically PROP).
Proof. solve_proper. Qed.
Global Instance intuitionistically_mono' :
Proper (() ==> ()) (@bi_intuitionistically PROP).
Proof. solve_proper. Qed.
Global Instance intuitionistically_flip_mono' :
Proper (flip () ==> flip ()) (@bi_intuitionistically PROP).
Proof. solve_proper. Qed.
Global Instance intuitionistically_affine P : Affine ( P).
Proof. rewrite /bi_intuitionistically. apply _. Qed.
Global Instance intuitionistically_persistent P : Persistent ( P).
Proof. rewrite /bi_intuitionistically. apply _. Qed.
Lemma intuitionistically_def P : ( P)%I = (<affine> <pers> P)%I.
Proof. done. Qed.
Lemma intuitionistically_elim P : P P.
Proof. apply persistently_and_emp_elim. Qed.
Lemma intuitionistically_elim_emp P : P emp.
Proof. rewrite /bi_intuitionistically affinely_elim_emp //. Qed.
Lemma intuitionistically_intro' P Q : ( P Q) P Q.
Proof.
intros <-. rewrite /bi_intuitionistically.
by rewrite persistently_affinely_elim persistently_idemp.
Qed.
Lemma intuitionistically_emp : emp ⊣⊢ emp.
Proof.
by rewrite /bi_intuitionistically -persistently_True_emp persistently_pure
affinely_True_emp.
Qed.
Lemma intuitionistically_False : False ⊣⊢ False.
Proof. by rewrite /bi_intuitionistically persistently_pure affinely_False. Qed.
Lemma intuitionistically_True_emp : True ⊣⊢ emp.
Proof.
rewrite -intuitionistically_emp /bi_intuitionistically
persistently_True_emp //.
Qed.
Lemma intuitionistically_and P Q : (P Q) ⊣⊢ P Q.
Proof. by rewrite /bi_intuitionistically persistently_and affinely_and. Qed.
Lemma intuitionistically_forall {A} (Φ : A PROP) : ( x, Φ x) x, Φ x.
Proof.
by rewrite /bi_intuitionistically persistently_forall_1 affinely_forall.
Qed.
Lemma intuitionistically_or P Q : (P Q) ⊣⊢ P Q.
Proof. by rewrite /bi_intuitionistically persistently_or affinely_or. Qed.
Lemma intuitionistically_exist {A} (Φ : A PROP) : ( x, Φ x) ⊣⊢ x, Φ x.
Proof. by rewrite /bi_intuitionistically persistently_exist affinely_exist. Qed.
Lemma intuitionistically_sep_2 P Q : P Q (P Q).
Proof. by rewrite /bi_intuitionistically affinely_sep_2 persistently_sep_2. Qed.
Lemma intuitionistically_sep `{!BiPositive PROP} P Q : (P Q) ⊣⊢ P Q.
Proof. by rewrite /bi_intuitionistically -affinely_sep -persistently_sep. Qed.
Lemma intuitionistically_idemp P : P ⊣⊢ P.
Proof.
rewrite /bi_intuitionistically.
by rewrite persistently_affinely_elim persistently_idemp.
Qed.
Lemma intuitionistically_into_persistently_1 P : P <pers> P.
Proof. rewrite /bi_intuitionistically affinely_elim //. Qed.
Lemma intuitionistically_persistently_elim P : <pers> P ⊣⊢ P.
Proof. rewrite /bi_intuitionistically persistently_idemp //. Qed.
Lemma intuitionistic_intuitionistically P :
Affine P Persistent P P ⊣⊢ P.
Proof.
intros. apply (anti_symm _); first exact: intuitionistically_elim.
rewrite -{1}(affine_affinely P) {1}(persistent P) //.
Qed.
Lemma intuitionistically_affinely P : P <affine> P.
Proof.
rewrite /bi_intuitionistically /bi_affinely. apply and_intro.
- rewrite and_elim_l //.
- apply persistently_and_emp_elim.
Qed.
Lemma intuitionistically_affinely_elim P : <affine> P ⊣⊢ P.
Proof. rewrite /bi_intuitionistically persistently_affinely_elim //. Qed.
Lemma persistently_and_intuitionistically_sep_l P Q : <pers> P Q ⊣⊢ P Q.
Proof.
rewrite /bi_intuitionistically. apply (anti_symm _).
- by rewrite /bi_affinely -(comm bi_and (<pers> P)%I)
-persistently_and_sep_assoc left_id.
- apply and_intro.
+ by rewrite affinely_elim sep_elim_l.
+ by rewrite affinely_elim_emp left_id.
Qed.
Lemma persistently_and_intuitionistically_sep_r P Q : P <pers> Q ⊣⊢ P Q.
Proof. by rewrite !(comm _ P) persistently_and_intuitionistically_sep_l. Qed.
Lemma and_sep_intuitionistically P Q : P Q ⊣⊢ P Q.
Proof.
rewrite -persistently_and_intuitionistically_sep_l.
by rewrite -affinely_and affinely_and_r.
Qed.
Lemma intuitionistically_sep_dup P : P ⊣⊢ P P.
Proof.
by rewrite -persistently_and_intuitionistically_sep_l affinely_and_r idemp.
Qed.
Lemma impl_wand_intuitionistically P Q : (<pers> P Q) ⊣⊢ ( P -∗ Q).
Proof.
apply (anti_symm ()).
- apply wand_intro_l.
by rewrite -persistently_and_intuitionistically_sep_l impl_elim_r.
- apply impl_intro_l.
by rewrite persistently_and_intuitionistically_sep_l wand_elim_r.
Qed.
Lemma intuitionistically_alt_fixpoint P :
P ⊣⊢ emp (P P).
Proof.
apply (anti_symm ()).
- apply and_intro; first exact: affinely_elim_emp.
rewrite {1}intuitionistically_sep_dup. apply sep_mono; last done.
apply intuitionistically_elim.
- apply and_mono; first done.
rewrite /bi_intuitionistically {2}persistently_alt_fixpoint.
apply sep_mono; first done. apply and_elim_r.
Qed.
Lemma intuitionistically_impl_wand_2 P Q : (P -∗ Q) (P Q).
Proof. by rewrite /bi_intuitionistically persistently_impl_wand_2. Qed.
Lemma impl_alt P Q : (P Q) ⊣⊢ R, R <pers> (P R -∗ Q).
Proof.
apply (anti_symm ()).
- rewrite -(right_id True%I bi_and (P Q)%I) -(exist_intro (P Q)%I).
apply and_mono_r. rewrite impl_elim_r -entails_wand //.
apply persistently_emp_intro.
- apply exist_elim=> R. apply impl_intro_l.
rewrite assoc persistently_and_intuitionistically_sep_r.
by rewrite intuitionistically_elim wand_elim_r.
Qed.
Section bi_affine_intuitionistically.
Context `{!BiAffine PROP}.
Lemma intuitionistically_into_persistently P : P ⊣⊢ <pers> P.
Proof. rewrite /bi_intuitionistically affine_affinely //. Qed.
End bi_affine_intuitionistically.
(* Conditional affinely modality *)
Global Instance affinely_if_ne p : NonExpansive (@bi_affinely_if PROP p).
Proof. solve_proper. Qed.
Global Instance affinely_if_proper p :
Proper ((⊣⊢) ==> (⊣⊢)) (@bi_affinely_if PROP p).
Proof. solve_proper. Qed.
Global Instance affinely_if_mono' p :
Proper (() ==> ()) (@bi_affinely_if PROP p).
Proof. solve_proper. Qed.
Global Instance affinely_if_flip_mono' p :
Proper (flip () ==> flip ()) (@bi_affinely_if PROP p).
Proof. solve_proper. Qed.
Global Instance affinely_if_affine p P : Affine P Affine (<affine>?p P).
Proof. destruct p; simpl; apply _. Qed.
Global Instance affinely_if_persistent p P :
Persistent P Persistent (<affine>?p P).
Proof. destruct p; simpl; apply _. Qed.
Lemma affinely_if_mono p P Q : (P Q) <affine>?p P <affine>?p Q.
Proof. by intros ->. Qed.
Lemma affinely_if_flag_mono (p q : bool) P :
(q p) <affine>?p P <affine>?q P.
Proof. destruct p, q; naive_solver auto using affinely_elim. Qed.
Lemma affinely_if_elim p P : <affine>?p P P.
Proof. destruct p; simpl; auto using affinely_elim. Qed.
Lemma affinely_affinely_if p P : <affine> P <affine>?p P.
Proof. destruct p; simpl; auto using affinely_elim. Qed.
Lemma affinely_if_intro' p P Q :
(<affine>?p P Q) <affine>?p P <affine>?p Q.
Proof. destruct p; simpl; auto using affinely_intro'. Qed.
Lemma affinely_if_emp p : <affine>?p emp ⊣⊢ emp.
Proof. destruct p; simpl; auto using affinely_emp. Qed.
Lemma affinely_if_and p P Q : <affine>?p (P Q) ⊣⊢ <affine>?p P <affine>?p Q.
Proof. destruct p; simpl; auto using affinely_and. Qed.
Lemma affinely_if_or p P Q : <affine>?p (P Q) ⊣⊢ <affine>?p P <affine>?p Q.
Proof. destruct p; simpl; auto using affinely_or. Qed.
Lemma affinely_if_exist {A} p (Ψ : A PROP) :
<affine>?p ( a, Ψ a) ⊣⊢ a, <affine>?p (Ψ a).
Proof. destruct p; simpl; auto using affinely_exist. Qed.
Lemma affinely_if_sep_2 p P Q : <affine>?p P <affine>?p Q <affine>?p (P Q).
Proof. destruct p; simpl; auto using affinely_sep_2. Qed.
Lemma affinely_if_sep `{!BiPositive PROP} p P Q :
<affine>?p (P Q) ⊣⊢ <affine>?p P <affine>?p Q.
Proof. destruct p; simpl; auto using affinely_sep. Qed.
Lemma affinely_if_idemp p P : <affine>?p <affine>?p P ⊣⊢ <affine>?p P.
Proof. destruct p; simpl; auto using affinely_idemp. Qed.
Lemma affinely_if_and_l p P Q : <affine>?p P Q ⊣⊢ <affine>?p (P Q).
Proof. destruct p; simpl; auto using affinely_and_l. Qed.
Lemma affinely_if_and_r p P Q : P <affine>?p Q ⊣⊢ <affine>?p (P Q).
Proof. destruct p; simpl; auto using affinely_and_r. Qed.
Lemma affinely_if_and_lr p P Q : <affine>?p P Q ⊣⊢ P <affine>?p Q.
Proof. destruct p; simpl; auto using affinely_and_lr. Qed.
(* Conditional absorbingly modality *)
Global Instance absorbingly_if_ne p : NonExpansive (@bi_absorbingly_if PROP p).
Proof. solve_proper. Qed.
Global Instance absorbingly_if_proper p :
Proper ((⊣⊢) ==> (⊣⊢)) (@bi_absorbingly_if PROP p).
Proof. solve_proper. Qed.
Global Instance absorbingly_if_mono' p :
Proper (() ==> ()) (@bi_absorbingly_if PROP p).
Proof. solve_proper. Qed.
Global Instance absorbingly_if_flip_mono' p :
Proper (flip () ==> flip ()) (@bi_absorbingly_if PROP p).
Proof. solve_proper. Qed.
Global Instance absorbingly_if_persistent p P :
Persistent P Persistent (<absorb>?p P).
Proof. destruct p; simpl; apply _. Qed.
Lemma absorbingly_if_absorbingly p P : <absorb>?p P <absorb> P.
Proof. destruct p; simpl; auto using absorbingly_intro. Qed.
Lemma absorbingly_if_intro p P : P <absorb>?p P.
Proof. destruct p; simpl; auto using absorbingly_intro. Qed.
Lemma absorbingly_if_mono p P Q : (P Q) <absorb>?p P <absorb>?p Q.
Proof. by intros ->. Qed.
Lemma absorbingly_if_flag_mono (p q : bool) P :
(p q) <absorb>?p P <absorb>?q P.
Proof. destruct p, q; try naive_solver auto using absorbingly_intro. Qed.
Lemma absorbingly_if_idemp p P : <absorb>?p <absorb>?p P ⊣⊢ <absorb>?p P.
Proof. destruct p; simpl; auto using absorbingly_idemp. Qed.
Lemma absorbingly_if_pure p φ : <absorb>?p φ ⊣⊢ φ ⌝.
Proof. destruct p; simpl; auto using absorbingly_pure. Qed.
Lemma absorbingly_if_or p P Q :
<absorb>?p (P Q) ⊣⊢ <absorb>?p P <absorb>?p Q.
Proof. destruct p; simpl; auto using absorbingly_or. Qed.
Lemma absorbingly_if_and_1 p P Q :
<absorb>?p (P Q) <absorb>?p P <absorb>?p Q.
Proof. destruct p; simpl; auto using absorbingly_and_1. Qed.
Lemma absorbingly_if_forall {A} p (Φ : A PROP) :
<absorb>?p ( a, Φ a) a, <absorb>?p (Φ a).
Proof. destruct p; simpl; auto using absorbingly_forall. Qed.
Lemma absorbingly_if_exist {A} p (Φ : A PROP) :
<absorb>?p ( a, Φ a) ⊣⊢ a, <absorb>?p (Φ a).
Proof. destruct p; simpl; auto using absorbingly_exist. Qed.
Lemma absorbingly_if_sep p P Q :
<absorb>?p (P Q) ⊣⊢ <absorb>?p P <absorb>?p Q.
Proof. destruct p; simpl; auto using absorbingly_sep. Qed.
Lemma absorbingly_if_wand p P Q :
<absorb>?p (P -∗ Q) <absorb>?p P -∗ <absorb>?p Q.
Proof. destruct p; simpl; auto using absorbingly_wand. Qed.
Lemma absorbingly_if_sep_l p P Q : <absorb>?p P Q ⊣⊢ <absorb>?p (P Q).
Proof. destruct p; simpl; auto using absorbingly_sep_l. Qed.
Lemma absorbingly_if_sep_r p P Q : P <absorb>?p Q ⊣⊢ <absorb>?p (P Q).
Proof. destruct p; simpl; auto using absorbingly_sep_r. Qed.
Lemma absorbingly_if_sep_lr p P Q : <absorb>?p P Q ⊣⊢ P <absorb>?p Q.
Proof. destruct p; simpl; auto using absorbingly_sep_lr. Qed.
Lemma affinely_if_absorbingly_if_elim `{!BiPositive PROP} p P :
<affine>?p <absorb>?p P ⊣⊢ <affine>?p P.
Proof. destruct p; simpl; auto using affinely_absorbingly_elim. Qed.
(* Conditional persistently *)
Global Instance persistently_if_ne p : NonExpansive (@bi_persistently_if PROP p).
Proof. solve_proper. Qed.
Global Instance persistently_if_proper p :
Proper ((⊣⊢) ==> (⊣⊢)) (@bi_persistently_if PROP p).
Proof. solve_proper. Qed.
Global Instance persistently_if_mono' p :
Proper (() ==> ()) (@bi_persistently_if PROP p).
Proof. solve_proper. Qed.
Global Instance persistently_if_flip_mono' p :
Proper (flip () ==> flip ()) (@bi_persistently_if PROP p).
Proof. solve_proper. Qed.
Global Instance persistently_if_absorbing P p :
Absorbing P Absorbing (<pers>?p P).
Proof. intros; destruct p; simpl; apply _. Qed.
Lemma persistently_if_mono p P Q : (P Q) <pers>?p P <pers>?p Q.
Proof. by intros ->. Qed.
Lemma persistently_if_pure p φ : <pers>?p φ ⊣⊢ φ⌝.
Proof. destruct p; simpl; auto using persistently_pure. Qed.
Lemma persistently_if_and p P Q : <pers>?p (P Q) ⊣⊢ <pers>?p P <pers>?p Q.
Proof. destruct p; simpl; auto using persistently_and. Qed.
Lemma persistently_if_or p P Q : <pers>?p (P Q) ⊣⊢ <pers>?p P <pers>?p Q.
Proof. destruct p; simpl; auto using persistently_or. Qed.
Lemma persistently_if_exist {A} p (Ψ : A PROP) :
(<pers>?p ( a, Ψ a)) ⊣⊢ a, <pers>?p (Ψ a).
Proof. destruct p; simpl; auto using persistently_exist. Qed.
Lemma persistently_if_sep_2 p P Q : <pers>?p P <pers>?p Q <pers>?p (P Q).
Proof. destruct p; simpl; auto using persistently_sep_2. Qed.
Lemma persistently_if_sep `{!BiPositive PROP} p P Q :
<pers>?p (P Q) ⊣⊢ <pers>?p P <pers>?p Q.
Proof. destruct p; simpl; auto using persistently_sep. Qed.
Lemma persistently_if_idemp p P : <pers>?p <pers>?p P ⊣⊢ <pers>?p P.
Proof. destruct p; simpl; auto using persistently_idemp. Qed.
(* Conditional intuitionistically *)
Global Instance intuitionistically_if_ne p :
NonExpansive (@bi_intuitionistically_if PROP p).
Proof. solve_proper. Qed.
Global Instance intuitionistically_if_proper p :
Proper ((⊣⊢) ==> (⊣⊢)) (@bi_intuitionistically_if PROP p).
Proof. solve_proper. Qed.
Global Instance intuitionistically_if_mono' p :
Proper (() ==> ()) (@bi_intuitionistically_if PROP p).
Proof. solve_proper. Qed.
Global Instance intuitionistically_if_flip_mono' p :
Proper (flip () ==> flip ()) (@bi_intuitionistically_if PROP p).
Proof. solve_proper. Qed.
Global Instance intuitionistically_if_affine p P : Affine P Affine (?p P).
Proof. destruct p; simpl; apply _. Qed.
Lemma intuitionistically_if_mono p P Q : (P Q) ?p P ?p Q.
Proof. by intros ->. Qed.
Lemma intuitionistically_if_flag_mono (p q : bool) P :
(q p) ?p P ?q P.
Proof. destruct p, q; naive_solver auto using intuitionistically_elim. Qed.
Lemma intuitionistically_if_elim p P : ?p P P.
Proof. destruct p; simpl; auto using intuitionistically_elim. Qed.
Lemma intuitionistically_intuitionistically_if p P : P ?p P.
Proof. destruct p; simpl; auto using intuitionistically_elim. Qed.
Lemma intuitionistically_if_intro' p P Q : (?p P Q) ?p P ?p Q.
Proof. destruct p; simpl; auto using intuitionistically_intro'. Qed.
Lemma intuitionistically_if_emp p : ?p emp ⊣⊢ emp.
Proof. destruct p; simpl; auto using intuitionistically_emp. Qed.
Lemma intuitionistically_if_False p : ?p False ⊣⊢ False.
Proof. destruct p; simpl; auto using intuitionistically_False. Qed.
Lemma intuitionistically_if_and p P Q : ?p (P Q) ⊣⊢ ?p P ?p Q.
Proof. destruct p; simpl; auto using intuitionistically_and. Qed.
Lemma intuitionistically_if_or p P Q : ?p (P Q) ⊣⊢ ?p P ?p Q.
Proof. destruct p; simpl; auto using intuitionistically_or. Qed.
Lemma intuitionistically_if_exist {A} p (Ψ : A PROP) :
(?p a, Ψ a) ⊣⊢ a, ?p Ψ a.
Proof. destruct p; simpl; auto using intuitionistically_exist. Qed.
Lemma intuitionistically_if_sep_2 p P Q : ?p P ?p Q ?p (P Q).
Proof. destruct p; simpl; auto using intuitionistically_sep_2. Qed.
Lemma intuitionistically_if_sep `{!BiPositive PROP} p P Q :
?p (P Q) ⊣⊢ ?p P ?p Q.
Proof. destruct p; simpl; auto using intuitionistically_sep. Qed.
Lemma intuitionistically_if_idemp p P : ?p ?p P ⊣⊢ ?p P.
Proof. destruct p; simpl; auto using intuitionistically_idemp. Qed.
Lemma intuitionistically_if_unfold p P : ?p P ⊣⊢ <affine>?p <pers>?p P.
Proof. by destruct p. Qed.
(* Properties of persistent propositions *)
Global Instance Persistent_proper : Proper (() ==> iff) (@Persistent PROP).
Proof. solve_proper. Qed.
Lemma persistent_persistently_2 P `{!Persistent P} : P <pers> P.
Proof. done. Qed.
Lemma persistent_persistently P `{!Persistent P, !Absorbing P} : <pers> P ⊣⊢ P.
Proof.
apply (anti_symm _); auto using persistent_persistently_2, persistently_elim.
Qed.
Lemma persistently_intro P Q `{!Persistent P} : (P Q) P <pers> Q.
Proof. intros HP. by rewrite (persistent P) HP. Qed.
Lemma persistent_and_affinely_sep_l_1 P Q `{!Persistent P} :
P Q <affine> P Q.
Proof.
rewrite {1}(persistent_persistently_2 P).
rewrite persistently_and_intuitionistically_sep_l.
rewrite intuitionistically_affinely //.
Qed.
Lemma persistent_and_affinely_sep_r_1 P Q `{!Persistent Q} :
P Q P <affine> Q.
Proof. by rewrite !(comm _ P) persistent_and_affinely_sep_l_1. Qed.
Lemma persistent_and_affinely_sep_l P Q `{!Persistent P, !Absorbing P} :
P Q ⊣⊢ <affine> P Q.
Proof.
rewrite -(persistent_persistently P).
by rewrite persistently_and_intuitionistically_sep_l.
Qed.
Lemma persistent_and_affinely_sep_r P Q `{!Persistent Q, !Absorbing Q} :
P Q ⊣⊢ P <affine> Q.
Proof.
rewrite -(persistent_persistently Q).
by rewrite persistently_and_intuitionistically_sep_r.
Qed.
Lemma persistent_and_sep_1 P Q `{HPQ : !TCOr (Persistent P) (Persistent Q)} :
P Q P Q.
Proof.
destruct HPQ.
- by rewrite persistent_and_affinely_sep_l_1 affinely_elim.
- by rewrite persistent_and_affinely_sep_r_1 affinely_elim.
Qed.
Lemma persistent_sep_dup P
`{HP : !TCOr (Affine P) (Absorbing P), !Persistent P} :
P ⊣⊢ P P.
Proof.
destruct HP; last by rewrite -(persistent_persistently P) -persistently_sep_dup.
apply (anti_symm ()).
- by rewrite -{1}(intuitionistic_intuitionistically P)
intuitionistically_sep_dup intuitionistically_elim.
- by rewrite {1}(affine P) left_id.
Qed.
Lemma persistent_entails_l P Q `{!Persistent Q} : (P Q) P Q P.
Proof. intros. rewrite -persistent_and_sep_1; auto. Qed.
Lemma persistent_entails_r P Q `{!Persistent Q} : (P Q) P P Q.
Proof. intros. rewrite -persistent_and_sep_1; auto. Qed.
Lemma absorbingly_intuitionistically_into_persistently P :
<absorb> P ⊣⊢ <pers> P.
Proof.
apply (anti_symm _).
- rewrite intuitionistically_into_persistently_1.
by rewrite absorbingly_elim_persistently.
- rewrite -{1}(idemp bi_and (<pers> _)%I).
rewrite persistently_and_intuitionistically_sep_r.
by rewrite {1} (True_intro (<pers> _)%I).
Qed.
Lemma persistent_absorbingly_affinely_2 P `{!Persistent P} :
P <absorb> <affine> P.
Proof.
rewrite {1}(persistent P) -absorbingly_intuitionistically_into_persistently.
by rewrite intuitionistically_affinely.
Qed.
Lemma persistent_absorbingly_affinely P `{!Persistent P, !Absorbing P} :
<absorb> <affine> P ⊣⊢ P.
Proof.
rewrite -(persistent_persistently P).
by rewrite absorbingly_intuitionistically_into_persistently.
Qed.
Lemma persistent_and_sep_assoc P `{!Persistent P, !Absorbing P} Q R :
P (Q R) ⊣⊢ (P Q) R.
Proof. by rewrite -(persistent_persistently P) persistently_and_sep_assoc. Qed.
Lemma persistent_impl_wand_affinely P `{!Persistent P, !Absorbing P} Q :
(P Q) ⊣⊢ (<affine> P -∗ Q).
Proof.
apply (anti_symm _).
- apply wand_intro_l. rewrite -persistent_and_affinely_sep_l impl_elim_r //.
- apply impl_intro_l. rewrite persistent_and_affinely_sep_l wand_elim_r //.
Qed.
Lemma impl_wand_2 P `{!Persistent P} Q : (P -∗ Q) P Q.
Proof. apply impl_intro_l. by rewrite persistent_and_sep_1 wand_elim_r. Qed.
(** We don't have a [Intuitionistic] typeclass, but if we did, this
would be its only field. *)
Lemma intuitionistic P `{!Persistent P, !Affine P} : P P.
Proof. rewrite intuitionistic_intuitionistically. done. Qed.
Lemma intuitionistically_intro P Q `{!Affine P, !Persistent P} : (P Q) P Q.
Proof. intros. apply: affinely_intro. by apply: persistently_intro. Qed.
Section persistent_bi_absorbing.
Context `{!BiAffine PROP}.
Lemma persistent_and_sep P Q `{HPQ : !TCOr (Persistent P) (Persistent Q)} :
P Q ⊣⊢ P Q.
Proof.
destruct HPQ.
- by rewrite -(persistent_persistently P) persistently_and_sep_l.
- by rewrite -(persistent_persistently Q) persistently_and_sep_r.
Qed.
Lemma impl_wand P `{!Persistent P} Q : (P Q) ⊣⊢ (P -∗ Q).
Proof. apply (anti_symm _); auto using impl_wand_1, impl_wand_2. Qed.
End persistent_bi_absorbing.
Global Instance impl_absorbing P Q :
Persistent P Absorbing P Absorbing Q Absorbing (P Q).
Proof.
intros. rewrite /Absorbing. apply impl_intro_l.
rewrite persistent_and_affinely_sep_l_1 absorbingly_sep_r.
by rewrite -persistent_and_affinely_sep_l impl_elim_r.
Qed.
(* For big ops *)
Global Instance bi_and_monoid : Monoid (@bi_and PROP) :=
{| monoid_unit := True%I |}.
Global Instance bi_or_monoid : Monoid (@bi_or PROP) :=
{| monoid_unit := False%I |}.
Global Instance bi_sep_monoid : Monoid (@bi_sep PROP) :=
{| monoid_unit := emp%I |}.
Global Instance bi_persistently_and_homomorphism :
MonoidHomomorphism bi_and bi_and () (@bi_persistently PROP).
Proof.
split; [split|]; try apply _.
- apply persistently_and.
- apply persistently_pure.
Qed.
Global Instance bi_persistently_or_homomorphism :
MonoidHomomorphism bi_or bi_or () (@bi_persistently PROP).
Proof.
split; [split|]; try apply _.
- apply persistently_or.
- apply persistently_pure.
Qed.
Global Instance bi_persistently_sep_weak_homomorphism `{!BiPositive PROP} :
WeakMonoidHomomorphism bi_sep bi_sep () (@bi_persistently PROP).
Proof. split; [by apply _ ..|]. apply persistently_sep. Qed.
Global Instance bi_persistently_sep_homomorphism `{!BiAffine PROP} :
MonoidHomomorphism bi_sep bi_sep () (@bi_persistently PROP).
Proof. split; [by apply _ ..|]. apply persistently_emp. Qed.
Global Instance bi_persistently_sep_entails_weak_homomorphism :
WeakMonoidHomomorphism bi_sep bi_sep (flip ()) (@bi_persistently PROP).
Proof. split; [by apply _ ..|]. intros P Q; by rewrite persistently_sep_2. Qed.
Global Instance bi_persistently_sep_entails_homomorphism :
MonoidHomomorphism bi_sep bi_sep (flip ()) (@bi_persistently PROP).
Proof. split; [by apply _ ..|]. simpl. apply persistently_emp_intro. Qed.
(* Limits *)
Lemma limit_preserving_entails {A : ofe} `{!Cofe A} (Φ Ψ : A PROP) :
NonExpansive Φ NonExpansive Ψ LimitPreserving (λ x, Φ x Ψ x).
Proof.
intros. apply limit_preserving_ext with (λ x, True ⊣⊢ (Φ x Ψ x)).
- intros x. rewrite entails_eq_True. naive_solver.
- apply limit_preserving_equiv; solve_proper.
Qed.
Lemma limit_preserving_emp_valid {A : ofe} `{!Cofe A} (Φ : A PROP) :
NonExpansive Φ LimitPreserving (λ x, Φ x).
Proof. intros. apply limit_preserving_entails; solve_proper. Qed.
Global Instance limit_preserving_Persistent {A : ofe} `{!Cofe A} (Φ : A PROP) :
NonExpansive Φ LimitPreserving (λ x, Persistent (Φ x)).
Proof. intros. apply limit_preserving_entails; solve_proper. Qed.
(* iterated modalities *)
Lemma iter_modal_intro (M : PROP PROP) P (n : nat) :
( Q, Q M Q)
P Nat.iter n M P.
Proof.
intros Hintro; induction n as [|n IHn]; simpl; first done.
etransitivity; first apply IHn. apply Hintro.
Qed.
Lemma iter_modal_mono (M : PROP PROP) P Q (n : nat) :
( P Q, (P -∗ Q) M P -∗ M Q)
(P -∗ Q)
Nat.iter n M P -∗ Nat.iter n M Q.
Proof.
intros Hmono; induction n as [|n IHn]; simpl; first done.
rewrite -Hmono //.
Qed.
End derived.
End bi.
From iris.algebra Require Import monoid.
From iris.bi Require Export derived_laws.
From iris.prelude Require Import options.
Module bi.
Import interface.bi.
Import derived_laws.bi.
Section later_derived.
Context {PROP : bi}.
Implicit Types φ : Prop.
Implicit Types P Q R : PROP.
Implicit Types Ps : list PROP.
Implicit Types A : Type.
(* Force implicit argument PROP *)
Notation "P ⊢ Q" := (P ⊢@{PROP} Q).
Notation "P ⊣⊢ Q" := (P ⊣⊢@{PROP} Q).
Local Hint Resolve or_elim or_intro_l' or_intro_r' True_intro False_elim : core.
Local Hint Resolve and_elim_l' and_elim_r' and_intro forall_intro : core.
Global Instance later_proper :
Proper ((⊣⊢) ==> (⊣⊢)) (@bi_later PROP) := ne_proper _.
(* Later derived *)
Local Hint Resolve later_mono : core.
Global Instance later_mono' : Proper (() ==> ()) (@bi_later PROP).
Proof. intros P Q; apply later_mono. Qed.
Global Instance later_flip_mono' :
Proper (flip () ==> flip ()) (@bi_later PROP).
Proof. intros P Q; apply later_mono. Qed.
Lemma later_True : True ⊣⊢ True.
Proof. apply (anti_symm ()); auto using later_intro. Qed.
Lemma later_emp `{!BiAffine PROP} : emp ⊣⊢ emp.
Proof. by rewrite -True_emp later_True. Qed.
Lemma later_forall {A} (Φ : A PROP) : ( a, Φ a) ⊣⊢ ( a, Φ a).
Proof.
apply (anti_symm _); auto using later_forall_2.
apply forall_intro=> x. by rewrite (forall_elim x).
Qed.
Lemma later_exist_2 {A} (Φ : A PROP) : ( a, Φ a) ( a, Φ a).
Proof. apply exist_elim; eauto using exist_intro. Qed.
Lemma later_exist_except_0 {A} (Φ : A PROP) : ( a, Φ a) ( a, Φ a).
Proof. apply later_exist_false. Qed.
Lemma later_exist `{Inhabited A} (Φ : A PROP) : ( a, Φ a) ⊣⊢ ( a, Φ a).
Proof.
apply: anti_symm; [|apply later_exist_2].
rewrite later_exist_false. apply or_elim; last done.
rewrite -(exist_intro inhabitant); auto.
Qed.
Lemma later_and P Q : (P Q) ⊣⊢ P Q.
Proof. rewrite !and_alt later_forall. by apply forall_proper=> -[]. Qed.
Lemma later_or P Q : (P Q) ⊣⊢ P Q.
Proof. rewrite !or_alt later_exist. by apply exist_proper=> -[]. Qed.
Lemma later_impl P Q : (P Q) P Q.
Proof. apply impl_intro_l. by rewrite -later_and impl_elim_r. Qed.
Lemma later_sep P Q : (P Q) ⊣⊢ P Q.
Proof. apply (anti_symm _); auto using later_sep_1, later_sep_2. Qed.
Lemma later_wand P Q : (P -∗ Q) P -∗ Q.
Proof. apply wand_intro_l. by rewrite -later_sep wand_elim_r. Qed.
Lemma later_iff P Q : (P Q) P Q.
Proof. by rewrite /bi_iff later_and !later_impl. Qed.
Lemma later_wand_iff P Q : (P ∗-∗ Q) P ∗-∗ Q.
Proof. by rewrite /bi_wand_iff later_and !later_wand. Qed.
Lemma later_persistently P : <pers> P ⊣⊢ <pers> P.
Proof. apply (anti_symm _); auto using later_persistently_1, later_persistently_2. Qed.
Lemma later_affinely_2 P : <affine> P <affine> P.
Proof. rewrite /bi_affinely later_and. auto using later_intro. Qed.
Lemma later_intuitionistically_2 P : P P.
Proof. by rewrite /bi_intuitionistically -later_persistently later_affinely_2. Qed.
Lemma later_intuitionistically_if_2 p P : ?p P ?p P.
Proof. destruct p; simpl; auto using later_intuitionistically_2. Qed.
Lemma later_absorbingly P : <absorb> P ⊣⊢ <absorb> P.
Proof. by rewrite /bi_absorbingly later_sep later_True. Qed.
Lemma later_affinely `{!BiAffine PROP} P : <affine> P ⊣⊢ <affine> P.
Proof. by rewrite !affine_affinely. Qed.
Lemma later_intuitionistically `{!BiAffine PROP} P : P ⊣⊢ P.
Proof. by rewrite !intuitionistically_into_persistently later_persistently. Qed.
Lemma later_intuitionistically_if `{!BiAffine PROP} p P : ?p P ⊣⊢ ?p P.
Proof. destruct p; simpl; auto using later_intuitionistically. Qed.
Global Instance later_persistent P : Persistent P Persistent ( P).
Proof. intros. by rewrite /Persistent -later_persistently {1}(persistent P). Qed.
Global Instance later_absorbing P : Absorbing P Absorbing ( P).
Proof. intros ?. by rewrite /Absorbing -later_absorbingly absorbing. Qed.
(** * Alternatives to Löb induction *)
(** We prove relations between the following statements:
1. [Contractive (▷)], later is contractive as expressed by [BiLaterContractive].
2. [(▷ P ⊢ P) → (True ⊢ P)], the external/"weak" of Löb as expressed by [BiLöb].
3. [(▷ P → P) ⊢ P], the internal version/"strong" of Löb.
4. [□ (□ ▷ P -∗ P) ⊢ P], an internal version of Löb with magic wand instead of
implication.
5. [□ (▷ P -∗ P) ⊢ P], a weaker version of the former statement, which does not
make the induction hypothesis intuitionistic.
We prove that:
- (1) implies (2) in all BI logics (lemma [later_contractive_bi_löb]).
- (2) and (3) are logically equivalent in all BI logics (lemma [löb_alt_strong]).
- (2) implies (4) and (5) in all BI logics (lemmas [löb_wand_intuitionistically]
and [löb_wand]).
- (5) and (2) are logically equivalent in affine BI logics (lemma [löb_alt_wand]).
In particular, this gives that (2), (3), (4) and (5) are logically equivalent in
affine BI logics such as Iris. *)
Lemma löb `{!BiLöb PROP} P : ( P P) P.
Proof.
apply entails_impl_True, löb_weak. apply impl_intro_r.
rewrite -{2}(idemp () ( P P))%I.
rewrite {2}(later_intro ( P P)).
rewrite later_impl.
rewrite assoc impl_elim_l.
rewrite impl_elim_r. done.
Qed.
Lemma löb_alt_strong : BiLöb PROP P, ( P P) P.
Proof. split; intros HLöb P; [by apply löb|]. by intros ->%entails_impl_True. Qed.
(** Proof following https://en.wikipedia.org/wiki/L%C3%B6b's_theorem#Proof_of_L%C3%B6b's_theorem.
Their [Ψ] is called [Q] in our proof. *)
Global Instance later_contractive_bi_löb : BiLaterContractive PROP BiLöb PROP.
Proof.
intros=> P.
pose (flöb_pre (P Q : PROP) := ( Q P)%I).
assert ( P, Contractive (flöb_pre P)) by solve_contractive.
set (Q := fixpoint (flöb_pre P)).
assert (Q ⊣⊢ ( Q P)) as HQ by (exact: fixpoint_unfold).
intros HP. rewrite -HP.
assert ( Q P) as HQP.
{ rewrite -HP. rewrite -(idemp () ( Q))%I {2}(later_intro ( Q)).
by rewrite {1}HQ {1}later_impl impl_elim_l. }
rewrite -HQP HQ -2!later_intro.
apply (entails_impl_True _ P). done.
Qed.
Lemma löb_wand_intuitionistically `{!BiLöb PROP} P : ( P -∗ P) P.
Proof.
rewrite -{3}(intuitionistically_elim P) -(löb ( P)). apply impl_intro_l.
rewrite {1}intuitionistically_into_persistently_1 later_persistently.
rewrite persistently_and_intuitionistically_sep_l.
rewrite -{1}(intuitionistically_idemp ( P)) intuitionistically_sep_2.
by rewrite wand_elim_r.
Qed.
Lemma löb_wand `{!BiLöb PROP} P : ( P -∗ P) P.
Proof.
by rewrite -(intuitionistically_elim ( P)) löb_wand_intuitionistically.
Qed.
(** The proof of the right-to-left direction relies on the BI being affine. It
is unclear how to generalize the lemma or proof to support non-affine BIs. *)
Lemma löb_alt_wand `{!BiAffine PROP} : BiLöb PROP P, ( P -∗ P) P.
Proof.
split; intros Hlöb; [by apply löb_wand|].
apply löb_alt_strong=> P.
rewrite bi.impl_alt. apply bi.exist_elim=> R. apply impl_elim_r'.
rewrite -(Hlöb (R P)%I) -intuitionistically_into_persistently.
apply intuitionistically_intro', wand_intro_l, impl_intro_l.
rewrite -persistently_and_intuitionistically_sep_r assoc
persistently_and_intuitionistically_sep_r intuitionistically_elim.
rewrite -{1}(idemp bi_and R) -(assoc _ R) {2}(later_intro R).
by rewrite -later_and impl_elim_r (comm _ R) wand_elim_r.
Qed.
(** A funny consequence of Löb induction.
This shows that Löb induction is incompatible with classical logic.
See [lib/counterexamples.v] for a fully worked-out proof of that fact. *)
Lemma not_not_later_False `{!BiLöb PROP} : ⊢@{PROP} ¬ ¬ False.
Proof. apply entails_impl, löb. Qed.
(* Iterated later modality *)
Global Instance laterN_ne m : NonExpansive (@bi_laterN PROP m).
Proof. induction m; simpl; [by intros ???|]. solve_proper. Qed.
Global Instance laterN_proper m :
Proper ((⊣⊢) ==> (⊣⊢)) (@bi_laterN PROP m) := ne_proper _.
Lemma laterN_0 P : ▷^0 P ⊣⊢ P.
Proof. done. Qed.
Lemma later_laterN n P : ▷^(S n) P ⊣⊢ ▷^n P.
Proof. done. Qed.
Lemma laterN_later n P : ▷^(S n) P ⊣⊢ ▷^n P.
Proof. induction n; f_equiv/=; auto. Qed.
Lemma laterN_add n1 n2 P : ▷^(n1 + n2) P ⊣⊢ ▷^n1 ▷^n2 P.
Proof. induction n1; f_equiv/=; auto. Qed.
Lemma laterN_le n1 n2 P : n1 n2 ▷^n1 P ▷^n2 P.
Proof. induction 1; simpl; by rewrite -?later_intro. Qed.
Lemma laterN_iter n P : (▷^n P)%I = Nat.iter n bi_later P.
Proof. induction n; f_equal/=; auto. Qed.
Lemma laterN_mono n P Q : (P Q) ▷^n P ▷^n Q.
Proof. induction n; simpl; auto. Qed.
Global Instance laterN_mono' n : Proper (() ==> ()) (@bi_laterN PROP n).
Proof. intros P Q; apply laterN_mono. Qed.
Global Instance laterN_flip_mono' n :
Proper (flip () ==> flip ()) (@bi_laterN PROP n).
Proof. intros P Q; apply laterN_mono. Qed.
Lemma laterN_intro n P : P ▷^n P.
Proof. induction n as [|n IH]; simpl; by rewrite -?later_intro. Qed.
Lemma laterN_True n : ▷^n True ⊣⊢ True.
Proof. apply (anti_symm ()); auto using laterN_intro, True_intro. Qed.
Lemma laterN_emp `{!BiAffine PROP} n : ▷^n emp ⊣⊢ emp.
Proof. by rewrite -True_emp laterN_True. Qed.
Lemma laterN_forall {A} n (Φ : A PROP) : (▷^n a, Φ a) ⊣⊢ ( a, ▷^n Φ a).
Proof. induction n as [|n IH]; simpl; rewrite -?later_forall ?IH; auto. Qed.
Lemma laterN_exist_2 {A} n (Φ : A PROP) : ( a, ▷^n Φ a) ▷^n ( a, Φ a).
Proof. apply exist_elim; eauto using exist_intro, laterN_mono. Qed.
Lemma laterN_exist {A} `{!Inhabited A} n (Φ : A PROP) :
(▷^n a, Φ a) ⊣⊢ a, ▷^n Φ a.
Proof. induction n as [|n IH]; simpl; rewrite -?later_exist ?IH; auto. Qed.
Lemma laterN_and n P Q : ▷^n (P Q) ⊣⊢ ▷^n P ▷^n Q.
Proof. induction n as [|n IH]; simpl; rewrite -?later_and ?IH; auto. Qed.
Lemma laterN_or n P Q : ▷^n (P Q) ⊣⊢ ▷^n P ▷^n Q.
Proof. induction n as [|n IH]; simpl; rewrite -?later_or ?IH; auto. Qed.
Lemma laterN_impl n P Q : ▷^n (P Q) ▷^n P ▷^n Q.
Proof. apply impl_intro_l. by rewrite -laterN_and impl_elim_r. Qed.
Lemma laterN_sep n P Q : ▷^n (P Q) ⊣⊢ ▷^n P ▷^n Q.
Proof. induction n as [|n IH]; simpl; rewrite -?later_sep ?IH; auto. Qed.
Lemma laterN_wand n P Q : ▷^n (P -∗ Q) ▷^n P -∗ ▷^n Q.
Proof. apply wand_intro_l. by rewrite -laterN_sep wand_elim_r. Qed.
Lemma laterN_iff n P Q : ▷^n (P Q) ▷^n P ▷^n Q.
Proof. by rewrite /bi_iff laterN_and !laterN_impl. Qed.
Lemma laterN_persistently n P : ▷^n <pers> P ⊣⊢ <pers> ▷^n P.
Proof. induction n as [|n IH]; simpl; auto. by rewrite IH later_persistently. Qed.
Lemma laterN_affinely_2 n P : <affine> ▷^n P ▷^n <affine> P.
Proof. rewrite /bi_affinely laterN_and. auto using laterN_intro. Qed.
Lemma laterN_intuitionistically_2 n P : ▷^n P ▷^n P.
Proof. by rewrite /bi_intuitionistically -laterN_persistently laterN_affinely_2. Qed.
Lemma laterN_intuitionistically_if_2 n p P : ?p ▷^n P ▷^n ?p P.
Proof. destruct p; simpl; auto using laterN_intuitionistically_2. Qed.
Lemma laterN_absorbingly n P : ▷^n <absorb> P ⊣⊢ <absorb> ▷^n P.
Proof. by rewrite /bi_absorbingly laterN_sep laterN_True. Qed.
Global Instance laterN_persistent n P : Persistent P Persistent (▷^n P).
Proof. induction n; apply _. Qed.
Global Instance laterN_absorbing n P : Absorbing P Absorbing (▷^n P).
Proof. induction n; apply _. Qed.
(* Except-0 *)
Global Instance except_0_ne : NonExpansive (@bi_except_0 PROP).
Proof. solve_proper. Qed.
Global Instance except_0_proper : Proper ((⊣⊢) ==> (⊣⊢)) (@bi_except_0 PROP).
Proof. solve_proper. Qed.
Global Instance except_0_mono' : Proper (() ==> ()) (@bi_except_0 PROP).
Proof. solve_proper. Qed.
Global Instance except_0_flip_mono' :
Proper (flip () ==> flip ()) (@bi_except_0 PROP).
Proof. solve_proper. Qed.
Lemma except_0_intro P : P P.
Proof. rewrite /bi_except_0; auto. Qed.
Lemma except_0_mono P Q : (P Q) P Q.
Proof. by intros ->. Qed.
Lemma except_0_idemp P : P ⊣⊢ P.
Proof. apply (anti_symm _); rewrite /bi_except_0; auto. Qed.
Lemma except_0_True : True ⊣⊢ True.
Proof. rewrite /bi_except_0. apply (anti_symm _); auto. Qed.
Lemma except_0_emp `{!BiAffine PROP} : emp ⊣⊢ emp.
Proof. by rewrite -True_emp except_0_True. Qed.
Lemma except_0_or P Q : (P Q) ⊣⊢ P Q.
Proof. rewrite /bi_except_0. apply (anti_symm _); auto. Qed.
Lemma except_0_and P Q : (P Q) ⊣⊢ P Q.
Proof. by rewrite /bi_except_0 or_and_l. Qed.
Lemma except_0_sep P Q : (P Q) ⊣⊢ P Q.
Proof.
rewrite /bi_except_0. apply (anti_symm _).
- apply or_elim; last by auto using sep_mono.
by rewrite -!or_intro_l -persistently_pure -later_sep -persistently_sep_dup.
- rewrite sep_or_r !sep_or_l {1}(later_intro P) {1}(later_intro Q).
rewrite -!later_sep !left_absorb right_absorb. auto.
Qed.
Lemma except_0_forall {A} (Φ : A PROP) : ( a, Φ a) ⊣⊢ a, Φ a.
Proof.
apply (anti_symm _).
{ apply forall_intro=> a. by rewrite (forall_elim a). }
trans ( ( a : A, Φ a) ( a : A, Φ a))%I.
{ apply and_intro, reflexivity. rewrite later_forall. apply forall_mono=> a.
apply or_elim; auto using later_intro. }
rewrite later_false_em and_or_r. apply or_elim.
{ rewrite and_elim_l. apply or_intro_l. }
apply or_intro_r', forall_intro=> a. rewrite !(forall_elim a).
by rewrite and_or_l impl_elim_l and_elim_r idemp.
Qed.
Lemma except_0_exist_2 {A} (Φ : A PROP) : ( a, Φ a) a, Φ a.
Proof. apply exist_elim=> a. by rewrite (exist_intro a). Qed.
Lemma except_0_exist `{Inhabited A} (Φ : A PROP) :
( a, Φ a) ⊣⊢ ( a, Φ a).
Proof.
apply (anti_symm _); [|by apply except_0_exist_2]. apply or_elim.
- rewrite -(exist_intro inhabitant). by apply or_intro_l.
- apply exist_mono=> a. apply except_0_intro.
Qed.
Lemma except_0_later P : P P.
Proof. by rewrite /bi_except_0 -later_or False_or. Qed.
Lemma except_0_laterN n P : ▷^n P ▷^n P.
Proof. by destruct n as [|n]; rewrite //= ?except_0_later -except_0_intro. Qed.
Lemma except_0_into_later P : P P.
Proof. by rewrite -except_0_later -later_intro. Qed.
Lemma except_0_persistently P : <pers> P ⊣⊢ <pers> P.
Proof.
by rewrite /bi_except_0 persistently_or -later_persistently persistently_pure.
Qed.
Lemma except_0_affinely_2 P : <affine> P <affine> P.
Proof. rewrite /bi_affinely except_0_and. auto using except_0_intro. Qed.
Lemma except_0_intuitionistically_2 P : P P.
Proof. by rewrite /bi_intuitionistically -except_0_persistently except_0_affinely_2. Qed.
Lemma except_0_intuitionistically_if_2 p P : ?p P ?p P.
Proof. destruct p; simpl; auto using except_0_intuitionistically_2. Qed.
Lemma except_0_absorbingly P : <absorb> P ⊣⊢ <absorb> P.
Proof. by rewrite /bi_absorbingly except_0_sep except_0_True. Qed.
Lemma except_0_frame_l P Q : P Q (P Q).
Proof. by rewrite {1}(except_0_intro P) except_0_sep. Qed.
Lemma except_0_frame_r P Q : P Q (P Q).
Proof. by rewrite {1}(except_0_intro Q) except_0_sep. Qed.
Lemma later_affinely_1 `{!Timeless (PROP:=PROP) emp} P : <affine> P <affine> P.
Proof.
rewrite /bi_affinely later_and (timeless emp) except_0_and.
by apply and_mono, except_0_intro.
Qed.
Global Instance except_0_persistent P : Persistent P Persistent ( P).
Proof. rewrite /bi_except_0; apply _. Qed.
Global Instance except_0_absorbing P : Absorbing P Absorbing ( P).
Proof. rewrite /bi_except_0; apply _. Qed.
(* Timeless instances *)
Global Instance Timeless_proper : Proper (() ==> iff) (@Timeless PROP).
Proof. solve_proper. Qed.
(* The left-to-right direction of this lemma shows that to prove a timeless
proposition [Q], we can additionally assume that we are at step-index 0, i.e.
we can add [▷ False] to our assumptions. The right-to-left direction shows
that this is in fact an exact characterization of timeless propositions.
See also the comment above the definition of [Timeless]. *)
Lemma timeless_alt `{!BiLöb PROP} Q :
Timeless Q ( P, ( False P Q) (P Q)).
Proof.
split; rewrite /Timeless => H.
* intros P Hpr.
rewrite -(löb Q). apply impl_intro_l.
rewrite H /bi_except_0 and_or_r. apply or_elim; auto.
* rewrite later_false_em.
apply or_mono; first done.
apply H, impl_elim_r.
Qed.
Global Instance pure_timeless φ : Timeless (PROP:=PROP) φ⌝.
Proof.
rewrite /Timeless /bi_except_0 pure_alt later_exist_false.
apply or_elim, exist_elim; [auto|]=> . rewrite -(exist_intro ). auto.
Qed.
Global Instance emp_timeless `{!BiAffine PROP} : Timeless (PROP:=PROP) emp.
Proof. rewrite -True_emp. apply _. Qed.
Global Instance and_timeless P Q : Timeless P Timeless Q Timeless (P Q).
Proof. intros; rewrite /Timeless except_0_and later_and; auto. Qed.
Global Instance or_timeless P Q : Timeless P Timeless Q Timeless (P Q).
Proof. intros; rewrite /Timeless except_0_or later_or; auto. Qed.
Global Instance impl_timeless `{!BiLöb PROP} P Q : Timeless Q Timeless (P Q).
Proof.
rewrite !timeless_alt=> HQ R HR. apply impl_intro_l, HQ.
rewrite assoc -(comm _ P) -assoc. by apply impl_elim_r'.
Qed.
Global Instance sep_timeless P Q: Timeless P Timeless Q Timeless (P Q).
Proof.
intros; rewrite /Timeless except_0_sep later_sep; auto using sep_mono.
Qed.
Global Instance wand_timeless `{!BiLöb PROP} P Q : Timeless Q Timeless (P -∗ Q).
Proof.
rewrite !timeless_alt=> HQ R HR. apply wand_intro_l, HQ.
rewrite persistent_and_affinely_sep_l assoc -(comm _ P) -assoc.
rewrite -persistent_and_affinely_sep_l. by apply wand_elim_r'.
Qed.
Global Instance forall_timeless {A} (Ψ : A PROP) :
( x, Timeless (Ψ x)) Timeless ( x, Ψ x).
Proof.
rewrite /Timeless=> HQ. rewrite except_0_forall later_forall.
apply forall_mono; auto.
Qed.
Global Instance exist_timeless {A} (Ψ : A PROP) :
( x, Timeless (Ψ x)) Timeless ( x, Ψ x).
Proof.
rewrite /Timeless=> ?. rewrite later_exist_false. apply or_elim.
- rewrite /bi_except_0; auto.
- apply exist_elim=> x. rewrite -(exist_intro x); auto.
Qed.
Global Instance persistently_timeless P : Timeless P Timeless (<pers> P).
Proof.
intros. rewrite /Timeless /bi_except_0 later_persistently_1.
by rewrite (timeless P) /bi_except_0 persistently_or {1}persistently_elim.
Qed.
Global Instance affinely_timeless P :
Timeless (PROP:=PROP) emp Timeless P Timeless (<affine> P).
Proof. rewrite /bi_affinely; apply _. Qed.
Global Instance absorbingly_timeless P : Timeless P Timeless (<absorb> P).
Proof. rewrite /bi_absorbingly; apply _. Qed.
Global Instance intuitionistically_timeless P :
Timeless (PROP:=PROP) emp Timeless P Timeless ( P).
Proof. rewrite /bi_intuitionistically; apply _. Qed.
Global Instance from_option_timeless {A} P (Ψ : A PROP) (mx : option A) :
( x, Timeless (Ψ x)) Timeless P Timeless (from_option Ψ P mx).
Proof. destruct mx; apply _. Qed.
(* Big op stuff *)
Global Instance bi_later_monoid_and_homomorphism :
MonoidHomomorphism bi_and bi_and () (@bi_later PROP).
Proof.
split; [split|]; try apply _.
- apply later_and.
- apply later_True.
Qed.
Global Instance bi_laterN_and_homomorphism n :
MonoidHomomorphism bi_and bi_and () (@bi_laterN PROP n).
Proof.
split; [split|]; try apply _.
- apply laterN_and.
- apply laterN_True.
Qed.
Global Instance bi_except_0_and_homomorphism :
MonoidHomomorphism bi_and bi_and () (@bi_except_0 PROP).
Proof.
split; [split|]; try apply _.
- apply except_0_and.
- apply except_0_True.
Qed.
Global Instance bi_later_monoid_or_homomorphism :
WeakMonoidHomomorphism bi_or bi_or () (@bi_later PROP).
Proof. split; try apply _. apply later_or. Qed.
Global Instance bi_laterN_or_homomorphism n :
WeakMonoidHomomorphism bi_or bi_or () (@bi_laterN PROP n).
Proof. split; try apply _. apply laterN_or. Qed.
Global Instance bi_except_0_or_homomorphism :
WeakMonoidHomomorphism bi_or bi_or () (@bi_except_0 PROP).
Proof. split; try apply _. apply except_0_or. Qed.
Global Instance bi_later_monoid_sep_weak_homomorphism :
WeakMonoidHomomorphism bi_sep bi_sep () (@bi_later PROP).
Proof. split; try apply _. apply later_sep. Qed.
Global Instance bi_laterN_sep_weak_homomorphism n :
WeakMonoidHomomorphism bi_sep bi_sep () (@bi_laterN PROP n).
Proof. split; try apply _. apply laterN_sep. Qed.
Global Instance bi_except_0_sep_weak_homomorphism :
WeakMonoidHomomorphism bi_sep bi_sep () (@bi_except_0 PROP).
Proof. split; try apply _. apply except_0_sep. Qed.
Global Instance bi_later_monoid_sep_homomorphism `{!BiAffine PROP} :
MonoidHomomorphism bi_sep bi_sep () (@bi_later PROP).
Proof. split; try apply _. apply later_emp. Qed.
Global Instance bi_laterN_sep_homomorphism `{!BiAffine PROP} n :
MonoidHomomorphism bi_sep bi_sep () (@bi_laterN PROP n).
Proof. split; try apply _. apply laterN_emp. Qed.
Global Instance bi_except_0_sep_homomorphism `{!BiAffine PROP} :
MonoidHomomorphism bi_sep bi_sep () (@bi_except_0 PROP).
Proof. split; try apply _. apply except_0_emp. Qed.
Global Instance bi_later_monoid_sep_entails_weak_homomorphism :
WeakMonoidHomomorphism bi_sep bi_sep (flip ()) (@bi_later PROP).
Proof. split; try apply _. intros P Q. by rewrite later_sep. Qed.
Global Instance bi_laterN_sep_entails_weak_homomorphism n :
WeakMonoidHomomorphism bi_sep bi_sep (flip ()) (@bi_laterN PROP n).
Proof. split; try apply _. intros P Q. by rewrite laterN_sep. Qed.
Global Instance bi_except_0_sep_entails_weak_homomorphism :
WeakMonoidHomomorphism bi_sep bi_sep (flip ()) (@bi_except_0 PROP).
Proof. split; try apply _. intros P Q. by rewrite except_0_sep. Qed.
Global Instance bi_later_monoid_sep_entails_homomorphism :
MonoidHomomorphism bi_sep bi_sep (flip ()) (@bi_later PROP).
Proof. split; try apply _. apply later_intro. Qed.
Global Instance bi_laterN_sep_entails_homomorphism n :
MonoidHomomorphism bi_sep bi_sep (flip ()) (@bi_laterN PROP n).
Proof. split; try apply _. apply laterN_intro. Qed.
Global Instance bi_except_0_sep_entails_homomorphism :
MonoidHomomorphism bi_sep bi_sep (flip ()) (@bi_except_0 PROP).
Proof. split; try apply _. apply except_0_intro. Qed.
End later_derived.
End bi.
From iris.algebra Require Import monoid.
From iris.bi Require Import interface derived_laws_later big_op.
From iris.bi Require Import plainly updates internal_eq.
From iris.prelude Require Import options.
(* 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*".
Class Embed (A B : Type) := embed : A B.
Global Arguments embed {_ _ _} _%_I : simpl never.
Notation "⎡ P ⎤" := (embed P) : bi_scope.
Global Instance: Params (@embed) 3 := {}.
Global Typeclasses Opaque embed.
Global Hint Mode Embed ! - : typeclass_instances.
Global Hint Mode Embed - ! : typeclass_instances.
(* Mixins allow us to create instances easily without having to use Program *)
Record BiEmbedMixin (PROP1 PROP2 : bi) `(Embed PROP1 PROP2) := {
bi_embed_mixin_ne : NonExpansive (embed (A:=PROP1) (B:=PROP2));
bi_embed_mixin_mono : Proper (() ==> ()) (embed (A:=PROP1) (B:=PROP2));
bi_embed_mixin_emp_valid_inj (P : PROP1) :
(⊢@{PROP2} P) P;
(** The following axiom expresses that the embedding is injective in the OFE
sense. Instead of this axiom being expressed in terms of [siProp] or
externally (i.e., as [Inj (dist n) (dist n) embed]), it is expressed using the
internal equality of _any other_ BI [PROP']. This is more general, as we do
not have any machinary to embed [siProp] into a BI with internal equality. *)
bi_embed_mixin_interal_inj {PROP' : bi} `{!BiInternalEq PROP'} (P Q : PROP1) :
P Q ⊢@{PROP'} (P Q);
bi_embed_mixin_emp_2 : emp ⊢@{PROP2} emp;
bi_embed_mixin_impl_2 (P Q : PROP1) :
(P Q) ⊢@{PROP2} P Q;
bi_embed_mixin_forall_2 A (Φ : A PROP1) :
( x, Φ x) ⊢@{PROP2} ⎡∀ x, Φ x;
bi_embed_mixin_exist_1 A (Φ : A PROP1) :
⎡∃ x, Φ x ⊢@{PROP2} x, Φ x;
bi_embed_mixin_sep (P Q : PROP1) :
P Q ⊣⊢@{PROP2} P Q;
bi_embed_mixin_wand_2 (P Q : PROP1) :
(P -∗ Q) ⊢@{PROP2} P -∗ Q;
bi_embed_mixin_persistently (P : PROP1) :
⎡<pers> P ⊣⊢@{PROP2} <pers> P
}.
Class BiEmbed (PROP1 PROP2 : bi) := {
#[global] bi_embed_embed :: Embed PROP1 PROP2;
bi_embed_mixin : BiEmbedMixin PROP1 PROP2 bi_embed_embed;
}.
Global Hint Mode BiEmbed ! - : typeclass_instances.
Global Hint Mode BiEmbed - ! : typeclass_instances.
Global Arguments bi_embed_embed : simpl never.
Class BiEmbedEmp (PROP1 PROP2 : bi) `{!BiEmbed PROP1 PROP2} :=
embed_emp_1 : emp : PROP1 emp.
Global Hint Mode BiEmbedEmp ! - - : typeclass_instances.
Global Hint Mode BiEmbedEmp - ! - : typeclass_instances.
Class BiEmbedLater (PROP1 PROP2 : bi) `{!BiEmbed PROP1 PROP2} :=
embed_later P : ⎡▷ P ⊣⊢@{PROP2} P⎤.
Global Hint Mode BiEmbedLater ! - - : typeclass_instances.
Global Hint Mode BiEmbedLater - ! - : typeclass_instances.
Class BiEmbedInternalEq (PROP1 PROP2 : bi)
`{!BiEmbed PROP1 PROP2, !BiInternalEq PROP1, !BiInternalEq PROP2} :=
embed_internal_eq_1 (A : ofe) (x y : A) : x y ⊢@{PROP2} x y.
Global Hint Mode BiEmbedInternalEq ! - - - - : typeclass_instances.
Global Hint Mode BiEmbedInternalEq - ! - - - : typeclass_instances.
Class BiEmbedBUpd (PROP1 PROP2 : bi)
`{!BiEmbed PROP1 PROP2, !BiBUpd PROP1, !BiBUpd PROP2} :=
embed_bupd P : ⎡|==> P ⊣⊢@{PROP2} |==> P⎤.
Global Hint Mode BiEmbedBUpd - ! - - - : typeclass_instances.
Global Hint Mode BiEmbedBUpd ! - - - - : typeclass_instances.
Class BiEmbedFUpd (PROP1 PROP2 : bi)
`{!BiEmbed PROP1 PROP2, !BiFUpd PROP1, !BiFUpd PROP2} :=
embed_fupd E1 E2 P : ⎡|={E1,E2}=> P ⊣⊢@{PROP2} |={E1,E2}=> P⎤.
Global Hint Mode BiEmbedFUpd - ! - - - : typeclass_instances.
Global Hint Mode BiEmbedFUpd ! - - - - : typeclass_instances.
Class BiEmbedPlainly (PROP1 PROP2 : bi)
`{!BiEmbed PROP1 PROP2, !BiPlainly PROP1, !BiPlainly PROP2} :=
embed_plainly (P : PROP1) : ⎡■ P ⊣⊢@{PROP2} P⎤.
Global Hint Mode BiEmbedPlainly - ! - - - : typeclass_instances.
Global Hint Mode BiEmbedPlainly ! - - - - : typeclass_instances.
Section embed_laws.
Context {PROP1 PROP2 : bi} `{!BiEmbed PROP1 PROP2}.
Local Notation embed := (embed (A:=bi_car PROP1) (B:=bi_car PROP2)).
Local Notation "⎡ P ⎤" := (embed P) : bi_scope.
Implicit Types P : PROP1.
Global Instance embed_ne : NonExpansive embed.
Proof. eapply bi_embed_mixin_ne, bi_embed_mixin. Qed.
Global Instance embed_mono : Proper (() ==> ()) embed.
Proof. eapply bi_embed_mixin_mono, bi_embed_mixin. Qed.
Lemma embed_emp_valid_inj P : (⊢@{PROP2} P) P.
Proof. eapply bi_embed_mixin_emp_valid_inj, bi_embed_mixin. Qed.
Lemma embed_interal_inj `{!BiInternalEq PROP'} (P Q : PROP1) :
P Q ⊢@{PROP'} (P Q).
Proof. eapply bi_embed_mixin_interal_inj, bi_embed_mixin. Qed.
Lemma embed_emp_2 : emp emp⎤.
Proof. eapply bi_embed_mixin_emp_2, bi_embed_mixin. Qed.
Lemma embed_impl_2 P Q : (P Q) P Q⎤.
Proof. eapply bi_embed_mixin_impl_2, bi_embed_mixin. Qed.
Lemma embed_forall_2 A (Φ : A PROP1) : ( x, Φ x) ⎡∀ x, Φ x⎤.
Proof. eapply bi_embed_mixin_forall_2, bi_embed_mixin. Qed.
Lemma embed_exist_1 A (Φ : A PROP1) : ⎡∃ x, Φ x x, Φ x⎤.
Proof. eapply bi_embed_mixin_exist_1, bi_embed_mixin. Qed.
Lemma embed_sep P Q : P Q ⊣⊢ P Q⎤.
Proof. eapply bi_embed_mixin_sep, bi_embed_mixin. Qed.
Lemma embed_wand_2 P Q : (P -∗ Q) P -∗ Q⎤.
Proof. eapply bi_embed_mixin_wand_2, bi_embed_mixin. Qed.
Lemma embed_persistently P : ⎡<pers> P ⊣⊢ <pers> P⎤.
Proof. eapply bi_embed_mixin_persistently, bi_embed_mixin. Qed.
End embed_laws.
Section embed.
Context {PROP1 PROP2 : bi} `{!BiEmbed PROP1 PROP2}.
Local Notation embed := (embed (A:=bi_car PROP1) (B:=bi_car PROP2)).
Local Notation "⎡ P ⎤" := (embed P) : bi_scope.
Implicit Types P Q R : PROP1.
Global Instance embed_proper : Proper (() ==> ()) embed.
Proof. apply (ne_proper _). Qed.
Global Instance embed_flip_mono : Proper (flip () ==> flip ()) embed.
Proof. solve_proper. Qed.
Global Instance embed_entails_inj : Inj () () embed.
Proof.
move=> P Q /bi.entails_wand. rewrite embed_wand_2.
by move=> /embed_emp_valid_inj /bi.wand_entails.
Qed.
Global Instance embed_inj : Inj () () embed.
Proof.
intros P Q EQ. apply bi.equiv_entails, conj; apply (inj embed); rewrite EQ //.
Qed.
Lemma embed_emp_valid (P : PROP1) : ( P) ( P).
Proof.
rewrite /bi_emp_valid. split=> HP.
- by apply embed_emp_valid_inj.
- by rewrite embed_emp_2 HP.
Qed.
Lemma embed_emp `{!BiEmbedEmp PROP1 PROP2} : emp ⊣⊢ emp.
Proof. apply (anti_symm _); eauto using embed_emp_1, embed_emp_2. Qed.
Lemma embed_forall A (Φ : A PROP1) : ⎡∀ x, Φ x ⊣⊢ x, Φ x⎤.
Proof.
apply bi.equiv_entails; split; [|apply embed_forall_2].
apply bi.forall_intro=>?. by rewrite bi.forall_elim.
Qed.
Lemma embed_exist A (Φ : A PROP1) : ⎡∃ x, Φ x ⊣⊢ x, Φ x⎤.
Proof.
apply bi.equiv_entails; split; [apply embed_exist_1|].
apply bi.exist_elim=>?. by rewrite -bi.exist_intro.
Qed.
Lemma embed_and P Q : P Q ⊣⊢ P Q⎤.
Proof. rewrite !bi.and_alt embed_forall. by f_equiv=>-[]. Qed.
Lemma embed_or P Q : P Q ⊣⊢ P Q⎤.
Proof. rewrite !bi.or_alt embed_exist. by f_equiv=>-[]. Qed.
Lemma embed_impl P Q : P Q ⊣⊢ (P Q).
Proof.
apply bi.equiv_entails; split; [|apply embed_impl_2].
apply bi.impl_intro_l. by rewrite -embed_and bi.impl_elim_r.
Qed.
Lemma embed_wand P Q : P -∗ Q ⊣⊢ (P -∗ Q).
Proof.
apply bi.equiv_entails; split; [|apply embed_wand_2].
apply bi.wand_intro_l. by rewrite -embed_sep bi.wand_elim_r.
Qed.
Lemma embed_pure φ : ⎡⌜φ⌝⎤ ⊣⊢ φ⌝.
Proof.
rewrite (@bi.pure_alt PROP1) (@bi.pure_alt PROP2) embed_exist.
do 2 f_equiv. apply bi.equiv_entails. split; [apply bi.True_intro|].
rewrite -(_ : (emp emp : PROP1) True) ?embed_impl;
last apply bi.True_intro.
apply bi.impl_intro_l. by rewrite right_id.
Qed.
Lemma embed_iff P Q : P Q ⊣⊢ (P Q).
Proof. by rewrite embed_and !embed_impl. Qed.
Lemma embed_wand_iff P Q : P ∗-∗ Q ⊣⊢ (P ∗-∗ Q).
Proof. by rewrite embed_and !embed_wand. Qed.
Lemma embed_affinely_2 P : <affine> P ⎡<affine> P⎤.
Proof. by rewrite embed_and -embed_emp_2. Qed.
Lemma embed_affinely `{!BiEmbedEmp PROP1 PROP2} P : ⎡<affine> P ⊣⊢ <affine> P⎤.
Proof. by rewrite /bi_intuitionistically embed_and embed_emp. Qed.
Lemma embed_absorbingly P : ⎡<absorb> P ⊣⊢ <absorb> P⎤.
Proof. by rewrite embed_sep embed_pure. Qed.
Lemma embed_intuitionistically_2 P : P ⎡□ P⎤.
Proof. by rewrite /bi_intuitionistically -embed_affinely_2 embed_persistently. Qed.
Lemma embed_intuitionistically `{!BiEmbedEmp PROP1 PROP2} P : ⎡□ P ⊣⊢ P⎤.
Proof. by rewrite /bi_intuitionistically embed_affinely embed_persistently. Qed.
Lemma embed_persistently_if P b : ⎡<pers>?b P ⊣⊢ <pers>?b P⎤.
Proof. destruct b; simpl; auto using embed_persistently. Qed.
Lemma embed_affinely_if_2 P b : <affine>?b P ⎡<affine>?b P⎤.
Proof. destruct b; simpl; auto using embed_affinely_2. Qed.
Lemma embed_affinely_if `{!BiEmbedEmp PROP1 PROP2} P b :
⎡<affine>?b P ⊣⊢ <affine>?b P⎤.
Proof. destruct b; simpl; auto using embed_affinely. Qed.
Lemma embed_absorbingly_if b P : ⎡<absorb>?b P ⊣⊢ <absorb>?b P⎤.
Proof. destruct b; simpl; auto using embed_absorbingly. Qed.
Lemma embed_intuitionistically_if_2 P b : ?b P ⎡□?b P⎤.
Proof. destruct b; simpl; auto using embed_intuitionistically_2. Qed.
Lemma embed_intuitionistically_if `{!BiEmbedEmp PROP1 PROP2} P b :
⎡□?b P ⊣⊢ ?b P⎤.
Proof. destruct b; simpl; auto using embed_intuitionistically. Qed.
Global Instance embed_persistent P : Persistent P Persistent P⎤.
Proof. intros ?. by rewrite /Persistent -embed_persistently -persistent. Qed.
Global Instance embed_affine `{!BiEmbedEmp PROP1 PROP2} P : Affine P Affine P⎤.
Proof. intros ?. by rewrite /Affine (affine P) embed_emp. Qed.
Global Instance embed_absorbing P : Absorbing P Absorbing P⎤.
Proof. intros ?. by rewrite /Absorbing -embed_absorbingly absorbing. Qed.
Global Instance embed_and_homomorphism :
MonoidHomomorphism bi_and bi_and () embed.
Proof.
by split; [split|]; try apply _;
[setoid_rewrite embed_and|rewrite embed_pure].
Qed.
Global Instance embed_or_homomorphism :
MonoidHomomorphism bi_or bi_or () embed.
Proof.
by split; [split|]; try apply _;
[setoid_rewrite embed_or|rewrite embed_pure].
Qed.
Global Instance embed_sep_entails_homomorphism :
MonoidHomomorphism bi_sep bi_sep (flip ()) embed.
Proof.
split; [split|]; simpl; try apply _;
[by setoid_rewrite embed_sep|by rewrite embed_emp_2].
Qed.
Lemma embed_big_sepL_2 {A} (Φ : nat A PROP1) l :
([ list] kx l, Φ k x) [ list] kx l, Φ k x⎤.
Proof. apply (big_opL_commute (R:=flip ()) _). Qed.
Lemma embed_big_sepM_2 `{Countable K} {A} (Φ : K A PROP1) (m : gmap K A) :
([ map] kx m, Φ k x) [ map] kx m, Φ k x⎤.
Proof. apply (big_opM_commute (R:=flip ()) _). Qed.
Lemma embed_big_sepS_2 `{Countable A} (Φ : A PROP1) (X : gset A) :
([ set] y X, Φ y) [ set] y X, Φ y⎤.
Proof. apply (big_opS_commute (R:=flip ()) _). Qed.
Lemma embed_big_sepMS_2 `{Countable A} (Φ : A PROP1) (X : gmultiset A) :
([ mset] y X, Φ y) [ mset] y X, Φ y⎤.
Proof. apply (big_opMS_commute (R:=flip ()) _). Qed.
Section big_ops_emp.
Context `{!BiEmbedEmp PROP1 PROP2}.
Global Instance embed_sep_homomorphism :
MonoidHomomorphism bi_sep bi_sep () embed.
Proof.
by split; [split|]; try apply _;
[setoid_rewrite embed_sep|rewrite embed_emp].
Qed.
Lemma embed_big_sepL {A} (Φ : nat A PROP1) l :
[ list] kx l, Φ k x ⊣⊢ [ list] kx l, Φ k x⎤.
Proof. apply (big_opL_commute _). Qed.
Lemma embed_big_sepM `{Countable K} {A} (Φ : K A PROP1) (m : gmap K A) :
[ map] kx m, Φ k x ⊣⊢ [ map] kx m, Φ k x⎤.
Proof. apply (big_opM_commute _). Qed.
Lemma embed_big_sepS `{Countable A} (Φ : A PROP1) (X : gset A) :
[ set] y X, Φ y ⊣⊢ [ set] y X, Φ y⎤.
Proof. apply (big_opS_commute _). Qed.
Lemma embed_big_sepMS `{Countable A} (Φ : A PROP1) (X : gmultiset A) :
[ mset] y X, Φ y ⊣⊢ [ mset] y X, Φ y⎤.
Proof. apply (big_opMS_commute _). Qed.
End big_ops_emp.
Section later.
Context `{!BiEmbedLater PROP1 PROP2}.
Lemma embed_laterN n P : ⎡▷^n P ⊣⊢ ▷^n P⎤.
Proof. induction n=>//=. rewrite embed_later. by f_equiv. Qed.
Lemma embed_except_0 P : ⎡◇ P ⊣⊢ P⎤.
Proof. by rewrite embed_or embed_later embed_pure. Qed.
Global Instance embed_timeless P : Timeless P Timeless P⎤.
Proof.
intros ?. by rewrite /Timeless -embed_except_0 -embed_later timeless.
Qed.
End later.
Section internal_eq.
Context `{!BiInternalEq PROP1, !BiInternalEq PROP2, !BiEmbedInternalEq PROP1 PROP2}.
Lemma embed_internal_eq (A : ofe) (x y : A) : x y ⊣⊢@{PROP2} x y.
Proof.
apply bi.equiv_entails; split; [apply embed_internal_eq_1|].
etrans; [apply (internal_eq_rewrite x y (λ y, x y⎤%I)); solve_proper|].
rewrite -(internal_eq_refl True%I) embed_pure.
eapply bi.impl_elim; [done|]. apply bi.True_intro.
Qed.
End internal_eq.
Section plainly.
Context `{!BiPlainly PROP1, !BiPlainly PROP2, !BiEmbedPlainly PROP1 PROP2}.
Lemma embed_plainly_if p P : ⎡■?p P ⊣⊢ ?p P⎤.
Proof. destruct p; simpl; auto using embed_plainly. Qed.
Lemma embed_plain (P : PROP1) : Plain P Plain (PROP:=PROP2) P⎤.
Proof. intros ?. by rewrite /Plain {1}(plain P) embed_plainly. Qed.
End plainly.
End embed.
(* Not defined using an ordinary [Instance] because the default
[class_apply @bi_embed_plainly] shelves the [BiPlainly] premise, making proof
search for the other premises fail. See the proof of [monPred_objectively_plain]
for an example where it would fail with a regular [Instance].*)
Global Hint Extern 4 (Plain _) =>
notypeclasses refine (embed_plain _ _) : typeclass_instances.
(** Transitive embedding: this constructs an embedding of [PROP1] into [PROP3]
by combining the embeddings of [PROP1] into [PROP2] and [PROP2] into [PROP3].
Note that declaring these instances globally can make TC search ambiguous or
diverging. These are only defined so that a user can conveniently use them to
manually combine embeddings. *)
Section embed_embed.
Context {PROP1 PROP2 PROP3 : bi} `{!BiEmbed PROP1 PROP2, !BiEmbed PROP2 PROP3}.
Local Instance embed_embed : Embed PROP1 PROP3 := λ P, P ⎤%I.
Lemma embed_embedding_mixin : BiEmbedMixin PROP1 PROP3 embed_embed.
Proof.
split; unfold embed, embed_embed.
- solve_proper.
- solve_proper.
- intros P. by rewrite !embed_emp_valid.
- intros PROP' ? P Q. by rewrite !embed_interal_inj.
- by rewrite -!embed_emp_2.
- intros P Q. by rewrite -!embed_impl.
- intros A Φ. by rewrite -!embed_forall.
- intros A Φ. by rewrite -!embed_exist.
- intros P Q. by rewrite -!embed_sep.
- intros P Q. by rewrite -!embed_wand.
- intros P. by rewrite -!embed_persistently.
Qed.
Local Instance embed_bi_embed : BiEmbed PROP1 PROP3 :=
{| bi_embed_mixin := embed_embedding_mixin |}.
Lemma embed_embed_alt (P : PROP1) : P ⊣⊢@{PROP3} P ⎤.
Proof. done. Qed.
Lemma embed_embed_emp :
BiEmbedEmp PROP1 PROP2 BiEmbedEmp PROP2 PROP3
BiEmbedEmp PROP1 PROP3.
Proof. rewrite /BiEmbedEmp !embed_embed_alt. by intros -> ->. Qed.
Lemma embed_embed_later :
BiEmbedLater PROP1 PROP2 BiEmbedLater PROP2 PROP3
BiEmbedLater PROP1 PROP3.
Proof. intros ?? P. by rewrite !embed_embed_alt !embed_later. Qed.
Lemma embed_embed_internal_eq
`{!BiInternalEq PROP1, !BiInternalEq PROP2, !BiInternalEq PROP3} :
BiEmbedInternalEq PROP1 PROP2 BiEmbedInternalEq PROP2 PROP3
BiEmbedInternalEq PROP1 PROP3.
Proof. intros ?? A x y. by rewrite !embed_embed_alt !embed_internal_eq. Qed.
Lemma embed_embed_bupd `{!BiBUpd PROP1, !BiBUpd PROP2, !BiBUpd PROP3} :
BiEmbedBUpd PROP1 PROP2 BiEmbedBUpd PROP2 PROP3
BiEmbedBUpd PROP1 PROP3.
Proof. intros ?? P. by rewrite !embed_embed_alt !embed_bupd. Qed.
Lemma embed_embed_fupd `{!BiFUpd PROP1, !BiFUpd PROP2, !BiFUpd PROP3} :
BiEmbedFUpd PROP1 PROP2 BiEmbedFUpd PROP2 PROP3
BiEmbedFUpd PROP1 PROP3.
Proof. intros ?? E1 E2 P. by rewrite !embed_embed_alt !embed_fupd. Qed.
Lemma embed_embed_plainly
`{!BiPlainly PROP1, !BiPlainly PROP2, !BiPlainly PROP3} :
BiEmbedPlainly PROP1 PROP2 BiEmbedPlainly PROP2 PROP3
BiEmbedPlainly PROP1 PROP3.
Proof. intros ?? P. by rewrite !embed_embed_alt !embed_plainly. Qed.
End embed_embed.
(** This file defines various extensions to the base BI interface, via
typeclasses that BIs can optionally implement. *)
From iris.bi Require Export derived_connectives.
From iris.prelude Require Import options.
Class BiAffine (PROP : bi) := absorbing_bi (Q : PROP) : Affine Q.
Global Hint Mode BiAffine ! : typeclass_instances.
Global Existing Instance absorbing_bi | 0.
Class BiPositive (PROP : bi) :=
bi_positive (P Q : PROP) : <affine> (P Q) <affine> P Q.
Global Hint Mode BiPositive ! : typeclass_instances.
(** The class [BiLöb] is required for the [iLöb] tactic. However, for most BI
logics [BiLaterContractive] should be used, which gives an instance of [BiLöb]
automatically (see [derived_laws_later]). A direct instance of [BiLöb] is useful
when considering a BI logic with a discrete OFE, instead of an OFE that takes
step-indexing of the logic in account.
The internal/"strong" version of Löb [(▷ P → P) ⊢ P] is derivable from [BiLöb].
It is provided by the lemma [löb] in [derived_laws_later]. *)
Class BiLöb (PROP : bi) :=
löb_weak (P : PROP) : ( P P) (True P).
Global Hint Mode BiLöb ! : typeclass_instances.
Global Arguments löb_weak {_ _} _ _.
Class BiLaterContractive (PROP : bi) :=
#[global] later_contractive :: Contractive (bi_later (PROP:=PROP)).
(** The class [BiPersistentlyForall] states that universal quantification
commutes with the persistently modality. The reverse direction of the entailment
described by this type class is derivable, so it is not included. *)
Class BiPersistentlyForall (PROP : bi) :=
persistently_forall_2 : {A} (Ψ : A PROP), ( a, <pers> (Ψ a)) <pers> ( a, Ψ a).
Global Hint Mode BiPersistentlyForall ! : typeclass_instances.
(** The class [BiPureForall] states that universal quantification commutes with
the embedding of pure propositions. The reverse direction of the entailment
described by this type class is derivable, so it is not included.
An instance of [BiPureForall] itself is derivable if we assume excluded middle
in Coq, see the lemma [bi_pure_forall_em] in [derived_laws]. *)
Class BiPureForall (PROP : bi) :=
pure_forall_2 : {A} (φ : A Prop), ( a, φ a ) ⊢@{PROP} a, φ a ⌝.
Global Hint Mode BiPureForall ! : typeclass_instances.
From iris.algebra Require Export ofe stepindex_finite.
From iris.bi Require Export notation.
From iris.prelude Require Import options.
(* 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.
Section bi_mixin.
Context {PROP : Type} `{!Dist PROP, !Equiv PROP}.
Context (bi_entails : PROP PROP Prop).
Context (bi_emp : PROP).
Context (bi_pure : Prop PROP).
Context (bi_and : PROP PROP PROP).
Context (bi_or : PROP PROP PROP).
Context (bi_impl : PROP PROP PROP).
Context (bi_forall : A, (A PROP) PROP).
Context (bi_exist : A, (A PROP) PROP).
Context (bi_sep : PROP PROP PROP).
Context (bi_wand : PROP PROP PROP).
Bind Scope bi_scope with PROP.
Local Infix "⊢" := bi_entails.
Local Notation "'emp'" := bi_emp : bi_scope.
Local Notation "'True'" := (bi_pure True) : bi_scope.
Local Notation "'False'" := (bi_pure False) : bi_scope.
Local Notation "'⌜' φ '⌝'" := (bi_pure φ%type%stdpp) : bi_scope.
Local Infix "∧" := bi_and : bi_scope.
Local Infix "∨" := bi_or : bi_scope.
Local Infix "→" := bi_impl : bi_scope.
Local Notation "∀ x .. y , P" :=
(bi_forall _ (λ x, .. (bi_forall _ (λ y, P%I)) ..)) : bi_scope.
Local Notation "∃ x .. y , P" :=
(bi_exist _ (λ x, .. (bi_exist _ (λ y, P%I)) ..)) : bi_scope.
Local Infix "∗" := bi_sep : bi_scope.
Local Infix "-∗" := bi_wand : bi_scope.
(** * Axioms for a general BI (logic of bunched implications) *)
(** The following axioms are satisifed by both affine and linear BIs, and BIs
that combine both kinds of resources. In particular, we have an "ordered RA"
model satisfying all these axioms. For this model, we extend RAs with an
arbitrary partial order, and up-close resources wrt. that order (instead of
extension order). We demand composition to be monotone wrt. the order: [x1 ≼
x2 → x1 ⋅ y ≼ x2 ⋅ y]. We define [emp := λ r, ε ≼ r]; persistently is still
defined with the core: [persistently P := λ r, P (core r)]. This is uplcosed
because the core is monotone. *)
Record BiMixin := {
bi_mixin_entails_po : PreOrder bi_entails;
bi_mixin_equiv_entails P Q : (P Q) (P Q) (Q P);
(** Non-expansiveness *)
bi_mixin_pure_ne n : Proper (iff ==> dist n) bi_pure;
bi_mixin_and_ne : NonExpansive2 bi_and;
bi_mixin_or_ne : NonExpansive2 bi_or;
bi_mixin_impl_ne : NonExpansive2 bi_impl;
bi_mixin_forall_ne A n :
Proper (pointwise_relation _ (dist n) ==> dist n) (bi_forall A);
bi_mixin_exist_ne A n :
Proper (pointwise_relation _ (dist n) ==> dist n) (bi_exist A);
bi_mixin_sep_ne : NonExpansive2 bi_sep;
bi_mixin_wand_ne : NonExpansive2 bi_wand;
(** Higher-order logic *)
bi_mixin_pure_intro (φ : Prop) P : φ P φ ;
bi_mixin_pure_elim' (φ : Prop) P : (φ True P) φ P;
bi_mixin_and_elim_l P Q : P Q P;
bi_mixin_and_elim_r P Q : P Q Q;
bi_mixin_and_intro P Q R : (P Q) (P R) P Q R;
bi_mixin_or_intro_l P Q : P P Q;
bi_mixin_or_intro_r P Q : Q P Q;
bi_mixin_or_elim P Q R : (P R) (Q R) P Q R;
bi_mixin_impl_intro_r P Q R : (P Q R) P Q R;
bi_mixin_impl_elim_l' P Q R : (P Q R) P Q R;
bi_mixin_forall_intro {A} P (Ψ : A PROP) : ( a, P Ψ a) P a, Ψ a;
bi_mixin_forall_elim {A} {Ψ : A PROP} a : ( a, Ψ a) Ψ a;
bi_mixin_exist_intro {A} {Ψ : A PROP} a : Ψ a a, Ψ a;
bi_mixin_exist_elim {A} (Φ : A PROP) Q : ( a, Φ a Q) ( a, Φ a) Q;
(** BI connectives *)
bi_mixin_sep_mono P P' Q Q' : (P Q) (P' Q') P P' Q Q';
bi_mixin_emp_sep_1 P : P emp P;
bi_mixin_emp_sep_2 P : emp P P;
bi_mixin_sep_comm' P Q : P Q Q P;
bi_mixin_sep_assoc' P Q R : (P Q) R P (Q R);
bi_mixin_wand_intro_r P Q R : (P Q R) P Q -∗ R;
bi_mixin_wand_elim_l' P Q R : (P Q -∗ R) P Q R;
}.
(** We require any BI to have a persistence modality that carves out the
intuitionistic fragment of the separation logic. For logics such as Iris,
the persistence modality has a non-trivial definition (involving the [core] of
the camera). It is not clear whether a trivial definition exists: while
[<pers> P := False] comes close, it does not satisfy [later_persistently_1].
However, for some simpler discrete BIs the persistence modality
can be defined as:
<pers> P := ⌜ emp ⊢ P ⌝
That is, [P] holds persistently if it holds without resources.
The nesting of the entailment below the pure embedding ⌜ ⌝ only works for
discrete BIs: Non-expansiveness of [<pers>] relies on [dist] ignoring the
step-index.
To prove the rule [<pers> (∃ a, Ψ a) ⊢ ∃ a, <pers> Ψ a] the BI furthermore
needs to satisfy the "existential property": [emp ⊢ ∃ x, Φ x] implies
[∃ x, emp ⊢ Φ x].
This construction is formalized by the smart constructor
[bi_persistently_mixin_discrete] for [BiPersistentlyMixin]. See
[tests/heapprop] and [tests/heapprop_affine] for examples of how to use this
smart constructor. *)
Context (bi_persistently : PROP PROP).
Local Notation "'<pers>' P" := (bi_persistently P) : bi_scope.
Record BiPersistentlyMixin := {
bi_mixin_persistently_ne : NonExpansive bi_persistently;
(* In the ordered RA model: Holds without further assumptions. *)
bi_mixin_persistently_mono P Q : (P Q) <pers> P <pers> Q;
(* In the ordered RA model: `core` is idempotent *)
bi_mixin_persistently_idemp_2 P : <pers> P <pers> <pers> P;
(* In the ordered RA model: [ε ≼ core x]. *)
bi_mixin_persistently_emp_2 : emp <pers> emp;
(* The laws of a "frame" (https://ncatlab.org/nlab/show/frame, not to be
confused with separation logic terminology): commuting with finite
conjunction and infinite disjunction.
The null-ary case, [persistently_True : True ⊢ <pers> True], is derivable from the
other laws. *)
bi_mixin_persistently_and_2 (P Q : PROP) :
(<pers> P) (<pers> Q) <pers> (P Q);
bi_mixin_persistently_exist_1 {A} (Ψ : A PROP) :
<pers> ( a, Ψ a) a, <pers> (Ψ a);
(* In the ordered RA model: [core x ≼ core (x ⋅ y)]. *)
bi_mixin_persistently_absorbing P Q : <pers> P Q <pers> P;
(* In the ordered RA model: [x ⋅ core x = x]. *)
bi_mixin_persistently_and_sep_elim P Q : <pers> P Q P Q;
}.
Lemma bi_persistently_mixin_discrete :
( n (P Q : PROP), P {n} Q P Q)
( {A} (Φ : A PROP), (emp x, Φ x) x, emp Φ x)
( P : PROP, (<pers> P)%I = emp P ⌝%I)
BiMixin
BiPersistentlyMixin.
Proof.
intros Hdiscrete Hex Hpers Hbi. pose proof (bi_mixin_entails_po Hbi).
split.
- (* [NonExpansive bi_persistently] *)
intros n P Q [HPQ HQP]%Hdiscrete%(bi_mixin_equiv_entails Hbi).
rewrite !Hpers. apply (bi_mixin_pure_ne Hbi). split=> ?; by etrans.
- (*[(P ⊢ Q) → <pers> P ⊢ <pers> Q] *)
intros P Q HPQ. rewrite !Hpers. apply (bi_mixin_pure_elim' Hbi)=> ?.
apply (bi_mixin_pure_intro Hbi). by trans P.
- (* [<pers> P ⊢ <pers> <pers> P] *)
intros P. rewrite !Hpers. apply (bi_mixin_pure_elim' Hbi)=> ?.
by do 2 apply (bi_mixin_pure_intro Hbi).
- (* [emp ⊢ <pers> emp] *)
rewrite Hpers. by apply (bi_mixin_pure_intro Hbi).
- (* [<pers> P ∧ <pers> Q ⊢ <pers> (P ∧ Q)] *)
intros P Q. rewrite !Hpers.
apply (bi_mixin_impl_elim_l' Hbi). apply (bi_mixin_pure_elim' Hbi)=> ?.
apply (bi_mixin_impl_intro_r Hbi).
etrans; [apply (bi_mixin_and_elim_r Hbi)|].
apply (bi_mixin_pure_elim' Hbi)=> ?.
apply (bi_mixin_pure_intro Hbi). by apply (bi_mixin_and_intro Hbi).
- (* [<pers> (∃ a, Ψ a) ⊢ ∃ a, <pers> Ψ a] *)
intros A Φ. rewrite !Hpers. apply (bi_mixin_pure_elim' Hbi)=> /Hex [x ?].
etrans; [|apply (bi_mixin_exist_intro Hbi x)]; simpl.
rewrite Hpers. by apply (bi_mixin_pure_intro Hbi).
- (* [<pers> P ∗ Q ⊢ <pers> P] *)
intros P Q. rewrite !Hpers.
apply (bi_mixin_wand_elim_l' Hbi). apply (bi_mixin_pure_elim' Hbi)=> ?.
apply (bi_mixin_wand_intro_r Hbi). by apply (bi_mixin_pure_intro Hbi).
- (* [<pers> P ∧ Q ⊢ P ∗ Q] *)
intros P Q. rewrite !Hpers.
apply (bi_mixin_impl_elim_l' Hbi). apply (bi_mixin_pure_elim' Hbi)=> ?.
apply (bi_mixin_impl_intro_r Hbi).
etrans; [apply (bi_mixin_and_elim_r Hbi)|].
etrans; [apply (bi_mixin_emp_sep_1 Hbi)|].
by apply (bi_mixin_sep_mono Hbi).
Qed.
(** We equip any BI with a later modality. This avoids an additional layer in
the BI hierarchy and improves performance significantly (see Iris issue #303).
For non step-indexed BIs the later modality can simply be defined as the
identity function, as the Löb axiom or contractiveness of later is not part of
[BiLaterMixin]. For step-indexed BIs one should separately prove an instance
of the class [BiLaterContractive PROP] or [BiLöb PROP]. (Note that there is an
instance [BiLaterContractive PROP → BiLöb PROP] in [derived_laws_later].)
For non step-indexed BIs one can get a "free" instance of [BiLaterMixin] using
the smart constructor [bi_later_mixin_id] below. *)
Context (bi_later : PROP PROP).
Local Notation "▷ P" := (bi_later P) : bi_scope.
Record BiLaterMixin := {
bi_mixin_later_ne : NonExpansive bi_later;
bi_mixin_later_mono P Q : (P Q) P Q;
bi_mixin_later_intro P : P P;
bi_mixin_later_forall_2 {A} (Φ : A PROP) : ( a, Φ a) a, Φ a;
bi_mixin_later_exist_false {A} (Φ : A PROP) :
( a, Φ a) False ( a, Φ a);
bi_mixin_later_sep_1 P Q : (P Q) P Q;
bi_mixin_later_sep_2 P Q : P Q (P Q);
bi_mixin_later_persistently_1 P : <pers> P <pers> P;
bi_mixin_later_persistently_2 P : <pers> P <pers> P;
(* In a step-index model, this law allows case distinctions on whether
the step-index is 0 (expressed as [▷ False] in the logic):
* If it is 0, the left side is true, and we know nothing about [P].
* If not, then it is [S n] for some [n], and [P] holds at [n]. By down-
closure, it also holds at [0]. Thus, we get to use [P], but only if
the step-index is 0 ([▷ False] is true). *)
bi_mixin_later_false_em P : P False ( False P);
}.
Lemma bi_later_mixin_id :
( (P : PROP), ( P)%I = P)
BiMixin BiLaterMixin.
Proof.
intros Hlater Hbi. pose proof (bi_mixin_entails_po Hbi).
split; repeat intro; rewrite ?Hlater //.
- apply (bi_mixin_forall_intro Hbi)=> a.
etrans; [apply (bi_mixin_forall_elim Hbi a)|]. by rewrite Hlater.
- etrans; [|apply (bi_mixin_or_intro_r Hbi)].
apply (bi_mixin_exist_elim Hbi)=> a.
etrans; [|apply (bi_mixin_exist_intro Hbi a)]. by rewrite /= Hlater.
- etrans; [|apply (bi_mixin_or_intro_r Hbi)].
apply (bi_mixin_impl_intro_r Hbi), (bi_mixin_and_elim_l Hbi).
Qed.
End bi_mixin.
Module Import universes.
(** The universe of the logic (PROP). *)
Universe Logic.
(** The universe of quantifiers in the logic. *)
Universe Quant.
End universes.
Structure bi := Bi {
bi_car :> Type@{Logic};
bi_dist : Dist bi_car;
bi_equiv : Equiv bi_car;
bi_entails : bi_car bi_car Prop;
bi_emp : bi_car;
bi_pure : Prop bi_car;
bi_and : bi_car bi_car bi_car;
bi_or : bi_car bi_car bi_car;
bi_impl : bi_car bi_car bi_car;
bi_forall : A : Type@{Quant}, (A bi_car) bi_car;
bi_exist : A : Type@{Quant}, (A bi_car) bi_car;
bi_sep : bi_car bi_car bi_car;
bi_wand : bi_car bi_car bi_car;
bi_persistently : bi_car bi_car;
bi_later : bi_car bi_car;
bi_ofe_mixin : OfeMixin bi_car;
bi_cofe_aux : Cofe (Ofe bi_car bi_ofe_mixin);
bi_bi_mixin : BiMixin bi_entails bi_emp bi_pure bi_and bi_or bi_impl bi_forall
bi_exist bi_sep bi_wand;
bi_bi_persistently_mixin :
BiPersistentlyMixin bi_entails bi_emp bi_and bi_exist bi_sep bi_persistently;
bi_bi_later_mixin : BiLaterMixin bi_entails bi_pure bi_or bi_impl
bi_forall bi_exist bi_sep bi_persistently bi_later;
}.
Bind Scope bi_scope with bi_car.
Coercion bi_ofeO (PROP : bi) : ofe := Ofe PROP (bi_ofe_mixin PROP).
Canonical Structure bi_ofeO.
(** The projection [bi_cofe_aux] is not registered as an instance because it has
the wrong type. Its result type is unfolded, i.e., [Cofe (Ofe PROP ...)], and
thus should never be used. The instance [bi_cofe] has the proper result type
[Cofe (bi_ofeO PROP)]. *)
Global Instance bi_cofe (PROP : bi) : Cofe PROP := bi_cofe_aux PROP.
Global Instance: Params (@bi_entails) 1 := {}.
Global Instance: Params (@bi_emp) 1 := {}.
Global Instance: Params (@bi_pure) 1 := {}.
Global Instance: Params (@bi_and) 1 := {}.
Global Instance: Params (@bi_or) 1 := {}.
Global Instance: Params (@bi_impl) 1 := {}.
Global Instance: Params (@bi_forall) 2 := {}.
Global Instance: Params (@bi_exist) 2 := {}.
Global Instance: Params (@bi_sep) 1 := {}.
Global Instance: Params (@bi_wand) 1 := {}.
Global Instance: Params (@bi_persistently) 1 := {}.
Global Instance: Params (@bi_later) 1 := {}.
Global Arguments bi_car : simpl never.
Global Arguments bi_dist : simpl never.
Global Arguments bi_equiv : simpl never.
Global Arguments bi_entails {PROP} _ _ : simpl never, rename.
Global Arguments bi_emp {PROP} : simpl never, rename.
Global Arguments bi_pure {PROP} _%_stdpp : simpl never, rename.
Global Arguments bi_and {PROP} _ _ : simpl never, rename.
Global Arguments bi_or {PROP} _ _ : simpl never, rename.
Global Arguments bi_impl {PROP} _ _ : simpl never, rename.
Global Arguments bi_forall {PROP _} _%_I : simpl never, rename.
Global Arguments bi_exist {PROP _} _%_I : simpl never, rename.
Global Arguments bi_sep {PROP} _ _ : simpl never, rename.
Global Arguments bi_wand {PROP} _ _ : simpl never, rename.
Global Arguments bi_persistently {PROP} _ : simpl never, rename.
Global Arguments bi_later {PROP} _ : simpl never, rename.
Global Hint Extern 0 (bi_entails _ _) => reflexivity : core.
(** We set this rewrite relation's cost above the stdlib's
([impl], [iff], [eq], ...) and [≡] but below [⊑].
[eq] (at 100) < [≡] (at 150) < [bi_entails _] (at 170) < [⊑] (at 200)
*)
Global Instance bi_rewrite_relation (PROP : bi) : RewriteRelation (@bi_entails PROP) | 170 := {}.
Global Instance bi_inhabited {PROP : bi} : Inhabited PROP := populate (bi_pure True).
Notation "'emp'" := (bi_emp) : bi_scope.
Notation "'⌜' φ '⌝'" := (bi_pure φ%type%stdpp) : bi_scope.
Notation "'True'" := (bi_pure True) : bi_scope.
Notation "'False'" := (bi_pure False) : bi_scope.
Infix "∧" := bi_and : bi_scope.
Notation "(∧)" := bi_and (only parsing) : bi_scope.
Infix "∨" := bi_or : bi_scope.
Notation "(∨)" := bi_or (only parsing) : bi_scope.
Infix "→" := bi_impl : bi_scope.
Notation "¬ P" := (P False)%I : bi_scope.
Infix "∗" := bi_sep : bi_scope.
Notation "(∗)" := bi_sep (only parsing) : bi_scope.
Notation "P -∗ Q" := (bi_wand P Q) : bi_scope.
Notation "∀ x .. y , P" :=
(bi_forall (λ x, .. (bi_forall (λ y, P%I)) ..)) : bi_scope.
Notation "∃ x .. y , P" :=
(bi_exist (λ x, .. (bi_exist (λ y, P%I)) ..)) : bi_scope.
Notation "'<pers>' P" := (bi_persistently P) : bi_scope.
Notation "▷ P" := (bi_later P) : bi_scope.
Notation "P ⊢ Q" := (bi_entails P%I Q%I) : stdpp_scope.
Notation "P '⊢@{' PROP } Q" := (bi_entails (PROP:=PROP) P%I Q%I) (only parsing) : stdpp_scope.
Notation "(⊢)" := bi_entails (only parsing) : stdpp_scope.
Notation "'(⊢@{' PROP } )" := (bi_entails (PROP:=PROP)) (only parsing) : stdpp_scope.
Notation "P ⊣⊢ Q" := (equiv (A:=bi_car _) P%I Q%I) : stdpp_scope.
Notation "P '⊣⊢@{' PROP } Q" := (equiv (A:=bi_car PROP) P%I Q%I) (only parsing) : stdpp_scope.
Notation "(⊣⊢)" := (equiv (A:=bi_car _)) (only parsing) : stdpp_scope.
Notation "'(⊣⊢@{' PROP } )" := (equiv (A:=bi_car PROP)) (only parsing) : stdpp_scope.
Notation "( P ⊣⊢.)" := (equiv (A:=bi_car _) P) (only parsing) : stdpp_scope.
Notation "(.⊣⊢ Q )" := (λ P, P ≡@{bi_car _} Q) (only parsing) : stdpp_scope.
Definition bi_emp_valid {PROP : bi} (P : PROP) : Prop := emp P.
Global Arguments bi_emp_valid {_} _%_I : simpl never.
Global Typeclasses Opaque bi_emp_valid.
Notation "⊢ Q" := (bi_emp_valid Q%I) : stdpp_scope.
Notation "'⊢@{' PROP } Q" := (bi_emp_valid (PROP:=PROP) Q%I) (only parsing) : stdpp_scope.
(** Work around parsing issues: see [notation.v] for details. *)
Notation "'(⊢@{' PROP } Q )" := (bi_emp_valid (PROP:=PROP) Q%I) (only parsing) : stdpp_scope.
Notation "(.⊢ Q )" := (λ P, P Q) (only parsing) : stdpp_scope.
Notation "( P ⊢.)" := (bi_entails P) (only parsing) : stdpp_scope.
Notation "P -∗ Q" := ( P -∗ Q) : stdpp_scope.
Module bi.
Section bi_laws.
Context {PROP : bi}.
Implicit Types φ : Prop.
Implicit Types P Q R : PROP.
Implicit Types A : Type.
(* About the entailment *)
Global Instance entails_po : PreOrder (@bi_entails PROP).
Proof. eapply bi_mixin_entails_po, bi_bi_mixin. Qed.
Lemma equiv_entails P Q : P Q (P Q) (Q P).
Proof. eapply bi_mixin_equiv_entails, bi_bi_mixin. Qed.
(* Non-expansiveness *)
Global Instance pure_ne n : Proper (iff ==> dist n) (@bi_pure PROP).
Proof. eapply bi_mixin_pure_ne, bi_bi_mixin. Qed.
Global Instance and_ne : NonExpansive2 (@bi_and PROP).
Proof. eapply bi_mixin_and_ne, bi_bi_mixin. Qed.
Global Instance or_ne : NonExpansive2 (@bi_or PROP).
Proof. eapply bi_mixin_or_ne, bi_bi_mixin. Qed.
Global Instance impl_ne : NonExpansive2 (@bi_impl PROP).
Proof. eapply bi_mixin_impl_ne, bi_bi_mixin. Qed.
Global Instance forall_ne A n :
Proper (pointwise_relation _ (dist n) ==> dist n) (@bi_forall PROP A).
Proof. eapply bi_mixin_forall_ne, bi_bi_mixin. Qed.
Global Instance exist_ne A n :
Proper (pointwise_relation _ (dist n) ==> dist n) (@bi_exist PROP A).
Proof. eapply bi_mixin_exist_ne, bi_bi_mixin. Qed.
Global Instance sep_ne : NonExpansive2 (@bi_sep PROP).
Proof. eapply bi_mixin_sep_ne, bi_bi_mixin. Qed.
Global Instance wand_ne : NonExpansive2 (@bi_wand PROP).
Proof. eapply bi_mixin_wand_ne, bi_bi_mixin. Qed.
Global Instance persistently_ne : NonExpansive (@bi_persistently PROP).
Proof. eapply bi_mixin_persistently_ne, bi_bi_persistently_mixin. Qed.
(* Higher-order logic *)
Lemma pure_intro (φ : Prop) P : φ P φ ⌝.
Proof. eapply bi_mixin_pure_intro, bi_bi_mixin. Qed.
Lemma pure_elim' (φ : Prop) P : (φ True P) φ P.
Proof. eapply bi_mixin_pure_elim', bi_bi_mixin. Qed.
Lemma and_elim_l P Q : P Q P.
Proof. eapply bi_mixin_and_elim_l, bi_bi_mixin. Qed.
Lemma and_elim_r P Q : P Q Q.
Proof. eapply bi_mixin_and_elim_r, bi_bi_mixin. Qed.
Lemma and_intro P Q R : (P Q) (P R) P Q R.
Proof. eapply bi_mixin_and_intro, bi_bi_mixin. Qed.
Lemma or_intro_l P Q : P P Q.
Proof. eapply bi_mixin_or_intro_l, bi_bi_mixin. Qed.
Lemma or_intro_r P Q : Q P Q.
Proof. eapply bi_mixin_or_intro_r, bi_bi_mixin. Qed.
Lemma or_elim P Q R : (P R) (Q R) P Q R.
Proof. eapply bi_mixin_or_elim, bi_bi_mixin. Qed.
Lemma impl_intro_r P Q R : (P Q R) P Q R.
Proof. eapply bi_mixin_impl_intro_r, bi_bi_mixin. Qed.
Lemma impl_elim_l' P Q R : (P Q R) P Q R.
Proof. eapply bi_mixin_impl_elim_l', bi_bi_mixin. Qed.
Lemma forall_intro {A} P (Ψ : A PROP) : ( a, P Ψ a) P a, Ψ a.
Proof. eapply bi_mixin_forall_intro, bi_bi_mixin. Qed.
Lemma forall_elim {A} {Ψ : A PROP} a : ( a, Ψ a) Ψ a.
Proof. eapply (bi_mixin_forall_elim bi_entails), bi_bi_mixin. Qed.
Lemma exist_intro {A} {Ψ : A PROP} a : Ψ a a, Ψ a.
Proof. eapply bi_mixin_exist_intro, bi_bi_mixin. Qed.
Lemma exist_elim {A} (Φ : A PROP) Q : ( a, Φ a Q) ( a, Φ a) Q.
Proof. eapply bi_mixin_exist_elim, bi_bi_mixin. Qed.
(* BI connectives *)
Lemma sep_mono P P' Q Q' : (P Q) (P' Q') P P' Q Q'.
Proof. eapply bi_mixin_sep_mono, bi_bi_mixin. Qed.
Lemma emp_sep_1 P : P emp P.
Proof. eapply bi_mixin_emp_sep_1, bi_bi_mixin. Qed.
Lemma emp_sep_2 P : emp P P.
Proof. eapply bi_mixin_emp_sep_2, bi_bi_mixin. Qed.
Lemma sep_comm' P Q : P Q Q P.
Proof. eapply (bi_mixin_sep_comm' bi_entails), bi_bi_mixin. Qed.
Lemma sep_assoc' P Q R : (P Q) R P (Q R).
Proof. eapply bi_mixin_sep_assoc', bi_bi_mixin. Qed.
Lemma wand_intro_r P Q R : (P Q R) P Q -∗ R.
Proof. eapply bi_mixin_wand_intro_r, bi_bi_mixin. Qed.
Lemma wand_elim_l' P Q R : (P Q -∗ R) P Q R.
Proof. eapply bi_mixin_wand_elim_l', bi_bi_mixin. Qed.
(* Persistently *)
Lemma persistently_mono P Q : (P Q) <pers> P <pers> Q.
Proof. eapply bi_mixin_persistently_mono, bi_bi_persistently_mixin. Qed.
Lemma persistently_idemp_2 P : <pers> P <pers> <pers> P.
Proof. eapply bi_mixin_persistently_idemp_2, bi_bi_persistently_mixin. Qed.
Lemma persistently_emp_2 : emp ⊢@{PROP} <pers> emp.
Proof. eapply bi_mixin_persistently_emp_2, bi_bi_persistently_mixin. Qed.
Lemma persistently_and_2 (P Q : PROP) :
((<pers> P) (<pers> Q)) <pers> (P Q).
Proof. eapply bi_mixin_persistently_and_2, bi_bi_persistently_mixin. Qed.
Lemma persistently_exist_1 {A} (Ψ : A PROP) :
<pers> ( a, Ψ a) a, <pers> (Ψ a).
Proof. eapply bi_mixin_persistently_exist_1, bi_bi_persistently_mixin. Qed.
Lemma persistently_absorbing P Q : <pers> P Q <pers> P.
Proof.
eapply (bi_mixin_persistently_absorbing bi_entails), bi_bi_persistently_mixin.
Qed.
Lemma persistently_and_sep_elim P Q : <pers> P Q P Q.
Proof.
eapply (bi_mixin_persistently_and_sep_elim bi_entails), bi_bi_persistently_mixin.
Qed.
(* Later *)
Global Instance later_ne : NonExpansive (@bi_later PROP).
Proof. eapply bi_mixin_later_ne, bi_bi_later_mixin. Qed.
Lemma later_mono P Q : (P Q) P Q.
Proof. eapply bi_mixin_later_mono, bi_bi_later_mixin. Qed.
Lemma later_intro P : P P.
Proof. eapply bi_mixin_later_intro, bi_bi_later_mixin. Qed.
Lemma later_forall_2 {A} (Φ : A PROP) : ( a, Φ a) a, Φ a.
Proof. eapply bi_mixin_later_forall_2, bi_bi_later_mixin. Qed.
Lemma later_exist_false {A} (Φ : A PROP) :
( a, Φ a) False ( a, Φ a).
Proof. eapply bi_mixin_later_exist_false, bi_bi_later_mixin. Qed.
Lemma later_sep_1 P Q : (P Q) P Q.
Proof. eapply bi_mixin_later_sep_1, bi_bi_later_mixin. Qed.
Lemma later_sep_2 P Q : P Q (P Q).
Proof. eapply bi_mixin_later_sep_2, bi_bi_later_mixin. Qed.
Lemma later_persistently_1 P : <pers> P <pers> P.
Proof. eapply (bi_mixin_later_persistently_1 bi_entails), bi_bi_later_mixin. Qed.
Lemma later_persistently_2 P : <pers> P <pers> P.
Proof. eapply (bi_mixin_later_persistently_2 bi_entails), bi_bi_later_mixin. Qed.
Lemma later_false_em P : P False ( False P).
Proof. eapply bi_mixin_later_false_em, bi_bi_later_mixin. Qed.
End bi_laws.
End bi.
From iris.bi Require Import derived_laws_later big_op.
From iris.prelude Require Import options.
From iris.algebra Require Import excl csum.
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.
(** This file defines a type class for BIs with a notion of internal equality.
Internal equality is not part of the [bi] canonical structure as [internal_eq]
can only be given a definition that satisfies [NonExpansive2 internal_eq] _and_
[▷ (x ≡ y) ⊢ Next x ≡ Next y] if the BI is step-indexed. *)
Class InternalEq (PROP : Type) :=
internal_eq : {A : ofe}, A A PROP.
Global Arguments internal_eq {_ _ _} _ _ : simpl never.
Global Hint Mode InternalEq ! : typeclass_instances.
Global Instance: Params (@internal_eq) 3 := {}.
Global Typeclasses Opaque internal_eq.
Infix "≡" := internal_eq : bi_scope.
Infix "≡@{ A }" := (internal_eq (A := A)) (only parsing) : bi_scope.
Notation "( X ≡.)" := (internal_eq X) (only parsing) : bi_scope.
Notation "(.≡ X )" := (λ Y, Y X)%I (only parsing) : bi_scope.
Notation "(≡@{ A } )" := (internal_eq (A:=A)) (only parsing) : bi_scope.
(* Mixins allow us to create instances easily without having to use Program *)
Record BiInternalEqMixin (PROP : bi) `(!InternalEq PROP) := {
bi_internal_eq_mixin_internal_eq_ne (A : ofe) : NonExpansive2 (@internal_eq PROP _ A);
bi_internal_eq_mixin_internal_eq_refl {A : ofe} (P : PROP) (a : A) : P a a;
bi_internal_eq_mixin_internal_eq_rewrite {A : ofe} a b (Ψ : A PROP) :
NonExpansive Ψ a b Ψ a Ψ b;
bi_internal_eq_mixin_fun_extI {A} {B : A ofe} (f g : discrete_fun B) :
( x, f x g x) ⊢@{PROP} f g;
bi_internal_eq_mixin_sig_equivI_1 {A : ofe} (P : A Prop) (x y : sig P) :
`x `y ⊢@{PROP} x y;
bi_internal_eq_mixin_discrete_eq_1 {A : ofe} (a b : A) :
Discrete a a b ⊢@{PROP} a b;
bi_internal_eq_mixin_later_equivI_1 {A : ofe} (x y : A) :
Next x Next y ⊢@{PROP} (x y);
bi_internal_eq_mixin_later_equivI_2 {A : ofe} (x y : A) :
(x y) ⊢@{PROP} Next x Next y;
}.
Class BiInternalEq (PROP : bi) := {
#[global] bi_internal_eq_internal_eq :: InternalEq PROP;
bi_internal_eq_mixin : BiInternalEqMixin PROP bi_internal_eq_internal_eq;
}.
Global Hint Mode BiInternalEq ! : typeclass_instances.
Global Arguments bi_internal_eq_internal_eq : simpl never.
Section internal_eq_laws.
Context {PROP : bi} `{!BiInternalEq PROP}.
Implicit Types P Q : PROP.
Global Instance internal_eq_ne (A : ofe) : NonExpansive2 (@internal_eq PROP _ A).
Proof. eapply bi_internal_eq_mixin_internal_eq_ne, (bi_internal_eq_mixin). Qed.
Lemma internal_eq_refl {A : ofe} P (a : A) : P a a.
Proof. eapply bi_internal_eq_mixin_internal_eq_refl, bi_internal_eq_mixin. Qed.
Lemma internal_eq_rewrite {A : ofe} a b (Ψ : A PROP) :
NonExpansive Ψ a b Ψ a Ψ b.
Proof. eapply bi_internal_eq_mixin_internal_eq_rewrite, bi_internal_eq_mixin. Qed.
Lemma fun_extI {A} {B : A ofe} (f g : discrete_fun B) :
( x, f x g x) ⊢@{PROP} f g.
Proof. eapply bi_internal_eq_mixin_fun_extI, bi_internal_eq_mixin. Qed.
Lemma sig_equivI_1 {A : ofe} (P : A Prop) (x y : sig P) :
`x `y ⊢@{PROP} x y.
Proof. eapply bi_internal_eq_mixin_sig_equivI_1, bi_internal_eq_mixin. Qed.
Lemma discrete_eq_1 {A : ofe} (a b : A) :
Discrete a a b ⊢@{PROP} a b⌝.
Proof. eapply bi_internal_eq_mixin_discrete_eq_1, bi_internal_eq_mixin. Qed.
Lemma later_equivI_1 {A : ofe} (x y : A) : Next x Next y ⊢@{PROP} (x y).
Proof. eapply bi_internal_eq_mixin_later_equivI_1, bi_internal_eq_mixin. Qed.
Lemma later_equivI_2 {A : ofe} (x y : A) : (x y) ⊢@{PROP} Next x Next y.
Proof. eapply bi_internal_eq_mixin_later_equivI_2, bi_internal_eq_mixin. Qed.
End internal_eq_laws.
(* Derived laws *)
Section internal_eq_derived.
Context {PROP : bi} `{!BiInternalEq PROP}.
Implicit Types P : PROP.
(* Force implicit argument PROP *)
Notation "P ⊢ Q" := (P ⊢@{PROP} Q).
Notation "P ⊣⊢ Q" := (P ⊣⊢@{PROP} Q).
Global Instance internal_eq_proper (A : ofe) :
Proper (() ==> () ==> (⊣⊢)) (@internal_eq PROP _ A) := ne_proper_2 _.
(* Equality *)
Local Hint Resolve or_elim or_intro_l' or_intro_r' True_intro False_elim : core.
Local Hint Resolve and_elim_l' and_elim_r' and_intro forall_intro : core.
Local Hint Resolve internal_eq_refl : core.
Local Hint Extern 100 (NonExpansive _) => solve_proper : core.
Lemma equiv_internal_eq {A : ofe} P (a b : A) : a b P a b.
Proof. intros ->. auto. Qed.
Lemma internal_eq_rewrite' {A : ofe} a b (Ψ : A PROP) P
{ : NonExpansive Ψ} : (P a b) (P Ψ a) P Ψ b.
Proof.
intros Heq HΨa. rewrite -(idemp bi_and P) {1}Heq HΨa.
apply impl_elim_l'. by apply internal_eq_rewrite.
Qed.
Lemma internal_eq_sym {A : ofe} (a b : A) : a b b a.
Proof. apply (internal_eq_rewrite' a b (λ b, b a)%I); auto. Qed.
Lemma internal_eq_trans {A : ofe} (a b c : A) : a b b c a c.
Proof.
apply (internal_eq_rewrite' b a (λ b, b c)%I); auto.
rewrite and_elim_l. apply internal_eq_sym.
Qed.
Lemma internal_eq_iff P Q : P Q P Q.
Proof. apply (internal_eq_rewrite' P Q (λ Q, P Q))%I; auto using iff_refl. Qed.
Lemma f_equivI {A B : ofe} (f : A B) `{!NonExpansive f} x y :
x y f x f y.
Proof. apply (internal_eq_rewrite' x y (λ y, f x f y)%I); auto. Qed.
Lemma f_equivI_contractive {A B : ofe} (f : A B) `{Hf : !Contractive f} x y :
(x y) f x f y.
Proof.
rewrite later_equivI_2. move: Hf=>/contractive_alt [g [? Hfg]]. rewrite !Hfg.
by apply f_equivI.
Qed.
Lemma prod_equivI {A B : ofe} (x y : A * B) : x y ⊣⊢ x.1 y.1 x.2 y.2.
Proof.
apply (anti_symm _).
- apply and_intro; apply f_equivI; apply _.
- rewrite {3}(surjective_pairing x) {3}(surjective_pairing y).
apply (internal_eq_rewrite' (x.1) (y.1) (λ a, (x.1,x.2) (a,y.2))%I); auto.
apply (internal_eq_rewrite' (x.2) (y.2) (λ b, (x.1,x.2) (x.1,b))%I); auto.
Qed.
Lemma sum_equivI {A B : ofe} (x y : A + B) :
x y ⊣⊢
match x, y with
| inl a, inl a' => a a' | inr b, inr b' => b b' | _, _ => False
end.
Proof.
apply (anti_symm _).
- apply (internal_eq_rewrite' x y (λ y,
match x, y with
| inl a, inl a' => a a' | inr b, inr b' => b b' | _, _ => False
end)%I); auto.
destruct x; auto.
- destruct x as [a|b], y as [a'|b']; auto; apply f_equivI, _.
Qed.
Lemma option_equivI {A : ofe} (x y : option A) :
x y ⊣⊢ match x, y with
| Some a, Some a' => a a' | None, None => True | _, _ => False
end.
Proof.
apply (anti_symm _).
- apply (internal_eq_rewrite' x y (λ y,
match x, y with
| Some a, Some a' => a a' | None, None => True | _, _ => False
end)%I); auto.
destruct x; auto.
- destruct x as [a|], y as [a'|]; auto. apply f_equivI, _.
Qed.
Lemma csum_equivI {A B : ofe} (sx sy : csum A B) :
sx sy ⊣⊢ match sx, sy with
| Cinl x, Cinl y => x y
| Cinr x, Cinr y => x y
| CsumInvalid, CsumInvalid => True
| _, _ => False
end.
Proof.
apply (anti_symm _).
- apply (internal_eq_rewrite' sx sy (λ sy',
match sx, sy' with
| Cinl x, Cinl y => x y
| Cinr x, Cinr y => x y
| CsumInvalid, CsumInvalid => True
| _, _ => False
end)%I); [solve_proper|auto|].
destruct sx; eauto.
- destruct sx; destruct sy; eauto;
apply f_equivI, _.
Qed.
Lemma excl_equivI {O : ofe} (x y : excl O) :
x y ⊣⊢ match x, y with
| Excl a, Excl b => a b
| ExclInvalid, ExclInvalid => True
| _, _ => False
end.
Proof.
apply (anti_symm _).
- apply (internal_eq_rewrite' x y (λ y',
match x, y' with
| Excl a, Excl b => a b
| ExclInvalid, ExclInvalid => True
| _, _ => False
end)%I); [solve_proper|auto|].
destruct x; eauto.
- destruct x as [e1|]; destruct y as [e2|]; [|by eauto..].
apply f_equivI, _.
Qed.
Lemma sig_equivI {A : ofe} (P : A Prop) (x y : sig P) : `x `y ⊣⊢ x y.
Proof.
apply (anti_symm _).
- apply sig_equivI_1.
- apply f_equivI, _.
Qed.
Section sigT_equivI.
Import EqNotations.
Lemma sigT_equivI {A : Type} {P : A ofe} (x y : sigT P) :
x y ⊣⊢
eq : projT1 x = projT1 y, rew eq in projT2 x projT2 y.
Proof.
apply (anti_symm _).
- apply (internal_eq_rewrite' x y (λ y,
eq : projT1 x = projT1 y,
rew eq in projT2 x projT2 y))%I;
[| done | exact: (exist_intro' _ _ eq_refl) ].
move => n [a pa] [b pb] [/=]; intros -> => /= Hab.
apply exist_ne => ?. by rewrite Hab.
- apply exist_elim. move: x y => [a pa] [b pb] /=. intros ->; simpl.
apply f_equivI, _.
Qed.
End sigT_equivI.
Lemma discrete_fun_equivI {A} {B : A ofe} (f g : discrete_fun B) : f g ⊣⊢ x, f x g x.
Proof.
apply (anti_symm _); auto using fun_extI.
apply (internal_eq_rewrite' f g (λ g, x : A, f x g x)%I); auto.
Qed.
Lemma ofe_morO_equivI {A B : ofe} (f g : A -n> B) : f g ⊣⊢ x, f x g x.
Proof.
apply (anti_symm _).
- apply (internal_eq_rewrite' f g (λ g, x : A, f x g x)%I); auto.
- rewrite -(discrete_fun_equivI (ofe_mor_car _ _ f) (ofe_mor_car _ _ g)).
set (h1 (f : A -n> B) :=
exist (λ f : A -d> B, NonExpansive (f : A B)) f (ofe_mor_ne A B f)).
set (h2 (f : sigO (λ f : A -d> B, NonExpansive (f : A B))) :=
@OfeMor _ A B (`f) (proj2_sig f)).
assert ( f, h2 (h1 f) = f) as Hh by (by intros []).
assert (NonExpansive h2) by (intros ??? EQ; apply EQ).
by rewrite -{2}[f]Hh -{2}[g]Hh -f_equivI -sig_equivI.
Qed.
Lemma pure_internal_eq {A : ofe} (x y : A) : x y x y.
Proof. apply pure_elim'=> ->. apply internal_eq_refl. Qed.
Lemma discrete_eq {A : ofe} (a b : A) : Discrete a a b ⊣⊢ a b⌝.
Proof.
intros. apply (anti_symm _); auto using discrete_eq_1, pure_internal_eq.
Qed.
Lemma absorbingly_internal_eq {A : ofe} (x y : A) : <absorb> (x y) ⊣⊢ x y.
Proof.
apply (anti_symm _), absorbingly_intro.
apply wand_elim_r', (internal_eq_rewrite' x y (λ y, True -∗ x y)%I); auto.
apply wand_intro_l, internal_eq_refl.
Qed.
Lemma persistently_internal_eq {A : ofe} (a b : A) : <pers> (a b) ⊣⊢ a b.
Proof.
apply (anti_symm ()).
{ by rewrite persistently_into_absorbingly absorbingly_internal_eq. }
apply (internal_eq_rewrite' a b (λ b, <pers> (a b))%I); auto.
rewrite -(internal_eq_refl emp%I a). apply persistently_emp_intro.
Qed.
Global Instance internal_eq_absorbing {A : ofe} (x y : A) :
Absorbing (PROP:=PROP) (x y).
Proof. by rewrite /Absorbing absorbingly_internal_eq. Qed.
Global Instance internal_eq_persistent {A : ofe} (a b : A) :
Persistent (PROP:=PROP) (a b).
Proof. by intros; rewrite /Persistent persistently_internal_eq. Qed.
(* Equality under a later. *)
Lemma internal_eq_rewrite_contractive {A : ofe} a b (Ψ : A PROP)
{ : Contractive Ψ} : (a b) Ψ a Ψ b.
Proof.
rewrite f_equivI_contractive. apply (internal_eq_rewrite (Ψ a) (Ψ b) id _).
Qed.
Lemma internal_eq_rewrite_contractive' {A : ofe} a b (Ψ : A PROP) P
{ : Contractive Ψ} : (P (a b)) (P Ψ a) P Ψ b.
Proof.
rewrite later_equivI_2. move: =>/contractive_alt [g [? ]]. rewrite !.
by apply internal_eq_rewrite'.
Qed.
Lemma later_equivI {A : ofe} (x y : A) : Next x Next y ⊣⊢ (x y).
Proof. apply (anti_symm _); auto using later_equivI_1, later_equivI_2. Qed.
Lemma later_equivI_prop_2 `{!Contractive (bi_later (PROP:=PROP))} P Q :
(P Q) ( P) ( Q).
Proof. apply (f_equivI_contractive _). Qed.
Global Instance eq_timeless {A : ofe} (a b : A) :
Discrete a Timeless (PROP:=PROP) (a b).
Proof. intros. rewrite /Discrete !discrete_eq. apply (timeless _). Qed.
End internal_eq_derived.
From stdpp Require Import coPset namespaces.
From iris.bi Require Export bi updates.
From iris.bi.lib Require Import fixpoint_mono.
From iris.proofmode Require Import coq_tactics proofmode reduction.
From iris.prelude Require Import options.
(** Conveniently split a conjunction on both assumption and conclusion. *)
Local Tactic Notation "iSplitWith" constr(H) :=
iApply (bi.and_parallel with H); iSplit; iIntros H.
Section definition.
Context {PROP : bi} `{!BiFUpd PROP} {TA TB : tele}.
Implicit Types
(Eo Ei : coPset) (* outer/inner masks *)
(α : TA PROP) (* atomic pre-condition *)
(P : PROP) (* abortion condition *)
(β : TA TB PROP) (* atomic post-condition *)
(Φ : TA TB PROP) (* post-condition *)
.
(** atomic_acc as the "introduction form" of atomic updates: An accessor
that can be aborted back to [P]. *)
Definition atomic_acc Eo Ei α P β Φ : PROP :=
|={Eo, Ei}=> .. x, α x
((α x ={Ei, Eo}=∗ P) (.. y, β x y ={Ei, Eo}=∗ Φ x y)).
Lemma atomic_acc_wand Eo Ei α P1 P2 β Φ1 Φ2 :
((P1 -∗ P2) (.. x y, Φ1 x y -∗ Φ2 x y)) -∗
(atomic_acc Eo Ei α P1 β Φ1 -∗ atomic_acc Eo Ei α P2 β Φ2).
Proof.
iIntros "HP12 AS". iMod "AS" as (x) "[Hα Hclose]".
iModIntro. iExists x. iFrame "Hα". iSplit.
- iIntros "Hα". iDestruct "Hclose" as "[Hclose _]".
iApply "HP12". iApply "Hclose". done.
- iIntros (y) "Hβ". iDestruct "Hclose" as "[_ Hclose]".
iApply "HP12". iApply "Hclose". done.
Qed.
Lemma atomic_acc_mask Eo Ed α P β Φ :
atomic_acc Eo (EoEd) α P β Φ ⊣⊢ E, Eo E atomic_acc E (EEd) α P β Φ.
Proof.
iSplit; last first.
{ iIntros "Hstep". iApply ("Hstep" with "[% //]"). }
iIntros "Hstep" (E HE).
iApply (fupd_mask_frame_acc with "Hstep"); first done.
iIntros "Hstep". iDestruct "Hstep" as (x) "[Hα Hclose]".
iIntros "!> Hclose'".
iExists x. iFrame. iSplitWith "Hclose".
- iIntros "Hα". iApply "Hclose'". iApply "Hclose". done.
- iIntros (y) "Hβ". iApply "Hclose'". iApply "Hclose". done.
Qed.
Lemma atomic_acc_mask_weaken Eo1 Eo2 Ei α P β Φ :
Eo1 Eo2
atomic_acc Eo1 Ei α P β Φ -∗ atomic_acc Eo2 Ei α P β Φ.
Proof.
iIntros (HE) "Hstep".
iMod (fupd_mask_subseteq Eo1) as "Hclose1"; first done.
iMod "Hstep" as (x) "[Hα Hclose2]". iIntros "!>". iExists x.
iFrame. iSplitWith "Hclose2".
- iIntros "Hα". iMod ("Hclose2" with "Hα") as "$". done.
- iIntros (y) "Hβ". iMod ("Hclose2" with "Hβ") as "$". done.
Qed.
(** atomic_update as a fixed-point of the equation
AU = atomic_acc α AU β Q
*)
Context Eo Ei α β Φ.
Definition atomic_update_pre (Ψ : () PROP) (_ : ()) : PROP :=
atomic_acc Eo Ei α (Ψ ()) β Φ.
Local Instance atomic_update_pre_mono : BiMonoPred atomic_update_pre.
Proof.
constructor.
- iIntros (P1 P2 ??) "#HP12". iIntros ([]) "AU".
iApply (atomic_acc_wand with "[HP12] AU").
iSplit; last by eauto. iApply "HP12".
- intros ??. solve_proper.
Qed.
Local Definition atomic_update_def :=
bi_greatest_fixpoint atomic_update_pre ().
End definition.
(** Seal it *)
Local Definition atomic_update_aux : seal (@atomic_update_def).
Proof. by eexists. Qed.
Definition atomic_update := atomic_update_aux.(unseal).
Global Arguments atomic_update {PROP _ TA TB}.
Local Definition atomic_update_unseal :
@atomic_update = _ := atomic_update_aux.(seal_eq).
Global Arguments atomic_acc {PROP _ TA TB} Eo Ei _ _ _ _ : simpl never.
Global Arguments atomic_update {PROP _ TA TB} Eo Ei _ _ _ : simpl never.
(** Notation: Atomic updates *)
(** We avoid '<<'/'>>' since those can also reasonably be infix operators
(and in fact Autosubst uses the latter). *)
Notation "'AU' '<{' ∃∃ x1 .. xn , α '}>' @ Eo , Ei '<{' ∀∀ y1 .. yn , β , 'COMM' Φ '}>'" :=
(* 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_update (TA:=TeleS (λ x1, .. (TeleS (λ xn, TeleO)) .. ))
(TB:=TeleS (λ y1, .. (TeleS (λ yn, TeleO)) .. ))
Eo Ei
(tele_app $ λ x1, .. (λ xn, α%I) ..)
(tele_app $ λ x1, .. (λ xn,
tele_app (λ y1, .. (λ yn, β%I) .. )
) .. )
(tele_app $ λ x1, .. (λ xn,
tele_app (λ y1, .. (λ yn, Φ%I) .. )
) .. )
)
(at level 20, Eo, Ei, α, β, Φ at level 200, x1 binder, xn binder, y1 binder, yn binder,
format "'[hv ' 'AU' '<{' '[' ∃∃ x1 .. xn , '/' α ']' '}>' '/' @ '[' Eo , '/' Ei ']' '/' '<{' '[' ∀∀ y1 .. yn , '/' β , '/' COMM Φ ']' '}>' ']'") : bi_scope.
Notation "'AU' '<{' ∃∃ x1 .. xn , α '}>' @ Eo , Ei '<{' β , 'COMM' Φ '}>'" :=
(atomic_update (TA:=TeleS (λ x1, .. (TeleS (λ xn, TeleO)) .. ))
(TB:=TeleO)
Eo Ei
(tele_app $ λ x1, .. (λ xn, α%I) ..)
(tele_app $ λ x1, .. (λ xn, tele_app β%I) .. )
(tele_app $ λ x1, .. (λ xn, tele_app Φ%I) .. )
)
(at level 20, Eo, Ei, α, β, Φ at level 200, x1 binder, xn binder,
format "'[hv ' 'AU' '<{' '[' ∃∃ x1 .. xn , '/' α ']' '}>' '/' @ '[' Eo , '/' Ei ']' '/' '<{' '[' β , '/' COMM Φ ']' '}>' ']'") : bi_scope.
Notation "'AU' '<{' α '}>' @ Eo , Ei '<{' ∀∀ y1 .. yn , β , 'COMM' Φ '}>'" :=
(atomic_update (TA:=TeleO)
(TB:=TeleS (λ y1, .. (TeleS (λ yn, TeleO)) .. ))
Eo Ei
(tele_app α%I)
(tele_app $ tele_app (λ y1, .. (λ yn, β%I) ..))
(tele_app $ tele_app (λ y1, .. (λ yn, Φ%I) ..))
)
(at level 20, Eo, Ei, α, β, Φ at level 200, y1 binder, yn binder,
format "'[hv ' 'AU' '<{' '[' α ']' '}>' '/' @ '[' Eo , '/' Ei ']' '/' '<{' '[' ∀∀ y1 .. yn , '/' β , '/' COMM Φ ']' '}>' ']'") : bi_scope.
Notation "'AU' '<{' α '}>' @ Eo , Ei '<{' β , 'COMM' Φ '}>'" :=
(atomic_update (TA:=TeleO) (TB:=TeleO)
Eo Ei
(tele_app α%I)
(tele_app $ tele_app β%I)
(tele_app $ tele_app Φ%I)
)
(at level 20, Eo, Ei, α, β, Φ at level 200,
format "'[hv ' 'AU' '<{' '[' α ']' '}>' '/' @ '[' Eo , '/' Ei ']' '/' '<{' '[' β , '/' COMM Φ ']' '}>' ']'") : bi_scope.
(** Notation: Atomic accessors *)
Notation "'AACC' '<{' ∃∃ x1 .. xn , α , 'ABORT' P '}>' @ Eo , Ei '<{' ∀∀ y1 .. yn , β , 'COMM' Φ '}>'" :=
(atomic_acc (TA:=TeleS (λ x1, .. (TeleS (λ xn, TeleO)) .. ))
(TB:=TeleS (λ y1, .. (TeleS (λ yn, TeleO)) .. ))
Eo Ei
(tele_app $ λ x1, .. (λ xn, α%I) ..)
P%I
(tele_app $ λ x1, .. (λ xn,
tele_app (λ y1, .. (λ yn, β%I) .. )
) .. )
(tele_app $ λ x1, .. (λ xn,
tele_app (λ y1, .. (λ yn, Φ%I) .. )
) .. )
)
(at level 20, Eo, Ei, α, P, β, Φ at level 200, x1 binder, xn binder, y1 binder, yn binder,
format "'[hv ' 'AACC' '<{' '[' ∃∃ x1 .. xn , '/' α , '/' ABORT P ']' '}>' '/' @ '[' Eo , '/' Ei ']' '/' '<{' '[' ∀∀ y1 .. yn , '/' β , '/' COMM Φ ']' '}>' ']'") : bi_scope.
Notation "'AACC' '<{' ∃∃ x1 .. xn , α , 'ABORT' P '}>' @ Eo , Ei '<{' β , 'COMM' Φ '}>'" :=
(atomic_acc (TA:=TeleS (λ x1, .. (TeleS (λ xn, TeleO)) .. ))
(TB:=TeleO)
Eo Ei
(tele_app $ λ x1, .. (λ xn, α%I) ..)
P%I
(tele_app $ λ x1, .. (λ xn, tele_app β%I) .. )
(tele_app $ λ x1, .. (λ xn, tele_app Φ%I) .. )
)
(at level 20, Eo, Ei, α, P, β, Φ at level 200, x1 binder, xn binder,
format "'[hv ' 'AACC' '<{' '[' ∃∃ x1 .. xn , '/' α , '/' ABORT P ']' '}>' '/' @ '[' Eo , '/' Ei ']' '/' '<{' '[' β , '/' COMM Φ ']' '}>' ']'") : bi_scope.
Notation "'AACC' '<{' α , 'ABORT' P '}>' @ Eo , Ei '<{' ∀∀ y1 .. yn , β , 'COMM' Φ '}>'" :=
(atomic_acc (TA:=TeleO)
(TB:=TeleS (λ y1, .. (TeleS (λ yn, TeleO)) .. ))
Eo Ei
(tele_app α%I)
P%I
(tele_app $ tele_app (λ y1, .. (λ yn, β%I) ..))
(tele_app $ tele_app (λ y1, .. (λ yn, Φ%I) ..))
)
(at level 20, Eo, Ei, α, P, β, Φ at level 200, y1 binder, yn binder,
format "'[hv ' 'AACC' '<{' '[' α , '/' ABORT P ']' '}>' '/' @ '[' Eo , '/' Ei ']' '/' '<{' '[' ∀∀ y1 .. yn , '/' β , '/' COMM Φ ']' '}>' ']'") : bi_scope.
Notation "'AACC' '<{' α , 'ABORT' P '}>' @ Eo , Ei '<{' β , 'COMM' Φ '}>'" :=
(atomic_acc (TA:=TeleO)
(TB:=TeleO)
Eo Ei
(tele_app α%I)
P%I
(tele_app $ tele_app β%I)
(tele_app $ tele_app Φ%I)
)
(at level 20, Eo, Ei, α, P, β, Φ at level 200,
format "'[hv ' 'AACC' '<{' '[' α , '/' ABORT P ']' '}>' '/' @ '[' Eo , '/' Ei ']' '/' '<{' '[' β , '/' COMM Φ ']' '}>' ']'") : bi_scope.
(** Lemmas about AU *)
Section lemmas.
Context `{BiFUpd PROP} {TA TB : tele}.
Implicit Types (α : TA PROP) (β Φ : TA TB PROP) (P : PROP).
Local Existing Instance atomic_update_pre_mono.
(* Can't be in the section above as that fixes the parameters *)
Global Instance atomic_acc_ne Eo Ei n :
Proper (
pointwise_relation TA (dist n) ==>
dist n ==>
pointwise_relation TA (pointwise_relation TB (dist n)) ==>
pointwise_relation TA (pointwise_relation TB (dist n)) ==>
dist n
) (atomic_acc (PROP:=PROP) Eo Ei).
Proof. solve_proper. Qed.
Global Instance atomic_update_ne Eo Ei n :
Proper (
pointwise_relation TA (dist n) ==>
pointwise_relation TA (pointwise_relation TB (dist n)) ==>
pointwise_relation TA (pointwise_relation TB (dist n)) ==>
dist n
) (atomic_update (PROP:=PROP) Eo Ei).
Proof.
rewrite atomic_update_unseal /atomic_update_def /atomic_update_pre. solve_proper.
Qed.
Lemma atomic_update_mask_weaken Eo1 Eo2 Ei α β Φ :
Eo1 Eo2
atomic_update Eo1 Ei α β Φ -∗ atomic_update Eo2 Ei α β Φ.
Proof.
rewrite atomic_update_unseal {2}/atomic_update_def /=.
iIntros (Heo) "HAU".
iApply (greatest_fixpoint_coiter _ (λ _, atomic_update_def Eo1 Ei α β Φ)); last done.
iIntros "!> *". rewrite {1}/atomic_update_def /= greatest_fixpoint_unfold.
iApply atomic_acc_mask_weaken. done.
Qed.
Local Lemma aupd_unfold Eo Ei α β Φ :
atomic_update Eo Ei α β Φ ⊣⊢
atomic_acc Eo Ei α (atomic_update Eo Ei α β Φ) β Φ.
Proof.
rewrite atomic_update_unseal /atomic_update_def /=. apply: greatest_fixpoint_unfold.
Qed.
(** The elimination form: an atomic accessor *)
Lemma aupd_aacc Eo Ei α β Φ :
atomic_update Eo Ei α β Φ
atomic_acc Eo Ei α (atomic_update Eo Ei α β Φ) β Φ.
Proof using Type*. by rewrite {1}aupd_unfold. Qed.
(* This lets you eliminate atomic updates with iMod. *)
Global Instance elim_mod_aupd φ Eo Ei E α β Φ Q Q' :
( R, ElimModal φ false false (|={E,Ei}=> R) R Q Q')
ElimModal (φ Eo E) false false
(atomic_update Eo Ei α β Φ)
(.. x, α x
(α x ={Ei,E}=∗ atomic_update Eo Ei α β Φ)
(.. y, β x y ={Ei,E}=∗ Φ x y))
Q Q'.
Proof.
intros ?. rewrite /ElimModal /= =>-[??]. iIntros "[AU Hcont]".
iPoseProof (aupd_aacc with "AU") as "AC".
iMod (atomic_acc_mask_weaken with "AC"); first done.
iApply "Hcont". done.
Qed.
(** The introduction lemma for atomic_update. This should usually not be used
directly; use the [iAuIntro] tactic instead. *)
Local Lemma aupd_intro P Q α β Eo Ei Φ :
Absorbing P Persistent P
(P Q atomic_acc Eo Ei α Q β Φ)
P Q atomic_update Eo Ei α β Φ.
Proof.
rewrite atomic_update_unseal {1}/atomic_update_def /=.
iIntros (?? HAU) "[#HP HQ]".
iApply (greatest_fixpoint_coiter _ (λ _, Q)); last done. iIntros "!>" ([]) "HQ".
iApply HAU. iSplit; by iFrame.
Qed.
Lemma aacc_intro Eo Ei α P β Φ :
Ei Eo (.. x, α x -∗
((α x ={Eo}=∗ P) (.. y, β x y ={Eo}=∗ Φ x y)) -∗
atomic_acc Eo Ei α P β Φ).
Proof.
iIntros (? x) "Hα Hclose".
iApply fupd_mask_intro; first set_solver. iIntros "Hclose'".
iExists x. iFrame. iSplitWith "Hclose".
- iIntros "Hα". iMod "Hclose'" as "_". iApply "Hclose". done.
- iIntros (y) "Hβ". iMod "Hclose'" as "_". iApply "Hclose". done.
Qed.
(* This lets you open invariants etc. when the goal is an atomic accessor. *)
Global Instance elim_acc_aacc {X} E1 E2 Ei (α' β' : X PROP) γ' α β Pas Φ :
ElimAcc (X:=X) True (fupd E1 E2) (fupd E2 E1) α' β' γ'
(atomic_acc E1 Ei α Pas β Φ)
(λ x', atomic_acc E2 Ei α (β' x' (γ' x' -∗? Pas))%I β
(λ.. x y, β' x' (γ' x' -∗? Φ x y))
)%I.
Proof.
(* FIXME: Is there any way to prevent maybe_wand from unfolding?
It gets unfolded by env_cbv in the proofmode, ideally we'd like that
to happen only if one argument is a constructor. *)
iIntros (_) "Hinner >Hacc". iDestruct "Hacc" as (x') "[Hα' Hclose]".
iMod ("Hinner" with "Hα'") as (x) "[Hα Hclose']".
iApply fupd_mask_intro; first set_solver. iIntros "Hclose''".
iExists x. iFrame. iSplitWith "Hclose'".
- iIntros "Hα". iMod "Hclose''" as "_".
iMod ("Hclose'" with "Hα") as "[Hβ' HPas]".
iMod ("Hclose" with "Hβ'") as "Hγ'".
iModIntro. destruct (γ' x'); iApply "HPas"; done.
- iIntros (y) "Hβ". iMod "Hclose''" as "_".
iMod ("Hclose'" with "Hβ") as "Hβ'".
(* FIXME: Using ssreflect rewrite does not work, see Coq bug #7773. *)
rewrite ->!tele_app_bind. iDestruct "Hβ'" as "[Hβ' HΦ]".
iMod ("Hclose" with "Hβ'") as "Hγ'".
iModIntro. destruct (γ' x'); iApply "HΦ"; done.
Qed.
(* Everything that fancy updates can eliminate without changing, atomic
accessors can eliminate as well. This is a forwarding instance needed because
atomic_acc is becoming opaque. *)
Global Instance elim_modal_acc p q φ P P' Eo Ei α Pas β Φ :
( Q, ElimModal φ p q P P' (|={Eo,Ei}=> Q) (|={Eo,Ei}=> Q))
ElimModal φ p q P P'
(atomic_acc Eo Ei α Pas β Φ)
(atomic_acc Eo Ei α Pas β Φ).
Proof. intros Helim. apply Helim. Qed.
(** Lemmas for directly proving one atomic accessor in terms of another (or an
atomic update). These are only really useful when the atomic accessor you
are trying to prove exactly corresponds to an atomic update/accessor you
have as an assumption -- which is not very common. *)
Lemma aacc_aacc {TA' TB' : tele} E1 E1' E2 E3
α P β Φ
(α' : TA' PROP) P' (β' Φ' : TA' TB' PROP) :
E1' E1
atomic_acc E1' E2 α P β Φ -∗
(.. x, α x -∗ atomic_acc E2 E3 α' (α x (P ={E1}=∗ P')) β'
(λ.. x' y', (α x (P ={E1}=∗ Φ' x' y'))
.. y, β x y (Φ x y ={E1}=∗ Φ' x' y'))) -∗
atomic_acc E1 E3 α' P' β' Φ'.
Proof.
iIntros (?) "Hupd Hstep".
iMod (atomic_acc_mask_weaken with "Hupd") as (x) "[Hα Hclose]"; first done.
iMod ("Hstep" with "Hα") as (x') "[Hα' Hclose']".
iModIntro. iExists x'. iFrame "Hα'". iSplit.
- iIntros "Hα'". iDestruct "Hclose'" as "[Hclose' _]".
iMod ("Hclose'" with "Hα'") as "[Hα Hupd]".
iDestruct "Hclose" as "[Hclose _]".
iMod ("Hclose" with "Hα"). iApply "Hupd". auto.
- iIntros (y') "Hβ'". iDestruct "Hclose'" as "[_ Hclose']".
iMod ("Hclose'" with "Hβ'") as "Hres".
(* FIXME: Using ssreflect rewrite does not work, see Coq bug #7773. *)
rewrite ->!tele_app_bind. iDestruct "Hres" as "[[Hα HΦ']|Hcont]".
+ (* Abort the step we are eliminating *)
iDestruct "Hclose" as "[Hclose _]".
iMod ("Hclose" with "Hα") as "HP".
iApply "HΦ'". done.
+ (* Complete the step we are eliminating *)
iDestruct "Hclose" as "[_ Hclose]".
iDestruct "Hcont" as (y) "[Hβ HΦ']".
iMod ("Hclose" with "Hβ") as "HΦ".
iApply "HΦ'". done.
Qed.
Lemma aacc_aupd {TA' TB' : tele} E1 E1' E2 E3
α β Φ
(α' : TA' PROP) P' (β' Φ' : TA' TB' PROP) :
E1' E1
atomic_update E1' E2 α β Φ -∗
(.. x, α x -∗ atomic_acc E2 E3 α' (α x (atomic_update E1' E2 α β Φ ={E1}=∗ P')) β'
(λ.. x' y', (α x (atomic_update E1' E2 α β Φ ={E1}=∗ Φ' x' y'))
.. y, β x y (Φ x y ={E1}=∗ Φ' x' y'))) -∗
atomic_acc E1 E3 α' P' β' Φ'.
Proof.
iIntros (?) "Hupd Hstep". iApply (aacc_aacc with "[Hupd] Hstep"); first done.
iApply aupd_aacc; done.
Qed.
Lemma aacc_aupd_commit {TA' TB' : tele} E1 E1' E2 E3
α β Φ
(α' : TA' PROP) P' (β' Φ' : TA' TB' PROP) :
E1' E1
atomic_update E1' E2 α β Φ -∗
(.. x, α x -∗ atomic_acc E2 E3 α' (α x (atomic_update E1' E2 α β Φ ={E1}=∗ P')) β'
(λ.. x' y', .. y, β x y (Φ x y ={E1}=∗ Φ' x' y'))) -∗
atomic_acc E1 E3 α' P' β' Φ'.
Proof.
iIntros (?) "Hupd Hstep". iApply (aacc_aupd with "Hupd"); first done.
iIntros (x) "Hα". iApply atomic_acc_wand; last first.
{ iApply "Hstep". done. }
(* FIXME: Using ssreflect rewrite does not work, see Coq bug #7773. *)
iSplit; first by eauto. iIntros (??) "?". rewrite ->!tele_app_bind. by iRight.
Qed.
Lemma aacc_aupd_abort {TA' TB' : tele} E1 E1' E2 E3
α β Φ
(α' : TA' PROP) P' (β' Φ' : TA' TB' PROP) :
E1' E1
atomic_update E1' E2 α β Φ -∗
(.. x, α x -∗ atomic_acc E2 E3 α' (α x (atomic_update E1' E2 α β Φ ={E1}=∗ P')) β'
(λ.. x' y', α x (atomic_update E1' E2 α β Φ ={E1}=∗ Φ' x' y'))) -∗
atomic_acc E1 E3 α' P' β' Φ'.
Proof.
iIntros (?) "Hupd Hstep". iApply (aacc_aupd with "Hupd"); first done.
iIntros (x) "Hα". iApply atomic_acc_wand; last first.
{ iApply "Hstep". done. }
(* FIXME: Using ssreflect rewrite does not work, see Coq bug #7773. *)
iSplit; first by eauto. iIntros (??) "?". rewrite ->!tele_app_bind. by iLeft.
Qed.
End lemmas.
(** ProofMode support for atomic updates. *)
Section proof_mode.
Context `{BiFUpd PROP} {TA TB : tele}.
Implicit Types (α : TA PROP) (β Φ : TA TB PROP) (P : PROP).
Lemma tac_aupd_intro Γp Γs n α β Eo Ei Φ P :
P = env_to_prop Γs
envs_entails (Envs Γp Γs n) (atomic_acc Eo Ei α P β Φ)
envs_entails (Envs Γp Γs n) (atomic_update Eo Ei α β Φ).
Proof.
intros ->. rewrite envs_entails_unseal of_envs_eq /atomic_acc /=.
setoid_rewrite env_to_prop_sound =>HAU.
rewrite assoc. apply: aupd_intro. by rewrite -assoc.
Qed.
End proof_mode.
(** * Now the coq-level tactics *)
Tactic Notation "iAuIntro" :=
match goal with
| |- envs_entails (Envs ?Γp ?Γs _) (atomic_update _ _ _ _ ) =>
notypeclasses refine (tac_aupd_intro Γp Γs _ _ _ _ _ Φ _ _ _); [
(* P = ...: make the P pretty *) pm_reflexivity
| (* the new proof mode goal *) ]
end.
(** Tactic to apply [aacc_intro]. This only really works well when you have
[α ?] already and pass it as [iAaccIntro with "Hα"]. Doing
[rewrite /atomic_acc /=] is an entirely legitimate alternative. *)
Tactic Notation "iAaccIntro" "with" constr(sel) :=
iStartProof; lazymatch goal with
| |- envs_entails _ (@atomic_acc ?PROP ?H ?TA ?TB ?Eo ?Ei ?P ) =>
iApply (@aacc_intro PROP H TA TB Eo Ei α P β Φ with sel);
first try solve_ndisj; last iSplit
| _ => fail "iAAccIntro: Goal is not an atomic accessor"
end.
(* From here on, prevent TC search from implicitly unfolding these. *)
Global Typeclasses Opaque atomic_acc atomic_update.
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.