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