Skip to content

Commit 8e6d79b

Browse files
Prove F_degree, M_deg_eq_F_deg, M_monic
Fixes #141
1 parent d56dd12 commit 8e6d79b

File tree

2 files changed

+34
-26
lines changed

2 files changed

+34
-26
lines changed

FLT/MathlibExperiments/FrobeniusRiou.lean

+28-25
Original file line numberDiff line numberDiff line change
@@ -139,13 +139,13 @@ theorem F_monic [Nontrivial B] (b : B) : (F G b).Monic := by
139139
rw [Monic, F_spec, finprod_eq_prod_of_fintype, leadingCoeff_prod'] <;> simp
140140

141141
variable (G) in
142-
theorem F_degree (b : B) [Nontrivial B] : (F G b).degree = Nat.card G := by
143-
unfold F
144-
-- need (∏ᶠ fᵢ).degree = ∑ᶠ (fᵢ.degree)
145-
-- then it should be easy
146-
sorry
142+
theorem F_degree [Nontrivial B] [NoZeroDivisors B] (b : B) : (F G b).degree = Nat.card G := by
143+
have := Fintype.ofFinite G
144+
rw [F_spec, finprod_eq_prod_of_fintype, Polynomial.degree_prod]
145+
simp only [degree_X_sub_C, Finset.sum_const, Finset.card_univ, Fintype.card_eq_nat_card,
146+
nsmul_eq_mul, mul_one]
147147

148-
theorem F_natdegree [Nontrivial B] (b : B) : (F G b).natDegree = Nat.card G := by
148+
theorem F_natdegree [Nontrivial B] [NoZeroDivisors B] (b : B) : (F G b).natDegree = Nat.card G := by
149149
rw [← degree_eq_iff_natDegree_eq_of_pos Nat.card_pos]
150150
exact F_degree G b
151151

@@ -220,7 +220,8 @@ theorem M_deg_le (b : B) : (M hFull b).degree ≤ (F G b).degree := by
220220
apply le_trans (degree_monomial_le n _) ?_
221221
exact le_degree_of_mem_supp n hn
222222

223-
variable [Nontrivial B] [Finite G]
223+
section
224+
variable [Nontrivial B] [NoZeroDivisors B] [Finite G]
224225

225226
theorem M_coeff_card (b : B) :
226227
(M hFull b).coeff (Nat.card G) = 1 := by
@@ -242,12 +243,15 @@ theorem M_coeff_card (b : B) :
242243
exact coeff_monomial_of_ne (splitting_of_full hFull ((F G b).coeff d)) hd
243244

244245
-- **IMPORTANT** `M` should be refactored before proving this. See commented out code below.
245-
theorem M_deg_eq_F_deg (b : B) : (M hFull b).degree = (F G b).degree := by
246+
theorem M_deg_eq_F_deg [Nontrivial A] (b : B) : (M hFull b).degree = (F G b).degree := by
246247
apply le_antisymm (M_deg_le hFull b)
247-
-- hopefully not too hard from previous two lemmas
248-
sorry
248+
rw [F_degree]
249+
have := M_coeff_card hFull b
250+
refine le_degree_of_ne_zero ?h
251+
rw [this]
252+
exact one_ne_zero
249253

250-
theorem M_deg (b : B) : (M hFull b).degree = Nat.card G := by
254+
theorem M_deg [Nontrivial A] (b : B) : (M hFull b).degree = Nat.card G := by
251255
rw [M_deg_eq_F_deg hFull b]
252256
exact F_degree G b
253257

@@ -259,12 +263,11 @@ theorem M_monic (b : B) : (M hFull b).Monic := by
259263
rw [← F_natdegree b] at this2 this3
260264
-- then the hypos say deg(M)<=n, coefficient of X^n is 1 in M
261265
have this4 : (M hFull b).natDegree ≤ (F G b).natDegree := natDegree_le_natDegree this1
262-
clear this1
263-
-- Now it's just a logic puzzle. If deg(F)=n>0 then we
264-
-- know deg(M)<=n and the coefficient of X^n is 1 in M
265-
sorry
266+
exact Polynomial.monic_of_natDegree_le_of_coeff_eq_one _ this4 this2
267+
end
268+
269+
variable [Finite G]
266270

267-
omit [Nontrivial B] in
268271
theorem M_spec (b : B) : ((M hFull b : A[X]) : B[X]) = F G b := by
269272
unfold M
270273
ext N
@@ -294,18 +297,16 @@ theorem M_spec' (b : B) : (map (algebraMap A B) (M G hFull b)) = F G b :=
294297
(F_descent_monic hFull b).choose_spec.1
295298
296299
theorem M_monic (b : B) : (M G hFull b).Monic := (F_descent_monic hFull b).choose_spec.2
297-
--/
300+
-/
298301

299-
omit [Nontrivial B] in
300302
theorem coe_poly_as_map (p : A[X]) : (p : B[X]) = map (algebraMap A B) p := rfl
301303

302304

303-
omit [Nontrivial B] in
304305
theorem M_eval_eq_zero (b : B) : (M hFull b).eval₂ (algebraMap A B) b = 0 := by
305306
rw [eval₂_eq_eval_map, ← coe_poly_as_map, M_spec, F_eval_eq_zero]
306307

307308
include hFull in
308-
theorem isIntegral : Algebra.IsIntegral A B where
309+
theorem isIntegral [Nontrivial B] [NoZeroDivisors B] : Algebra.IsIntegral A B where
309310
isIntegral b := ⟨M hFull b, M_monic hFull b, M_eval_eq_zero hFull b⟩
310311

311312
end full_descent
@@ -322,7 +323,8 @@ variable {A : Type*} [CommRing A]
322323
{B : Type*} [CommRing B] [Algebra A B] [Nontrivial B]
323324
{G : Type*} [Group G] [Finite G] [MulSemiringAction G B]
324325

325-
theorem isIntegral_of_Full (hFull : ∀ (b : B), (∀ (g : G), g • b = b) → ∃ a : A, b = a) :
326+
theorem isIntegral_of_Full [NoZeroDivisors B]
327+
(hFull : ∀ (b : B), (∀ (g : G), g • b = b) → ∃ a : A, b = a) :
326328
Algebra.IsIntegral A B where
327329
isIntegral b := ⟨M hFull b, M_monic hFull b, M_eval_eq_zero hFull b⟩
328330

@@ -341,8 +343,8 @@ theorem Nontrivial_of_exists_prime {R : Type*} [CommRing R]
341343
apply Subsingleton.elim
342344

343345
-- (Part a of Théorème 2 in section 2 of chapter 5 of Bourbaki Alg Comm)
344-
omit [Nontrivial B] in
345-
theorem part_a [SMulCommClass G A B]
346+
omit [Nontrivial B] in
347+
theorem part_a [SMulCommClass G A B] [NoZeroDivisors B]
346348
(hPQ : Ideal.comap (algebraMap A B) P = Ideal.comap (algebraMap A B) Q)
347349
(hFull' : ∀ (b : B), (∀ (g : G), g • b = b) → ∃ a : A, b = a) :
348350
∃ g : G, Q = g • P := by
@@ -710,8 +712,9 @@ theorem Algebra.isAlgebraic_of_subring_isAlgebraic {R k K : Type*} [CommRing R]
710712
-- this uses `Algebra.isAlgebraic_of_subring_isAlgebraic` but I think we're going to have
711713
-- to introduce `f` anyway because we need not just that the extension is algebraic but
712714
-- that every element satisfies a poly of degree <= |G|.
713-
theorem algebraic {A : Type*} [CommRing A] {B : Type*} [Nontrivial B] [CommRing B] [Algebra A B]
714-
[Algebra.IsIntegral A B] {G : Type*} [Group G] [Finite G] [MulSemiringAction G B] (Q : Ideal B)
715+
theorem algebraic {A : Type*} [CommRing A] {B : Type*} [Nontrivial B] [CommRing B]
716+
[NoZeroDivisors B] [Algebra A B] [Algebra.IsIntegral A B]
717+
{G : Type*} [Group G] [Finite G] [MulSemiringAction G B] (Q : Ideal B)
715718
[Q.IsPrime] (P : Ideal A) [P.IsPrime] [Algebra (A ⧸ P) (B ⧸ Q)]
716719
[IsScalarTower A (A ⧸ P) (B ⧸ Q)] (L : Type*) [Field L] [Algebra (B ⧸ Q) L]
717720
[IsFractionRing (B ⧸ Q) L] [Algebra (A ⧸ P) L] [IsScalarTower (A ⧸ P) (B ⧸ Q) L]

blueprint/src/chapter/FrobeniusProject.tex

+6-1
Original file line numberDiff line numberDiff line change
@@ -174,17 +174,22 @@ \section{The extension $B/A$.}
174174
\label{MulSemiringAction.CharacteristicPolynomial.M_deg}
175175
\lean{MulSemiringAction.CharacteristicPolynomial.M_deg}
176176
\uses{MulSemiringAction.CharacteristicPolynomial.M}
177+
\leanok
177178
$M_b$ has degree $n$.
178179
\end{lemma}
179-
\begin{proof} Exercise.
180+
\begin{proof}
181+
\leanok
182+
Exercise.
180183
\end{proof}
181184
\begin{lemma}
182185
\lean{MulSemiringAction.CharacteristicPolynomial.M_monic}
183186
\label{MulSemiringAction.CharacteristicPolynomial.M_monic}
184187
\uses{MulSemiringAction.CharacteristicPolynomial.M}
188+
\leanok
185189
$M_b$ is monic.
186190
\end{lemma}
187191
\begin{proof}
192+
\leanok
188193
Exercise.
189194
\end{proof}
190195

0 commit comments

Comments
 (0)