Sets of tuples with a fixed product #
This file defines the finite set of d
-tuples of natural numbers with a fixed product n
as
Nat.finMulAntidiag
.
Main Results #
- There are
d^(ω n)
ways to writen
as a product ofd
natural numbers, whenn
is squarefree (card_finMulAntidiag_of_squarefree
) - There are
3^(ω n)
pairs of natural numbers whoselcm
isn
, whenn
is squarefree (card_pair_lcm_eq
)
Equations
- One or more equations did not get rendered due to their size.
theorem
Nat.finMulAntidiag_eq_piFinset_divisors_filter
{d m n : ℕ}
(hmn : m ∣ n)
(hn : n ≠ 0)
:
d.finMulAntidiag m = Finset.filter (fun (f : Fin d → ℕ) => ∏ i : Fin d, f i = m) (Fintype.piFinset fun (x : Fin d) => n.divisors)
theorem
Nat.image_apply_finMulAntidiag
{d n : ℕ}
{i : Fin d}
(hd : d ≠ 1)
:
Finset.image (fun (f : Fin d → ℕ) => f i) (d.finMulAntidiag n) = n.divisors
theorem
Nat.image_piFinTwoEquiv_finMulAntidiag
{n : ℕ}
:
Finset.image (⇑(piFinTwoEquiv fun (x : Fin 2) => ℕ)) (Nat.finMulAntidiag 2 n) = n.divisorsAntidiagonal
theorem
Nat.finMulAntidiag_existsUnique_prime_dvd
{d n p : ℕ}
(hn : Squarefree n)
(hp : p ∈ n.primeFactorsList)
(f : Fin d → ℕ)
(hf : f ∈ d.finMulAntidiag n)
:
@[deprecated Nat.finMulAntidiag_existsUnique_prime_dvd]
theorem
Nat.finMulAntidiag_exists_unique_prime_dvd
{d n p : ℕ}
(hn : Squarefree n)
(hp : p ∈ n.primeFactorsList)
(f : Fin d → ℕ)
(hf : f ∈ d.finMulAntidiag n)
:
theorem
Nat.card_finMulAntidiag_of_squarefree
{d n : ℕ}
(hn : Squarefree n)
:
(d.finMulAntidiag n).card = d ^ ArithmeticFunction.cardDistinctFactors n
theorem
Nat.finMulAntidiag_three
{n : ℕ}
(a : Fin 3 → ℕ)
:
a ∈ Nat.finMulAntidiag 3 n → a 0 * a 1 * a 2 = n
The following private declarations are ingredients for the proof of card_pair_lcm_eq
.
theorem
Nat.card_pair_lcm_eq
{n : ℕ}
(hn : Squarefree n)
:
(Finset.filter (fun (p : ℕ × ℕ) => p.1.lcm p.2 = n) (n.divisors ×ˢ n.divisors)).card = 3 ^ ArithmeticFunction.cardDistinctFactors n