Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 3a0dddc

Browse files
committed
feat(algebra/order_functions): (min|max)_eq_iff (#8911)
Co-authored-by: Yakov Pechersky <pechersky@users.noreply.github.com>
1 parent 39cea43 commit 3a0dddc

File tree

1 file changed

+13
-0
lines changed

1 file changed

+13
-0
lines changed

src/algebra/order_functions.lean

Lines changed: 13 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -48,6 +48,19 @@ lemma min_le_max : min a b ≤ max a b := le_trans (min_le_left a b) (le_max_lef
4848
@[simp] lemma max_eq_left_iff : max a b = a ↔ b ≤ a := sup_eq_left
4949
@[simp] lemma max_eq_right_iff : max a b = b ↔ a ≤ b := sup_eq_right
5050

51+
lemma min_eq_iff : min a b = c ↔ a = c ∧ a ≤ b ∨ b = c ∧ b ≤ a :=
52+
begin
53+
split,
54+
{ intro h,
55+
refine or.imp (λ h', _) (λ h', _) (le_total a b);
56+
exact ⟨by simpa [h'] using h, h'⟩ },
57+
{ rintro (⟨rfl, h⟩|⟨rfl, h⟩);
58+
simp [h] }
59+
end
60+
61+
lemma max_eq_iff : max a b = c ↔ a = c ∧ b ≤ a ∨ b = c ∧ a ≤ b :=
62+
@min_eq_iff (order_dual α) _ a b c
63+
5164
/-- An instance asserting that `max a a = a` -/
5265
instance max_idem : is_idempotent α max := by apply_instance -- short-circuit type class inference
5366

0 commit comments

Comments
 (0)