Skip to content

Commit 4180b90

Browse files
committed
refactor(LinearAlgebra,RingTheory): Replace Fintype.card with Nat.card (#13637)
This PR replaces a few occurrences of `Fintype.card` with `Nat.card` in the `LinearAlgebra` and `RingTheory` folders.
1 parent 7733e60 commit 4180b90

File tree

4 files changed

+22
-57
lines changed

4 files changed

+22
-57
lines changed

Mathlib/LinearAlgebra/Isomorphisms.lean

Lines changed: 3 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -188,11 +188,10 @@ def quotientQuotientEquivQuotientSup : ((M ⧸ S) ⧸ T.map S.mkQ) ≃ₗ[R] M
188188
quotientQuotientEquivQuotient S (S ⊔ T) le_sup_left
189189

190190
/-- Corollary of the third isomorphism theorem: `[S : T] [M : S] = [M : T]` -/
191-
theorem card_quotient_mul_card_quotient (S T : Submodule R M) (hST : T ≤ S)
192-
[DecidablePred fun x => x ∈ S.map T.mkQ] [Fintype (M ⧸ S)] [Fintype (M ⧸ T)] :
193-
Fintype.card (S.map T.mkQ) * Fintype.card (M ⧸ S) = Fintype.card (M ⧸ T) := by
191+
theorem card_quotient_mul_card_quotient (S T : Submodule R M) (hST : T ≤ S) :
192+
Nat.card (S.map T.mkQ) * Nat.card (M ⧸ S) = Nat.card (M ⧸ T) := by
194193
rw [Submodule.card_eq_card_quotient_mul_card (map T.mkQ S),
195-
Fintype.card_eq.mpr ⟨(quotientQuotientEquivQuotient T S hST).toEquiv]
194+
Nat.card_congr (quotientQuotientEquivQuotient T S hST).toEquiv]
196195
#align submodule.card_quotient_mul_card_quotient Submodule.card_quotient_mul_card_quotient
197196

198197
end Submodule

Mathlib/LinearAlgebra/Quotient.lean

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -304,10 +304,10 @@ noncomputable instance Quotient.fintype [Fintype M] (S : Submodule R M) : Fintyp
304304
@_root_.Quotient.fintype _ _ _ fun _ _ => Classical.dec _
305305
#align submodule.quotient.fintype Submodule.Quotient.fintype
306306

307-
theorem card_eq_card_quotient_mul_card [Fintype M] (S : Submodule R M) [DecidablePred (· ∈ S)] :
308-
Fintype.card M = Fintype.card S * Fintype.card (M ⧸ S) := by
309-
rw [mul_comm, ← Fintype.card_prod]
310-
exact Fintype.card_congr AddSubgroup.addGroupEquivQuotientProdAddSubgroup
307+
theorem card_eq_card_quotient_mul_card (S : Submodule R M) :
308+
Nat.card M = Nat.card S * Nat.card (M ⧸ S) := by
309+
rw [mul_comm, ← Nat.card_prod]
310+
exact Nat.card_congr AddSubgroup.addGroupEquivQuotientProdAddSubgroup
311311
#align submodule.card_eq_card_quotient_mul_card Submodule.card_eq_card_quotient_mul_card
312312

313313
section

Mathlib/NumberTheory/Cyclotomic/Rat.lean

Lines changed: 1 addition & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -201,8 +201,7 @@ lemma card_quotient_toInteger_sub_one [NumberField K] {k : ℕ+} (hk : 1 < k)
201201
(Algebra.norm ℤ (hζ.toInteger - 1)).natAbs := by
202202
have := hζ.finite_quotient_toInteger_sub_one hk
203203
let _ := Fintype.ofFinite (𝓞 K ⧸ Ideal.span {hζ.toInteger - 1})
204-
rw [Nat.card_eq_fintype_card, ← Submodule.cardQuot_apply, ← Ideal.absNorm_apply,
205-
Ideal.absNorm_span_singleton]
204+
rw [← Submodule.cardQuot_apply, ← Ideal.absNorm_apply, Ideal.absNorm_span_singleton]
206205

207206
lemma toInteger_isPrimitiveRoot {k : ℕ+} (hζ : IsPrimitiveRoot ζ k) :
208207
IsPrimitiveRoot hζ.toInteger k :=

Mathlib/RingTheory/Ideal/Norm.lean

Lines changed: 14 additions & 47 deletions
Original file line numberDiff line numberDiff line change
@@ -66,12 +66,8 @@ noncomputable def cardQuot (S : Submodule R M) : ℕ :=
6666
AddSubgroup.index S.toAddSubgroup
6767
#align submodule.card_quot Submodule.cardQuot
6868

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
7571
#align submodule.card_quot_apply Submodule.cardQuot_apply
7672

7773
variable (R M)
@@ -81,7 +77,7 @@ theorem cardQuot_bot [Infinite M] : cardQuot (⊥ : Submodule R M) = 0 :=
8177
AddSubgroup.index_bot.trans Nat.card_eq_zero_of_infinite
8278
#align submodule.card_quot_bot Submodule.cardQuot_bot
8379

84-
-- @[simp] -- Porting note (#10618): simp can prove this
80+
@[simp]
8581
theorem cardQuot_top : cardQuot (⊤ : Submodule R M) = 1 :=
8682
AddSubgroup.index_top
8783
#align submodule.card_quot_top Submodule.cardQuot_top
@@ -106,27 +102,11 @@ open Submodule
106102
/-- Multiplicity of the ideal norm, for coprime ideals.
107103
This is essentially just a repackaging of the Chinese Remainder Theorem.
108104
-/
109-
theorem cardQuot_mul_of_coprime [Module.Free ℤ S] [Module.Finite ℤ S]
105+
theorem cardQuot_mul_of_coprime
110106
{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
127107
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]
130110
#align card_quot_mul_of_coprime cardQuot_mul_of_coprime
131111

132112
/-- 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'
184164
#align ideal.mul_add_mem_pow_succ_unique Ideal.mul_add_mem_pow_succ_unique
185165

186166
/-- 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 : ℕ} :
188168
cardQuot (P ^ i) = cardQuot P ^ i := by
189-
let _ := Module.Free.chooseBasis ℤ S
190-
classical
191169
induction' i with i ih
192170
· 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
196171
have : P ^ (i + 1) < P ^ i := Ideal.pow_succ_lt_pow hP i
197172
suffices hquot : map (P ^ i.succ).mkQ (P ^ i) ≃ S ⧸ P by
198173
rw [pow_succ' (cardQuot P), ← ih, cardQuot_apply (P ^ i.succ), ←
199174
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]
204176
choose a a_mem a_not_mem using SetLike.exists_of_lt this
205177
choose f g hg hf using fun c (hc : c ∈ P ^ i) =>
206178
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
228200
end PPrime
229201

230202
/-- 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) :
232204
cardQuot (I * J) = cardQuot I * cardQuot J := by
233205
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›
239206
haveI : Infinite S := Infinite.of_surjective _ b.repr.toEquiv.surjective
240207
exact UniqueFactorizationMonoid.multiplicative_of_coprime cardQuot I J (cardQuot_bot _ _)
241208
(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 ℤ
248215
#align card_quot_mul cardQuot_mul
249216

250217
/-- 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
253220
toFun := Submodule.cardQuot
254221
map_mul' I J := by dsimp only; rw [cardQuot_mul]
255222
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
331298
_ = Int.natAbs (Matrix.diagonal a).det := ?_
332299
_ = Int.natAbs (∏ i, a i) := by rw [Matrix.det_diagonal]
333300
_ = ∏ i, Int.natAbs (a i) := map_prod Int.natAbsHom a Finset.univ
334-
_ = Fintype.card (S ⧸ I) := ?_
301+
_ = Nat.card (S ⧸ I) := ?_
335302
_ = absNorm I := (Submodule.cardQuot_apply _).symm
336303
-- since `LinearMap.toMatrix b' b' f` is the diagonal matrix with `a` along the diagonal.
337304
· congr 2; ext i j
@@ -344,8 +311,8 @@ theorem natAbs_det_equiv (I : Ideal S) {E : Type*} [EquivLike E S I] [AddEquivCl
344311
-- which maps `(S ⧸ I)` to `Π i, ZMod (a i).nat_abs`.
345312
haveI : ∀ i, NeZero (a i).natAbs := fun i =>
346313
⟨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]
349316
#align ideal.nat_abs_det_equiv Ideal.natAbs_det_equiv
350317

351318
/-- Let `b` be a basis for `S` over `ℤ` and `bI` a basis for `I` over `ℤ` of the same dimension.

0 commit comments

Comments
 (0)