diff --git a/theories/bi/derived_laws.v b/theories/bi/derived_laws.v index e9191e1348cd180d4e228335ccadcbe1b41cd4a8..b8170151aca1b3a9cff8f70f0c7696db4d8f36fb 100644 --- a/theories/bi/derived_laws.v +++ b/theories/bi/derived_laws.v @@ -249,6 +249,25 @@ Proof. - rewrite -(exist_intro ()). done. Qed. +Lemma exist_exist {A B} (Ψ : A → B → PROP) : + (∃ x y, Ψ x y) ⊣⊢ (∃ y x, Ψ x y). +Proof. + apply (anti_symm (⊢)); + do 2 (apply exist_elim=>?); rewrite -2!exist_intro; eauto. +Qed. +Lemma forall_forall {A B} (Ψ : A → B → PROP) : + (∀ x y, Ψ x y) ⊣⊢ (∀ y x, Ψ x y). +Proof. + apply (anti_symm (⊢)); + do 2 (apply forall_intro=>?); rewrite 2!forall_elim; eauto. +Qed. +Lemma exist_forall {A B} (Ψ : A → B → PROP) : + (∃ x, ∀ y, Ψ x y) ⊢ (∀ y, ∃ x, Ψ x y). +Proof. + apply forall_intro=>?. apply exist_elim=>?. + rewrite -exist_intro forall_elim ; eauto. +Qed. + Lemma impl_curry P Q R : (P → Q → R) ⊣⊢ (P ∧ Q → R). Proof. apply (anti_symm _).