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