| author | wenzelm | 
| Wed, 04 Apr 2012 11:15:54 +0200 | |
| changeset 47333 | 8204b1023537 | 
| parent 47108 | 2a1953f0d20d | 
| child 52380 | 3cc46b8cca5e | 
| permissions | -rw-r--r-- | 
| 41959 | 1  | 
(* Title: HOL/Library/Poly_Deriv.thy  | 
| 
29985
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
2  | 
Author: Amine Chaieb  | 
| 41959 | 3  | 
Author: Brian Huffman  | 
| 
29985
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
4  | 
*)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
5  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
6  | 
header{* Polynomials and Differentiation *}
 | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
7  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
8  | 
theory Poly_Deriv  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
9  | 
imports Deriv Polynomial  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
10  | 
begin  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
11  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
12  | 
subsection {* Derivatives of univariate polynomials *}
 | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
13  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
14  | 
definition  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
15  | 
pderiv :: "'a::real_normed_field poly \<Rightarrow> 'a poly" where  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
16  | 
"pderiv = poly_rec 0 (\<lambda>a p p'. p + pCons 0 p')"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
17  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
18  | 
lemma pderiv_0 [simp]: "pderiv 0 = 0"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
19  | 
unfolding pderiv_def by (simp add: poly_rec_0)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
20  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
21  | 
lemma pderiv_pCons: "pderiv (pCons a p) = p + pCons 0 (pderiv p)"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
22  | 
unfolding pderiv_def by (simp add: poly_rec_pCons)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
23  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
24  | 
lemma coeff_pderiv: "coeff (pderiv p) n = of_nat (Suc n) * coeff p (Suc n)"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
25  | 
apply (induct p arbitrary: n, simp)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
26  | 
apply (simp add: pderiv_pCons coeff_pCons algebra_simps split: nat.split)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
27  | 
done  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
28  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
29  | 
lemma pderiv_eq_0_iff: "pderiv p = 0 \<longleftrightarrow> degree p = 0"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
30  | 
apply (rule iffI)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
31  | 
apply (cases p, simp)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
32  | 
apply (simp add: expand_poly_eq coeff_pderiv del: of_nat_Suc)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
33  | 
apply (simp add: expand_poly_eq coeff_pderiv coeff_eq_0)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
34  | 
done  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
35  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
36  | 
lemma degree_pderiv: "degree (pderiv p) = degree p - 1"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
37  | 
apply (rule order_antisym [OF degree_le])  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
38  | 
apply (simp add: coeff_pderiv coeff_eq_0)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
39  | 
apply (cases "degree p", simp)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
40  | 
apply (rule le_degree)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
41  | 
apply (simp add: coeff_pderiv del: of_nat_Suc)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
42  | 
apply (rule subst, assumption)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
43  | 
apply (rule leading_coeff_neq_0, clarsimp)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
44  | 
done  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
45  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
46  | 
lemma pderiv_singleton [simp]: "pderiv [:a:] = 0"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
47  | 
by (simp add: pderiv_pCons)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
48  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
49  | 
lemma pderiv_add: "pderiv (p + q) = pderiv p + pderiv q"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
50  | 
by (rule poly_ext, simp add: coeff_pderiv algebra_simps)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
51  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
52  | 
lemma pderiv_minus: "pderiv (- p) = - pderiv p"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
53  | 
by (rule poly_ext, simp add: coeff_pderiv)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
54  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
55  | 
lemma pderiv_diff: "pderiv (p - q) = pderiv p - pderiv q"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
56  | 
by (rule poly_ext, simp add: coeff_pderiv algebra_simps)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
57  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
58  | 
lemma pderiv_smult: "pderiv (smult a p) = smult a (pderiv p)"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
59  | 
by (rule poly_ext, simp add: coeff_pderiv algebra_simps)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
60  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
61  | 
lemma pderiv_mult: "pderiv (p * q) = p * pderiv q + q * pderiv p"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
62  | 
apply (induct p)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
63  | 
apply simp  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
64  | 
apply (simp add: pderiv_add pderiv_smult pderiv_pCons algebra_simps)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
65  | 
done  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
66  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
67  | 
lemma pderiv_power_Suc:  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
68  | 
"pderiv (p ^ Suc n) = smult (of_nat (Suc n)) (p ^ n) * pderiv p"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
69  | 
apply (induct n)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
70  | 
apply simp  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
71  | 
apply (subst power_Suc)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
72  | 
apply (subst pderiv_mult)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
73  | 
apply (erule ssubst)  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
44317 
diff
changeset
 | 
74  | 
apply (simp only: of_nat_Suc smult_add_left smult_1_left)  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
44317 
diff
changeset
 | 
75  | 
apply (simp add: algebra_simps) (* FIXME *)  | 
| 
29985
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
76  | 
done  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
77  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
78  | 
lemma DERIV_cmult2: "DERIV f x :> D ==> DERIV (%x. (f x) * c :: real) x :> D * c"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
79  | 
by (simp add: DERIV_cmult mult_commute [of _ c])  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
80  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
81  | 
lemma DERIV_pow2: "DERIV (%x. x ^ Suc n) x :> real (Suc n) * (x ^ n)"  | 
| 
44317
 
b7e9fa025f15
remove redundant lemma lemma_DERIV_subst in favor of DERIV_cong
 
huffman 
parents: 
41959 
diff
changeset
 | 
82  | 
by (rule DERIV_cong, rule DERIV_pow, simp)  | 
| 
29985
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
83  | 
declare DERIV_pow2 [simp] DERIV_pow [simp]  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
84  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
85  | 
lemma DERIV_add_const: "DERIV f x :> D ==> DERIV (%x. a + f x :: 'a::real_normed_field) x :> D"  | 
| 
44317
 
b7e9fa025f15
remove redundant lemma lemma_DERIV_subst in favor of DERIV_cong
 
huffman 
parents: 
41959 
diff
changeset
 | 
86  | 
by (rule DERIV_cong, rule DERIV_add, auto)  | 
| 
29985
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
87  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
88  | 
lemma poly_DERIV[simp]: "DERIV (%x. poly p x) x :> poly (pderiv p) x"  | 
| 31881 | 89  | 
by (induct p, auto intro!: DERIV_intros simp add: pderiv_pCons)  | 
| 
29985
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
90  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
91  | 
text{* Consequences of the derivative theorem above*}
 | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
92  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
93  | 
lemma poly_differentiable[simp]: "(%x. poly p x) differentiable (x::real)"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
94  | 
apply (simp add: differentiable_def)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
95  | 
apply (blast intro: poly_DERIV)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
96  | 
done  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
97  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
98  | 
lemma poly_isCont[simp]: "isCont (%x. poly p x) (x::real)"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
99  | 
by (rule poly_DERIV [THEN DERIV_isCont])  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
100  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
101  | 
lemma poly_IVT_pos: "[| a < b; poly p (a::real) < 0; 0 < poly p b |]  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
102  | 
==> \<exists>x. a < x & x < b & (poly p x = 0)"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
103  | 
apply (cut_tac f = "%x. poly p x" and a = a and b = b and y = 0 in IVT_objl)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
104  | 
apply (auto simp add: order_le_less)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
105  | 
done  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
106  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
107  | 
lemma poly_IVT_neg: "[| (a::real) < b; 0 < poly p a; poly p b < 0 |]  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
108  | 
==> \<exists>x. a < x & x < b & (poly p x = 0)"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
109  | 
by (insert poly_IVT_pos [where p = "- p" ]) simp  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
110  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
111  | 
lemma poly_MVT: "(a::real) < b ==>  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
112  | 
\<exists>x. a < x & x < b & (poly p b - poly p a = (b - a) * poly (pderiv p) x)"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
113  | 
apply (drule_tac f = "poly p" in MVT, auto)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
114  | 
apply (rule_tac x = z in exI)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
115  | 
apply (auto simp add: real_mult_left_cancel poly_DERIV [THEN DERIV_unique])  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
116  | 
done  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
117  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
118  | 
text{*Lemmas for Derivatives*}
 | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
119  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
120  | 
lemma order_unique_lemma:  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
121  | 
fixes p :: "'a::idom poly"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
122  | 
assumes "[:-a, 1:] ^ n dvd p \<and> \<not> [:-a, 1:] ^ Suc n dvd p"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
123  | 
shows "n = order a p"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
124  | 
unfolding Polynomial.order_def  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
125  | 
apply (rule Least_equality [symmetric])  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
126  | 
apply (rule assms [THEN conjunct2])  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
127  | 
apply (erule contrapos_np)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
128  | 
apply (rule power_le_dvd)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
129  | 
apply (rule assms [THEN conjunct1])  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
130  | 
apply simp  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
131  | 
done  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
132  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
133  | 
lemma lemma_order_pderiv1:  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
134  | 
"pderiv ([:- a, 1:] ^ Suc n * q) = [:- a, 1:] ^ Suc n * pderiv q +  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
135  | 
smult (of_nat (Suc n)) (q * [:- a, 1:] ^ n)"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
136  | 
apply (simp only: pderiv_mult pderiv_power_Suc)  | 
| 
30273
 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 
huffman 
parents: 
29985 
diff
changeset
 | 
137  | 
apply (simp del: power_Suc of_nat_Suc add: pderiv_pCons)  | 
| 
29985
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
138  | 
done  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
139  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
140  | 
lemma dvd_add_cancel1:  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
141  | 
fixes a b c :: "'a::comm_ring_1"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
142  | 
shows "a dvd b + c \<Longrightarrow> a dvd b \<Longrightarrow> a dvd c"  | 
| 
35050
 
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
 
haftmann 
parents: 
31881 
diff
changeset
 | 
143  | 
by (drule (1) Rings.dvd_diff, simp)  | 
| 
29985
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
144  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
145  | 
lemma lemma_order_pderiv [rule_format]:  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
146  | 
"\<forall>p q a. 0 < n &  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
147  | 
pderiv p \<noteq> 0 &  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
148  | 
p = [:- a, 1:] ^ n * q & ~ [:- a, 1:] dvd q  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
149  | 
--> n = Suc (order a (pderiv p))"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
150  | 
apply (cases "n", safe, rename_tac n p q a)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
151  | 
apply (rule order_unique_lemma)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
152  | 
apply (rule conjI)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
153  | 
apply (subst lemma_order_pderiv1)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
154  | 
apply (rule dvd_add)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
155  | 
apply (rule dvd_mult2)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
156  | 
apply (rule le_imp_power_dvd, simp)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
157  | 
apply (rule dvd_smult)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
158  | 
apply (rule dvd_mult)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
159  | 
apply (rule dvd_refl)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
160  | 
apply (subst lemma_order_pderiv1)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
161  | 
apply (erule contrapos_nn) back  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
162  | 
apply (subgoal_tac "[:- a, 1:] ^ Suc n dvd q * [:- a, 1:] ^ n")  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
163  | 
apply (simp del: mult_pCons_left)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
164  | 
apply (drule dvd_add_cancel1)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
165  | 
apply (simp del: mult_pCons_left)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
166  | 
apply (drule dvd_smult_cancel, simp del: of_nat_Suc)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
167  | 
apply assumption  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
168  | 
done  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
169  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
170  | 
lemma order_decomp:  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
171  | 
"p \<noteq> 0  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
172  | 
==> \<exists>q. p = [:-a, 1:] ^ (order a p) * q &  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
173  | 
~([:-a, 1:] dvd q)"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
174  | 
apply (drule order [where a=a])  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
175  | 
apply (erule conjE)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
176  | 
apply (erule dvdE)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
177  | 
apply (rule exI)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
178  | 
apply (rule conjI, assumption)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
179  | 
apply (erule contrapos_nn)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
180  | 
apply (erule ssubst) back  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
181  | 
apply (subst power_Suc2)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
182  | 
apply (erule mult_dvd_mono [OF dvd_refl])  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
183  | 
done  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
184  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
185  | 
lemma order_pderiv: "[| pderiv p \<noteq> 0; order a p \<noteq> 0 |]  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
186  | 
==> (order a p = Suc (order a (pderiv p)))"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
187  | 
apply (case_tac "p = 0", simp)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
188  | 
apply (drule_tac a = a and p = p in order_decomp)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
189  | 
using neq0_conv  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
190  | 
apply (blast intro: lemma_order_pderiv)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
191  | 
done  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
192  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
193  | 
lemma order_mult: "p * q \<noteq> 0 \<Longrightarrow> order a (p * q) = order a p + order a q"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
194  | 
proof -  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
195  | 
def i \<equiv> "order a p"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
196  | 
def j \<equiv> "order a q"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
197  | 
def t \<equiv> "[:-a, 1:]"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
198  | 
have t_dvd_iff: "\<And>u. t dvd u \<longleftrightarrow> poly u a = 0"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
199  | 
unfolding t_def by (simp add: dvd_iff_poly_eq_0)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
200  | 
assume "p * q \<noteq> 0"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
201  | 
then show "order a (p * q) = i + j"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
202  | 
apply clarsimp  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
203  | 
apply (drule order [where a=a and p=p, folded i_def t_def])  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
204  | 
apply (drule order [where a=a and p=q, folded j_def t_def])  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
205  | 
apply clarify  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
206  | 
apply (rule order_unique_lemma [symmetric], fold t_def)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
207  | 
apply (erule dvdE)+  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
208  | 
apply (simp add: power_add t_dvd_iff)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
209  | 
done  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
210  | 
qed  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
211  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
212  | 
text{*Now justify the standard squarefree decomposition, i.e. f / gcd(f,f'). *}
 | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
213  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
214  | 
lemma order_divides: "[:-a, 1:] ^ n dvd p \<longleftrightarrow> p = 0 \<or> n \<le> order a p"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
215  | 
apply (cases "p = 0", auto)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
216  | 
apply (drule order_2 [where a=a and p=p])  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
217  | 
apply (erule contrapos_np)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
218  | 
apply (erule power_le_dvd)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
219  | 
apply simp  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
220  | 
apply (erule power_le_dvd [OF order_1])  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
221  | 
done  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
222  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
223  | 
lemma poly_squarefree_decomp_order:  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
224  | 
assumes "pderiv p \<noteq> 0"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
225  | 
and p: "p = q * d"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
226  | 
and p': "pderiv p = e * d"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
227  | 
and d: "d = r * p + s * pderiv p"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
228  | 
shows "order a q = (if order a p = 0 then 0 else 1)"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
229  | 
proof (rule classical)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
230  | 
assume 1: "order a q \<noteq> (if order a p = 0 then 0 else 1)"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
231  | 
from `pderiv p \<noteq> 0` have "p \<noteq> 0" by auto  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
232  | 
with p have "order a p = order a q + order a d"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
233  | 
by (simp add: order_mult)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
234  | 
with 1 have "order a p \<noteq> 0" by (auto split: if_splits)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
235  | 
have "order a (pderiv p) = order a e + order a d"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
236  | 
using `pderiv p \<noteq> 0` `pderiv p = e * d` by (simp add: order_mult)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
237  | 
have "order a p = Suc (order a (pderiv p))"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
238  | 
using `pderiv p \<noteq> 0` `order a p \<noteq> 0` by (rule order_pderiv)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
239  | 
have "d \<noteq> 0" using `p \<noteq> 0` `p = q * d` by simp  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
240  | 
have "([:-a, 1:] ^ (order a (pderiv p))) dvd d"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
241  | 
apply (simp add: d)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
242  | 
apply (rule dvd_add)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
243  | 
apply (rule dvd_mult)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
244  | 
apply (simp add: order_divides `p \<noteq> 0`  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
245  | 
`order a p = Suc (order a (pderiv p))`)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
246  | 
apply (rule dvd_mult)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
247  | 
apply (simp add: order_divides)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
248  | 
done  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
249  | 
then have "order a (pderiv p) \<le> order a d"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
250  | 
using `d \<noteq> 0` by (simp add: order_divides)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
251  | 
show ?thesis  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
252  | 
using `order a p = order a q + order a d`  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
253  | 
using `order a (pderiv p) = order a e + order a d`  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
254  | 
using `order a p = Suc (order a (pderiv p))`  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
255  | 
using `order a (pderiv p) \<le> order a d`  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
256  | 
by auto  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
257  | 
qed  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
258  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
259  | 
lemma poly_squarefree_decomp_order2: "[| pderiv p \<noteq> 0;  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
260  | 
p = q * d;  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
261  | 
pderiv p = e * d;  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
262  | 
d = r * p + s * pderiv p  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
263  | 
|] ==> \<forall>a. order a q = (if order a p = 0 then 0 else 1)"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
264  | 
apply (blast intro: poly_squarefree_decomp_order)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
265  | 
done  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
266  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
267  | 
lemma order_pderiv2: "[| pderiv p \<noteq> 0; order a p \<noteq> 0 |]  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
268  | 
==> (order a (pderiv p) = n) = (order a p = Suc n)"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
269  | 
apply (auto dest: order_pderiv)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
270  | 
done  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
271  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
272  | 
definition  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
273  | 
rsquarefree :: "'a::idom poly => bool" where  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
274  | 
"rsquarefree p = (p \<noteq> 0 & (\<forall>a. (order a p = 0) | (order a p = 1)))"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
275  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
276  | 
lemma pderiv_iszero: "pderiv p = 0 \<Longrightarrow> \<exists>h. p = [:h:]"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
277  | 
apply (simp add: pderiv_eq_0_iff)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
278  | 
apply (case_tac p, auto split: if_splits)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
279  | 
done  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
280  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
281  | 
lemma rsquarefree_roots:  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
282  | 
"rsquarefree p = (\<forall>a. ~(poly p a = 0 & poly (pderiv p) a = 0))"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
283  | 
apply (simp add: rsquarefree_def)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
284  | 
apply (case_tac "p = 0", simp, simp)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
285  | 
apply (case_tac "pderiv p = 0")  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
286  | 
apply simp  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
287  | 
apply (drule pderiv_iszero, clarify)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
288  | 
apply simp  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
289  | 
apply (rule allI)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
290  | 
apply (cut_tac p = "[:h:]" and a = a in order_root)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
291  | 
apply simp  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
292  | 
apply (auto simp add: order_root order_pderiv2)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
293  | 
apply (erule_tac x="a" in allE, simp)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
294  | 
done  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
295  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
296  | 
lemma poly_squarefree_decomp:  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
297  | 
assumes "pderiv p \<noteq> 0"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
298  | 
and "p = q * d"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
299  | 
and "pderiv p = e * d"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
300  | 
and "d = r * p + s * pderiv p"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
301  | 
shows "rsquarefree q & (\<forall>a. (poly q a = 0) = (poly p a = 0))"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
302  | 
proof -  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
303  | 
from `pderiv p \<noteq> 0` have "p \<noteq> 0" by auto  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
304  | 
with `p = q * d` have "q \<noteq> 0" by simp  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
305  | 
have "\<forall>a. order a q = (if order a p = 0 then 0 else 1)"  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
306  | 
using assms by (rule poly_squarefree_decomp_order2)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
307  | 
with `p \<noteq> 0` `q \<noteq> 0` show ?thesis  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
308  | 
by (simp add: rsquarefree_def order_root)  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
309  | 
qed  | 
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
310  | 
|
| 
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents:  
diff
changeset
 | 
311  | 
end  |