mathlib3 documentation

core / init.data.int.basic

The integers, with addition, multiplication, and subtraction. #

the type, coercions, and notation

inductive int  :
Instances for int
@[protected, instance]
@[protected, instance]
Equations
@[protected]
Equations
@[protected, instance]
Equations
@[protected, instance]
Equations
@[protected]
theorem int.coe_nat_eq (n : ) :
@[protected]
def int.zero  :
Equations
@[protected]
def int.one  :
Equations
@[protected, instance]
Equations
@[protected, instance]
Equations
theorem int.of_nat_one  :

definitions of basic functions

def int.sub_nat_nat (m n : ) :
Equations
theorem int.sub_nat_nat_of_sub_eq_zero {m n : } (h : n - m = 0) :
theorem int.sub_nat_nat_of_sub_eq_succ {m n k : } (h : n - m = k.succ) :
@[protected]
def int.neg  :
Equations
@[protected]
def int.mul  :
Equations
@[protected, instance]
Equations
@[protected, instance]
Equations
@[protected, instance]
Equations
@[protected]
def int.sub  :
Equations
@[protected, instance]
Equations
@[protected]
theorem int.neg_zero  :
-0 = 0
theorem int.of_nat_add (n m : ) :
theorem int.of_nat_mul (n m : ) :
@[simp]
theorem int.of_nat_eq_coe (n : ) :
theorem int.neg_succ_of_nat_coe (n : ) :
-[1+ n] = -(n + 1)
@[protected]
theorem int.coe_nat_add (m n : ) :
(m + n) = m + n
@[protected]
theorem int.coe_nat_mul (m n : ) :
(m * n) = m * n
@[protected]
theorem int.coe_nat_zero  :
0 = 0
@[protected]
theorem int.coe_nat_one  :
1 = 1
@[protected]
theorem int.coe_nat_succ (n : ) :
(n.succ) = n + 1
@[protected]
theorem int.coe_nat_add_out (m n : ) :
m + n = m + n
@[protected]
theorem int.coe_nat_mul_out (m n : ) :
m * n = (m * n)
@[protected]
theorem int.coe_nat_add_one_out (n : ) :
n + 1 = (n.succ)

these are only for internal use

some basic functions and properties

@[protected]
theorem int.coe_nat_inj {m n : } (h : m = n) :
m = n
@[protected]
theorem int.coe_nat_eq_coe_nat_iff (m n : ) :
m = n m = n
theorem int.neg_succ_of_nat_inj_iff {m n : } :
-[1+ m] = -[1+ n] m = n
theorem int.neg_succ_of_nat_eq (n : ) :
-[1+ n] = -(n + 1)

neg

@[protected]
theorem int.neg_neg (a : ) :
--a = a
@[protected]
theorem int.neg_inj {a b : } (h : -a = -b) :
a = b
@[protected]
theorem int.sub_eq_add_neg {a b : } :
a - b = a + -b

basic properties of sub_nat_nat

theorem int.sub_nat_nat_elim (m n : ) (P : Prop) (hp : (i n : ), P (n + i) n (int.of_nat i)) (hn : (i m : ), P m (m + i + 1) -[1+ i]) :
P m n (int.sub_nat_nat m n)
theorem int.sub_nat_nat_add_right {m n : } :
int.sub_nat_nat m (m + n + 1) = -[1+ n]
theorem int.sub_nat_nat_add_add (m n k : ) :
theorem int.sub_nat_nat_of_le {m n : } (h : n m) :
theorem int.sub_nat_nat_of_lt {m n : } (h : m < n) :

nat_abs

@[simp]
Equations
@[simp, norm_cast]
theorem int.nat_abs_of_nat (n : ) :
theorem int.nat_abs_pos_of_ne_zero {a : } (h : a 0) :
@[simp]
theorem int.nat_abs_zero  :
@[simp]
theorem int.nat_abs_one  :
theorem int.nat_abs_mul_self {a : } :
(a.nat_abs * a.nat_abs) = a * a
@[simp]
theorem int.nat_abs_neg (a : ) :
theorem int.nat_abs_eq (a : ) :
theorem int.eq_coe_or_neg (a : ) :
(n : ), a = n a = -n

sign

def int.sign  :
Equations
@[simp]
theorem int.sign_zero  :
0.sign = 0
@[simp]
theorem int.sign_one  :
1.sign = 1
@[simp]
theorem int.sign_neg_one  :
(-1).sign = -1

Quotient and remainder

@[protected]
def int.div  :
Equations
@[protected]
def int.mod  :
Equations
Equations
Equations
@[protected, instance]
Equations
@[protected, instance]
Equations

gcd

def int.gcd (m n : ) :
Equations

addition

@[protected]
theorem int.add_comm (a b : ) :
a + b = b + a
@[protected]
theorem int.add_zero (a : ) :
a + 0 = a
@[protected]
theorem int.zero_add (a : ) :
0 + a = a
theorem int.sub_nat_nat_sub {m n : } (h : n m) (k : ) :
theorem int.add_assoc_aux1 (m n : ) (c : ) :
theorem int.add_assoc_aux2 (m n k : ) :
@[protected]
theorem int.add_assoc (a b c : ) :
a + b + c = a + (b + c)

negation

theorem int.sub_nat_self (n : ) :
@[protected]
theorem int.add_left_neg (a : ) :
-a + a = 0
@[protected]
theorem int.add_right_neg (a : ) :
a + -a = 0

multiplication

@[protected]
theorem int.mul_comm (a b : ) :
a * b = b * a
@[protected]
theorem int.mul_assoc (a b c : ) :
a * b * c = a * (b * c)
@[protected]
theorem int.mul_zero (a : ) :
a * 0 = 0
@[protected]
theorem int.zero_mul (a : ) :
0 * a = 0
@[protected]
theorem int.distrib_left (a b c : ) :
a * (b + c) = a * b + a * c
@[protected]
theorem int.distrib_right (a b c : ) :
(a + b) * c = a * c + b * c
@[protected]
theorem int.zero_ne_one  :
0 1
theorem int.of_nat_sub {n m : } (h : m n) :
@[protected]
theorem int.add_left_comm (a b c : ) :
a + (b + c) = b + (a + c)
@[protected]
theorem int.add_left_cancel {a b c : } (h : a + b = a + c) :
b = c
@[protected]
theorem int.neg_add {a b : } :
-(a + b) = -a + -b
theorem int.neg_succ_of_nat_coe' (n : ) :
-[1+ n] = -n - 1
@[protected]
theorem int.coe_nat_sub {n m : } :
n m (m - n) = m - n
@[protected]
theorem int.sub_nat_nat_eq_coe {m n : } :
def int.to_nat  :
Equations
theorem int.to_nat_sub (m n : ) :
(m - n).to_nat = m - n
def int.nat_mod (m n : ) :
Equations
@[protected]
theorem int.one_mul (a : ) :
1 * a = a
@[protected]
theorem int.mul_one (a : ) :
a * 1 = a
@[protected]
theorem int.neg_eq_neg_one_mul (a : ) :
-a = (-1) * a
theorem int.sign_mul_nat_abs (a : ) :
a.sign * (a.nat_abs) = a