diff --git a/theories/base_logic/lib/core.v b/theories/base_logic/lib/core.v index ba5de9b212160f408d7fa0d49f14e3e74a26d227..fe6a0ed7de8cc79324f354c4961dcf746af17bda 100644 --- a/theories/base_logic/lib/core.v +++ b/theories/base_logic/lib/core.v @@ -7,7 +7,7 @@ Import uPred. in the shallow embedding. *) Definition coreP {M : ucmraT} (P : uPred M) : uPred M := - (∀ (Q : uPred M) `(!PersistentP Q, !P -∗ Q), Q)%I. + (∀ `(!PersistentP Q, P ⊢ Q), Q)%I. Section core. Context {M : ucmraT}. @@ -23,11 +23,18 @@ Section core. iIntros (HPQ). iApply HQ. unshelve iApply ("HCP" $! Q). done. Qed. + Global Instance coreP_mono : Proper ((⊢) ==> (⊢)) (@coreP M). + Proof. + unfold coreP. iIntros (?? ENT) "H *". unshelve iApply ("H" $! Q). by etrans. + Qed. + + Global Instance coreP_proper : Proper ((⊣⊢) ==> (⊣⊢)) (@coreP M). + Proof. intros ??. rewrite !equiv_spec=>-[A B]. split; rewrite ?A // ?B //. Qed. + Lemma coreP_elim P : PersistentP P → coreP P -∗ P. Proof. iIntros (?) "HCP". unshelve iApply ("HCP" $! P). iIntros "P". done. Qed. End core. - - +Global Opaque coreP. \ No newline at end of file