Skip to content

Commit bd3e369

Browse files
authored
feat(Data/../ENNReal): add lemmas about ENNReal.ofReal (#8163)
Compare `ENNReal.ofReal r` to `Nat.cast n` and literals. --- <!-- The text above the `---` will become the commit message when your PR is merged. Please leave a blank newline before the `---`, otherwise GitHub will format the text above it as a title. To indicate co-authors, include lines at the bottom of the commit message (that is, before the `---`) using the following format: Co-authored-by: Author Name <author@email.com> Any other comments you want to keep out of the PR commit should go below the `---`, and placed outside this HTML comment, or else they will be invisible to reviewers. If this PR depends on other PRs, please list them below this comment, using the following format: - [ ] depends on: #abc [optional extra text] - [ ] depends on: #xyz [optional extra text] --> [![Open in Gitpod](https://gitpod.io/button/open-in-gitpod.svg)](https://gitpod.io/from-referrer/)
1 parent 64b26c9 commit bd3e369

File tree

1 file changed

+71
-1
lines changed

1 file changed

+71
-1
lines changed

Mathlib/Data/Real/ENNReal.lean

Lines changed: 71 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -707,7 +707,7 @@ theorem one_lt_coe_iff : 1 < (↑p : ℝ≥0∞) ↔ 1 < p := coe_lt_coe
707707
theorem coe_nat (n : ℕ) : ((n : ℝ≥0) : ℝ≥0∞) = n := rfl
708708
#align ennreal.coe_nat ENNReal.coe_nat
709709

710-
@[simp] theorem ofReal_coe_nat (n : ℕ) : ENNReal.ofReal n = n := by simp [ENNReal.ofReal]
710+
@[simp, norm_cast] lemma ofReal_coe_nat (n : ℕ) : ENNReal.ofReal n = n := by simp [ENNReal.ofReal]
711711
#align ennreal.of_real_coe_nat ENNReal.ofReal_coe_nat
712712

713713
@[simp] theorem ofReal_ofNat (n : ℕ) [n.AtLeastTwo] :
@@ -2146,6 +2146,12 @@ theorem ofReal_le_ofReal_iff {p q : ℝ} (h : 0 ≤ q) : ENNReal.ofReal p ≤ EN
21462146
by rw [ENNReal.ofReal, ENNReal.ofReal, coe_le_coe, Real.toNNReal_le_toNNReal_iff h]
21472147
#align ennreal.of_real_le_of_real_iff ENNReal.ofReal_le_ofReal_iff
21482148

2149+
lemma ofReal_le_ofReal_iff' {p q : ℝ} : ENNReal.ofReal p ≤ .ofReal q ↔ p ≤ q ∨ p ≤ 0 :=
2150+
coe_le_coe.trans Real.toNNReal_le_toNNReal_iff'
2151+
2152+
lemma ofReal_lt_ofReal_iff' {p q : ℝ} : ENNReal.ofReal p < .ofReal q ↔ p < q ∧ 0 < q :=
2153+
coe_lt_coe.trans Real.toNNReal_lt_toNNReal_iff'
2154+
21492155
@[simp]
21502156
theorem ofReal_eq_ofReal_iff {p q : ℝ} (hp : 0 ≤ p) (hq : 0 ≤ q) :
21512157
ENNReal.ofReal p = ENNReal.ofReal q ↔ p = q := by
@@ -2178,6 +2184,70 @@ theorem zero_eq_ofReal {p : ℝ} : 0 = ENNReal.ofReal p ↔ p ≤ 0 :=
21782184
alias ⟨_, ofReal_of_nonpos⟩ := ofReal_eq_zero
21792185
#align ennreal.of_real_of_nonpos ENNReal.ofReal_of_nonpos
21802186

2187+
@[simp]
2188+
lemma ofReal_lt_nat_cast {p : ℝ} {n : ℕ} (hn : n ≠ 0) : ENNReal.ofReal p < n ↔ p < n := by
2189+
exact_mod_cast ofReal_lt_ofReal_iff (Nat.cast_pos.2 hn.bot_lt)
2190+
2191+
@[simp]
2192+
lemma ofReal_lt_one {p : ℝ} : ENNReal.ofReal p < 1 ↔ p < 1 := by
2193+
exact_mod_cast ofReal_lt_nat_cast one_ne_zero
2194+
2195+
@[simp]
2196+
lemma ofReal_lt_ofNat {p : ℝ} {n : ℕ} [h : n.AtLeastTwo] :
2197+
ENNReal.ofReal p < no_index (OfNat.ofNat n) ↔ p < OfNat.ofNat n :=
2198+
ofReal_lt_nat_cast h.ne_zero
2199+
2200+
@[simp]
2201+
lemma nat_cast_le_ofReal {n : ℕ} {p : ℝ} (hn : n ≠ 0) : n ≤ ENNReal.ofReal p ↔ n ≤ p := by
2202+
simp only [← not_lt, ofReal_lt_nat_cast hn]
2203+
2204+
@[simp]
2205+
lemma one_le_ofReal {p : ℝ} : 1 ≤ ENNReal.ofReal p ↔ 1 ≤ p := by
2206+
exact_mod_cast nat_cast_le_ofReal one_ne_zero
2207+
2208+
@[simp]
2209+
lemma ofNat_le_ofReal {n : ℕ} [h : n.AtLeastTwo] {p : ℝ} :
2210+
no_index (OfNat.ofNat n) ≤ ENNReal.ofReal p ↔ OfNat.ofNat n ≤ p :=
2211+
nat_cast_le_ofReal h.ne_zero
2212+
2213+
@[simp]
2214+
lemma ofReal_le_nat_cast {r : ℝ} {n : ℕ} : ENNReal.ofReal r ≤ n ↔ r ≤ n :=
2215+
coe_le_coe.trans Real.toNNReal_le_nat_cast
2216+
2217+
@[simp]
2218+
lemma ofReal_le_one {r : ℝ} : ENNReal.ofReal r ≤ 1 ↔ r ≤ 1 :=
2219+
coe_le_coe.trans Real.toNNReal_le_one
2220+
2221+
@[simp]
2222+
lemma ofReal_le_ofNat {r : ℝ} {n : ℕ} [n.AtLeastTwo] :
2223+
ENNReal.ofReal r ≤ no_index (OfNat.ofNat n) ↔ r ≤ OfNat.ofNat n :=
2224+
ofReal_le_nat_cast
2225+
2226+
@[simp]
2227+
lemma nat_cast_lt_ofReal {n : ℕ} {r : ℝ} : n < ENNReal.ofReal r ↔ n < r :=
2228+
coe_lt_coe.trans Real.nat_cast_lt_toNNReal
2229+
2230+
@[simp]
2231+
lemma one_lt_ofReal {r : ℝ} : 1 < ENNReal.ofReal r ↔ 1 < r := coe_lt_coe.trans Real.one_lt_toNNReal
2232+
2233+
@[simp]
2234+
lemma ofNat_lt_ofReal {n : ℕ} [n.AtLeastTwo] {r : ℝ} :
2235+
no_index (OfNat.ofNat n) < ENNReal.ofReal r ↔ OfNat.ofNat n < r :=
2236+
nat_cast_lt_ofReal
2237+
2238+
@[simp]
2239+
lemma ofReal_eq_nat_cast {r : ℝ} {n : ℕ} (h : n ≠ 0) : ENNReal.ofReal r = n ↔ r = n :=
2240+
ENNReal.coe_eq_coe.trans <| Real.toNNReal_eq_nat_cast h
2241+
2242+
@[simp]
2243+
lemma ofReal_eq_one {r : ℝ} : ENNReal.ofReal r = 1 ↔ r = 1 :=
2244+
ENNReal.coe_eq_coe.trans Real.toNNReal_eq_one
2245+
2246+
@[simp]
2247+
lemma ofReal_eq_ofNat {r : ℝ} {n : ℕ} [h : n.AtLeastTwo] :
2248+
ENNReal.ofReal r = no_index (OfNat.ofNat n) ↔ r = OfNat.ofNat n :=
2249+
ofReal_eq_nat_cast h.ne_zero
2250+
21812251
theorem ofReal_sub (p : ℝ) {q : ℝ} (hq : 0 ≤ q) :
21822252
ENNReal.ofReal (p - q) = ENNReal.ofReal p - ENNReal.ofReal q := by
21832253
obtain h | h := le_total p q

0 commit comments

Comments
 (0)