Documentation

Mathlib.LinearAlgebra.FreeModule.Finite.CardQuotient

Cardinal of quotient of free finite -modules by submodules of full rank #

Main results #

theorem Submodule.natAbs_det_equiv {M : Type u_1} [AddCommGroup M] [Module.Free M] [Module.Finite M] (N : Submodule M) {E : Type u_2} [EquivLike E M N] [AddEquivClass E M N] (e : E) :

Let e : M ≃ N be an additive isomorphism (therefore a -linear equiv). Then an alternative way to compute the cardinality of the quotient M ⧸ N is given by taking the determinant of e. See natAbs_det_basis_change for a more familiar formulation of this result.

theorem Submodule.natAbs_det_basis_change {M : Type u_1} [AddCommGroup M] [Module.Free M] [Module.Finite M] {ι : Type u_2} [Fintype ι] [DecidableEq ι] (b : Basis ι M) (N : Submodule M) (bN : Basis ι N) :
(b.det (Subtype.val bN)).natAbs = Nat.card (M N)

Let b be a basis for M over and bN a basis for N over of the same dimension. Then an alternative way to compute the cardinality of M ⧸ N is given by taking the determinant of bN over b.

theorem AddSubgroup.index_eq_natAbs_det {E : Type u_1} [AddCommGroup E] {ι : Type u_2} [DecidableEq ι] [Fintype ι] (bE : Basis ι E) (N : AddSubgroup E) (bN : Basis ι N) :
N.index = (bE.det fun (x : ι) => (bN x)).natAbs
theorem AddSubgroup.relindex_eq_natAbs_det {E : Type u_1} [AddCommGroup E] (L₁ L₂ : AddSubgroup E) (H : L₁ L₂) {ι : Type u_2} [DecidableEq ι] [Fintype ι] (b₁ : Basis ι (toIntSubmodule L₁)) (b₂ : Basis ι (toIntSubmodule L₂)) :
L₁.relindex L₂ = (b₂.det fun (i : ι) => (b₁ i), ).natAbs
theorem AddSubgroup.relindex_eq_abs_det {E : Type u_1} [AddCommGroup E] [Module E] (L₁ L₂ : AddSubgroup E) (H : L₁ L₂) {ι : Type u_2} [DecidableEq ι] [Fintype ι] (b₁ b₂ : Basis ι E) (h₁ : L₁ = closure (Set.range b₁)) (h₂ : L₂ = closure (Set.range b₂)) :
(L₁.relindex L₂) = |b₂.det b₁|