@@ -90,6 +90,9 @@ variables {f}
90
90
91
91
instance adjoin_root.has_coe_t : has_coe_t R (adjoin_root f) := ⟨of f⟩
92
92
93
+ @[simp] lemma mk_eq_mk {g h : polynomial R} : mk f g = mk f h ↔ f ∣ g - h :=
94
+ ideal.quotient.eq.trans ideal.mem_span_singleton
95
+
93
96
@[simp] lemma mk_self : mk f f = 0 :=
94
97
quotient.sound' (mem_span_singleton.2 $ by simp)
95
98
@@ -205,6 +208,81 @@ end irreducible
205
208
206
209
section power_basis
207
210
211
+ variables [comm_ring R] {g : polynomial R}
212
+
213
+ lemma is_integral_root' (hg : g.monic) : is_integral R (root g) :=
214
+ ⟨g, hg, eval₂_root g⟩
215
+
216
+ /-- `adjoin_root.mod_by_monic_hom` sends the equivalence class of `f` mod `g` to `f %ₘ g`.
217
+
218
+ This is a well-defined right inverse to `adjoin_root.mk`, see `adjoin_root.mk_left_inverse`. -/
219
+ def mod_by_monic_hom [nontrivial R] (hg : g.monic) :
220
+ adjoin_root g →ₗ[R] polynomial R :=
221
+ (submodule.liftq _ (polynomial.mod_by_monic_hom hg)
222
+ (λ f (hf : f ∈ (ideal.span {g}).restrict_scalars R),
223
+ (mem_ker_mod_by_monic hg).mpr (ideal.mem_span_singleton.mp hf))).comp $
224
+ (submodule.quotient.restrict_scalars_equiv R (ideal.span {g} : ideal (polynomial R)))
225
+ .symm.to_linear_map
226
+
227
+ @[simp] lemma mod_by_monic_hom_mk [nontrivial R] (hg : g.monic) (f : polynomial R) :
228
+ mod_by_monic_hom hg (mk g f) = f %ₘ g := rfl
229
+
230
+ lemma mk_left_inverse [nontrivial R] (hg : g.monic) :
231
+ function.left_inverse (mk g) (mod_by_monic_hom hg) :=
232
+ λ f, induction_on g f $ λ f, begin
233
+ rw [mod_by_monic_hom_mk hg, mk_eq_mk, mod_by_monic_eq_sub_mul_div _ hg,
234
+ sub_sub_cancel_left, dvd_neg],
235
+ apply dvd_mul_right
236
+ end
237
+
238
+ lemma mk_surjective [nontrivial R] (hg : g.monic) : function.surjective (mk g) :=
239
+ (mk_left_inverse hg).surjective
240
+
241
+ /-- The elements `1, root g, ..., root g ^ (d - 1)` form a basis for `adjoin_root g`,
242
+ where `g` is a monic polynomial of degree `d`. -/
243
+ @[simps] def power_basis_aux' [nontrivial R] (hg : g.monic) :
244
+ basis (fin g.nat_degree) R (adjoin_root g) :=
245
+ basis.of_equiv_fun
246
+ { to_fun := λ f i, (mod_by_monic_hom hg f).coeff i,
247
+ inv_fun := λ c, mk g $ ∑ (i : fin g.nat_degree), monomial i (c i),
248
+ map_add' := λ f₁ f₂, funext $ λ i,
249
+ by simp only [(mod_by_monic_hom hg).map_add, coeff_add, pi.add_apply],
250
+ map_smul' := λ f₁ f₂, funext $ λ i,
251
+ by simp only [(mod_by_monic_hom hg).map_smul, coeff_smul, pi.smul_apply, ring_hom.id_apply],
252
+ left_inv := λ f, induction_on g f (λ f, eq.symm $ mk_eq_mk.mpr $
253
+ by { simp only [mod_by_monic_hom_mk, sum_mod_by_monic_coeff hg degree_le_nat_degree],
254
+ rw [mod_by_monic_eq_sub_mul_div _ hg, sub_sub_cancel],
255
+ exact dvd_mul_right _ _ }),
256
+ right_inv := λ x, funext $ λ i, begin
257
+ simp only [mod_by_monic_hom_mk],
258
+ rw [(mod_by_monic_eq_self_iff hg).mpr, finset_sum_coeff, finset.sum_eq_single i];
259
+ try { simp only [coeff_monomial, eq_self_iff_true, if_true] },
260
+ { intros j _ hj, exact if_neg (fin.coe_injective.ne hj) },
261
+ { intros, have := finset.mem_univ i, contradiction },
262
+ { refine (degree_sum_le _ _).trans_lt ((finset.sup_lt_iff _).mpr (λ j _, _)),
263
+ { exact bot_lt_iff_ne_bot.mpr (mt degree_eq_bot.mp hg.ne_zero) },
264
+ { refine (degree_monomial_le _ _).trans_lt _,
265
+ rw [degree_eq_nat_degree hg.ne_zero, with_bot.coe_lt_coe],
266
+ exact j.2 } },
267
+ end }
268
+
269
+ /-- The power basis `1, root g, ..., root g ^ (d - 1)` for `adjoin_root g`,
270
+ where `g` is a monic polynomial of degree `d`. -/
271
+ @[simps] def power_basis' [nontrivial R] (hg : g.monic) :
272
+ power_basis R (adjoin_root g) :=
273
+ { gen := root g,
274
+ dim := g.nat_degree,
275
+ basis := power_basis_aux' hg,
276
+ basis_eq_pow := λ i, begin
277
+ simp only [power_basis_aux', basis.coe_of_equiv_fun, linear_equiv.coe_symm_mk],
278
+ rw finset.sum_eq_single i,
279
+ { rw [function.update_same, monomial_one_right_eq_X_pow, (mk g).map_pow, mk_X] },
280
+ { intros j _ hj,
281
+ rw ← monomial_zero_right _,
282
+ convert congr_arg _ (function.update_noteq hj _ _) }, -- Fix `decidable_eq` mismatch
283
+ { intros, have := finset.mem_univ i, contradiction },
284
+ end }
285
+
208
286
variables [field K] {f : polynomial K}
209
287
210
288
lemma is_integral_root (hf : f ≠ 0 ) : is_integral K (root f) :=
0 commit comments