| author | wenzelm |
| Fri, 18 Mar 2016 17:58:19 +0100 | |
| changeset 62668 | 360d3464919c |
| parent 62422 | 4aa35fd6c152 |
| child 63027 | 8de0ebee3f1c |
| permissions | -rw-r--r-- |
| 41959 | 1 |
(* Title: HOL/Library/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 |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
11 |
imports Main "~~/src/HOL/Deriv" "~~/src/HOL/Library/More_List" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
12 |
"~~/src/HOL/Library/Infinite_Set" |
| 29451 | 13 |
begin |
14 |
||
| 60500 | 15 |
subsection \<open>Auxiliary: operations for lists (later) representing coefficients\<close> |
| 52380 | 16 |
|
17 |
definition cCons :: "'a::zero \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixr "##" 65) |
|
18 |
where |
|
19 |
"x ## xs = (if xs = [] \<and> x = 0 then [] else x # xs)" |
|
20 |
||
21 |
lemma cCons_0_Nil_eq [simp]: |
|
22 |
"0 ## [] = []" |
|
23 |
by (simp add: cCons_def) |
|
24 |
||
25 |
lemma cCons_Cons_eq [simp]: |
|
26 |
"x ## y # ys = x # y # ys" |
|
27 |
by (simp add: cCons_def) |
|
28 |
||
29 |
lemma cCons_append_Cons_eq [simp]: |
|
30 |
"x ## xs @ y # ys = x # xs @ y # ys" |
|
31 |
by (simp add: cCons_def) |
|
32 |
||
33 |
lemma cCons_not_0_eq [simp]: |
|
34 |
"x \<noteq> 0 \<Longrightarrow> x ## xs = x # xs" |
|
35 |
by (simp add: cCons_def) |
|
36 |
||
37 |
lemma strip_while_not_0_Cons_eq [simp]: |
|
38 |
"strip_while (\<lambda>x. x = 0) (x # xs) = x ## strip_while (\<lambda>x. x = 0) xs" |
|
39 |
proof (cases "x = 0") |
|
40 |
case False then show ?thesis by simp |
|
41 |
next |
|
42 |
case True show ?thesis |
|
43 |
proof (induct xs rule: rev_induct) |
|
44 |
case Nil with True show ?case by simp |
|
45 |
next |
|
46 |
case (snoc y ys) then show ?case |
|
47 |
by (cases "y = 0") (simp_all add: append_Cons [symmetric] del: append_Cons) |
|
48 |
qed |
|
49 |
qed |
|
50 |
||
51 |
lemma tl_cCons [simp]: |
|
52 |
"tl (x ## xs) = xs" |
|
53 |
by (simp add: cCons_def) |
|
54 |
||
| 61585 | 55 |
subsection \<open>Definition of type \<open>poly\<close>\<close> |
| 29451 | 56 |
|
| 61260 | 57 |
typedef (overloaded) 'a poly = "{f :: nat \<Rightarrow> 'a::zero. \<forall>\<^sub>\<infinity> n. f n = 0}"
|
|
59983
cd2efd7d06bd
replace almost_everywhere_zero by Infinite_Set.MOST
hoelzl
parents:
59815
diff
changeset
|
58 |
morphisms coeff Abs_poly by (auto intro!: ALL_MOST) |
| 29451 | 59 |
|
|
59487
adaa430fc0f7
default abstypes and default abstract equations make technical (no_code) annotation superfluous
haftmann
parents:
58881
diff
changeset
|
60 |
setup_lifting type_definition_poly |
| 52380 | 61 |
|
62 |
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
|
63 |
by (simp add: coeff_inject [symmetric] fun_eq_iff) |
| 29451 | 64 |
|
| 52380 | 65 |
lemma poly_eqI: "(\<And>n. coeff p n = coeff q n) \<Longrightarrow> p = q" |
66 |
by (simp add: poly_eq_iff) |
|
67 |
||
|
59983
cd2efd7d06bd
replace almost_everywhere_zero by Infinite_Set.MOST
hoelzl
parents:
59815
diff
changeset
|
68 |
lemma MOST_coeff_eq_0: "\<forall>\<^sub>\<infinity> n. coeff p n = 0" |
| 52380 | 69 |
using coeff [of p] by simp |
| 29451 | 70 |
|
71 |
||
| 60500 | 72 |
subsection \<open>Degree of a polynomial\<close> |
| 29451 | 73 |
|
| 52380 | 74 |
definition degree :: "'a::zero poly \<Rightarrow> nat" |
75 |
where |
|
| 29451 | 76 |
"degree p = (LEAST n. \<forall>i>n. coeff p i = 0)" |
77 |
||
| 52380 | 78 |
lemma coeff_eq_0: |
79 |
assumes "degree p < n" |
|
80 |
shows "coeff p n = 0" |
|
| 29451 | 81 |
proof - |
|
59983
cd2efd7d06bd
replace almost_everywhere_zero by Infinite_Set.MOST
hoelzl
parents:
59815
diff
changeset
|
82 |
have "\<exists>n. \<forall>i>n. coeff p i = 0" |
|
cd2efd7d06bd
replace almost_everywhere_zero by Infinite_Set.MOST
hoelzl
parents:
59815
diff
changeset
|
83 |
using MOST_coeff_eq_0 by (simp add: MOST_nat) |
| 52380 | 84 |
then have "\<forall>i>degree p. coeff p i = 0" |
| 29451 | 85 |
unfolding degree_def by (rule LeastI_ex) |
| 52380 | 86 |
with assms show ?thesis by simp |
| 29451 | 87 |
qed |
88 |
||
89 |
lemma le_degree: "coeff p n \<noteq> 0 \<Longrightarrow> n \<le> degree p" |
|
90 |
by (erule contrapos_np, rule coeff_eq_0, simp) |
|
91 |
||
92 |
lemma degree_le: "\<forall>i>n. coeff p i = 0 \<Longrightarrow> degree p \<le> n" |
|
93 |
unfolding degree_def by (erule Least_le) |
|
94 |
||
95 |
lemma less_degree_imp: "n < degree p \<Longrightarrow> \<exists>i>n. coeff p i \<noteq> 0" |
|
96 |
unfolding degree_def by (drule not_less_Least, simp) |
|
97 |
||
98 |
||
| 60500 | 99 |
subsection \<open>The zero polynomial\<close> |
| 29451 | 100 |
|
101 |
instantiation poly :: (zero) zero |
|
102 |
begin |
|
103 |
||
| 52380 | 104 |
lift_definition zero_poly :: "'a poly" |
|
59983
cd2efd7d06bd
replace almost_everywhere_zero by Infinite_Set.MOST
hoelzl
parents:
59815
diff
changeset
|
105 |
is "\<lambda>_. 0" by (rule MOST_I) simp |
| 29451 | 106 |
|
107 |
instance .. |
|
| 52380 | 108 |
|
| 29451 | 109 |
end |
110 |
||
| 52380 | 111 |
lemma coeff_0 [simp]: |
112 |
"coeff 0 n = 0" |
|
113 |
by transfer rule |
|
| 29451 | 114 |
|
| 52380 | 115 |
lemma degree_0 [simp]: |
116 |
"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 |
|
| 60500 | 124 |
from \<open>p \<noteq> 0\<close> have "\<exists>n. coeff p n \<noteq> 0" |
| 52380 | 125 |
by (simp add: poly_eq_iff) |
| 29451 | 126 |
then obtain n where "coeff p n \<noteq> 0" .. |
127 |
hence "n \<le> degree p" by (rule le_degree) |
|
| 60500 | 128 |
with \<open>coeff p n \<noteq> 0\<close> and \<open>degree p = 0\<close> |
| 29451 | 129 |
show "coeff p (degree p) \<noteq> 0" by simp |
130 |
next |
|
131 |
case (Suc n) |
|
| 60500 | 132 |
from \<open>degree p = Suc n\<close> have "n < degree p" by simp |
| 29451 | 133 |
hence "\<exists>i>n. coeff p i \<noteq> 0" by (rule less_degree_imp) |
134 |
then obtain i where "n < i" and "coeff p i \<noteq> 0" by fast |
|
| 60500 | 135 |
from \<open>degree p = Suc n\<close> and \<open>n < i\<close> have "degree p \<le> i" by simp |
136 |
also from \<open>coeff p i \<noteq> 0\<close> have "i \<le> degree p" by (rule le_degree) |
|
| 29451 | 137 |
finally have "degree p = i" . |
| 60500 | 138 |
with \<open>coeff p i \<noteq> 0\<close> show "coeff p (degree p) \<noteq> 0" by simp |
| 29451 | 139 |
qed |
140 |
||
| 52380 | 141 |
lemma leading_coeff_0_iff [simp]: |
142 |
"coeff p (degree p) = 0 \<longleftrightarrow> p = 0" |
|
| 29451 | 143 |
by (cases "p = 0", simp, simp add: leading_coeff_neq_0) |
144 |
||
145 |
||
| 60500 | 146 |
subsection \<open>List-style constructor for polynomials\<close> |
| 29451 | 147 |
|
| 52380 | 148 |
lift_definition pCons :: "'a::zero \<Rightarrow> 'a poly \<Rightarrow> 'a poly" |
| 55415 | 149 |
is "\<lambda>a p. case_nat a (coeff p)" |
|
59983
cd2efd7d06bd
replace almost_everywhere_zero by Infinite_Set.MOST
hoelzl
parents:
59815
diff
changeset
|
150 |
by (rule MOST_SucD) (simp add: MOST_coeff_eq_0) |
| 29451 | 151 |
|
| 52380 | 152 |
lemmas coeff_pCons = pCons.rep_eq |
| 29455 | 153 |
|
| 52380 | 154 |
lemma coeff_pCons_0 [simp]: |
155 |
"coeff (pCons a p) 0 = a" |
|
156 |
by transfer simp |
|
| 29455 | 157 |
|
| 52380 | 158 |
lemma coeff_pCons_Suc [simp]: |
159 |
"coeff (pCons a p) (Suc n) = coeff p n" |
|
| 29451 | 160 |
by (simp add: coeff_pCons) |
161 |
||
| 52380 | 162 |
lemma degree_pCons_le: |
163 |
"degree (pCons a p) \<le> Suc (degree p)" |
|
164 |
by (rule degree_le) (simp add: coeff_eq_0 coeff_pCons split: nat.split) |
|
| 29451 | 165 |
|
166 |
lemma degree_pCons_eq: |
|
167 |
"p \<noteq> 0 \<Longrightarrow> degree (pCons a p) = Suc (degree p)" |
|
| 52380 | 168 |
apply (rule order_antisym [OF degree_pCons_le]) |
169 |
apply (rule le_degree, simp) |
|
170 |
done |
|
| 29451 | 171 |
|
| 52380 | 172 |
lemma degree_pCons_0: |
173 |
"degree (pCons a 0) = 0" |
|
174 |
apply (rule order_antisym [OF _ le0]) |
|
175 |
apply (rule degree_le, simp add: coeff_pCons split: nat.split) |
|
176 |
done |
|
| 29451 | 177 |
|
|
29460
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
178 |
lemma degree_pCons_eq_if [simp]: |
| 29451 | 179 |
"degree (pCons a p) = (if p = 0 then 0 else Suc (degree p))" |
| 52380 | 180 |
apply (cases "p = 0", simp_all) |
181 |
apply (rule order_antisym [OF _ le0]) |
|
182 |
apply (rule degree_le, simp add: coeff_pCons split: nat.split) |
|
183 |
apply (rule order_antisym [OF degree_pCons_le]) |
|
184 |
apply (rule le_degree, simp) |
|
185 |
done |
|
| 29451 | 186 |
|
| 52380 | 187 |
lemma pCons_0_0 [simp]: |
188 |
"pCons 0 0 = 0" |
|
189 |
by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) |
|
| 29451 | 190 |
|
191 |
lemma pCons_eq_iff [simp]: |
|
192 |
"pCons a p = pCons b q \<longleftrightarrow> a = b \<and> p = q" |
|
| 52380 | 193 |
proof safe |
| 29451 | 194 |
assume "pCons a p = pCons b q" |
195 |
then have "coeff (pCons a p) 0 = coeff (pCons b q) 0" by simp |
|
196 |
then show "a = b" by simp |
|
197 |
next |
|
198 |
assume "pCons a p = pCons b q" |
|
199 |
then have "\<forall>n. coeff (pCons a p) (Suc n) = |
|
200 |
coeff (pCons b q) (Suc n)" by simp |
|
| 52380 | 201 |
then show "p = q" by (simp add: poly_eq_iff) |
| 29451 | 202 |
qed |
203 |
||
| 52380 | 204 |
lemma pCons_eq_0_iff [simp]: |
205 |
"pCons a p = 0 \<longleftrightarrow> a = 0 \<and> p = 0" |
|
| 29451 | 206 |
using pCons_eq_iff [of a p 0 0] by simp |
207 |
||
208 |
lemma pCons_cases [cases type: poly]: |
|
209 |
obtains (pCons) a q where "p = pCons a q" |
|
210 |
proof |
|
211 |
show "p = pCons (coeff p 0) (Abs_poly (\<lambda>n. coeff p (Suc n)))" |
|
| 52380 | 212 |
by transfer |
|
59983
cd2efd7d06bd
replace almost_everywhere_zero by Infinite_Set.MOST
hoelzl
parents:
59815
diff
changeset
|
213 |
(simp_all add: MOST_inj[where f=Suc and P="\<lambda>n. p n = 0" for p] fun_eq_iff Abs_poly_inverse |
|
cd2efd7d06bd
replace almost_everywhere_zero by Infinite_Set.MOST
hoelzl
parents:
59815
diff
changeset
|
214 |
split: nat.split) |
| 29451 | 215 |
qed |
216 |
||
217 |
lemma pCons_induct [case_names 0 pCons, induct type: poly]: |
|
218 |
assumes zero: "P 0" |
|
| 54856 | 219 |
assumes pCons: "\<And>a p. a \<noteq> 0 \<or> p \<noteq> 0 \<Longrightarrow> P p \<Longrightarrow> P (pCons a p)" |
| 29451 | 220 |
shows "P p" |
221 |
proof (induct p rule: measure_induct_rule [where f=degree]) |
|
222 |
case (less p) |
|
223 |
obtain a q where "p = pCons a q" by (rule pCons_cases) |
|
224 |
have "P q" |
|
225 |
proof (cases "q = 0") |
|
226 |
case True |
|
227 |
then show "P q" by (simp add: zero) |
|
228 |
next |
|
229 |
case False |
|
230 |
then have "degree (pCons a q) = Suc (degree q)" |
|
231 |
by (rule degree_pCons_eq) |
|
232 |
then have "degree q < degree p" |
|
| 60500 | 233 |
using \<open>p = pCons a q\<close> by simp |
| 29451 | 234 |
then show "P q" |
235 |
by (rule less.hyps) |
|
236 |
qed |
|
| 54856 | 237 |
have "P (pCons a q)" |
238 |
proof (cases "a \<noteq> 0 \<or> q \<noteq> 0") |
|
239 |
case True |
|
| 60500 | 240 |
with \<open>P q\<close> show ?thesis by (auto intro: pCons) |
| 54856 | 241 |
next |
242 |
case False |
|
243 |
with zero show ?thesis by simp |
|
244 |
qed |
|
| 29451 | 245 |
then show ?case |
| 60500 | 246 |
using \<open>p = pCons a q\<close> by simp |
| 29451 | 247 |
qed |
248 |
||
| 60570 | 249 |
lemma degree_eq_zeroE: |
250 |
fixes p :: "'a::zero poly" |
|
251 |
assumes "degree p = 0" |
|
252 |
obtains a where "p = pCons a 0" |
|
253 |
proof - |
|
254 |
obtain a q where p: "p = pCons a q" by (cases p) |
|
255 |
with assms have "q = 0" by (cases "q = 0") simp_all |
|
256 |
with p have "p = pCons a 0" by simp |
|
257 |
with that show thesis . |
|
258 |
qed |
|
259 |
||
| 29451 | 260 |
|
| 62422 | 261 |
subsection \<open>Quickcheck generator for polynomials\<close> |
262 |
||
263 |
quickcheck_generator poly constructors: "0 :: _ poly", pCons |
|
264 |
||
265 |
||
| 60500 | 266 |
subsection \<open>List-style syntax for polynomials\<close> |
| 52380 | 267 |
|
268 |
syntax |
|
269 |
"_poly" :: "args \<Rightarrow> 'a poly" ("[:(_):]")
|
|
270 |
||
271 |
translations |
|
272 |
"[:x, xs:]" == "CONST pCons x [:xs:]" |
|
273 |
"[:x:]" == "CONST pCons x 0" |
|
274 |
"[:x:]" <= "CONST pCons x (_constrain 0 t)" |
|
275 |
||
276 |
||
| 60500 | 277 |
subsection \<open>Representation of polynomials by lists of coefficients\<close> |
| 52380 | 278 |
|
279 |
primrec Poly :: "'a::zero list \<Rightarrow> 'a poly" |
|
280 |
where |
|
| 54855 | 281 |
[code_post]: "Poly [] = 0" |
282 |
| [code_post]: "Poly (a # as) = pCons a (Poly as)" |
|
| 52380 | 283 |
|
284 |
lemma Poly_replicate_0 [simp]: |
|
285 |
"Poly (replicate n 0) = 0" |
|
286 |
by (induct n) simp_all |
|
287 |
||
288 |
lemma Poly_eq_0: |
|
289 |
"Poly as = 0 \<longleftrightarrow> (\<exists>n. as = replicate n 0)" |
|
290 |
by (induct as) (auto simp add: Cons_replicate_eq) |
|
| 62065 | 291 |
|
292 |
lemma degree_Poly: "degree (Poly xs) \<le> length xs" |
|
293 |
by (induction xs) simp_all |
|
294 |
||
| 52380 | 295 |
definition coeffs :: "'a poly \<Rightarrow> 'a::zero list" |
296 |
where |
|
297 |
"coeffs p = (if p = 0 then [] else map (\<lambda>i. coeff p i) [0 ..< Suc (degree p)])" |
|
298 |
||
299 |
lemma coeffs_eq_Nil [simp]: |
|
300 |
"coeffs p = [] \<longleftrightarrow> p = 0" |
|
301 |
by (simp add: coeffs_def) |
|
302 |
||
303 |
lemma not_0_coeffs_not_Nil: |
|
304 |
"p \<noteq> 0 \<Longrightarrow> coeffs p \<noteq> []" |
|
305 |
by simp |
|
306 |
||
307 |
lemma coeffs_0_eq_Nil [simp]: |
|
308 |
"coeffs 0 = []" |
|
309 |
by simp |
|
|
29454
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
310 |
|
| 52380 | 311 |
lemma coeffs_pCons_eq_cCons [simp]: |
312 |
"coeffs (pCons a p) = a ## coeffs p" |
|
313 |
proof - |
|
314 |
{ fix ms :: "nat list" and f :: "nat \<Rightarrow> 'a" and x :: "'a"
|
|
315 |
assume "\<forall>m\<in>set ms. m > 0" |
|
| 55415 | 316 |
then have "map (case_nat x f) ms = map f (map (\<lambda>n. n - 1) ms)" |
|
58199
5fbe474b5da8
explicit theory with additional, less commonly used list operations
haftmann
parents:
57862
diff
changeset
|
317 |
by (induct ms) (auto split: nat.split) |
|
5fbe474b5da8
explicit theory with additional, less commonly used list operations
haftmann
parents:
57862
diff
changeset
|
318 |
} |
| 52380 | 319 |
note * = this |
320 |
show ?thesis |
|
| 60570 | 321 |
by (simp add: coeffs_def * upt_conv_Cons coeff_pCons map_decr_upt del: upt_Suc) |
| 52380 | 322 |
qed |
323 |
||
| 62065 | 324 |
lemma length_coeffs: "p \<noteq> 0 \<Longrightarrow> length (coeffs p) = degree p + 1" |
325 |
by (simp add: coeffs_def) |
|
326 |
||
327 |
lemma coeffs_nth: |
|
328 |
assumes "p \<noteq> 0" "n \<le> degree p" |
|
329 |
shows "coeffs p ! n = coeff p n" |
|
330 |
using assms unfolding coeffs_def by (auto simp del: upt_Suc) |
|
331 |
||
| 52380 | 332 |
lemma not_0_cCons_eq [simp]: |
333 |
"p \<noteq> 0 \<Longrightarrow> a ## coeffs p = a # coeffs p" |
|
334 |
by (simp add: cCons_def) |
|
335 |
||
336 |
lemma Poly_coeffs [simp, code abstype]: |
|
337 |
"Poly (coeffs p) = p" |
|
| 54856 | 338 |
by (induct p) auto |
| 52380 | 339 |
|
340 |
lemma coeffs_Poly [simp]: |
|
341 |
"coeffs (Poly as) = strip_while (HOL.eq 0) as" |
|
342 |
proof (induct as) |
|
343 |
case Nil then show ?case by simp |
|
344 |
next |
|
345 |
case (Cons a as) |
|
346 |
have "(\<forall>n. as \<noteq> replicate n 0) \<longleftrightarrow> (\<exists>a\<in>set as. a \<noteq> 0)" |
|
347 |
using replicate_length_same [of as 0] by (auto dest: sym [of _ as]) |
|
348 |
with Cons show ?case by auto |
|
349 |
qed |
|
350 |
||
351 |
lemma last_coeffs_not_0: |
|
352 |
"p \<noteq> 0 \<Longrightarrow> last (coeffs p) \<noteq> 0" |
|
353 |
by (induct p) (auto simp add: cCons_def) |
|
354 |
||
355 |
lemma strip_while_coeffs [simp]: |
|
356 |
"strip_while (HOL.eq 0) (coeffs p) = coeffs p" |
|
357 |
by (cases "p = 0") (auto dest: last_coeffs_not_0 intro: strip_while_not_last) |
|
358 |
||
359 |
lemma coeffs_eq_iff: |
|
360 |
"p = q \<longleftrightarrow> coeffs p = coeffs q" (is "?P \<longleftrightarrow> ?Q") |
|
361 |
proof |
|
362 |
assume ?P then show ?Q by simp |
|
363 |
next |
|
364 |
assume ?Q |
|
365 |
then have "Poly (coeffs p) = Poly (coeffs q)" by simp |
|
366 |
then show ?P by simp |
|
367 |
qed |
|
368 |
||
369 |
lemma coeff_Poly_eq: |
|
370 |
"coeff (Poly xs) n = nth_default 0 xs n" |
|
371 |
apply (induct xs arbitrary: n) apply simp_all |
|
|
55642
63beb38e9258
adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
blanchet
parents:
55417
diff
changeset
|
372 |
by (metis nat.case not0_implies_Suc nth_default_Cons_0 nth_default_Cons_Suc pCons.rep_eq) |
|
29454
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
373 |
|
| 52380 | 374 |
lemma nth_default_coeffs_eq: |
375 |
"nth_default 0 (coeffs p) = coeff p" |
|
376 |
by (simp add: fun_eq_iff coeff_Poly_eq [symmetric]) |
|
377 |
||
378 |
lemma [code]: |
|
379 |
"coeff p = nth_default 0 (coeffs p)" |
|
380 |
by (simp add: nth_default_coeffs_eq) |
|
381 |
||
382 |
lemma coeffs_eqI: |
|
383 |
assumes coeff: "\<And>n. coeff p n = nth_default 0 xs n" |
|
384 |
assumes zero: "xs \<noteq> [] \<Longrightarrow> last xs \<noteq> 0" |
|
385 |
shows "coeffs p = xs" |
|
386 |
proof - |
|
387 |
from coeff have "p = Poly xs" by (simp add: poly_eq_iff coeff_Poly_eq) |
|
388 |
with zero show ?thesis by simp (cases xs, simp_all) |
|
389 |
qed |
|
390 |
||
391 |
lemma degree_eq_length_coeffs [code]: |
|
392 |
"degree p = length (coeffs p) - 1" |
|
393 |
by (simp add: coeffs_def) |
|
394 |
||
395 |
lemma length_coeffs_degree: |
|
396 |
"p \<noteq> 0 \<Longrightarrow> length (coeffs p) = Suc (degree p)" |
|
397 |
by (induct p) (auto simp add: cCons_def) |
|
398 |
||
399 |
lemma [code abstract]: |
|
400 |
"coeffs 0 = []" |
|
401 |
by (fact coeffs_0_eq_Nil) |
|
402 |
||
403 |
lemma [code abstract]: |
|
404 |
"coeffs (pCons a p) = a ## coeffs p" |
|
405 |
by (fact coeffs_pCons_eq_cCons) |
|
406 |
||
407 |
instantiation poly :: ("{zero, equal}") equal
|
|
408 |
begin |
|
409 |
||
410 |
definition |
|
411 |
[code]: "HOL.equal (p::'a poly) q \<longleftrightarrow> HOL.equal (coeffs p) (coeffs q)" |
|
412 |
||
| 60679 | 413 |
instance |
414 |
by standard (simp add: equal equal_poly_def coeffs_eq_iff) |
|
| 52380 | 415 |
|
416 |
end |
|
417 |
||
| 60679 | 418 |
lemma [code nbe]: "HOL.equal (p :: _ poly) p \<longleftrightarrow> True" |
| 52380 | 419 |
by (fact equal_refl) |
|
29454
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
420 |
|
| 52380 | 421 |
definition is_zero :: "'a::zero poly \<Rightarrow> bool" |
422 |
where |
|
423 |
[code]: "is_zero p \<longleftrightarrow> List.null (coeffs p)" |
|
424 |
||
425 |
lemma is_zero_null [code_abbrev]: |
|
426 |
"is_zero p \<longleftrightarrow> p = 0" |
|
427 |
by (simp add: is_zero_def null_def) |
|
428 |
||
429 |
||
| 60500 | 430 |
subsection \<open>Fold combinator for polynomials\<close> |
| 52380 | 431 |
|
432 |
definition fold_coeffs :: "('a::zero \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a poly \<Rightarrow> 'b \<Rightarrow> 'b"
|
|
433 |
where |
|
434 |
"fold_coeffs f p = foldr f (coeffs p)" |
|
435 |
||
436 |
lemma fold_coeffs_0_eq [simp]: |
|
437 |
"fold_coeffs f 0 = id" |
|
438 |
by (simp add: fold_coeffs_def) |
|
439 |
||
440 |
lemma fold_coeffs_pCons_eq [simp]: |
|
441 |
"f 0 = id \<Longrightarrow> fold_coeffs f (pCons a p) = f a \<circ> fold_coeffs f p" |
|
442 |
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
|
443 |
|
| 52380 | 444 |
lemma fold_coeffs_pCons_0_0_eq [simp]: |
445 |
"fold_coeffs f (pCons 0 0) = id" |
|
446 |
by (simp add: fold_coeffs_def) |
|
447 |
||
448 |
lemma fold_coeffs_pCons_coeff_not_0_eq [simp]: |
|
449 |
"a \<noteq> 0 \<Longrightarrow> fold_coeffs f (pCons a p) = f a \<circ> fold_coeffs f p" |
|
450 |
by (simp add: fold_coeffs_def) |
|
451 |
||
452 |
lemma fold_coeffs_pCons_not_0_0_eq [simp]: |
|
453 |
"p \<noteq> 0 \<Longrightarrow> fold_coeffs f (pCons a p) = f a \<circ> fold_coeffs f p" |
|
454 |
by (simp add: fold_coeffs_def) |
|
455 |
||
456 |
||
| 60500 | 457 |
subsection \<open>Canonical morphism on polynomials -- evaluation\<close> |
| 52380 | 458 |
|
459 |
definition poly :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a" |
|
460 |
where |
|
| 61585 | 461 |
"poly p = fold_coeffs (\<lambda>a f x. a + x * f x) p (\<lambda>x. 0)" \<comment> \<open>The Horner Schema\<close> |
| 52380 | 462 |
|
463 |
lemma poly_0 [simp]: |
|
464 |
"poly 0 x = 0" |
|
465 |
by (simp add: poly_def) |
|
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
466 |
|
| 52380 | 467 |
lemma poly_pCons [simp]: |
468 |
"poly (pCons a p) x = a + x * poly p x" |
|
469 |
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
|
470 |
|
| 62065 | 471 |
lemma poly_altdef: |
472 |
"poly p (x :: 'a :: {comm_semiring_0, semiring_1}) = (\<Sum>i\<le>degree p. coeff p i * x ^ i)"
|
|
473 |
proof (induction p rule: pCons_induct) |
|
474 |
case (pCons a p) |
|
475 |
show ?case |
|
476 |
proof (cases "p = 0") |
|
477 |
case False |
|
478 |
let ?p' = "pCons a p" |
|
479 |
note poly_pCons[of a p x] |
|
480 |
also note pCons.IH |
|
481 |
also have "a + x * (\<Sum>i\<le>degree p. coeff p i * x ^ i) = |
|
482 |
coeff ?p' 0 * x^0 + (\<Sum>i\<le>degree p. coeff ?p' (Suc i) * x^Suc i)" |
|
483 |
by (simp add: field_simps setsum_right_distrib coeff_pCons) |
|
484 |
also note setsum_atMost_Suc_shift[symmetric] |
|
| 62072 | 485 |
also note degree_pCons_eq[OF \<open>p \<noteq> 0\<close>, of a, symmetric] |
| 62065 | 486 |
finally show ?thesis . |
487 |
qed simp |
|
488 |
qed simp |
|
489 |
||
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
490 |
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
|
491 |
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
|
492 |
|
|
29454
b0f586f38dd7
add recursion combinator poly_rec; define poly function using poly_rec
huffman
parents:
29453
diff
changeset
|
493 |
|
| 60500 | 494 |
subsection \<open>Monomials\<close> |
| 29451 | 495 |
|
| 52380 | 496 |
lift_definition monom :: "'a \<Rightarrow> nat \<Rightarrow> 'a::zero poly" |
497 |
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
|
498 |
by (simp add: MOST_iff_cofinite) |
| 52380 | 499 |
|
500 |
lemma coeff_monom [simp]: |
|
501 |
"coeff (monom a m) n = (if m = n then a else 0)" |
|
502 |
by transfer rule |
|
| 29451 | 503 |
|
| 52380 | 504 |
lemma monom_0: |
505 |
"monom a 0 = pCons a 0" |
|
506 |
by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) |
|
| 29451 | 507 |
|
| 52380 | 508 |
lemma monom_Suc: |
509 |
"monom a (Suc n) = pCons 0 (monom a n)" |
|
510 |
by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) |
|
| 29451 | 511 |
|
512 |
lemma monom_eq_0 [simp]: "monom 0 n = 0" |
|
| 52380 | 513 |
by (rule poly_eqI) simp |
| 29451 | 514 |
|
515 |
lemma monom_eq_0_iff [simp]: "monom a n = 0 \<longleftrightarrow> a = 0" |
|
| 52380 | 516 |
by (simp add: poly_eq_iff) |
| 29451 | 517 |
|
518 |
lemma monom_eq_iff [simp]: "monom a n = monom b n \<longleftrightarrow> a = b" |
|
| 52380 | 519 |
by (simp add: poly_eq_iff) |
| 29451 | 520 |
|
521 |
lemma degree_monom_le: "degree (monom a n) \<le> n" |
|
522 |
by (rule degree_le, simp) |
|
523 |
||
524 |
lemma degree_monom_eq: "a \<noteq> 0 \<Longrightarrow> degree (monom a n) = n" |
|
525 |
apply (rule order_antisym [OF degree_monom_le]) |
|
526 |
apply (rule le_degree, simp) |
|
527 |
done |
|
528 |
||
| 52380 | 529 |
lemma coeffs_monom [code abstract]: |
530 |
"coeffs (monom a n) = (if a = 0 then [] else replicate n 0 @ [a])" |
|
531 |
by (induct n) (simp_all add: monom_0 monom_Suc) |
|
532 |
||
533 |
lemma fold_coeffs_monom [simp]: |
|
534 |
"a \<noteq> 0 \<Longrightarrow> fold_coeffs f (monom a n) = f 0 ^^ n \<circ> f a" |
|
535 |
by (simp add: fold_coeffs_def coeffs_monom fun_eq_iff) |
|
536 |
||
537 |
lemma poly_monom: |
|
538 |
fixes a x :: "'a::{comm_semiring_1}"
|
|
539 |
shows "poly (monom a n) x = a * x ^ n" |
|
540 |
by (cases "a = 0", simp_all) |
|
541 |
(induct n, simp_all add: mult.left_commute poly_def) |
|
542 |
||
| 62065 | 543 |
|
| 60500 | 544 |
subsection \<open>Addition and subtraction\<close> |
| 29451 | 545 |
|
546 |
instantiation poly :: (comm_monoid_add) comm_monoid_add |
|
547 |
begin |
|
548 |
||
| 52380 | 549 |
lift_definition plus_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" |
550 |
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
|
551 |
proof - |
| 60679 | 552 |
fix q p :: "'a poly" |
553 |
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
|
554 |
using MOST_coeff_eq_0[of p] MOST_coeff_eq_0[of q] by eventually_elim simp |
| 52380 | 555 |
qed |
| 29451 | 556 |
|
| 60679 | 557 |
lemma coeff_add [simp]: "coeff (p + q) n = coeff p n + coeff q n" |
| 52380 | 558 |
by (simp add: plus_poly.rep_eq) |
| 29451 | 559 |
|
| 60679 | 560 |
instance |
561 |
proof |
|
| 29451 | 562 |
fix p q r :: "'a poly" |
563 |
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
|
564 |
by (simp add: poly_eq_iff add.assoc) |
| 29451 | 565 |
show "p + q = q + p" |
|
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
57482
diff
changeset
|
566 |
by (simp add: poly_eq_iff add.commute) |
| 29451 | 567 |
show "0 + p = p" |
| 52380 | 568 |
by (simp add: poly_eq_iff) |
| 29451 | 569 |
qed |
570 |
||
571 |
end |
|
572 |
||
|
59815
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
573 |
instantiation poly :: (cancel_comm_monoid_add) cancel_comm_monoid_add |
|
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
574 |
begin |
|
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
575 |
|
|
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
576 |
lift_definition minus_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" |
|
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
577 |
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
|
578 |
proof - |
| 60679 | 579 |
fix q p :: "'a poly" |
580 |
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
|
581 |
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
|
582 |
qed |
|
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
583 |
|
| 60679 | 584 |
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
|
585 |
by (simp add: minus_poly.rep_eq) |
|
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
586 |
|
| 60679 | 587 |
instance |
588 |
proof |
|
| 29540 | 589 |
fix p q r :: "'a poly" |
|
59815
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
590 |
show "p + q - p = q" |
| 52380 | 591 |
by (simp add: poly_eq_iff) |
|
59815
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
592 |
show "p - q - r = p - (q + r)" |
|
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
593 |
by (simp add: poly_eq_iff diff_diff_eq) |
| 29540 | 594 |
qed |
595 |
||
|
59815
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
596 |
end |
|
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
597 |
|
| 29451 | 598 |
instantiation poly :: (ab_group_add) ab_group_add |
599 |
begin |
|
600 |
||
| 52380 | 601 |
lift_definition uminus_poly :: "'a poly \<Rightarrow> 'a poly" |
602 |
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
|
603 |
proof - |
| 60679 | 604 |
fix p :: "'a poly" |
605 |
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
|
606 |
using MOST_coeff_eq_0 by simp |
| 52380 | 607 |
qed |
| 29451 | 608 |
|
609 |
lemma coeff_minus [simp]: "coeff (- p) n = - coeff p n" |
|
| 52380 | 610 |
by (simp add: uminus_poly.rep_eq) |
| 29451 | 611 |
|
| 60679 | 612 |
instance |
613 |
proof |
|
| 29451 | 614 |
fix p q :: "'a poly" |
615 |
show "- p + p = 0" |
|
| 52380 | 616 |
by (simp add: poly_eq_iff) |
| 29451 | 617 |
show "p - q = p + - q" |
|
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
52380
diff
changeset
|
618 |
by (simp add: poly_eq_iff) |
| 29451 | 619 |
qed |
620 |
||
621 |
end |
|
622 |
||
623 |
lemma add_pCons [simp]: |
|
624 |
"pCons a p + pCons b q = pCons (a + b) (p + q)" |
|
| 52380 | 625 |
by (rule poly_eqI, simp add: coeff_pCons split: nat.split) |
| 29451 | 626 |
|
627 |
lemma minus_pCons [simp]: |
|
628 |
"- pCons a p = pCons (- a) (- p)" |
|
| 52380 | 629 |
by (rule poly_eqI, simp add: coeff_pCons split: nat.split) |
| 29451 | 630 |
|
631 |
lemma diff_pCons [simp]: |
|
632 |
"pCons a p - pCons b q = pCons (a - b) (p - q)" |
|
| 52380 | 633 |
by (rule poly_eqI, simp add: coeff_pCons split: nat.split) |
| 29451 | 634 |
|
| 29539 | 635 |
lemma degree_add_le_max: "degree (p + q) \<le> max (degree p) (degree q)" |
| 29451 | 636 |
by (rule degree_le, auto simp add: coeff_eq_0) |
637 |
||
| 29539 | 638 |
lemma degree_add_le: |
639 |
"\<lbrakk>degree p \<le> n; degree q \<le> n\<rbrakk> \<Longrightarrow> degree (p + q) \<le> n" |
|
640 |
by (auto intro: order_trans degree_add_le_max) |
|
641 |
||
| 29453 | 642 |
lemma degree_add_less: |
643 |
"\<lbrakk>degree p < n; degree q < n\<rbrakk> \<Longrightarrow> degree (p + q) < n" |
|
| 29539 | 644 |
by (auto intro: le_less_trans degree_add_le_max) |
| 29453 | 645 |
|
| 29451 | 646 |
lemma degree_add_eq_right: |
647 |
"degree p < degree q \<Longrightarrow> degree (p + q) = degree q" |
|
648 |
apply (cases "q = 0", simp) |
|
649 |
apply (rule order_antisym) |
|
| 29539 | 650 |
apply (simp add: degree_add_le) |
| 29451 | 651 |
apply (rule le_degree) |
652 |
apply (simp add: coeff_eq_0) |
|
653 |
done |
|
654 |
||
655 |
lemma degree_add_eq_left: |
|
656 |
"degree q < degree p \<Longrightarrow> degree (p + q) = degree p" |
|
657 |
using degree_add_eq_right [of q p] |
|
|
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
57482
diff
changeset
|
658 |
by (simp add: add.commute) |
| 29451 | 659 |
|
|
59815
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
660 |
lemma degree_minus [simp]: |
|
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
661 |
"degree (- p) = degree p" |
| 29451 | 662 |
unfolding degree_def by simp |
663 |
||
|
59815
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
664 |
lemma degree_diff_le_max: |
|
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
665 |
fixes p q :: "'a :: ab_group_add poly" |
|
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
666 |
shows "degree (p - q) \<le> max (degree p) (degree q)" |
| 29451 | 667 |
using degree_add_le [where p=p and q="-q"] |
|
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
52380
diff
changeset
|
668 |
by simp |
| 29451 | 669 |
|
| 29539 | 670 |
lemma degree_diff_le: |
|
59815
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
671 |
fixes p q :: "'a :: ab_group_add poly" |
|
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
672 |
assumes "degree p \<le> n" and "degree q \<le> n" |
|
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
673 |
shows "degree (p - q) \<le> n" |
|
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
674 |
using assms degree_add_le [of p n "- q"] by simp |
| 29539 | 675 |
|
| 29453 | 676 |
lemma degree_diff_less: |
|
59815
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
677 |
fixes p q :: "'a :: ab_group_add poly" |
|
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
678 |
assumes "degree p < n" and "degree q < n" |
|
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
679 |
shows "degree (p - q) < n" |
|
cce82e360c2f
explicit commutative additive inverse operation;
haftmann
parents:
59557
diff
changeset
|
680 |
using assms degree_add_less [of p n "- q"] by simp |
| 29453 | 681 |
|
| 29451 | 682 |
lemma add_monom: "monom a n + monom b n = monom (a + b) n" |
| 52380 | 683 |
by (rule poly_eqI) simp |
| 29451 | 684 |
|
685 |
lemma diff_monom: "monom a n - monom b n = monom (a - b) n" |
|
| 52380 | 686 |
by (rule poly_eqI) simp |
| 29451 | 687 |
|
688 |
lemma minus_monom: "- monom a n = monom (-a) n" |
|
| 52380 | 689 |
by (rule poly_eqI) simp |
| 29451 | 690 |
|
691 |
lemma coeff_setsum: "coeff (\<Sum>x\<in>A. p x) i = (\<Sum>x\<in>A. coeff (p x) i)" |
|
692 |
by (cases "finite A", induct set: finite, simp_all) |
|
693 |
||
694 |
lemma monom_setsum: "monom (\<Sum>x\<in>A. a x) n = (\<Sum>x\<in>A. monom (a x) n)" |
|
| 52380 | 695 |
by (rule poly_eqI) (simp add: coeff_setsum) |
696 |
||
697 |
fun plus_coeffs :: "'a::comm_monoid_add list \<Rightarrow> 'a list \<Rightarrow> 'a list" |
|
698 |
where |
|
699 |
"plus_coeffs xs [] = xs" |
|
700 |
| "plus_coeffs [] ys = ys" |
|
701 |
| "plus_coeffs (x # xs) (y # ys) = (x + y) ## plus_coeffs xs ys" |
|
702 |
||
703 |
lemma coeffs_plus_eq_plus_coeffs [code abstract]: |
|
704 |
"coeffs (p + q) = plus_coeffs (coeffs p) (coeffs q)" |
|
705 |
proof - |
|
706 |
{ fix xs ys :: "'a list" and n
|
|
707 |
have "nth_default 0 (plus_coeffs xs ys) n = nth_default 0 xs n + nth_default 0 ys n" |
|
708 |
proof (induct xs ys arbitrary: n rule: plus_coeffs.induct) |
|
| 60679 | 709 |
case (3 x xs y ys n) |
710 |
then show ?case by (cases n) (auto simp add: cCons_def) |
|
| 52380 | 711 |
qed simp_all } |
712 |
note * = this |
|
713 |
{ fix xs ys :: "'a list"
|
|
714 |
assume "xs \<noteq> [] \<Longrightarrow> last xs \<noteq> 0" and "ys \<noteq> [] \<Longrightarrow> last ys \<noteq> 0" |
|
715 |
moreover assume "plus_coeffs xs ys \<noteq> []" |
|
716 |
ultimately have "last (plus_coeffs xs ys) \<noteq> 0" |
|
717 |
proof (induct xs ys rule: plus_coeffs.induct) |
|
718 |
case (3 x xs y ys) then show ?case by (auto simp add: cCons_def) metis |
|
719 |
qed simp_all } |
|
720 |
note ** = this |
|
721 |
show ?thesis |
|
722 |
apply (rule coeffs_eqI) |
|
723 |
apply (simp add: * nth_default_coeffs_eq) |
|
724 |
apply (rule **) |
|
725 |
apply (auto dest: last_coeffs_not_0) |
|
726 |
done |
|
727 |
qed |
|
728 |
||
729 |
lemma coeffs_uminus [code abstract]: |
|
730 |
"coeffs (- p) = map (\<lambda>a. - a) (coeffs p)" |
|
731 |
by (rule coeffs_eqI) |
|
732 |
(simp_all add: not_0_coeffs_not_Nil last_map last_coeffs_not_0 nth_default_map_eq nth_default_coeffs_eq) |
|
733 |
||
734 |
lemma [code]: |
|
735 |
fixes p q :: "'a::ab_group_add poly" |
|
736 |
shows "p - q = p + - q" |
|
| 59557 | 737 |
by (fact diff_conv_add_uminus) |
| 52380 | 738 |
|
739 |
lemma poly_add [simp]: "poly (p + q) x = poly p x + poly q x" |
|
740 |
apply (induct p arbitrary: q, simp) |
|
741 |
apply (case_tac q, simp, simp add: algebra_simps) |
|
742 |
done |
|
743 |
||
744 |
lemma poly_minus [simp]: |
|
745 |
fixes x :: "'a::comm_ring" |
|
746 |
shows "poly (- p) x = - poly p x" |
|
747 |
by (induct p) simp_all |
|
748 |
||
749 |
lemma poly_diff [simp]: |
|
750 |
fixes x :: "'a::comm_ring" |
|
751 |
shows "poly (p - q) x = poly p x - poly q x" |
|
|
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
52380
diff
changeset
|
752 |
using poly_add [of p "- q" x] by simp |
| 52380 | 753 |
|
754 |
lemma poly_setsum: "poly (\<Sum>k\<in>A. p k) x = (\<Sum>k\<in>A. poly (p k) x)" |
|
755 |
by (induct A rule: infinite_finite_induct) simp_all |
|
| 29451 | 756 |
|
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
757 |
lemma degree_setsum_le: "finite S \<Longrightarrow> (\<And> p . p \<in> S \<Longrightarrow> degree (f p) \<le> n) |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
758 |
\<Longrightarrow> degree (setsum f S) \<le> n" |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
759 |
proof (induct S rule: finite_induct) |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
760 |
case (insert p S) |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
761 |
hence "degree (setsum f S) \<le> n" "degree (f p) \<le> n" by auto |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
762 |
thus ?case unfolding setsum.insert[OF insert(1-2)] by (metis degree_add_le) |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
763 |
qed simp |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
764 |
|
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
765 |
lemma poly_as_sum_of_monoms': |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
766 |
assumes n: "degree p \<le> n" |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
767 |
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
|
768 |
proof - |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
769 |
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
|
770 |
by auto |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
771 |
show ?thesis |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
772 |
using n by (simp add: poly_eq_iff coeff_setsum coeff_eq_0 setsum.If_cases eq |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
773 |
if_distrib[where f="\<lambda>x. x * a" for a]) |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
774 |
qed |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
775 |
|
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
776 |
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
|
777 |
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
|
778 |
|
| 62065 | 779 |
lemma Poly_snoc: "Poly (xs @ [x]) = Poly xs + monom x (length xs)" |
780 |
by (induction xs) (simp_all add: monom_0 monom_Suc) |
|
781 |
||
| 29451 | 782 |
|
| 60500 | 783 |
subsection \<open>Multiplication by a constant, polynomial multiplication and the unit polynomial\<close> |
| 29451 | 784 |
|
| 52380 | 785 |
lift_definition smult :: "'a::comm_semiring_0 \<Rightarrow> 'a poly \<Rightarrow> 'a poly" |
786 |
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
|
787 |
proof - |
|
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59983
diff
changeset
|
788 |
fix a :: 'a and p :: "'a poly" show "\<forall>\<^sub>\<infinity> i. a * coeff p i = 0" |
|
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
59983
diff
changeset
|
789 |
using MOST_coeff_eq_0[of p] by eventually_elim simp |
| 52380 | 790 |
qed |
| 29451 | 791 |
|
| 52380 | 792 |
lemma coeff_smult [simp]: |
793 |
"coeff (smult a p) n = a * coeff p n" |
|
794 |
by (simp add: smult.rep_eq) |
|
| 29451 | 795 |
|
796 |
lemma degree_smult_le: "degree (smult a p) \<le> degree p" |
|
797 |
by (rule degree_le, simp add: coeff_eq_0) |
|
798 |
||
| 29472 | 799 |
lemma smult_smult [simp]: "smult a (smult b p) = smult (a * b) p" |
|
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
57482
diff
changeset
|
800 |
by (rule poly_eqI, simp add: mult.assoc) |
| 29451 | 801 |
|
802 |
lemma smult_0_right [simp]: "smult a 0 = 0" |
|
| 52380 | 803 |
by (rule poly_eqI, simp) |
| 29451 | 804 |
|
805 |
lemma smult_0_left [simp]: "smult 0 p = 0" |
|
| 52380 | 806 |
by (rule poly_eqI, simp) |
| 29451 | 807 |
|
808 |
lemma smult_1_left [simp]: "smult (1::'a::comm_semiring_1) p = p" |
|
| 52380 | 809 |
by (rule poly_eqI, simp) |
| 29451 | 810 |
|
811 |
lemma smult_add_right: |
|
812 |
"smult a (p + q) = smult a p + smult a q" |
|
| 52380 | 813 |
by (rule poly_eqI, simp add: algebra_simps) |
| 29451 | 814 |
|
815 |
lemma smult_add_left: |
|
816 |
"smult (a + b) p = smult a p + smult b p" |
|
| 52380 | 817 |
by (rule poly_eqI, simp add: algebra_simps) |
| 29451 | 818 |
|
| 29457 | 819 |
lemma smult_minus_right [simp]: |
| 29451 | 820 |
"smult (a::'a::comm_ring) (- p) = - smult a p" |
| 52380 | 821 |
by (rule poly_eqI, simp) |
| 29451 | 822 |
|
| 29457 | 823 |
lemma smult_minus_left [simp]: |
| 29451 | 824 |
"smult (- a::'a::comm_ring) p = - smult a p" |
| 52380 | 825 |
by (rule poly_eqI, simp) |
| 29451 | 826 |
|
827 |
lemma smult_diff_right: |
|
828 |
"smult (a::'a::comm_ring) (p - q) = smult a p - smult a q" |
|
| 52380 | 829 |
by (rule poly_eqI, simp add: algebra_simps) |
| 29451 | 830 |
|
831 |
lemma smult_diff_left: |
|
832 |
"smult (a - b::'a::comm_ring) p = smult a p - smult b p" |
|
| 52380 | 833 |
by (rule poly_eqI, simp add: algebra_simps) |
| 29451 | 834 |
|
| 29472 | 835 |
lemmas smult_distribs = |
836 |
smult_add_left smult_add_right |
|
837 |
smult_diff_left smult_diff_right |
|
838 |
||
| 29451 | 839 |
lemma smult_pCons [simp]: |
840 |
"smult a (pCons b p) = pCons (a * b) (smult a p)" |
|
| 52380 | 841 |
by (rule poly_eqI, simp add: coeff_pCons split: nat.split) |
| 29451 | 842 |
|
843 |
lemma smult_monom: "smult a (monom b n) = monom (a * b) n" |
|
844 |
by (induct n, simp add: monom_0, simp add: monom_Suc) |
|
845 |
||
| 29659 | 846 |
lemma degree_smult_eq [simp]: |
847 |
fixes a :: "'a::idom" |
|
848 |
shows "degree (smult a p) = (if a = 0 then 0 else degree p)" |
|
849 |
by (cases "a = 0", simp, simp add: degree_def) |
|
850 |
||
851 |
lemma smult_eq_0_iff [simp]: |
|
852 |
fixes a :: "'a::idom" |
|
853 |
shows "smult a p = 0 \<longleftrightarrow> a = 0 \<or> p = 0" |
|
| 52380 | 854 |
by (simp add: poly_eq_iff) |
| 29451 | 855 |
|
| 52380 | 856 |
lemma coeffs_smult [code abstract]: |
857 |
fixes p :: "'a::idom poly" |
|
858 |
shows "coeffs (smult a p) = (if a = 0 then [] else map (Groups.times a) (coeffs p))" |
|
859 |
by (rule coeffs_eqI) |
|
860 |
(auto simp add: not_0_coeffs_not_Nil last_map last_coeffs_not_0 nth_default_map_eq nth_default_coeffs_eq) |
|
| 29451 | 861 |
|
862 |
instantiation poly :: (comm_semiring_0) comm_semiring_0 |
|
863 |
begin |
|
864 |
||
865 |
definition |
|
| 52380 | 866 |
"p * q = fold_coeffs (\<lambda>a p. smult a q + pCons 0 p) p 0" |
| 29474 | 867 |
|
868 |
lemma mult_poly_0_left: "(0::'a poly) * q = 0" |
|
| 52380 | 869 |
by (simp add: times_poly_def) |
| 29474 | 870 |
|
871 |
lemma mult_pCons_left [simp]: |
|
872 |
"pCons a p * q = smult a q + pCons 0 (p * q)" |
|
| 52380 | 873 |
by (cases "p = 0 \<and> a = 0") (auto simp add: times_poly_def) |
| 29474 | 874 |
|
875 |
lemma mult_poly_0_right: "p * (0::'a poly) = 0" |
|
| 52380 | 876 |
by (induct p) (simp add: mult_poly_0_left, simp) |
| 29451 | 877 |
|
| 29474 | 878 |
lemma mult_pCons_right [simp]: |
879 |
"p * pCons a q = smult a p + pCons 0 (p * q)" |
|
| 52380 | 880 |
by (induct p) (simp add: mult_poly_0_left, simp add: algebra_simps) |
| 29474 | 881 |
|
882 |
lemmas mult_poly_0 = mult_poly_0_left mult_poly_0_right |
|
883 |
||
| 52380 | 884 |
lemma mult_smult_left [simp]: |
885 |
"smult a p * q = smult a (p * q)" |
|
886 |
by (induct p) (simp add: mult_poly_0, simp add: smult_add_right) |
|
| 29474 | 887 |
|
| 52380 | 888 |
lemma mult_smult_right [simp]: |
889 |
"p * smult a q = smult a (p * q)" |
|
890 |
by (induct q) (simp add: mult_poly_0, simp add: smult_add_right) |
|
| 29474 | 891 |
|
892 |
lemma mult_poly_add_left: |
|
893 |
fixes p q r :: "'a poly" |
|
894 |
shows "(p + q) * r = p * r + q * r" |
|
| 52380 | 895 |
by (induct r) (simp add: mult_poly_0, simp add: smult_distribs algebra_simps) |
| 29451 | 896 |
|
| 60679 | 897 |
instance |
898 |
proof |
|
| 29451 | 899 |
fix p q r :: "'a poly" |
900 |
show 0: "0 * p = 0" |
|
| 29474 | 901 |
by (rule mult_poly_0_left) |
| 29451 | 902 |
show "p * 0 = 0" |
| 29474 | 903 |
by (rule mult_poly_0_right) |
| 29451 | 904 |
show "(p + q) * r = p * r + q * r" |
| 29474 | 905 |
by (rule mult_poly_add_left) |
| 29451 | 906 |
show "(p * q) * r = p * (q * r)" |
| 29474 | 907 |
by (induct p, simp add: mult_poly_0, simp add: mult_poly_add_left) |
| 29451 | 908 |
show "p * q = q * p" |
| 29474 | 909 |
by (induct p, simp add: mult_poly_0, simp) |
| 29451 | 910 |
qed |
911 |
||
912 |
end |
|
913 |
||
| 29540 | 914 |
instance poly :: (comm_semiring_0_cancel) comm_semiring_0_cancel .. |
915 |
||
| 29474 | 916 |
lemma coeff_mult: |
917 |
"coeff (p * q) n = (\<Sum>i\<le>n. coeff p i * coeff q (n-i))" |
|
918 |
proof (induct p arbitrary: n) |
|
919 |
case 0 show ?case by simp |
|
920 |
next |
|
921 |
case (pCons a p n) thus ?case |
|
922 |
by (cases n, simp, simp add: setsum_atMost_Suc_shift |
|
923 |
del: setsum_atMost_Suc) |
|
924 |
qed |
|
| 29451 | 925 |
|
| 29474 | 926 |
lemma degree_mult_le: "degree (p * q) \<le> degree p + degree q" |
927 |
apply (rule degree_le) |
|
928 |
apply (induct p) |
|
929 |
apply simp |
|
930 |
apply (simp add: coeff_eq_0 coeff_pCons split: nat.split) |
|
| 29451 | 931 |
done |
932 |
||
933 |
lemma mult_monom: "monom a m * monom b n = monom (a * b) (m + n)" |
|
| 60679 | 934 |
by (induct m) (simp add: monom_0 smult_monom, simp add: monom_Suc) |
| 29451 | 935 |
|
936 |
instantiation poly :: (comm_semiring_1) comm_semiring_1 |
|
937 |
begin |
|
938 |
||
| 60679 | 939 |
definition one_poly_def: "1 = pCons 1 0" |
| 29451 | 940 |
|
| 60679 | 941 |
instance |
942 |
proof |
|
943 |
show "1 * p = p" for p :: "'a poly" |
|
| 52380 | 944 |
unfolding one_poly_def by simp |
| 29451 | 945 |
show "0 \<noteq> (1::'a poly)" |
946 |
unfolding one_poly_def by simp |
|
947 |
qed |
|
948 |
||
949 |
end |
|
950 |
||
| 52380 | 951 |
instance poly :: (comm_ring) comm_ring .. |
952 |
||
953 |
instance poly :: (comm_ring_1) comm_ring_1 .. |
|
954 |
||
| 29451 | 955 |
lemma coeff_1 [simp]: "coeff 1 n = (if n = 0 then 1 else 0)" |
956 |
unfolding one_poly_def |
|
957 |
by (simp add: coeff_pCons split: nat.split) |
|
958 |
||
| 60570 | 959 |
lemma monom_eq_1 [simp]: |
960 |
"monom 1 0 = 1" |
|
961 |
by (simp add: monom_0 one_poly_def) |
|
962 |
||
| 29451 | 963 |
lemma degree_1 [simp]: "degree 1 = 0" |
964 |
unfolding one_poly_def |
|
965 |
by (rule degree_pCons_0) |
|
966 |
||
| 52380 | 967 |
lemma coeffs_1_eq [simp, code abstract]: |
968 |
"coeffs 1 = [1]" |
|
969 |
by (simp add: one_poly_def) |
|
970 |
||
971 |
lemma degree_power_le: |
|
972 |
"degree (p ^ n) \<le> degree p * n" |
|
973 |
by (induct n) (auto intro: order_trans degree_mult_le) |
|
974 |
||
975 |
lemma poly_smult [simp]: |
|
976 |
"poly (smult a p) x = a * poly p x" |
|
977 |
by (induct p, simp, simp add: algebra_simps) |
|
978 |
||
979 |
lemma poly_mult [simp]: |
|
980 |
"poly (p * q) x = poly p x * poly q x" |
|
981 |
by (induct p, simp_all, simp add: algebra_simps) |
|
982 |
||
983 |
lemma poly_1 [simp]: |
|
984 |
"poly 1 x = 1" |
|
985 |
by (simp add: one_poly_def) |
|
986 |
||
987 |
lemma poly_power [simp]: |
|
988 |
fixes p :: "'a::{comm_semiring_1} poly"
|
|
989 |
shows "poly (p ^ n) x = poly p x ^ n" |
|
990 |
by (induct n) simp_all |
|
991 |
||
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
992 |
lemma poly_setprod: "poly (\<Prod>k\<in>A. p k) x = (\<Prod>k\<in>A. poly (p k) x)" |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
993 |
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
|
994 |
|
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
995 |
lemma degree_setprod_setsum_le: "finite S \<Longrightarrow> degree (setprod f S) \<le> setsum (degree o f) S" |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
996 |
proof (induct S rule: finite_induct) |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
997 |
case (insert a S) |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
998 |
show ?case unfolding setprod.insert[OF insert(1-2)] setsum.insert[OF insert(1-2)] |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
999 |
by (rule le_trans[OF degree_mult_le], insert insert, auto) |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1000 |
qed simp |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1001 |
|
| 62065 | 1002 |
subsection \<open>Conversions from natural numbers\<close> |
1003 |
||
1004 |
lemma of_nat_poly: "of_nat n = [:of_nat n :: 'a :: comm_semiring_1:]" |
|
1005 |
proof (induction n) |
|
1006 |
case (Suc n) |
|
1007 |
hence "of_nat (Suc n) = 1 + (of_nat n :: 'a poly)" |
|
1008 |
by simp |
|
1009 |
also have "(of_nat n :: 'a poly) = [: of_nat n :]" |
|
1010 |
by (subst Suc) (rule refl) |
|
1011 |
also have "1 = [:1:]" by (simp add: one_poly_def) |
|
1012 |
finally show ?case by (subst (asm) add_pCons) simp |
|
1013 |
qed simp |
|
1014 |
||
1015 |
lemma degree_of_nat [simp]: "degree (of_nat n) = 0" |
|
1016 |
by (simp add: of_nat_poly) |
|
1017 |
||
1018 |
lemma degree_numeral [simp]: "degree (numeral n) = 0" |
|
1019 |
by (subst of_nat_numeral [symmetric], subst of_nat_poly) simp |
|
1020 |
||
1021 |
lemma numeral_poly: "numeral n = [:numeral n:]" |
|
1022 |
by (subst of_nat_numeral [symmetric], subst of_nat_poly) simp |
|
| 52380 | 1023 |
|
| 60500 | 1024 |
subsection \<open>Lemmas about divisibility\<close> |
| 29979 | 1025 |
|
1026 |
lemma dvd_smult: "p dvd q \<Longrightarrow> p dvd smult a q" |
|
1027 |
proof - |
|
1028 |
assume "p dvd q" |
|
1029 |
then obtain k where "q = p * k" .. |
|
1030 |
then have "smult a q = p * smult a k" by simp |
|
1031 |
then show "p dvd smult a q" .. |
|
1032 |
qed |
|
1033 |
||
1034 |
lemma dvd_smult_cancel: |
|
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1035 |
fixes a :: "'a :: field" |
| 29979 | 1036 |
shows "p dvd smult a q \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> p dvd q" |
1037 |
by (drule dvd_smult [where a="inverse a"]) simp |
|
1038 |
||
1039 |
lemma dvd_smult_iff: |
|
1040 |
fixes a :: "'a::field" |
|
1041 |
shows "a \<noteq> 0 \<Longrightarrow> p dvd smult a q \<longleftrightarrow> p dvd q" |
|
1042 |
by (safe elim!: dvd_smult dvd_smult_cancel) |
|
1043 |
||
| 31663 | 1044 |
lemma smult_dvd_cancel: |
1045 |
"smult a p dvd q \<Longrightarrow> p dvd q" |
|
1046 |
proof - |
|
1047 |
assume "smult a p dvd q" |
|
1048 |
then obtain k where "q = smult a p * k" .. |
|
1049 |
then have "q = p * smult a k" by simp |
|
1050 |
then show "p dvd q" .. |
|
1051 |
qed |
|
1052 |
||
1053 |
lemma smult_dvd: |
|
1054 |
fixes a :: "'a::field" |
|
1055 |
shows "p dvd q \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> smult a p dvd q" |
|
1056 |
by (rule smult_dvd_cancel [where a="inverse a"]) simp |
|
1057 |
||
1058 |
lemma smult_dvd_iff: |
|
1059 |
fixes a :: "'a::field" |
|
1060 |
shows "smult a p dvd q \<longleftrightarrow> (if a = 0 then q = 0 else p dvd q)" |
|
1061 |
by (auto elim: smult_dvd smult_dvd_cancel) |
|
1062 |
||
| 29451 | 1063 |
|
| 60500 | 1064 |
subsection \<open>Polynomials form an integral domain\<close> |
| 29451 | 1065 |
|
1066 |
lemma coeff_mult_degree_sum: |
|
1067 |
"coeff (p * q) (degree p + degree q) = |
|
1068 |
coeff p (degree p) * coeff q (degree q)" |
|
| 29471 | 1069 |
by (induct p, simp, simp add: coeff_eq_0) |
| 29451 | 1070 |
|
1071 |
instance poly :: (idom) idom |
|
1072 |
proof |
|
1073 |
fix p q :: "'a poly" |
|
1074 |
assume "p \<noteq> 0" and "q \<noteq> 0" |
|
1075 |
have "coeff (p * q) (degree p + degree q) = |
|
1076 |
coeff p (degree p) * coeff q (degree q)" |
|
1077 |
by (rule coeff_mult_degree_sum) |
|
1078 |
also have "coeff p (degree p) * coeff q (degree q) \<noteq> 0" |
|
| 60500 | 1079 |
using \<open>p \<noteq> 0\<close> and \<open>q \<noteq> 0\<close> by simp |
| 29451 | 1080 |
finally have "\<exists>n. coeff (p * q) n \<noteq> 0" .. |
| 52380 | 1081 |
thus "p * q \<noteq> 0" by (simp add: poly_eq_iff) |
| 29451 | 1082 |
qed |
1083 |
||
1084 |
lemma degree_mult_eq: |
|
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1085 |
fixes p q :: "'a::semidom poly" |
| 29451 | 1086 |
shows "\<lbrakk>p \<noteq> 0; q \<noteq> 0\<rbrakk> \<Longrightarrow> degree (p * q) = degree p + degree q" |
1087 |
apply (rule order_antisym [OF degree_mult_le le_degree]) |
|
1088 |
apply (simp add: coeff_mult_degree_sum) |
|
1089 |
done |
|
1090 |
||
| 60570 | 1091 |
lemma degree_mult_right_le: |
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1092 |
fixes p q :: "'a::semidom poly" |
| 60570 | 1093 |
assumes "q \<noteq> 0" |
1094 |
shows "degree p \<le> degree (p * q)" |
|
1095 |
using assms by (cases "p = 0") (simp_all add: degree_mult_eq) |
|
1096 |
||
1097 |
lemma coeff_degree_mult: |
|
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1098 |
fixes p q :: "'a::semidom poly" |
| 60570 | 1099 |
shows "coeff (p * q) (degree (p * q)) = |
1100 |
coeff q (degree q) * coeff p (degree p)" |
|
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1101 |
by (cases "p = 0 \<or> q = 0") (auto simp add: degree_mult_eq coeff_mult_degree_sum mult_ac) |
| 60570 | 1102 |
|
| 29451 | 1103 |
lemma dvd_imp_degree_le: |
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1104 |
fixes p q :: "'a::semidom poly" |
| 29451 | 1105 |
shows "\<lbrakk>p dvd q; q \<noteq> 0\<rbrakk> \<Longrightarrow> degree p \<le> degree q" |
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1106 |
by (erule dvdE, hypsubst, subst degree_mult_eq) auto |
| 29451 | 1107 |
|
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1108 |
lemma divides_degree: |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1109 |
assumes pq: "p dvd (q :: 'a :: semidom poly)" |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1110 |
shows "degree p \<le> degree q \<or> q = 0" |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1111 |
by (metis dvd_imp_degree_le pq) |
| 29451 | 1112 |
|
| 60500 | 1113 |
subsection \<open>Polynomials form an ordered integral domain\<close> |
| 29878 | 1114 |
|
| 52380 | 1115 |
definition pos_poly :: "'a::linordered_idom poly \<Rightarrow> bool" |
| 29878 | 1116 |
where |
1117 |
"pos_poly p \<longleftrightarrow> 0 < coeff p (degree p)" |
|
1118 |
||
1119 |
lemma pos_poly_pCons: |
|
1120 |
"pos_poly (pCons a p) \<longleftrightarrow> pos_poly p \<or> (p = 0 \<and> 0 < a)" |
|
1121 |
unfolding pos_poly_def by simp |
|
1122 |
||
1123 |
lemma not_pos_poly_0 [simp]: "\<not> pos_poly 0" |
|
1124 |
unfolding pos_poly_def by simp |
|
1125 |
||
1126 |
lemma pos_poly_add: "\<lbrakk>pos_poly p; pos_poly q\<rbrakk> \<Longrightarrow> pos_poly (p + q)" |
|
1127 |
apply (induct p arbitrary: q, simp) |
|
1128 |
apply (case_tac q, force simp add: pos_poly_pCons add_pos_pos) |
|
1129 |
done |
|
1130 |
||
1131 |
lemma pos_poly_mult: "\<lbrakk>pos_poly p; pos_poly q\<rbrakk> \<Longrightarrow> pos_poly (p * q)" |
|
1132 |
unfolding pos_poly_def |
|
1133 |
apply (subgoal_tac "p \<noteq> 0 \<and> q \<noteq> 0") |
|
| 56544 | 1134 |
apply (simp add: degree_mult_eq coeff_mult_degree_sum) |
| 29878 | 1135 |
apply auto |
1136 |
done |
|
1137 |
||
1138 |
lemma pos_poly_total: "p = 0 \<or> pos_poly p \<or> pos_poly (- p)" |
|
1139 |
by (induct p) (auto simp add: pos_poly_pCons) |
|
1140 |
||
| 52380 | 1141 |
lemma last_coeffs_eq_coeff_degree: |
1142 |
"p \<noteq> 0 \<Longrightarrow> last (coeffs p) = coeff p (degree p)" |
|
1143 |
by (simp add: coeffs_def) |
|
1144 |
||
1145 |
lemma pos_poly_coeffs [code]: |
|
1146 |
"pos_poly p \<longleftrightarrow> (let as = coeffs p in as \<noteq> [] \<and> last as > 0)" (is "?P \<longleftrightarrow> ?Q") |
|
1147 |
proof |
|
1148 |
assume ?Q then show ?P by (auto simp add: pos_poly_def last_coeffs_eq_coeff_degree) |
|
1149 |
next |
|
1150 |
assume ?P then have *: "0 < coeff p (degree p)" by (simp add: pos_poly_def) |
|
1151 |
then have "p \<noteq> 0" by auto |
|
1152 |
with * show ?Q by (simp add: last_coeffs_eq_coeff_degree) |
|
1153 |
qed |
|
1154 |
||
|
35028
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
34973
diff
changeset
|
1155 |
instantiation poly :: (linordered_idom) linordered_idom |
| 29878 | 1156 |
begin |
1157 |
||
1158 |
definition |
|
| 37765 | 1159 |
"x < y \<longleftrightarrow> pos_poly (y - x)" |
| 29878 | 1160 |
|
1161 |
definition |
|
| 37765 | 1162 |
"x \<le> y \<longleftrightarrow> x = y \<or> pos_poly (y - x)" |
| 29878 | 1163 |
|
1164 |
definition |
|
| 61945 | 1165 |
"\<bar>x::'a poly\<bar> = (if x < 0 then - x else x)" |
| 29878 | 1166 |
|
1167 |
definition |
|
| 37765 | 1168 |
"sgn (x::'a poly) = (if x = 0 then 0 else if 0 < x then 1 else - 1)" |
| 29878 | 1169 |
|
| 60679 | 1170 |
instance |
1171 |
proof |
|
1172 |
fix x y z :: "'a poly" |
|
| 29878 | 1173 |
show "x < y \<longleftrightarrow> x \<le> y \<and> \<not> y \<le> x" |
1174 |
unfolding less_eq_poly_def less_poly_def |
|
1175 |
apply safe |
|
1176 |
apply simp |
|
1177 |
apply (drule (1) pos_poly_add) |
|
1178 |
apply simp |
|
1179 |
done |
|
| 60679 | 1180 |
show "x \<le> x" |
| 29878 | 1181 |
unfolding less_eq_poly_def by simp |
| 60679 | 1182 |
show "x \<le> y \<Longrightarrow> y \<le> z \<Longrightarrow> x \<le> z" |
| 29878 | 1183 |
unfolding less_eq_poly_def |
1184 |
apply safe |
|
1185 |
apply (drule (1) pos_poly_add) |
|
1186 |
apply (simp add: algebra_simps) |
|
1187 |
done |
|
| 60679 | 1188 |
show "x \<le> y \<Longrightarrow> y \<le> x \<Longrightarrow> x = y" |
| 29878 | 1189 |
unfolding less_eq_poly_def |
1190 |
apply safe |
|
1191 |
apply (drule (1) pos_poly_add) |
|
1192 |
apply simp |
|
1193 |
done |
|
| 60679 | 1194 |
show "x \<le> y \<Longrightarrow> z + x \<le> z + y" |
| 29878 | 1195 |
unfolding less_eq_poly_def |
1196 |
apply safe |
|
1197 |
apply (simp add: algebra_simps) |
|
1198 |
done |
|
1199 |
show "x \<le> y \<or> y \<le> x" |
|
1200 |
unfolding less_eq_poly_def |
|
1201 |
using pos_poly_total [of "x - y"] |
|
1202 |
by auto |
|
| 60679 | 1203 |
show "x < y \<Longrightarrow> 0 < z \<Longrightarrow> z * x < z * y" |
| 29878 | 1204 |
unfolding less_poly_def |
1205 |
by (simp add: right_diff_distrib [symmetric] pos_poly_mult) |
|
1206 |
show "\<bar>x\<bar> = (if x < 0 then - x else x)" |
|
1207 |
by (rule abs_poly_def) |
|
1208 |
show "sgn x = (if x = 0 then 0 else if 0 < x then 1 else - 1)" |
|
1209 |
by (rule sgn_poly_def) |
|
1210 |
qed |
|
1211 |
||
1212 |
end |
|
1213 |
||
| 60500 | 1214 |
text \<open>TODO: Simplification rules for comparisons\<close> |
| 29878 | 1215 |
|
1216 |
||
| 60500 | 1217 |
subsection \<open>Synthetic division and polynomial roots\<close> |
| 52380 | 1218 |
|
| 60500 | 1219 |
text \<open> |
| 52380 | 1220 |
Synthetic division is simply division by the linear polynomial @{term "x - c"}.
|
| 60500 | 1221 |
\<close> |
| 52380 | 1222 |
|
1223 |
definition synthetic_divmod :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a poly \<times> 'a" |
|
1224 |
where |
|
1225 |
"synthetic_divmod p c = fold_coeffs (\<lambda>a (q, r). (pCons r q, a + c * r)) p (0, 0)" |
|
1226 |
||
1227 |
definition synthetic_div :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a poly" |
|
1228 |
where |
|
1229 |
"synthetic_div p c = fst (synthetic_divmod p c)" |
|
1230 |
||
1231 |
lemma synthetic_divmod_0 [simp]: |
|
1232 |
"synthetic_divmod 0 c = (0, 0)" |
|
1233 |
by (simp add: synthetic_divmod_def) |
|
1234 |
||
1235 |
lemma synthetic_divmod_pCons [simp]: |
|
1236 |
"synthetic_divmod (pCons a p) c = (\<lambda>(q, r). (pCons r q, a + c * r)) (synthetic_divmod p c)" |
|
1237 |
by (cases "p = 0 \<and> a = 0") (auto simp add: synthetic_divmod_def) |
|
1238 |
||
1239 |
lemma synthetic_div_0 [simp]: |
|
1240 |
"synthetic_div 0 c = 0" |
|
1241 |
unfolding synthetic_div_def by simp |
|
1242 |
||
1243 |
lemma synthetic_div_unique_lemma: "smult c p = pCons a p \<Longrightarrow> p = 0" |
|
1244 |
by (induct p arbitrary: a) simp_all |
|
1245 |
||
1246 |
lemma snd_synthetic_divmod: |
|
1247 |
"snd (synthetic_divmod p c) = poly p c" |
|
1248 |
by (induct p, simp, simp add: split_def) |
|
1249 |
||
1250 |
lemma synthetic_div_pCons [simp]: |
|
1251 |
"synthetic_div (pCons a p) c = pCons (poly p c) (synthetic_div p c)" |
|
1252 |
unfolding synthetic_div_def |
|
1253 |
by (simp add: split_def snd_synthetic_divmod) |
|
1254 |
||
1255 |
lemma synthetic_div_eq_0_iff: |
|
1256 |
"synthetic_div p c = 0 \<longleftrightarrow> degree p = 0" |
|
1257 |
by (induct p, simp, case_tac p, simp) |
|
1258 |
||
1259 |
lemma degree_synthetic_div: |
|
1260 |
"degree (synthetic_div p c) = degree p - 1" |
|
1261 |
by (induct p, simp, simp add: synthetic_div_eq_0_iff) |
|
1262 |
||
1263 |
lemma synthetic_div_correct: |
|
1264 |
"p + smult c (synthetic_div p c) = pCons (poly p c) (synthetic_div p c)" |
|
1265 |
by (induct p) simp_all |
|
1266 |
||
1267 |
lemma synthetic_div_unique: |
|
1268 |
"p + smult c q = pCons r q \<Longrightarrow> r = poly p c \<and> q = synthetic_div p c" |
|
1269 |
apply (induct p arbitrary: q r) |
|
1270 |
apply (simp, frule synthetic_div_unique_lemma, simp) |
|
1271 |
apply (case_tac q, force) |
|
1272 |
done |
|
1273 |
||
1274 |
lemma synthetic_div_correct': |
|
1275 |
fixes c :: "'a::comm_ring_1" |
|
1276 |
shows "[:-c, 1:] * synthetic_div p c + [:poly p c:] = p" |
|
1277 |
using synthetic_div_correct [of p c] |
|
1278 |
by (simp add: algebra_simps) |
|
1279 |
||
1280 |
lemma poly_eq_0_iff_dvd: |
|
1281 |
fixes c :: "'a::idom" |
|
1282 |
shows "poly p c = 0 \<longleftrightarrow> [:-c, 1:] dvd p" |
|
1283 |
proof |
|
1284 |
assume "poly p c = 0" |
|
1285 |
with synthetic_div_correct' [of c p] |
|
1286 |
have "p = [:-c, 1:] * synthetic_div p c" by simp |
|
1287 |
then show "[:-c, 1:] dvd p" .. |
|
1288 |
next |
|
1289 |
assume "[:-c, 1:] dvd p" |
|
1290 |
then obtain k where "p = [:-c, 1:] * k" by (rule dvdE) |
|
1291 |
then show "poly p c = 0" by simp |
|
1292 |
qed |
|
1293 |
||
1294 |
lemma dvd_iff_poly_eq_0: |
|
1295 |
fixes c :: "'a::idom" |
|
1296 |
shows "[:c, 1:] dvd p \<longleftrightarrow> poly p (-c) = 0" |
|
1297 |
by (simp add: poly_eq_0_iff_dvd) |
|
1298 |
||
1299 |
lemma poly_roots_finite: |
|
1300 |
fixes p :: "'a::idom poly" |
|
1301 |
shows "p \<noteq> 0 \<Longrightarrow> finite {x. poly p x = 0}"
|
|
1302 |
proof (induct n \<equiv> "degree p" arbitrary: p) |
|
1303 |
case (0 p) |
|
1304 |
then obtain a where "a \<noteq> 0" and "p = [:a:]" |
|
1305 |
by (cases p, simp split: if_splits) |
|
1306 |
then show "finite {x. poly p x = 0}" by simp
|
|
1307 |
next |
|
1308 |
case (Suc n p) |
|
1309 |
show "finite {x. poly p x = 0}"
|
|
1310 |
proof (cases "\<exists>x. poly p x = 0") |
|
1311 |
case False |
|
1312 |
then show "finite {x. poly p x = 0}" by simp
|
|
1313 |
next |
|
1314 |
case True |
|
1315 |
then obtain a where "poly p a = 0" .. |
|
1316 |
then have "[:-a, 1:] dvd p" by (simp only: poly_eq_0_iff_dvd) |
|
1317 |
then obtain k where k: "p = [:-a, 1:] * k" .. |
|
| 60500 | 1318 |
with \<open>p \<noteq> 0\<close> have "k \<noteq> 0" by auto |
| 52380 | 1319 |
with k have "degree p = Suc (degree k)" |
1320 |
by (simp add: degree_mult_eq del: mult_pCons_left) |
|
| 60500 | 1321 |
with \<open>Suc n = degree p\<close> have "n = degree k" by simp |
1322 |
then have "finite {x. poly k x = 0}" using \<open>k \<noteq> 0\<close> by (rule Suc.hyps)
|
|
| 52380 | 1323 |
then have "finite (insert a {x. poly k x = 0})" by simp
|
1324 |
then show "finite {x. poly p x = 0}"
|
|
| 57862 | 1325 |
by (simp add: k Collect_disj_eq del: mult_pCons_left) |
| 52380 | 1326 |
qed |
1327 |
qed |
|
1328 |
||
1329 |
lemma poly_eq_poly_eq_iff: |
|
1330 |
fixes p q :: "'a::{idom,ring_char_0} poly"
|
|
1331 |
shows "poly p = poly q \<longleftrightarrow> p = q" (is "?P \<longleftrightarrow> ?Q") |
|
1332 |
proof |
|
1333 |
assume ?Q then show ?P by simp |
|
1334 |
next |
|
1335 |
{ fix p :: "'a::{idom,ring_char_0} poly"
|
|
1336 |
have "poly p = poly 0 \<longleftrightarrow> p = 0" |
|
1337 |
apply (cases "p = 0", simp_all) |
|
1338 |
apply (drule poly_roots_finite) |
|
1339 |
apply (auto simp add: infinite_UNIV_char_0) |
|
1340 |
done |
|
1341 |
} note this [of "p - q"] |
|
1342 |
moreover assume ?P |
|
1343 |
ultimately show ?Q by auto |
|
1344 |
qed |
|
1345 |
||
1346 |
lemma poly_all_0_iff_0: |
|
1347 |
fixes p :: "'a::{ring_char_0, idom} poly"
|
|
1348 |
shows "(\<forall>x. poly p x = 0) \<longleftrightarrow> p = 0" |
|
1349 |
by (auto simp add: poly_eq_poly_eq_iff [symmetric]) |
|
1350 |
||
1351 |
||
| 60500 | 1352 |
subsection \<open>Long division of polynomials\<close> |
| 29451 | 1353 |
|
| 52380 | 1354 |
definition pdivmod_rel :: "'a::field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<Rightarrow> bool" |
| 29451 | 1355 |
where |
| 29537 | 1356 |
"pdivmod_rel x y q r \<longleftrightarrow> |
| 29451 | 1357 |
x = q * y + r \<and> (if y = 0 then q = 0 else r = 0 \<or> degree r < degree y)" |
1358 |
||
| 29537 | 1359 |
lemma pdivmod_rel_0: |
1360 |
"pdivmod_rel 0 y 0 0" |
|
1361 |
unfolding pdivmod_rel_def by simp |
|
| 29451 | 1362 |
|
| 29537 | 1363 |
lemma pdivmod_rel_by_0: |
1364 |
"pdivmod_rel x 0 0 x" |
|
1365 |
unfolding pdivmod_rel_def by simp |
|
| 29451 | 1366 |
|
1367 |
lemma eq_zero_or_degree_less: |
|
1368 |
assumes "degree p \<le> n" and "coeff p n = 0" |
|
1369 |
shows "p = 0 \<or> degree p < n" |
|
1370 |
proof (cases n) |
|
1371 |
case 0 |
|
| 60500 | 1372 |
with \<open>degree p \<le> n\<close> and \<open>coeff p n = 0\<close> |
| 29451 | 1373 |
have "coeff p (degree p) = 0" by simp |
1374 |
then have "p = 0" by simp |
|
1375 |
then show ?thesis .. |
|
1376 |
next |
|
1377 |
case (Suc m) |
|
1378 |
have "\<forall>i>n. coeff p i = 0" |
|
| 60500 | 1379 |
using \<open>degree p \<le> n\<close> by (simp add: coeff_eq_0) |
| 29451 | 1380 |
then have "\<forall>i\<ge>n. coeff p i = 0" |
| 60500 | 1381 |
using \<open>coeff p n = 0\<close> by (simp add: le_less) |
| 29451 | 1382 |
then have "\<forall>i>m. coeff p i = 0" |
| 60500 | 1383 |
using \<open>n = Suc m\<close> by (simp add: less_eq_Suc_le) |
| 29451 | 1384 |
then have "degree p \<le> m" |
1385 |
by (rule degree_le) |
|
1386 |
then have "degree p < n" |
|
| 60500 | 1387 |
using \<open>n = Suc m\<close> by (simp add: less_Suc_eq_le) |
| 29451 | 1388 |
then show ?thesis .. |
1389 |
qed |
|
1390 |
||
| 29537 | 1391 |
lemma pdivmod_rel_pCons: |
1392 |
assumes rel: "pdivmod_rel x y q r" |
|
| 29451 | 1393 |
assumes y: "y \<noteq> 0" |
1394 |
assumes b: "b = coeff (pCons a r) (degree y) / coeff y (degree y)" |
|
| 29537 | 1395 |
shows "pdivmod_rel (pCons a x) y (pCons b q) (pCons a r - smult b y)" |
1396 |
(is "pdivmod_rel ?x y ?q ?r") |
|
| 29451 | 1397 |
proof - |
1398 |
have x: "x = q * y + r" and r: "r = 0 \<or> degree r < degree y" |
|
| 29537 | 1399 |
using assms unfolding pdivmod_rel_def by simp_all |
| 29451 | 1400 |
|
1401 |
have 1: "?x = ?q * y + ?r" |
|
1402 |
using b x by simp |
|
1403 |
||
1404 |
have 2: "?r = 0 \<or> degree ?r < degree y" |
|
1405 |
proof (rule eq_zero_or_degree_less) |
|
| 29539 | 1406 |
show "degree ?r \<le> degree y" |
1407 |
proof (rule degree_diff_le) |
|
| 29451 | 1408 |
show "degree (pCons a r) \<le> degree y" |
|
29460
ad87e5d1488b
new lemmas about synthetic_div; declare degree_pCons_eq_if [simp]
huffman
parents:
29457
diff
changeset
|
1409 |
using r by auto |
| 29451 | 1410 |
show "degree (smult b y) \<le> degree y" |
1411 |
by (rule degree_smult_le) |
|
1412 |
qed |
|
1413 |
next |
|
1414 |
show "coeff ?r (degree y) = 0" |
|
| 60500 | 1415 |
using \<open>y \<noteq> 0\<close> unfolding b by simp |
| 29451 | 1416 |
qed |
1417 |
||
1418 |
from 1 2 show ?thesis |
|
| 29537 | 1419 |
unfolding pdivmod_rel_def |
| 60500 | 1420 |
using \<open>y \<noteq> 0\<close> by simp |
| 29451 | 1421 |
qed |
1422 |
||
| 29537 | 1423 |
lemma pdivmod_rel_exists: "\<exists>q r. pdivmod_rel x y q r" |
| 29451 | 1424 |
apply (cases "y = 0") |
| 29537 | 1425 |
apply (fast intro!: pdivmod_rel_by_0) |
| 29451 | 1426 |
apply (induct x) |
| 29537 | 1427 |
apply (fast intro!: pdivmod_rel_0) |
1428 |
apply (fast intro!: pdivmod_rel_pCons) |
|
| 29451 | 1429 |
done |
1430 |
||
| 29537 | 1431 |
lemma pdivmod_rel_unique: |
1432 |
assumes 1: "pdivmod_rel x y q1 r1" |
|
1433 |
assumes 2: "pdivmod_rel x y q2 r2" |
|
| 29451 | 1434 |
shows "q1 = q2 \<and> r1 = r2" |
1435 |
proof (cases "y = 0") |
|
1436 |
assume "y = 0" with assms show ?thesis |
|
| 29537 | 1437 |
by (simp add: pdivmod_rel_def) |
| 29451 | 1438 |
next |
1439 |
assume [simp]: "y \<noteq> 0" |
|
1440 |
from 1 have q1: "x = q1 * y + r1" and r1: "r1 = 0 \<or> degree r1 < degree y" |
|
| 29537 | 1441 |
unfolding pdivmod_rel_def by simp_all |
| 29451 | 1442 |
from 2 have q2: "x = q2 * y + r2" and r2: "r2 = 0 \<or> degree r2 < degree y" |
| 29537 | 1443 |
unfolding pdivmod_rel_def by simp_all |
| 29451 | 1444 |
from q1 q2 have q3: "(q1 - q2) * y = r2 - r1" |
| 29667 | 1445 |
by (simp add: algebra_simps) |
| 29451 | 1446 |
from r1 r2 have r3: "(r2 - r1) = 0 \<or> degree (r2 - r1) < degree y" |
| 29453 | 1447 |
by (auto intro: degree_diff_less) |
| 29451 | 1448 |
|
1449 |
show "q1 = q2 \<and> r1 = r2" |
|
1450 |
proof (rule ccontr) |
|
1451 |
assume "\<not> (q1 = q2 \<and> r1 = r2)" |
|
1452 |
with q3 have "q1 \<noteq> q2" and "r1 \<noteq> r2" by auto |
|
1453 |
with r3 have "degree (r2 - r1) < degree y" by simp |
|
1454 |
also have "degree y \<le> degree (q1 - q2) + degree y" by simp |
|
1455 |
also have "\<dots> = degree ((q1 - q2) * y)" |
|
| 60500 | 1456 |
using \<open>q1 \<noteq> q2\<close> by (simp add: degree_mult_eq) |
| 29451 | 1457 |
also have "\<dots> = degree (r2 - r1)" |
1458 |
using q3 by simp |
|
1459 |
finally have "degree (r2 - r1) < degree (r2 - r1)" . |
|
1460 |
then show "False" by simp |
|
1461 |
qed |
|
1462 |
qed |
|
1463 |
||
| 29660 | 1464 |
lemma pdivmod_rel_0_iff: "pdivmod_rel 0 y q r \<longleftrightarrow> q = 0 \<and> r = 0" |
1465 |
by (auto dest: pdivmod_rel_unique intro: pdivmod_rel_0) |
|
1466 |
||
1467 |
lemma pdivmod_rel_by_0_iff: "pdivmod_rel x 0 q r \<longleftrightarrow> q = 0 \<and> r = x" |
|
1468 |
by (auto dest: pdivmod_rel_unique intro: pdivmod_rel_by_0) |
|
1469 |
||
| 45605 | 1470 |
lemmas pdivmod_rel_unique_div = pdivmod_rel_unique [THEN conjunct1] |
| 29451 | 1471 |
|
| 45605 | 1472 |
lemmas pdivmod_rel_unique_mod = pdivmod_rel_unique [THEN conjunct2] |
| 29451 | 1473 |
|
1474 |
instantiation poly :: (field) ring_div |
|
1475 |
begin |
|
1476 |
||
|
60352
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
haftmann
parents:
60040
diff
changeset
|
1477 |
definition divide_poly where |
|
60429
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
haftmann
parents:
60352
diff
changeset
|
1478 |
div_poly_def: "x div y = (THE q. \<exists>r. pdivmod_rel x y q r)" |
| 29451 | 1479 |
|
1480 |
definition mod_poly where |
|
| 37765 | 1481 |
"x mod y = (THE r. \<exists>q. pdivmod_rel x y q r)" |
| 29451 | 1482 |
|
1483 |
lemma div_poly_eq: |
|
|
60429
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
haftmann
parents:
60352
diff
changeset
|
1484 |
"pdivmod_rel x y q r \<Longrightarrow> x div y = q" |
| 29451 | 1485 |
unfolding div_poly_def |
| 29537 | 1486 |
by (fast elim: pdivmod_rel_unique_div) |
| 29451 | 1487 |
|
1488 |
lemma mod_poly_eq: |
|
| 29537 | 1489 |
"pdivmod_rel x y q r \<Longrightarrow> x mod y = r" |
| 29451 | 1490 |
unfolding mod_poly_def |
| 29537 | 1491 |
by (fast elim: pdivmod_rel_unique_mod) |
| 29451 | 1492 |
|
| 29537 | 1493 |
lemma pdivmod_rel: |
|
60429
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
haftmann
parents:
60352
diff
changeset
|
1494 |
"pdivmod_rel x y (x div y) (x mod y)" |
| 29451 | 1495 |
proof - |
| 29537 | 1496 |
from pdivmod_rel_exists |
1497 |
obtain q r where "pdivmod_rel x y q r" by fast |
|
| 29451 | 1498 |
thus ?thesis |
1499 |
by (simp add: div_poly_eq mod_poly_eq) |
|
1500 |
qed |
|
1501 |
||
| 60679 | 1502 |
instance |
1503 |
proof |
|
| 29451 | 1504 |
fix x y :: "'a poly" |
|
60429
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
haftmann
parents:
60352
diff
changeset
|
1505 |
show "x div y * y + x mod y = x" |
| 29537 | 1506 |
using pdivmod_rel [of x y] |
1507 |
by (simp add: pdivmod_rel_def) |
|
| 29451 | 1508 |
next |
1509 |
fix x :: "'a poly" |
|
| 29537 | 1510 |
have "pdivmod_rel x 0 0 x" |
1511 |
by (rule pdivmod_rel_by_0) |
|
|
60429
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
haftmann
parents:
60352
diff
changeset
|
1512 |
thus "x div 0 = 0" |
| 29451 | 1513 |
by (rule div_poly_eq) |
1514 |
next |
|
1515 |
fix y :: "'a poly" |
|
| 29537 | 1516 |
have "pdivmod_rel 0 y 0 0" |
1517 |
by (rule pdivmod_rel_0) |
|
|
60429
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
haftmann
parents:
60352
diff
changeset
|
1518 |
thus "0 div y = 0" |
| 29451 | 1519 |
by (rule div_poly_eq) |
1520 |
next |
|
1521 |
fix x y z :: "'a poly" |
|
1522 |
assume "y \<noteq> 0" |
|
|
60429
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
haftmann
parents:
60352
diff
changeset
|
1523 |
hence "pdivmod_rel (x + z * y) y (z + x div y) (x mod y)" |
| 29537 | 1524 |
using pdivmod_rel [of x y] |
|
49962
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents:
49834
diff
changeset
|
1525 |
by (simp add: pdivmod_rel_def distrib_right) |
|
60429
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
haftmann
parents:
60352
diff
changeset
|
1526 |
thus "(x + z * y) div y = z + x div y" |
| 29451 | 1527 |
by (rule div_poly_eq) |
| 30930 | 1528 |
next |
1529 |
fix x y z :: "'a poly" |
|
1530 |
assume "x \<noteq> 0" |
|
|
60429
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
haftmann
parents:
60352
diff
changeset
|
1531 |
show "(x * y) div (x * z) = y div z" |
| 30930 | 1532 |
proof (cases "y \<noteq> 0 \<and> z \<noteq> 0") |
1533 |
have "\<And>x::'a poly. pdivmod_rel x 0 0 x" |
|
1534 |
by (rule pdivmod_rel_by_0) |
|
|
60429
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
haftmann
parents:
60352
diff
changeset
|
1535 |
then have [simp]: "\<And>x::'a poly. x div 0 = 0" |
| 30930 | 1536 |
by (rule div_poly_eq) |
1537 |
have "\<And>x::'a poly. pdivmod_rel 0 x 0 0" |
|
1538 |
by (rule pdivmod_rel_0) |
|
|
60429
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
haftmann
parents:
60352
diff
changeset
|
1539 |
then have [simp]: "\<And>x::'a poly. 0 div x = 0" |
| 30930 | 1540 |
by (rule div_poly_eq) |
1541 |
case False then show ?thesis by auto |
|
1542 |
next |
|
1543 |
case True then have "y \<noteq> 0" and "z \<noteq> 0" by auto |
|
| 60500 | 1544 |
with \<open>x \<noteq> 0\<close> |
| 30930 | 1545 |
have "\<And>q r. pdivmod_rel y z q r \<Longrightarrow> pdivmod_rel (x * y) (x * z) q (x * r)" |
1546 |
by (auto simp add: pdivmod_rel_def algebra_simps) |
|
1547 |
(rule classical, simp add: degree_mult_eq) |
|
|
60429
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
haftmann
parents:
60352
diff
changeset
|
1548 |
moreover from pdivmod_rel have "pdivmod_rel y z (y div z) (y mod z)" . |
|
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
haftmann
parents:
60352
diff
changeset
|
1549 |
ultimately have "pdivmod_rel (x * y) (x * z) (y div z) (x * (y mod z))" . |
| 30930 | 1550 |
then show ?thesis by (simp add: div_poly_eq) |
1551 |
qed |
|
| 29451 | 1552 |
qed |
1553 |
||
1554 |
end |
|
1555 |
||
| 60570 | 1556 |
lemma is_unit_monom_0: |
1557 |
fixes a :: "'a::field" |
|
1558 |
assumes "a \<noteq> 0" |
|
1559 |
shows "is_unit (monom a 0)" |
|
1560 |
proof |
|
| 62351 | 1561 |
from assms show "1 = monom a 0 * monom (inverse a) 0" |
| 60570 | 1562 |
by (simp add: mult_monom) |
1563 |
qed |
|
1564 |
||
1565 |
lemma is_unit_triv: |
|
1566 |
fixes a :: "'a::field" |
|
1567 |
assumes "a \<noteq> 0" |
|
1568 |
shows "is_unit [:a:]" |
|
1569 |
using assms by (simp add: is_unit_monom_0 monom_0 [symmetric]) |
|
1570 |
||
1571 |
lemma is_unit_iff_degree: |
|
1572 |
assumes "p \<noteq> 0" |
|
1573 |
shows "is_unit p \<longleftrightarrow> degree p = 0" (is "?P \<longleftrightarrow> ?Q") |
|
1574 |
proof |
|
1575 |
assume ?Q |
|
1576 |
then obtain a where "p = [:a:]" by (rule degree_eq_zeroE) |
|
1577 |
with assms show ?P by (simp add: is_unit_triv) |
|
1578 |
next |
|
1579 |
assume ?P |
|
1580 |
then obtain q where "q \<noteq> 0" "p * q = 1" .. |
|
1581 |
then have "degree (p * q) = degree 1" |
|
1582 |
by simp |
|
1583 |
with \<open>p \<noteq> 0\<close> \<open>q \<noteq> 0\<close> have "degree p + degree q = 0" |
|
1584 |
by (simp add: degree_mult_eq) |
|
1585 |
then show ?Q by simp |
|
1586 |
qed |
|
1587 |
||
1588 |
lemma is_unit_pCons_iff: |
|
1589 |
"is_unit (pCons a p) \<longleftrightarrow> p = 0 \<and> a \<noteq> 0" (is "?P \<longleftrightarrow> ?Q") |
|
1590 |
by (cases "p = 0") (auto simp add: is_unit_triv is_unit_iff_degree) |
|
1591 |
||
1592 |
lemma is_unit_monom_trival: |
|
1593 |
fixes p :: "'a::field poly" |
|
1594 |
assumes "is_unit p" |
|
1595 |
shows "monom (coeff p (degree p)) 0 = p" |
|
1596 |
using assms by (cases p) (simp_all add: monom_0 is_unit_pCons_iff) |
|
1597 |
||
|
60685
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1598 |
lemma is_unit_polyE: |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1599 |
assumes "is_unit p" |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1600 |
obtains a where "p = monom a 0" and "a \<noteq> 0" |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1601 |
proof - |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1602 |
obtain a q where "p = pCons a q" by (cases p) |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1603 |
with assms have "p = [:a:]" and "a \<noteq> 0" |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1604 |
by (simp_all add: is_unit_pCons_iff) |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1605 |
with that show thesis by (simp add: monom_0) |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1606 |
qed |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1607 |
|
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1608 |
instantiation poly :: (field) normalization_semidom |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1609 |
begin |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1610 |
|
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1611 |
definition normalize_poly :: "'a poly \<Rightarrow> 'a poly" |
| 62351 | 1612 |
where "normalize_poly p = smult (inverse (coeff p (degree p))) p" |
|
60685
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1613 |
|
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1614 |
definition unit_factor_poly :: "'a poly \<Rightarrow> 'a poly" |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1615 |
where "unit_factor_poly p = monom (coeff p (degree p)) 0" |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1616 |
|
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1617 |
instance |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1618 |
proof |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1619 |
fix p :: "'a poly" |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1620 |
show "unit_factor p * normalize p = p" |
| 62351 | 1621 |
by (cases "p = 0") |
1622 |
(simp_all add: normalize_poly_def unit_factor_poly_def, |
|
1623 |
simp only: mult_smult_left [symmetric] smult_monom, simp) |
|
|
60685
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1624 |
next |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1625 |
show "normalize 0 = (0::'a poly)" |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1626 |
by (simp add: normalize_poly_def) |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1627 |
next |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1628 |
show "unit_factor 0 = (0::'a poly)" |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1629 |
by (simp add: unit_factor_poly_def) |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1630 |
next |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1631 |
fix p :: "'a poly" |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1632 |
assume "is_unit p" |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1633 |
then obtain a where "p = monom a 0" and "a \<noteq> 0" |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1634 |
by (rule is_unit_polyE) |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1635 |
then show "normalize p = 1" |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1636 |
by (auto simp add: normalize_poly_def smult_monom degree_monom_eq) |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1637 |
next |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1638 |
fix p q :: "'a poly" |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1639 |
assume "q \<noteq> 0" |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1640 |
from \<open>q \<noteq> 0\<close> have "is_unit (monom (coeff q (degree q)) 0)" |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1641 |
by (auto intro: is_unit_monom_0) |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1642 |
then show "is_unit (unit_factor q)" |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1643 |
by (simp add: unit_factor_poly_def) |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1644 |
next |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1645 |
fix p q :: "'a poly" |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1646 |
have "monom (coeff (p * q) (degree (p * q))) 0 = |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1647 |
monom (coeff p (degree p)) 0 * monom (coeff q (degree q)) 0" |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1648 |
by (simp add: monom_0 coeff_degree_mult) |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1649 |
then show "unit_factor (p * q) = |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1650 |
unit_factor p * unit_factor q" |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1651 |
by (simp add: unit_factor_poly_def) |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1652 |
qed |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1653 |
|
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1654 |
end |
|
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
haftmann
parents:
60679
diff
changeset
|
1655 |
|
| 62351 | 1656 |
lemma unit_factor_monom [simp]: |
1657 |
"unit_factor (monom a n) = |
|
1658 |
(if a = 0 then 0 else monom a 0)" |
|
1659 |
by (simp add: unit_factor_poly_def degree_monom_eq) |
|
1660 |
||
1661 |
lemma unit_factor_pCons [simp]: |
|
1662 |
"unit_factor (pCons a p) = |
|
1663 |
(if p = 0 then monom a 0 else unit_factor p)" |
|
1664 |
by (simp add: unit_factor_poly_def) |
|
1665 |
||
1666 |
lemma normalize_monom [simp]: |
|
1667 |
"normalize (monom a n) = |
|
1668 |
(if a = 0 then 0 else monom 1 n)" |
|
1669 |
by (simp add: normalize_poly_def degree_monom_eq smult_monom) |
|
1670 |
||
| 29451 | 1671 |
lemma degree_mod_less: |
1672 |
"y \<noteq> 0 \<Longrightarrow> x mod y = 0 \<or> degree (x mod y) < degree y" |
|
| 29537 | 1673 |
using pdivmod_rel [of x y] |
1674 |
unfolding pdivmod_rel_def by simp |
|
| 29451 | 1675 |
|
1676 |
lemma div_poly_less: "degree x < degree y \<Longrightarrow> x div y = 0" |
|
1677 |
proof - |
|
1678 |
assume "degree x < degree y" |
|
| 29537 | 1679 |
hence "pdivmod_rel x y 0 x" |
1680 |
by (simp add: pdivmod_rel_def) |
|
| 29451 | 1681 |
thus "x div y = 0" by (rule div_poly_eq) |
1682 |
qed |
|
1683 |
||
1684 |
lemma mod_poly_less: "degree x < degree y \<Longrightarrow> x mod y = x" |
|
1685 |
proof - |
|
1686 |
assume "degree x < degree y" |
|
| 29537 | 1687 |
hence "pdivmod_rel x y 0 x" |
1688 |
by (simp add: pdivmod_rel_def) |
|
| 29451 | 1689 |
thus "x mod y = x" by (rule mod_poly_eq) |
1690 |
qed |
|
1691 |
||
| 29659 | 1692 |
lemma pdivmod_rel_smult_left: |
1693 |
"pdivmod_rel x y q r |
|
1694 |
\<Longrightarrow> pdivmod_rel (smult a x) y (smult a q) (smult a r)" |
|
1695 |
unfolding pdivmod_rel_def by (simp add: smult_add_right) |
|
1696 |
||
1697 |
lemma div_smult_left: "(smult a x) div y = smult a (x div y)" |
|
1698 |
by (rule div_poly_eq, rule pdivmod_rel_smult_left, rule pdivmod_rel) |
|
1699 |
||
1700 |
lemma mod_smult_left: "(smult a x) mod y = smult a (x mod y)" |
|
1701 |
by (rule mod_poly_eq, rule pdivmod_rel_smult_left, rule pdivmod_rel) |
|
1702 |
||
| 30072 | 1703 |
lemma poly_div_minus_left [simp]: |
1704 |
fixes x y :: "'a::field poly" |
|
1705 |
shows "(- x) div y = - (x div y)" |
|
|
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54230
diff
changeset
|
1706 |
using div_smult_left [of "- 1::'a"] by simp |
| 30072 | 1707 |
|
1708 |
lemma poly_mod_minus_left [simp]: |
|
1709 |
fixes x y :: "'a::field poly" |
|
1710 |
shows "(- x) mod y = - (x mod y)" |
|
|
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54230
diff
changeset
|
1711 |
using mod_smult_left [of "- 1::'a"] by simp |
| 30072 | 1712 |
|
|
57482
60459c3853af
add lemmas: polynomial div/mod distribute over addition
huffman
parents:
56544
diff
changeset
|
1713 |
lemma pdivmod_rel_add_left: |
|
60459c3853af
add lemmas: polynomial div/mod distribute over addition
huffman
parents:
56544
diff
changeset
|
1714 |
assumes "pdivmod_rel x y q r" |
|
60459c3853af
add lemmas: polynomial div/mod distribute over addition
huffman
parents:
56544
diff
changeset
|
1715 |
assumes "pdivmod_rel x' y q' r'" |
|
60459c3853af
add lemmas: polynomial div/mod distribute over addition
huffman
parents:
56544
diff
changeset
|
1716 |
shows "pdivmod_rel (x + x') y (q + q') (r + r')" |
|
60459c3853af
add lemmas: polynomial div/mod distribute over addition
huffman
parents:
56544
diff
changeset
|
1717 |
using assms unfolding pdivmod_rel_def |
| 59557 | 1718 |
by (auto simp add: algebra_simps degree_add_less) |
|
57482
60459c3853af
add lemmas: polynomial div/mod distribute over addition
huffman
parents:
56544
diff
changeset
|
1719 |
|
|
60459c3853af
add lemmas: polynomial div/mod distribute over addition
huffman
parents:
56544
diff
changeset
|
1720 |
lemma poly_div_add_left: |
|
60459c3853af
add lemmas: polynomial div/mod distribute over addition
huffman
parents:
56544
diff
changeset
|
1721 |
fixes x y z :: "'a::field poly" |
|
60459c3853af
add lemmas: polynomial div/mod distribute over addition
huffman
parents:
56544
diff
changeset
|
1722 |
shows "(x + y) div z = x div z + y div z" |
|
60459c3853af
add lemmas: polynomial div/mod distribute over addition
huffman
parents:
56544
diff
changeset
|
1723 |
using pdivmod_rel_add_left [OF pdivmod_rel pdivmod_rel] |
|
60459c3853af
add lemmas: polynomial div/mod distribute over addition
huffman
parents:
56544
diff
changeset
|
1724 |
by (rule div_poly_eq) |
|
60459c3853af
add lemmas: polynomial div/mod distribute over addition
huffman
parents:
56544
diff
changeset
|
1725 |
|
|
60459c3853af
add lemmas: polynomial div/mod distribute over addition
huffman
parents:
56544
diff
changeset
|
1726 |
lemma poly_mod_add_left: |
|
60459c3853af
add lemmas: polynomial div/mod distribute over addition
huffman
parents:
56544
diff
changeset
|
1727 |
fixes x y z :: "'a::field poly" |
|
60459c3853af
add lemmas: polynomial div/mod distribute over addition
huffman
parents:
56544
diff
changeset
|
1728 |
shows "(x + y) mod z = x mod z + y mod z" |
|
60459c3853af
add lemmas: polynomial div/mod distribute over addition
huffman
parents:
56544
diff
changeset
|
1729 |
using pdivmod_rel_add_left [OF pdivmod_rel pdivmod_rel] |
|
60459c3853af
add lemmas: polynomial div/mod distribute over addition
huffman
parents:
56544
diff
changeset
|
1730 |
by (rule mod_poly_eq) |
|
60459c3853af
add lemmas: polynomial div/mod distribute over addition
huffman
parents:
56544
diff
changeset
|
1731 |
|
|
60459c3853af
add lemmas: polynomial div/mod distribute over addition
huffman
parents:
56544
diff
changeset
|
1732 |
lemma poly_div_diff_left: |
|
60459c3853af
add lemmas: polynomial div/mod distribute over addition
huffman
parents:
56544
diff
changeset
|
1733 |
fixes x y z :: "'a::field poly" |
|
60459c3853af
add lemmas: polynomial div/mod distribute over addition
huffman
parents:
56544
diff
changeset
|
1734 |
shows "(x - y) div z = x div z - y div z" |
|
60459c3853af
add lemmas: polynomial div/mod distribute over addition
huffman
parents:
56544
diff
changeset
|
1735 |
by (simp only: diff_conv_add_uminus poly_div_add_left poly_div_minus_left) |
|
60459c3853af
add lemmas: polynomial div/mod distribute over addition
huffman
parents:
56544
diff
changeset
|
1736 |
|
|
60459c3853af
add lemmas: polynomial div/mod distribute over addition
huffman
parents:
56544
diff
changeset
|
1737 |
lemma poly_mod_diff_left: |
|
60459c3853af
add lemmas: polynomial div/mod distribute over addition
huffman
parents:
56544
diff
changeset
|
1738 |
fixes x y z :: "'a::field poly" |
|
60459c3853af
add lemmas: polynomial div/mod distribute over addition
huffman
parents:
56544
diff
changeset
|
1739 |
shows "(x - y) mod z = x mod z - y mod z" |
|
60459c3853af
add lemmas: polynomial div/mod distribute over addition
huffman
parents:
56544
diff
changeset
|
1740 |
by (simp only: diff_conv_add_uminus poly_mod_add_left poly_mod_minus_left) |
|
60459c3853af
add lemmas: polynomial div/mod distribute over addition
huffman
parents:
56544
diff
changeset
|
1741 |
|
| 29659 | 1742 |
lemma pdivmod_rel_smult_right: |
1743 |
"\<lbrakk>a \<noteq> 0; pdivmod_rel x y q r\<rbrakk> |
|
1744 |
\<Longrightarrow> pdivmod_rel x (smult a y) (smult (inverse a) q) r" |
|
1745 |
unfolding pdivmod_rel_def by simp |
|
1746 |
||
1747 |
lemma div_smult_right: |
|
1748 |
"a \<noteq> 0 \<Longrightarrow> x div (smult a y) = smult (inverse a) (x div y)" |
|
1749 |
by (rule div_poly_eq, erule pdivmod_rel_smult_right, rule pdivmod_rel) |
|
1750 |
||
1751 |
lemma mod_smult_right: "a \<noteq> 0 \<Longrightarrow> x mod (smult a y) = x mod y" |
|
1752 |
by (rule mod_poly_eq, erule pdivmod_rel_smult_right, rule pdivmod_rel) |
|
1753 |
||
| 30072 | 1754 |
lemma poly_div_minus_right [simp]: |
1755 |
fixes x y :: "'a::field poly" |
|
1756 |
shows "x div (- y) = - (x div y)" |
|
|
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54230
diff
changeset
|
1757 |
using div_smult_right [of "- 1::'a"] by (simp add: nonzero_inverse_minus_eq) |
| 30072 | 1758 |
|
1759 |
lemma poly_mod_minus_right [simp]: |
|
1760 |
fixes x y :: "'a::field poly" |
|
1761 |
shows "x mod (- y) = x mod y" |
|
|
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54230
diff
changeset
|
1762 |
using mod_smult_right [of "- 1::'a"] by simp |
| 30072 | 1763 |
|
| 29660 | 1764 |
lemma pdivmod_rel_mult: |
1765 |
"\<lbrakk>pdivmod_rel x y q r; pdivmod_rel q z q' r'\<rbrakk> |
|
1766 |
\<Longrightarrow> pdivmod_rel x (y * z) q' (y * r' + r)" |
|
1767 |
apply (cases "z = 0", simp add: pdivmod_rel_def) |
|
1768 |
apply (cases "y = 0", simp add: pdivmod_rel_by_0_iff pdivmod_rel_0_iff) |
|
1769 |
apply (cases "r = 0") |
|
1770 |
apply (cases "r' = 0") |
|
1771 |
apply (simp add: pdivmod_rel_def) |
|
| 36350 | 1772 |
apply (simp add: pdivmod_rel_def field_simps degree_mult_eq) |
| 29660 | 1773 |
apply (cases "r' = 0") |
1774 |
apply (simp add: pdivmod_rel_def degree_mult_eq) |
|
| 36350 | 1775 |
apply (simp add: pdivmod_rel_def field_simps) |
| 29660 | 1776 |
apply (simp add: degree_mult_eq degree_add_less) |
1777 |
done |
|
1778 |
||
1779 |
lemma poly_div_mult_right: |
|
1780 |
fixes x y z :: "'a::field poly" |
|
1781 |
shows "x div (y * z) = (x div y) div z" |
|
1782 |
by (rule div_poly_eq, rule pdivmod_rel_mult, (rule pdivmod_rel)+) |
|
1783 |
||
1784 |
lemma poly_mod_mult_right: |
|
1785 |
fixes x y z :: "'a::field poly" |
|
1786 |
shows "x mod (y * z) = y * (x div y mod z) + x mod y" |
|
1787 |
by (rule mod_poly_eq, rule pdivmod_rel_mult, (rule pdivmod_rel)+) |
|
1788 |
||
| 29451 | 1789 |
lemma mod_pCons: |
1790 |
fixes a and x |
|
1791 |
assumes y: "y \<noteq> 0" |
|
1792 |
defines b: "b \<equiv> coeff (pCons a (x mod y)) (degree y) / coeff y (degree y)" |
|
1793 |
shows "(pCons a x) mod y = (pCons a (x mod y) - smult b y)" |
|
1794 |
unfolding b |
|
1795 |
apply (rule mod_poly_eq) |
|
| 29537 | 1796 |
apply (rule pdivmod_rel_pCons [OF pdivmod_rel y refl]) |
| 29451 | 1797 |
done |
1798 |
||
| 52380 | 1799 |
definition pdivmod :: "'a::field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<times> 'a poly" |
1800 |
where |
|
1801 |
"pdivmod p q = (p div q, p mod q)" |
|
| 31663 | 1802 |
|
| 52380 | 1803 |
lemma div_poly_code [code]: |
1804 |
"p div q = fst (pdivmod p q)" |
|
1805 |
by (simp add: pdivmod_def) |
|
| 31663 | 1806 |
|
| 52380 | 1807 |
lemma mod_poly_code [code]: |
1808 |
"p mod q = snd (pdivmod p q)" |
|
1809 |
by (simp add: pdivmod_def) |
|
| 31663 | 1810 |
|
| 52380 | 1811 |
lemma pdivmod_0: |
1812 |
"pdivmod 0 q = (0, 0)" |
|
1813 |
by (simp add: pdivmod_def) |
|
| 31663 | 1814 |
|
| 52380 | 1815 |
lemma pdivmod_pCons: |
1816 |
"pdivmod (pCons a p) q = |
|
1817 |
(if q = 0 then (0, pCons a p) else |
|
1818 |
(let (s, r) = pdivmod p q; |
|
1819 |
b = coeff (pCons a r) (degree q) / coeff q (degree q) |
|
1820 |
in (pCons b s, pCons a r - smult b q)))" |
|
1821 |
apply (simp add: pdivmod_def Let_def, safe) |
|
1822 |
apply (rule div_poly_eq) |
|
1823 |
apply (erule pdivmod_rel_pCons [OF pdivmod_rel _ refl]) |
|
1824 |
apply (rule mod_poly_eq) |
|
1825 |
apply (erule pdivmod_rel_pCons [OF pdivmod_rel _ refl]) |
|
| 29451 | 1826 |
done |
1827 |
||
| 52380 | 1828 |
lemma pdivmod_fold_coeffs [code]: |
1829 |
"pdivmod p q = (if q = 0 then (0, p) |
|
1830 |
else fold_coeffs (\<lambda>a (s, r). |
|
1831 |
let b = coeff (pCons a r) (degree q) / coeff q (degree q) |
|
1832 |
in (pCons b s, pCons a r - smult b q) |
|
1833 |
) p (0, 0))" |
|
1834 |
apply (cases "q = 0") |
|
1835 |
apply (simp add: pdivmod_def) |
|
1836 |
apply (rule sym) |
|
1837 |
apply (induct p) |
|
1838 |
apply (simp_all add: pdivmod_0 pdivmod_pCons) |
|
1839 |
apply (case_tac "a = 0 \<and> p = 0") |
|
1840 |
apply (auto simp add: pdivmod_def) |
|
1841 |
done |
|
| 29980 | 1842 |
|
1843 |
||
| 60500 | 1844 |
subsection \<open>Order of polynomial roots\<close> |
|
29977
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1845 |
|
| 52380 | 1846 |
definition order :: "'a::idom \<Rightarrow> 'a poly \<Rightarrow> nat" |
|
29977
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1847 |
where |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1848 |
"order a p = (LEAST n. \<not> [:-a, 1:] ^ Suc n dvd p)" |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1849 |
|
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1850 |
lemma coeff_linear_power: |
| 29979 | 1851 |
fixes a :: "'a::comm_semiring_1" |
|
29977
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1852 |
shows "coeff ([:a, 1:] ^ n) n = 1" |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1853 |
apply (induct n, simp_all) |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1854 |
apply (subst coeff_eq_0) |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1855 |
apply (auto intro: le_less_trans degree_power_le) |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1856 |
done |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1857 |
|
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1858 |
lemma degree_linear_power: |
| 29979 | 1859 |
fixes a :: "'a::comm_semiring_1" |
|
29977
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1860 |
shows "degree ([:a, 1:] ^ n) = n" |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1861 |
apply (rule order_antisym) |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1862 |
apply (rule ord_le_eq_trans [OF degree_power_le], simp) |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1863 |
apply (rule le_degree, simp add: coeff_linear_power) |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1864 |
done |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1865 |
|
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1866 |
lemma order_1: "[:-a, 1:] ^ order a p dvd p" |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1867 |
apply (cases "p = 0", simp) |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1868 |
apply (cases "order a p", simp) |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1869 |
apply (subgoal_tac "nat < (LEAST n. \<not> [:-a, 1:] ^ Suc n dvd p)") |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1870 |
apply (drule not_less_Least, simp) |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1871 |
apply (fold order_def, simp) |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1872 |
done |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1873 |
|
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1874 |
lemma order_2: "p \<noteq> 0 \<Longrightarrow> \<not> [:-a, 1:] ^ Suc (order a p) dvd p" |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1875 |
unfolding order_def |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1876 |
apply (rule LeastI_ex) |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1877 |
apply (rule_tac x="degree p" in exI) |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1878 |
apply (rule notI) |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1879 |
apply (drule (1) dvd_imp_degree_le) |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1880 |
apply (simp only: degree_linear_power) |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1881 |
done |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1882 |
|
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1883 |
lemma order: |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1884 |
"p \<noteq> 0 \<Longrightarrow> [:-a, 1:] ^ order a p dvd p \<and> \<not> [:-a, 1:] ^ Suc (order a p) dvd p" |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1885 |
by (rule conjI [OF order_1 order_2]) |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1886 |
|
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1887 |
lemma order_degree: |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1888 |
assumes p: "p \<noteq> 0" |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1889 |
shows "order a p \<le> degree p" |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1890 |
proof - |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1891 |
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
|
1892 |
by (simp only: degree_linear_power) |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1893 |
also have "\<dots> \<le> degree p" |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1894 |
using order_1 p by (rule dvd_imp_degree_le) |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1895 |
finally show ?thesis . |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1896 |
qed |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1897 |
|
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1898 |
lemma order_root: "poly p a = 0 \<longleftrightarrow> p = 0 \<or> order a p \<noteq> 0" |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1899 |
apply (cases "p = 0", simp_all) |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1900 |
apply (rule iffI) |
| 56383 | 1901 |
apply (metis order_2 not_gr0 poly_eq_0_iff_dvd power_0 power_Suc_0 power_one_right) |
1902 |
unfolding poly_eq_0_iff_dvd |
|
1903 |
apply (metis dvd_power dvd_trans order_1) |
|
|
29977
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1904 |
done |
|
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1905 |
|
| 62065 | 1906 |
lemma order_0I: "poly p a \<noteq> 0 \<Longrightarrow> order a p = 0" |
1907 |
by (subst (asm) order_root) auto |
|
1908 |
||
|
29977
d76b830366bc
move polynomial order stuff from Fundamental_Theorem_Algebra to Polynomial
huffman
parents:
29904
diff
changeset
|
1909 |
|
| 62065 | 1910 |
subsection \<open>Additional induction rules on polynomials\<close> |
1911 |
||
1912 |
text \<open> |
|
1913 |
An induction rule for induction over the roots of a polynomial with a certain property. |
|
1914 |
(e.g. all positive roots) |
|
1915 |
\<close> |
|
1916 |
lemma poly_root_induct [case_names 0 no_roots root]: |
|
1917 |
fixes p :: "'a :: idom poly" |
|
1918 |
assumes "Q 0" |
|
1919 |
assumes "\<And>p. (\<And>a. P a \<Longrightarrow> poly p a \<noteq> 0) \<Longrightarrow> Q p" |
|
1920 |
assumes "\<And>a p. P a \<Longrightarrow> Q p \<Longrightarrow> Q ([:a, -1:] * p)" |
|
1921 |
shows "Q p" |
|
1922 |
proof (induction "degree p" arbitrary: p rule: less_induct) |
|
1923 |
case (less p) |
|
1924 |
show ?case |
|
1925 |
proof (cases "p = 0") |
|
1926 |
assume nz: "p \<noteq> 0" |
|
1927 |
show ?case |
|
1928 |
proof (cases "\<exists>a. P a \<and> poly p a = 0") |
|
1929 |
case False |
|
1930 |
thus ?thesis by (intro assms(2)) blast |
|
1931 |
next |
|
1932 |
case True |
|
1933 |
then obtain a where a: "P a" "poly p a = 0" |
|
1934 |
by blast |
|
1935 |
hence "-[:-a, 1:] dvd p" |
|
1936 |
by (subst minus_dvd_iff) (simp add: poly_eq_0_iff_dvd) |
|
1937 |
then obtain q where q: "p = [:a, -1:] * q" by (elim dvdE) simp |
|
1938 |
with nz have q_nz: "q \<noteq> 0" by auto |
|
1939 |
have "degree p = Suc (degree q)" |
|
1940 |
by (subst q, subst degree_mult_eq) (simp_all add: q_nz) |
|
1941 |
hence "Q q" by (intro less) simp |
|
1942 |
from a(1) and this have "Q ([:a, -1:] * q)" |
|
1943 |
by (rule assms(3)) |
|
1944 |
with q show ?thesis by simp |
|
1945 |
qed |
|
1946 |
qed (simp add: assms(1)) |
|
1947 |
qed |
|
1948 |
||
1949 |
lemma dropWhile_replicate_append: |
|
1950 |
"dropWhile (op= a) (replicate n a @ ys) = dropWhile (op= a) ys" |
|
1951 |
by (induction n) simp_all |
|
1952 |
||
1953 |
lemma Poly_append_replicate_0: "Poly (xs @ replicate n 0) = Poly xs" |
|
1954 |
by (subst coeffs_eq_iff) (simp_all add: strip_while_def dropWhile_replicate_append) |
|
1955 |
||
1956 |
text \<open> |
|
1957 |
An induction rule for simultaneous induction over two polynomials, |
|
1958 |
prepending one coefficient in each step. |
|
1959 |
\<close> |
|
1960 |
lemma poly_induct2 [case_names 0 pCons]: |
|
1961 |
assumes "P 0 0" "\<And>a p b q. P p q \<Longrightarrow> P (pCons a p) (pCons b q)" |
|
1962 |
shows "P p q" |
|
1963 |
proof - |
|
1964 |
def n \<equiv> "max (length (coeffs p)) (length (coeffs q))" |
|
1965 |
def xs \<equiv> "coeffs p @ (replicate (n - length (coeffs p)) 0)" |
|
1966 |
def ys \<equiv> "coeffs q @ (replicate (n - length (coeffs q)) 0)" |
|
1967 |
have "length xs = length ys" |
|
1968 |
by (simp add: xs_def ys_def n_def) |
|
1969 |
hence "P (Poly xs) (Poly ys)" |
|
1970 |
by (induction rule: list_induct2) (simp_all add: assms) |
|
1971 |
also have "Poly xs = p" |
|
1972 |
by (simp add: xs_def Poly_append_replicate_0) |
|
1973 |
also have "Poly ys = q" |
|
1974 |
by (simp add: ys_def Poly_append_replicate_0) |
|
1975 |
finally show ?thesis . |
|
1976 |
qed |
|
1977 |
||
1978 |
||
| 60500 | 1979 |
subsection \<open>Composition of polynomials\<close> |
| 29478 | 1980 |
|
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1981 |
(* 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
|
1982 |
|
| 52380 | 1983 |
definition pcompose :: "'a::comm_semiring_0 poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" |
1984 |
where |
|
1985 |
"pcompose p q = fold_coeffs (\<lambda>a c. [:a:] + q * c) p 0" |
|
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 |
|
| 52380 | 1989 |
lemma pcompose_0 [simp]: |
1990 |
"pcompose 0 q = 0" |
|
1991 |
by (simp add: pcompose_def) |
|
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1992 |
|
| 52380 | 1993 |
lemma pcompose_pCons: |
1994 |
"pcompose (pCons a p) q = [:a:] + q * pcompose p q" |
|
1995 |
by (cases "p = 0 \<and> a = 0") (auto simp add: pcompose_def) |
|
1996 |
||
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1997 |
lemma pcompose_1: |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1998 |
fixes p :: "'a :: comm_semiring_1 poly" |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
1999 |
shows "pcompose 1 p = 1" |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
2000 |
unfolding one_poly_def by (auto simp: pcompose_pCons) |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
2001 |
|
| 52380 | 2002 |
lemma poly_pcompose: |
2003 |
"poly (pcompose p q) x = poly p (poly q x)" |
|
2004 |
by (induct p) (simp_all add: pcompose_pCons) |
|
2005 |
||
2006 |
lemma degree_pcompose_le: |
|
2007 |
"degree (pcompose p q) \<le> degree p * degree q" |
|
2008 |
apply (induct p, simp) |
|
2009 |
apply (simp add: pcompose_pCons, clarify) |
|
2010 |
apply (rule degree_add_le, simp) |
|
2011 |
apply (rule order_trans [OF degree_mult_le], simp) |
|
| 29478 | 2012 |
done |
2013 |
||
| 62065 | 2014 |
lemma pcompose_add: |
2015 |
fixes p q r :: "'a :: {comm_semiring_0, ab_semigroup_add} poly"
|
|
2016 |
shows "pcompose (p + q) r = pcompose p r + pcompose q r" |
|
2017 |
proof (induction p q rule: poly_induct2) |
|
2018 |
case (pCons a p b q) |
|
2019 |
have "pcompose (pCons a p + pCons b q) r = |
|
2020 |
[:a + b:] + r * pcompose p r + r * pcompose q r" |
|
2021 |
by (simp_all add: pcompose_pCons pCons.IH algebra_simps) |
|
2022 |
also have "[:a + b:] = [:a:] + [:b:]" by simp |
|
2023 |
also have "\<dots> + r * pcompose p r + r * pcompose q r = |
|
2024 |
pcompose (pCons a p) r + pcompose (pCons b q) r" |
|
2025 |
by (simp only: pcompose_pCons add_ac) |
|
2026 |
finally show ?case . |
|
2027 |
qed simp |
|
2028 |
||
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
2029 |
lemma pcompose_uminus: |
| 62065 | 2030 |
fixes p r :: "'a :: comm_ring poly" |
2031 |
shows "pcompose (-p) r = -pcompose p r" |
|
2032 |
by (induction p) (simp_all add: pcompose_pCons) |
|
2033 |
||
2034 |
lemma pcompose_diff: |
|
2035 |
fixes p q r :: "'a :: comm_ring poly" |
|
2036 |
shows "pcompose (p - q) r = pcompose p r - pcompose q r" |
|
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
2037 |
using pcompose_add[of p "-q"] by (simp add: pcompose_uminus) |
| 62065 | 2038 |
|
2039 |
lemma pcompose_smult: |
|
2040 |
fixes p r :: "'a :: comm_semiring_0 poly" |
|
2041 |
shows "pcompose (smult a p) r = smult a (pcompose p r)" |
|
2042 |
by (induction p) |
|
2043 |
(simp_all add: pcompose_pCons pcompose_add smult_add_right) |
|
2044 |
||
2045 |
lemma pcompose_mult: |
|
2046 |
fixes p q r :: "'a :: comm_semiring_0 poly" |
|
2047 |
shows "pcompose (p * q) r = pcompose p r * pcompose q r" |
|
2048 |
by (induction p arbitrary: q) |
|
2049 |
(simp_all add: pcompose_add pcompose_smult pcompose_pCons algebra_simps) |
|
2050 |
||
2051 |
lemma pcompose_assoc: |
|
2052 |
"pcompose p (pcompose q r :: 'a :: comm_semiring_0 poly ) = |
|
2053 |
pcompose (pcompose p q) r" |
|
2054 |
by (induction p arbitrary: q) |
|
2055 |
(simp_all add: pcompose_pCons pcompose_add pcompose_mult) |
|
2056 |
||
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
2057 |
lemma pcompose_idR[simp]: |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
2058 |
fixes p :: "'a :: comm_semiring_1 poly" |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
2059 |
shows "pcompose p [: 0, 1 :] = p" |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
2060 |
by (induct p; simp add: pcompose_pCons) |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
2061 |
|
| 62065 | 2062 |
|
2063 |
(* The remainder of this section and the next were contributed by Wenda Li *) |
|
2064 |
||
2065 |
lemma degree_mult_eq_0: |
|
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
2066 |
fixes p q:: "'a :: semidom poly" |
| 62065 | 2067 |
shows "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)" |
2068 |
by (auto simp add:degree_mult_eq) |
|
2069 |
||
2070 |
lemma pcompose_const[simp]:"pcompose [:a:] q = [:a:]" by (subst pcompose_pCons,simp) |
|
2071 |
||
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
2072 |
lemma pcompose_0': "pcompose 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
|
2073 |
by (induct p) (auto simp add:pcompose_pCons) |
| 62065 | 2074 |
|
2075 |
lemma degree_pcompose: |
|
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
2076 |
fixes p q:: "'a::semidom poly" |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
2077 |
shows "degree (pcompose p q) = degree p * degree q" |
| 62065 | 2078 |
proof (induct p) |
2079 |
case 0 |
|
2080 |
thus ?case by auto |
|
2081 |
next |
|
2082 |
case (pCons a p) |
|
2083 |
have "degree (q * pcompose p q) = 0 \<Longrightarrow> ?case" |
|
2084 |
proof (cases "p=0") |
|
2085 |
case True |
|
2086 |
thus ?thesis by auto |
|
2087 |
next |
|
2088 |
case False assume "degree (q * pcompose p q) = 0" |
|
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
2089 |
hence "degree q=0 \<or> pcompose p q=0" by (auto simp add: degree_mult_eq_0) |
| 62072 | 2090 |
moreover have "\<lbrakk>pcompose p q=0;degree q\<noteq>0\<rbrakk> \<Longrightarrow> False" using pCons.hyps(2) \<open>p\<noteq>0\<close> |
| 62065 | 2091 |
proof - |
2092 |
assume "pcompose p q=0" "degree q\<noteq>0" |
|
2093 |
hence "degree p=0" using pCons.hyps(2) by auto |
|
2094 |
then obtain a1 where "p=[:a1:]" |
|
2095 |
by (metis degree_pCons_eq_if old.nat.distinct(2) pCons_cases) |
|
| 62072 | 2096 |
thus False using \<open>pcompose p q=0\<close> \<open>p\<noteq>0\<close> by auto |
| 62065 | 2097 |
qed |
2098 |
ultimately have "degree (pCons a p) * degree q=0" by auto |
|
2099 |
moreover have "degree (pcompose (pCons a p) q) = 0" |
|
2100 |
proof - |
|
2101 |
have" 0 = max (degree [:a:]) (degree (q*pcompose p q))" |
|
| 62072 | 2102 |
using \<open>degree (q * pcompose p q) = 0\<close> by simp |
| 62065 | 2103 |
also have "... \<ge> degree ([:a:] + q * pcompose p q)" |
2104 |
by (rule degree_add_le_max) |
|
2105 |
finally show ?thesis by (auto simp add:pcompose_pCons) |
|
2106 |
qed |
|
2107 |
ultimately show ?thesis by simp |
|
2108 |
qed |
|
2109 |
moreover have "degree (q * pcompose p q)>0 \<Longrightarrow> ?case" |
|
2110 |
proof - |
|
2111 |
assume asm:"0 < degree (q * pcompose p q)" |
|
2112 |
hence "p\<noteq>0" "q\<noteq>0" "pcompose p q\<noteq>0" by auto |
|
2113 |
have "degree (pcompose (pCons a p) q) = degree ( q * pcompose p q)" |
|
2114 |
unfolding pcompose_pCons |
|
2115 |
using degree_add_eq_right[of "[:a:]" ] asm by auto |
|
2116 |
thus ?thesis |
|
| 62072 | 2117 |
using pCons.hyps(2) degree_mult_eq[OF \<open>q\<noteq>0\<close> \<open>pcompose p q\<noteq>0\<close>] by auto |
| 62065 | 2118 |
qed |
2119 |
ultimately show ?case by blast |
|
2120 |
qed |
|
2121 |
||
2122 |
lemma pcompose_eq_0: |
|
|
62128
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
2123 |
fixes p q:: "'a :: semidom poly" |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
2124 |
assumes "pcompose p q = 0" "degree q > 0" |
|
3201ddb00097
Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents:
62072
diff
changeset
|
2125 |
shows "p = 0" |
| 62065 | 2126 |
proof - |
2127 |
have "degree p=0" using assms degree_pcompose[of p q] by auto |
|
2128 |
then obtain a where "p=[:a:]" |
|
2129 |
by (metis degree_pCons_eq_if gr0_conv_Suc neq0_conv pCons_cases) |
|
2130 |
hence "a=0" using assms(1) by auto |
|
| 62072 | 2131 |
thus ?thesis using \<open>p=[:a:]\<close> by simp |
| 62065 | 2132 |
qed |
2133 |
||
2134 |
||
| 62072 | 2135 |
subsection \<open>Leading coefficient\<close> |
| 62065 | 2136 |
|
2137 |
definition lead_coeff:: "'a::zero poly \<Rightarrow> 'a" where |
|
2138 |
"lead_coeff p= coeff p (degree p)" |
|
2139 |
||
2140 |
lemma lead_coeff_pCons[simp]: |
|
2141 |
"p\<noteq>0 \<Longrightarrow>lead_coeff (pCons a p) = lead_coeff p" |
|
2142 |
"p=0 \<Longrightarrow> lead_coeff (pCons a p) = a" |
|
2143 |
unfolding lead_coeff_def by auto |
|
2144 |
||
2145 |
lemma lead_coeff_0[simp]:"lead_coeff 0 =0" |
|
2146 |
unfolding lead_coeff_def by auto |
|
2147 |
||
2148 |
lemma lead_coeff_mult: |
|
2149 |
fixes p q::"'a ::idom poly" |
|
2150 |
shows "lead_coeff (p * q) = lead_coeff p * lead_coeff q" |
|
2151 |
by (unfold lead_coeff_def,cases "p=0 \<or> q=0",auto simp add:coeff_mult_degree_sum degree_mult_eq) |
|
2152 |
||
2153 |
lemma lead_coeff_add_le: |
|
2154 |
assumes "degree p < degree q" |
|
2155 |
shows "lead_coeff (p+q) = lead_coeff q" |
|
2156 |
using assms unfolding lead_coeff_def |
|
2157 |
by (metis coeff_add coeff_eq_0 monoid_add_class.add.left_neutral degree_add_eq_right) |
|
2158 |
||
2159 |
lemma lead_coeff_minus: |
|
2160 |
"lead_coeff (-p) = - lead_coeff p" |
|
2161 |
by (metis coeff_minus degree_minus lead_coeff_def) |
|
2162 |
||
2163 |
||
2164 |
lemma lead_coeff_comp: |
|
2165 |
fixes p q:: "'a::idom poly" |
|
2166 |
assumes "degree q > 0" |
|
2167 |
shows "lead_coeff (pcompose p q) = lead_coeff p * lead_coeff q ^ (degree p)" |
|
2168 |
proof (induct p) |
|
2169 |
case 0 |
|
2170 |
thus ?case unfolding lead_coeff_def by auto |
|
2171 |
next |
|
2172 |
case (pCons a p) |
|
2173 |
have "degree ( q * pcompose p q) = 0 \<Longrightarrow> ?case" |
|
2174 |
proof - |
|
2175 |
assume "degree ( q * pcompose p q) = 0" |
|
2176 |
hence "pcompose p q = 0" by (metis assms degree_0 degree_mult_eq_0 neq0_conv) |
|
| 62072 | 2177 |
hence "p=0" using pcompose_eq_0[OF _ \<open>degree q > 0\<close>] by simp |
| 62065 | 2178 |
thus ?thesis by auto |
2179 |
qed |
|
2180 |
moreover have "degree ( q * pcompose p q) > 0 \<Longrightarrow> ?case" |
|
2181 |
proof - |
|
2182 |
assume "degree ( q * pcompose p q) > 0" |
|
2183 |
hence "lead_coeff (pcompose (pCons a p) q) =lead_coeff ( q * pcompose p q)" |
|
2184 |
by (auto simp add:pcompose_pCons lead_coeff_add_le) |
|
2185 |
also have "... = lead_coeff q * (lead_coeff p * lead_coeff q ^ degree p)" |
|
2186 |
using pCons.hyps(2) lead_coeff_mult[of q "pcompose p q"] by simp |
|
2187 |
also have "... = lead_coeff p * lead_coeff q ^ (degree p + 1)" |
|
2188 |
by auto |
|
2189 |
finally show ?thesis by auto |
|
2190 |
qed |
|
2191 |
ultimately show ?case by blast |
|
2192 |
qed |
|
2193 |
||
2194 |
lemma lead_coeff_smult: |
|
2195 |
"lead_coeff (smult c p :: 'a :: idom poly) = c * lead_coeff p" |
|
2196 |
proof - |
|
2197 |
have "smult c p = [:c:] * p" by simp |
|
2198 |
also have "lead_coeff \<dots> = c * lead_coeff p" |
|
2199 |
by (subst lead_coeff_mult) simp_all |
|
2200 |
finally show ?thesis . |
|
2201 |
qed |
|
2202 |
||
2203 |
lemma lead_coeff_1 [simp]: "lead_coeff 1 = 1" |
|
2204 |
by (simp add: lead_coeff_def) |
|
2205 |
||
2206 |
lemma lead_coeff_of_nat [simp]: |
|
2207 |
"lead_coeff (of_nat n) = (of_nat n :: 'a :: {comm_semiring_1,semiring_char_0})"
|
|
2208 |
by (induction n) (simp_all add: lead_coeff_def of_nat_poly) |
|
2209 |
||
2210 |
lemma lead_coeff_numeral [simp]: |
|
2211 |
"lead_coeff (numeral n) = numeral n" |
|
2212 |
unfolding lead_coeff_def |
|
2213 |
by (subst of_nat_numeral [symmetric], subst of_nat_poly) simp |
|
2214 |
||
2215 |
lemma lead_coeff_power: |
|
2216 |
"lead_coeff (p ^ n :: 'a :: idom poly) = lead_coeff p ^ n" |
|
2217 |
by (induction n) (simp_all add: lead_coeff_mult) |
|
2218 |
||
2219 |
lemma lead_coeff_nonzero: "p \<noteq> 0 \<Longrightarrow> lead_coeff p \<noteq> 0" |
|
2220 |
by (simp add: lead_coeff_def) |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2221 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2222 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2223 |
subsection \<open>Derivatives of univariate polynomials\<close> |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2224 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2225 |
function pderiv :: "('a :: semidom) poly \<Rightarrow> 'a poly"
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2226 |
where |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2227 |
[simp del]: "pderiv (pCons a p) = (if p = 0 then 0 else p + pCons 0 (pderiv p))" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2228 |
by (auto intro: pCons_cases) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2229 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2230 |
termination pderiv |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2231 |
by (relation "measure degree") simp_all |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2232 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2233 |
lemma pderiv_0 [simp]: |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2234 |
"pderiv 0 = 0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2235 |
using pderiv.simps [of 0 0] by simp |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2236 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2237 |
lemma pderiv_pCons: |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2238 |
"pderiv (pCons a p) = p + pCons 0 (pderiv p)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2239 |
by (simp add: pderiv.simps) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2240 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2241 |
lemma pderiv_1 [simp]: "pderiv 1 = 0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2242 |
unfolding one_poly_def by (simp add: pderiv_pCons) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2243 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2244 |
lemma pderiv_of_nat [simp]: "pderiv (of_nat n) = 0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2245 |
and pderiv_numeral [simp]: "pderiv (numeral m) = 0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2246 |
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
|
2247 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2248 |
lemma coeff_pderiv: "coeff (pderiv p) n = of_nat (Suc n) * coeff p (Suc n)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2249 |
by (induct p arbitrary: n) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2250 |
(auto simp add: pderiv_pCons coeff_pCons algebra_simps split: nat.split) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2251 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2252 |
fun pderiv_coeffs_code :: "('a :: semidom) \<Rightarrow> 'a list \<Rightarrow> 'a list" where
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2253 |
"pderiv_coeffs_code f (x # xs) = cCons (f * x) (pderiv_coeffs_code (f+1) xs)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2254 |
| "pderiv_coeffs_code f [] = []" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2255 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2256 |
definition pderiv_coeffs :: "('a :: semidom) list \<Rightarrow> 'a list" where
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2257 |
"pderiv_coeffs xs = pderiv_coeffs_code 1 (tl xs)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2258 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2259 |
(* Efficient code for pderiv contributed by René Thiemann and Akihisa Yamada *) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2260 |
lemma pderiv_coeffs_code: |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2261 |
"nth_default 0 (pderiv_coeffs_code f xs) n = (f + of_nat n) * (nth_default 0 xs n)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2262 |
proof (induct xs arbitrary: f n) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2263 |
case (Cons x xs f n) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2264 |
show ?case |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2265 |
proof (cases n) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2266 |
case 0 |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2267 |
thus ?thesis by (cases "pderiv_coeffs_code (f + 1) xs = [] \<and> f * x = 0", auto simp: cCons_def) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2268 |
next |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2269 |
case (Suc m) note n = this |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2270 |
show ?thesis |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2271 |
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
|
2272 |
case False |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2273 |
hence "nth_default 0 (pderiv_coeffs_code f (x # xs)) n = |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2274 |
nth_default 0 (pderiv_coeffs_code (f + 1) xs) m" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2275 |
by (auto simp: cCons_def n) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2276 |
also have "\<dots> = (f + of_nat n) * (nth_default 0 xs m)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2277 |
unfolding Cons by (simp add: n add_ac) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2278 |
finally show ?thesis by (simp add: n) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2279 |
next |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2280 |
case True |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2281 |
{
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2282 |
fix g |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2283 |
have "pderiv_coeffs_code g xs = [] \<Longrightarrow> g + of_nat m = 0 \<or> nth_default 0 xs m = 0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2284 |
proof (induct xs arbitrary: g m) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2285 |
case (Cons x xs g) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2286 |
from Cons(2) have empty: "pderiv_coeffs_code (g + 1) xs = []" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2287 |
and g: "(g = 0 \<or> x = 0)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2288 |
by (auto simp: cCons_def split: if_splits) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2289 |
note IH = Cons(1)[OF empty] |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2290 |
from IH[of m] IH[of "m - 1"] g |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2291 |
show ?case by (cases m, auto simp: field_simps) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2292 |
qed simp |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2293 |
} note empty = this |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2294 |
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
|
2295 |
by (auto simp: cCons_def n) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2296 |
moreover have "(f + of_nat n) * nth_default 0 (x # xs) n = 0" using True |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2297 |
by (simp add: n, insert empty[of "f+1"], auto simp: field_simps) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2298 |
ultimately show ?thesis by simp |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2299 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2300 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2301 |
qed simp |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2302 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2303 |
lemma map_upt_Suc: "map f [0 ..< Suc n] = f 0 # map (\<lambda> i. f (Suc i)) [0 ..< n]" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2304 |
by (induct n arbitrary: f, auto) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2305 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2306 |
lemma coeffs_pderiv_code [code abstract]: |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2307 |
"coeffs (pderiv p) = pderiv_coeffs (coeffs p)" unfolding pderiv_coeffs_def |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2308 |
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
|
2309 |
case (1 n) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2310 |
have id: "coeff p (Suc n) = nth_default 0 (map (\<lambda>i. coeff p (Suc i)) [0..<degree p]) n" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2311 |
by (cases "n < degree p", auto simp: nth_default_def coeff_eq_0) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2312 |
show ?case unfolding coeffs_def map_upt_Suc by (auto simp: id) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2313 |
next |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2314 |
case 2 |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2315 |
obtain n xs where id: "tl (coeffs p) = xs" "(1 :: 'a) = n" by auto |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2316 |
from 2 show ?case |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2317 |
unfolding id by (induct xs arbitrary: n, auto simp: cCons_def) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2318 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2319 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2320 |
context |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2321 |
assumes "SORT_CONSTRAINT('a::{semidom, semiring_char_0})"
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2322 |
begin |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2323 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2324 |
lemma pderiv_eq_0_iff: |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2325 |
"pderiv (p :: 'a poly) = 0 \<longleftrightarrow> degree p = 0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2326 |
apply (rule iffI) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2327 |
apply (cases p, simp) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2328 |
apply (simp add: poly_eq_iff coeff_pderiv del: of_nat_Suc) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2329 |
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
|
2330 |
done |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2331 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2332 |
lemma degree_pderiv: "degree (pderiv (p :: 'a poly)) = degree p - 1" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2333 |
apply (rule order_antisym [OF degree_le]) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2334 |
apply (simp add: coeff_pderiv coeff_eq_0) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2335 |
apply (cases "degree p", simp) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2336 |
apply (rule le_degree) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2337 |
apply (simp add: coeff_pderiv del: of_nat_Suc) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2338 |
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
|
2339 |
done |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2340 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2341 |
lemma not_dvd_pderiv: |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2342 |
assumes "degree (p :: 'a poly) \<noteq> 0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2343 |
shows "\<not> p dvd pderiv p" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2344 |
proof |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2345 |
assume dvd: "p dvd pderiv p" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2346 |
then obtain q where p: "pderiv p = p * q" unfolding dvd_def by auto |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2347 |
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
|
2348 |
by (simp add: assms dvd_imp_degree_le pderiv_eq_0_iff) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2349 |
from this[unfolded degree_pderiv] assms show False by auto |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2350 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2351 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2352 |
lemma dvd_pderiv_iff [simp]: "(p :: 'a poly) dvd pderiv p \<longleftrightarrow> degree p = 0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2353 |
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
|
2354 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2355 |
end |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2356 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2357 |
lemma pderiv_singleton [simp]: "pderiv [:a:] = 0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2358 |
by (simp add: pderiv_pCons) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2359 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2360 |
lemma pderiv_add: "pderiv (p + q) = pderiv p + pderiv q" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2361 |
by (rule poly_eqI, simp add: coeff_pderiv algebra_simps) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2362 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2363 |
lemma pderiv_minus: "pderiv (- p :: 'a :: idom poly) = - pderiv p" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2364 |
by (rule poly_eqI, simp add: coeff_pderiv algebra_simps) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2365 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2366 |
lemma pderiv_diff: "pderiv (p - q) = pderiv p - pderiv q" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2367 |
by (rule poly_eqI, simp add: coeff_pderiv algebra_simps) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2368 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2369 |
lemma pderiv_smult: "pderiv (smult a p) = smult a (pderiv p)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2370 |
by (rule poly_eqI, simp add: coeff_pderiv algebra_simps) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2371 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2372 |
lemma pderiv_mult: "pderiv (p * q) = p * pderiv q + q * pderiv p" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2373 |
by (induct p) (auto simp: pderiv_add pderiv_smult pderiv_pCons algebra_simps) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2374 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2375 |
lemma pderiv_power_Suc: |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2376 |
"pderiv (p ^ Suc n) = smult (of_nat (Suc n)) (p ^ n) * pderiv p" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2377 |
apply (induct n) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2378 |
apply simp |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2379 |
apply (subst power_Suc) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2380 |
apply (subst pderiv_mult) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2381 |
apply (erule ssubst) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2382 |
apply (simp only: of_nat_Suc smult_add_left smult_1_left) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2383 |
apply (simp add: algebra_simps) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2384 |
done |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2385 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2386 |
lemma pderiv_setprod: "pderiv (setprod f (as)) = |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2387 |
(\<Sum>a \<in> as. setprod f (as - {a}) * pderiv (f a))"
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2388 |
proof (induct as rule: infinite_finite_induct) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2389 |
case (insert a as) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2390 |
hence id: "setprod f (insert a as) = f a * setprod f as" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2391 |
"\<And> g. setsum g (insert a as) = g a + setsum g as" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2392 |
"insert a as - {a} = as"
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2393 |
by auto |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2394 |
{
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2395 |
fix b |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2396 |
assume "b \<in> as" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2397 |
hence id2: "insert a as - {b} = insert a (as - {b})" using \<open>a \<notin> as\<close> by auto
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2398 |
have "setprod f (insert a as - {b}) = f a * setprod f (as - {b})"
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2399 |
unfolding id2 |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2400 |
by (subst setprod.insert, insert insert, auto) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2401 |
} note id2 = this |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2402 |
show ?case |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2403 |
unfolding id pderiv_mult insert(3) setsum_right_distrib |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2404 |
by (auto simp add: ac_simps id2 intro!: setsum.cong) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2405 |
qed auto |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2406 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2407 |
lemma DERIV_pow2: "DERIV (%x. x ^ Suc n) x :> real (Suc n) * (x ^ n)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2408 |
by (rule DERIV_cong, rule DERIV_pow, simp) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2409 |
declare DERIV_pow2 [simp] DERIV_pow [simp] |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2410 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2411 |
lemma DERIV_add_const: "DERIV f x :> D ==> DERIV (%x. a + f x :: 'a::real_normed_field) x :> D" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2412 |
by (rule DERIV_cong, rule DERIV_add, auto) |
|
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 |
lemma poly_DERIV [simp]: "DERIV (%x. poly p x) x :> poly (pderiv p) x" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2415 |
by (induct p, auto intro!: derivative_eq_intros simp add: pderiv_pCons) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2416 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2417 |
lemma continuous_on_poly [continuous_intros]: |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2418 |
fixes p :: "'a :: {real_normed_field} poly"
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2419 |
assumes "continuous_on A f" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2420 |
shows "continuous_on A (\<lambda>x. poly p (f x))" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2421 |
proof - |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2422 |
have "continuous_on A (\<lambda>x. (\<Sum>i\<le>degree p. (f x) ^ i * coeff p i))" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2423 |
by (intro continuous_intros assms) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2424 |
also have "\<dots> = (\<lambda>x. poly p (f x))" by (intro ext) (simp add: poly_altdef mult_ac) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2425 |
finally show ?thesis . |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2426 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2427 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2428 |
text\<open>Consequences of the derivative theorem above\<close> |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2429 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2430 |
lemma poly_differentiable[simp]: "(%x. poly p x) differentiable (at x::real filter)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2431 |
apply (simp add: real_differentiable_def) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2432 |
apply (blast intro: poly_DERIV) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2433 |
done |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2434 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2435 |
lemma poly_isCont[simp]: "isCont (%x. poly p x) (x::real)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2436 |
by (rule poly_DERIV [THEN DERIV_isCont]) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2437 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2438 |
lemma poly_IVT_pos: "[| a < b; poly p (a::real) < 0; 0 < poly p b |] |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2439 |
==> \<exists>x. a < x & x < b & (poly p x = 0)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2440 |
using IVT_objl [of "poly p" a 0 b] |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2441 |
by (auto simp add: order_le_less) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2442 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2443 |
lemma poly_IVT_neg: "[| (a::real) < b; 0 < poly p a; poly p b < 0 |] |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2444 |
==> \<exists>x. a < x & x < b & (poly p x = 0)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2445 |
by (insert poly_IVT_pos [where p = "- p" ]) simp |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2446 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2447 |
lemma poly_IVT: |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2448 |
fixes p::"real poly" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2449 |
assumes "a<b" and "poly p a * poly p b < 0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2450 |
shows "\<exists>x>a. x < b \<and> poly p x = 0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2451 |
by (metis assms(1) assms(2) less_not_sym mult_less_0_iff poly_IVT_neg poly_IVT_pos) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2452 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2453 |
lemma poly_MVT: "(a::real) < b ==> |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2454 |
\<exists>x. a < x & x < b & (poly p b - poly p a = (b - a) * poly (pderiv p) x)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2455 |
using MVT [of a b "poly p"] |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2456 |
apply auto |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2457 |
apply (rule_tac x = z in exI) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2458 |
apply (auto simp add: mult_left_cancel poly_DERIV [THEN DERIV_unique]) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2459 |
done |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2460 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2461 |
lemma poly_MVT': |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2462 |
assumes "{min a b..max a b} \<subseteq> A"
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2463 |
shows "\<exists>x\<in>A. poly p b - poly p a = (b - a) * poly (pderiv p) (x::real)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2464 |
proof (cases a b rule: linorder_cases) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2465 |
case less |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2466 |
from poly_MVT[OF less, of p] guess x by (elim exE conjE) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2467 |
thus ?thesis by (intro bexI[of _ x]) (auto intro!: subsetD[OF assms]) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2468 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2469 |
next |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2470 |
case greater |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2471 |
from poly_MVT[OF greater, of p] guess x by (elim exE conjE) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2472 |
thus ?thesis by (intro bexI[of _ x]) (auto simp: algebra_simps intro!: subsetD[OF assms]) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2473 |
qed (insert assms, auto) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2474 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2475 |
lemma poly_pinfty_gt_lc: |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2476 |
fixes p:: "real poly" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2477 |
assumes "lead_coeff p > 0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2478 |
shows "\<exists> n. \<forall> x \<ge> n. poly p x \<ge> lead_coeff p" using assms |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2479 |
proof (induct p) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2480 |
case 0 |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2481 |
thus ?case by auto |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2482 |
next |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2483 |
case (pCons a p) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2484 |
have "\<lbrakk>a\<noteq>0;p=0\<rbrakk> \<Longrightarrow> ?case" by auto |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2485 |
moreover have "p\<noteq>0 \<Longrightarrow> ?case" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2486 |
proof - |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2487 |
assume "p\<noteq>0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2488 |
then obtain n1 where gte_lcoeff:"\<forall>x\<ge>n1. lead_coeff p \<le> poly p x" using that pCons by auto |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2489 |
have gt_0:"lead_coeff p >0" using pCons(3) \<open>p\<noteq>0\<close> by auto |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2490 |
def n\<equiv>"max n1 (1+ \<bar>a\<bar>/(lead_coeff p))" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2491 |
show ?thesis |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2492 |
proof (rule_tac x=n in exI,rule,rule) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2493 |
fix x assume "n \<le> x" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2494 |
hence "lead_coeff p \<le> poly p x" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2495 |
using gte_lcoeff unfolding n_def by auto |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2496 |
hence " \<bar>a\<bar>/(lead_coeff p) \<ge> \<bar>a\<bar>/(poly p x)" and "poly p x>0" using gt_0 |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2497 |
by (intro frac_le,auto) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2498 |
hence "x\<ge>1+ \<bar>a\<bar>/(poly p x)" using \<open>n\<le>x\<close>[unfolded n_def] by auto |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2499 |
thus "lead_coeff (pCons a p) \<le> poly (pCons a p) x" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2500 |
using \<open>lead_coeff p \<le> poly p x\<close> \<open>poly p x>0\<close> \<open>p\<noteq>0\<close> |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2501 |
by (auto simp add:field_simps) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2502 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2503 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2504 |
ultimately show ?case by fastforce |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2505 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2506 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2507 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2508 |
subsection \<open>Algebraic numbers\<close> |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2509 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2510 |
text \<open> |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2511 |
Algebraic numbers can be defined in two equivalent ways: all real numbers that are |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2512 |
roots of rational polynomials or of integer polynomials. The Algebraic-Numbers AFP entry |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2513 |
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
|
2514 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2515 |
The equivalence is obvious since any rational polynomial can be multiplied with the |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2516 |
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
|
2517 |
\<close> |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2518 |
subsection \<open>Algebraic numbers\<close> |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2519 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2520 |
definition algebraic :: "'a :: field_char_0 \<Rightarrow> bool" where |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2521 |
"algebraic x \<longleftrightarrow> (\<exists>p. (\<forall>i. coeff p i \<in> \<int>) \<and> p \<noteq> 0 \<and> poly p x = 0)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2522 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2523 |
lemma algebraicI: |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2524 |
assumes "\<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
|
2525 |
shows "algebraic x" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2526 |
using assms unfolding algebraic_def by blast |
| 62065 | 2527 |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2528 |
lemma algebraicE: |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2529 |
assumes "algebraic x" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2530 |
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
|
2531 |
using assms unfolding algebraic_def by blast |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2532 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2533 |
lemma quotient_of_denom_pos': "snd (quotient_of x) > 0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2534 |
using quotient_of_denom_pos[OF surjective_pairing] . |
| 62065 | 2535 |
|
|
62352
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2536 |
lemma of_int_div_in_Ints: |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2537 |
"b dvd a \<Longrightarrow> of_int a div of_int b \<in> (\<int> :: 'a :: ring_div set)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2538 |
proof (cases "of_int b = (0 :: 'a)") |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2539 |
assume "b dvd a" "of_int b \<noteq> (0::'a)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2540 |
then obtain c where "a = b * c" by (elim dvdE) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2541 |
with \<open>of_int b \<noteq> (0::'a)\<close> show ?thesis by simp |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2542 |
qed auto |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2543 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2544 |
lemma of_int_divide_in_Ints: |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2545 |
"b dvd a \<Longrightarrow> of_int a / of_int b \<in> (\<int> :: 'a :: field set)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2546 |
proof (cases "of_int b = (0 :: 'a)") |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2547 |
assume "b dvd a" "of_int b \<noteq> (0::'a)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2548 |
then obtain c where "a = b * c" by (elim dvdE) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2549 |
with \<open>of_int b \<noteq> (0::'a)\<close> show ?thesis by simp |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2550 |
qed auto |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2551 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2552 |
lemma algebraic_altdef: |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2553 |
fixes p :: "'a :: field_char_0 poly" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2554 |
shows "algebraic x \<longleftrightarrow> (\<exists>p. (\<forall>i. coeff p i \<in> \<rat>) \<and> p \<noteq> 0 \<and> poly p x = 0)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2555 |
proof safe |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2556 |
fix p assume rat: "\<forall>i. coeff p i \<in> \<rat>" and root: "poly p x = 0" and nz: "p \<noteq> 0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2557 |
def cs \<equiv> "coeffs p" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2558 |
from rat have "\<forall>c\<in>range (coeff p). \<exists>c'. c = of_rat c'" unfolding Rats_def by blast |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2559 |
then obtain f where f: "\<And>i. coeff p i = of_rat (f (coeff p i))" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2560 |
by (subst (asm) bchoice_iff) blast |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2561 |
def cs' \<equiv> "map (quotient_of \<circ> f) (coeffs p)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2562 |
def d \<equiv> "Lcm (set (map snd cs'))" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2563 |
def p' \<equiv> "smult (of_int d) p" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2564 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2565 |
have "\<forall>n. coeff p' n \<in> \<int>" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2566 |
proof |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2567 |
fix n :: nat |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2568 |
show "coeff p' n \<in> \<int>" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2569 |
proof (cases "n \<le> degree p") |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2570 |
case True |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2571 |
def c \<equiv> "coeff p n" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2572 |
def a \<equiv> "fst (quotient_of (f (coeff p n)))" and b \<equiv> "snd (quotient_of (f (coeff p n)))" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2573 |
have b_pos: "b > 0" unfolding b_def using quotient_of_denom_pos' by simp |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2574 |
have "coeff p' n = of_int d * coeff p n" by (simp add: p'_def) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2575 |
also have "coeff p n = of_rat (of_int a / of_int b)" unfolding a_def b_def |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2576 |
by (subst quotient_of_div [of "f (coeff p n)", symmetric]) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2577 |
(simp_all add: f [symmetric]) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2578 |
also have "of_int d * \<dots> = of_rat (of_int (a*d) / of_int b)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2579 |
by (simp add: of_rat_mult of_rat_divide) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2580 |
also from nz True have "b \<in> snd ` set cs'" unfolding cs'_def |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2581 |
by (force simp: o_def b_def coeffs_def simp del: upt_Suc) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2582 |
hence "b dvd (a * d)" unfolding d_def by simp |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2583 |
hence "of_int (a * d) / of_int b \<in> (\<int> :: rat set)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2584 |
by (rule of_int_divide_in_Ints) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2585 |
hence "of_rat (of_int (a * d) / of_int b) \<in> \<int>" by (elim Ints_cases) auto |
|
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 (auto simp: p'_def not_le coeff_eq_0) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2588 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2589 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2590 |
moreover have "set (map snd cs') \<subseteq> {0<..}"
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2591 |
unfolding cs'_def using quotient_of_denom_pos' by (auto simp: coeffs_def simp del: upt_Suc) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2592 |
hence "d \<noteq> 0" unfolding d_def by (induction cs') simp_all |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2593 |
with nz have "p' \<noteq> 0" by (simp add: p'_def) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2594 |
moreover from root have "poly p' x = 0" by (simp add: p'_def) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2595 |
ultimately show "algebraic x" unfolding algebraic_def by blast |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2596 |
next |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2597 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2598 |
assume "algebraic x" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2599 |
then obtain p where p: "\<And>i. coeff p i \<in> \<int>" "poly p x = 0" "p \<noteq> 0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2600 |
by (force simp: algebraic_def) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2601 |
moreover have "coeff p i \<in> \<int> \<Longrightarrow> coeff p i \<in> \<rat>" for i by (elim Ints_cases) simp |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2602 |
ultimately show "(\<exists>p. (\<forall>i. coeff p i \<in> \<rat>) \<and> p \<noteq> 0 \<and> poly p x = 0)" by auto |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2603 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2604 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2605 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2606 |
text\<open>Lemmas for Derivatives\<close> |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2607 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2608 |
lemma order_unique_lemma: |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2609 |
fixes p :: "'a::idom poly" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2610 |
assumes "[:-a, 1:] ^ n dvd p" "\<not> [:-a, 1:] ^ Suc n dvd p" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2611 |
shows "n = order a p" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2612 |
unfolding Polynomial.order_def |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2613 |
apply (rule Least_equality [symmetric]) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2614 |
apply (fact assms) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2615 |
apply (rule classical) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2616 |
apply (erule notE) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2617 |
unfolding not_less_eq_eq |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2618 |
using assms(1) apply (rule power_le_dvd) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2619 |
apply assumption |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2620 |
done |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2621 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2622 |
lemma lemma_order_pderiv1: |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2623 |
"pderiv ([:- a, 1:] ^ Suc n * q) = [:- a, 1:] ^ Suc n * pderiv q + |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2624 |
smult (of_nat (Suc n)) (q * [:- a, 1:] ^ n)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2625 |
apply (simp only: pderiv_mult pderiv_power_Suc) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2626 |
apply (simp del: power_Suc of_nat_Suc add: pderiv_pCons) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2627 |
done |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2628 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2629 |
lemma lemma_order_pderiv: |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2630 |
fixes p :: "'a :: field_char_0 poly" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2631 |
assumes n: "0 < n" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2632 |
and pd: "pderiv p \<noteq> 0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2633 |
and pe: "p = [:- a, 1:] ^ n * q" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2634 |
and nd: "~ [:- a, 1:] dvd q" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2635 |
shows "n = Suc (order a (pderiv p))" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2636 |
using n |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2637 |
proof - |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2638 |
have "pderiv ([:- a, 1:] ^ n * q) \<noteq> 0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2639 |
using assms by auto |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2640 |
obtain n' where "n = Suc n'" "0 < Suc n'" "pderiv ([:- a, 1:] ^ Suc n' * q) \<noteq> 0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2641 |
using assms by (cases n) auto |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2642 |
have *: "!!k l. k dvd k * pderiv q + smult (of_nat (Suc n')) l \<Longrightarrow> k dvd l" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2643 |
by (auto simp del: of_nat_Suc simp: dvd_add_right_iff dvd_smult_iff) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2644 |
have "n' = order a (pderiv ([:- a, 1:] ^ Suc n' * q))" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2645 |
proof (rule order_unique_lemma) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2646 |
show "[:- a, 1:] ^ n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2647 |
apply (subst lemma_order_pderiv1) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2648 |
apply (rule dvd_add) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2649 |
apply (metis dvdI dvd_mult2 power_Suc2) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2650 |
apply (metis dvd_smult dvd_triv_right) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2651 |
done |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2652 |
next |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2653 |
show "\<not> [:- a, 1:] ^ Suc n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2654 |
apply (subst lemma_order_pderiv1) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2655 |
by (metis * nd dvd_mult_cancel_right power_not_zero pCons_eq_0_iff power_Suc zero_neq_one) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2656 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2657 |
then show ?thesis |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2658 |
by (metis \<open>n = Suc n'\<close> pe) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2659 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2660 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2661 |
lemma order_decomp: |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2662 |
assumes "p \<noteq> 0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2663 |
shows "\<exists>q. p = [:- a, 1:] ^ order a p * q \<and> \<not> [:- a, 1:] dvd q" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2664 |
proof - |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2665 |
from assms have A: "[:- a, 1:] ^ order a p dvd p" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2666 |
and B: "\<not> [:- a, 1:] ^ Suc (order a p) dvd p" by (auto dest: order) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2667 |
from A obtain q where C: "p = [:- a, 1:] ^ order a p * q" .. |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2668 |
with B have "\<not> [:- a, 1:] ^ Suc (order a p) dvd [:- a, 1:] ^ order a p * q" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2669 |
by simp |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2670 |
then have "\<not> [:- a, 1:] ^ order a p * [:- a, 1:] dvd [:- a, 1:] ^ order a p * q" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2671 |
by simp |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2672 |
then have D: "\<not> [:- a, 1:] dvd q" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2673 |
using idom_class.dvd_mult_cancel_left [of "[:- a, 1:] ^ order a p" "[:- a, 1:]" q] |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2674 |
by auto |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2675 |
from C D show ?thesis by blast |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2676 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2677 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2678 |
lemma order_pderiv: |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2679 |
"\<lbrakk>pderiv p \<noteq> 0; order a (p :: 'a :: field_char_0 poly) \<noteq> 0\<rbrakk> \<Longrightarrow> |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2680 |
(order a p = Suc (order a (pderiv p)))" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2681 |
apply (case_tac "p = 0", simp) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2682 |
apply (drule_tac a = a and p = p in order_decomp) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2683 |
using neq0_conv |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2684 |
apply (blast intro: lemma_order_pderiv) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2685 |
done |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2686 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2687 |
lemma order_mult: "p * q \<noteq> 0 \<Longrightarrow> order a (p * q) = order a p + order a q" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2688 |
proof - |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2689 |
def i \<equiv> "order a p" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2690 |
def j \<equiv> "order a q" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2691 |
def t \<equiv> "[:-a, 1:]" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2692 |
have t_dvd_iff: "\<And>u. t dvd u \<longleftrightarrow> poly u a = 0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2693 |
unfolding t_def by (simp add: dvd_iff_poly_eq_0) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2694 |
assume "p * q \<noteq> 0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2695 |
then show "order a (p * q) = i + j" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2696 |
apply clarsimp |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2697 |
apply (drule order [where a=a and p=p, folded i_def t_def]) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2698 |
apply (drule order [where a=a and p=q, folded j_def t_def]) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2699 |
apply clarify |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2700 |
apply (erule dvdE)+ |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2701 |
apply (rule order_unique_lemma [symmetric], fold t_def) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2702 |
apply (simp_all add: power_add t_dvd_iff) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2703 |
done |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2704 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2705 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2706 |
lemma order_smult: |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2707 |
assumes "c \<noteq> 0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2708 |
shows "order x (smult c p) = order x p" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2709 |
proof (cases "p = 0") |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2710 |
case False |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2711 |
have "smult c p = [:c:] * p" by simp |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2712 |
also from assms False have "order x \<dots> = order x [:c:] + order x p" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2713 |
by (subst order_mult) simp_all |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2714 |
also from assms have "order x [:c:] = 0" by (intro order_0I) auto |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2715 |
finally show ?thesis by simp |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2716 |
qed simp |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2717 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2718 |
(* Next two lemmas contributed by Wenda Li *) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2719 |
lemma order_1_eq_0 [simp]:"order x 1 = 0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2720 |
by (metis order_root poly_1 zero_neq_one) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2721 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2722 |
lemma order_power_n_n: "order a ([:-a,1:]^n)=n" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2723 |
proof (induct n) (*might be proved more concisely using nat_less_induct*) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2724 |
case 0 |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2725 |
thus ?case by (metis order_root poly_1 power_0 zero_neq_one) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2726 |
next |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2727 |
case (Suc n) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2728 |
have "order a ([:- a, 1:] ^ Suc n)=order a ([:- a, 1:] ^ n) + order a [:-a,1:]" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2729 |
by (metis (no_types, hide_lams) One_nat_def add_Suc_right monoid_add_class.add.right_neutral |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2730 |
one_neq_zero order_mult pCons_eq_0_iff power_add power_eq_0_iff power_one_right) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2731 |
moreover have "order a [:-a,1:]=1" unfolding order_def |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2732 |
proof (rule Least_equality,rule ccontr) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2733 |
assume "\<not> \<not> [:- a, 1:] ^ Suc 1 dvd [:- a, 1:]" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2734 |
hence "[:- a, 1:] ^ Suc 1 dvd [:- a, 1:]" by simp |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2735 |
hence "degree ([:- a, 1:] ^ Suc 1) \<le> degree ([:- a, 1:] )" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2736 |
by (rule dvd_imp_degree_le,auto) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2737 |
thus False by auto |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2738 |
next |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2739 |
fix y assume asm:"\<not> [:- a, 1:] ^ Suc y dvd [:- a, 1:]" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2740 |
show "1 \<le> y" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2741 |
proof (rule ccontr) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2742 |
assume "\<not> 1 \<le> y" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2743 |
hence "y=0" by auto |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2744 |
hence "[:- a, 1:] ^ Suc y dvd [:- a, 1:]" by auto |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2745 |
thus False using asm by auto |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2746 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2747 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2748 |
ultimately show ?case using Suc by auto |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2749 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2750 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2751 |
text\<open>Now justify the standard squarefree decomposition, i.e. f / gcd(f,f').\<close> |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2752 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2753 |
lemma order_divides: "[:-a, 1:] ^ n dvd p \<longleftrightarrow> p = 0 \<or> n \<le> order a p" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2754 |
apply (cases "p = 0", auto) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2755 |
apply (drule order_2 [where a=a and p=p]) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2756 |
apply (metis not_less_eq_eq power_le_dvd) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2757 |
apply (erule power_le_dvd [OF order_1]) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2758 |
done |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2759 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2760 |
lemma poly_squarefree_decomp_order: |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2761 |
assumes "pderiv (p :: 'a :: field_char_0 poly) \<noteq> 0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2762 |
and p: "p = q * d" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2763 |
and p': "pderiv p = e * d" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2764 |
and d: "d = r * p + s * pderiv p" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2765 |
shows "order a q = (if order a p = 0 then 0 else 1)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2766 |
proof (rule classical) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2767 |
assume 1: "order a q \<noteq> (if order a p = 0 then 0 else 1)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2768 |
from \<open>pderiv p \<noteq> 0\<close> have "p \<noteq> 0" by auto |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2769 |
with p have "order a p = order a q + order a d" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2770 |
by (simp add: order_mult) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2771 |
with 1 have "order a p \<noteq> 0" by (auto split: if_splits) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2772 |
have "order a (pderiv p) = order a e + order a d" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2773 |
using \<open>pderiv p \<noteq> 0\<close> \<open>pderiv p = e * d\<close> by (simp add: order_mult) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2774 |
have "order a p = Suc (order a (pderiv p))" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2775 |
using \<open>pderiv p \<noteq> 0\<close> \<open>order a p \<noteq> 0\<close> by (rule order_pderiv) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2776 |
have "d \<noteq> 0" using \<open>p \<noteq> 0\<close> \<open>p = q * d\<close> by simp |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2777 |
have "([:-a, 1:] ^ (order a (pderiv p))) dvd d" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2778 |
apply (simp add: d) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2779 |
apply (rule dvd_add) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2780 |
apply (rule dvd_mult) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2781 |
apply (simp add: order_divides \<open>p \<noteq> 0\<close> |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2782 |
\<open>order a p = Suc (order a (pderiv p))\<close>) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2783 |
apply (rule dvd_mult) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2784 |
apply (simp add: order_divides) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2785 |
done |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2786 |
then have "order a (pderiv p) \<le> order a d" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2787 |
using \<open>d \<noteq> 0\<close> by (simp add: order_divides) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2788 |
show ?thesis |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2789 |
using \<open>order a p = order a q + order a d\<close> |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2790 |
using \<open>order a (pderiv p) = order a e + order a d\<close> |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2791 |
using \<open>order a p = Suc (order a (pderiv p))\<close> |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2792 |
using \<open>order a (pderiv p) \<le> order a d\<close> |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2793 |
by auto |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2794 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2795 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2796 |
lemma poly_squarefree_decomp_order2: |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2797 |
"\<lbrakk>pderiv p \<noteq> (0 :: 'a :: field_char_0 poly); |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2798 |
p = q * d; |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2799 |
pderiv p = e * d; |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2800 |
d = r * p + s * pderiv p |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2801 |
\<rbrakk> \<Longrightarrow> \<forall>a. order a q = (if order a p = 0 then 0 else 1)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2802 |
by (blast intro: poly_squarefree_decomp_order) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2803 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2804 |
lemma order_pderiv2: |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2805 |
"\<lbrakk>pderiv p \<noteq> 0; order a (p :: 'a :: field_char_0 poly) \<noteq> 0\<rbrakk> |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2806 |
\<Longrightarrow> (order a (pderiv p) = n) = (order a p = Suc n)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2807 |
by (auto dest: order_pderiv) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2808 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2809 |
definition |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2810 |
rsquarefree :: "'a::idom poly => bool" where |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2811 |
"rsquarefree p = (p \<noteq> 0 & (\<forall>a. (order a p = 0) | (order a p = 1)))" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2812 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2813 |
lemma pderiv_iszero: "pderiv p = 0 \<Longrightarrow> \<exists>h. p = [:h :: 'a :: {semidom,semiring_char_0}:]"
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2814 |
apply (simp add: pderiv_eq_0_iff) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2815 |
apply (case_tac p, auto split: if_splits) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2816 |
done |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2817 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2818 |
lemma rsquarefree_roots: |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2819 |
fixes p :: "'a :: field_char_0 poly" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2820 |
shows "rsquarefree p = (\<forall>a. \<not>(poly p a = 0 \<and> poly (pderiv p) a = 0))" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2821 |
apply (simp add: rsquarefree_def) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2822 |
apply (case_tac "p = 0", simp, simp) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2823 |
apply (case_tac "pderiv p = 0") |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2824 |
apply simp |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2825 |
apply (drule pderiv_iszero, clarsimp) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2826 |
apply (metis coeff_0 coeff_pCons_0 degree_pCons_0 le0 le_antisym order_degree) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2827 |
apply (force simp add: order_root order_pderiv2) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2828 |
done |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2829 |
|
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2830 |
lemma poly_squarefree_decomp: |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2831 |
assumes "pderiv (p :: 'a :: field_char_0 poly) \<noteq> 0" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2832 |
and "p = q * d" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2833 |
and "pderiv p = e * d" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2834 |
and "d = r * p + s * pderiv p" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2835 |
shows "rsquarefree q & (\<forall>a. (poly q a = 0) = (poly p a = 0))" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2836 |
proof - |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2837 |
from \<open>pderiv p \<noteq> 0\<close> have "p \<noteq> 0" by auto |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2838 |
with \<open>p = q * d\<close> have "q \<noteq> 0" by simp |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2839 |
have "\<forall>a. order a q = (if order a p = 0 then 0 else 1)" |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2840 |
using assms by (rule poly_squarefree_decomp_order2) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2841 |
with \<open>p \<noteq> 0\<close> \<open>q \<noteq> 0\<close> show ?thesis |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2842 |
by (simp add: rsquarefree_def order_root) |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2843 |
qed |
|
35a9e1cbb5b3
separated potentially conflicting type class instance into separate theory
haftmann
parents:
62351
diff
changeset
|
2844 |
|
| 52380 | 2845 |
|
2846 |
no_notation cCons (infixr "##" 65) |
|
| 31663 | 2847 |
|
| 29478 | 2848 |
end |