| 56215 |      1 | header {* polynomial functions: extremal behaviour and root counts *}
 | 
|  |      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 | 
 | 
|  |     12 | subsection{*Geometric progressions*}
 | 
|  |     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:
 | 
|  |     20 |  fixes x :: "'a::{comm_ring,division_ring_inverse_zero}"
 | 
|  |     21 |  shows   "(\<Sum>i\<le>n. x^i) = (if x = 1 then of_nat(n + 1) else (1 - x^Suc n) / (1 - x))"
 | 
|  |     22 | using setsum_gp_basic[of x n]
 | 
|  |     23 | apply (simp add: real_of_nat_def)
 | 
|  |     24 | by (metis eq_iff_diff_eq_0 mult_commute nonzero_eq_divide_eq)
 | 
|  |     25 | 
 | 
|  |     26 | lemma setsum_power_shift:
 | 
|  |     27 |   fixes x :: "'a::{comm_ring,monoid_mult}"
 | 
|  |     28 |   assumes "m \<le> n"
 | 
|  |     29 |   shows "(\<Sum>i=m..n. x^i) = x^m * (\<Sum>i\<le>n-m. x^i)"
 | 
|  |     30 | proof -
 | 
|  |     31 |   have "(\<Sum>i=m..n. x^i) = x^m * (\<Sum>i=m..n. x^(i-m))"
 | 
|  |     32 |     by (simp add: setsum_right_distrib power_add [symmetric])
 | 
|  |     33 |   also have "... = x^m * (\<Sum>i\<le>n-m. x^i)"
 | 
|  |     34 |     apply (subst setsum_reindex_cong [where f = "%i. i+m" and A = "{..n-m}"])
 | 
|  |     35 |     apply (auto simp: image_def)
 | 
|  |     36 |     apply (rule_tac x="x-m" in bexI, auto)
 | 
|  |     37 |     by (metis assms ordered_cancel_comm_monoid_diff_class.le_diff_conv2)
 | 
|  |     38 |   finally show ?thesis .
 | 
|  |     39 | qed
 | 
|  |     40 | 
 | 
|  |     41 | lemma setsum_gp_multiplied:
 | 
|  |     42 |   fixes x :: "'a::{comm_ring,monoid_mult}"
 | 
|  |     43 |   assumes "m \<le> n"
 | 
|  |     44 |   shows "(1 - x) * (\<Sum>i=m..n. x^i) = x^m - x^Suc n"
 | 
|  |     45 | proof -
 | 
|  |     46 |   have  "(1 - x) * (\<Sum>i=m..n. x^i) = x^m * (1 - x) * (\<Sum>i\<le>n-m. x^i)"
 | 
|  |     47 |     by (metis ab_semigroup_mult_class.mult_ac(1) assms mult_commute setsum_power_shift)
 | 
|  |     48 |   also have "... =x^m * (1 - x^Suc(n-m))"
 | 
|  |     49 |     by (metis ab_semigroup_mult_class.mult_ac(1) setsum_gp_basic)
 | 
|  |     50 |   also have "... = x^m - x^Suc n"
 | 
|  |     51 |     using assms
 | 
|  |     52 |     by (simp add: algebra_simps) (metis le_add_diff_inverse power_add)
 | 
|  |     53 |   finally show ?thesis .
 | 
|  |     54 | qed
 | 
|  |     55 | 
 | 
|  |     56 | lemma setsum_gp:
 | 
|  |     57 |   fixes x :: "'a::{comm_ring,division_ring_inverse_zero}"
 | 
|  |     58 |   shows   "(\<Sum>i=m..n. x^i) =
 | 
|  |     59 |                (if n < m then 0
 | 
|  |     60 |                 else if x = 1 then of_nat((n + 1) - m)
 | 
|  |     61 |                 else (x^m - x^Suc n) / (1 - x))"
 | 
|  |     62 | using setsum_gp_multiplied [of m n x] 
 | 
|  |     63 | apply (auto simp: real_of_nat_def)
 | 
|  |     64 | by (metis eq_iff_diff_eq_0 mult_commute nonzero_divide_eq_eq)
 | 
|  |     65 | 
 | 
|  |     66 | lemma setsum_gp_offset:
 | 
|  |     67 |   fixes x :: "'a::{comm_ring,division_ring_inverse_zero}"
 | 
|  |     68 |   shows   "(\<Sum>i=m..m+n. x^i) =
 | 
|  |     69 |        (if x = 1 then of_nat n + 1 else x^m * (1 - x^Suc n) / (1 - x))"
 | 
|  |     70 |   using setsum_gp [of x m "m+n"]
 | 
|  |     71 |   by (auto simp: power_add algebra_simps)
 | 
|  |     72 | 
 | 
|  |     73 | subsection{*Basics about polynomial functions: extremal behaviour and root counts.*}
 | 
|  |     74 | 
 | 
|  |     75 | lemma sub_polyfun:
 | 
|  |     76 |   fixes x :: "'a::{comm_ring,monoid_mult}"
 | 
|  |     77 |   shows   "(\<Sum>i\<le>n. a i * x^i) - (\<Sum>i\<le>n. a i * y^i) = 
 | 
|  |     78 |            (x - y) * (\<Sum>j<n. \<Sum>k= Suc j..n. a k * y^(k - Suc j) * x^j)"
 | 
|  |     79 | proof -
 | 
|  |     80 |   have "(\<Sum>i\<le>n. a i * x^i) - (\<Sum>i\<le>n. a i * y^i) = 
 | 
|  |     81 |         (\<Sum>i\<le>n. a i * (x^i - y^i))"
 | 
|  |     82 |     by (simp add: algebra_simps setsum_subtractf [symmetric])
 | 
|  |     83 |   also have "... = (\<Sum>i\<le>n. a i * (x - y) * (\<Sum>j<i. y^(i - Suc j) * x^j))"
 | 
|  |     84 |     by (simp add: power_diff_sumr2 mult_ac)
 | 
|  |     85 |   also have "... = (x - y) * (\<Sum>i\<le>n. (\<Sum>j<i. a i * y^(i - Suc j) * x^j))"
 | 
|  |     86 |     by (simp add: setsum_right_distrib mult_ac)
 | 
|  |     87 |   also have "... = (x - y) * (\<Sum>j<n. (\<Sum>i=Suc j..n. a i * y^(i - Suc j) * x^j))"
 | 
|  |     88 |     by (simp add: nested_setsum_swap')
 | 
|  |     89 |   finally show ?thesis .
 | 
|  |     90 | qed
 | 
|  |     91 | 
 | 
|  |     92 | lemma sub_polyfun_alt:
 | 
|  |     93 |   fixes x :: "'a::{comm_ring,monoid_mult}"
 | 
|  |     94 |   shows   "(\<Sum>i\<le>n. a i * x^i) - (\<Sum>i\<le>n. a i * y^i) = 
 | 
|  |     95 |            (x - y) * (\<Sum>j<n. \<Sum>k<n-j. a (j+k+1) * y^k * x^j)"
 | 
|  |     96 | proof -
 | 
|  |     97 |   { fix j
 | 
|  |     98 |     assume "j < n"
 | 
|  |     99 |     have "(\<Sum>k = Suc j..n. a k * y^(k - Suc j) * x^j) =
 | 
|  |    100 |           (\<Sum>k <n - j. a (Suc (j + k)) * y^k * x^j)"
 | 
|  |    101 |       apply (rule_tac f = "\<lambda>i. Suc j + i" in setsum_reindex_cong)
 | 
|  |    102 |       apply (auto simp: inj_on_def image_def atLeastLessThan_def lessThan_def)
 | 
|  |    103 |       apply (metis Suc_le_eq diff_add_inverse diff_less_mono le_add1 less_imp_Suc_add)
 | 
|  |    104 |       done }
 | 
|  |    105 |   then show ?thesis
 | 
|  |    106 |     by (simp add: sub_polyfun)
 | 
|  |    107 | qed
 | 
|  |    108 | 
 | 
|  |    109 | lemma polyfun_linear_factor:
 | 
|  |    110 |   fixes a :: "'a::{comm_ring,monoid_mult}"
 | 
|  |    111 |   shows  "\<exists>b. \<forall>z. (\<Sum>i\<le>n. c i * z^i) = 
 | 
|  |    112 |                   (z-a) * (\<Sum>i<n. b i * z^i) + (\<Sum>i\<le>n. c i * a^i)"
 | 
|  |    113 | proof -
 | 
|  |    114 |   { fix z
 | 
|  |    115 |     have "(\<Sum>i\<le>n. c i * z^i) - (\<Sum>i\<le>n. c i * a^i) = 
 | 
|  |    116 |           (z - a) * (\<Sum>j<n. (\<Sum>k = Suc j..n. c k * a^(k - Suc j)) * z^j)"
 | 
|  |    117 |       by (simp add: sub_polyfun setsum_left_distrib)
 | 
|  |    118 |     then have "(\<Sum>i\<le>n. c i * z^i) = 
 | 
|  |    119 |           (z - a) * (\<Sum>j<n. (\<Sum>k = Suc j..n. c k * a^(k - Suc j)) * z^j)
 | 
|  |    120 |           + (\<Sum>i\<le>n. c i * a^i)"
 | 
|  |    121 |       by (simp add: algebra_simps) }
 | 
|  |    122 |   then show ?thesis
 | 
|  |    123 |     by (intro exI allI) 
 | 
|  |    124 | qed
 | 
|  |    125 | 
 | 
|  |    126 | lemma polyfun_linear_factor_root:
 | 
|  |    127 |   fixes a :: "'a::{comm_ring,monoid_mult}"
 | 
|  |    128 |   assumes "(\<Sum>i\<le>n. c i * a^i) = 0"
 | 
|  |    129 |   shows  "\<exists>b. \<forall>z. (\<Sum>i\<le>n. c i * z^i) = (z-a) * (\<Sum>i<n. b i * z^i)"
 | 
|  |    130 |   using polyfun_linear_factor [of c n a] assms
 | 
|  |    131 |   by simp
 | 
|  |    132 | 
 | 
|  |    133 | lemma adhoc_norm_triangle: "a + norm(y) \<le> b ==> norm(x) \<le> a ==> norm(x + y) \<le> b"
 | 
|  |    134 |   by (metis norm_triangle_mono order.trans order_refl)
 | 
|  |    135 | 
 | 
|  |    136 | lemma polyfun_extremal_lemma:
 | 
|  |    137 |   fixes c :: "nat \<Rightarrow> 'a::real_normed_div_algebra"
 | 
|  |    138 |   assumes "e > 0"
 | 
|  |    139 |     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"
 | 
|  |    140 | proof (induction n)
 | 
|  |    141 |   case 0
 | 
|  |    142 |   show ?case 
 | 
|  |    143 |     by (rule exI [where x="norm (c 0) / e"]) (auto simp: mult_commute pos_divide_le_eq assms)
 | 
|  |    144 | next
 | 
|  |    145 |   case (Suc n)
 | 
|  |    146 |   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" ..
 | 
|  |    147 |   show ?case
 | 
|  |    148 |   proof (rule exI [where x="max 1 (max M ((e + norm(c(Suc n))) / e))"], clarify)
 | 
|  |    149 |     fix z::'a
 | 
|  |    150 |     assume "max 1 (max M ((e + norm (c (Suc n))) / e)) \<le> norm z"
 | 
|  |    151 |     then have norm1: "0 < norm z" "M \<le> norm z" "(e + norm (c (Suc n))) / e \<le> norm z"
 | 
|  |    152 |       by auto
 | 
|  |    153 |     then have norm2: "(e + norm (c (Suc n))) \<le> e * norm z"  "(norm z * norm z ^ n) > 0"
 | 
|  |    154 |       apply (metis assms less_divide_eq mult_commute not_le) 
 | 
|  |    155 |       using norm1 apply (metis mult_pos_pos zero_less_power)
 | 
|  |    156 |       done
 | 
|  |    157 |     have "e * (norm z * norm z ^ n) + norm (c (Suc n) * (z * z ^ n)) =
 | 
|  |    158 |           (e + norm (c (Suc n))) * (norm z * norm z ^ n)"
 | 
|  |    159 |       by (simp add: norm_mult norm_power algebra_simps)
 | 
|  |    160 |     also have "... \<le> (e * norm z) * (norm z * norm z ^ n)"
 | 
|  |    161 |       using norm2 by (metis real_mult_le_cancel_iff1) 
 | 
|  |    162 |     also have "... = e * (norm z * (norm z * norm z ^ n))"
 | 
|  |    163 |       by (simp add: algebra_simps)
 | 
|  |    164 |     finally have "e * (norm z * norm z ^ n) + norm (c (Suc n) * (z * z ^ n))
 | 
|  |    165 |                   \<le> e * (norm z * (norm z * norm z ^ n))" .
 | 
|  |    166 |     then show "norm (\<Sum>i\<le>Suc n. c i * z^i) \<le> e * norm z ^ Suc (Suc n)" using M norm1
 | 
|  |    167 |       by (drule_tac x=z in spec) (auto simp: intro!: adhoc_norm_triangle)
 | 
|  |    168 |     qed
 | 
|  |    169 | qed
 | 
|  |    170 | 
 | 
|  |    171 | lemma norm_lemma_xy: "\<lbrakk>abs b + 1 \<le> norm(y) - a; norm(x) \<le> a\<rbrakk> \<Longrightarrow> b \<le> norm(x + y)"
 | 
|  |    172 |   by (metis abs_add_one_not_less_self add_commute diff_le_eq dual_order.trans le_less_linear 
 | 
|  |    173 |          norm_diff_ineq)
 | 
|  |    174 | 
 | 
|  |    175 | lemma polyfun_extremal:
 | 
|  |    176 |   fixes c :: "nat \<Rightarrow> 'a::real_normed_div_algebra"
 | 
|  |    177 |   assumes "\<exists>k. k \<noteq> 0 \<and> k \<le> n \<and> c k \<noteq> 0"
 | 
|  |    178 |     shows "eventually (\<lambda>z. norm(\<Sum>i\<le>n. c i * z^i) \<ge> B) at_infinity"
 | 
|  |    179 | using assms
 | 
|  |    180 | proof (induction n)
 | 
|  |    181 |   case 0 then show ?case
 | 
|  |    182 |     by simp
 | 
|  |    183 | next
 | 
|  |    184 |   case (Suc n)
 | 
|  |    185 |   show ?case
 | 
|  |    186 |   proof (cases "c (Suc n) = 0")
 | 
|  |    187 |     case True
 | 
|  |    188 |     with Suc show ?thesis
 | 
|  |    189 |       by auto (metis diff_is_0_eq diffs0_imp_equal less_Suc_eq_le not_less_eq)
 | 
|  |    190 |   next
 | 
|  |    191 |     case False
 | 
|  |    192 |     with polyfun_extremal_lemma [of "norm(c (Suc n)) / 2" c n]
 | 
|  |    193 |     obtain M where M: "\<And>z. M \<le> norm z \<Longrightarrow> 
 | 
|  |    194 |                norm (\<Sum>i\<le>n. c i * z^i) \<le> norm (c (Suc n)) / 2 * norm z ^ Suc n"
 | 
|  |    195 |       by auto
 | 
|  |    196 |     show ?thesis
 | 
|  |    197 |     unfolding eventually_at_infinity
 | 
|  |    198 |     proof (rule exI [where x="max M (max 1 ((abs B + 1) / (norm (c (Suc n)) / 2)))"], clarsimp)
 | 
|  |    199 |       fix z::'a
 | 
|  |    200 |       assume les: "M \<le> norm z"  "1 \<le> norm z"  "(\<bar>B\<bar> * 2 + 2) / norm (c (Suc n)) \<le> norm z"
 | 
|  |    201 |       then have "\<bar>B\<bar> * 2 + 2 \<le> norm z * norm (c (Suc n))"
 | 
|  |    202 |         by (metis False pos_divide_le_eq zero_less_norm_iff)
 | 
|  |    203 |       then have "\<bar>B\<bar> * 2 + 2 \<le> norm z ^ (Suc n) * norm (c (Suc n))" 
 | 
|  |    204 |         by (metis `1 \<le> norm z` order.trans mult_right_mono norm_ge_zero self_le_power zero_less_Suc)
 | 
|  |    205 |       then show "B \<le> norm ((\<Sum>i\<le>n. c i * z^i) + c (Suc n) * (z * z ^ n))" using M les
 | 
|  |    206 |         apply auto
 | 
|  |    207 |         apply (rule norm_lemma_xy [where a = "norm (c (Suc n)) * norm z ^ (Suc n) / 2"])
 | 
|  |    208 |         apply (simp_all add: norm_mult norm_power)
 | 
|  |    209 |         done
 | 
|  |    210 |     qed
 | 
|  |    211 |   qed
 | 
|  |    212 | qed
 | 
|  |    213 | 
 | 
|  |    214 | lemma polyfun_rootbound:
 | 
|  |    215 |  fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
 | 
|  |    216 |  assumes "\<exists>k. k \<le> n \<and> c k \<noteq> 0"
 | 
|  |    217 |    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"
 | 
|  |    218 | using assms
 | 
|  |    219 | proof (induction n arbitrary: c)
 | 
|  |    220 |  case (Suc n) show ?case
 | 
|  |    221 |  proof (cases "{z. (\<Sum>i\<le>Suc n. c i * z^i) = 0} = {}")
 | 
|  |    222 |    case False
 | 
|  |    223 |    then obtain a where a: "(\<Sum>i\<le>Suc n. c i * a^i) = 0"
 | 
|  |    224 |      by auto
 | 
|  |    225 |    from polyfun_linear_factor_root [OF this]
 | 
|  |    226 |    obtain b where "\<And>z. (\<Sum>i\<le>Suc n. c i * z^i) = (z - a) * (\<Sum>i< Suc n. b i * z^i)"
 | 
|  |    227 |      by auto
 | 
|  |    228 |    then have b: "\<And>z. (\<Sum>i\<le>Suc n. c i * z^i) = (z - a) * (\<Sum>i\<le>n. b i * z^i)"
 | 
|  |    229 |      by (metis lessThan_Suc_atMost)
 | 
|  |    230 |    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}"
 | 
|  |    231 |      by auto
 | 
|  |    232 |    have c0: "c 0 = - (a * b 0)" using  b [of 0]
 | 
|  |    233 |      by simp
 | 
|  |    234 |    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"
 | 
|  |    235 |      by (metis Suc.prems le0 minus_zero mult_zero_right)
 | 
|  |    236 |    have "\<exists>k\<le>n. b k \<noteq> 0" 
 | 
|  |    237 |      apply (rule ccontr)
 | 
|  |    238 |      using polyfun_extremal [OF extr_prem, of 1]
 | 
|  |    239 |      apply (auto simp: eventually_at_infinity b simp del: setsum_atMost_Suc)
 | 
|  |    240 |      apply (drule_tac x="of_real ba" in spec, simp)
 | 
|  |    241 |      done
 | 
|  |    242 |    then show ?thesis using Suc.IH [of b] ins_ab
 | 
|  |    243 |      by (auto simp: card_insert_if)
 | 
|  |    244 |    qed simp
 | 
|  |    245 | qed simp
 | 
|  |    246 | 
 | 
|  |    247 | corollary
 | 
|  |    248 |   fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
 | 
|  |    249 |   assumes "\<exists>k. k \<le> n \<and> c k \<noteq> 0"
 | 
|  |    250 |     shows polyfun_rootbound_finite: "finite {z. (\<Sum>i\<le>n. c i * z^i) = 0}"
 | 
|  |    251 |       and polyfun_rootbound_card:   "card {z. (\<Sum>i\<le>n. c i * z^i) = 0} \<le> n"
 | 
|  |    252 | using polyfun_rootbound [OF assms] by auto
 | 
|  |    253 | 
 | 
|  |    254 | lemma polyfun_finite_roots:
 | 
|  |    255 |   fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
 | 
|  |    256 |     shows  "finite {z. (\<Sum>i\<le>n. c i * z^i) = 0} \<longleftrightarrow> (\<exists>k. k \<le> n \<and> c k \<noteq> 0)"
 | 
|  |    257 | proof (cases " \<exists>k\<le>n. c k \<noteq> 0")
 | 
|  |    258 |   case True then show ?thesis 
 | 
|  |    259 |     by (blast intro: polyfun_rootbound_finite)
 | 
|  |    260 | next
 | 
|  |    261 |   case False then show ?thesis 
 | 
|  |    262 |     by (auto simp: infinite_UNIV_char_0)
 | 
|  |    263 | qed
 | 
|  |    264 | 
 | 
|  |    265 | lemma polyfun_eq_0:
 | 
|  |    266 |   fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
 | 
|  |    267 |     shows  "(\<forall>z. (\<Sum>i\<le>n. c i * z^i) = 0) \<longleftrightarrow> (\<forall>k. k \<le> n \<longrightarrow> c k = 0)"
 | 
|  |    268 | proof (cases "(\<forall>z. (\<Sum>i\<le>n. c i * z^i) = 0)")
 | 
|  |    269 |   case True
 | 
|  |    270 |   then have "~ finite {z. (\<Sum>i\<le>n. c i * z^i) = 0}"
 | 
|  |    271 |     by (simp add: infinite_UNIV_char_0)
 | 
|  |    272 |   with True show ?thesis
 | 
|  |    273 |     by (metis (poly_guards_query) polyfun_rootbound_finite)
 | 
|  |    274 | next
 | 
|  |    275 |   case False
 | 
|  |    276 |   then show ?thesis
 | 
|  |    277 |     by auto
 | 
|  |    278 | qed
 | 
|  |    279 | 
 | 
|  |    280 | lemma polyfun_eq_const:
 | 
|  |    281 |   fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
 | 
|  |    282 |     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)"
 | 
|  |    283 | proof -
 | 
|  |    284 |   {fix z
 | 
|  |    285 |     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"
 | 
|  |    286 |       by (induct n) auto
 | 
|  |    287 |   } then
 | 
|  |    288 |   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)"
 | 
|  |    289 |     by auto
 | 
|  |    290 |   also have "... \<longleftrightarrow>  c 0 = k \<and> (\<forall>k. k \<noteq> 0 \<and> k \<le> n \<longrightarrow> c k = 0)"
 | 
|  |    291 |     by (auto simp: polyfun_eq_0)
 | 
|  |    292 |   finally show ?thesis .
 | 
|  |    293 | qed
 | 
|  |    294 | 
 | 
|  |    295 | end
 | 
|  |    296 | 
 |