| author | wenzelm | 
| Mon, 02 Nov 2015 09:43:20 +0100 | |
| changeset 61536 | 346aa2c5447f | 
| parent 61260 | e6f03fae14d5 | 
| child 61585 | a9599d3d7610 | 
| permissions | -rw-r--r-- | 
| 41959 | 1 | (* Title: HOL/Library/Polynomial.thy | 
| 29451 | 2 | Author: Brian Huffman | 
| 41959 | 3 | Author: Clemens Ballarin | 
| 52380 | 4 | Author: Florian Haftmann | 
| 29451 | 5 | *) | 
| 6 | ||
| 60500 | 7 | section \<open>Polynomials as type over a ring structure\<close> | 
| 29451 | 8 | |
| 9 | theory Polynomial | |
| 60685 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 10 | imports Main "~~/src/HOL/GCD" "~~/src/HOL/Library/More_List" "~~/src/HOL/Library/Infinite_Set" | 
| 29451 | 11 | begin | 
| 12 | ||
| 60500 | 13 | subsection \<open>Auxiliary: operations for lists (later) representing coefficients\<close> | 
| 52380 | 14 | |
| 15 | definition cCons :: "'a::zero \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixr "##" 65) | |
| 16 | where | |
| 17 | "x ## xs = (if xs = [] \<and> x = 0 then [] else x # xs)" | |
| 18 | ||
| 19 | lemma cCons_0_Nil_eq [simp]: | |
| 20 | "0 ## [] = []" | |
| 21 | by (simp add: cCons_def) | |
| 22 | ||
| 23 | lemma cCons_Cons_eq [simp]: | |
| 24 | "x ## y # ys = x # y # ys" | |
| 25 | by (simp add: cCons_def) | |
| 26 | ||
| 27 | lemma cCons_append_Cons_eq [simp]: | |
| 28 | "x ## xs @ y # ys = x # xs @ y # ys" | |
| 29 | by (simp add: cCons_def) | |
| 30 | ||
| 31 | lemma cCons_not_0_eq [simp]: | |
| 32 | "x \<noteq> 0 \<Longrightarrow> x ## xs = x # xs" | |
| 33 | by (simp add: cCons_def) | |
| 34 | ||
| 35 | lemma strip_while_not_0_Cons_eq [simp]: | |
| 36 | "strip_while (\<lambda>x. x = 0) (x # xs) = x ## strip_while (\<lambda>x. x = 0) xs" | |
| 37 | proof (cases "x = 0") | |
| 38 | case False then show ?thesis by simp | |
| 39 | next | |
| 40 | case True show ?thesis | |
| 41 | proof (induct xs rule: rev_induct) | |
| 42 | case Nil with True show ?case by simp | |
| 43 | next | |
| 44 | case (snoc y ys) then show ?case | |
| 45 | by (cases "y = 0") (simp_all add: append_Cons [symmetric] del: append_Cons) | |
| 46 | qed | |
| 47 | qed | |
| 48 | ||
| 49 | lemma tl_cCons [simp]: | |
| 50 | "tl (x ## xs) = xs" | |
| 51 | by (simp add: cCons_def) | |
| 52 | ||
| 60500 | 53 | subsection \<open>Definition of type @{text poly}\<close>
 | 
| 29451 | 54 | |
| 61260 | 55 | typedef (overloaded) 'a poly = "{f :: nat \<Rightarrow> 'a::zero. \<forall>\<^sub>\<infinity> n. f n = 0}"
 | 
| 59983 
cd2efd7d06bd
replace almost_everywhere_zero by Infinite_Set.MOST
 hoelzl parents: 
59815diff
changeset | 56 | morphisms coeff Abs_poly by (auto intro!: ALL_MOST) | 
| 29451 | 57 | |
| 59487 
adaa430fc0f7
default abstypes and default abstract equations make technical (no_code) annotation superfluous
 haftmann parents: 
58881diff
changeset | 58 | setup_lifting type_definition_poly | 
| 52380 | 59 | |
| 60 | lemma poly_eq_iff: "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 | 61 | by (simp add: coeff_inject [symmetric] fun_eq_iff) | 
| 29451 | 62 | |
| 52380 | 63 | lemma poly_eqI: "(\<And>n. coeff p n = coeff q n) \<Longrightarrow> p = q" | 
| 64 | by (simp add: poly_eq_iff) | |
| 65 | ||
| 59983 
cd2efd7d06bd
replace almost_everywhere_zero by Infinite_Set.MOST
 hoelzl parents: 
59815diff
changeset | 66 | lemma MOST_coeff_eq_0: "\<forall>\<^sub>\<infinity> n. coeff p n = 0" | 
| 52380 | 67 | using coeff [of p] by simp | 
| 29451 | 68 | |
| 69 | ||
| 60500 | 70 | subsection \<open>Degree of a polynomial\<close> | 
| 29451 | 71 | |
| 52380 | 72 | definition degree :: "'a::zero poly \<Rightarrow> nat" | 
| 73 | where | |
| 29451 | 74 | "degree p = (LEAST n. \<forall>i>n. coeff p i = 0)" | 
| 75 | ||
| 52380 | 76 | lemma coeff_eq_0: | 
| 77 | assumes "degree p < n" | |
| 78 | shows "coeff p n = 0" | |
| 29451 | 79 | proof - | 
| 59983 
cd2efd7d06bd
replace almost_everywhere_zero by Infinite_Set.MOST
 hoelzl parents: 
59815diff
changeset | 80 | have "\<exists>n. \<forall>i>n. coeff p i = 0" | 
| 
cd2efd7d06bd
replace almost_everywhere_zero by Infinite_Set.MOST
 hoelzl parents: 
59815diff
changeset | 81 | using MOST_coeff_eq_0 by (simp add: MOST_nat) | 
| 52380 | 82 | then have "\<forall>i>degree p. coeff p i = 0" | 
| 29451 | 83 | unfolding degree_def by (rule LeastI_ex) | 
| 52380 | 84 | with assms show ?thesis by simp | 
| 29451 | 85 | qed | 
| 86 | ||
| 87 | lemma le_degree: "coeff p n \<noteq> 0 \<Longrightarrow> n \<le> degree p" | |
| 88 | by (erule contrapos_np, rule coeff_eq_0, simp) | |
| 89 | ||
| 90 | lemma degree_le: "\<forall>i>n. coeff p i = 0 \<Longrightarrow> degree p \<le> n" | |
| 91 | unfolding degree_def by (erule Least_le) | |
| 92 | ||
| 93 | lemma less_degree_imp: "n < degree p \<Longrightarrow> \<exists>i>n. coeff p i \<noteq> 0" | |
| 94 | unfolding degree_def by (drule not_less_Least, simp) | |
| 95 | ||
| 96 | ||
| 60500 | 97 | subsection \<open>The zero polynomial\<close> | 
| 29451 | 98 | |
| 99 | instantiation poly :: (zero) zero | |
| 100 | begin | |
| 101 | ||
| 52380 | 102 | lift_definition zero_poly :: "'a poly" | 
| 59983 
cd2efd7d06bd
replace almost_everywhere_zero by Infinite_Set.MOST
 hoelzl parents: 
59815diff
changeset | 103 | is "\<lambda>_. 0" by (rule MOST_I) simp | 
| 29451 | 104 | |
| 105 | instance .. | |
| 52380 | 106 | |
| 29451 | 107 | end | 
| 108 | ||
| 52380 | 109 | lemma coeff_0 [simp]: | 
| 110 | "coeff 0 n = 0" | |
| 111 | by transfer rule | |
| 29451 | 112 | |
| 52380 | 113 | lemma degree_0 [simp]: | 
| 114 | "degree 0 = 0" | |
| 29451 | 115 | by (rule order_antisym [OF degree_le le0]) simp | 
| 116 | ||
| 117 | lemma leading_coeff_neq_0: | |
| 52380 | 118 | assumes "p \<noteq> 0" | 
| 119 | shows "coeff p (degree p) \<noteq> 0" | |
| 29451 | 120 | proof (cases "degree p") | 
| 121 | case 0 | |
| 60500 | 122 | from \<open>p \<noteq> 0\<close> have "\<exists>n. coeff p n \<noteq> 0" | 
| 52380 | 123 | by (simp add: poly_eq_iff) | 
| 29451 | 124 | then obtain n where "coeff p n \<noteq> 0" .. | 
| 125 | hence "n \<le> degree p" by (rule le_degree) | |
| 60500 | 126 | with \<open>coeff p n \<noteq> 0\<close> and \<open>degree p = 0\<close> | 
| 29451 | 127 | show "coeff p (degree p) \<noteq> 0" by simp | 
| 128 | next | |
| 129 | case (Suc n) | |
| 60500 | 130 | from \<open>degree p = Suc n\<close> have "n < degree p" by simp | 
| 29451 | 131 | hence "\<exists>i>n. coeff p i \<noteq> 0" by (rule less_degree_imp) | 
| 132 | then obtain i where "n < i" and "coeff p i \<noteq> 0" by fast | |
| 60500 | 133 | from \<open>degree p = Suc n\<close> and \<open>n < i\<close> have "degree p \<le> i" by simp | 
| 134 | also from \<open>coeff p i \<noteq> 0\<close> have "i \<le> degree p" by (rule le_degree) | |
| 29451 | 135 | finally have "degree p = i" . | 
| 60500 | 136 | with \<open>coeff p i \<noteq> 0\<close> show "coeff p (degree p) \<noteq> 0" by simp | 
| 29451 | 137 | qed | 
| 138 | ||
| 52380 | 139 | lemma leading_coeff_0_iff [simp]: | 
| 140 | "coeff p (degree p) = 0 \<longleftrightarrow> p = 0" | |
| 29451 | 141 | by (cases "p = 0", simp, simp add: leading_coeff_neq_0) | 
| 142 | ||
| 143 | ||
| 60500 | 144 | subsection \<open>List-style constructor for polynomials\<close> | 
| 29451 | 145 | |
| 52380 | 146 | lift_definition pCons :: "'a::zero \<Rightarrow> 'a poly \<Rightarrow> 'a poly" | 
| 55415 | 147 | is "\<lambda>a p. case_nat a (coeff p)" | 
| 59983 
cd2efd7d06bd
replace almost_everywhere_zero by Infinite_Set.MOST
 hoelzl parents: 
59815diff
changeset | 148 | by (rule MOST_SucD) (simp add: MOST_coeff_eq_0) | 
| 29451 | 149 | |
| 52380 | 150 | lemmas coeff_pCons = pCons.rep_eq | 
| 29455 | 151 | |
| 52380 | 152 | lemma coeff_pCons_0 [simp]: | 
| 153 | "coeff (pCons a p) 0 = a" | |
| 154 | by transfer simp | |
| 29455 | 155 | |
| 52380 | 156 | lemma coeff_pCons_Suc [simp]: | 
| 157 | "coeff (pCons a p) (Suc n) = coeff p n" | |
| 29451 | 158 | by (simp add: coeff_pCons) | 
| 159 | ||
| 52380 | 160 | lemma degree_pCons_le: | 
| 161 | "degree (pCons a p) \<le> Suc (degree p)" | |
| 162 | by (rule degree_le) (simp add: coeff_eq_0 coeff_pCons split: nat.split) | |
| 29451 | 163 | |
| 164 | lemma degree_pCons_eq: | |
| 165 | "p \<noteq> 0 \<Longrightarrow> degree (pCons a p) = Suc (degree p)" | |
| 52380 | 166 | apply (rule order_antisym [OF degree_pCons_le]) | 
| 167 | apply (rule le_degree, simp) | |
| 168 | done | |
| 29451 | 169 | |
| 52380 | 170 | lemma degree_pCons_0: | 
| 171 | "degree (pCons a 0) = 0" | |
| 172 | apply (rule order_antisym [OF _ le0]) | |
| 173 | apply (rule degree_le, simp add: coeff_pCons split: nat.split) | |
| 174 | done | |
| 29451 | 175 | |
| 29460 
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
 huffman parents: 
29457diff
changeset | 176 | lemma degree_pCons_eq_if [simp]: | 
| 29451 | 177 | "degree (pCons a p) = (if p = 0 then 0 else Suc (degree p))" | 
| 52380 | 178 | apply (cases "p = 0", simp_all) | 
| 179 | apply (rule order_antisym [OF _ le0]) | |
| 180 | apply (rule degree_le, simp add: coeff_pCons split: nat.split) | |
| 181 | apply (rule order_antisym [OF degree_pCons_le]) | |
| 182 | apply (rule le_degree, simp) | |
| 183 | done | |
| 29451 | 184 | |
| 52380 | 185 | lemma pCons_0_0 [simp]: | 
| 186 | "pCons 0 0 = 0" | |
| 187 | by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) | |
| 29451 | 188 | |
| 189 | lemma pCons_eq_iff [simp]: | |
| 190 | "pCons a p = pCons b q \<longleftrightarrow> a = b \<and> p = q" | |
| 52380 | 191 | proof safe | 
| 29451 | 192 | assume "pCons a p = pCons b q" | 
| 193 | then have "coeff (pCons a p) 0 = coeff (pCons b q) 0" by simp | |
| 194 | then show "a = b" by simp | |
| 195 | next | |
| 196 | assume "pCons a p = pCons b q" | |
| 197 | then have "\<forall>n. coeff (pCons a p) (Suc n) = | |
| 198 | coeff (pCons b q) (Suc n)" by simp | |
| 52380 | 199 | then show "p = q" by (simp add: poly_eq_iff) | 
| 29451 | 200 | qed | 
| 201 | ||
| 52380 | 202 | lemma pCons_eq_0_iff [simp]: | 
| 203 | "pCons a p = 0 \<longleftrightarrow> a = 0 \<and> p = 0" | |
| 29451 | 204 | using pCons_eq_iff [of a p 0 0] by simp | 
| 205 | ||
| 206 | lemma pCons_cases [cases type: poly]: | |
| 207 | obtains (pCons) a q where "p = pCons a q" | |
| 208 | proof | |
| 209 | show "p = pCons (coeff p 0) (Abs_poly (\<lambda>n. coeff p (Suc n)))" | |
| 52380 | 210 | by transfer | 
| 59983 
cd2efd7d06bd
replace almost_everywhere_zero by Infinite_Set.MOST
 hoelzl parents: 
59815diff
changeset | 211 | (simp_all add: MOST_inj[where f=Suc and P="\<lambda>n. p n = 0" for p] fun_eq_iff Abs_poly_inverse | 
| 
cd2efd7d06bd
replace almost_everywhere_zero by Infinite_Set.MOST
 hoelzl parents: 
59815diff
changeset | 212 | split: nat.split) | 
| 29451 | 213 | qed | 
| 214 | ||
| 215 | lemma pCons_induct [case_names 0 pCons, induct type: poly]: | |
| 216 | assumes zero: "P 0" | |
| 54856 | 217 | assumes pCons: "\<And>a p. a \<noteq> 0 \<or> p \<noteq> 0 \<Longrightarrow> P p \<Longrightarrow> P (pCons a p)" | 
| 29451 | 218 | shows "P p" | 
| 219 | proof (induct p rule: measure_induct_rule [where f=degree]) | |
| 220 | case (less p) | |
| 221 | obtain a q where "p = pCons a q" by (rule pCons_cases) | |
| 222 | have "P q" | |
| 223 | proof (cases "q = 0") | |
| 224 | case True | |
| 225 | then show "P q" by (simp add: zero) | |
| 226 | next | |
| 227 | case False | |
| 228 | then have "degree (pCons a q) = Suc (degree q)" | |
| 229 | by (rule degree_pCons_eq) | |
| 230 | then have "degree q < degree p" | |
| 60500 | 231 | using \<open>p = pCons a q\<close> by simp | 
| 29451 | 232 | then show "P q" | 
| 233 | by (rule less.hyps) | |
| 234 | qed | |
| 54856 | 235 | have "P (pCons a q)" | 
| 236 | proof (cases "a \<noteq> 0 \<or> q \<noteq> 0") | |
| 237 | case True | |
| 60500 | 238 | with \<open>P q\<close> show ?thesis by (auto intro: pCons) | 
| 54856 | 239 | next | 
| 240 | case False | |
| 241 | with zero show ?thesis by simp | |
| 242 | qed | |
| 29451 | 243 | then show ?case | 
| 60500 | 244 | using \<open>p = pCons a q\<close> by simp | 
| 29451 | 245 | qed | 
| 246 | ||
| 60570 | 247 | lemma degree_eq_zeroE: | 
| 248 | fixes p :: "'a::zero poly" | |
| 249 | assumes "degree p = 0" | |
| 250 | obtains a where "p = pCons a 0" | |
| 251 | proof - | |
| 252 | obtain a q where p: "p = pCons a q" by (cases p) | |
| 253 | with assms have "q = 0" by (cases "q = 0") simp_all | |
| 254 | with p have "p = pCons a 0" by simp | |
| 255 | with that show thesis . | |
| 256 | qed | |
| 257 | ||
| 29451 | 258 | |
| 60500 | 259 | subsection \<open>List-style syntax for polynomials\<close> | 
| 52380 | 260 | |
| 261 | syntax | |
| 262 |   "_poly" :: "args \<Rightarrow> 'a poly"  ("[:(_):]")
 | |
| 263 | ||
| 264 | translations | |
| 265 | "[:x, xs:]" == "CONST pCons x [:xs:]" | |
| 266 | "[:x:]" == "CONST pCons x 0" | |
| 267 | "[:x:]" <= "CONST pCons x (_constrain 0 t)" | |
| 268 | ||
| 269 | ||
| 60500 | 270 | subsection \<open>Representation of polynomials by lists of coefficients\<close> | 
| 52380 | 271 | |
| 272 | primrec Poly :: "'a::zero list \<Rightarrow> 'a poly" | |
| 273 | where | |
| 54855 | 274 | [code_post]: "Poly [] = 0" | 
| 275 | | [code_post]: "Poly (a # as) = pCons a (Poly as)" | |
| 52380 | 276 | |
| 277 | lemma Poly_replicate_0 [simp]: | |
| 278 | "Poly (replicate n 0) = 0" | |
| 279 | by (induct n) simp_all | |
| 280 | ||
| 281 | lemma Poly_eq_0: | |
| 282 | "Poly as = 0 \<longleftrightarrow> (\<exists>n. as = replicate n 0)" | |
| 283 | by (induct as) (auto simp add: Cons_replicate_eq) | |
| 284 | ||
| 285 | definition coeffs :: "'a poly \<Rightarrow> 'a::zero list" | |
| 286 | where | |
| 287 | "coeffs p = (if p = 0 then [] else map (\<lambda>i. coeff p i) [0 ..< Suc (degree p)])" | |
| 288 | ||
| 289 | lemma coeffs_eq_Nil [simp]: | |
| 290 | "coeffs p = [] \<longleftrightarrow> p = 0" | |
| 291 | by (simp add: coeffs_def) | |
| 292 | ||
| 293 | lemma not_0_coeffs_not_Nil: | |
| 294 | "p \<noteq> 0 \<Longrightarrow> coeffs p \<noteq> []" | |
| 295 | by simp | |
| 296 | ||
| 297 | lemma coeffs_0_eq_Nil [simp]: | |
| 298 | "coeffs 0 = []" | |
| 299 | by simp | |
| 29454 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 300 | |
| 52380 | 301 | lemma coeffs_pCons_eq_cCons [simp]: | 
| 302 | "coeffs (pCons a p) = a ## coeffs p" | |
| 303 | proof - | |
| 304 |   { fix ms :: "nat list" and f :: "nat \<Rightarrow> 'a" and x :: "'a"
 | |
| 305 | assume "\<forall>m\<in>set ms. m > 0" | |
| 55415 | 306 | then have "map (case_nat x f) ms = map f (map (\<lambda>n. n - 1) ms)" | 
| 58199 
5fbe474b5da8
explicit theory with additional, less commonly used list operations
 haftmann parents: 
57862diff
changeset | 307 | by (induct ms) (auto split: nat.split) | 
| 
5fbe474b5da8
explicit theory with additional, less commonly used list operations
 haftmann parents: 
57862diff
changeset | 308 | } | 
| 52380 | 309 | note * = this | 
| 310 | show ?thesis | |
| 60570 | 311 | by (simp add: coeffs_def * upt_conv_Cons coeff_pCons map_decr_upt del: upt_Suc) | 
| 52380 | 312 | qed | 
| 313 | ||
| 314 | lemma not_0_cCons_eq [simp]: | |
| 315 | "p \<noteq> 0 \<Longrightarrow> a ## coeffs p = a # coeffs p" | |
| 316 | by (simp add: cCons_def) | |
| 317 | ||
| 318 | lemma Poly_coeffs [simp, code abstype]: | |
| 319 | "Poly (coeffs p) = p" | |
| 54856 | 320 | by (induct p) auto | 
| 52380 | 321 | |
| 322 | lemma coeffs_Poly [simp]: | |
| 323 | "coeffs (Poly as) = strip_while (HOL.eq 0) as" | |
| 324 | proof (induct as) | |
| 325 | case Nil then show ?case by simp | |
| 326 | next | |
| 327 | case (Cons a as) | |
| 328 | have "(\<forall>n. as \<noteq> replicate n 0) \<longleftrightarrow> (\<exists>a\<in>set as. a \<noteq> 0)" | |
| 329 | using replicate_length_same [of as 0] by (auto dest: sym [of _ as]) | |
| 330 | with Cons show ?case by auto | |
| 331 | qed | |
| 332 | ||
| 333 | lemma last_coeffs_not_0: | |
| 334 | "p \<noteq> 0 \<Longrightarrow> last (coeffs p) \<noteq> 0" | |
| 335 | by (induct p) (auto simp add: cCons_def) | |
| 336 | ||
| 337 | lemma strip_while_coeffs [simp]: | |
| 338 | "strip_while (HOL.eq 0) (coeffs p) = coeffs p" | |
| 339 | by (cases "p = 0") (auto dest: last_coeffs_not_0 intro: strip_while_not_last) | |
| 340 | ||
| 341 | lemma coeffs_eq_iff: | |
| 342 | "p = q \<longleftrightarrow> coeffs p = coeffs q" (is "?P \<longleftrightarrow> ?Q") | |
| 343 | proof | |
| 344 | assume ?P then show ?Q by simp | |
| 345 | next | |
| 346 | assume ?Q | |
| 347 | then have "Poly (coeffs p) = Poly (coeffs q)" by simp | |
| 348 | then show ?P by simp | |
| 349 | qed | |
| 350 | ||
| 351 | lemma coeff_Poly_eq: | |
| 352 | "coeff (Poly xs) n = nth_default 0 xs n" | |
| 353 | apply (induct xs arbitrary: n) apply simp_all | |
| 55642 
63beb38e9258
adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
 blanchet parents: 
55417diff
changeset | 354 | by (metis nat.case not0_implies_Suc nth_default_Cons_0 nth_default_Cons_Suc pCons.rep_eq) | 
| 29454 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 355 | |
| 52380 | 356 | lemma nth_default_coeffs_eq: | 
| 357 | "nth_default 0 (coeffs p) = coeff p" | |
| 358 | by (simp add: fun_eq_iff coeff_Poly_eq [symmetric]) | |
| 359 | ||
| 360 | lemma [code]: | |
| 361 | "coeff p = nth_default 0 (coeffs p)" | |
| 362 | by (simp add: nth_default_coeffs_eq) | |
| 363 | ||
| 364 | lemma coeffs_eqI: | |
| 365 | assumes coeff: "\<And>n. coeff p n = nth_default 0 xs n" | |
| 366 | assumes zero: "xs \<noteq> [] \<Longrightarrow> last xs \<noteq> 0" | |
| 367 | shows "coeffs p = xs" | |
| 368 | proof - | |
| 369 | from coeff have "p = Poly xs" by (simp add: poly_eq_iff coeff_Poly_eq) | |
| 370 | with zero show ?thesis by simp (cases xs, simp_all) | |
| 371 | qed | |
| 372 | ||
| 373 | lemma degree_eq_length_coeffs [code]: | |
| 374 | "degree p = length (coeffs p) - 1" | |
| 375 | by (simp add: coeffs_def) | |
| 376 | ||
| 377 | lemma length_coeffs_degree: | |
| 378 | "p \<noteq> 0 \<Longrightarrow> length (coeffs p) = Suc (degree p)" | |
| 379 | by (induct p) (auto simp add: cCons_def) | |
| 380 | ||
| 381 | lemma [code abstract]: | |
| 382 | "coeffs 0 = []" | |
| 383 | by (fact coeffs_0_eq_Nil) | |
| 384 | ||
| 385 | lemma [code abstract]: | |
| 386 | "coeffs (pCons a p) = a ## coeffs p" | |
| 387 | by (fact coeffs_pCons_eq_cCons) | |
| 388 | ||
| 389 | instantiation poly :: ("{zero, equal}") equal
 | |
| 390 | begin | |
| 391 | ||
| 392 | definition | |
| 393 | [code]: "HOL.equal (p::'a poly) q \<longleftrightarrow> HOL.equal (coeffs p) (coeffs q)" | |
| 394 | ||
| 60679 | 395 | instance | 
| 396 | by standard (simp add: equal equal_poly_def coeffs_eq_iff) | |
| 52380 | 397 | |
| 398 | end | |
| 399 | ||
| 60679 | 400 | lemma [code nbe]: "HOL.equal (p :: _ poly) p \<longleftrightarrow> True" | 
| 52380 | 401 | by (fact equal_refl) | 
| 29454 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 402 | |
| 52380 | 403 | definition is_zero :: "'a::zero poly \<Rightarrow> bool" | 
| 404 | where | |
| 405 | [code]: "is_zero p \<longleftrightarrow> List.null (coeffs p)" | |
| 406 | ||
| 407 | lemma is_zero_null [code_abbrev]: | |
| 408 | "is_zero p \<longleftrightarrow> p = 0" | |
| 409 | by (simp add: is_zero_def null_def) | |
| 410 | ||
| 411 | ||
| 60500 | 412 | subsection \<open>Fold combinator for polynomials\<close> | 
| 52380 | 413 | |
| 414 | definition fold_coeffs :: "('a::zero \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a poly \<Rightarrow> 'b \<Rightarrow> 'b"
 | |
| 415 | where | |
| 416 | "fold_coeffs f p = foldr f (coeffs p)" | |
| 417 | ||
| 418 | lemma fold_coeffs_0_eq [simp]: | |
| 419 | "fold_coeffs f 0 = id" | |
| 420 | by (simp add: fold_coeffs_def) | |
| 421 | ||
| 422 | lemma fold_coeffs_pCons_eq [simp]: | |
| 423 | "f 0 = id \<Longrightarrow> fold_coeffs f (pCons a p) = f a \<circ> fold_coeffs f p" | |
| 424 | by (simp add: fold_coeffs_def cCons_def fun_eq_iff) | |
| 29454 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 425 | |
| 52380 | 426 | lemma fold_coeffs_pCons_0_0_eq [simp]: | 
| 427 | "fold_coeffs f (pCons 0 0) = id" | |
| 428 | by (simp add: fold_coeffs_def) | |
| 429 | ||
| 430 | lemma fold_coeffs_pCons_coeff_not_0_eq [simp]: | |
| 431 | "a \<noteq> 0 \<Longrightarrow> fold_coeffs f (pCons a p) = f a \<circ> fold_coeffs f p" | |
| 432 | by (simp add: fold_coeffs_def) | |
| 433 | ||
| 434 | lemma fold_coeffs_pCons_not_0_0_eq [simp]: | |
| 435 | "p \<noteq> 0 \<Longrightarrow> fold_coeffs f (pCons a p) = f a \<circ> fold_coeffs f p" | |
| 436 | by (simp add: fold_coeffs_def) | |
| 437 | ||
| 438 | ||
| 60500 | 439 | subsection \<open>Canonical morphism on polynomials -- evaluation\<close> | 
| 52380 | 440 | |
| 441 | definition poly :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a" | |
| 442 | where | |
| 60500 | 443 | "poly p = fold_coeffs (\<lambda>a f x. a + x * f x) p (\<lambda>x. 0)" -- \<open>The Horner Schema\<close> | 
| 52380 | 444 | |
| 445 | lemma poly_0 [simp]: | |
| 446 | "poly 0 x = 0" | |
| 447 | by (simp add: poly_def) | |
| 448 | ||
| 449 | lemma poly_pCons [simp]: | |
| 450 | "poly (pCons a p) x = a + x * poly p x" | |
| 451 | by (cases "p = 0 \<and> a = 0") (auto simp add: poly_def) | |
| 29454 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 452 | |
| 
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
 huffman parents: 
29453diff
changeset | 453 | |
| 60500 | 454 | subsection \<open>Monomials\<close> | 
| 29451 | 455 | |
| 52380 | 456 | lift_definition monom :: "'a \<Rightarrow> nat \<Rightarrow> 'a::zero poly" | 
| 457 | is "\<lambda>a m n. if m = n then a else 0" | |
| 59983 
cd2efd7d06bd
replace almost_everywhere_zero by Infinite_Set.MOST
 hoelzl parents: 
59815diff
changeset | 458 | by (simp add: MOST_iff_cofinite) | 
| 52380 | 459 | |
| 460 | lemma coeff_monom [simp]: | |
| 461 | "coeff (monom a m) n = (if m = n then a else 0)" | |
| 462 | by transfer rule | |
| 29451 | 463 | |
| 52380 | 464 | lemma monom_0: | 
| 465 | "monom a 0 = pCons a 0" | |
| 466 | by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) | |
| 29451 | 467 | |
| 52380 | 468 | lemma monom_Suc: | 
| 469 | "monom a (Suc n) = pCons 0 (monom a n)" | |
| 470 | by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) | |
| 29451 | 471 | |
| 472 | lemma monom_eq_0 [simp]: "monom 0 n = 0" | |
| 52380 | 473 | by (rule poly_eqI) simp | 
| 29451 | 474 | |
| 475 | lemma monom_eq_0_iff [simp]: "monom a n = 0 \<longleftrightarrow> a = 0" | |
| 52380 | 476 | by (simp add: poly_eq_iff) | 
| 29451 | 477 | |
| 478 | lemma monom_eq_iff [simp]: "monom a n = monom b n \<longleftrightarrow> a = b" | |
| 52380 | 479 | by (simp add: poly_eq_iff) | 
| 29451 | 480 | |
| 481 | lemma degree_monom_le: "degree (monom a n) \<le> n" | |
| 482 | by (rule degree_le, simp) | |
| 483 | ||
| 484 | lemma degree_monom_eq: "a \<noteq> 0 \<Longrightarrow> degree (monom a n) = n" | |
| 485 | apply (rule order_antisym [OF degree_monom_le]) | |
| 486 | apply (rule le_degree, simp) | |
| 487 | done | |
| 488 | ||
| 52380 | 489 | lemma coeffs_monom [code abstract]: | 
| 490 | "coeffs (monom a n) = (if a = 0 then [] else replicate n 0 @ [a])" | |
| 491 | by (induct n) (simp_all add: monom_0 monom_Suc) | |
| 492 | ||
| 493 | lemma fold_coeffs_monom [simp]: | |
| 494 | "a \<noteq> 0 \<Longrightarrow> fold_coeffs f (monom a n) = f 0 ^^ n \<circ> f a" | |
| 495 | by (simp add: fold_coeffs_def coeffs_monom fun_eq_iff) | |
| 496 | ||
| 497 | lemma poly_monom: | |
| 498 |   fixes a x :: "'a::{comm_semiring_1}"
 | |
| 499 | shows "poly (monom a n) x = a * x ^ n" | |
| 500 | by (cases "a = 0", simp_all) | |
| 501 | (induct n, simp_all add: mult.left_commute poly_def) | |
| 502 | ||
| 29451 | 503 | |
| 60500 | 504 | subsection \<open>Addition and subtraction\<close> | 
| 29451 | 505 | |
| 506 | instantiation poly :: (comm_monoid_add) comm_monoid_add | |
| 507 | begin | |
| 508 | ||
| 52380 | 509 | lift_definition plus_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" | 
| 510 | is "\<lambda>p q n. coeff p n + coeff q n" | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59983diff
changeset | 511 | proof - | 
| 60679 | 512 | fix q p :: "'a poly" | 
| 513 | show "\<forall>\<^sub>\<infinity>n. coeff p n + coeff q n = 0" | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59983diff
changeset | 514 | using MOST_coeff_eq_0[of p] MOST_coeff_eq_0[of q] by eventually_elim simp | 
| 52380 | 515 | qed | 
| 29451 | 516 | |
| 60679 | 517 | lemma coeff_add [simp]: "coeff (p + q) n = coeff p n + coeff q n" | 
| 52380 | 518 | by (simp add: plus_poly.rep_eq) | 
| 29451 | 519 | |
| 60679 | 520 | instance | 
| 521 | proof | |
| 29451 | 522 | fix p q r :: "'a poly" | 
| 523 | show "(p + q) + r = p + (q + r)" | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57482diff
changeset | 524 | by (simp add: poly_eq_iff add.assoc) | 
| 29451 | 525 | show "p + q = q + p" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57482diff
changeset | 526 | by (simp add: poly_eq_iff add.commute) | 
| 29451 | 527 | show "0 + p = p" | 
| 52380 | 528 | by (simp add: poly_eq_iff) | 
| 29451 | 529 | qed | 
| 530 | ||
| 531 | end | |
| 532 | ||
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59557diff
changeset | 533 | instantiation poly :: (cancel_comm_monoid_add) cancel_comm_monoid_add | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59557diff
changeset | 534 | begin | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59557diff
changeset | 535 | |
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59557diff
changeset | 536 | lift_definition minus_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59557diff
changeset | 537 | is "\<lambda>p q n. coeff p n - coeff q n" | 
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59983diff
changeset | 538 | proof - | 
| 60679 | 539 | fix q p :: "'a poly" | 
| 540 | show "\<forall>\<^sub>\<infinity>n. coeff p n - coeff q n = 0" | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59983diff
changeset | 541 | using MOST_coeff_eq_0[of p] MOST_coeff_eq_0[of q] by eventually_elim simp | 
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59557diff
changeset | 542 | qed | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59557diff
changeset | 543 | |
| 60679 | 544 | lemma coeff_diff [simp]: "coeff (p - q) n = coeff p n - coeff q n" | 
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59557diff
changeset | 545 | by (simp add: minus_poly.rep_eq) | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59557diff
changeset | 546 | |
| 60679 | 547 | instance | 
| 548 | proof | |
| 29540 | 549 | fix p q r :: "'a poly" | 
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59557diff
changeset | 550 | show "p + q - p = q" | 
| 52380 | 551 | by (simp add: poly_eq_iff) | 
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59557diff
changeset | 552 | show "p - q - r = p - (q + r)" | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59557diff
changeset | 553 | by (simp add: poly_eq_iff diff_diff_eq) | 
| 29540 | 554 | qed | 
| 555 | ||
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59557diff
changeset | 556 | end | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59557diff
changeset | 557 | |
| 29451 | 558 | instantiation poly :: (ab_group_add) ab_group_add | 
| 559 | begin | |
| 560 | ||
| 52380 | 561 | lift_definition uminus_poly :: "'a poly \<Rightarrow> 'a poly" | 
| 562 | is "\<lambda>p n. - coeff p n" | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59983diff
changeset | 563 | proof - | 
| 60679 | 564 | fix p :: "'a poly" | 
| 565 | show "\<forall>\<^sub>\<infinity>n. - coeff p n = 0" | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59983diff
changeset | 566 | using MOST_coeff_eq_0 by simp | 
| 52380 | 567 | qed | 
| 29451 | 568 | |
| 569 | lemma coeff_minus [simp]: "coeff (- p) n = - coeff p n" | |
| 52380 | 570 | by (simp add: uminus_poly.rep_eq) | 
| 29451 | 571 | |
| 60679 | 572 | instance | 
| 573 | proof | |
| 29451 | 574 | fix p q :: "'a poly" | 
| 575 | show "- p + p = 0" | |
| 52380 | 576 | by (simp add: poly_eq_iff) | 
| 29451 | 577 | show "p - q = p + - q" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
52380diff
changeset | 578 | by (simp add: poly_eq_iff) | 
| 29451 | 579 | qed | 
| 580 | ||
| 581 | end | |
| 582 | ||
| 583 | lemma add_pCons [simp]: | |
| 584 | "pCons a p + pCons b q = pCons (a + b) (p + q)" | |
| 52380 | 585 | by (rule poly_eqI, simp add: coeff_pCons split: nat.split) | 
| 29451 | 586 | |
| 587 | lemma minus_pCons [simp]: | |
| 588 | "- pCons a p = pCons (- a) (- p)" | |
| 52380 | 589 | by (rule poly_eqI, simp add: coeff_pCons split: nat.split) | 
| 29451 | 590 | |
| 591 | lemma diff_pCons [simp]: | |
| 592 | "pCons a p - pCons b q = pCons (a - b) (p - q)" | |
| 52380 | 593 | by (rule poly_eqI, simp add: coeff_pCons split: nat.split) | 
| 29451 | 594 | |
| 29539 | 595 | lemma degree_add_le_max: "degree (p + q) \<le> max (degree p) (degree q)" | 
| 29451 | 596 | by (rule degree_le, auto simp add: coeff_eq_0) | 
| 597 | ||
| 29539 | 598 | lemma degree_add_le: | 
| 599 | "\<lbrakk>degree p \<le> n; degree q \<le> n\<rbrakk> \<Longrightarrow> degree (p + q) \<le> n" | |
| 600 | by (auto intro: order_trans degree_add_le_max) | |
| 601 | ||
| 29453 | 602 | lemma degree_add_less: | 
| 603 | "\<lbrakk>degree p < n; degree q < n\<rbrakk> \<Longrightarrow> degree (p + q) < n" | |
| 29539 | 604 | by (auto intro: le_less_trans degree_add_le_max) | 
| 29453 | 605 | |
| 29451 | 606 | lemma degree_add_eq_right: | 
| 607 | "degree p < degree q \<Longrightarrow> degree (p + q) = degree q" | |
| 608 | apply (cases "q = 0", simp) | |
| 609 | apply (rule order_antisym) | |
| 29539 | 610 | apply (simp add: degree_add_le) | 
| 29451 | 611 | apply (rule le_degree) | 
| 612 | apply (simp add: coeff_eq_0) | |
| 613 | done | |
| 614 | ||
| 615 | lemma degree_add_eq_left: | |
| 616 | "degree q < degree p \<Longrightarrow> degree (p + q) = degree p" | |
| 617 | using degree_add_eq_right [of q p] | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57482diff
changeset | 618 | by (simp add: add.commute) | 
| 29451 | 619 | |
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59557diff
changeset | 620 | lemma degree_minus [simp]: | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59557diff
changeset | 621 | "degree (- p) = degree p" | 
| 29451 | 622 | unfolding degree_def by simp | 
| 623 | ||
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59557diff
changeset | 624 | lemma degree_diff_le_max: | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59557diff
changeset | 625 | fixes p q :: "'a :: ab_group_add poly" | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59557diff
changeset | 626 | shows "degree (p - q) \<le> max (degree p) (degree q)" | 
| 29451 | 627 | using degree_add_le [where p=p and q="-q"] | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
52380diff
changeset | 628 | by simp | 
| 29451 | 629 | |
| 29539 | 630 | lemma degree_diff_le: | 
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59557diff
changeset | 631 | fixes p q :: "'a :: ab_group_add poly" | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59557diff
changeset | 632 | assumes "degree p \<le> n" and "degree q \<le> n" | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59557diff
changeset | 633 | shows "degree (p - q) \<le> n" | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59557diff
changeset | 634 | using assms degree_add_le [of p n "- q"] by simp | 
| 29539 | 635 | |
| 29453 | 636 | lemma degree_diff_less: | 
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59557diff
changeset | 637 | fixes p q :: "'a :: ab_group_add poly" | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59557diff
changeset | 638 | assumes "degree p < n" and "degree q < n" | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59557diff
changeset | 639 | shows "degree (p - q) < n" | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59557diff
changeset | 640 | using assms degree_add_less [of p n "- q"] by simp | 
| 29453 | 641 | |
| 29451 | 642 | lemma add_monom: "monom a n + monom b n = monom (a + b) n" | 
| 52380 | 643 | by (rule poly_eqI) simp | 
| 29451 | 644 | |
| 645 | lemma diff_monom: "monom a n - monom b n = monom (a - b) n" | |
| 52380 | 646 | by (rule poly_eqI) simp | 
| 29451 | 647 | |
| 648 | lemma minus_monom: "- monom a n = monom (-a) n" | |
| 52380 | 649 | by (rule poly_eqI) simp | 
| 29451 | 650 | |
| 651 | lemma coeff_setsum: "coeff (\<Sum>x\<in>A. p x) i = (\<Sum>x\<in>A. coeff (p x) i)" | |
| 652 | by (cases "finite A", induct set: finite, simp_all) | |
| 653 | ||
| 654 | lemma monom_setsum: "monom (\<Sum>x\<in>A. a x) n = (\<Sum>x\<in>A. monom (a x) n)" | |
| 52380 | 655 | by (rule poly_eqI) (simp add: coeff_setsum) | 
| 656 | ||
| 657 | fun plus_coeffs :: "'a::comm_monoid_add list \<Rightarrow> 'a list \<Rightarrow> 'a list" | |
| 658 | where | |
| 659 | "plus_coeffs xs [] = xs" | |
| 660 | | "plus_coeffs [] ys = ys" | |
| 661 | | "plus_coeffs (x # xs) (y # ys) = (x + y) ## plus_coeffs xs ys" | |
| 662 | ||
| 663 | lemma coeffs_plus_eq_plus_coeffs [code abstract]: | |
| 664 | "coeffs (p + q) = plus_coeffs (coeffs p) (coeffs q)" | |
| 665 | proof - | |
| 666 |   { fix xs ys :: "'a list" and n
 | |
| 667 | have "nth_default 0 (plus_coeffs xs ys) n = nth_default 0 xs n + nth_default 0 ys n" | |
| 668 | proof (induct xs ys arbitrary: n rule: plus_coeffs.induct) | |
| 60679 | 669 | case (3 x xs y ys n) | 
| 670 | then show ?case by (cases n) (auto simp add: cCons_def) | |
| 52380 | 671 | qed simp_all } | 
| 672 | note * = this | |
| 673 |   { fix xs ys :: "'a list"
 | |
| 674 | assume "xs \<noteq> [] \<Longrightarrow> last xs \<noteq> 0" and "ys \<noteq> [] \<Longrightarrow> last ys \<noteq> 0" | |
| 675 | moreover assume "plus_coeffs xs ys \<noteq> []" | |
| 676 | ultimately have "last (plus_coeffs xs ys) \<noteq> 0" | |
| 677 | proof (induct xs ys rule: plus_coeffs.induct) | |
| 678 | case (3 x xs y ys) then show ?case by (auto simp add: cCons_def) metis | |
| 679 | qed simp_all } | |
| 680 | note ** = this | |
| 681 | show ?thesis | |
| 682 | apply (rule coeffs_eqI) | |
| 683 | apply (simp add: * nth_default_coeffs_eq) | |
| 684 | apply (rule **) | |
| 685 | apply (auto dest: last_coeffs_not_0) | |
| 686 | done | |
| 687 | qed | |
| 688 | ||
| 689 | lemma coeffs_uminus [code abstract]: | |
| 690 | "coeffs (- p) = map (\<lambda>a. - a) (coeffs p)" | |
| 691 | by (rule coeffs_eqI) | |
| 692 | (simp_all add: not_0_coeffs_not_Nil last_map last_coeffs_not_0 nth_default_map_eq nth_default_coeffs_eq) | |
| 693 | ||
| 694 | lemma [code]: | |
| 695 | fixes p q :: "'a::ab_group_add poly" | |
| 696 | shows "p - q = p + - q" | |
| 59557 | 697 | by (fact diff_conv_add_uminus) | 
| 52380 | 698 | |
| 699 | lemma poly_add [simp]: "poly (p + q) x = poly p x + poly q x" | |
| 700 | apply (induct p arbitrary: q, simp) | |
| 701 | apply (case_tac q, simp, simp add: algebra_simps) | |
| 702 | done | |
| 703 | ||
| 704 | lemma poly_minus [simp]: | |
| 705 | fixes x :: "'a::comm_ring" | |
| 706 | shows "poly (- p) x = - poly p x" | |
| 707 | by (induct p) simp_all | |
| 708 | ||
| 709 | lemma poly_diff [simp]: | |
| 710 | fixes x :: "'a::comm_ring" | |
| 711 | shows "poly (p - q) x = poly p x - poly q x" | |
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
52380diff
changeset | 712 | using poly_add [of p "- q" x] by simp | 
| 52380 | 713 | |
| 714 | lemma poly_setsum: "poly (\<Sum>k\<in>A. p k) x = (\<Sum>k\<in>A. poly (p k) x)" | |
| 715 | by (induct A rule: infinite_finite_induct) simp_all | |
| 29451 | 716 | |
| 717 | ||
| 60500 | 718 | subsection \<open>Multiplication by a constant, polynomial multiplication and the unit polynomial\<close> | 
| 29451 | 719 | |
| 52380 | 720 | lift_definition smult :: "'a::comm_semiring_0 \<Rightarrow> 'a poly \<Rightarrow> 'a poly" | 
| 721 | is "\<lambda>a p n. a * coeff p n" | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59983diff
changeset | 722 | proof - | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59983diff
changeset | 723 | fix a :: 'a and p :: "'a poly" show "\<forall>\<^sub>\<infinity> i. a * coeff p i = 0" | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59983diff
changeset | 724 | using MOST_coeff_eq_0[of p] by eventually_elim simp | 
| 52380 | 725 | qed | 
| 29451 | 726 | |
| 52380 | 727 | lemma coeff_smult [simp]: | 
| 728 | "coeff (smult a p) n = a * coeff p n" | |
| 729 | by (simp add: smult.rep_eq) | |
| 29451 | 730 | |
| 731 | lemma degree_smult_le: "degree (smult a p) \<le> degree p" | |
| 732 | by (rule degree_le, simp add: coeff_eq_0) | |
| 733 | ||
| 29472 | 734 | lemma smult_smult [simp]: "smult a (smult b p) = smult (a * b) p" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57482diff
changeset | 735 | by (rule poly_eqI, simp add: mult.assoc) | 
| 29451 | 736 | |
| 737 | lemma smult_0_right [simp]: "smult a 0 = 0" | |
| 52380 | 738 | by (rule poly_eqI, simp) | 
| 29451 | 739 | |
| 740 | lemma smult_0_left [simp]: "smult 0 p = 0" | |
| 52380 | 741 | by (rule poly_eqI, simp) | 
| 29451 | 742 | |
| 743 | lemma smult_1_left [simp]: "smult (1::'a::comm_semiring_1) p = p" | |
| 52380 | 744 | by (rule poly_eqI, simp) | 
| 29451 | 745 | |
| 746 | lemma smult_add_right: | |
| 747 | "smult a (p + q) = smult a p + smult a q" | |
| 52380 | 748 | by (rule poly_eqI, simp add: algebra_simps) | 
| 29451 | 749 | |
| 750 | lemma smult_add_left: | |
| 751 | "smult (a + b) p = smult a p + smult b p" | |
| 52380 | 752 | by (rule poly_eqI, simp add: algebra_simps) | 
| 29451 | 753 | |
| 29457 | 754 | lemma smult_minus_right [simp]: | 
| 29451 | 755 | "smult (a::'a::comm_ring) (- p) = - smult a p" | 
| 52380 | 756 | by (rule poly_eqI, simp) | 
| 29451 | 757 | |
| 29457 | 758 | lemma smult_minus_left [simp]: | 
| 29451 | 759 | "smult (- a::'a::comm_ring) p = - smult a p" | 
| 52380 | 760 | by (rule poly_eqI, simp) | 
| 29451 | 761 | |
| 762 | lemma smult_diff_right: | |
| 763 | "smult (a::'a::comm_ring) (p - q) = smult a p - smult a q" | |
| 52380 | 764 | by (rule poly_eqI, simp add: algebra_simps) | 
| 29451 | 765 | |
| 766 | lemma smult_diff_left: | |
| 767 | "smult (a - b::'a::comm_ring) p = smult a p - smult b p" | |
| 52380 | 768 | by (rule poly_eqI, simp add: algebra_simps) | 
| 29451 | 769 | |
| 29472 | 770 | lemmas smult_distribs = | 
| 771 | smult_add_left smult_add_right | |
| 772 | smult_diff_left smult_diff_right | |
| 773 | ||
| 29451 | 774 | lemma smult_pCons [simp]: | 
| 775 | "smult a (pCons b p) = pCons (a * b) (smult a p)" | |
| 52380 | 776 | by (rule poly_eqI, simp add: coeff_pCons split: nat.split) | 
| 29451 | 777 | |
| 778 | lemma smult_monom: "smult a (monom b n) = monom (a * b) n" | |
| 779 | by (induct n, simp add: monom_0, simp add: monom_Suc) | |
| 780 | ||
| 29659 | 781 | lemma degree_smult_eq [simp]: | 
| 782 | fixes a :: "'a::idom" | |
| 783 | shows "degree (smult a p) = (if a = 0 then 0 else degree p)" | |
| 784 | by (cases "a = 0", simp, simp add: degree_def) | |
| 785 | ||
| 786 | lemma smult_eq_0_iff [simp]: | |
| 787 | fixes a :: "'a::idom" | |
| 788 | shows "smult a p = 0 \<longleftrightarrow> a = 0 \<or> p = 0" | |
| 52380 | 789 | by (simp add: poly_eq_iff) | 
| 29451 | 790 | |
| 52380 | 791 | lemma coeffs_smult [code abstract]: | 
| 792 | fixes p :: "'a::idom poly" | |
| 793 | shows "coeffs (smult a p) = (if a = 0 then [] else map (Groups.times a) (coeffs p))" | |
| 794 | by (rule coeffs_eqI) | |
| 795 | (auto simp add: not_0_coeffs_not_Nil last_map last_coeffs_not_0 nth_default_map_eq nth_default_coeffs_eq) | |
| 29451 | 796 | |
| 797 | instantiation poly :: (comm_semiring_0) comm_semiring_0 | |
| 798 | begin | |
| 799 | ||
| 800 | definition | |
| 52380 | 801 | "p * q = fold_coeffs (\<lambda>a p. smult a q + pCons 0 p) p 0" | 
| 29474 | 802 | |
| 803 | lemma mult_poly_0_left: "(0::'a poly) * q = 0" | |
| 52380 | 804 | by (simp add: times_poly_def) | 
| 29474 | 805 | |
| 806 | lemma mult_pCons_left [simp]: | |
| 807 | "pCons a p * q = smult a q + pCons 0 (p * q)" | |
| 52380 | 808 | by (cases "p = 0 \<and> a = 0") (auto simp add: times_poly_def) | 
| 29474 | 809 | |
| 810 | lemma mult_poly_0_right: "p * (0::'a poly) = 0" | |
| 52380 | 811 | by (induct p) (simp add: mult_poly_0_left, simp) | 
| 29451 | 812 | |
| 29474 | 813 | lemma mult_pCons_right [simp]: | 
| 814 | "p * pCons a q = smult a p + pCons 0 (p * q)" | |
| 52380 | 815 | by (induct p) (simp add: mult_poly_0_left, simp add: algebra_simps) | 
| 29474 | 816 | |
| 817 | lemmas mult_poly_0 = mult_poly_0_left mult_poly_0_right | |
| 818 | ||
| 52380 | 819 | lemma mult_smult_left [simp]: | 
| 820 | "smult a p * q = smult a (p * q)" | |
| 821 | by (induct p) (simp add: mult_poly_0, simp add: smult_add_right) | |
| 29474 | 822 | |
| 52380 | 823 | lemma mult_smult_right [simp]: | 
| 824 | "p * smult a q = smult a (p * q)" | |
| 825 | by (induct q) (simp add: mult_poly_0, simp add: smult_add_right) | |
| 29474 | 826 | |
| 827 | lemma mult_poly_add_left: | |
| 828 | fixes p q r :: "'a poly" | |
| 829 | shows "(p + q) * r = p * r + q * r" | |
| 52380 | 830 | by (induct r) (simp add: mult_poly_0, simp add: smult_distribs algebra_simps) | 
| 29451 | 831 | |
| 60679 | 832 | instance | 
| 833 | proof | |
| 29451 | 834 | fix p q r :: "'a poly" | 
| 835 | show 0: "0 * p = 0" | |
| 29474 | 836 | by (rule mult_poly_0_left) | 
| 29451 | 837 | show "p * 0 = 0" | 
| 29474 | 838 | by (rule mult_poly_0_right) | 
| 29451 | 839 | show "(p + q) * r = p * r + q * r" | 
| 29474 | 840 | by (rule mult_poly_add_left) | 
| 29451 | 841 | show "(p * q) * r = p * (q * r)" | 
| 29474 | 842 | by (induct p, simp add: mult_poly_0, simp add: mult_poly_add_left) | 
| 29451 | 843 | show "p * q = q * p" | 
| 29474 | 844 | by (induct p, simp add: mult_poly_0, simp) | 
| 29451 | 845 | qed | 
| 846 | ||
| 847 | end | |
| 848 | ||
| 29540 | 849 | instance poly :: (comm_semiring_0_cancel) comm_semiring_0_cancel .. | 
| 850 | ||
| 29474 | 851 | lemma coeff_mult: | 
| 852 | "coeff (p * q) n = (\<Sum>i\<le>n. coeff p i * coeff q (n-i))" | |
| 853 | proof (induct p arbitrary: n) | |
| 854 | case 0 show ?case by simp | |
| 855 | next | |
| 856 | case (pCons a p n) thus ?case | |
| 857 | by (cases n, simp, simp add: setsum_atMost_Suc_shift | |
| 858 | del: setsum_atMost_Suc) | |
| 859 | qed | |
| 29451 | 860 | |
| 29474 | 861 | lemma degree_mult_le: "degree (p * q) \<le> degree p + degree q" | 
| 862 | apply (rule degree_le) | |
| 863 | apply (induct p) | |
| 864 | apply simp | |
| 865 | apply (simp add: coeff_eq_0 coeff_pCons split: nat.split) | |
| 29451 | 866 | done | 
| 867 | ||
| 868 | lemma mult_monom: "monom a m * monom b n = monom (a * b) (m + n)" | |
| 60679 | 869 | by (induct m) (simp add: monom_0 smult_monom, simp add: monom_Suc) | 
| 29451 | 870 | |
| 871 | instantiation poly :: (comm_semiring_1) comm_semiring_1 | |
| 872 | begin | |
| 873 | ||
| 60679 | 874 | definition one_poly_def: "1 = pCons 1 0" | 
| 29451 | 875 | |
| 60679 | 876 | instance | 
| 877 | proof | |
| 878 | show "1 * p = p" for p :: "'a poly" | |
| 52380 | 879 | unfolding one_poly_def by simp | 
| 29451 | 880 | show "0 \<noteq> (1::'a poly)" | 
| 881 | unfolding one_poly_def by simp | |
| 882 | qed | |
| 883 | ||
| 884 | end | |
| 885 | ||
| 52380 | 886 | instance poly :: (comm_ring) comm_ring .. | 
| 887 | ||
| 888 | instance poly :: (comm_ring_1) comm_ring_1 .. | |
| 889 | ||
| 29451 | 890 | lemma coeff_1 [simp]: "coeff 1 n = (if n = 0 then 1 else 0)" | 
| 891 | unfolding one_poly_def | |
| 892 | by (simp add: coeff_pCons split: nat.split) | |
| 893 | ||
| 60570 | 894 | lemma monom_eq_1 [simp]: | 
| 895 | "monom 1 0 = 1" | |
| 896 | by (simp add: monom_0 one_poly_def) | |
| 897 | ||
| 29451 | 898 | lemma degree_1 [simp]: "degree 1 = 0" | 
| 899 | unfolding one_poly_def | |
| 900 | by (rule degree_pCons_0) | |
| 901 | ||
| 52380 | 902 | lemma coeffs_1_eq [simp, code abstract]: | 
| 903 | "coeffs 1 = [1]" | |
| 904 | by (simp add: one_poly_def) | |
| 905 | ||
| 906 | lemma degree_power_le: | |
| 907 | "degree (p ^ n) \<le> degree p * n" | |
| 908 | by (induct n) (auto intro: order_trans degree_mult_le) | |
| 909 | ||
| 910 | lemma poly_smult [simp]: | |
| 911 | "poly (smult a p) x = a * poly p x" | |
| 912 | by (induct p, simp, simp add: algebra_simps) | |
| 913 | ||
| 914 | lemma poly_mult [simp]: | |
| 915 | "poly (p * q) x = poly p x * poly q x" | |
| 916 | by (induct p, simp_all, simp add: algebra_simps) | |
| 917 | ||
| 918 | lemma poly_1 [simp]: | |
| 919 | "poly 1 x = 1" | |
| 920 | by (simp add: one_poly_def) | |
| 921 | ||
| 922 | lemma poly_power [simp]: | |
| 923 |   fixes p :: "'a::{comm_semiring_1} poly"
 | |
| 924 | shows "poly (p ^ n) x = poly p x ^ n" | |
| 925 | by (induct n) simp_all | |
| 926 | ||
| 927 | ||
| 60500 | 928 | subsection \<open>Lemmas about divisibility\<close> | 
| 29979 | 929 | |
| 930 | lemma dvd_smult: "p dvd q \<Longrightarrow> p dvd smult a q" | |
| 931 | proof - | |
| 932 | assume "p dvd q" | |
| 933 | then obtain k where "q = p * k" .. | |
| 934 | then have "smult a q = p * smult a k" by simp | |
| 935 | then show "p dvd smult a q" .. | |
| 936 | qed | |
| 937 | ||
| 938 | lemma dvd_smult_cancel: | |
| 939 | fixes a :: "'a::field" | |
| 940 | shows "p dvd smult a q \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> p dvd q" | |
| 941 | by (drule dvd_smult [where a="inverse a"]) simp | |
| 942 | ||
| 943 | lemma dvd_smult_iff: | |
| 944 | fixes a :: "'a::field" | |
| 945 | shows "a \<noteq> 0 \<Longrightarrow> p dvd smult a q \<longleftrightarrow> p dvd q" | |
| 946 | by (safe elim!: dvd_smult dvd_smult_cancel) | |
| 947 | ||
| 31663 | 948 | lemma smult_dvd_cancel: | 
| 949 | "smult a p dvd q \<Longrightarrow> p dvd q" | |
| 950 | proof - | |
| 951 | assume "smult a p dvd q" | |
| 952 | then obtain k where "q = smult a p * k" .. | |
| 953 | then have "q = p * smult a k" by simp | |
| 954 | then show "p dvd q" .. | |
| 955 | qed | |
| 956 | ||
| 957 | lemma smult_dvd: | |
| 958 | fixes a :: "'a::field" | |
| 959 | shows "p dvd q \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> smult a p dvd q" | |
| 960 | by (rule smult_dvd_cancel [where a="inverse a"]) simp | |
| 961 | ||
| 962 | lemma smult_dvd_iff: | |
| 963 | fixes a :: "'a::field" | |
| 964 | shows "smult a p dvd q \<longleftrightarrow> (if a = 0 then q = 0 else p dvd q)" | |
| 965 | by (auto elim: smult_dvd smult_dvd_cancel) | |
| 966 | ||
| 29451 | 967 | |
| 60500 | 968 | subsection \<open>Polynomials form an integral domain\<close> | 
| 29451 | 969 | |
| 970 | lemma coeff_mult_degree_sum: | |
| 971 | "coeff (p * q) (degree p + degree q) = | |
| 972 | coeff p (degree p) * coeff q (degree q)" | |
| 29471 | 973 | by (induct p, simp, simp add: coeff_eq_0) | 
| 29451 | 974 | |
| 975 | instance poly :: (idom) idom | |
| 976 | proof | |
| 977 | fix p q :: "'a poly" | |
| 978 | assume "p \<noteq> 0" and "q \<noteq> 0" | |
| 979 | have "coeff (p * q) (degree p + degree q) = | |
| 980 | coeff p (degree p) * coeff q (degree q)" | |
| 981 | by (rule coeff_mult_degree_sum) | |
| 982 | also have "coeff p (degree p) * coeff q (degree q) \<noteq> 0" | |
| 60500 | 983 | using \<open>p \<noteq> 0\<close> and \<open>q \<noteq> 0\<close> by simp | 
| 29451 | 984 | finally have "\<exists>n. coeff (p * q) n \<noteq> 0" .. | 
| 52380 | 985 | thus "p * q \<noteq> 0" by (simp add: poly_eq_iff) | 
| 29451 | 986 | qed | 
| 987 | ||
| 988 | lemma degree_mult_eq: | |
| 989 | fixes p q :: "'a::idom poly" | |
| 990 | shows "\<lbrakk>p \<noteq> 0; q \<noteq> 0\<rbrakk> \<Longrightarrow> degree (p * q) = degree p + degree q" | |
| 991 | apply (rule order_antisym [OF degree_mult_le le_degree]) | |
| 992 | apply (simp add: coeff_mult_degree_sum) | |
| 993 | done | |
| 994 | ||
| 60570 | 995 | lemma degree_mult_right_le: | 
| 996 | fixes p q :: "'a::idom poly" | |
| 997 | assumes "q \<noteq> 0" | |
| 998 | shows "degree p \<le> degree (p * q)" | |
| 999 | using assms by (cases "p = 0") (simp_all add: degree_mult_eq) | |
| 1000 | ||
| 1001 | lemma coeff_degree_mult: | |
| 1002 | fixes p q :: "'a::idom poly" | |
| 1003 | shows "coeff (p * q) (degree (p * q)) = | |
| 1004 | coeff q (degree q) * coeff p (degree p)" | |
| 1005 | by (cases "p = 0 \<or> q = 0") (auto simp add: degree_mult_eq coeff_mult_degree_sum) | |
| 1006 | ||
| 29451 | 1007 | lemma dvd_imp_degree_le: | 
| 1008 | fixes p q :: "'a::idom poly" | |
| 1009 | shows "\<lbrakk>p dvd q; q \<noteq> 0\<rbrakk> \<Longrightarrow> degree p \<le> degree q" | |
| 1010 | by (erule dvdE, simp add: degree_mult_eq) | |
| 1011 | ||
| 1012 | ||
| 60500 | 1013 | subsection \<open>Polynomials form an ordered integral domain\<close> | 
| 29878 | 1014 | |
| 52380 | 1015 | definition pos_poly :: "'a::linordered_idom poly \<Rightarrow> bool" | 
| 29878 | 1016 | where | 
| 1017 | "pos_poly p \<longleftrightarrow> 0 < coeff p (degree p)" | |
| 1018 | ||
| 1019 | lemma pos_poly_pCons: | |
| 1020 | "pos_poly (pCons a p) \<longleftrightarrow> pos_poly p \<or> (p = 0 \<and> 0 < a)" | |
| 1021 | unfolding pos_poly_def by simp | |
| 1022 | ||
| 1023 | lemma not_pos_poly_0 [simp]: "\<not> pos_poly 0" | |
| 1024 | unfolding pos_poly_def by simp | |
| 1025 | ||
| 1026 | lemma pos_poly_add: "\<lbrakk>pos_poly p; pos_poly q\<rbrakk> \<Longrightarrow> pos_poly (p + q)" | |
| 1027 | apply (induct p arbitrary: q, simp) | |
| 1028 | apply (case_tac q, force simp add: pos_poly_pCons add_pos_pos) | |
| 1029 | done | |
| 1030 | ||
| 1031 | lemma pos_poly_mult: "\<lbrakk>pos_poly p; pos_poly q\<rbrakk> \<Longrightarrow> pos_poly (p * q)" | |
| 1032 | unfolding pos_poly_def | |
| 1033 | apply (subgoal_tac "p \<noteq> 0 \<and> q \<noteq> 0") | |
| 56544 | 1034 | apply (simp add: degree_mult_eq coeff_mult_degree_sum) | 
| 29878 | 1035 | apply auto | 
| 1036 | done | |
| 1037 | ||
| 1038 | lemma pos_poly_total: "p = 0 \<or> pos_poly p \<or> pos_poly (- p)" | |
| 1039 | by (induct p) (auto simp add: pos_poly_pCons) | |
| 1040 | ||
| 52380 | 1041 | lemma last_coeffs_eq_coeff_degree: | 
| 1042 | "p \<noteq> 0 \<Longrightarrow> last (coeffs p) = coeff p (degree p)" | |
| 1043 | by (simp add: coeffs_def) | |
| 1044 | ||
| 1045 | lemma pos_poly_coeffs [code]: | |
| 1046 | "pos_poly p \<longleftrightarrow> (let as = coeffs p in as \<noteq> [] \<and> last as > 0)" (is "?P \<longleftrightarrow> ?Q") | |
| 1047 | proof | |
| 1048 | assume ?Q then show ?P by (auto simp add: pos_poly_def last_coeffs_eq_coeff_degree) | |
| 1049 | next | |
| 1050 | assume ?P then have *: "0 < coeff p (degree p)" by (simp add: pos_poly_def) | |
| 1051 | then have "p \<noteq> 0" by auto | |
| 1052 | with * show ?Q by (simp add: last_coeffs_eq_coeff_degree) | |
| 1053 | qed | |
| 1054 | ||
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34973diff
changeset | 1055 | instantiation poly :: (linordered_idom) linordered_idom | 
| 29878 | 1056 | begin | 
| 1057 | ||
| 1058 | definition | |
| 37765 | 1059 | "x < y \<longleftrightarrow> pos_poly (y - x)" | 
| 29878 | 1060 | |
| 1061 | definition | |
| 37765 | 1062 | "x \<le> y \<longleftrightarrow> x = y \<or> pos_poly (y - x)" | 
| 29878 | 1063 | |
| 1064 | definition | |
| 37765 | 1065 | "abs (x::'a poly) = (if x < 0 then - x else x)" | 
| 29878 | 1066 | |
| 1067 | definition | |
| 37765 | 1068 | "sgn (x::'a poly) = (if x = 0 then 0 else if 0 < x then 1 else - 1)" | 
| 29878 | 1069 | |
| 60679 | 1070 | instance | 
| 1071 | proof | |
| 1072 | fix x y z :: "'a poly" | |
| 29878 | 1073 | show "x < y \<longleftrightarrow> x \<le> y \<and> \<not> y \<le> x" | 
| 1074 | unfolding less_eq_poly_def less_poly_def | |
| 1075 | apply safe | |
| 1076 | apply simp | |
| 1077 | apply (drule (1) pos_poly_add) | |
| 1078 | apply simp | |
| 1079 | done | |
| 60679 | 1080 | show "x \<le> x" | 
| 29878 | 1081 | unfolding less_eq_poly_def by simp | 
| 60679 | 1082 | show "x \<le> y \<Longrightarrow> y \<le> z \<Longrightarrow> x \<le> z" | 
| 29878 | 1083 | unfolding less_eq_poly_def | 
| 1084 | apply safe | |
| 1085 | apply (drule (1) pos_poly_add) | |
| 1086 | apply (simp add: algebra_simps) | |
| 1087 | done | |
| 60679 | 1088 | show "x \<le> y \<Longrightarrow> y \<le> x \<Longrightarrow> x = y" | 
| 29878 | 1089 | unfolding less_eq_poly_def | 
| 1090 | apply safe | |
| 1091 | apply (drule (1) pos_poly_add) | |
| 1092 | apply simp | |
| 1093 | done | |
| 60679 | 1094 | show "x \<le> y \<Longrightarrow> z + x \<le> z + y" | 
| 29878 | 1095 | unfolding less_eq_poly_def | 
| 1096 | apply safe | |
| 1097 | apply (simp add: algebra_simps) | |
| 1098 | done | |
| 1099 | show "x \<le> y \<or> y \<le> x" | |
| 1100 | unfolding less_eq_poly_def | |
| 1101 | using pos_poly_total [of "x - y"] | |
| 1102 | by auto | |
| 60679 | 1103 | show "x < y \<Longrightarrow> 0 < z \<Longrightarrow> z * x < z * y" | 
| 29878 | 1104 | unfolding less_poly_def | 
| 1105 | by (simp add: right_diff_distrib [symmetric] pos_poly_mult) | |
| 1106 | show "\<bar>x\<bar> = (if x < 0 then - x else x)" | |
| 1107 | by (rule abs_poly_def) | |
| 1108 | show "sgn x = (if x = 0 then 0 else if 0 < x then 1 else - 1)" | |
| 1109 | by (rule sgn_poly_def) | |
| 1110 | qed | |
| 1111 | ||
| 1112 | end | |
| 1113 | ||
| 60500 | 1114 | text \<open>TODO: Simplification rules for comparisons\<close> | 
| 29878 | 1115 | |
| 1116 | ||
| 60500 | 1117 | subsection \<open>Synthetic division and polynomial roots\<close> | 
| 52380 | 1118 | |
| 60500 | 1119 | text \<open> | 
| 52380 | 1120 |   Synthetic division is simply division by the linear polynomial @{term "x - c"}.
 | 
| 60500 | 1121 | \<close> | 
| 52380 | 1122 | |
| 1123 | definition synthetic_divmod :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a poly \<times> 'a" | |
| 1124 | where | |
| 1125 | "synthetic_divmod p c = fold_coeffs (\<lambda>a (q, r). (pCons r q, a + c * r)) p (0, 0)" | |
| 1126 | ||
| 1127 | definition synthetic_div :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a poly" | |
| 1128 | where | |
| 1129 | "synthetic_div p c = fst (synthetic_divmod p c)" | |
| 1130 | ||
| 1131 | lemma synthetic_divmod_0 [simp]: | |
| 1132 | "synthetic_divmod 0 c = (0, 0)" | |
| 1133 | by (simp add: synthetic_divmod_def) | |
| 1134 | ||
| 1135 | lemma synthetic_divmod_pCons [simp]: | |
| 1136 | "synthetic_divmod (pCons a p) c = (\<lambda>(q, r). (pCons r q, a + c * r)) (synthetic_divmod p c)" | |
| 1137 | by (cases "p = 0 \<and> a = 0") (auto simp add: synthetic_divmod_def) | |
| 1138 | ||
| 1139 | lemma synthetic_div_0 [simp]: | |
| 1140 | "synthetic_div 0 c = 0" | |
| 1141 | unfolding synthetic_div_def by simp | |
| 1142 | ||
| 1143 | lemma synthetic_div_unique_lemma: "smult c p = pCons a p \<Longrightarrow> p = 0" | |
| 1144 | by (induct p arbitrary: a) simp_all | |
| 1145 | ||
| 1146 | lemma snd_synthetic_divmod: | |
| 1147 | "snd (synthetic_divmod p c) = poly p c" | |
| 1148 | by (induct p, simp, simp add: split_def) | |
| 1149 | ||
| 1150 | lemma synthetic_div_pCons [simp]: | |
| 1151 | "synthetic_div (pCons a p) c = pCons (poly p c) (synthetic_div p c)" | |
| 1152 | unfolding synthetic_div_def | |
| 1153 | by (simp add: split_def snd_synthetic_divmod) | |
| 1154 | ||
| 1155 | lemma synthetic_div_eq_0_iff: | |
| 1156 | "synthetic_div p c = 0 \<longleftrightarrow> degree p = 0" | |
| 1157 | by (induct p, simp, case_tac p, simp) | |
| 1158 | ||
| 1159 | lemma degree_synthetic_div: | |
| 1160 | "degree (synthetic_div p c) = degree p - 1" | |
| 1161 | by (induct p, simp, simp add: synthetic_div_eq_0_iff) | |
| 1162 | ||
| 1163 | lemma synthetic_div_correct: | |
| 1164 | "p + smult c (synthetic_div p c) = pCons (poly p c) (synthetic_div p c)" | |
| 1165 | by (induct p) simp_all | |
| 1166 | ||
| 1167 | lemma synthetic_div_unique: | |
| 1168 | "p + smult c q = pCons r q \<Longrightarrow> r = poly p c \<and> q = synthetic_div p c" | |
| 1169 | apply (induct p arbitrary: q r) | |
| 1170 | apply (simp, frule synthetic_div_unique_lemma, simp) | |
| 1171 | apply (case_tac q, force) | |
| 1172 | done | |
| 1173 | ||
| 1174 | lemma synthetic_div_correct': | |
| 1175 | fixes c :: "'a::comm_ring_1" | |
| 1176 | shows "[:-c, 1:] * synthetic_div p c + [:poly p c:] = p" | |
| 1177 | using synthetic_div_correct [of p c] | |
| 1178 | by (simp add: algebra_simps) | |
| 1179 | ||
| 1180 | lemma poly_eq_0_iff_dvd: | |
| 1181 | fixes c :: "'a::idom" | |
| 1182 | shows "poly p c = 0 \<longleftrightarrow> [:-c, 1:] dvd p" | |
| 1183 | proof | |
| 1184 | assume "poly p c = 0" | |
| 1185 | with synthetic_div_correct' [of c p] | |
| 1186 | have "p = [:-c, 1:] * synthetic_div p c" by simp | |
| 1187 | then show "[:-c, 1:] dvd p" .. | |
| 1188 | next | |
| 1189 | assume "[:-c, 1:] dvd p" | |
| 1190 | then obtain k where "p = [:-c, 1:] * k" by (rule dvdE) | |
| 1191 | then show "poly p c = 0" by simp | |
| 1192 | qed | |
| 1193 | ||
| 1194 | lemma dvd_iff_poly_eq_0: | |
| 1195 | fixes c :: "'a::idom" | |
| 1196 | shows "[:c, 1:] dvd p \<longleftrightarrow> poly p (-c) = 0" | |
| 1197 | by (simp add: poly_eq_0_iff_dvd) | |
| 1198 | ||
| 1199 | lemma poly_roots_finite: | |
| 1200 | fixes p :: "'a::idom poly" | |
| 1201 |   shows "p \<noteq> 0 \<Longrightarrow> finite {x. poly p x = 0}"
 | |
| 1202 | proof (induct n \<equiv> "degree p" arbitrary: p) | |
| 1203 | case (0 p) | |
| 1204 | then obtain a where "a \<noteq> 0" and "p = [:a:]" | |
| 1205 | by (cases p, simp split: if_splits) | |
| 1206 |   then show "finite {x. poly p x = 0}" by simp
 | |
| 1207 | next | |
| 1208 | case (Suc n p) | |
| 1209 |   show "finite {x. poly p x = 0}"
 | |
| 1210 | proof (cases "\<exists>x. poly p x = 0") | |
| 1211 | case False | |
| 1212 |     then show "finite {x. poly p x = 0}" by simp
 | |
| 1213 | next | |
| 1214 | case True | |
| 1215 | then obtain a where "poly p a = 0" .. | |
| 1216 | then have "[:-a, 1:] dvd p" by (simp only: poly_eq_0_iff_dvd) | |
| 1217 | then obtain k where k: "p = [:-a, 1:] * k" .. | |
| 60500 | 1218 | with \<open>p \<noteq> 0\<close> have "k \<noteq> 0" by auto | 
| 52380 | 1219 | with k have "degree p = Suc (degree k)" | 
| 1220 | by (simp add: degree_mult_eq del: mult_pCons_left) | |
| 60500 | 1221 | with \<open>Suc n = degree p\<close> have "n = degree k" by simp | 
| 1222 |     then have "finite {x. poly k x = 0}" using \<open>k \<noteq> 0\<close> by (rule Suc.hyps)
 | |
| 52380 | 1223 |     then have "finite (insert a {x. poly k x = 0})" by simp
 | 
| 1224 |     then show "finite {x. poly p x = 0}"
 | |
| 57862 | 1225 | by (simp add: k Collect_disj_eq del: mult_pCons_left) | 
| 52380 | 1226 | qed | 
| 1227 | qed | |
| 1228 | ||
| 1229 | lemma poly_eq_poly_eq_iff: | |
| 1230 |   fixes p q :: "'a::{idom,ring_char_0} poly"
 | |
| 1231 | shows "poly p = poly q \<longleftrightarrow> p = q" (is "?P \<longleftrightarrow> ?Q") | |
| 1232 | proof | |
| 1233 | assume ?Q then show ?P by simp | |
| 1234 | next | |
| 1235 |   { fix p :: "'a::{idom,ring_char_0} poly"
 | |
| 1236 | have "poly p = poly 0 \<longleftrightarrow> p = 0" | |
| 1237 | apply (cases "p = 0", simp_all) | |
| 1238 | apply (drule poly_roots_finite) | |
| 1239 | apply (auto simp add: infinite_UNIV_char_0) | |
| 1240 | done | |
| 1241 | } note this [of "p - q"] | |
| 1242 | moreover assume ?P | |
| 1243 | ultimately show ?Q by auto | |
| 1244 | qed | |
| 1245 | ||
| 1246 | lemma poly_all_0_iff_0: | |
| 1247 |   fixes p :: "'a::{ring_char_0, idom} poly"
 | |
| 1248 | shows "(\<forall>x. poly p x = 0) \<longleftrightarrow> p = 0" | |
| 1249 | by (auto simp add: poly_eq_poly_eq_iff [symmetric]) | |
| 1250 | ||
| 1251 | ||
| 60500 | 1252 | subsection \<open>Long division of polynomials\<close> | 
| 29451 | 1253 | |
| 52380 | 1254 | definition pdivmod_rel :: "'a::field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<Rightarrow> bool" | 
| 29451 | 1255 | where | 
| 29537 | 1256 | "pdivmod_rel x y q r \<longleftrightarrow> | 
| 29451 | 1257 | x = q * y + r \<and> (if y = 0 then q = 0 else r = 0 \<or> degree r < degree y)" | 
| 1258 | ||
| 29537 | 1259 | lemma pdivmod_rel_0: | 
| 1260 | "pdivmod_rel 0 y 0 0" | |
| 1261 | unfolding pdivmod_rel_def by simp | |
| 29451 | 1262 | |
| 29537 | 1263 | lemma pdivmod_rel_by_0: | 
| 1264 | "pdivmod_rel x 0 0 x" | |
| 1265 | unfolding pdivmod_rel_def by simp | |
| 29451 | 1266 | |
| 1267 | lemma eq_zero_or_degree_less: | |
| 1268 | assumes "degree p \<le> n" and "coeff p n = 0" | |
| 1269 | shows "p = 0 \<or> degree p < n" | |
| 1270 | proof (cases n) | |
| 1271 | case 0 | |
| 60500 | 1272 | with \<open>degree p \<le> n\<close> and \<open>coeff p n = 0\<close> | 
| 29451 | 1273 | have "coeff p (degree p) = 0" by simp | 
| 1274 | then have "p = 0" by simp | |
| 1275 | then show ?thesis .. | |
| 1276 | next | |
| 1277 | case (Suc m) | |
| 1278 | have "\<forall>i>n. coeff p i = 0" | |
| 60500 | 1279 | using \<open>degree p \<le> n\<close> by (simp add: coeff_eq_0) | 
| 29451 | 1280 | then have "\<forall>i\<ge>n. coeff p i = 0" | 
| 60500 | 1281 | using \<open>coeff p n = 0\<close> by (simp add: le_less) | 
| 29451 | 1282 | then have "\<forall>i>m. coeff p i = 0" | 
| 60500 | 1283 | using \<open>n = Suc m\<close> by (simp add: less_eq_Suc_le) | 
| 29451 | 1284 | then have "degree p \<le> m" | 
| 1285 | by (rule degree_le) | |
| 1286 | then have "degree p < n" | |
| 60500 | 1287 | using \<open>n = Suc m\<close> by (simp add: less_Suc_eq_le) | 
| 29451 | 1288 | then show ?thesis .. | 
| 1289 | qed | |
| 1290 | ||
| 29537 | 1291 | lemma pdivmod_rel_pCons: | 
| 1292 | assumes rel: "pdivmod_rel x y q r" | |
| 29451 | 1293 | assumes y: "y \<noteq> 0" | 
| 1294 | assumes b: "b = coeff (pCons a r) (degree y) / coeff y (degree y)" | |
| 29537 | 1295 | shows "pdivmod_rel (pCons a x) y (pCons b q) (pCons a r - smult b y)" | 
| 1296 | (is "pdivmod_rel ?x y ?q ?r") | |
| 29451 | 1297 | proof - | 
| 1298 | have x: "x = q * y + r" and r: "r = 0 \<or> degree r < degree y" | |
| 29537 | 1299 | using assms unfolding pdivmod_rel_def by simp_all | 
| 29451 | 1300 | |
| 1301 | have 1: "?x = ?q * y + ?r" | |
| 1302 | using b x by simp | |
| 1303 | ||
| 1304 | have 2: "?r = 0 \<or> degree ?r < degree y" | |
| 1305 | proof (rule eq_zero_or_degree_less) | |
| 29539 | 1306 | show "degree ?r \<le> degree y" | 
| 1307 | proof (rule degree_diff_le) | |
| 29451 | 1308 | 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 | 1309 | using r by auto | 
| 29451 | 1310 | show "degree (smult b y) \<le> degree y" | 
| 1311 | by (rule degree_smult_le) | |
| 1312 | qed | |
| 1313 | next | |
| 1314 | show "coeff ?r (degree y) = 0" | |
| 60500 | 1315 | using \<open>y \<noteq> 0\<close> unfolding b by simp | 
| 29451 | 1316 | qed | 
| 1317 | ||
| 1318 | from 1 2 show ?thesis | |
| 29537 | 1319 | unfolding pdivmod_rel_def | 
| 60500 | 1320 | using \<open>y \<noteq> 0\<close> by simp | 
| 29451 | 1321 | qed | 
| 1322 | ||
| 29537 | 1323 | lemma pdivmod_rel_exists: "\<exists>q r. pdivmod_rel x y q r" | 
| 29451 | 1324 | apply (cases "y = 0") | 
| 29537 | 1325 | apply (fast intro!: pdivmod_rel_by_0) | 
| 29451 | 1326 | apply (induct x) | 
| 29537 | 1327 | apply (fast intro!: pdivmod_rel_0) | 
| 1328 | apply (fast intro!: pdivmod_rel_pCons) | |
| 29451 | 1329 | done | 
| 1330 | ||
| 29537 | 1331 | lemma pdivmod_rel_unique: | 
| 1332 | assumes 1: "pdivmod_rel x y q1 r1" | |
| 1333 | assumes 2: "pdivmod_rel x y q2 r2" | |
| 29451 | 1334 | shows "q1 = q2 \<and> r1 = r2" | 
| 1335 | proof (cases "y = 0") | |
| 1336 | assume "y = 0" with assms show ?thesis | |
| 29537 | 1337 | by (simp add: pdivmod_rel_def) | 
| 29451 | 1338 | next | 
| 1339 | assume [simp]: "y \<noteq> 0" | |
| 1340 | from 1 have q1: "x = q1 * y + r1" and r1: "r1 = 0 \<or> degree r1 < degree y" | |
| 29537 | 1341 | unfolding pdivmod_rel_def by simp_all | 
| 29451 | 1342 | from 2 have q2: "x = q2 * y + r2" and r2: "r2 = 0 \<or> degree r2 < degree y" | 
| 29537 | 1343 | unfolding pdivmod_rel_def by simp_all | 
| 29451 | 1344 | from q1 q2 have q3: "(q1 - q2) * y = r2 - r1" | 
| 29667 | 1345 | by (simp add: algebra_simps) | 
| 29451 | 1346 | from r1 r2 have r3: "(r2 - r1) = 0 \<or> degree (r2 - r1) < degree y" | 
| 29453 | 1347 | by (auto intro: degree_diff_less) | 
| 29451 | 1348 | |
| 1349 | show "q1 = q2 \<and> r1 = r2" | |
| 1350 | proof (rule ccontr) | |
| 1351 | assume "\<not> (q1 = q2 \<and> r1 = r2)" | |
| 1352 | with q3 have "q1 \<noteq> q2" and "r1 \<noteq> r2" by auto | |
| 1353 | with r3 have "degree (r2 - r1) < degree y" by simp | |
| 1354 | also have "degree y \<le> degree (q1 - q2) + degree y" by simp | |
| 1355 | also have "\<dots> = degree ((q1 - q2) * y)" | |
| 60500 | 1356 | using \<open>q1 \<noteq> q2\<close> by (simp add: degree_mult_eq) | 
| 29451 | 1357 | also have "\<dots> = degree (r2 - r1)" | 
| 1358 | using q3 by simp | |
| 1359 | finally have "degree (r2 - r1) < degree (r2 - r1)" . | |
| 1360 | then show "False" by simp | |
| 1361 | qed | |
| 1362 | qed | |
| 1363 | ||
| 29660 | 1364 | lemma pdivmod_rel_0_iff: "pdivmod_rel 0 y q r \<longleftrightarrow> q = 0 \<and> r = 0" | 
| 1365 | by (auto dest: pdivmod_rel_unique intro: pdivmod_rel_0) | |
| 1366 | ||
| 1367 | lemma pdivmod_rel_by_0_iff: "pdivmod_rel x 0 q r \<longleftrightarrow> q = 0 \<and> r = x" | |
| 1368 | by (auto dest: pdivmod_rel_unique intro: pdivmod_rel_by_0) | |
| 1369 | ||
| 45605 | 1370 | lemmas pdivmod_rel_unique_div = pdivmod_rel_unique [THEN conjunct1] | 
| 29451 | 1371 | |
| 45605 | 1372 | lemmas pdivmod_rel_unique_mod = pdivmod_rel_unique [THEN conjunct2] | 
| 29451 | 1373 | |
| 1374 | instantiation poly :: (field) ring_div | |
| 1375 | begin | |
| 1376 | ||
| 60352 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
60040diff
changeset | 1377 | definition divide_poly where | 
| 60429 
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
 haftmann parents: 
60352diff
changeset | 1378 | div_poly_def: "x div y = (THE q. \<exists>r. pdivmod_rel x y q r)" | 
| 29451 | 1379 | |
| 1380 | definition mod_poly where | |
| 37765 | 1381 | "x mod y = (THE r. \<exists>q. pdivmod_rel x y q r)" | 
| 29451 | 1382 | |
| 1383 | lemma div_poly_eq: | |
| 60429 
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
 haftmann parents: 
60352diff
changeset | 1384 | "pdivmod_rel x y q r \<Longrightarrow> x div y = q" | 
| 29451 | 1385 | unfolding div_poly_def | 
| 29537 | 1386 | by (fast elim: pdivmod_rel_unique_div) | 
| 29451 | 1387 | |
| 1388 | lemma mod_poly_eq: | |
| 29537 | 1389 | "pdivmod_rel x y q r \<Longrightarrow> x mod y = r" | 
| 29451 | 1390 | unfolding mod_poly_def | 
| 29537 | 1391 | by (fast elim: pdivmod_rel_unique_mod) | 
| 29451 | 1392 | |
| 29537 | 1393 | lemma pdivmod_rel: | 
| 60429 
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
 haftmann parents: 
60352diff
changeset | 1394 | "pdivmod_rel x y (x div y) (x mod y)" | 
| 29451 | 1395 | proof - | 
| 29537 | 1396 | from pdivmod_rel_exists | 
| 1397 | obtain q r where "pdivmod_rel x y q r" by fast | |
| 29451 | 1398 | thus ?thesis | 
| 1399 | by (simp add: div_poly_eq mod_poly_eq) | |
| 1400 | qed | |
| 1401 | ||
| 60679 | 1402 | instance | 
| 1403 | proof | |
| 29451 | 1404 | fix x y :: "'a poly" | 
| 60429 
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
 haftmann parents: 
60352diff
changeset | 1405 | show "x div y * y + x mod y = x" | 
| 29537 | 1406 | using pdivmod_rel [of x y] | 
| 1407 | by (simp add: pdivmod_rel_def) | |
| 29451 | 1408 | next | 
| 1409 | fix x :: "'a poly" | |
| 29537 | 1410 | have "pdivmod_rel x 0 0 x" | 
| 1411 | by (rule pdivmod_rel_by_0) | |
| 60429 
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
 haftmann parents: 
60352diff
changeset | 1412 | thus "x div 0 = 0" | 
| 29451 | 1413 | by (rule div_poly_eq) | 
| 1414 | next | |
| 1415 | fix y :: "'a poly" | |
| 29537 | 1416 | have "pdivmod_rel 0 y 0 0" | 
| 1417 | by (rule pdivmod_rel_0) | |
| 60429 
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
 haftmann parents: 
60352diff
changeset | 1418 | thus "0 div y = 0" | 
| 29451 | 1419 | by (rule div_poly_eq) | 
| 1420 | next | |
| 1421 | fix x y z :: "'a poly" | |
| 1422 | assume "y \<noteq> 0" | |
| 60429 
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
 haftmann parents: 
60352diff
changeset | 1423 | hence "pdivmod_rel (x + z * y) y (z + x div y) (x mod y)" | 
| 29537 | 1424 | using pdivmod_rel [of x y] | 
| 49962 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 webertj parents: 
49834diff
changeset | 1425 | by (simp add: pdivmod_rel_def distrib_right) | 
| 60429 
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
 haftmann parents: 
60352diff
changeset | 1426 | thus "(x + z * y) div y = z + x div y" | 
| 29451 | 1427 | by (rule div_poly_eq) | 
| 30930 | 1428 | next | 
| 1429 | fix x y z :: "'a poly" | |
| 1430 | assume "x \<noteq> 0" | |
| 60429 
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
 haftmann parents: 
60352diff
changeset | 1431 | show "(x * y) div (x * z) = y div z" | 
| 30930 | 1432 | proof (cases "y \<noteq> 0 \<and> z \<noteq> 0") | 
| 1433 | have "\<And>x::'a poly. pdivmod_rel x 0 0 x" | |
| 1434 | by (rule pdivmod_rel_by_0) | |
| 60429 
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
 haftmann parents: 
60352diff
changeset | 1435 | then have [simp]: "\<And>x::'a poly. x div 0 = 0" | 
| 30930 | 1436 | by (rule div_poly_eq) | 
| 1437 | have "\<And>x::'a poly. pdivmod_rel 0 x 0 0" | |
| 1438 | by (rule pdivmod_rel_0) | |
| 60429 
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
 haftmann parents: 
60352diff
changeset | 1439 | then have [simp]: "\<And>x::'a poly. 0 div x = 0" | 
| 30930 | 1440 | by (rule div_poly_eq) | 
| 1441 | case False then show ?thesis by auto | |
| 1442 | next | |
| 1443 | case True then have "y \<noteq> 0" and "z \<noteq> 0" by auto | |
| 60500 | 1444 | with \<open>x \<noteq> 0\<close> | 
| 30930 | 1445 | have "\<And>q r. pdivmod_rel y z q r \<Longrightarrow> pdivmod_rel (x * y) (x * z) q (x * r)" | 
| 1446 | by (auto simp add: pdivmod_rel_def algebra_simps) | |
| 1447 | (rule classical, simp add: degree_mult_eq) | |
| 60429 
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
 haftmann parents: 
60352diff
changeset | 1448 | moreover from pdivmod_rel have "pdivmod_rel y z (y div z) (y mod z)" . | 
| 
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
 haftmann parents: 
60352diff
changeset | 1449 | ultimately have "pdivmod_rel (x * y) (x * z) (y div z) (x * (y mod z))" . | 
| 30930 | 1450 | then show ?thesis by (simp add: div_poly_eq) | 
| 1451 | qed | |
| 29451 | 1452 | qed | 
| 1453 | ||
| 1454 | end | |
| 1455 | ||
| 60570 | 1456 | lemma is_unit_monom_0: | 
| 1457 | fixes a :: "'a::field" | |
| 1458 | assumes "a \<noteq> 0" | |
| 1459 | shows "is_unit (monom a 0)" | |
| 1460 | proof | |
| 1461 | from assms show "1 = monom a 0 * monom (1 / a) 0" | |
| 1462 | by (simp add: mult_monom) | |
| 1463 | qed | |
| 1464 | ||
| 1465 | lemma is_unit_triv: | |
| 1466 | fixes a :: "'a::field" | |
| 1467 | assumes "a \<noteq> 0" | |
| 1468 | shows "is_unit [:a:]" | |
| 1469 | using assms by (simp add: is_unit_monom_0 monom_0 [symmetric]) | |
| 1470 | ||
| 1471 | lemma is_unit_iff_degree: | |
| 1472 | assumes "p \<noteq> 0" | |
| 1473 | shows "is_unit p \<longleftrightarrow> degree p = 0" (is "?P \<longleftrightarrow> ?Q") | |
| 1474 | proof | |
| 1475 | assume ?Q | |
| 1476 | then obtain a where "p = [:a:]" by (rule degree_eq_zeroE) | |
| 1477 | with assms show ?P by (simp add: is_unit_triv) | |
| 1478 | next | |
| 1479 | assume ?P | |
| 1480 | then obtain q where "q \<noteq> 0" "p * q = 1" .. | |
| 1481 | then have "degree (p * q) = degree 1" | |
| 1482 | by simp | |
| 1483 | with \<open>p \<noteq> 0\<close> \<open>q \<noteq> 0\<close> have "degree p + degree q = 0" | |
| 1484 | by (simp add: degree_mult_eq) | |
| 1485 | then show ?Q by simp | |
| 1486 | qed | |
| 1487 | ||
| 1488 | lemma is_unit_pCons_iff: | |
| 1489 | "is_unit (pCons a p) \<longleftrightarrow> p = 0 \<and> a \<noteq> 0" (is "?P \<longleftrightarrow> ?Q") | |
| 1490 | by (cases "p = 0") (auto simp add: is_unit_triv is_unit_iff_degree) | |
| 1491 | ||
| 1492 | lemma is_unit_monom_trival: | |
| 1493 | fixes p :: "'a::field poly" | |
| 1494 | assumes "is_unit p" | |
| 1495 | shows "monom (coeff p (degree p)) 0 = p" | |
| 1496 | using assms by (cases p) (simp_all add: monom_0 is_unit_pCons_iff) | |
| 1497 | ||
| 60685 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1498 | lemma is_unit_polyE: | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1499 | assumes "is_unit p" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1500 | obtains a where "p = monom a 0" and "a \<noteq> 0" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1501 | proof - | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1502 | obtain a q where "p = pCons a q" by (cases p) | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1503 | with assms have "p = [:a:]" and "a \<noteq> 0" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1504 | by (simp_all add: is_unit_pCons_iff) | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1505 | with that show thesis by (simp add: monom_0) | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1506 | qed | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1507 | |
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1508 | instantiation poly :: (field) normalization_semidom | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1509 | begin | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1510 | |
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1511 | definition normalize_poly :: "'a poly \<Rightarrow> 'a poly" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1512 | where "normalize_poly p = smult (1 / coeff p (degree p)) p" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1513 | |
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1514 | definition unit_factor_poly :: "'a poly \<Rightarrow> 'a poly" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1515 | where "unit_factor_poly p = monom (coeff p (degree p)) 0" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1516 | |
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1517 | instance | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1518 | proof | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1519 | fix p :: "'a poly" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1520 | show "unit_factor p * normalize p = p" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1521 | by (simp add: normalize_poly_def unit_factor_poly_def) | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1522 | (simp only: mult_smult_left [symmetric] smult_monom, simp) | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1523 | next | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1524 | show "normalize 0 = (0::'a poly)" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1525 | by (simp add: normalize_poly_def) | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1526 | next | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1527 | show "unit_factor 0 = (0::'a poly)" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1528 | by (simp add: unit_factor_poly_def) | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1529 | next | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1530 | fix p :: "'a poly" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1531 | assume "is_unit p" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1532 | then obtain a where "p = monom a 0" and "a \<noteq> 0" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1533 | by (rule is_unit_polyE) | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1534 | then show "normalize p = 1" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1535 | by (auto simp add: normalize_poly_def smult_monom degree_monom_eq) | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1536 | next | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1537 | fix p q :: "'a poly" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1538 | assume "q \<noteq> 0" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1539 | from \<open>q \<noteq> 0\<close> have "is_unit (monom (coeff q (degree q)) 0)" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1540 | by (auto intro: is_unit_monom_0) | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1541 | then show "is_unit (unit_factor q)" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1542 | by (simp add: unit_factor_poly_def) | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1543 | next | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1544 | fix p q :: "'a poly" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1545 | have "monom (coeff (p * q) (degree (p * q))) 0 = | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1546 | monom (coeff p (degree p)) 0 * monom (coeff q (degree q)) 0" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1547 | by (simp add: monom_0 coeff_degree_mult) | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1548 | then show "unit_factor (p * q) = | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1549 | unit_factor p * unit_factor q" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1550 | by (simp add: unit_factor_poly_def) | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1551 | qed | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1552 | |
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1553 | end | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60679diff
changeset | 1554 | |
| 29451 | 1555 | lemma degree_mod_less: | 
| 1556 | "y \<noteq> 0 \<Longrightarrow> x mod y = 0 \<or> degree (x mod y) < degree y" | |
| 29537 | 1557 | using pdivmod_rel [of x y] | 
| 1558 | unfolding pdivmod_rel_def by simp | |
| 29451 | 1559 | |
| 1560 | lemma div_poly_less: "degree x < degree y \<Longrightarrow> x div y = 0" | |
| 1561 | proof - | |
| 1562 | assume "degree x < degree y" | |
| 29537 | 1563 | hence "pdivmod_rel x y 0 x" | 
| 1564 | by (simp add: pdivmod_rel_def) | |
| 29451 | 1565 | thus "x div y = 0" by (rule div_poly_eq) | 
| 1566 | qed | |
| 1567 | ||
| 1568 | lemma mod_poly_less: "degree x < degree y \<Longrightarrow> x mod y = x" | |
| 1569 | proof - | |
| 1570 | assume "degree x < degree y" | |
| 29537 | 1571 | hence "pdivmod_rel x y 0 x" | 
| 1572 | by (simp add: pdivmod_rel_def) | |
| 29451 | 1573 | thus "x mod y = x" by (rule mod_poly_eq) | 
| 1574 | qed | |
| 1575 | ||
| 29659 | 1576 | lemma pdivmod_rel_smult_left: | 
| 1577 | "pdivmod_rel x y q r | |
| 1578 | \<Longrightarrow> pdivmod_rel (smult a x) y (smult a q) (smult a r)" | |
| 1579 | unfolding pdivmod_rel_def by (simp add: smult_add_right) | |
| 1580 | ||
| 1581 | lemma div_smult_left: "(smult a x) div y = smult a (x div y)" | |
| 1582 | by (rule div_poly_eq, rule pdivmod_rel_smult_left, rule pdivmod_rel) | |
| 1583 | ||
| 1584 | lemma mod_smult_left: "(smult a x) mod y = smult a (x mod y)" | |
| 1585 | by (rule mod_poly_eq, rule pdivmod_rel_smult_left, rule pdivmod_rel) | |
| 1586 | ||
| 30072 | 1587 | lemma poly_div_minus_left [simp]: | 
| 1588 | fixes x y :: "'a::field poly" | |
| 1589 | shows "(- x) div y = - (x div y)" | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 1590 | using div_smult_left [of "- 1::'a"] by simp | 
| 30072 | 1591 | |
| 1592 | lemma poly_mod_minus_left [simp]: | |
| 1593 | fixes x y :: "'a::field poly" | |
| 1594 | shows "(- x) mod y = - (x mod y)" | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 1595 | using mod_smult_left [of "- 1::'a"] by simp | 
| 30072 | 1596 | |
| 57482 
60459c3853af
add lemmas: polynomial div/mod distribute over addition
 huffman parents: 
56544diff
changeset | 1597 | lemma pdivmod_rel_add_left: | 
| 
60459c3853af
add lemmas: polynomial div/mod distribute over addition
 huffman parents: 
56544diff
changeset | 1598 | assumes "pdivmod_rel x y q r" | 
| 
60459c3853af
add lemmas: polynomial div/mod distribute over addition
 huffman parents: 
56544diff
changeset | 1599 | assumes "pdivmod_rel x' y q' r'" | 
| 
60459c3853af
add lemmas: polynomial div/mod distribute over addition
 huffman parents: 
56544diff
changeset | 1600 | shows "pdivmod_rel (x + x') y (q + q') (r + r')" | 
| 
60459c3853af
add lemmas: polynomial div/mod distribute over addition
 huffman parents: 
56544diff
changeset | 1601 | using assms unfolding pdivmod_rel_def | 
| 59557 | 1602 | by (auto simp add: algebra_simps degree_add_less) | 
| 57482 
60459c3853af
add lemmas: polynomial div/mod distribute over addition
 huffman parents: 
56544diff
changeset | 1603 | |
| 
60459c3853af
add lemmas: polynomial div/mod distribute over addition
 huffman parents: 
56544diff
changeset | 1604 | lemma poly_div_add_left: | 
| 
60459c3853af
add lemmas: polynomial div/mod distribute over addition
 huffman parents: 
56544diff
changeset | 1605 | fixes x y z :: "'a::field poly" | 
| 
60459c3853af
add lemmas: polynomial div/mod distribute over addition
 huffman parents: 
56544diff
changeset | 1606 | shows "(x + y) div z = x div z + y div z" | 
| 
60459c3853af
add lemmas: polynomial div/mod distribute over addition
 huffman parents: 
56544diff
changeset | 1607 | using pdivmod_rel_add_left [OF pdivmod_rel pdivmod_rel] | 
| 
60459c3853af
add lemmas: polynomial div/mod distribute over addition
 huffman parents: 
56544diff
changeset | 1608 | by (rule div_poly_eq) | 
| 
60459c3853af
add lemmas: polynomial div/mod distribute over addition
 huffman parents: 
56544diff
changeset | 1609 | |
| 
60459c3853af
add lemmas: polynomial div/mod distribute over addition
 huffman parents: 
56544diff
changeset | 1610 | lemma poly_mod_add_left: | 
| 
60459c3853af
add lemmas: polynomial div/mod distribute over addition
 huffman parents: 
56544diff
changeset | 1611 | fixes x y z :: "'a::field poly" | 
| 
60459c3853af
add lemmas: polynomial div/mod distribute over addition
 huffman parents: 
56544diff
changeset | 1612 | shows "(x + y) mod z = x mod z + y mod z" | 
| 
60459c3853af
add lemmas: polynomial div/mod distribute over addition
 huffman parents: 
56544diff
changeset | 1613 | using pdivmod_rel_add_left [OF pdivmod_rel pdivmod_rel] | 
| 
60459c3853af
add lemmas: polynomial div/mod distribute over addition
 huffman parents: 
56544diff
changeset | 1614 | by (rule mod_poly_eq) | 
| 
60459c3853af
add lemmas: polynomial div/mod distribute over addition
 huffman parents: 
56544diff
changeset | 1615 | |
| 
60459c3853af
add lemmas: polynomial div/mod distribute over addition
 huffman parents: 
56544diff
changeset | 1616 | lemma poly_div_diff_left: | 
| 
60459c3853af
add lemmas: polynomial div/mod distribute over addition
 huffman parents: 
56544diff
changeset | 1617 | fixes x y z :: "'a::field poly" | 
| 
60459c3853af
add lemmas: polynomial div/mod distribute over addition
 huffman parents: 
56544diff
changeset | 1618 | shows "(x - y) div z = x div z - y div z" | 
| 
60459c3853af
add lemmas: polynomial div/mod distribute over addition
 huffman parents: 
56544diff
changeset | 1619 | by (simp only: diff_conv_add_uminus poly_div_add_left poly_div_minus_left) | 
| 
60459c3853af
add lemmas: polynomial div/mod distribute over addition
 huffman parents: 
56544diff
changeset | 1620 | |
| 
60459c3853af
add lemmas: polynomial div/mod distribute over addition
 huffman parents: 
56544diff
changeset | 1621 | lemma poly_mod_diff_left: | 
| 
60459c3853af
add lemmas: polynomial div/mod distribute over addition
 huffman parents: 
56544diff
changeset | 1622 | fixes x y z :: "'a::field poly" | 
| 
60459c3853af
add lemmas: polynomial div/mod distribute over addition
 huffman parents: 
56544diff
changeset | 1623 | shows "(x - y) mod z = x mod z - y mod z" | 
| 
60459c3853af
add lemmas: polynomial div/mod distribute over addition
 huffman parents: 
56544diff
changeset | 1624 | by (simp only: diff_conv_add_uminus poly_mod_add_left poly_mod_minus_left) | 
| 
60459c3853af
add lemmas: polynomial div/mod distribute over addition
 huffman parents: 
56544diff
changeset | 1625 | |
| 29659 | 1626 | lemma pdivmod_rel_smult_right: | 
| 1627 | "\<lbrakk>a \<noteq> 0; pdivmod_rel x y q r\<rbrakk> | |
| 1628 | \<Longrightarrow> pdivmod_rel x (smult a y) (smult (inverse a) q) r" | |
| 1629 | unfolding pdivmod_rel_def by simp | |
| 1630 | ||
| 1631 | lemma div_smult_right: | |
| 1632 | "a \<noteq> 0 \<Longrightarrow> x div (smult a y) = smult (inverse a) (x div y)" | |
| 1633 | by (rule div_poly_eq, erule pdivmod_rel_smult_right, rule pdivmod_rel) | |
| 1634 | ||
| 1635 | lemma mod_smult_right: "a \<noteq> 0 \<Longrightarrow> x mod (smult a y) = x mod y" | |
| 1636 | by (rule mod_poly_eq, erule pdivmod_rel_smult_right, rule pdivmod_rel) | |
| 1637 | ||
| 30072 | 1638 | lemma poly_div_minus_right [simp]: | 
| 1639 | fixes x y :: "'a::field poly" | |
| 1640 | shows "x div (- y) = - (x div y)" | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 1641 | using div_smult_right [of "- 1::'a"] by (simp add: nonzero_inverse_minus_eq) | 
| 30072 | 1642 | |
| 1643 | lemma poly_mod_minus_right [simp]: | |
| 1644 | fixes x y :: "'a::field poly" | |
| 1645 | shows "x mod (- y) = x mod y" | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 1646 | using mod_smult_right [of "- 1::'a"] by simp | 
| 30072 | 1647 | |
| 29660 | 1648 | lemma pdivmod_rel_mult: | 
| 1649 | "\<lbrakk>pdivmod_rel x y q r; pdivmod_rel q z q' r'\<rbrakk> | |
| 1650 | \<Longrightarrow> pdivmod_rel x (y * z) q' (y * r' + r)" | |
| 1651 | apply (cases "z = 0", simp add: pdivmod_rel_def) | |
| 1652 | apply (cases "y = 0", simp add: pdivmod_rel_by_0_iff pdivmod_rel_0_iff) | |
| 1653 | apply (cases "r = 0") | |
| 1654 | apply (cases "r' = 0") | |
| 1655 | apply (simp add: pdivmod_rel_def) | |
| 36350 | 1656 | apply (simp add: pdivmod_rel_def field_simps degree_mult_eq) | 
| 29660 | 1657 | apply (cases "r' = 0") | 
| 1658 | apply (simp add: pdivmod_rel_def degree_mult_eq) | |
| 36350 | 1659 | apply (simp add: pdivmod_rel_def field_simps) | 
| 29660 | 1660 | apply (simp add: degree_mult_eq degree_add_less) | 
| 1661 | done | |
| 1662 | ||
| 1663 | lemma poly_div_mult_right: | |
| 1664 | fixes x y z :: "'a::field poly" | |
| 1665 | shows "x div (y * z) = (x div y) div z" | |
| 1666 | by (rule div_poly_eq, rule pdivmod_rel_mult, (rule pdivmod_rel)+) | |
| 1667 | ||
| 1668 | lemma poly_mod_mult_right: | |
| 1669 | fixes x y z :: "'a::field poly" | |
| 1670 | shows "x mod (y * z) = y * (x div y mod z) + x mod y" | |
| 1671 | by (rule mod_poly_eq, rule pdivmod_rel_mult, (rule pdivmod_rel)+) | |
| 1672 | ||
| 29451 | 1673 | lemma mod_pCons: | 
| 1674 | fixes a and x | |
| 1675 | assumes y: "y \<noteq> 0" | |
| 1676 | defines b: "b \<equiv> coeff (pCons a (x mod y)) (degree y) / coeff y (degree y)" | |
| 1677 | shows "(pCons a x) mod y = (pCons a (x mod y) - smult b y)" | |
| 1678 | unfolding b | |
| 1679 | apply (rule mod_poly_eq) | |
| 29537 | 1680 | apply (rule pdivmod_rel_pCons [OF pdivmod_rel y refl]) | 
| 29451 | 1681 | done | 
| 1682 | ||
| 52380 | 1683 | definition pdivmod :: "'a::field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<times> 'a poly" | 
| 1684 | where | |
| 1685 | "pdivmod p q = (p div q, p mod q)" | |
| 31663 | 1686 | |
| 52380 | 1687 | lemma div_poly_code [code]: | 
| 1688 | "p div q = fst (pdivmod p q)" | |
| 1689 | by (simp add: pdivmod_def) | |
| 31663 | 1690 | |
| 52380 | 1691 | lemma mod_poly_code [code]: | 
| 1692 | "p mod q = snd (pdivmod p q)" | |
| 1693 | by (simp add: pdivmod_def) | |
| 31663 | 1694 | |
| 52380 | 1695 | lemma pdivmod_0: | 
| 1696 | "pdivmod 0 q = (0, 0)" | |
| 1697 | by (simp add: pdivmod_def) | |
| 31663 | 1698 | |
| 52380 | 1699 | lemma pdivmod_pCons: | 
| 1700 | "pdivmod (pCons a p) q = | |
| 1701 | (if q = 0 then (0, pCons a p) else | |
| 1702 | (let (s, r) = pdivmod p q; | |
| 1703 | b = coeff (pCons a r) (degree q) / coeff q (degree q) | |
| 1704 | in (pCons b s, pCons a r - smult b q)))" | |
| 1705 | apply (simp add: pdivmod_def Let_def, safe) | |
| 1706 | apply (rule div_poly_eq) | |
| 1707 | apply (erule pdivmod_rel_pCons [OF pdivmod_rel _ refl]) | |
| 1708 | apply (rule mod_poly_eq) | |
| 1709 | apply (erule pdivmod_rel_pCons [OF pdivmod_rel _ refl]) | |
| 29451 | 1710 | done | 
| 1711 | ||
| 52380 | 1712 | lemma pdivmod_fold_coeffs [code]: | 
| 1713 | "pdivmod p q = (if q = 0 then (0, p) | |
| 1714 | else fold_coeffs (\<lambda>a (s, r). | |
| 1715 | let b = coeff (pCons a r) (degree q) / coeff q (degree q) | |
| 1716 | in (pCons b s, pCons a r - smult b q) | |
| 1717 | ) p (0, 0))" | |
| 1718 | apply (cases "q = 0") | |
| 1719 | apply (simp add: pdivmod_def) | |
| 1720 | apply (rule sym) | |
| 1721 | apply (induct p) | |
| 1722 | apply (simp_all add: pdivmod_0 pdivmod_pCons) | |
| 1723 | apply (case_tac "a = 0 \<and> p = 0") | |
| 1724 | apply (auto simp add: pdivmod_def) | |
| 1725 | done | |
| 29980 | 1726 | |
| 1727 | ||
| 60500 | 1728 | subsection \<open>Order of polynomial roots\<close> | 
| 29977 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1729 | |
| 52380 | 1730 | definition order :: "'a::idom \<Rightarrow> 'a poly \<Rightarrow> nat" | 
| 29977 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1731 | where | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1732 | "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 | 1733 | |
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1734 | lemma coeff_linear_power: | 
| 29979 | 1735 | fixes a :: "'a::comm_semiring_1" | 
| 29977 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1736 | shows "coeff ([:a, 1:] ^ n) n = 1" | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1737 | apply (induct n, simp_all) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1738 | apply (subst coeff_eq_0) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1739 | apply (auto intro: le_less_trans degree_power_le) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1740 | done | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1741 | |
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1742 | lemma degree_linear_power: | 
| 29979 | 1743 | fixes a :: "'a::comm_semiring_1" | 
| 29977 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1744 | shows "degree ([:a, 1:] ^ n) = n" | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1745 | apply (rule order_antisym) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1746 | 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 | 1747 | apply (rule le_degree, simp add: coeff_linear_power) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1748 | done | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1749 | |
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1750 | 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 | 1751 | apply (cases "p = 0", simp) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1752 | apply (cases "order a p", simp) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1753 | 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 | 1754 | apply (drule not_less_Least, simp) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1755 | apply (fold order_def, simp) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1756 | done | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1757 | |
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1758 | 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 | 1759 | unfolding order_def | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1760 | apply (rule LeastI_ex) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1761 | apply (rule_tac x="degree p" in exI) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1762 | apply (rule notI) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1763 | apply (drule (1) dvd_imp_degree_le) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1764 | apply (simp only: degree_linear_power) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1765 | done | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1766 | |
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1767 | lemma order: | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1768 | "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 | 1769 | by (rule conjI [OF order_1 order_2]) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1770 | |
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1771 | lemma order_degree: | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1772 | assumes p: "p \<noteq> 0" | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1773 | shows "order a p \<le> degree p" | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1774 | proof - | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1775 | 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 | 1776 | by (simp only: degree_linear_power) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1777 | also have "\<dots> \<le> degree p" | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1778 | 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 | 1779 | finally show ?thesis . | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1780 | qed | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1781 | |
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1782 | 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 | 1783 | apply (cases "p = 0", simp_all) | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1784 | apply (rule iffI) | 
| 56383 | 1785 | apply (metis order_2 not_gr0 poly_eq_0_iff_dvd power_0 power_Suc_0 power_one_right) | 
| 1786 | unfolding poly_eq_0_iff_dvd | |
| 1787 | apply (metis dvd_power dvd_trans order_1) | |
| 29977 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1788 | done | 
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1789 | |
| 
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
 huffman parents: 
29904diff
changeset | 1790 | |
| 60500 | 1791 | subsection \<open>GCD of polynomials\<close> | 
| 29478 | 1792 | |
| 52380 | 1793 | instantiation poly :: (field) gcd | 
| 29478 | 1794 | begin | 
| 1795 | ||
| 52380 | 1796 | function gcd_poly :: "'a::field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" | 
| 1797 | where | |
| 1798 | "gcd (x::'a poly) 0 = smult (inverse (coeff x (degree x))) x" | |
| 1799 | | "y \<noteq> 0 \<Longrightarrow> gcd (x::'a poly) y = gcd y (x mod y)" | |
| 1800 | by auto | |
| 29478 | 1801 | |
| 52380 | 1802 | termination "gcd :: _ poly \<Rightarrow> _" | 
| 1803 | by (relation "measure (\<lambda>(x, y). if y = 0 then 0 else Suc (degree y))") | |
| 1804 | (auto dest: degree_mod_less) | |
| 1805 | ||
| 1806 | declare gcd_poly.simps [simp del] | |
| 1807 | ||
| 58513 | 1808 | definition lcm_poly :: "'a::field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" | 
| 1809 | where | |
| 1810 | "lcm_poly a b = a * b div smult (coeff a (degree a) * coeff b (degree b)) (gcd a b)" | |
| 1811 | ||
| 52380 | 1812 | instance .. | 
| 29478 | 1813 | |
| 29451 | 1814 | end | 
| 29478 | 1815 | |
| 52380 | 1816 | lemma | 
| 1817 | fixes x y :: "_ poly" | |
| 1818 | shows poly_gcd_dvd1 [iff]: "gcd x y dvd x" | |
| 1819 | and poly_gcd_dvd2 [iff]: "gcd x y dvd y" | |
| 1820 | apply (induct x y rule: gcd_poly.induct) | |
| 1821 | apply (simp_all add: gcd_poly.simps) | |
| 1822 | apply (fastforce simp add: smult_dvd_iff dest: inverse_zero_imp_zero) | |
| 1823 | apply (blast dest: dvd_mod_imp_dvd) | |
| 1824 | done | |
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37770diff
changeset | 1825 | |
| 52380 | 1826 | lemma poly_gcd_greatest: | 
| 1827 | fixes k x y :: "_ poly" | |
| 1828 | shows "k dvd x \<Longrightarrow> k dvd y \<Longrightarrow> k dvd gcd x y" | |
| 1829 | by (induct x y rule: gcd_poly.induct) | |
| 1830 | (simp_all add: gcd_poly.simps dvd_mod dvd_smult) | |
| 29478 | 1831 | |
| 52380 | 1832 | lemma dvd_poly_gcd_iff [iff]: | 
| 1833 | fixes k x y :: "_ poly" | |
| 1834 | shows "k dvd gcd x y \<longleftrightarrow> k dvd x \<and> k dvd y" | |
| 60686 | 1835 | by (auto intro!: poly_gcd_greatest intro: dvd_trans [of _ "gcd x y"]) | 
| 29478 | 1836 | |
| 52380 | 1837 | lemma poly_gcd_monic: | 
| 1838 | fixes x y :: "_ poly" | |
| 1839 | shows "coeff (gcd x y) (degree (gcd x y)) = | |
| 1840 | (if x = 0 \<and> y = 0 then 0 else 1)" | |
| 1841 | by (induct x y rule: gcd_poly.induct) | |
| 1842 | (simp_all add: gcd_poly.simps nonzero_imp_inverse_nonzero) | |
| 29478 | 1843 | |
| 52380 | 1844 | lemma poly_gcd_zero_iff [simp]: | 
| 1845 | fixes x y :: "_ poly" | |
| 1846 | shows "gcd x y = 0 \<longleftrightarrow> x = 0 \<and> y = 0" | |
| 1847 | by (simp only: dvd_0_left_iff [symmetric] dvd_poly_gcd_iff) | |
| 29478 | 1848 | |
| 52380 | 1849 | lemma poly_gcd_0_0 [simp]: | 
| 1850 | "gcd (0::_ poly) 0 = 0" | |
| 1851 | by simp | |
| 29478 | 1852 | |
| 52380 | 1853 | lemma poly_dvd_antisym: | 
| 1854 | fixes p q :: "'a::idom poly" | |
| 1855 | assumes coeff: "coeff p (degree p) = coeff q (degree q)" | |
| 1856 | assumes dvd1: "p dvd q" and dvd2: "q dvd p" shows "p = q" | |
| 1857 | proof (cases "p = 0") | |
| 1858 | case True with coeff show "p = q" by simp | |
| 1859 | next | |
| 1860 | case False with coeff have "q \<noteq> 0" by auto | |
| 1861 | have degree: "degree p = degree q" | |
| 60500 | 1862 | using \<open>p dvd q\<close> \<open>q dvd p\<close> \<open>p \<noteq> 0\<close> \<open>q \<noteq> 0\<close> | 
| 52380 | 1863 | by (intro order_antisym dvd_imp_degree_le) | 
| 29478 | 1864 | |
| 60500 | 1865 | from \<open>p dvd q\<close> obtain a where a: "q = p * a" .. | 
| 1866 | with \<open>q \<noteq> 0\<close> have "a \<noteq> 0" by auto | |
| 1867 | with degree a \<open>p \<noteq> 0\<close> have "degree a = 0" | |
| 52380 | 1868 | by (simp add: degree_mult_eq) | 
| 1869 | with coeff a show "p = q" | |
| 1870 | by (cases a, auto split: if_splits) | |
| 1871 | qed | |
| 29478 | 1872 | |
| 52380 | 1873 | lemma poly_gcd_unique: | 
| 1874 | fixes d x y :: "_ poly" | |
| 1875 | assumes dvd1: "d dvd x" and dvd2: "d dvd y" | |
| 1876 | and greatest: "\<And>k. k dvd x \<Longrightarrow> k dvd y \<Longrightarrow> k dvd d" | |
| 1877 | and monic: "coeff d (degree d) = (if x = 0 \<and> y = 0 then 0 else 1)" | |
| 1878 | shows "gcd x y = d" | |
| 1879 | proof - | |
| 1880 | have "coeff (gcd x y) (degree (gcd x y)) = coeff d (degree d)" | |
| 1881 | by (simp_all add: poly_gcd_monic monic) | |
| 1882 | moreover have "gcd x y dvd d" | |
| 1883 | using poly_gcd_dvd1 poly_gcd_dvd2 by (rule greatest) | |
| 1884 | moreover have "d dvd gcd x y" | |
| 1885 | using dvd1 dvd2 by (rule poly_gcd_greatest) | |
| 1886 | ultimately show ?thesis | |
| 1887 | by (rule poly_dvd_antisym) | |
| 1888 | qed | |
| 29478 | 1889 | |
| 52380 | 1890 | interpretation gcd_poly!: abel_semigroup "gcd :: _ poly \<Rightarrow> _" | 
| 1891 | proof | |
| 1892 | fix x y z :: "'a poly" | |
| 1893 | show "gcd (gcd x y) z = gcd x (gcd y z)" | |
| 1894 | by (rule poly_gcd_unique) (auto intro: dvd_trans simp add: poly_gcd_monic) | |
| 1895 | show "gcd x y = gcd y x" | |
| 1896 | by (rule poly_gcd_unique) (simp_all add: poly_gcd_monic) | |
| 1897 | qed | |
| 29478 | 1898 | |
| 52380 | 1899 | lemmas poly_gcd_assoc = gcd_poly.assoc | 
| 1900 | lemmas poly_gcd_commute = gcd_poly.commute | |
| 1901 | lemmas poly_gcd_left_commute = gcd_poly.left_commute | |
| 29478 | 1902 | |
| 52380 | 1903 | lemmas poly_gcd_ac = poly_gcd_assoc poly_gcd_commute poly_gcd_left_commute | 
| 1904 | ||
| 1905 | lemma poly_gcd_1_left [simp]: "gcd 1 y = (1 :: _ poly)" | |
| 1906 | by (rule poly_gcd_unique) simp_all | |
| 29478 | 1907 | |
| 52380 | 1908 | lemma poly_gcd_1_right [simp]: "gcd x 1 = (1 :: _ poly)" | 
| 1909 | by (rule poly_gcd_unique) simp_all | |
| 1910 | ||
| 1911 | lemma poly_gcd_minus_left [simp]: "gcd (- x) y = gcd x (y :: _ poly)" | |
| 1912 | by (rule poly_gcd_unique) (simp_all add: poly_gcd_monic) | |
| 29478 | 1913 | |
| 52380 | 1914 | lemma poly_gcd_minus_right [simp]: "gcd x (- y) = gcd x (y :: _ poly)" | 
| 1915 | by (rule poly_gcd_unique) (simp_all add: poly_gcd_monic) | |
| 29478 | 1916 | |
| 52380 | 1917 | lemma poly_gcd_code [code]: | 
| 1918 | "gcd x y = (if y = 0 then smult (inverse (coeff x (degree x))) x else gcd y (x mod (y :: _ poly)))" | |
| 1919 | by (simp add: gcd_poly.simps) | |
| 1920 | ||
| 1921 | ||
| 60500 | 1922 | subsection \<open>Composition of polynomials\<close> | 
| 29478 | 1923 | |
| 52380 | 1924 | definition pcompose :: "'a::comm_semiring_0 poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" | 
| 1925 | where | |
| 1926 | "pcompose p q = fold_coeffs (\<lambda>a c. [:a:] + q * c) p 0" | |
| 1927 | ||
| 1928 | lemma pcompose_0 [simp]: | |
| 1929 | "pcompose 0 q = 0" | |
| 1930 | by (simp add: pcompose_def) | |
| 1931 | ||
| 1932 | lemma pcompose_pCons: | |
| 1933 | "pcompose (pCons a p) q = [:a:] + q * pcompose p q" | |
| 1934 | by (cases "p = 0 \<and> a = 0") (auto simp add: pcompose_def) | |
| 1935 | ||
| 1936 | lemma poly_pcompose: | |
| 1937 | "poly (pcompose p q) x = poly p (poly q x)" | |
| 1938 | by (induct p) (simp_all add: pcompose_pCons) | |
| 1939 | ||
| 1940 | lemma degree_pcompose_le: | |
| 1941 | "degree (pcompose p q) \<le> degree p * degree q" | |
| 1942 | apply (induct p, simp) | |
| 1943 | apply (simp add: pcompose_pCons, clarify) | |
| 1944 | apply (rule degree_add_le, simp) | |
| 1945 | apply (rule order_trans [OF degree_mult_le], simp) | |
| 29478 | 1946 | done | 
| 1947 | ||
| 52380 | 1948 | |
| 1949 | no_notation cCons (infixr "##" 65) | |
| 31663 | 1950 | |
| 29478 | 1951 | end |