| author | Fabian Huch <huch@in.tum.de> | 
| Tue, 11 Jun 2024 14:27:04 +0200 | |
| changeset 80347 | 613ac8c77a84 | 
| parent 72569 | d56e4eeae967 | 
| child 82522 | 62afd98e3f3e | 
| permissions | -rw-r--r-- | 
| 56215 | 1 | (* Author: John Harrison and Valentina Bruno | 
| 2 | Ported from "hol_light/Multivariate/complexes.ml" by L C Paulson | |
| 3 | *) | |
| 4 | ||
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 5 | section \<open>Polynomial Functions: Extremal Behaviour and Root Counts\<close> | 
| 61560 | 6 | |
| 63594 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
62626diff
changeset | 7 | theory Poly_Roots | 
| 56215 | 8 | imports Complex_Main | 
| 9 | begin | |
| 10 | ||
| 69683 | 11 | subsection\<open>Basics about polynomial functions: extremal behaviour and root counts\<close> | 
| 56215 | 12 | |
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 13 | lemma sub_polyfun: | 
| 56215 | 14 |   fixes x :: "'a::{comm_ring,monoid_mult}"
 | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 15 | shows "(\<Sum>i\<le>n. a i * x^i) - (\<Sum>i\<le>n. a i * y^i) = | 
| 56215 | 16 | (x - y) * (\<Sum>j<n. \<Sum>k= Suc j..n. a k * y^(k - Suc j) * x^j)" | 
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 17 | proof - | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 18 | have "(\<Sum>i\<le>n. a i * x^i) - (\<Sum>i\<le>n. a i * y^i) = | 
| 56215 | 19 | (\<Sum>i\<le>n. a i * (x^i - y^i))" | 
| 64267 | 20 | by (simp add: algebra_simps sum_subtractf [symmetric]) | 
| 56215 | 21 | 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 | 22 | by (simp add: power_diff_sumr2 ac_simps) | 
| 56215 | 23 | also have "... = (x - y) * (\<Sum>i\<le>n. (\<Sum>j<i. a i * y^(i - Suc j) * x^j))" | 
| 64267 | 24 | by (simp add: sum_distrib_left ac_simps) | 
| 56215 | 25 | also have "... = (x - y) * (\<Sum>j<n. (\<Sum>i=Suc j..n. a i * y^(i - Suc j) * x^j))" | 
| 70113 
c8deb8ba6d05
Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
 paulson <lp15@cam.ac.uk> parents: 
70097diff
changeset | 26 | by (simp add: sum.nested_swap') | 
| 56215 | 27 | finally show ?thesis . | 
| 28 | qed | |
| 29 | ||
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 30 | lemma sub_polyfun_alt: | 
| 56215 | 31 |   fixes x :: "'a::{comm_ring,monoid_mult}"
 | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 32 | shows "(\<Sum>i\<le>n. a i * x^i) - (\<Sum>i\<le>n. a i * y^i) = | 
| 56215 | 33 | (x - y) * (\<Sum>j<n. \<Sum>k<n-j. a (j+k+1) * y^k * x^j)" | 
| 34 | proof - | |
| 35 |   { fix j
 | |
| 36 | have "(\<Sum>k = Suc j..n. a k * y^(k - Suc j) * x^j) = | |
| 37 | (\<Sum>k <n - j. a (Suc (j + k)) * y^k * x^j)" | |
| 64267 | 38 | by (rule sum.reindex_bij_witness[where i="\<lambda>i. i + Suc j" and j="\<lambda>i. i - Suc j"]) auto } | 
| 56215 | 39 | then show ?thesis | 
| 40 | by (simp add: sub_polyfun) | |
| 41 | qed | |
| 42 | ||
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 43 | lemma polyfun_linear_factor: | 
| 56215 | 44 |   fixes a :: "'a::{comm_ring,monoid_mult}"
 | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 45 | shows "\<exists>b. \<forall>z. (\<Sum>i\<le>n. c i * z^i) = | 
| 56215 | 46 | (z-a) * (\<Sum>i<n. b i * z^i) + (\<Sum>i\<le>n. c i * a^i)" | 
| 47 | proof - | |
| 48 |   { fix z
 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 49 | have "(\<Sum>i\<le>n. c i * z^i) - (\<Sum>i\<le>n. c i * a^i) = | 
| 56215 | 50 | (z - a) * (\<Sum>j<n. (\<Sum>k = Suc j..n. c k * a^(k - Suc j)) * z^j)" | 
| 64267 | 51 | by (simp add: sub_polyfun sum_distrib_right) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 52 | then have "(\<Sum>i\<le>n. c i * z^i) = | 
| 56215 | 53 | (z - a) * (\<Sum>j<n. (\<Sum>k = Suc j..n. c k * a^(k - Suc j)) * z^j) | 
| 54 | + (\<Sum>i\<le>n. c i * a^i)" | |
| 55 | by (simp add: algebra_simps) } | |
| 56 | then show ?thesis | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 57 | by (intro exI allI) | 
| 56215 | 58 | qed | 
| 59 | ||
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 60 | lemma polyfun_linear_factor_root: | 
| 56215 | 61 |   fixes a :: "'a::{comm_ring,monoid_mult}"
 | 
| 62 | assumes "(\<Sum>i\<le>n. c i * a^i) = 0" | |
| 63 | shows "\<exists>b. \<forall>z. (\<Sum>i\<le>n. c i * z^i) = (z-a) * (\<Sum>i<n. b i * z^i)" | |
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 64 | using polyfun_linear_factor [of c n a] assms | 
| 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 65 | by simp | 
| 56215 | 66 | |
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 67 | lemma adhoc_norm_triangle: "a + norm(y) \<le> b ==> norm(x) \<le> a ==> norm(x + y) \<le> b" | 
| 56215 | 68 | by (metis norm_triangle_mono order.trans order_refl) | 
| 69 | ||
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 70 | proposition polyfun_extremal_lemma: | 
| 56215 | 71 | fixes c :: "nat \<Rightarrow> 'a::real_normed_div_algebra" | 
| 72 | assumes "e > 0" | |
| 73 | 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" | |
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 74 | proof (induction n) | 
| 56215 | 75 | case 0 | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 76 | show ?case | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57129diff
changeset | 77 | by (rule exI [where x="norm (c 0) / e"]) (auto simp: mult.commute pos_divide_le_eq assms) | 
| 56215 | 78 | next | 
| 79 | case (Suc n) | |
| 80 | 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" .. | |
| 81 | show ?case | |
| 82 | proof (rule exI [where x="max 1 (max M ((e + norm(c(Suc n))) / e))"], clarify) | |
| 83 | fix z::'a | |
| 84 | assume "max 1 (max M ((e + norm (c (Suc n))) / e)) \<le> norm z" | |
| 85 | then have norm1: "0 < norm z" "M \<le> norm z" "(e + norm (c (Suc n))) / e \<le> norm z" | |
| 86 | by auto | |
| 87 | then have norm2: "(e + norm (c (Suc n))) \<le> e * norm z" "(norm z * norm z ^ n) > 0" | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 88 | apply (metis assms less_divide_eq mult.commute not_le) | 
| 56215 | 89 | using norm1 apply (metis mult_pos_pos zero_less_power) | 
| 90 | done | |
| 91 | have "e * (norm z * norm z ^ n) + norm (c (Suc n) * (z * z ^ n)) = | |
| 92 | (e + norm (c (Suc n))) * (norm z * norm z ^ n)" | |
| 93 | by (simp add: norm_mult norm_power algebra_simps) | |
| 94 | also have "... \<le> (e * norm z) * (norm z * norm z ^ n)" | |
| 72569 
d56e4eeae967
mult_le_cancel_iff1, mult_le_cancel_iff2, mult_less_iff1 generalised from the real_ versions
 paulson <lp15@cam.ac.uk> parents: 
70113diff
changeset | 95 | using norm2 | 
| 
d56e4eeae967
mult_le_cancel_iff1, mult_le_cancel_iff2, mult_less_iff1 generalised from the real_ versions
 paulson <lp15@cam.ac.uk> parents: 
70113diff
changeset | 96 | using assms mult_mono by fastforce | 
| 56215 | 97 | also have "... = e * (norm z * (norm z * norm z ^ n))" | 
| 98 | by (simp add: algebra_simps) | |
| 99 | finally have "e * (norm z * norm z ^ n) + norm (c (Suc n) * (z * z ^ n)) | |
| 100 | \<le> e * (norm z * (norm z * norm z ^ n))" . | |
| 101 | then show "norm (\<Sum>i\<le>Suc n. c i * z^i) \<le> e * norm z ^ Suc (Suc n)" using M norm1 | |
| 102 | by (drule_tac x=z in spec) (auto simp: intro!: adhoc_norm_triangle) | |
| 103 | qed | |
| 104 | qed | |
| 105 | ||
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 106 | lemma norm_lemma_xy: assumes "\<bar>b\<bar> + 1 \<le> norm(y) - a" "norm(x) \<le> a" shows "b \<le> norm(x + y)" | 
| 62626 
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
 paulson <lp15@cam.ac.uk> parents: 
61945diff
changeset | 107 | proof - | 
| 
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
 paulson <lp15@cam.ac.uk> parents: 
61945diff
changeset | 108 | have "b \<le> norm y - norm x" | 
| 
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
 paulson <lp15@cam.ac.uk> parents: 
61945diff
changeset | 109 | using assms by linarith | 
| 
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
 paulson <lp15@cam.ac.uk> parents: 
61945diff
changeset | 110 | then show ?thesis | 
| 
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
 paulson <lp15@cam.ac.uk> parents: 
61945diff
changeset | 111 | by (metis (no_types) add.commute norm_diff_ineq order_trans) | 
| 
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
 paulson <lp15@cam.ac.uk> parents: 
61945diff
changeset | 112 | qed | 
| 56215 | 113 | |
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 114 | proposition polyfun_extremal: | 
| 56215 | 115 | fixes c :: "nat \<Rightarrow> 'a::real_normed_div_algebra" | 
| 116 | assumes "\<exists>k. k \<noteq> 0 \<and> k \<le> n \<and> c k \<noteq> 0" | |
| 117 | shows "eventually (\<lambda>z. norm(\<Sum>i\<le>n. c i * z^i) \<ge> B) at_infinity" | |
| 118 | using assms | |
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 119 | proof (induction n) | 
| 56215 | 120 | case 0 then show ?case | 
| 121 | by simp | |
| 122 | next | |
| 123 | case (Suc n) | |
| 124 | show ?case | |
| 125 | proof (cases "c (Suc n) = 0") | |
| 126 | case True | |
| 127 | with Suc show ?thesis | |
| 128 | by auto (metis diff_is_0_eq diffs0_imp_equal less_Suc_eq_le not_less_eq) | |
| 129 | next | |
| 130 | case False | |
| 131 | with polyfun_extremal_lemma [of "norm(c (Suc n)) / 2" c n] | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 132 | obtain M where M: "\<And>z. M \<le> norm z \<Longrightarrow> | 
| 56215 | 133 | norm (\<Sum>i\<le>n. c i * z^i) \<le> norm (c (Suc n)) / 2 * norm z ^ Suc n" | 
| 134 | by auto | |
| 135 | show ?thesis | |
| 136 | unfolding eventually_at_infinity | |
| 61945 | 137 | proof (rule exI [where x="max M (max 1 ((\<bar>B\<bar> + 1) / (norm (c (Suc n)) / 2)))"], clarsimp) | 
| 56215 | 138 | fix z::'a | 
| 139 | assume les: "M \<le> norm z" "1 \<le> norm z" "(\<bar>B\<bar> * 2 + 2) / norm (c (Suc n)) \<le> norm z" | |
| 140 | then have "\<bar>B\<bar> * 2 + 2 \<le> norm z * norm (c (Suc n))" | |
| 141 | by (metis False pos_divide_le_eq zero_less_norm_iff) | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 142 | then have "\<bar>B\<bar> * 2 + 2 \<le> norm z ^ (Suc n) * norm (c (Suc n))" | 
| 60420 | 143 | by (metis \<open>1 \<le> norm z\<close> order.trans mult_right_mono norm_ge_zero self_le_power zero_less_Suc) | 
| 56215 | 144 | then show "B \<le> norm ((\<Sum>i\<le>n. c i * z^i) + c (Suc n) * (z * z ^ n))" using M les | 
| 145 | apply auto | |
| 146 | apply (rule norm_lemma_xy [where a = "norm (c (Suc n)) * norm z ^ (Suc n) / 2"]) | |
| 147 | apply (simp_all add: norm_mult norm_power) | |
| 148 | done | |
| 149 | qed | |
| 150 | qed | |
| 151 | qed | |
| 152 | ||
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 153 | proposition polyfun_rootbound: | 
| 56215 | 154 |  fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
 | 
| 155 | assumes "\<exists>k. k \<le> n \<and> c k \<noteq> 0" | |
| 156 |    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"
 | |
| 157 | using assms | |
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 158 | proof (induction n arbitrary: c) | 
| 56215 | 159 | case (Suc n) show ?case | 
| 160 |  proof (cases "{z. (\<Sum>i\<le>Suc n. c i * z^i) = 0} = {}")
 | |
| 161 | case False | |
| 162 | then obtain a where a: "(\<Sum>i\<le>Suc n. c i * a^i) = 0" | |
| 163 | by auto | |
| 164 | from polyfun_linear_factor_root [OF this] | |
| 165 | obtain b where "\<And>z. (\<Sum>i\<le>Suc n. c i * z^i) = (z - a) * (\<Sum>i< Suc n. b i * z^i)" | |
| 166 | by auto | |
| 167 | then have b: "\<And>z. (\<Sum>i\<le>Suc n. c i * z^i) = (z - a) * (\<Sum>i\<le>n. b i * z^i)" | |
| 168 | by (metis lessThan_Suc_atMost) | |
| 169 |    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}"
 | |
| 170 | by auto | |
| 171 | have c0: "c 0 = - (a * b 0)" using b [of 0] | |
| 172 | by simp | |
| 69508 | 173 | then have extr_prem: "\<not> (\<exists>k\<le>n. b k \<noteq> 0) \<Longrightarrow> \<exists>k. k \<noteq> 0 \<and> k \<le> Suc n \<and> c k \<noteq> 0" | 
| 56215 | 174 | by (metis Suc.prems le0 minus_zero mult_zero_right) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 175 | have "\<exists>k\<le>n. b k \<noteq> 0" | 
| 56215 | 176 | apply (rule ccontr) | 
| 177 | using polyfun_extremal [OF extr_prem, of 1] | |
| 70097 
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
 paulson <lp15@cam.ac.uk> parents: 
69739diff
changeset | 178 | apply (auto simp: eventually_at_infinity b simp del: sum.atMost_Suc) | 
| 56215 | 179 | apply (drule_tac x="of_real ba" in spec, simp) | 
| 180 | done | |
| 181 | then show ?thesis using Suc.IH [of b] ins_ab | |
| 182 | by (auto simp: card_insert_if) | |
| 183 | qed simp | |
| 184 | qed simp | |
| 185 | ||
| 69739 | 186 | corollary | 
| 56215 | 187 |   fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
 | 
| 188 | assumes "\<exists>k. k \<le> n \<and> c k \<noteq> 0" | |
| 189 |     shows polyfun_rootbound_finite: "finite {z. (\<Sum>i\<le>n. c i * z^i) = 0}"
 | |
| 190 |       and polyfun_rootbound_card:   "card {z. (\<Sum>i\<le>n. c i * z^i) = 0} \<le> n"
 | |
| 191 | using polyfun_rootbound [OF assms] by auto | |
| 192 | ||
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 193 | proposition polyfun_finite_roots: | 
| 56215 | 194 |   fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
 | 
| 195 |     shows  "finite {z. (\<Sum>i\<le>n. c i * z^i) = 0} \<longleftrightarrow> (\<exists>k. k \<le> n \<and> c k \<noteq> 0)"
 | |
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 196 | proof (cases " \<exists>k\<le>n. c k \<noteq> 0") | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 197 | case True then show ?thesis | 
| 56215 | 198 | by (blast intro: polyfun_rootbound_finite) | 
| 199 | next | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 200 | case False then show ?thesis | 
| 56215 | 201 | by (auto simp: infinite_UNIV_char_0) | 
| 202 | qed | |
| 203 | ||
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 204 | lemma polyfun_eq_0: | 
| 56215 | 205 |   fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
 | 
| 206 | shows "(\<forall>z. (\<Sum>i\<le>n. c i * z^i) = 0) \<longleftrightarrow> (\<forall>k. k \<le> n \<longrightarrow> c k = 0)" | |
| 207 | proof (cases "(\<forall>z. (\<Sum>i\<le>n. c i * z^i) = 0)") | |
| 208 | case True | |
| 69508 | 209 |   then have "\<not> finite {z. (\<Sum>i\<le>n. c i * z^i) = 0}"
 | 
| 56215 | 210 | by (simp add: infinite_UNIV_char_0) | 
| 211 | with True show ?thesis | |
| 212 | by (metis (poly_guards_query) polyfun_rootbound_finite) | |
| 213 | next | |
| 214 | case False | |
| 215 | then show ?thesis | |
| 216 | by auto | |
| 217 | qed | |
| 218 | ||
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 219 | theorem polyfun_eq_const: | 
| 56215 | 220 |   fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
 | 
| 221 | 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)" | |
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 222 | proof - | 
| 56215 | 223 |   {fix z
 | 
| 224 | 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" | |
| 225 | by (induct n) auto | |
| 226 | } then | |
| 227 | 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)" | |
| 228 | by auto | |
| 229 | also have "... \<longleftrightarrow> c 0 = k \<and> (\<forall>k. k \<noteq> 0 \<and> k \<le> n \<longrightarrow> c k = 0)" | |
| 230 | by (auto simp: polyfun_eq_0) | |
| 231 | finally show ?thesis . | |
| 232 | qed | |
| 233 | ||
| 234 | end | |
| 235 |