Skip to content

Commit 0806dd4

Browse files
authored
Merge pull request #126 from ImperialCollegeLondon/start_on_localization
Start on localization and immediately run into trouble
2 parents 69a6c0d + c3a4078 commit 0806dd4

File tree

4 files changed

+108
-40
lines changed

4 files changed

+108
-40
lines changed

FLT/FLT_files.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -3,7 +3,6 @@ import FLT.AutomorphicRepresentation.Example
33
import FLT.Basic.Reductions
44
import FLT.EllipticCurve.Torsion
55
import FLT.ForMathlib.ActionTopology
6-
import FLT.ForMathlib.FGModuleTopology
76
import FLT.ForMathlib.MiscLemmas
87
import FLT.GaloisRepresentation.Cyclotomic
98
import FLT.GaloisRepresentation.HardlyRamified
@@ -12,6 +11,7 @@ import FLT.GlobalLanglandsConjectures.GLzero
1211
import FLT.GroupScheme.FiniteFlat
1312
import FLT.HIMExperiments.ContinuousSMul_topology
1413
import FLT.HIMExperiments.dual_topology
14+
import FLT.HIMExperiments.FGModuleTopology
1515
import FLT.HIMExperiments.flatness
1616
import FLT.HIMExperiments.module_topology
1717
import FLT.HIMExperiments.right_module_topology

FLT/ForMathlib/ActionTopology.lean

+14-14
Original file line numberDiff line numberDiff line change
@@ -95,13 +95,13 @@ the action topology. See `actionTopology` for more discussion of the action topo
9595
class IsActionTopology [τA : TopologicalSpace A] : Prop where
9696
isActionTopology' : τA = actionTopology R A
9797

98-
lemma isActionTopology [τA : TopologicalSpace A] [IsActionTopology R A] :
98+
theorem isActionTopology [τA : TopologicalSpace A] [IsActionTopology R A] :
9999
τA = actionTopology R A :=
100100
IsActionTopology.isActionTopology' (R := R) (A := A)
101101

102102
/-- Scalar multiplication `• : R × A → A` is continuous if `R` is a topological
103103
ring, and `A` is an `R` module with the action topology. -/
104-
lemma ActionTopology.continuousSMul : @ContinuousSMul R A _ _ (actionTopology R A) :=
104+
theorem ActionTopology.continuousSMul : @ContinuousSMul R A _ _ (actionTopology R A) :=
105105
-- Proof: We need to prove that the product topology is finer than the pullback
106106
-- of the action topology. But the action topology is an Inf and thus a limit,
107107
-- and pullback is a right adjoint, so it preserves limits.
@@ -112,16 +112,16 @@ lemma ActionTopology.continuousSMul : @ContinuousSMul R A _ _ (actionTopology R
112112

113113
/-- Addition `+ : A × A → A` is continuous if `R` is a topological
114114
ring, and `A` is an `R` module with the action topology. -/
115-
lemma ActionTopology.continuousAdd : @ContinuousAdd A (actionTopology R A) _ :=
115+
theorem ActionTopology.continuousAdd : @ContinuousAdd A (actionTopology R A) _ :=
116116
continuousAdd_sInf <| fun _ _ ↦ by simp_all only [Set.mem_setOf_eq]
117117

118118
instance instIsActionTopology_continuousSMul [TopologicalSpace A] [IsActionTopology R A] :
119119
ContinuousSMul R A := isActionTopology R A ▸ ActionTopology.continuousSMul R A
120120

121-
lemma isActionTopology_continuousAdd [TopologicalSpace A] [IsActionTopology R A] :
121+
theorem isActionTopology_continuousAdd [TopologicalSpace A] [IsActionTopology R A] :
122122
ContinuousAdd A := isActionTopology R A ▸ ActionTopology.continuousAdd R A
123123

124-
lemma actionTopology_le [τA : TopologicalSpace A] [ContinuousSMul R A] [ContinuousAdd A] :
124+
theorem actionTopology_le [τA : TopologicalSpace A] [ContinuousSMul R A] [ContinuousAdd A] :
125125
actionTopology R A ≤ τA := sInf_le ⟨‹ContinuousSMul R A›, ‹ContinuousAdd A›⟩
126126

127127
end basics
@@ -195,7 +195,7 @@ variable {B : Type*} [AddCommMonoid B] [Module R B] [τB : TopologicalSpace B]
195195
-- this is horrible. Why isn't it easy?
196196
-- One reason: we are rolling our own continuous linear equivs!
197197
-- **TODO** Ask about making continuous linear equivs properly
198-
lemma iso (ehomeo : A ≃ₜ B) (elinear : A ≃ₗ[R] B) (he : ∀ a, ehomeo a = elinear a) :
198+
theorem iso (ehomeo : A ≃ₜ B) (elinear : A ≃ₗ[R] B) (he : ∀ a, ehomeo a = elinear a) :
199199
IsActionTopology R B where
200200
isActionTopology' := by
201201
simp_rw [ehomeo.symm.inducing.1, isActionTopology R A, actionTopology, induced_sInf]
@@ -237,7 +237,7 @@ variable {B : Type*} [AddCommMonoid B] [Module R B] [aB : TopologicalSpace B] [I
237237

238238
/-- Every `R`-linear map between two `R`-modules with the canonical topology is continuous. -/
239239
@[fun_prop, continuity]
240-
lemma continuous_of_distribMulActionHom (φ : A →+[R] B) : Continuous φ := by
240+
theorem continuous_of_distribMulActionHom (φ : A →+[R] B) : Continuous φ := by
241241
-- the proof: We know that `+ : B × B → B` and `• : R × B → B` are continuous for the action
242242
-- topology on `B`, and two earlier theorems (`induced_continuous_smul` and
243243
-- `induced_continuous_add`) say that hence `+` and `•` on `A` are continuous if `A`
@@ -249,11 +249,11 @@ lemma continuous_of_distribMulActionHom (φ : A →+[R] B) : Continuous φ := by
249249
induced_continuous_add φ.toAddMonoidHom⟩
250250

251251
@[fun_prop, continuity]
252-
lemma continuous_of_linearMap (φ : A →ₗ[R] B) : Continuous φ :=
252+
theorem continuous_of_linearMap (φ : A →ₗ[R] B) : Continuous φ :=
253253
continuous_of_distribMulActionHom φ.toDistribMulActionHom
254254

255255
variable (R) in
256-
lemma continuous_neg (C : Type*) [AddCommGroup C] [Module R C] [TopologicalSpace C]
256+
theorem continuous_neg (C : Type*) [AddCommGroup C] [Module R C] [TopologicalSpace C]
257257
[IsActionTopology R C] : Continuous (fun a ↦ -a : C → C) :=
258258
continuous_of_linearMap (LinearEquiv.neg R).toLinearMap
259259

@@ -267,7 +267,7 @@ variable {B : Type*} [AddCommGroup B] [Module R B] [aB : TopologicalSpace B] [Is
267267

268268
-- Here I need the lemma about how quotients are open so I do need groups
269269
-- because this relies on translates of an open being open
270-
lemma coinduced_of_surjective {φ : A →ₗ[R] B} (hφ : Function.Surjective φ) :
270+
theorem coinduced_of_surjective {φ : A →ₗ[R] B} (hφ : Function.Surjective φ) :
271271
TopologicalSpace.coinduced φ (actionTopology R A) = actionTopology R B := by
272272
have : Continuous φ := continuous_of_linearMap φ
273273
rw [continuous_iff_coinduced_le, isActionTopology R A, isActionTopology R B] at this
@@ -401,7 +401,7 @@ variable {A : Type*} [AddCommGroup A] [Module R A] [aA : TopologicalSpace A] [Is
401401
variable {B : Type*} [AddCommGroup B] [Module R B] [aB : TopologicalSpace B] [IsActionTopology R B]
402402
variable {C : Type*} [AddCommGroup C] [Module R C] [aC : TopologicalSpace C] [IsActionTopology R C]
403403

404-
lemma Module.continuous_bilinear_of_pi_finite (ι : Type*) [Finite ι]
404+
theorem Module.continuous_bilinear_of_pi_finite (ι : Type*) [Finite ι]
405405
(bil : (ι → R) →ₗ[R] B →ₗ[R] C) : Continuous (fun ab ↦ bil ab.1 ab.2 : ((ι → R) × B → C)) := by
406406
classical
407407
have foo : (fun fb ↦ bil fb.1 fb.2 : ((ι → R) × B → C)) =
@@ -431,7 +431,7 @@ lemma Module.continuous_bilinear_of_pi_finite (ι : Type*) [Finite ι]
431431
simp [Set.toFinite _]
432432

433433
-- Probably this can be beefed up to semirings.
434-
lemma Module.continuous_bilinear_of_finite_free [TopologicalSemiring R] [Module.Finite R A] [Module.Free R A]
434+
theorem Module.continuous_bilinear_of_finite_free [TopologicalSemiring R] [Module.Finite R A] [Module.Free R A]
435435
(bil : A →ₗ[R] B →ₗ[R] C) : Continuous (fun ab ↦ bil ab.1 ab.2 : (A × B → C)) := by
436436
let ι := Module.Free.ChooseBasisIndex R A
437437
let hι : Fintype ι := Module.Free.ChooseBasisIndex.fintype R A
@@ -460,7 +460,7 @@ variable {B : Type*} [AddCommGroup B] [Module R B] [aB : TopologicalSpace B] [Is
460460
variable {C : Type*} [AddCommGroup C] [Module R C] [aC : TopologicalSpace C] [IsActionTopology R C]
461461

462462
-- This needs rings though
463-
lemma Module.continuous_bilinear_of_finite [Module.Finite R A]
463+
theorem Module.continuous_bilinear_of_finite [Module.Finite R A]
464464
(bil : A →ₗ[R] B →ₗ[R] C) : Continuous (fun ab ↦ bil ab.1 ab.2 : (A × B → C)) := by
465465
obtain ⟨m, f, hf⟩ := Module.Finite.exists_fin' R A
466466
let bil' : (Fin m → R) →ₗ[R] B →ₗ[R] C := bil.comp f
@@ -494,7 +494,7 @@ variable [TopologicalSpace D] [IsActionTopology R D]
494494
open scoped TensorProduct
495495

496496
@[continuity, fun_prop]
497-
lemma continuous_mul'
497+
theorem continuous_mul'
498498
(R) [CommRing R] [TopologicalSpace R] [TopologicalRing R]
499499
(D : Type*) [Ring D] [Algebra R D] [Module.Finite R D] [Module.Free R D] [TopologicalSpace D]
500500
[IsActionTopology R D]: Continuous (fun ab ↦ ab.1 * ab.2 : D × D → D) := by

FLT/HIMExperiments/FGModuleTopology.lean

+16-16
Original file line numberDiff line numberDiff line change
@@ -50,7 +50,7 @@ class IsModuleTopology [τA : TopologicalSpace A] : Prop where
5050

5151
-- because default binders for isModuleTopology' are wrong and currently
5252
-- there is no way to change this. There's a Lean 4 issue for this IIRC **TODO** where?
53-
lemma isModuleTopology [τA : TopologicalSpace A] [IsModuleTopology R A] :
53+
theorem isModuleTopology [τA : TopologicalSpace A] [IsModuleTopology R A] :
5454
τA = moduleTopology R A :=
5555
IsModuleTopology.isModuleTopology' (R := R) (A := A)
5656

@@ -88,14 +88,14 @@ variable {B : Type*} [AddCommMonoid B] [Module R B] [aB : TopologicalSpace B] [I
8888

8989
/-- Every `A`-linear map between two `A`-modules with the canonical topology is continuous. -/
9090
@[continuity, fun_prop]
91-
lemma continuous_of_linearMap (f : A →ₗ[R] B) : Continuous f := by
91+
theorem continuous_of_linearMap (f : A →ₗ[R] B) : Continuous f := by
9292
rw [isModuleTopology R A, isModuleTopology R B, continuous_iff_le_induced]
9393
apply iSup_le <| fun n ↦ iSup_le <| fun φ ↦ ?_
9494
rw [← coinduced_le_iff_le_induced, coinduced_compose]
9595
exact le_iSup_of_le n <| le_iSup_of_le (f.comp φ) le_rfl
9696

9797
-- should this be in the API?
98-
lemma continuous_of_linearMap' {A : Type*} [AddCommMonoid A] [Module R A]
98+
theorem continuous_of_linearMap' {A : Type*} [AddCommMonoid A] [Module R A]
9999
{B : Type*} [AddCommMonoid B] [Module R B] (f : A →ₗ[R] B) :
100100
@Continuous _ _ (moduleTopology R A) (moduleTopology R B) f := by
101101
letI : TopologicalSpace A := moduleTopology R A
@@ -133,7 +133,7 @@ end linear_map
133133

134134
section continuous_neg
135135

136-
lemma continuous_neg
136+
theorem continuous_neg
137137
(R : Type*) [TopologicalSpace R] [Ring R]
138138
(A : Type*) [AddCommGroup A] [Module R A] [TopologicalSpace A] [IsModuleTopology R A] :
139139
Continuous (-· : A → A) :=
@@ -146,7 +146,7 @@ section surj
146146
variable {R : Type*} [τR : TopologicalSpace R] [Semiring R] [TopologicalSemiring R]
147147
variable {A : Type*} [AddCommMonoid A] [Module R A] [aA : TopologicalSpace A] [IsModuleTopology R A]
148148

149-
lemma coinduced_of_surjectivePiFin {n : ℕ} {φ : ((Fin n) → R) →ₗ[R] A} (hφ : Function.Surjective φ) :
149+
theorem coinduced_of_surjectivePiFin {n : ℕ} {φ : ((Fin n) → R) →ₗ[R] A} (hφ : Function.Surjective φ) :
150150
TopologicalSpace.coinduced φ Pi.topologicalSpace = moduleTopology R A := by
151151
apply le_antisymm
152152
· rw [← continuous_iff_coinduced_le, ← isModuleTopology R A]
@@ -161,7 +161,7 @@ lemma coinduced_of_surjectivePiFin {n : ℕ} {φ : ((Fin n) → R) →ₗ[R] A}
161161
exact continuous_of_linearMap α
162162

163163
/-- Any surjection between finite R-modules is coinducing for the R-module topology. -/
164-
lemma coinduced_of_surjective {B : Type*} [AddCommMonoid B] [aB : TopologicalSpace B] [Module R B]
164+
theorem coinduced_of_surjective {B : Type*} [AddCommMonoid B] [aB : TopologicalSpace B] [Module R B]
165165
[IsModuleTopology R B] [Module.Finite R A] {φ : A →ₗ[R] B} (hφ : Function.Surjective φ) :
166166
TopologicalSpace.coinduced φ (moduleTopology R A) = moduleTopology R B := by
167167
obtain ⟨n, f, hf⟩ := Module.Finite.exists_fin' R A
@@ -173,7 +173,7 @@ lemma coinduced_of_surjective {B : Type*} [AddCommMonoid B] [aB : TopologicalSpa
173173

174174
-- do I need this? Yes, I need (fin n) × fin m -> R
175175
-- **^TODO** why didn't have/let linter warn me
176-
lemma coinduced_of_surjectivePiFinite {ι : Type*} [Finite ι] {φ : (ι → R) →ₗ[R] A}
176+
theorem coinduced_of_surjectivePiFinite {ι : Type*} [Finite ι] {φ : (ι → R) →ₗ[R] A}
177177
(hφ : Function.Surjective φ) :
178178
TopologicalSpace.coinduced φ Pi.topologicalSpace = moduleTopology R A := by
179179
rw [(instPiFinite R ι).isModuleTopology']
@@ -188,7 +188,7 @@ variable {A : Type*} [AddCommGroup A] [Module R A] [aA : TopologicalSpace A] [Is
188188

189189
variable (R A) in
190190
@[continuity, fun_prop]
191-
lemma continuous_add [Module.Finite R A] : Continuous (fun ab ↦ ab.1 + ab.2 : A × A → A) := by
191+
theorem continuous_add [Module.Finite R A] : Continuous (fun ab ↦ ab.1 + ab.2 : A × A → A) := by
192192
rw [continuous_iff_coinduced_le, isModuleTopology R A]
193193
obtain ⟨n, f, hf⟩ := Module.Finite.exists_fin' R A
194194
rw [← coinduced_of_surjectivePiFin hf]
@@ -213,7 +213,7 @@ lemma continuous_add [Module.Finite R A] : Continuous (fun ab ↦ ab.1 + ab.2 :
213213

214214
variable (R A) in
215215
@[continuity, fun_prop]
216-
lemma continuous_sum_finset (ι : Type*) [DecidableEq ι] (s : Finset ι) [Module.Finite R A] :
216+
theorem continuous_sum_finset (ι : Type*) [DecidableEq ι] (s : Finset ι) [Module.Finite R A] :
217217
Continuous (fun as ↦ ∑ i ∈ s, as i : (∀ (_ : ι), A) → A) := by
218218
induction s using Finset.induction
219219
· simp only [Finset.sum_empty]
@@ -227,7 +227,7 @@ lemma continuous_sum_finset (ι : Type*) [DecidableEq ι] (s : Finset ι) [Modul
227227
attribute [local instance] Fintype.ofFinite
228228
variable (R A) in
229229
@[continuity, fun_prop]
230-
lemma continuous_sum_finite (ι : Type*) [Finite ι] [Module.Finite R A] :
230+
theorem continuous_sum_finite (ι : Type*) [Finite ι] [Module.Finite R A] :
231231
Continuous (fun as ↦ ∑ i, as i : (∀ (_ : ι), A) → A) := by
232232
classical
233233
exact continuous_sum_finset R A ι _
@@ -316,7 +316,7 @@ noncomputable def isom'' (R : Type*) [CommRing R] (m n : Type*) [Finite m] [Deci
316316

317317
variable (m n : Type*) [Finite m] [DecidableEq m] (a1 : m → R)
318318
(b1 : n → R) (f : (m → R) →ₗ[R] A) (g : (n → R) →ₗ[R] B) in
319-
lemma key : ((TensorProduct.map f g ∘ₗ
319+
theorem key : ((TensorProduct.map f g ∘ₗ
320320
(isom'' R m n).symm.toLinearMap) fun mn ↦ a1 mn.1 * b1 mn.2) = f a1 ⊗ₜ[R] g b1 := by
321321
sorry
322322

@@ -336,7 +336,7 @@ lemma key : ((TensorProduct.map f g ∘ₗ
336336
-- I don't know whether finiteness is needed on `B` here. Removing it here would enable
337337
-- removal in `continuous_bilinear`.
338338
@[continuity, fun_prop]
339-
lemma Module.continuous_tprod [Module.Finite R A] [Module.Finite R B] :
339+
theorem Module.continuous_tprod [Module.Finite R A] [Module.Finite R B] :
340340
Continuous (fun (ab : A × B) ↦ ab.1 ⊗ₜ ab.2 : A × B → A ⊗[R] B) := by
341341
-- reduce to R^m x R^n -> R^m ⊗ R^n
342342
-- then check explicitly
@@ -377,7 +377,7 @@ lemma Module.continuous_tprod [Module.Finite R A] [Module.Finite R B] :
377377
rfl
378378

379379
-- did I really use finiteness of B?
380-
lemma Module.continuous_bilinear [Module.Finite R A] [Module.Finite R B]
380+
theorem Module.continuous_bilinear [Module.Finite R A] [Module.Finite R B]
381381
{C : Type*} [AddCommGroup C] [Module R C] [TopologicalSpace C] [IsModuleTopology R C]
382382
(bil : A →ₗ[R] B →ₗ[R] C) : Continuous (fun ab ↦ bil ab.1 ab.2 : (A × B → C)) := by
383383
letI : TopologicalSpace (A ⊗[R] B) := moduleTopology R _
@@ -390,7 +390,7 @@ variable (D : Type*) [Ring D] [Algebra R D] [Module.Finite R D]
390390
variable [TopologicalSpace D] [IsModuleTopology R D]
391391

392392
@[continuity, fun_prop]
393-
lemma continuous_mul (D : Type*) [Ring D] [Algebra R D] [Module.Finite R D] [TopologicalSpace D]
393+
theorem continuous_mul (D : Type*) [Ring D] [Algebra R D] [Module.Finite R D] [TopologicalSpace D]
394394
[IsModuleTopology R D] : Continuous (fun ab ↦ ab.1 * ab.2 : D × D → D) := by
395395
letI : TopologicalSpace (D ⊗[R] D) := moduleTopology R _
396396
haveI : IsModuleTopology R (D ⊗[R] D) := { isModuleTopology' := rfl }
@@ -403,7 +403,7 @@ def Module.topologicalRing : TopologicalRing D where
403403

404404
end commutative
405405

406-
lemma continuousSMul (R : Type*) [CommRing R] [TopologicalSpace R] [TopologicalRing R]
406+
theorem continuousSMul (R : Type*) [CommRing R] [TopologicalSpace R] [TopologicalRing R]
407407
(A : Type*) [AddCommGroup A] [Module R A] [Module.Finite R A] [TopologicalSpace A]
408408
[IsModuleTopology R A] :
409409
ContinuousSMul R A := by
@@ -420,7 +420,7 @@ I can only prove that `SMul : R × A → A` is continuous for the module topolog
420420
commutative (because my proof uses tensor products) and if `A` is finite (because
421421
I reduce to a basis check ). Is it true in general? I'm assuming not.
422422
423-
lemma continuousSMul (R : Type*) [Ring R] [TopologicalSpace R] [TopologicalRing R]
423+
theorem continuousSMul (R : Type*) [Ring R] [TopologicalSpace R] [TopologicalRing R]
424424
(A : Type*) [AddCommGroup A] [Module R A] : @ContinuousSMul R A _ _ (moduleTopology R A) := by
425425
refine @ContinuousSMul.mk ?_ ?_ ?_ ?_ ?_ ?_
426426
haveI : IsModuleTopology R R := inferInstance

FLT/MathlibExperiments/FrobeniusRiou.lean

+77-9
Original file line numberDiff line numberDiff line change
@@ -1,13 +1,54 @@
1-
-- Frobenius elements following Joel Riou's idea to use a very general lemma
2-
-- from Bourbaki comm alg
3-
-- (Théorème 2 in section 2 of chapter 5 of Bourbaki Alg Comm)
4-
-- (see https://leanprover.zulipchat.com/#narrow/stream/416277-FLT/topic/Outstanding.20Tasks.2C.20V4/near/449562398)
5-
6-
1+
/-
2+
Copyright (c) 2024 Jou Glasheen. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Jou Glasheen, Amelia Livingston, Jujian Zhang, Kevin Buzzard
5+
-/
6+
import Mathlib.FieldTheory.Cardinality
7+
import Mathlib.RingTheory.DedekindDomain.Ideal
8+
import Mathlib.RingTheory.Ideal.Pointwise
9+
import Mathlib.RingTheory.Ideal.QuotientOperations
10+
import Mathlib.RingTheory.IntegralClosure.IntegralRestrict
711
import Mathlib.RingTheory.Ideal.Pointwise
812
import Mathlib.RingTheory.Ideal.Over
913
import Mathlib.FieldTheory.Normal
1014
import Mathlib
15+
import Mathlib.RingTheory.OreLocalization.Ring
16+
17+
/-!
18+
19+
# Frobenius elements.
20+
21+
This file proves a general result in commutative algebra which can be used to define Frobenius
22+
elements of Galois groups of local or fields (for example number fields).
23+
24+
KB was alerted to this very general result (which needs no Noetherian or finiteness assumptions
25+
on the rings, just on the Galois group) by Joel Riou
26+
here https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/Uses.20of.20Frobenius.20elements.20in.20mathematics/near/448934112
27+
28+
## Mathematical details
29+
30+
### The general commutative algebra result
31+
32+
This is Theorem 2 in Chapter V, Section 2, number 2 of Bourbaki Alg Comm. It says the following.
33+
Let `A → B` be commutative rings and let `G` be a finite group acting on `B` by ring automorphisms,
34+
such that the `G`-invariants of `B` are exactly the image of `A`.
35+
36+
The two claims of the theorem are:
37+
(i) If `P` is a prime ideal of `A` then `G` acts transitively on the primes of `B` above `P`.
38+
(ii) If `Q` is a prime ideal of `B` lying over `P` then the canonica map from the stabilizer of `Q`
39+
in `G` to the automorphism group of the residue field extension `Frac(B/Q) / Frac(A/P)` is
40+
surjective.
41+
42+
We say "automorphism group" rather than "Galois group" because the extension might not be
43+
separable (it is algebraic and normal however, but it can be infinite; however its maximal
44+
separable subextension is finite).
45+
46+
This result means that if the inertia group I_Q is trivial and the residue fields are finite,
47+
there's a Frobenius element Frob_Q in the stabiliser of Q, and a Frobenius conjugacy class
48+
Frob_P in G.
49+
50+
-/
51+
1152

1253
variable {A : Type*} [CommRing A]
1354
{B : Type*} [CommRing B] [Algebra A B] [Algebra.IsIntegral A B]
@@ -80,11 +121,11 @@ theorem part_a
80121

81122
end part_a
82123

83-
section normal_elements
124+
-- section normal_elements
84125

85-
variable (K : Type*) [Field K] {L : Type*} [Field L] [Algebra K L]
126+
-- variable (K : Type*) [Field K] {L : Type*} [Field L] [Algebra K L]
86127

87-
open Polynomial
128+
-- open Polynomial
88129

89130
-- I've commented out the next section because ACL suggested
90131
-- reading up-to-date Bourbaki here https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/poly.20whose.20roots.20are.20the.20products.20of.20the.20roots.20of.20polys/near/468585267
@@ -327,6 +368,33 @@ theorem main_result : Function.Surjective
327368
(baz2 Q P L K : MulAction.stabilizer G Q → (L ≃ₐ[K] L)) := by
328369
sorry
329370

371+
section localization
372+
373+
/-
374+
375+
In this section we reduce to the case where P and Q are maximal.
376+
More precisely, we set `S := A - P` and localize everything in sight by `S`
377+
We then construct a group isomophism
378+
`MulAction.stabilizer G Q ≃ MulAction.stabilizer G SQ` where `SQ` is the prime ideal of `S⁻¹B`
379+
obtained by localizing `Q` at `S`, and show that it commutes with the maps currently called
380+
`baz2 Q P L K` and `baz2 SQ SP L K`.
381+
382+
-/
383+
384+
abbrev S := P.primeCompl
385+
abbrev SA := A[(S P)⁻¹]
386+
387+
abbrev SB := B[(S P)⁻¹]
388+
389+
-- Currently stuck here
390+
--instance : Algebra (A[(S P)⁻¹]) (B[(S P)⁻¹]) where
391+
-- sorry--__ := OreLocalization.instModule (R := A) (X := B) (S := P.primeCompl)
392+
393+
/-
394+
failed to synthesize
395+
Semiring (OreLocalization (S P) B)
396+
-/
397+
end localization
330398

331399
-- In Frobenius2.lean in this dir (Jou's FM24 project) there's a proof of surjectivity
332400
-- assuming B/Q is finite and P is maximal.

0 commit comments

Comments
 (0)