Module docstring
{"Helper theorems for dvd simproc "}
{"Helper theorems for dvd simproc "}
Nat.dvd_refl
theorem
(a : Nat) : a ∣ a
@[simp] protected theorem dvd_refl (a : Nat) : a ∣ a := ⟨1, by simp⟩
Nat.dvd_zero
theorem
(a : Nat) : a ∣ 0
@[simp] protected theorem dvd_zero (a : Nat) : a ∣ 0 := ⟨0, by simp⟩
Nat.dvd_mul_left
theorem
(a b : Nat) : a ∣ b * a
protected theorem dvd_mul_left (a b : Nat) : a ∣ b * a := ⟨b, Nat.mul_comm b a⟩
Nat.dvd_mul_right
theorem
(a b : Nat) : a ∣ a * b
protected theorem dvd_mul_right (a b : Nat) : a ∣ a * b := ⟨b, rfl⟩
Nat.dvd_trans
theorem
{a b c : Nat} (h₁ : a ∣ b) (h₂ : b ∣ c) : a ∣ c
Nat.dvd_mul_left_of_dvd
theorem
{a b : Nat} (h : a ∣ b) (c : Nat) : a ∣ c * b
protected theorem dvd_mul_left_of_dvd {a b : Nat} (h : a ∣ b) (c : Nat) : a ∣ c * b :=
Nat.dvd_trans h (Nat.dvd_mul_left _ _)
Nat.dvd_mul_right_of_dvd
theorem
{a b : Nat} (h : a ∣ b) (c : Nat) : a ∣ b * c
protected theorem dvd_mul_right_of_dvd {a b : Nat} (h : a ∣ b) (c : Nat) : a ∣ b * c :=
Nat.dvd_trans h (Nat.dvd_mul_right _ _)
Nat.zero_dvd
theorem
{n : Nat} : 0 ∣ n ↔ n = 0
@[simp] protected theorem zero_dvd {n : Nat} : 0 ∣ n0 ∣ n ↔ n = 0 :=
⟨Nat.eq_zero_of_zero_dvd, fun h => h.symm ▸ Nat.dvd_zero 0⟩
Nat.dvd_add
theorem
{a b c : Nat} (h₁ : a ∣ b) (h₂ : a ∣ c) : a ∣ b + c
Nat.dvd_add_iff_right
theorem
{k m n : Nat} (h : k ∣ m) : k ∣ n ↔ k ∣ m + n
protected theorem dvd_add_iff_right {k m n : Nat} (h : k ∣ m) : k ∣ nk ∣ n ↔ k ∣ m + n :=
⟨Nat.dvd_add h,
match m, h with
| _, ⟨d, rfl⟩ => fun ⟨e, he⟩ =>
⟨e - d, by rw [Nat.mul_sub_left_distrib, ← he, Nat.add_sub_cancel_left]⟩⟩
Nat.dvd_add_iff_left
theorem
{k m n : Nat} (h : k ∣ n) : k ∣ m ↔ k ∣ m + n
protected theorem dvd_add_iff_left {k m n : Nat} (h : k ∣ n) : k ∣ mk ∣ m ↔ k ∣ m + n := by
rw [Nat.add_comm]; exact Nat.dvd_add_iff_right h
Nat.dvd_mod_iff
theorem
{k m n : Nat} (h : k ∣ n) : k ∣ m % n ↔ k ∣ m
theorem dvd_mod_iff {k m n : Nat} (h: k ∣ n) : k ∣ m % nk ∣ m % n ↔ k ∣ m := by
have := Nat.dvd_add_iff_left (m := m % n) <| Nat.dvd_trans h <| Nat.dvd_mul_right n (m / n)
rwa [mod_add_div] at this
Nat.le_of_dvd
theorem
{m n : Nat} (h : 0 < n) : m ∣ n → m ≤ n
theorem le_of_dvd {m n : Nat} (h : 0 < n) : m ∣ n → m ≤ n
| ⟨k, e⟩ => by
revert h
rw [e]
match k with
| 0 => intro hn; simp at hn
| pk+1 =>
intro
have := Nat.mul_le_mul_left m (succ_pos pk)
rwa [Nat.mul_one] at this
Nat.dvd_antisymm
theorem
: ∀ {m n : Nat}, m ∣ n → n ∣ m → m = n
protected theorem dvd_antisymm : ∀ {m n : Nat}, m ∣ n → n ∣ m → m = n
| _, 0, _, h₂ => Nat.eq_zero_of_zero_dvd h₂
| 0, _, h₁, _ => (Nat.eq_zero_of_zero_dvd h₁).symm
| _+1, _+1, h₁, h₂ => Nat.le_antisymm (le_of_dvd (succ_pos _) h₁) (le_of_dvd (succ_pos _) h₂)
Nat.pos_of_dvd_of_pos
theorem
{m n : Nat} (H1 : m ∣ n) (H2 : 0 < n) : 0 < m
theorem pos_of_dvd_of_pos {m n : Nat} (H1 : m ∣ n) (H2 : 0 < n) : 0 < m :=
Nat.pos_of_ne_zero fun m0 => Nat.ne_of_gt H2 <| Nat.eq_zero_of_zero_dvd (m0 ▸ H1)
Nat.one_dvd
theorem
(n : Nat) : 1 ∣ n
@[simp] protected theorem one_dvd (n : Nat) : 1 ∣ n := ⟨n, n.one_mul.symm⟩
Nat.mod_eq_zero_of_dvd
theorem
{m n : Nat} (H : m ∣ n) : n % m = 0
theorem mod_eq_zero_of_dvd {m n : Nat} (H : m ∣ n) : n % m = 0 := by
let ⟨z, H⟩ := H; rw [H, mul_mod_right]
Nat.dvd_of_mod_eq_zero
theorem
{m n : Nat} (H : n % m = 0) : m ∣ n
theorem dvd_of_mod_eq_zero {m n : Nat} (H : n % m = 0) : m ∣ n := by
exists n / m
have := (mod_add_div n m).symm
rwa [H, Nat.zero_add] at this
Nat.dvd_iff_mod_eq_zero
theorem
{m n : Nat} : m ∣ n ↔ n % m = 0
theorem dvd_iff_mod_eq_zero {m n : Nat} : m ∣ nm ∣ n ↔ n % m = 0 :=
⟨mod_eq_zero_of_dvd, dvd_of_mod_eq_zero⟩
Nat.decidable_dvd
instance
: @DecidableRel Nat Nat (· ∣ ·)
instance decidable_dvd : @DecidableRel Nat Nat (·∣·) :=
fun _ _ => decidable_of_decidable_of_iff dvd_iff_mod_eq_zero.symm
Nat.emod_pos_of_not_dvd
theorem
{a b : Nat} (h : ¬a ∣ b) : 0 < b % a
theorem emod_pos_of_not_dvd {a b : Nat} (h : ¬ a ∣ b) : 0 < b % a := by
rw [dvd_iff_mod_eq_zero] at h
exact Nat.pos_of_ne_zero h
Nat.mul_div_cancel'
theorem
{n m : Nat} (H : n ∣ m) : n * (m / n) = m
protected theorem mul_div_cancel' {n m : Nat} (H : n ∣ m) : n * (m / n) = m := by
have := mod_add_div m n
rwa [mod_eq_zero_of_dvd H, Nat.zero_add] at this
Nat.div_mul_cancel
theorem
{n m : Nat} (H : n ∣ m) : m / n * n = m
protected theorem div_mul_cancel {n m : Nat} (H : n ∣ m) : m / n * n = m := by
rw [Nat.mul_comm, Nat.mul_div_cancel' H]
Nat.mod_mod_of_dvd
theorem
(a : Nat) (h : c ∣ b) : a % b % c = a % c
@[simp] theorem mod_mod_of_dvd (a : Nat) (h : c ∣ b) : a % b % c = a % c := by
rw (occs := [2]) [← mod_add_div a b]
have ⟨x, h⟩ := h
subst h
rw [Nat.mul_assoc, add_mul_mod_self_left]
Nat.dvd_of_mul_dvd_mul_left
theorem
(kpos : 0 < k) (H : k * m ∣ k * n) : m ∣ n
protected theorem dvd_of_mul_dvd_mul_left
(kpos : 0 < k) (H : k * m ∣ k * n) : m ∣ n := by
let ⟨l, H⟩ := H
rw [Nat.mul_assoc] at H
exact ⟨_, Nat.eq_of_mul_eq_mul_left kpos H⟩
Nat.dvd_of_mul_dvd_mul_right
theorem
(kpos : 0 < k) (H : m * k ∣ n * k) : m ∣ n
protected theorem dvd_of_mul_dvd_mul_right (kpos : 0 < k) (H : m * k ∣ n * k) : m ∣ n := by
rw [Nat.mul_comm m k, Nat.mul_comm n k] at H; exact Nat.dvd_of_mul_dvd_mul_left kpos H
Nat.dvd_sub
theorem
{k m n : Nat} (h₁ : k ∣ m) (h₂ : k ∣ n) : k ∣ m - n
Nat.dvd_sub_iff_right
theorem
{m n k : Nat} (hkn : k ≤ n) (h : m ∣ n) : m ∣ n - k ↔ m ∣ k
theorem dvd_sub_iff_right {m n k : Nat} (hkn : k ≤ n) (h : m ∣ n) : m ∣ n - km ∣ n - k ↔ m ∣ k := by
refine ⟨?_, dvd_sub h⟩
let ⟨x, hx⟩ := h
cases hx
intro hy
let ⟨y, hy⟩ := hy
have hk : k = m * (x - y) := by
rw [Nat.sub_eq_iff_eq_add hkn] at hy
rw [Nat.mul_sub, hy, Nat.add_comm, Nat.add_sub_cancel]
exact hk ▸ Nat.dvd_mul_right _ _
Nat.dvd_sub_iff_left
theorem
{m n k : Nat} (hkn : k ≤ n) (h : m ∣ k) : m ∣ n - k ↔ m ∣ n
theorem dvd_sub_iff_left {m n k : Nat} (hkn : k ≤ n) (h : m ∣ k) : m ∣ n - km ∣ n - k ↔ m ∣ n := by
rw (occs := [2]) [← Nat.sub_add_cancel hkn]
exact Nat.dvd_add_iff_left h
Nat.mul_dvd_mul
theorem
{a b c d : Nat} : a ∣ b → c ∣ d → a * c ∣ b * d
protected theorem mul_dvd_mul {a b c d : Nat} : a ∣ b → c ∣ d → a * c ∣ b * d
| ⟨e, he⟩, ⟨f, hf⟩ =>
⟨e * f, by simp [he, hf, Nat.mul_assoc, Nat.mul_left_comm, Nat.mul_comm]⟩
Nat.mul_dvd_mul_left
theorem
(a : Nat) (h : b ∣ c) : a * b ∣ a * c
protected theorem mul_dvd_mul_left (a : Nat) (h : b ∣ c) : a * b ∣ a * c :=
Nat.mul_dvd_mul (Nat.dvd_refl a) h
Nat.mul_dvd_mul_right
theorem
(h : a ∣ b) (c : Nat) : a * c ∣ b * c
protected theorem mul_dvd_mul_right (h: a ∣ b) (c : Nat) : a * c ∣ b * c :=
Nat.mul_dvd_mul h (Nat.dvd_refl c)
Nat.dvd_one
theorem
{n : Nat} : n ∣ 1 ↔ n = 1
@[simp] theorem dvd_one {n : Nat} : n ∣ 1n ∣ 1 ↔ n = 1 :=
⟨eq_one_of_dvd_one, fun h => h.symm ▸ Nat.dvd_refl _⟩
Nat.mul_div_assoc
theorem
(m : Nat) (H : k ∣ n) : m * n / k = m * (n / k)
protected theorem mul_div_assoc (m : Nat) (H : k ∣ n) : m * n / k = m * (n / k) := by
match Nat.eq_zero_or_pos k with
| .inl h0 => rw [h0, Nat.div_zero, Nat.div_zero, Nat.mul_zero]
| .inr hpos =>
have h1 : m * n / k = m * (n / k * k) / k := by rw [Nat.div_mul_cancel H]
rw [h1, ← Nat.mul_assoc, Nat.mul_div_cancel _ hpos]
Nat.dvd_eq_true_of_mod_eq_zero
theorem
{m n : Nat} (h : n % m == 0) : (m ∣ n) = True
protected theorem dvd_eq_true_of_mod_eq_zero {m n : Nat} (h : n % m == 0) : (m ∣ n) = True := by
simp [Nat.dvd_of_mod_eq_zero, eq_of_beq h]
Nat.dvd_eq_false_of_mod_ne_zero
theorem
{m n : Nat} (h : n % m != 0) : (m ∣ n) = False
protected theorem dvd_eq_false_of_mod_ne_zero {m n : Nat} (h : n % m != 0) : (m ∣ n) = False := by
simp [eq_of_beq] at h
simp [dvd_iff_mod_eq_zero, h]