Cusps #
We define the cusps of a subgroup of GL(2, ℝ)
as the fixed points of parabolic elements.
theorem
OnePoint.exists_mem_SL2
{K : Type u_1}
[Field K]
[DecidableEq K]
(A : Type u_2)
[CommRing A]
[IsDomain A]
[Algebra A K]
[IsFractionRing A K]
[IsPrincipalIdealRing A]
(c : OnePoint K)
:
∃ (g : Matrix.SpecialLinearGroup (Fin 2) A), (Matrix.SpecialLinearGroup.mapGL K) g • infty = c
The modular group SL(2, A)
acts transitively on OnePoint K
, if A
is a PID whose fraction
field is K
. (This includes the case A = ℤ
, K = ℚ
.)
The cusps of SL(2, ℤ)
are precisely the elements of ℙ¹(ℚ)
.
theorem
isCusp_SL2Z_iff'
{c : OnePoint ℝ}
:
IsCusp c (Matrix.SpecialLinearGroup.mapGL ℝ).range ↔ ∃ (g : Matrix.SpecialLinearGroup (Fin 2) ℤ), c = (Matrix.SpecialLinearGroup.mapGL ℝ) g • OnePoint.infty
The cusps of SL(2, ℤ)
are precisely the SL(2, ℤ)
orbit of ∞
.
theorem
Subgroup.IsArithmetic.isCusp_iff_isCusp_SL2Z
(𝒢 : Subgroup (GL (Fin 2) ℝ))
[𝒢.IsArithmetic]
{c : OnePoint ℝ}
:
The cusps of any arithmetic subgroup are the same as those of SL(2, ℤ)
.
@[reducible, inline]
The type of cusp orbits of 𝒢
, i.e. orbits for the action of 𝒢
on its own cusps.
Equations
Instances For
Surjection from SL(2, ℤ) / (𝒢 ⊓ SL(2, ℤ))
to cusp orbits of 𝒢
. Mostly useful for showing
that CuspOrbits 𝒢
is finite for arithmetic subgroups.
Equations
- cosetToCuspOrbit 𝒢 = Quotient.lift (fun (g : Matrix.SpecialLinearGroup (Fin 2) ℤ) => ⟦⟨(Matrix.SpecialLinearGroup.mapGL ℝ) g⁻¹ • OnePoint.infty, ⋯⟩⟧) ⋯
Instances For
@[simp]
theorem
cosetToCuspOrbit_apply_mk
{𝒢 : Subgroup (GL (Fin 2) ℝ)}
[𝒢.IsArithmetic]
(g : Matrix.SpecialLinearGroup (Fin 2) ℤ)
:
instance
instFiniteCuspOrbitsOfIsArithmetic
(𝒢 : Subgroup (GL (Fin 2) ℝ))
[𝒢.IsArithmetic]
:
Finite (CuspOrbits 𝒢)
An arithmetic subgroup has finitely many cusp orbits.