Documentation

Mathlib.LinearAlgebra.SpecialLinearGroup

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 by 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.SpecialLinaerGruop.toLin'_equiv: the mul 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.

@[reducible, inline]
abbrev SpecialLinearGroup (R : Type u_1) (V : Type u_2) [CommRing R] [AddCommGroup V] [Module R V] :
Type u_2

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
Instances For
    theorem SpecialLinearGroup.det_eq_one {R : Type u_1} {V : Type u_2} [CommRing R] [AddCommGroup V] [Module R V] (u : SpecialLinearGroup R V) :
    LinearMap.det u = 1
    instance SpecialLinearGroup.instCoeFunForall {R : Type u_1} {V : Type u_2} [CommRing R] [AddCommGroup V] [Module R V] :
    CoeFun (SpecialLinearGroup R V) fun (x : SpecialLinearGroup R V) => VV
    Equations
    theorem SpecialLinearGroup.ext_iff {R : Type u_1} {V : Type u_2} [CommRing R] [AddCommGroup V] [Module R V] (u v : SpecialLinearGroup R V) :
    u = v ∀ (x : V), (fun (x : V) => u x) x = (fun (x : V) => v x) x
    theorem SpecialLinearGroup.ext {R : Type u_1} {V : Type u_2} [CommRing R] [AddCommGroup V] [Module R V] (u v : SpecialLinearGroup R V) :
    (∀ (x : V), (fun (x : V) => u x) x = (fun (x : V) => v x) x)u = v
    instance SpecialLinearGroup.instInv {R : Type u_1} {V : Type u_2} [CommRing R] [AddCommGroup V] [Module R V] :
    Equations
    instance SpecialLinearGroup.instMul {R : Type u_1} {V : Type u_2} [CommRing R] [AddCommGroup V] [Module R V] :
    Equations
    instance SpecialLinearGroup.instDiv {R : Type u_1} {V : Type u_2} [CommRing R] [AddCommGroup V] [Module R V] :
    Equations
    instance SpecialLinearGroup.instOne {R : Type u_1} {V : Type u_2} [CommRing R] [AddCommGroup V] [Module R V] :
    Equations
    instance SpecialLinearGroup.instPowNat {R : Type u_1} {V : Type u_2} [CommRing R] [AddCommGroup V] [Module R V] :
    Equations
    instance SpecialLinearGroup.instPowInt {R : Type u_1} {V : Type u_2} [CommRing R] [AddCommGroup V] [Module R V] :
    Equations

    The transpose of an element in SpecialLinearGroup R V.

    Equations
    Instances For

      The transpose of an element in SpecialLinearGroup R V.

      Equations
      Instances For
        theorem SpecialLinearGroup.coe_mk {R : Type u_1} {V : Type u_2} [CommRing R] [AddCommGroup V] [Module R V] (A : V ≃ₗ[R] V) (h : LinearEquiv.det A = 1) :
        A, h = A
        @[simp]
        theorem SpecialLinearGroup.coe_mul {R : Type u_1} {V : Type u_2} [CommRing R] [AddCommGroup V] [Module R V] (A B : SpecialLinearGroup R V) :
        ↑(A * B) = A * B
        @[simp]
        theorem SpecialLinearGroup.coe_div {R : Type u_1} {V : Type u_2} [CommRing R] [AddCommGroup V] [Module R V] (A B : SpecialLinearGroup R V) :
        ↑(A / B) = A / B
        @[simp]
        theorem SpecialLinearGroup.coe_inv {R : Type u_1} {V : Type u_2} [CommRing R] [AddCommGroup V] [Module R V] (A : SpecialLinearGroup R V) :
        A⁻¹ = (↑A)⁻¹
        @[simp]
        theorem SpecialLinearGroup.coe_one {R : Type u_1} {V : Type u_2} [CommRing R] [AddCommGroup V] [Module R V] :
        1 = 1
        @[simp]
        theorem SpecialLinearGroup.det_coe {R : Type u_1} {V : Type u_2} [CommRing R] [AddCommGroup V] [Module R V] (A : SpecialLinearGroup R V) :
        @[simp]
        theorem SpecialLinearGroup.coe_pow {R : Type u_1} {V : Type u_2} [CommRing R] [AddCommGroup V] [Module R V] (A : SpecialLinearGroup R V) (m : ) :
        ↑(A ^ m) = A ^ m
        @[simp]
        theorem SpecialLinearGroup.coe_zpow {R : Type u_1} {V : Type u_2} [CommRing R] [AddCommGroup V] [Module R V] (A : SpecialLinearGroup R V) (m : ) :
        ↑(A ^ m) = A ^ m
        @[simp]
        theorem SpecialLinearGroup.coe_dualMap {R : Type u_1} {V : Type u_2} [CommRing R] [AddCommGroup V] [Module R V] (A : SpecialLinearGroup R V) [Module.Free R V] [Module.Finite R V] :
        A.dualMap = (↑A).dualMap
        instance SpecialLinearGroup.instGroup {R : Type u_1} {V : Type u_2} [CommRing R] [AddCommGroup V] [Module R V] :
        Equations
        • One or more equations did not get rendered due to their size.

        A version of Matrix.toLin' A that produces linear equivalences.

        Equations
        Instances For
          theorem SpecialLinearGroup.toLinearEquiv_apply {R : Type u_1} {V : Type u_2} [CommRing R] [AddCommGroup V] [Module R V] (A : SpecialLinearGroup R V) (v : V) :
          (toLinearEquiv A) v = (fun (x : V) => A x) v
          theorem SpecialLinearGroup.toLinearEquiv_symm_apply {R : Type u_1} {V : Type u_2} [CommRing R] [AddCommGroup V] [Module R V] (A : SpecialLinearGroup R V) (v : V) :
          (toLinearEquiv A).symm v = (fun (x : V) => A⁻¹ x) v

          By base change, an R-algebra S induces a group homomorphism from SpecialLinearGroup R V to SpecialLinearGroup S (S ⊗[R] V).

          Equations
          Instances For
            @[simp]
            theorem SpecialLinearGroup.baseChange_apply_coe {R : Type u_1} {V : Type u_2} [CommRing R] [AddCommGroup V] [Module R V] {S : Type u_3} [CommRing S] [Algebra R S] [Module.Free R V] [Module.Finite R V] (g : SpecialLinearGroup R V) :
            (baseChange g) = LinearEquiv.baseChange R S V V g

            The isomorphism between special linear groups of isomorphic modules.

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              @[simp]
              theorem SpecialLinearGroup.congr_linearEquiv_coe_apply {R : Type u_1} {V : Type u_2} [CommRing R] [AddCommGroup V] [Module R V] {W : Type u_3} [AddCommGroup W] [Module R W] (e : V ≃ₗ[R] W) (f : SpecialLinearGroup R V) :
              @[simp]
              theorem SpecialLinearGroup.congr_linearEquiv_apply_apply {R : Type u_1} {V : Type u_2} [CommRing R] [AddCommGroup V] [Module R V] {W : Type u_3} [AddCommGroup W] [Module R W] (e : V ≃ₗ[R] W) (f : SpecialLinearGroup R V) (x : W) :
              (fun (x : W) => ((congr_linearEquiv e) f) x) x = e ((fun (x : V) => f x) (e.symm x))

              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
                noncomputable def Matrix.SpecialLinearGroup.toLin_equiv {R : Type u_1} {V : Type u_2} [CommRing R] [AddCommGroup V] [Module R V] {n : Type u_3} [Fintype n] [DecidableEq n] (b : Module.Basis n R V) :

                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