Basic lemmas on prime factorizations #
Basic facts about factorization #
Lemmas characterising when n.factorization p = 0
#
The only numbers with empty prime factorization are 0
and 1
Lemmas about factorizations of products and powers #
A product over n.factorization
can be written as a product over n.primeFactors
;
A product over n.primeFactors
can be written as a product over n.factorization
;
Lemmas about factorizations of primes and prime powers #
The multiplicity of prime p
in p
is 1
If the factorization of n
contains just one number p
then n
is a power of p
The only prime factor of prime p
is p
itself.
Equivalence between ℕ+
and ℕ →₀ ℕ
with support in the primes. #
Alias of Nat.ordProj_of_not_prime
.
Alias of Nat.ordCompl_of_not_prime
.
Alias of Nat.ordCompl_dvd
.
Alias of Nat.ordProj_pos
.
Alias of Nat.ordProj_le
.
Alias of Nat.ordCompl_pos
.
Alias of Nat.ordCompl_le
.
Alias of Nat.ordProj_mul_ordCompl_eq_self
.
Alias of Nat.ordProj_mul
.
Alias of Nat.ordCompl_mul
.
Factorization and divisibility #
A crude upper bound on n.factorization p
An upper bound on n.factorization p
Alias of Nat.Prime.pow_dvd_iff_dvd_ordProj
.
Alias of Nat.dvd_ordProj_of_dvd
.
Alias of Nat.not_dvd_ordCompl
.
Alias of Nat.coprime_ordCompl
.
Alias of Nat.factorization_ordCompl
.
Alias of Nat.dvd_ordCompl_of_dvd_not_dvd
.
Alias of Nat.ordProj_dvd_ordProj_of_dvd
.
Alias of Nat.ordProj_dvd_ordProj_iff_dvd
.
Alias of Nat.ordCompl_dvd_ordCompl_of_dvd
.
Alias of Nat.ordCompl_dvd_ordCompl_iff_dvd
.
The set of positive powers of prime p
that divide n
is exactly the set of
positive natural numbers up to n.factorization p
.
Factorization and coprimes #
If p
is a prime factor of a
then the power of p
in a
is the same that in a * b
,
for any b
coprime to a
.
If p
is a prime factor of b
then the power of p
in b
is the same that in a * b
,
for any a
coprime to b
.
Two positive naturals are equal if their prime padic valuations are equal
Lemmas about factorizations of particular functions #
Exactly n / p
naturals in [1, n]
are multiples of p
.
See Nat.card_multiples'
for an alternative spelling of the statement.
Exactly n / p
naturals in (0, n]
are multiples of p
.
There are exactly ⌊N/n⌋
positive multiples of n
that are ≤ N
.
See Nat.card_multiples
for a "shifted-by-one" version.