RunST.naught
From iris.base_logic Require Import upred derived primitive.
From iris.proofmode Require Export classes.
Import uPred_primitive.
Import uPred.
Require Import iris.proofmode.tactics.
(* Some interesting lemmas! *)
Lemma core_unit {M : ucmraT} : core (∅ : M) ≡ ∅.
Proof. by rewrite /core /core' ucmra_pcore_unit. Qed.
Lemma except_0_sep' (M : ucmraT) (P Q : uPred M) : (◇ (P ★ Q))%I ⊢ (◇ P ★ ◇ Q)%I.
Proof. by rewrite except_0_sep. Qed.
Lemma except_0_exist' (M : ucmraT) A `{Inhabited A} (P : A → uPred M) :
(◇ ∃ x, P x) ⊢ ∃ x, ◇ P x.
Proof.
rewrite /uPred_except_0; unseal; constructor => n x Hnx Hde.
destruct n.
- exists inhabitant; by left.
- destruct Hde as [Hde|Hde]; first done. destruct Hde as [z Hde].
by exists z; right.
Qed.
Lemma except_0_forall' (M : ucmraT) A (P : A → uPred M) :
(∀ x, ◇ P x) ⊢ (◇ ∀ x, P x).
Proof.
rewrite /uPred_except_0; unseal; constructor => n x Hnx Hde.
destruct n; first by left.
right=>z. specialize (Hde z); destruct Hde as [Hde|Hde]; done.
Qed.
Lemma later_n_except_0_exist A `{Inhabited A} n M (P : A → uPred M) :
Nat.iter n (λ Q, ▷ Q) (◇ (∃ x : A, P x)) ⊣⊢
∃ x : A, (Nat.iter n (λ Q, ▷ Q) (◇ P x)).
Proof.
induction n.
- simpl. iSplit.
apply except_0_exist'; auto.
iDestruct 1 as (x) "H". iApply except_0_mono; last eauto.
by iIntros "?"; iExists _.
- simpl. by rewrite IHn later_exist.
Qed.
Lemma later_n_except_0_forall A `{Inhabited A} n M (P : A → uPred M) :
Nat.iter n (λ Q, ▷ Q) (◇ (∀ x : A, P x)) ⊣⊢
∀ x : A, (Nat.iter n (λ Q, ▷ Q) (◇ P x)).
Proof.
induction n.
- simpl. iSplit; last apply except_0_forall'; auto.
iIntros "H"; iIntros (x). iApply except_0_mono; last eauto.
iIntros "H"; iApply "H".
- simpl. by rewrite IHn later_forall.
Qed.
(* Naught modality *)
Program Definition uPred_naught_def {M} (P : uPred M) : uPred M :=
{| uPred_holds n x := P n ∅ |}.
Next Obligation.
intros M; naive_solver eauto using uPred_mono, @cmra_core_monoN.
Qed.
Next Obligation. naive_solver eauto using uPred_closed, @ucmra_unit_validN. Qed.
Definition uPred_naught_aux : { x | x = @uPred_naught_def }. by eexists. Qed.
Definition uPred_naught {M} := proj1_sig uPred_naught_aux M.
Definition uPred_naught_eq :
@uPred_naught = @uPred_naught_def := proj2_sig uPred_naught_aux.
(* Notation for naught: ⋊ (\rtimes) *)
Notation "⋊ P" := (uPred_naught P)
(at level 20, right associativity) : uPred_scope.
(* A plain fact is one that holds independent of resources. *)
Class PlainP {M} (P : uPred M) := plainP : P ⊢ ⋊ P.
Arguments plainP {_} _ {_}.
Section naught_properties.
Context {M : ucmraT}.
Implicit Types φ : Prop.
Implicit Types P Q : uPred M.
Implicit Types A : Type.
Notation "P ⊢ Q" := (@uPred_entails M P%I Q%I). (* Force implicit argument M *)
Notation "P ⊣⊢ Q" := (equiv (A:=uPred M) P%I Q%I). (* Force implicit argument M *)
(* naught primitive properties *)
Global Instance naught_ne n : Proper (dist n ==> dist n) (@uPred_naught M).
Proof.
intros P1 P2 HP.
rewrite uPred_naught_eq; split=> n' x; split; apply HP;
eauto using @ucmra_unit_validN.
Qed.
Global Instance naught_proper :
Proper ((⊣⊢) ==> (⊣⊢)) (@uPred_naught M) := ne_proper _.
Lemma true_entails_naught P : (True ⊢ P) → (True ⊢ ⋊ P).
Proof.
unseal; rewrite uPred_naught_eq; intros HP. split=>n x Hx HT.
apply HP; last done; apply ucmra_unit_validN.
Qed.
Lemma naught_irrel P n x y : (⋊ P)%I n x → (⋊ P)%I n y.
Proof. by rewrite uPred_naught_eq. Qed.
Lemma naught_keeps_premises P Q : (P ⊢ ⋊ Q) → P ⊢ P ★ ⋊ Q.
Proof.
intros HPQ. split=> n x Hx HPx. unseal.
exists x, (∅ : M); repeat split; trivial; first by rewrite right_id.
apply HPQ in HPx; last trivial. eauto using naught_irrel.
Qed.
Lemma naught_mono P Q : (P ⊢ Q) → ⋊ P ⊢ ⋊ Q.
Proof.
intros HP; rewrite uPred_naught_eq; split=> n x ? /=.
apply HP, ucmra_unit_validN.
Qed.
Lemma naught_elim P : ⋊ P ⊢ P.
Proof.
rewrite uPred_naught_eq; split=> n x ? ? /=.
by eapply uPred_mono; last eapply cmra_included_includedN, @ucmra_unit_least.
Qed.
Lemma naught_idemp P : ⋊ P ⊢ ⋊ ⋊ P.
Proof. by rewrite uPred_naught_eq. Qed.
Lemma naught_pure_2 φ : ■ φ ⊢ ⋊ ■ φ.
Proof. by unseal; rewrite uPred_naught_eq. Qed.
Lemma naught_implies_always P : ⋊ P ⊢ □ P.
Proof.
unseal; rewrite uPred_naught_eq. split=> n x Hx.
cbv [uPred_naught_def uPred_always_def uPred_holds]=> HP.
eapply uPred_mono; eauto; apply ucmra_unit_leastN.
Qed.
Lemma naught_always_1 P : □ ⋊ P ⊢ ⋊ □ P.
Proof.
unseal; rewrite uPred_naught_eq; split=> n x ? /=.
cbv [uPred_naught_def uPred_always_def uPred_holds].
by rewrite core_unit.
Qed.
Lemma naught_always_2 P : ⋊ □ P ⊢ □ ⋊ P.
Proof.
unseal; rewrite uPred_naught_eq; split=> n x ? /=.
cbv [uPred_naught_def uPred_always_def uPred_holds].
by rewrite core_unit.
Qed.
Lemma naught_forall_2 {A} (Ψ : A → uPred M) : (∀ a, ⋊ Ψ a) ⊢ (⋊ ∀ a, Ψ a).
Proof. by unseal; rewrite uPred_naught_eq. Qed.
Lemma naught_exist_1 {A} (Ψ : A → uPred M) : (⋊ ∃ a, Ψ a) ⊢ (∃ a, ⋊ Ψ a).
Proof. by unseal; rewrite uPred_naught_eq. Qed.
Lemma naught_and_sep_1 P Q : ⋊ (P ∧ Q) ⊢ ⋊ (P ★ Q).
Proof.
unseal; rewrite uPred_naught_eq; split=> n x ? [??].
exists (∅ : M), (∅ : M); rewrite left_id; auto.
Qed.
Lemma naught_and_sep_l_1 P Q : ⋊ P ∧ Q ⊢ ⋊ P ★ Q.
Proof.
unseal; rewrite uPred_naught_eq; split=> n x ? [??]; exists ∅, x; simpl in *.
by rewrite left_id.
Qed.
Lemma naught_later P : ⋊ ▷ P ⊣⊢ ▷ ⋊ P.
Proof. by unseal; rewrite uPred_naught_eq. Qed.
Lemma naught_ownM : ⋊ uPred_ownM ∅ ⊣⊢ uPred_ownM ∅. (* perhaps quite useless!? *)
Proof.
intros; apply (anti_symm _); first by apply:naught_elim.
unseal; rewrite uPred_naught_eq; split=> n x ? ?.
by cbv [uPred_naught_def uPred_ownM_def uPred_holds].
Qed.
Lemma naught_cmra_valid_1 {A : cmraT} (a : A) : ✓ a ⊢ ⋊ ✓ a.
Proof. by unseal; rewrite uPred_naught_eq. Qed.
Hint Resolve pure_intro.
Hint Resolve or_elim or_intro_l' or_intro_r'.
Hint Resolve and_intro and_elim_l' and_elim_r'.
Hint Immediate True_intro False_elim.
(* Naught derived *)
Hint Resolve naught_mono naught_elim.
Global Instance naught_mono' : Proper ((⊢) ==> (⊢)) (@uPred_naught M).
Proof. intros P Q; apply naught_mono. Qed.
Global Instance naught_flip_mono' :
Proper (flip (⊢) ==> flip (⊢)) (@uPred_naught M).
Proof. intros P Q; apply naught_mono. Qed.
Lemma naught_intro' P Q : (⋊ P ⊢ Q) → ⋊ P ⊢ ⋊ Q.
Proof. intros <-. apply naught_idemp. Qed.
Lemma naught_pure φ : ⋊ ■ φ ⊣⊢ ■ φ.
Proof. apply (anti_symm _); auto using naught_pure_2. Qed.
Lemma naught_forall {A} (Ψ : A → uPred M) : (⋊ ∀ a, Ψ a) ⊣⊢ (∀ a, ⋊ Ψ a).
Proof.
apply (anti_symm _); auto using naught_forall_2.
apply forall_intro=> x. by rewrite (forall_elim x).
Qed.
Lemma naught_exist {A} (Ψ : A → uPred M) : (⋊ ∃ a, Ψ a) ⊣⊢ (∃ a, ⋊ Ψ a).
Proof.
apply (anti_symm _); auto using naught_exist_1.
apply exist_elim=> x. by rewrite (exist_intro x).
Qed.
Lemma naught_and P Q : ⋊ (P ∧ Q) ⊣⊢ ⋊ P ∧ ⋊ Q.
Proof. rewrite !and_alt naught_forall. by apply forall_proper=> -[]. Qed.
Lemma naught_or P Q : ⋊ (P ∨ Q) ⊣⊢ ⋊ P ∨ ⋊ Q.
Proof. rewrite !or_alt naught_exist. by apply exist_proper=> -[]. Qed.
Lemma naught_impl P Q : ⋊ (P → Q) ⊢ ⋊ P → ⋊ Q.
Proof.
apply impl_intro_l; rewrite -naught_and.
apply naught_mono, impl_elim with P; auto.
Qed.
Lemma naught_internal_eq {A:cofeT} (a b : A) : ⋊ (a ≡ b) ⊣⊢ a ≡ b.
Proof.
apply (anti_symm (⊢)); auto using naught_elim.
apply (internal_eq_rewrite a b (λ b, ⋊ (a ≡ b))%I); auto.
{ intros n; solve_proper. }
rewrite -(internal_eq_refl a) naught_pure; auto.
Qed.
Lemma naught_always P : ⋊ □ P ⊣⊢ □ ⋊ P.
Proof.
apply (anti_symm _); auto using naught_always_1, naught_always_2.
Qed.
Lemma naught_and_sep P Q : ⋊ (P ∧ Q) ⊣⊢ ⋊ (P ★ Q).
Proof.
apply (anti_symm (⊢)); last rewrite sep_and; auto using naught_and_sep_1.
Qed.
Lemma naught_and_sep_l' P Q : ⋊ P ∧ Q ⊣⊢ ⋊ P ★ Q.
Proof.
apply (anti_symm (⊢)); last rewrite sep_and; auto using naught_and_sep_l_1.
Qed.
Lemma naught_and_sep_r' P Q : P ∧ ⋊ Q ⊣⊢ P ★ ⋊ Q.
Proof. by rewrite !(comm _ P) naught_and_sep_l'. Qed.
Lemma naught_sep P Q : ⋊ (P ★ Q) ⊣⊢ ⋊ P ★ ⋊ Q.
Proof. by rewrite -naught_and_sep -naught_and_sep_l' naught_and. Qed.
Lemma naught_sep_dup' P : ⋊ P ⊣⊢ ⋊ P ★ ⋊ P.
Proof. by rewrite -naught_sep -naught_and_sep (idemp _). Qed.
Lemma naught_wand P Q : ⋊ (P -★ Q) ⊢ ⋊ P -★ ⋊ Q.
Proof. by apply wand_intro_r; rewrite -naught_sep wand_elim_l. Qed.
Lemma naught_wand_impl P Q : ⋊ (P -★ Q) ⊣⊢ ⋊ (P → Q).
Proof.
apply (anti_symm (⊢)); [|by rewrite -impl_wand].
apply naught_intro', impl_intro_r.
by rewrite naught_and_sep_l' naught_elim wand_elim_l.
Qed.
Lemma naught_entails_l' P Q : (P ⊢ ⋊ Q) → P ⊢ ⋊ Q ★ P.
Proof. intros; rewrite -naught_and_sep_l'; auto. Qed.
Lemma naught_entails_r' P Q : (P ⊢ ⋊ Q) → P ⊢ P ★ ⋊ Q.
Proof. intros; rewrite -naught_and_sep_r'; auto. Qed.
Lemma except_0_naught P : ◇ ⋊ P ⊣⊢ ⋊ ◇ P.
Proof. by rewrite /uPred_except_0 naught_or naught_later naught_pure. Qed.
Lemma naught_cmra_valid {A : cmraT} (a : A) : ⋊ ✓ a ⊣⊢ ✓ a.
Proof.
intros; apply (anti_symm _); first by apply:naught_elim.
apply:naught_cmra_valid_1.
Qed.
(* Plainness *)
Global Instance pure_plain φ : PlainP (■ φ : uPred M)%I.
Proof. by rewrite /PlainP naught_pure. Qed.
Global Instance naught_plain P : PlainP (⋊ P).
Proof. by intros; apply naught_intro'. Qed.
Global Instance and_plain P Q :
PlainP P → PlainP Q → PlainP (P ∧ Q).
Proof. by intros; rewrite /PlainP naught_and; apply and_mono. Qed.
Global Instance or_plain P Q :
PlainP P → PlainP Q → PlainP (P ∨ Q).
Proof. by intros; rewrite /PlainP naught_or; apply or_mono. Qed.
Global Instance sep_plain P Q :
PlainP P → PlainP Q → PlainP (P ★ Q).
Proof. by intros; rewrite /PlainP naught_sep; apply sep_mono. Qed.
Global Instance forall_plain {A} (Ψ : A → uPred M) :
(∀ x, PlainP (Ψ x)) → PlainP (∀ x, Ψ x).
Proof. by intros; rewrite /PlainP naught_forall; apply forall_mono. Qed.
Global Instance exist_plain {A} (Ψ : A → uPred M) :
(∀ x, PlainP (Ψ x)) → PlainP (∃ x, Ψ x).
Proof. by intros; rewrite /PlainP naught_exist; apply exist_mono. Qed.
Global Instance internal_eq_plain {A : cofeT} (a b : A) :
PlainP (a ≡ b : uPred M)%I.
Proof. by intros; rewrite /PlainP naught_internal_eq. Qed.
Global Instance cmra_valid_plain {A : cmraT} (a : A) :
PlainP (✓ a : uPred M)%I.
Proof. by intros; rewrite /PlainP naught_cmra_valid. Qed.
Global Instance later_plain P : PlainP P → PlainP (▷ P).
Proof. by intros; rewrite /PlainP naught_later; apply later_mono. Qed.
Global Instance ownM_plain : PlainP (@uPred_ownM M ∅). (* perhaps quite useless!? *)
Proof. intros. by rewrite /PlainP naught_ownM. Qed.
Global Instance from_option_plain {A} P (Ψ : A → uPred M) (mx : option A) :
(∀ x, PlainP (Ψ x)) → PlainP P → PlainP (from_option Ψ P mx).
Proof. destruct mx; apply _. Qed.
Global Instance except_0_plain P : PlainP P → PlainP (◇ P).
Proof. intros; by rewrite /PlainP -except_0_naught except_0_mono. Qed.
(* Plainness implies persistence *)
Global Instance plain_persistent P : PlainP P → PersistentP P.
Proof.
rewrite /PlainP /PersistentP => HP.
by etrans; last apply naught_implies_always.
Qed.
(* Derived lemmas for plainness *)
Lemma naught_naught P `{!PlainP P} : ⋊ P ⊣⊢ P.
Proof. apply (anti_symm (⊢)); auto. Qed.
Lemma naught_intro P Q `{!PlainP P} : (P ⊢ Q) → P ⊢ ⋊ Q.
Proof. rewrite -(naught_naught P); apply naught_intro'. Qed.
Lemma naught_and_sep_l P Q `{!PlainP P} : P ∧ Q ⊣⊢ P ★ Q.
Proof. by rewrite -(naught_naught P) naught_and_sep_l'. Qed.
Lemma naught_and_sep_r P Q `{!PlainP Q} : P ∧ Q ⊣⊢ P ★ Q.
Proof. by rewrite -(naught_naught Q) naught_and_sep_r'. Qed.
Lemma naught_sep_dup P `{!PlainP P} : P ⊣⊢ P ★ P.
Proof. by rewrite -(naught_naught P) -naught_sep_dup'. Qed.
Lemma naught_entails_l P Q `{!PlainP Q} : (P ⊢ Q) → P ⊢ Q ★ P.
Proof. by rewrite -(naught_naught Q); apply naught_entails_l'. Qed.
Lemma naught_entails_r P Q `{!PlainP Q} : (P ⊢ Q) → P ⊢ P ★ Q.
Proof. by rewrite -(naught_naught Q); apply naught_entails_r'. Qed.
(* important for proving properties in the model. *)
Lemma naught_irrel' P `{HPl: !PlainP P} n x y : ✓{n} x → ✓{n} y → P n x → P n y.
Proof.
intros Hx Hy Hnx; apply HPl in Hnx; last done.
apply naught_elim; first done. eapply naught_irrel; eauto.
Qed.
Lemma naught_keeps_premises' P Q `{HPl: !PlainP Q} : (P ⊢ Q) → P ⊢ P ★ Q.
Proof.
iIntros (HPQ) "HP".
assert (HPQ': P ⊢ ⋊ Q) by (by iIntros "HP"; iApply HPl; iApply HPQ).
iPoseProof (naught_keeps_premises _ _ HPQ' with "HP") as "HPQ".
iDestruct "HPQ" as "[HP HQ]"; iSplitL "HP"; first trivial.
by iApply naught_elim.
Qed.
Global Instance iter_Plain_Plain P `{Hpl: !PlainP P} n f :
(∀ Q, PlainP Q -> PlainP (f Q)) →
PlainP (Nat.iter n f P).
Proof.
revert P Hpl. induction n => P Hpl; first done.
rewrite Nat_iter_S_r; auto.
Qed.
End naught_properties.
Lemma bupd_extract_plain {M : ucmraT} (P Q : uPred M) `{!PlainP P} :
(Q ⊢ P) → (|==> Q) ⊢ P.
Proof.
intros HQP; unseal; split=> n x Hnx Hbp.
destruct (Hbp n ∅ (le_n _)) as [z [Hz1 Hz2]]; first by rewrite right_id.
revert Hz1; rewrite right_id=> Hz1.
apply HQP in Hz2; trivial. eauto using naught_irrel', ucmra_unit_validN.
Qed.
Section naught_proofmode_classes.
Context {M : ucmraT}.
Implicit Types P Q R : uPred M.
(* FromAssumption *)
Global Instance from_assumption_naught_l p P Q :
FromAssumption p P Q → FromAssumption p (⋊ P) Q.
Proof. rewrite /FromAssumption=><-. by rewrite naught_elim. Qed.
(* IntoPersistentP *)
Global Instance into_persistentP_naught P : IntoPersistentP (⋊ P) P | 1.
Proof. apply naught_implies_always. Qed.
Global Instance into_PlainP_persistent P :
PlainP P → IntoPersistentP P P | 100.
Proof.
rewrite /IntoPersistentP => HP; by etrans; [|apply naught_implies_always].
Qed.
(* FromAnd *)
Global Instance from_and_sep_Plain_l P1 P2 :
PlainP P1 → FromAnd (P1 ★ P2) P1 P2 | 9.
Proof. intros. by rewrite /FromAnd naught_and_sep_l. Qed.
Global Instance from_and_sep_Plain_r P1 P2 :
PlainP P2 → FromAnd (P1 ★ P2) P1 P2 | 10.
Proof. intros. by rewrite /FromAnd naught_and_sep_r. Qed.
Global Instance from_and_naught P Q1 Q2 :
FromAnd P Q1 Q2 → FromAnd (⋊ P) (⋊ Q1) (⋊ Q2).
Proof. rewrite /FromAnd=> <-. by rewrite naught_and. Qed.
(* FromSep *)
Global Instance from_sep_naught P Q1 Q2 :
FromSep P Q1 Q2 → FromSep (⋊ P) (⋊ Q1) (⋊ Q2).
Proof. rewrite /FromSep=> <-. by rewrite naught_sep. Qed.
(* IntoAnd *)
Global Instance into_and_and_plain_l P Q :
PlainP P → IntoAnd false (P ∧ Q) P Q.
Proof. intros; by rewrite /IntoAnd /= naught_and_sep_l. Qed.
Global Instance into_and_and_plain_r P Q :
PlainP Q → IntoAnd false (P ∧ Q) P Q.
Proof. intros; by rewrite /IntoAnd /= always_and_sep_r. Qed.
Global Instance into_exist_naught {A} P (Φ : A → uPred M) :
IntoExist P Φ → IntoExist (⋊ P) (λ a, ⋊ (Φ a))%I.
Proof. rewrite /IntoExist=> HP. by rewrite HP naught_exist. Qed.
(* ElimModal *)
Global Instance elim_modal_naught P Q : ElimModal (⋊ P) P Q Q.
Proof. intros. by rewrite /ElimModal naught_elim wand_elim_r. Qed.
End naught_proofmode_classes.
From iris.proofmode Require Export classes.
Import uPred_primitive.
Import uPred.
Require Import iris.proofmode.tactics.
(* Some interesting lemmas! *)
Lemma core_unit {M : ucmraT} : core (∅ : M) ≡ ∅.
Proof. by rewrite /core /core' ucmra_pcore_unit. Qed.
Lemma except_0_sep' (M : ucmraT) (P Q : uPred M) : (◇ (P ★ Q))%I ⊢ (◇ P ★ ◇ Q)%I.
Proof. by rewrite except_0_sep. Qed.
Lemma except_0_exist' (M : ucmraT) A `{Inhabited A} (P : A → uPred M) :
(◇ ∃ x, P x) ⊢ ∃ x, ◇ P x.
Proof.
rewrite /uPred_except_0; unseal; constructor => n x Hnx Hde.
destruct n.
- exists inhabitant; by left.
- destruct Hde as [Hde|Hde]; first done. destruct Hde as [z Hde].
by exists z; right.
Qed.
Lemma except_0_forall' (M : ucmraT) A (P : A → uPred M) :
(∀ x, ◇ P x) ⊢ (◇ ∀ x, P x).
Proof.
rewrite /uPred_except_0; unseal; constructor => n x Hnx Hde.
destruct n; first by left.
right=>z. specialize (Hde z); destruct Hde as [Hde|Hde]; done.
Qed.
Lemma later_n_except_0_exist A `{Inhabited A} n M (P : A → uPred M) :
Nat.iter n (λ Q, ▷ Q) (◇ (∃ x : A, P x)) ⊣⊢
∃ x : A, (Nat.iter n (λ Q, ▷ Q) (◇ P x)).
Proof.
induction n.
- simpl. iSplit.
apply except_0_exist'; auto.
iDestruct 1 as (x) "H". iApply except_0_mono; last eauto.
by iIntros "?"; iExists _.
- simpl. by rewrite IHn later_exist.
Qed.
Lemma later_n_except_0_forall A `{Inhabited A} n M (P : A → uPred M) :
Nat.iter n (λ Q, ▷ Q) (◇ (∀ x : A, P x)) ⊣⊢
∀ x : A, (Nat.iter n (λ Q, ▷ Q) (◇ P x)).
Proof.
induction n.
- simpl. iSplit; last apply except_0_forall'; auto.
iIntros "H"; iIntros (x). iApply except_0_mono; last eauto.
iIntros "H"; iApply "H".
- simpl. by rewrite IHn later_forall.
Qed.
(* Naught modality *)
Program Definition uPred_naught_def {M} (P : uPred M) : uPred M :=
{| uPred_holds n x := P n ∅ |}.
Next Obligation.
intros M; naive_solver eauto using uPred_mono, @cmra_core_monoN.
Qed.
Next Obligation. naive_solver eauto using uPred_closed, @ucmra_unit_validN. Qed.
Definition uPred_naught_aux : { x | x = @uPred_naught_def }. by eexists. Qed.
Definition uPred_naught {M} := proj1_sig uPred_naught_aux M.
Definition uPred_naught_eq :
@uPred_naught = @uPred_naught_def := proj2_sig uPred_naught_aux.
(* Notation for naught: ⋊ (\rtimes) *)
Notation "⋊ P" := (uPred_naught P)
(at level 20, right associativity) : uPred_scope.
(* A plain fact is one that holds independent of resources. *)
Class PlainP {M} (P : uPred M) := plainP : P ⊢ ⋊ P.
Arguments plainP {_} _ {_}.
Section naught_properties.
Context {M : ucmraT}.
Implicit Types φ : Prop.
Implicit Types P Q : uPred M.
Implicit Types A : Type.
Notation "P ⊢ Q" := (@uPred_entails M P%I Q%I). (* Force implicit argument M *)
Notation "P ⊣⊢ Q" := (equiv (A:=uPred M) P%I Q%I). (* Force implicit argument M *)
(* naught primitive properties *)
Global Instance naught_ne n : Proper (dist n ==> dist n) (@uPred_naught M).
Proof.
intros P1 P2 HP.
rewrite uPred_naught_eq; split=> n' x; split; apply HP;
eauto using @ucmra_unit_validN.
Qed.
Global Instance naught_proper :
Proper ((⊣⊢) ==> (⊣⊢)) (@uPred_naught M) := ne_proper _.
Lemma true_entails_naught P : (True ⊢ P) → (True ⊢ ⋊ P).
Proof.
unseal; rewrite uPred_naught_eq; intros HP. split=>n x Hx HT.
apply HP; last done; apply ucmra_unit_validN.
Qed.
Lemma naught_irrel P n x y : (⋊ P)%I n x → (⋊ P)%I n y.
Proof. by rewrite uPred_naught_eq. Qed.
Lemma naught_keeps_premises P Q : (P ⊢ ⋊ Q) → P ⊢ P ★ ⋊ Q.
Proof.
intros HPQ. split=> n x Hx HPx. unseal.
exists x, (∅ : M); repeat split; trivial; first by rewrite right_id.
apply HPQ in HPx; last trivial. eauto using naught_irrel.
Qed.
Lemma naught_mono P Q : (P ⊢ Q) → ⋊ P ⊢ ⋊ Q.
Proof.
intros HP; rewrite uPred_naught_eq; split=> n x ? /=.
apply HP, ucmra_unit_validN.
Qed.
Lemma naught_elim P : ⋊ P ⊢ P.
Proof.
rewrite uPred_naught_eq; split=> n x ? ? /=.
by eapply uPred_mono; last eapply cmra_included_includedN, @ucmra_unit_least.
Qed.
Lemma naught_idemp P : ⋊ P ⊢ ⋊ ⋊ P.
Proof. by rewrite uPred_naught_eq. Qed.
Lemma naught_pure_2 φ : ■ φ ⊢ ⋊ ■ φ.
Proof. by unseal; rewrite uPred_naught_eq. Qed.
Lemma naught_implies_always P : ⋊ P ⊢ □ P.
Proof.
unseal; rewrite uPred_naught_eq. split=> n x Hx.
cbv [uPred_naught_def uPred_always_def uPred_holds]=> HP.
eapply uPred_mono; eauto; apply ucmra_unit_leastN.
Qed.
Lemma naught_always_1 P : □ ⋊ P ⊢ ⋊ □ P.
Proof.
unseal; rewrite uPred_naught_eq; split=> n x ? /=.
cbv [uPred_naught_def uPred_always_def uPred_holds].
by rewrite core_unit.
Qed.
Lemma naught_always_2 P : ⋊ □ P ⊢ □ ⋊ P.
Proof.
unseal; rewrite uPred_naught_eq; split=> n x ? /=.
cbv [uPred_naught_def uPred_always_def uPred_holds].
by rewrite core_unit.
Qed.
Lemma naught_forall_2 {A} (Ψ : A → uPred M) : (∀ a, ⋊ Ψ a) ⊢ (⋊ ∀ a, Ψ a).
Proof. by unseal; rewrite uPred_naught_eq. Qed.
Lemma naught_exist_1 {A} (Ψ : A → uPred M) : (⋊ ∃ a, Ψ a) ⊢ (∃ a, ⋊ Ψ a).
Proof. by unseal; rewrite uPred_naught_eq. Qed.
Lemma naught_and_sep_1 P Q : ⋊ (P ∧ Q) ⊢ ⋊ (P ★ Q).
Proof.
unseal; rewrite uPred_naught_eq; split=> n x ? [??].
exists (∅ : M), (∅ : M); rewrite left_id; auto.
Qed.
Lemma naught_and_sep_l_1 P Q : ⋊ P ∧ Q ⊢ ⋊ P ★ Q.
Proof.
unseal; rewrite uPred_naught_eq; split=> n x ? [??]; exists ∅, x; simpl in *.
by rewrite left_id.
Qed.
Lemma naught_later P : ⋊ ▷ P ⊣⊢ ▷ ⋊ P.
Proof. by unseal; rewrite uPred_naught_eq. Qed.
Lemma naught_ownM : ⋊ uPred_ownM ∅ ⊣⊢ uPred_ownM ∅. (* perhaps quite useless!? *)
Proof.
intros; apply (anti_symm _); first by apply:naught_elim.
unseal; rewrite uPred_naught_eq; split=> n x ? ?.
by cbv [uPred_naught_def uPred_ownM_def uPred_holds].
Qed.
Lemma naught_cmra_valid_1 {A : cmraT} (a : A) : ✓ a ⊢ ⋊ ✓ a.
Proof. by unseal; rewrite uPred_naught_eq. Qed.
Hint Resolve pure_intro.
Hint Resolve or_elim or_intro_l' or_intro_r'.
Hint Resolve and_intro and_elim_l' and_elim_r'.
Hint Immediate True_intro False_elim.
(* Naught derived *)
Hint Resolve naught_mono naught_elim.
Global Instance naught_mono' : Proper ((⊢) ==> (⊢)) (@uPred_naught M).
Proof. intros P Q; apply naught_mono. Qed.
Global Instance naught_flip_mono' :
Proper (flip (⊢) ==> flip (⊢)) (@uPred_naught M).
Proof. intros P Q; apply naught_mono. Qed.
Lemma naught_intro' P Q : (⋊ P ⊢ Q) → ⋊ P ⊢ ⋊ Q.
Proof. intros <-. apply naught_idemp. Qed.
Lemma naught_pure φ : ⋊ ■ φ ⊣⊢ ■ φ.
Proof. apply (anti_symm _); auto using naught_pure_2. Qed.
Lemma naught_forall {A} (Ψ : A → uPred M) : (⋊ ∀ a, Ψ a) ⊣⊢ (∀ a, ⋊ Ψ a).
Proof.
apply (anti_symm _); auto using naught_forall_2.
apply forall_intro=> x. by rewrite (forall_elim x).
Qed.
Lemma naught_exist {A} (Ψ : A → uPred M) : (⋊ ∃ a, Ψ a) ⊣⊢ (∃ a, ⋊ Ψ a).
Proof.
apply (anti_symm _); auto using naught_exist_1.
apply exist_elim=> x. by rewrite (exist_intro x).
Qed.
Lemma naught_and P Q : ⋊ (P ∧ Q) ⊣⊢ ⋊ P ∧ ⋊ Q.
Proof. rewrite !and_alt naught_forall. by apply forall_proper=> -[]. Qed.
Lemma naught_or P Q : ⋊ (P ∨ Q) ⊣⊢ ⋊ P ∨ ⋊ Q.
Proof. rewrite !or_alt naught_exist. by apply exist_proper=> -[]. Qed.
Lemma naught_impl P Q : ⋊ (P → Q) ⊢ ⋊ P → ⋊ Q.
Proof.
apply impl_intro_l; rewrite -naught_and.
apply naught_mono, impl_elim with P; auto.
Qed.
Lemma naught_internal_eq {A:cofeT} (a b : A) : ⋊ (a ≡ b) ⊣⊢ a ≡ b.
Proof.
apply (anti_symm (⊢)); auto using naught_elim.
apply (internal_eq_rewrite a b (λ b, ⋊ (a ≡ b))%I); auto.
{ intros n; solve_proper. }
rewrite -(internal_eq_refl a) naught_pure; auto.
Qed.
Lemma naught_always P : ⋊ □ P ⊣⊢ □ ⋊ P.
Proof.
apply (anti_symm _); auto using naught_always_1, naught_always_2.
Qed.
Lemma naught_and_sep P Q : ⋊ (P ∧ Q) ⊣⊢ ⋊ (P ★ Q).
Proof.
apply (anti_symm (⊢)); last rewrite sep_and; auto using naught_and_sep_1.
Qed.
Lemma naught_and_sep_l' P Q : ⋊ P ∧ Q ⊣⊢ ⋊ P ★ Q.
Proof.
apply (anti_symm (⊢)); last rewrite sep_and; auto using naught_and_sep_l_1.
Qed.
Lemma naught_and_sep_r' P Q : P ∧ ⋊ Q ⊣⊢ P ★ ⋊ Q.
Proof. by rewrite !(comm _ P) naught_and_sep_l'. Qed.
Lemma naught_sep P Q : ⋊ (P ★ Q) ⊣⊢ ⋊ P ★ ⋊ Q.
Proof. by rewrite -naught_and_sep -naught_and_sep_l' naught_and. Qed.
Lemma naught_sep_dup' P : ⋊ P ⊣⊢ ⋊ P ★ ⋊ P.
Proof. by rewrite -naught_sep -naught_and_sep (idemp _). Qed.
Lemma naught_wand P Q : ⋊ (P -★ Q) ⊢ ⋊ P -★ ⋊ Q.
Proof. by apply wand_intro_r; rewrite -naught_sep wand_elim_l. Qed.
Lemma naught_wand_impl P Q : ⋊ (P -★ Q) ⊣⊢ ⋊ (P → Q).
Proof.
apply (anti_symm (⊢)); [|by rewrite -impl_wand].
apply naught_intro', impl_intro_r.
by rewrite naught_and_sep_l' naught_elim wand_elim_l.
Qed.
Lemma naught_entails_l' P Q : (P ⊢ ⋊ Q) → P ⊢ ⋊ Q ★ P.
Proof. intros; rewrite -naught_and_sep_l'; auto. Qed.
Lemma naught_entails_r' P Q : (P ⊢ ⋊ Q) → P ⊢ P ★ ⋊ Q.
Proof. intros; rewrite -naught_and_sep_r'; auto. Qed.
Lemma except_0_naught P : ◇ ⋊ P ⊣⊢ ⋊ ◇ P.
Proof. by rewrite /uPred_except_0 naught_or naught_later naught_pure. Qed.
Lemma naught_cmra_valid {A : cmraT} (a : A) : ⋊ ✓ a ⊣⊢ ✓ a.
Proof.
intros; apply (anti_symm _); first by apply:naught_elim.
apply:naught_cmra_valid_1.
Qed.
(* Plainness *)
Global Instance pure_plain φ : PlainP (■ φ : uPred M)%I.
Proof. by rewrite /PlainP naught_pure. Qed.
Global Instance naught_plain P : PlainP (⋊ P).
Proof. by intros; apply naught_intro'. Qed.
Global Instance and_plain P Q :
PlainP P → PlainP Q → PlainP (P ∧ Q).
Proof. by intros; rewrite /PlainP naught_and; apply and_mono. Qed.
Global Instance or_plain P Q :
PlainP P → PlainP Q → PlainP (P ∨ Q).
Proof. by intros; rewrite /PlainP naught_or; apply or_mono. Qed.
Global Instance sep_plain P Q :
PlainP P → PlainP Q → PlainP (P ★ Q).
Proof. by intros; rewrite /PlainP naught_sep; apply sep_mono. Qed.
Global Instance forall_plain {A} (Ψ : A → uPred M) :
(∀ x, PlainP (Ψ x)) → PlainP (∀ x, Ψ x).
Proof. by intros; rewrite /PlainP naught_forall; apply forall_mono. Qed.
Global Instance exist_plain {A} (Ψ : A → uPred M) :
(∀ x, PlainP (Ψ x)) → PlainP (∃ x, Ψ x).
Proof. by intros; rewrite /PlainP naught_exist; apply exist_mono. Qed.
Global Instance internal_eq_plain {A : cofeT} (a b : A) :
PlainP (a ≡ b : uPred M)%I.
Proof. by intros; rewrite /PlainP naught_internal_eq. Qed.
Global Instance cmra_valid_plain {A : cmraT} (a : A) :
PlainP (✓ a : uPred M)%I.
Proof. by intros; rewrite /PlainP naught_cmra_valid. Qed.
Global Instance later_plain P : PlainP P → PlainP (▷ P).
Proof. by intros; rewrite /PlainP naught_later; apply later_mono. Qed.
Global Instance ownM_plain : PlainP (@uPred_ownM M ∅). (* perhaps quite useless!? *)
Proof. intros. by rewrite /PlainP naught_ownM. Qed.
Global Instance from_option_plain {A} P (Ψ : A → uPred M) (mx : option A) :
(∀ x, PlainP (Ψ x)) → PlainP P → PlainP (from_option Ψ P mx).
Proof. destruct mx; apply _. Qed.
Global Instance except_0_plain P : PlainP P → PlainP (◇ P).
Proof. intros; by rewrite /PlainP -except_0_naught except_0_mono. Qed.
(* Plainness implies persistence *)
Global Instance plain_persistent P : PlainP P → PersistentP P.
Proof.
rewrite /PlainP /PersistentP => HP.
by etrans; last apply naught_implies_always.
Qed.
(* Derived lemmas for plainness *)
Lemma naught_naught P `{!PlainP P} : ⋊ P ⊣⊢ P.
Proof. apply (anti_symm (⊢)); auto. Qed.
Lemma naught_intro P Q `{!PlainP P} : (P ⊢ Q) → P ⊢ ⋊ Q.
Proof. rewrite -(naught_naught P); apply naught_intro'. Qed.
Lemma naught_and_sep_l P Q `{!PlainP P} : P ∧ Q ⊣⊢ P ★ Q.
Proof. by rewrite -(naught_naught P) naught_and_sep_l'. Qed.
Lemma naught_and_sep_r P Q `{!PlainP Q} : P ∧ Q ⊣⊢ P ★ Q.
Proof. by rewrite -(naught_naught Q) naught_and_sep_r'. Qed.
Lemma naught_sep_dup P `{!PlainP P} : P ⊣⊢ P ★ P.
Proof. by rewrite -(naught_naught P) -naught_sep_dup'. Qed.
Lemma naught_entails_l P Q `{!PlainP Q} : (P ⊢ Q) → P ⊢ Q ★ P.
Proof. by rewrite -(naught_naught Q); apply naught_entails_l'. Qed.
Lemma naught_entails_r P Q `{!PlainP Q} : (P ⊢ Q) → P ⊢ P ★ Q.
Proof. by rewrite -(naught_naught Q); apply naught_entails_r'. Qed.
(* important for proving properties in the model. *)
Lemma naught_irrel' P `{HPl: !PlainP P} n x y : ✓{n} x → ✓{n} y → P n x → P n y.
Proof.
intros Hx Hy Hnx; apply HPl in Hnx; last done.
apply naught_elim; first done. eapply naught_irrel; eauto.
Qed.
Lemma naught_keeps_premises' P Q `{HPl: !PlainP Q} : (P ⊢ Q) → P ⊢ P ★ Q.
Proof.
iIntros (HPQ) "HP".
assert (HPQ': P ⊢ ⋊ Q) by (by iIntros "HP"; iApply HPl; iApply HPQ).
iPoseProof (naught_keeps_premises _ _ HPQ' with "HP") as "HPQ".
iDestruct "HPQ" as "[HP HQ]"; iSplitL "HP"; first trivial.
by iApply naught_elim.
Qed.
Global Instance iter_Plain_Plain P `{Hpl: !PlainP P} n f :
(∀ Q, PlainP Q -> PlainP (f Q)) →
PlainP (Nat.iter n f P).
Proof.
revert P Hpl. induction n => P Hpl; first done.
rewrite Nat_iter_S_r; auto.
Qed.
End naught_properties.
Lemma bupd_extract_plain {M : ucmraT} (P Q : uPred M) `{!PlainP P} :
(Q ⊢ P) → (|==> Q) ⊢ P.
Proof.
intros HQP; unseal; split=> n x Hnx Hbp.
destruct (Hbp n ∅ (le_n _)) as [z [Hz1 Hz2]]; first by rewrite right_id.
revert Hz1; rewrite right_id=> Hz1.
apply HQP in Hz2; trivial. eauto using naught_irrel', ucmra_unit_validN.
Qed.
Section naught_proofmode_classes.
Context {M : ucmraT}.
Implicit Types P Q R : uPred M.
(* FromAssumption *)
Global Instance from_assumption_naught_l p P Q :
FromAssumption p P Q → FromAssumption p (⋊ P) Q.
Proof. rewrite /FromAssumption=><-. by rewrite naught_elim. Qed.
(* IntoPersistentP *)
Global Instance into_persistentP_naught P : IntoPersistentP (⋊ P) P | 1.
Proof. apply naught_implies_always. Qed.
Global Instance into_PlainP_persistent P :
PlainP P → IntoPersistentP P P | 100.
Proof.
rewrite /IntoPersistentP => HP; by etrans; [|apply naught_implies_always].
Qed.
(* FromAnd *)
Global Instance from_and_sep_Plain_l P1 P2 :
PlainP P1 → FromAnd (P1 ★ P2) P1 P2 | 9.
Proof. intros. by rewrite /FromAnd naught_and_sep_l. Qed.
Global Instance from_and_sep_Plain_r P1 P2 :
PlainP P2 → FromAnd (P1 ★ P2) P1 P2 | 10.
Proof. intros. by rewrite /FromAnd naught_and_sep_r. Qed.
Global Instance from_and_naught P Q1 Q2 :
FromAnd P Q1 Q2 → FromAnd (⋊ P) (⋊ Q1) (⋊ Q2).
Proof. rewrite /FromAnd=> <-. by rewrite naught_and. Qed.
(* FromSep *)
Global Instance from_sep_naught P Q1 Q2 :
FromSep P Q1 Q2 → FromSep (⋊ P) (⋊ Q1) (⋊ Q2).
Proof. rewrite /FromSep=> <-. by rewrite naught_sep. Qed.
(* IntoAnd *)
Global Instance into_and_and_plain_l P Q :
PlainP P → IntoAnd false (P ∧ Q) P Q.
Proof. intros; by rewrite /IntoAnd /= naught_and_sep_l. Qed.
Global Instance into_and_and_plain_r P Q :
PlainP Q → IntoAnd false (P ∧ Q) P Q.
Proof. intros; by rewrite /IntoAnd /= always_and_sep_r. Qed.
Global Instance into_exist_naught {A} P (Φ : A → uPred M) :
IntoExist P Φ → IntoExist (⋊ P) (λ a, ⋊ (Φ a))%I.
Proof. rewrite /IntoExist=> HP. by rewrite HP naught_exist. Qed.
(* ElimModal *)
Global Instance elim_modal_naught P Q : ElimModal (⋊ P) P Q Q.
Proof. intros. by rewrite /ElimModal naught_elim wand_elim_r. Qed.
End naught_proofmode_classes.