Skip to content

Commit aa6dd4c

Browse files
committed
auto forms are an addcommgroup
1 parent dd04633 commit aa6dd4c

File tree

2 files changed

+23
-20
lines changed

2 files changed

+23
-20
lines changed

FLT/AutomorphicForm/QuaternionAlgebra.lean

+18-15
Original file line numberDiff line numberDiff line change
@@ -28,6 +28,7 @@ open DedekindDomain
2828

2929
abbrev Dfx := (D ⊗[F] (FiniteAdeleRing (𝓞 F) F))ˣ
3030

31+
variable {F D} in
3132
noncomputable abbrev incl₁ : Dˣ →* Dfx F D :=
3233
Units.map Algebra.TensorProduct.includeLeftRingHom.toMonoidHom
3334

@@ -51,7 +52,7 @@ structure AutomorphicForm
5152
-- definition
5253
toFun : (D ⊗[F] (FiniteAdeleRing (𝓞 F) F))ˣ → W
5354
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)
5556
has_character : ∀ (g : (D ⊗[F] (FiniteAdeleRing (𝓞 F) F))ˣ) (z : (FiniteAdeleRing (𝓞 F) F)ˣ),
5657
toFun (g * incl₂ F D z) = χ z • toFun g
5758
right_invt : ∀ (g : (D ⊗[F] (FiniteAdeleRing (𝓞 F) F))ˣ),
@@ -62,7 +63,7 @@ namespace AutomorphicForm
6263
-- defined over R
6364
variable (R : Type*) [CommRing R]
6465
-- 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
6667
-- level
6768
(U : Subgroup (D ⊗[F] (FiniteAdeleRing (𝓞 F) F))ˣ) -- subgroup should be compact and open
6869
-- character
@@ -81,9 +82,9 @@ theorem ext (φ ψ : AutomorphicForm F D R W U χ) (h : ∀ x, φ x = ψ x) : φ
8182

8283
def zero : (AutomorphicForm F D R W U χ) where
8384
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
8788

8889
instance : Zero (AutomorphicForm F D R W U χ) where
8990
zero := zero
@@ -94,9 +95,9 @@ theorem zero_apply (x : (D ⊗[F] (FiniteAdeleRing (𝓞 F) F))ˣ) :
9495

9596
def neg (φ : AutomorphicForm F D R W U χ) : AutomorphicForm F D R W U χ where
9697
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]
100101

101102
instance : Neg (AutomorphicForm F D R W U χ) where
102103
neg := neg
@@ -107,9 +108,9 @@ theorem neg_apply (φ : AutomorphicForm F D R W U χ) (x : (D ⊗[F] (FiniteAdel
107108

108109
instance add (φ ψ : AutomorphicForm F D R W U χ) : AutomorphicForm F D R W U χ where
109110
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]
113114

114115
instance : Add (AutomorphicForm F D R W U χ) where
115116
add := add
@@ -136,9 +137,11 @@ variable [SMulCommClass R Dˣ W]
136137
def smul (r : R) (φ : AutomorphicForm F D R W U χ) :
137138
AutomorphicForm F D R W U χ where
138139
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]
142145

143146
instance : SMul R (AutomorphicForm F D R W U χ) where
144147
smul := smul
@@ -158,7 +161,7 @@ end AutomorphicForm
158161

159162
variable (R : Type*) [Field R]
160163
-- 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]
162165
-- level
163166
(U : Subgroup (D ⊗[F] (FiniteAdeleRing (𝓞 F) F))ˣ) -- subgroup should be compact and open
164167
-- character

blueprint/lean_decls

+5-5
Original file line numberDiff line numberDiff line change
@@ -73,12 +73,12 @@ MulSemiringAction.CharacteristicPolynomial.Mbar_monic
7373
MulSemiringAction.CharacteristicPolynomial.Mbar_eval_eq_zero
7474
MulSemiringAction.reduction_isIntegral
7575
Algebra.exists_dvd_nonzero_if_isIntegral
76-
TotallyDefiniteQuaternionAlgebra.AutomorphicForm
77-
TotallyDefiniteQuaternionAlgebra.AutomorphicForm.addCommGroup
78-
TotallyDefiniteQuaternionAlgebra.AutomorphicForm.module
79-
TotallyDefiniteQuaternionAlgebra.AutomorphicForm.finiteDimensional
80-
TotallyDefiniteQuaternionAlgebra.finiteDoubleCoset
8176
Bourbaki52222.f_exists
8277
Bourbaki52222.algebraic
8378
Bourbaki52222.normal
8479
Bourbaki52222.separableClosure_finrank_le
80+
TotallyDefiniteQuaternionAlgebra.AutomorphicForm
81+
TotallyDefiniteQuaternionAlgebra.AutomorphicForm.addCommGroup
82+
TotallyDefiniteQuaternionAlgebra.AutomorphicForm.module
83+
TotallyDefiniteQuaternionAlgebra.AutomorphicForm.finiteDimensional
84+
TotallyDefiniteQuaternionAlgebra.finiteDoubleCoset

0 commit comments

Comments
 (0)