From 611995eefd234c8c68392e56e52230d73052e23d Mon Sep 17 00:00:00 2001 From: Robbert Krebbers <mail@robbertkrebbers.nl> Date: Tue, 27 Dec 2016 12:44:26 +0100 Subject: [PATCH] Tweak uPred core. --- theories/base_logic/lib/core.v | 20 ++++++++------------ 1 file changed, 8 insertions(+), 12 deletions(-) diff --git a/theories/base_logic/lib/core.v b/theories/base_logic/lib/core.v index ccabc7ecf..219e758b6 100644 --- a/theories/base_logic/lib/core.v +++ b/theories/base_logic/lib/core.v @@ -8,33 +8,29 @@ Import uPred. Definition coreP {M : ucmraT} (P : uPred M) : uPred M := (∀ `(!PersistentP Q, P ⊢ Q), Q)%I. +Instance: Params (@coreP) 1. +Typeclasses Opaque coreP. Section core. Context {M : ucmraT}. Implicit Types P Q : uPred M. Lemma coreP_intro P : P -∗ coreP P. - Proof. iIntros "HP". iIntros (Q HQ HPQ). by iApply HPQ. Qed. + Proof. iIntros "HP". by iIntros (Q HQ ->). Qed. Global Instance coreP_persistent P : PersistentP (coreP P). - Proof. - iIntros "HCP". iApply always_forall. iIntros (Q). - iApply always_forall. iIntros (HQ). iApply always_forall. - iIntros (HPQ). iApply HQ. unshelve iApply ("HCP" $! Q). done. - Qed. + Proof. rewrite /coreP. apply _. Qed. Global Instance coreP_mono : Proper ((⊢) ==> (⊢)) (@coreP M). Proof. - unfold coreP. iIntros (?? ENT) "H *". unshelve iApply ("H" $! Q). by etrans. + rewrite /coreP. iIntros (P P' ?) "H"; iIntros (Q ??). + 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. + Proof. intros P Q. rewrite !equiv_spec=>-[??]. by split; apply coreP_mono. Qed. Lemma coreP_elim P : PersistentP P → coreP P -∗ P. - Proof. - iIntros (?) "HCP". unshelve iApply ("HCP" $! P). iIntros "P". done. - Qed. + Proof. rewrite /coreP. iIntros (?) "HCP". unshelve iApply ("HCP" $! P); auto. Qed. End core. -Global Opaque coreP. -- GitLab