From d4a687a83e303316ebc0c9e9824360f9c6296bfe Mon Sep 17 00:00:00 2001 From: Robbert Krebbers <mail@robbertkrebbers.nl> Date: Sun, 20 Nov 2016 15:33:51 +0100 Subject: [PATCH] =?UTF-8?q?Prove=20(=E2=96=A0=20=CF=86=20=E2=86=92=20P)=20?= =?UTF-8?q?=E2=8A=A3=E2=8A=A2=20(=E2=88=80=20=5F=20:=20=CF=86,=20P).?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- base_logic/derived.v | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/base_logic/derived.v b/base_logic/derived.v index 95150ab85..8a3b26acf 100644 --- a/base_logic/derived.v +++ b/base_logic/derived.v @@ -265,6 +265,12 @@ Proof. by intros ->. Qed. Lemma internal_eq_sym {A : cofeT} (a b : A) : a ≡ b ⊢ b ≡ a. Proof. apply (internal_eq_rewrite a b (λ b, b ≡ a)%I); auto. solve_proper. Qed. +Lemma pure_impl_forall φ P : (■φ → P) ⊣⊢ (∀ _ : φ, P). +Proof. + apply (anti_symm _). + - apply forall_intro=> ?. by rewrite pure_equiv // left_id. + - apply impl_intro_l, pure_elim_l=> Hφ. by rewrite (forall_elim Hφ). +Qed. Lemma pure_alt φ : ■φ ⊣⊢ ∃ _ : φ, True. Proof. apply (anti_symm _). -- GitLab