doc-next-gen

Init.Data.Int.Order

Module docstring

{"# Results about the order properties of the integers, and the integers as an ordered ring. ","## Order properties of the integers ","### toNat ","### toNat? ","## Order properties of the integers "}

Int.NonNeg.elim theorem
{a : Int} : NonNeg a → ∃ n : Nat, a = n
Full source
theorem NonNeg.elim {a : Int} : NonNeg a → ∃ n : Nat, a = n := nonneg_def.1
Non-Negative Integers are Natural Numbers
Informal description
For any integer $a$, if $a$ is non-negative, then there exists a natural number $n$ such that $a = n$.
Int.nonneg_or_nonneg_neg theorem
: ∀ (a : Int), NonNeg a ∨ NonNeg (-a)
Full source
theorem nonneg_or_nonneg_neg : ∀ (a : Int), NonNegNonNeg a ∨ NonNeg (-a)
  | (_:Nat) => .inl ⟨_⟩
  | -[_+1]  => .inr ⟨_⟩
Non-Negativity Dichotomy for Integers
Informal description
For every integer $a$, either $a$ is non-negative or $-a$ is non-negative.
Int.le_def theorem
{a b : Int} : a ≤ b ↔ NonNeg (b - a)
Full source
theorem le_def {a b : Int} : a ≤ b ↔ NonNeg (b - a) := .rfl
Characterization of Integer Inequality via Non-Negativity of Difference
Informal description
For any integers $a$ and $b$, the inequality $a \leq b$ holds if and only if the difference $b - a$ is non-negative.
Int.lt_iff_add_one_le theorem
{a b : Int} : a < b ↔ a + 1 ≤ b
Full source
theorem lt_iff_add_one_le {a b : Int} : a < b ↔ a + 1 ≤ b := .rfl
Characterization of Strict Inequality via Successor in Integers
Informal description
For any integers $a$ and $b$, the strict inequality $a < b$ holds if and only if $a + 1 \leq b$.
Int.le.intro_sub theorem
{a b : Int} (n : Nat) (h : b - a = n) : a ≤ b
Full source
theorem le.intro_sub {a b : Int} (n : Nat) (h : b - a = n) : a ≤ b := by
  simp [le_def, h]; constructor
Integer Inequality via Natural Difference
Informal description
For any integers $a$ and $b$, if there exists a natural number $n$ such that $b - a = n$, then $a \leq b$.
Int.le.intro theorem
{a b : Int} (n : Nat) (h : a + n = b) : a ≤ b
Full source
theorem le.intro {a b : Int} (n : Nat) (h : a + n = b) : a ≤ b :=
  le.intro_sub n <| by rw [← h, Int.add_comm]; simp [Int.sub_eq_add_neg, Int.add_assoc, Int.add_right_neg]
Integer Inequality via Natural Number Addition
Informal description
For any integers $a$ and $b$, if there exists a natural number $n$ such that $a + n = b$, then $a \leq b$.
Int.le.dest_sub theorem
{a b : Int} (h : a ≤ b) : ∃ n : Nat, b - a = n
Full source
theorem le.dest_sub {a b : Int} (h : a ≤ b) : ∃ n : Nat, b - a = n := nonneg_def.1 h
Existence of Natural Difference for Integer Inequality
Informal description
For any integers $a$ and $b$ such that $a \leq b$, there exists a natural number $n$ such that $b - a = n$.
Int.ofNat_le theorem
{m n : Nat} : (↑m : Int) ≤ ↑n ↔ m ≤ n
Full source
@[simp, norm_cast] theorem ofNat_le {m n : Nat} : (↑m : Int) ≤ ↑n ↔ m ≤ n :=
  ⟨fun h =>
    let ⟨k, hk⟩ := le.dest h
    Nat.le.intro <| Int.ofNat.inj <| (Int.ofNat_add m k).trans hk,
  fun h =>
    let ⟨k, (hk : m + k = n)⟩ := Nat.le.dest h
    le.intro k (by rw [← hk]; rfl)⟩
Preservation of Order under Canonical Embedding: $m \leq n$ in $\mathbb{Z}$ iff $m \leq n$ in $\mathbb{N}$
Informal description
For any natural numbers $m$ and $n$, the canonical image of $m$ in the integers is less than or equal to the canonical image of $n$ in the integers if and only if $m \leq n$ in the natural numbers. In other words, $m \leq n$ as integers if and only if $m \leq n$ as natural numbers.
Int.ofNat_zero_le theorem
(n : Nat) : 0 ≤ (↑n : Int)
Full source
@[simp] theorem ofNat_zero_le (n : Nat) : 0 ≤ (↑n : Int) := ofNat_le.2 n.zero_le
Nonnegativity of Natural Numbers Embedded in Integers: $0 \leq n$
Informal description
For any natural number $n$, the integer $0$ is less than or equal to the canonical image of $n$ in the integers, i.e., $0 \leq n$ where $n$ is considered as an integer.
Int.lt_add_succ theorem
(a : Int) (n : Nat) : a < a + (n + 1)
Full source
theorem lt_add_succ (a : Int) (n : Nat) : a < a + (n + 1) :=
  le.intro n <| by rw [Int.add_comm, Int.add_left_comm]
Integer Strict Inequality Under Successor Addition: $a < a + (n + 1)$
Informal description
For any integer $a$ and natural number $n$, we have $a < a + (n + 1)$.
Int.lt.intro theorem
{a b : Int} {n : Nat} (h : a + (n + 1) = b) : a < b
Full source
theorem lt.intro {a b : Int} {n : Nat} (h : a + (n + 1) = b) : a < b :=
  h ▸ lt_add_succ a n
Integer Strict Inequality from Successor Addition: $a + (n + 1) = b \implies a < b$
Informal description
For any integers $a$ and $b$ and any natural number $n$, if $a + (n + 1) = b$, then $a < b$.
Int.lt.dest theorem
{a b : Int} (h : a < b) : ∃ n : Nat, a + Nat.succ n = b
Full source
theorem lt.dest {a b : Int} (h : a < b) : ∃ n : Nat, a + Nat.succ n = b :=
  let ⟨n, h⟩ := le.dest h; ⟨n, by rwa [Int.add_comm, Int.add_left_comm] at h⟩
Existence of Natural Difference for Strict Integer Inequality: $a < b \implies \exists n \in \mathbb{N}, a + (n + 1) = b$
Informal description
For any integers $a$ and $b$ such that $a < b$, there exists a natural number $n$ such that $a + (n + 1) = b$.
Int.ofNat_lt theorem
{n m : Nat} : (↑n : Int) < ↑m ↔ n < m
Full source
@[simp, norm_cast] theorem ofNat_lt {n m : Nat} : (↑n : Int) < ↑m ↔ n < m := by
  rw [lt_iff_add_one_le, ← ofNat_succ, ofNat_le]; rfl
Preservation of Strict Order under Canonical Embedding: $n < m$ in $\mathbb{Z}$ iff $n < m$ in $\mathbb{N}$
Informal description
For any natural numbers $n$ and $m$, the canonical image of $n$ in the integers is strictly less than the canonical image of $m$ in the integers if and only if $n < m$ in the natural numbers. In other words, $n < m$ as integers if and only if $n < m$ as natural numbers.
Int.ofNat_pos theorem
{n : Nat} : 0 < (↑n : Int) ↔ 0 < n
Full source
@[simp, norm_cast] theorem ofNat_pos {n : Nat} : 0 < (↑n : Int) ↔ 0 < n := ofNat_lt
Positivity Preservation in Integer Embedding: $0 < n$ in $\mathbb{Z}$ iff $0 < n$ in $\mathbb{N}$
Informal description
For any natural number $n$, the integer obtained by casting $n$ to $\mathbb{Z}$ is strictly positive if and only if $n$ is strictly positive as a natural number, i.e., $0 < n$ in $\mathbb{Z}$ if and only if $0 < n$ in $\mathbb{N}$.
Int.ofNat_nonneg theorem
(n : Nat) : 0 ≤ (n : Int)
Full source
theorem ofNat_nonneg (n : Nat) : 0 ≤ (n : Int) := ⟨_⟩
Nonnegativity of Natural Number Cast to Integers
Informal description
For any natural number $n$, the integer obtained by casting $n$ to $\mathbb{Z}$ is nonnegative, i.e., $0 \leq n$ where $n$ is viewed as an integer.
Int.ofNat_succ_pos theorem
(n : Nat) : 0 < (succ n : Int)
Full source
theorem ofNat_succ_pos (n : Nat) : 0 < (succ n : Int) := ofNat_lt.2 <| Nat.succ_pos _
Positivity of Successor in Integers: $0 < (n + 1)$
Informal description
For any natural number $n$, the integer obtained by casting the successor of $n$ (i.e., $n + 1$) to $\mathbb{Z}$ is strictly positive, i.e., $0 < (n + 1)$ where $n + 1$ is viewed as an integer.
Int.le_refl theorem
(a : Int) : a ≤ a
Full source
@[simp] protected theorem le_refl (a : Int) : a ≤ a :=
  le.intro _ (Int.add_zero a)
Reflexivity of Integer Ordering
Informal description
For any integer $a$, the relation $a \leq a$ holds.
Int.le_trans theorem
{a b c : Int} (h₁ : a ≤ b) (h₂ : b ≤ c) : a ≤ c
Full source
protected theorem le_trans {a b c : Int} (h₁ : a ≤ b) (h₂ : b ≤ c) : a ≤ c :=
  let ⟨n, hn⟩ := le.dest h₁; let ⟨m, hm⟩ := le.dest h₂
  le.intro (n + m) <| by rw [← hm, ← hn, Int.add_assoc, ofNat_add]
Transitivity of Integer Ordering: $a \leq b \leq c \implies a \leq c$
Informal description
For any integers $a$, $b$, and $c$, if $a \leq b$ and $b \leq c$, then $a \leq c$.
Int.le_antisymm theorem
{a b : Int} (h₁ : a ≤ b) (h₂ : b ≤ a) : a = b
Full source
protected theorem le_antisymm {a b : Int} (h₁ : a ≤ b) (h₂ : b ≤ a) : a = b := by
  let ⟨n, hn⟩ := le.dest h₁; let ⟨m, hm⟩ := le.dest h₂
  have := hn; rw [← hm, Int.add_assoc, ← ofNat_add] at this
  have := Int.ofNat.inj <| Int.add_left_cancel <| this.trans (Int.add_zero _).symm
  rw [← hn, Nat.eq_zero_of_add_eq_zero_left this, ofNat_zero, Int.add_zero a]
Antisymmetry of Integer Ordering: $a \leq b \land b \leq a \implies a = b$
Informal description
For any integers $a$ and $b$, if $a \leq b$ and $b \leq a$, then $a = b$.
Int.ne_of_lt theorem
{a b : Int} (h : a < b) : a ≠ b
Full source
protected theorem ne_of_lt {a b : Int} (h : a < b) : a ≠ b := fun e => by
  cases e; exact Int.lt_irrefl _ h
Strict Inequality Implies Inequality in Integers
Informal description
For any integers $a$ and $b$, if $a < b$, then $a \neq b$.
Int.ne_of_gt theorem
{a b : Int} (h : b < a) : a ≠ b
Full source
protected theorem ne_of_gt {a b : Int} (h : b < a) : a ≠ b := (Int.ne_of_lt h).symm
Strict Inequality Implies Inequality in Integers (Reverse Order)
Informal description
For any integers $a$ and $b$, if $b < a$, then $a \neq b$.
Int.le_of_lt theorem
{a b : Int} (h : a < b) : a ≤ b
Full source
protected theorem le_of_lt {a b : Int} (h : a < b) : a ≤ b :=
  let ⟨_, hn⟩ := lt.dest h; le.intro _ hn
Strict Inequality Implies Non-Strict Inequality in Integers
Informal description
For any integers $a$ and $b$, if $a < b$, then $a \leq b$.
Int.lt_iff_le_and_ne theorem
{a b : Int} : a < b ↔ a ≤ b ∧ a ≠ b
Full source
protected theorem lt_iff_le_and_ne {a b : Int} : a < b ↔ a ≤ b ∧ a ≠ b := by
  refine ⟨fun h => ⟨Int.le_of_lt h, Int.ne_of_lt h⟩, fun ⟨aleb, aneb⟩ => ?_⟩
  let ⟨n, hn⟩ := le.dest aleb
  have : n ≠ 0 := aneb.imp fun eq => by rw [← hn, eq, ofNat_zero, Int.add_zero]
  apply lt.intro; rwa [← Nat.succ_pred_eq_of_pos (Nat.pos_of_ne_zero this)] at hn
Characterization of Strict Inequality in Integers: $a < b \iff a \leq b \land a \neq b$
Informal description
For any integers $a$ and $b$, the strict inequality $a < b$ holds if and only if $a \leq b$ and $a \neq b$.
Int.lt_succ theorem
(a : Int) : a < a + 1
Full source
protected theorem lt_succ (a : Int) : a < a + 1 := Int.le_refl _
Integer Successor Inequality: $a < a + 1$
Informal description
For any integer $a$, it holds that $a < a + 1$.
Int.zero_lt_one theorem
: (0 : Int) < 1
Full source
protected theorem zero_lt_one : (0 : Int) < 1 := ⟨_⟩
Zero is Less Than One in Integers
Informal description
The integer zero is strictly less than the integer one, i.e., $0 < 1$.
Int.lt_iff_le_not_le theorem
{a b : Int} : a < b ↔ a ≤ b ∧ ¬b ≤ a
Full source
protected theorem lt_iff_le_not_le {a b : Int} : a < b ↔ a ≤ b ∧ ¬b ≤ a := by
  rw [Int.lt_iff_le_and_ne]
  constructor <;> refine fun ⟨h, h'⟩ => ⟨h, h'.imp fun h' => ?_⟩
  · exact Int.le_antisymm h h'
  · subst h'; apply Int.le_refl
Characterization of Strict Integer Inequality via Non-Reversed Order
Informal description
For any integers $a$ and $b$, the strict inequality $a < b$ holds if and only if $a \leq b$ and it is not the case that $b \leq a$.
Int.not_le_of_gt theorem
{a b : Int} (h : b < a) : ¬a ≤ b
Full source
protected theorem not_le_of_gt {a b : Int} (h : b < a) : ¬a ≤ b :=
  (Int.lt_iff_le_not_le.mp h).right
Strict Inequality Implies Non-Reversed Order in Integers
Informal description
For any integers $a$ and $b$, if $b < a$, then it is not the case that $a \leq b$.
Int.not_lt theorem
{a b : Int} : ¬a < b ↔ b ≤ a
Full source
@[simp] protected theorem not_lt {a b : Int} : ¬a < b¬a < b ↔ b ≤ a :=
  by rw [← Int.not_le, Decidable.not_not]
Negation of Strict Inequality in Integers: $\neg(a < b) \leftrightarrow b \leq a$
Informal description
For any integers $a$ and $b$, the statement $\neg(a < b)$ is equivalent to $b \leq a$.
Int.le_of_not_gt theorem
{a b : Int} (h : ¬a > b) : a ≤ b
Full source
protected theorem le_of_not_gt {a b : Int} (h : ¬ a > b) : a ≤ b :=
  Int.not_lt.mp h
Non-greater Implies Less Than or Equal in Integers
Informal description
For any integers $a$ and $b$, if $a$ is not greater than $b$ (i.e., $\neg(a > b)$), then $a$ is less than or equal to $b$ (i.e., $a \leq b$).
Int.ne_iff_lt_or_gt theorem
{a b : Int} : a ≠ b ↔ a < b ∨ b < a
Full source
protected theorem ne_iff_lt_or_gt {a b : Int} : a ≠ ba ≠ b ↔ a < b ∨ b < a := by
  constructor
  · intro h
    cases Int.lt_trichotomy a b
    case inl lt => exact Or.inl lt
    case inr h =>
      cases h
      case inl =>simp_all
      case inr gt => exact Or.inr gt
  · intro h
    cases h
    case inl lt => exact Int.ne_of_lt lt
    case inr gt => exact Int.ne_of_gt gt
Inequality Equivalent to Strict Ordering in Integers
Informal description
For any integers $a$ and $b$, the inequality $a \neq b$ holds if and only if either $a < b$ or $b < a$.
Int.lt_or_gt_of_ne theorem
{a b : Int} : a ≠ b → a < b ∨ b < a
Full source
protected theorem lt_or_gt_of_ne {a b : Int} : a ≠ ba < b ∨ b < a:= Int.ne_iff_lt_or_gt.mp
Non-Equality Implies Strict Ordering in Integers
Informal description
For any integers $a$ and $b$, if $a \neq b$, then either $a < b$ or $b < a$.
Int.lt_of_le_of_lt theorem
{a b c : Int} (h₁ : a ≤ b) (h₂ : b < c) : a < c
Full source
protected theorem lt_of_le_of_lt {a b c : Int} (h₁ : a ≤ b) (h₂ : b < c) : a < c :=
  Int.not_le.1 fun h => Int.not_le.2 h₂ (Int.le_trans h h₁)
Transitivity of Mixed Integer Inequalities: $a \leq b < c \implies a < c$
Informal description
For any integers $a$, $b$, and $c$, if $a \leq b$ and $b < c$, then $a < c$.
Int.lt_of_lt_of_le theorem
{a b c : Int} (h₁ : a < b) (h₂ : b ≤ c) : a < c
Full source
protected theorem lt_of_lt_of_le {a b c : Int} (h₁ : a < b) (h₂ : b ≤ c) : a < c :=
  Int.not_le.1 fun h => Int.not_le.2 h₁ (Int.le_trans h₂ h)
Transitivity of Strict and Non-Strict Integer Ordering: $a < b \leq c \implies a < c$
Informal description
For any integers $a$, $b$, and $c$, if $a < b$ and $b \leq c$, then $a < c$.
Int.lt_trans theorem
{a b c : Int} (h₁ : a < b) (h₂ : b < c) : a < c
Full source
protected theorem lt_trans {a b c : Int} (h₁ : a < b) (h₂ : b < c) : a < c :=
  Int.lt_of_le_of_lt (Int.le_of_lt h₁) h₂
Transitivity of Strict Integer Ordering: $a < b < c \implies a < c$
Informal description
For any integers $a$, $b$, and $c$, if $a < b$ and $b < c$, then $a < c$.
Int.instTransLe instance
: Trans (α := Int) (· ≤ ·) (· ≤ ·) (· ≤ ·)
Full source
instance : Trans (α := Int) (· ≤ ·) (· ≤ ·) (· ≤ ·) := ⟨Int.le_trans⟩
Transitivity of Integer Ordering
Informal description
The relation $\leq$ on the integers is transitive. That is, for any integers $a$, $b$, and $c$, if $a \leq b$ and $b \leq c$, then $a \leq c$.
Int.instTransLtLe instance
: Trans (α := Int) (· < ·) (· ≤ ·) (· < ·)
Full source
instance : Trans (α := Int) (· < ·) (· ≤ ·) (· < ·) := ⟨Int.lt_of_lt_of_le⟩
Transitivity of Strict and Non-Strict Integer Ordering: $a < b \leq c \implies a < c$
Informal description
For integers, the relation $a < b$ followed by $b \leq c$ implies $a < c$.
Int.instTransLeLt instance
: Trans (α := Int) (· ≤ ·) (· < ·) (· < ·)
Full source
instance : Trans (α := Int) (· ≤ ·) (· < ·) (· < ·) := ⟨Int.lt_of_le_of_lt⟩
Transitivity of Mixed Integer Inequalities: $\leq$ and $<$ Imply $<$
Informal description
For integers, the relation $\leq$ followed by $<$ implies $<$. That is, for any integers $a$, $b$, and $c$, if $a \leq b$ and $b < c$, then $a < c$.
Int.instTransLt instance
: Trans (α := Int) (· < ·) (· < ·) (· < ·)
Full source
instance : Trans (α := Int) (· < ·) (· < ·) (· < ·) := ⟨Int.lt_trans⟩
Transitivity of Strict Order on Integers
Informal description
The strict order relation $<$ on the integers is transitive. That is, for any integers $a$, $b$, and $c$, if $a < b$ and $b < c$, then $a < c$.
Int.le_natAbs theorem
{a : Int} : a ≤ natAbs a
Full source
theorem le_natAbs {a : Int} : a ≤ natAbs a :=
  match Int.le_total 0 a with
  | .inl h => by rw [eq_natAbs_of_nonneg h]; apply Int.le_refl
  | .inr h => Int.le_trans h (ofNat_zero_le _)
Integer is Less Than or Equal to its Natural Absolute Value
Informal description
For any integer $a$, the inequality $a \leq |a|$ holds, where $|a|$ denotes the natural number absolute value of $a$.
Int.negSucc_lt_zero theorem
(n : Nat) : -[n+1] < 0
Full source
@[simp] theorem negSucc_lt_zero (n : Nat) : -[n+1] < 0 :=
  Int.not_le.1 fun h => let ⟨_, h⟩ := eq_ofNat_of_zero_le h; nomatch h
Negative Successor Integer is Less Than Zero
Informal description
For any natural number $n$, the integer $- (n + 1)$ is strictly less than zero, i.e., $- [n + 1] < 0$.
Int.negSucc_le_zero theorem
(n : Nat) : -[n+1] ≤ 0
Full source
theorem negSucc_le_zero (n : Nat) : -[n+1] ≤ 0 :=
  Int.le_of_lt (negSucc_lt_zero n)
Negative Successor Integer is Nonpositive
Informal description
For any natural number $n$, the integer $- (n + 1)$ is less than or equal to zero, i.e., $- [n + 1] \leq 0$.
Int.add_le_add_left theorem
{a b : Int} (h : a ≤ b) (c : Int) : c + a ≤ c + b
Full source
protected theorem add_le_add_left {a b : Int} (h : a ≤ b) (c : Int) : c + a ≤ c + b :=
  let ⟨n, hn⟩ := le.dest h; le.intro n <| by rw [Int.add_assoc, hn]
Left Addition Preserves Inequality in Integers
Informal description
For any integers $a$ and $b$ such that $a \leq b$, and for any integer $c$, it holds that $c + a \leq c + b$.
Int.add_le_add_right theorem
{a b : Int} (h : a ≤ b) (c : Int) : a + c ≤ b + c
Full source
protected theorem add_le_add_right {a b : Int} (h : a ≤ b) (c : Int) : a + c ≤ b + c :=
  Int.add_comm c a ▸ Int.add_comm c b ▸ Int.add_le_add_left h c
Right Addition Preserves Inequality in Integers
Informal description
For any integers $a$ and $b$ such that $a \leq b$, and for any integer $c$, it holds that $a + c \leq b + c$.
Int.add_lt_add_right theorem
{a b : Int} (h : a < b) (c : Int) : a + c < b + c
Full source
protected theorem add_lt_add_right {a b : Int} (h : a < b) (c : Int) : a + c < b + c :=
  Int.add_comm c a ▸ Int.add_comm c b ▸ Int.add_lt_add_left h c
Right Addition Preserves Strict Inequality in Integers
Informal description
For any integers $a$ and $b$ such that $a < b$, and for any integer $c$, it holds that $a + c < b + c$.
Int.le_of_add_le_add_left theorem
{a b c : Int} (h : a + b ≤ a + c) : b ≤ c
Full source
protected theorem le_of_add_le_add_left {a b c : Int} (h : a + b ≤ a + c) : b ≤ c := by
  have : -a + (a + b) ≤ -a + (a + c) := Int.add_le_add_left h _
  simp [Int.neg_add_cancel_left] at this
  assumption
Left Cancellation of Addition Preserves Order in Integers
Informal description
For any integers $a$, $b$, and $c$, if $a + b \leq a + c$, then $b \leq c$.
Int.le_of_add_le_add_right theorem
{a b c : Int} (h : a + b ≤ c + b) : a ≤ c
Full source
protected theorem le_of_add_le_add_right {a b c : Int} (h : a + b ≤ c + b) : a ≤ c :=
  Int.le_of_add_le_add_left (a := b) <| by rwa [Int.add_comm b a, Int.add_comm b c]
Right Cancellation of Addition Preserves Order in Integers
Informal description
For any integers $a$, $b$, and $c$, if $a + b \leq c + b$, then $a \leq c$.
Int.add_le_add theorem
{a b c d : Int} (h₁ : a ≤ b) (h₂ : c ≤ d) : a + c ≤ b + d
Full source
protected theorem add_le_add {a b c d : Int} (h₁ : a ≤ b) (h₂ : c ≤ d) : a + c ≤ b + d :=
  Int.le_trans (Int.add_le_add_right h₁ c) (Int.add_le_add_left h₂ b)
Addition Preserves Inequality in Integers: $a \leq b \land c \leq d \implies a + c \leq b + d$
Informal description
For any integers $a$, $b$, $c$, and $d$ such that $a \leq b$ and $c \leq d$, it holds that $a + c \leq b + d$.
Int.le_add_of_nonneg_right theorem
{a b : Int} (h : 0 ≤ b) : a ≤ a + b
Full source
protected theorem le_add_of_nonneg_right {a b : Int} (h : 0 ≤ b) : a ≤ a + b := by
  have : a + b ≥ a + 0 := Int.add_le_add_left h a
  rwa [Int.add_zero] at this
Nonnegative Addition Increases Value
Informal description
For any integers $a$ and $b$ with $b \geq 0$, we have $a \leq a + b$.
Int.le_add_of_nonneg_left theorem
{a b : Int} (h : 0 ≤ b) : a ≤ b + a
Full source
protected theorem le_add_of_nonneg_left {a b : Int} (h : 0 ≤ b) : a ≤ b + a := by
  have : 0 + a ≤ b + a := Int.add_le_add_right h a
  rwa [Int.zero_add] at this
Nonnegative Left Addition Increases Value
Informal description
For any integers $a$ and $b$ with $b \geq 0$, we have $a \leq b + a$.
Int.le_of_neg_le_neg theorem
{a b : Int} (h : -b ≤ -a) : a ≤ b
Full source
protected theorem le_of_neg_le_neg {a b : Int} (h : -b ≤ -a) : a ≤ b :=
  suffices - -a ≤ - -b by simp [Int.neg_neg] at this; assumption
  Int.neg_le_neg h
Inequality Reversal under Negation in Integers
Informal description
For any integers $a$ and $b$, if $-b \leq -a$, then $a \leq b$.
Int.neg_nonpos_of_nonneg theorem
{a : Int} (h : 0 ≤ a) : -a ≤ 0
Full source
protected theorem neg_nonpos_of_nonneg {a : Int} (h : 0 ≤ a) : -a ≤ 0 := by
  have : -a ≤ -0 := Int.neg_le_neg h
  rwa [Int.neg_zero] at this
Negation of Nonnegative Integer is Nonpositive
Informal description
For any integer $a$ such that $0 \leq a$, it holds that $-a \leq 0$.
Int.neg_nonneg_of_nonpos theorem
{a : Int} (h : a ≤ 0) : 0 ≤ -a
Full source
protected theorem neg_nonneg_of_nonpos {a : Int} (h : a ≤ 0) : 0 ≤ -a := by
  have : -0 ≤ -a := Int.neg_le_neg h
  rwa [Int.neg_zero] at this
Negation of Nonpositive Integer is Nonnegative
Informal description
For any integer $a$ such that $a \leq 0$, it holds that $0 \leq -a$.
Int.neg_neg_of_pos theorem
{a : Int} (h : 0 < a) : -a < 0
Full source
protected theorem neg_neg_of_pos {a : Int} (h : 0 < a) : -a < 0 := by
  have : -a < -0 := Int.neg_lt_neg h
  rwa [Int.neg_zero] at this
Negation of Positive Integer is Negative
Informal description
For any integer $a$ such that $0 < a$, it holds that $-a < 0$.
Int.neg_pos_of_neg theorem
{a : Int} (h : a < 0) : 0 < -a
Full source
protected theorem neg_pos_of_neg {a : Int} (h : a < 0) : 0 < -a := by
  have : -0 < -a := Int.neg_lt_neg h
  rwa [Int.neg_zero] at this
Negation of Negative Integer is Positive: $a < 0 \implies 0 < -a$
Informal description
For any integer $a$ such that $a < 0$, the negation of $a$ is positive, i.e., $0 < -a$.
Int.sub_nonneg_of_le theorem
{a b : Int} (h : b ≤ a) : 0 ≤ a - b
Full source
protected theorem sub_nonneg_of_le {a b : Int} (h : b ≤ a) : 0 ≤ a - b := by
  have h := Int.add_le_add_right h (-b)
  rwa [Int.add_right_neg] at h
Nonnegativity of Integer Difference from Inequality
Informal description
For any integers $a$ and $b$ such that $b \leq a$, the difference $a - b$ is nonnegative, i.e., $0 \leq a - b$.
Int.le_of_sub_nonneg theorem
{a b : Int} (h : 0 ≤ a - b) : b ≤ a
Full source
protected theorem le_of_sub_nonneg {a b : Int} (h : 0 ≤ a - b) : b ≤ a := by
  have h := Int.add_le_add_right h b
  rwa [Int.sub_add_cancel, Int.zero_add] at h
Nonnegative Difference Implies Inequality in Integers: $0 \leq a - b \implies b \leq a$
Informal description
For any integers $a$ and $b$, if the difference $a - b$ is nonnegative, then $b \leq a$.
Int.sub_pos_of_lt theorem
{a b : Int} (h : b < a) : 0 < a - b
Full source
protected theorem sub_pos_of_lt {a b : Int} (h : b < a) : 0 < a - b := by
  have h := Int.add_lt_add_right h (-b)
  rwa [Int.add_right_neg] at h
Positive Difference from Strict Inequality in Integers
Informal description
For any integers $a$ and $b$ such that $b < a$, the difference $a - b$ is positive, i.e., $0 < a - b$.
Int.lt_of_sub_pos theorem
{a b : Int} (h : 0 < a - b) : b < a
Full source
protected theorem lt_of_sub_pos {a b : Int} (h : 0 < a - b) : b < a := by
  have h := Int.add_lt_add_right h b
  rwa [Int.sub_add_cancel, Int.zero_add] at h
Positive Difference Implies Strict Inequality in Integers: $0 < a - b \implies b < a$
Informal description
For any integers $a$ and $b$, if the difference $a - b$ is positive, then $b$ is strictly less than $a$.
Int.sub_left_le_of_le_add theorem
{a b c : Int} (h : a ≤ b + c) : a - b ≤ c
Full source
protected theorem sub_left_le_of_le_add {a b c : Int} (h : a ≤ b + c) : a - b ≤ c := by
  have h := Int.add_le_add_right h (-b)
  rwa [Int.add_comm b c, Int.add_neg_cancel_right] at h
Subtraction Preserves Inequality: $a \leq b + c \implies a - b \leq c$
Informal description
For any integers $a$, $b$, and $c$ such that $a \leq b + c$, it holds that $a - b \leq c$.
Int.sub_le_self theorem
(a : Int) {b : Int} (h : 0 ≤ b) : a - b ≤ a
Full source
protected theorem sub_le_self (a : Int) {b : Int} (h : 0 ≤ b) : a - b ≤ a :=
  calc  a + -b
    _ ≤ a + 0 := Int.add_le_add_left (Int.neg_nonpos_of_nonneg h) _
    _ = a     := by rw [Int.add_zero]
Subtraction of Nonnegative Integer Decreases Value: $a - b \leq a$ for $b \geq 0$
Informal description
For any integer $a$ and any nonnegative integer $b$ (i.e., $0 \leq b$), the difference $a - b$ is less than or equal to $a$.
Int.sub_lt_self theorem
(a : Int) {b : Int} (h : 0 < b) : a - b < a
Full source
protected theorem sub_lt_self (a : Int) {b : Int} (h : 0 < b) : a - b < a :=
  calc  a + -b
    _ < a + 0 := Int.add_lt_add_left (Int.neg_neg_of_pos h) _
    _ = a     := by rw [Int.add_zero]
Subtraction of Positive Integer Decreases Value Strictly: $a - b < a$ for $b > 0$
Informal description
For any integer $a$ and any positive integer $b$ (i.e., $0 < b$), the difference $a - b$ is strictly less than $a$.
Int.add_one_le_of_lt theorem
{a b : Int} (H : a < b) : a + 1 ≤ b
Full source
theorem add_one_le_of_lt {a b : Int} (H : a < b) : a + 1 ≤ b := H
Integer Strict Inequality Implies Successor Inequality
Informal description
For any integers $a$ and $b$ such that $a < b$, it holds that $a + 1 \leq b$.
Int.min_def theorem
(n m : Int) : min n m = if n ≤ m then n else m
Full source
protected theorem min_def (n m : Int) : min n m = if n ≤ m then n else m := rfl
Definition of Minimum for Integers via Conditional
Informal description
For any integers $n$ and $m$, the minimum $\min(n, m)$ is equal to $n$ if $n \leq m$ and $m$ otherwise.
Int.max_def theorem
(n m : Int) : max n m = if n ≤ m then m else n
Full source
protected theorem max_def (n m : Int) : max n m = if n ≤ m then m else n := rfl
Definition of Maximum for Integers via Conditional
Informal description
For any integers $n$ and $m$, the maximum of $n$ and $m$ is equal to $m$ if $n \leq m$, and $n$ otherwise. In other words, $\max(n, m) = \begin{cases} m & \text{if } n \leq m \\ n & \text{otherwise} \end{cases}$.
Int.neg_min_neg theorem
(a b : Int) : min (-a) (-b) = -max a b
Full source
@[simp] protected theorem neg_min_neg (a b : Int) : min (-a) (-b) = -max a b := by
  rw [Int.min_def, Int.max_def]
  simp
  split <;> rename_i h₁ <;> split <;> rename_i h₂
  · simpa using Int.le_antisymm h₂ h₁
  · simp
  · simp
  · simp only [Int.not_le] at h₁ h₂
    exfalso
    exact Int.lt_irrefl _ (Int.lt_trans h₁ h₂)
Negation Swaps Minimum and Maximum: $\min(-a, -b) = -\max(a, b)$
Informal description
For any integers $a$ and $b$, the minimum of $-a$ and $-b$ is equal to the negation of the maximum of $a$ and $b$, i.e., $$\min(-a, -b) = -\max(a, b).$$
Int.min_add_right theorem
(a b c : Int) : min (a + c) (b + c) = min a b + c
Full source
@[simp] protected theorem min_add_right (a b c : Int) : min (a + c) (b + c) = min a b + c := by
  rw [Int.min_def, Int.min_def]
  simp only [Int.add_le_add_iff_right]
  split <;> simp
Minimum Preservation under Right Addition: $\min(a + c, b + c) = \min(a, b) + c$
Informal description
For any integers $a$, $b$, and $c$, the minimum of $a + c$ and $b + c$ is equal to the minimum of $a$ and $b$ plus $c$, i.e., $\min(a + c, b + c) = \min(a, b) + c$.
Int.min_add_left theorem
(a b c : Int) : min (a + b) (a + c) = a + min b c
Full source
@[simp] protected theorem min_add_left (a b c : Int) : min (a + b) (a + c) = a + min b c := by
  rw [Int.min_def, Int.min_def]
  simp only [Int.add_le_add_iff_left]
  split <;> simp
Left Additivity of Minimum for Integers: $\min(a + b, a + c) = a + \min(b, c)$
Informal description
For any integers $a, b, c$, the minimum of $a + b$ and $a + c$ is equal to $a$ plus the minimum of $b$ and $c$, i.e., $$\min(a + b, a + c) = a + \min(b, c).$$
Int.min_comm theorem
(a b : Int) : min a b = min b a
Full source
protected theorem min_comm (a b : Int) : min a b = min b a := by
  simp [Int.min_def]
  by_cases h₁ : a ≤ b <;> by_cases h₂ : b ≤ a <;> simp [h₁, h₂]
  · exact Int.le_antisymm h₁ h₂
  · cases not_or_intro h₁ h₂ <| Int.le_total ..
Commutativity of Minimum Operation on Integers: $\min(a, b) = \min(b, a)$
Informal description
For any integers $a$ and $b$, the minimum of $a$ and $b$ is equal to the minimum of $b$ and $a$, i.e., $\min(a, b) = \min(b, a)$.
Int.instCommutativeMin instance
: Std.Commutative (α := Int) min
Full source
instance : Std.Commutative (α := Int) min := ⟨Int.min_comm⟩
Commutativity of Minimum Operation on Integers
Informal description
The minimum operation $\min$ on the integers $\mathbb{Z}$ is commutative.
Int.min_le_right theorem
(a b : Int) : min a b ≤ b
Full source
protected theorem min_le_right (a b : Int) : min a b ≤ b := by rw [Int.min_def]; split <;> simp [*]
Minimum of Two Integers is Less Than or Equal to the Right Argument: $\min(a, b) \leq b$
Informal description
For any integers $a$ and $b$, the minimum of $a$ and $b$ is less than or equal to $b$, i.e., $\min(a, b) \leq b$.
Int.min_le_left theorem
(a b : Int) : min a b ≤ a
Full source
protected theorem min_le_left (a b : Int) : min a b ≤ a := Int.min_comm .. ▸ Int.min_le_right ..
Minimum of Two Integers is Less Than or Equal to the Left Argument: $\min(a, b) \leq a$
Informal description
For any integers $a$ and $b$, the minimum of $a$ and $b$ is less than or equal to $a$, i.e., $\min(a, b) \leq a$.
Int.min_eq_left theorem
{a b : Int} (h : a ≤ b) : min a b = a
Full source
protected theorem min_eq_left {a b : Int} (h : a ≤ b) : min a b = a := by simp [Int.min_def, h]
Minimum of Two Integers When Left is Smaller: $\min(a, b) = a$ if $a \leq b$
Informal description
For any integers $a$ and $b$, if $a \leq b$, then the minimum of $a$ and $b$ is equal to $a$, i.e., $\min(a, b) = a$.
Int.min_eq_right theorem
{a b : Int} (h : b ≤ a) : min a b = b
Full source
protected theorem min_eq_right {a b : Int} (h : b ≤ a) : min a b = b := by
  rw [Int.min_comm a b]; exact Int.min_eq_left h
Minimum of Two Integers When Right is Smaller: $\min(a, b) = b$ if $b \leq a$
Informal description
For any integers $a$ and $b$, if $b \leq a$, then the minimum of $a$ and $b$ is equal to $b$, i.e., $\min(a, b) = b$.
Int.le_min theorem
{a b c : Int} : a ≤ min b c ↔ a ≤ b ∧ a ≤ c
Full source
protected theorem le_min {a b c : Int} : a ≤ min b c ↔ a ≤ b ∧ a ≤ c :=
  ⟨fun h => ⟨Int.le_trans h (Int.min_le_left ..), Int.le_trans h (Int.min_le_right ..)⟩,
   fun ⟨h₁, h₂⟩ => by rw [Int.min_def]; split <;> assumption⟩
Characterization of Inequality with Minimum: $a \leq \min(b, c) \leftrightarrow a \leq b \land a \leq c$
Informal description
For any integers $a$, $b$, and $c$, the inequality $a \leq \min(b, c)$ holds if and only if both $a \leq b$ and $a \leq c$ hold.
Int.lt_min theorem
{a b c : Int} : a < min b c ↔ a < b ∧ a < c
Full source
protected theorem lt_min {a b c : Int} : a < min b c ↔ a < b ∧ a < c := Int.le_min
Characterization of Strict Inequality with Minimum: $a < \min(b, c) \leftrightarrow a < b \land a < c$
Informal description
For any integers $a$, $b$, and $c$, the inequality $a < \min(b, c)$ holds if and only if both $a < b$ and $a < c$ hold.
Int.neg_max_neg theorem
(a b : Int) : max (-a) (-b) = -min a b
Full source
@[simp] protected theorem neg_max_neg (a b : Int) : max (-a) (-b) = -min a b := by
  rw [Int.min_def, Int.max_def]
  simp
  split <;> rename_i h₁ <;> split <;> rename_i h₂
  · simpa using Int.le_antisymm h₁ h₂
  · simp
  · simp
  · simp only [Int.not_le] at h₁ h₂
    exfalso
    exact Int.lt_irrefl _ (Int.lt_trans h₁ h₂)
Maximum of Negatives Equals Negative of Minimum: $\max(-a, -b) = -\min(a, b)$
Informal description
For any integers $a$ and $b$, the maximum of $-a$ and $-b$ is equal to the negation of the minimum of $a$ and $b$, i.e., \[ \max(-a, -b) = -\min(a, b). \]
Int.max_add_right theorem
(a b c : Int) : max (a + c) (b + c) = max a b + c
Full source
@[simp] protected theorem max_add_right (a b c : Int) : max (a + c) (b + c) = max a b + c := by
  rw [Int.max_def, Int.max_def]
  simp only [Int.add_le_add_iff_right]
  split <;> simp
Right Addition Preserves Maximum for Integers: $\max(a + c, b + c) = \max(a, b) + c$
Informal description
For any integers $a$, $b$, and $c$, the maximum of $a + c$ and $b + c$ is equal to the maximum of $a$ and $b$ plus $c$, i.e., $\max(a + c, b + c) = \max(a, b) + c$.
Int.max_add_left theorem
(a b c : Int) : max (a + b) (a + c) = a + max b c
Full source
@[simp] protected theorem max_add_left (a b c : Int) : max (a + b) (a + c) = a + max b c := by
  rw [Int.max_def, Int.max_def]
  simp only [Int.add_le_add_iff_left]
  split <;> simp
Left Addition Distributes Over Maximum for Integers
Informal description
For any integers $a$, $b$, and $c$, the maximum of $a + b$ and $a + c$ is equal to $a$ plus the maximum of $b$ and $c$, i.e., \[ \max(a + b, a + c) = a + \max(b, c). \]
Int.max_comm theorem
(a b : Int) : max a b = max b a
Full source
protected theorem max_comm (a b : Int) : max a b = max b a := by
  simp only [Int.max_def]
  by_cases h₁ : a ≤ b <;> by_cases h₂ : b ≤ a <;> simp [h₁, h₂]
  · exact Int.le_antisymm h₂ h₁
  · cases not_or_intro h₁ h₂ <| Int.le_total ..
Commutativity of Maximum for Integers: $\max(a, b) = \max(b, a)$
Informal description
For any integers $a$ and $b$, the maximum of $a$ and $b$ is equal to the maximum of $b$ and $a$, i.e., $\max(a, b) = \max(b, a)$.
Int.instCommutativeMax instance
: Std.Commutative (α := Int) max
Full source
instance : Std.Commutative (α := Int) max := ⟨Int.max_comm⟩
Commutativity of Maximum Operation on Integers
Informal description
The maximum operation $\max$ on the integers $\mathbb{Z}$ is commutative.
Int.le_max_left theorem
(a b : Int) : a ≤ max a b
Full source
protected theorem le_max_left (a b : Int) : a ≤ max a b := by rw [Int.max_def]; split <;> simp [*]
$a \leq \max(a, b)$ for integers
Informal description
For any integers $a$ and $b$, the integer $a$ is less than or equal to the maximum of $a$ and $b$, i.e., $a \leq \max(a, b)$.
Int.le_max_right theorem
(a b : Int) : b ≤ max a b
Full source
protected theorem le_max_right (a b : Int) : b ≤ max a b := Int.max_comm .. ▸ Int.le_max_left ..
$b \leq \max(a, b)$ for integers
Informal description
For any integers $a$ and $b$, the integer $b$ is less than or equal to the maximum of $a$ and $b$, i.e., $b \leq \max(a, b)$.
Int.max_eq_right theorem
{a b : Int} (h : a ≤ b) : max a b = b
Full source
protected theorem max_eq_right {a b : Int} (h : a ≤ b) : max a b = b := by
  simp [Int.max_def, h, Int.not_lt.2 h]
Maximum of Two Integers When Right is Greater
Informal description
For any integers $a$ and $b$, if $a \leq b$, then the maximum of $a$ and $b$ is equal to $b$.
Int.max_eq_left theorem
{a b : Int} (h : b ≤ a) : max a b = a
Full source
protected theorem max_eq_left {a b : Int} (h : b ≤ a) : max a b = a := by
  rw [← Int.max_comm b a]; exact Int.max_eq_right h
Maximum of Two Integers When Left is Greater: $\max(a, b) = a$ if $b \leq a$
Informal description
For any integers $a$ and $b$, if $b \leq a$, then the maximum of $a$ and $b$ is equal to $a$.
Int.max_le theorem
{a b c : Int} : max a b ≤ c ↔ a ≤ c ∧ b ≤ c
Full source
protected theorem max_le {a b c : Int} : maxmax a b ≤ c ↔ a ≤ c ∧ b ≤ c :=
  ⟨fun h => ⟨Int.le_trans (Int.le_max_left ..) h, Int.le_trans (Int.le_max_right ..) h⟩,
   fun ⟨h₁, h₂⟩ => by rw [Int.max_def]; split <;> assumption⟩
Maximum of Two Integers is Bounded Above: $\max(a, b) \leq c \leftrightarrow a \leq c \land b \leq c$
Informal description
For any integers $a$, $b$, and $c$, the maximum of $a$ and $b$ is less than or equal to $c$ if and only if both $a \leq c$ and $b \leq c$ hold.
Int.max_lt theorem
{a b c : Int} : max a b < c ↔ a < c ∧ b < c
Full source
protected theorem max_lt {a b c : Int} : maxmax a b < c ↔ a < c ∧ b < c := by
  simp only [Int.lt_iff_add_one_le]
  simpa using Int.max_le (a := a + 1) (b := b + 1) (c := c)
Maximum of Two Integers is Strictly Bounded Above: $\max(a, b) < c \leftrightarrow a < c \land b < c$
Informal description
For any integers $a$, $b$, and $c$, the maximum of $a$ and $b$ is strictly less than $c$ if and only if both $a < c$ and $b < c$ hold.
Int.ofNat_max_zero theorem
(n : Nat) : (max (n : Int) 0) = n
Full source
@[simp] theorem ofNat_max_zero (n : Nat) : (max (n : Int) 0) = n := by
  rw [Int.max_eq_left (ofNat_zero_le n)]
Maximum of Natural Number and Zero in Integers is the Number Itself
Informal description
For any natural number $n$, the maximum of $n$ (considered as an integer) and $0$ is equal to $n$, i.e., $\max(n, 0) = n$.
Int.zero_max_ofNat theorem
(n : Nat) : (max 0 (n : Int)) = n
Full source
@[simp] theorem zero_max_ofNat (n : Nat) : (max 0 (n : Int)) = n := by
  rw [Int.max_eq_right (ofNat_zero_le n)]
Maximum of Zero and Natural Number as Integer is the Number Itself
Informal description
For any natural number $n$, the maximum of $0$ and $n$ (considered as an integer) equals $n$, i.e., $\max(0, n) = n$.
Int.negSucc_max_zero theorem
(n : Nat) : (max (Int.negSucc n) 0) = 0
Full source
@[simp] theorem negSucc_max_zero (n : Nat) : (max (Int.negSucc n) 0) = 0 := by
  rw [Int.max_eq_right (negSucc_le_zero _)]
Maximum of Negative Successor and Zero is Zero
Informal description
For any natural number $n$, the maximum of the negative successor integer $- (n + 1)$ and $0$ is equal to $0$, i.e., $\max(-[n+1], 0) = 0$.
Int.zero_max_negSucc theorem
(n : Nat) : (max 0 (Int.negSucc n)) = 0
Full source
@[simp] theorem zero_max_negSucc (n : Nat) : (max 0 (Int.negSucc n)) = 0 := by
  rw [Int.max_eq_left (negSucc_le_zero _)]
Maximum of Zero and Negative Successor Integer is Zero
Informal description
For any natural number $n$, the maximum of $0$ and the negative successor integer $- (n + 1)$ is equal to $0$, i.e., $\max(0, -[n + 1]) = 0$.
Int.min_self theorem
(a : Int) : min a a = a
Full source
@[simp] protected theorem min_self (a : Int) : min a a = a := Int.min_eq_left (Int.le_refl _)
Idempotence of Minimum Operation on Integers: $\min(a, a) = a$
Informal description
For any integer $a$, the minimum of $a$ with itself is equal to $a$, i.e., $\min(a, a) = a$.
Int.instIdempotentOpMin instance
: Std.IdempotentOp (α := Int) min
Full source
instance : Std.IdempotentOp (α := Int) min := ⟨Int.min_self⟩
Idempotence of Minimum Operation on Integers
Informal description
The minimum operation $\min$ on the integers is idempotent, meaning that for any integer $a$, $\min(a, a) = a$.
Int.max_self theorem
(a : Int) : max a a = a
Full source
@[simp] protected theorem max_self (a : Int) : max a a = a := Int.max_eq_right (Int.le_refl _)
Idempotence of Maximum Operation on Integers
Informal description
For any integer $a$, the maximum of $a$ with itself is equal to $a$, i.e., $\max(a, a) = a$.
Int.instIdempotentOpMax instance
: Std.IdempotentOp (α := Int) max
Full source
instance : Std.IdempotentOp (α := Int) max := ⟨Int.max_self⟩
Idempotence of Maximum Operation on Integers
Informal description
The maximum operation $\max$ on the integers is idempotent, meaning that for any integer $a$, $\max(a, a) = a$.
Int.mul_nonneg theorem
{a b : Int} (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a * b
Full source
protected theorem mul_nonneg {a b : Int} (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a * b := by
  let ⟨n, hn⟩ := eq_ofNat_of_zero_le ha
  let ⟨m, hm⟩ := eq_ofNat_of_zero_le hb
  rw [hn, hm, ← ofNat_mul]; apply ofNat_nonneg
Nonnegativity of Product of Nonnegative Integers
Informal description
For any integers $a$ and $b$ such that $0 \leq a$ and $0 \leq b$, the product $a \cdot b$ satisfies $0 \leq a \cdot b$.
Int.mul_pos theorem
{a b : Int} (ha : 0 < a) (hb : 0 < b) : 0 < a * b
Full source
protected theorem mul_pos {a b : Int} (ha : 0 < a) (hb : 0 < b) : 0 < a * b := by
  let ⟨n, hn⟩ := eq_succ_of_zero_lt ha
  let ⟨m, hm⟩ := eq_succ_of_zero_lt hb
  rw [hn, hm]
  apply ofNat_succ_pos
Positivity of Product of Positive Integers
Informal description
For any integers $a$ and $b$ such that $0 < a$ and $0 < b$, the product $a \cdot b$ satisfies $0 < a \cdot b$.
Int.mul_lt_mul_of_pos_left theorem
{a b c : Int} (h₁ : a < b) (h₂ : 0 < c) : c * a < c * b
Full source
protected theorem mul_lt_mul_of_pos_left {a b c : Int}
  (h₁ : a < b) (h₂ : 0 < c) : c * a < c * b := by
  have : 0 < c * (b - a) := Int.mul_pos h₂ (Int.sub_pos_of_lt h₁)
  rw [Int.mul_sub] at this
  exact Int.lt_of_sub_pos this
Left Multiplication Preserves Strict Inequality for Positive Integers: $a < b \land c > 0 \implies c \cdot a < c \cdot b$
Informal description
For any integers $a$, $b$, and $c$ such that $a < b$ and $0 < c$, the product $c \cdot a$ is strictly less than $c \cdot b$.
Int.mul_lt_mul_of_pos_right theorem
{a b c : Int} (h₁ : a < b) (h₂ : 0 < c) : a * c < b * c
Full source
protected theorem mul_lt_mul_of_pos_right {a b c : Int}
  (h₁ : a < b) (h₂ : 0 < c) : a * c < b * c := by
  have : 0 < b - a := Int.sub_pos_of_lt h₁
  have : 0 < (b - a) * c := Int.mul_pos this h₂
  rw [Int.sub_mul] at this
  exact Int.lt_of_sub_pos this
Right Multiplication Preserves Strict Inequality for Positive Integers: $a < b \land 0 < c \implies a \cdot c < b \cdot c$
Informal description
For any integers $a$, $b$, and $c$ such that $a < b$ and $0 < c$, the product $a \cdot c$ is strictly less than $b \cdot c$.
Int.mul_le_mul_of_nonneg_left theorem
{a b c : Int} (h₁ : a ≤ b) (h₂ : 0 ≤ c) : c * a ≤ c * b
Full source
protected theorem mul_le_mul_of_nonneg_left {a b c : Int}
    (h₁ : a ≤ b) (h₂ : 0 ≤ c) : c * a ≤ c * b :=
  if hba : b ≤ a then by
    rw [Int.le_antisymm hba h₁]; apply Int.le_refl
  else if hc0 : c ≤ 0 then by
    simp [Int.le_antisymm hc0 h₂, Int.zero_mul]
  else by
    exact Int.le_of_lt <| Int.mul_lt_mul_of_pos_left
      (Int.lt_iff_le_not_le.2 ⟨h₁, hba⟩) (Int.lt_iff_le_not_le.2 ⟨h₂, hc0⟩)
Left Multiplication Preserves Non-Strict Inequality for Non-Negative Integers: $a \leq b \land c \geq 0 \implies c \cdot a \leq c \cdot b$
Informal description
For any integers $a$, $b$, and $c$ such that $a \leq b$ and $0 \leq c$, the product $c \cdot a$ is less than or equal to $c \cdot b$.
Int.mul_le_mul_of_nonneg_right theorem
{a b c : Int} (h₁ : a ≤ b) (h₂ : 0 ≤ c) : a * c ≤ b * c
Full source
protected theorem mul_le_mul_of_nonneg_right {a b c : Int}
    (h₁ : a ≤ b) (h₂ : 0 ≤ c) : a * c ≤ b * c := by
  rw [Int.mul_comm, Int.mul_comm b]; exact Int.mul_le_mul_of_nonneg_left h₁ h₂
Right Multiplication Preserves Non-Strict Inequality for Non-Negative Integers: $a \leq b \land c \geq 0 \implies a \cdot c \leq b \cdot c$
Informal description
For any integers $a$, $b$, and $c$ such that $a \leq b$ and $0 \leq c$, the product $a \cdot c$ is less than or equal to $b \cdot c$.
Int.mul_le_mul theorem
{a b c d : Int} (hac : a ≤ c) (hbd : b ≤ d) (nn_b : 0 ≤ b) (nn_c : 0 ≤ c) : a * b ≤ c * d
Full source
protected theorem mul_le_mul {a b c d : Int}
    (hac : a ≤ c) (hbd : b ≤ d) (nn_b : 0 ≤ b) (nn_c : 0 ≤ c) : a * b ≤ c * d :=
  Int.le_trans (Int.mul_le_mul_of_nonneg_right hac nn_b) (Int.mul_le_mul_of_nonneg_left hbd nn_c)
Multiplication Preserves Non-Strict Inequalities for Non-Negative Integers: $a \leq c \land b \leq d \land b \geq 0 \land c \geq 0 \implies a \cdot b \leq c \cdot d$
Informal description
For any integers $a$, $b$, $c$, and $d$ such that $a \leq c$, $b \leq d$, $b \geq 0$, and $c \geq 0$, the product $a \cdot b$ is less than or equal to $c \cdot d$.
Int.mul_nonpos_of_nonneg_of_nonpos theorem
{a b : Int} (ha : 0 ≤ a) (hb : b ≤ 0) : a * b ≤ 0
Full source
protected theorem mul_nonpos_of_nonneg_of_nonpos {a b : Int}
  (ha : 0 ≤ a) (hb : b ≤ 0) : a * b ≤ 0 := by
  have h : a * b ≤ a * 0 := Int.mul_le_mul_of_nonneg_left hb ha
  rwa [Int.mul_zero] at h
Nonnegative Times Nonpositive is Nonpositive in Integers: $a \geq 0 \land b \leq 0 \implies a \cdot b \leq 0$
Informal description
For any integers $a$ and $b$ such that $a \geq 0$ and $b \leq 0$, the product $a \cdot b$ is less than or equal to $0$.
Int.mul_nonpos_of_nonpos_of_nonneg theorem
{a b : Int} (ha : a ≤ 0) (hb : 0 ≤ b) : a * b ≤ 0
Full source
protected theorem mul_nonpos_of_nonpos_of_nonneg {a b : Int}
  (ha : a ≤ 0) (hb : 0 ≤ b) : a * b ≤ 0 := by
  have h : a * b ≤ 0 * b := Int.mul_le_mul_of_nonneg_right ha hb
  rwa [Int.zero_mul] at h
Nonpositive Times Nonnegative is Nonpositive in Integers: $a \leq 0 \land b \geq 0 \implies a \cdot b \leq 0$
Informal description
For any integers $a$ and $b$ such that $a \leq 0$ and $0 \leq b$, the product $a \cdot b$ is less than or equal to $0$.
Int.mul_le_mul_of_nonpos_right theorem
{a b c : Int} (h : b ≤ a) (hc : c ≤ 0) : a * c ≤ b * c
Full source
protected theorem mul_le_mul_of_nonpos_right {a b c : Int}
    (h : b ≤ a) (hc : c ≤ 0) : a * c ≤ b * c :=
  have : -c ≥ 0 := Int.neg_nonneg_of_nonpos hc
  have : b * -c ≤ a * -c := Int.mul_le_mul_of_nonneg_right h this
  Int.le_of_neg_le_neg <| by rwa [← Int.neg_mul_eq_mul_neg, ← Int.neg_mul_eq_mul_neg] at this
Right Multiplication by Nonpositive Integer Reverses Inequality: $b \leq a \land c \leq 0 \implies a \cdot c \leq b \cdot c$
Informal description
For any integers $a$, $b$, and $c$ such that $b \leq a$ and $c \leq 0$, the product $a \cdot c$ is less than or equal to $b \cdot c$.
Int.mul_le_mul_of_nonpos_left theorem
{a b c : Int} (ha : a ≤ 0) (h : c ≤ b) : a * b ≤ a * c
Full source
protected theorem mul_le_mul_of_nonpos_left {a b c : Int}
    (ha : a ≤ 0) (h : c ≤ b) : a * b ≤ a * c := by
  rw [Int.mul_comm a b, Int.mul_comm a c]
  apply Int.mul_le_mul_of_nonpos_right h ha
Left Multiplication by Nonpositive Integer Reverses Inequality: $a \leq 0 \land c \leq b \implies a \cdot b \leq a \cdot c$
Informal description
For any integers $a$, $b$, and $c$ such that $a \leq 0$ and $c \leq b$, the product $a \cdot b$ is less than or equal to $a \cdot c$.
Int.natAbs_ofNat theorem
(n : Nat) : natAbs ↑n = n
Full source
@[simp, norm_cast] theorem natAbs_ofNat (n : Nat) : natAbs ↑n = n := rfl
Absolute Value of Natural Number Cast to Integer Equals Itself
Informal description
For any natural number $n$, the absolute value (as a natural number) of the integer obtained by casting $n$ to $\mathbb{Z}$ is equal to $n$ itself, i.e., $\text{natAbs}(n) = n$.
Int.natAbs_negSucc theorem
(n : Nat) : natAbs -[n+1] = n.succ
Full source
@[simp] theorem natAbs_negSucc (n : Nat) : natAbs -[n+1] = n.succ := rfl
Absolute Value of Negative Successor Integer Equals Successor
Informal description
For any natural number $n$, the absolute value (as a natural number) of the negative successor integer $-n-1$ is equal to the successor of $n$, i.e., $\text{natAbs}(-n-1) = n + 1$.
Int.natAbs_zero theorem
: natAbs (0 : Int) = (0 : Nat)
Full source
@[simp] theorem natAbs_zero : natAbs (0 : Int) = (0 : Nat) := rfl
Absolute Value of Integer Zero Equals Zero
Informal description
The absolute value (as a natural number) of the integer $0$ is equal to the natural number $0$, i.e., $\text{natAbs}(0) = 0$.
Int.natAbs_one theorem
: natAbs (1 : Int) = (1 : Nat)
Full source
@[simp] theorem natAbs_one : natAbs (1 : Int) = (1 : Nat) := rfl
Absolute Value of Integer One Equals One
Informal description
The absolute value (as a natural number) of the integer $1$ is equal to the natural number $1$, i.e., $\text{natAbs}(1) = 1$.
Int.natAbs_pos theorem
: 0 < natAbs a ↔ a ≠ 0
Full source
theorem natAbs_pos : 0 < natAbs a ↔ a ≠ 0 := by rw [Nat.pos_iff_ne_zero, Ne, natAbs_eq_zero]
Positivity of Integer Absolute Value Criterion: $0 < \text{natAbs}(a) \leftrightarrow a \neq 0$
Informal description
For any integer $a$, the natural number absolute value of $a$ is positive if and only if $a$ is non-zero, i.e., $0 < \text{natAbs}(a) \leftrightarrow a \neq 0$.
Int.natAbs_neg theorem
: ∀ (a : Int), natAbs (-a) = natAbs a
Full source
@[simp] theorem natAbs_neg : ∀ (a : Int), natAbs (-a) = natAbs a
  | 0      => rfl
  | succ _ => rfl
  | -[_+1] => rfl
Absolute Value of Negative Integer Equals Absolute Value of Original Integer
Informal description
For any integer $a$, the absolute value of $-a$ as a natural number is equal to the absolute value of $a$ as a natural number, i.e., $\text{natAbs}(-a) = \text{natAbs}(a)$.
Int.natAbs_eq theorem
: ∀ (a : Int), a = natAbs a ∨ a = -↑(natAbs a)
Full source
theorem natAbs_eq : ∀ (a : Int), a = natAbs a ∨ a = -↑(natAbs a)
  | ofNat _ => Or.inl rfl
  | -[_+1]  => Or.inr rfl
Integer as Absolute Value or Its Negation
Informal description
For any integer $a$, either $a$ is equal to its natural absolute value $\text{natAbs}(a)$, or $a$ is equal to the negation of its natural absolute value, i.e., $a = \text{natAbs}(a) \lor a = -\text{natAbs}(a)$.
Int.natAbs_negOfNat theorem
(n : Nat) : natAbs (negOfNat n) = n
Full source
theorem natAbs_negOfNat (n : Nat) : natAbs (negOfNat n) = n := by
  cases n <;> rfl
Absolute Value of Negated Natural Number Equals Original
Informal description
For any natural number $n$, the absolute value (as a natural number) of the integer obtained by negating $n$ is equal to $n$ itself, i.e., $\text{natAbs}(\text{negOfNat}(n)) = n$.
Int.natAbs_mul theorem
(a b : Int) : natAbs (a * b) = natAbs a * natAbs b
Full source
theorem natAbs_mul (a b : Int) : natAbs (a * b) = natAbs a * natAbs b := by
  cases a <;> cases b <;>
    simp only [← Int.mul_def, Int.mul, natAbs_negOfNat] <;> simp only [natAbs]
Absolute Value of Integer Product: $\text{natAbs}(a \cdot b) = \text{natAbs}(a) \cdot \text{natAbs}(b)$
Informal description
For any integers $a$ and $b$, the absolute value of their product equals the product of their absolute values, i.e., $\text{natAbs}(a \cdot b) = \text{natAbs}(a) \cdot \text{natAbs}(b)$.
Int.natAbs_eq_natAbs_iff theorem
{a b : Int} : a.natAbs = b.natAbs ↔ a = b ∨ a = -b
Full source
theorem natAbs_eq_natAbs_iff {a b : Int} : a.natAbs = b.natAbs ↔ a = b ∨ a = -b := by
  constructor <;> intro h
  · cases Int.natAbs_eq a with
    | inl h₁ | inr h₁ =>
      cases Int.natAbs_eq b with
      | inl h₂ | inr h₂ => rw [h₁, h₂]; simp [h]
  · cases h with (subst a; try rfl)
    | inr h => rw [Int.natAbs_neg]
Equality of Integer Absolute Values: $\text{natAbs}(a) = \text{natAbs}(b) \leftrightarrow a = b \lor a = -b$
Informal description
For any integers $a$ and $b$, the natural number absolute values of $a$ and $b$ are equal if and only if $a$ equals $b$ or $a$ equals $-b$. That is, $\text{natAbs}(a) = \text{natAbs}(b) \leftrightarrow (a = b \lor a = -b)$.
Int.natAbs_of_nonneg theorem
{a : Int} (H : 0 ≤ a) : (natAbs a : Int) = a
Full source
theorem natAbs_of_nonneg {a : Int} (H : 0 ≤ a) : (natAbs a : Int) = a :=
  match a, eq_ofNat_of_zero_le H with
  | _, ⟨_, rfl⟩ => rfl
Nonnegative Integer Absolute Value Identity
Informal description
For any integer $a$ such that $0 \leq a$, the canonical image of the natural number absolute value of $a$ in the integers is equal to $a$ itself, i.e., $\text{natAbs}(a) = a$.
Int.ofNat_natAbs_of_nonpos theorem
{a : Int} (H : a ≤ 0) : (natAbs a : Int) = -a
Full source
theorem ofNat_natAbs_of_nonpos {a : Int} (H : a ≤ 0) : (natAbs a : Int) = -a := by
  rw [← natAbs_neg, natAbs_of_nonneg (Int.neg_nonneg_of_nonpos H)]
Absolute Value Identity for Nonpositive Integers: $\text{natAbs}(a) = -a$ when $a \leq 0$
Informal description
For any integer $a$ such that $a \leq 0$, the canonical image of the natural number absolute value of $a$ in the integers is equal to $-a$, i.e., $\text{natAbs}(a) = -a$.
Int.natAbs_sub_of_nonneg_of_le theorem
{a b : Int} (h₁ : 0 ≤ b) (h₂ : b ≤ a) : (a - b).natAbs = a.natAbs - b.natAbs
Full source
theorem natAbs_sub_of_nonneg_of_le {a b : Int} (h₁ : 0 ≤ b) (h₂ : b ≤ a) :
    (a - b).natAbs = a.natAbs - b.natAbs := by
  rw [← Int.ofNat_inj]
  rw [natAbs_of_nonneg, ofNat_sub, natAbs_of_nonneg (Int.le_trans h₁ h₂), natAbs_of_nonneg h₁]
  · rwa [← Int.ofNat_le, natAbs_of_nonneg h₁, natAbs_of_nonneg (Int.le_trans h₁ h₂)]
  · exact Int.sub_nonneg_of_le h₂
Absolute Value of Nonnegative Integer Difference: $\text{natAbs}(a - b) = \text{natAbs}(a) - \text{natAbs}(b)$ under $0 \leq b \leq a$
Informal description
For any integers $a$ and $b$ such that $0 \leq b$ and $b \leq a$, the natural number absolute value of their difference equals the difference of their natural number absolute values, i.e., $\text{natAbs}(a - b) = \text{natAbs}(a) - \text{natAbs}(b)$.
Int.toNat_eq_max theorem
: ∀ a : Int, (toNat a : Int) = max a 0
Full source
theorem toNat_eq_max : ∀ a : Int, (toNat a : Int) = max a 0
  | (n : Nat) => (Int.max_eq_left (ofNat_zero_le n)).symm
  | -[n+1] => (Int.max_eq_right (Int.le_of_lt (negSucc_lt_zero n))).symm
Natural Number Conversion as Maximum with Zero: $\text{toNat}(a) = \max(a, 0)$
Informal description
For any integer $a$, the natural number obtained by applying the `toNat` function to $a$ is equal to the maximum of $a$ and $0$ when viewed as an integer, i.e., $\text{toNat}(a) = \max(a, 0)$.
Int.toNat_zero theorem
: (0 : Int).toNat = 0
Full source
@[simp] theorem toNat_zero : (0 : Int).toNat = 0 := rfl
$\text{toNat}(0) = 0$
Informal description
The natural number obtained by applying the `toNat` function to the integer $0$ is equal to $0$, i.e., $\text{toNat}(0) = 0$.
Int.toNat_one theorem
: (1 : Int).toNat = 1
Full source
@[simp] theorem toNat_one : (1 : Int).toNat = 1 := rfl
$\text{toNat}(1) = 1$
Informal description
The natural number obtained by applying the `toNat` function to the integer $1$ is equal to $1$, i.e., $\text{toNat}(1) = 1$.
Int.toNat_of_nonneg theorem
{a : Int} (h : 0 ≤ a) : (toNat a : Int) = a
Full source
theorem toNat_of_nonneg {a : Int} (h : 0 ≤ a) : (toNat a : Int) = a := by
  rw [toNat_eq_max, Int.max_eq_left h]
Natural Number Conversion Preserves Nonnegative Integers: $\text{toNat}(a) = a$ for $0 \leq a$
Informal description
For any integer $a$ such that $0 \leq a$, the natural number obtained by applying the `toNat` function to $a$ is equal to $a$ when viewed as an integer, i.e., $\text{toNat}(a) = a$.
Int.toNat_ofNat theorem
(n : Nat) : toNat ↑n = n
Full source
@[simp] theorem toNat_ofNat (n : Nat) : toNat ↑n = n := rfl
Preservation of Natural Numbers under Integer Conversion
Informal description
For any natural number $n$, the conversion of $n$ to an integer and then back to a natural number via the `toNat` function returns $n$ itself, i.e., $\text{toNat}(n) = n$.
Int.toNat_negSucc theorem
(n : Nat) : (Int.negSucc n).toNat = 0
Full source
@[simp] theorem toNat_negSucc (n : Nat) : (Int.negSucc n).toNat = 0 := by
  simp [toNat]
Conversion of Negative Successor to Natural Number Yields Zero
Informal description
For any natural number $n$, the conversion of the negative successor $-n-1$ (represented as `Int.negSucc n`) to a natural number via the `toNat` function yields $0$.
Int.toNat_ofNat_add_one theorem
{n : Nat} : ((n : Int) + 1).toNat = n + 1
Full source
@[simp] theorem toNat_ofNat_add_one {n : Nat} : ((n : Int) + 1).toNat = n + 1 := rfl
Preservation of Successor under Integer-to-Natural Conversion
Informal description
For any natural number $n$, the conversion of the integer $(n + 1)$ to a natural number via the `toNat` function yields $n + 1$, i.e., $\text{toNat}(n + 1) = n + 1$.
Int.ofNat_toNat theorem
(a : Int) : (a.toNat : Int) = max a 0
Full source
@[simp] theorem ofNat_toNat (a : Int) : (a.toNat : Int) = max a 0 := by
  match a with
  | (n : Nat) => simp
  | -(n + 1 : Nat) => norm_cast
Natural Part of Integer Equals Maximum with Zero
Informal description
For any integer $a$, the canonical embedding of its natural number part $a.\text{toNat}$ back into the integers is equal to the maximum of $a$ and $0$, i.e., $(a.\text{toNat} : \mathbb{Z}) = \max(a, 0)$.
Int.self_le_toNat theorem
(a : Int) : a ≤ toNat a
Full source
theorem self_le_toNat (a : Int) : a ≤ toNat a := by rw [toNat_eq_max]; apply Int.le_max_left
Integer is Less Than or Equal to its Natural Truncation
Informal description
For any integer $a$, we have $a \leq \text{toNat}(a)$, where $\text{toNat}(a)$ is the natural number obtained by truncating $a$ (i.e., taking the non-negative part of $a$).
Int.le_toNat theorem
{n : Nat} {z : Int} (h : 0 ≤ z) : n ≤ z.toNat ↔ (n : Int) ≤ z
Full source
@[simp] theorem le_toNat {n : Nat} {z : Int} (h : 0 ≤ z) : n ≤ z.toNat ↔ (n : Int) ≤ z := by
  rw [← Int.ofNat_le, Int.toNat_of_nonneg h]
Natural Number Comparison via Integer Conversion: $n \leq z.\text{toNat} \leftrightarrow n \leq z$ for $0 \leq z$
Informal description
For any natural number $n$ and any non-negative integer $z$, the inequality $n \leq z.\text{toNat}$ holds if and only if the integer $n$ is less than or equal to $z$ when viewed as an integer.
Int.toNat_lt theorem
{n : Nat} {z : Int} (h : 0 ≤ z) : z.toNat < n ↔ z < (n : Int)
Full source
@[simp] theorem toNat_lt {n : Nat} {z : Int} (h : 0 ≤ z) : z.toNat < n ↔ z < (n : Int) := by
  rw [← Int.not_le, ← Nat.not_le, Int.le_toNat h]
Natural Truncation Preserves Strict Inequality for Non-Negative Integers: $z.\text{toNat} < n \leftrightarrow z < n$
Informal description
For any natural number $n$ and any non-negative integer $z$, the inequality $z.\text{toNat} < n$ holds if and only if $z < n$ when viewed as an integer.
Int.toNat_add theorem
{a b : Int} (ha : 0 ≤ a) (hb : 0 ≤ b) : (a + b).toNat = a.toNat + b.toNat
Full source
theorem toNat_add {a b : Int} (ha : 0 ≤ a) (hb : 0 ≤ b) : (a + b).toNat = a.toNat + b.toNat :=
  match a, b, eq_ofNat_of_zero_le ha, eq_ofNat_of_zero_le hb with
  | _, _, ⟨_, rfl⟩, ⟨_, rfl⟩ => rfl
Conversion of Integer Addition to Natural Numbers for Non-Negative Integers
Informal description
For any non-negative integers $a$ and $b$, the natural number obtained by converting the sum $a + b$ (as integers) to a natural number equals the sum of the natural number conversions of $a$ and $b$ individually. That is, $(a + b).\text{toNat} = a.\text{toNat} + b.\text{toNat}$.
Int.toNat_add_nat theorem
{a : Int} (ha : 0 ≤ a) (n : Nat) : (a + n).toNat = a.toNat + n
Full source
theorem toNat_add_nat {a : Int} (ha : 0 ≤ a) (n : Nat) : (a + n).toNat = a.toNat + n :=
  match a, eq_ofNat_of_zero_le ha with | _, ⟨_, rfl⟩ => rfl
Natural Number Correspondence Preserves Addition with Natural Numbers
Informal description
For any integer $a \geq 0$ and natural number $n$, the natural number corresponding to $a + n$ is equal to the sum of the natural numbers corresponding to $a$ and $n$, i.e., $(a + n).\text{toNat} = a.\text{toNat} + n$.
Int.pred_toNat theorem
: ∀ i : Int, (i - 1).toNat = i.toNat - 1
Full source
@[simp] theorem pred_toNat : ∀ i : Int, (i - 1).toNat = i.toNat - 1
  | 0 => rfl
  | (_+1:Nat) => by simp [ofNat_add]
  | -[_+1] => rfl
Natural Conversion of Integer Predecessor: $(i - 1).\text{toNat} = i.\text{toNat} - 1$
Informal description
For any integer $i$, the natural number conversion of $i - 1$ equals the natural number conversion of $i$ minus one, i.e., $(i - 1).\text{toNat} = i.\text{toNat} - 1$.
Int.toNat_sub_toNat_neg theorem
: ∀ n : Int, ↑n.toNat - ↑(-n).toNat = n
Full source
theorem toNat_sub_toNat_neg : ∀ n : Int, ↑n.toNat - ↑(-n).toNat = n
  | 0 => rfl
  | (_+1:Nat) => Int.sub_zero _
  | -[_+1] => Int.zero_sub _
Difference of Natural Conversions Equals Original Integer
Informal description
For any integer $n$, the difference between the canonical natural number conversion of $n$ and the canonical natural number conversion of $-n$ equals $n$, i.e., $n.\text{toNat} - (-n).\text{toNat} = n$.
Int.toNat_add_toNat_neg_eq_natAbs theorem
: ∀ n : Int, n.toNat + (-n).toNat = n.natAbs
Full source
@[simp] theorem toNat_add_toNat_neg_eq_natAbs : ∀ n : Int, n.toNat + (-n).toNat = n.natAbs
  | 0 => rfl
  | (_+1:Nat) => Nat.add_zero _
  | -[_+1] => Nat.zero_add _
Sum of Natural Conversions Equals Natural Absolute Value
Informal description
For any integer $n$, the sum of the natural number conversions of $n$ and $-n$ equals the natural absolute value of $n$, i.e., $\text{toNat}(n) + \text{toNat}(-n) = \text{natAbs}(n)$.
Int.toNat_neg_nat theorem
: ∀ n : Nat, (-(n : Int)).toNat = 0
Full source
@[simp] theorem toNat_neg_nat : ∀ n : Nat, (-(n : Int)).toNat = 0
  | 0 => rfl
  | _+1 => rfl
Conversion of Negative Natural Number to Natural Number Yields Zero
Informal description
For any natural number $n$, the canonical conversion of the negative integer $-n$ to a natural number via `toNat` yields zero, i.e., $(-n).\text{toNat} = 0$.
Int.mem_toNat? theorem
: ∀ {a : Int} {n : Nat}, toNat? a = some n ↔ a = n
Full source
theorem mem_toNat? : ∀ {a : Int} {n : Nat}, toNat?toNat? a = some n ↔ a = n
  | (m : Nat), n => by simp [toNat?, Int.ofNat_inj]
  | -[m+1], n => by constructor <;> nofun
Characterization of Integer-to-Natural Conversion: $\text{toNat?}(a) = \text{some }n \leftrightarrow a = n$
Informal description
For any integer $a$ and natural number $n$, the optional natural number conversion `toNat? a` equals `some n` if and only if $a$ is equal to $n$.
Int.mem_toNat' abbrev
Full source
@[deprecated mem_toNat? (since := "2025-03-11")]
abbrev mem_toNat' := @mem_toNat?
Characterization of Integer-to-Natural Conversion via `toNat'`: $\text{toNat'}(a) = n \leftrightarrow a = n$
Informal description
For any integer $a$ and natural number $n$, the conversion of $a$ to a natural number via `toNat'` equals $n$ if and only if $a$ is equal to $n$.
Int.le_of_not_le theorem
{a b : Int} : ¬a ≤ b → b ≤ a
Full source
protected theorem le_of_not_le {a b : Int} : ¬ a ≤ b → b ≤ a := (Int.le_total a b).resolve_left
Negation of Inequality Implies Reverse Inequality in Integers
Informal description
For any integers $a$ and $b$, if $a$ is not less than or equal to $b$, then $b$ is less than or equal to $a$.
Int.negSucc_not_pos theorem
(n : Nat) : 0 < -[n+1] ↔ False
Full source
@[simp] theorem negSucc_not_pos (n : Nat) : 0 < -[n+1] ↔ False := by
  simp only [Int.not_lt, iff_false]; constructor
Negative Successor is Not Positive: $0 < -[n+1] \leftrightarrow \text{False}$
Informal description
For any natural number $n$, the negative successor $- [n+1]$ is not positive, i.e., the inequality $0 < -[n+1]$ is equivalent to $\text{False}$.
Int.lt_of_add_lt_add_left theorem
{a b c : Int} (h : a + b < a + c) : b < c
Full source
protected theorem lt_of_add_lt_add_left {a b c : Int} (h : a + b < a + c) : b < c := by
  have : -a + (a + b) < -a + (a + c) := Int.add_lt_add_left h _
  simp [Int.neg_add_cancel_left] at this
  assumption
Left Cancellation of Addition Preserves Strict Inequality in Integers
Informal description
For any integers $a$, $b$, and $c$, if $a + b < a + c$, then $b < c$.
Int.lt_of_add_lt_add_right theorem
{a b c : Int} (h : a + b < c + b) : a < c
Full source
protected theorem lt_of_add_lt_add_right {a b c : Int} (h : a + b < c + b) : a < c :=
  Int.lt_of_add_lt_add_left (a := b) <| by rwa [Int.add_comm b a, Int.add_comm b c]
Right Cancellation of Addition Preserves Strict Inequality in Integers
Informal description
For any integers $a$, $b$, and $c$, if $a + b < c + b$, then $a < c$.
Int.add_lt_add theorem
{a b c d : Int} (h₁ : a < b) (h₂ : c < d) : a + c < b + d
Full source
protected theorem add_lt_add {a b c d : Int} (h₁ : a < b) (h₂ : c < d) : a + c < b + d :=
  Int.lt_trans (Int.add_lt_add_right h₁ c) (Int.add_lt_add_left h₂ b)
Addition Preserves Strict Inequalities in Integers: $a < b$ and $c < d$ implies $a + c < b + d$
Informal description
For any integers $a$, $b$, $c$, and $d$ such that $a < b$ and $c < d$, it holds that $a + c < b + d$.
Int.add_lt_add_of_le_of_lt theorem
{a b c d : Int} (h₁ : a ≤ b) (h₂ : c < d) : a + c < b + d
Full source
protected theorem add_lt_add_of_le_of_lt {a b c d : Int} (h₁ : a ≤ b) (h₂ : c < d) :
    a + c < b + d :=
  Int.lt_of_le_of_lt (Int.add_le_add_right h₁ c) (Int.add_lt_add_left h₂ b)
Sum of Mixed Inequalities in Integers: $a \leq b$ and $c < d$ implies $a + c < b + d$
Informal description
For any integers $a, b, c, d$ such that $a \leq b$ and $c < d$, it holds that $a + c < b + d$.
Int.add_lt_add_of_lt_of_le theorem
{a b c d : Int} (h₁ : a < b) (h₂ : c ≤ d) : a + c < b + d
Full source
protected theorem add_lt_add_of_lt_of_le {a b c d : Int} (h₁ : a < b) (h₂ : c ≤ d) :
    a + c < b + d :=
  Int.lt_of_lt_of_le (Int.add_lt_add_right h₁ c) (Int.add_le_add_left h₂ b)
Addition Preserves Mixed Inequalities in Integers: $a < b$ and $c \leq d$ implies $a + c < b + d$
Informal description
For any integers $a$, $b$, $c$, and $d$ such that $a < b$ and $c \leq d$, it holds that $a + c < b + d$.
Int.lt_add_of_pos_right theorem
(a : Int) {b : Int} (h : 0 < b) : a < a + b
Full source
protected theorem lt_add_of_pos_right (a : Int) {b : Int} (h : 0 < b) : a < a + b := by
  have : a + 0 < a + b := Int.add_lt_add_left h a
  rwa [Int.add_zero] at this
Positive Addition Increases Integer Value: $a < a + b$ for $b > 0$
Informal description
For any integer $a$ and any positive integer $b$ (i.e., $0 < b$), it holds that $a < a + b$.
Int.lt_add_of_pos_left theorem
(a : Int) {b : Int} (h : 0 < b) : a < b + a
Full source
protected theorem lt_add_of_pos_left (a : Int) {b : Int} (h : 0 < b) : a < b + a := by
  have : 0 + a < b + a := Int.add_lt_add_right h a
  rwa [Int.zero_add] at this
Positive Addition Increases Integer Value on the Left: $a < b + a$ for $b > 0$
Informal description
For any integer $a$ and any positive integer $b$ (i.e., $0 < b$), it holds that $a < b + a$.
Int.add_nonneg theorem
{a b : Int} (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a + b
Full source
protected theorem add_nonneg {a b : Int} (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a + b :=
  Int.zero_add 0 ▸ Int.add_le_add ha hb
Nonnegative Integers are Closed Under Addition
Informal description
For any integers $a$ and $b$ such that $0 \leq a$ and $0 \leq b$, it holds that $0 \leq a + b$.
Int.add_pos theorem
{a b : Int} (ha : 0 < a) (hb : 0 < b) : 0 < a + b
Full source
protected theorem add_pos {a b : Int} (ha : 0 < a) (hb : 0 < b) : 0 < a + b :=
  Int.zero_add 0 ▸ Int.add_lt_add ha hb
Sum of Positive Integers is Positive: $a > 0 \land b > 0 \implies a + b > 0$
Informal description
For any integers $a$ and $b$ such that $a > 0$ and $b > 0$, their sum satisfies $a + b > 0$.
Int.add_pos_of_pos_of_nonneg theorem
{a b : Int} (ha : 0 < a) (hb : 0 ≤ b) : 0 < a + b
Full source
protected theorem add_pos_of_pos_of_nonneg {a b : Int} (ha : 0 < a) (hb : 0 ≤ b) : 0 < a + b :=
  Int.zero_add 0 ▸ Int.add_lt_add_of_lt_of_le ha hb
Positivity of Sum: $a > 0$ and $b \geq 0$ implies $a + b > 0$
Informal description
For any integers $a$ and $b$ such that $a > 0$ and $b \geq 0$, their sum satisfies $a + b > 0$.
Int.add_pos_of_nonneg_of_pos theorem
{a b : Int} (ha : 0 ≤ a) (hb : 0 < b) : 0 < a + b
Full source
protected theorem add_pos_of_nonneg_of_pos {a b : Int} (ha : 0 ≤ a) (hb : 0 < b) : 0 < a + b :=
  Int.zero_add 0 ▸ Int.add_lt_add_of_le_of_lt ha hb
Sum of Nonnegative and Positive Integers is Positive: $0 \leq a \land 0 < b \implies 0 < a + b$
Informal description
For any integers $a$ and $b$ such that $a$ is nonnegative ($0 \leq a$) and $b$ is positive ($0 < b$), their sum is positive ($0 < a + b$).
Int.add_nonpos theorem
{a b : Int} (ha : a ≤ 0) (hb : b ≤ 0) : a + b ≤ 0
Full source
protected theorem add_nonpos {a b : Int} (ha : a ≤ 0) (hb : b ≤ 0) : a + b ≤ 0 :=
  Int.zero_add 0 ▸ Int.add_le_add ha hb
Sum of Nonpositive Integers is Nonpositive: $a \leq 0 \land b \leq 0 \implies a + b \leq 0$
Informal description
For any integers $a$ and $b$ such that $a \leq 0$ and $b \leq 0$, their sum satisfies $a + b \leq 0$.
Int.add_neg theorem
{a b : Int} (ha : a < 0) (hb : b < 0) : a + b < 0
Full source
protected theorem add_neg {a b : Int} (ha : a < 0) (hb : b < 0) : a + b < 0 :=
  Int.zero_add 0 ▸ Int.add_lt_add ha hb
Sum of Negative Integers is Negative: $a < 0 \land b < 0 \implies a + b < 0$
Informal description
For any integers $a$ and $b$ such that $a < 0$ and $b < 0$, their sum satisfies $a + b < 0$.
Int.add_neg_of_neg_of_nonpos theorem
{a b : Int} (ha : a < 0) (hb : b ≤ 0) : a + b < 0
Full source
protected theorem add_neg_of_neg_of_nonpos {a b : Int} (ha : a < 0) (hb : b ≤ 0) : a + b < 0 :=
  Int.zero_add 0 ▸ Int.add_lt_add_of_lt_of_le ha hb
Sum of Negative and Nonpositive Integers is Negative: $a < 0 \land b \leq 0 \implies a + b < 0$
Informal description
For any integers $a$ and $b$ such that $a < 0$ and $b \leq 0$, their sum satisfies $a + b < 0$.
Int.add_neg_of_nonpos_of_neg theorem
{a b : Int} (ha : a ≤ 0) (hb : b < 0) : a + b < 0
Full source
protected theorem add_neg_of_nonpos_of_neg {a b : Int} (ha : a ≤ 0) (hb : b < 0) : a + b < 0 :=
  Int.zero_add 0 ▸ Int.add_lt_add_of_le_of_lt ha hb
Sum of Nonpositive and Negative Integers is Negative
Informal description
For any integers $a$ and $b$ such that $a \leq 0$ and $b < 0$, the sum $a + b$ is strictly less than $0$.
Int.lt_add_of_le_of_pos theorem
{a b c : Int} (hbc : b ≤ c) (ha : 0 < a) : b < c + a
Full source
protected theorem lt_add_of_le_of_pos {a b c : Int} (hbc : b ≤ c) (ha : 0 < a) : b < c + a :=
  Int.add_zero b ▸ Int.add_lt_add_of_le_of_lt hbc ha
Inequality Preservation Under Addition of Positive Integer: $b \leq c$ and $0 < a$ implies $b < c + a$
Informal description
For any integers $a, b, c$ such that $b \leq c$ and $0 < a$, it holds that $b < c + a$.
Int.add_one_le_iff theorem
{a b : Int} : a + 1 ≤ b ↔ a < b
Full source
theorem add_one_le_iff {a b : Int} : a + 1 ≤ b ↔ a < b := .rfl
Integer Inequality: $a + 1 \leq b \leftrightarrow a < b$
Informal description
For any integers $a$ and $b$, the inequality $a + 1 \leq b$ holds if and only if $a < b$.
Int.lt_add_one_iff theorem
{a b : Int} : a < b + 1 ↔ a ≤ b
Full source
theorem lt_add_one_iff {a b : Int} : a < b + 1 ↔ a ≤ b := Int.add_le_add_iff_right _
Integer Inequality: $a < b + 1 \leftrightarrow a \leq b$
Informal description
For any integers $a$ and $b$, the inequality $a < b + 1$ holds if and only if $a \leq b$.
Int.succ_ofNat_pos theorem
(n : Nat) : 0 < (n : Int) + 1
Full source
@[simp] theorem succ_ofNat_pos (n : Nat) : 0 < (n : Int) + 1 :=
  lt_add_one_iff.mpr (ofNat_zero_le _)
Positivity of Successor of Natural Number in Integers: $0 < n + 1$
Informal description
For any natural number $n$, the integer $0$ is strictly less than $n + 1$, where $n$ is considered as an integer.
Int.not_ofNat_neg theorem
(n : Nat) : ¬((n : Int) < 0)
Full source
theorem not_ofNat_neg (n : Nat) : ¬((n : Int) < 0) :=
  Int.not_lt.mpr (ofNat_zero_le ..)
Non-negativity of Natural Numbers in Integers: $\neg(n < 0)$ for $n \in \mathbb{N}$
Informal description
For any natural number $n$, the integer obtained by casting $n$ to $\mathbb{Z}$ is not less than zero, i.e., $\neg(n < 0)$ where $n$ is considered as an integer.
Int.le_add_one theorem
{a b : Int} (h : a ≤ b) : a ≤ b + 1
Full source
theorem le_add_one {a b : Int} (h : a ≤ b) : a ≤ b + 1 :=
  Int.le_of_lt (Int.lt_add_one_iff.2 h)
Non-strict Inequality Preservation Under Integer Successor
Informal description
For any integers $a$ and $b$ such that $a \leq b$, it holds that $a \leq b + 1$.
Int.nonneg_of_neg_nonpos theorem
{a : Int} (h : -a ≤ 0) : 0 ≤ a
Full source
protected theorem nonneg_of_neg_nonpos {a : Int} (h : -a ≤ 0) : 0 ≤ a :=
  Int.le_of_neg_le_neg <| by rwa [Int.neg_zero]
Nonnegativity from Nonpositive Negation in Integers
Informal description
For any integer $a$, if $-a \leq 0$, then $0 \leq a$.
Int.nonpos_of_neg_nonneg theorem
{a : Int} (h : 0 ≤ -a) : a ≤ 0
Full source
protected theorem nonpos_of_neg_nonneg {a : Int} (h : 0 ≤ -a) : a ≤ 0 :=
  Int.le_of_neg_le_neg <| by rwa [Int.neg_zero]
Nonpositivity from Nonnegative Negation in Integers
Informal description
For any integer $a$, if the negation of $a$ is nonnegative (i.e., $0 \leq -a$), then $a$ is nonpositive (i.e., $a \leq 0$).
Int.lt_of_neg_lt_neg theorem
{a b : Int} (h : -b < -a) : a < b
Full source
protected theorem lt_of_neg_lt_neg {a b : Int} (h : -b < -a) : a < b :=
  Int.neg_neg a ▸ Int.neg_neg b ▸ Int.neg_lt_neg h
Strict Inequality Reversal under Negation in Integers: $-b < -a \implies a < b$
Informal description
For any integers $a$ and $b$, if $-b < -a$, then $a < b$.
Int.pos_of_neg_neg theorem
{a : Int} (h : -a < 0) : 0 < a
Full source
protected theorem pos_of_neg_neg {a : Int} (h : -a < 0) : 0 < a :=
  Int.lt_of_neg_lt_neg <| by rwa [Int.neg_zero]
Positivity from Negative Negation in Integers
Informal description
For any integer $a$, if $-a$ is strictly negative, then $a$ is strictly positive, i.e., $-a < 0 \implies 0 < a$.
Int.neg_of_neg_pos theorem
{a : Int} (h : 0 < -a) : a < 0
Full source
protected theorem neg_of_neg_pos {a : Int} (h : 0 < -a) : a < 0 :=
  have : -0 < -a := by rwa [Int.neg_zero]
  Int.lt_of_neg_lt_neg this
Negation Reverses Positive to Negative: $0 < -a \implies a < 0$
Informal description
For any integer $a$, if $0 < -a$, then $a < 0$.
Int.le_neg_of_le_neg theorem
{a b : Int} (h : a ≤ -b) : b ≤ -a
Full source
protected theorem le_neg_of_le_neg {a b : Int} (h : a ≤ -b) : b ≤ -a := by
  have h := Int.neg_le_neg h
  rwa [Int.neg_neg] at h
Inequality Reversal under Negation: $a \leq -b \implies b \leq -a$
Informal description
For any integers $a$ and $b$, if $a \leq -b$, then $b \leq -a$.
Int.neg_le_of_neg_le theorem
{a b : Int} (h : -a ≤ b) : -b ≤ a
Full source
protected theorem neg_le_of_neg_le {a b : Int} (h : -a ≤ b) : -b ≤ a := by
  have h := Int.neg_le_neg h
  rwa [Int.neg_neg] at h
Negation Reverses Inequality: $-a \leq b \implies -b \leq a$
Informal description
For any integers $a$ and $b$, if $-a \leq b$, then $-b \leq a$.
Int.lt_neg_of_lt_neg theorem
{a b : Int} (h : a < -b) : b < -a
Full source
protected theorem lt_neg_of_lt_neg {a b : Int} (h : a < -b) : b < -a := by
  have h := Int.neg_lt_neg h
  rwa [Int.neg_neg] at h
Negation Reverses Strict Inequality: $a < -b \implies b < -a$
Informal description
For any integers $a$ and $b$ such that $a < -b$, it holds that $b < -a$.
Int.neg_lt_of_neg_lt theorem
{a b : Int} (h : -a < b) : -b < a
Full source
protected theorem neg_lt_of_neg_lt {a b : Int} (h : -a < b) : -b < a := by
  have h := Int.neg_lt_neg h
  rwa [Int.neg_neg] at h
Negation Reverses Strict Inequality: $-a < b \implies -b < a$
Informal description
For any integers $a$ and $b$, if $-a < b$, then $-b < a$.
Int.sub_nonpos_of_le theorem
{a b : Int} (h : a ≤ b) : a - b ≤ 0
Full source
protected theorem sub_nonpos_of_le {a b : Int} (h : a ≤ b) : a - b ≤ 0 := by
  have h := Int.add_le_add_right h (-b)
  rwa [Int.add_right_neg] at h
Difference Nonpositivity from Integer Inequality
Informal description
For any integers $a$ and $b$ such that $a \leq b$, the difference $a - b$ is less than or equal to zero, i.e., $a - b \leq 0$.
Int.le_of_sub_nonpos theorem
{a b : Int} (h : a - b ≤ 0) : a ≤ b
Full source
protected theorem le_of_sub_nonpos {a b : Int} (h : a - b ≤ 0) : a ≤ b := by
  have h := Int.add_le_add_right h b
  rwa [Int.sub_add_cancel, Int.zero_add] at h
Nonpositive Difference Implies Inequality: $a - b \leq 0 \Rightarrow a \leq b$
Informal description
For any integers $a$ and $b$, if $a - b \leq 0$, then $a \leq b$.
Int.sub_neg_of_lt theorem
{a b : Int} (h : a < b) : a - b < 0
Full source
protected theorem sub_neg_of_lt {a b : Int} (h : a < b) : a - b < 0 := by
  have h := Int.add_lt_add_right h (-b)
  rwa [Int.add_right_neg] at h
Negative Difference from Strict Integer Inequality
Informal description
For any integers $a$ and $b$ such that $a < b$, the difference $a - b$ is negative, i.e., $a - b < 0$.
Int.lt_of_sub_neg theorem
{a b : Int} (h : a - b < 0) : a < b
Full source
protected theorem lt_of_sub_neg {a b : Int} (h : a - b < 0) : a < b := by
  have h := Int.add_lt_add_right h b
  rwa [Int.sub_add_cancel, Int.zero_add] at h
Negative Difference Implies Strict Inequality in Integers: $a - b < 0 \Rightarrow a < b$
Informal description
For any integers $a$ and $b$, if the difference $a - b$ is negative, then $a$ is strictly less than $b$, i.e., $a < b$.
Int.add_le_of_le_neg_add theorem
{a b c : Int} (h : b ≤ -a + c) : a + b ≤ c
Full source
protected theorem add_le_of_le_neg_add {a b c : Int} (h : b ≤ -a + c) : a + b ≤ c := by
  have h := Int.add_le_add_left h a
  rwa [Int.add_neg_cancel_left] at h
Inequality Transformation via Negated Addition in Integers
Informal description
For any integers $a$, $b$, and $c$, if $b \leq -a + c$, then $a + b \leq c$.
Int.le_neg_add_of_add_le theorem
{a b c : Int} (h : a + b ≤ c) : b ≤ -a + c
Full source
protected theorem le_neg_add_of_add_le {a b c : Int} (h : a + b ≤ c) : b ≤ -a + c := by
  have h := Int.add_le_add_left h (-a)
  rwa [Int.neg_add_cancel_left] at h
Inequality Transformation via Negation and Addition in Integers
Informal description
For any integers $a$, $b$, and $c$, if $a + b \leq c$, then $b \leq -a + c$.
Int.add_le_of_le_sub_left theorem
{a b c : Int} (h : b ≤ c - a) : a + b ≤ c
Full source
protected theorem add_le_of_le_sub_left {a b c : Int} (h : b ≤ c - a) : a + b ≤ c := by
  have h := Int.add_le_add_left h a
  rwa [← Int.add_sub_assoc, Int.add_comm a c, Int.add_sub_cancel] at h
Inequality Preservation under Addition and Subtraction: $b \leq c - a \implies a + b \leq c$
Informal description
For any integers $a$, $b$, and $c$, if $b \leq c - a$, then $a + b \leq c$.
Int.le_sub_left_of_add_le theorem
{a b c : Int} (h : a + b ≤ c) : b ≤ c - a
Full source
protected theorem le_sub_left_of_add_le {a b c : Int} (h : a + b ≤ c) : b ≤ c - a := by
  have h := Int.add_le_add_right h (-a)
  rwa [Int.add_comm a b, Int.add_neg_cancel_right] at h
Subtraction Inequality from Addition: $a + b \leq c \implies b \leq c - a$
Informal description
For any integers $a$, $b$, and $c$, if $a + b \leq c$, then $b \leq c - a$.
Int.add_le_of_le_sub_right theorem
{a b c : Int} (h : a ≤ c - b) : a + b ≤ c
Full source
protected theorem add_le_of_le_sub_right {a b c : Int} (h : a ≤ c - b) : a + b ≤ c := by
  have h := Int.add_le_add_right h b
  rwa [Int.sub_add_cancel] at h
Inequality Preservation under Addition and Subtraction: $a \leq c - b \implies a + b \leq c$
Informal description
For any integers $a$, $b$, and $c$, if $a \leq c - b$, then $a + b \leq c$.
Int.le_sub_right_of_add_le theorem
{a b c : Int} (h : a + b ≤ c) : a ≤ c - b
Full source
protected theorem le_sub_right_of_add_le {a b c : Int} (h : a + b ≤ c) : a ≤ c - b := by
  have h := Int.add_le_add_right h (-b)
  rwa [Int.add_neg_cancel_right] at h
Subtraction Inequality from Addition: $a + b \leq c \implies a \leq c - b$
Informal description
For any integers $a$, $b$, and $c$, if $a + b \leq c$, then $a \leq c - b$.
Int.le_add_of_neg_add_le theorem
{a b c : Int} (h : -b + a ≤ c) : a ≤ b + c
Full source
protected theorem le_add_of_neg_add_le {a b c : Int} (h : -b + a ≤ c) : a ≤ b + c := by
  have h := Int.add_le_add_left h b
  rwa [Int.add_neg_cancel_left] at h
Inequality Transformation via Negated Addition in Integers
Informal description
For any integers $a$, $b$, and $c$, if $-b + a \leq c$, then $a \leq b + c$.
Int.neg_add_le_of_le_add theorem
{a b c : Int} (h : a ≤ b + c) : -b + a ≤ c
Full source
protected theorem neg_add_le_of_le_add {a b c : Int} (h : a ≤ b + c) : -b + a ≤ c := by
  have h := Int.add_le_add_left h (-b)
  rwa [Int.neg_add_cancel_left] at h
Left Translation of Inequality via Negation: $-b + a \leq c$ when $a \leq b + c$
Informal description
For any integers $a$, $b$, and $c$ such that $a \leq b + c$, it holds that $-b + a \leq c$.
Int.le_add_of_sub_left_le theorem
{a b c : Int} (h : a - b ≤ c) : a ≤ b + c
Full source
protected theorem le_add_of_sub_left_le {a b c : Int} (h : a - b ≤ c) : a ≤ b + c := by
  have h := Int.add_le_add_right h b
  rwa [Int.sub_add_cancel, Int.add_comm] at h
Inequality Transformation via Left Subtraction in Integers
Informal description
For any integers $a$, $b$, and $c$, if $a - b \leq c$, then $a \leq b + c$.
Int.le_add_of_sub_right_le theorem
{a b c : Int} (h : a - c ≤ b) : a ≤ b + c
Full source
protected theorem le_add_of_sub_right_le {a b c : Int} (h : a - c ≤ b) : a ≤ b + c := by
  have h := Int.add_le_add_right h c
  rwa [Int.sub_add_cancel] at h
Inequality from Subtraction on the Right: $a - c \leq b \implies a \leq b + c$
Informal description
For any integers $a$, $b$, and $c$, if $a - c \leq b$, then $a \leq b + c$.
Int.sub_right_le_of_le_add theorem
{a b c : Int} (h : a ≤ b + c) : a - c ≤ b
Full source
protected theorem sub_right_le_of_le_add {a b c : Int} (h : a ≤ b + c) : a - c ≤ b := by
  have h := Int.add_le_add_right h (-c)
  rwa [Int.add_neg_cancel_right] at h
Subtraction Inequality from Addition Inequality in Integers: $a \leq b + c \implies a - c \leq b$
Informal description
For any integers $a$, $b$, and $c$ such that $a \leq b + c$, it follows that $a - c \leq b$.
Int.le_add_of_neg_add_le_left theorem
{a b c : Int} (h : -b + a ≤ c) : a ≤ b + c
Full source
protected theorem le_add_of_neg_add_le_left {a b c : Int} (h : -b + a ≤ c) : a ≤ b + c := by
  rw [Int.add_comm] at h
  exact Int.le_add_of_sub_left_le h
Inequality Transformation via Left Negation and Addition in Integers: $-b + a \leq c \implies a \leq b + c$
Informal description
For any integers $a$, $b$, and $c$, if $-b + a \leq c$, then $a \leq b + c$.
Int.neg_add_le_left_of_le_add theorem
{a b c : Int} (h : a ≤ b + c) : -b + a ≤ c
Full source
protected theorem neg_add_le_left_of_le_add {a b c : Int} (h : a ≤ b + c) : -b + a ≤ c := by
  rw [Int.add_comm]
  exact Int.sub_left_le_of_le_add h
Left Negated Addition Inequality from Addition Inequality in Integers: $-b + a \leq c$ if $a \leq b + c$
Informal description
For any integers $a$, $b$, and $c$ such that $a \leq b + c$, it holds that $-b + a \leq c$.
Int.le_add_of_neg_add_le_right theorem
{a b c : Int} (h : -c + a ≤ b) : a ≤ b + c
Full source
protected theorem le_add_of_neg_add_le_right {a b c : Int} (h : -c + a ≤ b) : a ≤ b + c := by
  rw [Int.add_comm] at h
  exact Int.le_add_of_sub_right_le h
Inequality from Negative Addition on the Right: $-c + a \leq b \implies a \leq b + c$
Informal description
For any integers $a$, $b$, and $c$, if $-c + a \leq b$, then $a \leq b + c$.
Int.neg_add_le_right_of_le_add theorem
{a b c : Int} (h : a ≤ b + c) : -c + a ≤ b
Full source
protected theorem neg_add_le_right_of_le_add {a b c : Int} (h : a ≤ b + c) : -c + a ≤ b := by
  rw [Int.add_comm] at h
  exact Int.neg_add_le_left_of_le_add h
Right Negated Addition Inequality from Addition Inequality in Integers: $-c + a \leq b$ if $a \leq b + c$
Informal description
For any integers $a$, $b$, and $c$ such that $a \leq b + c$, it holds that $-c + a \leq b$.
Int.le_add_of_neg_le_sub_left theorem
{a b c : Int} (h : -a ≤ b - c) : c ≤ a + b
Full source
protected theorem le_add_of_neg_le_sub_left {a b c : Int} (h : -a ≤ b - c) : c ≤ a + b :=
  Int.le_add_of_neg_add_le_left (Int.add_le_of_le_sub_right h)
Inequality from Negative and Subtraction on the Left: $-a \leq b - c \implies c \leq a + b$
Informal description
For any integers $a$, $b$, and $c$, if $-a \leq b - c$, then $c \leq a + b$.
Int.neg_le_sub_left_of_le_add theorem
{a b c : Int} (h : c ≤ a + b) : -a ≤ b - c
Full source
protected theorem neg_le_sub_left_of_le_add {a b c : Int} (h : c ≤ a + b) : -a ≤ b - c := by
  have h := Int.le_neg_add_of_add_le (Int.sub_left_le_of_le_add h)
  rwa [Int.add_comm] at h
Negation and Subtraction Inequality from Sum Inequality in Integers
Informal description
For any integers $a$, $b$, and $c$ such that $c \leq a + b$, it holds that $-a \leq b - c$.
Int.le_add_of_neg_le_sub_right theorem
{a b c : Int} (h : -b ≤ a - c) : c ≤ a + b
Full source
protected theorem le_add_of_neg_le_sub_right {a b c : Int} (h : -b ≤ a - c) : c ≤ a + b :=
  Int.le_add_of_sub_right_le (Int.add_le_of_le_sub_left h)
Inequality from Negative and Subtraction on the Right: $-b \leq a - c \implies c \leq a + b$
Informal description
For any integers $a$, $b$, and $c$, if $-b \leq a - c$, then $c \leq a + b$.
Int.neg_le_sub_right_of_le_add theorem
{a b c : Int} (h : c ≤ a + b) : -b ≤ a - c
Full source
protected theorem neg_le_sub_right_of_le_add {a b c : Int} (h : c ≤ a + b) : -b ≤ a - c :=
  Int.le_sub_left_of_add_le (Int.sub_right_le_of_le_add h)
Negative Bound from Addition Inequality in Integers: $c \leq a + b \implies -b \leq a - c$
Informal description
For any integers $a$, $b$, and $c$ such that $c \leq a + b$, it follows that $-b \leq a - c$.
Int.sub_le_of_sub_le theorem
{a b c : Int} (h : a - b ≤ c) : a - c ≤ b
Full source
protected theorem sub_le_of_sub_le {a b c : Int} (h : a - b ≤ c) : a - c ≤ b :=
  Int.sub_left_le_of_le_add (Int.le_add_of_sub_right_le h)
Subtraction Inequality Transfer: $a - b \leq c \implies a - c \leq b$
Informal description
For any integers $a$, $b$, and $c$, if $a - b \leq c$, then $a - c \leq b$.
Int.sub_le_sub_left theorem
{a b : Int} (h : a ≤ b) (c : Int) : c - b ≤ c - a
Full source
protected theorem sub_le_sub_left {a b : Int} (h : a ≤ b) (c : Int) : c - b ≤ c - a :=
  Int.add_le_add_left (Int.neg_le_neg h) c
Left Subtraction Reverses Inequality in Integers
Informal description
For any integers $a$ and $b$ such that $a \leq b$, and for any integer $c$, it holds that $c - b \leq c - a$.
Int.sub_le_sub_right theorem
{a b : Int} (h : a ≤ b) (c : Int) : a - c ≤ b - c
Full source
protected theorem sub_le_sub_right {a b : Int} (h : a ≤ b) (c : Int) : a - c ≤ b - c :=
  Int.add_le_add_right h (-c)
Right Subtraction Preserves Inequality in Integers
Informal description
For any integers $a$ and $b$ such that $a \leq b$, and for any integer $c$, it holds that $a - c \leq b - c$.
Int.sub_le_sub theorem
{a b c d : Int} (hab : a ≤ b) (hcd : c ≤ d) : a - d ≤ b - c
Full source
protected theorem sub_le_sub {a b c d : Int} (hab : a ≤ b) (hcd : c ≤ d) : a - d ≤ b - c :=
  Int.add_le_add hab (Int.neg_le_neg hcd)
Subtraction Preserves Inequality in Integers: $a \leq b \land c \leq d \implies a - d \leq b - c$
Informal description
For any integers $a$, $b$, $c$, and $d$ such that $a \leq b$ and $c \leq d$, it holds that $a - d \leq b - c$.
Int.le_of_sub_le_sub_left theorem
{a b c : Int} (h : c - a ≤ c - b) : b ≤ a
Full source
protected theorem le_of_sub_le_sub_left {a b c : Int} (h : c - a ≤ c - b) : b ≤ a :=
  Int.le_of_neg_le_neg <| Int.le_of_add_le_add_left h
Left Subtraction Preserves Order in Integers
Informal description
For any integers $a$, $b$, and $c$, if $c - a \leq c - b$, then $b \leq a$.
Int.le_of_sub_le_sub_right theorem
{a b c : Int} (h : a - c ≤ b - c) : a ≤ b
Full source
protected theorem le_of_sub_le_sub_right {a b c : Int} (h : a - c ≤ b - c) : a ≤ b :=
  Int.le_of_add_le_add_right h
Right Subtraction Preserves Order in Integers
Informal description
For any integers $a$, $b$, and $c$, if $a - c \leq b - c$, then $a \leq b$.
Int.add_lt_of_lt_neg_add theorem
{a b c : Int} (h : b < -a + c) : a + b < c
Full source
protected theorem add_lt_of_lt_neg_add {a b c : Int} (h : b < -a + c) : a + b < c := by
  have h := Int.add_lt_add_left h a
  rwa [Int.add_neg_cancel_left] at h
Inequality Transformation via Negated Addition: $b < -a + c \implies a + b < c$
Informal description
For any integers $a$, $b$, and $c$, if $b < -a + c$, then $a + b < c$.
Int.lt_neg_add_of_add_lt theorem
{a b c : Int} (h : a + b < c) : b < -a + c
Full source
protected theorem lt_neg_add_of_add_lt {a b c : Int} (h : a + b < c) : b < -a + c := by
  have h := Int.add_lt_add_left h (-a)
  rwa [Int.neg_add_cancel_left] at h
Inequality Transformation: $a + b < c$ implies $b < -a + c$ in integers
Informal description
For any integers $a$, $b$, and $c$, if $a + b < c$, then $b < -a + c$.
Int.lt_sub_left_of_add_lt theorem
{a b c : Int} (h : a + b < c) : b < c - a
Full source
protected theorem lt_sub_left_of_add_lt {a b c : Int} (h : a + b < c) : b < c - a := by
  have h := Int.add_lt_add_right h (-a)
  rwa [Int.add_comm a b, Int.add_neg_cancel_right] at h
Inequality Transformation: $a + b < c$ implies $b < c - a$ in integers
Informal description
For any integers $a$, $b$, and $c$, if $a + b < c$, then $b < c - a$.
Int.add_lt_of_lt_sub_right theorem
{a b c : Int} (h : a < c - b) : a + b < c
Full source
protected theorem add_lt_of_lt_sub_right {a b c : Int} (h : a < c - b) : a + b < c := by
  have h := Int.add_lt_add_right h b
  rwa [Int.sub_add_cancel] at h
Sum Inequality from Right Subtraction Inequality in Integers
Informal description
For any integers $a$, $b$, and $c$, if $a < c - b$, then $a + b < c$.
Int.lt_sub_right_of_add_lt theorem
{a b c : Int} (h : a + b < c) : a < c - b
Full source
protected theorem lt_sub_right_of_add_lt {a b c : Int} (h : a + b < c) : a < c - b := by
  have h := Int.add_lt_add_right h (-b)
  rwa [Int.add_neg_cancel_right] at h
Subtraction Preserves Inequality from Right in Integers
Informal description
For any integers $a$, $b$, and $c$, if $a + b < c$, then $a < c - b$.
Int.lt_add_of_neg_add_lt theorem
{a b c : Int} (h : -b + a < c) : a < b + c
Full source
protected theorem lt_add_of_neg_add_lt {a b c : Int} (h : -b + a < c) : a < b + c := by
  have h := Int.add_lt_add_left h b
  rwa [Int.add_neg_cancel_left] at h
Inequality Transformation: $-b + a < c$ implies $a < b + c$
Informal description
For any integers $a$, $b$, and $c$, if $-b + a < c$, then $a < b + c$.
Int.neg_add_lt_of_lt_add theorem
{a b c : Int} (h : a < b + c) : -b + a < c
Full source
protected theorem neg_add_lt_of_lt_add {a b c : Int} (h : a < b + c) : -b + a < c := by
  have h := Int.add_lt_add_left h (-b)
  rwa [Int.neg_add_cancel_left] at h
Negated Addition Preserves Inequality in Integers
Informal description
For any integers $a$, $b$, and $c$ such that $a < b + c$, it holds that $-b + a < c$.
Int.lt_add_of_sub_left_lt theorem
{a b c : Int} (h : a - b < c) : a < b + c
Full source
protected theorem lt_add_of_sub_left_lt {a b c : Int} (h : a - b < c) : a < b + c := by
  have h := Int.add_lt_add_right h b
  rwa [Int.sub_add_cancel, Int.add_comm] at h
Inequality Transformation via Left Subtraction: $a - b < c$ implies $a < b + c$
Informal description
For any integers $a$, $b$, and $c$, if $a - b < c$, then $a < b + c$.
Int.sub_left_lt_of_lt_add theorem
{a b c : Int} (h : a < b + c) : a - b < c
Full source
protected theorem sub_left_lt_of_lt_add {a b c : Int} (h : a < b + c) : a - b < c := by
  have h := Int.add_lt_add_right h (-b)
  rwa [Int.add_comm b c, Int.add_neg_cancel_right] at h
Subtraction Inequality from Addition Inequality in Integers: $a < b + c \implies a - b < c$
Informal description
For any integers $a$, $b$, and $c$ such that $a < b + c$, it holds that $a - b < c$.
Int.lt_add_of_sub_right_lt theorem
{a b c : Int} (h : a - c < b) : a < b + c
Full source
protected theorem lt_add_of_sub_right_lt {a b c : Int} (h : a - c < b) : a < b + c := by
  have h := Int.add_lt_add_right h c
  rwa [Int.sub_add_cancel] at h
Inequality from Subtraction: $a - c < b$ implies $a < b + c$
Informal description
For any integers $a$, $b$, and $c$, if $a - c < b$, then $a < b + c$.
Int.sub_right_lt_of_lt_add theorem
{a b c : Int} (h : a < b + c) : a - c < b
Full source
protected theorem sub_right_lt_of_lt_add {a b c : Int} (h : a < b + c) : a - c < b := by
  have h := Int.add_lt_add_right h (-c)
  rwa [Int.add_neg_cancel_right] at h
Subtraction from Right Preserves Inequality in Integers: $a < b + c \to a - c < b$
Informal description
For any integers $a$, $b$, and $c$ such that $a < b + c$, it holds that $a - c < b$.
Int.lt_add_of_neg_add_lt_left theorem
{a b c : Int} (h : -b + a < c) : a < b + c
Full source
protected theorem lt_add_of_neg_add_lt_left {a b c : Int} (h : -b + a < c) : a < b + c := by
  rw [Int.add_comm] at h
  exact Int.lt_add_of_sub_left_lt h
Inequality from Negative Addition: $-b + a < c$ implies $a < b + c$
Informal description
For any integers $a$, $b$, and $c$, if $-b + a < c$, then $a < b + c$.
Int.neg_add_lt_left_of_lt_add theorem
{a b c : Int} (h : a < b + c) : -b + a < c
Full source
protected theorem neg_add_lt_left_of_lt_add {a b c : Int} (h : a < b + c) : -b + a < c := by
  rw [Int.add_comm]
  exact Int.sub_left_lt_of_lt_add h
Left Negation-Addition Inequality from Addition Inequality in Integers: $a < b + c \implies -b + a < c$
Informal description
For any integers $a$, $b$, and $c$ such that $a < b + c$, it holds that $-b + a < c$.
Int.lt_add_of_neg_add_lt_right theorem
{a b c : Int} (h : -c + a < b) : a < b + c
Full source
protected theorem lt_add_of_neg_add_lt_right {a b c : Int} (h : -c + a < b) : a < b + c := by
  rw [Int.add_comm] at h
  exact Int.lt_add_of_sub_right_lt h
Inequality from Negative Addition: $-c + a < b$ implies $a < b + c$
Informal description
For any integers $a$, $b$, and $c$, if $-c + a < b$, then $a < b + c$.
Int.neg_add_lt_right_of_lt_add theorem
{a b c : Int} (h : a < b + c) : -c + a < b
Full source
protected theorem neg_add_lt_right_of_lt_add {a b c : Int} (h : a < b + c) : -c + a < b := by
  rw [Int.add_comm] at h
  exact Int.neg_add_lt_left_of_lt_add h
Right Negation-Addition Inequality from Sum Inequality in Integers: $a < b + c \implies -c + a < b$
Informal description
For any integers $a$, $b$, and $c$ such that $a < b + c$, it holds that $-c + a < b$.
Int.neg_lt_sub_left_of_lt_add theorem
{a b c : Int} (h : c < a + b) : -a < b - c
Full source
protected theorem neg_lt_sub_left_of_lt_add {a b c : Int} (h : c < a + b) : -a < b - c := by
  have h := Int.lt_neg_add_of_add_lt (Int.sub_left_lt_of_lt_add h)
  rwa [Int.add_comm] at h
Negation-Subtraction Inequality from Sum Inequality in Integers: $-a < b - c$ when $c < a + b$
Informal description
For any integers $a$, $b$, and $c$ such that $c < a + b$, it holds that $-a < b - c$.
Int.lt_add_of_neg_lt_sub_right theorem
{a b c : Int} (h : -b < a - c) : c < a + b
Full source
protected theorem lt_add_of_neg_lt_sub_right {a b c : Int} (h : -b < a - c) : c < a + b :=
  Int.lt_add_of_sub_right_lt (Int.add_lt_of_lt_sub_left h)
Inequality from Negative Subtraction: $-b < a - c$ implies $c < a + b$
Informal description
For any integers $a$, $b$, and $c$, if $-b < a - c$, then $c < a + b$.
Int.neg_lt_sub_right_of_lt_add theorem
{a b c : Int} (h : c < a + b) : -b < a - c
Full source
protected theorem neg_lt_sub_right_of_lt_add {a b c : Int} (h : c < a + b) : -b < a - c :=
  Int.lt_sub_left_of_add_lt (Int.sub_right_lt_of_lt_add h)
Negation and Subtraction Inequality from Sum Inequality in Integers
Informal description
For any integers $a$, $b$, and $c$ such that $c < a + b$, it holds that $-b < a - c$.
Int.sub_lt_of_sub_lt theorem
{a b c : Int} (h : a - b < c) : a - c < b
Full source
protected theorem sub_lt_of_sub_lt {a b c : Int} (h : a - b < c) : a - c < b :=
  Int.sub_left_lt_of_lt_add (Int.lt_add_of_sub_right_lt h)
Subtraction Inequality Transformation: $a - b < c \implies a - c < b$
Informal description
For any integers $a$, $b$, and $c$, if $a - b < c$, then $a - c < b$.
Int.sub_lt_sub_left theorem
{a b : Int} (h : a < b) (c : Int) : c - b < c - a
Full source
protected theorem sub_lt_sub_left {a b : Int} (h : a < b) (c : Int) : c - b < c - a :=
  Int.add_lt_add_left (Int.neg_lt_neg h) c
Left Subtraction Reverses Strict Inequality in Integers
Informal description
For any integers $a$ and $b$ such that $a < b$, and for any integer $c$, it holds that $c - b < c - a$.
Int.sub_lt_sub_right theorem
{a b : Int} (h : a < b) (c : Int) : a - c < b - c
Full source
protected theorem sub_lt_sub_right {a b : Int} (h : a < b) (c : Int) : a - c < b - c :=
  Int.add_lt_add_right h (-c)
Right Subtraction Preserves Strict Inequality in Integers
Informal description
For any integers $a$ and $b$ such that $a < b$, and for any integer $c$, it holds that $a - c < b - c$.
Int.sub_lt_sub theorem
{a b c d : Int} (hab : a < b) (hcd : c < d) : a - d < b - c
Full source
protected theorem sub_lt_sub {a b c d : Int} (hab : a < b) (hcd : c < d) : a - d < b - c :=
  Int.add_lt_add hab (Int.neg_lt_neg hcd)
Subtraction of Strict Inequalities in Integers: $a < b$ and $c < d$ implies $a - d < b - c$
Informal description
For any integers $a$, $b$, $c$, and $d$ such that $a < b$ and $c < d$, it holds that $a - d < b - c$.
Int.lt_of_sub_lt_sub_left theorem
{a b c : Int} (h : c - a < c - b) : b < a
Full source
protected theorem lt_of_sub_lt_sub_left {a b c : Int} (h : c - a < c - b) : b < a :=
  Int.lt_of_neg_lt_neg <| Int.lt_of_add_lt_add_left h
Left Subtraction Reverses Strict Inequality in Integers: $c - a < c - b \implies b < a$
Informal description
For any integers $a$, $b$, and $c$, if $c - a < c - b$, then $b < a$.
Int.lt_of_sub_lt_sub_right theorem
{a b c : Int} (h : a - c < b - c) : a < b
Full source
protected theorem lt_of_sub_lt_sub_right {a b c : Int} (h : a - c < b - c) : a < b :=
  Int.lt_of_add_lt_add_right h
Right Subtraction Preserves Strict Inequality in Integers
Informal description
For any integers $a$, $b$, and $c$, if $a - c < b - c$, then $a < b$.
Int.sub_lt_sub_of_le_of_lt theorem
{a b c d : Int} (hab : a ≤ b) (hcd : c < d) : a - d < b - c
Full source
protected theorem sub_lt_sub_of_le_of_lt {a b c d : Int}
    (hab : a ≤ b) (hcd : c < d) : a - d < b - c :=
  Int.add_lt_add_of_le_of_lt hab (Int.neg_lt_neg hcd)
Subtraction of Mixed Inequalities in Integers: $a \leq b$ and $c < d$ implies $a - d < b - c$
Informal description
For any integers $a, b, c, d$ such that $a \leq b$ and $c < d$, it holds that $a - d < b - c$.
Int.sub_lt_sub_of_lt_of_le theorem
{a b c d : Int} (hab : a < b) (hcd : c ≤ d) : a - d < b - c
Full source
protected theorem sub_lt_sub_of_lt_of_le {a b c d : Int}
    (hab : a < b) (hcd : c ≤ d) : a - d < b - c :=
  Int.add_lt_add_of_lt_of_le hab (Int.neg_le_neg hcd)
Subtraction Preserves Mixed Inequalities in Integers: $a < b$ and $c \leq d$ implies $a - d < b - c$
Informal description
For any integers $a$, $b$, $c$, and $d$ such that $a < b$ and $c \leq d$, it holds that $a - d < b - c$.
Int.add_le_add_three theorem
{a b c d e f : Int} (h₁ : a ≤ d) (h₂ : b ≤ e) (h₃ : c ≤ f) : a + b + c ≤ d + e + f
Full source
protected theorem add_le_add_three {a b c d e f : Int}
    (h₁ : a ≤ d) (h₂ : b ≤ e) (h₃ : c ≤ f) : a + b + c ≤ d + e + f :=
  Int.add_le_add (Int.add_le_add h₁ h₂) h₃
Addition Preserves Inequality for Three Pairs of Integers
Informal description
For any integers $a, b, c, d, e, f$ such that $a \leq d$, $b \leq e$, and $c \leq f$, it holds that $a + b + c \leq d + e + f$.
Int.lt_of_add_one_le theorem
{a b : Int} (H : a + 1 ≤ b) : a < b
Full source
theorem lt_of_add_one_le {a b : Int} (H : a + 1 ≤ b) : a < b := H
Strict Inequality from Successor Inequality in Integers
Informal description
For any integers $a$ and $b$, if $a + 1 \leq b$, then $a < b$.
Int.lt_add_one_of_le theorem
{a b : Int} (H : a ≤ b) : a < b + 1
Full source
theorem lt_add_one_of_le {a b : Int} (H : a ≤ b) : a < b + 1 := Int.add_le_add_right H 1
Strict Inequality from Non-Strict Inequality in Integers via Successor
Informal description
For any integers $a$ and $b$ such that $a \leq b$, it holds that $a < b + 1$.
Int.le_of_lt_add_one theorem
{a b : Int} (H : a < b + 1) : a ≤ b
Full source
theorem le_of_lt_add_one {a b : Int} (H : a < b + 1) : a ≤ b := Int.le_of_add_le_add_right H
Non-Strict Inequality from Strict Inequality via Successor in Integers
Informal description
For any integers $a$ and $b$, if $a < b + 1$, then $a \leq b$.
Int.sub_one_lt_of_le theorem
{a b : Int} (H : a ≤ b) : a - 1 < b
Full source
theorem sub_one_lt_of_le {a b : Int} (H : a ≤ b) : a - 1 < b :=
  Int.sub_right_lt_of_lt_add <| lt_add_one_of_le H
Subtraction of One Preserves Inequality in Integers: $a \leq b \to a - 1 < b$
Informal description
For any integers $a$ and $b$ such that $a \leq b$, it holds that $a - 1 < b$.
Int.le_of_sub_one_lt theorem
{a b : Int} (H : a - 1 < b) : a ≤ b
Full source
theorem le_of_sub_one_lt {a b : Int} (H : a - 1 < b) : a ≤ b :=
  le_of_lt_add_one <| Int.lt_add_of_sub_right_lt H
Non-Strict Inequality from Subtraction and Strict Inequality: $a - 1 < b \implies a \leq b$
Informal description
For any integers $a$ and $b$, if $a - 1 < b$, then $a \leq b$.
Int.le_sub_one_of_lt theorem
{a b : Int} (H : a < b) : a ≤ b - 1
Full source
theorem le_sub_one_of_lt {a b : Int} (H : a < b) : a ≤ b - 1 := Int.le_sub_right_of_add_le H
Integer Strict Inequality Implies Non-Strict Inequality with Decrement: $a < b \implies a \leq b - 1$
Informal description
For any integers $a$ and $b$, if $a < b$, then $a \leq b - 1$.
Int.lt_of_le_sub_one theorem
{a b : Int} (H : a ≤ b - 1) : a < b
Full source
theorem lt_of_le_sub_one {a b : Int} (H : a ≤ b - 1) : a < b := Int.add_le_of_le_sub_right H
Strict Inequality from Subtracting One: $a \leq b - 1 \implies a < b$
Informal description
For any integers $a$ and $b$, if $a \leq b - 1$, then $a < b$.
Int.mul_lt_mul theorem
{a b c d : Int} (h₁ : a < c) (h₂ : b ≤ d) (h₃ : 0 < b) (h₄ : 0 ≤ c) : a * b < c * d
Full source
protected theorem mul_lt_mul {a b c d : Int}
    (h₁ : a < c) (h₂ : b ≤ d) (h₃ : 0 < b) (h₄ : 0 ≤ c) : a * b < c * d :=
  Int.lt_of_lt_of_le (Int.mul_lt_mul_of_pos_right h₁ h₃) (Int.mul_le_mul_of_nonneg_left h₂ h₄)
Strict Inequality Preservation under Integer Multiplication: $a < c$, $b \leq d$, $0 < b$, $0 \leq c$ $\implies$ $a \cdot b < c \cdot d$
Informal description
For any integers $a$, $b$, $c$, and $d$ such that $a < c$, $b \leq d$, $0 < b$, and $0 \leq c$, the product $a \cdot b$ is strictly less than $c \cdot d$, i.e., $a \cdot b < c \cdot d$.
Int.mul_lt_mul' theorem
{a b c d : Int} (h₁ : a ≤ c) (h₂ : b < d) (h₃ : 0 ≤ b) (h₄ : 0 < c) : a * b < c * d
Full source
protected theorem mul_lt_mul' {a b c d : Int}
    (h₁ : a ≤ c) (h₂ : b < d) (h₃ : 0 ≤ b) (h₄ : 0 < c) : a * b < c * d :=
  Int.lt_of_le_of_lt (Int.mul_le_mul_of_nonneg_right h₁ h₃) (Int.mul_lt_mul_of_pos_left h₂ h₄)
Strict Inequality Preservation under Integer Multiplication with Mixed Conditions: $a \leq c$, $b < d$, $0 \leq b$, $0 < c$ $\implies$ $a \cdot b < c \cdot d$
Informal description
For any integers $a$, $b$, $c$, and $d$ such that $a \leq c$, $b < d$, $0 \leq b$, and $0 < c$, the product $a \cdot b$ is strictly less than $c \cdot d$, i.e., $a \cdot b < c \cdot d$.
Int.mul_neg_of_pos_of_neg theorem
{a b : Int} (ha : 0 < a) (hb : b < 0) : a * b < 0
Full source
protected theorem mul_neg_of_pos_of_neg {a b : Int} (ha : 0 < a) (hb : b < 0) : a * b < 0 := by
  have h : a * b < a * 0 := Int.mul_lt_mul_of_pos_left hb ha
  rwa [Int.mul_zero] at h
Product of Positive and Negative Integers is Negative: $a > 0 \land b < 0 \implies a \cdot b < 0$
Informal description
For any integers $a$ and $b$ such that $a > 0$ and $b < 0$, the product $a \cdot b$ is strictly less than zero, i.e., $a \cdot b < 0$.
Int.mul_neg_of_neg_of_pos theorem
{a b : Int} (ha : a < 0) (hb : 0 < b) : a * b < 0
Full source
protected theorem mul_neg_of_neg_of_pos {a b : Int} (ha : a < 0) (hb : 0 < b) : a * b < 0 := by
  have h : a * b < 0 * b := Int.mul_lt_mul_of_pos_right ha hb
  rwa [Int.zero_mul] at h
Product of Negative and Positive Integers is Negative: $a < 0 \land 0 < b \implies a \cdot b < 0$
Informal description
For any integers $a$ and $b$ such that $a < 0$ and $0 < b$, the product $a \cdot b$ is strictly less than zero, i.e., $a \cdot b < 0$.
Int.mul_nonneg_of_nonpos_of_nonpos theorem
{a b : Int} (ha : a ≤ 0) (hb : b ≤ 0) : 0 ≤ a * b
Full source
protected theorem mul_nonneg_of_nonpos_of_nonpos {a b : Int}
  (ha : a ≤ 0) (hb : b ≤ 0) : 0 ≤ a * b := by
  have : 0 * b ≤ a * b := Int.mul_le_mul_of_nonpos_right ha hb
  rwa [Int.zero_mul] at this
Product of Nonpositive Integers is Nonnegative: $a \leq 0 \land b \leq 0 \implies 0 \leq a \cdot b$
Informal description
For any integers $a$ and $b$ such that $a \leq 0$ and $b \leq 0$, the product $a \cdot b$ is nonnegative, i.e., $0 \leq a \cdot b$.
Int.mul_lt_mul_of_neg_left theorem
{a b c : Int} (h : b < a) (hc : c < 0) : c * a < c * b
Full source
protected theorem mul_lt_mul_of_neg_left {a b c : Int} (h : b < a) (hc : c < 0) : c * a < c * b :=
  have : -c > 0 := Int.neg_pos_of_neg hc
  have : -c * b < -c * a := Int.mul_lt_mul_of_pos_left h this
  have : -(c * b) < -(c * a) := by
    rwa [← Int.neg_mul_eq_neg_mul, ← Int.neg_mul_eq_neg_mul] at this
  Int.lt_of_neg_lt_neg this
Left Multiplication Reverses Strict Inequality for Negative Integers: $b < a \land c < 0 \implies c \cdot a < c \cdot b$
Informal description
For any integers $a$, $b$, and $c$ such that $b < a$ and $c < 0$, the product $c \cdot a$ is strictly less than $c \cdot b$.
Int.mul_lt_mul_of_neg_right theorem
{a b c : Int} (h : b < a) (hc : c < 0) : a * c < b * c
Full source
protected theorem mul_lt_mul_of_neg_right {a b c : Int} (h : b < a) (hc : c < 0) : a * c < b * c :=
  have : -c > 0 := Int.neg_pos_of_neg hc
  have : b * -c < a * -c := Int.mul_lt_mul_of_pos_right h this
  have : -(b * c) < -(a * c) := by
    rwa [← Int.neg_mul_eq_mul_neg, ← Int.neg_mul_eq_mul_neg] at this
  Int.lt_of_neg_lt_neg this
Right Multiplication Reverses Strict Inequality for Negative Integers: $b < a \land c < 0 \implies a \cdot c < b \cdot c$
Informal description
For any integers $a$, $b$, and $c$ such that $b < a$ and $c < 0$, the product $a \cdot c$ is strictly less than $b \cdot c$.
Int.mul_pos_of_neg_of_neg theorem
{a b : Int} (ha : a < 0) (hb : b < 0) : 0 < a * b
Full source
protected theorem mul_pos_of_neg_of_neg {a b : Int} (ha : a < 0) (hb : b < 0) : 0 < a * b := by
  have : 0 * b < a * b := Int.mul_lt_mul_of_neg_right ha hb
  rwa [Int.zero_mul] at this
Product of Two Negative Integers is Positive: $a < 0 \land b < 0 \implies 0 < a \cdot b$
Informal description
For any integers $a$ and $b$ such that $a < 0$ and $b < 0$, the product $a \cdot b$ is strictly positive, i.e., $0 < a \cdot b$.
Int.mul_self_le_mul_self theorem
{a b : Int} (h1 : 0 ≤ a) (h2 : a ≤ b) : a * a ≤ b * b
Full source
protected theorem mul_self_le_mul_self {a b : Int} (h1 : 0 ≤ a) (h2 : a ≤ b) : a * a ≤ b * b :=
  Int.mul_le_mul h2 h2 h1 (Int.le_trans h1 h2)
Square Monotonicity for Non-Negative Integers: $0 \leq a \leq b \implies a^2 \leq b^2$
Informal description
For any integers $a$ and $b$ such that $0 \leq a$ and $a \leq b$, it holds that $a^2 \leq b^2$.
Int.mul_self_lt_mul_self theorem
{a b : Int} (h1 : 0 ≤ a) (h2 : a < b) : a * a < b * b
Full source
protected theorem mul_self_lt_mul_self {a b : Int} (h1 : 0 ≤ a) (h2 : a < b) : a * a < b * b :=
  Int.mul_lt_mul' (Int.le_of_lt h2) h2 h1 (Int.lt_of_le_of_lt h1 h2)
Square Preserves Strict Inequality for Non-Negative Integers: $0 \leq a < b \implies a^2 < b^2$
Informal description
For any integers $a$ and $b$ such that $0 \leq a$ and $a < b$, it holds that $a^2 < b^2$.
Int.nonneg_of_mul_nonneg_left theorem
{a b : Int} (h : 0 ≤ a * b) (hb : 0 < b) : 0 ≤ a
Full source
protected theorem nonneg_of_mul_nonneg_left {a b : Int}
    (h : 0 ≤ a * b) (hb : 0 < b) : 0 ≤ a :=
  Int.le_of_not_gt fun ha => Int.not_le_of_gt (Int.mul_neg_of_neg_of_pos ha hb) h
Non-negativity of Factor Implied by Non-negative Product and Positive Other Factor in Integers
Informal description
For any integers $a$ and $b$, if the product $a \cdot b$ is non-negative ($0 \leq a \cdot b$) and $b$ is positive ($0 < b$), then $a$ is non-negative ($0 \leq a$).
Int.nonneg_of_mul_nonneg_right theorem
{a b : Int} (h : 0 ≤ a * b) (ha : 0 < a) : 0 ≤ b
Full source
protected theorem nonneg_of_mul_nonneg_right {a b : Int}
    (h : 0 ≤ a * b) (ha : 0 < a) : 0 ≤ b :=
  Int.le_of_not_gt fun hb => Int.not_le_of_gt (Int.mul_neg_of_pos_of_neg ha hb) h
Non-negativity of Right Factor in Non-negative Product with Positive Left Factor
Informal description
For any integers $a$ and $b$, if the product $a \cdot b$ is non-negative (i.e., $0 \leq a \cdot b$) and $a$ is strictly positive (i.e., $0 < a$), then $b$ is non-negative (i.e., $0 \leq b$).
Int.nonpos_of_mul_nonpos_left theorem
{a b : Int} (h : a * b ≤ 0) (hb : 0 < b) : a ≤ 0
Full source
protected theorem nonpos_of_mul_nonpos_left {a b : Int}
    (h : a * b ≤ 0) (hb : 0 < b) : a ≤ 0 :=
  Int.le_of_not_gt fun ha : a > 0 => Int.not_le_of_gt (Int.mul_pos ha hb) h
Non-positivity of Left Factor in Non-positive Product with Positive Right Factor
Informal description
For any integers $a$ and $b$, if the product $a \cdot b$ is non-positive (i.e., $a \cdot b \leq 0$) and $b$ is strictly positive (i.e., $0 < b$), then $a$ is non-positive (i.e., $a \leq 0$).
Int.nonpos_of_mul_nonpos_right theorem
{a b : Int} (h : a * b ≤ 0) (ha : 0 < a) : b ≤ 0
Full source
protected theorem nonpos_of_mul_nonpos_right {a b : Int}
    (h : a * b ≤ 0) (ha : 0 < a) : b ≤ 0 :=
  Int.le_of_not_gt fun hb : b > 0 => Int.not_le_of_gt (Int.mul_pos ha hb) h
Non-positivity of Right Factor in Non-positive Product with Positive Left Factor
Informal description
For any integers $a$ and $b$, if the product $a \cdot b$ is non-positive (i.e., $a \cdot b \leq 0$) and $a$ is strictly positive (i.e., $0 < a$), then $b$ is non-positive (i.e., $b \leq 0$).
Int.sign_zero theorem
: sign 0 = 0
Full source
@[simp] theorem sign_zero : sign 0 = 0 := rfl
Sign of Zero: $\text{sign}(0) = 0$
Informal description
The sign function evaluated at $0$ equals $0$, i.e., $\text{sign}(0) = 0$.
Int.sign_one theorem
: sign 1 = 1
Full source
@[simp] theorem sign_one : sign 1 = 1 := rfl
Sign of One: $\text{sign}(1) = 1$
Informal description
The sign function evaluated at $1$ equals $1$, i.e., $\text{sign}(1) = 1$.
Int.sign_neg_one theorem
: sign (-1) = -1
Full source
theorem sign_neg_one : sign (-1) = -1 := rfl
Sign of Negative One: $\text{sign}(-1) = -1$
Informal description
The sign function evaluated at $-1$ equals $-1$, i.e., $\text{sign}(-1) = -1$.
Int.sign_of_add_one theorem
(x : Nat) : Int.sign (x + 1) = 1
Full source
@[simp] theorem sign_of_add_one (x : Nat) : Int.sign (x + 1) = 1 := rfl
Sign of Successor: $\text{sign}(x + 1) = 1$ for natural $x$
Informal description
For any natural number $x$, the sign function evaluated at $x + 1$ equals $1$, i.e., $\text{sign}(x + 1) = 1$.
Int.sign_negSucc theorem
(x : Nat) : Int.sign (Int.negSucc x) = -1
Full source
@[simp] theorem sign_negSucc (x : Nat) : Int.sign (Int.negSucc x) = -1 := rfl
Sign of Negative Successor: $\text{sign}(-x - 1) = -1$
Informal description
For any natural number $x$, the sign of the negative successor of $x$ (i.e., $-x - 1$) is $-1$, i.e., $\text{sign}(-x - 1) = -1$.
Int.natAbs_sign theorem
(z : Int) : z.sign.natAbs = if z = 0 then 0 else 1
Full source
theorem natAbs_sign (z : Int) : z.sign.natAbs = if z = 0 then 0 else 1 :=
  match z with | 0 | succ _ | -[_+1] => rfl
Natural Absolute Value of Integer Sign: $\text{natAbs}(\text{sign}(z)) = \text{if } z = 0 \text{ then } 0 \text{ else } 1$
Informal description
For any integer $z$, the natural absolute value of its sign is equal to $0$ if $z = 0$ and $1$ otherwise, i.e., \[ \text{natAbs}(\text{sign}(z)) = \begin{cases} 0 & \text{if } z = 0, \\ 1 & \text{otherwise.} \end{cases} \]
Int.natAbs_sign_of_nonzero theorem
{z : Int} (hz : z ≠ 0) : z.sign.natAbs = 1
Full source
theorem natAbs_sign_of_nonzero {z : Int} (hz : z ≠ 0) : z.sign.natAbs = 1 := by
  rw [Int.natAbs_sign, if_neg hz]
Natural Absolute Value of Sign for Nonzero Integers: $\text{natAbs}(\text{sign}(z)) = 1$ when $z \neq 0$
Informal description
For any nonzero integer $z$, the natural absolute value of its sign is equal to $1$, i.e., $\text{natAbs}(\text{sign}(z)) = 1$.
Int.sign_ofNat_of_nonzero theorem
{n : Nat} (hn : n ≠ 0) : Int.sign n = 1
Full source
theorem sign_ofNat_of_nonzero {n : Nat} (hn : n ≠ 0) : Int.sign n = 1 :=
  match n, Nat.exists_eq_succ_of_ne_zero hn with
  | _, ⟨n, rfl⟩ => Int.sign_of_add_one n
Sign of Nonzero Natural Number: $\text{sign}(n) = 1$ for $n \neq 0$
Informal description
For any natural number $n \neq 0$, the sign function evaluated at $n$ equals $1$, i.e., $\text{sign}(n) = 1$.
Int.sign_neg theorem
(z : Int) : Int.sign (-z) = -Int.sign z
Full source
@[simp] theorem sign_neg (z : Int) : Int.sign (-z) = -Int.sign z := by
  match z with | 0 | succ _ | -[_+1] => rfl
Sign of Negated Integer: $\text{sign}(-z) = -\text{sign}(z)$
Informal description
For any integer $z$, the sign of $-z$ is equal to the negation of the sign of $z$, i.e., $\text{sign}(-z) = -\text{sign}(z)$.
Int.sign_mul_natAbs theorem
: ∀ a : Int, sign a * natAbs a = a
Full source
theorem sign_mul_natAbs : ∀ a : Int, sign a * natAbs a = a
  | 0      => rfl
  | succ _ => Int.one_mul _
  | -[_+1] => (Int.neg_eq_neg_one_mul _).symm
Sign-Natural Absolute Value Decomposition of Integers: $\text{sign}(a) \cdot \text{natAbs}(a) = a$
Informal description
For any integer $a$, the product of the sign of $a$ and the natural absolute value of $a$ equals $a$, i.e., $\text{sign}(a) \cdot \text{natAbs}(a) = a$.
Int.sign_mul theorem
: ∀ a b, sign (a * b) = sign a * sign b
Full source
@[simp] theorem sign_mul : ∀ a b, sign (a * b) = sign a * sign b
  | a, 0 | 0, b => by simp [Int.mul_zero, Int.zero_mul]
  | succ _, succ _ | succ _, -[_+1] | -[_+1], succ _ | -[_+1], -[_+1] => rfl
Sign of Product Equals Product of Signs in Integers: $\text{sign}(a \times b) = \text{sign}(a) \times \text{sign}(b)$
Informal description
For any integers $a$ and $b$, the sign of the product $a \times b$ is equal to the product of the signs of $a$ and $b$, i.e., $\text{sign}(a \times b) = \text{sign}(a) \times \text{sign}(b)$.
Int.sign_eq_one_of_pos theorem
{a : Int} (h : 0 < a) : sign a = 1
Full source
theorem sign_eq_one_of_pos {a : Int} (h : 0 < a) : sign a = 1 :=
  match a, eq_succ_of_zero_lt h with
  | _, ⟨_, rfl⟩ => rfl
Sign of Positive Integer is $1$
Informal description
For any integer $a$, if $a$ is positive (i.e., $0 < a$), then the sign of $a$ is $1$.
Int.sign_eq_neg_one_of_neg theorem
{a : Int} (h : a < 0) : sign a = -1
Full source
theorem sign_eq_neg_one_of_neg {a : Int} (h : a < 0) : sign a = -1 :=
  match a, eq_negSucc_of_lt_zero h with
  | _, ⟨_, rfl⟩ => rfl
Sign of Negative Integer is $-1$
Informal description
For any integer $a$, if $a$ is negative (i.e., $a < 0$), then the sign of $a$ is $-1$.
Int.pos_of_sign_eq_one theorem
: ∀ {a : Int}, sign a = 1 → 0 < a
Full source
theorem pos_of_sign_eq_one : ∀ {a : Int}, sign a = 1 → 0 < a
  | (_ + 1 : Nat), _ => ofNat_lt.2 (Nat.succ_pos _)
Positive Integers Have Sign One
Informal description
For any integer $a$, if the sign of $a$ is equal to $1$, then $a$ is positive, i.e., $0 < a$.
Int.neg_of_sign_eq_neg_one theorem
: ∀ {a : Int}, sign a = -1 → a < 0
Full source
theorem neg_of_sign_eq_neg_one : ∀ {a : Int}, sign a = -1 → a < 0
  | (_ + 1 : Nat), h => nomatch h
  | 0, h => nomatch h
  | -[_+1], _ => negSucc_lt_zero _
Negative Integers Have Sign Negative One
Informal description
For any integer $a$, if the sign of $a$ is $-1$, then $a$ is negative, i.e., $a < 0$.
Int.sign_sign theorem
: sign (sign x) = sign x
Full source
@[simp] theorem sign_sign : sign (sign x) = sign x := by
  match x with
  | 0 => rfl
  | .ofNat (_ + 1) => rfl
  | .negSucc _ => rfl
Idempotence of the Sign Function on Integers: $\text{sign}(\text{sign}(x)) = \text{sign}(x)$
Informal description
For any integer $x$, the sign of the sign of $x$ is equal to the sign of $x$, i.e., $\text{sign}(\text{sign}(x)) = \text{sign}(x)$.
Int.sign_nonneg_iff theorem
: 0 ≤ sign x ↔ 0 ≤ x
Full source
@[simp] theorem sign_nonneg_iff : 0 ≤ sign x ↔ 0 ≤ x := by
  match x with
  | 0 => rfl
  | .ofNat (_ + 1) =>
    simp +decide only [sign, true_iff]
    exact Int.le_add_one (ofNat_nonneg _)
  | .negSucc _ => simp +decide [sign]
Nonnegativity Criterion for Integer Sign Function: $0 \leq \text{sign}(x) \leftrightarrow 0 \leq x$
Informal description
For any integer $x$, the sign of $x$ is nonnegative if and only if $x$ is nonnegative, i.e., $0 \leq \text{sign}(x) \leftrightarrow 0 \leq x$.
Int.sign_nonneg abbrev
Full source
@[deprecated sign_nonneg_iff (since := "2025-03-11")] abbrev sign_nonneg := @sign_nonneg_iff
Nonnegativity of Sign for Nonnegative Integers: $0 \leq x \implies 0 \leq \text{sign}(x)$
Informal description
For any integer $x$, if $x$ is nonnegative (i.e., $0 \leq x$), then the sign of $x$ is also nonnegative (i.e., $0 \leq \text{sign}(x)$).
Int.sign_pos_iff theorem
: 0 < sign x ↔ 0 < x
Full source
@[simp] theorem sign_pos_iff : 0 < sign x ↔ 0 < x := by
  match x with
  | 0
  | .ofNat (_ + 1) => simp
  | .negSucc x => simp
Sign Positivity Criterion for Integers: $0 < \text{sign}(x) \leftrightarrow 0 < x$
Informal description
For any integer $x$, the sign of $x$ is positive if and only if $x$ is positive, i.e., $0 < \text{sign}(x) \leftrightarrow 0 < x$.
Int.sign_nonpos_iff theorem
: sign x ≤ 0 ↔ x ≤ 0
Full source
@[simp] theorem sign_nonpos_iff : signsign x ≤ 0 ↔ x ≤ 0 := by
  match x with
  | 0 => rfl
  | .ofNat (_ + 1) => simp
  | .negSucc _ => simpa using negSucc_le_zero _
Sign Nonpositivity Criterion for Integers: $\text{sign}(x) \leq 0 \leftrightarrow x \leq 0$
Informal description
For any integer $x$, the sign of $x$ is nonpositive if and only if $x$ is nonpositive, i.e., $\text{sign}(x) \leq 0 \leftrightarrow x \leq 0$.
Int.sign_neg_iff theorem
: sign x < 0 ↔ x < 0
Full source
@[simp] theorem sign_neg_iff : signsign x < 0 ↔ x < 0 := by
  match x with
  | 0 => simp
  | .ofNat (_ + 1) => simpa using le.intro_sub _ rfl
  | .negSucc _ => simp
Sign of Integer is Negative if and only if Integer is Negative
Informal description
For any integer $x$, the sign of $x$ is negative if and only if $x$ is negative, i.e., $\text{sign}(x) < 0 \leftrightarrow x < 0$.
Int.mul_sign_self theorem
: ∀ i : Int, i * sign i = natAbs i
Full source
@[simp] theorem mul_sign_self : ∀ i : Int, i * sign i = natAbs i
  | succ _ => Int.mul_one _
  | 0 => Int.mul_zero _
  | -[_+1] => Int.mul_neg_one _
Integer Sign-Absolute Value Identity: $i \cdot \text{sign}(i) = |i|_{\mathbb{N}}$
Informal description
For any integer $i$, the product of $i$ and its sign equals the absolute value of $i$ as a natural number, i.e., $i \cdot \text{sign}(i) = |i|_{\mathbb{N}}$.
Int.mul_sign abbrev
Full source
@[deprecated mul_sign_self (since := "2025-02-24")] abbrev mul_sign := @mul_sign_self
Commutativity of Integer Sign Multiplication: $i \cdot \text{sign}(i) = \text{sign}(i) \cdot i$
Informal description
For any integer $i$, the product of $i$ and its sign equals the product of its sign and $i$, i.e., $i \cdot \text{sign}(i) = \text{sign}(i) \cdot i$.
Int.sign_mul_self theorem
: sign i * i = natAbs i
Full source
@[simp] theorem sign_mul_self : sign i * i = natAbs i := by
  rw [Int.mul_comm, mul_sign_self]
Sign-Absolute Value Identity for Integers: $\text{sign}(i) \cdot i = |i|_{\mathbb{N}}$
Informal description
For any integer $i$, the product of the sign of $i$ and $i$ itself equals the absolute value of $i$ as a natural number, i.e., $\text{sign}(i) \cdot i = |i|_{\mathbb{N}}$.
Int.sign_trichotomy theorem
(a : Int) : sign a = 1 ∨ sign a = 0 ∨ sign a = -1
Full source
theorem sign_trichotomy (a : Int) : signsign a = 1 ∨ sign a = 0 ∨ sign a = -1 := by
  match a with
  | 0 => simp
  | .ofNat (_ + 1) => simp
  | .negSucc _ => simp
Trichotomy of Integer Sign Function: $\text{sign}(a) \in \{1, 0, -1\}$
Informal description
For any integer $a$, the sign of $a$ is either $1$, $0$, or $-1$.
Int.natAbs_ne_zero theorem
{a : Int} : a.natAbs ≠ 0 ↔ a ≠ 0
Full source
theorem natAbs_ne_zero {a : Int} : a.natAbs ≠ 0a.natAbs ≠ 0 ↔ a ≠ 0 := not_congr Int.natAbs_eq_zero
Nonzero Absolute Value Criterion for Integers
Informal description
For any integer $a$, the absolute value of $a$ as a natural number is nonzero if and only if $a$ is nonzero, i.e., $\text{natAbs}(a) \neq 0 \leftrightarrow a \neq 0$.
Int.natAbs_mul_self theorem
: ∀ {a : Int}, ↑(natAbs a * natAbs a) = a * a
Full source
theorem natAbs_mul_self : ∀ {a : Int}, ↑(natAbs a * natAbs a) = a * a
  | ofNat _ => rfl
  | -[_+1]  => rfl
Square of Integer's Absolute Value Equals Square of Integer
Informal description
For any integer $a$, the natural number obtained by casting the product of the absolute values of $a$ with itself equals the product of $a$ with itself, i.e., $\text{natAbs}(a) \cdot \text{natAbs}(a) = a \cdot a$.
Int.natAbs_mul_natAbs_eq theorem
{a b : Int} {c : Nat} (h : a * b = (c : Int)) : a.natAbs * b.natAbs = c
Full source
theorem natAbs_mul_natAbs_eq {a b : Int} {c : Nat}
    (h : a * b = (c : Int)) : a.natAbs * b.natAbs = c := by rw [← natAbs_mul, h, natAbs.eq_def]
Absolute Value Product Condition for Integer Multiplication: $\text{natAbs}(a) \cdot \text{natAbs}(b) = c$ when $a \cdot b = c$
Informal description
For any integers $a$ and $b$ and any natural number $c$, if the product $a \cdot b$ equals the canonical image of $c$ in the integers, then the product of the absolute values of $a$ and $b$ (as natural numbers) equals $c$. That is, if $a \cdot b = c$ (as integers), then $\text{natAbs}(a) \cdot \text{natAbs}(b) = c$.
Int.natAbs_mul_self' theorem
(a : Int) : (natAbs a * natAbs a : Int) = a * a
Full source
@[simp] theorem natAbs_mul_self' (a : Int) : (natAbs a * natAbs a : Int) = a * a := by
  rw [← Int.ofNat_mul, natAbs_mul_self]
Square of Integer's Absolute Value Equals Square of Integer (Cast Version)
Informal description
For any integer $a$, the product of the natural number absolute value of $a$ with itself, when cast to the integers, equals the product of $a$ with itself. That is, $(\text{natAbs}(a) \cdot \text{natAbs}(a) : \mathbb{Z}) = a \cdot a$.
Int.natAbs_eq_iff theorem
{a : Int} {n : Nat} : a.natAbs = n ↔ a = n ∨ a = -↑n
Full source
theorem natAbs_eq_iff {a : Int} {n : Nat} : a.natAbs = n ↔ a = n ∨ a = -↑n := by
  rw [← Int.natAbs_eq_natAbs_iff, Int.natAbs_ofNat]
Integer Absolute Value Equivalence: $\text{natAbs}(a) = n \leftrightarrow a = n \lor a = -n$
Informal description
For any integer $a$ and natural number $n$, the absolute value of $a$ (as a natural number) equals $n$ if and only if $a$ is equal to $n$ or $a$ is equal to $-n$. That is, $\text{natAbs}(a) = n \leftrightarrow (a = n \lor a = -n)$.
Int.natAbs_add_le theorem
(a b : Int) : natAbs (a + b) ≤ natAbs a + natAbs b
Full source
theorem natAbs_add_le (a b : Int) : natAbs (a + b) ≤ natAbs a + natAbs b := by
  suffices ∀ a b : Nat, natAbs (subNatNat a b.succ) ≤ (a + b).succ by
    match a, b with
    | (a:Nat), (b:Nat) => rw [← ofNat_add, natAbs_ofNat]; apply Nat.le_refl
    | (a:Nat), -[b+1]  => rw [natAbs_ofNat, natAbs_negSucc]; apply this
    | -[a+1],  (b:Nat) =>
      rw [natAbs_negSucc, natAbs_ofNat, Nat.succ_add, Nat.add_comm a b]; apply this
    | -[a+1],  -[b+1]  => rw [natAbs_negSucc, succ_add]; apply Nat.le_refl
  refine fun a b => subNatNat_elim a b.succ
    (fun m n i => n = b.succnatAbs i ≤ (m + b).succ) ?_
    (fun i n (e : (n + i).succ = _) => ?_) rfl
  · intro i n h
    subst h
    rw [Nat.add_comm _ i, Nat.add_assoc]
    exact Nat.le_add_right i (b.succ + b).succ
  · apply succ_le_succ
    rw [← succ.inj e, ← Nat.add_assoc, Nat.add_comm]
    apply Nat.le_add_right
Triangle Inequality for Integer Absolute Values: $\text{natAbs}(a + b) \leq \text{natAbs}(a) + \text{natAbs}(b)$
Informal description
For any integers $a$ and $b$, the absolute value (as a natural number) of their sum satisfies $\text{natAbs}(a + b) \leq \text{natAbs}(a) + \text{natAbs}(b)$.
Int.natAbs_sub_le theorem
(a b : Int) : natAbs (a - b) ≤ natAbs a + natAbs b
Full source
theorem natAbs_sub_le (a b : Int) : natAbs (a - b) ≤ natAbs a + natAbs b := by
  rw [← Int.natAbs_neg b]; apply natAbs_add_le
Triangle Inequality for Integer Absolute Values of Differences: $\text{natAbs}(a - b) \leq \text{natAbs}(a) + \text{natAbs}(b)$
Informal description
For any integers $a$ and $b$, the absolute value (as a natural number) of their difference satisfies $\text{natAbs}(a - b) \leq \text{natAbs}(a) + \text{natAbs}(b)$.
Int.negSucc_eq' theorem
(m : Nat) : -[m+1] = -m - 1
Full source
@[deprecated negSucc_eq (since := "2025-03-11")]
theorem negSucc_eq' (m : Nat) : -[m+1] = -m - 1 := by simp only [negSucc_eq, Int.neg_add]; rfl
Negative Successor Identity: $- [m+1] = -m - 1$
Informal description
For any natural number $m$, the negative successor of $m$ (denoted $- [m+1]$) is equal to $-m - 1$.
Int.natAbs_lt_natAbs_of_nonneg_of_lt theorem
{a b : Int} (w₁ : 0 ≤ a) (w₂ : a < b) : a.natAbs < b.natAbs
Full source
theorem natAbs_lt_natAbs_of_nonneg_of_lt {a b : Int}
    (w₁ : 0 ≤ a) (w₂ : a < b) : a.natAbs < b.natAbs :=
  match a, b, eq_ofNat_of_zero_le w₁, eq_ofNat_of_zero_le (Int.le_trans w₁ (Int.le_of_lt w₂)) with
  | _, _, ⟨_, rfl⟩, ⟨_, rfl⟩ => ofNat_lt.1 w₂
Absolute Value Inequality for Nonnegative Integers: $0 \leq a < b \implies \text{natAbs}(a) < \text{natAbs}(b)$
Informal description
For any integers $a$ and $b$ such that $0 \leq a$ and $a < b$, the absolute value of $a$ (as a natural number) is strictly less than the absolute value of $b$ (as a natural number), i.e., $\text{natAbs}(a) < \text{natAbs}(b)$.
Int.natAbs_eq_iff_mul_eq_zero theorem
: natAbs a = n ↔ (a - n) * (a + n) = 0
Full source
theorem natAbs_eq_iff_mul_eq_zero : natAbsnatAbs a = n ↔ (a - n) * (a + n) = 0 := by
  rw [natAbs_eq_iff, Int.mul_eq_zero, ← Int.sub_neg, Int.sub_eq_zero, Int.sub_eq_zero]
Absolute Value Equivalence via Zero Product: $\text{natAbs}(a) = n \leftrightarrow (a - n)(a + n) = 0$
Informal description
For any integer $a$ and natural number $n$, the absolute value of $a$ (as a natural number) equals $n$ if and only if the product $(a - n)(a + n)$ equals zero. That is, $\text{natAbs}(a) = n \leftrightarrow (a - n)(a + n) = 0$.