Skip to content

Commit 05a0e5e

Browse files
create M_spec_map (#152)
1 parent ea954b8 commit 05a0e5e

File tree

1 file changed

+4
-27
lines changed

1 file changed

+4
-27
lines changed

FLT/MathlibExperiments/FrobeniusRiou.lean

+4-27
Original file line numberDiff line numberDiff line change
@@ -271,36 +271,13 @@ theorem M_spec (b : B) : ((M hFull b : A[X]) : B[X]) = F G b := by
271271
simp_rw [finset_sum_coeff, ← lcoeff_apply, lcoeff_apply, coeff_monomial]
272272
aesop
273273

274-
/-
275-
private theorem F_descent_monic
276-
(hFull : ∀ (b : B), (∀ (g : G), g • b = b) → ∃ a : A, b = a) (b : B) :
277-
∃ M : A[X], (M : B[X]) = F G b ∧ Monic M := by
278-
have : F G b ∈ Polynomial.lifts (algebraMap A B) := by
279-
choose M hM using F_descent hFull b
280-
use M; exact hM
281-
choose M hM using lifts_and_degree_eq_and_monic this (F_monic b)
282-
use M
283-
exact ⟨hM.1, hM.2.2⟩
284-
285-
variable (G) in
286-
noncomputable def M [Finite G] (b : B) : A[X] := (F_descent_monic hFull b).choose
287-
288-
theorem M_spec (b : B) : ((M G hFull b : A[X]) : B[X]) = F G b :=
289-
(F_descent_monic hFull b).choose_spec.1
290-
291-
theorem M_spec' (b : B) : (map (algebraMap A B) (M G hFull b)) = F G b :=
292-
(F_descent_monic hFull b).choose_spec.1
293-
294-
theorem M_monic (b : B) : (M G hFull b).Monic := (F_descent_monic hFull b).choose_spec.2
295-
-/
296-
297274
omit [Nontrivial B] in
298-
theorem coe_poly_as_map (p : A[X]) : (p : B[X]) = map (algebraMap A B) p := rfl
299-
275+
theorem M_spec_map (b : B) : (map (algebraMap A B) (M hFull b)) = F G b := by
276+
rw [← M_spec hFull b]; rfl
300277

301278
omit [Nontrivial B] in
302279
theorem M_eval_eq_zero (b : B) : (M hFull b).eval₂ (algebraMap A B) b = 0 := by
303-
rw [eval₂_eq_eval_map, ← coe_poly_as_map, M_spec, F_eval_eq_zero]
280+
rw [eval₂_eq_eval_map, M_spec_map, F_eval_eq_zero]
304281

305282
include hFull in
306283
theorem isIntegral : Algebra.IsIntegral A B where
@@ -740,7 +717,7 @@ theorem algebraic {A : Type*} [CommRing A] {B : Type*} [Nontrivial B] [CommRing
740717
exact M_monic hFull b
741718
. rw [← hb, algebraMap_cast, map_map, ← IsScalarTower.algebraMap_eq]
742719
rw [algebraMap_algebraMap, aeval_def, eval₂_eq_eval_map, map_map, ← IsScalarTower.algebraMap_eq]
743-
rw [IsScalarTower.algebraMap_eq A B L, ← map_map, ← coe_poly_as_map (M hFull b), M_spec]
720+
rw [IsScalarTower.algebraMap_eq A B L, ← map_map, M_spec_map]
744721
rw [eval_map, eval₂_hom, F_eval_eq_zero]
745722
exact algebraMap.coe_zero
746723

0 commit comments

Comments
 (0)