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