1
1
import Mathlib.RingTheory.DedekindDomain.FiniteAdeleRing
2
2
import Mathlib.NumberTheory.NumberField.Basic
3
+ import Mathlib
3
4
4
5
universe u
5
6
6
- variable (K : Type *) [Field K] [NumberField K]
7
-
8
7
section PRed
9
8
-- Don't try anything in this section because we are missing a definition.
10
9
@@ -13,6 +12,8 @@ def NumberField.AdeleRing (K : Type u) [Field K] [NumberField K] : Type u := sor
13
12
14
13
open NumberField
15
14
15
+ variable (K : Type *) [Field K] [NumberField K]
16
+
16
17
-- All these are already done in mathlib PRs, so can be assumed for now.
17
18
-- When they're in mathlib master we can update mathlib and then get these theorems/definitions
18
19
-- for free.
@@ -41,6 +42,8 @@ section TODO
41
42
-- these theorems can't be worked on yet because we don't have an actual definition
42
43
-- of the adele ring. We can work out a strategy when we do.
43
44
45
+ variable (K : Type *) [Field K] [NumberField K]
46
+
44
47
-- Maybe this isn't even stated in the best way?
45
48
theorem NumberField.AdeleRing.discrete : ∀ k : K, ∃ U : Set (AdeleRing K),
46
49
IsOpen U ∧ (algebraMap K (AdeleRing K)) ⁻¹' U = {k} := sorry
@@ -56,48 +59,89 @@ end TODO
56
59
57
60
-- ************ *** Sorries which are not impossible start here *** ************ **
58
61
59
- -- recall we already had `variable (K : Type*) [Field K] [NumberField K]`
62
+ variable (A K L B : Type *) [CommRing A] [CommRing B] [Algebra A B] [Field K] [Field L]
63
+ [Algebra A K] [IsFractionRing A K] [Algebra B L] [IsDedekindDomain A]
64
+ [Algebra K L] [Algebra A L] [IsScalarTower A B L] [IsScalarTower A K L]
65
+ [IsIntegralClosure B A L] [FiniteDimensional K L]
66
+
67
+ variable [IsDomain B] -- is this necessary?
68
+ variable [Algebra.IsSeparable K L]
60
69
61
- variable (L : Type *) [Field L] [Algebra K L] [FiniteDimensional K L]
70
+ -- example : IsDedekindDomain B := IsIntegralClosure.isDedekindDomain A K L B
62
71
63
- instance : NumberField L := sorry -- inferInstance fails
72
+ variable [IsDedekindDomain B]
64
73
65
- -- We want the map `L ⊗[K] 𝔸_K → 𝔸_L`, but we don't have a definition of adeles.
66
- -- However we do have a definition of finite adeles, so let's work there.
74
+ -- example : IsFractionRing B L := IsIntegralClosure.isFractionRing_of_finite_extension A K L B
67
75
76
+ variable [IsFractionRing B L]
77
+
78
+ -- example : Algebra.IsIntegral A B := IsIntegralClosure.isIntegral_algebra A L
79
+
80
+ variable [Algebra.IsIntegral A B]
81
+
82
+ -- Conjecture: in this generality there's a natural isomorphism `L ⊗[K] 𝔸_K^∞ → 𝔸_L^∞`
83
+ -- Note however that we don't have a definition of adeles in this generality.
84
+
85
+ -- One key input is the purely local statement that if L/K
86
+ -- is a finite extension of number fields, and v is a place of K,
87
+ -- then `L ⊗[K] Kᵥ ≅ ⊕_w|v L_w`, where the sum is over the places w
88
+ -- of L above v.
68
89
namespace DedekindDomain
69
90
70
- open scoped TensorProduct NumberField -- ⊗ notation for tensor product and 𝓞 K notation for ring of ints
91
+ open scoped TensorProduct -- ⊗ notation for tensor product
71
92
72
- noncomputable instance : Algebra K (ProdAdicCompletions (𝓞 L) L) := RingHom.toAlgebra <|
73
- (algebraMap L (ProdAdicCompletions (𝓞 L) L)).comp (algebraMap K L)
93
+ noncomputable instance : Algebra K (ProdAdicCompletions B L) := RingHom.toAlgebra <|
94
+ (algebraMap L (ProdAdicCompletions B L)).comp (algebraMap K L)
74
95
75
- -- Need to pull back elements of HeightOneSpectrum (𝓞 L) to HeightOneSpectrum (𝓞 K)
96
+ -- Need to pull back elements of HeightOneSpectrum B to HeightOneSpectrum A
76
97
open IsDedekindDomain
77
98
78
- variable {L} in
79
- def HeightOneSpectrum.comap (w : HeightOneSpectrum (𝓞 L)) : HeightOneSpectrum (𝓞 K) where
80
- asIdeal := w.asIdeal.comap (algebraMap (𝓞 K) (𝓞 L))
81
- isPrime := Ideal.comap_isPrime (algebraMap (𝓞 K) (𝓞 L)) w.asIdeal
82
- ne_bot := sorry -- pullback of nonzero prime is nonzero? Certainly true for number fields.
83
- -- Idea of proof: show that if B/A is integral and B is an integral domain, then
84
- -- any nonzero element of B divides a nonzero element of A.
85
- -- In fact this definition should be made in that generality.
86
- -- Use `Algebra.exists_dvd_nonzero_if_isIntegral` in the FLT file MathlibExperiments/FrobeniusRiou.lean
87
-
88
- variable {L} in
89
- noncomputable def HeightOneSpectrum.of_comap (w : HeightOneSpectrum (𝓞 L)) :
90
- (HeightOneSpectrum.adicCompletion K (comap K w)) →+* (HeightOneSpectrum.adicCompletion L w) :=
91
- letI : UniformSpace K := (comap K w).adicValued.toUniformSpace;
99
+ namespace HeightOneSpectrum
100
+ -- first need a map from finite places of 𝓞L to finite places of 𝓞K
101
+ variable {B L} in
102
+ def comap (w : HeightOneSpectrum B) : HeightOneSpectrum A where
103
+ asIdeal := w.asIdeal.comap (algebraMap A B)
104
+ isPrime := Ideal.comap_isPrime (algebraMap A B) w.asIdeal
105
+ ne_bot := mt Ideal.eq_bot_of_comap_eq_bot w.ne_bot
106
+
107
+ open scoped algebraMap
108
+
109
+ -- homework
110
+ def valuation_comap (w : HeightOneSpectrum B) (x : K) :
111
+ (comap A w).valuation x = w.valuation (algebraMap K L x) ^ (Ideal.ramificationIdx (algebraMap A B) (comap A w).asIdeal w.asIdeal) := sorry
112
+
113
+ -- Say w is a finite place of L lying above v a finite places
114
+ -- of K. Then there's a ring hom K_v -> L_w.
115
+ variable {B L} in
116
+ noncomputable def of_comap (w : HeightOneSpectrum B) :
117
+ (HeightOneSpectrum.adicCompletion K (comap A w)) →+* (HeightOneSpectrum.adicCompletion L w) :=
118
+ letI : UniformSpace K := (comap A w).adicValued.toUniformSpace;
92
119
letI : UniformSpace L := w.adicValued.toUniformSpace;
93
120
UniformSpace.Completion.mapRingHom (algebraMap K L) <| by
121
+ -- question is the following:
122
+ -- if L/K is a finite extension of sensible fields (e.g. number fields)
123
+ -- and if w is a finite place of L lying above v a finite place of K,
124
+ -- and if we give L the w-adic topology and K the v-adic topology,
125
+ -- then the map K → L is continuous
126
+ refine continuous_of_continuousAt_zero (algebraMap K L) ?hf
127
+ delta ContinuousAt
128
+ simp only [map_zero]
129
+ rw [(@Valued.hasBasis_nhds_zero K _ _ _ (comap A w).adicValued).tendsto_iff
130
+ (@Valued.hasBasis_nhds_zero L _ _ _ w.adicValued)]
131
+ simp only [HeightOneSpectrum.adicValued_apply, Set.mem_setOf_eq, true_and, true_implies]
132
+ refine fun i ↦ ⟨i ^ (1 / Ideal.ramificationIdx (algebraMap A B) (comap A w).asIdeal w.asIdeal), fun _ h ↦ ?_⟩
133
+ -- last line is nonsense but use `valuation_comap` to finish
134
+ push_cast at h
94
135
sorry
95
136
96
- open HeightOneSpectrum in
137
+ end HeightOneSpectrum
138
+
139
+ open HeightOneSpectrum
140
+
97
141
noncomputable def ProdAdicCompletions.baseChange :
98
- L ⊗[K] ProdAdicCompletions (𝓞 K) K →ₗ[K] ProdAdicCompletions (𝓞 L) L := TensorProduct.lift <| {
142
+ L ⊗[K] ProdAdicCompletions A K →ₗ[K] ProdAdicCompletions B L := TensorProduct.lift <| {
99
143
toFun := fun l ↦ {
100
- toFun := fun kv w ↦ l • (of_comap K w (kv (comap K w)))
144
+ toFun := fun kv w ↦ l • (of_comap A K w (kv (comap A w)))
101
145
map_add' := sorry
102
146
map_smul' := sorry
103
147
}
@@ -106,8 +150,10 @@ noncomputable def ProdAdicCompletions.baseChange :
106
150
}
107
151
108
152
-- hard but hopefully enough (this proof will be a lot of work)
109
- theorem ProdAdicCompletions.baseChange_iso (x : ProdAdicCompletions (𝓞 K) K) :
153
+ theorem ProdAdicCompletions.baseChange_iso (x : ProdAdicCompletions A K) :
110
154
ProdAdicCompletions.IsFiniteAdele x ↔
111
- ProdAdicCompletions.IsFiniteAdele (ProdAdicCompletions.baseChange K L (1 ⊗ₜ x)) := sorry
155
+ ProdAdicCompletions.IsFiniteAdele (ProdAdicCompletions.baseChange A K L B (1 ⊗ₜ x)) := sorry
156
+
157
+ end DedekindDomain
112
158
113
159
-- Can we now write down the isomorphism L ⊗ 𝔸_K^∞ = 𝔸_L^∞ ?
0 commit comments