Skip to content

Commit f7fc80f

Browse files
authored
Merge pull request #109 from scholzhannah/QuaternionAlgebra2
Quaternion algebra2
2 parents 9d03193 + 72b3377 commit f7fc80f

File tree

1 file changed

+46
-34
lines changed

1 file changed

+46
-34
lines changed

FLT/AutomorphicForm/QuaternionAlgebra.lean

+46-34
Original file line numberDiff line numberDiff line change
@@ -1,12 +1,13 @@
11
/-
22
Copyright (c) 2024 Kevin Buzzard. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
4-
Authors: Kevin Buzzard
4+
Authors: Kevin Buzzard, Ludwig Monnerjahn, Hannah Scholz
55
-/
66
import Mathlib.Geometry.Manifold.Instances.UnitsOfNormedAlgebra
77
import Mathlib.LinearAlgebra.Matrix.GeneralLinearGroup.Basic
88
import Mathlib.NumberTheory.NumberField.Basic
99
import Mathlib.RingTheory.DedekindDomain.FiniteAdeleRing
10+
import Mathlib.Algebra.Group.Subgroup.Pointwise
1011
import FLT.HIMExperiments.module_topology
1112

1213
/-
@@ -20,27 +21,20 @@ variable (F : Type*) [Field F] [NumberField F]
2021

2122
variable (D : Type*) [Ring D] [Algebra F D] [FiniteDimensional F D]
2223

23-
#check DedekindDomain.FiniteAdeleRing
24-
2524
open DedekindDomain
2625

2726
open scoped NumberField
2827

29-
#check FiniteAdeleRing (𝓞 F) F
30-
3128
-- my work (two PRs)
3229
instance : TopologicalSpace (FiniteAdeleRing (𝓞 F) F) := sorry
3330
instance : TopologicalRing (FiniteAdeleRing (𝓞 F) F) := sorry
3431

3532
open scoped TensorProduct
3633

37-
#check D ⊗[F] (FiniteAdeleRing (𝓞 F) F)
38-
3934
section missing_instances
4035

4136
variable {R D A : Type*} [CommRing R] [Ring D] [CommRing A] [Algebra R D] [Algebra R A]
4237

43-
--TODO:
4438
instance : Algebra A (D ⊗[R] A) :=
4539
Algebra.TensorProduct.includeRight.toRingHom.toAlgebra' (by
4640
simp only [AlgHom.toRingHom_eq_coe, RingHom.coe_coe, Algebra.TensorProduct.includeRight_apply]
@@ -54,10 +48,7 @@ instance : Algebra A (D ⊗[R] A) :=
5448
rw [left_distrib, hx, hy, right_distrib]
5549
)
5650

57-
58-
5951
instance [Module.Finite R D] : Module.Finite A (D ⊗[R] A) := sorry
60-
6152
instance [Module.Free R D] : Module.Free A (D ⊗[R] A) := sorry
6253

6354
-- #synth Ring (D ⊗[F] FiniteAdeleRing (𝓞 F) F)
@@ -165,34 +156,55 @@ instance addCommGroup : AddCommGroup (AutomorphicForm F D M) where
165156
add_left_neg := by intros; ext; simp
166157
add_comm := by intros; ext; simp [add_comm]
167158

168-
-- this should be a SMul instance first, and then a simp lemma SMul_eval, and then one_smul etc are easy
169-
instance : MulAction (Dfx F D) (AutomorphicForm F D M) where
170-
smul g φ := {
171-
toFun := fun x => φ (x*g),
159+
open ConjAct
160+
open scoped Pointwise
161+
162+
lemma conjAct_mem {G: Type*} [Group G] (U: Subgroup G) (g: G) (x : G):
163+
x ∈ toConjAct g • U ↔ ∃ u ∈ U, g * u * g⁻¹ = x := by rfl
164+
165+
lemma toConjAct_open {G : Type*} [Group G] [TopologicalSpace G] [TopologicalGroup G]
166+
(U : Subgroup G) (hU : IsOpen (U : Set G)) (g : G) : IsOpen (toConjAct g • U : Set G) := by
167+
have this1 := continuous_mul_left g⁻¹
168+
have this2 := continuous_mul_right g
169+
rw [continuous_def] at this1 this2
170+
specialize this2 U hU
171+
specialize this1 _ this2
172+
convert this1 using 1
173+
ext x
174+
convert conjAct_mem _ _ _ using 1
175+
simp only [Set.mem_preimage, SetLike.mem_coe]
176+
refine ⟨?_, ?_⟩ <;> intro h
177+
· use g⁻¹ * x * g -- duh
178+
simp [h]
179+
group
180+
· rcases h with ⟨u, hu, rfl⟩
181+
group
182+
exact hu
183+
184+
instance : SMul (Dfx F D) (AutomorphicForm F D M) where
185+
smul g φ := { -- (g • f) (x) := f(xg) -- x(gf)=(xg)f
186+
toFun := fun x => φ (x * g)
172187
left_invt := by
173188
intros d x
174-
simp only [← φ.left_invt d x]
175-
rw [mul_assoc]
189+
simp only [← φ.left_invt d x, mul_assoc]
176190
exact φ.left_invt d (x * g)
177191
loc_cst := by
178192
rcases φ.loc_cst with ⟨U, openU, hU⟩
179-
use U -- not mathematically correct
193+
use toConjAct g • U
180194
constructor
181-
· exact openU
195+
· apply toConjAct_open _ openU
182196
· intros x u umem
183197
simp only
184-
sorry
185-
} -- (g • f) (x) := f(xg) -- x(gf)=(xg)f
186-
one_smul := by
187-
intros φ
188-
ext df
189-
-- missing simp lemma
190-
change φ (df * 1) = φ df
191-
simp
192-
mul_smul := by
193-
intros g h φ
194-
sorry
195-
-- if M is an R-module (e.g. if M = R!), then Automorphic forms are also an R-module
196-
-- with the action being 0on the coefficients.
197-
198-
example(a b c :ℝ ): a * b * c = (a * b) * c := rfl
198+
rw[conjAct_mem] at umem
199+
obtain ⟨ugu, hugu, eq⟩ := umem
200+
rw[←eq, ←mul_assoc, ←mul_assoc, inv_mul_cancel_right, hU (x*g) ugu hugu]
201+
}
202+
203+
@[simp]
204+
lemma sMul_eval (g : Dfx F D) (f : AutomorphicForm F D M) (x : (D ⊗[F] FiniteAdeleRing (𝓞 F) F)ˣ) :
205+
(g • f) x = f (x * g) := rfl
206+
207+
instance : MulAction (Dfx F D) (AutomorphicForm F D M) where
208+
smul := (. • .)
209+
one_smul := by intros; ext; simp only [sMul_eval, mul_one]
210+
mul_smul := by intros; ext; simp only [sMul_eval, mul_assoc]

0 commit comments

Comments
 (0)