Documentation

Mathlib.RingTheory.UniqueFactorizationDomain

Unique factorization #

Main Definitions #

To do #

class WfDvdMonoid (α : Type u_2) [CommMonoidWithZero α] :

Well-foundedness of the strict version of |, which is equivalent to the descending chain condition on divisibility and to the ascending chain condition on principal ideals in an integral domain.

Instances
    theorem WfDvdMonoid.exists_irreducible_factor {α : Type u_1} [CommMonoidWithZero α] [WfDvdMonoid α] {a : α} (ha : ¬IsUnit a) (ha0 : a 0) :
    i, Irreducible i i a
    theorem WfDvdMonoid.induction_on_irreducible {α : Type u_1} [CommMonoidWithZero α] [WfDvdMonoid α] {P : αProp} (a : α) (h0 : P 0) (hu : (u : α) → IsUnit uP u) (hi : (a i : α) → a 0Irreducible iP aP (i * a)) :
    P a
    theorem WfDvdMonoid.exists_factors {α : Type u_1} [CommMonoidWithZero α] [WfDvdMonoid α] (a : α) :
    a 0f, (∀ (b : α), b fIrreducible b) Associated (Multiset.prod f) a
    theorem WfDvdMonoid.not_unit_iff_exists_factors_eq {α : Type u_1} [CommMonoidWithZero α] [WfDvdMonoid α] (a : α) (hn0 : a 0) :
    ¬IsUnit a f, (∀ (b : α), b fIrreducible b) Multiset.prod f = a f

    unique factorization monoids.

    These are defined as CancelCommMonoidWithZeros with well-founded strict divisibility relations, but this is equivalent to more familiar definitions:

    Each element (except zero) is uniquely represented as a multiset of irreducible factors. Uniqueness is only up to associated elements.

    Each element (except zero) is non-uniquely represented as a multiset of prime factors.

    To define a UFD using the definition in terms of multisets of irreducible factors, use the definition of_exists_unique_irreducible_factors

    To define a UFD using the definition in terms of multisets of prime factors, use the definition of_exists_prime_factors

    Instances
      @[reducible]

      Can't be an instance because it would cause a loop ufmWfDvdMonoidufm → ....

      theorem UniqueFactorizationMonoid.exists_prime_factors {α : Type u_1} [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] (a : α) :
      a 0f, (∀ (b : α), b fPrime b) Associated (Multiset.prod f) a
      theorem UniqueFactorizationMonoid.induction_on_prime {α : Type u_1} [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] {P : αProp} (a : α) (h₁ : P 0) (h₂ : (x : α) → IsUnit xP x) (h₃ : (a p : α) → a 0Prime pP aP (p * a)) :
      P a
      theorem prime_factors_unique {α : Type u_1} [CancelCommMonoidWithZero α] {f : Multiset α} {g : Multiset α} :
      (∀ (x : α), x fPrime x) → (∀ (x : α), x gPrime x) → Associated (Multiset.prod f) (Multiset.prod g)Multiset.Rel Associated f g
      theorem UniqueFactorizationMonoid.factors_unique {α : Type u_1} [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] {f : Multiset α} {g : Multiset α} (hf : ∀ (x : α), x fIrreducible x) (hg : ∀ (x : α), x gIrreducible x) (h : Associated (Multiset.prod f) (Multiset.prod g)) :
      Multiset.Rel Associated f g
      theorem prime_factors_irreducible {α : Type u_1} [CancelCommMonoidWithZero α] {a : α} {f : Multiset α} (ha : Irreducible a) (pfa : (∀ (b : α), b fPrime b) Associated (Multiset.prod f) a) :
      p, Associated a p f = {p}

      If an irreducible has a prime factorization, then it is an associate of one of its prime factors.

      theorem WfDvdMonoid.of_exists_prime_factors {α : Type u_1} [CancelCommMonoidWithZero α] (pf : ∀ (a : α), a 0f, (∀ (b : α), b fPrime b) Associated (Multiset.prod f) a) :
      theorem irreducible_iff_prime_of_exists_prime_factors {α : Type u_1} [CancelCommMonoidWithZero α] (pf : ∀ (a : α), a 0f, (∀ (b : α), b fPrime b) Associated (Multiset.prod f) a) {p : α} :
      theorem UniqueFactorizationMonoid.of_exists_prime_factors {α : Type u_1} [CancelCommMonoidWithZero α] (pf : ∀ (a : α), a 0f, (∀ (b : α), b fPrime b) Associated (Multiset.prod f) a) :
      theorem irreducible_iff_prime_of_exists_unique_irreducible_factors {α : Type u_1} [CancelCommMonoidWithZero α] (eif : ∀ (a : α), a 0f, (∀ (b : α), b fIrreducible b) Associated (Multiset.prod f) a) (uif : ∀ (f g : Multiset α), (∀ (x : α), x fIrreducible x) → (∀ (x : α), x gIrreducible x) → Associated (Multiset.prod f) (Multiset.prod g)Multiset.Rel Associated f g) (p : α) :
      theorem UniqueFactorizationMonoid.of_exists_unique_irreducible_factors {α : Type u_1} [CancelCommMonoidWithZero α] (eif : ∀ (a : α), a 0f, (∀ (b : α), b fIrreducible b) Associated (Multiset.prod f) a) (uif : ∀ (f g : Multiset α), (∀ (x : α), x fIrreducible x) → (∀ (x : α), x gIrreducible x) → Associated (Multiset.prod f) (Multiset.prod g)Multiset.Rel Associated f g) :

      Noncomputably determines the multiset of prime factors.

      Equations
      Instances For

        Noncomputably determines the multiset of prime factors.

        Equations
        Instances For
          @[simp]

          An arbitrary choice of factors of x : M is exactly the (unique) normalized set of factors, if M has a trivial group of units.

          Noncomputably defines a normalizationMonoid structure on a UniqueFactorizationMonoid.

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For
            Equations
            • UniqueFactorizationMonoid.instInhabitedNormalizationMonoid = { default := UniqueFactorizationMonoid.normalizationMonoid }
            theorem UniqueFactorizationMonoid.no_factors_of_no_prime_factors {R : Type u_2} [CancelCommMonoidWithZero R] [UniqueFactorizationMonoid R] {a : R} {b : R} (ha : a 0) (h : ∀ {d : R}, d ad b¬Prime d) {d : R} :
            d ad bIsUnit d
            theorem UniqueFactorizationMonoid.dvd_of_dvd_mul_left_of_no_prime_factors {R : Type u_2} [CancelCommMonoidWithZero R] [UniqueFactorizationMonoid R] {a : R} {b : R} {c : R} (ha : a 0) :
            (∀ {d : R}, d ad c¬Prime d) → a b * ca b

            Euclid's lemma: if a ∣ b * c and a and c have no common prime factors, a ∣ b. Compare IsCoprime.dvd_of_dvd_mul_left.

            theorem UniqueFactorizationMonoid.dvd_of_dvd_mul_right_of_no_prime_factors {R : Type u_2} [CancelCommMonoidWithZero R] [UniqueFactorizationMonoid R] {a : R} {b : R} {c : R} (ha : a 0) (no_factors : ∀ {d : R}, d ad b¬Prime d) :
            a b * ca c

            Euclid's lemma: if a ∣ b * c and a and b have no common prime factors, a ∣ c. Compare IsCoprime.dvd_of_dvd_mul_right.

            theorem UniqueFactorizationMonoid.exists_reduced_factors {R : Type u_2} [CancelCommMonoidWithZero R] [UniqueFactorizationMonoid R] (a : R) :
            a 0∀ (b : R), a' b' c', (∀ {d : R}, d a'd b'IsUnit d) c' * a' = a c' * b' = b

            If a ≠ 0, b are elements of a unique factorization domain, then dividing out their common factor c' gives a' and b' with no factors in common.

            theorem UniqueFactorizationMonoid.exists_reduced_factors' {R : Type u_2} [CancelCommMonoidWithZero R] [UniqueFactorizationMonoid R] (a : R) (b : R) (hb : b 0) :
            a' b' c', (∀ {d : R}, d a'd b'IsUnit d) c' * a' = a c' * b' = b
            theorem UniqueFactorizationMonoid.pow_eq_pow_iff {R : Type u_2} [CancelCommMonoidWithZero R] [UniqueFactorizationMonoid R] {a : R} (ha0 : a 0) (ha1 : ¬IsUnit a) {i : } {j : } :
            a ^ i = a ^ j i = j

            The multiplicity of an irreducible factor of a nonzero element is exactly the number of times the normalized factor occurs in the normalizedFactors.

            See also count_normalizedFactors_eq which expands the definition of multiplicity to produce a specification for count (normalizedFactors _) _..

            theorem UniqueFactorizationMonoid.count_normalizedFactors_eq {R : Type u_2} [CancelCommMonoidWithZero R] [UniqueFactorizationMonoid R] [Nontrivial R] [NormalizationMonoid R] [DecidableEq R] {p : R} {x : R} (hp : Irreducible p) (hnorm : normalize p = p) {n : } (hle : p ^ n x) (hlt : ¬p ^ (n + 1) x) :

            The number of times an irreducible factor p appears in normalizedFactors x is defined by the number of times it divides x.

            See also multiplicity_eq_count_normalizedFactors if n is given by multiplicity p x.

            theorem UniqueFactorizationMonoid.count_normalizedFactors_eq' {R : Type u_2} [CancelCommMonoidWithZero R] [UniqueFactorizationMonoid R] [Nontrivial R] [NormalizationMonoid R] [DecidableEq R] {p : R} {x : R} (hp : p = 0 Irreducible p) (hnorm : normalize p = p) {n : } (hle : p ^ n x) (hlt : ¬p ^ (n + 1) x) :

            The number of times an irreducible factor p appears in normalizedFactors x is defined by the number of times it divides x. This is a slightly more general version of UniqueFactorizationMonoid.count_normalizedFactors_eq that allows p = 0.

            See also multiplicity_eq_count_normalizedFactors if n is given by multiplicity p x.

            theorem UniqueFactorizationMonoid.max_power_factor {R : Type u_2} [CancelCommMonoidWithZero R] [UniqueFactorizationMonoid R] [Nontrivial R] [NormalizationMonoid R] [dec_dvd : DecidableRel Dvd.dvd] {a₀ : R} {x : R} (h : a₀ 0) (hx : Irreducible x) :
            n a, ¬x a a₀ = x ^ n * a
            theorem UniqueFactorizationMonoid.prime_pow_coprime_prod_of_coprime_insert {α : Type u_1} [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] [DecidableEq α] {s : Finset α} (i : α) (p : α) (hps : ¬p s) (is_prime : ∀ (q : α), q insert p sPrime q) (is_coprime : ∀ (q : α), q insert p s∀ (q' : α), q' insert p sq q'q = q') (q : α) :
            q p ^ i p(q Finset.prod s fun p' => p' ^ i p') → IsUnit q
            theorem UniqueFactorizationMonoid.induction_on_prime_power {α : Type u_1} [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] {P : αProp} (s : Finset α) (i : α) (is_prime : ∀ (p : α), p sPrime p) (is_coprime : ∀ (p : α), p s∀ (q : α), q sp qp = q) (h1 : {x : α} → IsUnit xP x) (hpr : {p : α} → (i : ) → Prime pP (p ^ i)) (hcp : {x y : α} → (∀ (p : α), p xp yIsUnit p) → P xP yP (x * y)) :
            P (Finset.prod s fun p => p ^ i p)

            If P holds for units and powers of primes, and P x ∧ P y for coprime x, y implies P (x * y), then P holds on a product of powers of distinct primes.

            theorem UniqueFactorizationMonoid.induction_on_coprime {α : Type u_1} [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] {P : αProp} (a : α) (h0 : P 0) (h1 : {x : α} → IsUnit xP x) (hpr : {p : α} → (i : ) → Prime pP (p ^ i)) (hcp : {x y : α} → (∀ (p : α), p xp yIsUnit p) → P xP yP (x * y)) :
            P a

            If P holds for 0, units and powers of primes, and P x ∧ P y for coprime x, y implies P (x * y), then P holds on all a : α.

            theorem UniqueFactorizationMonoid.multiplicative_prime_power {α : Type u_1} [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] {β : Type u_3} [CancelCommMonoidWithZero β] {f : αβ} (s : Finset α) (i : α) (j : α) (is_prime : ∀ (p : α), p sPrime p) (is_coprime : ∀ (p : α), p s∀ (q : α), q sp qp = q) (h1 : ∀ {x y : α}, IsUnit yf (x * y) = f x * f y) (hpr : ∀ {p : α} (i : ), Prime pf (p ^ i) = f p ^ i) (hcp : ∀ {x y : α}, (∀ (p : α), p xp yIsUnit p) → f (x * y) = f x * f y) :
            f (Finset.prod s fun p => p ^ (i p + j p)) = f (Finset.prod s fun p => p ^ i p) * f (Finset.prod s fun p => p ^ j p)

            If f maps p ^ i to (f p) ^ i for primes p, and f is multiplicative on coprime elements, then f is multiplicative on all products of primes.

            theorem UniqueFactorizationMonoid.multiplicative_of_coprime {α : Type u_1} [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] {β : Type u_3} [CancelCommMonoidWithZero β] (f : αβ) (a : α) (b : α) (h0 : f 0 = 0) (h1 : ∀ {x y : α}, IsUnit yf (x * y) = f x * f y) (hpr : ∀ {p : α} (i : ), Prime pf (p ^ i) = f p ^ i) (hcp : ∀ {x y : α}, (∀ (p : α), p xp yIsUnit p) → f (x * y) = f x * f y) :
            f (a * b) = f a * f b

            If f maps p ^ i to (f p) ^ i for primes p, and f is multiplicative on coprime elements, then f is multiplicative everywhere.

            @[reducible]

            FactorSet α representation elements of unique factorization domain as multisets. Multiset α produced by normalizedFactors are only unique up to associated elements, while the multisets in FactorSet α are unique by equality and restricted to irreducible elements. This gives us a representation of each element as a unique multisets (or the added ⊤ for 0), which has a complete lattice structure. Infimum is the greatest common divisor and supremum is the least common multiple.

            Equations
            Instances For
              theorem Associates.FactorSet.coe_add {α : Type u_1} [CancelCommMonoidWithZero α] {a : Multiset { a // Irreducible a }} {b : Multiset { a // Irreducible a }} :
              ↑(a + b) = a + b

              Evaluates the product of a FactorSet to be the product of the corresponding multiset, or 0 if there is none.

              Equations
              Instances For

                bcount p s is the multiplicity of p in the FactorSet s (with bundled p)

                Equations
                Instances For

                  count p s is the multiplicity of the irreducible p in the FactorSet s.

                  If p is not irreducible, count p s is defined to be 0.

                  Equations
                  Instances For
                    @[simp]
                    theorem Associates.count_some {α : Type u_1} [CancelCommMonoidWithZero α] [dec_irr : (p : Associates α) → Decidable (Irreducible p)] [DecidableEq (Associates α)] {p : Associates α} (hp : Irreducible p) (s : Multiset { a // Irreducible a }) :
                    Associates.count p (some s) = Multiset.count { val := p, property := hp } s
                    @[simp]
                    theorem Associates.count_zero {α : Type u_1} [CancelCommMonoidWithZero α] [dec_irr : (p : Associates α) → Decidable (Irreducible p)] [DecidableEq (Associates α)] {p : Associates α} (hp : Irreducible p) :

                    membership in a FactorSet (bundled version)

                    Equations
                    Instances For

                      FactorSetMem p s is the predicate that the irreducible p is a member of s : FactorSet α.

                      If p is not irreducible, p is not a member of any FactorSet.

                      Equations
                      Instances For
                        Equations
                        • Associates.instMembershipAssociatesToMonoidToMonoidWithZeroToCommMonoidWithZeroFactorSet = { mem := Associates.FactorSetMem }
                        theorem Associates.mem_factorSet_top {α : Type u_1} [CancelCommMonoidWithZero α] [dec_irr : (p : Associates α) → Decidable (Irreducible p)] {p : Associates α} {hp : Irreducible p} :
                        theorem Associates.mem_factorSet_some {α : Type u_1} [CancelCommMonoidWithZero α] [dec_irr : (p : Associates α) → Decidable (Irreducible p)] {p : Associates α} {hp : Irreducible p} {l : Multiset { a // Irreducible a }} :
                        p l { val := p, property := hp } l
                        theorem Associates.unique' {α : Type u_1} [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] {p : Multiset (Associates α)} {q : Multiset (Associates α)} :
                        (∀ (a : Associates α), a pIrreducible a) → (∀ (a : Associates α), a qIrreducible a) → Multiset.prod p = Multiset.prod qp = q
                        theorem Associates.prod_le_prod_iff_le {α : Type u_1} [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] [Nontrivial α] {p : Multiset (Associates α)} {q : Multiset (Associates α)} (hp : ∀ (a : Associates α), a pIrreducible a) (hq : ∀ (a : Associates α), a qIrreducible a) :
                        noncomputable def Associates.factors' {α : Type u_1} [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] [dec : DecidableEq α] (a : α) :

                        This returns the multiset of irreducible factors as a FactorSet, a multiset of irreducible associates WithTop.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For

                          This returns the multiset of irreducible factors of an associate as a FactorSet, a multiset of irreducible associates WithTop.

                          Equations
                          • One or more equations did not get rendered due to their size.
                          Instances For
                            theorem Associates.eq_of_eq_counts {α : Type u_1} [CancelCommMonoidWithZero α] [dec_irr : (p : Associates α) → Decidable (Irreducible p)] [UniqueFactorizationMonoid α] [dec : DecidableEq α] [dec' : DecidableEq (Associates α)] {a : Associates α} {b : Associates α} (ha : a 0) (hb : b 0) (h : ∀ (p : Associates α), Irreducible pAssociates.count p (Associates.factors a) = Associates.count p (Associates.factors b)) :
                            a = b
                            Equations
                            Equations
                            Equations
                            • One or more equations did not get rendered due to their size.
                            theorem Associates.sup_mul_inf {α : Type u_1} [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] [dec : DecidableEq α] [dec' : DecidableEq (Associates α)] (a : Associates α) (b : Associates α) :
                            (a b) * (a b) = a * b
                            theorem Associates.dvd_of_mem_factors {α : Type u_1} [CancelCommMonoidWithZero α] [dec_irr : (p : Associates α) → Decidable (Irreducible p)] [UniqueFactorizationMonoid α] [dec : DecidableEq α] [dec' : DecidableEq (Associates α)] {a : Associates α} {p : Associates α} {hp : Irreducible p} (hm : p Associates.factors a) :
                            p a
                            theorem Associates.dvd_of_mem_factors' {α : Type u_1} [CancelCommMonoidWithZero α] [dec_irr : (p : Associates α) → Decidable (Irreducible p)] [UniqueFactorizationMonoid α] [dec : DecidableEq α] {a : α} {p : Associates α} {hp : Irreducible p} {hz : a 0} (h_mem : { val := p, property := hp } Associates.factors' a) :
                            theorem Associates.mem_factors'_of_dvd {α : Type u_1} [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] [dec : DecidableEq α] {a : α} {p : α} (ha0 : a 0) (hp : Irreducible p) (hd : p a) :
                            { val := Associates.mk p, property := (_ : Irreducible (Associates.mk p)) } Associates.factors' a
                            theorem Associates.mem_factors'_iff_dvd {α : Type u_1} [CancelCommMonoidWithZero α] [dec_irr : (p : Associates α) → Decidable (Irreducible p)] [UniqueFactorizationMonoid α] [dec : DecidableEq α] {a : α} {p : α} (ha0 : a 0) (hp : Irreducible p) :
                            { val := Associates.mk p, property := (_ : Irreducible (Associates.mk p)) } Associates.factors' a p a
                            theorem Associates.mem_factors_of_dvd {α : Type u_1} [CancelCommMonoidWithZero α] [dec_irr : (p : Associates α) → Decidable (Irreducible p)] [UniqueFactorizationMonoid α] [dec : DecidableEq α] [dec' : DecidableEq (Associates α)] {a : α} {p : α} (ha0 : a 0) (hp : Irreducible p) (hd : p a) :
                            theorem Associates.mem_factors_iff_dvd {α : Type u_1} [CancelCommMonoidWithZero α] [dec_irr : (p : Associates α) → Decidable (Irreducible p)] [UniqueFactorizationMonoid α] [dec : DecidableEq α] [dec' : DecidableEq (Associates α)] {a : α} {p : α} (ha0 : a 0) (hp : Irreducible p) :
                            theorem Associates.exists_prime_dvd_of_not_inf_one {α : Type u_1} [CancelCommMonoidWithZero α] [dec_irr : (p : Associates α) → Decidable (Irreducible p)] [UniqueFactorizationMonoid α] [dec : DecidableEq α] [dec' : DecidableEq (Associates α)] {a : α} {b : α} (ha : a 0) (hb : b 0) (h : Associates.mk a Associates.mk b 1) :
                            p, Prime p p a p b
                            theorem Associates.coprime_iff_inf_one {α : Type u_1} [CancelCommMonoidWithZero α] [dec_irr : (p : Associates α) → Decidable (Irreducible p)] [UniqueFactorizationMonoid α] [dec : DecidableEq α] [dec' : DecidableEq (Associates α)] {a : α} {b : α} (ha0 : a 0) (hb0 : b 0) :
                            Associates.mk a Associates.mk b = 1 ∀ {d : α}, d ad b¬Prime d
                            theorem Associates.factors_self {α : Type u_1} [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] [dec : DecidableEq α] [dec' : DecidableEq (Associates α)] [Nontrivial α] {p : Associates α} (hp : Irreducible p) :
                            Associates.factors p = some {{ val := p, property := hp }}
                            theorem Associates.factors_prime_pow {α : Type u_1} [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] [dec : DecidableEq α] [dec' : DecidableEq (Associates α)] [Nontrivial α] {p : Associates α} (hp : Irreducible p) (k : ) :
                            Associates.factors (p ^ k) = some (Multiset.replicate k { val := p, property := hp })
                            theorem Associates.prime_pow_dvd_iff_le {α : Type u_1} [CancelCommMonoidWithZero α] [dec_irr : (p : Associates α) → Decidable (Irreducible p)] [UniqueFactorizationMonoid α] [dec : DecidableEq α] [dec' : DecidableEq (Associates α)] [Nontrivial α] {m : Associates α} {p : Associates α} (h₁ : m 0) (h₂ : Irreducible p) {k : } :
                            theorem Associates.le_of_count_ne_zero {α : Type u_1} [CancelCommMonoidWithZero α] [dec_irr : (p : Associates α) → Decidable (Irreducible p)] [UniqueFactorizationMonoid α] [dec : DecidableEq α] [dec' : DecidableEq (Associates α)] {m : Associates α} {p : Associates α} (h0 : m 0) (hp : Irreducible p) :
                            theorem Associates.count_eq_zero_of_ne {α : Type u_1} [CancelCommMonoidWithZero α] [dec_irr : (p : Associates α) → Decidable (Irreducible p)] [UniqueFactorizationMonoid α] [dec : DecidableEq α] [dec' : DecidableEq (Associates α)] {p : Associates α} {q : Associates α} (hp : Irreducible p) (hq : Irreducible q) (h : p q) :
                            theorem Associates.count_of_coprime {α : Type u_1} [CancelCommMonoidWithZero α] [dec_irr : (p : Associates α) → Decidable (Irreducible p)] [UniqueFactorizationMonoid α] [dec : DecidableEq α] [dec' : DecidableEq (Associates α)] {a : Associates α} (ha : a 0) {b : Associates α} (hb : b 0) (hab : ∀ (d : Associates α), d ad b¬Prime d) {p : Associates α} (hp : Irreducible p) :
                            theorem Associates.count_mul_of_coprime {α : Type u_1} [CancelCommMonoidWithZero α] [dec_irr : (p : Associates α) → Decidable (Irreducible p)] [UniqueFactorizationMonoid α] [dec : DecidableEq α] [dec' : DecidableEq (Associates α)] {a : Associates α} {b : Associates α} (hb : b 0) {p : Associates α} (hp : Irreducible p) (hab : ∀ (d : Associates α), d ad b¬Prime d) :
                            theorem Associates.dvd_count_of_dvd_count_mul {α : Type u_1} [CancelCommMonoidWithZero α] [dec_irr : (p : Associates α) → Decidable (Irreducible p)] [UniqueFactorizationMonoid α] [dec : DecidableEq α] [dec' : DecidableEq (Associates α)] {a : Associates α} {b : Associates α} (hb : b 0) {p : Associates α} (hp : Irreducible p) (hab : ∀ (d : Associates α), d ad b¬Prime d) {k : } (habk : k Associates.count p (Associates.factors (a * b))) :
                            theorem Associates.count_pow {α : Type u_1} [CancelCommMonoidWithZero α] [dec_irr : (p : Associates α) → Decidable (Irreducible p)] [UniqueFactorizationMonoid α] [dec : DecidableEq α] [dec' : DecidableEq (Associates α)] [Nontrivial α] {a : Associates α} (ha : a 0) {p : Associates α} (hp : Irreducible p) (k : ) :
                            theorem Associates.dvd_count_pow {α : Type u_1} [CancelCommMonoidWithZero α] [dec_irr : (p : Associates α) → Decidable (Irreducible p)] [UniqueFactorizationMonoid α] [dec : DecidableEq α] [dec' : DecidableEq (Associates α)] [Nontrivial α] {a : Associates α} (ha : a 0) {p : Associates α} (hp : Irreducible p) (k : ) :
                            theorem Associates.is_pow_of_dvd_count {α : Type u_1} [CancelCommMonoidWithZero α] [dec_irr : (p : Associates α) → Decidable (Irreducible p)] [UniqueFactorizationMonoid α] [dec : DecidableEq α] [dec' : DecidableEq (Associates α)] [Nontrivial α] {a : Associates α} (ha : a 0) {k : } (hk : ∀ (p : Associates α), Irreducible pk Associates.count p (Associates.factors a)) :
                            b, a = b ^ k
                            theorem Associates.eq_pow_count_factors_of_dvd_pow {α : Type u_1} [CancelCommMonoidWithZero α] [dec_irr : (p : Associates α) → Decidable (Irreducible p)] [UniqueFactorizationMonoid α] [dec : DecidableEq α] [dec' : DecidableEq (Associates α)] {p : Associates α} {a : Associates α} (hp : Irreducible p) {n : } (h : a p ^ n) :

                            The only divisors of prime powers are prime powers. See eq_pow_find_of_dvd_irreducible_pow for an explicit expression as a p-power (without using count).

                            theorem Associates.count_factors_eq_find_of_dvd_pow {α : Type u_1} [CancelCommMonoidWithZero α] [dec_irr : (p : Associates α) → Decidable (Irreducible p)] [UniqueFactorizationMonoid α] [dec : DecidableEq α] [dec' : DecidableEq (Associates α)] {a : Associates α} {p : Associates α} (hp : Irreducible p) [(n : ) → Decidable (a p ^ n)] {n : } (h : a p ^ n) :
                            Nat.find (_ : n, (fun n => a p ^ n) n) = Associates.count p (Associates.factors a)
                            theorem Associates.eq_pow_of_mul_eq_pow {α : Type u_1} [CancelCommMonoidWithZero α] [dec_irr : (p : Associates α) → Decidable (Irreducible p)] [UniqueFactorizationMonoid α] [dec : DecidableEq α] [dec' : DecidableEq (Associates α)] [Nontrivial α] {a : Associates α} {b : Associates α} {c : Associates α} (ha : a 0) (hb : b 0) (hab : ∀ (d : Associates α), d ad b¬Prime d) {k : } (h : a * b = c ^ k) :
                            d, a = d ^ k
                            theorem Associates.eq_pow_find_of_dvd_irreducible_pow {α : Type u_1} [CancelCommMonoidWithZero α] [dec_irr : (p : Associates α) → Decidable (Irreducible p)] [UniqueFactorizationMonoid α] [dec : DecidableEq α] [dec' : DecidableEq (Associates α)] {a : Associates α} {p : Associates α} (hp : Irreducible p) [(n : ) → Decidable (a p ^ n)] {n : } (h : a p ^ n) :
                            a = p ^ Nat.find (_ : n, (fun n => a p ^ n) n)

                            The only divisors of prime powers are prime powers.

                            toGCDMonoid constructs a GCD monoid out of a unique factorization domain.

                            Equations
                            • One or more equations did not get rendered due to their size.
                            Instances For

                              toNormalizedGCDMonoid constructs a GCD monoid out of a normalization on a unique factorization domain.

                              Equations
                              • One or more equations did not get rendered due to their size.
                              Instances For
                                noncomputable def UniqueFactorizationMonoid.fintypeSubtypeDvd {M : Type u_2} [CancelCommMonoidWithZero M] [UniqueFactorizationMonoid M] [Fintype Mˣ] (y : M) (hy : y 0) :
                                Fintype { x // x y }

                                If y is a nonzero element of a unique factorization monoid with finitely many units (e.g. , Ideal (ring_of_integers K)), it has finitely many divisors.

                                Equations
                                • One or more equations did not get rendered due to their size.
                                Instances For

                                  This returns the multiset of irreducible factors as a Finsupp.

                                  Equations
                                  Instances For
                                    @[simp]

                                    The support of factorization n is exactly the Finset of normalized factors

                                    @[simp]
                                    theorem factorization_mul {α : Type u_1} [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] [NormalizationMonoid α] [DecidableEq α] {a : α} {b : α} (ha : a 0) (hb : b 0) :

                                    For nonzero a and b, the power of p in a * b is the sum of the powers in a and b

                                    For any p, the power of p in x^n is n times the power in x