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