Documentation

Mathlib.NumberTheory.ModularForms.Basic

Modular forms #

This file defines modular forms and proves some basic properties about them. Including constructing the graded ring of modular forms.

We begin by defining modular forms and cusp forms as extension of SlashInvariantForms then we define the space of modular forms, cusp forms and prove that the product of two modular forms is a modular form.

The matrix [-1, 0; 0, 1], which defines an anti-holomorphic involution of via τ ↦ -conj τ.

Equations
Instances For
    @[simp]
    theorem UpperHalfPlane.val_J :
    J = !![-1, 0; 0, 1]
    @[simp]
    theorem UpperHalfPlane.J_sq :
    J ^ 2 = 1
    @[simp]

    The weight k slash action of GL(2, ℝ) preserves holomorphic functions.

    structure ModularForm (Γ : Subgroup (GL (Fin 2) )) (k : ) extends SlashInvariantForm Γ k :

    These are SlashInvariantForm's that are holomorphic and bounded at infinity.

    Instances For
      structure CuspForm (Γ : Subgroup (GL (Fin 2) )) (k : ) extends SlashInvariantForm Γ k :

      These are SlashInvariantForms that are holomorphic and zero at infinity.

      Instances For

        ModularFormClass F Γ k says that F is a type of bundled functions that extend SlashInvariantFormClass by requiring that the functions be holomorphic and bounded at all cusps.

        Instances
          class CuspFormClass (F : Type u_2) (Γ : outParam (Subgroup (GL (Fin 2) ))) (k : outParam ) [FunLike F UpperHalfPlane ] extends SlashInvariantFormClass F Γ k :

          CuspFormClass F Γ k says that F is a type of bundled functions that extend SlashInvariantFormClass by requiring that the functions be holomorphic and zero at all cusps.

          Instances
            @[instance 100]
            Equations
            @[instance 100]
            @[instance 100]
            instance CuspForm.funLike (Γ : Subgroup (GL (Fin 2) )) (k : ) :
            Equations
            @[instance 100]
            instance CuspFormClass.cuspForm (Γ : Subgroup (GL (Fin 2) )) (k : ) :
            theorem ModularForm.toFun_eq_coe {Γ : Subgroup (GL (Fin 2) )} {k : } (f : ModularForm Γ k) :
            f.toFun = f
            @[simp]
            theorem CuspForm.toFun_eq_coe {Γ : Subgroup (GL (Fin 2) )} {k : } {f : CuspForm Γ k} :
            f.toFun = f
            @[simp]
            theorem CuspForm.toSlashInvariantForm_coe {Γ : Subgroup (GL (Fin 2) )} {k : } (f : CuspForm Γ k) :
            theorem ModularForm.ext {Γ : Subgroup (GL (Fin 2) )} {k : } {f g : ModularForm Γ k} (h : ∀ (x : UpperHalfPlane), f x = g x) :
            f = g
            theorem ModularForm.ext_iff {Γ : Subgroup (GL (Fin 2) )} {k : } {f g : ModularForm Γ k} :
            f = g ∀ (x : UpperHalfPlane), f x = g x
            theorem CuspForm.ext {Γ : Subgroup (GL (Fin 2) )} {k : } {f g : CuspForm Γ k} (h : ∀ (x : UpperHalfPlane), f x = g x) :
            f = g
            theorem CuspForm.ext_iff {Γ : Subgroup (GL (Fin 2) )} {k : } {f g : CuspForm Γ k} :
            f = g ∀ (x : UpperHalfPlane), f x = g x
            def ModularForm.copy {Γ : Subgroup (GL (Fin 2) )} {k : } (f : ModularForm Γ k) (f' : UpperHalfPlane) (h : f' = f) :

            Copy of a ModularForm with a new toFun equal to the old one. Useful to fix definitional equalities.

            Equations
            • f.copy f' h = { toSlashInvariantForm := f.copy f' h, holo' := , bdd_at_cusps' := }
            Instances For
              def CuspForm.copy {Γ : Subgroup (GL (Fin 2) )} {k : } (f : CuspForm Γ k) (f' : UpperHalfPlane) (h : f' = f) :

              Copy of a CuspForm with a new toFun equal to the old one. Useful to fix definitional equalities.

              Equations
              • f.copy f' h = { toSlashInvariantForm := f.copy f' h, holo' := , zero_at_cusps' := }
              Instances For
                instance ModularForm.add {Γ : Subgroup (GL (Fin 2) )} {k : } :
                Equations
                • One or more equations did not get rendered due to their size.
                @[simp]
                theorem ModularForm.coe_add {Γ : Subgroup (GL (Fin 2) )} {k : } (f g : ModularForm Γ k) :
                ⇑(f + g) = f + g
                @[simp]
                theorem ModularForm.add_apply {Γ : Subgroup (GL (Fin 2) )} {k : } (f g : ModularForm Γ k) (z : UpperHalfPlane) :
                (f + g) z = f z + g z
                instance ModularForm.instZero {Γ : Subgroup (GL (Fin 2) )} {k : } :
                Equations
                @[simp]
                theorem ModularForm.coe_zero {Γ : Subgroup (GL (Fin 2) )} {k : } :
                0 = 0
                @[simp]
                theorem ModularForm.zero_apply {Γ : Subgroup (GL (Fin 2) )} {k : } (z : UpperHalfPlane) :
                0 z = 0
                instance ModularForm.instSMulℝ {Γ : Subgroup (GL (Fin 2) )} {k : } {α : Type u_1} [SMul α ] [SMul α ] [IsScalarTower α ] :
                SMul α (ModularForm Γ k)
                Equations
                • One or more equations did not get rendered due to their size.
                @[simp]
                theorem ModularForm.coe_smul {Γ : Subgroup (GL (Fin 2) )} {k : } {α : Type u_1} [SMul α ] [SMul α ] [IsScalarTower α ] (f : ModularForm Γ k) (n : α) :
                ⇑(n f) = n f
                @[simp]
                theorem ModularForm.smul_apply {Γ : Subgroup (GL (Fin 2) )} {k : } {α : Type u_1} [SMul α ] [SMul α ] [IsScalarTower α ] (f : ModularForm Γ k) (n : α) (z : UpperHalfPlane) :
                (n f) z = n f z
                instance ModularForm.instSMulℂ {Γ : Subgroup (GL (Fin 2) )} {k : } {α : Type u_1} [SMul α ] [IsScalarTower α ] [Γ.HasDetOne] :
                SMul α (ModularForm Γ k)
                Equations
                • One or more equations did not get rendered due to their size.
                @[simp]
                theorem ModularForm.IsGLPos.coe_smul {Γ : Subgroup (GL (Fin 2) )} {k : } {α : Type u_1} [SMul α ] [IsScalarTower α ] [Γ.HasDetOne] (f : ModularForm Γ k) (n : α) :
                ⇑(n f) = n f
                @[simp]
                theorem ModularForm.IsGLPos.smul_apply {Γ : Subgroup (GL (Fin 2) )} {k : } {α : Type u_1} [SMul α ] [IsScalarTower α ] [Γ.HasDetOne] (f : ModularForm Γ k) (n : α) (z : UpperHalfPlane) :
                (n f) z = n f z
                instance ModularForm.instNeg {Γ : Subgroup (GL (Fin 2) )} {k : } :
                Equations
                • One or more equations did not get rendered due to their size.
                @[simp]
                theorem ModularForm.coe_neg {Γ : Subgroup (GL (Fin 2) )} {k : } (f : ModularForm Γ k) :
                ⇑(-f) = -f
                @[simp]
                theorem ModularForm.neg_apply {Γ : Subgroup (GL (Fin 2) )} {k : } (f : ModularForm Γ k) (z : UpperHalfPlane) :
                (-f) z = -f z
                instance ModularForm.instSub {Γ : Subgroup (GL (Fin 2) )} {k : } :
                Equations
                @[simp]
                theorem ModularForm.coe_sub {Γ : Subgroup (GL (Fin 2) )} {k : } (f g : ModularForm Γ k) :
                ⇑(f - g) = f - g
                @[simp]
                theorem ModularForm.sub_apply {Γ : Subgroup (GL (Fin 2) )} {k : } (f g : ModularForm Γ k) (z : UpperHalfPlane) :
                (f - g) z = f z - g z
                Equations

                Additive coercion from ModularForm to ℍ → ℂ.

                Equations
                Instances For
                  @[simp]
                  theorem ModularForm.coeHom_apply {Γ : Subgroup (GL (Fin 2) )} {k : } (f : ModularForm Γ k) (a : UpperHalfPlane) :
                  coeHom f a = f a
                  instance ModularForm.instInhabited {Γ : Subgroup (GL (Fin 2) )} {k : } :
                  Equations
                  def ModularForm.mul {Γ : Subgroup (GL (Fin 2) )} {k_1 k_2 : } [Γ.HasDetPlusMinusOne] (f : ModularForm Γ k_1) (g : ModularForm Γ k_2) :
                  ModularForm Γ (k_1 + k_2)

                  The modular form of weight k_1 + k_2 given by the product of two modular forms of weights k_1 and k_2.

                  Equations
                  Instances For
                    @[simp]
                    theorem ModularForm.mul_coe {Γ : Subgroup (GL (Fin 2) )} [Γ.HasDetPlusMinusOne] {k_1 k_2 : } (f : ModularForm Γ k_1) (g : ModularForm Γ k_2) :
                    (f.mul g) = f * g
                    def ModularForm.const {Γ : Subgroup (GL (Fin 2) )} (x : ) [Γ.HasDetOne] :

                    The constant function with value x : ℂ as a modular form of weight 0 and any level.

                    Equations
                    Instances For
                      @[simp]
                      theorem ModularForm.const_apply {Γ : Subgroup (GL (Fin 2) )} [Γ.HasDetOne] (x : ) (τ : UpperHalfPlane) :
                      (const x) τ = x

                      The constant function with value x : ℂ as a modular form of weight 0 and any level.

                      Equations
                      Instances For
                        @[simp]
                        theorem ModularForm.constℝ_apply {Γ : Subgroup (GL (Fin 2) )} [Γ.HasDetPlusMinusOne] (x : ) (τ : UpperHalfPlane) :
                        (constℝ x) τ = x
                        Equations
                        • One or more equations did not get rendered due to their size.
                        @[simp]
                        @[simp]
                        theorem ModularForm.coe_natCast {Γ : Subgroup (GL (Fin 2) )} [Γ.HasDetPlusMinusOne] (n : ) :
                        n = n
                        @[simp]
                        theorem ModularForm.coe_intCast {Γ : Subgroup (GL (Fin 2) )} [Γ.HasDetPlusMinusOne] (z : ) :
                        z = z
                        instance CuspForm.hasAdd {Γ : Subgroup (GL (Fin 2) )} {k : } :
                        Add (CuspForm Γ k)
                        Equations
                        • One or more equations did not get rendered due to their size.
                        @[simp]
                        theorem CuspForm.coe_add {Γ : Subgroup (GL (Fin 2) )} {k : } (f g : CuspForm Γ k) :
                        ⇑(f + g) = f + g
                        @[simp]
                        theorem CuspForm.add_apply {Γ : Subgroup (GL (Fin 2) )} {k : } (f g : CuspForm Γ k) (z : UpperHalfPlane) :
                        (f + g) z = f z + g z
                        instance CuspForm.instZero {Γ : Subgroup (GL (Fin 2) )} {k : } :
                        Zero (CuspForm Γ k)
                        Equations
                        @[simp]
                        theorem CuspForm.coe_zero {Γ : Subgroup (GL (Fin 2) )} {k : } :
                        0 = 0
                        @[simp]
                        theorem CuspForm.zero_apply {Γ : Subgroup (GL (Fin 2) )} {k : } (z : UpperHalfPlane) :
                        0 z = 0
                        instance CuspForm.instSMul {Γ : Subgroup (GL (Fin 2) )} {k : } {α : Type u_2} [SMul α ] [SMul α ] [IsScalarTower α ] :
                        SMul α (CuspForm Γ k)
                        Equations
                        • One or more equations did not get rendered due to their size.
                        @[simp]
                        theorem CuspForm.coe_smul {Γ : Subgroup (GL (Fin 2) )} {k : } {α : Type u_2} [SMul α ] [SMul α ] [IsScalarTower α ] (f : CuspForm Γ k) (n : α) :
                        ⇑(n f) = n f
                        @[simp]
                        theorem CuspForm.smul_apply {Γ : Subgroup (GL (Fin 2) )} {k : } {α : Type u_2} [SMul α ] [SMul α ] [IsScalarTower α ] (f : CuspForm Γ k) (n : α) {z : UpperHalfPlane} :
                        (n f) z = n f z
                        instance CuspForm.IsGLPos.instSMul {Γ : Subgroup (GL (Fin 2) )} {k : } {α : Type u_2} [SMul α ] [IsScalarTower α ] [Γ.HasDetOne] :
                        SMul α (CuspForm Γ k)
                        Equations
                        • One or more equations did not get rendered due to their size.
                        @[simp]
                        theorem CuspForm.IsGLPos.coe_smul {Γ : Subgroup (GL (Fin 2) )} {k : } {α : Type u_2} [SMul α ] [IsScalarTower α ] [Γ.HasDetOne] (f : CuspForm Γ k) (n : α) :
                        ⇑(n f) = n f
                        @[simp]
                        theorem CuspForm.IsGLPos.smul_apply {Γ : Subgroup (GL (Fin 2) )} {k : } {α : Type u_2} [SMul α ] [IsScalarTower α ] [Γ.HasDetOne] (f : CuspForm Γ k) (n : α) {z : UpperHalfPlane} :
                        (n f) z = n f z
                        instance CuspForm.instNeg {Γ : Subgroup (GL (Fin 2) )} {k : } :
                        Neg (CuspForm Γ k)
                        Equations
                        • One or more equations did not get rendered due to their size.
                        @[simp]
                        theorem CuspForm.coe_neg {Γ : Subgroup (GL (Fin 2) )} {k : } (f : CuspForm Γ k) :
                        ⇑(-f) = -f
                        @[simp]
                        theorem CuspForm.neg_apply {Γ : Subgroup (GL (Fin 2) )} {k : } (f : CuspForm Γ k) (z : UpperHalfPlane) :
                        (-f) z = -f z
                        instance CuspForm.instSub {Γ : Subgroup (GL (Fin 2) )} {k : } :
                        Sub (CuspForm Γ k)
                        Equations
                        @[simp]
                        theorem CuspForm.coe_sub {Γ : Subgroup (GL (Fin 2) )} {k : } (f g : CuspForm Γ k) :
                        ⇑(f - g) = f - g
                        @[simp]
                        theorem CuspForm.sub_apply {Γ : Subgroup (GL (Fin 2) )} {k : } (f g : CuspForm Γ k) (z : UpperHalfPlane) :
                        (f - g) z = f z - g z
                        instance CuspForm.instAddCommGroup {Γ : Subgroup (GL (Fin 2) )} {k : } :
                        Equations
                        def CuspForm.coeHom {Γ : Subgroup (GL (Fin 2) )} {k : } :

                        Additive coercion from CuspForm to ℍ → ℂ.

                        Equations
                        Instances For
                          @[simp]
                          theorem CuspForm.coeHom_apply {Γ : Subgroup (GL (Fin 2) )} {k : } (f : CuspForm Γ k) (a : UpperHalfPlane) :
                          coeHom f a = f a
                          instance CuspForm.instInhabited {Γ : Subgroup (GL (Fin 2) )} {k : } :
                          Equations
                          @[instance 99]
                          def ModularForm.mcast {a b : } {Γ : Subgroup (GL (Fin 2) )} (h : a = b) (f : ModularForm Γ a) :

                          Cast for modular forms, which is useful for avoiding Heqs.

                          Equations
                          • ModularForm.mcast h f = { toFun := f, slash_action_eq' := , holo' := , bdd_at_cusps' := }
                          Instances For
                            theorem ModularForm.gradedMonoid_eq_of_cast {Γ : Subgroup (GL (Fin 2) )} {a b : GradedMonoid (ModularForm Γ)} (h : a.fst = b.fst) (h2 : mcast h a.snd = b.snd) :
                            a = b
                            Equations
                            • One or more equations did not get rendered due to their size.
                            Equations
                            noncomputable def ModularForm.translate {k : } {Γ : Subgroup (GL (Fin 2) )} {F : Type u_1} [FunLike F UpperHalfPlane ] (f : F) [ModularFormClass F Γ k] (g : GL (Fin 2) ) :

                            Translating a ModularForm by GL(2, ℝ), to obtain a new ModularForm.

                            Equations
                            Instances For
                              @[simp]
                              theorem ModularForm.coe_translate {k : } {Γ : Subgroup (GL (Fin 2) )} {F : Type u_1} [FunLike F UpperHalfPlane ] (f : F) [ModularFormClass F Γ k] (g : GL (Fin 2) ) :
                              (translate f g) = SlashAction.map k g f
                              noncomputable def CuspForm.translate {k : } {Γ : Subgroup (GL (Fin 2) )} {F : Type u_1} [FunLike F UpperHalfPlane ] (f : F) [CuspFormClass F Γ k] (g : GL (Fin 2) ) :

                              Translating a CuspForm by SL(2, ℤ), to obtain a new CuspForm.

                              Equations
                              Instances For