Documentation

Mathlib.Data.ENat.Basic

Definition and basic properties of extended natural numbers #

In this file we define ENat (notation: ℕ∞) to be WithTop and prove some basic lemmas about this type.

Implementation details #

There are two natural coercions from to WithTop ℕ = ENat: WithTop.some and Nat.cast. In Lean 3, this difference was hidden in typeclass instances. Since these instances were definitionally equal, we did not duplicate generic lemmas about WithTop α and WithTop.some coercion for ENat and Nat.cast coercion. If you need to apply a lemma about WithTop, you may either rewrite back and forth using ENat.some_eq_coe, or restate the lemma for ENat.

TODO #

Unify ENat.add_iSup/ENat.iSup_add with ENNReal.add_iSup/ENNReal.iSup_add. The key property of ENat and ENNReal we are using is that all a are either absorbing for addition (a + b = a for all b), or that it's order-cancellable (a + b ≤ a + c → b ≤ c for all b, c), and similarly for multiplication.

Equations
Equations
Equations
Equations
Equations
Equations
@[simp]
theorem ENat.some_eq_coe :
WithTop.some = Nat.cast

Lemmas about WithTop expect (and can output) WithTop.some but the normal form for coercion ℕ → ℕ∞ is Nat.cast.

theorem ENat.coe_inj {a b : } :
a = b a = b
theorem ENat.coe_zero :
0 = 0
theorem ENat.coe_one :
1 = 1
theorem ENat.coe_add (m n : ) :
(m + n) = m + n
@[simp]
theorem ENat.coe_sub (m n : ) :
(m - n) = m - n
@[simp]
theorem ENat.coe_mul (m n : ) :
(m * n) = m * n
@[simp]
theorem ENat.mul_top {m : ℕ∞} (hm : m 0) :
@[simp]
theorem ENat.top_mul {m : ℕ∞} (hm : m 0) :
theorem ENat.top_pow {n : } (n_pos : 0 < n) :
def ENat.lift (x : ℕ∞) (h : x < ) :

Convert a ℕ∞ to a using a proof that it is not infinite.

Equations
Instances For
    @[simp]
    theorem ENat.coe_lift (x : ℕ∞) (h : x < ) :
    (x.lift h) = x
    @[simp]
    theorem ENat.lift_coe (n : ) :
    (↑n).lift = n
    @[simp]
    theorem ENat.lift_lt_iff {x : ℕ∞} {h : x < } {n : } :
    x.lift h < n x < n
    @[simp]
    theorem ENat.lift_le_iff {x : ℕ∞} {h : x < } {n : } :
    x.lift h n x n
    @[simp]
    theorem ENat.lt_lift_iff {x : } {n : ℕ∞} {h : n < } :
    x < n.lift h x < n
    @[simp]
    theorem ENat.le_lift_iff {x : } {n : ℕ∞} {h : n < } :
    x n.lift h x n
    @[simp]
    theorem ENat.lift_zero :
    ENat.lift 0 = 0
    @[simp]
    theorem ENat.lift_one :
    ENat.lift 1 = 1
    @[simp]
    theorem ENat.lift_ofNat (n : ) [n.AtLeastTwo] :
    (OfNat.ofNat n).lift = OfNat.ofNat n
    @[simp]
    theorem ENat.add_lt_top {a b : ℕ∞} :
    a + b < a < b <
    @[simp]
    theorem ENat.lift_add (a b : ℕ∞) (h : a + b < ) :
    (a + b).lift h = a.lift + b.lift
    theorem ENat.canLift :
    CanLift ℕ∞ Nat.cast fun (x : ℕ∞) => x

    Conversion of ℕ∞ to sending to 0.

    Equations
    Instances For

      Homomorphism from ℕ∞ to sending to 0.

      Equations
      Instances For
        theorem ENat.toNatHom_apply (n : ) :
        ENat.toNatHom n = (↑n).toNat
        @[simp]
        theorem ENat.toNat_coe (n : ) :
        (↑n).toNat = n
        @[simp]
        @[simp]
        @[simp]
        theorem ENat.toNat_ofNat (n : ) [n.AtLeastTwo] :
        (OfNat.ofNat n).toNat = n
        @[simp]
        theorem ENat.toNat_top :
        .toNat = 0
        @[simp]
        theorem ENat.toNat_eq_zero {n : ℕ∞} :
        n.toNat = 0 n = 0 n =
        @[simp]
        theorem ENat.recTopCoe_zero {C : ℕ∞Sort u_1} (d : C ) (f : (a : ) → C a) :
        ENat.recTopCoe d f 0 = f 0
        @[simp]
        theorem ENat.recTopCoe_one {C : ℕ∞Sort u_1} (d : C ) (f : (a : ) → C a) :
        ENat.recTopCoe d f 1 = f 1
        @[simp]
        theorem ENat.recTopCoe_ofNat {C : ℕ∞Sort u_1} (d : C ) (f : (a : ) → C a) (x : ) [x.AtLeastTwo] :
        @[simp]
        theorem ENat.top_ne_coe (a : ) :
        a
        @[simp]
        theorem ENat.top_ne_ofNat (a : ) [a.AtLeastTwo] :
        @[simp]
        @[simp]
        @[simp]
        theorem ENat.coe_ne_top (a : ) :
        a
        @[simp]
        theorem ENat.ofNat_ne_top (a : ) [a.AtLeastTwo] :
        @[simp]
        @[simp]
        @[simp]
        theorem ENat.top_sub_coe (a : ) :
        - a =
        @[simp]
        @[simp]
        theorem ENat.top_sub_ofNat (a : ) [a.AtLeastTwo] :
        @[simp]
        theorem ENat.top_pos :
        0 <
        @[deprecated ENat.top_pos]

        Alias of ENat.top_pos.

        theorem ENat.sub_top (a : ℕ∞) :
        a - = 0
        @[simp]
        theorem ENat.coe_toNat_eq_self {n : ℕ∞} :
        n.toNat = n n
        theorem ENat.coe_toNat {n : ℕ∞} :
        n n.toNat = n

        Alias of the reverse direction of ENat.coe_toNat_eq_self.

        theorem ENat.coe_toNat_le_self (n : ℕ∞) :
        n.toNat n
        theorem ENat.toNat_add {m n : ℕ∞} (hm : m ) (hn : n ) :
        (m + n).toNat = m.toNat + n.toNat
        theorem ENat.toNat_sub {n : ℕ∞} (hn : n ) (m : ℕ∞) :
        (m - n).toNat = m.toNat - n.toNat
        theorem ENat.toNat_eq_iff {m : ℕ∞} {n : } (hn : n 0) :
        m.toNat = n m = n
        theorem ENat.toNat_le_of_le_coe {m : ℕ∞} {n : } (h : m n) :
        m.toNat n
        theorem ENat.toNat_le_toNat {m n : ℕ∞} (h : m n) (hn : n ) :
        m.toNat n.toNat
        @[simp]
        theorem ENat.succ_def (m : ℕ∞) :
        Order.succ m = m + 1
        @[deprecated Order.add_one_le_of_lt]
        theorem ENat.add_one_le_of_lt {m n : ℕ∞} (h : m < n) :
        m + 1 n
        theorem ENat.add_one_le_iff {m n : ℕ∞} (hm : m ) :
        m + 1 n m < n
        @[deprecated Order.one_le_iff_pos]
        theorem ENat.one_le_iff_pos {n : ℕ∞} :
        1 n 0 < n
        theorem ENat.lt_one_iff_eq_zero {n : ℕ∞} :
        n < 1 n = 0
        @[deprecated Order.le_of_lt_add_one]
        theorem ENat.le_of_lt_add_one {m n : ℕ∞} (h : m < n + 1) :
        m n
        theorem ENat.lt_add_one_iff {m n : ℕ∞} (hm : n ) :
        m < n + 1 m n
        theorem ENat.le_coe_iff {n : ℕ∞} {k : } :
        n k ∃ (n₀ : ), n = n₀ n₀ k
        @[simp]
        theorem ENat.not_lt_zero (n : ℕ∞) :
        ¬n < 0
        @[simp]
        theorem ENat.coe_lt_top (n : ) :
        n <
        theorem ENat.coe_lt_coe {n m : } :
        n < m n < m
        theorem ENat.coe_le_coe {n m : } :
        n m n m
        theorem ENat.nat_induction {P : ℕ∞Prop} (a : ℕ∞) (h0 : P 0) (hsuc : ∀ (n : ), P nP n.succ) (htop : (∀ (n : ), P n)P ) :
        P a
        theorem ENat.add_one_pos {n : ℕ∞} :
        0 < n + 1
        theorem ENat.add_lt_add_iff_right {m n k : ℕ∞} (h : k ) :
        n + k < m + k n < m
        theorem ENat.add_lt_add_iff_left {m n k : ℕ∞} (h : k ) :
        k + n < k + m n < m
        theorem ENat.exists_nat_gt {n : ℕ∞} (hn : n ) :
        ∃ (m : ), n < m
        @[simp]
        theorem ENat.sub_eq_top_iff {a b : ℕ∞} :
        a - b = a = b
        theorem ENat.le_sub_of_add_le_left {a b c : ℕ∞} (ha : a ) :
        a + b cb c - a
        theorem ENat.sub_sub_cancel {a b : ℕ∞} (h : a ) (h2 : b a) :
        a - (a - b) = b
        theorem ENat.add_one_natCast_le_withTop_of_lt {m : } {n : WithTop ℕ∞} (h : m < n) :
        (m + 1) n
        @[simp]
        theorem ENat.coe_top_add_one :
        + 1 =
        @[simp]
        theorem ENat.add_one_eq_coe_top_iff {n : WithTop ℕ∞} :
        n + 1 = n =
        @[simp]
        theorem ENat.natCast_ne_coe_top (n : ) :
        n
        @[deprecated ENat.natCast_ne_coe_top]
        theorem ENat.nat_ne_coe_top (n : ) :
        n

        Alias of ENat.natCast_ne_coe_top.

        theorem ENat.natCast_le_of_coe_top_le_withTop {N : WithTop ℕ∞} (hN : N) (n : ) :
        n N
        theorem ENat.natCast_lt_of_coe_top_le_withTop {N : WithTop ℕ∞} (hN : N) (n : ) :
        n < N