@@ -231,6 +231,17 @@ theorem xor_comm (a b) : xor a b = xor b a := by simp [xor, and_comm, or_comm]
231
231
232
232
/-! ### Declarations about `and` -/
233
233
234
+ theorem and_symm_right (a b : α) (p : Prop ) : p ∧ a = b <-> p ∧ b = a :=
235
+ Iff.intro
236
+ (fun ⟨hp, a_eq_b⟩ => ⟨hp, a_eq_b.symm⟩)
237
+ (fun ⟨hp, b_eq_a⟩ => ⟨hp, b_eq_a.symm⟩)
238
+
239
+ theorem and_symm_left (a b : α) (p : Prop ) : a = b ∧ p <-> b = a ∧ p :=
240
+ Iff.intro
241
+ (fun ⟨a_eq_b, hp⟩ => ⟨a_eq_b.symm, hp⟩)
242
+ (fun ⟨b_eq_a, hp⟩ => ⟨b_eq_a.symm, hp⟩)
243
+
244
+
234
245
theorem and_congr_left (h : c → (a ↔ b)) : a ∧ c ↔ b ∧ c :=
235
246
And.comm.trans $ (and_congr_right h).trans And.comm
236
247
@@ -689,6 +700,23 @@ by simp [And.comm]
689
700
@[simp] theorem exists_eq_left' {p : α → Prop } {a' : α} : (∃ a, a' = a ∧ p a) ↔ p a' :=
690
701
by simp [@eq_comm _ a']
691
702
703
+ @[simp] theorem exists_eq_right_right {α : Sort _} {p : α → Prop } {b : Prop } {a' : α} :
704
+ (∃ (a : α), p a ∧ b ∧ a = a') ↔ p a' ∧ b :=
705
+ Iff.intro
706
+ (fun ⟨_, ⟨p_a, hb, a_eq_a'⟩⟩ => And.intro (a_eq_a' ▸ p_a) hb)
707
+ (fun ⟨p_a', hb⟩ => Exists.intro a' ⟨p_a', hb, (rfl : a' = a')⟩)
708
+
709
+ @[simp] theorem exists_eq_right_right' {α : Sort _} {p : α → Prop } {b : Prop } {a' : α} :
710
+ (∃ (a : α), p a ∧ b ∧ a' = a) ↔ p a' ∧ b :=
711
+ Iff.intro
712
+ (fun ⟨_, ⟨p_a, hb, a_eq_a'⟩⟩ => And.intro (a_eq_a' ▸ p_a) hb)
713
+ (fun ⟨p_a', hb⟩ => Exists.intro a' ⟨p_a', hb, (rfl : a' = a')⟩)
714
+
715
+
716
+ @[simp]
717
+ theorem exists_prop {p q : Prop } : (∃ h : p, q) ↔ p ∧ q :=
718
+ Iff.intro (fun ⟨hp, hq⟩ => And.intro hp hq) (fun ⟨hp, hq⟩ => Exists.intro hp hq)
719
+
692
720
@[simp] theorem exists_apply_eq_apply {α β : Type _} (f : α → β) (a' : α) : ∃ a, f a = f a' :=
693
721
⟨a', rfl⟩
694
722
0 commit comments