Documentation

Mathlib.NumberTheory.ModularForms.Cusps

Cusps #

We define the cusps of a subgroup of GL(2, ℝ) as the fixed points of parabolic elements.

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 = ℚ.)

def IsCusp (c : OnePoint ) (𝒢 : Subgroup (GL (Fin 2) )) :

The cusps of a subgroup of GL(2, ℝ) are the fixed points of parabolic elements of g.

Equations
Instances For
    theorem IsCusp.smul {c : OnePoint } {𝒢 : Subgroup (GL (Fin 2) )} (hc : IsCusp c 𝒢) (g : GL (Fin 2) ) :
    theorem IsCusp.smul_of_mem {c : OnePoint } {𝒢 : Subgroup (GL (Fin 2) )} (hc : IsCusp c 𝒢) {g : GL (Fin 2) } (hg : g 𝒢) :
    IsCusp (g c) 𝒢
    theorem isCusp_iff_of_relIndex_ne_zero {𝒢 𝒢' : Subgroup (GL (Fin 2) )} (h𝒢 : 𝒢' 𝒢) (h𝒢' : 𝒢'.relIndex 𝒢 0) (c : OnePoint ) :
    IsCusp c 𝒢' IsCusp c 𝒢
    @[deprecated isCusp_iff_of_relIndex_ne_zero (since := "2025-09-13")]
    theorem isCusp_iff_of_relindex_ne_zero {𝒢 𝒢' : Subgroup (GL (Fin 2) )} (h𝒢 : 𝒢' 𝒢) (h𝒢' : 𝒢'.relIndex 𝒢 0) (c : OnePoint ) :
    IsCusp c 𝒢' IsCusp c 𝒢

    Alias of isCusp_iff_of_relIndex_ne_zero.

    theorem Commensurable.isCusp_iff {𝒢 𝒢' : Subgroup (GL (Fin 2) )} (h𝒢 : Commensurable 𝒢 𝒢') {c : OnePoint } :
    IsCusp c 𝒢 IsCusp c 𝒢'

    The cusps of SL(2, ℤ) are precisely the elements of ℙ¹(ℚ).

    The cusps of SL(2, ℤ) are precisely the SL(2, ℤ) orbit of .

    The cusps of any arithmetic subgroup are the same as those of SL(2, ℤ).

    def cusps_subMulAction (𝒢 : Subgroup (GL (Fin 2) )) :

    The action of 𝒢 on its own cusps.

    Equations
    Instances For
      @[reducible, inline]
      abbrev CuspOrbits (𝒢 : Subgroup (GL (Fin 2) )) :

      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
        Instances For

          An arithmetic subgroup has finitely many cusp orbits.