Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - refactor(ring_theory/unique_factorization_domain): completes the refactor of unique_factorization_domain #4156

Closed
wants to merge 15 commits into from
4 changes: 2 additions & 2 deletions docs/100.yaml
Original file line number Diff line number Diff line change
Expand Up @@ -271,8 +271,8 @@
- nat.factors_unique
- int.euclidean_domain
- euclidean_domain.to_principal_ideal_domain
- unique_factorization_domain
- unique_factorization_domain.unique
- unique_factorization_monoid
- unique_factorization_monoid.factors_unique
author : mathlib (Chris Hughes)
note : "it also has a generalized version, by showing that every Euclidean domain is a unique factorization domain, and showing that the integers form a Euclidean domain."
81:
Expand Down
2 changes: 1 addition & 1 deletion docs/overview.yaml
Original file line number Diff line number Diff line change
Expand Up @@ -63,7 +63,7 @@ General algebra:
Divisibility in integral domains:
irreducible elements: 'irreducible'
coprime elements: 'is_coprime'
unique factorisation domain: 'unique_factorization_domain'
unique factorisation domain: 'unique_factorization_monoid'
greatest common divisor: 'gcd_monoid.gcd'
least common multiple: 'gcd_monoid.lcm'
principal ideal domain: 'submodule.is_principal'
Expand Down
6 changes: 3 additions & 3 deletions scripts/nolints.txt
Original file line number Diff line number Diff line change
Expand Up @@ -1211,11 +1211,11 @@ apply_nolint multiplicity.finite doc_blame
-- ring_theory/ring_invo.lean
apply_nolint ring_invo.id doc_blame

-- ring_theory/unique_factorization_domain.lean
-- ring_theory/unique_factorization_monoid.lean
awainverse marked this conversation as resolved.
Show resolved Hide resolved
apply_nolint associates.factor_set.prod doc_blame
apply_nolint associates.factors doc_blame
apply_nolint associates.factors' doc_blame
apply_nolint unique_factorization_domain.of_unique_irreducible_factorization doc_blame
apply_nolint unique_factorization_monoid.of_unique_irreducible_factorization doc_blame
apply_nolint unique_irreducible_factorization doc_blame has_inhabited_instance

-- set_theory/cofinality.lean
Expand Down Expand Up @@ -1891,4 +1891,4 @@ apply_nolint uniform_space.completion.map₂ doc_blame

-- topology/uniform_space/uniform_embedding.lean
apply_nolint uniform_embedding doc_blame
apply_nolint uniform_inducing doc_blame
apply_nolint uniform_inducing doc_blame
5 changes: 0 additions & 5 deletions src/algebra/associated.lean
Original file line number Diff line number Diff line change
Expand Up @@ -626,11 +626,6 @@ instance : no_zero_divisors (associates α) :=
have a = 0 ∨ b = 0, from mul_eq_zero.1 this,
this.imp (assume h, h.symm ▸ rfl) (assume h, h.symm ▸ rfl))⟩

theorem prod_eq_zero_iff [nontrivial α] {s : multiset (associates α)} :
s.prod = 0 ↔ (0 : associates α) ∈ s :=
multiset.induction_on s (by simp) $
assume a s, by simp [mul_eq_zero, @eq_comm _ 0 a] {contextual := tt}

theorem irreducible_iff_prime_iff :
(∀ a : α, irreducible a ↔ prime a) ↔ (∀ a : (associates α), irreducible a ↔ prime a) :=
begin
Expand Down
7 changes: 5 additions & 2 deletions src/algebra/gcd_monoid.lean
Original file line number Diff line number Diff line change
Expand Up @@ -81,14 +81,17 @@ def normalize : α →* α :=
classical.by_cases (λ hy : y = 0, by rw [hy, mul_zero, zero_mul, mul_zero]) $ λ hy,
by simp only [norm_unit_mul hx hy, units.coe_mul]; simp only [mul_assoc, mul_left_comm y], }

@[simp] lemma normalize_apply {x : α} : normalize x = x * norm_unit x := rfl

theorem associated_normalize {x : α} : associated x (normalize x) :=
⟨_, rfl⟩

theorem normalize_associated {x : α} : associated (normalize x) x :=
associated_normalize.symm

lemma associates.mk_normalize {x : α} : associates.mk (normalize x) = associates.mk x :=
associates.mk_eq_mk_iff_associated.2 normalize_associated

@[simp] lemma normalize_apply {x : α} : normalize x = x * norm_unit x := rfl

@[simp] lemma normalize_zero : normalize (0 : α) = 0 := by simp

@[simp] lemma normalize_one : normalize (1 : α) = 1 := normalize.map_one
Expand Down
2 changes: 1 addition & 1 deletion src/data/list/basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1972,7 +1972,7 @@ theorem prod_erase [decidable_eq α] [comm_monoid α] {a} :
{ simp only [list.erase, if_neg (mt eq.symm ne), prod_cons, prod_erase h, mul_left_comm a b] }
end

lemma dvd_prod [comm_semiring α] {a} {l : list α} (ha : a ∈ l) : a ∣ l.prod :=
lemma dvd_prod [comm_monoid α] {a} {l : list α} (ha : a ∈ l) : a ∣ l.prod :=
let ⟨s, t, h⟩ := mem_split ha in
by rw [h, prod_append, prod_cons, mul_left_comm]; exact dvd_mul_right _ _

Expand Down
8 changes: 7 additions & 1 deletion src/data/multiset/basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -833,9 +833,15 @@ theorem prod_hom_rel [comm_monoid β] [comm_monoid γ] (s : multiset α) {r : β
quotient.induction_on s $ λ l,
by simp only [l.prod_hom_rel h₁ h₂, quot_mk_to_coe, coe_map, coe_prod]

lemma dvd_prod [comm_semiring α] {a : α} {s : multiset α} : a ∈ s → a ∣ s.prod :=
lemma dvd_prod [comm_monoid α] {a : α} {s : multiset α} : a ∈ s → a ∣ s.prod :=
quotient.induction_on s (λ l a h, by simpa using list.dvd_prod h) a

theorem prod_eq_zero_iff [comm_cancel_monoid_with_zero α] [nontrivial α]
{s : multiset α} :
s.prod = 0 ↔ (0 : α) ∈ s :=
multiset.induction_on s (by simp) $
assume a s, by simp [mul_eq_zero, @eq_comm _ 0 a] {contextual := tt}

lemma le_sum_of_subadditive [add_comm_monoid α] [ordered_add_comm_monoid β]
(f : α → β) (h_zero : f 0 = 0) (h_add : ∀x y, f (x + y) ≤ f x + f y) (s : multiset α) :
f s.sum ≤ (s.map f).sum :=
Expand Down
12 changes: 6 additions & 6 deletions src/field_theory/splitting_field.lean
Original file line number Diff line number Diff line change
Expand Up @@ -191,10 +191,10 @@ end

section UFD

local attribute [instance, priority 10] principal_ideal_ring.to_unique_factorization_domain
local attribute [instance, priority 10] principal_ideal_ring.to_unique_factorization_monoid
local infix ` ~ᵤ ` : 50 := associated

open unique_factorization_domain associates
open unique_factorization_monoid associates

lemma splits_of_exists_multiset {f : polynomial α} {s : multiset β}
(hs : f.map i = C (i f.leading_coeff) * (s.map (λ a : β, (X : polynomial β) - C a)).prod) :
Expand All @@ -204,8 +204,8 @@ else
or.inr $ λ p hp hdp,
have ht : multiset.rel associated
(factors (f.map i)) (s.map (λ a : β, (X : polynomial β) - C a)) :=
unique
(λ p hp, irreducible_factors (map_ne_zero hf0) _ hp)
factors_unique
(λ p hp, irreducible_of_factor _ hp)
(λ p' m, begin
obtain ⟨a,m,rfl⟩ := multiset.mem_map.1 m,
exact irreducible_of_degree_eq_one (degree_X_sub_C _),
Expand All @@ -214,15 +214,15 @@ else
⟨(units.map' C : units β →* units (polynomial β)) (units.mk0 (f.map i).leading_coeff
(mt leading_coeff_eq_zero.1 (map_ne_zero hf0))),
by conv_rhs {rw [hs, ← leading_coeff_map i, mul_comm]}; refl⟩
... ~ᵤ _ : associated.symm (unique_factorization_domain.factors_prod (by simpa using hf0))),
... ~ᵤ _ : associated.symm (unique_factorization_monoid.factors_prod (by simpa using hf0))),
let ⟨q, hq, hpq⟩ := exists_mem_factors_of_dvd (by simpa) hp hdp in
let ⟨q', hq', hqq'⟩ := multiset.exists_mem_of_rel_of_mem ht hq in
let ⟨a, ha⟩ := multiset.mem_map.1 hq' in
by rw [← degree_X_sub_C a, ha.2];
exact degree_eq_degree_of_associated (hpq.trans hqq')

lemma splits_of_splits_id {f : polynomial α} : splits (ring_hom.id _) f → splits i f :=
unique_factorization_domain.induction_on_prime f (λ _, splits_zero _)
unique_factorization_monoid.induction_on_prime f (λ _, splits_zero _)
(λ _ hu _, splits_of_degree_le_one _
((is_unit_iff_degree_eq_zero.1 hu).symm ▸ dec_trivial))
(λ a p ha0 hp ih hfi, splits_mul _
Expand Down
51 changes: 25 additions & 26 deletions src/ring_theory/discrete_valuation_ring.lean
Original file line number Diff line number Diff line change
Expand Up @@ -174,18 +174,15 @@ end

/-- Implementation detail: an integral domain in which there is a unit `p`
such that every nonzero element is associated to a power of `p` is a unique factorization domain.

See `discrete_valuation_ring.of_has_unit_mul_pow_irreducible_factorization`. -/
noncomputable def ufd : unique_factorization_domain R :=
theorem ufd : unique_factorization_monoid R :=
awainverse marked this conversation as resolved.
Show resolved Hide resolved
let p := classical.some hR in
let spec := classical.some_spec hR in
{ factors := λ x, if h : x = 0 then 0 else multiset.repeat p (classical.some (spec.2 h)),
factors_prod := λ x hx,
by { rw [dif_neg hx, multiset.prod_repeat], exact (classical.some_spec (spec.2 hx)), },
prime_factors :=
begin
intros x hx q hq,
rw dif_neg hx at hq,
unique_factorization_monoid_of_exists_prime_of_factor $ λ x hx,
begin
use multiset.repeat p (classical.some (spec.2 hx)),
split,
{ intros q hq,
have hpq := multiset.eq_of_mem_repeat hq,
rw hpq,
refine ⟨spec.1.ne_zero, spec.1.not_unit, _⟩,
Expand All @@ -201,21 +198,23 @@ let spec := classical.some_spec hR in
{ simp only [hm, one_mul, pow_zero] at h ⊢, right, exact h },
left,
obtain ⟨m, rfl⟩ := nat.exists_eq_succ_of_ne_zero hm,
apply dvd_mul_of_dvd_left (dvd_refl _) _,
end }
apply dvd_mul_of_dvd_left (dvd_refl _) _ },
{ rw [multiset.prod_repeat], exact (classical.some_spec (spec.2 hx)), }
end

omit hR

lemma of_ufd_of_unique_irreducible [unique_factorization_domain R]
lemma of_ufd_of_unique_irreducible [unique_factorization_monoid R]
(h₁ : ∃ p : R, irreducible p)
(h₂ : ∀ ⦃p q : R⦄, irreducible p → irreducible q → associated p q) :
has_unit_mul_pow_irreducible_factorization R :=
begin
obtain ⟨p, hp⟩ := h₁,
refine ⟨p, hp, _⟩,
intros x hx,
refine ⟨(unique_factorization_domain.factors x).card, _⟩,
have H := unique_factorization_domain.factors_prod hx,
cases wf_dvd_monoid.exists_factors x hx with fx hfx,
refine ⟨fx.card, _⟩,
have H := hfx.2,
rw ← associates.mk_eq_mk_iff_associated at H ⊢,
rw [← H, ← associates.prod_mk, associates.mk_pow, ← multiset.prod_repeat],
congr' 1,
Expand All @@ -225,13 +224,13 @@ begin
multiset.mem_map, exists_imp_distrib],
rintros _ q hq rfl,
rw associates.mk_eq_mk_iff_associated,
apply h₂ (unique_factorization_domain.irreducible_factors hx _ hq) hp,
apply h₂ (hfx.1 _ hq) hp,
end

end has_unit_mul_pow_irreducible_factorization

lemma aux_pid_of_ufd_of_unique_irreducible
(R : Type u) [integral_domain R] [unique_factorization_domain R]
(R : Type u) [integral_domain R] [unique_factorization_monoid R]
(h₁ : ∃ p : R, irreducible p)
(h₂ : ∀ ⦃p q : R⦄, irreducible p → irreducible q → associated p q) :
is_principal_ideal_ring R :=
Expand Down Expand Up @@ -267,7 +266,7 @@ A unique factorization domain with at least one irreducible element
in which all irreducible elements are associated
is a discrete valuation ring.
-/
lemma of_ufd_of_unique_irreducible {R : Type u} [integral_domain R] [unique_factorization_domain R]
lemma of_ufd_of_unique_irreducible {R : Type u} [integral_domain R] [unique_factorization_monoid R]
(h₁ : ∃ p : R, irreducible p)
(h₂ : ∀ ⦃p q : R⦄, irreducible p → irreducible q → associated p q) :
discrete_valuation_ring R :=
Expand All @@ -279,14 +278,14 @@ begin
{ rw submodule.ne_bot_iff,
refine ⟨p, ideal.mem_span_singleton.mpr (dvd_refl p), hp.ne_zero⟩, },
{ rwa [ideal.span_singleton_prime hp.ne_zero,
unique_factorization_domain.irreducible_iff_prime], },
unique_factorization_monoid.irreducible_iff_prime], },
{ intro I,
rw ← submodule.is_principal.span_singleton_generator I,
rintro ⟨I0, hI⟩,
apply span_singleton_eq_span_singleton.mpr,
apply h₂ _ hp,
erw [ne.def, span_singleton_eq_bot] at I0,
rwa [unique_factorization_domain.irreducible_iff_prime, ← ideal.span_singleton_prime I0], },
rwa [unique_factorization_monoid.irreducible_iff_prime, ← ideal.span_singleton_prime I0], },
end

/--
Expand All @@ -298,7 +297,7 @@ lemma of_has_unit_mul_pow_irreducible_factorization {R : Type u} [integral_domai
(hR : has_unit_mul_pow_irreducible_factorization R) :
discrete_valuation_ring R :=
begin
letI : unique_factorization_domain R := hR.ufd,
letI : unique_factorization_monoid R := hR.ufd,
apply of_ufd_of_unique_irreducible _ hR.unique_irreducible,
unfreezingI { obtain ⟨p, hp, H⟩ := hR, exact ⟨p, hp⟩, },
end
Expand All @@ -312,9 +311,10 @@ variable {R}
lemma associated_pow_irreducible {x : R} (hx : x ≠ 0) {ϖ : R} (hirr : irreducible ϖ) :
∃ (n : ℕ), associated x (ϖ ^ n) :=
begin
have : unique_factorization_domain R := principal_ideal_ring.to_unique_factorization_domain,
unfreezingI { use (unique_factorization_domain.factors x).card },
have H := unique_factorization_domain.factors_prod hx,
have : wf_dvd_monoid R := is_noetherian_ring.wf_dvd_monoid,
cases wf_dvd_monoid.exists_factors x hx with fx hfx,
unfreezingI { use fx.card },
have H := hfx.2,
rw ← associates.mk_eq_mk_iff_associated at H ⊢,
rw [← H, ← associates.prod_mk, associates.mk_pow, ← multiset.prod_repeat],
congr' 1,
Expand All @@ -324,7 +324,7 @@ begin
rintros _ _ _ rfl,
rw associates.mk_eq_mk_iff_associated,
refine associated_of_irreducible _ _ hirr,
apply unique_factorization_domain.irreducible_factors hx,
apply hfx.1,
assumption
end

Expand All @@ -351,8 +351,7 @@ begin
refine ⟨u * v⁻¹, _⟩,
simp only [units.coe_mul],
rw [mul_left_comm, ← mul_assoc, h, mul_right_comm, units.mul_inv, one_mul], },
letI := @principal_ideal_ring.to_unique_factorization_domain R _ _,
have := multiset.card_eq_card_of_rel (unique_factorization_domain.unique _ _ key),
have := multiset.card_eq_card_of_rel (unique_factorization_monoid.factors_unique _ _ key),
{ simpa only [multiset.card_repeat] },
all_goals
{ intros x hx, replace hx := multiset.eq_of_mem_repeat hx,
Expand Down
6 changes: 3 additions & 3 deletions src/ring_theory/localization.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1249,15 +1249,15 @@ end

section num_denom

variables [unique_factorization_domain A] (φ : fraction_map A K)
variables [unique_factorization_monoid A] (φ : fraction_map A K)

lemma exists_reduced_fraction (x : φ.codomain) :
∃ (a : A) (b : non_zero_divisors A),
(∀ {d}, d ∣ a → d ∣ b → is_unit d) ∧ φ.mk' a b = x :=
begin
obtain ⟨⟨b, b_nonzero⟩, a, hab⟩ := φ.exists_integer_multiple x,
obtain ⟨a', b', c', no_factor, rfl, rfl⟩ :=
unique_factorization_domain.exists_reduced_factors' a b
unique_factorization_monoid.exists_reduced_factors' a b
(mem_non_zero_divisors_iff_ne_zero.mp b_nonzero),
obtain ⟨c'_nonzero, b'_nonzero⟩ := mul_mem_non_zero_divisors.mp b_nonzero,
refine ⟨a', ⟨b', b'_nonzero⟩, @no_factor, _⟩,
Expand Down Expand Up @@ -1447,4 +1447,4 @@ noncomputable def field_equiv_of_quotient {K : Type*} [field K] (f : fraction_ma
fraction_ring A ≃+* K :=
localization.ring_equiv_of_quotient f

end fraction_ring
end fraction_ring
14 changes: 7 additions & 7 deletions src/ring_theory/polynomial/rational_root.lean
Original file line number Diff line number Diff line change
Expand Up @@ -42,7 +42,7 @@ begin
end

lemma num_is_root_scale_roots_of_aeval_eq_zero
[unique_factorization_domain A] (g : fraction_map A K)
[unique_factorization_monoid A] (g : fraction_map A K)
{p : polynomial A} {x : g.codomain} (hr : aeval x p = 0) :
is_root (scale_roots p (g.denom x)) (g.num x) :=
begin
Expand All @@ -56,10 +56,10 @@ end scale_roots

section rational_root_theorem

variables {A K : Type*} [integral_domain A] [unique_factorization_domain A] [field K]
variables {A K : Type*} [integral_domain A] [unique_factorization_monoid A] [field K]
variables {f : fraction_map A K}

open polynomial unique_factorization_domain
open polynomial unique_factorization_monoid

/-- Rational root theorem part 1:
if `r : f.codomain` is a root of a polynomial over the ufd `A`,
Expand All @@ -74,7 +74,7 @@ begin
{ obtain ⟨u, hu⟩ := is_unit_pow p.nat_degree (f.is_unit_denom_of_num_eq_zero hr),
rw ←hu at this,
exact units.dvd_mul_right.mp this },
{ refine dvd_of_dvd_mul_left_of_no_prime_factors hr _ this,
{ refine dvd_of_dvd_mul_left_of_no_prime_of_factor hr _ this,
intros q dvd_num dvd_denom_pow hq,
apply hq.not_unit,
exact f.num_denom_reduced r dvd_num (hq.dvd_of_dvd_pow dvd_denom_pow) } },
Expand All @@ -93,7 +93,7 @@ theorem denom_dvd_of_is_root {p : polynomial A} {r : f.codomain} (hr : aeval r p
(f.denom r : A) ∣ p.leading_coeff :=
begin
suffices : (f.denom r : A) ∣ p.leading_coeff * f.num r ^ p.nat_degree,
{ refine dvd_of_dvd_mul_left_of_no_prime_factors
{ refine dvd_of_dvd_mul_left_of_no_prime_of_factor
(mem_non_zero_divisors_iff_ne_zero.mp (f.denom r).2) _ this,
intros q dvd_denom dvd_num_pow hq,
apply hq.not_unit,
Expand All @@ -118,7 +118,7 @@ theorem is_integer_of_is_root_of_monic {p : polynomial A} (hp : monic p) {r : f.
(hr : aeval r p = 0) : f.is_integer r :=
f.is_integer_of_is_unit_denom (is_unit_of_dvd_one _ (hp ▸ denom_dvd_of_is_root hr))

namespace unique_factorization_domain
namespace unique_factorization_monoid

lemma integer_of_integral {x : f.codomain} :
is_integral A x → f.is_integer x :=
Expand All @@ -127,6 +127,6 @@ lemma integer_of_integral {x : f.codomain} :
lemma integrally_closed : integral_closure A f.codomain = ⊥ :=
eq_bot_iff.mpr (λ x hx, algebra.mem_bot.mpr (integer_of_integral hx))

end unique_factorization_domain
end unique_factorization_monoid

end rational_root_theorem
Loading