@@ -96,8 +96,8 @@ theorem coe_ne_bot : (a : WithBot α) ≠ ⊥ :=
96
96
/-- Recursor for `WithBot` using the preferred forms `⊥` and `↑a`. -/
97
97
@[elab_as_elim]
98
98
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
101
101
#align with_bot.rec_bot_coe WithBot.recBotCoe
102
102
103
103
@[simp]
@@ -239,18 +239,18 @@ theorem coe_le : ∀ {o : Option α}, b ∈ o → ((a : WithBot α) ≤ o ↔ a
239
239
#align with_bot.coe_le WithBot.coe_le
240
240
241
241
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
244
244
#align with_bot.coe_le_iff WithBot.coe_le_iff
245
245
246
246
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
249
249
#align with_bot.le_coe_iff WithBot.le_coe_iff
250
250
251
251
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
254
254
#align is_max.with_bot IsMax.withBot
255
255
256
256
theorem le_unbot_iff {a : α} {b : WithBot α} (h : b ≠ ⊥) :
@@ -305,13 +305,13 @@ theorem not_lt_none (a : WithBot α) : ¬@LT.lt (WithBot α) _ a none :=
305
305
#align with_bot.not_lt_none WithBot.not_lt_none
306
306
307
307
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
310
310
#align with_bot.lt_iff_exists_coe WithBot.lt_iff_exists_coe
311
311
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]
315
315
#align with_bot.lt_coe_iff WithBot.lt_coe_iff
316
316
317
317
/-- 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
1300
1300
1301
1301
instance trichotomous.lt [Preorder α] [IsTrichotomous α (· < ·)] :
1302
1302
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⟩
1309
1308
#align with_top.trichotomous.lt WithTop.trichotomous.lt
1310
1309
1311
1310
instance IsWellOrder.lt [Preorder α] [IsWellOrder α (· < ·)] : IsWellOrder (WithTop α) (· < ·) where
1312
1311
#align with_top.is_well_order.lt WithTop.IsWellOrder.lt
1313
1312
1314
1313
instance trichotomous.gt [Preorder α] [IsTrichotomous α (· > ·)] :
1315
1314
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 _
1322
1316
#align with_top.trichotomous.gt WithTop.trichotomous.gt
1323
1317
1324
1318
instance IsWellOrder.gt [Preorder α] [IsWellOrder α (· > ·)] : IsWellOrder (WithTop α) (· > ·) where
0 commit comments