Documentation

Mathlib.Algebra.Order.Antidiag.Nat

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 #

Equations
  • One or more equations did not get rendered due to their size.
def Nat.finMulAntidiag (d n : ) :
Finset (Fin d)

The Finset of all d-tuples of natural numbers whose product is n. Defined to be when n = 0.

Equations
  • One or more equations did not get rendered due to their size.
Instances For
    @[simp]
    theorem Nat.mem_finMulAntidiag {d n : } {f : Fin d} :
    f d.finMulAntidiag n i : Fin d, f i = n n 0
    @[simp]
    theorem Nat.finMulAntidiag_zero_right (d : ) :
    d.finMulAntidiag 0 =
    theorem Nat.finMulAntidiag_one {d : } :
    d.finMulAntidiag 1 = {fun (x : Fin d) => 1}
    theorem Nat.dvd_of_mem_finMulAntidiag {n d : } {f : Fin d} (hf : f d.finMulAntidiag n) (i : Fin d) :
    f i n
    theorem Nat.ne_zero_of_mem_finMulAntidiag {d n : } {f : Fin d} (hf : f d.finMulAntidiag n) (i : Fin d) :
    f i 0
    theorem Nat.prod_eq_of_mem_finMulAntidiag {d n : } {f : Fin d} (hf : f d.finMulAntidiag n) :
    i : Fin d, f i = n
    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) :
    ∃! i : Fin d, p f i
    @[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) :
    ∃! i : Fin d, p f i

    Alias of Nat.finMulAntidiag_existsUnique_prime_dvd.

    theorem Nat.finMulAntidiag_three {n : } (a : Fin 3) :
    a Nat.finMulAntidiag 3 na 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