src/HOL/Library/Fundamental_Theorem_Algebra.thy
changeset 56778 cb0929421ca6
parent 56776 309e1a61ee7c
child 56795 e8cce2bd23e5
equal deleted inserted replaced
56777:9c3f0ae99532 56778:cb0929421ca6
    15        if 0 \<le> Re z then Complex (sqrt(Re z)) 0
    15        if 0 \<le> Re z then Complex (sqrt(Re z)) 0
    16        else Complex 0 (sqrt(- Re z))
    16        else Complex 0 (sqrt(- Re z))
    17      else Complex (sqrt((cmod z + Re z) /2)) ((Im z / abs(Im z)) * sqrt((cmod z - Re z) /2)))"
    17      else Complex (sqrt((cmod z + Re z) /2)) ((Im z / abs(Im z)) * sqrt((cmod z - Re z) /2)))"
    18 
    18 
    19 lemma csqrt[algebra]: "(csqrt z)\<^sup>2 = z"
    19 lemma csqrt[algebra]: "(csqrt z)\<^sup>2 = z"
    20 proof-
    20 proof -
    21   obtain x y where xy: "z = Complex x y" by (cases z)
    21   obtain x y where xy: "z = Complex x y" by (cases z)
    22   {assume y0: "y = 0"
    22   {
    23     {assume x0: "x \<ge> 0"
    23     assume y0: "y = 0"
    24       then have ?thesis using y0 xy real_sqrt_pow2[OF x0]
    24     {
    25         by (simp add: csqrt_def power2_eq_square)}
    25       assume x0: "x \<ge> 0"
       
    26       then have ?thesis
       
    27         using y0 xy real_sqrt_pow2[OF x0]
       
    28         by (simp add: csqrt_def power2_eq_square)
       
    29     }
    26     moreover
    30     moreover
    27     {assume "\<not> x \<ge> 0" hence x0: "- x \<ge> 0" by arith
    31     {
    28       then have ?thesis using y0 xy real_sqrt_pow2[OF x0]
    32       assume "\<not> x \<ge> 0"
    29         by (simp add: csqrt_def power2_eq_square) }
    33       then have x0: "- x \<ge> 0" by arith
    30     ultimately have ?thesis by blast}
    34       then have ?thesis
       
    35         using y0 xy real_sqrt_pow2[OF x0]
       
    36         by (simp add: csqrt_def power2_eq_square)
       
    37     }
       
    38     ultimately have ?thesis by blast
       
    39   }
    31   moreover
    40   moreover
    32   {assume y0: "y\<noteq>0"
    41   {
    33     {fix x y
    42     assume y0: "y \<noteq> 0"
       
    43     {
       
    44       fix x y
    34       let ?z = "Complex x y"
    45       let ?z = "Complex x y"
    35       from abs_Re_le_cmod[of ?z] have tha: "abs x \<le> cmod ?z" by auto
    46       from abs_Re_le_cmod[of ?z] have tha: "abs x \<le> cmod ?z"
    36       hence "cmod ?z - x \<ge> 0" "cmod ?z + x \<ge> 0" by arith+
    47         by auto
    37       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) }
    48       then have "cmod ?z - x \<ge> 0" "cmod ?z + x \<ge> 0"
       
    49         by arith+
       
    50       then have "(sqrt (x * x + y * y) + x) / 2 \<ge> 0" "(sqrt (x * x + y * y) - x) / 2 \<ge> 0"
       
    51         by (simp_all add: power2_eq_square)
       
    52     }
    38     note th = this
    53     note th = this
    39     have sq4: "\<And>x::real. x\<^sup>2 / 4 = (x / 2)\<^sup>2"
    54     have sq4: "\<And>x::real. x\<^sup>2 / 4 = (x / 2)\<^sup>2"
    40       by (simp add: power2_eq_square)
    55       by (simp add: power2_eq_square)
    41     from th[of x y]
    56     from th[of x y]
    42     have sq4': "sqrt (((sqrt (x * x + y * y) + x)\<^sup>2 / 4)) = (sqrt (x * x + y * y) + x) / 2"
    57     have sq4': "sqrt (((sqrt (x * x + y * y) + x)\<^sup>2 / 4)) = (sqrt (x * x + y * y) + x) / 2"
    43       "sqrt (((sqrt (x * x + y * y) - x)\<^sup>2 / 4)) = (sqrt (x * x + y * y) - x) / 2"
    58       "sqrt (((sqrt (x * x + y * y) - x)\<^sup>2 / 4)) = (sqrt (x * x + y * y) - x) / 2"
    44       unfolding sq4 by simp_all
    59       unfolding sq4 by simp_all
    45     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"
    60     then have th1: "sqrt ((sqrt (x * x + y * y) + x) * (sqrt (x * x + y * y) + x) / 4) -
       
    61         sqrt ((sqrt (x * x + y * y) - x) * (sqrt (x * x + y * y) - x) / 4) = x"
    46       unfolding power2_eq_square by simp
    62       unfolding power2_eq_square by simp
    47     have "sqrt 4 = sqrt (2\<^sup>2)" by simp
    63     have "sqrt 4 = sqrt (2\<^sup>2)"
    48     hence sqrt4: "sqrt 4 = 2" by (simp only: real_sqrt_abs)
    64       by simp
    49     have th2: "2 *(y * sqrt ((sqrt (x * x + y * y) - x) * (sqrt (x * x + y * y) + x) / 4)) / \<bar>y\<bar> = y"
    65     then have sqrt4: "sqrt 4 = 2"
       
    66       by (simp only: real_sqrt_abs)
       
    67     have th2: "2 * (y * sqrt ((sqrt (x * x + y * y) - x) * (sqrt (x * x + y * y) + x) / 4)) / \<bar>y\<bar> = y"
    50       using iffD2[OF real_sqrt_pow2_iff sum_power2_ge_zero[of x y]] y0
    68       using iffD2[OF real_sqrt_pow2_iff sum_power2_ge_zero[of x y]] y0
    51       unfolding power2_eq_square
    69       unfolding power2_eq_square
    52       by (simp add: algebra_simps real_sqrt_divide sqrt4)
    70       by (simp add: algebra_simps real_sqrt_divide sqrt4)
    53      from y0 xy have ?thesis  apply (simp add: csqrt_def power2_eq_square)
    71     from y0 xy have ?thesis
    54        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])
    72       apply (simp add: csqrt_def power2_eq_square)
    55       using th1 th2  ..}
    73       apply (simp add: real_sqrt_sum_squares_mult_ge_zero[of x y]
       
    74         real_sqrt_pow2[OF th(1)[of x y], unfolded power2_eq_square]
       
    75         real_sqrt_pow2[OF th(2)[of x y], unfolded power2_eq_square]
       
    76         real_sqrt_mult[symmetric])
       
    77       using th1 th2  ..
       
    78   }
    56   ultimately show ?thesis by blast
    79   ultimately show ?thesis by blast
    57 qed
    80 qed
    58 
    81 
    59 lemma csqrt_Complex: "x \<ge> 0 \<Longrightarrow> csqrt (Complex x 0) = Complex (sqrt x) 0"
    82 lemma csqrt_Complex: "x \<ge> 0 \<Longrightarrow> csqrt (Complex x 0) = Complex (sqrt x) 0"
    60   by (simp add: csqrt_def)
    83   by (simp add: csqrt_def)
    72     using real_sqrt_sum_squares_ge1 [of "x" y]
    95     using real_sqrt_sum_squares_ge1 [of "x" y]
    73           real_sqrt_sum_squares_ge1 [of "-x" y]
    96           real_sqrt_sum_squares_ge1 [of "-x" y]
    74           real_sqrt_sum_squares_eq_cancel [of x y]
    97           real_sqrt_sum_squares_eq_cancel [of x y]
    75     apply (auto simp: csqrt_def intro!: Rings.ordered_ring_class.split_mult_pos_le)
    98     apply (auto simp: csqrt_def intro!: Rings.ordered_ring_class.split_mult_pos_le)
    76     apply (metis add_commute diff_add_cancel le_add_same_cancel1 real_sqrt_sum_squares_ge1)
    99     apply (metis add_commute diff_add_cancel le_add_same_cancel1 real_sqrt_sum_squares_ge1)
    77     by (metis add_commute less_eq_real_def power_minus_Bit0 real_0_less_add_iff real_sqrt_sum_squares_eq_cancel)
   100     apply (metis add_commute less_eq_real_def power_minus_Bit0
       
   101             real_0_less_add_iff real_sqrt_sum_squares_eq_cancel)
       
   102     done
    78 qed
   103 qed
    79 
   104 
    80 lemma Re_csqrt: "0 \<le> Re(csqrt z)"
   105 lemma Re_csqrt: "0 \<le> Re(csqrt z)"
    81   by (metis csqrt_principal le_less)
   106   by (metis csqrt_principal le_less)
    82 
   107 
    83 lemma csqrt_square: "(0 < Re z | Re z = 0 & 0 \<le> Im z) \<Longrightarrow> csqrt (z^2) = z"
   108 lemma csqrt_square: "0 < Re z \<or> Re z = 0 \<and> 0 \<le> Im z \<Longrightarrow> csqrt (z\<^sup>2) = z"
    84   using csqrt [of "z^2"] csqrt_principal [of "z^2"]
   109   using csqrt [of "z\<^sup>2"] csqrt_principal [of "z\<^sup>2"]
    85   by (cases z) (auto simp: power2_eq_iff)
   110   by (cases z) (auto simp: power2_eq_iff)
    86 
   111 
    87 lemma csqrt_eq_0 [simp]: "csqrt z = 0 \<longleftrightarrow> z = 0"
   112 lemma csqrt_eq_0 [simp]: "csqrt z = 0 \<longleftrightarrow> z = 0"
    88   by auto (metis csqrt power_eq_0_iff)
   113   by auto (metis csqrt power_eq_0_iff)
    89 
   114 
    90 lemma csqrt_eq_1 [simp]: "csqrt z = 1 \<longleftrightarrow> z = 1"
   115 lemma csqrt_eq_1 [simp]: "csqrt z = 1 \<longleftrightarrow> z = 1"
    91   by auto (metis csqrt power2_eq_1_iff)
   116   by auto (metis csqrt power2_eq_1_iff)
    92 
   117 
    93 subsection{* More lemmas about module of complex numbers *}
   118 
       
   119 subsection {* More lemmas about module of complex numbers *}
    94 
   120 
    95 lemma complex_of_real_power: "complex_of_real x ^ n = complex_of_real (x^n)"
   121 lemma complex_of_real_power: "complex_of_real x ^ n = complex_of_real (x^n)"
    96   by (rule of_real_power [symmetric])
   122   by (rule of_real_power [symmetric])
    97 
   123 
    98 text{* The triangle inequality for cmod *}
   124 text{* The triangle inequality for cmod *}
    99 lemma complex_mod_triangle_sub: "cmod w \<le> cmod (w + z) + norm z"
   125 lemma complex_mod_triangle_sub: "cmod w \<le> cmod (w + z) + norm z"
   100   using complex_mod_triangle_ineq2[of "w + z" "-z"] by auto
   126   using complex_mod_triangle_ineq2[of "w + z" "-z"] by auto
   101 
   127 
   102 subsection{* Basic lemmas about polynomials *}
   128 
       
   129 subsection {* Basic lemmas about polynomials *}
   103 
   130 
   104 lemma poly_bound_exists:
   131 lemma poly_bound_exists:
   105   fixes p:: "('a::{comm_semiring_0,real_normed_div_algebra}) poly" 
   132   fixes p :: "'a::{comm_semiring_0,real_normed_div_algebra} poly"
   106   shows "\<exists>m. m > 0 \<and> (\<forall>z. norm z <= r \<longrightarrow> norm (poly p z) \<le> m)"
   133   shows "\<exists>m. m > 0 \<and> (\<forall>z. norm z \<le> r \<longrightarrow> norm (poly p z) \<le> m)"
   107 proof(induct p)
   134 proof (induct p)
   108   case 0 thus ?case by (rule exI[where x=1], simp)
   135   case 0
       
   136   then show ?case by (rule exI[where x=1]) simp
   109 next
   137 next
   110   case (pCons c cs)
   138   case (pCons c cs)
   111   from pCons.hyps obtain m where m: "\<forall>z. norm z \<le> r \<longrightarrow> norm (poly cs z) \<le> m"
   139   from pCons.hyps obtain m where m: "\<forall>z. norm z \<le> r \<longrightarrow> norm (poly cs z) \<le> m"
   112     by blast
   140     by blast
   113   let ?k = " 1 + norm c + \<bar>r * m\<bar>"
   141   let ?k = " 1 + norm c + \<bar>r * m\<bar>"
   114   have kp: "?k > 0" using abs_ge_zero[of "r*m"] norm_ge_zero[of c] by arith
   142   have kp: "?k > 0" using abs_ge_zero[of "r*m"] norm_ge_zero[of c] by arith
   115   {fix z :: 'a 
   143   {
       
   144     fix z :: 'a
   116     assume H: "norm z \<le> r"
   145     assume H: "norm z \<le> r"
   117     from m H have th: "norm (poly cs z) \<le> m" by blast
   146     from m H have th: "norm (poly cs z) \<le> m"
   118     from H have rp: "r \<ge> 0" using norm_ge_zero[of z] by arith
   147       by blast
       
   148     from H have rp: "r \<ge> 0" using norm_ge_zero[of z]
       
   149       by arith
   119     have "norm (poly (pCons c cs) z) \<le> norm c + norm (z* poly cs z)"
   150     have "norm (poly (pCons c cs) z) \<le> norm c + norm (z* poly cs z)"
   120       using norm_triangle_ineq[of c "z* poly cs z"] by simp
   151       using norm_triangle_ineq[of c "z* poly cs z"] by simp
   121     also have "\<dots> \<le> norm c + r*m" using mult_mono[OF H th rp norm_ge_zero[of "poly cs z"]]
   152     also have "\<dots> \<le> norm c + r * m"
       
   153       using mult_mono[OF H th rp norm_ge_zero[of "poly cs z"]]
   122       by (simp add: norm_mult)
   154       by (simp add: norm_mult)
   123     also have "\<dots> \<le> ?k" by simp
   155     also have "\<dots> \<le> ?k"
   124     finally have "norm (poly (pCons c cs) z) \<le> ?k" .}
   156       by simp
       
   157     finally have "norm (poly (pCons c cs) z) \<le> ?k" .
       
   158   }
   125   with kp show ?case by blast
   159   with kp show ?case by blast
   126 qed
   160 qed
   127 
   161 
   128 
   162 
   129 text{* Offsetting the variable in a polynomial gives another of same degree *}
   163 text{* Offsetting the variable in a polynomial gives another of same degree *}
   130 
   164 
   131 definition offset_poly :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a poly"
   165 definition offset_poly :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a poly"
   132 where
   166   where "offset_poly p h = fold_coeffs (\<lambda>a q. smult h q + pCons a q) p 0"
   133   "offset_poly p h = fold_coeffs (\<lambda>a q. smult h q + pCons a q) p 0"
       
   134 
   167 
   135 lemma offset_poly_0: "offset_poly 0 h = 0"
   168 lemma offset_poly_0: "offset_poly 0 h = 0"
   136   by (simp add: offset_poly_def)
   169   by (simp add: offset_poly_def)
   137 
   170 
   138 lemma offset_poly_pCons:
   171 lemma offset_poly_pCons:
   139   "offset_poly (pCons a p) h =
   172   "offset_poly (pCons a p) h =
   140     smult h (offset_poly p h) + pCons a (offset_poly p h)"
   173     smult h (offset_poly p h) + pCons a (offset_poly p h)"
   141   by (cases "p = 0 \<and> a = 0") (auto simp add: offset_poly_def)
   174   by (cases "p = 0 \<and> a = 0") (auto simp add: offset_poly_def)
   142 
   175 
   143 lemma offset_poly_single: "offset_poly [:a:] h = [:a:]"
   176 lemma offset_poly_single: "offset_poly [:a:] h = [:a:]"
   144 by (simp add: offset_poly_pCons offset_poly_0)
   177   by (simp add: offset_poly_pCons offset_poly_0)
   145 
   178 
   146 lemma poly_offset_poly: "poly (offset_poly p h) x = poly p (h + x)"
   179 lemma poly_offset_poly: "poly (offset_poly p h) x = poly p (h + x)"
   147 apply (induct p)
   180   apply (induct p)
   148 apply (simp add: offset_poly_0)
   181   apply (simp add: offset_poly_0)
   149 apply (simp add: offset_poly_pCons algebra_simps)
   182   apply (simp add: offset_poly_pCons algebra_simps)
   150 done
   183   done
   151 
   184 
   152 lemma offset_poly_eq_0_lemma: "smult c p + pCons a p = 0 \<Longrightarrow> p = 0"
   185 lemma offset_poly_eq_0_lemma: "smult c p + pCons a p = 0 \<Longrightarrow> p = 0"
   153 by (induct p arbitrary: a, simp, force)
   186   by (induct p arbitrary: a) (simp, force)
   154 
   187 
   155 lemma offset_poly_eq_0_iff: "offset_poly p h = 0 \<longleftrightarrow> p = 0"
   188 lemma offset_poly_eq_0_iff: "offset_poly p h = 0 \<longleftrightarrow> p = 0"
   156 apply (safe intro!: offset_poly_0)
   189   apply (safe intro!: offset_poly_0)
   157 apply (induct p, simp)
   190   apply (induct p, simp)
   158 apply (simp add: offset_poly_pCons)
   191   apply (simp add: offset_poly_pCons)
   159 apply (frule offset_poly_eq_0_lemma, simp)
   192   apply (frule offset_poly_eq_0_lemma, simp)
   160 done
   193   done
   161 
   194 
   162 lemma degree_offset_poly: "degree (offset_poly p h) = degree p"
   195 lemma degree_offset_poly: "degree (offset_poly p h) = degree p"
   163 apply (induct p)
   196   apply (induct p)
   164 apply (simp add: offset_poly_0)
   197   apply (simp add: offset_poly_0)
   165 apply (case_tac "p = 0")
   198   apply (case_tac "p = 0")
   166 apply (simp add: offset_poly_0 offset_poly_pCons)
   199   apply (simp add: offset_poly_0 offset_poly_pCons)
   167 apply (simp add: offset_poly_pCons)
   200   apply (simp add: offset_poly_pCons)
   168 apply (subst degree_add_eq_right)
   201   apply (subst degree_add_eq_right)
   169 apply (rule le_less_trans [OF degree_smult_le])
   202   apply (rule le_less_trans [OF degree_smult_le])
   170 apply (simp add: offset_poly_eq_0_iff)
   203   apply (simp add: offset_poly_eq_0_iff)
   171 apply (simp add: offset_poly_eq_0_iff)
   204   apply (simp add: offset_poly_eq_0_iff)
   172 done
   205   done
   173 
   206 
   174 definition
   207 definition "psize p = (if p = 0 then 0 else Suc (degree p))"
   175   "psize p = (if p = 0 then 0 else Suc (degree p))"
       
   176 
   208 
   177 lemma psize_eq_0_iff [simp]: "psize p = 0 \<longleftrightarrow> p = 0"
   209 lemma psize_eq_0_iff [simp]: "psize p = 0 \<longleftrightarrow> p = 0"
   178   unfolding psize_def by simp
   210   unfolding psize_def by simp
   179 
   211 
   180 lemma poly_offset: 
   212 lemma poly_offset:
   181   fixes p:: "('a::comm_ring_1) poly" 
   213   fixes p :: "'a::comm_ring_1 poly"
   182   shows "\<exists> q. psize q = psize p \<and> (\<forall>x. poly q x = poly p (a + x))"
   214   shows "\<exists>q. psize q = psize p \<and> (\<forall>x. poly q x = poly p (a + x))"
   183 proof (intro exI conjI)
   215 proof (intro exI conjI)
   184   show "psize (offset_poly p a) = psize p"
   216   show "psize (offset_poly p a) = psize p"
   185     unfolding psize_def
   217     unfolding psize_def
   186     by (simp add: offset_poly_eq_0_iff degree_offset_poly)
   218     by (simp add: offset_poly_eq_0_iff degree_offset_poly)
   187   show "\<forall>x. poly (offset_poly p a) x = poly p (a + x)"
   219   show "\<forall>x. poly (offset_poly p a) x = poly p (a + x)"
   188     by (simp add: poly_offset_poly)
   220     by (simp add: poly_offset_poly)
   189 qed
   221 qed
   190 
   222 
   191 text{* An alternative useful formulation of completeness of the reals *}
   223 text{* An alternative useful formulation of completeness of the reals *}
   192 lemma real_sup_exists: assumes ex: "\<exists>x. P x" and bz: "\<exists>z. \<forall>x. P x \<longrightarrow> x < z"
   224 lemma real_sup_exists:
   193   shows "\<exists>(s::real). \<forall>y. (\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < s"
   225   assumes ex: "\<exists>x. P x"
       
   226     and bz: "\<exists>z. \<forall>x. P x \<longrightarrow> x < z"
       
   227   shows "\<exists>s::real. \<forall>y. (\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < s"
   194 proof
   228 proof
   195   from bz have "bdd_above (Collect P)"
   229   from bz have "bdd_above (Collect P)"
   196     by (force intro: less_imp_le)
   230     by (force intro: less_imp_le)
   197   then show "\<forall>y. (\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < Sup (Collect P)"
   231   then show "\<forall>y. (\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < Sup (Collect P)"
   198     using ex bz by (subst less_cSup_iff) auto
   232     using ex bz by (subst less_cSup_iff) auto
   200 
   234 
   201 subsection {* Fundamental theorem of algebra *}
   235 subsection {* Fundamental theorem of algebra *}
   202 lemma  unimodular_reduce_norm:
   236 lemma  unimodular_reduce_norm:
   203   assumes md: "cmod z = 1"
   237   assumes md: "cmod z = 1"
   204   shows "cmod (z + 1) < 1 \<or> cmod (z - 1) < 1 \<or> cmod (z + ii) < 1 \<or> cmod (z - ii) < 1"
   238   shows "cmod (z + 1) < 1 \<or> cmod (z - 1) < 1 \<or> cmod (z + ii) < 1 \<or> cmod (z - ii) < 1"
   205 proof-
   239 proof -
   206   obtain x y where z: "z = Complex x y " by (cases z, auto)
   240   obtain x y where z: "z = Complex x y "
   207   from md z have xy: "x\<^sup>2 + y\<^sup>2 = 1" by (simp add: cmod_def)
   241     by (cases z) auto
   208   {assume C: "cmod (z + 1) \<ge> 1" "cmod (z - 1) \<ge> 1" "cmod (z + ii) \<ge> 1" "cmod (z - ii) \<ge> 1"
   242   from md z have xy: "x\<^sup>2 + y\<^sup>2 = 1"
   209     from C z xy have "2*x \<le> 1" "2*x \<ge> -1" "2*y \<le> 1" "2*y \<ge> -1"
   243     by (simp add: cmod_def)
       
   244   {
       
   245     assume C: "cmod (z + 1) \<ge> 1" "cmod (z - 1) \<ge> 1" "cmod (z + ii) \<ge> 1" "cmod (z - ii) \<ge> 1"
       
   246     from C z xy have "2 * x \<le> 1" "2 * x \<ge> -1" "2 * y \<le> 1" "2 * y \<ge> -1"
   210       by (simp_all add: cmod_def power2_eq_square algebra_simps)
   247       by (simp_all add: cmod_def power2_eq_square algebra_simps)
   211     hence "abs (2*x) \<le> 1" "abs (2*y) \<le> 1" by simp_all
   248     then have "abs (2 * x) \<le> 1" "abs (2 * y) \<le> 1"
   212     hence "(abs (2 * x))\<^sup>2 <= 1\<^sup>2" "(abs (2 * y))\<^sup>2 <= 1\<^sup>2"
   249       by simp_all
       
   250     then have "(abs (2 * x))\<^sup>2 \<le> 1\<^sup>2" "(abs (2 * y))\<^sup>2 \<le> 1\<^sup>2"
   213       by - (rule power_mono, simp, simp)+
   251       by - (rule power_mono, simp, simp)+
   214     hence th0: "4*x\<^sup>2 \<le> 1" "4*y\<^sup>2 \<le> 1"
   252     then have th0: "4 * x\<^sup>2 \<le> 1" "4 * y\<^sup>2 \<le> 1"
   215       by (simp_all add: power_mult_distrib)
   253       by (simp_all add: power_mult_distrib)
   216     from add_mono[OF th0] xy have False by simp }
   254     from add_mono[OF th0] xy have False by simp
   217   thus ?thesis unfolding linorder_not_le[symmetric] by blast
   255   }
       
   256   then show ?thesis
       
   257     unfolding linorder_not_le[symmetric] by blast
   218 qed
   258 qed
   219 
   259 
   220 text{* Hence we can always reduce modulus of @{text "1 + b z^n"} if nonzero *}
   260 text{* Hence we can always reduce modulus of @{text "1 + b z^n"} if nonzero *}
   221 lemma reduce_poly_simple:
   261 lemma reduce_poly_simple:
   222  assumes b: "b \<noteq> 0" and n: "n\<noteq>0"
   262   assumes b: "b \<noteq> 0"
       
   263     and n: "n \<noteq> 0"
   223   shows "\<exists>z. cmod (1 + b * z^n) < 1"
   264   shows "\<exists>z. cmod (1 + b * z^n) < 1"
   224 using n
   265   using n
   225 proof(induct n rule: nat_less_induct)
   266 proof (induct n rule: nat_less_induct)
   226   fix n
   267   fix n
   227   assume IH: "\<forall>m<n. m \<noteq> 0 \<longrightarrow> (\<exists>z. cmod (1 + b * z ^ m) < 1)" and n: "n \<noteq> 0"
   268   assume IH: "\<forall>m<n. m \<noteq> 0 \<longrightarrow> (\<exists>z. cmod (1 + b * z ^ m) < 1)"
       
   269   assume n: "n \<noteq> 0"
   228   let ?P = "\<lambda>z n. cmod (1 + b * z ^ n) < 1"
   270   let ?P = "\<lambda>z n. cmod (1 + b * z ^ n) < 1"
   229   {assume e: "even n"
   271   {
   230     hence "\<exists>m. n = 2*m" by presburger
   272     assume e: "even n"
   231     then obtain m where m: "n = 2*m" by blast
   273     then have "\<exists>m. n = 2 * m"
   232     from n m have "m\<noteq>0" "m < n" by presburger+
   274       by presburger
   233     with IH[rule_format, of m] obtain z where z: "?P z m" by blast
   275     then obtain m where m: "n = 2 * m"
       
   276       by blast
       
   277     from n m have "m \<noteq> 0" "m < n"
       
   278       by presburger+
       
   279     with IH[rule_format, of m] obtain z where z: "?P z m"
       
   280       by blast
   234     from z have "?P (csqrt z) n" by (simp add: m power_mult csqrt)
   281     from z have "?P (csqrt z) n" by (simp add: m power_mult csqrt)
   235     hence "\<exists>z. ?P z n" ..}
   282     then have "\<exists>z. ?P z n" ..
       
   283   }
   236   moreover
   284   moreover
   237   {assume o: "odd n"
   285   {
       
   286     assume o: "odd n"
   238     have th0: "cmod (complex_of_real (cmod b) / b) = 1"
   287     have th0: "cmod (complex_of_real (cmod b) / b) = 1"
   239       using b by (simp add: norm_divide)
   288       using b by (simp add: norm_divide)
   240     from o have "\<exists>m. n = Suc (2*m)" by presburger+
   289     from o have "\<exists>m. n = Suc (2 * m)"
   241     then obtain m where m: "n = Suc (2*m)" by blast
   290       by presburger+
       
   291     then obtain m where m: "n = Suc (2*m)"
       
   292       by blast
   242     from unimodular_reduce_norm[OF th0] o
   293     from unimodular_reduce_norm[OF th0] o
   243     have "\<exists>v. cmod (complex_of_real (cmod b) / b + v^n) < 1"
   294     have "\<exists>v. cmod (complex_of_real (cmod b) / b + v^n) < 1"
   244       apply (cases "cmod (complex_of_real (cmod b) / b + 1) < 1", rule_tac x="1" in exI, simp)
   295       apply (cases "cmod (complex_of_real (cmod b) / b + 1) < 1", rule_tac x="1" in exI, simp)
   245       apply (cases "cmod (complex_of_real (cmod b) / b - 1) < 1", rule_tac x="-1" in exI, simp)
   296       apply (cases "cmod (complex_of_real (cmod b) / b - 1) < 1", rule_tac x="-1" in exI, simp)
   246       apply (cases "cmod (complex_of_real (cmod b) / b + ii) < 1")
   297       apply (cases "cmod (complex_of_real (cmod b) / b + ii) < 1")
   249       apply (cases "even m", rule_tac x="- ii" in exI, simp add: m power_mult)
   300       apply (cases "even m", rule_tac x="- ii" in exI, simp add: m power_mult)
   250       apply (auto simp add: m power_mult)
   301       apply (auto simp add: m power_mult)
   251       apply (rule_tac x="ii" in exI)
   302       apply (rule_tac x="ii" in exI)
   252       apply (auto simp add: m power_mult)
   303       apply (auto simp add: m power_mult)
   253       done
   304       done
   254     then obtain v where v: "cmod (complex_of_real (cmod b) / b + v^n) < 1" by blast
   305     then obtain v where v: "cmod (complex_of_real (cmod b) / b + v^n) < 1"
       
   306       by blast
   255     let ?w = "v / complex_of_real (root n (cmod b))"
   307     let ?w = "v / complex_of_real (root n (cmod b))"
   256     from odd_real_root_pow[OF o, of "cmod b"]
   308     from odd_real_root_pow[OF o, of "cmod b"]
   257     have th1: "?w ^ n = v^n / complex_of_real (cmod b)"
   309     have th1: "?w ^ n = v^n / complex_of_real (cmod b)"
   258       by (simp add: power_divide complex_of_real_power)
   310       by (simp add: power_divide complex_of_real_power)
   259     have th2:"cmod (complex_of_real (cmod b) / b) = 1" using b by (simp add: norm_divide)
   311     have th2:"cmod (complex_of_real (cmod b) / b) = 1"
   260     hence th3: "cmod (complex_of_real (cmod b) / b) \<ge> 0" by simp
   312       using b by (simp add: norm_divide)
       
   313     then have th3: "cmod (complex_of_real (cmod b) / b) \<ge> 0"
       
   314       by simp
   261     have th4: "cmod (complex_of_real (cmod b) / b) *
   315     have th4: "cmod (complex_of_real (cmod b) / b) *
   262    cmod (1 + b * (v ^ n / complex_of_real (cmod b)))
   316         cmod (1 + b * (v ^ n / complex_of_real (cmod b))) <
   263    < cmod (complex_of_real (cmod b) / b) * 1"
   317         cmod (complex_of_real (cmod b) / b) * 1"
   264       apply (simp only: norm_mult[symmetric] distrib_left)
   318       apply (simp only: norm_mult[symmetric] distrib_left)
   265       using b v by (simp add: th2)
   319       using b v
   266 
   320       apply (simp add: th2)
       
   321       done
   267     from mult_less_imp_less_left[OF th4 th3]
   322     from mult_less_imp_less_left[OF th4 th3]
   268     have "?P ?w n" unfolding th1 .
   323     have "?P ?w n" unfolding th1 .
   269     hence "\<exists>z. ?P z n" .. }
   324     then have "\<exists>z. ?P z n" ..
       
   325   }
   270   ultimately show "\<exists>z. ?P z n" by blast
   326   ultimately show "\<exists>z. ?P z n" by blast
   271 qed
   327 qed
   272 
   328 
   273 text{* Bolzano-Weierstrass type property for closed disc in complex plane. *}
   329 text{* Bolzano-Weierstrass type property for closed disc in complex plane. *}
   274 
   330 
   275 lemma metric_bound_lemma: "cmod (x - y) <= \<bar>Re x - Re y\<bar> + \<bar>Im x - Im y\<bar>"
   331 lemma metric_bound_lemma: "cmod (x - y) \<le> \<bar>Re x - Re y\<bar> + \<bar>Im x - Im y\<bar>"
   276   using real_sqrt_sum_squares_triangle_ineq[of "Re x - Re y" 0 0 "Im x - Im y" ]
   332   using real_sqrt_sum_squares_triangle_ineq[of "Re x - Re y" 0 0 "Im x - Im y" ]
   277   unfolding cmod_def by simp
   333   unfolding cmod_def by simp
   278 
   334 
   279 lemma bolzano_weierstrass_complex_disc:
   335 lemma bolzano_weierstrass_complex_disc:
   280   assumes r: "\<forall>n. cmod (s n) \<le> r"
   336   assumes r: "\<forall>n. cmod (s n) \<le> r"
   281   shows "\<exists>f z. subseq f \<and> (\<forall>e >0. \<exists>N. \<forall>n \<ge> N. cmod (s (f n) - z) < e)"
   337   shows "\<exists>f z. subseq f \<and> (\<forall>e >0. \<exists>N. \<forall>n \<ge> N. cmod (s (f n) - z) < e)"
   282 proof-
   338 proof-
   283   from seq_monosub[of "Re o s"]
   339   from seq_monosub[of "Re \<circ> s"]
   284   obtain f where f: "subseq f" "monoseq (\<lambda>n. Re (s (f n)))"
   340   obtain f where f: "subseq f" "monoseq (\<lambda>n. Re (s (f n)))"
   285     unfolding o_def by blast
   341     unfolding o_def by blast
   286   from seq_monosub[of "Im o s o f"]
   342   from seq_monosub[of "Im \<circ> s \<circ> f"]
   287   obtain g where g: "subseq g" "monoseq (\<lambda>n. Im (s(f(g n))))" unfolding o_def by blast
   343   obtain g where g: "subseq g" "monoseq (\<lambda>n. Im (s (f (g n))))"
   288   let ?h = "f o g"
   344     unfolding o_def by blast
   289   from r[rule_format, of 0] have rp: "r \<ge> 0" using norm_ge_zero[of "s 0"] by arith
   345   let ?h = "f \<circ> g"
   290   have th:"\<forall>n. r + 1 \<ge> \<bar> Re (s n)\<bar>"
   346   from r[rule_format, of 0] have rp: "r \<ge> 0"
       
   347     using norm_ge_zero[of "s 0"] by arith
       
   348   have th: "\<forall>n. r + 1 \<ge> \<bar>Re (s n)\<bar>"
   291   proof
   349   proof
   292     fix n
   350     fix n
   293     from abs_Re_le_cmod[of "s n"] r[rule_format, of n]  show "\<bar>Re (s n)\<bar> \<le> r + 1" by arith
   351     from abs_Re_le_cmod[of "s n"] r[rule_format, of n]
       
   352     show "\<bar>Re (s n)\<bar> \<le> r + 1" by arith
   294   qed
   353   qed
   295   have conv1: "convergent (\<lambda>n. Re (s ( f n)))"
   354   have conv1: "convergent (\<lambda>n. Re (s (f n)))"
   296     apply (rule Bseq_monoseq_convergent)
   355     apply (rule Bseq_monoseq_convergent)
   297     apply (simp add: Bseq_def)
   356     apply (simp add: Bseq_def)
   298     apply (metis gt_ex le_less_linear less_trans order.trans th)
   357     apply (metis gt_ex le_less_linear less_trans order.trans th)
   299     using f(2) .
   358     apply (rule f(2))
   300   have th:"\<forall>n. r + 1 \<ge> \<bar> Im (s n)\<bar>"
   359     done
       
   360   have th: "\<forall>n. r + 1 \<ge> \<bar>Im (s n)\<bar>"
   301   proof
   361   proof
   302     fix n
   362     fix n
   303     from abs_Im_le_cmod[of "s n"] r[rule_format, of n]  show "\<bar>Im (s n)\<bar> \<le> r + 1" by arith
   363     from abs_Im_le_cmod[of "s n"] r[rule_format, of n]
       
   364     show "\<bar>Im (s n)\<bar> \<le> r + 1"
       
   365       by arith
   304   qed
   366   qed
   305 
   367 
   306   have conv2: "convergent (\<lambda>n. Im (s (f (g n))))"
   368   have conv2: "convergent (\<lambda>n. Im (s (f (g n))))"
   307     apply (rule Bseq_monoseq_convergent)
   369     apply (rule Bseq_monoseq_convergent)
   308     apply (simp add: Bseq_def)
   370     apply (simp add: Bseq_def)
   309     apply (metis gt_ex le_less_linear less_trans order.trans th)
   371     apply (metis gt_ex le_less_linear less_trans order.trans th)
   310     using g(2) .
   372     apply (rule g(2))
       
   373     done
   311 
   374 
   312   from conv1[unfolded convergent_def] obtain x where "LIMSEQ (\<lambda>n. Re (s (f n))) x"
   375   from conv1[unfolded convergent_def] obtain x where "LIMSEQ (\<lambda>n. Re (s (f n))) x"
   313     by blast
   376     by blast
   314   hence  x: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar> Re (s (f n)) - x \<bar> < r"
   377   then have x: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar> Re (s (f n)) - x \<bar> < r"
   315     unfolding LIMSEQ_iff real_norm_def .
   378     unfolding LIMSEQ_iff real_norm_def .
   316 
   379 
   317   from conv2[unfolded convergent_def] obtain y where "LIMSEQ (\<lambda>n. Im (s (f (g n)))) y"
   380   from conv2[unfolded convergent_def] obtain y where "LIMSEQ (\<lambda>n. Im (s (f (g n)))) y"
   318     by blast
   381     by blast
   319   hence  y: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar> Im (s (f (g n))) - y \<bar> < r"
   382   then have y: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar> Im (s (f (g n))) - y \<bar> < r"
   320     unfolding LIMSEQ_iff real_norm_def .
   383     unfolding LIMSEQ_iff real_norm_def .
   321   let ?w = "Complex x y"
   384   let ?w = "Complex x y"
   322   from f(1) g(1) have hs: "subseq ?h" unfolding subseq_def by auto
   385   from f(1) g(1) have hs: "subseq ?h"
   323   {fix e assume ep: "e > (0::real)"
   386     unfolding subseq_def by auto
   324     hence e2: "e/2 > 0" by simp
   387   {
       
   388     fix e :: real
       
   389     assume ep: "e > 0"
       
   390     then have e2: "e/2 > 0" by simp
   325     from x[rule_format, OF e2] y[rule_format, OF e2]
   391     from x[rule_format, OF e2] y[rule_format, OF e2]
   326     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
   392     obtain N1 N2 where N1: "\<forall>n\<ge>N1. \<bar>Re (s (f n)) - x\<bar> < e / 2"
   327     {fix n assume nN12: "n \<ge> N1 + N2"
   393       and N2: "\<forall>n\<ge>N2. \<bar>Im (s (f (g n))) - y\<bar> < e / 2" by blast
   328       hence nN1: "g n \<ge> N1" and nN2: "n \<ge> N2" using seq_suble[OF g(1), of n] by arith+
   394     {
       
   395       fix n
       
   396       assume nN12: "n \<ge> N1 + N2"
       
   397       then have nN1: "g n \<ge> N1" and nN2: "n \<ge> N2"
       
   398         using seq_suble[OF g(1), of n] by arith+
   329       from add_strict_mono[OF N1[rule_format, OF nN1] N2[rule_format, OF nN2]]
   399       from add_strict_mono[OF N1[rule_format, OF nN1] N2[rule_format, OF nN2]]
   330       have "cmod (s (?h n) - ?w) < e"
   400       have "cmod (s (?h n) - ?w) < e"
   331         using metric_bound_lemma[of "s (f (g n))" ?w] by simp }
   401         using metric_bound_lemma[of "s (f (g n))" ?w] by simp
   332     hence "\<exists>N. \<forall>n\<ge>N. cmod (s (?h n) - ?w) < e" by blast }
   402     }
   333   with hs show ?thesis  by blast
   403     then have "\<exists>N. \<forall>n\<ge>N. cmod (s (?h n) - ?w) < e" by blast
       
   404   }
       
   405   with hs show ?thesis by blast
   334 qed
   406 qed
   335 
   407 
   336 text{* Polynomial is continuous. *}
   408 text{* Polynomial is continuous. *}
   337 
   409 
   338 lemma poly_cont:
   410 lemma poly_cont:
   339   fixes p:: "('a::{comm_semiring_0,real_normed_div_algebra}) poly" 
   411   fixes p :: "'a::{comm_semiring_0,real_normed_div_algebra} poly"
   340   assumes ep: "e > 0"
   412   assumes ep: "e > 0"
   341   shows "\<exists>d >0. \<forall>w. 0 < norm (w - z) \<and> norm (w - z) < d \<longrightarrow> norm (poly p w - poly p z) < e"
   413   shows "\<exists>d >0. \<forall>w. 0 < norm (w - z) \<and> norm (w - z) < d \<longrightarrow> norm (poly p w - poly p z) < e"
   342 proof-
   414 proof -
   343   obtain q where q: "degree q = degree p" "\<And>x. poly q x = poly p (z + x)"
   415   obtain q where q: "degree q = degree p" "\<And>x. poly q x = poly p (z + x)"
   344   proof
   416   proof
   345     show "degree (offset_poly p z) = degree p"
   417     show "degree (offset_poly p z) = degree p"
   346       by (rule degree_offset_poly)
   418       by (rule degree_offset_poly)
   347     show "\<And>x. poly (offset_poly p z) x = poly p (z + x)"
   419     show "\<And>x. poly (offset_poly p z) x = poly p (z + x)"
   348       by (rule poly_offset_poly)
   420       by (rule poly_offset_poly)
   349   qed
   421   qed
   350   {fix w
   422   have th: "\<And>w. poly q (w - z) = poly p w"
   351     note q(2)[of "w - z", simplified]}
   423     using q(2)[of "w - z" for w] by simp
   352   note th = this
       
   353   show ?thesis unfolding th[symmetric]
   424   show ?thesis unfolding th[symmetric]
   354   proof(induct q)
   425   proof (induct q)
   355     case 0 thus ?case  using ep by auto
   426     case 0
       
   427     then show ?case
       
   428       using ep by auto
   356   next
   429   next
   357     case (pCons c cs)
   430     case (pCons c cs)
   358     from poly_bound_exists[of 1 "cs"]
   431     from poly_bound_exists[of 1 "cs"]
   359     obtain m where m: "m > 0" "\<And>z. norm z \<le> 1 \<Longrightarrow> norm (poly cs z) \<le> m" by blast
   432     obtain m where m: "m > 0" "\<And>z. norm z \<le> 1 \<Longrightarrow> norm (poly cs z) \<le> m"
   360     from ep m(1) have em0: "e/m > 0" by (simp add: field_simps)
   433       by blast
   361     have one0: "1 > (0::real)"  by arith
   434     from ep m(1) have em0: "e/m > 0"
       
   435       by (simp add: field_simps)
       
   436     have one0: "1 > (0::real)"
       
   437       by arith
   362     from real_lbound_gt_zero[OF one0 em0]
   438     from real_lbound_gt_zero[OF one0 em0]
   363     obtain d where d: "d >0" "d < 1" "d < e / m" by blast
   439     obtain d where d: "d > 0" "d < 1" "d < e / m"
   364     from d(1,3) m(1) have dm: "d*m > 0" "d*m < e"
   440       by blast
       
   441     from d(1,3) m(1) have dm: "d * m > 0" "d * m < e"
   365       by (simp_all add: field_simps)
   442       by (simp_all add: field_simps)
   366     show ?case
   443     show ?case
   367       proof(rule ex_forward[OF real_lbound_gt_zero[OF one0 em0]], clarsimp simp add: norm_mult)
   444     proof (rule ex_forward[OF real_lbound_gt_zero[OF one0 em0]], clarsimp simp add: norm_mult)
   368         fix d w
   445       fix d w
   369         assume H: "d > 0" "d < 1" "d < e/m" "w\<noteq>z" "norm (w-z) < d"
   446       assume H: "d > 0" "d < 1" "d < e/m" "w \<noteq> z" "norm (w - z) < d"
   370         hence d1: "norm (w-z) \<le> 1" "d \<ge> 0" by simp_all
   447       then have d1: "norm (w-z) \<le> 1" "d \<ge> 0"
   371         from H(3) m(1) have dme: "d*m < e" by (simp add: field_simps)
   448         by simp_all
   372         from H have th: "norm (w-z) \<le> d" by simp
   449       from H(3) m(1) have dme: "d*m < e"
   373         from mult_mono[OF th m(2)[OF d1(1)] d1(2) norm_ge_zero] dme
   450         by (simp add: field_simps)
   374         show "norm (w - z) * norm (poly cs (w - z)) < e" by simp
   451       from H have th: "norm (w - z) \<le> d"
   375       qed
   452         by simp
       
   453       from mult_mono[OF th m(2)[OF d1(1)] d1(2) norm_ge_zero] dme
       
   454       show "norm (w - z) * norm (poly cs (w - z)) < e"
       
   455         by simp
   376     qed
   456     qed
       
   457   qed
   377 qed
   458 qed
   378 
   459 
   379 text{* Hence a polynomial attains minimum on a closed disc
   460 text{* Hence a polynomial attains minimum on a closed disc
   380   in the complex plane. *}
   461   in the complex plane. *}
   381 lemma  poly_minimum_modulus_disc:
   462 lemma poly_minimum_modulus_disc: "\<exists>z. \<forall>w. cmod w \<le> r \<longrightarrow> cmod (poly p z) \<le> cmod (poly p w)"
   382   "\<exists>z. \<forall>w. cmod w \<le> r \<longrightarrow> cmod (poly p z) \<le> cmod (poly p w)"
   463 proof -
   383 proof-
   464   {
   384   {assume "\<not> r \<ge> 0" hence ?thesis
   465     assume "\<not> r \<ge> 0"
   385     by (metis norm_ge_zero order.trans)}
   466     then have ?thesis
       
   467       by (metis norm_ge_zero order.trans)
       
   468   }
   386   moreover
   469   moreover
   387   {assume rp: "r \<ge> 0"
   470   {
   388     from rp have "cmod 0 \<le> r \<and> cmod (poly p 0) = - (- cmod (poly p 0))" by simp
   471     assume rp: "r \<ge> 0"
   389     hence mth1: "\<exists>x z. cmod z \<le> r \<and> cmod (poly p z) = - x"  by blast
   472     from rp have "cmod 0 \<le> r \<and> cmod (poly p 0) = - (- cmod (poly p 0))"
   390     {fix x z
   473       by simp
   391       assume H: "cmod z \<le> r" "cmod (poly p z) = - x" "\<not>x < 1"
   474     then have mth1: "\<exists>x z. cmod z \<le> r \<and> cmod (poly p z) = - x"
   392       hence "- x < 0 " by arith
   475       by blast
   393       with H(2) norm_ge_zero[of "poly p z"]  have False by simp }
   476     {
   394     then have mth2: "\<exists>z. \<forall>x. (\<exists>z. cmod z \<le> r \<and> cmod (poly p z) = - x) \<longrightarrow> x < z" by blast
   477       fix x z
       
   478       assume H: "cmod z \<le> r" "cmod (poly p z) = - x" "\<not> x < 1"
       
   479       then have "- x < 0 "
       
   480         by arith
       
   481       with H(2) norm_ge_zero[of "poly p z"] have False
       
   482         by simp
       
   483     }
       
   484     then have mth2: "\<exists>z. \<forall>x. (\<exists>z. cmod z \<le> r \<and> cmod (poly p z) = - x) \<longrightarrow> x < z"
       
   485       by blast
   395     from real_sup_exists[OF mth1 mth2] obtain s where
   486     from real_sup_exists[OF mth1 mth2] obtain s where
   396       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
   487       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
   397     let ?m = "-s"
   488     let ?m = "- s"
   398     {fix y
   489     {
   399       from s[rule_format, of "-y"] have
   490       fix y
   400     "(\<exists>z x. cmod z \<le> r \<and> -(- cmod (poly p z)) < y) \<longleftrightarrow> ?m < y"
   491       from s[rule_format, of "-y"]
   401         unfolding minus_less_iff[of y ] equation_minus_iff by blast }
   492       have "(\<exists>z x. cmod z \<le> r \<and> - (- cmod (poly p z)) < y) \<longleftrightarrow> ?m < y"
       
   493         unfolding minus_less_iff[of y ] equation_minus_iff by blast
       
   494     }
   402     note s1 = this[unfolded minus_minus]
   495     note s1 = this[unfolded minus_minus]
   403     from s1[of ?m] have s1m: "\<And>z x. cmod z \<le> r \<Longrightarrow> cmod (poly p z) \<ge> ?m"
   496     from s1[of ?m] have s1m: "\<And>z x. cmod z \<le> r \<Longrightarrow> cmod (poly p z) \<ge> ?m"
   404       by auto
   497       by auto
   405     {fix n::nat
   498     {
       
   499       fix n :: nat
   406       from s1[rule_format, of "?m + 1/real (Suc n)"]
   500       from s1[rule_format, of "?m + 1/real (Suc n)"]
   407       have "\<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)"
   501       have "\<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)"
   408         by simp}
   502         by simp
   409     hence th: "\<forall>n. \<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)" ..
   503     }
       
   504     then have th: "\<forall>n. \<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)" ..
   410     from choice[OF th] obtain g where
   505     from choice[OF th] obtain g where
   411       g: "\<forall>n. cmod (g n) \<le> r" "\<forall>n. cmod (poly p (g n)) <?m+1 /real(Suc n)"
   506         g: "\<forall>n. cmod (g n) \<le> r" "\<forall>n. cmod (poly p (g n)) <?m + 1 /real(Suc n)"
   412       by blast
   507       by blast
   413     from bolzano_weierstrass_complex_disc[OF g(1)]
   508     from bolzano_weierstrass_complex_disc[OF g(1)]
   414     obtain f z where fz: "subseq f" "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. cmod (g (f n) - z) < e"
   509     obtain f z where fz: "subseq f" "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. cmod (g (f n) - z) < e"
   415       by blast
   510       by blast
   416     {fix w
   511     {
       
   512       fix w
   417       assume wr: "cmod w \<le> r"
   513       assume wr: "cmod w \<le> r"
   418       let ?e = "\<bar>cmod (poly p z) - ?m\<bar>"
   514       let ?e = "\<bar>cmod (poly p z) - ?m\<bar>"
   419       {assume e: "?e > 0"
   515       {
   420         hence e2: "?e/2 > 0" by simp
   516         assume e: "?e > 0"
       
   517         then have e2: "?e/2 > 0" by simp
   421         from poly_cont[OF e2, of z p] obtain d where
   518         from poly_cont[OF e2, of z p] obtain d where
   422           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
   519             d: "d > 0" "\<forall>w. 0<cmod (w - z)\<and> cmod(w - z) < d \<longrightarrow> cmod(poly p w - poly p z) < ?e/2"
   423         {fix w assume w: "cmod (w - z) < d"
   520           by blast
       
   521         {
       
   522           fix w
       
   523           assume w: "cmod (w - z) < d"
   424           have "cmod(poly p w - poly p z) < ?e / 2"
   524           have "cmod(poly p w - poly p z) < ?e / 2"
   425             using d(2)[rule_format, of w] w e by (cases "w=z", simp_all)}
   525             using d(2)[rule_format, of w] w e by (cases "w = z") simp_all
       
   526         }
   426         note th1 = this
   527         note th1 = this
   427 
   528 
   428         from fz(2)[rule_format, OF d(1)] obtain N1 where
   529         from fz(2) d(1) obtain N1 where N1: "\<forall>n\<ge>N1. cmod (g (f n) - z) < d"
   429           N1: "\<forall>n\<ge>N1. cmod (g (f n) - z) < d" by blast
   530           by blast
   430         from reals_Archimedean2[of "2/?e"] obtain N2::nat where
   531         from reals_Archimedean2[of "2/?e"] obtain N2 :: nat where N2: "2/?e < real N2"
   431           N2: "2/?e < real N2" by blast
   532           by blast
   432         have th2: "cmod(poly p (g(f(N1 + N2))) - poly p z) < ?e/2"
   533         have th2: "cmod (poly p (g (f (N1 + N2))) - poly p z) < ?e/2"
   433           using N1[rule_format, of "N1 + N2"] th1 by simp
   534           using N1[rule_format, of "N1 + N2"] th1 by simp
   434         {fix a b e2 m :: real
   535         {
   435         have "a < e2 \<Longrightarrow> abs(b - m) < e2 \<Longrightarrow> 2 * e2 <= abs(b - m) + a
   536           fix a b e2 m :: real
   436           ==> False" by arith}
   537           have "a < e2 \<Longrightarrow> \<bar>b - m\<bar> < e2 \<Longrightarrow> 2 * e2 \<le> \<bar>b - m\<bar> + a \<Longrightarrow> False"
   437       note th0 = this
   538             by arith
   438       have ath:
   539         }
   439         "\<And>m x e. m <= x \<Longrightarrow>  x < m + e ==> abs(x - m::real) < e" by arith
   540         note th0 = this
   440       from s1m[OF g(1)[rule_format]]
   541         have ath: "\<And>m x e::real. m \<le> x \<Longrightarrow> x < m + e \<Longrightarrow> \<bar>x - m\<bar> < e"
   441       have th31: "?m \<le> cmod(poly p (g (f (N1 + N2))))" .
   542           by arith
   442       from seq_suble[OF fz(1), of "N1+N2"]
   543         from s1m[OF g(1)[rule_format]] have th31: "?m \<le> cmod(poly p (g (f (N1 + N2))))" .
   443       have th00: "real (Suc (N1+N2)) \<le> real (Suc (f (N1+N2)))" by simp
   544         from seq_suble[OF fz(1), of "N1+N2"]
   444       have th000: "0 \<le> (1::real)" "(1::real) \<le> 1" "real (Suc (N1+N2)) > 0"
   545         have th00: "real (Suc (N1 + N2)) \<le> real (Suc (f (N1 + N2)))"
   445         using N2 by auto
   546           by simp
   446       from frac_le[OF th000 th00] have th00: "?m +1 / real (Suc (f (N1 + N2))) \<le> ?m + 1 / real (Suc (N1 + N2))" by simp
   547         have th000: "0 \<le> (1::real)" "(1::real) \<le> 1" "real (Suc (N1 + N2)) > 0"
   447       from g(2)[rule_format, of "f (N1 + N2)"]
   548           using N2 by auto
   448       have th01:"cmod (poly p (g (f (N1 + N2)))) < - s + 1 / real (Suc (f (N1 + N2)))" .
   549         from frac_le[OF th000 th00]
   449       from order_less_le_trans[OF th01 th00]
   550         have th00: "?m +1 / real (Suc (f (N1 + N2))) \<le> ?m + 1 / real (Suc (N1 + N2))"
   450       have th32: "cmod(poly p (g (f (N1 + N2)))) < ?m + (1/ real(Suc (N1 + N2)))" .
   551           by simp
   451       from N2 have "2/?e < real (Suc (N1 + N2))" by arith
   552         from g(2)[rule_format, of "f (N1 + N2)"]
   452       with e2 less_imp_inverse_less[of "2/?e" "real (Suc (N1 + N2))"]
   553         have th01:"cmod (poly p (g (f (N1 + N2)))) < - s + 1 / real (Suc (f (N1 + N2)))" .
   453       have "?e/2 > 1/ real (Suc (N1 + N2))" by (simp add: inverse_eq_divide)
   554         from order_less_le_trans[OF th01 th00]
   454       with ath[OF th31 th32]
   555         have th32: "cmod(poly p (g (f (N1 + N2)))) < ?m + (1/ real(Suc (N1 + N2)))" .
   455       have thc1:"\<bar>cmod(poly p (g (f (N1 + N2)))) - ?m\<bar>< ?e/2" by arith
   556         from N2 have "2/?e < real (Suc (N1 + N2))"
   456       have ath2: "\<And>(a::real) b c m. \<bar>a - b\<bar> <= c ==> \<bar>b - m\<bar> <= \<bar>a - m\<bar> + c"
   557           by arith
   457         by arith
   558         with e2 less_imp_inverse_less[of "2/?e" "real (Suc (N1 + N2))"]
   458       have th22: "\<bar>cmod (poly p (g (f (N1 + N2)))) - cmod (poly p z)\<bar>
   559         have "?e/2 > 1/ real (Suc (N1 + N2))"
   459 \<le> cmod (poly p (g (f (N1 + N2))) - poly p z)"
   560           by (simp add: inverse_eq_divide)
   460         by (simp add: norm_triangle_ineq3)
   561         with ath[OF th31 th32]
   461       from ath2[OF th22, of ?m]
   562         have thc1: "\<bar>cmod(poly p (g (f (N1 + N2)))) - ?m\<bar> < ?e/2"
   462       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
   563           by arith
   463       from th0[OF th2 thc1 thc2] have False .}
   564         have ath2: "\<And>a b c m::real. \<bar>a - b\<bar> \<le> c \<Longrightarrow> \<bar>b - m\<bar> \<le> \<bar>a - m\<bar> + c"
   464       hence "?e = 0" by auto
   565           by arith
   465       then have "cmod (poly p z) = ?m" by simp
   566         have th22: "\<bar>cmod (poly p (g (f (N1 + N2)))) - cmod (poly p z)\<bar> \<le>
   466       with s1m[OF wr]
   567             cmod (poly p (g (f (N1 + N2))) - poly p z)"
   467       have "cmod (poly p z) \<le> cmod (poly p w)" by simp }
   568           by (simp add: norm_triangle_ineq3)
   468     hence ?thesis by blast}
   569         from ath2[OF th22, of ?m]
       
   570         have thc2: "2 * (?e/2) \<le>
       
   571             \<bar>cmod(poly p (g (f (N1 + N2)))) - ?m\<bar> + cmod (poly p (g (f (N1 + N2))) - poly p z)"
       
   572           by simp
       
   573         from th0[OF th2 thc1 thc2] have False .
       
   574       }
       
   575       then have "?e = 0"
       
   576         by auto
       
   577       then have "cmod (poly p z) = ?m"
       
   578         by simp
       
   579       with s1m[OF wr] have "cmod (poly p z) \<le> cmod (poly p w)"
       
   580         by simp
       
   581     }
       
   582     then have ?thesis by blast
       
   583   }
   469   ultimately show ?thesis by blast
   584   ultimately show ?thesis by blast
   470 qed
   585 qed
   471 
   586 
   472 lemma "(rcis (sqrt (abs r)) (a/2))\<^sup>2 = rcis (abs r) a"
   587 lemma "(rcis (sqrt (abs r)) (a/2))\<^sup>2 = rcis (abs r) a"
   473   unfolding power2_eq_square
   588   unfolding power2_eq_square
   474   apply (simp add: rcis_mult)
   589   apply (simp add: rcis_mult)
   475   apply (simp add: power2_eq_square[symmetric])
   590   apply (simp add: power2_eq_square[symmetric])
   476   done
   591   done
   477 
   592 
   478 lemma cispi: "cis pi = -1"
   593 lemma cispi: "cis pi = -1"
   479   unfolding cis_def
   594   by (simp add: cis_def)
   480   by simp
       
   481 
   595 
   482 lemma "(rcis (sqrt (abs r)) ((pi + a)/2))\<^sup>2 = rcis (- abs r) a"
   596 lemma "(rcis (sqrt (abs r)) ((pi + a)/2))\<^sup>2 = rcis (- abs r) a"
   483   unfolding power2_eq_square
   597   unfolding power2_eq_square
   484   apply (simp add: rcis_mult add_divide_distrib)
   598   apply (simp add: rcis_mult add_divide_distrib)
   485   apply (simp add: power2_eq_square[symmetric] rcis_def cispi cis_mult[symmetric])
   599   apply (simp add: power2_eq_square[symmetric] rcis_def cispi cis_mult[symmetric])
   486   done
   600   done
   487 
   601 
   488 text {* Nonzero polynomial in z goes to infinity as z does. *}
   602 text {* Nonzero polynomial in z goes to infinity as z does. *}
   489 
   603 
   490 lemma poly_infinity:
   604 lemma poly_infinity:
   491   fixes p:: "('a::{comm_semiring_0,real_normed_div_algebra}) poly" 
   605   fixes p:: "'a::{comm_semiring_0,real_normed_div_algebra} poly"
   492   assumes ex: "p \<noteq> 0"
   606   assumes ex: "p \<noteq> 0"
   493   shows "\<exists>r. \<forall>z. r \<le> norm z \<longrightarrow> d \<le> norm (poly (pCons a p) z)"
   607   shows "\<exists>r. \<forall>z. r \<le> norm z \<longrightarrow> d \<le> norm (poly (pCons a p) z)"
   494 using ex
   608   using ex
   495 proof(induct p arbitrary: a d)
   609 proof (induct p arbitrary: a d)
   496   case (pCons c cs a d)
   610   case (pCons c cs a d)
   497   {assume H: "cs \<noteq> 0"
   611   {
   498     with pCons.hyps obtain r where r: "\<forall>z. r \<le> norm z \<longrightarrow> d + norm a \<le> norm (poly (pCons c cs) z)" by blast
   612     assume H: "cs \<noteq> 0"
       
   613     with pCons.hyps obtain r where r: "\<forall>z. r \<le> norm z \<longrightarrow> d + norm a \<le> norm (poly (pCons c cs) z)"
       
   614       by blast
   499     let ?r = "1 + \<bar>r\<bar>"
   615     let ?r = "1 + \<bar>r\<bar>"
   500     {fix z::'a assume h: "1 + \<bar>r\<bar> \<le> norm z"
   616     {
       
   617       fix z::'a
       
   618       assume h: "1 + \<bar>r\<bar> \<le> norm z"
   501       have r0: "r \<le> norm z" using h by arith
   619       have r0: "r \<le> norm z" using h by arith
   502       from r[rule_format, OF r0]
   620       from r[rule_format, OF r0] have th0: "d + norm a \<le> 1 * norm(poly (pCons c cs) z)"
   503       have th0: "d + norm a \<le> 1 * norm(poly (pCons c cs) z)" by arith
   621         by arith
   504       from h have z1: "norm z \<ge> 1" by arith
   622       from h have z1: "norm z \<ge> 1"
       
   623         by arith
   505       from order_trans[OF th0 mult_right_mono[OF z1 norm_ge_zero[of "poly (pCons c cs) z"]]]
   624       from order_trans[OF th0 mult_right_mono[OF z1 norm_ge_zero[of "poly (pCons c cs) z"]]]
   506       have th1: "d \<le> norm(z * poly (pCons c cs) z) - norm a"
   625       have th1: "d \<le> norm(z * poly (pCons c cs) z) - norm a"
   507         unfolding norm_mult by (simp add: algebra_simps)
   626         unfolding norm_mult by (simp add: algebra_simps)
   508       from norm_diff_ineq[of "z * poly (pCons c cs) z" a]
   627       from norm_diff_ineq[of "z * poly (pCons c cs) z" a]
   509       have th2: "norm(z * poly (pCons c cs) z) - norm a \<le> norm (poly (pCons a (pCons c cs)) z)"
   628       have th2: "norm(z * poly (pCons c cs) z) - norm a \<le> norm (poly (pCons a (pCons c cs)) z)"
   510         by (simp add: algebra_simps)
   629         by (simp add: algebra_simps)
   511       from th1 th2 have "d \<le> norm (poly (pCons a (pCons c cs)) z)"  by arith}
   630       from th1 th2 have "d \<le> norm (poly (pCons a (pCons c cs)) z)" by arith
   512     hence ?case by blast}
   631     }
       
   632     then have ?case by blast
       
   633   }
   513   moreover
   634   moreover
   514   {assume cs0: "\<not> (cs \<noteq> 0)"
   635   {
   515     with pCons.prems have c0: "c \<noteq> 0" by simp
   636     assume cs0: "\<not> (cs \<noteq> 0)"
   516     from cs0 have cs0': "cs = 0" by simp
   637     with pCons.prems have c0: "c \<noteq> 0"
   517     {fix z::'a
   638       by simp
       
   639     from cs0 have cs0': "cs = 0"
       
   640       by simp
       
   641     {
       
   642       fix z::'a
   518       assume h: "(\<bar>d\<bar> + norm a) / norm c \<le> norm z"
   643       assume h: "(\<bar>d\<bar> + norm a) / norm c \<le> norm z"
   519       from c0 have "norm c > 0" by simp
   644       from c0 have "norm c > 0"
       
   645         by simp
   520       from h c0 have th0: "\<bar>d\<bar> + norm a \<le> norm (z * c)"
   646       from h c0 have th0: "\<bar>d\<bar> + norm a \<le> norm (z * c)"
   521         by (simp add: field_simps norm_mult)
   647         by (simp add: field_simps norm_mult)
   522       have ath: "\<And>mzh mazh ma. mzh <= mazh + ma ==> abs(d) + ma <= mzh ==> d <= mazh" by arith
   648       have ath: "\<And>mzh mazh ma. mzh \<le> mazh + ma \<Longrightarrow> \<bar>d\<bar> + ma \<le> mzh \<Longrightarrow> d \<le> mazh"
   523       from norm_diff_ineq[of "z * c" a ]
   649         by arith
   524       have th1: "norm (z * c) \<le> norm (a + z * c) + norm a"
   650       from norm_diff_ineq[of "z * c" a] have th1: "norm (z * c) \<le> norm (a + z * c) + norm a"
   525         by (simp add: algebra_simps)
   651         by (simp add: algebra_simps)
   526       from ath[OF th1 th0] have "d \<le> norm (poly (pCons a (pCons c cs)) z)"
   652       from ath[OF th1 th0] have "d \<le> norm (poly (pCons a (pCons c cs)) z)"
   527         using cs0' by simp}
   653         using cs0' by simp
   528     then have ?case  by blast}
   654     }
       
   655     then have ?case  by blast
       
   656   }
   529   ultimately show ?case by blast
   657   ultimately show ?case by blast
   530 qed simp
   658 qed simp
   531 
   659 
   532 text {* Hence polynomial's modulus attains its minimum somewhere. *}
   660 text {* Hence polynomial's modulus attains its minimum somewhere. *}
   533 lemma poly_minimum_modulus:
   661 lemma poly_minimum_modulus: "\<exists>z.\<forall>w. cmod (poly p z) \<le> cmod (poly p w)"
   534   "\<exists>z.\<forall>w. cmod (poly p z) \<le> cmod (poly p w)"
   662 proof (induct p)
   535 proof(induct p)
   663   case 0
       
   664   then show ?case by simp
       
   665 next
   536   case (pCons c cs)
   666   case (pCons c cs)
   537   {assume cs0: "cs \<noteq> 0"
   667   show ?case
   538     from poly_infinity[OF cs0, of "cmod (poly (pCons c cs) 0)" c]
   668   proof (cases "cs = 0")
   539     obtain r where r: "\<And>z. r \<le> cmod z \<Longrightarrow> cmod (poly (pCons c cs) 0) \<le> cmod (poly (pCons c cs) z)" by blast
   669     case False
   540     have ath: "\<And>z r. r \<le> cmod z \<or> cmod z \<le> \<bar>r\<bar>" by arith
   670     from poly_infinity[OF False, of "cmod (poly (pCons c cs) 0)" c]
       
   671     obtain r where r: "\<And>z. r \<le> cmod z \<Longrightarrow> cmod (poly (pCons c cs) 0) \<le> cmod (poly (pCons c cs) z)"
       
   672       by blast
       
   673     have ath: "\<And>z r. r \<le> cmod z \<or> cmod z \<le> \<bar>r\<bar>"
       
   674       by arith
   541     from poly_minimum_modulus_disc[of "\<bar>r\<bar>" "pCons c cs"]
   675     from poly_minimum_modulus_disc[of "\<bar>r\<bar>" "pCons c cs"]
   542     obtain v where v: "\<And>w. cmod w \<le> \<bar>r\<bar> \<Longrightarrow> cmod (poly (pCons c cs) v) \<le> cmod (poly (pCons c cs) w)" by blast
   676     obtain v where v: "\<And>w. cmod w \<le> \<bar>r\<bar> \<Longrightarrow> cmod (poly (pCons c cs) v) \<le> cmod (poly (pCons c cs) w)"
   543     {fix z assume z: "r \<le> cmod z"
   677       by blast
   544       from v[of 0] r[OF z]
   678     {
   545       have "cmod (poly (pCons c cs) v) \<le> cmod (poly (pCons c cs) z)"
   679       fix z
   546         by simp }
   680       assume z: "r \<le> cmod z"
       
   681       from v[of 0] r[OF z] have "cmod (poly (pCons c cs) v) \<le> cmod (poly (pCons c cs) z)"
       
   682         by simp
       
   683     }
   547     note v0 = this
   684     note v0 = this
   548     from v0 v ath[of r] have ?case by blast}
   685     from v0 v ath[of r] show ?thesis
   549   moreover
   686       by blast
   550   {assume cs0: "\<not> (cs \<noteq> 0)"
   687   next
   551     hence th:"cs = 0" by simp
   688     case True
   552     from th pCons.hyps have ?case by simp}
   689     with pCons.hyps show ?thesis by simp
   553   ultimately show ?case by blast
   690   qed
   554 qed simp
   691 qed
   555 
   692 
   556 text{* Constant function (non-syntactic characterization). *}
   693 text{* Constant function (non-syntactic characterization). *}
   557 definition "constant f = (\<forall>x y. f x = f y)"
   694 definition "constant f = (\<forall>x y. f x = f y)"
   558 
   695 
   559 lemma nonconstant_length: "\<not> (constant (poly p)) \<Longrightarrow> psize p \<ge> 2"
   696 lemma nonconstant_length: "\<not> constant (poly p) \<Longrightarrow> psize p \<ge> 2"
   560   unfolding constant_def psize_def
   697   by (induct p) (auto simp: constant_def psize_def)
   561   apply (induct p, auto)
       
   562   done
       
   563 
   698 
   564 lemma poly_replicate_append:
   699 lemma poly_replicate_append:
   565   "poly (monom 1 n * p) (x::'a::{comm_ring_1}) = x^n * poly p x"
   700   "poly (monom 1 n * p) (x::'a::{comm_ring_1}) = x^n * poly p x"
   566   by (simp add: poly_monom)
   701   by (simp add: poly_monom)
   567 
   702 
   568 text {* Decomposition of polynomial, skipping zero coefficients
   703 text {* Decomposition of polynomial, skipping zero coefficients
   569   after the first.  *}
   704   after the first.  *}
   570 
   705 
   571 lemma poly_decompose_lemma:
   706 lemma poly_decompose_lemma:
   572  assumes nz: "\<not>(\<forall>z. z\<noteq>0 \<longrightarrow> poly p z = (0::'a::{idom}))"
   707   assumes nz: "\<not> (\<forall>z. z \<noteq> 0 \<longrightarrow> poly p z = (0::'a::idom))"
   573   shows "\<exists>k a q. a\<noteq>0 \<and> Suc (psize q + k) = psize p \<and>
   708   shows "\<exists>k a q. a \<noteq> 0 \<and> Suc (psize q + k) = psize p \<and>
   574                  (\<forall>z. poly p z = z^k * poly (pCons a q) z)"
   709     (\<forall>z. poly p z = z^k * poly (pCons a q) z)"
   575 unfolding psize_def
   710   unfolding psize_def
   576 using nz
   711   using nz
   577 proof(induct p)
   712 proof (induct p)
   578   case 0 thus ?case by simp
   713   case 0
       
   714   then show ?case by simp
   579 next
   715 next
   580   case (pCons c cs)
   716   case (pCons c cs)
   581   {assume c0: "c = 0"
   717   show ?case
   582     from pCons.hyps pCons.prems c0 have ?case
   718   proof (cases "c = 0")
       
   719     case True
       
   720     from pCons.hyps pCons.prems True show ?thesis
   583       apply (auto)
   721       apply (auto)
   584       apply (rule_tac x="k+1" in exI)
   722       apply (rule_tac x="k+1" in exI)
   585       apply (rule_tac x="a" in exI, clarsimp)
   723       apply (rule_tac x="a" in exI, clarsimp)
   586       apply (rule_tac x="q" in exI)
   724       apply (rule_tac x="q" in exI)
   587       by (auto)}
   725       apply auto
   588   moreover
   726       done
   589   {assume c0: "c\<noteq>0"
   727   next
   590     have ?case 
   728     case False
       
   729     show ?thesis
   591       apply (rule exI[where x=0])
   730       apply (rule exI[where x=0])
   592       apply (rule exI[where x=c], auto simp add: c0)
   731       apply (rule exI[where x=c], auto simp add: False)
   593       done}
   732       done
   594   ultimately show ?case by blast
   733   qed
   595 qed
   734 qed
   596 
   735 
   597 lemma poly_decompose:
   736 lemma poly_decompose:
   598   assumes nc: "\<not> constant (poly p)"
   737   assumes nc: "\<not> constant (poly p)"
   599   shows "\<exists>k a q. a \<noteq> (0::'a::{idom}) \<and> k \<noteq> 0 \<and>
   738   shows "\<exists>k a q. a \<noteq> (0::'a::idom) \<and> k \<noteq> 0 \<and>
   600                psize q + k + 1 = psize p \<and>
   739                psize q + k + 1 = psize p \<and>
   601               (\<forall>z. poly p z = poly p 0 + z^k * poly (pCons a q) z)"
   740               (\<forall>z. poly p z = poly p 0 + z^k * poly (pCons a q) z)"
   602   using nc
   741   using nc
   603 proof (induct p)
   742 proof (induct p)
   604   case 0
   743   case 0
   611     {
   750     {
   612       fix x y
   751       fix x y
   613       from C have "poly (pCons c cs) x = poly (pCons c cs) y"
   752       from C have "poly (pCons c cs) x = poly (pCons c cs) y"
   614         by (cases "x = 0") auto
   753         by (cases "x = 0") auto
   615     }
   754     }
   616     with pCons.prems have False by (auto simp add: constant_def)
   755     with pCons.prems have False
       
   756       by (auto simp add: constant_def)
   617   }
   757   }
   618   then have th: "\<not> (\<forall>z. z \<noteq> 0 \<longrightarrow> poly cs z = 0)" ..
   758   then have th: "\<not> (\<forall>z. z \<noteq> 0 \<longrightarrow> poly cs z = 0)" ..
   619   from poly_decompose_lemma[OF th]
   759   from poly_decompose_lemma[OF th]
   620   show ?case
   760   show ?case
   621     apply clarsimp
   761     apply clarsimp
   639   let ?ths = "\<exists>z. ?p z = 0"
   779   let ?ths = "\<exists>z. ?p z = 0"
   640 
   780 
   641   from nonconstant_length[OF less(2)] have n2: "psize p \<ge> 2" .
   781   from nonconstant_length[OF less(2)] have n2: "psize p \<ge> 2" .
   642   from poly_minimum_modulus obtain c where c: "\<forall>w. cmod (?p c) \<le> cmod (?p w)"
   782   from poly_minimum_modulus obtain c where c: "\<forall>w. cmod (?p c) \<le> cmod (?p w)"
   643     by blast
   783     by blast
   644   {assume pc: "?p c = 0" hence ?ths by blast}
   784 
   645   moreover
   785   show ?ths
   646   {assume pc0: "?p c \<noteq> 0"
   786   proof (cases "?p c = 0")
   647     from poly_offset[of p c] obtain q where
   787     case True
   648       q: "psize q = psize p" "\<forall>x. poly q x = ?p (c+x)" by blast
   788     then show ?thesis by blast
   649     {assume h: "constant (poly q)"
   789   next
       
   790     case False
       
   791     note pc0 = this
       
   792     from poly_offset[of p c] obtain q where q: "psize q = psize p" "\<forall>x. poly q x = ?p (c + x)"
       
   793       by blast
       
   794     {
       
   795       assume h: "constant (poly q)"
   650       from q(2) have th: "\<forall>x. poly q (x - c) = ?p x" by auto
   796       from q(2) have th: "\<forall>x. poly q (x - c) = ?p x" by auto
   651       {fix x y
   797       {
       
   798         fix x y
   652         from th have "?p x = poly q (x - c)" by auto
   799         from th have "?p x = poly q (x - c)" by auto
   653         also have "\<dots> = poly q (y - c)"
   800         also have "\<dots> = poly q (y - c)"
   654           using h unfolding constant_def by blast
   801           using h unfolding constant_def by blast
   655         also have "\<dots> = ?p y" using th by auto
   802         also have "\<dots> = ?p y" using th by auto
   656         finally have "?p x = ?p y" .}
   803         finally have "?p x = ?p y" .
   657       with less(2) have False unfolding constant_def by blast }
   804       }
   658     hence qnc: "\<not> constant (poly q)" by blast
   805       with less(2) have False
   659     from q(2) have pqc0: "?p c = poly q 0" by simp
   806         unfolding constant_def by blast
   660     from c pqc0 have cq0: "\<forall>w. cmod (poly q 0) \<le> cmod (?p w)" by simp
   807     }
       
   808     then have qnc: "\<not> constant (poly q)"
       
   809       by blast
       
   810     from q(2) have pqc0: "?p c = poly q 0"
       
   811       by simp
       
   812     from c pqc0 have cq0: "\<forall>w. cmod (poly q 0) \<le> cmod (?p w)"
       
   813       by simp
   661     let ?a0 = "poly q 0"
   814     let ?a0 = "poly q 0"
   662     from pc0 pqc0 have a00: "?a0 \<noteq> 0" by simp
   815     from pc0 pqc0 have a00: "?a0 \<noteq> 0"
   663     from a00
   816       by simp
   664     have qr: "\<forall>z. poly q z = poly (smult (inverse ?a0) q) z * ?a0"
   817     from a00 have qr: "\<forall>z. poly q z = poly (smult (inverse ?a0) q) z * ?a0"
   665       by simp
   818       by simp
   666     let ?r = "smult (inverse ?a0) q"
   819     let ?r = "smult (inverse ?a0) q"
   667     have lgqr: "psize q = psize ?r"
   820     have lgqr: "psize q = psize ?r"
   668       using a00 unfolding psize_def degree_def
   821       using a00
       
   822       unfolding psize_def degree_def
   669       by (simp add: poly_eq_iff)
   823       by (simp add: poly_eq_iff)
   670     {assume h: "\<And>x y. poly ?r x = poly ?r y"
   824     {
   671       {fix x y
   825       assume h: "\<And>x y. poly ?r x = poly ?r y"
   672         from qr[rule_format, of x]
   826       {
   673         have "poly q x = poly ?r x * ?a0" by auto
   827         fix x y
   674         also have "\<dots> = poly ?r y * ?a0" using h by simp
   828         from qr[rule_format, of x] have "poly q x = poly ?r x * ?a0"
   675         also have "\<dots> = poly q y" using qr[rule_format, of y] by simp
   829           by auto
   676         finally have "poly q x = poly q y" .}
   830         also have "\<dots> = poly ?r y * ?a0"
   677       with qnc have False unfolding constant_def by blast}
   831           using h by simp
   678     hence rnc: "\<not> constant (poly ?r)" unfolding constant_def by blast
   832         also have "\<dots> = poly q y"
   679     from qr[rule_format, of 0] a00  have r01: "poly ?r 0 = 1" by auto
   833           using qr[rule_format, of y] by simp
   680     {fix w
   834         finally have "poly q x = poly q y" .
       
   835       }
       
   836       with qnc have False unfolding constant_def by blast
       
   837     }
       
   838     then have rnc: "\<not> constant (poly ?r)"
       
   839       unfolding constant_def by blast
       
   840     from qr[rule_format, of 0] a00 have r01: "poly ?r 0 = 1"
       
   841       by auto
       
   842     {
       
   843       fix w
   681       have "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w / ?a0) < 1"
   844       have "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w / ?a0) < 1"
   682         using qr[rule_format, of w] a00 by (simp add: divide_inverse mult_ac)
   845         using qr[rule_format, of w] a00 by (simp add: divide_inverse mult_ac)
   683       also have "\<dots> \<longleftrightarrow> cmod (poly q w) < cmod ?a0"
   846       also have "\<dots> \<longleftrightarrow> cmod (poly q w) < cmod ?a0"
   684         using a00 unfolding norm_divide by (simp add: field_simps)
   847         using a00 unfolding norm_divide by (simp add: field_simps)
   685       finally have "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w) < cmod ?a0" .}
   848       finally have "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w) < cmod ?a0" .
       
   849     }
   686     note mrmq_eq = this
   850     note mrmq_eq = this
   687     from poly_decompose[OF rnc] obtain k a s where
   851     from poly_decompose[OF rnc] obtain k a s where
   688       kas: "a\<noteq>0" "k\<noteq>0" "psize s + k + 1 = psize ?r"
   852       kas: "a \<noteq> 0" "k \<noteq> 0" "psize s + k + 1 = psize ?r"
   689       "\<forall>z. poly ?r z = poly ?r 0 + z^k* poly (pCons a s) z" by blast
   853         "\<forall>z. poly ?r z = poly ?r 0 + z^k* poly (pCons a s) z" by blast
   690     {assume "psize p = k + 1"
   854     {
   691       with kas(3) lgqr[symmetric] q(1) have s0:"s=0" by auto
   855       assume "psize p = k + 1"
   692       {fix w
   856       with kas(3) lgqr[symmetric] q(1) have s0: "s = 0"
       
   857         by auto
       
   858       {
       
   859         fix w
   693         have "cmod (poly ?r w) = cmod (1 + a * w ^ k)"
   860         have "cmod (poly ?r w) = cmod (1 + a * w ^ k)"
   694           using kas(4)[rule_format, of w] s0 r01 by (simp add: algebra_simps)}
   861           using kas(4)[rule_format, of w] s0 r01 by (simp add: algebra_simps)
       
   862       }
   695       note hth = this [symmetric]
   863       note hth = this [symmetric]
   696         from reduce_poly_simple[OF kas(1,2)]
   864       from reduce_poly_simple[OF kas(1,2)] have "\<exists>w. cmod (poly ?r w) < 1"
   697       have "\<exists>w. cmod (poly ?r w) < 1" unfolding hth by blast}
   865         unfolding hth by blast
       
   866     }
   698     moreover
   867     moreover
   699     {assume kn: "psize p \<noteq> k+1"
   868     {
   700       from kn kas(3) q(1) lgqr have k1n: "k + 1 < psize p" by simp
   869       assume kn: "psize p \<noteq> k + 1"
       
   870       from kn kas(3) q(1) lgqr have k1n: "k + 1 < psize p"
       
   871         by simp
   701       have th01: "\<not> constant (poly (pCons 1 (monom a (k - 1))))"
   872       have th01: "\<not> constant (poly (pCons 1 (monom a (k - 1))))"
   702         unfolding constant_def poly_pCons poly_monom
   873         unfolding constant_def poly_pCons poly_monom
   703         using kas(1) apply simp
   874         using kas(1) apply simp
   704         by (rule exI[where x=0], rule exI[where x=1], simp)
   875         apply (rule exI[where x=0])
   705       from kas(1) kas(2) have th02: "k+1 = psize (pCons 1 (monom a (k - 1)))"
   876         apply (rule exI[where x=1])
       
   877         apply simp
       
   878         done
       
   879       from kas(1) kas(2) have th02: "k + 1 = psize (pCons 1 (monom a (k - 1)))"
   706         by (simp add: psize_def degree_monom_eq)
   880         by (simp add: psize_def degree_monom_eq)
   707       from less(1) [OF k1n [simplified th02] th01]
   881       from less(1) [OF k1n [simplified th02] th01]
   708       obtain w where w: "1 + w^k * a = 0"
   882       obtain w where w: "1 + w^k * a = 0"
   709         unfolding poly_pCons poly_monom
   883         unfolding poly_pCons poly_monom
   710         using kas(2) by (cases k, auto simp add: algebra_simps)
   884         using kas(2) by (cases k) (auto simp add: algebra_simps)
   711       from poly_bound_exists[of "cmod w" s] obtain m where
   885       from poly_bound_exists[of "cmod w" s] obtain m where
   712         m: "m > 0" "\<forall>z. cmod z \<le> cmod w \<longrightarrow> cmod (poly s z) \<le> m" by blast
   886         m: "m > 0" "\<forall>z. cmod z \<le> cmod w \<longrightarrow> cmod (poly s z) \<le> m" by blast
   713       have w0: "w\<noteq>0" using kas(2) w by (auto simp add: power_0_left)
   887       have w0: "w \<noteq> 0" using kas(2) w
   714       from w have "(1 + w ^ k * a) - 1 = 0 - 1" by simp
   888         by (auto simp add: power_0_left)
   715       then have wm1: "w^k * a = - 1" by simp
   889       from w have "(1 + w ^ k * a) - 1 = 0 - 1"
       
   890         by simp
       
   891       then have wm1: "w^k * a = - 1"
       
   892         by simp
   716       have inv0: "0 < inverse (cmod w ^ (k + 1) * m)"
   893       have inv0: "0 < inverse (cmod w ^ (k + 1) * m)"
   717         using norm_ge_zero[of w] w0 m(1)
   894         using norm_ge_zero[of w] w0 m(1)
   718           by (simp add: inverse_eq_divide zero_less_mult_iff)
   895         by (simp add: inverse_eq_divide zero_less_mult_iff)
   719       with real_lbound_gt_zero[OF zero_less_one] obtain t where
   896       with real_lbound_gt_zero[OF zero_less_one] obtain t where
   720         t: "t > 0" "t < 1" "t < inverse (cmod w ^ (k + 1) * m)" by blast
   897         t: "t > 0" "t < 1" "t < inverse (cmod w ^ (k + 1) * m)" by blast
   721       let ?ct = "complex_of_real t"
   898       let ?ct = "complex_of_real t"
   722       let ?w = "?ct * w"
   899       let ?w = "?ct * w"
   723       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: algebra_simps power_mult_distrib)
   900       have "1 + ?w^k * (a + ?w * poly s ?w) = 1 + ?ct^k * (w^k * a) + ?w^k * ?w * poly s ?w"
       
   901         using kas(1) by (simp add: algebra_simps power_mult_distrib)
   724       also have "\<dots> = complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w"
   902       also have "\<dots> = complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w"
   725         unfolding wm1 by (simp)
   903         unfolding wm1 by simp
   726       finally have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) = cmod (complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w)"
   904       finally have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) =
       
   905         cmod (complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w)"
   727         by metis
   906         by metis
   728       with norm_triangle_ineq[of "complex_of_real (1 - t^k)" "?w^k * ?w * poly s ?w"]
   907       with norm_triangle_ineq[of "complex_of_real (1 - t^k)" "?w^k * ?w * poly s ?w"]
   729       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 norm_of_real by simp
   908       have th11: "cmod (1 + ?w^k * (a + ?w * poly s ?w)) \<le> \<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w)"
   730       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
   909         unfolding norm_of_real by simp
   731       have "t * cmod w \<le> 1 * cmod w" apply (rule mult_mono) using t(1,2) by auto
   910       have ath: "\<And>x t::real. 0 \<le> x \<Longrightarrow> x < t \<Longrightarrow> t \<le> 1 \<Longrightarrow> \<bar>1 - t\<bar> + x < 1"
   732       then have tw: "cmod ?w \<le> cmod w" using t(1) by (simp add: norm_mult)
   911         by arith
   733       from t inv0 have "t* (cmod w ^ (k + 1) * m) < 1"
   912       have "t * cmod w \<le> 1 * cmod w"
       
   913         apply (rule mult_mono)
       
   914         using t(1,2)
       
   915         apply auto
       
   916         done
       
   917       then have tw: "cmod ?w \<le> cmod w"
       
   918         using t(1) by (simp add: norm_mult)
       
   919       from t inv0 have "t * (cmod w ^ (k + 1) * m) < 1"
   734         by (simp add: inverse_eq_divide field_simps)
   920         by (simp add: inverse_eq_divide field_simps)
   735       with zero_less_power[OF t(1), of k]
   921       with zero_less_power[OF t(1), of k] have th30: "t^k * (t* (cmod w ^ (k + 1) * m)) < t^k * 1"
   736       have th30: "t^k * (t* (cmod w ^ (k + 1) * m)) < t^k * 1"
       
   737         by (metis comm_mult_strict_left_mono)
   922         by (metis comm_mult_strict_left_mono)
   738       have "cmod (?w^k * ?w * poly s ?w) = t^k * (t* (cmod w ^ (k+1) * cmod (poly s ?w)))"  using w0 t(1)
   923       have "cmod (?w^k * ?w * poly s ?w) = t^k * (t* (cmod w ^ (k + 1) * cmod (poly s ?w)))"
       
   924         using w0 t(1)
   739         by (simp add: algebra_simps power_mult_distrib norm_power norm_mult)
   925         by (simp add: algebra_simps power_mult_distrib norm_power norm_mult)
   740       then have "cmod (?w^k * ?w * poly s ?w) \<le> t^k * (t* (cmod w ^ (k + 1) * m))"
   926       then have "cmod (?w^k * ?w * poly s ?w) \<le> t^k * (t* (cmod w ^ (k + 1) * m))"
   741         using t(1,2) m(2)[rule_format, OF tw] w0
   927         using t(1,2) m(2)[rule_format, OF tw] w0
   742         by auto
   928         by auto
   743       with th30 have th120: "cmod (?w^k * ?w * poly s ?w) < t^k" by simp
   929       with th30 have th120: "cmod (?w^k * ?w * poly s ?w) < t^k"
       
   930         by simp
   744       from power_strict_mono[OF t(2), of k] t(1) kas(2) have th121: "t^k \<le> 1"
   931       from power_strict_mono[OF t(2), of k] t(1) kas(2) have th121: "t^k \<le> 1"
   745         by auto
   932         by auto
   746       from ath[OF norm_ge_zero[of "?w^k * ?w * poly s ?w"] th120 th121]
   933       from ath[OF norm_ge_zero[of "?w^k * ?w * poly s ?w"] th120 th121]
   747       have th12: "\<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w) < 1" .
   934       have th12: "\<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w) < 1" .
   748       from th11 th12
   935       from th11 th12 have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) < 1"
   749       have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) < 1"  by arith
   936         by arith
   750       then have "cmod (poly ?r ?w) < 1"
   937       then have "cmod (poly ?r ?w) < 1"
   751         unfolding kas(4)[rule_format, of ?w] r01 by simp
   938         unfolding kas(4)[rule_format, of ?w] r01 by simp
   752       then have "\<exists>w. cmod (poly ?r w) < 1" by blast}
   939       then have "\<exists>w. cmod (poly ?r w) < 1"
   753     ultimately have cr0_contr: "\<exists>w. cmod (poly ?r w) < 1" by blast
   940         by blast
   754     from cr0_contr cq0 q(2)
   941     }
   755     have ?ths unfolding mrmq_eq not_less[symmetric] by auto}
   942     ultimately have cr0_contr: "\<exists>w. cmod (poly ?r w) < 1"
   756   ultimately show ?ths by blast
   943       by blast
       
   944     from cr0_contr cq0 q(2) show ?thesis
       
   945       unfolding mrmq_eq not_less[symmetric] by auto
       
   946   qed
   757 qed
   947 qed
   758 
   948 
   759 text {* Alternative version with a syntactic notion of constant polynomial. *}
   949 text {* Alternative version with a syntactic notion of constant polynomial. *}
   760 
   950 
   761 lemma fundamental_theorem_of_algebra_alt:
   951 lemma fundamental_theorem_of_algebra_alt:
   762   assumes nc: "~(\<exists>a l. a\<noteq> 0 \<and> l = 0 \<and> p = pCons a l)"
   952   assumes nc: "\<not> (\<exists>a l. a \<noteq> 0 \<and> l = 0 \<and> p = pCons a l)"
   763   shows "\<exists>z. poly p z = (0::complex)"
   953   shows "\<exists>z. poly p z = (0::complex)"
   764 using nc
   954   using nc
   765 proof(induct p)
   955 proof (induct p)
       
   956   case 0
       
   957   then show ?case by simp
       
   958 next
   766   case (pCons c cs)
   959   case (pCons c cs)
   767   {assume "c=0" hence ?case by auto}
   960   show ?case
   768   moreover
   961   proof (cases "c = 0")
   769   {assume c0: "c\<noteq>0"
   962     case True
   770     {assume nc: "constant (poly (pCons c cs))"
   963     then show ?thesis by auto
       
   964   next
       
   965     case False
       
   966     {
       
   967       assume nc: "constant (poly (pCons c cs))"
   771       from nc[unfolded constant_def, rule_format, of 0]
   968       from nc[unfolded constant_def, rule_format, of 0]
   772       have "\<forall>w. w \<noteq> 0 \<longrightarrow> poly cs w = 0" by auto
   969       have "\<forall>w. w \<noteq> 0 \<longrightarrow> poly cs w = 0" by auto
   773       hence "cs = 0"
   970       then have "cs = 0"
   774         proof(induct cs)
   971       proof (induct cs)
   775           case (pCons d ds)
   972         case 0
   776           {assume "d=0" hence ?case using pCons.prems pCons.hyps by simp}
   973         then show ?case by simp
   777           moreover
   974       next
   778           {assume d0: "d\<noteq>0"
   975         case (pCons d ds)
   779             from poly_bound_exists[of 1 ds] obtain m where
   976         show ?case
   780               m: "m > 0" "\<forall>z. \<forall>z. cmod z \<le> 1 \<longrightarrow> cmod (poly ds z) \<le> m" by blast
   977         proof (cases "d = 0")
   781             have dm: "cmod d / m > 0" using d0 m(1) by (simp add: field_simps)
   978           case True
   782             from real_lbound_gt_zero[OF dm zero_less_one] obtain x where
   979           then show ?thesis using pCons.prems pCons.hyps by simp
   783               x: "x > 0" "x < cmod d / m" "x < 1" by blast
   980         next
   784             let ?x = "complex_of_real x"
   981           case False
   785             from x have cx: "?x \<noteq> 0"  "cmod ?x \<le> 1" by simp_all
   982           from poly_bound_exists[of 1 ds] obtain m where
   786             from pCons.prems[rule_format, OF cx(1)]
   983             m: "m > 0" "\<forall>z. \<forall>z. cmod z \<le> 1 \<longrightarrow> cmod (poly ds z) \<le> m" by blast
   787             have cth: "cmod (?x*poly ds ?x) = cmod d" by (simp add: eq_diff_eq[symmetric])
   984           have dm: "cmod d / m > 0" using False m(1) by (simp add: field_simps)
   788             from m(2)[rule_format, OF cx(2)] x(1)
   985           from real_lbound_gt_zero[OF dm zero_less_one] obtain x where
   789             have th0: "cmod (?x*poly ds ?x) \<le> x*m"
   986             x: "x > 0" "x < cmod d / m" "x < 1" by blast
   790               by (simp add: norm_mult)
   987           let ?x = "complex_of_real x"
   791             from x(2) m(1) have "x*m < cmod d" by (simp add: field_simps)
   988           from x have cx: "?x \<noteq> 0"  "cmod ?x \<le> 1" by simp_all
   792             with th0 have "cmod (?x*poly ds ?x) \<noteq> cmod d" by auto
   989           from pCons.prems[rule_format, OF cx(1)]
   793             with cth  have ?case by blast}
   990           have cth: "cmod (?x*poly ds ?x) = cmod d" by (simp add: eq_diff_eq[symmetric])
   794           ultimately show ?case by blast
   991           from m(2)[rule_format, OF cx(2)] x(1)
   795         qed simp}
   992           have th0: "cmod (?x*poly ds ?x) \<le> x*m"
   796       then have nc: "\<not> constant (poly (pCons c cs))" using pCons.prems c0
   993             by (simp add: norm_mult)
   797         by blast
   994           from x(2) m(1) have "x*m < cmod d" by (simp add: field_simps)
   798       from fundamental_theorem_of_algebra[OF nc] have ?case .}
   995           with th0 have "cmod (?x*poly ds ?x) \<noteq> cmod d" by auto
   799   ultimately show ?case by blast
   996           with cth show ?thesis by blast
   800 qed simp
   997         qed
       
   998       qed
       
   999     }
       
  1000     then have nc: "\<not> constant (poly (pCons c cs))" using pCons.prems False
       
  1001       by blast
       
  1002     from fundamental_theorem_of_algebra[OF nc] show ?thesis .
       
  1003   qed
       
  1004 qed
   801 
  1005 
   802 
  1006 
   803 subsection{* Nullstellensatz, degrees and divisibility of polynomials *}
  1007 subsection{* Nullstellensatz, degrees and divisibility of polynomials *}
   804 
  1008 
   805 lemma nullstellensatz_lemma:
  1009 lemma nullstellensatz_lemma:
   814   fix p q :: "complex poly"
  1018   fix p q :: "complex poly"
   815   assume IH: "\<forall>m<n. \<forall>p q.
  1019   assume IH: "\<forall>m<n. \<forall>p q.
   816                  (\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longrightarrow>
  1020                  (\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longrightarrow>
   817                  degree p = m \<longrightarrow> m \<noteq> 0 \<longrightarrow> p dvd (q ^ m)"
  1021                  degree p = m \<longrightarrow> m \<noteq> 0 \<longrightarrow> p dvd (q ^ m)"
   818     and pq0: "\<forall>x. poly p x = 0 \<longrightarrow> poly q x = 0"
  1022     and pq0: "\<forall>x. poly p x = 0 \<longrightarrow> poly q x = 0"
   819     and dpn: "degree p = n" and n0: "n \<noteq> 0"
  1023     and dpn: "degree p = n"
       
  1024     and n0: "n \<noteq> 0"
   820   from dpn n0 have pne: "p \<noteq> 0" by auto
  1025   from dpn n0 have pne: "p \<noteq> 0" by auto
   821   let ?ths = "p dvd (q ^ n)"
  1026   let ?ths = "p dvd (q ^ n)"
   822   {fix a assume a: "poly p a = 0"
  1027   {
   823     {assume oa: "order a p \<noteq> 0"
  1028     fix a
       
  1029     assume a: "poly p a = 0"
       
  1030     {
       
  1031       assume oa: "order a p \<noteq> 0"
   824       let ?op = "order a p"
  1032       let ?op = "order a p"
   825       from pne have ap: "([:- a, 1:] ^ ?op) dvd p"
  1033       from pne have ap: "([:- a, 1:] ^ ?op) dvd p" "\<not> [:- a, 1:] ^ (Suc ?op) dvd p"
   826         "\<not> [:- a, 1:] ^ (Suc ?op) dvd p" using order by blast+
  1034         using order by blast+
   827       note oop = order_degree[OF pne, unfolded dpn]
  1035       note oop = order_degree[OF pne, unfolded dpn]
   828       {assume q0: "q = 0"
  1036       {
   829         hence ?ths using n0
  1037         assume q0: "q = 0"
   830           by (simp add: power_0_left)}
  1038         then have ?ths using n0
       
  1039           by (simp add: power_0_left)
       
  1040       }
   831       moreover
  1041       moreover
   832       {assume q0: "q \<noteq> 0"
  1042       {
       
  1043         assume q0: "q \<noteq> 0"
   833         from pq0[rule_format, OF a, unfolded poly_eq_0_iff_dvd]
  1044         from pq0[rule_format, OF a, unfolded poly_eq_0_iff_dvd]
   834         obtain r where r: "q = [:- a, 1:] * r" by (rule dvdE)
  1045         obtain r where r: "q = [:- a, 1:] * r" by (rule dvdE)
   835         from ap(1) obtain s where
  1046         from ap(1) obtain s where s: "p = [:- a, 1:] ^ ?op * s"
   836           s: "p = [:- a, 1:] ^ ?op * s" by (rule dvdE)
  1047           by (rule dvdE)
   837         have sne: "s \<noteq> 0"
  1048         have sne: "s \<noteq> 0" using s pne by auto
   838           using s pne by auto
  1049         {
   839         {assume ds0: "degree s = 0"
  1050           assume ds0: "degree s = 0"
   840           from ds0 obtain k where kpn: "s = [:k:]"
  1051           from ds0 obtain k where kpn: "s = [:k:]"
   841             by (cases s) (auto split: if_splits)
  1052             by (cases s) (auto split: if_splits)
   842           from sne kpn have k: "k \<noteq> 0" by simp
  1053           from sne kpn have k: "k \<noteq> 0" by simp
   843           let ?w = "([:1/k:] * ([:-a,1:] ^ (n - ?op))) * (r ^ n)"
  1054           let ?w = "([:1/k:] * ([:-a,1:] ^ (n - ?op))) * (r ^ n)"
   844           have "q ^ n = p * ?w"
  1055           have "q ^ n = p * ?w"
   845             apply (subst r, subst s, subst kpn)
  1056             apply (subst r, subst s, subst kpn)
   846             using k oop [of a] 
  1057             using k oop [of a]
   847             apply (subst power_mult_distrib, simp)
  1058             apply (subst power_mult_distrib, simp)
   848             apply (subst power_add [symmetric], simp)
  1059             apply (subst power_add [symmetric], simp)
   849             done
  1060             done
   850           hence ?ths unfolding dvd_def by blast}
  1061           then have ?ths unfolding dvd_def by blast
       
  1062         }
   851         moreover
  1063         moreover
   852         {assume ds0: "degree s \<noteq> 0"
  1064         {
       
  1065           assume ds0: "degree s \<noteq> 0"
   853           from ds0 sne dpn s oa
  1066           from ds0 sne dpn s oa
   854             have dsn: "degree s < n" apply auto
  1067             have dsn: "degree s < n"
       
  1068               apply auto
   855               apply (erule ssubst)
  1069               apply (erule ssubst)
   856               apply (simp add: degree_mult_eq degree_linear_power)
  1070               apply (simp add: degree_mult_eq degree_linear_power)
   857               done
  1071               done
   858             {fix x assume h: "poly s x = 0"
  1072             {
   859               {assume xa: "x = a"
  1073               fix x assume h: "poly s x = 0"
   860                 from h[unfolded xa poly_eq_0_iff_dvd] obtain u where
  1074               {
   861                   u: "s = [:- a, 1:] * u" by (rule dvdE)
  1075                 assume xa: "x = a"
       
  1076                 from h[unfolded xa poly_eq_0_iff_dvd] obtain u where u: "s = [:- a, 1:] * u"
       
  1077                   by (rule dvdE)
   862                 have "p = [:- a, 1:] ^ (Suc ?op) * u"
  1078                 have "p = [:- a, 1:] ^ (Suc ?op) * u"
   863                   by (subst s, subst u, simp only: power_Suc mult_ac)
  1079                   by (subst s, subst u, simp only: power_Suc mult_ac)
   864                 with ap(2)[unfolded dvd_def] have False by blast}
  1080                 with ap(2)[unfolded dvd_def] have False by blast
       
  1081               }
   865               note xa = this
  1082               note xa = this
   866               from h have "poly p x = 0" by (subst s, simp)
  1083               from h have "poly p x = 0" by (subst s) simp
   867               with pq0 have "poly q x = 0" by blast
  1084               with pq0 have "poly q x = 0" by blast
   868               with r xa have "poly r x = 0"
  1085               with r xa have "poly r x = 0"
   869                 by auto}
  1086                 by auto
       
  1087             }
   870             note impth = this
  1088             note impth = this
   871             from IH[rule_format, OF dsn, of s r] impth ds0
  1089             from IH[rule_format, OF dsn, of s r] impth ds0
   872             have "s dvd (r ^ (degree s))" by blast
  1090             have "s dvd (r ^ (degree s))" by blast
   873             then obtain u where u: "r ^ (degree s) = s * u" ..
  1091             then obtain u where u: "r ^ (degree s) = s * u" ..
   874             hence u': "\<And>x. poly s x * poly u x = poly r x ^ degree s"
  1092             then have u': "\<And>x. poly s x * poly u x = poly r x ^ degree s"
   875               by (simp only: poly_mult[symmetric] poly_power[symmetric])
  1093               by (simp only: poly_mult[symmetric] poly_power[symmetric])
   876             let ?w = "(u * ([:-a,1:] ^ (n - ?op))) * (r ^ (n - degree s))"
  1094             let ?w = "(u * ([:-a,1:] ^ (n - ?op))) * (r ^ (n - degree s))"
   877             from oop[of a] dsn have "q ^ n = p * ?w"
  1095             from oop[of a] dsn have "q ^ n = p * ?w"
   878               apply -
  1096               apply -
   879               apply (subst s, subst r)
  1097               apply (subst s, subst r)
   882               apply (subst mult_assoc [where a=u])
  1100               apply (subst mult_assoc [where a=u])
   883               apply (subst mult_assoc [where b=u, symmetric])
  1101               apply (subst mult_assoc [where b=u, symmetric])
   884               apply (subst u [symmetric])
  1102               apply (subst u [symmetric])
   885               apply (simp add: mult_ac power_add [symmetric])
  1103               apply (simp add: mult_ac power_add [symmetric])
   886               done
  1104               done
   887             hence ?ths unfolding dvd_def by blast}
  1105             then have ?ths unfolding dvd_def by blast
   888       ultimately have ?ths by blast }
  1106         }
   889       ultimately have ?ths by blast}
  1107         ultimately have ?ths by blast
   890     then have ?ths using a order_root pne by blast}
  1108       }
       
  1109       ultimately have ?ths by blast
       
  1110     }
       
  1111     then have ?ths using a order_root pne by blast
       
  1112   }
   891   moreover
  1113   moreover
   892   {assume exa: "\<not> (\<exists>a. poly p a = 0)"
  1114   {
   893     from fundamental_theorem_of_algebra_alt[of p] exa obtain c where
  1115     assume exa: "\<not> (\<exists>a. poly p a = 0)"
   894       ccs: "c\<noteq>0" "p = pCons c 0" by blast
  1116     from fundamental_theorem_of_algebra_alt[of p] exa
   895 
  1117     obtain c where ccs: "c \<noteq> 0" "p = pCons c 0"
   896     then have pp: "\<And>x. poly p x =  c" by simp
  1118       by blast
       
  1119     then have pp: "\<And>x. poly p x = c"
       
  1120       by simp
   897     let ?w = "[:1/c:] * (q ^ n)"
  1121     let ?w = "[:1/c:] * (q ^ n)"
   898     from ccs have "(q ^ n) = (p * ?w)" by simp
  1122     from ccs have "(q ^ n) = (p * ?w)"
   899     hence ?ths unfolding dvd_def by blast}
  1123       by simp
       
  1124     then have ?ths
       
  1125       unfolding dvd_def by blast
       
  1126   }
   900   ultimately show ?ths by blast
  1127   ultimately show ?ths by blast
   901 qed
  1128 qed
   902 
  1129 
   903 lemma nullstellensatz_univariate:
  1130 lemma nullstellensatz_univariate:
   904   "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow>
  1131   "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow>
   905     p dvd (q ^ (degree p)) \<or> (p = 0 \<and> q = 0)"
  1132     p dvd (q ^ (degree p)) \<or> (p = 0 \<and> q = 0)"
   906 proof -
  1133 proof -
   907   {assume pe: "p = 0"
  1134   {
   908     hence eq: "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow> q = 0"
  1135     assume pe: "p = 0"
       
  1136     then have eq: "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow> q = 0"
   909       by (auto simp add: poly_all_0_iff_0)
  1137       by (auto simp add: poly_all_0_iff_0)
   910     {assume "p dvd (q ^ (degree p))"
  1138     {
       
  1139       assume "p dvd (q ^ (degree p))"
   911       then obtain r where r: "q ^ (degree p) = p * r" ..
  1140       then obtain r where r: "q ^ (degree p) = p * r" ..
   912       from r pe have False by simp}
  1141       from r pe have False by simp
   913     with eq pe have ?thesis by blast}
  1142     }
       
  1143     with eq pe have ?thesis by blast
       
  1144   }
   914   moreover
  1145   moreover
   915   {assume pe: "p \<noteq> 0"
  1146   {
   916     {assume dp: "degree p = 0"
  1147     assume pe: "p \<noteq> 0"
   917       then obtain k where k: "p = [:k:]" "k\<noteq>0" using pe
  1148     {
       
  1149       assume dp: "degree p = 0"
       
  1150       then obtain k where k: "p = [:k:]" "k \<noteq> 0" using pe
   918         by (cases p) (simp split: if_splits)
  1151         by (cases p) (simp split: if_splits)
   919       hence th1: "\<forall>x. poly p x \<noteq> 0" by simp
  1152       then have th1: "\<forall>x. poly p x \<noteq> 0"
       
  1153         by simp
   920       from k dp have "q ^ (degree p) = p * [:1/k:]"
  1154       from k dp have "q ^ (degree p) = p * [:1/k:]"
   921         by (simp add: one_poly_def)
  1155         by (simp add: one_poly_def)
   922       hence th2: "p dvd (q ^ (degree p))" ..
  1156       then have th2: "p dvd (q ^ (degree p))" ..
   923       from th1 th2 pe have ?thesis by blast}
  1157       from th1 th2 pe have ?thesis by blast
       
  1158     }
   924     moreover
  1159     moreover
   925     {assume dp: "degree p \<noteq> 0"
  1160     {
   926       then obtain n where n: "degree p = Suc n " by (cases "degree p", auto)
  1161       assume dp: "degree p \<noteq> 0"
   927       {assume "p dvd (q ^ (Suc n))"
  1162       then obtain n where n: "degree p = Suc n "
       
  1163         by (cases "degree p") auto
       
  1164       {
       
  1165         assume "p dvd (q ^ (Suc n))"
   928         then obtain u where u: "q ^ (Suc n) = p * u" ..
  1166         then obtain u where u: "q ^ (Suc n) = p * u" ..
   929         {fix x assume h: "poly p x = 0" "poly q x \<noteq> 0"
  1167         {
   930           hence "poly (q ^ (Suc n)) x \<noteq> 0" by simp
  1168           fix x
   931           hence False using u h(1) by (simp only: poly_mult) simp}}
  1169           assume h: "poly p x = 0" "poly q x \<noteq> 0"
   932         with n nullstellensatz_lemma[of p q "degree p"] dp
  1170           then have "poly (q ^ (Suc n)) x \<noteq> 0"
   933         have ?thesis by auto}
  1171             by simp
   934     ultimately have ?thesis by blast}
  1172           then have False using u h(1)
       
  1173             by (simp only: poly_mult) simp
       
  1174         }
       
  1175       }
       
  1176       with n nullstellensatz_lemma[of p q "degree p"] dp
       
  1177       have ?thesis by auto
       
  1178     }
       
  1179     ultimately have ?thesis by blast
       
  1180   }
   935   ultimately show ?thesis by blast
  1181   ultimately show ?thesis by blast
   936 qed
  1182 qed
   937 
  1183 
   938 text{* Useful lemma *}
  1184 text{* Useful lemma *}
   939 
  1185 
   965   by (metis dvd_imp_degree_le pq)
  1211   by (metis dvd_imp_degree_le pq)
   966 
  1212 
   967 (* Arithmetic operations on multivariate polynomials.                        *)
  1213 (* Arithmetic operations on multivariate polynomials.                        *)
   968 
  1214 
   969 lemma mpoly_base_conv:
  1215 lemma mpoly_base_conv:
   970   fixes x :: "'a::comm_ring_1" 
  1216   fixes x :: "'a::comm_ring_1"
   971   shows "0 = poly 0 x" "c = poly [:c:] x" "x = poly [:0,1:] x"
  1217   shows "0 = poly 0 x" "c = poly [:c:] x" "x = poly [:0,1:] x"
   972   by simp_all
  1218   by simp_all
   973 
  1219 
   974 lemma mpoly_norm_conv:
  1220 lemma mpoly_norm_conv:
   975   fixes x :: "'a::comm_ring_1" 
  1221   fixes x :: "'a::comm_ring_1"
   976   shows "poly [:0:] x = poly 0 x" "poly [:poly 0 y:] x = poly 0 x"
  1222   shows "poly [:0:] x = poly 0 x" "poly [:poly 0 y:] x = poly 0 x"
   977   by simp_all
  1223   by simp_all
   978 
  1224 
   979 lemma mpoly_sub_conv:
  1225 lemma mpoly_sub_conv:
   980   fixes x :: "'a::comm_ring_1" 
  1226   fixes x :: "'a::comm_ring_1"
   981   shows "poly p x - poly q x = poly p x + -1 * poly q x"
  1227   shows "poly p x - poly q x = poly p x + -1 * poly q x"
   982   by simp
  1228   by simp
   983 
  1229 
   984 lemma poly_pad_rule: "poly p x = 0 ==> poly (pCons 0 p) x = 0" by simp
  1230 lemma poly_pad_rule: "poly p x = 0 \<Longrightarrow> poly (pCons 0 p) x = 0"
       
  1231   by simp
   985 
  1232 
   986 lemma poly_cancel_eq_conv:
  1233 lemma poly_cancel_eq_conv:
   987   fixes x :: "'a::field" 
  1234   fixes x :: "'a::field"
   988   shows "x = 0 \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> (y = 0) = (a * y - b * x = 0)" 
  1235   shows "x = 0 \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> (y = 0) = (a * y - b * x = 0)"
   989   by auto
  1236   by auto
   990 
  1237 
   991 lemma poly_divides_pad_rule:
  1238 lemma poly_divides_pad_rule:
   992   fixes p:: "('a::comm_ring_1) poly" 
  1239   fixes p:: "('a::comm_ring_1) poly"
   993   assumes pq: "p dvd q"
  1240   assumes pq: "p dvd q"
   994 shows "p dvd (pCons 0 q)"
  1241   shows "p dvd (pCons 0 q)"
   995 proof-
  1242 proof -
   996   have "pCons 0 q = q * [:0,1:]" by simp
  1243   have "pCons 0 q = q * [:0,1:]" by simp
   997   then have "q dvd (pCons 0 q)" ..
  1244   then have "q dvd (pCons 0 q)" ..
   998   with pq show ?thesis by (rule dvd_trans)
  1245   with pq show ?thesis by (rule dvd_trans)
   999 qed
  1246 qed
  1000 
  1247 
  1001 lemma poly_divides_conv0:
  1248 lemma poly_divides_conv0:
  1002   fixes p:: "'a::field poly" 
  1249   fixes p:: "'a::field poly"
  1003   assumes lgpq: "degree q < degree p"
  1250   assumes lgpq: "degree q < degree p"
  1004     and lq: "p \<noteq> 0"
  1251     and lq: "p \<noteq> 0"
  1005   shows "p dvd q \<longleftrightarrow> q = 0" (is "?lhs \<longleftrightarrow> ?rhs")
  1252   shows "p dvd q \<longleftrightarrow> q = 0" (is "?lhs \<longleftrightarrow> ?rhs")
  1006 proof
  1253 proof
  1007   assume r: ?rhs
  1254   assume r: ?rhs
  1008   then have "q = p * 0" by simp
  1255   then have "q = p * 0" by simp
  1009   then show ?lhs ..
  1256   then show ?lhs ..
  1010 next
  1257 next
  1011   assume l: ?lhs
  1258   assume l: ?lhs
  1012   {
  1259   show ?rhs
  1013     assume q0: "q = 0"
  1260   proof (cases "q = 0")
  1014     then have ?rhs by simp
  1261     case True
  1015   }
  1262     then show ?thesis by simp
  1016   moreover
  1263   next
  1017   {
       
  1018     assume q0: "q \<noteq> 0"
  1264     assume q0: "q \<noteq> 0"
  1019     from l q0 have "degree p \<le> degree q"
  1265     from l q0 have "degree p \<le> degree q"
  1020       by (rule dvd_imp_degree_le)
  1266       by (rule dvd_imp_degree_le)
  1021     with lgpq have ?rhs by simp
  1267     with lgpq show ?thesis by simp
  1022   }
  1268   qed
  1023   ultimately show ?rhs by blast
       
  1024 qed
  1269 qed
  1025 
  1270 
  1026 lemma poly_divides_conv1:
  1271 lemma poly_divides_conv1:
  1027   fixes p :: "('a::field) poly" 
  1272   fixes p :: "'a::field poly"
  1028   assumes a0: "a \<noteq> 0"
  1273   assumes a0: "a \<noteq> 0"
  1029     and pp': "p dvd p'"
  1274     and pp': "p dvd p'"
  1030     and qrp': "smult a q - p' = r"
  1275     and qrp': "smult a q - p' = r"
  1031   shows "p dvd q \<longleftrightarrow> p dvd r" (is "?lhs \<longleftrightarrow> ?rhs")
  1276   shows "p dvd q \<longleftrightarrow> p dvd r" (is "?lhs \<longleftrightarrow> ?rhs")
  1032 proof
  1277 proof
  1096   then show "p dvd (q ^ n) \<longleftrightarrow> p dvd r"
  1341   then show "p dvd (q ^ n) \<longleftrightarrow> p dvd r"
  1097     by simp
  1342     by simp
  1098 qed
  1343 qed
  1099 
  1344 
  1100 lemma poly_const_conv:
  1345 lemma poly_const_conv:
  1101   fixes x :: "'a::comm_ring_1" 
  1346   fixes x :: "'a::comm_ring_1"
  1102   shows "poly [:c:] x = y \<longleftrightarrow> c = y"
  1347   shows "poly [:c:] x = y \<longleftrightarrow> c = y"
  1103   by simp
  1348   by simp
  1104 
  1349 
  1105 end
  1350 end