Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit abe81bc

Browse files
committed
feat(linear_algebra/matrix/general_linear_group): GL(n, R) (#8466)
added this file which contains definition of the general linear group as well as the subgroup of matrices with positive determinant.
1 parent edb22fe commit abe81bc

File tree

3 files changed

+196
-0
lines changed

3 files changed

+196
-0
lines changed
Lines changed: 170 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,170 @@
1+
/-
2+
Copyright (c) 2021 Chris Birkbeck. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Chris Birkbeck
5+
-/
6+
import linear_algebra.matrix
7+
import linear_algebra.matrix.nonsingular_inverse
8+
import linear_algebra.special_linear_group
9+
import linear_algebra.determinant
10+
11+
/-!
12+
# The General Linear group $GL(n, R)$
13+
This file defines the elements of the General Linear group `general_linear_group n R`,
14+
consisting of all invertible `n` by `n` `R`-matrices.
15+
## Main definitions
16+
* `matrix.general_linear_group` is the type of matrices over R which are units in the matrix ring.
17+
* `matrix.GL_pos` gives the subgroup of matrices with
18+
positive determinant (over a linear ordered ring).
19+
## Tags
20+
matrix group, group, matrix inverse
21+
-/
22+
23+
namespace matrix
24+
universes u v
25+
open_locale matrix
26+
open linear_map
27+
28+
/-- `GL n R` is the group of `n` by `n` `R`-matrices with unit determinant.
29+
Defined as a subtype of matrices-/
30+
abbreviation general_linear_group (n : Type u) (R : Type v)
31+
[decidable_eq n] [fintype n] [comm_ring R] : Type* := units (matrix n n R)
32+
33+
notation `GL` := general_linear_group
34+
35+
namespace general_linear_group
36+
37+
variables {n : Type u} [decidable_eq n] [fintype n] {R : Type v} [comm_ring R]
38+
39+
/-- The determinant of a unit matrix is itself a unit. -/
40+
@[simps]
41+
def det : GL n R →* units R :=
42+
{ to_fun := λ A,
43+
{ val := (↑A : matrix n n R).det,
44+
inv := (↑(A⁻¹) : matrix n n R).det,
45+
val_inv := by rw [←det_mul, ←mul_eq_mul, A.mul_inv, det_one],
46+
inv_val := by rw [←det_mul, ←mul_eq_mul, A.inv_mul, det_one]},
47+
map_one' := units.ext det_one,
48+
map_mul' := λ A B, units.ext $ det_mul _ _ }
49+
50+
/--The `GL n R` and `general_linear_group R n` groups are multiplicatively equivalent-/
51+
def to_lin : (GL n R) ≃* (linear_map.general_linear_group R (n → R)) :=
52+
units.map_equiv to_lin_alg_equiv'.to_mul_equiv
53+
54+
/--Given a matrix with invertible determinant we get an element of `GL n R`-/
55+
def mk' (A : matrix n n R) (h : invertible (matrix.det A)) : GL n R :=
56+
unit_of_det_invertible A
57+
58+
/--Given a matrix with unit determinant we get an element of `GL n R`-/
59+
noncomputable def mk'' (A : matrix n n R) (h : is_unit (matrix.det A)) : GL n R :=
60+
nonsing_inv_unit A h
61+
62+
instance coe_fun : has_coe_to_fun (GL n R) :=
63+
{ F := λ _, n → n → R,
64+
coe := λ A, A.val }
65+
66+
lemma ext_iff (A B : GL n R) : A = B ↔ (∀ i j, (A : matrix n n R) i j = (B : matrix n n R) i j) :=
67+
units.ext_iff.trans matrix.ext_iff.symm
68+
69+
/-- Not marked `@[ext]` as the `ext` tactic already solves this. -/
70+
lemma ext ⦃A B : GL n R⦄ (h : ∀ i j, (A : matrix n n R) i j = (B : matrix n n R) i j) :
71+
A = B :=
72+
units.ext $ matrix.ext h
73+
74+
section coe_lemmas
75+
76+
variables (A B : GL n R)
77+
78+
@[simp] lemma coe_fn_eq_coe : ⇑A = (↑A : matrix n n R) := rfl
79+
80+
@[simp] lemma coe_mul : ↑(A * B) = (↑A : matrix n n R) ⬝ (↑B : matrix n n R) := rfl
81+
82+
@[simp] lemma coe_one : ↑(1 : GL n R) = (1 : matrix n n R) := rfl
83+
84+
lemma coe_inv : ↑(A⁻¹) = (↑A : matrix n n R)⁻¹ :=
85+
begin
86+
letI := A.invertible,
87+
exact inv_eq_nonsing_inv_of_invertible (↑A : matrix n n R),
88+
end
89+
90+
end coe_lemmas
91+
92+
end general_linear_group
93+
94+
namespace special_linear_group
95+
96+
variables {n : Type u} [decidable_eq n] [fintype n] {R : Type v} [comm_ring R]
97+
98+
instance has_coe_to_general_linear_group : has_coe (special_linear_group n R) (GL n R) :=
99+
⟨λ A, ⟨↑A, ↑(A⁻¹), congr_arg coe (mul_right_inv A), congr_arg coe (mul_left_inv A)⟩⟩
100+
101+
end special_linear_group
102+
103+
section
104+
105+
variables {n : Type u} {R : Type v} [decidable_eq n] [fintype n] [linear_ordered_comm_ring R ]
106+
107+
section
108+
variables (n R)
109+
110+
/-- This is the subgroup of `nxn` matrices with entries over a
111+
linear ordered ring and positive determinant. -/
112+
def GL_pos : subgroup (GL n R) :=
113+
(units.pos_subgroup R).comap general_linear_group.det
114+
end
115+
116+
@[simp] lemma mem_GL_pos (A : GL n R) : A ∈ GL_pos n R ↔ 0 < (A.det : R) := iff.rfl
117+
end
118+
119+
section has_neg
120+
121+
variables {n : Type u} {R : Type v} [decidable_eq n] [fintype n] [linear_ordered_comm_ring R ]
122+
[fact (even (fintype.card n))]
123+
124+
/-- Formal operation of negation on general linear group on even cardinality `n` given by negating
125+
each element. -/
126+
instance : has_neg (GL_pos n R) :=
127+
⟨λ g,
128+
⟨- g,
129+
begin
130+
simp only [mem_GL_pos, general_linear_group.coe_det_apply, units.coe_neg],
131+
have := det_smul g (-1),
132+
simp only [general_linear_group.coe_fn_eq_coe, one_smul, coe_fn_coe_base, neg_smul] at this,
133+
rw this,
134+
simp [nat.neg_one_pow_of_even (fact.out (even (fintype.card n)))],
135+
have gdet := g.property,
136+
simp only [mem_GL_pos, general_linear_group.coe_det_apply, subtype.val_eq_coe] at gdet,
137+
exact gdet,
138+
end⟩⟩
139+
140+
@[simp] lemma GL_pos_coe_neg (g : GL_pos n R) : ↑(- g) = - (↑g : matrix n n R) :=
141+
rfl
142+
143+
@[simp]lemma GL_pos_neg_elt (g : GL_pos n R): ∀ i j, ( ↑(-g): matrix n n R) i j= - (g i j):=
144+
begin
145+
simp,
146+
end
147+
148+
end has_neg
149+
150+
namespace special_linear_group
151+
152+
variables {n : Type u} [decidable_eq n] [fintype n] {R : Type v} [linear_ordered_comm_ring R]
153+
154+
/-- `special_linear_group n R` embeds into `GL_pos n R` -/
155+
def to_GL_pos : special_linear_group n R →* GL_pos n R :=
156+
{ to_fun := λ A, ⟨(A : GL n R), show 0 < (↑A : matrix n n R).det, from A.prop.symm ▸ zero_lt_one⟩,
157+
map_one' := subtype.ext $ units.ext $ rfl,
158+
map_mul' := λ A₁ A₂, subtype.ext $ units.ext $ rfl }
159+
160+
instance : has_coe (special_linear_group n R) (GL_pos n R) := ⟨to_GL_pos⟩
161+
162+
lemma coe_eq_to_GL_pos : (coe : special_linear_group n R → GL_pos n R) = to_GL_pos := rfl
163+
164+
lemma to_GL_pos_injective :
165+
function.injective (to_GL_pos : special_linear_group n R → GL_pos n R) :=
166+
(show function.injective ((coe : GL_pos n R → matrix n n R) ∘ to_GL_pos),
167+
from subtype.coe_injective).of_comp
168+
169+
end special_linear_group
170+
end matrix

src/ring_theory/subring.lean

Lines changed: 13 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -983,3 +983,16 @@ S.to_subsemiring.module
983983
end subring
984984

985985
end actions
986+
987+
-- while this definition is not about subrings, this is the earliest we have
988+
-- both ordered ring structures and submonoids available
989+
990+
/-- The subgroup of positive units of a linear ordered commutative ring. -/
991+
def units.pos_subgroup (R : Type*) [linear_ordered_comm_ring R] [nontrivial R] :
992+
subgroup (units R) :=
993+
{ carrier := {x | (0 : R) < x},
994+
inv_mem' := λ x (hx : (0 : R) < x), (zero_lt_mul_left hx).mp $ x.mul_inv.symm ▸ zero_lt_one,
995+
..(pos_submonoid R).comap (units.coe_hom R)}
996+
997+
@[simp] lemma units.mem_pos_subgroup {R : Type*} [linear_ordered_comm_ring R] [nontrivial R]
998+
(u : units R) : u ∈ units.pos_subgroup R ↔ (0 : R) < u := iff.rfl

src/ring_theory/subsemiring.lean

Lines changed: 13 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -831,3 +831,16 @@ instance [add_comm_monoid α] [module R' α] (S : subsemiring R') : module S α
831831
end subsemiring
832832

833833
end actions
834+
835+
-- While this definition is not about `subsemiring`s, this is the earliest we have
836+
-- both `ordered_semiring` and `submonoid` available.
837+
838+
839+
/-- Submonoid of positive elements of an ordered semiring. -/
840+
def pos_submonoid (R : Type*) [ordered_semiring R] [nontrivial R] : submonoid R :=
841+
{ carrier := {x | 0 < x},
842+
one_mem' := show (0 : R) < 1, from zero_lt_one,
843+
mul_mem' := λ x y (hx : 0 < x) (hy : 0 < y), mul_pos hx hy }
844+
845+
@[simp] lemma mem_pos_monoid {R : Type*} [ordered_semiring R] [nontrivial R] (u : units R) :
846+
↑u ∈ pos_submonoid R ↔ (0 : R) < u := iff.rfl

0 commit comments

Comments
 (0)