Skip to content

Commit

Permalink
Mention upstreaming PR (or remove if not used)
Browse files Browse the repository at this point in the history
  • Loading branch information
nomeata committed Jul 30, 2024
1 parent 6deb401 commit da9eefb
Showing 1 changed file with 3 additions and 10 deletions.
13 changes: 3 additions & 10 deletions Carleson/ToMathlib/Height.lean
Original file line number Diff line number Diff line change
Expand Up @@ -42,13 +42,12 @@ set. Some results
-/


-- https://github.com/leanprover-community/mathlib4/pull/15341
theorem ENat.lt_add_one_iff (n m : ℕ∞) (hm : n ≠ ⊤) : m < n + 1 ↔ m ≤ n := by
cases n <;> cases m <;> try contradiction
· norm_num
· norm_cast; omega


lemma ENat.iSup_eq_coe_iff' {α : Type*} [Nonempty α] (f : α → ℕ∞) (n : ℕ) :
(⨆ x, f x = n) ↔ (∃ x, f x = n) ∧ (∀ y, f y ≤ n) := by
constructor
Expand Down Expand Up @@ -86,18 +85,12 @@ lemma ENat.iSup_eq_coe_iff {α : Type*} [Nonempty α] (f : α → ℕ) (n : Nat)
(⨆ x, (f x : ℕ∞) = n) ↔ (∃ x, f x = n) ∧ (∀ y, f y ≤ n) := by
simp [ENat.iSup_eq_coe_iff']

@[simp]
theorem ENat.iSup_eq_zero {ι : Type*} (s : ι → ℕ∞) : iSup s = 0 ↔ ∀ i, s i = 0 := iSup_eq_bot


-- https://github.com/leanprover-community/mathlib4/pull/15342
@[simp]
lemma ENat.not_lt_zero (n : ℕ∞) : ¬ n < 0 := by
cases n <;> simp

@[simp]
lemma ENat.coe_lt_top (n : ℕ) : (n : ℕ∞) < ⊤ := by
exact Batteries.compareOfLessAndEq_eq_lt.mp rfl

-- https://github.com/leanprover-community/mathlib4/pull/15344
lemma ENat.isup_add (ι : Type*) [Nonempty ι] (f : ι → ℕ∞) (n : ℕ∞) :
(⨆ x, f x) + n = (⨆ x, f x + n) := by
cases n; simp; next n =>
Expand Down

0 comments on commit da9eefb

Please sign in to comment.