| author | wenzelm | 
| Sat, 09 Jan 2016 22:22:17 +0100 | |
| changeset 62114 | a7cf464933f7 | 
| parent 62072 | bf3d9f113474 | 
| child 62128 | 3201ddb00097 | 
| 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 | |
| 60500 | 6 | section\<open>Polynomials and Differentiation\<close> | 
| 29985 
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 | |
| 60500 | 12 | subsection \<open>Derivatives of univariate polynomials\<close> | 
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 13 | |
| 52380 | 14 | function pderiv :: "'a::real_normed_field poly \<Rightarrow> 'a poly" | 
| 15 | where | |
| 16 | [simp del]: "pderiv (pCons a p) = (if p = 0 then 0 else p + pCons 0 (pderiv p))" | |
| 17 | by (auto intro: pCons_cases) | |
| 18 | ||
| 19 | termination pderiv | |
| 20 | by (relation "measure degree") simp_all | |
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 21 | |
| 52380 | 22 | lemma pderiv_0 [simp]: | 
| 23 | "pderiv 0 = 0" | |
| 24 | using pderiv.simps [of 0 0] by simp | |
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 25 | |
| 52380 | 26 | lemma pderiv_pCons: | 
| 27 | "pderiv (pCons a p) = p + pCons 0 (pderiv p)" | |
| 28 | by (simp add: pderiv.simps) | |
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 29 | |
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 30 | lemma coeff_pderiv: "coeff (pderiv p) n = of_nat (Suc n) * coeff p (Suc n)" | 
| 56383 | 31 | by (induct p arbitrary: n) | 
| 32 | (auto simp add: pderiv_pCons coeff_pCons algebra_simps split: nat.split) | |
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 33 | |
| 52380 | 34 | primrec pderiv_coeffs :: "'a::comm_monoid_add list \<Rightarrow> 'a list" | 
| 35 | where | |
| 36 | "pderiv_coeffs [] = []" | |
| 37 | | "pderiv_coeffs (x # xs) = plus_coeffs xs (cCons 0 (pderiv_coeffs xs))" | |
| 38 | ||
| 39 | lemma coeffs_pderiv [code abstract]: | |
| 40 | "coeffs (pderiv p) = pderiv_coeffs (coeffs p)" | |
| 41 | by (rule sym, induct p) (simp_all add: pderiv_pCons coeffs_plus_eq_plus_coeffs cCons_def) | |
| 42 | ||
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 43 | 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 | 44 | apply (rule iffI) | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 45 | apply (cases p, simp) | 
| 52380 | 46 | apply (simp add: poly_eq_iff coeff_pderiv del: of_nat_Suc) | 
| 47 | apply (simp add: poly_eq_iff coeff_pderiv coeff_eq_0) | |
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 48 | done | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 49 | |
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 50 | 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 | 51 | apply (rule order_antisym [OF degree_le]) | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 52 | 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 | 53 | apply (cases "degree p", simp) | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 54 | apply (rule le_degree) | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 55 | apply (simp add: coeff_pderiv del: of_nat_Suc) | 
| 56383 | 56 | apply (metis degree_0 leading_coeff_0_iff nat.distinct(1)) | 
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 57 | done | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 58 | |
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 59 | lemma pderiv_singleton [simp]: "pderiv [:a:] = 0" | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 60 | by (simp add: pderiv_pCons) | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 61 | |
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 62 | lemma pderiv_add: "pderiv (p + q) = pderiv p + pderiv q" | 
| 52380 | 63 | by (rule poly_eqI, simp add: coeff_pderiv algebra_simps) | 
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 64 | |
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 65 | lemma pderiv_minus: "pderiv (- p) = - pderiv p" | 
| 52380 | 66 | by (rule poly_eqI, simp add: coeff_pderiv) | 
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 67 | |
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 68 | lemma pderiv_diff: "pderiv (p - q) = pderiv p - pderiv q" | 
| 52380 | 69 | by (rule poly_eqI, simp add: coeff_pderiv algebra_simps) | 
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 70 | |
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 71 | lemma pderiv_smult: "pderiv (smult a p) = smult a (pderiv p)" | 
| 52380 | 72 | by (rule poly_eqI, simp add: coeff_pderiv algebra_simps) | 
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 73 | |
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 74 | lemma pderiv_mult: "pderiv (p * q) = p * pderiv q + q * pderiv p" | 
| 56383 | 75 | by (induct p) (auto simp: pderiv_add pderiv_smult pderiv_pCons algebra_simps) | 
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 76 | |
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 77 | lemma pderiv_power_Suc: | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 78 | "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 | 79 | apply (induct n) | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 80 | apply simp | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 81 | apply (subst power_Suc) | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 82 | apply (subst pderiv_mult) | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 83 | apply (erule ssubst) | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
44317diff
changeset | 84 | apply (simp only: of_nat_Suc smult_add_left smult_1_left) | 
| 56383 | 85 | apply (simp add: algebra_simps) | 
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 86 | done | 
| 
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 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: 
41959diff
changeset | 89 | 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 | 90 | declare DERIV_pow2 [simp] DERIV_pow [simp] | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 91 | |
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 92 | 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: 
41959diff
changeset | 93 | 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 | 94 | |
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 95 | lemma poly_DERIV[simp]: "DERIV (%x. poly p x) x :> poly (pderiv p) x" | 
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56217diff
changeset | 96 | by (induct p, auto intro!: derivative_eq_intros simp add: pderiv_pCons) | 
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 97 | |
| 62065 | 98 | lemma continuous_on_poly [continuous_intros]: | 
| 99 |   fixes p :: "'a :: {real_normed_field} poly"
 | |
| 100 | assumes "continuous_on A f" | |
| 101 | shows "continuous_on A (\<lambda>x. poly p (f x))" | |
| 102 | proof - | |
| 103 | have "continuous_on A (\<lambda>x. (\<Sum>i\<le>degree p. (f x) ^ i * coeff p i))" | |
| 104 | by (intro continuous_intros assms) | |
| 105 | also have "\<dots> = (\<lambda>x. poly p (f x))" by (intro ext) (simp add: poly_altdef mult_ac) | |
| 106 | finally show ?thesis . | |
| 107 | qed | |
| 108 | ||
| 60500 | 109 | text\<open>Consequences of the derivative theorem above\<close> | 
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 110 | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
52380diff
changeset | 111 | lemma poly_differentiable[simp]: "(%x. poly p x) differentiable (at x::real filter)" | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
52380diff
changeset | 112 | apply (simp add: real_differentiable_def) | 
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 113 | apply (blast intro: poly_DERIV) | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 114 | done | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 115 | |
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 116 | 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 | 117 | by (rule poly_DERIV [THEN DERIV_isCont]) | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 118 | |
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 119 | 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 | 120 | ==> \<exists>x. a < x & x < b & (poly p x = 0)" | 
| 56383 | 121 | using IVT_objl [of "poly p" a 0 b] | 
| 122 | by (auto simp add: order_le_less) | |
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 123 | |
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 124 | 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 | 125 | ==> \<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 | 126 | 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 | 127 | |
| 62065 | 128 | lemma poly_IVT: | 
| 129 | fixes p::"real poly" | |
| 130 | assumes "a<b" and "poly p a * poly p b < 0" | |
| 131 | shows "\<exists>x>a. x < b \<and> poly p x = 0" | |
| 132 | by (metis assms(1) assms(2) less_not_sym mult_less_0_iff poly_IVT_neg poly_IVT_pos) | |
| 133 | ||
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 134 | lemma poly_MVT: "(a::real) < b ==> | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 135 | \<exists>x. a < x & x < b & (poly p b - poly p a = (b - a) * poly (pderiv p) x)" | 
| 56383 | 136 | using MVT [of a b "poly p"] | 
| 137 | apply auto | |
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 138 | apply (rule_tac x = z in exI) | 
| 56217 
dc429a5b13c4
Some rationalisation of basic lemmas
 paulson <lp15@cam.ac.uk> parents: 
56181diff
changeset | 139 | apply (auto simp add: mult_left_cancel poly_DERIV [THEN DERIV_unique]) | 
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 140 | done | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 141 | |
| 62065 | 142 | lemma poly_MVT': | 
| 143 |   assumes "{min a b..max a b} \<subseteq> A"
 | |
| 144 | shows "\<exists>x\<in>A. poly p b - poly p a = (b - a) * poly (pderiv p) (x::real)" | |
| 145 | proof (cases a b rule: linorder_cases) | |
| 146 | case less | |
| 147 | from poly_MVT[OF less, of p] guess x by (elim exE conjE) | |
| 148 | thus ?thesis by (intro bexI[of _ x]) (auto intro!: subsetD[OF assms]) | |
| 149 | ||
| 150 | next | |
| 151 | case greater | |
| 152 | from poly_MVT[OF greater, of p] guess x by (elim exE conjE) | |
| 153 | thus ?thesis by (intro bexI[of _ x]) (auto simp: algebra_simps intro!: subsetD[OF assms]) | |
| 154 | qed (insert assms, auto) | |
| 155 | ||
| 156 | lemma poly_pinfty_gt_lc: | |
| 157 | fixes p:: "real poly" | |
| 158 | assumes "lead_coeff p > 0" | |
| 159 | shows "\<exists> n. \<forall> x \<ge> n. poly p x \<ge> lead_coeff p" using assms | |
| 160 | proof (induct p) | |
| 161 | case 0 | |
| 162 | thus ?case by auto | |
| 163 | next | |
| 164 | case (pCons a p) | |
| 165 | have "\<lbrakk>a\<noteq>0;p=0\<rbrakk> \<Longrightarrow> ?case" by auto | |
| 166 | moreover have "p\<noteq>0 \<Longrightarrow> ?case" | |
| 167 | proof - | |
| 168 | assume "p\<noteq>0" | |
| 169 | then obtain n1 where gte_lcoeff:"\<forall>x\<ge>n1. lead_coeff p \<le> poly p x" using that pCons by auto | |
| 62072 | 170 | have gt_0:"lead_coeff p >0" using pCons(3) \<open>p\<noteq>0\<close> by auto | 
| 62065 | 171 | def n\<equiv>"max n1 (1+ \<bar>a\<bar>/(lead_coeff p))" | 
| 172 | show ?thesis | |
| 173 | proof (rule_tac x=n in exI,rule,rule) | |
| 174 | fix x assume "n \<le> x" | |
| 175 | hence "lead_coeff p \<le> poly p x" | |
| 176 | using gte_lcoeff unfolding n_def by auto | |
| 177 | hence " \<bar>a\<bar>/(lead_coeff p) \<ge> \<bar>a\<bar>/(poly p x)" and "poly p x>0" using gt_0 | |
| 178 | by (intro frac_le,auto) | |
| 62072 | 179 | hence "x\<ge>1+ \<bar>a\<bar>/(poly p x)" using \<open>n\<le>x\<close>[unfolded n_def] by auto | 
| 62065 | 180 | thus "lead_coeff (pCons a p) \<le> poly (pCons a p) x" | 
| 62072 | 181 | using \<open>lead_coeff p \<le> poly p x\<close> \<open>poly p x>0\<close> \<open>p\<noteq>0\<close> | 
| 62065 | 182 | by (auto simp add:field_simps) | 
| 183 | qed | |
| 184 | qed | |
| 185 | ultimately show ?case by fastforce | |
| 186 | qed | |
| 187 | ||
| 188 | ||
| 60500 | 189 | text\<open>Lemmas for Derivatives\<close> | 
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 190 | |
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 191 | lemma order_unique_lemma: | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 192 | fixes p :: "'a::idom poly" | 
| 56383 | 193 | assumes "[:-a, 1:] ^ n dvd p" "\<not> [:-a, 1:] ^ Suc n dvd p" | 
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 194 | shows "n = order a p" | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 195 | unfolding Polynomial.order_def | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 196 | apply (rule Least_equality [symmetric]) | 
| 58199 
5fbe474b5da8
explicit theory with additional, less commonly used list operations
 haftmann parents: 
56383diff
changeset | 197 | apply (fact assms) | 
| 
5fbe474b5da8
explicit theory with additional, less commonly used list operations
 haftmann parents: 
56383diff
changeset | 198 | apply (rule classical) | 
| 
5fbe474b5da8
explicit theory with additional, less commonly used list operations
 haftmann parents: 
56383diff
changeset | 199 | apply (erule notE) | 
| 
5fbe474b5da8
explicit theory with additional, less commonly used list operations
 haftmann parents: 
56383diff
changeset | 200 | unfolding not_less_eq_eq | 
| 
5fbe474b5da8
explicit theory with additional, less commonly used list operations
 haftmann parents: 
56383diff
changeset | 201 | using assms(1) apply (rule power_le_dvd) | 
| 
5fbe474b5da8
explicit theory with additional, less commonly used list operations
 haftmann parents: 
56383diff
changeset | 202 | apply assumption | 
| 
5fbe474b5da8
explicit theory with additional, less commonly used list operations
 haftmann parents: 
56383diff
changeset | 203 | done | 
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 204 | |
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 205 | lemma lemma_order_pderiv1: | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 206 | "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 | 207 | 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 | 208 | apply (simp only: pderiv_mult pderiv_power_Suc) | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
29985diff
changeset | 209 | 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 | 210 | done | 
| 
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 | lemma dvd_add_cancel1: | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 213 | 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 | 214 | 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: 
31881diff
changeset | 215 | 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 | 216 | |
| 56383 | 217 | lemma lemma_order_pderiv: | 
| 218 | assumes n: "0 < n" | |
| 219 | and pd: "pderiv p \<noteq> 0" | |
| 220 | and pe: "p = [:- a, 1:] ^ n * q" | |
| 221 | and nd: "~ [:- a, 1:] dvd q" | |
| 222 | shows "n = Suc (order a (pderiv p))" | |
| 223 | using n | |
| 224 | proof - | |
| 225 | have "pderiv ([:- a, 1:] ^ n * q) \<noteq> 0" | |
| 226 | using assms by auto | |
| 227 | obtain n' where "n = Suc n'" "0 < Suc n'" "pderiv ([:- a, 1:] ^ Suc n' * q) \<noteq> 0" | |
| 228 | using assms by (cases n) auto | |
| 229 | then have *: "!!k l. k dvd k * pderiv q + smult (of_nat (Suc n')) l \<Longrightarrow> k dvd l" | |
| 230 | by (metis dvd_add_cancel1 dvd_smult_iff dvd_triv_left of_nat_eq_0_iff old.nat.distinct(2)) | |
| 231 | have "n' = order a (pderiv ([:- a, 1:] ^ Suc n' * q))" | |
| 232 | proof (rule order_unique_lemma) | |
| 233 | show "[:- a, 1:] ^ n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)" | |
| 234 | apply (subst lemma_order_pderiv1) | |
| 235 | apply (rule dvd_add) | |
| 236 | apply (metis dvdI dvd_mult2 power_Suc2) | |
| 237 | apply (metis dvd_smult dvd_triv_right) | |
| 238 | done | |
| 239 | next | |
| 240 | show "\<not> [:- a, 1:] ^ Suc n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)" | |
| 241 | apply (subst lemma_order_pderiv1) | |
| 60867 | 242 | by (metis * nd dvd_mult_cancel_right power_not_zero pCons_eq_0_iff power_Suc zero_neq_one) | 
| 56383 | 243 | qed | 
| 244 | then show ?thesis | |
| 60500 | 245 | by (metis \<open>n = Suc n'\<close> pe) | 
| 56383 | 246 | qed | 
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 247 | |
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 248 | lemma order_decomp: | 
| 60688 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60500diff
changeset | 249 | assumes "p \<noteq> 0" | 
| 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60500diff
changeset | 250 | shows "\<exists>q. p = [:- a, 1:] ^ order a p * q \<and> \<not> [:- a, 1:] dvd q" | 
| 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60500diff
changeset | 251 | proof - | 
| 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60500diff
changeset | 252 | from assms have A: "[:- a, 1:] ^ order a p dvd p" | 
| 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60500diff
changeset | 253 | and B: "\<not> [:- a, 1:] ^ Suc (order a p) dvd p" by (auto dest: order) | 
| 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60500diff
changeset | 254 | from A obtain q where C: "p = [:- a, 1:] ^ order a p * q" .. | 
| 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60500diff
changeset | 255 | with B have "\<not> [:- a, 1:] ^ Suc (order a p) dvd [:- a, 1:] ^ order a p * q" | 
| 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60500diff
changeset | 256 | by simp | 
| 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60500diff
changeset | 257 | then have "\<not> [:- a, 1:] ^ order a p * [:- a, 1:] dvd [:- a, 1:] ^ order a p * q" | 
| 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60500diff
changeset | 258 | by simp | 
| 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60500diff
changeset | 259 | then have D: "\<not> [:- a, 1:] dvd q" | 
| 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60500diff
changeset | 260 | using idom_class.dvd_mult_cancel_left [of "[:- a, 1:] ^ order a p" "[:- a, 1:]" q] | 
| 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60500diff
changeset | 261 | by auto | 
| 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60500diff
changeset | 262 | from C D show ?thesis by blast | 
| 
01488b559910
avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
 haftmann parents: 
60500diff
changeset | 263 | qed | 
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 264 | |
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 265 | 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 | 266 | ==> (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 | 267 | apply (case_tac "p = 0", simp) | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 268 | 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 | 269 | using neq0_conv | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 270 | apply (blast intro: lemma_order_pderiv) | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 271 | done | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 272 | |
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 273 | 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 | 274 | proof - | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 275 | def i \<equiv> "order a p" | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 276 | def j \<equiv> "order a q" | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 277 | def t \<equiv> "[:-a, 1:]" | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 278 | 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 | 279 | 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 | 280 | assume "p * q \<noteq> 0" | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 281 | 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 | 282 | apply clarsimp | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 283 | 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 | 284 | 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 | 285 | apply clarify | 
| 56383 | 286 | apply (erule dvdE)+ | 
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 287 | apply (rule order_unique_lemma [symmetric], fold t_def) | 
| 56383 | 288 | apply (simp_all add: power_add t_dvd_iff) | 
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 289 | done | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 290 | qed | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 291 | |
| 62065 | 292 | lemma order_smult: | 
| 293 | assumes "c \<noteq> 0" | |
| 294 | shows "order x (smult c p) = order x p" | |
| 295 | proof (cases "p = 0") | |
| 296 | case False | |
| 297 | have "smult c p = [:c:] * p" by simp | |
| 298 | also from assms False have "order x \<dots> = order x [:c:] + order x p" | |
| 299 | by (subst order_mult) simp_all | |
| 300 | also from assms have "order x [:c:] = 0" by (intro order_0I) auto | |
| 301 | finally show ?thesis by simp | |
| 302 | qed simp | |
| 303 | ||
| 304 | (* Next two lemmas contributed by Wenda Li *) | |
| 305 | lemma order_1_eq_0 [simp]:"order x 1 = 0" | |
| 306 | by (metis order_root poly_1 zero_neq_one) | |
| 307 | ||
| 308 | lemma order_power_n_n: "order a ([:-a,1:]^n)=n" | |
| 309 | proof (induct n) (*might be proved more concisely using nat_less_induct*) | |
| 310 | case 0 | |
| 311 | thus ?case by (metis order_root poly_1 power_0 zero_neq_one) | |
| 312 | next | |
| 313 | case (Suc n) | |
| 314 | have "order a ([:- a, 1:] ^ Suc n)=order a ([:- a, 1:] ^ n) + order a [:-a,1:]" | |
| 315 | by (metis (no_types, hide_lams) One_nat_def add_Suc_right monoid_add_class.add.right_neutral | |
| 316 | one_neq_zero order_mult pCons_eq_0_iff power_add power_eq_0_iff power_one_right) | |
| 317 | moreover have "order a [:-a,1:]=1" unfolding order_def | |
| 318 | proof (rule Least_equality,rule ccontr) | |
| 319 | assume "\<not> \<not> [:- a, 1:] ^ Suc 1 dvd [:- a, 1:]" | |
| 320 | hence "[:- a, 1:] ^ Suc 1 dvd [:- a, 1:]" by simp | |
| 321 | hence "degree ([:- a, 1:] ^ Suc 1) \<le> degree ([:- a, 1:] )" | |
| 322 | by (rule dvd_imp_degree_le,auto) | |
| 323 | thus False by auto | |
| 324 | next | |
| 325 | fix y assume asm:"\<not> [:- a, 1:] ^ Suc y dvd [:- a, 1:]" | |
| 326 | show "1 \<le> y" | |
| 327 | proof (rule ccontr) | |
| 328 | assume "\<not> 1 \<le> y" | |
| 329 | hence "y=0" by auto | |
| 330 | hence "[:- a, 1:] ^ Suc y dvd [:- a, 1:]" by auto | |
| 331 | thus False using asm by auto | |
| 332 | qed | |
| 333 | qed | |
| 334 | ultimately show ?case using Suc by auto | |
| 335 | qed | |
| 336 | ||
| 60500 | 337 | text\<open>Now justify the standard squarefree decomposition, i.e. f / gcd(f,f').\<close> | 
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 338 | |
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 339 | 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 | 340 | apply (cases "p = 0", auto) | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 341 | apply (drule order_2 [where a=a and p=p]) | 
| 56383 | 342 | apply (metis not_less_eq_eq power_le_dvd) | 
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 343 | 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 | 344 | done | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 345 | |
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 346 | lemma poly_squarefree_decomp_order: | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 347 | assumes "pderiv p \<noteq> 0" | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 348 | and p: "p = q * d" | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 349 | and p': "pderiv p = e * d" | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 350 | 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 | 351 | 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 | 352 | proof (rule classical) | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 353 | assume 1: "order a q \<noteq> (if order a p = 0 then 0 else 1)" | 
| 60500 | 354 | from \<open>pderiv p \<noteq> 0\<close> have "p \<noteq> 0" by auto | 
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 355 | 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 | 356 | by (simp add: order_mult) | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 357 | 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 | 358 | have "order a (pderiv p) = order a e + order a d" | 
| 60500 | 359 | using \<open>pderiv p \<noteq> 0\<close> \<open>pderiv p = e * d\<close> by (simp add: order_mult) | 
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 360 | have "order a p = Suc (order a (pderiv p))" | 
| 60500 | 361 | using \<open>pderiv p \<noteq> 0\<close> \<open>order a p \<noteq> 0\<close> by (rule order_pderiv) | 
| 362 | have "d \<noteq> 0" using \<open>p \<noteq> 0\<close> \<open>p = q * d\<close> by simp | |
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 363 | 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 | 364 | apply (simp add: d) | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 365 | apply (rule dvd_add) | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 366 | apply (rule dvd_mult) | 
| 60500 | 367 | apply (simp add: order_divides \<open>p \<noteq> 0\<close> | 
| 368 | \<open>order a p = Suc (order a (pderiv p))\<close>) | |
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 369 | apply (rule dvd_mult) | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 370 | apply (simp add: order_divides) | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 371 | done | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 372 | then have "order a (pderiv p) \<le> order a d" | 
| 60500 | 373 | using \<open>d \<noteq> 0\<close> by (simp add: order_divides) | 
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 374 | show ?thesis | 
| 60500 | 375 | using \<open>order a p = order a q + order a d\<close> | 
| 376 | using \<open>order a (pderiv p) = order a e + order a d\<close> | |
| 377 | using \<open>order a p = Suc (order a (pderiv p))\<close> | |
| 378 | using \<open>order a (pderiv p) \<le> order a d\<close> | |
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 379 | by auto | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 380 | qed | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 381 | |
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 382 | 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 | 383 | p = q * d; | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 384 | pderiv p = e * d; | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 385 | d = r * p + s * pderiv p | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 386 | |] ==> \<forall>a. order a q = (if order a p = 0 then 0 else 1)" | 
| 56383 | 387 | by (blast intro: poly_squarefree_decomp_order) | 
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 388 | |
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 389 | 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 | 390 | ==> (order a (pderiv p) = n) = (order a p = Suc n)" | 
| 56383 | 391 | by (auto dest: order_pderiv) | 
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 392 | |
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 393 | definition | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 394 | rsquarefree :: "'a::idom poly => bool" where | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 395 | "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 | 396 | |
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 397 | 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 | 398 | apply (simp add: pderiv_eq_0_iff) | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 399 | 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 | 400 | done | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 401 | |
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 402 | lemma rsquarefree_roots: | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 403 | "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 | 404 | apply (simp add: rsquarefree_def) | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 405 | apply (case_tac "p = 0", simp, simp) | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 406 | apply (case_tac "pderiv p = 0") | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 407 | apply simp | 
| 56383 | 408 | apply (drule pderiv_iszero, clarsimp) | 
| 409 | apply (metis coeff_0 coeff_pCons_0 degree_pCons_0 le0 le_antisym order_degree) | |
| 410 | apply (force simp add: order_root order_pderiv2) | |
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 411 | done | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 412 | |
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 413 | lemma poly_squarefree_decomp: | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 414 | assumes "pderiv p \<noteq> 0" | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 415 | and "p = q * d" | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 416 | and "pderiv p = e * d" | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 417 | and "d = r * p + s * pderiv p" | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 418 | 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 | 419 | proof - | 
| 60500 | 420 | from \<open>pderiv p \<noteq> 0\<close> have "p \<noteq> 0" by auto | 
| 421 | with \<open>p = q * d\<close> have "q \<noteq> 0" by simp | |
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 422 | 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 | 423 | using assms by (rule poly_squarefree_decomp_order2) | 
| 60500 | 424 | with \<open>p \<noteq> 0\<close> \<open>q \<noteq> 0\<close> show ?thesis | 
| 29985 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 425 | by (simp add: rsquarefree_def order_root) | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 426 | qed | 
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 427 | |
| 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 huffman parents: diff
changeset | 428 | end |