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