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
@[irreducible]
@[irreducible]
theorem
Nat.primeFactorsList_chain
{n a : ℕ}
:
(∀ (p : ℕ), Prime p → p ∣ n → a ≤ p) → List.Chain (fun (x1 x2 : ℕ) => x1 ≤ x2) a n.primeFactorsList
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
primeFactorsList
can be constructed inductively by extracting minFac
, for sufficiently
large n
.
theorem
Nat.eq_of_perm_primeFactorsList
{a b : ℕ}
(ha : a ≠ 0)
(hb : b ≠ 0)
(h : a.primeFactorsList.Perm b.primeFactorsList)
:
theorem
Nat.primeFactorsList_unique
{n : ℕ}
{l : List ℕ}
(h₁ : l.prod = n)
(h₂ : ∀ p ∈ l, Prime p)
:
l.Perm n.primeFactorsList
Fundamental theorem of arithmetic
theorem
Nat.perm_primeFactorsList_mul
{a b : ℕ}
(ha : a ≠ 0)
(hb : b ≠ 0)
:
(a * b).primeFactorsList.Perm (a.primeFactorsList ++ b.primeFactorsList)
For positive a
and b
, the prime factors of a * b
are the union of those of a
and b
theorem
Nat.perm_primeFactorsList_mul_of_coprime
{a b : ℕ}
(hab : a.Coprime b)
:
(a * b).primeFactorsList.Perm (a.primeFactorsList ++ b.primeFactorsList)
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.dvd_of_primeFactorsList_subperm
{a b : ℕ}
(ha : a ≠ 0)
(h : a.primeFactorsList.Subperm b.primeFactorsList)
:
The sets of factors of coprime a
and b
are disjoint
If p
is a prime factor of a
then p
is also a prime factor of a * b
for any b > 0
theorem
Nat.mem_primeFactorsList_mul_right
{p a b : ℕ}
(hpb : p ∈ b.primeFactorsList)
(ha : a ≠ 0)
:
If p
is a prime factor of b
then p
is also a prime factor of a * b
for any a > 0