| author | wenzelm | 
| Wed, 25 Jun 2008 22:11:17 +0200 | |
| changeset 27364 | a8672b0e2b15 | 
| parent 26316 | 9e9e67e33557 | 
| permissions | -rw-r--r-- | 
| 12224 | 1 | (* Title: Poly.thy | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 2 | ID: $Id$ | 
| 12224 | 3 | Author: Jacques D. Fleuriot | 
| 4 | Copyright: 2000 University of Edinburgh | |
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 5 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 6 | Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4 | 
| 12224 | 7 | *) | 
| 8 | ||
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 9 | header{*Univariate Real Polynomials*}
 | 
| 12224 | 10 | |
| 15131 | 11 | theory Poly | 
| 23256 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 12 | imports Deriv | 
| 15131 | 13 | begin | 
| 12224 | 14 | |
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 15 | text{*Application of polynomial as a real function.*}
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 16 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 17 | consts poly :: "real list => real => real" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 18 | primrec | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 19 | poly_Nil: "poly [] x = 0" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 20 | poly_Cons: "poly (h#t) x = h + x * poly t x" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 21 | |
| 12224 | 22 | |
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 23 | subsection{*Arithmetic Operations on Polynomials*}
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 24 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 25 | text{*addition*}
 | 
| 22808 | 26 | consts padd :: "[real list, real list] => real list" (infixl "+++" 65) | 
| 12224 | 27 | primrec | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 28 | padd_Nil: "[] +++ l2 = l2" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 29 | padd_Cons: "(h#t) +++ l2 = (if l2 = [] then h#t | 
| 12224 | 30 | else (h + hd l2)#(t +++ tl l2))" | 
| 31 | ||
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 32 | text{*Multiplication by a constant*}
 | 
| 22808 | 33 | consts cmult :: "[real, real list] => real list" (infixl "%*" 70) | 
| 12224 | 34 | primrec | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 35 | cmult_Nil: "c %* [] = []" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 36 | cmult_Cons: "c %* (h#t) = (c * h)#(c %* t)" | 
| 12224 | 37 | |
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 38 | text{*Multiplication by a polynomial*}
 | 
| 22808 | 39 | consts pmult :: "[real list, real list] => real list" (infixl "***" 70) | 
| 12224 | 40 | primrec | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 41 | pmult_Nil: "[] *** l2 = []" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 42 | pmult_Cons: "(h#t) *** l2 = (if t = [] then h %* l2 | 
| 12224 | 43 | else (h %* l2) +++ ((0) # (t *** l2)))" | 
| 44 | ||
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 45 | text{*Repeated multiplication by a polynomial*}
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 46 | consts mulexp :: "[nat, real list, real list] => real list" | 
| 12224 | 47 | primrec | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 48 | mulexp_zero: "mulexp 0 p q = q" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 49 | mulexp_Suc: "mulexp (Suc n) p q = p *** mulexp n p q" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 50 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 51 | text{*Exponential*}
 | 
| 22808 | 52 | consts pexp :: "[real list, nat] => real list" (infixl "%^" 80) | 
| 12224 | 53 | primrec | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 54 | pexp_0: "p %^ 0 = [1]" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 55 | pexp_Suc: "p %^ (Suc n) = p *** (p %^ n)" | 
| 12224 | 56 | |
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 57 | text{*Quotient related value of dividing a polynomial by x + a*}
 | 
| 12224 | 58 | (* Useful for divisor properties in inductive proofs *) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 59 | consts "pquot" :: "[real list, real] => real list" | 
| 12224 | 60 | primrec | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 61 | pquot_Nil: "pquot [] a= []" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 62 | pquot_Cons: "pquot (h#t) a = (if t = [] then [h] | 
| 12224 | 63 | else (inverse(a) * (h - hd( pquot t a)))#(pquot t a))" | 
| 64 | ||
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 65 | text{*Differentiation of polynomials (needs an auxiliary function).*}
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 66 | consts pderiv_aux :: "nat => real list => real list" | 
| 12224 | 67 | primrec | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 68 | pderiv_aux_Nil: "pderiv_aux n [] = []" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 69 | pderiv_aux_Cons: "pderiv_aux n (h#t) = | 
| 12224 | 70 | (real n * h)#(pderiv_aux (Suc n) t)" | 
| 71 | ||
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 72 | text{*normalization of polynomials (remove extra 0 coeff)*}
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 73 | consts pnormalize :: "real list => real list" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 74 | primrec | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 75 | pnormalize_Nil: "pnormalize [] = []" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 76 | pnormalize_Cons: "pnormalize (h#p) = (if ( (pnormalize p) = []) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 77 | then (if (h = 0) then [] else [h]) | 
| 12224 | 78 | else (h#(pnormalize p)))" | 
| 79 | ||
| 23256 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 80 | definition "pnormal p = ((pnormalize p = p) \<and> p \<noteq> [])" | 
| 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 81 | definition "nonconstant p = (pnormal p \<and> (\<forall>x. p \<noteq> [x]))" | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 82 | text{*Other definitions*}
 | 
| 12224 | 83 | |
| 19765 | 84 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20898diff
changeset | 85 |   poly_minus :: "real list => real list"      ("-- _" [80] 80) where
 | 
| 19765 | 86 | "-- p = (- 1) %* p" | 
| 12224 | 87 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20898diff
changeset | 88 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20898diff
changeset | 89 | pderiv :: "real list => real list" where | 
| 19765 | 90 | "pderiv p = (if p = [] then [] else pderiv_aux 1 (tl p))" | 
| 12224 | 91 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20898diff
changeset | 92 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20898diff
changeset | 93 | divides :: "[real list,real list] => bool" (infixl "divides" 70) where | 
| 19765 | 94 | "p1 divides p2 = (\<exists>q. poly p2 = poly(p1 *** q))" | 
| 12224 | 95 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20898diff
changeset | 96 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20898diff
changeset | 97 | order :: "real => real list => nat" where | 
| 19765 | 98 |     --{*order of a polynomial*}
 | 
| 99 | "order a p = (SOME n. ([-a, 1] %^ n) divides p & | |
| 12224 | 100 | ~ (([-a, 1] %^ (Suc n)) divides p))" | 
| 101 | ||
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20898diff
changeset | 102 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20898diff
changeset | 103 | degree :: "real list => nat" where | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 104 |      --{*degree of a polynomial*}
 | 
| 23256 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 105 | "degree p = length (pnormalize p) - 1" | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 106 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20898diff
changeset | 107 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20898diff
changeset | 108 | rsquarefree :: "real list => bool" where | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 109 |      --{*squarefree polynomials --- NB with respect to real roots only.*}
 | 
| 19765 | 110 | "rsquarefree p = (poly p \<noteq> poly [] & | 
| 111 | (\<forall>a. (order a p = 0) | (order a p = 1)))" | |
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 112 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 113 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 114 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 115 | lemma padd_Nil2: "p +++ [] = p" | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20898diff
changeset | 116 | by (induct p) auto | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 117 | declare padd_Nil2 [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 118 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 119 | lemma padd_Cons_Cons: "(h1 # p1) +++ (h2 # p2) = (h1 + h2) # (p1 +++ p2)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 120 | by auto | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 121 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 122 | lemma pminus_Nil: "-- [] = []" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 123 | by (simp add: poly_minus_def) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 124 | declare pminus_Nil [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 125 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 126 | lemma pmult_singleton: "[h1] *** p1 = h1 %* p1" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 127 | by simp | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 128 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 129 | lemma poly_ident_mult: "1 %* t = t" | 
| 15251 | 130 | by (induct "t", auto) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 131 | declare poly_ident_mult [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 132 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 133 | lemma poly_simple_add_Cons: "[a] +++ ((0)#t) = (a#t)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 134 | by simp | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 135 | declare poly_simple_add_Cons [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 136 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 137 | text{*Handy general properties*}
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 138 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 139 | lemma padd_commut: "b +++ a = a +++ b" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 140 | apply (subgoal_tac "\<forall>a. b +++ a = a +++ b") | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 141 | apply (induct_tac [2] "b", auto) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 142 | apply (rule padd_Cons [THEN ssubst]) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 143 | apply (case_tac "aa", auto) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 144 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 145 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 146 | lemma padd_assoc [rule_format]: "\<forall>b c. (a +++ b) +++ c = a +++ (b +++ c)" | 
| 15251 | 147 | apply (induct "a", simp, clarify) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 148 | apply (case_tac b, simp_all) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 149 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 150 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 151 | lemma poly_cmult_distr [rule_format]: | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 152 | "\<forall>q. a %* ( p +++ q) = (a %* p +++ a %* q)" | 
| 15251 | 153 | apply (induct "p", simp, clarify) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 154 | apply (case_tac "q") | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 155 | apply (simp_all add: right_distrib) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 156 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 157 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 158 | lemma pmult_by_x: "[0, 1] *** t = ((0)#t)" | 
| 15251 | 159 | apply (induct "t", simp) | 
| 160 | apply (auto simp add: poly_ident_mult padd_commut) | |
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 161 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 162 | declare pmult_by_x [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 163 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 164 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 165 | text{*properties of evaluation of polynomials.*}
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 166 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 167 | lemma poly_add: "poly (p1 +++ p2) x = poly p1 x + poly p2 x" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 168 | apply (subgoal_tac "\<forall>p2. poly (p1 +++ p2) x = poly (p1) x + poly (p2) x") | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 169 | apply (induct_tac [2] "p1", auto) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 170 | apply (case_tac "p2") | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 171 | apply (auto simp add: right_distrib) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 172 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 173 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 174 | lemma poly_cmult: "poly (c %* p) x = c * poly p x" | 
| 15251 | 175 | apply (induct "p") | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 176 | apply (case_tac [2] "x=0") | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 177 | apply (auto simp add: right_distrib mult_ac) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 178 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 179 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 180 | lemma poly_minus: "poly (-- p) x = - (poly p x)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 181 | apply (simp add: poly_minus_def) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 182 | apply (auto simp add: poly_cmult) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 183 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 184 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 185 | lemma poly_mult: "poly (p1 *** p2) x = poly p1 x * poly p2 x" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 186 | apply (subgoal_tac "\<forall>p2. poly (p1 *** p2) x = poly p1 x * poly p2 x") | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 187 | apply (simp (no_asm_simp)) | 
| 15251 | 188 | apply (induct "p1") | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 189 | apply (auto simp add: poly_cmult) | 
| 15251 | 190 | apply (case_tac p1) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 191 | apply (auto simp add: poly_cmult poly_add left_distrib right_distrib mult_ac) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 192 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 193 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 194 | lemma poly_exp: "poly (p %^ n) x = (poly p x) ^ n" | 
| 15251 | 195 | apply (induct "n") | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 196 | apply (auto simp add: poly_cmult poly_mult) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 197 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 198 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 199 | text{*More Polynomial Evaluation Lemmas*}
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 200 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 201 | lemma poly_add_rzero: "poly (a +++ []) x = poly a x" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 202 | by simp | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 203 | declare poly_add_rzero [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 204 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 205 | lemma poly_mult_assoc: "poly ((a *** b) *** c) x = poly (a *** (b *** c)) x" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 206 | by (simp add: poly_mult real_mult_assoc) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 207 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 208 | lemma poly_mult_Nil2: "poly (p *** []) x = 0" | 
| 15251 | 209 | by (induct "p", auto) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 210 | declare poly_mult_Nil2 [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 211 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 212 | lemma poly_exp_add: "poly (p %^ (n + d)) x = poly( p %^ n *** p %^ d) x" | 
| 15251 | 213 | apply (induct "n") | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 214 | apply (auto simp add: poly_mult real_mult_assoc) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 215 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 216 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 217 | text{*The derivative*}
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 218 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 219 | lemma pderiv_Nil: "pderiv [] = []" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 220 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 221 | apply (simp add: pderiv_def) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 222 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 223 | declare pderiv_Nil [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 224 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 225 | lemma pderiv_singleton: "pderiv [c] = []" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 226 | by (simp add: pderiv_def) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 227 | declare pderiv_singleton [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 228 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 229 | lemma pderiv_Cons: "pderiv (h#t) = pderiv_aux 1 t" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 230 | by (simp add: pderiv_def) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 231 | |
| 20792 | 232 | lemma DERIV_cmult2: "DERIV f x :> D ==> DERIV (%x. (f x) * c :: real) x :> D * c" | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 233 | by (simp add: DERIV_cmult mult_commute [of _ c]) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 234 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 235 | lemma DERIV_pow2: "DERIV (%x. x ^ Suc n) x :> real (Suc n) * (x ^ n)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 236 | by (rule lemma_DERIV_subst, rule DERIV_pow, simp) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 237 | declare DERIV_pow2 [simp] DERIV_pow [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 238 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 239 | lemma lemma_DERIV_poly1: "\<forall>n. DERIV (%x. (x ^ (Suc n) * poly p x)) x :> | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 240 | x ^ n * poly (pderiv_aux (Suc n) p) x " | 
| 15251 | 241 | apply (induct "p") | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 242 | apply (auto intro!: DERIV_add DERIV_cmult2 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 243 | simp add: pderiv_def right_distrib real_mult_assoc [symmetric] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 244 | simp del: realpow_Suc) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 245 | apply (subst mult_commute) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 246 | apply (simp del: realpow_Suc) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 247 | apply (simp add: mult_commute realpow_Suc [symmetric] del: realpow_Suc) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 248 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 249 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 250 | lemma lemma_DERIV_poly: "DERIV (%x. (x ^ (Suc n) * poly p x)) x :> | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 251 | x ^ n * poly (pderiv_aux (Suc n) p) x " | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 252 | by (simp add: lemma_DERIV_poly1 del: realpow_Suc) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 253 | |
| 20792 | 254 | lemma DERIV_add_const: "DERIV f x :> D ==> DERIV (%x. a + f x :: real) x :> D" | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 255 | by (rule lemma_DERIV_subst, rule DERIV_add, auto) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 256 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 257 | lemma poly_DERIV: "DERIV (%x. poly p x) x :> poly (pderiv p) x" | 
| 15251 | 258 | apply (induct "p") | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 259 | apply (auto simp add: pderiv_Cons) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 260 | apply (rule DERIV_add_const) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 261 | apply (rule lemma_DERIV_subst) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 262 | apply (rule lemma_DERIV_poly [where n=0, simplified], simp) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 263 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 264 | declare poly_DERIV [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 265 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 266 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 267 | text{* Consequences of the derivative theorem above*}
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 268 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 269 | lemma poly_differentiable: "(%x. poly p x) differentiable x" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 270 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 271 | apply (simp add: differentiable_def) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 272 | apply (blast intro: poly_DERIV) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 273 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 274 | declare poly_differentiable [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 275 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 276 | lemma poly_isCont: "isCont (%x. poly p x) x" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 277 | by (rule poly_DERIV [THEN DERIV_isCont]) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 278 | declare poly_isCont [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 279 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 280 | lemma poly_IVT_pos: "[| a < b; poly p a < 0; 0 < poly p b |] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 281 | ==> \<exists>x. a < x & x < b & (poly p x = 0)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 282 | apply (cut_tac f = "%x. poly p x" and a = a and b = b and y = 0 in IVT_objl) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 283 | apply (auto simp add: order_le_less) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 284 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 285 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 286 | lemma poly_IVT_neg: "[| a < b; 0 < poly p a; poly p b < 0 |] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 287 | ==> \<exists>x. a < x & x < b & (poly p x = 0)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 288 | apply (insert poly_IVT_pos [where p = "-- p" ]) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 289 | apply (simp add: poly_minus neg_less_0_iff_less) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 290 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 291 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 292 | lemma poly_MVT: "a < b ==> | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 293 | \<exists>x. a < x & x < b & (poly p b - poly p a = (b - a) * poly (pderiv p) x)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 294 | apply (drule_tac f = "poly p" in MVT, auto) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 295 | apply (rule_tac x = z in exI) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 296 | apply (auto simp add: real_mult_left_cancel poly_DERIV [THEN DERIV_unique]) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 297 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 298 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 299 | text{*Lemmas for Derivatives*}
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 300 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 301 | lemma lemma_poly_pderiv_aux_add: "\<forall>p2 n. poly (pderiv_aux n (p1 +++ p2)) x = | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 302 | poly (pderiv_aux n p1 +++ pderiv_aux n p2) x" | 
| 15251 | 303 | apply (induct "p1", simp, clarify) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 304 | apply (case_tac "p2") | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 305 | apply (auto simp add: right_distrib) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 306 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 307 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 308 | lemma poly_pderiv_aux_add: "poly (pderiv_aux n (p1 +++ p2)) x = | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 309 | poly (pderiv_aux n p1 +++ pderiv_aux n p2) x" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 310 | apply (simp add: lemma_poly_pderiv_aux_add) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 311 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 312 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 313 | lemma lemma_poly_pderiv_aux_cmult: "\<forall>n. poly (pderiv_aux n (c %* p)) x = poly (c %* pderiv_aux n p) x" | 
| 15251 | 314 | apply (induct "p") | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 315 | apply (auto simp add: poly_cmult mult_ac) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 316 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 317 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 318 | lemma poly_pderiv_aux_cmult: "poly (pderiv_aux n (c %* p)) x = poly (c %* pderiv_aux n p) x" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 319 | by (simp add: lemma_poly_pderiv_aux_cmult) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 320 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 321 | lemma poly_pderiv_aux_minus: | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 322 | "poly (pderiv_aux n (-- p)) x = poly (-- pderiv_aux n p) x" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 323 | apply (simp add: poly_minus_def poly_pderiv_aux_cmult) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 324 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 325 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 326 | lemma lemma_poly_pderiv_aux_mult1: "\<forall>n. poly (pderiv_aux (Suc n) p) x = poly ((pderiv_aux n p) +++ p) x" | 
| 15251 | 327 | apply (induct "p") | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 328 | apply (auto simp add: real_of_nat_Suc left_distrib) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 329 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 330 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 331 | lemma lemma_poly_pderiv_aux_mult: "poly (pderiv_aux (Suc n) p) x = poly ((pderiv_aux n p) +++ p) x" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 332 | by (simp add: lemma_poly_pderiv_aux_mult1) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 333 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 334 | lemma lemma_poly_pderiv_add: "\<forall>q. poly (pderiv (p +++ q)) x = poly (pderiv p +++ pderiv q) x" | 
| 15251 | 335 | apply (induct "p", simp, clarify) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 336 | apply (case_tac "q") | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 337 | apply (auto simp add: poly_pderiv_aux_add poly_add pderiv_def) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 338 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 339 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 340 | lemma poly_pderiv_add: "poly (pderiv (p +++ q)) x = poly (pderiv p +++ pderiv q) x" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 341 | by (simp add: lemma_poly_pderiv_add) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 342 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 343 | lemma poly_pderiv_cmult: "poly (pderiv (c %* p)) x = poly (c %* (pderiv p)) x" | 
| 15251 | 344 | apply (induct "p") | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 345 | apply (auto simp add: poly_pderiv_aux_cmult poly_cmult pderiv_def) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 346 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 347 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 348 | lemma poly_pderiv_minus: "poly (pderiv (--p)) x = poly (--(pderiv p)) x" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 349 | by (simp add: poly_minus_def poly_pderiv_cmult) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 350 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 351 | lemma lemma_poly_mult_pderiv: | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 352 | "poly (pderiv (h#t)) x = poly ((0 # (pderiv t)) +++ t) x" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 353 | apply (simp add: pderiv_def) | 
| 15251 | 354 | apply (induct "t") | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 355 | apply (auto simp add: poly_add lemma_poly_pderiv_aux_mult) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 356 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 357 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 358 | lemma poly_pderiv_mult: "\<forall>q. poly (pderiv (p *** q)) x = | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 359 | poly (p *** (pderiv q) +++ q *** (pderiv p)) x" | 
| 15251 | 360 | apply (induct "p") | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 361 | apply (auto simp add: poly_add poly_cmult poly_pderiv_cmult poly_pderiv_add poly_mult) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 362 | apply (rule lemma_poly_mult_pderiv [THEN ssubst]) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 363 | apply (rule lemma_poly_mult_pderiv [THEN ssubst]) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 364 | apply (rule poly_add [THEN ssubst]) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 365 | apply (rule poly_add [THEN ssubst]) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 366 | apply (simp (no_asm_simp) add: poly_mult right_distrib add_ac mult_ac) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 367 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 368 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 369 | lemma poly_pderiv_exp: "poly (pderiv (p %^ (Suc n))) x = | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 370 | poly ((real (Suc n)) %* (p %^ n) *** pderiv p) x" | 
| 15251 | 371 | apply (induct "n") | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 372 | apply (auto simp add: poly_add poly_pderiv_cmult poly_cmult poly_pderiv_mult | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 373 | real_of_nat_zero poly_mult real_of_nat_Suc | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 374 | right_distrib left_distrib mult_ac) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 375 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 376 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 377 | lemma poly_pderiv_exp_prime: "poly (pderiv ([-a, 1] %^ (Suc n))) x = | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 378 | poly (real (Suc n) %* ([-a, 1] %^ n)) x" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 379 | apply (simp add: poly_pderiv_exp poly_mult del: pexp_Suc) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 380 | apply (simp add: poly_cmult pderiv_def) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 381 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 382 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 383 | subsection{*Key Property: if @{term "f(a) = 0"} then @{term "(x - a)"} divides
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 384 |  @{term "p(x)"} *}
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 385 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 386 | lemma lemma_poly_linear_rem: "\<forall>h. \<exists>q r. h#t = [r] +++ [-a, 1] *** q" | 
| 15251 | 387 | apply (induct "t", safe) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 388 | apply (rule_tac x = "[]" in exI) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 389 | apply (rule_tac x = h in exI, simp) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 390 | apply (drule_tac x = aa in spec, safe) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 391 | apply (rule_tac x = "r#q" in exI) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 392 | apply (rule_tac x = "a*r + h" in exI) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 393 | apply (case_tac "q", auto) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 394 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 395 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 396 | lemma poly_linear_rem: "\<exists>q r. h#t = [r] +++ [-a, 1] *** q" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 397 | by (cut_tac t = t and a = a in lemma_poly_linear_rem, auto) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 398 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 399 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 400 | lemma poly_linear_divides: "(poly p a = 0) = ((p = []) | (\<exists>q. p = [-a, 1] *** q))" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 401 | apply (auto simp add: poly_add poly_cmult right_distrib) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 402 | apply (case_tac "p", simp) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 403 | apply (cut_tac h = aa and t = list and a = a in poly_linear_rem, safe) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 404 | apply (case_tac "q", auto) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 405 | apply (drule_tac x = "[]" in spec, simp) | 
| 23286 | 406 | apply (auto simp add: poly_add poly_cmult add_assoc) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 407 | apply (drule_tac x = "aa#lista" in spec, auto) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 408 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 409 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 410 | lemma lemma_poly_length_mult: "\<forall>h k a. length (k %* p +++ (h # (a %* p))) = Suc (length p)" | 
| 15251 | 411 | by (induct "p", auto) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 412 | declare lemma_poly_length_mult [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 413 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 414 | lemma lemma_poly_length_mult2: "\<forall>h k. length (k %* p +++ (h # p)) = Suc (length p)" | 
| 15251 | 415 | by (induct "p", auto) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 416 | declare lemma_poly_length_mult2 [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 417 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 418 | lemma poly_length_mult: "length([-a,1] *** q) = Suc (length q)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 419 | by auto | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 420 | declare poly_length_mult [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 421 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 422 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 423 | subsection{*Polynomial length*}
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 424 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 425 | lemma poly_cmult_length: "length (a %* p) = length p" | 
| 15251 | 426 | by (induct "p", auto) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 427 | declare poly_cmult_length [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 428 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 429 | lemma poly_add_length [rule_format]: | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 430 | "\<forall>p2. length (p1 +++ p2) = | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 431 | (if (length p1 < length p2) then length p2 else length p1)" | 
| 20256 
5024ba0831a6
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
20217diff
changeset | 432 | apply (induct "p1", simp_all) | 
| 20432 
07ec57376051
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20256diff
changeset | 433 | apply arith | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 434 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 435 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 436 | lemma poly_root_mult_length: "length([a,b] *** p) = Suc (length p)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 437 | by (simp add: poly_cmult_length poly_add_length) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 438 | declare poly_root_mult_length [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 439 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 440 | lemma poly_mult_not_eq_poly_Nil: "(poly (p *** q) x \<noteq> poly [] x) = | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 441 | (poly p x \<noteq> poly [] x & poly q x \<noteq> poly [] x)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 442 | apply (auto simp add: poly_mult) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 443 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 444 | declare poly_mult_not_eq_poly_Nil [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 445 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 446 | lemma poly_mult_eq_zero_disj: "(poly (p *** q) x = 0) = (poly p x = 0 | poly q x = 0)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 447 | by (auto simp add: poly_mult) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 448 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 449 | text{*Normalisation Properties*}
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 450 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 451 | lemma poly_normalized_nil: "(pnormalize p = []) --> (poly p x = 0)" | 
| 15251 | 452 | by (induct "p", auto) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 453 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 454 | text{*A nontrivial polynomial of degree n has no more than n roots*}
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 455 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 456 | lemma poly_roots_index_lemma [rule_format]: | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 457 | "\<forall>p x. poly p x \<noteq> poly [] x & length p = n | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 458 | --> (\<exists>i. \<forall>x. (poly p x = (0::real)) --> (\<exists>m. (m \<le> n & x = i m)))" | 
| 15251 | 459 | apply (induct "n", safe) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 460 | apply (rule ccontr) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 461 | apply (subgoal_tac "\<exists>a. poly p a = 0", safe) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 462 | apply (drule poly_linear_divides [THEN iffD1], safe) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 463 | apply (drule_tac x = q in spec) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 464 | apply (drule_tac x = x in spec) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 465 | apply (simp del: poly_Nil pmult_Cons) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 466 | apply (erule exE) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 467 | apply (drule_tac x = "%m. if m = Suc n then a else i m" in spec, safe) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 468 | apply (drule poly_mult_eq_zero_disj [THEN iffD1], safe) | 
| 15251 | 469 | apply (drule_tac x = "Suc (length q)" in spec) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 470 | apply simp | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 471 | apply (drule_tac x = xa in spec, safe) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 472 | apply (drule_tac x = m in spec, simp, blast) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 473 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 474 | lemmas poly_roots_index_lemma2 = conjI [THEN poly_roots_index_lemma, standard] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 475 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 476 | lemma poly_roots_index_length: "poly p x \<noteq> poly [] x ==> | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 477 | \<exists>i. \<forall>x. (poly p x = 0) --> (\<exists>n. n \<le> length p & x = i n)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 478 | by (blast intro: poly_roots_index_lemma2) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 479 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 480 | lemma poly_roots_finite_lemma: "poly p x \<noteq> poly [] x ==> | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 481 | \<exists>N i. \<forall>x. (poly p x = 0) --> (\<exists>n. (n::nat) < N & x = i n)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 482 | apply (drule poly_roots_index_length, safe) | 
| 15251 | 483 | apply (rule_tac x = "Suc (length p)" in exI) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 484 | apply (rule_tac x = i in exI) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 485 | apply (simp add: less_Suc_eq_le) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 486 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 487 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 488 | (* annoying proof *) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 489 | lemma real_finite_lemma [rule_format (no_asm)]: | 
| 15251 | 490 | "\<forall>P. (\<forall>x. P x --> (\<exists>n. n < N & x = (j::nat=>real) n)) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 491 | --> (\<exists>a. \<forall>x. P x --> x < a)" | 
| 15251 | 492 | apply (induct "N", simp, safe) | 
| 493 | apply (drule_tac x = "%z. P z & (z \<noteq> j N)" in spec) | |
| 494 | apply (auto simp add: less_Suc_eq) | |
| 495 | apply (rename_tac N P a) | |
| 496 | apply (rule_tac x = "abs a + abs (j N) + 1" in exI) | |
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 497 | apply safe | 
| 15251 | 498 | apply (drule_tac x = x in spec, safe) | 
| 23315 | 499 | apply (drule_tac x = "j n" in spec) | 
| 500 | apply arith | |
| 501 | apply arith | |
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 502 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 503 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 504 | lemma poly_roots_finite: "(poly p \<noteq> poly []) = | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 505 | (\<exists>N j. \<forall>x. poly p x = 0 --> (\<exists>n. (n::nat) < N & x = j n))" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 506 | apply safe | 
| 18585 | 507 | apply (erule contrapos_np, rule ext) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 508 | apply (rule ccontr) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 509 | apply (clarify dest!: poly_roots_finite_lemma) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 510 | apply (clarify dest!: real_finite_lemma) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 511 | apply (drule_tac x = a in fun_cong, auto) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 512 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 513 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 514 | text{*Entirety and Cancellation for polynomials*}
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 515 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 516 | lemma poly_entire_lemma: "[| poly p \<noteq> poly [] ; poly q \<noteq> poly [] |] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 517 | ==> poly (p *** q) \<noteq> poly []" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 518 | apply (auto simp add: poly_roots_finite) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 519 | apply (rule_tac x = "N + Na" in exI) | 
| 15251 | 520 | apply (rule_tac x = "%n. if n < N then j n else ja (n - N)" in exI) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 521 | apply (auto simp add: poly_mult_eq_zero_disj, force) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 522 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 523 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 524 | lemma poly_entire: "(poly (p *** q) = poly []) = ((poly p = poly []) | (poly q = poly []))" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 525 | apply (auto intro: ext dest: fun_cong simp add: poly_entire_lemma poly_mult) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 526 | apply (blast intro: ccontr dest: poly_entire_lemma poly_mult [THEN subst]) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 527 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 528 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 529 | lemma poly_entire_neg: "(poly (p *** q) \<noteq> poly []) = ((poly p \<noteq> poly []) & (poly q \<noteq> poly []))" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 530 | by (simp add: poly_entire) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 531 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 532 | lemma fun_eq: " (f = g) = (\<forall>x. f x = g x)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 533 | by (auto intro!: ext) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 534 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 535 | lemma poly_add_minus_zero_iff: "(poly (p +++ -- q) = poly []) = (poly p = poly q)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 536 | by (auto simp add: poly_add poly_minus_def fun_eq poly_cmult) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 537 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 538 | lemma poly_add_minus_mult_eq: "poly (p *** q +++ --(p *** r)) = poly (p *** (q +++ -- r))" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 539 | by (auto simp add: poly_add poly_minus_def fun_eq poly_mult poly_cmult right_distrib) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 540 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 541 | lemma poly_mult_left_cancel: "(poly (p *** q) = poly (p *** r)) = (poly p = poly [] | poly q = poly r)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 542 | apply (rule_tac p1 = "p *** q" in poly_add_minus_zero_iff [THEN subst]) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 543 | apply (auto intro: ext simp add: poly_add_minus_mult_eq poly_entire poly_add_minus_zero_iff) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 544 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 545 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 546 | lemma real_mult_zero_disj_iff: "(x * y = 0) = (x = (0::real) | y = 0)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 547 | by simp | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 548 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 549 | lemma poly_exp_eq_zero: | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 550 | "(poly (p %^ n) = poly []) = (poly p = poly [] & n \<noteq> 0)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 551 | apply (simp only: fun_eq add: all_simps [symmetric]) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 552 | apply (rule arg_cong [where f = All]) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 553 | apply (rule ext) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 554 | apply (induct_tac "n") | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 555 | apply (auto simp add: poly_mult real_mult_zero_disj_iff) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 556 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 557 | declare poly_exp_eq_zero [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 558 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 559 | lemma poly_prime_eq_zero: "poly [a,1] \<noteq> poly []" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 560 | apply (simp add: fun_eq) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 561 | apply (rule_tac x = "1 - a" in exI, simp) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 562 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 563 | declare poly_prime_eq_zero [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 564 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 565 | lemma poly_exp_prime_eq_zero: "(poly ([a, 1] %^ n) \<noteq> poly [])" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 566 | by auto | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 567 | declare poly_exp_prime_eq_zero [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 568 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 569 | text{*A more constructive notion of polynomials being trivial*}
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 570 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 571 | lemma poly_zero_lemma: "poly (h # t) = poly [] ==> h = 0 & poly t = poly []" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 572 | apply (simp add: fun_eq) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 573 | apply (case_tac "h = 0") | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 574 | apply (drule_tac [2] x = 0 in spec, auto) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 575 | apply (case_tac "poly t = poly []", simp) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 576 | apply (auto simp add: poly_roots_finite real_mult_zero_disj_iff) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 577 | apply (drule real_finite_lemma, safe) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 578 | apply (drule_tac x = "abs a + 1" in spec)+ | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 579 | apply arith | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 580 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 581 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 582 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 583 | lemma poly_zero: "(poly p = poly []) = list_all (%c. c = 0) p" | 
| 15251 | 584 | apply (induct "p", simp) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 585 | apply (rule iffI) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 586 | apply (drule poly_zero_lemma, auto) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 587 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 588 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 589 | declare real_mult_zero_disj_iff [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 590 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 591 | lemma pderiv_aux_iszero [rule_format, simp]: | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 592 | "\<forall>n. list_all (%c. c = 0) (pderiv_aux (Suc n) p) = list_all (%c. c = 0) p" | 
| 15251 | 593 | by (induct "p", auto) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 594 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 595 | lemma pderiv_aux_iszero_num: "(number_of n :: nat) \<noteq> 0 | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 596 | ==> (list_all (%c. c = 0) (pderiv_aux (number_of n) p) = | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 597 | list_all (%c. c = 0) p)" | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 598 | apply(drule not0_implies_Suc, clarify) | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 599 | apply (rule_tac n1 = "m" in pderiv_aux_iszero [THEN subst]) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 600 | apply (simp (no_asm_simp) del: pderiv_aux_iszero) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 601 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 602 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 603 | lemma pderiv_iszero [rule_format]: | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 604 | "poly (pderiv p) = poly [] --> (\<exists>h. poly p = poly [h])" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 605 | apply (simp add: poly_zero) | 
| 15251 | 606 | apply (induct "p", force) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 607 | apply (simp add: pderiv_Cons pderiv_aux_iszero_num del: poly_Cons) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 608 | apply (auto simp add: poly_zero [symmetric]) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 609 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 610 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 611 | lemma pderiv_zero_obj: "poly p = poly [] --> (poly (pderiv p) = poly [])" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 612 | apply (simp add: poly_zero) | 
| 15251 | 613 | apply (induct "p", force) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 614 | apply (simp add: pderiv_Cons pderiv_aux_iszero_num del: poly_Cons) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 615 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 616 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 617 | lemma pderiv_zero: "poly p = poly [] ==> (poly (pderiv p) = poly [])" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 618 | by (blast elim: pderiv_zero_obj [THEN impE]) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 619 | declare pderiv_zero [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 620 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 621 | lemma poly_pderiv_welldef: "poly p = poly q ==> (poly (pderiv p) = poly (pderiv q))" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 622 | apply (cut_tac p = "p +++ --q" in pderiv_zero_obj) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 623 | apply (simp add: fun_eq poly_add poly_minus poly_pderiv_add poly_pderiv_minus del: pderiv_zero) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 624 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 625 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 626 | text{*Basics of divisibility.*}
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 627 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 628 | lemma poly_primes: "([a, 1] divides (p *** q)) = ([a, 1] divides p | [a, 1] divides q)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 629 | apply (auto simp add: divides_def fun_eq poly_mult poly_add poly_cmult left_distrib [symmetric]) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 630 | apply (drule_tac x = "-a" in spec) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 631 | apply (auto simp add: poly_linear_divides poly_add poly_cmult left_distrib [symmetric]) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 632 | apply (rule_tac x = "qa *** q" in exI) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 633 | apply (rule_tac [2] x = "p *** qa" in exI) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 634 | apply (auto simp add: poly_add poly_mult poly_cmult mult_ac) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 635 | done | 
| 12224 | 636 | |
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 637 | lemma poly_divides_refl: "p divides p" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 638 | apply (simp add: divides_def) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 639 | apply (rule_tac x = "[1]" in exI) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 640 | apply (auto simp add: poly_mult fun_eq) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 641 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 642 | declare poly_divides_refl [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 643 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 644 | lemma poly_divides_trans: "[| p divides q; q divides r |] ==> p divides r" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 645 | apply (simp add: divides_def, safe) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 646 | apply (rule_tac x = "qa *** qaa" in exI) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 647 | apply (auto simp add: poly_mult fun_eq real_mult_assoc) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 648 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 649 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 650 | lemma poly_divides_exp: "m \<le> n ==> (p %^ m) divides (p %^ n)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 651 | apply (auto simp add: le_iff_add) | 
| 17688 | 652 | apply (induct_tac k) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 653 | apply (rule_tac [2] poly_divides_trans) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 654 | apply (auto simp add: divides_def) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 655 | apply (rule_tac x = p in exI) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 656 | apply (auto simp add: poly_mult fun_eq mult_ac) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 657 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 658 | |
| 17688 | 659 | lemma poly_exp_divides: "[| (p %^ n) divides q; m\<le>n |] ==> (p %^ m) divides q" | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 660 | by (blast intro: poly_divides_exp poly_divides_trans) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 661 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 662 | lemma poly_divides_add: | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 663 | "[| p divides q; p divides r |] ==> p divides (q +++ r)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 664 | apply (simp add: divides_def, auto) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 665 | apply (rule_tac x = "qa +++ qaa" in exI) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 666 | apply (auto simp add: poly_add fun_eq poly_mult right_distrib) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 667 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 668 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 669 | lemma poly_divides_diff: | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 670 | "[| p divides q; p divides (q +++ r) |] ==> p divides r" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 671 | apply (simp add: divides_def, auto) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 672 | apply (rule_tac x = "qaa +++ -- qa" in exI) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 673 | apply (auto simp add: poly_add fun_eq poly_mult poly_minus right_diff_distrib compare_rls add_ac) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 674 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 675 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 676 | lemma poly_divides_diff2: "[| p divides r; p divides (q +++ r) |] ==> p divides q" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 677 | apply (erule poly_divides_diff) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 678 | apply (auto simp add: poly_add fun_eq poly_mult divides_def add_ac) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 679 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 680 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 681 | lemma poly_divides_zero: "poly p = poly [] ==> q divides p" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 682 | apply (simp add: divides_def) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 683 | apply (auto simp add: fun_eq poly_mult) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 684 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 685 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 686 | lemma poly_divides_zero2: "q divides []" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 687 | apply (simp add: divides_def) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 688 | apply (rule_tac x = "[]" in exI) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 689 | apply (auto simp add: fun_eq) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 690 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 691 | declare poly_divides_zero2 [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 692 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 693 | text{*At last, we can consider the order of a root.*}
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 694 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 695 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 696 | lemma poly_order_exists_lemma [rule_format]: | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 697 | "\<forall>p. length p = d --> poly p \<noteq> poly [] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 698 | --> (\<exists>n q. p = mulexp n [-a, 1] q & poly q a \<noteq> 0)" | 
| 15251 | 699 | apply (induct "d") | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 700 | apply (simp add: fun_eq, safe) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 701 | apply (case_tac "poly p a = 0") | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 702 | apply (drule_tac poly_linear_divides [THEN iffD1], safe) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 703 | apply (drule_tac x = q in spec) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 704 | apply (drule_tac poly_entire_neg [THEN iffD1], safe, force, blast) | 
| 15251 | 705 | apply (rule_tac x = "Suc n" in exI) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 706 | apply (rule_tac x = qa in exI) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 707 | apply (simp del: pmult_Cons) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 708 | apply (rule_tac x = 0 in exI, force) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 709 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 710 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 711 | (* FIXME: Tidy up *) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 712 | lemma poly_order_exists: | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 713 | "[| length p = d; poly p \<noteq> poly [] |] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 714 | ==> \<exists>n. ([-a, 1] %^ n) divides p & | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 715 | ~(([-a, 1] %^ (Suc n)) divides p)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 716 | apply (drule poly_order_exists_lemma [where a=a], assumption, clarify) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 717 | apply (rule_tac x = n in exI, safe) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 718 | apply (unfold divides_def) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 719 | apply (rule_tac x = q in exI) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 720 | apply (induct_tac "n", simp) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 721 | apply (simp (no_asm_simp) add: poly_add poly_cmult poly_mult right_distrib mult_ac) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 722 | apply safe | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 723 | apply (subgoal_tac "poly (mulexp n [- a, 1] q) \<noteq> poly ([- a, 1] %^ Suc n *** qa)") | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 724 | apply simp | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 725 | apply (induct_tac "n") | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 726 | apply (simp del: pmult_Cons pexp_Suc) | 
| 18585 | 727 | apply (erule_tac Q = "poly q a = 0" in contrapos_np) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 728 | apply (simp add: poly_add poly_cmult) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 729 | apply (rule pexp_Suc [THEN ssubst]) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 730 | apply (rule ccontr) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 731 | apply (simp add: poly_mult_left_cancel poly_mult_assoc del: pmult_Cons pexp_Suc) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 732 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 733 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 734 | lemma poly_one_divides: "[1] divides p" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 735 | by (simp add: divides_def, auto) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 736 | declare poly_one_divides [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 737 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 738 | lemma poly_order: "poly p \<noteq> poly [] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 739 | ==> EX! n. ([-a, 1] %^ n) divides p & | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 740 | ~(([-a, 1] %^ (Suc n)) divides p)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 741 | apply (auto intro: poly_order_exists simp add: less_linear simp del: pmult_Cons pexp_Suc) | 
| 26316 
9e9e67e33557
removed redundant less_trans, less_linear, le_imp_less_or_eq, le_less_trans, less_le_trans (cf. Orderings.thy);
 wenzelm parents: 
25162diff
changeset | 742 | apply (metis Suc_leI less_linear poly_exp_divides) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 743 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 744 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 745 | text{*Order*}
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 746 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 747 | lemma some1_equalityD: "[| n = (@n. P n); EX! n. P n |] ==> P n" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 748 | by (blast intro: someI2) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 749 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 750 | lemma order: | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 751 | "(([-a, 1] %^ n) divides p & | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 752 | ~(([-a, 1] %^ (Suc n)) divides p)) = | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 753 | ((n = order a p) & ~(poly p = poly []))" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 754 | apply (unfold order_def) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 755 | apply (rule iffI) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 756 | apply (blast dest: poly_divides_zero intro!: some1_equality [symmetric] poly_order) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 757 | apply (blast intro!: poly_order [THEN [2] some1_equalityD]) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 758 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 759 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 760 | lemma order2: "[| poly p \<noteq> poly [] |] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 761 | ==> ([-a, 1] %^ (order a p)) divides p & | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 762 | ~(([-a, 1] %^ (Suc(order a p))) divides p)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 763 | by (simp add: order del: pexp_Suc) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 764 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 765 | lemma order_unique: "[| poly p \<noteq> poly []; ([-a, 1] %^ n) divides p; | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 766 | ~(([-a, 1] %^ (Suc n)) divides p) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 767 | |] ==> (n = order a p)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 768 | by (insert order [of a n p], auto) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 769 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 770 | lemma order_unique_lemma: "(poly p \<noteq> poly [] & ([-a, 1] %^ n) divides p & | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 771 | ~(([-a, 1] %^ (Suc n)) divides p)) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 772 | ==> (n = order a p)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 773 | by (blast intro: order_unique) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 774 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 775 | lemma order_poly: "poly p = poly q ==> order a p = order a q" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 776 | by (auto simp add: fun_eq divides_def poly_mult order_def) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 777 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 778 | lemma pexp_one: "p %^ (Suc 0) = p" | 
| 15251 | 779 | apply (induct "p") | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 780 | apply (auto simp add: numeral_1_eq_1) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 781 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 782 | declare pexp_one [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 783 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 784 | lemma lemma_order_root [rule_format]: | 
| 25162 | 785 | "\<forall>p a. n > 0 & [- a, 1] %^ n divides p & ~ [- a, 1] %^ (Suc n) divides p | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 786 | --> poly p a = 0" | 
| 15251 | 787 | apply (induct "n", blast) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 788 | apply (auto simp add: divides_def poly_mult simp del: pmult_Cons) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 789 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 790 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 791 | lemma order_root: "(poly p a = 0) = ((poly p = poly []) | order a p \<noteq> 0)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 792 | apply (case_tac "poly p = poly []", auto) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 793 | apply (simp add: poly_linear_divides del: pmult_Cons, safe) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 794 | apply (drule_tac [!] a = a in order2) | 
| 25162 | 795 | apply (rule ccontr) | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 796 | apply (simp add: divides_def poly_mult fun_eq del: pmult_Cons, blast) | 
| 25162 | 797 | apply (blast intro: lemma_order_root) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 798 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 799 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 800 | lemma order_divides: "(([-a, 1] %^ n) divides p) = ((poly p = poly []) | n \<le> order a p)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 801 | apply (case_tac "poly p = poly []", auto) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 802 | apply (simp add: divides_def fun_eq poly_mult) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 803 | apply (rule_tac x = "[]" in exI) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 804 | apply (auto dest!: order2 [where a=a] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 805 | intro: poly_exp_divides simp del: pexp_Suc) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 806 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 807 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 808 | lemma order_decomp: | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 809 | "poly p \<noteq> poly [] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 810 | ==> \<exists>q. (poly p = poly (([-a, 1] %^ (order a p)) *** q)) & | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 811 | ~([-a, 1] divides q)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 812 | apply (unfold divides_def) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 813 | apply (drule order2 [where a = a]) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 814 | apply (simp add: divides_def del: pexp_Suc pmult_Cons, safe) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 815 | apply (rule_tac x = q in exI, safe) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 816 | apply (drule_tac x = qa in spec) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 817 | apply (auto simp add: poly_mult fun_eq poly_exp mult_ac simp del: pmult_Cons) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 818 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 819 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 820 | text{*Important composition properties of orders.*}
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 821 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 822 | lemma order_mult: "poly (p *** q) \<noteq> poly [] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 823 | ==> order a (p *** q) = order a p + order a q" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 824 | apply (cut_tac a = a and p = "p***q" and n = "order a p + order a q" in order) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 825 | apply (auto simp add: poly_entire simp del: pmult_Cons) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 826 | apply (drule_tac a = a in order2)+ | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 827 | apply safe | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 828 | apply (simp add: divides_def fun_eq poly_exp_add poly_mult del: pmult_Cons, safe) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 829 | apply (rule_tac x = "qa *** qaa" in exI) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 830 | apply (simp add: poly_mult mult_ac del: pmult_Cons) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 831 | apply (drule_tac a = a in order_decomp)+ | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 832 | apply safe | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 833 | apply (subgoal_tac "[-a,1] divides (qa *** qaa) ") | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 834 | apply (simp add: poly_primes del: pmult_Cons) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 835 | apply (auto simp add: divides_def simp del: pmult_Cons) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 836 | apply (rule_tac x = qb in exI) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 837 | apply (subgoal_tac "poly ([-a, 1] %^ (order a p) *** (qa *** qaa)) = poly ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))") | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 838 | apply (drule poly_mult_left_cancel [THEN iffD1], force) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 839 | apply (subgoal_tac "poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** (qa *** qaa))) = poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))) ") | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 840 | apply (drule poly_mult_left_cancel [THEN iffD1], force) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 841 | apply (simp add: fun_eq poly_exp_add poly_mult mult_ac del: pmult_Cons) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 842 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 843 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 844 | (* FIXME: too too long! *) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 845 | lemma lemma_order_pderiv [rule_format]: | 
| 25162 | 846 | "\<forall>p q a. n > 0 & | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 847 | poly (pderiv p) \<noteq> poly [] & | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 848 | poly p = poly ([- a, 1] %^ n *** q) & ~ [- a, 1] divides q | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 849 | --> n = Suc (order a (pderiv p))" | 
| 15251 | 850 | apply (induct "n", safe) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 851 | apply (rule order_unique_lemma, rule conjI, assumption) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 852 | apply (subgoal_tac "\<forall>r. r divides (pderiv p) = r divides (pderiv ([-a, 1] %^ Suc n *** q))") | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 853 | apply (drule_tac [2] poly_pderiv_welldef) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 854 | prefer 2 apply (simp add: divides_def del: pmult_Cons pexp_Suc) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 855 | apply (simp del: pmult_Cons pexp_Suc) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 856 | apply (rule conjI) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 857 | apply (simp add: divides_def fun_eq del: pmult_Cons pexp_Suc) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 858 | apply (rule_tac x = "[-a, 1] *** (pderiv q) +++ real (Suc n) %* q" in exI) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 859 | apply (simp add: poly_pderiv_mult poly_pderiv_exp_prime poly_add poly_mult poly_cmult right_distrib mult_ac del: pmult_Cons pexp_Suc) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 860 | apply (simp add: poly_mult right_distrib left_distrib mult_ac del: pmult_Cons) | 
| 15251 | 861 | apply (erule_tac V = "\<forall>r. r divides pderiv p = r divides pderiv ([- a, 1] %^ Suc n *** q)" in thin_rl) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 862 | apply (unfold divides_def) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 863 | apply (simp (no_asm) add: poly_pderiv_mult poly_pderiv_exp_prime fun_eq poly_add poly_mult del: pmult_Cons pexp_Suc) | 
| 18585 | 864 | apply (rule contrapos_np, assumption) | 
| 865 | apply (rotate_tac 3, erule contrapos_np) | |
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 866 | apply (simp del: pmult_Cons pexp_Suc, safe) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 867 | apply (rule_tac x = "inverse (real (Suc n)) %* (qa +++ -- (pderiv q))" in exI) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 868 | apply (subgoal_tac "poly ([-a, 1] %^ n *** q) = poly ([-a, 1] %^ n *** ([-a, 1] *** (inverse (real (Suc n)) %* (qa +++ -- (pderiv q))))) ") | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 869 | apply (drule poly_mult_left_cancel [THEN iffD1], simp) | 
| 23497 | 870 | apply (simp add: fun_eq poly_mult poly_add poly_cmult poly_minus del: pmult_Cons mult_cancel_left, safe) | 
| 15251 | 871 | apply (rule_tac c1 = "real (Suc n)" in real_mult_left_cancel [THEN iffD1]) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 872 | apply (simp (no_asm)) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 873 | apply (subgoal_tac "real (Suc n) * (poly ([- a, 1] %^ n) xa * poly q xa) = | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 874 | (poly qa xa + - poly (pderiv q) xa) * | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 875 | (poly ([- a, 1] %^ n) xa * | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 876 | ((- a + xa) * (inverse (real (Suc n)) * real (Suc n))))") | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 877 | apply (simp only: mult_ac) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 878 | apply (rotate_tac 2) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 879 | apply (drule_tac x = xa in spec) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 880 | apply (simp add: left_distrib mult_ac del: pmult_Cons) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 881 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 882 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 883 | lemma order_pderiv: "[| poly (pderiv p) \<noteq> poly []; order a p \<noteq> 0 |] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 884 | ==> (order a p = Suc (order a (pderiv p)))" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 885 | apply (case_tac "poly p = poly []") | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 886 | apply (auto dest: pderiv_zero) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 887 | apply (drule_tac a = a and p = p in order_decomp) | 
| 25157 | 888 | apply (metis lemma_order_pderiv length_0_conv length_greater_0_conv) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 889 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 890 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 891 | text{*Now justify the standard squarefree decomposition, i.e. f / gcd(f,f').    *)
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 892 | (* `a la Harrison*} | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 893 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 894 | lemma poly_squarefree_decomp_order: "[| poly (pderiv p) \<noteq> poly []; | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 895 | poly p = poly (q *** d); | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 896 | poly (pderiv p) = poly (e *** d); | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 897 | poly d = poly (r *** p +++ s *** pderiv p) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 898 | |] ==> order a q = (if order a p = 0 then 0 else 1)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 899 | apply (subgoal_tac "order a p = order a q + order a d") | 
| 15251 | 900 | apply (rule_tac [2] s = "order a (q *** d)" in trans) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 901 | prefer 2 apply (blast intro: order_poly) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 902 | apply (rule_tac [2] order_mult) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 903 | prefer 2 apply force | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 904 | apply (case_tac "order a p = 0", simp) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 905 | apply (subgoal_tac "order a (pderiv p) = order a e + order a d") | 
| 15251 | 906 | apply (rule_tac [2] s = "order a (e *** d)" in trans) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 907 | prefer 2 apply (blast intro: order_poly) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 908 | apply (rule_tac [2] order_mult) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 909 | prefer 2 apply force | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 910 | apply (case_tac "poly p = poly []") | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 911 | apply (drule_tac p = p in pderiv_zero, simp) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 912 | apply (drule order_pderiv, assumption) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 913 | apply (subgoal_tac "order a (pderiv p) \<le> order a d") | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 914 | apply (subgoal_tac [2] " ([-a, 1] %^ (order a (pderiv p))) divides d") | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 915 | prefer 2 apply (simp add: poly_entire order_divides) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 916 | apply (subgoal_tac [2] " ([-a, 1] %^ (order a (pderiv p))) divides p & ([-a, 1] %^ (order a (pderiv p))) divides (pderiv p) ") | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 917 | prefer 3 apply (simp add: order_divides) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 918 | prefer 2 apply (simp add: divides_def del: pexp_Suc pmult_Cons, safe) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 919 | apply (rule_tac x = "r *** qa +++ s *** qaa" in exI) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 920 | apply (simp add: fun_eq poly_add poly_mult left_distrib right_distrib mult_ac del: pexp_Suc pmult_Cons, auto) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 921 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 922 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 923 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 924 | lemma poly_squarefree_decomp_order2: "[| poly (pderiv p) \<noteq> poly []; | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 925 | poly p = poly (q *** d); | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 926 | poly (pderiv p) = poly (e *** d); | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 927 | poly d = poly (r *** p +++ s *** pderiv p) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 928 | |] ==> \<forall>a. order a q = (if order a p = 0 then 0 else 1)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 929 | apply (blast intro: poly_squarefree_decomp_order) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 930 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 931 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 932 | lemma order_root2: "poly p \<noteq> poly [] ==> (poly p a = 0) = (order a p \<noteq> 0)" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 933 | by (rule order_root [THEN ssubst], auto) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 934 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 935 | lemma order_pderiv2: "[| poly (pderiv p) \<noteq> poly []; order a p \<noteq> 0 |] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 936 | ==> (order a (pderiv p) = n) = (order a p = Suc n)" | 
| 25157 | 937 | by (metis Suc_Suc_eq order_pderiv) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 938 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 939 | lemma rsquarefree_roots: | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 940 | "rsquarefree p = (\<forall>a. ~(poly p a = 0 & poly (pderiv p) a = 0))" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 941 | apply (simp add: rsquarefree_def) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 942 | apply (case_tac "poly p = poly []", simp, simp) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 943 | apply (case_tac "poly (pderiv p) = poly []") | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 944 | apply simp | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 945 | apply (drule pderiv_iszero, clarify) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 946 | apply (subgoal_tac "\<forall>a. order a p = order a [h]") | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 947 | apply (simp add: fun_eq) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 948 | apply (rule allI) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 949 | apply (cut_tac p = "[h]" and a = a in order_root) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 950 | apply (simp add: fun_eq) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 951 | apply (blast intro: order_poly) | 
| 25157 | 952 | apply (metis One_nat_def order_pderiv2 order_root rsquarefree_def) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 953 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 954 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 955 | lemma pmult_one: "[1] *** p = p" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 956 | by auto | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 957 | declare pmult_one [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 958 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 959 | lemma poly_Nil_zero: "poly [] = poly [0]" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 960 | by (simp add: fun_eq) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 961 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 962 | lemma rsquarefree_decomp: | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 963 | "[| rsquarefree p; poly p a = 0 |] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 964 | ==> \<exists>q. (poly p = poly ([-a, 1] *** q)) & poly q a \<noteq> 0" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 965 | apply (simp add: rsquarefree_def, safe) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 966 | apply (frule_tac a = a in order_decomp) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 967 | apply (drule_tac x = a in spec) | 
| 20898 | 968 | apply (drule_tac a = a in order_root2 [symmetric]) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 969 | apply (auto simp del: pmult_Cons) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 970 | apply (rule_tac x = q in exI, safe) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 971 | apply (simp add: poly_mult fun_eq) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 972 | apply (drule_tac p1 = q in poly_linear_divides [THEN iffD1]) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 973 | apply (simp add: divides_def del: pmult_Cons, safe) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 974 | apply (drule_tac x = "[]" in spec) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 975 | apply (auto simp add: fun_eq) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 976 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 977 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 978 | lemma poly_squarefree_decomp: "[| poly (pderiv p) \<noteq> poly []; | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 979 | poly p = poly (q *** d); | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 980 | poly (pderiv p) = poly (e *** d); | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 981 | poly d = poly (r *** p +++ s *** pderiv p) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 982 | |] ==> rsquarefree q & (\<forall>a. (poly q a = 0) = (poly p a = 0))" | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 983 | apply (frule poly_squarefree_decomp_order2, assumption+) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 984 | apply (case_tac "poly p = poly []") | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 985 | apply (blast dest: pderiv_zero) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 986 | apply (simp (no_asm) add: rsquarefree_def order_root del: pmult_Cons) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 987 | apply (simp add: poly_entire del: pmult_Cons) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 988 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 989 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 990 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 991 | text{*Normalization of a polynomial.*}
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 992 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 993 | lemma poly_normalize: "poly (pnormalize p) = poly p" | 
| 15251 | 994 | apply (induct "p") | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 995 | apply (auto simp add: fun_eq) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 996 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 997 | declare poly_normalize [simp] | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 998 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 999 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 1000 | text{*The degree of a polynomial.*}
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 1001 | |
| 23256 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 1002 | lemma lemma_degree_zero: | 
| 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 1003 | "list_all (%c. c = 0) p \<longleftrightarrow> pnormalize p = []" | 
| 15251 | 1004 | by (induct "p", auto) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 1005 | |
| 23256 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 1006 | lemma degree_zero: "(poly p = poly []) \<Longrightarrow> (degree p = 0)" | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 1007 | apply (simp add: degree_def) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 1008 | apply (case_tac "pnormalize p = []") | 
| 23256 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 1009 | apply (auto simp add: poly_zero lemma_degree_zero ) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 1010 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 1011 | |
| 23256 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 1012 | lemma pnormalize_sing: "(pnormalize [x] = [x]) \<longleftrightarrow> x \<noteq> 0" by simp | 
| 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 1013 | lemma pnormalize_pair: "y \<noteq> 0 \<longleftrightarrow> (pnormalize [x, y] = [x, y])" by simp | 
| 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 1014 | lemma pnormal_cons: "pnormal p \<Longrightarrow> pnormal (c#p)" | 
| 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 1015 | unfolding pnormal_def by simp | 
| 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 1016 | lemma pnormal_tail: "p\<noteq>[] \<Longrightarrow> pnormal (c#p) \<Longrightarrow> pnormal p" | 
| 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 1017 | unfolding pnormal_def | 
| 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 1018 | apply (cases "pnormalize p = []", auto) | 
| 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 1019 | by (cases "c = 0", auto) | 
| 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 1020 | lemma pnormal_last_nonzero: "pnormal p ==> last p \<noteq> 0" | 
| 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 1021 | apply (induct p, auto simp add: pnormal_def) | 
| 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 1022 | apply (case_tac "pnormalize p = []", auto) | 
| 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 1023 | by (case_tac "a=0", auto) | 
| 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 1024 | lemma pnormal_length: "pnormal p \<Longrightarrow> 0 < length p" | 
| 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 1025 | unfolding pnormal_def length_greater_0_conv by blast | 
| 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 1026 | lemma pnormal_last_length: "\<lbrakk>0 < length p ; last p \<noteq> 0\<rbrakk> \<Longrightarrow> pnormal p" | 
| 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 1027 | apply (induct p, auto) | 
| 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 1028 | apply (case_tac "p = []", auto) | 
| 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 1029 | apply (simp add: pnormal_def) | 
| 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 1030 | by (rule pnormal_cons, auto) | 
| 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 1031 | lemma pnormal_id: "pnormal p \<longleftrightarrow> (0 < length p \<and> last p \<noteq> 0)" | 
| 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 1032 | using pnormal_last_length pnormal_length pnormal_last_nonzero by blast | 
| 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 1033 | |
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 1034 | text{*Tidier versions of finiteness of roots.*}
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 1035 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 1036 | lemma poly_roots_finite_set: "poly p \<noteq> poly [] ==> finite {x. poly p x = 0}"
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 1037 | apply (auto simp add: poly_roots_finite) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 1038 | apply (rule_tac B = "{x::real. \<exists>n. (n::nat) < N & (x = j n) }" in finite_subset)
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 1039 | apply (induct_tac [2] "N", auto) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 1040 | apply (subgoal_tac "{x::real. \<exists>na. na < Suc n & (x = j na) } = { (j n) } Un {x. \<exists>na. na < n & (x = j na) }") 
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 1041 | apply (auto simp add: less_Suc_eq) | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 1042 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 1043 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 1044 | text{*bound for polynomial.*}
 | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 1045 | |
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 1046 | lemma poly_mono: "abs(x) \<le> k ==> abs(poly p x) \<le> poly (map abs p) k" | 
| 15251 | 1047 | apply (induct "p", auto) | 
| 1048 | apply (rule_tac j = "abs a + abs (x * poly p x)" in real_le_trans) | |
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 1049 | apply (rule abs_triangle_ineq) | 
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19765diff
changeset | 1050 | apply (auto intro!: mult_mono simp add: abs_mult) | 
| 14435 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 1051 | done | 
| 
9e22eeccf129
Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
 paulson parents: 
12224diff
changeset | 1052 | |
| 23256 
d797768d5655
Polynomials now only depend on Deriv; Definition of degree changed
 chaieb parents: 
22983diff
changeset | 1053 | lemma poly_Sing: "poly [c] x = c" by simp | 
| 12224 | 1054 | end |