@@ -29,13 +29,13 @@ open_locale big_operators
29
29
/-- `multiplicity a b` returns the largest natural number `n` such that
30
30
`a ^ n ∣ b`, as an `part_enat` or natural with infinity. If `∀ n, a ^ n ∣ b`,
31
31
then it returns `⊤`-/
32
- def multiplicity [comm_monoid α] [decidable_rel ((∣) : α → α → Prop )] (a b : α) : part_enat :=
32
+ def multiplicity [monoid α] [decidable_rel ((∣) : α → α → Prop )] (a b : α) : part_enat :=
33
33
part_enat.find $ λ n, ¬a ^ (n + 1 ) ∣ b
34
34
35
35
namespace multiplicity
36
36
37
- section comm_monoid
38
- variables [comm_monoid α]
37
+ section monoid
38
+ variables [monoid α]
39
39
40
40
/-- `multiplicity.finite a b` indicates that the multiplicity of `a` in `b` is finite. -/
41
41
@[reducible] def finite (a b : α) : Prop := ∃ n : ℕ, ¬a ^ (n + 1 ) ∣ b
@@ -45,6 +45,9 @@ lemma finite_iff_dom [decidable_rel ((∣) : α → α → Prop)] {a b : α} :
45
45
46
46
lemma finite_def {a b : α} : finite a b ↔ ∃ n : ℕ, ¬a ^ (n + 1 ) ∣ b := iff.rfl
47
47
48
+ lemma not_dvd_one_of_finite_one_right {a : α} : finite a 1 → ¬a ∣ 1 :=
49
+ λ ⟨n, hn⟩ ⟨d, hd⟩, hn ⟨d ^ (n + 1 ), (pow_mul_pow_eq_one (n + 1 ) hd.symm).symm⟩
50
+
48
51
@[norm_cast]
49
52
theorem int.coe_nat_multiplicity (a b : ℕ) :
50
53
multiplicity (a : ℤ) (b : ℤ) = multiplicity a b :=
@@ -61,14 +64,10 @@ lemma not_finite_iff_forall {a b : α} : (¬ finite a b) ↔ ∀ n : ℕ, a ^ n
61
64
by simp [finite, multiplicity, not_not]; tauto⟩
62
65
63
66
lemma not_unit_of_finite {a b : α} (h : finite a b) : ¬is_unit a :=
64
- let ⟨n, hn⟩ := h in mt (is_unit_iff_forall_dvd.1 ∘ is_unit.pow (n + 1 )) $
65
- λ h, hn (h b)
66
-
67
- lemma finite_of_finite_mul_left {a b c : α} : finite a (b * c) → finite a c :=
68
- λ ⟨n, hn⟩, ⟨n, λ h, hn (h.trans (by simp [mul_pow]))⟩
67
+ let ⟨n, hn⟩ := h in hn ∘ is_unit.dvd ∘ is_unit.pow (n + 1 )
69
68
70
69
lemma finite_of_finite_mul_right {a b c : α} : finite a (b * c) → finite a b :=
71
- by rw mul_comm; exact finite_of_finite_mul_left
70
+ λ ⟨n, hn⟩, ⟨n, λ h, hn (h.trans (dvd_mul_right _ _))⟩
72
71
73
72
variable [decidable_rel ((∣) : α → α → Prop )]
74
73
@@ -133,29 +132,19 @@ by { simp only [not_not],
133
132
exact ⟨λ h n, nat.cases_on n (by { rw pow_zero, exact one_dvd _}) (λ n, h _), λ h n, h _⟩ }
134
133
135
134
@[simp] lemma is_unit_left {a : α} (b : α) (ha : is_unit a) : multiplicity a b = ⊤ :=
136
- eq_top_iff.2 (λ _, is_unit_iff_forall_dvd.1 (ha.pow _) _)
137
-
138
- lemma is_unit_right {a b : α} (ha : ¬is_unit a) (hb : is_unit b) :
139
- multiplicity a b = 0 :=
140
- eq_coe_iff.2 ⟨show a ^ 0 ∣ b, by simp only [pow_zero, one_dvd],
141
- by { rw pow_one, exact λ h, mt (is_unit_of_dvd_unit h) ha hb }⟩
135
+ eq_top_iff.2 (λ _, is_unit.dvd (ha.pow _))
142
136
143
137
@[simp] lemma one_left (b : α) : multiplicity 1 b = ⊤ := is_unit_left b is_unit_one
144
138
145
- lemma one_right {a : α} (ha : ¬is_unit a) : multiplicity a 1 = 0 := is_unit_right ha is_unit_one
146
-
147
139
@[simp] lemma get_one_right {a : α} (ha : finite a 1 ) : get (multiplicity a 1 ) ha = 0 :=
148
140
begin
149
141
rw [part_enat.get_eq_iff_eq_coe, eq_coe_iff, pow_zero],
150
- simpa [is_unit_iff_dvd_one.symm] using not_unit_of_finite ha ,
142
+ simp [not_dvd_one_of_finite_one_right ha] ,
151
143
end
152
144
153
145
@[simp] lemma unit_left (a : α) (u : αˣ) : multiplicity (u : α) a = ⊤ :=
154
146
is_unit_left a u.is_unit
155
147
156
- lemma unit_right {a : α} (ha : ¬is_unit a) (u : αˣ) : multiplicity a u = 0 :=
157
- is_unit_right ha u.is_unit
158
-
159
148
lemma multiplicity_eq_zero_of_not_dvd {a b : α} (ha : ¬a ∣ b) : multiplicity a b = 0 :=
160
149
by { rw [← nat.cast_zero, eq_coe_iff], simpa }
161
150
@@ -192,14 +181,11 @@ lemma multiplicity_le_multiplicity_iff {a b c d : α} : multiplicity a b ≤ mul
192
181
by rw [eq_top_iff_not_finite.2 hab, eq_top_iff_not_finite.2
193
182
(not_finite_iff_forall.2 this )]⟩
194
183
195
- lemma multiplicity_le_multiplicity_of_dvd_left {a b c : α} (hdvd : a ∣ b) :
196
- multiplicity b c ≤ multiplicity a c :=
197
- multiplicity_le_multiplicity_iff.2 $ λ n h, (pow_dvd_pow_of_dvd hdvd n).trans h
198
-
199
- lemma eq_of_associated_left {a b c : α} (h : associated a b) :
200
- multiplicity b c = multiplicity a c :=
201
- le_antisymm (multiplicity_le_multiplicity_of_dvd_left h.dvd)
202
- (multiplicity_le_multiplicity_of_dvd_left h.symm.dvd)
184
+ lemma multiplicity_eq_multiplicity_iff {a b c d : α} : multiplicity a b = multiplicity c d ↔
185
+ (∀ n : ℕ, a ^ n ∣ b ↔ c ^ n ∣ d) :=
186
+ ⟨λ h n, ⟨multiplicity_le_multiplicity_iff.mp h.le n, multiplicity_le_multiplicity_iff.mp h.ge n⟩,
187
+ λ h, le_antisymm (multiplicity_le_multiplicity_iff.mpr (λ n, (h n).mp))
188
+ (multiplicity_le_multiplicity_iff.mpr (λ n, (h n).mpr))⟩
203
189
204
190
lemma multiplicity_le_multiplicity_of_dvd_right {a b c : α} (h : b ∣ c) :
205
191
multiplicity a b ≤ multiplicity a c :=
@@ -240,11 +226,44 @@ end
240
226
241
227
alias dvd_iff_multiplicity_pos ↔ _ _root_.has_dvd.dvd.multiplicity_pos
242
228
229
+ end monoid
230
+
231
+ section comm_monoid
232
+ variables [comm_monoid α]
233
+
234
+ lemma finite_of_finite_mul_left {a b c : α} : finite a (b * c) → finite a c :=
235
+ by rw mul_comm; exact finite_of_finite_mul_right
236
+
237
+ variable [decidable_rel ((∣) : α → α → Prop )]
238
+
239
+ lemma is_unit_right {a b : α} (ha : ¬is_unit a) (hb : is_unit b) :
240
+ multiplicity a b = 0 :=
241
+ eq_coe_iff.2 ⟨show a ^ 0 ∣ b, by simp only [pow_zero, one_dvd],
242
+ by { rw pow_one, exact λ h, mt (is_unit_of_dvd_unit h) ha hb }⟩
243
+
244
+ lemma one_right {a : α} (ha : ¬is_unit a) : multiplicity a 1 = 0 := is_unit_right ha is_unit_one
245
+
246
+ lemma unit_right {a : α} (ha : ¬is_unit a) (u : αˣ) : multiplicity a u = 0 :=
247
+ is_unit_right ha u.is_unit
248
+
249
+ open_locale classical
250
+
251
+ lemma multiplicity_le_multiplicity_of_dvd_left {a b c : α} (hdvd : a ∣ b) :
252
+ multiplicity b c ≤ multiplicity a c :=
253
+ multiplicity_le_multiplicity_iff.2 $ λ n h, (pow_dvd_pow_of_dvd hdvd n).trans h
254
+
255
+ lemma eq_of_associated_left {a b c : α} (h : associated a b) :
256
+ multiplicity b c = multiplicity a c :=
257
+ le_antisymm (multiplicity_le_multiplicity_of_dvd_left h.dvd)
258
+ (multiplicity_le_multiplicity_of_dvd_left h.symm.dvd)
259
+
260
+ alias dvd_iff_multiplicity_pos ↔ _ _root_.has_dvd.dvd.multiplicity_pos
261
+
243
262
end comm_monoid
244
263
245
- section comm_monoid_with_zero
264
+ section monoid_with_zero
246
265
247
- variable [comm_monoid_with_zero α]
266
+ variable [monoid_with_zero α]
248
267
249
268
lemma ne_zero_of_finite {a b : α} (h : finite a b) : b ≠ 0 :=
250
269
let ⟨n, hn⟩ := h in λ hb, by simpa [hb] using hn
@@ -260,6 +279,14 @@ begin
260
279
rwa zero_dvd_iff,
261
280
end
262
281
282
+ end monoid_with_zero
283
+
284
+ section comm_monoid_with_zero
285
+
286
+ variable [comm_monoid_with_zero α]
287
+
288
+ variable [decidable_rel ((∣) : α → α → Prop )]
289
+
263
290
lemma multiplicity_mk_eq_multiplicity [decidable_rel ((∣) : associates α → associates α → Prop )]
264
291
{a b : α} : multiplicity (associates.mk a) (associates.mk b) = multiplicity a b :=
265
292
begin
@@ -274,15 +301,15 @@ begin
274
301
{ suffices : ¬ (finite (associates.mk a) (associates.mk b)),
275
302
{ rw [finite_iff_dom, part_enat.not_dom_iff_eq_top] at h this ,
276
303
rw [h, this ] },
277
- refine not_finite_iff_forall.mpr (λ n, by {rw [← associates.mk_pow, associates.mk_dvd_mk],
304
+ refine not_finite_iff_forall.mpr (λ n, by { rw [← associates.mk_pow, associates.mk_dvd_mk],
278
305
exact not_finite_iff_forall.mp h n }) }
279
306
end
280
307
281
308
end comm_monoid_with_zero
282
309
283
- section comm_semiring
310
+ section semiring
284
311
285
- variables [comm_semiring α] [decidable_rel ((∣) : α → α → Prop )]
312
+ variables [semiring α] [decidable_rel ((∣) : α → α → Prop )]
286
313
287
314
lemma min_le_multiplicity_add {p a b : α} :
288
315
min (multiplicity p a) (multiplicity p b) ≤ multiplicity p (a + b) :=
@@ -292,13 +319,11 @@ lemma min_le_multiplicity_add {p a b : α} :
292
319
(λ h, by rw [min_eq_right h, multiplicity_le_multiplicity_iff];
293
320
exact λ n hn, dvd_add (multiplicity_le_multiplicity_iff.1 h n hn) hn)
294
321
295
- end comm_semiring
296
-
297
- section comm_ring
322
+ end semiring
298
323
299
- variables [comm_ring α] [decidable_rel ((∣) : α → α → Prop )]
324
+ section ring
300
325
301
- open_locale classical
326
+ variables [ring α] [decidable_rel ((∣) : α → α → Prop )]
302
327
303
328
@[simp] protected lemma neg (a b : α) : multiplicity a (-b) = multiplicity a b :=
304
329
part.ext' (by simp only [multiplicity, part_enat.find, dvd_neg])
@@ -340,7 +365,7 @@ begin
340
365
{ rw [multiplicity_add_of_gt hab, min_eq_right], exact le_of_lt hab},
341
366
end
342
367
343
- end comm_ring
368
+ end ring
344
369
345
370
section cancel_comm_monoid_with_zero
346
371
@@ -476,7 +501,6 @@ lemma multiplicity_pow_self_of_prime {p : α} (hp : prime p) (n : ℕ) :
476
501
multiplicity p (p ^ n) = n :=
477
502
multiplicity_pow_self hp.ne_zero hp.not_unit n
478
503
479
-
480
504
end cancel_comm_monoid_with_zero
481
505
482
506
section valuation
0 commit comments