| author | wenzelm | 
| Fri, 19 Jun 2015 20:43:34 +0200 | |
| changeset 60524 | ffc1ee11759c | 
| parent 60420 | 884f54e01427 | 
| child 61560 | 7c985fd653c5 | 
| child 61609 | 77b453bd616f | 
| permissions | -rw-r--r-- | 
| 60420 | 1 | section \<open>polynomial functions: extremal behaviour and root counts\<close> | 
| 56215 | 2 | |
| 3 | (* Author: John Harrison and Valentina Bruno | |
| 4 | Ported from "hol_light/Multivariate/complexes.ml" by L C Paulson | |
| 5 | *) | |
| 6 | ||
| 7 | theory PolyRoots | |
| 8 | imports Complex_Main | |
| 9 | ||
| 10 | begin | |
| 11 | ||
| 60420 | 12 | subsection\<open>Geometric progressions\<close> | 
| 56215 | 13 | |
| 14 | lemma setsum_gp_basic: | |
| 15 |   fixes x :: "'a::{comm_ring,monoid_mult}"
 | |
| 16 | shows "(1 - x) * (\<Sum>i\<le>n. x^i) = 1 - x^Suc n" | |
| 17 | by (simp only: one_diff_power_eq [of "Suc n" x] lessThan_Suc_atMost) | |
| 18 | ||
| 19 | lemma setsum_gp0: | |
| 59867 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
59615diff
changeset | 20 |   fixes x :: "'a::{comm_ring,division_ring}"
 | 
| 59615 
fdfdf89a83a6
A few new lemmas and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
58877diff
changeset | 21 | shows "(\<Sum>i\<le>n. x^i) = (if x = 1 then of_nat(n + 1) else (1 - x^Suc n) / (1 - x))" | 
| 
fdfdf89a83a6
A few new lemmas and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
58877diff
changeset | 22 | using setsum_gp_basic[of x n] | 
| 
fdfdf89a83a6
A few new lemmas and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
58877diff
changeset | 23 | by (simp add: real_of_nat_def mult.commute divide_simps) | 
| 
fdfdf89a83a6
A few new lemmas and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
58877diff
changeset | 24 | |
| 
fdfdf89a83a6
A few new lemmas and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
58877diff
changeset | 25 | lemma setsum_power_add: | 
| 
fdfdf89a83a6
A few new lemmas and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
58877diff
changeset | 26 |   fixes x :: "'a::{comm_ring,monoid_mult}"
 | 
| 
fdfdf89a83a6
A few new lemmas and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
58877diff
changeset | 27 | shows "(\<Sum>i\<in>I. x^(m+i)) = x^m * (\<Sum>i\<in>I. x^i)" | 
| 
fdfdf89a83a6
A few new lemmas and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
58877diff
changeset | 28 | by (simp add: setsum_right_distrib power_add) | 
| 56215 | 29 | |
| 30 | lemma setsum_power_shift: | |
| 31 |   fixes x :: "'a::{comm_ring,monoid_mult}"
 | |
| 32 | assumes "m \<le> n" | |
| 33 | shows "(\<Sum>i=m..n. x^i) = x^m * (\<Sum>i\<le>n-m. x^i)" | |
| 34 | proof - | |
| 35 | have "(\<Sum>i=m..n. x^i) = x^m * (\<Sum>i=m..n. x^(i-m))" | |
| 36 | by (simp add: setsum_right_distrib power_add [symmetric]) | |
| 57129 
7edb7550663e
introduce more powerful reindexing rules for big operators
 hoelzl parents: 
56215diff
changeset | 37 | also have "(\<Sum>i=m..n. x^(i-m)) = (\<Sum>i\<le>n-m. x^i)" | 
| 60420 | 38 | using \<open>m \<le> n\<close> by (intro setsum.reindex_bij_witness[where j="\<lambda>i. i - m" and i="\<lambda>i. i + m"]) auto | 
| 56215 | 39 | finally show ?thesis . | 
| 40 | qed | |
| 41 | ||
| 42 | lemma setsum_gp_multiplied: | |
| 43 |   fixes x :: "'a::{comm_ring,monoid_mult}"
 | |
| 44 | assumes "m \<le> n" | |
| 45 | shows "(1 - x) * (\<Sum>i=m..n. x^i) = x^m - x^Suc n" | |
| 46 | proof - | |
| 47 | have "(1 - x) * (\<Sum>i=m..n. x^i) = x^m * (1 - x) * (\<Sum>i\<le>n-m. x^i)" | |
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 48 | by (metis mult.assoc mult.commute assms setsum_power_shift) | 
| 56215 | 49 | also have "... =x^m * (1 - x^Suc(n-m))" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 50 | by (metis mult.assoc setsum_gp_basic) | 
| 56215 | 51 | also have "... = x^m - x^Suc n" | 
| 52 | using assms | |
| 53 | by (simp add: algebra_simps) (metis le_add_diff_inverse power_add) | |
| 54 | finally show ?thesis . | |
| 55 | qed | |
| 56 | ||
| 57 | lemma setsum_gp: | |
| 59867 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
59615diff
changeset | 58 |   fixes x :: "'a::{comm_ring,division_ring}"
 | 
| 56215 | 59 | shows "(\<Sum>i=m..n. x^i) = | 
| 60 | (if n < m then 0 | |
| 61 | else if x = 1 then of_nat((n + 1) - m) | |
| 62 | else (x^m - x^Suc n) / (1 - x))" | |
| 63 | using setsum_gp_multiplied [of m n x] | |
| 64 | apply (auto simp: real_of_nat_def) | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57129diff
changeset | 65 | by (metis eq_iff_diff_eq_0 mult.commute nonzero_divide_eq_eq) | 
| 56215 | 66 | |
| 67 | lemma setsum_gp_offset: | |
| 59867 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
59615diff
changeset | 68 |   fixes x :: "'a::{comm_ring,division_ring}"
 | 
| 56215 | 69 | shows "(\<Sum>i=m..m+n. x^i) = | 
| 70 | (if x = 1 then of_nat n + 1 else x^m * (1 - x^Suc n) / (1 - x))" | |
| 71 | using setsum_gp [of x m "m+n"] | |
| 72 | by (auto simp: power_add algebra_simps) | |
| 73 | ||
| 59615 
fdfdf89a83a6
A few new lemmas and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
58877diff
changeset | 74 | lemma setsum_gp_strict: | 
| 59867 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
59615diff
changeset | 75 |   fixes x :: "'a::{comm_ring,division_ring}"
 | 
| 59615 
fdfdf89a83a6
A few new lemmas and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
58877diff
changeset | 76 | shows "(\<Sum>i<n. x^i) = (if x = 1 then of_nat n else (1 - x^n) / (1 - x))" | 
| 
fdfdf89a83a6
A few new lemmas and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
58877diff
changeset | 77 | by (induct n) (auto simp: algebra_simps divide_simps) | 
| 
fdfdf89a83a6
A few new lemmas and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
58877diff
changeset | 78 | |
| 60420 | 79 | subsection\<open>Basics about polynomial functions: extremal behaviour and root counts.\<close> | 
| 56215 | 80 | |
| 81 | lemma sub_polyfun: | |
| 82 |   fixes x :: "'a::{comm_ring,monoid_mult}"
 | |
| 83 | shows "(\<Sum>i\<le>n. a i * x^i) - (\<Sum>i\<le>n. a i * y^i) = | |
| 84 | (x - y) * (\<Sum>j<n. \<Sum>k= Suc j..n. a k * y^(k - Suc j) * x^j)" | |
| 85 | proof - | |
| 86 | have "(\<Sum>i\<le>n. a i * x^i) - (\<Sum>i\<le>n. a i * y^i) = | |
| 87 | (\<Sum>i\<le>n. a i * (x^i - y^i))" | |
| 88 | by (simp add: algebra_simps setsum_subtractf [symmetric]) | |
| 89 | also have "... = (\<Sum>i\<le>n. a i * (x - y) * (\<Sum>j<i. y^(i - Suc j) * x^j))" | |
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 90 | by (simp add: power_diff_sumr2 ac_simps) | 
| 56215 | 91 | also have "... = (x - y) * (\<Sum>i\<le>n. (\<Sum>j<i. a i * y^(i - Suc j) * x^j))" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 92 | by (simp add: setsum_right_distrib ac_simps) | 
| 56215 | 93 | also have "... = (x - y) * (\<Sum>j<n. (\<Sum>i=Suc j..n. a i * y^(i - Suc j) * x^j))" | 
| 94 | by (simp add: nested_setsum_swap') | |
| 95 | finally show ?thesis . | |
| 96 | qed | |
| 97 | ||
| 98 | lemma sub_polyfun_alt: | |
| 99 |   fixes x :: "'a::{comm_ring,monoid_mult}"
 | |
| 100 | shows "(\<Sum>i\<le>n. a i * x^i) - (\<Sum>i\<le>n. a i * y^i) = | |
| 101 | (x - y) * (\<Sum>j<n. \<Sum>k<n-j. a (j+k+1) * y^k * x^j)" | |
| 102 | proof - | |
| 103 |   { fix j
 | |
| 104 | have "(\<Sum>k = Suc j..n. a k * y^(k - Suc j) * x^j) = | |
| 105 | (\<Sum>k <n - j. a (Suc (j + k)) * y^k * x^j)" | |
| 57129 
7edb7550663e
introduce more powerful reindexing rules for big operators
 hoelzl parents: 
56215diff
changeset | 106 | by (rule setsum.reindex_bij_witness[where i="\<lambda>i. i + Suc j" and j="\<lambda>i. i - Suc j"]) auto } | 
| 56215 | 107 | then show ?thesis | 
| 108 | by (simp add: sub_polyfun) | |
| 109 | qed | |
| 110 | ||
| 111 | lemma polyfun_linear_factor: | |
| 112 |   fixes a :: "'a::{comm_ring,monoid_mult}"
 | |
| 113 | shows "\<exists>b. \<forall>z. (\<Sum>i\<le>n. c i * z^i) = | |
| 114 | (z-a) * (\<Sum>i<n. b i * z^i) + (\<Sum>i\<le>n. c i * a^i)" | |
| 115 | proof - | |
| 116 |   { fix z
 | |
| 117 | have "(\<Sum>i\<le>n. c i * z^i) - (\<Sum>i\<le>n. c i * a^i) = | |
| 118 | (z - a) * (\<Sum>j<n. (\<Sum>k = Suc j..n. c k * a^(k - Suc j)) * z^j)" | |
| 119 | by (simp add: sub_polyfun setsum_left_distrib) | |
| 120 | then have "(\<Sum>i\<le>n. c i * z^i) = | |
| 121 | (z - a) * (\<Sum>j<n. (\<Sum>k = Suc j..n. c k * a^(k - Suc j)) * z^j) | |
| 122 | + (\<Sum>i\<le>n. c i * a^i)" | |
| 123 | by (simp add: algebra_simps) } | |
| 124 | then show ?thesis | |
| 125 | by (intro exI allI) | |
| 126 | qed | |
| 127 | ||
| 128 | lemma polyfun_linear_factor_root: | |
| 129 |   fixes a :: "'a::{comm_ring,monoid_mult}"
 | |
| 130 | assumes "(\<Sum>i\<le>n. c i * a^i) = 0" | |
| 131 | shows "\<exists>b. \<forall>z. (\<Sum>i\<le>n. c i * z^i) = (z-a) * (\<Sum>i<n. b i * z^i)" | |
| 132 | using polyfun_linear_factor [of c n a] assms | |
| 133 | by simp | |
| 134 | ||
| 135 | lemma adhoc_norm_triangle: "a + norm(y) \<le> b ==> norm(x) \<le> a ==> norm(x + y) \<le> b" | |
| 136 | by (metis norm_triangle_mono order.trans order_refl) | |
| 137 | ||
| 138 | lemma polyfun_extremal_lemma: | |
| 139 | fixes c :: "nat \<Rightarrow> 'a::real_normed_div_algebra" | |
| 140 | assumes "e > 0" | |
| 141 | shows "\<exists>M. \<forall>z. M \<le> norm z \<longrightarrow> norm(\<Sum>i\<le>n. c i * z^i) \<le> e * norm(z) ^ Suc n" | |
| 142 | proof (induction n) | |
| 143 | case 0 | |
| 144 | show ?case | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57129diff
changeset | 145 | by (rule exI [where x="norm (c 0) / e"]) (auto simp: mult.commute pos_divide_le_eq assms) | 
| 56215 | 146 | next | 
| 147 | case (Suc n) | |
| 148 | then obtain M where M: "\<forall>z. M \<le> norm z \<longrightarrow> norm (\<Sum>i\<le>n. c i * z^i) \<le> e * norm z ^ Suc n" .. | |
| 149 | show ?case | |
| 150 | proof (rule exI [where x="max 1 (max M ((e + norm(c(Suc n))) / e))"], clarify) | |
| 151 | fix z::'a | |
| 152 | assume "max 1 (max M ((e + norm (c (Suc n))) / e)) \<le> norm z" | |
| 153 | then have norm1: "0 < norm z" "M \<le> norm z" "(e + norm (c (Suc n))) / e \<le> norm z" | |
| 154 | by auto | |
| 155 | then have norm2: "(e + norm (c (Suc n))) \<le> e * norm z" "(norm z * norm z ^ n) > 0" | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57129diff
changeset | 156 | apply (metis assms less_divide_eq mult.commute not_le) | 
| 56215 | 157 | using norm1 apply (metis mult_pos_pos zero_less_power) | 
| 158 | done | |
| 159 | have "e * (norm z * norm z ^ n) + norm (c (Suc n) * (z * z ^ n)) = | |
| 160 | (e + norm (c (Suc n))) * (norm z * norm z ^ n)" | |
| 161 | by (simp add: norm_mult norm_power algebra_simps) | |
| 162 | also have "... \<le> (e * norm z) * (norm z * norm z ^ n)" | |
| 163 | using norm2 by (metis real_mult_le_cancel_iff1) | |
| 164 | also have "... = e * (norm z * (norm z * norm z ^ n))" | |
| 165 | by (simp add: algebra_simps) | |
| 166 | finally have "e * (norm z * norm z ^ n) + norm (c (Suc n) * (z * z ^ n)) | |
| 167 | \<le> e * (norm z * (norm z * norm z ^ n))" . | |
| 168 | then show "norm (\<Sum>i\<le>Suc n. c i * z^i) \<le> e * norm z ^ Suc (Suc n)" using M norm1 | |
| 169 | by (drule_tac x=z in spec) (auto simp: intro!: adhoc_norm_triangle) | |
| 170 | qed | |
| 171 | qed | |
| 172 | ||
| 173 | lemma norm_lemma_xy: "\<lbrakk>abs b + 1 \<le> norm(y) - a; norm(x) \<le> a\<rbrakk> \<Longrightarrow> b \<le> norm(x + y)" | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57129diff
changeset | 174 | by (metis abs_add_one_not_less_self add.commute diff_le_eq dual_order.trans le_less_linear | 
| 56215 | 175 | norm_diff_ineq) | 
| 176 | ||
| 177 | lemma polyfun_extremal: | |
| 178 | fixes c :: "nat \<Rightarrow> 'a::real_normed_div_algebra" | |
| 179 | assumes "\<exists>k. k \<noteq> 0 \<and> k \<le> n \<and> c k \<noteq> 0" | |
| 180 | shows "eventually (\<lambda>z. norm(\<Sum>i\<le>n. c i * z^i) \<ge> B) at_infinity" | |
| 181 | using assms | |
| 182 | proof (induction n) | |
| 183 | case 0 then show ?case | |
| 184 | by simp | |
| 185 | next | |
| 186 | case (Suc n) | |
| 187 | show ?case | |
| 188 | proof (cases "c (Suc n) = 0") | |
| 189 | case True | |
| 190 | with Suc show ?thesis | |
| 191 | by auto (metis diff_is_0_eq diffs0_imp_equal less_Suc_eq_le not_less_eq) | |
| 192 | next | |
| 193 | case False | |
| 194 | with polyfun_extremal_lemma [of "norm(c (Suc n)) / 2" c n] | |
| 195 | obtain M where M: "\<And>z. M \<le> norm z \<Longrightarrow> | |
| 196 | norm (\<Sum>i\<le>n. c i * z^i) \<le> norm (c (Suc n)) / 2 * norm z ^ Suc n" | |
| 197 | by auto | |
| 198 | show ?thesis | |
| 199 | unfolding eventually_at_infinity | |
| 200 | proof (rule exI [where x="max M (max 1 ((abs B + 1) / (norm (c (Suc n)) / 2)))"], clarsimp) | |
| 201 | fix z::'a | |
| 202 | assume les: "M \<le> norm z" "1 \<le> norm z" "(\<bar>B\<bar> * 2 + 2) / norm (c (Suc n)) \<le> norm z" | |
| 203 | then have "\<bar>B\<bar> * 2 + 2 \<le> norm z * norm (c (Suc n))" | |
| 204 | by (metis False pos_divide_le_eq zero_less_norm_iff) | |
| 205 | then have "\<bar>B\<bar> * 2 + 2 \<le> norm z ^ (Suc n) * norm (c (Suc n))" | |
| 60420 | 206 | by (metis \<open>1 \<le> norm z\<close> order.trans mult_right_mono norm_ge_zero self_le_power zero_less_Suc) | 
| 56215 | 207 | then show "B \<le> norm ((\<Sum>i\<le>n. c i * z^i) + c (Suc n) * (z * z ^ n))" using M les | 
| 208 | apply auto | |
| 209 | apply (rule norm_lemma_xy [where a = "norm (c (Suc n)) * norm z ^ (Suc n) / 2"]) | |
| 210 | apply (simp_all add: norm_mult norm_power) | |
| 211 | done | |
| 212 | qed | |
| 213 | qed | |
| 214 | qed | |
| 215 | ||
| 216 | lemma polyfun_rootbound: | |
| 217 |  fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
 | |
| 218 | assumes "\<exists>k. k \<le> n \<and> c k \<noteq> 0" | |
| 219 |    shows "finite {z. (\<Sum>i\<le>n. c i * z^i) = 0} \<and> card {z. (\<Sum>i\<le>n. c i * z^i) = 0} \<le> n"
 | |
| 220 | using assms | |
| 221 | proof (induction n arbitrary: c) | |
| 222 | case (Suc n) show ?case | |
| 223 |  proof (cases "{z. (\<Sum>i\<le>Suc n. c i * z^i) = 0} = {}")
 | |
| 224 | case False | |
| 225 | then obtain a where a: "(\<Sum>i\<le>Suc n. c i * a^i) = 0" | |
| 226 | by auto | |
| 227 | from polyfun_linear_factor_root [OF this] | |
| 228 | obtain b where "\<And>z. (\<Sum>i\<le>Suc n. c i * z^i) = (z - a) * (\<Sum>i< Suc n. b i * z^i)" | |
| 229 | by auto | |
| 230 | then have b: "\<And>z. (\<Sum>i\<le>Suc n. c i * z^i) = (z - a) * (\<Sum>i\<le>n. b i * z^i)" | |
| 231 | by (metis lessThan_Suc_atMost) | |
| 232 |    then have ins_ab: "{z. (\<Sum>i\<le>Suc n. c i * z^i) = 0} = insert a {z. (\<Sum>i\<le>n. b i * z^i) = 0}"
 | |
| 233 | by auto | |
| 234 | have c0: "c 0 = - (a * b 0)" using b [of 0] | |
| 235 | by simp | |
| 236 | then have extr_prem: "~ (\<exists>k\<le>n. b k \<noteq> 0) \<Longrightarrow> \<exists>k. k \<noteq> 0 \<and> k \<le> Suc n \<and> c k \<noteq> 0" | |
| 237 | by (metis Suc.prems le0 minus_zero mult_zero_right) | |
| 238 | have "\<exists>k\<le>n. b k \<noteq> 0" | |
| 239 | apply (rule ccontr) | |
| 240 | using polyfun_extremal [OF extr_prem, of 1] | |
| 241 | apply (auto simp: eventually_at_infinity b simp del: setsum_atMost_Suc) | |
| 242 | apply (drule_tac x="of_real ba" in spec, simp) | |
| 243 | done | |
| 244 | then show ?thesis using Suc.IH [of b] ins_ab | |
| 245 | by (auto simp: card_insert_if) | |
| 246 | qed simp | |
| 247 | qed simp | |
| 248 | ||
| 249 | corollary | |
| 250 |   fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
 | |
| 251 | assumes "\<exists>k. k \<le> n \<and> c k \<noteq> 0" | |
| 252 |     shows polyfun_rootbound_finite: "finite {z. (\<Sum>i\<le>n. c i * z^i) = 0}"
 | |
| 253 |       and polyfun_rootbound_card:   "card {z. (\<Sum>i\<le>n. c i * z^i) = 0} \<le> n"
 | |
| 254 | using polyfun_rootbound [OF assms] by auto | |
| 255 | ||
| 256 | lemma polyfun_finite_roots: | |
| 257 |   fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
 | |
| 258 |     shows  "finite {z. (\<Sum>i\<le>n. c i * z^i) = 0} \<longleftrightarrow> (\<exists>k. k \<le> n \<and> c k \<noteq> 0)"
 | |
| 259 | proof (cases " \<exists>k\<le>n. c k \<noteq> 0") | |
| 260 | case True then show ?thesis | |
| 261 | by (blast intro: polyfun_rootbound_finite) | |
| 262 | next | |
| 263 | case False then show ?thesis | |
| 264 | by (auto simp: infinite_UNIV_char_0) | |
| 265 | qed | |
| 266 | ||
| 267 | lemma polyfun_eq_0: | |
| 268 |   fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
 | |
| 269 | shows "(\<forall>z. (\<Sum>i\<le>n. c i * z^i) = 0) \<longleftrightarrow> (\<forall>k. k \<le> n \<longrightarrow> c k = 0)" | |
| 270 | proof (cases "(\<forall>z. (\<Sum>i\<le>n. c i * z^i) = 0)") | |
| 271 | case True | |
| 272 |   then have "~ finite {z. (\<Sum>i\<le>n. c i * z^i) = 0}"
 | |
| 273 | by (simp add: infinite_UNIV_char_0) | |
| 274 | with True show ?thesis | |
| 275 | by (metis (poly_guards_query) polyfun_rootbound_finite) | |
| 276 | next | |
| 277 | case False | |
| 278 | then show ?thesis | |
| 279 | by auto | |
| 280 | qed | |
| 281 | ||
| 282 | lemma polyfun_eq_const: | |
| 283 |   fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
 | |
| 284 | shows "(\<forall>z. (\<Sum>i\<le>n. c i * z^i) = k) \<longleftrightarrow> c 0 = k \<and> (\<forall>k. k \<noteq> 0 \<and> k \<le> n \<longrightarrow> c k = 0)" | |
| 285 | proof - | |
| 286 |   {fix z
 | |
| 287 | have "(\<Sum>i\<le>n. c i * z^i) = (\<Sum>i\<le>n. (if i = 0 then c 0 - k else c i) * z^i) + k" | |
| 288 | by (induct n) auto | |
| 289 | } then | |
| 290 | have "(\<forall>z. (\<Sum>i\<le>n. c i * z^i) = k) \<longleftrightarrow> (\<forall>z. (\<Sum>i\<le>n. (if i = 0 then c 0 - k else c i) * z^i) = 0)" | |
| 291 | by auto | |
| 292 | also have "... \<longleftrightarrow> c 0 = k \<and> (\<forall>k. k \<noteq> 0 \<and> k \<le> n \<longrightarrow> c k = 0)" | |
| 293 | by (auto simp: polyfun_eq_0) | |
| 294 | finally show ?thesis . | |
| 295 | qed | |
| 296 | ||
| 297 | end | |
| 298 |