author | huffman |
Thu, 15 Jan 2009 12:43:12 -0800 | |
changeset 29539 | abfe2af6883e |
parent 29537 | 50345a0f9df8 |
child 29540 | 8858d197a9b6 |
permissions | -rw-r--r-- |
29451 | 1 |
(* Title: HOL/Polynomial.thy |
2 |
Author: Brian Huffman |
|
3 |
Based on an earlier development by Clemens Ballarin |
|
4 |
*) |
|
5 |
||
6 |
header {* Univariate Polynomials *} |
|
7 |
||
8 |
theory Polynomial |
|
9 |
imports Plain SetInterval |
|
10 |
begin |
|
11 |
||
12 |
subsection {* Definition of type @{text poly} *} |
|
13 |
||
14 |
typedef (Poly) 'a poly = "{f::nat \<Rightarrow> 'a::zero. \<exists>n. \<forall>i>n. f i = 0}" |
|
15 |
morphisms coeff Abs_poly |
|
16 |
by auto |
|
17 |
||
18 |
lemma expand_poly_eq: "p = q \<longleftrightarrow> (\<forall>n. coeff p n = coeff q n)" |
|
19 |
by (simp add: coeff_inject [symmetric] expand_fun_eq) |
|
20 |
||
21 |
lemma poly_ext: "(\<And>n. coeff p n = coeff q n) \<Longrightarrow> p = q" |
|
22 |
by (simp add: expand_poly_eq) |
|
23 |
||
24 |
||
25 |
subsection {* Degree of a polynomial *} |
|
26 |
||
27 |
definition |
|
28 |
degree :: "'a::zero poly \<Rightarrow> nat" where |
|
29 |
"degree p = (LEAST n. \<forall>i>n. coeff p i = 0)" |
|
30 |
||
31 |
lemma coeff_eq_0: "degree p < n \<Longrightarrow> coeff p n = 0" |
|
32 |
proof - |
|
33 |
have "coeff p \<in> Poly" |
|
34 |
by (rule coeff) |
|
35 |
hence "\<exists>n. \<forall>i>n. coeff p i = 0" |
|
36 |
unfolding Poly_def by simp |
|
37 |
hence "\<forall>i>degree p. coeff p i = 0" |
|
38 |
unfolding degree_def by (rule LeastI_ex) |
|
39 |
moreover assume "degree p < n" |
|
40 |
ultimately show ?thesis by simp |
|
41 |
qed |
|
42 |
||
43 |
lemma le_degree: "coeff p n \<noteq> 0 \<Longrightarrow> n \<le> degree p" |
|
44 |
by (erule contrapos_np, rule coeff_eq_0, simp) |
|
45 |
||
46 |
lemma degree_le: "\<forall>i>n. coeff p i = 0 \<Longrightarrow> degree p \<le> n" |
|
47 |
unfolding degree_def by (erule Least_le) |
|
48 |
||
49 |
lemma less_degree_imp: "n < degree p \<Longrightarrow> \<exists>i>n. coeff p i \<noteq> 0" |
|
50 |
unfolding degree_def by (drule not_less_Least, simp) |
|
51 |
||
52 |
||
53 |
subsection {* The zero polynomial *} |
|
54 |
||
55 |
instantiation poly :: (zero) zero |
|
56 |
begin |
|
57 |
||
58 |
definition |
|
59 |
zero_poly_def: "0 = Abs_poly (\<lambda>n. 0)" |
|
60 |
||
61 |
instance .. |
|
62 |
end |
|
63 |
||
64 |
lemma coeff_0 [simp]: "coeff 0 n = 0" |
|
65 |
unfolding zero_poly_def |
|
66 |
by (simp add: Abs_poly_inverse Poly_def) |
|
67 |
||
68 |
lemma degree_0 [simp]: "degree 0 = 0" |
|
69 |
by (rule order_antisym [OF degree_le le0]) simp |
|
70 |
||
71 |
lemma leading_coeff_neq_0: |
|
72 |
assumes "p \<noteq> 0" shows "coeff p (degree p) \<noteq> 0" |
|
73 |
proof (cases "degree p") |
|
74 |
case 0 |
|
75 |
from `p \<noteq> 0` have "\<exists>n. coeff p n \<noteq> 0" |
|
76 |
by (simp add: expand_poly_eq) |
|
77 |
then obtain n where "coeff p n \<noteq> 0" .. |
|
78 |
hence "n \<le> degree p" by (rule le_degree) |
|
79 |
with `coeff p n \<noteq> 0` and `degree p = 0` |
|
80 |
show "coeff p (degree p) \<noteq> 0" by simp |
|
81 |
next |
|
82 |
case (Suc n) |
|
83 |
from `degree p = Suc n` have "n < degree p" by simp |
|
84 |
hence "\<exists>i>n. coeff p i \<noteq> 0" by (rule less_degree_imp) |
|
85 |
then obtain i where "n < i" and "coeff p i \<noteq> 0" by fast |
|
86 |
from `degree p = Suc n` and `n < i` have "degree p \<le> i" by simp |
|
87 |
also from `coeff p i \<noteq> 0` have "i \<le> degree p" by (rule le_degree) |
|
88 |
finally have "degree p = i" . |
|
89 |
with `coeff p i \<noteq> 0` show "coeff p (degree p) \<noteq> 0" by simp |
|
90 |
qed |
|
91 |
||
92 |
lemma leading_coeff_0_iff [simp]: "coeff p (degree p) = 0 \<longleftrightarrow> p = 0" |
|
93 |
by (cases "p = 0", simp, simp add: leading_coeff_neq_0) |
|
94 |
||
95 |
||
96 |
subsection {* List-style constructor for polynomials *} |
|
97 |
||
98 |
definition |
|
99 |
pCons :: "'a::zero \<Rightarrow> 'a poly \<Rightarrow> 'a poly" |
|
100 |
where |
|
101 |
[code del]: "pCons a p = Abs_poly (nat_case a (coeff p))" |
|
102 |
||
29455 | 103 |
syntax |
104 |
"_poly" :: "args \<Rightarrow> 'a poly" ("[:(_):]") |
|
105 |
||
106 |
translations |
|
107 |
"[:x, xs:]" == "CONST pCons x [:xs:]" |
|
108 |
"[:x:]" == "CONST pCons x 0" |
|
109 |
||
29451 | 110 |
lemma Poly_nat_case: "f \<in> Poly \<Longrightarrow> nat_case a f \<in> Poly" |
111 |
unfolding Poly_def by (auto split: nat.split) |
|
112 |
||
113 |
lemma coeff_pCons: |
|
114 |
"coeff (pCons a p) = nat_case a (coeff p)" |
|
115 |
unfolding pCons_def |
|
116 |
by (simp add: Abs_poly_inverse Poly_nat_case coeff) |
|
117 |
||
118 |
lemma coeff_pCons_0 [simp]: "coeff (pCons a p) 0 = a" |
|
119 |
by (simp add: coeff_pCons) |
|
120 |
||
121 |
lemma coeff_pCons_Suc [simp]: "coeff (pCons a p) (Suc n) = coeff p n" |
|
122 |
by (simp add: coeff_pCons) |
|
123 |
||
124 |
lemma degree_pCons_le: "degree (pCons a p) \<le> Suc (degree p)" |
|
125 |
by (rule degree_le, simp add: coeff_eq_0 coeff_pCons split: nat.split) |
|
126 |
||
127 |
lemma degree_pCons_eq: |
|
128 |
"p \<noteq> 0 \<Longrightarrow> degree (pCons a p) = Suc (degree p)" |
|
129 |
apply (rule order_antisym [OF degree_pCons_le]) |
|
130 |
apply (rule le_degree, simp) |
|
131 |
done |
|
132 |
||
133 |
lemma degree_pCons_0: "degree (pCons a 0) = 0" |
|
134 |
apply (rule order_antisym [OF _ le0]) |
|
135 |
apply (rule degree_le, simp add: coeff_pCons split: nat.split) |
|
136 |
done |
|
137 |
||
29460
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
138 |
lemma degree_pCons_eq_if [simp]: |
29451 | 139 |
"degree (pCons a p) = (if p = 0 then 0 else Suc (degree p))" |
140 |
apply (cases "p = 0", simp_all) |
|
141 |
apply (rule order_antisym [OF _ le0]) |
|
142 |
apply (rule degree_le, simp add: coeff_pCons split: nat.split) |
|
143 |
apply (rule order_antisym [OF degree_pCons_le]) |
|
144 |
apply (rule le_degree, simp) |
|
145 |
done |
|
146 |
||
147 |
lemma pCons_0_0 [simp]: "pCons 0 0 = 0" |
|
148 |
by (rule poly_ext, simp add: coeff_pCons split: nat.split) |
|
149 |
||
150 |
lemma pCons_eq_iff [simp]: |
|
151 |
"pCons a p = pCons b q \<longleftrightarrow> a = b \<and> p = q" |
|
152 |
proof (safe) |
|
153 |
assume "pCons a p = pCons b q" |
|
154 |
then have "coeff (pCons a p) 0 = coeff (pCons b q) 0" by simp |
|
155 |
then show "a = b" by simp |
|
156 |
next |
|
157 |
assume "pCons a p = pCons b q" |
|
158 |
then have "\<forall>n. coeff (pCons a p) (Suc n) = |
|
159 |
coeff (pCons b q) (Suc n)" by simp |
|
160 |
then show "p = q" by (simp add: expand_poly_eq) |
|
161 |
qed |
|
162 |
||
163 |
lemma pCons_eq_0_iff [simp]: "pCons a p = 0 \<longleftrightarrow> a = 0 \<and> p = 0" |
|
164 |
using pCons_eq_iff [of a p 0 0] by simp |
|
165 |
||
166 |
lemma Poly_Suc: "f \<in> Poly \<Longrightarrow> (\<lambda>n. f (Suc n)) \<in> Poly" |
|
167 |
unfolding Poly_def |
|
168 |
by (clarify, rule_tac x=n in exI, simp) |
|
169 |
||
170 |
lemma pCons_cases [cases type: poly]: |
|
171 |
obtains (pCons) a q where "p = pCons a q" |
|
172 |
proof |
|
173 |
show "p = pCons (coeff p 0) (Abs_poly (\<lambda>n. coeff p (Suc n)))" |
|
174 |
by (rule poly_ext) |
|
175 |
(simp add: Abs_poly_inverse Poly_Suc coeff coeff_pCons |
|
176 |
split: nat.split) |
|
177 |
qed |
|
178 |
||
179 |
lemma pCons_induct [case_names 0 pCons, induct type: poly]: |
|
180 |
assumes zero: "P 0" |
|
181 |
assumes pCons: "\<And>a p. P p \<Longrightarrow> P (pCons a p)" |
|
182 |
shows "P p" |
|
183 |
proof (induct p rule: measure_induct_rule [where f=degree]) |
|
184 |
case (less p) |
|
185 |
obtain a q where "p = pCons a q" by (rule pCons_cases) |
|
186 |
have "P q" |
|
187 |
proof (cases "q = 0") |
|
188 |
case True |
|
189 |
then show "P q" by (simp add: zero) |
|
190 |
next |
|
191 |
case False |
|
192 |
then have "degree (pCons a q) = Suc (degree q)" |
|
193 |
by (rule degree_pCons_eq) |
|
194 |
then have "degree q < degree p" |
|
195 |
using `p = pCons a q` by simp |
|
196 |
then show "P q" |
|
197 |
by (rule less.hyps) |
|
198 |
qed |
|
199 |
then have "P (pCons a q)" |
|
200 |
by (rule pCons) |
|
201 |
then show ?case |
|
202 |
using `p = pCons a q` by simp |
|
203 |
qed |
|
204 |
||
205 |
||
29454
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
206 |
subsection {* Recursion combinator for polynomials *} |
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
207 |
|
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
208 |
function |
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
209 |
poly_rec :: "'b \<Rightarrow> ('a::zero \<Rightarrow> 'a poly \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a poly \<Rightarrow> 'b" |
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
210 |
where |
29475 | 211 |
poly_rec_pCons_eq_if [simp del, code del]: |
29454
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
212 |
"poly_rec z f (pCons a p) = f a p (if p = 0 then z else poly_rec z f p)" |
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
213 |
by (case_tac x, rename_tac q, case_tac q, auto) |
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
214 |
|
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
215 |
termination poly_rec |
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
216 |
by (relation "measure (degree \<circ> snd \<circ> snd)", simp) |
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
217 |
(simp add: degree_pCons_eq) |
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
218 |
|
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
219 |
lemma poly_rec_0: |
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
220 |
"f 0 0 z = z \<Longrightarrow> poly_rec z f 0 = z" |
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
221 |
using poly_rec_pCons_eq_if [of z f 0 0] by simp |
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
222 |
|
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
223 |
lemma poly_rec_pCons: |
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
224 |
"f 0 0 z = z \<Longrightarrow> poly_rec z f (pCons a p) = f a p (poly_rec z f p)" |
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
225 |
by (simp add: poly_rec_pCons_eq_if poly_rec_0) |
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
226 |
|
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
227 |
|
29451 | 228 |
subsection {* Monomials *} |
229 |
||
230 |
definition |
|
231 |
monom :: "'a \<Rightarrow> nat \<Rightarrow> 'a::zero poly" where |
|
232 |
"monom a m = Abs_poly (\<lambda>n. if m = n then a else 0)" |
|
233 |
||
234 |
lemma coeff_monom [simp]: "coeff (monom a m) n = (if m=n then a else 0)" |
|
235 |
unfolding monom_def |
|
236 |
by (subst Abs_poly_inverse, auto simp add: Poly_def) |
|
237 |
||
238 |
lemma monom_0: "monom a 0 = pCons a 0" |
|
239 |
by (rule poly_ext, simp add: coeff_pCons split: nat.split) |
|
240 |
||
241 |
lemma monom_Suc: "monom a (Suc n) = pCons 0 (monom a n)" |
|
242 |
by (rule poly_ext, simp add: coeff_pCons split: nat.split) |
|
243 |
||
244 |
lemma monom_eq_0 [simp]: "monom 0 n = 0" |
|
245 |
by (rule poly_ext) simp |
|
246 |
||
247 |
lemma monom_eq_0_iff [simp]: "monom a n = 0 \<longleftrightarrow> a = 0" |
|
248 |
by (simp add: expand_poly_eq) |
|
249 |
||
250 |
lemma monom_eq_iff [simp]: "monom a n = monom b n \<longleftrightarrow> a = b" |
|
251 |
by (simp add: expand_poly_eq) |
|
252 |
||
253 |
lemma degree_monom_le: "degree (monom a n) \<le> n" |
|
254 |
by (rule degree_le, simp) |
|
255 |
||
256 |
lemma degree_monom_eq: "a \<noteq> 0 \<Longrightarrow> degree (monom a n) = n" |
|
257 |
apply (rule order_antisym [OF degree_monom_le]) |
|
258 |
apply (rule le_degree, simp) |
|
259 |
done |
|
260 |
||
261 |
||
262 |
subsection {* Addition and subtraction *} |
|
263 |
||
264 |
instantiation poly :: (comm_monoid_add) comm_monoid_add |
|
265 |
begin |
|
266 |
||
267 |
definition |
|
268 |
plus_poly_def [code del]: |
|
269 |
"p + q = Abs_poly (\<lambda>n. coeff p n + coeff q n)" |
|
270 |
||
271 |
lemma Poly_add: |
|
272 |
fixes f g :: "nat \<Rightarrow> 'a" |
|
273 |
shows "\<lbrakk>f \<in> Poly; g \<in> Poly\<rbrakk> \<Longrightarrow> (\<lambda>n. f n + g n) \<in> Poly" |
|
274 |
unfolding Poly_def |
|
275 |
apply (clarify, rename_tac m n) |
|
276 |
apply (rule_tac x="max m n" in exI, simp) |
|
277 |
done |
|
278 |
||
279 |
lemma coeff_add [simp]: |
|
280 |
"coeff (p + q) n = coeff p n + coeff q n" |
|
281 |
unfolding plus_poly_def |
|
282 |
by (simp add: Abs_poly_inverse coeff Poly_add) |
|
283 |
||
284 |
instance proof |
|
285 |
fix p q r :: "'a poly" |
|
286 |
show "(p + q) + r = p + (q + r)" |
|
287 |
by (simp add: expand_poly_eq add_assoc) |
|
288 |
show "p + q = q + p" |
|
289 |
by (simp add: expand_poly_eq add_commute) |
|
290 |
show "0 + p = p" |
|
291 |
by (simp add: expand_poly_eq) |
|
292 |
qed |
|
293 |
||
294 |
end |
|
295 |
||
296 |
instantiation poly :: (ab_group_add) ab_group_add |
|
297 |
begin |
|
298 |
||
299 |
definition |
|
300 |
uminus_poly_def [code del]: |
|
301 |
"- p = Abs_poly (\<lambda>n. - coeff p n)" |
|
302 |
||
303 |
definition |
|
304 |
minus_poly_def [code del]: |
|
305 |
"p - q = Abs_poly (\<lambda>n. coeff p n - coeff q n)" |
|
306 |
||
307 |
lemma Poly_minus: |
|
308 |
fixes f :: "nat \<Rightarrow> 'a" |
|
309 |
shows "f \<in> Poly \<Longrightarrow> (\<lambda>n. - f n) \<in> Poly" |
|
310 |
unfolding Poly_def by simp |
|
311 |
||
312 |
lemma Poly_diff: |
|
313 |
fixes f g :: "nat \<Rightarrow> 'a" |
|
314 |
shows "\<lbrakk>f \<in> Poly; g \<in> Poly\<rbrakk> \<Longrightarrow> (\<lambda>n. f n - g n) \<in> Poly" |
|
315 |
unfolding diff_minus by (simp add: Poly_add Poly_minus) |
|
316 |
||
317 |
lemma coeff_minus [simp]: "coeff (- p) n = - coeff p n" |
|
318 |
unfolding uminus_poly_def |
|
319 |
by (simp add: Abs_poly_inverse coeff Poly_minus) |
|
320 |
||
321 |
lemma coeff_diff [simp]: |
|
322 |
"coeff (p - q) n = coeff p n - coeff q n" |
|
323 |
unfolding minus_poly_def |
|
324 |
by (simp add: Abs_poly_inverse coeff Poly_diff) |
|
325 |
||
326 |
instance proof |
|
327 |
fix p q :: "'a poly" |
|
328 |
show "- p + p = 0" |
|
329 |
by (simp add: expand_poly_eq) |
|
330 |
show "p - q = p + - q" |
|
331 |
by (simp add: expand_poly_eq diff_minus) |
|
332 |
qed |
|
333 |
||
334 |
end |
|
335 |
||
336 |
lemma add_pCons [simp]: |
|
337 |
"pCons a p + pCons b q = pCons (a + b) (p + q)" |
|
338 |
by (rule poly_ext, simp add: coeff_pCons split: nat.split) |
|
339 |
||
340 |
lemma minus_pCons [simp]: |
|
341 |
"- pCons a p = pCons (- a) (- p)" |
|
342 |
by (rule poly_ext, simp add: coeff_pCons split: nat.split) |
|
343 |
||
344 |
lemma diff_pCons [simp]: |
|
345 |
"pCons a p - pCons b q = pCons (a - b) (p - q)" |
|
346 |
by (rule poly_ext, simp add: coeff_pCons split: nat.split) |
|
347 |
||
29539 | 348 |
lemma degree_add_le_max: "degree (p + q) \<le> max (degree p) (degree q)" |
29451 | 349 |
by (rule degree_le, auto simp add: coeff_eq_0) |
350 |
||
29539 | 351 |
lemma degree_add_le: |
352 |
"\<lbrakk>degree p \<le> n; degree q \<le> n\<rbrakk> \<Longrightarrow> degree (p + q) \<le> n" |
|
353 |
by (auto intro: order_trans degree_add_le_max) |
|
354 |
||
29453 | 355 |
lemma degree_add_less: |
356 |
"\<lbrakk>degree p < n; degree q < n\<rbrakk> \<Longrightarrow> degree (p + q) < n" |
|
29539 | 357 |
by (auto intro: le_less_trans degree_add_le_max) |
29453 | 358 |
|
29451 | 359 |
lemma degree_add_eq_right: |
360 |
"degree p < degree q \<Longrightarrow> degree (p + q) = degree q" |
|
361 |
apply (cases "q = 0", simp) |
|
362 |
apply (rule order_antisym) |
|
29539 | 363 |
apply (simp add: degree_add_le) |
29451 | 364 |
apply (rule le_degree) |
365 |
apply (simp add: coeff_eq_0) |
|
366 |
done |
|
367 |
||
368 |
lemma degree_add_eq_left: |
|
369 |
"degree q < degree p \<Longrightarrow> degree (p + q) = degree p" |
|
370 |
using degree_add_eq_right [of q p] |
|
371 |
by (simp add: add_commute) |
|
372 |
||
373 |
lemma degree_minus [simp]: "degree (- p) = degree p" |
|
374 |
unfolding degree_def by simp |
|
375 |
||
29539 | 376 |
lemma degree_diff_le_max: "degree (p - q) \<le> max (degree p) (degree q)" |
29451 | 377 |
using degree_add_le [where p=p and q="-q"] |
378 |
by (simp add: diff_minus) |
|
379 |
||
29539 | 380 |
lemma degree_diff_le: |
381 |
"\<lbrakk>degree p \<le> n; degree q \<le> n\<rbrakk> \<Longrightarrow> degree (p - q) \<le> n" |
|
382 |
by (simp add: diff_minus degree_add_le) |
|
383 |
||
29453 | 384 |
lemma degree_diff_less: |
385 |
"\<lbrakk>degree p < n; degree q < n\<rbrakk> \<Longrightarrow> degree (p - q) < n" |
|
29539 | 386 |
by (simp add: diff_minus degree_add_less) |
29453 | 387 |
|
29451 | 388 |
lemma add_monom: "monom a n + monom b n = monom (a + b) n" |
389 |
by (rule poly_ext) simp |
|
390 |
||
391 |
lemma diff_monom: "monom a n - monom b n = monom (a - b) n" |
|
392 |
by (rule poly_ext) simp |
|
393 |
||
394 |
lemma minus_monom: "- monom a n = monom (-a) n" |
|
395 |
by (rule poly_ext) simp |
|
396 |
||
397 |
lemma coeff_setsum: "coeff (\<Sum>x\<in>A. p x) i = (\<Sum>x\<in>A. coeff (p x) i)" |
|
398 |
by (cases "finite A", induct set: finite, simp_all) |
|
399 |
||
400 |
lemma monom_setsum: "monom (\<Sum>x\<in>A. a x) n = (\<Sum>x\<in>A. monom (a x) n)" |
|
401 |
by (rule poly_ext) (simp add: coeff_setsum) |
|
402 |
||
403 |
||
404 |
subsection {* Multiplication by a constant *} |
|
405 |
||
406 |
definition |
|
407 |
smult :: "'a::comm_semiring_0 \<Rightarrow> 'a poly \<Rightarrow> 'a poly" where |
|
408 |
"smult a p = Abs_poly (\<lambda>n. a * coeff p n)" |
|
409 |
||
410 |
lemma Poly_smult: |
|
411 |
fixes f :: "nat \<Rightarrow> 'a::comm_semiring_0" |
|
412 |
shows "f \<in> Poly \<Longrightarrow> (\<lambda>n. a * f n) \<in> Poly" |
|
413 |
unfolding Poly_def |
|
414 |
by (clarify, rule_tac x=n in exI, simp) |
|
415 |
||
416 |
lemma coeff_smult [simp]: "coeff (smult a p) n = a * coeff p n" |
|
417 |
unfolding smult_def |
|
418 |
by (simp add: Abs_poly_inverse Poly_smult coeff) |
|
419 |
||
420 |
lemma degree_smult_le: "degree (smult a p) \<le> degree p" |
|
421 |
by (rule degree_le, simp add: coeff_eq_0) |
|
422 |
||
29472 | 423 |
lemma smult_smult [simp]: "smult a (smult b p) = smult (a * b) p" |
29451 | 424 |
by (rule poly_ext, simp add: mult_assoc) |
425 |
||
426 |
lemma smult_0_right [simp]: "smult a 0 = 0" |
|
427 |
by (rule poly_ext, simp) |
|
428 |
||
429 |
lemma smult_0_left [simp]: "smult 0 p = 0" |
|
430 |
by (rule poly_ext, simp) |
|
431 |
||
432 |
lemma smult_1_left [simp]: "smult (1::'a::comm_semiring_1) p = p" |
|
433 |
by (rule poly_ext, simp) |
|
434 |
||
435 |
lemma smult_add_right: |
|
436 |
"smult a (p + q) = smult a p + smult a q" |
|
437 |
by (rule poly_ext, simp add: ring_simps) |
|
438 |
||
439 |
lemma smult_add_left: |
|
440 |
"smult (a + b) p = smult a p + smult b p" |
|
441 |
by (rule poly_ext, simp add: ring_simps) |
|
442 |
||
29457 | 443 |
lemma smult_minus_right [simp]: |
29451 | 444 |
"smult (a::'a::comm_ring) (- p) = - smult a p" |
445 |
by (rule poly_ext, simp) |
|
446 |
||
29457 | 447 |
lemma smult_minus_left [simp]: |
29451 | 448 |
"smult (- a::'a::comm_ring) p = - smult a p" |
449 |
by (rule poly_ext, simp) |
|
450 |
||
451 |
lemma smult_diff_right: |
|
452 |
"smult (a::'a::comm_ring) (p - q) = smult a p - smult a q" |
|
453 |
by (rule poly_ext, simp add: ring_simps) |
|
454 |
||
455 |
lemma smult_diff_left: |
|
456 |
"smult (a - b::'a::comm_ring) p = smult a p - smult b p" |
|
457 |
by (rule poly_ext, simp add: ring_simps) |
|
458 |
||
29472 | 459 |
lemmas smult_distribs = |
460 |
smult_add_left smult_add_right |
|
461 |
smult_diff_left smult_diff_right |
|
462 |
||
29451 | 463 |
lemma smult_pCons [simp]: |
464 |
"smult a (pCons b p) = pCons (a * b) (smult a p)" |
|
465 |
by (rule poly_ext, simp add: coeff_pCons split: nat.split) |
|
466 |
||
467 |
lemma smult_monom: "smult a (monom b n) = monom (a * b) n" |
|
468 |
by (induct n, simp add: monom_0, simp add: monom_Suc) |
|
469 |
||
470 |
||
471 |
subsection {* Multiplication of polynomials *} |
|
472 |
||
29474 | 473 |
text {* TODO: move to SetInterval.thy *} |
29451 | 474 |
lemma setsum_atMost_Suc_shift: |
475 |
fixes f :: "nat \<Rightarrow> 'a::comm_monoid_add" |
|
476 |
shows "(\<Sum>i\<le>Suc n. f i) = f 0 + (\<Sum>i\<le>n. f (Suc i))" |
|
477 |
proof (induct n) |
|
478 |
case 0 show ?case by simp |
|
479 |
next |
|
480 |
case (Suc n) note IH = this |
|
481 |
have "(\<Sum>i\<le>Suc (Suc n). f i) = (\<Sum>i\<le>Suc n. f i) + f (Suc (Suc n))" |
|
482 |
by (rule setsum_atMost_Suc) |
|
483 |
also have "(\<Sum>i\<le>Suc n. f i) = f 0 + (\<Sum>i\<le>n. f (Suc i))" |
|
484 |
by (rule IH) |
|
485 |
also have "f 0 + (\<Sum>i\<le>n. f (Suc i)) + f (Suc (Suc n)) = |
|
486 |
f 0 + ((\<Sum>i\<le>n. f (Suc i)) + f (Suc (Suc n)))" |
|
487 |
by (rule add_assoc) |
|
488 |
also have "(\<Sum>i\<le>n. f (Suc i)) + f (Suc (Suc n)) = (\<Sum>i\<le>Suc n. f (Suc i))" |
|
489 |
by (rule setsum_atMost_Suc [symmetric]) |
|
490 |
finally show ?case . |
|
491 |
qed |
|
492 |
||
493 |
instantiation poly :: (comm_semiring_0) comm_semiring_0 |
|
494 |
begin |
|
495 |
||
496 |
definition |
|
29475 | 497 |
times_poly_def [code del]: |
29474 | 498 |
"p * q = poly_rec 0 (\<lambda>a p pq. smult a q + pCons 0 pq) p" |
499 |
||
500 |
lemma mult_poly_0_left: "(0::'a poly) * q = 0" |
|
501 |
unfolding times_poly_def by (simp add: poly_rec_0) |
|
502 |
||
503 |
lemma mult_pCons_left [simp]: |
|
504 |
"pCons a p * q = smult a q + pCons 0 (p * q)" |
|
505 |
unfolding times_poly_def by (simp add: poly_rec_pCons) |
|
506 |
||
507 |
lemma mult_poly_0_right: "p * (0::'a poly) = 0" |
|
508 |
by (induct p, simp add: mult_poly_0_left, simp) |
|
29451 | 509 |
|
29474 | 510 |
lemma mult_pCons_right [simp]: |
511 |
"p * pCons a q = smult a p + pCons 0 (p * q)" |
|
512 |
by (induct p, simp add: mult_poly_0_left, simp add: ring_simps) |
|
513 |
||
514 |
lemmas mult_poly_0 = mult_poly_0_left mult_poly_0_right |
|
515 |
||
516 |
lemma mult_smult_left [simp]: "smult a p * q = smult a (p * q)" |
|
517 |
by (induct p, simp add: mult_poly_0, simp add: smult_add_right) |
|
518 |
||
519 |
lemma mult_smult_right [simp]: "p * smult a q = smult a (p * q)" |
|
520 |
by (induct q, simp add: mult_poly_0, simp add: smult_add_right) |
|
521 |
||
522 |
lemma mult_poly_add_left: |
|
523 |
fixes p q r :: "'a poly" |
|
524 |
shows "(p + q) * r = p * r + q * r" |
|
525 |
by (induct r, simp add: mult_poly_0, |
|
526 |
simp add: smult_distribs group_simps) |
|
29451 | 527 |
|
528 |
instance proof |
|
529 |
fix p q r :: "'a poly" |
|
530 |
show 0: "0 * p = 0" |
|
29474 | 531 |
by (rule mult_poly_0_left) |
29451 | 532 |
show "p * 0 = 0" |
29474 | 533 |
by (rule mult_poly_0_right) |
29451 | 534 |
show "(p + q) * r = p * r + q * r" |
29474 | 535 |
by (rule mult_poly_add_left) |
29451 | 536 |
show "(p * q) * r = p * (q * r)" |
29474 | 537 |
by (induct p, simp add: mult_poly_0, simp add: mult_poly_add_left) |
29451 | 538 |
show "p * q = q * p" |
29474 | 539 |
by (induct p, simp add: mult_poly_0, simp) |
29451 | 540 |
qed |
541 |
||
542 |
end |
|
543 |
||
29474 | 544 |
lemma coeff_mult: |
545 |
"coeff (p * q) n = (\<Sum>i\<le>n. coeff p i * coeff q (n-i))" |
|
546 |
proof (induct p arbitrary: n) |
|
547 |
case 0 show ?case by simp |
|
548 |
next |
|
549 |
case (pCons a p n) thus ?case |
|
550 |
by (cases n, simp, simp add: setsum_atMost_Suc_shift |
|
551 |
del: setsum_atMost_Suc) |
|
552 |
qed |
|
29451 | 553 |
|
29474 | 554 |
lemma degree_mult_le: "degree (p * q) \<le> degree p + degree q" |
555 |
apply (rule degree_le) |
|
556 |
apply (induct p) |
|
557 |
apply simp |
|
558 |
apply (simp add: coeff_eq_0 coeff_pCons split: nat.split) |
|
29451 | 559 |
done |
560 |
||
561 |
lemma mult_monom: "monom a m * monom b n = monom (a * b) (m + n)" |
|
562 |
by (induct m, simp add: monom_0 smult_monom, simp add: monom_Suc) |
|
563 |
||
564 |
||
565 |
subsection {* The unit polynomial and exponentiation *} |
|
566 |
||
567 |
instantiation poly :: (comm_semiring_1) comm_semiring_1 |
|
568 |
begin |
|
569 |
||
570 |
definition |
|
571 |
one_poly_def: |
|
572 |
"1 = pCons 1 0" |
|
573 |
||
574 |
instance proof |
|
575 |
fix p :: "'a poly" show "1 * p = p" |
|
576 |
unfolding one_poly_def |
|
577 |
by simp |
|
578 |
next |
|
579 |
show "0 \<noteq> (1::'a poly)" |
|
580 |
unfolding one_poly_def by simp |
|
581 |
qed |
|
582 |
||
583 |
end |
|
584 |
||
585 |
lemma coeff_1 [simp]: "coeff 1 n = (if n = 0 then 1 else 0)" |
|
586 |
unfolding one_poly_def |
|
587 |
by (simp add: coeff_pCons split: nat.split) |
|
588 |
||
589 |
lemma degree_1 [simp]: "degree 1 = 0" |
|
590 |
unfolding one_poly_def |
|
591 |
by (rule degree_pCons_0) |
|
592 |
||
593 |
instantiation poly :: (comm_semiring_1) recpower |
|
594 |
begin |
|
595 |
||
596 |
primrec power_poly where |
|
597 |
power_poly_0: "(p::'a poly) ^ 0 = 1" |
|
598 |
| power_poly_Suc: "(p::'a poly) ^ (Suc n) = p * p ^ n" |
|
599 |
||
600 |
instance |
|
601 |
by default simp_all |
|
602 |
||
603 |
end |
|
604 |
||
605 |
instance poly :: (comm_ring) comm_ring .. |
|
606 |
||
607 |
instance poly :: (comm_ring_1) comm_ring_1 .. |
|
608 |
||
609 |
instantiation poly :: (comm_ring_1) number_ring |
|
610 |
begin |
|
611 |
||
612 |
definition |
|
613 |
"number_of k = (of_int k :: 'a poly)" |
|
614 |
||
615 |
instance |
|
616 |
by default (rule number_of_poly_def) |
|
617 |
||
618 |
end |
|
619 |
||
620 |
||
621 |
subsection {* Polynomials form an integral domain *} |
|
622 |
||
623 |
lemma coeff_mult_degree_sum: |
|
624 |
"coeff (p * q) (degree p + degree q) = |
|
625 |
coeff p (degree p) * coeff q (degree q)" |
|
29471 | 626 |
by (induct p, simp, simp add: coeff_eq_0) |
29451 | 627 |
|
628 |
instance poly :: (idom) idom |
|
629 |
proof |
|
630 |
fix p q :: "'a poly" |
|
631 |
assume "p \<noteq> 0" and "q \<noteq> 0" |
|
632 |
have "coeff (p * q) (degree p + degree q) = |
|
633 |
coeff p (degree p) * coeff q (degree q)" |
|
634 |
by (rule coeff_mult_degree_sum) |
|
635 |
also have "coeff p (degree p) * coeff q (degree q) \<noteq> 0" |
|
636 |
using `p \<noteq> 0` and `q \<noteq> 0` by simp |
|
637 |
finally have "\<exists>n. coeff (p * q) n \<noteq> 0" .. |
|
638 |
thus "p * q \<noteq> 0" by (simp add: expand_poly_eq) |
|
639 |
qed |
|
640 |
||
641 |
lemma degree_mult_eq: |
|
642 |
fixes p q :: "'a::idom poly" |
|
643 |
shows "\<lbrakk>p \<noteq> 0; q \<noteq> 0\<rbrakk> \<Longrightarrow> degree (p * q) = degree p + degree q" |
|
644 |
apply (rule order_antisym [OF degree_mult_le le_degree]) |
|
645 |
apply (simp add: coeff_mult_degree_sum) |
|
646 |
done |
|
647 |
||
648 |
lemma dvd_imp_degree_le: |
|
649 |
fixes p q :: "'a::idom poly" |
|
650 |
shows "\<lbrakk>p dvd q; q \<noteq> 0\<rbrakk> \<Longrightarrow> degree p \<le> degree q" |
|
651 |
by (erule dvdE, simp add: degree_mult_eq) |
|
652 |
||
653 |
||
654 |
subsection {* Long division of polynomials *} |
|
655 |
||
656 |
definition |
|
29537 | 657 |
pdivmod_rel :: "'a::field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<Rightarrow> bool" |
29451 | 658 |
where |
29475 | 659 |
[code del]: |
29537 | 660 |
"pdivmod_rel x y q r \<longleftrightarrow> |
29451 | 661 |
x = q * y + r \<and> (if y = 0 then q = 0 else r = 0 \<or> degree r < degree y)" |
662 |
||
29537 | 663 |
lemma pdivmod_rel_0: |
664 |
"pdivmod_rel 0 y 0 0" |
|
665 |
unfolding pdivmod_rel_def by simp |
|
29451 | 666 |
|
29537 | 667 |
lemma pdivmod_rel_by_0: |
668 |
"pdivmod_rel x 0 0 x" |
|
669 |
unfolding pdivmod_rel_def by simp |
|
29451 | 670 |
|
671 |
lemma eq_zero_or_degree_less: |
|
672 |
assumes "degree p \<le> n" and "coeff p n = 0" |
|
673 |
shows "p = 0 \<or> degree p < n" |
|
674 |
proof (cases n) |
|
675 |
case 0 |
|
676 |
with `degree p \<le> n` and `coeff p n = 0` |
|
677 |
have "coeff p (degree p) = 0" by simp |
|
678 |
then have "p = 0" by simp |
|
679 |
then show ?thesis .. |
|
680 |
next |
|
681 |
case (Suc m) |
|
682 |
have "\<forall>i>n. coeff p i = 0" |
|
683 |
using `degree p \<le> n` by (simp add: coeff_eq_0) |
|
684 |
then have "\<forall>i\<ge>n. coeff p i = 0" |
|
685 |
using `coeff p n = 0` by (simp add: le_less) |
|
686 |
then have "\<forall>i>m. coeff p i = 0" |
|
687 |
using `n = Suc m` by (simp add: less_eq_Suc_le) |
|
688 |
then have "degree p \<le> m" |
|
689 |
by (rule degree_le) |
|
690 |
then have "degree p < n" |
|
691 |
using `n = Suc m` by (simp add: less_Suc_eq_le) |
|
692 |
then show ?thesis .. |
|
693 |
qed |
|
694 |
||
29537 | 695 |
lemma pdivmod_rel_pCons: |
696 |
assumes rel: "pdivmod_rel x y q r" |
|
29451 | 697 |
assumes y: "y \<noteq> 0" |
698 |
assumes b: "b = coeff (pCons a r) (degree y) / coeff y (degree y)" |
|
29537 | 699 |
shows "pdivmod_rel (pCons a x) y (pCons b q) (pCons a r - smult b y)" |
700 |
(is "pdivmod_rel ?x y ?q ?r") |
|
29451 | 701 |
proof - |
702 |
have x: "x = q * y + r" and r: "r = 0 \<or> degree r < degree y" |
|
29537 | 703 |
using assms unfolding pdivmod_rel_def by simp_all |
29451 | 704 |
|
705 |
have 1: "?x = ?q * y + ?r" |
|
706 |
using b x by simp |
|
707 |
||
708 |
have 2: "?r = 0 \<or> degree ?r < degree y" |
|
709 |
proof (rule eq_zero_or_degree_less) |
|
29539 | 710 |
show "degree ?r \<le> degree y" |
711 |
proof (rule degree_diff_le) |
|
29451 | 712 |
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
|
713 |
using r by auto |
29451 | 714 |
show "degree (smult b y) \<le> degree y" |
715 |
by (rule degree_smult_le) |
|
716 |
qed |
|
717 |
next |
|
718 |
show "coeff ?r (degree y) = 0" |
|
719 |
using `y \<noteq> 0` unfolding b by simp |
|
720 |
qed |
|
721 |
||
722 |
from 1 2 show ?thesis |
|
29537 | 723 |
unfolding pdivmod_rel_def |
29451 | 724 |
using `y \<noteq> 0` by simp |
725 |
qed |
|
726 |
||
29537 | 727 |
lemma pdivmod_rel_exists: "\<exists>q r. pdivmod_rel x y q r" |
29451 | 728 |
apply (cases "y = 0") |
29537 | 729 |
apply (fast intro!: pdivmod_rel_by_0) |
29451 | 730 |
apply (induct x) |
29537 | 731 |
apply (fast intro!: pdivmod_rel_0) |
732 |
apply (fast intro!: pdivmod_rel_pCons) |
|
29451 | 733 |
done |
734 |
||
29537 | 735 |
lemma pdivmod_rel_unique: |
736 |
assumes 1: "pdivmod_rel x y q1 r1" |
|
737 |
assumes 2: "pdivmod_rel x y q2 r2" |
|
29451 | 738 |
shows "q1 = q2 \<and> r1 = r2" |
739 |
proof (cases "y = 0") |
|
740 |
assume "y = 0" with assms show ?thesis |
|
29537 | 741 |
by (simp add: pdivmod_rel_def) |
29451 | 742 |
next |
743 |
assume [simp]: "y \<noteq> 0" |
|
744 |
from 1 have q1: "x = q1 * y + r1" and r1: "r1 = 0 \<or> degree r1 < degree y" |
|
29537 | 745 |
unfolding pdivmod_rel_def by simp_all |
29451 | 746 |
from 2 have q2: "x = q2 * y + r2" and r2: "r2 = 0 \<or> degree r2 < degree y" |
29537 | 747 |
unfolding pdivmod_rel_def by simp_all |
29451 | 748 |
from q1 q2 have q3: "(q1 - q2) * y = r2 - r1" |
749 |
by (simp add: ring_simps) |
|
750 |
from r1 r2 have r3: "(r2 - r1) = 0 \<or> degree (r2 - r1) < degree y" |
|
29453 | 751 |
by (auto intro: degree_diff_less) |
29451 | 752 |
|
753 |
show "q1 = q2 \<and> r1 = r2" |
|
754 |
proof (rule ccontr) |
|
755 |
assume "\<not> (q1 = q2 \<and> r1 = r2)" |
|
756 |
with q3 have "q1 \<noteq> q2" and "r1 \<noteq> r2" by auto |
|
757 |
with r3 have "degree (r2 - r1) < degree y" by simp |
|
758 |
also have "degree y \<le> degree (q1 - q2) + degree y" by simp |
|
759 |
also have "\<dots> = degree ((q1 - q2) * y)" |
|
760 |
using `q1 \<noteq> q2` by (simp add: degree_mult_eq) |
|
761 |
also have "\<dots> = degree (r2 - r1)" |
|
762 |
using q3 by simp |
|
763 |
finally have "degree (r2 - r1) < degree (r2 - r1)" . |
|
764 |
then show "False" by simp |
|
765 |
qed |
|
766 |
qed |
|
767 |
||
29537 | 768 |
lemmas pdivmod_rel_unique_div = |
769 |
pdivmod_rel_unique [THEN conjunct1, standard] |
|
29451 | 770 |
|
29537 | 771 |
lemmas pdivmod_rel_unique_mod = |
772 |
pdivmod_rel_unique [THEN conjunct2, standard] |
|
29451 | 773 |
|
774 |
instantiation poly :: (field) ring_div |
|
775 |
begin |
|
776 |
||
777 |
definition div_poly where |
|
29537 | 778 |
[code del]: "x div y = (THE q. \<exists>r. pdivmod_rel x y q r)" |
29451 | 779 |
|
780 |
definition mod_poly where |
|
29537 | 781 |
[code del]: "x mod y = (THE r. \<exists>q. pdivmod_rel x y q r)" |
29451 | 782 |
|
783 |
lemma div_poly_eq: |
|
29537 | 784 |
"pdivmod_rel x y q r \<Longrightarrow> x div y = q" |
29451 | 785 |
unfolding div_poly_def |
29537 | 786 |
by (fast elim: pdivmod_rel_unique_div) |
29451 | 787 |
|
788 |
lemma mod_poly_eq: |
|
29537 | 789 |
"pdivmod_rel x y q r \<Longrightarrow> x mod y = r" |
29451 | 790 |
unfolding mod_poly_def |
29537 | 791 |
by (fast elim: pdivmod_rel_unique_mod) |
29451 | 792 |
|
29537 | 793 |
lemma pdivmod_rel: |
794 |
"pdivmod_rel x y (x div y) (x mod y)" |
|
29451 | 795 |
proof - |
29537 | 796 |
from pdivmod_rel_exists |
797 |
obtain q r where "pdivmod_rel x y q r" by fast |
|
29451 | 798 |
thus ?thesis |
799 |
by (simp add: div_poly_eq mod_poly_eq) |
|
800 |
qed |
|
801 |
||
802 |
instance proof |
|
803 |
fix x y :: "'a poly" |
|
804 |
show "x div y * y + x mod y = x" |
|
29537 | 805 |
using pdivmod_rel [of x y] |
806 |
by (simp add: pdivmod_rel_def) |
|
29451 | 807 |
next |
808 |
fix x :: "'a poly" |
|
29537 | 809 |
have "pdivmod_rel x 0 0 x" |
810 |
by (rule pdivmod_rel_by_0) |
|
29451 | 811 |
thus "x div 0 = 0" |
812 |
by (rule div_poly_eq) |
|
813 |
next |
|
814 |
fix y :: "'a poly" |
|
29537 | 815 |
have "pdivmod_rel 0 y 0 0" |
816 |
by (rule pdivmod_rel_0) |
|
29451 | 817 |
thus "0 div y = 0" |
818 |
by (rule div_poly_eq) |
|
819 |
next |
|
820 |
fix x y z :: "'a poly" |
|
821 |
assume "y \<noteq> 0" |
|
29537 | 822 |
hence "pdivmod_rel (x + z * y) y (z + x div y) (x mod y)" |
823 |
using pdivmod_rel [of x y] |
|
824 |
by (simp add: pdivmod_rel_def left_distrib) |
|
29451 | 825 |
thus "(x + z * y) div y = z + x div y" |
826 |
by (rule div_poly_eq) |
|
827 |
qed |
|
828 |
||
829 |
end |
|
830 |
||
831 |
lemma degree_mod_less: |
|
832 |
"y \<noteq> 0 \<Longrightarrow> x mod y = 0 \<or> degree (x mod y) < degree y" |
|
29537 | 833 |
using pdivmod_rel [of x y] |
834 |
unfolding pdivmod_rel_def by simp |
|
29451 | 835 |
|
836 |
lemma div_poly_less: "degree x < degree y \<Longrightarrow> x div y = 0" |
|
837 |
proof - |
|
838 |
assume "degree x < degree y" |
|
29537 | 839 |
hence "pdivmod_rel x y 0 x" |
840 |
by (simp add: pdivmod_rel_def) |
|
29451 | 841 |
thus "x div y = 0" by (rule div_poly_eq) |
842 |
qed |
|
843 |
||
844 |
lemma mod_poly_less: "degree x < degree y \<Longrightarrow> x mod y = x" |
|
845 |
proof - |
|
846 |
assume "degree x < degree y" |
|
29537 | 847 |
hence "pdivmod_rel x y 0 x" |
848 |
by (simp add: pdivmod_rel_def) |
|
29451 | 849 |
thus "x mod y = x" by (rule mod_poly_eq) |
850 |
qed |
|
851 |
||
852 |
lemma mod_pCons: |
|
853 |
fixes a and x |
|
854 |
assumes y: "y \<noteq> 0" |
|
855 |
defines b: "b \<equiv> coeff (pCons a (x mod y)) (degree y) / coeff y (degree y)" |
|
856 |
shows "(pCons a x) mod y = (pCons a (x mod y) - smult b y)" |
|
857 |
unfolding b |
|
858 |
apply (rule mod_poly_eq) |
|
29537 | 859 |
apply (rule pdivmod_rel_pCons [OF pdivmod_rel y refl]) |
29451 | 860 |
done |
861 |
||
862 |
||
863 |
subsection {* Evaluation of polynomials *} |
|
864 |
||
865 |
definition |
|
29454
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
866 |
poly :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a" where |
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
867 |
"poly = poly_rec (\<lambda>x. 0) (\<lambda>a p f x. a + x * f x)" |
29451 | 868 |
|
869 |
lemma poly_0 [simp]: "poly 0 x = 0" |
|
29454
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
870 |
unfolding poly_def by (simp add: poly_rec_0) |
29451 | 871 |
|
872 |
lemma poly_pCons [simp]: "poly (pCons a p) x = a + x * poly p x" |
|
29454
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
873 |
unfolding poly_def by (simp add: poly_rec_pCons) |
29451 | 874 |
|
875 |
lemma poly_1 [simp]: "poly 1 x = 1" |
|
876 |
unfolding one_poly_def by simp |
|
877 |
||
29454
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
878 |
lemma poly_monom: |
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
879 |
fixes a x :: "'a::{comm_semiring_1,recpower}" |
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
880 |
shows "poly (monom a n) x = a * x ^ n" |
29451 | 881 |
by (induct n, simp add: monom_0, simp add: monom_Suc power_Suc mult_ac) |
882 |
||
883 |
lemma poly_add [simp]: "poly (p + q) x = poly p x + poly q x" |
|
884 |
apply (induct p arbitrary: q, simp) |
|
885 |
apply (case_tac q, simp, simp add: ring_simps) |
|
886 |
done |
|
887 |
||
888 |
lemma poly_minus [simp]: |
|
29454
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
889 |
fixes x :: "'a::comm_ring" |
29451 | 890 |
shows "poly (- p) x = - poly p x" |
891 |
by (induct p, simp_all) |
|
892 |
||
893 |
lemma poly_diff [simp]: |
|
29454
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
894 |
fixes x :: "'a::comm_ring" |
29451 | 895 |
shows "poly (p - q) x = poly p x - poly q x" |
896 |
by (simp add: diff_minus) |
|
897 |
||
898 |
lemma poly_setsum: "poly (\<Sum>k\<in>A. p k) x = (\<Sum>k\<in>A. poly (p k) x)" |
|
899 |
by (cases "finite A", induct set: finite, simp_all) |
|
900 |
||
901 |
lemma poly_smult [simp]: "poly (smult a p) x = a * poly p x" |
|
902 |
by (induct p, simp, simp add: ring_simps) |
|
903 |
||
904 |
lemma poly_mult [simp]: "poly (p * q) x = poly p x * poly q x" |
|
905 |
by (induct p, simp_all, simp add: ring_simps) |
|
906 |
||
29462 | 907 |
lemma poly_power [simp]: |
908 |
fixes p :: "'a::{comm_semiring_1,recpower} poly" |
|
909 |
shows "poly (p ^ n) x = poly p x ^ n" |
|
910 |
by (induct n, simp, simp add: power_Suc) |
|
911 |
||
29456 | 912 |
|
913 |
subsection {* Synthetic division *} |
|
914 |
||
915 |
definition |
|
916 |
synthetic_divmod :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a poly \<times> 'a" |
|
29478 | 917 |
where [code del]: |
29456 | 918 |
"synthetic_divmod p c = |
919 |
poly_rec (0, 0) (\<lambda>a p (q, r). (pCons r q, a + c * r)) p" |
|
920 |
||
921 |
definition |
|
922 |
synthetic_div :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a poly" |
|
923 |
where |
|
924 |
"synthetic_div p c = fst (synthetic_divmod p c)" |
|
925 |
||
926 |
lemma synthetic_divmod_0 [simp]: |
|
927 |
"synthetic_divmod 0 c = (0, 0)" |
|
928 |
unfolding synthetic_divmod_def |
|
929 |
by (simp add: poly_rec_0) |
|
930 |
||
931 |
lemma synthetic_divmod_pCons [simp]: |
|
932 |
"synthetic_divmod (pCons a p) c = |
|
933 |
(\<lambda>(q, r). (pCons r q, a + c * r)) (synthetic_divmod p c)" |
|
934 |
unfolding synthetic_divmod_def |
|
935 |
by (simp add: poly_rec_pCons) |
|
936 |
||
937 |
lemma snd_synthetic_divmod: "snd (synthetic_divmod p c) = poly p c" |
|
938 |
by (induct p, simp, simp add: split_def) |
|
939 |
||
940 |
lemma synthetic_div_0 [simp]: "synthetic_div 0 c = 0" |
|
941 |
unfolding synthetic_div_def by simp |
|
942 |
||
943 |
lemma synthetic_div_pCons [simp]: |
|
944 |
"synthetic_div (pCons a p) c = pCons (poly p c) (synthetic_div p c)" |
|
945 |
unfolding synthetic_div_def |
|
946 |
by (simp add: split_def snd_synthetic_divmod) |
|
947 |
||
29460
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
948 |
lemma synthetic_div_eq_0_iff: |
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
949 |
"synthetic_div p c = 0 \<longleftrightarrow> degree p = 0" |
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
950 |
by (induct p, simp, case_tac p, simp) |
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
951 |
|
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
952 |
lemma degree_synthetic_div: |
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
953 |
"degree (synthetic_div p c) = degree p - 1" |
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
954 |
by (induct p, simp, simp add: synthetic_div_eq_0_iff) |
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
955 |
|
29457 | 956 |
lemma synthetic_div_correct: |
29456 | 957 |
"p + smult c (synthetic_div p c) = pCons (poly p c) (synthetic_div p c)" |
958 |
by (induct p) simp_all |
|
959 |
||
29457 | 960 |
lemma synthetic_div_unique_lemma: "smult c p = pCons a p \<Longrightarrow> p = 0" |
961 |
by (induct p arbitrary: a) simp_all |
|
962 |
||
963 |
lemma synthetic_div_unique: |
|
964 |
"p + smult c q = pCons r q \<Longrightarrow> r = poly p c \<and> q = synthetic_div p c" |
|
965 |
apply (induct p arbitrary: q r) |
|
966 |
apply (simp, frule synthetic_div_unique_lemma, simp) |
|
967 |
apply (case_tac q, force) |
|
968 |
done |
|
969 |
||
970 |
lemma synthetic_div_correct': |
|
971 |
fixes c :: "'a::comm_ring_1" |
|
972 |
shows "[:-c, 1:] * synthetic_div p c + [:poly p c:] = p" |
|
973 |
using synthetic_div_correct [of p c] |
|
974 |
by (simp add: group_simps) |
|
975 |
||
29460
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
976 |
lemma poly_eq_0_iff_dvd: |
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
977 |
fixes c :: "'a::idom" |
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
978 |
shows "poly p c = 0 \<longleftrightarrow> [:-c, 1:] dvd p" |
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
979 |
proof |
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
980 |
assume "poly p c = 0" |
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
981 |
with synthetic_div_correct' [of c p] |
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
982 |
have "p = [:-c, 1:] * synthetic_div p c" by simp |
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
983 |
then show "[:-c, 1:] dvd p" .. |
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
984 |
next |
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
985 |
assume "[:-c, 1:] dvd p" |
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
986 |
then obtain k where "p = [:-c, 1:] * k" by (rule dvdE) |
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
987 |
then show "poly p c = 0" by simp |
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
988 |
qed |
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
989 |
|
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
990 |
lemma dvd_iff_poly_eq_0: |
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
991 |
fixes c :: "'a::idom" |
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
992 |
shows "[:c, 1:] dvd p \<longleftrightarrow> poly p (-c) = 0" |
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
993 |
by (simp add: poly_eq_0_iff_dvd) |
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
994 |
|
29462 | 995 |
lemma poly_roots_finite: |
996 |
fixes p :: "'a::idom poly" |
|
997 |
shows "p \<noteq> 0 \<Longrightarrow> finite {x. poly p x = 0}" |
|
998 |
proof (induct n \<equiv> "degree p" arbitrary: p) |
|
999 |
case (0 p) |
|
1000 |
then obtain a where "a \<noteq> 0" and "p = [:a:]" |
|
1001 |
by (cases p, simp split: if_splits) |
|
1002 |
then show "finite {x. poly p x = 0}" by simp |
|
1003 |
next |
|
1004 |
case (Suc n p) |
|
1005 |
show "finite {x. poly p x = 0}" |
|
1006 |
proof (cases "\<exists>x. poly p x = 0") |
|
1007 |
case False |
|
1008 |
then show "finite {x. poly p x = 0}" by simp |
|
1009 |
next |
|
1010 |
case True |
|
1011 |
then obtain a where "poly p a = 0" .. |
|
1012 |
then have "[:-a, 1:] dvd p" by (simp only: poly_eq_0_iff_dvd) |
|
1013 |
then obtain k where k: "p = [:-a, 1:] * k" .. |
|
1014 |
with `p \<noteq> 0` have "k \<noteq> 0" by auto |
|
1015 |
with k have "degree p = Suc (degree k)" |
|
1016 |
by (simp add: degree_mult_eq del: mult_pCons_left) |
|
1017 |
with `Suc n = degree p` have "n = degree k" by simp |
|
1018 |
with `k \<noteq> 0` have "finite {x. poly k x = 0}" by (rule Suc.hyps) |
|
1019 |
then have "finite (insert a {x. poly k x = 0})" by simp |
|
1020 |
then show "finite {x. poly p x = 0}" |
|
1021 |
by (simp add: k uminus_add_conv_diff Collect_disj_eq |
|
1022 |
del: mult_pCons_left) |
|
1023 |
qed |
|
1024 |
qed |
|
1025 |
||
29478 | 1026 |
|
1027 |
subsection {* Configuration of the code generator *} |
|
1028 |
||
1029 |
code_datatype "0::'a::zero poly" pCons |
|
1030 |
||
29480 | 1031 |
declare pCons_0_0 [code post] |
1032 |
||
29478 | 1033 |
instantiation poly :: ("{zero,eq}") eq |
1034 |
begin |
|
1035 |
||
1036 |
definition [code del]: |
|
1037 |
"eq_class.eq (p::'a poly) q \<longleftrightarrow> p = q" |
|
1038 |
||
1039 |
instance |
|
1040 |
by default (rule eq_poly_def) |
|
1041 |
||
29451 | 1042 |
end |
29478 | 1043 |
|
1044 |
lemma eq_poly_code [code]: |
|
1045 |
"eq_class.eq (0::_ poly) (0::_ poly) \<longleftrightarrow> True" |
|
1046 |
"eq_class.eq (0::_ poly) (pCons b q) \<longleftrightarrow> eq_class.eq 0 b \<and> eq_class.eq 0 q" |
|
1047 |
"eq_class.eq (pCons a p) (0::_ poly) \<longleftrightarrow> eq_class.eq a 0 \<and> eq_class.eq p 0" |
|
1048 |
"eq_class.eq (pCons a p) (pCons b q) \<longleftrightarrow> eq_class.eq a b \<and> eq_class.eq p q" |
|
1049 |
unfolding eq by simp_all |
|
1050 |
||
1051 |
lemmas coeff_code [code] = |
|
1052 |
coeff_0 coeff_pCons_0 coeff_pCons_Suc |
|
1053 |
||
1054 |
lemmas degree_code [code] = |
|
1055 |
degree_0 degree_pCons_eq_if |
|
1056 |
||
1057 |
lemmas monom_poly_code [code] = |
|
1058 |
monom_0 monom_Suc |
|
1059 |
||
1060 |
lemma add_poly_code [code]: |
|
1061 |
"0 + q = (q :: _ poly)" |
|
1062 |
"p + 0 = (p :: _ poly)" |
|
1063 |
"pCons a p + pCons b q = pCons (a + b) (p + q)" |
|
1064 |
by simp_all |
|
1065 |
||
1066 |
lemma minus_poly_code [code]: |
|
1067 |
"- 0 = (0 :: _ poly)" |
|
1068 |
"- pCons a p = pCons (- a) (- p)" |
|
1069 |
by simp_all |
|
1070 |
||
1071 |
lemma diff_poly_code [code]: |
|
1072 |
"0 - q = (- q :: _ poly)" |
|
1073 |
"p - 0 = (p :: _ poly)" |
|
1074 |
"pCons a p - pCons b q = pCons (a - b) (p - q)" |
|
1075 |
by simp_all |
|
1076 |
||
1077 |
lemmas smult_poly_code [code] = |
|
1078 |
smult_0_right smult_pCons |
|
1079 |
||
1080 |
lemma mult_poly_code [code]: |
|
1081 |
"0 * q = (0 :: _ poly)" |
|
1082 |
"pCons a p * q = smult a q + pCons 0 (p * q)" |
|
1083 |
by simp_all |
|
1084 |
||
1085 |
lemmas poly_code [code] = |
|
1086 |
poly_0 poly_pCons |
|
1087 |
||
1088 |
lemmas synthetic_divmod_code [code] = |
|
1089 |
synthetic_divmod_0 synthetic_divmod_pCons |
|
1090 |
||
1091 |
text {* code generator setup for div and mod *} |
|
1092 |
||
1093 |
definition |
|
29537 | 1094 |
pdivmod :: "'a::field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<times> 'a poly" |
29478 | 1095 |
where |
29537 | 1096 |
[code del]: "pdivmod x y = (x div y, x mod y)" |
29478 | 1097 |
|
29537 | 1098 |
lemma div_poly_code [code]: "x div y = fst (pdivmod x y)" |
1099 |
unfolding pdivmod_def by simp |
|
29478 | 1100 |
|
29537 | 1101 |
lemma mod_poly_code [code]: "x mod y = snd (pdivmod x y)" |
1102 |
unfolding pdivmod_def by simp |
|
29478 | 1103 |
|
29537 | 1104 |
lemma pdivmod_0 [code]: "pdivmod 0 y = (0, 0)" |
1105 |
unfolding pdivmod_def by simp |
|
29478 | 1106 |
|
29537 | 1107 |
lemma pdivmod_pCons [code]: |
1108 |
"pdivmod (pCons a x) y = |
|
29478 | 1109 |
(if y = 0 then (0, pCons a x) else |
29537 | 1110 |
(let (q, r) = pdivmod x y; |
29478 | 1111 |
b = coeff (pCons a r) (degree y) / coeff y (degree y) |
1112 |
in (pCons b q, pCons a r - smult b y)))" |
|
29537 | 1113 |
apply (simp add: pdivmod_def Let_def, safe) |
29478 | 1114 |
apply (rule div_poly_eq) |
29537 | 1115 |
apply (erule pdivmod_rel_pCons [OF pdivmod_rel _ refl]) |
29478 | 1116 |
apply (rule mod_poly_eq) |
29537 | 1117 |
apply (erule pdivmod_rel_pCons [OF pdivmod_rel _ refl]) |
29478 | 1118 |
done |
1119 |
||
1120 |
end |