Newer
Older
From stdpp Require Import nat_cancel.
From iris.proofmode Require Export modality_instances classes.
Set Default Proof Using "Type".
Import bi.
Section bi_instances.
Context {PROP : bi}.
Implicit Types P Q R : PROP.
(* FromAffinely *)
Global Instance from_affinely_affine P : Affine P → FromAffinely P P.
Proof. intros. by rewrite /FromAffinely affinely_elim. Qed.
Global Instance from_affinely_default P : FromAffinely (bi_affinely P) P | 100.
Proof. by rewrite /FromAffinely. Qed.
(* IntoAbsorbingly *)
Global Instance into_absorbingly_True : @IntoAbsorbingly PROP True emp | 0.
Proof. by rewrite /IntoAbsorbingly -absorbingly_True_emp absorbingly_pure. Qed.
Global Instance into_absorbingly_absorbing P : Absorbing P → IntoAbsorbingly P P | 1.
Proof. intros. by rewrite /IntoAbsorbingly absorbing_absorbingly. Qed.
Global Instance into_absorbingly_default P : IntoAbsorbingly (bi_absorbingly P) P | 100.
Proof. by rewrite /IntoAbsorbingly. Qed.
(* FromAssumption *)
Global Instance from_assumption_exact p P : FromAssumption p P P | 0.
Proof. by rewrite /FromAssumption /= affinely_persistently_if_elim. Qed.
Global Instance from_assumption_persistently_r P Q :
FromAssumption true P Q → FromAssumption true P (bi_persistently Q).
by rewrite -{1}affinely_persistently_idemp affinely_elim.
Global Instance from_assumption_affinely_r P Q :
FromAssumption true P Q → FromAssumption true P (bi_affinely Q).
Proof. rewrite /FromAssumption /= =><-. by rewrite affinely_idemp. Qed.
Global Instance from_assumption_absorbingly_r p P Q :
FromAssumption p P Q → FromAssumption p P (bi_absorbingly Q).
Proof. rewrite /FromAssumption /= =><-. apply absorbingly_intro. Qed.
Global Instance from_assumption_affinely_persistently_l p P Q :
FromAssumption true P Q → FromAssumption p (□ P) Q.
Proof. rewrite /FromAssumption /= =><-. by rewrite affinely_persistently_if_elim. Qed.
Global Instance from_assumption_persistently_l_true P Q :
FromAssumption true P Q → FromAssumption true (bi_persistently P) Q.
Proof. rewrite /FromAssumption /= =><-. by rewrite persistently_idemp. Qed.
Global Instance from_assumption_persistently_l_false `{BiAffine PROP} P Q :
FromAssumption true P Q → FromAssumption false (bi_persistently P) Q.
Proof. rewrite /FromAssumption /= =><-. by rewrite affine_affinely. Qed.
Global Instance from_assumption_affinely_l_true p P Q :
FromAssumption p P Q → FromAssumption p (bi_affinely P) Q.
Proof. rewrite /FromAssumption /= =><-. by rewrite affinely_elim. Qed.
Global Instance from_assumption_forall {A} p (Φ : A → PROP) Q x :
FromAssumption p (Φ x) Q → FromAssumption p (∀ x, Φ x) Q.
Proof. rewrite /FromAssumption=> <-. by rewrite forall_elim. Qed.
(* IntoPure *)
Global Instance into_pure_pure φ : @IntoPure PROP ⌜φ⌝ φ.
Proof. by rewrite /IntoPure. Qed.
Proof. rewrite /IntoPure pure_and. by intros -> ->. Qed.
Proof. rewrite /IntoPure pure_or. by intros -> ->. Qed.
Global Instance into_pure_pure_impl (φ1 φ2 : Prop) P1 P2 :
FromPure false P1 φ1 → IntoPure P2 φ2 → IntoPure (P1 → P2) (φ1 → φ2).
Proof. rewrite /FromPure /IntoPure pure_impl=> <- -> //. Qed.
Global Instance into_pure_exist {A} (Φ : A → PROP) (φ : A → Prop) :
(∀ x, IntoPure (Φ x) (φ x)) → IntoPure (∃ x, Φ x) (∃ x, φ x).
Proof. rewrite /IntoPure=>Hx. rewrite pure_exist. by setoid_rewrite Hx. Qed.
Global Instance into_pure_forall {A} (Φ : A → PROP) (φ : A → Prop) :
(∀ x, IntoPure (Φ x) (φ x)) → IntoPure (∀ x, Φ x) (∀ x, φ x).
Proof. rewrite /IntoPure=>Hx. rewrite -pure_forall_2. by setoid_rewrite Hx. Qed.
Global Instance into_pure_pure_sep (φ1 φ2 : Prop) P1 P2 :
IntoPure P1 φ1 → IntoPure P2 φ2 → IntoPure (P1 ∗ P2) (φ1 ∧ φ2).
Proof. rewrite /IntoPure=> -> ->. by rewrite sep_and pure_and. Qed.
Global Instance into_pure_pure_wand (φ1 φ2 : Prop) P1 P2 :
FromPure true P1 φ1 → IntoPure P2 φ2 → IntoPure (P1 -∗ P2) (φ1 → φ2).
Proof.
rewrite /FromPure /IntoPure=> <- -> /=.
rewrite pure_impl -impl_wand_2. apply bi.wand_intro_l.
rewrite {1}(persistent_absorbingly_affinely ⌜φ1⌝%I) absorbingly_sep_l
bi.wand_elim_r absorbing //.
Qed.
Global Instance into_pure_affinely P φ :
IntoPure P φ → IntoPure (bi_affinely P) φ.
Proof. rewrite /IntoPure=> ->. apply affinely_elim. Qed.
Global Instance into_pure_absorbingly P φ : IntoPure P φ → IntoPure (bi_absorbingly P) φ.
Proof. rewrite /IntoPure=> ->. by rewrite absorbingly_pure. Qed.
Global Instance into_pure_persistently P φ :
IntoPure P φ → IntoPure (bi_persistently P) φ.
Proof. rewrite /IntoPure=> ->. apply: persistently_elim. Qed.
Global Instance into_pure_embed `{BiEmbed PROP PROP'} P φ :
IntoPure P φ → IntoPure ⎡P⎤ φ.
Proof. rewrite /IntoPure=> ->. by rewrite embed_pure. Qed.
(* FromPure *)
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
Global Instance from_pure_pure a φ : @FromPure PROP a ⌜φ⌝ φ.
Proof. rewrite /FromPure. apply affinely_if_elim. Qed.
Global Instance from_pure_pure_and a (φ1 φ2 : Prop) P1 P2 :
FromPure a P1 φ1 → FromPure a P2 φ2 → FromPure a (P1 ∧ P2) (φ1 ∧ φ2).
Proof. rewrite /FromPure pure_and=> <- <- /=. by rewrite affinely_if_and. Qed.
Global Instance from_pure_pure_or a (φ1 φ2 : Prop) P1 P2 :
FromPure a P1 φ1 → FromPure a P2 φ2 → FromPure a (P1 ∨ P2) (φ1 ∨ φ2).
Proof. by rewrite /FromPure pure_or affinely_if_or=><- <-. Qed.
Global Instance from_pure_pure_impl a (φ1 φ2 : Prop) P1 P2 :
IntoPure P1 φ1 → FromPure a P2 φ2 → FromPure a (P1 → P2) (φ1 → φ2).
Proof.
rewrite /FromPure /IntoPure pure_impl=> -> <-. destruct a=>//=.
apply bi.impl_intro_l. by rewrite affinely_and_r bi.impl_elim_r.
Qed.
Global Instance from_pure_exist {A} a (Φ : A → PROP) (φ : A → Prop) :
(∀ x, FromPure a (Φ x) (φ x)) → FromPure a (∃ x, Φ x) (∃ x, φ x).
Proof.
rewrite /FromPure=>Hx. rewrite pure_exist affinely_if_exist.
by setoid_rewrite Hx.
Qed.
Global Instance from_pure_forall {A} a (Φ : A → PROP) (φ : A → Prop) :
(∀ x, FromPure a (Φ x) (φ x)) → FromPure a (∀ x, Φ x) (∀ x, φ x).
Proof.
rewrite /FromPure=>Hx. rewrite pure_forall. setoid_rewrite <-Hx.
destruct a=>//=. apply affinely_forall.
Qed.
Global Instance from_pure_pure_sep_true (φ1 φ2 : Prop) P1 P2 :
FromPure true P1 φ1 → FromPure true P2 φ2 → FromPure true (P1 ∗ P2) (φ1 ∧ φ2).
Proof.
rewrite /FromPure=> <- <- /=.
by rewrite -persistent_and_affinely_sep_l affinely_and_r pure_and.
Qed.
Global Instance from_pure_pure_sep_false_l (φ1 φ2 : Prop) P1 P2 :
FromPure false P1 φ1 → FromPure true P2 φ2 → FromPure false (P1 ∗ P2) (φ1 ∧ φ2).
Proof.
rewrite /FromPure=> <- <- /=. by rewrite -persistent_and_affinely_sep_r pure_and.
Qed.
Global Instance from_pure_pure_sep_false_r (φ1 φ2 : Prop) P1 P2 :
FromPure true P1 φ1 → FromPure false P2 φ2 → FromPure false (P1 ∗ P2) (φ1 ∧ φ2).
Proof.
rewrite /FromPure=> <- <- /=. by rewrite -persistent_and_affinely_sep_l pure_and.
Qed.
Global Instance from_pure_pure_wand (φ1 φ2 : Prop) a P1 P2 :
IntoPure P1 φ1 → FromPure false P2 φ2 → FromPure a (P1 -∗ P2) (φ1 → φ2).
Proof.
rewrite /FromPure /IntoPure=> -> <- /=.
by rewrite bi.affinely_if_elim pure_wand_forall pure_impl pure_impl_forall.
Global Instance from_pure_persistently P a φ :
FromPure true P φ → FromPure a (bi_persistently P) φ.
Proof.
rewrite /FromPure=> <- /=.
by rewrite persistently_affinely affinely_if_elim persistently_pure.
Qed.
Global Instance from_pure_affinely_true P φ :
FromPure true P φ → FromPure true (bi_affinely P) φ.
Proof. rewrite /FromPure=><- /=. by rewrite affinely_idemp. Qed.
Global Instance from_pure_affinely_false P φ `{!Affine P} :
FromPure false P φ → FromPure false (bi_affinely P) φ.
Proof. rewrite /FromPure /= affine_affinely //. Qed.
Global Instance from_pure_absorbingly P φ :
FromPure true P φ → FromPure false (bi_absorbingly P) φ.
Proof. rewrite /FromPure=> <- /=. apply persistent_absorbingly_affinely, _. Qed.
Global Instance from_pure_embed `{BiEmbed PROP PROP'} a P φ :
FromPure a P φ → FromPure a ⎡P⎤ φ.
Proof. rewrite /FromPure=> <-. by rewrite embed_affinely_if embed_pure. Qed.
(* IntoPersistent *)
Global Instance into_persistent_persistently p P Q :
IntoPersistent true P Q → IntoPersistent p (bi_persistently P) Q | 0.
Proof.
rewrite /IntoPersistent /= => ->.
destruct p; simpl; auto using persistently_idemp_1.
Qed.
Global Instance into_persistent_affinely p P Q :
IntoPersistent p P Q → IntoPersistent p (bi_affinely P) Q | 0.
Proof. rewrite /IntoPersistent /= => <-. by rewrite affinely_elim. Qed.
Global Instance into_persistent_embed `{BiEmbed PROP PROP'} p P Q :
IntoPersistent p P Q → IntoPersistent p ⎡P⎤ ⎡Q⎤ | 0.
Proof.
rewrite /IntoPersistent -embed_persistently -embed_persistently_if=> -> //.
Qed.
Global Instance into_persistent_here P : IntoPersistent true P P | 1.
Global Instance into_persistent_persistent P :
Persistent P → IntoPersistent false P P | 100.
Proof. intros. by rewrite /IntoPersistent. Qed.
(* FromModal *)
Global Instance from_modal_affinely P :
FromModal modality_affinely (bi_affinely P) (bi_affinely P) P | 2.
Proof. by rewrite /FromModal. Qed.
Global Instance from_modal_persistently P :
FromModal modality_persistently (bi_persistently P) (bi_persistently P) P | 2.
Proof. by rewrite /FromModal. Qed.
Global Instance from_modal_affinely_persistently P :
FromModal modality_affinely_persistently (□ P) (□ P) P | 1.
Proof. by rewrite /FromModal. Qed.
Global Instance from_modal_affinely_persistently_affine_bi P :
BiAffine PROP → FromModal modality_persistently (□ P) (□ P) P | 0.
Proof. intros. by rewrite /FromModal /= affine_affinely. Qed.
Global Instance from_modal_absorbingly P :
FromModal modality_id (bi_absorbingly P) (bi_absorbingly P) P.
Proof. by rewrite /FromModal /= -absorbingly_intro. Qed.
(* When having a modality nested in an embedding, e.g. [ ⎡|==> P⎤ ], we prefer
the embedding over the modality. *)
Global Instance from_modal_embed `{BiEmbed PROP PROP'} (P : PROP) :
FromModal (@modality_embed PROP PROP' _) ⎡P⎤ ⎡P⎤ P.
Proof. by rewrite /FromModal. Qed.
Global Instance from_modal_id_embed `{BiEmbed PROP PROP'} `(sel : A) P Q :
FromModal modality_id sel P Q →
FromModal modality_id sel ⎡P⎤ ⎡Q⎤ | 100.
Proof. by rewrite /FromModal /= =><-. Qed.
Global Instance from_modal_affinely_embed `{BiEmbed PROP PROP'} `(sel : A) P Q :
FromModal modality_affinely sel P Q →
FromModal modality_affinely sel ⎡P⎤ ⎡Q⎤ | 100.
Proof. rewrite /FromModal /= =><-. by rewrite embed_affinely. Qed.
Global Instance from_modal_persistently_embed `{BiEmbed PROP PROP'} `(sel : A) P Q :
FromModal modality_persistently sel P Q →
FromModal modality_persistently sel ⎡P⎤ ⎡Q⎤ | 100.
Proof. rewrite /FromModal /= =><-. by rewrite embed_persistently. Qed.
Global Instance from_modal_affinely_persistently_embed `{BiEmbed PROP PROP'} `(sel : A) P Q :
FromModal modality_affinely_persistently sel P Q →
FromModal modality_affinely_persistently sel ⎡P⎤ ⎡Q⎤ | 100.
rewrite /FromModal /= =><-. by rewrite embed_affinely embed_persistently.
(* IntoWand *)
Global Instance into_wand_wand p q P Q P' :
FromAssumption q P P' → IntoWand p q (P' -∗ Q) P Q.
rewrite /FromAssumption /IntoWand=> HP. by rewrite HP affinely_persistently_if_elim.
Global Instance into_wand_impl_false_false P Q P' :
Absorbing P' → Absorbing (P' → Q) →
FromAssumption false P P' → IntoWand false false (P' → Q) P Q.
Proof.
rewrite /FromAssumption /IntoWand /= => ?? ->. apply wand_intro_r.
by rewrite sep_and impl_elim_l.
Qed.
Global Instance into_wand_impl_false_true P Q P' :
Absorbing P' → FromAssumption true P P' →
rewrite /IntoWand /FromAssumption /= => ? HP. apply wand_intro_l.
rewrite -(affinely_persistently_idemp P) HP.
by rewrite -persistently_and_affinely_sep_l persistently_elim impl_elim_r.
Global Instance into_wand_impl_true_false P Q P' :
Affine P' → FromAssumption false P P' →
rewrite /FromAssumption /IntoWand /= => ? HP. apply wand_intro_r.
rewrite -persistently_and_affinely_sep_l HP -{2}(affine_affinely P') -affinely_and_lr.
by rewrite affinely_persistently_elim impl_elim_l.
Global Instance into_wand_impl_true_true P Q P' :
FromAssumption true P P' → IntoWand true true (P' → Q) P Q.
rewrite /FromAssumption /IntoWand /= => <-. apply wand_intro_l.
rewrite -{1}(affinely_persistently_idemp P) -and_sep_affinely_persistently.
by rewrite -affinely_persistently_and impl_elim_r affinely_persistently_elim.
Global Instance into_wand_and_l p q R1 R2 P' Q' :
IntoWand p q R1 P' Q' → IntoWand p q (R1 ∧ R2) P' Q'.
Proof. rewrite /IntoWand=> ?. by rewrite /bi_wand_iff and_elim_l. Qed.
Global Instance into_wand_and_r p q R1 R2 P' Q' :
IntoWand p q R2 Q' P' → IntoWand p q (R1 ∧ R2) Q' P'.
Proof. rewrite /IntoWand=> ?. by rewrite /bi_wand_iff and_elim_r. Qed.
Global Instance into_wand_forall_prop_true p (φ : Prop) P :
IntoWand p true (∀ _ : φ, P) ⌜ φ ⌝ P.
Proof.
rewrite /IntoWand (affinely_persistently_if_elim p) /=
-impl_wand_affinely_persistently -pure_impl_forall
bi.persistently_elim //.
Qed.
Global Instance into_wand_forall_prop_false p (φ : Prop) P :
Absorbing P → IntoWand p false (∀ _ : φ, P) ⌜ φ ⌝ P.
Proof.
intros ?.
rewrite /IntoWand (affinely_persistently_if_elim p) /= pure_wand_forall //.
Qed.
Global Instance into_wand_forall {A} p q (Φ : A → PROP) P Q x :
IntoWand p q (Φ x) P Q → IntoWand p q (∀ x, Φ x) P Q.
Proof. rewrite /IntoWand=> <-. by rewrite (forall_elim x). Qed.
Global Instance into_wand_affinely_persistently p q R P Q :
IntoWand p q R P Q → IntoWand p q (□ R) P Q.
Proof. by rewrite /IntoWand affinely_persistently_elim. Qed.
Global Instance into_wand_persistently_true q R P Q :
IntoWand true q R P Q → IntoWand true q (bi_persistently R) P Q.
Proof. by rewrite /IntoWand /= persistently_idemp. Qed.
Global Instance into_wand_persistently_false `{!BiAffine PROP} q R P Q :
IntoWand false q R P Q → IntoWand false q (bi_persistently R) P Q.
Proof. by rewrite /IntoWand persistently_elim. Qed.
Global Instance into_wand_embed `{BiEmbed PROP PROP'} p q R P Q :
IntoWand p q R P Q → IntoWand p q ⎡R⎤ ⎡P⎤ ⎡Q⎤.
Proof.
rewrite /IntoWand -!embed_persistently_if -!embed_affinely_if
-embed_wand => -> //.
Qed.
Jacques-Henri Jourdan
committed
(* FromWand *)
Global Instance from_wand_wand P1 P2 : FromWand (P1 -∗ P2) P1 P2.
Proof. by rewrite /FromWand. Qed.
Global Instance from_wand_embed `{BiEmbed PROP PROP'} P Q1 Q2 :
Jacques-Henri Jourdan
committed
FromWand P Q1 Q2 → FromWand ⎡P⎤ ⎡Q1⎤ ⎡Q2⎤.
Proof. by rewrite /FromWand -embed_wand => <-. Qed.
Jacques-Henri Jourdan
committed
(* FromImpl *)
Global Instance from_impl_impl P1 P2 : FromImpl (P1 → P2) P1 P2.
Proof. by rewrite /FromImpl. Qed.
Global Instance from_impl_embed `{BiEmbed PROP PROP'} P Q1 Q2 :
Jacques-Henri Jourdan
committed
FromImpl P Q1 Q2 → FromImpl ⎡P⎤ ⎡Q1⎤ ⎡Q2⎤.
Proof. by rewrite /FromImpl -embed_impl => <-. Qed.
Jacques-Henri Jourdan
committed
(* FromAnd *)
Global Instance from_and_and P1 P2 : FromAnd (P1 ∧ P2) P1 P2 | 100.
Proof. by rewrite /FromAnd. Qed.
Global Instance from_and_sep_persistent_l P1 P1' P2 :
FromAffinely P1 P1' → Persistent P1' → FromAnd (P1 ∗ P2) P1' P2 | 9.
rewrite /FromAffinely /FromAnd=> <- ?. by rewrite persistent_and_affinely_sep_l_1.
Global Instance from_and_sep_persistent_r P1 P2 P2' :
FromAffinely P2 P2' → Persistent P2' → FromAnd (P1 ∗ P2) P1 P2' | 10.
rewrite /FromAffinely /FromAnd=> <- ?. by rewrite persistent_and_affinely_sep_r_1.
Global Instance from_and_sep_persistent P1 P2 :
Persistent P1 → Persistent P2 → FromAnd (P1 ∗ P2) P1 P2 | 11.
rewrite /FromAffinely /FromAnd. intros ??. by rewrite -persistent_and_sep_1.
Global Instance from_and_pure φ ψ : @FromAnd PROP ⌜φ ∧ ψ⌝ ⌜φ⌝ ⌜ψ⌝.
Proof. by rewrite /FromAnd pure_and. Qed.
Global Instance from_and_persistently P Q1 Q2 :
FromAnd P Q1 Q2 →
FromAnd (bi_persistently P) (bi_persistently Q1) (bi_persistently Q2).
Proof. rewrite /FromAnd=> <-. by rewrite persistently_and. Qed.
Global Instance from_and_persistently_sep P Q1 Q2 :
FromSep P Q1 Q2 →
FromAnd (bi_persistently P) (bi_persistently Q1) (bi_persistently Q2) | 11.
Proof. rewrite /FromAnd=> <-. by rewrite -persistently_and persistently_and_sep. Qed.
Global Instance from_and_embed `{BiEmbed PROP PROP'} P Q1 Q2 :
FromAnd P Q1 Q2 → FromAnd ⎡P⎤ ⎡Q1⎤ ⎡Q2⎤.
Proof. by rewrite /FromAnd -embed_and => <-. Qed.
Global Instance from_and_big_sepL_cons_persistent {A} (Φ : nat → A → PROP) x l :
Persistent (Φ 0 x) →
FromAnd ([∗ list] k ↦ y ∈ x :: l, Φ k y) (Φ 0 x) ([∗ list] k ↦ y ∈ l, Φ (S k) y).
Proof. intros. by rewrite /FromAnd big_opL_cons persistent_and_sep_1. Qed.
Global Instance from_and_big_sepL_app_persistent {A} (Φ : nat → A → PROP) l1 l2 :
(∀ k y, Persistent (Φ k y)) →
FromAnd ([∗ list] k ↦ y ∈ l1 ++ l2, Φ k y)
([∗ list] k ↦ y ∈ l1, Φ k y) ([∗ list] k ↦ y ∈ l2, Φ (length l1 + k) y).
Proof. intros. by rewrite /FromAnd big_opL_app persistent_and_sep_1. Qed.
(* FromSep *)
Global Instance from_sep_sep P1 P2 : FromSep (P1 ∗ P2) P1 P2 | 100.
Proof. by rewrite /FromSep. Qed.
Global Instance from_sep_and P1 P2 :
TCOr (Affine P1) (Absorbing P2) → TCOr (Absorbing P1) (Affine P2) →
FromSep (P1 ∧ P2) P1 P2 | 101.
Proof. intros. by rewrite /FromSep sep_and. Qed.
Global Instance from_sep_pure φ ψ : @FromSep PROP ⌜φ ∧ ψ⌝ ⌜φ⌝ ⌜ψ⌝.
Proof. by rewrite /FromSep pure_and sep_and. Qed.
Global Instance from_sep_affinely P Q1 Q2 :
FromSep P Q1 Q2 → FromSep (bi_affinely P) (bi_affinely Q1) (bi_affinely Q2).
Proof. rewrite /FromSep=> <-. by rewrite affinely_sep_2. Qed.
Global Instance from_sep_absorbingly P Q1 Q2 :
FromSep P Q1 Q2 → FromSep (bi_absorbingly P) (bi_absorbingly Q1) (bi_absorbingly Q2).
Proof. rewrite /FromSep=> <-. by rewrite absorbingly_sep. Qed.
Global Instance from_sep_persistently P Q1 Q2 :
FromSep P Q1 Q2 →
FromSep (bi_persistently P) (bi_persistently Q1) (bi_persistently Q2).
Proof. rewrite /FromSep=> <-. by rewrite persistently_sep_2. Qed.
Global Instance from_sep_embed `{BiEmbed PROP PROP'} P Q1 Q2 :
FromSep P Q1 Q2 → FromSep ⎡P⎤ ⎡Q1⎤ ⎡Q2⎤.
Proof. by rewrite /FromSep -embed_sep => <-. Qed.
Global Instance from_sep_big_sepL_cons {A} (Φ : nat → A → PROP) x l :
FromSep ([∗ list] k ↦ y ∈ x :: l, Φ k y) (Φ 0 x) ([∗ list] k ↦ y ∈ l, Φ (S k) y).
Proof. by rewrite /FromSep big_sepL_cons. Qed.
Global Instance from_sep_big_sepL_app {A} (Φ : nat → A → PROP) l1 l2 :
FromSep ([∗ list] k ↦ y ∈ l1 ++ l2, Φ k y)
([∗ list] k ↦ y ∈ l1, Φ k y) ([∗ list] k ↦ y ∈ l2, Φ (length l1 + k) y).
Proof. by rewrite /FromSep big_opL_app. Qed.
Global Instance into_and_and p P Q : IntoAnd p (P ∧ Q) P Q | 10.
Proof. by rewrite /IntoAnd affinely_persistently_if_and. Qed.
Global Instance into_and_and_affine_l P Q Q' :
Affine P → FromAffinely Q' Q → IntoAnd false (P ∧ Q) P Q'.
Proof.
intros. rewrite /IntoAnd /=.
by rewrite -(affine_affinely P) affinely_and_l affinely_and (from_affinely Q').
Qed.
Global Instance into_and_and_affine_r P P' Q :
Affine Q → FromAffinely P' P → IntoAnd false (P ∧ Q) P' Q.
Proof.
intros. rewrite /IntoAnd /=.
by rewrite -(affine_affinely Q) affinely_and_r affinely_and (from_affinely P').
Global Instance into_and_sep `{BiPositive PROP} P Q : IntoAnd true (P ∗ Q) P Q.
Proof.
by rewrite /IntoAnd /= persistently_sep -and_sep_persistently persistently_and.
Qed.
Global Instance into_and_sep_affine P Q :
TCOr (Affine P) (Absorbing Q) → TCOr (Absorbing P) (Affine Q) →
IntoAnd true (P ∗ Q) P Q.
Proof. intros. by rewrite /IntoAnd /= sep_and. Qed.
Global Instance into_and_pure p φ ψ : @IntoAnd PROP p ⌜φ ∧ ψ⌝ ⌜φ⌝ ⌜ψ⌝.
Proof. by rewrite /IntoAnd pure_and affinely_persistently_if_and. Qed.
Global Instance into_and_affinely p P Q1 Q2 :
IntoAnd p P Q1 Q2 → IntoAnd p (bi_affinely P) (bi_affinely Q1) (bi_affinely Q2).
- by rewrite -affinely_and !persistently_affinely.
- intros ->. by rewrite affinely_and.
Global Instance into_and_persistently p P Q1 Q2 :
IntoAnd p P Q1 Q2 →
IntoAnd p (bi_persistently P) (bi_persistently Q1) (bi_persistently Q2).
rewrite /IntoAnd /=. destruct p; simpl.
- by rewrite -persistently_and !persistently_idemp.
- intros ->. by rewrite persistently_and.
Global Instance into_and_embed `{BiEmbed PROP PROP'} p P Q1 Q2 :
IntoAnd p P Q1 Q2 → IntoAnd p ⎡P⎤ ⎡Q1⎤ ⎡Q2⎤.
Proof.
rewrite /IntoAnd -embed_and -!embed_persistently_if
-!embed_affinely_if=> -> //.
Qed.
Global Instance into_sep_sep P Q : IntoSep (P ∗ Q) P Q.
Proof. by rewrite /IntoSep. Qed.
Inductive AndIntoSep : PROP → PROP → PROP → PROP → Prop :=
| and_into_sep_affine P Q Q' : Affine P → FromAffinely Q' Q → AndIntoSep P P Q Q'
| and_into_sep P Q : AndIntoSep P (bi_affinely P)%I Q Q.
Existing Class AndIntoSep.
Global Existing Instance and_into_sep_affine | 0.
Global Existing Instance and_into_sep | 2.
Global Instance into_sep_and_persistent_l P P' Q Q' :
Persistent P → AndIntoSep P P' Q Q' → IntoSep (P ∧ Q) P' Q'.
destruct 2 as [P Q Q'|P Q]; rewrite /IntoSep.
- rewrite -(from_affinely Q') -(affine_affinely P) affinely_and_lr.
by rewrite persistent_and_affinely_sep_l_1.
- by rewrite persistent_and_affinely_sep_l_1.
Global Instance into_sep_and_persistent_r P P' Q Q' :
Persistent Q → AndIntoSep Q Q' P P' → IntoSep (P ∧ Q) P' Q'.
Proof.
destruct 2 as [Q P P'|Q P]; rewrite /IntoSep.
- rewrite -(from_affinely P') -(affine_affinely Q) -affinely_and_lr.
by rewrite persistent_and_affinely_sep_r_1.
- by rewrite persistent_and_affinely_sep_r_1.
Qed.
Global Instance into_sep_pure φ ψ : @IntoSep PROP ⌜φ ∧ ψ⌝ ⌜φ⌝ ⌜ψ⌝.
Proof. by rewrite /IntoSep pure_and persistent_and_sep_1. Qed.
Global Instance into_sep_embed `{BiEmbed PROP PROP'} P Q1 Q2 :
IntoSep P Q1 Q2 → IntoSep ⎡P⎤ ⎡Q1⎤ ⎡Q2⎤.
Proof. rewrite /IntoSep -embed_sep=> -> //. Qed.
Global Instance into_sep_affinely `{BiPositive PROP} P Q1 Q2 :
IntoSep P Q1 Q2 → IntoSep (bi_affinely P) (bi_affinely Q1) (bi_affinely Q2) | 0.
Proof. rewrite /IntoSep /= => ->. by rewrite affinely_sep. Qed.
(* FIXME: This instance is kind of strange, it just gets rid of the bi_affinely.
Also, it overlaps with `into_sep_affinely_later`, and hence has lower
precedence. *)
Global Instance into_sep_affinely_trim P Q1 Q2 :
IntoSep P Q1 Q2 → IntoSep (bi_affinely P) Q1 Q2 | 20.
Proof. rewrite /IntoSep /= => ->. by rewrite affinely_elim. Qed.
Global Instance into_sep_persistently `{BiPositive PROP} P Q1 Q2 :
IntoSep P Q1 Q2 →
IntoSep (bi_persistently P) (bi_persistently Q1) (bi_persistently Q2).
Proof. rewrite /IntoSep /= => ->. by rewrite persistently_sep. Qed.
Global Instance into_sep_persistently_affine P Q1 Q2 :
IntoSep P Q1 Q2 →
TCOr (Affine Q1) (Absorbing Q2) → TCOr (Absorbing Q1) (Affine Q2) →
IntoSep (bi_persistently P) (bi_persistently Q1) (bi_persistently Q2).
Proof.
by rewrite sep_and persistently_and persistently_and_sep_l_1.
Qed.
Global Instance into_sep_affinely_persistently_affine P Q1 Q2 :
IntoSep P Q1 Q2 →
TCOr (Affine Q1) (Absorbing Q2) → TCOr (Absorbing Q1) (Affine Q2) →
IntoSep (□ P) (□ Q1) (□ Q2).
Proof.
by rewrite sep_and affinely_persistently_and and_sep_affinely_persistently.
Qed.
Robbert Krebbers
committed
(* We use [IsCons] and [IsApp] to make sure that [frame_big_sepL_cons] and
[frame_big_sepL_app] cannot be applied repeatedly often when having
[ [∗ list] k ↦ x ∈ ?e, Φ k x] with [?e] an evar. *)
Global Instance into_sep_big_sepL_cons {A} (Φ : nat → A → PROP) l x l' :
IntoSep ([∗ list] k ↦ y ∈ l, Φ k y)
(Φ 0 x) ([∗ list] k ↦ y ∈ l', Φ (S k) y).
Proof. rewrite /IsCons=>->. by rewrite /IntoSep big_sepL_cons. Qed.
Global Instance into_sep_big_sepL_app {A} (Φ : nat → A → PROP) l l1 l2 :
IntoSep ([∗ list] k ↦ y ∈ l, Φ k y)
([∗ list] k ↦ y ∈ l1, Φ k y) ([∗ list] k ↦ y ∈ l2, Φ (length l1 + k) y).
Proof. rewrite /IsApp=>->. by rewrite /IntoSep big_sepL_app. Qed.
(* FromOr *)
Global Instance from_or_or P1 P2 : FromOr (P1 ∨ P2) P1 P2.
Proof. by rewrite /FromOr. Qed.
Global Instance from_or_pure φ ψ : @FromOr PROP ⌜φ ∨ ψ⌝ ⌜φ⌝ ⌜ψ⌝.
Proof. by rewrite /FromOr pure_or. Qed.
Global Instance from_or_affinely P Q1 Q2 :
FromOr P Q1 Q2 → FromOr (bi_affinely P) (bi_affinely Q1) (bi_affinely Q2).
Proof. rewrite /FromOr=> <-. by rewrite affinely_or. Qed.
Global Instance from_or_absorbingly P Q1 Q2 :
FromOr P Q1 Q2 → FromOr (bi_absorbingly P) (bi_absorbingly Q1) (bi_absorbingly Q2).
Proof. rewrite /FromOr=> <-. by rewrite absorbingly_or. Qed.
FromOr P Q1 Q2 →
FromOr (bi_persistently P) (bi_persistently Q1) (bi_persistently Q2).
Proof. rewrite /FromOr=> <-. by rewrite persistently_or. Qed.
Global Instance from_or_embed `{BiEmbed PROP PROP'} P Q1 Q2 :
FromOr P Q1 Q2 → FromOr ⎡P⎤ ⎡Q1⎤ ⎡Q2⎤.
Proof. by rewrite /FromOr -embed_or => <-. Qed.
(* IntoOr *)
Global Instance into_or_or P Q : IntoOr (P ∨ Q) P Q.
Proof. by rewrite /IntoOr. Qed.
Global Instance into_or_pure φ ψ : @IntoOr PROP ⌜φ ∨ ψ⌝ ⌜φ⌝ ⌜ψ⌝.
Proof. by rewrite /IntoOr pure_or. Qed.
Global Instance into_or_affinely P Q1 Q2 :
IntoOr P Q1 Q2 → IntoOr (bi_affinely P) (bi_affinely Q1) (bi_affinely Q2).
Proof. rewrite /IntoOr=>->. by rewrite affinely_or. Qed.
Global Instance into_or_absorbingly P Q1 Q2 :
IntoOr P Q1 Q2 → IntoOr (bi_absorbingly P) (bi_absorbingly Q1) (bi_absorbingly Q2).
Proof. rewrite /IntoOr=>->. by rewrite absorbingly_or. Qed.
IntoOr P Q1 Q2 →
IntoOr (bi_persistently P) (bi_persistently Q1) (bi_persistently Q2).
Proof. rewrite /IntoOr=>->. by rewrite persistently_or. Qed.
Global Instance into_or_embed `{BiEmbed PROP PROP'} P Q1 Q2 :
IntoOr P Q1 Q2 → IntoOr ⎡P⎤ ⎡Q1⎤ ⎡Q2⎤.
Proof. by rewrite /IntoOr -embed_or => <-. Qed.
(* FromExist *)
Global Instance from_exist_exist {A} (Φ : A → PROP): FromExist (∃ a, Φ a) Φ.
Proof. by rewrite /FromExist. Qed.
Global Instance from_exist_pure {A} (φ : A → Prop) :
@FromExist PROP A ⌜∃ x, φ x⌝ (λ a, ⌜φ a⌝)%I.
Proof. by rewrite /FromExist pure_exist. Qed.
Global Instance from_exist_affinely {A} P (Φ : A → PROP) :
FromExist P Φ → FromExist (bi_affinely P) (λ a, bi_affinely (Φ a))%I.
Proof. rewrite /FromExist=> <-. by rewrite affinely_exist. Qed.
Global Instance from_exist_absorbingly {A} P (Φ : A → PROP) :
FromExist P Φ → FromExist (bi_absorbingly P) (λ a, bi_absorbingly (Φ a))%I.
Proof. rewrite /FromExist=> <-. by rewrite absorbingly_exist. Qed.
Global Instance from_exist_persistently {A} P (Φ : A → PROP) :
FromExist P Φ → FromExist (bi_persistently P) (λ a, bi_persistently (Φ a))%I.
Proof. rewrite /FromExist=> <-. by rewrite persistently_exist. Qed.
Global Instance from_exist_embed `{BiEmbed PROP PROP'} {A} P (Φ : A → PROP) :
FromExist P Φ → FromExist ⎡P⎤ (λ a, ⎡Φ a⎤%I).
Proof. by rewrite /FromExist -embed_exist => <-. Qed.
(* IntoExist *)
Global Instance into_exist_exist {A} (Φ : A → PROP) : IntoExist (∃ a, Φ a) Φ.
Proof. by rewrite /IntoExist. Qed.
Global Instance into_exist_pure {A} (φ : A → Prop) :
@IntoExist PROP A ⌜∃ x, φ x⌝ (λ a, ⌜φ a⌝)%I.
Proof. by rewrite /IntoExist pure_exist. Qed.
Global Instance into_exist_affinely {A} P (Φ : A → PROP) :
IntoExist P Φ → IntoExist (bi_affinely P) (λ a, bi_affinely (Φ a))%I.
Proof. rewrite /IntoExist=> HP. by rewrite HP affinely_exist. Qed.
Global Instance into_exist_and_pure P Q φ :
IntoPureT P φ → IntoExist (P ∧ Q) (λ _ : φ, Q).
Proof.
intros (φ'&->&?). rewrite /IntoExist (into_pure P).
apply pure_elim_l=> Hφ. by rewrite -(exist_intro Hφ).
Qed.
Global Instance into_exist_sep_pure P Q φ :
TCOr (Affine P) (Absorbing Q) → IntoPureT P φ → IntoExist (P ∗ Q) (λ _ : φ, Q).
Proof.
intros ? (φ'&->&?). rewrite /IntoExist.
eapply (pure_elim φ'); [by rewrite (into_pure P); apply sep_elim_l, _|]=>?.
rewrite -exist_intro //. apply sep_elim_r, _.
Qed.
Global Instance into_exist_absorbingly {A} P (Φ : A → PROP) :
IntoExist P Φ → IntoExist (bi_absorbingly P) (λ a, bi_absorbingly (Φ a))%I.
Proof. rewrite /IntoExist=> HP. by rewrite HP absorbingly_exist. Qed.
Global Instance into_exist_persistently {A} P (Φ : A → PROP) :
IntoExist P Φ → IntoExist (bi_persistently P) (λ a, bi_persistently (Φ a))%I.
Proof. rewrite /IntoExist=> HP. by rewrite HP persistently_exist. Qed.
Global Instance into_exist_embed `{BiEmbed PROP PROP'} {A} P (Φ : A → PROP) :
IntoExist P Φ → IntoExist ⎡P⎤ (λ a, ⎡Φ a⎤%I).
Proof. by rewrite /IntoExist -embed_exist => <-. Qed.
(* IntoForall *)
Global Instance into_forall_forall {A} (Φ : A → PROP) : IntoForall (∀ a, Φ a) Φ.
Proof. by rewrite /IntoForall. Qed.
Global Instance into_forall_affinely {A} P (Φ : A → PROP) :
IntoForall P Φ → IntoForall (bi_affinely P) (λ a, bi_affinely (Φ a))%I.
Proof. rewrite /IntoForall=> HP. by rewrite HP affinely_forall. Qed.
Global Instance into_forall_persistently {A} P (Φ : A → PROP) :
IntoForall P Φ → IntoForall (bi_persistently P) (λ a, bi_persistently (Φ a))%I.
Proof. rewrite /IntoForall=> HP. by rewrite HP persistently_forall. Qed.
Global Instance into_forall_embed `{BiEmbed PROP PROP'} {A} P (Φ : A → PROP) :
IntoForall P Φ → IntoForall ⎡P⎤ (λ a, ⎡Φ a⎤%I).
Proof. by rewrite /IntoForall -embed_forall => <-. Qed.
(* FromForall *)
Global Instance from_forall_forall {A} (Φ : A → PROP) :
FromForall (∀ x, Φ x)%I Φ.
Proof. by rewrite /FromForall. Qed.
Global Instance from_forall_pure {A} (φ : A → Prop) :
@FromForall PROP A (⌜∀ a : A, φ a⌝)%I (λ a, ⌜ φ a ⌝)%I.
Proof. by rewrite /FromForall pure_forall. Qed.
Global Instance from_forall_pure_not (φ : Prop) :
@FromForall PROP φ (⌜¬ φ⌝)%I (λ a : φ, False)%I.
Proof. by rewrite /FromForall pure_forall. Qed.
Global Instance from_forall_impl_pure P Q φ :
IntoPureT P φ → FromForall (P → Q)%I (λ _ : φ, Q)%I.
Proof.
intros (φ'&->&?). by rewrite /FromForall -pure_impl_forall (into_pure P).
Qed.
Global Instance from_forall_wand_pure P Q φ :
TCOr (Affine P) (Absorbing Q) → IntoPureT P φ →
FromForall (P -∗ Q)%I (λ _ : φ, Q)%I.
Proof.
intros [|] (φ'&->&?); rewrite /FromForall; apply wand_intro_r.
- rewrite -(affine_affinely P) (into_pure P) -persistent_and_affinely_sep_r.
apply pure_elim_r=>?. by rewrite forall_elim.
- by rewrite (into_pure P) -pure_wand_forall wand_elim_l.
Qed.
Global Instance from_forall_affinely `{BiAffine PROP} {A} P (Φ : A → PROP) :
FromForall P Φ → FromForall (bi_affinely P)%I (λ a, bi_affinely (Φ a))%I.
Proof.
rewrite /FromForall=> <-. rewrite affine_affinely. by setoid_rewrite affinely_elim.
Qed.
Global Instance from_forall_persistently {A} P (Φ : A → PROP) :
FromForall P Φ → FromForall (bi_persistently P)%I (λ a, bi_persistently (Φ a))%I.
Proof. rewrite /FromForall=> <-. by rewrite persistently_forall. Qed.
Global Instance from_forall_embed `{BiEmbed PROP PROP'} {A} P (Φ : A → PROP) :
FromForall P Φ → FromForall ⎡P⎤%I (λ a, ⎡Φ a⎤%I).
Proof. by rewrite /FromForall -embed_forall => <-. Qed.
(* IntoInv *)
Global Instance into_inv_embed {PROP' : bi} `{BiEmbed PROP PROP'} P N :
IntoInv P N → IntoInv ⎡P⎤ N.
Global Instance elim_modal_wand φ P P' Q Q' R :
ElimModal φ P P' Q Q' → ElimModal φ P P' (R -∗ Q) (R -∗ Q').
rewrite /ElimModal=> H Hφ. apply wand_intro_r.
rewrite wand_curry -assoc (comm _ P') -wand_curry wand_elim_l; auto.
Global Instance elim_modal_forall {A} φ P P' (Φ Ψ : A → PROP) :
(∀ x, ElimModal φ P P' (Φ x) (Ψ x)) → ElimModal φ P P' (∀ x, Φ x) (∀ x, Ψ x).
rewrite /ElimModal=> H ?. apply forall_intro=> a. rewrite (forall_elim a); auto.
Global Instance elim_modal_absorbingly_here P Q :
Absorbing Q → ElimModal True (bi_absorbingly P) P Q Q.
by rewrite absorbingly_sep_l wand_elim_r absorbing_absorbingly.
(* AddModal *)
Global Instance add_modal_wand P P' Q R :
AddModal P P' Q → AddModal P P' (R -∗ Q).
Proof.
rewrite /AddModal=> H. apply wand_intro_r.
by rewrite wand_curry -assoc (comm _ P') -wand_curry wand_elim_l.
Qed.
Global Instance add_modal_forall {A} P P' (Φ : A → PROP) :
(∀ x, AddModal P P' (Φ x)) → AddModal P P' (∀ x, Φ x).
Proof.
rewrite /AddModal=> H. apply forall_intro=> a. by rewrite (forall_elim a).
Qed.
Global Instance frame_here_absorbing p R : Absorbing R → Frame p R R True | 0.
Proof. intros. by rewrite /Frame affinely_persistently_if_elim sep_elim_l. Qed.
Global Instance frame_here p R : Frame p R R emp | 1.
Proof. intros. by rewrite /Frame affinely_persistently_if_elim sep_elim_l. Qed.
Global Instance frame_affinely_here_absorbing p R :
Absorbing R → Frame p (bi_affinely R) R True | 0.
intros. rewrite /Frame affinely_persistently_if_elim affinely_elim.
apply sep_elim_l, _.
Global Instance frame_affinely_here p R : Frame p (bi_affinely R) R emp | 1.
intros. rewrite /Frame affinely_persistently_if_elim affinely_elim.
apply sep_elim_l, _.
Global Instance frame_here_pure p φ Q : FromPure false Q φ → Frame p ⌜φ⌝ Q True.
rewrite /FromPure /Frame=> <-.
by rewrite affinely_persistently_if_elim sep_elim_l.
Global Instance make_embed_pure `{BiEmbed PROP PROP'} φ :
Jacques-Henri Jourdan
committed
KnownMakeEmbed ⌜φ⌝ ⌜φ⌝.
Proof. apply embed_pure. Qed.
Global Instance make_embed_emp `{BiEmbed PROP PROP'} :
Jacques-Henri Jourdan
committed
KnownMakeEmbed emp emp.
Proof. apply embed_emp. Qed.
Global Instance make_embed_default `{BiEmbed PROP PROP'} P :
MakeEmbed P ⎡P⎤ | 100.
Proof. by rewrite /MakeEmbed. Qed.
Global Instance frame_embed `{BiEmbed PROP PROP'} p P Q (Q' : PROP') R :
Frame p R P Q → MakeEmbed Q Q' → Frame p ⎡R⎤ ⎡P⎤ Q'.
Proof.
rewrite /Frame /MakeEmbed => <- <-.
rewrite embed_sep embed_affinely_if embed_persistently_if => //.
Qed.
Global Instance frame_pure_embed `{BiEmbed PROP PROP'} p P Q (Q' : PROP') φ :
Frame p ⌜φ⌝ P Q → MakeEmbed Q Q' → Frame p ⌜φ⌝ ⎡P⎤ Q'.
Proof. rewrite /Frame /MakeEmbed -embed_pure. apply (frame_embed p P Q). Qed.
Jacques-Henri Jourdan
committed
Global Instance make_sep_emp_l P : KnownLMakeSep emp P P.
Proof. apply left_id, _. Qed.
Global Instance make_sep_emp_r P : KnownRMakeSep P emp P.
Proof. apply right_id, _. Qed.
Global Instance make_sep_true_l P : Absorbing P → KnownLMakeSep True P P.
Proof. intros. apply True_sep, _. Qed.
Global Instance make_and_emp_l_absorbingly P : KnownLMakeSep True P (bi_absorbingly P) | 10.
Proof. intros. by rewrite /KnownLMakeSep /MakeSep. Qed.
Global Instance make_sep_true_r P : Absorbing P → KnownRMakeSep P True P.
Proof. intros. by rewrite /KnownRMakeSep /MakeSep sep_True. Qed.
Global Instance make_and_emp_r_absorbingly P : KnownRMakeSep P True (bi_absorbingly P) | 10.
Proof. intros. by rewrite /KnownRMakeSep /MakeSep comm. Qed.
Global Instance make_sep_default P Q : MakeSep P Q (P ∗ Q) | 100.
Global Instance frame_sep_persistent_l progress R P1 P2 Q1 Q2 Q' :
Frame true R P1 Q1 → MaybeFrame true R P2 Q2 progress → MakeSep Q1 Q2 Q' →
Frame true R (P1 ∗ P2) Q' | 9.
rewrite /Frame /MaybeFrame /MakeSep /= => <- <- <-.
rewrite {1}(affinely_persistently_sep_dup R). solve_sep_entails.
Global Instance frame_sep_l R P1 P2 Q Q' :
Frame false R P1 Q → MakeSep Q P2 Q' → Frame false R (P1 ∗ P2) Q' | 9.
Proof. rewrite /Frame /MakeSep => <- <-. by rewrite assoc. Qed.
Global Instance frame_sep_r p R P1 P2 Q Q' :
Frame p R P2 Q → MakeSep P1 Q Q' → Frame p R (P1 ∗ P2) Q' | 10.
rewrite /Frame /MakeSep => <- <-. by rewrite assoc -(comm _ P1) assoc.
Global Instance frame_big_sepL_cons {A} p (Φ : nat → A → PROP) R Q l x l' :
IsCons l x l' →
Frame p R (Φ 0 x ∗ [∗ list] k ↦ y ∈ l', Φ (S k) y) Q →
Frame p R ([∗ list] k ↦ y ∈ l, Φ k y) Q.
Proof. rewrite /IsCons=>->. by rewrite /Frame big_sepL_cons. Qed.
Global Instance frame_big_sepL_app {A} p (Φ : nat → A → PROP) R Q l l1 l2 :
Frame p R (([∗ list] k ↦ y ∈ l1, Φ k y) ∗
[∗ list] k ↦ y ∈ l2, Φ (length l1 + k) y) Q →
Frame p R ([∗ list] k ↦ y ∈ l, Φ k y) Q.
Proof. rewrite /IsApp=>->. by rewrite /Frame big_opL_app. Qed.
Jacques-Henri Jourdan
committed
Global Instance make_and_true_l P : KnownLMakeAnd True P P.
Proof. apply left_id, _. Qed.
Global Instance make_and_true_r P : KnownRMakeAnd P True P.
Proof. by rewrite /KnownRMakeAnd /MakeAnd right_id. Qed.
Global Instance make_and_emp_l P : Affine P → KnownLMakeAnd emp P P.
Proof. intros. by rewrite /KnownLMakeAnd /MakeAnd emp_and. Qed.
Global Instance make_and_emp_l_affinely P : KnownLMakeAnd emp P (bi_affinely P) | 10.
Proof. intros. by rewrite /KnownLMakeAnd /MakeAnd. Qed.
Global Instance make_and_emp_r P : Affine P → KnownRMakeAnd P emp P.
Proof. intros. by rewrite /KnownRMakeAnd /MakeAnd and_emp. Qed.
Global Instance make_and_emp_r_affinely P : KnownRMakeAnd P emp (bi_affinely P) | 10.
Proof. intros. by rewrite /KnownRMakeAnd /MakeAnd comm. Qed.
Global Instance make_and_default P Q : MakeAnd P Q (P ∧ Q) | 100.
Global Instance frame_and p progress1 progress2 R P1 P2 Q1 Q2 Q' :
MaybeFrame p R P1 Q1 progress1 →
MaybeFrame p R P2 Q2 progress2 →
TCEq (progress1 || progress2) true →
MakeAnd Q1 Q2 Q' →
Frame p R (P1 ∧ P2) Q' | 9.
rewrite /MaybeFrame /Frame /MakeAnd => <- <- _ <-. apply and_intro;
[rewrite and_elim_l|rewrite and_elim_r]; done.
Jacques-Henri Jourdan
committed
Global Instance make_or_true_l P : KnownLMakeOr True P True.
Proof. apply left_absorb, _. Qed.
Global Instance make_or_true_r P : KnownRMakeOr P True True.
Proof. by rewrite /KnownRMakeOr /MakeOr right_absorb. Qed.
Global Instance make_or_emp_l P : Affine P → KnownLMakeOr emp P emp.
Proof. intros. by rewrite /KnownLMakeOr /MakeOr emp_or. Qed.
Global Instance make_or_emp_r P : Affine P → KnownRMakeOr P emp emp.
Proof. intros. by rewrite /KnownRMakeOr /MakeOr or_emp. Qed.
Global Instance make_or_default P Q : MakeOr P Q (P ∨ Q) | 100.
(* We could in principle write the instance [frame_or_spatial] by a bunch of
instances, i.e. (omitting the parameter [p = false]):
Frame R P1 Q1 → Frame R P2 Q2 → Frame R (P1 ∨ P2) (Q1 ∨ Q2)
Frame R P1 True → Frame R (P1 ∨ P2) P2
Frame R P2 True → Frame R (P1 ∨ P2) P1
The problem here is that Coq will try to infer [Frame R P1 ?] and [Frame R P2 ?]
multiple times, whereas the current solution makes sure that said inference
appears at most once.
If Coq would memorize the results of type class resolution, the solution with
multiple instances would be preferred (and more Prolog-like). *)
Global Instance frame_or_spatial progress1 progress2 R P1 P2 Q1 Q2 Q :
MaybeFrame false R P1 Q1 progress1 → MaybeFrame false R P2 Q2 progress2 →
TCOr (TCEq (progress1 && progress2) true) (TCOr
(TCAnd (TCEq progress1 true) (TCEq Q1 True%I))
(TCAnd (TCEq progress2 true) (TCEq Q2 True%I))) →
MakeOr Q1 Q2 Q →
Frame false R (P1 ∨ P2) Q | 9.
Proof. rewrite /Frame /MakeOr => <- <- _ <-. by rewrite -sep_or_l. Qed.
Global Instance frame_or_persistent progress1 progress2 R P1 P2 Q1 Q2 Q :
MaybeFrame true R P1 Q1 progress1 → MaybeFrame true R P2 Q2 progress2 →
TCEq (progress1 || progress2) true →
MakeOr Q1 Q2 Q → Frame true R (P1 ∨ P2) Q | 9.
Proof. rewrite /Frame /MakeOr => <- <- _ <-. by rewrite -sep_or_l. Qed.
Global Instance frame_wand p R P1 P2 Q2 :
Frame p R P2 Q2 → Frame p R (P1 -∗ P2) (P1 -∗ Q2).
rewrite /Frame=> ?. apply wand_intro_l.
by rewrite assoc (comm _ P1) -assoc wand_elim_r.
Qed.
Jacques-Henri Jourdan
committed
Global Instance make_affinely_True : @KnownMakeAffinely PROP True emp | 0.
Proof. by rewrite /KnownMakeAffinely /MakeAffinely affinely_True_emp affinely_emp. Qed.
Global Instance make_affinely_affine P : Affine P → KnownMakeAffinely P P | 1.
Proof. intros. by rewrite /KnownMakeAffinely /MakeAffinely affine_affinely. Qed.
Global Instance make_affinely_default P : MakeAffinely P (bi_affinely P) | 100.
Proof. by rewrite /MakeAffinely. Qed.
Global Instance frame_affinely R P Q Q' :
Frame true R P Q → MakeAffinely Q Q' → Frame true R (bi_affinely P) Q'.
rewrite /Frame /MakeAffinely=> <- <- /=.
by rewrite -{1}affinely_idemp affinely_sep_2.
Jacques-Henri Jourdan
committed
Global Instance make_absorbingly_emp : @KnownMakeAbsorbingly PROP emp True | 0.
Proof.
by rewrite /KnownMakeAbsorbingly /MakeAbsorbingly
-absorbingly_True_emp absorbingly_pure.
Qed.
(* Note: there is no point in having an instance `Absorbing P → MakeAbsorbingly P P`
because framing will never turn a proposition that is not absorbing into
something that is absorbing. *)
Global Instance make_absorbingly_default P : MakeAbsorbingly P (bi_absorbingly P) | 100.
Proof. by rewrite /MakeAbsorbingly. Qed.
Global Instance frame_absorbingly p R P Q Q' :
Frame p R P Q → MakeAbsorbingly Q Q' → Frame p R (bi_absorbingly P) Q'.
rewrite /Frame /MakeAbsorbingly=> <- <- /=. by rewrite absorbingly_sep_r.
Jacques-Henri Jourdan
committed
Global Instance make_persistently_true : @KnownMakePersistently PROP True True.
Proof. by rewrite /KnownMakePersistently /MakePersistently persistently_pure. Qed.
Global Instance make_persistently_emp : @KnownMakePersistently PROP emp True.
Proof.
by rewrite /KnownMakePersistently /MakePersistently
-persistently_True_emp persistently_pure.
Qed.
Global Instance make_persistently_default P :
MakePersistently P (bi_persistently P) | 100.
Global Instance frame_persistently R P Q Q' :
Frame true R P Q → MakePersistently Q Q' → Frame true R (bi_persistently P) Q'.
rewrite /Frame /MakePersistently=> <- <- /=.
rewrite -persistently_and_affinely_sep_l.
by rewrite -persistently_sep_2 -persistently_and_sep_l_1 persistently_affinely
persistently_idemp.
Global Instance frame_exist {A} p R (Φ Ψ : A → PROP) :
(∀ a, Frame p R (Φ a) (Ψ a)) → Frame p R (∃ x, Φ x) (∃ x, Ψ x).
Proof. rewrite /Frame=> ?. by rewrite sep_exist_l; apply exist_mono. Qed.
Global Instance frame_forall {A} p R (Φ Ψ : A → PROP) :
(∀ a, Frame p R (Φ a) (Ψ a)) → Frame p R (∀ x, Φ x) (∀ x, Ψ x).
Proof. rewrite /Frame=> ?. by rewrite sep_forall_l; apply forall_mono. Qed.
Global Instance frame_impl_persistent R P1 P2 Q2 :
Frame true R P2 Q2 → Frame true R (P1 → P2) (P1 → Q2).
rewrite /Frame /= => ?. apply impl_intro_l.
by rewrite -persistently_and_affinely_sep_l assoc (comm _ P1) -assoc impl_elim_r
persistently_and_affinely_sep_l.
Qed.
Global Instance frame_impl R P1 P2 Q2 :
Persistent P1 → Absorbing P1 →
Frame false R P2 Q2 → Frame false R (P1 → P2) (P1 → Q2).
rewrite /Frame /==> ???. apply impl_intro_l.
rewrite {1}(persistent P1) persistently_and_affinely_sep_l assoc.
rewrite (comm _ (□ P1)%I) -assoc -persistently_and_affinely_sep_l.
rewrite persistently_elim impl_elim_r //.
(* IntoEmbed *)
Global Instance into_embed_embed {PROP' : bi} `{BiEmbed PROP PROP'} P :
IntoEmbed ⎡P⎤ P.
Proof. by rewrite /IntoEmbed. Qed.
(* AsValid *)
Global Instance as_valid_valid {PROP : bi} (P : PROP) : AsValid0 (bi_valid P) P | 0.
Proof. by rewrite /AsValid. Qed.
Global Instance as_valid_entails {PROP : bi} (P Q : PROP) : AsValid0 (P ⊢ Q) (P -∗ Q).
Proof. split. apply bi.entails_wand. apply bi.wand_entails. Qed.
Global Instance as_valid_equiv {PROP : bi} (P Q : PROP) : AsValid0 (P ≡ Q) (P ∗-∗ Q).
Proof. split. apply bi.equiv_wand_iff. apply bi.wand_iff_equiv. Qed.
Global Instance as_valid_forall {A : Type} (φ : A → Prop) (P : A → PROP) :
(∀ x, AsValid (φ x) (P x)) → AsValid (∀ x, φ x) (∀ x, P x).
Proof.
rewrite /AsValid=>H1. split=>H2.
- apply bi.forall_intro=>?. apply H1, H2.
- intros x. apply H1. revert H2. by rewrite (bi.forall_elim x).
Qed.
Global Instance as_valid_embed `{BiEmbed PROP PROP'} (φ : Prop) (P : PROP) :
AsValid0 φ P → AsValid φ ⎡P⎤.
Proof. rewrite /AsValid0 /AsValid=> ->. rewrite embed_valid //. Qed.
End bi_instances.
Section sbi_instances.
Context {PROP : sbi}.
Implicit Types P Q R : PROP.
(* FromAssumption *)
Global Instance from_assumption_later p P Q :
FromAssumption p P Q → FromAssumption p P (▷ Q)%I.
Proof. rewrite /FromAssumption=>->. apply later_intro. Qed.
Global Instance from_assumption_laterN n p P Q :
FromAssumption p P Q → FromAssumption p P (▷^n Q)%I.
Proof. rewrite /FromAssumption=>->. apply laterN_intro. Qed.
Global Instance from_assumption_except_0 p P Q :
FromAssumption p P Q → FromAssumption p P (◇ Q)%I.
Proof. rewrite /FromAssumption=>->. apply except_0_intro. Qed.
Global Instance from_assumption_bupd `{BiBUpd PROP} p P Q :
FromAssumption p P Q → FromAssumption p P (|==> Q).
Proof. rewrite /FromAssumption=>->. apply bupd_intro. Qed.
Global Instance from_assumption_fupd `{BiBUpdFUpd PROP} E p P Q :
FromAssumption p P (|==> Q) → FromAssumption p P (|={E}=> Q)%I.
Proof. rewrite /FromAssumption=>->. apply bupd_fupd. Qed.
Global Instance from_assumption_plainly_l_true `{BiPlainly PROP} P Q :
FromAssumption true P Q → FromAssumption true (■ P) Q.