@@ -9,6 +9,7 @@ import order.hom.basic
9
9
import ring_theory.dedekind_domain.basic
10
10
import ring_theory.fractional_ideal
11
11
import ring_theory.principal_ideal_domain
12
+ import ring_theory.chain_of_divisors
12
13
13
14
/-!
14
15
# Dedekind domains and ideals
@@ -925,13 +926,12 @@ begin
925
926
map_comap_of_surjective J^.quotient.mk quotient.mk_surjective, map_map],
926
927
end
927
928
928
- variables [is_domain R] [is_dedekind_domain R]
929
+ variables [is_domain R] [is_dedekind_domain R] (f : R ⧸ I ≃+* A ⧸ J)
929
930
930
931
/-- The bijection between ideals of `R` dividing `I` and the ideals of `A` dividing `J` induced by
931
932
an isomorphism `f : R/I ≅ A/J`. -/
932
933
@[simps]
933
- def ideal_factors_equiv_of_quot_equiv (f : R ⧸ I ≃+* A ⧸ J) :
934
- {p : ideal R | p ∣ I} ≃o {p : ideal A | p ∣ J} :=
934
+ def ideal_factors_equiv_of_quot_equiv : {p : ideal R | p ∣ I} ≃o {p : ideal A | p ∣ J} :=
935
935
order_iso.of_hom_inv
936
936
(ideal_factors_fun_of_quot_hom (show function.surjective
937
937
(f : R ⧸I →+* A ⧸ J), from f.surjective))
@@ -946,6 +946,71 @@ order_iso.of_hom_inv
946
946
← ring_equiv.to_ring_hom_eq_coe, ← ring_equiv.to_ring_hom_trans, ring_equiv.self_trans_symm,
947
947
ring_equiv.to_ring_hom_refl])
948
948
949
+ lemma ideal_factors_equiv_of_quot_equiv_symm :
950
+ (ideal_factors_equiv_of_quot_equiv f).symm = ideal_factors_equiv_of_quot_equiv f.symm := rfl
951
+
952
+ lemma ideal_factors_equiv_of_quot_equiv_is_dvd_iso {L M : ideal R} (hL : L ∣ I) (hM : M ∣ I) :
953
+ (ideal_factors_equiv_of_quot_equiv f ⟨L, hL⟩ : ideal A) ∣
954
+ ideal_factors_equiv_of_quot_equiv f ⟨M, hM⟩ ↔ L ∣ M :=
955
+ begin
956
+ suffices : ideal_factors_equiv_of_quot_equiv f ⟨M, hM⟩ ≤
957
+ ideal_factors_equiv_of_quot_equiv f ⟨L, hL⟩ ↔ (⟨M, hM⟩ : {p : ideal R | p ∣ I}) ≤ ⟨L, hL⟩,
958
+ { rw [dvd_iff_le, dvd_iff_le, subtype.coe_le_coe, this , subtype.mk_le_mk] },
959
+ exact (ideal_factors_equiv_of_quot_equiv f).le_iff_le,
960
+ end
961
+
962
+ open unique_factorization_monoid
963
+
964
+ variables [decidable_eq (ideal R)] [decidable_eq (ideal A)]
965
+
966
+ lemma ideal_factors_equiv_of_quot_equiv_mem_normalized_factors_of_mem_normalized_factors
967
+ (hJ : J ≠ ⊥) {L : ideal R} (hL : L ∈ normalized_factors I) :
968
+ ↑(ideal_factors_equiv_of_quot_equiv f
969
+ ⟨L, dvd_of_mem_normalized_factors hL⟩) ∈ normalized_factors J :=
970
+ begin
971
+ by_cases hI : I = ⊥,
972
+ { exfalso,
973
+ rw [hI, bot_eq_zero, normalized_factors_zero, ← multiset.empty_eq_zero] at hL,
974
+ exact hL, },
975
+ { apply mem_normalized_factors_factor_dvd_iso_of_mem_normalized_factors hI hJ hL _,
976
+ rintros ⟨l, hl⟩ ⟨l', hl'⟩,
977
+ rw [subtype.coe_mk, subtype.coe_mk],
978
+ apply ideal_factors_equiv_of_quot_equiv_is_dvd_iso f }
979
+ end
980
+
981
+ /-- The bijection between the sets of normalized factors of I and J induced by a ring
982
+ isomorphism `f : R/I ≅ A/J`. -/
983
+ @[simps apply]
984
+ def normalized_factors_equiv_of_quot_equiv (hI : I ≠ ⊥) (hJ : J ≠ ⊥) :
985
+ {L : ideal R | L ∈ normalized_factors I } ≃ {M : ideal A | M ∈ normalized_factors J } :=
986
+ { to_fun := λ j, ⟨ideal_factors_equiv_of_quot_equiv f ⟨↑j, dvd_of_mem_normalized_factors j.prop⟩,
987
+ ideal_factors_equiv_of_quot_equiv_mem_normalized_factors_of_mem_normalized_factors f hJ j.prop⟩,
988
+ inv_fun := λ j, ⟨(ideal_factors_equiv_of_quot_equiv f).symm
989
+ ⟨↑j, dvd_of_mem_normalized_factors j.prop⟩, by { rw ideal_factors_equiv_of_quot_equiv_symm,
990
+ exact ideal_factors_equiv_of_quot_equiv_mem_normalized_factors_of_mem_normalized_factors
991
+ f.symm hI j.prop} ⟩,
992
+ left_inv := λ ⟨j, hj⟩, by simp,
993
+ right_inv := λ ⟨j, hj⟩, by simp }
994
+
995
+ @[simp]
996
+ lemma normalized_factors_equiv_of_quot_equiv_symm (hI : I ≠ ⊥) (hJ : J ≠ ⊥) :
997
+ (normalized_factors_equiv_of_quot_equiv f hI hJ).symm =
998
+ normalized_factors_equiv_of_quot_equiv f.symm hJ hI :=
999
+ rfl
1000
+
1001
+ variable [decidable_rel ((∣) : ideal R → ideal R → Prop )]
1002
+ variable [decidable_rel ((∣) : ideal A → ideal A → Prop )]
1003
+
1004
+ /-- The map `normalized_factors_equiv_of_quot_equiv` preserves multiplicities. -/
1005
+ lemma normalized_factors_equiv_of_quot_equiv_multiplicity_eq_multiplicity (hI : I ≠ ⊥) (hJ : J ≠ ⊥)
1006
+ (L : ideal R) (hL : L ∈ normalized_factors I) :
1007
+ multiplicity ↑(normalized_factors_equiv_of_quot_equiv f hI hJ ⟨L, hL⟩) J = multiplicity L I :=
1008
+ begin
1009
+ rw [normalized_factors_equiv_of_quot_equiv, equiv.coe_fn_mk, subtype.coe_mk],
1010
+ exact multiplicity_factor_dvd_iso_eq_multiplicity_of_mem_normalized_factor hI hJ hL
1011
+ (λ ⟨l, hl⟩ ⟨l', hl'⟩, ideal_factors_equiv_of_quot_equiv_is_dvd_iso f hl hl'),
1012
+ end
1013
+
949
1014
end
950
1015
951
1016
section chinese_remainder
@@ -1148,13 +1213,33 @@ begin
1148
1213
exact (prime_of_normalized_factor a ha).ne_zero (span_singleton_eq_bot.mp h) },
1149
1214
end
1150
1215
1216
+ lemma multiplicity_eq_multiplicity_span [decidable_rel ((∣) : R → R → Prop )]
1217
+ [decidable_rel ((∣) : ideal R → ideal R → Prop )] {a b : R} :
1218
+ multiplicity (ideal.span {a}) (ideal.span ({b} : set R)) = multiplicity a b :=
1219
+ begin
1220
+ by_cases h : finite a b,
1221
+ { rw ← part_enat.coe_get (finite_iff_dom.mp h),
1222
+ refine (multiplicity.unique
1223
+ (show (ideal.span {a})^(((multiplicity a b).get h)) ∣ (ideal.span {b}), from _) _).symm ;
1224
+ rw [ideal.span_singleton_pow, span_singleton_dvd_span_singleton_iff_dvd],
1225
+ exact pow_multiplicity_dvd h ,
1226
+ { exact multiplicity.is_greatest ((part_enat.lt_coe_iff _ _).mpr (exists.intro
1227
+ (finite_iff_dom.mp h) (nat.lt_succ_self _))) } },
1228
+ { suffices : ¬ (finite (ideal.span ({a} : set R)) (ideal.span ({b} : set R))),
1229
+ { rw [finite_iff_dom, part_enat.not_dom_iff_eq_top] at h this ,
1230
+ rw [h, this ] },
1231
+ refine not_finite_iff_forall.mpr (λ n, by {rw [ideal.span_singleton_pow,
1232
+ span_singleton_dvd_span_singleton_iff_dvd], exact not_finite_iff_forall.mp h n }) }
1233
+ end
1234
+
1235
+ variables [decidable_eq R] [decidable_eq (ideal R)] [normalization_monoid R]
1236
+
1151
1237
/-- The bijection between the (normalized) prime factors of `r` and the (normalized) prime factors
1152
1238
of `span {r}` -/
1153
1239
@[simps]
1154
- noncomputable def normalized_factors_equiv_span_normalized_factors [normalization_monoid R]
1155
- [decidable_eq R] [decidable_eq (ideal R)] {r : R} (hr : r ≠ 0 ) :
1240
+ noncomputable def normalized_factors_equiv_span_normalized_factors {r : R} (hr : r ≠ 0 ) :
1156
1241
{d : R | d ∈ normalized_factors r} ≃
1157
- {I : ideal R | I ∈ normalized_factors (ideal.span ({r} : set R))} :=
1242
+ {I : ideal R | I ∈ normalized_factors (ideal.span ({r} : set R))} :=
1158
1243
equiv.of_bijective
1159
1244
(λ d, ⟨ideal.span {↑d}, singleton_span_mem_normalized_factors_of_mem_normalized_factors d.prop⟩)
1160
1245
begin
@@ -1176,23 +1261,29 @@ begin
1176
1261
dvd_iff_le.mp (dvd_of_mem_normalized_factors hi))) (mem_span_singleton.mpr (dvd_refl r))) } }
1177
1262
end
1178
1263
1179
- lemma multiplicity_eq_multiplicity_span [decidable_rel ((∣) : R → R → Prop )]
1180
- [decidable_rel ((∣) : ideal R → ideal R → Prop )] {a b : R} :
1181
- multiplicity (ideal.span {a}) (ideal.span ({b} : set R)) = multiplicity a b :=
1264
+ variables [decidable_rel ((∣) : R → R → Prop )] [decidable_rel ((∣) : ideal R → ideal R → Prop )]
1265
+
1266
+ /-- The bijection `normalized_factors_equiv_span_normalized_factors` between the set of prime
1267
+ factors of `r` and the set of prime factors of the ideal `⟨r⟩` preserves multiplicities. -/
1268
+ lemma multiplicity_normalized_factors_equiv_span_normalized_factors_eq_multiplicity {r d: R}
1269
+ (hr : r ≠ 0 ) (hd : d ∈ normalized_factors r) :
1270
+ multiplicity d r =
1271
+ multiplicity (normalized_factors_equiv_span_normalized_factors hr ⟨d, hd⟩ : ideal R)
1272
+ (ideal.span {r}) :=
1273
+ by simp only [normalized_factors_equiv_span_normalized_factors, multiplicity_eq_multiplicity_span,
1274
+ subtype.coe_mk, equiv.of_bijective_apply]
1275
+
1276
+ /-- The bijection `normalized_factors_equiv_span_normalized_factors.symm` between the set of prime
1277
+ factors of the ideal `⟨r⟩` and the set of prime factors of `r` preserves multiplicities. -/
1278
+ lemma multiplicity_normalized_factors_equiv_span_normalized_factors_symm_eq_multiplicity
1279
+ {r : R} (hr : r ≠ 0 ) (I : {I : ideal R | I ∈ normalized_factors (ideal.span ({r} : set R))}) :
1280
+ multiplicity ((normalized_factors_equiv_span_normalized_factors hr).symm I : R) r =
1281
+ multiplicity (I : ideal R) (ideal.span {r}) :=
1182
1282
begin
1183
- by_cases h : finite a b,
1184
- { rw ← part_enat.coe_get (finite_iff_dom.mp h),
1185
- refine (multiplicity.unique
1186
- (show (ideal.span {a})^(((multiplicity a b).get h)) ∣ (ideal.span {b}), from _) _).symm ;
1187
- rw [ideal.span_singleton_pow, span_singleton_dvd_span_singleton_iff_dvd],
1188
- exact pow_multiplicity_dvd h ,
1189
- { exact multiplicity.is_greatest ((part_enat.lt_coe_iff _ _).mpr (exists.intro
1190
- (finite_iff_dom.mp h) (nat.lt_succ_self _))) } },
1191
- { suffices : ¬ (finite (ideal.span ({a} : set R)) (ideal.span ({b} : set R))),
1192
- { rw [finite_iff_dom, part_enat.not_dom_iff_eq_top] at h this ,
1193
- rw [h, this ] },
1194
- refine not_finite_iff_forall.mpr (λ n, by {rw [ideal.span_singleton_pow,
1195
- span_singleton_dvd_span_singleton_iff_dvd], exact not_finite_iff_forall.mp h n }) }
1283
+ obtain ⟨x, hx⟩ := (normalized_factors_equiv_span_normalized_factors hr).surjective I,
1284
+ obtain ⟨a, ha⟩ := x,
1285
+ rw [hx.symm, equiv.symm_apply_apply, subtype.coe_mk,
1286
+ multiplicity_normalized_factors_equiv_span_normalized_factors_eq_multiplicity hr ha, hx],
1196
1287
end
1197
1288
1198
1289
end PID
0 commit comments