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