| author | nipkow | 
| Thu, 13 Apr 2023 15:36:07 +1000 | |
| changeset 77830 | 0f2baf04b782 | 
| parent 69597 | ff784d5a5bfb | 
| child 80098 | c06c95576ea9 | 
| permissions | -rw-r--r-- | 
| 33268 
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
33153diff
changeset | 1 | (* Title: HOL/Decision_Procs/Polynomial_List.thy | 
| 
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
33153diff
changeset | 2 | Author: Amine Chaieb | 
| 33153 | 3 | *) | 
| 4 | ||
| 60533 | 5 | section \<open>Univariate Polynomials as lists\<close> | 
| 33153 | 6 | |
| 7 | theory Polynomial_List | |
| 54219 | 8 | imports Complex_Main | 
| 33153 | 9 | begin | 
| 10 | ||
| 60536 | 11 | text \<open>Application of polynomial as a function.\<close> | 
| 33153 | 12 | |
| 54219 | 13 | primrec (in semiring_0) poly :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a" | 
| 52778 | 14 | where | 
| 60536 | 15 | poly_Nil: "poly [] x = 0" | 
| 16 | | poly_Cons: "poly (h # t) x = h + x * poly t x" | |
| 33153 | 17 | |
| 18 | ||
| 60536 | 19 | subsection \<open>Arithmetic Operations on Polynomials\<close> | 
| 33153 | 20 | |
| 60536 | 21 | text \<open>Addition\<close> | 
| 69597 | 22 | primrec (in semiring_0) padd :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixl \<open>+++\<close> 65) | 
| 52778 | 23 | where | 
| 60536 | 24 | padd_Nil: "[] +++ l2 = l2" | 
| 25 | | padd_Cons: "(h # t) +++ l2 = (if l2 = [] then h # t else (h + hd l2) # (t +++ tl l2))" | |
| 33153 | 26 | |
| 60536 | 27 | text \<open>Multiplication by a constant\<close> | 
| 69597 | 28 | primrec (in semiring_0) cmult :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixl \<open>%*\<close> 70) where | 
| 60536 | 29 | cmult_Nil: "c %* [] = []" | 
| 54219 | 30 | | cmult_Cons: "c %* (h#t) = (c * h)#(c %* t)" | 
| 33153 | 31 | |
| 60536 | 32 | text \<open>Multiplication by a polynomial\<close> | 
| 69597 | 33 | primrec (in semiring_0) pmult :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixl \<open>***\<close> 70) | 
| 52778 | 34 | where | 
| 60536 | 35 | pmult_Nil: "[] *** l2 = []" | 
| 36 | | pmult_Cons: "(h # t) *** l2 = (if t = [] then h %* l2 else (h %* l2) +++ (0 # (t *** l2)))" | |
| 33153 | 37 | |
| 60536 | 38 | text \<open>Repeated multiplication by a polynomial\<close> | 
| 39 | primrec (in semiring_0) mulexp :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" | |
| 40 | where | |
| 41 | mulexp_zero: "mulexp 0 p q = q" | |
| 42 | | mulexp_Suc: "mulexp (Suc n) p q = p *** mulexp n p q" | |
| 33153 | 43 | |
| 60536 | 44 | text \<open>Exponential\<close> | 
| 69597 | 45 | primrec (in semiring_1) pexp :: "'a list \<Rightarrow> nat \<Rightarrow> 'a list" (infixl \<open>%^\<close> 80) | 
| 60536 | 46 | where | 
| 47 | pexp_0: "p %^ 0 = [1]" | |
| 39246 | 48 | | pexp_Suc: "p %^ (Suc n) = p *** (p %^ n)" | 
| 33153 | 49 | |
| 60536 | 50 | text \<open>Quotient related value of dividing a polynomial by x + a. | 
| 51 | Useful for divisor properties in inductive proofs.\<close> | |
| 54219 | 52 | primrec (in field) "pquot" :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a list" | 
| 52778 | 53 | where | 
| 60536 | 54 | pquot_Nil: "pquot [] a = []" | 
| 55 | | pquot_Cons: "pquot (h # t) a = | |
| 56 | (if t = [] then [h] else (inverse a * (h - hd( pquot t a))) # pquot t a)" | |
| 33153 | 57 | |
| 60536 | 58 | text \<open>Normalization of polynomials (remove extra 0 coeff).\<close> | 
| 59 | primrec (in semiring_0) pnormalize :: "'a list \<Rightarrow> 'a list" | |
| 60 | where | |
| 61 | pnormalize_Nil: "pnormalize [] = []" | |
| 62 | | pnormalize_Cons: "pnormalize (h # p) = | |
| 54219 | 63 | (if pnormalize p = [] then (if h = 0 then [] else [h]) else h # pnormalize p)" | 
| 33153 | 64 | |
| 60536 | 65 | definition (in semiring_0) "pnormal p \<longleftrightarrow> pnormalize p = p \<and> p \<noteq> []" | 
| 66 | definition (in semiring_0) "nonconstant p \<longleftrightarrow> pnormal p \<and> (\<forall>x. p \<noteq> [x])" | |
| 67 | ||
| 68 | text \<open>Other definitions.\<close> | |
| 33153 | 69 | |
| 69597 | 70 | definition (in ring_1) poly_minus :: "'a list \<Rightarrow> 'a list" (\<open>-- _\<close> [80] 80) | 
| 52778 | 71 | where "-- p = (- 1) %* p" | 
| 33153 | 72 | |
| 69597 | 73 | definition (in semiring_0) divides :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infixl \<open>divides\<close> 70) | 
| 60536 | 74 | where "p1 divides p2 \<longleftrightarrow> (\<exists>q. poly p2 = poly(p1 *** q))" | 
| 54219 | 75 | |
| 60536 | 76 | lemma (in semiring_0) dividesI: "poly p2 = poly (p1 *** q) \<Longrightarrow> p1 divides p2" | 
| 54219 | 77 | by (auto simp add: divides_def) | 
| 33153 | 78 | |
| 54219 | 79 | lemma (in semiring_0) dividesE: | 
| 80 | assumes "p1 divides p2" | |
| 81 | obtains q where "poly p2 = poly (p1 *** q)" | |
| 82 | using assms by (auto simp add: divides_def) | |
| 33153 | 83 | |
| 61586 | 84 | \<comment> \<open>order of a polynomial\<close> | 
| 60536 | 85 | definition (in ring_1) order :: "'a \<Rightarrow> 'a list \<Rightarrow> nat" | 
| 86 | where "order a p = (SOME n. ([-a, 1] %^ n) divides p \<and> \<not> (([-a, 1] %^ (Suc n)) divides p))" | |
| 54219 | 87 | |
| 61586 | 88 | \<comment> \<open>degree of a polynomial\<close> | 
| 54219 | 89 | definition (in semiring_0) degree :: "'a list \<Rightarrow> nat" | 
| 52778 | 90 | where "degree p = length (pnormalize p) - 1" | 
| 33153 | 91 | |
| 61586 | 92 | \<comment> \<open>squarefree polynomials --- NB with respect to real roots only\<close> | 
| 54219 | 93 | definition (in ring_1) rsquarefree :: "'a list \<Rightarrow> bool" | 
| 94 | where "rsquarefree p \<longleftrightarrow> poly p \<noteq> poly [] \<and> (\<forall>a. order a p = 0 \<or> order a p = 1)" | |
| 33153 | 95 | |
| 54219 | 96 | context semiring_0 | 
| 97 | begin | |
| 98 | ||
| 99 | lemma padd_Nil2[simp]: "p +++ [] = p" | |
| 52778 | 100 | by (induct p) auto | 
| 33153 | 101 | |
| 102 | lemma padd_Cons_Cons: "(h1 # p1) +++ (h2 # p2) = (h1 + h2) # (p1 +++ p2)" | |
| 52778 | 103 | by auto | 
| 33153 | 104 | |
| 54219 | 105 | lemma pminus_Nil: "-- [] = []" | 
| 52778 | 106 | by (simp add: poly_minus_def) | 
| 33153 | 107 | |
| 54219 | 108 | lemma pmult_singleton: "[h1] *** p1 = h1 %* p1" by simp | 
| 109 | ||
| 110 | end | |
| 33153 | 111 | |
| 60536 | 112 | lemma (in semiring_1) poly_ident_mult[simp]: "1 %* t = t" | 
| 113 | by (induct t) auto | |
| 33153 | 114 | |
| 60536 | 115 | lemma (in semiring_0) poly_simple_add_Cons[simp]: "[a] +++ (0 # t) = a # t" | 
| 52778 | 116 | by simp | 
| 33153 | 117 | |
| 60536 | 118 | |
| 119 | text \<open>Handy general properties.\<close> | |
| 33153 | 120 | |
| 54219 | 121 | lemma (in comm_semiring_0) padd_commut: "b +++ a = a +++ b" | 
| 122 | proof (induct b arbitrary: a) | |
| 123 | case Nil | |
| 60536 | 124 | then show ?case | 
| 125 | by auto | |
| 54219 | 126 | next | 
| 127 | case (Cons b bs a) | |
| 60536 | 128 | then show ?case | 
| 129 | by (cases a) (simp_all add: add.commute) | |
| 54219 | 130 | qed | 
| 131 | ||
| 60698 | 132 | lemma (in comm_semiring_0) padd_assoc: "(a +++ b) +++ c = a +++ (b +++ c)" | 
| 133 | proof (induct a arbitrary: b c) | |
| 134 | case Nil | |
| 135 | then show ?case | |
| 136 | by simp | |
| 137 | next | |
| 138 | case Cons | |
| 139 | then show ?case | |
| 140 | by (cases b) (simp_all add: ac_simps) | |
| 141 | qed | |
| 33153 | 142 | |
| 60536 | 143 | lemma (in semiring_0) poly_cmult_distr: "a %* (p +++ q) = a %* p +++ a %* q" | 
| 60698 | 144 | proof (induct p arbitrary: q) | 
| 145 | case Nil | |
| 146 | then show ?case | |
| 147 | by simp | |
| 148 | next | |
| 149 | case Cons | |
| 150 | then show ?case | |
| 151 | by (cases q) (simp_all add: distrib_left) | |
| 152 | qed | |
| 33153 | 153 | |
| 60536 | 154 | lemma (in ring_1) pmult_by_x[simp]: "[0, 1] *** t = 0 # t" | 
| 60698 | 155 | proof (induct t) | 
| 156 | case Nil | |
| 157 | then show ?case | |
| 158 | by simp | |
| 159 | next | |
| 160 | case (Cons a t) | |
| 161 | then show ?case | |
| 162 | by (cases t) (auto simp add: padd_commut) | |
| 163 | qed | |
| 60536 | 164 | |
| 165 | text \<open>Properties of evaluation of polynomials.\<close> | |
| 33153 | 166 | |
| 54219 | 167 | lemma (in semiring_0) poly_add: "poly (p1 +++ p2) x = poly p1 x + poly p2 x" | 
| 60536 | 168 | proof (induct p1 arbitrary: p2) | 
| 54219 | 169 | case Nil | 
| 60536 | 170 | then show ?case | 
| 171 | by simp | |
| 54219 | 172 | next | 
| 173 | case (Cons a as p2) | |
| 60536 | 174 | then show ?case | 
| 175 | by (cases p2) (simp_all add: ac_simps distrib_left) | |
| 54219 | 176 | qed | 
| 33153 | 177 | |
| 54219 | 178 | lemma (in comm_semiring_0) poly_cmult: "poly (c %* p) x = c * poly p x" | 
| 60698 | 179 | proof (induct p) | 
| 180 | case Nil | |
| 181 | then show ?case | |
| 182 | by simp | |
| 183 | next | |
| 184 | case Cons | |
| 185 | then show ?case | |
| 186 | by (cases "x = zero") (auto simp add: distrib_left ac_simps) | |
| 187 | qed | |
| 33153 | 188 | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
67399diff
changeset | 189 | lemma (in comm_semiring_0) poly_cmult_map: "poly (map ((*) c) p) x = c * poly p x" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 190 | by (induct p) (auto simp add: distrib_left ac_simps) | 
| 33153 | 191 | |
| 54219 | 192 | lemma (in comm_ring_1) poly_minus: "poly (-- p) x = - (poly p x)" | 
| 60698 | 193 | by (simp add: poly_minus_def) (auto simp add: poly_cmult) | 
| 33153 | 194 | |
| 54219 | 195 | lemma (in comm_semiring_0) poly_mult: "poly (p1 *** p2) x = poly p1 x * poly p2 x" | 
| 196 | proof (induct p1 arbitrary: p2) | |
| 197 | case Nil | |
| 60536 | 198 | then show ?case | 
| 199 | by simp | |
| 54219 | 200 | next | 
| 60698 | 201 | case (Cons a as) | 
| 60536 | 202 | then show ?case | 
| 203 | by (cases as) (simp_all add: poly_cmult poly_add distrib_right distrib_left ac_simps) | |
| 54219 | 204 | qed | 
| 205 | ||
| 206 | class idom_char_0 = idom + ring_char_0 | |
| 207 | ||
| 208 | subclass (in field_char_0) idom_char_0 .. | |
| 209 | ||
| 210 | lemma (in comm_ring_1) poly_exp: "poly (p %^ n) x = (poly p x) ^ n" | |
| 52881 | 211 | by (induct n) (auto simp add: poly_cmult poly_mult) | 
| 33153 | 212 | |
| 60536 | 213 | |
| 214 | text \<open>More Polynomial Evaluation lemmas.\<close> | |
| 33153 | 215 | |
| 54219 | 216 | lemma (in semiring_0) poly_add_rzero[simp]: "poly (a +++ []) x = poly a x" | 
| 52778 | 217 | by simp | 
| 33153 | 218 | |
| 54219 | 219 | lemma (in comm_semiring_0) poly_mult_assoc: "poly ((a *** b) *** c) x = poly (a *** (b *** c)) x" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
55417diff
changeset | 220 | by (simp add: poly_mult mult.assoc) | 
| 33153 | 221 | |
| 54219 | 222 | lemma (in semiring_0) poly_mult_Nil2[simp]: "poly (p *** []) x = 0" | 
| 52881 | 223 | by (induct p) auto | 
| 33153 | 224 | |
| 60536 | 225 | lemma (in comm_semiring_1) poly_exp_add: "poly (p %^ (n + d)) x = poly (p %^ n *** p %^ d) x" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
55417diff
changeset | 226 | by (induct n) (auto simp add: poly_mult mult.assoc) | 
| 33153 | 227 | |
| 60536 | 228 | |
| 69597 | 229 | subsection \<open>Key Property: if \<^term>\<open>f a = 0\<close> then \<^term>\<open>(x - a)\<close> divides \<^term>\<open>p(x)\<close>.\<close> | 
| 33153 | 230 | |
| 60698 | 231 | lemma (in comm_ring_1) lemma_poly_linear_rem: "\<exists>q r. h#t = [r] +++ [-a, 1] *** q" | 
| 232 | proof (induct t arbitrary: h) | |
| 54219 | 233 | case Nil | 
| 60698 | 234 | have "[h] = [h] +++ [- a, 1] *** []" by simp | 
| 60536 | 235 | then show ?case by blast | 
| 54219 | 236 | next | 
| 237 | case (Cons x xs) | |
| 60698 | 238 | have "\<exists>q r. h # x # xs = [r] +++ [-a, 1] *** q" | 
| 60536 | 239 | proof - | 
| 60698 | 240 | from Cons obtain q r where qr: "x # xs = [r] +++ [- a, 1] *** q" | 
| 60536 | 241 | by blast | 
| 242 | have "h # x # xs = [a * r + h] +++ [-a, 1] *** (r # q)" | |
| 54219 | 243 | using qr by (cases q) (simp_all add: algebra_simps) | 
| 60536 | 244 | then show ?thesis by blast | 
| 245 | qed | |
| 246 | then show ?case by blast | |
| 54219 | 247 | qed | 
| 248 | ||
| 249 | lemma (in comm_ring_1) poly_linear_rem: "\<exists>q r. h#t = [r] +++ [-a, 1] *** q" | |
| 250 | using lemma_poly_linear_rem [where t = t and a = a] by auto | |
| 251 | ||
| 60698 | 252 | lemma (in comm_ring_1) poly_linear_divides: "poly p a = 0 \<longleftrightarrow> p = [] \<or> (\<exists>q. p = [-a, 1] *** q)" | 
| 60536 | 253 | proof (cases p) | 
| 254 | case Nil | |
| 255 | then show ?thesis by simp | |
| 256 | next | |
| 257 | case (Cons x xs) | |
| 258 | have "poly p a = 0" if "p = [-a, 1] *** q" for q | |
| 259 | using that by (simp add: poly_add poly_cmult) | |
| 54219 | 260 | moreover | 
| 60536 | 261 | have "\<exists>q. p = [- a, 1] *** q" if p0: "poly p a = 0" | 
| 262 | proof - | |
| 263 | from poly_linear_rem[of x xs a] obtain q r where qr: "x#xs = [r] +++ [- a, 1] *** q" | |
| 264 | by blast | |
| 265 | have "r = 0" | |
| 266 | using p0 by (simp only: Cons qr poly_mult poly_add) simp | |
| 267 | with Cons qr show ?thesis | |
| 268 | apply - | |
| 269 | apply (rule exI[where x = q]) | |
| 270 | apply auto | |
| 271 | apply (cases q) | |
| 272 | apply auto | |
| 273 | done | |
| 274 | qed | |
| 275 | ultimately show ?thesis using Cons by blast | |
| 54219 | 276 | qed | 
| 33153 | 277 | |
| 60536 | 278 | lemma (in semiring_0) lemma_poly_length_mult[simp]: | 
| 60698 | 279 | "length (k %* p +++ (h # (a %* p))) = Suc (length p)" | 
| 280 | by (induct p arbitrary: h k a) auto | |
| 33153 | 281 | |
| 60536 | 282 | lemma (in semiring_0) lemma_poly_length_mult2[simp]: | 
| 60698 | 283 | "length (k %* p +++ (h # p)) = Suc (length p)" | 
| 284 | by (induct p arbitrary: h k) auto | |
| 33153 | 285 | |
| 54219 | 286 | lemma (in ring_1) poly_length_mult[simp]: "length([-a,1] *** q) = Suc (length q)" | 
| 52778 | 287 | by auto | 
| 33153 | 288 | |
| 60536 | 289 | |
| 290 | subsection \<open>Polynomial length\<close> | |
| 33153 | 291 | |
| 54219 | 292 | lemma (in semiring_0) poly_cmult_length[simp]: "length (a %* p) = length p" | 
| 52778 | 293 | by (induct p) auto | 
| 33153 | 294 | |
| 54219 | 295 | lemma (in semiring_0) poly_add_length: "length (p1 +++ p2) = max (length p1) (length p2)" | 
| 60698 | 296 | by (induct p1 arbitrary: p2) auto | 
| 33153 | 297 | |
| 60698 | 298 | lemma (in semiring_0) poly_root_mult_length[simp]: "length ([a, b] *** p) = Suc (length p)" | 
| 54219 | 299 | by (simp add: poly_add_length) | 
| 33153 | 300 | |
| 54219 | 301 | lemma (in idom) poly_mult_not_eq_poly_Nil[simp]: | 
| 302 | "poly (p *** q) x \<noteq> poly [] x \<longleftrightarrow> poly p x \<noteq> poly [] x \<and> poly q x \<noteq> poly [] x" | |
| 52778 | 303 | by (auto simp add: poly_mult) | 
| 33153 | 304 | |
| 54219 | 305 | lemma (in idom) poly_mult_eq_zero_disj: "poly (p *** q) x = 0 \<longleftrightarrow> poly p x = 0 \<or> poly q x = 0" | 
| 52778 | 306 | by (auto simp add: poly_mult) | 
| 33153 | 307 | |
| 308 | ||
| 60536 | 309 | text \<open>Normalisation Properties.\<close> | 
| 310 | ||
| 311 | lemma (in semiring_0) poly_normalized_nil: "pnormalize p = [] \<longrightarrow> poly p x = 0" | |
| 52778 | 312 | by (induct p) auto | 
| 33153 | 313 | |
| 60536 | 314 | text \<open>A nontrivial polynomial of degree n has no more than n roots.\<close> | 
| 54219 | 315 | lemma (in idom) poly_roots_index_lemma: | 
| 60698 | 316 | assumes "poly p x \<noteq> poly [] x" | 
| 317 | and "length p = n" | |
| 54219 | 318 | shows "\<exists>i. \<forall>x. poly p x = 0 \<longrightarrow> (\<exists>m\<le>n. x = i m)" | 
| 60698 | 319 | using assms | 
| 54219 | 320 | proof (induct n arbitrary: p x) | 
| 321 | case 0 | |
| 60536 | 322 | then show ?case by simp | 
| 54219 | 323 | next | 
| 60698 | 324 | case (Suc n) | 
| 60536 | 325 | have False if C: "\<And>i. \<exists>x. poly p x = 0 \<and> (\<forall>m\<le>Suc n. x \<noteq> i m)" | 
| 326 | proof - | |
| 327 | from Suc.prems have p0: "poly p x \<noteq> 0" "p \<noteq> []" | |
| 328 | by auto | |
| 54219 | 329 | from p0(1)[unfolded poly_linear_divides[of p x]] | 
| 60536 | 330 | have "\<forall>q. p \<noteq> [- x, 1] *** q" | 
| 331 | by blast | |
| 332 | from C obtain a where a: "poly p a = 0" | |
| 333 | by blast | |
| 334 | from a[unfolded poly_linear_divides[of p a]] p0(2) obtain q where q: "p = [-a, 1] *** q" | |
| 335 | by blast | |
| 336 | have lg: "length q = n" | |
| 337 | using q Suc.prems(2) by simp | |
| 54219 | 338 | from q p0 have qx: "poly q x \<noteq> poly [] x" | 
| 339 | by (auto simp add: poly_mult poly_add poly_cmult) | |
| 60698 | 340 | from Suc.hyps[OF qx lg] obtain i where i: "\<And>x. poly q x = 0 \<longrightarrow> (\<exists>m\<le>n. x = i m)" | 
| 60536 | 341 | by blast | 
| 54219 | 342 | let ?i = "\<lambda>m. if m = Suc n then a else i m" | 
| 343 | from C[of ?i] obtain y where y: "poly p y = 0" "\<forall>m\<le> Suc n. y \<noteq> ?i m" | |
| 344 | by blast | |
| 345 | from y have "y = a \<or> poly q y = 0" | |
| 346 | by (simp only: q poly_mult_eq_zero_disj poly_add) (simp add: algebra_simps) | |
| 60698 | 347 | with i[of y] y(1) y(2) show ?thesis | 
| 54219 | 348 | apply auto | 
| 349 | apply (erule_tac x = "m" in allE) | |
| 350 | apply auto | |
| 351 | done | |
| 60536 | 352 | qed | 
| 353 | then show ?case by blast | |
| 54219 | 354 | qed | 
| 33153 | 355 | |
| 356 | ||
| 54219 | 357 | lemma (in idom) poly_roots_index_length: | 
| 60698 | 358 | "poly p x \<noteq> poly [] x \<Longrightarrow> \<exists>i. \<forall>x. poly p x = 0 \<longrightarrow> (\<exists>n. n \<le> length p \<and> x = i n)" | 
| 54219 | 359 | by (blast intro: poly_roots_index_lemma) | 
| 33153 | 360 | |
| 54219 | 361 | lemma (in idom) poly_roots_finite_lemma1: | 
| 60698 | 362 | "poly p x \<noteq> poly [] x \<Longrightarrow> \<exists>N i. \<forall>x. poly p x = 0 \<longrightarrow> (\<exists>n::nat. n < N \<and> x = i n)" | 
| 363 | apply (drule poly_roots_index_length) | |
| 364 | apply safe | |
| 52778 | 365 | apply (rule_tac x = "Suc (length p)" in exI) | 
| 366 | apply (rule_tac x = i in exI) | |
| 367 | apply (simp add: less_Suc_eq_le) | |
| 368 | done | |
| 33153 | 369 | |
| 54219 | 370 | lemma (in idom) idom_finite_lemma: | 
| 60536 | 371 | assumes "\<forall>x. P x \<longrightarrow> (\<exists>n. n < length j \<and> x = j!n)" | 
| 54219 | 372 |   shows "finite {x. P x}"
 | 
| 52778 | 373 | proof - | 
| 60698 | 374 |   from assms have "{x. P x} \<subseteq> set j"
 | 
| 375 | by auto | |
| 376 | then show ?thesis | |
| 377 | using finite_subset by auto | |
| 33153 | 378 | qed | 
| 379 | ||
| 54219 | 380 | lemma (in idom) poly_roots_finite_lemma2: | 
| 381 | "poly p x \<noteq> poly [] x \<Longrightarrow> \<exists>i. \<forall>x. poly p x = 0 \<longrightarrow> x \<in> set i" | |
| 60536 | 382 | apply (drule poly_roots_index_length) | 
| 383 | apply safe | |
| 384 | apply (rule_tac x = "map (\<lambda>n. i n) [0 ..< Suc (length p)]" in exI) | |
| 54219 | 385 | apply (auto simp add: image_iff) | 
| 60536 | 386 | apply (erule_tac x="x" in allE) | 
| 387 | apply clarsimp | |
| 54219 | 388 | apply (case_tac "n = length p") | 
| 389 | apply (auto simp add: order_le_less) | |
| 52778 | 390 | done | 
| 33153 | 391 | |
| 60536 | 392 | lemma (in ring_char_0) UNIV_ring_char_0_infinte: "\<not> finite (UNIV :: 'a set)" | 
| 54219 | 393 | proof | 
| 394 | assume F: "finite (UNIV :: 'a set)" | |
| 395 | have "finite (UNIV :: nat set)" | |
| 396 | proof (rule finite_imageD) | |
| 60698 | 397 | have "of_nat ` UNIV \<subseteq> UNIV" | 
| 398 | by simp | |
| 60536 | 399 | then show "finite (of_nat ` UNIV :: 'a set)" | 
| 400 | using F by (rule finite_subset) | |
| 401 | show "inj (of_nat :: nat \<Rightarrow> 'a)" | |
| 402 | by (simp add: inj_on_def) | |
| 54219 | 403 | qed | 
| 404 | with infinite_UNIV_nat show False .. | |
| 33153 | 405 | qed | 
| 406 | ||
| 54219 | 407 | lemma (in idom_char_0) poly_roots_finite: "poly p \<noteq> poly [] \<longleftrightarrow> finite {x. poly p x = 0}"
 | 
| 60536 | 408 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 33153 | 409 | proof | 
| 60536 | 410 | show ?rhs if ?lhs | 
| 411 | using that | |
| 33153 | 412 | apply - | 
| 60536 | 413 | apply (erule contrapos_np) | 
| 414 | apply (rule ext) | |
| 33153 | 415 | apply (rule ccontr) | 
| 54219 | 416 | apply (clarify dest!: poly_roots_finite_lemma2) | 
| 33153 | 417 | using finite_subset | 
| 52778 | 418 | proof - | 
| 33153 | 419 | fix x i | 
| 60536 | 420 |     assume F: "\<not> finite {x. poly p x = 0}"
 | 
| 421 | and P: "\<forall>x. poly p x = 0 \<longrightarrow> x \<in> set i" | |
| 422 |     from P have "{x. poly p x = 0} \<subseteq> set i"
 | |
| 423 | by auto | |
| 424 | with finite_subset F show False | |
| 425 | by auto | |
| 33153 | 426 | qed | 
| 60536 | 427 | show ?lhs if ?rhs | 
| 428 | using UNIV_ring_char_0_infinte that by auto | |
| 33153 | 429 | qed | 
| 430 | ||
| 60536 | 431 | |
| 432 | text \<open>Entirety and Cancellation for polynomials\<close> | |
| 33153 | 433 | |
| 54219 | 434 | lemma (in idom_char_0) poly_entire_lemma2: | 
| 435 | assumes p0: "poly p \<noteq> poly []" | |
| 436 | and q0: "poly q \<noteq> poly []" | |
| 437 | shows "poly (p***q) \<noteq> poly []" | |
| 438 | proof - | |
| 439 |   let ?S = "\<lambda>p. {x. poly p x = 0}"
 | |
| 60536 | 440 | have "?S (p *** q) = ?S p \<union> ?S q" | 
| 441 | by (auto simp add: poly_mult) | |
| 442 | with p0 q0 show ?thesis | |
| 443 | unfolding poly_roots_finite by auto | |
| 54219 | 444 | qed | 
| 33153 | 445 | |
| 54219 | 446 | lemma (in idom_char_0) poly_entire: | 
| 447 | "poly (p *** q) = poly [] \<longleftrightarrow> poly p = poly [] \<or> poly q = poly []" | |
| 448 | using poly_entire_lemma2[of p q] | |
| 449 | by (auto simp add: fun_eq_iff poly_mult) | |
| 33153 | 450 | |
| 54219 | 451 | lemma (in idom_char_0) poly_entire_neg: | 
| 452 | "poly (p *** q) \<noteq> poly [] \<longleftrightarrow> poly p \<noteq> poly [] \<and> poly q \<noteq> poly []" | |
| 52778 | 453 | by (simp add: poly_entire) | 
| 33153 | 454 | |
| 54219 | 455 | lemma (in comm_ring_1) poly_add_minus_zero_iff: | 
| 456 | "poly (p +++ -- q) = poly [] \<longleftrightarrow> poly p = poly q" | |
| 60536 | 457 | by (auto simp add: algebra_simps poly_add poly_minus_def fun_eq_iff poly_cmult) | 
| 33153 | 458 | |
| 54219 | 459 | lemma (in comm_ring_1) poly_add_minus_mult_eq: | 
| 460 | "poly (p *** q +++ --(p *** r)) = poly (p *** (q +++ -- r))" | |
| 60536 | 461 | by (auto simp add: poly_add poly_minus_def fun_eq_iff poly_mult poly_cmult algebra_simps) | 
| 33153 | 462 | |
| 54219 | 463 | subclass (in idom_char_0) comm_ring_1 .. | 
| 33153 | 464 | |
| 54219 | 465 | lemma (in idom_char_0) poly_mult_left_cancel: | 
| 466 | "poly (p *** q) = poly (p *** r) \<longleftrightarrow> poly p = poly [] \<or> poly q = poly r" | |
| 467 | proof - | |
| 468 | have "poly (p *** q) = poly (p *** r) \<longleftrightarrow> poly (p *** q +++ -- (p *** r)) = poly []" | |
| 469 | by (simp only: poly_add_minus_zero_iff) | |
| 470 | also have "\<dots> \<longleftrightarrow> poly p = poly [] \<or> poly q = poly r" | |
| 471 | by (auto intro: simp add: poly_add_minus_mult_eq poly_entire poly_add_minus_zero_iff) | |
| 472 | finally show ?thesis . | |
| 473 | qed | |
| 474 | ||
| 60536 | 475 | lemma (in idom) poly_exp_eq_zero[simp]: "poly (p %^ n) = poly [] \<longleftrightarrow> poly p = poly [] \<and> n \<noteq> 0" | 
| 476 | apply (simp only: fun_eq_iff add: HOL.all_simps [symmetric]) | |
| 52778 | 477 | apply (rule arg_cong [where f = All]) | 
| 478 | apply (rule ext) | |
| 54219 | 479 | apply (induct n) | 
| 480 | apply (auto simp add: poly_exp poly_mult) | |
| 52778 | 481 | done | 
| 33153 | 482 | |
| 60536 | 483 | lemma (in comm_ring_1) poly_prime_eq_zero[simp]: "poly [a, 1] \<noteq> poly []" | 
| 484 | apply (simp add: fun_eq_iff) | |
| 54219 | 485 | apply (rule_tac x = "minus one a" in exI) | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
55417diff
changeset | 486 | apply (simp add: add.commute [of a]) | 
| 52778 | 487 | done | 
| 33153 | 488 | |
| 54219 | 489 | lemma (in idom) poly_exp_prime_eq_zero: "poly ([a, 1] %^ n) \<noteq> poly []" | 
| 52778 | 490 | by auto | 
| 33153 | 491 | |
| 60536 | 492 | |
| 493 | text \<open>A more constructive notion of polynomials being trivial.\<close> | |
| 33153 | 494 | |
| 54219 | 495 | lemma (in idom_char_0) poly_zero_lemma': "poly (h # t) = poly [] \<Longrightarrow> h = 0 \<and> poly t = poly []" | 
| 60536 | 496 | apply (simp add: fun_eq_iff) | 
| 54219 | 497 | apply (case_tac "h = zero") | 
| 60536 | 498 | apply (drule_tac [2] x = zero in spec) | 
| 499 | apply auto | |
| 500 | apply (cases "poly t = poly []") | |
| 501 | apply simp | |
| 52778 | 502 | proof - | 
| 33153 | 503 | fix x | 
| 60536 | 504 | assume H: "\<forall>x. x = 0 \<or> poly t x = 0" | 
| 505 | assume pnz: "poly t \<noteq> poly []" | |
| 33153 | 506 |   let ?S = "{x. poly t x = 0}"
 | 
| 60536 | 507 | from H have "\<forall>x. x \<noteq> 0 \<longrightarrow> poly t x = 0" | 
| 508 | by blast | |
| 509 |   then have th: "?S \<supseteq> UNIV - {0}"
 | |
| 510 | by auto | |
| 511 | from poly_roots_finite pnz have th': "finite ?S" | |
| 512 | by blast | |
| 513 | from finite_subset[OF th th'] UNIV_ring_char_0_infinte show "poly t x = 0" | |
| 54219 | 514 | by simp | 
| 52778 | 515 | qed | 
| 33153 | 516 | |
| 60537 | 517 | lemma (in idom_char_0) poly_zero: "poly p = poly [] \<longleftrightarrow> (\<forall>c \<in> set p. c = 0)" | 
| 60698 | 518 | proof (induct p) | 
| 519 | case Nil | |
| 520 | then show ?case by simp | |
| 521 | next | |
| 522 | case Cons | |
| 523 | show ?case | |
| 524 | apply (rule iffI) | |
| 525 | apply (drule poly_zero_lemma') | |
| 526 | using Cons | |
| 527 | apply auto | |
| 528 | done | |
| 529 | qed | |
| 33153 | 530 | |
| 60537 | 531 | lemma (in idom_char_0) poly_0: "\<forall>c \<in> set p. c = 0 \<Longrightarrow> poly p x = 0" | 
| 54219 | 532 | unfolding poly_zero[symmetric] by simp | 
| 533 | ||
| 534 | ||
| 60536 | 535 | text \<open>Basics of divisibility.\<close> | 
| 33153 | 536 | |
| 60536 | 537 | lemma (in idom) poly_primes: "[a, 1] divides (p *** q) \<longleftrightarrow> [a, 1] divides p \<or> [a, 1] divides q" | 
| 538 | apply (auto simp add: divides_def fun_eq_iff poly_mult poly_add poly_cmult distrib_right [symmetric]) | |
| 54219 | 539 | apply (drule_tac x = "uminus a" in spec) | 
| 540 | apply (simp add: poly_linear_divides poly_add poly_cmult distrib_right [symmetric]) | |
| 541 | apply (cases "p = []") | |
| 542 | apply (rule exI[where x="[]"]) | |
| 543 | apply simp | |
| 544 | apply (cases "q = []") | |
| 60536 | 545 | apply (erule allE[where x="[]"]) | 
| 546 | apply simp | |
| 54219 | 547 | |
| 548 | apply clarsimp | |
| 60536 | 549 | apply (cases "\<exists>q. p = a %* q +++ (0 # q)") | 
| 54219 | 550 | apply (clarsimp simp add: poly_add poly_cmult) | 
| 60536 | 551 | apply (rule_tac x = qa in exI) | 
| 54219 | 552 | apply (simp add: distrib_right [symmetric]) | 
| 553 | apply clarsimp | |
| 554 | ||
| 52778 | 555 | apply (auto simp add: poly_linear_divides poly_add poly_cmult distrib_right [symmetric]) | 
| 54219 | 556 | apply (rule_tac x = "pmult qa q" in exI) | 
| 557 | apply (rule_tac [2] x = "pmult p qa" in exI) | |
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 558 | apply (auto simp add: poly_add poly_mult poly_cmult ac_simps) | 
| 52778 | 559 | done | 
| 33153 | 560 | |
| 54219 | 561 | lemma (in comm_semiring_1) poly_divides_refl[simp]: "p divides p" | 
| 52778 | 562 | apply (simp add: divides_def) | 
| 54219 | 563 | apply (rule_tac x = "[one]" in exI) | 
| 60536 | 564 | apply (auto simp add: poly_mult fun_eq_iff) | 
| 52778 | 565 | done | 
| 33153 | 566 | |
| 54219 | 567 | lemma (in comm_semiring_1) poly_divides_trans: "p divides q \<Longrightarrow> q divides r \<Longrightarrow> p divides r" | 
| 60536 | 568 | apply (simp add: divides_def) | 
| 569 | apply safe | |
| 54219 | 570 | apply (rule_tac x = "pmult qa qaa" in exI) | 
| 60536 | 571 | apply (auto simp add: poly_mult fun_eq_iff mult.assoc) | 
| 52778 | 572 | done | 
| 33153 | 573 | |
| 54219 | 574 | lemma (in comm_semiring_1) poly_divides_exp: "m \<le> n \<Longrightarrow> (p %^ m) divides (p %^ n)" | 
| 62378 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
61945diff
changeset | 575 | by (auto simp: le_iff_add divides_def poly_exp_add fun_eq_iff) | 
| 33153 | 576 | |
| 60536 | 577 | lemma (in comm_semiring_1) poly_exp_divides: "(p %^ n) divides q \<Longrightarrow> m \<le> n \<Longrightarrow> (p %^ m) divides q" | 
| 52778 | 578 | by (blast intro: poly_divides_exp poly_divides_trans) | 
| 33153 | 579 | |
| 60536 | 580 | lemma (in comm_semiring_0) poly_divides_add: "p divides q \<Longrightarrow> p divides r \<Longrightarrow> p divides (q +++ r)" | 
| 581 | apply (auto simp add: divides_def) | |
| 54219 | 582 | apply (rule_tac x = "padd qa qaa" in exI) | 
| 60536 | 583 | apply (auto simp add: poly_add fun_eq_iff poly_mult distrib_left) | 
| 52778 | 584 | done | 
| 33153 | 585 | |
| 60536 | 586 | lemma (in comm_ring_1) poly_divides_diff: "p divides q \<Longrightarrow> p divides (q +++ r) \<Longrightarrow> p divides r" | 
| 587 | apply (auto simp add: divides_def) | |
| 54219 | 588 | apply (rule_tac x = "padd qaa (poly_minus qa)" in exI) | 
| 60536 | 589 | apply (auto simp add: poly_add fun_eq_iff poly_mult poly_minus algebra_simps) | 
| 52778 | 590 | done | 
| 33153 | 591 | |
| 60536 | 592 | lemma (in comm_ring_1) poly_divides_diff2: "p divides r \<Longrightarrow> p divides (q +++ r) \<Longrightarrow> p divides q" | 
| 52778 | 593 | apply (erule poly_divides_diff) | 
| 60536 | 594 | apply (auto simp add: poly_add fun_eq_iff poly_mult divides_def ac_simps) | 
| 52778 | 595 | done | 
| 33153 | 596 | |
| 54219 | 597 | lemma (in semiring_0) poly_divides_zero: "poly p = poly [] \<Longrightarrow> q divides p" | 
| 52778 | 598 | apply (simp add: divides_def) | 
| 60536 | 599 | apply (rule exI[where x = "[]"]) | 
| 600 | apply (auto simp add: fun_eq_iff poly_mult) | |
| 52778 | 601 | done | 
| 33153 | 602 | |
| 54219 | 603 | lemma (in semiring_0) poly_divides_zero2 [simp]: "q divides []" | 
| 52778 | 604 | apply (simp add: divides_def) | 
| 605 | apply (rule_tac x = "[]" in exI) | |
| 60536 | 606 | apply (auto simp add: fun_eq_iff) | 
| 52778 | 607 | done | 
| 33153 | 608 | |
| 60536 | 609 | |
| 610 | text \<open>At last, we can consider the order of a root.\<close> | |
| 33153 | 611 | |
| 54219 | 612 | lemma (in idom_char_0) poly_order_exists_lemma: | 
| 60698 | 613 | assumes "length p = d" | 
| 614 | and "poly p \<noteq> poly []" | |
| 54219 | 615 | shows "\<exists>n q. p = mulexp n [-a, 1] q \<and> poly q a \<noteq> 0" | 
| 60698 | 616 | using assms | 
| 54219 | 617 | proof (induct d arbitrary: p) | 
| 618 | case 0 | |
| 60536 | 619 | then show ?case by simp | 
| 54219 | 620 | next | 
| 621 | case (Suc n p) | |
| 622 | show ?case | |
| 623 | proof (cases "poly p a = 0") | |
| 624 | case True | |
| 60536 | 625 | from Suc.prems have h: "length p = Suc n" "poly p \<noteq> poly []" | 
| 626 | by auto | |
| 627 | then have pN: "p \<noteq> []" | |
| 628 | by auto | |
| 54219 | 629 | from True[unfolded poly_linear_divides] pN obtain q where q: "p = [-a, 1] *** q" | 
| 630 | by blast | |
| 631 | from q h True have qh: "length q = n" "poly q \<noteq> poly []" | |
| 60698 | 632 | apply simp_all | 
| 60536 | 633 | apply (simp only: fun_eq_iff) | 
| 54219 | 634 | apply (rule ccontr) | 
| 60536 | 635 | apply (simp add: fun_eq_iff poly_add poly_cmult) | 
| 54219 | 636 | done | 
| 637 | from Suc.hyps[OF qh] obtain m r where mr: "q = mulexp m [-a,1] r" "poly r a \<noteq> 0" | |
| 638 | by blast | |
| 60698 | 639 | from mr q have "p = mulexp (Suc m) [-a,1] r \<and> poly r a \<noteq> 0" | 
| 640 | by simp | |
| 54219 | 641 | then show ?thesis by blast | 
| 642 | next | |
| 643 | case False | |
| 644 | then show ?thesis | |
| 645 | using Suc.prems | |
| 646 | apply simp | |
| 647 | apply (rule exI[where x="0::nat"]) | |
| 648 | apply simp | |
| 649 | done | |
| 650 | qed | |
| 651 | qed | |
| 652 | ||
| 653 | ||
| 654 | lemma (in comm_semiring_1) poly_mulexp: "poly (mulexp n p q) x = (poly p x) ^ n * poly q x" | |
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 655 | by (induct n) (auto simp add: poly_mult ac_simps) | 
| 54219 | 656 | |
| 657 | lemma (in comm_semiring_1) divides_left_mult: | |
| 60536 | 658 | assumes "(p *** q) divides r" | 
| 659 | shows "p divides r \<and> q divides r" | |
| 54219 | 660 | proof- | 
| 60536 | 661 | from assms obtain t where "poly r = poly (p *** q *** t)" | 
| 54219 | 662 | unfolding divides_def by blast | 
| 60536 | 663 | then have "poly r = poly (p *** (q *** t))" and "poly r = poly (q *** (p *** t))" | 
| 664 | by (auto simp add: fun_eq_iff poly_mult ac_simps) | |
| 665 | then show ?thesis | |
| 666 | unfolding divides_def by blast | |
| 54219 | 667 | qed | 
| 668 | ||
| 33153 | 669 | |
| 670 | (* FIXME: Tidy up *) | |
| 54219 | 671 | |
| 672 | lemma (in semiring_1) zero_power_iff: "0 ^ n = (if n = 0 then 1 else 0)" | |
| 673 | by (induct n) simp_all | |
| 33153 | 674 | |
| 54219 | 675 | lemma (in idom_char_0) poly_order_exists: | 
| 60536 | 676 | assumes "length p = d" | 
| 677 | and "poly p \<noteq> poly []" | |
| 54219 | 678 | shows "\<exists>n. [- a, 1] %^ n divides p \<and> \<not> [- a, 1] %^ Suc n divides p" | 
| 679 | proof - | |
| 680 | from assms have "\<exists>n q. p = mulexp n [- a, 1] q \<and> poly q a \<noteq> 0" | |
| 681 | by (rule poly_order_exists_lemma) | |
| 60536 | 682 | then obtain n q where p: "p = mulexp n [- a, 1] q" and "poly q a \<noteq> 0" | 
| 683 | by blast | |
| 54219 | 684 | have "[- a, 1] %^ n divides mulexp n [- a, 1] q" | 
| 685 | proof (rule dividesI) | |
| 686 | show "poly (mulexp n [- a, 1] q) = poly ([- a, 1] %^ n *** q)" | |
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54219diff
changeset | 687 | by (induct n) (simp_all add: poly_add poly_cmult poly_mult algebra_simps) | 
| 54219 | 688 | qed | 
| 689 | moreover have "\<not> [- a, 1] %^ Suc n divides mulexp n [- a, 1] q" | |
| 690 | proof | |
| 691 | assume "[- a, 1] %^ Suc n divides mulexp n [- a, 1] q" | |
| 692 | then obtain m where "poly (mulexp n [- a, 1] q) = poly ([- a, 1] %^ Suc n *** m)" | |
| 693 | by (rule dividesE) | |
| 694 | moreover have "poly (mulexp n [- a, 1] q) \<noteq> poly ([- a, 1] %^ Suc n *** m)" | |
| 695 | proof (induct n) | |
| 60536 | 696 | case 0 | 
| 697 | show ?case | |
| 54219 | 698 | proof (rule ccontr) | 
| 60698 | 699 | assume "\<not> ?thesis" | 
| 54219 | 700 | then have "poly q a = 0" | 
| 701 | by (simp add: poly_add poly_cmult) | |
| 60536 | 702 | with \<open>poly q a \<noteq> 0\<close> show False | 
| 703 | by simp | |
| 54219 | 704 | qed | 
| 705 | next | |
| 60536 | 706 | case (Suc n) | 
| 707 | show ?case | |
| 60698 | 708 | by (rule pexp_Suc [THEN ssubst]) | 
| 54219 | 709 | (simp add: poly_mult_left_cancel poly_mult_assoc Suc del: pmult_Cons pexp_Suc) | 
| 710 | qed | |
| 711 | ultimately show False by simp | |
| 712 | qed | |
| 60536 | 713 | ultimately show ?thesis | 
| 714 | by (auto simp add: p) | |
| 54219 | 715 | qed | 
| 33153 | 716 | |
| 54219 | 717 | lemma (in semiring_1) poly_one_divides[simp]: "[1] divides p" | 
| 718 | by (auto simp add: divides_def) | |
| 719 | ||
| 720 | lemma (in idom_char_0) poly_order: | |
| 721 | "poly p \<noteq> poly [] \<Longrightarrow> \<exists>!n. ([-a, 1] %^ n) divides p \<and> \<not> (([-a, 1] %^ Suc n) divides p)" | |
| 52778 | 722 | apply (auto intro: poly_order_exists simp add: less_linear simp del: pmult_Cons pexp_Suc) | 
| 723 | apply (cut_tac x = y and y = n in less_linear) | |
| 724 | apply (drule_tac m = n in poly_exp_divides) | |
| 725 | apply (auto dest: Suc_le_eq [THEN iffD2, THEN [2] poly_exp_divides] | |
| 60536 | 726 | simp del: pmult_Cons pexp_Suc) | 
| 52778 | 727 | done | 
| 33153 | 728 | |
| 60536 | 729 | |
| 730 | text \<open>Order\<close> | |
| 33153 | 731 | |
| 54219 | 732 | lemma some1_equalityD: "n = (SOME n. P n) \<Longrightarrow> \<exists>!n. P n \<Longrightarrow> P n" | 
| 52778 | 733 | by (blast intro: someI2) | 
| 33153 | 734 | |
| 54219 | 735 | lemma (in idom_char_0) order: | 
| 60536 | 736 | "([-a, 1] %^ n) divides p \<and> \<not> (([-a, 1] %^ Suc n) divides p) \<longleftrightarrow> | 
| 737 | n = order a p \<and> poly p \<noteq> poly []" | |
| 738 | unfolding order_def | |
| 52778 | 739 | apply (rule iffI) | 
| 740 | apply (blast dest: poly_divides_zero intro!: some1_equality [symmetric] poly_order) | |
| 741 | apply (blast intro!: poly_order [THEN [2] some1_equalityD]) | |
| 742 | done | |
| 33153 | 743 | |
| 54219 | 744 | lemma (in idom_char_0) order2: | 
| 745 | "poly p \<noteq> poly [] \<Longrightarrow> | |
| 60536 | 746 | ([-a, 1] %^ (order a p)) divides p \<and> \<not> ([-a, 1] %^ Suc (order a p)) divides p" | 
| 52778 | 747 | by (simp add: order del: pexp_Suc) | 
| 33153 | 748 | |
| 54219 | 749 | lemma (in idom_char_0) order_unique: | 
| 60536 | 750 | "poly p \<noteq> poly [] \<Longrightarrow> ([-a, 1] %^ n) divides p \<Longrightarrow> \<not> ([-a, 1] %^ (Suc n)) divides p \<Longrightarrow> | 
| 54219 | 751 | n = order a p" | 
| 52778 | 752 | using order [of a n p] by auto | 
| 33153 | 753 | |
| 54219 | 754 | lemma (in idom_char_0) order_unique_lemma: | 
| 60536 | 755 | "poly p \<noteq> poly [] \<and> ([-a, 1] %^ n) divides p \<and> \<not> ([-a, 1] %^ (Suc n)) divides p \<Longrightarrow> | 
| 52881 | 756 | n = order a p" | 
| 52778 | 757 | by (blast intro: order_unique) | 
| 33153 | 758 | |
| 54219 | 759 | lemma (in ring_1) order_poly: "poly p = poly q \<Longrightarrow> order a p = order a q" | 
| 60536 | 760 | by (auto simp add: fun_eq_iff divides_def poly_mult order_def) | 
| 33153 | 761 | |
| 54219 | 762 | lemma (in semiring_1) pexp_one[simp]: "p %^ (Suc 0) = p" | 
| 60536 | 763 | by (induct p) auto | 
| 54219 | 764 | |
| 765 | lemma (in comm_ring_1) lemma_order_root: | |
| 60536 | 766 | "0 < n \<and> [- a, 1] %^ n divides p \<and> \<not> [- a, 1] %^ (Suc n) divides p \<Longrightarrow> poly p a = 0" | 
| 54219 | 767 | by (induct n arbitrary: a p) (auto simp add: divides_def poly_mult simp del: pmult_Cons) | 
| 33153 | 768 | |
| 60536 | 769 | lemma (in idom_char_0) order_root: "poly p a = 0 \<longleftrightarrow> poly p = poly [] \<or> order a p \<noteq> 0" | 
| 54219 | 770 | apply (cases "poly p = poly []") | 
| 771 | apply auto | |
| 60536 | 772 | apply (simp add: poly_linear_divides del: pmult_Cons) | 
| 773 | apply safe | |
| 54219 | 774 | apply (drule_tac [!] a = a in order2) | 
| 775 | apply (rule ccontr) | |
| 60536 | 776 | apply (simp add: divides_def poly_mult fun_eq_iff del: pmult_Cons) | 
| 777 | apply blast | |
| 778 | using neq0_conv apply (blast intro: lemma_order_root) | |
| 52778 | 779 | done | 
| 33153 | 780 | |
| 54219 | 781 | lemma (in idom_char_0) order_divides: | 
| 782 | "([-a, 1] %^ n) divides p \<longleftrightarrow> poly p = poly [] \<or> n \<le> order a p" | |
| 52881 | 783 | apply (cases "poly p = poly []") | 
| 784 | apply auto | |
| 60536 | 785 | apply (simp add: divides_def fun_eq_iff poly_mult) | 
| 52778 | 786 | apply (rule_tac x = "[]" in exI) | 
| 54219 | 787 | apply (auto dest!: order2 [where a=a] intro: poly_exp_divides simp del: pexp_Suc) | 
| 52778 | 788 | done | 
| 33153 | 789 | |
| 54219 | 790 | lemma (in idom_char_0) order_decomp: | 
| 60536 | 791 | "poly p \<noteq> poly [] \<Longrightarrow> \<exists>q. poly p = poly (([-a, 1] %^ order a p) *** q) \<and> \<not> [-a, 1] divides q" | 
| 792 | unfolding divides_def | |
| 52778 | 793 | apply (drule order2 [where a = a]) | 
| 60536 | 794 | apply (simp add: divides_def del: pexp_Suc pmult_Cons) | 
| 795 | apply safe | |
| 796 | apply (rule_tac x = q in exI) | |
| 797 | apply safe | |
| 52778 | 798 | apply (drule_tac x = qa in spec) | 
| 60536 | 799 | apply (auto simp add: poly_mult fun_eq_iff poly_exp ac_simps simp del: pmult_Cons) | 
| 52778 | 800 | done | 
| 33153 | 801 | |
| 60536 | 802 | text \<open>Important composition properties of orders.\<close> | 
| 54219 | 803 | lemma order_mult: | 
| 60536 | 804 | fixes a :: "'a::idom_char_0" | 
| 805 | shows "poly (p *** q) \<noteq> poly [] \<Longrightarrow> order a (p *** q) = order a p + order a q" | |
| 54219 | 806 | apply (cut_tac a = a and p = "p *** q" and n = "order a p + order a q" in order) | 
| 52778 | 807 | apply (auto simp add: poly_entire simp del: pmult_Cons) | 
| 808 | apply (drule_tac a = a in order2)+ | |
| 809 | apply safe | |
| 60536 | 810 | apply (simp add: divides_def fun_eq_iff poly_exp_add poly_mult del: pmult_Cons, safe) | 
| 52778 | 811 | apply (rule_tac x = "qa *** qaa" in exI) | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 812 | apply (simp add: poly_mult ac_simps del: pmult_Cons) | 
| 52778 | 813 | apply (drule_tac a = a in order_decomp)+ | 
| 814 | apply safe | |
| 60536 | 815 | apply (subgoal_tac "[-a, 1] divides (qa *** qaa) ") | 
| 52778 | 816 | apply (simp add: poly_primes del: pmult_Cons) | 
| 817 | apply (auto simp add: divides_def simp del: pmult_Cons) | |
| 818 | apply (rule_tac x = qb in exI) | |
| 60536 | 819 | apply (subgoal_tac "poly ([-a, 1] %^ (order a p) *** (qa *** qaa)) = | 
| 820 | poly ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))") | |
| 821 | apply (drule poly_mult_left_cancel [THEN iffD1]) | |
| 822 | apply force | |
| 823 | apply (subgoal_tac "poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** (qa *** qaa))) = | |
| 824 | poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))) ") | |
| 825 | apply (drule poly_mult_left_cancel [THEN iffD1]) | |
| 826 | apply force | |
| 827 | apply (simp add: fun_eq_iff poly_exp_add poly_mult ac_simps del: pmult_Cons) | |
| 52778 | 828 | done | 
| 33153 | 829 | |
| 54219 | 830 | lemma (in idom_char_0) order_mult: | 
| 831 | assumes "poly (p *** q) \<noteq> poly []" | |
| 832 | shows "order a (p *** q) = order a p + order a q" | |
| 833 | using assms | |
| 834 | apply (cut_tac a = a and p = "pmult p q" and n = "order a p + order a q" in order) | |
| 835 | apply (auto simp add: poly_entire simp del: pmult_Cons) | |
| 836 | apply (drule_tac a = a in order2)+ | |
| 837 | apply safe | |
| 60536 | 838 | apply (simp add: divides_def fun_eq_iff poly_exp_add poly_mult del: pmult_Cons) | 
| 839 | apply safe | |
| 54219 | 840 | apply (rule_tac x = "pmult qa qaa" in exI) | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 841 | apply (simp add: poly_mult ac_simps del: pmult_Cons) | 
| 54219 | 842 | apply (drule_tac a = a in order_decomp)+ | 
| 843 | apply safe | |
| 844 | apply (subgoal_tac "[uminus a, one] divides pmult qa qaa") | |
| 845 | apply (simp add: poly_primes del: pmult_Cons) | |
| 846 | apply (auto simp add: divides_def simp del: pmult_Cons) | |
| 847 | apply (rule_tac x = qb in exI) | |
| 848 | apply (subgoal_tac "poly (pmult (pexp [uminus a, one] (order a p)) (pmult qa qaa)) = | |
| 59807 | 849 | poly (pmult (pexp [uminus a, one] (order a p)) (pmult [uminus a, one] qb))") | 
| 54219 | 850 | apply (drule poly_mult_left_cancel [THEN iffD1], force) | 
| 851 | apply (subgoal_tac "poly (pmult (pexp [uminus a, one] (order a q)) | |
| 852 | (pmult (pexp [uminus a, one] (order a p)) (pmult qa qaa))) = | |
| 853 | poly (pmult (pexp [uminus a, one] (order a q)) | |
| 854 | (pmult (pexp [uminus a, one] (order a p)) (pmult [uminus a, one] qb)))") | |
| 855 | apply (drule poly_mult_left_cancel [THEN iffD1], force) | |
| 60536 | 856 | apply (simp add: fun_eq_iff poly_exp_add poly_mult ac_simps del: pmult_Cons) | 
| 54219 | 857 | done | 
| 858 | ||
| 859 | lemma (in idom_char_0) order_root2: "poly p \<noteq> poly [] \<Longrightarrow> poly p a = 0 \<longleftrightarrow> order a p \<noteq> 0" | |
| 52881 | 860 | by (rule order_root [THEN ssubst]) auto | 
| 33153 | 861 | |
| 60536 | 862 | lemma (in semiring_1) pmult_one[simp]: "[1] *** p = p" | 
| 863 | by auto | |
| 33153 | 864 | |
| 54219 | 865 | lemma (in semiring_0) poly_Nil_zero: "poly [] = poly [0]" | 
| 60536 | 866 | by (simp add: fun_eq_iff) | 
| 33153 | 867 | |
| 54219 | 868 | lemma (in idom_char_0) rsquarefree_decomp: | 
| 60536 | 869 | "rsquarefree p \<Longrightarrow> poly p a = 0 \<Longrightarrow> \<exists>q. poly p = poly ([-a, 1] *** q) \<and> poly q a \<noteq> 0" | 
| 870 | apply (simp add: rsquarefree_def) | |
| 871 | apply safe | |
| 52778 | 872 | apply (frule_tac a = a in order_decomp) | 
| 873 | apply (drule_tac x = a in spec) | |
| 874 | apply (drule_tac a = a in order_root2 [symmetric]) | |
| 875 | apply (auto simp del: pmult_Cons) | |
| 54219 | 876 | apply (rule_tac x = q in exI, safe) | 
| 60536 | 877 | apply (simp add: poly_mult fun_eq_iff) | 
| 52778 | 878 | apply (drule_tac p1 = q in poly_linear_divides [THEN iffD1]) | 
| 54219 | 879 | apply (simp add: divides_def del: pmult_Cons, safe) | 
| 52778 | 880 | apply (drule_tac x = "[]" in spec) | 
| 60536 | 881 | apply (auto simp add: fun_eq_iff) | 
| 52778 | 882 | done | 
| 33153 | 883 | |
| 884 | ||
| 60536 | 885 | text \<open>Normalization of a polynomial.\<close> | 
| 33153 | 886 | |
| 54219 | 887 | lemma (in semiring_0) poly_normalize[simp]: "poly (pnormalize p) = poly p" | 
| 60536 | 888 | by (induct p) (auto simp add: fun_eq_iff) | 
| 33153 | 889 | |
| 60536 | 890 | text \<open>The degree of a polynomial.\<close> | 
| 33153 | 891 | |
| 60537 | 892 | lemma (in semiring_0) lemma_degree_zero: "(\<forall>c \<in> set p. c = 0) \<longleftrightarrow> pnormalize p = []" | 
| 52778 | 893 | by (induct p) auto | 
| 33153 | 894 | |
| 54219 | 895 | lemma (in idom_char_0) degree_zero: | 
| 896 | assumes "poly p = poly []" | |
| 897 | shows "degree p = 0" | |
| 898 | using assms | |
| 899 | by (cases "pnormalize p = []") (auto simp add: degree_def poly_zero lemma_degree_zero) | |
| 33153 | 900 | |
| 60536 | 901 | lemma (in semiring_0) pnormalize_sing: "pnormalize [x] = [x] \<longleftrightarrow> x \<noteq> 0" | 
| 54219 | 902 | by simp | 
| 903 | ||
| 60536 | 904 | lemma (in semiring_0) pnormalize_pair: "y \<noteq> 0 \<longleftrightarrow> pnormalize [x, y] = [x, y]" | 
| 52881 | 905 | by simp | 
| 52778 | 906 | |
| 60536 | 907 | lemma (in semiring_0) pnormal_cons: "pnormal p \<Longrightarrow> pnormal (c # p)" | 
| 33153 | 908 | unfolding pnormal_def by simp | 
| 52778 | 909 | |
| 60536 | 910 | lemma (in semiring_0) pnormal_tail: "p \<noteq> [] \<Longrightarrow> pnormal (c # p) \<Longrightarrow> pnormal p" | 
| 62390 | 911 | unfolding pnormal_def by (auto split: if_split_asm) | 
| 54219 | 912 | |
| 913 | lemma (in semiring_0) pnormal_last_nonzero: "pnormal p \<Longrightarrow> last p \<noteq> 0" | |
| 62390 | 914 | by (induct p) (simp_all add: pnormal_def split: if_split_asm) | 
| 54219 | 915 | |
| 916 | lemma (in semiring_0) pnormal_length: "pnormal p \<Longrightarrow> 0 < length p" | |
| 917 | unfolding pnormal_def length_greater_0_conv by blast | |
| 918 | ||
| 919 | lemma (in semiring_0) pnormal_last_length: "0 < length p \<Longrightarrow> last p \<noteq> 0 \<Longrightarrow> pnormal p" | |
| 62390 | 920 | by (induct p) (auto simp: pnormal_def split: if_split_asm) | 
| 54219 | 921 | |
| 922 | lemma (in semiring_0) pnormal_id: "pnormal p \<longleftrightarrow> 0 < length p \<and> last p \<noteq> 0" | |
| 923 | using pnormal_last_length pnormal_length pnormal_last_nonzero by blast | |
| 924 | ||
| 60698 | 925 | lemma (in idom_char_0) poly_Cons_eq: "poly (c # cs) = poly (d # ds) \<longleftrightarrow> c = d \<and> poly cs = poly ds" | 
| 54219 | 926 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 927 | proof | |
| 60536 | 928 | show ?rhs if ?lhs | 
| 929 | proof - | |
| 930 | from that have "poly ((c # cs) +++ -- (d # ds)) x = 0" for x | |
| 931 | by (simp only: poly_minus poly_add algebra_simps) (simp add: algebra_simps) | |
| 932 | then have "poly ((c # cs) +++ -- (d # ds)) = poly []" | |
| 933 | by (simp add: fun_eq_iff) | |
| 60537 | 934 | then have "c = d" and "\<forall>x \<in> set (cs +++ -- ds). x = 0" | 
| 60536 | 935 | unfolding poly_zero by (simp_all add: poly_minus_def algebra_simps) | 
| 936 | from this(2) have "poly (cs +++ -- ds) x = 0" for x | |
| 937 | unfolding poly_zero[symmetric] by simp | |
| 938 | with \<open>c = d\<close> show ?thesis | |
| 939 | by (simp add: poly_minus poly_add algebra_simps fun_eq_iff) | |
| 940 | qed | |
| 941 | show ?lhs if ?rhs | |
| 942 | using that by (simp add:fun_eq_iff) | |
| 54219 | 943 | qed | 
| 944 | ||
| 945 | lemma (in idom_char_0) pnormalize_unique: "poly p = poly q \<Longrightarrow> pnormalize p = pnormalize q" | |
| 946 | proof (induct q arbitrary: p) | |
| 947 | case Nil | |
| 60536 | 948 | then show ?case | 
| 949 | by (simp only: poly_zero lemma_degree_zero) simp | |
| 54219 | 950 | next | 
| 951 | case (Cons c cs p) | |
| 60536 | 952 | then show ?case | 
| 54219 | 953 | proof (induct p) | 
| 954 | case Nil | |
| 60536 | 955 | then have "poly [] = poly (c # cs)" | 
| 956 | by blast | |
| 957 | then have "poly (c#cs) = poly []" | |
| 958 | by simp | |
| 959 | then show ?case | |
| 960 | by (simp only: poly_zero lemma_degree_zero) simp | |
| 54219 | 961 | next | 
| 962 | case (Cons d ds) | |
| 60536 | 963 | then have eq: "poly (d # ds) = poly (c # cs)" | 
| 964 | by blast | |
| 965 | then have eq': "\<And>x. poly (d # ds) x = poly (c # cs) x" | |
| 966 | by simp | |
| 967 | then have "poly (d # ds) 0 = poly (c # cs) 0" | |
| 968 | by blast | |
| 969 | then have dc: "d = c" | |
| 970 | by auto | |
| 54219 | 971 | with eq have "poly ds = poly cs" | 
| 972 | unfolding poly_Cons_eq by simp | |
| 60536 | 973 | with Cons.prems have "pnormalize ds = pnormalize cs" | 
| 974 | by blast | |
| 975 | with dc show ?case | |
| 976 | by simp | |
| 54219 | 977 | qed | 
| 978 | qed | |
| 979 | ||
| 980 | lemma (in idom_char_0) degree_unique: | |
| 981 | assumes pq: "poly p = poly q" | |
| 982 | shows "degree p = degree q" | |
| 983 | using pnormalize_unique[OF pq] unfolding degree_def by simp | |
| 984 | ||
| 60536 | 985 | lemma (in semiring_0) pnormalize_length: "length (pnormalize p) \<le> length p" | 
| 986 | by (induct p) auto | |
| 54219 | 987 | |
| 988 | lemma (in semiring_0) last_linear_mul_lemma: | |
| 60536 | 989 | "last ((a %* p) +++ (x # (b %* p))) = (if p = [] then x else b * last p)" | 
| 54219 | 990 | apply (induct p arbitrary: a x b) | 
| 52881 | 991 | apply auto | 
| 60698 | 992 | subgoal for a p c x b | 
| 993 | apply (subgoal_tac "padd (cmult c p) (times b a # cmult b p) \<noteq> []") | |
| 994 | apply simp | |
| 995 | apply (induct p) | |
| 996 | apply auto | |
| 997 | done | |
| 52778 | 998 | done | 
| 999 | ||
| 54219 | 1000 | lemma (in semiring_1) last_linear_mul: | 
| 1001 | assumes p: "p \<noteq> []" | |
| 60536 | 1002 | shows "last ([a, 1] *** p) = last p" | 
| 54219 | 1003 | proof - | 
| 60536 | 1004 | from p obtain c cs where cs: "p = c # cs" | 
| 1005 | by (cases p) auto | |
| 1006 | from cs have eq: "[a, 1] *** p = (a %* (c # cs)) +++ (0 # (1 %* (c # cs)))" | |
| 54219 | 1007 | by (simp add: poly_cmult_distr) | 
| 60536 | 1008 | show ?thesis | 
| 1009 | using cs unfolding eq last_linear_mul_lemma by simp | |
| 54219 | 1010 | qed | 
| 1011 | ||
| 1012 | lemma (in semiring_0) pnormalize_eq: "last p \<noteq> 0 \<Longrightarrow> pnormalize p = p" | |
| 62390 | 1013 | by (induct p) (auto split: if_split_asm) | 
| 54219 | 1014 | |
| 1015 | lemma (in semiring_0) last_pnormalize: "pnormalize p \<noteq> [] \<Longrightarrow> last (pnormalize p) \<noteq> 0" | |
| 1016 | by (induct p) auto | |
| 1017 | ||
| 1018 | lemma (in semiring_0) pnormal_degree: "last p \<noteq> 0 \<Longrightarrow> degree p = length p - 1" | |
| 1019 | using pnormalize_eq[of p] unfolding degree_def by simp | |
| 52778 | 1020 | |
| 54219 | 1021 | lemma (in semiring_0) poly_Nil_ext: "poly [] = (\<lambda>x. 0)" | 
| 60536 | 1022 | by auto | 
| 54219 | 1023 | |
| 1024 | lemma (in idom_char_0) linear_mul_degree: | |
| 1025 | assumes p: "poly p \<noteq> poly []" | |
| 60536 | 1026 | shows "degree ([a, 1] *** p) = degree p + 1" | 
| 54219 | 1027 | proof - | 
| 1028 | from p have pnz: "pnormalize p \<noteq> []" | |
| 1029 | unfolding poly_zero lemma_degree_zero . | |
| 1030 | ||
| 1031 | from last_linear_mul[OF pnz, of a] last_pnormalize[OF pnz] | |
| 1032 | have l0: "last ([a, 1] *** pnormalize p) \<noteq> 0" by simp | |
| 60536 | 1033 | |
| 54219 | 1034 | from last_pnormalize[OF pnz] last_linear_mul[OF pnz, of a] | 
| 1035 | pnormal_degree[OF l0] pnormal_degree[OF last_pnormalize[OF pnz]] pnz | |
| 1036 | have th: "degree ([a,1] *** pnormalize p) = degree (pnormalize p) + 1" | |
| 1037 | by simp | |
| 1038 | ||
| 1039 | have eqs: "poly ([a,1] *** pnormalize p) = poly ([a,1] *** p)" | |
| 1040 | by (rule ext) (simp add: poly_mult poly_add poly_cmult) | |
| 60536 | 1041 | from degree_unique[OF eqs] th show ?thesis | 
| 1042 | by (simp add: degree_unique[OF poly_normalize]) | |
| 54219 | 1043 | qed | 
| 52778 | 1044 | |
| 54219 | 1045 | lemma (in idom_char_0) linear_pow_mul_degree: | 
| 1046 | "degree([a,1] %^n *** p) = (if poly p = poly [] then 0 else degree p + n)" | |
| 1047 | proof (induct n arbitrary: a p) | |
| 1048 | case (0 a p) | |
| 1049 | show ?case | |
| 1050 | proof (cases "poly p = poly []") | |
| 1051 | case True | |
| 1052 | then show ?thesis | |
| 1053 | using degree_unique[OF True] by (simp add: degree_def) | |
| 1054 | next | |
| 1055 | case False | |
| 60536 | 1056 | then show ?thesis | 
| 1057 | by (auto simp add: poly_Nil_ext) | |
| 54219 | 1058 | qed | 
| 1059 | next | |
| 1060 | case (Suc n a p) | |
| 60536 | 1061 | have eq: "poly ([a, 1] %^(Suc n) *** p) = poly ([a, 1] %^ n *** ([a, 1] *** p))" | 
| 54219 | 1062 | apply (rule ext) | 
| 1063 | apply (simp add: poly_mult poly_add poly_cmult) | |
| 60536 | 1064 | apply (simp add: ac_simps distrib_left) | 
| 54219 | 1065 | done | 
| 1066 | note deq = degree_unique[OF eq] | |
| 1067 | show ?case | |
| 1068 | proof (cases "poly p = poly []") | |
| 1069 | case True | |
| 60536 | 1070 | with eq have eq': "poly ([a, 1] %^(Suc n) *** p) = poly []" | 
| 1071 | by (auto simp add: poly_mult poly_cmult poly_add) | |
| 54219 | 1072 | from degree_unique[OF eq'] True show ?thesis | 
| 1073 | by (simp add: degree_def) | |
| 1074 | next | |
| 1075 | case False | |
| 1076 | then have ap: "poly ([a,1] *** p) \<noteq> poly []" | |
| 1077 | using poly_mult_not_eq_poly_Nil unfolding poly_entire by auto | |
| 60536 | 1078 | have eq: "poly ([a, 1] %^(Suc n) *** p) = poly ([a, 1]%^n *** ([a, 1] *** p))" | 
| 1079 | by (auto simp add: poly_mult poly_add poly_exp poly_cmult algebra_simps) | |
| 1080 | from ap have ap': "poly ([a, 1] *** p) = poly [] \<longleftrightarrow> False" | |
| 54219 | 1081 | by blast | 
| 60536 | 1082 | have th0: "degree ([a, 1]%^n *** ([a, 1] *** p)) = degree ([a, 1] *** p) + n" | 
| 54219 | 1083 | apply (simp only: Suc.hyps[of a "pmult [a,one] p"] ap') | 
| 1084 | apply simp | |
| 1085 | done | |
| 1086 | from degree_unique[OF eq] ap False th0 linear_mul_degree[OF False, of a] | |
| 60536 | 1087 | show ?thesis | 
| 1088 | by (auto simp del: poly.simps) | |
| 54219 | 1089 | qed | 
| 1090 | qed | |
| 52778 | 1091 | |
| 54219 | 1092 | lemma (in idom_char_0) order_degree: | 
| 1093 | assumes p0: "poly p \<noteq> poly []" | |
| 1094 | shows "order a p \<le> degree p" | |
| 1095 | proof - | |
| 1096 | from order2[OF p0, unfolded divides_def] | |
| 60536 | 1097 | obtain q where q: "poly p = poly ([- a, 1]%^ (order a p) *** q)" | 
| 1098 | by blast | |
| 1099 | with q p0 have "poly q \<noteq> poly []" | |
| 1100 | by (simp add: poly_mult poly_entire) | |
| 54219 | 1101 | with degree_unique[OF q, unfolded linear_pow_mul_degree] show ?thesis | 
| 1102 | by auto | |
| 1103 | qed | |
| 33153 | 1104 | |
| 1105 | ||
| 60536 | 1106 | text \<open>Tidier versions of finiteness of roots.\<close> | 
| 54219 | 1107 | lemma (in idom_char_0) poly_roots_finite_set: | 
| 1108 |   "poly p \<noteq> poly [] \<Longrightarrow> finite {x. poly p x = 0}"
 | |
| 52778 | 1109 | unfolding poly_roots_finite . | 
| 33153 | 1110 | |
| 1111 | ||
| 60536 | 1112 | text \<open>Bound for polynomial.\<close> | 
| 1113 | lemma poly_mono: | |
| 1114 | fixes x :: "'a::linordered_idom" | |
| 61945 | 1115 | shows "\<bar>x\<bar> \<le> k \<Longrightarrow> \<bar>poly p x\<bar> \<le> poly (map abs p) k" | 
| 60698 | 1116 | proof (induct p) | 
| 1117 | case Nil | |
| 1118 | then show ?case by simp | |
| 1119 | next | |
| 1120 | case (Cons a p) | |
| 1121 | then show ?case | |
| 1122 | apply auto | |
| 61945 | 1123 | apply (rule_tac y = "\<bar>a\<bar> + \<bar>x * poly p x\<bar>" in order_trans) | 
| 60698 | 1124 | apply (rule abs_triangle_ineq) | 
| 1125 | apply (auto intro!: mult_mono simp add: abs_mult) | |
| 1126 | done | |
| 1127 | qed | |
| 33153 | 1128 | |
| 60536 | 1129 | lemma (in semiring_0) poly_Sing: "poly [c] x = c" | 
| 1130 | by simp | |
| 33268 
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
33153diff
changeset | 1131 | |
| 33153 | 1132 | end |