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