| author | Christian Sternagel | 
| Thu, 30 Aug 2012 13:06:04 +0900 | |
| changeset 49088 | 5cd8b4426a57 | 
| parent 47382 | 5b902eeb2a29 | 
| child 49834 | b27bbb021df1 | 
| permissions | -rw-r--r-- | 
| 41959 | 1 | (* Title: HOL/Library/Polynomial.thy | 
| 29451 | 2 | Author: Brian Huffman | 
| 41959 | 3 | Author: Clemens Ballarin | 
| 29451 | 4 | *) | 
| 5 | ||
| 6 | header {* Univariate Polynomials *}
 | |
| 7 | ||
| 8 | theory Polynomial | |
| 30738 | 9 | imports Main | 
| 29451 | 10 | begin | 
| 11 | ||
| 12 | subsection {* Definition of type @{text poly} *}
 | |
| 13 | ||
| 45694 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
45605diff
changeset | 14 | definition "Poly = {f::nat \<Rightarrow> 'a::zero. \<exists>n. \<forall>i>n. f i = 0}"
 | 
| 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
45605diff
changeset | 15 | |
| 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
45605diff
changeset | 16 | typedef (open) 'a poly = "Poly :: (nat => 'a::zero) set" | 
| 29451 | 17 | morphisms coeff Abs_poly | 
| 45694 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
45605diff
changeset | 18 | unfolding Poly_def by auto | 
| 29451 | 19 | |
| 47002 | 20 | (* FIXME should be named poly_eq_iff *) | 
| 29451 | 21 | lemma expand_poly_eq: "p = q \<longleftrightarrow> (\<forall>n. coeff p n = coeff q n)" | 
| 45694 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
45605diff
changeset | 22 | by (simp add: coeff_inject [symmetric] fun_eq_iff) | 
| 29451 | 23 | |
| 47002 | 24 | (* FIXME should be named poly_eqI *) | 
| 29451 | 25 | lemma poly_ext: "(\<And>n. coeff p n = coeff q n) \<Longrightarrow> p = q" | 
| 45694 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
45605diff
changeset | 26 | by (simp add: expand_poly_eq) | 
| 29451 | 27 | |
| 28 | ||
| 29 | subsection {* Degree of a polynomial *}
 | |
| 30 | ||
| 31 | definition | |
| 32 | degree :: "'a::zero poly \<Rightarrow> nat" where | |
| 33 | "degree p = (LEAST n. \<forall>i>n. coeff p i = 0)" | |
| 34 | ||
| 35 | lemma coeff_eq_0: "degree p < n \<Longrightarrow> coeff p n = 0" | |
| 36 | proof - | |
| 37 | have "coeff p \<in> Poly" | |
| 38 | by (rule coeff) | |
| 39 | hence "\<exists>n. \<forall>i>n. coeff p i = 0" | |
| 40 | unfolding Poly_def by simp | |
| 41 | hence "\<forall>i>degree p. coeff p i = 0" | |
| 42 | unfolding degree_def by (rule LeastI_ex) | |
| 43 | moreover assume "degree p < n" | |
| 44 | ultimately show ?thesis by simp | |
| 45 | qed | |
| 46 | ||
| 47 | lemma le_degree: "coeff p n \<noteq> 0 \<Longrightarrow> n \<le> degree p" | |
| 48 | by (erule contrapos_np, rule coeff_eq_0, simp) | |
| 49 | ||
| 50 | lemma degree_le: "\<forall>i>n. coeff p i = 0 \<Longrightarrow> degree p \<le> n" | |
| 51 | unfolding degree_def by (erule Least_le) | |
| 52 | ||
| 53 | lemma less_degree_imp: "n < degree p \<Longrightarrow> \<exists>i>n. coeff p i \<noteq> 0" | |
| 54 | unfolding degree_def by (drule not_less_Least, simp) | |
| 55 | ||
| 56 | ||
| 57 | subsection {* The zero polynomial *}
 | |
| 58 | ||
| 59 | instantiation poly :: (zero) zero | |
| 60 | begin | |
| 61 | ||
| 62 | definition | |
| 63 | zero_poly_def: "0 = Abs_poly (\<lambda>n. 0)" | |
| 64 | ||
| 65 | instance .. | |
| 66 | end | |
| 67 | ||
| 68 | lemma coeff_0 [simp]: "coeff 0 n = 0" | |
| 69 | unfolding zero_poly_def | |
| 70 | by (simp add: Abs_poly_inverse Poly_def) | |
| 71 | ||
| 72 | lemma degree_0 [simp]: "degree 0 = 0" | |
| 73 | by (rule order_antisym [OF degree_le le0]) simp | |
| 74 | ||
| 75 | lemma leading_coeff_neq_0: | |
| 76 | assumes "p \<noteq> 0" shows "coeff p (degree p) \<noteq> 0" | |
| 77 | proof (cases "degree p") | |
| 78 | case 0 | |
| 79 | from `p \<noteq> 0` have "\<exists>n. coeff p n \<noteq> 0" | |
| 80 | by (simp add: expand_poly_eq) | |
| 81 | then obtain n where "coeff p n \<noteq> 0" .. | |
| 82 | hence "n \<le> degree p" by (rule le_degree) | |
| 83 | with `coeff p n \<noteq> 0` and `degree p = 0` | |
| 84 | show "coeff p (degree p) \<noteq> 0" by simp | |
| 85 | next | |
| 86 | case (Suc n) | |
| 87 | from `degree p = Suc n` have "n < degree p" by simp | |
| 88 | hence "\<exists>i>n. coeff p i \<noteq> 0" by (rule less_degree_imp) | |
| 89 | then obtain i where "n < i" and "coeff p i \<noteq> 0" by fast | |
| 90 | from `degree p = Suc n` and `n < i` have "degree p \<le> i" by simp | |
| 91 | also from `coeff p i \<noteq> 0` have "i \<le> degree p" by (rule le_degree) | |
| 92 | finally have "degree p = i" . | |
| 93 | with `coeff p i \<noteq> 0` show "coeff p (degree p) \<noteq> 0" by simp | |
| 94 | qed | |
| 95 | ||
| 96 | lemma leading_coeff_0_iff [simp]: "coeff p (degree p) = 0 \<longleftrightarrow> p = 0" | |
| 97 | by (cases "p = 0", simp, simp add: leading_coeff_neq_0) | |
| 98 | ||
| 99 | ||
| 100 | subsection {* List-style constructor for polynomials *}
 | |
| 101 | ||
| 102 | definition | |
| 103 | pCons :: "'a::zero \<Rightarrow> 'a poly \<Rightarrow> 'a poly" | |
| 104 | where | |
| 37765 | 105 | "pCons a p = Abs_poly (nat_case a (coeff p))" | 
| 29451 | 106 | |
| 29455 | 107 | syntax | 
| 108 |   "_poly" :: "args \<Rightarrow> 'a poly"  ("[:(_):]")
 | |
| 109 | ||
| 110 | translations | |
| 111 | "[:x, xs:]" == "CONST pCons x [:xs:]" | |
| 112 | "[:x:]" == "CONST pCons x 0" | |
| 30155 
f65dc19cd5f0
make list-style polynomial syntax work when show_sorts is on
 huffman parents: 
30072diff
changeset | 113 | "[:x:]" <= "CONST pCons x (_constrain 0 t)" | 
| 29455 | 114 | |
| 29451 | 115 | lemma Poly_nat_case: "f \<in> Poly \<Longrightarrow> nat_case a f \<in> Poly" | 
| 116 | unfolding Poly_def by (auto split: nat.split) | |
| 117 | ||
| 118 | lemma coeff_pCons: | |
| 119 | "coeff (pCons a p) = nat_case a (coeff p)" | |
| 120 | unfolding pCons_def | |
| 121 | by (simp add: Abs_poly_inverse Poly_nat_case coeff) | |
| 122 | ||
| 123 | lemma coeff_pCons_0 [simp]: "coeff (pCons a p) 0 = a" | |
| 124 | by (simp add: coeff_pCons) | |
| 125 | ||
| 126 | lemma coeff_pCons_Suc [simp]: "coeff (pCons a p) (Suc n) = coeff p n" | |
| 127 | by (simp add: coeff_pCons) | |
| 128 | ||
| 129 | lemma degree_pCons_le: "degree (pCons a p) \<le> Suc (degree p)" | |
| 130 | by (rule degree_le, simp add: coeff_eq_0 coeff_pCons split: nat.split) | |
| 131 | ||
| 132 | lemma degree_pCons_eq: | |
| 133 | "p \<noteq> 0 \<Longrightarrow> degree (pCons a p) = Suc (degree p)" | |
| 134 | apply (rule order_antisym [OF degree_pCons_le]) | |
| 135 | apply (rule le_degree, simp) | |
| 136 | done | |
| 137 | ||
| 138 | lemma degree_pCons_0: "degree (pCons a 0) = 0" | |
| 139 | apply (rule order_antisym [OF _ le0]) | |
| 140 | apply (rule degree_le, simp add: coeff_pCons split: nat.split) | |
| 141 | done | |
| 142 | ||
| 29460 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 143 | lemma degree_pCons_eq_if [simp]: | 
| 29451 | 144 | "degree (pCons a p) = (if p = 0 then 0 else Suc (degree p))" | 
| 145 | apply (cases "p = 0", simp_all) | |
| 146 | apply (rule order_antisym [OF _ le0]) | |
| 147 | apply (rule degree_le, simp add: coeff_pCons split: nat.split) | |
| 148 | apply (rule order_antisym [OF degree_pCons_le]) | |
| 149 | apply (rule le_degree, simp) | |
| 150 | done | |
| 151 | ||
| 46031 | 152 | lemma pCons_0_0 [simp, code_post]: "pCons 0 0 = 0" | 
| 29451 | 153 | by (rule poly_ext, simp add: coeff_pCons split: nat.split) | 
| 154 | ||
| 155 | lemma pCons_eq_iff [simp]: | |
| 156 | "pCons a p = pCons b q \<longleftrightarrow> a = b \<and> p = q" | |
| 157 | proof (safe) | |
| 158 | assume "pCons a p = pCons b q" | |
| 159 | then have "coeff (pCons a p) 0 = coeff (pCons b q) 0" by simp | |
| 160 | then show "a = b" by simp | |
| 161 | next | |
| 162 | assume "pCons a p = pCons b q" | |
| 163 | then have "\<forall>n. coeff (pCons a p) (Suc n) = | |
| 164 | coeff (pCons b q) (Suc n)" by simp | |
| 165 | then show "p = q" by (simp add: expand_poly_eq) | |
| 166 | qed | |
| 167 | ||
| 168 | lemma pCons_eq_0_iff [simp]: "pCons a p = 0 \<longleftrightarrow> a = 0 \<and> p = 0" | |
| 169 | using pCons_eq_iff [of a p 0 0] by simp | |
| 170 | ||
| 171 | lemma Poly_Suc: "f \<in> Poly \<Longrightarrow> (\<lambda>n. f (Suc n)) \<in> Poly" | |
| 172 | unfolding Poly_def | |
| 173 | by (clarify, rule_tac x=n in exI, simp) | |
| 174 | ||
| 175 | lemma pCons_cases [cases type: poly]: | |
| 176 | obtains (pCons) a q where "p = pCons a q" | |
| 177 | proof | |
| 178 | show "p = pCons (coeff p 0) (Abs_poly (\<lambda>n. coeff p (Suc n)))" | |
| 179 | by (rule poly_ext) | |
| 180 | (simp add: Abs_poly_inverse Poly_Suc coeff coeff_pCons | |
| 181 | split: nat.split) | |
| 182 | qed | |
| 183 | ||
| 184 | lemma pCons_induct [case_names 0 pCons, induct type: poly]: | |
| 185 | assumes zero: "P 0" | |
| 186 | assumes pCons: "\<And>a p. P p \<Longrightarrow> P (pCons a p)" | |
| 187 | shows "P p" | |
| 188 | proof (induct p rule: measure_induct_rule [where f=degree]) | |
| 189 | case (less p) | |
| 190 | obtain a q where "p = pCons a q" by (rule pCons_cases) | |
| 191 | have "P q" | |
| 192 | proof (cases "q = 0") | |
| 193 | case True | |
| 194 | then show "P q" by (simp add: zero) | |
| 195 | next | |
| 196 | case False | |
| 197 | then have "degree (pCons a q) = Suc (degree q)" | |
| 198 | by (rule degree_pCons_eq) | |
| 199 | then have "degree q < degree p" | |
| 200 | using `p = pCons a q` by simp | |
| 201 | then show "P q" | |
| 202 | by (rule less.hyps) | |
| 203 | qed | |
| 204 | then have "P (pCons a q)" | |
| 205 | by (rule pCons) | |
| 206 | then show ?case | |
| 207 | using `p = pCons a q` by simp | |
| 208 | qed | |
| 209 | ||
| 210 | ||
| 29454 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 211 | subsection {* Recursion combinator for polynomials *}
 | 
| 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 212 | |
| 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 213 | function | 
| 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 214 |   poly_rec :: "'b \<Rightarrow> ('a::zero \<Rightarrow> 'a poly \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a poly \<Rightarrow> 'b"
 | 
| 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 215 | where | 
| 37765 | 216 | poly_rec_pCons_eq_if [simp del]: | 
| 29454 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 217 | "poly_rec z f (pCons a p) = f a p (if p = 0 then z else poly_rec z f p)" | 
| 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 218 | by (case_tac x, rename_tac q, case_tac q, auto) | 
| 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 219 | |
| 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 220 | termination poly_rec | 
| 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 221 | by (relation "measure (degree \<circ> snd \<circ> snd)", simp) | 
| 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 222 | (simp add: degree_pCons_eq) | 
| 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 223 | |
| 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 224 | lemma poly_rec_0: | 
| 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 225 | "f 0 0 z = z \<Longrightarrow> poly_rec z f 0 = z" | 
| 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 226 | using poly_rec_pCons_eq_if [of z f 0 0] by simp | 
| 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 227 | |
| 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 228 | lemma poly_rec_pCons: | 
| 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 229 | "f 0 0 z = z \<Longrightarrow> poly_rec z f (pCons a p) = f a p (poly_rec z f p)" | 
| 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 230 | by (simp add: poly_rec_pCons_eq_if poly_rec_0) | 
| 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 231 | |
| 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 232 | |
| 29451 | 233 | subsection {* Monomials *}
 | 
| 234 | ||
| 235 | definition | |
| 236 | monom :: "'a \<Rightarrow> nat \<Rightarrow> 'a::zero poly" where | |
| 237 | "monom a m = Abs_poly (\<lambda>n. if m = n then a else 0)" | |
| 238 | ||
| 239 | lemma coeff_monom [simp]: "coeff (monom a m) n = (if m=n then a else 0)" | |
| 240 | unfolding monom_def | |
| 241 | by (subst Abs_poly_inverse, auto simp add: Poly_def) | |
| 242 | ||
| 243 | lemma monom_0: "monom a 0 = pCons a 0" | |
| 244 | by (rule poly_ext, simp add: coeff_pCons split: nat.split) | |
| 245 | ||
| 246 | lemma monom_Suc: "monom a (Suc n) = pCons 0 (monom a n)" | |
| 247 | by (rule poly_ext, simp add: coeff_pCons split: nat.split) | |
| 248 | ||
| 249 | lemma monom_eq_0 [simp]: "monom 0 n = 0" | |
| 250 | by (rule poly_ext) simp | |
| 251 | ||
| 252 | lemma monom_eq_0_iff [simp]: "monom a n = 0 \<longleftrightarrow> a = 0" | |
| 253 | by (simp add: expand_poly_eq) | |
| 254 | ||
| 255 | lemma monom_eq_iff [simp]: "monom a n = monom b n \<longleftrightarrow> a = b" | |
| 256 | by (simp add: expand_poly_eq) | |
| 257 | ||
| 258 | lemma degree_monom_le: "degree (monom a n) \<le> n" | |
| 259 | by (rule degree_le, simp) | |
| 260 | ||
| 261 | lemma degree_monom_eq: "a \<noteq> 0 \<Longrightarrow> degree (monom a n) = n" | |
| 262 | apply (rule order_antisym [OF degree_monom_le]) | |
| 263 | apply (rule le_degree, simp) | |
| 264 | done | |
| 265 | ||
| 266 | ||
| 267 | subsection {* Addition and subtraction *}
 | |
| 268 | ||
| 269 | instantiation poly :: (comm_monoid_add) comm_monoid_add | |
| 270 | begin | |
| 271 | ||
| 272 | definition | |
| 37765 | 273 | plus_poly_def: | 
| 29451 | 274 | "p + q = Abs_poly (\<lambda>n. coeff p n + coeff q n)" | 
| 275 | ||
| 276 | lemma Poly_add: | |
| 277 | fixes f g :: "nat \<Rightarrow> 'a" | |
| 278 | shows "\<lbrakk>f \<in> Poly; g \<in> Poly\<rbrakk> \<Longrightarrow> (\<lambda>n. f n + g n) \<in> Poly" | |
| 279 | unfolding Poly_def | |
| 280 | apply (clarify, rename_tac m n) | |
| 281 | apply (rule_tac x="max m n" in exI, simp) | |
| 282 | done | |
| 283 | ||
| 284 | lemma coeff_add [simp]: | |
| 285 | "coeff (p + q) n = coeff p n + coeff q n" | |
| 286 | unfolding plus_poly_def | |
| 287 | by (simp add: Abs_poly_inverse coeff Poly_add) | |
| 288 | ||
| 289 | instance proof | |
| 290 | fix p q r :: "'a poly" | |
| 291 | show "(p + q) + r = p + (q + r)" | |
| 292 | by (simp add: expand_poly_eq add_assoc) | |
| 293 | show "p + q = q + p" | |
| 294 | by (simp add: expand_poly_eq add_commute) | |
| 295 | show "0 + p = p" | |
| 296 | by (simp add: expand_poly_eq) | |
| 297 | qed | |
| 298 | ||
| 299 | end | |
| 300 | ||
| 29904 | 301 | instance poly :: (cancel_comm_monoid_add) cancel_comm_monoid_add | 
| 29540 | 302 | proof | 
| 303 | fix p q r :: "'a poly" | |
| 304 | assume "p + q = p + r" thus "q = r" | |
| 305 | by (simp add: expand_poly_eq) | |
| 306 | qed | |
| 307 | ||
| 29451 | 308 | instantiation poly :: (ab_group_add) ab_group_add | 
| 309 | begin | |
| 310 | ||
| 311 | definition | |
| 37765 | 312 | uminus_poly_def: | 
| 29451 | 313 | "- p = Abs_poly (\<lambda>n. - coeff p n)" | 
| 314 | ||
| 315 | definition | |
| 37765 | 316 | minus_poly_def: | 
| 29451 | 317 | "p - q = Abs_poly (\<lambda>n. coeff p n - coeff q n)" | 
| 318 | ||
| 319 | lemma Poly_minus: | |
| 320 | fixes f :: "nat \<Rightarrow> 'a" | |
| 321 | shows "f \<in> Poly \<Longrightarrow> (\<lambda>n. - f n) \<in> Poly" | |
| 322 | unfolding Poly_def by simp | |
| 323 | ||
| 324 | lemma Poly_diff: | |
| 325 | fixes f g :: "nat \<Rightarrow> 'a" | |
| 326 | shows "\<lbrakk>f \<in> Poly; g \<in> Poly\<rbrakk> \<Longrightarrow> (\<lambda>n. f n - g n) \<in> Poly" | |
| 327 | unfolding diff_minus by (simp add: Poly_add Poly_minus) | |
| 328 | ||
| 329 | lemma coeff_minus [simp]: "coeff (- p) n = - coeff p n" | |
| 330 | unfolding uminus_poly_def | |
| 331 | by (simp add: Abs_poly_inverse coeff Poly_minus) | |
| 332 | ||
| 333 | lemma coeff_diff [simp]: | |
| 334 | "coeff (p - q) n = coeff p n - coeff q n" | |
| 335 | unfolding minus_poly_def | |
| 336 | by (simp add: Abs_poly_inverse coeff Poly_diff) | |
| 337 | ||
| 338 | instance proof | |
| 339 | fix p q :: "'a poly" | |
| 340 | show "- p + p = 0" | |
| 341 | by (simp add: expand_poly_eq) | |
| 342 | show "p - q = p + - q" | |
| 343 | by (simp add: expand_poly_eq diff_minus) | |
| 344 | qed | |
| 345 | ||
| 346 | end | |
| 347 | ||
| 348 | lemma add_pCons [simp]: | |
| 349 | "pCons a p + pCons b q = pCons (a + b) (p + q)" | |
| 350 | by (rule poly_ext, simp add: coeff_pCons split: nat.split) | |
| 351 | ||
| 352 | lemma minus_pCons [simp]: | |
| 353 | "- pCons a p = pCons (- a) (- p)" | |
| 354 | by (rule poly_ext, simp add: coeff_pCons split: nat.split) | |
| 355 | ||
| 356 | lemma diff_pCons [simp]: | |
| 357 | "pCons a p - pCons b q = pCons (a - b) (p - q)" | |
| 358 | by (rule poly_ext, simp add: coeff_pCons split: nat.split) | |
| 359 | ||
| 29539 | 360 | lemma degree_add_le_max: "degree (p + q) \<le> max (degree p) (degree q)" | 
| 29451 | 361 | by (rule degree_le, auto simp add: coeff_eq_0) | 
| 362 | ||
| 29539 | 363 | lemma degree_add_le: | 
| 364 | "\<lbrakk>degree p \<le> n; degree q \<le> n\<rbrakk> \<Longrightarrow> degree (p + q) \<le> n" | |
| 365 | by (auto intro: order_trans degree_add_le_max) | |
| 366 | ||
| 29453 | 367 | lemma degree_add_less: | 
| 368 | "\<lbrakk>degree p < n; degree q < n\<rbrakk> \<Longrightarrow> degree (p + q) < n" | |
| 29539 | 369 | by (auto intro: le_less_trans degree_add_le_max) | 
| 29453 | 370 | |
| 29451 | 371 | lemma degree_add_eq_right: | 
| 372 | "degree p < degree q \<Longrightarrow> degree (p + q) = degree q" | |
| 373 | apply (cases "q = 0", simp) | |
| 374 | apply (rule order_antisym) | |
| 29539 | 375 | apply (simp add: degree_add_le) | 
| 29451 | 376 | apply (rule le_degree) | 
| 377 | apply (simp add: coeff_eq_0) | |
| 378 | done | |
| 379 | ||
| 380 | lemma degree_add_eq_left: | |
| 381 | "degree q < degree p \<Longrightarrow> degree (p + q) = degree p" | |
| 382 | using degree_add_eq_right [of q p] | |
| 383 | by (simp add: add_commute) | |
| 384 | ||
| 385 | lemma degree_minus [simp]: "degree (- p) = degree p" | |
| 386 | unfolding degree_def by simp | |
| 387 | ||
| 29539 | 388 | lemma degree_diff_le_max: "degree (p - q) \<le> max (degree p) (degree q)" | 
| 29451 | 389 | using degree_add_le [where p=p and q="-q"] | 
| 390 | by (simp add: diff_minus) | |
| 391 | ||
| 29539 | 392 | lemma degree_diff_le: | 
| 393 | "\<lbrakk>degree p \<le> n; degree q \<le> n\<rbrakk> \<Longrightarrow> degree (p - q) \<le> n" | |
| 394 | by (simp add: diff_minus degree_add_le) | |
| 395 | ||
| 29453 | 396 | lemma degree_diff_less: | 
| 397 | "\<lbrakk>degree p < n; degree q < n\<rbrakk> \<Longrightarrow> degree (p - q) < n" | |
| 29539 | 398 | by (simp add: diff_minus degree_add_less) | 
| 29453 | 399 | |
| 29451 | 400 | lemma add_monom: "monom a n + monom b n = monom (a + b) n" | 
| 401 | by (rule poly_ext) simp | |
| 402 | ||
| 403 | lemma diff_monom: "monom a n - monom b n = monom (a - b) n" | |
| 404 | by (rule poly_ext) simp | |
| 405 | ||
| 406 | lemma minus_monom: "- monom a n = monom (-a) n" | |
| 407 | by (rule poly_ext) simp | |
| 408 | ||
| 409 | lemma coeff_setsum: "coeff (\<Sum>x\<in>A. p x) i = (\<Sum>x\<in>A. coeff (p x) i)" | |
| 410 | by (cases "finite A", induct set: finite, simp_all) | |
| 411 | ||
| 412 | lemma monom_setsum: "monom (\<Sum>x\<in>A. a x) n = (\<Sum>x\<in>A. monom (a x) n)" | |
| 413 | by (rule poly_ext) (simp add: coeff_setsum) | |
| 414 | ||
| 415 | ||
| 416 | subsection {* Multiplication by a constant *}
 | |
| 417 | ||
| 418 | definition | |
| 419 | smult :: "'a::comm_semiring_0 \<Rightarrow> 'a poly \<Rightarrow> 'a poly" where | |
| 420 | "smult a p = Abs_poly (\<lambda>n. a * coeff p n)" | |
| 421 | ||
| 422 | lemma Poly_smult: | |
| 423 | fixes f :: "nat \<Rightarrow> 'a::comm_semiring_0" | |
| 424 | shows "f \<in> Poly \<Longrightarrow> (\<lambda>n. a * f n) \<in> Poly" | |
| 425 | unfolding Poly_def | |
| 426 | by (clarify, rule_tac x=n in exI, simp) | |
| 427 | ||
| 428 | lemma coeff_smult [simp]: "coeff (smult a p) n = a * coeff p n" | |
| 429 | unfolding smult_def | |
| 430 | by (simp add: Abs_poly_inverse Poly_smult coeff) | |
| 431 | ||
| 432 | lemma degree_smult_le: "degree (smult a p) \<le> degree p" | |
| 433 | by (rule degree_le, simp add: coeff_eq_0) | |
| 434 | ||
| 29472 | 435 | lemma smult_smult [simp]: "smult a (smult b p) = smult (a * b) p" | 
| 29451 | 436 | by (rule poly_ext, simp add: mult_assoc) | 
| 437 | ||
| 438 | lemma smult_0_right [simp]: "smult a 0 = 0" | |
| 439 | by (rule poly_ext, simp) | |
| 440 | ||
| 441 | lemma smult_0_left [simp]: "smult 0 p = 0" | |
| 442 | by (rule poly_ext, simp) | |
| 443 | ||
| 444 | lemma smult_1_left [simp]: "smult (1::'a::comm_semiring_1) p = p" | |
| 445 | by (rule poly_ext, simp) | |
| 446 | ||
| 447 | lemma smult_add_right: | |
| 448 | "smult a (p + q) = smult a p + smult a q" | |
| 29667 | 449 | by (rule poly_ext, simp add: algebra_simps) | 
| 29451 | 450 | |
| 451 | lemma smult_add_left: | |
| 452 | "smult (a + b) p = smult a p + smult b p" | |
| 29667 | 453 | by (rule poly_ext, simp add: algebra_simps) | 
| 29451 | 454 | |
| 29457 | 455 | lemma smult_minus_right [simp]: | 
| 29451 | 456 | "smult (a::'a::comm_ring) (- p) = - smult a p" | 
| 457 | by (rule poly_ext, simp) | |
| 458 | ||
| 29457 | 459 | lemma smult_minus_left [simp]: | 
| 29451 | 460 | "smult (- a::'a::comm_ring) p = - smult a p" | 
| 461 | by (rule poly_ext, simp) | |
| 462 | ||
| 463 | lemma smult_diff_right: | |
| 464 | "smult (a::'a::comm_ring) (p - q) = smult a p - smult a q" | |
| 29667 | 465 | by (rule poly_ext, simp add: algebra_simps) | 
| 29451 | 466 | |
| 467 | lemma smult_diff_left: | |
| 468 | "smult (a - b::'a::comm_ring) p = smult a p - smult b p" | |
| 29667 | 469 | by (rule poly_ext, simp add: algebra_simps) | 
| 29451 | 470 | |
| 29472 | 471 | lemmas smult_distribs = | 
| 472 | smult_add_left smult_add_right | |
| 473 | smult_diff_left smult_diff_right | |
| 474 | ||
| 29451 | 475 | lemma smult_pCons [simp]: | 
| 476 | "smult a (pCons b p) = pCons (a * b) (smult a p)" | |
| 477 | by (rule poly_ext, simp add: coeff_pCons split: nat.split) | |
| 478 | ||
| 479 | lemma smult_monom: "smult a (monom b n) = monom (a * b) n" | |
| 480 | by (induct n, simp add: monom_0, simp add: monom_Suc) | |
| 481 | ||
| 29659 | 482 | lemma degree_smult_eq [simp]: | 
| 483 | fixes a :: "'a::idom" | |
| 484 | shows "degree (smult a p) = (if a = 0 then 0 else degree p)" | |
| 485 | by (cases "a = 0", simp, simp add: degree_def) | |
| 486 | ||
| 487 | lemma smult_eq_0_iff [simp]: | |
| 488 | fixes a :: "'a::idom" | |
| 489 | shows "smult a p = 0 \<longleftrightarrow> a = 0 \<or> p = 0" | |
| 490 | by (simp add: expand_poly_eq) | |
| 491 | ||
| 29451 | 492 | |
| 493 | subsection {* Multiplication of polynomials *}
 | |
| 494 | ||
| 47382 | 495 | (* TODO: move to Set_Interval.thy *) | 
| 29451 | 496 | lemma setsum_atMost_Suc_shift: | 
| 497 | fixes f :: "nat \<Rightarrow> 'a::comm_monoid_add" | |
| 498 | shows "(\<Sum>i\<le>Suc n. f i) = f 0 + (\<Sum>i\<le>n. f (Suc i))" | |
| 499 | proof (induct n) | |
| 500 | case 0 show ?case by simp | |
| 501 | next | |
| 502 | case (Suc n) note IH = this | |
| 503 | have "(\<Sum>i\<le>Suc (Suc n). f i) = (\<Sum>i\<le>Suc n. f i) + f (Suc (Suc n))" | |
| 504 | by (rule setsum_atMost_Suc) | |
| 505 | also have "(\<Sum>i\<le>Suc n. f i) = f 0 + (\<Sum>i\<le>n. f (Suc i))" | |
| 506 | by (rule IH) | |
| 507 | also have "f 0 + (\<Sum>i\<le>n. f (Suc i)) + f (Suc (Suc n)) = | |
| 508 | f 0 + ((\<Sum>i\<le>n. f (Suc i)) + f (Suc (Suc n)))" | |
| 509 | by (rule add_assoc) | |
| 510 | also have "(\<Sum>i\<le>n. f (Suc i)) + f (Suc (Suc n)) = (\<Sum>i\<le>Suc n. f (Suc i))" | |
| 511 | by (rule setsum_atMost_Suc [symmetric]) | |
| 512 | finally show ?case . | |
| 513 | qed | |
| 514 | ||
| 515 | instantiation poly :: (comm_semiring_0) comm_semiring_0 | |
| 516 | begin | |
| 517 | ||
| 518 | definition | |
| 37765 | 519 | times_poly_def: | 
| 29474 | 520 | "p * q = poly_rec 0 (\<lambda>a p pq. smult a q + pCons 0 pq) p" | 
| 521 | ||
| 522 | lemma mult_poly_0_left: "(0::'a poly) * q = 0" | |
| 523 | unfolding times_poly_def by (simp add: poly_rec_0) | |
| 524 | ||
| 525 | lemma mult_pCons_left [simp]: | |
| 526 | "pCons a p * q = smult a q + pCons 0 (p * q)" | |
| 527 | unfolding times_poly_def by (simp add: poly_rec_pCons) | |
| 528 | ||
| 529 | lemma mult_poly_0_right: "p * (0::'a poly) = 0" | |
| 530 | by (induct p, simp add: mult_poly_0_left, simp) | |
| 29451 | 531 | |
| 29474 | 532 | lemma mult_pCons_right [simp]: | 
| 533 | "p * pCons a q = smult a p + pCons 0 (p * q)" | |
| 29667 | 534 | by (induct p, simp add: mult_poly_0_left, simp add: algebra_simps) | 
| 29474 | 535 | |
| 536 | lemmas mult_poly_0 = mult_poly_0_left mult_poly_0_right | |
| 537 | ||
| 538 | lemma mult_smult_left [simp]: "smult a p * q = smult a (p * q)" | |
| 539 | by (induct p, simp add: mult_poly_0, simp add: smult_add_right) | |
| 540 | ||
| 541 | lemma mult_smult_right [simp]: "p * smult a q = smult a (p * q)" | |
| 542 | by (induct q, simp add: mult_poly_0, simp add: smult_add_right) | |
| 543 | ||
| 544 | lemma mult_poly_add_left: | |
| 545 | fixes p q r :: "'a poly" | |
| 546 | shows "(p + q) * r = p * r + q * r" | |
| 547 | by (induct r, simp add: mult_poly_0, | |
| 29667 | 548 | simp add: smult_distribs algebra_simps) | 
| 29451 | 549 | |
| 550 | instance proof | |
| 551 | fix p q r :: "'a poly" | |
| 552 | show 0: "0 * p = 0" | |
| 29474 | 553 | by (rule mult_poly_0_left) | 
| 29451 | 554 | show "p * 0 = 0" | 
| 29474 | 555 | by (rule mult_poly_0_right) | 
| 29451 | 556 | show "(p + q) * r = p * r + q * r" | 
| 29474 | 557 | by (rule mult_poly_add_left) | 
| 29451 | 558 | show "(p * q) * r = p * (q * r)" | 
| 29474 | 559 | by (induct p, simp add: mult_poly_0, simp add: mult_poly_add_left) | 
| 29451 | 560 | show "p * q = q * p" | 
| 29474 | 561 | by (induct p, simp add: mult_poly_0, simp) | 
| 29451 | 562 | qed | 
| 563 | ||
| 564 | end | |
| 565 | ||
| 29540 | 566 | instance poly :: (comm_semiring_0_cancel) comm_semiring_0_cancel .. | 
| 567 | ||
| 29474 | 568 | lemma coeff_mult: | 
| 569 | "coeff (p * q) n = (\<Sum>i\<le>n. coeff p i * coeff q (n-i))" | |
| 570 | proof (induct p arbitrary: n) | |
| 571 | case 0 show ?case by simp | |
| 572 | next | |
| 573 | case (pCons a p n) thus ?case | |
| 574 | by (cases n, simp, simp add: setsum_atMost_Suc_shift | |
| 575 | del: setsum_atMost_Suc) | |
| 576 | qed | |
| 29451 | 577 | |
| 29474 | 578 | lemma degree_mult_le: "degree (p * q) \<le> degree p + degree q" | 
| 579 | apply (rule degree_le) | |
| 580 | apply (induct p) | |
| 581 | apply simp | |
| 582 | apply (simp add: coeff_eq_0 coeff_pCons split: nat.split) | |
| 29451 | 583 | done | 
| 584 | ||
| 585 | lemma mult_monom: "monom a m * monom b n = monom (a * b) (m + n)" | |
| 586 | by (induct m, simp add: monom_0 smult_monom, simp add: monom_Suc) | |
| 587 | ||
| 588 | ||
| 589 | subsection {* The unit polynomial and exponentiation *}
 | |
| 590 | ||
| 591 | instantiation poly :: (comm_semiring_1) comm_semiring_1 | |
| 592 | begin | |
| 593 | ||
| 594 | definition | |
| 595 | one_poly_def: | |
| 596 | "1 = pCons 1 0" | |
| 597 | ||
| 598 | instance proof | |
| 599 | fix p :: "'a poly" show "1 * p = p" | |
| 600 | unfolding one_poly_def | |
| 601 | by simp | |
| 602 | next | |
| 603 | show "0 \<noteq> (1::'a poly)" | |
| 604 | unfolding one_poly_def by simp | |
| 605 | qed | |
| 606 | ||
| 607 | end | |
| 608 | ||
| 29540 | 609 | instance poly :: (comm_semiring_1_cancel) comm_semiring_1_cancel .. | 
| 610 | ||
| 29451 | 611 | lemma coeff_1 [simp]: "coeff 1 n = (if n = 0 then 1 else 0)" | 
| 612 | unfolding one_poly_def | |
| 613 | by (simp add: coeff_pCons split: nat.split) | |
| 614 | ||
| 615 | lemma degree_1 [simp]: "degree 1 = 0" | |
| 616 | unfolding one_poly_def | |
| 617 | by (rule degree_pCons_0) | |
| 618 | ||
| 29979 | 619 | text {* Lemmas about divisibility *}
 | 
| 620 | ||
| 621 | lemma dvd_smult: "p dvd q \<Longrightarrow> p dvd smult a q" | |
| 622 | proof - | |
| 623 | assume "p dvd q" | |
| 624 | then obtain k where "q = p * k" .. | |
| 625 | then have "smult a q = p * smult a k" by simp | |
| 626 | then show "p dvd smult a q" .. | |
| 627 | qed | |
| 628 | ||
| 629 | lemma dvd_smult_cancel: | |
| 630 | fixes a :: "'a::field" | |
| 631 | shows "p dvd smult a q \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> p dvd q" | |
| 632 | by (drule dvd_smult [where a="inverse a"]) simp | |
| 633 | ||
| 634 | lemma dvd_smult_iff: | |
| 635 | fixes a :: "'a::field" | |
| 636 | shows "a \<noteq> 0 \<Longrightarrow> p dvd smult a q \<longleftrightarrow> p dvd q" | |
| 637 | by (safe elim!: dvd_smult dvd_smult_cancel) | |
| 638 | ||
| 31663 | 639 | lemma smult_dvd_cancel: | 
| 640 | "smult a p dvd q \<Longrightarrow> p dvd q" | |
| 641 | proof - | |
| 642 | assume "smult a p dvd q" | |
| 643 | then obtain k where "q = smult a p * k" .. | |
| 644 | then have "q = p * smult a k" by simp | |
| 645 | then show "p dvd q" .. | |
| 646 | qed | |
| 647 | ||
| 648 | lemma smult_dvd: | |
| 649 | fixes a :: "'a::field" | |
| 650 | shows "p dvd q \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> smult a p dvd q" | |
| 651 | by (rule smult_dvd_cancel [where a="inverse a"]) simp | |
| 652 | ||
| 653 | lemma smult_dvd_iff: | |
| 654 | fixes a :: "'a::field" | |
| 655 | shows "smult a p dvd q \<longleftrightarrow> (if a = 0 then q = 0 else p dvd q)" | |
| 656 | by (auto elim: smult_dvd smult_dvd_cancel) | |
| 657 | ||
| 29979 | 658 | lemma degree_power_le: "degree (p ^ n) \<le> degree p * n" | 
| 659 | by (induct n, simp, auto intro: order_trans degree_mult_le) | |
| 660 | ||
| 29451 | 661 | instance poly :: (comm_ring) comm_ring .. | 
| 662 | ||
| 663 | instance poly :: (comm_ring_1) comm_ring_1 .. | |
| 664 | ||
| 665 | ||
| 666 | subsection {* Polynomials form an integral domain *}
 | |
| 667 | ||
| 668 | lemma coeff_mult_degree_sum: | |
| 669 | "coeff (p * q) (degree p + degree q) = | |
| 670 | coeff p (degree p) * coeff q (degree q)" | |
| 29471 | 671 | by (induct p, simp, simp add: coeff_eq_0) | 
| 29451 | 672 | |
| 673 | instance poly :: (idom) idom | |
| 674 | proof | |
| 675 | fix p q :: "'a poly" | |
| 676 | assume "p \<noteq> 0" and "q \<noteq> 0" | |
| 677 | have "coeff (p * q) (degree p + degree q) = | |
| 678 | coeff p (degree p) * coeff q (degree q)" | |
| 679 | by (rule coeff_mult_degree_sum) | |
| 680 | also have "coeff p (degree p) * coeff q (degree q) \<noteq> 0" | |
| 681 | using `p \<noteq> 0` and `q \<noteq> 0` by simp | |
| 682 | finally have "\<exists>n. coeff (p * q) n \<noteq> 0" .. | |
| 683 | thus "p * q \<noteq> 0" by (simp add: expand_poly_eq) | |
| 684 | qed | |
| 685 | ||
| 686 | lemma degree_mult_eq: | |
| 687 | fixes p q :: "'a::idom poly" | |
| 688 | shows "\<lbrakk>p \<noteq> 0; q \<noteq> 0\<rbrakk> \<Longrightarrow> degree (p * q) = degree p + degree q" | |
| 689 | apply (rule order_antisym [OF degree_mult_le le_degree]) | |
| 690 | apply (simp add: coeff_mult_degree_sum) | |
| 691 | done | |
| 692 | ||
| 693 | lemma dvd_imp_degree_le: | |
| 694 | fixes p q :: "'a::idom poly" | |
| 695 | shows "\<lbrakk>p dvd q; q \<noteq> 0\<rbrakk> \<Longrightarrow> degree p \<le> degree q" | |
| 696 | by (erule dvdE, simp add: degree_mult_eq) | |
| 697 | ||
| 698 | ||
| 29878 | 699 | subsection {* Polynomials form an ordered integral domain *}
 | 
| 700 | ||
| 701 | definition | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34973diff
changeset | 702 | pos_poly :: "'a::linordered_idom poly \<Rightarrow> bool" | 
| 29878 | 703 | where | 
| 704 | "pos_poly p \<longleftrightarrow> 0 < coeff p (degree p)" | |
| 705 | ||
| 706 | lemma pos_poly_pCons: | |
| 707 | "pos_poly (pCons a p) \<longleftrightarrow> pos_poly p \<or> (p = 0 \<and> 0 < a)" | |
| 708 | unfolding pos_poly_def by simp | |
| 709 | ||
| 710 | lemma not_pos_poly_0 [simp]: "\<not> pos_poly 0" | |
| 711 | unfolding pos_poly_def by simp | |
| 712 | ||
| 713 | lemma pos_poly_add: "\<lbrakk>pos_poly p; pos_poly q\<rbrakk> \<Longrightarrow> pos_poly (p + q)" | |
| 714 | apply (induct p arbitrary: q, simp) | |
| 715 | apply (case_tac q, force simp add: pos_poly_pCons add_pos_pos) | |
| 716 | done | |
| 717 | ||
| 718 | lemma pos_poly_mult: "\<lbrakk>pos_poly p; pos_poly q\<rbrakk> \<Longrightarrow> pos_poly (p * q)" | |
| 719 | unfolding pos_poly_def | |
| 720 | apply (subgoal_tac "p \<noteq> 0 \<and> q \<noteq> 0") | |
| 721 | apply (simp add: degree_mult_eq coeff_mult_degree_sum mult_pos_pos) | |
| 722 | apply auto | |
| 723 | done | |
| 724 | ||
| 725 | lemma pos_poly_total: "p = 0 \<or> pos_poly p \<or> pos_poly (- p)" | |
| 726 | by (induct p) (auto simp add: pos_poly_pCons) | |
| 727 | ||
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34973diff
changeset | 728 | instantiation poly :: (linordered_idom) linordered_idom | 
| 29878 | 729 | begin | 
| 730 | ||
| 731 | definition | |
| 37765 | 732 | "x < y \<longleftrightarrow> pos_poly (y - x)" | 
| 29878 | 733 | |
| 734 | definition | |
| 37765 | 735 | "x \<le> y \<longleftrightarrow> x = y \<or> pos_poly (y - x)" | 
| 29878 | 736 | |
| 737 | definition | |
| 37765 | 738 | "abs (x::'a poly) = (if x < 0 then - x else x)" | 
| 29878 | 739 | |
| 740 | definition | |
| 37765 | 741 | "sgn (x::'a poly) = (if x = 0 then 0 else if 0 < x then 1 else - 1)" | 
| 29878 | 742 | |
| 743 | instance proof | |
| 744 | fix x y :: "'a poly" | |
| 745 | show "x < y \<longleftrightarrow> x \<le> y \<and> \<not> y \<le> x" | |
| 746 | unfolding less_eq_poly_def less_poly_def | |
| 747 | apply safe | |
| 748 | apply simp | |
| 749 | apply (drule (1) pos_poly_add) | |
| 750 | apply simp | |
| 751 | done | |
| 752 | next | |
| 753 | fix x :: "'a poly" show "x \<le> x" | |
| 754 | unfolding less_eq_poly_def by simp | |
| 755 | next | |
| 756 | fix x y z :: "'a poly" | |
| 757 | assume "x \<le> y" and "y \<le> z" thus "x \<le> z" | |
| 758 | unfolding less_eq_poly_def | |
| 759 | apply safe | |
| 760 | apply (drule (1) pos_poly_add) | |
| 761 | apply (simp add: algebra_simps) | |
| 762 | done | |
| 763 | next | |
| 764 | fix x y :: "'a poly" | |
| 765 | assume "x \<le> y" and "y \<le> x" thus "x = y" | |
| 766 | unfolding less_eq_poly_def | |
| 767 | apply safe | |
| 768 | apply (drule (1) pos_poly_add) | |
| 769 | apply simp | |
| 770 | done | |
| 771 | next | |
| 772 | fix x y z :: "'a poly" | |
| 773 | assume "x \<le> y" thus "z + x \<le> z + y" | |
| 774 | unfolding less_eq_poly_def | |
| 775 | apply safe | |
| 776 | apply (simp add: algebra_simps) | |
| 777 | done | |
| 778 | next | |
| 779 | fix x y :: "'a poly" | |
| 780 | show "x \<le> y \<or> y \<le> x" | |
| 781 | unfolding less_eq_poly_def | |
| 782 | using pos_poly_total [of "x - y"] | |
| 783 | by auto | |
| 784 | next | |
| 785 | fix x y z :: "'a poly" | |
| 786 | assume "x < y" and "0 < z" | |
| 787 | thus "z * x < z * y" | |
| 788 | unfolding less_poly_def | |
| 789 | by (simp add: right_diff_distrib [symmetric] pos_poly_mult) | |
| 790 | next | |
| 791 | fix x :: "'a poly" | |
| 792 | show "\<bar>x\<bar> = (if x < 0 then - x else x)" | |
| 793 | by (rule abs_poly_def) | |
| 794 | next | |
| 795 | fix x :: "'a poly" | |
| 796 | show "sgn x = (if x = 0 then 0 else if 0 < x then 1 else - 1)" | |
| 797 | by (rule sgn_poly_def) | |
| 798 | qed | |
| 799 | ||
| 800 | end | |
| 801 | ||
| 802 | text {* TODO: Simplification rules for comparisons *}
 | |
| 803 | ||
| 804 | ||
| 29451 | 805 | subsection {* Long division of polynomials *}
 | 
| 806 | ||
| 807 | definition | |
| 29537 | 808 | pdivmod_rel :: "'a::field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<Rightarrow> bool" | 
| 29451 | 809 | where | 
| 29537 | 810 | "pdivmod_rel x y q r \<longleftrightarrow> | 
| 29451 | 811 | x = q * y + r \<and> (if y = 0 then q = 0 else r = 0 \<or> degree r < degree y)" | 
| 812 | ||
| 29537 | 813 | lemma pdivmod_rel_0: | 
| 814 | "pdivmod_rel 0 y 0 0" | |
| 815 | unfolding pdivmod_rel_def by simp | |
| 29451 | 816 | |
| 29537 | 817 | lemma pdivmod_rel_by_0: | 
| 818 | "pdivmod_rel x 0 0 x" | |
| 819 | unfolding pdivmod_rel_def by simp | |
| 29451 | 820 | |
| 821 | lemma eq_zero_or_degree_less: | |
| 822 | assumes "degree p \<le> n" and "coeff p n = 0" | |
| 823 | shows "p = 0 \<or> degree p < n" | |
| 824 | proof (cases n) | |
| 825 | case 0 | |
| 826 | with `degree p \<le> n` and `coeff p n = 0` | |
| 827 | have "coeff p (degree p) = 0" by simp | |
| 828 | then have "p = 0" by simp | |
| 829 | then show ?thesis .. | |
| 830 | next | |
| 831 | case (Suc m) | |
| 832 | have "\<forall>i>n. coeff p i = 0" | |
| 833 | using `degree p \<le> n` by (simp add: coeff_eq_0) | |
| 834 | then have "\<forall>i\<ge>n. coeff p i = 0" | |
| 835 | using `coeff p n = 0` by (simp add: le_less) | |
| 836 | then have "\<forall>i>m. coeff p i = 0" | |
| 837 | using `n = Suc m` by (simp add: less_eq_Suc_le) | |
| 838 | then have "degree p \<le> m" | |
| 839 | by (rule degree_le) | |
| 840 | then have "degree p < n" | |
| 841 | using `n = Suc m` by (simp add: less_Suc_eq_le) | |
| 842 | then show ?thesis .. | |
| 843 | qed | |
| 844 | ||
| 29537 | 845 | lemma pdivmod_rel_pCons: | 
| 846 | assumes rel: "pdivmod_rel x y q r" | |
| 29451 | 847 | assumes y: "y \<noteq> 0" | 
| 848 | assumes b: "b = coeff (pCons a r) (degree y) / coeff y (degree y)" | |
| 29537 | 849 | shows "pdivmod_rel (pCons a x) y (pCons b q) (pCons a r - smult b y)" | 
| 850 | (is "pdivmod_rel ?x y ?q ?r") | |
| 29451 | 851 | proof - | 
| 852 | have x: "x = q * y + r" and r: "r = 0 \<or> degree r < degree y" | |
| 29537 | 853 | using assms unfolding pdivmod_rel_def by simp_all | 
| 29451 | 854 | |
| 855 | have 1: "?x = ?q * y + ?r" | |
| 856 | using b x by simp | |
| 857 | ||
| 858 | have 2: "?r = 0 \<or> degree ?r < degree y" | |
| 859 | proof (rule eq_zero_or_degree_less) | |
| 29539 | 860 | show "degree ?r \<le> degree y" | 
| 861 | proof (rule degree_diff_le) | |
| 29451 | 862 | show "degree (pCons a r) \<le> degree y" | 
| 29460 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 863 | using r by auto | 
| 29451 | 864 | show "degree (smult b y) \<le> degree y" | 
| 865 | by (rule degree_smult_le) | |
| 866 | qed | |
| 867 | next | |
| 868 | show "coeff ?r (degree y) = 0" | |
| 869 | using `y \<noteq> 0` unfolding b by simp | |
| 870 | qed | |
| 871 | ||
| 872 | from 1 2 show ?thesis | |
| 29537 | 873 | unfolding pdivmod_rel_def | 
| 29451 | 874 | using `y \<noteq> 0` by simp | 
| 875 | qed | |
| 876 | ||
| 29537 | 877 | lemma pdivmod_rel_exists: "\<exists>q r. pdivmod_rel x y q r" | 
| 29451 | 878 | apply (cases "y = 0") | 
| 29537 | 879 | apply (fast intro!: pdivmod_rel_by_0) | 
| 29451 | 880 | apply (induct x) | 
| 29537 | 881 | apply (fast intro!: pdivmod_rel_0) | 
| 882 | apply (fast intro!: pdivmod_rel_pCons) | |
| 29451 | 883 | done | 
| 884 | ||
| 29537 | 885 | lemma pdivmod_rel_unique: | 
| 886 | assumes 1: "pdivmod_rel x y q1 r1" | |
| 887 | assumes 2: "pdivmod_rel x y q2 r2" | |
| 29451 | 888 | shows "q1 = q2 \<and> r1 = r2" | 
| 889 | proof (cases "y = 0") | |
| 890 | assume "y = 0" with assms show ?thesis | |
| 29537 | 891 | by (simp add: pdivmod_rel_def) | 
| 29451 | 892 | next | 
| 893 | assume [simp]: "y \<noteq> 0" | |
| 894 | from 1 have q1: "x = q1 * y + r1" and r1: "r1 = 0 \<or> degree r1 < degree y" | |
| 29537 | 895 | unfolding pdivmod_rel_def by simp_all | 
| 29451 | 896 | from 2 have q2: "x = q2 * y + r2" and r2: "r2 = 0 \<or> degree r2 < degree y" | 
| 29537 | 897 | unfolding pdivmod_rel_def by simp_all | 
| 29451 | 898 | from q1 q2 have q3: "(q1 - q2) * y = r2 - r1" | 
| 29667 | 899 | by (simp add: algebra_simps) | 
| 29451 | 900 | from r1 r2 have r3: "(r2 - r1) = 0 \<or> degree (r2 - r1) < degree y" | 
| 29453 | 901 | by (auto intro: degree_diff_less) | 
| 29451 | 902 | |
| 903 | show "q1 = q2 \<and> r1 = r2" | |
| 904 | proof (rule ccontr) | |
| 905 | assume "\<not> (q1 = q2 \<and> r1 = r2)" | |
| 906 | with q3 have "q1 \<noteq> q2" and "r1 \<noteq> r2" by auto | |
| 907 | with r3 have "degree (r2 - r1) < degree y" by simp | |
| 908 | also have "degree y \<le> degree (q1 - q2) + degree y" by simp | |
| 909 | also have "\<dots> = degree ((q1 - q2) * y)" | |
| 910 | using `q1 \<noteq> q2` by (simp add: degree_mult_eq) | |
| 911 | also have "\<dots> = degree (r2 - r1)" | |
| 912 | using q3 by simp | |
| 913 | finally have "degree (r2 - r1) < degree (r2 - r1)" . | |
| 914 | then show "False" by simp | |
| 915 | qed | |
| 916 | qed | |
| 917 | ||
| 29660 | 918 | lemma pdivmod_rel_0_iff: "pdivmod_rel 0 y q r \<longleftrightarrow> q = 0 \<and> r = 0" | 
| 919 | by (auto dest: pdivmod_rel_unique intro: pdivmod_rel_0) | |
| 920 | ||
| 921 | lemma pdivmod_rel_by_0_iff: "pdivmod_rel x 0 q r \<longleftrightarrow> q = 0 \<and> r = x" | |
| 922 | by (auto dest: pdivmod_rel_unique intro: pdivmod_rel_by_0) | |
| 923 | ||
| 45605 | 924 | lemmas pdivmod_rel_unique_div = pdivmod_rel_unique [THEN conjunct1] | 
| 29451 | 925 | |
| 45605 | 926 | lemmas pdivmod_rel_unique_mod = pdivmod_rel_unique [THEN conjunct2] | 
| 29451 | 927 | |
| 928 | instantiation poly :: (field) ring_div | |
| 929 | begin | |
| 930 | ||
| 931 | definition div_poly where | |
| 37765 | 932 | "x div y = (THE q. \<exists>r. pdivmod_rel x y q r)" | 
| 29451 | 933 | |
| 934 | definition mod_poly where | |
| 37765 | 935 | "x mod y = (THE r. \<exists>q. pdivmod_rel x y q r)" | 
| 29451 | 936 | |
| 937 | lemma div_poly_eq: | |
| 29537 | 938 | "pdivmod_rel x y q r \<Longrightarrow> x div y = q" | 
| 29451 | 939 | unfolding div_poly_def | 
| 29537 | 940 | by (fast elim: pdivmod_rel_unique_div) | 
| 29451 | 941 | |
| 942 | lemma mod_poly_eq: | |
| 29537 | 943 | "pdivmod_rel x y q r \<Longrightarrow> x mod y = r" | 
| 29451 | 944 | unfolding mod_poly_def | 
| 29537 | 945 | by (fast elim: pdivmod_rel_unique_mod) | 
| 29451 | 946 | |
| 29537 | 947 | lemma pdivmod_rel: | 
| 948 | "pdivmod_rel x y (x div y) (x mod y)" | |
| 29451 | 949 | proof - | 
| 29537 | 950 | from pdivmod_rel_exists | 
| 951 | obtain q r where "pdivmod_rel x y q r" by fast | |
| 29451 | 952 | thus ?thesis | 
| 953 | by (simp add: div_poly_eq mod_poly_eq) | |
| 954 | qed | |
| 955 | ||
| 956 | instance proof | |
| 957 | fix x y :: "'a poly" | |
| 958 | show "x div y * y + x mod y = x" | |
| 29537 | 959 | using pdivmod_rel [of x y] | 
| 960 | by (simp add: pdivmod_rel_def) | |
| 29451 | 961 | next | 
| 962 | fix x :: "'a poly" | |
| 29537 | 963 | have "pdivmod_rel x 0 0 x" | 
| 964 | by (rule pdivmod_rel_by_0) | |
| 29451 | 965 | thus "x div 0 = 0" | 
| 966 | by (rule div_poly_eq) | |
| 967 | next | |
| 968 | fix y :: "'a poly" | |
| 29537 | 969 | have "pdivmod_rel 0 y 0 0" | 
| 970 | by (rule pdivmod_rel_0) | |
| 29451 | 971 | thus "0 div y = 0" | 
| 972 | by (rule div_poly_eq) | |
| 973 | next | |
| 974 | fix x y z :: "'a poly" | |
| 975 | assume "y \<noteq> 0" | |
| 29537 | 976 | hence "pdivmod_rel (x + z * y) y (z + x div y) (x mod y)" | 
| 977 | using pdivmod_rel [of x y] | |
| 978 | by (simp add: pdivmod_rel_def left_distrib) | |
| 29451 | 979 | thus "(x + z * y) div y = z + x div y" | 
| 980 | by (rule div_poly_eq) | |
| 30930 | 981 | next | 
| 982 | fix x y z :: "'a poly" | |
| 983 | assume "x \<noteq> 0" | |
| 984 | show "(x * y) div (x * z) = y div z" | |
| 985 | proof (cases "y \<noteq> 0 \<and> z \<noteq> 0") | |
| 986 | have "\<And>x::'a poly. pdivmod_rel x 0 0 x" | |
| 987 | by (rule pdivmod_rel_by_0) | |
| 988 | then have [simp]: "\<And>x::'a poly. x div 0 = 0" | |
| 989 | by (rule div_poly_eq) | |
| 990 | have "\<And>x::'a poly. pdivmod_rel 0 x 0 0" | |
| 991 | by (rule pdivmod_rel_0) | |
| 992 | then have [simp]: "\<And>x::'a poly. 0 div x = 0" | |
| 993 | by (rule div_poly_eq) | |
| 994 | case False then show ?thesis by auto | |
| 995 | next | |
| 996 | case True then have "y \<noteq> 0" and "z \<noteq> 0" by auto | |
| 997 | with `x \<noteq> 0` | |
| 998 | have "\<And>q r. pdivmod_rel y z q r \<Longrightarrow> pdivmod_rel (x * y) (x * z) q (x * r)" | |
| 999 | by (auto simp add: pdivmod_rel_def algebra_simps) | |
| 1000 | (rule classical, simp add: degree_mult_eq) | |
| 1001 | moreover from pdivmod_rel have "pdivmod_rel y z (y div z) (y mod z)" . | |
| 1002 | ultimately have "pdivmod_rel (x * y) (x * z) (y div z) (x * (y mod z))" . | |
| 1003 | then show ?thesis by (simp add: div_poly_eq) | |
| 1004 | qed | |
| 29451 | 1005 | qed | 
| 1006 | ||
| 1007 | end | |
| 1008 | ||
| 1009 | lemma degree_mod_less: | |
| 1010 | "y \<noteq> 0 \<Longrightarrow> x mod y = 0 \<or> degree (x mod y) < degree y" | |
| 29537 | 1011 | using pdivmod_rel [of x y] | 
| 1012 | unfolding pdivmod_rel_def by simp | |
| 29451 | 1013 | |
| 1014 | lemma div_poly_less: "degree x < degree y \<Longrightarrow> x div y = 0" | |
| 1015 | proof - | |
| 1016 | assume "degree x < degree y" | |
| 29537 | 1017 | hence "pdivmod_rel x y 0 x" | 
| 1018 | by (simp add: pdivmod_rel_def) | |
| 29451 | 1019 | thus "x div y = 0" by (rule div_poly_eq) | 
| 1020 | qed | |
| 1021 | ||
| 1022 | lemma mod_poly_less: "degree x < degree y \<Longrightarrow> x mod y = x" | |
| 1023 | proof - | |
| 1024 | assume "degree x < degree y" | |
| 29537 | 1025 | hence "pdivmod_rel x y 0 x" | 
| 1026 | by (simp add: pdivmod_rel_def) | |
| 29451 | 1027 | thus "x mod y = x" by (rule mod_poly_eq) | 
| 1028 | qed | |
| 1029 | ||
| 29659 | 1030 | lemma pdivmod_rel_smult_left: | 
| 1031 | "pdivmod_rel x y q r | |
| 1032 | \<Longrightarrow> pdivmod_rel (smult a x) y (smult a q) (smult a r)" | |
| 1033 | unfolding pdivmod_rel_def by (simp add: smult_add_right) | |
| 1034 | ||
| 1035 | lemma div_smult_left: "(smult a x) div y = smult a (x div y)" | |
| 1036 | by (rule div_poly_eq, rule pdivmod_rel_smult_left, rule pdivmod_rel) | |
| 1037 | ||
| 1038 | lemma mod_smult_left: "(smult a x) mod y = smult a (x mod y)" | |
| 1039 | by (rule mod_poly_eq, rule pdivmod_rel_smult_left, rule pdivmod_rel) | |
| 1040 | ||
| 30072 | 1041 | lemma poly_div_minus_left [simp]: | 
| 1042 | fixes x y :: "'a::field poly" | |
| 1043 | shows "(- x) div y = - (x div y)" | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
47002diff
changeset | 1044 | using div_smult_left [of "- 1::'a"] by (simp del: minus_one) (* FIXME *) | 
| 30072 | 1045 | |
| 1046 | lemma poly_mod_minus_left [simp]: | |
| 1047 | fixes x y :: "'a::field poly" | |
| 1048 | shows "(- x) mod y = - (x mod y)" | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
47002diff
changeset | 1049 | using mod_smult_left [of "- 1::'a"] by (simp del: minus_one) (* FIXME *) | 
| 30072 | 1050 | |
| 29659 | 1051 | lemma pdivmod_rel_smult_right: | 
| 1052 | "\<lbrakk>a \<noteq> 0; pdivmod_rel x y q r\<rbrakk> | |
| 1053 | \<Longrightarrow> pdivmod_rel x (smult a y) (smult (inverse a) q) r" | |
| 1054 | unfolding pdivmod_rel_def by simp | |
| 1055 | ||
| 1056 | lemma div_smult_right: | |
| 1057 | "a \<noteq> 0 \<Longrightarrow> x div (smult a y) = smult (inverse a) (x div y)" | |
| 1058 | by (rule div_poly_eq, erule pdivmod_rel_smult_right, rule pdivmod_rel) | |
| 1059 | ||
| 1060 | lemma mod_smult_right: "a \<noteq> 0 \<Longrightarrow> x mod (smult a y) = x mod y" | |
| 1061 | by (rule mod_poly_eq, erule pdivmod_rel_smult_right, rule pdivmod_rel) | |
| 1062 | ||
| 30072 | 1063 | lemma poly_div_minus_right [simp]: | 
| 1064 | fixes x y :: "'a::field poly" | |
| 1065 | shows "x div (- y) = - (x div y)" | |
| 1066 | using div_smult_right [of "- 1::'a"] | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
47002diff
changeset | 1067 | by (simp add: nonzero_inverse_minus_eq del: minus_one) (* FIXME *) | 
| 30072 | 1068 | |
| 1069 | lemma poly_mod_minus_right [simp]: | |
| 1070 | fixes x y :: "'a::field poly" | |
| 1071 | shows "x mod (- y) = x mod y" | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
47002diff
changeset | 1072 | using mod_smult_right [of "- 1::'a"] by (simp del: minus_one) (* FIXME *) | 
| 30072 | 1073 | |
| 29660 | 1074 | lemma pdivmod_rel_mult: | 
| 1075 | "\<lbrakk>pdivmod_rel x y q r; pdivmod_rel q z q' r'\<rbrakk> | |
| 1076 | \<Longrightarrow> pdivmod_rel x (y * z) q' (y * r' + r)" | |
| 1077 | apply (cases "z = 0", simp add: pdivmod_rel_def) | |
| 1078 | apply (cases "y = 0", simp add: pdivmod_rel_by_0_iff pdivmod_rel_0_iff) | |
| 1079 | apply (cases "r = 0") | |
| 1080 | apply (cases "r' = 0") | |
| 1081 | apply (simp add: pdivmod_rel_def) | |
| 36350 | 1082 | apply (simp add: pdivmod_rel_def field_simps degree_mult_eq) | 
| 29660 | 1083 | apply (cases "r' = 0") | 
| 1084 | apply (simp add: pdivmod_rel_def degree_mult_eq) | |
| 36350 | 1085 | apply (simp add: pdivmod_rel_def field_simps) | 
| 29660 | 1086 | apply (simp add: degree_mult_eq degree_add_less) | 
| 1087 | done | |
| 1088 | ||
| 1089 | lemma poly_div_mult_right: | |
| 1090 | fixes x y z :: "'a::field poly" | |
| 1091 | shows "x div (y * z) = (x div y) div z" | |
| 1092 | by (rule div_poly_eq, rule pdivmod_rel_mult, (rule pdivmod_rel)+) | |
| 1093 | ||
| 1094 | lemma poly_mod_mult_right: | |
| 1095 | fixes x y z :: "'a::field poly" | |
| 1096 | shows "x mod (y * z) = y * (x div y mod z) + x mod y" | |
| 1097 | by (rule mod_poly_eq, rule pdivmod_rel_mult, (rule pdivmod_rel)+) | |
| 1098 | ||
| 29451 | 1099 | lemma mod_pCons: | 
| 1100 | fixes a and x | |
| 1101 | assumes y: "y \<noteq> 0" | |
| 1102 | defines b: "b \<equiv> coeff (pCons a (x mod y)) (degree y) / coeff y (degree y)" | |
| 1103 | shows "(pCons a x) mod y = (pCons a (x mod y) - smult b y)" | |
| 1104 | unfolding b | |
| 1105 | apply (rule mod_poly_eq) | |
| 29537 | 1106 | apply (rule pdivmod_rel_pCons [OF pdivmod_rel y refl]) | 
| 29451 | 1107 | done | 
| 1108 | ||
| 1109 | ||
| 31663 | 1110 | subsection {* GCD of polynomials *}
 | 
| 1111 | ||
| 1112 | function | |
| 1113 | poly_gcd :: "'a::field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" where | |
| 1114 | "poly_gcd x 0 = smult (inverse (coeff x (degree x))) x" | |
| 1115 | | "y \<noteq> 0 \<Longrightarrow> poly_gcd x y = poly_gcd y (x mod y)" | |
| 1116 | by auto | |
| 1117 | ||
| 1118 | termination poly_gcd | |
| 1119 | by (relation "measure (\<lambda>(x, y). if y = 0 then 0 else Suc (degree y))") | |
| 1120 | (auto dest: degree_mod_less) | |
| 1121 | ||
| 37765 | 1122 | declare poly_gcd.simps [simp del] | 
| 31663 | 1123 | |
| 1124 | lemma poly_gcd_dvd1 [iff]: "poly_gcd x y dvd x" | |
| 1125 | and poly_gcd_dvd2 [iff]: "poly_gcd x y dvd y" | |
| 1126 | apply (induct x y rule: poly_gcd.induct) | |
| 1127 | apply (simp_all add: poly_gcd.simps) | |
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
41959diff
changeset | 1128 | apply (fastforce simp add: smult_dvd_iff dest: inverse_zero_imp_zero) | 
| 31663 | 1129 | apply (blast dest: dvd_mod_imp_dvd) | 
| 1130 | done | |
| 1131 | ||
| 1132 | lemma poly_gcd_greatest: "k dvd x \<Longrightarrow> k dvd y \<Longrightarrow> k dvd poly_gcd x y" | |
| 1133 | by (induct x y rule: poly_gcd.induct) | |
| 1134 | (simp_all add: poly_gcd.simps dvd_mod dvd_smult) | |
| 1135 | ||
| 1136 | lemma dvd_poly_gcd_iff [iff]: | |
| 1137 | "k dvd poly_gcd x y \<longleftrightarrow> k dvd x \<and> k dvd y" | |
| 1138 | by (blast intro!: poly_gcd_greatest intro: dvd_trans) | |
| 1139 | ||
| 1140 | lemma poly_gcd_monic: | |
| 1141 | "coeff (poly_gcd x y) (degree (poly_gcd x y)) = | |
| 1142 | (if x = 0 \<and> y = 0 then 0 else 1)" | |
| 1143 | by (induct x y rule: poly_gcd.induct) | |
| 1144 | (simp_all add: poly_gcd.simps nonzero_imp_inverse_nonzero) | |
| 1145 | ||
| 1146 | lemma poly_gcd_zero_iff [simp]: | |
| 1147 | "poly_gcd x y = 0 \<longleftrightarrow> x = 0 \<and> y = 0" | |
| 1148 | by (simp only: dvd_0_left_iff [symmetric] dvd_poly_gcd_iff) | |
| 1149 | ||
| 1150 | lemma poly_gcd_0_0 [simp]: "poly_gcd 0 0 = 0" | |
| 1151 | by simp | |
| 1152 | ||
| 1153 | lemma poly_dvd_antisym: | |
| 1154 | fixes p q :: "'a::idom poly" | |
| 1155 | assumes coeff: "coeff p (degree p) = coeff q (degree q)" | |
| 1156 | assumes dvd1: "p dvd q" and dvd2: "q dvd p" shows "p = q" | |
| 1157 | proof (cases "p = 0") | |
| 1158 | case True with coeff show "p = q" by simp | |
| 1159 | next | |
| 1160 | case False with coeff have "q \<noteq> 0" by auto | |
| 1161 | have degree: "degree p = degree q" | |
| 1162 | using `p dvd q` `q dvd p` `p \<noteq> 0` `q \<noteq> 0` | |
| 1163 | by (intro order_antisym dvd_imp_degree_le) | |
| 1164 | ||
| 1165 | from `p dvd q` obtain a where a: "q = p * a" .. | |
| 1166 | with `q \<noteq> 0` have "a \<noteq> 0" by auto | |
| 1167 | with degree a `p \<noteq> 0` have "degree a = 0" | |
| 1168 | by (simp add: degree_mult_eq) | |
| 1169 | with coeff a show "p = q" | |
| 1170 | by (cases a, auto split: if_splits) | |
| 1171 | qed | |
| 1172 | ||
| 1173 | lemma poly_gcd_unique: | |
| 1174 | assumes dvd1: "d dvd x" and dvd2: "d dvd y" | |
| 1175 | and greatest: "\<And>k. k dvd x \<Longrightarrow> k dvd y \<Longrightarrow> k dvd d" | |
| 1176 | and monic: "coeff d (degree d) = (if x = 0 \<and> y = 0 then 0 else 1)" | |
| 1177 | shows "poly_gcd x y = d" | |
| 1178 | proof - | |
| 1179 | have "coeff (poly_gcd x y) (degree (poly_gcd x y)) = coeff d (degree d)" | |
| 1180 | by (simp_all add: poly_gcd_monic monic) | |
| 1181 | moreover have "poly_gcd x y dvd d" | |
| 1182 | using poly_gcd_dvd1 poly_gcd_dvd2 by (rule greatest) | |
| 1183 | moreover have "d dvd poly_gcd x y" | |
| 1184 | using dvd1 dvd2 by (rule poly_gcd_greatest) | |
| 1185 | ultimately show ?thesis | |
| 1186 | by (rule poly_dvd_antisym) | |
| 1187 | qed | |
| 1188 | ||
| 37770 
cddb3106adb8
avoid explicit mandatory prefix markers when prefixes are mandatory implicitly
 haftmann parents: 
37765diff
changeset | 1189 | interpretation poly_gcd: abel_semigroup poly_gcd | 
| 34973 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34915diff
changeset | 1190 | proof | 
| 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34915diff
changeset | 1191 | fix x y z :: "'a poly" | 
| 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34915diff
changeset | 1192 | show "poly_gcd (poly_gcd x y) z = poly_gcd x (poly_gcd y z)" | 
| 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34915diff
changeset | 1193 | by (rule poly_gcd_unique) (auto intro: dvd_trans simp add: poly_gcd_monic) | 
| 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34915diff
changeset | 1194 | show "poly_gcd x y = poly_gcd y x" | 
| 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34915diff
changeset | 1195 | by (rule poly_gcd_unique) (simp_all add: poly_gcd_monic) | 
| 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34915diff
changeset | 1196 | qed | 
| 31663 | 1197 | |
| 34973 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34915diff
changeset | 1198 | lemmas poly_gcd_assoc = poly_gcd.assoc | 
| 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34915diff
changeset | 1199 | lemmas poly_gcd_commute = poly_gcd.commute | 
| 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34915diff
changeset | 1200 | lemmas poly_gcd_left_commute = poly_gcd.left_commute | 
| 31663 | 1201 | |
| 1202 | lemmas poly_gcd_ac = poly_gcd_assoc poly_gcd_commute poly_gcd_left_commute | |
| 1203 | ||
| 1204 | lemma poly_gcd_1_left [simp]: "poly_gcd 1 y = 1" | |
| 1205 | by (rule poly_gcd_unique) simp_all | |
| 1206 | ||
| 1207 | lemma poly_gcd_1_right [simp]: "poly_gcd x 1 = 1" | |
| 1208 | by (rule poly_gcd_unique) simp_all | |
| 1209 | ||
| 1210 | lemma poly_gcd_minus_left [simp]: "poly_gcd (- x) y = poly_gcd x y" | |
| 1211 | by (rule poly_gcd_unique) (simp_all add: poly_gcd_monic) | |
| 1212 | ||
| 1213 | lemma poly_gcd_minus_right [simp]: "poly_gcd x (- y) = poly_gcd x y" | |
| 1214 | by (rule poly_gcd_unique) (simp_all add: poly_gcd_monic) | |
| 1215 | ||
| 1216 | ||
| 29451 | 1217 | subsection {* Evaluation of polynomials *}
 | 
| 1218 | ||
| 1219 | definition | |
| 29454 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 1220 | poly :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a" where | 
| 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 1221 | "poly = poly_rec (\<lambda>x. 0) (\<lambda>a p f x. a + x * f x)" | 
| 29451 | 1222 | |
| 1223 | lemma poly_0 [simp]: "poly 0 x = 0" | |
| 29454 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 1224 | unfolding poly_def by (simp add: poly_rec_0) | 
| 29451 | 1225 | |
| 1226 | lemma poly_pCons [simp]: "poly (pCons a p) x = a + x * poly p x" | |
| 29454 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 1227 | unfolding poly_def by (simp add: poly_rec_pCons) | 
| 29451 | 1228 | |
| 1229 | lemma poly_1 [simp]: "poly 1 x = 1" | |
| 1230 | unfolding one_poly_def by simp | |
| 1231 | ||
| 29454 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 1232 | lemma poly_monom: | 
| 31021 | 1233 |   fixes a x :: "'a::{comm_semiring_1}"
 | 
| 29454 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 1234 | shows "poly (monom a n) x = a * x ^ n" | 
| 29451 | 1235 | by (induct n, simp add: monom_0, simp add: monom_Suc power_Suc mult_ac) | 
| 1236 | ||
| 1237 | lemma poly_add [simp]: "poly (p + q) x = poly p x + poly q x" | |
| 1238 | apply (induct p arbitrary: q, simp) | |
| 29667 | 1239 | apply (case_tac q, simp, simp add: algebra_simps) | 
| 29451 | 1240 | done | 
| 1241 | ||
| 1242 | lemma poly_minus [simp]: | |
| 29454 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 1243 | fixes x :: "'a::comm_ring" | 
| 29451 | 1244 | shows "poly (- p) x = - poly p x" | 
| 1245 | by (induct p, simp_all) | |
| 1246 | ||
| 1247 | lemma poly_diff [simp]: | |
| 29454 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 1248 | fixes x :: "'a::comm_ring" | 
| 29451 | 1249 | shows "poly (p - q) x = poly p x - poly q x" | 
| 1250 | by (simp add: diff_minus) | |
| 1251 | ||
| 1252 | lemma poly_setsum: "poly (\<Sum>k\<in>A. p k) x = (\<Sum>k\<in>A. poly (p k) x)" | |
| 1253 | by (cases "finite A", induct set: finite, simp_all) | |
| 1254 | ||
| 1255 | lemma poly_smult [simp]: "poly (smult a p) x = a * poly p x" | |
| 29667 | 1256 | by (induct p, simp, simp add: algebra_simps) | 
| 29451 | 1257 | |
| 1258 | lemma poly_mult [simp]: "poly (p * q) x = poly p x * poly q x" | |
| 29667 | 1259 | by (induct p, simp_all, simp add: algebra_simps) | 
| 29451 | 1260 | |
| 29462 | 1261 | lemma poly_power [simp]: | 
| 31021 | 1262 |   fixes p :: "'a::{comm_semiring_1} poly"
 | 
| 29462 | 1263 | shows "poly (p ^ n) x = poly p x ^ n" | 
| 1264 | by (induct n, simp, simp add: power_Suc) | |
| 1265 | ||
| 29456 | 1266 | |
| 1267 | subsection {* Synthetic division *}
 | |
| 1268 | ||
| 29980 | 1269 | text {*
 | 
| 1270 | Synthetic division is simply division by the | |
| 1271 |   linear polynomial @{term "x - c"}.
 | |
| 1272 | *} | |
| 1273 | ||
| 29456 | 1274 | definition | 
| 1275 | synthetic_divmod :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a poly \<times> 'a" | |
| 37765 | 1276 | where | 
| 29456 | 1277 | "synthetic_divmod p c = | 
| 1278 | poly_rec (0, 0) (\<lambda>a p (q, r). (pCons r q, a + c * r)) p" | |
| 1279 | ||
| 1280 | definition | |
| 1281 | synthetic_div :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a poly" | |
| 1282 | where | |
| 1283 | "synthetic_div p c = fst (synthetic_divmod p c)" | |
| 1284 | ||
| 1285 | lemma synthetic_divmod_0 [simp]: | |
| 1286 | "synthetic_divmod 0 c = (0, 0)" | |
| 1287 | unfolding synthetic_divmod_def | |
| 1288 | by (simp add: poly_rec_0) | |
| 1289 | ||
| 1290 | lemma synthetic_divmod_pCons [simp]: | |
| 1291 | "synthetic_divmod (pCons a p) c = | |
| 1292 | (\<lambda>(q, r). (pCons r q, a + c * r)) (synthetic_divmod p c)" | |
| 1293 | unfolding synthetic_divmod_def | |
| 1294 | by (simp add: poly_rec_pCons) | |
| 1295 | ||
| 1296 | lemma snd_synthetic_divmod: "snd (synthetic_divmod p c) = poly p c" | |
| 1297 | by (induct p, simp, simp add: split_def) | |
| 1298 | ||
| 1299 | lemma synthetic_div_0 [simp]: "synthetic_div 0 c = 0" | |
| 1300 | unfolding synthetic_div_def by simp | |
| 1301 | ||
| 1302 | lemma synthetic_div_pCons [simp]: | |
| 1303 | "synthetic_div (pCons a p) c = pCons (poly p c) (synthetic_div p c)" | |
| 1304 | unfolding synthetic_div_def | |
| 1305 | by (simp add: split_def snd_synthetic_divmod) | |
| 1306 | ||
| 29460 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 1307 | lemma synthetic_div_eq_0_iff: | 
| 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 1308 | "synthetic_div p c = 0 \<longleftrightarrow> degree p = 0" | 
| 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 1309 | by (induct p, simp, case_tac p, simp) | 
| 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 1310 | |
| 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 1311 | lemma degree_synthetic_div: | 
| 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 1312 | "degree (synthetic_div p c) = degree p - 1" | 
| 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 1313 | by (induct p, simp, simp add: synthetic_div_eq_0_iff) | 
| 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 1314 | |
| 29457 | 1315 | lemma synthetic_div_correct: | 
| 29456 | 1316 | "p + smult c (synthetic_div p c) = pCons (poly p c) (synthetic_div p c)" | 
| 1317 | by (induct p) simp_all | |
| 1318 | ||
| 29457 | 1319 | lemma synthetic_div_unique_lemma: "smult c p = pCons a p \<Longrightarrow> p = 0" | 
| 1320 | by (induct p arbitrary: a) simp_all | |
| 1321 | ||
| 1322 | lemma synthetic_div_unique: | |
| 1323 | "p + smult c q = pCons r q \<Longrightarrow> r = poly p c \<and> q = synthetic_div p c" | |
| 1324 | apply (induct p arbitrary: q r) | |
| 1325 | apply (simp, frule synthetic_div_unique_lemma, simp) | |
| 1326 | apply (case_tac q, force) | |
| 1327 | done | |
| 1328 | ||
| 1329 | lemma synthetic_div_correct': | |
| 1330 | fixes c :: "'a::comm_ring_1" | |
| 1331 | shows "[:-c, 1:] * synthetic_div p c + [:poly p c:] = p" | |
| 1332 | using synthetic_div_correct [of p c] | |
| 29667 | 1333 | by (simp add: algebra_simps) | 
| 29457 | 1334 | |
| 29460 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 1335 | lemma poly_eq_0_iff_dvd: | 
| 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 1336 | fixes c :: "'a::idom" | 
| 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 1337 | shows "poly p c = 0 \<longleftrightarrow> [:-c, 1:] dvd p" | 
| 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 1338 | proof | 
| 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 1339 | assume "poly p c = 0" | 
| 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 1340 | with synthetic_div_correct' [of c p] | 
| 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 1341 | have "p = [:-c, 1:] * synthetic_div p c" by simp | 
| 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 1342 | then show "[:-c, 1:] dvd p" .. | 
| 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 1343 | next | 
| 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 1344 | assume "[:-c, 1:] dvd p" | 
| 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 1345 | then obtain k where "p = [:-c, 1:] * k" by (rule dvdE) | 
| 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 1346 | then show "poly p c = 0" by simp | 
| 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 1347 | qed | 
| 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 1348 | |
| 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 1349 | lemma dvd_iff_poly_eq_0: | 
| 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 1350 | fixes c :: "'a::idom" | 
| 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 1351 | shows "[:c, 1:] dvd p \<longleftrightarrow> poly p (-c) = 0" | 
| 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 1352 | by (simp add: poly_eq_0_iff_dvd) | 
| 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 1353 | |
| 29462 | 1354 | lemma poly_roots_finite: | 
| 1355 | fixes p :: "'a::idom poly" | |
| 1356 |   shows "p \<noteq> 0 \<Longrightarrow> finite {x. poly p x = 0}"
 | |
| 1357 | proof (induct n \<equiv> "degree p" arbitrary: p) | |
| 1358 | case (0 p) | |
| 1359 | then obtain a where "a \<noteq> 0" and "p = [:a:]" | |
| 1360 | by (cases p, simp split: if_splits) | |
| 1361 |   then show "finite {x. poly p x = 0}" by simp
 | |
| 1362 | next | |
| 1363 | case (Suc n p) | |
| 1364 |   show "finite {x. poly p x = 0}"
 | |
| 1365 | proof (cases "\<exists>x. poly p x = 0") | |
| 1366 | case False | |
| 1367 |     then show "finite {x. poly p x = 0}" by simp
 | |
| 1368 | next | |
| 1369 | case True | |
| 1370 | then obtain a where "poly p a = 0" .. | |
| 1371 | then have "[:-a, 1:] dvd p" by (simp only: poly_eq_0_iff_dvd) | |
| 1372 | then obtain k where k: "p = [:-a, 1:] * k" .. | |
| 1373 | with `p \<noteq> 0` have "k \<noteq> 0" by auto | |
| 1374 | with k have "degree p = Suc (degree k)" | |
| 1375 | by (simp add: degree_mult_eq del: mult_pCons_left) | |
| 1376 | with `Suc n = degree p` have "n = degree k" by simp | |
| 34915 | 1377 |     then have "finite {x. poly k x = 0}" using `k \<noteq> 0` by (rule Suc.hyps)
 | 
| 29462 | 1378 |     then have "finite (insert a {x. poly k x = 0})" by simp
 | 
| 1379 |     then show "finite {x. poly p x = 0}"
 | |
| 1380 | by (simp add: k uminus_add_conv_diff Collect_disj_eq | |
| 1381 | del: mult_pCons_left) | |
| 1382 | qed | |
| 1383 | qed | |
| 1384 | ||
| 29979 | 1385 | lemma poly_zero: | 
| 1386 |   fixes p :: "'a::{idom,ring_char_0} poly"
 | |
| 1387 | shows "poly p = poly 0 \<longleftrightarrow> p = 0" | |
| 1388 | apply (cases "p = 0", simp_all) | |
| 1389 | apply (drule poly_roots_finite) | |
| 1390 | apply (auto simp add: infinite_UNIV_char_0) | |
| 1391 | done | |
| 1392 | ||
| 1393 | lemma poly_eq_iff: | |
| 1394 |   fixes p q :: "'a::{idom,ring_char_0} poly"
 | |
| 1395 | shows "poly p = poly q \<longleftrightarrow> p = q" | |
| 1396 | using poly_zero [of "p - q"] | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 1397 | by (simp add: fun_eq_iff) | 
| 29979 | 1398 | |
| 29478 | 1399 | |
| 29980 | 1400 | subsection {* Composition of polynomials *}
 | 
| 1401 | ||
| 1402 | definition | |
| 1403 | pcompose :: "'a::comm_semiring_0 poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" | |
| 1404 | where | |
| 1405 | "pcompose p q = poly_rec 0 (\<lambda>a _ c. [:a:] + q * c) p" | |
| 1406 | ||
| 1407 | lemma pcompose_0 [simp]: "pcompose 0 q = 0" | |
| 1408 | unfolding pcompose_def by (simp add: poly_rec_0) | |
| 1409 | ||
| 1410 | lemma pcompose_pCons: | |
| 1411 | "pcompose (pCons a p) q = [:a:] + q * pcompose p q" | |
| 1412 | unfolding pcompose_def by (simp add: poly_rec_pCons) | |
| 1413 | ||
| 1414 | lemma poly_pcompose: "poly (pcompose p q) x = poly p (poly q x)" | |
| 1415 | by (induct p) (simp_all add: pcompose_pCons) | |
| 1416 | ||
| 1417 | lemma degree_pcompose_le: | |
| 1418 | "degree (pcompose p q) \<le> degree p * degree q" | |
| 1419 | apply (induct p, simp) | |
| 1420 | apply (simp add: pcompose_pCons, clarify) | |
| 1421 | apply (rule degree_add_le, simp) | |
| 1422 | apply (rule order_trans [OF degree_mult_le], simp) | |
| 1423 | done | |
| 1424 | ||
| 1425 | ||
| 29977 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1426 | subsection {* Order of polynomial roots *}
 | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1427 | |
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1428 | definition | 
| 29979 | 1429 | order :: "'a::idom \<Rightarrow> 'a poly \<Rightarrow> nat" | 
| 29977 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1430 | where | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1431 | "order a p = (LEAST n. \<not> [:-a, 1:] ^ Suc n dvd p)" | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1432 | |
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1433 | lemma coeff_linear_power: | 
| 29979 | 1434 | fixes a :: "'a::comm_semiring_1" | 
| 29977 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1435 | shows "coeff ([:a, 1:] ^ n) n = 1" | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1436 | apply (induct n, simp_all) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1437 | apply (subst coeff_eq_0) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1438 | apply (auto intro: le_less_trans degree_power_le) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1439 | done | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1440 | |
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1441 | lemma degree_linear_power: | 
| 29979 | 1442 | fixes a :: "'a::comm_semiring_1" | 
| 29977 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1443 | shows "degree ([:a, 1:] ^ n) = n" | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1444 | apply (rule order_antisym) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1445 | apply (rule ord_le_eq_trans [OF degree_power_le], simp) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1446 | apply (rule le_degree, simp add: coeff_linear_power) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1447 | done | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1448 | |
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1449 | lemma order_1: "[:-a, 1:] ^ order a p dvd p" | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1450 | apply (cases "p = 0", simp) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1451 | apply (cases "order a p", simp) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1452 | apply (subgoal_tac "nat < (LEAST n. \<not> [:-a, 1:] ^ Suc n dvd p)") | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1453 | apply (drule not_less_Least, simp) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1454 | apply (fold order_def, simp) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1455 | done | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1456 | |
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1457 | lemma order_2: "p \<noteq> 0 \<Longrightarrow> \<not> [:-a, 1:] ^ Suc (order a p) dvd p" | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1458 | unfolding order_def | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1459 | apply (rule LeastI_ex) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1460 | apply (rule_tac x="degree p" in exI) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1461 | apply (rule notI) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1462 | apply (drule (1) dvd_imp_degree_le) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1463 | apply (simp only: degree_linear_power) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1464 | done | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1465 | |
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1466 | lemma order: | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1467 | "p \<noteq> 0 \<Longrightarrow> [:-a, 1:] ^ order a p dvd p \<and> \<not> [:-a, 1:] ^ Suc (order a p) dvd p" | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1468 | by (rule conjI [OF order_1 order_2]) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1469 | |
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1470 | lemma order_degree: | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1471 | assumes p: "p \<noteq> 0" | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1472 | shows "order a p \<le> degree p" | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1473 | proof - | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1474 | have "order a p = degree ([:-a, 1:] ^ order a p)" | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1475 | by (simp only: degree_linear_power) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1476 | also have "\<dots> \<le> degree p" | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1477 | using order_1 p by (rule dvd_imp_degree_le) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1478 | finally show ?thesis . | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1479 | qed | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1480 | |
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1481 | lemma order_root: "poly p a = 0 \<longleftrightarrow> p = 0 \<or> order a p \<noteq> 0" | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1482 | apply (cases "p = 0", simp_all) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1483 | apply (rule iffI) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1484 | apply (rule ccontr, simp) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1485 | apply (frule order_2 [where a=a], simp) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1486 | apply (simp add: poly_eq_0_iff_dvd) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1487 | apply (simp add: poly_eq_0_iff_dvd) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1488 | apply (simp only: order_def) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1489 | apply (drule not_less_Least, simp) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1490 | done | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1491 | |
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1492 | |
| 29478 | 1493 | subsection {* Configuration of the code generator *}
 | 
| 1494 | ||
| 1495 | code_datatype "0::'a::zero poly" pCons | |
| 1496 | ||
| 45928 
874845660119
adding quickcheck generators in some HOL-Library theories
 bulwahn parents: 
45694diff
changeset | 1497 | quickcheck_generator poly constructors: "0::'a::zero poly", pCons | 
| 
874845660119
adding quickcheck generators in some HOL-Library theories
 bulwahn parents: 
45694diff
changeset | 1498 | |
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37770diff
changeset | 1499 | instantiation poly :: ("{zero, equal}") equal
 | 
| 29478 | 1500 | begin | 
| 1501 | ||
| 37765 | 1502 | definition | 
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37770diff
changeset | 1503 | "HOL.equal (p::'a poly) q \<longleftrightarrow> p = q" | 
| 29478 | 1504 | |
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37770diff
changeset | 1505 | instance proof | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37770diff
changeset | 1506 | qed (rule equal_poly_def) | 
| 29478 | 1507 | |
| 29451 | 1508 | end | 
| 29478 | 1509 | |
| 1510 | lemma eq_poly_code [code]: | |
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37770diff
changeset | 1511 | "HOL.equal (0::_ poly) (0::_ poly) \<longleftrightarrow> True" | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37770diff
changeset | 1512 | "HOL.equal (0::_ poly) (pCons b q) \<longleftrightarrow> HOL.equal 0 b \<and> HOL.equal 0 q" | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37770diff
changeset | 1513 | "HOL.equal (pCons a p) (0::_ poly) \<longleftrightarrow> HOL.equal a 0 \<and> HOL.equal p 0" | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37770diff
changeset | 1514 | "HOL.equal (pCons a p) (pCons b q) \<longleftrightarrow> HOL.equal a b \<and> HOL.equal p q" | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37770diff
changeset | 1515 | by (simp_all add: equal) | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37770diff
changeset | 1516 | |
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37770diff
changeset | 1517 | lemma [code nbe]: | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37770diff
changeset | 1518 | "HOL.equal (p :: _ poly) p \<longleftrightarrow> True" | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37770diff
changeset | 1519 | by (fact equal_refl) | 
| 29478 | 1520 | |
| 1521 | lemmas coeff_code [code] = | |
| 1522 | coeff_0 coeff_pCons_0 coeff_pCons_Suc | |
| 1523 | ||
| 1524 | lemmas degree_code [code] = | |
| 1525 | degree_0 degree_pCons_eq_if | |
| 1526 | ||
| 1527 | lemmas monom_poly_code [code] = | |
| 1528 | monom_0 monom_Suc | |
| 1529 | ||
| 1530 | lemma add_poly_code [code]: | |
| 1531 | "0 + q = (q :: _ poly)" | |
| 1532 | "p + 0 = (p :: _ poly)" | |
| 1533 | "pCons a p + pCons b q = pCons (a + b) (p + q)" | |
| 1534 | by simp_all | |
| 1535 | ||
| 1536 | lemma minus_poly_code [code]: | |
| 1537 | "- 0 = (0 :: _ poly)" | |
| 1538 | "- pCons a p = pCons (- a) (- p)" | |
| 1539 | by simp_all | |
| 1540 | ||
| 1541 | lemma diff_poly_code [code]: | |
| 1542 | "0 - q = (- q :: _ poly)" | |
| 1543 | "p - 0 = (p :: _ poly)" | |
| 1544 | "pCons a p - pCons b q = pCons (a - b) (p - q)" | |
| 1545 | by simp_all | |
| 1546 | ||
| 1547 | lemmas smult_poly_code [code] = | |
| 1548 | smult_0_right smult_pCons | |
| 1549 | ||
| 1550 | lemma mult_poly_code [code]: | |
| 1551 | "0 * q = (0 :: _ poly)" | |
| 1552 | "pCons a p * q = smult a q + pCons 0 (p * q)" | |
| 1553 | by simp_all | |
| 1554 | ||
| 1555 | lemmas poly_code [code] = | |
| 1556 | poly_0 poly_pCons | |
| 1557 | ||
| 1558 | lemmas synthetic_divmod_code [code] = | |
| 1559 | synthetic_divmod_0 synthetic_divmod_pCons | |
| 1560 | ||
| 1561 | text {* code generator setup for div and mod *}
 | |
| 1562 | ||
| 1563 | definition | |
| 29537 | 1564 | pdivmod :: "'a::field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<times> 'a poly" | 
| 29478 | 1565 | where | 
| 37765 | 1566 | "pdivmod x y = (x div y, x mod y)" | 
| 29478 | 1567 | |
| 29537 | 1568 | lemma div_poly_code [code]: "x div y = fst (pdivmod x y)" | 
| 1569 | unfolding pdivmod_def by simp | |
| 29478 | 1570 | |
| 29537 | 1571 | lemma mod_poly_code [code]: "x mod y = snd (pdivmod x y)" | 
| 1572 | unfolding pdivmod_def by simp | |
| 29478 | 1573 | |
| 29537 | 1574 | lemma pdivmod_0 [code]: "pdivmod 0 y = (0, 0)" | 
| 1575 | unfolding pdivmod_def by simp | |
| 29478 | 1576 | |
| 29537 | 1577 | lemma pdivmod_pCons [code]: | 
| 1578 | "pdivmod (pCons a x) y = | |
| 29478 | 1579 | (if y = 0 then (0, pCons a x) else | 
| 29537 | 1580 | (let (q, r) = pdivmod x y; | 
| 29478 | 1581 | b = coeff (pCons a r) (degree y) / coeff y (degree y) | 
| 1582 | in (pCons b q, pCons a r - smult b y)))" | |
| 29537 | 1583 | apply (simp add: pdivmod_def Let_def, safe) | 
| 29478 | 1584 | apply (rule div_poly_eq) | 
| 29537 | 1585 | apply (erule pdivmod_rel_pCons [OF pdivmod_rel _ refl]) | 
| 29478 | 1586 | apply (rule mod_poly_eq) | 
| 29537 | 1587 | apply (erule pdivmod_rel_pCons [OF pdivmod_rel _ refl]) | 
| 29478 | 1588 | done | 
| 1589 | ||
| 31663 | 1590 | lemma poly_gcd_code [code]: | 
| 1591 | "poly_gcd x y = | |
| 1592 | (if y = 0 then smult (inverse (coeff x (degree x))) x | |
| 1593 | else poly_gcd y (x mod y))" | |
| 1594 | by (simp add: poly_gcd.simps) | |
| 1595 | ||
| 29478 | 1596 | end |