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