From b3891ed55f15504d80d2ce02d5da32220f0eadb7 Mon Sep 17 00:00:00 2001 From: Robbert Krebbers <mail@robbertkrebbers.nl> Date: Wed, 31 Aug 2016 21:07:40 +0200 Subject: [PATCH] Group iff properties in uPred. --- algebra/upred.v | 38 +++++++++++++++++++------------------- 1 file changed, 19 insertions(+), 19 deletions(-) diff --git a/algebra/upred.v b/algebra/upred.v index 773186211..5d0701852 100644 --- a/algebra/upred.v +++ b/algebra/upred.v @@ -563,11 +563,6 @@ Proof. Qed. (* Derived logical stuff *) -Global Instance iff_ne n : Proper (dist n ==> dist n ==> dist n) (@uPred_iff M). -Proof. unfold uPred_iff; solve_proper. Qed. -Global Instance iff_proper : - Proper ((⊣⊢) ==> (⊣⊢) ==> (⊣⊢)) (@uPred_iff M) := ne_proper_2 _. - Lemma False_elim P : False ⊢ P. Proof. by apply (pure_elim False). Qed. Lemma True_intro P : P ⊢ True. @@ -606,16 +601,6 @@ Proof. intros HPQ; apply impl_elim with P; rewrite -?HPQ; auto. Qed. Lemma entails_impl P Q : (P ⊢ Q) → True ⊢ P → Q. Proof. auto using impl_intro_l. Qed. -Lemma iff_refl Q P : Q ⊢ P ↔ P. -Proof. rewrite /uPred_iff; apply and_intro; apply impl_intro_l; auto. Qed. -Lemma iff_equiv P Q : (True ⊢ P ↔ Q) → (P ⊣⊢ Q). -Proof. - intros HPQ; apply (anti_symm (⊢)); - apply impl_entails; rewrite HPQ /uPred_iff; auto. -Qed. -Lemma equiv_iff P Q : (P ⊣⊢ Q) → True ⊢ P ↔ Q. -Proof. intros ->; apply iff_refl. Qed. - Lemma and_mono P P' Q Q' : (P ⊢ Q) → (P' ⊢ Q') → P ∧ P' ⊢ Q ∧ Q'. Proof. auto. Qed. Lemma and_mono_l P P' Q : (P ⊢ Q) → P ∧ P' ⊢ Q ∧ P'. @@ -783,10 +768,6 @@ Lemma equiv_eq {A : cofeT} P (a b : A) : a ≡ b → P ⊢ a ≡ b. Proof. by intros ->. Qed. Lemma eq_sym {A : cofeT} (a b : A) : a ≡ b ⊢ b ≡ a. Proof. apply (eq_rewrite a b (λ b, b ≡ a)%I); auto. solve_proper. Qed. -Lemma eq_iff P Q : P ≡ Q ⊢ P ↔ Q. -Proof. - apply (eq_rewrite P Q (λ Q, P ↔ Q))%I; first solve_proper; auto using iff_refl. -Qed. Lemma pure_alt φ : ■φ ⊣⊢ ∃ _ : φ, True. Proof. @@ -805,6 +786,25 @@ Proof. apply or_elim. by rewrite -(exist_intro true). by rewrite -(exist_intro false). Qed. +Global Instance iff_ne n : Proper (dist n ==> dist n ==> dist n) (@uPred_iff M). +Proof. unfold uPred_iff; solve_proper. Qed. +Global Instance iff_proper : + Proper ((⊣⊢) ==> (⊣⊢) ==> (⊣⊢)) (@uPred_iff M) := ne_proper_2 _. + +Lemma iff_refl Q P : Q ⊢ P ↔ P. +Proof. rewrite /uPred_iff; apply and_intro; apply impl_intro_l; auto. Qed. +Lemma iff_equiv P Q : (True ⊢ P ↔ Q) → (P ⊣⊢ Q). +Proof. + intros HPQ; apply (anti_symm (⊢)); + apply impl_entails; rewrite HPQ /uPred_iff; auto. +Qed. +Lemma equiv_iff P Q : (P ⊣⊢ Q) → True ⊢ P ↔ Q. +Proof. intros ->; apply iff_refl. Qed. +Lemma eq_iff P Q : P ≡ Q ⊢ P ↔ Q. +Proof. + apply (eq_rewrite P Q (λ Q, P ↔ Q))%I; first solve_proper; auto using iff_refl. +Qed. + (* BI connectives *) Lemma sep_mono P P' Q Q' : (P ⊢ Q) → (P' ⊢ Q') → P ★ P' ⊢ Q ★ Q'. Proof. -- GitLab