@@ -115,24 +115,30 @@ theorem integralNormalization_degree : (integralNormalization p).degree = p.degr
115
115
· rw [← degree_scaleRoots, ← integralNormalization_mul_C_leadingCoeff]
116
116
exact (degree_mul_le _ _).trans (add_le_of_nonpos_right degree_C_le)
117
117
118
- variable [CommSemiring S]
118
+ variable {A : Type *} [CommSemiring S] [Semiring A ]
119
119
120
120
theorem leadingCoeff_smul_integralNormalization (p : S[X]) :
121
121
p.leadingCoeff • integralNormalization p = scaleRoots p p.leadingCoeff := by
122
122
rw [Algebra.smul_def, algebraMap_eq, mul_comm, integralNormalization_mul_C_leadingCoeff]
123
123
124
- theorem integralNormalization_eval₂_leadingCoeff_mul (h : 1 ≤ p.natDegree) (f : R →+* S) (x : S) :
124
+ theorem integralNormalization_eval₂_leadingCoeff_mul_of_commute (h : 1 ≤ p.natDegree) (f : R →+* A)
125
+ (x : A) (h₁ : Commute (f p.leadingCoeff) x) (h₂ : ∀ {r r'}, Commute (f r) (f r')) :
125
126
(integralNormalization p).eval₂ f (f p.leadingCoeff * x) =
126
127
f p.leadingCoeff ^ (p.natDegree - 1 ) * p.eval₂ f x := by
127
128
rw [eval₂_eq_sum_range, eval₂_eq_sum_range, Finset.mul_sum]
128
129
apply Finset.sum_congr
129
130
· rw [natDegree_eq_of_degree_eq p.integralNormalization_degree]
130
131
intro n _hn
131
- rw [mul_pow, ← mul_assoc, ← f.map_pow, ← f.map_mul,
132
- integralNormalization_coeff_mul_leadingCoeff_pow _ h, f.map_mul, f.map_pow]
133
- ring
132
+ rw [h₁.mul_pow, ← mul_assoc, ← f.map_pow, ← f.map_mul,
133
+ integralNormalization_coeff_mul_leadingCoeff_pow _ h, f.map_mul, h₂.eq, f.map_pow, mul_assoc]
134
134
135
- theorem integralNormalization_eval₂_eq_zero {p : R[X]} (f : R →+* S) {z : S} (hz : eval₂ f z p = 0 )
135
+ theorem integralNormalization_eval₂_leadingCoeff_mul (h : 1 ≤ p.natDegree) (f : R →+* S) (x : S) :
136
+ (integralNormalization p).eval₂ f (f p.leadingCoeff * x) =
137
+ f p.leadingCoeff ^ (p.natDegree - 1 ) * p.eval₂ f x :=
138
+ integralNormalization_eval₂_leadingCoeff_mul_of_commute h _ _ (.all _ _) (.all _ _)
139
+
140
+ theorem integralNormalization_eval₂_eq_zero_of_commute {p : R[X]} (f : R →+* A) {z : A}
141
+ (hz : eval₂ f z p = 0 ) (h₁ : Commute (f p.leadingCoeff) z) (h₂ : ∀ {r r'}, Commute (f r) (f r'))
136
142
(inj : ∀ x : R, f x = 0 → x = 0 ) :
137
143
eval₂ f (f p.leadingCoeff * z) (integralNormalization p) = 0 := by
138
144
obtain (h | h) := p.natDegree.eq_zero_or_pos
@@ -141,20 +147,20 @@ theorem integralNormalization_eval₂_eq_zero {p : R[X]} (f : R →+* S) {z : S}
141
147
simp [h0]
142
148
· rw [eq_C_of_natDegree_eq_zero h, eval₂_C] at hz
143
149
exact absurd (inj _ hz) h0
144
- · rw [integralNormalization_eval₂_leadingCoeff_mul h , hz, mul_zero]
150
+ · rw [integralNormalization_eval₂_leadingCoeff_mul_of_commute h _ _ h₁ h₂ , hz, mul_zero]
145
151
146
- end Semiring
147
-
148
- section CommSemiring
149
-
150
- variable [CommSemiring R] [CommSemiring S]
152
+ theorem integralNormalization_eval₂_eq_zero {p : R[X]} (f : R →+* S) {z : S} (hz : eval₂ f z p = 0 )
153
+ (inj : ∀ x : R, f x = 0 → x = 0 ) :
154
+ eval₂ f (f p.leadingCoeff * z) (integralNormalization p) = 0 :=
155
+ integralNormalization_eval₂_eq_zero_of_commute _ hz (.all _ _) (.all _ _) inj
151
156
152
- theorem integralNormalization_aeval_eq_zero [Algebra R S] {f : R[X]} {z : S} (hz : aeval z f = 0 )
153
- (inj : ∀ x : R, algebraMap R S x = 0 → x = 0 ) :
154
- aeval (algebraMap R S f.leadingCoeff * z) (integralNormalization f) = 0 :=
155
- integralNormalization_eval₂_eq_zero (algebraMap R S) hz inj
157
+ theorem integralNormalization_aeval_eq_zero [Algebra S A] {f : S[X]} {z : A} (hz : aeval z f = 0 )
158
+ (inj : ∀ x : S, algebraMap S A x = 0 → x = 0 ) :
159
+ aeval (algebraMap S A f.leadingCoeff * z) (integralNormalization f) = 0 :=
160
+ integralNormalization_eval₂_eq_zero_of_commute (algebraMap S A) hz
161
+ (Algebra.commute_algebraMap_left _ _) (.map (.all _ _) _) inj
156
162
157
- end CommSemiring
163
+ end Semiring
158
164
159
165
section IsCancelMulZero
160
166
0 commit comments