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