Skip to content

Commit

Permalink
feat(Data/Real/EReal): add simp theorems involving the sum of and…
Browse files Browse the repository at this point in the history
… `x` (#14102)

- [x] Add the theorem `top_add_of_ne_bot` which states that for any extended real number `x` which is not `⊥`, the sum of `⊤` and `x` is equal to `⊤`.
- [x] Add the theorem `add_top_of_ne_bot` which states that for any extended real number `x` which is not `⊥`, the sum of `x` and `⊤` is equal to `⊤`.
- [x]  Add the theorem `add_pos` which states that for any two extended real numbers `a` and `b`, if both `a` and `b` are greater than `0`, then their sum is also greater than `0`.
- [x] Add the theorem `mul_pos` which states that the product of two positive extended real numbers is positive. 
- [x] Add the theorem `add_top_iff_ne_bot` which states that for any extended real number `x`, the sum of `x` and `⊤` is equal to `⊤` if and only if `x` is not `⊥`.
- [x] Add the theorem `top_add_iff_ne_bot` which states that for any extended real number `x`, the sum of `⊤` and `x` is equal to `⊤` if and only if `x` is not `⊥`.
 
Co-authored-by: @D-Thomine 

-- 

[![Open in Gitpod](https://gitpod.io/button/open-in-gitpod.svg)](https://gitpod.io/from-referrer/)
  • Loading branch information
pitmonticone committed Jun 25, 2024
1 parent 0512c53 commit c4f1cc8
Showing 1 changed file with 51 additions and 0 deletions.
51 changes: 51 additions & 0 deletions Mathlib/Data/Real/EReal.lean
Original file line number Diff line number Diff line change
Expand Up @@ -800,6 +800,47 @@ theorem top_add_coe (x : ℝ) : (⊤ : EReal) + x = ⊤ :=
rfl
#align ereal.top_add_coe EReal.top_add_coe

/-- For any extended real number `x` which is not `⊥`, the sum of `⊤` and `x` is equal to `⊤`. -/
@[simp]
theorem top_add_of_ne_bot {x : EReal} (h : x ≠ ⊥) : ⊤ + x = ⊤ := by
induction x using EReal.rec
· exfalso; exact h (Eq.refl ⊥)
· exact top_add_coe _
· exact top_add_top

/-- For any extended real number `x`, the sum of `⊤` and `x` is equal to `⊤`
if and only if `x` is not `⊥`. -/
theorem top_add_iff_ne_bot {x : EReal} : ⊤ + x = ⊤ ↔ x ≠ ⊥ := by
constructor <;> intro h
· by_contra h'
rw [h', add_bot] at h
exact bot_ne_top h
· cases x
case h_bot => contradiction
case h_top => rfl
case h_real r => exact top_add_of_ne_bot h

/-- For any extended real number `x` which is not `⊥`, the sum of `x` and `⊤` is equal to `⊤`. -/
@[simp]
theorem add_top_of_ne_bot {x : EReal} (h : x ≠ ⊥) : x + ⊤ = ⊤ := by
rw [add_comm, top_add_of_ne_bot h]

/-- For any extended real number `x`, the sum of `x` and `⊤` is equal to `⊤`
if and only if `x` is not `⊥`. -/
theorem add_top_iff_ne_bot {x : EReal} : x + ⊤ = ⊤ ↔ x ≠ ⊥ := by rw [add_comm, top_add_iff_ne_bot]

/-- For any two extended real numbers `a` and `b`, if both `a` and `b` are greater than `0`,
then their sum is also greater than `0`. -/
theorem add_pos {a b : EReal} (ha : 0 < a) (hb : 0 < b) : 0 < a + b := by
induction a using EReal.rec
· exfalso; exact not_lt_bot ha
· induction b using EReal.rec
· exfalso; exact not_lt_bot hb
· norm_cast at *; exact Left.add_pos ha hb
· exact add_top_of_ne_bot (bot_lt_zero.trans ha).ne' ▸ hb
· rw [top_add_of_ne_bot (bot_lt_zero.trans hb).ne']
exact ha

@[simp]
theorem coe_add_top (x : ℝ) : (x : EReal) + ⊤ = ⊤ :=
rfl
Expand Down Expand Up @@ -1081,6 +1122,16 @@ theorem top_mul_of_pos {x : EReal} (h : 0 < x) : ⊤ * x = ⊤ := by
exact mul_top_of_pos h
#align ereal.top_mul_of_pos EReal.top_mul_of_pos

/-- The product of two positive extended real numbers is positive. -/
theorem mul_pos {a b : EReal} (ha : 0 < a) (hb : 0 < b) : 0 < a * b := by
induction' a using EReal.rec with a
· exfalso; exact not_lt_bot ha
· induction' b using EReal.rec with b
· exfalso; exact not_lt_bot hb
· norm_cast at *; exact Left.mul_pos ha hb
· rw [EReal.mul_comm, top_mul_of_pos ha]; exact hb
· rw [top_mul_of_pos hb]; exact ha

theorem top_mul_of_neg {x : EReal} (h : x < 0) : ⊤ * x = ⊥ := by
rw [EReal.mul_comm]
exact mul_top_of_neg h
Expand Down

0 comments on commit c4f1cc8

Please sign in to comment.