Skip to content

Commit

Permalink
feat(logic/basic): add two simp lemmas (#8148)
Browse files Browse the repository at this point in the history
  • Loading branch information
fpvandoorn committed Jul 1, 2021
1 parent 8986c4f commit 6e0d2d3
Showing 1 changed file with 6 additions and 0 deletions.
6 changes: 6 additions & 0 deletions src/logic/basic.lean
Expand Up @@ -384,6 +384,12 @@ iff.intro and.right (λ hb, ⟨h hb, hb⟩)
@[simp] theorem and_iff_right_iff_imp {a b : Prop} : ((a ∧ b) ↔ b) ↔ (b → a) :=
⟨λ h ha, (h.2 ha).1, and_iff_right_of_imp⟩

@[simp] lemma iff_self_and {p q : Prop} : (p ↔ p ∧ q) ↔ (p → q) :=
by rw [@iff.comm p, and_iff_left_iff_imp]

@[simp] lemma iff_and_self {p q : Prop} : (p ↔ q ∧ p) ↔ (p → q) :=
by rw [and_comm, iff_self_and]

@[simp] lemma and.congr_right_iff : (a ∧ b ↔ a ∧ c) ↔ (a → (b ↔ c)) :=
⟨λ h ha, by simp [ha] at h; exact h, and_congr_right⟩

Expand Down

0 comments on commit 6e0d2d3

Please sign in to comment.