Documentation

Mathlib.Data.Int.Defs

Basic operations on the integers #

This file contains some basic lemmas about integers.

See note [foundational algebra order theory].

TODO #

Split this file into:

theorem Int.one_pos :
0 < 1
@[simp]
theorem Int.ofNat_eq_natCast {n : } :
Int.ofNat n = n
@[deprecated Int.ofNat_eq_natCast]
theorem Int.natCast_eq_ofNat (n : ) :
n = Int.ofNat n
theorem Int.natCast_inj {m : } {n : } :
m = n m = n
@[simp]
theorem Int.natAbs_cast (n : ) :
Int.natAbs n = n
theorem Int.natCast_sub {n : } {m : } :
n m(m - n) = m - n
theorem Int.natCast_eq_zero {n : } :
n = 0 n = 0
theorem Int.natCast_ne_zero {n : } :
n 0 n 0
theorem Int.natCast_ne_zero_iff_pos {n : } :
n 0 0 < n
theorem Int.natCast_nonneg (n : ) :
0 n
@[simp]
theorem Int.sign_natCast_add_one (n : ) :
Int.sign (n + 1) = 1
theorem Int.two_mul (n : ) :
2 * n = n + n
@[deprecated]
theorem Int.ofNat_bit0 (n : ) :
(bit0 n) = bit0 n
@[deprecated]
theorem Int.ofNat_bit1 (n : ) :
(bit1 n) = bit1 n

succ and pred #

def Int.succ (a : ) :

Immediate successor of an integer: succ n = n + 1

Equations
Instances For
    def Int.pred (a : ) :

    Immediate predecessor of an integer: pred n = n - 1

    Equations
    Instances For
      theorem Int.natCast_succ (n : ) :
      (Nat.succ n) = Int.succ n
      theorem Int.pred_nat_succ (n : ) :
      Int.pred (Nat.succ n) = n
      theorem Int.neg_nat_succ (n : ) :
      -(Nat.succ n) = Int.pred (-n)
      theorem Int.natCast_pred_of_pos {n : } (h : 0 < n) :
      (n - 1) = n - 1
      theorem Int.induction_on {p : Prop} (i : ) (hz : p 0) (hp : ∀ (i : ), p ip (i + 1)) (hn : ∀ (i : ), p (-i)p (-i - 1)) :
      p i

      nat abs #

      theorem Int.natAbs_add_of_nonneg {a : } {b : } :
      0 a0 bInt.natAbs (a + b) = Int.natAbs a + Int.natAbs b
      theorem Int.natAbs_add_of_nonpos {a : } {b : } (ha : a 0) (hb : b 0) :
      @[deprecated Int.natAbs_ne_zero]
      theorem Int.natAbs_sq (x : ) :
      (Int.natAbs x) ^ 2 = x ^ 2
      theorem Int.natAbs_pow_two (x : ) :
      (Int.natAbs x) ^ 2 = x ^ 2

      Alias of Int.natAbs_sq.

      / #

      @[simp]
      theorem Int.natCast_div (m : ) (n : ) :
      (m / n) = m / n
      theorem Int.natCast_ediv (m : ) (n : ) :
      (m / n) = Int.ediv m n
      theorem Int.ediv_of_neg_of_pos {a : } {b : } (Ha : a < 0) (Hb : 0 < b) :
      Int.ediv a b = -((-a - 1) / b + 1)

      mod #

      @[simp]
      theorem Int.natCast_mod (m : ) (n : ) :
      (m % n) = m % n

      properties of / and % #

      theorem Int.sign_natCast_of_ne_zero {n : } (hn : n 0) :
      Int.sign n = 1

      toNat #

      @[simp]
      theorem Int.toNat_natCast (n : ) :
      Int.toNat n = n
      @[simp]
      theorem Int.toNat_natCast_add_one {n : } :
      Int.toNat (n + 1) = n + 1
      @[simp]
      theorem Int.pow_eq (m : ) (n : ) :
      Int.pow m n = m ^ n
      @[deprecated Int.ofNat_eq_natCast]
      theorem Int.ofNat_eq_cast {n : } :
      Int.ofNat n = n

      Alias of Int.ofNat_eq_natCast.

      @[deprecated Int.natCast_inj]
      theorem Int.cast_eq_cast_iff_Nat {m : } {n : } :
      m = n m = n

      Alias of Int.natCast_inj.

      @[deprecated Int.natCast_sub]
      theorem Int.coe_nat_sub {n : } {m : } :
      n m(m - n) = m - n

      Alias of Int.natCast_sub.

      @[deprecated Int.natCast_nonneg]
      theorem Int.coe_nat_nonneg (n : ) :
      0 n

      Alias of Int.natCast_nonneg.

      @[deprecated Int.sign_natCast_add_one]
      theorem Int.sign_coe_add_one (n : ) :
      Int.sign (n + 1) = 1

      Alias of Int.sign_natCast_add_one.

      @[deprecated Int.natCast_succ]
      theorem Int.nat_succ_eq_int_succ (n : ) :
      (Nat.succ n) = Int.succ n

      Alias of Int.natCast_succ.

      @[deprecated Int.succ_neg_natCast_succ]
      theorem Int.succ_neg_nat_succ (n : ) :
      Int.succ (-(Nat.succ n)) = -n

      Alias of Int.succ_neg_natCast_succ.

      @[deprecated Int.natCast_pred_of_pos]
      theorem Int.coe_pred_of_pos {n : } (h : 0 < n) :
      (n - 1) = n - 1

      Alias of Int.natCast_pred_of_pos.

      @[deprecated Int.natCast_div]
      theorem Int.coe_nat_div (m : ) (n : ) :
      (m / n) = m / n

      Alias of Int.natCast_div.

      @[deprecated Int.natCast_ediv]
      theorem Int.coe_nat_ediv (m : ) (n : ) :
      (m / n) = Int.ediv m n

      Alias of Int.natCast_ediv.

      @[deprecated Int.sign_natCast_of_ne_zero]
      theorem Int.sign_coe_nat_of_nonzero {n : } (hn : n 0) :
      Int.sign n = 1

      Alias of Int.sign_natCast_of_ne_zero.

      @[deprecated Int.toNat_natCast]
      theorem Int.toNat_coe_nat (n : ) :
      Int.toNat n = n

      Alias of Int.toNat_natCast.

      @[deprecated Int.toNat_natCast_add_one]
      theorem Int.toNat_coe_nat_add_one {n : } :
      Int.toNat (n + 1) = n + 1

      Alias of Int.toNat_natCast_add_one.