Newer
Older
From iris.proofmode Require Export classes.
Set Default Proof Using "Type".
Import bi.
Section bi_instances.
Context {PROP : bi}.
Implicit Types P Q R : PROP.
(* IntoInternalEq *)
Global Instance into_internal_eq_internal_eq {A : ofeT} (x y : A) :
@IntoInternalEq PROP A (x ≡ y) x y.
Proof. by rewrite /IntoInternalEq. Qed.
Global Instance into_internal_eq_bare {A : ofeT} (x y : A) P :
IntoInternalEq P x y → IntoInternalEq (■ P) x y.
Proof. rewrite /IntoInternalEq=> ->. by rewrite bare_elim. Qed.
Global Instance into_internal_eq_sink {A : ofeT} (x y : A) P :
IntoInternalEq P x y → IntoInternalEq (▲ P) x y.
Proof. rewrite /IntoInternalEq=> ->. by rewrite sink_internal_eq. Qed.
Global Instance into_internal_eq_persistently {A : ofeT} (x y : A) P :
IntoInternalEq P x y → IntoInternalEq (□ P) x y.
Proof. rewrite /IntoInternalEq=> ->. by rewrite persistently_elim. Qed.
(* FromBare *)
Global Instance from_bare_affine P : Affine P → FromBare P P.
Proof. intros. by rewrite /FromBare bare_elim. Qed.
Global Instance from_bare_default P : FromBare (■ P) P | 100.
Proof. by rewrite /FromBare. Qed.
(* FromAssumption *)
Global Instance from_assumption_exact p P : FromAssumption p P P | 0.
Proof. by rewrite /FromAssumption /= bare_persistently_if_elim. Qed.
Global Instance from_assumption_persistently_r P Q :
FromAssumption true P Q → FromAssumption true P (□ Q).
Proof.
rewrite /FromAssumption /= =><-.
by rewrite -{1}bare_persistently_idemp bare_elim.
Qed.
Global Instance from_assumption_bare_r P Q :
FromAssumption true P Q → FromAssumption true P (■ Q).
Proof. rewrite /FromAssumption /= =><-. by rewrite bare_idemp. Qed.
Global Instance from_assumption_sink_r p P Q :
FromAssumption p P Q → FromAssumption p P (▲ Q).
Proof. rewrite /FromAssumption /= =><-. apply sink_intro. Qed.
Global Instance from_assumption_bare_persistently_l p P Q :
FromAssumption true P Q → FromAssumption p (⬕ P) Q.
Proof. rewrite /FromAssumption /= =><-. by rewrite bare_persistently_if_elim. Qed.
Global Instance from_assumption_persistently_l_true P Q :
FromAssumption true P Q → FromAssumption true (□ P) Q.
Proof. rewrite /FromAssumption /= =><-. by rewrite persistently_idemp. Qed.
Global Instance from_assumption_persistently_l_false `{AffineBI PROP} P Q :
FromAssumption true P Q → FromAssumption false (□ P) Q.
Proof. rewrite /FromAssumption /= =><-. by rewrite affine_bare. Qed.
Global Instance from_assumption_bare_l_true p P Q :
FromAssumption p P Q → FromAssumption p (■ P) Q.
Proof. rewrite /FromAssumption /= =><-. by rewrite bare_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.
Global Instance into_pure_eq {A : ofeT} (a b : A) :
Discrete a → @IntoPure M (a ≡ b) (a ≡ b).
Proof. intros. by rewrite /IntoPure discrete_eq. 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 :
Proof. rewrite /FromPure /IntoPure pure_impl. by intros -> ->. 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 :
Proof. rewrite /FromPure /IntoPure=> <- ->. by rewrite pure_impl impl_wand_2. Qed.
Global Instance into_pure_bare P φ : IntoPure P φ → IntoPure (■ P) φ.
Proof. rewrite /IntoPure=> ->. apply bare_elim. Qed.
Global Instance into_pure_sink P φ : IntoPure P φ → IntoPure (▲ P) φ.
Proof. rewrite /IntoPure=> ->. by rewrite sink_pure. Qed.
Global Instance into_pure_persistently P φ : IntoPure P φ → IntoPure (□ P) φ.
Proof. rewrite /IntoPure=> ->. apply: persistently_elim. Qed.
(* FromPure *)
Global Instance from_pure_pure φ : @FromPure PROP ⌜φ⌝ φ.
Proof. by rewrite /FromPure. Qed.
Global Instance from_pure_internal_eq {A : ofeT} (a b : A) :
@FromPure PROP (a ≡ b) (a ≡ b).
Proof. by rewrite /FromPure pure_internal_eq. Qed.
FromPure P1 φ1 → FromPure P2 φ2 → FromPure (P1 ∧ P2) (φ1 ∧ φ2).
Proof. rewrite /FromPure pure_and. by intros -> ->. Qed.
FromPure P1 φ1 → FromPure P2 φ2 → FromPure (P1 ∨ P2) (φ1 ∨ φ2).
Proof. rewrite /FromPure pure_or. by intros -> ->. Qed.
Global Instance from_pure_pure_impl (φ1 φ2 : Prop) P1 P2 :
IntoPure P1 φ1 → FromPure P2 φ2 → FromPure (P1 → P2) (φ1 → φ2).
Proof. rewrite /FromPure /IntoPure pure_impl. by intros -> ->. Qed.
Global Instance from_pure_exist {A} (Φ : A → PROP) (φ : A → Prop) :
(∀ x, FromPure (Φ x) (φ x)) → FromPure (∃ x, Φ x) (∃ x, φ x).
Proof. rewrite /FromPure=>Hx. rewrite pure_exist. by setoid_rewrite Hx. Qed.
Global Instance from_pure_forall {A} (Φ : A → PROP) (φ : A → Prop) :
(∀ x, FromPure (Φ x) (φ x)) → FromPure (∀ x, Φ x) (∀ x, φ x).
Proof. rewrite /FromPure=>Hx. rewrite pure_forall. by setoid_rewrite Hx. Qed.
Global Instance from_pure_pure_sep (φ1 φ2 : Prop) P1 P2 :
FromPure P1 φ1 → FromPure P2 φ2 → FromPure (P1 ∗ P2) (φ1 ∧ φ2).
Proof. rewrite /FromPure=> <- <-. by rewrite pure_and persistent_and_sep_1. Qed.
Global Instance from_pure_pure_wand (φ1 φ2 : Prop) P1 P2 :
IntoPure P1 φ1 → FromPure P2 φ2 → FromPure (P1 -∗ P2) (φ1 → φ2).
rewrite /FromPure /IntoPure=> -> <-.
by rewrite pure_wand_forall pure_impl pure_impl_forall.
Global Instance from_pure_persistently P φ : FromPure P φ → FromPure (□ P) φ.
Proof. rewrite /FromPure=> <-. by rewrite persistently_pure. Qed.
Global Instance from_pure_sink P φ : FromPure P φ → FromPure (▲ P) φ.
Proof. rewrite /FromPure=> <-. by rewrite sink_pure. Qed.
(* IntoPersistent *)
Global Instance into_persistent_persistently p P Q :
IntoPersistent true P Q → IntoPersistent p (□ P) Q | 0.
Proof.
rewrite /IntoPersistent /= => ->.
destruct p; simpl; auto using persistently_idemp_1.
Qed.
Global Instance into_persistent_bare p P Q :
IntoPersistent p P Q → IntoPersistent p (■ P) Q | 0.
Proof. rewrite /IntoPersistent /= => <-. by rewrite bare_elim. 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.
Global Instance from_persistent_here P : FromPersistent false false P P | 1.
Global Instance from_persistent_persistently a P Q :
FromPersistent a false P Q → FromPersistent false true (□ P) Q | 0.
Proof.
rewrite /FromPersistent /= => <-. by destruct a; rewrite /= ?persistently_bare.
Qed.
Global Instance from_persistent_bare a p P Q :
FromPersistent a p P Q → FromPersistent true p (■ P) Q | 0.
Proof. rewrite /FromPersistent /= => <-. destruct a; by rewrite /= ?bare_idemp. Qed.
(* 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 bare_persistently_if_elim.
Global Instance into_wand_impl_false_false `{!AffineBI PROP} P Q P' :
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.
by rewrite -persistently_and_bare_sep_l persistently_elim impl_elim_r.
Qed.
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_bare_sep_l HP -{2}(affine_bare P') -bare_and_lr.
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}(bare_persistently_idemp P) -and_sep_bare_persistently.
by rewrite -bare_persistently_and impl_elim_r bare_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 {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_persistently_true q R P Q :
IntoWand true q R P Q → IntoWand true q (□ R) P Q.
Proof. by rewrite /IntoWand /= persistently_idemp. Qed.
Global Instance into_wand_persistently_false `{!AffineBI PROP} q R P Q :
IntoWand false q R P Q → IntoWand false q (□ R) P Q.
Proof. by rewrite /IntoWand persistently_elim. Qed.
Global Instance into_wand_bare_persistently p q R P Q :
IntoWand p q R P Q → IntoWand p q (⬕ R) P Q.
Proof. by rewrite /IntoWand bare_persistently_elim. Qed.
(* 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 :
FromBare P1 P1' → Persistent P1' → FromAnd (P1 ∗ P2) P1' P2 | 9.
rewrite /FromBare /FromAnd=> <- ?. by rewrite persistent_and_bare_sep_l_1.
Global Instance from_and_sep_persistent_r P1 P2 P2' :
FromBare P2 P2' → Persistent P2' → FromAnd (P1 ∗ P2) P1 P2' | 10.
rewrite /FromBare /FromAnd=> <- ?. by rewrite persistent_and_bare_sep_r_1.
Global Instance from_and_sep_persistent P1 P2 :
Persistent P1 → Persistent P2 → FromAnd (P1 ∗ P2) P1 P2 | 11.
rewrite /FromBare /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 (□ P) (□ Q1) (□ Q2).
Proof. rewrite /FromAnd=> <-. by rewrite persistently_and. Qed.
Global Instance from_and_persistently_sep P Q1 Q2 :
FromSep P Q1 Q2 → FromAnd (□ P) (□ Q1) (□ Q2) | 11.
Proof. rewrite /FromAnd=> <-. by rewrite -persistently_and persistently_and_sep. 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 (TCAnd (Affine P1) (Affine P2)) (TCAnd (Absorbing P1) (Absorbing 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_bare P Q1 Q2 :
FromSep P Q1 Q2 → FromSep (■ P) (■ Q1) (■ Q2).
Proof. rewrite /FromSep=> <-. by rewrite bare_sep_2. Qed.
Global Instance from_sep_sink P Q1 Q2 :
FromSep P Q1 Q2 → FromSep (▲ P) (▲ Q1) (▲ Q2).
Proof. rewrite /FromSep=> <-. by rewrite sink_sep. Qed.
Global Instance from_sep_persistently P Q1 Q2 :
FromSep P Q1 Q2 → FromSep (□ P) (□ Q1) (□ Q2).
Proof. rewrite /FromSep=> <-. by rewrite persistently_sep_2. 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 bare_persistently_if_and. Qed.
Global Instance into_and_and_affine_l P Q Q' :
Affine P → FromBare Q' Q → IntoAnd false (P ∧ Q) P Q'.
Proof.
intros. rewrite /IntoAnd /=.
by rewrite -(affine_bare P) bare_and_l bare_and (from_bare Q').
Qed.
Global Instance into_and_and_affine_r P P' Q :
Affine Q → FromBare P' P → IntoAnd false (P ∧ Q) P' Q.
Proof.
intros. rewrite /IntoAnd /=.
by rewrite -(affine_bare Q) bare_and_r bare_and (from_bare P').
Global Instance into_and_sep `{PositiveBI 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_pure p φ ψ : @IntoAnd PROP p ⌜φ ∧ ψ⌝ ⌜φ⌝ ⌜ψ⌝.
Proof. by rewrite /IntoAnd pure_and bare_persistently_if_and. Qed.
Global Instance into_and_bare p P Q1 Q2 :
IntoAnd p P Q1 Q2 → IntoAnd p (■ P) (■ Q1) (■ Q2).
rewrite /IntoAnd. destruct p; simpl.
- by rewrite -bare_and !persistently_bare.
- intros ->. by rewrite bare_and.
Global Instance into_and_persistently p P Q1 Q2 :
IntoAnd p P Q1 Q2 → IntoAnd p (□ P) (□ Q1) (□ Q2).
rewrite /IntoAnd /=. destruct p; simpl.
- by rewrite -persistently_and !persistently_idemp.
- intros ->. by rewrite persistently_and.
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 → FromBare Q' Q → AndIntoSep P P Q Q'
| and_into_sep P Q : AndIntoSep P (■ 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_bare Q') -(affine_bare P) bare_and_lr.
by rewrite persistent_and_bare_sep_l_1.
- by rewrite persistent_and_bare_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_bare P') -(affine_bare Q) -bare_and_lr.
by rewrite persistent_and_bare_sep_r_1.
- by rewrite persistent_and_bare_sep_r_1.
Qed.
Global Instance into_sep_pure φ ψ : @IntoSep PROP ⌜φ ∧ ψ⌝ ⌜φ⌝ ⌜ψ⌝.
Proof. by rewrite /IntoSep pure_and persistent_and_sep_1. Qed.
(* FIXME: This instance is kind of strange, it just gets rid of the ■. Also, it
overlaps with `into_sep_bare_later`, and hence has lower precedence. *)
Global Instance into_sep_bare P Q1 Q2 :
IntoSep P Q1 Q2 → IntoSep (■ P) Q1 Q2 | 20.
Proof. rewrite /IntoSep /= => ->. by rewrite bare_elim. Qed.
Global Instance into_sep_persistently `{PositiveBI PROP} P Q1 Q2 :
IntoSep P Q1 Q2 → IntoSep (□ P) (□ Q1) (□ Q2).
Proof. rewrite /IntoSep /= => ->. by rewrite persistently_sep. 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_bare P Q1 Q2 :
FromOr P Q1 Q2 → FromOr (■ P) (■ Q1) (■ Q2).
Proof. rewrite /FromOr=> <-. by rewrite bare_or. Qed.
Global Instance from_or_sink P Q1 Q2 :
FromOr P Q1 Q2 → FromOr (▲ P) (▲ Q1) (▲ Q2).
Proof. rewrite /FromOr=> <-. by rewrite sink_or. Qed.
Global Instance from_or_persistently P Q1 Q2 :
FromOr P Q1 Q2 → FromOr (□ P) (□ Q1) (□ Q2).
Proof. rewrite /FromOr=> <-. by rewrite persistently_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_bare P Q1 Q2 :
IntoOr P Q1 Q2 → IntoOr (■ P) (■ Q1) (■ Q2).
Proof. rewrite /IntoOr=>->. by rewrite bare_or. Qed.
Global Instance into_or_sink P Q1 Q2 :
IntoOr P Q1 Q2 → IntoOr (▲ P) (▲ Q1) (▲ Q2).
Proof. rewrite /IntoOr=>->. by rewrite sink_or. Qed.
Global Instance into_or_persistently P Q1 Q2 :
IntoOr P Q1 Q2 → IntoOr (□ P) (□ Q1) (□ Q2).
Proof. rewrite /IntoOr=>->. by rewrite persistently_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_bare {A} P (Φ : A → PROP) :
FromExist P Φ → FromExist (■ P) (λ a, ■ (Φ a))%I.
Proof. rewrite /FromExist=> <-. by rewrite bare_exist. Qed.
Global Instance from_exist_sink {A} P (Φ : A → PROP) :
FromExist P Φ → FromExist (▲ P) (λ a, ▲ (Φ a))%I.
Proof. rewrite /FromExist=> <-. by rewrite sink_exist. Qed.
Global Instance from_exist_persistently {A} P (Φ : A → PROP) :
FromExist P Φ → FromExist (□ P) (λ a, □ (Φ a))%I.
Proof. rewrite /FromExist=> <-. by rewrite persistently_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_bare {A} P (Φ : A → PROP) :
IntoExist P Φ → IntoExist (■ P) (λ a, ■ (Φ a))%I.
Proof. rewrite /IntoExist=> HP. by rewrite HP bare_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_sink {A} P (Φ : A → PROP) :
IntoExist P Φ → IntoExist (▲ P) (λ a, ▲ (Φ a))%I.
Proof. rewrite /IntoExist=> HP. by rewrite HP sink_exist. Qed.
464
465
466
467
468
469
470
471
472
473
474
475
476
477
478
479
480
481
482
483
484
485
486
487
488
489
490
491
492
493
494
495
496
497
498
499
500
501
502
503
504
505
506
507
508
509
510
511
512
513
514
515
Global Instance into_exist_persistently {A} P (Φ : A → PROP) :
IntoExist P Φ → IntoExist (□ P) (λ a, □ (Φ a))%I.
Proof. rewrite /IntoExist=> HP. by rewrite HP persistently_exist. Qed.
(* IntoForall *)
Global Instance into_forall_forall {A} (Φ : A → PROP) : IntoForall (∀ a, Φ a) Φ.
Proof. by rewrite /IntoForall. Qed.
Global Instance into_forall_bare {A} P (Φ : A → PROP) :
IntoForall P Φ → IntoForall (■ P) (λ a, ■ (Φ a))%I.
Proof. rewrite /IntoForall=> HP. by rewrite HP bare_forall. Qed.
Global Instance into_forall_persistently {A} P (Φ : A → PROP) :
IntoForall P Φ → IntoForall (□ P) (λ a, □ (Φ a))%I.
Proof. rewrite /IntoForall=> HP. by rewrite HP persistently_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_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_bare P) (into_pure P) -persistent_and_bare_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_persistently {A} P (Φ : A → PROP) :
FromForall P Φ → FromForall (□ P)%I (λ a, □ (Φ a))%I.
Proof. rewrite /FromForall=> <-. by rewrite persistently_forall. Qed.
(* ElimModal *)
Global Instance elim_modal_wand P P' Q Q' R :
ElimModal P P' Q Q' → ElimModal P P' (R -∗ Q) (R -∗ Q').
Proof.
rewrite /ElimModal=> H. apply wand_intro_r.
by rewrite wand_curry -assoc (comm _ P') -wand_curry wand_elim_l.
Qed.
Global Instance forall_modal_wand {A} P P' (Φ Ψ : A → PROP) :
(∀ x, ElimModal P P' (Φ x) (Ψ x)) → ElimModal P P' (∀ x, Φ x) (∀ x, Ψ x).
Proof.
rewrite /ElimModal=> H. apply forall_intro=> a. by rewrite (forall_elim a).
Qed.
Global Instance elim_modal_sink P Q : Absorbing Q → ElimModal (▲ P) P Q Q.
Proof.
rewrite /ElimModal=> H. by rewrite sink_sep_l wand_elim_r absorbing_sink.
Qed.
(* Frame *)
Global Instance frame_here_absorbing p R : Absorbing R → Frame p R R True | 0.
Proof. intros. by rewrite /Frame bare_persistently_if_elim sep_elim_l. Qed.
Global Instance frame_here p R : Frame p R R emp | 1.
Proof. intros. by rewrite /Frame bare_persistently_if_elim sep_elim_l. Qed.
Global Instance frame_bare_here_absorbing p R : Absorbing R → Frame p (■ R) R True | 0.
Proof. intros. by rewrite /Frame bare_persistently_if_elim bare_elim sep_elim_l. Qed.
Global Instance frame_bare_here p R : Frame p (■ R) R emp | 1.
Proof. intros. by rewrite /Frame bare_persistently_if_elim bare_elim sep_elim_l. Qed.
Global Instance frame_here_pure p φ Q : FromPure Q φ → Frame p ⌜φ⌝ Q True.
Proof.
rewrite /FromPure /Frame=> <-. by rewrite bare_persistently_if_elim sep_elim_l.
Qed.
Class MakeSep (P Q PQ : PROP) := make_sep : P ∗ Q ⊣⊢ PQ.
Arguments MakeSep _%I _%I _%I.
Global Instance make_sep_emp_l P : MakeSep emp P P.
Proof. by rewrite /MakeSep left_id. Qed.
Global Instance make_sep_emp_r P : MakeSep P emp P.
Proof. by rewrite /MakeSep right_id. Qed.
Global Instance make_sep_true_l P : Absorbing P → MakeSep True P P.
Proof. intros. by rewrite /MakeSep True_sep. Qed.
Global Instance make_and_emp_l_sink P : MakeSep True P (▲ P) | 10.
Proof. intros. by rewrite /MakeSep. Qed.
Global Instance make_sep_true_r P : Absorbing P → MakeSep P True P.
Proof. intros. by rewrite /MakeSep sep_True. Qed.
Global Instance make_and_emp_r_sink P : MakeSep P True (▲ P) | 10.
Proof. intros. by rewrite /MakeSep comm. Qed.
Global Instance make_sep_default P Q : MakeSep P Q (P ∗ Q) | 100.
Global Instance frame_sep_persistent_l R P1 P2 Q1 Q2 Q' :
Frame true R P1 Q1 → MaybeFrame true R P2 Q2 → MakeSep Q1 Q2 Q' →
Frame true R (P1 ∗ P2) Q' | 9.
Proof.
rewrite /Frame /MaybeFrame /MakeSep /= => <- <- <-.
rewrite {1}(bare_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.
Proof. rewrite /Frame /MakeSep => <- <-. by rewrite assoc -(comm _ P1) assoc. Qed.
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.
Class MakeAnd (P Q PQ : PROP) := make_and : P ∧ Q ⊣⊢ PQ.
Arguments MakeAnd _%I _%I _%I.
Global Instance make_and_true_l P : MakeAnd True P P.
Proof. by rewrite /MakeAnd left_id. Qed.
Global Instance make_and_true_r P : MakeAnd P True P.
Proof. by rewrite /MakeAnd right_id. Qed.
Global Instance make_and_emp_l P : Affine P → MakeAnd emp P P.
Proof. intros. by rewrite /MakeAnd emp_and. Qed.
Global Instance make_and_emp_l_bare P : MakeAnd emp P (■ P) | 10.
Proof. intros. by rewrite /MakeAnd. Qed.
Global Instance make_and_emp_r P : Affine P → MakeAnd P emp P.
Proof. intros. by rewrite /MakeAnd and_emp. Qed.
Global Instance make_and_emp_r_bare P : MakeAnd P emp (■ P) | 10.
Proof. intros. by rewrite /MakeAnd comm. Qed.
Global Instance make_and_default P Q : MakeAnd P Q (P ∧ Q) | 100.
594
595
596
597
598
599
600
601
602
603
604
605
606
607
608
609
610
611
612
613
614
615
616
617
618
619
620
621
Proof. by rewrite /MakeAnd. Qed.
Global Instance frame_and_l p R P1 P2 Q1 Q2 Q :
Frame p R P1 Q1 → MaybeFrame p R P2 Q2 →
MakeAnd Q1 Q2 Q → Frame p R (P1 ∧ P2) Q | 9.
Proof.
rewrite /Frame /MakeAnd => <- <- <- /=.
auto using and_intro, and_elim_l, and_elim_r, sep_mono.
Qed.
Global Instance frame_and_persistent_r R P1 P2 Q2 Q :
Frame true R P2 Q2 →
MakeAnd P1 Q2 Q → Frame true R (P1 ∧ P2) Q | 10.
Proof.
rewrite /Frame /MakeAnd => <- <- /=. rewrite -!persistently_and_bare_sep_l.
auto using and_intro, and_elim_l', and_elim_r'.
Qed.
Global Instance frame_and_r R P1 P2 Q2 Q :
TCOr (Affine R) (Absorbing P1) →
Frame false R P2 Q2 →
MakeAnd P1 Q2 Q → Frame false R (P1 ∧ P2) Q | 10.
Proof.
rewrite /Frame /MakeAnd=> ? <- <- /=. apply and_intro.
- by rewrite and_elim_l sep_elim_r.
- by rewrite and_elim_r.
Qed.
Class MakeOr (P Q PQ : PROP) := make_or : P ∨ Q ⊣⊢ PQ.
Arguments MakeOr _%I _%I _%I.
Global Instance make_or_true_l P : MakeOr True P True.
Proof. by rewrite /MakeOr left_absorb. Qed.
Global Instance make_or_true_r P : MakeOr P True True.
Proof. by rewrite /MakeOr right_absorb. Qed.
Global Instance make_or_emp_l P : Affine P → MakeOr emp P emp.
Proof. intros. by rewrite /MakeOr emp_or. Qed.
Global Instance make_or_emp_r P : Affine P → MakeOr P emp emp.
Proof. intros. by rewrite /MakeOr or_emp. Qed.
Global Instance make_or_default P Q : MakeOr P Q (P ∨ Q) | 100.
Global Instance frame_or_persistent_l R P1 P2 Q1 Q2 Q :
Frame true R P1 Q1 → MaybeFrame true R P2 Q2 → MakeOr Q1 Q2 Q →
Frame true R (P1 ∨ P2) Q | 9.
Proof. rewrite /Frame /MakeOr => <- <- <-. by rewrite -sep_or_l. Qed.
Global Instance frame_or_persistent_r R P1 P2 Q1 Q2 Q :
MaybeFrame true R P2 Q2 → MakeOr P1 Q2 Q →
Frame true R (P1 ∨ P2) Q | 10.
Proof.
rewrite /Frame /MaybeFrame /MakeOr => <- <- /=.
by rewrite sep_or_l sep_elim_r.
Global Instance frame_or R P1 P2 Q1 Q2 Q :
Frame false R P1 Q1 → Frame false R P2 Q2 → MakeOr Q1 Q2 Q →
Frame false R (P1 ∨ P2) Q.
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).
Proof.
rewrite /Frame=> ?. apply wand_intro_l.
by rewrite assoc (comm _ P1) -assoc wand_elim_r.
Qed.
Class MakeBare (P Q : PROP) := make_bare : ■ P ⊣⊢ Q.
Arguments MakeBare _%I _%I.
Global Instance make_bare_True : MakeBare True emp | 0.
Proof. by rewrite /MakeBare bare_True_emp bare_emp. Qed.
Global Instance make_bare_affine P : Affine P → MakeBare P P | 1.
Proof. intros. by rewrite /MakeBare affine_bare. Qed.
Global Instance make_bare_default P : MakeBare P (■ P) | 100.
Proof. by rewrite /MakeBare. Qed.
Global Instance frame_bare R P Q Q' :
Frame true R P Q → MakeBare Q Q' → Frame true R (■ P) Q'.
Proof.
rewrite /Frame /MakeBare=> <- <- /=.
by rewrite -{1}bare_idemp bare_sep_2.
Qed.
Class MakeSink (P Q : PROP) := make_sink : ▲ P ⊣⊢ Q.
Arguments MakeSink _%I _%I.
Global Instance make_sink_emp : MakeSink emp True | 0.
Proof. by rewrite /MakeSink -sink_True_emp sink_pure. Qed.
(* Note: there is no point in having an instance `Absorbing P → MakeSink P P`
because framing will never turn a proposition that is not absorbing into
something that is absorbing. *)
Global Instance make_sink_default P : MakeSink P (▲ P) | 100.
Proof. by rewrite /MakeSink. Qed.
Global Instance frame_sink p R P Q Q' :
Frame p R P Q → MakeSink Q Q' → Frame p R (▲ P) Q'.
Proof. rewrite /Frame /MakeSink=> <- <- /=. by rewrite sink_sep_r. Qed.
Class MakePersistently (P Q : PROP) := make_persistently : □ P ⊣⊢ Q.
Arguments MakePersistently _%I _%I.
Global Instance make_persistently_true : MakePersistently True True.
Proof. by rewrite /MakePersistently persistently_pure. Qed.
Global Instance make_persistently_emp : MakePersistently emp True.
Proof. by rewrite /MakePersistently -persistently_True_emp persistently_pure. Qed.
Global Instance make_persistently_default P : MakePersistently P (□ P) | 100.
Global Instance frame_persistently R P Q Q' :
Frame true R P Q → MakePersistently Q Q' → Frame true R (□ P) Q'.
rewrite /Frame /MakePersistently=> <- <- /=. rewrite -persistently_and_bare_sep_l.
by rewrite -persistently_sep_2 -persistently_and_sep_l_1 persistently_bare
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.
(* FromModal *)
Global Instance from_modal_sink P : FromModal (▲ P) P.
Proof. apply sink_intro. Qed.
716
717
718
719
720
721
722
723
724
725
726
727
728
729
730
731
732
733
734
735
736
737
738
739
740
741
742
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.
(* FromPure *)
Global Instance from_pure_later P φ : FromPure P φ → FromPure (▷ P) φ.
Proof. rewrite /FromPure=> ->. apply later_intro. Qed.
Global Instance from_pure_laterN n P φ : FromPure P φ → FromPure (▷^n P) φ.
Proof. rewrite /FromPure=> ->. apply laterN_intro. Qed.
Global Instance from_pure_except_0 P φ : FromPure P φ → FromPure (◇ P) φ.
Proof. rewrite /FromPure=> ->. apply except_0_intro. Qed.
(* IntoWand *)
Global Instance into_wand_later p q R P Q :
IntoWand p q R P Q → IntoWand p q (▷ R) (▷ P) (▷ Q).
Proof.
rewrite /IntoWand /= => HR. by rewrite !later_bare_persistently_if_2 -later_wand HR.
Qed.
Global Instance into_wand_later_args p q R P Q :
IntoWand p q R P Q → IntoWand' p q R (▷ P) (▷ Q).
Proof.
rewrite /IntoWand' /IntoWand /= => HR.
by rewrite !later_bare_persistently_if_2 (later_intro (⬕?p R)%I) -later_wand HR.
Qed.
Global Instance into_wand_laterN n p q R P Q :
IntoWand p q R P Q → IntoWand p q (▷^n R) (▷^n P) (▷^n Q).
Proof.
rewrite /IntoWand /= => HR. by rewrite !laterN_bare_persistently_if_2 -laterN_wand HR.
Qed.
Global Instance into_wand_laterN_args n p q R P Q :
IntoWand p q R P Q → IntoWand' p q R (▷^n P) (▷^n Q).
Proof.
rewrite /IntoWand' /IntoWand /= => HR.
by rewrite !laterN_bare_persistently_if_2 (laterN_intro _ (⬕?p R)%I) -laterN_wand HR.
766
767
768
769
770
771
772
773
774
775
776
777
778
779
780
781
782
783
784
785
786
787
788
789
790
791
792
(* FromAnd *)
Global Instance from_and_later P Q1 Q2 :
FromAnd P Q1 Q2 → FromAnd (▷ P) (▷ Q1) (▷ Q2).
Proof. rewrite /FromAnd=> <-. by rewrite later_and. Qed.
Global Instance from_and_laterN n P Q1 Q2 :
FromAnd P Q1 Q2 → FromAnd (▷^n P) (▷^n Q1) (▷^n Q2).
Proof. rewrite /FromAnd=> <-. by rewrite laterN_and. Qed.
Global Instance from_and_except_0 P Q1 Q2 :
FromAnd P Q1 Q2 → FromAnd (◇ P) (◇ Q1) (◇ Q2).
Proof. rewrite /FromAnd=><-. by rewrite except_0_and. Qed.
(* FromSep *)
Global Instance from_sep_later P Q1 Q2 :
FromSep P Q1 Q2 → FromSep (▷ P) (▷ Q1) (▷ Q2).
Proof. rewrite /FromSep=> <-. by rewrite later_sep. Qed.
Global Instance from_sep_laterN n P Q1 Q2 :
FromSep P Q1 Q2 → FromSep (▷^n P) (▷^n Q1) (▷^n Q2).
Proof. rewrite /FromSep=> <-. by rewrite laterN_sep. Qed.
Global Instance from_sep_except_0 P Q1 Q2 :
FromSep P Q1 Q2 → FromSep (◇ P) (◇ Q1) (◇ Q2).
Proof. rewrite /FromSep=><-. by rewrite except_0_sep. Qed.
(* IntoAnd *)
Global Instance into_and_later p P Q1 Q2 :
IntoAnd p P Q1 Q2 → IntoAnd p (▷ P) (▷ Q1) (▷ Q2).
Proof.
rewrite /IntoAnd=> HP. apply bare_persistently_if_intro'.
by rewrite later_bare_persistently_if_2 HP bare_persistently_if_elim later_and.
Qed.
Global Instance into_and_laterN n p P Q1 Q2 :
IntoAnd p P Q1 Q2 → IntoAnd p (▷^n P) (▷^n Q1) (▷^n Q2).
Proof.
rewrite /IntoAnd=> HP. apply bare_persistently_if_intro'.
by rewrite laterN_bare_persistently_if_2 HP bare_persistently_if_elim laterN_and.
Qed.
Global Instance into_and_except_0 p P Q1 Q2 :
IntoAnd p P Q1 Q2 → IntoAnd p (◇ P) (◇ Q1) (◇ Q2).
Proof.
rewrite /IntoAnd=> HP. apply bare_persistently_if_intro'.
by rewrite except_0_bare_persistently_if_2 HP bare_persistently_if_elim except_0_and.
Global Instance into_sep_later P Q1 Q2 :
IntoSep P Q1 Q2 → IntoSep (▷ P) (▷ Q1) (▷ Q2).
Proof. rewrite /IntoSep=> ->. by rewrite later_sep. Qed.
Global Instance into_sep_laterN n P Q1 Q2 :
IntoSep P Q1 Q2 → IntoSep (▷^n P) (▷^n Q1) (▷^n Q2).
Proof. rewrite /IntoSep=> ->. by rewrite laterN_sep. Qed.
Global Instance into_sep_except_0 P Q1 Q2 :
IntoSep P Q1 Q2 → IntoSep (◇ P) (◇ Q1) (◇ Q2).
Proof. rewrite /IntoSep=> ->. by rewrite except_0_sep. Qed.
(* FIXME: This instance is overly specific, generalize it. *)
Global Instance into_sep_bare_later `{!Timeless (emp%I : PROP)} P Q1 Q2 :
Affine Q1 → Affine Q2 → IntoSep P Q1 Q2 → IntoSep (■ ▷ P) (■ ▷ Q1) (■ ▷ Q2).
Proof.
rewrite /IntoSep /= => ?? ->.
rewrite -{1}(affine_bare Q1) -{1}(affine_bare Q2) later_sep !later_bare_1.
rewrite -except_0_sep /bi_except_0 bare_or. apply or_elim, bare_elim.
rewrite -(idemp bi_and (■ ▷ False)%I) persistent_and_sep_1.
by rewrite -(False_elim Q1) -(False_elim Q2).
Qed.
Global Instance from_or_later P Q1 Q2 :
FromOr P Q1 Q2 → FromOr (▷ P) (▷ Q1) (▷ Q2).
Proof. rewrite /FromOr=><-. by rewrite later_or. Qed.
Global Instance from_or_laterN n P Q1 Q2 :
FromOr P Q1 Q2 → FromOr (▷^n P) (▷^n Q1) (▷^n Q2).
Proof. rewrite /FromOr=><-. by rewrite laterN_or. Qed.
Global Instance from_or_except_0 P Q1 Q2 :
FromOr P Q1 Q2 → FromOr (◇ P) (◇ Q1) (◇ Q2).
Proof. rewrite /FromOr=><-. by rewrite except_0_or. Qed.
(* IntoOr *)
Global Instance into_or_later P Q1 Q2 :
IntoOr P Q1 Q2 → IntoOr (▷ P) (▷ Q1) (▷ Q2).
Proof. rewrite /IntoOr=>->. by rewrite later_or. Qed.
Global Instance into_or_laterN n P Q1 Q2 :
IntoOr P Q1 Q2 → IntoOr (▷^n P) (▷^n Q1) (▷^n Q2).
Proof. rewrite /IntoOr=>->. by rewrite laterN_or. Qed.
Global Instance into_or_except_0 P Q1 Q2 :
IntoOr P Q1 Q2 → IntoOr (◇ P) (◇ Q1) (◇ Q2).
Proof. rewrite /IntoOr=>->. by rewrite except_0_or. Qed.
(* FromExist *)
Global Instance from_exist_later {A} P (Φ : A → PROP) :
FromExist P Φ → FromExist (▷ P) (λ a, ▷ (Φ a))%I.
Proof.
rewrite /FromExist=> <-. apply exist_elim=>x. apply later_mono, exist_intro.
Qed.
Global Instance from_exist_laterN {A} n P (Φ : A → PROP) :
FromExist P Φ → FromExist (▷^n P) (λ a, ▷^n (Φ a))%I.
Proof.
rewrite /FromExist=> <-. apply exist_elim=>x. apply laterN_mono, exist_intro.
Qed.
Global Instance from_exist_except_0 {A} P (Φ : A → PROP) :
FromExist P Φ → FromExist (◇ P) (λ a, ◇ (Φ a))%I.
Proof. rewrite /FromExist=> <-. by rewrite except_0_exist_2. Qed.
(* IntoExist *)
Global Instance into_exist_later {A} P (Φ : A → PROP) :
IntoExist P Φ → Inhabited A → IntoExist (▷ P) (λ a, ▷ (Φ a))%I.
Proof. rewrite /IntoExist=> HP ?. by rewrite HP later_exist. Qed.
Global Instance into_exist_laterN {A} n P (Φ : A → PROP) :
IntoExist P Φ → Inhabited A → IntoExist (▷^n P) (λ a, ▷^n (Φ a))%I.
Proof. rewrite /IntoExist=> HP ?. by rewrite HP laterN_exist. Qed.
Global Instance into_exist_except_0 {A} P (Φ : A → PROP) :
IntoExist P Φ → Inhabited A → IntoExist (◇ P) (λ a, ◇ (Φ a))%I.
Proof. rewrite /IntoExist=> HP ?. by rewrite HP except_0_exist. Qed.
Global Instance into_forall_later {A} P (Φ : A → PROP) :
IntoForall P Φ → IntoForall (▷ P) (λ a, ▷ (Φ a))%I.
Proof. rewrite /IntoForall=> HP. by rewrite HP later_forall. Qed.
(* FromForall *)
Global Instance from_forall_later {A} P (Φ : A → PROP) :
FromForall P Φ → FromForall (▷ P)%I (λ a, ▷ (Φ a))%I.
Proof. rewrite /FromForall=> <-. by rewrite later_forall. Qed.
(* IsExcept0 *)
Global Instance is_except_0_except_0 P : IsExcept0 (◇ P).
Proof. by rewrite /IsExcept0 except_0_idemp. Qed.
Global Instance is_except_0_later P : IsExcept0 (▷ P).
Proof. by rewrite /IsExcept0 except_0_later. Qed.
(* FromModal *)
Global Instance from_modal_later P : FromModal (▷ P) P.
Proof. apply later_intro. Qed.
Global Instance from_modal_except_0 P : FromModal (◇ P) P.
Proof. apply except_0_intro. Qed.
(* IntoExcept0 *)
Global Instance into_timeless_except_0 P : IntoExcept0 (◇ P) P.
Proof. by rewrite /IntoExcept0. Qed.
Global Instance into_timeless_later P : Timeless P → IntoExcept0 (▷ P) P.
Proof. by rewrite /IntoExcept0. Qed.
Global Instance into_timeless_later_if p P : Timeless P → IntoExcept0 (▷?p P) P.
Proof. rewrite /IntoExcept0. destruct p; auto using except_0_intro. Qed.
Global Instance into_timeless_bare P Q : IntoExcept0 P Q → IntoExcept0 (■ P) (■ Q).
Proof. rewrite /IntoExcept0=> ->. by rewrite except_0_bare_2. Qed.
Global Instance into_timeless_sink P Q : IntoExcept0 P Q → IntoExcept0 (▲ P) (▲ Q).
Proof. rewrite /IntoExcept0=> ->. by rewrite except_0_sink. Qed.
Global Instance into_timeless_persistently P Q : IntoExcept0 P Q → IntoExcept0 (□ P) (□ Q).
Proof. rewrite /IntoExcept0=> ->. by rewrite except_0_persistently. Qed.
(* ElimModal *)
Global Instance elim_modal_timeless P Q :
IntoExcept0 P P' → IsExcept0 Q → ElimModal P P' Q Q.
intros. rewrite /ElimModal (except_0_intro (_ -∗ _)%I).
by rewrite (into_except_0 P) -except_0_sep wand_elim_r.
(* Frame *)
Class MakeLater (P lP : PROP) := make_later : ▷ P ⊣⊢ lP.
Arguments MakeLater _%I _%I.
Global Instance make_later_true : MakeLater True True.
Proof. by rewrite /MakeLater later_True. Qed.
Global Instance make_later_default P : MakeLater P (▷ P) | 100.
Proof. by rewrite /MakeLater. Qed.
Global Instance frame_later p R R' P Q Q' :
IntoLaterN 1 R' R → Frame p R P Q → MakeLater Q Q' → Frame p R' (▷ P) Q'.
Proof.
rewrite /Frame /MakeLater /IntoLaterN=>-> <- <- /=.
by rewrite later_bare_persistently_if_2 later_sep.
Qed.
Class MakeLaterN (n : nat) (P lP : PROP) := make_laterN : ▷^n P ⊣⊢ lP.
Arguments MakeLaterN _%nat _%I _%I.
Global Instance make_laterN_true n : MakeLaterN n True True.
Proof. by rewrite /MakeLaterN laterN_True. Qed.
Global Instance make_laterN_default P : MakeLaterN n P (▷^n P) | 100.
Proof. by rewrite /MakeLaterN. Qed.
Global Instance frame_laterN p n R R' P Q Q' :
IntoLaterN n R' R → Frame p R P Q → MakeLaterN n Q Q' → Frame p R' (▷^n P) Q'.
Proof.
rewrite /Frame /MakeLaterN /IntoLaterN=>-> <- <-.
by rewrite laterN_bare_persistently_if_2 laterN_sep.
Qed.
Class MakeExcept0 (P Q : PROP) := make_except_0 : ◇ P ⊣⊢ Q.
Arguments MakeExcept0 _%I _%I.
Global Instance make_except_0_True : MakeExcept0 True True.
Proof. by rewrite /MakeExcept0 except_0_True. Qed.
Global Instance make_except_0_default P : MakeExcept0 P (◇ P) | 100.
Proof. by rewrite /MakeExcept0. Qed.
Global Instance frame_except_0 p R P Q Q' :
Frame p R P Q → MakeExcept0 Q Q' → Frame p R (◇ P) Q'.
rewrite /Frame /MakeExcept0=><- <-.
by rewrite except_0_sep -(except_0_intro (⬕?p R)%I).
969
970
971
972
973
974
975
976
977
978
979
980
981
982
983
984
985
986
987
988
989
990
991
992
993
994
995
996
997
998
999
1000
(* IntoLater *)
Global Instance into_laterN_later n P Q :
IntoLaterN n P Q → IntoLaterN' (S n) (▷ P) Q.
Proof. by rewrite /IntoLaterN' /IntoLaterN =>->. Qed.
Global Instance into_laterN_laterN n P : IntoLaterN' n (▷^n P) P.
Proof. by rewrite /IntoLaterN' /IntoLaterN. Qed.
Global Instance into_laterN_laterN_plus n m P Q :
IntoLaterN m P Q → IntoLaterN' (n + m) (▷^n P) Q.
Proof. rewrite /IntoLaterN' /IntoLaterN=>->. by rewrite laterN_plus. Qed.
Global Instance into_laterN_and_l n P1 P2 Q1 Q2 :
IntoLaterN' n P1 Q1 → IntoLaterN n P2 Q2 →
IntoLaterN' n (P1 ∧ P2) (Q1 ∧ Q2) | 10.
Proof. rewrite /IntoLaterN' /IntoLaterN=> -> ->. by rewrite laterN_and. Qed.
Global Instance into_laterN_and_r n P P2 Q2 :
IntoLaterN' n P2 Q2 → IntoLaterN' n (P ∧ P2) (P ∧ Q2) | 11.
Proof.
rewrite /IntoLaterN' /IntoLaterN=> ->. by rewrite laterN_and -(laterN_intro _ P).
Qed.
Global Instance into_laterN_or_l n P1 P2 Q1 Q2 :
IntoLaterN' n P1 Q1 → IntoLaterN n P2 Q2 →
IntoLaterN' n (P1 ∨ P2) (Q1 ∨ Q2) | 10.
Proof. rewrite /IntoLaterN' /IntoLaterN=> -> ->. by rewrite laterN_or. Qed.
Global Instance into_laterN_or_r n P P2 Q2 :
IntoLaterN' n P2 Q2 →
IntoLaterN' n (P ∨ P2) (P ∨ Q2) | 11.
Proof.
rewrite /IntoLaterN' /IntoLaterN=> ->. by rewrite laterN_or -(laterN_intro _ P).
Qed.
Global Instance into_laterN_forall {A} n (Φ Ψ : A → PROP) :