@@ -66,12 +66,8 @@ noncomputable def cardQuot (S : Submodule R M) : ℕ :=
66
66
AddSubgroup.index S.toAddSubgroup
67
67
#align submodule.card_quot Submodule.cardQuot
68
68
69
- @[simp]
70
- theorem cardQuot_apply (S : Submodule R M) [h : Fintype (M ⧸ S)] :
71
- cardQuot S = Fintype.card (M ⧸ S) := by
72
- -- Porting note: original proof was AddSubgroup.index_eq_card _
73
- suffices Fintype (M ⧸ S.toAddSubgroup) by convert AddSubgroup.index_eq_card S.toAddSubgroup
74
- convert h
69
+ theorem cardQuot_apply (S : Submodule R M) : cardQuot S = Nat.card (M ⧸ S) := by
70
+ rfl
75
71
#align submodule.card_quot_apply Submodule.cardQuot_apply
76
72
77
73
variable (R M)
@@ -81,7 +77,7 @@ theorem cardQuot_bot [Infinite M] : cardQuot (⊥ : Submodule R M) = 0 :=
81
77
AddSubgroup.index_bot.trans Nat.card_eq_zero_of_infinite
82
78
#align submodule.card_quot_bot Submodule.cardQuot_bot
83
79
84
- -- @[ simp ] -- Porting note (#10618): simp can prove this
80
+ @[simp]
85
81
theorem cardQuot_top : cardQuot (⊤ : Submodule R M) = 1 :=
86
82
AddSubgroup.index_top
87
83
#align submodule.card_quot_top Submodule.cardQuot_top
@@ -106,27 +102,11 @@ open Submodule
106
102
/-- Multiplicity of the ideal norm, for coprime ideals.
107
103
This is essentially just a repackaging of the Chinese Remainder Theorem.
108
104
-/
109
- theorem cardQuot_mul_of_coprime [Module.Free ℤ S] [Module.Finite ℤ S]
105
+ theorem cardQuot_mul_of_coprime
110
106
{I J : Ideal S} (coprime : IsCoprime I J) : cardQuot (I * J) = cardQuot I * cardQuot J := by
111
- let b := Module.Free.chooseBasis ℤ S
112
- cases isEmpty_or_nonempty (Module.Free.ChooseBasisIndex ℤ S)
113
- · haveI : Subsingleton S := Function.Surjective.subsingleton b.repr.toEquiv.symm.surjective
114
- nontriviality S
115
- exfalso
116
- exact not_nontrivial_iff_subsingleton.mpr ‹Subsingleton S› ‹Nontrivial S›
117
- haveI : Infinite S := Infinite.of_surjective _ b.repr.toEquiv.surjective
118
- by_cases hI : I = ⊥
119
- · rw [hI, Submodule.bot_mul, cardQuot_bot, zero_mul]
120
- by_cases hJ : J = ⊥
121
- · rw [hJ, Submodule.mul_bot, cardQuot_bot, mul_zero]
122
- have hIJ : I * J ≠ ⊥ := mt Ideal.mul_eq_bot.mp (not_or_of_not hI hJ)
123
- letI := Classical.decEq (Module.Free.ChooseBasisIndex ℤ S)
124
- letI := I.fintypeQuotientOfFreeOfNeBot hI
125
- letI := J.fintypeQuotientOfFreeOfNeBot hJ
126
- letI := (I * J).fintypeQuotientOfFreeOfNeBot hIJ
127
107
rw [cardQuot_apply, cardQuot_apply, cardQuot_apply,
128
- Fintype.card_eq.mpr ⟨ (Ideal.quotientMulEquivQuotientProd I J coprime).toEquiv⟩ ,
129
- Fintype .card_prod]
108
+ Nat.card_congr (Ideal.quotientMulEquivQuotientProd I J coprime).toEquiv,
109
+ Nat .card_prod]
130
110
#align card_quot_mul_of_coprime cardQuot_mul_of_coprime
131
111
132
112
/-- If the `d` from `Ideal.exists_mul_add_mem_pow_succ` is unique, up to `P`,
@@ -184,23 +164,15 @@ theorem Ideal.mul_add_mem_pow_succ_unique [IsDedekindDomain S] {i : ℕ} (a d d'
184
164
#align ideal.mul_add_mem_pow_succ_unique Ideal.mul_add_mem_pow_succ_unique
185
165
186
166
/-- Multiplicity of the ideal norm, for powers of prime ideals. -/
187
- theorem cardQuot_pow_of_prime [IsDedekindDomain S] [Module.Finite ℤ S] [Module.Free ℤ S] {i : ℕ} :
167
+ theorem cardQuot_pow_of_prime [IsDedekindDomain S] {i : ℕ} :
188
168
cardQuot (P ^ i) = cardQuot P ^ i := by
189
- let _ := Module.Free.chooseBasis ℤ S
190
- classical
191
169
induction' i with i ih
192
170
· simp
193
- letI := Ideal.fintypeQuotientOfFreeOfNeBot (P ^ i.succ) (pow_ne_zero _ hP)
194
- letI := Ideal.fintypeQuotientOfFreeOfNeBot (P ^ i) (pow_ne_zero _ hP)
195
- letI := Ideal.fintypeQuotientOfFreeOfNeBot P hP
196
171
have : P ^ (i + 1 ) < P ^ i := Ideal.pow_succ_lt_pow hP i
197
172
suffices hquot : map (P ^ i.succ).mkQ (P ^ i) ≃ S ⧸ P by
198
173
rw [pow_succ' (cardQuot P), ← ih, cardQuot_apply (P ^ i.succ), ←
199
174
card_quotient_mul_card_quotient (P ^ i) (P ^ i.succ) this.le, cardQuot_apply (P ^ i),
200
- cardQuot_apply P]
201
- congr 1
202
- rw [Fintype.card_eq]
203
- exact ⟨hquot⟩
175
+ cardQuot_apply P, Nat.card_congr hquot]
204
176
choose a a_mem a_not_mem using SetLike.exists_of_lt this
205
177
choose f g hg hf using fun c (hc : c ∈ P ^ i) =>
206
178
Ideal.exists_mul_add_mem_pow_succ hP a c a_mem a_not_mem hc
@@ -228,14 +200,9 @@ theorem cardQuot_pow_of_prime [IsDedekindDomain S] [Module.Finite ℤ S] [Module
228
200
end PPrime
229
201
230
202
/-- Multiplicativity of the ideal norm in number rings. -/
231
- theorem cardQuot_mul [IsDedekindDomain S] [Module.Free ℤ S] [Module.Finite ℤ S] (I J : Ideal S) :
203
+ theorem cardQuot_mul [IsDedekindDomain S] [Module.Free ℤ S] (I J : Ideal S) :
232
204
cardQuot (I * J) = cardQuot I * cardQuot J := by
233
205
let b := Module.Free.chooseBasis ℤ S
234
- cases isEmpty_or_nonempty (Module.Free.ChooseBasisIndex ℤ S)
235
- · haveI : Subsingleton S := Function.Surjective.subsingleton b.repr.toEquiv.symm.surjective
236
- nontriviality S
237
- exfalso
238
- exact not_nontrivial_iff_subsingleton.mpr ‹Subsingleton S› ‹Nontrivial S›
239
206
haveI : Infinite S := Infinite.of_surjective _ b.repr.toEquiv.surjective
240
207
exact UniqueFactorizationMonoid.multiplicative_of_coprime cardQuot I J (cardQuot_bot _ _)
241
208
(fun {I J} hI => by simp [Ideal.isUnit_iff.mp hI, Ideal.mul_top])
@@ -248,8 +215,8 @@ theorem cardQuot_mul [IsDedekindDomain S] [Module.Free ℤ S] [Module.Finite ℤ
248
215
#align card_quot_mul cardQuot_mul
249
216
250
217
/-- The absolute norm of the ideal `I : Ideal R` is the cardinality of the quotient `R ⧸ I`. -/
251
- noncomputable def Ideal.absNorm [Nontrivial S] [IsDedekindDomain S] [Module.Free ℤ S]
252
- [Module.Finite ℤ S] : Ideal S →*₀ ℕ where
218
+ noncomputable def Ideal.absNorm [Nontrivial S] [IsDedekindDomain S] [Module.Free ℤ S] :
219
+ Ideal S →*₀ ℕ where
253
220
toFun := Submodule.cardQuot
254
221
map_mul' I J := by dsimp only; rw [cardQuot_mul]
255
222
map_one' := by dsimp only; rw [Ideal.one_eq_top, cardQuot_top]
@@ -331,7 +298,7 @@ theorem natAbs_det_equiv (I : Ideal S) {E : Type*} [EquivLike E S I] [AddEquivCl
331
298
_ = Int.natAbs (Matrix.diagonal a).det := ?_
332
299
_ = Int.natAbs (∏ i, a i) := by rw [Matrix.det_diagonal]
333
300
_ = ∏ i, Int.natAbs (a i) := map_prod Int.natAbsHom a Finset.univ
334
- _ = Fintype .card (S ⧸ I) := ?_
301
+ _ = Nat .card (S ⧸ I) := ?_
335
302
_ = absNorm I := (Submodule.cardQuot_apply _).symm
336
303
-- since `LinearMap.toMatrix b' b' f` is the diagonal matrix with `a` along the diagonal.
337
304
· congr 2 ; ext i j
@@ -344,8 +311,8 @@ theorem natAbs_det_equiv (I : Ideal S) {E : Type*} [EquivLike E S I] [AddEquivCl
344
311
-- which maps `(S ⧸ I)` to `Π i, ZMod (a i).nat_abs`.
345
312
haveI : ∀ i, NeZero (a i).natAbs := fun i =>
346
313
⟨Int.natAbs_ne_zero.mpr (Ideal.smithCoeffs_ne_zero b I hI i)⟩
347
- simp_rw [Fintype.card_eq.mpr ⟨ (Ideal.quotientEquivPiZMod I b hI).toEquiv⟩, Fintype .card_pi,
348
- ZMod.card ]
314
+ simp_rw [Nat.card_congr (Ideal.quotientEquivPiZMod I b hI).toEquiv, Nat .card_pi,
315
+ Nat.card_zmod ]
349
316
#align ideal.nat_abs_det_equiv Ideal.natAbs_det_equiv
350
317
351
318
/-- Let `b` be a basis for `S` over `ℤ` and `bI` a basis for `I` over `ℤ` of the same dimension.
0 commit comments