Theory Winding_Number_Eval.Missing_Algebraic
section ‹Some useful lemmas in algebra›
theory Missing_Algebraic imports
"HOL-Computational_Algebra.Polynomial_Factorial"
"HOL-Computational_Algebra.Fundamental_Theorem_Algebra"
"HOL-Complex_Analysis.Complex_Analysis"
Missing_Topology
"Budan_Fourier.BF_Misc"
begin
subsection ‹Misc›
lemma poly_holomorphic_on[simp]: "(poly p) holomorphic_on s"
by (meson field_differentiable_def has_field_derivative_at_within holomorphic_onI poly_DERIV)
lemma order_zorder:
fixes p::"complex poly" and z::complex
assumes "p≠0"
shows "order z p = nat (zorder (poly p) z)"
proof -
define n where "n=nat (zorder (poly p) z)"
define h where "h=zor_poly (poly p) z"
have "∃w. poly p w ≠ 0" using assms poly_all_0_iff_0 by auto
then obtain r where "0 < r" "cball z r ⊆ UNIV" and
h_holo: "h holomorphic_on cball z r" and
poly_prod:"(∀w∈cball z r. poly p w = h w * (w - z) ^ n ∧ h w ≠ 0)"
using zorder_exist_zero[of "poly p" UNIV z,folded h_def] poly_holomorphic_on
unfolding n_def by auto
then have "h holomorphic_on ball z r"
and "(∀w∈ball z r. poly p w = h w * (w - z) ^ n)"
and "h z≠0"
by auto
then have "order z p = n" using ‹p≠0›
proof (induct n arbitrary:p h)
case 0
then have "poly p z=h z" using ‹r>0› by auto
then have "poly p z≠0" using ‹h z≠0› by auto
then show ?case using order_root by blast
next
case (Suc n)
define sn where "sn=Suc n"
define h' where "h'≡ λw. deriv h w * (w-z)+ sn * h w"
have "(poly p has_field_derivative poly (pderiv p) w) (at w)" for w
using poly_DERIV[of p w] .
moreover have "(poly p has_field_derivative (h' w)*(w-z)^n ) (at w)" when "w∈ball z r" for w
proof (subst DERIV_cong_ev[of w w "poly p" "λw. h w * (w - z) ^ Suc n" ],simp_all)
show "∀⇩F x in nhds w. poly p x = h x * ((x - z) * (x - z) ^ n)"
unfolding eventually_nhds using Suc(3) ‹w∈ball z r›
by (metis Elementary_Metric_Spaces.open_ball power_Suc)
next
have "(h has_field_derivative deriv h w) (at w)"
using ‹h holomorphic_on ball z r› ‹w∈ball z r› holomorphic_on_imp_differentiable_at
by (simp add: holomorphic_derivI)
then have "((λw. h w * ((w - z) ^ sn))
has_field_derivative h' w * (w - z) ^ (sn - 1)) (at w)"
unfolding h'_def
apply (auto intro!: derivative_eq_intros simp add:field_simps)
by (auto simp add:field_simps sn_def)
then show "((λw. h w * ((w - z) * (w - z) ^ n))
has_field_derivative h' w * (w - z) ^ n) (at w)"
unfolding sn_def by auto
qed
ultimately have "∀w∈ball z r. poly (pderiv p) w = h' w * (w - z) ^ n"
using DERIV_unique by blast
moreover have "h' holomorphic_on ball z r"
unfolding h'_def using ‹h holomorphic_on ball z r›
by (auto intro!: holomorphic_intros)
moreover have "h' z≠0" unfolding h'_def sn_def using ‹h z ≠ 0› of_nat_neq_0 by auto
moreover have "pderiv p ≠ 0"
proof
assume "pderiv p = 0"
obtain c where "p=[:c:]" using ‹pderiv p = 0› using pderiv_iszero by blast
then have "c=0"
using Suc(3)[rule_format,of z] ‹r>0› by auto
then show False using ‹p≠0› using ‹p=[:c:]› by auto
qed
ultimately have "order z (pderiv p) = n" by (auto elim: Suc.hyps)
moreover have "order z p ≠ 0"
using Suc(3)[rule_format,of z] ‹r>0› order_root ‹p≠0› by auto
ultimately show ?case using order_pderiv[OF ‹pderiv p ≠ 0›] by auto
qed
then show ?thesis unfolding n_def .
qed
lemma poly_field_differentiable_at[simp]:
"poly p field_differentiable (at x within s)"
using field_differentiable_at_within field_differentiable_def poly_DERIV by blast
lemma deriv_pderiv: "deriv (poly p) = poly (pderiv p)"
by (meson ext DERIV_imp_deriv poly_DERIV)
lemma poly_linepath_comp:
fixes a::"'a::{real_normed_vector,comm_semiring_0,real_algebra_1}"
shows "poly p o (linepath a b) = poly (p ∘⇩p [:a, b-a:]) o of_real"
by (force simp add: poly_pcompose linepath_def scaleR_conv_of_real algebra_simps)
subsection ‹roots / zeros of a univariate function›
definition roots_within::"('a ⇒ 'b::zero) ⇒ 'a set ⇒ 'a set" where
"roots_within f s = {x∈s. f x = 0}"
abbreviation roots::"('a ⇒ 'b::zero) ⇒ 'a set" where
"roots f ≡ roots_within f UNIV"
subsection ‹The argument principle specialised to polynomials.›
lemma argument_principle_poly:
assumes "p≠0" and valid:"valid_path g" and loop: "pathfinish g = pathstart g"
and no_proots:"path_image g ⊆ - proots p"
shows "contour_integral g (λx. deriv (poly p) x / poly p x) = 2 * of_real pi * 𝗂 *
(∑x∈proots p. winding_number g x * of_nat (order x p))"
proof -
have "contour_integral g (λx. deriv (poly p) x / poly p x) = 2 * of_real pi * 𝗂 *
(∑x | poly p x = 0. winding_number g x * of_int (zorder (poly p) x))"
apply (rule argument_principle[of UNIV "poly p" "{}" "λ_. 1" g,simplified,OF _ valid loop])
using no_proots[unfolded proots_def] by (auto simp add:poly_roots_finite[OF ‹p≠0›] )
also have "… = 2 * of_real pi * 𝗂 * (∑x∈proots p. winding_number g x * of_nat (order x p))"
proof -
have "nat (zorder (poly p) x) = order x p" when "x∈proots p" for x
using order_zorder[OF ‹p≠0›] that unfolding proots_def by auto
then show ?thesis unfolding proots_def
apply (auto intro!: sum.cong)
by (metis assms(1) nat_eq_iff2 of_nat_nat order_root)
qed
finally show ?thesis .
qed
end