@@ -171,18 +171,15 @@ end
171
171
172
172
/-- Implementation detail: an integral domain in which there is a unit `p`
173
173
such that every nonzero element is associated to a power of `p` is a unique factorization domain.
174
-
175
174
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 :=
177
176
let p := classical.some hR in
178
177
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,
186
183
have hpq := multiset.eq_of_mem_repeat hq,
187
184
rw hpq,
188
185
refine ⟨spec.1 .ne_zero, spec.1 .not_unit, _⟩,
@@ -198,21 +195,23 @@ let spec := classical.some_spec hR in
198
195
{ simp only [hm, one_mul, pow_zero] at h ⊢, right, exact h },
199
196
left,
200
197
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
203
201
204
202
omit hR
205
203
206
- lemma of_ufd_of_unique_irreducible [unique_factorization_domain R]
204
+ lemma of_ufd_of_unique_irreducible [unique_factorization_monoid R]
207
205
(h₁ : ∃ p : R, irreducible p)
208
206
(h₂ : ∀ ⦃p q : R⦄, irreducible p → irreducible q → associated p q) :
209
207
has_unit_mul_pow_irreducible_factorization R :=
210
208
begin
211
209
obtain ⟨p, hp⟩ := h₁,
212
210
refine ⟨p, hp, _⟩,
213
211
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 ,
216
215
rw ← associates.mk_eq_mk_iff_associated at H ⊢,
217
216
rw [← H, ← associates.prod_mk, associates.mk_pow, ← multiset.prod_repeat],
218
217
congr' 1 ,
@@ -222,13 +221,13 @@ begin
222
221
multiset.mem_map, exists_imp_distrib],
223
222
rintros _ q hq rfl,
224
223
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,
226
225
end
227
226
228
227
end has_unit_mul_pow_irreducible_factorization
229
228
230
229
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]
232
231
(h₁ : ∃ p : R, irreducible p)
233
232
(h₂ : ∀ ⦃p q : R⦄, irreducible p → irreducible q → associated p q) :
234
233
is_principal_ideal_ring R :=
@@ -264,7 +263,7 @@ A unique factorization domain with at least one irreducible element
264
263
in which all irreducible elements are associated
265
264
is a discrete valuation ring.
266
265
-/
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]
268
267
(h₁ : ∃ p : R, irreducible p)
269
268
(h₂ : ∀ ⦃p q : R⦄, irreducible p → irreducible q → associated p q) :
270
269
discrete_valuation_ring R :=
@@ -276,14 +275,14 @@ begin
276
275
{ rw submodule.ne_bot_iff,
277
276
refine ⟨p, ideal.mem_span_singleton.mpr (dvd_refl p), hp.ne_zero⟩, },
278
277
{ rwa [ideal.span_singleton_prime hp.ne_zero,
279
- ← unique_factorization_domain .irreducible_iff_prime], },
278
+ ← unique_factorization_monoid .irreducible_iff_prime], },
280
279
{ intro I,
281
280
rw ← submodule.is_principal.span_singleton_generator I,
282
281
rintro ⟨I0, hI⟩,
283
282
apply span_singleton_eq_span_singleton.mpr,
284
283
apply h₂ _ hp,
285
284
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], },
287
286
end
288
287
289
288
/--
@@ -295,7 +294,7 @@ lemma of_has_unit_mul_pow_irreducible_factorization {R : Type u} [integral_domai
295
294
(hR : has_unit_mul_pow_irreducible_factorization R) :
296
295
discrete_valuation_ring R :=
297
296
begin
298
- letI : unique_factorization_domain R := hR.ufd,
297
+ letI : unique_factorization_monoid R := hR.ufd,
299
298
apply of_ufd_of_unique_irreducible _ hR.unique_irreducible,
300
299
unfreezingI { obtain ⟨p, hp, H⟩ := hR, exact ⟨p, hp⟩, },
301
300
end
@@ -309,9 +308,10 @@ variable {R}
309
308
lemma associated_pow_irreducible {x : R} (hx : x ≠ 0 ) {ϖ : R} (hirr : irreducible ϖ) :
310
309
∃ (n : ℕ), associated x (ϖ ^ n) :=
311
310
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 ,
315
315
rw ← associates.mk_eq_mk_iff_associated at H ⊢,
316
316
rw [← H, ← associates.prod_mk, associates.mk_pow, ← multiset.prod_repeat],
317
317
congr' 1 ,
@@ -321,7 +321,7 @@ begin
321
321
rintros _ _ _ rfl,
322
322
rw associates.mk_eq_mk_iff_associated,
323
323
refine associated_of_irreducible _ _ hirr,
324
- apply unique_factorization_domain.irreducible_factors hx ,
324
+ apply hfx. 1 ,
325
325
assumption
326
326
end
327
327
@@ -348,8 +348,7 @@ begin
348
348
refine ⟨u * v⁻¹, _⟩,
349
349
simp only [units.coe_mul],
350
350
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),
353
352
{ simpa only [multiset.card_repeat] },
354
353
all_goals
355
354
{ intros x hx, replace hx := multiset.eq_of_mem_repeat hx,
0 commit comments