diff --git a/CHANGELOG.md b/CHANGELOG.md index 81f2ec071944eb05dc9646b4a85fd33485903f48..6583178012847822c8a86b0a4de62a50ec0a4bb9 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -159,6 +159,8 @@ Changes in Coq: fractional camera (`frac_auth`) with unbounded fractions. * Changed `frac_auth` notation from `â—!`/`â—¯!` to `â—F`/`â—¯F`. sed script: `s/â—¯!/â—¯F/g; s/â—!/â—F/g;`. +* Lemma `prop_ext` works in both directions; its default direction is the + opposite of what it used to be. * Rename `C` suffixes into `O` since we no longer use COFEs but OFEs. Also rename `ofe_fun` into `discrete_fun` and the corresponding notation `-c>` into `-d>`. The renaming can be automatically done using the following script (on diff --git a/theories/base_logic/bi.v b/theories/base_logic/bi.v index 1d3758f02a1bceae2473ca4906726732970aa97e..11a1a6a4032883379eb8edc3a8bc861a95b81c69 100644 --- a/theories/base_logic/bi.v +++ b/theories/base_logic/bi.v @@ -121,7 +121,7 @@ Proof. etrans; first exact: sep_comm'. apply sep_mono; last done. exact: pure_intro. - - exact: prop_ext. + - exact: prop_ext_2. - exact: later_plainly_1. - exact: later_plainly_2. Qed. diff --git a/theories/base_logic/upred.v b/theories/base_logic/upred.v index e1b07ab63f6032f6d81c6813174ba5388c1661fc..b861d2284d29cffe1db9c0363b6cad57bab70679 100644 --- a/theories/base_logic/upred.v +++ b/theories/base_logic/upred.v @@ -617,7 +617,7 @@ Proof. by unseal. Qed. Lemma plainly_exist_1 {A} (Ψ : A → uPred M) : (■∃ a, Ψ a) ⊢ (∃ a, ■Ψ a). Proof. by unseal. Qed. -Lemma prop_ext P Q : â– ((P -∗ Q) ∧ (Q -∗ P)) ⊢ P ≡ Q. +Lemma prop_ext_2 P Q : â– ((P -∗ Q) ∧ (Q -∗ P)) ⊢ P ≡ Q. Proof. unseal; split=> n x ? /= HPQ. split=> n' x' ??. move: HPQ=> [] /(_ n' x'); rewrite !left_id=> ?. diff --git a/theories/bi/monpred.v b/theories/bi/monpred.v index 9e9fe6a05908dac779199b51d81746b852a1383c..0cb22093cf7bf68f5940b473c00454662ed26622 100644 --- a/theories/bi/monpred.v +++ b/theories/bi/monpred.v @@ -923,7 +923,7 @@ Proof. - intros P. split=> i /=. apply bi.forall_intro=>_. by apply plainly_emp_intro. - intros P Q. split=> i. apply bi.sep_elim_l, _. - intros P Q. split=> i /=. rewrite (monPred_equivI P Q). f_equiv=> j. - by rewrite -prop_ext !(bi.forall_elim j) !bi.pure_True // !bi.True_impl. + by rewrite prop_ext !(bi.forall_elim j) !bi.pure_True // !bi.True_impl. - intros P. split=> i /=. rewrite bi.later_forall. f_equiv=> j. by rewrite -later_plainly_1. - intros P. split=> i /=. diff --git a/theories/bi/plainly.v b/theories/bi/plainly.v index b28934bcb2ee6d1e4d2bbbd40f675c5b4579b482..c270a8422591c468b67ca6010b9ee00be056be8e 100644 --- a/theories/bi/plainly.v +++ b/theories/bi/plainly.v @@ -29,7 +29,7 @@ Record BiPlainlyMixin (PROP : sbi) `(Plainly PROP) := { bi_plainly_mixin_plainly_emp_intro (P : PROP) : P ⊢ â– emp; bi_plainly_mixin_plainly_absorb (P Q : PROP) : â– P ∗ Q ⊢ â– P; - bi_plainly_mixin_prop_ext (P Q : PROP) : â– ((P -∗ Q) ∧ (Q -∗ P)) ⊢ P ≡ Q; + bi_plainly_mixin_prop_ext_2 (P Q : PROP) : â– ((P -∗ Q) ∧ (Q -∗ P)) ⊢ P ≡ Q; bi_plainly_mixin_later_plainly_1 (P : PROP) : â–· â– P ⊢ â– â–· P; bi_plainly_mixin_later_plainly_2 (P : PROP) : â– â–· P ⊢ â–· â– P; @@ -74,8 +74,8 @@ Section plainly_laws. Lemma plainly_emp_intro P : P ⊢ â– emp. Proof. eapply bi_plainly_mixin_plainly_emp_intro, bi_plainly_mixin. Qed. - Lemma prop_ext P Q : â– ((P -∗ Q) ∧ (Q -∗ P)) ⊢ P ≡ Q. - Proof. eapply bi_plainly_mixin_prop_ext, bi_plainly_mixin. Qed. + Lemma prop_ext_2 P Q : â– ((P -∗ Q) ∧ (Q -∗ P)) ⊢ P ≡ Q. + Proof. eapply bi_plainly_mixin_prop_ext_2, bi_plainly_mixin. Qed. Lemma later_plainly_1 P : â–· â– P ⊢ â– (â–· P). Proof. eapply bi_plainly_mixin_later_plainly_1, bi_plainly_mixin. Qed. @@ -467,10 +467,19 @@ Proof. rewrite -(internal_eq_refl True%I a) plainly_pure; auto. Qed. +Lemma prop_ext P Q : P ≡ Q ⊣⊢ â– (P ∗-∗ Q). +Proof. + apply (anti_symm (⊢)); last exact: prop_ext_2. + apply (internal_eq_rewrite' P Q (λ Q, â– (P ∗-∗ Q))%I); + [ solve_proper | done | ]. + rewrite (plainly_emp_intro (P ≡ Q)%I). + apply plainly_mono, wand_iff_refl. +Qed. + Lemma plainly_alt P : â– P ⊣⊢ <affine> P ≡ emp. Proof. rewrite -plainly_affinely_elim. apply (anti_symm (⊢)). - - rewrite -prop_ext. apply plainly_mono, and_intro; apply wand_intro_l. + - rewrite prop_ext. apply plainly_mono, and_intro; apply wand_intro_l. + by rewrite affinely_elim_emp left_id. + by rewrite left_id. - rewrite internal_eq_sym (internal_eq_rewrite _ _ plainly). @@ -480,7 +489,7 @@ Qed. Lemma plainly_alt_absorbing P `{!Absorbing P} : â– P ⊣⊢ P ≡ True. Proof. apply (anti_symm (⊢)). - - rewrite -prop_ext. apply plainly_mono, and_intro; apply wand_intro_l; auto. + - rewrite prop_ext. apply plainly_mono, and_intro; apply wand_intro_l; auto. - rewrite internal_eq_sym (internal_eq_rewrite _ _ plainly). by rewrite plainly_pure True_impl. Qed. @@ -488,7 +497,7 @@ Qed. Lemma plainly_True_alt P : â– (True -∗ P) ⊣⊢ P ≡ True. Proof. apply (anti_symm (⊢)). - - rewrite -prop_ext. apply plainly_mono, and_intro; apply wand_intro_l; auto. + - rewrite prop_ext. apply plainly_mono, and_intro; apply wand_intro_l; auto. by rewrite wand_elim_r. - rewrite internal_eq_sym (internal_eq_rewrite _ _ (λ Q, â– (True -∗ Q))%I ltac:(shelve)); last solve_proper.