mathlib documentation

data.​nat.​multiplicity

data.​nat.​multiplicity

Natural number multiplicity

This file contains lemmas about the multiplicity function (the maximum prime power divding a number).

Main results

There are natural number versions of some basic lemmas about multiplicity.

There are also lemmas about the multiplicity of primes in factorials and in binomial coefficients.

theorem nat.​multiplicity_eq_card_pow_dvd {m n b : } :
m 10 < nn bmultiplicity m n = ((finset.filter (λ (i : ), m ^ i n) (finset.Ico 1 b)).card)

The multiplicity of a divisor m of n, is the cardinality of the set of positive natural numbers i such that p ^ i divides n. The set is expressed by filtering Ico 1 b where b is any bound at least n

theorem nat.​prime.​multiplicity_fact {p : } (hp : nat.prime p) {n b : } :
n bmultiplicity p n.fact = ((finset.Ico 1 b).sum (λ (i : ), n / p ^ i))

The multiplicity of a prime in fact n is the sum of the quotients n / p ^ i. This sum is expressed over the set Ico 1 b where b is any bound at least n

theorem nat.​prime.​pow_dvd_fact_iff {p n r b : } :
nat.prime pn b(p ^ r n.fact r (finset.Ico 1 b).sum (λ (i : ), n / p ^ i))

A prime power divides fact 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 at least n

theorem nat.​prime.​multiplicity_choose_aux {p n b k : } :
nat.prime pk n(finset.Ico 1 b).sum (λ (i : ), n / p ^ i) = (finset.Ico 1 b).sum (λ (i : ), k / p ^ i) + (finset.Ico 1 b).sum (λ (i : ), (n - k) / p ^ i) + (finset.filter (λ (i : ), p ^ i k % p ^ i + (n - k) % p ^ i) (finset.Ico 1 b)).card

theorem nat.​prime.​multiplicity_choose {p n k b : } :
nat.prime pk nn bmultiplicity p (n.choose k) = ((finset.filter (λ (i : ), p ^ i k % p ^ i + (n - k) % p ^ i) (finset.Ico 1 b)).card)

The multiplity 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 at least n.

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

theorem nat.​prime.​multiplicity_choose_prime_pow {p n k : } :
nat.prime pk p ^ n0 < kmultiplicity p ((p ^ n).choose k) + multiplicity p k = n