author | huffman |
Tue, 15 Nov 2011 12:39:49 +0100 | |
changeset 45503 | 44790ec65f70 |
parent 45129 | 1fce03e3e8ad |
child 46991 | 196f2d9406c4 |
permissions | -rw-r--r-- |
33154 | 1 |
(* Title: HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy |
2 |
Author: Amine Chaieb |
|
3 |
*) |
|
4 |
||
35046 | 5 |
header {* Implementation and verification of multivariate polynomials *} |
33154 | 6 |
|
7 |
theory Reflected_Multivariate_Polynomial |
|
41413
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
wenzelm
parents:
41404
diff
changeset
|
8 |
imports Complex_Main "~~/src/HOL/Library/Abstract_Rat" Polynomial_List |
33154 | 9 |
begin |
10 |
||
35046 | 11 |
(* Implementation *) |
33154 | 12 |
|
13 |
subsection{* Datatype of polynomial expressions *} |
|
14 |
||
15 |
datatype poly = C Num| Bound nat| Add poly poly|Sub poly poly |
|
16 |
| Mul poly poly| Neg poly| Pw poly nat| CN poly nat poly |
|
17 |
||
35054 | 18 |
abbreviation poly_0 :: "poly" ("0\<^sub>p") where "0\<^sub>p \<equiv> C (0\<^sub>N)" |
19 |
abbreviation poly_p :: "int \<Rightarrow> poly" ("_\<^sub>p") where "i\<^sub>p \<equiv> C (i\<^sub>N)" |
|
33154 | 20 |
|
21 |
subsection{* Boundedness, substitution and all that *} |
|
39246 | 22 |
primrec polysize:: "poly \<Rightarrow> nat" where |
33154 | 23 |
"polysize (C c) = 1" |
39246 | 24 |
| "polysize (Bound n) = 1" |
25 |
| "polysize (Neg p) = 1 + polysize p" |
|
26 |
| "polysize (Add p q) = 1 + polysize p + polysize q" |
|
27 |
| "polysize (Sub p q) = 1 + polysize p + polysize q" |
|
28 |
| "polysize (Mul p q) = 1 + polysize p + polysize q" |
|
29 |
| "polysize (Pw p n) = 1 + polysize p" |
|
30 |
| "polysize (CN c n p) = 4 + polysize c + polysize p" |
|
33154 | 31 |
|
39246 | 32 |
primrec polybound0:: "poly \<Rightarrow> bool" (* a poly is INDEPENDENT of Bound 0 *) where |
33154 | 33 |
"polybound0 (C c) = True" |
39246 | 34 |
| "polybound0 (Bound n) = (n>0)" |
35 |
| "polybound0 (Neg a) = polybound0 a" |
|
36 |
| "polybound0 (Add a b) = (polybound0 a \<and> polybound0 b)" |
|
37 |
| "polybound0 (Sub a b) = (polybound0 a \<and> polybound0 b)" |
|
38 |
| "polybound0 (Mul a b) = (polybound0 a \<and> polybound0 b)" |
|
39 |
| "polybound0 (Pw p n) = (polybound0 p)" |
|
40 |
| "polybound0 (CN c n p) = (n \<noteq> 0 \<and> polybound0 c \<and> polybound0 p)" |
|
41 |
||
42 |
primrec polysubst0:: "poly \<Rightarrow> poly \<Rightarrow> poly" (* substitute a poly into a poly for Bound 0 *) where |
|
33154 | 43 |
"polysubst0 t (C c) = (C c)" |
39246 | 44 |
| "polysubst0 t (Bound n) = (if n=0 then t else Bound n)" |
45 |
| "polysubst0 t (Neg a) = Neg (polysubst0 t a)" |
|
46 |
| "polysubst0 t (Add a b) = Add (polysubst0 t a) (polysubst0 t b)" |
|
47 |
| "polysubst0 t (Sub a b) = Sub (polysubst0 t a) (polysubst0 t b)" |
|
48 |
| "polysubst0 t (Mul a b) = Mul (polysubst0 t a) (polysubst0 t b)" |
|
49 |
| "polysubst0 t (Pw p n) = Pw (polysubst0 t p) n" |
|
50 |
| "polysubst0 t (CN c n p) = (if n=0 then Add (polysubst0 t c) (Mul t (polysubst0 t p)) |
|
33154 | 51 |
else CN (polysubst0 t c) n (polysubst0 t p))" |
52 |
||
41808 | 53 |
fun decrpoly:: "poly \<Rightarrow> poly" |
54 |
where |
|
33154 | 55 |
"decrpoly (Bound n) = Bound (n - 1)" |
41808 | 56 |
| "decrpoly (Neg a) = Neg (decrpoly a)" |
57 |
| "decrpoly (Add a b) = Add (decrpoly a) (decrpoly b)" |
|
58 |
| "decrpoly (Sub a b) = Sub (decrpoly a) (decrpoly b)" |
|
59 |
| "decrpoly (Mul a b) = Mul (decrpoly a) (decrpoly b)" |
|
60 |
| "decrpoly (Pw p n) = Pw (decrpoly p) n" |
|
61 |
| "decrpoly (CN c n p) = CN (decrpoly c) (n - 1) (decrpoly p)" |
|
62 |
| "decrpoly a = a" |
|
33154 | 63 |
|
64 |
subsection{* Degrees and heads and coefficients *} |
|
65 |
||
41808 | 66 |
fun degree:: "poly \<Rightarrow> nat" |
67 |
where |
|
33154 | 68 |
"degree (CN c 0 p) = 1 + degree p" |
41808 | 69 |
| "degree p = 0" |
33154 | 70 |
|
41808 | 71 |
fun head:: "poly \<Rightarrow> poly" |
72 |
where |
|
33154 | 73 |
"head (CN c 0 p) = head p" |
41808 | 74 |
| "head p = p" |
75 |
||
76 |
(* More general notions of degree and head *) |
|
77 |
fun degreen:: "poly \<Rightarrow> nat \<Rightarrow> nat" |
|
78 |
where |
|
33154 | 79 |
"degreen (CN c n p) = (\<lambda>m. if n=m then 1 + degreen p n else 0)" |
41808 | 80 |
|"degreen p = (\<lambda>m. 0)" |
33154 | 81 |
|
41808 | 82 |
fun headn:: "poly \<Rightarrow> nat \<Rightarrow> poly" |
83 |
where |
|
84 |
"headn (CN c n p) = (\<lambda>m. if n \<le> m then headn p m else CN c n p)" |
|
85 |
| "headn p = (\<lambda>m. p)" |
|
33154 | 86 |
|
41808 | 87 |
fun coefficients:: "poly \<Rightarrow> poly list" |
88 |
where |
|
89 |
"coefficients (CN c 0 p) = c#(coefficients p)" |
|
90 |
| "coefficients p = [p]" |
|
33154 | 91 |
|
41808 | 92 |
fun isconstant:: "poly \<Rightarrow> bool" |
93 |
where |
|
94 |
"isconstant (CN c 0 p) = False" |
|
95 |
| "isconstant p = True" |
|
33154 | 96 |
|
41808 | 97 |
fun behead:: "poly \<Rightarrow> poly" |
98 |
where |
|
99 |
"behead (CN c 0 p) = (let p' = behead p in if p' = 0\<^sub>p then c else CN c 0 p')" |
|
100 |
| "behead p = 0\<^sub>p" |
|
101 |
||
102 |
fun headconst:: "poly \<Rightarrow> Num" |
|
103 |
where |
|
33154 | 104 |
"headconst (CN c n p) = headconst p" |
41808 | 105 |
| "headconst (C n) = n" |
33154 | 106 |
|
107 |
subsection{* Operations for normalization *} |
|
41812 | 108 |
|
109 |
||
110 |
declare if_cong[fundef_cong del] |
|
111 |
declare let_cong[fundef_cong del] |
|
112 |
||
113 |
fun polyadd :: "poly \<Rightarrow> poly \<Rightarrow> poly" (infixl "+\<^sub>p" 60) |
|
114 |
where |
|
115 |
"polyadd (C c) (C c') = C (c+\<^sub>Nc')" |
|
116 |
| "polyadd (C c) (CN c' n' p') = CN (polyadd (C c) c') n' p'" |
|
117 |
| "polyadd (CN c n p) (C c') = CN (polyadd c (C c')) n p" |
|
118 |
| "polyadd (CN c n p) (CN c' n' p') = |
|
119 |
(if n < n' then CN (polyadd c (CN c' n' p')) n p |
|
120 |
else if n'<n then CN (polyadd (CN c n p) c') n' p' |
|
121 |
else (let cc' = polyadd c c' ; |
|
122 |
pp' = polyadd p p' |
|
33154 | 123 |
in (if pp' = 0\<^sub>p then cc' else CN cc' n pp')))" |
41812 | 124 |
| "polyadd a b = Add a b" |
125 |
||
33154 | 126 |
|
41808 | 127 |
fun polyneg :: "poly \<Rightarrow> poly" ("~\<^sub>p") |
128 |
where |
|
33154 | 129 |
"polyneg (C c) = C (~\<^sub>N c)" |
41808 | 130 |
| "polyneg (CN c n p) = CN (polyneg c) n (polyneg p)" |
131 |
| "polyneg a = Neg a" |
|
33154 | 132 |
|
41814 | 133 |
definition polysub :: "poly \<Rightarrow> poly \<Rightarrow> poly" (infixl "-\<^sub>p" 60) |
134 |
where |
|
135 |
"p -\<^sub>p q = polyadd p (polyneg q)" |
|
41813 | 136 |
|
137 |
fun polymul :: "poly \<Rightarrow> poly \<Rightarrow> poly" (infixl "*\<^sub>p" 60) |
|
138 |
where |
|
139 |
"polymul (C c) (C c') = C (c*\<^sub>Nc')" |
|
140 |
| "polymul (C c) (CN c' n' p') = |
|
141 |
(if c = 0\<^sub>N then 0\<^sub>p else CN (polymul (C c) c') n' (polymul (C c) p'))" |
|
142 |
| "polymul (CN c n p) (C c') = |
|
143 |
(if c' = 0\<^sub>N then 0\<^sub>p else CN (polymul c (C c')) n (polymul p (C c')))" |
|
144 |
| "polymul (CN c n p) (CN c' n' p') = |
|
145 |
(if n<n' then CN (polymul c (CN c' n' p')) n (polymul p (CN c' n' p')) |
|
33154 | 146 |
else if n' < n |
41813 | 147 |
then CN (polymul (CN c n p) c') n' (polymul (CN c n p) p') |
148 |
else polyadd (polymul (CN c n p) c') (CN 0\<^sub>p n' (polymul (CN c n p) p')))" |
|
149 |
| "polymul a b = Mul a b" |
|
41808 | 150 |
|
41812 | 151 |
declare if_cong[fundef_cong] |
152 |
declare let_cong[fundef_cong] |
|
153 |
||
41808 | 154 |
fun polypow :: "nat \<Rightarrow> poly \<Rightarrow> poly" |
155 |
where |
|
33154 | 156 |
"polypow 0 = (\<lambda>p. 1\<^sub>p)" |
41813 | 157 |
| "polypow n = (\<lambda>p. let q = polypow (n div 2) p ; d = polymul q q in |
158 |
if even n then d else polymul p d)" |
|
33154 | 159 |
|
35054 | 160 |
abbreviation poly_pow :: "poly \<Rightarrow> nat \<Rightarrow> poly" (infixl "^\<^sub>p" 60) |
161 |
where "a ^\<^sub>p k \<equiv> polypow k a" |
|
33154 | 162 |
|
41808 | 163 |
function polynate :: "poly \<Rightarrow> poly" |
164 |
where |
|
33154 | 165 |
"polynate (Bound n) = CN 0\<^sub>p n 1\<^sub>p" |
41808 | 166 |
| "polynate (Add p q) = (polynate p +\<^sub>p polynate q)" |
167 |
| "polynate (Sub p q) = (polynate p -\<^sub>p polynate q)" |
|
168 |
| "polynate (Mul p q) = (polynate p *\<^sub>p polynate q)" |
|
169 |
| "polynate (Neg p) = (~\<^sub>p (polynate p))" |
|
170 |
| "polynate (Pw p n) = ((polynate p) ^\<^sub>p n)" |
|
171 |
| "polynate (CN c n p) = polynate (Add c (Mul (Bound n) p))" |
|
172 |
| "polynate (C c) = C (normNum c)" |
|
173 |
by pat_completeness auto |
|
174 |
termination by (relation "measure polysize") auto |
|
33154 | 175 |
|
176 |
fun poly_cmul :: "Num \<Rightarrow> poly \<Rightarrow> poly" where |
|
177 |
"poly_cmul y (C x) = C (y *\<^sub>N x)" |
|
178 |
| "poly_cmul y (CN c n p) = CN (poly_cmul y c) n (poly_cmul y p)" |
|
179 |
| "poly_cmul y p = C y *\<^sub>p p" |
|
180 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35054
diff
changeset
|
181 |
definition monic :: "poly \<Rightarrow> (poly \<times> bool)" where |
33154 | 182 |
"monic p \<equiv> (let h = headconst p in if h = 0\<^sub>N then (p,False) else ((C (Ninv h)) *\<^sub>p p, 0>\<^sub>N h))" |
183 |
||
184 |
subsection{* Pseudo-division *} |
|
185 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35054
diff
changeset
|
186 |
definition shift1 :: "poly \<Rightarrow> poly" where |
33154 | 187 |
"shift1 p \<equiv> CN 0\<^sub>p 0 p" |
188 |
||
39246 | 189 |
abbreviation funpow :: "nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a)" where |
190 |
"funpow \<equiv> compow" |
|
191 |
||
41403
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
192 |
partial_function (tailrec) polydivide_aux :: "poly \<Rightarrow> nat \<Rightarrow> poly \<Rightarrow> nat \<Rightarrow> poly \<Rightarrow> nat \<times> poly" |
33154 | 193 |
where |
41403
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
194 |
"polydivide_aux a n p k s = |
33154 | 195 |
(if s = 0\<^sub>p then (k,s) |
196 |
else (let b = head s; m = degree s in |
|
197 |
(if m < n then (k,s) else |
|
198 |
(let p'= funpow (m - n) shift1 p in |
|
41403
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
199 |
(if a = b then polydivide_aux a n p k (s -\<^sub>p p') |
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
200 |
else polydivide_aux a n p (Suc k) ((a *\<^sub>p s) -\<^sub>p (b *\<^sub>p p')))))))" |
33154 | 201 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35054
diff
changeset
|
202 |
definition polydivide :: "poly \<Rightarrow> poly \<Rightarrow> (nat \<times> poly)" where |
41403
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
203 |
"polydivide s p \<equiv> polydivide_aux (head p) (degree p) p 0 s" |
33154 | 204 |
|
205 |
fun poly_deriv_aux :: "nat \<Rightarrow> poly \<Rightarrow> poly" where |
|
206 |
"poly_deriv_aux n (CN c 0 p) = CN (poly_cmul ((int n)\<^sub>N) c) 0 (poly_deriv_aux (n + 1) p)" |
|
207 |
| "poly_deriv_aux n p = poly_cmul ((int n)\<^sub>N) p" |
|
208 |
||
209 |
fun poly_deriv :: "poly \<Rightarrow> poly" where |
|
210 |
"poly_deriv (CN c 0 p) = poly_deriv_aux 1 p" |
|
211 |
| "poly_deriv p = 0\<^sub>p" |
|
212 |
||
213 |
subsection{* Semantics of the polynomial representation *} |
|
214 |
||
39246 | 215 |
primrec Ipoly :: "'a list \<Rightarrow> poly \<Rightarrow> 'a::{field_char_0, field_inverse_zero, power}" where |
33154 | 216 |
"Ipoly bs (C c) = INum c" |
39246 | 217 |
| "Ipoly bs (Bound n) = bs!n" |
218 |
| "Ipoly bs (Neg a) = - Ipoly bs a" |
|
219 |
| "Ipoly bs (Add a b) = Ipoly bs a + Ipoly bs b" |
|
220 |
| "Ipoly bs (Sub a b) = Ipoly bs a - Ipoly bs b" |
|
221 |
| "Ipoly bs (Mul a b) = Ipoly bs a * Ipoly bs b" |
|
222 |
| "Ipoly bs (Pw t n) = (Ipoly bs t) ^ n" |
|
223 |
| "Ipoly bs (CN c n p) = (Ipoly bs c) + (bs!n)*(Ipoly bs p)" |
|
224 |
||
35054 | 225 |
abbreviation |
36409 | 226 |
Ipoly_syntax :: "poly \<Rightarrow> 'a list \<Rightarrow>'a::{field_char_0, field_inverse_zero, power}" ("\<lparr>_\<rparr>\<^sub>p\<^bsup>_\<^esup>") |
35054 | 227 |
where "\<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<equiv> Ipoly bs p" |
33154 | 228 |
|
229 |
lemma Ipoly_CInt: "Ipoly bs (C (i,1)) = of_int i" |
|
230 |
by (simp add: INum_def) |
|
231 |
lemma Ipoly_CRat: "Ipoly bs (C (i, j)) = of_int i / of_int j" |
|
232 |
by (simp add: INum_def) |
|
233 |
||
234 |
lemmas RIpoly_eqs = Ipoly.simps(2-7) Ipoly_CInt Ipoly_CRat |
|
235 |
||
236 |
subsection {* Normal form and normalization *} |
|
237 |
||
41808 | 238 |
fun isnpolyh:: "poly \<Rightarrow> nat \<Rightarrow> bool" |
239 |
where |
|
33154 | 240 |
"isnpolyh (C c) = (\<lambda>k. isnormNum c)" |
41808 | 241 |
| "isnpolyh (CN c n p) = (\<lambda>k. n \<ge> k \<and> (isnpolyh c (Suc n)) \<and> (isnpolyh p n) \<and> (p \<noteq> 0\<^sub>p))" |
242 |
| "isnpolyh p = (\<lambda>k. False)" |
|
33154 | 243 |
|
244 |
lemma isnpolyh_mono: "\<lbrakk>n' \<le> n ; isnpolyh p n\<rbrakk> \<Longrightarrow> isnpolyh p n'" |
|
245 |
by (induct p rule: isnpolyh.induct, auto) |
|
246 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35054
diff
changeset
|
247 |
definition isnpoly :: "poly \<Rightarrow> bool" where |
33154 | 248 |
"isnpoly p \<equiv> isnpolyh p 0" |
249 |
||
250 |
text{* polyadd preserves normal forms *} |
|
251 |
||
252 |
lemma polyadd_normh: "\<lbrakk>isnpolyh p n0 ; isnpolyh q n1\<rbrakk> |
|
41812 | 253 |
\<Longrightarrow> isnpolyh (polyadd p q) (min n0 n1)" |
33154 | 254 |
proof(induct p q arbitrary: n0 n1 rule: polyadd.induct) |
41812 | 255 |
case (2 ab c' n' p' n0 n1) |
41815 | 256 |
from 2 have th1: "isnpolyh (C ab) (Suc n')" by simp |
41807 | 257 |
from 2(3) have th2: "isnpolyh c' (Suc n')" and nplen1: "n' \<ge> n1" by simp_all |
33154 | 258 |
with isnpolyh_mono have cp: "isnpolyh c' (Suc n')" by simp |
41815 | 259 |
with 2(1)[OF th1 th2] have th3:"isnpolyh (C ab +\<^sub>p c') (Suc n')" by simp |
33154 | 260 |
from nplen1 have n01len1: "min n0 n1 \<le> n'" by simp |
41807 | 261 |
thus ?case using 2 th3 by simp |
33154 | 262 |
next |
41812 | 263 |
case (3 c' n' p' ab n1 n0) |
41815 | 264 |
from 3 have th1: "isnpolyh (C ab) (Suc n')" by simp |
41807 | 265 |
from 3(2) have th2: "isnpolyh c' (Suc n')" and nplen1: "n' \<ge> n1" by simp_all |
33154 | 266 |
with isnpolyh_mono have cp: "isnpolyh c' (Suc n')" by simp |
41815 | 267 |
with 3(1)[OF th2 th1] have th3:"isnpolyh (c' +\<^sub>p C ab) (Suc n')" by simp |
33154 | 268 |
from nplen1 have n01len1: "min n0 n1 \<le> n'" by simp |
41807 | 269 |
thus ?case using 3 th3 by simp |
33154 | 270 |
next |
271 |
case (4 c n p c' n' p' n0 n1) |
|
272 |
hence nc: "isnpolyh c (Suc n)" and np: "isnpolyh p n" by simp_all |
|
41807 | 273 |
from 4 have nc': "isnpolyh c' (Suc n')" and np': "isnpolyh p' n'" by simp_all |
274 |
from 4 have ngen0: "n \<ge> n0" by simp |
|
275 |
from 4 have n'gen1: "n' \<ge> n1" by simp |
|
33154 | 276 |
have "n < n' \<or> n' < n \<or> n = n'" by auto |
41763 | 277 |
moreover {assume eq: "n = n'" |
41812 | 278 |
with "4.hyps"(3)[OF nc nc'] |
33154 | 279 |
have ncc':"isnpolyh (c +\<^sub>p c') (Suc n)" by auto |
280 |
hence ncc'n01: "isnpolyh (c +\<^sub>p c') (min n0 n1)" |
|
281 |
using isnpolyh_mono[where n'="min n0 n1" and n="Suc n"] ngen0 n'gen1 by auto |
|
41812 | 282 |
from eq "4.hyps"(4)[OF np np'] have npp': "isnpolyh (p +\<^sub>p p') n" by simp |
33154 | 283 |
have minle: "min n0 n1 \<le> n'" using ngen0 n'gen1 eq by simp |
41815 | 284 |
from minle npp' ncc'n01 4 eq ngen0 n'gen1 ncc' have ?case by (simp add: Let_def)} |
33154 | 285 |
moreover {assume lt: "n < n'" |
286 |
have "min n0 n1 \<le> n0" by simp |
|
41815 | 287 |
with 4 lt have th1:"min n0 n1 \<le> n" by auto |
41807 | 288 |
from 4 have th21: "isnpolyh c (Suc n)" by simp |
289 |
from 4 have th22: "isnpolyh (CN c' n' p') n'" by simp |
|
33154 | 290 |
from lt have th23: "min (Suc n) n' = Suc n" by arith |
41812 | 291 |
from "4.hyps"(1)[OF th21 th22] |
292 |
have "isnpolyh (polyadd c (CN c' n' p')) (Suc n)" using th23 by simp |
|
41815 | 293 |
with 4 lt th1 have ?case by simp } |
33154 | 294 |
moreover {assume gt: "n' < n" hence gt': "n' < n \<and> \<not> n < n'" by simp |
295 |
have "min n0 n1 \<le> n1" by simp |
|
41815 | 296 |
with 4 gt have th1:"min n0 n1 \<le> n'" by auto |
41807 | 297 |
from 4 have th21: "isnpolyh c' (Suc n')" by simp_all |
298 |
from 4 have th22: "isnpolyh (CN c n p) n" by simp |
|
33154 | 299 |
from gt have th23: "min n (Suc n') = Suc n'" by arith |
41812 | 300 |
from "4.hyps"(2)[OF th22 th21] |
301 |
have "isnpolyh (polyadd (CN c n p) c') (Suc n')" using th23 by simp |
|
41807 | 302 |
with 4 gt th1 have ?case by simp} |
33154 | 303 |
ultimately show ?case by blast |
304 |
qed auto |
|
305 |
||
41812 | 306 |
lemma polyadd[simp]: "Ipoly bs (polyadd p q) = Ipoly bs p + Ipoly bs q" |
36349 | 307 |
by (induct p q rule: polyadd.induct, auto simp add: Let_def field_simps right_distrib[symmetric] simp del: right_distrib) |
33154 | 308 |
|
41812 | 309 |
lemma polyadd_norm: "\<lbrakk> isnpoly p ; isnpoly q\<rbrakk> \<Longrightarrow> isnpoly (polyadd p q)" |
33154 | 310 |
using polyadd_normh[of "p" "0" "q" "0"] isnpoly_def by simp |
311 |
||
41404 | 312 |
text{* The degree of addition and other general lemmas needed for the normal form of polymul *} |
33154 | 313 |
|
314 |
lemma polyadd_different_degreen: |
|
315 |
"\<lbrakk>isnpolyh p n0 ; isnpolyh q n1; degreen p m \<noteq> degreen q m ; m \<le> min n0 n1\<rbrakk> \<Longrightarrow> |
|
41812 | 316 |
degreen (polyadd p q) m = max (degreen p m) (degreen q m)" |
33154 | 317 |
proof (induct p q arbitrary: m n0 n1 rule: polyadd.induct) |
318 |
case (4 c n p c' n' p' m n0 n1) |
|
41763 | 319 |
have "n' = n \<or> n < n' \<or> n' < n" by arith |
320 |
thus ?case |
|
321 |
proof (elim disjE) |
|
322 |
assume [simp]: "n' = n" |
|
41812 | 323 |
from 4(4)[of n n m] 4(3)[of "Suc n" "Suc n" m] 4(5-7) |
41763 | 324 |
show ?thesis by (auto simp: Let_def) |
325 |
next |
|
326 |
assume "n < n'" |
|
327 |
with 4 show ?thesis by auto |
|
328 |
next |
|
329 |
assume "n' < n" |
|
330 |
with 4 show ?thesis by auto |
|
331 |
qed |
|
332 |
qed auto |
|
33154 | 333 |
|
334 |
lemma headnz[simp]: "\<lbrakk>isnpolyh p n ; p \<noteq> 0\<^sub>p\<rbrakk> \<Longrightarrow> headn p m \<noteq> 0\<^sub>p" |
|
335 |
by (induct p arbitrary: n rule: headn.induct, auto) |
|
336 |
lemma degree_isnpolyh_Suc[simp]: "isnpolyh p (Suc n) \<Longrightarrow> degree p = 0" |
|
337 |
by (induct p arbitrary: n rule: degree.induct, auto) |
|
338 |
lemma degreen_0[simp]: "isnpolyh p n \<Longrightarrow> m < n \<Longrightarrow> degreen p m = 0" |
|
339 |
by (induct p arbitrary: n rule: degreen.induct, auto) |
|
340 |
||
341 |
lemma degree_isnpolyh_Suc': "n > 0 \<Longrightarrow> isnpolyh p n \<Longrightarrow> degree p = 0" |
|
342 |
by (induct p arbitrary: n rule: degree.induct, auto) |
|
343 |
||
344 |
lemma degree_npolyhCN[simp]: "isnpolyh (CN c n p) n0 \<Longrightarrow> degree c = 0" |
|
345 |
using degree_isnpolyh_Suc by auto |
|
346 |
lemma degreen_npolyhCN[simp]: "isnpolyh (CN c n p) n0 \<Longrightarrow> degreen c n = 0" |
|
347 |
using degreen_0 by auto |
|
348 |
||
349 |
||
350 |
lemma degreen_polyadd: |
|
351 |
assumes np: "isnpolyh p n0" and nq: "isnpolyh q n1" and m: "m \<le> max n0 n1" |
|
352 |
shows "degreen (p +\<^sub>p q) m \<le> max (degreen p m) (degreen q m)" |
|
353 |
using np nq m |
|
354 |
proof (induct p q arbitrary: n0 n1 m rule: polyadd.induct) |
|
355 |
case (2 c c' n' p' n0 n1) thus ?case by (cases n', simp_all) |
|
356 |
next |
|
357 |
case (3 c n p c' n0 n1) thus ?case by (cases n, auto) |
|
358 |
next |
|
359 |
case (4 c n p c' n' p' n0 n1 m) |
|
41763 | 360 |
have "n' = n \<or> n < n' \<or> n' < n" by arith |
361 |
thus ?case |
|
362 |
proof (elim disjE) |
|
363 |
assume [simp]: "n' = n" |
|
41812 | 364 |
from 4(4)[of n n m] 4(3)[of "Suc n" "Suc n" m] 4(5-7) |
41763 | 365 |
show ?thesis by (auto simp: Let_def) |
366 |
qed simp_all |
|
33154 | 367 |
qed auto |
368 |
||
41812 | 369 |
lemma polyadd_eq_const_degreen: "\<lbrakk> isnpolyh p n0 ; isnpolyh q n1 ; polyadd p q = C c\<rbrakk> |
33154 | 370 |
\<Longrightarrow> degreen p m = degreen q m" |
371 |
proof (induct p q arbitrary: m n0 n1 c rule: polyadd.induct) |
|
372 |
case (4 c n p c' n' p' m n0 n1 x) |
|
41807 | 373 |
{assume nn': "n' < n" hence ?case using 4 by simp} |
33154 | 374 |
moreover |
375 |
{assume nn':"\<not> n' < n" hence "n < n' \<or> n = n'" by arith |
|
41807 | 376 |
moreover {assume "n < n'" with 4 have ?case by simp } |
377 |
moreover {assume eq: "n = n'" hence ?case using 4 |
|
41763 | 378 |
apply (cases "p +\<^sub>p p' = 0\<^sub>p") |
379 |
apply (auto simp add: Let_def) |
|
380 |
by blast |
|
381 |
} |
|
33154 | 382 |
ultimately have ?case by blast} |
383 |
ultimately show ?case by blast |
|
384 |
qed simp_all |
|
385 |
||
386 |
lemma polymul_properties: |
|
36409 | 387 |
assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" |
33154 | 388 |
and np: "isnpolyh p n0" and nq: "isnpolyh q n1" and m: "m \<le> min n0 n1" |
389 |
shows "isnpolyh (p *\<^sub>p q) (min n0 n1)" |
|
390 |
and "(p *\<^sub>p q = 0\<^sub>p) = (p = 0\<^sub>p \<or> q = 0\<^sub>p)" |
|
391 |
and "degreen (p *\<^sub>p q) m = (if (p = 0\<^sub>p \<or> q = 0\<^sub>p) then 0 |
|
392 |
else degreen p m + degreen q m)" |
|
393 |
using np nq m |
|
394 |
proof(induct p q arbitrary: n0 n1 m rule: polymul.induct) |
|
41813 | 395 |
case (2 c c' n' p') |
33154 | 396 |
{ case (1 n0 n1) |
41813 | 397 |
with "2.hyps"(4-6)[of n' n' n'] |
398 |
and "2.hyps"(1-3)[of "Suc n'" "Suc n'" n'] |
|
41811 | 399 |
show ?case by (auto simp add: min_def) |
33154 | 400 |
next |
401 |
case (2 n0 n1) thus ?case by auto |
|
402 |
next |
|
403 |
case (3 n0 n1) thus ?case using "2.hyps" by auto } |
|
404 |
next |
|
41813 | 405 |
case (3 c n p c') |
41811 | 406 |
{ case (1 n0 n1) |
41813 | 407 |
with "3.hyps"(4-6)[of n n n] |
408 |
"3.hyps"(1-3)[of "Suc n" "Suc n" n] |
|
41811 | 409 |
show ?case by (auto simp add: min_def) |
33154 | 410 |
next |
41811 | 411 |
case (2 n0 n1) thus ?case by auto |
33154 | 412 |
next |
413 |
case (3 n0 n1) thus ?case using "3.hyps" by auto } |
|
414 |
next |
|
415 |
case (4 c n p c' n' p') |
|
416 |
let ?cnp = "CN c n p" let ?cnp' = "CN c' n' p'" |
|
41811 | 417 |
{ |
418 |
case (1 n0 n1) |
|
33154 | 419 |
hence cnp: "isnpolyh ?cnp n" and cnp': "isnpolyh ?cnp' n'" |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
420 |
and np: "isnpolyh p n" and nc: "isnpolyh c (Suc n)" |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
421 |
and np': "isnpolyh p' n'" and nc': "isnpolyh c' (Suc n')" |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
422 |
and nn0: "n \<ge> n0" and nn1:"n' \<ge> n1" |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
423 |
by simp_all |
41811 | 424 |
{ assume "n < n'" |
41813 | 425 |
with "4.hyps"(4-5)[OF np cnp', of n] |
426 |
"4.hyps"(1)[OF nc cnp', of n] nn0 cnp |
|
41811 | 427 |
have ?case by (simp add: min_def) |
428 |
} moreover { |
|
429 |
assume "n' < n" |
|
41813 | 430 |
with "4.hyps"(16-17)[OF cnp np', of "n'"] |
431 |
"4.hyps"(13)[OF cnp nc', of "Suc n'"] nn1 cnp' |
|
41811 | 432 |
have ?case |
433 |
by (cases "Suc n' = n", simp_all add: min_def) |
|
434 |
} moreover { |
|
435 |
assume "n' = n" |
|
41813 | 436 |
with "4.hyps"(16-17)[OF cnp np', of n] |
437 |
"4.hyps"(13)[OF cnp nc', of n] cnp cnp' nn1 nn0 |
|
41811 | 438 |
have ?case |
439 |
apply (auto intro!: polyadd_normh) |
|
440 |
apply (simp_all add: min_def isnpolyh_mono[OF nn0]) |
|
441 |
done |
|
442 |
} |
|
443 |
ultimately show ?case by arith |
|
444 |
next |
|
445 |
fix n0 n1 m |
|
33154 | 446 |
assume np: "isnpolyh ?cnp n0" and np':"isnpolyh ?cnp' n1" |
447 |
and m: "m \<le> min n0 n1" |
|
448 |
let ?d = "degreen (?cnp *\<^sub>p ?cnp') m" |
|
449 |
let ?d1 = "degreen ?cnp m" |
|
450 |
let ?d2 = "degreen ?cnp' m" |
|
451 |
let ?eq = "?d = (if ?cnp = 0\<^sub>p \<or> ?cnp' = 0\<^sub>p then 0 else ?d1 + ?d2)" |
|
452 |
have "n'<n \<or> n < n' \<or> n' = n" by auto |
|
453 |
moreover |
|
454 |
{assume "n' < n \<or> n < n'" |
|
41813 | 455 |
with "4.hyps"(3,6,18) np np' m |
41811 | 456 |
have ?eq by auto } |
33154 | 457 |
moreover |
41811 | 458 |
{assume nn': "n' = n" hence nn:"\<not> n' < n \<and> \<not> n < n'" by arith |
41813 | 459 |
from "4.hyps"(16,18)[of n n' n] |
460 |
"4.hyps"(13,14)[of n "Suc n'" n] |
|
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
461 |
np np' nn' |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
462 |
have norm: "isnpolyh ?cnp n" "isnpolyh c' (Suc n)" "isnpolyh (?cnp *\<^sub>p c') n" |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
463 |
"isnpolyh p' n" "isnpolyh (?cnp *\<^sub>p p') n" "isnpolyh (CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n" |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
464 |
"(?cnp *\<^sub>p c' = 0\<^sub>p) = (c' = 0\<^sub>p)" |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
465 |
"?cnp *\<^sub>p p' \<noteq> 0\<^sub>p" by (auto simp add: min_def) |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
466 |
{assume mn: "m = n" |
41813 | 467 |
from "4.hyps"(17,18)[OF norm(1,4), of n] |
468 |
"4.hyps"(13,15)[OF norm(1,2), of n] norm nn' mn |
|
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
469 |
have degs: "degreen (?cnp *\<^sub>p c') n = |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
470 |
(if c'=0\<^sub>p then 0 else ?d1 + degreen c' n)" |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
471 |
"degreen (?cnp *\<^sub>p p') n = ?d1 + degreen p' n" by (simp_all add: min_def) |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
472 |
from degs norm |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
473 |
have th1: "degreen(?cnp *\<^sub>p c') n < degreen (CN 0\<^sub>p n (?cnp *\<^sub>p p')) n" by simp |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
474 |
hence neq: "degreen (?cnp *\<^sub>p c') n \<noteq> degreen (CN 0\<^sub>p n (?cnp *\<^sub>p p')) n" |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
475 |
by simp |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
476 |
have nmin: "n \<le> min n n" by (simp add: min_def) |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
477 |
from polyadd_different_degreen[OF norm(3,6) neq nmin] th1 |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
478 |
have deg: "degreen (CN c n p *\<^sub>p c' +\<^sub>p CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n = degreen (CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n" by simp |
41813 | 479 |
from "4.hyps"(16-18)[OF norm(1,4), of n] |
480 |
"4.hyps"(13-15)[OF norm(1,2), of n] |
|
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
481 |
mn norm m nn' deg |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
482 |
have ?eq by simp} |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
483 |
moreover |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
484 |
{assume mn: "m \<noteq> n" hence mn': "m < n" using m np by auto |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
485 |
from nn' m np have max1: "m \<le> max n n" by simp |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
486 |
hence min1: "m \<le> min n n" by simp |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
487 |
hence min2: "m \<le> min n (Suc n)" by simp |
41813 | 488 |
from "4.hyps"(16-18)[OF norm(1,4) min1] |
489 |
"4.hyps"(13-15)[OF norm(1,2) min2] |
|
41811 | 490 |
degreen_polyadd[OF norm(3,6) max1] |
33154 | 491 |
|
41811 | 492 |
have "degreen (?cnp *\<^sub>p c' +\<^sub>p CN 0\<^sub>p n (?cnp *\<^sub>p p')) m |
493 |
\<le> max (degreen (?cnp *\<^sub>p c') m) (degreen (CN 0\<^sub>p n (?cnp *\<^sub>p p')) m)" |
|
494 |
using mn nn' np np' by simp |
|
41813 | 495 |
with "4.hyps"(16-18)[OF norm(1,4) min1] |
496 |
"4.hyps"(13-15)[OF norm(1,2) min2] |
|
41811 | 497 |
degreen_0[OF norm(3) mn'] |
498 |
have ?eq using nn' mn np np' by clarsimp} |
|
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
499 |
ultimately have ?eq by blast} |
33154 | 500 |
ultimately show ?eq by blast} |
501 |
{ case (2 n0 n1) |
|
502 |
hence np: "isnpolyh ?cnp n0" and np': "isnpolyh ?cnp' n1" |
|
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
503 |
and m: "m \<le> min n0 n1" by simp_all |
33154 | 504 |
hence mn: "m \<le> n" by simp |
505 |
let ?c0p = "CN 0\<^sub>p n (?cnp *\<^sub>p p')" |
|
506 |
{assume C: "?cnp *\<^sub>p c' +\<^sub>p ?c0p = 0\<^sub>p" "n' = n" |
|
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
507 |
hence nn: "\<not>n' < n \<and> \<not> n<n'" by simp |
41813 | 508 |
from "4.hyps"(16-18) [of n n n] |
509 |
"4.hyps"(13-15)[of n "Suc n" n] |
|
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
510 |
np np' C(2) mn |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
511 |
have norm: "isnpolyh ?cnp n" "isnpolyh c' (Suc n)" "isnpolyh (?cnp *\<^sub>p c') n" |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
512 |
"isnpolyh p' n" "isnpolyh (?cnp *\<^sub>p p') n" "isnpolyh (CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n" |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
513 |
"(?cnp *\<^sub>p c' = 0\<^sub>p) = (c' = 0\<^sub>p)" |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
514 |
"?cnp *\<^sub>p p' \<noteq> 0\<^sub>p" |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
515 |
"degreen (?cnp *\<^sub>p c') n = (if c'=0\<^sub>p then 0 else degreen ?cnp n + degreen c' n)" |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
516 |
"degreen (?cnp *\<^sub>p p') n = degreen ?cnp n + degreen p' n" |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
517 |
by (simp_all add: min_def) |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
518 |
|
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
519 |
from norm have cn: "isnpolyh (CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n" by simp |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
520 |
have degneq: "degreen (?cnp *\<^sub>p c') n < degreen (CN 0\<^sub>p n (?cnp *\<^sub>p p')) n" |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
521 |
using norm by simp |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
522 |
from polyadd_eq_const_degreen[OF norm(3) cn C(1), where m="n"] degneq |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
523 |
have "False" by simp } |
33154 | 524 |
thus ?case using "4.hyps" by clarsimp} |
525 |
qed auto |
|
526 |
||
527 |
lemma polymul[simp]: "Ipoly bs (p *\<^sub>p q) = (Ipoly bs p) * (Ipoly bs q)" |
|
36349 | 528 |
by(induct p q rule: polymul.induct, auto simp add: field_simps) |
33154 | 529 |
|
530 |
lemma polymul_normh: |
|
36409 | 531 |
assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" |
33154 | 532 |
shows "\<lbrakk>isnpolyh p n0 ; isnpolyh q n1\<rbrakk> \<Longrightarrow> isnpolyh (p *\<^sub>p q) (min n0 n1)" |
533 |
using polymul_properties(1) by blast |
|
534 |
lemma polymul_eq0_iff: |
|
36409 | 535 |
assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" |
33154 | 536 |
shows "\<lbrakk> isnpolyh p n0 ; isnpolyh q n1\<rbrakk> \<Longrightarrow> (p *\<^sub>p q = 0\<^sub>p) = (p = 0\<^sub>p \<or> q = 0\<^sub>p) " |
537 |
using polymul_properties(2) by blast |
|
538 |
lemma polymul_degreen: |
|
36409 | 539 |
assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" |
33154 | 540 |
shows "\<lbrakk> isnpolyh p n0 ; isnpolyh q n1 ; m \<le> min n0 n1\<rbrakk> \<Longrightarrow> degreen (p *\<^sub>p q) m = (if (p = 0\<^sub>p \<or> q = 0\<^sub>p) then 0 else degreen p m + degreen q m)" |
541 |
using polymul_properties(3) by blast |
|
542 |
lemma polymul_norm: |
|
36409 | 543 |
assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" |
41813 | 544 |
shows "\<lbrakk> isnpoly p; isnpoly q\<rbrakk> \<Longrightarrow> isnpoly (polymul p q)" |
33154 | 545 |
using polymul_normh[of "p" "0" "q" "0"] isnpoly_def by simp |
546 |
||
547 |
lemma headconst_zero: "isnpolyh p n0 \<Longrightarrow> headconst p = 0\<^sub>N \<longleftrightarrow> p = 0\<^sub>p" |
|
548 |
by (induct p arbitrary: n0 rule: headconst.induct, auto) |
|
549 |
||
550 |
lemma headconst_isnormNum: "isnpolyh p n0 \<Longrightarrow> isnormNum (headconst p)" |
|
551 |
by (induct p arbitrary: n0, auto) |
|
552 |
||
553 |
lemma monic_eqI: assumes np: "isnpolyh p n0" |
|
36409 | 554 |
shows "INum (headconst p) * Ipoly bs (fst (monic p)) = (Ipoly bs p ::'a::{field_char_0, field_inverse_zero, power})" |
33154 | 555 |
unfolding monic_def Let_def |
556 |
proof(cases "headconst p = 0\<^sub>N", simp_all add: headconst_zero[OF np]) |
|
557 |
let ?h = "headconst p" |
|
558 |
assume pz: "p \<noteq> 0\<^sub>p" |
|
559 |
{assume hz: "INum ?h = (0::'a)" |
|
560 |
from headconst_isnormNum[OF np] have norm: "isnormNum ?h" "isnormNum 0\<^sub>N" by simp_all |
|
561 |
from isnormNum_unique[where ?'a = 'a, OF norm] hz have "?h = 0\<^sub>N" by simp |
|
562 |
with headconst_zero[OF np] have "p =0\<^sub>p" by blast with pz have "False" by blast} |
|
563 |
thus "INum (headconst p) = (0::'a) \<longrightarrow> \<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> = 0" by blast |
|
564 |
qed |
|
565 |
||
566 |
||
41404 | 567 |
text{* polyneg is a negation and preserves normal forms *} |
33154 | 568 |
|
569 |
lemma polyneg[simp]: "Ipoly bs (polyneg p) = - Ipoly bs p" |
|
570 |
by (induct p rule: polyneg.induct, auto) |
|
571 |
||
572 |
lemma polyneg0: "isnpolyh p n \<Longrightarrow> ((~\<^sub>p p) = 0\<^sub>p) = (p = 0\<^sub>p)" |
|
573 |
by (induct p arbitrary: n rule: polyneg.induct, auto simp add: Nneg_def) |
|
574 |
lemma polyneg_polyneg: "isnpolyh p n0 \<Longrightarrow> ~\<^sub>p (~\<^sub>p p) = p" |
|
575 |
by (induct p arbitrary: n0 rule: polyneg.induct, auto) |
|
576 |
lemma polyneg_normh: "\<And>n. isnpolyh p n \<Longrightarrow> isnpolyh (polyneg p) n " |
|
577 |
by (induct p rule: polyneg.induct, auto simp add: polyneg0) |
|
578 |
||
579 |
lemma polyneg_norm: "isnpoly p \<Longrightarrow> isnpoly (polyneg p)" |
|
580 |
using isnpoly_def polyneg_normh by simp |
|
581 |
||
582 |
||
41404 | 583 |
text{* polysub is a substraction and preserves normal forms *} |
584 |
||
41814 | 585 |
lemma polysub[simp]: "Ipoly bs (polysub p q) = (Ipoly bs p) - (Ipoly bs q)" |
33154 | 586 |
by (simp add: polysub_def polyneg polyadd) |
41814 | 587 |
lemma polysub_normh: "\<And> n0 n1. \<lbrakk> isnpolyh p n0 ; isnpolyh q n1\<rbrakk> \<Longrightarrow> isnpolyh (polysub p q) (min n0 n1)" |
33154 | 588 |
by (simp add: polysub_def polyneg_normh polyadd_normh) |
589 |
||
41814 | 590 |
lemma polysub_norm: "\<lbrakk> isnpoly p; isnpoly q\<rbrakk> \<Longrightarrow> isnpoly (polysub p q)" |
33154 | 591 |
using polyadd_norm polyneg_norm by (simp add: polysub_def) |
36409 | 592 |
lemma polysub_same_0[simp]: assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" |
41814 | 593 |
shows "isnpolyh p n0 \<Longrightarrow> polysub p p = 0\<^sub>p" |
33154 | 594 |
unfolding polysub_def split_def fst_conv snd_conv |
595 |
by (induct p arbitrary: n0,auto simp add: Let_def Nsub0[simplified Nsub_def]) |
|
596 |
||
597 |
lemma polysub_0: |
|
36409 | 598 |
assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" |
33154 | 599 |
shows "\<lbrakk> isnpolyh p n0 ; isnpolyh q n1\<rbrakk> \<Longrightarrow> (p -\<^sub>p q = 0\<^sub>p) = (p = q)" |
600 |
unfolding polysub_def split_def fst_conv snd_conv |
|
41763 | 601 |
by (induct p q arbitrary: n0 n1 rule:polyadd.induct) |
602 |
(auto simp: Nsub0[simplified Nsub_def] Let_def) |
|
33154 | 603 |
|
604 |
text{* polypow is a power function and preserves normal forms *} |
|
41404 | 605 |
|
36409 | 606 |
lemma polypow[simp]: "Ipoly bs (polypow n p) = ((Ipoly bs p :: 'a::{field_char_0, field_inverse_zero})) ^ n" |
33154 | 607 |
proof(induct n rule: polypow.induct) |
608 |
case 1 thus ?case by simp |
|
609 |
next |
|
610 |
case (2 n) |
|
611 |
let ?q = "polypow ((Suc n) div 2) p" |
|
41813 | 612 |
let ?d = "polymul ?q ?q" |
33154 | 613 |
have "odd (Suc n) \<or> even (Suc n)" by simp |
614 |
moreover |
|
615 |
{assume odd: "odd (Suc n)" |
|
616 |
have th: "(Suc (Suc (Suc (0\<Colon>nat)) * (Suc n div Suc (Suc (0\<Colon>nat))))) = Suc n div 2 + Suc n div 2 + 1" by arith |
|
41813 | 617 |
from odd have "Ipoly bs (p ^\<^sub>p Suc n) = Ipoly bs (polymul p ?d)" by (simp add: Let_def) |
33154 | 618 |
also have "\<dots> = (Ipoly bs p) * (Ipoly bs p)^(Suc n div 2)*(Ipoly bs p)^(Suc n div 2)" |
619 |
using "2.hyps" by simp |
|
620 |
also have "\<dots> = (Ipoly bs p) ^ (Suc n div 2 + Suc n div 2 + 1)" |
|
621 |
apply (simp only: power_add power_one_right) by simp |
|
622 |
also have "\<dots> = (Ipoly bs p) ^ (Suc (Suc (Suc (0\<Colon>nat)) * (Suc n div Suc (Suc (0\<Colon>nat)))))" |
|
623 |
by (simp only: th) |
|
624 |
finally have ?case |
|
625 |
using odd_nat_div_two_times_two_plus_one[OF odd, symmetric] by simp } |
|
626 |
moreover |
|
627 |
{assume even: "even (Suc n)" |
|
628 |
have th: "(Suc (Suc (0\<Colon>nat))) * (Suc n div Suc (Suc (0\<Colon>nat))) = Suc n div 2 + Suc n div 2" by arith |
|
629 |
from even have "Ipoly bs (p ^\<^sub>p Suc n) = Ipoly bs ?d" by (simp add: Let_def) |
|
630 |
also have "\<dots> = (Ipoly bs p) ^ (Suc n div 2 + Suc n div 2)" |
|
631 |
using "2.hyps" apply (simp only: power_add) by simp |
|
632 |
finally have ?case using even_nat_div_two_times_two[OF even] by (simp only: th)} |
|
633 |
ultimately show ?case by blast |
|
634 |
qed |
|
635 |
||
636 |
lemma polypow_normh: |
|
36409 | 637 |
assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" |
33154 | 638 |
shows "isnpolyh p n \<Longrightarrow> isnpolyh (polypow k p) n" |
639 |
proof (induct k arbitrary: n rule: polypow.induct) |
|
640 |
case (2 k n) |
|
641 |
let ?q = "polypow (Suc k div 2) p" |
|
41813 | 642 |
let ?d = "polymul ?q ?q" |
41807 | 643 |
from 2 have th1:"isnpolyh ?q n" and th2: "isnpolyh p n" by blast+ |
33154 | 644 |
from polymul_normh[OF th1 th1] have dn: "isnpolyh ?d n" by simp |
41813 | 645 |
from polymul_normh[OF th2 dn] have on: "isnpolyh (polymul p ?d) n" by simp |
33154 | 646 |
from dn on show ?case by (simp add: Let_def) |
647 |
qed auto |
|
648 |
||
649 |
lemma polypow_norm: |
|
36409 | 650 |
assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" |
33154 | 651 |
shows "isnpoly p \<Longrightarrow> isnpoly (polypow k p)" |
652 |
by (simp add: polypow_normh isnpoly_def) |
|
653 |
||
41404 | 654 |
text{* Finally the whole normalization *} |
33154 | 655 |
|
36409 | 656 |
lemma polynate[simp]: "Ipoly bs (polynate p) = (Ipoly bs p :: 'a ::{field_char_0, field_inverse_zero})" |
33154 | 657 |
by (induct p rule:polynate.induct, auto) |
658 |
||
659 |
lemma polynate_norm[simp]: |
|
36409 | 660 |
assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" |
33154 | 661 |
shows "isnpoly (polynate p)" |
662 |
by (induct p rule: polynate.induct, simp_all add: polyadd_norm polymul_norm polysub_norm polyneg_norm polypow_norm) (simp_all add: isnpoly_def) |
|
663 |
||
664 |
text{* shift1 *} |
|
665 |
||
666 |
||
667 |
lemma shift1: "Ipoly bs (shift1 p) = Ipoly bs (Mul (Bound 0) p)" |
|
668 |
by (simp add: shift1_def polymul) |
|
669 |
||
670 |
lemma shift1_isnpoly: |
|
671 |
assumes pn: "isnpoly p" and pnz: "p \<noteq> 0\<^sub>p" shows "isnpoly (shift1 p) " |
|
672 |
using pn pnz by (simp add: shift1_def isnpoly_def ) |
|
673 |
||
674 |
lemma shift1_nz[simp]:"shift1 p \<noteq> 0\<^sub>p" |
|
675 |
by (simp add: shift1_def) |
|
676 |
lemma funpow_shift1_isnpoly: |
|
677 |
"\<lbrakk> isnpoly p ; p \<noteq> 0\<^sub>p\<rbrakk> \<Longrightarrow> isnpoly (funpow n shift1 p)" |
|
39246 | 678 |
by (induct n arbitrary: p) (auto simp add: shift1_isnpoly funpow_swap1) |
33154 | 679 |
|
680 |
lemma funpow_isnpolyh: |
|
681 |
assumes f: "\<And> p. isnpolyh p n \<Longrightarrow> isnpolyh (f p) n "and np: "isnpolyh p n" |
|
682 |
shows "isnpolyh (funpow k f p) n" |
|
683 |
using f np by (induct k arbitrary: p, auto) |
|
684 |
||
36409 | 685 |
lemma funpow_shift1: "(Ipoly bs (funpow n shift1 p) :: 'a :: {field_char_0, field_inverse_zero}) = Ipoly bs (Mul (Pw (Bound 0) n) p)" |
33154 | 686 |
by (induct n arbitrary: p, simp_all add: shift1_isnpoly shift1 power_Suc ) |
687 |
||
688 |
lemma shift1_isnpolyh: "isnpolyh p n0 \<Longrightarrow> p\<noteq> 0\<^sub>p \<Longrightarrow> isnpolyh (shift1 p) 0" |
|
689 |
using isnpolyh_mono[where n="n0" and n'="0" and p="p"] by (simp add: shift1_def) |
|
690 |
||
691 |
lemma funpow_shift1_1: |
|
36409 | 692 |
"(Ipoly bs (funpow n shift1 p) :: 'a :: {field_char_0, field_inverse_zero}) = Ipoly bs (funpow n shift1 1\<^sub>p *\<^sub>p p)" |
33154 | 693 |
by (simp add: funpow_shift1) |
694 |
||
695 |
lemma poly_cmul[simp]: "Ipoly bs (poly_cmul c p) = Ipoly bs (Mul (C c) p)" |
|
45129
1fce03e3e8ad
tuned proofs -- eliminated vacuous "induct arbitrary: ..." situations;
wenzelm
parents:
41842
diff
changeset
|
696 |
by (induct p rule: poly_cmul.induct) (auto simp add: field_simps) |
33154 | 697 |
|
698 |
lemma behead: |
|
699 |
assumes np: "isnpolyh p n" |
|
36409 | 700 |
shows "Ipoly bs (Add (Mul (head p) (Pw (Bound 0) (degree p))) (behead p)) = (Ipoly bs p :: 'a :: {field_char_0, field_inverse_zero})" |
33154 | 701 |
using np |
702 |
proof (induct p arbitrary: n rule: behead.induct) |
|
703 |
case (1 c p n) hence pn: "isnpolyh p n" by simp |
|
41807 | 704 |
from 1(1)[OF pn] |
33154 | 705 |
have th:"Ipoly bs (Add (Mul (head p) (Pw (Bound 0) (degree p))) (behead p)) = Ipoly bs p" . |
706 |
then show ?case using "1.hyps" apply (simp add: Let_def,cases "behead p = 0\<^sub>p") |
|
36349 | 707 |
by (simp_all add: th[symmetric] field_simps power_Suc) |
33154 | 708 |
qed (auto simp add: Let_def) |
709 |
||
710 |
lemma behead_isnpolyh: |
|
711 |
assumes np: "isnpolyh p n" shows "isnpolyh (behead p) n" |
|
712 |
using np by (induct p rule: behead.induct, auto simp add: Let_def isnpolyh_mono) |
|
713 |
||
41404 | 714 |
subsection{* Miscellaneous lemmas about indexes, decrementation, substitution etc ... *} |
33154 | 715 |
lemma isnpolyh_polybound0: "isnpolyh p (Suc n) \<Longrightarrow> polybound0 p" |
39246 | 716 |
proof(induct p arbitrary: n rule: poly.induct, auto) |
33154 | 717 |
case (goal1 c n p n') |
718 |
hence "n = Suc (n - 1)" by simp |
|
719 |
hence "isnpolyh p (Suc (n - 1))" using `isnpolyh p n` by simp |
|
41807 | 720 |
with goal1(2) show ?case by simp |
33154 | 721 |
qed |
722 |
||
723 |
lemma isconstant_polybound0: "isnpolyh p n0 \<Longrightarrow> isconstant p \<longleftrightarrow> polybound0 p" |
|
724 |
by (induct p arbitrary: n0 rule: isconstant.induct, auto simp add: isnpolyh_polybound0) |
|
725 |
||
726 |
lemma decrpoly_zero[simp]: "decrpoly p = 0\<^sub>p \<longleftrightarrow> p = 0\<^sub>p" by (induct p, auto) |
|
727 |
||
728 |
lemma decrpoly_normh: "isnpolyh p n0 \<Longrightarrow> polybound0 p \<Longrightarrow> isnpolyh (decrpoly p) (n0 - 1)" |
|
729 |
apply (induct p arbitrary: n0, auto) |
|
730 |
apply (atomize) |
|
731 |
apply (erule_tac x = "Suc nat" in allE) |
|
732 |
apply auto |
|
733 |
done |
|
734 |
||
735 |
lemma head_polybound0: "isnpolyh p n0 \<Longrightarrow> polybound0 (head p)" |
|
736 |
by (induct p arbitrary: n0 rule: head.induct, auto intro: isnpolyh_polybound0) |
|
737 |
||
738 |
lemma polybound0_I: |
|
739 |
assumes nb: "polybound0 a" |
|
740 |
shows "Ipoly (b#bs) a = Ipoly (b'#bs) a" |
|
741 |
using nb |
|
39246 | 742 |
by (induct a rule: poly.induct) auto |
33154 | 743 |
lemma polysubst0_I: |
744 |
shows "Ipoly (b#bs) (polysubst0 a t) = Ipoly ((Ipoly (b#bs) a)#bs) t" |
|
745 |
by (induct t) simp_all |
|
746 |
||
747 |
lemma polysubst0_I': |
|
748 |
assumes nb: "polybound0 a" |
|
749 |
shows "Ipoly (b#bs) (polysubst0 a t) = Ipoly ((Ipoly (b'#bs) a)#bs) t" |
|
750 |
by (induct t) (simp_all add: polybound0_I[OF nb, where b="b" and b'="b'"]) |
|
751 |
||
752 |
lemma decrpoly: assumes nb: "polybound0 t" |
|
753 |
shows "Ipoly (x#bs) t = Ipoly bs (decrpoly t)" |
|
754 |
using nb by (induct t rule: decrpoly.induct, simp_all) |
|
755 |
||
756 |
lemma polysubst0_polybound0: assumes nb: "polybound0 t" |
|
757 |
shows "polybound0 (polysubst0 t a)" |
|
39246 | 758 |
using nb by (induct a rule: poly.induct, auto) |
33154 | 759 |
|
760 |
lemma degree0_polybound0: "isnpolyh p n \<Longrightarrow> degree p = 0 \<Longrightarrow> polybound0 p" |
|
761 |
by (induct p arbitrary: n rule: degree.induct, auto simp add: isnpolyh_polybound0) |
|
762 |
||
39246 | 763 |
primrec maxindex :: "poly \<Rightarrow> nat" where |
33154 | 764 |
"maxindex (Bound n) = n + 1" |
765 |
| "maxindex (CN c n p) = max (n + 1) (max (maxindex c) (maxindex p))" |
|
766 |
| "maxindex (Add p q) = max (maxindex p) (maxindex q)" |
|
767 |
| "maxindex (Sub p q) = max (maxindex p) (maxindex q)" |
|
768 |
| "maxindex (Mul p q) = max (maxindex p) (maxindex q)" |
|
769 |
| "maxindex (Neg p) = maxindex p" |
|
770 |
| "maxindex (Pw p n) = maxindex p" |
|
771 |
| "maxindex (C x) = 0" |
|
772 |
||
773 |
definition wf_bs :: "'a list \<Rightarrow> poly \<Rightarrow> bool" where |
|
774 |
"wf_bs bs p = (length bs \<ge> maxindex p)" |
|
775 |
||
776 |
lemma wf_bs_coefficients: "wf_bs bs p \<Longrightarrow> \<forall> c \<in> set (coefficients p). wf_bs bs c" |
|
777 |
proof(induct p rule: coefficients.induct) |
|
778 |
case (1 c p) |
|
779 |
show ?case |
|
780 |
proof |
|
781 |
fix x assume xc: "x \<in> set (coefficients (CN c 0 p))" |
|
782 |
hence "x = c \<or> x \<in> set (coefficients p)" by simp |
|
783 |
moreover |
|
784 |
{assume "x = c" hence "wf_bs bs x" using "1.prems" unfolding wf_bs_def by simp} |
|
785 |
moreover |
|
786 |
{assume H: "x \<in> set (coefficients p)" |
|
787 |
from "1.prems" have "wf_bs bs p" unfolding wf_bs_def by simp |
|
788 |
with "1.hyps" H have "wf_bs bs x" by blast } |
|
789 |
ultimately show "wf_bs bs x" by blast |
|
790 |
qed |
|
791 |
qed simp_all |
|
792 |
||
793 |
lemma maxindex_coefficients: " \<forall>c\<in> set (coefficients p). maxindex c \<le> maxindex p" |
|
794 |
by (induct p rule: coefficients.induct, auto) |
|
795 |
||
796 |
lemma wf_bs_I: "wf_bs bs p ==> Ipoly (bs@bs') p = Ipoly bs p" |
|
797 |
unfolding wf_bs_def by (induct p, auto simp add: nth_append) |
|
798 |
||
799 |
lemma take_maxindex_wf: assumes wf: "wf_bs bs p" |
|
800 |
shows "Ipoly (take (maxindex p) bs) p = Ipoly bs p" |
|
801 |
proof- |
|
802 |
let ?ip = "maxindex p" |
|
803 |
let ?tbs = "take ?ip bs" |
|
804 |
from wf have "length ?tbs = ?ip" unfolding wf_bs_def by simp |
|
805 |
hence wf': "wf_bs ?tbs p" unfolding wf_bs_def by simp |
|
806 |
have eq: "bs = ?tbs @ (drop ?ip bs)" by simp |
|
807 |
from wf_bs_I[OF wf', of "drop ?ip bs"] show ?thesis using eq by simp |
|
808 |
qed |
|
809 |
||
810 |
lemma decr_maxindex: "polybound0 p \<Longrightarrow> maxindex (decrpoly p) = maxindex p - 1" |
|
811 |
by (induct p, auto) |
|
812 |
||
813 |
lemma wf_bs_insensitive: "length bs = length bs' \<Longrightarrow> wf_bs bs p = wf_bs bs' p" |
|
814 |
unfolding wf_bs_def by simp |
|
815 |
||
816 |
lemma wf_bs_insensitive': "wf_bs (x#bs) p = wf_bs (y#bs) p" |
|
817 |
unfolding wf_bs_def by simp |
|
818 |
||
819 |
||
820 |
||
821 |
lemma wf_bs_coefficients': "\<forall>c \<in> set (coefficients p). wf_bs bs c \<Longrightarrow> wf_bs (x#bs) p" |
|
822 |
by(induct p rule: coefficients.induct, auto simp add: wf_bs_def) |
|
823 |
lemma coefficients_Nil[simp]: "coefficients p \<noteq> []" |
|
824 |
by (induct p rule: coefficients.induct, simp_all) |
|
825 |
||
826 |
||
827 |
lemma coefficients_head: "last (coefficients p) = head p" |
|
828 |
by (induct p rule: coefficients.induct, auto) |
|
829 |
||
830 |
lemma wf_bs_decrpoly: "wf_bs bs (decrpoly p) \<Longrightarrow> wf_bs (x#bs) p" |
|
831 |
unfolding wf_bs_def by (induct p rule: decrpoly.induct, auto) |
|
832 |
||
833 |
lemma length_le_list_ex: "length xs \<le> n \<Longrightarrow> \<exists> ys. length (xs @ ys) = n" |
|
834 |
apply (rule exI[where x="replicate (n - length xs) z"]) |
|
835 |
by simp |
|
836 |
lemma isnpolyh_Suc_const:"isnpolyh p (Suc n) \<Longrightarrow> isconstant p" |
|
837 |
by (cases p, auto) (case_tac "nat", simp_all) |
|
838 |
||
839 |
lemma wf_bs_polyadd: "wf_bs bs p \<and> wf_bs bs q \<longrightarrow> wf_bs bs (p +\<^sub>p q)" |
|
840 |
unfolding wf_bs_def |
|
841 |
apply (induct p q rule: polyadd.induct) |
|
842 |
apply (auto simp add: Let_def) |
|
843 |
done |
|
844 |
||
845 |
lemma wf_bs_polyul: "wf_bs bs p \<Longrightarrow> wf_bs bs q \<Longrightarrow> wf_bs bs (p *\<^sub>p q)" |
|
41811 | 846 |
unfolding wf_bs_def |
33154 | 847 |
apply (induct p q arbitrary: bs rule: polymul.induct) |
848 |
apply (simp_all add: wf_bs_polyadd) |
|
849 |
apply clarsimp |
|
850 |
apply (rule wf_bs_polyadd[unfolded wf_bs_def, rule_format]) |
|
851 |
apply auto |
|
852 |
done |
|
853 |
||
854 |
lemma wf_bs_polyneg: "wf_bs bs p \<Longrightarrow> wf_bs bs (~\<^sub>p p)" |
|
855 |
unfolding wf_bs_def by (induct p rule: polyneg.induct, auto) |
|
856 |
||
857 |
lemma wf_bs_polysub: "wf_bs bs p \<Longrightarrow> wf_bs bs q \<Longrightarrow> wf_bs bs (p -\<^sub>p q)" |
|
858 |
unfolding polysub_def split_def fst_conv snd_conv using wf_bs_polyadd wf_bs_polyneg by blast |
|
859 |
||
860 |
subsection{* Canonicity of polynomial representation, see lemma isnpolyh_unique*} |
|
861 |
||
862 |
definition "polypoly bs p = map (Ipoly bs) (coefficients p)" |
|
863 |
definition "polypoly' bs p = map ((Ipoly bs o decrpoly)) (coefficients p)" |
|
864 |
definition "poly_nate bs p = map ((Ipoly bs o decrpoly)) (coefficients (polynate p))" |
|
865 |
||
866 |
lemma coefficients_normh: "isnpolyh p n0 \<Longrightarrow> \<forall> q \<in> set (coefficients p). isnpolyh q n0" |
|
867 |
proof (induct p arbitrary: n0 rule: coefficients.induct) |
|
868 |
case (1 c p n0) |
|
869 |
have cp: "isnpolyh (CN c 0 p) n0" by fact |
|
870 |
hence norm: "isnpolyh c 0" "isnpolyh p 0" "p \<noteq> 0\<^sub>p" "n0 = 0" |
|
871 |
by (auto simp add: isnpolyh_mono[where n'=0]) |
|
872 |
from "1.hyps"[OF norm(2)] norm(1) norm(4) show ?case by simp |
|
873 |
qed auto |
|
874 |
||
875 |
lemma coefficients_isconst: |
|
876 |
"isnpolyh p n \<Longrightarrow> \<forall>q\<in>set (coefficients p). isconstant q" |
|
877 |
by (induct p arbitrary: n rule: coefficients.induct, |
|
878 |
auto simp add: isnpolyh_Suc_const) |
|
879 |
||
880 |
lemma polypoly_polypoly': |
|
881 |
assumes np: "isnpolyh p n0" |
|
882 |
shows "polypoly (x#bs) p = polypoly' bs p" |
|
883 |
proof- |
|
884 |
let ?cf = "set (coefficients p)" |
|
885 |
from coefficients_normh[OF np] have cn_norm: "\<forall> q\<in> ?cf. isnpolyh q n0" . |
|
886 |
{fix q assume q: "q \<in> ?cf" |
|
887 |
from q cn_norm have th: "isnpolyh q n0" by blast |
|
888 |
from coefficients_isconst[OF np] q have "isconstant q" by blast |
|
889 |
with isconstant_polybound0[OF th] have "polybound0 q" by blast} |
|
890 |
hence "\<forall>q \<in> ?cf. polybound0 q" .. |
|
891 |
hence "\<forall>q \<in> ?cf. Ipoly (x#bs) q = Ipoly bs (decrpoly q)" |
|
892 |
using polybound0_I[where b=x and bs=bs and b'=y] decrpoly[where x=x and bs=bs] |
|
893 |
by auto |
|
894 |
||
895 |
thus ?thesis unfolding polypoly_def polypoly'_def by simp |
|
896 |
qed |
|
897 |
||
898 |
lemma polypoly_poly: |
|
899 |
assumes np: "isnpolyh p n0" shows "Ipoly (x#bs) p = poly (polypoly (x#bs) p) x" |
|
900 |
using np |
|
901 |
by (induct p arbitrary: n0 bs rule: coefficients.induct, auto simp add: polypoly_def) |
|
902 |
||
903 |
lemma polypoly'_poly: |
|
904 |
assumes np: "isnpolyh p n0" shows "\<lparr>p\<rparr>\<^sub>p\<^bsup>x # bs\<^esup> = poly (polypoly' bs p) x" |
|
905 |
using polypoly_poly[OF np, simplified polypoly_polypoly'[OF np]] . |
|
906 |
||
907 |
||
908 |
lemma polypoly_poly_polybound0: |
|
909 |
assumes np: "isnpolyh p n0" and nb: "polybound0 p" |
|
910 |
shows "polypoly bs p = [Ipoly bs p]" |
|
911 |
using np nb unfolding polypoly_def |
|
912 |
by (cases p, auto, case_tac nat, auto) |
|
913 |
||
914 |
lemma head_isnpolyh: "isnpolyh p n0 \<Longrightarrow> isnpolyh (head p) n0" |
|
915 |
by (induct p rule: head.induct, auto) |
|
916 |
||
917 |
lemma headn_nz[simp]: "isnpolyh p n0 \<Longrightarrow> (headn p m = 0\<^sub>p) = (p = 0\<^sub>p)" |
|
918 |
by (cases p,auto) |
|
919 |
||
920 |
lemma head_eq_headn0: "head p = headn p 0" |
|
921 |
by (induct p rule: head.induct, simp_all) |
|
922 |
||
923 |
lemma head_nz[simp]: "isnpolyh p n0 \<Longrightarrow> (head p = 0\<^sub>p) = (p = 0\<^sub>p)" |
|
924 |
by (simp add: head_eq_headn0) |
|
925 |
||
926 |
lemma isnpolyh_zero_iff: |
|
36409 | 927 |
assumes nq: "isnpolyh p n0" and eq :"\<forall>bs. wf_bs bs p \<longrightarrow> \<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (0::'a::{field_char_0, field_inverse_zero, power})" |
33154 | 928 |
shows "p = 0\<^sub>p" |
929 |
using nq eq |
|
34915 | 930 |
proof (induct "maxindex p" arbitrary: p n0 rule: less_induct) |
931 |
case less |
|
932 |
note np = `isnpolyh p n0` and zp = `\<forall>bs. wf_bs bs p \<longrightarrow> \<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (0::'a)` |
|
933 |
{assume nz: "maxindex p = 0" |
|
934 |
then obtain c where "p = C c" using np by (cases p, auto) |
|
33154 | 935 |
with zp np have "p = 0\<^sub>p" unfolding wf_bs_def by simp} |
936 |
moreover |
|
34915 | 937 |
{assume nz: "maxindex p \<noteq> 0" |
33154 | 938 |
let ?h = "head p" |
939 |
let ?hd = "decrpoly ?h" |
|
940 |
let ?ihd = "maxindex ?hd" |
|
941 |
from head_isnpolyh[OF np] head_polybound0[OF np] have h:"isnpolyh ?h n0" "polybound0 ?h" |
|
942 |
by simp_all |
|
943 |
hence nhd: "isnpolyh ?hd (n0 - 1)" using decrpoly_normh by blast |
|
944 |
||
945 |
from maxindex_coefficients[of p] coefficients_head[of p, symmetric] |
|
34915 | 946 |
have mihn: "maxindex ?h \<le> maxindex p" by auto |
947 |
with decr_maxindex[OF h(2)] nz have ihd_lt_n: "?ihd < maxindex p" by auto |
|
33154 | 948 |
{fix bs:: "'a list" assume bs: "wf_bs bs ?hd" |
949 |
let ?ts = "take ?ihd bs" |
|
950 |
let ?rs = "drop ?ihd bs" |
|
951 |
have ts: "wf_bs ?ts ?hd" using bs unfolding wf_bs_def by simp |
|
952 |
have bs_ts_eq: "?ts@ ?rs = bs" by simp |
|
953 |
from wf_bs_decrpoly[OF ts] have tsh: " \<forall>x. wf_bs (x#?ts) ?h" by simp |
|
34915 | 954 |
from ihd_lt_n have "ALL x. length (x#?ts) \<le> maxindex p" by simp |
955 |
with length_le_list_ex obtain xs where xs:"length ((x#?ts) @ xs) = maxindex p" by blast |
|
956 |
hence "\<forall> x. wf_bs ((x#?ts) @ xs) p" unfolding wf_bs_def by simp |
|
33154 | 957 |
with zp have "\<forall> x. Ipoly ((x#?ts) @ xs) p = 0" by blast |
958 |
hence "\<forall> x. Ipoly (x#(?ts @ xs)) p = 0" by simp |
|
959 |
with polypoly_poly[OF np, where ?'a = 'a] polypoly_polypoly'[OF np, where ?'a = 'a] |
|
960 |
have "\<forall>x. poly (polypoly' (?ts @ xs) p) x = poly [] x" by simp |
|
961 |
hence "poly (polypoly' (?ts @ xs) p) = poly []" by (auto intro: ext) |
|
962 |
hence "\<forall> c \<in> set (coefficients p). Ipoly (?ts @ xs) (decrpoly c) = 0" |
|
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
963 |
using poly_zero[where ?'a='a] by (simp add: polypoly'_def list_all_iff) |
33154 | 964 |
with coefficients_head[of p, symmetric] |
965 |
have th0: "Ipoly (?ts @ xs) ?hd = 0" by simp |
|
966 |
from bs have wf'': "wf_bs ?ts ?hd" unfolding wf_bs_def by simp |
|
967 |
with th0 wf_bs_I[of ?ts ?hd xs] have "Ipoly ?ts ?hd = 0" by simp |
|
968 |
with wf'' wf_bs_I[of ?ts ?hd ?rs] bs_ts_eq have "\<lparr>?hd\<rparr>\<^sub>p\<^bsup>bs\<^esup> = 0" by simp } |
|
969 |
then have hdz: "\<forall>bs. wf_bs bs ?hd \<longrightarrow> \<lparr>?hd\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (0::'a)" by blast |
|
970 |
||
34915 | 971 |
from less(1)[OF ihd_lt_n nhd] hdz have "?hd = 0\<^sub>p" by blast |
33154 | 972 |
hence "?h = 0\<^sub>p" by simp |
973 |
with head_nz[OF np] have "p = 0\<^sub>p" by simp} |
|
974 |
ultimately show "p = 0\<^sub>p" by blast |
|
975 |
qed |
|
976 |
||
977 |
lemma isnpolyh_unique: |
|
978 |
assumes np:"isnpolyh p n0" and nq: "isnpolyh q n1" |
|
36409 | 979 |
shows "(\<forall>bs. \<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (\<lparr>q\<rparr>\<^sub>p\<^bsup>bs\<^esup> :: 'a::{field_char_0, field_inverse_zero, power})) \<longleftrightarrow> p = q" |
33154 | 980 |
proof(auto) |
981 |
assume H: "\<forall>bs. (\<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> ::'a)= \<lparr>q\<rparr>\<^sub>p\<^bsup>bs\<^esup>" |
|
982 |
hence "\<forall>bs.\<lparr>p -\<^sub>p q\<rparr>\<^sub>p\<^bsup>bs\<^esup>= (0::'a)" by simp |
|
983 |
hence "\<forall>bs. wf_bs bs (p -\<^sub>p q) \<longrightarrow> \<lparr>p -\<^sub>p q\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (0::'a)" |
|
984 |
using wf_bs_polysub[where p=p and q=q] by auto |
|
985 |
with isnpolyh_zero_iff[OF polysub_normh[OF np nq]] polysub_0[OF np nq] |
|
986 |
show "p = q" by blast |
|
987 |
qed |
|
988 |
||
989 |
||
41404 | 990 |
text{* consequences of unicity on the algorithms for polynomial normalization *} |
33154 | 991 |
|
36409 | 992 |
lemma polyadd_commute: assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" |
33154 | 993 |
and np: "isnpolyh p n0" and nq: "isnpolyh q n1" shows "p +\<^sub>p q = q +\<^sub>p p" |
994 |
using isnpolyh_unique[OF polyadd_normh[OF np nq] polyadd_normh[OF nq np]] by simp |
|
995 |
||
996 |
lemma zero_normh: "isnpolyh 0\<^sub>p n" by simp |
|
997 |
lemma one_normh: "isnpolyh 1\<^sub>p n" by simp |
|
998 |
lemma polyadd_0[simp]: |
|
36409 | 999 |
assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" |
33154 | 1000 |
and np: "isnpolyh p n0" shows "p +\<^sub>p 0\<^sub>p = p" and "0\<^sub>p +\<^sub>p p = p" |
1001 |
using isnpolyh_unique[OF polyadd_normh[OF np zero_normh] np] |
|
1002 |
isnpolyh_unique[OF polyadd_normh[OF zero_normh np] np] by simp_all |
|
1003 |
||
1004 |
lemma polymul_1[simp]: |
|
36409 | 1005 |
assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" |
33154 | 1006 |
and np: "isnpolyh p n0" shows "p *\<^sub>p 1\<^sub>p = p" and "1\<^sub>p *\<^sub>p p = p" |
1007 |
using isnpolyh_unique[OF polymul_normh[OF np one_normh] np] |
|
1008 |
isnpolyh_unique[OF polymul_normh[OF one_normh np] np] by simp_all |
|
1009 |
lemma polymul_0[simp]: |
|
36409 | 1010 |
assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" |
33154 | 1011 |
and np: "isnpolyh p n0" shows "p *\<^sub>p 0\<^sub>p = 0\<^sub>p" and "0\<^sub>p *\<^sub>p p = 0\<^sub>p" |
1012 |
using isnpolyh_unique[OF polymul_normh[OF np zero_normh] zero_normh] |
|
1013 |
isnpolyh_unique[OF polymul_normh[OF zero_normh np] zero_normh] by simp_all |
|
1014 |
||
1015 |
lemma polymul_commute: |
|
36409 | 1016 |
assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" |
33154 | 1017 |
and np:"isnpolyh p n0" and nq: "isnpolyh q n1" |
1018 |
shows "p *\<^sub>p q = q *\<^sub>p p" |
|
36409 | 1019 |
using isnpolyh_unique[OF polymul_normh[OF np nq] polymul_normh[OF nq np], where ?'a = "'a\<Colon>{field_char_0, field_inverse_zero, power}"] by simp |
33154 | 1020 |
|
1021 |
declare polyneg_polyneg[simp] |
|
1022 |
||
1023 |
lemma isnpolyh_polynate_id[simp]: |
|
36409 | 1024 |
assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" |
33154 | 1025 |
and np:"isnpolyh p n0" shows "polynate p = p" |
36409 | 1026 |
using isnpolyh_unique[where ?'a= "'a::{field_char_0, field_inverse_zero}", OF polynate_norm[of p, unfolded isnpoly_def] np] polynate[where ?'a = "'a::{field_char_0, field_inverse_zero}"] by simp |
33154 | 1027 |
|
1028 |
lemma polynate_idempotent[simp]: |
|
36409 | 1029 |
assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" |
33154 | 1030 |
shows "polynate (polynate p) = polynate p" |
1031 |
using isnpolyh_polynate_id[OF polynate_norm[of p, unfolded isnpoly_def]] . |
|
1032 |
||
1033 |
lemma poly_nate_polypoly': "poly_nate bs p = polypoly' bs (polynate p)" |
|
1034 |
unfolding poly_nate_def polypoly'_def .. |
|
36409 | 1035 |
lemma poly_nate_poly: shows "poly (poly_nate bs p) = (\<lambda>x:: 'a ::{field_char_0, field_inverse_zero}. \<lparr>p\<rparr>\<^sub>p\<^bsup>x # bs\<^esup>)" |
33154 | 1036 |
using polypoly'_poly[OF polynate_norm[unfolded isnpoly_def], symmetric, of bs p] |
1037 |
unfolding poly_nate_polypoly' by (auto intro: ext) |
|
1038 |
||
1039 |
subsection{* heads, degrees and all that *} |
|
1040 |
lemma degree_eq_degreen0: "degree p = degreen p 0" |
|
1041 |
by (induct p rule: degree.induct, simp_all) |
|
1042 |
||
1043 |
lemma degree_polyneg: assumes n: "isnpolyh p n" |
|
1044 |
shows "degree (polyneg p) = degree p" |
|
1045 |
using n |
|
1046 |
by (induct p arbitrary: n rule: polyneg.induct, simp_all) (case_tac na, auto) |
|
1047 |
||
1048 |
lemma degree_polyadd: |
|
1049 |
assumes np: "isnpolyh p n0" and nq: "isnpolyh q n1" |
|
1050 |
shows "degree (p +\<^sub>p q) \<le> max (degree p) (degree q)" |
|
1051 |
using degreen_polyadd[OF np nq, where m= "0"] degree_eq_degreen0 by simp |
|
1052 |
||
1053 |
||
1054 |
lemma degree_polysub: assumes np: "isnpolyh p n0" and nq: "isnpolyh q n1" |
|
1055 |
shows "degree (p -\<^sub>p q) \<le> max (degree p) (degree q)" |
|
1056 |
proof- |
|
1057 |
from nq have nq': "isnpolyh (~\<^sub>p q) n1" using polyneg_normh by simp |
|
1058 |
from degree_polyadd[OF np nq'] show ?thesis by (simp add: polysub_def degree_polyneg[OF nq]) |
|
1059 |
qed |
|
1060 |
||
1061 |
lemma degree_polysub_samehead: |
|
36409 | 1062 |
assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" |
33154 | 1063 |
and np: "isnpolyh p n0" and nq: "isnpolyh q n1" and h: "head p = head q" |
1064 |
and d: "degree p = degree q" |
|
1065 |
shows "degree (p -\<^sub>p q) < degree p \<or> (p -\<^sub>p q = 0\<^sub>p)" |
|
1066 |
unfolding polysub_def split_def fst_conv snd_conv |
|
1067 |
using np nq h d |
|
1068 |
proof(induct p q rule:polyadd.induct) |
|
41812 | 1069 |
case (1 c c') thus ?case by (simp add: Nsub_def Nsub0[simplified Nsub_def]) |
33154 | 1070 |
next |
41812 | 1071 |
case (2 c c' n' p') |
41815 | 1072 |
from 2 have "degree (C c) = degree (CN c' n' p')" by simp |
33154 | 1073 |
hence nz:"n' > 0" by (cases n', auto) |
1074 |
hence "head (CN c' n' p') = CN c' n' p'" by (cases n', auto) |
|
41807 | 1075 |
with 2 show ?case by simp |
33154 | 1076 |
next |
41812 | 1077 |
case (3 c n p c') |
41815 | 1078 |
hence "degree (C c') = degree (CN c n p)" by simp |
33154 | 1079 |
hence nz:"n > 0" by (cases n, auto) |
1080 |
hence "head (CN c n p) = CN c n p" by (cases n, auto) |
|
41807 | 1081 |
with 3 show ?case by simp |
33154 | 1082 |
next |
1083 |
case (4 c n p c' n' p') |
|
1084 |
hence H: "isnpolyh (CN c n p) n0" "isnpolyh (CN c' n' p') n1" |
|
1085 |
"head (CN c n p) = head (CN c' n' p')" "degree (CN c n p) = degree (CN c' n' p')" by simp+ |
|
1086 |
hence degc: "degree c = 0" and degc': "degree c' = 0" by simp_all |
|
1087 |
hence degnc: "degree (~\<^sub>p c) = 0" and degnc': "degree (~\<^sub>p c') = 0" |
|
1088 |
using H(1-2) degree_polyneg by auto |
|
1089 |
from H have cnh: "isnpolyh c (Suc n)" and c'nh: "isnpolyh c' (Suc n')" by simp+ |
|
1090 |
from degree_polysub[OF cnh c'nh, simplified polysub_def] degc degc' have degcmc': "degree (c +\<^sub>p ~\<^sub>pc') = 0" by simp |
|
1091 |
from H have pnh: "isnpolyh p n" and p'nh: "isnpolyh p' n'" by auto |
|
1092 |
have "n = n' \<or> n < n' \<or> n > n'" by arith |
|
1093 |
moreover |
|
1094 |
{assume nn': "n = n'" |
|
1095 |
have "n = 0 \<or> n >0" by arith |
|
41807 | 1096 |
moreover {assume nz: "n = 0" hence ?case using 4 nn' by (auto simp add: Let_def degcmc')} |
33154 | 1097 |
moreover {assume nz: "n > 0" |
1098 |
with nn' H(3) have cc':"c = c'" and pp': "p = p'" by (cases n, auto)+ |
|
41815 | 1099 |
hence ?case using polysub_same_0[OF p'nh, simplified polysub_def split_def fst_conv snd_conv] polysub_same_0[OF c'nh, simplified polysub_def] using nn' 4 by (simp add: Let_def)} |
33154 | 1100 |
ultimately have ?case by blast} |
1101 |
moreover |
|
1102 |
{assume nn': "n < n'" hence n'p: "n' > 0" by simp |
|
1103 |
hence headcnp':"head (CN c' n' p') = CN c' n' p'" by (cases n', simp_all) |
|
41815 | 1104 |
have degcnp': "degree (CN c' n' p') = 0" and degcnpeq: "degree (CN c n p) = degree (CN c' n' p')" using 4 nn' by (cases n', simp_all) |
33154 | 1105 |
hence "n > 0" by (cases n, simp_all) |
1106 |
hence headcnp: "head (CN c n p) = CN c n p" by (cases n, auto) |
|
1107 |
from H(3) headcnp headcnp' nn' have ?case by auto} |
|
1108 |
moreover |
|
1109 |
{assume nn': "n > n'" hence np: "n > 0" by simp |
|
1110 |
hence headcnp:"head (CN c n p) = CN c n p" by (cases n, simp_all) |
|
41807 | 1111 |
from 4 have degcnpeq: "degree (CN c' n' p') = degree (CN c n p)" by simp |
33154 | 1112 |
from np have degcnp: "degree (CN c n p) = 0" by (cases n, simp_all) |
1113 |
with degcnpeq have "n' > 0" by (cases n', simp_all) |
|
1114 |
hence headcnp': "head (CN c' n' p') = CN c' n' p'" by (cases n', auto) |
|
1115 |
from H(3) headcnp headcnp' nn' have ?case by auto} |
|
1116 |
ultimately show ?case by blast |
|
41812 | 1117 |
qed auto |
33154 | 1118 |
|
1119 |
lemma shift1_head : "isnpolyh p n0 \<Longrightarrow> head (shift1 p) = head p" |
|
1120 |
by (induct p arbitrary: n0 rule: head.induct, simp_all add: shift1_def) |
|
1121 |
||
1122 |
lemma funpow_shift1_head: "isnpolyh p n0 \<Longrightarrow> p\<noteq> 0\<^sub>p \<Longrightarrow> head (funpow k shift1 p) = head p" |
|
1123 |
proof(induct k arbitrary: n0 p) |
|
1124 |
case (Suc k n0 p) hence "isnpolyh (shift1 p) 0" by (simp add: shift1_isnpolyh) |
|
41807 | 1125 |
with Suc have "head (funpow k shift1 (shift1 p)) = head (shift1 p)" |
33154 | 1126 |
and "head (shift1 p) = head p" by (simp_all add: shift1_head) |
39246 | 1127 |
thus ?case by (simp add: funpow_swap1) |
33154 | 1128 |
qed auto |
1129 |
||
1130 |
lemma shift1_degree: "degree (shift1 p) = 1 + degree p" |
|
1131 |
by (simp add: shift1_def) |
|
1132 |
lemma funpow_shift1_degree: "degree (funpow k shift1 p) = k + degree p " |
|
1133 |
by (induct k arbitrary: p, auto simp add: shift1_degree) |
|
1134 |
||
1135 |
lemma funpow_shift1_nz: "p \<noteq> 0\<^sub>p \<Longrightarrow> funpow n shift1 p \<noteq> 0\<^sub>p" |
|
1136 |
by (induct n arbitrary: p, simp_all add: funpow_def) |
|
1137 |
||
1138 |
lemma head_isnpolyh_Suc[simp]: "isnpolyh p (Suc n) \<Longrightarrow> head p = p" |
|
1139 |
by (induct p arbitrary: n rule: degree.induct, auto) |
|
1140 |
lemma headn_0[simp]: "isnpolyh p n \<Longrightarrow> m < n \<Longrightarrow> headn p m = p" |
|
1141 |
by (induct p arbitrary: n rule: degreen.induct, auto) |
|
1142 |
lemma head_isnpolyh_Suc': "n > 0 \<Longrightarrow> isnpolyh p n \<Longrightarrow> head p = p" |
|
1143 |
by (induct p arbitrary: n rule: degree.induct, auto) |
|
1144 |
lemma head_head[simp]: "isnpolyh p n0 \<Longrightarrow> head (head p) = head p" |
|
1145 |
by (induct p rule: head.induct, auto) |
|
1146 |
||
1147 |
lemma polyadd_eq_const_degree: |
|
41812 | 1148 |
"\<lbrakk> isnpolyh p n0 ; isnpolyh q n1 ; polyadd p q = C c\<rbrakk> \<Longrightarrow> degree p = degree q" |
33154 | 1149 |
using polyadd_eq_const_degreen degree_eq_degreen0 by simp |
1150 |
||
1151 |
lemma polyadd_head: assumes np: "isnpolyh p n0" and nq: "isnpolyh q n1" |
|
1152 |
and deg: "degree p \<noteq> degree q" |
|
1153 |
shows "head (p +\<^sub>p q) = (if degree p < degree q then head q else head p)" |
|
1154 |
using np nq deg |
|
1155 |
apply(induct p q arbitrary: n0 n1 rule: polyadd.induct,simp_all) |
|
1156 |
apply (case_tac n', simp, simp) |
|
1157 |
apply (case_tac n, simp, simp) |
|
1158 |
apply (case_tac n, case_tac n', simp add: Let_def) |
|
1159 |
apply (case_tac "pa +\<^sub>p p' = 0\<^sub>p") |
|
41763 | 1160 |
apply (auto simp add: polyadd_eq_const_degree) |
1161 |
apply (metis head_nz) |
|
1162 |
apply (metis head_nz) |
|
1163 |
apply (metis degree.simps(9) gr0_conv_Suc head.simps(1) less_Suc0 not_less_eq) |
|
1164 |
by (metis degree.simps(9) gr0_conv_Suc nat_less_le order_le_less_trans) |
|
33154 | 1165 |
|
1166 |
lemma polymul_head_polyeq: |
|
36409 | 1167 |
assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" |
33154 | 1168 |
shows "\<lbrakk>isnpolyh p n0; isnpolyh q n1 ; p \<noteq> 0\<^sub>p ; q \<noteq> 0\<^sub>p \<rbrakk> \<Longrightarrow> head (p *\<^sub>p q) = head p *\<^sub>p head q" |
1169 |
proof (induct p q arbitrary: n0 n1 rule: polymul.induct) |
|
41813 | 1170 |
case (2 c c' n' p' n0 n1) |
1171 |
hence "isnpolyh (head (CN c' n' p')) n1" "isnormNum c" by (simp_all add: head_isnpolyh) |
|
41815 | 1172 |
thus ?case using 2 by (cases n', auto) |
33154 | 1173 |
next |
41813 | 1174 |
case (3 c n p c' n0 n1) |
1175 |
hence "isnpolyh (head (CN c n p)) n0" "isnormNum c'" by (simp_all add: head_isnpolyh) |
|
41815 | 1176 |
thus ?case using 3 by (cases n, auto) |
33154 | 1177 |
next |
1178 |
case (4 c n p c' n' p' n0 n1) |
|
1179 |
hence norm: "isnpolyh p n" "isnpolyh c (Suc n)" "isnpolyh p' n'" "isnpolyh c' (Suc n')" |
|
1180 |
"isnpolyh (CN c n p) n" "isnpolyh (CN c' n' p') n'" |
|
1181 |
by simp_all |
|
1182 |
have "n < n' \<or> n' < n \<or> n = n'" by arith |
|
1183 |
moreover |
|
1184 |
{assume nn': "n < n'" hence ?case |
|
1185 |
using norm |
|
41813 | 1186 |
"4.hyps"(2)[OF norm(1,6)] |
1187 |
"4.hyps"(1)[OF norm(2,6)] by (simp, cases n, simp,cases n', simp_all)} |
|
33154 | 1188 |
moreover {assume nn': "n'< n" |
41813 | 1189 |
hence ?case using norm "4.hyps"(6) [OF norm(5,3)] |
1190 |
"4.hyps"(5)[OF norm(5,4)] |
|
33154 | 1191 |
by (simp,cases n',simp,cases n,auto)} |
1192 |
moreover {assume nn': "n' = n" |
|
1193 |
from nn' polymul_normh[OF norm(5,4)] |
|
1194 |
have ncnpc':"isnpolyh (CN c n p *\<^sub>p c') n" by (simp add: min_def) |
|
1195 |
from nn' polymul_normh[OF norm(5,3)] norm |
|
1196 |
have ncnpp':"isnpolyh (CN c n p *\<^sub>p p') n" by simp |
|
1197 |
from nn' ncnpp' polymul_eq0_iff[OF norm(5,3)] norm(6) |
|
1198 |
have ncnpp0':"isnpolyh (CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n" by simp |
|
1199 |
from polyadd_normh[OF ncnpc' ncnpp0'] |
|
1200 |
have nth: "isnpolyh ((CN c n p *\<^sub>p c') +\<^sub>p (CN 0\<^sub>p n (CN c n p *\<^sub>p p'))) n" |
|
1201 |
by (simp add: min_def) |
|
1202 |
{assume np: "n > 0" |
|
1203 |
with nn' head_isnpolyh_Suc'[OF np nth] |
|
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1204 |
head_isnpolyh_Suc'[OF np norm(5)] head_isnpolyh_Suc'[OF np norm(6)[simplified nn']] |
33154 | 1205 |
have ?case by simp} |
1206 |
moreover |
|
1207 |
{moreover assume nz: "n = 0" |
|
1208 |
from polymul_degreen[OF norm(5,4), where m="0"] |
|
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1209 |
polymul_degreen[OF norm(5,3), where m="0"] nn' nz degree_eq_degreen0 |
33154 | 1210 |
norm(5,6) degree_npolyhCN[OF norm(6)] |
1211 |
have dth:"degree (CN c 0 p *\<^sub>p c') < degree (CN 0\<^sub>p 0 (CN c 0 p *\<^sub>p p'))" by simp |
|
1212 |
hence dth':"degree (CN c 0 p *\<^sub>p c') \<noteq> degree (CN 0\<^sub>p 0 (CN c 0 p *\<^sub>p p'))" by simp |
|
1213 |
from polyadd_head[OF ncnpc'[simplified nz] ncnpp0'[simplified nz] dth'] dth |
|
41813 | 1214 |
have ?case using norm "4.hyps"(6)[OF norm(5,3)] |
1215 |
"4.hyps"(5)[OF norm(5,4)] nn' nz by simp } |
|
33154 | 1216 |
ultimately have ?case by (cases n) auto} |
1217 |
ultimately show ?case by blast |
|
1218 |
qed simp_all |
|
1219 |
||
1220 |
lemma degree_coefficients: "degree p = length (coefficients p) - 1" |
|
1221 |
by(induct p rule: degree.induct, auto) |
|
1222 |
||
1223 |
lemma degree_head[simp]: "degree (head p) = 0" |
|
1224 |
by (induct p rule: head.induct, auto) |
|
1225 |
||
41812 | 1226 |
lemma degree_CN: "isnpolyh p n \<Longrightarrow> degree (CN c n p) \<le> 1 + degree p" |
33154 | 1227 |
by (cases n, simp_all) |
1228 |
lemma degree_CN': "isnpolyh p n \<Longrightarrow> degree (CN c n p) \<ge> degree p" |
|
1229 |
by (cases n, simp_all) |
|
1230 |
||
41812 | 1231 |
lemma polyadd_different_degree: "\<lbrakk>isnpolyh p n0 ; isnpolyh q n1; degree p \<noteq> degree q\<rbrakk> \<Longrightarrow> degree (polyadd p q) = max (degree p) (degree q)" |
33154 | 1232 |
using polyadd_different_degreen degree_eq_degreen0 by simp |
1233 |
||
1234 |
lemma degreen_polyneg: "isnpolyh p n0 \<Longrightarrow> degreen (~\<^sub>p p) m = degreen p m" |
|
1235 |
by (induct p arbitrary: n0 rule: polyneg.induct, auto) |
|
1236 |
||
1237 |
lemma degree_polymul: |
|
36409 | 1238 |
assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" |
33154 | 1239 |
and np: "isnpolyh p n0" and nq: "isnpolyh q n1" |
1240 |
shows "degree (p *\<^sub>p q) \<le> degree p + degree q" |
|
1241 |
using polymul_degreen[OF np nq, where m="0"] degree_eq_degreen0 by simp |
|
1242 |
||
1243 |
lemma polyneg_degree: "isnpolyh p n \<Longrightarrow> degree (polyneg p) = degree p" |
|
1244 |
by (induct p arbitrary: n rule: degree.induct, auto) |
|
1245 |
||
1246 |
lemma polyneg_head: "isnpolyh p n \<Longrightarrow> head(polyneg p) = polyneg (head p)" |
|
1247 |
by (induct p arbitrary: n rule: degree.induct, auto) |
|
1248 |
||
1249 |
subsection {* Correctness of polynomial pseudo division *} |
|
1250 |
||
1251 |
lemma polydivide_aux_properties: |
|
36409 | 1252 |
assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" |
33154 | 1253 |
and np: "isnpolyh p n0" and ns: "isnpolyh s n1" |
1254 |
and ap: "head p = a" and ndp: "degree p = n" and pnz: "p \<noteq> 0\<^sub>p" |
|
41403
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
1255 |
shows "(polydivide_aux a n p k s = (k',r) \<longrightarrow> (k' \<ge> k) \<and> (degree r = 0 \<or> degree r < degree p) |
33154 | 1256 |
\<and> (\<exists>nr. isnpolyh r nr) \<and> (\<exists>q n1. isnpolyh q n1 \<and> ((polypow (k' - k) a) *\<^sub>p s = p *\<^sub>p q +\<^sub>p r)))" |
1257 |
using ns |
|
34915 | 1258 |
proof(induct "degree s" arbitrary: s k k' r n1 rule: less_induct) |
1259 |
case less |
|
33154 | 1260 |
let ?qths = "\<exists>q n1. isnpolyh q n1 \<and> (a ^\<^sub>p (k' - k) *\<^sub>p s = p *\<^sub>p q +\<^sub>p r)" |
41403
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
1261 |
let ?ths = "polydivide_aux a n p k s = (k', r) \<longrightarrow> k \<le> k' \<and> (degree r = 0 \<or> degree r < degree p) |
33154 | 1262 |
\<and> (\<exists>nr. isnpolyh r nr) \<and> ?qths" |
1263 |
let ?b = "head s" |
|
34915 | 1264 |
let ?p' = "funpow (degree s - n) shift1 p" |
1265 |
let ?xdn = "funpow (degree s - n) shift1 1\<^sub>p" |
|
33154 | 1266 |
let ?akk' = "a ^\<^sub>p (k' - k)" |
34915 | 1267 |
note ns = `isnpolyh s n1` |
33154 | 1268 |
from np have np0: "isnpolyh p 0" |
1269 |
using isnpolyh_mono[where n="n0" and n'="0" and p="p"] by simp |
|
34915 | 1270 |
have np': "isnpolyh ?p' 0" using funpow_shift1_isnpoly[OF np0[simplified isnpoly_def[symmetric]] pnz, where n="degree s - n"] isnpoly_def by simp |
33154 | 1271 |
have headp': "head ?p' = head p" using funpow_shift1_head[OF np pnz] by simp |
1272 |
from funpow_shift1_isnpoly[where p="1\<^sub>p"] have nxdn: "isnpolyh ?xdn 0" by (simp add: isnpoly_def) |
|
1273 |
from polypow_normh [OF head_isnpolyh[OF np0], where k="k' - k"] ap |
|
1274 |
have nakk':"isnpolyh ?akk' 0" by blast |
|
1275 |
{assume sz: "s = 0\<^sub>p" |
|
41403
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
1276 |
hence ?ths using np polydivide_aux.simps apply clarsimp by (rule exI[where x="0\<^sub>p"], simp) } |
33154 | 1277 |
moreover |
1278 |
{assume sz: "s \<noteq> 0\<^sub>p" |
|
34915 | 1279 |
{assume dn: "degree s < n" |
41403
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
1280 |
hence "?ths" using ns ndp np polydivide_aux.simps by auto (rule exI[where x="0\<^sub>p"],simp) } |
33154 | 1281 |
moreover |
34915 | 1282 |
{assume dn': "\<not> degree s < n" hence dn: "degree s \<ge> n" by arith |
33154 | 1283 |
have degsp': "degree s = degree ?p'" |
34915 | 1284 |
using dn ndp funpow_shift1_degree[where k = "degree s - n" and p="p"] by simp |
33154 | 1285 |
{assume ba: "?b = a" |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1286 |
hence headsp': "head s = head ?p'" using ap headp' by simp |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1287 |
have nr: "isnpolyh (s -\<^sub>p ?p') 0" using polysub_normh[OF ns np'] by simp |
34915 | 1288 |
from degree_polysub_samehead[OF ns np' headsp' degsp'] |
1289 |
have "degree (s -\<^sub>p ?p') < degree s \<or> s -\<^sub>p ?p' = 0\<^sub>p" by simp |
|
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1290 |
moreover |
34915 | 1291 |
{assume deglt:"degree (s -\<^sub>p ?p') < degree s" |
41403
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
1292 |
from polydivide_aux.simps sz dn' ba |
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
1293 |
have eq: "polydivide_aux a n p k s = polydivide_aux a n p k (s -\<^sub>p ?p')" |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1294 |
by (simp add: Let_def) |
41403
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
1295 |
{assume h1: "polydivide_aux a n p k s = (k', r)" |
34915 | 1296 |
from less(1)[OF deglt nr, of k k' r] |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1297 |
trans[OF eq[symmetric] h1] |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1298 |
have kk': "k \<le> k'" and nr:"\<exists>nr. isnpolyh r nr" and dr: "degree r = 0 \<or> degree r < degree p" |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1299 |
and q1:"\<exists>q nq. isnpolyh q nq \<and> (a ^\<^sub>p k' - k *\<^sub>p (s -\<^sub>p ?p') = p *\<^sub>p q +\<^sub>p r)" by auto |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1300 |
from q1 obtain q n1 where nq: "isnpolyh q n1" |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1301 |
and asp:"a^\<^sub>p (k' - k) *\<^sub>p (s -\<^sub>p ?p') = p *\<^sub>p q +\<^sub>p r" by blast |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1302 |
from nr obtain nr where nr': "isnpolyh r nr" by blast |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1303 |
from polymul_normh[OF nakk' ns] have nakks': "isnpolyh (a ^\<^sub>p (k' - k) *\<^sub>p s) 0" by simp |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1304 |
from polyadd_normh[OF polymul_normh[OF nakk' nxdn] nq] |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1305 |
have nq': "isnpolyh (?akk' *\<^sub>p ?xdn +\<^sub>p q) 0" by simp |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1306 |
from polyadd_normh[OF polymul_normh[OF np |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1307 |
polyadd_normh[OF polymul_normh[OF nakk' nxdn] nq]] nr'] |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1308 |
have nqr': "isnpolyh (p *\<^sub>p (?akk' *\<^sub>p ?xdn +\<^sub>p q) +\<^sub>p r) 0" by simp |
36409 | 1309 |
from asp have "\<forall> (bs:: 'a::{field_char_0, field_inverse_zero} list). Ipoly bs (a^\<^sub>p (k' - k) *\<^sub>p (s -\<^sub>p ?p')) = |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1310 |
Ipoly bs (p *\<^sub>p q +\<^sub>p r)" by simp |
36409 | 1311 |
hence " \<forall>(bs:: 'a::{field_char_0, field_inverse_zero} list). Ipoly bs (a^\<^sub>p (k' - k)*\<^sub>p s) = |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1312 |
Ipoly bs (a^\<^sub>p (k' - k)) * Ipoly bs ?p' + Ipoly bs p * Ipoly bs q + Ipoly bs r" |
36349 | 1313 |
by (simp add: field_simps) |
36409 | 1314 |
hence " \<forall>(bs:: 'a::{field_char_0, field_inverse_zero} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = |
34915 | 1315 |
Ipoly bs (a^\<^sub>p (k' - k)) * Ipoly bs (funpow (degree s - n) shift1 1\<^sub>p *\<^sub>p p) |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1316 |
+ Ipoly bs p * Ipoly bs q + Ipoly bs r" |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1317 |
by (auto simp only: funpow_shift1_1) |
36409 | 1318 |
hence "\<forall>(bs:: 'a::{field_char_0, field_inverse_zero} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = |
34915 | 1319 |
Ipoly bs p * (Ipoly bs (a^\<^sub>p (k' - k)) * Ipoly bs (funpow (degree s - n) shift1 1\<^sub>p) |
36349 | 1320 |
+ Ipoly bs q) + Ipoly bs r" by (simp add: field_simps) |
36409 | 1321 |
hence "\<forall>(bs:: 'a::{field_char_0, field_inverse_zero} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = |
34915 | 1322 |
Ipoly bs (p *\<^sub>p ((a^\<^sub>p (k' - k)) *\<^sub>p (funpow (degree s - n) shift1 1\<^sub>p) +\<^sub>p q) +\<^sub>p r)" by simp |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1323 |
with isnpolyh_unique[OF nakks' nqr'] |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1324 |
have "a ^\<^sub>p (k' - k) *\<^sub>p s = |
34915 | 1325 |
p *\<^sub>p ((a^\<^sub>p (k' - k)) *\<^sub>p (funpow (degree s - n) shift1 1\<^sub>p) +\<^sub>p q) +\<^sub>p r" by blast |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1326 |
hence ?qths using nq' |
34915 | 1327 |
apply (rule_tac x="(a^\<^sub>p (k' - k)) *\<^sub>p (funpow (degree s - n) shift1 1\<^sub>p) +\<^sub>p q" in exI) |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1328 |
apply (rule_tac x="0" in exI) by simp |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1329 |
with kk' nr dr have "k \<le> k' \<and> (degree r = 0 \<or> degree r < degree p) \<and> (\<exists>nr. isnpolyh r nr) \<and> ?qths" |
41403
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
1330 |
by blast } hence ?ths by blast } |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1331 |
moreover |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1332 |
{assume spz:"s -\<^sub>p ?p' = 0\<^sub>p" |
36409 | 1333 |
from spz isnpolyh_unique[OF polysub_normh[OF ns np'], where q="0\<^sub>p", symmetric, where ?'a = "'a::{field_char_0, field_inverse_zero}"] |
1334 |
have " \<forall>(bs:: 'a::{field_char_0, field_inverse_zero} list). Ipoly bs s = Ipoly bs ?p'" by simp |
|
1335 |
hence "\<forall>(bs:: 'a::{field_char_0, field_inverse_zero} list). Ipoly bs s = Ipoly bs (?xdn *\<^sub>p p)" using np nxdn apply simp |
|
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1336 |
by (simp only: funpow_shift1_1) simp |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1337 |
hence sp': "s = ?xdn *\<^sub>p p" using isnpolyh_unique[OF ns polymul_normh[OF nxdn np]] by blast |
41403
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
1338 |
{assume h1: "polydivide_aux a n p k s = (k',r)" |
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
1339 |
from polydivide_aux.simps sz dn' ba |
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
1340 |
have eq: "polydivide_aux a n p k s = polydivide_aux a n p k (s -\<^sub>p ?p')" |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1341 |
by (simp add: Let_def) |
41403
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
1342 |
also have "\<dots> = (k,0\<^sub>p)" using polydivide_aux.simps spz by simp |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1343 |
finally have "(k',r) = (k,0\<^sub>p)" using h1 by simp |
34915 | 1344 |
with sp'[symmetric] ns np nxdn polyadd_0(1)[OF polymul_normh[OF np nxdn]] |
41403
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
1345 |
polyadd_0(2)[OF polymul_normh[OF np nxdn]] have ?ths |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1346 |
apply auto |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1347 |
apply (rule exI[where x="?xdn"]) |
34915 | 1348 |
apply (auto simp add: polymul_commute[of p]) |
41403
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
1349 |
done} } |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1350 |
ultimately have ?ths by blast } |
33154 | 1351 |
moreover |
1352 |
{assume ba: "?b \<noteq> a" |
|
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1353 |
from polysub_normh[OF polymul_normh[OF head_isnpolyh[OF np0, simplified ap] ns] |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1354 |
polymul_normh[OF head_isnpolyh[OF ns] np']] |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1355 |
have nth: "isnpolyh ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) 0" by(simp add: min_def) |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1356 |
have nzths: "a *\<^sub>p s \<noteq> 0\<^sub>p" "?b *\<^sub>p ?p' \<noteq> 0\<^sub>p" |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1357 |
using polymul_eq0_iff[OF head_isnpolyh[OF np0, simplified ap] ns] |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1358 |
polymul_eq0_iff[OF head_isnpolyh[OF ns] np']head_nz[OF np0] ap pnz sz head_nz[OF ns] |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1359 |
funpow_shift1_nz[OF pnz] by simp_all |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1360 |
from polymul_head_polyeq[OF head_isnpolyh[OF np] ns] head_nz[OF np] sz ap head_head[OF np] pnz |
34915 | 1361 |
polymul_head_polyeq[OF head_isnpolyh[OF ns] np'] head_nz [OF ns] sz funpow_shift1_nz[OF pnz, where n="degree s - n"] |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1362 |
have hdth: "head (a *\<^sub>p s) = head (?b *\<^sub>p ?p')" |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1363 |
using head_head[OF ns] funpow_shift1_head[OF np pnz] |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1364 |
polymul_commute[OF head_isnpolyh[OF np] head_isnpolyh[OF ns]] |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1365 |
by (simp add: ap) |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1366 |
from polymul_degreen[OF head_isnpolyh[OF np] ns, where m="0"] |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1367 |
head_nz[OF np] pnz sz ap[symmetric] |
34915 | 1368 |
funpow_shift1_nz[OF pnz, where n="degree s - n"] |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1369 |
polymul_degreen[OF head_isnpolyh[OF ns] np', where m="0"] head_nz[OF ns] |
34915 | 1370 |
ndp dn |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1371 |
have degth: "degree (a *\<^sub>p s) = degree (?b *\<^sub>p ?p') " |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1372 |
by (simp add: degree_eq_degreen0[symmetric] funpow_shift1_degree) |
34915 | 1373 |
{assume dth: "degree ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) < degree s" |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1374 |
from polysub_normh[OF polymul_normh[OF head_isnpolyh[OF np] ns] polymul_normh[OF head_isnpolyh[OF ns]np']] |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1375 |
ap have nasbp': "isnpolyh ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) 0" by simp |
41403
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
1376 |
{assume h1:"polydivide_aux a n p k s = (k', r)" |
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
1377 |
from h1 polydivide_aux.simps sz dn' ba |
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
1378 |
have eq:"polydivide_aux a n p (Suc k) ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) = (k',r)" |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1379 |
by (simp add: Let_def) |
34915 | 1380 |
with less(1)[OF dth nasbp', of "Suc k" k' r] |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1381 |
obtain q nq nr where kk': "Suc k \<le> k'" and nr: "isnpolyh r nr" and nq: "isnpolyh q nq" |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1382 |
and dr: "degree r = 0 \<or> degree r < degree p" |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1383 |
and qr: "a ^\<^sub>p (k' - Suc k) *\<^sub>p ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) = p *\<^sub>p q +\<^sub>p r" by auto |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1384 |
from kk' have kk'':"Suc (k' - Suc k) = k' - k" by arith |
36409 | 1385 |
{fix bs:: "'a::{field_char_0, field_inverse_zero} list" |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1386 |
|
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1387 |
from qr isnpolyh_unique[OF polypow_normh[OF head_isnpolyh[OF np], where k="k' - Suc k", simplified ap] nasbp', symmetric] |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1388 |
have "Ipoly bs (a ^\<^sub>p (k' - Suc k) *\<^sub>p ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p'))) = Ipoly bs (p *\<^sub>p q +\<^sub>p r)" by simp |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1389 |
hence "Ipoly bs a ^ (Suc (k' - Suc k)) * Ipoly bs s = Ipoly bs p * Ipoly bs q + Ipoly bs a ^ (k' - Suc k) * Ipoly bs ?b * Ipoly bs ?p' + Ipoly bs r" |
36349 | 1390 |
by (simp add: field_simps power_Suc) |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1391 |
hence "Ipoly bs a ^ (k' - k) * Ipoly bs s = Ipoly bs p * Ipoly bs q + Ipoly bs a ^ (k' - Suc k) * Ipoly bs ?b * Ipoly bs ?xdn * Ipoly bs p + Ipoly bs r" |
34915 | 1392 |
by (simp add:kk'' funpow_shift1_1[where n="degree s - n" and p="p"]) |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1393 |
hence "Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = Ipoly bs p * (Ipoly bs q + Ipoly bs a ^ (k' - Suc k) * Ipoly bs ?b * Ipoly bs ?xdn) + Ipoly bs r" |
36349 | 1394 |
by (simp add: field_simps)} |
36409 | 1395 |
hence ieq:"\<forall>(bs :: 'a::{field_char_0, field_inverse_zero} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1396 |
Ipoly bs (p *\<^sub>p (q +\<^sub>p (a ^\<^sub>p (k' - Suc k) *\<^sub>p ?b *\<^sub>p ?xdn)) +\<^sub>p r)" by auto |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1397 |
let ?q = "q +\<^sub>p (a ^\<^sub>p (k' - Suc k) *\<^sub>p ?b *\<^sub>p ?xdn)" |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1398 |
from polyadd_normh[OF nq polymul_normh[OF polymul_normh[OF polypow_normh[OF head_isnpolyh[OF np], where k="k' - Suc k"] head_isnpolyh[OF ns], simplified ap ] nxdn]] |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1399 |
have nqw: "isnpolyh ?q 0" by simp |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1400 |
from ieq isnpolyh_unique[OF polymul_normh[OF polypow_normh[OF head_isnpolyh[OF np], where k="k' - k"] ns, simplified ap] polyadd_normh[OF polymul_normh[OF np nqw] nr]] |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1401 |
have asth: "(a ^\<^sub>p (k' - k) *\<^sub>p s) = p *\<^sub>p (q +\<^sub>p (a ^\<^sub>p (k' - Suc k) *\<^sub>p ?b *\<^sub>p ?xdn)) +\<^sub>p r" by blast |
41403
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
1402 |
from dr kk' nr h1 asth nqw have ?ths apply simp |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1403 |
apply (rule conjI) |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1404 |
apply (rule exI[where x="nr"], simp) |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1405 |
apply (rule exI[where x="(q +\<^sub>p (a ^\<^sub>p (k' - Suc k) *\<^sub>p ?b *\<^sub>p ?xdn))"], simp) |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1406 |
apply (rule exI[where x="0"], simp) |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1407 |
done} |
41403
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
1408 |
hence ?ths by blast } |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1409 |
moreover |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1410 |
{assume spz: "a *\<^sub>p s -\<^sub>p (?b *\<^sub>p ?p') = 0\<^sub>p" |
36409 | 1411 |
{fix bs :: "'a::{field_char_0, field_inverse_zero} list" |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1412 |
from isnpolyh_unique[OF nth, where ?'a="'a" and q="0\<^sub>p",simplified,symmetric] spz |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1413 |
have "Ipoly bs (a*\<^sub>p s) = Ipoly bs ?b * Ipoly bs ?p'" by simp |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1414 |
hence "Ipoly bs (a*\<^sub>p s) = Ipoly bs (?b *\<^sub>p ?xdn) * Ipoly bs p" |
34915 | 1415 |
by (simp add: funpow_shift1_1[where n="degree s - n" and p="p"]) |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1416 |
hence "Ipoly bs (a*\<^sub>p s) = Ipoly bs (p *\<^sub>p (?b *\<^sub>p ?xdn))" by simp |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1417 |
} |
36409 | 1418 |
hence hth: "\<forall> (bs:: 'a::{field_char_0, field_inverse_zero} list). Ipoly bs (a*\<^sub>p s) = Ipoly bs (p *\<^sub>p (?b *\<^sub>p ?xdn))" .. |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1419 |
from hth |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1420 |
have asq: "a *\<^sub>p s = p *\<^sub>p (?b *\<^sub>p ?xdn)" |
36409 | 1421 |
using isnpolyh_unique[where ?'a = "'a::{field_char_0, field_inverse_zero}", OF polymul_normh[OF head_isnpolyh[OF np] ns] |
33154 | 1422 |
polymul_normh[OF np polymul_normh[OF head_isnpolyh[OF ns] nxdn]], |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1423 |
simplified ap] by simp |
41403
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
1424 |
{assume h1: "polydivide_aux a n p k s = (k', r)" |
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
1425 |
from h1 sz ba dn' spz polydivide_aux.simps polydivide_aux.simps |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1426 |
have "(k',r) = (Suc k, 0\<^sub>p)" by (simp add: Let_def) |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1427 |
with h1 np head_isnpolyh[OF np, simplified ap] ns polymul_normh[OF head_isnpolyh[OF ns] nxdn] |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1428 |
polymul_normh[OF np polymul_normh[OF head_isnpolyh[OF ns] nxdn]] asq |
41403
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
1429 |
have ?ths apply (clarsimp simp add: Let_def) |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1430 |
apply (rule exI[where x="?b *\<^sub>p ?xdn"]) apply simp |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1431 |
apply (rule exI[where x="0"], simp) |
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1432 |
done} |
41403
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
1433 |
hence ?ths by blast} |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1434 |
ultimately have ?ths using degree_polysub_samehead[OF polymul_normh[OF head_isnpolyh[OF np0, simplified ap] ns] polymul_normh[OF head_isnpolyh[OF ns] np'] hdth degth] polymul_degreen[OF head_isnpolyh[OF np] ns, where m="0"] |
34915 | 1435 |
head_nz[OF np] pnz sz ap[symmetric] |
33268
02de0317f66f
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
33154
diff
changeset
|
1436 |
by (simp add: degree_eq_degreen0[symmetric]) blast } |
33154 | 1437 |
ultimately have ?ths by blast |
1438 |
} |
|
1439 |
ultimately have ?ths by blast} |
|
1440 |
ultimately show ?ths by blast |
|
1441 |
qed |
|
1442 |
||
1443 |
lemma polydivide_properties: |
|
36409 | 1444 |
assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" |
33154 | 1445 |
and np: "isnpolyh p n0" and ns: "isnpolyh s n1" and pnz: "p \<noteq> 0\<^sub>p" |
1446 |
shows "(\<exists> k r. polydivide s p = (k,r) \<and> (\<exists>nr. isnpolyh r nr) \<and> (degree r = 0 \<or> degree r < degree p) |
|
1447 |
\<and> (\<exists>q n1. isnpolyh q n1 \<and> ((polypow k (head p)) *\<^sub>p s = p *\<^sub>p q +\<^sub>p r)))" |
|
1448 |
proof- |
|
1449 |
have trv: "head p = head p" "degree p = degree p" by simp_all |
|
41403
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
1450 |
from polydivide_def[where s="s" and p="p"] |
33154 | 1451 |
have ex: "\<exists> k r. polydivide s p = (k,r)" by auto |
1452 |
then obtain k r where kr: "polydivide s p = (k,r)" by blast |
|
41403
7eba049f7310
partial_function (tailrec) replaces function (tailrec);
krauss
parents:
39246
diff
changeset
|
1453 |
from trans[OF meta_eq_to_obj_eq[OF polydivide_def[where s="s"and p="p"], symmetric] kr] |
33154 | 1454 |
polydivide_aux_properties[OF np ns trv pnz, where k="0" and k'="k" and r="r"] |
1455 |
have "(degree r = 0 \<or> degree r < degree p) \<and> |
|
1456 |
(\<exists>nr. isnpolyh r nr) \<and> (\<exists>q n1. isnpolyh q n1 \<and> head p ^\<^sub>p k - 0 *\<^sub>p s = p *\<^sub>p q +\<^sub>p r)" by blast |
|
1457 |
with kr show ?thesis |
|
1458 |
apply - |
|
1459 |
apply (rule exI[where x="k"]) |
|
1460 |
apply (rule exI[where x="r"]) |
|
1461 |
apply simp |
|
1462 |
done |
|
1463 |
qed |
|
1464 |
||
1465 |
subsection{* More about polypoly and pnormal etc *} |
|
1466 |
||
1467 |
definition "isnonconstant p = (\<not> isconstant p)" |
|
1468 |
||
1469 |
lemma isnonconstant_pnormal_iff: assumes nc: "isnonconstant p" |
|
1470 |
shows "pnormal (polypoly bs p) \<longleftrightarrow> Ipoly bs (head p) \<noteq> 0" |
|
1471 |
proof |
|
1472 |
let ?p = "polypoly bs p" |
|
1473 |
assume H: "pnormal ?p" |
|
1474 |
have csz: "coefficients p \<noteq> []" using nc by (cases p, auto) |
|
1475 |
||
1476 |
from coefficients_head[of p] last_map[OF csz, of "Ipoly bs"] |
|
1477 |
pnormal_last_nonzero[OF H] |
|
1478 |
show "Ipoly bs (head p) \<noteq> 0" by (simp add: polypoly_def) |
|
1479 |
next |
|
1480 |
assume h: "\<lparr>head p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<noteq> 0" |
|
1481 |
let ?p = "polypoly bs p" |
|
1482 |
have csz: "coefficients p \<noteq> []" using nc by (cases p, auto) |
|
1483 |
hence pz: "?p \<noteq> []" by (simp add: polypoly_def) |
|
1484 |
hence lg: "length ?p > 0" by simp |
|
1485 |
from h coefficients_head[of p] last_map[OF csz, of "Ipoly bs"] |
|
1486 |
have lz: "last ?p \<noteq> 0" by (simp add: polypoly_def) |
|
1487 |
from pnormal_last_length[OF lg lz] show "pnormal ?p" . |
|
1488 |
qed |
|
1489 |
||
1490 |
lemma isnonconstant_coefficients_length: "isnonconstant p \<Longrightarrow> length (coefficients p) > 1" |
|
1491 |
unfolding isnonconstant_def |
|
1492 |
apply (cases p, simp_all) |
|
1493 |
apply (case_tac nat, auto) |
|
1494 |
done |
|
1495 |
lemma isnonconstant_nonconstant: assumes inc: "isnonconstant p" |
|
1496 |
shows "nonconstant (polypoly bs p) \<longleftrightarrow> Ipoly bs (head p) \<noteq> 0" |
|
1497 |
proof |
|
1498 |
let ?p = "polypoly bs p" |
|
1499 |
assume nc: "nonconstant ?p" |
|
1500 |
from isnonconstant_pnormal_iff[OF inc, of bs] nc |
|
1501 |
show "\<lparr>head p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<noteq> 0" unfolding nonconstant_def by blast |
|
1502 |
next |
|
1503 |
let ?p = "polypoly bs p" |
|
1504 |
assume h: "\<lparr>head p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<noteq> 0" |
|
1505 |
from isnonconstant_pnormal_iff[OF inc, of bs] h |
|
1506 |
have pn: "pnormal ?p" by blast |
|
1507 |
{fix x assume H: "?p = [x]" |
|
1508 |
from H have "length (coefficients p) = 1" unfolding polypoly_def by auto |
|
1509 |
with isnonconstant_coefficients_length[OF inc] have False by arith} |
|
1510 |
thus "nonconstant ?p" using pn unfolding nonconstant_def by blast |
|
1511 |
qed |
|
1512 |
||
1513 |
lemma pnormal_length: "p\<noteq>[] \<Longrightarrow> pnormal p \<longleftrightarrow> length (pnormalize p) = length p" |
|
1514 |
unfolding pnormal_def |
|
39246 | 1515 |
apply (induct p) |
1516 |
apply (simp_all, case_tac "p=[]", simp_all) |
|
33154 | 1517 |
done |
1518 |
||
1519 |
lemma degree_degree: assumes inc: "isnonconstant p" |
|
1520 |
shows "degree p = Polynomial_List.degree (polypoly bs p) \<longleftrightarrow> \<lparr>head p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<noteq> 0" |
|
1521 |
proof |
|
1522 |
let ?p = "polypoly bs p" |
|
1523 |
assume H: "degree p = Polynomial_List.degree ?p" |
|
1524 |
from isnonconstant_coefficients_length[OF inc] have pz: "?p \<noteq> []" |
|
1525 |
unfolding polypoly_def by auto |
|
1526 |
from H degree_coefficients[of p] isnonconstant_coefficients_length[OF inc] |
|
1527 |
have lg:"length (pnormalize ?p) = length ?p" |
|
1528 |
unfolding Polynomial_List.degree_def polypoly_def by simp |
|
1529 |
hence "pnormal ?p" using pnormal_length[OF pz] by blast |
|
1530 |
with isnonconstant_pnormal_iff[OF inc] |
|
1531 |
show "\<lparr>head p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<noteq> 0" by blast |
|
1532 |
next |
|
1533 |
let ?p = "polypoly bs p" |
|
1534 |
assume H: "\<lparr>head p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<noteq> 0" |
|
1535 |
with isnonconstant_pnormal_iff[OF inc] have "pnormal ?p" by blast |
|
1536 |
with degree_coefficients[of p] isnonconstant_coefficients_length[OF inc] |
|
1537 |
show "degree p = Polynomial_List.degree ?p" |
|
1538 |
unfolding polypoly_def pnormal_def Polynomial_List.degree_def by auto |
|
1539 |
qed |
|
1540 |
||
1541 |
section{* Swaps ; Division by a certain variable *} |
|
39246 | 1542 |
primrec swap:: "nat \<Rightarrow> nat \<Rightarrow> poly \<Rightarrow> poly" where |
33154 | 1543 |
"swap n m (C x) = C x" |
39246 | 1544 |
| "swap n m (Bound k) = Bound (if k = n then m else if k=m then n else k)" |
1545 |
| "swap n m (Neg t) = Neg (swap n m t)" |
|
1546 |
| "swap n m (Add s t) = Add (swap n m s) (swap n m t)" |
|
1547 |
| "swap n m (Sub s t) = Sub (swap n m s) (swap n m t)" |
|
1548 |
| "swap n m (Mul s t) = Mul (swap n m s) (swap n m t)" |
|
1549 |
| "swap n m (Pw t k) = Pw (swap n m t) k" |
|
1550 |
| "swap n m (CN c k p) = CN (swap n m c) (if k = n then m else if k=m then n else k) |
|
33154 | 1551 |
(swap n m p)" |
1552 |
||
1553 |
lemma swap: assumes nbs: "n < length bs" and mbs: "m < length bs" |
|
1554 |
shows "Ipoly bs (swap n m t) = Ipoly ((bs[n:= bs!m])[m:= bs!n]) t" |
|
1555 |
proof (induct t) |
|
1556 |
case (Bound k) thus ?case using nbs mbs by simp |
|
1557 |
next |
|
1558 |
case (CN c k p) thus ?case using nbs mbs by simp |
|
1559 |
qed simp_all |
|
1560 |
lemma swap_swap_id[simp]: "swap n m (swap m n t) = t" |
|
1561 |
by (induct t,simp_all) |
|
1562 |
||
1563 |
lemma swap_commute: "swap n m p = swap m n p" by (induct p, simp_all) |
|
1564 |
||
1565 |
lemma swap_same_id[simp]: "swap n n t = t" |
|
1566 |
by (induct t, simp_all) |
|
1567 |
||
1568 |
definition "swapnorm n m t = polynate (swap n m t)" |
|
1569 |
||
1570 |
lemma swapnorm: assumes nbs: "n < length bs" and mbs: "m < length bs" |
|
36409 | 1571 |
shows "((Ipoly bs (swapnorm n m t) :: 'a\<Colon>{field_char_0, field_inverse_zero})) = Ipoly ((bs[n:= bs!m])[m:= bs!n]) t" |
41807 | 1572 |
using swap[OF assms] swapnorm_def by simp |
33154 | 1573 |
|
1574 |
lemma swapnorm_isnpoly[simp]: |
|
36409 | 1575 |
assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" |
33154 | 1576 |
shows "isnpoly (swapnorm n m p)" |
1577 |
unfolding swapnorm_def by simp |
|
1578 |
||
1579 |
definition "polydivideby n s p = |
|
1580 |
(let ss = swapnorm 0 n s ; sp = swapnorm 0 n p ; h = head sp; (k,r) = polydivide ss sp |
|
1581 |
in (k,swapnorm 0 n h,swapnorm 0 n r))" |
|
1582 |
||
1583 |
lemma swap_nz [simp]: " (swap n m p = 0\<^sub>p) = (p = 0\<^sub>p)" by (induct p, simp_all) |
|
1584 |
||
41808 | 1585 |
fun isweaknpoly :: "poly \<Rightarrow> bool" |
1586 |
where |
|
33154 | 1587 |
"isweaknpoly (C c) = True" |
41808 | 1588 |
| "isweaknpoly (CN c n p) \<longleftrightarrow> isweaknpoly c \<and> isweaknpoly p" |
1589 |
| "isweaknpoly p = False" |
|
33154 | 1590 |
|
1591 |
lemma isnpolyh_isweaknpoly: "isnpolyh p n0 \<Longrightarrow> isweaknpoly p" |
|
1592 |
by (induct p arbitrary: n0, auto) |
|
1593 |
||
1594 |
lemma swap_isweanpoly: "isweaknpoly p \<Longrightarrow> isweaknpoly (swap n m p)" |
|
1595 |
by (induct p, auto) |
|
1596 |
||
1597 |
end |