Documentation

Mathlib.Data.Int.Init

Basic operations on the integers #

This file contains some basic lemmas about integers.

See note [foundational algebra order theory].

This file should not depend on anything defined in Mathlib (except for notation), so that it can be upstreamed to Batteries easily.

theorem Int.neg_eq_neg {a b : } (h : -a = -b) :
a = b
@[deprecated Int.neg_nonpos_iff (since := "2025-03-07")]
theorem Int.neg_nonpos_iff_nonneg (i : ) :
-i 0 0 i

Alias of Int.neg_nonpos_iff.

succ and pred #

def Int.succ (a : ) :

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

Equations
def Int.pred (a : ) :

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

Equations
theorem Int.pred_succ (a : ) :
a.succ.pred = a
theorem Int.succ_pred (a : ) :
a.pred.succ = a
theorem Int.neg_succ (a : ) :
-a.succ = (-a).pred
theorem Int.succ_neg_succ (a : ) :
(-a.succ).succ = -a
theorem Int.neg_pred (a : ) :
-a.pred = (-a).succ
theorem Int.pred_neg_pred (a : ) :
(-a.pred).pred = -a
theorem Int.pred_nat_succ (n : ) :
(↑n.succ).pred = n
theorem Int.neg_nat_succ (n : ) :
-n.succ = (-n).pred
theorem Int.succ_neg_natCast_succ (n : ) :
(-n.succ).succ = -n
theorem Int.natCast_pred_of_pos {n : } (h : 0 < n) :
↑(n - 1) = n - 1
theorem Int.lt_succ_self (a : ) :
a < a.succ
theorem Int.pred_self_lt (a : ) :
a.pred < a
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

Induction on integers: prove a proposition p i by proving the base case p 0, the upwards induction step p i → p (i + 1) and the downwards induction step p (-i) → p (-i - 1).

It is used as the default induction principle for the induction tactic.

def Int.inductionOn' {C : Sort u_1} (z b : ) (H0 : C b) (Hs : (k : ) → b kC kC (k + 1)) (Hp : (k : ) → k bC kC (k - 1)) :
C z

Inductively define a function on by defining it at b, for the succ of a number greater than b, and the pred of a number less than b.

Equations
def Int.inductionOn'.pos {C : Sort u_1} (b : ) (H0 : C b) (Hs : (k : ) → b kC kC (k + 1)) (n : ) :
C (b + n)

The positive case of Int.inductionOn'.

Equations
def Int.inductionOn'.neg {C : Sort u_1} (b : ) (H0 : C b) (Hp : (k : ) → k bC kC (k - 1)) (n : ) :
C (b + negSucc n)

The negative case of Int.inductionOn'.

Equations
theorem Int.inductionOn'_self {C : Sort u_1} {b : } {H0 : C b} {Hs : (k : ) → b kC kC (k + 1)} {Hp : (k : ) → k bC kC (k - 1)} :
b.inductionOn' b H0 Hs Hp = H0
theorem Int.inductionOn'_sub_one {C : Sort u_1} {z b : } {H0 : C b} {Hs : (k : ) → b kC kC (k + 1)} {Hp : (k : ) → k bC kC (k - 1)} (hz : z b) :
(z - 1).inductionOn' b H0 Hs Hp = Hp z hz (z.inductionOn' b H0 Hs Hp)
def Int.negInduction {C : Sort u_1} (nat : (n : ) → C n) (neg : ((n : ) → C n)(n : ) → C (-n)) (n : ) :
C n

Inductively define a function on by defining it on and extending it from n to -n.

Equations
theorem Int.le_induction {P : Prop} {m : } (h0 : P m) (h1 : ∀ (n : ), m nP nP (n + 1)) (n : ) :
m nP n

See Int.inductionOn' for an induction in both directions.

theorem Int.le_induction_down {P : Prop} {m : } (h0 : P m) (h1 : ∀ (n : ), n mP nP (n - 1)) (n : ) :
n mP n

See Int.inductionOn' for an induction in both directions.

def Int.strongRec {m : } {P : Sort u_1} (lt : (n : ) → n < mP n) (ge : (n : ) → n m((k : ) → k < nP k)P n) (n : ) :
P n

A strong recursor for Int that specifies explicit values for integers below a threshold, and is analogous to Nat.strongRec for integers on or above the threshold.

Equations
  • One or more equations did not get rendered due to their size.
theorem Int.strongRec_of_lt {m n : } {P : Sort u_1} {lt : (n : ) → n < mP n} {ge : (n : ) → n m((k : ) → k < nP k)P n} (hn : n < m) :
Int.strongRec lt ge n = lt n hn

mul #

natAbs #

theorem Int.natAbs_sq (a : ) :
a.natAbs ^ 2 = a ^ 2

Alias of Int.natAbs_pow_two.

/ #

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

mod #

@[simp]
theorem Int.natCast_mod (m n : ) :
↑(m % n) = m % n
@[deprecated Int.add_emod_eq_add_emod_right (since := "2025-04-16")]
theorem Int.add_emod_eq_add_mod_right {m n k : } (i : ) (H : m % n = k % n) :
(m + i) % n = (k + i) % n

Alias of Int.add_emod_eq_add_emod_right.

theorem Int.div_le_iff_of_dvd_of_pos {a b c : } (hb : 0 < b) (hba : b a) :
a / b c a b * c
theorem Int.div_le_iff_of_dvd_of_neg {a b c : } (hb : b < 0) (hba : b a) :
a / b c b * c a
theorem Int.div_lt_iff_of_dvd_of_pos {a b c : } (hb : 0 < b) (hba : b a) :
a / b < c a < b * c
theorem Int.div_lt_iff_of_dvd_of_neg {a b c : } (hb : b < 0) (hba : b a) :
a / b < c b * c < a
theorem Int.le_div_iff_of_dvd_of_pos {a b c : } (hc : 0 < c) (hcb : c b) :
a b / c c * a b
theorem Int.le_div_iff_of_dvd_of_neg {a b c : } (hc : c < 0) (hcb : c b) :
a b / c b c * a
theorem Int.lt_div_iff_of_dvd_of_pos {a b c : } (hc : 0 < c) (hcb : c b) :
a < b / c c * a < b
theorem Int.lt_div_iff_of_dvd_of_neg {a b c : } (hc : c < 0) (hcb : c b) :
a < b / c b < c * a
theorem Int.div_le_div_iff_of_dvd_of_pos_of_pos {a b c d : } (hb : 0 < b) (hd : 0 < d) (hba : b a) (hdc : d c) :
a / b c / d d * a c * b
theorem Int.div_le_div_iff_of_dvd_of_pos_of_neg {a b c d : } (hb : 0 < b) (hd : d < 0) (hba : b a) (hdc : d c) :
a / b c / d c * b d * a
theorem Int.div_le_div_iff_of_dvd_of_neg_of_pos {a b c d : } (hb : b < 0) (hd : 0 < d) (hba : b a) (hdc : d c) :
a / b c / d c * b d * a
theorem Int.div_le_div_iff_of_dvd_of_neg_of_neg {a b c d : } (hb : b < 0) (hd : d < 0) (hba : b a) (hdc : d c) :
a / b c / d d * a c * b
theorem Int.div_lt_div_iff_of_dvd_of_pos {a b c d : } (hb : 0 < b) (hd : 0 < d) (hba : b a) (hdc : d c) :
a / b < c / d d * a < c * b
theorem Int.div_lt_div_iff_of_dvd_of_pos_of_neg {a b c d : } (hb : 0 < b) (hd : d < 0) (hba : b a) (hdc : d c) :
a / b < c / d c * b < d * a
theorem Int.div_lt_div_iff_of_dvd_of_neg_of_pos {a b c d : } (hb : b < 0) (hd : 0 < d) (hba : b a) (hdc : d c) :
a / b < c / d c * b < d * a
theorem Int.div_lt_div_iff_of_dvd_of_neg_of_neg {a b c d : } (hb : b < 0) (hd : d < 0) (hba : b a) (hdc : d c) :
a / b < c / d d * a < c * b

properties of / and % #

theorem Int.emod_two_eq_zero_or_one (n : ) :
n % 2 = 0 n % 2 = 1

dvd #

theorem Int.dvd_mul_of_div_dvd {a b c : } (h : b a) (hdiv : a / b c) :
a b * c
theorem Int.div_dvd_iff_dvd_mul {a b c : } (h : b a) (hb : b 0) :
a / b c a b * c
theorem Int.mul_dvd_of_dvd_div {a b c : } (hcb : c b) (h : a b / c) :
c * a b
theorem Int.dvd_div_of_mul_dvd {a b c : } (h : a * b c) :
b c / a
theorem Int.dvd_div_iff_mul_dvd {a b c : } (hbc : c b) :
a b / c c * a b
theorem Int.exists_lt_and_lt_iff_not_dvd {n : } (m : ) (hn : 0 < n) :
( (k : ), n * k < m m < n * (k + 1)) ¬n m

If n > 0 then m is not divisible by n iff it is between n * k and n * (k + 1) for some k.

theorem Int.eq_mul_div_of_mul_eq_mul_of_dvd_left {a b c d : } (hb : b 0) (hbc : b c) (h : b * a = c * d) :
a = c / b * d
theorem Int.ofNat_add_negSucc_of_ge {m n : } (h : n.succ m) :
ofNat m + negSucc n = ofNat (m - n.succ)

/ and ordering #

theorem Int.le_iff_pos_of_dvd {a b : } (ha : 0 < a) (hab : a b) :
a b 0 < b
theorem Int.le_add_iff_lt_of_dvd_sub {a b c : } (ha : 0 < a) (hab : a c - b) :
a + b c b < c

sign #

theorem Int.sign_add_eq_of_sign_eq {m n : } :
m.sign = n.sign(m + n).sign = n.sign

toNat #

@[simp]
theorem Int.toNat_pred_coe_of_pos {i : } (h : 0 < i) :
↑(i.toNat - 1) = i - 1
theorem Int.toNat_lt'' {m : } {n : } (hn : n 0) :
m.toNat < n m < n
def Int.natMod (m n : ) :

The modulus of an integer by another as a natural. Uses the E-rounding convention.

Equations
theorem Int.natMod_lt {m : } {n : } (hn : n 0) :
m.natMod n < n
@[simp]
theorem Int.pow_eq (m : ) (n : ) :
m.pow n = m ^ n

For use in Mathlib.Tactic.NormNum.Pow