| author | wenzelm |
| Fri, 28 Apr 2017 14:50:18 +0200 | |
| changeset 65609 | 9917b8e3b5c1 |
| parent 65577 | 32d4117ad6e8 |
| child 65811 | 2653f1cd8775 |
| permissions | -rw-r--r-- |
| 65435 | 1 |
(* Title: HOL/Computational_Algebra/Polynomial.thy |
| 29451 | 2 |
Author: Brian Huffman |
| 41959 | 3 |
Author: Clemens Ballarin |
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
4 |
Author: Amine Chaieb |
| 52380 | 5 |
Author: Florian Haftmann |
| 29451 | 6 |
*) |
7 |
||
| 60500 | 8 |
section \<open>Polynomials as type over a ring structure\<close> |
| 29451 | 9 |
|
10 |
theory Polynomial |
|
| 65417 | 11 |
imports |
12 |
"~~/src/HOL/Deriv" |
|
13 |
"~~/src/HOL/Library/More_List" |
|
14 |
"~~/src/HOL/Library/Infinite_Set" |
|
| 29451 | 15 |
begin |
16 |
||
| 60500 | 17 |
subsection \<open>Auxiliary: operations for lists (later) representing coefficients\<close> |
| 52380 | 18 |
|
19 |
definition cCons :: "'a::zero \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixr "##" 65) |
|
| 65346 | 20 |
where "x ## xs = (if xs = [] \<and> x = 0 then [] else x # xs)" |
21 |
||
22 |
lemma cCons_0_Nil_eq [simp]: "0 ## [] = []" |
|
| 52380 | 23 |
by (simp add: cCons_def) |
24 |
||
| 65346 | 25 |
lemma cCons_Cons_eq [simp]: "x ## y # ys = x # y # ys" |
| 52380 | 26 |
by (simp add: cCons_def) |
27 |
||
| 65346 | 28 |
lemma cCons_append_Cons_eq [simp]: "x ## xs @ y # ys = x # xs @ y # ys" |
| 52380 | 29 |
by (simp add: cCons_def) |
30 |
||
| 65346 | 31 |
lemma cCons_not_0_eq [simp]: "x \<noteq> 0 \<Longrightarrow> x ## xs = x # xs" |
| 52380 | 32 |
by (simp add: cCons_def) |
33 |
||
34 |
lemma strip_while_not_0_Cons_eq [simp]: |
|
35 |
"strip_while (\<lambda>x. x = 0) (x # xs) = x ## strip_while (\<lambda>x. x = 0) xs" |
|
36 |
proof (cases "x = 0") |
|
| 65346 | 37 |
case False |
38 |
then show ?thesis by simp |
|
| 52380 | 39 |
next |
| 65346 | 40 |
case True |
41 |
show ?thesis |
|
| 52380 | 42 |
proof (induct xs rule: rev_induct) |
| 65346 | 43 |
case Nil |
44 |
with True show ?case by simp |
|
| 52380 | 45 |
next |
| 65346 | 46 |
case (snoc y ys) |
47 |
then show ?case |
|
| 52380 | 48 |
by (cases "y = 0") (simp_all add: append_Cons [symmetric] del: append_Cons) |
49 |
qed |
|
50 |
qed |
|
51 |
||
| 65346 | 52 |
lemma tl_cCons [simp]: "tl (x ## xs) = xs" |
| 52380 | 53 |
by (simp add: cCons_def) |
54 |
||
| 65346 | 55 |
|
| 61585 | 56 |
subsection \<open>Definition of type \<open>poly\<close>\<close> |
| 29451 | 57 |
|
| 61260 | 58 |
typedef (overloaded) 'a poly = "{f :: nat \<Rightarrow> 'a::zero. \<forall>\<^sub>\<infinity> n. f n = 0}"
|
| 63433 | 59 |
morphisms coeff Abs_poly |
60 |
by (auto intro!: ALL_MOST) |
|
| 29451 | 61 |
|
|
59487
adaa430fc0f7
default abstypes and default abstract equations make technical (no_code) annotation superfluous
haftmann
parents:
58881
diff
changeset
|
62 |
setup_lifting type_definition_poly |
| 52380 | 63 |
|
64 |
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
|
65 |
by (simp add: coeff_inject [symmetric] fun_eq_iff) |
| 29451 | 66 |
|
| 52380 | 67 |
lemma poly_eqI: "(\<And>n. coeff p n = coeff q n) \<Longrightarrow> p = q" |
68 |
by (simp add: poly_eq_iff) |
|
69 |
||
|
59983
cd2efd7d06bd
replace almost_everywhere_zero by Infinite_Set.MOST
hoelzl
parents:
59815
diff
changeset
|
70 |
lemma MOST_coeff_eq_0: "\<forall>\<^sub>\<infinity> n. coeff p n = 0" |
| 52380 | 71 |
using coeff [of p] by simp |
| 29451 | 72 |
|
73 |
||
| 60500 | 74 |
subsection \<open>Degree of a polynomial\<close> |
| 29451 | 75 |
|
| 52380 | 76 |
definition degree :: "'a::zero poly \<Rightarrow> nat" |
| 65346 | 77 |
where "degree p = (LEAST n. \<forall>i>n. coeff p i = 0)" |
| 29451 | 78 |
|
| 52380 | 79 |
lemma coeff_eq_0: |
80 |
assumes "degree p < n" |
|
81 |
shows "coeff p n = 0" |
|
| 29451 | 82 |
proof - |
|
59983
cd2efd7d06bd
replace almost_everywhere_zero by Infinite_Set.MOST
hoelzl
parents:
59815
diff
changeset
|
83 |
have "\<exists>n. \<forall>i>n. coeff p i = 0" |
|
cd2efd7d06bd
replace almost_everywhere_zero by Infinite_Set.MOST
hoelzl
parents:
59815
diff
changeset
|
84 |
using MOST_coeff_eq_0 by (simp add: MOST_nat) |
| 52380 | 85 |
then have "\<forall>i>degree p. coeff p i = 0" |
| 29451 | 86 |
unfolding degree_def by (rule LeastI_ex) |
| 52380 | 87 |
with assms show ?thesis by simp |
| 29451 | 88 |
qed |
89 |
||
90 |
lemma le_degree: "coeff p n \<noteq> 0 \<Longrightarrow> n \<le> degree p" |
|
91 |
by (erule contrapos_np, rule coeff_eq_0, simp) |
|
92 |
||
93 |
lemma degree_le: "\<forall>i>n. coeff p i = 0 \<Longrightarrow> degree p \<le> n" |
|
94 |
unfolding degree_def by (erule Least_le) |
|
95 |
||
96 |
lemma less_degree_imp: "n < degree p \<Longrightarrow> \<exists>i>n. coeff p i \<noteq> 0" |
|
97 |
unfolding degree_def by (drule not_less_Least, simp) |
|
98 |
||
99 |
||
| 60500 | 100 |
subsection \<open>The zero polynomial\<close> |
| 29451 | 101 |
|
102 |
instantiation poly :: (zero) zero |
|
103 |
begin |
|
104 |
||
| 52380 | 105 |
lift_definition zero_poly :: "'a poly" |
| 65390 | 106 |
is "\<lambda>_. 0" |
107 |
by (rule MOST_I) simp |
|
| 29451 | 108 |
|
109 |
instance .. |
|
| 52380 | 110 |
|
| 29451 | 111 |
end |
112 |
||
| 65346 | 113 |
lemma coeff_0 [simp]: "coeff 0 n = 0" |
| 52380 | 114 |
by transfer rule |
| 29451 | 115 |
|
| 65346 | 116 |
lemma degree_0 [simp]: "degree 0 = 0" |
| 29451 | 117 |
by (rule order_antisym [OF degree_le le0]) simp |
118 |
||
119 |
lemma leading_coeff_neq_0: |
|
| 52380 | 120 |
assumes "p \<noteq> 0" |
121 |
shows "coeff p (degree p) \<noteq> 0" |
|
| 29451 | 122 |
proof (cases "degree p") |
123 |
case 0 |
|
| 65346 | 124 |
from \<open>p \<noteq> 0\<close> obtain n where "coeff p n \<noteq> 0" |
125 |
by (auto simp add: poly_eq_iff) |
|
126 |
then have "n \<le> degree p" |
|
127 |
by (rule le_degree) |
|
128 |
with \<open>coeff p n \<noteq> 0\<close> and \<open>degree p = 0\<close> show "coeff p (degree p) \<noteq> 0" |
|
129 |
by simp |
|
| 29451 | 130 |
next |
131 |
case (Suc n) |
|
| 65346 | 132 |
from \<open>degree p = Suc n\<close> have "n < degree p" |
133 |
by simp |
|
134 |
then have "\<exists>i>n. coeff p i \<noteq> 0" |
|
135 |
by (rule less_degree_imp) |
|
136 |
then obtain i where "n < i" and "coeff p i \<noteq> 0" |
|
137 |
by blast |
|
138 |
from \<open>degree p = Suc n\<close> and \<open>n < i\<close> have "degree p \<le> i" |
|
139 |
by simp |
|
140 |
also from \<open>coeff p i \<noteq> 0\<close> have "i \<le> degree p" |
|
141 |
by (rule le_degree) |
|
| 29451 | 142 |
finally have "degree p = i" . |
| 60500 | 143 |
with \<open>coeff p i \<noteq> 0\<close> show "coeff p (degree p) \<noteq> 0" by simp |
| 29451 | 144 |
qed |
145 |
||
| 65346 | 146 |
lemma leading_coeff_0_iff [simp]: "coeff p (degree p) = 0 \<longleftrightarrow> p = 0" |
147 |
by (cases "p = 0") (simp_all add: leading_coeff_neq_0) |
|
| 29451 | 148 |
|
| 64795 | 149 |
lemma eq_zero_or_degree_less: |
150 |
assumes "degree p \<le> n" and "coeff p n = 0" |
|
151 |
shows "p = 0 \<or> degree p < n" |
|
152 |
proof (cases n) |
|
153 |
case 0 |
|
| 65346 | 154 |
with \<open>degree p \<le> n\<close> and \<open>coeff p n = 0\<close> have "coeff p (degree p) = 0" |
155 |
by simp |
|
| 64795 | 156 |
then have "p = 0" by simp |
157 |
then show ?thesis .. |
|
158 |
next |
|
159 |
case (Suc m) |
|
| 65346 | 160 |
from \<open>degree p \<le> n\<close> have "\<forall>i>n. coeff p i = 0" |
161 |
by (simp add: coeff_eq_0) |
|
162 |
with \<open>coeff p n = 0\<close> have "\<forall>i\<ge>n. coeff p i = 0" |
|
163 |
by (simp add: le_less) |
|
164 |
with \<open>n = Suc m\<close> have "\<forall>i>m. coeff p i = 0" |
|
165 |
by (simp add: less_eq_Suc_le) |
|
| 64795 | 166 |
then have "degree p \<le> m" |
167 |
by (rule degree_le) |
|
| 65346 | 168 |
with \<open>n = Suc m\<close> have "degree p < n" |
169 |
by (simp add: less_Suc_eq_le) |
|
| 64795 | 170 |
then show ?thesis .. |
171 |
qed |
|
172 |
||
173 |
lemma coeff_0_degree_minus_1: "coeff rrr dr = 0 \<Longrightarrow> degree rrr \<le> dr \<Longrightarrow> degree rrr \<le> dr - 1" |
|
174 |
using eq_zero_or_degree_less by fastforce |
|
175 |
||
| 29451 | 176 |
|
| 60500 | 177 |
subsection \<open>List-style constructor for polynomials\<close> |
| 29451 | 178 |
|
| 52380 | 179 |
lift_definition pCons :: "'a::zero \<Rightarrow> 'a poly \<Rightarrow> 'a poly" |
| 55415 | 180 |
is "\<lambda>a p. case_nat a (coeff p)" |
|
59983
cd2efd7d06bd
replace almost_everywhere_zero by Infinite_Set.MOST
hoelzl
parents:
59815
diff
changeset
|
181 |
by (rule MOST_SucD) (simp add: MOST_coeff_eq_0) |
| 29451 | 182 |
|
| 52380 | 183 |
lemmas coeff_pCons = pCons.rep_eq |
| 29455 | 184 |
|
| 65346 | 185 |
lemma coeff_pCons_0 [simp]: "coeff (pCons a p) 0 = a" |
| 52380 | 186 |
by transfer simp |
| 29455 | 187 |
|
| 65346 | 188 |
lemma coeff_pCons_Suc [simp]: "coeff (pCons a p) (Suc n) = coeff p n" |
| 29451 | 189 |
by (simp add: coeff_pCons) |
190 |
||
| 65346 | 191 |
lemma degree_pCons_le: "degree (pCons a p) \<le> Suc (degree p)" |
| 52380 | 192 |
by (rule degree_le) (simp add: coeff_eq_0 coeff_pCons split: nat.split) |
| 29451 | 193 |
|
| 65346 | 194 |
lemma degree_pCons_eq: "p \<noteq> 0 \<Longrightarrow> degree (pCons a p) = Suc (degree p)" |
| 52380 | 195 |
apply (rule order_antisym [OF degree_pCons_le]) |
196 |
apply (rule le_degree, simp) |
|
197 |
done |
|
| 29451 | 198 |
|
| 65346 | 199 |
lemma degree_pCons_0: "degree (pCons a 0) = 0" |
| 52380 | 200 |
apply (rule order_antisym [OF _ le0]) |
201 |
apply (rule degree_le, simp add: coeff_pCons split: nat.split) |
|
202 |
done |
|
| 29451 | 203 |
|
| 65346 | 204 |
lemma degree_pCons_eq_if [simp]: "degree (pCons a p) = (if p = 0 then 0 else Suc (degree p))" |
| 52380 | 205 |
apply (cases "p = 0", simp_all) |
206 |
apply (rule order_antisym [OF _ le0]) |
|
207 |
apply (rule degree_le, simp add: coeff_pCons split: nat.split) |
|
208 |
apply (rule order_antisym [OF degree_pCons_le]) |
|
209 |
apply (rule le_degree, simp) |
|
210 |
done |
|
| 29451 | 211 |
|
| 65346 | 212 |
lemma pCons_0_0 [simp]: "pCons 0 0 = 0" |
| 52380 | 213 |
by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) |
| 29451 | 214 |
|
| 65346 | 215 |
lemma pCons_eq_iff [simp]: "pCons a p = pCons b q \<longleftrightarrow> a = b \<and> p = q" |
| 52380 | 216 |
proof safe |
| 29451 | 217 |
assume "pCons a p = pCons b q" |
| 65346 | 218 |
then have "coeff (pCons a p) 0 = coeff (pCons b q) 0" |
219 |
by simp |
|
220 |
then show "a = b" |
|
221 |
by simp |
|
| 29451 | 222 |
next |
223 |
assume "pCons a p = pCons b q" |
|
| 65346 | 224 |
then have "coeff (pCons a p) (Suc n) = coeff (pCons b q) (Suc n)" for n |
225 |
by simp |
|
226 |
then show "p = q" |
|
227 |
by (simp add: poly_eq_iff) |
|
| 29451 | 228 |
qed |
229 |
||
| 65346 | 230 |
lemma pCons_eq_0_iff [simp]: "pCons a p = 0 \<longleftrightarrow> a = 0 \<and> p = 0" |
| 29451 | 231 |
using pCons_eq_iff [of a p 0 0] by simp |
232 |
||
233 |
lemma pCons_cases [cases type: poly]: |
|
234 |
obtains (pCons) a q where "p = pCons a q" |
|
235 |
proof |
|
236 |
show "p = pCons (coeff p 0) (Abs_poly (\<lambda>n. coeff p (Suc n)))" |
|
| 52380 | 237 |
by transfer |
| 65346 | 238 |
(simp_all add: MOST_inj[where f=Suc and P="\<lambda>n. p n = 0" for p] fun_eq_iff Abs_poly_inverse |
239 |
split: nat.split) |
|
| 29451 | 240 |
qed |
241 |
||
242 |
lemma pCons_induct [case_names 0 pCons, induct type: poly]: |
|
243 |
assumes zero: "P 0" |
|
| 54856 | 244 |
assumes pCons: "\<And>a p. a \<noteq> 0 \<or> p \<noteq> 0 \<Longrightarrow> P p \<Longrightarrow> P (pCons a p)" |
| 29451 | 245 |
shows "P p" |
246 |
proof (induct p rule: measure_induct_rule [where f=degree]) |
|
247 |
case (less p) |
|
248 |
obtain a q where "p = pCons a q" by (rule pCons_cases) |
|
249 |
have "P q" |
|
250 |
proof (cases "q = 0") |
|
251 |
case True |
|
252 |
then show "P q" by (simp add: zero) |
|
253 |
next |
|
254 |
case False |
|
255 |
then have "degree (pCons a q) = Suc (degree q)" |
|
256 |
by (rule degree_pCons_eq) |
|
| 65346 | 257 |
with \<open>p = pCons a q\<close> have "degree q < degree p" |
258 |
by simp |
|
| 29451 | 259 |
then show "P q" |
260 |
by (rule less.hyps) |
|
261 |
qed |
|
| 54856 | 262 |
have "P (pCons a q)" |
263 |
proof (cases "a \<noteq> 0 \<or> q \<noteq> 0") |
|
264 |
case True |
|
| 60500 | 265 |
with \<open>P q\<close> show ?thesis by (auto intro: pCons) |
| 54856 | 266 |
next |
267 |
case False |
|
268 |
with zero show ?thesis by simp |
|
269 |
qed |
|
| 65346 | 270 |
with \<open>p = pCons a q\<close> show ?case |
271 |
by simp |
|
| 29451 | 272 |
qed |
273 |
||
| 60570 | 274 |
lemma degree_eq_zeroE: |
275 |
fixes p :: "'a::zero poly" |
|
276 |
assumes "degree p = 0" |
|
277 |
obtains a where "p = pCons a 0" |
|
278 |
proof - |
|
| 65346 | 279 |
obtain a q where p: "p = pCons a q" |
280 |
by (cases p) |
|
281 |
with assms have "q = 0" |
|
282 |
by (cases "q = 0") simp_all |
|
283 |
with p have "p = pCons a 0" |
|
284 |
by simp |
|
285 |
then show thesis .. |
|
| 60570 | 286 |
qed |
287 |
||
| 29451 | 288 |
|
| 62422 | 289 |
subsection \<open>Quickcheck generator for polynomials\<close> |
290 |
||
291 |
quickcheck_generator poly constructors: "0 :: _ poly", pCons |
|
292 |
||
293 |
||
| 60500 | 294 |
subsection \<open>List-style syntax for polynomials\<close> |
| 52380 | 295 |
|
| 65346 | 296 |
syntax "_poly" :: "args \<Rightarrow> 'a poly" ("[:(_):]")
|
| 52380 | 297 |
translations |
| 65346 | 298 |
"[:x, xs:]" \<rightleftharpoons> "CONST pCons x [:xs:]" |
299 |
"[:x:]" \<rightleftharpoons> "CONST pCons x 0" |
|
300 |
"[:x:]" \<leftharpoondown> "CONST pCons x (_constrain 0 t)" |
|
| 52380 | 301 |
|
302 |
||
| 60500 | 303 |
subsection \<open>Representation of polynomials by lists of coefficients\<close> |
| 52380 | 304 |
|
305 |
primrec Poly :: "'a::zero list \<Rightarrow> 'a poly" |
|
| 65346 | 306 |
where |
307 |
[code_post]: "Poly [] = 0" |
|
308 |
| [code_post]: "Poly (a # as) = pCons a (Poly as)" |
|
309 |
||
310 |
lemma Poly_replicate_0 [simp]: "Poly (replicate n 0) = 0" |
|
| 52380 | 311 |
by (induct n) simp_all |
312 |
||
| 65346 | 313 |
lemma Poly_eq_0: "Poly as = 0 \<longleftrightarrow> (\<exists>n. as = replicate n 0)" |
| 52380 | 314 |
by (induct as) (auto simp add: Cons_replicate_eq) |
|
63027
8de0ebee3f1c
several updates on polynomial long division and pseudo division
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
62422
diff
changeset
|
315 |
|
| 65346 | 316 |
lemma Poly_append_replicate_zero [simp]: "Poly (as @ replicate n 0) = Poly as" |
|
63027
8de0ebee3f1c
several updates on polynomial long division and pseudo division
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
62422
diff
changeset
|
317 |
by (induct as) simp_all |
|
8de0ebee3f1c
several updates on polynomial long division and pseudo division
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
62422
diff
changeset
|
318 |
|
| 65346 | 319 |
lemma Poly_snoc_zero [simp]: "Poly (as @ [0]) = Poly as" |
|
63027
8de0ebee3f1c
several updates on polynomial long division and pseudo division
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
62422
diff
changeset
|
320 |
using Poly_append_replicate_zero [of as 1] by simp |
|
8de0ebee3f1c
several updates on polynomial long division and pseudo division
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
62422
diff
changeset
|
321 |
|
| 65346 | 322 |
lemma Poly_cCons_eq_pCons_Poly [simp]: "Poly (a ## p) = pCons a (Poly p)" |
|
63027
8de0ebee3f1c
several updates on polynomial long division and pseudo division
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
62422
diff
changeset
|
323 |
by (simp add: cCons_def) |
|
8de0ebee3f1c
several updates on polynomial long division and pseudo division
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
62422
diff
changeset
|
324 |
|
| 65346 | 325 |
lemma Poly_on_rev_starting_with_0 [simp]: "hd as = 0 \<Longrightarrow> Poly (rev (tl as)) = Poly (rev as)" |
326 |
by (cases as) simp_all |
|
|
63027
8de0ebee3f1c
several updates on polynomial long division and pseudo division
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
62422
diff
changeset
|
327 |
|
| 62065 | 328 |
lemma degree_Poly: "degree (Poly xs) \<le> length xs" |
| 65346 | 329 |
by (induct xs) simp_all |
330 |
||
331 |
lemma coeff_Poly_eq [simp]: "coeff (Poly xs) = nth_default 0 xs" |
|
|
63027
8de0ebee3f1c
several updates on polynomial long division and pseudo division
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
62422
diff
changeset
|
332 |
by (induct xs) (simp_all add: fun_eq_iff coeff_pCons split: nat.splits) |
|
8de0ebee3f1c
several updates on polynomial long division and pseudo division
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
62422
diff
changeset
|
333 |
|
| 52380 | 334 |
definition coeffs :: "'a poly \<Rightarrow> 'a::zero list" |
| 65346 | 335 |
where "coeffs p = (if p = 0 then [] else map (\<lambda>i. coeff p i) [0 ..< Suc (degree p)])" |
336 |
||
337 |
lemma coeffs_eq_Nil [simp]: "coeffs p = [] \<longleftrightarrow> p = 0" |
|
| 52380 | 338 |
by (simp add: coeffs_def) |
339 |
||
| 65346 | 340 |
lemma not_0_coeffs_not_Nil: "p \<noteq> 0 \<Longrightarrow> coeffs p \<noteq> []" |
| 52380 | 341 |
by simp |
342 |
||
| 65346 | 343 |
lemma coeffs_0_eq_Nil [simp]: "coeffs 0 = []" |
| 52380 | 344 |
by simp |
|
29454
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
345 |
|
| 65346 | 346 |
lemma coeffs_pCons_eq_cCons [simp]: "coeffs (pCons a p) = a ## coeffs p" |
| 52380 | 347 |
proof - |
| 65346 | 348 |
have *: "\<forall>m\<in>set ms. m > 0 \<Longrightarrow> map (case_nat x f) ms = map f (map (\<lambda>n. n - 1) ms)" |
349 |
for ms :: "nat list" and f :: "nat \<Rightarrow> 'a" and x :: "'a" |
|
350 |
by (induct ms) (auto split: nat.split) |
|
| 52380 | 351 |
show ?thesis |
| 65346 | 352 |
by (simp add: * coeffs_def upt_conv_Cons coeff_pCons map_decr_upt del: upt_Suc) |
| 52380 | 353 |
qed |
354 |
||
| 62065 | 355 |
lemma length_coeffs: "p \<noteq> 0 \<Longrightarrow> length (coeffs p) = degree p + 1" |
356 |
by (simp add: coeffs_def) |
|
| 64860 | 357 |
|
| 65346 | 358 |
lemma coeffs_nth: "p \<noteq> 0 \<Longrightarrow> n \<le> degree p \<Longrightarrow> coeffs p ! n = coeff p n" |
359 |
by (auto simp: coeffs_def simp del: upt_Suc) |
|
360 |
||
361 |
lemma coeff_in_coeffs: "p \<noteq> 0 \<Longrightarrow> n \<le> degree p \<Longrightarrow> coeff p n \<in> set (coeffs p)" |
|
362 |
using coeffs_nth [of p n, symmetric] by (simp add: length_coeffs) |
|
363 |
||
364 |
lemma not_0_cCons_eq [simp]: "p \<noteq> 0 \<Longrightarrow> a ## coeffs p = a # coeffs p" |
|
| 52380 | 365 |
by (simp add: cCons_def) |
366 |
||
| 65346 | 367 |
lemma Poly_coeffs [simp, code abstype]: "Poly (coeffs p) = p" |
| 54856 | 368 |
by (induct p) auto |
| 52380 | 369 |
|
| 65346 | 370 |
lemma coeffs_Poly [simp]: "coeffs (Poly as) = strip_while (HOL.eq 0) as" |
| 52380 | 371 |
proof (induct as) |
| 65346 | 372 |
case Nil |
373 |
then show ?case by simp |
|
| 52380 | 374 |
next |
375 |
case (Cons a as) |
|
| 65346 | 376 |
from replicate_length_same [of as 0] have "(\<forall>n. as \<noteq> replicate n 0) \<longleftrightarrow> (\<exists>a\<in>set as. a \<noteq> 0)" |
377 |
by (auto dest: sym [of _ as]) |
|
| 52380 | 378 |
with Cons show ?case by auto |
379 |
qed |
|
380 |
||
| 65390 | 381 |
lemma no_trailing_coeffs [simp]: |
382 |
"no_trailing (HOL.eq 0) (coeffs p)" |
|
383 |
by (induct p) auto |
|
384 |
||
385 |
lemma strip_while_coeffs [simp]: |
|
386 |
"strip_while (HOL.eq 0) (coeffs p) = coeffs p" |
|
387 |
by simp |
|
| 52380 | 388 |
|
| 65346 | 389 |
lemma coeffs_eq_iff: "p = q \<longleftrightarrow> coeffs p = coeffs q" |
390 |
(is "?P \<longleftrightarrow> ?Q") |
|
| 52380 | 391 |
proof |
| 65346 | 392 |
assume ?P |
393 |
then show ?Q by simp |
|
| 52380 | 394 |
next |
395 |
assume ?Q |
|
396 |
then have "Poly (coeffs p) = Poly (coeffs q)" by simp |
|
397 |
then show ?P by simp |
|
398 |
qed |
|
399 |
||
| 65346 | 400 |
lemma nth_default_coeffs_eq: "nth_default 0 (coeffs p) = coeff p" |
| 52380 | 401 |
by (simp add: fun_eq_iff coeff_Poly_eq [symmetric]) |
402 |
||
| 65346 | 403 |
lemma [code]: "coeff p = nth_default 0 (coeffs p)" |
| 52380 | 404 |
by (simp add: nth_default_coeffs_eq) |
405 |
||
406 |
lemma coeffs_eqI: |
|
407 |
assumes coeff: "\<And>n. coeff p n = nth_default 0 xs n" |
|
| 65390 | 408 |
assumes zero: "no_trailing (HOL.eq 0) xs" |
| 52380 | 409 |
shows "coeffs p = xs" |
410 |
proof - |
|
| 65390 | 411 |
from coeff have "p = Poly xs" |
412 |
by (simp add: poly_eq_iff) |
|
413 |
with zero show ?thesis by simp |
|
| 52380 | 414 |
qed |
415 |
||
| 65346 | 416 |
lemma degree_eq_length_coeffs [code]: "degree p = length (coeffs p) - 1" |
| 52380 | 417 |
by (simp add: coeffs_def) |
418 |
||
| 65346 | 419 |
lemma length_coeffs_degree: "p \<noteq> 0 \<Longrightarrow> length (coeffs p) = Suc (degree p)" |
420 |
by (induct p) (auto simp: cCons_def) |
|
421 |
||
422 |
lemma [code abstract]: "coeffs 0 = []" |
|
| 52380 | 423 |
by (fact coeffs_0_eq_Nil) |
424 |
||
| 65346 | 425 |
lemma [code abstract]: "coeffs (pCons a p) = a ## coeffs p" |
| 52380 | 426 |
by (fact coeffs_pCons_eq_cCons) |
427 |
||
428 |
instantiation poly :: ("{zero, equal}") equal
|
|
429 |
begin |
|
430 |
||
| 65346 | 431 |
definition [code]: "HOL.equal (p::'a poly) q \<longleftrightarrow> HOL.equal (coeffs p) (coeffs q)" |
| 52380 | 432 |
|
| 60679 | 433 |
instance |
434 |
by standard (simp add: equal equal_poly_def coeffs_eq_iff) |
|
| 52380 | 435 |
|
436 |
end |
|
437 |
||
| 60679 | 438 |
lemma [code nbe]: "HOL.equal (p :: _ poly) p \<longleftrightarrow> True" |
| 52380 | 439 |
by (fact equal_refl) |
|
29454
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
440 |
|
| 52380 | 441 |
definition is_zero :: "'a::zero poly \<Rightarrow> bool" |
| 65346 | 442 |
where [code]: "is_zero p \<longleftrightarrow> List.null (coeffs p)" |
443 |
||
444 |
lemma is_zero_null [code_abbrev]: "is_zero p \<longleftrightarrow> p = 0" |
|
| 52380 | 445 |
by (simp add: is_zero_def null_def) |
446 |
||
| 65346 | 447 |
|
|
63027
8de0ebee3f1c
several updates on polynomial long division and pseudo division
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
62422
diff
changeset
|
448 |
subsubsection \<open>Reconstructing the polynomial from the list\<close> |
| 63145 | 449 |
\<comment> \<open>contributed by Sebastiaan J.C. Joosten and René Thiemann\<close> |
|
63027
8de0ebee3f1c
several updates on polynomial long division and pseudo division
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
62422
diff
changeset
|
450 |
|
|
8de0ebee3f1c
several updates on polynomial long division and pseudo division
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
62422
diff
changeset
|
451 |
definition poly_of_list :: "'a::comm_monoid_add list \<Rightarrow> 'a poly" |
| 65346 | 452 |
where [simp]: "poly_of_list = Poly" |
453 |
||
454 |
lemma poly_of_list_impl [code abstract]: "coeffs (poly_of_list as) = strip_while (HOL.eq 0) as" |
|
|
63027
8de0ebee3f1c
several updates on polynomial long division and pseudo division
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
62422
diff
changeset
|
455 |
by simp |
|
8de0ebee3f1c
several updates on polynomial long division and pseudo division
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
62422
diff
changeset
|
456 |
|
| 52380 | 457 |
|
| 60500 | 458 |
subsection \<open>Fold combinator for polynomials\<close> |
| 52380 | 459 |
|
460 |
definition fold_coeffs :: "('a::zero \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a poly \<Rightarrow> 'b \<Rightarrow> 'b"
|
|
| 65346 | 461 |
where "fold_coeffs f p = foldr f (coeffs p)" |
462 |
||
463 |
lemma fold_coeffs_0_eq [simp]: "fold_coeffs f 0 = id" |
|
| 52380 | 464 |
by (simp add: fold_coeffs_def) |
465 |
||
| 65346 | 466 |
lemma fold_coeffs_pCons_eq [simp]: "f 0 = id \<Longrightarrow> fold_coeffs f (pCons a p) = f a \<circ> fold_coeffs f p" |
| 52380 | 467 |
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
|
468 |
|
| 65346 | 469 |
lemma fold_coeffs_pCons_0_0_eq [simp]: "fold_coeffs f (pCons 0 0) = id" |
| 52380 | 470 |
by (simp add: fold_coeffs_def) |
471 |
||
472 |
lemma fold_coeffs_pCons_coeff_not_0_eq [simp]: |
|
473 |
"a \<noteq> 0 \<Longrightarrow> fold_coeffs f (pCons a p) = f a \<circ> fold_coeffs f p" |
|
474 |
by (simp add: fold_coeffs_def) |
|
475 |
||
476 |
lemma fold_coeffs_pCons_not_0_0_eq [simp]: |
|
477 |
"p \<noteq> 0 \<Longrightarrow> fold_coeffs f (pCons a p) = f a \<circ> fold_coeffs f p" |
|
478 |
by (simp add: fold_coeffs_def) |
|
479 |
||
| 64795 | 480 |
|
| 60500 | 481 |
subsection \<open>Canonical morphism on polynomials -- evaluation\<close> |
| 52380 | 482 |
|
483 |
definition poly :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a" |
|
| 65346 | 484 |
where "poly p = fold_coeffs (\<lambda>a f x. a + x * f x) p (\<lambda>x. 0)" \<comment> \<open>The Horner Schema\<close> |
485 |
||
486 |
lemma poly_0 [simp]: "poly 0 x = 0" |
|
| 52380 | 487 |
by (simp add: poly_def) |
| 65346 | 488 |
|
489 |
lemma poly_pCons [simp]: "poly (pCons a p) x = a + x * poly p x" |
|
| 52380 | 490 |
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
|
491 |
|
| 65346 | 492 |
lemma poly_altdef: "poly p x = (\<Sum>i\<le>degree p. coeff p i * x ^ i)" |
493 |
for x :: "'a::{comm_semiring_0,semiring_1}"
|
|
| 62065 | 494 |
proof (induction p rule: pCons_induct) |
| 65346 | 495 |
case 0 |
496 |
then show ?case |
|
497 |
by simp |
|
498 |
next |
|
| 62065 | 499 |
case (pCons a p) |
| 65346 | 500 |
show ?case |
501 |
proof (cases "p = 0") |
|
502 |
case True |
|
503 |
then show ?thesis by simp |
|
504 |
next |
|
505 |
case False |
|
506 |
let ?p' = "pCons a p" |
|
507 |
note poly_pCons[of a p x] |
|
508 |
also note pCons.IH |
|
509 |
also have "a + x * (\<Sum>i\<le>degree p. coeff p i * x ^ i) = |
|
510 |
coeff ?p' 0 * x^0 + (\<Sum>i\<le>degree p. coeff ?p' (Suc i) * x^Suc i)" |
|
511 |
by (simp add: field_simps sum_distrib_left coeff_pCons) |
|
512 |
also note sum_atMost_Suc_shift[symmetric] |
|
513 |
also note degree_pCons_eq[OF \<open>p \<noteq> 0\<close>, of a, symmetric] |
|
514 |
finally show ?thesis . |
|
515 |
qed |
|
516 |
qed |
|
| 62065 | 517 |
|
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
518 |
lemma poly_0_coeff_0: "poly p 0 = coeff p 0" |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
519 |
by (cases p) (auto simp: poly_altdef) |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
520 |
|
|
29454
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
521 |
|
| 60500 | 522 |
subsection \<open>Monomials\<close> |
| 29451 | 523 |
|
| 52380 | 524 |
lift_definition monom :: "'a \<Rightarrow> nat \<Rightarrow> 'a::zero poly" |
525 |
is "\<lambda>a m n. if m = n then a else 0" |
|
|
59983
cd2efd7d06bd
replace almost_everywhere_zero by Infinite_Set.MOST
hoelzl
parents:
59815
diff
changeset
|
526 |
by (simp add: MOST_iff_cofinite) |
| 52380 | 527 |
|
| 65346 | 528 |
lemma coeff_monom [simp]: "coeff (monom a m) n = (if m = n then a else 0)" |
| 52380 | 529 |
by transfer rule |
| 29451 | 530 |
|
| 65346 | 531 |
lemma monom_0: "monom a 0 = pCons a 0" |
| 52380 | 532 |
by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) |
| 29451 | 533 |
|
| 65346 | 534 |
lemma monom_Suc: "monom a (Suc n) = pCons 0 (monom a n)" |
| 52380 | 535 |
by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) |
| 29451 | 536 |
|
537 |
lemma monom_eq_0 [simp]: "monom 0 n = 0" |
|
| 52380 | 538 |
by (rule poly_eqI) simp |
| 29451 | 539 |
|
540 |
lemma monom_eq_0_iff [simp]: "monom a n = 0 \<longleftrightarrow> a = 0" |
|
| 52380 | 541 |
by (simp add: poly_eq_iff) |
| 29451 | 542 |
|
543 |
lemma monom_eq_iff [simp]: "monom a n = monom b n \<longleftrightarrow> a = b" |
|
| 52380 | 544 |
by (simp add: poly_eq_iff) |
| 29451 | 545 |
|
546 |
lemma degree_monom_le: "degree (monom a n) \<le> n" |
|
547 |
by (rule degree_le, simp) |
|
548 |
||
549 |
lemma degree_monom_eq: "a \<noteq> 0 \<Longrightarrow> degree (monom a n) = n" |
|
550 |
apply (rule order_antisym [OF degree_monom_le]) |
|
| 65346 | 551 |
apply (rule le_degree) |
552 |
apply simp |
|
| 29451 | 553 |
done |
554 |
||
| 52380 | 555 |
lemma coeffs_monom [code abstract]: |
556 |
"coeffs (monom a n) = (if a = 0 then [] else replicate n 0 @ [a])" |
|
557 |
by (induct n) (simp_all add: monom_0 monom_Suc) |
|
558 |
||
| 65346 | 559 |
lemma fold_coeffs_monom [simp]: "a \<noteq> 0 \<Longrightarrow> fold_coeffs f (monom a n) = f 0 ^^ n \<circ> f a" |
| 52380 | 560 |
by (simp add: fold_coeffs_def coeffs_monom fun_eq_iff) |
561 |
||
| 65346 | 562 |
lemma poly_monom: "poly (monom a n) x = a * x ^ n" |
563 |
for a x :: "'a::comm_semiring_1" |
|
564 |
by (cases "a = 0", simp_all) (induct n, simp_all add: mult.left_commute poly_def) |
|
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
565 |
|
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
566 |
lemma monom_eq_iff': "monom c n = monom d m \<longleftrightarrow> c = d \<and> (c = 0 \<or> n = m)" |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
567 |
by (auto simp: poly_eq_iff) |
| 65346 | 568 |
|
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
569 |
lemma monom_eq_const_iff: "monom c n = [:d:] \<longleftrightarrow> c = d \<and> (c = 0 \<or> n = 0)" |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
570 |
using monom_eq_iff'[of c n d 0] by (simp add: monom_0) |
| 64795 | 571 |
|
572 |
||
573 |
subsection \<open>Leading coefficient\<close> |
|
574 |
||
575 |
abbreviation lead_coeff:: "'a::zero poly \<Rightarrow> 'a" |
|
576 |
where "lead_coeff p \<equiv> coeff p (degree p)" |
|
577 |
||
578 |
lemma lead_coeff_pCons[simp]: |
|
579 |
"p \<noteq> 0 \<Longrightarrow> lead_coeff (pCons a p) = lead_coeff p" |
|
580 |
"p = 0 \<Longrightarrow> lead_coeff (pCons a p) = a" |
|
581 |
by auto |
|
582 |
||
583 |
lemma lead_coeff_monom [simp]: "lead_coeff (monom c n) = c" |
|
584 |
by (cases "c = 0") (simp_all add: degree_monom_eq) |
|
585 |
||
586 |
||
| 60500 | 587 |
subsection \<open>Addition and subtraction\<close> |
| 29451 | 588 |
|
589 |
instantiation poly :: (comm_monoid_add) comm_monoid_add |
|
590 |
begin |
|
591 |
||
| 52380 | 592 |
lift_definition plus_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" |
593 |
is "\<lambda>p q n. coeff p n + coeff q n" |
|
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59983
diff
changeset
|
594 |
proof - |
| 60679 | 595 |
fix q p :: "'a poly" |
596 |
show "\<forall>\<^sub>\<infinity>n. coeff p n + coeff q n = 0" |
|
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59983
diff
changeset
|
597 |
using MOST_coeff_eq_0[of p] MOST_coeff_eq_0[of q] by eventually_elim simp |
| 52380 | 598 |
qed |
| 29451 | 599 |
|
| 60679 | 600 |
lemma coeff_add [simp]: "coeff (p + q) n = coeff p n + coeff q n" |
| 52380 | 601 |
by (simp add: plus_poly.rep_eq) |
| 29451 | 602 |
|
| 60679 | 603 |
instance |
604 |
proof |
|
| 29451 | 605 |
fix p q r :: "'a poly" |
606 |
show "(p + q) + r = p + (q + r)" |
|
|
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
57482
diff
changeset
|
607 |
by (simp add: poly_eq_iff add.assoc) |
| 29451 | 608 |
show "p + q = q + p" |
|
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
57482
diff
changeset
|
609 |
by (simp add: poly_eq_iff add.commute) |
| 29451 | 610 |
show "0 + p = p" |
| 52380 | 611 |
by (simp add: poly_eq_iff) |
| 29451 | 612 |
qed |
613 |
||
614 |
end |
|
615 |
||
|
59815
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
616 |
instantiation poly :: (cancel_comm_monoid_add) cancel_comm_monoid_add |
|
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
617 |
begin |
|
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
618 |
|
|
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
619 |
lift_definition minus_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" |
|
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
620 |
is "\<lambda>p q n. coeff p n - coeff q n" |
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59983
diff
changeset
|
621 |
proof - |
| 60679 | 622 |
fix q p :: "'a poly" |
623 |
show "\<forall>\<^sub>\<infinity>n. coeff p n - coeff q n = 0" |
|
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59983
diff
changeset
|
624 |
using MOST_coeff_eq_0[of p] MOST_coeff_eq_0[of q] by eventually_elim simp |
|
59815
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
625 |
qed |
|
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
626 |
|
| 60679 | 627 |
lemma coeff_diff [simp]: "coeff (p - q) n = coeff p n - coeff q n" |
|
59815
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
628 |
by (simp add: minus_poly.rep_eq) |
|
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
629 |
|
| 60679 | 630 |
instance |
631 |
proof |
|
| 29540 | 632 |
fix p q r :: "'a poly" |
|
59815
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
633 |
show "p + q - p = q" |
| 52380 | 634 |
by (simp add: poly_eq_iff) |
|
59815
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
635 |
show "p - q - r = p - (q + r)" |
|
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
636 |
by (simp add: poly_eq_iff diff_diff_eq) |
| 29540 | 637 |
qed |
638 |
||
|
59815
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
639 |
end |
|
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
640 |
|
| 29451 | 641 |
instantiation poly :: (ab_group_add) ab_group_add |
642 |
begin |
|
643 |
||
| 52380 | 644 |
lift_definition uminus_poly :: "'a poly \<Rightarrow> 'a poly" |
645 |
is "\<lambda>p n. - coeff p n" |
|
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59983
diff
changeset
|
646 |
proof - |
| 60679 | 647 |
fix p :: "'a poly" |
648 |
show "\<forall>\<^sub>\<infinity>n. - coeff p n = 0" |
|
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59983
diff
changeset
|
649 |
using MOST_coeff_eq_0 by simp |
| 52380 | 650 |
qed |
| 29451 | 651 |
|
652 |
lemma coeff_minus [simp]: "coeff (- p) n = - coeff p n" |
|
| 52380 | 653 |
by (simp add: uminus_poly.rep_eq) |
| 29451 | 654 |
|
| 60679 | 655 |
instance |
656 |
proof |
|
| 29451 | 657 |
fix p q :: "'a poly" |
658 |
show "- p + p = 0" |
|
| 52380 | 659 |
by (simp add: poly_eq_iff) |
| 29451 | 660 |
show "p - q = p + - q" |
|
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
52380
diff
changeset
|
661 |
by (simp add: poly_eq_iff) |
| 29451 | 662 |
qed |
663 |
||
664 |
end |
|
665 |
||
| 65346 | 666 |
lemma add_pCons [simp]: "pCons a p + pCons b q = pCons (a + b) (p + q)" |
667 |
by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) |
|
668 |
||
669 |
lemma minus_pCons [simp]: "- pCons a p = pCons (- a) (- p)" |
|
670 |
by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) |
|
671 |
||
672 |
lemma diff_pCons [simp]: "pCons a p - pCons b q = pCons (a - b) (p - q)" |
|
673 |
by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) |
|
| 29451 | 674 |
|
| 29539 | 675 |
lemma degree_add_le_max: "degree (p + q) \<le> max (degree p) (degree q)" |
| 65346 | 676 |
by (rule degree_le) (auto simp add: coeff_eq_0) |
677 |
||
678 |
lemma degree_add_le: "degree p \<le> n \<Longrightarrow> degree q \<le> n \<Longrightarrow> degree (p + q) \<le> n" |
|
| 29539 | 679 |
by (auto intro: order_trans degree_add_le_max) |
680 |
||
| 65346 | 681 |
lemma degree_add_less: "degree p < n \<Longrightarrow> degree q < n \<Longrightarrow> degree (p + q) < n" |
| 29539 | 682 |
by (auto intro: le_less_trans degree_add_le_max) |
| 29453 | 683 |
|
| 65346 | 684 |
lemma degree_add_eq_right: "degree p < degree q \<Longrightarrow> degree (p + q) = degree q" |
685 |
apply (cases "q = 0") |
|
686 |
apply simp |
|
| 29451 | 687 |
apply (rule order_antisym) |
| 65346 | 688 |
apply (simp add: degree_add_le) |
| 29451 | 689 |
apply (rule le_degree) |
690 |
apply (simp add: coeff_eq_0) |
|
691 |
done |
|
692 |
||
| 65346 | 693 |
lemma degree_add_eq_left: "degree q < degree p \<Longrightarrow> degree (p + q) = degree p" |
694 |
using degree_add_eq_right [of q p] by (simp add: add.commute) |
|
695 |
||
696 |
lemma degree_minus [simp]: "degree (- p) = degree p" |
|
697 |
by (simp add: degree_def) |
|
698 |
||
699 |
lemma lead_coeff_add_le: "degree p < degree q \<Longrightarrow> lead_coeff (p + q) = lead_coeff q" |
|
| 64795 | 700 |
by (metis coeff_add coeff_eq_0 monoid_add_class.add.left_neutral degree_add_eq_right) |
701 |
||
| 65346 | 702 |
lemma lead_coeff_minus: "lead_coeff (- p) = - lead_coeff p" |
| 64795 | 703 |
by (metis coeff_minus degree_minus) |
704 |
||
| 65346 | 705 |
lemma degree_diff_le_max: "degree (p - q) \<le> max (degree p) (degree q)" |
706 |
for p q :: "'a::ab_group_add poly" |
|
707 |
using degree_add_le [where p=p and q="-q"] by simp |
|
708 |
||
709 |
lemma degree_diff_le: "degree p \<le> n \<Longrightarrow> degree q \<le> n \<Longrightarrow> degree (p - q) \<le> n" |
|
710 |
for p q :: "'a::ab_group_add poly" |
|
711 |
using degree_add_le [of p n "- q"] by simp |
|
712 |
||
713 |
lemma degree_diff_less: "degree p < n \<Longrightarrow> degree q < n \<Longrightarrow> degree (p - q) < n" |
|
714 |
for p q :: "'a::ab_group_add poly" |
|
715 |
using degree_add_less [of p n "- q"] by simp |
|
| 29453 | 716 |
|
| 29451 | 717 |
lemma add_monom: "monom a n + monom b n = monom (a + b) n" |
| 52380 | 718 |
by (rule poly_eqI) simp |
| 29451 | 719 |
|
720 |
lemma diff_monom: "monom a n - monom b n = monom (a - b) n" |
|
| 52380 | 721 |
by (rule poly_eqI) simp |
| 29451 | 722 |
|
| 65346 | 723 |
lemma minus_monom: "- monom a n = monom (- a) n" |
| 52380 | 724 |
by (rule poly_eqI) simp |
| 29451 | 725 |
|
| 64267 | 726 |
lemma coeff_sum: "coeff (\<Sum>x\<in>A. p x) i = (\<Sum>x\<in>A. coeff (p x) i)" |
| 65346 | 727 |
by (induct A rule: infinite_finite_induct) simp_all |
| 29451 | 728 |
|
| 64267 | 729 |
lemma monom_sum: "monom (\<Sum>x\<in>A. a x) n = (\<Sum>x\<in>A. monom (a x) n)" |
730 |
by (rule poly_eqI) (simp add: coeff_sum) |
|
| 52380 | 731 |
|
732 |
fun plus_coeffs :: "'a::comm_monoid_add list \<Rightarrow> 'a list \<Rightarrow> 'a list" |
|
| 65346 | 733 |
where |
734 |
"plus_coeffs xs [] = xs" |
|
735 |
| "plus_coeffs [] ys = ys" |
|
736 |
| "plus_coeffs (x # xs) (y # ys) = (x + y) ## plus_coeffs xs ys" |
|
| 52380 | 737 |
|
738 |
lemma coeffs_plus_eq_plus_coeffs [code abstract]: |
|
739 |
"coeffs (p + q) = plus_coeffs (coeffs p) (coeffs q)" |
|
740 |
proof - |
|
| 65346 | 741 |
have *: "nth_default 0 (plus_coeffs xs ys) n = nth_default 0 xs n + nth_default 0 ys n" |
742 |
for xs ys :: "'a list" and n |
|
743 |
proof (induct xs ys arbitrary: n rule: plus_coeffs.induct) |
|
| 65390 | 744 |
case (3 x xs y ys n) |
745 |
then show ?case |
|
746 |
by (cases n) (auto simp add: cCons_def) |
|
| 65346 | 747 |
qed simp_all |
| 65390 | 748 |
have **: "no_trailing (HOL.eq 0) (plus_coeffs xs ys)" |
749 |
if "no_trailing (HOL.eq 0) xs" and "no_trailing (HOL.eq 0) ys" |
|
750 |
for xs ys :: "'a list" |
|
751 |
using that by (induct xs ys rule: plus_coeffs.induct) (simp_all add: cCons_def) |
|
| 52380 | 752 |
show ?thesis |
| 65390 | 753 |
by (rule coeffs_eqI) (auto simp add: * nth_default_coeffs_eq intro: **) |
| 52380 | 754 |
qed |
755 |
||
| 65390 | 756 |
lemma coeffs_uminus [code abstract]: |
757 |
"coeffs (- p) = map uminus (coeffs p)" |
|
758 |
proof - |
|
759 |
have eq_0: "HOL.eq 0 \<circ> uminus = HOL.eq (0::'a)" |
|
760 |
by (simp add: fun_eq_iff) |
|
761 |
show ?thesis |
|
762 |
by (rule coeffs_eqI) (simp_all add: nth_default_map_eq nth_default_coeffs_eq no_trailing_map eq_0) |
|
763 |
qed |
|
| 52380 | 764 |
|
| 65346 | 765 |
lemma [code]: "p - q = p + - q" |
766 |
for p q :: "'a::ab_group_add poly" |
|
| 59557 | 767 |
by (fact diff_conv_add_uminus) |
| 52380 | 768 |
|
769 |
lemma poly_add [simp]: "poly (p + q) x = poly p x + poly q x" |
|
| 65346 | 770 |
apply (induct p arbitrary: q) |
771 |
apply simp |
|
| 52380 | 772 |
apply (case_tac q, simp, simp add: algebra_simps) |
773 |
done |
|
774 |
||
| 65346 | 775 |
lemma poly_minus [simp]: "poly (- p) x = - poly p x" |
776 |
for x :: "'a::comm_ring" |
|
| 52380 | 777 |
by (induct p) simp_all |
778 |
||
| 65346 | 779 |
lemma poly_diff [simp]: "poly (p - q) x = poly p x - poly q x" |
780 |
for x :: "'a::comm_ring" |
|
|
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
52380
diff
changeset
|
781 |
using poly_add [of p "- q" x] by simp |
| 52380 | 782 |
|
| 64267 | 783 |
lemma poly_sum: "poly (\<Sum>k\<in>A. p k) x = (\<Sum>k\<in>A. poly (p k) x)" |
| 52380 | 784 |
by (induct A rule: infinite_finite_induct) simp_all |
| 29451 | 785 |
|
| 65346 | 786 |
lemma degree_sum_le: "finite S \<Longrightarrow> (\<And>p. p \<in> S \<Longrightarrow> degree (f p) \<le> n) \<Longrightarrow> degree (sum f S) \<le> n" |
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
787 |
proof (induct S rule: finite_induct) |
| 65346 | 788 |
case empty |
789 |
then show ?case by simp |
|
790 |
next |
|
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
791 |
case (insert p S) |
| 65346 | 792 |
then have "degree (sum f S) \<le> n" "degree (f p) \<le> n" |
793 |
by auto |
|
794 |
then show ?case |
|
795 |
unfolding sum.insert[OF insert(1-2)] by (metis degree_add_le) |
|
796 |
qed |
|
797 |
||
798 |
lemma poly_as_sum_of_monoms': |
|
799 |
assumes "degree p \<le> n" |
|
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
800 |
shows "(\<Sum>i\<le>n. monom (coeff p i) i) = p" |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
801 |
proof - |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
802 |
have eq: "\<And>i. {..n} \<inter> {i} = (if i \<le> n then {i} else {})"
|
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
803 |
by auto |
| 65346 | 804 |
from assms show ?thesis |
805 |
by (simp add: poly_eq_iff coeff_sum coeff_eq_0 sum.If_cases eq |
|
806 |
if_distrib[where f="\<lambda>x. x * a" for a]) |
|
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
807 |
qed |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
808 |
|
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
809 |
lemma poly_as_sum_of_monoms: "(\<Sum>i\<le>degree p. monom (coeff p i) i) = p" |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
810 |
by (intro poly_as_sum_of_monoms' order_refl) |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
811 |
|
| 62065 | 812 |
lemma Poly_snoc: "Poly (xs @ [x]) = Poly xs + monom x (length xs)" |
| 65346 | 813 |
by (induct xs) (simp_all add: monom_0 monom_Suc) |
| 62065 | 814 |
|
| 29451 | 815 |
|
| 60500 | 816 |
subsection \<open>Multiplication by a constant, polynomial multiplication and the unit polynomial\<close> |
| 29451 | 817 |
|
| 52380 | 818 |
lift_definition smult :: "'a::comm_semiring_0 \<Rightarrow> 'a poly \<Rightarrow> 'a poly" |
819 |
is "\<lambda>a p n. a * coeff p n" |
|
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59983
diff
changeset
|
820 |
proof - |
| 65346 | 821 |
fix a :: 'a and p :: "'a poly" |
822 |
show "\<forall>\<^sub>\<infinity> i. a * coeff p i = 0" |
|
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59983
diff
changeset
|
823 |
using MOST_coeff_eq_0[of p] by eventually_elim simp |
| 52380 | 824 |
qed |
| 29451 | 825 |
|
| 65346 | 826 |
lemma coeff_smult [simp]: "coeff (smult a p) n = a * coeff p n" |
| 52380 | 827 |
by (simp add: smult.rep_eq) |
| 29451 | 828 |
|
829 |
lemma degree_smult_le: "degree (smult a p) \<le> degree p" |
|
| 65346 | 830 |
by (rule degree_le) (simp add: coeff_eq_0) |
| 29451 | 831 |
|
| 29472 | 832 |
lemma smult_smult [simp]: "smult a (smult b p) = smult (a * b) p" |
| 65346 | 833 |
by (rule poly_eqI) (simp add: mult.assoc) |
| 29451 | 834 |
|
835 |
lemma smult_0_right [simp]: "smult a 0 = 0" |
|
| 65346 | 836 |
by (rule poly_eqI) simp |
| 29451 | 837 |
|
838 |
lemma smult_0_left [simp]: "smult 0 p = 0" |
|
| 65346 | 839 |
by (rule poly_eqI) simp |
| 29451 | 840 |
|
841 |
lemma smult_1_left [simp]: "smult (1::'a::comm_semiring_1) p = p" |
|
| 65346 | 842 |
by (rule poly_eqI) simp |
843 |
||
844 |
lemma smult_add_right: "smult a (p + q) = smult a p + smult a q" |
|
845 |
by (rule poly_eqI) (simp add: algebra_simps) |
|
846 |
||
847 |
lemma smult_add_left: "smult (a + b) p = smult a p + smult b p" |
|
848 |
by (rule poly_eqI) (simp add: algebra_simps) |
|
849 |
||
850 |
lemma smult_minus_right [simp]: "smult a (- p) = - smult a p" |
|
851 |
for a :: "'a::comm_ring" |
|
852 |
by (rule poly_eqI) simp |
|
853 |
||
854 |
lemma smult_minus_left [simp]: "smult (- a) p = - smult a p" |
|
855 |
for a :: "'a::comm_ring" |
|
856 |
by (rule poly_eqI) simp |
|
857 |
||
858 |
lemma smult_diff_right: "smult a (p - q) = smult a p - smult a q" |
|
859 |
for a :: "'a::comm_ring" |
|
860 |
by (rule poly_eqI) (simp add: algebra_simps) |
|
861 |
||
862 |
lemma smult_diff_left: "smult (a - b) p = smult a p - smult b p" |
|
863 |
for a b :: "'a::comm_ring" |
|
864 |
by (rule poly_eqI) (simp add: algebra_simps) |
|
| 29451 | 865 |
|
| 29472 | 866 |
lemmas smult_distribs = |
867 |
smult_add_left smult_add_right |
|
868 |
smult_diff_left smult_diff_right |
|
869 |
||
| 65346 | 870 |
lemma smult_pCons [simp]: "smult a (pCons b p) = pCons (a * b) (smult a p)" |
871 |
by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) |
|
| 29451 | 872 |
|
873 |
lemma smult_monom: "smult a (monom b n) = monom (a * b) n" |
|
| 65346 | 874 |
by (induct n) (simp_all add: monom_0 monom_Suc) |
| 29451 | 875 |
|
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
876 |
lemma smult_Poly: "smult c (Poly xs) = Poly (map (op * c) xs)" |
| 65346 | 877 |
by (auto simp: poly_eq_iff nth_default_def) |
878 |
||
879 |
lemma degree_smult_eq [simp]: "degree (smult a p) = (if a = 0 then 0 else degree p)" |
|
880 |
for a :: "'a::{comm_semiring_0,semiring_no_zero_divisors}"
|
|
881 |
by (cases "a = 0") (simp_all add: degree_def) |
|
882 |
||
883 |
lemma smult_eq_0_iff [simp]: "smult a p = 0 \<longleftrightarrow> a = 0 \<or> p = 0" |
|
884 |
for a :: "'a::{comm_semiring_0,semiring_no_zero_divisors}"
|
|
| 52380 | 885 |
by (simp add: poly_eq_iff) |
| 65346 | 886 |
|
| 52380 | 887 |
lemma coeffs_smult [code abstract]: |
| 65346 | 888 |
"coeffs (smult a p) = (if a = 0 then [] else map (Groups.times a) (coeffs p))" |
889 |
for p :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
|
|
| 65390 | 890 |
proof - |
891 |
have eq_0: "HOL.eq 0 \<circ> times a = HOL.eq (0::'a)" if "a \<noteq> 0" |
|
892 |
using that by (simp add: fun_eq_iff) |
|
893 |
show ?thesis |
|
894 |
by (rule coeffs_eqI) (auto simp add: no_trailing_map nth_default_map_eq nth_default_coeffs_eq eq_0) |
|
895 |
qed |
|
| 64795 | 896 |
|
897 |
lemma smult_eq_iff: |
|
| 65346 | 898 |
fixes b :: "'a :: field" |
899 |
assumes "b \<noteq> 0" |
|
900 |
shows "smult a p = smult b q \<longleftrightarrow> smult (a / b) p = q" |
|
901 |
(is "?lhs \<longleftrightarrow> ?rhs") |
|
| 64795 | 902 |
proof |
| 65346 | 903 |
assume ?lhs |
904 |
also from assms have "smult (inverse b) \<dots> = q" |
|
905 |
by simp |
|
906 |
finally show ?rhs |
|
907 |
by (simp add: field_simps) |
|
908 |
next |
|
909 |
assume ?rhs |
|
910 |
with assms show ?lhs by auto |
|
911 |
qed |
|
| 64795 | 912 |
|
| 29451 | 913 |
instantiation poly :: (comm_semiring_0) comm_semiring_0 |
914 |
begin |
|
915 |
||
| 65346 | 916 |
definition "p * q = fold_coeffs (\<lambda>a p. smult a q + pCons 0 p) p 0" |
| 29474 | 917 |
|
918 |
lemma mult_poly_0_left: "(0::'a poly) * q = 0" |
|
| 52380 | 919 |
by (simp add: times_poly_def) |
| 29474 | 920 |
|
| 65346 | 921 |
lemma mult_pCons_left [simp]: "pCons a p * q = smult a q + pCons 0 (p * q)" |
| 52380 | 922 |
by (cases "p = 0 \<and> a = 0") (auto simp add: times_poly_def) |
| 29474 | 923 |
|
924 |
lemma mult_poly_0_right: "p * (0::'a poly) = 0" |
|
| 65346 | 925 |
by (induct p) (simp_all add: mult_poly_0_left) |
926 |
||
927 |
lemma mult_pCons_right [simp]: "p * pCons a q = smult a p + pCons 0 (p * q)" |
|
928 |
by (induct p) (simp_all add: mult_poly_0_left algebra_simps) |
|
| 29474 | 929 |
|
930 |
lemmas mult_poly_0 = mult_poly_0_left mult_poly_0_right |
|
931 |
||
| 65346 | 932 |
lemma mult_smult_left [simp]: "smult a p * q = smult a (p * q)" |
933 |
by (induct p) (simp_all add: mult_poly_0 smult_add_right) |
|
934 |
||
935 |
lemma mult_smult_right [simp]: "p * smult a q = smult a (p * q)" |
|
936 |
by (induct q) (simp_all add: mult_poly_0 smult_add_right) |
|
937 |
||
938 |
lemma mult_poly_add_left: "(p + q) * r = p * r + q * r" |
|
939 |
for p q r :: "'a poly" |
|
940 |
by (induct r) (simp_all add: mult_poly_0 smult_distribs algebra_simps) |
|
| 29451 | 941 |
|
| 60679 | 942 |
instance |
943 |
proof |
|
| 29451 | 944 |
fix p q r :: "'a poly" |
945 |
show 0: "0 * p = 0" |
|
| 29474 | 946 |
by (rule mult_poly_0_left) |
| 29451 | 947 |
show "p * 0 = 0" |
| 29474 | 948 |
by (rule mult_poly_0_right) |
| 29451 | 949 |
show "(p + q) * r = p * r + q * r" |
| 29474 | 950 |
by (rule mult_poly_add_left) |
| 29451 | 951 |
show "(p * q) * r = p * (q * r)" |
| 65346 | 952 |
by (induct p) (simp_all add: mult_poly_0 mult_poly_add_left) |
| 29451 | 953 |
show "p * q = q * p" |
| 65346 | 954 |
by (induct p) (simp_all add: mult_poly_0) |
| 29451 | 955 |
qed |
956 |
||
957 |
end |
|
958 |
||
| 63498 | 959 |
lemma coeff_mult_degree_sum: |
| 65346 | 960 |
"coeff (p * q) (degree p + degree q) = coeff p (degree p) * coeff q (degree q)" |
961 |
by (induct p) (simp_all add: coeff_eq_0) |
|
| 63498 | 962 |
|
963 |
instance poly :: ("{comm_semiring_0,semiring_no_zero_divisors}") semiring_no_zero_divisors
|
|
964 |
proof |
|
965 |
fix p q :: "'a poly" |
|
966 |
assume "p \<noteq> 0" and "q \<noteq> 0" |
|
| 65346 | 967 |
have "coeff (p * q) (degree p + degree q) = coeff p (degree p) * coeff q (degree q)" |
| 63498 | 968 |
by (rule coeff_mult_degree_sum) |
| 65346 | 969 |
also from \<open>p \<noteq> 0\<close> \<open>q \<noteq> 0\<close> have "coeff p (degree p) * coeff q (degree q) \<noteq> 0" |
970 |
by simp |
|
| 63498 | 971 |
finally have "\<exists>n. coeff (p * q) n \<noteq> 0" .. |
| 65346 | 972 |
then show "p * q \<noteq> 0" |
973 |
by (simp add: poly_eq_iff) |
|
| 63498 | 974 |
qed |
975 |
||
| 29540 | 976 |
instance poly :: (comm_semiring_0_cancel) comm_semiring_0_cancel .. |
977 |
||
| 65346 | 978 |
lemma coeff_mult: "coeff (p * q) n = (\<Sum>i\<le>n. coeff p i * coeff q (n-i))" |
| 29474 | 979 |
proof (induct p arbitrary: n) |
| 65346 | 980 |
case 0 |
981 |
show ?case by simp |
|
| 29474 | 982 |
next |
| 65346 | 983 |
case (pCons a p n) |
984 |
then show ?case |
|
985 |
by (cases n) (simp_all add: sum_atMost_Suc_shift del: sum_atMost_Suc) |
|
| 29474 | 986 |
qed |
| 29451 | 987 |
|
| 29474 | 988 |
lemma degree_mult_le: "degree (p * q) \<le> degree p + degree q" |
| 65346 | 989 |
apply (rule degree_le) |
990 |
apply (induct p) |
|
991 |
apply simp |
|
992 |
apply (simp add: coeff_eq_0 coeff_pCons split: nat.split) |
|
993 |
done |
|
| 29451 | 994 |
|
995 |
lemma mult_monom: "monom a m * monom b n = monom (a * b) (m + n)" |
|
| 60679 | 996 |
by (induct m) (simp add: monom_0 smult_monom, simp add: monom_Suc) |
| 29451 | 997 |
|
998 |
instantiation poly :: (comm_semiring_1) comm_semiring_1 |
|
999 |
begin |
|
1000 |
||
| 65486 | 1001 |
lift_definition one_poly :: "'a poly" |
1002 |
is "\<lambda>n. of_bool (n = 0)" |
|
1003 |
by (rule MOST_SucD) simp |
|
1004 |
||
1005 |
lemma coeff_1 [simp]: |
|
1006 |
"coeff 1 n = of_bool (n = 0)" |
|
1007 |
by (simp add: one_poly.rep_eq) |
|
1008 |
||
1009 |
lemma one_pCons: |
|
1010 |
"1 = [:1:]" |
|
1011 |
by (simp add: poly_eq_iff coeff_pCons split: nat.splits) |
|
1012 |
||
1013 |
lemma pCons_one: |
|
1014 |
"[:1:] = 1" |
|
1015 |
by (simp add: one_pCons) |
|
| 29451 | 1016 |
|
| 60679 | 1017 |
instance |
| 65486 | 1018 |
by standard (simp_all add: one_pCons) |
| 29451 | 1019 |
|
1020 |
end |
|
1021 |
||
| 65486 | 1022 |
lemma poly_1 [simp]: |
1023 |
"poly 1 x = 1" |
|
1024 |
by (simp add: one_pCons) |
|
1025 |
||
1026 |
lemma one_poly_eq_simps [simp]: |
|
1027 |
"1 = [:1:] \<longleftrightarrow> True" |
|
1028 |
"[:1:] = 1 \<longleftrightarrow> True" |
|
1029 |
by (simp_all add: one_pCons) |
|
1030 |
||
1031 |
lemma degree_1 [simp]: |
|
1032 |
"degree 1 = 0" |
|
1033 |
by (simp add: one_pCons) |
|
1034 |
||
1035 |
lemma coeffs_1_eq [simp, code abstract]: |
|
1036 |
"coeffs 1 = [1]" |
|
1037 |
by (simp add: one_pCons) |
|
1038 |
||
1039 |
lemma smult_one [simp]: |
|
1040 |
"smult c 1 = [:c:]" |
|
1041 |
by (simp add: one_pCons) |
|
1042 |
||
1043 |
lemma monom_eq_1 [simp]: |
|
1044 |
"monom 1 0 = 1" |
|
1045 |
by (simp add: monom_0 one_pCons) |
|
1046 |
||
1047 |
lemma monom_eq_1_iff: |
|
1048 |
"monom c n = 1 \<longleftrightarrow> c = 1 \<and> n = 0" |
|
1049 |
using monom_eq_const_iff [of c n 1] by auto |
|
1050 |
||
1051 |
lemma monom_altdef: |
|
1052 |
"monom c n = smult c ([:0, 1:] ^ n)" |
|
1053 |
by (induct n) (simp_all add: monom_0 monom_Suc) |
|
1054 |
||
| 63498 | 1055 |
instance poly :: ("{comm_semiring_1,semiring_1_no_zero_divisors}") semiring_1_no_zero_divisors ..
|
| 52380 | 1056 |
instance poly :: (comm_ring) comm_ring .. |
1057 |
instance poly :: (comm_ring_1) comm_ring_1 .. |
|
| 63498 | 1058 |
instance poly :: (comm_ring_1) comm_semiring_1_cancel .. |
1059 |
||
| 65346 | 1060 |
lemma degree_power_le: "degree (p ^ n) \<le> degree p * n" |
| 52380 | 1061 |
by (induct n) (auto intro: order_trans degree_mult_le) |
1062 |
||
| 65346 | 1063 |
lemma coeff_0_power: "coeff (p ^ n) 0 = coeff p 0 ^ n" |
1064 |
by (induct n) (simp_all add: coeff_mult) |
|
1065 |
||
1066 |
lemma poly_smult [simp]: "poly (smult a p) x = a * poly p x" |
|
1067 |
by (induct p) (simp_all add: algebra_simps) |
|
1068 |
||
1069 |
lemma poly_mult [simp]: "poly (p * q) x = poly p x * poly q x" |
|
1070 |
by (induct p) (simp_all add: algebra_simps) |
|
1071 |
||
1072 |
lemma poly_power [simp]: "poly (p ^ n) x = poly p x ^ n" |
|
1073 |
for p :: "'a::comm_semiring_1 poly" |
|
| 52380 | 1074 |
by (induct n) simp_all |
1075 |
||
| 64272 | 1076 |
lemma poly_prod: "poly (\<Prod>k\<in>A. p k) x = (\<Prod>k\<in>A. poly (p k) x)" |
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1077 |
by (induct A rule: infinite_finite_induct) simp_all |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1078 |
|
| 64272 | 1079 |
lemma degree_prod_sum_le: "finite S \<Longrightarrow> degree (prod f S) \<le> sum (degree o f) S" |
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1080 |
proof (induct S rule: finite_induct) |
| 65346 | 1081 |
case empty |
1082 |
then show ?case by simp |
|
1083 |
next |
|
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1084 |
case (insert a S) |
| 65346 | 1085 |
show ?case |
1086 |
unfolding prod.insert[OF insert(1-2)] sum.insert[OF insert(1-2)] |
|
1087 |
by (rule le_trans[OF degree_mult_le]) (use insert in auto) |
|
1088 |
qed |
|
1089 |
||
1090 |
lemma coeff_0_prod_list: "coeff (prod_list xs) 0 = prod_list (map (\<lambda>p. coeff p 0) xs)" |
|
1091 |
by (induct xs) (simp_all add: coeff_mult) |
|
1092 |
||
1093 |
lemma coeff_monom_mult: "coeff (monom c n * p) k = (if k < n then 0 else c * coeff p (k - n))" |
|
| 64795 | 1094 |
proof - |
1095 |
have "coeff (monom c n * p) k = (\<Sum>i\<le>k. (if n = i then c else 0) * coeff p (k - i))" |
|
1096 |
by (simp add: coeff_mult) |
|
1097 |
also have "\<dots> = (\<Sum>i\<le>k. (if n = i then c * coeff p (k - i) else 0))" |
|
1098 |
by (intro sum.cong) simp_all |
|
| 65346 | 1099 |
also have "\<dots> = (if k < n then 0 else c * coeff p (k - n))" |
1100 |
by (simp add: sum.delta') |
|
| 64795 | 1101 |
finally show ?thesis . |
1102 |
qed |
|
1103 |
||
| 65346 | 1104 |
lemma monom_1_dvd_iff': "monom 1 n dvd p \<longleftrightarrow> (\<forall>k<n. coeff p k = 0)" |
| 64795 | 1105 |
proof |
1106 |
assume "monom 1 n dvd p" |
|
| 65346 | 1107 |
then obtain r where "p = monom 1 n * r" |
1108 |
by (rule dvdE) |
|
1109 |
then show "\<forall>k<n. coeff p k = 0" |
|
1110 |
by (simp add: coeff_mult) |
|
| 64795 | 1111 |
next |
1112 |
assume zero: "(\<forall>k<n. coeff p k = 0)" |
|
1113 |
define r where "r = Abs_poly (\<lambda>k. coeff p (k + n))" |
|
1114 |
have "\<forall>\<^sub>\<infinity>k. coeff p (k + n) = 0" |
|
| 65346 | 1115 |
by (subst cofinite_eq_sequentially, subst eventually_sequentially_seg, |
| 64795 | 1116 |
subst cofinite_eq_sequentially [symmetric]) transfer |
| 65346 | 1117 |
then have coeff_r [simp]: "coeff r k = coeff p (k + n)" for k |
1118 |
unfolding r_def by (subst poly.Abs_poly_inverse) simp_all |
|
| 64795 | 1119 |
have "p = monom 1 n * r" |
| 65346 | 1120 |
by (rule poly_eqI, subst coeff_monom_mult) (simp_all add: zero) |
1121 |
then show "monom 1 n dvd p" by simp |
|
| 64795 | 1122 |
qed |
1123 |
||
|
64591
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1124 |
|
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1125 |
subsection \<open>Mapping polynomials\<close> |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1126 |
|
| 65346 | 1127 |
definition map_poly :: "('a :: zero \<Rightarrow> 'b :: zero) \<Rightarrow> 'a poly \<Rightarrow> 'b poly"
|
1128 |
where "map_poly f p = Poly (map f (coeffs p))" |
|
|
64591
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1129 |
|
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1130 |
lemma map_poly_0 [simp]: "map_poly f 0 = 0" |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1131 |
by (simp add: map_poly_def) |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1132 |
|
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1133 |
lemma map_poly_1: "map_poly f 1 = [:f 1:]" |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1134 |
by (simp add: map_poly_def) |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1135 |
|
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1136 |
lemma map_poly_1' [simp]: "f 1 = 1 \<Longrightarrow> map_poly f 1 = 1" |
| 65486 | 1137 |
by (simp add: map_poly_def one_pCons) |
|
64591
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1138 |
|
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1139 |
lemma coeff_map_poly: |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1140 |
assumes "f 0 = 0" |
| 65346 | 1141 |
shows "coeff (map_poly f p) n = f (coeff p n)" |
1142 |
by (auto simp: assms map_poly_def nth_default_def coeffs_def not_less Suc_le_eq coeff_eq_0 |
|
1143 |
simp del: upt_Suc) |
|
1144 |
||
1145 |
lemma coeffs_map_poly [code abstract]: |
|
1146 |
"coeffs (map_poly f p) = strip_while (op = 0) (map f (coeffs p))" |
|
|
64591
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1147 |
by (simp add: map_poly_def) |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1148 |
|
| 65346 | 1149 |
lemma coeffs_map_poly': |
1150 |
assumes "\<And>x. x \<noteq> 0 \<Longrightarrow> f x \<noteq> 0" |
|
1151 |
shows "coeffs (map_poly f p) = map f (coeffs p)" |
|
| 65390 | 1152 |
using assms by (simp add: coeffs_map_poly no_trailing_map strip_while_idem_iff) |
1153 |
(metis comp_def no_leading_def no_trailing_coeffs) |
|
1154 |
||
1155 |
lemma set_coeffs_map_poly: |
|
1156 |
"(\<And>x. f x = 0 \<longleftrightarrow> x = 0) \<Longrightarrow> set (coeffs (map_poly f p)) = f ` set (coeffs p)" |
|
1157 |
by (simp add: coeffs_map_poly') |
|
|
64591
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1158 |
|
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1159 |
lemma degree_map_poly: |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1160 |
assumes "\<And>x. x \<noteq> 0 \<Longrightarrow> f x \<noteq> 0" |
| 65346 | 1161 |
shows "degree (map_poly f p) = degree p" |
|
64591
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1162 |
by (simp add: degree_eq_length_coeffs coeffs_map_poly' assms) |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1163 |
|
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1164 |
lemma map_poly_eq_0_iff: |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1165 |
assumes "f 0 = 0" "\<And>x. x \<in> set (coeffs p) \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> f x \<noteq> 0" |
| 65346 | 1166 |
shows "map_poly f p = 0 \<longleftrightarrow> p = 0" |
|
64591
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1167 |
proof - |
| 65346 | 1168 |
have "(coeff (map_poly f p) n = 0) = (coeff p n = 0)" for n |
1169 |
proof - |
|
1170 |
have "coeff (map_poly f p) n = f (coeff p n)" |
|
1171 |
by (simp add: coeff_map_poly assms) |
|
|
64591
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1172 |
also have "\<dots> = 0 \<longleftrightarrow> coeff p n = 0" |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1173 |
proof (cases "n < length (coeffs p)") |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1174 |
case True |
| 65346 | 1175 |
then have "coeff p n \<in> set (coeffs p)" |
1176 |
by (auto simp: coeffs_def simp del: upt_Suc) |
|
1177 |
with assms show "f (coeff p n) = 0 \<longleftrightarrow> coeff p n = 0" |
|
1178 |
by auto |
|
1179 |
next |
|
1180 |
case False |
|
1181 |
then show ?thesis |
|
1182 |
by (auto simp: assms length_coeffs nth_default_coeffs_eq [symmetric] nth_default_def) |
|
1183 |
qed |
|
1184 |
finally show ?thesis . |
|
1185 |
qed |
|
1186 |
then show ?thesis by (auto simp: poly_eq_iff) |
|
|
64591
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1187 |
qed |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1188 |
|
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1189 |
lemma map_poly_smult: |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1190 |
assumes "f 0 = 0""\<And>c x. f (c * x) = f c * f x" |
| 65346 | 1191 |
shows "map_poly f (smult c p) = smult (f c) (map_poly f p)" |
|
64591
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1192 |
by (intro poly_eqI) (simp_all add: assms coeff_map_poly) |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1193 |
|
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1194 |
lemma map_poly_pCons: |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1195 |
assumes "f 0 = 0" |
| 65346 | 1196 |
shows "map_poly f (pCons c p) = pCons (f c) (map_poly f p)" |
|
64591
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1197 |
by (intro poly_eqI) (simp_all add: assms coeff_map_poly coeff_pCons split: nat.splits) |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1198 |
|
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1199 |
lemma map_poly_map_poly: |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1200 |
assumes "f 0 = 0" "g 0 = 0" |
| 65346 | 1201 |
shows "map_poly f (map_poly g p) = map_poly (f \<circ> g) p" |
|
64591
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1202 |
by (intro poly_eqI) (simp add: coeff_map_poly assms) |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1203 |
|
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1204 |
lemma map_poly_id [simp]: "map_poly id p = p" |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1205 |
by (simp add: map_poly_def) |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1206 |
|
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1207 |
lemma map_poly_id' [simp]: "map_poly (\<lambda>x. x) p = p" |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1208 |
by (simp add: map_poly_def) |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1209 |
|
| 65346 | 1210 |
lemma map_poly_cong: |
|
64591
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1211 |
assumes "(\<And>x. x \<in> set (coeffs p) \<Longrightarrow> f x = g x)" |
| 65346 | 1212 |
shows "map_poly f p = map_poly g p" |
|
64591
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1213 |
proof - |
| 65346 | 1214 |
from assms have "map f (coeffs p) = map g (coeffs p)" |
1215 |
by (intro map_cong) simp_all |
|
1216 |
then show ?thesis |
|
1217 |
by (simp only: coeffs_eq_iff coeffs_map_poly) |
|
|
64591
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1218 |
qed |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1219 |
|
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1220 |
lemma map_poly_monom: "f 0 = 0 \<Longrightarrow> map_poly f (monom c n) = monom (f c) n" |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1221 |
by (intro poly_eqI) (simp_all add: coeff_map_poly) |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1222 |
|
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1223 |
lemma map_poly_idI: |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1224 |
assumes "\<And>x. x \<in> set (coeffs p) \<Longrightarrow> f x = x" |
| 65346 | 1225 |
shows "map_poly f p = p" |
|
64591
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1226 |
using map_poly_cong[OF assms, of _ id] by simp |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1227 |
|
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1228 |
lemma map_poly_idI': |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1229 |
assumes "\<And>x. x \<in> set (coeffs p) \<Longrightarrow> f x = x" |
| 65346 | 1230 |
shows "p = map_poly f p" |
|
64591
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1231 |
using map_poly_cong[OF assms, of _ id] by simp |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1232 |
|
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1233 |
lemma smult_conv_map_poly: "smult c p = map_poly (\<lambda>x. c * x) p" |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1234 |
by (intro poly_eqI) (simp_all add: coeff_map_poly) |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1235 |
|
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1236 |
|
| 65484 | 1237 |
subsection \<open>Conversions\<close> |
1238 |
||
1239 |
lemma of_nat_poly: |
|
1240 |
"of_nat n = [:of_nat n:]" |
|
| 65486 | 1241 |
by (induct n) (simp_all add: one_pCons) |
| 65484 | 1242 |
|
1243 |
lemma of_nat_monom: |
|
1244 |
"of_nat n = monom (of_nat n) 0" |
|
1245 |
by (simp add: of_nat_poly monom_0) |
|
1246 |
||
1247 |
lemma degree_of_nat [simp]: |
|
1248 |
"degree (of_nat n) = 0" |
|
| 62065 | 1249 |
by (simp add: of_nat_poly) |
1250 |
||
| 64795 | 1251 |
lemma lead_coeff_of_nat [simp]: |
| 65484 | 1252 |
"lead_coeff (of_nat n) = of_nat n" |
| 64795 | 1253 |
by (simp add: of_nat_poly) |
1254 |
||
| 65484 | 1255 |
lemma of_int_poly: |
1256 |
"of_int k = [:of_int k:]" |
|
| 64793 | 1257 |
by (simp only: of_int_of_nat of_nat_poly) simp |
1258 |
||
| 65484 | 1259 |
lemma of_int_monom: |
1260 |
"of_int k = monom (of_int k) 0" |
|
1261 |
by (simp add: of_int_poly monom_0) |
|
1262 |
||
1263 |
lemma degree_of_int [simp]: |
|
1264 |
"degree (of_int k) = 0" |
|
| 64795 | 1265 |
by (simp add: of_int_poly) |
1266 |
||
1267 |
lemma lead_coeff_of_int [simp]: |
|
| 65484 | 1268 |
"lead_coeff (of_int k) = of_int k" |
| 64793 | 1269 |
by (simp add: of_int_poly) |
| 62065 | 1270 |
|
1271 |
lemma numeral_poly: "numeral n = [:numeral n:]" |
|
| 65484 | 1272 |
proof - |
1273 |
have "numeral n = of_nat (numeral n)" |
|
1274 |
by simp |
|
1275 |
also have "\<dots> = [:of_nat (numeral n):]" |
|
1276 |
by (simp add: of_nat_poly) |
|
1277 |
finally show ?thesis |
|
1278 |
by simp |
|
1279 |
qed |
|
1280 |
||
1281 |
lemma numeral_monom: |
|
1282 |
"numeral n = monom (numeral n) 0" |
|
1283 |
by (simp add: numeral_poly monom_0) |
|
1284 |
||
1285 |
lemma degree_numeral [simp]: |
|
1286 |
"degree (numeral n) = 0" |
|
1287 |
by (simp add: numeral_poly) |
|
| 52380 | 1288 |
|
| 65346 | 1289 |
lemma lead_coeff_numeral [simp]: |
| 64795 | 1290 |
"lead_coeff (numeral n) = numeral n" |
1291 |
by (simp add: numeral_poly) |
|
1292 |
||
|
64591
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1293 |
|
| 60500 | 1294 |
subsection \<open>Lemmas about divisibility\<close> |
| 29979 | 1295 |
|
| 65346 | 1296 |
lemma dvd_smult: |
1297 |
assumes "p dvd q" |
|
1298 |
shows "p dvd smult a q" |
|
| 29979 | 1299 |
proof - |
| 65346 | 1300 |
from assms obtain k where "q = p * k" .. |
| 29979 | 1301 |
then have "smult a q = p * smult a k" by simp |
1302 |
then show "p dvd smult a q" .. |
|
1303 |
qed |
|
1304 |
||
| 65346 | 1305 |
lemma dvd_smult_cancel: "p dvd smult a q \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> p dvd q" |
1306 |
for a :: "'a::field" |
|
| 29979 | 1307 |
by (drule dvd_smult [where a="inverse a"]) simp |
1308 |
||
| 65346 | 1309 |
lemma dvd_smult_iff: "a \<noteq> 0 \<Longrightarrow> p dvd smult a q \<longleftrightarrow> p dvd q" |
1310 |
for a :: "'a::field" |
|
| 29979 | 1311 |
by (safe elim!: dvd_smult dvd_smult_cancel) |
1312 |
||
| 31663 | 1313 |
lemma smult_dvd_cancel: |
| 65346 | 1314 |
assumes "smult a p dvd q" |
1315 |
shows "p dvd q" |
|
| 31663 | 1316 |
proof - |
| 65346 | 1317 |
from assms obtain k where "q = smult a p * k" .. |
| 31663 | 1318 |
then have "q = p * smult a k" by simp |
1319 |
then show "p dvd q" .. |
|
1320 |
qed |
|
1321 |
||
| 65346 | 1322 |
lemma smult_dvd: "p dvd q \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> smult a p dvd q" |
1323 |
for a :: "'a::field" |
|
| 31663 | 1324 |
by (rule smult_dvd_cancel [where a="inverse a"]) simp |
1325 |
||
| 65346 | 1326 |
lemma smult_dvd_iff: "smult a p dvd q \<longleftrightarrow> (if a = 0 then q = 0 else p dvd q)" |
1327 |
for a :: "'a::field" |
|
| 31663 | 1328 |
by (auto elim: smult_dvd smult_dvd_cancel) |
1329 |
||
| 64795 | 1330 |
lemma is_unit_smult_iff: "smult c p dvd 1 \<longleftrightarrow> c dvd 1 \<and> p dvd 1" |
1331 |
proof - |
|
1332 |
have "smult c p = [:c:] * p" by simp |
|
1333 |
also have "\<dots> dvd 1 \<longleftrightarrow> c dvd 1 \<and> p dvd 1" |
|
1334 |
proof safe |
|
| 65346 | 1335 |
assume *: "[:c:] * p dvd 1" |
1336 |
then show "p dvd 1" |
|
1337 |
by (rule dvd_mult_right) |
|
1338 |
from * obtain q where q: "1 = [:c:] * p * q" |
|
1339 |
by (rule dvdE) |
|
1340 |
have "c dvd c * (coeff p 0 * coeff q 0)" |
|
1341 |
by simp |
|
1342 |
also have "\<dots> = coeff ([:c:] * p * q) 0" |
|
1343 |
by (simp add: mult.assoc coeff_mult) |
|
1344 |
also note q [symmetric] |
|
1345 |
finally have "c dvd coeff 1 0" . |
|
1346 |
then show "c dvd 1" by simp |
|
| 64795 | 1347 |
next |
1348 |
assume "c dvd 1" "p dvd 1" |
|
| 65346 | 1349 |
from this(1) obtain d where "1 = c * d" |
1350 |
by (rule dvdE) |
|
1351 |
then have "1 = [:c:] * [:d:]" |
|
| 65486 | 1352 |
by (simp add: one_pCons ac_simps) |
| 65346 | 1353 |
then have "[:c:] dvd 1" |
1354 |
by (rule dvdI) |
|
1355 |
from mult_dvd_mono[OF this \<open>p dvd 1\<close>] show "[:c:] * p dvd 1" |
|
1356 |
by simp |
|
| 64795 | 1357 |
qed |
1358 |
finally show ?thesis . |
|
1359 |
qed |
|
1360 |
||
| 29451 | 1361 |
|
| 60500 | 1362 |
subsection \<open>Polynomials form an integral domain\<close> |
| 29451 | 1363 |
|
| 63498 | 1364 |
instance poly :: (idom) idom .. |
| 29451 | 1365 |
|
|
65577
32d4117ad6e8
instance for polynomial rings with characteristic zero
haftmann
parents:
65486
diff
changeset
|
1366 |
instance poly :: ("{ring_char_0, comm_ring_1}") ring_char_0
|
|
32d4117ad6e8
instance for polynomial rings with characteristic zero
haftmann
parents:
65486
diff
changeset
|
1367 |
by standard (auto simp add: of_nat_poly intro: injI) |
|
32d4117ad6e8
instance for polynomial rings with characteristic zero
haftmann
parents:
65486
diff
changeset
|
1368 |
|
| 65346 | 1369 |
lemma degree_mult_eq: "p \<noteq> 0 \<Longrightarrow> q \<noteq> 0 \<Longrightarrow> degree (p * q) = degree p + degree q" |
1370 |
for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
|
|
1371 |
by (rule order_antisym [OF degree_mult_le le_degree]) (simp add: coeff_mult_degree_sum) |
|
| 29451 | 1372 |
|
|
64591
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1373 |
lemma degree_mult_eq_0: |
| 65346 | 1374 |
"degree (p * q) = 0 \<longleftrightarrow> p = 0 \<or> q = 0 \<or> (p \<noteq> 0 \<and> q \<noteq> 0 \<and> degree p = 0 \<and> degree q = 0)" |
1375 |
for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
|
|
1376 |
by (auto simp: degree_mult_eq) |
|
|
64591
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
1377 |
|
| 60570 | 1378 |
lemma degree_mult_right_le: |
| 63498 | 1379 |
fixes p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
|
| 60570 | 1380 |
assumes "q \<noteq> 0" |
1381 |
shows "degree p \<le> degree (p * q)" |
|
1382 |
using assms by (cases "p = 0") (simp_all add: degree_mult_eq) |
|
1383 |
||
| 65346 | 1384 |
lemma coeff_degree_mult: "coeff (p * q) (degree (p * q)) = coeff q (degree q) * coeff p (degree p)" |
1385 |
for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
|
|
1386 |
by (cases "p = 0 \<or> q = 0") (auto simp: degree_mult_eq coeff_mult_degree_sum mult_ac) |
|
1387 |
||
1388 |
lemma dvd_imp_degree_le: "p dvd q \<Longrightarrow> q \<noteq> 0 \<Longrightarrow> degree p \<le> degree q" |
|
1389 |
for p q :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly"
|
|
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1390 |
by (erule dvdE, hypsubst, subst degree_mult_eq) auto |
| 29451 | 1391 |
|
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1392 |
lemma divides_degree: |
| 65346 | 1393 |
fixes p q :: "'a ::{comm_semiring_1,semiring_no_zero_divisors} poly"
|
1394 |
assumes "p dvd q" |
|
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1395 |
shows "degree p \<le> degree q \<or> q = 0" |
| 65346 | 1396 |
by (metis dvd_imp_degree_le assms) |
1397 |
||
| 63498 | 1398 |
lemma const_poly_dvd_iff: |
| 65346 | 1399 |
fixes c :: "'a::{comm_semiring_1,semiring_no_zero_divisors}"
|
| 63498 | 1400 |
shows "[:c:] dvd p \<longleftrightarrow> (\<forall>n. c dvd coeff p n)" |
1401 |
proof (cases "c = 0 \<or> p = 0") |
|
| 65346 | 1402 |
case True |
1403 |
then show ?thesis |
|
1404 |
by (auto intro!: poly_eqI) |
|
1405 |
next |
|
| 63498 | 1406 |
case False |
1407 |
show ?thesis |
|
1408 |
proof |
|
1409 |
assume "[:c:] dvd p" |
|
| 65346 | 1410 |
then show "\<forall>n. c dvd coeff p n" |
1411 |
by (auto elim!: dvdE simp: coeffs_def) |
|
| 63498 | 1412 |
next |
1413 |
assume *: "\<forall>n. c dvd coeff p n" |
|
| 65346 | 1414 |
define mydiv where "mydiv x y = (SOME z. x = y * z)" for x y :: 'a |
| 63498 | 1415 |
have mydiv: "x = y * mydiv x y" if "y dvd x" for x y |
1416 |
using that unfolding mydiv_def dvd_def by (rule someI_ex) |
|
1417 |
define q where "q = Poly (map (\<lambda>a. mydiv a c) (coeffs p))" |
|
1418 |
from False * have "p = q * [:c:]" |
|
| 65346 | 1419 |
by (intro poly_eqI) |
1420 |
(auto simp: q_def nth_default_def not_less length_coeffs_degree coeffs_nth |
|
1421 |
intro!: coeff_eq_0 mydiv) |
|
1422 |
then show "[:c:] dvd p" |
|
1423 |
by (simp only: dvd_triv_right) |
|
| 63498 | 1424 |
qed |
| 65346 | 1425 |
qed |
1426 |
||
1427 |
lemma const_poly_dvd_const_poly_iff [simp]: "[:a:] dvd [:b:] \<longleftrightarrow> a dvd b" |
|
1428 |
for a b :: "'a::{comm_semiring_1,semiring_no_zero_divisors}"
|
|
| 63498 | 1429 |
by (subst const_poly_dvd_iff) (auto simp: coeff_pCons split: nat.splits) |
1430 |
||
| 65346 | 1431 |
lemma lead_coeff_mult: "lead_coeff (p * q) = lead_coeff p * lead_coeff q" |
1432 |
for p q :: "'a::{comm_semiring_0, semiring_no_zero_divisors} poly"
|
|
1433 |
by (cases "p = 0 \<or> q = 0") (auto simp: coeff_mult_degree_sum degree_mult_eq) |
|
1434 |
||
1435 |
lemma lead_coeff_smult: "lead_coeff (smult c p) = c * lead_coeff p" |
|
1436 |
for p :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
|
|
| 64795 | 1437 |
proof - |
1438 |
have "smult c p = [:c:] * p" by simp |
|
1439 |
also have "lead_coeff \<dots> = c * lead_coeff p" |
|
1440 |
by (subst lead_coeff_mult) simp_all |
|
1441 |
finally show ?thesis . |
|
1442 |
qed |
|
1443 |
||
1444 |
lemma lead_coeff_1 [simp]: "lead_coeff 1 = 1" |
|
1445 |
by simp |
|
1446 |
||
| 65346 | 1447 |
lemma lead_coeff_power: "lead_coeff (p ^ n) = lead_coeff p ^ n" |
1448 |
for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly"
|
|
1449 |
by (induct n) (simp_all add: lead_coeff_mult) |
|
| 64795 | 1450 |
|
| 29451 | 1451 |
|
| 60500 | 1452 |
subsection \<open>Polynomials form an ordered integral domain\<close> |
| 29878 | 1453 |
|
| 63498 | 1454 |
definition pos_poly :: "'a::linordered_semidom poly \<Rightarrow> bool" |
| 65346 | 1455 |
where "pos_poly p \<longleftrightarrow> 0 < coeff p (degree p)" |
1456 |
||
1457 |
lemma pos_poly_pCons: "pos_poly (pCons a p) \<longleftrightarrow> pos_poly p \<or> (p = 0 \<and> 0 < a)" |
|
1458 |
by (simp add: pos_poly_def) |
|
| 29878 | 1459 |
|
1460 |
lemma not_pos_poly_0 [simp]: "\<not> pos_poly 0" |
|
| 65346 | 1461 |
by (simp add: pos_poly_def) |
1462 |
||
1463 |
lemma pos_poly_add: "pos_poly p \<Longrightarrow> pos_poly q \<Longrightarrow> pos_poly (p + q)" |
|
1464 |
apply (induct p arbitrary: q) |
|
1465 |
apply simp |
|
1466 |
apply (case_tac q) |
|
1467 |
apply (force simp add: pos_poly_pCons add_pos_pos) |
|
| 29878 | 1468 |
done |
1469 |
||
| 65346 | 1470 |
lemma pos_poly_mult: "pos_poly p \<Longrightarrow> pos_poly q \<Longrightarrow> pos_poly (p * q)" |
| 29878 | 1471 |
unfolding pos_poly_def |
1472 |
apply (subgoal_tac "p \<noteq> 0 \<and> q \<noteq> 0") |
|
| 65346 | 1473 |
apply (simp add: degree_mult_eq coeff_mult_degree_sum) |
| 29878 | 1474 |
apply auto |
1475 |
done |
|
1476 |
||
| 65346 | 1477 |
lemma pos_poly_total: "p = 0 \<or> pos_poly p \<or> pos_poly (- p)" |
1478 |
for p :: "'a::linordered_idom poly" |
|
1479 |
by (induct p) (auto simp: pos_poly_pCons) |
|
1480 |
||
1481 |
lemma last_coeffs_eq_coeff_degree: "p \<noteq> 0 \<Longrightarrow> last (coeffs p) = coeff p (degree p)" |
|
| 52380 | 1482 |
by (simp add: coeffs_def) |
1483 |
||
| 65346 | 1484 |
lemma pos_poly_coeffs [code]: "pos_poly p \<longleftrightarrow> (let as = coeffs p in as \<noteq> [] \<and> last as > 0)" |
1485 |
(is "?lhs \<longleftrightarrow> ?rhs") |
|
| 52380 | 1486 |
proof |
| 65346 | 1487 |
assume ?rhs |
1488 |
then show ?lhs |
|
1489 |
by (auto simp add: pos_poly_def last_coeffs_eq_coeff_degree) |
|
| 52380 | 1490 |
next |
| 65346 | 1491 |
assume ?lhs |
1492 |
then have *: "0 < coeff p (degree p)" |
|
1493 |
by (simp add: pos_poly_def) |
|
1494 |
then have "p \<noteq> 0" |
|
1495 |
by auto |
|
1496 |
with * show ?rhs |
|
1497 |
by (simp add: last_coeffs_eq_coeff_degree) |
|
| 52380 | 1498 |
qed |
1499 |
||
|
35028
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
34973
diff
changeset
|
1500 |
instantiation poly :: (linordered_idom) linordered_idom |
| 29878 | 1501 |
begin |
1502 |
||
| 65346 | 1503 |
definition "x < y \<longleftrightarrow> pos_poly (y - x)" |
1504 |
||
1505 |
definition "x \<le> y \<longleftrightarrow> x = y \<or> pos_poly (y - x)" |
|
1506 |
||
1507 |
definition "\<bar>x::'a poly\<bar> = (if x < 0 then - x else x)" |
|
1508 |
||
1509 |
definition "sgn (x::'a poly) = (if x = 0 then 0 else if 0 < x then 1 else - 1)" |
|
| 29878 | 1510 |
|
| 60679 | 1511 |
instance |
1512 |
proof |
|
1513 |
fix x y z :: "'a poly" |
|
| 29878 | 1514 |
show "x < y \<longleftrightarrow> x \<le> y \<and> \<not> y \<le> x" |
1515 |
unfolding less_eq_poly_def less_poly_def |
|
1516 |
apply safe |
|
| 65346 | 1517 |
apply simp |
| 29878 | 1518 |
apply (drule (1) pos_poly_add) |
1519 |
apply simp |
|
1520 |
done |
|
| 60679 | 1521 |
show "x \<le> x" |
| 65346 | 1522 |
by (simp add: less_eq_poly_def) |
| 60679 | 1523 |
show "x \<le> y \<Longrightarrow> y \<le> z \<Longrightarrow> x \<le> z" |
| 29878 | 1524 |
unfolding less_eq_poly_def |
1525 |
apply safe |
|
1526 |
apply (drule (1) pos_poly_add) |
|
1527 |
apply (simp add: algebra_simps) |
|
1528 |
done |
|
| 60679 | 1529 |
show "x \<le> y \<Longrightarrow> y \<le> x \<Longrightarrow> x = y" |
| 29878 | 1530 |
unfolding less_eq_poly_def |
1531 |
apply safe |
|
1532 |
apply (drule (1) pos_poly_add) |
|
1533 |
apply simp |
|
1534 |
done |
|
| 60679 | 1535 |
show "x \<le> y \<Longrightarrow> z + x \<le> z + y" |
| 29878 | 1536 |
unfolding less_eq_poly_def |
1537 |
apply safe |
|
1538 |
apply (simp add: algebra_simps) |
|
1539 |
done |
|
1540 |
show "x \<le> y \<or> y \<le> x" |
|
1541 |
unfolding less_eq_poly_def |
|
1542 |
using pos_poly_total [of "x - y"] |
|
1543 |
by auto |
|
| 60679 | 1544 |
show "x < y \<Longrightarrow> 0 < z \<Longrightarrow> z * x < z * y" |
| 65346 | 1545 |
by (simp add: less_poly_def right_diff_distrib [symmetric] pos_poly_mult) |
| 29878 | 1546 |
show "\<bar>x\<bar> = (if x < 0 then - x else x)" |
1547 |
by (rule abs_poly_def) |
|
1548 |
show "sgn x = (if x = 0 then 0 else if 0 < x then 1 else - 1)" |
|
1549 |
by (rule sgn_poly_def) |
|
1550 |
qed |
|
1551 |
||
1552 |
end |
|
1553 |
||
| 60500 | 1554 |
text \<open>TODO: Simplification rules for comparisons\<close> |
| 29878 | 1555 |
|
1556 |
||
| 60500 | 1557 |
subsection \<open>Synthetic division and polynomial roots\<close> |
| 52380 | 1558 |
|
| 65346 | 1559 |
subsubsection \<open>Synthetic division\<close> |
1560 |
||
1561 |
text \<open>Synthetic division is simply division by the linear polynomial @{term "x - c"}.\<close>
|
|
| 52380 | 1562 |
|
1563 |
definition synthetic_divmod :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a poly \<times> 'a" |
|
| 65346 | 1564 |
where "synthetic_divmod p c = fold_coeffs (\<lambda>a (q, r). (pCons r q, a + c * r)) p (0, 0)" |
| 52380 | 1565 |
|
1566 |
definition synthetic_div :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a poly" |
|
| 65346 | 1567 |
where "synthetic_div p c = fst (synthetic_divmod p c)" |
1568 |
||
1569 |
lemma synthetic_divmod_0 [simp]: "synthetic_divmod 0 c = (0, 0)" |
|
| 52380 | 1570 |
by (simp add: synthetic_divmod_def) |
1571 |
||
1572 |
lemma synthetic_divmod_pCons [simp]: |
|
1573 |
"synthetic_divmod (pCons a p) c = (\<lambda>(q, r). (pCons r q, a + c * r)) (synthetic_divmod p c)" |
|
1574 |
by (cases "p = 0 \<and> a = 0") (auto simp add: synthetic_divmod_def) |
|
1575 |
||
| 65346 | 1576 |
lemma synthetic_div_0 [simp]: "synthetic_div 0 c = 0" |
1577 |
by (simp add: synthetic_div_def) |
|
| 52380 | 1578 |
|
1579 |
lemma synthetic_div_unique_lemma: "smult c p = pCons a p \<Longrightarrow> p = 0" |
|
| 65346 | 1580 |
by (induct p arbitrary: a) simp_all |
1581 |
||
1582 |
lemma snd_synthetic_divmod: "snd (synthetic_divmod p c) = poly p c" |
|
1583 |
by (induct p) (simp_all add: split_def) |
|
| 52380 | 1584 |
|
1585 |
lemma synthetic_div_pCons [simp]: |
|
1586 |
"synthetic_div (pCons a p) c = pCons (poly p c) (synthetic_div p c)" |
|
| 65346 | 1587 |
by (simp add: synthetic_div_def split_def snd_synthetic_divmod) |
1588 |
||
1589 |
lemma synthetic_div_eq_0_iff: "synthetic_div p c = 0 \<longleftrightarrow> degree p = 0" |
|
| 63649 | 1590 |
proof (induct p) |
1591 |
case 0 |
|
1592 |
then show ?case by simp |
|
1593 |
next |
|
1594 |
case (pCons a p) |
|
1595 |
then show ?case by (cases p) simp |
|
1596 |
qed |
|
| 52380 | 1597 |
|
| 65346 | 1598 |
lemma degree_synthetic_div: "degree (synthetic_div p c) = degree p - 1" |
| 63649 | 1599 |
by (induct p) (simp_all add: synthetic_div_eq_0_iff) |
| 52380 | 1600 |
|
1601 |
lemma synthetic_div_correct: |
|
1602 |
"p + smult c (synthetic_div p c) = pCons (poly p c) (synthetic_div p c)" |
|
1603 |
by (induct p) simp_all |
|
1604 |
||
| 65346 | 1605 |
lemma synthetic_div_unique: "p + smult c q = pCons r q \<Longrightarrow> r = poly p c \<and> q = synthetic_div p c" |
1606 |
apply (induct p arbitrary: q r) |
|
1607 |
apply simp |
|
1608 |
apply (frule synthetic_div_unique_lemma) |
|
1609 |
apply simp |
|
1610 |
apply (case_tac q, force) |
|
1611 |
done |
|
1612 |
||
1613 |
lemma synthetic_div_correct': "[:-c, 1:] * synthetic_div p c + [:poly p c:] = p" |
|
1614 |
for c :: "'a::comm_ring_1" |
|
1615 |
using synthetic_div_correct [of p c] by (simp add: algebra_simps) |
|
1616 |
||
1617 |
||
| 64795 | 1618 |
subsubsection \<open>Polynomial roots\<close> |
| 65346 | 1619 |
|
1620 |
lemma poly_eq_0_iff_dvd: "poly p c = 0 \<longleftrightarrow> [:- c, 1:] dvd p" |
|
1621 |
(is "?lhs \<longleftrightarrow> ?rhs") |
|
1622 |
for c :: "'a::comm_ring_1" |
|
| 52380 | 1623 |
proof |
| 65346 | 1624 |
assume ?lhs |
1625 |
with synthetic_div_correct' [of c p] have "p = [:-c, 1:] * synthetic_div p c" by simp |
|
1626 |
then show ?rhs .. |
|
| 52380 | 1627 |
next |
| 65346 | 1628 |
assume ?rhs |
| 52380 | 1629 |
then obtain k where "p = [:-c, 1:] * k" by (rule dvdE) |
| 65346 | 1630 |
then show ?lhs by simp |
| 52380 | 1631 |
qed |
1632 |
||
| 65346 | 1633 |
lemma dvd_iff_poly_eq_0: "[:c, 1:] dvd p \<longleftrightarrow> poly p (- c) = 0" |
1634 |
for c :: "'a::comm_ring_1" |
|
| 52380 | 1635 |
by (simp add: poly_eq_0_iff_dvd) |
1636 |
||
| 65346 | 1637 |
lemma poly_roots_finite: "p \<noteq> 0 \<Longrightarrow> finite {x. poly p x = 0}"
|
1638 |
for p :: "'a::{comm_ring_1,ring_no_zero_divisors} poly"
|
|
| 52380 | 1639 |
proof (induct n \<equiv> "degree p" arbitrary: p) |
| 65346 | 1640 |
case 0 |
| 52380 | 1641 |
then obtain a where "a \<noteq> 0" and "p = [:a:]" |
| 65346 | 1642 |
by (cases p) (simp split: if_splits) |
1643 |
then show "finite {x. poly p x = 0}"
|
|
1644 |
by simp |
|
| 52380 | 1645 |
next |
| 65346 | 1646 |
case (Suc n) |
| 52380 | 1647 |
show "finite {x. poly p x = 0}"
|
1648 |
proof (cases "\<exists>x. poly p x = 0") |
|
1649 |
case False |
|
1650 |
then show "finite {x. poly p x = 0}" by simp
|
|
1651 |
next |
|
1652 |
case True |
|
1653 |
then obtain a where "poly p a = 0" .. |
|
| 65346 | 1654 |
then have "[:-a, 1:] dvd p" |
1655 |
by (simp only: poly_eq_0_iff_dvd) |
|
| 52380 | 1656 |
then obtain k where k: "p = [:-a, 1:] * k" .. |
| 65346 | 1657 |
with \<open>p \<noteq> 0\<close> have "k \<noteq> 0" |
1658 |
by auto |
|
| 52380 | 1659 |
with k have "degree p = Suc (degree k)" |
1660 |
by (simp add: degree_mult_eq del: mult_pCons_left) |
|
| 65346 | 1661 |
with \<open>Suc n = degree p\<close> have "n = degree k" |
1662 |
by simp |
|
1663 |
from this \<open>k \<noteq> 0\<close> have "finite {x. poly k x = 0}"
|
|
1664 |
by (rule Suc.hyps) |
|
1665 |
then have "finite (insert a {x. poly k x = 0})"
|
|
1666 |
by simp |
|
| 52380 | 1667 |
then show "finite {x. poly p x = 0}"
|
| 57862 | 1668 |
by (simp add: k Collect_disj_eq del: mult_pCons_left) |
| 52380 | 1669 |
qed |
1670 |
qed |
|
1671 |
||
| 65346 | 1672 |
lemma poly_eq_poly_eq_iff: "poly p = poly q \<longleftrightarrow> p = q" |
1673 |
(is "?lhs \<longleftrightarrow> ?rhs") |
|
1674 |
for p q :: "'a::{comm_ring_1,ring_no_zero_divisors,ring_char_0} poly"
|
|
| 52380 | 1675 |
proof |
| 65346 | 1676 |
assume ?rhs |
1677 |
then show ?lhs by simp |
|
| 52380 | 1678 |
next |
| 65346 | 1679 |
assume ?lhs |
1680 |
have "poly p = poly 0 \<longleftrightarrow> p = 0" for p :: "'a poly" |
|
1681 |
apply (cases "p = 0") |
|
1682 |
apply simp_all |
|
1683 |
apply (drule poly_roots_finite) |
|
1684 |
apply (auto simp add: infinite_UNIV_char_0) |
|
1685 |
done |
|
1686 |
from \<open>?lhs\<close> and this [of "p - q"] show ?rhs |
|
1687 |
by auto |
|
| 52380 | 1688 |
qed |
1689 |
||
| 65346 | 1690 |
lemma poly_all_0_iff_0: "(\<forall>x. poly p x = 0) \<longleftrightarrow> p = 0" |
1691 |
for p :: "'a::{ring_char_0,comm_ring_1,ring_no_zero_divisors} poly"
|
|
| 52380 | 1692 |
by (auto simp add: poly_eq_poly_eq_iff [symmetric]) |
1693 |
||
| 65346 | 1694 |
|
| 64795 | 1695 |
subsubsection \<open>Order of polynomial roots\<close> |
|
29977
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1696 |
|
| 52380 | 1697 |
definition order :: "'a::idom \<Rightarrow> 'a poly \<Rightarrow> nat" |
| 65346 | 1698 |
where "order a p = (LEAST n. \<not> [:-a, 1:] ^ Suc n dvd p)" |
1699 |
||
1700 |
lemma coeff_linear_power: "coeff ([:a, 1:] ^ n) n = 1" |
|
1701 |
for a :: "'a::comm_semiring_1" |
|
1702 |
apply (induct n) |
|
1703 |
apply simp_all |
|
1704 |
apply (subst coeff_eq_0) |
|
1705 |
apply (auto intro: le_less_trans degree_power_le) |
|
1706 |
done |
|
1707 |
||
1708 |
lemma degree_linear_power: "degree ([:a, 1:] ^ n) = n" |
|
1709 |
for a :: "'a::comm_semiring_1" |
|
1710 |
apply (rule order_antisym) |
|
1711 |
apply (rule ord_le_eq_trans [OF degree_power_le]) |
|
1712 |
apply simp |
|
1713 |
apply (rule le_degree) |
|
1714 |
apply (simp add: coeff_linear_power) |
|
1715 |
done |
|
|
29977
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1716 |
|
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1717 |
lemma order_1: "[:-a, 1:] ^ order a p dvd p" |
| 65346 | 1718 |
apply (cases "p = 0") |
1719 |
apply simp |
|
1720 |
apply (cases "order a p") |
|
1721 |
apply simp |
|
1722 |
apply (subgoal_tac "nat < (LEAST n. \<not> [:-a, 1:] ^ Suc n dvd p)") |
|
1723 |
apply (drule not_less_Least) |
|
1724 |
apply simp |
|
1725 |
apply (fold order_def) |
|
1726 |
apply simp |
|
1727 |
done |
|
|
29977
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1728 |
|
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1729 |
lemma order_2: "p \<noteq> 0 \<Longrightarrow> \<not> [:-a, 1:] ^ Suc (order a p) dvd p" |
| 65346 | 1730 |
unfolding order_def |
1731 |
apply (rule LeastI_ex) |
|
1732 |
apply (rule_tac x="degree p" in exI) |
|
1733 |
apply (rule notI) |
|
1734 |
apply (drule (1) dvd_imp_degree_le) |
|
1735 |
apply (simp only: degree_linear_power) |
|
1736 |
done |
|
1737 |
||
1738 |
lemma order: "p \<noteq> 0 \<Longrightarrow> [:-a, 1:] ^ order a p dvd p \<and> \<not> [:-a, 1:] ^ Suc (order a p) dvd p" |
|
1739 |
by (rule conjI [OF order_1 order_2]) |
|
|
29977
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1740 |
|
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1741 |
lemma order_degree: |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1742 |
assumes p: "p \<noteq> 0" |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1743 |
shows "order a p \<le> degree p" |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1744 |
proof - |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1745 |
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
|
1746 |
by (simp only: degree_linear_power) |
| 65346 | 1747 |
also from order_1 p have "\<dots> \<le> degree p" |
1748 |
by (rule dvd_imp_degree_le) |
|
|
29977
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1749 |
finally show ?thesis . |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1750 |
qed |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1751 |
|
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1752 |
lemma order_root: "poly p a = 0 \<longleftrightarrow> p = 0 \<or> order a p \<noteq> 0" |
| 65346 | 1753 |
apply (cases "p = 0") |
1754 |
apply simp_all |
|
1755 |
apply (rule iffI) |
|
1756 |
apply (metis order_2 not_gr0 poly_eq_0_iff_dvd power_0 power_Suc_0 power_one_right) |
|
1757 |
unfolding poly_eq_0_iff_dvd |
|
1758 |
apply (metis dvd_power dvd_trans order_1) |
|
1759 |
done |
|
|
29977
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1760 |
|
| 62065 | 1761 |
lemma order_0I: "poly p a \<noteq> 0 \<Longrightarrow> order a p = 0" |
1762 |
by (subst (asm) order_root) auto |
|
1763 |
||
| 64795 | 1764 |
lemma order_unique_lemma: |
1765 |
fixes p :: "'a::idom poly" |
|
1766 |
assumes "[:-a, 1:] ^ n dvd p" "\<not> [:-a, 1:] ^ Suc n dvd p" |
|
1767 |
shows "n = order a p" |
|
| 65346 | 1768 |
unfolding Polynomial.order_def |
1769 |
apply (rule Least_equality [symmetric]) |
|
1770 |
apply (fact assms) |
|
1771 |
apply (rule classical) |
|
1772 |
apply (erule notE) |
|
1773 |
unfolding not_less_eq_eq |
|
1774 |
using assms(1) |
|
1775 |
apply (rule power_le_dvd) |
|
1776 |
apply assumption |
|
| 64795 | 1777 |
done |
| 65346 | 1778 |
|
| 64795 | 1779 |
lemma order_mult: "p * q \<noteq> 0 \<Longrightarrow> order a (p * q) = order a p + order a q" |
1780 |
proof - |
|
1781 |
define i where "i = order a p" |
|
1782 |
define j where "j = order a q" |
|
1783 |
define t where "t = [:-a, 1:]" |
|
1784 |
have t_dvd_iff: "\<And>u. t dvd u \<longleftrightarrow> poly u a = 0" |
|
| 65346 | 1785 |
by (simp add: t_def dvd_iff_poly_eq_0) |
| 64795 | 1786 |
assume "p * q \<noteq> 0" |
1787 |
then show "order a (p * q) = i + j" |
|
1788 |
apply clarsimp |
|
1789 |
apply (drule order [where a=a and p=p, folded i_def t_def]) |
|
1790 |
apply (drule order [where a=a and p=q, folded j_def t_def]) |
|
1791 |
apply clarify |
|
1792 |
apply (erule dvdE)+ |
|
1793 |
apply (rule order_unique_lemma [symmetric], fold t_def) |
|
| 65346 | 1794 |
apply (simp_all add: power_add t_dvd_iff) |
| 64795 | 1795 |
done |
1796 |
qed |
|
1797 |
||
1798 |
lemma order_smult: |
|
| 65346 | 1799 |
assumes "c \<noteq> 0" |
| 64795 | 1800 |
shows "order x (smult c p) = order x p" |
1801 |
proof (cases "p = 0") |
|
| 65346 | 1802 |
case True |
1803 |
then show ?thesis |
|
1804 |
by simp |
|
1805 |
next |
|
| 64795 | 1806 |
case False |
1807 |
have "smult c p = [:c:] * p" by simp |
|
| 65346 | 1808 |
also from assms False have "order x \<dots> = order x [:c:] + order x p" |
| 64795 | 1809 |
by (subst order_mult) simp_all |
| 65346 | 1810 |
also have "order x [:c:] = 0" |
1811 |
by (rule order_0I) (use assms in auto) |
|
1812 |
finally show ?thesis |
|
1813 |
by simp |
|
1814 |
qed |
|
| 64795 | 1815 |
|
1816 |
(* Next two lemmas contributed by Wenda Li *) |
|
| 65346 | 1817 |
lemma order_1_eq_0 [simp]:"order x 1 = 0" |
| 64795 | 1818 |
by (metis order_root poly_1 zero_neq_one) |
1819 |
||
| 65346 | 1820 |
lemma order_power_n_n: "order a ([:-a,1:]^n)=n" |
| 64795 | 1821 |
proof (induct n) (*might be proved more concisely using nat_less_induct*) |
1822 |
case 0 |
|
| 65346 | 1823 |
then show ?case |
1824 |
by (metis order_root poly_1 power_0 zero_neq_one) |
|
1825 |
next |
|
| 64795 | 1826 |
case (Suc n) |
| 65346 | 1827 |
have "order a ([:- a, 1:] ^ Suc n) = order a ([:- a, 1:] ^ n) + order a [:-a,1:]" |
1828 |
by (metis (no_types, hide_lams) One_nat_def add_Suc_right monoid_add_class.add.right_neutral |
|
| 64795 | 1829 |
one_neq_zero order_mult pCons_eq_0_iff power_add power_eq_0_iff power_one_right) |
| 65346 | 1830 |
moreover have "order a [:-a,1:] = 1" |
1831 |
unfolding order_def |
|
1832 |
proof (rule Least_equality, rule notI) |
|
1833 |
assume "[:- a, 1:] ^ Suc 1 dvd [:- a, 1:]" |
|
1834 |
then have "degree ([:- a, 1:] ^ Suc 1) \<le> degree ([:- a, 1:])" |
|
1835 |
by (rule dvd_imp_degree_le) auto |
|
1836 |
then show False |
|
1837 |
by auto |
|
1838 |
next |
|
1839 |
fix y |
|
1840 |
assume *: "\<not> [:- a, 1:] ^ Suc y dvd [:- a, 1:]" |
|
1841 |
show "1 \<le> y" |
|
1842 |
proof (rule ccontr) |
|
1843 |
assume "\<not> 1 \<le> y" |
|
1844 |
then have "y = 0" by auto |
|
1845 |
then have "[:- a, 1:] ^ Suc y dvd [:- a, 1:]" by auto |
|
1846 |
with * show False by auto |
|
| 64795 | 1847 |
qed |
| 65346 | 1848 |
qed |
1849 |
ultimately show ?case |
|
1850 |
using Suc by auto |
|
| 64795 | 1851 |
qed |
1852 |
||
| 65346 | 1853 |
lemma order_0_monom [simp]: "c \<noteq> 0 \<Longrightarrow> order 0 (monom c n) = n" |
1854 |
using order_power_n_n[of 0 n] by (simp add: monom_altdef order_smult) |
|
1855 |
||
1856 |
lemma dvd_imp_order_le: "q \<noteq> 0 \<Longrightarrow> p dvd q \<Longrightarrow> Polynomial.order a p \<le> Polynomial.order a q" |
|
| 64795 | 1857 |
by (auto simp: order_mult elim: dvdE) |
1858 |
||
| 65346 | 1859 |
text \<open>Now justify the standard squarefree decomposition, i.e. \<open>f / gcd f f'\<close>.\<close> |
| 64795 | 1860 |
|
1861 |
lemma order_divides: "[:-a, 1:] ^ n dvd p \<longleftrightarrow> p = 0 \<or> n \<le> order a p" |
|
| 65346 | 1862 |
apply (cases "p = 0") |
1863 |
apply auto |
|
1864 |
apply (drule order_2 [where a=a and p=p]) |
|
1865 |
apply (metis not_less_eq_eq power_le_dvd) |
|
1866 |
apply (erule power_le_dvd [OF order_1]) |
|
1867 |
done |
|
| 64795 | 1868 |
|
1869 |
lemma order_decomp: |
|
1870 |
assumes "p \<noteq> 0" |
|
1871 |
shows "\<exists>q. p = [:- a, 1:] ^ order a p * q \<and> \<not> [:- a, 1:] dvd q" |
|
1872 |
proof - |
|
| 65346 | 1873 |
from assms have *: "[:- a, 1:] ^ order a p dvd p" |
1874 |
and **: "\<not> [:- a, 1:] ^ Suc (order a p) dvd p" |
|
1875 |
by (auto dest: order) |
|
1876 |
from * obtain q where q: "p = [:- a, 1:] ^ order a p * q" .. |
|
1877 |
with ** have "\<not> [:- a, 1:] ^ Suc (order a p) dvd [:- a, 1:] ^ order a p * q" |
|
| 64795 | 1878 |
by simp |
1879 |
then have "\<not> [:- a, 1:] ^ order a p * [:- a, 1:] dvd [:- a, 1:] ^ order a p * q" |
|
1880 |
by simp |
|
| 65346 | 1881 |
with idom_class.dvd_mult_cancel_left [of "[:- a, 1:] ^ order a p" "[:- a, 1:]" q] |
1882 |
have "\<not> [:- a, 1:] dvd q" by auto |
|
1883 |
with q show ?thesis by blast |
|
| 64795 | 1884 |
qed |
1885 |
||
| 65346 | 1886 |
lemma monom_1_dvd_iff: "p \<noteq> 0 \<Longrightarrow> monom 1 n dvd p \<longleftrightarrow> n \<le> order 0 p" |
1887 |
using order_divides[of 0 n p] by (simp add: monom_altdef) |
|
| 64795 | 1888 |
|
|
29977
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1889 |
|
| 62065 | 1890 |
subsection \<open>Additional induction rules on polynomials\<close> |
1891 |
||
1892 |
text \<open> |
|
| 65346 | 1893 |
An induction rule for induction over the roots of a polynomial with a certain property. |
| 62065 | 1894 |
(e.g. all positive roots) |
1895 |
\<close> |
|
1896 |
lemma poly_root_induct [case_names 0 no_roots root]: |
|
1897 |
fixes p :: "'a :: idom poly" |
|
1898 |
assumes "Q 0" |
|
| 65346 | 1899 |
and "\<And>p. (\<And>a. P a \<Longrightarrow> poly p a \<noteq> 0) \<Longrightarrow> Q p" |
1900 |
and "\<And>a p. P a \<Longrightarrow> Q p \<Longrightarrow> Q ([:a, -1:] * p)" |
|
1901 |
shows "Q p" |
|
| 62065 | 1902 |
proof (induction "degree p" arbitrary: p rule: less_induct) |
1903 |
case (less p) |
|
1904 |
show ?case |
|
1905 |
proof (cases "p = 0") |
|
| 65346 | 1906 |
case True |
1907 |
with assms(1) show ?thesis by simp |
|
1908 |
next |
|
1909 |
case False |
|
1910 |
show ?thesis |
|
| 62065 | 1911 |
proof (cases "\<exists>a. P a \<and> poly p a = 0") |
1912 |
case False |
|
| 65346 | 1913 |
then show ?thesis by (intro assms(2)) blast |
| 62065 | 1914 |
next |
1915 |
case True |
|
| 65346 | 1916 |
then obtain a where a: "P a" "poly p a = 0" |
| 62065 | 1917 |
by blast |
| 65346 | 1918 |
then have "-[:-a, 1:] dvd p" |
| 62065 | 1919 |
by (subst minus_dvd_iff) (simp add: poly_eq_0_iff_dvd) |
1920 |
then obtain q where q: "p = [:a, -1:] * q" by (elim dvdE) simp |
|
| 65346 | 1921 |
with False have "q \<noteq> 0" by auto |
| 62065 | 1922 |
have "degree p = Suc (degree q)" |
| 65346 | 1923 |
by (subst q, subst degree_mult_eq) (simp_all add: \<open>q \<noteq> 0\<close>) |
1924 |
then have "Q q" by (intro less) simp |
|
1925 |
with a(1) have "Q ([:a, -1:] * q)" |
|
| 62065 | 1926 |
by (rule assms(3)) |
1927 |
with q show ?thesis by simp |
|
1928 |
qed |
|
| 65346 | 1929 |
qed |
| 62065 | 1930 |
qed |
1931 |
||
| 65346 | 1932 |
lemma dropWhile_replicate_append: |
1933 |
"dropWhile (op = a) (replicate n a @ ys) = dropWhile (op = a) ys" |
|
1934 |
by (induct n) simp_all |
|
| 62065 | 1935 |
|
1936 |
lemma Poly_append_replicate_0: "Poly (xs @ replicate n 0) = Poly xs" |
|
1937 |
by (subst coeffs_eq_iff) (simp_all add: strip_while_def dropWhile_replicate_append) |
|
1938 |
||
1939 |
text \<open> |
|
| 65346 | 1940 |
An induction rule for simultaneous induction over two polynomials, |
| 62065 | 1941 |
prepending one coefficient in each step. |
1942 |
\<close> |
|
1943 |
lemma poly_induct2 [case_names 0 pCons]: |
|
1944 |
assumes "P 0 0" "\<And>a p b q. P p q \<Longrightarrow> P (pCons a p) (pCons b q)" |
|
| 65346 | 1945 |
shows "P p q" |
| 62065 | 1946 |
proof - |
| 63040 | 1947 |
define n where "n = max (length (coeffs p)) (length (coeffs q))" |
1948 |
define xs where "xs = coeffs p @ (replicate (n - length (coeffs p)) 0)" |
|
1949 |
define ys where "ys = coeffs q @ (replicate (n - length (coeffs q)) 0)" |
|
| 65346 | 1950 |
have "length xs = length ys" |
| 62065 | 1951 |
by (simp add: xs_def ys_def n_def) |
| 65346 | 1952 |
then have "P (Poly xs) (Poly ys)" |
1953 |
by (induct rule: list_induct2) (simp_all add: assms) |
|
1954 |
also have "Poly xs = p" |
|
| 62065 | 1955 |
by (simp add: xs_def Poly_append_replicate_0) |
| 65346 | 1956 |
also have "Poly ys = q" |
| 62065 | 1957 |
by (simp add: ys_def Poly_append_replicate_0) |
1958 |
finally show ?thesis . |
|
1959 |
qed |
|
1960 |
||
| 65346 | 1961 |
|
| 60500 | 1962 |
subsection \<open>Composition of polynomials\<close> |
| 29478 | 1963 |
|
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1964 |
(* Several lemmas contributed by René Thiemann and Akihisa Yamada *) |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1965 |
|
| 52380 | 1966 |
definition pcompose :: "'a::comm_semiring_0 poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" |
| 65346 | 1967 |
where "pcompose p q = fold_coeffs (\<lambda>a c. [:a:] + q * c) p 0" |
| 52380 | 1968 |
|
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1969 |
notation pcompose (infixl "\<circ>\<^sub>p" 71) |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1970 |
|
| 65346 | 1971 |
lemma pcompose_0 [simp]: "pcompose 0 q = 0" |
| 52380 | 1972 |
by (simp add: pcompose_def) |
| 65346 | 1973 |
|
1974 |
lemma pcompose_pCons: "pcompose (pCons a p) q = [:a:] + q * pcompose p q" |
|
| 52380 | 1975 |
by (cases "p = 0 \<and> a = 0") (auto simp add: pcompose_def) |
1976 |
||
| 65346 | 1977 |
lemma pcompose_1: "pcompose 1 p = 1" |
1978 |
for p :: "'a::comm_semiring_1 poly" |
|
| 65486 | 1979 |
by (auto simp: one_pCons pcompose_pCons) |
| 65346 | 1980 |
|
1981 |
lemma poly_pcompose: "poly (pcompose p q) x = poly p (poly q x)" |
|
| 52380 | 1982 |
by (induct p) (simp_all add: pcompose_pCons) |
1983 |
||
| 65346 | 1984 |
lemma degree_pcompose_le: "degree (pcompose p q) \<le> degree p * degree q" |
1985 |
apply (induct p) |
|
1986 |
apply simp |
|
1987 |
apply (simp add: pcompose_pCons) |
|
1988 |
apply clarify |
|
1989 |
apply (rule degree_add_le) |
|
1990 |
apply simp |
|
1991 |
apply (rule order_trans [OF degree_mult_le]) |
|
1992 |
apply simp |
|
1993 |
done |
|
1994 |
||
1995 |
lemma pcompose_add: "pcompose (p + q) r = pcompose p r + pcompose q r" |
|
1996 |
for p q r :: "'a::{comm_semiring_0, ab_semigroup_add} poly"
|
|
| 62065 | 1997 |
proof (induction p q rule: poly_induct2) |
| 65346 | 1998 |
case 0 |
1999 |
then show ?case by simp |
|
2000 |
next |
|
| 62065 | 2001 |
case (pCons a p b q) |
| 65346 | 2002 |
have "pcompose (pCons a p + pCons b q) r = [:a + b:] + r * pcompose p r + r * pcompose q r" |
| 62065 | 2003 |
by (simp_all add: pcompose_pCons pCons.IH algebra_simps) |
2004 |
also have "[:a + b:] = [:a:] + [:b:]" by simp |
|
| 65346 | 2005 |
also have "\<dots> + r * pcompose p r + r * pcompose q r = |
2006 |
pcompose (pCons a p) r + pcompose (pCons b q) r" |
|
| 62065 | 2007 |
by (simp only: pcompose_pCons add_ac) |
2008 |
finally show ?case . |
|
| 65346 | 2009 |
qed |
2010 |
||
2011 |
lemma pcompose_uminus: "pcompose (-p) r = -pcompose p r" |
|
2012 |
for p r :: "'a::comm_ring poly" |
|
2013 |
by (induct p) (simp_all add: pcompose_pCons) |
|
2014 |
||
2015 |
lemma pcompose_diff: "pcompose (p - q) r = pcompose p r - pcompose q r" |
|
2016 |
for p q r :: "'a::comm_ring poly" |
|
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
2017 |
using pcompose_add[of p "-q"] by (simp add: pcompose_uminus) |
| 62065 | 2018 |
|
| 65346 | 2019 |
lemma pcompose_smult: "pcompose (smult a p) r = smult a (pcompose p r)" |
2020 |
for p r :: "'a::comm_semiring_0 poly" |
|
2021 |
by (induct p) (simp_all add: pcompose_pCons pcompose_add smult_add_right) |
|
2022 |
||
2023 |
lemma pcompose_mult: "pcompose (p * q) r = pcompose p r * pcompose q r" |
|
2024 |
for p q r :: "'a::comm_semiring_0 poly" |
|
2025 |
by (induct p arbitrary: q) (simp_all add: pcompose_add pcompose_smult pcompose_pCons algebra_simps) |
|
2026 |
||
2027 |
lemma pcompose_assoc: "pcompose p (pcompose q r) = pcompose (pcompose p q) r" |
|
2028 |
for p q r :: "'a::comm_semiring_0 poly" |
|
2029 |
by (induct p arbitrary: q) (simp_all add: pcompose_pCons pcompose_add pcompose_mult) |
|
2030 |
||
2031 |
lemma pcompose_idR[simp]: "pcompose p [: 0, 1 :] = p" |
|
2032 |
for p :: "'a::comm_semiring_1 poly" |
|
2033 |
by (induct p) (simp_all add: pcompose_pCons) |
|
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
2034 |
|
| 64267 | 2035 |
lemma pcompose_sum: "pcompose (sum f A) p = sum (\<lambda>i. pcompose (f i) p) A" |
| 65346 | 2036 |
by (induct A rule: infinite_finite_induct) (simp_all add: pcompose_1 pcompose_add) |
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2037 |
|
| 64272 | 2038 |
lemma pcompose_prod: "pcompose (prod f A) p = prod (\<lambda>i. pcompose (f i) p) A" |
| 65346 | 2039 |
by (induct A rule: infinite_finite_induct) (simp_all add: pcompose_1 pcompose_mult) |
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2040 |
|
|
64591
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
2041 |
lemma pcompose_const [simp]: "pcompose [:a:] q = [:a:]" |
|
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
2042 |
by (subst pcompose_pCons) simp |
| 62065 | 2043 |
|
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
2044 |
lemma pcompose_0': "pcompose p 0 = [:coeff p 0:]" |
|
64591
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64272
diff
changeset
|
2045 |
by (induct p) (auto simp add: pcompose_pCons) |
| 62065 | 2046 |
|
| 65346 | 2047 |
lemma degree_pcompose: "degree (pcompose p q) = degree p * degree q" |
2048 |
for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
|
|
| 62065 | 2049 |
proof (induct p) |
2050 |
case 0 |
|
| 65346 | 2051 |
then show ?case by auto |
| 62065 | 2052 |
next |
2053 |
case (pCons a p) |
|
| 65346 | 2054 |
consider "degree (q * pcompose p q) = 0" | "degree (q * pcompose p q) > 0" |
2055 |
by blast |
|
2056 |
then show ?case |
|
2057 |
proof cases |
|
2058 |
case prems: 1 |
|
2059 |
show ?thesis |
|
2060 |
proof (cases "p = 0") |
|
| 62065 | 2061 |
case True |
| 65346 | 2062 |
then show ?thesis by auto |
| 62065 | 2063 |
next |
| 65346 | 2064 |
case False |
2065 |
from prems have "degree q = 0 \<or> pcompose p q = 0" |
|
2066 |
by (auto simp add: degree_mult_eq_0) |
|
2067 |
moreover have False if "pcompose p q = 0" "degree q \<noteq> 0" |
|
2068 |
proof - |
|
2069 |
from pCons.hyps(2) that have "degree p = 0" |
|
2070 |
by auto |
|
2071 |
then obtain a1 where "p = [:a1:]" |
|
2072 |
by (metis degree_pCons_eq_if old.nat.distinct(2) pCons_cases) |
|
2073 |
with \<open>pcompose p q = 0\<close> \<open>p \<noteq> 0\<close> show False |
|
2074 |
by auto |
|
2075 |
qed |
|
2076 |
ultimately have "degree (pCons a p) * degree q = 0" |
|
2077 |
by auto |
|
2078 |
moreover have "degree (pcompose (pCons a p) q) = 0" |
|
2079 |
proof - |
|
2080 |
from prems have "0 = max (degree [:a:]) (degree (q * pcompose p q))" |
|
2081 |
by simp |
|
2082 |
also have "\<dots> \<ge> degree ([:a:] + q * pcompose p q)" |
|
2083 |
by (rule degree_add_le_max) |
|
2084 |
finally show ?thesis |
|
2085 |
by (auto simp add: pcompose_pCons) |
|
2086 |
qed |
|
| 62065 | 2087 |
ultimately show ?thesis by simp |
2088 |
qed |
|
| 65346 | 2089 |
next |
2090 |
case prems: 2 |
|
2091 |
then have "p \<noteq> 0" "q \<noteq> 0" "pcompose p q \<noteq> 0" |
|
2092 |
by auto |
|
2093 |
from prems degree_add_eq_right [of "[:a:]"] |
|
2094 |
have "degree (pcompose (pCons a p) q) = degree (q * pcompose p q)" |
|
2095 |
by (auto simp: pcompose_pCons) |
|
2096 |
with pCons.hyps(2) degree_mult_eq[OF \<open>q\<noteq>0\<close> \<open>pcompose p q\<noteq>0\<close>] show ?thesis |
|
2097 |
by auto |
|
2098 |
qed |
|
| 62065 | 2099 |
qed |
2100 |
||
2101 |
lemma pcompose_eq_0: |
|
| 65346 | 2102 |
fixes p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
|
2103 |
assumes "pcompose p q = 0" "degree q > 0" |
|
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
2104 |
shows "p = 0" |
| 62065 | 2105 |
proof - |
| 65346 | 2106 |
from assms degree_pcompose [of p q] have "degree p = 0" |
2107 |
by auto |
|
2108 |
then obtain a where "p = [:a:]" |
|
| 62065 | 2109 |
by (metis degree_pCons_eq_if gr0_conv_Suc neq0_conv pCons_cases) |
| 65346 | 2110 |
with assms(1) have "a = 0" |
2111 |
by auto |
|
2112 |
with \<open>p = [:a:]\<close> show ?thesis |
|
2113 |
by simp |
|
| 62065 | 2114 |
qed |
2115 |
||
2116 |
lemma lead_coeff_comp: |
|
| 65346 | 2117 |
fixes p q :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly"
|
2118 |
assumes "degree q > 0" |
|
| 62065 | 2119 |
shows "lead_coeff (pcompose p q) = lead_coeff p * lead_coeff q ^ (degree p)" |
2120 |
proof (induct p) |
|
2121 |
case 0 |
|
| 65346 | 2122 |
then show ?case by auto |
| 62065 | 2123 |
next |
2124 |
case (pCons a p) |
|
| 65346 | 2125 |
consider "degree (q * pcompose p q) = 0" | "degree (q * pcompose p q) > 0" |
2126 |
by blast |
|
2127 |
then show ?case |
|
2128 |
proof cases |
|
2129 |
case prems: 1 |
|
2130 |
then have "pcompose p q = 0" |
|
2131 |
by (metis assms degree_0 degree_mult_eq_0 neq0_conv) |
|
2132 |
with pcompose_eq_0[OF _ \<open>degree q > 0\<close>] have "p = 0" |
|
2133 |
by simp |
|
2134 |
then show ?thesis |
|
2135 |
by auto |
|
2136 |
next |
|
2137 |
case prems: 2 |
|
2138 |
then have "degree [:a:] < degree (q * pcompose p q)" |
|
2139 |
by simp |
|
2140 |
then have "lead_coeff ([:a:] + q * p \<circ>\<^sub>p q) = lead_coeff (q * p \<circ>\<^sub>p q)" |
|
2141 |
by (rule lead_coeff_add_le) |
|
2142 |
then have "lead_coeff (pcompose (pCons a p) q) = lead_coeff (q * pcompose p q)" |
|
2143 |
by (simp add: pcompose_pCons) |
|
2144 |
also have "\<dots> = lead_coeff q * (lead_coeff p * lead_coeff q ^ degree p)" |
|
2145 |
using pCons.hyps(2) lead_coeff_mult[of q "pcompose p q"] by simp |
|
2146 |
also have "\<dots> = lead_coeff p * lead_coeff q ^ (degree p + 1)" |
|
2147 |
by (auto simp: mult_ac) |
|
2148 |
finally show ?thesis by auto |
|
2149 |
qed |
|
| 62065 | 2150 |
qed |
2151 |
||
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2152 |
|
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2153 |
subsection \<open>Shifting polynomials\<close> |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2154 |
|
| 65346 | 2155 |
definition poly_shift :: "nat \<Rightarrow> 'a::zero poly \<Rightarrow> 'a poly" |
2156 |
where "poly_shift n p = Abs_poly (\<lambda>i. coeff p (i + n))" |
|
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2157 |
|
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2158 |
lemma nth_default_drop: "nth_default x (drop n xs) m = nth_default x xs (m + n)" |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2159 |
by (auto simp add: nth_default_def add_ac) |
| 65346 | 2160 |
|
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2161 |
lemma nth_default_take: "nth_default x (take n xs) m = (if m < n then nth_default x xs m else x)" |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2162 |
by (auto simp add: nth_default_def add_ac) |
| 65346 | 2163 |
|
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2164 |
lemma coeff_poly_shift: "coeff (poly_shift n p) i = coeff p (i + n)" |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2165 |
proof - |
| 65346 | 2166 |
from MOST_coeff_eq_0[of p] obtain m where "\<forall>k>m. coeff p k = 0" |
2167 |
by (auto simp: MOST_nat) |
|
2168 |
then have "\<forall>k>m. coeff p (k + n) = 0" |
|
2169 |
by auto |
|
2170 |
then have "\<forall>\<^sub>\<infinity>k. coeff p (k + n) = 0" |
|
2171 |
by (auto simp: MOST_nat) |
|
2172 |
then show ?thesis |
|
2173 |
by (simp add: poly_shift_def poly.Abs_poly_inverse) |
|
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2174 |
qed |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2175 |
|
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2176 |
lemma poly_shift_id [simp]: "poly_shift 0 = (\<lambda>x. x)" |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2177 |
by (simp add: poly_eq_iff fun_eq_iff coeff_poly_shift) |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2178 |
|
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2179 |
lemma poly_shift_0 [simp]: "poly_shift n 0 = 0" |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2180 |
by (simp add: poly_eq_iff coeff_poly_shift) |
| 65346 | 2181 |
|
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2182 |
lemma poly_shift_1: "poly_shift n 1 = (if n = 0 then 1 else 0)" |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2183 |
by (simp add: poly_eq_iff coeff_poly_shift) |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2184 |
|
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2185 |
lemma poly_shift_monom: "poly_shift n (monom c m) = (if m \<ge> n then monom c (m - n) else 0)" |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2186 |
by (auto simp add: poly_eq_iff coeff_poly_shift) |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2187 |
|
| 65390 | 2188 |
lemma coeffs_shift_poly [code abstract]: |
2189 |
"coeffs (poly_shift n p) = drop n (coeffs p)" |
|
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2190 |
proof (cases "p = 0") |
| 65346 | 2191 |
case True |
2192 |
then show ?thesis by simp |
|
2193 |
next |
|
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2194 |
case False |
| 65346 | 2195 |
then show ?thesis |
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2196 |
by (intro coeffs_eqI) |
| 65390 | 2197 |
(simp_all add: coeff_poly_shift nth_default_drop nth_default_coeffs_eq) |
| 65346 | 2198 |
qed |
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2199 |
|
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2200 |
|
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2201 |
subsection \<open>Truncating polynomials\<close> |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2202 |
|
| 65346 | 2203 |
definition poly_cutoff |
2204 |
where "poly_cutoff n p = Abs_poly (\<lambda>k. if k < n then coeff p k else 0)" |
|
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2205 |
|
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2206 |
lemma coeff_poly_cutoff: "coeff (poly_cutoff n p) k = (if k < n then coeff p k else 0)" |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2207 |
unfolding poly_cutoff_def |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2208 |
by (subst poly.Abs_poly_inverse) (auto simp: MOST_nat intro: exI[of _ n]) |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2209 |
|
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2210 |
lemma poly_cutoff_0 [simp]: "poly_cutoff n 0 = 0" |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2211 |
by (simp add: poly_eq_iff coeff_poly_cutoff) |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2212 |
|
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2213 |
lemma poly_cutoff_1 [simp]: "poly_cutoff n 1 = (if n = 0 then 0 else 1)" |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2214 |
by (simp add: poly_eq_iff coeff_poly_cutoff) |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2215 |
|
| 65346 | 2216 |
lemma coeffs_poly_cutoff [code abstract]: |
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2217 |
"coeffs (poly_cutoff n p) = strip_while (op = 0) (take n (coeffs p))" |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2218 |
proof (cases "strip_while (op = 0) (take n (coeffs p)) = []") |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2219 |
case True |
| 65346 | 2220 |
then have "coeff (poly_cutoff n p) k = 0" for k |
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2221 |
unfolding coeff_poly_cutoff |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2222 |
by (auto simp: nth_default_coeffs_eq [symmetric] nth_default_def set_conv_nth) |
| 65346 | 2223 |
then have "poly_cutoff n p = 0" |
2224 |
by (simp add: poly_eq_iff) |
|
2225 |
then show ?thesis |
|
2226 |
by (subst True) simp_all |
|
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2227 |
next |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2228 |
case False |
| 65346 | 2229 |
have "no_trailing (op = 0) (strip_while (op = 0) (take n (coeffs p)))" |
2230 |
by simp |
|
2231 |
with False have "last (strip_while (op = 0) (take n (coeffs p))) \<noteq> 0" |
|
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2232 |
unfolding no_trailing_unfold by auto |
| 65346 | 2233 |
then show ?thesis |
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2234 |
by (intro coeffs_eqI) |
| 65390 | 2235 |
(simp_all add: coeff_poly_cutoff nth_default_take nth_default_coeffs_eq) |
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2236 |
qed |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2237 |
|
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2238 |
|
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2239 |
subsection \<open>Reflecting polynomials\<close> |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2240 |
|
| 65346 | 2241 |
definition reflect_poly :: "'a::zero poly \<Rightarrow> 'a poly" |
2242 |
where "reflect_poly p = Poly (rev (coeffs p))" |
|
2243 |
||
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2244 |
lemma coeffs_reflect_poly [code abstract]: |
| 65346 | 2245 |
"coeffs (reflect_poly p) = rev (dropWhile (op = 0) (coeffs p))" |
2246 |
by (simp add: reflect_poly_def) |
|
2247 |
||
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2248 |
lemma reflect_poly_0 [simp]: "reflect_poly 0 = 0" |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2249 |
by (simp add: reflect_poly_def) |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2250 |
|
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2251 |
lemma reflect_poly_1 [simp]: "reflect_poly 1 = 1" |
| 65486 | 2252 |
by (simp add: reflect_poly_def one_pCons) |
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2253 |
|
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2254 |
lemma coeff_reflect_poly: |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2255 |
"coeff (reflect_poly p) n = (if n > degree p then 0 else coeff p (degree p - n))" |
| 65346 | 2256 |
by (cases "p = 0") |
2257 |
(auto simp add: reflect_poly_def nth_default_def |
|
2258 |
rev_nth degree_eq_length_coeffs coeffs_nth not_less |
|
2259 |
dest: le_imp_less_Suc) |
|
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2260 |
|
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2261 |
lemma coeff_0_reflect_poly_0_iff [simp]: "coeff (reflect_poly p) 0 = 0 \<longleftrightarrow> p = 0" |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2262 |
by (simp add: coeff_reflect_poly) |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2263 |
|
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2264 |
lemma reflect_poly_at_0_eq_0_iff [simp]: "poly (reflect_poly p) 0 = 0 \<longleftrightarrow> p = 0" |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2265 |
by (simp add: coeff_reflect_poly poly_0_coeff_0) |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2266 |
|
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2267 |
lemma reflect_poly_pCons': |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2268 |
"p \<noteq> 0 \<Longrightarrow> reflect_poly (pCons c p) = reflect_poly p + monom c (Suc (degree p))" |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2269 |
by (intro poly_eqI) |
| 65346 | 2270 |
(auto simp: coeff_reflect_poly coeff_pCons not_less Suc_diff_le split: nat.split) |
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2271 |
|
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2272 |
lemma reflect_poly_const [simp]: "reflect_poly [:a:] = [:a:]" |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2273 |
by (cases "a = 0") (simp_all add: reflect_poly_def) |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2274 |
|
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2275 |
lemma poly_reflect_poly_nz: |
| 65346 | 2276 |
"x \<noteq> 0 \<Longrightarrow> poly (reflect_poly p) x = x ^ degree p * poly p (inverse x)" |
2277 |
for x :: "'a::field" |
|
2278 |
by (induct rule: pCons_induct) (simp_all add: field_simps reflect_poly_pCons' poly_monom) |
|
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2279 |
|
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2280 |
lemma coeff_0_reflect_poly [simp]: "coeff (reflect_poly p) 0 = lead_coeff p" |
| 64794 | 2281 |
by (simp add: coeff_reflect_poly) |
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2282 |
|
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2283 |
lemma poly_reflect_poly_0 [simp]: "poly (reflect_poly p) 0 = lead_coeff p" |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2284 |
by (simp add: poly_0_coeff_0) |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2285 |
|
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2286 |
lemma reflect_poly_reflect_poly [simp]: "coeff p 0 \<noteq> 0 \<Longrightarrow> reflect_poly (reflect_poly p) = p" |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2287 |
by (cases p rule: pCons_cases) (simp add: reflect_poly_def ) |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2288 |
|
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2289 |
lemma degree_reflect_poly_le: "degree (reflect_poly p) \<le> degree p" |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2290 |
by (simp add: degree_eq_length_coeffs coeffs_reflect_poly length_dropWhile_le diff_le_mono) |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2291 |
|
| 65346 | 2292 |
lemma reflect_poly_pCons: "a \<noteq> 0 \<Longrightarrow> reflect_poly (pCons a p) = Poly (rev (a # coeffs p))" |
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2293 |
by (subst coeffs_eq_iff) (simp add: coeffs_reflect_poly) |
| 65346 | 2294 |
|
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2295 |
lemma degree_reflect_poly_eq [simp]: "coeff p 0 \<noteq> 0 \<Longrightarrow> degree (reflect_poly p) = degree p" |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2296 |
by (cases p rule: pCons_cases) (simp add: reflect_poly_pCons degree_eq_length_coeffs) |
| 65346 | 2297 |
|
| 63498 | 2298 |
(* TODO: does this work with zero divisors as well? Probably not. *) |
| 65346 | 2299 |
lemma reflect_poly_mult: "reflect_poly (p * q) = reflect_poly p * reflect_poly q" |
2300 |
for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
|
|
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2301 |
proof (cases "p = 0 \<or> q = 0") |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2302 |
case False |
| 65346 | 2303 |
then have [simp]: "p \<noteq> 0" "q \<noteq> 0" by auto |
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2304 |
show ?thesis |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2305 |
proof (rule poly_eqI) |
| 65346 | 2306 |
show "coeff (reflect_poly (p * q)) i = coeff (reflect_poly p * reflect_poly q) i" for i |
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2307 |
proof (cases "i \<le> degree (p * q)") |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2308 |
case True |
| 64811 | 2309 |
define A where "A = {..i} \<inter> {i - degree q..degree p}"
|
2310 |
define B where "B = {..degree p} \<inter> {degree p - i..degree (p*q) - i}"
|
|
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2311 |
let ?f = "\<lambda>j. degree p - j" |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2312 |
|
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2313 |
from True have "coeff (reflect_poly (p * q)) i = coeff (p * q) (degree (p * q) - i)" |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2314 |
by (simp add: coeff_reflect_poly) |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2315 |
also have "\<dots> = (\<Sum>j\<le>degree (p * q) - i. coeff p j * coeff q (degree (p * q) - i - j))" |
| 65346 | 2316 |
by (simp add: coeff_mult) |
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2317 |
also have "\<dots> = (\<Sum>j\<in>B. coeff p j * coeff q (degree (p * q) - i - j))" |
| 64267 | 2318 |
by (intro sum.mono_neutral_right) (auto simp: B_def degree_mult_eq not_le coeff_eq_0) |
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2319 |
also from True have "\<dots> = (\<Sum>j\<in>A. coeff p (degree p - j) * coeff q (degree q - (i - j)))" |
| 64267 | 2320 |
by (intro sum.reindex_bij_witness[of _ ?f ?f]) |
| 65346 | 2321 |
(auto simp: A_def B_def degree_mult_eq add_ac) |
2322 |
also have "\<dots> = |
|
2323 |
(\<Sum>j\<le>i. |
|
2324 |
if j \<in> {i - degree q..degree p}
|
|
2325 |
then coeff p (degree p - j) * coeff q (degree q - (i - j)) |
|
2326 |
else 0)" |
|
| 64267 | 2327 |
by (subst sum.inter_restrict [symmetric]) (simp_all add: A_def) |
| 65346 | 2328 |
also have "\<dots> = coeff (reflect_poly p * reflect_poly q) i" |
2329 |
by (fastforce simp: coeff_mult coeff_reflect_poly intro!: sum.cong) |
|
2330 |
finally show ?thesis . |
|
| 64267 | 2331 |
qed (auto simp: coeff_mult coeff_reflect_poly coeff_eq_0 degree_mult_eq intro!: sum.neutral) |
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2332 |
qed |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2333 |
qed auto |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2334 |
|
| 65346 | 2335 |
lemma reflect_poly_smult: "reflect_poly (smult c p) = smult c (reflect_poly p)" |
2336 |
for p :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly"
|
|
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2337 |
using reflect_poly_mult[of "[:c:]" p] by simp |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2338 |
|
| 65346 | 2339 |
lemma reflect_poly_power: "reflect_poly (p ^ n) = reflect_poly p ^ n" |
2340 |
for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly"
|
|
2341 |
by (induct n) (simp_all add: reflect_poly_mult) |
|
2342 |
||
2343 |
lemma reflect_poly_prod: "reflect_poly (prod f A) = prod (\<lambda>x. reflect_poly (f x)) A" |
|
2344 |
for f :: "_ \<Rightarrow> _::{comm_semiring_0,semiring_no_zero_divisors} poly"
|
|
2345 |
by (induct A rule: infinite_finite_induct) (simp_all add: reflect_poly_mult) |
|
2346 |
||
2347 |
lemma reflect_poly_prod_list: "reflect_poly (prod_list xs) = prod_list (map reflect_poly xs)" |
|
2348 |
for xs :: "_::{comm_semiring_0,semiring_no_zero_divisors} poly list"
|
|
2349 |
by (induct xs) (simp_all add: reflect_poly_mult) |
|
2350 |
||
| 65390 | 2351 |
lemma reflect_poly_Poly_nz: |
2352 |
"no_trailing (HOL.eq 0) xs \<Longrightarrow> reflect_poly (Poly xs) = Poly (rev xs)" |
|
| 65346 | 2353 |
by (simp add: reflect_poly_def) |
2354 |
||
2355 |
lemmas reflect_poly_simps = |
|
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2356 |
reflect_poly_0 reflect_poly_1 reflect_poly_const reflect_poly_smult reflect_poly_mult |
| 64272 | 2357 |
reflect_poly_power reflect_poly_prod reflect_poly_prod_list |
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2358 |
|
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
2359 |
|
| 64795 | 2360 |
subsection \<open>Derivatives\<close> |
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2361 |
|
| 63498 | 2362 |
function pderiv :: "('a :: {comm_semiring_1,semiring_no_zero_divisors}) poly \<Rightarrow> 'a poly"
|
| 65346 | 2363 |
where "pderiv (pCons a p) = (if p = 0 then 0 else p + pCons 0 (pderiv p))" |
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2364 |
by (auto intro: pCons_cases) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2365 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2366 |
termination pderiv |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2367 |
by (relation "measure degree") simp_all |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2368 |
|
|
63027
8de0ebee3f1c
several updates on polynomial long division and pseudo division
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
62422
diff
changeset
|
2369 |
declare pderiv.simps[simp del] |
|
8de0ebee3f1c
several updates on polynomial long division and pseudo division
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
62422
diff
changeset
|
2370 |
|
| 65346 | 2371 |
lemma pderiv_0 [simp]: "pderiv 0 = 0" |
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2372 |
using pderiv.simps [of 0 0] by simp |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2373 |
|
| 65346 | 2374 |
lemma pderiv_pCons: "pderiv (pCons a p) = p + pCons 0 (pderiv p)" |
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2375 |
by (simp add: pderiv.simps) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2376 |
|
| 65346 | 2377 |
lemma pderiv_1 [simp]: "pderiv 1 = 0" |
| 65486 | 2378 |
by (simp add: one_pCons pderiv_pCons) |
| 65346 | 2379 |
|
2380 |
lemma pderiv_of_nat [simp]: "pderiv (of_nat n) = 0" |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2381 |
and pderiv_numeral [simp]: "pderiv (numeral m) = 0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2382 |
by (simp_all add: of_nat_poly numeral_poly pderiv_pCons) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2383 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2384 |
lemma coeff_pderiv: "coeff (pderiv p) n = of_nat (Suc n) * coeff p (Suc n)" |
| 65346 | 2385 |
by (induct p arbitrary: n) |
2386 |
(auto simp add: pderiv_pCons coeff_pCons algebra_simps split: nat.split) |
|
2387 |
||
2388 |
fun pderiv_coeffs_code :: "'a::{comm_semiring_1,semiring_no_zero_divisors} \<Rightarrow> 'a list \<Rightarrow> 'a list"
|
|
2389 |
where |
|
2390 |
"pderiv_coeffs_code f (x # xs) = cCons (f * x) (pderiv_coeffs_code (f+1) xs)" |
|
2391 |
| "pderiv_coeffs_code f [] = []" |
|
2392 |
||
2393 |
definition pderiv_coeffs :: "'a::{comm_semiring_1,semiring_no_zero_divisors} list \<Rightarrow> 'a list"
|
|
2394 |
where "pderiv_coeffs xs = pderiv_coeffs_code 1 (tl xs)" |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2395 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2396 |
(* Efficient code for pderiv contributed by René Thiemann and Akihisa Yamada *) |
| 65346 | 2397 |
lemma pderiv_coeffs_code: |
2398 |
"nth_default 0 (pderiv_coeffs_code f xs) n = (f + of_nat n) * nth_default 0 xs n" |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2399 |
proof (induct xs arbitrary: f n) |
| 65346 | 2400 |
case Nil |
2401 |
then show ?case by simp |
|
2402 |
next |
|
2403 |
case (Cons x xs) |
|
2404 |
show ?case |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2405 |
proof (cases n) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2406 |
case 0 |
| 65346 | 2407 |
then show ?thesis |
2408 |
by (cases "pderiv_coeffs_code (f + 1) xs = [] \<and> f * x = 0") (auto simp: cCons_def) |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2409 |
next |
| 65346 | 2410 |
case n: (Suc m) |
2411 |
show ?thesis |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2412 |
proof (cases "pderiv_coeffs_code (f + 1) xs = [] \<and> f * x = 0") |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2413 |
case False |
| 65346 | 2414 |
then have "nth_default 0 (pderiv_coeffs_code f (x # xs)) n = |
2415 |
nth_default 0 (pderiv_coeffs_code (f + 1) xs) m" |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2416 |
by (auto simp: cCons_def n) |
| 65346 | 2417 |
also have "\<dots> = (f + of_nat n) * nth_default 0 xs m" |
2418 |
by (simp add: Cons n add_ac) |
|
2419 |
finally show ?thesis |
|
2420 |
by (simp add: n) |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2421 |
next |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2422 |
case True |
| 65346 | 2423 |
have empty: "pderiv_coeffs_code g xs = [] \<Longrightarrow> g + of_nat m = 0 \<or> nth_default 0 xs m = 0" for g |
2424 |
proof (induct xs arbitrary: g m) |
|
2425 |
case Nil |
|
2426 |
then show ?case by simp |
|
2427 |
next |
|
2428 |
case (Cons x xs) |
|
2429 |
from Cons(2) have empty: "pderiv_coeffs_code (g + 1) xs = []" and g: "g = 0 \<or> x = 0" |
|
2430 |
by (auto simp: cCons_def split: if_splits) |
|
2431 |
note IH = Cons(1)[OF empty] |
|
2432 |
from IH[of m] IH[of "m - 1"] g show ?case |
|
2433 |
by (cases m) (auto simp: field_simps) |
|
2434 |
qed |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2435 |
from True have "nth_default 0 (pderiv_coeffs_code f (x # xs)) n = 0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2436 |
by (auto simp: cCons_def n) |
| 65346 | 2437 |
moreover from True have "(f + of_nat n) * nth_default 0 (x # xs) n = 0" |
2438 |
by (simp add: n) (use empty[of "f+1"] in \<open>auto simp: field_simps\<close>) |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2439 |
ultimately show ?thesis by simp |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2440 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2441 |
qed |
| 65346 | 2442 |
qed |
2443 |
||
2444 |
lemma map_upt_Suc: "map f [0 ..< Suc n] = f 0 # map (\<lambda>i. f (Suc i)) [0 ..< n]" |
|
2445 |
by (induct n arbitrary: f) auto |
|
2446 |
||
2447 |
lemma coeffs_pderiv_code [code abstract]: "coeffs (pderiv p) = pderiv_coeffs (coeffs p)" |
|
2448 |
unfolding pderiv_coeffs_def |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2449 |
proof (rule coeffs_eqI, unfold pderiv_coeffs_code coeff_pderiv, goal_cases) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2450 |
case (1 n) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2451 |
have id: "coeff p (Suc n) = nth_default 0 (map (\<lambda>i. coeff p (Suc i)) [0..<degree p]) n" |
| 65346 | 2452 |
by (cases "n < degree p") (auto simp: nth_default_def coeff_eq_0) |
2453 |
show ?case |
|
2454 |
unfolding coeffs_def map_upt_Suc by (auto simp: id) |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2455 |
next |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2456 |
case 2 |
| 65346 | 2457 |
obtain n :: 'a and xs where defs: "tl (coeffs p) = xs" "1 = n" |
2458 |
by simp |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2459 |
from 2 show ?case |
| 65346 | 2460 |
unfolding defs by (induct xs arbitrary: n) (auto simp: cCons_def) |
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2461 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2462 |
|
| 65346 | 2463 |
lemma pderiv_eq_0_iff: "pderiv p = 0 \<longleftrightarrow> degree p = 0" |
2464 |
for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly"
|
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2465 |
apply (rule iffI) |
| 65346 | 2466 |
apply (cases p) |
2467 |
apply simp |
|
2468 |
apply (simp add: poly_eq_iff coeff_pderiv del: of_nat_Suc) |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2469 |
apply (simp add: poly_eq_iff coeff_pderiv coeff_eq_0) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2470 |
done |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2471 |
|
| 65346 | 2472 |
lemma degree_pderiv: "degree (pderiv p) = degree p - 1" |
2473 |
for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly"
|
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2474 |
apply (rule order_antisym [OF degree_le]) |
| 65346 | 2475 |
apply (simp add: coeff_pderiv coeff_eq_0) |
2476 |
apply (cases "degree p") |
|
2477 |
apply simp |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2478 |
apply (rule le_degree) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2479 |
apply (simp add: coeff_pderiv del: of_nat_Suc) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2480 |
apply (metis degree_0 leading_coeff_0_iff nat.distinct(1)) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2481 |
done |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2482 |
|
| 65346 | 2483 |
lemma not_dvd_pderiv: |
2484 |
fixes p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly"
|
|
2485 |
assumes "degree p \<noteq> 0" |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2486 |
shows "\<not> p dvd pderiv p" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2487 |
proof |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2488 |
assume dvd: "p dvd pderiv p" |
| 65346 | 2489 |
then obtain q where p: "pderiv p = p * q" |
2490 |
unfolding dvd_def by auto |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2491 |
from dvd have le: "degree p \<le> degree (pderiv p)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2492 |
by (simp add: assms dvd_imp_degree_le pderiv_eq_0_iff) |
| 65346 | 2493 |
from assms and this [unfolded degree_pderiv] |
2494 |
show False by auto |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2495 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2496 |
|
| 65346 | 2497 |
lemma dvd_pderiv_iff [simp]: "p dvd pderiv p \<longleftrightarrow> degree p = 0" |
2498 |
for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly"
|
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2499 |
using not_dvd_pderiv[of p] by (auto simp: pderiv_eq_0_iff [symmetric]) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2500 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2501 |
lemma pderiv_singleton [simp]: "pderiv [:a:] = 0" |
| 65346 | 2502 |
by (simp add: pderiv_pCons) |
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2503 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2504 |
lemma pderiv_add: "pderiv (p + q) = pderiv p + pderiv q" |
| 65346 | 2505 |
by (rule poly_eqI) (simp add: coeff_pderiv algebra_simps) |
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2506 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2507 |
lemma pderiv_minus: "pderiv (- p :: 'a :: idom poly) = - pderiv p" |
| 65346 | 2508 |
by (rule poly_eqI) (simp add: coeff_pderiv algebra_simps) |
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2509 |
|
| 63498 | 2510 |
lemma pderiv_diff: "pderiv ((p :: _ :: idom poly) - q) = pderiv p - pderiv q" |
| 65346 | 2511 |
by (rule poly_eqI) (simp add: coeff_pderiv algebra_simps) |
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2512 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2513 |
lemma pderiv_smult: "pderiv (smult a p) = smult a (pderiv p)" |
| 65346 | 2514 |
by (rule poly_eqI) (simp add: coeff_pderiv algebra_simps) |
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2515 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2516 |
lemma pderiv_mult: "pderiv (p * q) = p * pderiv q + q * pderiv p" |
| 65346 | 2517 |
by (induct p) (auto simp: pderiv_add pderiv_smult pderiv_pCons algebra_simps) |
2518 |
||
2519 |
lemma pderiv_power_Suc: "pderiv (p ^ Suc n) = smult (of_nat (Suc n)) (p ^ n) * pderiv p" |
|
2520 |
apply (induct n) |
|
2521 |
apply simp |
|
2522 |
apply (subst power_Suc) |
|
2523 |
apply (subst pderiv_mult) |
|
2524 |
apply (erule ssubst) |
|
2525 |
apply (simp only: of_nat_Suc smult_add_left smult_1_left) |
|
2526 |
apply (simp add: algebra_simps) |
|
2527 |
done |
|
2528 |
||
2529 |
lemma pderiv_prod: "pderiv (prod f (as)) = (\<Sum>a\<in>as. prod f (as - {a}) * pderiv (f a))"
|
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2530 |
proof (induct as rule: infinite_finite_induct) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2531 |
case (insert a as) |
| 65346 | 2532 |
then have id: "prod f (insert a as) = f a * prod f as" |
2533 |
"\<And>g. sum g (insert a as) = g a + sum g as" |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2534 |
"insert a as - {a} = as"
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2535 |
by auto |
| 65346 | 2536 |
have "prod f (insert a as - {b}) = f a * prod f (as - {b})" if "b \<in> as" for b
|
2537 |
proof - |
|
2538 |
from \<open>a \<notin> as\<close> that have *: "insert a as - {b} = insert a (as - {b})"
|
|
2539 |
by auto |
|
2540 |
show ?thesis |
|
2541 |
unfolding * by (subst prod.insert) (use insert in auto) |
|
2542 |
qed |
|
2543 |
then show ?case |
|
| 64267 | 2544 |
unfolding id pderiv_mult insert(3) sum_distrib_left |
| 65346 | 2545 |
by (auto simp add: ac_simps intro!: sum.cong) |
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2546 |
qed auto |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2547 |
|
| 65346 | 2548 |
lemma DERIV_pow2: "DERIV (\<lambda>x. x ^ Suc n) x :> real (Suc n) * (x ^ n)" |
2549 |
by (rule DERIV_cong, rule DERIV_pow) simp |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2550 |
declare DERIV_pow2 [simp] DERIV_pow [simp] |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2551 |
|
| 65346 | 2552 |
lemma DERIV_add_const: "DERIV f x :> D \<Longrightarrow> DERIV (\<lambda>x. a + f x :: 'a::real_normed_field) x :> D" |
2553 |
by (rule DERIV_cong, rule DERIV_add) auto |
|
2554 |
||
2555 |
lemma poly_DERIV [simp]: "DERIV (\<lambda>x. poly p x) x :> poly (pderiv p) x" |
|
2556 |
by (induct p) (auto intro!: derivative_eq_intros simp add: pderiv_pCons) |
|
2557 |
||
2558 |
lemma continuous_on_poly [continuous_intros]: |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2559 |
fixes p :: "'a :: {real_normed_field} poly"
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2560 |
assumes "continuous_on A f" |
| 65346 | 2561 |
shows "continuous_on A (\<lambda>x. poly p (f x))" |
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2562 |
proof - |
| 65346 | 2563 |
have "continuous_on A (\<lambda>x. (\<Sum>i\<le>degree p. (f x) ^ i * coeff p i))" |
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2564 |
by (intro continuous_intros assms) |
| 65346 | 2565 |
also have "\<dots> = (\<lambda>x. poly p (f x))" |
2566 |
by (rule ext) (simp add: poly_altdef mult_ac) |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2567 |
finally show ?thesis . |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2568 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2569 |
|
| 65346 | 2570 |
text \<open>Consequences of the derivative theorem above.\<close> |
2571 |
||
2572 |
lemma poly_differentiable[simp]: "(\<lambda>x. poly p x) differentiable (at x)" |
|
2573 |
for x :: real |
|
2574 |
by (simp add: real_differentiable_def) (blast intro: poly_DERIV) |
|
2575 |
||
2576 |
lemma poly_isCont[simp]: "isCont (\<lambda>x. poly p x) x" |
|
2577 |
for x :: real |
|
2578 |
by (rule poly_DERIV [THEN DERIV_isCont]) |
|
2579 |
||
2580 |
lemma poly_IVT_pos: "a < b \<Longrightarrow> poly p a < 0 \<Longrightarrow> 0 < poly p b \<Longrightarrow> \<exists>x. a < x \<and> x < b \<and> poly p x = 0" |
|
2581 |
for a b :: real |
|
2582 |
using IVT_objl [of "poly p" a 0 b] by (auto simp add: order_le_less) |
|
2583 |
||
2584 |
lemma poly_IVT_neg: "a < b \<Longrightarrow> 0 < poly p a \<Longrightarrow> poly p b < 0 \<Longrightarrow> \<exists>x. a < x \<and> x < b \<and> poly p x = 0" |
|
2585 |
for a b :: real |
|
2586 |
using poly_IVT_pos [where p = "- p"] by simp |
|
2587 |
||
2588 |
lemma poly_IVT: "a < b \<Longrightarrow> poly p a * poly p b < 0 \<Longrightarrow> \<exists>x>a. x < b \<and> poly p x = 0" |
|
2589 |
for p :: "real poly" |
|
2590 |
by (metis less_not_sym mult_less_0_iff poly_IVT_neg poly_IVT_pos) |
|
2591 |
||
2592 |
lemma poly_MVT: "a < b \<Longrightarrow> \<exists>x. a < x \<and> x < b \<and> poly p b - poly p a = (b - a) * poly (pderiv p) x" |
|
2593 |
for a b :: real |
|
2594 |
using MVT [of a b "poly p"] |
|
2595 |
apply auto |
|
2596 |
apply (rule_tac x = z in exI) |
|
2597 |
apply (auto simp add: mult_left_cancel poly_DERIV [THEN DERIV_unique]) |
|
2598 |
done |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2599 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2600 |
lemma poly_MVT': |
| 65346 | 2601 |
fixes a b :: real |
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2602 |
assumes "{min a b..max a b} \<subseteq> A"
|
| 65346 | 2603 |
shows "\<exists>x\<in>A. poly p b - poly p a = (b - a) * poly (pderiv p) x" |
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2604 |
proof (cases a b rule: linorder_cases) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2605 |
case less |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2606 |
from poly_MVT[OF less, of p] guess x by (elim exE conjE) |
| 65346 | 2607 |
then show ?thesis by (intro bexI[of _ x]) (auto intro!: subsetD[OF assms]) |
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2608 |
next |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2609 |
case greater |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2610 |
from poly_MVT[OF greater, of p] guess x by (elim exE conjE) |
| 65346 | 2611 |
then show ?thesis by (intro bexI[of _ x]) (auto simp: algebra_simps intro!: subsetD[OF assms]) |
2612 |
qed (use assms in auto) |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2613 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2614 |
lemma poly_pinfty_gt_lc: |
| 63649 | 2615 |
fixes p :: "real poly" |
| 65346 | 2616 |
assumes "lead_coeff p > 0" |
| 65347 | 2617 |
shows "\<exists>n. \<forall> x \<ge> n. poly p x \<ge> lead_coeff p" |
| 63649 | 2618 |
using assms |
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2619 |
proof (induct p) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2620 |
case 0 |
| 63649 | 2621 |
then show ?case by auto |
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2622 |
next |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2623 |
case (pCons a p) |
| 63649 | 2624 |
from this(1) consider "a \<noteq> 0" "p = 0" | "p \<noteq> 0" by auto |
2625 |
then show ?case |
|
2626 |
proof cases |
|
2627 |
case 1 |
|
2628 |
then show ?thesis by auto |
|
2629 |
next |
|
2630 |
case 2 |
|
2631 |
with pCons obtain n1 where gte_lcoeff: "\<forall>x\<ge>n1. lead_coeff p \<le> poly p x" |
|
2632 |
by auto |
|
2633 |
from pCons(3) \<open>p \<noteq> 0\<close> have gt_0: "lead_coeff p > 0" by auto |
|
2634 |
define n where "n = max n1 (1 + \<bar>a\<bar> / lead_coeff p)" |
|
2635 |
have "lead_coeff (pCons a p) \<le> poly (pCons a p) x" if "n \<le> x" for x |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2636 |
proof - |
| 63649 | 2637 |
from gte_lcoeff that have "lead_coeff p \<le> poly p x" |
2638 |
by (auto simp: n_def) |
|
2639 |
with gt_0 have "\<bar>a\<bar> / lead_coeff p \<ge> \<bar>a\<bar> / poly p x" and "poly p x > 0" |
|
2640 |
by (auto intro: frac_le) |
|
| 65346 | 2641 |
with \<open>n \<le> x\<close>[unfolded n_def] have "x \<ge> 1 + \<bar>a\<bar> / poly p x" |
| 63649 | 2642 |
by auto |
2643 |
with \<open>lead_coeff p \<le> poly p x\<close> \<open>poly p x > 0\<close> \<open>p \<noteq> 0\<close> |
|
2644 |
show "lead_coeff (pCons a p) \<le> poly (pCons a p) x" |
|
2645 |
by (auto simp: field_simps) |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2646 |
qed |
| 63649 | 2647 |
then show ?thesis by blast |
2648 |
qed |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2649 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2650 |
|
| 64795 | 2651 |
lemma lemma_order_pderiv1: |
2652 |
"pderiv ([:- a, 1:] ^ Suc n * q) = [:- a, 1:] ^ Suc n * pderiv q + |
|
2653 |
smult (of_nat (Suc n)) (q * [:- a, 1:] ^ n)" |
|
| 65346 | 2654 |
by (simp only: pderiv_mult pderiv_power_Suc) (simp del: power_Suc of_nat_Suc add: pderiv_pCons) |
| 64795 | 2655 |
|
2656 |
lemma lemma_order_pderiv: |
|
2657 |
fixes p :: "'a :: field_char_0 poly" |
|
| 65346 | 2658 |
assumes n: "0 < n" |
2659 |
and pd: "pderiv p \<noteq> 0" |
|
2660 |
and pe: "p = [:- a, 1:] ^ n * q" |
|
2661 |
and nd: "\<not> [:- a, 1:] dvd q" |
|
2662 |
shows "n = Suc (order a (pderiv p))" |
|
| 64795 | 2663 |
proof - |
| 65346 | 2664 |
from assms have "pderiv ([:- a, 1:] ^ n * q) \<noteq> 0" |
2665 |
by auto |
|
2666 |
from assms obtain n' where "n = Suc n'" "0 < Suc n'" "pderiv ([:- a, 1:] ^ Suc n' * q) \<noteq> 0" |
|
2667 |
by (cases n) auto |
|
2668 |
have *: "k dvd k * pderiv q + smult (of_nat (Suc n')) l \<Longrightarrow> k dvd l" for k l |
|
| 64795 | 2669 |
by (auto simp del: of_nat_Suc simp: dvd_add_right_iff dvd_smult_iff) |
| 65346 | 2670 |
have "n' = order a (pderiv ([:- a, 1:] ^ Suc n' * q))" |
| 64795 | 2671 |
proof (rule order_unique_lemma) |
2672 |
show "[:- a, 1:] ^ n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)" |
|
2673 |
apply (subst lemma_order_pderiv1) |
|
2674 |
apply (rule dvd_add) |
|
| 65346 | 2675 |
apply (metis dvdI dvd_mult2 power_Suc2) |
| 64795 | 2676 |
apply (metis dvd_smult dvd_triv_right) |
2677 |
done |
|
2678 |
show "\<not> [:- a, 1:] ^ Suc n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)" |
|
| 65346 | 2679 |
apply (subst lemma_order_pderiv1) |
2680 |
apply (metis * nd dvd_mult_cancel_right power_not_zero pCons_eq_0_iff power_Suc zero_neq_one) |
|
2681 |
done |
|
| 64795 | 2682 |
qed |
2683 |
then show ?thesis |
|
2684 |
by (metis \<open>n = Suc n'\<close> pe) |
|
2685 |
qed |
|
2686 |
||
| 65346 | 2687 |
lemma order_pderiv: "pderiv p \<noteq> 0 \<Longrightarrow> order a p \<noteq> 0 \<Longrightarrow> order a p = Suc (order a (pderiv p))" |
2688 |
for p :: "'a::field_char_0 poly" |
|
2689 |
apply (cases "p = 0") |
|
2690 |
apply simp |
|
2691 |
apply (drule_tac a = a and p = p in order_decomp) |
|
2692 |
using neq0_conv |
|
2693 |
apply (blast intro: lemma_order_pderiv) |
|
2694 |
done |
|
| 64795 | 2695 |
|
2696 |
lemma poly_squarefree_decomp_order: |
|
| 65346 | 2697 |
fixes p :: "'a::field_char_0 poly" |
2698 |
assumes "pderiv p \<noteq> 0" |
|
2699 |
and p: "p = q * d" |
|
2700 |
and p': "pderiv p = e * d" |
|
2701 |
and d: "d = r * p + s * pderiv p" |
|
| 64795 | 2702 |
shows "order a q = (if order a p = 0 then 0 else 1)" |
2703 |
proof (rule classical) |
|
| 65346 | 2704 |
assume 1: "\<not> ?thesis" |
| 64795 | 2705 |
from \<open>pderiv p \<noteq> 0\<close> have "p \<noteq> 0" by auto |
2706 |
with p have "order a p = order a q + order a d" |
|
2707 |
by (simp add: order_mult) |
|
| 65346 | 2708 |
with 1 have "order a p \<noteq> 0" |
2709 |
by (auto split: if_splits) |
|
2710 |
from \<open>pderiv p \<noteq> 0\<close> \<open>pderiv p = e * d\<close> have "order a (pderiv p) = order a e + order a d" |
|
2711 |
by (simp add: order_mult) |
|
2712 |
from \<open>pderiv p \<noteq> 0\<close> \<open>order a p \<noteq> 0\<close> have "order a p = Suc (order a (pderiv p))" |
|
2713 |
by (rule order_pderiv) |
|
2714 |
from \<open>p \<noteq> 0\<close> \<open>p = q * d\<close> have "d \<noteq> 0" |
|
2715 |
by simp |
|
| 64795 | 2716 |
have "([:-a, 1:] ^ (order a (pderiv p))) dvd d" |
2717 |
apply (simp add: d) |
|
2718 |
apply (rule dvd_add) |
|
| 65346 | 2719 |
apply (rule dvd_mult) |
2720 |
apply (simp add: order_divides \<open>p \<noteq> 0\<close> \<open>order a p = Suc (order a (pderiv p))\<close>) |
|
| 64795 | 2721 |
apply (rule dvd_mult) |
2722 |
apply (simp add: order_divides) |
|
2723 |
done |
|
| 65346 | 2724 |
with \<open>d \<noteq> 0\<close> have "order a (pderiv p) \<le> order a d" |
2725 |
by (simp add: order_divides) |
|
| 64795 | 2726 |
show ?thesis |
2727 |
using \<open>order a p = order a q + order a d\<close> |
|
| 65346 | 2728 |
and \<open>order a (pderiv p) = order a e + order a d\<close> |
2729 |
and \<open>order a p = Suc (order a (pderiv p))\<close> |
|
2730 |
and \<open>order a (pderiv p) \<le> order a d\<close> |
|
| 64795 | 2731 |
by auto |
2732 |
qed |
|
2733 |
||
| 65346 | 2734 |
lemma poly_squarefree_decomp_order2: |
| 65347 | 2735 |
"pderiv p \<noteq> 0 \<Longrightarrow> p = q * d \<Longrightarrow> pderiv p = e * d \<Longrightarrow> |
2736 |
d = r * p + s * pderiv p \<Longrightarrow> \<forall>a. order a q = (if order a p = 0 then 0 else 1)" |
|
2737 |
for p :: "'a::field_char_0 poly" |
|
2738 |
by (blast intro: poly_squarefree_decomp_order) |
|
| 64795 | 2739 |
|
| 65346 | 2740 |
lemma order_pderiv2: |
| 65347 | 2741 |
"pderiv p \<noteq> 0 \<Longrightarrow> order a p \<noteq> 0 \<Longrightarrow> order a (pderiv p) = n \<longleftrightarrow> order a p = Suc n" |
2742 |
for p :: "'a::field_char_0 poly" |
|
2743 |
by (auto dest: order_pderiv) |
|
| 64795 | 2744 |
|
2745 |
definition rsquarefree :: "'a::idom poly \<Rightarrow> bool" |
|
2746 |
where "rsquarefree p \<longleftrightarrow> p \<noteq> 0 \<and> (\<forall>a. order a p = 0 \<or> order a p = 1)" |
|
2747 |
||
| 65347 | 2748 |
lemma pderiv_iszero: "pderiv p = 0 \<Longrightarrow> \<exists>h. p = [:h:]" |
2749 |
for p :: "'a::{semidom,semiring_char_0} poly"
|
|
| 64795 | 2750 |
by (cases p) (auto simp: pderiv_eq_0_iff split: if_splits) |
2751 |
||
| 65347 | 2752 |
lemma rsquarefree_roots: "rsquarefree p \<longleftrightarrow> (\<forall>a. \<not> (poly p a = 0 \<and> poly (pderiv p) a = 0))" |
2753 |
for p :: "'a::field_char_0 poly" |
|
2754 |
apply (simp add: rsquarefree_def) |
|
2755 |
apply (case_tac "p = 0") |
|
2756 |
apply simp |
|
2757 |
apply simp |
|
2758 |
apply (case_tac "pderiv p = 0") |
|
2759 |
apply simp |
|
2760 |
apply (drule pderiv_iszero, clarsimp) |
|
2761 |
apply (metis coeff_0 coeff_pCons_0 degree_pCons_0 le0 le_antisym order_degree) |
|
2762 |
apply (force simp add: order_root order_pderiv2) |
|
| 64795 | 2763 |
done |
2764 |
||
2765 |
lemma poly_squarefree_decomp: |
|
| 65347 | 2766 |
fixes p :: "'a::field_char_0 poly" |
2767 |
assumes "pderiv p \<noteq> 0" |
|
| 64795 | 2768 |
and "p = q * d" |
2769 |
and "pderiv p = e * d" |
|
2770 |
and "d = r * p + s * pderiv p" |
|
| 65347 | 2771 |
shows "rsquarefree q \<and> (\<forall>a. poly q a = 0 \<longleftrightarrow> poly p a = 0)" |
| 64795 | 2772 |
proof - |
2773 |
from \<open>pderiv p \<noteq> 0\<close> have "p \<noteq> 0" by auto |
|
2774 |
with \<open>p = q * d\<close> have "q \<noteq> 0" by simp |
|
| 65347 | 2775 |
from assms have "\<forall>a. order a q = (if order a p = 0 then 0 else 1)" |
2776 |
by (rule poly_squarefree_decomp_order2) |
|
| 64795 | 2777 |
with \<open>p \<noteq> 0\<close> \<open>q \<noteq> 0\<close> show ?thesis |
2778 |
by (simp add: rsquarefree_def order_root) |
|
2779 |
qed |
|
2780 |
||
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2781 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2782 |
subsection \<open>Algebraic numbers\<close> |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2783 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2784 |
text \<open> |
| 65346 | 2785 |
Algebraic numbers can be defined in two equivalent ways: all real numbers that are |
2786 |
roots of rational polynomials or of integer polynomials. The Algebraic-Numbers AFP entry |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2787 |
uses the rational definition, but we need the integer definition. |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2788 |
|
| 65346 | 2789 |
The equivalence is obvious since any rational polynomial can be multiplied with the |
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2790 |
LCM of its coefficients, yielding an integer polynomial with the same roots. |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2791 |
\<close> |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2792 |
|
| 65347 | 2793 |
definition algebraic :: "'a :: field_char_0 \<Rightarrow> bool" |
2794 |
where "algebraic x \<longleftrightarrow> (\<exists>p. (\<forall>i. coeff p i \<in> \<int>) \<and> p \<noteq> 0 \<and> poly p x = 0)" |
|
2795 |
||
2796 |
lemma algebraicI: "(\<And>i. coeff p i \<in> \<int>) \<Longrightarrow> p \<noteq> 0 \<Longrightarrow> poly p x = 0 \<Longrightarrow> algebraic x" |
|
2797 |
unfolding algebraic_def by blast |
|
| 65346 | 2798 |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2799 |
lemma algebraicE: |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2800 |
assumes "algebraic x" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2801 |
obtains p where "\<And>i. coeff p i \<in> \<int>" "p \<noteq> 0" "poly p x = 0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2802 |
using assms unfolding algebraic_def by blast |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2803 |
|
| 65347 | 2804 |
lemma algebraic_altdef: "algebraic x \<longleftrightarrow> (\<exists>p. (\<forall>i. coeff p i \<in> \<rat>) \<and> p \<noteq> 0 \<and> poly p x = 0)" |
2805 |
for p :: "'a::field_char_0 poly" |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2806 |
proof safe |
| 65347 | 2807 |
fix p |
2808 |
assume rat: "\<forall>i. coeff p i \<in> \<rat>" and root: "poly p x = 0" and nz: "p \<noteq> 0" |
|
| 63040 | 2809 |
define cs where "cs = coeffs p" |
| 65347 | 2810 |
from rat have "\<forall>c\<in>range (coeff p). \<exists>c'. c = of_rat c'" |
2811 |
unfolding Rats_def by blast |
|
| 63060 | 2812 |
then obtain f where f: "coeff p i = of_rat (f (coeff p i))" for i |
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2813 |
by (subst (asm) bchoice_iff) blast |
| 63040 | 2814 |
define cs' where "cs' = map (quotient_of \<circ> f) (coeffs p)" |
2815 |
define d where "d = Lcm (set (map snd cs'))" |
|
2816 |
define p' where "p' = smult (of_int d) p" |
|
| 65346 | 2817 |
|
| 65347 | 2818 |
have "coeff p' n \<in> \<int>" for n |
2819 |
proof (cases "n \<le> degree p") |
|
2820 |
case True |
|
2821 |
define c where "c = coeff p n" |
|
2822 |
define a where "a = fst (quotient_of (f (coeff p n)))" |
|
2823 |
define b where "b = snd (quotient_of (f (coeff p n)))" |
|
2824 |
have b_pos: "b > 0" |
|
2825 |
unfolding b_def using quotient_of_denom_pos' by simp |
|
2826 |
have "coeff p' n = of_int d * coeff p n" |
|
2827 |
by (simp add: p'_def) |
|
2828 |
also have "coeff p n = of_rat (of_int a / of_int b)" |
|
2829 |
unfolding a_def b_def |
|
2830 |
by (subst quotient_of_div [of "f (coeff p n)", symmetric]) (simp_all add: f [symmetric]) |
|
2831 |
also have "of_int d * \<dots> = of_rat (of_int (a*d) / of_int b)" |
|
2832 |
by (simp add: of_rat_mult of_rat_divide) |
|
2833 |
also from nz True have "b \<in> snd ` set cs'" |
|
2834 |
by (force simp: cs'_def o_def b_def coeffs_def simp del: upt_Suc) |
|
2835 |
then have "b dvd (a * d)" |
|
2836 |
by (simp add: d_def) |
|
2837 |
then have "of_int (a * d) / of_int b \<in> (\<int> :: rat set)" |
|
2838 |
by (rule of_int_divide_in_Ints) |
|
2839 |
then have "of_rat (of_int (a * d) / of_int b) \<in> \<int>" by (elim Ints_cases) auto |
|
2840 |
finally show ?thesis . |
|
2841 |
next |
|
2842 |
case False |
|
2843 |
then show ?thesis |
|
2844 |
by (auto simp: p'_def not_le coeff_eq_0) |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2845 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2846 |
moreover have "set (map snd cs') \<subseteq> {0<..}"
|
| 65346 | 2847 |
unfolding cs'_def using quotient_of_denom_pos' by (auto simp: coeffs_def simp del: upt_Suc) |
| 65347 | 2848 |
then have "d \<noteq> 0" |
2849 |
unfolding d_def by (induct cs') simp_all |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2850 |
with nz have "p' \<noteq> 0" by (simp add: p'_def) |
| 65347 | 2851 |
moreover from root have "poly p' x = 0" |
2852 |
by (simp add: p'_def) |
|
2853 |
ultimately show "algebraic x" |
|
2854 |
unfolding algebraic_def by blast |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2855 |
next |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2856 |
assume "algebraic x" |
| 63060 | 2857 |
then obtain p where p: "coeff p i \<in> \<int>" "poly p x = 0" "p \<noteq> 0" for i |
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2858 |
by (force simp: algebraic_def) |
| 65347 | 2859 |
moreover have "coeff p i \<in> \<int> \<Longrightarrow> coeff p i \<in> \<rat>" for i |
2860 |
by (elim Ints_cases) simp |
|
2861 |
ultimately show "\<exists>p. (\<forall>i. coeff p i \<in> \<rat>) \<and> p \<noteq> 0 \<and> poly p x = 0" by auto |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2862 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2863 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2864 |
|
| 64795 | 2865 |
subsection \<open>Content and primitive part of a polynomial\<close> |
2866 |
||
| 65347 | 2867 |
definition content :: "'a::semiring_gcd poly \<Rightarrow> 'a" |
2868 |
where "content p = gcd_list (coeffs p)" |
|
2869 |
||
2870 |
lemma content_eq_fold_coeffs [code]: "content p = fold_coeffs gcd p 0" |
|
| 64860 | 2871 |
by (simp add: content_def Gcd_fin.set_eq_fold fold_coeffs_def foldr_fold fun_eq_iff ac_simps) |
| 64795 | 2872 |
|
2873 |
lemma content_0 [simp]: "content 0 = 0" |
|
2874 |
by (simp add: content_def) |
|
2875 |
||
2876 |
lemma content_1 [simp]: "content 1 = 1" |
|
2877 |
by (simp add: content_def) |
|
2878 |
||
2879 |
lemma content_const [simp]: "content [:c:] = normalize c" |
|
2880 |
by (simp add: content_def cCons_def) |
|
2881 |
||
| 65347 | 2882 |
lemma const_poly_dvd_iff_dvd_content: "[:c:] dvd p \<longleftrightarrow> c dvd content p" |
2883 |
for c :: "'a::semiring_gcd" |
|
| 64795 | 2884 |
proof (cases "p = 0") |
| 65347 | 2885 |
case True |
2886 |
then show ?thesis by simp |
|
2887 |
next |
|
2888 |
case False |
|
2889 |
have "[:c:] dvd p \<longleftrightarrow> (\<forall>n. c dvd coeff p n)" |
|
2890 |
by (rule const_poly_dvd_iff) |
|
| 64795 | 2891 |
also have "\<dots> \<longleftrightarrow> (\<forall>a\<in>set (coeffs p). c dvd a)" |
2892 |
proof safe |
|
| 65347 | 2893 |
fix n :: nat |
2894 |
assume "\<forall>a\<in>set (coeffs p). c dvd a" |
|
| 65346 | 2895 |
then show "c dvd coeff p n" |
| 64795 | 2896 |
by (cases "n \<le> degree p") (auto simp: coeff_eq_0 coeffs_def split: if_splits) |
2897 |
qed (auto simp: coeffs_def simp del: upt_Suc split: if_splits) |
|
2898 |
also have "\<dots> \<longleftrightarrow> c dvd content p" |
|
| 64860 | 2899 |
by (simp add: content_def dvd_Gcd_fin_iff dvd_mult_unit_iff) |
| 64795 | 2900 |
finally show ?thesis . |
| 65347 | 2901 |
qed |
| 64795 | 2902 |
|
2903 |
lemma content_dvd [simp]: "[:content p:] dvd p" |
|
2904 |
by (subst const_poly_dvd_iff_dvd_content) simp_all |
|
| 65346 | 2905 |
|
| 64795 | 2906 |
lemma content_dvd_coeff [simp]: "content p dvd coeff p n" |
| 64860 | 2907 |
proof (cases "p = 0") |
2908 |
case True |
|
2909 |
then show ?thesis |
|
2910 |
by simp |
|
2911 |
next |
|
2912 |
case False |
|
2913 |
then show ?thesis |
|
2914 |
by (cases "n \<le> degree p") |
|
2915 |
(auto simp add: content_def not_le coeff_eq_0 coeff_in_coeffs intro: Gcd_fin_dvd) |
|
2916 |
qed |
|
| 65346 | 2917 |
|
| 64795 | 2918 |
lemma content_dvd_coeffs: "c \<in> set (coeffs p) \<Longrightarrow> content p dvd c" |
| 64860 | 2919 |
by (simp add: content_def Gcd_fin_dvd) |
| 64795 | 2920 |
|
2921 |
lemma normalize_content [simp]: "normalize (content p) = content p" |
|
2922 |
by (simp add: content_def) |
|
2923 |
||
2924 |
lemma is_unit_content_iff [simp]: "is_unit (content p) \<longleftrightarrow> content p = 1" |
|
2925 |
proof |
|
2926 |
assume "is_unit (content p)" |
|
| 65346 | 2927 |
then have "normalize (content p) = 1" by (simp add: is_unit_normalize del: normalize_content) |
2928 |
then show "content p = 1" by simp |
|
| 64795 | 2929 |
qed auto |
2930 |
||
2931 |
lemma content_smult [simp]: "content (smult c p) = normalize c * content p" |
|
| 64860 | 2932 |
by (simp add: content_def coeffs_smult Gcd_fin_mult) |
| 64795 | 2933 |
|
2934 |
lemma content_eq_zero_iff [simp]: "content p = 0 \<longleftrightarrow> p = 0" |
|
2935 |
by (auto simp: content_def simp: poly_eq_iff coeffs_def) |
|
2936 |
||
| 65347 | 2937 |
definition primitive_part :: "'a :: semiring_gcd poly \<Rightarrow> 'a poly" |
2938 |
where "primitive_part p = map_poly (\<lambda>x. x div content p) p" |
|
| 64860 | 2939 |
|
| 64795 | 2940 |
lemma primitive_part_0 [simp]: "primitive_part 0 = 0" |
2941 |
by (simp add: primitive_part_def) |
|
2942 |
||
| 65347 | 2943 |
lemma content_times_primitive_part [simp]: "smult (content p) (primitive_part p) = p" |
2944 |
for p :: "'a :: semiring_gcd poly" |
|
| 64795 | 2945 |
proof (cases "p = 0") |
| 65347 | 2946 |
case True |
2947 |
then show ?thesis by simp |
|
2948 |
next |
|
| 64795 | 2949 |
case False |
| 65346 | 2950 |
then show ?thesis |
| 64795 | 2951 |
unfolding primitive_part_def |
| 65346 | 2952 |
by (auto simp: smult_conv_map_poly map_poly_map_poly o_def content_dvd_coeffs |
| 65347 | 2953 |
intro: map_poly_idI) |
2954 |
qed |
|
| 64795 | 2955 |
|
2956 |
lemma primitive_part_eq_0_iff [simp]: "primitive_part p = 0 \<longleftrightarrow> p = 0" |
|
2957 |
proof (cases "p = 0") |
|
| 65347 | 2958 |
case True |
2959 |
then show ?thesis by simp |
|
2960 |
next |
|
| 64795 | 2961 |
case False |
| 65346 | 2962 |
then have "primitive_part p = map_poly (\<lambda>x. x div content p) p" |
| 64795 | 2963 |
by (simp add: primitive_part_def) |
2964 |
also from False have "\<dots> = 0 \<longleftrightarrow> p = 0" |
|
2965 |
by (intro map_poly_eq_0_iff) (auto simp: dvd_div_eq_0_iff content_dvd_coeffs) |
|
| 65347 | 2966 |
finally show ?thesis |
2967 |
using False by simp |
|
2968 |
qed |
|
| 64795 | 2969 |
|
2970 |
lemma content_primitive_part [simp]: |
|
2971 |
assumes "p \<noteq> 0" |
|
| 65347 | 2972 |
shows "content (primitive_part p) = 1" |
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2973 |
proof - |
| 65347 | 2974 |
have "p = smult (content p) (primitive_part p)" |
2975 |
by simp |
|
| 65346 | 2976 |
also have "content \<dots> = content (primitive_part p) * content p" |
| 64860 | 2977 |
by (simp del: content_times_primitive_part add: ac_simps) |
2978 |
finally have "1 * content p = content (primitive_part p) * content p" |
|
2979 |
by simp |
|
2980 |
then have "1 * content p div content p = content (primitive_part p) * content p div content p" |
|
2981 |
by simp |
|
2982 |
with assms show ?thesis |
|
2983 |
by simp |
|
| 64795 | 2984 |
qed |
2985 |
||
2986 |
lemma content_decompose: |
|
| 65347 | 2987 |
obtains p' :: "'a::semiring_gcd poly" where "p = smult (content p) p'" "content p' = 1" |
| 64795 | 2988 |
proof (cases "p = 0") |
2989 |
case True |
|
| 65346 | 2990 |
then show ?thesis by (intro that[of 1]) simp_all |
| 64795 | 2991 |
next |
2992 |
case False |
|
| 65347 | 2993 |
from content_dvd[of p] obtain r where r: "p = [:content p:] * r" |
2994 |
by (rule dvdE) |
|
2995 |
have "content p * 1 = content p * content r" |
|
2996 |
by (subst r) simp |
|
2997 |
with False have "content r = 1" |
|
2998 |
by (subst (asm) mult_left_cancel) simp_all |
|
2999 |
with r show ?thesis |
|
3000 |
by (intro that[of r]) simp_all |
|
| 64795 | 3001 |
qed |
3002 |
||
| 65347 | 3003 |
lemma content_dvd_contentI [intro]: "p dvd q \<Longrightarrow> content p dvd content q" |
| 64795 | 3004 |
using const_poly_dvd_iff_dvd_content content_dvd dvd_trans by blast |
| 65346 | 3005 |
|
| 64795 | 3006 |
lemma primitive_part_const_poly [simp]: "primitive_part [:x:] = [:unit_factor x:]" |
3007 |
by (simp add: primitive_part_def map_poly_pCons) |
|
| 65346 | 3008 |
|
| 64795 | 3009 |
lemma primitive_part_prim: "content p = 1 \<Longrightarrow> primitive_part p = p" |
3010 |
by (auto simp: primitive_part_def) |
|
| 65346 | 3011 |
|
| 64795 | 3012 |
lemma degree_primitive_part [simp]: "degree (primitive_part p) = degree p" |
3013 |
proof (cases "p = 0") |
|
| 65347 | 3014 |
case True |
3015 |
then show ?thesis by simp |
|
3016 |
next |
|
| 64795 | 3017 |
case False |
| 65347 | 3018 |
have "p = smult (content p) (primitive_part p)" |
3019 |
by simp |
|
| 64795 | 3020 |
also from False have "degree \<dots> = degree (primitive_part p)" |
3021 |
by (subst degree_smult_eq) simp_all |
|
3022 |
finally show ?thesis .. |
|
| 65347 | 3023 |
qed |
| 64795 | 3024 |
|
3025 |
||
3026 |
subsection \<open>Division of polynomials\<close> |
|
3027 |
||
3028 |
subsubsection \<open>Division in general\<close> |
|
| 65346 | 3029 |
|
| 64795 | 3030 |
instantiation poly :: (idom_divide) idom_divide |
3031 |
begin |
|
3032 |
||
| 65347 | 3033 |
fun divide_poly_main :: "'a \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a poly" |
3034 |
where |
|
3035 |
"divide_poly_main lc q r d dr (Suc n) = |
|
3036 |
(let cr = coeff r dr; a = cr div lc; mon = monom a n in |
|
3037 |
if False \<or> a * lc = cr then (* False \<or> is only because of problem in function-package *) |
|
3038 |
divide_poly_main |
|
3039 |
lc |
|
3040 |
(q + mon) |
|
3041 |
(r - mon * d) |
|
3042 |
d (dr - 1) n else 0)" |
|
3043 |
| "divide_poly_main lc q r d dr 0 = q" |
|
3044 |
||
3045 |
definition divide_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" |
|
3046 |
where "divide_poly f g = |
|
3047 |
(if g = 0 then 0 |
|
3048 |
else |
|
3049 |
divide_poly_main (coeff g (degree g)) 0 f g (degree f) |
|
3050 |
(1 + length (coeffs f) - length (coeffs g)))" |
|
| 64795 | 3051 |
|
3052 |
lemma divide_poly_main: |
|
3053 |
assumes d: "d \<noteq> 0" "lc = coeff d (degree d)" |
|
| 65347 | 3054 |
and "degree (d * r) \<le> dr" "divide_poly_main lc q (d * r) d dr n = q'" |
3055 |
and "n = 1 + dr - degree d \<or> dr = 0 \<and> n = 0 \<and> d * r = 0" |
|
| 64795 | 3056 |
shows "q' = q + r" |
| 65347 | 3057 |
using assms(3-) |
| 64795 | 3058 |
proof (induct n arbitrary: q r dr) |
| 65347 | 3059 |
case (Suc n) |
| 64795 | 3060 |
let ?rr = "d * r" |
3061 |
let ?a = "coeff ?rr dr" |
|
3062 |
let ?qq = "?a div lc" |
|
3063 |
define b where [simp]: "b = monom ?qq n" |
|
3064 |
let ?rrr = "d * (r - b)" |
|
3065 |
let ?qqq = "q + b" |
|
3066 |
note res = Suc(3) |
|
| 65347 | 3067 |
from Suc(4) have dr: "dr = n + degree d" by auto |
3068 |
from d have lc: "lc \<noteq> 0" by auto |
|
| 64795 | 3069 |
have "coeff (b * d) dr = coeff b n * coeff d (degree d)" |
3070 |
proof (cases "?qq = 0") |
|
| 65347 | 3071 |
case True |
3072 |
then show ?thesis by simp |
|
3073 |
next |
|
| 64795 | 3074 |
case False |
| 65347 | 3075 |
then have n: "n = degree b" |
3076 |
by (simp add: degree_monom_eq) |
|
3077 |
show ?thesis |
|
3078 |
unfolding n dr by (simp add: coeff_mult_degree_sum) |
|
3079 |
qed |
|
3080 |
also have "\<dots> = lc * coeff b n" |
|
3081 |
by (simp add: d) |
|
| 64795 | 3082 |
finally have c2: "coeff (b * d) dr = lc * coeff b n" . |
| 65347 | 3083 |
have rrr: "?rrr = ?rr - b * d" |
3084 |
by (simp add: field_simps) |
|
| 64795 | 3085 |
have c1: "coeff (d * r) dr = lc * coeff r n" |
3086 |
proof (cases "degree r = n") |
|
3087 |
case True |
|
| 65347 | 3088 |
with Suc(2) show ?thesis |
3089 |
unfolding dr using coeff_mult_degree_sum[of d r] d by (auto simp: ac_simps) |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
3090 |
next |
| 64795 | 3091 |
case False |
| 65347 | 3092 |
from dr Suc(2) have "degree r \<le> n" |
3093 |
by auto |
|
3094 |
(metis add.commute add_le_cancel_left d(1) degree_0 degree_mult_eq |
|
3095 |
diff_is_0_eq diff_zero le_cases) |
|
3096 |
with False have r_n: "degree r < n" |
|
3097 |
by auto |
|
3098 |
then have right: "lc * coeff r n = 0" |
|
3099 |
by (simp add: coeff_eq_0) |
|
3100 |
have "coeff (d * r) dr = coeff (d * r) (degree d + n)" |
|
3101 |
by (simp add: dr ac_simps) |
|
3102 |
also from r_n have "\<dots> = 0" |
|
| 65346 | 3103 |
by (metis False Suc.prems(1) add.commute add_left_imp_eq coeff_degree_mult coeff_eq_0 |
| 64795 | 3104 |
coeff_mult_degree_sum degree_mult_le dr le_eq_less_or_eq) |
| 65347 | 3105 |
finally show ?thesis |
3106 |
by (simp only: right) |
|
| 64795 | 3107 |
qed |
| 65346 | 3108 |
have c0: "coeff ?rrr dr = 0" |
| 65347 | 3109 |
and id: "lc * (coeff (d * r) dr div lc) = coeff (d * r) dr" |
3110 |
unfolding rrr coeff_diff c2 |
|
| 64795 | 3111 |
unfolding b_def coeff_monom coeff_smult c1 using lc by auto |
3112 |
from res[unfolded divide_poly_main.simps[of lc q] Let_def] id |
|
| 65346 | 3113 |
have res: "divide_poly_main lc ?qqq ?rrr d (dr - 1) n = q'" |
| 64795 | 3114 |
by (simp del: divide_poly_main.simps add: field_simps) |
| 65346 | 3115 |
note IH = Suc(1)[OF _ res] |
| 65347 | 3116 |
from Suc(4) have dr: "dr = n + degree d" by auto |
3117 |
from Suc(2) have deg_rr: "degree ?rr \<le> dr" by auto |
|
| 64795 | 3118 |
have deg_bd: "degree (b * d) \<le> dr" |
| 65347 | 3119 |
unfolding dr b_def by (rule order.trans[OF degree_mult_le]) (auto simp: degree_monom_le) |
3120 |
have "degree ?rrr \<le> dr" |
|
3121 |
unfolding rrr by (rule degree_diff_le[OF deg_rr deg_bd]) |
|
| 64795 | 3122 |
with c0 have deg_rrr: "degree ?rrr \<le> (dr - 1)" |
3123 |
by (rule coeff_0_degree_minus_1) |
|
| 65346 | 3124 |
have "n = 1 + (dr - 1) - degree d \<or> dr - 1 = 0 \<and> n = 0 \<and> ?rrr = 0" |
| 64795 | 3125 |
proof (cases dr) |
3126 |
case 0 |
|
| 65347 | 3127 |
with Suc(4) have 0: "dr = 0" "n = 0" "degree d = 0" |
3128 |
by auto |
|
3129 |
with deg_rrr have "degree ?rrr = 0" |
|
3130 |
by simp |
|
3131 |
from degree_eq_zeroE[OF this] obtain a where rrr: "?rrr = [:a:]" |
|
3132 |
by metis |
|
3133 |
show ?thesis |
|
3134 |
unfolding 0 using c0 unfolding rrr 0 by simp |
|
3135 |
next |
|
3136 |
case _: Suc |
|
3137 |
with Suc(4) show ?thesis by auto |
|
3138 |
qed |
|
3139 |
from IH[OF deg_rrr this] show ?case |
|
3140 |
by simp |
|
| 64795 | 3141 |
next |
| 65347 | 3142 |
case 0 |
| 65346 | 3143 |
show ?case |
| 64795 | 3144 |
proof (cases "r = 0") |
3145 |
case True |
|
| 65347 | 3146 |
with 0 show ?thesis by auto |
| 64795 | 3147 |
next |
3148 |
case False |
|
| 65347 | 3149 |
from d False have "degree (d * r) = degree d + degree r" |
3150 |
by (subst degree_mult_eq) auto |
|
3151 |
with 0 d show ?thesis by auto |
|
| 64795 | 3152 |
qed |
| 65346 | 3153 |
qed |
| 64795 | 3154 |
|
3155 |
lemma divide_poly_main_0: "divide_poly_main 0 0 r d dr n = 0" |
|
3156 |
proof (induct n arbitrary: r d dr) |
|
| 65347 | 3157 |
case 0 |
3158 |
then show ?case by simp |
|
3159 |
next |
|
3160 |
case Suc |
|
3161 |
show ?case |
|
3162 |
unfolding divide_poly_main.simps[of _ _ r] Let_def |
|
| 64795 | 3163 |
by (simp add: Suc del: divide_poly_main.simps) |
| 65347 | 3164 |
qed |
| 64795 | 3165 |
|
3166 |
lemma divide_poly: |
|
3167 |
assumes g: "g \<noteq> 0" |
|
| 65346 | 3168 |
shows "(f * g) div g = (f :: 'a poly)" |
3169 |
proof - |
|
| 65347 | 3170 |
have len: "length (coeffs f) = Suc (degree f)" if "f \<noteq> 0" for f :: "'a poly" |
3171 |
using that unfolding degree_eq_length_coeffs by auto |
|
3172 |
have "divide_poly_main (coeff g (degree g)) 0 (g * f) g (degree (g * f)) |
|
3173 |
(1 + length (coeffs (g * f)) - length (coeffs g)) = (f * g) div g" |
|
3174 |
by (simp add: divide_poly_def Let_def ac_simps) |
|
| 64795 | 3175 |
note main = divide_poly_main[OF g refl le_refl this] |
3176 |
have "(f * g) div g = 0 + f" |
|
3177 |
proof (rule main, goal_cases) |
|
3178 |
case 1 |
|
3179 |
show ?case |
|
3180 |
proof (cases "f = 0") |
|
3181 |
case True |
|
| 65347 | 3182 |
with g show ?thesis |
3183 |
by (auto simp: degree_eq_length_coeffs) |
|
| 64795 | 3184 |
next |
3185 |
case False |
|
3186 |
with g have fg: "g * f \<noteq> 0" by auto |
|
| 65347 | 3187 |
show ?thesis |
3188 |
unfolding len[OF fg] len[OF g] by auto |
|
| 64795 | 3189 |
qed |
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
3190 |
qed |
| 65346 | 3191 |
then show ?thesis by simp |
| 64795 | 3192 |
qed |
3193 |
||
| 65347 | 3194 |
lemma divide_poly_0: "f div 0 = 0" |
3195 |
for f :: "'a poly" |
|
| 64795 | 3196 |
by (simp add: divide_poly_def Let_def divide_poly_main_0) |
3197 |
||
3198 |
instance |
|
3199 |
by standard (auto simp: divide_poly divide_poly_0) |
|
3200 |
||
3201 |
end |
|
3202 |
||
3203 |
instance poly :: (idom_divide) algebraic_semidom .. |
|
3204 |
||
| 65346 | 3205 |
lemma div_const_poly_conv_map_poly: |
| 64795 | 3206 |
assumes "[:c:] dvd p" |
| 65347 | 3207 |
shows "p div [:c:] = map_poly (\<lambda>x. x div c) p" |
| 64795 | 3208 |
proof (cases "c = 0") |
| 65347 | 3209 |
case True |
3210 |
then show ?thesis |
|
3211 |
by (auto intro!: poly_eqI simp: coeff_map_poly) |
|
3212 |
next |
|
| 64795 | 3213 |
case False |
| 65347 | 3214 |
from assms obtain q where p: "p = [:c:] * q" by (rule dvdE) |
| 64795 | 3215 |
moreover {
|
| 65347 | 3216 |
have "smult c q = [:c:] * q" |
3217 |
by simp |
|
3218 |
also have "\<dots> div [:c:] = q" |
|
3219 |
by (rule nonzero_mult_div_cancel_left) (use False in auto) |
|
| 64795 | 3220 |
finally have "smult c q div [:c:] = q" . |
3221 |
} |
|
3222 |
ultimately show ?thesis by (intro poly_eqI) (auto simp: coeff_map_poly False) |
|
| 65347 | 3223 |
qed |
| 64795 | 3224 |
|
3225 |
lemma is_unit_monom_0: |
|
3226 |
fixes a :: "'a::field" |
|
3227 |
assumes "a \<noteq> 0" |
|
3228 |
shows "is_unit (monom a 0)" |
|
3229 |
proof |
|
3230 |
from assms show "1 = monom a 0 * monom (inverse a) 0" |
|
3231 |
by (simp add: mult_monom) |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
3232 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
3233 |
|
| 65347 | 3234 |
lemma is_unit_triv: "a \<noteq> 0 \<Longrightarrow> is_unit [:a:]" |
3235 |
for a :: "'a::field" |
|
3236 |
by (simp add: is_unit_monom_0 monom_0 [symmetric]) |
|
| 64795 | 3237 |
|
3238 |
lemma is_unit_iff_degree: |
|
| 65347 | 3239 |
fixes p :: "'a::field poly" |
3240 |
assumes "p \<noteq> 0" |
|
3241 |
shows "is_unit p \<longleftrightarrow> degree p = 0" |
|
3242 |
(is "?lhs \<longleftrightarrow> ?rhs") |
|
| 64795 | 3243 |
proof |
| 65347 | 3244 |
assume ?rhs |
3245 |
then obtain a where "p = [:a:]" |
|
3246 |
by (rule degree_eq_zeroE) |
|
3247 |
with assms show ?lhs |
|
3248 |
by (simp add: is_unit_triv) |
|
| 64795 | 3249 |
next |
| 65347 | 3250 |
assume ?lhs |
| 64795 | 3251 |
then obtain q where "q \<noteq> 0" "p * q = 1" .. |
3252 |
then have "degree (p * q) = degree 1" |
|
3253 |
by simp |
|
3254 |
with \<open>p \<noteq> 0\<close> \<open>q \<noteq> 0\<close> have "degree p + degree q = 0" |
|
3255 |
by (simp add: degree_mult_eq) |
|
| 65347 | 3256 |
then show ?rhs by simp |
| 64795 | 3257 |
qed |
3258 |
||
| 65347 | 3259 |
lemma is_unit_pCons_iff: "is_unit (pCons a p) \<longleftrightarrow> p = 0 \<and> a \<noteq> 0" |
3260 |
for p :: "'a::field poly" |
|
3261 |
by (cases "p = 0") (auto simp: is_unit_triv is_unit_iff_degree) |
|
3262 |
||
3263 |
lemma is_unit_monom_trival: "is_unit p \<Longrightarrow> monom (coeff p (degree p)) 0 = p" |
|
3264 |
for p :: "'a::field poly" |
|
3265 |
by (cases p) (simp_all add: monom_0 is_unit_pCons_iff) |
|
3266 |
||
3267 |
lemma is_unit_const_poly_iff: "[:c:] dvd 1 \<longleftrightarrow> c dvd 1" |
|
3268 |
for c :: "'a::{comm_semiring_1,semiring_no_zero_divisors}"
|
|
| 65486 | 3269 |
by (auto simp: one_pCons) |
| 64795 | 3270 |
|
3271 |
lemma is_unit_polyE: |
|
3272 |
fixes p :: "'a :: {comm_semiring_1,semiring_no_zero_divisors} poly"
|
|
| 65347 | 3273 |
assumes "p dvd 1" |
3274 |
obtains c where "p = [:c:]" "c dvd 1" |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
3275 |
proof - |
| 64795 | 3276 |
from assms obtain q where "1 = p * q" |
3277 |
by (rule dvdE) |
|
3278 |
then have "p \<noteq> 0" and "q \<noteq> 0" |
|
3279 |
by auto |
|
3280 |
from \<open>1 = p * q\<close> have "degree 1 = degree (p * q)" |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
3281 |
by simp |
| 64795 | 3282 |
also from \<open>p \<noteq> 0\<close> and \<open>q \<noteq> 0\<close> have "\<dots> = degree p + degree q" |
3283 |
by (simp add: degree_mult_eq) |
|
3284 |
finally have "degree p = 0" by simp |
|
3285 |
with degree_eq_zeroE obtain c where c: "p = [:c:]" . |
|
| 65347 | 3286 |
with \<open>p dvd 1\<close> have "c dvd 1" |
| 64795 | 3287 |
by (simp add: is_unit_const_poly_iff) |
| 65347 | 3288 |
with c show thesis .. |
| 64795 | 3289 |
qed |
3290 |
||
3291 |
lemma is_unit_polyE': |
|
| 65347 | 3292 |
fixes p :: "'a::field poly" |
3293 |
assumes "is_unit p" |
|
| 64795 | 3294 |
obtains a where "p = monom a 0" and "a \<noteq> 0" |
3295 |
proof - |
|
| 65347 | 3296 |
obtain a q where "p = pCons a q" |
3297 |
by (cases p) |
|
| 64795 | 3298 |
with assms have "p = [:a:]" and "a \<noteq> 0" |
3299 |
by (simp_all add: is_unit_pCons_iff) |
|
3300 |
with that show thesis by (simp add: monom_0) |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
3301 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
3302 |
|
| 65347 | 3303 |
lemma is_unit_poly_iff: "p dvd 1 \<longleftrightarrow> (\<exists>c. p = [:c:] \<and> c dvd 1)" |
3304 |
for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly"
|
|
| 64795 | 3305 |
by (auto elim: is_unit_polyE simp add: is_unit_const_poly_iff) |
3306 |
||
| 65346 | 3307 |
|
| 64795 | 3308 |
subsubsection \<open>Pseudo-Division\<close> |
3309 |
||
| 65347 | 3310 |
text \<open>This part is by René Thiemann and Akihisa Yamada.\<close> |
3311 |
||
3312 |
fun pseudo_divmod_main :: |
|
3313 |
"'a :: comm_ring_1 \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a poly \<times> 'a poly" |
|
3314 |
where |
|
3315 |
"pseudo_divmod_main lc q r d dr (Suc n) = |
|
3316 |
(let |
|
3317 |
rr = smult lc r; |
|
3318 |
qq = coeff r dr; |
|
3319 |
rrr = rr - monom qq n * d; |
|
3320 |
qqq = smult lc q + monom qq n |
|
3321 |
in pseudo_divmod_main lc qqq rrr d (dr - 1) n)" |
|
3322 |
| "pseudo_divmod_main lc q r d dr 0 = (q,r)" |
|
3323 |
||
3324 |
definition pseudo_divmod :: "'a :: comm_ring_1 poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<times> 'a poly" |
|
3325 |
where "pseudo_divmod p q \<equiv> |
|
3326 |
if q = 0 then (0, p) |
|
3327 |
else |
|
3328 |
pseudo_divmod_main (coeff q (degree q)) 0 p q (degree p) |
|
3329 |
(1 + length (coeffs p) - length (coeffs q))" |
|
3330 |
||
3331 |
lemma pseudo_divmod_main: |
|
3332 |
assumes d: "d \<noteq> 0" "lc = coeff d (degree d)" |
|
3333 |
and "degree r \<le> dr" "pseudo_divmod_main lc q r d dr n = (q',r')" |
|
3334 |
and "n = 1 + dr - degree d \<or> dr = 0 \<and> n = 0 \<and> r = 0" |
|
| 64795 | 3335 |
shows "(r' = 0 \<or> degree r' < degree d) \<and> smult (lc^n) (d * q + r) = d * q' + r'" |
| 65347 | 3336 |
using assms(3-) |
| 64795 | 3337 |
proof (induct n arbitrary: q r dr) |
| 65347 | 3338 |
case 0 |
3339 |
then show ?case by auto |
|
3340 |
next |
|
3341 |
case (Suc n) |
|
| 64795 | 3342 |
let ?rr = "smult lc r" |
3343 |
let ?qq = "coeff r dr" |
|
3344 |
define b where [simp]: "b = monom ?qq n" |
|
3345 |
let ?rrr = "?rr - b * d" |
|
3346 |
let ?qqq = "smult lc q + b" |
|
3347 |
note res = Suc(3) |
|
| 65346 | 3348 |
from res[unfolded pseudo_divmod_main.simps[of lc q] Let_def] |
3349 |
have res: "pseudo_divmod_main lc ?qqq ?rrr d (dr - 1) n = (q',r')" |
|
| 64795 | 3350 |
by (simp del: pseudo_divmod_main.simps) |
| 65347 | 3351 |
from Suc(4) have dr: "dr = n + degree d" by auto |
| 64795 | 3352 |
have "coeff (b * d) dr = coeff b n * coeff d (degree d)" |
3353 |
proof (cases "?qq = 0") |
|
| 65347 | 3354 |
case True |
3355 |
then show ?thesis by auto |
|
3356 |
next |
|
| 64795 | 3357 |
case False |
| 65347 | 3358 |
then have n: "n = degree b" |
3359 |
by (simp add: degree_monom_eq) |
|
3360 |
show ?thesis |
|
3361 |
unfolding n dr by (simp add: coeff_mult_degree_sum) |
|
3362 |
qed |
|
3363 |
also have "\<dots> = lc * coeff b n" by (simp add: d) |
|
| 64795 | 3364 |
finally have "coeff (b * d) dr = lc * coeff b n" . |
| 65347 | 3365 |
moreover have "coeff ?rr dr = lc * coeff r dr" |
3366 |
by simp |
|
3367 |
ultimately have c0: "coeff ?rrr dr = 0" |
|
3368 |
by auto |
|
3369 |
from Suc(4) have dr: "dr = n + degree d" by auto |
|
3370 |
have deg_rr: "degree ?rr \<le> dr" |
|
3371 |
using Suc(2) degree_smult_le dual_order.trans by blast |
|
| 64795 | 3372 |
have deg_bd: "degree (b * d) \<le> dr" |
| 65347 | 3373 |
unfolding dr by (rule order.trans[OF degree_mult_le]) (auto simp: degree_monom_le) |
| 64795 | 3374 |
have "degree ?rrr \<le> dr" |
3375 |
using degree_diff_le[OF deg_rr deg_bd] by auto |
|
| 65347 | 3376 |
with c0 have deg_rrr: "degree ?rrr \<le> (dr - 1)" |
3377 |
by (rule coeff_0_degree_minus_1) |
|
| 64795 | 3378 |
have "n = 1 + (dr - 1) - degree d \<or> dr - 1 = 0 \<and> n = 0 \<and> ?rrr = 0" |
3379 |
proof (cases dr) |
|
3380 |
case 0 |
|
3381 |
with Suc(4) have 0: "dr = 0" "n = 0" "degree d = 0" by auto |
|
3382 |
with deg_rrr have "degree ?rrr = 0" by simp |
|
| 65347 | 3383 |
then have "\<exists>a. ?rrr = [:a:]" |
3384 |
by (metis degree_pCons_eq_if old.nat.distinct(2) pCons_cases) |
|
3385 |
from this obtain a where rrr: "?rrr = [:a:]" |
|
3386 |
by auto |
|
3387 |
show ?thesis |
|
3388 |
unfolding 0 using c0 unfolding rrr 0 by simp |
|
3389 |
next |
|
3390 |
case _: Suc |
|
3391 |
with Suc(4) show ?thesis by auto |
|
3392 |
qed |
|
| 64795 | 3393 |
note IH = Suc(1)[OF deg_rrr res this] |
3394 |
show ?case |
|
3395 |
proof (intro conjI) |
|
| 65347 | 3396 |
from IH show "r' = 0 \<or> degree r' < degree d" |
3397 |
by blast |
|
| 64795 | 3398 |
show "smult (lc ^ Suc n) (d * q + r) = d * q' + r'" |
3399 |
unfolding IH[THEN conjunct2,symmetric] |
|
3400 |
by (simp add: field_simps smult_add_right) |
|
3401 |
qed |
|
| 65347 | 3402 |
qed |
| 64795 | 3403 |
|
3404 |
lemma pseudo_divmod: |
|
| 65347 | 3405 |
assumes g: "g \<noteq> 0" |
3406 |
and *: "pseudo_divmod f g = (q,r)" |
|
3407 |
shows "smult (coeff g (degree g) ^ (Suc (degree f) - degree g)) f = g * q + r" (is ?A) |
|
3408 |
and "r = 0 \<or> degree r < degree g" (is ?B) |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
3409 |
proof - |
| 64795 | 3410 |
from *[unfolded pseudo_divmod_def Let_def] |
| 65347 | 3411 |
have "pseudo_divmod_main (coeff g (degree g)) 0 f g (degree f) |
3412 |
(1 + length (coeffs f) - length (coeffs g)) = (q, r)" |
|
3413 |
by (auto simp: g) |
|
| 64795 | 3414 |
note main = pseudo_divmod_main[OF _ _ _ this, OF g refl le_refl] |
| 65347 | 3415 |
from g have "1 + length (coeffs f) - length (coeffs g) = 1 + degree f - degree g \<or> |
3416 |
degree f = 0 \<and> 1 + length (coeffs f) - length (coeffs g) = 0 \<and> f = 0" |
|
3417 |
by (cases "f = 0"; cases "coeffs g") (auto simp: degree_eq_length_coeffs) |
|
3418 |
note main' = main[OF this] |
|
3419 |
then show "r = 0 \<or> degree r < degree g" by auto |
|
| 65346 | 3420 |
show "smult (coeff g (degree g) ^ (Suc (degree f) - degree g)) f = g * q + r" |
| 65347 | 3421 |
by (subst main'[THEN conjunct2, symmetric], simp add: degree_eq_length_coeffs, |
3422 |
cases "f = 0"; cases "coeffs g", use g in auto) |
|
| 64795 | 3423 |
qed |
| 65346 | 3424 |
|
| 64795 | 3425 |
definition "pseudo_mod_main lc r d dr n = snd (pseudo_divmod_main lc 0 r d dr n)" |
3426 |
||
3427 |
lemma snd_pseudo_divmod_main: |
|
3428 |
"snd (pseudo_divmod_main lc q r d dr n) = snd (pseudo_divmod_main lc q' r d dr n)" |
|
| 65347 | 3429 |
by (induct n arbitrary: q q' lc r d dr) (simp_all add: Let_def) |
3430 |
||
3431 |
definition pseudo_mod :: "'a::{comm_ring_1,semiring_1_no_zero_divisors} poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly"
|
|
3432 |
where "pseudo_mod f g = snd (pseudo_divmod f g)" |
|
| 65346 | 3433 |
|
| 64795 | 3434 |
lemma pseudo_mod: |
| 65347 | 3435 |
fixes f g :: "'a::{comm_ring_1,semiring_1_no_zero_divisors} poly"
|
| 64795 | 3436 |
defines "r \<equiv> pseudo_mod f g" |
3437 |
assumes g: "g \<noteq> 0" |
|
| 65347 | 3438 |
shows "\<exists>a q. a \<noteq> 0 \<and> smult a f = g * q + r" "r = 0 \<or> degree r < degree g" |
| 65346 | 3439 |
proof - |
| 64795 | 3440 |
let ?cg = "coeff g (degree g)" |
3441 |
let ?cge = "?cg ^ (Suc (degree f) - degree g)" |
|
3442 |
define a where "a = ?cge" |
|
| 65347 | 3443 |
from r_def[unfolded pseudo_mod_def] obtain q where pdm: "pseudo_divmod f g = (q, r)" |
3444 |
by (cases "pseudo_divmod f g") auto |
|
| 65346 | 3445 |
from pseudo_divmod[OF g pdm] have id: "smult a f = g * q + r" and "r = 0 \<or> degree r < degree g" |
| 65347 | 3446 |
by (auto simp: a_def) |
| 64795 | 3447 |
show "r = 0 \<or> degree r < degree g" by fact |
| 65347 | 3448 |
from g have "a \<noteq> 0" |
3449 |
by (auto simp: a_def) |
|
3450 |
with id show "\<exists>a q. a \<noteq> 0 \<and> smult a f = g * q + r" |
|
3451 |
by auto |
|
| 64795 | 3452 |
qed |
| 65346 | 3453 |
|
| 64795 | 3454 |
lemma fst_pseudo_divmod_main_as_divide_poly_main: |
3455 |
assumes d: "d \<noteq> 0" |
|
3456 |
defines lc: "lc \<equiv> coeff d (degree d)" |
|
| 65347 | 3457 |
shows "fst (pseudo_divmod_main lc q r d dr n) = |
3458 |
divide_poly_main lc (smult (lc^n) q) (smult (lc^n) r) d dr n" |
|
3459 |
proof (induct n arbitrary: q r dr) |
|
3460 |
case 0 |
|
3461 |
then show ?case by simp |
|
| 64795 | 3462 |
next |
3463 |
case (Suc n) |
|
| 65347 | 3464 |
note lc0 = leading_coeff_neq_0[OF d, folded lc] |
3465 |
then have "pseudo_divmod_main lc q r d dr (Suc n) = |
|
| 64795 | 3466 |
pseudo_divmod_main lc (smult lc q + monom (coeff r dr) n) |
3467 |
(smult lc r - monom (coeff r dr) n * d) d (dr - 1) n" |
|
3468 |
by (simp add: Let_def ac_simps) |
|
| 65347 | 3469 |
also have "fst \<dots> = divide_poly_main lc |
| 64795 | 3470 |
(smult (lc^n) (smult lc q + monom (coeff r dr) n)) |
3471 |
(smult (lc^n) (smult lc r - monom (coeff r dr) n * d)) |
|
3472 |
d (dr - 1) n" |
|
| 65347 | 3473 |
by (simp only: Suc[unfolded divide_poly_main.simps Let_def]) |
3474 |
also have "\<dots> = divide_poly_main lc (smult (lc ^ Suc n) q) (smult (lc ^ Suc n) r) d dr (Suc n)" |
|
3475 |
unfolding smult_monom smult_distribs mult_smult_left[symmetric] |
|
3476 |
using lc0 by (simp add: Let_def ac_simps) |
|
3477 |
finally show ?case . |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
3478 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
3479 |
|
| 64795 | 3480 |
|
3481 |
subsubsection \<open>Division in polynomials over fields\<close> |
|
3482 |
||
3483 |
lemma pseudo_divmod_field: |
|
| 65347 | 3484 |
fixes g :: "'a::field poly" |
3485 |
assumes g: "g \<noteq> 0" |
|
3486 |
and *: "pseudo_divmod f g = (q,r)" |
|
| 64795 | 3487 |
defines "c \<equiv> coeff g (degree g) ^ (Suc (degree f) - degree g)" |
3488 |
shows "f = g * smult (1/c) q + smult (1/c) r" |
|
3489 |
proof - |
|
| 65347 | 3490 |
from leading_coeff_neq_0[OF g] have c0: "c \<noteq> 0" |
3491 |
by (auto simp: c_def) |
|
3492 |
from pseudo_divmod(1)[OF g *, folded c_def] have "smult c f = g * q + r" |
|
3493 |
by auto |
|
3494 |
also have "smult (1 / c) \<dots> = g * smult (1 / c) q + smult (1 / c) r" |
|
3495 |
by (simp add: smult_add_right) |
|
3496 |
finally show ?thesis |
|
3497 |
using c0 by auto |
|
| 64795 | 3498 |
qed |
3499 |
||
3500 |
lemma divide_poly_main_field: |
|
| 65347 | 3501 |
fixes d :: "'a::field poly" |
3502 |
assumes d: "d \<noteq> 0" |
|
| 64795 | 3503 |
defines lc: "lc \<equiv> coeff d (degree d)" |
| 65347 | 3504 |
shows "divide_poly_main lc q r d dr n = |
3505 |
fst (pseudo_divmod_main lc (smult ((1 / lc)^n) q) (smult ((1 / lc)^n) r) d dr n)" |
|
3506 |
unfolding lc by (subst fst_pseudo_divmod_main_as_divide_poly_main) (auto simp: d power_one_over) |
|
| 64795 | 3507 |
|
3508 |
lemma divide_poly_field: |
|
| 65347 | 3509 |
fixes f g :: "'a::field poly" |
| 64795 | 3510 |
defines "f' \<equiv> smult ((1 / coeff g (degree g)) ^ (Suc (degree f) - degree g)) f" |
| 65347 | 3511 |
shows "f div g = fst (pseudo_divmod f' g)" |
| 64795 | 3512 |
proof (cases "g = 0") |
| 65347 | 3513 |
case True |
3514 |
show ?thesis |
|
3515 |
unfolding divide_poly_def pseudo_divmod_def Let_def f'_def True |
|
3516 |
by (simp add: divide_poly_main_0) |
|
| 64795 | 3517 |
next |
3518 |
case False |
|
| 65347 | 3519 |
from leading_coeff_neq_0[OF False] have "degree f' = degree f" |
3520 |
by (auto simp: f'_def) |
|
3521 |
then show ?thesis |
|
3522 |
using length_coeffs_degree[of f'] length_coeffs_degree[of f] |
|
3523 |
unfolding divide_poly_def pseudo_divmod_def Let_def |
|
3524 |
divide_poly_main_field[OF False] |
|
3525 |
length_coeffs_degree[OF False] |
|
3526 |
f'_def |
|
3527 |
by force |
|
| 64795 | 3528 |
qed |
3529 |
||
| 65347 | 3530 |
instantiation poly :: ("{semidom_divide_unit_factor,idom_divide}") normalization_semidom
|
| 64795 | 3531 |
begin |
3532 |
||
3533 |
definition unit_factor_poly :: "'a poly \<Rightarrow> 'a poly" |
|
|
64848
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3534 |
where "unit_factor_poly p = [:unit_factor (lead_coeff p):]" |
| 64795 | 3535 |
|
3536 |
definition normalize_poly :: "'a poly \<Rightarrow> 'a poly" |
|
|
64848
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3537 |
where "normalize p = p div [:unit_factor (lead_coeff p):]" |
| 64795 | 3538 |
|
| 65347 | 3539 |
instance |
3540 |
proof |
|
| 64795 | 3541 |
fix p :: "'a poly" |
3542 |
show "unit_factor p * normalize p = p" |
|
|
64848
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3543 |
proof (cases "p = 0") |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3544 |
case True |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3545 |
then show ?thesis |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3546 |
by (simp add: unit_factor_poly_def normalize_poly_def) |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3547 |
next |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3548 |
case False |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3549 |
then have "lead_coeff p \<noteq> 0" |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3550 |
by simp |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3551 |
then have *: "unit_factor (lead_coeff p) \<noteq> 0" |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3552 |
using unit_factor_is_unit [of "lead_coeff p"] by auto |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3553 |
then have "unit_factor (lead_coeff p) dvd 1" |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3554 |
by (auto intro: unit_factor_is_unit) |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3555 |
then have **: "unit_factor (lead_coeff p) dvd c" for c |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3556 |
by (rule dvd_trans) simp |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3557 |
have ***: "unit_factor (lead_coeff p) * (c div unit_factor (lead_coeff p)) = c" for c |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3558 |
proof - |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3559 |
from ** obtain b where "c = unit_factor (lead_coeff p) * b" .. |
| 65347 | 3560 |
with False * show ?thesis by simp |
|
64848
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3561 |
qed |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3562 |
have "p div [:unit_factor (lead_coeff p):] = |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3563 |
map_poly (\<lambda>c. c div unit_factor (lead_coeff p)) p" |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3564 |
by (simp add: const_poly_dvd_iff div_const_poly_conv_map_poly **) |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3565 |
then show ?thesis |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3566 |
by (simp add: normalize_poly_def unit_factor_poly_def |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3567 |
smult_conv_map_poly map_poly_map_poly o_def ***) |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3568 |
qed |
| 64795 | 3569 |
next |
3570 |
fix p :: "'a poly" |
|
3571 |
assume "is_unit p" |
|
|
64848
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3572 |
then obtain c where p: "p = [:c:]" "c dvd 1" |
| 64795 | 3573 |
by (auto simp: is_unit_poly_iff) |
|
64848
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3574 |
then show "unit_factor p = p" |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3575 |
by (simp add: unit_factor_poly_def monom_0 is_unit_unit_factor) |
| 64795 | 3576 |
next |
| 65347 | 3577 |
fix p :: "'a poly" |
3578 |
assume "p \<noteq> 0" |
|
|
64848
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3579 |
then show "is_unit (unit_factor p)" |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3580 |
by (simp add: unit_factor_poly_def monom_0 is_unit_poly_iff unit_factor_is_unit) |
| 64795 | 3581 |
qed (simp_all add: normalize_poly_def unit_factor_poly_def monom_0 lead_coeff_mult unit_factor_mult) |
3582 |
||
3583 |
end |
|
3584 |
||
| 65347 | 3585 |
lemma normalize_poly_eq_map_poly: "normalize p = map_poly (\<lambda>x. x div unit_factor (lead_coeff p)) p" |
|
64848
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3586 |
proof - |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3587 |
have "[:unit_factor (lead_coeff p):] dvd p" |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3588 |
by (metis unit_factor_poly_def unit_factor_self) |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3589 |
then show ?thesis |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3590 |
by (simp add: normalize_poly_def div_const_poly_conv_map_poly) |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3591 |
qed |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3592 |
|
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3593 |
lemma coeff_normalize [simp]: |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3594 |
"coeff (normalize p) n = coeff p n div unit_factor (lead_coeff p)" |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3595 |
by (simp add: normalize_poly_eq_map_poly coeff_map_poly) |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3596 |
|
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3597 |
class field_unit_factor = field + unit_factor + |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3598 |
assumes unit_factor_field [simp]: "unit_factor = id" |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3599 |
begin |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3600 |
|
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3601 |
subclass semidom_divide_unit_factor |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3602 |
proof |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3603 |
fix a |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3604 |
assume "a \<noteq> 0" |
| 65347 | 3605 |
then have "1 = a * inverse a" by simp |
|
64848
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3606 |
then have "a dvd 1" .. |
| 65347 | 3607 |
then show "unit_factor a dvd 1" by simp |
|
64848
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3608 |
qed simp_all |
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3609 |
|
|
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3610 |
end |
| 64795 | 3611 |
|
3612 |
lemma unit_factor_pCons: |
|
|
64848
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3613 |
"unit_factor (pCons a p) = (if p = 0 then [:unit_factor a:] else unit_factor p)" |
| 64795 | 3614 |
by (simp add: unit_factor_poly_def) |
3615 |
||
| 65347 | 3616 |
lemma normalize_monom [simp]: "normalize (monom a n) = monom (normalize a) n" |
|
64848
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3617 |
by (cases "a = 0") (simp_all add: map_poly_monom normalize_poly_eq_map_poly degree_monom_eq) |
| 64795 | 3618 |
|
| 65347 | 3619 |
lemma unit_factor_monom [simp]: "unit_factor (monom a n) = [:unit_factor a:]" |
| 64795 | 3620 |
by (cases "a = 0") (simp_all add: unit_factor_poly_def degree_monom_eq) |
3621 |
||
3622 |
lemma normalize_const_poly: "normalize [:c:] = [:normalize c:]" |
|
|
64848
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents:
64811
diff
changeset
|
3623 |
by (simp add: normalize_poly_eq_map_poly map_poly_pCons) |
| 64795 | 3624 |
|
3625 |
lemma normalize_smult: "normalize (smult c p) = smult (normalize c) (normalize p)" |
|
3626 |
proof - |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
3627 |
have "smult c p = [:c:] * p" by simp |
| 64795 | 3628 |
also have "normalize \<dots> = smult (normalize c) (normalize p)" |
3629 |
by (subst normalize_mult) (simp add: normalize_const_poly) |
|
3630 |
finally show ?thesis . |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
3631 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
3632 |
|
| 64795 | 3633 |
lemma smult_content_normalize_primitive_part [simp]: |
3634 |
"smult (content p) (normalize (primitive_part p)) = normalize p" |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
3635 |
proof - |
| 65346 | 3636 |
have "smult (content p) (normalize (primitive_part p)) = |
| 65347 | 3637 |
normalize ([:content p:] * primitive_part p)" |
| 64795 | 3638 |
by (subst normalize_mult) (simp_all add: normalize_const_poly) |
3639 |
also have "[:content p:] * primitive_part p = p" by simp |
|
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
3640 |
finally show ?thesis . |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
3641 |
qed |
|
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
3642 |
|
| 64795 | 3643 |
inductive eucl_rel_poly :: "'a::field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<times> 'a poly \<Rightarrow> bool" |
| 65347 | 3644 |
where |
3645 |
eucl_rel_poly_by0: "eucl_rel_poly x 0 (0, x)" |
|
| 64795 | 3646 |
| eucl_rel_poly_dividesI: "y \<noteq> 0 \<Longrightarrow> x = q * y \<Longrightarrow> eucl_rel_poly x y (q, 0)" |
| 65347 | 3647 |
| eucl_rel_poly_remainderI: |
3648 |
"y \<noteq> 0 \<Longrightarrow> degree r < degree y \<Longrightarrow> x = q * y + r \<Longrightarrow> eucl_rel_poly x y (q, r)" |
|
| 65346 | 3649 |
|
| 64795 | 3650 |
lemma eucl_rel_poly_iff: |
3651 |
"eucl_rel_poly x y (q, r) \<longleftrightarrow> |
|
| 65347 | 3652 |
x = q * y + r \<and> (if y = 0 then q = 0 else r = 0 \<or> degree r < degree y)" |
| 64795 | 3653 |
by (auto elim: eucl_rel_poly.cases |
| 65347 | 3654 |
intro: eucl_rel_poly_by0 eucl_rel_poly_dividesI eucl_rel_poly_remainderI) |
3655 |
||
3656 |
lemma eucl_rel_poly_0: "eucl_rel_poly 0 y (0, 0)" |
|
3657 |
by (simp add: eucl_rel_poly_iff) |
|
3658 |
||
3659 |
lemma eucl_rel_poly_by_0: "eucl_rel_poly x 0 (0, x)" |
|
3660 |
by (simp add: eucl_rel_poly_iff) |
|
| 64795 | 3661 |
|
3662 |
lemma eucl_rel_poly_pCons: |
|
3663 |
assumes rel: "eucl_rel_poly x y (q, r)" |
|
3664 |
assumes y: "y \<noteq> 0" |
|
3665 |
assumes b: "b = coeff (pCons a r) (degree y) / coeff y (degree y)" |
|
3666 |
shows "eucl_rel_poly (pCons a x) y (pCons b q, pCons a r - smult b y)" |
|
3667 |
(is "eucl_rel_poly ?x y (?q, ?r)") |
|
3668 |
proof - |
|
| 65347 | 3669 |
from assms have x: "x = q * y + r" and r: "r = 0 \<or> degree r < degree y" |
3670 |
by (simp_all add: eucl_rel_poly_iff) |
|
3671 |
from b x have "?x = ?q * y + ?r" by simp |
|
3672 |
moreover |
|
3673 |
have "?r = 0 \<or> degree ?r < degree y" |
|
| 64795 | 3674 |
proof (rule eq_zero_or_degree_less) |
3675 |
show "degree ?r \<le> degree y" |
|
3676 |
proof (rule degree_diff_le) |
|
| 65347 | 3677 |
from r show "degree (pCons a r) \<le> degree y" |
3678 |
by auto |
|
| 64795 | 3679 |
show "degree (smult b y) \<le> degree y" |
3680 |
by (rule degree_smult_le) |
|
3681 |
qed |
|
| 65347 | 3682 |
from \<open>y \<noteq> 0\<close> show "coeff ?r (degree y) = 0" |
3683 |
by (simp add: b) |
|
| 64795 | 3684 |
qed |
| 65347 | 3685 |
ultimately show ?thesis |
3686 |
unfolding eucl_rel_poly_iff using \<open>y \<noteq> 0\<close> by simp |
|
| 64795 | 3687 |
qed |
3688 |
||
3689 |
lemma eucl_rel_poly_exists: "\<exists>q r. eucl_rel_poly x y (q, r)" |
|
| 65347 | 3690 |
apply (cases "y = 0") |
3691 |
apply (fast intro!: eucl_rel_poly_by_0) |
|
3692 |
apply (induct x) |
|
3693 |
apply (fast intro!: eucl_rel_poly_0) |
|
3694 |
apply (fast intro!: eucl_rel_poly_pCons) |
|
3695 |
done |
|
| 64795 | 3696 |
|
3697 |
lemma eucl_rel_poly_unique: |
|
3698 |
assumes 1: "eucl_rel_poly x y (q1, r1)" |
|
3699 |
assumes 2: "eucl_rel_poly x y (q2, r2)" |
|
3700 |
shows "q1 = q2 \<and> r1 = r2" |
|
3701 |
proof (cases "y = 0") |
|
| 65347 | 3702 |
assume "y = 0" |
3703 |
with assms show ?thesis |
|
| 64795 | 3704 |
by (simp add: eucl_rel_poly_iff) |
3705 |
next |
|
3706 |
assume [simp]: "y \<noteq> 0" |
|
3707 |
from 1 have q1: "x = q1 * y + r1" and r1: "r1 = 0 \<or> degree r1 < degree y" |
|
3708 |
unfolding eucl_rel_poly_iff by simp_all |
|
3709 |
from 2 have q2: "x = q2 * y + r2" and r2: "r2 = 0 \<or> degree r2 < degree y" |
|
3710 |
unfolding eucl_rel_poly_iff by simp_all |
|
3711 |
from q1 q2 have q3: "(q1 - q2) * y = r2 - r1" |
|
3712 |
by (simp add: algebra_simps) |
|
3713 |
from r1 r2 have r3: "(r2 - r1) = 0 \<or> degree (r2 - r1) < degree y" |
|
3714 |
by (auto intro: degree_diff_less) |
|
3715 |
show "q1 = q2 \<and> r1 = r2" |
|
| 65347 | 3716 |
proof (rule classical) |
3717 |
assume "\<not> ?thesis" |
|
| 64795 | 3718 |
with q3 have "q1 \<noteq> q2" and "r1 \<noteq> r2" by auto |
3719 |
with r3 have "degree (r2 - r1) < degree y" by simp |
|
3720 |
also have "degree y \<le> degree (q1 - q2) + degree y" by simp |
|
| 65347 | 3721 |
also from \<open>q1 \<noteq> q2\<close> have "\<dots> = degree ((q1 - q2) * y)" |
3722 |
by (simp add: degree_mult_eq) |
|
3723 |
also from q3 have "\<dots> = degree (r2 - r1)" |
|
3724 |
by simp |
|
| 64795 | 3725 |
finally have "degree (r2 - r1) < degree (r2 - r1)" . |
| 65347 | 3726 |
then show ?thesis by simp |
| 64795 | 3727 |
qed |
3728 |
qed |
|
3729 |
||
3730 |
lemma eucl_rel_poly_0_iff: "eucl_rel_poly 0 y (q, r) \<longleftrightarrow> q = 0 \<and> r = 0" |
|
| 65347 | 3731 |
by (auto dest: eucl_rel_poly_unique intro: eucl_rel_poly_0) |
| 64795 | 3732 |
|
3733 |
lemma eucl_rel_poly_by_0_iff: "eucl_rel_poly x 0 (q, r) \<longleftrightarrow> q = 0 \<and> r = x" |
|
| 65347 | 3734 |
by (auto dest: eucl_rel_poly_unique intro: eucl_rel_poly_by_0) |
| 64795 | 3735 |
|
3736 |
lemmas eucl_rel_poly_unique_div = eucl_rel_poly_unique [THEN conjunct1] |
|
3737 |
||
3738 |
lemmas eucl_rel_poly_unique_mod = eucl_rel_poly_unique [THEN conjunct2] |
|
3739 |
||
| 64861 | 3740 |
instantiation poly :: (field) semidom_modulo |
| 64795 | 3741 |
begin |
| 65346 | 3742 |
|
| 64861 | 3743 |
definition modulo_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" |
| 65347 | 3744 |
where mod_poly_def: "f mod g = |
3745 |
(if g = 0 then f else pseudo_mod (smult ((1 / lead_coeff g) ^ (Suc (degree f) - degree g)) f) g)" |
|
3746 |
||
3747 |
instance |
|
3748 |
proof |
|
| 64861 | 3749 |
fix x y :: "'a poly" |
3750 |
show "x div y * y + x mod y = x" |
|
3751 |
proof (cases "y = 0") |
|
| 65347 | 3752 |
case True |
3753 |
then show ?thesis |
|
| 64861 | 3754 |
by (simp add: divide_poly_0 mod_poly_def) |
| 64795 | 3755 |
next |
3756 |
case False |
|
| 64861 | 3757 |
then have "pseudo_divmod (smult ((1 / lead_coeff y) ^ (Suc (degree x) - degree y)) x) y = |
| 65347 | 3758 |
(x div y, x mod y)" |
| 64861 | 3759 |
by (simp add: divide_poly_field mod_poly_def pseudo_mod_def) |
| 65347 | 3760 |
with False pseudo_divmod [OF False this] show ?thesis |
| 64861 | 3761 |
by (simp add: power_mult_distrib [symmetric] ac_simps) |
| 64795 | 3762 |
qed |
| 64861 | 3763 |
qed |
| 65346 | 3764 |
|
| 64861 | 3765 |
end |
| 65346 | 3766 |
|
| 64861 | 3767 |
lemma eucl_rel_poly: "eucl_rel_poly x y (x div y, x mod y)" |
| 65347 | 3768 |
unfolding eucl_rel_poly_iff |
3769 |
proof |
|
| 64861 | 3770 |
show "x = x div y * y + x mod y" |
3771 |
by (simp add: div_mult_mod_eq) |
|
| 64795 | 3772 |
show "if y = 0 then x div y = 0 else x mod y = 0 \<or> degree (x mod y) < degree y" |
3773 |
proof (cases "y = 0") |
|
| 65347 | 3774 |
case True |
3775 |
then show ?thesis by auto |
|
| 64795 | 3776 |
next |
3777 |
case False |
|
| 65347 | 3778 |
with pseudo_mod[OF this] show ?thesis |
3779 |
by (simp add: mod_poly_def) |
|
| 64795 | 3780 |
qed |
3781 |
qed |
|
3782 |
||
| 65347 | 3783 |
lemma div_poly_eq: "eucl_rel_poly x y (q, r) \<Longrightarrow> x div y = q" |
3784 |
for x :: "'a::field poly" |
|
3785 |
by (rule eucl_rel_poly_unique_div [OF eucl_rel_poly]) |
|
3786 |
||
3787 |
lemma mod_poly_eq: "eucl_rel_poly x y (q, r) \<Longrightarrow> x mod y = r" |
|
3788 |
for x :: "'a::field poly" |
|
| 64861 | 3789 |
by (rule eucl_rel_poly_unique_mod [OF eucl_rel_poly]) |
3790 |
||
3791 |
instance poly :: (field) ring_div |
|
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
3792 |
proof |
| 64795 | 3793 |
fix x y z :: "'a poly" |
3794 |
assume "y \<noteq> 0" |
|
| 65347 | 3795 |
with eucl_rel_poly [of x y] have "eucl_rel_poly (x + z * y) y (z + x div y, x mod y)" |
| 64795 | 3796 |
by (simp add: eucl_rel_poly_iff distrib_right) |
| 65346 | 3797 |
then show "(x + z * y) div y = z + x div y" |
| 64795 | 3798 |
by (rule div_poly_eq) |
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
3799 |
next |
| 64795 | 3800 |
fix x y z :: "'a poly" |
3801 |
assume "x \<noteq> 0" |
|
3802 |
show "(x * y) div (x * z) = y div z" |
|
3803 |
proof (cases "y \<noteq> 0 \<and> z \<noteq> 0") |
|
3804 |
have "\<And>x::'a poly. eucl_rel_poly x 0 (0, x)" |
|
3805 |
by (rule eucl_rel_poly_by_0) |
|
3806 |
then have [simp]: "\<And>x::'a poly. x div 0 = 0" |
|
3807 |
by (rule div_poly_eq) |
|
3808 |
have "\<And>x::'a poly. eucl_rel_poly 0 x (0, 0)" |
|
3809 |
by (rule eucl_rel_poly_0) |
|
3810 |
then have [simp]: "\<And>x::'a poly. 0 div x = 0" |
|
3811 |
by (rule div_poly_eq) |
|
| 65347 | 3812 |
case False |
3813 |
then show ?thesis by auto |
|
| 64795 | 3814 |
next |
| 65347 | 3815 |
case True |
3816 |
then have "y \<noteq> 0" and "z \<noteq> 0" by auto |
|
3817 |
with \<open>x \<noteq> 0\<close> have "\<And>q r. eucl_rel_poly y z (q, r) \<Longrightarrow> eucl_rel_poly (x * y) (x * z) (q, x * r)" |
|
3818 |
by (auto simp: eucl_rel_poly_iff algebra_simps) (rule classical, simp add: degree_mult_eq) |
|
| 64795 | 3819 |
moreover from eucl_rel_poly have "eucl_rel_poly y z (y div z, y mod z)" . |
3820 |
ultimately have "eucl_rel_poly (x * y) (x * z) (y div z, x * (y mod z))" . |
|
| 65347 | 3821 |
then show ?thesis |
3822 |
by (simp add: div_poly_eq) |
|
| 64795 | 3823 |
qed |
3824 |
qed |
|
3825 |
||
| 64811 | 3826 |
lemma div_pCons_eq: |
| 65347 | 3827 |
"pCons a p div q = |
3828 |
(if q = 0 then 0 |
|
3829 |
else pCons (coeff (pCons a (p mod q)) (degree q) / lead_coeff q) (p div q))" |
|
| 64811 | 3830 |
using eucl_rel_poly_pCons [OF eucl_rel_poly _ refl, of q a p] |
3831 |
by (auto intro: div_poly_eq) |
|
3832 |
||
3833 |
lemma mod_pCons_eq: |
|
| 65347 | 3834 |
"pCons a p mod q = |
3835 |
(if q = 0 then pCons a p |
|
3836 |
else pCons a (p mod q) - smult (coeff (pCons a (p mod q)) (degree q) / lead_coeff q) q)" |
|
| 64811 | 3837 |
using eucl_rel_poly_pCons [OF eucl_rel_poly _ refl, of q a p] |
3838 |
by (auto intro: mod_poly_eq) |
|
3839 |
||
3840 |
lemma div_mod_fold_coeffs: |
|
| 65347 | 3841 |
"(p div q, p mod q) = |
3842 |
(if q = 0 then (0, p) |
|
3843 |
else |
|
3844 |
fold_coeffs |
|
3845 |
(\<lambda>a (s, r). |
|
3846 |
let b = coeff (pCons a r) (degree q) / coeff q (degree q) |
|
3847 |
in (pCons b s, pCons a r - smult b q)) p (0, 0))" |
|
3848 |
by (rule sym, induct p) (auto simp: div_pCons_eq mod_pCons_eq Let_def) |
|
3849 |
||
3850 |
lemma degree_mod_less: "y \<noteq> 0 \<Longrightarrow> x mod y = 0 \<or> degree (x mod y) < degree y" |
|
3851 |
using eucl_rel_poly [of x y] unfolding eucl_rel_poly_iff by simp |
|
| 64795 | 3852 |
|
3853 |
lemma degree_mod_less': "b \<noteq> 0 \<Longrightarrow> a mod b \<noteq> 0 \<Longrightarrow> degree (a mod b) < degree b" |
|
3854 |
using degree_mod_less[of b a] by auto |
|
3855 |
||
| 65347 | 3856 |
lemma div_poly_less: |
3857 |
fixes x :: "'a::field poly" |
|
3858 |
assumes "degree x < degree y" |
|
3859 |
shows "x div y = 0" |
|
| 64795 | 3860 |
proof - |
| 65347 | 3861 |
from assms have "eucl_rel_poly x y (0, x)" |
| 64795 | 3862 |
by (simp add: eucl_rel_poly_iff) |
| 65347 | 3863 |
then show "x div y = 0" |
3864 |
by (rule div_poly_eq) |
|
| 64795 | 3865 |
qed |
3866 |
||
| 65347 | 3867 |
lemma mod_poly_less: |
3868 |
assumes "degree x < degree y" |
|
3869 |
shows "x mod y = x" |
|
| 64795 | 3870 |
proof - |
| 65347 | 3871 |
from assms have "eucl_rel_poly x y (0, x)" |
| 64795 | 3872 |
by (simp add: eucl_rel_poly_iff) |
| 65347 | 3873 |
then show "x mod y = x" |
3874 |
by (rule mod_poly_eq) |
|
| 64795 | 3875 |
qed |
3876 |
||
3877 |
lemma eucl_rel_poly_smult_left: |
|
| 65347 | 3878 |
"eucl_rel_poly x y (q, r) \<Longrightarrow> eucl_rel_poly (smult a x) y (smult a q, smult a r)" |
3879 |
by (simp add: eucl_rel_poly_iff smult_add_right) |
|
3880 |
||
3881 |
lemma div_smult_left: "(smult a x) div y = smult a (x div y)" |
|
3882 |
for x y :: "'a::field poly" |
|
| 64795 | 3883 |
by (rule div_poly_eq, rule eucl_rel_poly_smult_left, rule eucl_rel_poly) |
3884 |
||
3885 |
lemma mod_smult_left: "(smult a x) mod y = smult a (x mod y)" |
|
3886 |
by (rule mod_poly_eq, rule eucl_rel_poly_smult_left, rule eucl_rel_poly) |
|
3887 |
||
| 65347 | 3888 |
lemma poly_div_minus_left [simp]: "(- x) div y = - (x div y)" |
3889 |
for x y :: "'a::field poly" |
|
| 64795 | 3890 |
using div_smult_left [of "- 1::'a"] by simp |
3891 |
||
| 65347 | 3892 |
lemma poly_mod_minus_left [simp]: "(- x) mod y = - (x mod y)" |
3893 |
for x y :: "'a::field poly" |
|
| 64795 | 3894 |
using mod_smult_left [of "- 1::'a"] by simp |
3895 |
||
3896 |
lemma eucl_rel_poly_add_left: |
|
3897 |
assumes "eucl_rel_poly x y (q, r)" |
|
3898 |
assumes "eucl_rel_poly x' y (q', r')" |
|
3899 |
shows "eucl_rel_poly (x + x') y (q + q', r + r')" |
|
3900 |
using assms unfolding eucl_rel_poly_iff |
|
| 65347 | 3901 |
by (auto simp: algebra_simps degree_add_less) |
3902 |
||
3903 |
lemma poly_div_add_left: "(x + y) div z = x div z + y div z" |
|
3904 |
for x y z :: "'a::field poly" |
|
| 64795 | 3905 |
using eucl_rel_poly_add_left [OF eucl_rel_poly eucl_rel_poly] |
3906 |
by (rule div_poly_eq) |
|
3907 |
||
| 65347 | 3908 |
lemma poly_mod_add_left: "(x + y) mod z = x mod z + y mod z" |
3909 |
for x y z :: "'a::field poly" |
|
| 64795 | 3910 |
using eucl_rel_poly_add_left [OF eucl_rel_poly eucl_rel_poly] |
3911 |
by (rule mod_poly_eq) |
|
3912 |
||
| 65347 | 3913 |
lemma poly_div_diff_left: "(x - y) div z = x div z - y div z" |
3914 |
for x y z :: "'a::field poly" |
|
| 64795 | 3915 |
by (simp only: diff_conv_add_uminus poly_div_add_left poly_div_minus_left) |
3916 |
||
| 65347 | 3917 |
lemma poly_mod_diff_left: "(x - y) mod z = x mod z - y mod z" |
3918 |
for x y z :: "'a::field poly" |
|
| 64795 | 3919 |
by (simp only: diff_conv_add_uminus poly_mod_add_left poly_mod_minus_left) |
3920 |
||
3921 |
lemma eucl_rel_poly_smult_right: |
|
| 65347 | 3922 |
"a \<noteq> 0 \<Longrightarrow> eucl_rel_poly x y (q, r) \<Longrightarrow> eucl_rel_poly x (smult a y) (smult (inverse a) q, r)" |
3923 |
by (simp add: eucl_rel_poly_iff) |
|
3924 |
||
3925 |
lemma div_smult_right: "a \<noteq> 0 \<Longrightarrow> x div (smult a y) = smult (inverse a) (x div y)" |
|
3926 |
for x y :: "'a::field poly" |
|
| 64795 | 3927 |
by (rule div_poly_eq, erule eucl_rel_poly_smult_right, rule eucl_rel_poly) |
3928 |
||
3929 |
lemma mod_smult_right: "a \<noteq> 0 \<Longrightarrow> x mod (smult a y) = x mod y" |
|
3930 |
by (rule mod_poly_eq, erule eucl_rel_poly_smult_right, rule eucl_rel_poly) |
|
3931 |
||
| 65347 | 3932 |
lemma poly_div_minus_right [simp]: "x div (- y) = - (x div y)" |
3933 |
for x y :: "'a::field poly" |
|
| 64795 | 3934 |
using div_smult_right [of "- 1::'a"] by (simp add: nonzero_inverse_minus_eq) |
3935 |
||
| 65347 | 3936 |
lemma poly_mod_minus_right [simp]: "x mod (- y) = x mod y" |
3937 |
for x y :: "'a::field poly" |
|
| 64795 | 3938 |
using mod_smult_right [of "- 1::'a"] by simp |
3939 |
||
3940 |
lemma eucl_rel_poly_mult: |
|
| 65347 | 3941 |
"eucl_rel_poly x y (q, r) \<Longrightarrow> eucl_rel_poly q z (q', r') \<Longrightarrow> |
3942 |
eucl_rel_poly x (y * z) (q', y * r' + r)" |
|
3943 |
apply (cases "z = 0", simp add: eucl_rel_poly_iff) |
|
3944 |
apply (cases "y = 0", simp add: eucl_rel_poly_by_0_iff eucl_rel_poly_0_iff) |
|
3945 |
apply (cases "r = 0") |
|
3946 |
apply (cases "r' = 0") |
|
3947 |
apply (simp add: eucl_rel_poly_iff) |
|
3948 |
apply (simp add: eucl_rel_poly_iff field_simps degree_mult_eq) |
|
3949 |
apply (cases "r' = 0") |
|
3950 |
apply (simp add: eucl_rel_poly_iff degree_mult_eq) |
|
3951 |
apply (simp add: eucl_rel_poly_iff field_simps) |
|
3952 |
apply (simp add: degree_mult_eq degree_add_less) |
|
3953 |
done |
|
3954 |
||
3955 |
lemma poly_div_mult_right: "x div (y * z) = (x div y) div z" |
|
3956 |
for x y z :: "'a::field poly" |
|
| 64795 | 3957 |
by (rule div_poly_eq, rule eucl_rel_poly_mult, (rule eucl_rel_poly)+) |
3958 |
||
| 65347 | 3959 |
lemma poly_mod_mult_right: "x mod (y * z) = y * (x div y mod z) + x mod y" |
3960 |
for x y z :: "'a::field poly" |
|
| 64795 | 3961 |
by (rule mod_poly_eq, rule eucl_rel_poly_mult, (rule eucl_rel_poly)+) |
3962 |
||
3963 |
lemma mod_pCons: |
|
| 65347 | 3964 |
fixes a :: "'a::field" |
3965 |
and x y :: "'a::field poly" |
|
| 64795 | 3966 |
assumes y: "y \<noteq> 0" |
| 65347 | 3967 |
defines "b \<equiv> coeff (pCons a (x mod y)) (degree y) / coeff y (degree y)" |
3968 |
shows "(pCons a x) mod y = pCons a (x mod y) - smult b y" |
|
3969 |
unfolding b_def |
|
3970 |
by (rule mod_poly_eq, rule eucl_rel_poly_pCons [OF eucl_rel_poly y refl]) |
|
| 64795 | 3971 |
|
| 65346 | 3972 |
|
| 64795 | 3973 |
subsubsection \<open>List-based versions for fast implementation\<close> |
3974 |
(* Subsection by: |
|
3975 |
Sebastiaan Joosten |
|
3976 |
René Thiemann |
|
3977 |
Akihisa Yamada |
|
3978 |
*) |
|
| 65347 | 3979 |
fun minus_poly_rev_list :: "'a :: group_add list \<Rightarrow> 'a list \<Rightarrow> 'a list" |
3980 |
where |
|
3981 |
"minus_poly_rev_list (x # xs) (y # ys) = (x - y) # (minus_poly_rev_list xs ys)" |
|
3982 |
| "minus_poly_rev_list xs [] = xs" |
|
3983 |
| "minus_poly_rev_list [] (y # ys) = []" |
|
3984 |
||
3985 |
fun pseudo_divmod_main_list :: |
|
3986 |
"'a::comm_ring_1 \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> nat \<Rightarrow> 'a list \<times> 'a list" |
|
3987 |
where |
|
3988 |
"pseudo_divmod_main_list lc q r d (Suc n) = |
|
3989 |
(let |
|
3990 |
rr = map (op * lc) r; |
|
3991 |
a = hd r; |
|
3992 |
qqq = cCons a (map (op * lc) q); |
|
3993 |
rrr = tl (if a = 0 then rr else minus_poly_rev_list rr (map (op * a) d)) |
|
3994 |
in pseudo_divmod_main_list lc qqq rrr d n)" |
|
3995 |
| "pseudo_divmod_main_list lc q r d 0 = (q, r)" |
|
3996 |
||
3997 |
fun pseudo_mod_main_list :: "'a::comm_ring_1 \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> nat \<Rightarrow> 'a list" |
|
3998 |
where |
|
3999 |
"pseudo_mod_main_list lc r d (Suc n) = |
|
4000 |
(let |
|
4001 |
rr = map (op * lc) r; |
|
4002 |
a = hd r; |
|
4003 |
rrr = tl (if a = 0 then rr else minus_poly_rev_list rr (map (op * a) d)) |
|
4004 |
in pseudo_mod_main_list lc rrr d n)" |
|
4005 |
| "pseudo_mod_main_list lc r d 0 = r" |
|
4006 |
||
4007 |
||
4008 |
fun divmod_poly_one_main_list :: |
|
4009 |
"'a::comm_ring_1 list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> nat \<Rightarrow> 'a list \<times> 'a list" |
|
4010 |
where |
|
4011 |
"divmod_poly_one_main_list q r d (Suc n) = |
|
4012 |
(let |
|
4013 |
a = hd r; |
|
4014 |
qqq = cCons a q; |
|
4015 |
rr = tl (if a = 0 then r else minus_poly_rev_list r (map (op * a) d)) |
|
4016 |
in divmod_poly_one_main_list qqq rr d n)" |
|
4017 |
| "divmod_poly_one_main_list q r d 0 = (q, r)" |
|
4018 |
||
4019 |
fun mod_poly_one_main_list :: "'a::comm_ring_1 list \<Rightarrow> 'a list \<Rightarrow> nat \<Rightarrow> 'a list" |
|
4020 |
where |
|
4021 |
"mod_poly_one_main_list r d (Suc n) = |
|
4022 |
(let |
|
4023 |
a = hd r; |
|
4024 |
rr = tl (if a = 0 then r else minus_poly_rev_list r (map (op * a) d)) |
|
4025 |
in mod_poly_one_main_list rr d n)" |
|
4026 |
| "mod_poly_one_main_list r d 0 = r" |
|
4027 |
||
4028 |
definition pseudo_divmod_list :: "'a::comm_ring_1 list \<Rightarrow> 'a list \<Rightarrow> 'a list \<times> 'a list" |
|
4029 |
where "pseudo_divmod_list p q = |
|
4030 |
(if q = [] then ([], p) |
|
4031 |
else |
|
4032 |
(let rq = rev q; |
|
4033 |
(qu,re) = pseudo_divmod_main_list (hd rq) [] (rev p) rq (1 + length p - length q) |
|
4034 |
in (qu, rev re)))" |
|
4035 |
||
4036 |
definition pseudo_mod_list :: "'a::comm_ring_1 list \<Rightarrow> 'a list \<Rightarrow> 'a list" |
|
4037 |
where "pseudo_mod_list p q = |
|
4038 |
(if q = [] then p |
|
4039 |
else |
|
4040 |
(let |
|
4041 |
rq = rev q; |
|
4042 |
re = pseudo_mod_main_list (hd rq) (rev p) rq (1 + length p - length q) |
|
4043 |
in rev re))" |
|
4044 |
||
4045 |
lemma minus_zero_does_nothing: "minus_poly_rev_list x (map (op * 0) y) = x" |
|
4046 |
for x :: "'a::ring list" |
|
4047 |
by (induct x y rule: minus_poly_rev_list.induct) auto |
|
4048 |
||
4049 |
lemma length_minus_poly_rev_list [simp]: "length (minus_poly_rev_list xs ys) = length xs" |
|
4050 |
by (induct xs ys rule: minus_poly_rev_list.induct) auto |
|
| 64795 | 4051 |
|
4052 |
lemma if_0_minus_poly_rev_list: |
|
| 65347 | 4053 |
"(if a = 0 then x else minus_poly_rev_list x (map (op * a) y)) = |
4054 |
minus_poly_rev_list x (map (op * a) y)" |
|
4055 |
for a :: "'a::ring" |
|
4056 |
by(cases "a = 0") (simp_all add: minus_zero_does_nothing) |
|
4057 |
||
4058 |
lemma Poly_append: "Poly (a @ b) = Poly a + monom 1 (length a) * Poly b" |
|
4059 |
for a :: "'a::comm_semiring_1 list" |
|
4060 |
by (induct a) (auto simp: monom_0 monom_Suc) |
|
4061 |
||
4062 |
lemma minus_poly_rev_list: "length p \<ge> length q \<Longrightarrow> |
|
4063 |
Poly (rev (minus_poly_rev_list (rev p) (rev q))) = |
|
4064 |
Poly p - monom 1 (length p - length q) * Poly q" |
|
4065 |
for p q :: "'a :: comm_ring_1 list" |
|
| 64795 | 4066 |
proof (induct "rev p" "rev q" arbitrary: p q rule: minus_poly_rev_list.induct) |
| 65346 | 4067 |
case (1 x xs y ys) |
| 65347 | 4068 |
then have "length (rev q) \<le> length (rev p)" |
4069 |
by simp |
|
4070 |
from this[folded 1(2,3)] have ys_xs: "length ys \<le> length xs" |
|
4071 |
by simp |
|
4072 |
then have *: "Poly (rev (minus_poly_rev_list xs ys)) = |
|
4073 |
Poly (rev xs) - monom 1 (length xs - length ys) * Poly (rev ys)" |
|
4074 |
by (subst "1.hyps"(1)[of "rev xs" "rev ys", unfolded rev_rev_ident length_rev]) auto |
|
4075 |
have "Poly p - monom 1 (length p - length q) * Poly q = |
|
4076 |
Poly (rev (rev p)) - monom 1 (length (rev (rev p)) - length (rev (rev q))) * Poly (rev (rev q))" |
|
| 64795 | 4077 |
by simp |
| 65347 | 4078 |
also have "\<dots> = |
4079 |
Poly (rev (x # xs)) - monom 1 (length (x # xs) - length (y # ys)) * Poly (rev (y # ys))" |
|
| 64795 | 4080 |
unfolding 1(2,3) by simp |
| 65347 | 4081 |
also from ys_xs have "\<dots> = |
4082 |
Poly (rev xs) + monom x (length xs) - |
|
4083 |
(monom 1 (length xs - length ys) * Poly (rev ys) + monom y (length xs))" |
|
4084 |
by (simp add: Poly_append distrib_left mult_monom smult_monom) |
|
| 64795 | 4085 |
also have "\<dots> = Poly (rev (minus_poly_rev_list xs ys)) + monom (x - y) (length xs)" |
| 65347 | 4086 |
unfolding * diff_monom[symmetric] by simp |
| 64795 | 4087 |
finally show ?case |
| 65347 | 4088 |
by (simp add: 1(2,3)[symmetric] smult_monom Poly_append) |
| 64795 | 4089 |
qed auto |
4090 |
||
4091 |
lemma smult_monom_mult: "smult a (monom b n * f) = monom (a * b) n * f" |
|
4092 |
using smult_monom [of a _ n] by (metis mult_smult_left) |
|
4093 |
||
4094 |
lemma head_minus_poly_rev_list: |
|
| 65347 | 4095 |
"length d \<le> length r \<Longrightarrow> d \<noteq> [] \<Longrightarrow> |
4096 |
hd (minus_poly_rev_list (map (op * (last d)) r) (map (op * (hd r)) (rev d))) = 0" |
|
4097 |
for d r :: "'a::comm_ring list" |
|
4098 |
proof (induct r) |
|
4099 |
case Nil |
|
4100 |
then show ?case by simp |
|
4101 |
next |
|
| 64795 | 4102 |
case (Cons a rs) |
| 65347 | 4103 |
then show ?case by (cases "rev d") (simp_all add: ac_simps) |
4104 |
qed |
|
| 64795 | 4105 |
|
4106 |
lemma Poly_map: "Poly (map (op * a) p) = smult a (Poly p)" |
|
4107 |
proof (induct p) |
|
| 65347 | 4108 |
case Nil |
4109 |
then show ?case by simp |
|
4110 |
next |
|
4111 |
case (Cons x xs) |
|
4112 |
then show ?case by (cases "Poly xs = 0") auto |
|
4113 |
qed |
|
| 64795 | 4114 |
|
4115 |
lemma last_coeff_is_hd: "xs \<noteq> [] \<Longrightarrow> coeff (Poly xs) (length xs - 1) = hd (rev xs)" |
|
4116 |
by (simp_all add: hd_conv_nth rev_nth nth_default_nth nth_append) |
|
4117 |
||
| 65347 | 4118 |
lemma pseudo_divmod_main_list_invar: |
4119 |
assumes leading_nonzero: "last d \<noteq> 0" |
|
4120 |
and lc: "last d = lc" |
|
4121 |
and "d \<noteq> []" |
|
4122 |
and "pseudo_divmod_main_list lc q (rev r) (rev d) n = (q', rev r')" |
|
4123 |
and "n = 1 + length r - length d" |
|
4124 |
shows "pseudo_divmod_main lc (monom 1 n * Poly q) (Poly r) (Poly d) (length r - 1) n = |
|
4125 |
(Poly q', Poly r')" |
|
4126 |
using assms(4-) |
|
4127 |
proof (induct n arbitrary: r q) |
|
4128 |
case (Suc n) |
|
4129 |
from Suc.prems have *: "\<not> Suc (length r) \<le> length d" |
|
4130 |
by simp |
|
4131 |
with \<open>d \<noteq> []\<close> have "r \<noteq> []" |
|
4132 |
using Suc_leI length_greater_0_conv list.size(3) by fastforce |
|
| 64795 | 4133 |
let ?a = "(hd (rev r))" |
4134 |
let ?rr = "map (op * lc) (rev r)" |
|
4135 |
let ?rrr = "rev (tl (minus_poly_rev_list ?rr (map (op * ?a) (rev d))))" |
|
4136 |
let ?qq = "cCons ?a (map (op * lc) q)" |
|
| 65347 | 4137 |
from * Suc(3) have n: "n = (1 + length r - length d - 1)" |
4138 |
by simp |
|
4139 |
from * have rr_val:"(length ?rrr) = (length r - 1)" |
|
4140 |
by auto |
|
4141 |
with \<open>r \<noteq> []\<close> * have rr_smaller: "(1 + length r - length d - 1) = (1 + length ?rrr - length d)" |
|
4142 |
by auto |
|
4143 |
from * have id: "Suc (length r) - length d = Suc (length r - length d)" |
|
4144 |
by auto |
|
4145 |
from Suc.prems * |
|
| 64795 | 4146 |
have "pseudo_divmod_main_list lc ?qq (rev ?rrr) (rev d) (1 + length r - length d - 1) = (q', rev r')" |
| 65347 | 4147 |
by (simp add: Let_def if_0_minus_poly_rev_list id) |
4148 |
with n have v: "pseudo_divmod_main_list lc ?qq (rev ?rrr) (rev d) n = (q', rev r')" |
|
4149 |
by auto |
|
4150 |
from * have sucrr:"Suc (length r) - length d = Suc (length r - length d)" |
|
4151 |
using Suc_diff_le not_less_eq_eq by blast |
|
4152 |
from Suc(3) \<open>r \<noteq> []\<close> have n_ok : "n = 1 + (length ?rrr) - length d" |
|
4153 |
by simp |
|
| 65346 | 4154 |
have cong: "\<And>x1 x2 x3 x4 y1 y2 y3 y4. x1 = y1 \<Longrightarrow> x2 = y2 \<Longrightarrow> x3 = y3 \<Longrightarrow> x4 = y4 \<Longrightarrow> |
| 65347 | 4155 |
pseudo_divmod_main lc x1 x2 x3 x4 n = pseudo_divmod_main lc y1 y2 y3 y4 n" |
4156 |
by simp |
|
4157 |
have hd_rev: "coeff (Poly r) (length r - Suc 0) = hd (rev r)" |
|
4158 |
using last_coeff_is_hd[OF \<open>r \<noteq> []\<close>] by simp |
|
4159 |
show ?case |
|
4160 |
unfolding Suc.hyps(1)[OF v n_ok, symmetric] pseudo_divmod_main.simps Let_def |
|
| 64795 | 4161 |
proof (rule cong[OF _ _ refl], goal_cases) |
| 65346 | 4162 |
case 1 |
| 65347 | 4163 |
show ?case |
4164 |
by (simp add: monom_Suc hd_rev[symmetric] smult_monom Poly_map) |
|
| 64795 | 4165 |
next |
| 65346 | 4166 |
case 2 |
4167 |
show ?case |
|
| 64795 | 4168 |
proof (subst Poly_on_rev_starting_with_0, goal_cases) |
4169 |
show "hd (minus_poly_rev_list (map (op * lc) (rev r)) (map (op * (hd (rev r))) (rev d))) = 0" |
|
| 65347 | 4170 |
by (fold lc, subst head_minus_poly_rev_list, insert * \<open>d \<noteq> []\<close>, auto) |
4171 |
from * have "length d \<le> length r" |
|
4172 |
by simp |
|
| 64795 | 4173 |
then show "smult lc (Poly r) - monom (coeff (Poly r) (length r - 1)) n * Poly d = |
| 65347 | 4174 |
Poly (rev (minus_poly_rev_list (map (op * lc) (rev r)) (map (op * (hd (rev r))) (rev d))))" |
| 64795 | 4175 |
by (fold rev_map) (auto simp add: n smult_monom_mult Poly_map hd_rev [symmetric] |
| 65347 | 4176 |
minus_poly_rev_list) |
| 64795 | 4177 |
qed |
4178 |
qed simp |
|
4179 |
qed simp |
|
4180 |
||
| 65390 | 4181 |
lemma pseudo_divmod_impl [code]: |
4182 |
"pseudo_divmod f g = map_prod poly_of_list poly_of_list (pseudo_divmod_list (coeffs f) (coeffs g))" |
|
4183 |
for f g :: "'a::comm_ring_1 poly" |
|
| 65347 | 4184 |
proof (cases "g = 0") |
4185 |
case False |
|
| 65390 | 4186 |
then have "last (coeffs g) \<noteq> 0" |
4187 |
and "last (coeffs g) = lead_coeff g" |
|
4188 |
and "coeffs g \<noteq> []" |
|
4189 |
by (simp_all add: last_coeffs_eq_coeff_degree) |
|
4190 |
moreover obtain q r where qr: "pseudo_divmod_main_list |
|
4191 |
(last (coeffs g)) (rev []) |
|
4192 |
(rev (coeffs f)) (rev (coeffs g)) |
|
4193 |
(1 + length (coeffs f) - |
|
4194 |
length (coeffs g)) = (q, rev (rev r))" |
|
| 65347 | 4195 |
by force |
| 65390 | 4196 |
ultimately have "(Poly q, Poly (rev r)) = pseudo_divmod_main (lead_coeff g) 0 f g |
4197 |
(length (coeffs f) - Suc 0) (Suc (length (coeffs f)) - length (coeffs g))" |
|
4198 |
by (subst pseudo_divmod_main_list_invar [symmetric]) auto |
|
4199 |
moreover have "pseudo_divmod_main_list |
|
4200 |
(hd (rev (coeffs g))) [] |
|
4201 |
(rev (coeffs f)) (rev (coeffs g)) |
|
4202 |
(1 + length (coeffs f) - |
|
4203 |
length (coeffs g)) = (q, r)" |
|
4204 |
using qr hd_rev [OF \<open>coeffs g \<noteq> []\<close>] by simp |
|
4205 |
ultimately show ?thesis |
|
4206 |
by (auto simp: degree_eq_length_coeffs pseudo_divmod_def pseudo_divmod_list_def Let_def) |
|
| 64795 | 4207 |
next |
4208 |
case True |
|
| 65347 | 4209 |
then show ?thesis |
| 65390 | 4210 |
by (auto simp add: pseudo_divmod_def pseudo_divmod_list_def) |
| 64795 | 4211 |
qed |
4212 |
||
| 65347 | 4213 |
lemma pseudo_mod_main_list: |
4214 |
"snd (pseudo_divmod_main_list l q xs ys n) = pseudo_mod_main_list l xs ys n" |
|
4215 |
by (induct n arbitrary: l q xs ys) (auto simp: Let_def) |
|
4216 |
||
4217 |
lemma pseudo_mod_impl[code]: "pseudo_mod f g = poly_of_list (pseudo_mod_list (coeffs f) (coeffs g))" |
|
| 64795 | 4218 |
proof - |
| 65346 | 4219 |
have snd_case: "\<And>f g p. snd ((\<lambda>(x,y). (f x, g y)) p) = g (snd p)" |
| 64795 | 4220 |
by auto |
4221 |
show ?thesis |
|
| 65347 | 4222 |
unfolding pseudo_mod_def pseudo_divmod_impl pseudo_divmod_list_def |
4223 |
pseudo_mod_list_def Let_def |
|
4224 |
by (simp add: snd_case pseudo_mod_main_list) |
|
| 64795 | 4225 |
qed |
4226 |
||
4227 |
||
4228 |
subsubsection \<open>Improved Code-Equations for Polynomial (Pseudo) Division\<close> |
|
4229 |
||
| 64811 | 4230 |
lemma pdivmod_pdivmodrel: "eucl_rel_poly p q (r, s) \<longleftrightarrow> (p div q, p mod q) = (r, s)" |
4231 |
by (metis eucl_rel_poly eucl_rel_poly_unique) |
|
4232 |
||
| 65347 | 4233 |
lemma pdivmod_via_pseudo_divmod: |
4234 |
"(f div g, f mod g) = |
|
4235 |
(if g = 0 then (0, f) |
|
4236 |
else |
|
4237 |
let |
|
4238 |
ilc = inverse (coeff g (degree g)); |
|
4239 |
h = smult ilc g; |
|
4240 |
(q,r) = pseudo_divmod f h |
|
4241 |
in (smult ilc q, r))" |
|
4242 |
(is "?l = ?r") |
|
| 64795 | 4243 |
proof (cases "g = 0") |
| 65347 | 4244 |
case True |
4245 |
then show ?thesis by simp |
|
4246 |
next |
|
| 64795 | 4247 |
case False |
4248 |
define lc where "lc = inverse (coeff g (degree g))" |
|
4249 |
define h where "h = smult lc g" |
|
| 65347 | 4250 |
from False have h1: "coeff h (degree h) = 1" and lc: "lc \<noteq> 0" |
4251 |
by (auto simp: h_def lc_def) |
|
4252 |
then have h0: "h \<noteq> 0" |
|
4253 |
by auto |
|
4254 |
obtain q r where p: "pseudo_divmod f h = (q, r)" |
|
4255 |
by force |
|
| 65346 | 4256 |
from False have id: "?r = (smult lc q, r)" |
| 65347 | 4257 |
by (auto simp: Let_def h_def[symmetric] lc_def[symmetric] p) |
| 65346 | 4258 |
from pseudo_divmod[OF h0 p, unfolded h1] |
| 65347 | 4259 |
have f: "f = h * q + r" and r: "r = 0 \<or> degree r < degree h" |
4260 |
by auto |
|
4261 |
from f r h0 have "eucl_rel_poly f h (q, r)" |
|
4262 |
by (auto simp: eucl_rel_poly_iff) |
|
4263 |
then have "(f div h, f mod h) = (q, r)" |
|
4264 |
by (simp add: pdivmod_pdivmodrel) |
|
4265 |
with lc have "(f div g, f mod g) = (smult lc q, r)" |
|
4266 |
by (auto simp: h_def div_smult_right[OF lc] mod_smult_right[OF lc]) |
|
4267 |
with id show ?thesis |
|
4268 |
by auto |
|
4269 |
qed |
|
4270 |
||
4271 |
lemma pdivmod_via_pseudo_divmod_list: |
|
4272 |
"(f div g, f mod g) = |
|
4273 |
(let cg = coeffs g in |
|
4274 |
if cg = [] then (0, f) |
|
4275 |
else |
|
4276 |
let |
|
4277 |
cf = coeffs f; |
|
4278 |
ilc = inverse (last cg); |
|
4279 |
ch = map (op * ilc) cg; |
|
4280 |
(q, r) = pseudo_divmod_main_list 1 [] (rev cf) (rev ch) (1 + length cf - length cg) |
|
4281 |
in (poly_of_list (map (op * ilc) q), poly_of_list (rev r)))" |
|
| 64795 | 4282 |
proof - |
| 65347 | 4283 |
note d = pdivmod_via_pseudo_divmod pseudo_divmod_impl pseudo_divmod_list_def |
| 64795 | 4284 |
show ?thesis |
4285 |
proof (cases "g = 0") |
|
| 65347 | 4286 |
case True |
4287 |
with d show ?thesis by auto |
|
| 64795 | 4288 |
next |
4289 |
case False |
|
4290 |
define ilc where "ilc = inverse (coeff g (degree g))" |
|
| 65347 | 4291 |
from False have ilc: "ilc \<noteq> 0" |
4292 |
by (auto simp: ilc_def) |
|
4293 |
with False have id: "g = 0 \<longleftrightarrow> False" "coeffs g = [] \<longleftrightarrow> False" |
|
| 65346 | 4294 |
"last (coeffs g) = coeff g (degree g)" |
| 65347 | 4295 |
"coeffs (smult ilc g) = [] \<longleftrightarrow> False" |
| 65346 | 4296 |
by (auto simp: last_coeffs_eq_coeff_degree) |
4297 |
have id2: "hd (rev (coeffs (smult ilc g))) = 1" |
|
| 64795 | 4298 |
by (subst hd_rev, insert id ilc, auto simp: coeffs_smult, subst last_map, auto simp: id ilc_def) |
| 65346 | 4299 |
have id3: "length (coeffs (smult ilc g)) = length (coeffs g)" |
| 65347 | 4300 |
"rev (coeffs (smult ilc g)) = rev (map (op * ilc) (coeffs g))" |
4301 |
unfolding coeffs_smult using ilc by auto |
|
4302 |
obtain q r where pair: |
|
4303 |
"pseudo_divmod_main_list 1 [] (rev (coeffs f)) (rev (map (op * ilc) (coeffs g))) |
|
4304 |
(1 + length (coeffs f) - length (coeffs g)) = (q, r)" |
|
4305 |
by force |
|
4306 |
show ?thesis |
|
4307 |
unfolding d Let_def id if_False ilc_def[symmetric] map_prod_def[symmetric] id2 |
|
4308 |
unfolding id3 pair map_prod_def split |
|
4309 |
by (auto simp: Poly_map) |
|
| 64795 | 4310 |
qed |
4311 |
qed |
|
4312 |
||
4313 |
lemma pseudo_divmod_main_list_1: "pseudo_divmod_main_list 1 = divmod_poly_one_main_list" |
|
4314 |
proof (intro ext, goal_cases) |
|
4315 |
case (1 q r d n) |
|
| 65347 | 4316 |
have *: "map (op * 1) xs = xs" for xs :: "'a list" |
4317 |
by (induct xs) auto |
|
4318 |
show ?case |
|
4319 |
by (induct n arbitrary: q r d) (auto simp: * Let_def) |
|
| 64795 | 4320 |
qed |
4321 |
||
| 65347 | 4322 |
fun divide_poly_main_list :: "'a::idom_divide \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> nat \<Rightarrow> 'a list" |
4323 |
where |
|
4324 |
"divide_poly_main_list lc q r d (Suc n) = |
|
4325 |
(let |
|
4326 |
cr = hd r |
|
4327 |
in if cr = 0 then divide_poly_main_list lc (cCons cr q) (tl r) d n else let |
|
4328 |
a = cr div lc; |
|
4329 |
qq = cCons a q; |
|
4330 |
rr = minus_poly_rev_list r (map (op * a) d) |
|
4331 |
in if hd rr = 0 then divide_poly_main_list lc qq (tl rr) d n else [])" |
|
4332 |
| "divide_poly_main_list lc q r d 0 = q" |
|
4333 |
||
4334 |
lemma divide_poly_main_list_simp [simp]: |
|
4335 |
"divide_poly_main_list lc q r d (Suc n) = |
|
4336 |
(let |
|
4337 |
cr = hd r; |
|
4338 |
a = cr div lc; |
|
4339 |
qq = cCons a q; |
|
4340 |
rr = minus_poly_rev_list r (map (op * a) d) |
|
| 64795 | 4341 |
in if hd rr = 0 then divide_poly_main_list lc qq (tl rr) d n else [])" |
4342 |
by (simp add: Let_def minus_zero_does_nothing) |
|
4343 |
||
4344 |
declare divide_poly_main_list.simps(1)[simp del] |
|
4345 |
||
| 65347 | 4346 |
definition divide_poly_list :: "'a::idom_divide poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" |
4347 |
where "divide_poly_list f g = |
|
4348 |
(let cg = coeffs g in |
|
4349 |
if cg = [] then g |
|
4350 |
else |
|
4351 |
let |
|
4352 |
cf = coeffs f; |
|
4353 |
cgr = rev cg |
|
4354 |
in poly_of_list (divide_poly_main_list (hd cgr) [] (rev cf) cgr (1 + length cf - length cg)))" |
|
| 64795 | 4355 |
|
| 64811 | 4356 |
lemmas pdivmod_via_divmod_list = pdivmod_via_pseudo_divmod_list[unfolded pseudo_divmod_main_list_1] |
| 64795 | 4357 |
|
4358 |
lemma mod_poly_one_main_list: "snd (divmod_poly_one_main_list q r d n) = mod_poly_one_main_list r d n" |
|
| 65347 | 4359 |
by (induct n arbitrary: q r d) (auto simp: Let_def) |
4360 |
||
4361 |
lemma mod_poly_code [code]: |
|
4362 |
"f mod g = |
|
4363 |
(let cg = coeffs g in |
|
4364 |
if cg = [] then f |
|
4365 |
else |
|
4366 |
let |
|
4367 |
cf = coeffs f; |
|
4368 |
ilc = inverse (last cg); |
|
4369 |
ch = map (op * ilc) cg; |
|
4370 |
r = mod_poly_one_main_list (rev cf) (rev ch) (1 + length cf - length cg) |
|
4371 |
in poly_of_list (rev r))" |
|
4372 |
(is "_ = ?rhs") |
|
| 64795 | 4373 |
proof - |
| 65347 | 4374 |
have "snd (f div g, f mod g) = ?rhs" |
4375 |
unfolding pdivmod_via_divmod_list Let_def mod_poly_one_main_list [symmetric, of _ _ _ Nil] |
|
4376 |
by (auto split: prod.splits) |
|
4377 |
then show ?thesis by simp |
|
| 64795 | 4378 |
qed |
4379 |
||
| 65347 | 4380 |
definition div_field_poly_impl :: "'a :: field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" |
4381 |
where "div_field_poly_impl f g = |
|
4382 |
(let cg = coeffs g in |
|
4383 |
if cg = [] then 0 |
|
4384 |
else |
|
4385 |
let |
|
4386 |
cf = coeffs f; |
|
4387 |
ilc = inverse (last cg); |
|
4388 |
ch = map (op * ilc) cg; |
|
4389 |
q = fst (divmod_poly_one_main_list [] (rev cf) (rev ch) (1 + length cf - length cg)) |
|
4390 |
in poly_of_list ((map (op * ilc) q)))" |
|
| 64795 | 4391 |
|
| 65346 | 4392 |
text \<open>We do not declare the following lemma as code equation, since then polynomial division |
4393 |
on non-fields will no longer be executable. However, a code-unfold is possible, since |
|
| 64795 | 4394 |
\<open>div_field_poly_impl\<close> is a bit more efficient than the generic polynomial division.\<close> |
4395 |
lemma div_field_poly_impl[code_unfold]: "op div = div_field_poly_impl" |
|
4396 |
proof (intro ext) |
|
4397 |
fix f g :: "'a poly" |
|
| 65347 | 4398 |
have "fst (f div g, f mod g) = div_field_poly_impl f g" |
4399 |
unfolding div_field_poly_impl_def pdivmod_via_divmod_list Let_def |
|
4400 |
by (auto split: prod.splits) |
|
| 64811 | 4401 |
then show "f div g = div_field_poly_impl f g" |
4402 |
by simp |
|
| 64795 | 4403 |
qed |
4404 |
||
4405 |
lemma divide_poly_main_list: |
|
4406 |
assumes lc0: "lc \<noteq> 0" |
|
| 65347 | 4407 |
and lc: "last d = lc" |
4408 |
and d: "d \<noteq> []" |
|
4409 |
and "n = (1 + length r - length d)" |
|
4410 |
shows "Poly (divide_poly_main_list lc q (rev r) (rev d) n) = |
|
4411 |
divide_poly_main lc (monom 1 n * Poly q) (Poly r) (Poly d) (length r - 1) n" |
|
4412 |
using assms(4-) |
|
4413 |
proof (induct "n" arbitrary: r q) |
|
4414 |
case (Suc n) |
|
4415 |
from Suc.prems have ifCond: "\<not> Suc (length r) \<le> length d" |
|
4416 |
by simp |
|
4417 |
with d have r: "r \<noteq> []" |
|
4418 |
using Suc_leI length_greater_0_conv list.size(3) by fastforce |
|
4419 |
then obtain rr lcr where r: "r = rr @ [lcr]" |
|
4420 |
by (cases r rule: rev_cases) auto |
|
| 65346 | 4421 |
from d lc obtain dd where d: "d = dd @ [lc]" |
| 65347 | 4422 |
by (cases d rule: rev_cases) auto |
4423 |
from Suc(2) ifCond have n: "n = 1 + length rr - length d" |
|
4424 |
by (auto simp: r) |
|
4425 |
from ifCond have len: "length dd \<le> length rr" |
|
4426 |
by (simp add: r d) |
|
| 64795 | 4427 |
show ?case |
4428 |
proof (cases "lcr div lc * lc = lcr") |
|
4429 |
case False |
|
| 65347 | 4430 |
with r d show ?thesis |
4431 |
unfolding Suc(2)[symmetric] |
|
| 64795 | 4432 |
by (auto simp add: Let_def nth_default_append) |
4433 |
next |
|
4434 |
case True |
|
| 65347 | 4435 |
with r d have id: |
4436 |
"?thesis \<longleftrightarrow> |
|
4437 |
Poly (divide_poly_main_list lc (cCons (lcr div lc) q) |
|
4438 |
(rev (rev (minus_poly_rev_list (rev rr) (rev (map (op * (lcr div lc)) dd))))) (rev d) n) = |
|
4439 |
divide_poly_main lc |
|
4440 |
(monom 1 (Suc n) * Poly q + monom (lcr div lc) n) |
|
4441 |
(Poly r - monom (lcr div lc) n * Poly d) |
|
4442 |
(Poly d) (length rr - 1) n" |
|
4443 |
by (cases r rule: rev_cases; cases "d" rule: rev_cases) |
|
4444 |
(auto simp add: Let_def rev_map nth_default_append) |
|
| 65346 | 4445 |
have cong: "\<And>x1 x2 x3 x4 y1 y2 y3 y4. x1 = y1 \<Longrightarrow> x2 = y2 \<Longrightarrow> x3 = y3 \<Longrightarrow> x4 = y4 \<Longrightarrow> |
| 65347 | 4446 |
divide_poly_main lc x1 x2 x3 x4 n = divide_poly_main lc y1 y2 y3 y4 n" |
4447 |
by simp |
|
4448 |
show ?thesis |
|
4449 |
unfolding id |
|
| 64795 | 4450 |
proof (subst Suc(1), simp add: n, |
| 65347 | 4451 |
subst minus_poly_rev_list, force simp: len, rule cong[OF _ _ refl], goal_cases) |
| 65346 | 4452 |
case 2 |
| 64795 | 4453 |
have "monom lcr (length rr) = monom (lcr div lc) (length rr - length dd) * monom lc (length dd)" |
4454 |
by (simp add: mult_monom len True) |
|
| 65346 | 4455 |
then show ?case unfolding r d Poly_append n ring_distribs |
| 64795 | 4456 |
by (auto simp: Poly_map smult_monom smult_monom_mult) |
4457 |
qed (auto simp: len monom_Suc smult_monom) |
|
4458 |
qed |
|
4459 |
qed simp |
|
4460 |
||
| 65346 | 4461 |
lemma divide_poly_list[code]: "f div g = divide_poly_list f g" |
| 64795 | 4462 |
proof - |
4463 |
note d = divide_poly_def divide_poly_list_def |
|
4464 |
show ?thesis |
|
4465 |
proof (cases "g = 0") |
|
4466 |
case True |
|
| 65347 | 4467 |
show ?thesis by (auto simp: d True) |
| 64795 | 4468 |
next |
4469 |
case False |
|
| 65347 | 4470 |
then obtain cg lcg where cg: "coeffs g = cg @ [lcg]" |
4471 |
by (cases "coeffs g" rule: rev_cases) auto |
|
4472 |
with False have id: "(g = 0) = False" "(cg @ [lcg] = []) = False" |
|
4473 |
by auto |
|
| 65346 | 4474 |
from cg False have lcg: "coeff g (degree g) = lcg" |
| 64795 | 4475 |
using last_coeffs_eq_coeff_degree last_snoc by force |
| 65347 | 4476 |
with False have "lcg \<noteq> 0" by auto |
4477 |
from cg Poly_coeffs [of g] have ltp: "Poly (cg @ [lcg]) = g" |
|
4478 |
by auto |
|
4479 |
show ?thesis |
|
4480 |
unfolding d cg Let_def id if_False poly_of_list_def |
|
4481 |
by (subst divide_poly_main_list, insert False cg \<open>lcg \<noteq> 0\<close>) |
|
4482 |
(auto simp: lcg ltp, simp add: degree_eq_length_coeffs) |
|
| 64795 | 4483 |
qed |
|
63317
ca187a9f66da
Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
63145
diff
changeset
|
4484 |
qed |
| 52380 | 4485 |
|
4486 |
no_notation cCons (infixr "##" 65) |
|
| 31663 | 4487 |
|
| 29478 | 4488 |
end |