| author | wenzelm | 
| Thu, 08 Nov 2018 22:29:09 +0100 | |
| changeset 69272 | 15e9ed5b28fb | 
| parent 68790 | 851a9d9746c6 | 
| child 71398 | e0237f2eb49d | 
| permissions | -rw-r--r-- | 
| 65435 | 1 | (* Title: HOL/Computational_Algebra/Polynomial_Factorial.thy | 
| 63764 | 2 | Author: Manuel Eberl | 
| 3 | *) | |
| 4 | ||
| 66805 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 5 | section \<open>Polynomials, fractions and rings\<close> | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 6 | |
| 63498 | 7 | theory Polynomial_Factorial | 
| 8 | imports | |
| 9 | Complex_Main | |
| 65366 | 10 | Polynomial | 
| 11 | Normalized_Fraction | |
| 63498 | 12 | begin | 
| 13 | ||
| 14 | subsection \<open>Lifting elements into the field of fractions\<close> | |
| 15 | ||
| 66805 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 16 | definition to_fract :: "'a :: idom \<Rightarrow> 'a fract" | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 17 | where "to_fract x = Fract x 1" | 
| 66808 
1907167b6038
elementary definition of division on natural numbers
 haftmann parents: 
66806diff
changeset | 18 | \<comment> \<open>FIXME: more idiomatic name, abbreviation\<close> | 
| 63498 | 19 | |
| 20 | lemma to_fract_0 [simp]: "to_fract 0 = 0" | |
| 21 | by (simp add: to_fract_def eq_fract Zero_fract_def) | |
| 22 | ||
| 23 | lemma to_fract_1 [simp]: "to_fract 1 = 1" | |
| 24 | by (simp add: to_fract_def eq_fract One_fract_def) | |
| 25 | ||
| 26 | lemma to_fract_add [simp]: "to_fract (x + y) = to_fract x + to_fract y" | |
| 27 | by (simp add: to_fract_def) | |
| 28 | ||
| 29 | lemma to_fract_diff [simp]: "to_fract (x - y) = to_fract x - to_fract y" | |
| 30 | by (simp add: to_fract_def) | |
| 31 | ||
| 32 | lemma to_fract_uminus [simp]: "to_fract (-x) = -to_fract x" | |
| 33 | by (simp add: to_fract_def) | |
| 34 | ||
| 35 | lemma to_fract_mult [simp]: "to_fract (x * y) = to_fract x * to_fract y" | |
| 36 | by (simp add: to_fract_def) | |
| 37 | ||
| 38 | lemma to_fract_eq_iff [simp]: "to_fract x = to_fract y \<longleftrightarrow> x = y" | |
| 39 | by (simp add: to_fract_def eq_fract) | |
| 40 | ||
| 41 | lemma to_fract_eq_0_iff [simp]: "to_fract x = 0 \<longleftrightarrow> x = 0" | |
| 42 | by (simp add: to_fract_def Zero_fract_def eq_fract) | |
| 43 | ||
| 44 | lemma snd_quot_of_fract_nonzero [simp]: "snd (quot_of_fract x) \<noteq> 0" | |
| 45 | by transfer simp | |
| 46 | ||
| 47 | lemma Fract_quot_of_fract [simp]: "Fract (fst (quot_of_fract x)) (snd (quot_of_fract x)) = x" | |
| 48 | by transfer (simp del: fractrel_iff, subst fractrel_normalize_quot_left, simp) | |
| 49 | ||
| 50 | lemma to_fract_quot_of_fract: | |
| 51 | assumes "snd (quot_of_fract x) = 1" | |
| 52 | shows "to_fract (fst (quot_of_fract x)) = x" | |
| 53 | proof - | |
| 54 | have "x = Fract (fst (quot_of_fract x)) (snd (quot_of_fract x))" by simp | |
| 55 | also note assms | |
| 56 | finally show ?thesis by (simp add: to_fract_def) | |
| 57 | qed | |
| 58 | ||
| 59 | lemma snd_quot_of_fract_Fract_whole: | |
| 60 | assumes "y dvd x" | |
| 61 | shows "snd (quot_of_fract (Fract x y)) = 1" | |
| 62 | using assms by transfer (auto simp: normalize_quot_def Let_def gcd_proj2_if_dvd) | |
| 63 | ||
| 64 | lemma Fract_conv_to_fract: "Fract a b = to_fract a / to_fract b" | |
| 65 | by (simp add: to_fract_def) | |
| 66 | ||
| 67 | lemma quot_of_fract_to_fract [simp]: "quot_of_fract (to_fract x) = (x, 1)" | |
| 68 | unfolding to_fract_def by transfer (simp add: normalize_quot_def) | |
| 69 | ||
| 70 | lemma fst_quot_of_fract_eq_0_iff [simp]: "fst (quot_of_fract x) = 0 \<longleftrightarrow> x = 0" | |
| 71 | by transfer simp | |
| 72 | ||
| 73 | lemma snd_quot_of_fract_to_fract [simp]: "snd (quot_of_fract (to_fract x)) = 1" | |
| 74 | unfolding to_fract_def by (rule snd_quot_of_fract_Fract_whole) simp_all | |
| 75 | ||
| 76 | lemma coprime_quot_of_fract: | |
| 77 | "coprime (fst (quot_of_fract x)) (snd (quot_of_fract x))" | |
| 78 | by transfer (simp add: coprime_normalize_quot) | |
| 79 | ||
| 80 | lemma unit_factor_snd_quot_of_fract: "unit_factor (snd (quot_of_fract x)) = 1" | |
| 81 | using quot_of_fract_in_normalized_fracts[of x] | |
| 82 | by (simp add: normalized_fracts_def case_prod_unfold) | |
| 83 | ||
| 84 | lemma unit_factor_1_imp_normalized: "unit_factor x = 1 \<Longrightarrow> normalize x = x" | |
| 85 | by (subst (2) normalize_mult_unit_factor [symmetric, of x]) | |
| 86 | (simp del: normalize_mult_unit_factor) | |
| 87 | ||
| 88 | lemma normalize_snd_quot_of_fract: "normalize (snd (quot_of_fract x)) = snd (quot_of_fract x)" | |
| 89 | by (intro unit_factor_1_imp_normalized unit_factor_snd_quot_of_fract) | |
| 90 | ||
| 91 | ||
| 92 | subsection \<open>Lifting polynomial coefficients to the field of fractions\<close> | |
| 93 | ||
| 94 | abbreviation (input) fract_poly | |
| 95 | where "fract_poly \<equiv> map_poly to_fract" | |
| 96 | ||
| 97 | abbreviation (input) unfract_poly | |
| 98 | where "unfract_poly \<equiv> map_poly (fst \<circ> quot_of_fract)" | |
| 99 | ||
| 100 | lemma fract_poly_smult [simp]: "fract_poly (smult c p) = smult (to_fract c) (fract_poly p)" | |
| 101 | by (simp add: smult_conv_map_poly map_poly_map_poly o_def) | |
| 102 | ||
| 103 | lemma fract_poly_0 [simp]: "fract_poly 0 = 0" | |
| 104 | by (simp add: poly_eqI coeff_map_poly) | |
| 105 | ||
| 106 | lemma fract_poly_1 [simp]: "fract_poly 1 = 1" | |
| 65486 | 107 | by (simp add: map_poly_pCons) | 
| 63498 | 108 | |
| 109 | lemma fract_poly_add [simp]: | |
| 110 | "fract_poly (p + q) = fract_poly p + fract_poly q" | |
| 111 | by (intro poly_eqI) (simp_all add: coeff_map_poly) | |
| 112 | ||
| 113 | lemma fract_poly_diff [simp]: | |
| 114 | "fract_poly (p - q) = fract_poly p - fract_poly q" | |
| 115 | by (intro poly_eqI) (simp_all add: coeff_map_poly) | |
| 116 | ||
| 64267 | 117 | lemma to_fract_sum [simp]: "to_fract (sum f A) = sum (\<lambda>x. to_fract (f x)) A" | 
| 63498 | 118 | by (cases "finite A", induction A rule: finite_induct) simp_all | 
| 119 | ||
| 120 | lemma fract_poly_mult [simp]: | |
| 121 | "fract_poly (p * q) = fract_poly p * fract_poly q" | |
| 122 | by (intro poly_eqI) (simp_all add: coeff_map_poly coeff_mult) | |
| 123 | ||
| 124 | lemma fract_poly_eq_iff [simp]: "fract_poly p = fract_poly q \<longleftrightarrow> p = q" | |
| 125 | by (auto simp: poly_eq_iff coeff_map_poly) | |
| 126 | ||
| 127 | lemma fract_poly_eq_0_iff [simp]: "fract_poly p = 0 \<longleftrightarrow> p = 0" | |
| 128 | using fract_poly_eq_iff[of p 0] by (simp del: fract_poly_eq_iff) | |
| 129 | ||
| 130 | lemma fract_poly_dvd: "p dvd q \<Longrightarrow> fract_poly p dvd fract_poly q" | |
| 131 | by (auto elim!: dvdE) | |
| 132 | ||
| 63830 | 133 | lemma prod_mset_fract_poly: | 
| 65390 | 134 | "(\<Prod>x\<in>#A. map_poly to_fract (f x)) = fract_poly (prod_mset (image_mset f A))" | 
| 135 | by (induct A) (simp_all add: ac_simps) | |
| 63498 | 136 | |
| 137 | lemma is_unit_fract_poly_iff: | |
| 138 | "p dvd 1 \<longleftrightarrow> fract_poly p dvd 1 \<and> content p = 1" | |
| 139 | proof safe | |
| 140 | assume A: "p dvd 1" | |
| 65389 | 141 | with fract_poly_dvd [of p 1] show "is_unit (fract_poly p)" | 
| 142 | by simp | |
| 63498 | 143 | from A show "content p = 1" | 
| 144 | by (auto simp: is_unit_poly_iff normalize_1_iff) | |
| 145 | next | |
| 146 | assume A: "fract_poly p dvd 1" and B: "content p = 1" | |
| 147 | from A obtain c where c: "fract_poly p = [:c:]" by (auto simp: is_unit_poly_iff) | |
| 148 |   {
 | |
| 149 | fix n :: nat assume "n > 0" | |
| 150 | have "to_fract (coeff p n) = coeff (fract_poly p) n" by (simp add: coeff_map_poly) | |
| 151 | also note c | |
| 152 | also from \<open>n > 0\<close> have "coeff [:c:] n = 0" by (simp add: coeff_pCons split: nat.splits) | |
| 153 | finally have "coeff p n = 0" by simp | |
| 154 | } | |
| 155 | hence "degree p \<le> 0" by (intro degree_le) simp_all | |
| 156 | with B show "p dvd 1" by (auto simp: is_unit_poly_iff normalize_1_iff elim!: degree_eq_zeroE) | |
| 157 | qed | |
| 158 | ||
| 159 | lemma fract_poly_is_unit: "p dvd 1 \<Longrightarrow> fract_poly p dvd 1" | |
| 160 | using fract_poly_dvd[of p 1] by simp | |
| 161 | ||
| 162 | lemma fract_poly_smult_eqE: | |
| 163 |   fixes c :: "'a :: {idom_divide,ring_gcd} fract"
 | |
| 164 | assumes "fract_poly p = smult c (fract_poly q)" | |
| 165 | obtains a b | |
| 166 | where "c = to_fract b / to_fract a" "smult a p = smult b q" "coprime a b" "normalize a = a" | |
| 167 | proof - | |
| 168 | define a b where "a = fst (quot_of_fract c)" and "b = snd (quot_of_fract c)" | |
| 169 | have "smult (to_fract a) (fract_poly q) = smult (to_fract b) (fract_poly p)" | |
| 170 | by (subst smult_eq_iff) (simp_all add: a_def b_def Fract_conv_to_fract [symmetric] assms) | |
| 171 | hence "fract_poly (smult a q) = fract_poly (smult b p)" by (simp del: fract_poly_eq_iff) | |
| 172 | hence "smult b p = smult a q" by (simp only: fract_poly_eq_iff) | |
| 173 | moreover have "c = to_fract a / to_fract b" "coprime b a" "normalize b = b" | |
| 67051 | 174 | by (simp_all add: a_def b_def coprime_quot_of_fract [of c] ac_simps | 
| 63498 | 175 | normalize_snd_quot_of_fract Fract_conv_to_fract [symmetric]) | 
| 176 | ultimately show ?thesis by (intro that[of a b]) | |
| 177 | qed | |
| 178 | ||
| 179 | ||
| 180 | subsection \<open>Fractional content\<close> | |
| 181 | ||
| 182 | abbreviation (input) Lcm_coeff_denoms | |
| 183 |     :: "'a :: {semiring_Gcd,idom_divide,ring_gcd} fract poly \<Rightarrow> 'a"
 | |
| 184 | where "Lcm_coeff_denoms p \<equiv> Lcm (snd ` quot_of_fract ` set (coeffs p))" | |
| 185 | ||
| 186 | definition fract_content :: | |
| 187 |       "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} fract poly \<Rightarrow> 'a fract" where
 | |
| 188 | "fract_content p = | |
| 189 | (let d = Lcm_coeff_denoms p in Fract (content (unfract_poly (smult (to_fract d) p))) d)" | |
| 190 | ||
| 191 | definition primitive_part_fract :: | |
| 192 |       "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} fract poly \<Rightarrow> 'a poly" where
 | |
| 193 | "primitive_part_fract p = | |
| 194 | primitive_part (unfract_poly (smult (to_fract (Lcm_coeff_denoms p)) p))" | |
| 195 | ||
| 196 | lemma primitive_part_fract_0 [simp]: "primitive_part_fract 0 = 0" | |
| 197 | by (simp add: primitive_part_fract_def) | |
| 198 | ||
| 199 | lemma fract_content_eq_0_iff [simp]: | |
| 200 | "fract_content p = 0 \<longleftrightarrow> p = 0" | |
| 201 | unfolding fract_content_def Let_def Zero_fract_def | |
| 202 | by (subst eq_fract) (auto simp: Lcm_0_iff map_poly_eq_0_iff) | |
| 203 | ||
| 204 | lemma content_primitive_part_fract [simp]: "p \<noteq> 0 \<Longrightarrow> content (primitive_part_fract p) = 1" | |
| 205 | unfolding primitive_part_fract_def | |
| 206 | by (rule content_primitive_part) | |
| 207 | (auto simp: primitive_part_fract_def map_poly_eq_0_iff Lcm_0_iff) | |
| 208 | ||
| 209 | lemma content_times_primitive_part_fract: | |
| 210 | "smult (fract_content p) (fract_poly (primitive_part_fract p)) = p" | |
| 211 | proof - | |
| 212 | define p' where "p' = unfract_poly (smult (to_fract (Lcm_coeff_denoms p)) p)" | |
| 213 | have "fract_poly p' = | |
| 214 | map_poly (to_fract \<circ> fst \<circ> quot_of_fract) (smult (to_fract (Lcm_coeff_denoms p)) p)" | |
| 215 | unfolding primitive_part_fract_def p'_def | |
| 216 | by (subst map_poly_map_poly) (simp_all add: o_assoc) | |
| 217 | also have "\<dots> = smult (to_fract (Lcm_coeff_denoms p)) p" | |
| 218 | proof (intro map_poly_idI, unfold o_apply) | |
| 219 | fix c assume "c \<in> set (coeffs (smult (to_fract (Lcm_coeff_denoms p)) p))" | |
| 220 | then obtain c' where c: "c' \<in> set (coeffs p)" "c = to_fract (Lcm_coeff_denoms p) * c'" | |
| 221 | by (auto simp add: Lcm_0_iff coeffs_smult split: if_splits) | |
| 222 | note c(2) | |
| 223 | also have "c' = Fract (fst (quot_of_fract c')) (snd (quot_of_fract c'))" | |
| 224 | by simp | |
| 225 | also have "to_fract (Lcm_coeff_denoms p) * \<dots> = | |
| 226 | Fract (Lcm_coeff_denoms p * fst (quot_of_fract c')) (snd (quot_of_fract c'))" | |
| 227 | unfolding to_fract_def by (subst mult_fract) simp_all | |
| 228 | also have "snd (quot_of_fract \<dots>) = 1" | |
| 229 | by (intro snd_quot_of_fract_Fract_whole dvd_mult2 dvd_Lcm) (insert c(1), auto) | |
| 230 | finally show "to_fract (fst (quot_of_fract c)) = c" | |
| 231 | by (rule to_fract_quot_of_fract) | |
| 232 | qed | |
| 233 | also have "p' = smult (content p') (primitive_part p')" | |
| 234 | by (rule content_times_primitive_part [symmetric]) | |
| 235 | also have "primitive_part p' = primitive_part_fract p" | |
| 236 | by (simp add: primitive_part_fract_def p'_def) | |
| 237 | also have "fract_poly (smult (content p') (primitive_part_fract p)) = | |
| 238 | smult (to_fract (content p')) (fract_poly (primitive_part_fract p))" by simp | |
| 239 | finally have "smult (to_fract (content p')) (fract_poly (primitive_part_fract p)) = | |
| 240 | smult (to_fract (Lcm_coeff_denoms p)) p" . | |
| 241 | thus ?thesis | |
| 242 | by (subst (asm) smult_eq_iff) | |
| 243 | (auto simp add: Let_def p'_def Fract_conv_to_fract field_simps Lcm_0_iff fract_content_def) | |
| 244 | qed | |
| 245 | ||
| 246 | lemma fract_content_fract_poly [simp]: "fract_content (fract_poly p) = to_fract (content p)" | |
| 247 | proof - | |
| 248 | have "Lcm_coeff_denoms (fract_poly p) = 1" | |
| 63905 | 249 | by (auto simp: set_coeffs_map_poly) | 
| 63498 | 250 | hence "fract_content (fract_poly p) = | 
| 251 | to_fract (content (map_poly (fst \<circ> quot_of_fract \<circ> to_fract) p))" | |
| 252 | by (simp add: fract_content_def to_fract_def fract_collapse map_poly_map_poly del: Lcm_1_iff) | |
| 253 | also have "map_poly (fst \<circ> quot_of_fract \<circ> to_fract) p = p" | |
| 254 | by (intro map_poly_idI) simp_all | |
| 255 | finally show ?thesis . | |
| 256 | qed | |
| 257 | ||
| 258 | lemma content_decompose_fract: | |
| 259 |   fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} fract poly"
 | |
| 260 | obtains c p' where "p = smult c (map_poly to_fract p')" "content p' = 1" | |
| 261 | proof (cases "p = 0") | |
| 262 | case True | |
| 263 | hence "p = smult 0 (map_poly to_fract 1)" "content 1 = 1" by simp_all | |
| 264 | thus ?thesis .. | |
| 265 | next | |
| 266 | case False | |
| 267 | thus ?thesis | |
| 268 | by (rule that[OF content_times_primitive_part_fract [symmetric] content_primitive_part_fract]) | |
| 269 | qed | |
| 270 | ||
| 271 | ||
| 272 | subsection \<open>More properties of content and primitive part\<close> | |
| 273 | ||
| 274 | lemma lift_prime_elem_poly: | |
| 63633 | 275 | assumes "prime_elem (c :: 'a :: semidom)" | 
| 276 | shows "prime_elem [:c:]" | |
| 277 | proof (rule prime_elemI) | |
| 63498 | 278 | fix a b assume *: "[:c:] dvd a * b" | 
| 279 | from * have dvd: "c dvd coeff (a * b) n" for n | |
| 280 | by (subst (asm) const_poly_dvd_iff) blast | |
| 281 |   {
 | |
| 282 | define m where "m = (GREATEST m. \<not>c dvd coeff b m)" | |
| 283 | assume "\<not>[:c:] dvd b" | |
| 284 | hence A: "\<exists>i. \<not>c dvd coeff b i" by (subst (asm) const_poly_dvd_iff) blast | |
| 65963 | 285 | have B: "\<forall>i. \<not>c dvd coeff b i \<longrightarrow> i \<le> degree b" | 
| 286 | by (auto intro: le_degree) | |
| 65965 | 287 | have coeff_m: "\<not>c dvd coeff b m" unfolding m_def by (rule GreatestI_ex_nat[OF A B]) | 
| 63498 | 288 | have "i \<le> m" if "\<not>c dvd coeff b i" for i | 
| 65965 | 289 | unfolding m_def by (rule Greatest_le_nat[OF that B]) | 
| 63498 | 290 | hence dvd_b: "c dvd coeff b i" if "i > m" for i using that by force | 
| 291 | ||
| 292 | have "c dvd coeff a i" for i | |
| 293 | proof (induction i rule: nat_descend_induct[of "degree a"]) | |
| 294 | case (base i) | |
| 295 | thus ?case by (simp add: coeff_eq_0) | |
| 296 | next | |
| 297 | case (descend i) | |
| 298 |       let ?A = "{..i+m} - {i}"
 | |
| 299 | have "c dvd coeff (a * b) (i + m)" by (rule dvd) | |
| 300 | also have "coeff (a * b) (i + m) = (\<Sum>k\<le>i + m. coeff a k * coeff b (i + m - k))" | |
| 301 | by (simp add: coeff_mult) | |
| 302 |       also have "{..i+m} = insert i ?A" by auto
 | |
| 303 | also have "(\<Sum>k\<in>\<dots>. coeff a k * coeff b (i + m - k)) = | |
| 304 | coeff a i * coeff b m + (\<Sum>k\<in>?A. coeff a k * coeff b (i + m - k))" | |
| 305 | (is "_ = _ + ?S") | |
| 64267 | 306 | by (subst sum.insert) simp_all | 
| 63498 | 307 | finally have eq: "c dvd coeff a i * coeff b m + ?S" . | 
| 308 | moreover have "c dvd ?S" | |
| 64267 | 309 | proof (rule dvd_sum) | 
| 63498 | 310 |         fix k assume k: "k \<in> {..i+m} - {i}"
 | 
| 311 | show "c dvd coeff a k * coeff b (i + m - k)" | |
| 312 | proof (cases "k < i") | |
| 313 | case False | |
| 314 | with k have "c dvd coeff a k" by (intro descend.IH) simp | |
| 315 | thus ?thesis by simp | |
| 316 | next | |
| 317 | case True | |
| 318 | hence "c dvd coeff b (i + m - k)" by (intro dvd_b) simp | |
| 319 | thus ?thesis by simp | |
| 320 | qed | |
| 321 | qed | |
| 322 | ultimately have "c dvd coeff a i * coeff b m" | |
| 323 | by (simp add: dvd_add_left_iff) | |
| 324 | with assms coeff_m show "c dvd coeff a i" | |
| 63633 | 325 | by (simp add: prime_elem_dvd_mult_iff) | 
| 63498 | 326 | qed | 
| 327 | hence "[:c:] dvd a" by (subst const_poly_dvd_iff) blast | |
| 328 | } | |
| 65486 | 329 | then show "[:c:] dvd a \<or> [:c:] dvd b" by blast | 
| 330 | next | |
| 331 | from assms show "[:c:] \<noteq> 0" and "\<not> [:c:] dvd 1" | |
| 332 | by (simp_all add: prime_elem_def is_unit_const_poly_iff) | |
| 333 | qed | |
| 63498 | 334 | |
| 335 | lemma prime_elem_const_poly_iff: | |
| 336 | fixes c :: "'a :: semidom" | |
| 63633 | 337 | shows "prime_elem [:c:] \<longleftrightarrow> prime_elem c" | 
| 63498 | 338 | proof | 
| 63633 | 339 | assume A: "prime_elem [:c:]" | 
| 340 | show "prime_elem c" | |
| 341 | proof (rule prime_elemI) | |
| 63498 | 342 | fix a b assume "c dvd a * b" | 
| 343 | hence "[:c:] dvd [:a:] * [:b:]" by (simp add: mult_ac) | |
| 63633 | 344 | from A and this have "[:c:] dvd [:a:] \<or> [:c:] dvd [:b:]" by (rule prime_elem_dvd_multD) | 
| 63498 | 345 | thus "c dvd a \<or> c dvd b" by simp | 
| 63633 | 346 | qed (insert A, auto simp: prime_elem_def is_unit_poly_iff) | 
| 63498 | 347 | qed (auto intro: lift_prime_elem_poly) | 
| 348 | ||
| 349 | lemma fract_poly_dvdD: | |
| 350 |   fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} poly"
 | |
| 351 | assumes "fract_poly p dvd fract_poly q" "content p = 1" | |
| 352 | shows "p dvd q" | |
| 353 | proof - | |
| 354 | from assms(1) obtain r where r: "fract_poly q = fract_poly p * r" by (erule dvdE) | |
| 355 | from content_decompose_fract[of r] guess c r' . note r' = this | |
| 356 | from r r' have eq: "fract_poly q = smult c (fract_poly (p * r'))" by simp | |
| 357 | from fract_poly_smult_eqE[OF this] guess a b . note ab = this | |
| 358 | have "content (smult a q) = content (smult b (p * r'))" by (simp only: ab(2)) | |
| 359 | hence eq': "normalize b = a * content q" by (simp add: assms content_mult r' ab(4)) | |
| 360 | have "1 = gcd a (normalize b)" by (simp add: ab) | |
| 361 | also note eq' | |
| 362 | also have "gcd a (a * content q) = a" by (simp add: gcd_proj1_if_dvd ab(4)) | |
| 363 | finally have [simp]: "a = 1" by simp | |
| 364 | from eq ab have "q = p * ([:b:] * r')" by simp | |
| 365 | thus ?thesis by (rule dvdI) | |
| 366 | qed | |
| 367 | ||
| 368 | ||
| 369 | subsection \<open>Polynomials over a field are a Euclidean ring\<close> | |
| 370 | ||
| 66805 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 371 | context | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 372 | begin | 
| 63498 | 373 | |
| 374 | interpretation field_poly: | |
| 66817 | 375 | normalization_euclidean_semiring where zero = "0 :: 'a :: field poly" | 
| 376 | and one = 1 and plus = plus and minus = minus | |
| 64164 
38c407446400
separate type class for arbitrary quotient and remainder partitions
 haftmann parents: 
63954diff
changeset | 377 | and times = times | 
| 66805 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 378 | and normalize = "\<lambda>p. smult (inverse (lead_coeff p)) p" | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 379 | and unit_factor = "\<lambda>p. [:lead_coeff p:]" | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 380 | and euclidean_size = "\<lambda>p. if p = 0 then 0 else 2 ^ degree p" | 
| 64164 
38c407446400
separate type class for arbitrary quotient and remainder partitions
 haftmann parents: 
63954diff
changeset | 381 | and divide = divide and modulo = modulo | 
| 66805 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 382 | rewrites "dvd.dvd (times :: 'a poly \<Rightarrow> _) = Rings.dvd" | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 383 | and "comm_monoid_mult.prod_mset times 1 = prod_mset" | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 384 | and "comm_semiring_1.irreducible times 1 0 = irreducible" | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 385 | and "comm_semiring_1.prime_elem times 1 0 = prime_elem" | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 386 | proof - | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 387 | show "dvd.dvd (times :: 'a poly \<Rightarrow> _) = Rings.dvd" | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 388 | by (simp add: dvd_dict) | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 389 | show "comm_monoid_mult.prod_mset times 1 = prod_mset" | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 390 | by (simp add: prod_mset_dict) | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 391 | show "comm_semiring_1.irreducible times 1 0 = irreducible" | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 392 | by (simp add: irreducible_dict) | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 393 | show "comm_semiring_1.prime_elem times 1 0 = prime_elem" | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 394 | by (simp add: prime_elem_dict) | 
| 66817 | 395 | show "class.normalization_euclidean_semiring divide plus minus (0 :: 'a poly) times 1 | 
| 396 | modulo (\<lambda>p. if p = 0 then 0 else 2 ^ degree p) | |
| 397 | (\<lambda>p. [:lead_coeff p:]) (\<lambda>p. smult (inverse (lead_coeff p)) p)" | |
| 66805 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 398 | proof (standard, fold dvd_dict) | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 399 | fix p :: "'a poly" | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 400 | show "[:lead_coeff p:] * smult (inverse (lead_coeff p)) p = p" | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 401 | by (cases "p = 0") simp_all | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 402 | next | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 403 | fix p :: "'a poly" assume "is_unit p" | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 404 | then show "[:lead_coeff p:] = p" | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 405 | by (elim is_unit_polyE) (auto simp: monom_0 one_poly_def field_simps) | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 406 | next | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 407 | fix p :: "'a poly" assume "p \<noteq> 0" | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 408 | then show "is_unit [:lead_coeff p:]" | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 409 | by (simp add: is_unit_pCons_iff) | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 410 | qed (auto simp: lead_coeff_mult Rings.div_mult_mod_eq intro!: degree_mod_less' degree_mult_right_le) | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 411 | qed | 
| 63498 | 412 | |
| 63722 
b9c8da46443b
Deprivatisation of lemmas in Polynomial_Factorial
 Manuel Eberl <eberlm@in.tum.de> parents: 
63705diff
changeset | 413 | lemma field_poly_irreducible_imp_prime: | 
| 66805 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 414 | "prime_elem p" if "irreducible p" for p :: "'a :: field poly" | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 415 | using that by (fact field_poly.irreducible_imp_prime_elem) | 
| 63498 | 416 | |
| 63830 | 417 | lemma field_poly_prod_mset_prime_factorization: | 
| 66805 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 418 | "prod_mset (field_poly.prime_factorization p) = smult (inverse (lead_coeff p)) p" | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 419 | if "p \<noteq> 0" for p :: "'a :: field poly" | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 420 | using that by (fact field_poly.prod_mset_prime_factorization) | 
| 63498 | 421 | |
| 63722 
b9c8da46443b
Deprivatisation of lemmas in Polynomial_Factorial
 Manuel Eberl <eberlm@in.tum.de> parents: 
63705diff
changeset | 422 | lemma field_poly_in_prime_factorization_imp_prime: | 
| 66805 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 423 | "prime_elem p" if "p \<in># field_poly.prime_factorization x" | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 424 | for p :: "'a :: field poly" | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 425 | by (rule field_poly.prime_imp_prime_elem, rule field_poly.in_prime_factors_imp_prime) | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 426 | (fact that) | 
| 63498 | 427 | |
| 428 | ||
| 429 | subsection \<open>Primality and irreducibility in polynomial rings\<close> | |
| 430 | ||
| 431 | lemma nonconst_poly_irreducible_iff: | |
| 432 |   fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} poly"
 | |
| 433 | assumes "degree p \<noteq> 0" | |
| 434 | shows "irreducible p \<longleftrightarrow> irreducible (fract_poly p) \<and> content p = 1" | |
| 435 | proof safe | |
| 436 | assume p: "irreducible p" | |
| 437 | ||
| 438 | from content_decompose[of p] guess p' . note p' = this | |
| 439 | hence "p = [:content p:] * p'" by simp | |
| 440 | from p this have "[:content p:] dvd 1 \<or> p' dvd 1" by (rule irreducibleD) | |
| 441 | moreover have "\<not>p' dvd 1" | |
| 442 | proof | |
| 443 | assume "p' dvd 1" | |
| 444 | hence "degree p = 0" by (subst p') (auto simp: is_unit_poly_iff) | |
| 445 | with assms show False by contradiction | |
| 446 | qed | |
| 447 | ultimately show [simp]: "content p = 1" by (simp add: is_unit_const_poly_iff) | |
| 448 | ||
| 449 | show "irreducible (map_poly to_fract p)" | |
| 450 | proof (rule irreducibleI) | |
| 451 | have "fract_poly p = 0 \<longleftrightarrow> p = 0" by (intro map_poly_eq_0_iff) auto | |
| 452 | with assms show "map_poly to_fract p \<noteq> 0" by auto | |
| 453 | next | |
| 454 | show "\<not>is_unit (fract_poly p)" | |
| 455 | proof | |
| 456 | assume "is_unit (map_poly to_fract p)" | |
| 457 | hence "degree (map_poly to_fract p) = 0" | |
| 458 | by (auto simp: is_unit_poly_iff) | |
| 459 | hence "degree p = 0" by (simp add: degree_map_poly) | |
| 460 | with assms show False by contradiction | |
| 461 | qed | |
| 462 | next | |
| 463 | fix q r assume qr: "fract_poly p = q * r" | |
| 464 | from content_decompose_fract[of q] guess cg q' . note q = this | |
| 465 | from content_decompose_fract[of r] guess cr r' . note r = this | |
| 466 | from qr q r p have nz: "cg \<noteq> 0" "cr \<noteq> 0" by auto | |
| 467 | from qr have eq: "fract_poly p = smult (cr * cg) (fract_poly (q' * r'))" | |
| 468 | by (simp add: q r) | |
| 469 | from fract_poly_smult_eqE[OF this] guess a b . note ab = this | |
| 470 | hence "content (smult a p) = content (smult b (q' * r'))" by (simp only:) | |
| 471 | with ab(4) have a: "a = normalize b" by (simp add: content_mult q r) | |
| 67051 | 472 | then have "normalize b = gcd a b" | 
| 473 | by simp | |
| 474 | with \<open>coprime a b\<close> have "normalize b = 1" | |
| 475 | by simp | |
| 476 | then have "a = 1" "is_unit b" | |
| 477 | by (simp_all add: a normalize_1_iff) | |
| 63498 | 478 | |
| 479 | note eq | |
| 480 | also from ab(1) \<open>a = 1\<close> have "cr * cg = to_fract b" by simp | |
| 481 | also have "smult \<dots> (fract_poly (q' * r')) = fract_poly (smult b (q' * r'))" by simp | |
| 482 | finally have "p = ([:b:] * q') * r'" by (simp del: fract_poly_smult) | |
| 483 | from p and this have "([:b:] * q') dvd 1 \<or> r' dvd 1" by (rule irreducibleD) | |
| 484 | hence "q' dvd 1 \<or> r' dvd 1" by (auto dest: dvd_mult_right simp del: mult_pCons_left) | |
| 485 | hence "fract_poly q' dvd 1 \<or> fract_poly r' dvd 1" by (auto simp: fract_poly_is_unit) | |
| 486 | with q r show "is_unit q \<or> is_unit r" | |
| 487 | by (auto simp add: is_unit_smult_iff dvd_field_iff nz) | |
| 488 | qed | |
| 489 | ||
| 490 | next | |
| 491 | ||
| 492 | assume irred: "irreducible (fract_poly p)" and primitive: "content p = 1" | |
| 493 | show "irreducible p" | |
| 494 | proof (rule irreducibleI) | |
| 495 | from irred show "p \<noteq> 0" by auto | |
| 496 | next | |
| 497 | from irred show "\<not>p dvd 1" | |
| 498 | by (auto simp: irreducible_def dest: fract_poly_is_unit) | |
| 499 | next | |
| 500 | fix q r assume qr: "p = q * r" | |
| 501 | hence "fract_poly p = fract_poly q * fract_poly r" by simp | |
| 502 | from irred and this have "fract_poly q dvd 1 \<or> fract_poly r dvd 1" | |
| 503 | by (rule irreducibleD) | |
| 504 | with primitive qr show "q dvd 1 \<or> r dvd 1" | |
| 505 | by (auto simp: content_prod_eq_1_iff is_unit_fract_poly_iff) | |
| 506 | qed | |
| 507 | qed | |
| 508 | ||
| 509 | private lemma irreducible_imp_prime_poly: | |
| 510 |   fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} poly"
 | |
| 511 | assumes "irreducible p" | |
| 63633 | 512 | shows "prime_elem p" | 
| 63498 | 513 | proof (cases "degree p = 0") | 
| 514 | case True | |
| 515 | with assms show ?thesis | |
| 516 | by (auto simp: prime_elem_const_poly_iff irreducible_const_poly_iff | |
| 63633 | 517 | intro!: irreducible_imp_prime_elem elim!: degree_eq_zeroE) | 
| 63498 | 518 | next | 
| 519 | case False | |
| 520 | from assms False have irred: "irreducible (fract_poly p)" and primitive: "content p = 1" | |
| 521 | by (simp_all add: nonconst_poly_irreducible_iff) | |
| 63633 | 522 | from irred have prime: "prime_elem (fract_poly p)" by (rule field_poly_irreducible_imp_prime) | 
| 63498 | 523 | show ?thesis | 
| 63633 | 524 | proof (rule prime_elemI) | 
| 63498 | 525 | fix q r assume "p dvd q * r" | 
| 526 | hence "fract_poly p dvd fract_poly (q * r)" by (rule fract_poly_dvd) | |
| 527 | hence "fract_poly p dvd fract_poly q * fract_poly r" by simp | |
| 528 | from prime and this have "fract_poly p dvd fract_poly q \<or> fract_poly p dvd fract_poly r" | |
| 63633 | 529 | by (rule prime_elem_dvd_multD) | 
| 63498 | 530 | with primitive show "p dvd q \<or> p dvd r" by (auto dest: fract_poly_dvdD) | 
| 531 | qed (insert assms, auto simp: irreducible_def) | |
| 532 | qed | |
| 533 | ||
| 534 | lemma degree_primitive_part_fract [simp]: | |
| 535 | "degree (primitive_part_fract p) = degree p" | |
| 536 | proof - | |
| 537 | have "p = smult (fract_content p) (fract_poly (primitive_part_fract p))" | |
| 538 | by (simp add: content_times_primitive_part_fract) | |
| 539 | also have "degree \<dots> = degree (primitive_part_fract p)" | |
| 540 | by (auto simp: degree_map_poly) | |
| 541 | finally show ?thesis .. | |
| 542 | qed | |
| 543 | ||
| 544 | lemma irreducible_primitive_part_fract: | |
| 545 |   fixes p :: "'a :: {idom_divide, ring_gcd, factorial_semiring, semiring_Gcd} fract poly"
 | |
| 546 | assumes "irreducible p" | |
| 547 | shows "irreducible (primitive_part_fract p)" | |
| 548 | proof - | |
| 549 | from assms have deg: "degree (primitive_part_fract p) \<noteq> 0" | |
| 550 | by (intro notI) | |
| 551 | (auto elim!: degree_eq_zeroE simp: irreducible_def is_unit_poly_iff dvd_field_iff) | |
| 552 | hence [simp]: "p \<noteq> 0" by auto | |
| 553 | ||
| 554 | note \<open>irreducible p\<close> | |
| 555 | also have "p = [:fract_content p:] * fract_poly (primitive_part_fract p)" | |
| 556 | by (simp add: content_times_primitive_part_fract) | |
| 557 | also have "irreducible \<dots> \<longleftrightarrow> irreducible (fract_poly (primitive_part_fract p))" | |
| 558 | by (intro irreducible_mult_unit_left) (simp_all add: is_unit_poly_iff dvd_field_iff) | |
| 559 | finally show ?thesis using deg | |
| 560 | by (simp add: nonconst_poly_irreducible_iff) | |
| 561 | qed | |
| 562 | ||
| 63633 | 563 | lemma prime_elem_primitive_part_fract: | 
| 63498 | 564 |   fixes p :: "'a :: {idom_divide, ring_gcd, factorial_semiring, semiring_Gcd} fract poly"
 | 
| 63633 | 565 | shows "irreducible p \<Longrightarrow> prime_elem (primitive_part_fract p)" | 
| 63498 | 566 | by (intro irreducible_imp_prime_poly irreducible_primitive_part_fract) | 
| 567 | ||
| 568 | lemma irreducible_linear_field_poly: | |
| 569 | fixes a b :: "'a::field" | |
| 570 | assumes "b \<noteq> 0" | |
| 571 | shows "irreducible [:a,b:]" | |
| 572 | proof (rule irreducibleI) | |
| 573 | fix p q assume pq: "[:a,b:] = p * q" | |
| 63539 | 574 | also from pq assms have "degree \<dots> = degree p + degree q" | 
| 63498 | 575 | by (intro degree_mult_eq) auto | 
| 576 | finally have "degree p = 0 \<or> degree q = 0" using assms by auto | |
| 577 | with assms pq show "is_unit p \<or> is_unit q" | |
| 578 | by (auto simp: is_unit_const_poly_iff dvd_field_iff elim!: degree_eq_zeroE) | |
| 579 | qed (insert assms, auto simp: is_unit_poly_iff) | |
| 580 | ||
| 63633 | 581 | lemma prime_elem_linear_field_poly: | 
| 582 | "(b :: 'a :: field) \<noteq> 0 \<Longrightarrow> prime_elem [:a,b:]" | |
| 63498 | 583 | by (rule field_poly_irreducible_imp_prime, rule irreducible_linear_field_poly) | 
| 584 | ||
| 585 | lemma irreducible_linear_poly: | |
| 586 |   fixes a b :: "'a::{idom_divide,ring_gcd,factorial_semiring,semiring_Gcd}"
 | |
| 587 | shows "b \<noteq> 0 \<Longrightarrow> coprime a b \<Longrightarrow> irreducible [:a,b:]" | |
| 588 | by (auto intro!: irreducible_linear_field_poly | |
| 589 | simp: nonconst_poly_irreducible_iff content_def map_poly_pCons) | |
| 590 | ||
| 63633 | 591 | lemma prime_elem_linear_poly: | 
| 63498 | 592 |   fixes a b :: "'a::{idom_divide,ring_gcd,factorial_semiring,semiring_Gcd}"
 | 
| 63633 | 593 | shows "b \<noteq> 0 \<Longrightarrow> coprime a b \<Longrightarrow> prime_elem [:a,b:]" | 
| 63498 | 594 | by (rule irreducible_imp_prime_poly, rule irreducible_linear_poly) | 
| 595 | ||
| 64591 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 haftmann parents: 
64267diff
changeset | 596 | |
| 63498 | 597 | subsection \<open>Prime factorisation of polynomials\<close> | 
| 598 | ||
| 599 | private lemma poly_prime_factorization_exists_content_1: | |
| 600 |   fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} poly"
 | |
| 601 | assumes "p \<noteq> 0" "content p = 1" | |
| 63830 | 602 | shows "\<exists>A. (\<forall>p. p \<in># A \<longrightarrow> prime_elem p) \<and> prod_mset A = normalize p" | 
| 63498 | 603 | proof - | 
| 604 | let ?P = "field_poly.prime_factorization (fract_poly p)" | |
| 63830 | 605 | define c where "c = prod_mset (image_mset fract_content ?P)" | 
| 63498 | 606 | define c' where "c' = c * to_fract (lead_coeff p)" | 
| 63830 | 607 | define e where "e = prod_mset (image_mset primitive_part_fract ?P)" | 
| 63498 | 608 | define A where "A = image_mset (normalize \<circ> primitive_part_fract) ?P" | 
| 609 | have "content e = (\<Prod>x\<in>#field_poly.prime_factorization (map_poly to_fract p). | |
| 610 | content (primitive_part_fract x))" | |
| 63830 | 611 | by (simp add: e_def content_prod_mset multiset.map_comp o_def) | 
| 63498 | 612 | also have "image_mset (\<lambda>x. content (primitive_part_fract x)) ?P = image_mset (\<lambda>_. 1) ?P" | 
| 613 | by (intro image_mset_cong content_primitive_part_fract) auto | |
| 64591 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 haftmann parents: 
64267diff
changeset | 614 | finally have content_e: "content e = 1" | 
| 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 haftmann parents: 
64267diff
changeset | 615 | by simp | 
| 63498 | 616 | |
| 66805 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 617 | from \<open>p \<noteq> 0\<close> have "fract_poly p = [:lead_coeff (fract_poly p):] * | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 618 | smult (inverse (lead_coeff (fract_poly p))) (fract_poly p)" | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 619 | by simp | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 620 | also have "[:lead_coeff (fract_poly p):] = [:to_fract (lead_coeff p):]" | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 621 | by (simp add: monom_0 degree_map_poly coeff_map_poly) | 
| 
274b4edca859
Polynomial_Factorial does not depend on Field_as_Ring as such
 haftmann parents: 
65965diff
changeset | 622 | also from assms have "smult (inverse (lead_coeff (fract_poly p))) (fract_poly p) = prod_mset ?P" | 
| 63830 | 623 | by (subst field_poly_prod_mset_prime_factorization) simp_all | 
| 624 | also have "\<dots> = prod_mset (image_mset id ?P)" by simp | |
| 63498 | 625 | also have "image_mset id ?P = | 
| 626 | image_mset (\<lambda>x. [:fract_content x:] * fract_poly (primitive_part_fract x)) ?P" | |
| 627 | by (intro image_mset_cong) (auto simp: content_times_primitive_part_fract) | |
| 63830 | 628 | also have "prod_mset \<dots> = smult c (fract_poly e)" | 
| 64591 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 haftmann parents: 
64267diff
changeset | 629 | by (subst prod_mset.distrib) (simp_all add: prod_mset_fract_poly prod_mset_const_poly c_def e_def) | 
| 63498 | 630 | also have "[:to_fract (lead_coeff p):] * \<dots> = smult c' (fract_poly e)" | 
| 631 | by (simp add: c'_def) | |
| 632 | finally have eq: "fract_poly p = smult c' (fract_poly e)" . | |
| 633 | also obtain b where b: "c' = to_fract b" "is_unit b" | |
| 634 | proof - | |
| 635 | from fract_poly_smult_eqE[OF eq] guess a b . note ab = this | |
| 636 | from ab(2) have "content (smult a p) = content (smult b e)" by (simp only: ) | |
| 637 | with assms content_e have "a = normalize b" by (simp add: ab(4)) | |
| 67051 | 638 | with ab have ab': "a = 1" "is_unit b" | 
| 639 | by (simp_all add: normalize_1_iff) | |
| 63498 | 640 | with ab ab' have "c' = to_fract b" by auto | 
| 641 | from this and \<open>is_unit b\<close> show ?thesis by (rule that) | |
| 642 | qed | |
| 643 | hence "smult c' (fract_poly e) = fract_poly (smult b e)" by simp | |
| 644 | finally have "p = smult b e" by (simp only: fract_poly_eq_iff) | |
| 645 | hence "p = [:b:] * e" by simp | |
| 646 | with b have "normalize p = normalize e" | |
| 647 | by (simp only: normalize_mult) (simp add: is_unit_normalize is_unit_poly_iff) | |
| 63830 | 648 | also have "normalize e = prod_mset A" | 
| 649 | by (simp add: multiset.map_comp e_def A_def normalize_prod_mset) | |
| 650 | finally have "prod_mset A = normalize p" .. | |
| 63498 | 651 | |
| 63633 | 652 | have "prime_elem p" if "p \<in># A" for p | 
| 653 | using that by (auto simp: A_def prime_elem_primitive_part_fract prime_elem_imp_irreducible | |
| 63498 | 654 | dest!: field_poly_in_prime_factorization_imp_prime ) | 
| 63830 | 655 | from this and \<open>prod_mset A = normalize p\<close> show ?thesis | 
| 63498 | 656 | by (intro exI[of _ A]) blast | 
| 657 | qed | |
| 658 | ||
| 659 | lemma poly_prime_factorization_exists: | |
| 660 |   fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} poly"
 | |
| 661 | assumes "p \<noteq> 0" | |
| 63830 | 662 | shows "\<exists>A. (\<forall>p. p \<in># A \<longrightarrow> prime_elem p) \<and> prod_mset A = normalize p" | 
| 63498 | 663 | proof - | 
| 664 | define B where "B = image_mset (\<lambda>x. [:x:]) (prime_factorization (content p))" | |
| 63830 | 665 | have "\<exists>A. (\<forall>p. p \<in># A \<longrightarrow> prime_elem p) \<and> prod_mset A = normalize (primitive_part p)" | 
| 63498 | 666 | by (rule poly_prime_factorization_exists_content_1) (insert assms, simp_all) | 
| 667 | then guess A by (elim exE conjE) note A = this | |
| 63830 | 668 | moreover from assms have "prod_mset B = [:content p:]" | 
| 669 | by (simp add: B_def prod_mset_const_poly prod_mset_prime_factorization) | |
| 63633 | 670 | moreover have "\<forall>p. p \<in># B \<longrightarrow> prime_elem p" | 
| 63905 | 671 | by (auto simp: B_def intro!: lift_prime_elem_poly dest: in_prime_factors_imp_prime) | 
| 63498 | 672 | ultimately show ?thesis by (intro exI[of _ "B + A"]) auto | 
| 673 | qed | |
| 674 | ||
| 675 | end | |
| 676 | ||
| 677 | ||
| 678 | subsection \<open>Typeclass instances\<close> | |
| 679 | ||
| 680 | instance poly :: (factorial_ring_gcd) factorial_semiring | |
| 681 | by standard (rule poly_prime_factorization_exists) | |
| 682 | ||
| 683 | instantiation poly :: (factorial_ring_gcd) factorial_ring_gcd | |
| 684 | begin | |
| 685 | ||
| 686 | definition gcd_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" where | |
| 687 | [code del]: "gcd_poly = gcd_factorial" | |
| 688 | ||
| 689 | definition lcm_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" where | |
| 690 | [code del]: "lcm_poly = lcm_factorial" | |
| 691 | ||
| 692 | definition Gcd_poly :: "'a poly set \<Rightarrow> 'a poly" where | |
| 693 | [code del]: "Gcd_poly = Gcd_factorial" | |
| 694 | ||
| 695 | definition Lcm_poly :: "'a poly set \<Rightarrow> 'a poly" where | |
| 696 | [code del]: "Lcm_poly = Lcm_factorial" | |
| 697 | ||
| 698 | instance by standard (simp_all add: gcd_poly_def lcm_poly_def Gcd_poly_def Lcm_poly_def) | |
| 699 | ||
| 700 | end | |
| 701 | ||
| 66817 | 702 | instantiation poly :: ("{field,factorial_ring_gcd}") "{unique_euclidean_ring, normalization_euclidean_semiring}"
 | 
| 63498 | 703 | begin | 
| 704 | ||
| 64784 
5cb5e7ecb284
reshaped euclidean semiring into hierarchy of euclidean semirings culminating in uniquely determined euclidean divion
 haftmann parents: 
64591diff
changeset | 705 | definition euclidean_size_poly :: "'a poly \<Rightarrow> nat" | 
| 
5cb5e7ecb284
reshaped euclidean semiring into hierarchy of euclidean semirings culminating in uniquely determined euclidean divion
 haftmann parents: 
64591diff
changeset | 706 | where "euclidean_size_poly p = (if p = 0 then 0 else 2 ^ degree p)" | 
| 
5cb5e7ecb284
reshaped euclidean semiring into hierarchy of euclidean semirings culminating in uniquely determined euclidean divion
 haftmann parents: 
64591diff
changeset | 707 | |
| 66838 
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
 haftmann parents: 
66817diff
changeset | 708 | definition division_segment_poly :: "'a poly \<Rightarrow> 'a poly" | 
| 
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
 haftmann parents: 
66817diff
changeset | 709 | where [simp]: "division_segment_poly p = 1" | 
| 63498 | 710 | |
| 66806 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
66805diff
changeset | 711 | instance proof | 
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
66805diff
changeset | 712 | show "(q * p + r) div p = q" if "p \<noteq> 0" | 
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
66805diff
changeset | 713 | and "euclidean_size r < euclidean_size p" for q p r :: "'a poly" | 
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
66805diff
changeset | 714 | proof - | 
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
66805diff
changeset | 715 | from \<open>p \<noteq> 0\<close> eucl_rel_poly [of r p] have "eucl_rel_poly (r + q * p) p (q + r div p, r mod p)" | 
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
66805diff
changeset | 716 | by (simp add: eucl_rel_poly_iff distrib_right) | 
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
66805diff
changeset | 717 | then have "(r + q * p) div p = q + r div p" | 
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
66805diff
changeset | 718 | by (rule div_poly_eq) | 
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
66805diff
changeset | 719 | with that show ?thesis | 
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
66805diff
changeset | 720 | by (cases "r = 0") (simp_all add: euclidean_size_poly_def div_poly_less ac_simps) | 
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
66805diff
changeset | 721 | qed | 
| 66840 | 722 | qed (auto simp: euclidean_size_poly_def Rings.div_mult_mod_eq div_poly_less degree_mult_eq power_add | 
| 723 | intro!: degree_mod_less' split: if_splits) | |
| 64784 
5cb5e7ecb284
reshaped euclidean semiring into hierarchy of euclidean semirings culminating in uniquely determined euclidean divion
 haftmann parents: 
64591diff
changeset | 724 | |
| 63498 | 725 | end | 
| 726 | ||
| 66817 | 727 | instance poly :: ("{field, normalization_euclidean_semiring, factorial_ring_gcd}") euclidean_ring_gcd
 | 
| 728 | by (rule euclidean_ring_gcd_class.intro, rule factorial_euclidean_semiring_gcdI) standard | |
| 63498 | 729 | |
| 730 | ||
| 731 | subsection \<open>Polynomial GCD\<close> | |
| 732 | ||
| 733 | lemma gcd_poly_decompose: | |
| 734 | fixes p q :: "'a :: factorial_ring_gcd poly" | |
| 735 | shows "gcd p q = | |
| 736 | smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q))" | |
| 737 | proof (rule sym, rule gcdI) | |
| 738 | have "[:gcd (content p) (content q):] * gcd (primitive_part p) (primitive_part q) dvd | |
| 739 | [:content p:] * primitive_part p" by (intro mult_dvd_mono) simp_all | |
| 740 | thus "smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q)) dvd p" | |
| 741 | by simp | |
| 742 | next | |
| 743 | have "[:gcd (content p) (content q):] * gcd (primitive_part p) (primitive_part q) dvd | |
| 744 | [:content q:] * primitive_part q" by (intro mult_dvd_mono) simp_all | |
| 745 | thus "smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q)) dvd q" | |
| 746 | by simp | |
| 747 | next | |
| 748 | fix d assume "d dvd p" "d dvd q" | |
| 749 | hence "[:content d:] * primitive_part d dvd | |
| 750 | [:gcd (content p) (content q):] * gcd (primitive_part p) (primitive_part q)" | |
| 751 | by (intro mult_dvd_mono) auto | |
| 752 | thus "d dvd smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q))" | |
| 753 | by simp | |
| 754 | qed (auto simp: normalize_smult) | |
| 755 | ||
| 756 | ||
| 757 | lemma gcd_poly_pseudo_mod: | |
| 758 | fixes p q :: "'a :: factorial_ring_gcd poly" | |
| 759 | assumes nz: "q \<noteq> 0" and prim: "content p = 1" "content q = 1" | |
| 760 | shows "gcd p q = gcd q (primitive_part (pseudo_mod p q))" | |
| 761 | proof - | |
| 762 | define r s where "r = fst (pseudo_divmod p q)" and "s = snd (pseudo_divmod p q)" | |
| 763 | define a where "a = [:coeff q (degree q) ^ (Suc (degree p) - degree q):]" | |
| 764 | have [simp]: "primitive_part a = unit_factor a" | |
| 765 | by (simp add: a_def unit_factor_poly_def unit_factor_power monom_0) | |
| 766 | from nz have [simp]: "a \<noteq> 0" by (auto simp: a_def) | |
| 767 | ||
| 768 | have rs: "pseudo_divmod p q = (r, s)" by (simp add: r_def s_def) | |
| 769 | have "gcd (q * r + s) q = gcd q s" | |
| 770 | using gcd_add_mult[of q r s] by (simp add: gcd.commute add_ac mult_ac) | |
| 771 | with pseudo_divmod(1)[OF nz rs] | |
| 772 | have "gcd (p * a) q = gcd q s" by (simp add: a_def) | |
| 773 | also from prim have "gcd (p * a) q = gcd p q" | |
| 774 | by (subst gcd_poly_decompose) | |
| 775 | (auto simp: primitive_part_mult gcd_mult_unit1 primitive_part_prim | |
| 776 | simp del: mult_pCons_right ) | |
| 777 | also from prim have "gcd q s = gcd q (primitive_part s)" | |
| 778 | by (subst gcd_poly_decompose) (simp_all add: primitive_part_prim) | |
| 779 | also have "s = pseudo_mod p q" by (simp add: s_def pseudo_mod_def) | |
| 780 | finally show ?thesis . | |
| 781 | qed | |
| 782 | ||
| 783 | lemma degree_pseudo_mod_less: | |
| 784 | assumes "q \<noteq> 0" "pseudo_mod p q \<noteq> 0" | |
| 785 | shows "degree (pseudo_mod p q) < degree q" | |
| 786 | using pseudo_mod(2)[of q p] assms by auto | |
| 787 | ||
| 788 | function gcd_poly_code_aux :: "'a :: factorial_ring_gcd poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" where | |
| 789 | "gcd_poly_code_aux p q = | |
| 790 | (if q = 0 then normalize p else gcd_poly_code_aux q (primitive_part (pseudo_mod p q)))" | |
| 791 | by auto | |
| 792 | termination | |
| 793 | by (relation "measure ((\<lambda>p. if p = 0 then 0 else Suc (degree p)) \<circ> snd)") | |
| 64164 
38c407446400
separate type class for arbitrary quotient and remainder partitions
 haftmann parents: 
63954diff
changeset | 794 | (auto simp: degree_pseudo_mod_less) | 
| 63498 | 795 | |
| 796 | declare gcd_poly_code_aux.simps [simp del] | |
| 797 | ||
| 798 | lemma gcd_poly_code_aux_correct: | |
| 799 | assumes "content p = 1" "q = 0 \<or> content q = 1" | |
| 800 | shows "gcd_poly_code_aux p q = gcd p q" | |
| 801 | using assms | |
| 802 | proof (induction p q rule: gcd_poly_code_aux.induct) | |
| 803 | case (1 p q) | |
| 804 | show ?case | |
| 805 | proof (cases "q = 0") | |
| 806 | case True | |
| 807 | thus ?thesis by (subst gcd_poly_code_aux.simps) auto | |
| 808 | next | |
| 809 | case False | |
| 810 | hence "gcd_poly_code_aux p q = gcd_poly_code_aux q (primitive_part (pseudo_mod p q))" | |
| 811 | by (subst gcd_poly_code_aux.simps) simp_all | |
| 812 | also from "1.prems" False | |
| 813 | have "primitive_part (pseudo_mod p q) = 0 \<or> | |
| 814 | content (primitive_part (pseudo_mod p q)) = 1" | |
| 815 | by (cases "pseudo_mod p q = 0") auto | |
| 816 | with "1.prems" False | |
| 817 | have "gcd_poly_code_aux q (primitive_part (pseudo_mod p q)) = | |
| 818 | gcd q (primitive_part (pseudo_mod p q))" | |
| 819 | by (intro 1) simp_all | |
| 820 | also from "1.prems" False | |
| 821 | have "\<dots> = gcd p q" by (intro gcd_poly_pseudo_mod [symmetric]) auto | |
| 822 | finally show ?thesis . | |
| 823 | qed | |
| 824 | qed | |
| 825 | ||
| 826 | definition gcd_poly_code | |
| 827 | :: "'a :: factorial_ring_gcd poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" | |
| 828 | where "gcd_poly_code p q = | |
| 829 | (if p = 0 then normalize q else if q = 0 then normalize p else | |
| 830 | smult (gcd (content p) (content q)) | |
| 831 | (gcd_poly_code_aux (primitive_part p) (primitive_part q)))" | |
| 832 | ||
| 64591 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 haftmann parents: 
64267diff
changeset | 833 | lemma gcd_poly_code [code]: "gcd p q = gcd_poly_code p q" | 
| 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 haftmann parents: 
64267diff
changeset | 834 | by (simp add: gcd_poly_code_def gcd_poly_code_aux_correct gcd_poly_decompose [symmetric]) | 
| 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 haftmann parents: 
64267diff
changeset | 835 | |
| 63498 | 836 | lemma lcm_poly_code [code]: | 
| 837 | fixes p q :: "'a :: factorial_ring_gcd poly" | |
| 838 | shows "lcm p q = normalize (p * q) div gcd p q" | |
| 64591 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 haftmann parents: 
64267diff
changeset | 839 | by (fact lcm_gcd) | 
| 63498 | 840 | |
| 64850 | 841 | lemmas Gcd_poly_set_eq_fold [code] = Gcd_set_eq_fold [where ?'a = "'a :: factorial_ring_gcd poly"] | 
| 842 | lemmas Lcm_poly_set_eq_fold [code] = Lcm_set_eq_fold [where ?'a = "'a :: factorial_ring_gcd poly"] | |
| 64860 | 843 | |
| 64591 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 haftmann parents: 
64267diff
changeset | 844 | text \<open>Example: | 
| 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 haftmann parents: 
64267diff
changeset | 845 |   @{lemma "Lcm {[:1, 2, 3:], [:2, 3, 4:]} = [:[:2:], [:7:], [:16:], [:17:], [:12 :: int:]:]" by eval}
 | 
| 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 haftmann parents: 
64267diff
changeset | 846 | \<close> | 
| 63498 | 847 | |
| 63764 | 848 | end |