@@ -28,6 +28,7 @@ open DedekindDomain
28
28
29
29
abbrev Dfx := (D ⊗[F] (FiniteAdeleRing (𝓞 F) F))ˣ
30
30
31
+ variable {F D} in
31
32
noncomputable abbrev incl₁ : Dˣ →* Dfx F D :=
32
33
Units.map Algebra.TensorProduct.includeLeftRingHom.toMonoidHom
33
34
@@ -51,7 +52,7 @@ structure AutomorphicForm
51
52
-- definition
52
53
toFun : (D ⊗[F] (FiniteAdeleRing (𝓞 F) F))ˣ → W
53
54
left_invt : ∀ (δ : Dˣ) (g : (D ⊗[F] (FiniteAdeleRing (𝓞 F) F))ˣ),
54
- toFun (Units.map Algebra.TensorProduct.includeLeftRingHom.toMonoidHom δ * g) = δ • (toFun g)
55
+ toFun (incl₁ δ * g) = δ • (toFun g)
55
56
has_character : ∀ (g : (D ⊗[F] (FiniteAdeleRing (𝓞 F) F))ˣ) (z : (FiniteAdeleRing (𝓞 F) F)ˣ),
56
57
toFun (g * incl₂ F D z) = χ z • toFun g
57
58
right_invt : ∀ (g : (D ⊗[F] (FiniteAdeleRing (𝓞 F) F))ˣ),
@@ -62,7 +63,7 @@ namespace AutomorphicForm
62
63
-- defined over R
63
64
variable (R : Type *) [CommRing R]
64
65
-- weight
65
- (W : Type *) [AddCommGroup W] [Module R W] [MulAction Dˣ W] -- actions should commute in practice
66
+ (W : Type *) [AddCommGroup W] [Module R W] [DistribMulAction Dˣ W] -- actions should commute in practice
66
67
-- level
67
68
(U : Subgroup (D ⊗[F] (FiniteAdeleRing (𝓞 F) F))ˣ) -- subgroup should be compact and open
68
69
-- character
@@ -81,9 +82,9 @@ theorem ext (φ ψ : AutomorphicForm F D R W U χ) (h : ∀ x, φ x = ψ x) : φ
81
82
82
83
def zero : (AutomorphicForm F D R W U χ) where
83
84
toFun := 0
84
- left_invt := sorry
85
- has_character := sorry
86
- right_invt := sorry
85
+ left_invt δ _ := by simp
86
+ has_character _ z := by simp
87
+ right_invt _ u _ := by simp
87
88
88
89
instance : Zero (AutomorphicForm F D R W U χ) where
89
90
zero := zero
@@ -94,9 +95,9 @@ theorem zero_apply (x : (D ⊗[F] (FiniteAdeleRing (𝓞 F) F))ˣ) :
94
95
95
96
def neg (φ : AutomorphicForm F D R W U χ) : AutomorphicForm F D R W U χ where
96
97
toFun x := - φ x
97
- left_invt := sorry
98
- has_character := sorry
99
- right_invt := sorry
98
+ left_invt δ g := by simp [left_invt]
99
+ has_character g z := by simp [has_character]
100
+ right_invt g u hu := by simp_all [right_invt]
100
101
101
102
instance : Neg (AutomorphicForm F D R W U χ) where
102
103
neg := neg
@@ -107,9 +108,9 @@ theorem neg_apply (φ : AutomorphicForm F D R W U χ) (x : (D ⊗[F] (FiniteAdel
107
108
108
109
instance add (φ ψ : AutomorphicForm F D R W U χ) : AutomorphicForm F D R W U χ where
109
110
toFun x := φ x + ψ x
110
- left_invt := sorry
111
- has_character := sorry
112
- right_invt := sorry
111
+ left_invt := by simp [left_invt]
112
+ has_character := by simp [has_character]
113
+ right_invt := by simp_all [right_invt]
113
114
114
115
instance : Add (AutomorphicForm F D R W U χ) where
115
116
add := add
@@ -136,9 +137,11 @@ variable [SMulCommClass R Dˣ W]
136
137
def smul (r : R) (φ : AutomorphicForm F D R W U χ) :
137
138
AutomorphicForm F D R W U χ where
138
139
toFun g := r • φ g
139
- left_invt := sorry
140
- has_character := sorry
141
- right_invt := sorry
140
+ left_invt := by simp [left_invt, smul_comm]
141
+ has_character g z := by
142
+ simp_all [has_character]
143
+ rw [smul_comm] -- makes simp loop :-/
144
+ right_invt := by simp_all [right_invt]
142
145
143
146
instance : SMul R (AutomorphicForm F D R W U χ) where
144
147
smul := smul
@@ -158,7 +161,7 @@ end AutomorphicForm
158
161
159
162
variable (R : Type *) [Field R]
160
163
-- weight
161
- (W : Type *) [AddCommGroup W] [Module R W] [MulAction Dˣ W] -- actions should commute in practice
164
+ (W : Type *) [AddCommGroup W] [Module R W] [DistribMulAction Dˣ W] [SMulCommClass R Dˣ W]
162
165
-- level
163
166
(U : Subgroup (D ⊗[F] (FiniteAdeleRing (𝓞 F) F))ˣ) -- subgroup should be compact and open
164
167
-- character
0 commit comments