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