Prime numbers #
This file deals with the factors of natural numbers.
Important declarations #
Nat.factors n
: the prime factorization ofn
Nat.factors_unique
: uniqueness of the prime factorisation
@[irreducible]
primeFactorsList n
is the prime factorization of n
, listed in increasing order.
Equations
- Nat.primeFactorsList 0 = []
- Nat.primeFactorsList 1 = []
- k.succ.succ.primeFactorsList = (k + 2).minFac :: ((k + 2) / (k + 2).minFac).primeFactorsList
Instances For
@[deprecated Nat.primeFactorsList]
Alias of Nat.primeFactorsList
.
primeFactorsList n
is the prime factorization of n
, listed in increasing order.
Equations
Instances For
@[irreducible]
@[irreducible]
theorem
Nat.primeFactorsList_chain_2
(n : ℕ)
:
List.Chain (fun (x1 x2 : ℕ) => x1 ≤ x2) 2 n.primeFactorsList
theorem
Nat.primeFactorsList_chain'
(n : ℕ)
:
List.Chain' (fun (x1 x2 : ℕ) => x1 ≤ x2) n.primeFactorsList
theorem
Nat.primeFactorsList_sorted
(n : ℕ)
:
List.Sorted (fun (x1 x2 : ℕ) => x1 ≤ x2) n.primeFactorsList
theorem
Nat.Prime.primeFactorsList_pow
{p : ℕ}
(hp : Nat.Prime p)
(n : ℕ)
:
(p ^ n).primeFactorsList = List.replicate n p
For coprime a
and b
, the prime factors of a * b
are the union of those of a
and b
theorem
Nat.primeFactorsList_sublist_right
{n k : ℕ}
(h : k ≠ 0)
:
n.primeFactorsList.Sublist (n * k).primeFactorsList
theorem
Nat.primeFactorsList_sublist_of_dvd
{n k : ℕ}
(h : n ∣ k)
(h' : k ≠ 0)
:
n.primeFactorsList.Sublist k.primeFactorsList
theorem
Nat.dvd_of_primeFactorsList_subperm
{a b : ℕ}
(ha : a ≠ 0)
(h : a.primeFactorsList.Subperm b.primeFactorsList)
:
a ∣ b
theorem
Nat.replicate_subperm_primeFactorsList_iff
{a b n : ℕ}
(ha : Nat.Prime a)
(hb : b ≠ 0)
:
(List.replicate n a).Subperm b.primeFactorsList ↔ a ^ n ∣ b
theorem
Nat.coprime_primeFactorsList_disjoint
{a b : ℕ}
(hab : a.Coprime b)
:
a.primeFactorsList.Disjoint b.primeFactorsList
The sets of factors of coprime a
and b
are disjoint