Skip to content

Commit 9cc09f6

Browse files
committed
chore(WithTop): less abuse of defeq to Option _ (#9986)
Also reuse proofs here and there.
1 parent 15e555e commit 9cc09f6

File tree

2 files changed

+27
-43
lines changed

2 files changed

+27
-43
lines changed

Mathlib/Algebra/Order/Monoid/WithTop.lean

Lines changed: 8 additions & 18 deletions
Original file line numberDiff line numberDiff line change
@@ -155,22 +155,19 @@ theorem add_lt_top [LT α] {a b : WithTop α} : a + b < ⊤ ↔ a < ⊤ ∧ b <
155155

156156
theorem add_eq_coe :
157157
∀ {a b : WithTop α} {c : α}, a + b = c ↔ ∃ a' b' : α, ↑a' = a ∧ ↑b' = b ∧ a' + b' = c
158-
| none, b, c => by simp [none_eq_top]
159-
| Option.some a, none, c => by simp [none_eq_top]
160-
| Option.some a, Option.some b, c =>
161-
by simp only [some_eq_coe, ← coe_add, coe_eq_coe, exists_and_left, exists_eq_left]
158+
| ⊤, b, c => by simp
159+
| some a, ⊤, c => by simp
160+
| some a, some b, c => by norm_cast; simp
162161
#align with_top.add_eq_coe WithTop.add_eq_coe
163162

164163
-- Porting note: simp can already prove this.
165164
-- @[simp]
166-
theorem add_coe_eq_top_iff {x : WithTop α} {y : α} : x + y = ⊤ ↔ x = ⊤ := by
167-
induction x using WithTop.recTopCoe <;> simp [← coe_add]
165+
theorem add_coe_eq_top_iff {x : WithTop α} {y : α} : x + y = ⊤ ↔ x = ⊤ := by simp
168166
#align with_top.add_coe_eq_top_iff WithTop.add_coe_eq_top_iff
169167

170168
-- Porting note: simp can already prove this.
171169
-- @[simp]
172-
theorem coe_add_eq_top_iff {y : WithTop α} : ↑x + y = ⊤ ↔ y = ⊤ := by
173-
induction y using WithTop.recTopCoe <;> simp [← coe_add]
170+
theorem coe_add_eq_top_iff {y : WithTop α} : ↑x + y = ⊤ ↔ y = ⊤ := by simp
174171
#align with_top.coe_add_eq_top_iff WithTop.coe_add_eq_top_iff
175172

176173
theorem add_right_cancel_iff [IsRightCancelAdd α] (ha : a ≠ ⊤) : b + a = c + a ↔ b = c := by
@@ -364,8 +361,7 @@ instance addMonoidWithOne : AddMonoidWithOne (WithTop α) :=
364361
rw [Nat.cast_zero, WithTop.coe_zero],
365362
natCast_succ := fun n => by
366363
simp only -- Porting note: Had to add this...?
367-
rw [Nat.cast_add_one, WithTop.coe_add, WithTop.coe_one]
368-
}
364+
rw [Nat.cast_add_one, WithTop.coe_add, WithTop.coe_one] }
369365

370366
@[simp, norm_cast] lemma coe_nat (n : ℕ) : ((n : α) : WithTop α) = n := rfl
371367
#align with_top.coe_nat WithTop.coe_nat
@@ -385,14 +381,8 @@ instance charZero [AddMonoidWithOne α] [CharZero α] : CharZero (WithTop α) :=
385381
instance addCommMonoidWithOne [AddCommMonoidWithOne α] : AddCommMonoidWithOne (WithTop α) :=
386382
{ WithTop.addMonoidWithOne, WithTop.addCommMonoid with }
387383

388-
instance orderedAddCommMonoid [OrderedAddCommMonoid α] : OrderedAddCommMonoid (WithTop α) :=
389-
{ WithTop.partialOrder, WithTop.addCommMonoid with
390-
add_le_add_left := by
391-
rintro a b h (_ | c); · simp [none_eq_top]
392-
rcases b with (_ | b); · simp [none_eq_top]
393-
rcases le_coe_iff.1 h with ⟨a, rfl, _⟩
394-
simp only [some_eq_coe, ← coe_add, coe_le_coe] at h ⊢
395-
exact add_le_add_left h c }
384+
instance orderedAddCommMonoid [OrderedAddCommMonoid α] : OrderedAddCommMonoid (WithTop α) where
385+
add_le_add_left _ _ := add_le_add_left
396386

397387
instance linearOrderedAddCommMonoidWithTop [LinearOrderedAddCommMonoid α] :
398388
LinearOrderedAddCommMonoidWithTop (WithTop α) :=

Mathlib/Order/WithBot.lean

Lines changed: 19 additions & 25 deletions
Original file line numberDiff line numberDiff line change
@@ -96,8 +96,8 @@ theorem coe_ne_bot : (a : WithBot α) ≠ ⊥ :=
9696
/-- Recursor for `WithBot` using the preferred forms `⊥` and `↑a`. -/
9797
@[elab_as_elim]
9898
def recBotCoe {C : WithBot α → Sort*} (bot : C ⊥) (coe : ∀ a : α, C a) : ∀ n : WithBot α, C n
99-
| none => bot
100-
| Option.some a => coe a
99+
| => bot
100+
| (a : α) => coe a
101101
#align with_bot.rec_bot_coe WithBot.recBotCoe
102102

103103
@[simp]
@@ -239,18 +239,18 @@ theorem coe_le : ∀ {o : Option α}, b ∈ o → ((a : WithBot α) ≤ o ↔ a
239239
#align with_bot.coe_le WithBot.coe_le
240240

241241
theorem coe_le_iff : ∀ {x : WithBot α}, (a : WithBot α) ≤ x ↔ ∃ b : α, x = b ∧ a ≤ b
242-
| Option.some x => by simp [some_eq_coe]
243-
| none => iff_of_false (not_coe_le_bot _) <| by simp [none_eq_bot]
242+
| (x : α) => by simp
243+
| => iff_of_false (not_coe_le_bot _) <| by simp
244244
#align with_bot.coe_le_iff WithBot.coe_le_iff
245245

246246
theorem le_coe_iff : ∀ {x : WithBot α}, x ≤ b ↔ ∀ a : α, x = ↑a → a ≤ b
247-
| Option.some b => by simp [some_eq_coe, coe_eq_coe]
248-
| none => by simp [none_eq_bot]
247+
| (b : α) => by simp
248+
| => by simp
249249
#align with_bot.le_coe_iff WithBot.le_coe_iff
250250

251251
protected theorem _root_.IsMax.withBot (h : IsMax a) : IsMax (a : WithBot α)
252-
| none, _ => bot_le
253-
| Option.some _, hb => some_le_some.2 <| h <| some_le_some.1 hb
252+
| , _ => bot_le
253+
| (_ : α), hb => some_le_some.2 <| h <| some_le_some.1 hb
254254
#align is_max.with_bot IsMax.withBot
255255

256256
theorem le_unbot_iff {a : α} {b : WithBot α} (h : b ≠ ⊥) :
@@ -305,13 +305,13 @@ theorem not_lt_none (a : WithBot α) : ¬@LT.lt (WithBot α) _ a none :=
305305
#align with_bot.not_lt_none WithBot.not_lt_none
306306

307307
theorem lt_iff_exists_coe : ∀ {a b : WithBot α}, a < b ↔ ∃ p : α, b = p ∧ a < p
308-
| a, Option.some b => by simp [some_eq_coe, coe_eq_coe]
309-
| a, none => iff_of_false (not_lt_none _) <| by simp [none_eq_bot]
308+
| a, some b => by simp [coe_eq_coe]
309+
| a, => iff_of_false (not_lt_none _) <| by simp
310310
#align with_bot.lt_iff_exists_coe WithBot.lt_iff_exists_coe
311311

312-
theorem lt_coe_iff : ∀ {x : WithBot α}, x < b ↔ ∀ a : α, x = a → a < b
313-
| Option.some b => by simp [some_eq_coe, coe_eq_coe, coe_lt_coe]
314-
| none => by simp [none_eq_bot, bot_lt_coe]
312+
theorem lt_coe_iff : ∀ {x : WithBot α}, x < b ↔ ∀ a : α, x = a → a < b
313+
| (_ : α) => by simp
314+
| => by simp [bot_lt_coe]
315315
#align with_bot.lt_coe_iff WithBot.lt_coe_iff
316316

317317
/-- A version of `bot_lt_iff_ne_bot` for `WithBot` that only requires `LT α`, not
@@ -1300,25 +1300,19 @@ instance instWellFoundedGT [LT α] [WellFoundedGT α] : WellFoundedGT (WithTop
13001300

13011301
instance trichotomous.lt [Preorder α] [IsTrichotomous α (· < ·)] :
13021302
IsTrichotomous (WithTop α) (· < ·) :=
1303-
by
1304-
rintro (a | a) (b | b)
1305-
· simp
1306-
· simp
1307-
· simp
1308-
· simpa [some_eq_coe, IsTrichotomous, coe_eq_coe] using @trichotomous α (· < ·) _ a b⟩
1303+
fun
1304+
| (a : α), (b : α) => by simp [trichotomous]
1305+
| ⊤, (b : α) => by simp
1306+
| (a : α), ⊤ => by simp
1307+
| ⊤, ⊤ => by simp⟩
13091308
#align with_top.trichotomous.lt WithTop.trichotomous.lt
13101309

13111310
instance IsWellOrder.lt [Preorder α] [IsWellOrder α (· < ·)] : IsWellOrder (WithTop α) (· < ·) where
13121311
#align with_top.is_well_order.lt WithTop.IsWellOrder.lt
13131312

13141313
instance trichotomous.gt [Preorder α] [IsTrichotomous α (· > ·)] :
13151314
IsTrichotomous (WithTop α) (· > ·) :=
1316-
by
1317-
rintro (a | a) (b | b)
1318-
· simp
1319-
· simp
1320-
· simp
1321-
· simpa [some_eq_coe, IsTrichotomous, coe_eq_coe] using @trichotomous α (· > ·) _ a b⟩
1315+
have : IsTrichotomous α (· < ·) := .swap _; .swap _
13221316
#align with_top.trichotomous.gt WithTop.trichotomous.gt
13231317

13241318
instance IsWellOrder.gt [Preorder α] [IsWellOrder α (· > ·)] : IsWellOrder (WithTop α) (· > ·) where

0 commit comments

Comments
 (0)