The special linear group of a module #
If R is a commutative ring and V is an R-module,
we define SpecialLinearGroup R V as the subtype of
linear equivalences V ≃ₗ[R] V with determinant 1.
When V doesn't have a finite basis, the determinant is defined to be 1
and the definition gives V ≃ₗ[R] V.
The interest of this definition is that SpecialLinearGroup R V
has a group structure. (Starting from linear maps wouldn't have worked.)
The file is constructed parallel to the one defining Matrix.SpecialLinearGroup.
We provide SpecialLinearGroup.toLinearEquiv: the canonical map
from SpecialLinearGroup R V to V ≃ₗ[R] V, as a monoid hom.
When V is free and finite over R, we define
We define Matrix.SpecialLinearGroup.toLin'_equiv: the multiplicative equivalence
from Matrix.SpecialLinearGroup n R to SpecialLinearGroup R (n → R)
and its variant
Matrix.SpecialLinearGroup.toLin_equiv,
from Matrix.SpecialLinearGroup n R to SpecialLinearGroup R V,
associated with a finite basis of V.
The special linear group of a module.
This is only meaningful when the module is finite and free, for otherwise, it coincides with the group of linear equivalences.
Equations
- SpecialLinearGroup R V = { u : V ≃ₗ[R] V // LinearEquiv.det u = 1 }
Instances For
Equations
- SpecialLinearGroup.instCoeFunForall = { coe := fun (u : SpecialLinearGroup R V) (x : V) => ↑u x }
Equations
- SpecialLinearGroup.instInv = { inv := fun (A : SpecialLinearGroup R V) => ⟨(↑A)⁻¹, ⋯⟩ }
Equations
- SpecialLinearGroup.instMul = { mul := fun (A B : SpecialLinearGroup R V) => ⟨↑A * ↑B, ⋯⟩ }
Equations
- SpecialLinearGroup.instDiv = { div := fun (A B : SpecialLinearGroup R V) => ⟨↑A / ↑B, ⋯⟩ }
Equations
- SpecialLinearGroup.instPowNat = { pow := fun (x : SpecialLinearGroup R V) (n : ℕ) => ⟨↑x ^ n, ⋯⟩ }
Equations
- SpecialLinearGroup.instPowInt = { pow := fun (x : SpecialLinearGroup R V) (n : ℤ) => ⟨↑x ^ n, ⋯⟩ }
Equations
- SpecialLinearGroup.instInhabited = { default := 1 }
The transpose of an element in SpecialLinearGroup R V.
Instances For
The transpose of an element in SpecialLinearGroup R V.
Equations
- SpecialLinearGroup.«term_ᵀ» = Lean.ParserDescr.trailingNode `SpecialLinearGroup.«term_ᵀ» 1024 1024 (Lean.ParserDescr.symbol "ᵀ")
Instances For
Equations
- One or more equations did not get rendered due to their size.
A version of Matrix.toLin' A that produces linear equivalences.
Equations
- SpecialLinearGroup.toLinearEquiv = { toFun := fun (A : SpecialLinearGroup R V) => ↑A, map_one' := ⋯, map_mul' := ⋯ }
Instances For
The canonical group morphism from the special linear group to the general linear group.
Equations
Instances For
By base change, an R-algebra S induces a group homomorphism from
SpecialLinearGroup R V to SpecialLinearGroup S (S ⊗[R] V).
Equations
- SpecialLinearGroup.baseChange = { toFun := fun (g : SpecialLinearGroup R V) => ⟨LinearEquiv.baseChange R S V V ↑g, ⋯⟩, map_one' := ⋯, map_mul' := ⋯ }
Instances For
The isomorphism between special linear groups of isomorphic modules.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The canonical isomorphism from SL(n, R) to the special linear group of the module n → R.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The isomorphism from Matrix.SpecialLinearGroup n R
to the special linear group of a module associated with a basis of that module.
Equations
Instances For
The inverse map for the equivalence SpecialLinearGroup.centerEquivRootsOfUnity.
Equations
Instances For
The isomorphism between the roots of unity and the center of the special linear group.
Equations
- One or more equations did not get rendered due to their size.