@@ -139,13 +139,13 @@ theorem F_monic [Nontrivial B] (b : B) : (F G b).Monic := by
139
139
rw [Monic, F_spec, finprod_eq_prod_of_fintype, leadingCoeff_prod'] <;> simp
140
140
141
141
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]
147
147
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
149
149
rw [← degree_eq_iff_natDegree_eq_of_pos Nat.card_pos]
150
150
exact F_degree G b
151
151
@@ -220,7 +220,8 @@ theorem M_deg_le (b : B) : (M hFull b).degree ≤ (F G b).degree := by
220
220
apply le_trans (degree_monomial_le n _) ?_
221
221
exact le_degree_of_mem_supp n hn
222
222
223
- variable [Nontrivial B] [Finite G]
223
+ section
224
+ variable [Nontrivial B] [NoZeroDivisors B] [Finite G]
224
225
225
226
theorem M_coeff_card (b : B) :
226
227
(M hFull b).coeff (Nat.card G) = 1 := by
@@ -242,12 +243,15 @@ theorem M_coeff_card (b : B) :
242
243
exact coeff_monomial_of_ne (splitting_of_full hFull ((F G b).coeff d)) hd
243
244
244
245
-- **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
246
247
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
249
253
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
251
255
rw [M_deg_eq_F_deg hFull b]
252
256
exact F_degree G b
253
257
@@ -259,12 +263,11 @@ theorem M_monic (b : B) : (M hFull b).Monic := by
259
263
rw [← F_natdegree b] at this2 this3
260
264
-- then the hypos say deg(M)<=n, coefficient of X^n is 1 in M
261
265
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]
266
270
267
- omit [Nontrivial B] in
268
271
theorem M_spec (b : B) : ((M hFull b : A[X]) : B[X]) = F G b := by
269
272
unfold M
270
273
ext N
@@ -294,18 +297,16 @@ theorem M_spec' (b : B) : (map (algebraMap A B) (M G hFull b)) = F G b :=
294
297
(F_descent_monic hFull b).choose_spec.1
295
298
296
299
theorem M_monic (b : B) : (M G hFull b).Monic := (F_descent_monic hFull b).choose_spec.2
297
- -- /
300
+ -/
298
301
299
- omit [Nontrivial B] in
300
302
theorem coe_poly_as_map (p : A[X]) : (p : B[X]) = map (algebraMap A B) p := rfl
301
303
302
304
303
- omit [Nontrivial B] in
304
305
theorem M_eval_eq_zero (b : B) : (M hFull b).eval₂ (algebraMap A B) b = 0 := by
305
306
rw [eval₂_eq_eval_map, ← coe_poly_as_map, M_spec, F_eval_eq_zero]
306
307
307
308
include hFull in
308
- theorem isIntegral : Algebra.IsIntegral A B where
309
+ theorem isIntegral [Nontrivial B] [NoZeroDivisors B] : Algebra.IsIntegral A B where
309
310
isIntegral b := ⟨M hFull b, M_monic hFull b, M_eval_eq_zero hFull b⟩
310
311
311
312
end full_descent
@@ -322,7 +323,8 @@ variable {A : Type*} [CommRing A]
322
323
{B : Type *} [CommRing B] [Algebra A B] [Nontrivial B]
323
324
{G : Type *} [Group G] [Finite G] [MulSemiringAction G B]
324
325
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) :
326
328
Algebra.IsIntegral A B where
327
329
isIntegral b := ⟨M hFull b, M_monic hFull b, M_eval_eq_zero hFull b⟩
328
330
@@ -341,8 +343,8 @@ theorem Nontrivial_of_exists_prime {R : Type*} [CommRing R]
341
343
apply Subsingleton.elim
342
344
343
345
-- (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]
346
348
(hPQ : Ideal.comap (algebraMap A B) P = Ideal.comap (algebraMap A B) Q)
347
349
(hFull' : ∀ (b : B), (∀ (g : G), g • b = b) → ∃ a : A, b = a) :
348
350
∃ g : G, Q = g • P := by
@@ -710,8 +712,9 @@ theorem Algebra.isAlgebraic_of_subring_isAlgebraic {R k K : Type*} [CommRing R]
710
712
-- this uses `Algebra.isAlgebraic_of_subring_isAlgebraic` but I think we're going to have
711
713
-- to introduce `f` anyway because we need not just that the extension is algebraic but
712
714
-- 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)
715
718
[Q.IsPrime] (P : Ideal A) [P.IsPrime] [Algebra (A ⧸ P) (B ⧸ Q)]
716
719
[IsScalarTower A (A ⧸ P) (B ⧸ Q)] (L : Type *) [Field L] [Algebra (B ⧸ Q) L]
717
720
[IsFractionRing (B ⧸ Q) L] [Algebra (A ⧸ P) L] [IsScalarTower (A ⧸ P) (B ⧸ Q) L]
0 commit comments