Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 2ae199f

Browse files
committed
refactor(ring_theory/unique_factorization_domain): completes the refactor of unique_factorization_domain (#4156)
Refactors `unique_factorization_domain` to `unique_factorization_monoid` `unique_factorization_monoid` is a predicate `unique_factorization_monoid` now requires no additive/subtractive structure Co-authored-by: Aaron Anderson <65780815+awainverse@users.noreply.github.com>
1 parent 480c92c commit 2ae199f

File tree

12 files changed

+472
-363
lines changed

12 files changed

+472
-363
lines changed

docs/100.yaml

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -271,8 +271,8 @@
271271
- nat.factors_unique
272272
- int.euclidean_domain
273273
- euclidean_domain.to_principal_ideal_domain
274-
- unique_factorization_domain
275-
- unique_factorization_domain.unique
274+
- unique_factorization_monoid
275+
- unique_factorization_monoid.factors_unique
276276
author : mathlib (Chris Hughes)
277277
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."
278278
81:

docs/overview.yaml

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -63,7 +63,7 @@ General algebra:
6363
Divisibility in integral domains:
6464
irreducible elements: 'irreducible'
6565
coprime elements: 'is_coprime'
66-
unique factorisation domain: 'unique_factorization_domain'
66+
unique factorisation domain: 'unique_factorization_monoid'
6767
greatest common divisor: 'gcd_monoid.gcd'
6868
least common multiple: 'gcd_monoid.lcm'
6969
principal ideal domain: 'submodule.is_principal'

scripts/nolints.txt

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1848,4 +1848,4 @@ apply_nolint uniform_space.completion.map₂ doc_blame
18481848

18491849
-- topology/uniform_space/uniform_embedding.lean
18501850
apply_nolint uniform_embedding doc_blame
1851-
apply_nolint uniform_inducing doc_blame
1851+
apply_nolint uniform_inducing doc_blame

src/algebra/associated.lean

Lines changed: 0 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -632,11 +632,6 @@ instance : no_zero_divisors (associates α) :=
632632
have a = 0 ∨ b = 0, from mul_eq_zero.1 this,
633633
this.imp (assume h, h.symm ▸ rfl) (assume h, h.symm ▸ rfl))⟩
634634

635-
theorem prod_eq_zero_iff [nontrivial α] {s : multiset (associates α)} :
636-
s.prod = 0 ↔ (0 : associates α) ∈ s :=
637-
multiset.induction_on s (by simp) $
638-
assume a s, by simp [mul_eq_zero, @eq_comm _ 0 a] {contextual := tt}
639-
640635
theorem irreducible_iff_prime_iff :
641636
(∀ a : α, irreducible a ↔ prime a) ↔ (∀ a : (associates α), irreducible a ↔ prime a) :=
642637
begin

src/algebra/gcd_monoid.lean

Lines changed: 5 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -78,14 +78,17 @@ def normalize : α →* α :=
7878
classical.by_cases (λ hy : y = 0, by rw [hy, mul_zero, zero_mul, mul_zero]) $ λ hy,
7979
by simp only [norm_unit_mul hx hy, units.coe_mul]; simp only [mul_assoc, mul_left_comm y], }
8080

81-
@[simp] lemma normalize_apply {x : α} : normalize x = x * norm_unit x := rfl
82-
8381
theorem associated_normalize {x : α} : associated x (normalize x) :=
8482
⟨_, rfl⟩
8583

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

87+
lemma associates.mk_normalize {x : α} : associates.mk (normalize x) = associates.mk x :=
88+
associates.mk_eq_mk_iff_associated.2 normalize_associated
89+
90+
@[simp] lemma normalize_apply {x : α} : normalize x = x * norm_unit x := rfl
91+
8992
@[simp] lemma normalize_zero : normalize (0 : α) = 0 := by simp
9093

9194
@[simp] lemma normalize_one : normalize (1 : α) = 1 := normalize.map_one

src/data/multiset/basic.lean

Lines changed: 6 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -839,6 +839,12 @@ quotient.induction_on s $ λ l,
839839
lemma dvd_prod [comm_monoid α] {a : α} {s : multiset α} : a ∈ s → a ∣ s.prod :=
840840
quotient.induction_on s (λ l a h, by simpa using list.dvd_prod h) a
841841

842+
theorem prod_eq_zero_iff [comm_cancel_monoid_with_zero α] [nontrivial α]
843+
{s : multiset α} :
844+
s.prod = 0 ↔ (0 : α) ∈ s :=
845+
multiset.induction_on s (by simp) $
846+
assume a s, by simp [mul_eq_zero, @eq_comm _ 0 a] {contextual := tt}
847+
842848
lemma le_sum_of_subadditive [add_comm_monoid α] [ordered_add_comm_monoid β]
843849
(f : α → β) (h_zero : f 0 = 0) (h_add : ∀x y, f (x + y) ≤ f x + f y) (s : multiset α) :
844850
f s.sum ≤ (s.map f).sum :=

src/field_theory/splitting_field.lean

Lines changed: 6 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -255,10 +255,10 @@ by rw [degree_eq_nat_degree p_ne_zero, nat_degree_eq_card_roots hsplit]
255255

256256
section UFD
257257

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

261-
open unique_factorization_domain associates
261+
open unique_factorization_monoid associates
262262

263263
lemma splits_of_exists_multiset {f : polynomial α} {s : multiset β}
264264
(hs : f.map i = C (i f.leading_coeff) * (s.map (λ a : β, (X : polynomial β) - C a)).prod) :
@@ -268,8 +268,8 @@ else
268268
or.inr $ λ p hp hdp,
269269
have ht : multiset.rel associated
270270
(factors (f.map i)) (s.map (λ a : β, (X : polynomial β) - C a)) :=
271-
unique
272-
(λ p hp, irreducible_factors (map_ne_zero hf0) _ hp)
271+
factors_unique
272+
(λ p hp, irreducible_of_factor _ hp)
273273
(λ p' m, begin
274274
obtain ⟨a,m,rfl⟩ := multiset.mem_map.1 m,
275275
exact irreducible_of_degree_eq_one (degree_X_sub_C _),
@@ -278,15 +278,15 @@ else
278278
⟨(units.map' C : units β →* units (polynomial β)) (units.mk0 (f.map i).leading_coeff
279279
(mt leading_coeff_eq_zero.1 (map_ne_zero hf0))),
280280
by conv_rhs {rw [hs, ← leading_coeff_map i, mul_comm]}; refl⟩
281-
... ~ᵤ _ : associated.symm (unique_factorization_domain.factors_prod (by simpa using hf0))),
281+
... ~ᵤ _ : associated.symm (unique_factorization_monoid.factors_prod (by simpa using hf0))),
282282
let ⟨q, hq, hpq⟩ := exists_mem_factors_of_dvd (by simpa) hp hdp in
283283
let ⟨q', hq', hqq'⟩ := multiset.exists_mem_of_rel_of_mem ht hq in
284284
let ⟨a, ha⟩ := multiset.mem_map.1 hq' in
285285
by rw [← degree_X_sub_C a, ha.2];
286286
exact degree_eq_degree_of_associated (hpq.trans hqq')
287287

288288
lemma splits_of_splits_id {f : polynomial α} : splits (ring_hom.id _) f → splits i f :=
289-
unique_factorization_domain.induction_on_prime f (λ _, splits_zero _)
289+
unique_factorization_monoid.induction_on_prime f (λ _, splits_zero _)
290290
(λ _ hu _, splits_of_degree_le_one _
291291
((is_unit_iff_degree_eq_zero.1 hu).symm ▸ dec_trivial))
292292
(λ a p ha0 hp ih hfi, splits_mul _

src/ring_theory/discrete_valuation_ring.lean

Lines changed: 25 additions & 26 deletions
Original file line numberDiff line numberDiff line change
@@ -171,18 +171,15 @@ end
171171

172172
/-- Implementation detail: an integral domain in which there is a unit `p`
173173
such that every nonzero element is associated to a power of `p` is a unique factorization domain.
174-
175174
See `discrete_valuation_ring.of_has_unit_mul_pow_irreducible_factorization`. -/
176-
noncomputable def ufd : unique_factorization_domain R :=
175+
theorem ufd : unique_factorization_monoid R :=
177176
let p := classical.some hR in
178177
let spec := classical.some_spec hR in
179-
{ factors := λ x, if h : x = 0 then 0 else multiset.repeat p (classical.some (spec.2 h)),
180-
factors_prod := λ x hx,
181-
by { rw [dif_neg hx, multiset.prod_repeat], exact (classical.some_spec (spec.2 hx)), },
182-
prime_factors :=
183-
begin
184-
intros x hx q hq,
185-
rw dif_neg hx at hq,
178+
unique_factorization_monoid_of_exists_prime_of_factor $ λ x hx,
179+
begin
180+
use multiset.repeat p (classical.some (spec.2 hx)),
181+
split,
182+
{ intros q hq,
186183
have hpq := multiset.eq_of_mem_repeat hq,
187184
rw hpq,
188185
refine ⟨spec.1.ne_zero, spec.1.not_unit, _⟩,
@@ -198,21 +195,23 @@ let spec := classical.some_spec hR in
198195
{ simp only [hm, one_mul, pow_zero] at h ⊢, right, exact h },
199196
left,
200197
obtain ⟨m, rfl⟩ := nat.exists_eq_succ_of_ne_zero hm,
201-
apply dvd_mul_of_dvd_left (dvd_refl _) _,
202-
end }
198+
apply dvd_mul_of_dvd_left (dvd_refl _) _ },
199+
{ rw [multiset.prod_repeat], exact (classical.some_spec (spec.2 hx)), }
200+
end
203201

204202
omit hR
205203

206-
lemma of_ufd_of_unique_irreducible [unique_factorization_domain R]
204+
lemma of_ufd_of_unique_irreducible [unique_factorization_monoid R]
207205
(h₁ : ∃ p : R, irreducible p)
208206
(h₂ : ∀ ⦃p q : R⦄, irreducible p → irreducible q → associated p q) :
209207
has_unit_mul_pow_irreducible_factorization R :=
210208
begin
211209
obtain ⟨p, hp⟩ := h₁,
212210
refine ⟨p, hp, _⟩,
213211
intros x hx,
214-
refine ⟨(unique_factorization_domain.factors x).card, _⟩,
215-
have H := unique_factorization_domain.factors_prod hx,
212+
cases wf_dvd_monoid.exists_factors x hx with fx hfx,
213+
refine ⟨fx.card, _⟩,
214+
have H := hfx.2,
216215
rw ← associates.mk_eq_mk_iff_associated at H ⊢,
217216
rw [← H, ← associates.prod_mk, associates.mk_pow, ← multiset.prod_repeat],
218217
congr' 1,
@@ -222,13 +221,13 @@ begin
222221
multiset.mem_map, exists_imp_distrib],
223222
rintros _ q hq rfl,
224223
rw associates.mk_eq_mk_iff_associated,
225-
apply h₂ (unique_factorization_domain.irreducible_factors hx _ hq) hp,
224+
apply h₂ (hfx.1 _ hq) hp,
226225
end
227226

228227
end has_unit_mul_pow_irreducible_factorization
229228

230229
lemma aux_pid_of_ufd_of_unique_irreducible
231-
(R : Type u) [integral_domain R] [unique_factorization_domain R]
230+
(R : Type u) [integral_domain R] [unique_factorization_monoid R]
232231
(h₁ : ∃ p : R, irreducible p)
233232
(h₂ : ∀ ⦃p q : R⦄, irreducible p → irreducible q → associated p q) :
234233
is_principal_ideal_ring R :=
@@ -264,7 +263,7 @@ A unique factorization domain with at least one irreducible element
264263
in which all irreducible elements are associated
265264
is a discrete valuation ring.
266265
-/
267-
lemma of_ufd_of_unique_irreducible {R : Type u} [integral_domain R] [unique_factorization_domain R]
266+
lemma of_ufd_of_unique_irreducible {R : Type u} [integral_domain R] [unique_factorization_monoid R]
268267
(h₁ : ∃ p : R, irreducible p)
269268
(h₂ : ∀ ⦃p q : R⦄, irreducible p → irreducible q → associated p q) :
270269
discrete_valuation_ring R :=
@@ -276,14 +275,14 @@ begin
276275
{ rw submodule.ne_bot_iff,
277276
refine ⟨p, ideal.mem_span_singleton.mpr (dvd_refl p), hp.ne_zero⟩, },
278277
{ rwa [ideal.span_singleton_prime hp.ne_zero,
279-
unique_factorization_domain.irreducible_iff_prime], },
278+
unique_factorization_monoid.irreducible_iff_prime], },
280279
{ intro I,
281280
rw ← submodule.is_principal.span_singleton_generator I,
282281
rintro ⟨I0, hI⟩,
283282
apply span_singleton_eq_span_singleton.mpr,
284283
apply h₂ _ hp,
285284
erw [ne.def, span_singleton_eq_bot] at I0,
286-
rwa [unique_factorization_domain.irreducible_iff_prime, ← ideal.span_singleton_prime I0], },
285+
rwa [unique_factorization_monoid.irreducible_iff_prime, ← ideal.span_singleton_prime I0], },
287286
end
288287

289288
/--
@@ -295,7 +294,7 @@ lemma of_has_unit_mul_pow_irreducible_factorization {R : Type u} [integral_domai
295294
(hR : has_unit_mul_pow_irreducible_factorization R) :
296295
discrete_valuation_ring R :=
297296
begin
298-
letI : unique_factorization_domain R := hR.ufd,
297+
letI : unique_factorization_monoid R := hR.ufd,
299298
apply of_ufd_of_unique_irreducible _ hR.unique_irreducible,
300299
unfreezingI { obtain ⟨p, hp, H⟩ := hR, exact ⟨p, hp⟩, },
301300
end
@@ -309,9 +308,10 @@ variable {R}
309308
lemma associated_pow_irreducible {x : R} (hx : x ≠ 0) {ϖ : R} (hirr : irreducible ϖ) :
310309
∃ (n : ℕ), associated x (ϖ ^ n) :=
311310
begin
312-
have : unique_factorization_domain R := principal_ideal_ring.to_unique_factorization_domain,
313-
unfreezingI { use (unique_factorization_domain.factors x).card },
314-
have H := unique_factorization_domain.factors_prod hx,
311+
have : wf_dvd_monoid R := is_noetherian_ring.wf_dvd_monoid,
312+
cases wf_dvd_monoid.exists_factors x hx with fx hfx,
313+
unfreezingI { use fx.card },
314+
have H := hfx.2,
315315
rw ← associates.mk_eq_mk_iff_associated at H ⊢,
316316
rw [← H, ← associates.prod_mk, associates.mk_pow, ← multiset.prod_repeat],
317317
congr' 1,
@@ -321,7 +321,7 @@ begin
321321
rintros _ _ _ rfl,
322322
rw associates.mk_eq_mk_iff_associated,
323323
refine associated_of_irreducible _ _ hirr,
324-
apply unique_factorization_domain.irreducible_factors hx,
324+
apply hfx.1,
325325
assumption
326326
end
327327

@@ -348,8 +348,7 @@ begin
348348
refine ⟨u * v⁻¹, _⟩,
349349
simp only [units.coe_mul],
350350
rw [mul_left_comm, ← mul_assoc, h, mul_right_comm, units.mul_inv, one_mul], },
351-
letI := @principal_ideal_ring.to_unique_factorization_domain R _ _,
352-
have := multiset.card_eq_card_of_rel (unique_factorization_domain.unique _ _ key),
351+
have := multiset.card_eq_card_of_rel (unique_factorization_monoid.factors_unique _ _ key),
353352
{ simpa only [multiset.card_repeat] },
354353
all_goals
355354
{ intros x hx, replace hx := multiset.eq_of_mem_repeat hx,

src/ring_theory/localization.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -1316,15 +1316,15 @@ end
13161316

13171317
section num_denom
13181318

1319-
variables [unique_factorization_domain A] (φ : fraction_map A K)
1319+
variables [unique_factorization_monoid A] (φ : fraction_map A K)
13201320

13211321
lemma exists_reduced_fraction (x : φ.codomain) :
13221322
∃ (a : A) (b : non_zero_divisors A),
13231323
(∀ {d}, d ∣ a → d ∣ b → is_unit d) ∧ φ.mk' a b = x :=
13241324
begin
13251325
obtain ⟨⟨b, b_nonzero⟩, a, hab⟩ := φ.exists_integer_multiple x,
13261326
obtain ⟨a', b', c', no_factor, rfl, rfl⟩ :=
1327-
unique_factorization_domain.exists_reduced_factors' a b
1327+
unique_factorization_monoid.exists_reduced_factors' a b
13281328
(mem_non_zero_divisors_iff_ne_zero.mp b_nonzero),
13291329
obtain ⟨c'_nonzero, b'_nonzero⟩ := mul_mem_non_zero_divisors.mp b_nonzero,
13301330
refine ⟨a', ⟨b', b'_nonzero⟩, @no_factor, _⟩,

src/ring_theory/polynomial/rational_root.lean

Lines changed: 7 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -42,7 +42,7 @@ begin
4242
end
4343

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

5757
section rational_root_theorem
5858

59-
variables {A K : Type*} [integral_domain A] [unique_factorization_domain A] [field K]
59+
variables {A K : Type*} [integral_domain A] [unique_factorization_monoid A] [field K]
6060
variables {f : fraction_map A K}
6161

62-
open polynomial unique_factorization_domain
62+
open polynomial unique_factorization_monoid
6363

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

121-
namespace unique_factorization_domain
121+
namespace unique_factorization_monoid
122122

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

130-
end unique_factorization_domain
130+
end unique_factorization_monoid
131131

132132
end rational_root_theorem

0 commit comments

Comments
 (0)