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