doc-next-gen

Init.Data.Nat.Basic

Module docstring

{"# Helper \"packing\" theorems ","# Helper Bool relation theorems ","# Nat.add theorems ","# Nat.mul theorems ","# Inequalities ","### le/lt ","## zero/one/two ","## succ/pred ","# Basic theorems for comparing numerals ","# mul + order ","# power ","# min/max ","# Auxiliary theorems for well-founded recursion ","# pred theorems ","# sub theorems ","## Mul sub distrib ","# Helper normalization theorems "}

Nat.recAux abbrev
{motive : Nat → Sort u} (zero : motive 0) (succ : (n : Nat) → motive n → motive (n + 1)) (t : Nat) : motive t
Full source
/--
A recursor for `Nat` that uses the notations `0` for `Nat.zero` and `n + 1` for `Nat.succ`.

It is otherwise identical to the default recursor `Nat.rec`. It is used by the `induction` tactic
by default for `Nat`.
-/
@[elab_as_elim, induction_eliminator]
protected abbrev recAux {motive : Nat → Sort u} (zero : motive 0) (succ : (n : Nat) → motive n → motive (n + 1)) (t : Nat) : motive t :=
  Nat.rec zero succ t
Auxiliary Natural Number Recursor
Informal description
The auxiliary recursion principle for natural numbers states that given: - A motive `motive : ℕ → Sort u` (a type-valued function on natural numbers) - A base case `zero : motive 0` - An inductive step `succ : ∀ (n : ℕ), motive n → motive (n + 1)` - A natural number `t : ℕ` Then we can construct an element of type `motive t`.
Nat.casesAuxOn abbrev
{motive : Nat → Sort u} (t : Nat) (zero : motive 0) (succ : (n : Nat) → motive (n + 1)) : motive t
Full source
/--
A case analysis principle for `Nat` that uses the notations `0` for `Nat.zero` and `n + 1` for
`Nat.succ`.

It is otherwise identical to the default recursor `Nat.casesOn`. It is used as the default `Nat`
case analysis principle for `Nat` by the `cases` tactic.
-/
@[elab_as_elim, cases_eliminator]
protected abbrev casesAuxOn {motive : Nat → Sort u} (t : Nat) (zero : motive 0) (succ : (n : Nat) → motive (n + 1)) : motive t :=
  Nat.casesOn t zero succ
Case Analysis Principle for Natural Numbers (Zero/Successor Form)
Informal description
Given a type family `motive` indexed by natural numbers, a natural number `t`, a term `zero` of type `motive 0`, and a function `succ` that for any natural number `n` produces a term of type `motive (n + 1)` from a term of type `motive n`, the function `Nat.casesAuxOn` produces a term of type `motive t`. This serves as a case analysis principle for natural numbers, distinguishing between the base case `0` and the successor case `n + 1`.
Nat.repeat definition
{α : Type u} (f : α → α) : (n : Nat) → (a : α) → α
Full source
/--
Applies a function to a starting value the specified number of times.

In other words, `f` is iterated `n` times on `a`.

Examples:
* `Nat.repeat f 3 a = f <| f <| f <| a`
* `Nat.repeat (· ++ "!") 4 "Hello" = "Hello!!!!"`
-/
@[specialize] def repeat {α : Type u} (f : α → α) : (n : Nat) → (a : α) → α
  | 0,      a => a
  | succ n, a => f (repeat f n a)
Iterated function application
Informal description
Given a function \( f : \alpha \to \alpha \) and a natural number \( n \), the function `Nat.repeat f n` applies \( f \) to a starting value \( a \) exactly \( n \) times. That is, `Nat.repeat f n a = f (f ... (f a)...)` with \( n \) applications of \( f \).
Nat.repeatTR definition
{α : Type u} (f : α → α) (n : Nat) (a : α) : α
Full source
/--
Applies a function to a starting value the specified number of times.

In other words, `f` is iterated `n` times on `a`.

This is a tail-recursive version of `Nat.repeat` that's used at runtime.

Examples:
* `Nat.repeatTR f 3 a = f <| f <| f <| a`
* `Nat.repeatTR (· ++ "!") 4 "Hello" = "Hello!!!!"`
-/
@[inline] def repeatTR {α : Type u} (f : α → α) (n : Nat) (a : α) : α :=
  let rec @[specialize] loop
    | 0,      a => a
    | succ n, a => loop n (f a)
  loop n a
Tail-recursive function iteration
Informal description
Given a function \( f : \alpha \to \alpha \), a natural number \( n \), and a starting value \( a \), the function `Nat.repeatTR f n a` applies \( f \) to \( a \) exactly \( n \) times in a tail-recursive manner. That is, it computes \( f(f(\dots f(a)\dots)) \) with \( n \) applications of \( f \).
Nat.blt definition
(a b : Nat) : Bool
Full source
/--
The Boolean less-than comparison on natural numbers.

This function is overridden in both the kernel and the compiler to efficiently evaluate using the
arbitrary-precision arithmetic library. The definition provided here is the logical model.

Examples:
 * `Nat.blt 2 5 = true`
 * `Nat.blt 5 2 = false`
 * `Nat.blt 5 5 = false`
-/
def blt (a b : Nat) : Bool :=
  ble a.succ b
Boolean less-than for natural numbers
Informal description
The Boolean less-than comparison on natural numbers. Given two natural numbers \( a \) and \( b \), it returns `true` if \( a < b \) and `false` otherwise. This is implemented as `Nat.ble (a + 1) b`, where `Nat.ble` is the Boolean less-than-or-equal-to comparison. Examples: - `Nat.blt 2 5 = true` - `Nat.blt 5 2 = false` - `Nat.blt 5 5 = false`
Nat.and_forall_add_one theorem
{p : Nat → Prop} : p 0 ∧ (∀ n, p (n + 1)) ↔ ∀ n, p n
Full source
theorem and_forall_add_one {p : Nat → Prop} : p 0 ∧ (∀ n, p (n + 1))p 0 ∧ (∀ n, p (n + 1)) ↔ ∀ n, p n :=
  ⟨fun h n => Nat.casesOn n h.1 h.2, fun h => ⟨h _, fun _ => h _⟩⟩
Induction Principle for Natural Numbers via Predicate and Successor
Informal description
For any predicate $p$ on natural numbers, $p(0)$ holds and for every natural number $n$, $p(n+1)$ holds if and only if $p(n)$ holds for all natural numbers $n$.
Nat.or_exists_add_one theorem
: p 0 ∨ (Exists fun n => p (n + 1)) ↔ Exists p
Full source
theorem or_exists_add_one : p 0 ∨ (Exists fun n => p (n + 1))p 0 ∨ (Exists fun n => p (n + 1)) ↔ Exists p :=
  ⟨fun h => h.elim (fun h0 => ⟨0, h0⟩) fun ⟨n, hn⟩ => ⟨n + 1, hn⟩,
    fun ⟨n, h⟩ => match n with | 0 => Or.inl h | n+1 => Or.inr ⟨n, h⟩⟩
Existence of Natural Number Satisfying Predicate via Base Case or Successor
Informal description
For any predicate $p$ on natural numbers, the statement "$p(0)$ or there exists a natural number $n$ such that $p(n+1)$" is equivalent to the existence of a natural number $n$ satisfying $p(n)$.
Nat.zero_eq theorem
: Nat.zero = 0
Full source
@[simp] theorem zero_eq : Nat.zero = 0 := rfl
Zero as Natural Number: $\text{Nat.zero} = 0$
Informal description
The natural number `Nat.zero` is equal to $0$.
Nat.add_eq theorem
: Nat.add x y = x + y
Full source
@[simp] theorem add_eq : Nat.add x y = x + y := rfl
Equivalence of `Nat.add` and Addition: $\text{Nat.add}(x, y) = x + y$
Informal description
For any natural numbers $x$ and $y$, the function `Nat.add` applied to $x$ and $y$ is equal to the sum $x + y$.
Nat.mul_eq theorem
: Nat.mul x y = x * y
Full source
@[simp] theorem mul_eq : Nat.mul x y = x * y := rfl
Equivalence of `Nat.mul` and Multiplication: $\text{Nat.mul}(x, y) = x * y$
Informal description
For any natural numbers $x$ and $y$, the multiplication function `Nat.mul` applied to $x$ and $y$ is equal to the product $x * y$.
Nat.sub_eq theorem
: Nat.sub x y = x - y
Full source
@[simp] theorem sub_eq : Nat.sub x y = x - y := rfl
Truncated Subtraction Equivalence for Natural Numbers
Informal description
For any natural numbers $x$ and $y$, the truncated subtraction operation `Nat.sub x y` is equal to the standard mathematical notation $x - y$.
Nat.lt_eq theorem
: Nat.lt x y = (x < y)
Full source
@[simp] theorem lt_eq : Nat.lt x y = (x < y) := rfl
Equivalence of Natural Number Less-Than Relation and Standard Notation
Informal description
For any natural numbers $x$ and $y$, the relation `Nat.lt x y` is equivalent to the standard mathematical notation $x < y$.
Nat.le_eq theorem
: Nat.le x y = (x ≤ y)
Full source
@[simp] theorem le_eq : Nat.le x y = (x ≤ y) := rfl
Equivalence of Natural Number Less-Than-Or-Equal Relation and Standard Notation
Informal description
For any natural numbers $x$ and $y$, the relation `Nat.le x y` is equivalent to the standard mathematical notation $x \leq y$.
Nat.beq_refl theorem
(a : Nat) : Nat.beq a a = true
Full source
@[simp] theorem beq_refl (a : Nat) : Nat.beq a a = true := by
  induction a with simp [Nat.beq]
  | succ a ih => simp [ih]
Reflexivity of Natural Number Equality Test: $\text{Nat.beq}(a, a) = \text{true}$
Informal description
For any natural number $a$, the equality test `Nat.beq` evaluates to `true` when comparing $a$ with itself, i.e., $a = a$ holds.
Nat.beq_eq theorem
: (Nat.beq x y = true) = (x = y)
Full source
@[simp] theorem beq_eq : (Nat.beq x y = true) = (x = y) := propext <| Iff.intro Nat.eq_of_beq_eq_true (fun h => h ▸ (Nat.beq_refl x))
Equivalence of Natural Number Equality Test and Equality: $\text{Nat.beq}(x, y) = \text{true} \leftrightarrow x = y$
Informal description
For any natural numbers $x$ and $y$, the equality test `Nat.beq x y` returns `true` if and only if $x = y$.
Nat.ble_eq theorem
: (Nat.ble x y = true) = (x ≤ y)
Full source
@[simp] theorem ble_eq : (Nat.ble x y = true) = (x ≤ y) := propext <| Iff.intro Nat.le_of_ble_eq_true Nat.ble_eq_true_of_le
Boolean Less-Than-or-Equal-to Test Equivalence: $\text{Nat.ble}(x, y) = \text{true} \leftrightarrow x \leq y$
Informal description
For any natural numbers $x$ and $y$, the Boolean comparison `Nat.ble x y` evaluates to `true` if and only if $x \leq y$.
Nat.blt_eq theorem
: (Nat.blt x y = true) = (x < y)
Full source
@[simp] theorem blt_eq : (Nat.blt x y = true) = (x < y) := propext <| Iff.intro Nat.le_of_ble_eq_true Nat.ble_eq_true_of_le
Boolean Less-Than Equivalence: $\text{Nat.blt}(x, y) = \text{true} \leftrightarrow x < y$
Informal description
For any natural numbers $x$ and $y$, the Boolean less-than comparison `Nat.blt x y` evaluates to `true` if and only if $x < y$.
Nat.instLawfulBEq instance
: LawfulBEq Nat
Full source
instance : LawfulBEq Nat where
  eq_of_beq h := by simpa using h
  rfl := by simp [BEq.beq]
Lawful Boolean Equality for Natural Numbers
Informal description
The boolean equality operator `==` on natural numbers is lawful, meaning it coincides with propositional equality. Specifically: 1. For any natural numbers $x$ and $y$, if `x == y` evaluates to `true`, then $x = y$ holds. 2. For any natural number $x$, the test `x == x` evaluates to `true$.
Nat.beq_eq_true_eq theorem
(a b : Nat) : ((a == b) = true) = (a = b)
Full source
theorem beq_eq_true_eq (a b : Nat) : ((a == b) = true) = (a = b) := by simp
Boolean Equality Test for Natural Numbers: $(a == b) = \text{true} \leftrightarrow a = b$
Informal description
For any natural numbers $a$ and $b$, the boolean equality test `a == b` evaluates to `true` if and only if $a = b$.
Nat.not_beq_eq_true_eq theorem
(a b : Nat) : ((!(a == b)) = true) = ¬(a = b)
Full source
theorem not_beq_eq_true_eq (a b : Nat) : ((!(a == b)) = true) = ¬(a = b) := by simp
Negated Boolean Equality Test for Natural Numbers: $\neg(a == b) = \text{true} \leftrightarrow a \neq b$
Informal description
For any natural numbers $a$ and $b$, the boolean negation of the equality test `a == b` evaluates to `true` if and only if $a \neq b$.
Nat.zero_add theorem
: ∀ (n : Nat), 0 + n = n
Full source
@[simp] protected theorem zero_add : ∀ (n : Nat), 0 + n = n
  | 0   => rfl
  | n+1 => congrArg succ (Nat.zero_add n)
Additive Identity for Natural Numbers: $0 + n = n$
Informal description
For any natural number $n$, the sum of $0$ and $n$ equals $n$, i.e., $0 + n = n$.
Nat.instLawfulIdentityHAddOfNat instance
: Std.LawfulIdentity (α := Nat) (· + ·) 0
Full source
instance : Std.LawfulIdentity (α := Nat) (· + ·) 0 where
  left_id := Nat.zero_add
  right_id := Nat.add_zero
$0$ is the Additive Identity for Natural Numbers
Informal description
The natural number $0$ is a verified left and right identity for addition on $\mathbb{N}$. That is, for any natural number $n$, both $0 + n = n$ and $n + 0 = n$ hold.
Nat.succ_add theorem
: ∀ (n m : Nat), (succ n) + m = succ (n + m)
Full source
theorem succ_add : ∀ (n m : Nat), (succ n) + m = succ (n + m)
  | _, 0   => rfl
  | n, m+1 => congrArg succ (succ_add n m)
Successor Addition Identity: $(n + 1) + m = (n + m) + 1$
Informal description
For any natural numbers $n$ and $m$, the sum of the successor of $n$ and $m$ is equal to the successor of the sum of $n$ and $m$, i.e., $(n + 1) + m = (n + m) + 1$.
Nat.add_succ theorem
(n m : Nat) : n + succ m = succ (n + m)
Full source
theorem add_succ (n m : Nat) : n + succ m = succ (n + m) :=
  rfl
Successor Addition Identity: $n + (m + 1) = (n + m) + 1$
Informal description
For any natural numbers $n$ and $m$, the sum of $n$ and the successor of $m$ equals the successor of the sum of $n$ and $m$, i.e., $n + (m + 1) = (n + m) + 1$.
Nat.add_one theorem
(n : Nat) : n + 1 = succ n
Full source
theorem add_one (n : Nat) : n + 1 = succ n :=
  rfl
Successor as Addition of One in Natural Numbers
Informal description
For any natural number $n$, the sum of $n$ and $1$ is equal to the successor of $n$, i.e., $n + 1 = \text{succ}(n)$.
Nat.succ_eq_add_one theorem
(n : Nat) : succ n = n + 1
Full source
@[simp] theorem succ_eq_add_one (n : Nat) : succ n = n + 1 :=
  rfl
Successor as Addition by One
Informal description
For any natural number $n$, the successor of $n$ is equal to $n + 1$.
Nat.add_one_ne_zero theorem
(n : Nat) : n + 1 ≠ 0
Full source
@[simp] theorem add_one_ne_zero (n : Nat) : n + 1 ≠ 0 := nofun
Nonzero Successor: $n + 1 \neq 0$ for all $n \in \mathbb{N}$
Informal description
For any natural number $n$, the sum $n + 1$ is not equal to zero.
Nat.zero_ne_add_one theorem
(n : Nat) : 0 ≠ n + 1
Full source
theorem zero_ne_add_one (n : Nat) : 0 ≠ n + 1 := by simp
Zero Not Equal to Successor of Any Natural Number
Informal description
For any natural number $n$, the number $0$ is not equal to $n + 1$.
Nat.add_comm theorem
: ∀ (n m : Nat), n + m = m + n
Full source
protected theorem add_comm : ∀ (n m : Nat), n + m = m + n
  | n, 0   => Eq.symm (Nat.zero_add n)
  | n, m+1 => by
    have : succ (n + m) = succ (m + n) := by apply congrArg; apply Nat.add_comm
    rw [succ_add m n]
    apply this
Commutativity of Addition on Natural Numbers: $n + m = m + n$
Informal description
For any natural numbers $n$ and $m$, the sum $n + m$ is equal to $m + n$.
Nat.instCommutativeHAdd instance
: Std.Commutative (α := Nat) (· + ·)
Full source
instance : Std.Commutative (α := Nat) (· + ·) := ⟨Nat.add_comm⟩
Commutativity of Natural Number Addition
Informal description
The addition operation $+$ on natural numbers is commutative.
Nat.add_assoc theorem
: ∀ (n m k : Nat), (n + m) + k = n + (m + k)
Full source
protected theorem add_assoc : ∀ (n m k : Nat), (n + m) + k = n + (m + k)
  | _, _, 0      => rfl
  | n, m, succ k => congrArg succ (Nat.add_assoc n m k)
Associativity of Addition on Natural Numbers
Informal description
For any natural numbers $n$, $m$, and $k$, the addition operation satisfies the associative property: $(n + m) + k = n + (m + k)$.
Nat.instAssociativeHAdd instance
: Std.Associative (α := Nat) (· + ·)
Full source
instance : Std.Associative (α := Nat) (· + ·) := ⟨Nat.add_assoc⟩
Associativity of Natural Number Addition
Informal description
The addition operation $+$ on natural numbers is associative.
Nat.add_left_comm theorem
(n m k : Nat) : n + (m + k) = m + (n + k)
Full source
protected theorem add_left_comm (n m k : Nat) : n + (m + k) = m + (n + k) := by
  rw [← Nat.add_assoc, Nat.add_comm n m, Nat.add_assoc]
Left Commutativity of Addition in Natural Numbers: $n + (m + k) = m + (n + k)$
Informal description
For any natural numbers $n$, $m$, and $k$, the addition operation satisfies the left commutative property: $n + (m + k) = m + (n + k)$.
Nat.add_right_comm theorem
(n m k : Nat) : (n + m) + k = (n + k) + m
Full source
protected theorem add_right_comm (n m k : Nat) : (n + m) + k = (n + k) + m := by
  rw [Nat.add_assoc, Nat.add_comm m k, ← Nat.add_assoc]
Right Commutativity of Addition in Natural Numbers: $(n + m) + k = (n + k) + m$
Informal description
For any natural numbers $n$, $m$, and $k$, the following equality holds: $(n + m) + k = (n + k) + m$.
Nat.add_left_cancel theorem
{n m k : Nat} : n + m = n + k → m = k
Full source
protected theorem add_left_cancel {n m k : Nat} : n + m = n + k → m = k := by
  induction n with
  | zero => simp
  | succ n ih => simp [succ_add, succ.injEq]; intro h; apply ih h
Left Cancellation Property of Addition on Natural Numbers
Informal description
For any natural numbers $n, m, k$, if $n + m = n + k$, then $m = k$.
Nat.add_right_cancel theorem
{n m k : Nat} (h : n + m = k + m) : n = k
Full source
protected theorem add_right_cancel {n m k : Nat} (h : n + m = k + m) : n = k := by
  rw [Nat.add_comm n m, Nat.add_comm k m] at h
  apply Nat.add_left_cancel h
Right Cancellation Property of Addition on Natural Numbers
Informal description
For any natural numbers $n, m, k$, if $n + m = k + m$, then $n = k$.
Nat.mul_zero theorem
(n : Nat) : n * 0 = 0
Full source
@[simp] protected theorem mul_zero (n : Nat) : n * 0 = 0 :=
  rfl
Multiplication by Zero in Natural Numbers: $n \times 0 = 0$
Informal description
For any natural number $n$, the product of $n$ and zero is zero, i.e., $n \times 0 = 0$.
Nat.mul_succ theorem
(n m : Nat) : n * succ m = n * m + n
Full source
theorem mul_succ (n m : Nat) : n * succ m = n * m + n :=
  rfl
Recursive Definition of Multiplication by Successor in Natural Numbers
Informal description
For any natural numbers $n$ and $m$, the product of $n$ and the successor of $m$ is equal to the sum of the product of $n$ and $m$ with $n$, i.e., $n \cdot (m + 1) = n \cdot m + n$.
Nat.mul_add_one theorem
(n m : Nat) : n * (m + 1) = n * m + n
Full source
theorem mul_add_one (n m : Nat) : n * (m + 1) = n * m + n :=
  rfl
Distributive Property of Multiplication over Successor in Natural Numbers
Informal description
For any natural numbers $n$ and $m$, the product of $n$ and the successor of $m$ is equal to the sum of the product $n \times m$ and $n$, i.e., $n \times (m + 1) = n \times m + n$.
Nat.zero_mul theorem
: ∀ (n : Nat), 0 * n = 0
Full source
@[simp] protected theorem zero_mul : ∀ (n : Nat), 0 * n = 0
  | 0      => rfl
  | succ n => mul_succ 0 n ▸ (Nat.zero_mul n).symm ▸ rfl
Zero Multiplication Property in Natural Numbers
Informal description
For any natural number $n$, the product of $0$ and $n$ is equal to $0$, i.e., $0 \cdot n = 0$.
Nat.succ_mul theorem
(n m : Nat) : (succ n) * m = (n * m) + m
Full source
theorem succ_mul (n m : Nat) : (succ n) * m = (n * m) + m := by
  induction m with
  | zero => rfl
  | succ m ih => rw [mul_succ, add_succ, ih, mul_succ, add_succ, Nat.add_right_comm]
Successor Multiplication Identity: $(n + 1) \cdot m = n \cdot m + m$
Informal description
For any natural numbers $n$ and $m$, the product of the successor of $n$ and $m$ is equal to the sum of the product of $n$ and $m$ with $m$, i.e., $(n + 1) \cdot m = n \cdot m + m$.
Nat.add_one_mul theorem
(n m : Nat) : (n + 1) * m = (n * m) + m
Full source
theorem add_one_mul (n m : Nat) : (n + 1) * m = (n * m) + m := succ_mul n m
Multiplication by Successor Identity: $(n + 1) \cdot m = n \cdot m + m$
Informal description
For any natural numbers $n$ and $m$, the product of $n + 1$ and $m$ is equal to the sum of the product of $n$ and $m$ with $m$, i.e., $(n + 1) \cdot m = n \cdot m + m$.
Nat.mul_comm theorem
: ∀ (n m : Nat), n * m = m * n
Full source
protected theorem mul_comm : ∀ (n m : Nat), n * m = m * n
  | n, 0      => (Nat.zero_mul n).symm ▸ (Nat.mul_zero n).symm ▸ rfl
  | n, succ m => (mul_succ n m).symm ▸ (succ_mul m n).symm ▸ (Nat.mul_comm n m).symm ▸ rfl
Commutativity of Multiplication in Natural Numbers: $n \times m = m \times n$
Informal description
For any natural numbers $n$ and $m$, the product $n \times m$ is equal to $m \times n$.
Nat.instCommutativeHMul instance
: Std.Commutative (α := Nat) (· * ·)
Full source
instance : Std.Commutative (α := Nat) (· * ·) := ⟨Nat.mul_comm⟩
Commutativity of Natural Number Multiplication
Informal description
The multiplication operation on natural numbers is commutative. That is, for any natural numbers $n$ and $m$, we have $n \times m = m \times n$.
Nat.mul_one theorem
: ∀ (n : Nat), n * 1 = n
Full source
@[simp] protected theorem mul_one : ∀ (n : Nat), n * 1 = n :=
  Nat.zero_add
Multiplicative Identity for Natural Numbers: $n \times 1 = n$
Informal description
For any natural number $n$, the product of $n$ and $1$ equals $n$, i.e., $n \times 1 = n$.
Nat.one_mul theorem
(n : Nat) : 1 * n = n
Full source
@[simp] protected theorem one_mul (n : Nat) : 1 * n = n :=
  Nat.mul_comm n 1 ▸ Nat.mul_one n
Left Multiplicative Identity for Natural Numbers: $1 \times n = n$
Informal description
For any natural number $n$, the product of $1$ and $n$ equals $n$, i.e., $1 \times n = n$.
Nat.instLawfulIdentityHMulOfNat instance
: Std.LawfulIdentity (α := Nat) (· * ·) 1
Full source
instance : Std.LawfulIdentity (α := Nat) (· * ·) 1 where
  left_id := Nat.one_mul
  right_id := Nat.mul_one
$1$ is the Multiplicative Identity for Natural Numbers
Informal description
The natural number $1$ is a verified left and right identity element for multiplication on $\mathbb{N}$. That is, for any natural number $n$, we have $1 \cdot n = n$ and $n \cdot 1 = n$.
Nat.left_distrib theorem
(n m k : Nat) : n * (m + k) = n * m + n * k
Full source
protected theorem left_distrib (n m k : Nat) : n * (m + k) = n * m + n * k := by
  induction n with
  | zero      => repeat rw [Nat.zero_mul]
  | succ n ih => simp [succ_mul, ih]; rw [Nat.add_assoc, Nat.add_assoc (n*m)]; apply congrArg; apply Nat.add_left_comm
Left Distributivity of Multiplication over Addition in Natural Numbers: $n \cdot (m + k) = n \cdot m + n \cdot k$
Informal description
For any natural numbers $n$, $m$, and $k$, the multiplication operation distributes over addition on the left: $n \cdot (m + k) = n \cdot m + n \cdot k$.
Nat.right_distrib theorem
(n m k : Nat) : (n + m) * k = n * k + m * k
Full source
protected theorem right_distrib (n m k : Nat) : (n + m) * k = n * k + m * k := by
  rw [Nat.mul_comm, Nat.left_distrib]; simp [Nat.mul_comm]
Right Distributivity of Multiplication over Addition in Natural Numbers: $(n + m) \cdot k = n \cdot k + m \cdot k$
Informal description
For any natural numbers $n$, $m$, and $k$, the multiplication operation distributes over addition on the right: $(n + m) \cdot k = n \cdot k + m \cdot k$.
Nat.mul_add theorem
(n m k : Nat) : n * (m + k) = n * m + n * k
Full source
protected theorem mul_add (n m k : Nat) : n * (m + k) = n * m + n * k :=
  Nat.left_distrib n m k
Left Distributivity of Multiplication over Addition in Natural Numbers: $n \cdot (m + k) = n \cdot m + n \cdot k$
Informal description
For any natural numbers $n$, $m$, and $k$, the product of $n$ and the sum of $m$ and $k$ is equal to the sum of the products $n \cdot m$ and $n \cdot k$, i.e., $n \cdot (m + k) = n \cdot m + n \cdot k$.
Nat.add_mul theorem
(n m k : Nat) : (n + m) * k = n * k + m * k
Full source
protected theorem add_mul (n m k : Nat) : (n + m) * k = n * k + m * k :=
  Nat.right_distrib n m k
Right Distributivity of Multiplication over Addition in Natural Numbers: $(n + m) \cdot k = n \cdot k + m \cdot k$
Informal description
For any natural numbers $n$, $m$, and $k$, the product of the sum $n + m$ and $k$ is equal to the sum of the products $n \cdot k$ and $m \cdot k$, i.e., $(n + m) \cdot k = n \cdot k + m \cdot k$.
Nat.mul_assoc theorem
: ∀ (n m k : Nat), (n * m) * k = n * (m * k)
Full source
protected theorem mul_assoc : ∀ (n m k : Nat), (n * m) * k = n * (m * k)
  | _, _, 0      => rfl
  | n, m, succ k => by simp [mul_succ, Nat.mul_assoc n m k, Nat.left_distrib]
Associativity of Multiplication in Natural Numbers: $(n \cdot m) \cdot k = n \cdot (m \cdot k)$
Informal description
For any natural numbers $n$, $m$, and $k$, the multiplication operation is associative, i.e., $(n \cdot m) \cdot k = n \cdot (m \cdot k)$.
Nat.instAssociativeHMul instance
: Std.Associative (α := Nat) (· * ·)
Full source
instance : Std.Associative (α := Nat) (· * ·) := ⟨Nat.mul_assoc⟩
Associativity of Natural Number Multiplication
Informal description
The multiplication operation on natural numbers is associative, meaning that for any natural numbers $n$, $m$, and $k$, we have $(n \cdot m) \cdot k = n \cdot (m \cdot k)$.
Nat.mul_left_comm theorem
(n m k : Nat) : n * (m * k) = m * (n * k)
Full source
protected theorem mul_left_comm (n m k : Nat) : n * (m * k) = m * (n * k) := by
  rw [← Nat.mul_assoc, Nat.mul_comm n m, Nat.mul_assoc]
Left Commutativity of Multiplication in Natural Numbers: $n \cdot (m \cdot k) = m \cdot (n \cdot k)$
Informal description
For any natural numbers $n$, $m$, and $k$, the following equality holds: $$n \cdot (m \cdot k) = m \cdot (n \cdot k).$$
Nat.mul_two theorem
(n) : n * 2 = n + n
Full source
protected theorem mul_two (n) : n * 2 = n + n := by rw [Nat.mul_succ, Nat.mul_one]
Multiplication by Two Equals Self-Addition in Natural Numbers
Informal description
For any natural number $n$, the product of $n$ and $2$ equals the sum of $n$ with itself, i.e., $n \times 2 = n + n$.
Nat.two_mul theorem
(n) : 2 * n = n + n
Full source
protected theorem two_mul (n) : 2 * n = n + n := by rw [Nat.succ_mul, Nat.one_mul]
Double of a Natural Number Equals Its Sum with Itself
Informal description
For any natural number $n$, twice $n$ is equal to the sum of $n$ with itself, i.e., $2 \times n = n + n$.
Nat.succ_lt_succ theorem
{n m : Nat} : n < m → succ n < succ m
Full source
theorem succ_lt_succ {n m : Nat} : n < m → succ n < succ m := succ_le_succ
Successor Preserves Strict Order on Natural Numbers
Informal description
For any natural numbers $n$ and $m$, if $n < m$, then $\text{succ}(n) < \text{succ}(m)$.
Nat.lt_succ_of_le theorem
{n m : Nat} : n ≤ m → n < succ m
Full source
theorem lt_succ_of_le {n m : Nat} : n ≤ m → n < succ m := succ_le_succ
Strict Inequality from Non-Strict with Successor
Informal description
For any natural numbers $n$ and $m$, if $n \leq m$, then $n < m + 1$.
Nat.le_of_lt_add_one theorem
{n m : Nat} : n < m + 1 → n ≤ m
Full source
theorem le_of_lt_add_one {n m : Nat} : n < m + 1 → n ≤ m := le_of_succ_le_succ
Strict Inequality with Successor Implies Non-Strict Inequality
Informal description
For any natural numbers $n$ and $m$, if $n < m + 1$, then $n \leq m$.
Nat.lt_add_one_of_le theorem
{n m : Nat} : n ≤ m → n < m + 1
Full source
theorem lt_add_one_of_le {n m : Nat} : n ≤ m → n < m + 1 := succ_le_succ
Strict Inequality from Non-Strict with Successor: $n \leq m \Rightarrow n < m + 1$
Informal description
For any natural numbers $n$ and $m$, if $n \leq m$, then $n < m + 1$.
Nat.sub_zero theorem
(n : Nat) : n - 0 = n
Full source
@[simp] protected theorem sub_zero (n : Nat) : n - 0 = n := rfl
Subtraction of Zero from Natural Number is Identity
Informal description
For any natural number $n$, subtracting zero from $n$ leaves $n$ unchanged, i.e., $n - 0 = n$.
Nat.not_add_one_le_zero theorem
(n : Nat) : ¬n + 1 ≤ 0
Full source
theorem not_add_one_le_zero (n : Nat) : ¬ n + 1 ≤ 0 := nofun
Non-negativity of successor: $n + 1 \not\leq 0$
Informal description
For any natural number $n$, it is not true that $n + 1 \leq 0$.
Nat.not_add_one_le_self theorem
: (n : Nat) → ¬n + 1 ≤ n
Full source
theorem not_add_one_le_self : (n : Nat) → ¬ n + 1 ≤ n := Nat.not_succ_le_self
Non-reflexivity of successor: $n + 1 \not\leq n$
Informal description
For any natural number $n$, it is not true that $n + 1 \leq n$.
Nat.add_one_pos theorem
(n : Nat) : 0 < n + 1
Full source
theorem add_one_pos (n : Nat) : 0 < n + 1 := Nat.zero_lt_succ n
Positivity of Successor: $0 < n + 1$
Informal description
For any natural number $n$, the successor of $n$ is positive, i.e., $0 < n + 1$.
Nat.succ_sub_succ_eq_sub theorem
(n m : Nat) : succ n - succ m = n - m
Full source
theorem succ_sub_succ_eq_sub (n m : Nat) : succ n - succ m = n - m := by
  induction m with
  | zero      => exact rfl
  | succ m ih => apply congrArg pred ih
Successor Subtraction Identity: $(n+1)-(m+1) = n-m$
Informal description
For any natural numbers $n$ and $m$, the difference between the successors of $n$ and $m$ equals the difference between $n$ and $m$, i.e., $(n + 1) - (m + 1) = n - m$.
Nat.pred_le theorem
: ∀ (n : Nat), pred n ≤ n
Full source
theorem pred_le : ∀ (n : Nat), pred n ≤ n
  | zero   => Nat.le.refl
  | succ _ => le_succ _
Predecessor is Less Than or Equal to Original Number
Informal description
For any natural number $n$, the predecessor of $n$ is less than or equal to $n$, i.e., $\mathrm{pred}(n) \leq n$.
Nat.pred_lt theorem
: ∀ {n : Nat}, n ≠ 0 → pred n < n
Full source
theorem pred_lt : ∀ {n : Nat}, n ≠ 0pred n < n
  | zero,   h => absurd rfl h
  | succ _, _ => lt_succ_of_le (Nat.le_refl _)
Predecessor is Strictly Less Than Nonzero Natural Number
Informal description
For any natural number $n$ such that $n \neq 0$, the predecessor of $n$ is strictly less than $n$, i.e., $\mathrm{pred}(n) < n$.
Nat.sub_one_lt theorem
: ∀ {n : Nat}, n ≠ 0 → n - 1 < n
Full source
theorem sub_one_lt : ∀ {n : Nat}, n ≠ 0 → n - 1 < n := pred_lt
Subtraction of One from Nonzero Natural Number is Strictly Less Than Original Number
Informal description
For any nonzero natural number $n$, we have $n - 1 < n$.
Nat.sub_le theorem
(n m : Nat) : n - m ≤ n
Full source
@[simp] theorem sub_le (n m : Nat) : n - m ≤ n := by
  induction m with
  | zero      => exact Nat.le_refl (n - 0)
  | succ m ih => apply Nat.le_trans (pred_le (n - m)) ih
Truncated Subtraction is Bounded Above by the Minuend
Informal description
For any natural numbers $n$ and $m$, the truncated subtraction $n - m$ is less than or equal to $n$.
Nat.sub_lt_of_lt theorem
{a b c : Nat} (h : a < c) : a - b < c
Full source
theorem sub_lt_of_lt {a b c : Nat} (h : a < c) : a - b < c :=
  Nat.lt_of_le_of_lt (Nat.sub_le _ _) h
Subtraction Preserves Strict Inequality in Natural Numbers
Informal description
For any natural numbers $a$, $b$, and $c$, if $a < c$, then $a - b < c$.
Nat.sub_lt theorem
: ∀ {n m : Nat}, 0 < n → 0 < m → n - m < n
Full source
theorem sub_lt : ∀ {n m : Nat}, 0 < n → 0 < m → n - m < n
  | 0,   _,   h1, _  => absurd h1 (Nat.lt_irrefl 0)
  | _+1, 0,   _, h2  => absurd h2 (Nat.lt_irrefl 0)
  | n+1, m+1, _,  _  =>
    Eq.symm (succ_sub_succ_eq_sub n m) ▸
      show n - m < succ n from
      lt_succ_of_le (sub_le n m)
Truncated Subtraction Decreases Positive Natural Numbers
Informal description
For any natural numbers $n$ and $m$ such that $0 < n$ and $0 < m$, the truncated subtraction $n - m$ is strictly less than $n$.
Nat.sub_succ theorem
(n m : Nat) : n - succ m = pred (n - m)
Full source
theorem sub_succ (n m : Nat) : n - succ m = pred (n - m) := rfl
Subtraction of Successor Equals Predecessor of Difference
Informal description
For any natural numbers $n$ and $m$, the difference $n - (m + 1)$ is equal to the predecessor of $n - m$, i.e., $n - (m + 1) = \mathrm{pred}(n - m)$.
Nat.succ_sub_succ theorem
(n m : Nat) : succ n - succ m = n - m
Full source
theorem succ_sub_succ (n m : Nat) : succ n - succ m = n - m :=
  succ_sub_succ_eq_sub n m
Successor Subtraction Identity: $(n+1)-(m+1) = n-m$
Informal description
For any natural numbers $n$ and $m$, the difference between the successors of $n$ and $m$ equals the difference between $n$ and $m$, i.e., $(n + 1) - (m + 1) = n - m$.
Nat.sub_self theorem
: ∀ (n : Nat), n - n = 0
Full source
@[simp] protected theorem sub_self : ∀ (n : Nat), n - n = 0
  | 0        => by rw [Nat.sub_zero]
  | (succ n) => by rw [succ_sub_succ, Nat.sub_self n]
Subtraction of a Natural Number from Itself Yields Zero
Informal description
For any natural number $n$, the difference $n - n$ equals zero, i.e., $n - n = 0$.
Nat.sub_add_eq theorem
(a b c : Nat) : a - (b + c) = a - b - c
Full source
theorem sub_add_eq (a b c : Nat) : a - (b + c) = a - b - c := by
  induction c with
  | zero => simp
  | succ c ih => simp only [Nat.add_succ, Nat.sub_succ, ih]
Subtraction Distributes Over Addition: $a - (b + c) = (a - b) - c$
Informal description
For any natural numbers $a$, $b$, and $c$, the difference $a - (b + c)$ is equal to $(a - b) - c$.
Nat.lt_of_lt_of_le theorem
{n m k : Nat} : n < m → m ≤ k → n < k
Full source
protected theorem lt_of_lt_of_le {n m k : Nat} : n < m → m ≤ k → n < k :=
  Nat.le_trans
Transitivity of Strict and Non-Strict Inequalities for Natural Numbers
Informal description
For any natural numbers $n$, $m$, and $k$, if $n < m$ and $m \leq k$, then $n < k$.
Nat.lt_of_lt_of_eq theorem
{n m k : Nat} : n < m → m = k → n < k
Full source
protected theorem lt_of_lt_of_eq {n m k : Nat} : n < m → m = k → n < k :=
  fun h₁ h₂ => h₂ ▸ h₁
Transitivity of Strict Inequality with Equality for Natural Numbers
Informal description
For any natural numbers $n$, $m$, and $k$, if $n < m$ and $m = k$, then $n < k$.
Nat.instTransLt instance
: Trans (. < . : Nat → Nat → Prop) (. < . : Nat → Nat → Prop) (. < . : Nat → Nat → Prop)
Full source
instance : Trans (. < . : NatNat → Prop) (. < . : NatNat → Prop) (. < . : NatNat → Prop) where
  trans := Nat.lt_trans
Transitivity of the Strict Order on Natural Numbers
Informal description
The strict order relation $<$ on natural numbers is transitive. That is, for any natural numbers $n$, $m$, and $k$, if $n < m$ and $m < k$, then $n < k$.
Nat.instTransLe instance
: Trans (. ≤ . : Nat → Nat → Prop) (. ≤ . : Nat → Nat → Prop) (. ≤ . : Nat → Nat → Prop)
Full source
instance : Trans (. ≤ . : NatNat → Prop) (. ≤ . : NatNat → Prop) (. ≤ . : NatNat → Prop) where
  trans := Nat.le_trans
Transitivity of the Natural Number Order $\leq$
Informal description
The relation $\leq$ on natural numbers is transitive. That is, for any natural numbers $n$, $m$, and $k$, if $n \leq m$ and $m \leq k$, then $n \leq k$.
Nat.instTransLtLe instance
: Trans (. < . : Nat → Nat → Prop) (. ≤ . : Nat → Nat → Prop) (. < . : Nat → Nat → Prop)
Full source
instance : Trans (. < . : NatNat → Prop) (. ≤ . : NatNat → Prop) (. < . : NatNat → Prop) where
  trans := Nat.lt_of_lt_of_le
Transitivity of $<$ and $\leq$ on Natural Numbers
Informal description
The relation $<$ on natural numbers is transitive with respect to $\leq$. That is, for any natural numbers $n$, $m$, and $k$, if $n < m$ and $m \leq k$, then $n < k$.
Nat.instTransLeLt instance
: Trans (. ≤ . : Nat → Nat → Prop) (. < . : Nat → Nat → Prop) (. < . : Nat → Nat → Prop)
Full source
instance : Trans (. ≤ . : NatNat → Prop) (. < . : NatNat → Prop) (. < . : NatNat → Prop) where
  trans := Nat.lt_of_le_of_lt
Transitivity of $\leq$ and $<$ on Natural Numbers
Informal description
The relation $\leq$ on natural numbers is transitive with respect to $<$. That is, for any natural numbers $n$, $m$, and $k$, if $n \leq m$ and $m < k$, then $n < k$.
Nat.le_of_eq theorem
{n m : Nat} (p : n = m) : n ≤ m
Full source
protected theorem le_of_eq {n m : Nat} (p : n = m) : n ≤ m :=
  p ▸ Nat.le_refl n
Equality Implies Inequality in Natural Numbers
Informal description
For any natural numbers $n$ and $m$, if $n = m$, then $n \leq m$.
Nat.lt.step theorem
{n m : Nat} : n < m → n < succ m
Full source
theorem lt.step {n m : Nat} : n < m → n < succ m := le_step
Strict Inequality Preservation Under Successor
Informal description
For any natural numbers $n$ and $m$, if $n < m$, then $n < m + 1$.
Nat.le_of_succ_le theorem
{n m : Nat} (h : succ n ≤ m) : n ≤ m
Full source
theorem le_of_succ_le {n m : Nat} (h : succ n ≤ m) : n ≤ m := Nat.le_trans (le_succ n) h
Successor Inequality Implies Original Inequality in Natural Numbers
Informal description
For any natural numbers $n$ and $m$, if $n + 1 \leq m$, then $n \leq m$.
Nat.lt_of_succ_lt theorem
{n m : Nat} : succ n < m → n < m
Full source
theorem lt_of_succ_lt      {n m : Nat} : succ n < m → n < m := le_of_succ_le
Successor Strict Inequality Implies Original Strict Inequality in Natural Numbers
Informal description
For any natural numbers $n$ and $m$, if $n + 1 < m$, then $n < m$.
Nat.le_of_lt theorem
{n m : Nat} : n < m → n ≤ m
Full source
protected theorem le_of_lt {n m : Nat} : n < m → n ≤ m := le_of_succ_le
Strict Inequality Implies Non-Strict Inequality in Natural Numbers
Informal description
For any natural numbers $n$ and $m$, if $n < m$, then $n \leq m$.
Nat.lt_of_succ_lt_succ theorem
{n m : Nat} : succ n < succ m → n < m
Full source
theorem lt_of_succ_lt_succ {n m : Nat} : succ n < succ m → n < m := le_of_succ_le_succ
Successor Preserves Strict Inequality in Natural Numbers
Informal description
For any natural numbers $n$ and $m$, if the successor of $n$ is strictly less than the successor of $m$, then $n$ is strictly less than $m$.
Nat.lt_of_succ_le theorem
{n m : Nat} (h : succ n ≤ m) : n < m
Full source
theorem lt_of_succ_le {n m : Nat} (h : succ n ≤ m) : n < m := h
Successor Inequality Implies Strict Inequality in Natural Numbers
Informal description
For any natural numbers $n$ and $m$, if $n + 1 \leq m$, then $n < m$.
Nat.succ_le_of_lt theorem
{n m : Nat} (h : n < m) : succ n ≤ m
Full source
theorem succ_le_of_lt {n m : Nat} (h : n < m) : succ n ≤ m := h
Successor preserves inequality to non-strict inequality
Informal description
For any natural numbers $n$ and $m$, if $n < m$ then $n + 1 \leq m$.
Nat.pos_of_ne_zero theorem
{n : Nat} : n ≠ 0 → 0 < n
Full source
protected theorem pos_of_ne_zero {n : Nat} : n ≠ 0 → 0 < n := (eq_zero_or_pos n).resolve_left
Nonzero Natural Numbers are Positive
Informal description
For any natural number $n$, if $n$ is not equal to zero, then $0 < n$.
Nat.pos_of_neZero theorem
(n : Nat) [NeZero n] : 0 < n
Full source
theorem pos_of_neZero (n : Nat) [NeZero n] : 0 < n := Nat.pos_of_ne_zero (NeZero.ne _)
Non-zero natural numbers are positive
Informal description
For any natural number $n$ that is non-zero (i.e., $n \neq 0$), we have $0 < n$.
Nat.lt.base theorem
(n : Nat) : n < succ n
Full source
theorem lt.base (n : Nat) : n < succ n := Nat.le_refl (succ n)
Successor is greater than original number
Informal description
For any natural number $n$, we have $n < n + 1$.
Nat.lt_succ_self theorem
(n : Nat) : n < succ n
Full source
theorem lt_succ_self (n : Nat) : n < succ n := lt.base n
Successor is strictly greater than original number
Informal description
For any natural number $n$, we have $n < n + 1$.
Nat.lt_add_one theorem
(n : Nat) : n < n + 1
Full source
@[simp] protected theorem lt_add_one (n : Nat) : n < n + 1 := lt.base n
Natural number is less than its successor
Informal description
For any natural number $n$, we have $n < n + 1$.
Nat.zero_lt_of_lt theorem
: {a b : Nat} → a < b → 0 < b
Full source
theorem zero_lt_of_lt : {a b : Nat} → a < b → 0 < b
  | 0,   _, h => h
  | a+1, b, h =>
    have : a < b := Nat.lt_trans (Nat.lt_succ_self _) h
    zero_lt_of_lt this
Nonzero Lower Bound from Strict Inequality in Natural Numbers
Informal description
For any natural numbers $a$ and $b$, if $a < b$, then $0 < b$.
Nat.zero_lt_of_ne_zero theorem
{a : Nat} (h : a ≠ 0) : 0 < a
Full source
theorem zero_lt_of_ne_zero {a : Nat} (h : a ≠ 0) : 0 < a := by
  match a with
  | 0 => contradiction
  | a+1 => apply Nat.zero_lt_succ
Nonzero Natural Numbers are Positive: $a \neq 0 \Rightarrow 0 < a$
Informal description
For any natural number $a$, if $a$ is not equal to $0$, then $0 < a$.
Nat.ne_of_lt theorem
{a b : Nat} (h : a < b) : a ≠ b
Full source
theorem ne_of_lt {a b : Nat} (h : a < b) : a ≠ b := fun he => absurd (he ▸ h) (Nat.lt_irrefl a)
Strict Inequality Implies Inequality in Natural Numbers
Informal description
For any natural numbers $a$ and $b$, if $a < b$, then $a \neq b$.
Nat.le_or_eq_of_le_succ theorem
{m n : Nat} (h : m ≤ succ n) : m ≤ n ∨ m = succ n
Full source
theorem le_or_eq_of_le_succ {m n : Nat} (h : m ≤ succ n) : m ≤ n ∨ m = succ n :=
  Decidable.byCases
    (fun (h' : m = succ n) => Or.inr h')
    (fun (h' : m ≠ succ n) =>
       have : m < succ n := Nat.lt_of_le_of_ne h h'
       have : succ m ≤ succ n := succ_le_of_lt this
       Or.inl (le_of_succ_le_succ this))
Decomposition of Successor Inequality in Natural Numbers: $m \leq n + 1 \Rightarrow m \leq n \lor m = n + 1$
Informal description
For any natural numbers $m$ and $n$, if $m \leq n + 1$, then either $m \leq n$ or $m = n + 1$.
Nat.le_or_eq_of_le_add_one theorem
{m n : Nat} (h : m ≤ n + 1) : m ≤ n ∨ m = n + 1
Full source
theorem le_or_eq_of_le_add_one {m n : Nat} (h : m ≤ n + 1) : m ≤ n ∨ m = n + 1 :=
  le_or_eq_of_le_succ h
Decomposition of Inequality with Successor: $m \leq n + 1 \Rightarrow m \leq n \lor m = n + 1$
Informal description
For any natural numbers $m$ and $n$, if $m \leq n + 1$, then either $m \leq n$ or $m = n + 1$.
Nat.le_add_right theorem
: ∀ (n k : Nat), n ≤ n + k
Full source
@[simp] theorem le_add_right : ∀ (n k : Nat), n ≤ n + k
  | n, 0   => Nat.le_refl n
  | n, k+1 => le_succ_of_le (le_add_right n k)
Right Addition Preserves Natural Order
Informal description
For any natural numbers $n$ and $k$, the inequality $n \leq n + k$ holds.
Nat.le_add_left theorem
(n m : Nat) : n ≤ m + n
Full source
@[simp] theorem le_add_left (n m : Nat): n ≤ m + n :=
  Nat.add_comm n m ▸ le_add_right n m
Left Addition Preserves Natural Order
Informal description
For any natural numbers $n$ and $m$, the inequality $n \leq m + n$ holds.
Nat.le_of_add_right_le theorem
{n m k : Nat} (h : n + k ≤ m) : n ≤ m
Full source
theorem le_of_add_right_le {n m k : Nat} (h : n + k ≤ m) : n ≤ m :=
  Nat.le_trans (le_add_right n k) h
Left Term in Sum Inequality Implies Individual Term Inequality
Informal description
For any natural numbers $n$, $m$, and $k$, if $n + k \leq m$, then $n \leq m$.
Nat.le_add_right_of_le theorem
{n m k : Nat} (h : n ≤ m) : n ≤ m + k
Full source
theorem le_add_right_of_le {n m k : Nat} (h : n ≤ m) : n ≤ m + k :=
  Nat.le_trans h (le_add_right m k)
Right Addition Preserves Order in Natural Numbers
Informal description
For any natural numbers $n$, $m$, and $k$, if $n \leq m$, then $n \leq m + k$.
Nat.lt_of_add_one_le theorem
{n m : Nat} (h : n + 1 ≤ m) : n < m
Full source
theorem lt_of_add_one_le {n m : Nat} (h : n + 1 ≤ m) : n < m := h
Inequality from Successor Comparison: $n + 1 \leq m \Rightarrow n < m$
Informal description
For any natural numbers $n$ and $m$, if $n + 1 \leq m$, then $n < m$.
Nat.add_one_le_of_lt theorem
{n m : Nat} (h : n < m) : n + 1 ≤ m
Full source
theorem add_one_le_of_lt {n m : Nat} (h : n < m) : n + 1 ≤ m := h
Strict Inequality Implies Successor Inequality in Natural Numbers
Informal description
For any natural numbers $n$ and $m$, if $n < m$, then $n + 1 \leq m$.
Nat.lt_add_left theorem
(c : Nat) (h : a < b) : a < c + b
Full source
protected theorem lt_add_left (c : Nat) (h : a < b) : a < c + b :=
  Nat.lt_of_lt_of_le h (Nat.le_add_left ..)
Left Addition Preserves Strict Inequality in Natural Numbers
Informal description
For any natural numbers $a$, $b$, and $c$, if $a < b$, then $a < c + b$.
Nat.lt_add_right theorem
(c : Nat) (h : a < b) : a < b + c
Full source
protected theorem lt_add_right (c : Nat) (h : a < b) : a < b + c :=
  Nat.lt_of_lt_of_le h (Nat.le_add_right ..)
Right Addition Preserves Strict Inequality in Natural Numbers
Informal description
For any natural numbers $a$, $b$, and $c$, if $a < b$, then $a < b + c$.
Nat.lt_of_add_right_lt theorem
{n m k : Nat} (h : n + k < m) : n < m
Full source
theorem lt_of_add_right_lt {n m k : Nat} (h : n + k < m) : n < m :=
  Nat.lt_of_le_of_lt (Nat.le_add_right ..) h
Strict Inequality from Right Addition Inequality in Natural Numbers
Informal description
For any natural numbers $n$, $m$, and $k$, if $n + k < m$, then $n < m$.
Nat.lt_of_add_left_lt theorem
{n m k : Nat} (h : k + n < m) : n < m
Full source
theorem lt_of_add_left_lt {n m k : Nat} (h : k + n < m) : n < m :=
  lt_of_add_right_lt (Nat.add_comm _ _ ▸ h)
Strict Inequality from Left Addition Inequality in Natural Numbers
Informal description
For any natural numbers $n$, $m$, and $k$, if $k + n < m$, then $n < m$.
Nat.le.dest theorem
: ∀ {n m : Nat}, n ≤ m → Exists (fun k => n + k = m)
Full source
theorem le.dest : ∀ {n m : Nat}, n ≤ m → Exists (fun k => n + k = m)
  | zero,   zero,   _ => ⟨0, rfl⟩
  | zero,   succ n, _ => ⟨succ n, Nat.add_comm 0 (succ n) ▸ rfl⟩
  | succ _, zero,   h => absurd h (not_succ_le_zero _)
  | succ n, succ m, h =>
    have : n ≤ m := Nat.le_of_succ_le_succ h
    have : Exists (fun k => n + k = m) := dest this
    match this with
    | ⟨k, h⟩ => ⟨k, show succ n + k = succ m from ((succ_add n k).symm ▸ h ▸ rfl)⟩
Existence of Difference for Natural Number Order: $n \leq m \implies \exists k, n + k = m$
Informal description
For any natural numbers $n$ and $m$, if $n \leq m$, then there exists a natural number $k$ such that $n + k = m$.
Nat.le.intro theorem
{n m k : Nat} (h : n + k = m) : n ≤ m
Full source
theorem le.intro {n m k : Nat} (h : n + k = m) : n ≤ m :=
  h ▸ le_add_right n k
Introduction Rule for Natural Number Order via Addition
Informal description
For any natural numbers $n$, $m$, and $k$, if $n + k = m$, then $n$ is less than or equal to $m$ (i.e., $n \leq m$).
Nat.not_le_of_gt theorem
{n m : Nat} (h : n > m) : ¬n ≤ m
Full source
protected theorem not_le_of_gt {n m : Nat} (h : n > m) : ¬ n ≤ m := fun h₁ =>
  match Nat.lt_or_ge n m with
  | Or.inl h₂ => absurd (Nat.lt_trans h h₂) (Nat.lt_irrefl _)
  | Or.inr h₂ =>
    have Heq : n = m := Nat.le_antisymm h₁ h₂
    absurd (@Eq.subst _ _ _ _ Heq h) (Nat.lt_irrefl m)
Strictly Greater Implies Not Less Than or Equal in Natural Numbers
Informal description
For any natural numbers $n$ and $m$, if $n > m$, then it is not the case that $n \leq m$.
Nat.not_le_of_lt theorem
: ∀ {a b : Nat}, a < b → ¬(b ≤ a)
Full source
protected theorem not_le_of_lt : ∀{a b : Nat}, a < b → ¬(b ≤ a) := Nat.not_le_of_gt
Strictly Less Implies Not Greater Than or Equal in Natural Numbers
Informal description
For any natural numbers $a$ and $b$, if $a < b$, then it is not the case that $b \leq a$.
Nat.not_lt_of_ge theorem
: ∀ {a b : Nat}, b ≥ a → ¬(b < a)
Full source
protected theorem not_lt_of_ge : ∀{a b : Nat}, b ≥ a → ¬(b < a) := flip Nat.not_le_of_gt
Non-strict Inequality Implies Negation of Strict Inequality in Natural Numbers
Informal description
For any natural numbers $a$ and $b$, if $b \geq a$, then it is not the case that $b < a$.
Nat.not_lt_of_le theorem
: ∀ {a b : Nat}, a ≤ b → ¬(b < a)
Full source
protected theorem not_lt_of_le : ∀{a b : Nat}, a ≤ b → ¬(b < a) := flip Nat.not_le_of_gt
Non-strict Inequality Implies No Strict Reverse Inequality in Natural Numbers
Informal description
For any natural numbers $a$ and $b$, if $a \leq b$, then it is not the case that $b < a$.
Nat.lt_le_asymm theorem
: ∀ {a b : Nat}, a < b → ¬(b ≤ a)
Full source
protected theorem lt_le_asymm : ∀{a b : Nat}, a < b → ¬(b ≤ a) := Nat.not_le_of_gt
Asymmetry of Strict and Non-Strict Order on Natural Numbers
Informal description
For any natural numbers $a$ and $b$, if $a < b$ then it is not the case that $b \leq a$.
Nat.le_lt_asymm theorem
: ∀ {a b : Nat}, a ≤ b → ¬(b < a)
Full source
protected theorem le_lt_asymm : ∀{a b : Nat}, a ≤ b → ¬(b < a) := flip Nat.not_le_of_gt
Asymmetry of Weak and Strict Inequality in Natural Numbers
Informal description
For any natural numbers $a$ and $b$, if $a \leq b$, then it is not the case that $b < a$.
Nat.gt_of_not_le theorem
{n m : Nat} (h : ¬n ≤ m) : n > m
Full source
theorem gt_of_not_le {n m : Nat} (h : ¬ n ≤ m) : n > m := (Nat.lt_or_ge m n).resolve_right h
$n > m$ if $n \nleq m$ for natural numbers
Informal description
For any natural numbers $n$ and $m$, if $n$ is not less than or equal to $m$, then $n$ is greater than $m$.
Nat.lt_of_not_ge theorem
: ∀ {a b : Nat}, ¬(b ≥ a) → b < a
Full source
protected theorem lt_of_not_ge : ∀{a b : Nat}, ¬(b ≥ a) → b < a := Nat.gt_of_not_le
Strict Inequality from Negation of Weak Inequality in Natural Numbers
Informal description
For any natural numbers $a$ and $b$, if $b$ is not greater than or equal to $a$, then $b$ is less than $a$.
Nat.lt_of_not_le theorem
: ∀ {a b : Nat}, ¬(a ≤ b) → b < a
Full source
protected theorem lt_of_not_le : ∀{a b : Nat}, ¬(a ≤ b) → b < a := Nat.gt_of_not_le
Strict Inequality from Non-Leq in Natural Numbers
Informal description
For any natural numbers $a$ and $b$, if $a$ is not less than or equal to $b$, then $b$ is strictly less than $a$.
Nat.ge_of_not_lt theorem
{n m : Nat} (h : ¬n < m) : n ≥ m
Full source
theorem ge_of_not_lt {n m : Nat} (h : ¬ n < m) : n ≥ m := (Nat.lt_or_ge n m).resolve_left h
Negation of Strict Inequality Implies Weak Inequality in Natural Numbers
Informal description
For any natural numbers $n$ and $m$, if it is not the case that $n < m$, then $n \geq m$.
Nat.le_of_not_gt theorem
: ∀ {a b : Nat}, ¬(b > a) → b ≤ a
Full source
protected theorem le_of_not_gt : ∀{a b : Nat}, ¬(b > a) → b ≤ a := Nat.ge_of_not_lt
Negation of Greater Than Implies Less Than or Equal in Natural Numbers
Informal description
For any natural numbers $a$ and $b$, if it is not the case that $b > a$, then $b \leq a$.
Nat.le_of_not_lt theorem
: ∀ {a b : Nat}, ¬(a < b) → b ≤ a
Full source
protected theorem le_of_not_lt : ∀{a b : Nat}, ¬(a < b) → b ≤ a := Nat.ge_of_not_lt
Negation of Strict Inequality Implies Weak Inequality in Natural Numbers (Reverse Order)
Informal description
For any natural numbers $a$ and $b$, if it is not the case that $a < b$, then $b \leq a$.
Nat.ne_of_gt theorem
{a b : Nat} (h : b < a) : a ≠ b
Full source
theorem ne_of_gt {a b : Nat} (h : b < a) : a ≠ b := (ne_of_lt h).symm
Strict Inequality Implies Inequality (Reverse Order) in Natural Numbers
Informal description
For any natural numbers $a$ and $b$, if $b < a$, then $a \neq b$.
Nat.ne_of_lt' theorem
: ∀ {a b : Nat}, a < b → b ≠ a
Full source
protected theorem ne_of_lt' : ∀{a b : Nat}, a < b → b ≠ a := ne_of_gt
Inequality from Strict Order in Natural Numbers
Informal description
For any natural numbers $a$ and $b$, if $a < b$, then $b \neq a$.
Nat.not_lt theorem
{a b : Nat} : ¬a < b ↔ b ≤ a
Full source
@[simp] protected theorem not_lt {a b : Nat} : ¬ a < b¬ a < b ↔ b ≤ a :=
  Iff.intro Nat.ge_of_not_lt (flip Nat.not_le_of_gt)
Negation of Strict Inequality Equivalent to Weak Inequality in Natural Numbers
Informal description
For any natural numbers $a$ and $b$, the statement that $a$ is not less than $b$ is equivalent to $b$ being less than or equal to $a$, i.e., $\neg (a < b) \leftrightarrow b \leq a$.
Nat.le_of_not_le theorem
{a b : Nat} (h : ¬b ≤ a) : a ≤ b
Full source
protected theorem le_of_not_le {a b : Nat} (h : ¬ b ≤ a) : a ≤ b := Nat.le_of_lt (Nat.not_le.1 h)
Non-strict Order from Negated Non-strict Order in Natural Numbers
Informal description
For any natural numbers $a$ and $b$, if $b$ is not less than or equal to $a$, then $a$ is less than or equal to $b$.
Nat.le_of_not_ge theorem
: ∀ {a b : Nat}, ¬(a ≥ b) → a ≤ b
Full source
protected theorem le_of_not_ge : ∀{a b : Nat}, ¬(a ≥ b) → a ≤ b:= @Nat.le_of_not_le
Non-strict Order from Negated Non-strict Order in Natural Numbers (variant)
Informal description
For any natural numbers $a$ and $b$, if $a$ is not greater than or equal to $b$, then $a$ is less than or equal to $b$.
Nat.lt_trichotomy theorem
(a b : Nat) : a < b ∨ a = b ∨ b < a
Full source
protected theorem lt_trichotomy (a b : Nat) : a < b ∨ a = b ∨ b < a :=
  match Nat.lt_or_ge a b with
  | .inl h => .inl h
  | .inr h =>
    match Nat.eq_or_lt_of_le h with
    | .inl h => .inr (.inl h.symm)
    | .inr h => .inr (.inr h)
Trichotomy Property for Natural Numbers: $a < b \lor a = b \lor b < a$
Informal description
For any two natural numbers $a$ and $b$, exactly one of the following holds: $a < b$, $a = b$, or $b < a$.
Nat.lt_or_lt_of_ne theorem
: ∀ {a b : Nat}, a ≠ b → a < b ∨ b < a
Full source
protected theorem lt_or_lt_of_ne : ∀{a b : Nat}, a ≠ ba < b ∨ b < a := Nat.lt_or_gt_of_ne
Strict Order Trichotomy for Natural Numbers: $a \neq b \implies a < b \lor b < a$
Informal description
For any two natural numbers $a$ and $b$ such that $a \neq b$, either $a < b$ or $b < a$.
Nat.instAntisymmLe instance
: Std.Antisymm (. ≤ . : Nat → Nat → Prop)
Full source
instance : Std.Antisymm ( . ≤ . : NatNat → Prop) where
  antisymm _ _ h₁ h₂ := Nat.le_antisymm h₁ h₂
Antisymmetry of the Natural Order on $\mathbb{N}$
Informal description
The less-than-or-equal-to relation $\leq$ on the natural numbers $\mathbb{N}$ is antisymmetric. That is, for any natural numbers $n$ and $m$, if $n \leq m$ and $m \leq n$, then $n = m$.
Nat.instAntisymmNotLt instance
: Std.Antisymm (¬. < . : Nat → Nat → Prop)
Full source
instance : Std.Antisymm (¬ . < . : NatNat → Prop) where
  antisymm _ _ h₁ h₂ := Nat.le_antisymm (Nat.ge_of_not_lt h₂) (Nat.ge_of_not_lt h₁)
Antisymmetry of the Negated Strict Order on Natural Numbers
Informal description
The relation $\neg(n < m)$ on natural numbers $n$ and $m$ is antisymmetric, meaning that if both $\neg(n < m)$ and $\neg(m < n)$ hold, then $n = m$.
Nat.add_le_add_left theorem
{n m : Nat} (h : n ≤ m) (k : Nat) : k + n ≤ k + m
Full source
protected theorem add_le_add_left {n m : Nat} (h : n ≤ m) (k : Nat) : k + n ≤ k + m :=
  match le.dest h with
  | ⟨w, hw⟩ =>
    have h₁ : k + n + w = k + (n + w) := Nat.add_assoc ..
    have h₂ : k + (n + w) = k + m     := congrArg _ hw
    le.intro <| h₁.trans h₂
Left Addition Preserves Order in Natural Numbers
Informal description
For any natural numbers $n$ and $m$ such that $n \leq m$, and for any natural number $k$, it holds that $k + n \leq k + m$.
Nat.add_le_add_right theorem
{n m : Nat} (h : n ≤ m) (k : Nat) : n + k ≤ m + k
Full source
protected theorem add_le_add_right {n m : Nat} (h : n ≤ m) (k : Nat) : n + k ≤ m + k := by
  rw [Nat.add_comm n k, Nat.add_comm m k]
  apply Nat.add_le_add_left
  assumption
Right Addition Preserves Order in Natural Numbers
Informal description
For any natural numbers $n$ and $m$ such that $n \leq m$, and for any natural number $k$, it holds that $n + k \leq m + k$.
Nat.add_lt_add_left theorem
{n m : Nat} (h : n < m) (k : Nat) : k + n < k + m
Full source
protected theorem add_lt_add_left {n m : Nat} (h : n < m) (k : Nat) : k + n < k + m :=
  lt_of_succ_le (add_succ k n ▸ Nat.add_le_add_left (succ_le_of_lt h) k)
Left Addition Preserves Strict Inequality in Natural Numbers
Informal description
For any natural numbers $n$ and $m$ such that $n < m$, and for any natural number $k$, it holds that $k + n < k + m$.
Nat.add_lt_add_right theorem
{n m : Nat} (h : n < m) (k : Nat) : n + k < m + k
Full source
protected theorem add_lt_add_right {n m : Nat} (h : n < m) (k : Nat) : n + k < m + k :=
  Nat.add_comm k m ▸ Nat.add_comm k n ▸ Nat.add_lt_add_left h k
Right Addition Preserves Strict Inequality in Natural Numbers
Informal description
For any natural numbers $n$ and $m$ such that $n < m$, and for any natural number $k$, it holds that $n + k < m + k$.
Nat.lt_add_of_pos_left theorem
(h : 0 < k) : n < k + n
Full source
protected theorem lt_add_of_pos_left (h : 0 < k) : n < k + n := by
  rw [Nat.add_comm]
  exact Nat.add_lt_add_left h n
Strict Inequality Under Left Addition with Positive Natural Number
Informal description
For any natural numbers $n$ and $k$ such that $0 < k$, it holds that $n < k + n$.
Nat.lt_add_of_pos_right theorem
(h : 0 < k) : n < n + k
Full source
protected theorem lt_add_of_pos_right (h : 0 < k) : n < n + k :=
  Nat.add_lt_add_left h n
Right Addition of Positive Number Increases Natural Number
Informal description
For any natural number $n$ and any positive natural number $k$ (i.e., $0 < k$), it holds that $n < n + k$.
Nat.zero_lt_one theorem
: 0 < (1 : Nat)
Full source
protected theorem zero_lt_one : 0 < (1:Nat) :=
  zero_lt_succ 0
Zero is Less Than One in Natural Numbers
Informal description
The natural number $0$ is strictly less than $1$, i.e., $0 < 1$.
Nat.pos_iff_ne_zero theorem
: 0 < n ↔ n ≠ 0
Full source
protected theorem pos_iff_ne_zero : 0 < n ↔ n ≠ 0 := ⟨ne_of_gt, Nat.pos_of_ne_zero⟩
Equivalence of Positivity and Non-Zero for Natural Numbers
Informal description
For any natural number $n$, the following are equivalent: 1. $0 < n$ (i.e., $n$ is positive) 2. $n \neq 0$ (i.e., $n$ is non-zero)
Nat.add_le_add theorem
{a b c d : Nat} (h₁ : a ≤ b) (h₂ : c ≤ d) : a + c ≤ b + d
Full source
theorem add_le_add {a b c d : Nat} (h₁ : a ≤ b) (h₂ : c ≤ d) : a + c ≤ b + d :=
  Nat.le_trans (Nat.add_le_add_right h₁ c) (Nat.add_le_add_left h₂ b)
Addition Preserves Order in Natural Numbers
Informal description
For any natural numbers $a, b, c, d$ such that $a \leq b$ and $c \leq d$, it holds that $a + c \leq b + d$.
Nat.add_lt_add theorem
{a b c d : Nat} (h₁ : a < b) (h₂ : c < d) : a + c < b + d
Full source
theorem add_lt_add {a b c d : Nat} (h₁ : a < b) (h₂ : c < d) : a + c < b + d :=
  Nat.lt_trans (Nat.add_lt_add_right h₁ c) (Nat.add_lt_add_left h₂ b)
Addition Preserves Strict Inequality in Natural Numbers
Informal description
For any natural numbers $a, b, c, d$ such that $a < b$ and $c < d$, it holds that $a + c < b + d$.
Nat.le_of_add_le_add_left theorem
{a b c : Nat} (h : a + b ≤ a + c) : b ≤ c
Full source
protected theorem le_of_add_le_add_left {a b c : Nat} (h : a + b ≤ a + c) : b ≤ c := by
  match le.dest h with
  | ⟨d, hd⟩ =>
    apply @le.intro _ _ d
    rw [Nat.add_assoc] at hd
    apply Nat.add_left_cancel hd
Left Addition Preserves Order in Natural Numbers
Informal description
For any natural numbers $a$, $b$, and $c$, if $a + b \leq a + c$, then $b \leq c$.
Nat.le_of_add_le_add_right theorem
{a b c : Nat} : a + b ≤ c + b → a ≤ c
Full source
protected theorem le_of_add_le_add_right {a b c : Nat} : a + b ≤ c + b → a ≤ c := by
  rw [Nat.add_comm _ b, Nat.add_comm _ b]
  apply Nat.le_of_add_le_add_left
Right Addition Preserves Order in Natural Numbers
Informal description
For any natural numbers $a$, $b$, and $c$, if $a + b \leq c + b$, then $a \leq c$.
Nat.lt_asymm theorem
{a b : Nat} (h : a < b) : ¬b < a
Full source
protected theorem lt_asymm {a b : Nat} (h : a < b) : ¬ b < a := Nat.not_lt.2 (Nat.le_of_lt h)
Asymmetry of Strict Order on Natural Numbers
Informal description
For any natural numbers $a$ and $b$, if $a < b$, then it is not the case that $b < a$.
Nat.not_lt_of_gt abbrev
Full source
/-- Alias for `Nat.lt_asymm`. -/
protected abbrev not_lt_of_gt := @Nat.lt_asymm
Asymmetry of Strict Greater-Than Relation on Natural Numbers
Informal description
For any natural numbers $a$ and $b$, if $a > b$, then it is not the case that $b > a$.
Nat.not_lt_of_lt abbrev
Full source
/-- Alias for `Nat.lt_asymm`. -/
protected abbrev not_lt_of_lt := @Nat.lt_asymm
Asymmetry of Strict Inequality on Natural Numbers
Informal description
For any natural numbers $a$ and $b$, if $a < b$, then it is not the case that $b < a$.
Nat.lt_iff_le_not_le theorem
{m n : Nat} : m < n ↔ m ≤ n ∧ ¬n ≤ m
Full source
protected theorem lt_iff_le_not_le {m n : Nat} : m < n ↔ m ≤ n ∧ ¬ n ≤ m :=
  ⟨fun h => ⟨Nat.le_of_lt h, Nat.not_le_of_gt h⟩, fun ⟨_, h⟩ => Nat.lt_of_not_ge h⟩
Characterization of Strict Inequality via Non-Strict Inequality in Natural Numbers
Informal description
For any natural numbers $m$ and $n$, the strict inequality $m < n$ holds if and only if $m \leq n$ and it is not the case that $n \leq m$.
Nat.lt_iff_le_and_not_ge abbrev
Full source
/-- Alias for `Nat.lt_iff_le_not_le`. -/
protected abbrev lt_iff_le_and_not_ge := @Nat.lt_iff_le_not_le
Characterization of Strict Inequality via Non-Strict Inequality and Non-Reversed Inequality in Natural Numbers
Informal description
For any natural numbers $m$ and $n$, the strict inequality $m < n$ holds if and only if $m \leq n$ and it is not the case that $n \leq m$.
Nat.lt_or_gt abbrev
Full source
/-- Alias for `Nat.ne_iff_lt_or_gt`. -/
protected abbrev lt_or_gt := @Nat.ne_iff_lt_or_gt
Inequality Equivalent to Strict Ordering in Natural Numbers
Informal description
For any two natural numbers $a$ and $b$, either $a < b$ or $b < a$ holds if and only if $a \neq b$.
Nat.le_or_ge abbrev
Full source
/-- Alias for `Nat.le_total`. -/
protected abbrev le_or_ge := @Nat.le_total
Total Order Property of Natural Numbers: $m \leq n$ or $n \leq m$
Informal description
For any two natural numbers $m$ and $n$, either $m \leq n$ or $n \leq m$ holds.
Nat.le_or_le abbrev
Full source
/-- Alias for `Nat.le_total`. -/
protected abbrev le_or_le := @Nat.le_total
Total Order Property of Natural Numbers: $m \leq n$ or $n \leq m$
Informal description
For any two natural numbers $m$ and $n$, either $m \leq n$ or $n \leq m$ holds.
Nat.lt_or_eq_of_le theorem
{n m : Nat} (h : n ≤ m) : n < m ∨ n = m
Full source
protected theorem lt_or_eq_of_le {n m : Nat} (h : n ≤ m) : n < m ∨ n = m :=
  (Nat.lt_or_ge ..).imp_right (Nat.le_antisymm h)
Strict Inequality or Equality from Non-Strict Inequality in Natural Numbers
Informal description
For any natural numbers $n$ and $m$ such that $n \leq m$, either $n < m$ or $n = m$ holds.
Nat.lt_succ_iff theorem
: m < succ n ↔ m ≤ n
Full source
protected theorem lt_succ_iff : m < succ n ↔ m ≤ n := ⟨le_of_lt_succ, lt_succ_of_le⟩
Strict Inequality with Successor Equivalent to Non-Strict Inequality
Informal description
For any natural numbers $m$ and $n$, the inequality $m < n + 1$ holds if and only if $m \leq n$.
Nat.lt_add_one_iff theorem
: m < n + 1 ↔ m ≤ n
Full source
protected theorem lt_add_one_iff : m < n + 1 ↔ m ≤ n := ⟨le_of_lt_succ, lt_succ_of_le⟩
Strict Inequality with Successor Equivalent to Non-Strict Inequality
Informal description
For any natural numbers $m$ and $n$, the inequality $m < n + 1$ holds if and only if $m \leq n$.
Nat.one_pos abbrev
Full source
/-- Alias for `Nat.zero_lt_one`. -/
protected abbrev one_pos := @Nat.zero_lt_one
Positivity of One: $0 < 1$
Informal description
The natural number $1$ is strictly greater than $0$, i.e., $0 < 1$.
Nat.two_pos theorem
: 0 < 2
Full source
protected theorem two_pos : 0 < 2 := Nat.zero_lt_succ _
Positivity of Two: $0 < 2$
Informal description
The natural number $2$ is strictly greater than $0$, i.e., $0 < 2$.
Nat.zero_lt_two theorem
: 0 < 2
Full source
protected theorem zero_lt_two : 0 < 2 := Nat.zero_lt_succ _
Positivity of Two: $0 < 2$
Informal description
The natural number $0$ is strictly less than $2$, i.e., $0 < 2$.
Nat.one_lt_two theorem
: 1 < 2
Full source
protected theorem one_lt_two : 1 < 2 := Nat.succ_lt_succ Nat.zero_lt_one
Inequality of Numerals: $1 < 2$
Informal description
The natural number $1$ is strictly less than $2$, i.e., $1 < 2$.
Nat.succ_ne_self theorem
(n) : succ n ≠ n
Full source
theorem succ_ne_self (n) : succsucc n ≠ n := Nat.ne_of_gt (lt_succ_self n)
Successor Inequality: $n + 1 \neq n$
Informal description
For any natural number $n$, its successor $n + 1$ is not equal to $n$ itself, i.e., $n + 1 \neq n$.
Nat.add_one_ne_self theorem
(n) : n + 1 ≠ n
Full source
theorem add_one_ne_self (n) : n + 1 ≠ n := Nat.ne_of_gt (lt_succ_self n)
Successor Inequality: $n + 1 \neq n$
Informal description
For any natural number $n$, the sum $n + 1$ is not equal to $n$.
Nat.succ_le theorem
: succ n ≤ m ↔ n < m
Full source
theorem succ_le : succsucc n ≤ m ↔ n < m := .rfl
Successor Inequality: $n + 1 \leq m \leftrightarrow n < m$
Informal description
For any natural numbers $n$ and $m$, the successor of $n$ is less than or equal to $m$ if and only if $n$ is strictly less than $m$, i.e., $n + 1 \leq m \leftrightarrow n < m$.
Nat.add_one_le_iff theorem
: n + 1 ≤ m ↔ n < m
Full source
theorem add_one_le_iff : n + 1 ≤ m ↔ n < m := .rfl
Successor Inequality: $n + 1 \leq m \leftrightarrow n < m$
Informal description
For any natural numbers $n$ and $m$, the inequality $n + 1 \leq m$ holds if and only if $n < m$.
Nat.lt_succ theorem
: m < succ n ↔ m ≤ n
Full source
theorem lt_succ : m < succ n ↔ m ≤ n := ⟨le_of_lt_succ, lt_succ_of_le⟩
Strict Inequality with Successor Equivalent to Non-Strict Inequality
Informal description
For any natural numbers $m$ and $n$, the strict inequality $m < n + 1$ holds if and only if $m \leq n$.
Nat.lt_succ_of_lt theorem
(h : a < b) : a < succ b
Full source
theorem lt_succ_of_lt (h : a < b) : a < succ b := le_succ_of_le h
Successor Preserves Strict Inequality in Natural Numbers
Informal description
For any natural numbers $a$ and $b$, if $a < b$, then $a < b + 1$.
Nat.lt_add_one_of_lt theorem
(h : a < b) : a < b + 1
Full source
theorem lt_add_one_of_lt (h : a < b) : a < b + 1 := le_succ_of_le h
Successor Preserves Strict Inequality in Natural Numbers
Informal description
For any natural numbers $a$ and $b$, if $a < b$, then $a < b + 1$.
Nat.lt_one_iff theorem
: n < 1 ↔ n = 0
Full source
@[simp] theorem lt_one_iff : n < 1 ↔ n = 0 := Nat.lt_succ_iff.trans <| by rw [le_zero_eq]
Natural Number Strictly Less Than One is Zero ($n < 1 \leftrightarrow n = 0$)
Informal description
For any natural number $n$, the inequality $n < 1$ holds if and only if $n = 0$.
Nat.succ_pred_eq_of_ne_zero theorem
: ∀ {n}, n ≠ 0 → succ (pred n) = n
Full source
theorem succ_pred_eq_of_ne_zero : ∀ {n}, n ≠ 0succ (pred n) = n
  | _+1, _ => rfl
Successor of Predecessor Equals Original for Nonzero Natural Numbers
Informal description
For any natural number $n \neq 0$, the successor of the predecessor of $n$ equals $n$, i.e., $\mathrm{succ}(\mathrm{pred}(n)) = n$.
Nat.succ_inj' theorem
: succ a = succ b ↔ a = b
Full source
theorem succ_inj' : succsucc a = succ b ↔ a = b := (Nat.succ.injEq a b).to_iff
Injectivity of Successor Function on Natural Numbers
Informal description
For any natural numbers $a$ and $b$, the successor of $a$ equals the successor of $b$ if and only if $a = b$. That is, $a + 1 = b + 1 \leftrightarrow a = b$.
Nat.succ_le_succ_iff theorem
: succ a ≤ succ b ↔ a ≤ b
Full source
theorem succ_le_succ_iff : succsucc a ≤ succ b ↔ a ≤ b := ⟨le_of_succ_le_succ, succ_le_succ⟩
Successor Preserves Order Bi-implication on Natural Numbers
Informal description
For any natural numbers $a$ and $b$, the successor of $a$ is less than or equal to the successor of $b$ if and only if $a \leq b$. That is, $a + 1 \leq b + 1 \leftrightarrow a \leq b$.
Nat.succ_lt_succ_iff theorem
: succ a < succ b ↔ a < b
Full source
theorem succ_lt_succ_iff : succsucc a < succ b ↔ a < b := ⟨lt_of_succ_lt_succ, succ_lt_succ⟩
Successor Preserves Strict Order Bi-implication on Natural Numbers
Informal description
For any natural numbers $a$ and $b$, the successor of $a$ is strictly less than the successor of $b$ if and only if $a$ is strictly less than $b$. That is, $a + 1 < b + 1 \leftrightarrow a < b$.
Nat.add_one_inj theorem
: a + 1 = b + 1 ↔ a = b
Full source
theorem add_one_inj : a + 1 = b + 1 ↔ a = b := succ_inj'
Injectivity of Successor via Addition: $a + 1 = b + 1 \leftrightarrow a = b$
Informal description
For any natural numbers $a$ and $b$, the equality $a + 1 = b + 1$ holds if and only if $a = b$.
Nat.ne_add_one theorem
(n : Nat) : n ≠ n + 1
Full source
theorem ne_add_one (n : Nat) : n ≠ n + 1 := fun h => by cases h
Natural Number Inequality: $n \neq n + 1$
Informal description
For any natural number $n$, it holds that $n \neq n + 1$.
Nat.add_one_ne theorem
(n : Nat) : n + 1 ≠ n
Full source
theorem add_one_ne (n : Nat) : n + 1 ≠ n := fun h => by cases h
Successor Inequality: $n + 1 \neq n$ for Natural Numbers
Informal description
For any natural number $n$, the successor of $n$ (i.e., $n + 1$) is not equal to $n$.
Nat.add_one_le_add_one_iff theorem
: a + 1 ≤ b + 1 ↔ a ≤ b
Full source
theorem add_one_le_add_one_iff : a + 1 ≤ b + 1 ↔ a ≤ b := succ_le_succ_iff
Successor Preserves Order: $a+1 \leq b+1 \leftrightarrow a \leq b$
Informal description
For any natural numbers $a$ and $b$, the inequality $a + 1 \leq b + 1$ holds if and only if $a \leq b$.
Nat.add_one_lt_add_one_iff theorem
: a + 1 < b + 1 ↔ a < b
Full source
theorem add_one_lt_add_one_iff : a + 1 < b + 1 ↔ a < b := succ_lt_succ_iff
Successor Preserves Strict Order: $a+1 < b+1 \leftrightarrow a < b$
Informal description
For any natural numbers $a$ and $b$, the inequality $a + 1 < b + 1$ holds if and only if $a < b$.
Nat.pred_inj theorem
: ∀ {a b}, 0 < a → 0 < b → pred a = pred b → a = b
Full source
theorem pred_inj : ∀ {a b}, 0 < a → 0 < b → pred a = pred b → a = b
  | _+1, _+1, _, _ => congrArg _
Injectivity of Predecessor on Positive Natural Numbers
Informal description
For any positive natural numbers $a$ and $b$, if the predecessors of $a$ and $b$ are equal, then $a = b$.
Nat.pred_ne_self theorem
: ∀ {a}, a ≠ 0 → pred a ≠ a
Full source
theorem pred_ne_self : ∀ {a}, a ≠ 0predpred a ≠ a
  | _+1, _ => (succ_ne_self _).symm
Predecessor Inequality: $\mathrm{pred}(a) \neq a$ for $a \neq 0$
Informal description
For any natural number $a \neq 0$, the predecessor of $a$ is not equal to $a$ itself, i.e., $\mathrm{pred}(a) \neq a$.
Nat.sub_one_ne_self theorem
: ∀ {a}, a ≠ 0 → a - 1 ≠ a
Full source
theorem sub_one_ne_self : ∀ {a}, a ≠ 0a - 1 ≠ a
  | _+1, _ => (succ_ne_self _).symm
Predecessor Inequality: $a - 1 \neq a$ for $a \neq 0$
Informal description
For any nonzero natural number $a$, the predecessor $a - 1$ is not equal to $a$ itself, i.e., $a - 1 \neq a$.
Nat.pred_lt_self theorem
: ∀ {a}, 0 < a → pred a < a
Full source
theorem pred_lt_self : ∀ {a}, 0 < a → pred a < a
  | _+1, _ => lt_succ_self _
Predecessor is strictly less than positive natural numbers
Informal description
For any natural number $a$ such that $0 < a$, the predecessor of $a$ is strictly less than $a$, i.e., $\mathrm{pred}(a) < a$.
Nat.pred_lt_pred theorem
: ∀ {n m}, n ≠ 0 → n < m → pred n < pred m
Full source
theorem pred_lt_pred : ∀ {n m}, n ≠ 0 → n < m → pred n < pred m
  | _+1, _+1, _, h => lt_of_succ_lt_succ h
Predecessor Preserves Strict Inequality for Nonzero Natural Numbers
Informal description
For any natural numbers $n$ and $m$ such that $n \neq 0$ and $n < m$, it holds that $\mathrm{pred}(n) < \mathrm{pred}(m)$.
Nat.pred_le_iff_le_succ theorem
: ∀ {n m}, pred n ≤ m ↔ n ≤ succ m
Full source
theorem pred_le_iff_le_succ : ∀ {n m}, predpred n ≤ m ↔ n ≤ succ m
  | 0, _ => ⟨fun _ => Nat.zero_le _, fun _ => Nat.zero_le _⟩
  | _+1, _ => Nat.succ_le_succ_iff.symm
Predecessor-Leq-Successor Equivalence for Natural Numbers
Informal description
For any natural numbers $n$ and $m$, the predecessor of $n$ is less than or equal to $m$ if and only if $n$ is less than or equal to the successor of $m$. That is, $\mathrm{pred}(n) \leq m \leftrightarrow n \leq m + 1$.
Nat.le_succ_of_pred_le theorem
: pred n ≤ m → n ≤ succ m
Full source
theorem le_succ_of_pred_le : pred n ≤ m → n ≤ succ m := pred_le_iff_le_succ.1
Implication from Predecessor Inequality to Successor Inequality in Natural Numbers
Informal description
For any natural numbers $n$ and $m$, if the predecessor of $n$ is less than or equal to $m$, then $n$ is less than or equal to the successor of $m$. That is, $\mathrm{pred}(n) \leq m$ implies $n \leq m + 1$.
Nat.pred_le_of_le_succ theorem
: n ≤ succ m → pred n ≤ m
Full source
theorem pred_le_of_le_succ : n ≤ succ m → pred n ≤ m := pred_le_iff_le_succ.2
Predecessor Bound from Successor Inequality
Informal description
For any natural numbers $n$ and $m$, if $n \leq m + 1$, then the predecessor of $n$ is less than or equal to $m$.
Nat.lt_pred_iff_succ_lt theorem
: ∀ {n m}, n < pred m ↔ succ n < m
Full source
theorem lt_pred_iff_succ_lt : ∀ {n m}, n < pred m ↔ succ n < m
  | _, 0 => ⟨nofun, nofun⟩
  | _, _+1 => Nat.succ_lt_succ_iff.symm
Predecessor-Successor Inequality Equivalence for Natural Numbers
Informal description
For any natural numbers $n$ and $m$, the inequality $n < \mathrm{pred}(m)$ holds if and only if $\mathrm{succ}(n) < m$ holds.
Nat.succ_lt_of_lt_pred theorem
: n < pred m → succ n < m
Full source
theorem succ_lt_of_lt_pred : n < pred m → succ n < m := lt_pred_iff_succ_lt.1
Successor Inequality from Predecessor Inequality in Natural Numbers
Informal description
For any natural numbers $n$ and $m$, if $n$ is less than the predecessor of $m$, then the successor of $n$ is less than $m$.
Nat.lt_pred_of_succ_lt theorem
: succ n < m → n < pred m
Full source
theorem lt_pred_of_succ_lt : succ n < m → n < pred m := lt_pred_iff_succ_lt.2
Implication from Successor Inequality to Predecessor Inequality for Natural Numbers
Informal description
For any natural numbers $n$ and $m$, if $\mathrm{succ}(n) < m$ then $n < \mathrm{pred}(m)$.
Nat.le_pred_iff_lt theorem
: ∀ {n m}, 0 < m → (n ≤ pred m ↔ n < m)
Full source
theorem le_pred_iff_lt : ∀ {n m}, 0 < m → (n ≤ pred m ↔ n < m)
  | 0, _+1, _ => ⟨fun _ => Nat.zero_lt_succ _, fun _ => Nat.zero_le _⟩
  | _+1, _+1, _ => Nat.lt_pred_iff_succ_lt
Predecessor Inequality Equivalence for Positive Natural Numbers: $n \leq m-1 \leftrightarrow n < m$
Informal description
For any natural numbers $n$ and $m$ with $m > 0$, the inequality $n \leq \mathrm{pred}(m)$ holds if and only if $n < m$.
Nat.le_pred_of_lt theorem
(h : n < m) : n ≤ pred m
Full source
theorem le_pred_of_lt (h : n < m) : n ≤ pred m := (le_pred_iff_lt (Nat.zero_lt_of_lt h)).2 h
Predecessor Bound from Strict Inequality in Natural Numbers
Informal description
For any natural numbers $n$ and $m$, if $n < m$, then $n \leq \mathrm{pred}(m)$.
Nat.le_sub_one_of_lt theorem
: a < b → a ≤ b - 1
Full source
theorem le_sub_one_of_lt : a < b → a ≤ b - 1 := Nat.le_pred_of_lt
Inequality Implication: $a < b \Rightarrow a \leq b - 1$ for Natural Numbers
Informal description
For any natural numbers $a$ and $b$, if $a < b$, then $a \leq b - 1$.
Nat.lt_of_le_pred theorem
(h : 0 < m) : n ≤ pred m → n < m
Full source
theorem lt_of_le_pred (h : 0 < m) : n ≤ pred m → n < m := (le_pred_iff_lt h).1
Inequality from Predecessor Bound: $n \leq m-1 \Rightarrow n < m$
Informal description
For any natural numbers $n$ and $m$ with $m > 0$, if $n \leq \mathrm{pred}(m)$, then $n < m$.
Nat.lt_of_le_sub_one theorem
(h : 0 < m) : n ≤ m - 1 → n < m
Full source
theorem lt_of_le_sub_one (h : 0 < m) : n ≤ m - 1 → n < m := (le_pred_iff_lt h).1
Inequality from Predecessor in Natural Numbers: $n \leq m-1 \Rightarrow n < m$
Informal description
For any natural numbers $n$ and $m$ with $m > 0$, if $n \leq m - 1$, then $n < m$.
Nat.exists_eq_succ_of_ne_zero theorem
: ∀ {n}, n ≠ 0 → Exists fun k => n = succ k
Full source
theorem exists_eq_succ_of_ne_zero : ∀ {n}, n ≠ 0Exists fun k => n = succ k
  | _+1, _ => ⟨_, rfl⟩
Nonzero Natural Numbers are Successors
Informal description
For any natural number $n \neq 0$, there exists a natural number $k$ such that $n = k + 1$.
Nat.exists_eq_add_one_of_ne_zero theorem
: ∀ {n}, n ≠ 0 → Exists fun k => n = k + 1
Full source
theorem exists_eq_add_one_of_ne_zero : ∀ {n}, n ≠ 0Exists fun k => n = k + 1
  | _+1, _ => ⟨_, rfl⟩
Nonzero Natural Numbers Are Successors: $n \neq 0 \implies \exists k, n = k + 1$
Informal description
For any natural number $n \neq 0$, there exists a natural number $k$ such that $n = k + 1$.
Nat.ctor_eq_zero theorem
: Nat.zero = 0
Full source
theorem ctor_eq_zero : Nat.zero = 0 :=
  rfl
Constructor Zero Equals Numeric Zero
Informal description
The natural number zero defined via the constructor `Nat.zero` is equal to the numeric literal `0`, i.e., $\mathtt{Nat.zero} = 0$.
Nat.one_ne_zero theorem
: 1 ≠ (0 : Nat)
Full source
protected theorem one_ne_zero : 1 ≠ (0 : Nat) :=
  fun h => Nat.noConfusion h
Inequality of One and Zero in Natural Numbers
Informal description
The natural number $1$ is not equal to $0$, i.e., $1 \neq 0$.
Nat.zero_ne_one theorem
: 0 ≠ (1 : Nat)
Full source
protected theorem zero_ne_one : 0 ≠ (1 : Nat) :=
  fun h => Nat.noConfusion h
Inequality of Zero and One in Natural Numbers
Informal description
The natural numbers $0$ and $1$ are not equal, i.e., $0 \neq 1$.
Nat.succ_ne_zero theorem
(n : Nat) : succ n ≠ 0
Full source
theorem succ_ne_zero (n : Nat) : succsucc n ≠ 0 := by simp
Successor of Natural Number is Nonzero
Informal description
For any natural number $n$, the successor of $n$ is not equal to zero, i.e., $n + 1 \neq 0$.
Nat.instNeZeroSucc instance
{n : Nat} : NeZero (n + 1)
Full source
instance instNeZeroSucc {n : Nat} : NeZero (n + 1) := ⟨succ_ne_zero n⟩
Successor of Natural Number is Nonzero
Informal description
For any natural number $n$, the successor $n + 1$ is nonzero.
Nat.mul_le_mul_left theorem
{n m : Nat} (k : Nat) (h : n ≤ m) : k * n ≤ k * m
Full source
theorem mul_le_mul_left {n m : Nat} (k : Nat) (h : n ≤ m) : k * n ≤ k * m :=
  match le.dest h with
  | ⟨l, hl⟩ =>
    have : k * n + k * l = k * m := Nat.left_distrib k n l ▸ hl.symmrfl
    le.intro this
Left Multiplication Preserves Order in Natural Numbers
Informal description
For any natural numbers $n$ and $m$ such that $n \leq m$, and for any natural number $k$, it holds that $k \cdot n \leq k \cdot m$.
Nat.mul_le_mul_right theorem
{n m : Nat} (k : Nat) (h : n ≤ m) : n * k ≤ m * k
Full source
theorem mul_le_mul_right {n m : Nat} (k : Nat) (h : n ≤ m) : n * k ≤ m * k :=
  Nat.mul_comm k m ▸ Nat.mul_comm k n ▸ mul_le_mul_left k h
Right Multiplication Preserves Order in Natural Numbers
Informal description
For any natural numbers $n$ and $m$ such that $n \leq m$, and for any natural number $k$, it holds that $n \cdot k \leq m \cdot k$.
Nat.mul_le_mul theorem
{n₁ m₁ n₂ m₂ : Nat} (h₁ : n₁ ≤ n₂) (h₂ : m₁ ≤ m₂) : n₁ * m₁ ≤ n₂ * m₂
Full source
protected theorem mul_le_mul {n₁ m₁ n₂ m₂ : Nat} (h₁ : n₁ ≤ n₂) (h₂ : m₁ ≤ m₂) : n₁ * m₁ ≤ n₂ * m₂ :=
  Nat.le_trans (mul_le_mul_right _ h₁) (mul_le_mul_left _ h₂)
Multiplication Preserves Order in Natural Numbers
Informal description
For any natural numbers $n_1, n_2, m_1, m_2$ such that $n_1 \leq n_2$ and $m_1 \leq m_2$, it holds that $n_1 \cdot m_1 \leq n_2 \cdot m_2$.
Nat.mul_lt_mul_of_pos_left theorem
{n m k : Nat} (h : n < m) (hk : k > 0) : k * n < k * m
Full source
protected theorem mul_lt_mul_of_pos_left {n m k : Nat} (h : n < m) (hk : k > 0) : k * n < k * m :=
  Nat.lt_of_lt_of_le (Nat.add_lt_add_left hk _) (Nat.mul_succ k n ▸ Nat.mul_le_mul_left k (succ_le_of_lt h))
Left Multiplication by Positive Natural Number Preserves Strict Inequality
Informal description
For any natural numbers $n$, $m$, and $k$ such that $n < m$ and $k > 0$, it holds that $k \cdot n < k \cdot m$.
Nat.mul_lt_mul_of_pos_right theorem
{n m k : Nat} (h : n < m) (hk : k > 0) : n * k < m * k
Full source
protected theorem mul_lt_mul_of_pos_right {n m k : Nat} (h : n < m) (hk : k > 0) : n * k < m * k :=
  Nat.mul_comm k m ▸ Nat.mul_comm k n ▸ Nat.mul_lt_mul_of_pos_left h hk
Right Multiplication by Positive Natural Number Preserves Strict Inequality
Informal description
For any natural numbers $n$, $m$, and $k$ such that $n < m$ and $k > 0$, it holds that $n \cdot k < m \cdot k$.
Nat.mul_pos theorem
{n m : Nat} (ha : n > 0) (hb : m > 0) : n * m > 0
Full source
protected theorem mul_pos {n m : Nat} (ha : n > 0) (hb : m > 0) : n * m > 0 :=
  have h : 0 * m < n * m := Nat.mul_lt_mul_of_pos_right ha hb
  Nat.zero_mul m ▸ h
Product of Positive Natural Numbers is Positive
Informal description
For any natural numbers $n$ and $m$ such that $n > 0$ and $m > 0$, their product satisfies $n \cdot m > 0$.
Nat.le_of_mul_le_mul_left theorem
{a b c : Nat} (h : c * a ≤ c * b) (hc : 0 < c) : a ≤ b
Full source
protected theorem le_of_mul_le_mul_left {a b c : Nat} (h : c * a ≤ c * b) (hc : 0 < c) : a ≤ b :=
  Nat.ge_of_not_lt fun hlt : b < a =>
    have h' : c * b < c * a := Nat.mul_lt_mul_of_pos_left hlt hc
    absurd h (Nat.not_le_of_gt h')
Left Multiplication by Positive Natural Number Preserves Inequality
Informal description
For any natural numbers $a$, $b$, and $c$ such that $c > 0$, if $c \cdot a \leq c \cdot b$, then $a \leq b$.
Nat.le_of_mul_le_mul_right theorem
{a b c : Nat} (h : a * c ≤ b * c) (hc : 0 < c) : a ≤ b
Full source
protected theorem le_of_mul_le_mul_right {a b c : Nat} (h : a * c ≤ b * c) (hc : 0 < c) : a ≤ b := by
  rw [Nat.mul_comm a c, Nat.mul_comm b c] at h
  exact Nat.le_of_mul_le_mul_left h hc
Right Multiplication by Positive Natural Number Preserves Inequality
Informal description
For any natural numbers $a$, $b$, and $c$ such that $c > 0$, if $a \cdot c \leq b \cdot c$, then $a \leq b$.
Nat.mul_le_mul_left_iff theorem
{n m k : Nat} (w : 0 < k) : k * n ≤ k * m ↔ n ≤ m
Full source
protected theorem mul_le_mul_left_iff {n m k : Nat} (w : 0 < k) : k * n ≤ k * m ↔ n ≤ m :=
  ⟨fun h => Nat.le_of_mul_le_mul_left h w, fun h => mul_le_mul_left _ h⟩
Left Multiplication by Positive Natural Number Preserves Inequality (Equivalence)
Informal description
For any natural numbers $n$, $m$, and $k$ with $k > 0$, the inequality $k \cdot n \leq k \cdot m$ holds if and only if $n \leq m$.
Nat.mul_le_mul_right_iff theorem
{n m k : Nat} (w : 0 < k) : n * k ≤ m * k ↔ n ≤ m
Full source
protected theorem mul_le_mul_right_iff {n m k : Nat} (w : 0 < k) : n * k ≤ m * k ↔ n ≤ m :=
  ⟨fun h => Nat.le_of_mul_le_mul_right h w, fun h => mul_le_mul_right _ h⟩
Right Multiplication by Positive Natural Number Preserves Inequality (Equivalence)
Informal description
For any natural numbers $n$, $m$, and $k$ with $k > 0$, the inequality $n \cdot k \leq m \cdot k$ holds if and only if $n \leq m$.
Nat.pow_succ theorem
(n m : Nat) : n ^ (succ m) = n ^ m * n
Full source
protected theorem pow_succ (n m : Nat) : n^(succ m) = n^m * n :=
  rfl
Power of Successor: $n^{m+1} = n^m \cdot n$
Informal description
For any natural numbers $n$ and $m$, the power $n$ raised to the successor of $m$ equals $n^m \cdot n$.
Nat.pow_add_one theorem
(n m : Nat) : n ^ (m + 1) = n ^ m * n
Full source
protected theorem pow_add_one (n m : Nat) : n^(m + 1) = n^m * n :=
  rfl
Power Recursion Formula: $n^{m+1} = n^m \cdot n$
Informal description
For any natural numbers $n$ and $m$, the power $n^{m+1}$ is equal to $n^m \cdot n$.
Nat.pow_zero theorem
(n : Nat) : n ^ 0 = 1
Full source
@[simp] protected theorem pow_zero (n : Nat) : n^0 = 1 := rfl
Natural Number Zeroth Power Identity: $n^0 = 1$
Informal description
For any natural number $n$, the zeroth power of $n$ is equal to $1$, i.e., $n^0 = 1$.
Nat.pow_one theorem
(a : Nat) : a ^ 1 = a
Full source
@[simp] protected theorem pow_one (a : Nat) : a ^ 1 = a := by
  simp [Nat.pow_succ]
First Power Identity for Natural Numbers: $a^1 = a$
Informal description
For any natural number $a$, the first power of $a$ equals $a$ itself, i.e., $a^1 = a$.
Nat.pow_le_pow_left theorem
{n m : Nat} (h : n ≤ m) : ∀ (i : Nat), n ^ i ≤ m ^ i
Full source
protected theorem pow_le_pow_left {n m : Nat} (h : n ≤ m) : ∀ (i : Nat), n^i ≤ m^i
  | 0      => Nat.le_refl _
  | succ i => Nat.mul_le_mul (Nat.pow_le_pow_left h i) h
Exponentiation Preserves Order in Natural Numbers
Informal description
For any natural numbers $n$ and $m$ such that $n \leq m$, and for any natural number $i$, it holds that $n^i \leq m^i$.
Nat.pow_le_pow_right theorem
{n : Nat} (hx : n > 0) {i : Nat} : ∀ {j}, i ≤ j → n ^ i ≤ n ^ j
Full source
protected theorem pow_le_pow_right {n : Nat} (hx : n > 0) {i : Nat} : ∀ {j}, i ≤ j → n^i ≤ n^j
  | 0,      h =>
    have : i = 0 := eq_zero_of_le_zero h
    this.symm ▸ Nat.le_refl _
  | succ j, h =>
    match le_or_eq_of_le_succ h with
    | Or.inl h => show n^i ≤ n^j * n from
      have : n^i * 1 ≤ n^j * n := Nat.mul_le_mul (Nat.pow_le_pow_right hx h) hx
      Nat.mul_one (n^i) ▸ this
    | Or.inr h =>
      h.symm ▸ Nat.le_refl _
Monotonicity of Natural Number Exponentiation: $n > 0 \land i \leq j \Rightarrow n^i \leq n^j$
Informal description
For any natural number $n > 0$ and any natural numbers $i, j$ such that $i \leq j$, it holds that $n^i \leq n^j$.
Nat.pow_le_pow_of_le_left abbrev
Full source
@[deprecated Nat.pow_le_pow_left (since := "2025-02-17")]
abbrev pow_le_pow_of_le_left := @Nat.pow_le_pow_left
Exponentiation Preserves Order in Natural Numbers (Left Argument)
Informal description
For any natural numbers $n$ and $m$ such that $n \leq m$, and for any natural number $i$, it holds that $n^i \leq m^i$.
Nat.pow_le_pow_of_le_right abbrev
Full source
@[deprecated Nat.pow_le_pow_right (since := "2025-02-17")]
abbrev pow_le_pow_of_le_right := @Nat.pow_le_pow_right
Monotonicity of Exponentiation with Respect to Base: $a \leq b \Rightarrow a^n \leq b^n$
Informal description
For any natural numbers $a$, $b$, and $n$ such that $a \leq b$, it holds that $a^n \leq b^n$.
Nat.pow_pos theorem
(h : 0 < a) : 0 < a ^ n
Full source
protected theorem pow_pos (h : 0 < a) : 0 < a^n :=
  match n with
  | 0 => Nat.zero_lt_one
  | _ + 1 => Nat.mul_pos (Nat.pow_pos h) h
Positivity of Natural Number Powers: $a > 0 \implies a^n > 0$
Informal description
For any natural numbers $a$ and $n$, if $a > 0$, then $a^n > 0$.
Nat.pos_pow_of_pos abbrev
Full source
@[deprecated Nat.pow_pos (since := "2025-02-17")]
abbrev pos_pow_of_pos := @Nat.pow_pos
Positivity of Positive Natural Number Powers: $a > 0 \land n > 0 \implies a^n > 0$
Informal description
For any natural numbers $a$ and $n$, if $a > 0$ and $n > 0$, then $a^n > 0$.
Nat.zero_pow_of_pos theorem
(n : Nat) (h : 0 < n) : 0 ^ n = 0
Full source
@[simp] theorem zero_pow_of_pos (n : Nat) (h : 0 < n) : 0 ^ n = 0 := by
  cases n with
  | zero => cases h
  | succ n => simp [Nat.pow_succ]
Zero to Positive Power Equals Zero ($0^n = 0$ for $n > 0$)
Informal description
For any natural number $n > 0$, the $n$-th power of zero equals zero, i.e., $0^n = 0$.
Nat.two_pow_pos theorem
(w : Nat) : 0 < 2 ^ w
Full source
protected theorem two_pow_pos (w : Nat) : 0 < 2^w := Nat.pow_pos (by decide)
Positivity of Powers of Two: $0 < 2^w$
Informal description
For any natural number $w$, the power $2^w$ is strictly positive, i.e., $0 < 2^w$.
Nat.mul_pow theorem
(a b n : Nat) : (a * b) ^ n = a ^ n * b ^ n
Full source
protected theorem mul_pow (a b n : Nat) : (a * b) ^ n = a ^ n * b ^ n := by
  induction n with
  | zero => simp [Nat.pow_zero]
  | succ n ih =>
    rw [Nat.pow_succ, ih, Nat.pow_succ, Nat.pow_succ, ← Nat.mul_assoc, ← Nat.mul_assoc]
    congr 1
    rw [Nat.mul_assoc, Nat.mul_assoc, Nat.mul_comm _ a]
Power of Product Equals Product of Powers: $(a \cdot b)^n = a^n \cdot b^n$
Informal description
For any natural numbers $a$, $b$, and $n$, the power of the product $(a \cdot b)^n$ is equal to the product of the powers $a^n \cdot b^n$.
Nat.pow_lt_pow_left theorem
{a b n : Nat} (hab : a < b) (h : n ≠ 0) : a ^ n < b ^ n
Full source
protected theorem pow_lt_pow_left {a b n : Nat} (hab : a < b) (h : n ≠ 0) : a ^ n < b ^ n := by
  cases n with
  | zero => simp at h
  | succ n =>
    clear h
    induction n with
    | zero => simpa
    | succ n ih =>
      rw [Nat.pow_succ a, Nat.pow_succ b]
      exact Nat.lt_of_le_of_lt (Nat.mul_le_mul_left _ (Nat.le_of_lt hab))
        (Nat.mul_lt_mul_of_pos_right ih (Nat.lt_of_le_of_lt (Nat.zero_le _) hab))
Strict Monotonicity of Natural Number Exponentiation: $a < b \land n \neq 0 \implies a^n < b^n$
Informal description
For any natural numbers $a$, $b$, and $n$ such that $a < b$ and $n \neq 0$, it holds that $a^n < b^n$.
Nat.pow_left_inj theorem
{a b n : Nat} (hn : n ≠ 0) : a ^ n = b ^ n ↔ a = b
Full source
protected theorem pow_left_inj {a b n : Nat} (hn : n ≠ 0) : a ^ n = b ^ n ↔ a = b := by
  refine ⟨fun h => ?_, (· ▸ rfl)⟩
  match Nat.lt_trichotomy a b with
  | Or.inl hab => exact False.elim (absurd h (ne_of_lt (Nat.pow_lt_pow_left hab hn)))
  | Or.inr (Or.inl hab) => exact hab
  | Or.inr (Or.inr hab) => exact False.elim (absurd h (Nat.ne_of_lt' (Nat.pow_lt_pow_left hab hn)))
Injectivity of Natural Number Exponentiation: $a^n = b^n \leftrightarrow a = b$ for $n \neq 0$
Informal description
For any natural numbers $a$, $b$, and $n$ with $n \neq 0$, the equality $a^n = b^n$ holds if and only if $a = b$.
Nat.min abbrev
(n m : Nat)
Full source
/--
Returns the lesser of two natural numbers. Usually accessed via `Min.min`.

Returns `n` if `n ≤ m`, or `m` if `m ≤ n`.

Examples:
* `min 0 5 = 0`
* `min 4 5 = 4`
* `min 4 3 = 3`
* `min 8 8 = 8`
-/
protected abbrev min (n m : Nat) := min n m
Minimum of Two Natural Numbers
Informal description
The minimum of two natural numbers $n$ and $m$, denoted $\min(n, m)$, is defined to be $n$ if $n \leq m$ and $m$ otherwise.
Nat.min_def theorem
{n m : Nat} : min n m = if n ≤ m then n else m
Full source
protected theorem min_def {n m : Nat} : min n m = if n ≤ m then n else m := rfl
Definition of Minimum for Natural Numbers via Conditional
Informal description
For any natural numbers $n$ and $m$, the minimum $\min(n, m)$ is equal to $n$ if $n \leq m$ and $m$ otherwise.
Nat.instMax instance
: Max Nat
Full source
instance : Max Nat := maxOfLe
The Maximum Operation on Natural Numbers
Informal description
The natural numbers $\mathbb{N}$ are equipped with a canonical maximum operation $\max$ defined by $\max(n, m) = m$ if $n \leq m$ and $\max(n, m) = n$ otherwise.
Nat.max abbrev
(n m : Nat)
Full source
/--
Returns the greater of two natural numbers. Usually accessed via `Max.max`.

Returns `m` if `n ≤ m`, or `n` if `m ≤ n`.

Examples:
* `max 0 5 = 5`
* `max 4 5 = 5`
* `max 4 3 = 4`
* `max 8 8 = 8`
-/
protected abbrev max (n m : Nat) := max n m
Maximum of Two Natural Numbers
Informal description
The maximum of two natural numbers $n$ and $m$, defined as $m$ if $n \leq m$ and $n$ otherwise.
Nat.max_def theorem
{n m : Nat} : max n m = if n ≤ m then m else n
Full source
protected theorem max_def {n m : Nat} : max n m = if n ≤ m then m else n := rfl
Definition of Maximum for Natural Numbers via Conditional
Informal description
For any natural numbers $n$ and $m$, the maximum $\max(n, m)$ is equal to $m$ if $n \leq m$, and equal to $n$ otherwise.
Nat.ne_zero_of_lt theorem
(h : b < a) : a ≠ 0
Full source
protected theorem ne_zero_of_lt (h : b < a) : a ≠ 0 := by
  cases a
  exact absurd h (Nat.not_lt_zero _)
  apply Nat.noConfusion
Nonzero Natural Number from Strict Inequality
Informal description
For any natural numbers $a$ and $b$, if $b < a$ then $a \neq 0$.
Nat.not_eq_zero_of_lt theorem
(h : b < a) : a ≠ 0
Full source
@[deprecated Nat.ne_zero_of_lt (since := "2025-02-06")]
theorem not_eq_zero_of_lt (h : b < a) : a ≠ 0 := Nat.ne_zero_of_lt h
Nonzero Natural Number from Strict Inequality
Informal description
For any natural numbers $a$ and $b$, if $b < a$ then $a \neq 0$.
Nat.pred_lt_of_lt theorem
{n m : Nat} (h : m < n) : pred n < n
Full source
theorem pred_lt_of_lt {n m : Nat} (h : m < n) : pred n < n :=
  pred_lt (Nat.ne_zero_of_lt h)
Predecessor Strictly Less Than Number Under Inequality
Informal description
For any natural numbers $n$ and $m$ such that $m < n$, the predecessor of $n$ is strictly less than $n$, i.e., $\mathrm{pred}(n) < n$.
Nat.sub_one_lt_of_lt theorem
{n m : Nat} (h : m < n) : n - 1 < n
Full source
theorem sub_one_lt_of_lt {n m : Nat} (h : m < n) : n - 1 < n :=
  sub_one_lt (Nat.ne_zero_of_lt h)
Subtraction of One Preserves Strict Inequality in Natural Numbers
Informal description
For any natural numbers $n$ and $m$ such that $m < n$, we have $n - 1 < n$.
Nat.pred_zero theorem
: pred 0 = 0
Full source
protected theorem pred_zero : pred 0 = 0 := rfl
Predecessor of Zero is Zero
Informal description
The predecessor of the natural number $0$ is $0$, i.e., $\mathrm{pred}(0) = 0$.
Nat.pred_succ theorem
(n : Nat) : pred n.succ = n
Full source
protected theorem pred_succ (n : Nat) : pred n.succ = n := rfl
Predecessor-Successor Identity: $\mathrm{pred}(n + 1) = n$
Informal description
For any natural number $n$, the predecessor of the successor of $n$ is equal to $n$, i.e., $\mathrm{pred}(n + 1) = n$.
Nat.zero_sub_one theorem
: 0 - 1 = 0
Full source
@[simp] protected theorem zero_sub_one : 0 - 1 = 0 := rfl
Truncated Subtraction Identity: $0 - 1 = 0$ in $\mathbb{N}$
Informal description
The truncated subtraction of $1$ from $0$ in the natural numbers yields $0$, i.e., $0 - 1 = 0$.
Nat.add_one_sub_one theorem
(n : Nat) : n + 1 - 1 = n
Full source
@[simp] protected theorem add_one_sub_one (n : Nat) : n + 1 - 1 = n := rfl
Subtraction of One from Successor: $(n + 1) - 1 = n$
Informal description
For any natural number $n$, we have $(n + 1) - 1 = n$.
Nat.sub_one_eq_self theorem
{n : Nat} : n - 1 = n ↔ n = 0
Full source
theorem sub_one_eq_self {n : Nat} : n - 1 = n ↔ n = 0 := by cases n <;> simp [ne_add_one]
Subtraction of One Equals Original Number if and only if Zero: $n - 1 = n \leftrightarrow n = 0$
Informal description
For any natural number $n$, the equation $n - 1 = n$ holds if and only if $n = 0$.
Nat.succ_pred theorem
{a : Nat} (h : a ≠ 0) : a.pred.succ = a
Full source
theorem succ_pred {a : Nat} (h : a ≠ 0) : a.pred.succ = a := by
  induction a with
  | zero => contradiction
  | succ => rfl
Successor of Predecessor Equals Original Number for Nonzero Natural Numbers
Informal description
For any nonzero natural number $a$, the successor of the predecessor of $a$ equals $a$, i.e., $\mathrm{succ}(\mathrm{pred}(a)) = a$.
Nat.sub_one_add_one theorem
{a : Nat} (h : a ≠ 0) : a - 1 + 1 = a
Full source
theorem sub_one_add_one {a : Nat} (h : a ≠ 0) : a - 1 + 1 = a := by
  induction a with
  | zero => contradiction
  | succ => rfl
Recovery of Natural Number via Subtraction and Addition: $(a - 1) + 1 = a$ for $a \neq 0$
Informal description
For any nonzero natural number $a$, the sum of $a - 1$ and $1$ equals $a$, i.e., $(a - 1) + 1 = a$.
Nat.succ_pred_eq_of_pos theorem
: ∀ {n}, 0 < n → succ (pred n) = n
Full source
theorem succ_pred_eq_of_pos : ∀ {n}, 0 < n → succ (pred n) = n
  | _+1, _ => rfl
Successor of Predecessor Equals Original Number for Positive Natural Numbers
Informal description
For any natural number $n$ such that $0 < n$, the successor of the predecessor of $n$ equals $n$, i.e., $\mathrm{succ}(\mathrm{pred}(n)) = n$.
Nat.sub_one_add_one_eq_of_pos theorem
: ∀ {n}, 0 < n → (n - 1) + 1 = n
Full source
theorem sub_one_add_one_eq_of_pos : ∀ {n}, 0 < n → (n - 1) + 1 = n
  | _+1, _ => rfl
Recovery of Natural Number via Subtraction and Addition: $(n - 1) + 1 = n$ for $n > 0$
Informal description
For any natural number $n$ such that $0 < n$, it holds that $(n - 1) + 1 = n$.
Nat.pred_eq_sub_one theorem
: pred n = n - 1
Full source
@[simp] theorem pred_eq_sub_one : pred n = n - 1 := rfl
Predecessor as Subtraction by One in Natural Numbers
Informal description
For any natural number $n$, the predecessor of $n$ is equal to $n$ minus one, i.e., $\mathrm{pred}(n) = n - 1$.
Nat.add_sub_self_left theorem
(a b : Nat) : (a + b) - a = b
Full source
theorem add_sub_self_left (a b : Nat) : (a + b) - a = b := by
  induction a with
  | zero => simp
  | succ a ih =>
    rw [Nat.succ_add, Nat.succ_sub_succ]
    apply ih
Subtraction of Summand from Sum in Natural Numbers: $(a + b) - a = b$
Informal description
For any natural numbers $a$ and $b$, the difference between the sum $a + b$ and $a$ equals $b$, i.e., $(a + b) - a = b$.
Nat.add_sub_self_right theorem
(a b : Nat) : (a + b) - b = a
Full source
theorem add_sub_self_right (a b : Nat) : (a + b) - b = a := by
  rw [Nat.add_comm]; apply add_sub_self_left
Subtraction of Summand from Sum in Natural Numbers: $(a + b) - b = a$
Informal description
For any natural numbers $a$ and $b$, the difference between the sum $a + b$ and $b$ equals $a$, i.e., $(a + b) - b = a$.
Nat.sub_le_succ_sub theorem
(a i : Nat) : a - i ≤ a.succ - i
Full source
theorem sub_le_succ_sub (a i : Nat) : a - i ≤ a.succ - i := by
  cases i with
  | zero => apply Nat.le_of_lt; apply Nat.lt_succ_self
  | succ i => rw [Nat.sub_succ, Nat.succ_sub_succ]; apply Nat.pred_le
Monotonicity of Subtraction with Successor: $a - i \leq (a + 1) - i$
Informal description
For any natural numbers $a$ and $i$, the difference $a - i$ is less than or equal to the difference $(a + 1) - i$.
Nat.zero_lt_sub_of_lt theorem
(h : i < a) : 0 < a - i
Full source
theorem zero_lt_sub_of_lt (h : i < a) : 0 < a - i := by
  induction a with
  | zero => contradiction
  | succ a ih =>
    match Nat.eq_or_lt_of_le h with
    | Or.inl h => injection h with h; subst h; rw [Nat.add_sub_self_left]; decide
    | Or.inr h =>
      have : 0 < a - i := ih (Nat.lt_of_succ_lt_succ h)
      exact Nat.lt_of_lt_of_le this (Nat.sub_le_succ_sub _ _)
Positivity of Natural Number Difference Under Strict Inequality
Informal description
For any natural numbers $i$ and $a$, if $i < a$, then the difference $a - i$ is strictly positive, i.e., $0 < a - i$.
Nat.sub_succ_lt_self theorem
(a i : Nat) (h : i < a) : a - (i + 1) < a - i
Full source
theorem sub_succ_lt_self (a i : Nat) (h : i < a) : a - (i + 1) < a - i := by
  rw [Nat.add_succ, Nat.sub_succ]
  apply Nat.pred_lt
  apply Nat.ne_zero_of_lt
  apply Nat.zero_lt_sub_of_lt
  assumption
Strict Monotonicity of Successor Subtraction in Natural Numbers: $a - (i + 1) < a - i$ when $i < a$
Informal description
For any natural numbers $a$ and $i$ such that $i < a$, the difference $a - (i + 1)$ is strictly less than $a - i$.
Nat.sub_ne_zero_of_lt theorem
: {a b : Nat} → a < b → b - a ≠ 0
Full source
theorem sub_ne_zero_of_lt : {a b : Nat} → a < b → b - a ≠ 0
  | 0, 0, h      => absurd h (Nat.lt_irrefl 0)
  | 0, succ b, _ => by simp only [Nat.sub_zero, ne_eq, not_false_eq_true, Nat.succ_ne_zero]
  | succ a, 0, h => absurd h (Nat.not_lt_zero a.succ)
  | succ a, succ b, h => by rw [Nat.succ_sub_succ]; exact sub_ne_zero_of_lt (Nat.lt_of_succ_lt_succ h)
Nonzero Difference of Natural Numbers Under Strict Inequality
Informal description
For any natural numbers $a$ and $b$, if $a < b$, then the difference $b - a$ is not equal to zero.
Nat.add_sub_of_le theorem
{a b : Nat} (h : a ≤ b) : a + (b - a) = b
Full source
theorem add_sub_of_le {a b : Nat} (h : a ≤ b) : a + (b - a) = b := by
  induction a with
  | zero => simp
  | succ a ih =>
    have hne : b - a ≠ 0 := Nat.sub_ne_zero_of_lt h
    have : a ≤ b := Nat.le_of_succ_le h
    rw [sub_succ, Nat.succ_add, ← Nat.add_succ, Nat.succ_pred hne, ih this]
Addition-Subtraction Identity for Natural Numbers: $a + (b - a) = b$ when $a \leq b$
Informal description
For any natural numbers $a$ and $b$ such that $a \leq b$, the sum of $a$ and the difference $b - a$ equals $b$, i.e., $a + (b - a) = b$.
Nat.sub_one_cancel theorem
: ∀ {a b : Nat}, 0 < a → 0 < b → a - 1 = b - 1 → a = b
Full source
theorem sub_one_cancel : ∀ {a b : Nat}, 0 < a → 0 < b → a - 1 = b - 1 → a = b
  | _+1, _+1, _, _ => congrArg _
Cancellation of Subtraction by One for Positive Natural Numbers
Informal description
For any positive natural numbers $a$ and $b$, if $a - 1 = b - 1$, then $a = b$.
Nat.sub_add_cancel theorem
{n m : Nat} (h : m ≤ n) : n - m + m = n
Full source
@[simp] protected theorem sub_add_cancel {n m : Nat} (h : m ≤ n) : n - m + m = n := by
  rw [Nat.add_comm, Nat.add_sub_of_le h]
Subtraction-Additive Cancellation for Natural Numbers: $(n - m) + m = n$ when $m \leq n$
Informal description
For any natural numbers $n$ and $m$ such that $m \leq n$, the sum of the difference $n - m$ and $m$ equals $n$, i.e., $(n - m) + m = n$.
Nat.add_sub_add_right theorem
(n k m : Nat) : (n + k) - (m + k) = n - m
Full source
protected theorem add_sub_add_right (n k m : Nat) : (n + k) - (m + k) = n - m := by
  induction k with
  | zero => simp
  | succ k ih => simp [← Nat.add_assoc, succ_sub_succ_eq_sub, ih]
Right Addition Cancels in Subtraction: $(n + k) - (m + k) = n - m$
Informal description
For any natural numbers $n$, $k$, and $m$, the difference $(n + k) - (m + k)$ equals $n - m$.
Nat.add_sub_add_left theorem
(k n m : Nat) : (k + n) - (k + m) = n - m
Full source
protected theorem add_sub_add_left (k n m : Nat) : (k + n) - (k + m) = n - m := by
  rw [Nat.add_comm k n, Nat.add_comm k m, Nat.add_sub_add_right]
Left Addition Cancels in Subtraction: $(k + n) - (k + m) = n - m$
Informal description
For any natural numbers $k$, $n$, and $m$, the difference $(k + n) - (k + m)$ equals $n - m$.
Nat.add_sub_cancel theorem
(n m : Nat) : n + m - m = n
Full source
@[simp] protected theorem add_sub_cancel (n m : Nat) : n + m - m = n :=
  suffices n + m - (0 + m) = n by rw [Nat.zero_add] at this; assumption
  by rw [Nat.add_sub_add_right, Nat.sub_zero]
Right Cancellation in Natural Number Subtraction: $(n + m) - m = n$
Informal description
For any natural numbers $n$ and $m$, the difference $(n + m) - m$ equals $n$.
Nat.add_sub_cancel_left theorem
(n m : Nat) : n + m - n = m
Full source
protected theorem add_sub_cancel_left (n m : Nat) : n + m - n = m :=
  show n + m - (n + 0) = m from
  by rw [Nat.add_sub_add_left, Nat.sub_zero]
Left Cancellation in Natural Number Subtraction: $(n + m) - n = m$
Informal description
For any natural numbers $n$ and $m$, the difference $(n + m) - n$ equals $m$.
Nat.add_sub_assoc theorem
{m k : Nat} (h : k ≤ m) (n : Nat) : n + m - k = n + (m - k)
Full source
protected theorem add_sub_assoc {m k : Nat} (h : k ≤ m) (n : Nat) : n + m - k = n + (m - k) := by
 cases Nat.le.dest h
 rename_i l hl
 rw [← hl, Nat.add_sub_cancel_left, Nat.add_comm k, ← Nat.add_assoc, Nat.add_sub_cancel]
Associativity of Addition and Subtraction in Natural Numbers under Order Condition
Informal description
For any natural numbers $m$, $k$, and $n$, if $k \leq m$, then $(n + m) - k = n + (m - k)$.
Nat.sub_eq_of_eq_add theorem
{a b c : Nat} (h : a = c + b) : a - b = c
Full source
protected theorem sub_eq_of_eq_add {a b c : Nat} (h : a = c + b) : a - b = c := by
  rw [h, Nat.add_sub_cancel]
Subtraction from Sum: $a = c + b \Rightarrow a - b = c$
Informal description
For any natural numbers $a$, $b$, and $c$, if $a = c + b$, then $a - b = c$.
Nat.le_add_of_sub_le theorem
{a b c : Nat} (h : a - b ≤ c) : a ≤ c + b
Full source
theorem le_add_of_sub_le {a b c : Nat} (h : a - b ≤ c) : a ≤ c + b := by
  match le.dest h, Nat.le_total a b with
  | _, Or.inl hle =>
    exact Nat.le_trans hle (Nat.le_add_left ..)
  | ⟨d, hd⟩, Or.inr hge =>
    apply @le.intro _ _ d
    rw [Nat.add_comm, ← Nat.add_sub_assoc hge] at hd
    have hd := Nat.eq_add_of_sub_eq (Nat.le_trans hge (Nat.le_add_left ..)) hd
    rw [Nat.add_comm, hd]
Inequality from Subtraction: $a - b \leq c \Rightarrow a \leq c + b$
Informal description
For any natural numbers $a$, $b$, and $c$, if $a - b \leq c$, then $a \leq c + b$.
Nat.sub_lt_sub_left theorem
: ∀ {k m n : Nat}, k < m → k < n → m - n < m - k
Full source
protected theorem sub_lt_sub_left : ∀ {k m n : Nat}, k < m → k < n → m - n < m - k
  | 0, m+1, n+1, _, _ => by rw [Nat.add_sub_add_right]; exact lt_succ_of_le (Nat.sub_le _ _)
  | k+1, m+1, n+1, h1, h2 => by
    rw [Nat.add_sub_add_right, Nat.add_sub_add_right]
    exact Nat.sub_lt_sub_left (Nat.lt_of_succ_lt_succ h1) (Nat.lt_of_succ_lt_succ h2)
Left Subtraction Preserves Strict Inequality in Natural Numbers
Informal description
For any natural numbers $k$, $m$, and $n$, if $k < m$ and $k < n$, then $m - n < m - k$.
Nat.zero_sub theorem
(n : Nat) : 0 - n = 0
Full source
@[simp] protected theorem zero_sub (n : Nat) : 0 - n = 0 := by
  induction n with
  | zero => rfl
  | succ n ih => simp only [ih, Nat.sub_succ]; decide
Zero Minus Any Natural Number is Zero
Informal description
For any natural number $n$, the truncated subtraction $0 - n$ equals $0$.
Nat.sub_lt_sub_right theorem
: ∀ {a b c : Nat}, c ≤ a → a < b → a - c < b - c
Full source
protected theorem sub_lt_sub_right : ∀ {a b c : Nat}, c ≤ a → a < b → a - c < b - c
  | 0, _, _, hle, h => by
    rw [Nat.eq_zero_of_le_zero hle, Nat.sub_zero, Nat.sub_zero]
    exact h
  | _, _, 0, _, h => by
    rw [Nat.sub_zero, Nat.sub_zero]
    exact h
  | _+1, _+1, _+1, hle, h => by
    rw [Nat.add_sub_add_right, Nat.add_sub_add_right]
    exact Nat.sub_lt_sub_right (le_of_succ_le_succ hle) (lt_of_succ_lt_succ h)
Right Subtraction Preserves Strict Inequality in Natural Numbers
Informal description
For any natural numbers $a$, $b$, and $c$ such that $c \leq a$ and $a < b$, it holds that $a - c < b - c$.
Nat.sub_self_add theorem
(n m : Nat) : n - (n + m) = 0
Full source
protected theorem sub_self_add (n m : Nat) : n - (n + m) = 0 := by
  show (n + 0) - (n + m) = 0
  rw [Nat.add_sub_add_left, Nat.zero_sub]
Self-Addition Cancels in Subtraction: $n - (n + m) = 0$
Informal description
For any natural numbers $n$ and $m$, the truncated subtraction $n - (n + m)$ equals $0$.
Nat.sub_eq_zero_of_le theorem
{n m : Nat} (h : n ≤ m) : n - m = 0
Full source
protected theorem sub_eq_zero_of_le {n m : Nat} (h : n ≤ m) : n - m = 0 := by
  match le.dest h with
  | ⟨k, hk⟩ => rw [← hk, Nat.sub_self_add]
Truncated Subtraction Yields Zero When First Argument is Less Than or Equal to Second
Informal description
For any natural numbers $n$ and $m$ such that $n \leq m$, the truncated subtraction $n - m$ equals $0$.
Nat.sub_le_of_le_add theorem
{a b c : Nat} (h : a ≤ c + b) : a - b ≤ c
Full source
theorem sub_le_of_le_add {a b c : Nat} (h : a ≤ c + b) : a - b ≤ c := by
  match le.dest h, Nat.le_total a b with
  | _, Or.inl hle =>
    rw [Nat.sub_eq_zero_of_le hle]
    apply Nat.zero_le
  | ⟨d, hd⟩, Or.inr hge =>
    apply @le.intro _ _ d
    have hd := Nat.sub_eq_of_eq_add hd
    rw [Nat.add_comm, ← Nat.add_sub_assoc hge, Nat.add_comm]
    exact hd
Subtraction Preserves Order Under Addition: $a \leq c + b \implies a - b \leq c$
Informal description
For any natural numbers $a$, $b$, and $c$, if $a \leq c + b$, then $a - b \leq c$.
Nat.add_le_of_le_sub theorem
{a b c : Nat} (hle : b ≤ c) (h : a ≤ c - b) : a + b ≤ c
Full source
theorem add_le_of_le_sub {a b c : Nat} (hle : b ≤ c) (h : a ≤ c - b) : a + b ≤ c := by
  match le.dest h with
  | ⟨d, hd⟩ =>
    apply @le.intro _ _ d
    rw [Nat.eq_add_of_sub_eq hle hd.symm]
    simp [Nat.add_comm, Nat.add_assoc, Nat.add_left_comm]
Addition Preserves Order Under Subtraction: $b \leq c \land a \leq c - b \implies a + b \leq c$
Informal description
For any natural numbers $a$, $b$, and $c$ such that $b \leq c$ and $a \leq c - b$, it follows that $a + b \leq c$.
Nat.le_sub_of_add_le theorem
{a b c : Nat} (h : a + b ≤ c) : a ≤ c - b
Full source
theorem le_sub_of_add_le {a b c : Nat} (h : a + b ≤ c) : a ≤ c - b := by
  match le.dest h with
  | ⟨d, hd⟩ =>
    apply @le.intro _ _ d
    have hd : a + d + b = c := by simp [← hd, Nat.add_comm, Nat.add_assoc, Nat.add_left_comm]
    have hd := Nat.sub_eq_of_eq_add hd.symm
    exact hd.symm
Subtraction Preserves Order Under Addition: $a + b \leq c \implies a \leq c - b$
Informal description
For any natural numbers $a$, $b$, and $c$, if $a + b \leq c$, then $a \leq c - b$.
Nat.add_lt_of_lt_sub theorem
{a b c : Nat} (h : a < c - b) : a + b < c
Full source
theorem add_lt_of_lt_sub {a b c : Nat} (h : a < c - b) : a + b < c := by
  have hle : b ≤ c := by
    apply Nat.ge_of_not_lt
    intro hgt
    apply Nat.not_lt_zero a
    rw [Nat.sub_eq_zero_of_le (Nat.le_of_lt hgt)] at h
    exact h
  have : a.succ + b ≤ c := add_le_of_le_sub hle h
  simp [Nat.succ_add] at this
  exact this
Addition Preserves Strict Inequality Under Subtraction: $a < c - b \implies a + b < c$
Informal description
For any natural numbers $a$, $b$, and $c$, if $a < c - b$, then $a + b < c$.
Nat.lt_sub_of_add_lt theorem
{a b c : Nat} (h : a + b < c) : a < c - b
Full source
theorem lt_sub_of_add_lt {a b c : Nat} (h : a + b < c) : a < c - b :=
  have : a.succ + b ≤ c := by simp [Nat.succ_add]; exact h
  le_sub_of_add_le this
Subtraction Preserves Strict Order Under Addition: $a + b < c \implies a < c - b$
Informal description
For any natural numbers $a$, $b$, and $c$, if $a + b < c$, then $a < c - b$.
Nat.sub.elim theorem
{motive : Nat → Prop} (x y : Nat) (h₁ : y ≤ x → (k : Nat) → x = y + k → motive k) (h₂ : x < y → motive 0) : motive (x - y)
Full source
theorem sub.elim {motive : Nat → Prop}
    (x y : Nat)
    (h₁ : y ≤ x → (k : Nat) → x = y + k → motive k)
    (h₂ : x < y → motive 0)
    : motive (x - y) := by
  cases Nat.lt_or_ge x y with
  | inl hlt => rw [Nat.sub_eq_zero_of_le (Nat.le_of_lt hlt)]; exact h₂ hlt
  | inr hle => exact h₁ hle (x - y) (Nat.add_sub_of_le hle).symm
Elimination Principle for Natural Number Subtraction
Informal description
For any natural numbers $x$ and $y$, and any property `motive` on natural numbers, if: 1. For all $k \in \mathbb{N}$ such that $y \leq x$ and $x = y + k$, the property `motive` holds for $k$; and 2. When $x < y$, the property `motive` holds for $0$; then the property `motive` holds for $x - y$.
Nat.succ_sub theorem
{m n : Nat} (h : n ≤ m) : succ m - n = succ (m - n)
Full source
theorem succ_sub {m n : Nat} (h : n ≤ m) : succ m - n = succ (m - n) := by
  let ⟨k, hk⟩ := Nat.le.dest h
  rw [← hk, Nat.add_sub_cancel_left, ← add_succ, Nat.add_sub_cancel_left]
Successor Subtraction Identity: $(m + 1) - n = (m - n) + 1$ under $n \leq m$
Informal description
For any natural numbers $m$ and $n$ such that $n \leq m$, the successor of $m$ minus $n$ equals the successor of $m$ minus $n$, i.e., $(m + 1) - n = (m - n) + 1$.
Nat.sub_pos_of_lt theorem
(h : m < n) : 0 < n - m
Full source
protected theorem sub_pos_of_lt (h : m < n) : 0 < n - m :=
  Nat.pos_iff_ne_zero.2 (Nat.sub_ne_zero_of_lt h)
Positivity of Difference under Strict Inequality in Natural Numbers
Informal description
For any natural numbers $m$ and $n$ such that $m < n$, the difference $n - m$ is positive, i.e., $0 < n - m$.
Nat.sub_sub theorem
(n m k : Nat) : n - m - k = n - (m + k)
Full source
protected theorem sub_sub (n m k : Nat) : n - m - k = n - (m + k) := by
  induction k with
  | zero => simp
  | succ k ih => rw [Nat.add_succ, Nat.sub_succ, Nat.add_succ, Nat.sub_succ, ih]
Subtraction Associativity: $n - m - k = n - (m + k)$
Informal description
For any natural numbers $n$, $m$, and $k$, the difference $n - m - k$ equals $n - (m + k)$.
Nat.sub_le_sub_left theorem
(h : n ≤ m) (k : Nat) : k - m ≤ k - n
Full source
protected theorem sub_le_sub_left (h : n ≤ m) (k : Nat) : k - m ≤ k - n :=
  match m, le.dest h with
  | _, ⟨a, rfl⟩ => by rw [← Nat.sub_sub]; apply sub_le
Left Subtraction Reverses Order in Natural Numbers
Informal description
For any natural numbers $n$ and $m$ such that $n \leq m$, and for any natural number $k$, the truncated subtraction satisfies $k - m \leq k - n$.
Nat.sub_le_sub_right theorem
{n m : Nat} (h : n ≤ m) : ∀ k, n - k ≤ m - k
Full source
protected theorem sub_le_sub_right {n m : Nat} (h : n ≤ m) : ∀ k, n - k ≤ m - k
  | 0   => h
  | z+1 => pred_le_pred (Nat.sub_le_sub_right h z)
Right Subtraction Preserves Order in Natural Numbers
Informal description
For any natural numbers $n$ and $m$ such that $n \leq m$, and for any natural number $k$, the truncated subtraction satisfies $n - k \leq m - k$.
Nat.sub_le_add_right_sub theorem
(a i j : Nat) : a - i ≤ a + j - i
Full source
protected theorem sub_le_add_right_sub (a i j : Nat) : a - i ≤ a + j - i :=
  Nat.sub_le_sub_right (Nat.le_add_right ..) ..
Right Addition Preserves Truncated Subtraction Inequality in Natural Numbers
Informal description
For any natural numbers $a$, $i$, and $j$, the truncated subtraction satisfies $a - i \leq (a + j) - i$.
Nat.lt_of_sub_ne_zero theorem
(h : n - m ≠ 0) : m < n
Full source
protected theorem lt_of_sub_ne_zero (h : n - m ≠ 0) : m < n :=
  Nat.not_le.1 (mt Nat.sub_eq_zero_of_le h)
Nonzero Truncated Subtraction Implies Strict Inequality in Natural Numbers
Informal description
For any natural numbers $n$ and $m$, if the truncated subtraction $n - m$ is not zero, then $m < n$.
Nat.lt_of_sub_pos theorem
(h : 0 < n - m) : m < n
Full source
protected theorem lt_of_sub_pos (h : 0 < n - m) : m < n :=
  Nat.lt_of_sub_ne_zero (Nat.pos_iff_ne_zero.1 h)
Positivity of Difference Implies Strict Inequality in Natural Numbers
Informal description
For any natural numbers $n$ and $m$, if the truncated subtraction $n - m$ is positive (i.e., $0 < n - m$), then $m < n$.
Nat.lt_of_sub_eq_succ theorem
(h : m - n = succ l) : n < m
Full source
protected theorem lt_of_sub_eq_succ (h : m - n = succ l) : n < m :=
  Nat.lt_of_sub_pos (h ▸ Nat.zero_lt_succ _)
Successor Difference Implies Strict Inequality in Natural Numbers
Informal description
For any natural numbers $m$, $n$, and $l$, if the truncated subtraction $m - n$ equals the successor of $l$ (i.e., $m - n = l + 1$), then $n < m$.
Nat.lt_of_sub_eq_sub_one theorem
(h : m - n = l + 1) : n < m
Full source
protected theorem lt_of_sub_eq_sub_one (h : m - n = l + 1) : n < m :=
  Nat.lt_of_sub_pos (h ▸ Nat.zero_lt_succ _)
Difference Equals Successor Implies Strict Inequality in Natural Numbers
Informal description
For any natural numbers $m$, $n$, and $l$, if the difference $m - n$ equals $l + 1$, then $n$ is strictly less than $m$.
Nat.sub_lt_left_of_lt_add theorem
{n k m : Nat} (H : n ≤ k) (h : k < n + m) : k - n < m
Full source
protected theorem sub_lt_left_of_lt_add {n k m : Nat} (H : n ≤ k) (h : k < n + m) : k - n < m := by
  have := Nat.sub_le_sub_right (succ_le_of_lt h) n
  rwa [Nat.add_sub_cancel_left, Nat.succ_sub H] at this
Left Subtraction Inequality: $k - n < m$ under $n \leq k$ and $k < n + m$
Informal description
For any natural numbers $n$, $k$, and $m$, if $n \leq k$ and $k < n + m$, then $k - n < m$.
Nat.sub_lt_right_of_lt_add theorem
{n k m : Nat} (H : n ≤ k) (h : k < m + n) : k - n < m
Full source
protected theorem sub_lt_right_of_lt_add {n k m : Nat} (H : n ≤ k) (h : k < m + n) : k - n < m :=
  Nat.sub_lt_left_of_lt_add H (Nat.add_comm .. ▸ h)
Right Subtraction Inequality: $k - n < m$ under $n \leq k$ and $k < m + n$
Informal description
For any natural numbers $n$, $k$, and $m$, if $n \leq k$ and $k < m + n$, then $k - n < m$.
Nat.le_of_sub_eq_zero theorem
: ∀ {n m}, n - m = 0 → n ≤ m
Full source
protected theorem le_of_sub_eq_zero : ∀ {n m}, n - m = 0 → n ≤ m
  | 0, _, _ => Nat.zero_le ..
  | _+1, _+1, h => Nat.succ_le_succ <| Nat.le_of_sub_eq_zero (Nat.succ_sub_succ .. ▸ h)
Implication from Zero Subtraction to Inequality: $n - m = 0 \implies n \leq m$
Informal description
For any natural numbers $n$ and $m$, if the truncated subtraction $n - m$ equals zero, then $n \leq m$.
Nat.le_of_sub_le_sub_right theorem
: ∀ {n m k : Nat}, k ≤ m → n - k ≤ m - k → n ≤ m
Full source
protected theorem le_of_sub_le_sub_right : ∀ {n m k : Nat}, k ≤ m → n - k ≤ m - k → n ≤ m
  | 0, _, _, _, _ => Nat.zero_le ..
  | _+1, _, 0, _, h₁ => h₁
  | _+1, _+1, _+1, h₀, h₁ => by
    simp only [Nat.succ_sub_succ] at h₁
    exact succ_le_succ <| Nat.le_of_sub_le_sub_right (le_of_succ_le_succ h₀) h₁
Order Preservation under Subtraction: $k \leq m \land n - k \leq m - k \to n \leq m$
Informal description
For any natural numbers $n$, $m$, and $k$ such that $k \leq m$, if the difference $n - k$ is less than or equal to the difference $m - k$, then $n \leq m$.
Nat.sub_le_sub_iff_right theorem
{n : Nat} (h : k ≤ m) : n - k ≤ m - k ↔ n ≤ m
Full source
protected theorem sub_le_sub_iff_right {n : Nat} (h : k ≤ m) : n - k ≤ m - k ↔ n ≤ m :=
  ⟨Nat.le_of_sub_le_sub_right h, fun h => Nat.sub_le_sub_right h _⟩
Right Subtraction Preserves Order in Natural Numbers: $n - k \leq m - k \leftrightarrow n \leq m$
Informal description
For any natural numbers $n$, $m$, and $k$ such that $k \leq m$, the inequality $n - k \leq m - k$ holds if and only if $n \leq m$.
Nat.sub_eq_iff_eq_add theorem
{c : Nat} (h : b ≤ a) : a - b = c ↔ a = c + b
Full source
protected theorem sub_eq_iff_eq_add {c : Nat} (h : b ≤ a) : a - b = c ↔ a = c + b :=
  ⟨fun | rfl => by rw [Nat.sub_add_cancel h], fun heq => by rw [heq, Nat.add_sub_cancel]⟩
Subtraction-Additive Equivalence: $a - b = c \leftrightarrow a = c + b$ when $b \leq a$
Informal description
For any natural numbers $a$, $b$, and $c$ such that $b \leq a$, the equation $a - b = c$ holds if and only if $a = c + b$.
Nat.sub_eq_iff_eq_add' theorem
{c : Nat} (h : b ≤ a) : a - b = c ↔ a = b + c
Full source
protected theorem sub_eq_iff_eq_add' {c : Nat} (h : b ≤ a) : a - b = c ↔ a = b + c := by
  rw [Nat.add_comm, Nat.sub_eq_iff_eq_add h]
Subtraction-Additive Equivalence (variant): $a - b = c \leftrightarrow a = b + c$ when $b \leq a$
Informal description
For any natural numbers $a$, $b$, and $c$ such that $b \leq a$, the equation $a - b = c$ holds if and only if $a = b + c$.
Nat.pred_mul theorem
(n m : Nat) : pred n * m = n * m - m
Full source
theorem pred_mul (n m : Nat) : pred n * m = n * m - m := by
  cases n with
  | zero   => simp
  | succ n => rw [Nat.pred_succ, succ_mul, Nat.add_sub_cancel]
Predecessor Multiplication Identity: $\mathrm{pred}(n) \cdot m = n \cdot m - m$
Informal description
For any natural numbers $n$ and $m$, the product of the predecessor of $n$ and $m$ equals the difference between the product of $n$ and $m$ and $m$, i.e., $\mathrm{pred}(n) \cdot m = n \cdot m - m$.
Nat.sub_one_mul theorem
(n m : Nat) : (n - 1) * m = n * m - m
Full source
protected theorem sub_one_mul  (n m : Nat) : (n - 1) * m = n * m - m := by
  cases n with
  | zero   => simp
  | succ n =>
    rw [Nat.add_sub_cancel, add_one_mul, Nat.add_sub_cancel]
Multiplication by Predecessor Identity: $(n - 1) \cdot m = n \cdot m - m$
Informal description
For any natural numbers $n$ and $m$, the product of $n-1$ and $m$ equals the difference between the product of $n$ and $m$ and $m$, i.e., $(n - 1) \cdot m = n \cdot m - m$.
Nat.mul_pred theorem
(n m : Nat) : n * pred m = n * m - n
Full source
theorem mul_pred (n m : Nat) : n * pred m = n * m - n := by
  rw [Nat.mul_comm, pred_mul, Nat.mul_comm]
Multiplication-Predecessor Identity: $n \cdot \mathrm{pred}(m) = n \cdot m - n$
Informal description
For any natural numbers $n$ and $m$, the product of $n$ and the predecessor of $m$ equals the difference between the product of $n$ and $m$ and $n$, i.e., $n \cdot \mathrm{pred}(m) = n \cdot m - n$.
Nat.mul_sub_one theorem
(n m : Nat) : n * (m - 1) = n * m - n
Full source
theorem mul_sub_one (n m : Nat) : n * (m - 1) = n * m - n := by
  rw [Nat.mul_comm, Nat.sub_one_mul , Nat.mul_comm]
Multiplication by Predecessor Identity: $n \cdot (m - 1) = n \cdot m - n$
Informal description
For any natural numbers $n$ and $m$, the product of $n$ and $m-1$ equals the difference between the product of $n$ and $m$ and $n$, i.e., $n \cdot (m - 1) = n \cdot m - n$.
Nat.mul_sub_right_distrib theorem
(n m k : Nat) : (n - m) * k = n * k - m * k
Full source
protected theorem mul_sub_right_distrib (n m k : Nat) : (n - m) * k = n * k - m * k := by
  induction m with
  | zero => simp
  | succ m ih => rw [Nat.sub_succ, Nat.pred_mul, ih, succ_mul, Nat.sub_sub]; done
Right Distributivity of Multiplication over Subtraction in Natural Numbers: $(n - m) \cdot k = n \cdot k - m \cdot k$
Informal description
For any natural numbers $n$, $m$, and $k$, the product of the difference $n - m$ and $k$ equals the difference of the products $n \cdot k$ and $m \cdot k$, i.e., $(n - m) \cdot k = n \cdot k - m \cdot k$.
Nat.sub_mul theorem
(n m k : Nat) : (n - m) * k = n * k - m * k
Full source
protected theorem sub_mul (n m k : Nat) : (n - m) * k = n * k - m * k :=
  Nat.mul_sub_right_distrib n m k
Right Distributivity of Multiplication over Subtraction in Natural Numbers: $(n - m) \cdot k = n \cdot k - m \cdot k$
Informal description
For any natural numbers $n$, $m$, and $k$, the product of the difference $n - m$ and $k$ equals the difference of the products $n \cdot k$ and $m \cdot k$, i.e., $(n - m) \cdot k = n \cdot k - m \cdot k$.
Nat.mul_sub_left_distrib theorem
(n m k : Nat) : n * (m - k) = n * m - n * k
Full source
protected theorem mul_sub_left_distrib (n m k : Nat) : n * (m - k) = n * m - n * k := by
  rw [Nat.mul_comm, Nat.mul_sub_right_distrib, Nat.mul_comm m n, Nat.mul_comm n k]
Left Distributivity of Multiplication over Subtraction in Natural Numbers: $n \cdot (m - k) = n \cdot m - n \cdot k$
Informal description
For any natural numbers $n$, $m$, and $k$, the product of $n$ and the difference $m - k$ equals the difference of the products $n \cdot m$ and $n \cdot k$, i.e., $n \cdot (m - k) = n \cdot m - n \cdot k$.
Nat.mul_sub theorem
(n m k : Nat) : n * (m - k) = n * m - n * k
Full source
protected theorem mul_sub (n m k : Nat) : n * (m - k) = n * m - n * k :=
  Nat.mul_sub_left_distrib n m k
Left Distributivity of Multiplication over Subtraction in Natural Numbers: $n \cdot (m - k) = n \cdot m - n \cdot k$
Informal description
For any natural numbers $n$, $m$, and $k$, the product of $n$ and the difference $m - k$ equals the difference of the products $n \cdot m$ and $n \cdot k$, i.e., $n \cdot (m - k) = n \cdot m - n \cdot k$.
Nat.not_le_eq theorem
(a b : Nat) : (¬(a ≤ b)) = (b + 1 ≤ a)
Full source
theorem not_le_eq (a b : Nat) : (¬ (a ≤ b)) = (b + 1 ≤ a) :=
  Eq.propIntro Nat.gt_of_not_le Nat.not_le_of_gt
Negation of Natural Number Order Relation: $\neg(a \leq b) \leftrightarrow b+1 \leq a$
Informal description
For any natural numbers $a$ and $b$, the negation of $a \leq b$ is equivalent to $b + 1 \leq a$. In other words, $\neg(a \leq b) \leftrightarrow (b + 1 \leq a)$.
Nat.not_ge_eq theorem
(a b : Nat) : (¬(a ≥ b)) = (a + 1 ≤ b)
Full source
theorem not_ge_eq (a b : Nat) : (¬ (a ≥ b)) = (a + 1 ≤ b) :=
  not_le_eq b a
Negation of Natural Number Order Relation: $\neg(a \geq b) \leftrightarrow a+1 \leq b$
Informal description
For any natural numbers $a$ and $b$, the negation of $a \geq b$ is equivalent to $a + 1 \leq b$, i.e., $\neg(a \geq b) \leftrightarrow (a + 1 \leq b)$.
Nat.not_lt_eq theorem
(a b : Nat) : (¬(a < b)) = (b ≤ a)
Full source
theorem not_lt_eq (a b : Nat) : (¬ (a < b)) = (b ≤ a) :=
  Eq.propIntro Nat.le_of_not_lt Nat.not_lt_of_le
Negation of Strict Inequality in Natural Numbers: $\neg(a < b) \leftrightarrow b \leq a$
Informal description
For any natural numbers $a$ and $b$, the negation of $a < b$ is equivalent to $b \leq a$. In other words, $\neg(a < b) \leftrightarrow b \leq a$.
Nat.not_gt_eq theorem
(a b : Nat) : (¬(a > b)) = (a ≤ b)
Full source
theorem not_gt_eq (a b : Nat) : (¬ (a > b)) = (a ≤ b) :=
  not_lt_eq b a
Negation of Greater-Than Relation in Natural Numbers: $\neg(a > b) \leftrightarrow a \leq b$
Informal description
For any natural numbers $a$ and $b$, the negation of $a > b$ is equivalent to $a \leq b$, i.e., $\neg(a > b) \leftrightarrow a \leq b$.
Nat.repeat_eq_repeatTR theorem
: @repeat = @repeatTR
Full source
@[csimp] theorem repeat_eq_repeatTR : @repeat = @repeatTR :=
  funext fun α => funext fun f => funext fun n => funext fun init =>
  let rec go : ∀ m n, repeatTR.loop f m (repeat f n init) = repeat f (m + n) init
    | 0,      n => by simp [repeatTR.loop]
    | succ m, n => by rw [repeatTR.loop, succ_add]; exact go m (succ n)
  (go n 0).symm
Equivalence of Standard and Tail-Recursive Function Iteration
Informal description
For any type $\alpha$ and function $f : \alpha \to \alpha$, the standard iterated function application `Nat.repeat` is equal to its tail-recursive version `Nat.repeatTR`. That is, for all natural numbers $n$ and starting values $a \in \alpha$, we have $\text{repeat}\,f\,n\,a = \text{repeatTR}\,f\,n\,a$.