Documentation

Mathlib.RingTheory.UniqueFactorizationDomain.Basic

Basic results un unique factorization monoids #

Main results #

@[deprecated WfDvdMonoid.wellFoundedLT_associates]
theorem WfDvdMonoid.wellFounded_associates {α : Type u_1} [CommMonoidWithZero α] [WfDvdMonoid α] :
WellFounded fun (x1 x2 : Associates α) => x1 < x2
@[deprecated WfDvdMonoid.of_wellFoundedLT_associates]
theorem WfDvdMonoid.of_wellFounded_associates {α : Type u_1} [CancelCommMonoidWithZero α] (h : WellFounded fun (x1 x2 : Associates α) => x1 < x2) :
theorem prime_factors_unique {α : Type u_1} [CancelCommMonoidWithZero α] {f g : Multiset α} :
(∀ xf, Prime x)(∀ xg, Prime x)Associated f.prod g.prodMultiset.Rel Associated f g
theorem UniqueFactorizationMonoid.factors_unique {α : Type u_1} [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] {f g : Multiset α} (hf : xf, Irreducible x) (hg : xg, Irreducible x) (h : Associated f.prod g.prod) :
Multiset.Rel Associated f g
theorem prime_factors_irreducible {α : Type u_1} [CancelCommMonoidWithZero α] {a : α} {f : Multiset α} (ha : Irreducible a) (pfa : (∀ bf, Prime b) Associated f.prod 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 irreducible_iff_prime_of_exists_unique_irreducible_factors {α : Type u_1} [CancelCommMonoidWithZero α] (eif : ∀ (a : α), a 0∃ (f : Multiset α), (∀ bf, Irreducible b) Associated f.prod a) (uif : ∀ (f g : Multiset α), (∀ xf, Irreducible x)(∀ xg, Irreducible x)Associated f.prod g.prodMultiset.Rel Associated f g) (p : α) :
theorem Associates.unique' {α : Type u_1} [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] {p q : Multiset (Associates α)} :
(∀ ap, Irreducible a)(∀ aq, Irreducible a)p.prod = q.prodp = q
theorem Associates.prod_le_prod_iff_le {α : Type u_1} [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] [Nontrivial α] {p q : Multiset (Associates α)} (hp : ap, Irreducible a) (hq : aq, Irreducible a) :
p.prod q.prod p q
theorem WfDvdMonoid.of_exists_prime_factors {α : Type u_1} [CancelCommMonoidWithZero α] (pf : ∀ (a : α), a 0∃ (f : Multiset α), (∀ bf, Prime b) Associated f.prod a) :
theorem irreducible_iff_prime_of_exists_prime_factors {α : Type u_1} [CancelCommMonoidWithZero α] (pf : ∀ (a : α), a 0∃ (f : Multiset α), (∀ bf, Prime b) Associated f.prod a) {p : α} :
theorem UniqueFactorizationMonoid.of_exists_prime_factors {α : Type u_1} [CancelCommMonoidWithZero α] (pf : ∀ (a : α), a 0∃ (f : Multiset α), (∀ bf, Prime b) Associated f.prod a) :
theorem UniqueFactorizationMonoid.iff_exists_prime_factors {α : Type u_1} [CancelCommMonoidWithZero α] :
UniqueFactorizationMonoid α ∀ (a : α), a 0∃ (f : Multiset α), (∀ bf, Prime b) Associated f.prod a
theorem UniqueFactorizationMonoid.of_exists_unique_irreducible_factors {α : Type u_1} [CancelCommMonoidWithZero α] (eif : ∀ (a : α), a 0∃ (f : Multiset α), (∀ bf, Irreducible b) Associated f.prod a) (uif : ∀ (f g : Multiset α), (∀ xf, Irreducible x)(∀ xg, Irreducible x)Associated f.prod g.prodMultiset.Rel Associated f g) :
theorem UniqueFactorizationMonoid.dvd_of_dvd_mul_left_of_no_prime_factors {R : Type u_2} [CancelCommMonoidWithZero R] [UniqueFactorizationMonoid R] {a b c : R} (ha : a 0) (h : ∀ ⦃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 b 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' : R) (b' : R) (c' : R), IsRelPrime a' b' 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 b : R) (hb : b 0) :
∃ (a' : R) (b' : R) (c' : R), IsRelPrime a' b' c' * a' = a c' * b' = b
@[deprecated pow_injective_of_not_isUnit]
theorem UniqueFactorizationMonoid.pow_right_injective {M : Type u_1} [CancelCommMonoidWithZero M] {q : M} (hq : ¬IsUnit q) (hq' : q 0) :
Function.Injective fun (n : ) => q ^ n

Alias of pow_injective_of_not_isUnit.

@[deprecated pow_inj_of_not_isUnit]
theorem UniqueFactorizationMonoid.pow_eq_pow_iff {M : Type u_1} [CancelCommMonoidWithZero M] {q : M} (hq : ¬IsUnit q) (hq' : q 0) {m n : } :
q ^ m = q ^ n m = n

Alias of pow_inj_of_not_isUnit.