Skip to content

Commit cb1b3f8

Browse files
committed
add more details to blueprint about Frobenius project
1 parent af4548c commit cb1b3f8

File tree

3 files changed

+48
-68
lines changed

3 files changed

+48
-68
lines changed

FLT/MathlibExperiments/Frobenius2.lean

-57
Original file line numberDiff line numberDiff line change
@@ -290,60 +290,3 @@ lemma Frob_Q : Frob A Q isGalois P • Q = Q := by
290290
have this := Q.sub_mem hy2 <| Frob_spec A Q isGalois P
291291
simp only [sub_sub_cancel] at this
292292
exact y_not_in_Q A Q <| Ideal.IsPrime.mem_of_pow_mem (show Q.IsPrime by infer_instance) _ this
293-
294-
lemma _root_.Ideal.Quotient.coe_eq_coe_iff_sub_mem {R : Type*} [CommRing R] {I : Ideal R} (x y : R) :
295-
(x : R ⧸ I) = y ↔ x - y ∈ I := Ideal.Quotient.mk_eq_mk_iff_sub_mem _ _
296-
297-
lemma coething (A B : Type*) [CommSemiring A] [Ring B] [Algebra A B] (a : A) (n : ℕ) :
298-
(a ^ n : A) = (a : B) ^ n := by
299-
--change algebraMap A B a ^ n = algebraMap A B (a ^ n)
300-
--symm
301-
exact map_pow _ _ _
302-
-- simp only [map_pow]
303-
304-
--attribute [norm_cast] map_pow
305-
306-
#exit
307-
308-
lemma Frob_Q_eq_pow_card (x : B) : Frob A Q isGalois P x - x^(Fintype.card (A⧸P)) ∈ Q := by
309-
by_cases hx : x ∈ Q
310-
· refine Q.sub_mem ?_ (Q.pow_mem_of_mem hx _ Fintype.card_pos)
311-
nth_rw 2 [← Frob_Q A Q isGalois P]
312-
change (Frob A Q isGalois P) • x ∈ _
313-
rw [Ideal.map_eq_comap_symm, Ideal.mem_comap]
314-
convert hx
315-
exact inv_smul_smul _ _
316-
·
317-
have foo : (x : B ⧸ Q) ≠ 0 :=
318-
mt (fun h ↦ (Submodule.Quotient.mk_eq_zero Q).mp h) hx
319-
let foobar : Field (B⧸Q) := ((Ideal.Quotient.maximal_ideal_iff_isField_quotient Q).mp ‹_›).toField
320-
let xbar : (B ⧸ Q)ˣ := Units.mk0 (x : B ⧸ Q) foo
321-
obtain ⟨n, (hn : g Q ^ n = xbar)⟩ := g_spec Q xbar
322-
have hn2 : (g Q : B ⧸ Q) ^ n = xbar := by exact_mod_cast hn
323-
change _ = (x : B ⧸ Q) at hn2
324-
rw [← Ideal.Quotient.mk_eq_mk_iff_sub_mem]
325-
change ((Frob A Q isGalois P) x : B ⧸ Q) = x ^ Fintype.card (A ⧸ P)
326-
rw [← hn2]
327-
have fact1 := Frob_spec A Q isGalois P
328-
have fact2 : y A Q = (g Q : B⧸Q) := y_mod_Q A Q
329-
rw [← fact2]
330-
rw [← Ideal.Quotient.coe_eq_coe_iff_sub_mem] at fact1
331-
push_cast at fact1
332-
rw [pow_right_comm]
333-
rw [← fact1]
334-
norm_cast
335-
rw [Ideal.Quotient.coe_eq_coe_iff_sub_mem]
336-
have fact3 := Frob_Q A Q isGalois P
337-
rw [← smul_pow']
338-
change (Frob A Q isGalois P) • x - _ ∈ _
339-
rw [← smul_sub]
340-
nth_rw 3 [ ← fact3]
341-
suffices (x - y A Q ^ n) ∈ Q by
342-
exact? says exact Ideal.smul_mem_pointwise_smul_iff.mpr this
343-
sorry
344-
345-
/- maths proof:
346-
347-
2) σ is x ^ #A/P mod Q
348-
3) Application to number fields
349-
-/

FLT/MathlibExperiments/FrobeniusRiou.lean

+1-3
Original file line numberDiff line numberDiff line change
@@ -183,7 +183,7 @@ section full_descent
183183

184184
variable (hFull : ∀ (b : B), (∀ (g : G), g • b = b) → ∃ a : A, b = a)
185185

186-
-- **IMPORTANT** The `splitting_of_full` approach is lousy and should be
186+
-- Remark: the `splitting_of_full` approach is lousy and should be
187187
-- replaced by the commented-out code below (lines 275-296 currently)
188188

189189
/-- This "splitting" function from B to A will only ever be evaluated on
@@ -241,7 +241,6 @@ theorem M_coeff_card (b : B) :
241241
· intro d _ hd
242242
exact coeff_monomial_of_ne (splitting_of_full hFull ((F G b).coeff d)) hd
243243

244-
-- **IMPORTANT** `M` should be refactored before proving this. See commented out code below.
245244
theorem M_deg_eq_F_deg [Nontrivial A] (b : B) : (M hFull b).degree = (F G b).degree := by
246245
apply le_antisymm (M_deg_le hFull b)
247246
rw [F_degree]
@@ -254,7 +253,6 @@ theorem M_deg [Nontrivial A] (b : B) : (M hFull b).degree = Nat.card G := by
254253
rw [M_deg_eq_F_deg hFull b]
255254
exact F_degree G b
256255

257-
-- **IMPORTANT** `M` should be refactored before proving this. See commented out code below.
258256
theorem M_monic (b : B) : (M hFull b).Monic := by
259257
have this1 := M_deg_le hFull b
260258
have this2 := M_coeff_card hFull b

blueprint/src/chapter/FrobeniusProject.tex

+47-8
Original file line numberDiff line numberDiff line change
@@ -340,30 +340,69 @@ \section{Proof of surjectivity.}
340340
with $M$ the maximal separable subextension of $L/K$.
341341
\end{definition}
342342

343+
NB we do use nonzeroness at some point, and $y$ can be zero in the case $L=K$
344+
(this seems to have been missed by Bourbaki).
345+
343346
Note that the existence of some $\lambda\in L$ with this property just comes from finite
344347
and separable implies simple; we can use the ``clear denominator'' technique introduced
345348
earlier to scale this by some nonzero $\alpha\in A$ into $B/Q$, as
346349
$K(\lambda)=K(\alpha\lambda)$.
347350

348-
Here is a slightly delicate result whose proof I haven't thought too hard about.
351+
Our next goal is the following result:
349352
\begin{theorem} There exists some $x\in B$ and $\alpha\in A$ with the following
350353
properties.
351354
\begin{enumerate}
352-
\item $x=\alpha y$ mod $Q$ and $\alpha$ isn't zero mod $Q$.
355+
\item $x=\alpha y$ mod $Q$ and $\alpha$ isn't zero mod $P$.
353356
\item $x\in Q'$ for all $Q'\not=Q$ in the $G$-orbit of $Q$.
354357
\end{enumerate}
355358
\end{theorem}
359+
360+
Bourbaki only prove this in the case that $P$ is maximal (and implicitly use $\alpha=1$),
361+
but this generalisation seems to be fine. To prove it, we need to talk a little about localization.
362+
363+
Let $S$ be the complement of $P$, so $S$ is a multiplicative subset of $A$.
364+
Write $A[1/S]$ for the localisation of $A$ at $S$, and write $B[1/S]$
365+
for the localisation of $B$ at (the image of) $S$. I suspect that this is the
366+
same as $B\otimes_AA[1/S]$.
367+
368+
\begin{lemma}
369+
The prime ideals of $B[1/S]$ over $P[1/S] \subseteq A[1/S]$ biject naturally with
370+
the prime ideals of $B$ over $P$. More precisely, the $B$-algebra $B[1/S]$
371+
gives a canonical map $B\to B[1/S]$ and hence a canonical map from prime ideals
372+
of $B[1/S]$ to prime ideals of $B$. The claim is that this map induces
373+
a bijection between the primes of $B[1/S]$ above $P[1/S]$ and the primes
374+
of $B$ above $P$.
375+
\end{lemma}
376+
\begin{proof}
377+
Hopefully this is in mathlib in some form already. In general $\Spec(B[1/S])$ is just the subset of $\Spec(B)$
378+
consisting of primes of $B$ which miss $S$ (i.e., whose intersection with $A$ is a subset of $P$).
379+
\end{proof}
380+
381+
\begin{lemma}
382+
The primes of $B[1/S]$ above $P[1/S]$ are all maximal.
383+
\end{lemma}
356384
\begin{proof}
357-
Idea. Localise away from P, then all the $Q_i$ are maximal, use CRT and then clear denominators.
385+
This follows from {\tt Algebra.IsIntegral.isField\_iff\_isField} and the fact
386+
that an ideal is maximal iff the quotient by it is a field.
358387
\end{proof}
359388

360-
We now choose some $x\in B[1/S]$ which is $y$ modulo $Q$ and $0$ modulo all the other
361-
primes of $B$ above $P$, and consider the monic degree $|G|$ polynomial $f$ in $K[X]$
362-
with $x$ and its conjugates as roots. If $\sigma\in\Aut_K(L)$ then $\sigma(\overline{x})$
389+
\begin{proof}(of theorem):
390+
Because all these ideals of $B[1/S]$ are maximal, they're pairwise coprime.
391+
So by the Chinese Remainder Theorem we can find an element of $B[1/S]$ which
392+
is equal to $y$ modulo $Q[1/S]$ and equal to $0$ modulo all the other primes.
393+
This element is of the form $x/\alpha$ for some $x\in B$ and $\alpha\in S$,
394+
and one now checks that everything works.
395+
\end{proof}
396+
397+
It is probably worth remarking that the rest of the surjectivity argument was done
398+
by Jou Glasheen in the special case of number fields, in the file {\tt Frobenius2.lean}
399+
400+
Briefly: we consider the monic degree $|G|$ polynomial $f_x$ in $K[X]$ or $\overline{M}_x$ in $(A/P)[X]$,
401+
which has $x$ and its G-conjugates as roots. If $\sigma\in\Aut_K(L)$ then $\sigma(\overline{x})$
363402
is a root of $f$ as $\sigma$ fixes $K$ pointwise. Hence $\sigma(\overline{x})=\overline{g(x)}$
364-
for some $g\in G$, and because $\sigma(\overline{x})\not=0$ we have $\overline{g(x)}\not=0$
403+
for some $g\in G$ (because we're over an integral domain), and because $\sigma(\overline{x})\not=0$ we have $\overline{g(x)}\not=0$
365404
and hence $g(x)\notin Q[1/S]$. Hence $x\notin g^{-1} Q[1/S]$ and thus $g^{-1}Q=Q$ and $g\in SD_Q$.
366405
Finally we have $\phi_g=\sigma$ on $K$ and on $y$, so they are equal on $M$ and hence on $L$ as
367406
$L/M$ is purely inseparable.
368407

369-
This part of the argument seems weak.
408+
TODO: break this up into smaller pieces.

0 commit comments

Comments
 (0)