diff -r 2444c8b85aac -r 0c3dcb3a17f6 src/HOL/Analysis/Poly_Roots.thy --- a/src/HOL/Analysis/Poly_Roots.thy Wed Jan 23 17:20:35 2019 +0100 +++ b/src/HOL/Analysis/Poly_Roots.thy Thu Jan 24 00:28:29 2019 +0000 @@ -2,7 +2,7 @@ Ported from "hol_light/Multivariate/complexes.ml" by L C Paulson *) -section%important \Polynomial Functions: Extremal Behaviour and Root Counts\ +section \Polynomial Functions: Extremal Behaviour and Root Counts\ theory Poly_Roots imports Complex_Main @@ -10,11 +10,11 @@ subsection\Basics about polynomial functions: extremal behaviour and root counts\ -lemma%important sub_polyfun: +lemma sub_polyfun: fixes x :: "'a::{comm_ring,monoid_mult}" shows "(\i\n. a i * x^i) - (\i\n. a i * y^i) = (x - y) * (\jk= Suc j..n. a k * y^(k - Suc j) * x^j)" -proof%unimportant - +proof - have "(\i\n. a i * x^i) - (\i\n. a i * y^i) = (\i\n. a i * (x^i - y^i))" by (simp add: algebra_simps sum_subtractf [symmetric]) @@ -27,7 +27,7 @@ finally show ?thesis . qed -lemma%unimportant sub_polyfun_alt: +lemma sub_polyfun_alt: fixes x :: "'a::{comm_ring,monoid_mult}" shows "(\i\n. a i * x^i) - (\i\n. a i * y^i) = (x - y) * (\jkb. \z. (\i\n. c i * z^i) = (z-a) * (\ii\n. c i * a^i)" @@ -57,21 +57,21 @@ by (intro exI allI) qed -lemma%important polyfun_linear_factor_root: +lemma polyfun_linear_factor_root: fixes a :: "'a::{comm_ring,monoid_mult}" assumes "(\i\n. c i * a^i) = 0" shows "\b. \z. (\i\n. c i * z^i) = (z-a) * (\i b ==> norm(x) \ a ==> norm(x + y) \ b" +lemma adhoc_norm_triangle: "a + norm(y) \ b ==> norm(x) \ a ==> norm(x + y) \ b" by (metis norm_triangle_mono order.trans order_refl) -lemma%important polyfun_extremal_lemma: +proposition polyfun_extremal_lemma: fixes c :: "nat \ 'a::real_normed_div_algebra" assumes "e > 0" shows "\M. \z. M \ norm z \ norm(\i\n. c i * z^i) \ e * norm(z) ^ Suc n" -proof%unimportant (induction n) +proof (induction n) case 0 show ?case by (rule exI [where x="norm (c 0) / e"]) (auto simp: mult.commute pos_divide_le_eq assms) @@ -102,7 +102,7 @@ qed qed -lemma%unimportant norm_lemma_xy: assumes "\b\ + 1 \ norm(y) - a" "norm(x) \ a" shows "b \ norm(x + y)" +lemma norm_lemma_xy: assumes "\b\ + 1 \ norm(y) - a" "norm(x) \ a" shows "b \ norm(x + y)" proof - have "b \ norm y - norm x" using assms by linarith @@ -110,12 +110,12 @@ by (metis (no_types) add.commute norm_diff_ineq order_trans) qed -lemma%important polyfun_extremal: +proposition polyfun_extremal: fixes c :: "nat \ 'a::real_normed_div_algebra" assumes "\k. k \ 0 \ k \ n \ c k \ 0" shows "eventually (\z. norm(\i\n. c i * z^i) \ B) at_infinity" using assms -proof%unimportant (induction n) +proof (induction n) case 0 then show ?case by simp next @@ -149,12 +149,12 @@ qed qed -lemma%important polyfun_rootbound: +proposition polyfun_rootbound: fixes c :: "nat \ 'a::{comm_ring,real_normed_div_algebra}" assumes "\k. k \ n \ c k \ 0" shows "finite {z. (\i\n. c i * z^i) = 0} \ card {z. (\i\n. c i * z^i) = 0} \ n" using assms -proof%unimportant (induction n arbitrary: c) +proof (induction n arbitrary: c) case (Suc n) show ?case proof (cases "{z. (\i\Suc n. c i * z^i) = 0} = {}") case False @@ -182,17 +182,17 @@ qed simp qed simp -corollary%important (*FIX ME needs name *) +corollary (*FIX ME needs name *) fixes c :: "nat \ 'a::{comm_ring,real_normed_div_algebra}" assumes "\k. k \ n \ c k \ 0" shows polyfun_rootbound_finite: "finite {z. (\i\n. c i * z^i) = 0}" and polyfun_rootbound_card: "card {z. (\i\n. c i * z^i) = 0} \ n" using polyfun_rootbound [OF assms] by auto -lemma%important polyfun_finite_roots: +proposition polyfun_finite_roots: fixes c :: "nat \ 'a::{comm_ring,real_normed_div_algebra}" shows "finite {z. (\i\n. c i * z^i) = 0} \ (\k. k \ n \ c k \ 0)" -proof%unimportant (cases " \k\n. c k \ 0") +proof (cases " \k\n. c k \ 0") case True then show ?thesis by (blast intro: polyfun_rootbound_finite) next @@ -200,7 +200,7 @@ by (auto simp: infinite_UNIV_char_0) qed -lemma%unimportant polyfun_eq_0: +lemma polyfun_eq_0: fixes c :: "nat \ 'a::{comm_ring,real_normed_div_algebra}" shows "(\z. (\i\n. c i * z^i) = 0) \ (\k. k \ n \ c k = 0)" proof (cases "(\z. (\i\n. c i * z^i) = 0)") @@ -215,10 +215,10 @@ by auto qed -lemma%important polyfun_eq_const: +theorem polyfun_eq_const: fixes c :: "nat \ 'a::{comm_ring,real_normed_div_algebra}" shows "(\z. (\i\n. c i * z^i) = k) \ c 0 = k \ (\k. k \ 0 \ k \ n \ c k = 0)" -proof%unimportant - +proof - {fix z have "(\i\n. c i * z^i) = (\i\n. (if i = 0 then c 0 - k else c i) * z^i) + k" by (induct n) auto