Skip to content

Commit e29e9b5

Browse files
CBirkbeckloefflerd
andcommitted
feat: add definition of Eisenstein series (#9783)
This defines Eisenstein series of level N and shows that they are a slash invariant form. Also adds some pre-requisites for proving they are modular forms, which will be in the next PRs Co-authored-by: David Loeffler <d.loeffler.01@cantab.net> Co-authored-by: Chris Birkbeck <c.birkbeck@uea.ac.uk>
1 parent 5160f99 commit e29e9b5

File tree

3 files changed

+192
-0
lines changed

3 files changed

+192
-0
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -2815,6 +2815,7 @@ import Mathlib.NumberTheory.MaricaSchoenheim
28152815
import Mathlib.NumberTheory.Modular
28162816
import Mathlib.NumberTheory.ModularForms.Basic
28172817
import Mathlib.NumberTheory.ModularForms.CongruenceSubgroups
2818+
import Mathlib.NumberTheory.ModularForms.EisensteinSeries.Basic
28182819
import Mathlib.NumberTheory.ModularForms.JacobiTheta.Manifold
28192820
import Mathlib.NumberTheory.ModularForms.JacobiTheta.OneVariable
28202821
import Mathlib.NumberTheory.ModularForms.JacobiTheta.TwoVariable

Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean

Lines changed: 69 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -124,6 +124,13 @@ instance : Pow (SpecialLinearGroup n R) ℕ where
124124
instance : Inhabited (SpecialLinearGroup n R) :=
125125
1
126126

127+
/-- The transpose of a matrix in `SL(n, R)` -/
128+
def transpose (A : SpecialLinearGroup n R) : SpecialLinearGroup n R :=
129+
⟨A.1.transpose, A.1.det_transpose ▸ A.2
130+
131+
@[inherit_doc]
132+
scoped postfix:1024 "ᵀ" => SpecialLinearGroup.transpose
133+
127134
section CoeLemmas
128135

129136
variable (A B : SpecialLinearGroup n R)
@@ -158,6 +165,10 @@ theorem coe_pow (m : ℕ) : ↑ₘ(A ^ m) = ↑ₘA ^ m :=
158165
rfl
159166
#align matrix.special_linear_group.coe_pow Matrix.SpecialLinearGroup.coe_pow
160167

168+
@[simp]
169+
lemma coe_transpose (A : SpecialLinearGroup n R) : ↑ₘAᵀ = (↑ₘA)ᵀ :=
170+
rfl
171+
161172
theorem det_ne_zero [Nontrivial R] (g : SpecialLinearGroup n R) : det ↑ₘg ≠ 0 := by
162173
rw [g.det_coe]
163174
norm_num
@@ -319,12 +330,70 @@ theorem fin_two_exists_eq_mk_of_apply_zero_one_eq_zero {R : Type*} [Field R] (g
319330
simp_rw [eq_inv_of_mul_eq_one_right had, hg]
320331
#align matrix.special_linear_group.fin_two_exists_eq_mk_of_apply_zero_one_eq_zero Matrix.SpecialLinearGroup.fin_two_exists_eq_mk_of_apply_zero_one_eq_zero
321332

333+
lemma isCoprime_row (A : SL(2, R)) (i : Fin 2): IsCoprime (A i 0) (A i 1) := by
334+
refine match i with
335+
| 0 => ⟨A 1 1, -(A 1 0), ?_⟩
336+
| 1 => ⟨-(A 0 1), A 0 0, ?_⟩ <;>
337+
· simp_rw [det_coe A ▸ det_fin_two A.1]
338+
ring
339+
340+
lemma isCoprime_col (A : SL(2, R)) (j : Fin 2): IsCoprime (A 0 j) (A 1 j) := by
341+
refine match j with
342+
| 0 => ⟨A 1 1, -(A 0 1), ?_⟩
343+
| 1 => ⟨-(A 1 0), A 0 0, ?_⟩ <;>
344+
· simp_rw [det_coe A ▸ det_fin_two A.1]
345+
ring
346+
322347
end SpecialCases
323348

324349
end SpecialLinearGroup
325350

326351
end Matrix
327352

353+
namespace IsCoprime
354+
355+
open Matrix MatrixGroups SpecialLinearGroup
356+
357+
variable {R : Type*} [CommRing R]
358+
359+
/-- Given any pair of coprime elements of `R`, there exists a matrix in `SL(2, R)` having those
360+
entries as its left or right column. -/
361+
lemma exists_SL2_col {a b : R} (hab : IsCoprime a b) (j : Fin 2):
362+
∃ g : SL(2, R), g 0 j = a ∧ g 1 j = b := by
363+
obtain ⟨u, v, h⟩ := hab
364+
refine match j with
365+
| 0 => ⟨⟨!![a, -v; b, u], ?_⟩, rfl, rfl⟩
366+
| 1 => ⟨⟨!![v, a; -u, b], ?_⟩, rfl, rfl⟩ <;>
367+
· rw [Matrix.det_fin_two_of, ← h]
368+
ring
369+
370+
/-- Given any pair of coprime elements of `R`, there exists a matrix in `SL(2, R)` having those
371+
entries as its top or bottom row. -/
372+
lemma exists_SL2_row {a b : R} (hab : IsCoprime a b) (i : Fin 2):
373+
∃ g : SL(2, R), g i 0 = a ∧ g i 1 = b := by
374+
obtain ⟨u, v, h⟩ := hab
375+
refine match i with
376+
| 0 => ⟨⟨!![a, b; -v, u], ?_⟩, rfl, rfl⟩
377+
| 1 => ⟨⟨!![v, -u; a, b], ?_⟩, rfl, rfl⟩ <;>
378+
· rw [Matrix.det_fin_two_of, ← h]
379+
ring
380+
381+
/-- A vector with coprime entries, right-multiplied by a matrix in `SL(2, R)`, has
382+
coprime entries. -/
383+
lemma vecMulSL {v : Fin 2 → R} (hab : IsCoprime (v 0) (v 1)) (A : SL(2, R)) :
384+
IsCoprime (vecMul v A.1 0) (vecMul v A.1 1) := by
385+
obtain ⟨g, hg⟩ := hab.exists_SL2_row 0
386+
have : v = g 0 := funext fun t ↦ by { fin_cases t <;> tauto }
387+
simpa only [this] using isCoprime_row (g * A) 0
388+
389+
/-- A vector with coprime entries, left-multiplied by a matrix in `SL(2, R)`, has
390+
coprime entries. -/
391+
lemma mulVecSL {v : Fin 2 → R} (hab : IsCoprime (v 0) (v 1)) (A : SL(2, R)) :
392+
IsCoprime (mulVec A.1 v 0) (mulVec A.1 v 1) := by
393+
simpa only [← vecMul_transpose] using hab.vecMulSL A.transpose
394+
395+
end IsCoprime
396+
328397
namespace ModularGroup
329398

330399
open MatrixGroups
Lines changed: 122 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,122 @@
1+
/-
2+
Copyright (c) 2024 Chris Birkbeck. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Chris Birkbeck, David Loeffler
5+
-/
6+
import Mathlib.NumberTheory.ModularForms.SlashInvariantForms
7+
import Mathlib.NumberTheory.ModularForms.CongruenceSubgroups
8+
9+
/-!
10+
# Eisenstein Series
11+
12+
## Main definitions
13+
14+
* We define Eisenstein series of level `Γ(N)` for any `N : ℕ` and weight `k : ℤ` as the infinite sum
15+
`∑' v : (Fin 2 → ℤ), (1 / (v 0 * z + v 1) ^ k)`, where `z : ℍ` and `v` ranges over all pairs of
16+
coprime integers congruent to a fixed pair `(a, b)` modulo `N`. Note that by using `(Fin 2 → ℤ)`
17+
instead of `ℤ × ℤ` we can state all of the required equivalences using matrices and vectors, which
18+
makes working with them more convenient.
19+
20+
* We show that they define a slash invariant form of level `Γ(N)` and weight `k`.
21+
22+
## References
23+
* [F. Diamond and J. Shurman, *A First Course in Modular Forms*][diamondshurman2005]
24+
-/
25+
26+
noncomputable section
27+
28+
open ModularForm UpperHalfPlane Complex Matrix
29+
30+
open scoped MatrixGroups
31+
32+
namespace EisensteinSeries
33+
34+
variable (N : ℕ) (a : Fin 2 → ZMod N)
35+
36+
section gammaSet_def
37+
38+
/-- The set of pairs of coprime integers congruent to `a` mod `N`. -/
39+
def gammaSet := {v : Fin 2 → ℤ | (↑) ∘ v = a ∧ IsCoprime (v 0) (v 1)}
40+
41+
lemma pairwise_disjoint_gammaSet : Pairwise (Disjoint on gammaSet N) := by
42+
refine fun u v huv ↦ ?_
43+
contrapose! huv
44+
obtain ⟨f, hf⟩ := Set.not_disjoint_iff.mp huv
45+
exact hf.1.1.symm.trans hf.2.1
46+
47+
/-- For level `N = 1`, the gamma sets are all equal. -/
48+
lemma gammaSet_one_eq (a a' : Fin 2 → ZMod 1) : gammaSet 1 a = gammaSet 1 a' :=
49+
congr_arg _ (Subsingleton.elim _ _)
50+
51+
/-- For level `N = 1`, the gamma sets are all equivalent; this is the equivalence. -/
52+
def gammaSet_one_equiv (a a' : Fin 2 → ZMod 1) : gammaSet 1 a ≃ gammaSet 1 a' :=
53+
Equiv.Set.ofEq (gammaSet_one_eq a a')
54+
55+
end gammaSet_def
56+
57+
variable {N a}
58+
59+
section gamma_action
60+
61+
/-- Right-multiplying by `γ ∈ SL(2, ℤ)` sends `gammaSet N a` to `gammaSet N (vecMul a γ)`. -/
62+
lemma vecMul_SL2_mem_gammaSet {v : Fin 2 → ℤ} (hv : v ∈ gammaSet N a) (γ : SL(2, ℤ)) :
63+
vecMul v γ ∈ gammaSet N (vecMul a γ) := by
64+
refine ⟨?_, hv.2.vecMulSL γ⟩
65+
have := RingHom.map_vecMul (m := Fin 2) (n := Fin 2) (Int.castRingHom (ZMod N)) γ v
66+
simp only [eq_intCast, Int.coe_castRingHom] at this
67+
simp_rw [Function.comp, this, hv.1]
68+
simp
69+
70+
variable (a) in
71+
/-- The bijection between `GammaSets` given by multiplying by an element of `SL(2, ℤ)`. -/
72+
def gammaSetEquiv (γ : SL(2, ℤ)) : gammaSet N a ≃ gammaSet N (vecMul a γ) where
73+
toFun v := ⟨vecMul v.1 γ, vecMul_SL2_mem_gammaSet v.2 γ⟩
74+
invFun v := ⟨vecMul v.1 ↑(γ⁻¹), by
75+
have := vecMul_SL2_mem_gammaSet v.2 γ⁻¹
76+
rw [vecMul_vecMul, ← SpecialLinearGroup.coe_mul] at this
77+
simpa only [SpecialLinearGroup.map_apply_coe, RingHom.mapMatrix_apply, Int.coe_castRingHom,
78+
map_inv, mul_right_inv, SpecialLinearGroup.coe_one, vecMul_one]⟩
79+
left_inv v := by simp_rw [vecMul_vecMul, ← SpecialLinearGroup.coe_mul, mul_inv_self,
80+
SpecialLinearGroup.coe_one, vecMul_one]
81+
right_inv v := by simp_rw [vecMul_vecMul, ← SpecialLinearGroup.coe_mul, inv_mul_self,
82+
SpecialLinearGroup.coe_one, vecMul_one]
83+
84+
end gamma_action
85+
86+
section eisSummand
87+
88+
/-- The function on `(Fin 2 → ℤ)` whose sum defines an Eisenstein series.-/
89+
def eisSummand (k : ℤ) (v : Fin 2 → ℤ) (z : ℍ) : ℂ := 1 / (v 0 * z.1 + v 1) ^ k
90+
91+
/-- How the `eisSummand` function changes under the Moebius action. -/
92+
theorem eisSummand_SL2_apply (k : ℤ) (i : (Fin 2 → ℤ)) (A : SL(2, ℤ)) (z : ℍ) :
93+
eisSummand k i (A • z) = (z.denom A) ^ k * eisSummand k (vecMul i A) z := by
94+
simp only [eisSummand, specialLinearGroup_apply, algebraMap_int_eq, eq_intCast, ofReal_int_cast,
95+
one_div, vecMul, vec2_dotProduct, Int.cast_add, Int.cast_mul]
96+
have h (a b c d u v : ℂ) (hc : c * z + d ≠ 0) : ((u * ((a * z + b) / (c * z + d)) + v) ^ k)⁻¹ =
97+
(c * z + d) ^ k * (((u * a + v * c) * z + (u * b + v * d)) ^ k)⁻¹
98+
· field_simp [hc]
99+
ring_nf
100+
apply h (hc := z.denom_ne_zero A)
101+
102+
end eisSummand
103+
104+
variable (a)
105+
106+
/-- An Eisenstein series of weight `k` and level `Γ(N)`, with congruence condition `a`. -/
107+
def eisensteinSeries (k : ℤ) (z : ℍ) : ℂ := ∑' x : gammaSet N a, eisSummand k x z
108+
109+
lemma eisensteinSeries_slash_apply (k : ℤ) (γ : SL(2, ℤ)) :
110+
eisensteinSeries a k ∣[k] γ = eisensteinSeries (vecMul a γ) k := by
111+
ext1 z
112+
simp_rw [SL_slash, slash_def, slash, det_coe', ofReal_one, one_zpow, mul_one, zpow_neg,
113+
mul_inv_eq_iff_eq_mul₀ (zpow_ne_zero _ <| z.denom_ne_zero _), mul_comm,
114+
eisensteinSeries, ← UpperHalfPlane.sl_moeb, eisSummand_SL2_apply, tsum_mul_left]
115+
erw [(gammaSetEquiv a γ).tsum_eq (eisSummand k · z)]
116+
117+
/-- The SlashInvariantForm defined by an Eisenstein series of weight `k : ℤ`, level `Γ(N)`,
118+
and congruence condition given by `a : Fin 2 → ZMod N`. -/
119+
def eisensteinSeries_SIF (k : ℤ) : SlashInvariantForm (Gamma N) k where
120+
toFun := eisensteinSeries a k
121+
slash_action_eq' A := by rw [subgroup_slash, ← SL_slash, eisensteinSeries_slash_apply,
122+
(Gamma_mem' N A).mp A.2, SpecialLinearGroup.coe_one, vecMul_one]

0 commit comments

Comments
 (0)