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