Documentation

Mathlib.Data.Nat.Multiplicity

Natural number multiplicity #

This file contains lemmas about the multiplicity function (the maximum prime power dividing a number) when applied to naturals, in particular calculating it for factorials and binomial coefficients.

Multiplicity calculations #

Other declarations #

Tags #

Legendre, p-adic

theorem Nat.multiplicity_eq_card_pow_dvd {m : } {n : } {b : } (hm : m 1) (hn : 0 < n) (hb : m.log n < b) :
multiplicity m n = (Finset.filter (fun (i : ) => m ^ i n) (Finset.Ico 1 b)).card

The multiplicity of m in n is the number of positive natural numbers i such that m ^ i divides n. This set is expressed by filtering Ico 1 b where b is any bound greater than log m n.

theorem Nat.Prime.multiplicity_one {p : } (hp : p.Prime) :
theorem Nat.Prime.multiplicity_mul {p : } {m : } {n : } (hp : p.Prime) :
theorem Nat.Prime.multiplicity_pow {p : } {m : } {n : } (hp : p.Prime) :
theorem Nat.Prime.multiplicity_self {p : } (hp : p.Prime) :
theorem Nat.Prime.multiplicity_pow_self {p : } {n : } (hp : p.Prime) :
multiplicity p (p ^ n) = n
theorem Nat.Prime.multiplicity_factorial {p : } (hp : p.Prime) {n : } {b : } :
p.log n < bmultiplicity p n.factorial = ((Finset.Ico 1 b).sum fun (i : ) => n / p ^ i)

Legendre's Theorem

The multiplicity of a prime in n! is the sum of the quotients n / p ^ i. This sum is expressed over the finset Ico 1 b where b is any bound greater than log p n.

theorem Nat.Prime.sub_one_mul_multiplicity_factorial {n : } {p : } (hp : p.Prime) :
(p - 1) * (multiplicity p n.factorial).get = n - (p.digits n).sum

For a prime number p, taking (p - 1) times the multiplicity of p in n! equals n minus the sum of base p digits of n.

theorem Nat.Prime.multiplicity_factorial_mul_succ {n : } {p : } (hp : p.Prime) :
multiplicity p (p * (n + 1)).factorial = multiplicity p (p * n).factorial + multiplicity p (n + 1) + 1

The multiplicity of p in (p * (n + 1))! is one more than the sum of the multiplicities of p in (p * n)! and n + 1.

theorem Nat.Prime.multiplicity_factorial_mul {n : } {p : } (hp : p.Prime) :
multiplicity p (p * n).factorial = multiplicity p n.factorial + n

The multiplicity of p in (p * n)! is n more than that of n!.

theorem Nat.Prime.pow_dvd_factorial_iff {p : } {n : } {r : } {b : } (hp : p.Prime) (hbn : p.log n < b) :
p ^ r n.factorial r (Finset.Ico 1 b).sum fun (i : ) => n / p ^ i

A prime power divides n! iff it is at most the sum of the quotients n / p ^ i. This sum is expressed over the set Ico 1 b where b is any bound greater than log p n

theorem Nat.Prime.multiplicity_factorial_le_div_pred {p : } (hp : p.Prime) (n : ) :
multiplicity p n.factorial (n / (p - 1))
theorem Nat.Prime.multiplicity_choose_aux {p : } {n : } {b : } {k : } (hp : p.Prime) (hkn : k n) :
((Finset.Ico 1 b).sum fun (i : ) => n / p ^ i) = (((Finset.Ico 1 b).sum fun (i : ) => k / p ^ i) + (Finset.Ico 1 b).sum fun (i : ) => (n - k) / p ^ i) + (Finset.filter (fun (i : ) => p ^ i k % p ^ i + (n - k) % p ^ i) (Finset.Ico 1 b)).card
theorem Nat.Prime.multiplicity_choose' {p : } {n : } {k : } {b : } (hp : p.Prime) (hnb : p.log (n + k) < b) :
multiplicity p ((n + k).choose k) = (Finset.filter (fun (i : ) => p ^ i k % p ^ i + n % p ^ i) (Finset.Ico 1 b)).card

The multiplicity of p in choose (n + k) k is the number of carries when k and n are added in base p. The set is expressed by filtering Ico 1 b where b is any bound greater than log p (n + k).

theorem Nat.Prime.multiplicity_choose {p : } {n : } {k : } {b : } (hp : p.Prime) (hkn : k n) (hnb : p.log n < b) :
multiplicity p (n.choose k) = (Finset.filter (fun (i : ) => p ^ i k % p ^ i + (n - k) % p ^ i) (Finset.Ico 1 b)).card

The multiplicity of p in choose n k is the number of carries when k and n - k are added in base p. The set is expressed by filtering Ico 1 b where b is any bound greater than log p n.

theorem Nat.Prime.multiplicity_le_multiplicity_choose_add {p : } (hp : p.Prime) (n : ) (k : ) :
multiplicity p n multiplicity p (n.choose k) + multiplicity p k

A lower bound on the multiplicity of p in choose n k.

theorem Nat.Prime.multiplicity_choose_prime_pow_add_multiplicity {p : } {n : } {k : } (hp : p.Prime) (hkn : k p ^ n) (hk0 : k 0) :
multiplicity p ((p ^ n).choose k) + multiplicity p k = n
theorem Nat.Prime.multiplicity_choose_prime_pow {p : } {n : } {k : } (hp : p.Prime) (hkn : k p ^ n) (hk0 : k 0) :
multiplicity p ((p ^ n).choose k) = (n - (multiplicity p k).get )
theorem Nat.Prime.dvd_choose_pow {p : } {n : } {k : } (hp : p.Prime) (hk : k 0) (hkp : k p ^ n) :
p (p ^ n).choose k
theorem Nat.Prime.dvd_choose_pow_iff {p : } {n : } {k : } (hp : p.Prime) :
p (p ^ n).choose k k 0 k p ^ n
theorem Nat.multiplicity_two_factorial_lt {n : } :
n 0multiplicity 2 n.factorial < n