| 26123 |      1 | (*  Title:       Fundamental_Theorem_Algebra.thy
 | 
|  |      2 |     ID:          $Id$
 | 
|  |      3 |     Author:      Amine Chaieb
 | 
|  |      4 | *)
 | 
|  |      5 | 
 | 
|  |      6 | header{*Fundamental Theorem of Algebra*}
 | 
|  |      7 | 
 | 
|  |      8 | theory Fundamental_Theorem_Algebra
 | 
| 26162 |      9 |   imports  Univ_Poly Dense_Linear_Order Complex
 | 
| 26123 |     10 | begin
 | 
|  |     11 | 
 | 
|  |     12 | section {* Square root of complex numbers *}
 | 
|  |     13 | definition csqrt :: "complex \<Rightarrow> complex" where
 | 
|  |     14 | "csqrt z = (if Im z = 0 then
 | 
|  |     15 |             if 0 \<le> Re z then Complex (sqrt(Re z)) 0
 | 
|  |     16 |             else Complex 0 (sqrt(- Re z))
 | 
|  |     17 |            else Complex (sqrt((cmod z + Re z) /2))
 | 
|  |     18 |                         ((Im z / abs(Im z)) * sqrt((cmod z - Re z) /2)))"
 | 
|  |     19 | 
 | 
|  |     20 | lemma csqrt: "csqrt z ^ 2 = z"
 | 
|  |     21 | proof-
 | 
|  |     22 |   obtain x y where xy: "z = Complex x y" by (cases z, simp_all)
 | 
|  |     23 |   {assume y0: "y = 0"
 | 
|  |     24 |     {assume x0: "x \<ge> 0" 
 | 
|  |     25 |       then have ?thesis using y0 xy real_sqrt_pow2[OF x0]
 | 
|  |     26 | 	by (simp add: csqrt_def power2_eq_square)}
 | 
|  |     27 |     moreover
 | 
|  |     28 |     {assume "\<not> x \<ge> 0" hence x0: "- x \<ge> 0" by arith
 | 
|  |     29 |       then have ?thesis using y0 xy real_sqrt_pow2[OF x0] 
 | 
|  |     30 | 	by (simp add: csqrt_def power2_eq_square) }
 | 
|  |     31 |     ultimately have ?thesis by blast}
 | 
|  |     32 |   moreover
 | 
|  |     33 |   {assume y0: "y\<noteq>0"
 | 
|  |     34 |     {fix x y
 | 
|  |     35 |       let ?z = "Complex x y"
 | 
|  |     36 |       from abs_Re_le_cmod[of ?z] have tha: "abs x \<le> cmod ?z" by auto
 | 
|  |     37 |       hence "cmod ?z - x \<ge> 0" "cmod ?z + x \<ge> 0" by (cases "x \<ge> 0", arith+)
 | 
|  |     38 |       hence "(sqrt (x * x + y * y) + x) / 2 \<ge> 0" "(sqrt (x * x + y * y) - x) / 2 \<ge> 0" by (simp_all add: power2_eq_square) }
 | 
|  |     39 |     note th = this
 | 
|  |     40 |     have sq4: "\<And>x::real. x^2 / 4 = (x / 2) ^ 2"
 | 
|  |     41 |       by (simp add: power2_eq_square) 
 | 
|  |     42 |     from th[of x y]
 | 
|  |     43 |     have sq4': "sqrt (((sqrt (x * x + y * y) + x)^2 / 4)) = (sqrt (x * x + y * y) + x) / 2" "sqrt (((sqrt (x * x + y * y) - x)^2 / 4)) = (sqrt (x * x + y * y) - x) / 2" unfolding sq4 by simp_all
 | 
|  |     44 |     then have th1: "sqrt ((sqrt (x * x + y * y) + x) * (sqrt (x * x + y * y) + x) / 4) - sqrt ((sqrt (x * x + y * y) - x) * (sqrt (x * x + y * y) - x) / 4) = x"
 | 
|  |     45 |       unfolding power2_eq_square by simp 
 | 
|  |     46 |     have "sqrt 4 = sqrt (2^2)" by simp 
 | 
|  |     47 |     hence sqrt4: "sqrt 4 = 2" by (simp only: real_sqrt_abs)
 | 
|  |     48 |     have th2: "2 *(y * sqrt ((sqrt (x * x + y * y) - x) * (sqrt (x * x + y * y) + x) / 4)) / \<bar>y\<bar> = y"
 | 
|  |     49 |       using iffD2[OF real_sqrt_pow2_iff sum_power2_ge_zero[of x y]] y0
 | 
|  |     50 |       unfolding power2_eq_square 
 | 
|  |     51 |       by (simp add: ring_simps real_sqrt_divide sqrt4)
 | 
|  |     52 |      from y0 xy have ?thesis  apply (simp add: csqrt_def power2_eq_square)
 | 
|  |     53 |        apply (simp add: real_sqrt_sum_squares_mult_ge_zero[of x y] real_sqrt_pow2[OF th(1)[of x y], unfolded power2_eq_square] real_sqrt_pow2[OF th(2)[of x y], unfolded power2_eq_square] real_sqrt_mult[symmetric])
 | 
|  |     54 |       using th1 th2  ..}
 | 
|  |     55 |   ultimately show ?thesis by blast
 | 
|  |     56 | qed
 | 
|  |     57 | 
 | 
|  |     58 | 
 | 
|  |     59 | section{* More lemmas about module of complex numbers *}
 | 
|  |     60 | 
 | 
|  |     61 | lemma complex_of_real_power: "complex_of_real x ^ n = complex_of_real (x^n)"
 | 
|  |     62 |   by (induct n, auto)
 | 
|  |     63 | 
 | 
|  |     64 | lemma cmod_pos: "cmod z \<ge> 0" by simp
 | 
|  |     65 | lemma complex_mod_triangle_ineq: "cmod (z + w) \<le> cmod z + cmod w"
 | 
|  |     66 |   using complex_mod_triangle_ineq2[of z w] by (simp add: ring_simps)
 | 
|  |     67 | 
 | 
|  |     68 | lemma cmod_mult: "cmod (z*w) = cmod z * cmod w"
 | 
|  |     69 | proof-
 | 
|  |     70 |   from rcis_Ex[of z] rcis_Ex[of w]
 | 
|  |     71 |   obtain rz az rw aw where z: "z = rcis rz az" and w: "w = rcis rw aw"  by blast
 | 
|  |     72 |   thus ?thesis by (simp add: rcis_mult abs_mult)
 | 
|  |     73 | qed
 | 
|  |     74 | 
 | 
|  |     75 | lemma cmod_divide: "cmod (z/w) = cmod z / cmod w"
 | 
|  |     76 | proof-
 | 
|  |     77 |   from rcis_Ex[of z] rcis_Ex[of w]
 | 
|  |     78 |   obtain rz az rw aw where z: "z = rcis rz az" and w: "w = rcis rw aw"  by blast
 | 
|  |     79 |   thus ?thesis by (simp add: rcis_divide)
 | 
|  |     80 | qed
 | 
|  |     81 | 
 | 
|  |     82 | lemma cmod_inverse: "cmod (inverse z) = inverse (cmod z)"
 | 
|  |     83 |   using cmod_divide[of 1 z] by (simp add: inverse_eq_divide)
 | 
|  |     84 | 
 | 
|  |     85 | lemma cmod_uminus: "cmod (- z) = cmod z"
 | 
|  |     86 |   unfolding cmod_def by simp
 | 
|  |     87 | lemma cmod_abs_norm: "\<bar>cmod w - cmod z\<bar> \<le> cmod (w - z)"
 | 
|  |     88 | proof-
 | 
|  |     89 |   have ath: "\<And>(a::real) b x. a - b <= x \<Longrightarrow> b - a <= x ==> abs(a - b) <= x"
 | 
|  |     90 |     by arith
 | 
|  |     91 |   from complex_mod_triangle_ineq2[of "w - z" z]
 | 
|  |     92 |   have th1: "cmod w - cmod z \<le> cmod (w - z)" by simp
 | 
|  |     93 |   from complex_mod_triangle_ineq2[of "- (w - z)" "w"] 
 | 
|  |     94 |   have th2: "cmod z - cmod w \<le> cmod (w - z)" using cmod_uminus [of "w - z"]
 | 
|  |     95 |     by simp
 | 
|  |     96 |   from ath[OF th1 th2] show ?thesis .
 | 
|  |     97 | qed
 | 
|  |     98 | 
 | 
|  |     99 | lemma cmod_power: "cmod (z ^n) = cmod z ^ n" by (induct n, auto simp add: cmod_mult)
 | 
|  |    100 | lemma real_down2: "(0::real) < d1 \<Longrightarrow> 0 < d2 ==> EX e. 0 < e & e < d1 & e < d2"
 | 
|  |    101 |   apply ferrack apply arith done
 | 
|  |    102 | 
 | 
|  |    103 | lemma cmod_complex_of_real: "cmod (complex_of_real x) = \<bar>x\<bar>"
 | 
|  |    104 |   unfolding cmod_def by auto
 | 
|  |    105 | 
 | 
|  |    106 | 
 | 
|  |    107 | text{* The triangle inequality for cmod *}
 | 
|  |    108 | lemma complex_mod_triangle_sub: "cmod w \<le> cmod (w + z) + norm z"
 | 
|  |    109 |   using complex_mod_triangle_ineq2[of "w + z" "-z"] by auto
 | 
|  |    110 | 
 | 
|  |    111 | section{* Basic lemmas about complex polynomials *}
 | 
|  |    112 | 
 | 
|  |    113 | lemma poly_bound_exists:
 | 
|  |    114 |   shows "\<exists>m. m > 0 \<and> (\<forall>z. cmod z <= r \<longrightarrow> cmod (poly p z) \<le> m)"
 | 
|  |    115 | proof(induct p)
 | 
|  |    116 |   case Nil thus ?case by (rule exI[where x=1], simp) 
 | 
|  |    117 | next
 | 
|  |    118 |   case (Cons c cs)
 | 
|  |    119 |   from Cons.hyps obtain m where m: "\<forall>z. cmod z \<le> r \<longrightarrow> cmod (poly cs z) \<le> m"
 | 
|  |    120 |     by blast
 | 
|  |    121 |   let ?k = " 1 + cmod c + \<bar>r * m\<bar>"
 | 
|  |    122 |   have kp: "?k > 0" using abs_ge_zero[of "r*m"] cmod_pos[of c] by arith
 | 
|  |    123 |   {fix z
 | 
|  |    124 |     assume H: "cmod z \<le> r"
 | 
|  |    125 |     from m H have th: "cmod (poly cs z) \<le> m" by blast
 | 
|  |    126 |     from H have rp: "r \<ge> 0" using cmod_pos[of z] by arith
 | 
|  |    127 |     have "cmod (poly (c # cs) z) \<le> cmod c + cmod (z* poly cs z)"
 | 
|  |    128 |       using complex_mod_triangle_ineq[of c "z* poly cs z"] by simp
 | 
|  |    129 |     also have "\<dots> \<le> cmod c + r*m" using mult_mono[OF H th rp cmod_pos[of "poly cs z"]] by (simp add: cmod_mult)
 | 
|  |    130 |     also have "\<dots> \<le> ?k" by simp
 | 
|  |    131 |     finally have "cmod (poly (c # cs) z) \<le> ?k" .}
 | 
|  |    132 |   with kp show ?case by blast
 | 
|  |    133 | qed
 | 
|  |    134 | 
 | 
|  |    135 | 
 | 
|  |    136 | text{* Offsetting the variable in a polynomial gives another of same degree *}
 | 
| 26135 |    137 |   (* FIXME : Lemma holds also in locale --- fix it later *)
 | 
| 26123 |    138 | lemma  poly_offset_lemma:
 | 
|  |    139 |   shows "\<exists>b q. (length q = length p) \<and> (\<forall>x. poly (b#q) (x::complex) = (a + x) * poly p x)"
 | 
|  |    140 | proof(induct p)
 | 
|  |    141 |   case Nil thus ?case by simp
 | 
|  |    142 | next
 | 
|  |    143 |   case (Cons c cs)
 | 
|  |    144 |   from Cons.hyps obtain b q where 
 | 
|  |    145 |     bq: "length q = length cs" "\<forall>x. poly (b # q) x = (a + x) * poly cs x"
 | 
|  |    146 |     by blast
 | 
|  |    147 |   let ?b = "a*c"
 | 
|  |    148 |   let ?q = "(b+c)#q"
 | 
|  |    149 |   have lg: "length ?q = length (c#cs)" using bq(1) by simp
 | 
|  |    150 |   {fix x
 | 
|  |    151 |     from bq(2)[rule_format, of x]
 | 
|  |    152 |     have "x*poly (b # q) x = x*((a + x) * poly cs x)" by simp
 | 
|  |    153 |     hence "poly (?b# ?q) x = (a + x) * poly (c # cs) x"
 | 
|  |    154 |       by (simp add: ring_simps)}
 | 
|  |    155 |   with lg  show ?case by blast 
 | 
|  |    156 | qed
 | 
|  |    157 | 
 | 
|  |    158 |     (* FIXME : This one too*)
 | 
|  |    159 | lemma poly_offset: "\<exists> q. length q = length p \<and> (\<forall>x. poly q (x::complex) = poly p (a + x))"
 | 
|  |    160 | proof (induct p)
 | 
|  |    161 |   case Nil thus ?case by simp
 | 
|  |    162 | next
 | 
|  |    163 |   case (Cons c cs)
 | 
|  |    164 |   from Cons.hyps obtain q where q: "length q = length cs" "\<forall>x. poly q x = poly cs (a + x)" by blast
 | 
|  |    165 |   from poly_offset_lemma[of q a] obtain b p where 
 | 
|  |    166 |     bp: "length p = length q" "\<forall>x. poly (b # p) x = (a + x) * poly q x"
 | 
|  |    167 |     by blast
 | 
|  |    168 |   thus ?case using q bp by - (rule exI[where x="(c + b)#p"], simp)
 | 
|  |    169 | qed
 | 
|  |    170 | 
 | 
|  |    171 | text{* An alternative useful formulation of completeness of the reals *}
 | 
|  |    172 | lemma real_sup_exists: assumes ex: "\<exists>x. P x" and bz: "\<exists>z. \<forall>x. P x \<longrightarrow> x < z"
 | 
|  |    173 |   shows "\<exists>(s::real). \<forall>y. (\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < s"
 | 
|  |    174 | proof-
 | 
|  |    175 |   from ex bz obtain x Y where x: "P x" and Y: "\<And>x. P x \<Longrightarrow> x < Y"  by blast
 | 
|  |    176 |   from ex have thx:"\<exists>x. x \<in> Collect P" by blast
 | 
|  |    177 |   from bz have thY: "\<exists>Y. isUb UNIV (Collect P) Y" 
 | 
|  |    178 |     by(auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def order_le_less)
 | 
|  |    179 |   from reals_complete[OF thx thY] obtain L where L: "isLub UNIV (Collect P) L"
 | 
|  |    180 |     by blast
 | 
|  |    181 |   from Y[OF x] have xY: "x < Y" .
 | 
|  |    182 |   from L have L': "\<forall>x. P x \<longrightarrow> x \<le> L" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def)  
 | 
|  |    183 |   from Y have Y': "\<forall>x. P x \<longrightarrow> x \<le> Y" 
 | 
|  |    184 |     apply (clarsimp, atomize (full)) by auto 
 | 
|  |    185 |   from L Y' have "L \<le> Y" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def)
 | 
|  |    186 |   {fix y
 | 
|  |    187 |     {fix z assume z: "P z" "y < z"
 | 
|  |    188 |       from L' z have "y < L" by auto }
 | 
|  |    189 |     moreover
 | 
|  |    190 |     {assume yL: "y < L" "\<forall>z. P z \<longrightarrow> \<not> y < z"
 | 
|  |    191 |       hence nox: "\<forall>z. P z \<longrightarrow> y \<ge> z" by auto
 | 
|  |    192 |       from nox L have "y \<ge> L" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def) 
 | 
|  |    193 |       with yL(1) have False  by arith}
 | 
|  |    194 |     ultimately have "(\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < L" by blast}
 | 
|  |    195 |   thus ?thesis by blast
 | 
|  |    196 | qed
 | 
|  |    197 | 
 | 
|  |    198 | 
 | 
|  |    199 | section{* Some theorems about Sequences*}
 | 
|  |    200 | text{* Given a binary function @{text "f:: nat \<Rightarrow> 'a \<Rightarrow> 'a"}, its values are uniquely determined by a function g *}
 | 
|  |    201 | 
 | 
|  |    202 | lemma num_Axiom: "EX! g. g 0 = e \<and> (\<forall>n. g (Suc n) = f n (g n))"
 | 
|  |    203 |   unfolding Ex1_def
 | 
|  |    204 |   apply (rule_tac x="nat_rec e f" in exI)
 | 
|  |    205 |   apply (rule conjI)+
 | 
|  |    206 | apply (rule def_nat_rec_0, simp)
 | 
|  |    207 | apply (rule allI, rule def_nat_rec_Suc, simp)
 | 
|  |    208 | apply (rule allI, rule impI, rule ext)
 | 
|  |    209 | apply (erule conjE)
 | 
|  |    210 | apply (induct_tac x)
 | 
|  |    211 | apply (simp add: nat_rec_0)
 | 
|  |    212 | apply (erule_tac x="n" in allE)
 | 
|  |    213 | apply (simp)
 | 
|  |    214 | done
 | 
|  |    215 | 
 | 
|  |    216 |  text{* An equivalent formulation of monotony -- Not used here, but might be useful *}
 | 
|  |    217 | lemma mono_Suc: "mono f = (\<forall>n. (f n :: 'a :: order) \<le> f (Suc n))"
 | 
|  |    218 | unfolding mono_def
 | 
|  |    219 | proof auto
 | 
|  |    220 |   fix A B :: nat
 | 
|  |    221 |   assume H: "\<forall>n. f n \<le> f (Suc n)" "A \<le> B"
 | 
|  |    222 |   hence "\<exists>k. B = A + k" apply -  apply (thin_tac "\<forall>n. f n \<le> f (Suc n)") 
 | 
|  |    223 |     by presburger
 | 
|  |    224 |   then obtain k where k: "B = A + k" by blast
 | 
|  |    225 |   {fix a k
 | 
|  |    226 |     have "f a \<le> f (a + k)"
 | 
|  |    227 |     proof (induct k)
 | 
|  |    228 |       case 0 thus ?case by simp
 | 
|  |    229 |     next
 | 
|  |    230 |       case (Suc k)
 | 
|  |    231 |       from Suc.hyps H(1)[rule_format, of "a + k"] show ?case by simp
 | 
|  |    232 |     qed}
 | 
|  |    233 |   with k show "f A \<le> f B" by blast
 | 
|  |    234 | qed
 | 
|  |    235 | 
 | 
|  |    236 | text{* for any sequence, there is a mootonic subsequence *}
 | 
|  |    237 | lemma seq_monosub: "\<exists>f. subseq f \<and> monoseq (\<lambda> n. (s (f n)))"
 | 
|  |    238 | proof-
 | 
|  |    239 |   {assume H: "\<forall>n. \<exists>p >n. \<forall> m\<ge>p. s m \<le> s p"
 | 
|  |    240 |     let ?P = "\<lambda> p n. p > n \<and> (\<forall>m \<ge> p. s m \<le> s p)"
 | 
|  |    241 |     from num_Axiom[of "SOME p. ?P p 0" "\<lambda>p n. SOME p. ?P p n"]
 | 
|  |    242 |     obtain f where f: "f 0 = (SOME p. ?P p 0)" "\<forall>n. f (Suc n) = (SOME p. ?P p (f n))" by blast
 | 
|  |    243 |     have "?P (f 0) 0"  unfolding f(1) some_eq_ex[of "\<lambda>p. ?P p 0"]
 | 
|  |    244 |       using H apply - 
 | 
|  |    245 |       apply (erule allE[where x=0], erule exE, rule_tac x="p" in exI) 
 | 
|  |    246 |       unfolding order_le_less by blast 
 | 
|  |    247 |     hence f0: "f 0 > 0" "\<forall>m \<ge> f 0. s m \<le> s (f 0)" by blast+
 | 
|  |    248 |     {fix n
 | 
|  |    249 |       have "?P (f (Suc n)) (f n)" 
 | 
|  |    250 | 	unfolding f(2)[rule_format, of n] some_eq_ex[of "\<lambda>p. ?P p (f n)"]
 | 
|  |    251 | 	using H apply - 
 | 
|  |    252 |       apply (erule allE[where x="f n"], erule exE, rule_tac x="p" in exI) 
 | 
|  |    253 |       unfolding order_le_less by blast 
 | 
|  |    254 |     hence "f (Suc n) > f n" "\<forall>m \<ge> f (Suc n). s m \<le> s (f (Suc n))" by blast+}
 | 
|  |    255 |   note fSuc = this
 | 
|  |    256 |     {fix p q assume pq: "p \<ge> f q"
 | 
|  |    257 |       have "s p \<le> s(f(q))"  using f0(2)[rule_format, of p] pq fSuc
 | 
|  |    258 | 	by (cases q, simp_all) }
 | 
|  |    259 |     note pqth = this
 | 
|  |    260 |     {fix q
 | 
|  |    261 |       have "f (Suc q) > f q" apply (induct q) 
 | 
|  |    262 | 	using f0(1) fSuc(1)[of 0] apply simp by (rule fSuc(1))}
 | 
|  |    263 |     note fss = this
 | 
|  |    264 |     from fss have th1: "subseq f" unfolding subseq_Suc_iff ..
 | 
|  |    265 |     {fix a b 
 | 
|  |    266 |       have "f a \<le> f (a + b)"
 | 
|  |    267 |       proof(induct b)
 | 
|  |    268 | 	case 0 thus ?case by simp
 | 
|  |    269 |       next
 | 
|  |    270 | 	case (Suc b)
 | 
|  |    271 | 	from fSuc(1)[of "a + b"] Suc.hyps show ?case by simp
 | 
|  |    272 |       qed}
 | 
|  |    273 |     note fmon0 = this
 | 
|  |    274 |     have "monoseq (\<lambda>n. s (f n))" 
 | 
|  |    275 |     proof-
 | 
|  |    276 |       {fix n
 | 
|  |    277 | 	have "s (f n) \<ge> s (f (Suc n))" 
 | 
|  |    278 | 	proof(cases n)
 | 
|  |    279 | 	  case 0
 | 
|  |    280 | 	  assume n0: "n = 0"
 | 
|  |    281 | 	  from fSuc(1)[of 0] have th0: "f 0 \<le> f (Suc 0)" by simp
 | 
|  |    282 | 	  from f0(2)[rule_format, OF th0] show ?thesis  using n0 by simp
 | 
|  |    283 | 	next
 | 
|  |    284 | 	  case (Suc m)
 | 
|  |    285 | 	  assume m: "n = Suc m"
 | 
|  |    286 | 	  from fSuc(1)[of n] m have th0: "f (Suc m) \<le> f (Suc (Suc m))" by simp
 | 
|  |    287 | 	  from m fSuc(2)[rule_format, OF th0] show ?thesis by simp 
 | 
|  |    288 | 	qed}
 | 
|  |    289 |       thus "monoseq (\<lambda>n. s (f n))" unfolding monoseq_Suc by blast 
 | 
|  |    290 |     qed
 | 
|  |    291 |     with th1 have ?thesis by blast}
 | 
|  |    292 |   moreover
 | 
|  |    293 |   {fix N assume N: "\<forall>p >N. \<exists> m\<ge>p. s m > s p"
 | 
|  |    294 |     {fix p assume p: "p \<ge> Suc N" 
 | 
|  |    295 |       hence pN: "p > N" by arith with N obtain m where m: "m \<ge> p" "s m > s p" by blast
 | 
|  |    296 |       have "m \<noteq> p" using m(2) by auto 
 | 
|  |    297 |       with m have "\<exists>m>p. s p < s m" by - (rule exI[where x=m], auto)}
 | 
|  |    298 |     note th0 = this
 | 
|  |    299 |     let ?P = "\<lambda>m x. m > x \<and> s x < s m"
 | 
|  |    300 |     from num_Axiom[of "SOME x. ?P x (Suc N)" "\<lambda>m x. SOME y. ?P y x"]
 | 
|  |    301 |     obtain f where f: "f 0 = (SOME x. ?P x (Suc N))" 
 | 
|  |    302 |       "\<forall>n. f (Suc n) = (SOME m. ?P m (f n))" by blast
 | 
|  |    303 |     have "?P (f 0) (Suc N)"  unfolding f(1) some_eq_ex[of "\<lambda>p. ?P p (Suc N)"]
 | 
|  |    304 |       using N apply - 
 | 
|  |    305 |       apply (erule allE[where x="Suc N"], clarsimp)
 | 
|  |    306 |       apply (rule_tac x="m" in exI)
 | 
|  |    307 |       apply auto
 | 
|  |    308 |       apply (subgoal_tac "Suc N \<noteq> m")
 | 
|  |    309 |       apply simp
 | 
|  |    310 |       apply (rule ccontr, simp)
 | 
|  |    311 |       done
 | 
|  |    312 |     hence f0: "f 0 > Suc N" "s (Suc N) < s (f 0)" by blast+
 | 
|  |    313 |     {fix n
 | 
|  |    314 |       have "f n > N \<and> ?P (f (Suc n)) (f n)"
 | 
|  |    315 | 	unfolding f(2)[rule_format, of n] some_eq_ex[of "\<lambda>p. ?P p (f n)"]
 | 
|  |    316 |       proof (induct n)
 | 
|  |    317 | 	case 0 thus ?case
 | 
|  |    318 | 	  using f0 N apply auto 
 | 
|  |    319 | 	  apply (erule allE[where x="f 0"], clarsimp) 
 | 
|  |    320 | 	  apply (rule_tac x="m" in exI, simp)
 | 
|  |    321 | 	  by (subgoal_tac "f 0 \<noteq> m", auto)
 | 
|  |    322 |       next
 | 
|  |    323 | 	case (Suc n)
 | 
|  |    324 | 	from Suc.hyps have Nfn: "N < f n" by blast
 | 
|  |    325 | 	from Suc.hyps obtain m where m: "m > f n" "s (f n) < s m" by blast
 | 
|  |    326 | 	with Nfn have mN: "m > N" by arith
 | 
|  |    327 | 	note key = Suc.hyps[unfolded some_eq_ex[of "\<lambda>p. ?P p (f n)", symmetric] f(2)[rule_format, of n, symmetric]]
 | 
|  |    328 | 	
 | 
|  |    329 | 	from key have th0: "f (Suc n) > N" by simp
 | 
|  |    330 | 	from N[rule_format, OF th0]
 | 
|  |    331 | 	obtain m' where m': "m' \<ge> f (Suc n)" "s (f (Suc n)) < s m'" by blast
 | 
|  |    332 | 	have "m' \<noteq> f (Suc (n))" apply (rule ccontr) using m'(2) by auto
 | 
|  |    333 | 	hence "m' > f (Suc n)" using m'(1) by simp
 | 
|  |    334 | 	with key m'(2) show ?case by auto
 | 
|  |    335 |       qed}
 | 
|  |    336 |     note fSuc = this
 | 
|  |    337 |     {fix n
 | 
|  |    338 |       have "f n \<ge> Suc N \<and> f(Suc n) > f n \<and> s(f n) < s(f(Suc n))" using fSuc[of n] by auto 
 | 
|  |    339 |       hence "f n \<ge> Suc N" "f(Suc n) > f n" "s(f n) < s(f(Suc n))" by blast+}
 | 
|  |    340 |     note thf = this
 | 
|  |    341 |     have sqf: "subseq f" unfolding subseq_Suc_iff using thf by simp
 | 
|  |    342 |     have "monoseq (\<lambda>n. s (f n))"  unfolding monoseq_Suc using thf
 | 
|  |    343 |       apply -
 | 
|  |    344 |       apply (rule disjI1)
 | 
|  |    345 |       apply auto
 | 
|  |    346 |       apply (rule order_less_imp_le)
 | 
|  |    347 |       apply blast
 | 
|  |    348 |       done
 | 
|  |    349 |     then have ?thesis  using sqf by blast}
 | 
|  |    350 |   ultimately show ?thesis unfolding linorder_not_less[symmetric] by blast
 | 
|  |    351 | qed
 | 
|  |    352 | 
 | 
|  |    353 | lemma seq_suble: assumes sf: "subseq f" shows "n \<le> f n"
 | 
|  |    354 | proof(induct n)
 | 
|  |    355 |   case 0 thus ?case by simp
 | 
|  |    356 | next
 | 
|  |    357 |   case (Suc n)
 | 
|  |    358 |   from sf[unfolded subseq_Suc_iff, rule_format, of n] Suc.hyps
 | 
|  |    359 |   have "n < f (Suc n)" by arith 
 | 
|  |    360 |   thus ?case by arith
 | 
|  |    361 | qed
 | 
|  |    362 | 
 | 
|  |    363 | section {* Fundamental theorem of algebra *}
 | 
|  |    364 | lemma  unimodular_reduce_norm:
 | 
|  |    365 |   assumes md: "cmod z = 1"
 | 
|  |    366 |   shows "cmod (z + 1) < 1 \<or> cmod (z - 1) < 1 \<or> cmod (z + ii) < 1 \<or> cmod (z - ii) < 1"
 | 
|  |    367 | proof-
 | 
|  |    368 |   obtain x y where z: "z = Complex x y " by (cases z, auto)
 | 
|  |    369 |   from md z have xy: "x^2 + y^2 = 1" by (simp add: cmod_def)
 | 
|  |    370 |   {assume C: "cmod (z + 1) \<ge> 1" "cmod (z - 1) \<ge> 1" "cmod (z + ii) \<ge> 1" "cmod (z - ii) \<ge> 1"
 | 
|  |    371 |     from C z xy have "2*x \<le> 1" "2*x \<ge> -1" "2*y \<le> 1" "2*y \<ge> -1"
 | 
|  |    372 |       by (simp_all add: cmod_def power2_eq_square ring_simps)
 | 
|  |    373 |     hence "abs (2*x) \<le> 1" "abs (2*y) \<le> 1" by simp_all
 | 
|  |    374 |     hence "(abs (2 * x))^2 <= 1^2" "(abs (2 * y)) ^2 <= 1^2"
 | 
|  |    375 |       by - (rule power_mono, simp, simp)+
 | 
|  |    376 |     hence th0: "4*x^2 \<le> 1" "4*y^2 \<le> 1" 
 | 
|  |    377 |       by (simp_all  add: power2_abs power_mult_distrib)
 | 
|  |    378 |     from add_mono[OF th0] xy have False by simp }
 | 
|  |    379 |   thus ?thesis unfolding linorder_not_le[symmetric] by blast
 | 
|  |    380 | qed
 | 
|  |    381 | 
 | 
| 26135 |    382 | text{* Hence we can always reduce modulus of @{text "1 + b z^n"} if nonzero *}
 | 
| 26123 |    383 | lemma reduce_poly_simple:
 | 
|  |    384 |  assumes b: "b \<noteq> 0" and n: "n\<noteq>0"
 | 
|  |    385 |   shows "\<exists>z. cmod (1 + b * z^n) < 1"
 | 
|  |    386 | using n
 | 
|  |    387 | proof(induct n rule: nat_less_induct)
 | 
|  |    388 |   fix n
 | 
|  |    389 |   assume IH: "\<forall>m<n. m \<noteq> 0 \<longrightarrow> (\<exists>z. cmod (1 + b * z ^ m) < 1)" and n: "n \<noteq> 0"
 | 
|  |    390 |   let ?P = "\<lambda>z n. cmod (1 + b * z ^ n) < 1"
 | 
|  |    391 |   {assume e: "even n"
 | 
|  |    392 |     hence "\<exists>m. n = 2*m" by presburger
 | 
|  |    393 |     then obtain m where m: "n = 2*m" by blast
 | 
|  |    394 |     from n m have "m\<noteq>0" "m < n" by presburger+
 | 
|  |    395 |     with IH[rule_format, of m] obtain z where z: "?P z m" by blast
 | 
|  |    396 |     from z have "?P (csqrt z) n" by (simp add: m power_mult csqrt)
 | 
|  |    397 |     hence "\<exists>z. ?P z n" ..}
 | 
|  |    398 |   moreover
 | 
|  |    399 |   {assume o: "odd n"
 | 
|  |    400 |     from b have b': "b^2 \<noteq> 0" unfolding power2_eq_square by simp
 | 
|  |    401 |     have "Im (inverse b) * (Im (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) +
 | 
|  |    402 |     Re (inverse b) * (Re (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) = 
 | 
|  |    403 |     ((Re (inverse b))^2 + (Im (inverse b))^2) * \<bar>Im b * Im b + Re b * Re b\<bar>" by algebra
 | 
|  |    404 |     also have "\<dots> = cmod (inverse b) ^2 * cmod b ^ 2" 
 | 
|  |    405 |       apply (simp add: cmod_def) using realpow_two_le_add_order[of "Re b" "Im b"]
 | 
|  |    406 |       by (simp add: power2_eq_square)
 | 
|  |    407 |     finally 
 | 
|  |    408 |     have th0: "Im (inverse b) * (Im (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) +
 | 
|  |    409 |     Re (inverse b) * (Re (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) =
 | 
|  |    410 |     1" 
 | 
|  |    411 |       apply (simp add: power2_eq_square cmod_mult[symmetric] cmod_inverse[symmetric])
 | 
|  |    412 |       using right_inverse[OF b']
 | 
|  |    413 |       by (simp add: power2_eq_square[symmetric] power_inverse[symmetric] ring_simps)
 | 
|  |    414 |     have th0: "cmod (complex_of_real (cmod b) / b) = 1"
 | 
|  |    415 |       apply (simp add: complex_Re_mult cmod_def power2_eq_square Re_complex_of_real Im_complex_of_real divide_inverse ring_simps )
 | 
|  |    416 |       by (simp add: real_sqrt_mult[symmetric] th0)        
 | 
|  |    417 |     from o have "\<exists>m. n = Suc (2*m)" by presburger+
 | 
|  |    418 |     then obtain m where m: "n = Suc (2*m)" by blast
 | 
|  |    419 |     from unimodular_reduce_norm[OF th0] o
 | 
|  |    420 |     have "\<exists>v. cmod (complex_of_real (cmod b) / b + v^n) < 1"
 | 
|  |    421 |       apply (cases "cmod (complex_of_real (cmod b) / b + 1) < 1", rule_tac x="1" in exI, simp)
 | 
|  |    422 |       apply (cases "cmod (complex_of_real (cmod b) / b - 1) < 1", rule_tac x="-1" in exI, simp add: diff_def)
 | 
|  |    423 |       apply (cases "cmod (complex_of_real (cmod b) / b + ii) < 1")
 | 
|  |    424 |       apply (cases "even m", rule_tac x="ii" in exI, simp add: m power_mult)
 | 
|  |    425 |       apply (rule_tac x="- ii" in exI, simp add: m power_mult)
 | 
|  |    426 |       apply (cases "even m", rule_tac x="- ii" in exI, simp add: m power_mult diff_def)
 | 
|  |    427 |       apply (rule_tac x="ii" in exI, simp add: m power_mult diff_def)
 | 
|  |    428 |       done
 | 
|  |    429 |     then obtain v where v: "cmod (complex_of_real (cmod b) / b + v^n) < 1" by blast
 | 
|  |    430 |     let ?w = "v / complex_of_real (root n (cmod b))"
 | 
|  |    431 |     from odd_real_root_pow[OF o, of "cmod b"]
 | 
|  |    432 |     have th1: "?w ^ n = v^n / complex_of_real (cmod b)" 
 | 
|  |    433 |       by (simp add: power_divide complex_of_real_power)
 | 
|  |    434 |     have th2:"cmod (complex_of_real (cmod b) / b) = 1" using b by (simp add: cmod_divide)
 | 
|  |    435 |     hence th3: "cmod (complex_of_real (cmod b) / b) \<ge> 0" by simp
 | 
|  |    436 |     have th4: "cmod (complex_of_real (cmod b) / b) *
 | 
|  |    437 |    cmod (1 + b * (v ^ n / complex_of_real (cmod b)))
 | 
|  |    438 |    < cmod (complex_of_real (cmod b) / b) * 1"
 | 
|  |    439 |       apply (simp only: cmod_mult[symmetric] right_distrib)
 | 
|  |    440 |       using b v by (simp add: th2)
 | 
|  |    441 | 
 | 
|  |    442 |     from mult_less_imp_less_left[OF th4 th3]
 | 
|  |    443 |     have "?P ?w n" unfolding th1 . 
 | 
|  |    444 |     hence "\<exists>z. ?P z n" .. }
 | 
|  |    445 |   ultimately show "\<exists>z. ?P z n" by blast
 | 
|  |    446 | qed
 | 
|  |    447 | 
 | 
|  |    448 | 
 | 
|  |    449 | text{* Bolzano-Weierstrass type property for closed disc in complex plane. *}
 | 
|  |    450 | 
 | 
|  |    451 | lemma metric_bound_lemma: "cmod (x - y) <= \<bar>Re x - Re y\<bar> + \<bar>Im x - Im y\<bar>"
 | 
|  |    452 |   using real_sqrt_sum_squares_triangle_ineq[of "Re x - Re y" 0 0 "Im x - Im y" ]
 | 
|  |    453 |   unfolding cmod_def by simp
 | 
|  |    454 | 
 | 
|  |    455 | lemma bolzano_weierstrass_complex_disc:
 | 
|  |    456 |   assumes r: "\<forall>n. cmod (s n) \<le> r"
 | 
|  |    457 |   shows "\<exists>f z. subseq f \<and> (\<forall>e >0. \<exists>N. \<forall>n \<ge> N. cmod (s (f n) - z) < e)"
 | 
|  |    458 | proof-
 | 
|  |    459 |   from seq_monosub[of "Re o s"] 
 | 
|  |    460 |   obtain f g where f: "subseq f" "monoseq (\<lambda>n. Re (s (f n)))" 
 | 
|  |    461 |     unfolding o_def by blast
 | 
|  |    462 |   from seq_monosub[of "Im o s o f"] 
 | 
|  |    463 |   obtain g where g: "subseq g" "monoseq (\<lambda>n. Im (s(f(g n))))" unfolding o_def by blast  
 | 
|  |    464 |   let ?h = "f o g"
 | 
|  |    465 |   from r[rule_format, of 0] have rp: "r \<ge> 0" using cmod_pos[of "s 0"] by arith 
 | 
|  |    466 |   have th:"\<forall>n. r + 1 \<ge> \<bar> Re (s n)\<bar>" 
 | 
|  |    467 |   proof
 | 
|  |    468 |     fix n
 | 
|  |    469 |     from abs_Re_le_cmod[of "s n"] r[rule_format, of n]  show "\<bar>Re (s n)\<bar> \<le> r + 1" by arith
 | 
|  |    470 |   qed
 | 
|  |    471 |   have conv1: "convergent (\<lambda>n. Re (s ( f n)))"
 | 
|  |    472 |     apply (rule Bseq_monoseq_convergent)
 | 
|  |    473 |     apply (simp add: Bseq_def)
 | 
|  |    474 |     apply (rule exI[where x= "r + 1"])
 | 
|  |    475 |     using th rp apply simp
 | 
|  |    476 |     using f(2) .
 | 
|  |    477 |   have th:"\<forall>n. r + 1 \<ge> \<bar> Im (s n)\<bar>" 
 | 
|  |    478 |   proof
 | 
|  |    479 |     fix n
 | 
|  |    480 |     from abs_Im_le_cmod[of "s n"] r[rule_format, of n]  show "\<bar>Im (s n)\<bar> \<le> r + 1" by arith
 | 
|  |    481 |   qed
 | 
|  |    482 | 
 | 
|  |    483 |   have conv2: "convergent (\<lambda>n. Im (s (f (g n))))"
 | 
|  |    484 |     apply (rule Bseq_monoseq_convergent)
 | 
|  |    485 |     apply (simp add: Bseq_def)
 | 
|  |    486 |     apply (rule exI[where x= "r + 1"])
 | 
|  |    487 |     using th rp apply simp
 | 
|  |    488 |     using g(2) .
 | 
|  |    489 | 
 | 
|  |    490 |   from conv1[unfolded convergent_def] obtain x where "LIMSEQ (\<lambda>n. Re (s (f n))) x" 
 | 
|  |    491 |     by blast 
 | 
|  |    492 |   hence  x: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar> Re (s (f n)) - x \<bar> < r" 
 | 
|  |    493 |     unfolding LIMSEQ_def real_norm_def .
 | 
|  |    494 | 
 | 
|  |    495 |   from conv2[unfolded convergent_def] obtain y where "LIMSEQ (\<lambda>n. Im (s (f (g n)))) y" 
 | 
|  |    496 |     by blast 
 | 
|  |    497 |   hence  y: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar> Im (s (f (g n))) - y \<bar> < r" 
 | 
|  |    498 |     unfolding LIMSEQ_def real_norm_def .
 | 
|  |    499 |   let ?w = "Complex x y"
 | 
|  |    500 |   from f(1) g(1) have hs: "subseq ?h" unfolding subseq_def by auto 
 | 
|  |    501 |   {fix e assume ep: "e > (0::real)"
 | 
|  |    502 |     hence e2: "e/2 > 0" by simp
 | 
|  |    503 |     from x[rule_format, OF e2] y[rule_format, OF e2]
 | 
|  |    504 |     obtain N1 N2 where N1: "\<forall>n\<ge>N1. \<bar>Re (s (f n)) - x\<bar> < e / 2" and N2: "\<forall>n\<ge>N2. \<bar>Im (s (f (g n))) - y\<bar> < e / 2" by blast
 | 
|  |    505 |     {fix n assume nN12: "n \<ge> N1 + N2"
 | 
|  |    506 |       hence nN1: "g n \<ge> N1" and nN2: "n \<ge> N2" using seq_suble[OF g(1), of n] by arith+
 | 
|  |    507 |       from add_strict_mono[OF N1[rule_format, OF nN1] N2[rule_format, OF nN2]]
 | 
|  |    508 |       have "cmod (s (?h n) - ?w) < e" 
 | 
|  |    509 | 	using metric_bound_lemma[of "s (f (g n))" ?w] by simp }
 | 
|  |    510 |     hence "\<exists>N. \<forall>n\<ge>N. cmod (s (?h n) - ?w) < e" by blast }
 | 
|  |    511 |   with hs show ?thesis  by blast  
 | 
|  |    512 | qed
 | 
|  |    513 | 
 | 
|  |    514 | text{* Polynomial is continuous. *}
 | 
|  |    515 | 
 | 
|  |    516 | lemma poly_cont:
 | 
|  |    517 |   assumes ep: "e > 0" 
 | 
|  |    518 |   shows "\<exists>d >0. \<forall>w. 0 < cmod (w - z) \<and> cmod (w - z) < d \<longrightarrow> cmod (poly p w - poly p z) < e"
 | 
|  |    519 | proof-
 | 
|  |    520 |   from poly_offset[of p z] obtain q where q: "length q = length p" "\<And>x. poly q x = poly p (z + x)" by blast
 | 
|  |    521 |   {fix w
 | 
|  |    522 |     note q(2)[of "w - z", simplified]}
 | 
|  |    523 |   note th = this
 | 
|  |    524 |   show ?thesis unfolding th[symmetric]
 | 
|  |    525 |   proof(induct q)
 | 
|  |    526 |     case Nil thus ?case  using ep by auto
 | 
|  |    527 |   next
 | 
|  |    528 |     case (Cons c cs)
 | 
|  |    529 |     from poly_bound_exists[of 1 "cs"] 
 | 
|  |    530 |     obtain m where m: "m > 0" "\<And>z. cmod z \<le> 1 \<Longrightarrow> cmod (poly cs z) \<le> m" by blast
 | 
|  |    531 |     from ep m(1) have em0: "e/m > 0" by (simp add: field_simps)
 | 
|  |    532 |     have one0: "1 > (0::real)"  by arith
 | 
|  |    533 |     from real_lbound_gt_zero[OF one0 em0] 
 | 
|  |    534 |     obtain d where d: "d >0" "d < 1" "d < e / m" by blast
 | 
|  |    535 |     from d(1,3) m(1) have dm: "d*m > 0" "d*m < e" 
 | 
|  |    536 |       by (simp_all add: field_simps real_mult_order)
 | 
|  |    537 |     show ?case 
 | 
|  |    538 |       proof(rule ex_forward[OF real_lbound_gt_zero[OF one0 em0]], clarsimp simp add: cmod_mult)
 | 
|  |    539 | 	fix d w
 | 
|  |    540 | 	assume H: "d > 0" "d < 1" "d < e/m" "w\<noteq>z" "cmod (w-z) < d"
 | 
|  |    541 | 	hence d1: "cmod (w-z) \<le> 1" "d \<ge> 0" by simp_all
 | 
|  |    542 | 	from H(3) m(1) have dme: "d*m < e" by (simp add: field_simps)
 | 
|  |    543 | 	from H have th: "cmod (w-z) \<le> d" by simp 
 | 
|  |    544 | 	from mult_mono[OF th m(2)[OF d1(1)] d1(2) cmod_pos] dme
 | 
|  |    545 | 	show "cmod (w - z) * cmod (poly cs (w - z)) < e" by simp
 | 
|  |    546 |       qed  
 | 
|  |    547 |     qed
 | 
|  |    548 | qed
 | 
|  |    549 | 
 | 
|  |    550 | text{* Hence a polynomial attains minimum on a closed disc 
 | 
|  |    551 |   in the complex plane. *}
 | 
|  |    552 | lemma  poly_minimum_modulus_disc:
 | 
|  |    553 |   "\<exists>z. \<forall>w. cmod w \<le> r \<longrightarrow> cmod (poly p z) \<le> cmod (poly p w)"
 | 
|  |    554 | proof-
 | 
|  |    555 |   {assume "\<not> r \<ge> 0" hence ?thesis unfolding linorder_not_le
 | 
|  |    556 |       apply -
 | 
|  |    557 |       apply (rule exI[where x=0]) 
 | 
|  |    558 |       apply auto
 | 
|  |    559 |       apply (subgoal_tac "cmod w < 0")
 | 
|  |    560 |       apply simp
 | 
|  |    561 |       apply arith
 | 
|  |    562 |       done }
 | 
|  |    563 |   moreover
 | 
|  |    564 |   {assume rp: "r \<ge> 0"
 | 
|  |    565 |     from rp have "cmod 0 \<le> r \<and> cmod (poly p 0) = - (- cmod (poly p 0))" by simp 
 | 
|  |    566 |     hence mth1: "\<exists>x z. cmod z \<le> r \<and> cmod (poly p z) = - x"  by blast
 | 
|  |    567 |     {fix x z
 | 
|  |    568 |       assume H: "cmod z \<le> r" "cmod (poly p z) = - x" "\<not>x < 1"
 | 
|  |    569 |       hence "- x < 0 " by arith
 | 
|  |    570 |       with H(2) cmod_pos[of "poly p z"]  have False by simp }
 | 
|  |    571 |     then have mth2: "\<exists>z. \<forall>x. (\<exists>z. cmod z \<le> r \<and> cmod (poly p z) = - x) \<longrightarrow> x < z" by blast
 | 
|  |    572 |     from real_sup_exists[OF mth1 mth2] obtain s where 
 | 
|  |    573 |       s: "\<forall>y. (\<exists>x. (\<exists>z. cmod z \<le> r \<and> cmod (poly p z) = - x) \<and> y < x) \<longleftrightarrow>(y < s)" by blast
 | 
|  |    574 |     let ?m = "-s"
 | 
|  |    575 |     {fix y
 | 
|  |    576 |       from s[rule_format, of "-y"] have 
 | 
|  |    577 |     "(\<exists>z x. cmod z \<le> r \<and> -(- cmod (poly p z)) < y) \<longleftrightarrow> ?m < y" 
 | 
|  |    578 | 	unfolding minus_less_iff[of y ] equation_minus_iff by blast }
 | 
|  |    579 |     note s1 = this[unfolded minus_minus]
 | 
|  |    580 |     from s1[of ?m] have s1m: "\<And>z x. cmod z \<le> r \<Longrightarrow> cmod (poly p z) \<ge> ?m" 
 | 
|  |    581 |       by auto
 | 
|  |    582 |     {fix n::nat
 | 
|  |    583 |       from s1[rule_format, of "?m + 1/real (Suc n)"] 
 | 
|  |    584 |       have "\<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)"
 | 
|  |    585 | 	by simp}
 | 
|  |    586 |     hence th: "\<forall>n. \<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)" ..
 | 
|  |    587 |     from choice[OF th] obtain g where 
 | 
|  |    588 |       g: "\<forall>n. cmod (g n) \<le> r" "\<forall>n. cmod (poly p (g n)) <?m+1 /real(Suc n)" 
 | 
|  |    589 |       by blast
 | 
|  |    590 |     from bolzano_weierstrass_complex_disc[OF g(1)] 
 | 
|  |    591 |     obtain f z where fz: "subseq f" "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. cmod (g (f n) - z) < e"
 | 
|  |    592 |       by blast    
 | 
|  |    593 |     {fix w 
 | 
|  |    594 |       assume wr: "cmod w \<le> r"
 | 
|  |    595 |       let ?e = "\<bar>cmod (poly p z) - ?m\<bar>"
 | 
|  |    596 |       {assume e: "?e > 0"
 | 
|  |    597 | 	hence e2: "?e/2 > 0" by simp
 | 
|  |    598 | 	from poly_cont[OF e2, of z p] obtain d where
 | 
|  |    599 | 	  d: "d>0" "\<forall>w. 0<cmod (w - z)\<and> cmod(w - z) < d \<longrightarrow> cmod(poly p w - poly p z) < ?e/2" by blast
 | 
|  |    600 | 	{fix w assume w: "cmod (w - z) < d"
 | 
|  |    601 | 	  have "cmod(poly p w - poly p z) < ?e / 2"
 | 
|  |    602 | 	    using d(2)[rule_format, of w] w e by (cases "w=z", simp_all)}
 | 
|  |    603 | 	note th1 = this
 | 
|  |    604 | 	
 | 
|  |    605 | 	from fz(2)[rule_format, OF d(1)] obtain N1 where 
 | 
|  |    606 | 	  N1: "\<forall>n\<ge>N1. cmod (g (f n) - z) < d" by blast
 | 
|  |    607 | 	from reals_Archimedean2[of "2/?e"] obtain N2::nat where
 | 
|  |    608 | 	  N2: "2/?e < real N2" by blast
 | 
|  |    609 | 	have th2: "cmod(poly p (g(f(N1 + N2))) - poly p z) < ?e/2"
 | 
|  |    610 | 	  using N1[rule_format, of "N1 + N2"] th1 by simp
 | 
|  |    611 | 	{fix a b e2 m :: real
 | 
|  |    612 | 	have "a < e2 \<Longrightarrow> abs(b - m) < e2 \<Longrightarrow> 2 * e2 <= abs(b - m) + a
 | 
|  |    613 |           ==> False" by arith}
 | 
|  |    614 |       note th0 = this
 | 
|  |    615 |       have ath: 
 | 
|  |    616 | 	"\<And>m x e. m <= x \<Longrightarrow>  x < m + e ==> abs(x - m::real) < e" by arith
 | 
|  |    617 |       from s1m[OF g(1)[rule_format]]
 | 
|  |    618 |       have th31: "?m \<le> cmod(poly p (g (f (N1 + N2))))" .
 | 
|  |    619 |       from seq_suble[OF fz(1), of "N1+N2"]
 | 
|  |    620 |       have th00: "real (Suc (N1+N2)) \<le> real (Suc (f (N1+N2)))" by simp
 | 
|  |    621 |       have th000: "0 \<le> (1::real)" "(1::real) \<le> 1" "real (Suc (N1+N2)) > 0"  
 | 
|  |    622 | 	using N2 by auto
 | 
|  |    623 |       from frac_le[OF th000 th00] have th00: "?m +1 / real (Suc (f (N1 + N2))) \<le> ?m + 1 / real (Suc (N1 + N2))" by simp
 | 
|  |    624 |       from g(2)[rule_format, of "f (N1 + N2)"]
 | 
|  |    625 |       have th01:"cmod (poly p (g (f (N1 + N2)))) < - s + 1 / real (Suc (f (N1 + N2)))" .
 | 
|  |    626 |       from order_less_le_trans[OF th01 th00]
 | 
|  |    627 |       have th32: "cmod(poly p (g (f (N1 + N2)))) < ?m + (1/ real(Suc (N1 + N2)))" .
 | 
|  |    628 |       from N2 have "2/?e < real (Suc (N1 + N2))" by arith
 | 
|  |    629 |       with e2 less_imp_inverse_less[of "2/?e" "real (Suc (N1 + N2))"]
 | 
|  |    630 |       have "?e/2 > 1/ real (Suc (N1 + N2))" by (simp add: inverse_eq_divide)
 | 
|  |    631 |       with ath[OF th31 th32]
 | 
|  |    632 |       have thc1:"\<bar>cmod(poly p (g (f (N1 + N2)))) - ?m\<bar>< ?e/2" by arith  
 | 
|  |    633 |       have ath2: "\<And>(a::real) b c m. \<bar>a - b\<bar> <= c ==> \<bar>b - m\<bar> <= \<bar>a - m\<bar> + c" 
 | 
|  |    634 | 	by arith
 | 
|  |    635 |       have th22: "\<bar>cmod (poly p (g (f (N1 + N2)))) - cmod (poly p z)\<bar>
 | 
|  |    636 | \<le> cmod (poly p (g (f (N1 + N2))) - poly p z)" 
 | 
|  |    637 | 	by (simp add: cmod_abs_norm)
 | 
|  |    638 |       from ath2[OF th22, of ?m]
 | 
|  |    639 |       have thc2: "2*(?e/2) \<le> \<bar>cmod(poly p (g (f (N1 + N2)))) - ?m\<bar> + cmod (poly p (g (f (N1 + N2))) - poly p z)" by simp
 | 
|  |    640 |       from th0[OF th2 thc1 thc2] have False .}
 | 
|  |    641 |       hence "?e = 0" by auto
 | 
|  |    642 |       then have "cmod (poly p z) = ?m" by simp  
 | 
|  |    643 |       with s1m[OF wr]
 | 
|  |    644 |       have "cmod (poly p z) \<le> cmod (poly p w)" by simp }
 | 
|  |    645 |     hence ?thesis by blast}
 | 
|  |    646 |   ultimately show ?thesis by blast
 | 
|  |    647 | qed
 | 
|  |    648 | 
 | 
|  |    649 | lemma "(rcis (sqrt (abs r)) (a/2)) ^ 2 = rcis (abs r) a"
 | 
|  |    650 |   unfolding power2_eq_square
 | 
|  |    651 |   apply (simp add: rcis_mult)
 | 
|  |    652 |   apply (simp add: power2_eq_square[symmetric])
 | 
|  |    653 |   done
 | 
|  |    654 | 
 | 
|  |    655 | lemma cispi: "cis pi = -1" 
 | 
|  |    656 |   unfolding cis_def
 | 
|  |    657 |   by simp
 | 
|  |    658 | 
 | 
|  |    659 | lemma "(rcis (sqrt (abs r)) ((pi + a)/2)) ^ 2 = rcis (- abs r) a"
 | 
|  |    660 |   unfolding power2_eq_square
 | 
|  |    661 |   apply (simp add: rcis_mult add_divide_distrib)
 | 
|  |    662 |   apply (simp add: power2_eq_square[symmetric] rcis_def cispi cis_mult[symmetric])
 | 
|  |    663 |   done
 | 
|  |    664 | 
 | 
|  |    665 | text {* Nonzero polynomial in z goes to infinity as z does. *}
 | 
|  |    666 | 
 | 
|  |    667 | instance complex::idom_char_0 by (intro_classes)
 | 
|  |    668 | instance complex :: recpower_idom_char_0 by intro_classes
 | 
|  |    669 | 
 | 
|  |    670 | lemma poly_infinity:
 | 
|  |    671 |   assumes ex: "list_ex (\<lambda>c. c \<noteq> 0) p"
 | 
|  |    672 |   shows "\<exists>r. \<forall>z. r \<le> cmod z \<longrightarrow> d \<le> cmod (poly (a#p) z)"
 | 
|  |    673 | using ex
 | 
|  |    674 | proof(induct p arbitrary: a d)
 | 
|  |    675 |   case (Cons c cs a d) 
 | 
|  |    676 |   {assume H: "list_ex (\<lambda>c. c\<noteq>0) cs"
 | 
|  |    677 |     with Cons.hyps obtain r where r: "\<forall>z. r \<le> cmod z \<longrightarrow> d + cmod a \<le> cmod (poly (c # cs) z)" by blast
 | 
|  |    678 |     let ?r = "1 + \<bar>r\<bar>"
 | 
|  |    679 |     {fix z assume h: "1 + \<bar>r\<bar> \<le> cmod z"
 | 
|  |    680 |       have r0: "r \<le> cmod z" using h by arith
 | 
|  |    681 |       from r[rule_format, OF r0]
 | 
|  |    682 |       have th0: "d + cmod a \<le> 1 * cmod(poly (c#cs) z)" by arith
 | 
|  |    683 |       from h have z1: "cmod z \<ge> 1" by arith
 | 
|  |    684 |       from order_trans[OF th0 mult_right_mono[OF z1 cmod_pos[of "poly (c#cs) z"]]]
 | 
|  |    685 |       have th1: "d \<le> cmod(z * poly (c#cs) z) - cmod a"
 | 
|  |    686 | 	unfolding cmod_mult by (simp add: ring_simps)
 | 
|  |    687 |       from complex_mod_triangle_sub[of "z * poly (c#cs) z" a]
 | 
|  |    688 |       have th2: "cmod(z * poly (c#cs) z) - cmod a \<le> cmod (poly (a#c#cs) z)" 
 | 
|  |    689 | 	by (simp add: diff_le_eq ring_simps) 
 | 
|  |    690 |       from th1 th2 have "d \<le> cmod (poly (a#c#cs) z)"  by arith}
 | 
|  |    691 |     hence ?case by blast}
 | 
|  |    692 |   moreover
 | 
|  |    693 |   {assume cs0: "\<not> (list_ex (\<lambda>c. c \<noteq> 0) cs)"
 | 
|  |    694 |     with Cons.prems have c0: "c \<noteq> 0" by simp
 | 
|  |    695 |     from cs0 have cs0': "list_all (\<lambda>c. c = 0) cs" 
 | 
|  |    696 |       by (auto simp add: list_all_iff list_ex_iff)
 | 
|  |    697 |     {fix z
 | 
|  |    698 |       assume h: "(\<bar>d\<bar> + cmod a) / cmod c \<le> cmod z"
 | 
|  |    699 |       from c0 have "cmod c > 0" by simp
 | 
|  |    700 |       from h c0 have th0: "\<bar>d\<bar> + cmod a \<le> cmod (z*c)" 
 | 
|  |    701 | 	by (simp add: field_simps cmod_mult)
 | 
|  |    702 |       have ath: "\<And>mzh mazh ma. mzh <= mazh + ma ==> abs(d) + ma <= mzh ==> d <= mazh" by arith
 | 
|  |    703 |       from complex_mod_triangle_sub[of "z*c" a ]
 | 
|  |    704 |       have th1: "cmod (z * c) \<le> cmod (a + z * c) + cmod a"
 | 
|  |    705 | 	by (simp add: ring_simps)
 | 
|  |    706 |       from ath[OF th1 th0] have "d \<le> cmod (poly (a # c # cs) z)" 
 | 
|  |    707 | 	using poly_0[OF cs0'] by simp}
 | 
|  |    708 |     then have ?case  by blast}
 | 
|  |    709 |   ultimately show ?case by blast
 | 
|  |    710 | qed simp
 | 
|  |    711 | 
 | 
|  |    712 | text {* Hence polynomial's modulus attains its minimum somewhere. *}
 | 
|  |    713 | lemma poly_minimum_modulus:
 | 
|  |    714 |   "\<exists>z.\<forall>w. cmod (poly p z) \<le> cmod (poly p w)"
 | 
|  |    715 | proof(induct p)
 | 
|  |    716 |   case (Cons c cs) 
 | 
|  |    717 |   {assume cs0: "list_ex (\<lambda>c. c \<noteq> 0) cs"
 | 
|  |    718 |     from poly_infinity[OF cs0, of "cmod (poly (c#cs) 0)" c]
 | 
|  |    719 |     obtain r where r: "\<And>z. r \<le> cmod z \<Longrightarrow> cmod (poly (c # cs) 0) \<le> cmod (poly (c # cs) z)" by blast
 | 
|  |    720 |     have ath: "\<And>z r. r \<le> cmod z \<or> cmod z \<le> \<bar>r\<bar>" by arith
 | 
|  |    721 |     from poly_minimum_modulus_disc[of "\<bar>r\<bar>" "c#cs"] 
 | 
|  |    722 |     obtain v where v: "\<And>w. cmod w \<le> \<bar>r\<bar> \<Longrightarrow> cmod (poly (c # cs) v) \<le> cmod (poly (c # cs) w)" by blast
 | 
|  |    723 |     {fix z assume z: "r \<le> cmod z"
 | 
|  |    724 |       from v[of 0] r[OF z] 
 | 
|  |    725 |       have "cmod (poly (c # cs) v) \<le> cmod (poly (c # cs) z)"
 | 
|  |    726 | 	by simp }
 | 
|  |    727 |     note v0 = this
 | 
|  |    728 |     from v0 v ath[of r] have ?case by blast}
 | 
|  |    729 |   moreover
 | 
|  |    730 |   {assume cs0: "\<not> (list_ex (\<lambda>c. c\<noteq>0) cs)"
 | 
|  |    731 |     hence th:"list_all (\<lambda>c. c = 0) cs" by (simp add: list_all_iff list_ex_iff)
 | 
|  |    732 |     from poly_0[OF th] Cons.hyps have ?case by simp}
 | 
|  |    733 |   ultimately show ?case by blast
 | 
|  |    734 | qed simp
 | 
|  |    735 | 
 | 
|  |    736 | text{* Constant function (non-syntactic characterization). *}
 | 
|  |    737 | definition "constant f = (\<forall>x y. f x = f y)"
 | 
|  |    738 | 
 | 
|  |    739 | lemma nonconstant_length: "\<not> (constant (poly p)) \<Longrightarrow> length p \<ge> 2"
 | 
|  |    740 |   unfolding constant_def
 | 
|  |    741 |   apply (induct p, auto)
 | 
|  |    742 |   apply (unfold not_less[symmetric])
 | 
|  |    743 |   apply simp
 | 
|  |    744 |   apply (rule ccontr)
 | 
|  |    745 |   apply auto
 | 
|  |    746 |   done
 | 
|  |    747 |  
 | 
|  |    748 | lemma poly_replicate_append:
 | 
|  |    749 |   "poly ((replicate n 0)@p) (x::'a::{recpower, comm_ring}) = x^n * poly p x"
 | 
|  |    750 |   by(induct n, auto simp add: power_Suc ring_simps)
 | 
|  |    751 | 
 | 
|  |    752 | text {* Decomposition of polynomial, skipping zero coefficients 
 | 
|  |    753 |   after the first.  *}
 | 
|  |    754 | 
 | 
|  |    755 | lemma poly_decompose_lemma:
 | 
|  |    756 |  assumes nz: "\<not>(\<forall>z. z\<noteq>0 \<longrightarrow> poly p z = (0::'a::{recpower,idom}))"
 | 
|  |    757 |   shows "\<exists>k a q. a\<noteq>0 \<and> Suc (length q + k) = length p \<and> 
 | 
|  |    758 |                  (\<forall>z. poly p z = z^k * poly (a#q) z)"
 | 
|  |    759 | using nz
 | 
|  |    760 | proof(induct p)
 | 
|  |    761 |   case Nil thus ?case by simp
 | 
|  |    762 | next
 | 
|  |    763 |   case (Cons c cs)
 | 
|  |    764 |   {assume c0: "c = 0"
 | 
|  |    765 |     
 | 
|  |    766 |     from Cons.hyps Cons.prems c0 have ?case apply auto
 | 
|  |    767 |       apply (rule_tac x="k+1" in exI)
 | 
|  |    768 |       apply (rule_tac x="a" in exI, clarsimp)
 | 
|  |    769 |       apply (rule_tac x="q" in exI)
 | 
|  |    770 |       by (auto simp add: power_Suc)}
 | 
|  |    771 |   moreover
 | 
|  |    772 |   {assume c0: "c\<noteq>0"
 | 
|  |    773 |     hence ?case apply-
 | 
|  |    774 |       apply (rule exI[where x=0])
 | 
|  |    775 |       apply (rule exI[where x=c], clarsimp)
 | 
|  |    776 |       apply (rule exI[where x=cs])
 | 
|  |    777 |       apply auto
 | 
|  |    778 |       done}
 | 
|  |    779 |   ultimately show ?case by blast
 | 
|  |    780 | qed
 | 
|  |    781 | 
 | 
|  |    782 | lemma poly_decompose:
 | 
|  |    783 |   assumes nc: "~constant(poly p)"
 | 
|  |    784 |   shows "\<exists>k a q. a\<noteq>(0::'a::{recpower,idom}) \<and> k\<noteq>0 \<and>
 | 
|  |    785 |                length q + k + 1 = length p \<and> 
 | 
|  |    786 |               (\<forall>z. poly p z = poly p 0 + z^k * poly (a#q) z)"
 | 
|  |    787 | using nc 
 | 
|  |    788 | proof(induct p)
 | 
|  |    789 |   case Nil thus ?case by (simp add: constant_def)
 | 
|  |    790 | next
 | 
|  |    791 |   case (Cons c cs)
 | 
|  |    792 |   {assume C:"\<forall>z. z \<noteq> 0 \<longrightarrow> poly cs z = 0"
 | 
|  |    793 |     {fix x y
 | 
|  |    794 |       from C have "poly (c#cs) x = poly (c#cs) y" by (cases "x=0", auto)}
 | 
|  |    795 |     with Cons.prems have False by (auto simp add: constant_def)}
 | 
|  |    796 |   hence th: "\<not> (\<forall>z. z \<noteq> 0 \<longrightarrow> poly cs z = 0)" ..
 | 
|  |    797 |   from poly_decompose_lemma[OF th] 
 | 
|  |    798 |   show ?case 
 | 
|  |    799 |     apply clarsimp    
 | 
|  |    800 |     apply (rule_tac x="k+1" in exI)
 | 
|  |    801 |     apply (rule_tac x="a" in exI)
 | 
|  |    802 |     apply simp
 | 
|  |    803 |     apply (rule_tac x="q" in exI)
 | 
|  |    804 |     apply (auto simp add: power_Suc)
 | 
|  |    805 |     done
 | 
|  |    806 | qed
 | 
|  |    807 | 
 | 
|  |    808 | text{* Fundamental theorem of algebral *}
 | 
|  |    809 | 
 | 
|  |    810 | lemma fundamental_theorem_of_algebra:
 | 
|  |    811 |   assumes nc: "~constant(poly p)"
 | 
|  |    812 |   shows "\<exists>z::complex. poly p z = 0"
 | 
|  |    813 | using nc
 | 
|  |    814 | proof(induct n\<equiv> "length p" arbitrary: p rule: nat_less_induct)
 | 
|  |    815 |   fix n fix p :: "complex list"
 | 
|  |    816 |   let ?p = "poly p"
 | 
|  |    817 |   assume H: "\<forall>m<n. \<forall>p. \<not> constant (poly p) \<longrightarrow> m = length p \<longrightarrow> (\<exists>(z::complex). poly p z = 0)" and nc: "\<not> constant ?p" and n: "n = length p"
 | 
|  |    818 |   let ?ths = "\<exists>z. ?p z = 0"
 | 
|  |    819 | 
 | 
|  |    820 |   from nonconstant_length[OF nc] have n2: "n\<ge> 2" by (simp add: n)
 | 
|  |    821 |   from poly_minimum_modulus obtain c where 
 | 
|  |    822 |     c: "\<forall>w. cmod (?p c) \<le> cmod (?p w)" by blast
 | 
|  |    823 |   {assume pc: "?p c = 0" hence ?ths by blast}
 | 
|  |    824 |   moreover
 | 
|  |    825 |   {assume pc0: "?p c \<noteq> 0"
 | 
|  |    826 |     from poly_offset[of p c] obtain q where
 | 
|  |    827 |       q: "length q = length p" "\<forall>x. poly q x = ?p (c+x)" by blast
 | 
|  |    828 |     {assume h: "constant (poly q)"
 | 
|  |    829 |       from q(2) have th: "\<forall>x. poly q (x - c) = ?p x" by auto
 | 
|  |    830 |       {fix x y
 | 
|  |    831 | 	from th have "?p x = poly q (x - c)" by auto 
 | 
|  |    832 | 	also have "\<dots> = poly q (y - c)" 
 | 
|  |    833 | 	  using h unfolding constant_def by blast
 | 
|  |    834 | 	also have "\<dots> = ?p y" using th by auto
 | 
|  |    835 | 	finally have "?p x = ?p y" .}
 | 
|  |    836 |       with nc have False unfolding constant_def by blast }
 | 
|  |    837 |     hence qnc: "\<not> constant (poly q)" by blast
 | 
|  |    838 |     from q(2) have pqc0: "?p c = poly q 0" by simp
 | 
|  |    839 |     from c pqc0 have cq0: "\<forall>w. cmod (poly q 0) \<le> cmod (?p w)" by simp 
 | 
|  |    840 |     let ?a0 = "poly q 0"
 | 
|  |    841 |     from pc0 pqc0 have a00: "?a0 \<noteq> 0" by simp 
 | 
|  |    842 |     from a00 
 | 
|  |    843 |     have qr: "\<forall>z. poly q z = poly (map (op * (inverse ?a0)) q) z * ?a0"
 | 
|  |    844 |       by (simp add: poly_cmult_map)
 | 
|  |    845 |     let ?r = "map (op * (inverse ?a0)) q"
 | 
|  |    846 |     have lgqr: "length q = length ?r" by simp 
 | 
|  |    847 |     {assume h: "\<And>x y. poly ?r x = poly ?r y"
 | 
|  |    848 |       {fix x y
 | 
|  |    849 | 	from qr[rule_format, of x] 
 | 
|  |    850 | 	have "poly q x = poly ?r x * ?a0" by auto
 | 
|  |    851 | 	also have "\<dots> = poly ?r y * ?a0" using h by simp
 | 
|  |    852 | 	also have "\<dots> = poly q y" using qr[rule_format, of y] by simp
 | 
|  |    853 | 	finally have "poly q x = poly q y" .} 
 | 
|  |    854 |       with qnc have False unfolding constant_def by blast}
 | 
|  |    855 |     hence rnc: "\<not> constant (poly ?r)" unfolding constant_def by blast
 | 
|  |    856 |     from qr[rule_format, of 0] a00  have r01: "poly ?r 0 = 1" by auto
 | 
|  |    857 |     {fix w 
 | 
|  |    858 |       have "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w / ?a0) < 1"
 | 
|  |    859 | 	using qr[rule_format, of w] a00 by simp
 | 
|  |    860 |       also have "\<dots> \<longleftrightarrow> cmod (poly q w) < cmod ?a0"
 | 
|  |    861 | 	using a00 unfolding cmod_divide by (simp add: field_simps)
 | 
|  |    862 |       finally have "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w) < cmod ?a0" .}
 | 
|  |    863 |     note mrmq_eq = this
 | 
|  |    864 |     from poly_decompose[OF rnc] obtain k a s where 
 | 
|  |    865 |       kas: "a\<noteq>0" "k\<noteq>0" "length s + k + 1 = length ?r" 
 | 
|  |    866 |       "\<forall>z. poly ?r z = poly ?r 0 + z^k* poly (a#s) z" by blast
 | 
|  |    867 |     {assume "k + 1 = n"
 | 
|  |    868 |       with kas(3) lgqr[symmetric] q(1) n[symmetric] have s0:"s=[]" by auto
 | 
|  |    869 |       {fix w
 | 
|  |    870 | 	have "cmod (poly ?r w) = cmod (1 + a * w ^ k)" 
 | 
|  |    871 | 	  using kas(4)[rule_format, of w] s0 r01 by (simp add: ring_simps)}
 | 
|  |    872 |       note hth = this [symmetric]
 | 
|  |    873 | 	from reduce_poly_simple[OF kas(1,2)] 
 | 
|  |    874 |       have "\<exists>w. cmod (poly ?r w) < 1" unfolding hth by blast}
 | 
|  |    875 |     moreover
 | 
|  |    876 |     {assume kn: "k+1 \<noteq> n"
 | 
|  |    877 |       from kn kas(3) q(1) n[symmetric] have k1n: "k + 1 < n" by simp
 | 
|  |    878 |       have th01: "\<not> constant (poly (1#((replicate (k - 1) 0)@[a])))" 
 | 
|  |    879 | 	unfolding constant_def poly_Nil poly_Cons poly_replicate_append
 | 
|  |    880 | 	using kas(1) apply simp 
 | 
|  |    881 | 	by (rule exI[where x=0], rule exI[where x=1], simp)
 | 
|  |    882 |       from kas(2) have th02: "k+1 = length (1#((replicate (k - 1) 0)@[a]))" 
 | 
|  |    883 | 	by simp
 | 
|  |    884 |       from H[rule_format, OF k1n th01 th02]
 | 
|  |    885 |       obtain w where w: "1 + w^k * a = 0"
 | 
|  |    886 | 	unfolding poly_Nil poly_Cons poly_replicate_append
 | 
|  |    887 | 	using kas(2) by (auto simp add: power_Suc[symmetric, of _ "k - Suc 0"] 
 | 
|  |    888 | 	  mult_assoc[of _ _ a, symmetric])
 | 
|  |    889 |       from poly_bound_exists[of "cmod w" s] obtain m where 
 | 
|  |    890 | 	m: "m > 0" "\<forall>z. cmod z \<le> cmod w \<longrightarrow> cmod (poly s z) \<le> m" by blast
 | 
|  |    891 |       have w0: "w\<noteq>0" using kas(2) w by (auto simp add: power_0_left)
 | 
|  |    892 |       from w have "(1 + w ^ k * a) - 1 = 0 - 1" by simp
 | 
|  |    893 |       then have wm1: "w^k * a = - 1" by simp
 | 
|  |    894 |       have inv0: "0 < inverse (cmod w ^ (k + 1) * m)" 
 | 
|  |    895 | 	using cmod_pos[of w] w0 m(1)
 | 
|  |    896 | 	  by (simp add: inverse_eq_divide zero_less_mult_iff)
 | 
|  |    897 |       with real_down2[OF zero_less_one] obtain t where
 | 
|  |    898 | 	t: "t > 0" "t < 1" "t < inverse (cmod w ^ (k + 1) * m)" by blast
 | 
|  |    899 |       let ?ct = "complex_of_real t"
 | 
|  |    900 |       let ?w = "?ct * w"
 | 
|  |    901 |       have "1 + ?w^k * (a + ?w * poly s ?w) = 1 + ?ct^k * (w^k * a) + ?w^k * ?w * poly s ?w" using kas(1) by (simp add: ring_simps power_mult_distrib)
 | 
|  |    902 |       also have "\<dots> = complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w"
 | 
|  |    903 | 	unfolding wm1 by (simp)
 | 
|  |    904 |       finally have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) = cmod (complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w)" 
 | 
|  |    905 | 	apply -
 | 
|  |    906 | 	apply (rule cong[OF refl[of cmod]])
 | 
|  |    907 | 	apply assumption
 | 
|  |    908 | 	done
 | 
|  |    909 |       with complex_mod_triangle_ineq[of "complex_of_real (1 - t^k)" "?w^k * ?w * poly s ?w"] 
 | 
|  |    910 |       have th11: "cmod (1 + ?w^k * (a + ?w * poly s ?w)) \<le> \<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w)" unfolding cmod_complex_of_real by simp 
 | 
|  |    911 |       have ath: "\<And>x (t::real). 0\<le> x \<Longrightarrow> x < t \<Longrightarrow> t\<le>1 \<Longrightarrow> \<bar>1 - t\<bar> + x < 1" by arith
 | 
|  |    912 |       have "t *cmod w \<le> 1 * cmod w" apply (rule mult_mono) using t(1,2) by auto
 | 
|  |    913 |       then have tw: "cmod ?w \<le> cmod w" using t(1) by (simp add: cmod_mult) 
 | 
|  |    914 |       from t inv0 have "t* (cmod w ^ (k + 1) * m) < 1"
 | 
|  |    915 | 	by (simp add: inverse_eq_divide field_simps)
 | 
|  |    916 |       with zero_less_power[OF t(1), of k] 
 | 
|  |    917 |       have th30: "t^k * (t* (cmod w ^ (k + 1) * m)) < t^k * 1" 
 | 
|  |    918 | 	apply - apply (rule mult_strict_left_mono) by simp_all
 | 
|  |    919 |       have "cmod (?w^k * ?w * poly s ?w) = t^k * (t* (cmod w ^ (k+1) * cmod (poly s ?w)))"  using w0 t(1)
 | 
|  |    920 | 	by (simp add: ring_simps power_mult_distrib cmod_complex_of_real cmod_power cmod_mult)
 | 
|  |    921 |       then have "cmod (?w^k * ?w * poly s ?w) \<le> t^k * (t* (cmod w ^ (k + 1) * m))"
 | 
|  |    922 | 	using t(1,2) m(2)[rule_format, OF tw] w0
 | 
|  |    923 | 	apply (simp only: )
 | 
|  |    924 | 	apply auto
 | 
|  |    925 | 	apply (rule mult_mono, simp_all add: cmod_pos)+
 | 
|  |    926 | 	apply (simp add: zero_le_mult_iff zero_le_power)
 | 
|  |    927 | 	done
 | 
|  |    928 |       with th30 have th120: "cmod (?w^k * ?w * poly s ?w) < t^k" by simp 
 | 
|  |    929 |       from power_strict_mono[OF t(2), of k] t(1) kas(2) have th121: "t^k \<le> 1" 
 | 
|  |    930 | 	by auto
 | 
|  |    931 |       from ath[OF cmod_pos[of "?w^k * ?w * poly s ?w"] th120 th121]
 | 
|  |    932 |       have th12: "\<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w) < 1" . 
 | 
|  |    933 |       from th11 th12
 | 
|  |    934 |       have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) < 1"  by arith 
 | 
|  |    935 |       then have "cmod (poly ?r ?w) < 1" 
 | 
|  |    936 | 	unfolding kas(4)[rule_format, of ?w] r01 by simp 
 | 
|  |    937 |       then have "\<exists>w. cmod (poly ?r w) < 1" by blast}
 | 
|  |    938 |     ultimately have cr0_contr: "\<exists>w. cmod (poly ?r w) < 1" by blast
 | 
|  |    939 |     from cr0_contr cq0 q(2)
 | 
|  |    940 |     have ?ths unfolding mrmq_eq not_less[symmetric] by auto}
 | 
|  |    941 |   ultimately show ?ths by blast
 | 
|  |    942 | qed
 | 
|  |    943 | 
 | 
|  |    944 | text {* Alternative version with a syntactic notion of constant polynomial. *}
 | 
|  |    945 | 
 | 
|  |    946 | lemma fundamental_theorem_of_algebra_alt:
 | 
|  |    947 |   assumes nc: "~(\<exists>a l. a\<noteq> 0 \<and> list_all(\<lambda>b. b = 0) l \<and> p = a#l)"
 | 
|  |    948 |   shows "\<exists>z. poly p z = (0::complex)"
 | 
|  |    949 | using nc
 | 
|  |    950 | proof(induct p)
 | 
|  |    951 |   case (Cons c cs)
 | 
|  |    952 |   {assume "c=0" hence ?case by auto}
 | 
|  |    953 |   moreover
 | 
|  |    954 |   {assume c0: "c\<noteq>0"
 | 
|  |    955 |     {assume nc: "constant (poly (c#cs))"
 | 
|  |    956 |       from nc[unfolded constant_def, rule_format, of 0] 
 | 
|  |    957 |       have "\<forall>w. w \<noteq> 0 \<longrightarrow> poly cs w = 0" by auto 
 | 
|  |    958 |       hence "list_all (\<lambda>c. c=0) cs"
 | 
|  |    959 | 	proof(induct cs)
 | 
|  |    960 | 	  case (Cons d ds)
 | 
|  |    961 | 	  {assume "d=0" hence ?case using Cons.prems Cons.hyps by simp}
 | 
|  |    962 | 	  moreover
 | 
|  |    963 | 	  {assume d0: "d\<noteq>0"
 | 
|  |    964 | 	    from poly_bound_exists[of 1 ds] obtain m where 
 | 
|  |    965 | 	      m: "m > 0" "\<forall>z. \<forall>z. cmod z \<le> 1 \<longrightarrow> cmod (poly ds z) \<le> m" by blast
 | 
|  |    966 | 	    have dm: "cmod d / m > 0" using d0 m(1) by (simp add: field_simps)
 | 
|  |    967 | 	    from real_down2[OF dm zero_less_one] obtain x where 
 | 
|  |    968 | 	      x: "x > 0" "x < cmod d / m" "x < 1" by blast
 | 
|  |    969 | 	    let ?x = "complex_of_real x"
 | 
|  |    970 | 	    from x have cx: "?x \<noteq> 0"  "cmod ?x \<le> 1" by simp_all
 | 
|  |    971 | 	    from Cons.prems[rule_format, OF cx(1)]
 | 
|  |    972 | 	    have cth: "cmod (?x*poly ds ?x) = cmod d" by (simp add: eq_diff_eq[symmetric])
 | 
|  |    973 | 	    from m(2)[rule_format, OF cx(2)] x(1)
 | 
|  |    974 | 	    have th0: "cmod (?x*poly ds ?x) \<le> x*m"
 | 
|  |    975 | 	      by (simp add: cmod_mult)
 | 
|  |    976 | 	    from x(2) m(1) have "x*m < cmod d" by (simp add: field_simps)
 | 
|  |    977 | 	    with th0 have "cmod (?x*poly ds ?x) \<noteq> cmod d" by auto
 | 
|  |    978 | 	    with cth  have ?case by blast}
 | 
|  |    979 | 	  ultimately show ?case by blast 
 | 
|  |    980 | 	qed simp}
 | 
|  |    981 |       then have nc: "\<not> constant (poly (c#cs))" using Cons.prems c0 
 | 
|  |    982 | 	by blast
 | 
|  |    983 |       from fundamental_theorem_of_algebra[OF nc] have ?case .}
 | 
|  |    984 |   ultimately show ?case by blast  
 | 
|  |    985 | qed simp
 | 
|  |    986 | 
 | 
|  |    987 | section{* Nullstellenstatz, degrees and divisibility of polynomials *}
 | 
|  |    988 | 
 | 
|  |    989 | lemma nullstellensatz_lemma:
 | 
|  |    990 |   fixes p :: "complex list"
 | 
|  |    991 |   assumes "\<forall>x. poly p x = 0 \<longrightarrow> poly q x = 0"
 | 
|  |    992 |   and "degree p = n" and "n \<noteq> 0"
 | 
|  |    993 |   shows "p divides (pexp q n)"
 | 
|  |    994 | using prems
 | 
|  |    995 | proof(induct n arbitrary: p q rule: nat_less_induct)
 | 
|  |    996 |   fix n::nat fix p q :: "complex list"
 | 
|  |    997 |   assume IH: "\<forall>m<n. \<forall>p q.
 | 
|  |    998 |                  (\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longrightarrow>
 | 
|  |    999 |                  degree p = m \<longrightarrow> m \<noteq> 0 \<longrightarrow> p divides (q %^ m)"
 | 
|  |   1000 |     and pq0: "\<forall>x. poly p x = 0 \<longrightarrow> poly q x = 0" 
 | 
|  |   1001 |     and dpn: "degree p = n" and n0: "n \<noteq> 0"
 | 
|  |   1002 |   let ?ths = "p divides (q %^ n)"
 | 
|  |   1003 |   {fix a assume a: "poly p a = 0"
 | 
|  |   1004 |     {assume p0: "poly p = poly []" 
 | 
|  |   1005 |       hence ?ths unfolding divides_def  using pq0 n0
 | 
|  |   1006 | 	apply - apply (rule exI[where x="[]"], rule ext)
 | 
|  |   1007 | 	by (auto simp add: poly_mult poly_exp)}
 | 
|  |   1008 |     moreover
 | 
|  |   1009 |     {assume p0: "poly p \<noteq> poly []" 
 | 
|  |   1010 |       and oa: "order  a p \<noteq> 0"
 | 
|  |   1011 |       from p0 have pne: "p \<noteq> []" by auto
 | 
|  |   1012 |       let ?op = "order a p"
 | 
|  |   1013 |       from p0 have ap: "([- a, 1] %^ ?op) divides p" 
 | 
|  |   1014 | 	"\<not> pexp [- a, 1] (Suc ?op) divides p" using order by blast+ 
 | 
|  |   1015 |       note oop = order_degree[OF p0, unfolded dpn]
 | 
|  |   1016 |       {assume q0: "q = []"
 | 
|  |   1017 | 	hence ?ths using n0 unfolding divides_def 
 | 
|  |   1018 | 	  apply simp
 | 
|  |   1019 | 	  apply (rule exI[where x="[]"], rule ext)
 | 
|  |   1020 | 	  by (simp add: divides_def poly_exp poly_mult)}
 | 
|  |   1021 |       moreover
 | 
|  |   1022 |       {assume q0: "q\<noteq>[]"
 | 
|  |   1023 | 	from pq0[rule_format, OF a, unfolded poly_linear_divides] q0
 | 
|  |   1024 | 	obtain r where r: "q = pmult [- a, 1] r" by blast
 | 
|  |   1025 | 	from ap[unfolded divides_def] obtain s where
 | 
|  |   1026 | 	  s: "poly p = poly (pmult (pexp [- a, 1] ?op) s)" by blast
 | 
|  |   1027 | 	have s0: "poly s \<noteq> poly []"
 | 
|  |   1028 | 	  using s p0 by (simp add: poly_entire)
 | 
|  |   1029 | 	hence pns0: "poly (pnormalize s) \<noteq> poly []" and sne: "s\<noteq>[]" by auto
 | 
|  |   1030 | 	{assume ds0: "degree s = 0"
 | 
|  |   1031 | 	  from ds0 pns0 have "\<exists>k. pnormalize s = [k]" unfolding degree_def 
 | 
|  |   1032 | 	    by (cases "pnormalize s", auto)
 | 
|  |   1033 | 	  then obtain k where kpn: "pnormalize s = [k]" by blast
 | 
|  |   1034 | 	  from pns0[unfolded poly_zero] kpn have k: "k \<noteq>0" "poly s = poly [k]"
 | 
|  |   1035 | 	    using poly_normalize[of s] by simp_all
 | 
|  |   1036 | 	  let ?w = "pmult (pmult [1/k] (pexp [-a,1] (n - ?op))) (pexp r n)"
 | 
|  |   1037 | 	  from k r s oop have "poly (pexp q n) = poly (pmult p ?w)"
 | 
|  |   1038 | 	    by - (rule ext, simp add: poly_mult poly_exp poly_cmult poly_add power_add[symmetric] ring_simps power_mult_distrib[symmetric])
 | 
|  |   1039 | 	  hence ?ths unfolding divides_def by blast}
 | 
|  |   1040 | 	moreover
 | 
|  |   1041 | 	{assume ds0: "degree s \<noteq> 0"
 | 
|  |   1042 | 	  from ds0 s0 dpn degree_unique[OF s, unfolded linear_pow_mul_degree] oa
 | 
|  |   1043 | 	    have dsn: "degree s < n" by auto 
 | 
|  |   1044 | 	    {fix x assume h: "poly s x = 0"
 | 
|  |   1045 | 	      {assume xa: "x = a"
 | 
|  |   1046 | 		from h[unfolded xa poly_linear_divides] sne obtain u where
 | 
|  |   1047 | 		  u: "s = pmult [- a, 1] u" by blast
 | 
|  |   1048 | 		have "poly p = poly (pmult (pexp [- a, 1] (Suc ?op)) u)"
 | 
|  |   1049 | 		  unfolding s u
 | 
|  |   1050 | 		  apply (rule ext)
 | 
|  |   1051 | 		  by (simp add: ring_simps power_mult_distrib[symmetric] poly_mult poly_cmult poly_add poly_exp)
 | 
|  |   1052 | 		with ap(2)[unfolded divides_def] have False by blast}
 | 
|  |   1053 | 	      note xa = this
 | 
|  |   1054 | 	      from h s have "poly p x = 0" by (simp add: poly_mult)
 | 
|  |   1055 | 	      with pq0 have "poly q x = 0" by blast
 | 
|  |   1056 | 	      with r xa have "poly r x = 0"
 | 
|  |   1057 | 		by (auto simp add: poly_mult poly_add poly_cmult eq_diff_eq[symmetric])}
 | 
|  |   1058 | 	    note impth = this
 | 
|  |   1059 | 	    from IH[rule_format, OF dsn, of s r] impth ds0
 | 
|  |   1060 | 	    have "s divides (pexp r (degree s))" by blast
 | 
|  |   1061 | 	    then obtain u where u: "poly (pexp r (degree s)) = poly (pmult s u)"
 | 
|  |   1062 | 	      unfolding divides_def by blast
 | 
|  |   1063 | 	    hence u': "\<And>x. poly s x * poly u x = poly r x ^ degree s"
 | 
|  |   1064 | 	      by (simp add: poly_mult[symmetric] poly_exp[symmetric])
 | 
|  |   1065 | 	    let ?w = "pmult (pmult u (pexp [-a,1] (n - ?op))) (pexp r (n - degree s))"
 | 
|  |   1066 | 	    from u' s r oop[of a] dsn have "poly (pexp q n) = poly (pmult p ?w)"
 | 
|  |   1067 | 	      apply - apply (rule ext)
 | 
|  |   1068 | 	      apply (simp only:  power_mult_distrib power_add[symmetric] poly_add poly_mult poly_exp poly_cmult ring_simps)
 | 
|  |   1069 | 	      
 | 
|  |   1070 | 	      apply (simp add:  power_mult_distrib power_add[symmetric] poly_add poly_mult poly_exp poly_cmult mult_assoc[symmetric])
 | 
|  |   1071 | 	      done
 | 
|  |   1072 | 	    hence ?ths unfolding divides_def by blast}
 | 
|  |   1073 |       ultimately have ?ths by blast }
 | 
|  |   1074 |       ultimately have ?ths by blast}
 | 
|  |   1075 |     ultimately have ?ths using a order_root by blast}
 | 
|  |   1076 |   moreover
 | 
|  |   1077 |   {assume exa: "\<not> (\<exists>a. poly p a = 0)"
 | 
|  |   1078 |     from fundamental_theorem_of_algebra_alt[of p] exa obtain c cs where
 | 
|  |   1079 |       ccs: "c\<noteq>0" "list_all (\<lambda>c. c = 0) cs" "p = c#cs" by blast
 | 
|  |   1080 |     
 | 
|  |   1081 |     from poly_0[OF ccs(2)] ccs(3) 
 | 
|  |   1082 |     have pp: "\<And>x. poly p x =  c" by simp
 | 
|  |   1083 |     let ?w = "pmult [1/c] (pexp q n)"
 | 
|  |   1084 |     from pp ccs(1) 
 | 
|  |   1085 |     have "poly (pexp q n) = poly (pmult p ?w) "
 | 
|  |   1086 |       apply - apply (rule ext)
 | 
|  |   1087 |       unfolding poly_mult_assoc[symmetric] by (simp add: poly_mult)
 | 
|  |   1088 |     hence ?ths unfolding divides_def by blast}
 | 
|  |   1089 |   ultimately show ?ths by blast
 | 
|  |   1090 | qed
 | 
|  |   1091 | 
 | 
|  |   1092 | lemma nullstellensatz_univariate:
 | 
|  |   1093 |   "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow> 
 | 
|  |   1094 |     p divides (q %^ (degree p)) \<or> (poly p = poly [] \<and> poly q = poly [])"
 | 
|  |   1095 | proof-
 | 
|  |   1096 |   {assume pe: "poly p = poly []"
 | 
|  |   1097 |     hence eq: "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow> poly q = poly []"
 | 
|  |   1098 |       apply auto
 | 
|  |   1099 |       by (rule ext, simp)
 | 
|  |   1100 |     {assume "p divides (pexp q (degree p))"
 | 
|  |   1101 |       then obtain r where r: "poly (pexp q (degree p)) = poly (pmult p r)" 
 | 
|  |   1102 | 	unfolding divides_def by blast
 | 
|  |   1103 |       from cong[OF r refl] pe degree_unique[OF pe]
 | 
|  |   1104 |       have False by (simp add: poly_mult degree_def)}
 | 
|  |   1105 |     with eq pe have ?thesis by blast}
 | 
|  |   1106 |   moreover
 | 
|  |   1107 |   {assume pe: "poly p \<noteq> poly []"
 | 
|  |   1108 |     have p0: "poly [0] = poly []" by (rule ext, simp)
 | 
|  |   1109 |     {assume dp: "degree p = 0"
 | 
|  |   1110 |       then obtain k where "pnormalize p = [k]" using pe poly_normalize[of p]
 | 
|  |   1111 | 	unfolding degree_def by (cases "pnormalize p", auto)
 | 
|  |   1112 |       hence k: "pnormalize p = [k]" "poly p = poly [k]" "k\<noteq>0"
 | 
|  |   1113 | 	using pe poly_normalize[of p] by (auto simp add: p0)
 | 
|  |   1114 |       hence th1: "\<forall>x. poly p x \<noteq> 0" by simp
 | 
|  |   1115 |       from k(2,3) dp have "poly (pexp q (degree p)) = poly (pmult p [1/k]) "
 | 
|  |   1116 | 	by - (rule ext, simp add: poly_mult poly_exp)
 | 
|  |   1117 |       hence th2: "p divides (pexp q (degree p))" unfolding divides_def by blast
 | 
|  |   1118 |       from th1 th2 pe have ?thesis by blast}
 | 
|  |   1119 |     moreover
 | 
|  |   1120 |     {assume dp: "degree p \<noteq> 0"
 | 
|  |   1121 |       then obtain n where n: "degree p = Suc n " by (cases "degree p", auto)
 | 
|  |   1122 |       {assume "p divides (pexp q (Suc n))"
 | 
|  |   1123 | 	then obtain u where u: "poly (pexp q (Suc n)) = poly (pmult p u)"
 | 
|  |   1124 | 	  unfolding divides_def by blast
 | 
|  |   1125 | 	hence u' :"\<And>x. poly (pexp q (Suc n)) x = poly (pmult p u) x" by simp_all
 | 
|  |   1126 | 	{fix x assume h: "poly p x = 0" "poly q x \<noteq> 0"
 | 
|  |   1127 | 	  hence "poly (pexp q (Suc n)) x \<noteq> 0" by (simp only: poly_exp) simp	  
 | 
|  |   1128 | 	  hence False using u' h(1) by (simp only: poly_mult poly_exp) simp}}
 | 
|  |   1129 | 	with n nullstellensatz_lemma[of p q "degree p"] dp 
 | 
|  |   1130 | 	have ?thesis by auto}
 | 
|  |   1131 |     ultimately have ?thesis by blast}
 | 
|  |   1132 |   ultimately show ?thesis by blast
 | 
|  |   1133 | qed
 | 
|  |   1134 | 
 | 
|  |   1135 | text{* Useful lemma *}
 | 
|  |   1136 | 
 | 
|  |   1137 | lemma (in idom_char_0) constant_degree: "constant (poly p) \<longleftrightarrow> degree p = 0" (is "?lhs = ?rhs")
 | 
|  |   1138 | proof
 | 
|  |   1139 |   assume l: ?lhs
 | 
|  |   1140 |   from l[unfolded constant_def, rule_format, of _ "zero"]
 | 
|  |   1141 |   have th: "poly p = poly [poly p 0]" apply - by (rule ext, simp)
 | 
|  |   1142 |   from degree_unique[OF th] show ?rhs by (simp add: degree_def)
 | 
|  |   1143 | next
 | 
|  |   1144 |   assume r: ?rhs
 | 
|  |   1145 |   from r have "pnormalize p = [] \<or> (\<exists>k. pnormalize p = [k])"
 | 
|  |   1146 |     unfolding degree_def by (cases "pnormalize p", auto)
 | 
|  |   1147 |   then show ?lhs unfolding constant_def poly_normalize[of p, symmetric]
 | 
|  |   1148 |     by (auto simp del: poly_normalize)
 | 
|  |   1149 | qed
 | 
|  |   1150 | 
 | 
|  |   1151 | (* It would be nicer to prove this without using algebraic closure...        *)
 | 
|  |   1152 | 
 | 
|  |   1153 | lemma divides_degree_lemma: assumes dpn: "degree (p::complex list) = n"
 | 
|  |   1154 |   shows "n \<le> degree (p *** q) \<or> poly (p *** q) = poly []"
 | 
|  |   1155 |   using dpn
 | 
|  |   1156 | proof(induct n arbitrary: p q)
 | 
|  |   1157 |   case 0 thus ?case by simp
 | 
|  |   1158 | next
 | 
|  |   1159 |   case (Suc n p q)
 | 
|  |   1160 |   from Suc.prems fundamental_theorem_of_algebra[of p] constant_degree[of p]
 | 
|  |   1161 |   obtain a where a: "poly p a = 0" by auto
 | 
|  |   1162 |   then obtain r where r: "p = pmult [-a, 1] r" unfolding poly_linear_divides
 | 
|  |   1163 |     using Suc.prems by (auto simp add: degree_def)
 | 
|  |   1164 |   {assume h: "poly (pmult r q) = poly []"
 | 
|  |   1165 |     hence "poly (pmult p q) = poly []" using r
 | 
|  |   1166 |       apply - apply (rule ext)  by (auto simp add: poly_entire poly_mult poly_add poly_cmult) hence ?case by blast}
 | 
|  |   1167 |   moreover
 | 
|  |   1168 |   {assume h: "poly (pmult r q) \<noteq> poly []" 
 | 
|  |   1169 |     hence r0: "poly r \<noteq> poly []" and q0: "poly q \<noteq> poly []"
 | 
|  |   1170 |       by (auto simp add: poly_entire)
 | 
|  |   1171 |     have eq: "poly (pmult p q) = poly (pmult [-a, 1] (pmult r q))"
 | 
|  |   1172 |       apply - apply (rule ext)
 | 
|  |   1173 |       by (simp add: r poly_mult poly_add poly_cmult ring_simps)
 | 
|  |   1174 |     from linear_mul_degree[OF h, of "- a"]
 | 
|  |   1175 |     have dqe: "degree (pmult p q) = degree (pmult r q) + 1"
 | 
|  |   1176 |       unfolding degree_unique[OF eq] .
 | 
|  |   1177 |     from linear_mul_degree[OF r0, of "- a", unfolded r[symmetric]] r Suc.prems 
 | 
|  |   1178 |     have dr: "degree r = n" by auto
 | 
|  |   1179 |     from  Suc.hyps[OF dr, of q] have "Suc n \<le> degree (pmult p q)"
 | 
|  |   1180 |       unfolding dqe using h by (auto simp del: poly.simps) 
 | 
|  |   1181 |     hence ?case by blast}
 | 
|  |   1182 |   ultimately show ?case by blast
 | 
|  |   1183 | qed
 | 
|  |   1184 | 
 | 
|  |   1185 | lemma divides_degree: assumes pq: "p divides (q:: complex list)"
 | 
|  |   1186 |   shows "degree p \<le> degree q \<or> poly q = poly []"
 | 
|  |   1187 | using pq  divides_degree_lemma[OF refl, of p]
 | 
|  |   1188 | apply (auto simp add: divides_def poly_entire)
 | 
|  |   1189 | apply atomize
 | 
|  |   1190 | apply (erule_tac x="qa" in allE, auto)
 | 
|  |   1191 | apply (subgoal_tac "degree q = degree (p *** qa)", simp)
 | 
|  |   1192 | apply (rule degree_unique, simp)
 | 
|  |   1193 | done
 | 
|  |   1194 | 
 | 
|  |   1195 | (* Arithmetic operations on multivariate polynomials.                        *)
 | 
|  |   1196 | 
 | 
|  |   1197 | lemma mpoly_base_conv: 
 | 
|  |   1198 |   "(0::complex) \<equiv> poly [] x" "c \<equiv> poly [c] x" "x \<equiv> poly [0,1] x" by simp_all
 | 
|  |   1199 | 
 | 
|  |   1200 | lemma mpoly_norm_conv: 
 | 
|  |   1201 |   "poly [0] (x::complex) \<equiv> poly [] x" "poly [poly [] y] x \<equiv> poly [] x" by simp_all
 | 
|  |   1202 | 
 | 
|  |   1203 | lemma mpoly_sub_conv: 
 | 
|  |   1204 |   "poly p (x::complex) - poly q x \<equiv> poly p x + -1 * poly q x"
 | 
|  |   1205 |   by (simp add: diff_def)
 | 
|  |   1206 | 
 | 
|  |   1207 | lemma poly_pad_rule: "poly p x = 0 ==> poly (0#p) x = (0::complex)" by simp
 | 
|  |   1208 | 
 | 
|  |   1209 | lemma poly_cancel_eq_conv: "p = (0::complex) \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> (q = 0) \<equiv> (a * q - b * p = 0)" apply (atomize (full)) by auto
 | 
|  |   1210 | 
 | 
|  |   1211 | lemma resolve_eq_raw:  "poly [] x \<equiv> 0" "poly [c] x \<equiv> (c::complex)" by auto
 | 
|  |   1212 | lemma  resolve_eq_then: "(P \<Longrightarrow> (Q \<equiv> Q1)) \<Longrightarrow> (\<not>P \<Longrightarrow> (Q \<equiv> Q2))
 | 
|  |   1213 |   \<Longrightarrow> Q \<equiv> P \<and> Q1 \<or> \<not>P\<and> Q2" apply (atomize (full)) by blast 
 | 
|  |   1214 | lemma expand_ex_beta_conv: "list_ex P [c] \<equiv> P c" by simp
 | 
|  |   1215 | 
 | 
|  |   1216 | lemma poly_divides_pad_rule: 
 | 
|  |   1217 |   fixes p q :: "complex list"
 | 
|  |   1218 |   assumes pq: "p divides q"
 | 
|  |   1219 |   shows "p divides ((0::complex)#q)"
 | 
|  |   1220 | proof-
 | 
|  |   1221 |   from pq obtain r where r: "poly q = poly (p *** r)" unfolding divides_def by blast
 | 
|  |   1222 |   hence "poly (0#q) = poly (p *** ([0,1] *** r))" 
 | 
|  |   1223 |     by - (rule ext, simp add: poly_mult poly_cmult poly_add)
 | 
|  |   1224 |   thus ?thesis unfolding divides_def by blast
 | 
|  |   1225 | qed
 | 
|  |   1226 | 
 | 
|  |   1227 | lemma poly_divides_pad_const_rule: 
 | 
|  |   1228 |   fixes p q :: "complex list"
 | 
|  |   1229 |   assumes pq: "p divides q"
 | 
|  |   1230 |   shows "p divides (a %* q)"
 | 
|  |   1231 | proof-
 | 
|  |   1232 |   from pq obtain r where r: "poly q = poly (p *** r)" unfolding divides_def by blast
 | 
|  |   1233 |   hence "poly (a %* q) = poly (p *** (a %* r))" 
 | 
|  |   1234 |     by - (rule ext, simp add: poly_mult poly_cmult poly_add)
 | 
|  |   1235 |   thus ?thesis unfolding divides_def by blast
 | 
|  |   1236 | qed
 | 
|  |   1237 | 
 | 
|  |   1238 | 
 | 
|  |   1239 | lemma poly_divides_conv0:  
 | 
|  |   1240 |   fixes p :: "complex list"
 | 
|  |   1241 |   assumes lgpq: "length q < length p" and lq:"last p \<noteq> 0"
 | 
|  |   1242 |   shows "p divides q \<equiv> (\<not> (list_ex (\<lambda>c. c \<noteq> 0) q))" (is "?lhs \<equiv> ?rhs")
 | 
|  |   1243 | proof-
 | 
|  |   1244 |   {assume r: ?rhs 
 | 
|  |   1245 |     hence eq: "poly q = poly []" unfolding poly_zero 
 | 
|  |   1246 |       by (simp add: list_all_iff list_ex_iff)
 | 
|  |   1247 |     hence "poly q = poly (p *** [])" by - (rule ext, simp add: poly_mult)
 | 
|  |   1248 |     hence ?lhs unfolding divides_def  by blast}
 | 
|  |   1249 |   moreover
 | 
|  |   1250 |   {assume l: ?lhs
 | 
|  |   1251 |     have ath: "\<And>lq lp dq::nat. lq < lp ==> lq \<noteq> 0 \<Longrightarrow> dq <= lq - 1 ==> dq < lp - 1"
 | 
|  |   1252 |       by arith
 | 
|  |   1253 |     {assume q0: "length q = 0"
 | 
|  |   1254 |       hence "q = []" by simp
 | 
|  |   1255 |       hence ?rhs by simp}
 | 
|  |   1256 |     moreover
 | 
|  |   1257 |     {assume lgq0: "length q \<noteq> 0"
 | 
|  |   1258 |       from pnormalize_length[of q] have dql: "degree q \<le> length q - 1" 
 | 
|  |   1259 | 	unfolding degree_def by simp
 | 
|  |   1260 |       from ath[OF lgpq lgq0 dql, unfolded pnormal_degree[OF lq, symmetric]] divides_degree[OF l] have "poly q = poly []" by auto
 | 
|  |   1261 |       hence ?rhs unfolding poly_zero by (simp add: list_all_iff list_ex_iff)}
 | 
|  |   1262 |     ultimately have ?rhs by blast }
 | 
|  |   1263 |   ultimately show "?lhs \<equiv> ?rhs" by - (atomize (full), blast) 
 | 
|  |   1264 | qed
 | 
|  |   1265 | 
 | 
|  |   1266 | lemma poly_divides_conv1: 
 | 
|  |   1267 |   assumes a0: "a\<noteq> (0::complex)" and pp': "(p::complex list) divides p'"
 | 
|  |   1268 |   and qrp': "\<And>x. a * poly q x - poly p' x \<equiv> poly r x"
 | 
|  |   1269 |   shows "p divides q \<equiv> p divides (r::complex list)" (is "?lhs \<equiv> ?rhs")
 | 
|  |   1270 | proof-
 | 
|  |   1271 |   {
 | 
|  |   1272 |   from pp' obtain t where t: "poly p' = poly (p *** t)" 
 | 
|  |   1273 |     unfolding divides_def by blast
 | 
|  |   1274 |   {assume l: ?lhs
 | 
|  |   1275 |     then obtain u where u: "poly q = poly (p *** u)" unfolding divides_def by blast
 | 
|  |   1276 |      have "poly r = poly (p *** ((a %* u) +++ (-- t)))"
 | 
|  |   1277 |        using u qrp' t
 | 
|  |   1278 |        by - (rule ext, 
 | 
|  |   1279 | 	 simp add: poly_add poly_mult poly_cmult poly_minus ring_simps)
 | 
|  |   1280 |      then have ?rhs unfolding divides_def by blast}
 | 
|  |   1281 |   moreover
 | 
|  |   1282 |   {assume r: ?rhs
 | 
|  |   1283 |     then obtain u where u: "poly r = poly (p *** u)" unfolding divides_def by blast
 | 
|  |   1284 |     from u t qrp' a0 have "poly q = poly (p *** ((1/a) %* (u +++ t)))"
 | 
|  |   1285 |       by - (rule ext, atomize (full), simp add: poly_mult poly_add poly_cmult field_simps)
 | 
|  |   1286 |     hence ?lhs  unfolding divides_def by blast}
 | 
|  |   1287 |   ultimately have "?lhs = ?rhs" by blast }
 | 
|  |   1288 | thus "?lhs \<equiv> ?rhs"  by - (atomize(full), blast) 
 | 
|  |   1289 | qed
 | 
|  |   1290 | 
 | 
|  |   1291 | lemma basic_cqe_conv1:
 | 
|  |   1292 |   "(\<exists>x. poly p x = 0 \<and> poly [] x \<noteq> 0) \<equiv> False"
 | 
|  |   1293 |   "(\<exists>x. poly [] x \<noteq> 0) \<equiv> False"
 | 
|  |   1294 |   "(\<exists>x. poly [c] x \<noteq> 0) \<equiv> c\<noteq>0"
 | 
|  |   1295 |   "(\<exists>x. poly [] x = 0) \<equiv> True"
 | 
|  |   1296 |   "(\<exists>x. poly [c] x = 0) \<equiv> c = 0" by simp_all
 | 
|  |   1297 | 
 | 
|  |   1298 | lemma basic_cqe_conv2: 
 | 
|  |   1299 |   assumes l:"last (a#b#p) \<noteq> 0" 
 | 
|  |   1300 |   shows "(\<exists>x. poly (a#b#p) x = (0::complex)) \<equiv> True"
 | 
|  |   1301 | proof-
 | 
|  |   1302 |   {fix h t
 | 
|  |   1303 |     assume h: "h\<noteq>0" "list_all (\<lambda>c. c=(0::complex)) t"  "a#b#p = h#t"
 | 
|  |   1304 |     hence "list_all (\<lambda>c. c= 0) (b#p)" by simp
 | 
|  |   1305 |     moreover have "last (b#p) \<in> set (b#p)" by simp
 | 
|  |   1306 |     ultimately have "last (b#p) = 0" by (simp add: list_all_iff)
 | 
|  |   1307 |     with l have False by simp}
 | 
|  |   1308 |   hence th: "\<not> (\<exists> h t. h\<noteq>0 \<and> list_all (\<lambda>c. c=0) t \<and> a#b#p = h#t)"
 | 
|  |   1309 |     by blast
 | 
|  |   1310 |   from fundamental_theorem_of_algebra_alt[OF th] 
 | 
|  |   1311 |   show "(\<exists>x. poly (a#b#p) x = (0::complex)) \<equiv> True" by auto
 | 
|  |   1312 | qed
 | 
|  |   1313 | 
 | 
|  |   1314 | lemma  basic_cqe_conv_2b: "(\<exists>x. poly p x \<noteq> (0::complex)) \<equiv> (list_ex (\<lambda>c. c \<noteq> 0) p)"
 | 
|  |   1315 | proof-
 | 
|  |   1316 |   have "\<not> (list_ex (\<lambda>c. c \<noteq> 0) p) \<longleftrightarrow> poly p = poly []" 
 | 
|  |   1317 |     by (simp add: poly_zero list_all_iff list_ex_iff)
 | 
|  |   1318 |   also have "\<dots> \<longleftrightarrow> (\<not> (\<exists>x. poly p x \<noteq> 0))" by (auto intro: ext)
 | 
|  |   1319 |   finally show "(\<exists>x. poly p x \<noteq> (0::complex)) \<equiv> (list_ex (\<lambda>c. c \<noteq> 0) p)"
 | 
|  |   1320 |     by - (atomize (full), blast)
 | 
|  |   1321 | qed
 | 
|  |   1322 | 
 | 
|  |   1323 | lemma basic_cqe_conv3:
 | 
|  |   1324 |   fixes p q :: "complex list"
 | 
|  |   1325 |   assumes l: "last (a#p) \<noteq> 0" 
 | 
|  |   1326 |   shows "(\<exists>x. poly (a#p) x =0 \<and> poly q x \<noteq> 0) \<equiv> \<not> ((a#p) divides (q %^ (length p)))"
 | 
|  |   1327 | proof-
 | 
|  |   1328 |   note np = pnormalize_eq[OF l]
 | 
|  |   1329 |   {assume "poly (a#p) = poly []" hence False using l
 | 
|  |   1330 |       unfolding poly_zero apply (auto simp add: list_all_iff del: last.simps)
 | 
|  |   1331 |       apply (cases p, simp_all) done}
 | 
|  |   1332 |   then have p0: "poly (a#p) \<noteq> poly []"  by blast
 | 
|  |   1333 |   from np have dp:"degree (a#p) = length p" by (simp add: degree_def)
 | 
|  |   1334 |   from nullstellensatz_univariate[of "a#p" q] p0 dp
 | 
|  |   1335 |   show "(\<exists>x. poly (a#p) x =0 \<and> poly q x \<noteq> 0) \<equiv> \<not> ((a#p) divides (q %^ (length p)))"
 | 
|  |   1336 |     by - (atomize (full), auto)
 | 
|  |   1337 | qed
 | 
|  |   1338 | 
 | 
|  |   1339 | lemma basic_cqe_conv4:
 | 
|  |   1340 |   fixes p q :: "complex list"
 | 
|  |   1341 |   assumes h: "\<And>x. poly (q %^ n) x \<equiv> poly r x"
 | 
|  |   1342 |   shows "p divides (q %^ n) \<equiv> p divides r"
 | 
|  |   1343 | proof-
 | 
|  |   1344 |   from h have "poly (q %^ n) = poly r" by (auto intro: ext)  
 | 
|  |   1345 |   thus "p divides (q %^ n) \<equiv> p divides r" unfolding divides_def by simp
 | 
|  |   1346 | qed
 | 
|  |   1347 | 
 | 
|  |   1348 | lemma pmult_Cons_Cons: "((a::complex)#b#p) *** q = (a %*q) +++ (0#((b#p) *** q))"
 | 
|  |   1349 |   by simp
 | 
|  |   1350 | 
 | 
|  |   1351 | lemma elim_neg_conv: "- z \<equiv> (-1) * (z::complex)" by simp
 | 
|  |   1352 | lemma eqT_intr: "PROP P \<Longrightarrow> (True \<Longrightarrow> PROP P )" "PROP P \<Longrightarrow> True" by blast+
 | 
|  |   1353 | lemma negate_negate_rule: "Trueprop P \<equiv> \<not> P \<equiv> False" by (atomize (full), auto)
 | 
|  |   1354 | lemma last_simps: "last [x] = x" "last (x#y#ys) = last (y#ys)" by simp_all
 | 
|  |   1355 | lemma length_simps: "length [] = 0" "length (x#y#xs) = length xs + 2" "length [x] = 1" by simp_all
 | 
|  |   1356 | 
 | 
|  |   1357 | lemma complex_entire: "(z::complex) \<noteq> 0 \<and> w \<noteq> 0 \<equiv> z*w \<noteq> 0" by simp
 | 
|  |   1358 | lemma resolve_eq_ne: "(P \<equiv> True) \<equiv> (\<not>P \<equiv> False)" "(P \<equiv> False) \<equiv> (\<not>P \<equiv> True)" 
 | 
|  |   1359 |   by (atomize (full)) simp_all
 | 
|  |   1360 | lemma cqe_conv1: "poly [] x = 0 \<longleftrightarrow> True"  by simp
 | 
|  |   1361 | lemma cqe_conv2: "(p \<Longrightarrow> (q \<equiv> r)) \<equiv> ((p \<and> q) \<equiv> (p \<and> r))"  (is "?l \<equiv> ?r")
 | 
|  |   1362 | proof
 | 
|  |   1363 |   assume "p \<Longrightarrow> q \<equiv> r" thus "p \<and> q \<equiv> p \<and> r" apply - apply (atomize (full)) by blast
 | 
|  |   1364 | next
 | 
|  |   1365 |   assume "p \<and> q \<equiv> p \<and> r" "p"
 | 
|  |   1366 |   thus "q \<equiv> r" apply - apply (atomize (full)) apply blast done
 | 
|  |   1367 | qed
 | 
|  |   1368 | lemma poly_const_conv: "poly [c] (x::complex) = y \<longleftrightarrow> c = y" by simp
 | 
|  |   1369 | 
 | 
|  |   1370 | end |