Skip to content

Commit

Permalink
refactor(linear_algebra/basic): extract content about linear_map.gene…
Browse files Browse the repository at this point in the history
…ral_linear_group

This makes the file a little shorter.
  • Loading branch information
eric-wieser committed Feb 1, 2023
1 parent ae690b0 commit 04f0c28
Show file tree
Hide file tree
Showing 7 changed files with 78 additions and 51 deletions.
1 change: 1 addition & 0 deletions src/linear_algebra/affine_space/affine_equiv.lean
Original file line number Diff line number Diff line change
Expand Up @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Yury G. Kudryashov
-/
import linear_algebra.affine_space.affine_map
import linear_algebra.general_linear_group
import algebra.invertible

/-!
Expand Down
51 changes: 0 additions & 51 deletions src/linear_algebra/basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -25,7 +25,6 @@ Many of the relevant definitions, including `module`, `submodule`, and `linear_m
* Many constructors for (semi)linear maps
* The kernel `ker` and range `range` of a linear map are submodules of the domain and codomain
respectively.
* The general linear group is defined to be the group of invertible linear maps from `M` to itself.
See `linear_algebra.span` for the span of a set (as a submodule),
and `linear_algebra.quotient` for quotients by submodules.
Expand Down Expand Up @@ -2162,54 +2161,4 @@ end linear_equiv

end fun_left

namespace linear_map

variables [semiring R] [add_comm_monoid M] [module R M]
variables (R M)

/-- The group of invertible linear maps from `M` to itself -/
@[reducible] def general_linear_group := (M →ₗ[R] M)ˣ

namespace general_linear_group
variables {R M}

instance : has_coe_to_fun (general_linear_group R M) (λ _, M → M) := by apply_instance

/-- An invertible linear map `f` determines an equivalence from `M` to itself. -/
def to_linear_equiv (f : general_linear_group R M) : (M ≃ₗ[R] M) :=
{ inv_fun := f.inv.to_fun,
left_inv := λ m, show (f.inv * f.val) m = m,
by erw f.inv_val; simp,
right_inv := λ m, show (f.val * f.inv) m = m,
by erw f.val_inv; simp,
..f.val }

/-- An equivalence from `M` to itself determines an invertible linear map. -/
def of_linear_equiv (f : (M ≃ₗ[R] M)) : general_linear_group R M :=
{ val := f,
inv := (f.symm : M →ₗ[R] M),
val_inv := linear_map.ext $ λ _, f.apply_symm_apply _,
inv_val := linear_map.ext $ λ _, f.symm_apply_apply _ }

variables (R M)

/-- The general linear group on `R` and `M` is multiplicatively equivalent to the type of linear
equivalences between `M` and itself. -/
def general_linear_equiv : general_linear_group R M ≃* (M ≃ₗ[R] M) :=
{ to_fun := to_linear_equiv,
inv_fun := of_linear_equiv,
left_inv := λ f, by { ext, refl },
right_inv := λ f, by { ext, refl },
map_mul' := λ x y, by {ext, refl} }

@[simp] lemma general_linear_equiv_to_linear_map (f : general_linear_group R M) :
(general_linear_equiv R M f : M →ₗ[R] M) = f :=
by {ext, refl}

@[simp] lemma coe_fn_general_linear_equiv (f : general_linear_group R M) :
⇑(general_linear_equiv R M f) = (f : M → M) :=
rfl

end general_linear_group

end linear_map
1 change: 1 addition & 0 deletions src/linear_algebra/determinant.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3,6 +3,7 @@ Copyright (c) 2019 Johannes Hölzl. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Johannes Hölzl, Patrick Massot, Casper Putz, Anne Baanen
-/
import linear_algebra.general_linear_group
import linear_algebra.matrix.reindex
import tactic.field_simp
import linear_algebra.matrix.nonsingular_inverse
Expand Down
1 change: 1 addition & 0 deletions src/linear_algebra/eigenspace.lean
Original file line number Diff line number Diff line change
Expand Up @@ -7,6 +7,7 @@ Authors: Alexander Bentkamp
import algebra.algebra.spectrum
import order.hom.basic
import linear_algebra.free_module.finite.basic
import linear_algebra.general_linear_group

/-!
# Eigenvectors and eigenvalues
Expand Down
73 changes: 73 additions & 0 deletions src/linear_algebra/general_linear_group.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,73 @@
/-
Copyright (c) 2019 Johan Commelin. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Johan Commelin
-/
import algebra.module.equiv

/-!
# The general linear group of linear maps
The general linear group is defined to be the group of invertible linear maps from `M` to itself.
See also `matrix.general_linear_group`
## Main definitions
* `linear_map.general_linear_group`
-/

variables (R M : Type*)

namespace linear_map

variables [semiring R] [add_comm_monoid M] [module R M]
variables (R M)

/-- The group of invertible linear maps from `M` to itself -/
@[reducible] def general_linear_group := (M →ₗ[R] M)ˣ

namespace general_linear_group
variables {R M}

instance : has_coe_to_fun (general_linear_group R M) (λ _, M → M) := by apply_instance

/-- An invertible linear map `f` determines an equivalence from `M` to itself. -/
def to_linear_equiv (f : general_linear_group R M) : (M ≃ₗ[R] M) :=
{ inv_fun := f.inv.to_fun,
left_inv := λ m, show (f.inv * f.val) m = m,
by erw f.inv_val; simp,
right_inv := λ m, show (f.val * f.inv) m = m,
by erw f.val_inv; simp,
..f.val }

/-- An equivalence from `M` to itself determines an invertible linear map. -/
def of_linear_equiv (f : (M ≃ₗ[R] M)) : general_linear_group R M :=
{ val := f,
inv := (f.symm : M →ₗ[R] M),
val_inv := linear_map.ext $ λ _, f.apply_symm_apply _,
inv_val := linear_map.ext $ λ _, f.symm_apply_apply _ }

variables (R M)

/-- The general linear group on `R` and `M` is multiplicatively equivalent to the type of linear
equivalences between `M` and itself. -/
def general_linear_equiv : general_linear_group R M ≃* (M ≃ₗ[R] M) :=
{ to_fun := to_linear_equiv,
inv_fun := of_linear_equiv,
left_inv := λ f, by { ext, refl },
right_inv := λ f, by { ext, refl },
map_mul' := λ x y, by {ext, refl} }

@[simp] lemma general_linear_equiv_to_linear_map (f : general_linear_group R M) :
(general_linear_equiv R M f : M →ₗ[R] M) = f :=
by {ext, refl}

@[simp] lemma coe_fn_general_linear_equiv (f : general_linear_group R M) :
⇑(general_linear_equiv R M f) = (f : M → M) :=
rfl

end general_linear_group

end linear_map
1 change: 1 addition & 0 deletions src/linear_algebra/matrix/general_linear_group.lean
Original file line number Diff line number Diff line change
Expand Up @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Chris Birkbeck
-/
import linear_algebra.matrix.nonsingular_inverse
import linear_algebra.matrix.general_linear_group
import linear_algebra.matrix.special_linear_group

/-!
Expand Down
1 change: 1 addition & 0 deletions src/number_theory/modular.lean
Original file line number Diff line number Diff line change
Expand Up @@ -6,6 +6,7 @@ Authors: Alex Kontorovich, Heather Macbeth, Marc Masdeu

import analysis.complex.upper_half_plane.basic
import analysis.normed_space.finite_dimension
import linear_algebra.general_linear_group
import linear_algebra.matrix.general_linear_group

/-!
Expand Down

0 comments on commit 04f0c28

Please sign in to comment.