Documentation

Mathlib.NumberTheory.Divisors

Divisor Finsets #

This file defines sets of divisors of a natural number. This is particularly useful as background for defining Dirichlet convolution.

Main Definitions #

Let n : ℕ. All of the following definitions are in the Nat namespace:

Conventions #

Since 0 has infinitely many divisors, none of the definitions in this file make sense for it. Therefore we adopt the convention that Nat.divisors 0, Nat.properDivisors 0, Nat.divisorsAntidiagonal 0 and Int.divisorsAntidiag 0 are all .

Tags #

divisors, perfect numbers

divisors n is the Finset of divisors of n. By convention, we set divisors 0 = ∅.

Equations
Instances For

    properDivisors n is the Finset of divisors of n, other than n. By convention, we set properDivisors 0 = ∅.

    Equations
    Instances For

      Pairs of divisors of a natural number as a finset.

      n.divisorsAntidiagonal is the finset of pairs (a, b) : ℕ × ℕ such that a * b = n. By convention, we set Nat.divisorsAntidiagonal 0 = ∅.

      O(n).

      Equations
      Instances For
        @[simp]
        theorem Nat.filter_dvd_eq_divisors {n : } (h : n 0) :
        Finset.filter (fun (d : ) => d n) (Finset.range n.succ) = n.divisors
        @[simp]
        theorem Nat.filter_dvd_eq_properDivisors {n : } (h : n 0) :
        @[simp]
        theorem Nat.mem_properDivisors {n m : } :
        @[simp]
        theorem Nat.mem_divisors {n m : } :
        n m.divisors n m m 0
        theorem Nat.mem_divisors_self (n : ) (h : n 0) :
        theorem Nat.dvd_of_mem_divisors {n m : } (h : n m.divisors) :
        n m
        @[simp]
        theorem Nat.mem_divisorsAntidiagonal {n : } {x : × } :
        x n.divisorsAntidiagonal x.1 * x.2 = n n 0
        theorem Nat.divisor_le {n m : } :
        n m.divisorsn m
        theorem Nat.divisors_subset_of_dvd {n m : } (hzero : n 0) (h : m n) :
        theorem Nat.divisors_subset_properDivisors {n m : } (hzero : n 0) (h : m n) (hdiff : m n) :
        theorem Nat.divisors_filter_dvd_of_dvd {n m : } (hn : n 0) (hm : m n) :
        Finset.filter (fun (d : ) => d m) n.divisors = m.divisors
        @[simp]
        theorem Nat.divisors_eq_empty {n : } :
        @[simp]
        theorem Nat.pos_of_mem_divisors {n m : } (h : m n.divisors) :
        0 < m
        @[simp]
        theorem Nat.sup_divisors_id (n : ) :
        theorem Nat.mem_properDivisors_iff_exists {m n : } (hn : n 0) :
        m n.properDivisors k > 1, n = m * k

        See also Nat.mem_properDivisors.

        @[deprecated Nat.swap_mem_divisorsAntidiagonal (since := "2025-02-17")]

        Nat.swap_mem_divisorsAntidiagonal with the LHS in simp normal form.

        theorem Nat.map_div_right_divisors {n : } :
        Finset.map { toFun := fun (d : ) => (d, n / d), inj' := } n.divisors = n.divisorsAntidiagonal
        theorem Nat.map_div_left_divisors {n : } :
        Finset.map { toFun := fun (d : ) => (n / d, d), inj' := } n.divisors = n.divisorsAntidiagonal
        def Nat.Perfect (n : ) :

        n : ℕ is perfect if and only the sum of the proper divisors of n is n and n is positive.

        Equations
        Instances For
          theorem Nat.perfect_iff_sum_properDivisors {n : } (h : 0 < n) :
          n.Perfect in.properDivisors, i = n
          theorem Nat.perfect_iff_sum_divisors_eq_two_mul {n : } (h : 0 < n) :
          n.Perfect in.divisors, i = 2 * n
          theorem Nat.mem_divisors_prime_pow {p : } (pp : Prime p) (k : ) {x : } :
          x (p ^ k).divisors jk, x = p ^ j
          theorem Nat.Prime.divisors {p : } (pp : Prime p) :
          theorem Nat.divisors_prime_pow {p : } (pp : Prime p) (k : ) :
          (p ^ k).divisors = Finset.map { toFun := fun (x : ) => p ^ x, inj' := } (Finset.range (k + 1))
          @[simp]
          theorem Nat.divisors_inj {a b : } :
          theorem Nat.eq_properDivisors_of_subset_of_sum_eq_sum {n : } {s : Finset } (hsub : s n.properDivisors) :
          xs, x = xn.properDivisors, xs = n.properDivisors
          theorem Nat.sum_properDivisors_dvd {n : } (h : xn.properDivisors, x n) :
          xn.properDivisors, x = 1 xn.properDivisors, x = n
          @[simp]
          theorem Nat.Prime.prod_properDivisors {α : Type u_1} [CommMonoid α] {p : } {f : α} (h : Prime p) :
          xp.properDivisors, f x = f 1
          @[simp]
          theorem Nat.Prime.sum_properDivisors {α : Type u_1} [AddCommMonoid α] {p : } {f : α} (h : Prime p) :
          xp.properDivisors, f x = f 1
          @[simp]
          theorem Nat.Prime.prod_divisors {α : Type u_1} [CommMonoid α] {p : } {f : α} (h : Prime p) :
          xp.divisors, f x = f p * f 1
          @[simp]
          theorem Nat.Prime.sum_divisors {α : Type u_1} [AddCommMonoid α] {p : } {f : α} (h : Prime p) :
          xp.divisors, f x = f p + f 1
          theorem Nat.mem_properDivisors_prime_pow {p : } (pp : Prime p) (k : ) {x : } :
          x (p ^ k).properDivisors ∃ (j : ) (_ : j < k), x = p ^ j
          theorem Nat.properDivisors_prime_pow {p : } (pp : Prime p) (k : ) :
          (p ^ k).properDivisors = Finset.map { toFun := fun (x : ) => p ^ x, inj' := } (Finset.range k)
          @[simp]
          theorem Nat.prod_properDivisors_prime_pow {α : Type u_1} [CommMonoid α] {k p : } {f : α} (h : Prime p) :
          x(p ^ k).properDivisors, f x = xFinset.range k, f (p ^ x)
          @[simp]
          theorem Nat.sum_properDivisors_prime_nsmul {α : Type u_1} [AddCommMonoid α] {k p : } {f : α} (h : Prime p) :
          x(p ^ k).properDivisors, f x = xFinset.range k, f (p ^ x)
          @[simp]
          theorem Nat.prod_divisors_prime_pow {α : Type u_1} [CommMonoid α] {k p : } {f : α} (h : Prime p) :
          x(p ^ k).divisors, f x = xFinset.range (k + 1), f (p ^ x)
          @[simp]
          theorem Nat.sum_divisors_prime_pow {α : Type u_1} [AddCommMonoid α] {k p : } {f : α} (h : Prime p) :
          x(p ^ k).divisors, f x = xFinset.range (k + 1), f (p ^ x)
          theorem Nat.prod_divisorsAntidiagonal {M : Type u_1} [CommMonoid M] (f : M) {n : } :
          in.divisorsAntidiagonal, f i.1 i.2 = in.divisors, f i (n / i)
          theorem Nat.sum_divisorsAntidiagonal {M : Type u_1} [AddCommMonoid M] (f : M) {n : } :
          in.divisorsAntidiagonal, f i.1 i.2 = in.divisors, f i (n / i)
          theorem Nat.prod_divisorsAntidiagonal' {M : Type u_1} [CommMonoid M] (f : M) {n : } :
          in.divisorsAntidiagonal, f i.1 i.2 = in.divisors, f (n / i) i
          theorem Nat.sum_divisorsAntidiagonal' {M : Type u_1} [AddCommMonoid M] (f : M) {n : } :
          in.divisorsAntidiagonal, f i.1 i.2 = in.divisors, f (n / i) i

          The factors of n are the prime divisors

          theorem Nat.primeFactors_filter_dvd_of_dvd {m n : } (hn : n 0) (hmn : m n) :
          @[simp]
          theorem Nat.prod_div_divisors {α : Type u_1} [CommMonoid α] (n : ) (f : α) :
          dn.divisors, f (n / d) = n.divisors.prod f
          theorem Nat.sum_div_divisors {α : Type u_1} [AddCommMonoid α] (n : ) (f : α) :
          dn.divisors, f (n / d) = n.divisors.sum f

          Pairs of divisors of an integer as a finset.

          z.divisorsAntidiag is the finset of pairs (a, b) : ℤ × ℤ such that a * b = z. By convention, we set Int.divisorsAntidiag 0 = ∅.

          O(|z|). Computed from Nat.divisorsAntidiagonal.

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For
            @[simp]
            theorem Int.mem_divisorsAntidiag {z : } {xy : × } :
            xy z.divisorsAntidiag xy.1 * xy.2 = z z 0
            theorem Int.prodMk_mem_divisorsAntidiag {x y z : } (hz : z 0) :
            theorem Int.mul_mem_one_two_three_iff {a b : } :
            a * b {1, 2, 3} (a, b) {(1, 1), (-1, -1), (1, 2), (2, 1), (-1, -2), (-2, -1), (1, 3), (3, 1), (-1, -3), (-3, -1)}

            This lemma justifies its existence from its utility in crystallographic root system theory.

            theorem Int.mul_mem_zero_one_two_three_four_iff {a b : } (h₀ : a = 0 b = 0) :
            a * b {0, 1, 2, 3, 4} (a, b) {(0, 0), (1, 1), (-1, -1), (1, 2), (2, 1), (-1, -2), (-2, -1), (1, 3), (3, 1), (-1, -3), (-3, -1), (4, 1), (1, 4), (-4, -1), (-1, -4), (2, 2), (-2, -2)}

            This lemma justifies its existence from its utility in crystallographic root system theory.