doc-next-gen

Init.Data.Nat.Div.Basic

Module docstring

{}

Nat.instDvd instance
: Dvd Nat
Full source
/--
Divisibility of natural numbers. `a ∣ b` (typed as `\|`) says that
there is some `c` such that `b = a * c`.
-/
instance : Dvd Nat where
  dvd a b := Exists (fun c => b = a * c)
Divisibility Relation on Natural Numbers
Informal description
The natural numbers $\mathbb{N}$ are equipped with a divisibility relation $\mid$, where $a \mid b$ means there exists some $c \in \mathbb{N}$ such that $b = a \cdot c$.
Nat.div_rec_lemma theorem
{x y : Nat} : 0 < y ∧ y ≤ x → x - y < x
Full source
theorem div_rec_lemma {x y : Nat} : 0 < y ∧ y ≤ x → x - y < x :=
  fun ⟨ypos, ylex⟩ => sub_lt (Nat.lt_of_lt_of_le ypos ylex) ypos
Truncated Subtraction Decreases Natural Numbers Under Divisibility Conditions
Informal description
For any natural numbers $x$ and $y$, if $0 < y$ and $y \leq x$, then $x - y < x$.
Nat.div_rec_fuel_lemma theorem
{x y fuel : Nat} (hy : 0 < y) (hle : y ≤ x) (hfuel : x < fuel + 1) : x - y < fuel
Full source
theorem div_rec_fuel_lemma {x y fuel : Nat} (hy : 0 < y) (hle : y ≤ x) (hfuel : x < fuel + 1) :
    x - y < fuel :=
  Nat.lt_of_lt_of_le (div_rec_lemma ⟨hy, hle⟩) (Nat.le_of_lt_succ hfuel)
Truncated Subtraction Bounds Under Fuel Condition
Informal description
For any natural numbers $x$, $y$, and $\mathit{fuel}$, if $0 < y$, $y \leq x$, and $x < \mathit{fuel} + 1$, then $x - y < \mathit{fuel}$.
Nat.div definition
(x y : @& Nat) : Nat
Full source
/--
Division of natural numbers, discarding the remainder. Division by `0` returns `0`. Usually accessed
via the `/` operator.

This operation is sometimes called “floor division.”

This function is overridden at runtime with an efficient implementation. This definition is
the logical model.

Examples:
 * `21 / 3 = 7`
 * `21 / 5 = 4`
 * `0 / 22 = 0`
 * `5 / 0 = 0`
-/
@[extern "lean_nat_div", irreducible]
protected def div (x y : @& Nat) : Nat :=
  if hy : 0 < y then
    let rec
      go (fuel : Nat) (x : Nat) (hfuel : x < fuel) : Nat :=
      match fuel with
      | 0 => by contradiction
      | succ fuel =>
        if h : y ≤ x then
          go fuel (x - y) (div_rec_fuel_lemma hy h hfuel) + 1
        else
          0
      termination_by structural fuel
    go (x + 1) x (Nat.lt_succ_self _)
  else
    0
Floor division of natural numbers
Informal description
The function $\mathrm{div} : \mathbb{N} \times \mathbb{N} \to \mathbb{N}$ computes the floor division of natural numbers $x$ and $y$, defined as the largest natural number $q$ such that $q \cdot y \leq x$. Division by zero is defined to return zero. More precisely: - If $y > 0$, then $x / y$ is computed recursively by repeatedly subtracting $y$ from $x$ until $y > x$, counting the number of successful subtractions. - If $y = 0$, then $x / y = 0$. Examples: - $21 / 3 = 7$ - $21 / 5 = 4$ - $0 / 22 = 0$ - $5 / 0 = 0$
Nat.instDiv instance
: Div Nat
Full source
instance instDiv : Div Nat := ⟨Nat.div⟩
Floor Division on Natural Numbers
Informal description
The natural numbers $\mathbb{N}$ are equipped with a division operation $/$ defined as floor division, where for any $x, y \in \mathbb{N}$, $x / y$ is the largest natural number $q$ such that $q \cdot y \leq x$, with division by zero defined to return zero.
Nat.div_eq theorem
(x y : Nat) : x / y = if 0 < y ∧ y ≤ x then (x - y) / y + 1 else 0
Full source
theorem div_eq (x y : Nat) : x / y = if 0 < y ∧ y ≤ x then (x - y) / y + 1 else 0 := by
  show Nat.div _ _ = ite _ (Nat.div _ _ + 1) _
  unfold Nat.div
  split
  next =>
    rw [Nat.div.go]
    split
    next =>
      simp only [and_self, ↓reduceIte, *]
      congr 1
      apply div.go.fuel_congr
    next =>
      simp only [and_false, ↓reduceIte, *]
  next =>
    simp only [false_and, ↓reduceIte, *]
Recursive Definition of Natural Number Division: $x / y = \text{if } 0 < y \land y \leq x \text{ then } (x - y)/y + 1 \text{ else } 0$
Informal description
For any natural numbers $x$ and $y$, the floor division $x / y$ satisfies: \[ x / y = \begin{cases} (x - y) / y + 1 & \text{if } 0 < y \text{ and } y \leq x, \\ 0 & \text{otherwise.} \end{cases} \]
Nat.div.inductionOn definition
{motive : Nat → Nat → Sort u} (x y : Nat) (ind : ∀ x y, 0 < y ∧ y ≤ x → motive (x - y) y → motive x y) (base : ∀ x y, ¬(0 < y ∧ y ≤ x) → motive x y) : motive x y
Full source
/--
An induction principle customized for reasoning about the recursion pattern of natural number
division by iterated subtraction.
-/
def div.inductionOn.{u}
      {motive : NatNat → Sort u}
      (x y : Nat)
      (ind  : ∀ x y, 0 < y ∧ y ≤ x → motive (x - y) y → motive x y)
      (base : ∀ x y, ¬(0 < y ∧ y ≤ x) → motive x y)
      : motive x y :=
  if h : 0 < y ∧ y ≤ x then
    ind x y h (inductionOn (x - y) y ind base)
  else
    base x y h
decreasing_by apply div_rec_lemma; assumption
Induction principle for natural number division
Informal description
Given a motive `motive : ℕ → ℕ → Sort u` and natural numbers `x` and `y`, the induction principle `Nat.div.inductionOn` allows proving `motive x y` by: 1. An inductive step `ind` that assumes `0 < y ∧ y ≤ x` and a proof of `motive (x - y) y` to construct a proof of `motive x y`. 2. A base case `base` that assumes `¬(0 < y ∧ y ≤ x)` to construct a proof of `motive x y`. This principle is specifically designed for reasoning about natural number division via iterated subtraction.
Nat.div_le_self theorem
(n k : Nat) : n / k ≤ n
Full source
theorem div_le_self (n k : Nat) : n / k ≤ n := by
  induction n using Nat.strongRecOn with
  | ind n ih =>
    rw [div_eq]
    -- Note: manual split to avoid Classical.em which is not yet defined
    cases (inferInstance : Decidable (0 < k ∧ k ≤ n)) with
    | isFalse h => simp [h]
    | isTrue h =>
      suffices (n - k) / k + 1 ≤ n by simp [h, this]
      have ⟨hK, hKN⟩ := h
      have hSub : n - k < n := sub_lt (Nat.lt_of_lt_of_le hK hKN) hK
      have : (n - k) / k ≤ n - k := ih (n - k) hSub
      exact succ_le_of_lt (Nat.lt_of_le_of_lt this hSub)
Floor Division Bounded by Dividend: $n / k \leq n$
Informal description
For any natural numbers $n$ and $k$, the floor division $n / k$ is less than or equal to $n$.
Nat.div_lt_self theorem
{n k : Nat} (hLtN : 0 < n) (hLtK : 1 < k) : n / k < n
Full source
theorem div_lt_self {n k : Nat} (hLtN : 0 < n) (hLtK : 1 < k) : n / k < n := by
  rw [div_eq]
  cases (inferInstance : Decidable (0 < k ∧ k ≤ n)) with
  | isFalse h => simp [hLtN, h]
  | isTrue h =>
    suffices (n - k) / k + 1 < n by simp [h, this]
    have ⟨_, hKN⟩ := h
    have : (n - k) / k ≤ n - k := div_le_self (n - k) k
    have := Nat.add_le_of_le_sub hKN this
    exact Nat.lt_of_lt_of_le (Nat.add_lt_add_left hLtK _) this
Floor Division Strictly Decreases Dividend: $n / k < n$ when $n > 0$ and $k > 1$
Informal description
For any natural numbers $n$ and $k$ such that $n > 0$ and $k > 1$, the floor division $n / k$ is strictly less than $n$.
Nat.modCore definition
(x y : Nat) : Nat
Full source
/--
The modulo operator, which computes the remainder when dividing one natural number by another.
Usually accessed via the `%` operator. When the divisor is `0`, the result is the dividend rather
than an error.

This is the core implementation of `Nat.mod`. It computes the correct result for any two closed
natural numbers, but it does not have some convenient [definitional
reductions](lean-manual://section/type-system) when the `Nat`s contain free variables. The wrapper
`Nat.mod` handles those cases specially and then calls `Nat.modCore`.

This function is overridden at runtime with an efficient implementation. This definition is the
logical model.
-/
@[extern "lean_nat_mod", irreducible]
protected noncomputable def modCore (x y : Nat) : Nat :=
  if hy : 0 < y then
    let rec
      go (fuel : Nat) (x : Nat) (hfuel : x < fuel) : Nat :=
      match fuel with
      | 0 => by contradiction
      | succ fuel =>
        if h : y ≤ x then
          go fuel (x - y) (div_rec_fuel_lemma hy h hfuel)
        else
          x
      termination_by structural fuel
    go (x + 1) x (Nat.lt_succ_self _)
  else
    x
Natural number modulo (core implementation)
Informal description
The function computes the remainder when dividing one natural number \( x \) by another \( y \). If \( y = 0 \), the result is \( x \). For \( y > 0 \), it recursively subtracts \( y \) from \( x \) until \( x < y \), returning the final value of \( x \).
Nat.modCore_eq theorem
(x y : Nat) : Nat.modCore x y = if 0 < y ∧ y ≤ x then Nat.modCore (x - y) y else x
Full source
protected theorem modCore_eq (x y : Nat) : Nat.modCore x y =
  if 0 < y ∧ y ≤ x then
    Nat.modCore (x - y) y
  else
    x := by
  unfold Nat.modCore
  split
  next =>
    rw [Nat.modCore.go]
    split
    next =>
      simp only [and_self, ↓reduceIte, *]
      apply modCore.go.fuel_congr
    next =>
      simp only [and_false, ↓reduceIte, *]
  next =>
    simp only [false_and, ↓reduceIte, *]
Recursive Definition of Natural Number Modulo Core Function
Informal description
For any natural numbers $x$ and $y$, the core modulo function satisfies: \[ \text{modCore}(x, y) = \begin{cases} \text{modCore}(x - y, y) & \text{if } 0 < y \text{ and } y \leq x, \\ x & \text{otherwise.} \end{cases} \]
Nat.mod definition
: @& Nat → @& Nat → Nat
Full source
/--
The modulo operator, which computes the remainder when dividing one natural number by another.
Usually accessed via the `%` operator. When the divisor is `0`, the result is the dividend rather
than an error.

`Nat.mod` is a wrapper around `Nat.modCore` that special-cases two situations, giving better
definitional reductions:
 * `Nat.mod 0 m` should reduce to `m`, for all terms `m : Nat`.
 * `Nat.mod n (m + n + 1)` should reduce to `n` for concrete `Nat` literals `n`.

These reductions help `Fin n` literals work well, because the `OfNat` instance for `Fin` uses
`Nat.mod`. In particular, `(0 : Fin (n + 1)).val` should reduce definitionally to `0`. `Nat.modCore`
can handle all numbers, but its definitional reductions are not as convenient.

This function is overridden at runtime with an efficient implementation. This definition is the
logical model.

Examples:
 * `7 % 2 = 1`
 * `9 % 3 = 0`
 * `5 % 7 = 5`
 * `5 % 0 = 5`
 * `show ∀ (n : Nat), 0 % n = 0 from fun _ => rfl`
 * `show ∀ (m : Nat), 5 % (m + 6) = 5 from fun _ => rfl`
-/
@[extern "lean_nat_mod"]
protected def mod : @& Nat → @& NatNat
  /-
  Nat.modCore is defined with fuel and thus does not reduce with open terms very well.
  Nevertheless it is desirable for trivial `Nat.mod` calculations, namely
  * `Nat.mod 0 m` for all `m`
  * `Nat.mod n (m + n + 1)` for concrete literals `n`,
  to reduce definitionally.
  This property is desirable for `Fin n` literals, as it means `(ofNat 0 : Fin n).val = 0` by
  definition.
   -/
  | 0, _ => 0
  | n@(_ + 1), m =>
    if m ≤ n -- NB: if n < m does not reduce as well as `m ≤ n`!
    then Nat.modCore n m
    else n
Natural number modulo operation
Informal description
The modulo operation on natural numbers, denoted $n \% m$, computes the remainder when dividing $n$ by $m$. If $m = 0$, the result is $n$. For $m > 0$, it returns the unique natural number $r$ such that $0 \leq r < m$ and $n = qm + r$ for some $q \in \mathbb{N}$.
Nat.instMod instance
: Mod Nat
Full source
instance instMod : Mod Nat := ⟨Nat.mod⟩
The Modulo Operation on Natural Numbers
Informal description
The natural numbers $\mathbb{N}$ are equipped with a canonical modulo operation, where for any $n, m \in \mathbb{N}$, $n \% m$ computes the remainder when $n$ is divided by $m$. If $m = 0$, the result is $n$; otherwise, it returns the unique natural number $r$ such that $0 \leq r < m$ and $n = qm + r$ for some $q \in \mathbb{N}$.
Nat.modCore_eq_mod theorem
(n m : Nat) : Nat.modCore n m = n % m
Full source
protected theorem modCore_eq_mod (n m : Nat) : Nat.modCore n m = n % m := by
  show Nat.modCore n m = Nat.mod n m
  match n, m with
  | 0, _ =>
    rw [Nat.modCore_eq]
    exact if_neg fun ⟨hlt, hle⟩ => Nat.lt_irrefl _ (Nat.lt_of_lt_of_le hlt hle)
  | (_ + 1), _ =>
    rw [Nat.mod]; dsimp
    refine iteInduction (fun _ => rfl) (fun h => ?false) -- cannot use `split` this early yet
    rw [Nat.modCore_eq]
    exact if_neg fun ⟨_hlt, hle⟩ => h hle
Equivalence of Core Modulo and Standard Modulo Operations on Natural Numbers
Informal description
For any natural numbers $n$ and $m$, the core modulo function $\text{modCore}(n, m)$ is equal to the modulo operation $n \% m$.
Nat.mod_eq theorem
(x y : Nat) : x % y = if 0 < y ∧ y ≤ x then (x - y) % y else x
Full source
theorem mod_eq (x y : Nat) : x % y = if 0 < y ∧ y ≤ x then (x - y) % y else x := by
  rw [←Nat.modCore_eq_mod, ←Nat.modCore_eq_mod, Nat.modCore_eq]
Recursive Definition of Natural Number Modulo Operation
Informal description
For any natural numbers $x$ and $y$, the modulo operation satisfies: \[ x \% y = \begin{cases} (x - y) \% y & \text{if } 0 < y \text{ and } y \leq x, \\ x & \text{otherwise.} \end{cases} \]
Nat.mod.inductionOn definition
{motive : Nat → Nat → Sort u} (x y : Nat) (ind : ∀ x y, 0 < y ∧ y ≤ x → motive (x - y) y → motive x y) (base : ∀ x y, ¬(0 < y ∧ y ≤ x) → motive x y) : motive x y
Full source
/--
An induction principle customized for reasoning about the recursion pattern of `Nat.mod`.
-/
def mod.inductionOn.{u}
      {motive : NatNat → Sort u}
      (x y  : Nat)
      (ind  : ∀ x y, 0 < y ∧ y ≤ x → motive (x - y) y → motive x y)
      (base : ∀ x y, ¬(0 < y ∧ y ≤ x) → motive x y)
      : motive x y :=
  div.inductionOn x y ind base
Induction principle for natural number modulo
Informal description
Given a motive `motive : ℕ → ℕ → Sort u` and natural numbers `x` and `y`, the induction principle `Nat.mod.inductionOn` allows proving `motive x y` by: 1. An inductive step `ind` that assumes `0 < y ∧ y ≤ x` and a proof of `motive (x - y) y` to construct a proof of `motive x y`. 2. A base case `base` that assumes `¬(0 < y ∧ y ≤ x)` to construct a proof of `motive x y`. This principle is specifically designed for reasoning about natural number modulo via iterated subtraction.
Nat.mod_eq_of_lt theorem
{a b : Nat} (h : a < b) : a % b = a
Full source
theorem mod_eq_of_lt {a b : Nat} (h : a < b) : a % b = a :=
  have : (if 0 < b ∧ b ≤ a then (a - b) % b else a) = a :=
    have h' : ¬(0 < b ∧ b ≤ a) := fun ⟨_, h₁⟩ => absurd h₁ (Nat.not_le_of_gt h)
    if_neg h'
  (mod_eq a b).symm ▸ this
Remainder Identity for Natural Numbers Under Strict Inequality: $a < b \implies a \% b = a$
Informal description
For any natural numbers $a$ and $b$ such that $a < b$, the remainder of $a$ divided by $b$ is equal to $a$, i.e., $a \% b = a$.
Nat.one_mod_eq_zero_iff theorem
{n : Nat} : 1 % n = 0 ↔ n = 1
Full source
@[simp] theorem one_mod_eq_zero_iff {n : Nat} : 1 % n = 0 ↔ n = 1 := by
  match n with
  | 0 => simp
  | 1 => simp
  | n + 2 =>
    rw [mod_eq_of_lt (by exact Nat.lt_of_sub_eq_succ rfl)]
    simp only [add_one_ne_zero, false_iff, ne_eq]
    exact ne_of_beq_eq_false rfl
Remainder Condition: $1 \% n = 0 \leftrightarrow n = 1$
Informal description
For any natural number $n$, the remainder of $1$ divided by $n$ equals $0$ if and only if $n = 1$.
Nat.zero_eq_one_mod_iff theorem
{n : Nat} : 0 = 1 % n ↔ n = 1
Full source
@[simp] theorem zero_eq_one_mod_iff {n : Nat} : 0 = 1 % n ↔ n = 1 := by
  rw [eq_comm]
  simp
$0 = 1 \% n \leftrightarrow n = 1$ for natural numbers
Informal description
For any natural number $n$, the equality $0 = 1 \% n$ holds if and only if $n = 1$.
Nat.mod_eq_sub_mod theorem
{a b : Nat} (h : a ≥ b) : a % b = (a - b) % b
Full source
theorem mod_eq_sub_mod {a b : Nat} (h : a ≥ b) : a % b = (a - b) % b :=
  match eq_zero_or_pos b with
  | Or.inl h₁ => h₁.symm ▸ (Nat.sub_zero a).symmrfl
  | Or.inr h₁ => (mod_eq a b).symmif_pos ⟨h₁, h⟩
Modulo Equals Subtraction Modulo for Natural Numbers
Informal description
For any natural numbers $a$ and $b$ such that $a \geq b$, the remainder when $a$ is divided by $b$ is equal to the remainder when $(a - b)$ is divided by $b$, i.e., $a \% b = (a - b) \% b$.
Nat.mod_lt theorem
(x : Nat) {y : Nat} : y > 0 → x % y < y
Full source
theorem mod_lt (x : Nat) {y : Nat} : y > 0 → x % y < y := by
  induction x, y using mod.inductionOn with
  | base x y h₁ =>
    intro h₂
    have h₁ : ¬ 0 < y¬ 0 < y ∨ ¬ y ≤ x := Decidable.not_and_iff_or_not.mp h₁
    match h₁ with
    | Or.inl h₁ => exact absurd h₂ h₁
    | Or.inr h₁ =>
      have hgt : y > x := gt_of_not_le h₁
      have heq : x % y = x := mod_eq_of_lt hgt
      rw [← heq] at hgt
      exact hgt
  | ind x y h h₂ =>
    intro h₃
    have ⟨_, h₁⟩ := h
    rw [mod_eq_sub_mod h₁]
    exact h₂ h₃
Remainder is Less Than Divisor for Natural Numbers
Informal description
For any natural numbers $x$ and $y$ with $y > 0$, the remainder of $x$ divided by $y$ satisfies $x \% y < y$.
Nat.sub_mod_add_mod_cancel theorem
(a b : Nat) [NeZero a] : a - b % a + b % a = a
Full source
@[simp] protected theorem sub_mod_add_mod_cancel (a b : Nat) [NeZero a] : a - b % a + b % a = a := by
  rw [Nat.sub_add_cancel]
  cases a with
  | zero => simp_all
  | succ a =>
    exact Nat.le_of_lt (mod_lt b (zero_lt_succ a))
Subtraction-Modulo Cancellation: $(a - (b \% a)) + (b \% a) = a$ for $a \neq 0$
Informal description
For any natural numbers $a$ and $b$ with $a \neq 0$, the sum of $(a - (b \% a))$ and $(b \% a)$ equals $a$, i.e., $(a - (b \% a)) + (b \% a) = a$.
Nat.mod_le theorem
(x y : Nat) : x % y ≤ x
Full source
theorem mod_le (x y : Nat) : x % y ≤ x := by
  match Nat.lt_or_ge x y with
  | Or.inl h₁ => rw [mod_eq_of_lt h₁]; apply Nat.le_refl
  | Or.inr h₁ => match eq_zero_or_pos y with
    | Or.inl h₂ => rw [h₂, Nat.mod_zero x]; apply Nat.le_refl
    | Or.inr h₂ => exact Nat.le_trans (Nat.le_of_lt (mod_lt _ h₂)) h₁
Remainder Bound: $x \% y \leq x$ for natural numbers
Informal description
For any natural numbers $x$ and $y$, the remainder of $x$ divided by $y$ is less than or equal to $x$, i.e., $x \% y \leq x$.
Nat.mod_lt_of_lt theorem
{a b c : Nat} (h : a < c) : a % b < c
Full source
theorem mod_lt_of_lt {a b c : Nat} (h : a < c) : a % b < c :=
  Nat.lt_of_le_of_lt (Nat.mod_le _ _) h
Remainder Bound Under Strict Inequality: $a < c \implies a \% b < c$
Informal description
For any natural numbers $a$, $b$, and $c$, if $a < c$, then the remainder of $a$ divided by $b$ is less than $c$, i.e., $a \% b < c$.
Nat.zero_mod theorem
(b : Nat) : 0 % b = 0
Full source
@[simp] theorem zero_mod (b : Nat) : 0 % b = 0 := by
  rw [mod_eq]
  have : ¬ (0 < b ∧ b = 0) := by
    intro ⟨h₁, h₂⟩
    simp_all
  simp [this]
Modulo Zero Property: $0 \% b = 0$ for any natural number $b$
Informal description
For any natural number $b$, the remainder when $0$ is divided by $b$ is $0$, i.e., $0 \% b = 0$.
Nat.mod_self theorem
(n : Nat) : n % n = 0
Full source
@[simp] theorem mod_self (n : Nat) : n % n = 0 := by
  rw [mod_eq_sub_mod (Nat.le_refl _), Nat.sub_self, zero_mod]
Self-Modulo Property: $n \% n = 0$ for Natural Numbers
Informal description
For any natural number $n$, the remainder when $n$ is divided by itself is zero, i.e., $n \% n = 0$.
Nat.mod_one theorem
(x : Nat) : x % 1 = 0
Full source
theorem mod_one (x : Nat) : x % 1 = 0 := by
  have h : x % 1 < 1 := mod_lt x (by decide)
  have : (y : Nat) → y < 1 → y = 0 := by
    intro y
    cases y with
    | zero   => intro _; rfl
    | succ y => intro h; apply absurd (Nat.lt_of_succ_lt_succ h) (Nat.not_lt_zero y)
  exact this _ h
Remainder Modulo One is Zero for Natural Numbers
Informal description
For any natural number $x$, the remainder when $x$ is divided by $1$ is $0$, i.e., $x \% 1 = 0$.
Nat.div_add_mod theorem
(m n : Nat) : n * (m / n) + m % n = m
Full source
theorem div_add_mod (m n : Nat) : n * (m / n) + m % n = m := by
  rw [div_eq, mod_eq]
  have h : Decidable (0 < n ∧ n ≤ m) := inferInstance
  cases h with
  | isFalse h => simp [h]
  | isTrue h =>
    simp [h]
    have ih := div_add_mod (m - n) n
    rw [Nat.left_distrib, Nat.mul_one, Nat.add_assoc, Nat.add_left_comm, ih, Nat.add_comm, Nat.sub_add_cancel h.2]
decreasing_by apply div_rec_lemma; assumption
Division-Remainder Decomposition for Natural Numbers
Informal description
For any natural numbers $m$ and $n$, the sum of $n$ multiplied by the integer division of $m$ by $n$ and the remainder of $m$ divided by $n$ equals $m$, i.e., \[ n \cdot (m / n) + (m \% n) = m. \]
Nat.div_eq_sub_div theorem
(h₁ : 0 < b) (h₂ : b ≤ a) : a / b = (a - b) / b + 1
Full source
theorem div_eq_sub_div (h₁ : 0 < b) (h₂ : b ≤ a) : a / b = (a - b) / b + 1 := by
 rw [div_eq a, if_pos]; constructor <;> assumption
Recursive Division Formula for Natural Numbers: $a / b = (a - b) / b + 1$ under $0 < b \leq a$
Informal description
For any natural numbers $a$ and $b$ such that $0 < b$ and $b \leq a$, the division $a / b$ satisfies the recursive relation: \[ a / b = (a - b) / b + 1. \]
Nat.mod_add_div theorem
(m k : Nat) : m % k + k * (m / k) = m
Full source
theorem mod_add_div (m k : Nat) : m % k + k * (m / k) = m := by
  induction m, k using mod.inductionOn with rw [div_eq, mod_eq]
  | base x y h => simp [h]
  | ind x y h IH => simp [h]; rw [Nat.mul_succ, ← Nat.add_assoc, IH, Nat.sub_add_cancel h.2]
Division-Remainder Decomposition for Natural Numbers: $m \% k + k \cdot (m / k) = m$
Informal description
For any natural numbers $m$ and $k$, the sum of the remainder $m \% k$ and the product $k \cdot (m / k)$ equals $m$, i.e., \[ m \% k + k \cdot (m / k) = m. \]
Nat.mod_def theorem
(m k : Nat) : m % k = m - k * (m / k)
Full source
theorem mod_def (m k : Nat) : m % k = m - k * (m / k) := by
  rw [Nat.sub_eq_of_eq_add]
  apply (Nat.mod_add_div _ _).symm
Modulo as Remainder: $m \% k = m - k \cdot (m / k)$
Informal description
For any natural numbers $m$ and $k$, the modulo operation satisfies $m \% k = m - k \cdot (m / k)$, where $/$ denotes floor division.
Nat.mod_eq_sub_mul_div theorem
{x k : Nat} : x % k = x - k * (x / k)
Full source
theorem mod_eq_sub_mul_div {x k : Nat} : x % k = x - k * (x / k) := mod_def _ _
Modulo as Remainder: $x \% k = x - k \cdot (x / k)$
Informal description
For any natural numbers $x$ and $k$, the remainder when $x$ is divided by $k$ satisfies $x \% k = x - k \cdot (x / k)$, where $/$ denotes floor division.
Nat.mod_eq_sub_div_mul theorem
{x k : Nat} : x % k = x - (x / k) * k
Full source
theorem mod_eq_sub_div_mul {x k : Nat} : x % k = x - (x / k) * k := by
  rw [mod_eq_sub_mul_div, Nat.mul_comm]
Modulo as Remainder: $x \% k = x - (x / k) \cdot k$
Informal description
For any natural numbers $x$ and $k$, the remainder when $x$ is divided by $k$ satisfies $x \% k = x - (x / k) \cdot k$, where $/$ denotes floor division.
Nat.div_one theorem
(n : Nat) : n / 1 = n
Full source
@[simp] protected theorem div_one (n : Nat) : n / 1 = n := by
  have := mod_add_div n 1
  rwa [mod_one, Nat.zero_add, Nat.one_mul] at this
Division by One Identity: $n / 1 = n$
Informal description
For any natural number $n$, the division of $n$ by $1$ equals $n$, i.e., $n / 1 = n$.
Nat.div_zero theorem
(n : Nat) : n / 0 = 0
Full source
@[simp] protected theorem div_zero (n : Nat) : n / 0 = 0 := by
  rw [div_eq]; simp [Nat.lt_irrefl]
Division by Zero in Natural Numbers: $n / 0 = 0$
Informal description
For any natural number $n$, the division of $n$ by zero equals zero, i.e., $n / 0 = 0$.
Nat.zero_div theorem
(b : Nat) : 0 / b = 0
Full source
@[simp] protected theorem zero_div (b : Nat) : 0 / b = 0 :=
  (div_eq 0 b).trans <| if_neg <| And.rec Nat.not_le_of_gt
Division of Zero by Any Natural Number is Zero
Informal description
For any natural number $b$, the division of $0$ by $b$ equals $0$, i.e., $0 / b = 0$.
Nat.le_div_iff_mul_le theorem
(k0 : 0 < k) : x ≤ y / k ↔ x * k ≤ y
Full source
theorem le_div_iff_mul_le (k0 : 0 < k) : x ≤ y / k ↔ x * k ≤ y := by
  induction y, k using mod.inductionOn generalizing x with
    (rw [div_eq]; simp [h]; cases x with | zero => simp [zero_le] | succ x => ?_)
  | base y k h =>
    simp only [add_one, succ_mul, false_iff, Nat.not_le, Nat.succ_ne_zero]
    refine Nat.lt_of_lt_of_le ?_ (Nat.le_add_left ..)
    exact Nat.not_le.1 fun h' => h ⟨k0, h'⟩
  | ind y k h IH =>
    rw [Nat.add_le_add_iff_right, IH k0, succ_mul,
        ← Nat.add_sub_cancel (x*k) k, Nat.sub_le_sub_iff_right h.2, Nat.add_sub_cancel]
Natural Number Division Inequality: $x \leq y/k \leftrightarrow x \cdot k \leq y$
Informal description
For any natural numbers $x$, $y$, and $k$ with $0 < k$, the inequality $x \leq y / k$ holds if and only if $x \cdot k \leq y$.
Nat.div_div_eq_div_mul theorem
(m n k : Nat) : m / n / k = m / (n * k)
Full source
protected theorem div_div_eq_div_mul (m n k : Nat) : m / n / k = m / (n * k) := by
  cases eq_zero_or_pos k with
  | inl k0 => rw [k0, Nat.mul_zero, Nat.div_zero, Nat.div_zero] | inr kpos => ?_
  cases eq_zero_or_pos n with
  | inl n0 => rw [n0, Nat.zero_mul, Nat.div_zero, Nat.zero_div] | inr npos => ?_

  apply Nat.le_antisymm

  apply (le_div_iff_mul_le (Nat.mul_pos npos kpos)).2
  rw [Nat.mul_comm n k, ← Nat.mul_assoc]
  apply (le_div_iff_mul_le npos).1
  apply (le_div_iff_mul_le kpos).1
  (apply Nat.le_refl)

  apply (le_div_iff_mul_le kpos).2
  apply (le_div_iff_mul_le npos).2
  rw [Nat.mul_assoc, Nat.mul_comm n k]
  apply (le_div_iff_mul_le (Nat.mul_pos kpos npos)).1
  apply Nat.le_refl
Division of Division Equals Division by Product in Natural Numbers: $(m / n) / k = m / (n \cdot k)$
Informal description
For any natural numbers $m$, $n$, and $k$, the result of dividing $m$ by $n$ and then dividing the quotient by $k$ is equal to dividing $m$ by the product $n \cdot k$, i.e., $(m / n) / k = m / (n \cdot k)$.
Nat.div_mul_le_self theorem
: ∀ (m n : Nat), m / n * n ≤ m
Full source
theorem div_mul_le_self : ∀ (m n : Nat), m / n * n ≤ m
  | m, 0   => by simp
  | _, _+1 => (le_div_iff_mul_le (Nat.succ_pos _)).1 (Nat.le_refl _)
Floor Division Bound: $(m / n) \cdot n \leq m$ for Natural Numbers
Informal description
For any natural numbers $m$ and $n$, the product of the floor division $m / n$ and $n$ is less than or equal to $m$, i.e., $(m / n) \cdot n \leq m$.
Nat.div_lt_iff_lt_mul theorem
(Hk : 0 < k) : x / k < y ↔ x < y * k
Full source
theorem div_lt_iff_lt_mul (Hk : 0 < k) : x / k < y ↔ x < y * k := by
  rw [← Nat.not_le, ← Nat.not_le]; exact not_congr (le_div_iff_mul_le Hk)
Natural Number Division Inequality: $x / k < y \leftrightarrow x < y \cdot k$
Informal description
For any natural numbers $x$, $y$, and $k$ with $0 < k$, the inequality $x / k < y$ holds if and only if $x < y \cdot k$.
Nat.pos_of_div_pos theorem
{a b : Nat} (h : 0 < a / b) : 0 < a
Full source
theorem pos_of_div_pos {a b : Nat} (h : 0 < a / b) : 0 < a := by
  cases b with
  | zero => simp at h
  | succ b =>
    match a, h with
    | 0, h => simp at h
    | a + 1, _ => exact zero_lt_succ a
Positivity of Dividend from Positive Quotient in Natural Numbers
Informal description
For any natural numbers $a$ and $b$, if the quotient $a / b$ is positive (i.e., $0 < a / b$), then $a$ is positive (i.e., $0 < a$).
Nat.add_div_right theorem
(x : Nat) {z : Nat} (H : 0 < z) : (x + z) / z = (x / z) + 1
Full source
@[simp] theorem add_div_right (x : Nat) {z : Nat} (H : 0 < z) : (x + z) / z = (x / z) + 1 := by
  rw [div_eq_sub_div H (Nat.le_add_left _ _), Nat.add_sub_cancel]
Division of Sum by Nonzero Natural Number: $(x + z)/z = x/z + 1$
Informal description
For any natural numbers $x$ and $z$ with $z > 0$, the division $(x + z) / z$ equals $(x / z) + 1$.
Nat.add_div_left theorem
(x : Nat) {z : Nat} (H : 0 < z) : (z + x) / z = (x / z) + 1
Full source
@[simp] theorem add_div_left (x : Nat) {z : Nat} (H : 0 < z) : (z + x) / z = (x / z) + 1 := by
  rw [Nat.add_comm, add_div_right x H]
Division of Sum by Nonzero Natural Number: $(z + x)/z = x/z + 1$
Informal description
For any natural numbers $x$ and $z$ with $z > 0$, the division $(z + x) / z$ equals $(x / z) + 1$.
Nat.add_mul_div_left theorem
(x z : Nat) {y : Nat} (H : 0 < y) : (x + y * z) / y = x / y + z
Full source
theorem add_mul_div_left (x z : Nat) {y : Nat} (H : 0 < y) : (x + y * z) / y = x / y + z := by
  induction z with
  | zero => rw [Nat.mul_zero, Nat.add_zero, Nat.add_zero]
  | succ z ih => rw [mul_succ, ← Nat.add_assoc, add_div_right _ H, ih]; rfl
Division of Sum with Multiple: $(x + y \cdot z)/y = x/y + z$
Informal description
For any natural numbers $x$, $y$, and $z$ with $y > 0$, the division $(x + y \cdot z) / y$ equals $(x / y) + z$.
Nat.add_mul_div_right theorem
(x y : Nat) {z : Nat} (H : 0 < z) : (x + y * z) / z = x / z + y
Full source
theorem add_mul_div_right (x y : Nat) {z : Nat} (H : 0 < z) : (x + y * z) / z = x / z + y := by
  rw [Nat.mul_comm, add_mul_div_left _ _ H]
Division of Sum with Multiple (Right Version): $(x + y \cdot z)/z = x/z + y$
Informal description
For any natural numbers $x$, $y$, and $z$ with $z > 0$, the division $(x + y \cdot z) / z$ equals $(x / z) + y$.
Nat.add_mod_right theorem
(x z : Nat) : (x + z) % z = x % z
Full source
@[simp] theorem add_mod_right (x z : Nat) : (x + z) % z = x % z := by
  rw [mod_eq_sub_mod (Nat.le_add_left ..), Nat.add_sub_cancel]
Right Addition Modulo Identity: $(x + z) \% z = x \% z$
Informal description
For any natural numbers $x$ and $z$, the remainder when $(x + z)$ is divided by $z$ equals the remainder when $x$ is divided by $z$, i.e., $(x + z) \% z = x \% z$.
Nat.add_mod_left theorem
(x z : Nat) : (x + z) % x = z % x
Full source
@[simp] theorem add_mod_left (x z : Nat) : (x + z) % x = z % x := by
  rw [Nat.add_comm, add_mod_right]
Left Addition Modulo Identity: $(x + z) \% x = z \% x$
Informal description
For any natural numbers $x$ and $z$, the remainder when $(x + z)$ is divided by $x$ equals the remainder when $z$ is divided by $x$, i.e., $(x + z) \% x = z \% x$.
Nat.add_mul_mod_self_left theorem
(x y z : Nat) : (x + y * z) % y = x % y
Full source
@[simp] theorem add_mul_mod_self_left (x y z : Nat) : (x + y * z) % y = x % y := by
  match z with
  | 0 => rw [Nat.mul_zero, Nat.add_zero]
  | succ z => rw [mul_succ, ← Nat.add_assoc, add_mod_right, add_mul_mod_self_left (z := z)]
Left Multiplication-Addition Modulo Identity: $(x + y \cdot z) \% y = x \% y$
Informal description
For any natural numbers $x$, $y$, and $z$, the remainder when $(x + y \cdot z)$ is divided by $y$ equals the remainder when $x$ is divided by $y$, i.e., $(x + y \cdot z) \% y = x \% y$.
Nat.mul_add_mod_self_left theorem
(a b c : Nat) : (a * b + c) % a = c % a
Full source
@[simp] theorem mul_add_mod_self_left (a b c : Nat) : (a * b + c) % a = c % a := by
  rw [Nat.add_comm, Nat.add_mul_mod_self_left]
Left Multiplication-Addition Modulo Identity: $(a \cdot b + c) \% a = c \% a$
Informal description
For any natural numbers $a$, $b$, and $c$, the remainder when $(a \cdot b + c)$ is divided by $a$ equals the remainder when $c$ is divided by $a$, i.e., $(a \cdot b + c) \% a = c \% a$.
Nat.add_mul_mod_self_right theorem
(x y z : Nat) : (x + y * z) % z = x % z
Full source
@[simp] theorem add_mul_mod_self_right (x y z : Nat) : (x + y * z) % z = x % z := by
  rw [Nat.mul_comm, add_mul_mod_self_left]
Right Multiplication-Addition Modulo Identity: $(x + y \cdot z) \% z = x \% z$
Informal description
For any natural numbers $x$, $y$, and $z$, the remainder when $(x + y \cdot z)$ is divided by $z$ equals the remainder when $x$ is divided by $z$, i.e., $(x + y \cdot z) \% z = x \% z$.
Nat.mul_add_mod_self_right theorem
(a b c : Nat) : (a * b + c) % b = c % b
Full source
@[simp] theorem mul_add_mod_self_right (a b c : Nat) : (a * b + c) % b = c % b := by
  rw [Nat.add_comm, Nat.add_mul_mod_self_right]
Right Multiplication-Addition Modulo Identity: $(a \cdot b + c) \% b = c \% b$
Informal description
For any natural numbers $a$, $b$, and $c$, the remainder when $(a \cdot b + c)$ is divided by $b$ equals the remainder when $c$ is divided by $b$, i.e., $(a \cdot b + c) \% b = c \% b$.
Nat.mul_mod_right theorem
(m n : Nat) : (m * n) % m = 0
Full source
@[simp] theorem mul_mod_right (m n : Nat) : (m * n) % m = 0 := by
  rw [← Nat.zero_add (m * n), add_mul_mod_self_left, zero_mod]
Modulo Identity for Multiplication: $(m \cdot n) \% m = 0$
Informal description
For any natural numbers $m$ and $n$, the remainder when $m \cdot n$ is divided by $m$ is $0$, i.e., $(m \cdot n) \% m = 0$.
Nat.mul_mod_left theorem
(m n : Nat) : (m * n) % n = 0
Full source
@[simp] theorem mul_mod_left (m n : Nat) : (m * n) % n = 0 := by
  rw [Nat.mul_comm, mul_mod_right]
Modulo Identity for Multiplication: $(m \cdot n) \% n = 0$
Informal description
For any natural numbers $m$ and $n$, the remainder when $m \cdot n$ is divided by $n$ is $0$, i.e., $(m \cdot n) \% n = 0$.
Nat.div_eq_of_lt_le theorem
(lo : k * n ≤ m) (hi : m < (k + 1) * n) : m / n = k
Full source
protected theorem div_eq_of_lt_le (lo : k * n ≤ m) (hi : m < (k + 1) * n) : m / n = k :=
have npos : 0 < n := (eq_zero_or_pos _).resolve_left fun hn => by
  rw [hn, Nat.mul_zero] at hi lo; exact absurd lo (Nat.not_le_of_gt hi)
Nat.le_antisymm
  (le_of_lt_succ ((Nat.div_lt_iff_lt_mul npos).2 hi))
  ((Nat.le_div_iff_mul_le npos).2 lo)
Division of Natural Numbers in Interval: $m / n = k$ when $k \cdot n \leq m < (k + 1) \cdot n$
Informal description
For any natural numbers $k$, $n$, and $m$ such that $k \cdot n \leq m$ and $m < (k + 1) \cdot n$, the division of $m$ by $n$ equals $k$, i.e., $m / n = k$.
Nat.sub_mul_div theorem
(x n p : Nat) (h₁ : n * p ≤ x) : (x - n * p) / n = x / n - p
Full source
theorem sub_mul_div (x n p : Nat) (h₁ : n*p ≤ x) : (x - n*p) / n = x / n - p := by
  match eq_zero_or_pos n with
  | .inl h₀ => rw [h₀, Nat.div_zero, Nat.div_zero, Nat.zero_sub]
  | .inr h₀ => induction p with
    | zero => rw [Nat.mul_zero, Nat.sub_zero, Nat.sub_zero]
    | succ p IH =>
      have h₂ : n * p ≤ x := Nat.le_trans (Nat.mul_le_mul_left _ (le_succ _)) h₁
      have h₃ : x - n * p ≥ n := by
        apply Nat.le_of_add_le_add_right
        rw [Nat.sub_add_cancel h₂, Nat.add_comm]
        rw [mul_succ] at h₁
        exact h₁
      rw [sub_succ, ← IH h₂, div_eq_sub_div h₀ h₃]
      simp [Nat.pred_succ, mul_succ, Nat.sub_sub]
Subtraction-Multiplication-Division Identity for Natural Numbers: $(x - np)/n = x/n - p$ under $np \leq x$
Informal description
For any natural numbers $x$, $n$, and $p$ such that $n \cdot p \leq x$, the following equality holds: \[ (x - n \cdot p) / n = x / n - p. \]
Nat.mul_sub_div theorem
(x n p : Nat) (h₁ : x < n * p) : (n * p - (x + 1)) / n = p - ((x / n) + 1)
Full source
theorem mul_sub_div (x n p : Nat) (h₁ : x < n*p) : (n * p - (x + 1)) / n = p - ((x / n) + 1) := by
  have npos : 0 < n := (eq_zero_or_pos _).resolve_left fun n0 => by
    rw [n0, Nat.zero_mul] at h₁; exact not_lt_zero _ h₁
  apply Nat.div_eq_of_lt_le
  focus
    rw [Nat.mul_sub_right_distrib, Nat.mul_comm]
    exact Nat.sub_le_sub_left ((div_lt_iff_lt_mul npos).1 (lt_succ_self _)) _
  focus
    show succ (pred (n * p - x)) ≤ (succ (pred (p - x / n))) * n
    rw [succ_pred_eq_of_pos (Nat.sub_pos_of_lt h₁),
      fun h => succ_pred_eq_of_pos (Nat.sub_pos_of_lt h)] -- TODO: why is the function needed?
    focus
      rw [Nat.mul_sub_right_distrib, Nat.mul_comm]
      exact Nat.sub_le_sub_left (div_mul_le_self ..) _
    focus
      rwa [div_lt_iff_lt_mul npos, Nat.mul_comm]
Division Identity for Natural Numbers: $(np - (x+1))/n = p - (x/n + 1)$ under $x < np$
Informal description
For any natural numbers $x$, $n$, and $p$ such that $x < n \cdot p$, the following equality holds: \[ (n \cdot p - (x + 1)) / n = p - (x / n + 1). \]
Nat.mul_mod_mul_left theorem
(z x y : Nat) : (z * x) % (z * y) = z * (x % y)
Full source
theorem mul_mod_mul_left (z x y : Nat) : (z * x) % (z * y) = z * (x % y) :=
  if y0 : y = 0 then by
    rw [y0, Nat.mul_zero, mod_zero, mod_zero]
  else if z0 : z = 0 then by
    rw [z0, Nat.zero_mul, Nat.zero_mul, Nat.zero_mul, mod_zero]
  else by
    induction x using Nat.strongRecOn with
    | _ n IH =>
      have y0 : y > 0 := Nat.pos_of_ne_zero y0
      have z0 : z > 0 := Nat.pos_of_ne_zero z0
      cases Nat.lt_or_ge n y with
      | inl yn => rw [mod_eq_of_lt yn, mod_eq_of_lt (Nat.mul_lt_mul_of_pos_left yn z0)]
      | inr yn =>
        rw [mod_eq_sub_mod yn, mod_eq_sub_mod (Nat.mul_le_mul_left z yn),
          ← Nat.mul_sub_left_distrib]
        exact IH _ (sub_lt (Nat.lt_of_lt_of_le y0 yn) y0)
Left Multiplication Preserves Modulo Operation: $(zx) \% (zy) = z(x \% y)$
Informal description
For any natural numbers $z$, $x$, and $y$, the remainder of the product $z \cdot x$ divided by $z \cdot y$ is equal to $z$ multiplied by the remainder of $x$ divided by $y$, i.e., \[ (z \cdot x) \% (z \cdot y) = z \cdot (x \% y). \]
Nat.div_eq_of_lt theorem
(h₀ : a < b) : a / b = 0
Full source
theorem div_eq_of_lt (h₀ : a < b) : a / b = 0 := by
  rw [div_eq a, if_neg]
  intro h₁
  apply Nat.not_le_of_gt h₀ h₁.right
Division of Smaller Natural Number Yields Zero: $a < b \Rightarrow a / b = 0$
Informal description
For any natural numbers $a$ and $b$ such that $a < b$, the integer division $a / b$ equals $0$.
Nat.mul_div_cancel theorem
(m : Nat) {n : Nat} (H : 0 < n) : m * n / n = m
Full source
protected theorem mul_div_cancel (m : Nat) {n : Nat} (H : 0 < n) : m * n / n = m := by
  let t := add_mul_div_right 0 m H
  rwa [Nat.zero_add, Nat.zero_div, Nat.zero_add] at t
Cancellation of Multiplication and Division: $(m \cdot n)/n = m$ for $n > 0$
Informal description
For any natural numbers $m$ and $n$ with $n > 0$, the division $(m \cdot n) / n$ equals $m$.
Nat.mul_div_cancel_left theorem
(m : Nat) {n : Nat} (H : 0 < n) : n * m / n = m
Full source
protected theorem mul_div_cancel_left (m : Nat) {n : Nat} (H : 0 < n) : n * m / n = m := by
  rw [Nat.mul_comm, Nat.mul_div_cancel _ H]
Left Cancellation of Multiplication and Division: $(n \cdot m)/n = m$ for $n > 0$
Informal description
For any natural numbers $m$ and $n$ with $n > 0$, the division $(n \cdot m) / n$ equals $m$.
Nat.div_le_of_le_mul theorem
{m n : Nat} : ∀ {k}, m ≤ k * n → m / k ≤ n
Full source
protected theorem div_le_of_le_mul {m n : Nat} : ∀ {k}, m ≤ k * n → m / k ≤ n
  | 0, _ => by simp [Nat.div_zero, n.zero_le]
  | succ k, h => by
    suffices succ k * (m / succ k) ≤ succ k * n from
      Nat.le_of_mul_le_mul_left this (zero_lt_succ _)
    have h1 : succ k * (m / succ k) ≤ m % succ k + succ k * (m / succ k) := Nat.le_add_left _ _
    have h2 : m % succ k + succ k * (m / succ k) = m := by rw [mod_add_div]
    have h3 : m ≤ succ k * n := h
    rw [← h2] at h3
    exact Nat.le_trans h1 h3
Floor Division Bounds: $m \leq k \cdot n \implies m / k \leq n$
Informal description
For any natural numbers $m$, $n$, and $k$, if $m \leq k \cdot n$, then the floor division $m / k$ is less than or equal to $n$.
Nat.mul_div_right theorem
(n : Nat) {m : Nat} (H : 0 < m) : m * n / m = n
Full source
@[simp] theorem mul_div_right (n : Nat) {m : Nat} (H : 0 < m) : m * n / m = n := by
  induction n <;> simp_all [mul_succ]
Division Cancellation: $(m \cdot n)/m = n$ for $m > 0$
Informal description
For any natural numbers $n$ and $m$ with $m > 0$, the division $(m \cdot n) / m$ equals $n$.
Nat.mul_div_left theorem
(m : Nat) {n : Nat} (H : 0 < n) : m * n / n = m
Full source
@[simp] theorem mul_div_left (m : Nat) {n : Nat} (H : 0 < n) : m * n / n = m := by
  rw [Nat.mul_comm, mul_div_right _ H]
Left Multiplication Cancellation in Division: $(m \cdot n)/n = m$ for $n > 0$
Informal description
For any natural numbers $m$ and $n$ with $n > 0$, the division $(m \cdot n) / n$ equals $m$.
Nat.div_self theorem
(H : 0 < n) : n / n = 1
Full source
protected theorem div_self (H : 0 < n) : n / n = 1 := by
  let t := add_div_right 0 H
  rwa [Nat.zero_add, Nat.zero_div] at t
Self-Division of Positive Natural Number: $n / n = 1$
Informal description
For any natural number $n > 0$, the division of $n$ by itself equals $1$, i.e., $n / n = 1$.
Nat.div_eq_of_eq_mul_left theorem
(H1 : 0 < n) (H2 : m = k * n) : m / n = k
Full source
protected theorem div_eq_of_eq_mul_left (H1 : 0 < n) (H2 : m = k * n) : m / n = k :=
by rw [H2, Nat.mul_div_cancel _ H1]
Division of Product by Factor Equals Other Factor: $m = k \cdot n \implies m / n = k$ for $n > 0$
Informal description
For any natural numbers $m$, $n$, and $k$ with $n > 0$, if $m$ is equal to $k \cdot n$, then the division $m / n$ equals $k$.
Nat.div_eq_of_eq_mul_right theorem
(H1 : 0 < n) (H2 : m = n * k) : m / n = k
Full source
protected theorem div_eq_of_eq_mul_right (H1 : 0 < n) (H2 : m = n * k) : m / n = k :=
by rw [H2, Nat.mul_div_cancel_left _ H1]
Division of Product by Factor Equals Other Factor: $m = n \cdot k \implies m / n = k$ for $n > 0$
Informal description
For any natural numbers $m$, $n$, and $k$ with $n > 0$, if $m$ is equal to $n \cdot k$, then the division $m / n$ equals $k$.
Nat.mul_div_mul_left theorem
{m : Nat} (n k : Nat) (H : 0 < m) : m * n / (m * k) = n / k
Full source
protected theorem mul_div_mul_left {m : Nat} (n k : Nat) (H : 0 < m) :
    m * n / (m * k) = n / k := by rw [← Nat.div_div_eq_div_mul, Nat.mul_div_cancel_left _ H]
Left Cancellation of Common Factor in Division: $(m \cdot n)/(m \cdot k) = n/k$ for $m > 0$
Informal description
For any natural numbers $m > 0$, $n$, and $k$, the division $(m \cdot n) / (m \cdot k)$ equals $n / k$.
Nat.mul_div_mul_right theorem
{m : Nat} (n k : Nat) (H : 0 < m) : n * m / (k * m) = n / k
Full source
protected theorem mul_div_mul_right {m : Nat} (n k : Nat) (H : 0 < m) :
    n * m / (k * m) = n / k := by rw [Nat.mul_comm, Nat.mul_comm k, Nat.mul_div_mul_left _ _ H]
Right Cancellation of Common Factor in Division: $(n \cdot m)/(k \cdot m) = n/k$ for $m > 0$
Informal description
For any natural numbers $m > 0$, $n$, and $k$, the division $(n \cdot m) / (k \cdot m)$ equals $n / k$.
Nat.mul_div_le theorem
(m n : Nat) : n * (m / n) ≤ m
Full source
theorem mul_div_le (m n : Nat) : n * (m / n) ≤ m := by
  match n, Nat.eq_zero_or_pos n with
  | _, Or.inl rfl => rw [Nat.zero_mul]; exact m.zero_le
  | n, Or.inr h => rw [Nat.mul_comm, ← Nat.le_div_iff_mul_le h]; exact Nat.le_refl _
Upper Bound on Natural Number Division: $n \cdot (m / n) \leq m$
Informal description
For any natural numbers $m$ and $n$, the product $n \cdot (m / n)$ is less than or equal to $m$.