| author | nipkow | 
| Thu, 18 Jul 2024 16:00:40 +0200 | |
| changeset 80578 | 27e66a8323b2 | 
| parent 80061 | 4c1347e172b1 | 
| permissions | -rw-r--r-- | 
| 65435 | 1 | (* Title: HOL/Computational_Algebra/Fundamental_Theorem_Algebra.thy | 
| 2 | Author: Amine Chaieb, TU Muenchen | |
| 3 | *) | |
| 26123 | 4 | |
| 60424 | 5 | section \<open>Fundamental Theorem of Algebra\<close> | 
| 26123 | 6 | |
| 7 | theory Fundamental_Theorem_Algebra | |
| 51537 | 8 | imports Polynomial Complex_Main | 
| 26123 | 9 | begin | 
| 10 | ||
| 60424 | 11 | subsection \<open>More lemmas about module of complex numbers\<close> | 
| 26123 | 12 | |
| 60424 | 13 | text \<open>The triangle inequality for cmod\<close> | 
| 14 | ||
| 26123 | 15 | lemma complex_mod_triangle_sub: "cmod w \<le> cmod (w + z) + norm z" | 
| 77341 | 16 | by (metis add_diff_cancel norm_triangle_ineq4) | 
| 26123 | 17 | |
| 60424 | 18 | |
| 19 | subsection \<open>Basic lemmas about polynomials\<close> | |
| 26123 | 20 | |
| 21 | lemma poly_bound_exists: | |
| 56778 | 22 |   fixes p :: "'a::{comm_semiring_0,real_normed_div_algebra} poly"
 | 
| 23 | shows "\<exists>m. m > 0 \<and> (\<forall>z. norm z \<le> r \<longrightarrow> norm (poly p z) \<le> m)" | |
| 24 | proof (induct p) | |
| 25 | case 0 | |
| 26 | then show ?case by (rule exI[where x=1]) simp | |
| 26123 | 27 | next | 
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 28 | case (pCons c cs) | 
| 55735 
81ba62493610
generalised some results using type classes
 paulson <lp15@cam.ac.uk> parents: 
55734diff
changeset | 29 | from pCons.hyps obtain m where m: "\<forall>z. norm z \<le> r \<longrightarrow> norm (poly cs z) \<le> m" | 
| 26123 | 30 | by blast | 
| 55735 
81ba62493610
generalised some results using type classes
 paulson <lp15@cam.ac.uk> parents: 
55734diff
changeset | 31 | let ?k = " 1 + norm c + \<bar>r * m\<bar>" | 
| 56795 | 32 | have kp: "?k > 0" | 
| 33 | using abs_ge_zero[of "r*m"] norm_ge_zero[of c] by arith | |
| 60424 | 34 | have "norm (poly (pCons c cs) z) \<le> ?k" if H: "norm z \<le> r" for z | 
| 35 | proof - | |
| 56778 | 36 | from m H have th: "norm (poly cs z) \<le> m" | 
| 37 | by blast | |
| 56795 | 38 | from H have rp: "r \<ge> 0" | 
| 39 | using norm_ge_zero[of z] by arith | |
| 40 | have "norm (poly (pCons c cs) z) \<le> norm c + norm (z * poly cs z)" | |
| 27514 | 41 | using norm_triangle_ineq[of c "z* poly cs z"] by simp | 
| 77303 | 42 | also have "\<dots> \<le> ?k" | 
| 56778 | 43 | using mult_mono[OF H th rp norm_ge_zero[of "poly cs z"]] | 
| 55735 
81ba62493610
generalised some results using type classes
 paulson <lp15@cam.ac.uk> parents: 
55734diff
changeset | 44 | by (simp add: norm_mult) | 
| 60424 | 45 | finally show ?thesis . | 
| 46 | qed | |
| 26123 | 47 | with kp show ?case by blast | 
| 48 | qed | |
| 49 | ||
| 50 | ||
| 60424 | 51 | text \<open>Offsetting the variable in a polynomial gives another of same degree\<close> | 
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 52 | |
| 52380 | 53 | definition offset_poly :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a poly" | 
| 56778 | 54 | where "offset_poly p h = fold_coeffs (\<lambda>a q. smult h q + pCons a q) p 0" | 
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 55 | |
| 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 56 | lemma offset_poly_0: "offset_poly 0 h = 0" | 
| 52380 | 57 | by (simp add: offset_poly_def) | 
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 58 | |
| 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 59 | lemma offset_poly_pCons: | 
| 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 60 | "offset_poly (pCons a p) h = | 
| 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 61 | smult h (offset_poly p h) + pCons a (offset_poly p h)" | 
| 52380 | 62 | by (cases "p = 0 \<and> a = 0") (auto simp add: offset_poly_def) | 
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 63 | |
| 77282 | 64 | lemma offset_poly_single [simp]: "offset_poly [:a:] h = [:a:]" | 
| 56778 | 65 | by (simp add: offset_poly_pCons offset_poly_0) | 
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 66 | |
| 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 67 | lemma poly_offset_poly: "poly (offset_poly p h) x = poly p (h + x)" | 
| 77282 | 68 | by (induct p) (auto simp add: offset_poly_0 offset_poly_pCons algebra_simps) | 
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 69 | |
| 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 70 | lemma offset_poly_eq_0_lemma: "smult c p + pCons a p = 0 \<Longrightarrow> p = 0" | 
| 56778 | 71 | by (induct p arbitrary: a) (simp, force) | 
| 26123 | 72 | |
| 77282 | 73 | lemma offset_poly_eq_0_iff [simp]: "offset_poly p h = 0 \<longleftrightarrow> p = 0" | 
| 74 | proof | |
| 75 | show "offset_poly p h = 0 \<Longrightarrow> p = 0" | |
| 76 | proof(induction p) | |
| 77 | case 0 | |
| 78 | then show ?case by blast | |
| 79 | next | |
| 80 | case (pCons a p) | |
| 81 | then show ?case | |
| 82 | by (metis offset_poly_eq_0_lemma offset_poly_pCons offset_poly_single) | |
| 83 | qed | |
| 84 | qed (simp add: offset_poly_0) | |
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 85 | |
| 77282 | 86 | lemma degree_offset_poly [simp]: "degree (offset_poly p h) = degree p" | 
| 87 | proof(induction p) | |
| 88 | case 0 | |
| 89 | then show ?case | |
| 90 | by (simp add: offset_poly_0) | |
| 91 | next | |
| 92 | case (pCons a p) | |
| 93 | have "p \<noteq> 0 \<Longrightarrow> degree (offset_poly (pCons a p) h) = Suc (degree p)" | |
| 94 | by (metis degree_add_eq_right degree_pCons_eq degree_smult_le le_imp_less_Suc offset_poly_eq_0_iff offset_poly_pCons pCons.IH) | |
| 95 | then show ?case | |
| 96 | by simp | |
| 97 | qed | |
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 98 | |
| 56778 | 99 | definition "psize p = (if p = 0 then 0 else Suc (degree p))" | 
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 100 | |
| 29538 | 101 | lemma psize_eq_0_iff [simp]: "psize p = 0 \<longleftrightarrow> p = 0" | 
| 102 | unfolding psize_def by simp | |
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 103 | |
| 56778 | 104 | lemma poly_offset: | 
| 105 | fixes p :: "'a::comm_ring_1 poly" | |
| 106 | shows "\<exists>q. psize q = psize p \<and> (\<forall>x. poly q x = poly p (a + x))" | |
| 77282 | 107 | by (metis degree_offset_poly offset_poly_eq_0_iff poly_offset_poly psize_def) | 
| 26123 | 108 | |
| 60424 | 109 | text \<open>An alternative useful formulation of completeness of the reals\<close> | 
| 56778 | 110 | lemma real_sup_exists: | 
| 111 | assumes ex: "\<exists>x. P x" | |
| 112 | and bz: "\<exists>z. \<forall>x. P x \<longrightarrow> x < z" | |
| 113 | shows "\<exists>s::real. \<forall>y. (\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < s" | |
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54230diff
changeset | 114 | proof | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54230diff
changeset | 115 | from bz have "bdd_above (Collect P)" | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54230diff
changeset | 116 | by (force intro: less_imp_le) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54230diff
changeset | 117 | then show "\<forall>y. (\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < Sup (Collect P)" | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54230diff
changeset | 118 | using ex bz by (subst less_cSup_iff) auto | 
| 26123 | 119 | qed | 
| 120 | ||
| 60424 | 121 | |
| 122 | subsection \<open>Fundamental theorem of algebra\<close> | |
| 123 | ||
| 124 | lemma unimodular_reduce_norm: | |
| 26123 | 125 | assumes md: "cmod z = 1" | 
| 63589 | 126 | shows "cmod (z + 1) < 1 \<or> cmod (z - 1) < 1 \<or> cmod (z + \<i>) < 1 \<or> cmod (z - \<i>) < 1" | 
| 56778 | 127 | proof - | 
| 128 | obtain x y where z: "z = Complex x y " | |
| 129 | by (cases z) auto | |
| 130 | from md z have xy: "x\<^sup>2 + y\<^sup>2 = 1" | |
| 131 | by (simp add: cmod_def) | |
| 63589 | 132 | have False if "cmod (z + 1) \<ge> 1" "cmod (z - 1) \<ge> 1" "cmod (z + \<i>) \<ge> 1" "cmod (z - \<i>) \<ge> 1" | 
| 60557 | 133 | proof - | 
| 77303 | 134 | from that z xy have *: "2 * x \<le> 1" "2 * x \<ge> -1" "2 * y \<le> 1" "2 * y \<ge> -1" | 
| 29667 | 135 | by (simp_all add: cmod_def power2_eq_square algebra_simps) | 
| 61945 | 136 | then have "\<bar>2 * x\<bar> \<le> 1" "\<bar>2 * y\<bar> \<le> 1" | 
| 56778 | 137 | by simp_all | 
| 61945 | 138 | then have "\<bar>2 * x\<bar>\<^sup>2 \<le> 1\<^sup>2" "\<bar>2 * y\<bar>\<^sup>2 \<le> 1\<^sup>2" | 
| 77282 | 139 | by (metis abs_square_le_1 one_power2 power2_abs)+ | 
| 77303 | 140 | with xy * show ?thesis | 
| 141 | by (smt (verit, best) four_x_squared square_le_1) | |
| 60557 | 142 | qed | 
| 56778 | 143 | then show ?thesis | 
| 77303 | 144 | by force | 
| 26123 | 145 | qed | 
| 146 | ||
| 61585 | 147 | text \<open>Hence we can always reduce modulus of \<open>1 + b z^n\<close> if nonzero\<close> | 
| 26123 | 148 | lemma reduce_poly_simple: | 
| 56778 | 149 | assumes b: "b \<noteq> 0" | 
| 150 | and n: "n \<noteq> 0" | |
| 26123 | 151 | shows "\<exists>z. cmod (1 + b * z^n) < 1" | 
| 56778 | 152 | using n | 
| 153 | proof (induct n rule: nat_less_induct) | |
| 26123 | 154 | fix n | 
| 56778 | 155 | assume IH: "\<forall>m<n. m \<noteq> 0 \<longrightarrow> (\<exists>z. cmod (1 + b * z ^ m) < 1)" | 
| 156 | assume n: "n \<noteq> 0" | |
| 26123 | 157 | let ?P = "\<lambda>z n. cmod (1 + b * z ^ n) < 1" | 
| 60457 | 158 | show "\<exists>z. ?P z n" | 
| 159 | proof cases | |
| 77282 | 160 | assume "even n" | 
| 161 | then obtain m where m: "n = 2 * m" and "m \<noteq> 0" "m < n" | |
| 162 | using n by auto | |
| 163 | with IH obtain z where z: "?P z m" | |
| 56778 | 164 | by blast | 
| 56795 | 165 | from z have "?P (csqrt z) n" | 
| 60457 | 166 | by (simp add: m power_mult) | 
| 167 | then show ?thesis .. | |
| 168 | next | |
| 169 | assume "odd n" | |
| 170 | then have "\<exists>m. n = Suc (2 * m)" | |
| 56778 | 171 | by presburger+ | 
| 56795 | 172 | then obtain m where m: "n = Suc (2 * m)" | 
| 56778 | 173 | by blast | 
| 77303 | 174 | have 0: "cmod (complex_of_real (cmod b) / b) = 1" | 
| 60457 | 175 | using b by (simp add: norm_divide) | 
| 26123 | 176 | have "\<exists>v. cmod (complex_of_real (cmod b) / b + v^n) < 1" | 
| 77282 | 177 | proof (cases "cmod (complex_of_real (cmod b) / b + 1) < 1") | 
| 178 | case True | |
| 179 | then show ?thesis | |
| 180 | by (metis power_one) | |
| 181 | next | |
| 77303 | 182 | case F1: False | 
| 77282 | 183 | show ?thesis | 
| 184 | proof (cases "cmod (complex_of_real (cmod b) / b - 1) < 1") | |
| 185 | case True | |
| 186 | with \<open>odd n\<close> show ?thesis | |
| 187 | by (metis add_uminus_conv_diff neg_one_odd_power) | |
| 188 | next | |
| 77303 | 189 | case F2: False | 
| 77282 | 190 | show ?thesis | 
| 191 | proof (cases "cmod (complex_of_real (cmod b) / b + \<i>) < 1") | |
| 77303 | 192 | case T1: True | 
| 77282 | 193 | show ?thesis | 
| 194 | proof (cases "even m") | |
| 195 | case True | |
| 196 | with T1 show ?thesis | |
| 197 | by (rule_tac x="\<i>" in exI) (simp add: m power_mult) | |
| 198 | next | |
| 199 | case False | |
| 200 | with T1 show ?thesis | |
| 201 | by (rule_tac x="- \<i>" in exI) (simp add: m power_mult) | |
| 202 | qed | |
| 203 | next | |
| 204 | case False | |
| 77341 | 205 | then have lt1: "cmod (of_real (cmod b) / b - \<i>) < 1" | 
| 206 | using "0" F1 F2 unimodular_reduce_norm by blast | |
| 207 | show ?thesis | |
| 208 | proof (cases "even m") | |
| 209 | case True | |
| 210 | with m lt1 show ?thesis | |
| 211 | by (rule_tac x="- \<i>" in exI) (simp add: power_mult) | |
| 212 | next | |
| 213 | case False | |
| 214 | with m lt1 show ?thesis | |
| 215 | by (rule_tac x="\<i>" in exI) (simp add: power_mult) | |
| 216 | qed | |
| 77282 | 217 | qed | 
| 218 | qed | |
| 219 | qed | |
| 56778 | 220 | then obtain v where v: "cmod (complex_of_real (cmod b) / b + v^n) < 1" | 
| 221 | by blast | |
| 26123 | 222 | let ?w = "v / complex_of_real (root n (cmod b))" | 
| 60457 | 223 | from odd_real_root_pow[OF \<open>odd n\<close>, of "cmod b"] | 
| 77303 | 224 | have 1: "?w ^ n = v^n / complex_of_real (cmod b)" | 
| 56889 
48a745e1bde7
avoid the Complex constructor, use the more natural Re/Im view; moved csqrt to Complex.
 hoelzl parents: 
56795diff
changeset | 225 | by (simp add: power_divide of_real_power[symmetric]) | 
| 77303 | 226 | have 2:"cmod (complex_of_real (cmod b) / b) = 1" | 
| 56778 | 227 | using b by (simp add: norm_divide) | 
| 77303 | 228 | then have 3: "cmod (complex_of_real (cmod b) / b) \<ge> 0" | 
| 56778 | 229 | by simp | 
| 77303 | 230 | have 4: "cmod (complex_of_real (cmod b) / b) * | 
| 56778 | 231 | cmod (1 + b * (v ^ n / complex_of_real (cmod b))) < | 
| 232 | cmod (complex_of_real (cmod b) / b) * 1" | |
| 49962 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 webertj parents: 
46240diff
changeset | 233 | apply (simp only: norm_mult[symmetric] distrib_left) | 
| 56778 | 234 | using b v | 
| 77303 | 235 | apply (simp add: 2) | 
| 56778 | 236 | done | 
| 77282 | 237 | show ?thesis | 
| 77303 | 238 | by (metis 1 mult_left_less_imp_less[OF 4 3]) | 
| 60457 | 239 | qed | 
| 26123 | 240 | qed | 
| 241 | ||
| 60424 | 242 | text \<open>Bolzano-Weierstrass type property for closed disc in complex plane.\<close> | 
| 26123 | 243 | |
| 56778 | 244 | lemma metric_bound_lemma: "cmod (x - y) \<le> \<bar>Re x - Re y\<bar> + \<bar>Im x - Im y\<bar>" | 
| 56795 | 245 | using real_sqrt_sum_squares_triangle_ineq[of "Re x - Re y" 0 0 "Im x - Im y"] | 
| 26123 | 246 | unfolding cmod_def by simp | 
| 247 | ||
| 69529 | 248 | lemma Bolzano_Weierstrass_complex_disc: | 
| 26123 | 249 | assumes r: "\<forall>n. cmod (s n) \<le> r" | 
| 66447 
a1f5c5c26fa6
Replaced subseq with strict_mono
 eberlm <eberlm@in.tum.de> parents: 
65486diff
changeset | 250 | shows "\<exists>f z. strict_mono (f :: nat \<Rightarrow> nat) \<and> (\<forall>e >0. \<exists>N. \<forall>n \<ge> N. cmod (s (f n) - z) < e)" | 
| 60424 | 251 | proof - | 
| 56778 | 252 | from seq_monosub[of "Re \<circ> s"] | 
| 66447 
a1f5c5c26fa6
Replaced subseq with strict_mono
 eberlm <eberlm@in.tum.de> parents: 
65486diff
changeset | 253 | obtain f where f: "strict_mono f" "monoseq (\<lambda>n. Re (s (f n)))" | 
| 26123 | 254 | unfolding o_def by blast | 
| 56778 | 255 | from seq_monosub[of "Im \<circ> s \<circ> f"] | 
| 66447 
a1f5c5c26fa6
Replaced subseq with strict_mono
 eberlm <eberlm@in.tum.de> parents: 
65486diff
changeset | 256 | obtain g where g: "strict_mono g" "monoseq (\<lambda>n. Im (s (f (g n))))" | 
| 56778 | 257 | unfolding o_def by blast | 
| 258 | let ?h = "f \<circ> g" | |
| 77303 | 259 | have "r \<ge> 0" | 
| 260 | by (meson norm_ge_zero order_trans r) | |
| 261 | have "\<forall>n. r + 1 \<ge> \<bar>Re (s n)\<bar>" | |
| 77282 | 262 | by (smt (verit, ccfv_threshold) abs_Re_le_cmod r) | 
| 77303 | 263 | then have conv1: "convergent (\<lambda>n. Re (s (f n)))" | 
| 264 | by (metis Bseq_monoseq_convergent f(2) BseqI' real_norm_def) | |
| 265 | have "\<forall>n. r + 1 \<ge> \<bar>Im (s n)\<bar>" | |
| 77282 | 266 | by (smt (verit) abs_Im_le_cmod r) | 
| 77303 | 267 | then have conv2: "convergent (\<lambda>n. Im (s (f (g n))))" | 
| 268 | by (metis Bseq_monoseq_convergent g(2) BseqI' real_norm_def) | |
| 26123 | 269 | |
| 77303 | 270 | obtain x where x: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar>Re (s (f n)) - x\<bar> < r" | 
| 271 | using conv1[unfolded convergent_def] LIMSEQ_iff real_norm_def by metis | |
| 77282 | 272 | obtain y where y: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar>Im (s (f (g n))) - y\<bar> < r" | 
| 77303 | 273 | using conv2[unfolded convergent_def] LIMSEQ_iff real_norm_def by metis | 
| 26123 | 274 | let ?w = "Complex x y" | 
| 66447 
a1f5c5c26fa6
Replaced subseq with strict_mono
 eberlm <eberlm@in.tum.de> parents: 
65486diff
changeset | 275 | from f(1) g(1) have hs: "strict_mono ?h" | 
| 
a1f5c5c26fa6
Replaced subseq with strict_mono
 eberlm <eberlm@in.tum.de> parents: 
65486diff
changeset | 276 | unfolding strict_mono_def by auto | 
| 60557 | 277 | have "\<exists>N. \<forall>n\<ge>N. cmod (s (?h n) - ?w) < e" if "e > 0" for e | 
| 278 | proof - | |
| 279 | from that have e2: "e/2 > 0" | |
| 56795 | 280 | by simp | 
| 77282 | 281 | from x y e2 | 
| 56778 | 282 | obtain N1 N2 where N1: "\<forall>n\<ge>N1. \<bar>Re (s (f n)) - x\<bar> < e / 2" | 
| 56795 | 283 | and N2: "\<forall>n\<ge>N2. \<bar>Im (s (f (g n))) - y\<bar> < e / 2" | 
| 284 | by blast | |
| 60557 | 285 | have "cmod (s (?h n) - ?w) < e" if "n \<ge> N1 + N2" for n | 
| 286 | proof - | |
| 287 | from that have nN1: "g n \<ge> N1" and nN2: "n \<ge> N2" | |
| 56778 | 288 | using seq_suble[OF g(1), of n] by arith+ | 
| 60557 | 289 | show ?thesis | 
| 77282 | 290 | using metric_bound_lemma[of "s (f (g n))" ?w] N1 N2 nN1 nN2 by fastforce | 
| 60557 | 291 | qed | 
| 292 | then show ?thesis by blast | |
| 293 | qed | |
| 56778 | 294 | with hs show ?thesis by blast | 
| 26123 | 295 | qed | 
| 296 | ||
| 60424 | 297 | text \<open>Polynomial is continuous.\<close> | 
| 26123 | 298 | |
| 299 | lemma poly_cont: | |
| 56778 | 300 |   fixes p :: "'a::{comm_semiring_0,real_normed_div_algebra} poly"
 | 
| 30488 | 301 | assumes ep: "e > 0" | 
| 55735 
81ba62493610
generalised some results using type classes
 paulson <lp15@cam.ac.uk> parents: 
55734diff
changeset | 302 | shows "\<exists>d >0. \<forall>w. 0 < norm (w - z) \<and> norm (w - z) < d \<longrightarrow> norm (poly p w - poly p z) < e" | 
| 56778 | 303 | proof - | 
| 77341 | 304 | obtain q where "degree q = degree p" and q: "\<And>w. poly p w = poly q (w - z)" | 
| 305 | by (metis add.commute degree_offset_poly diff_add_cancel poly_offset_poly) | |
| 306 | show ?thesis unfolding q | |
| 56778 | 307 | proof (induct q) | 
| 308 | case 0 | |
| 309 | then show ?case | |
| 310 | using ep by auto | |
| 26123 | 311 | next | 
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 312 | case (pCons c cs) | 
| 63060 | 313 | obtain m where m: "m > 0" "norm z \<le> 1 \<Longrightarrow> norm (poly cs z) \<le> m" for z | 
| 77341 | 314 | using poly_bound_exists[of 1 "cs"] by blast | 
| 77303 | 315 | with ep have em0: "e/m > 0" | 
| 56778 | 316 | by (simp add: field_simps) | 
| 317 | obtain d where d: "d > 0" "d < 1" "d < e / m" | |
| 77341 | 318 | by (meson em0 field_lbound_gt_zero zero_less_one) | 
| 319 | then have "\<And>w. norm (w - z) < d \<Longrightarrow> norm (w - z) * norm (poly cs (w - z)) < e" | |
| 320 | by (smt (verit, del_insts) m mult_left_mono norm_ge_zero pos_less_divide_eq) | |
| 321 | with d show ?case | |
| 322 | by (force simp add: norm_mult) | |
| 56778 | 323 | qed | 
| 26123 | 324 | qed | 
| 325 | ||
| 60424 | 326 | text \<open>Hence a polynomial attains minimum on a closed disc | 
| 327 | in the complex plane.\<close> | |
| 56778 | 328 | lemma poly_minimum_modulus_disc: "\<exists>z. \<forall>w. cmod w \<le> r \<longrightarrow> cmod (poly p z) \<le> cmod (poly p w)" | 
| 329 | proof - | |
| 60424 | 330 | show ?thesis | 
| 331 | proof (cases "r \<ge> 0") | |
| 332 | case False | |
| 333 | then show ?thesis | |
| 56778 | 334 | by (metis norm_ge_zero order.trans) | 
| 60424 | 335 | next | 
| 336 | case True | |
| 56778 | 337 | then have mth1: "\<exists>x z. cmod z \<le> r \<and> cmod (poly p z) = - x" | 
| 77303 | 338 | by (metis add.inverse_inverse norm_zero) | 
| 339 | obtain s where s: "\<forall>y. (\<exists>x. (\<exists>z. cmod z \<le> r \<and> cmod (poly p z) = - x) \<and> y < x) \<longleftrightarrow> y < s" | |
| 77341 | 340 | by (smt (verit, del_insts) real_sup_exists[OF mth1] norm_zero zero_less_norm_iff) | 
| 77303 | 341 | |
| 56778 | 342 | let ?m = "- s" | 
| 77341 | 343 | have s1: "(\<exists>z. cmod z \<le> r \<and> - (- cmod (poly p z)) < y) \<longleftrightarrow> ?m < y" for y | 
| 77282 | 344 | by (metis add.inverse_inverse minus_less_iff s) | 
| 77341 | 345 | then have s1m: "\<And>z. cmod z \<le> r \<Longrightarrow> cmod (poly p z) \<ge> ?m" | 
| 346 | by force | |
| 60557 | 347 | have "\<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)" for n | 
| 77341 | 348 | using s1[of "?m + 1/real (Suc n)"] by simp | 
| 77303 | 349 | then obtain g where g: "\<forall>n. cmod (g n) \<le> r" "\<forall>n. cmod (poly p (g n)) <?m + 1 /real(Suc n)" | 
| 350 | by metis | |
| 69529 | 351 | from Bolzano_Weierstrass_complex_disc[OF g(1)] | 
| 77341 | 352 | obtain f::"nat \<Rightarrow> nat" and z where fz: "strict_mono f" "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. cmod (g (f n) - z) < e" | 
| 30488 | 353 | by blast | 
| 56778 | 354 |     {
 | 
| 355 | fix w | |
| 26123 | 356 | assume wr: "cmod w \<le> r" | 
| 357 | let ?e = "\<bar>cmod (poly p z) - ?m\<bar>" | |
| 56778 | 358 |       {
 | 
| 359 | assume e: "?e > 0" | |
| 56795 | 360 | then have e2: "?e/2 > 0" | 
| 361 | by simp | |
| 77341 | 362 | with poly_cont obtain d | 
| 363 | where "d > 0" and d: "\<And>w. 0<cmod (w - z)\<and> cmod(w - z) < d \<longrightarrow> cmod(poly p w - poly p z) < ?e/2" | |
| 56778 | 364 | by blast | 
| 77303 | 365 | have 1: "cmod(poly p w - poly p z) < ?e / 2" if w: "cmod (w - z) < d" for w | 
| 77341 | 366 | using d[of w] w e by (cases "w = z") simp_all | 
| 367 | from fz(2) \<open>d > 0\<close> obtain N1 where N1: "\<forall>n\<ge>N1. cmod (g (f n) - z) < d" | |
| 56778 | 368 | by blast | 
| 77303 | 369 | from reals_Archimedean2 obtain N2 :: nat where N2: "2/?e < real N2" | 
| 56778 | 370 | by blast | 
| 77303 | 371 | have 2: "cmod (poly p (g (f (N1 + N2))) - poly p z) < ?e/2" | 
| 372 | using N1 1 by auto | |
| 373 | have 0: "a < e2 \<Longrightarrow> \<bar>b - m\<bar> < e2 \<Longrightarrow> 2 * e2 \<le> \<bar>b - m\<bar> + a \<Longrightarrow> False" | |
| 60424 | 374 | for a b e2 m :: real | 
| 375 | by arith | |
| 56795 | 376 | from seq_suble[OF fz(1), of "N1 + N2"] | 
| 77303 | 377 | have 00: "?m + 1 / real (Suc (f (N1 + N2))) \<le> ?m + 1 / real (Suc (N1 + N2))" | 
| 378 | by (simp add: frac_le) | |
| 379 | from N2 e2 less_imp_inverse_less[of "2/?e" "real (Suc (N1 + N2))"] | |
| 56778 | 380 | have "?e/2 > 1/ real (Suc (N1 + N2))" | 
| 381 | by (simp add: inverse_eq_divide) | |
| 77303 | 382 | with order_less_le_trans[OF _ 00] | 
| 383 | have 1: "\<bar>cmod (poly p (g (f (N1 + N2)))) - ?m\<bar> < ?e/2" | |
| 77341 | 384 | using g s1 by (smt (verit)) | 
| 77303 | 385 | with 0[OF 2] have False | 
| 77282 | 386 | by (smt (verit) field_sum_of_halves norm_triangle_ineq3) | 
| 56778 | 387 | } | 
| 388 | then have "?e = 0" | |
| 389 | by auto | |
| 390 | with s1m[OF wr] have "cmod (poly p z) \<le> cmod (poly p w)" | |
| 391 | by simp | |
| 392 | } | |
| 60424 | 393 | then show ?thesis by blast | 
| 394 | qed | |
| 26123 | 395 | qed | 
| 396 | ||
| 60424 | 397 | text \<open>Nonzero polynomial in z goes to infinity as z does.\<close> | 
| 26123 | 398 | |
| 399 | lemma poly_infinity: | |
| 56778 | 400 |   fixes p:: "'a::{comm_semiring_0,real_normed_div_algebra} poly"
 | 
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 401 | assumes ex: "p \<noteq> 0" | 
| 55735 
81ba62493610
generalised some results using type classes
 paulson <lp15@cam.ac.uk> parents: 
55734diff
changeset | 402 | shows "\<exists>r. \<forall>z. r \<le> norm z \<longrightarrow> d \<le> norm (poly (pCons a p) z)" | 
| 56778 | 403 | using ex | 
| 404 | proof (induct p arbitrary: a d) | |
| 56795 | 405 | case 0 | 
| 406 | then show ?case by simp | |
| 407 | next | |
| 30488 | 408 | case (pCons c cs a d) | 
| 56795 | 409 | show ?case | 
| 410 | proof (cases "cs = 0") | |
| 411 | case False | |
| 56778 | 412 | with pCons.hyps obtain r where r: "\<forall>z. r \<le> norm z \<longrightarrow> d + norm a \<le> norm (poly (pCons c cs) z)" | 
| 413 | by blast | |
| 26123 | 414 | let ?r = "1 + \<bar>r\<bar>" | 
| 60557 | 415 | have "d \<le> norm (poly (pCons a (pCons c cs)) z)" if "1 + \<bar>r\<bar> \<le> norm z" for z | 
| 416 | proof - | |
| 77282 | 417 | have "d \<le> norm(z * poly (pCons c cs) z) - norm a" | 
| 77303 | 418 | by (smt (verit, best) norm_ge_zero mult_less_cancel_right2 norm_mult r that) | 
| 419 | with norm_diff_ineq add.commute | |
| 420 | show ?thesis | |
| 421 | by (metis order.trans poly_pCons) | |
| 60557 | 422 | qed | 
| 56795 | 423 | then show ?thesis by blast | 
| 424 | next | |
| 425 | case True | |
| 60424 | 426 | have "d \<le> norm (poly (pCons a (pCons c cs)) z)" | 
| 77303 | 427 | if "(\<bar>d\<bar> + norm a) / norm c \<le> norm z" for z :: 'a | 
| 60424 | 428 | proof - | 
| 77303 | 429 | have "\<bar>d\<bar> + norm a \<le> norm (z * c)" | 
| 430 | by (metis that True norm_mult pCons.hyps(1) pos_divide_le_eq zero_less_norm_iff) | |
| 431 | also have "\<dots> \<le> norm (a + z * c) + norm a" | |
| 432 | by (simp add: add.commute norm_add_leD) | |
| 433 | finally show ?thesis | |
| 434 | using True by auto | |
| 60424 | 435 | qed | 
| 56795 | 436 | then show ?thesis by blast | 
| 437 | qed | |
| 438 | qed | |
| 26123 | 439 | |
| 60424 | 440 | text \<open>Hence polynomial's modulus attains its minimum somewhere.\<close> | 
| 56778 | 441 | lemma poly_minimum_modulus: "\<exists>z.\<forall>w. cmod (poly p z) \<le> cmod (poly p w)" | 
| 442 | proof (induct p) | |
| 443 | case 0 | |
| 444 | then show ?case by simp | |
| 445 | next | |
| 30488 | 446 | case (pCons c cs) | 
| 56778 | 447 | show ?case | 
| 448 | proof (cases "cs = 0") | |
| 449 | case False | |
| 450 | from poly_infinity[OF False, of "cmod (poly (pCons c cs) 0)" c] | |
| 63060 | 451 | obtain r where r: "cmod (poly (pCons c cs) 0) \<le> cmod (poly (pCons c cs) z)" | 
| 452 | if "r \<le> cmod z" for z | |
| 56778 | 453 | by blast | 
| 77282 | 454 | from poly_minimum_modulus_disc[of "\<bar>r\<bar>" "pCons c cs"] show ?thesis | 
| 77303 | 455 | by (smt (verit, del_insts) order.trans linorder_linear r) | 
| 77282 | 456 | qed (use pCons.hyps in auto) | 
| 56778 | 457 | qed | 
| 26123 | 458 | |
| 60424 | 459 | text \<open>Constant function (non-syntactic characterization).\<close> | 
| 56795 | 460 | definition "constant f \<longleftrightarrow> (\<forall>x y. f x = f y)" | 
| 26123 | 461 | |
| 56778 | 462 | lemma nonconstant_length: "\<not> constant (poly p) \<Longrightarrow> psize p \<ge> 2" | 
| 463 | by (induct p) (auto simp: constant_def psize_def) | |
| 30488 | 464 | |
| 56795 | 465 | lemma poly_replicate_append: "poly (monom 1 n * p) (x::'a::comm_ring_1) = x^n * poly p x" | 
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 466 | by (simp add: poly_monom) | 
| 26123 | 467 | |
| 60424 | 468 | text \<open>Decomposition of polynomial, skipping zero coefficients after the first.\<close> | 
| 26123 | 469 | |
| 470 | lemma poly_decompose_lemma: | |
| 56778 | 471 | assumes nz: "\<not> (\<forall>z. z \<noteq> 0 \<longrightarrow> poly p z = (0::'a::idom))" | 
| 56795 | 472 | shows "\<exists>k a q. a \<noteq> 0 \<and> Suc (psize q + k) = psize p \<and> (\<forall>z. poly p z = z^k * poly (pCons a q) z)" | 
| 56778 | 473 | unfolding psize_def | 
| 474 | using nz | |
| 475 | proof (induct p) | |
| 476 | case 0 | |
| 477 | then show ?case by simp | |
| 26123 | 478 | next | 
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 479 | case (pCons c cs) | 
| 56778 | 480 | show ?case | 
| 481 | proof (cases "c = 0") | |
| 482 | case True | |
| 483 | from pCons.hyps pCons.prems True show ?thesis | |
| 60424 | 484 | apply auto | 
| 26123 | 485 | apply (rule_tac x="k+1" in exI) | 
| 60557 | 486 | apply (rule_tac x="a" in exI) | 
| 487 | apply clarsimp | |
| 26123 | 488 | apply (rule_tac x="q" in exI) | 
| 56778 | 489 | apply auto | 
| 490 | done | |
| 77303 | 491 | qed force | 
| 26123 | 492 | qed | 
| 493 | ||
| 494 | lemma poly_decompose: | |
| 77303 | 495 | fixes p :: "'a::idom poly" | 
| 56776 | 496 | assumes nc: "\<not> constant (poly p)" | 
| 77303 | 497 | shows "\<exists>k a q. a \<noteq> 0 \<and> k \<noteq> 0 \<and> | 
| 30488 | 498 | psize q + k + 1 = psize p \<and> | 
| 77303 | 499 | (\<forall>z. poly p z = poly p 0 + z^k * poly (pCons a q) z)" | 
| 56776 | 500 | using nc | 
| 501 | proof (induct p) | |
| 502 | case 0 | |
| 503 | then show ?case | |
| 504 | by (simp add: constant_def) | |
| 26123 | 505 | next | 
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 506 | case (pCons c cs) | 
| 60557 | 507 | have "\<not> (\<forall>z. z \<noteq> 0 \<longrightarrow> poly cs z = 0)" | 
| 77282 | 508 | by (smt (verit) constant_def mult_eq_0_iff pCons.prems poly_pCons) | 
| 60557 | 509 | from poly_decompose_lemma[OF this] | 
| 77303 | 510 | obtain k a q where *: "a \<noteq> 0 \<and> | 
| 511 | Suc (psize q + k) = psize cs \<and> (\<forall>z. poly cs z = z ^ k * poly (pCons a q) z)" | |
| 512 | by blast | |
| 513 | then have "psize q + k + 2 = psize (pCons c cs)" | |
| 514 | by (auto simp add: psize_def split: if_splits) | |
| 515 | then show ?case | |
| 516 | using "*" by force | |
| 26123 | 517 | qed | 
| 518 | ||
| 60424 | 519 | text \<open>Fundamental theorem of algebra\<close> | 
| 26123 | 520 | |
| 80061 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 521 | theorem fundamental_theorem_of_algebra: | 
| 56776 | 522 | assumes nc: "\<not> constant (poly p)" | 
| 26123 | 523 | shows "\<exists>z::complex. poly p z = 0" | 
| 56776 | 524 | using nc | 
| 525 | proof (induct "psize p" arbitrary: p rule: less_induct) | |
| 34915 | 526 | case less | 
| 26123 | 527 | let ?p = "poly p" | 
| 528 | let ?ths = "\<exists>z. ?p z = 0" | |
| 529 | ||
| 34915 | 530 | from nonconstant_length[OF less(2)] have n2: "psize p \<ge> 2" . | 
| 56776 | 531 | from poly_minimum_modulus obtain c where c: "\<forall>w. cmod (?p c) \<le> cmod (?p w)" | 
| 532 | by blast | |
| 56778 | 533 | |
| 534 | show ?ths | |
| 535 | proof (cases "?p c = 0") | |
| 536 | case True | |
| 537 | then show ?thesis by blast | |
| 538 | next | |
| 539 | case False | |
| 77303 | 540 | obtain q where q: "psize q = psize p" "\<forall>x. poly q x = ?p (c + x)" | 
| 541 | using poly_offset[of p c] by blast | |
| 56778 | 542 | then have qnc: "\<not> constant (poly q)" | 
| 77303 | 543 | by (metis (no_types, opaque_lifting) add.commute constant_def diff_add_cancel less.prems) | 
| 56778 | 544 | from q(2) have pqc0: "?p c = poly q 0" | 
| 545 | by simp | |
| 546 | from c pqc0 have cq0: "\<forall>w. cmod (poly q 0) \<le> cmod (?p w)" | |
| 547 | by simp | |
| 26123 | 548 | let ?a0 = "poly q 0" | 
| 60424 | 549 | from False pqc0 have a00: "?a0 \<noteq> 0" | 
| 56778 | 550 | by simp | 
| 551 | from a00 have qr: "\<forall>z. poly q z = poly (smult (inverse ?a0) q) z * ?a0" | |
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 552 | by simp | 
| 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 553 | let ?r = "smult (inverse ?a0) q" | 
| 29538 | 554 | have lgqr: "psize q = psize ?r" | 
| 77282 | 555 | by (simp add: a00 psize_def) | 
| 77303 | 556 | have rnc: "\<not> constant (poly ?r)" | 
| 557 | using constant_def qnc qr by fastforce | |
| 558 | have r01: "poly ?r 0 = 1" | |
| 559 | by (simp add: a00) | |
| 60424 | 560 | have mrmq_eq: "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w) < cmod ?a0" for w | 
| 77282 | 561 | by (smt (verit, del_insts) a00 mult_less_cancel_right2 norm_mult qr zero_less_norm_iff) | 
| 30488 | 562 | from poly_decompose[OF rnc] obtain k a s where | 
| 56778 | 563 | kas: "a \<noteq> 0" "k \<noteq> 0" "psize s + k + 1 = psize ?r" | 
| 564 | "\<forall>z. poly ?r z = poly ?r 0 + z^k* poly (pCons a s) z" by blast | |
| 60424 | 565 | have "\<exists>w. cmod (poly ?r w) < 1" | 
| 566 | proof (cases "psize p = k + 1") | |
| 77341 | 567 | case True | 
| 568 | with kas q have s0: "s = 0" | |
| 569 | by (simp add: lgqr) | |
| 77303 | 570 | with reduce_poly_simple kas show ?thesis | 
| 571 | by (metis mult.commute mult.right_neutral poly_1 poly_smult r01 smult_one) | |
| 60424 | 572 | next | 
| 573 | case False note kn = this | |
| 56778 | 574 | from kn kas(3) q(1) lgqr have k1n: "k + 1 < psize p" | 
| 575 | by simp | |
| 77303 | 576 | have 01: "\<not> constant (poly (pCons 1 (monom a (k - 1))))" | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32456diff
changeset | 577 | unfolding constant_def poly_pCons poly_monom | 
| 77282 | 578 | by (metis add_cancel_left_right kas(1) mult.commute mult_cancel_right2 power_one) | 
| 77341 | 579 | have 02: "k + 1 = psize (pCons 1 (monom a (k - 1)))" | 
| 580 | using kas by (simp add: psize_def degree_monom_eq) | |
| 581 | from less(1) [OF _ 01] k1n 02 | |
| 26123 | 582 | obtain w where w: "1 + w^k * a = 0" | 
| 77303 | 583 | by (metis kas(2) mult.commute mult.left_commute poly_monom poly_pCons power_eq_if) | 
| 30488 | 584 | from poly_bound_exists[of "cmod w" s] obtain m where | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32456diff
changeset | 585 | m: "m > 0" "\<forall>z. cmod z \<le> cmod w \<longrightarrow> cmod (poly s z) \<le> m" by blast | 
| 77303 | 586 | have "w \<noteq> 0" | 
| 56795 | 587 | using kas(2) w by (auto simp add: power_0_left) | 
| 77303 | 588 | from w have wm1: "w^k * a = - 1" | 
| 589 | by (simp add: add_eq_0_iff) | |
| 30488 | 590 | have inv0: "0 < inverse (cmod w ^ (k + 1) * m)" | 
| 77303 | 591 | by (simp add: \<open>w \<noteq> 0\<close> m(1)) | 
| 68527 
2f4e2aab190a
Generalising and renaming some basic results
 paulson <lp15@cam.ac.uk> parents: 
66447diff
changeset | 592 | with field_lbound_gt_zero[OF zero_less_one] obtain t where | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32456diff
changeset | 593 | t: "t > 0" "t < 1" "t < inverse (cmod w ^ (k + 1) * m)" by blast | 
| 26123 | 594 | let ?ct = "complex_of_real t" | 
| 595 | let ?w = "?ct * w" | |
| 56778 | 596 | have "1 + ?w^k * (a + ?w * poly s ?w) = 1 + ?ct^k * (w^k * a) + ?w^k * ?w * poly s ?w" | 
| 597 | using kas(1) by (simp add: algebra_simps power_mult_distrib) | |
| 26123 | 598 | also have "\<dots> = complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w" | 
| 56778 | 599 | unfolding wm1 by simp | 
| 600 | finally have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) = | |
| 601 | cmod (complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w)" | |
| 55358 
85d81bc281d0
Simplified some proofs, deleting a lot of strange unused material at the end of the theory.
 paulson <lp15@cam.ac.uk> parents: 
54489diff
changeset | 602 | by metis | 
| 30488 | 603 | with norm_triangle_ineq[of "complex_of_real (1 - t^k)" "?w^k * ?w * poly s ?w"] | 
| 77303 | 604 | have 11: "cmod (1 + ?w^k * (a + ?w * poly s ?w)) \<le> \<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w)" | 
| 56778 | 605 | unfolding norm_of_real by simp | 
| 606 | have ath: "\<And>x t::real. 0 \<le> x \<Longrightarrow> x < t \<Longrightarrow> t \<le> 1 \<Longrightarrow> \<bar>1 - t\<bar> + x < 1" | |
| 607 | by arith | |
| 77303 | 608 | have tw: "cmod ?w \<le> cmod w" | 
| 609 | by (smt (verit) mult_le_cancel_right2 norm_ge_zero norm_mult norm_of_real t) | |
| 610 | have "t * (cmod w ^ (k + 1) * m) < 1" | |
| 611 | by (smt (verit, best) inv0 inverse_positive_iff_positive left_inverse mult_strict_right_mono t(3)) | |
| 77341 | 612 | with zero_less_power[OF t(1), of k] have 30: "t^k * (t* (cmod w ^ (k + 1) * m)) < t^k" | 
| 59557 | 613 | by simp | 
| 56778 | 614 | have "cmod (?w^k * ?w * poly s ?w) = t^k * (t* (cmod w ^ (k + 1) * cmod (poly s ?w)))" | 
| 77303 | 615 | using \<open>w \<noteq> 0\<close> t(1) by (simp add: algebra_simps norm_power norm_mult) | 
| 616 | with 30 have 120: "cmod (?w^k * ?w * poly s ?w) < t^k" | |
| 617 | by (smt (verit, ccfv_SIG) m(2) mult_left_mono norm_ge_zero t(1) tw zero_le_power) | |
| 618 | from power_strict_mono[OF t(2), of k] t(1) kas(2) have 121: "t^k \<le> 1" | |
| 55358 
85d81bc281d0
Simplified some proofs, deleting a lot of strange unused material at the end of the theory.
 paulson <lp15@cam.ac.uk> parents: 
54489diff
changeset | 619 | by auto | 
| 77303 | 620 | from ath[OF norm_ge_zero[of "?w^k * ?w * poly s ?w"] 120 121] | 
| 77341 | 621 | show ?thesis | 
| 77303 | 622 | by (smt (verit) "11" kas(4) poly_pCons r01) | 
| 60424 | 623 | qed | 
| 624 | with cq0 q(2) show ?thesis | |
| 77303 | 625 | by (smt (verit) mrmq_eq) | 
| 56778 | 626 | qed | 
| 26123 | 627 | qed | 
| 628 | ||
| 60424 | 629 | text \<open>Alternative version with a syntactic notion of constant polynomial.\<close> | 
| 26123 | 630 | |
| 631 | lemma fundamental_theorem_of_algebra_alt: | |
| 56778 | 632 | assumes nc: "\<not> (\<exists>a l. a \<noteq> 0 \<and> l = 0 \<and> p = pCons a l)" | 
| 26123 | 633 | shows "\<exists>z. poly p z = (0::complex)" | 
| 77341 | 634 | proof (rule ccontr) | 
| 635 | assume N: "\<nexists>z. poly p z = 0" | |
| 636 | then have "\<not> constant (poly p)" | |
| 637 | unfolding constant_def | |
| 638 | by (metis (no_types, opaque_lifting) nc poly_pcompose pcompose_0' pcompose_const poly_0_coeff_0 | |
| 639 | poly_all_0_iff_0 poly_diff right_minus_eq) | |
| 640 | then show False | |
| 641 | using N fundamental_theorem_of_algebra by blast | |
| 56778 | 642 | qed | 
| 26123 | 643 | |
| 60424 | 644 | subsection \<open>Nullstellensatz, degrees and divisibility of polynomials\<close> | 
| 26123 | 645 | |
| 646 | lemma nullstellensatz_lemma: | |
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 647 | fixes p :: "complex poly" | 
| 26123 | 648 | assumes "\<forall>x. poly p x = 0 \<longrightarrow> poly q x = 0" | 
| 56776 | 649 | and "degree p = n" | 
| 650 | and "n \<noteq> 0" | |
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 651 | shows "p dvd (q ^ n)" | 
| 56776 | 652 | using assms | 
| 653 | proof (induct n arbitrary: p q rule: nat_less_induct) | |
| 654 | fix n :: nat | |
| 655 | fix p q :: "complex poly" | |
| 26123 | 656 | assume IH: "\<forall>m<n. \<forall>p q. | 
| 657 | (\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longrightarrow> | |
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 658 | degree p = m \<longrightarrow> m \<noteq> 0 \<longrightarrow> p dvd (q ^ m)" | 
| 30488 | 659 | and pq0: "\<forall>x. poly p x = 0 \<longrightarrow> poly q x = 0" | 
| 56778 | 660 | and dpn: "degree p = n" | 
| 661 | and n0: "n \<noteq> 0" | |
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 662 | from dpn n0 have pne: "p \<noteq> 0" by auto | 
| 60557 | 663 | show "p dvd (q ^ n)" | 
| 664 | proof (cases "\<exists>a. poly p a = 0") | |
| 665 | case True | |
| 666 | then obtain a where a: "poly p a = 0" .. | |
| 667 | have ?thesis if oa: "order a p \<noteq> 0" | |
| 60424 | 668 | proof - | 
| 26123 | 669 | let ?op = "order a p" | 
| 56778 | 670 | from pne have ap: "([:- a, 1:] ^ ?op) dvd p" "\<not> [:- a, 1:] ^ (Suc ?op) dvd p" | 
| 671 | using order by blast+ | |
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 672 | note oop = order_degree[OF pne, unfolded dpn] | 
| 60424 | 673 | show ?thesis | 
| 674 | proof (cases "q = 0") | |
| 675 | case True | |
| 676 | with n0 show ?thesis by (simp add: power_0_left) | |
| 677 | next | |
| 678 | case False | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32456diff
changeset | 679 | from pq0[rule_format, OF a, unfolded poly_eq_0_iff_dvd] | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32456diff
changeset | 680 | obtain r where r: "q = [:- a, 1:] * r" by (rule dvdE) | 
| 56778 | 681 | from ap(1) obtain s where s: "p = [:- a, 1:] ^ ?op * s" | 
| 682 | by (rule dvdE) | |
| 60424 | 683 | have sne: "s \<noteq> 0" | 
| 684 | using s pne by auto | |
| 685 | show ?thesis | |
| 686 | proof (cases "degree s = 0") | |
| 687 | case True | |
| 688 | then obtain k where kpn: "s = [:k:]" | |
| 51541 | 689 | by (cases s) (auto split: if_splits) | 
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 690 | from sne kpn have k: "k \<noteq> 0" by simp | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32456diff
changeset | 691 | let ?w = "([:1/k:] * ([:-a,1:] ^ (n - ?op))) * (r ^ n)" | 
| 77341 | 692 | have "q^n = [:- a, 1:] ^ n * r ^ n" | 
| 693 | using power_mult_distrib r by blast | |
| 694 | also have "... = [:- a, 1:] ^ order a p * [:k:] * ([:1 / k:] * [:- a, 1:] ^ (n - order a p) * r ^ n)" | |
| 695 | using k oop [of a] by (simp flip: power_add) | |
| 696 | also have "... = p * ?w" | |
| 697 | by (metis s kpn) | |
| 698 | finally show ?thesis | |
| 56795 | 699 | unfolding dvd_def by blast | 
| 60424 | 700 | next | 
| 701 | case False | |
| 702 | with sne dpn s oa have dsn: "degree s < n" | |
| 77341 | 703 | by (metis add_diff_cancel_right' degree_0 degree_linear_power degree_mult_eq gr0I zero_less_diff) | 
| 60557 | 704 | have "poly r x = 0" if h: "poly s x = 0" for x | 
| 705 | proof - | |
| 77341 | 706 | have "x \<noteq> a" | 
| 707 | by (metis ap(2) dvd_refl mult_dvd_mono poly_eq_0_iff_dvd power_Suc power_commutes s that) | |
| 708 | moreover have "poly p x = 0" | |
| 709 | by (metis (no_types) mult_eq_0_iff poly_mult s that) | |
| 710 | ultimately show ?thesis | |
| 711 | using pq0 r by auto | |
| 60557 | 712 | qed | 
| 77303 | 713 | with False IH dsn obtain u where u: "r ^ (degree s) = s * u" | 
| 60557 | 714 | by blast | 
| 715 | then have u': "\<And>x. poly s x * poly u x = poly r x ^ degree s" | |
| 716 | by (simp only: poly_mult[symmetric] poly_power[symmetric]) | |
| 77341 | 717 | have "q^n = [:- a, 1:] ^ n * r ^ n" | 
| 718 | using power_mult_distrib r by blast | |
| 719 | also have "... = [:- a, 1:] ^ order a p * (s * u * ([:- a, 1:] ^ (n - order a p) * r ^ (n - degree s)))" | |
| 720 | by (smt (verit, del_insts) s u mult_ac power_add add_diff_cancel_right' degree_linear_power degree_mult_eq dpn mult_zero_left) | |
| 721 | also have "... = p * (u * ([:-a,1:] ^ (n - ?op))) * (r ^ (n - degree s))" | |
| 722 | using s by force | |
| 723 | finally show ?thesis | |
| 724 | unfolding dvd_def by auto | |
| 60424 | 725 | qed | 
| 726 | qed | |
| 727 | qed | |
| 60557 | 728 | then show ?thesis | 
| 729 | using a order_root pne by blast | |
| 730 | next | |
| 731 | case False | |
| 732 | then show ?thesis | |
| 77282 | 733 | using dpn n0 fundamental_theorem_of_algebra_alt[of p] | 
| 734 | by fastforce | |
| 60557 | 735 | qed | 
| 26123 | 736 | qed | 
| 737 | ||
| 738 | lemma nullstellensatz_univariate: | |
| 30488 | 739 | "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow> | 
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 740 | p dvd (q ^ (degree p)) \<or> (p = 0 \<and> q = 0)" | 
| 56776 | 741 | proof - | 
| 60457 | 742 | consider "p = 0" | "p \<noteq> 0" "degree p = 0" | n where "p \<noteq> 0" "degree p = Suc n" | 
| 743 | by (cases "degree p") auto | |
| 744 | then show ?thesis | |
| 745 | proof cases | |
| 60567 | 746 | case p: 1 | 
| 77303 | 747 | then have "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow> q = 0" | 
| 52380 | 748 | by (auto simp add: poly_all_0_iff_0) | 
| 77303 | 749 | with p show ?thesis | 
| 750 | by force | |
| 60424 | 751 | next | 
| 60567 | 752 | case dp: 2 | 
| 77282 | 753 | then show ?thesis | 
| 77303 | 754 | by (meson dvd_trans is_unit_iff_degree poly_eq_0_iff_dvd unit_imp_dvd) | 
| 60457 | 755 | next | 
| 60567 | 756 | case dp: 3 | 
| 77282 | 757 | have False if "p dvd (q ^ (Suc n))" "poly p x = 0" "poly q x \<noteq> 0" for x | 
| 758 | by (metis dvd_trans poly_eq_0_iff_dvd poly_power power_eq_0_iff that) | |
| 60567 | 759 | with dp nullstellensatz_lemma[of p q "degree p"] show ?thesis | 
| 760 | by auto | |
| 60424 | 761 | qed | 
| 26123 | 762 | qed | 
| 763 | ||
| 60424 | 764 | text \<open>Useful lemma\<close> | 
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 765 | lemma constant_degree: | 
| 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 766 |   fixes p :: "'a::{idom,ring_char_0} poly"
 | 
| 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 767 | shows "constant (poly p) \<longleftrightarrow> degree p = 0" (is "?lhs = ?rhs") | 
| 26123 | 768 | proof | 
| 60557 | 769 | show ?rhs if ?lhs | 
| 770 | proof - | |
| 771 | from that[unfolded constant_def, rule_format, of _ "0"] | |
| 77282 | 772 | have "poly p = poly [:poly p 0:]" | 
| 60557 | 773 | by auto | 
| 774 | then show ?thesis | |
| 77282 | 775 | by (metis degree_pCons_0 poly_eq_poly_eq_iff) | 
| 60557 | 776 | qed | 
| 777 | show ?lhs if ?rhs | |
| 77303 | 778 | unfolding constant_def | 
| 779 | by (metis degree_eq_zeroE pcompose_const poly_0 poly_pcompose that) | |
| 26123 | 780 | qed | 
| 781 | ||
| 80061 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 782 | lemma complex_poly_decompose: | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 783 | "smult (lead_coeff p) (\<Prod>z|poly p z = 0. [:-z, 1:] ^ order z p) = (p :: complex poly)" | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 784 | proof (induction p rule: poly_root_order_induct) | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 785 | case (no_roots p) | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 786 | show ?case | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 787 | proof (cases "degree p = 0") | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 788 | case False | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 789 | hence "\<not>constant (poly p)" by (subst constant_degree) | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 790 | with fundamental_theorem_of_algebra and no_roots show ?thesis by blast | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 791 | qed (auto elim!: degree_eq_zeroE) | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 792 | next | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 793 | case (root p x n) | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 794 |   from root have *: "{z. poly ([:- x, 1:] ^ n * p) z = 0} = insert x {z. poly p z = 0}"
 | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 795 | by auto | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 796 | have "smult (lead_coeff ([:-x, 1:] ^ n * p)) | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 797 | (\<Prod>z|poly ([:-x,1:] ^ n * p) z = 0. [:-z, 1:] ^ order z ([:- x, 1:] ^ n * p)) = | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 798 | [:- x, 1:] ^ order x ([:- x, 1:] ^ n * p) * | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 799 |           smult (lead_coeff p) (\<Prod>z\<in>{z. poly p z = 0}. [:- z, 1:] ^ order z ([:- x, 1:] ^ n * p))"
 | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 800 | by (subst *, subst prod.insert) | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 801 | (insert root, auto intro: poly_roots_finite simp: mult_ac lead_coeff_mult lead_coeff_power) | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 802 | also have "order x ([:- x, 1:] ^ n * p) = n" | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 803 | using root by (subst order_mult) (auto simp: order_power_n_n order_0I) | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 804 |   also have "(\<Prod>z\<in>{z. poly p z = 0}. [:- z, 1:] ^ order z ([:- x, 1:] ^ n * p)) =
 | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 805 |                (\<Prod>z\<in>{z. poly p z = 0}. [:- z, 1:] ^ order z p)"
 | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 806 | proof (intro prod.cong refl, goal_cases) | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 807 | case (1 y) | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 808 | with root have "order y ([:-x,1:] ^ n) = 0" by (intro order_0I) auto | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 809 | thus ?case using root by (subst order_mult) auto | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 810 | qed | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 811 | also note root.IH | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 812 | finally show ?case . | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 813 | qed simp_all | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 814 | |
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 815 | instance complex :: alg_closed_field | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 816 | by standard (use fundamental_theorem_of_algebra constant_degree neq0_conv in blast) | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 817 | |
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 818 | lemma size_proots_complex: "size (proots (p :: complex poly)) = degree p" | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 819 | proof (cases "p = 0") | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 820 | case [simp]: False | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 821 | show "size (proots p) = degree p" | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 822 | by (subst (1 2) complex_poly_decompose [symmetric]) | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 823 | (simp add: proots_prod proots_power degree_prod_sum_eq degree_power_eq) | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 824 | qed auto | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 825 | |
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 826 | lemma complex_poly_decompose_multiset: | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 827 | "smult (lead_coeff p) (\<Prod>x\<in>#proots p. [:-x, 1:]) = (p :: complex poly)" | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 828 | proof (cases "p = 0") | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 829 | case False | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 830 | hence "(\<Prod>x\<in>#proots p. [:-x, 1:]) = (\<Prod>x | poly p x = 0. [:-x, 1:] ^ order x p)" | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 831 | by (subst image_prod_mset_multiplicity) simp_all | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 832 | also have "smult (lead_coeff p) \<dots> = p" | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 833 | by (rule complex_poly_decompose) | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 834 | finally show ?thesis . | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 835 | qed auto | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 836 | |
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 837 | lemma complex_poly_decompose': | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 838 | obtains root where "smult (lead_coeff p) (\<Prod>i<degree p. [:-root i, 1:]) = (p :: complex poly)" | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 839 | proof - | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 840 | obtain roots where roots: "mset roots = proots p" | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 841 | using ex_mset by blast | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 842 | |
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 843 | have "p = smult (lead_coeff p) (\<Prod>x\<in>#proots p. [:-x, 1:])" | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 844 | by (rule complex_poly_decompose_multiset [symmetric]) | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 845 | also have "(\<Prod>x\<in>#proots p. [:-x, 1:]) = (\<Prod>x\<leftarrow>roots. [:-x, 1:])" | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 846 | by (subst prod_mset_prod_list [symmetric]) (simp add: roots) | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 847 | also have "\<dots> = (\<Prod>i<length roots. [:-roots ! i, 1:])" | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 848 | by (subst prod.list_conv_set_nth) (auto simp: atLeast0LessThan) | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 849 | finally have eq: "p = smult (lead_coeff p) (\<Prod>i<length roots. [:-roots ! i, 1:])" . | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 850 | also have [simp]: "degree p = length roots" | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 851 | using roots by (subst eq) (auto simp: degree_prod_sum_eq) | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 852 | finally show ?thesis by (intro that[of "\<lambda>i. roots ! i"]) auto | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 853 | qed | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 854 | |
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 855 | lemma complex_poly_decompose_rsquarefree: | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 856 | assumes "rsquarefree p" | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 857 | shows "smult (lead_coeff p) (\<Prod>z|poly p z = 0. [:-z, 1:]) = (p :: complex poly)" | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 858 | proof (cases "p = 0") | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 859 | case False | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 860 | have "(\<Prod>z|poly p z = 0. [:-z, 1:]) = (\<Prod>z|poly p z = 0. [:-z, 1:] ^ order z p)" | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 861 | using assms False by (intro prod.cong) (auto simp: rsquarefree_root_order) | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 862 | also have "smult (lead_coeff p) \<dots> = p" | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 863 | by (rule complex_poly_decompose) | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 864 | finally show ?thesis . | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 865 | qed auto | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 866 | |
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
77341diff
changeset | 867 | |
| 60424 | 868 | text \<open>Arithmetic operations on multivariate polynomials.\<close> | 
| 26123 | 869 | |
| 30488 | 870 | lemma mpoly_base_conv: | 
| 56778 | 871 | fixes x :: "'a::comm_ring_1" | 
| 55735 
81ba62493610
generalised some results using type classes
 paulson <lp15@cam.ac.uk> parents: 
55734diff
changeset | 872 | shows "0 = poly 0 x" "c = poly [:c:] x" "x = poly [:0,1:] x" | 
| 
81ba62493610
generalised some results using type classes
 paulson <lp15@cam.ac.uk> parents: 
55734diff
changeset | 873 | by simp_all | 
| 26123 | 874 | |
| 30488 | 875 | lemma mpoly_norm_conv: | 
| 56778 | 876 | fixes x :: "'a::comm_ring_1" | 
| 56776 | 877 | shows "poly [:0:] x = poly 0 x" "poly [:poly 0 y:] x = poly 0 x" | 
| 878 | by simp_all | |
| 26123 | 879 | |
| 30488 | 880 | lemma mpoly_sub_conv: | 
| 56778 | 881 | fixes x :: "'a::comm_ring_1" | 
| 55735 
81ba62493610
generalised some results using type classes
 paulson <lp15@cam.ac.uk> parents: 
55734diff
changeset | 882 | shows "poly p x - poly q x = poly p x + -1 * poly q x" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53077diff
changeset | 883 | by simp | 
| 26123 | 884 | |
| 56778 | 885 | lemma poly_pad_rule: "poly p x = 0 \<Longrightarrow> poly (pCons 0 p) x = 0" | 
| 886 | by simp | |
| 26123 | 887 | |
| 55735 
81ba62493610
generalised some results using type classes
 paulson <lp15@cam.ac.uk> parents: 
55734diff
changeset | 888 | lemma poly_cancel_eq_conv: | 
| 56778 | 889 | fixes x :: "'a::field" | 
| 56795 | 890 | shows "x = 0 \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> y = 0 \<longleftrightarrow> a * y - b * x = 0" | 
| 55735 
81ba62493610
generalised some results using type classes
 paulson <lp15@cam.ac.uk> parents: 
55734diff
changeset | 891 | by auto | 
| 26123 | 892 | |
| 30488 | 893 | lemma poly_divides_pad_rule: | 
| 56778 | 894 |   fixes p:: "('a::comm_ring_1) poly"
 | 
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 895 | assumes pq: "p dvd q" | 
| 56778 | 896 | shows "p dvd (pCons 0 q)" | 
| 77282 | 897 | by (metis add_0 dvd_def mult_pCons_right pq smult_0_left) | 
| 26123 | 898 | |
| 30488 | 899 | lemma poly_divides_conv0: | 
| 56778 | 900 | fixes p:: "'a::field poly" | 
| 77282 | 901 | assumes lgpq: "degree q < degree p" and lq: "p \<noteq> 0" | 
| 902 | shows "p dvd q \<longleftrightarrow> q = 0" | |
| 903 | using lgpq mod_poly_less by fastforce | |
| 26123 | 904 | |
| 30488 | 905 | lemma poly_divides_conv1: | 
| 56778 | 906 | fixes p :: "'a::field poly" | 
| 56776 | 907 | assumes a0: "a \<noteq> 0" | 
| 908 | and pp': "p dvd p'" | |
| 909 | and qrp': "smult a q - p' = r" | |
| 77282 | 910 | shows "p dvd q \<longleftrightarrow> p dvd r" | 
| 911 | by (metis a0 diff_add_cancel dvd_add_left_iff dvd_smult_iff pp' qrp') | |
| 26123 | 912 | |
| 913 | lemma basic_cqe_conv1: | |
| 55358 
85d81bc281d0
Simplified some proofs, deleting a lot of strange unused material at the end of the theory.
 paulson <lp15@cam.ac.uk> parents: 
54489diff
changeset | 914 | "(\<exists>x. poly p x = 0 \<and> poly 0 x \<noteq> 0) \<longleftrightarrow> False" | 
| 
85d81bc281d0
Simplified some proofs, deleting a lot of strange unused material at the end of the theory.
 paulson <lp15@cam.ac.uk> parents: 
54489diff
changeset | 915 | "(\<exists>x. poly 0 x \<noteq> 0) \<longleftrightarrow> False" | 
| 56776 | 916 | "(\<exists>x. poly [:c:] x \<noteq> 0) \<longleftrightarrow> c \<noteq> 0" | 
| 55358 
85d81bc281d0
Simplified some proofs, deleting a lot of strange unused material at the end of the theory.
 paulson <lp15@cam.ac.uk> parents: 
54489diff
changeset | 917 | "(\<exists>x. poly 0 x = 0) \<longleftrightarrow> True" | 
| 56776 | 918 | "(\<exists>x. poly [:c:] x = 0) \<longleftrightarrow> c = 0" | 
| 919 | by simp_all | |
| 26123 | 920 | |
| 30488 | 921 | lemma basic_cqe_conv2: | 
| 56795 | 922 | assumes l: "p \<noteq> 0" | 
| 923 | shows "\<exists>x. poly (pCons a (pCons b p)) x = (0::complex)" | |
| 77282 | 924 | by (meson fundamental_theorem_of_algebra_alt l pCons_eq_0_iff pCons_eq_iff) | 
| 26123 | 925 | |
| 56776 | 926 | lemma basic_cqe_conv_2b: "(\<exists>x. poly p x \<noteq> (0::complex)) \<longleftrightarrow> p \<noteq> 0" | 
| 927 | by (metis poly_all_0_iff_0) | |
| 26123 | 928 | |
| 929 | lemma basic_cqe_conv3: | |
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 930 | fixes p q :: "complex poly" | 
| 30488 | 931 | assumes l: "p \<noteq> 0" | 
| 56795 | 932 | shows "(\<exists>x. poly (pCons a p) x = 0 \<and> poly q x \<noteq> 0) \<longleftrightarrow> \<not> (pCons a p) dvd (q ^ psize p)" | 
| 77282 | 933 | by (metis degree_pCons_eq_if l nullstellensatz_univariate pCons_eq_0_iff psize_def) | 
| 26123 | 934 | |
| 935 | lemma basic_cqe_conv4: | |
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 936 | fixes p q :: "complex poly" | 
| 55358 
85d81bc281d0
Simplified some proofs, deleting a lot of strange unused material at the end of the theory.
 paulson <lp15@cam.ac.uk> parents: 
54489diff
changeset | 937 | assumes h: "\<And>x. poly (q ^ n) x = poly r x" | 
| 
85d81bc281d0
Simplified some proofs, deleting a lot of strange unused material at the end of the theory.
 paulson <lp15@cam.ac.uk> parents: 
54489diff
changeset | 938 | shows "p dvd (q ^ n) \<longleftrightarrow> p dvd r" | 
| 77341 | 939 | by (metis (no_types) basic_cqe_conv_2b h poly_diff right_minus_eq) | 
| 26123 | 940 | |
| 55735 
81ba62493610
generalised some results using type classes
 paulson <lp15@cam.ac.uk> parents: 
55734diff
changeset | 941 | lemma poly_const_conv: | 
| 56778 | 942 | fixes x :: "'a::comm_ring_1" | 
| 56776 | 943 | shows "poly [:c:] x = y \<longleftrightarrow> c = y" | 
| 944 | by simp | |
| 26123 | 945 | |
| 29464 
c0d225a7f6ff
convert Fundamental_Theorem_Algebra.thy to use new Polynomial library
 huffman parents: 
29292diff
changeset | 946 | end |