diff --git a/theories/base_logic/derived.v b/theories/base_logic/derived.v index 80f7d2e175a13de4f60b7c4423b80cb5b9a395f5..11772f13cc2eeb9022bba7ff36d1053a2553eee8 100644 --- a/theories/base_logic/derived.v +++ b/theories/base_logic/derived.v @@ -437,6 +437,12 @@ Proof. intros P; apply (anti_symm _); auto. Qed. Global Instance False_sep : RightAbsorb (⊣⊢) False%I (@uPred_sep M). Proof. intros P; apply (anti_symm _); auto. Qed. +Lemma entails_equiv_and P Q : (P ⊣⊢ Q ∧ P) ↔ (P ⊢ Q). +Proof. + split; first by intros ->; apply uPred.and_elim_l. intros PQ. + apply uPred.equiv_spec; split; + [by apply uPred.and_intro|apply uPred.and_elim_r]. +Qed. Lemma sep_and_l P Q R : P ∗ (Q ∧ R) ⊢ (P ∗ Q) ∧ (P ∗ R). Proof. auto. Qed. Lemma sep_and_r P Q R : (P ∧ Q) ∗ R ⊢ (P ∗ R) ∧ (Q ∗ R).