diff --git a/src/HOL/Computational_Algebra/Euclidean_Algorithm.thy b/src/HOL/Computational_Algebra/Euclidean_Algorithm.thy --- a/src/HOL/Computational_Algebra/Euclidean_Algorithm.thy +++ b/src/HOL/Computational_Algebra/Euclidean_Algorithm.thy @@ -1,655 +1,678 @@ (* Title: HOL/Computational_Algebra/Euclidean_Algorithm.thy Author: Manuel Eberl, TU Muenchen *) section \Abstract euclidean algorithm in euclidean (semi)rings\ theory Euclidean_Algorithm imports Factorial_Ring begin subsection \Generic construction of the (simple) euclidean algorithm\ class normalization_euclidean_semiring = euclidean_semiring + normalization_semidom begin lemma euclidean_size_normalize [simp]: "euclidean_size (normalize a) = euclidean_size a" proof (cases "a = 0") case True then show ?thesis by simp next case [simp]: False have "euclidean_size (normalize a) \ euclidean_size (normalize a * unit_factor a)" by (rule size_mult_mono) simp moreover have "euclidean_size a \ euclidean_size (a * (1 div unit_factor a))" by (rule size_mult_mono) simp ultimately show ?thesis by simp qed context begin qualified function gcd :: "'a \ 'a \ 'a" where "gcd a b = (if b = 0 then normalize a else gcd b (a mod b))" by pat_completeness simp termination by (relation "measure (euclidean_size \ snd)") (simp_all add: mod_size_less) declare gcd.simps [simp del] lemma eucl_induct [case_names zero mod]: assumes H1: "\b. P b 0" and H2: "\a b. b \ 0 \ P b (a mod b) \ P a b" shows "P a b" proof (induct a b rule: gcd.induct) case (1 a b) show ?case proof (cases "b = 0") case True then show "P a b" by simp (rule H1) next case False then have "P b (a mod b)" by (rule "1.hyps") with \b \ 0\ show "P a b" by (blast intro: H2) qed qed qualified lemma gcd_0: "gcd a 0 = normalize a" by (simp add: gcd.simps [of a 0]) qualified lemma gcd_mod: "a \ 0 \ gcd a (b mod a) = gcd b a" by (simp add: gcd.simps [of b 0] gcd.simps [of b a]) qualified definition lcm :: "'a \ 'a \ 'a" - where "lcm a b = normalize (a * b) div gcd a b" + where "lcm a b = normalize (a * b div gcd a b)" qualified definition Lcm :: "'a set \ 'a" \ \Somewhat complicated definition of Lcm that has the advantage of working for infinite sets as well\ where [code del]: "Lcm A = (if \l. l \ 0 \ (\a\A. a dvd l) then let l = SOME l. l \ 0 \ (\a\A. a dvd l) \ euclidean_size l = (LEAST n. \l. l \ 0 \ (\a\A. a dvd l) \ euclidean_size l = n) in normalize l else 0)" qualified definition Gcd :: "'a set \ 'a" where [code del]: "Gcd A = Lcm {d. \a\A. d dvd a}" end lemma semiring_gcd: "class.semiring_gcd one zero times gcd lcm divide plus minus unit_factor normalize" proof show "gcd a b dvd a" and "gcd a b dvd b" for a b by (induct a b rule: eucl_induct) (simp_all add: local.gcd_0 local.gcd_mod dvd_mod_iff) next show "c dvd a \ c dvd b \ c dvd gcd a b" for a b c proof (induct a b rule: eucl_induct) case (zero a) from \c dvd a\ show ?case by (rule dvd_trans) (simp add: local.gcd_0) next case (mod a b) then show ?case by (simp add: local.gcd_mod dvd_mod_iff) qed next show "normalize (gcd a b) = gcd a b" for a b by (induct a b rule: eucl_induct) (simp_all add: local.gcd_0 local.gcd_mod) next - show "lcm a b = normalize (a * b) div gcd a b" for a b + show "lcm a b = normalize (a * b div gcd a b)" for a b by (fact local.lcm_def) qed interpretation semiring_gcd one zero times gcd lcm divide plus minus unit_factor normalize by (fact semiring_gcd) lemma semiring_Gcd: "class.semiring_Gcd one zero times gcd lcm Gcd Lcm divide plus minus unit_factor normalize" proof - show ?thesis proof have "(\a\A. a dvd Lcm A) \ (\b. (\a\A. a dvd b) \ Lcm A dvd b)" for A proof (cases "\l. l \ 0 \ (\a\A. a dvd l)") case False then have "Lcm A = 0" by (auto simp add: local.Lcm_def) with False show ?thesis by auto next case True then obtain l\<^sub>0 where l\<^sub>0_props: "l\<^sub>0 \ 0" "\a\A. a dvd l\<^sub>0" by blast define n where "n = (LEAST n. \l. l \ 0 \ (\a\A. a dvd l) \ euclidean_size l = n)" define l where "l = (SOME l. l \ 0 \ (\a\A. a dvd l) \ euclidean_size l = n)" have "\l. l \ 0 \ (\a\A. a dvd l) \ euclidean_size l = n" apply (subst n_def) apply (rule LeastI [of _ "euclidean_size l\<^sub>0"]) apply (rule exI [of _ l\<^sub>0]) apply (simp add: l\<^sub>0_props) done from someI_ex [OF this] have "l \ 0" and "\a\A. a dvd l" and "euclidean_size l = n" unfolding l_def by simp_all { fix l' assume "\a\A. a dvd l'" with \\a\A. a dvd l\ have "\a\A. a dvd gcd l l'" by (auto intro: gcd_greatest) moreover from \l \ 0\ have "gcd l l' \ 0" by simp ultimately have "\b. b \ 0 \ (\a\A. a dvd b) \ euclidean_size b = euclidean_size (gcd l l')" by (intro exI [of _ "gcd l l'"], auto) then have "euclidean_size (gcd l l') \ n" by (subst n_def) (rule Least_le) moreover have "euclidean_size (gcd l l') \ n" proof - have "gcd l l' dvd l" by simp then obtain a where "l = gcd l l' * a" .. with \l \ 0\ have "a \ 0" by auto hence "euclidean_size (gcd l l') \ euclidean_size (gcd l l' * a)" by (rule size_mult_mono) also have "gcd l l' * a = l" using \l = gcd l l' * a\ .. also note \euclidean_size l = n\ finally show "euclidean_size (gcd l l') \ n" . qed ultimately have *: "euclidean_size l = euclidean_size (gcd l l')" by (intro le_antisym, simp_all add: \euclidean_size l = n\) from \l \ 0\ have "l dvd gcd l l'" by (rule dvd_euclidean_size_eq_imp_dvd) (auto simp add: *) hence "l dvd l'" by (rule dvd_trans [OF _ gcd_dvd2]) } with \\a\A. a dvd l\ and \l \ 0\ have "(\a\A. a dvd normalize l) \ (\l'. (\a\A. a dvd l') \ normalize l dvd l')" by auto also from True have "normalize l = Lcm A" by (simp add: local.Lcm_def Let_def n_def l_def) finally show ?thesis . qed then show dvd_Lcm: "a \ A \ a dvd Lcm A" and Lcm_least: "(\a. a \ A \ a dvd b) \ Lcm A dvd b" for A and a b by auto show "a \ A \ Gcd A dvd a" for A and a by (auto simp add: local.Gcd_def intro: Lcm_least) show "(\a. a \ A \ b dvd a) \ b dvd Gcd A" for A and b by (auto simp add: local.Gcd_def intro: dvd_Lcm) show [simp]: "normalize (Lcm A) = Lcm A" for A by (simp add: local.Lcm_def) show "normalize (Gcd A) = Gcd A" for A by (simp add: local.Gcd_def) qed qed interpretation semiring_Gcd one zero times gcd lcm Gcd Lcm divide plus minus unit_factor normalize by (fact semiring_Gcd) subclass factorial_semiring proof - show "class.factorial_semiring divide plus minus zero times one unit_factor normalize" proof (standard, rule factorial_semiring_altI_aux) \ \FIXME rule\ fix x assume "x \ 0" thus "finite {p. p dvd x \ normalize p = p}" proof (induction "euclidean_size x" arbitrary: x rule: less_induct) case (less x) show ?case proof (cases "\y. y dvd x \ \x dvd y \ \is_unit y") case False have "{p. p dvd x \ normalize p = p} \ {1, normalize x}" proof fix p assume p: "p \ {p. p dvd x \ normalize p = p}" with False have "is_unit p \ x dvd p" by blast thus "p \ {1, normalize x}" proof (elim disjE) assume "is_unit p" hence "normalize p = 1" by (simp add: is_unit_normalize) with p show ?thesis by simp next assume "x dvd p" with p have "normalize p = normalize x" by (intro associatedI) simp_all with p show ?thesis by simp qed qed moreover have "finite \" by simp ultimately show ?thesis by (rule finite_subset) next case True then obtain y where y: "y dvd x" "\x dvd y" "\is_unit y" by blast define z where "z = x div y" let ?fctrs = "\x. {p. p dvd x \ normalize p = p}" from y have x: "x = y * z" by (simp add: z_def) with less.prems have "y \ 0" "z \ 0" by auto have normalized_factors_product: - "{p. p dvd a * b \ normalize p = p} = - (\(x,y). x * y) ` ({p. p dvd a \ normalize p = p} \ {p. p dvd b \ normalize p = p})" + "{p. p dvd a * b \ normalize p = p} \ + (\(x,y). normalize (x * y)) ` ({p. p dvd a \ normalize p = p} \ {p. p dvd b \ normalize p = p})" for a b proof safe fix p assume p: "p dvd a * b" "normalize p = p" from p(1) obtain x y where xy: "p = x * y" "x dvd a" "y dvd b" by (rule dvd_productE) define x' y' where "x' = normalize x" and "y' = normalize y" - have "p = x' * y'" - by (subst p(2) [symmetric]) (simp add: xy x'_def y'_def normalize_mult) - moreover from xy have "normalize x' = x'" "normalize y' = y'" "x' dvd a" "y' dvd b" - by (simp_all add: x'_def y'_def) - ultimately show "p \ (\(x, y). x * y) ` - ({p. p dvd a \ normalize p = p} \ {p. p dvd b \ normalize p = p})" - by blast - qed (auto simp: normalize_mult mult_dvd_mono) + have "p = normalize (x' * y')" + using p by (simp add: xy x'_def y'_def) + moreover have "x' dvd a \ normalize x' = x'" and "y' dvd b \ normalize y' = y'" + using xy by (auto simp: x'_def y'_def) + ultimately show "p \ (\(x, y). normalize (x * y)) ` + ({p. p dvd a \ normalize p = p} \ {p. p dvd b \ normalize p = p})" by fast + qed from x y have "\is_unit z" by (auto simp: mult_unit_dvd_iff) - have "?fctrs x = (\(p,p'). p * p') ` (?fctrs y \ ?fctrs z)" + have "?fctrs x \ (\(p,p'). normalize (p * p')) ` (?fctrs y \ ?fctrs z)" by (subst x) (rule normalized_factors_product) - also have "\y * z dvd y * 1" "\y * z dvd 1 * z" + moreover have "\y * z dvd y * 1" "\y * z dvd 1 * z" by (subst dvd_times_left_cancel_iff dvd_times_right_cancel_iff; fact)+ - hence "finite ((\(p,p'). p * p') ` (?fctrs y \ ?fctrs z))" + hence "finite ((\(p,p'). normalize (p * p')) ` (?fctrs y \ ?fctrs z))" by (intro finite_imageI finite_cartesian_product less dvd_proper_imp_size_less) (auto simp: x) - finally show ?thesis . + ultimately show ?thesis by (rule finite_subset) qed qed next fix p assume "irreducible p" then show "prime_elem p" by (rule irreducible_imp_prime_elem_gcd) qed qed lemma Gcd_eucl_set [code]: "Gcd (set xs) = fold gcd xs 0" by (fact Gcd_set_eq_fold) lemma Lcm_eucl_set [code]: "Lcm (set xs) = fold lcm xs 1" by (fact Lcm_set_eq_fold) end hide_const (open) gcd lcm Gcd Lcm lemma prime_elem_int_abs_iff [simp]: fixes p :: int shows "prime_elem \p\ \ prime_elem p" using prime_elem_normalize_iff [of p] by simp lemma prime_elem_int_minus_iff [simp]: fixes p :: int shows "prime_elem (- p) \ prime_elem p" using prime_elem_normalize_iff [of "- p"] by simp lemma prime_int_iff: fixes p :: int shows "prime p \ p > 0 \ prime_elem p" by (auto simp add: prime_def dest: prime_elem_not_zeroI) subsection \The (simple) euclidean algorithm as gcd computation\ class euclidean_semiring_gcd = normalization_euclidean_semiring + gcd + Gcd + assumes gcd_eucl: "Euclidean_Algorithm.gcd = GCD.gcd" and lcm_eucl: "Euclidean_Algorithm.lcm = GCD.lcm" assumes Gcd_eucl: "Euclidean_Algorithm.Gcd = GCD.Gcd" and Lcm_eucl: "Euclidean_Algorithm.Lcm = GCD.Lcm" begin subclass semiring_gcd unfolding gcd_eucl [symmetric] lcm_eucl [symmetric] by (fact semiring_gcd) subclass semiring_Gcd unfolding gcd_eucl [symmetric] lcm_eucl [symmetric] Gcd_eucl [symmetric] Lcm_eucl [symmetric] by (fact semiring_Gcd) subclass factorial_semiring_gcd proof show "gcd a b = gcd_factorial a b" for a b apply (rule sym) apply (rule gcdI) apply (fact gcd_lcm_factorial)+ done then show "lcm a b = lcm_factorial a b" for a b by (simp add: lcm_factorial_gcd_factorial lcm_gcd) show "Gcd A = Gcd_factorial A" for A apply (rule sym) apply (rule GcdI) apply (fact gcd_lcm_factorial)+ done show "Lcm A = Lcm_factorial A" for A apply (rule sym) apply (rule LcmI) apply (fact gcd_lcm_factorial)+ done qed lemma gcd_mod_right [simp]: "a \ 0 \ gcd a (b mod a) = gcd a b" unfolding gcd.commute [of a b] by (simp add: gcd_eucl [symmetric] local.gcd_mod) lemma gcd_mod_left [simp]: "b \ 0 \ gcd (a mod b) b = gcd a b" by (drule gcd_mod_right [of _ a]) (simp add: gcd.commute) lemma euclidean_size_gcd_le1 [simp]: assumes "a \ 0" shows "euclidean_size (gcd a b) \ euclidean_size a" proof - from gcd_dvd1 obtain c where A: "a = gcd a b * c" .. with assms have "c \ 0" by auto moreover from this have "euclidean_size (gcd a b) \ euclidean_size (gcd a b * c)" by (rule size_mult_mono) with A show ?thesis by simp qed lemma euclidean_size_gcd_le2 [simp]: "b \ 0 \ euclidean_size (gcd a b) \ euclidean_size b" by (subst gcd.commute, rule euclidean_size_gcd_le1) lemma euclidean_size_gcd_less1: assumes "a \ 0" and "\ a dvd b" shows "euclidean_size (gcd a b) < euclidean_size a" proof (rule ccontr) assume "\euclidean_size (gcd a b) < euclidean_size a" with \a \ 0\ have A: "euclidean_size (gcd a b) = euclidean_size a" by (intro le_antisym, simp_all) have "a dvd gcd a b" by (rule dvd_euclidean_size_eq_imp_dvd) (simp_all add: assms A) hence "a dvd b" using dvd_gcdD2 by blast with \\ a dvd b\ show False by contradiction qed lemma euclidean_size_gcd_less2: assumes "b \ 0" and "\ b dvd a" shows "euclidean_size (gcd a b) < euclidean_size b" using assms by (subst gcd.commute, rule euclidean_size_gcd_less1) lemma euclidean_size_lcm_le1: assumes "a \ 0" and "b \ 0" shows "euclidean_size a \ euclidean_size (lcm a b)" proof - have "a dvd lcm a b" by (rule dvd_lcm1) then obtain c where A: "lcm a b = a * c" .. with \a \ 0\ and \b \ 0\ have "c \ 0" by (auto simp: lcm_eq_0_iff) then show ?thesis by (subst A, intro size_mult_mono) qed lemma euclidean_size_lcm_le2: "a \ 0 \ b \ 0 \ euclidean_size b \ euclidean_size (lcm a b)" using euclidean_size_lcm_le1 [of b a] by (simp add: ac_simps) lemma euclidean_size_lcm_less1: assumes "b \ 0" and "\ b dvd a" shows "euclidean_size a < euclidean_size (lcm a b)" proof (rule ccontr) from assms have "a \ 0" by auto assume "\euclidean_size a < euclidean_size (lcm a b)" with \a \ 0\ and \b \ 0\ have "euclidean_size (lcm a b) = euclidean_size a" by (intro le_antisym, simp, intro euclidean_size_lcm_le1) with assms have "lcm a b dvd a" by (rule_tac dvd_euclidean_size_eq_imp_dvd) (auto simp: lcm_eq_0_iff) hence "b dvd a" by (rule lcm_dvdD2) with \\b dvd a\ show False by contradiction qed lemma euclidean_size_lcm_less2: assumes "a \ 0" and "\ a dvd b" shows "euclidean_size b < euclidean_size (lcm a b)" using assms euclidean_size_lcm_less1 [of a b] by (simp add: ac_simps) end lemma factorial_euclidean_semiring_gcdI: "OFCLASS('a::{factorial_semiring_gcd, normalization_euclidean_semiring}, euclidean_semiring_gcd_class)" proof interpret semiring_Gcd 1 0 times Euclidean_Algorithm.gcd Euclidean_Algorithm.lcm Euclidean_Algorithm.Gcd Euclidean_Algorithm.Lcm divide plus minus unit_factor normalize rewrites "dvd.dvd (*) = Rings.dvd" by (fact semiring_Gcd) (simp add: dvd.dvd_def dvd_def fun_eq_iff) show [simp]: "Euclidean_Algorithm.gcd = (gcd :: 'a \ _)" proof (rule ext)+ fix a b :: 'a show "Euclidean_Algorithm.gcd a b = gcd a b" proof (induct a b rule: eucl_induct) case zero then show ?case by simp next case (mod a b) moreover have "gcd b (a mod b) = gcd b a" using GCD.gcd_add_mult [of b "a div b" "a mod b", symmetric] by (simp add: div_mult_mod_eq) ultimately show ?case by (simp add: Euclidean_Algorithm.gcd_mod ac_simps) qed qed show [simp]: "Euclidean_Algorithm.Lcm = (Lcm :: 'a set \ _)" by (auto intro!: Lcm_eqI GCD.dvd_Lcm GCD.Lcm_least) show "Euclidean_Algorithm.lcm = (lcm :: 'a \ _)" by (simp add: fun_eq_iff Euclidean_Algorithm.lcm_def semiring_gcd_class.lcm_gcd) show "Euclidean_Algorithm.Gcd = (Gcd :: 'a set \ _)" by (simp add: fun_eq_iff Euclidean_Algorithm.Gcd_def semiring_Gcd_class.Gcd_Lcm) qed subsection \The extended euclidean algorithm\ class euclidean_ring_gcd = euclidean_semiring_gcd + idom begin subclass euclidean_ring .. subclass ring_gcd .. subclass factorial_ring_gcd .. function euclid_ext_aux :: "'a \ 'a \ 'a \ 'a \ 'a \ 'a \ ('a \ 'a) \ 'a" where "euclid_ext_aux s' s t' t r' r = ( if r = 0 then let c = 1 div unit_factor r' in ((s' * c, t' * c), normalize r') else let q = r' div r in euclid_ext_aux s (s' - q * s) t (t' - q * t) r (r' mod r))" by auto termination by (relation "measure (\(_, _, _, _, _, b). euclidean_size b)") (simp_all add: mod_size_less) abbreviation (input) euclid_ext :: "'a \ 'a \ ('a \ 'a) \ 'a" where "euclid_ext \ euclid_ext_aux 1 0 0 1" lemma assumes "gcd r' r = gcd a b" assumes "s' * a + t' * b = r'" assumes "s * a + t * b = r" assumes "euclid_ext_aux s' s t' t r' r = ((x, y), c)" shows euclid_ext_aux_eq_gcd: "c = gcd a b" and euclid_ext_aux_bezout: "x * a + y * b = gcd a b" proof - have "case euclid_ext_aux s' s t' t r' r of ((x, y), c) \ x * a + y * b = c \ c = gcd a b" (is "?P (euclid_ext_aux s' s t' t r' r)") using assms(1-3) proof (induction s' s t' t r' r rule: euclid_ext_aux.induct) case (1 s' s t' t r' r) show ?case proof (cases "r = 0") case True hence "euclid_ext_aux s' s t' t r' r = ((s' div unit_factor r', t' div unit_factor r'), normalize r')" by (subst euclid_ext_aux.simps) (simp add: Let_def) also have "?P \" proof safe have "s' div unit_factor r' * a + t' div unit_factor r' * b = (s' * a + t' * b) div unit_factor r'" by (cases "r' = 0") (simp_all add: unit_div_commute) also have "s' * a + t' * b = r'" by fact also have "\ div unit_factor r' = normalize r'" by simp finally show "s' div unit_factor r' * a + t' div unit_factor r' * b = normalize r'" . next from "1.prems" True show "normalize r' = gcd a b" by simp qed finally show ?thesis . next case False hence "euclid_ext_aux s' s t' t r' r = euclid_ext_aux s (s' - r' div r * s) t (t' - r' div r * t) r (r' mod r)" by (subst euclid_ext_aux.simps) (simp add: Let_def) also from "1.prems" False have "?P \" proof (intro "1.IH") have "(s' - r' div r * s) * a + (t' - r' div r * t) * b = (s' * a + t' * b) - r' div r * (s * a + t * b)" by (simp add: algebra_simps) also have "s' * a + t' * b = r'" by fact also have "s * a + t * b = r" by fact also have "r' - r' div r * r = r' mod r" using div_mult_mod_eq [of r' r] by (simp add: algebra_simps) finally show "(s' - r' div r * s) * a + (t' - r' div r * t) * b = r' mod r" . qed (auto simp: algebra_simps minus_mod_eq_div_mult [symmetric] gcd.commute) finally show ?thesis . qed qed with assms(4) show "c = gcd a b" "x * a + y * b = gcd a b" by simp_all qed declare euclid_ext_aux.simps [simp del] definition bezout_coefficients :: "'a \ 'a \ 'a \ 'a" where [code]: "bezout_coefficients a b = fst (euclid_ext a b)" lemma bezout_coefficients_0: "bezout_coefficients a 0 = (1 div unit_factor a, 0)" by (simp add: bezout_coefficients_def euclid_ext_aux.simps) lemma bezout_coefficients_left_0: "bezout_coefficients 0 a = (0, 1 div unit_factor a)" by (simp add: bezout_coefficients_def euclid_ext_aux.simps) lemma bezout_coefficients: assumes "bezout_coefficients a b = (x, y)" shows "x * a + y * b = gcd a b" using assms by (simp add: bezout_coefficients_def euclid_ext_aux_bezout [of a b a b 1 0 0 1 x y] prod_eq_iff) lemma bezout_coefficients_fst_snd: "fst (bezout_coefficients a b) * a + snd (bezout_coefficients a b) * b = gcd a b" by (rule bezout_coefficients) simp lemma euclid_ext_eq [simp]: "euclid_ext a b = (bezout_coefficients a b, gcd a b)" (is "?p = ?q") proof show "fst ?p = fst ?q" by (simp add: bezout_coefficients_def) have "snd (euclid_ext_aux 1 0 0 1 a b) = gcd a b" by (rule euclid_ext_aux_eq_gcd [of a b a b 1 0 0 1]) (simp_all add: prod_eq_iff) then show "snd ?p = snd ?q" by simp qed declare euclid_ext_eq [symmetric, code_unfold] end +class normalization_euclidean_semiring_multiplicative = + normalization_euclidean_semiring + normalization_semidom_multiplicative +begin + +subclass factorial_semiring_multiplicative .. + +end + +class field_gcd = + field + unique_euclidean_ring + euclidean_ring_gcd + normalization_semidom_multiplicative +begin + +subclass normalization_euclidean_semiring_multiplicative .. + +subclass normalization_euclidean_semiring .. + +subclass semiring_gcd_mult_normalize .. + +end + subsection \Typical instances\ instance nat :: normalization_euclidean_semiring .. instance nat :: euclidean_semiring_gcd proof interpret semiring_Gcd 1 0 times "Euclidean_Algorithm.gcd" "Euclidean_Algorithm.lcm" "Euclidean_Algorithm.Gcd" "Euclidean_Algorithm.Lcm" divide plus minus unit_factor normalize rewrites "dvd.dvd (*) = Rings.dvd" by (fact semiring_Gcd) (simp add: dvd.dvd_def dvd_def fun_eq_iff) show [simp]: "(Euclidean_Algorithm.gcd :: nat \ _) = gcd" proof (rule ext)+ fix m n :: nat show "Euclidean_Algorithm.gcd m n = gcd m n" proof (induct m n rule: eucl_induct) case zero then show ?case by simp next case (mod m n) then have "gcd n (m mod n) = gcd n m" using gcd_nat.simps [of m n] by (simp add: ac_simps) with mod show ?case by (simp add: Euclidean_Algorithm.gcd_mod ac_simps) qed qed show [simp]: "(Euclidean_Algorithm.Lcm :: nat set \ _) = Lcm" by (auto intro!: ext Lcm_eqI) show "(Euclidean_Algorithm.lcm :: nat \ _) = lcm" by (simp add: fun_eq_iff Euclidean_Algorithm.lcm_def semiring_gcd_class.lcm_gcd) show "(Euclidean_Algorithm.Gcd :: nat set \ _) = Gcd" by (simp add: fun_eq_iff Euclidean_Algorithm.Gcd_def semiring_Gcd_class.Gcd_Lcm) qed +instance nat :: normalization_euclidean_semiring_multiplicative .. + lemma prime_factorization_Suc_0 [simp]: "prime_factorization (Suc 0) = {#}" unfolding One_nat_def [symmetric] using prime_factorization_1 . instance int :: normalization_euclidean_semiring .. instance int :: euclidean_ring_gcd proof interpret semiring_Gcd 1 0 times "Euclidean_Algorithm.gcd" "Euclidean_Algorithm.lcm" "Euclidean_Algorithm.Gcd" "Euclidean_Algorithm.Lcm" divide plus minus unit_factor normalize rewrites "dvd.dvd (*) = Rings.dvd" by (fact semiring_Gcd) (simp add: dvd.dvd_def dvd_def fun_eq_iff) show [simp]: "(Euclidean_Algorithm.gcd :: int \ _) = gcd" proof (rule ext)+ fix k l :: int show "Euclidean_Algorithm.gcd k l = gcd k l" proof (induct k l rule: eucl_induct) case zero then show ?case by simp next case (mod k l) have "gcd l (k mod l) = gcd l k" proof (cases l "0::int" rule: linorder_cases) case less then show ?thesis using gcd_non_0_int [of "- l" "- k"] by (simp add: ac_simps) next case equal with mod show ?thesis by simp next case greater then show ?thesis using gcd_non_0_int [of l k] by (simp add: ac_simps) qed with mod show ?case by (simp add: Euclidean_Algorithm.gcd_mod ac_simps) qed qed show [simp]: "(Euclidean_Algorithm.Lcm :: int set \ _) = Lcm" by (auto intro!: ext Lcm_eqI) show "(Euclidean_Algorithm.lcm :: int \ _) = lcm" by (simp add: fun_eq_iff Euclidean_Algorithm.lcm_def semiring_gcd_class.lcm_gcd) show "(Euclidean_Algorithm.Gcd :: int set \ _) = Gcd" by (simp add: fun_eq_iff Euclidean_Algorithm.Gcd_def semiring_Gcd_class.Gcd_Lcm) qed +instance int :: normalization_euclidean_semiring_multiplicative .. + end diff --git a/src/HOL/Computational_Algebra/Factorial_Ring.thy b/src/HOL/Computational_Algebra/Factorial_Ring.thy --- a/src/HOL/Computational_Algebra/Factorial_Ring.thy +++ b/src/HOL/Computational_Algebra/Factorial_Ring.thy @@ -1,1947 +1,2131 @@ (* Title: HOL/Computational_Algebra/Factorial_Ring.thy Author: Manuel Eberl, TU Muenchen Author: Florian Haftmann, TU Muenchen *) section \Factorial (semi)rings\ theory Factorial_Ring imports Main "HOL-Library.Multiset" begin subsection \Irreducible and prime elements\ context comm_semiring_1 begin definition irreducible :: "'a \ bool" where "irreducible p \ p \ 0 \ \p dvd 1 \ (\a b. p = a * b \ a dvd 1 \ b dvd 1)" lemma not_irreducible_zero [simp]: "\irreducible 0" by (simp add: irreducible_def) lemma irreducible_not_unit: "irreducible p \ \p dvd 1" by (simp add: irreducible_def) lemma not_irreducible_one [simp]: "\irreducible 1" by (simp add: irreducible_def) lemma irreducibleI: "p \ 0 \ \p dvd 1 \ (\a b. p = a * b \ a dvd 1 \ b dvd 1) \ irreducible p" by (simp add: irreducible_def) lemma irreducibleD: "irreducible p \ p = a * b \ a dvd 1 \ b dvd 1" by (simp add: irreducible_def) +lemma irreducible_mono: + assumes irr: "irreducible b" and "a dvd b" "\a dvd 1" + shows "irreducible a" +proof (rule irreducibleI) + fix c d assume "a = c * d" + from assms obtain k where [simp]: "b = a * k" by auto + from \a = c * d\ have "b = c * d * k" + by simp + hence "c dvd 1 \ (d * k) dvd 1" + using irreducibleD[OF irr, of c "d * k"] by (auto simp: mult.assoc) + thus "c dvd 1 \ d dvd 1" + by auto +qed (use assms in \auto simp: irreducible_def\) + definition prime_elem :: "'a \ bool" where "prime_elem p \ p \ 0 \ \p dvd 1 \ (\a b. p dvd (a * b) \ p dvd a \ p dvd b)" lemma not_prime_elem_zero [simp]: "\prime_elem 0" by (simp add: prime_elem_def) lemma prime_elem_not_unit: "prime_elem p \ \p dvd 1" by (simp add: prime_elem_def) lemma prime_elemI: "p \ 0 \ \p dvd 1 \ (\a b. p dvd (a * b) \ p dvd a \ p dvd b) \ prime_elem p" by (simp add: prime_elem_def) lemma prime_elem_dvd_multD: "prime_elem p \ p dvd (a * b) \ p dvd a \ p dvd b" by (simp add: prime_elem_def) lemma prime_elem_dvd_mult_iff: "prime_elem p \ p dvd (a * b) \ p dvd a \ p dvd b" by (auto simp: prime_elem_def) lemma not_prime_elem_one [simp]: "\ prime_elem 1" by (auto dest: prime_elem_not_unit) lemma prime_elem_not_zeroI: assumes "prime_elem p" shows "p \ 0" using assms by (auto intro: ccontr) lemma prime_elem_dvd_power: "prime_elem p \ p dvd x ^ n \ p dvd x" by (induction n) (auto dest: prime_elem_dvd_multD intro: dvd_trans[of _ 1]) lemma prime_elem_dvd_power_iff: "prime_elem p \ n > 0 \ p dvd x ^ n \ p dvd x" by (auto dest: prime_elem_dvd_power intro: dvd_trans) lemma prime_elem_imp_nonzero [simp]: "ASSUMPTION (prime_elem x) \ x \ 0" unfolding ASSUMPTION_def by (rule prime_elem_not_zeroI) lemma prime_elem_imp_not_one [simp]: "ASSUMPTION (prime_elem x) \ x \ 1" unfolding ASSUMPTION_def by auto end + +lemma (in normalization_semidom) irreducible_cong: + assumes "normalize a = normalize b" + shows "irreducible a \ irreducible b" +proof (cases "a = 0 \ a dvd 1") + case True + hence "\irreducible a" by (auto simp: irreducible_def) + from True have "normalize a = 0 \ normalize a dvd 1" + by auto + also note assms + finally have "b = 0 \ b dvd 1" by simp + hence "\irreducible b" by (auto simp: irreducible_def) + with \\irreducible a\ show ?thesis by simp +next + case False + hence b: "b \ 0" "\is_unit b" using assms + by (auto simp: is_unit_normalize[of b]) + show ?thesis + proof + assume "irreducible a" + thus "irreducible b" + by (rule irreducible_mono) (use assms False b in \auto dest: associatedD2\) + next + assume "irreducible b" + thus "irreducible a" + by (rule irreducible_mono) (use assms False b in \auto dest: associatedD1\) + qed +qed + +lemma (in normalization_semidom) associatedE1: + assumes "normalize a = normalize b" + obtains u where "is_unit u" "a = u * b" +proof (cases "a = 0") + case [simp]: False + from assms have [simp]: "b \ 0" by auto + show ?thesis + proof (rule that) + show "is_unit (unit_factor a div unit_factor b)" + by auto + have "unit_factor a div unit_factor b * b = unit_factor a * (b div unit_factor b)" + using \b \ 0\ unit_div_commute unit_div_mult_swap unit_factor_is_unit by metis + also have "b div unit_factor b = normalize b" by simp + finally show "a = unit_factor a div unit_factor b * b" + by (metis assms unit_factor_mult_normalize) + qed +next + case [simp]: True + hence [simp]: "b = 0" + using assms[symmetric] by auto + show ?thesis + by (intro that[of 1]) auto +qed + +lemma (in normalization_semidom) associatedE2: + assumes "normalize a = normalize b" + obtains u where "is_unit u" "b = u * a" +proof - + from assms have "normalize b = normalize a" + by simp + then obtain u where "is_unit u" "b = u * a" + by (elim associatedE1) + thus ?thesis using that by blast +qed + + +(* TODO Move *) +lemma (in normalization_semidom) normalize_power_normalize: + "normalize (normalize x ^ n) = normalize (x ^ n)" +proof (induction n) + case (Suc n) + have "normalize (normalize x ^ Suc n) = normalize (x * normalize (normalize x ^ n))" + by simp + also note Suc.IH + finally show ?case by simp +qed auto + context algebraic_semidom begin lemma prime_elem_imp_irreducible: assumes "prime_elem p" shows "irreducible p" proof (rule irreducibleI) fix a b assume p_eq: "p = a * b" with assms have nz: "a \ 0" "b \ 0" by auto from p_eq have "p dvd a * b" by simp with \prime_elem p\ have "p dvd a \ p dvd b" by (rule prime_elem_dvd_multD) with \p = a * b\ have "a * b dvd 1 * b \ a * b dvd a * 1" by auto thus "a dvd 1 \ b dvd 1" by (simp only: dvd_times_left_cancel_iff[OF nz(1)] dvd_times_right_cancel_iff[OF nz(2)]) qed (insert assms, simp_all add: prime_elem_def) lemma (in algebraic_semidom) unit_imp_no_irreducible_divisors: assumes "is_unit x" "irreducible p" shows "\p dvd x" proof (rule notI) assume "p dvd x" with \is_unit x\ have "is_unit p" by (auto intro: dvd_trans) with \irreducible p\ show False by (simp add: irreducible_not_unit) qed lemma unit_imp_no_prime_divisors: assumes "is_unit x" "prime_elem p" shows "\p dvd x" using unit_imp_no_irreducible_divisors[OF assms(1) prime_elem_imp_irreducible[OF assms(2)]] . lemma prime_elem_mono: assumes "prime_elem p" "\q dvd 1" "q dvd p" shows "prime_elem q" proof - from \q dvd p\ obtain r where r: "p = q * r" by (elim dvdE) hence "p dvd q * r" by simp with \prime_elem p\ have "p dvd q \ p dvd r" by (rule prime_elem_dvd_multD) hence "p dvd q" proof assume "p dvd r" then obtain s where s: "r = p * s" by (elim dvdE) from r have "p * 1 = p * (q * s)" by (subst (asm) s) (simp add: mult_ac) with \prime_elem p\ have "q dvd 1" by (subst (asm) mult_cancel_left) auto with \\q dvd 1\ show ?thesis by contradiction qed show ?thesis proof (rule prime_elemI) fix a b assume "q dvd (a * b)" with \p dvd q\ have "p dvd (a * b)" by (rule dvd_trans) with \prime_elem p\ have "p dvd a \ p dvd b" by (rule prime_elem_dvd_multD) with \q dvd p\ show "q dvd a \ q dvd b" by (blast intro: dvd_trans) qed (insert assms, auto) qed lemma irreducibleD': assumes "irreducible a" "b dvd a" shows "a dvd b \ is_unit b" proof - from assms obtain c where c: "a = b * c" by (elim dvdE) from irreducibleD[OF assms(1) this] have "is_unit b \ is_unit c" . thus ?thesis by (auto simp: c mult_unit_dvd_iff) qed lemma irreducibleI': assumes "a \ 0" "\is_unit a" "\b. b dvd a \ a dvd b \ is_unit b" shows "irreducible a" proof (rule irreducibleI) fix b c assume a_eq: "a = b * c" hence "a dvd b \ is_unit b" by (intro assms) simp_all thus "is_unit b \ is_unit c" proof assume "a dvd b" hence "b * c dvd b * 1" by (simp add: a_eq) moreover from \a \ 0\ a_eq have "b \ 0" by auto ultimately show ?thesis by (subst (asm) dvd_times_left_cancel_iff) auto qed blast qed (simp_all add: assms(1,2)) lemma irreducible_altdef: "irreducible x \ x \ 0 \ \is_unit x \ (\b. b dvd x \ x dvd b \ is_unit b)" using irreducibleI'[of x] irreducibleD'[of x] irreducible_not_unit[of x] by auto lemma prime_elem_multD: assumes "prime_elem (a * b)" shows "is_unit a \ is_unit b" proof - from assms have "a \ 0" "b \ 0" by (auto dest!: prime_elem_not_zeroI) moreover from assms prime_elem_dvd_multD [of "a * b"] have "a * b dvd a \ a * b dvd b" by auto ultimately show ?thesis using dvd_times_left_cancel_iff [of a b 1] dvd_times_right_cancel_iff [of b a 1] by auto qed lemma prime_elemD2: assumes "prime_elem p" and "a dvd p" and "\ is_unit a" shows "p dvd a" proof - from \a dvd p\ obtain b where "p = a * b" .. with \prime_elem p\ prime_elem_multD \\ is_unit a\ have "is_unit b" by auto with \p = a * b\ show ?thesis by (auto simp add: mult_unit_dvd_iff) qed lemma prime_elem_dvd_prod_msetE: assumes "prime_elem p" assumes dvd: "p dvd prod_mset A" obtains a where "a \# A" and "p dvd a" proof - from dvd have "\a. a \# A \ p dvd a" proof (induct A) case empty then show ?case using \prime_elem p\ by (simp add: prime_elem_not_unit) next case (add a A) then have "p dvd a * prod_mset A" by simp with \prime_elem p\ consider (A) "p dvd prod_mset A" | (B) "p dvd a" by (blast dest: prime_elem_dvd_multD) then show ?case proof cases case B then show ?thesis by auto next case A with add.hyps obtain b where "b \# A" "p dvd b" by auto then show ?thesis by auto qed qed with that show thesis by blast qed context begin private lemma prime_elem_powerD: assumes "prime_elem (p ^ n)" shows "prime_elem p \ n = 1" proof (cases n) case (Suc m) note assms also from Suc have "p ^ n = p * p^m" by simp finally have "is_unit p \ is_unit (p^m)" by (rule prime_elem_multD) moreover from assms have "\is_unit p" by (simp add: prime_elem_def is_unit_power_iff) ultimately have "is_unit (p ^ m)" by simp with \\is_unit p\ have "m = 0" by (simp add: is_unit_power_iff) with Suc assms show ?thesis by simp qed (insert assms, simp_all) lemma prime_elem_power_iff: "prime_elem (p ^ n) \ prime_elem p \ n = 1" by (auto dest: prime_elem_powerD) end lemma irreducible_mult_unit_left: "is_unit a \ irreducible (a * p) \ irreducible p" by (auto simp: irreducible_altdef mult.commute[of a] is_unit_mult_iff mult_unit_dvd_iff dvd_mult_unit_iff) lemma prime_elem_mult_unit_left: "is_unit a \ prime_elem (a * p) \ prime_elem p" by (auto simp: prime_elem_def mult.commute[of a] is_unit_mult_iff mult_unit_dvd_iff) lemma prime_elem_dvd_cases: assumes pk: "p*k dvd m*n" and p: "prime_elem p" shows "(\x. k dvd x*n \ m = p*x) \ (\y. k dvd m*y \ n = p*y)" proof - have "p dvd m*n" using dvd_mult_left pk by blast then consider "p dvd m" | "p dvd n" using p prime_elem_dvd_mult_iff by blast then show ?thesis proof cases case 1 then obtain a where "m = p * a" by (metis dvd_mult_div_cancel) then have "\x. k dvd x * n \ m = p * x" using p pk by (auto simp: mult.assoc) then show ?thesis .. next case 2 then obtain b where "n = p * b" by (metis dvd_mult_div_cancel) with p pk have "\y. k dvd m*y \ n = p*y" by (metis dvd_mult_right dvd_times_left_cancel_iff mult.left_commute mult_zero_left) then show ?thesis .. qed qed lemma prime_elem_power_dvd_prod: assumes pc: "p^c dvd m*n" and p: "prime_elem p" shows "\a b. a+b = c \ p^a dvd m \ p^b dvd n" using pc proof (induct c arbitrary: m n) case 0 show ?case by simp next case (Suc c) consider x where "p^c dvd x*n" "m = p*x" | y where "p^c dvd m*y" "n = p*y" using prime_elem_dvd_cases [of _ "p^c", OF _ p] Suc.prems by force then show ?case proof cases case (1 x) with Suc.hyps[of x n] obtain a b where "a + b = c \ p ^ a dvd x \ p ^ b dvd n" by blast with 1 have "Suc a + b = Suc c \ p ^ Suc a dvd m \ p ^ b dvd n" by (auto intro: mult_dvd_mono) thus ?thesis by blast next case (2 y) with Suc.hyps[of m y] obtain a b where "a + b = c \ p ^ a dvd m \ p ^ b dvd y" by blast with 2 have "a + Suc b = Suc c \ p ^ a dvd m \ p ^ Suc b dvd n" by (auto intro: mult_dvd_mono) with Suc.hyps [of m y] show "\a b. a + b = Suc c \ p ^ a dvd m \ p ^ b dvd n" by blast qed qed lemma prime_elem_power_dvd_cases: assumes "p ^ c dvd m * n" and "a + b = Suc c" and "prime_elem p" shows "p ^ a dvd m \ p ^ b dvd n" proof - from assms obtain r s where "r + s = c \ p ^ r dvd m \ p ^ s dvd n" by (blast dest: prime_elem_power_dvd_prod) moreover with assms have "a \ r \ b \ s" by arith ultimately show ?thesis by (auto intro: power_le_dvd) qed lemma prime_elem_not_unit' [simp]: "ASSUMPTION (prime_elem x) \ \is_unit x" unfolding ASSUMPTION_def by (rule prime_elem_not_unit) lemma prime_elem_dvd_power_iff: assumes "prime_elem p" shows "p dvd a ^ n \ p dvd a \ n > 0" using assms by (induct n) (auto dest: prime_elem_not_unit prime_elem_dvd_multD) lemma prime_power_dvd_multD: assumes "prime_elem p" assumes "p ^ n dvd a * b" and "n > 0" and "\ p dvd a" shows "p ^ n dvd b" using \p ^ n dvd a * b\ and \n > 0\ proof (induct n arbitrary: b) case 0 then show ?case by simp next case (Suc n) show ?case proof (cases "n = 0") case True with Suc \prime_elem p\ \\ p dvd a\ show ?thesis by (simp add: prime_elem_dvd_mult_iff) next case False then have "n > 0" by simp from \prime_elem p\ have "p \ 0" by auto from Suc.prems have *: "p * p ^ n dvd a * b" by simp then have "p dvd a * b" by (rule dvd_mult_left) with Suc \prime_elem p\ \\ p dvd a\ have "p dvd b" by (simp add: prime_elem_dvd_mult_iff) moreover define c where "c = b div p" ultimately have b: "b = p * c" by simp with * have "p * p ^ n dvd p * (a * c)" by (simp add: ac_simps) with \p \ 0\ have "p ^ n dvd a * c" by simp with Suc.hyps \n > 0\ have "p ^ n dvd c" by blast with \p \ 0\ show ?thesis by (simp add: b) qed qed end subsection \Generalized primes: normalized prime elements\ context normalization_semidom begin lemma irreducible_normalized_divisors: assumes "irreducible x" "y dvd x" "normalize y = y" shows "y = 1 \ y = normalize x" proof - from assms have "is_unit y \ x dvd y" by (auto simp: irreducible_altdef) thus ?thesis proof (elim disjE) assume "is_unit y" hence "normalize y = 1" by (simp add: is_unit_normalize) with assms show ?thesis by simp next assume "x dvd y" with \y dvd x\ have "normalize y = normalize x" by (rule associatedI) with assms show ?thesis by simp qed qed lemma irreducible_normalize_iff [simp]: "irreducible (normalize x) = irreducible x" using irreducible_mult_unit_left[of "1 div unit_factor x" x] by (cases "x = 0") (simp_all add: unit_div_commute) lemma prime_elem_normalize_iff [simp]: "prime_elem (normalize x) = prime_elem x" using prime_elem_mult_unit_left[of "1 div unit_factor x" x] by (cases "x = 0") (simp_all add: unit_div_commute) lemma prime_elem_associated: assumes "prime_elem p" and "prime_elem q" and "q dvd p" shows "normalize q = normalize p" using \q dvd p\ proof (rule associatedI) from \prime_elem q\ have "\ is_unit q" by (auto simp add: prime_elem_not_unit) with \prime_elem p\ \q dvd p\ show "p dvd q" by (blast intro: prime_elemD2) qed definition prime :: "'a \ bool" where "prime p \ prime_elem p \ normalize p = p" lemma not_prime_0 [simp]: "\prime 0" by (simp add: prime_def) lemma not_prime_unit: "is_unit x \ \prime x" using prime_elem_not_unit[of x] by (auto simp add: prime_def) lemma not_prime_1 [simp]: "\prime 1" by (simp add: not_prime_unit) lemma primeI: "prime_elem x \ normalize x = x \ prime x" by (simp add: prime_def) lemma prime_imp_prime_elem [dest]: "prime p \ prime_elem p" by (simp add: prime_def) lemma normalize_prime: "prime p \ normalize p = p" by (simp add: prime_def) lemma prime_normalize_iff [simp]: "prime (normalize p) \ prime_elem p" by (auto simp add: prime_def) lemma prime_power_iff: "prime (p ^ n) \ prime p \ n = 1" by (auto simp: prime_def prime_elem_power_iff) lemma prime_imp_nonzero [simp]: "ASSUMPTION (prime x) \ x \ 0" unfolding ASSUMPTION_def prime_def by auto lemma prime_imp_not_one [simp]: "ASSUMPTION (prime x) \ x \ 1" unfolding ASSUMPTION_def by auto lemma prime_not_unit' [simp]: "ASSUMPTION (prime x) \ \is_unit x" unfolding ASSUMPTION_def prime_def by auto lemma prime_normalize' [simp]: "ASSUMPTION (prime x) \ normalize x = x" unfolding ASSUMPTION_def prime_def by simp lemma unit_factor_prime: "prime x \ unit_factor x = 1" using unit_factor_normalize[of x] unfolding prime_def by auto lemma unit_factor_prime' [simp]: "ASSUMPTION (prime x) \ unit_factor x = 1" unfolding ASSUMPTION_def by (rule unit_factor_prime) lemma prime_imp_prime_elem' [simp]: "ASSUMPTION (prime x) \ prime_elem x" by (simp add: prime_def ASSUMPTION_def) lemma prime_dvd_multD: "prime p \ p dvd a * b \ p dvd a \ p dvd b" by (intro prime_elem_dvd_multD) simp_all lemma prime_dvd_mult_iff: "prime p \ p dvd a * b \ p dvd a \ p dvd b" by (auto dest: prime_dvd_multD) lemma prime_dvd_power: "prime p \ p dvd x ^ n \ p dvd x" by (auto dest!: prime_elem_dvd_power simp: prime_def) lemma prime_dvd_power_iff: "prime p \ n > 0 \ p dvd x ^ n \ p dvd x" by (subst prime_elem_dvd_power_iff) simp_all lemma prime_dvd_prod_mset_iff: "prime p \ p dvd prod_mset A \ (\x. x \# A \ p dvd x)" by (induction A) (simp_all add: prime_elem_dvd_mult_iff prime_imp_prime_elem, blast+) lemma prime_dvd_prod_iff: "finite A \ prime p \ p dvd prod f A \ (\x\A. p dvd f x)" by (auto simp: prime_dvd_prod_mset_iff prod_unfold_prod_mset) lemma primes_dvd_imp_eq: assumes "prime p" "prime q" "p dvd q" shows "p = q" proof - from assms have "irreducible q" by (simp add: prime_elem_imp_irreducible prime_def) from irreducibleD'[OF this \p dvd q\] assms have "q dvd p" by simp with \p dvd q\ have "normalize p = normalize q" by (rule associatedI) with assms show "p = q" by simp qed lemma prime_dvd_prod_mset_primes_iff: assumes "prime p" "\q. q \# A \ prime q" shows "p dvd prod_mset A \ p \# A" proof - from assms(1) have "p dvd prod_mset A \ (\x. x \# A \ p dvd x)" by (rule prime_dvd_prod_mset_iff) also from assms have "\ \ p \# A" by (auto dest: primes_dvd_imp_eq) finally show ?thesis . qed lemma prod_mset_primes_dvd_imp_subset: assumes "prod_mset A dvd prod_mset B" "\p. p \# A \ prime p" "\p. p \# B \ prime p" shows "A \# B" using assms proof (induction A arbitrary: B) case empty thus ?case by simp next case (add p A B) hence p: "prime p" by simp define B' where "B' = B - {#p#}" from add.prems have "p dvd prod_mset B" by (simp add: dvd_mult_left) with add.prems have "p \# B" by (subst (asm) (2) prime_dvd_prod_mset_primes_iff) simp_all hence B: "B = B' + {#p#}" by (simp add: B'_def) from add.prems p have "A \# B'" by (intro add.IH) (simp_all add: B) thus ?case by (simp add: B) qed -lemma normalize_prod_mset_primes: - "(\p. p \# A \ prime p) \ normalize (prod_mset A) = prod_mset A" -proof (induction A) - case (add p A) - hence "prime p" by simp - hence "normalize p = p" by simp - with add show ?case by (simp add: normalize_mult) -qed simp_all - lemma prod_mset_dvd_prod_mset_primes_iff: assumes "\x. x \# A \ prime x" "\x. x \# B \ prime x" shows "prod_mset A dvd prod_mset B \ A \# B" using assms by (auto intro: prod_mset_subset_imp_dvd prod_mset_primes_dvd_imp_subset) lemma is_unit_prod_mset_primes_iff: assumes "\x. x \# A \ prime x" shows "is_unit (prod_mset A) \ A = {#}" by (auto simp add: is_unit_prod_mset_iff) (meson all_not_in_conv assms not_prime_unit set_mset_eq_empty_iff) lemma prod_mset_primes_irreducible_imp_prime: assumes irred: "irreducible (prod_mset A)" assumes A: "\x. x \# A \ prime x" assumes B: "\x. x \# B \ prime x" assumes C: "\x. x \# C \ prime x" assumes dvd: "prod_mset A dvd prod_mset B * prod_mset C" shows "prod_mset A dvd prod_mset B \ prod_mset A dvd prod_mset C" proof - from dvd have "prod_mset A dvd prod_mset (B + C)" by simp with A B C have subset: "A \# B + C" by (subst (asm) prod_mset_dvd_prod_mset_primes_iff) auto define A1 and A2 where "A1 = A \# B" and "A2 = A - A1" have "A = A1 + A2" unfolding A1_def A2_def by (rule sym, intro subset_mset.add_diff_inverse) simp_all from subset have "A1 \# B" "A2 \# C" by (auto simp: A1_def A2_def Multiset.subset_eq_diff_conv Multiset.union_commute) from \A = A1 + A2\ have "prod_mset A = prod_mset A1 * prod_mset A2" by simp from irred and this have "is_unit (prod_mset A1) \ is_unit (prod_mset A2)" by (rule irreducibleD) with A have "A1 = {#} \ A2 = {#}" unfolding A1_def A2_def by (subst (asm) (1 2) is_unit_prod_mset_primes_iff) (auto dest: Multiset.in_diffD) with dvd \A = A1 + A2\ \A1 \# B\ \A2 \# C\ show ?thesis by (auto intro: prod_mset_subset_imp_dvd) qed lemma prod_mset_primes_finite_divisor_powers: assumes A: "\x. x \# A \ prime x" assumes B: "\x. x \# B \ prime x" assumes "A \ {#}" shows "finite {n. prod_mset A ^ n dvd prod_mset B}" proof - from \A \ {#}\ obtain x where x: "x \# A" by blast define m where "m = count B x" have "{n. prod_mset A ^ n dvd prod_mset B} \ {..m}" proof safe fix n assume dvd: "prod_mset A ^ n dvd prod_mset B" from x have "x ^ n dvd prod_mset A ^ n" by (intro dvd_power_same dvd_prod_mset) also note dvd also have "x ^ n = prod_mset (replicate_mset n x)" by simp finally have "replicate_mset n x \# B" by (rule prod_mset_primes_dvd_imp_subset) (insert A B x, simp_all split: if_splits) thus "n \ m" by (simp add: count_le_replicate_mset_subset_eq m_def) qed moreover have "finite {..m}" by simp ultimately show ?thesis by (rule finite_subset) qed end subsection \In a semiring with GCD, each irreducible element is a prime element\ context semiring_gcd begin lemma irreducible_imp_prime_elem_gcd: assumes "irreducible x" shows "prime_elem x" proof (rule prime_elemI) fix a b assume "x dvd a * b" from dvd_productE[OF this] obtain y z where yz: "x = y * z" "y dvd a" "z dvd b" . from \irreducible x\ and \x = y * z\ have "is_unit y \ is_unit z" by (rule irreducibleD) with yz show "x dvd a \ x dvd b" by (auto simp: mult_unit_dvd_iff mult_unit_dvd_iff') qed (insert assms, auto simp: irreducible_not_unit) lemma prime_elem_imp_coprime: assumes "prime_elem p" "\p dvd n" shows "coprime p n" proof (rule coprimeI) fix d assume "d dvd p" "d dvd n" show "is_unit d" proof (rule ccontr) assume "\is_unit d" from \prime_elem p\ and \d dvd p\ and this have "p dvd d" by (rule prime_elemD2) from this and \d dvd n\ have "p dvd n" by (rule dvd_trans) with \\p dvd n\ show False by contradiction qed qed lemma prime_imp_coprime: assumes "prime p" "\p dvd n" shows "coprime p n" using assms by (simp add: prime_elem_imp_coprime) lemma prime_elem_imp_power_coprime: "prime_elem p \ \ p dvd a \ coprime a (p ^ m)" by (cases "m > 0") (auto dest: prime_elem_imp_coprime simp add: ac_simps) lemma prime_imp_power_coprime: "prime p \ \ p dvd a \ coprime a (p ^ m)" by (rule prime_elem_imp_power_coprime) simp_all lemma prime_elem_divprod_pow: assumes p: "prime_elem p" and ab: "coprime a b" and pab: "p^n dvd a * b" shows "p^n dvd a \ p^n dvd b" using assms proof - from p have "\ is_unit p" by simp with ab p have "\ p dvd a \ \ p dvd b" using not_coprimeI by blast with p have "coprime (p ^ n) a \ coprime (p ^ n) b" by (auto dest: prime_elem_imp_power_coprime simp add: ac_simps) with pab show ?thesis by (auto simp add: coprime_dvd_mult_left_iff coprime_dvd_mult_right_iff) qed lemma primes_coprime: "prime p \ prime q \ p \ q \ coprime p q" using prime_imp_coprime primes_dvd_imp_eq by blast end subsection \Factorial semirings: algebraic structures with unique prime factorizations\ class factorial_semiring = normalization_semidom + assumes prime_factorization_exists: - "x \ 0 \ \A. (\x. x \# A \ prime_elem x) \ prod_mset A = normalize x" + "x \ 0 \ \A. (\x. x \# A \ prime_elem x) \ normalize (prod_mset A) = normalize x" text \Alternative characterization\ lemma (in normalization_semidom) factorial_semiring_altI_aux: assumes finite_divisors: "\x. x \ 0 \ finite {y. y dvd x \ normalize y = y}" assumes irreducible_imp_prime_elem: "\x. irreducible x \ prime_elem x" assumes "x \ 0" - shows "\A. (\x. x \# A \ prime_elem x) \ prod_mset A = normalize x" + shows "\A. (\x. x \# A \ prime_elem x) \ normalize (prod_mset A) = normalize x" using \x \ 0\ proof (induction "card {b. b dvd x \ normalize b = b}" arbitrary: x rule: less_induct) case (less a) let ?fctrs = "\a. {b. b dvd a \ normalize b = b}" show ?case proof (cases "is_unit a") case True thus ?thesis by (intro exI[of _ "{#}"]) (auto simp: is_unit_normalize) next case False show ?thesis proof (cases "\b. b dvd a \ \is_unit b \ \a dvd b") case False with \\is_unit a\ less.prems have "irreducible a" by (auto simp: irreducible_altdef) hence "prime_elem a" by (rule irreducible_imp_prime_elem) thus ?thesis by (intro exI[of _ "{#normalize a#}"]) auto next case True then guess b by (elim exE conjE) note b = this from b have "?fctrs b \ ?fctrs a" by (auto intro: dvd_trans) moreover from b have "normalize a \ ?fctrs b" "normalize a \ ?fctrs a" by simp_all hence "?fctrs b \ ?fctrs a" by blast ultimately have "?fctrs b \ ?fctrs a" by (subst subset_not_subset_eq) blast with finite_divisors[OF \a \ 0\] have "card (?fctrs b) < card (?fctrs a)" by (rule psubset_card_mono) moreover from \a \ 0\ b have "b \ 0" by auto - ultimately have "\A. (\x. x \# A \ prime_elem x) \ prod_mset A = normalize b" + ultimately have "\A. (\x. x \# A \ prime_elem x) \ normalize (prod_mset A) = normalize b" by (intro less) auto then guess A .. note A = this define c where "c = a div b" from b have c: "a = b * c" by (simp add: c_def) from less.prems c have "c \ 0" by auto from b c have "?fctrs c \ ?fctrs a" by (auto intro: dvd_trans) moreover have "normalize a \ ?fctrs c" proof safe assume "normalize a dvd c" hence "b * c dvd 1 * c" by (simp add: c) hence "b dvd 1" by (subst (asm) dvd_times_right_cancel_iff) fact+ with b show False by simp qed with \normalize a \ ?fctrs a\ have "?fctrs a \ ?fctrs c" by blast ultimately have "?fctrs c \ ?fctrs a" by (subst subset_not_subset_eq) blast with finite_divisors[OF \a \ 0\] have "card (?fctrs c) < card (?fctrs a)" by (rule psubset_card_mono) - with \c \ 0\ have "\A. (\x. x \# A \ prime_elem x) \ prod_mset A = normalize c" + with \c \ 0\ have "\A. (\x. x \# A \ prime_elem x) \ normalize (prod_mset A) = normalize c" by (intro less) auto then guess B .. note B = this - from A B show ?thesis by (intro exI[of _ "A + B"]) (auto simp: c normalize_mult) + show ?thesis + proof (rule exI[of _ "A + B"]; safe) + have "normalize (prod_mset (A + B)) = + normalize (normalize (prod_mset A) * normalize (prod_mset B))" + by simp + also have "\ = normalize (b * c)" + by (simp only: A B) auto + also have "b * c = a" + using c by simp + finally show "normalize (prod_mset (A + B)) = normalize a" . + next + qed (use A B in auto) qed qed qed lemma factorial_semiring_altI: assumes finite_divisors: "\x::'a. x \ 0 \ finite {y. y dvd x \ normalize y = y}" assumes irreducible_imp_prime: "\x::'a. irreducible x \ prime_elem x" shows "OFCLASS('a :: normalization_semidom, factorial_semiring_class)" by intro_classes (rule factorial_semiring_altI_aux[OF assms]) text \Properties\ context factorial_semiring begin lemma prime_factorization_exists': assumes "x \ 0" - obtains A where "\x. x \# A \ prime x" "prod_mset A = normalize x" + obtains A where "\x. x \# A \ prime x" "normalize (prod_mset A) = normalize x" proof - from prime_factorization_exists[OF assms] obtain A - where A: "\x. x \# A \ prime_elem x" "prod_mset A = normalize x" by blast + where A: "\x. x \# A \ prime_elem x" "normalize (prod_mset A) = normalize x" by blast define A' where "A' = image_mset normalize A" - have "prod_mset A' = normalize (prod_mset A)" - by (simp add: A'_def normalize_prod_mset) + have "normalize (prod_mset A') = normalize (prod_mset A)" + by (simp add: A'_def normalize_prod_mset_normalize) also note A(2) - finally have "prod_mset A' = normalize x" by simp + finally have "normalize (prod_mset A') = normalize x" by simp moreover from A(1) have "\x. x \# A' \ prime x" by (auto simp: prime_def A'_def) ultimately show ?thesis by (intro that[of A']) blast qed lemma irreducible_imp_prime_elem: assumes "irreducible x" shows "prime_elem x" proof (rule prime_elemI) fix a b assume dvd: "x dvd a * b" from assms have "x \ 0" by auto show "x dvd a \ x dvd b" proof (cases "a = 0 \ b = 0") case False hence "a \ 0" "b \ 0" by blast+ note nz = \x \ 0\ this from nz[THEN prime_factorization_exists'] guess A B C . note ABC = this - from assms ABC have "irreducible (prod_mset A)" by simp - from dvd prod_mset_primes_irreducible_imp_prime[of A B C, OF this ABC(1,3,5)] ABC(2,4,6) - show ?thesis by (simp add: normalize_mult [symmetric]) + + have "irreducible (prod_mset A)" + by (subst irreducible_cong[OF ABC(2)]) fact + moreover have "normalize (prod_mset A) dvd + normalize (normalize (prod_mset B) * normalize (prod_mset C))" + unfolding ABC using dvd by simp + hence "prod_mset A dvd prod_mset B * prod_mset C" + unfolding normalize_mult_normalize_left normalize_mult_normalize_right by simp + ultimately have "prod_mset A dvd prod_mset B \ prod_mset A dvd prod_mset C" + by (intro prod_mset_primes_irreducible_imp_prime) (use ABC in auto) + hence "normalize (prod_mset A) dvd normalize (prod_mset B) \ + normalize (prod_mset A) dvd normalize (prod_mset C)" by simp + thus ?thesis unfolding ABC by simp qed auto qed (insert assms, simp_all add: irreducible_def) lemma finite_divisor_powers: assumes "y \ 0" "\is_unit x" shows "finite {n. x ^ n dvd y}" proof (cases "x = 0") case True with assms have "{n. x ^ n dvd y} = {0}" by (auto simp: power_0_left) thus ?thesis by simp next case False note nz = this \y \ 0\ from nz[THEN prime_factorization_exists'] guess A B . note AB = this from AB assms have "A \ {#}" by (auto simp: normalize_1_iff) from AB(2,4) prod_mset_primes_finite_divisor_powers [of A B, OF AB(1,3) this] - show ?thesis by (simp add: normalize_power [symmetric]) + have "finite {n. prod_mset A ^ n dvd prod_mset B}" by simp + also have "{n. prod_mset A ^ n dvd prod_mset B} = + {n. normalize (normalize (prod_mset A) ^ n) dvd normalize (prod_mset B)}" + unfolding normalize_power_normalize by simp + also have "\ = {n. x ^ n dvd y}" + unfolding AB unfolding normalize_power_normalize by simp + finally show ?thesis . qed lemma finite_prime_divisors: assumes "x \ 0" shows "finite {p. prime p \ p dvd x}" proof - from prime_factorization_exists'[OF assms] guess A . note A = this have "{p. prime p \ p dvd x} \ set_mset A" proof safe fix p assume p: "prime p" and dvd: "p dvd x" from dvd have "p dvd normalize x" by simp - also from A have "normalize x = prod_mset A" by simp - finally show "p \# A" using p A by (subst (asm) prime_dvd_prod_mset_primes_iff) + also from A have "normalize x = normalize (prod_mset A)" by simp + finally have "p dvd prod_mset A" + by simp + thus "p \# A" using p A + by (subst (asm) prime_dvd_prod_mset_primes_iff) qed moreover have "finite (set_mset A)" by simp ultimately show ?thesis by (rule finite_subset) qed lemma prime_elem_iff_irreducible: "prime_elem x \ irreducible x" by (blast intro: irreducible_imp_prime_elem prime_elem_imp_irreducible) lemma prime_divisor_exists: assumes "a \ 0" "\is_unit a" shows "\b. b dvd a \ prime b" proof - from prime_factorization_exists'[OF assms(1)] guess A . note A = this moreover from A and assms have "A \ {#}" by auto then obtain x where "x \# A" by blast - with A(1) have *: "x dvd prod_mset A" "prime x" by (auto simp: dvd_prod_mset) - with A have "x dvd a" by simp + with A(1) have *: "x dvd normalize (prod_mset A)" "prime x" + by (auto simp: dvd_prod_mset) + hence "x dvd a" unfolding A by simp with * show ?thesis by blast qed lemma prime_divisors_induct [case_names zero unit factor]: assumes "P 0" "\x. is_unit x \ P x" "\p x. prime p \ P x \ P (p * x)" shows "P x" proof (cases "x = 0") case False from prime_factorization_exists'[OF this] guess A . note A = this - from A(1) have "P (unit_factor x * prod_mset A)" + from A obtain u where u: "is_unit u" "x = u * prod_mset A" + by (elim associatedE2) + + from A(1) have "P (u * prod_mset A)" proof (induction A) case (add p A) from add.prems have "prime p" by simp - moreover from add.prems have "P (unit_factor x * prod_mset A)" by (intro add.IH) simp_all - ultimately have "P (p * (unit_factor x * prod_mset A))" by (rule assms(3)) + moreover from add.prems have "P (u * prod_mset A)" by (intro add.IH) simp_all + ultimately have "P (p * (u * prod_mset A))" by (rule assms(3)) thus ?case by (simp add: mult_ac) - qed (simp_all add: assms False) - with A show ?thesis by simp + qed (simp_all add: assms False u) + with A u show ?thesis by simp qed (simp_all add: assms(1)) lemma no_prime_divisors_imp_unit: assumes "a \ 0" "\b. b dvd a \ normalize b = b \ \ prime_elem b" shows "is_unit a" proof (rule ccontr) assume "\is_unit a" from prime_divisor_exists[OF assms(1) this] guess b by (elim exE conjE) with assms(2)[of b] show False by (simp add: prime_def) qed lemma prime_divisorE: assumes "a \ 0" and "\ is_unit a" obtains p where "prime p" and "p dvd a" using assms no_prime_divisors_imp_unit unfolding prime_def by blast definition multiplicity :: "'a \ 'a \ nat" where "multiplicity p x = (if finite {n. p ^ n dvd x} then Max {n. p ^ n dvd x} else 0)" lemma multiplicity_dvd: "p ^ multiplicity p x dvd x" proof (cases "finite {n. p ^ n dvd x}") case True hence "multiplicity p x = Max {n. p ^ n dvd x}" by (simp add: multiplicity_def) also have "\ \ {n. p ^ n dvd x}" by (rule Max_in) (auto intro!: True exI[of _ "0::nat"]) finally show ?thesis by simp qed (simp add: multiplicity_def) lemma multiplicity_dvd': "n \ multiplicity p x \ p ^ n dvd x" by (rule dvd_trans[OF le_imp_power_dvd multiplicity_dvd]) context fixes x p :: 'a assumes xp: "x \ 0" "\is_unit p" begin lemma multiplicity_eq_Max: "multiplicity p x = Max {n. p ^ n dvd x}" using finite_divisor_powers[OF xp] by (simp add: multiplicity_def) lemma multiplicity_geI: assumes "p ^ n dvd x" shows "multiplicity p x \ n" proof - from assms have "n \ Max {n. p ^ n dvd x}" by (intro Max_ge finite_divisor_powers xp) simp_all thus ?thesis by (subst multiplicity_eq_Max) qed lemma multiplicity_lessI: assumes "\p ^ n dvd x" shows "multiplicity p x < n" proof (rule ccontr) assume "\(n > multiplicity p x)" hence "p ^ n dvd x" by (intro multiplicity_dvd') simp with assms show False by contradiction qed lemma power_dvd_iff_le_multiplicity: "p ^ n dvd x \ n \ multiplicity p x" using multiplicity_geI[of n] multiplicity_lessI[of n] by (cases "p ^ n dvd x") auto lemma multiplicity_eq_zero_iff: shows "multiplicity p x = 0 \ \p dvd x" using power_dvd_iff_le_multiplicity[of 1] by auto lemma multiplicity_gt_zero_iff: shows "multiplicity p x > 0 \ p dvd x" using power_dvd_iff_le_multiplicity[of 1] by auto lemma multiplicity_decompose: "\p dvd (x div p ^ multiplicity p x)" proof assume *: "p dvd x div p ^ multiplicity p x" have "x = x div p ^ multiplicity p x * (p ^ multiplicity p x)" using multiplicity_dvd[of p x] by simp also from * have "x div p ^ multiplicity p x = (x div p ^ multiplicity p x div p) * p" by simp also have "x div p ^ multiplicity p x div p * p * p ^ multiplicity p x = x div p ^ multiplicity p x div p * p ^ Suc (multiplicity p x)" by (simp add: mult_assoc) also have "p ^ Suc (multiplicity p x) dvd \" by (rule dvd_triv_right) finally show False by (subst (asm) power_dvd_iff_le_multiplicity) simp qed lemma multiplicity_decompose': obtains y where "x = p ^ multiplicity p x * y" "\p dvd y" using that[of "x div p ^ multiplicity p x"] by (simp add: multiplicity_decompose multiplicity_dvd) end lemma multiplicity_zero [simp]: "multiplicity p 0 = 0" by (simp add: multiplicity_def) lemma prime_elem_multiplicity_eq_zero_iff: "prime_elem p \ x \ 0 \ multiplicity p x = 0 \ \p dvd x" by (rule multiplicity_eq_zero_iff) simp_all lemma prime_multiplicity_other: assumes "prime p" "prime q" "p \ q" shows "multiplicity p q = 0" using assms by (subst prime_elem_multiplicity_eq_zero_iff) (auto dest: primes_dvd_imp_eq) lemma prime_multiplicity_gt_zero_iff: "prime_elem p \ x \ 0 \ multiplicity p x > 0 \ p dvd x" by (rule multiplicity_gt_zero_iff) simp_all lemma multiplicity_unit_left: "is_unit p \ multiplicity p x = 0" by (simp add: multiplicity_def is_unit_power_iff unit_imp_dvd) lemma multiplicity_unit_right: assumes "is_unit x" shows "multiplicity p x = 0" proof (cases "is_unit p \ x = 0") case False with multiplicity_lessI[of x p 1] this assms show ?thesis by (auto dest: dvd_unit_imp_unit) qed (auto simp: multiplicity_unit_left) lemma multiplicity_one [simp]: "multiplicity p 1 = 0" by (rule multiplicity_unit_right) simp_all lemma multiplicity_eqI: assumes "p ^ n dvd x" "\p ^ Suc n dvd x" shows "multiplicity p x = n" proof - consider "x = 0" | "is_unit p" | "x \ 0" "\is_unit p" by blast thus ?thesis proof cases assume xp: "x \ 0" "\is_unit p" from xp assms(1) have "multiplicity p x \ n" by (intro multiplicity_geI) moreover from assms(2) xp have "multiplicity p x < Suc n" by (intro multiplicity_lessI) ultimately show ?thesis by simp next assume "is_unit p" hence "is_unit (p ^ Suc n)" by (simp add: is_unit_power_iff del: power_Suc) hence "p ^ Suc n dvd x" by (rule unit_imp_dvd) with \\p ^ Suc n dvd x\ show ?thesis by contradiction qed (insert assms, simp_all) qed context fixes x p :: 'a assumes xp: "x \ 0" "\is_unit p" begin lemma multiplicity_times_same: assumes "p \ 0" shows "multiplicity p (p * x) = Suc (multiplicity p x)" proof (rule multiplicity_eqI) show "p ^ Suc (multiplicity p x) dvd p * x" by (auto intro!: mult_dvd_mono multiplicity_dvd) from xp assms show "\ p ^ Suc (Suc (multiplicity p x)) dvd p * x" using power_dvd_iff_le_multiplicity[OF xp, of "Suc (multiplicity p x)"] by simp qed end lemma multiplicity_same_power': "multiplicity p (p ^ n) = (if p = 0 \ is_unit p then 0 else n)" proof - consider "p = 0" | "is_unit p" |"p \ 0" "\is_unit p" by blast thus ?thesis proof cases assume "p \ 0" "\is_unit p" thus ?thesis by (induction n) (simp_all add: multiplicity_times_same) qed (simp_all add: power_0_left multiplicity_unit_left) qed lemma multiplicity_same_power: "p \ 0 \ \is_unit p \ multiplicity p (p ^ n) = n" by (simp add: multiplicity_same_power') lemma multiplicity_prime_elem_times_other: assumes "prime_elem p" "\p dvd q" shows "multiplicity p (q * x) = multiplicity p x" proof (cases "x = 0") case False show ?thesis proof (rule multiplicity_eqI) have "1 * p ^ multiplicity p x dvd q * x" by (intro mult_dvd_mono multiplicity_dvd) simp_all thus "p ^ multiplicity p x dvd q * x" by simp next define n where "n = multiplicity p x" from assms have "\is_unit p" by simp from multiplicity_decompose'[OF False this] guess y . note y = this [folded n_def] from y have "p ^ Suc n dvd q * x \ p ^ n * p dvd p ^ n * (q * y)" by (simp add: mult_ac) also from assms have "\ \ p dvd q * y" by simp also have "\ \ p dvd q \ p dvd y" by (rule prime_elem_dvd_mult_iff) fact+ also from assms y have "\ \ False" by simp finally show "\(p ^ Suc n dvd q * x)" by blast qed qed simp_all lemma multiplicity_self: assumes "p \ 0" "\is_unit p" shows "multiplicity p p = 1" proof - from assms have "multiplicity p p = Max {n. p ^ n dvd p}" by (simp add: multiplicity_eq_Max) also from assms have "p ^ n dvd p \ n \ 1" for n using dvd_power_iff[of p n 1] by auto hence "{n. p ^ n dvd p} = {..1}" by auto also have "\ = {0,1}" by auto finally show ?thesis by simp qed lemma multiplicity_times_unit_left: assumes "is_unit c" shows "multiplicity (c * p) x = multiplicity p x" proof - from assms have "{n. (c * p) ^ n dvd x} = {n. p ^ n dvd x}" by (subst mult.commute) (simp add: mult_unit_dvd_iff power_mult_distrib is_unit_power_iff) thus ?thesis by (simp add: multiplicity_def) qed lemma multiplicity_times_unit_right: assumes "is_unit c" shows "multiplicity p (c * x) = multiplicity p x" proof - from assms have "{n. p ^ n dvd c * x} = {n. p ^ n dvd x}" by (subst mult.commute) (simp add: dvd_mult_unit_iff) thus ?thesis by (simp add: multiplicity_def) qed lemma multiplicity_normalize_left [simp]: "multiplicity (normalize p) x = multiplicity p x" proof (cases "p = 0") case [simp]: False have "normalize p = (1 div unit_factor p) * p" by (simp add: unit_div_commute is_unit_unit_factor) also have "multiplicity \ x = multiplicity p x" by (rule multiplicity_times_unit_left) (simp add: is_unit_unit_factor) finally show ?thesis . qed simp_all lemma multiplicity_normalize_right [simp]: "multiplicity p (normalize x) = multiplicity p x" proof (cases "x = 0") case [simp]: False have "normalize x = (1 div unit_factor x) * x" by (simp add: unit_div_commute is_unit_unit_factor) also have "multiplicity p \ = multiplicity p x" by (rule multiplicity_times_unit_right) (simp add: is_unit_unit_factor) finally show ?thesis . qed simp_all lemma multiplicity_prime [simp]: "prime_elem p \ multiplicity p p = 1" by (rule multiplicity_self) auto lemma multiplicity_prime_power [simp]: "prime_elem p \ multiplicity p (p ^ n) = n" by (subst multiplicity_same_power') auto lift_definition prime_factorization :: "'a \ 'a multiset" is "\x p. if prime p then multiplicity p x else 0" unfolding multiset_def proof clarify fix x :: 'a show "finite {p. 0 < (if prime p then multiplicity p x else 0)}" (is "finite ?A") proof (cases "x = 0") case False from False have "?A \ {p. prime p \ p dvd x}" by (auto simp: multiplicity_gt_zero_iff) moreover from False have "finite {p. prime p \ p dvd x}" by (rule finite_prime_divisors) ultimately show ?thesis by (rule finite_subset) qed simp_all qed abbreviation prime_factors :: "'a \ 'a set" where "prime_factors a \ set_mset (prime_factorization a)" lemma count_prime_factorization_nonprime: "\prime p \ count (prime_factorization x) p = 0" by transfer simp lemma count_prime_factorization_prime: "prime p \ count (prime_factorization x) p = multiplicity p x" by transfer simp lemma count_prime_factorization: "count (prime_factorization x) p = (if prime p then multiplicity p x else 0)" by transfer simp lemma dvd_imp_multiplicity_le: assumes "a dvd b" "b \ 0" shows "multiplicity p a \ multiplicity p b" proof (cases "is_unit p") case False with assms show ?thesis by (intro multiplicity_geI ) (auto intro: dvd_trans[OF multiplicity_dvd' assms(1)]) qed (insert assms, auto simp: multiplicity_unit_left) lemma prime_power_inj: assumes "prime a" "a ^ m = a ^ n" shows "m = n" proof - have "multiplicity a (a ^ m) = multiplicity a (a ^ n)" by (simp only: assms) thus ?thesis using assms by (subst (asm) (1 2) multiplicity_prime_power) simp_all qed lemma prime_power_inj': assumes "prime p" "prime q" assumes "p ^ m = q ^ n" "m > 0" "n > 0" shows "p = q" "m = n" proof - from assms have "p ^ 1 dvd p ^ m" by (intro le_imp_power_dvd) simp also have "p ^ m = q ^ n" by fact finally have "p dvd q ^ n" by simp with assms have "p dvd q" using prime_dvd_power[of p q] by simp with assms show "p = q" by (simp add: primes_dvd_imp_eq) with assms show "m = n" by (simp add: prime_power_inj) qed lemma prime_power_eq_one_iff [simp]: "prime p \ p ^ n = 1 \ n = 0" using prime_power_inj[of p n 0] by auto lemma one_eq_prime_power_iff [simp]: "prime p \ 1 = p ^ n \ n = 0" using prime_power_inj[of p 0 n] by auto lemma prime_power_inj'': assumes "prime p" "prime q" shows "p ^ m = q ^ n \ (m = 0 \ n = 0) \ (p = q \ m = n)" using assms by (cases "m = 0"; cases "n = 0") (auto dest: prime_power_inj'[OF assms]) lemma prime_factorization_0 [simp]: "prime_factorization 0 = {#}" by (simp add: multiset_eq_iff count_prime_factorization) lemma prime_factorization_empty_iff: "prime_factorization x = {#} \ x = 0 \ is_unit x" proof assume *: "prime_factorization x = {#}" { assume x: "x \ 0" "\is_unit x" { fix p assume p: "prime p" have "count (prime_factorization x) p = 0" by (simp add: *) also from p have "count (prime_factorization x) p = multiplicity p x" by (rule count_prime_factorization_prime) also from x p have "\ = 0 \ \p dvd x" by (simp add: multiplicity_eq_zero_iff) finally have "\p dvd x" . } with prime_divisor_exists[OF x] have False by blast } thus "x = 0 \ is_unit x" by blast next assume "x = 0 \ is_unit x" thus "prime_factorization x = {#}" proof assume x: "is_unit x" { fix p assume p: "prime p" from p x have "multiplicity p x = 0" by (subst multiplicity_eq_zero_iff) (auto simp: multiplicity_eq_zero_iff dest: unit_imp_no_prime_divisors) } thus ?thesis by (simp add: multiset_eq_iff count_prime_factorization) qed simp_all qed lemma prime_factorization_unit: assumes "is_unit x" shows "prime_factorization x = {#}" proof (rule multiset_eqI) fix p :: 'a show "count (prime_factorization x) p = count {#} p" proof (cases "prime p") case True with assms have "multiplicity p x = 0" by (subst multiplicity_eq_zero_iff) (auto simp: multiplicity_eq_zero_iff dest: unit_imp_no_prime_divisors) with True show ?thesis by (simp add: count_prime_factorization_prime) qed (simp_all add: count_prime_factorization_nonprime) qed lemma prime_factorization_1 [simp]: "prime_factorization 1 = {#}" by (simp add: prime_factorization_unit) lemma prime_factorization_times_prime: assumes "x \ 0" "prime p" shows "prime_factorization (p * x) = {#p#} + prime_factorization x" proof (rule multiset_eqI) fix q :: 'a consider "\prime q" | "p = q" | "prime q" "p \ q" by blast thus "count (prime_factorization (p * x)) q = count ({#p#} + prime_factorization x) q" proof cases assume q: "prime q" "p \ q" with assms primes_dvd_imp_eq[of q p] have "\q dvd p" by auto with q assms show ?thesis by (simp add: multiplicity_prime_elem_times_other count_prime_factorization) qed (insert assms, auto simp: count_prime_factorization multiplicity_times_same) qed -lemma prod_mset_prime_factorization: +lemma prod_mset_prime_factorization_weak: assumes "x \ 0" - shows "prod_mset (prime_factorization x) = normalize x" + shows "normalize (prod_mset (prime_factorization x)) = normalize x" using assms - by (induction x rule: prime_divisors_induct) - (simp_all add: prime_factorization_unit prime_factorization_times_prime - is_unit_normalize normalize_mult) +proof (induction x rule: prime_divisors_induct) + case (factor p x) + have "normalize (prod_mset (prime_factorization (p * x))) = + normalize (p * normalize (prod_mset (prime_factorization x)))" + using factor.prems factor.hyps by (simp add: prime_factorization_times_prime) + also have "normalize (prod_mset (prime_factorization x)) = normalize x" + by (rule factor.IH) (use factor in auto) + finally show ?case by simp +qed (auto simp: prime_factorization_unit is_unit_normalize) lemma in_prime_factors_iff: "p \ prime_factors x \ x \ 0 \ p dvd x \ prime p" proof - have "p \ prime_factors x \ count (prime_factorization x) p > 0" by simp also have "\ \ x \ 0 \ p dvd x \ prime p" by (subst count_prime_factorization, cases "x = 0") (auto simp: multiplicity_eq_zero_iff multiplicity_gt_zero_iff) finally show ?thesis . qed lemma in_prime_factors_imp_prime [intro]: "p \ prime_factors x \ prime p" by (simp add: in_prime_factors_iff) lemma in_prime_factors_imp_dvd [dest]: "p \ prime_factors x \ p dvd x" by (simp add: in_prime_factors_iff) lemma prime_factorsI: "x \ 0 \ prime p \ p dvd x \ p \ prime_factors x" by (auto simp: in_prime_factors_iff) lemma prime_factors_dvd: "x \ 0 \ prime_factors x = {p. prime p \ p dvd x}" by (auto intro: prime_factorsI) lemma prime_factors_multiplicity: "prime_factors n = {p. prime p \ multiplicity p n > 0}" by (cases "n = 0") (auto simp add: prime_factors_dvd prime_multiplicity_gt_zero_iff) lemma prime_factorization_prime: assumes "prime p" shows "prime_factorization p = {#p#}" proof (rule multiset_eqI) fix q :: 'a consider "\prime q" | "q = p" | "prime q" "q \ p" by blast thus "count (prime_factorization p) q = count {#p#} q" by cases (insert assms, auto dest: primes_dvd_imp_eq simp: count_prime_factorization multiplicity_self multiplicity_eq_zero_iff) qed lemma prime_factorization_prod_mset_primes: assumes "\p. p \# A \ prime p" shows "prime_factorization (prod_mset A) = A" using assms proof (induction A) case (add p A) from add.prems[of 0] have "0 \# A" by auto hence "prod_mset A \ 0" by auto with add show ?case by (simp_all add: mult_ac prime_factorization_times_prime Multiset.union_commute) qed simp_all lemma prime_factorization_cong: "normalize x = normalize y \ prime_factorization x = prime_factorization y" by (simp add: multiset_eq_iff count_prime_factorization multiplicity_normalize_right [of _ x, symmetric] multiplicity_normalize_right [of _ y, symmetric] del: multiplicity_normalize_right) lemma prime_factorization_unique: assumes "x \ 0" "y \ 0" shows "prime_factorization x = prime_factorization y \ normalize x = normalize y" proof assume "prime_factorization x = prime_factorization y" hence "prod_mset (prime_factorization x) = prod_mset (prime_factorization y)" by simp - with assms show "normalize x = normalize y" by (simp add: prod_mset_prime_factorization) + hence "normalize (prod_mset (prime_factorization x)) = + normalize (prod_mset (prime_factorization y))" + by (simp only: ) + with assms show "normalize x = normalize y" + by (simp add: prod_mset_prime_factorization_weak) qed (rule prime_factorization_cong) -lemma prime_factorization_eqI: +lemma prime_factorization_normalize [simp]: + "prime_factorization (normalize x) = prime_factorization x" + by (cases "x = 0", simp, subst prime_factorization_unique) auto + +lemma prime_factorization_eqI_strong: assumes "\p. p \# P \ prime p" "prod_mset P = n" shows "prime_factorization n = P" using prime_factorization_prod_mset_primes[of P] assms by simp +lemma prime_factorization_eqI: + assumes "\p. p \# P \ prime p" "normalize (prod_mset P) = normalize n" + shows "prime_factorization n = P" +proof - + have "P = prime_factorization (normalize (prod_mset P))" + using prime_factorization_prod_mset_primes[of P] assms(1) by simp + with assms(2) show ?thesis by simp +qed + lemma prime_factorization_mult: assumes "x \ 0" "y \ 0" shows "prime_factorization (x * y) = prime_factorization x + prime_factorization y" proof - - have "prime_factorization (prod_mset (prime_factorization (x * y))) = - prime_factorization (prod_mset (prime_factorization x + prime_factorization y))" - by (simp add: prod_mset_prime_factorization assms normalize_mult) - also have "prime_factorization (prod_mset (prime_factorization (x * y))) = - prime_factorization (x * y)" - by (rule prime_factorization_prod_mset_primes) (simp_all add: in_prime_factors_imp_prime) - also have "prime_factorization (prod_mset (prime_factorization x + prime_factorization y)) = - prime_factorization x + prime_factorization y" - by (rule prime_factorization_prod_mset_primes) (auto simp: in_prime_factors_imp_prime) - finally show ?thesis . + have "normalize (prod_mset (prime_factorization x) * prod_mset (prime_factorization y)) = + normalize (normalize (prod_mset (prime_factorization x)) * + normalize (prod_mset (prime_factorization y)))" + by (simp only: normalize_mult_normalize_left normalize_mult_normalize_right) + also have "\ = normalize (x * y)" + by (subst (1 2) prod_mset_prime_factorization_weak) (use assms in auto) + finally show ?thesis + by (intro prime_factorization_eqI) auto qed lemma prime_factorization_prod: assumes "finite A" "\x. x \ A \ f x \ 0" shows "prime_factorization (prod f A) = (\n\A. prime_factorization (f n))" using assms by (induction A rule: finite_induct) (auto simp: Sup_multiset_empty prime_factorization_mult) lemma prime_elem_multiplicity_mult_distrib: assumes "prime_elem p" "x \ 0" "y \ 0" shows "multiplicity p (x * y) = multiplicity p x + multiplicity p y" proof - have "multiplicity p (x * y) = count (prime_factorization (x * y)) (normalize p)" by (subst count_prime_factorization_prime) (simp_all add: assms) also from assms have "prime_factorization (x * y) = prime_factorization x + prime_factorization y" by (intro prime_factorization_mult) also have "count \ (normalize p) = count (prime_factorization x) (normalize p) + count (prime_factorization y) (normalize p)" by simp also have "\ = multiplicity p x + multiplicity p y" by (subst (1 2) count_prime_factorization_prime) (simp_all add: assms) finally show ?thesis . qed lemma prime_elem_multiplicity_prod_mset_distrib: assumes "prime_elem p" "0 \# A" shows "multiplicity p (prod_mset A) = sum_mset (image_mset (multiplicity p) A)" using assms by (induction A) (auto simp: prime_elem_multiplicity_mult_distrib) lemma prime_elem_multiplicity_power_distrib: assumes "prime_elem p" "x \ 0" shows "multiplicity p (x ^ n) = n * multiplicity p x" using assms prime_elem_multiplicity_prod_mset_distrib [of p "replicate_mset n x"] by simp lemma prime_elem_multiplicity_prod_distrib: assumes "prime_elem p" "0 \ f ` A" "finite A" shows "multiplicity p (prod f A) = (\x\A. multiplicity p (f x))" proof - have "multiplicity p (prod f A) = (\x\#mset_set A. multiplicity p (f x))" using assms by (subst prod_unfold_prod_mset) (simp_all add: prime_elem_multiplicity_prod_mset_distrib sum_unfold_sum_mset multiset.map_comp o_def) also from \finite A\ have "\ = (\x\A. multiplicity p (f x))" by (induction A rule: finite_induct) simp_all finally show ?thesis . qed lemma multiplicity_distinct_prime_power: "prime p \ prime q \ p \ q \ multiplicity p (q ^ n) = 0" by (subst prime_elem_multiplicity_power_distrib) (auto simp: prime_multiplicity_other) lemma prime_factorization_prime_power: "prime p \ prime_factorization (p ^ n) = replicate_mset n p" by (induction n) (simp_all add: prime_factorization_mult prime_factorization_prime Multiset.union_commute) -lemma prime_decomposition: "unit_factor x * prod_mset (prime_factorization x) = x" - by (cases "x = 0") (simp_all add: prod_mset_prime_factorization) - lemma prime_factorization_subset_iff_dvd: assumes [simp]: "x \ 0" "y \ 0" shows "prime_factorization x \# prime_factorization y \ x dvd y" proof - - have "x dvd y \ prod_mset (prime_factorization x) dvd prod_mset (prime_factorization y)" - by (simp add: prod_mset_prime_factorization) + have "x dvd y \ + normalize (prod_mset (prime_factorization x)) dvd normalize (prod_mset (prime_factorization y))" + using assms by (subst (1 2) prod_mset_prime_factorization_weak) auto also have "\ \ prime_factorization x \# prime_factorization y" by (auto intro!: prod_mset_primes_dvd_imp_subset prod_mset_subset_imp_dvd) finally show ?thesis .. qed lemma prime_factorization_subset_imp_dvd: "x \ 0 \ (prime_factorization x \# prime_factorization y) \ x dvd y" by (cases "y = 0") (simp_all add: prime_factorization_subset_iff_dvd) lemma prime_factorization_divide: assumes "b dvd a" shows "prime_factorization (a div b) = prime_factorization a - prime_factorization b" proof (cases "a = 0") case [simp]: False from assms have [simp]: "b \ 0" by auto have "prime_factorization ((a div b) * b) = prime_factorization (a div b) + prime_factorization b" by (intro prime_factorization_mult) (insert assms, auto elim!: dvdE) with assms show ?thesis by simp qed simp_all lemma zero_not_in_prime_factors [simp]: "0 \ prime_factors x" by (auto dest: in_prime_factors_imp_prime) lemma prime_prime_factors: "prime p \ prime_factors p = {p}" by (drule prime_factorization_prime) simp -lemma prod_prime_factors: - assumes "x \ 0" - shows "(\p \ prime_factors x. p ^ multiplicity p x) = normalize x" -proof - - have "normalize x = prod_mset (prime_factorization x)" - by (simp add: prod_mset_prime_factorization assms) - also have "\ = (\p \ prime_factors x. p ^ count (prime_factorization x) p)" - by (subst prod_mset_multiplicity) simp_all - also have "\ = (\p \ prime_factors x. p ^ multiplicity p x)" - by (intro prod.cong) - (simp_all add: assms count_prime_factorization_prime in_prime_factors_imp_prime) - finally show ?thesis .. -qed - -lemma prime_factorization_unique'': - assumes S_eq: "S = {p. 0 < f p}" - and "finite S" - and S: "\p\S. prime p" "normalize n = (\p\S. p ^ f p)" - shows "S = prime_factors n \ (\p. prime p \ f p = multiplicity p n)" -proof - define A where "A = Abs_multiset f" - from \finite S\ S(1) have "(\p\S. p ^ f p) \ 0" by auto - with S(2) have nz: "n \ 0" by auto - from S_eq \finite S\ have count_A: "count A = f" - unfolding A_def by (subst multiset.Abs_multiset_inverse) (simp_all add: multiset_def) - from S_eq count_A have set_mset_A: "set_mset A = S" - by (simp only: set_mset_def) - from S(2) have "normalize n = (\p\S. p ^ f p)" . - also have "\ = prod_mset A" by (simp add: prod_mset_multiplicity S_eq set_mset_A count_A) - also from nz have "normalize n = prod_mset (prime_factorization n)" - by (simp add: prod_mset_prime_factorization) - finally have "prime_factorization (prod_mset A) = - prime_factorization (prod_mset (prime_factorization n))" by simp - also from S(1) have "prime_factorization (prod_mset A) = A" - by (intro prime_factorization_prod_mset_primes) (auto simp: set_mset_A) - also have "prime_factorization (prod_mset (prime_factorization n)) = prime_factorization n" - by (intro prime_factorization_prod_mset_primes) auto - finally show "S = prime_factors n" by (simp add: set_mset_A [symmetric]) - - show "(\p. prime p \ f p = multiplicity p n)" - proof safe - fix p :: 'a assume p: "prime p" - have "multiplicity p n = multiplicity p (normalize n)" by simp - also have "normalize n = prod_mset A" - by (simp add: prod_mset_multiplicity S_eq set_mset_A count_A S) - also from p set_mset_A S(1) - have "multiplicity p \ = sum_mset (image_mset (multiplicity p) A)" - by (intro prime_elem_multiplicity_prod_mset_distrib) auto - also from S(1) p - have "image_mset (multiplicity p) A = image_mset (\q. if p = q then 1 else 0) A" - by (intro image_mset_cong) (auto simp: set_mset_A multiplicity_self prime_multiplicity_other) - also have "sum_mset \ = f p" - by (simp add: semiring_1_class.sum_mset_delta' count_A) - finally show "f p = multiplicity p n" .. - qed -qed - lemma prime_factors_product: "x \ 0 \ y \ 0 \ prime_factors (x * y) = prime_factors x \ prime_factors y" by (simp add: prime_factorization_mult) lemma dvd_prime_factors [intro]: "y \ 0 \ x dvd y \ prime_factors x \ prime_factors y" by (intro set_mset_mono, subst prime_factorization_subset_iff_dvd) auto (* RENAMED multiplicity_dvd *) lemma multiplicity_le_imp_dvd: assumes "x \ 0" "\p. prime p \ multiplicity p x \ multiplicity p y" shows "x dvd y" proof (cases "y = 0") case False from assms this have "prime_factorization x \# prime_factorization y" by (intro mset_subset_eqI) (auto simp: count_prime_factorization) with assms False show ?thesis by (subst (asm) prime_factorization_subset_iff_dvd) qed auto lemma dvd_multiplicity_eq: "x \ 0 \ y \ 0 \ x dvd y \ (\p. multiplicity p x \ multiplicity p y)" by (auto intro: dvd_imp_multiplicity_le multiplicity_le_imp_dvd) lemma multiplicity_eq_imp_eq: assumes "x \ 0" "y \ 0" assumes "\p. prime p \ multiplicity p x = multiplicity p y" shows "normalize x = normalize y" using assms by (intro associatedI multiplicity_le_imp_dvd) simp_all lemma prime_factorization_unique': assumes "\p \# M. prime p" "\p \# N. prime p" "(\i \# M. i) = (\i \# N. i)" shows "M = N" proof - have "prime_factorization (\i \# M. i) = prime_factorization (\i \# N. i)" by (simp only: assms) also from assms have "prime_factorization (\i \# M. i) = M" by (subst prime_factorization_prod_mset_primes) simp_all also from assms have "prime_factorization (\i \# N. i) = N" by (subst prime_factorization_prod_mset_primes) simp_all finally show ?thesis . qed +lemma prime_factorization_unique'': + assumes "\p \# M. prime p" "\p \# N. prime p" "normalize (\i \# M. i) = normalize (\i \# N. i)" + shows "M = N" +proof - + have "prime_factorization (normalize (\i \# M. i)) = + prime_factorization (normalize (\i \# N. i))" + by (simp only: assms) + also from assms have "prime_factorization (normalize (\i \# M. i)) = M" + by (subst prime_factorization_normalize, subst prime_factorization_prod_mset_primes) simp_all + also from assms have "prime_factorization (normalize (\i \# N. i)) = N" + by (subst prime_factorization_normalize, subst prime_factorization_prod_mset_primes) simp_all + finally show ?thesis . +qed + lemma multiplicity_cong: "(\r. p ^ r dvd a \ p ^ r dvd b) \ multiplicity p a = multiplicity p b" by (simp add: multiplicity_def) lemma not_dvd_imp_multiplicity_0: assumes "\p dvd x" shows "multiplicity p x = 0" proof - from assms have "multiplicity p x < 1" by (intro multiplicity_lessI) auto thus ?thesis by simp qed lemma inj_on_Prod_primes: assumes "\P p. P \ A \ p \ P \ prime p" assumes "\P. P \ A \ finite P" shows "inj_on Prod A" proof (rule inj_onI) fix P Q assume PQ: "P \ A" "Q \ A" "\P = \Q" with prime_factorization_unique'[of "mset_set P" "mset_set Q"] assms[of P] assms[of Q] have "mset_set P = mset_set Q" by (auto simp: prod_unfold_prod_mset) with assms[of P] assms[of Q] PQ show "P = Q" by simp qed -lemma divides_primepow: +lemma divides_primepow_weak: assumes "prime p" and "a dvd p ^ n" - obtains m where "m \ n" and "normalize a = p ^ m" + obtains m where "m \ n" and "normalize a = normalize (p ^ m)" proof - from assms have "a \ 0" by auto with assms - have "prod_mset (prime_factorization a) dvd prod_mset (prime_factorization (p ^ n))" - by (simp add: prod_mset_prime_factorization) + have "normalize (prod_mset (prime_factorization a)) dvd + normalize (prod_mset (prime_factorization (p ^ n)))" + by (subst (1 2) prod_mset_prime_factorization_weak) auto then have "prime_factorization a \# prime_factorization (p ^ n)" by (simp add: in_prime_factors_imp_prime prod_mset_dvd_prod_mset_primes_iff) with assms have "prime_factorization a \# replicate_mset n p" by (simp add: prime_factorization_prime_power) then obtain m where "m \ n" and "prime_factorization a = replicate_mset m p" by (rule msubseteq_replicate_msetE) - then have "prod_mset (prime_factorization a) = prod_mset (replicate_mset m p)" + then have *: "normalize (prod_mset (prime_factorization a)) = + normalize (prod_mset (replicate_mset m p))" by metis + also have "normalize (prod_mset (prime_factorization a)) = normalize a" + using \a \ 0\ by (simp add: prod_mset_prime_factorization_weak) + also have "prod_mset (replicate_mset m p) = p ^ m" by simp - with \a \ 0\ have "normalize a = p ^ m" - by (simp add: prod_mset_prime_factorization) - with \m \ n\ show thesis .. + finally show ?thesis using \m \ n\ + by (intro that[of m]) qed lemma divide_out_primepow_ex: assumes "n \ 0" "\p\prime_factors n. P p" obtains p k n' where "P p" "prime p" "p dvd n" "\p dvd n'" "k > 0" "n = p ^ k * n'" proof - from assms obtain p where p: "P p" "prime p" "p dvd n" by auto define k where "k = multiplicity p n" define n' where "n' = n div p ^ k" have n': "n = p ^ k * n'" "\p dvd n'" using assms p multiplicity_decompose[of n p] by (auto simp: n'_def k_def multiplicity_dvd) from n' p have "k > 0" by (intro Nat.gr0I) auto with n' p that[of p n' k] show ?thesis by auto qed lemma divide_out_primepow: assumes "n \ 0" "\is_unit n" obtains p k n' where "prime p" "p dvd n" "\p dvd n'" "k > 0" "n = p ^ k * n'" using divide_out_primepow_ex[OF assms(1), of "\_. True"] prime_divisor_exists[OF assms] assms prime_factorsI by metis - -lemma Ex_other_prime_factor: - assumes "n \ 0" and "\(\k. normalize n = p ^ k)" "prime p" - shows "\q\prime_factors n. q \ p" -proof (rule ccontr) - assume *: "\(\q\prime_factors n. q \ p)" - have "normalize n = (\p\prime_factors n. p ^ multiplicity p n)" - using assms(1) by (intro prod_prime_factors [symmetric]) auto - also from * have "\ = (\p\{p}. p ^ multiplicity p n)" - using assms(3) by (intro prod.mono_neutral_left) (auto simp: prime_factors_multiplicity) - finally have "normalize n = p ^ multiplicity p n" by auto - with assms show False by auto -qed subsection \GCD and LCM computation with unique factorizations\ definition "gcd_factorial a b = (if a = 0 then normalize b else if b = 0 then normalize a - else prod_mset (prime_factorization a \# prime_factorization b))" + else normalize (prod_mset (prime_factorization a \# prime_factorization b)))" definition "lcm_factorial a b = (if a = 0 \ b = 0 then 0 - else prod_mset (prime_factorization a \# prime_factorization b))" + else normalize (prod_mset (prime_factorization a \# prime_factorization b)))" definition "Gcd_factorial A = - (if A \ {0} then 0 else prod_mset (Inf (prime_factorization ` (A - {0}))))" + (if A \ {0} then 0 else normalize (prod_mset (Inf (prime_factorization ` (A - {0})))))" definition "Lcm_factorial A = (if A = {} then 1 else if 0 \ A \ subset_mset.bdd_above (prime_factorization ` (A - {0})) then - prod_mset (Sup (prime_factorization ` A)) + normalize (prod_mset (Sup (prime_factorization ` A))) else 0)" lemma prime_factorization_gcd_factorial: assumes [simp]: "a \ 0" "b \ 0" shows "prime_factorization (gcd_factorial a b) = prime_factorization a \# prime_factorization b" proof - have "prime_factorization (gcd_factorial a b) = prime_factorization (prod_mset (prime_factorization a \# prime_factorization b))" by (simp add: gcd_factorial_def) also have "\ = prime_factorization a \# prime_factorization b" by (subst prime_factorization_prod_mset_primes) auto finally show ?thesis . qed lemma prime_factorization_lcm_factorial: assumes [simp]: "a \ 0" "b \ 0" shows "prime_factorization (lcm_factorial a b) = prime_factorization a \# prime_factorization b" proof - have "prime_factorization (lcm_factorial a b) = prime_factorization (prod_mset (prime_factorization a \# prime_factorization b))" by (simp add: lcm_factorial_def) also have "\ = prime_factorization a \# prime_factorization b" by (subst prime_factorization_prod_mset_primes) auto finally show ?thesis . qed lemma prime_factorization_Gcd_factorial: assumes "\A \ {0}" shows "prime_factorization (Gcd_factorial A) = Inf (prime_factorization ` (A - {0}))" proof - from assms obtain x where x: "x \ A - {0}" by auto hence "Inf (prime_factorization ` (A - {0})) \# prime_factorization x" by (intro subset_mset.cInf_lower) simp_all hence "\y. y \# Inf (prime_factorization ` (A - {0})) \ y \ prime_factors x" by (auto dest: mset_subset_eqD) with in_prime_factors_imp_prime[of _ x] have "\p. p \# Inf (prime_factorization ` (A - {0})) \ prime p" by blast with assms show ?thesis by (simp add: Gcd_factorial_def prime_factorization_prod_mset_primes) qed lemma prime_factorization_Lcm_factorial: assumes "0 \ A" "subset_mset.bdd_above (prime_factorization ` A)" shows "prime_factorization (Lcm_factorial A) = Sup (prime_factorization ` A)" proof (cases "A = {}") case True hence "prime_factorization ` A = {}" by auto also have "Sup \ = {#}" by (simp add: Sup_multiset_empty) finally show ?thesis by (simp add: Lcm_factorial_def) next case False have "\y. y \# Sup (prime_factorization ` A) \ prime y" by (auto simp: in_Sup_multiset_iff assms) with assms False show ?thesis by (simp add: Lcm_factorial_def prime_factorization_prod_mset_primes) qed lemma gcd_factorial_commute: "gcd_factorial a b = gcd_factorial b a" by (simp add: gcd_factorial_def multiset_inter_commute) lemma gcd_factorial_dvd1: "gcd_factorial a b dvd a" proof (cases "a = 0 \ b = 0") case False hence "gcd_factorial a b \ 0" by (auto simp: gcd_factorial_def) with False show ?thesis by (subst prime_factorization_subset_iff_dvd [symmetric]) (auto simp: prime_factorization_gcd_factorial) qed (auto simp: gcd_factorial_def) lemma gcd_factorial_dvd2: "gcd_factorial a b dvd b" by (subst gcd_factorial_commute) (rule gcd_factorial_dvd1) -lemma normalize_gcd_factorial: "normalize (gcd_factorial a b) = gcd_factorial a b" -proof - - have "normalize (prod_mset (prime_factorization a \# prime_factorization b)) = - prod_mset (prime_factorization a \# prime_factorization b)" - by (intro normalize_prod_mset_primes) auto - thus ?thesis by (simp add: gcd_factorial_def) -qed +lemma normalize_gcd_factorial [simp]: "normalize (gcd_factorial a b) = gcd_factorial a b" + by (simp add: gcd_factorial_def) + +lemma normalize_lcm_factorial [simp]: "normalize (lcm_factorial a b) = lcm_factorial a b" + by (simp add: lcm_factorial_def) lemma gcd_factorial_greatest: "c dvd gcd_factorial a b" if "c dvd a" "c dvd b" for a b c proof (cases "a = 0 \ b = 0") case False with that have [simp]: "c \ 0" by auto let ?p = "prime_factorization" from that False have "?p c \# ?p a" "?p c \# ?p b" by (simp_all add: prime_factorization_subset_iff_dvd) hence "prime_factorization c \# prime_factorization (prod_mset (prime_factorization a \# prime_factorization b))" using False by (subst prime_factorization_prod_mset_primes) auto with False show ?thesis by (auto simp: gcd_factorial_def prime_factorization_subset_iff_dvd [symmetric]) qed (auto simp: gcd_factorial_def that) lemma lcm_factorial_gcd_factorial: - "lcm_factorial a b = normalize (a * b) div gcd_factorial a b" for a b + "lcm_factorial a b = normalize (a * b div gcd_factorial a b)" for a b proof (cases "a = 0 \ b = 0") case False let ?p = "prime_factorization" - from False have "prod_mset (?p (a * b)) div gcd_factorial a b = - prod_mset (?p a + ?p b - ?p a \# ?p b)" - by (subst prod_mset_diff) (auto simp: lcm_factorial_def gcd_factorial_def - prime_factorization_mult subset_mset.le_infI1) - also from False have "prod_mset (?p (a * b)) = normalize (a * b)" - by (intro prod_mset_prime_factorization) simp_all - also from False have "prod_mset (?p a + ?p b - ?p a \# ?p b) = lcm_factorial a b" - by (simp add: union_diff_inter_eq_sup lcm_factorial_def) - finally show ?thesis .. + have 1: "normalize x * normalize y dvd z \ x * y dvd z" for x y z :: 'a + proof - + have "normalize (normalize x * normalize y) dvd z \ x * y dvd z" + unfolding normalize_mult_normalize_left normalize_mult_normalize_right by simp + thus ?thesis unfolding normalize_dvd_iff by simp + qed + + have "?p (a * b) = (?p a \# ?p b) + (?p a \# ?p b)" + using False by (subst prime_factorization_mult) (auto intro!: multiset_eqI) + hence "normalize (prod_mset (?p (a * b))) = + normalize (prod_mset ((?p a \# ?p b) + (?p a \# ?p b)))" + by (simp only:) + hence *: "normalize (a * b) = normalize (lcm_factorial a b * gcd_factorial a b)" using False + by (subst (asm) prod_mset_prime_factorization_weak) + (auto simp: lcm_factorial_def gcd_factorial_def) + + have [simp]: "gcd_factorial a b dvd a * b" "lcm_factorial a b dvd a * b" + using associatedD2[OF *] by auto + from False have [simp]: "gcd_factorial a b \ 0" "lcm_factorial a b \ 0" + by (auto simp: gcd_factorial_def lcm_factorial_def) + + show ?thesis + by (rule associated_eqI) + (use * in \auto simp: dvd_div_iff_mult div_dvd_iff_mult dest: associatedD1 associatedD2\) qed (auto simp: lcm_factorial_def) lemma normalize_Gcd_factorial: "normalize (Gcd_factorial A) = Gcd_factorial A" -proof (cases "A \ {0}") - case False - then obtain x where "x \ A" "x \ 0" by blast - hence "Inf (prime_factorization ` (A - {0})) \# prime_factorization x" - by (intro subset_mset.cInf_lower) auto - hence "prime p" if "p \# Inf (prime_factorization ` (A - {0}))" for p - using that by (auto dest: mset_subset_eqD) - with False show ?thesis - by (auto simp add: Gcd_factorial_def normalize_prod_mset_primes) -qed (simp_all add: Gcd_factorial_def) + by (simp add: Gcd_factorial_def) lemma Gcd_factorial_eq_0_iff: "Gcd_factorial A = 0 \ A \ {0}" by (auto simp: Gcd_factorial_def in_Inf_multiset_iff split: if_splits) lemma Gcd_factorial_dvd: assumes "x \ A" shows "Gcd_factorial A dvd x" proof (cases "x = 0") case False with assms have "prime_factorization (Gcd_factorial A) = Inf (prime_factorization ` (A - {0}))" by (intro prime_factorization_Gcd_factorial) auto also from False assms have "\ \# prime_factorization x" by (intro subset_mset.cInf_lower) auto finally show ?thesis by (subst (asm) prime_factorization_subset_iff_dvd) (insert assms False, auto simp: Gcd_factorial_eq_0_iff) qed simp_all lemma Gcd_factorial_greatest: assumes "\y. y \ A \ x dvd y" shows "x dvd Gcd_factorial A" proof (cases "A \ {0}") case False from False obtain y where "y \ A" "y \ 0" by auto with assms[of y] have nz: "x \ 0" by auto from nz assms have "prime_factorization x \# prime_factorization y" if "y \ A - {0}" for y using that by (subst prime_factorization_subset_iff_dvd) auto with False have "prime_factorization x \# Inf (prime_factorization ` (A - {0}))" by (intro subset_mset.cInf_greatest) auto also from False have "\ = prime_factorization (Gcd_factorial A)" by (rule prime_factorization_Gcd_factorial [symmetric]) finally show ?thesis by (subst (asm) prime_factorization_subset_iff_dvd) (insert nz False, auto simp: Gcd_factorial_eq_0_iff) qed (simp_all add: Gcd_factorial_def) lemma normalize_Lcm_factorial: "normalize (Lcm_factorial A) = Lcm_factorial A" -proof (cases "subset_mset.bdd_above (prime_factorization ` A)") - case True - hence "normalize (prod_mset (Sup (prime_factorization ` A))) = - prod_mset (Sup (prime_factorization ` A))" - by (intro normalize_prod_mset_primes) - (auto simp: in_Sup_multiset_iff) - with True show ?thesis by (simp add: Lcm_factorial_def) -qed (auto simp: Lcm_factorial_def) + by (simp add: Lcm_factorial_def) lemma Lcm_factorial_eq_0_iff: "Lcm_factorial A = 0 \ 0 \ A \ \subset_mset.bdd_above (prime_factorization ` A)" by (auto simp: Lcm_factorial_def in_Sup_multiset_iff) lemma dvd_Lcm_factorial: assumes "x \ A" shows "x dvd Lcm_factorial A" proof (cases "0 \ A \ subset_mset.bdd_above (prime_factorization ` A)") case True with assms have [simp]: "0 \ A" "x \ 0" "A \ {}" by auto from assms True have "prime_factorization x \# Sup (prime_factorization ` A)" by (intro subset_mset.cSup_upper) auto also have "\ = prime_factorization (Lcm_factorial A)" by (rule prime_factorization_Lcm_factorial [symmetric]) (insert True, simp_all) finally show ?thesis by (subst (asm) prime_factorization_subset_iff_dvd) (insert True, auto simp: Lcm_factorial_eq_0_iff) qed (insert assms, auto simp: Lcm_factorial_def) lemma Lcm_factorial_least: assumes "\y. y \ A \ y dvd x" shows "Lcm_factorial A dvd x" proof - consider "A = {}" | "0 \ A" | "x = 0" | "A \ {}" "0 \ A" "x \ 0" by blast thus ?thesis proof cases assume *: "A \ {}" "0 \ A" "x \ 0" hence nz: "x \ 0" if "x \ A" for x using that by auto from * have bdd: "subset_mset.bdd_above (prime_factorization ` A)" by (intro subset_mset.bdd_aboveI[of _ "prime_factorization x"]) (auto simp: prime_factorization_subset_iff_dvd nz dest: assms) have "prime_factorization (Lcm_factorial A) = Sup (prime_factorization ` A)" by (rule prime_factorization_Lcm_factorial) fact+ also from * have "\ \# prime_factorization x" by (intro subset_mset.cSup_least) (auto simp: prime_factorization_subset_iff_dvd nz dest: assms) finally show ?thesis by (subst (asm) prime_factorization_subset_iff_dvd) (insert * bdd, auto simp: Lcm_factorial_eq_0_iff) qed (auto simp: Lcm_factorial_def dest: assms) qed lemmas gcd_lcm_factorial = gcd_factorial_dvd1 gcd_factorial_dvd2 gcd_factorial_greatest normalize_gcd_factorial lcm_factorial_gcd_factorial normalize_Gcd_factorial Gcd_factorial_dvd Gcd_factorial_greatest normalize_Lcm_factorial dvd_Lcm_factorial Lcm_factorial_least end class factorial_semiring_gcd = factorial_semiring + gcd + Gcd + assumes gcd_eq_gcd_factorial: "gcd a b = gcd_factorial a b" and lcm_eq_lcm_factorial: "lcm a b = lcm_factorial a b" and Gcd_eq_Gcd_factorial: "Gcd A = Gcd_factorial A" and Lcm_eq_Lcm_factorial: "Lcm A = Lcm_factorial A" begin lemma prime_factorization_gcd: assumes [simp]: "a \ 0" "b \ 0" shows "prime_factorization (gcd a b) = prime_factorization a \# prime_factorization b" by (simp add: gcd_eq_gcd_factorial prime_factorization_gcd_factorial) lemma prime_factorization_lcm: assumes [simp]: "a \ 0" "b \ 0" shows "prime_factorization (lcm a b) = prime_factorization a \# prime_factorization b" by (simp add: lcm_eq_lcm_factorial prime_factorization_lcm_factorial) lemma prime_factorization_Gcd: assumes "Gcd A \ 0" shows "prime_factorization (Gcd A) = Inf (prime_factorization ` (A - {0}))" using assms by (simp add: prime_factorization_Gcd_factorial Gcd_eq_Gcd_factorial Gcd_factorial_eq_0_iff) lemma prime_factorization_Lcm: assumes "Lcm A \ 0" shows "prime_factorization (Lcm A) = Sup (prime_factorization ` A)" using assms by (simp add: prime_factorization_Lcm_factorial Lcm_eq_Lcm_factorial Lcm_factorial_eq_0_iff) lemma prime_factors_gcd [simp]: "a \ 0 \ b \ 0 \ prime_factors (gcd a b) = prime_factors a \ prime_factors b" by (subst prime_factorization_gcd) auto lemma prime_factors_lcm [simp]: "a \ 0 \ b \ 0 \ prime_factors (lcm a b) = prime_factors a \ prime_factors b" by (subst prime_factorization_lcm) auto subclass semiring_gcd by (standard, unfold gcd_eq_gcd_factorial lcm_eq_lcm_factorial) (rule gcd_lcm_factorial; assumption)+ subclass semiring_Gcd by (standard, unfold Gcd_eq_Gcd_factorial Lcm_eq_Lcm_factorial) (rule gcd_lcm_factorial; assumption)+ lemma assumes "x \ 0" "y \ 0" shows gcd_eq_factorial': - "gcd x y = (\p \ prime_factors x \ prime_factors y. + "gcd x y = normalize (\p \ prime_factors x \ prime_factors y. p ^ min (multiplicity p x) (multiplicity p y))" (is "_ = ?rhs1") and lcm_eq_factorial': - "lcm x y = (\p \ prime_factors x \ prime_factors y. + "lcm x y = normalize (\p \ prime_factors x \ prime_factors y. p ^ max (multiplicity p x) (multiplicity p y))" (is "_ = ?rhs2") proof - have "gcd x y = gcd_factorial x y" by (rule gcd_eq_gcd_factorial) also have "\ = ?rhs1" by (auto simp: gcd_factorial_def assms prod_mset_multiplicity - count_prime_factorization_prime dest: in_prime_factors_imp_prime intro!: prod.cong) + count_prime_factorization_prime + intro!: arg_cong[of _ _ normalize] dest: in_prime_factors_imp_prime intro!: prod.cong) finally show "gcd x y = ?rhs1" . have "lcm x y = lcm_factorial x y" by (rule lcm_eq_lcm_factorial) also have "\ = ?rhs2" by (auto simp: lcm_factorial_def assms prod_mset_multiplicity - count_prime_factorization_prime dest: in_prime_factors_imp_prime intro!: prod.cong) + count_prime_factorization_prime intro!: arg_cong[of _ _ normalize] + dest: in_prime_factors_imp_prime intro!: prod.cong) finally show "lcm x y = ?rhs2" . qed lemma assumes "x \ 0" "y \ 0" "prime p" shows multiplicity_gcd: "multiplicity p (gcd x y) = min (multiplicity p x) (multiplicity p y)" and multiplicity_lcm: "multiplicity p (lcm x y) = max (multiplicity p x) (multiplicity p y)" proof - have "gcd x y = gcd_factorial x y" by (rule gcd_eq_gcd_factorial) also from assms have "multiplicity p \ = min (multiplicity p x) (multiplicity p y)" by (simp add: count_prime_factorization_prime [symmetric] prime_factorization_gcd_factorial) finally show "multiplicity p (gcd x y) = min (multiplicity p x) (multiplicity p y)" . have "lcm x y = lcm_factorial x y" by (rule lcm_eq_lcm_factorial) also from assms have "multiplicity p \ = max (multiplicity p x) (multiplicity p y)" by (simp add: count_prime_factorization_prime [symmetric] prime_factorization_lcm_factorial) finally show "multiplicity p (lcm x y) = max (multiplicity p x) (multiplicity p y)" . qed lemma gcd_lcm_distrib: "gcd x (lcm y z) = lcm (gcd x y) (gcd x z)" proof (cases "x = 0 \ y = 0 \ z = 0") case True thus ?thesis by (auto simp: lcm_proj1_if_dvd lcm_proj2_if_dvd) next case False hence "normalize (gcd x (lcm y z)) = normalize (lcm (gcd x y) (gcd x z))" by (intro associatedI prime_factorization_subset_imp_dvd) (auto simp: lcm_eq_0_iff prime_factorization_gcd prime_factorization_lcm subset_mset.inf_sup_distrib1) thus ?thesis by simp qed lemma lcm_gcd_distrib: "lcm x (gcd y z) = gcd (lcm x y) (lcm x z)" proof (cases "x = 0 \ y = 0 \ z = 0") case True thus ?thesis by (auto simp: lcm_proj1_if_dvd lcm_proj2_if_dvd) next case False hence "normalize (lcm x (gcd y z)) = normalize (gcd (lcm x y) (lcm x z))" by (intro associatedI prime_factorization_subset_imp_dvd) (auto simp: lcm_eq_0_iff prime_factorization_gcd prime_factorization_lcm subset_mset.sup_inf_distrib1) thus ?thesis by simp qed end class factorial_ring_gcd = factorial_semiring_gcd + idom begin subclass ring_gcd .. subclass idom_divide .. end + +class factorial_semiring_multiplicative = + factorial_semiring + normalization_semidom_multiplicative +begin + +lemma normalize_prod_mset_primes: + "(\p. p \# A \ prime p) \ normalize (prod_mset A) = prod_mset A" +proof (induction A) + case (add p A) + hence "prime p" by simp + hence "normalize p = p" by simp + with add show ?case by (simp add: normalize_mult) +qed simp_all + +lemma prod_mset_prime_factorization: + assumes "x \ 0" + shows "prod_mset (prime_factorization x) = normalize x" + using assms + by (induction x rule: prime_divisors_induct) + (simp_all add: prime_factorization_unit prime_factorization_times_prime + is_unit_normalize normalize_mult) + +lemma prime_decomposition: "unit_factor x * prod_mset (prime_factorization x) = x" + by (cases "x = 0") (simp_all add: prod_mset_prime_factorization) + +lemma prod_prime_factors: + assumes "x \ 0" + shows "(\p \ prime_factors x. p ^ multiplicity p x) = normalize x" +proof - + have "normalize x = prod_mset (prime_factorization x)" + by (simp add: prod_mset_prime_factorization assms) + also have "\ = (\p \ prime_factors x. p ^ count (prime_factorization x) p)" + by (subst prod_mset_multiplicity) simp_all + also have "\ = (\p \ prime_factors x. p ^ multiplicity p x)" + by (intro prod.cong) + (simp_all add: assms count_prime_factorization_prime in_prime_factors_imp_prime) + finally show ?thesis .. +qed + +lemma prime_factorization_unique'': + assumes S_eq: "S = {p. 0 < f p}" + and "finite S" + and S: "\p\S. prime p" "normalize n = (\p\S. p ^ f p)" + shows "S = prime_factors n \ (\p. prime p \ f p = multiplicity p n)" +proof + define A where "A = Abs_multiset f" + from \finite S\ S(1) have "(\p\S. p ^ f p) \ 0" by auto + with S(2) have nz: "n \ 0" by auto + from S_eq \finite S\ have count_A: "count A = f" + unfolding A_def by (subst multiset.Abs_multiset_inverse) (simp_all add: multiset_def) + from S_eq count_A have set_mset_A: "set_mset A = S" + by (simp only: set_mset_def) + from S(2) have "normalize n = (\p\S. p ^ f p)" . + also have "\ = prod_mset A" by (simp add: prod_mset_multiplicity S_eq set_mset_A count_A) + also from nz have "normalize n = prod_mset (prime_factorization n)" + by (simp add: prod_mset_prime_factorization) + finally have "prime_factorization (prod_mset A) = + prime_factorization (prod_mset (prime_factorization n))" by simp + also from S(1) have "prime_factorization (prod_mset A) = A" + by (intro prime_factorization_prod_mset_primes) (auto simp: set_mset_A) + also have "prime_factorization (prod_mset (prime_factorization n)) = prime_factorization n" + by (intro prime_factorization_prod_mset_primes) auto + finally show "S = prime_factors n" by (simp add: set_mset_A [symmetric]) + + show "(\p. prime p \ f p = multiplicity p n)" + proof safe + fix p :: 'a assume p: "prime p" + have "multiplicity p n = multiplicity p (normalize n)" by simp + also have "normalize n = prod_mset A" + by (simp add: prod_mset_multiplicity S_eq set_mset_A count_A S) + also from p set_mset_A S(1) + have "multiplicity p \ = sum_mset (image_mset (multiplicity p) A)" + by (intro prime_elem_multiplicity_prod_mset_distrib) auto + also from S(1) p + have "image_mset (multiplicity p) A = image_mset (\q. if p = q then 1 else 0) A" + by (intro image_mset_cong) (auto simp: set_mset_A multiplicity_self prime_multiplicity_other) + also have "sum_mset \ = f p" + by (simp add: semiring_1_class.sum_mset_delta' count_A) + finally show "f p = multiplicity p n" .. + qed +qed + +lemma divides_primepow: + assumes "prime p" and "a dvd p ^ n" + obtains m where "m \ n" and "normalize a = p ^ m" + using divides_primepow_weak[OF assms] that assms + by (auto simp add: normalize_power) + +lemma Ex_other_prime_factor: + assumes "n \ 0" and "\(\k. normalize n = p ^ k)" "prime p" + shows "\q\prime_factors n. q \ p" +proof (rule ccontr) + assume *: "\(\q\prime_factors n. q \ p)" + have "normalize n = (\p\prime_factors n. p ^ multiplicity p n)" + using assms(1) by (intro prod_prime_factors [symmetric]) auto + also from * have "\ = (\p\{p}. p ^ multiplicity p n)" + using assms(3) by (intro prod.mono_neutral_left) (auto simp: prime_factors_multiplicity) + finally have "normalize n = p ^ multiplicity p n" by auto + with assms show False by auto +qed + end + +end diff --git a/src/HOL/Computational_Algebra/Field_as_Ring.thy b/src/HOL/Computational_Algebra/Field_as_Ring.thy --- a/src/HOL/Computational_Algebra/Field_as_Ring.thy +++ b/src/HOL/Computational_Algebra/Field_as_Ring.thy @@ -1,120 +1,131 @@ (* Title: HOL/Computational_Algebra/Field_as_Ring.thy Author: Manuel Eberl *) theory Field_as_Ring imports Complex_Main Euclidean_Algorithm begin context field begin subclass idom_divide .. definition normalize_field :: "'a \ 'a" where [simp]: "normalize_field x = (if x = 0 then 0 else 1)" definition unit_factor_field :: "'a \ 'a" where [simp]: "unit_factor_field x = x" definition euclidean_size_field :: "'a \ nat" where [simp]: "euclidean_size_field x = (if x = 0 then 0 else 1)" definition mod_field :: "'a \ 'a \ 'a" where [simp]: "mod_field x y = (if y = 0 then x else 0)" end -instantiation real :: "{unique_euclidean_ring, normalization_euclidean_semiring}" +instantiation real :: + "{unique_euclidean_ring, normalization_euclidean_semiring, normalization_semidom_multiplicative}" begin definition [simp]: "normalize_real = (normalize_field :: real \ _)" definition [simp]: "unit_factor_real = (unit_factor_field :: real \ _)" definition [simp]: "modulo_real = (mod_field :: real \ _)" definition [simp]: "euclidean_size_real = (euclidean_size_field :: real \ _)" definition [simp]: "division_segment (x :: real) = 1" instance by standard (simp_all add: dvd_field_iff field_split_simps split: if_splits) end instantiation real :: euclidean_ring_gcd begin definition gcd_real :: "real \ real \ real" where "gcd_real = Euclidean_Algorithm.gcd" definition lcm_real :: "real \ real \ real" where "lcm_real = Euclidean_Algorithm.lcm" definition Gcd_real :: "real set \ real" where "Gcd_real = Euclidean_Algorithm.Gcd" definition Lcm_real :: "real set \ real" where "Lcm_real = Euclidean_Algorithm.Lcm" instance by standard (simp_all add: gcd_real_def lcm_real_def Gcd_real_def Lcm_real_def) end -instantiation rat :: "{unique_euclidean_ring, normalization_euclidean_semiring}" +instance real :: field_gcd .. + + +instantiation rat :: + "{unique_euclidean_ring, normalization_euclidean_semiring, normalization_semidom_multiplicative}" begin definition [simp]: "normalize_rat = (normalize_field :: rat \ _)" definition [simp]: "unit_factor_rat = (unit_factor_field :: rat \ _)" definition [simp]: "modulo_rat = (mod_field :: rat \ _)" definition [simp]: "euclidean_size_rat = (euclidean_size_field :: rat \ _)" definition [simp]: "division_segment (x :: rat) = 1" instance by standard (simp_all add: dvd_field_iff field_split_simps split: if_splits) end instantiation rat :: euclidean_ring_gcd begin definition gcd_rat :: "rat \ rat \ rat" where "gcd_rat = Euclidean_Algorithm.gcd" definition lcm_rat :: "rat \ rat \ rat" where "lcm_rat = Euclidean_Algorithm.lcm" definition Gcd_rat :: "rat set \ rat" where "Gcd_rat = Euclidean_Algorithm.Gcd" definition Lcm_rat :: "rat set \ rat" where "Lcm_rat = Euclidean_Algorithm.Lcm" instance by standard (simp_all add: gcd_rat_def lcm_rat_def Gcd_rat_def Lcm_rat_def) end -instantiation complex :: "{unique_euclidean_ring, normalization_euclidean_semiring}" +instance rat :: field_gcd .. + + +instantiation complex :: + "{unique_euclidean_ring, normalization_euclidean_semiring, normalization_semidom_multiplicative}" begin definition [simp]: "normalize_complex = (normalize_field :: complex \ _)" definition [simp]: "unit_factor_complex = (unit_factor_field :: complex \ _)" definition [simp]: "modulo_complex = (mod_field :: complex \ _)" definition [simp]: "euclidean_size_complex = (euclidean_size_field :: complex \ _)" definition [simp]: "division_segment (x :: complex) = 1" instance by standard (simp_all add: dvd_field_iff field_split_simps split: if_splits) end instantiation complex :: euclidean_ring_gcd begin definition gcd_complex :: "complex \ complex \ complex" where "gcd_complex = Euclidean_Algorithm.gcd" definition lcm_complex :: "complex \ complex \ complex" where "lcm_complex = Euclidean_Algorithm.lcm" definition Gcd_complex :: "complex set \ complex" where "Gcd_complex = Euclidean_Algorithm.Gcd" definition Lcm_complex :: "complex set \ complex" where "Lcm_complex = Euclidean_Algorithm.Lcm" instance by standard (simp_all add: gcd_complex_def lcm_complex_def Gcd_complex_def Lcm_complex_def) end +instance complex :: field_gcd .. + end diff --git a/src/HOL/Computational_Algebra/Formal_Power_Series.thy b/src/HOL/Computational_Algebra/Formal_Power_Series.thy --- a/src/HOL/Computational_Algebra/Formal_Power_Series.thy +++ b/src/HOL/Computational_Algebra/Formal_Power_Series.thy @@ -1,6458 +1,6465 @@ (* Title: HOL/Computational_Algebra/Formal_Power_Series.thy Author: Amine Chaieb, University of Cambridge Author: Jeremy Sylvestre, University of Alberta (Augustana Campus) Author: Manuel Eberl, TU München *) section \A formalization of formal power series\ theory Formal_Power_Series imports Complex_Main Euclidean_Algorithm begin subsection \The type of formal power series\ typedef 'a fps = "{f :: nat \ 'a. True}" morphisms fps_nth Abs_fps by simp notation fps_nth (infixl "$" 75) lemma expand_fps_eq: "p = q \ (\n. p $ n = q $ n)" by (simp add: fps_nth_inject [symmetric] fun_eq_iff) lemmas fps_eq_iff = expand_fps_eq lemma fps_ext: "(\n. p $ n = q $ n) \ p = q" by (simp add: expand_fps_eq) lemma fps_nth_Abs_fps [simp]: "Abs_fps f $ n = f n" by (simp add: Abs_fps_inverse) text \Definition of the basic elements 0 and 1 and the basic operations of addition, negation and multiplication.\ instantiation fps :: (zero) zero begin definition fps_zero_def: "0 = Abs_fps (\n. 0)" instance .. end lemma fps_zero_nth [simp]: "0 $ n = 0" unfolding fps_zero_def by simp lemma fps_nonzero_nth: "f \ 0 \ (\ n. f $n \ 0)" by (simp add: expand_fps_eq) lemma fps_nonzero_nth_minimal: "f \ 0 \ (\n. f $ n \ 0 \ (\m < n. f $ m = 0))" (is "?lhs \ ?rhs") proof let ?n = "LEAST n. f $ n \ 0" show ?rhs if ?lhs proof - from that have "\n. f $ n \ 0" by (simp add: fps_nonzero_nth) then have "f $ ?n \ 0" by (rule LeastI_ex) moreover have "\m 0 \ (\m 0 \ f \ 0" by auto instantiation fps :: ("{one, zero}") one begin definition fps_one_def: "1 = Abs_fps (\n. if n = 0 then 1 else 0)" instance .. end lemma fps_one_nth [simp]: "1 $ n = (if n = 0 then 1 else 0)" unfolding fps_one_def by simp instantiation fps :: (plus) plus begin definition fps_plus_def: "(+) = (\f g. Abs_fps (\n. f $ n + g $ n))" instance .. end lemma fps_add_nth [simp]: "(f + g) $ n = f $ n + g $ n" unfolding fps_plus_def by simp instantiation fps :: (minus) minus begin definition fps_minus_def: "(-) = (\f g. Abs_fps (\n. f $ n - g $ n))" instance .. end lemma fps_sub_nth [simp]: "(f - g) $ n = f $ n - g $ n" unfolding fps_minus_def by simp instantiation fps :: (uminus) uminus begin definition fps_uminus_def: "uminus = (\f. Abs_fps (\n. - (f $ n)))" instance .. end lemma fps_neg_nth [simp]: "(- f) $ n = - (f $ n)" unfolding fps_uminus_def by simp lemma fps_neg_0 [simp]: "-(0::'a::group_add fps) = 0" by (rule iffD2, rule fps_eq_iff, auto) instantiation fps :: ("{comm_monoid_add, times}") times begin definition fps_times_def: "(*) = (\f g. Abs_fps (\n. \i=0..n. f $ i * g $ (n - i)))" instance .. end lemma fps_mult_nth: "(f * g) $ n = (\i=0..n. f$i * g$(n - i))" unfolding fps_times_def by simp lemma fps_mult_nth_0 [simp]: "(f * g) $ 0 = f $ 0 * g $ 0" unfolding fps_times_def by simp lemma fps_mult_nth_1 [simp]: "(f * g) $ 1 = f$0 * g$1 + f$1 * g$0" by (simp add: fps_mult_nth) lemmas mult_nth_0 = fps_mult_nth_0 lemmas mult_nth_1 = fps_mult_nth_1 instance fps :: ("{comm_monoid_add, mult_zero}") mult_zero proof fix a :: "'a fps" show "0 * a = 0" by (simp add: fps_ext fps_mult_nth) show "a * 0 = 0" by (simp add: fps_ext fps_mult_nth) qed declare atLeastAtMost_iff [presburger] declare Bex_def [presburger] declare Ball_def [presburger] lemma mult_delta_left: fixes x y :: "'a::mult_zero" shows "(if b then x else 0) * y = (if b then x * y else 0)" by simp lemma mult_delta_right: fixes x y :: "'a::mult_zero" shows "x * (if b then y else 0) = (if b then x * y else 0)" by simp lemma fps_one_mult: fixes f :: "'a::{comm_monoid_add, mult_zero, monoid_mult} fps" shows "1 * f = f" and "f * 1 = f" by (simp_all add: fps_ext fps_mult_nth mult_delta_left mult_delta_right) subsection \Subdegrees\ definition subdegree :: "('a::zero) fps \ nat" where "subdegree f = (if f = 0 then 0 else LEAST n. f$n \ 0)" lemma subdegreeI: assumes "f $ d \ 0" and "\i. i < d \ f $ i = 0" shows "subdegree f = d" proof- from assms(1) have "f \ 0" by auto moreover from assms(1) have "(LEAST i. f $ i \ 0) = d" proof (rule Least_equality) fix e assume "f $ e \ 0" with assms(2) have "\(e < d)" by blast thus "e \ d" by simp qed ultimately show ?thesis unfolding subdegree_def by simp qed lemma nth_subdegree_nonzero [simp,intro]: "f \ 0 \ f $ subdegree f \ 0" proof- assume "f \ 0" hence "subdegree f = (LEAST n. f $ n \ 0)" by (simp add: subdegree_def) also from \f \ 0\ have "\n. f$n \ 0" using fps_nonzero_nth by blast from LeastI_ex[OF this] have "f $ (LEAST n. f $ n \ 0) \ 0" . finally show ?thesis . qed lemma nth_less_subdegree_zero [dest]: "n < subdegree f \ f $ n = 0" proof (cases "f = 0") assume "f \ 0" and less: "n < subdegree f" note less also from \f \ 0\ have "subdegree f = (LEAST n. f $ n \ 0)" by (simp add: subdegree_def) finally show "f $ n = 0" using not_less_Least by blast qed simp_all lemma subdegree_geI: assumes "f \ 0" "\i. i < n \ f$i = 0" shows "subdegree f \ n" proof (rule ccontr) assume "\(subdegree f \ n)" with assms(2) have "f $ subdegree f = 0" by simp moreover from assms(1) have "f $ subdegree f \ 0" by simp ultimately show False by contradiction qed lemma subdegree_greaterI: assumes "f \ 0" "\i. i \ n \ f$i = 0" shows "subdegree f > n" proof (rule ccontr) assume "\(subdegree f > n)" with assms(2) have "f $ subdegree f = 0" by simp moreover from assms(1) have "f $ subdegree f \ 0" by simp ultimately show False by contradiction qed lemma subdegree_leI: "f $ n \ 0 \ subdegree f \ n" by (rule leI) auto lemma subdegree_0 [simp]: "subdegree 0 = 0" by (simp add: subdegree_def) lemma subdegree_1 [simp]: "subdegree 1 = 0" by (cases "(1::'a) = 0") (auto intro: subdegreeI fps_ext simp: subdegree_def) lemma subdegree_eq_0_iff: "subdegree f = 0 \ f = 0 \ f $ 0 \ 0" proof (cases "f = 0") assume "f \ 0" thus ?thesis using nth_subdegree_nonzero[OF \f \ 0\] by (fastforce intro!: subdegreeI) qed simp_all lemma subdegree_eq_0 [simp]: "f $ 0 \ 0 \ subdegree f = 0" by (simp add: subdegree_eq_0_iff) lemma nth_subdegree_zero_iff [simp]: "f $ subdegree f = 0 \ f = 0" by (cases "f = 0") auto lemma fps_nonzero_subdegree_nonzeroI: "subdegree f > 0 \ f \ 0" by auto lemma subdegree_uminus [simp]: "subdegree (-(f::('a::group_add) fps)) = subdegree f" proof (cases "f=0") case False thus ?thesis by (force intro: subdegreeI) qed simp lemma subdegree_minus_commute [simp]: "subdegree (f-(g::('a::group_add) fps)) = subdegree (g - f)" proof (-, cases "g-f=0") case True have "\n. (f - g) $ n = -((g - f) $ n)" by simp with True have "f - g = 0" by (intro fps_ext) simp with True show ?thesis by simp next case False show ?thesis using nth_subdegree_nonzero[OF False] by (fastforce intro: subdegreeI) qed lemma subdegree_add_ge': fixes f g :: "'a::monoid_add fps" assumes "f + g \ 0" shows "subdegree (f + g) \ min (subdegree f) (subdegree g)" using assms by (force intro: subdegree_geI) lemma subdegree_add_ge: assumes "f \ -(g :: ('a :: group_add) fps)" shows "subdegree (f + g) \ min (subdegree f) (subdegree g)" proof (rule subdegree_add_ge') have "f + g = 0 \ False" proof- assume fg: "f + g = 0" have "\n. f $ n = - g $ n" proof- fix n from fg have "(f + g) $ n = 0" by simp hence "f $ n + g $ n - g $ n = - g $ n" by simp thus "f $ n = - g $ n" by simp qed with assms show False by (auto intro: fps_ext) qed thus "f + g \ 0" by fast qed lemma subdegree_add_eq1: assumes "f \ 0" and "subdegree f < subdegree (g :: 'a::monoid_add fps)" shows "subdegree (f + g) = subdegree f" using assms by (auto intro: subdegreeI simp: nth_less_subdegree_zero) lemma subdegree_add_eq2: assumes "g \ 0" and "subdegree g < subdegree (f :: 'a :: monoid_add fps)" shows "subdegree (f + g) = subdegree g" using assms by (auto intro: subdegreeI simp: nth_less_subdegree_zero) lemma subdegree_diff_eq1: assumes "f \ 0" and "subdegree f < subdegree (g :: 'a :: group_add fps)" shows "subdegree (f - g) = subdegree f" using assms by (auto intro: subdegreeI simp: nth_less_subdegree_zero) lemma subdegree_diff_eq1_cancel: assumes "f \ 0" and "subdegree f < subdegree (g :: 'a :: cancel_comm_monoid_add fps)" shows "subdegree (f - g) = subdegree f" using assms by (auto intro: subdegreeI simp: nth_less_subdegree_zero) lemma subdegree_diff_eq2: assumes "g \ 0" and "subdegree g < subdegree (f :: 'a :: group_add fps)" shows "subdegree (f - g) = subdegree g" using assms by (auto intro: subdegreeI simp: nth_less_subdegree_zero) lemma subdegree_diff_ge [simp]: assumes "f \ (g :: 'a :: group_add fps)" shows "subdegree (f - g) \ min (subdegree f) (subdegree g)" proof- from assms have "f = - (- g) \ False" using expand_fps_eq by fastforce hence "f \ - (- g)" by fast moreover have "f + - g = f - g" by (simp add: fps_ext) ultimately show ?thesis using subdegree_add_ge[of f "-g"] by simp qed lemma subdegree_diff_ge': fixes f g :: "'a :: comm_monoid_diff fps" assumes "f - g \ 0" shows "subdegree (f - g) \ subdegree f" using assms by (auto intro: subdegree_geI simp: nth_less_subdegree_zero) lemma nth_subdegree_mult_left [simp]: fixes f g :: "('a :: {mult_zero,comm_monoid_add}) fps" shows "(f * g) $ (subdegree f) = f $ subdegree f * g $ 0" by (cases "subdegree f") (simp_all add: fps_mult_nth nth_less_subdegree_zero) lemma nth_subdegree_mult_right [simp]: fixes f g :: "('a :: {mult_zero,comm_monoid_add}) fps" shows "(f * g) $ (subdegree g) = f $ 0 * g $ subdegree g" by (cases "subdegree g") (simp_all add: fps_mult_nth nth_less_subdegree_zero sum.atLeast_Suc_atMost) lemma nth_subdegree_mult [simp]: fixes f g :: "('a :: {mult_zero,comm_monoid_add}) fps" shows "(f * g) $ (subdegree f + subdegree g) = f $ subdegree f * g $ subdegree g" proof- let ?n = "subdegree f + subdegree g" have "(f * g) $ ?n = (\i=0..?n. f$i * g$(?n-i))" by (simp add: fps_mult_nth) also have "... = (\i=0..?n. if i = subdegree f then f$i * g$(?n-i) else 0)" proof (intro sum.cong) fix x assume x: "x \ {0..?n}" hence "x = subdegree f \ x < subdegree f \ ?n - x < subdegree g" by auto thus "f $ x * g $ (?n - x) = (if x = subdegree f then f $ x * g $ (?n - x) else 0)" by (elim disjE conjE) auto qed auto also have "... = f $ subdegree f * g $ subdegree g" by simp finally show ?thesis . qed lemma fps_mult_nth_eq0: fixes f g :: "'a::{comm_monoid_add,mult_zero} fps" assumes "n < subdegree f + subdegree g" shows "(f*g) $ n = 0" proof- have "\i. i\{0..n} \ f$i * g$(n - i) = 0" proof- fix i assume i: "i\{0..n}" show "f$i * g$(n - i) = 0" proof (cases "i < subdegree f \ n - i < subdegree g") case False with assms i show ?thesis by auto qed (auto simp: nth_less_subdegree_zero) qed thus "(f * g) $ n = 0" by (simp add: fps_mult_nth) qed lemma fps_mult_subdegree_ge: fixes f g :: "'a::{comm_monoid_add,mult_zero} fps" assumes "f*g \ 0" shows "subdegree (f*g) \ subdegree f + subdegree g" using assms fps_mult_nth_eq0 by (intro subdegree_geI) simp lemma subdegree_mult': fixes f g :: "'a::{comm_monoid_add,mult_zero} fps" assumes "f $ subdegree f * g $ subdegree g \ 0" shows "subdegree (f*g) = subdegree f + subdegree g" proof- from assms have "(f * g) $ (subdegree f + subdegree g) \ 0" by simp hence "f*g \ 0" by fastforce hence "subdegree (f*g) \ subdegree f + subdegree g" using fps_mult_subdegree_ge by fast moreover from assms have "subdegree (f*g) \ subdegree f + subdegree g" by (intro subdegree_leI) simp ultimately show ?thesis by simp qed lemma subdegree_mult [simp]: fixes f g :: "'a :: {semiring_no_zero_divisors} fps" assumes "f \ 0" "g \ 0" shows "subdegree (f * g) = subdegree f + subdegree g" using assms by (intro subdegree_mult') simp lemma fps_mult_nth_conv_upto_subdegree_left: fixes f g :: "('a :: {mult_zero,comm_monoid_add}) fps" shows "(f * g) $ n = (\i=subdegree f..n. f $ i * g $ (n - i))" proof (cases "subdegree f \ n") case True hence "{0..n} = {0.. {subdegree f..n}" by auto moreover have "{0.. {subdegree f..n} = {}" by auto ultimately show ?thesis using nth_less_subdegree_zero[of _ f] by (simp add: fps_mult_nth sum.union_disjoint) qed (simp add: fps_mult_nth nth_less_subdegree_zero) lemma fps_mult_nth_conv_upto_subdegree_right: fixes f g :: "('a :: {mult_zero,comm_monoid_add}) fps" shows "(f * g) $ n = (\i=0..n - subdegree g. f $ i * g $ (n - i))" proof- have "{0..n} = {0..n - subdegree g} \ {n - subdegree g<..n}" by auto moreover have "{0..n - subdegree g} \ {n - subdegree g<..n} = {}" by auto moreover have "\i\{n - subdegree g<..n}. g $ (n - i) = 0" using nth_less_subdegree_zero[of _ g] by auto ultimately show ?thesis by (simp add: fps_mult_nth sum.union_disjoint) qed lemma fps_mult_nth_conv_inside_subdegrees: fixes f g :: "('a :: {mult_zero,comm_monoid_add}) fps" shows "(f * g) $ n = (\i=subdegree f..n - subdegree g. f $ i * g $ (n - i))" proof (cases "subdegree f \ n - subdegree g") case True hence "{subdegree f..n} = {subdegree f..n - subdegree g} \ {n - subdegree g<..n}" by auto moreover have "{subdegree f..n - subdegree g} \ {n - subdegree g<..n} = {}" by auto moreover have "\i\{n - subdegree g<..n}. f $ i * g $ (n - i) = 0" using nth_less_subdegree_zero[of _ g] by auto ultimately show ?thesis using fps_mult_nth_conv_upto_subdegree_left[of f g n] by (simp add: sum.union_disjoint) next case False hence 1: "subdegree f > n - subdegree g" by simp show ?thesis proof (cases "f*g = 0") case False with 1 have "n < subdegree (f*g)" using fps_mult_subdegree_ge[of f g] by simp with 1 show ?thesis by auto qed (simp add: 1) qed lemma fps_mult_nth_outside_subdegrees: fixes f g :: "('a :: {mult_zero,comm_monoid_add}) fps" shows "n < subdegree f \ (f * g) $ n = 0" and "n < subdegree g \ (f * g) $ n = 0" by (auto simp: fps_mult_nth_conv_inside_subdegrees) subsection \Formal power series form a commutative ring with unity, if the range of sequences they represent is a commutative ring with unity\ instance fps :: (semigroup_add) semigroup_add proof fix a b c :: "'a fps" show "a + b + c = a + (b + c)" by (simp add: fps_ext add.assoc) qed instance fps :: (ab_semigroup_add) ab_semigroup_add proof fix a b :: "'a fps" show "a + b = b + a" by (simp add: fps_ext add.commute) qed instance fps :: (monoid_add) monoid_add proof fix a :: "'a fps" show "0 + a = a" by (simp add: fps_ext) show "a + 0 = a" by (simp add: fps_ext) qed instance fps :: (comm_monoid_add) comm_monoid_add proof fix a :: "'a fps" show "0 + a = a" by (simp add: fps_ext) qed instance fps :: (cancel_semigroup_add) cancel_semigroup_add proof fix a b c :: "'a fps" show "b = c" if "a + b = a + c" using that by (simp add: expand_fps_eq) show "b = c" if "b + a = c + a" using that by (simp add: expand_fps_eq) qed instance fps :: (cancel_ab_semigroup_add) cancel_ab_semigroup_add proof fix a b c :: "'a fps" show "a + b - a = b" by (simp add: expand_fps_eq) show "a - b - c = a - (b + c)" by (simp add: expand_fps_eq diff_diff_eq) qed instance fps :: (cancel_comm_monoid_add) cancel_comm_monoid_add .. instance fps :: (group_add) group_add proof fix a b :: "'a fps" show "- a + a = 0" by (simp add: fps_ext) show "a + - b = a - b" by (simp add: fps_ext) qed instance fps :: (ab_group_add) ab_group_add proof fix a b :: "'a fps" show "- a + a = 0" by (simp add: fps_ext) show "a - b = a + - b" by (simp add: fps_ext) qed instance fps :: (zero_neq_one) zero_neq_one by standard (simp add: expand_fps_eq) lemma fps_mult_assoc_lemma: fixes k :: nat and f :: "nat \ nat \ nat \ 'a::comm_monoid_add" shows "(\j=0..k. \i=0..j. f i (j - i) (n - j)) = (\j=0..k. \i=0..k - j. f j i (n - j - i))" by (induct k) (simp_all add: Suc_diff_le sum.distrib add.assoc) instance fps :: (semiring_0) semiring_0 proof fix a b c :: "'a fps" show "(a + b) * c = a * c + b * c" by (simp add: expand_fps_eq fps_mult_nth distrib_right sum.distrib) show "a * (b + c) = a * b + a * c" by (simp add: expand_fps_eq fps_mult_nth distrib_left sum.distrib) show "(a * b) * c = a * (b * c)" proof (rule fps_ext) fix n :: nat have "(\j=0..n. \i=0..j. a$i * b$(j - i) * c$(n - j)) = (\j=0..n. \i=0..n - j. a$j * b$i * c$(n - j - i))" by (rule fps_mult_assoc_lemma) then show "((a * b) * c) $ n = (a * (b * c)) $ n" by (simp add: fps_mult_nth sum_distrib_left sum_distrib_right mult.assoc) qed qed instance fps :: (semiring_0_cancel) semiring_0_cancel .. lemma fps_mult_commute_lemma: fixes n :: nat and f :: "nat \ nat \ 'a::comm_monoid_add" shows "(\i=0..n. f i (n - i)) = (\i=0..n. f (n - i) i)" by (rule sum.reindex_bij_witness[where i="(-) n" and j="(-) n"]) auto instance fps :: (comm_semiring_0) comm_semiring_0 proof fix a b c :: "'a fps" show "a * b = b * a" proof (rule fps_ext) fix n :: nat have "(\i=0..n. a$i * b$(n - i)) = (\i=0..n. a$(n - i) * b$i)" by (rule fps_mult_commute_lemma) then show "(a * b) $ n = (b * a) $ n" by (simp add: fps_mult_nth mult.commute) qed qed (simp add: distrib_right) instance fps :: (comm_semiring_0_cancel) comm_semiring_0_cancel .. instance fps :: (semiring_1) semiring_1 proof fix a :: "'a fps" show "1 * a = a" "a * 1 = a" by (simp_all add: fps_one_mult) qed instance fps :: (comm_semiring_1) comm_semiring_1 by standard simp instance fps :: (semiring_1_cancel) semiring_1_cancel .. subsection \Selection of the nth power of the implicit variable in the infinite sum\ lemma fps_square_nth: "(f^2) $ n = (\k\n. f $ k * f $ (n - k))" by (simp add: power2_eq_square fps_mult_nth atLeast0AtMost) lemma fps_sum_nth: "sum f S $ n = sum (\k. (f k) $ n) S" proof (cases "finite S") case True then show ?thesis by (induct set: finite) auto next case False then show ?thesis by simp qed subsection \Injection of the basic ring elements and multiplication by scalars\ definition "fps_const c = Abs_fps (\n. if n = 0 then c else 0)" lemma fps_nth_fps_const [simp]: "fps_const c $ n = (if n = 0 then c else 0)" unfolding fps_const_def by simp lemma fps_const_0_eq_0 [simp]: "fps_const 0 = 0" by (simp add: fps_ext) lemma fps_const_nonzero_eq_nonzero: "c \ 0 \ fps_const c \ 0" using fps_nonzeroI[of "fps_const c" 0] by simp lemma fps_const_1_eq_1 [simp]: "fps_const 1 = 1" by (simp add: fps_ext) lemma subdegree_fps_const [simp]: "subdegree (fps_const c) = 0" by (cases "c = 0") (auto intro!: subdegreeI) lemma fps_const_neg [simp]: "- (fps_const (c::'a::group_add)) = fps_const (- c)" by (simp add: fps_ext) lemma fps_const_add [simp]: "fps_const (c::'a::monoid_add) + fps_const d = fps_const (c + d)" by (simp add: fps_ext) lemma fps_const_add_left: "fps_const (c::'a::monoid_add) + f = Abs_fps (\n. if n = 0 then c + f$0 else f$n)" by (simp add: fps_ext) lemma fps_const_add_right: "f + fps_const (c::'a::monoid_add) = Abs_fps (\n. if n = 0 then f$0 + c else f$n)" by (simp add: fps_ext) lemma fps_const_sub [simp]: "fps_const (c::'a::group_add) - fps_const d = fps_const (c - d)" by (simp add: fps_ext) lemmas fps_const_minus = fps_const_sub lemma fps_const_mult[simp]: fixes c d :: "'a::{comm_monoid_add,mult_zero}" shows "fps_const c * fps_const d = fps_const (c * d)" by (simp add: fps_eq_iff fps_mult_nth sum.neutral) lemma fps_const_mult_left: "fps_const (c::'a::{comm_monoid_add,mult_zero}) * f = Abs_fps (\n. c * f$n)" unfolding fps_eq_iff fps_mult_nth by (simp add: fps_const_def mult_delta_left) lemma fps_const_mult_right: "f * fps_const (c::'a::{comm_monoid_add,mult_zero}) = Abs_fps (\n. f$n * c)" unfolding fps_eq_iff fps_mult_nth by (simp add: fps_const_def mult_delta_right) lemma fps_mult_left_const_nth [simp]: "(fps_const (c::'a::{comm_monoid_add,mult_zero}) * f)$n = c* f$n" by (simp add: fps_mult_nth mult_delta_left) lemma fps_mult_right_const_nth [simp]: "(f * fps_const (c::'a::{comm_monoid_add,mult_zero}))$n = f$n * c" by (simp add: fps_mult_nth mult_delta_right) lemma fps_const_power [simp]: "fps_const c ^ n = fps_const (c^n)" by (induct n) auto subsection \Formal power series form an integral domain\ instance fps :: (ring) ring .. instance fps :: (comm_ring) comm_ring .. instance fps :: (ring_1) ring_1 .. instance fps :: (comm_ring_1) comm_ring_1 .. instance fps :: (semiring_no_zero_divisors) semiring_no_zero_divisors proof fix a b :: "'a fps" assume "a \ 0" and "b \ 0" hence "(a * b) $ (subdegree a + subdegree b) \ 0" by simp thus "a * b \ 0" using fps_nonzero_nth by fast qed instance fps :: (semiring_1_no_zero_divisors) semiring_1_no_zero_divisors .. instance fps :: ("{cancel_semigroup_add,semiring_no_zero_divisors_cancel}") semiring_no_zero_divisors_cancel proof fix a b c :: "'a fps" show "(a * c = b * c) = (c = 0 \ a = b)" proof assume ab: "a * c = b * c" have "c \ 0 \ a = b" proof (rule fps_ext) fix n assume c: "c \ 0" show "a $ n = b $ n" proof (induct n rule: nat_less_induct) case (1 n) with ab c show ?case using fps_mult_nth_conv_upto_subdegree_right[of a c "subdegree c + n"] fps_mult_nth_conv_upto_subdegree_right[of b c "subdegree c + n"] by (cases n) auto qed qed thus "c = 0 \ a = b" by fast qed auto show "(c * a = c * b) = (c = 0 \ a = b)" proof assume ab: "c * a = c * b" have "c \ 0 \ a = b" proof (rule fps_ext) fix n assume c: "c \ 0" show "a $ n = b $ n" proof (induct n rule: nat_less_induct) case (1 n) moreover have "\i\{Suc (subdegree c)..subdegree c + n}. subdegree c + n - i < n" by auto ultimately show ?case using ab c fps_mult_nth_conv_upto_subdegree_left[of c a "subdegree c + n"] fps_mult_nth_conv_upto_subdegree_left[of c b "subdegree c + n"] by (simp add: sum.atLeast_Suc_atMost) qed qed thus "c = 0 \ a = b" by fast qed auto qed instance fps :: (ring_no_zero_divisors) ring_no_zero_divisors .. instance fps :: (ring_1_no_zero_divisors) ring_1_no_zero_divisors .. instance fps :: (idom) idom .. lemma fps_numeral_fps_const: "numeral k = fps_const (numeral k)" by (induct k) (simp_all only: numeral.simps fps_const_1_eq_1 fps_const_add [symmetric]) lemmas numeral_fps_const = fps_numeral_fps_const lemma neg_numeral_fps_const: "(- numeral k :: 'a :: ring_1 fps) = fps_const (- numeral k)" by (simp add: numeral_fps_const) lemma fps_numeral_nth: "numeral n $ i = (if i = 0 then numeral n else 0)" by (simp add: numeral_fps_const) lemma fps_numeral_nth_0 [simp]: "numeral n $ 0 = numeral n" by (simp add: numeral_fps_const) lemma subdegree_numeral [simp]: "subdegree (numeral n) = 0" by (simp add: numeral_fps_const) lemma fps_of_nat: "fps_const (of_nat c) = of_nat c" by (induction c) (simp_all add: fps_const_add [symmetric] del: fps_const_add) lemma fps_of_int: "fps_const (of_int c) = of_int c" by (induction c) (simp_all add: fps_const_minus [symmetric] fps_of_nat fps_const_neg [symmetric] del: fps_const_minus fps_const_neg) lemma fps_nth_of_nat [simp]: "(of_nat c) $ n = (if n=0 then of_nat c else 0)" by (simp add: fps_of_nat[symmetric]) lemma fps_nth_of_int [simp]: "(of_int c) $ n = (if n=0 then of_int c else 0)" by (simp add: fps_of_int[symmetric]) lemma fps_mult_of_nat_nth [simp]: shows "(of_nat k * f) $ n = of_nat k * f$n" and "(f * of_nat k ) $ n = f$n * of_nat k" by (simp_all add: fps_of_nat[symmetric]) lemma fps_mult_of_int_nth [simp]: shows "(of_int k * f) $ n = of_int k * f$n" and "(f * of_int k ) $ n = f$n * of_int k" by (simp_all add: fps_of_int[symmetric]) lemma numeral_neq_fps_zero [simp]: "(numeral f :: 'a :: field_char_0 fps) \ 0" proof assume "numeral f = (0 :: 'a fps)" from arg_cong[of _ _ "\F. F $ 0", OF this] show False by simp qed lemma subdegree_power_ge: "f^n \ 0 \ subdegree (f^n) \ n * subdegree f" proof (induct n) case (Suc n) thus ?case using fps_mult_subdegree_ge by fastforce qed simp lemma fps_pow_nth_below_subdegree: "k < n * subdegree f \ (f^n) $ k = 0" proof (cases "f^n = 0") case False assume "k < n * subdegree f" with False have "k < subdegree (f^n)" using subdegree_power_ge[of f n] by simp thus "(f^n) $ k = 0" by auto qed simp lemma fps_pow_base [simp]: "(f ^ n) $ (n * subdegree f) = (f $ subdegree f) ^ n" proof (induct n) case (Suc n) show ?case proof (cases "Suc n * subdegree f < subdegree f + subdegree (f^n)") case True with Suc show ?thesis by (auto simp: fps_mult_nth_eq0 distrib_right) next case False hence "\i\{Suc (subdegree f)..Suc n * subdegree f - subdegree (f ^ n)}. f ^ n $ (Suc n * subdegree f - i) = 0" by (auto simp: fps_pow_nth_below_subdegree) with False Suc show ?thesis using fps_mult_nth_conv_inside_subdegrees[of f "f^n" "Suc n * subdegree f"] sum.atLeast_Suc_atMost[of "subdegree f" "Suc n * subdegree f - subdegree (f ^ n)" "\i. f $ i * f ^ n $ (Suc n * subdegree f - i)" ] by simp qed qed simp lemma subdegree_power_eqI: fixes f :: "'a::semiring_1 fps" shows "(f $ subdegree f) ^ n \ 0 \ subdegree (f ^ n) = n * subdegree f" proof (induct n) case (Suc n) from Suc have 1: "subdegree (f ^ n) = n * subdegree f" by fastforce with Suc(2) have "f $ subdegree f * f ^ n $ subdegree (f ^ n) \ 0" by simp with 1 show ?case using subdegree_mult'[of f "f^n"] by simp qed simp lemma subdegree_power [simp]: "subdegree ((f :: ('a :: semiring_1_no_zero_divisors) fps) ^ n) = n * subdegree f" by (cases "f = 0"; induction n) simp_all subsection \The efps_Xtractor series fps_X\ lemma minus_one_power_iff: "(- (1::'a::ring_1)) ^ n = (if even n then 1 else - 1)" by (induct n) auto definition "fps_X = Abs_fps (\n. if n = 1 then 1 else 0)" lemma subdegree_fps_X [simp]: "subdegree (fps_X :: ('a :: zero_neq_one) fps) = 1" by (auto intro!: subdegreeI simp: fps_X_def) lemma fps_X_mult_nth [simp]: fixes f :: "'a::{comm_monoid_add,mult_zero,monoid_mult} fps" shows "(fps_X * f) $ n = (if n = 0 then 0 else f $ (n - 1))" proof (cases n) case (Suc m) moreover have "(fps_X * f) $ Suc m = f $ (Suc m - 1)" proof (cases m) case 0 thus ?thesis using fps_mult_nth_1[of "fps_X" f] by (simp add: fps_X_def) next case (Suc k) thus ?thesis by (simp add: fps_mult_nth fps_X_def sum.atLeast_Suc_atMost) qed ultimately show ?thesis by simp qed (simp add: fps_X_def) lemma fps_X_mult_right_nth [simp]: fixes a :: "'a::{comm_monoid_add,mult_zero,monoid_mult} fps" shows "(a * fps_X) $ n = (if n = 0 then 0 else a $ (n - 1))" proof (cases n) case (Suc m) moreover have "(a * fps_X) $ Suc m = a $ (Suc m - 1)" proof (cases m) case 0 thus ?thesis using fps_mult_nth_1[of a "fps_X"] by (simp add: fps_X_def) next case (Suc k) hence "(a * fps_X) $ Suc m = (\i=0..k. a$i * fps_X$(Suc m - i)) + a$(Suc k)" by (simp add: fps_mult_nth fps_X_def) moreover have "\i\{0..k}. a$i * fps_X$(Suc m - i) = 0" by (auto simp: Suc fps_X_def) ultimately show ?thesis by (simp add: Suc) qed ultimately show ?thesis by simp qed (simp add: fps_X_def) lemma fps_mult_fps_X_commute: fixes a :: "'a::{comm_monoid_add,mult_zero,monoid_mult} fps" shows "fps_X * a = a * fps_X" by (simp add: fps_eq_iff) lemma fps_mult_fps_X_power_commute: "fps_X ^ k * a = a * fps_X ^ k" proof (induct k) case (Suc k) hence "fps_X ^ Suc k * a = a * fps_X * fps_X ^ k" by (simp add: mult.assoc fps_mult_fps_X_commute[symmetric]) thus ?case by (simp add: mult.assoc) qed simp lemma fps_subdegree_mult_fps_X: fixes f :: "'a::{comm_monoid_add,mult_zero,monoid_mult} fps" assumes "f \ 0" shows "subdegree (fps_X * f) = subdegree f + 1" and "subdegree (f * fps_X) = subdegree f + 1" proof- show "subdegree (fps_X * f) = subdegree f + 1" proof (intro subdegreeI) fix i :: nat assume i: "i < subdegree f + 1" show "(fps_X * f) $ i = 0" proof (cases "i=0") case False with i show ?thesis by (simp add: nth_less_subdegree_zero) next case True thus ?thesis using fps_X_mult_nth[of f i] by simp qed qed (simp add: assms) thus "subdegree (f * fps_X) = subdegree f + 1" by (simp add: fps_mult_fps_X_commute) qed lemma fps_mult_fps_X_nonzero: fixes f :: "'a::{comm_monoid_add,mult_zero,monoid_mult} fps" assumes "f \ 0" shows "fps_X * f \ 0" and "f * fps_X \ 0" using assms fps_subdegree_mult_fps_X[of f] fps_nonzero_subdegree_nonzeroI[of "fps_X * f"] fps_nonzero_subdegree_nonzeroI[of "f * fps_X"] by auto lemma fps_mult_fps_X_power_nonzero: assumes "f \ 0" shows "fps_X ^ n * f \ 0" and "f * fps_X ^ n \ 0" proof - show "fps_X ^ n * f \ 0" by (induct n) (simp_all add: assms mult.assoc fps_mult_fps_X_nonzero(1)) thus "f * fps_X ^ n \ 0" by (simp add: fps_mult_fps_X_power_commute) qed lemma fps_X_power_iff: "fps_X ^ n = Abs_fps (\m. if m = n then 1 else 0)" by (induction n) (auto simp: fps_eq_iff) lemma fps_X_nth[simp]: "fps_X$n = (if n = 1 then 1 else 0)" by (simp add: fps_X_def) lemma fps_X_power_nth[simp]: "(fps_X^k) $n = (if n = k then 1 else 0)" by (simp add: fps_X_power_iff) lemma fps_X_power_subdegree: "subdegree (fps_X^n) = n" by (auto intro: subdegreeI) lemma fps_X_power_mult_nth: "(fps_X^k * f) $ n = (if n < k then 0 else f $ (n - k))" by (cases "n 0" shows "subdegree (fps_X ^ n * f) = subdegree f + n" and "subdegree (f * fps_X ^ n) = subdegree f + n" proof - from assms show "subdegree (fps_X ^ n * f) = subdegree f + n" by (induct n) (simp_all add: algebra_simps fps_subdegree_mult_fps_X(1) fps_mult_fps_X_power_nonzero(1)) thus "subdegree (f * fps_X ^ n) = subdegree f + n" by (simp add: fps_mult_fps_X_power_commute) qed lemma fps_mult_fps_X_plus_1_nth: "((1+fps_X)*a) $n = (if n = 0 then (a$n :: 'a::semiring_1) else a$n + a$(n - 1))" proof (cases n) case 0 then show ?thesis by (simp add: fps_mult_nth) next case (Suc m) have "((1 + fps_X)*a) $ n = sum (\i. (1 + fps_X) $ i * a $ (n - i)) {0..n}" by (simp add: fps_mult_nth) also have "\ = sum (\i. (1+fps_X)$i * a$(n-i)) {0.. 1}" unfolding Suc by (rule sum.mono_neutral_right) auto also have "\ = (if n = 0 then a$n else a$n + a$(n - 1))" by (simp add: Suc) finally show ?thesis . qed lemma fps_mult_right_fps_X_plus_1_nth: fixes a :: "'a :: semiring_1 fps" shows "(a*(1+fps_X)) $ n = (if n = 0 then a$n else a$n + a$(n - 1))" using fps_mult_fps_X_plus_1_nth by (simp add: distrib_left fps_mult_fps_X_commute distrib_right) lemma fps_X_neq_fps_const [simp]: "(fps_X :: 'a :: zero_neq_one fps) \ fps_const c" proof assume "(fps_X::'a fps) = fps_const (c::'a)" hence "fps_X$1 = (fps_const (c::'a))$1" by (simp only:) thus False by auto qed lemma fps_X_neq_zero [simp]: "(fps_X :: 'a :: zero_neq_one fps) \ 0" by (simp only: fps_const_0_eq_0[symmetric] fps_X_neq_fps_const) simp lemma fps_X_neq_one [simp]: "(fps_X :: 'a :: zero_neq_one fps) \ 1" by (simp only: fps_const_1_eq_1[symmetric] fps_X_neq_fps_const) simp lemma fps_X_neq_numeral [simp]: "fps_X \ numeral c" by (simp only: numeral_fps_const fps_X_neq_fps_const) simp lemma fps_X_pow_eq_fps_X_pow_iff [simp]: "fps_X ^ m = fps_X ^ n \ m = n" proof assume "(fps_X :: 'a fps) ^ m = fps_X ^ n" hence "(fps_X :: 'a fps) ^ m $ m = fps_X ^ n $ m" by (simp only:) thus "m = n" by (simp split: if_split_asm) qed simp_all subsection \Shifting and slicing\ definition fps_shift :: "nat \ 'a fps \ 'a fps" where "fps_shift n f = Abs_fps (\i. f $ (i + n))" lemma fps_shift_nth [simp]: "fps_shift n f $ i = f $ (i + n)" by (simp add: fps_shift_def) lemma fps_shift_0 [simp]: "fps_shift 0 f = f" by (intro fps_ext) (simp add: fps_shift_def) lemma fps_shift_zero [simp]: "fps_shift n 0 = 0" by (intro fps_ext) (simp add: fps_shift_def) lemma fps_shift_one: "fps_shift n 1 = (if n = 0 then 1 else 0)" by (intro fps_ext) (simp add: fps_shift_def) lemma fps_shift_fps_const: "fps_shift n (fps_const c) = (if n = 0 then fps_const c else 0)" by (intro fps_ext) (simp add: fps_shift_def) lemma fps_shift_numeral: "fps_shift n (numeral c) = (if n = 0 then numeral c else 0)" by (simp add: numeral_fps_const fps_shift_fps_const) lemma fps_shift_fps_X [simp]: "n \ 1 \ fps_shift n fps_X = (if n = 1 then 1 else 0)" by (intro fps_ext) (auto simp: fps_X_def) lemma fps_shift_fps_X_power [simp]: "n \ m \ fps_shift n (fps_X ^ m) = fps_X ^ (m - n)" by (intro fps_ext) auto lemma fps_shift_subdegree [simp]: "n \ subdegree f \ subdegree (fps_shift n f) = subdegree f - n" by (cases "f=0") (auto intro: subdegreeI simp: nth_less_subdegree_zero) lemma fps_shift_fps_shift: "fps_shift (m + n) f = fps_shift m (fps_shift n f)" by (rule fps_ext) (simp add: add_ac) lemma fps_shift_fps_shift_reorder: "fps_shift m (fps_shift n f) = fps_shift n (fps_shift m f)" using fps_shift_fps_shift[of m n f] fps_shift_fps_shift[of n m f] by (simp add: add.commute) lemma fps_shift_rev_shift: "m \ n \ fps_shift n (Abs_fps (\k. if k n \ fps_shift n (Abs_fps (\k. if kk. if k n" thus "fps_shift n (Abs_fps (\k. if k n" hence "\k. k \ m-n \ k+n-m = k - (m-n)" by auto thus "fps_shift n (Abs_fps (\k. if kk. if k subdegree (g :: 'b :: {comm_monoid_add, mult_zero} fps)" shows "fps_shift n (h*g) = h * fps_shift n g" proof- have case1: "\a b::'b fps. 1 \ subdegree b \ fps_shift 1 (a*b) = a * fps_shift 1 b" proof (rule fps_ext) fix a b :: "'b fps" and n :: nat assume b: "1 \ subdegree b" have "\i. i \ n \ n + 1 - i = (n-i) + 1" by (simp add: algebra_simps) with b show "fps_shift 1 (a*b) $ n = (a * fps_shift 1 b) $ n" by (simp add: fps_mult_nth nth_less_subdegree_zero) qed have "n \ subdegree g \ fps_shift n (h*g) = h * fps_shift n g" proof (induct n) case (Suc n) have "fps_shift (Suc n) (h*g) = fps_shift 1 (fps_shift n (h*g))" by (simp add: fps_shift_fps_shift[symmetric]) also have "\ = h * (fps_shift 1 (fps_shift n g))" using Suc case1 by force finally show ?case by (simp add: fps_shift_fps_shift[symmetric]) qed simp with assms show ?thesis by fast qed lemma fps_shift_mult_right_noncomm: assumes "n \ subdegree (g :: 'b :: {comm_monoid_add, mult_zero} fps)" shows "fps_shift n (g*h) = fps_shift n g * h" proof- have case1: "\a b::'b fps. 1 \ subdegree a \ fps_shift 1 (a*b) = fps_shift 1 a * b" proof (rule fps_ext) fix a b :: "'b fps" and n :: nat assume "1 \ subdegree a" hence "fps_shift 1 (a*b) $ n = (\i=Suc 0..Suc n. a$i * b$(n+1-i))" using sum.atLeast_Suc_atMost[of 0 "n+1" "\i. a$i * b$(n+1-i)"] by (simp add: fps_mult_nth nth_less_subdegree_zero) thus "fps_shift 1 (a*b) $ n = (fps_shift 1 a * b) $ n" using sum.shift_bounds_cl_Suc_ivl[of "\i. a$i * b$(n+1-i)" 0 n] by (simp add: fps_mult_nth) qed have "n \ subdegree g \ fps_shift n (g*h) = fps_shift n g * h" proof (induct n) case (Suc n) have "fps_shift (Suc n) (g*h) = fps_shift 1 (fps_shift n (g*h))" by (simp add: fps_shift_fps_shift[symmetric]) also have "\ = (fps_shift 1 (fps_shift n g)) * h" using Suc case1 by force finally show ?case by (simp add: fps_shift_fps_shift[symmetric]) qed simp with assms show ?thesis by fast qed lemma fps_shift_mult_right: assumes "n \ subdegree (g :: 'b :: comm_semiring_0 fps)" shows "fps_shift n (g*h) = h * fps_shift n g" by (simp add: assms fps_shift_mult_right_noncomm mult.commute) lemma fps_shift_mult_both: fixes f g :: "'a::{comm_monoid_add, mult_zero} fps" assumes "m \ subdegree f" "n \ subdegree g" shows "fps_shift m f * fps_shift n g = fps_shift (m+n) (f*g)" using assms by (simp add: fps_shift_mult fps_shift_mult_right_noncomm fps_shift_fps_shift) lemma fps_shift_subdegree_zero_iff [simp]: "fps_shift (subdegree f) f = 0 \ f = 0" by (subst (1) nth_subdegree_zero_iff[symmetric], cases "f = 0") (simp_all del: nth_subdegree_zero_iff) lemma fps_shift_times_fps_X: fixes f g :: "'a::{comm_monoid_add,mult_zero,monoid_mult} fps" shows "1 \ subdegree f \ fps_shift 1 f * fps_X = f" by (intro fps_ext) (simp add: nth_less_subdegree_zero) lemma fps_shift_times_fps_X' [simp]: fixes f :: "'a::{comm_monoid_add,mult_zero,monoid_mult} fps" shows "fps_shift 1 (f * fps_X) = f" by (intro fps_ext) (simp add: nth_less_subdegree_zero) lemma fps_shift_times_fps_X'': fixes f :: "'a::{comm_monoid_add,mult_zero,monoid_mult} fps" shows "1 \ n \ fps_shift n (f * fps_X) = fps_shift (n - 1) f" by (intro fps_ext) (simp add: nth_less_subdegree_zero) lemma fps_shift_times_fps_X_power: "n \ subdegree f \ fps_shift n f * fps_X ^ n = f" by (intro fps_ext) (simp add: fps_X_power_mult_right_nth nth_less_subdegree_zero) lemma fps_shift_times_fps_X_power' [simp]: "fps_shift n (f * fps_X^n) = f" by (intro fps_ext) (simp add: fps_X_power_mult_right_nth nth_less_subdegree_zero) lemma fps_shift_times_fps_X_power'': "m \ n \ fps_shift n (f * fps_X^m) = fps_shift (n - m) f" by (intro fps_ext) (simp add: fps_X_power_mult_right_nth nth_less_subdegree_zero) lemma fps_shift_times_fps_X_power''': "m > n \ fps_shift n (f * fps_X^m) = f * fps_X^(m - n)" proof (cases "f=0") case False assume m: "m>n" hence "m = n + (m-n)" by auto with False m show ?thesis using power_add[of "fps_X::'a fps" n "m-n"] fps_shift_mult_right_noncomm[of n "f * fps_X^n" "fps_X^(m-n)"] by (simp add: mult.assoc fps_subdegree_mult_fps_X_power(2)) qed simp lemma subdegree_decompose: "f = fps_shift (subdegree f) f * fps_X ^ subdegree f" by (rule fps_ext) (auto simp: fps_X_power_mult_right_nth) lemma subdegree_decompose': "n \ subdegree f \ f = fps_shift n f * fps_X^n" by (rule fps_ext) (auto simp: fps_X_power_mult_right_nth intro!: nth_less_subdegree_zero) instantiation fps :: (zero) unit_factor begin definition fps_unit_factor_def [simp]: "unit_factor f = fps_shift (subdegree f) f" instance .. end lemma fps_unit_factor_zero_iff: "unit_factor (f::'a::zero fps) = 0 \ f = 0" by simp lemma fps_unit_factor_nth_0: "f \ 0 \ unit_factor f $ 0 \ 0" by simp lemma fps_X_unit_factor: "unit_factor (fps_X :: 'a :: zero_neq_one fps) = 1" by (intro fps_ext) auto lemma fps_X_power_unit_factor: "unit_factor (fps_X ^ n) = 1" proof- define X :: "'a fps" where "X \ fps_X" hence "unit_factor (X^n) = fps_shift n (X^n)" by (simp add: fps_X_power_subdegree) moreover have "fps_shift n (X^n) = 1" by (auto intro: fps_ext simp: fps_X_power_iff X_def) ultimately show ?thesis by (simp add: X_def) qed lemma fps_unit_factor_decompose: "f = unit_factor f * fps_X ^ subdegree f" by (simp add: subdegree_decompose) lemma fps_unit_factor_decompose': "f = fps_X ^ subdegree f * unit_factor f" using fps_unit_factor_decompose by (simp add: fps_mult_fps_X_power_commute) lemma fps_unit_factor_uminus: "unit_factor (-f) = - unit_factor (f::'a::group_add fps)" by (simp add: fps_shift_uminus) lemma fps_unit_factor_shift: assumes "n \ subdegree f" shows "unit_factor (fps_shift n f) = unit_factor f" by (simp add: assms fps_shift_fps_shift[symmetric]) lemma fps_unit_factor_mult_fps_X: fixes f :: "'a::{comm_monoid_add,monoid_mult,mult_zero} fps" shows "unit_factor (fps_X * f) = unit_factor f" and "unit_factor (f * fps_X) = unit_factor f" proof - show "unit_factor (fps_X * f) = unit_factor f" by (cases "f=0") (auto intro: fps_ext simp: fps_subdegree_mult_fps_X(1)) thus "unit_factor (f * fps_X) = unit_factor f" by (simp add: fps_mult_fps_X_commute) qed lemma fps_unit_factor_mult_fps_X_power: shows "unit_factor (fps_X ^ n * f) = unit_factor f" and "unit_factor (f * fps_X ^ n) = unit_factor f" proof - show "unit_factor (fps_X ^ n * f) = unit_factor f" proof (induct n) case (Suc m) thus ?case using fps_unit_factor_mult_fps_X(1)[of "fps_X ^ m * f"] by (simp add: mult.assoc) qed simp thus "unit_factor (f * fps_X ^ n) = unit_factor f" by (simp add: fps_mult_fps_X_power_commute) qed lemma fps_unit_factor_mult_unit_factor: fixes f g :: "'a::{comm_monoid_add,mult_zero} fps" shows "unit_factor (f * unit_factor g) = unit_factor (f * g)" and "unit_factor (unit_factor f * g) = unit_factor (f * g)" proof - show "unit_factor (f * unit_factor g) = unit_factor (f * g)" proof (cases "f*g = 0") case False thus ?thesis using fps_mult_subdegree_ge[of f g] fps_unit_factor_shift[of "subdegree g" "f*g"] by (simp add: fps_shift_mult) next case True moreover have "f * unit_factor g = fps_shift (subdegree g) (f*g)" by (simp add: fps_shift_mult) ultimately show ?thesis by simp qed show "unit_factor (unit_factor f * g) = unit_factor (f * g)" proof (cases "f*g = 0") case False thus ?thesis using fps_mult_subdegree_ge[of f g] fps_unit_factor_shift[of "subdegree f" "f*g"] by (simp add: fps_shift_mult_right_noncomm) next case True moreover have "unit_factor f * g = fps_shift (subdegree f) (f*g)" by (simp add: fps_shift_mult_right_noncomm) ultimately show ?thesis by simp qed qed lemma fps_unit_factor_mult_both_unit_factor: fixes f g :: "'a::{comm_monoid_add,mult_zero} fps" shows "unit_factor (unit_factor f * unit_factor g) = unit_factor (f * g)" using fps_unit_factor_mult_unit_factor(1)[of "unit_factor f" g] fps_unit_factor_mult_unit_factor(2)[of f g] by simp lemma fps_unit_factor_mult': fixes f g :: "'a::{comm_monoid_add,mult_zero} fps" assumes "f $ subdegree f * g $ subdegree g \ 0" shows "unit_factor (f * g) = unit_factor f * unit_factor g" using assms by (simp add: subdegree_mult' fps_shift_mult_both) lemma fps_unit_factor_mult: fixes f g :: "'a::semiring_no_zero_divisors fps" shows "unit_factor (f * g) = unit_factor f * unit_factor g" using fps_unit_factor_mult'[of f g] by (cases "f=0 \ g=0") auto definition "fps_cutoff n f = Abs_fps (\i. if i < n then f$i else 0)" lemma fps_cutoff_nth [simp]: "fps_cutoff n f $ i = (if i < n then f$i else 0)" unfolding fps_cutoff_def by simp lemma fps_cutoff_zero_iff: "fps_cutoff n f = 0 \ (f = 0 \ n \ subdegree f)" proof assume A: "fps_cutoff n f = 0" thus "f = 0 \ n \ subdegree f" proof (cases "f = 0") assume "f \ 0" with A have "n \ subdegree f" by (intro subdegree_geI) (simp_all add: fps_eq_iff split: if_split_asm) thus ?thesis .. qed simp qed (auto simp: fps_eq_iff intro: nth_less_subdegree_zero) lemma fps_cutoff_0 [simp]: "fps_cutoff 0 f = 0" by (simp add: fps_eq_iff) lemma fps_cutoff_zero [simp]: "fps_cutoff n 0 = 0" by (simp add: fps_eq_iff) lemma fps_cutoff_one: "fps_cutoff n 1 = (if n = 0 then 0 else 1)" by (simp add: fps_eq_iff) lemma fps_cutoff_fps_const: "fps_cutoff n (fps_const c) = (if n = 0 then 0 else fps_const c)" by (simp add: fps_eq_iff) lemma fps_cutoff_numeral: "fps_cutoff n (numeral c) = (if n = 0 then 0 else numeral c)" by (simp add: numeral_fps_const fps_cutoff_fps_const) lemma fps_shift_cutoff: "fps_shift n f * fps_X^n + fps_cutoff n f = f" by (simp add: fps_eq_iff fps_X_power_mult_right_nth) lemma fps_shift_cutoff': "fps_X^n * fps_shift n f + fps_cutoff n f = f" by (simp add: fps_eq_iff fps_X_power_mult_nth) lemma fps_cutoff_left_mult_nth: "k < n \ (fps_cutoff n f * g) $ k = (f * g) $ k" by (simp add: fps_mult_nth) lemma fps_cutoff_right_mult_nth: assumes "k < n" shows "(f * fps_cutoff n g) $ k = (f * g) $ k" proof- from assms have "\i\{0..k}. fps_cutoff n g $ (k - i) = g $ (k - i)" by auto thus ?thesis by (simp add: fps_mult_nth) qed subsection \Formal Power series form a metric space\ instantiation fps :: ("{minus,zero}") dist begin definition dist_fps_def: "dist (a :: 'a fps) b = (if a = b then 0 else inverse (2 ^ subdegree (a - b)))" lemma dist_fps_ge0: "dist (a :: 'a fps) b \ 0" by (simp add: dist_fps_def) instance .. end instantiation fps :: (group_add) metric_space begin definition uniformity_fps_def [code del]: "(uniformity :: ('a fps \ 'a fps) filter) = (INF e\{0 <..}. principal {(x, y). dist x y < e})" definition open_fps_def' [code del]: "open (U :: 'a fps set) \ (\x\U. eventually (\(x', y). x' = x \ y \ U) uniformity)" lemma dist_fps_sym: "dist (a :: 'a fps) b = dist b a" by (simp add: dist_fps_def) instance proof show th: "dist a b = 0 \ a = b" for a b :: "'a fps" by (simp add: dist_fps_def split: if_split_asm) then have th'[simp]: "dist a a = 0" for a :: "'a fps" by simp fix a b c :: "'a fps" consider "a = b" | "c = a \ c = b" | "a \ b" "a \ c" "b \ c" by blast then show "dist a b \ dist a c + dist b c" proof cases case 1 then show ?thesis by (simp add: dist_fps_def) next case 2 then show ?thesis by (cases "c = a") (simp_all add: th dist_fps_sym) next case neq: 3 have False if "dist a b > dist a c + dist b c" proof - let ?n = "subdegree (a - b)" from neq have "dist a b > 0" "dist b c > 0" and "dist a c > 0" by (simp_all add: dist_fps_def) with that have "dist a b > dist a c" and "dist a b > dist b c" by simp_all with neq have "?n < subdegree (a - c)" and "?n < subdegree (b - c)" by (simp_all add: dist_fps_def field_simps) hence "(a - c) $ ?n = 0" and "(b - c) $ ?n = 0" by (simp_all only: nth_less_subdegree_zero) hence "(a - b) $ ?n = 0" by simp moreover from neq have "(a - b) $ ?n \ 0" by (intro nth_subdegree_nonzero) simp_all ultimately show False by contradiction qed thus ?thesis by (auto simp add: not_le[symmetric]) qed qed (rule open_fps_def' uniformity_fps_def)+ end declare uniformity_Abort[where 'a="'a :: group_add fps", code] lemma open_fps_def: "open (S :: 'a::group_add fps set) = (\a \ S. \r. r >0 \ {y. dist y a < r} \ S)" unfolding open_dist subset_eq by simp text \The infinite sums and justification of the notation in textbooks.\ lemma reals_power_lt_ex: fixes x y :: real assumes xp: "x > 0" and y1: "y > 1" shows "\k>0. (1/y)^k < x" proof - have yp: "y > 0" using y1 by simp from reals_Archimedean2[of "max 0 (- log y x) + 1"] obtain k :: nat where k: "real k > max 0 (- log y x) + 1" by blast from k have kp: "k > 0" by simp from k have "real k > - log y x" by simp then have "ln y * real k > - ln x" unfolding log_def using ln_gt_zero_iff[OF yp] y1 by (simp add: minus_divide_left field_simps del: minus_divide_left[symmetric]) then have "ln y * real k + ln x > 0" by simp then have "exp (real k * ln y + ln x) > exp 0" by (simp add: ac_simps) then have "y ^ k * x > 1" unfolding exp_zero exp_add exp_of_nat_mult exp_ln [OF xp] exp_ln [OF yp] by simp then have "x > (1 / y)^k" using yp by (simp add: field_simps) then show ?thesis using kp by blast qed lemma fps_sum_rep_nth: "(sum (\i. fps_const(a$i)*fps_X^i) {0..m})$n = (if n \ m then a$n else 0)" by (simp add: fps_sum_nth if_distrib cong del: if_weak_cong) lemma fps_notation: "(\n. sum (\i. fps_const(a$i) * fps_X^i) {0..n}) \ a" (is "?s \ a") proof - have "\n0. \n \ n0. dist (?s n) a < r" if "r > 0" for r proof - obtain n0 where n0: "(1/2)^n0 < r" "n0 > 0" using reals_power_lt_ex[OF \r > 0\, of 2] by auto show ?thesis proof - have "dist (?s n) a < r" if nn0: "n \ n0" for n proof - from that have thnn0: "(1/2)^n \ (1/2 :: real)^n0" by (simp add: field_split_simps) show ?thesis proof (cases "?s n = a") case True then show ?thesis unfolding dist_eq_0_iff[of "?s n" a, symmetric] using \r > 0\ by (simp del: dist_eq_0_iff) next case False from False have dth: "dist (?s n) a = (1/2)^subdegree (?s n - a)" by (simp add: dist_fps_def field_simps) from False have kn: "subdegree (?s n - a) > n" by (intro subdegree_greaterI) (simp_all add: fps_sum_rep_nth) then have "dist (?s n) a < (1/2)^n" by (simp add: field_simps dist_fps_def) also have "\ \ (1/2)^n0" using nn0 by (simp add: field_split_simps) also have "\ < r" using n0 by simp finally show ?thesis . qed qed then show ?thesis by blast qed qed then show ?thesis unfolding lim_sequentially by blast qed subsection \Inverses and division of formal power series\ declare sum.cong[fundef_cong] fun fps_left_inverse_constructor :: "'a::{comm_monoid_add,times,uminus} fps \ 'a \ nat \ 'a" where "fps_left_inverse_constructor f a 0 = a" | "fps_left_inverse_constructor f a (Suc n) = - sum (\i. fps_left_inverse_constructor f a i * f$(Suc n - i)) {0..n} * a" \ \This will construct a left inverse for f in case that x * f$0 = 1\ abbreviation "fps_left_inverse \ (\f x. Abs_fps (fps_left_inverse_constructor f x))" fun fps_right_inverse_constructor :: "'a::{comm_monoid_add,times,uminus} fps \ 'a \ nat \ 'a" where "fps_right_inverse_constructor f a 0 = a" | "fps_right_inverse_constructor f a n = - a * sum (\i. f$i * fps_right_inverse_constructor f a (n - i)) {1..n}" \ \This will construct a right inverse for f in case that f$0 * y = 1\ abbreviation "fps_right_inverse \ (\f y. Abs_fps (fps_right_inverse_constructor f y))" instantiation fps :: ("{comm_monoid_add,inverse,times,uminus}") inverse begin \ \For backwards compatibility.\ abbreviation natfun_inverse:: "'a fps \ nat \ 'a" where "natfun_inverse f \ fps_right_inverse_constructor f (inverse (f$0))" definition fps_inverse_def: "inverse f = Abs_fps (natfun_inverse f)" \ \ With scalars from a (possibly non-commutative) ring, this defines a right inverse. Furthermore, if scalars are of class @{class mult_zero} and satisfy condition @{term "inverse 0 = 0"}, then this will evaluate to zero when the zeroth term is zero. \ definition fps_divide_def: "f div g = fps_shift (subdegree g) (f * inverse (unit_factor g))" \ \ If scalars are of class @{class mult_zero} and satisfy condition @{term "inverse 0 = 0"}, then div by zero will equal zero. \ instance .. end lemma fps_lr_inverse_0_iff: "(fps_left_inverse f x) $ 0 = 0 \ x = 0" "(fps_right_inverse f x) $ 0 = 0 \ x = 0" by auto lemma fps_inverse_0_iff': "(inverse f) $ 0 = 0 \ inverse (f $ 0) = 0" by (simp add: fps_inverse_def fps_lr_inverse_0_iff(2)) lemma fps_inverse_0_iff[simp]: "(inverse f) $ 0 = (0::'a::division_ring) \ f $ 0 = 0" by (simp add: fps_inverse_0_iff') lemma fps_lr_inverse_nth_0: "(fps_left_inverse f x) $ 0 = x" "(fps_right_inverse f x) $ 0 = x" by auto lemma fps_inverse_nth_0 [simp]: "(inverse f) $ 0 = inverse (f $ 0)" by (simp add: fps_inverse_def) lemma fps_lr_inverse_starting0: fixes f :: "'a::{comm_monoid_add,mult_zero,uminus} fps" and g :: "'b::{ab_group_add,mult_zero} fps" shows "fps_left_inverse f 0 = 0" and "fps_right_inverse g 0 = 0" proof- show "fps_left_inverse f 0 = 0" proof (rule fps_ext) fix n show "fps_left_inverse f 0 $ n = 0 $ n" by (cases n) (simp_all add: fps_inverse_def) qed show "fps_right_inverse g 0 = 0" proof (rule fps_ext) fix n show "fps_right_inverse g 0 $ n = 0 $ n" by (cases n) (simp_all add: fps_inverse_def) qed qed lemma fps_lr_inverse_eq0_imp_starting0: "fps_left_inverse f x = 0 \ x = 0" "fps_right_inverse f x = 0 \ x = 0" proof- assume A: "fps_left_inverse f x = 0" have "0 = fps_left_inverse f x $ 0" by (subst A) simp thus "x = 0" by simp next assume A: "fps_right_inverse f x = 0" have "0 = fps_right_inverse f x $ 0" by (subst A) simp thus "x = 0" by simp qed lemma fps_lr_inverse_eq_0_iff: fixes x :: "'a::{comm_monoid_add,mult_zero,uminus}" and y :: "'b::{ab_group_add,mult_zero}" shows "fps_left_inverse f x = 0 \ x = 0" and "fps_right_inverse g y = 0 \ y = 0" using fps_lr_inverse_starting0 fps_lr_inverse_eq0_imp_starting0 by auto lemma fps_inverse_eq_0_iff': fixes f :: "'a::{ab_group_add,inverse,mult_zero} fps" shows "inverse f = 0 \ inverse (f $ 0) = 0" by (simp add: fps_inverse_def fps_lr_inverse_eq_0_iff(2)) lemma fps_inverse_eq_0_iff[simp]: "inverse f = (0:: ('a::division_ring) fps) \ f $ 0 = 0" using fps_inverse_eq_0_iff'[of f] by simp lemmas fps_inverse_eq_0' = iffD2[OF fps_inverse_eq_0_iff'] lemmas fps_inverse_eq_0 = iffD2[OF fps_inverse_eq_0_iff] lemma fps_const_lr_inverse: fixes a :: "'a::{ab_group_add,mult_zero}" and b :: "'b::{comm_monoid_add,mult_zero,uminus}" shows "fps_left_inverse (fps_const a) x = fps_const x" and "fps_right_inverse (fps_const b) y = fps_const y" proof- show "fps_left_inverse (fps_const a) x = fps_const x" proof (rule fps_ext) fix n show "fps_left_inverse (fps_const a) x $ n = fps_const x $ n" by (cases n) auto qed show "fps_right_inverse (fps_const b) y = fps_const y" proof (rule fps_ext) fix n show "fps_right_inverse (fps_const b) y $ n = fps_const y $ n" by (cases n) auto qed qed lemma fps_const_inverse: fixes a :: "'a::{comm_monoid_add,inverse,mult_zero,uminus}" shows "inverse (fps_const a) = fps_const (inverse a)" unfolding fps_inverse_def by (simp add: fps_const_lr_inverse(2)) lemma fps_lr_inverse_zero: fixes x :: "'a::{ab_group_add,mult_zero}" and y :: "'b::{comm_monoid_add,mult_zero,uminus}" shows "fps_left_inverse 0 x = fps_const x" and "fps_right_inverse 0 y = fps_const y" using fps_const_lr_inverse[of 0] by simp_all lemma fps_inverse_zero_conv_fps_const: "inverse (0::'a::{comm_monoid_add,mult_zero,uminus,inverse} fps) = fps_const (inverse 0)" using fps_lr_inverse_zero(2)[of "inverse (0::'a)"] by (simp add: fps_inverse_def) lemma fps_inverse_zero': assumes "inverse (0::'a::{comm_monoid_add,inverse,mult_zero,uminus}) = 0" shows "inverse (0::'a fps) = 0" by (simp add: assms fps_inverse_zero_conv_fps_const) lemma fps_inverse_zero [simp]: "inverse (0::'a::division_ring fps) = 0" by (rule fps_inverse_zero'[OF inverse_zero]) lemma fps_lr_inverse_one: fixes x :: "'a::{ab_group_add,mult_zero,one}" and y :: "'b::{comm_monoid_add,mult_zero,uminus,one}" shows "fps_left_inverse 1 x = fps_const x" and "fps_right_inverse 1 y = fps_const y" using fps_const_lr_inverse[of 1] by simp_all lemma fps_lr_inverse_one_one: "fps_left_inverse 1 1 = (1::'a::{ab_group_add,mult_zero,one} fps)" "fps_right_inverse 1 1 = (1::'b::{comm_monoid_add,mult_zero,uminus,one} fps)" by (simp_all add: fps_lr_inverse_one) lemma fps_inverse_one': assumes "inverse (1::'a::{comm_monoid_add,inverse,mult_zero,uminus,one}) = 1" shows "inverse (1 :: 'a fps) = 1" using assms fps_lr_inverse_one_one(2) by (simp add: fps_inverse_def) lemma fps_inverse_one [simp]: "inverse (1 :: 'a :: division_ring fps) = 1" by (rule fps_inverse_one'[OF inverse_1]) lemma fps_lr_inverse_minus: fixes f :: "'a::ring_1 fps" shows "fps_left_inverse (-f) (-x) = - fps_left_inverse f x" and "fps_right_inverse (-f) (-x) = - fps_right_inverse f x" proof- show "fps_left_inverse (-f) (-x) = - fps_left_inverse f x" proof (intro fps_ext) fix n show "fps_left_inverse (-f) (-x) $ n = - fps_left_inverse f x $ n" proof (induct n rule: nat_less_induct) case (1 n) thus ?case by (cases n) (simp_all add: sum_negf algebra_simps) qed qed show "fps_right_inverse (-f) (-x) = - fps_right_inverse f x" proof (intro fps_ext) fix n show "fps_right_inverse (-f) (-x) $ n = - fps_right_inverse f x $ n" proof (induct n rule: nat_less_induct) case (1 n) show ?case proof (cases n) case (Suc m) with 1 have "\i\{1..Suc m}. fps_right_inverse (-f) (-x) $ (Suc m - i) = - fps_right_inverse f x $ (Suc m - i)" by auto with Suc show ?thesis by (simp add: sum_negf algebra_simps) qed simp qed qed qed lemma fps_inverse_minus [simp]: "inverse (-f) = -inverse (f :: 'a :: division_ring fps)" by (simp add: fps_inverse_def fps_lr_inverse_minus(2)) lemma fps_left_inverse: fixes f :: "'a::ring_1 fps" assumes f0: "x * f$0 = 1" shows "fps_left_inverse f x * f = 1" proof (rule fps_ext) fix n show "(fps_left_inverse f x * f) $ n = 1 $ n" by (cases n) (simp_all add: f0 fps_mult_nth mult.assoc) qed lemma fps_right_inverse: fixes f :: "'a::ring_1 fps" assumes f0: "f$0 * y = 1" shows "f * fps_right_inverse f y = 1" proof (rule fps_ext) fix n show "(f * fps_right_inverse f y) $ n = 1 $ n" proof (cases n) case (Suc k) moreover from Suc have "fps_right_inverse f y $ n = - y * sum (\i. f$i * fps_right_inverse_constructor f y (n - i)) {1..n}" by simp hence "(f * fps_right_inverse f y) $ n = - 1 * sum (\i. f$i * fps_right_inverse_constructor f y (n - i)) {1..n} + sum (\i. f$i * (fps_right_inverse_constructor f y (n - i))) {1..n}" by (simp add: fps_mult_nth sum.atLeast_Suc_atMost mult.assoc f0[symmetric]) thus "(f * fps_right_inverse f y) $ n = 1 $ n" by (simp add: Suc) qed (simp add: f0 fps_inverse_def) qed \ \ It is possible in a ring for an element to have a left inverse but not a right inverse, or vice versa. But when an element has both, they must be the same. \ lemma fps_left_inverse_eq_fps_right_inverse: fixes f :: "'a::ring_1 fps" assumes f0: "x * f$0 = 1" "f $ 0 * y = 1" \ \These assumptions imply x equals y, but no need to assume that.\ shows "fps_left_inverse f x = fps_right_inverse f y" proof- from f0(2) have "f * fps_right_inverse f y = 1" by (simp add: fps_right_inverse) hence "fps_left_inverse f x * f * fps_right_inverse f y = fps_left_inverse f x" by (simp add: mult.assoc) moreover from f0(1) have "fps_left_inverse f x * f * fps_right_inverse f y = fps_right_inverse f y" by (simp add: fps_left_inverse) ultimately show ?thesis by simp qed lemma fps_left_inverse_eq_fps_right_inverse_comm: fixes f :: "'a::comm_ring_1 fps" assumes f0: "x * f$0 = 1" shows "fps_left_inverse f x = fps_right_inverse f x" using assms fps_left_inverse_eq_fps_right_inverse[of x f x] by (simp add: mult.commute) lemma fps_left_inverse': fixes f :: "'a::ring_1 fps" assumes "x * f$0 = 1" "f$0 * y = 1" \ \These assumptions imply x equals y, but no need to assume that.\ shows "fps_right_inverse f y * f = 1" using assms fps_left_inverse_eq_fps_right_inverse[of x f y] fps_left_inverse[of x f] by simp lemma fps_right_inverse': fixes f :: "'a::ring_1 fps" assumes "x * f$0 = 1" "f$0 * y = 1" \ \These assumptions imply x equals y, but no need to assume that.\ shows "f * fps_left_inverse f x = 1" using assms fps_left_inverse_eq_fps_right_inverse[of x f y] fps_right_inverse[of f y] by simp lemma inverse_mult_eq_1 [intro]: assumes "f$0 \ (0::'a::division_ring)" shows "inverse f * f = 1" using fps_left_inverse'[of "inverse (f$0)"] by (simp add: assms fps_inverse_def) lemma inverse_mult_eq_1': assumes "f$0 \ (0::'a::division_ring)" shows "f * inverse f = 1" using assms fps_right_inverse by (force simp: fps_inverse_def) lemma fps_mult_left_inverse_unit_factor: fixes f :: "'a::ring_1 fps" assumes "x * f $ subdegree f = 1" shows "fps_left_inverse (unit_factor f) x * f = fps_X ^ subdegree f" proof- have "fps_left_inverse (unit_factor f) x * f = fps_left_inverse (unit_factor f) x * unit_factor f * fps_X ^ subdegree f" using fps_unit_factor_decompose[of f] by (simp add: mult.assoc) with assms show ?thesis by (simp add: fps_left_inverse) qed lemma fps_mult_right_inverse_unit_factor: fixes f :: "'a::ring_1 fps" assumes "f $ subdegree f * y = 1" shows "f * fps_right_inverse (unit_factor f) y = fps_X ^ subdegree f" proof- have "f * fps_right_inverse (unit_factor f) y = fps_X ^ subdegree f * (unit_factor f * fps_right_inverse (unit_factor f) y)" using fps_unit_factor_decompose'[of f] by (simp add: mult.assoc[symmetric]) with assms show ?thesis by (simp add: fps_right_inverse) qed lemma fps_mult_right_inverse_unit_factor_divring: "(f :: 'a::division_ring fps) \ 0 \ f * inverse (unit_factor f) = fps_X ^ subdegree f" using fps_mult_right_inverse_unit_factor[of f] by (simp add: fps_inverse_def) lemma fps_left_inverse_idempotent_ring1: fixes f :: "'a::ring_1 fps" assumes "x * f$0 = 1" "y * x = 1" \ \These assumptions imply y equals f$0, but no need to assume that.\ shows "fps_left_inverse (fps_left_inverse f x) y = f" proof- from assms(1) have "fps_left_inverse (fps_left_inverse f x) y * fps_left_inverse f x * f = fps_left_inverse (fps_left_inverse f x) y" by (simp add: fps_left_inverse mult.assoc) moreover from assms(2) have "fps_left_inverse (fps_left_inverse f x) y * fps_left_inverse f x = 1" by (simp add: fps_left_inverse) ultimately show ?thesis by simp qed lemma fps_left_inverse_idempotent_comm_ring1: fixes f :: "'a::comm_ring_1 fps" assumes "x * f$0 = 1" shows "fps_left_inverse (fps_left_inverse f x) (f$0) = f" using assms fps_left_inverse_idempotent_ring1[of x f "f$0"] by (simp add: mult.commute) lemma fps_right_inverse_idempotent_ring1: fixes f :: "'a::ring_1 fps" assumes "f$0 * x = 1" "x * y = 1" \ \These assumptions imply y equals f$0, but no need to assume that.\ shows "fps_right_inverse (fps_right_inverse f x) y = f" proof- from assms(1) have "f * (fps_right_inverse f x * fps_right_inverse (fps_right_inverse f x) y) = fps_right_inverse (fps_right_inverse f x) y" by (simp add: fps_right_inverse mult.assoc[symmetric]) moreover from assms(2) have "fps_right_inverse f x * fps_right_inverse (fps_right_inverse f x) y = 1" by (simp add: fps_right_inverse) ultimately show ?thesis by simp qed lemma fps_right_inverse_idempotent_comm_ring1: fixes f :: "'a::comm_ring_1 fps" assumes "f$0 * x = 1" shows "fps_right_inverse (fps_right_inverse f x) (f$0) = f" using assms fps_right_inverse_idempotent_ring1[of f x "f$0"] by (simp add: mult.commute) lemma fps_inverse_idempotent[intro, simp]: "f$0 \ (0::'a::division_ring) \ inverse (inverse f) = f" using fps_right_inverse_idempotent_ring1[of f] by (simp add: fps_inverse_def) lemma fps_lr_inverse_unique_ring1: fixes f g :: "'a :: ring_1 fps" assumes fg: "f * g = 1" "g$0 * f$0 = 1" shows "fps_left_inverse g (f$0) = f" and "fps_right_inverse f (g$0) = g" proof- show "fps_left_inverse g (f$0) = f" proof (intro fps_ext) fix n show "fps_left_inverse g (f$0) $ n = f $ n" proof (induct n rule: nat_less_induct) case (1 n) show ?case proof (cases n) case (Suc k) hence "\i\{0..k}. fps_left_inverse g (f$0) $ i = f $ i" using 1 by simp hence "fps_left_inverse g (f$0) $ Suc k = f $ Suc k - 1 $ Suc k * f$0" by (simp add: fps_mult_nth fg(1)[symmetric] distrib_right mult.assoc fg(2)) with Suc show ?thesis by simp qed simp qed qed show "fps_right_inverse f (g$0) = g" proof (intro fps_ext) fix n show "fps_right_inverse f (g$0) $ n = g $ n" proof (induct n rule: nat_less_induct) case (1 n) show ?case proof (cases n) case (Suc k) hence "\i\{1..Suc k}. fps_right_inverse f (g$0) $ (Suc k - i) = g $ (Suc k - i)" using 1 by auto hence "fps_right_inverse f (g$0) $ Suc k = 1 * g $ Suc k - g$0 * 1 $ Suc k" by (simp add: fps_mult_nth fg(1)[symmetric] algebra_simps fg(2)[symmetric] sum.atLeast_Suc_atMost) with Suc show ?thesis by simp qed simp qed qed qed lemma fps_lr_inverse_unique_divring: fixes f g :: "'a ::division_ring fps" assumes fg: "f * g = 1" shows "fps_left_inverse g (f$0) = f" and "fps_right_inverse f (g$0) = g" proof- from fg have "f$0 * g$0 = 1" using fps_mult_nth_0[of f g] by simp hence "g$0 * f$0 = 1" using inverse_unique[of "f$0"] left_inverse[of "f$0"] by force thus "fps_left_inverse g (f$0) = f" "fps_right_inverse f (g$0) = g" using fg fps_lr_inverse_unique_ring1 by auto qed lemma fps_inverse_unique: fixes f g :: "'a :: division_ring fps" assumes fg: "f * g = 1" shows "inverse f = g" proof - from fg have if0: "inverse (f$0) = g$0" "f$0 \ 0" using inverse_unique[of "f$0"] fps_mult_nth_0[of f g] by auto with fg have "fps_right_inverse f (g$0) = g" using left_inverse[of "f$0"] by (intro fps_lr_inverse_unique_ring1(2)) simp_all with if0(1) show ?thesis by (simp add: fps_inverse_def) qed lemma inverse_fps_numeral: "inverse (numeral n :: ('a :: field_char_0) fps) = fps_const (inverse (numeral n))" by (intro fps_inverse_unique fps_ext) (simp_all add: fps_numeral_nth) lemma inverse_fps_of_nat: "inverse (of_nat n :: 'a :: {semiring_1,times,uminus,inverse} fps) = fps_const (inverse (of_nat n))" by (simp add: fps_of_nat fps_const_inverse[symmetric]) lemma sum_zero_lemma: fixes n::nat assumes "0 < n" shows "(\i = 0..n. if n = i then 1 else if n - i = 1 then - 1 else 0) = (0::'a::field)" proof - let ?f = "\i. if n = i then 1 else if n - i = 1 then - 1 else 0" let ?g = "\i. if i = n then 1 else if i = n - 1 then - 1 else 0" let ?h = "\i. if i=n - 1 then - 1 else 0" have th1: "sum ?f {0..n} = sum ?g {0..n}" by (rule sum.cong) auto have th2: "sum ?g {0..n - 1} = sum ?h {0..n - 1}" apply (rule sum.cong) using assms apply auto done have eq: "{0 .. n} = {0.. n - 1} \ {n}" by auto from assms have d: "{0.. n - 1} \ {n} = {}" by auto have f: "finite {0.. n - 1}" "finite {n}" by auto show ?thesis unfolding th1 apply (simp add: sum.union_disjoint[OF f d, unfolded eq[symmetric]] del: One_nat_def) unfolding th2 apply simp done qed lemma fps_lr_inverse_mult_ring1: fixes f g :: "'a::ring_1 fps" assumes x: "x * f$0 = 1" "f$0 * x = 1" and y: "y * g$0 = 1" "g$0 * y = 1" shows "fps_left_inverse (f * g) (y*x) = fps_left_inverse g y * fps_left_inverse f x" and "fps_right_inverse (f * g) (y*x) = fps_right_inverse g y * fps_right_inverse f x" proof - define h where "h \ fps_left_inverse g y * fps_left_inverse f x" hence h0: "h$0 = y*x" by simp have "fps_left_inverse (f*g) (h$0) = h" proof (intro fps_lr_inverse_unique_ring1(1)) from h_def have "h * (f * g) = fps_left_inverse g y * (fps_left_inverse f x * f) * g" by (simp add: mult.assoc) thus "h * (f * g) = 1" using fps_left_inverse[OF x(1)] fps_left_inverse[OF y(1)] by simp from h_def have "(f*g)$0 * h$0 = f$0 * 1 * x" by (simp add: mult.assoc y(2)[symmetric]) with x(2) show "(f * g) $ 0 * h $ 0 = 1" by simp qed with h_def show "fps_left_inverse (f * g) (y*x) = fps_left_inverse g y * fps_left_inverse f x" by simp next define h where "h \ fps_right_inverse g y * fps_right_inverse f x" hence h0: "h$0 = y*x" by simp have "fps_right_inverse (f*g) (h$0) = h" proof (intro fps_lr_inverse_unique_ring1(2)) from h_def have "f * g * h = f * (g * fps_right_inverse g y) * fps_right_inverse f x" by (simp add: mult.assoc) thus "f * g * h = 1" using fps_right_inverse[OF x(2)] fps_right_inverse[OF y(2)] by simp from h_def have "h$0 * (f*g)$0 = y * 1 * g$0" by (simp add: mult.assoc x(1)[symmetric]) with y(1) show "h$0 * (f*g)$0 = 1" by simp qed with h_def show "fps_right_inverse (f * g) (y*x) = fps_right_inverse g y * fps_right_inverse f x" by simp qed lemma fps_lr_inverse_mult_divring: fixes f g :: "'a::division_ring fps" shows "fps_left_inverse (f * g) (inverse ((f*g)$0)) = fps_left_inverse g (inverse (g$0)) * fps_left_inverse f (inverse (f$0))" and "fps_right_inverse (f * g) (inverse ((f*g)$0)) = fps_right_inverse g (inverse (g$0)) * fps_right_inverse f (inverse (f$0))" proof- show "fps_left_inverse (f * g) (inverse ((f*g)$0)) = fps_left_inverse g (inverse (g$0)) * fps_left_inverse f (inverse (f$0))" proof (cases "f$0 = 0 \ g$0 = 0") case True hence "fps_left_inverse (f * g) (inverse ((f*g)$0)) = 0" by (simp add: fps_lr_inverse_eq_0_iff(1)) moreover from True have "fps_left_inverse g (inverse (g$0)) * fps_left_inverse f (inverse (f$0)) = 0" by (auto simp: fps_lr_inverse_eq_0_iff(1)) ultimately show ?thesis by simp next case False hence "fps_left_inverse (f * g) (inverse (g$0) * inverse (f$0)) = fps_left_inverse g (inverse (g$0)) * fps_left_inverse f (inverse (f$0))" by (intro fps_lr_inverse_mult_ring1(1)) simp_all with False show ?thesis by (simp add: nonzero_inverse_mult_distrib) qed show "fps_right_inverse (f * g) (inverse ((f*g)$0)) = fps_right_inverse g (inverse (g$0)) * fps_right_inverse f (inverse (f$0))" proof (cases "f$0 = 0 \ g$0 = 0") case True from True have "fps_right_inverse (f * g) (inverse ((f*g)$0)) = 0" by (simp add: fps_lr_inverse_eq_0_iff(2)) moreover from True have "fps_right_inverse g (inverse (g$0)) * fps_right_inverse f (inverse (f$0)) = 0" by (auto simp: fps_lr_inverse_eq_0_iff(2)) ultimately show ?thesis by simp next case False hence "fps_right_inverse (f * g) (inverse (g$0) * inverse (f$0)) = fps_right_inverse g (inverse (g$0)) * fps_right_inverse f (inverse (f$0))" by (intro fps_lr_inverse_mult_ring1(2)) simp_all with False show ?thesis by (simp add: nonzero_inverse_mult_distrib) qed qed lemma fps_inverse_mult_divring: "inverse (f * g) = inverse g * inverse (f :: 'a::division_ring fps)" using fps_lr_inverse_mult_divring(2) by (simp add: fps_inverse_def) lemma fps_inverse_mult: "inverse (f * g :: 'a::field fps) = inverse f * inverse g" by (simp add: fps_inverse_mult_divring) lemma fps_lr_inverse_gp_ring1: fixes ones ones_inv :: "'a :: ring_1 fps" defines "ones \ Abs_fps (\n. 1)" and "ones_inv \ Abs_fps (\n. if n=0 then 1 else if n=1 then - 1 else 0)" shows "fps_left_inverse ones 1 = ones_inv" and "fps_right_inverse ones 1 = ones_inv" proof- show "fps_left_inverse ones 1 = ones_inv" proof (rule fps_ext) fix n show "fps_left_inverse ones 1 $ n = ones_inv $ n" proof (induct n rule: nat_less_induct) case (1 n) show ?case proof (cases n) case (Suc m) have m: "n = Suc m" by fact moreover have "fps_left_inverse ones 1 $ Suc m = ones_inv $ Suc m" proof (cases m) case (Suc k) thus ?thesis using Suc m 1 by (simp add: ones_def ones_inv_def sum.atLeast_Suc_atMost) qed (simp add: ones_def ones_inv_def) ultimately show ?thesis by simp qed (simp add: ones_inv_def) qed qed moreover have "fps_right_inverse ones 1 = fps_left_inverse ones 1" by (auto intro: fps_left_inverse_eq_fps_right_inverse[symmetric] simp: ones_def) ultimately show "fps_right_inverse ones 1 = ones_inv" by simp qed lemma fps_lr_inverse_gp_ring1': fixes ones :: "'a :: ring_1 fps" defines "ones \ Abs_fps (\n. 1)" shows "fps_left_inverse ones 1 = 1 - fps_X" and "fps_right_inverse ones 1 = 1 - fps_X" proof- define ones_inv :: "'a :: ring_1 fps" where "ones_inv \ Abs_fps (\n. if n=0 then 1 else if n=1 then - 1 else 0)" hence "fps_left_inverse ones 1 = ones_inv" and "fps_right_inverse ones 1 = ones_inv" using ones_def fps_lr_inverse_gp_ring1 by auto thus "fps_left_inverse ones 1 = 1 - fps_X" and "fps_right_inverse ones 1 = 1 - fps_X" by (auto intro: fps_ext simp: ones_inv_def) qed lemma fps_inverse_gp: "inverse (Abs_fps(\n. (1::'a::division_ring))) = Abs_fps (\n. if n= 0 then 1 else if n=1 then - 1 else 0)" using fps_lr_inverse_gp_ring1(2) by (simp add: fps_inverse_def) lemma fps_inverse_gp': "inverse (Abs_fps (\n. 1::'a::division_ring)) = 1 - fps_X" by (simp add: fps_inverse_def fps_lr_inverse_gp_ring1'(2)) lemma fps_lr_inverse_one_minus_fps_X: fixes ones :: "'a :: ring_1 fps" defines "ones \ Abs_fps (\n. 1)" shows "fps_left_inverse (1 - fps_X) 1 = ones" and "fps_right_inverse (1 - fps_X) 1 = ones" proof- have "fps_left_inverse ones 1 = 1 - fps_X" using fps_lr_inverse_gp_ring1'(1) by (simp add: ones_def) thus "fps_left_inverse (1 - fps_X) 1 = ones" using fps_left_inverse_idempotent_ring1[of 1 ones 1] by (simp add: ones_def) have "fps_right_inverse ones 1 = 1 - fps_X" using fps_lr_inverse_gp_ring1'(2) by (simp add: ones_def) thus "fps_right_inverse (1 - fps_X) 1 = ones" using fps_right_inverse_idempotent_ring1[of ones 1 1] by (simp add: ones_def) qed lemma fps_inverse_one_minus_fps_X: fixes ones :: "'a :: division_ring fps" defines "ones \ Abs_fps (\n. 1)" shows "inverse (1 - fps_X) = ones" by (simp add: fps_inverse_def assms fps_lr_inverse_one_minus_fps_X(2)) lemma fps_lr_one_over_one_minus_fps_X_squared: shows "fps_left_inverse ((1 - fps_X)^2) (1::'a::ring_1) = Abs_fps (\n. of_nat (n+1))" "fps_right_inverse ((1 - fps_X)^2) (1::'a) = Abs_fps (\n. of_nat (n+1))" proof- define f invf2 :: "'a fps" where "f \ (1 - fps_X)" and "invf2 \ Abs_fps (\n. of_nat (n+1))" have f2_nth_simps: "f^2 $ 1 = - of_nat 2" "f^2 $ 2 = 1" "\n. n>2 \ f^2 $ n = 0" by (simp_all add: power2_eq_square f_def fps_mult_nth sum.atLeast_Suc_atMost) show "fps_left_inverse (f^2) 1 = invf2" proof (intro fps_ext) fix n show "fps_left_inverse (f^2) 1 $ n = invf2 $ n" proof (induct n rule: nat_less_induct) case (1 t) hence induct_assm: "\m. m < t \ fps_left_inverse (f\<^sup>2) 1 $ m = invf2 $ m" by fast show ?case proof (cases t) case (Suc m) have m: "t = Suc m" by fact moreover have "fps_left_inverse (f^2) 1 $ Suc m = invf2 $ Suc m" proof (cases m) case 0 thus ?thesis using f2_nth_simps(1) by (simp add: invf2_def) next case (Suc l) have l: "m = Suc l" by fact moreover have "fps_left_inverse (f^2) 1 $ Suc (Suc l) = invf2 $ Suc (Suc l)" proof (cases l) case 0 thus ?thesis using f2_nth_simps(1,2) by (simp add: Suc_1[symmetric] invf2_def) next case (Suc k) from Suc l m have A: "fps_left_inverse (f\<^sup>2) 1 $ Suc (Suc k) = invf2 $ Suc (Suc k)" and B: "fps_left_inverse (f\<^sup>2) 1 $ Suc k = invf2 $ Suc k" using induct_assm[of "Suc k"] induct_assm[of "Suc (Suc k)"] by auto have times2: "\a::nat. 2*a = a + a" by simp have "\i\{0..k}. (f^2)$(Suc (Suc (Suc k)) - i) = 0" using f2_nth_simps(3) by auto hence "fps_left_inverse (f^2) 1 $ Suc (Suc (Suc k)) = fps_left_inverse (f\<^sup>2) 1 $ Suc (Suc k) * of_nat 2 - fps_left_inverse (f\<^sup>2) 1 $ Suc k" using sum.ub_add_nat f2_nth_simps(1,2) by simp also have "\ = of_nat (2 * Suc (Suc (Suc k))) - of_nat (Suc (Suc k))" by (subst A, subst B) (simp add: invf2_def mult.commute) also have "\ = of_nat (Suc (Suc (Suc k)) + 1)" by (subst times2[of "Suc (Suc (Suc k))"]) simp finally have "fps_left_inverse (f^2) 1 $ Suc (Suc (Suc k)) = invf2 $ Suc (Suc (Suc k))" by (simp add: invf2_def) with Suc show ?thesis by simp qed ultimately show ?thesis by simp qed ultimately show ?thesis by simp qed (simp add: invf2_def) qed qed moreover have "fps_right_inverse (f^2) 1 = fps_left_inverse (f^2) 1" by (auto intro: fps_left_inverse_eq_fps_right_inverse[symmetric] simp: f_def power2_eq_square ) ultimately show "fps_right_inverse (f^2) 1 = invf2" by simp qed lemma fps_one_over_one_minus_fps_X_squared': assumes "inverse (1::'a::{ring_1,inverse}) = 1" shows "inverse ((1 - fps_X)^2 :: 'a fps) = Abs_fps (\n. of_nat (n+1))" using assms fps_lr_one_over_one_minus_fps_X_squared(2) by (simp add: fps_inverse_def power2_eq_square) lemma fps_one_over_one_minus_fps_X_squared: "inverse ((1 - fps_X)^2 :: 'a :: division_ring fps) = Abs_fps (\n. of_nat (n+1))" by (rule fps_one_over_one_minus_fps_X_squared'[OF inverse_1]) lemma fps_lr_inverse_fps_X_plus1: "fps_left_inverse (1 + fps_X) (1::'a::ring_1) = Abs_fps (\n. (-1)^n)" "fps_right_inverse (1 + fps_X) (1::'a) = Abs_fps (\n. (-1)^n)" proof- show "fps_left_inverse (1 + fps_X) (1::'a) = Abs_fps (\n. (-1)^n)" proof (rule fps_ext) fix n show "fps_left_inverse (1 + fps_X) (1::'a) $ n = Abs_fps (\n. (-1)^n) $ n" proof (induct n rule: nat_less_induct) case (1 n) show ?case proof (cases n) case (Suc m) have m: "n = Suc m" by fact from Suc 1 have A: "fps_left_inverse (1 + fps_X) (1::'a) $ n = - (\i=0..m. (- 1)^i * (1 + fps_X) $ (Suc m - i))" by simp show ?thesis proof (cases m) case (Suc l) have "\i\{0..l}. ((1::'a fps) + fps_X) $ (Suc (Suc l) - i) = 0" by auto with Suc A m show ?thesis by simp qed (simp add: m A) qed simp qed qed moreover have "fps_right_inverse (1 + fps_X) (1::'a) = fps_left_inverse (1 + fps_X) 1" by (intro fps_left_inverse_eq_fps_right_inverse[symmetric]) simp_all ultimately show "fps_right_inverse (1 + fps_X) (1::'a) = Abs_fps (\n. (-1)^n)" by simp qed lemma fps_inverse_fps_X_plus1': assumes "inverse (1::'a::{ring_1,inverse}) = 1" shows "inverse (1 + fps_X) = Abs_fps (\n. (- (1::'a)) ^ n)" using assms fps_lr_inverse_fps_X_plus1(2) by (simp add: fps_inverse_def) lemma fps_inverse_fps_X_plus1: "inverse (1 + fps_X) = Abs_fps (\n. (- (1::'a::division_ring)) ^ n)" by (rule fps_inverse_fps_X_plus1'[OF inverse_1]) lemma subdegree_lr_inverse: fixes x :: "'a::{comm_monoid_add,mult_zero,uminus}" and y :: "'b::{ab_group_add,mult_zero}" shows "subdegree (fps_left_inverse f x) = 0" and "subdegree (fps_right_inverse g y) = 0" proof- show "subdegree (fps_left_inverse f x) = 0" using fps_lr_inverse_eq_0_iff(1) subdegree_eq_0_iff by fastforce show "subdegree (fps_right_inverse g y) = 0" using fps_lr_inverse_eq_0_iff(2) subdegree_eq_0_iff by fastforce qed lemma subdegree_inverse [simp]: fixes f :: "'a::{ab_group_add,inverse,mult_zero} fps" shows "subdegree (inverse f) = 0" using subdegree_lr_inverse(2) by (simp add: fps_inverse_def) lemma fps_div_zero [simp]: "0 div (g :: 'a :: {comm_monoid_add,inverse,mult_zero,uminus} fps) = 0" by (simp add: fps_divide_def) lemma fps_div_by_zero': fixes g :: "'a::{comm_monoid_add,inverse,mult_zero,uminus} fps" assumes "inverse (0::'a) = 0" shows "g div 0 = 0" by (simp add: fps_divide_def assms fps_inverse_zero') lemma fps_div_by_zero [simp]: "(g::'a::division_ring fps) div 0 = 0" by (rule fps_div_by_zero'[OF inverse_zero]) lemma fps_divide_unit': "subdegree g = 0 \ f div g = f * inverse g" by (simp add: fps_divide_def) lemma fps_divide_unit: "g$0 \ 0 \ f div g = f * inverse g" by (intro fps_divide_unit') (simp add: subdegree_eq_0_iff) lemma fps_divide_nth_0': "subdegree (g::'a::division_ring fps) = 0 \ (f div g) $ 0 = f $ 0 / (g $ 0)" by (simp add: fps_divide_unit' divide_inverse) lemma fps_divide_nth_0 [simp]: "g $ 0 \ 0 \ (f div g) $ 0 = f $ 0 / (g $ 0 :: _ :: division_ring)" by (simp add: fps_divide_nth_0') lemma fps_divide_nth_below: fixes f g :: "'a::{comm_monoid_add,uminus,mult_zero,inverse} fps" shows "n < subdegree f - subdegree g \ (f div g) $ n = 0" by (simp add: fps_divide_def fps_mult_nth_eq0) lemma fps_divide_nth_base: fixes f g :: "'a::division_ring fps" assumes "subdegree g \ subdegree f" shows "(f div g) $ (subdegree f - subdegree g) = f $ subdegree f * inverse (g $ subdegree g)" by (simp add: assms fps_divide_def fps_divide_unit') lemma fps_divide_subdegree_ge: fixes f g :: "'a::{comm_monoid_add,uminus,mult_zero,inverse} fps" assumes "f / g \ 0" shows "subdegree (f / g) \ subdegree f - subdegree g" by (intro subdegree_geI) (simp_all add: assms fps_divide_nth_below) lemma fps_divide_subdegree: fixes f g :: "'a::division_ring fps" assumes "f \ 0" "g \ 0" "subdegree g \ subdegree f" shows "subdegree (f / g) = subdegree f - subdegree g" proof (intro antisym) from assms have 1: "(f div g) $ (subdegree f - subdegree g) \ 0" using fps_divide_nth_base[of g f] by simp thus "subdegree (f / g) \ subdegree f - subdegree g" by (intro subdegree_leI) simp from 1 have "f / g \ 0" by (auto intro: fps_nonzeroI) thus "subdegree f - subdegree g \ subdegree (f / g)" by (rule fps_divide_subdegree_ge) qed lemma fps_divide_shift_numer: fixes f g :: "'a::{inverse,comm_monoid_add,uminus,mult_zero} fps" assumes "n \ subdegree f" shows "fps_shift n f / g = fps_shift n (f/g)" using assms fps_shift_mult_right_noncomm[of n f "inverse (unit_factor g)"] fps_shift_fps_shift_reorder[of "subdegree g" n "f * inverse (unit_factor g)"] by (simp add: fps_divide_def) lemma fps_divide_shift_denom: fixes f g :: "'a::{inverse,comm_monoid_add,uminus,mult_zero} fps" assumes "n \ subdegree g" "subdegree g \ subdegree f" shows "f / fps_shift n g = Abs_fps (\k. if kk. if k subdegree f" shows "f / unit_factor g = Abs_fps (\k. if k subdegree f" shows "unit_factor f / unit_factor g = fps_shift (subdegree f - subdegree g) (f / g)" using assms fps_divide_unit_factor_numer[of f "unit_factor g"] fps_divide_unit_factor_denom[of g f] fps_shift_rev_shift(1)[of "subdegree g" "subdegree f" "f/g"] by simp lemma fps_divide_unit_factor_both: fixes f g :: "'a::division_ring fps" assumes "subdegree g \ subdegree f" shows "unit_factor f / unit_factor g = unit_factor (f / g)" using assms fps_divide_unit_factor_both'[of g f] fps_divide_subdegree[of f g] by (cases "f=0 \ g=0") auto lemma fps_divide_self: "(f::'a::division_ring fps) \ 0 \ f / f = 1" using fps_mult_right_inverse_unit_factor_divring[of f] by (simp add: fps_divide_def) lemma fps_divide_add: fixes f g h :: "'a::{semiring_0,inverse,uminus} fps" shows "(f + g) / h = f / h + g / h" by (simp add: fps_divide_def algebra_simps fps_shift_add) lemma fps_divide_diff: fixes f g h :: "'a::{ring,inverse} fps" shows "(f - g) / h = f / h - g / h" by (simp add: fps_divide_def algebra_simps fps_shift_diff) lemma fps_divide_uminus: fixes f g h :: "'a::{ring,inverse} fps" shows "(- f) / g = - (f / g)" by (simp add: fps_divide_def algebra_simps fps_shift_uminus) lemma fps_divide_uminus': fixes f g h :: "'a::division_ring fps" shows "f / (- g) = - (f / g)" by (simp add: fps_divide_def fps_unit_factor_uminus fps_shift_uminus) lemma fps_divide_times: fixes f g h :: "'a::{semiring_0,inverse,uminus} fps" assumes "subdegree h \ subdegree g" shows "(f * g) / h = f * (g / h)" using assms fps_mult_subdegree_ge[of g "inverse (unit_factor h)"] fps_shift_mult[of "subdegree h" "g * inverse (unit_factor h)" f] by (fastforce simp add: fps_divide_def mult.assoc) lemma fps_divide_times2: fixes f g h :: "'a::{comm_semiring_0,inverse,uminus} fps" assumes "subdegree h \ subdegree f" shows "(f * g) / h = (f / h) * g" using assms fps_divide_times[of h f g] by (simp add: mult.commute) lemma fps_times_divide_eq: fixes f g :: "'a::field fps" assumes "g \ 0" and "subdegree f \ subdegree g" shows "f div g * g = f" using assms fps_divide_times2[of g f g] by (simp add: fps_divide_times fps_divide_self) lemma fps_divide_times_eq: "(g :: 'a::division_ring fps) \ 0 \ (f * g) div g = f" by (simp add: fps_divide_times fps_divide_self) lemma fps_divide_by_mult': fixes f g h :: "'a :: division_ring fps" assumes "subdegree h \ subdegree f" shows "f / (g * h) = f / h / g" proof (cases "f=0 \ g=0 \ h=0") case False with assms show ?thesis using fps_unit_factor_mult[of g h] by (auto simp: fps_divide_def fps_shift_fps_shift fps_inverse_mult_divring mult.assoc fps_shift_mult_right_noncomm ) qed auto lemma fps_divide_by_mult: fixes f g h :: "'a :: field fps" assumes "subdegree g \ subdegree f" shows "f / (g * h) = f / g / h" proof- have "f / (g * h) = f / (h * g)" by (simp add: mult.commute) also have "\ = f / g / h" using fps_divide_by_mult'[OF assms] by simp finally show ?thesis by simp qed lemma fps_divide_cancel: fixes f g h :: "'a :: division_ring fps" shows "h \ 0 \ (f * h) div (g * h) = f div g" by (cases "f=0") (auto simp: fps_divide_by_mult' fps_divide_times_eq) lemma fps_divide_1': fixes a :: "'a::{comm_monoid_add,inverse,mult_zero,uminus,zero_neq_one,monoid_mult} fps" assumes "inverse (1::'a) = 1" shows "a / 1 = a" using assms fps_inverse_one' fps_one_mult(2)[of a] by (force simp: fps_divide_def) lemma fps_divide_1 [simp]: "(a :: 'a::division_ring fps) / 1 = a" by (rule fps_divide_1'[OF inverse_1]) lemma fps_divide_X': fixes f :: "'a::{comm_monoid_add,inverse,mult_zero,uminus,zero_neq_one,monoid_mult} fps" assumes "inverse (1::'a) = 1" shows "f / fps_X = fps_shift 1 f" using assms fps_one_mult(2)[of f] by (simp add: fps_divide_def fps_X_unit_factor fps_inverse_one') lemma fps_divide_X [simp]: "a / fps_X = fps_shift 1 (a::'a::division_ring fps)" by (rule fps_divide_X'[OF inverse_1]) lemma fps_divide_X_power': fixes f :: "'a::{semiring_1,inverse,uminus} fps" assumes "inverse (1::'a) = 1" shows "f / (fps_X ^ n) = fps_shift n f" using fps_inverse_one'[OF assms] fps_one_mult(2)[of f] by (simp add: fps_divide_def fps_X_power_subdegree) lemma fps_divide_X_power [simp]: "a / (fps_X ^ n) = fps_shift n (a::'a::division_ring fps)" by (rule fps_divide_X_power'[OF inverse_1]) lemma fps_divide_shift_denom_conv_times_fps_X_power: fixes f g :: "'a::{semiring_1,inverse,uminus} fps" assumes "n \ subdegree g" "subdegree g \ subdegree f" shows "f / fps_shift n g = f / g * fps_X ^ n" using assms by (intro fps_ext) (simp_all add: fps_divide_shift_denom fps_X_power_mult_right_nth) lemma fps_divide_unit_factor_denom_conv_times_fps_X_power: fixes f g :: "'a::{semiring_1,inverse,uminus} fps" assumes "subdegree g \ subdegree f" shows "f / unit_factor g = f / g * fps_X ^ subdegree g" by (simp add: assms fps_divide_shift_denom_conv_times_fps_X_power) lemma fps_shift_altdef': fixes f :: "'a::{semiring_1,inverse,uminus} fps" assumes "inverse (1::'a) = 1" shows "fps_shift n f = f div fps_X^n" using assms by (simp add: fps_divide_def fps_X_power_subdegree fps_X_power_unit_factor fps_inverse_one' ) lemma fps_shift_altdef: "fps_shift n f = (f :: 'a :: division_ring fps) div fps_X^n" by (rule fps_shift_altdef'[OF inverse_1]) lemma fps_div_fps_X_power_nth': fixes f :: "'a::{semiring_1,inverse,uminus} fps" assumes "inverse (1::'a) = 1" shows "(f div fps_X^n) $ k = f $ (k + n)" using assms by (simp add: fps_shift_altdef' [symmetric]) lemma fps_div_fps_X_power_nth: "((f :: 'a :: division_ring fps) div fps_X^n) $ k = f $ (k + n)" by (rule fps_div_fps_X_power_nth'[OF inverse_1]) lemma fps_div_fps_X_nth': fixes f :: "'a::{semiring_1,inverse,uminus} fps" assumes "inverse (1::'a) = 1" shows "(f div fps_X) $ k = f $ Suc k" using assms fps_div_fps_X_power_nth'[of f 1] by simp lemma fps_div_fps_X_nth: "((f :: 'a :: division_ring fps) div fps_X) $ k = f $ Suc k" by (rule fps_div_fps_X_nth'[OF inverse_1]) lemma divide_fps_const': fixes c :: "'a :: {inverse,comm_monoid_add,uminus,mult_zero}" shows "f / fps_const c = f * fps_const (inverse c)" by (simp add: fps_divide_def fps_const_inverse) lemma divide_fps_const [simp]: fixes c :: "'a :: {comm_semiring_0,inverse,uminus}" shows "f / fps_const c = fps_const (inverse c) * f" by (simp add: divide_fps_const' mult.commute) lemma fps_const_divide: "fps_const (x :: _ :: division_ring) / fps_const y = fps_const (x / y)" by (simp add: fps_divide_def fps_const_inverse divide_inverse) lemma fps_numeral_divide_divide: "x / numeral b / numeral c = (x / numeral (b * c) :: 'a :: field fps)" by (simp add: fps_divide_by_mult[symmetric]) lemma fps_numeral_mult_divide: "numeral b * x / numeral c = (numeral b / numeral c * x :: 'a :: field fps)" by (simp add: fps_divide_times2) lemmas fps_numeral_simps = fps_numeral_divide_divide fps_numeral_mult_divide inverse_fps_numeral neg_numeral_fps_const lemma fps_is_left_unit_iff_zeroth_is_left_unit: fixes f :: "'a :: ring_1 fps" shows "(\g. 1 = f * g) \ (\k. 1 = f$0 * k)" proof assume "\g. 1 = f * g" then obtain g where "1 = f * g" by fast hence "1 = f$0 * g$0" using fps_mult_nth_0[of f g] by simp thus "\k. 1 = f$0 * k" by auto next assume "\k. 1 = f$0 * k" then obtain k where "1 = f$0 * k" by fast hence "1 = f * fps_right_inverse f k" using fps_right_inverse by simp thus "\g. 1 = f * g" by fast qed lemma fps_is_right_unit_iff_zeroth_is_right_unit: fixes f :: "'a :: ring_1 fps" shows "(\g. 1 = g * f) \ (\k. 1 = k * f$0)" proof assume "\g. 1 = g * f" then obtain g where "1 = g * f" by fast hence "1 = g$0 * f$0" using fps_mult_nth_0[of g f] by simp thus "\k. 1 = k * f$0" by auto next assume "\k. 1 = k * f$0" then obtain k where "1 = k * f$0" by fast hence "1 = fps_left_inverse f k * f" using fps_left_inverse by simp thus "\g. 1 = g * f" by fast qed lemma fps_is_unit_iff [simp]: "(f :: 'a :: field fps) dvd 1 \ f $ 0 \ 0" proof assume "f dvd 1" then obtain g where "1 = f * g" by (elim dvdE) from this[symmetric] have "(f*g) $ 0 = 1" by simp thus "f $ 0 \ 0" by auto next assume A: "f $ 0 \ 0" thus "f dvd 1" by (simp add: inverse_mult_eq_1[OF A, symmetric]) qed lemma subdegree_eq_0_left: fixes f :: "'a::{comm_monoid_add,zero_neq_one,mult_zero} fps" assumes "\g. 1 = f * g" shows "subdegree f = 0" proof (intro subdegree_eq_0) from assms obtain g where "1 = f * g" by fast hence "f$0 * g$0 = 1" using fps_mult_nth_0[of f g] by simp thus "f$0 \ 0" by auto qed lemma subdegree_eq_0_right: fixes f :: "'a::{comm_monoid_add,zero_neq_one,mult_zero} fps" assumes "\g. 1 = g * f" shows "subdegree f = 0" proof (intro subdegree_eq_0) from assms obtain g where "1 = g * f" by fast hence "g$0 * f$0 = 1" using fps_mult_nth_0[of g f] by simp thus "f$0 \ 0" by auto qed lemma subdegree_eq_0' [simp]: "(f :: 'a :: field fps) dvd 1 \ subdegree f = 0" by simp lemma fps_dvd1_left_trivial_unit_factor: fixes f :: "'a::{comm_monoid_add, zero_neq_one, mult_zero} fps" assumes "\g. 1 = f * g" shows "unit_factor f = f" using assms subdegree_eq_0_left by fastforce lemma fps_dvd1_right_trivial_unit_factor: fixes f :: "'a::{comm_monoid_add, zero_neq_one, mult_zero} fps" assumes "\g. 1 = g * f" shows "unit_factor f = f" using assms subdegree_eq_0_right by fastforce lemma fps_dvd1_trivial_unit_factor: "(f :: 'a::comm_semiring_1 fps) dvd 1 \ unit_factor f = f" unfolding dvd_def by (rule fps_dvd1_left_trivial_unit_factor) simp lemma fps_unit_dvd_left: fixes f :: "'a :: division_ring fps" assumes "f $ 0 \ 0" shows "\g. 1 = f * g" using assms fps_is_left_unit_iff_zeroth_is_left_unit right_inverse by fastforce lemma fps_unit_dvd_right: fixes f :: "'a :: division_ring fps" assumes "f $ 0 \ 0" shows "\g. 1 = g * f" using assms fps_is_right_unit_iff_zeroth_is_right_unit left_inverse by fastforce lemma fps_unit_dvd [simp]: "(f $ 0 :: 'a :: field) \ 0 \ f dvd g" using fps_unit_dvd_left dvd_trans[of f 1] by simp lemma dvd_left_imp_subdegree_le: fixes f g :: "'a::{comm_monoid_add,mult_zero} fps" assumes "\k. g = f * k" "g \ 0" shows "subdegree f \ subdegree g" using assms fps_mult_subdegree_ge by fastforce lemma dvd_right_imp_subdegree_le: fixes f g :: "'a::{comm_monoid_add,mult_zero} fps" assumes "\k. g = k * f" "g \ 0" shows "subdegree f \ subdegree g" using assms fps_mult_subdegree_ge by fastforce lemma dvd_imp_subdegree_le: "f dvd g \ g \ 0 \ subdegree f \ subdegree g" using dvd_left_imp_subdegree_le by fast lemma subdegree_le_imp_dvd_left_ring1: fixes f g :: "'a :: ring_1 fps" assumes "\y. f $ subdegree f * y = 1" "subdegree f \ subdegree g" shows "\k. g = f * k" proof- define h :: "'a fps" where "h \ fps_X ^ (subdegree g - subdegree f)" from assms(1) obtain y where "f $ subdegree f * y = 1" by fast hence "unit_factor f $ 0 * y = 1" by simp from this obtain k where "1 = unit_factor f * k" using fps_is_left_unit_iff_zeroth_is_left_unit[of "unit_factor f"] by auto hence "fps_X ^ subdegree f = fps_X ^ subdegree f * unit_factor f * k" by (simp add: mult.assoc) moreover have "fps_X ^ subdegree f * unit_factor f = f" by (rule fps_unit_factor_decompose'[symmetric]) ultimately have "fps_X ^ (subdegree f + (subdegree g - subdegree f)) = f * k * h" by (simp add: power_add h_def) hence "g = f * (k * h * unit_factor g)" using fps_unit_factor_decompose'[of g] by (simp add: assms(2) mult.assoc) thus ?thesis by fast qed lemma subdegree_le_imp_dvd_left_divring: fixes f g :: "'a :: division_ring fps" assumes "f \ 0" "subdegree f \ subdegree g" shows "\k. g = f * k" proof (intro subdegree_le_imp_dvd_left_ring1) from assms(1) have "f $ subdegree f \ 0" by simp thus "\y. f $ subdegree f * y = 1" using right_inverse by blast qed (rule assms(2)) lemma subdegree_le_imp_dvd_right_ring1: fixes f g :: "'a :: ring_1 fps" assumes "\x. x * f $ subdegree f = 1" "subdegree f \ subdegree g" shows "\k. g = k * f" proof- define h :: "'a fps" where "h \ fps_X ^ (subdegree g - subdegree f)" from assms(1) obtain x where "x * f $ subdegree f = 1" by fast hence "x * unit_factor f $ 0 = 1" by simp from this obtain k where "1 = k * unit_factor f" using fps_is_right_unit_iff_zeroth_is_right_unit[of "unit_factor f"] by auto hence "fps_X ^ subdegree f = k * (unit_factor f * fps_X ^ subdegree f)" by (simp add: mult.assoc[symmetric]) moreover have "unit_factor f * fps_X ^ subdegree f = f" by (rule fps_unit_factor_decompose[symmetric]) ultimately have "fps_X ^ (subdegree g - subdegree f + subdegree f) = h * k * f" by (simp add: power_add h_def mult.assoc) hence "g = unit_factor g * h * k * f" using fps_unit_factor_decompose[of g] by (simp add: assms(2) mult.assoc) thus ?thesis by fast qed lemma subdegree_le_imp_dvd_right_divring: fixes f g :: "'a :: division_ring fps" assumes "f \ 0" "subdegree f \ subdegree g" shows "\k. g = k * f" proof (intro subdegree_le_imp_dvd_right_ring1) from assms(1) have "f $ subdegree f \ 0" by simp thus "\x. x * f $ subdegree f = 1" using left_inverse by blast qed (rule assms(2)) lemma fps_dvd_iff: assumes "(f :: 'a :: field fps) \ 0" "g \ 0" shows "f dvd g \ subdegree f \ subdegree g" proof assume "subdegree f \ subdegree g" with assms show "f dvd g" using subdegree_le_imp_dvd_left_divring by (auto intro: dvdI) qed (simp add: assms dvd_imp_subdegree_le) lemma subdegree_div': fixes p q :: "'a::division_ring fps" assumes "\k. p = k * q" shows "subdegree (p div q) = subdegree p - subdegree q" proof (cases "p = 0") case False from assms(1) obtain k where k: "p = k * q" by blast with False have "subdegree (p div q) = subdegree k" by (simp add: fps_divide_times_eq) moreover have "k $ subdegree k * q $ subdegree q \ 0" proof assume "k $ subdegree k * q $ subdegree q = 0" hence "k $ subdegree k * q $ subdegree q * inverse (q $ subdegree q) = 0" by simp with False k show False by (simp add: mult.assoc) qed ultimately show ?thesis by (simp add: k subdegree_mult') qed simp lemma subdegree_div: fixes p q :: "'a :: field fps" assumes "q dvd p" shows "subdegree (p div q) = subdegree p - subdegree q" using assms unfolding dvd_def by (auto intro: subdegree_div') lemma subdegree_div_unit': fixes p q :: "'a :: {ab_group_add,mult_zero,inverse} fps" assumes "q $ 0 \ 0" "p $ subdegree p * inverse (q $ 0) \ 0" shows "subdegree (p div q) = subdegree p" using assms subdegree_mult'[of p "inverse q"] by (auto simp add: fps_divide_unit) lemma subdegree_div_unit'': fixes p q :: "'a :: {ring_no_zero_divisors,inverse} fps" assumes "q $ 0 \ 0" "inverse (q $ 0) \ 0" shows "subdegree (p div q) = subdegree p" by (cases "p = 0") (auto intro: subdegree_div_unit' simp: assms) lemma subdegree_div_unit: fixes p q :: "'a :: division_ring fps" assumes "q $ 0 \ 0" shows "subdegree (p div q) = subdegree p" by (intro subdegree_div_unit'') (simp_all add: assms) instantiation fps :: ("{comm_semiring_1,inverse,uminus}") modulo begin definition fps_mod_def: "f mod g = (if g = 0 then f else let h = unit_factor g in fps_cutoff (subdegree g) (f * inverse h) * h)" instance .. end lemma fps_mod_zero [simp]: "(f::'a::{comm_semiring_1,inverse,uminus} fps) mod 0 = f" by (simp add: fps_mod_def) lemma fps_mod_eq_zero: assumes "g \ 0" and "subdegree f \ subdegree g" shows "f mod g = 0" proof (cases "f * inverse (unit_factor g) = 0") case False have "fps_cutoff (subdegree g) (f * inverse (unit_factor g)) = 0" using False assms(2) fps_mult_subdegree_ge fps_cutoff_zero_iff by force with assms(1) show ?thesis by (simp add: fps_mod_def Let_def) qed (simp add: assms fps_mod_def) lemma fps_mod_unit [simp]: "g$0 \ 0 \ f mod g = 0" by (intro fps_mod_eq_zero) auto lemma subdegree_mod: assumes "subdegree (f::'a::field fps) < subdegree g" shows "subdegree (f mod g) = subdegree f" proof (cases "f = 0") case False with assms show ?thesis by (intro subdegreeI) (auto simp: inverse_mult_eq_1 fps_mod_def Let_def fps_cutoff_left_mult_nth mult.assoc) qed (simp add: fps_mod_def) instance fps :: (field) idom_modulo proof fix f g :: "'a fps" define n where "n = subdegree g" define h where "h = f * inverse (unit_factor g)" show "f div g * g + f mod g = f" proof (cases "g = 0") case False with n_def h_def have "f div g * g + f mod g = (fps_shift n h * fps_X ^ n + fps_cutoff n h) * unit_factor g" by (simp add: fps_divide_def fps_mod_def Let_def subdegree_decompose algebra_simps) with False show ?thesis by (simp add: fps_shift_cutoff h_def inverse_mult_eq_1) qed auto qed (rule fps_divide_times_eq, simp_all add: fps_divide_def) -instantiation fps :: (field) normalization_semidom +instantiation fps :: (field) normalization_semidom_multiplicative begin definition fps_normalize_def [simp]: "normalize f = (if f = 0 then 0 else fps_X ^ subdegree f)" instance proof fix f g :: "'a fps" - show "unit_factor (f * g) = unit_factor f * unit_factor g" - using fps_unit_factor_mult by simp + assume "is_unit f" + thus "unit_factor (f * g) = f * unit_factor g" + using fps_unit_factor_mult[of f g] by simp +next + fix f g :: "'a fps" show "unit_factor f * normalize f = f" by (simp add: fps_shift_times_fps_X_power) +next + fix f g :: "'a fps" + show "unit_factor (f * g) = unit_factor f * unit_factor g" + using fps_unit_factor_mult[of f g] by simp qed (simp_all add: fps_divide_def Let_def) end subsection \Formal power series form a Euclidean ring\ instantiation fps :: (field) euclidean_ring_cancel begin definition fps_euclidean_size_def: "euclidean_size f = (if f = 0 then 0 else 2 ^ subdegree f)" instance proof fix f g :: "'a fps" assume [simp]: "g \ 0" show "euclidean_size f \ euclidean_size (f * g)" by (cases "f = 0") (simp_all add: fps_euclidean_size_def) show "euclidean_size (f mod g) < euclidean_size g" apply (cases "f = 0", simp add: fps_euclidean_size_def) apply (rule disjE[OF le_less_linear[of "subdegree g" "subdegree f"]]) apply (simp_all add: fps_mod_eq_zero fps_euclidean_size_def subdegree_mod) done next fix f g h :: "'a fps" assume [simp]: "h \ 0" show "(h * f) div (h * g) = f div g" by (simp add: fps_divide_cancel mult.commute) show "(f + g * h) div h = g + f div h" by (simp add: fps_divide_add fps_divide_times_eq) qed (simp add: fps_euclidean_size_def) end instance fps :: (field) normalization_euclidean_semiring .. instantiation fps :: (field) euclidean_ring_gcd begin definition fps_gcd_def: "(gcd :: 'a fps \ _) = Euclidean_Algorithm.gcd" definition fps_lcm_def: "(lcm :: 'a fps \ _) = Euclidean_Algorithm.lcm" definition fps_Gcd_def: "(Gcd :: 'a fps set \ _) = Euclidean_Algorithm.Gcd" definition fps_Lcm_def: "(Lcm :: 'a fps set \ _) = Euclidean_Algorithm.Lcm" instance by standard (simp_all add: fps_gcd_def fps_lcm_def fps_Gcd_def fps_Lcm_def) end lemma fps_gcd: assumes [simp]: "f \ 0" "g \ 0" shows "gcd f g = fps_X ^ min (subdegree f) (subdegree g)" proof - let ?m = "min (subdegree f) (subdegree g)" show "gcd f g = fps_X ^ ?m" proof (rule sym, rule gcdI) fix d assume "d dvd f" "d dvd g" thus "d dvd fps_X ^ ?m" by (cases "d = 0") (simp_all add: fps_dvd_iff) qed (simp_all add: fps_dvd_iff) qed lemma fps_gcd_altdef: "gcd f g = (if f = 0 \ g = 0 then 0 else if f = 0 then fps_X ^ subdegree g else if g = 0 then fps_X ^ subdegree f else fps_X ^ min (subdegree f) (subdegree g))" by (simp add: fps_gcd) lemma fps_lcm: assumes [simp]: "f \ 0" "g \ 0" shows "lcm f g = fps_X ^ max (subdegree f) (subdegree g)" proof - let ?m = "max (subdegree f) (subdegree g)" show "lcm f g = fps_X ^ ?m" proof (rule sym, rule lcmI) fix d assume "f dvd d" "g dvd d" thus "fps_X ^ ?m dvd d" by (cases "d = 0") (simp_all add: fps_dvd_iff) qed (simp_all add: fps_dvd_iff) qed lemma fps_lcm_altdef: "lcm f g = (if f = 0 \ g = 0 then 0 else fps_X ^ max (subdegree f) (subdegree g))" by (simp add: fps_lcm) lemma fps_Gcd: assumes "A - {0} \ {}" shows "Gcd A = fps_X ^ (INF f\A-{0}. subdegree f)" proof (rule sym, rule GcdI) fix f assume "f \ A" thus "fps_X ^ (INF f\A - {0}. subdegree f) dvd f" by (cases "f = 0") (auto simp: fps_dvd_iff intro!: cINF_lower) next fix d assume d: "\f. f \ A \ d dvd f" from assms obtain f where "f \ A - {0}" by auto with d[of f] have [simp]: "d \ 0" by auto from d assms have "subdegree d \ (INF f\A-{0}. subdegree f)" by (intro cINF_greatest) (simp_all add: fps_dvd_iff[symmetric]) with d assms show "d dvd fps_X ^ (INF f\A-{0}. subdegree f)" by (simp add: fps_dvd_iff) qed simp_all lemma fps_Gcd_altdef: "Gcd A = (if A \ {0} then 0 else fps_X ^ (INF f\A-{0}. subdegree f))" using fps_Gcd by auto lemma fps_Lcm: assumes "A \ {}" "0 \ A" "bdd_above (subdegree`A)" shows "Lcm A = fps_X ^ (SUP f\A. subdegree f)" proof (rule sym, rule LcmI) fix f assume "f \ A" moreover from assms(3) have "bdd_above (subdegree ` A)" by auto ultimately show "f dvd fps_X ^ (SUP f\A. subdegree f)" using assms(2) by (cases "f = 0") (auto simp: fps_dvd_iff intro!: cSUP_upper) next fix d assume d: "\f. f \ A \ f dvd d" from assms obtain f where f: "f \ A" "f \ 0" by auto show "fps_X ^ (SUP f\A. subdegree f) dvd d" proof (cases "d = 0") assume "d \ 0" moreover from d have "\f. f \ A \ f \ 0 \ f dvd d" by blast ultimately have "subdegree d \ (SUP f\A. subdegree f)" using assms by (intro cSUP_least) (auto simp: fps_dvd_iff) with \d \ 0\ show ?thesis by (simp add: fps_dvd_iff) qed simp_all qed simp_all lemma fps_Lcm_altdef: "Lcm A = (if 0 \ A \ \bdd_above (subdegree`A) then 0 else if A = {} then 1 else fps_X ^ (SUP f\A. subdegree f))" proof (cases "bdd_above (subdegree`A)") assume unbounded: "\bdd_above (subdegree`A)" have "Lcm A = 0" proof (rule ccontr) assume "Lcm A \ 0" from unbounded obtain f where f: "f \ A" "subdegree (Lcm A) < subdegree f" unfolding bdd_above_def by (auto simp: not_le) moreover from f and \Lcm A \ 0\ have "subdegree f \ subdegree (Lcm A)" by (intro dvd_imp_subdegree_le dvd_Lcm) simp_all ultimately show False by simp qed with unbounded show ?thesis by simp qed (simp_all add: fps_Lcm Lcm_eq_0_I) subsection \Formal Derivatives, and the MacLaurin theorem around 0\ definition "fps_deriv f = Abs_fps (\n. of_nat (n + 1) * f $ (n + 1))" lemma fps_deriv_nth[simp]: "fps_deriv f $ n = of_nat (n + 1) * f $ (n + 1)" by (simp add: fps_deriv_def) lemma fps_0th_higher_deriv: "(fps_deriv ^^ n) f $ 0 = fact n * f $ n" by (induction n arbitrary: f) (simp_all add: funpow_Suc_right mult_of_nat_commute algebra_simps del: funpow.simps) lemma fps_deriv_mult[simp]: "fps_deriv (f * g) = f * fps_deriv g + fps_deriv f * g" proof (intro fps_ext) fix n have LHS: "fps_deriv (f * g) $ n = (\i=0..Suc n. of_nat (n+1) * f$i * g$(Suc n - i))" by (simp add: fps_mult_nth sum_distrib_left algebra_simps) have "\i\{1..n}. n - (i - 1) = n - i + 1" by auto moreover have "(\i=0..n. of_nat (i+1) * f$(i+1) * g$(n - i)) = (\i=1..Suc n. of_nat i * f$i * g$(n - (i - 1)))" by (intro sum.reindex_bij_witness[where i="\x. x-1" and j="\x. x+1"]) auto ultimately have "(f * fps_deriv g + fps_deriv f * g) $ n = of_nat (Suc n) * f$0 * g$(Suc n) + (\i=1..n. (of_nat (n - i + 1) + of_nat i) * f $ i * g $ (n - i + 1)) + of_nat (Suc n) * f$(Suc n) * g$0" by (simp add: fps_mult_nth algebra_simps mult_of_nat_commute sum.atLeast_Suc_atMost sum.distrib) moreover have "\i\{1..n}. (of_nat (n - i + 1) + of_nat i) * f $ i * g $ (n - i + 1) = of_nat (n + 1) * f $ i * g $ (Suc n - i)" proof fix i assume i: "i \ {1..n}" from i have "of_nat (n - i + 1) + (of_nat i :: 'a) = of_nat (n + 1)" using of_nat_add[of "n-i+1" i,symmetric] by simp moreover from i have "Suc n - i = n - i + 1" by auto ultimately show "(of_nat (n - i + 1) + of_nat i) * f $ i * g $ (n - i + 1) = of_nat (n + 1) * f $ i * g $ (Suc n - i)" by simp qed ultimately have "(f * fps_deriv g + fps_deriv f * g) $ n = (\i=0..Suc n. of_nat (Suc n) * f $ i * g $ (Suc n - i))" by (simp add: sum.atLeast_Suc_atMost) with LHS show "fps_deriv (f * g) $ n = (f * fps_deriv g + fps_deriv f * g) $ n" by simp qed lemma fps_deriv_fps_X[simp]: "fps_deriv fps_X = 1" by (simp add: fps_deriv_def fps_X_def fps_eq_iff) lemma fps_deriv_neg[simp]: "fps_deriv (- (f:: 'a::ring_1 fps)) = - (fps_deriv f)" by (simp add: fps_eq_iff fps_deriv_def) lemma fps_deriv_add[simp]: "fps_deriv (f + g) = fps_deriv f + fps_deriv g" by (auto intro: fps_ext simp: algebra_simps) lemma fps_deriv_sub[simp]: "fps_deriv ((f:: 'a::ring_1 fps) - g) = fps_deriv f - fps_deriv g" using fps_deriv_add [of f "- g"] by simp lemma fps_deriv_const[simp]: "fps_deriv (fps_const c) = 0" by (simp add: fps_ext fps_deriv_def fps_const_def) lemma fps_deriv_of_nat [simp]: "fps_deriv (of_nat n) = 0" by (simp add: fps_of_nat [symmetric]) lemma fps_deriv_of_int [simp]: "fps_deriv (of_int n) = 0" by (simp add: fps_of_int [symmetric]) lemma fps_deriv_numeral [simp]: "fps_deriv (numeral n) = 0" by (simp add: numeral_fps_const) lemma fps_deriv_mult_const_left[simp]: "fps_deriv (fps_const c * f) = fps_const c * fps_deriv f" by simp lemma fps_deriv_linear[simp]: "fps_deriv (fps_const a * f + fps_const b * g) = fps_const a * fps_deriv f + fps_const b * fps_deriv g" by simp lemma fps_deriv_0[simp]: "fps_deriv 0 = 0" by (simp add: fps_deriv_def fps_eq_iff) lemma fps_deriv_1[simp]: "fps_deriv 1 = 0" by (simp add: fps_deriv_def fps_eq_iff) lemma fps_deriv_mult_const_right[simp]: "fps_deriv (f * fps_const c) = fps_deriv f * fps_const c" by simp lemma fps_deriv_sum: "fps_deriv (sum f S) = sum (\i. fps_deriv (f i)) S" proof (cases "finite S") case False then show ?thesis by simp next case True show ?thesis by (induct rule: finite_induct [OF True]) simp_all qed lemma fps_deriv_eq_0_iff [simp]: "fps_deriv f = 0 \ f = fps_const (f$0 :: 'a::{semiring_no_zero_divisors,semiring_char_0})" proof assume f: "fps_deriv f = 0" show "f = fps_const (f$0)" proof (intro fps_ext) fix n show "f $ n = fps_const (f$0) $ n" proof (cases n) case (Suc m) have "(of_nat (Suc m) :: 'a) \ 0" by (rule of_nat_neq_0) with f Suc show ?thesis using fps_deriv_nth[of f] by auto qed simp qed next show "f = fps_const (f$0) \ fps_deriv f = 0" using fps_deriv_const[of "f$0"] by simp qed lemma fps_deriv_eq_iff: fixes f g :: "'a::{ring_1_no_zero_divisors,semiring_char_0} fps" shows "fps_deriv f = fps_deriv g \ (f = fps_const(f$0 - g$0) + g)" proof - have "fps_deriv f = fps_deriv g \ fps_deriv (f - g) = 0" using fps_deriv_sub[of f g] by simp also have "\ \ f - g = fps_const ((f - g) $ 0)" unfolding fps_deriv_eq_0_iff .. finally show ?thesis by (simp add: field_simps) qed lemma fps_deriv_eq_iff_ex: fixes f g :: "'a::{ring_1_no_zero_divisors,semiring_char_0} fps" shows "(fps_deriv f = fps_deriv g) \ (\c. f = fps_const c + g)" by (auto simp: fps_deriv_eq_iff) fun fps_nth_deriv :: "nat \ 'a::semiring_1 fps \ 'a fps" where "fps_nth_deriv 0 f = f" | "fps_nth_deriv (Suc n) f = fps_nth_deriv n (fps_deriv f)" lemma fps_nth_deriv_commute: "fps_nth_deriv (Suc n) f = fps_deriv (fps_nth_deriv n f)" by (induct n arbitrary: f) auto lemma fps_nth_deriv_linear[simp]: "fps_nth_deriv n (fps_const a * f + fps_const b * g) = fps_const a * fps_nth_deriv n f + fps_const b * fps_nth_deriv n g" by (induct n arbitrary: f g) auto lemma fps_nth_deriv_neg[simp]: "fps_nth_deriv n (- (f :: 'a::ring_1 fps)) = - (fps_nth_deriv n f)" by (induct n arbitrary: f) simp_all lemma fps_nth_deriv_add[simp]: "fps_nth_deriv n ((f :: 'a::ring_1 fps) + g) = fps_nth_deriv n f + fps_nth_deriv n g" using fps_nth_deriv_linear[of n 1 f 1 g] by simp lemma fps_nth_deriv_sub[simp]: "fps_nth_deriv n ((f :: 'a::ring_1 fps) - g) = fps_nth_deriv n f - fps_nth_deriv n g" using fps_nth_deriv_add [of n f "- g"] by simp lemma fps_nth_deriv_0[simp]: "fps_nth_deriv n 0 = 0" by (induct n) simp_all lemma fps_nth_deriv_1[simp]: "fps_nth_deriv n 1 = (if n = 0 then 1 else 0)" by (induct n) simp_all lemma fps_nth_deriv_const[simp]: "fps_nth_deriv n (fps_const c) = (if n = 0 then fps_const c else 0)" by (cases n) simp_all lemma fps_nth_deriv_mult_const_left[simp]: "fps_nth_deriv n (fps_const c * f) = fps_const c * fps_nth_deriv n f" using fps_nth_deriv_linear[of n "c" f 0 0 ] by simp lemma fps_nth_deriv_mult_const_right[simp]: "fps_nth_deriv n (f * fps_const c) = fps_nth_deriv n f * fps_const c" by (induct n arbitrary: f) auto lemma fps_nth_deriv_sum: "fps_nth_deriv n (sum f S) = sum (\i. fps_nth_deriv n (f i :: 'a::ring_1 fps)) S" proof (cases "finite S") case True show ?thesis by (induct rule: finite_induct [OF True]) simp_all next case False then show ?thesis by simp qed lemma fps_deriv_maclauren_0: "(fps_nth_deriv k (f :: 'a::comm_semiring_1 fps)) $ 0 = of_nat (fact k) * f $ k" by (induct k arbitrary: f) (simp_all add: field_simps) lemma fps_deriv_lr_inverse: fixes x y :: "'a::ring_1" assumes "x * f$0 = 1" "f$0 * y = 1" \ \These assumptions imply x equals y, but no need to assume that.\ shows "fps_deriv (fps_left_inverse f x) = - fps_left_inverse f x * fps_deriv f * fps_left_inverse f x" and "fps_deriv (fps_right_inverse f y) = - fps_right_inverse f y * fps_deriv f * fps_right_inverse f y" proof- define L where "L \ fps_left_inverse f x" hence "fps_deriv (L * f) = 0" using fps_left_inverse[OF assms(1)] by simp with assms show "fps_deriv L = - L * fps_deriv f * L" using fps_right_inverse'[OF assms] by (simp add: minus_unique mult.assoc L_def) define R where "R \ fps_right_inverse f y" hence "fps_deriv (f * R) = 0" using fps_right_inverse[OF assms(2)] by simp hence 1: "f * fps_deriv R + fps_deriv f * R = 0" by simp have "R * f * fps_deriv R = - R * fps_deriv f * R" using iffD2[OF eq_neg_iff_add_eq_0, OF 1] by (simp add: mult.assoc) thus "fps_deriv R = - R * fps_deriv f * R" using fps_left_inverse'[OF assms] by (simp add: R_def) qed lemma fps_deriv_lr_inverse_comm: fixes x :: "'a::comm_ring_1" assumes "x * f$0 = 1" shows "fps_deriv (fps_left_inverse f x) = - fps_deriv f * (fps_left_inverse f x)\<^sup>2" and "fps_deriv (fps_right_inverse f x) = - fps_deriv f * (fps_right_inverse f x)\<^sup>2" using assms fps_deriv_lr_inverse[of x f x] by (simp_all add: mult.commute power2_eq_square) lemma fps_inverse_deriv_divring: fixes a :: "'a::division_ring fps" assumes "a$0 \ 0" shows "fps_deriv (inverse a) = - inverse a * fps_deriv a * inverse a" using assms fps_deriv_lr_inverse(2)[of "inverse (a$0)" a "inverse (a$0)"] by (simp add: fps_inverse_def) lemma fps_inverse_deriv: fixes a :: "'a::field fps" assumes "a$0 \ 0" shows "fps_deriv (inverse a) = - fps_deriv a * (inverse a)\<^sup>2" using assms fps_deriv_lr_inverse_comm(2)[of "inverse (a$0)" a] by (simp add: fps_inverse_def) lemma fps_inverse_deriv': fixes a :: "'a::field fps" assumes a0: "a $ 0 \ 0" shows "fps_deriv (inverse a) = - fps_deriv a / a\<^sup>2" using fps_inverse_deriv[OF a0] a0 by (simp add: fps_divide_unit power2_eq_square fps_inverse_mult) (* FIXME: The last part of this proof should go through by simp once we have a proper theorem collection for simplifying division on rings *) lemma fps_divide_deriv: assumes "b dvd (a :: 'a :: field fps)" shows "fps_deriv (a / b) = (fps_deriv a * b - a * fps_deriv b) / b^2" proof - have eq_divide_imp: "c \ 0 \ a * c = b \ a = b div c" for a b c :: "'a :: field fps" by (drule sym) (simp add: mult.assoc) from assms have "a = a / b * b" by simp also have "fps_deriv (a / b * b) = fps_deriv (a / b) * b + a / b * fps_deriv b" by simp finally have "fps_deriv (a / b) * b^2 = fps_deriv a * b - a * fps_deriv b" using assms by (simp add: power2_eq_square algebra_simps) thus ?thesis by (cases "b = 0") (simp_all add: eq_divide_imp) qed lemma fps_nth_deriv_fps_X[simp]: "fps_nth_deriv n fps_X = (if n = 0 then fps_X else if n=1 then 1 else 0)" by (cases n) simp_all subsection \Powers\ lemma fps_power_zeroth: "(a^n) $ 0 = (a$0)^n" by (induct n) auto lemma fps_power_zeroth_eq_one: "a$0 = 1 \ a^n $ 0 = 1" by (simp add: fps_power_zeroth) lemma fps_power_first: fixes a :: "'a::comm_semiring_1 fps" shows "(a^n) $ 1 = of_nat n * (a$0)^(n-1) * a$1" proof (cases n) case (Suc m) have "(a ^ Suc m) $ 1 = of_nat (Suc m) * (a$0)^(Suc m - 1) * a$1" proof (induct m) case (Suc k) hence "(a ^ Suc (Suc k)) $ 1 = a$0 * of_nat (Suc k) * (a $ 0)^k * a$1 + a$1 * ((a$0)^(Suc k))" using fps_mult_nth_1[of a] by (simp add: fps_power_zeroth[symmetric] mult.assoc) thus ?case by (simp add: algebra_simps) qed simp with Suc show ?thesis by simp qed simp lemma fps_power_first_eq: "a $ 0 = 1 \ a^n $ 1 = of_nat n * a$1" proof (induct n) case (Suc n) show ?case unfolding power_Suc fps_mult_nth using Suc.hyps[OF \a$0 = 1\] \a$0 = 1\ fps_power_zeroth_eq_one[OF \a$0=1\] by (simp add: algebra_simps) qed simp lemma fps_power_first_eq': assumes "a $ 1 = 1" shows "a^n $ 1 = of_nat n * (a$0)^(n-1)" proof (cases n) case (Suc m) from assms have "(a ^ Suc m) $ 1 = of_nat (Suc m) * (a$0)^(Suc m - 1)" using fps_mult_nth_1[of a] by (induct m) (simp_all add: algebra_simps mult_of_nat_commute fps_power_zeroth) with Suc show ?thesis by simp qed simp lemmas startsby_one_power = fps_power_zeroth_eq_one lemma startsby_zero_power: "a $ 0 = 0 \ n > 0 \ a^n $0 = 0" by (simp add: fps_power_zeroth zero_power) lemma startsby_power: "a $0 = v \ a^n $0 = v^n" by (simp add: fps_power_zeroth) lemma startsby_nonzero_power: fixes a :: "'a::semiring_1_no_zero_divisors fps" shows "a $ 0 \ 0 \ a^n $ 0 \ 0" by (simp add: startsby_power) lemma startsby_zero_power_iff[simp]: "a^n $0 = (0::'a::semiring_1_no_zero_divisors) \ n \ 0 \ a$0 = 0" proof show "a ^ n $ 0 = 0 \ n \ 0 \ a $ 0 = 0" proof assume a: "a^n $ 0 = 0" thus "a $ 0 = 0" using startsby_nonzero_power by auto have "n = 0 \ a^n $ 0 = 1" by simp with a show "n \ 0" by fastforce qed show "n \ 0 \ a $ 0 = 0 \ a ^ n $ 0 = 0" by (cases n) auto qed lemma startsby_zero_power_prefix: assumes a0: "a $ 0 = 0" shows "\n < k. a ^ k $ n = 0" proof (induct k rule: nat_less_induct, clarify) case (1 k) fix j :: nat assume j: "j < k" show "a ^ k $ j = 0" proof (cases k) case 0 with j show ?thesis by simp next case (Suc i) with 1 j have "\m\{0<..j}. a ^ i $ (j - m) = 0" by auto with Suc a0 show ?thesis by (simp add: fps_mult_nth sum.atLeast_Suc_atMost) qed qed lemma startsby_zero_sum_depends: assumes a0: "a $0 = 0" and kn: "n \ k" shows "sum (\i. (a ^ i)$k) {0 .. n} = sum (\i. (a ^ i)$k) {0 .. k}" apply (rule sum.mono_neutral_right) using kn apply auto apply (rule startsby_zero_power_prefix[rule_format, OF a0]) apply arith done lemma startsby_zero_power_nth_same: assumes a0: "a$0 = 0" shows "a^n $ n = (a$1) ^ n" proof (induct n) case (Suc n) have "\i\{Suc 1..Suc n}. a ^ n $ (Suc n - i) = 0" using a0 startsby_zero_power_prefix[of a n] by auto thus ?case using a0 Suc sum.atLeast_Suc_atMost[of 0 "Suc n" "\i. a $ i * a ^ n $ (Suc n - i)"] sum.atLeast_Suc_atMost[of 1 "Suc n" "\i. a $ i * a ^ n $ (Suc n - i)"] by (simp add: fps_mult_nth) qed simp lemma fps_lr_inverse_power: fixes a :: "'a::ring_1 fps" assumes "x * a$0 = 1" "a$0 * x = 1" shows "fps_left_inverse (a^n) (x^n) = fps_left_inverse a x ^ n" and "fps_right_inverse (a^n) (x^n) = fps_right_inverse a x ^ n" proof- from assms have xn: "\n. x^n * (a^n $ 0) = 1" "\n. (a^n $ 0) * x^n = 1" by (simp_all add: left_right_inverse_power fps_power_zeroth) show "fps_left_inverse (a^n) (x^n) = fps_left_inverse a x ^ n" proof (induct n) case 0 then show ?case by (simp add: fps_lr_inverse_one_one(1)) next case (Suc n) with assms show ?case using xn fps_lr_inverse_mult_ring1(1)[of x a "x^n" "a^n"] by (simp add: power_Suc2[symmetric]) qed moreover have "fps_right_inverse (a^n) (x^n) = fps_left_inverse (a^n) (x^n)" using xn by (intro fps_left_inverse_eq_fps_right_inverse[symmetric]) moreover have "fps_right_inverse a x = fps_left_inverse a x" using assms by (intro fps_left_inverse_eq_fps_right_inverse[symmetric]) ultimately show "fps_right_inverse (a^n) (x^n) = fps_right_inverse a x ^ n" by simp qed lemma fps_inverse_power: fixes a :: "'a::division_ring fps" shows "inverse (a^n) = inverse a ^ n" proof (cases "n=0" "a$0 = 0" rule: case_split[case_product case_split]) case False_True hence LHS: "inverse (a^n) = 0" and RHS: "inverse a ^ n = 0" by (simp_all add: startsby_zero_power) show ?thesis using trans_sym[OF LHS RHS] by fast next case False_False from False_False(2) show ?thesis by (simp add: fps_inverse_def fps_power_zeroth power_inverse fps_lr_inverse_power(2)[symmetric] ) qed auto lemma fps_deriv_power': fixes a :: "'a::comm_semiring_1 fps" shows "fps_deriv (a ^ n) = (of_nat n) * fps_deriv a * a ^ (n - 1)" proof (cases n) case (Suc m) moreover have "fps_deriv (a^Suc m) = of_nat (Suc m) * fps_deriv a * a^m" by (induct m) (simp_all add: algebra_simps) ultimately show ?thesis by simp qed simp lemma fps_deriv_power: fixes a :: "'a::comm_semiring_1 fps" shows "fps_deriv (a ^ n) = fps_const (of_nat n) * fps_deriv a * a ^ (n - 1)" by (simp add: fps_deriv_power' fps_of_nat) subsection \Integration\ definition fps_integral :: "'a::{semiring_1,inverse} fps \ 'a \ 'a fps" where "fps_integral a a0 = Abs_fps (\n. if n=0 then a0 else inverse (of_nat n) * a$(n - 1))" abbreviation "fps_integral0 a \ fps_integral a 0" lemma fps_integral_nth_0_Suc [simp]: fixes a :: "'a::{semiring_1,inverse} fps" shows "fps_integral a a0 $ 0 = a0" and "fps_integral a a0 $ Suc n = inverse (of_nat (Suc n)) * a $ n" by (auto simp: fps_integral_def) lemma fps_integral_conv_plus_const: "fps_integral a a0 = fps_integral a 0 + fps_const a0" unfolding fps_integral_def by (intro fps_ext) simp lemma fps_deriv_fps_integral: fixes a :: "'a::{division_ring,ring_char_0} fps" shows "fps_deriv (fps_integral a a0) = a" proof (intro fps_ext) fix n have "(of_nat (Suc n) :: 'a) \ 0" by (rule of_nat_neq_0) hence "of_nat (Suc n) * inverse (of_nat (Suc n) :: 'a) = 1" by simp moreover have "fps_deriv (fps_integral a a0) $ n = of_nat (Suc n) * inverse (of_nat (Suc n)) * a $ n" by (simp add: mult.assoc) ultimately show "fps_deriv (fps_integral a a0) $ n = a $ n" by simp qed lemma fps_integral0_deriv: fixes a :: "'a::{division_ring,ring_char_0} fps" shows "fps_integral0 (fps_deriv a) = a - fps_const (a$0)" proof (intro fps_ext) fix n show "fps_integral0 (fps_deriv a) $ n = (a - fps_const (a$0)) $ n" proof (cases n) case (Suc m) have "(of_nat (Suc m) :: 'a) \ 0" by (rule of_nat_neq_0) hence "inverse (of_nat (Suc m) :: 'a) * of_nat (Suc m) = 1" by simp moreover have "fps_integral0 (fps_deriv a) $ Suc m = inverse (of_nat (Suc m)) * of_nat (Suc m) * a $ (Suc m)" by (simp add: mult.assoc) ultimately show ?thesis using Suc by simp qed simp qed lemma fps_integral_deriv: fixes a :: "'a::{division_ring,ring_char_0} fps" shows "fps_integral (fps_deriv a) (a$0) = a" using fps_integral_conv_plus_const[of "fps_deriv a" "a$0"] by (simp add: fps_integral0_deriv) lemma fps_integral0_zero: "fps_integral0 (0::'a::{semiring_1,inverse} fps) = 0" by (intro fps_ext) (simp add: fps_integral_def) lemma fps_integral0_fps_const': fixes c :: "'a::{semiring_1,inverse}" assumes "inverse (1::'a) = 1" shows "fps_integral0 (fps_const c) = fps_const c * fps_X" proof (intro fps_ext) fix n show "fps_integral0 (fps_const c) $ n = (fps_const c * fps_X) $ n" by (cases n) (simp_all add: assms mult_delta_right) qed lemma fps_integral0_fps_const: fixes c :: "'a::division_ring" shows "fps_integral0 (fps_const c) = fps_const c * fps_X" by (rule fps_integral0_fps_const'[OF inverse_1]) lemma fps_integral0_one': assumes "inverse (1::'a::{semiring_1,inverse}) = 1" shows "fps_integral0 (1::'a fps) = fps_X" using assms fps_integral0_fps_const'[of "1::'a"] by simp lemma fps_integral0_one: "fps_integral0 (1::'a::division_ring fps) = fps_X" by (rule fps_integral0_one'[OF inverse_1]) lemma fps_integral0_fps_const_mult_left: fixes a :: "'a::division_ring fps" shows "fps_integral0 (fps_const c * a) = fps_const c * fps_integral0 a" proof (intro fps_ext) fix n show "fps_integral0 (fps_const c * a) $ n = (fps_const c * fps_integral0 a) $ n" using mult_inverse_of_nat_commute[of n c, symmetric] mult.assoc[of "inverse (of_nat n)" c "a$(n-1)"] mult.assoc[of c "inverse (of_nat n)" "a$(n-1)"] by (simp add: fps_integral_def) qed lemma fps_integral0_fps_const_mult_right: fixes a :: "'a::{semiring_1,inverse} fps" shows "fps_integral0 (a * fps_const c) = fps_integral0 a * fps_const c" by (intro fps_ext) (simp add: fps_integral_def algebra_simps) lemma fps_integral0_neg: fixes a :: "'a::{ring_1,inverse} fps" shows "fps_integral0 (-a) = - fps_integral0 a" using fps_integral0_fps_const_mult_right[of a "-1"] by (simp add: fps_const_neg[symmetric]) lemma fps_integral0_add: "fps_integral0 (a+b) = fps_integral0 a + fps_integral0 b" by (intro fps_ext) (simp add: fps_integral_def algebra_simps) lemma fps_integral0_linear: fixes a b :: "'a::division_ring" shows "fps_integral0 (fps_const a * f + fps_const b * g) = fps_const a * fps_integral0 f + fps_const b * fps_integral0 g" by (simp add: fps_integral0_add fps_integral0_fps_const_mult_left) lemma fps_integral0_linear2: "fps_integral0 (f * fps_const a + g * fps_const b) = fps_integral0 f * fps_const a + fps_integral0 g * fps_const b" by (simp add: fps_integral0_add fps_integral0_fps_const_mult_right) lemma fps_integral_linear: fixes a b a0 b0 :: "'a::division_ring" shows "fps_integral (fps_const a * f + fps_const b * g) (a*a0 + b*b0) = fps_const a * fps_integral f a0 + fps_const b * fps_integral g b0" using fps_integral_conv_plus_const[of "fps_const a * f + fps_const b * g" "a*a0 + b*b0" ] fps_integral_conv_plus_const[of f a0] fps_integral_conv_plus_const[of g b0] by (simp add: fps_integral0_linear algebra_simps) lemma fps_integral0_sub: fixes a b :: "'a::{ring_1,inverse} fps" shows "fps_integral0 (a-b) = fps_integral0 a - fps_integral0 b" using fps_integral0_linear2[of a 1 b "-1"] by (simp add: fps_const_neg[symmetric]) lemma fps_integral0_of_nat: "fps_integral0 (of_nat n :: 'a::division_ring fps) = of_nat n * fps_X" using fps_integral0_fps_const[of "of_nat n :: 'a"] by (simp add: fps_of_nat) lemma fps_integral0_sum: "fps_integral0 (sum f S) = sum (\i. fps_integral0 (f i)) S" proof (cases "finite S") case True show ?thesis by (induct rule: finite_induct [OF True]) (simp_all add: fps_integral0_zero fps_integral0_add) qed (simp add: fps_integral0_zero) lemma fps_integral0_by_parts: fixes a b :: "'a::{division_ring,ring_char_0} fps" shows "fps_integral0 (a * b) = a * fps_integral0 b - fps_integral0 (fps_deriv a * fps_integral0 b)" proof- have "fps_integral0 (fps_deriv (a * fps_integral0 b)) = a * fps_integral0 b" using fps_integral0_deriv[of "(a * fps_integral0 b)"] by simp moreover have "fps_integral0 (a * b) = fps_integral0 (fps_deriv (a * fps_integral0 b)) - fps_integral0 (fps_deriv a * fps_integral0 b)" by (auto simp: fps_deriv_fps_integral fps_integral0_sub[symmetric]) ultimately show ?thesis by simp qed lemma fps_integral0_fps_X: "fps_integral0 (fps_X::'a::{semiring_1,inverse} fps) = fps_const (inverse (of_nat 2)) * fps_X\<^sup>2" by (intro fps_ext) (auto simp: fps_integral_def) lemma fps_integral0_fps_X_power: "fps_integral0 ((fps_X::'a::{semiring_1,inverse} fps) ^ n) = fps_const (inverse (of_nat (Suc n))) * fps_X ^ Suc n" proof (intro fps_ext) fix k show "fps_integral0 ((fps_X::'a fps) ^ n) $ k = (fps_const (inverse (of_nat (Suc n))) * fps_X ^ Suc n) $ k" by (cases k) simp_all qed subsection \Composition of FPSs\ definition fps_compose :: "'a::semiring_1 fps \ 'a fps \ 'a fps" (infixl "oo" 55) where "a oo b = Abs_fps (\n. sum (\i. a$i * (b^i$n)) {0..n})" lemma fps_compose_nth: "(a oo b)$n = sum (\i. a$i * (b^i$n)) {0..n}" by (simp add: fps_compose_def) lemma fps_compose_nth_0 [simp]: "(f oo g) $ 0 = f $ 0" by (simp add: fps_compose_nth) lemma fps_compose_fps_X[simp]: "a oo fps_X = (a :: 'a::comm_ring_1 fps)" by (simp add: fps_ext fps_compose_def mult_delta_right) lemma fps_const_compose[simp]: "fps_const (a::'a::comm_ring_1) oo b = fps_const a" by (simp add: fps_eq_iff fps_compose_nth mult_delta_left) lemma numeral_compose[simp]: "(numeral k :: 'a::comm_ring_1 fps) oo b = numeral k" unfolding numeral_fps_const by simp lemma neg_numeral_compose[simp]: "(- numeral k :: 'a::comm_ring_1 fps) oo b = - numeral k" unfolding neg_numeral_fps_const by simp lemma fps_X_fps_compose_startby0[simp]: "a$0 = 0 \ fps_X oo a = (a :: 'a::comm_ring_1 fps)" by (simp add: fps_eq_iff fps_compose_def mult_delta_left not_le) subsection \Rules from Herbert Wilf's Generatingfunctionology\ subsubsection \Rule 1\ (* {a_{n+k}}_0^infty Corresponds to (f - sum (\i. a_i * x^i))/x^h, for h>0*) lemma fps_power_mult_eq_shift: "fps_X^Suc k * Abs_fps (\n. a (n + Suc k)) = Abs_fps a - sum (\i. fps_const (a i :: 'a::comm_ring_1) * fps_X^i) {0 .. k}" (is "?lhs = ?rhs") proof - have "?lhs $ n = ?rhs $ n" for n :: nat proof - have "?lhs $ n = (if n < Suc k then 0 else a n)" unfolding fps_X_power_mult_nth by auto also have "\ = ?rhs $ n" proof (induct k) case 0 then show ?case by (simp add: fps_sum_nth) next case (Suc k) have "(Abs_fps a - sum (\i. fps_const (a i :: 'a) * fps_X^i) {0 .. Suc k})$n = (Abs_fps a - sum (\i. fps_const (a i :: 'a) * fps_X^i) {0 .. k} - fps_const (a (Suc k)) * fps_X^ Suc k) $ n" by (simp add: field_simps) also have "\ = (if n < Suc k then 0 else a n) - (fps_const (a (Suc k)) * fps_X^ Suc k)$n" using Suc.hyps[symmetric] unfolding fps_sub_nth by simp also have "\ = (if n < Suc (Suc k) then 0 else a n)" unfolding fps_X_power_mult_right_nth apply (auto simp add: not_less fps_const_def) apply (rule cong[of a a, OF refl]) apply arith done finally show ?case by simp qed finally show ?thesis . qed then show ?thesis by (simp add: fps_eq_iff) qed subsubsection \Rule 2\ (* We can not reach the form of Wilf, but still near to it using rewrite rules*) (* If f reprents {a_n} and P is a polynomial, then P(xD) f represents {P(n) a_n}*) definition "fps_XD = (*) fps_X \ fps_deriv" lemma fps_XD_add[simp]:"fps_XD (a + b) = fps_XD a + fps_XD (b :: 'a::comm_ring_1 fps)" by (simp add: fps_XD_def field_simps) lemma fps_XD_mult_const[simp]:"fps_XD (fps_const (c::'a::comm_ring_1) * a) = fps_const c * fps_XD a" by (simp add: fps_XD_def field_simps) lemma fps_XD_linear[simp]: "fps_XD (fps_const c * a + fps_const d * b) = fps_const c * fps_XD a + fps_const d * fps_XD (b :: 'a::comm_ring_1 fps)" by simp lemma fps_XDN_linear: "(fps_XD ^^ n) (fps_const c * a + fps_const d * b) = fps_const c * (fps_XD ^^ n) a + fps_const d * (fps_XD ^^ n) (b :: 'a::comm_ring_1 fps)" by (induct n) simp_all lemma fps_mult_fps_X_deriv_shift: "fps_X* fps_deriv a = Abs_fps (\n. of_nat n* a$n)" by (simp add: fps_eq_iff) lemma fps_mult_fps_XD_shift: "(fps_XD ^^ k) (a :: 'a::comm_ring_1 fps) = Abs_fps (\n. (of_nat n ^ k) * a$n)" by (induct k arbitrary: a) (simp_all add: fps_XD_def fps_eq_iff field_simps del: One_nat_def) subsubsection \Rule 3\ text \Rule 3 is trivial and is given by \fps_times_def\.\ subsubsection \Rule 5 --- summation and "division" by (1 - fps_X)\ lemma fps_divide_fps_X_minus1_sum_lemma: "a = ((1::'a::ring_1 fps) - fps_X) * Abs_fps (\n. sum (\i. a $ i) {0..n})" proof (rule fps_ext) define f g :: "'a fps" where "f \ 1 - fps_X" and "g \ Abs_fps (\n. sum (\i. a $ i) {0..n})" fix n show "a $ n= (f * g) $ n" proof (cases n) case (Suc m) hence "(f * g) $ n = g $ Suc m - g $ m" using fps_mult_nth[of f g "Suc m"] sum.atLeast_Suc_atMost[of 0 "Suc m" "\i. f $ i * g $ (Suc m - i)"] sum.atLeast_Suc_atMost[of 1 "Suc m" "\i. f $ i * g $ (Suc m - i)"] by (simp add: f_def) with Suc show ?thesis by (simp add: g_def) qed (simp add: f_def g_def) qed lemma fps_divide_fps_X_minus1_sum_ring1: assumes "inverse 1 = (1::'a::{ring_1,inverse})" shows "a /((1::'a fps) - fps_X) = Abs_fps (\n. sum (\i. a $ i) {0..n})" proof- from assms have "a /((1::'a fps) - fps_X) = a * Abs_fps (\n. 1)" by (simp add: fps_divide_def fps_inverse_def fps_lr_inverse_one_minus_fps_X(2)) thus ?thesis by (auto intro: fps_ext simp: fps_mult_nth) qed lemma fps_divide_fps_X_minus1_sum: "a /((1::'a::division_ring fps) - fps_X) = Abs_fps (\n. sum (\i. a $ i) {0..n})" using fps_divide_fps_X_minus1_sum_ring1[of a] by simp subsubsection \Rule 4 in its more general form: generalizes Rule 3 for an arbitrary finite product of FPS, also the relvant instance of powers of a FPS\ definition "natpermute n k = {l :: nat list. length l = k \ sum_list l = n}" lemma natlist_trivial_1: "natpermute n 1 = {[n]}" apply (auto simp add: natpermute_def) apply (case_tac x) apply auto done lemma append_natpermute_less_eq: assumes "xs @ ys \ natpermute n k" shows "sum_list xs \ n" and "sum_list ys \ n" proof - from assms have "sum_list (xs @ ys) = n" by (simp add: natpermute_def) then have "sum_list xs + sum_list ys = n" by simp then show "sum_list xs \ n" and "sum_list ys \ n" by simp_all qed lemma natpermute_split: assumes "h \ k" shows "natpermute n k = (\m \{0..n}. {l1 @ l2 |l1 l2. l1 \ natpermute m h \ l2 \ natpermute (n - m) (k - h)})" (is "?L = ?R" is "_ = (\m \{0..n}. ?S m)") proof show "?R \ ?L" proof fix l assume l: "l \ ?R" from l obtain m xs ys where h: "m \ {0..n}" and xs: "xs \ natpermute m h" and ys: "ys \ natpermute (n - m) (k - h)" and leq: "l = xs@ys" by blast from xs have xs': "sum_list xs = m" by (simp add: natpermute_def) from ys have ys': "sum_list ys = n - m" by (simp add: natpermute_def) show "l \ ?L" using leq xs ys h apply (clarsimp simp add: natpermute_def) unfolding xs' ys' using assms xs ys unfolding natpermute_def apply simp done qed show "?L \ ?R" proof fix l assume l: "l \ natpermute n k" let ?xs = "take h l" let ?ys = "drop h l" let ?m = "sum_list ?xs" from l have ls: "sum_list (?xs @ ?ys) = n" by (simp add: natpermute_def) have xs: "?xs \ natpermute ?m h" using l assms by (simp add: natpermute_def) have l_take_drop: "sum_list l = sum_list (take h l @ drop h l)" by simp then have ys: "?ys \ natpermute (n - ?m) (k - h)" using l assms ls by (auto simp add: natpermute_def simp del: append_take_drop_id) from ls have m: "?m \ {0..n}" by (simp add: l_take_drop del: append_take_drop_id) from xs ys ls show "l \ ?R" apply auto apply (rule bexI [where x = "?m"]) apply (rule exI [where x = "?xs"]) apply (rule exI [where x = "?ys"]) using ls l apply (auto simp add: natpermute_def l_take_drop simp del: append_take_drop_id) apply simp done qed qed lemma natpermute_0: "natpermute n 0 = (if n = 0 then {[]} else {})" by (auto simp add: natpermute_def) lemma natpermute_0'[simp]: "natpermute 0 k = (if k = 0 then {[]} else {replicate k 0})" apply (auto simp add: set_replicate_conv_if natpermute_def) apply (rule nth_equalityI) apply simp_all done lemma natpermute_finite: "finite (natpermute n k)" proof (induct k arbitrary: n) case 0 then show ?case apply (subst natpermute_split[of 0 0, simplified]) apply (simp add: natpermute_0) done next case (Suc k) then show ?case unfolding natpermute_split [of k "Suc k", simplified] apply - apply (rule finite_UN_I) apply simp unfolding One_nat_def[symmetric] natlist_trivial_1 apply simp done qed lemma natpermute_contain_maximal: "{xs \ natpermute n (k + 1). n \ set xs} = (\i\{0 .. k}. {(replicate (k + 1) 0) [i:=n]})" (is "?A = ?B") proof show "?A \ ?B" proof fix xs assume "xs \ ?A" then have H: "xs \ natpermute n (k + 1)" and n: "n \ set xs" by blast+ then obtain i where i: "i \ {0.. k}" "xs!i = n" unfolding in_set_conv_nth by (auto simp add: less_Suc_eq_le natpermute_def) have eqs: "({0..k} - {i}) \ {i} = {0..k}" using i by auto have f: "finite({0..k} - {i})" "finite {i}" by auto have d: "({0..k} - {i}) \ {i} = {}" using i by auto from H have "n = sum (nth xs) {0..k}" apply (simp add: natpermute_def) apply (auto simp add: atLeastLessThanSuc_atLeastAtMost sum_list_sum_nth) done also have "\ = n + sum (nth xs) ({0..k} - {i})" unfolding sum.union_disjoint[OF f d, unfolded eqs] using i by simp finally have zxs: "\ j\ {0..k} - {i}. xs!j = 0" by auto from H have xsl: "length xs = k+1" by (simp add: natpermute_def) from i have i': "i < length (replicate (k+1) 0)" "i < k+1" unfolding length_replicate by presburger+ have "xs = (replicate (k+1) 0) [i := n]" proof (rule nth_equalityI) show "length xs = length ((replicate (k + 1) 0)[i := n])" by (metis length_list_update length_replicate xsl) show "xs ! j = (replicate (k + 1) 0)[i := n] ! j" if "j < length xs" for j proof (cases "j = i") case True then show ?thesis by (metis i'(1) i(2) nth_list_update) next case False with that show ?thesis by (simp add: xsl zxs del: replicate.simps split: nat.split) qed qed then show "xs \ ?B" using i by blast qed show "?B \ ?A" proof fix xs assume "xs \ ?B" then obtain i where i: "i \ {0..k}" and xs: "xs = (replicate (k + 1) 0) [i:=n]" by auto have nxs: "n \ set xs" unfolding xs apply (rule set_update_memI) using i apply simp done have xsl: "length xs = k + 1" by (simp only: xs length_replicate length_list_update) have "sum_list xs = sum (nth xs) {0.. = sum (\j. if j = i then n else 0) {0..< k+1}" by (rule sum.cong) (simp_all add: xs del: replicate.simps) also have "\ = n" using i by simp finally have "xs \ natpermute n (k + 1)" using xsl unfolding natpermute_def mem_Collect_eq by blast then show "xs \ ?A" using nxs by blast qed qed text \The general form.\ lemma fps_prod_nth: fixes m :: nat and a :: "nat \ 'a::comm_ring_1 fps" shows "(prod a {0 .. m}) $ n = sum (\v. prod (\j. (a j) $ (v!j)) {0..m}) (natpermute n (m+1))" (is "?P m n") proof (induct m arbitrary: n rule: nat_less_induct) fix m n assume H: "\m' < m. \n. ?P m' n" show "?P m n" proof (cases m) case 0 then show ?thesis apply simp unfolding natlist_trivial_1[where n = n, unfolded One_nat_def] apply simp done next case (Suc k) then have km: "k < m" by arith have u0: "{0 .. k} \ {m} = {0..m}" using Suc by (simp add: set_eq_iff) presburger have f0: "finite {0 .. k}" "finite {m}" by auto have d0: "{0 .. k} \ {m} = {}" using Suc by auto have "(prod a {0 .. m}) $ n = (prod a {0 .. k} * a m) $ n" unfolding prod.union_disjoint[OF f0 d0, unfolded u0] by simp also have "\ = (\i = 0..n. (\v\natpermute i (k + 1). \j\{0..k}. a j $ v ! j) * a m $ (n - i))" unfolding fps_mult_nth H[rule_format, OF km] .. also have "\ = (\v\natpermute n (m + 1). \j\{0..m}. a j $ v ! j)" apply (simp add: Suc) unfolding natpermute_split[of m "m + 1", simplified, of n, unfolded natlist_trivial_1[unfolded One_nat_def] Suc] apply (subst sum.UNION_disjoint) apply simp apply simp unfolding image_Collect[symmetric] apply clarsimp apply (rule finite_imageI) apply (rule natpermute_finite) apply (clarsimp simp add: set_eq_iff) apply auto apply (rule sum.cong) apply (rule refl) unfolding sum_distrib_right apply (rule sym) apply (rule_tac l = "\xs. xs @ [n - x]" in sum.reindex_cong) apply (simp add: inj_on_def) apply auto unfolding prod.union_disjoint[OF f0 d0, unfolded u0, unfolded Suc] apply (clarsimp simp add: natpermute_def nth_append) done finally show ?thesis . qed qed text \The special form for powers.\ lemma fps_power_nth_Suc: fixes m :: nat and a :: "'a::comm_ring_1 fps" shows "(a ^ Suc m)$n = sum (\v. prod (\j. a $ (v!j)) {0..m}) (natpermute n (m+1))" proof - have th0: "a^Suc m = prod (\i. a) {0..m}" by (simp add: prod_constant) show ?thesis unfolding th0 fps_prod_nth .. qed lemma fps_power_nth: fixes m :: nat and a :: "'a::comm_ring_1 fps" shows "(a ^m)$n = (if m=0 then 1$n else sum (\v. prod (\j. a $ (v!j)) {0..m - 1}) (natpermute n m))" by (cases m) (simp_all add: fps_power_nth_Suc del: power_Suc) lemmas fps_nth_power_0 = fps_power_zeroth lemma natpermute_max_card: assumes n0: "n \ 0" shows "card {xs \ natpermute n (k + 1). n \ set xs} = k + 1" unfolding natpermute_contain_maximal proof - let ?A = "\i. {(replicate (k + 1) 0)[i := n]}" let ?K = "{0 ..k}" have fK: "finite ?K" by simp have fAK: "\i\?K. finite (?A i)" by auto have d: "\i\ ?K. \j\ ?K. i \ j \ {(replicate (k + 1) 0)[i := n]} \ {(replicate (k + 1) 0)[j := n]} = {}" proof clarify fix i j assume i: "i \ ?K" and j: "j \ ?K" and ij: "i \ j" have False if eq: "(replicate (k+1) 0)[i:=n] = (replicate (k+1) 0)[j:= n]" proof - have "(replicate (k+1) 0) [i:=n] ! i = n" using i by (simp del: replicate.simps) moreover have "(replicate (k+1) 0) [j:=n] ! i = 0" using i ij by (simp del: replicate.simps) ultimately show ?thesis using eq n0 by (simp del: replicate.simps) qed then show "{(replicate (k + 1) 0)[i := n]} \ {(replicate (k + 1) 0)[j := n]} = {}" by auto qed from card_UN_disjoint[OF fK fAK d] show "card (\i\{0..k}. {(replicate (k + 1) 0)[i := n]}) = k + 1" by simp qed lemma fps_power_Suc_nth: fixes f :: "'a :: comm_ring_1 fps" assumes k: "k > 0" shows "(f ^ Suc m) $ k = of_nat (Suc m) * (f $ k * (f $ 0) ^ m) + (\v\{v\natpermute k (m+1). k \ set v}. \j = 0..m. f $ v ! j)" proof - define A B where "A = {v\natpermute k (m+1). k \ set v}" and "B = {v\natpermute k (m+1). k \ set v}" have [simp]: "finite A" "finite B" "A \ B = {}" by (auto simp: A_def B_def natpermute_finite) from natpermute_max_card[of k m] k have card_A: "card A = m + 1" by (simp add: A_def) { fix v assume v: "v \ A" from v have [simp]: "length v = Suc m" by (simp add: A_def natpermute_def) from v have "\j. j \ m \ v ! j = k" by (auto simp: set_conv_nth A_def natpermute_def less_Suc_eq_le) then guess j by (elim exE conjE) note j = this from v have "k = sum_list v" by (simp add: A_def natpermute_def) also have "\ = (\i=0..m. v ! i)" by (simp add: sum_list_sum_nth atLeastLessThanSuc_atLeastAtMost del: sum.op_ivl_Suc) also from j have "{0..m} = insert j ({0..m}-{j})" by auto also from j have "(\i\\. v ! i) = k + (\i\{0..m}-{j}. v ! i)" by (subst sum.insert) simp_all finally have "(\i\{0..m}-{j}. v ! i) = 0" by simp hence zero: "v ! i = 0" if "i \ {0..m}-{j}" for i using that by (subst (asm) sum_eq_0_iff) auto from j have "{0..m} = insert j ({0..m} - {j})" by auto also from j have "(\i\\. f $ (v ! i)) = f $ k * (\i\{0..m} - {j}. f $ (v ! i))" by (subst prod.insert) auto also have "(\i\{0..m} - {j}. f $ (v ! i)) = (\i\{0..m} - {j}. f $ 0)" by (intro prod.cong) (simp_all add: zero) also from j have "\ = (f $ 0) ^ m" by (subst prod_constant) simp_all finally have "(\j = 0..m. f $ (v ! j)) = f $ k * (f $ 0) ^ m" . } note A = this have "(f ^ Suc m) $ k = (\v\natpermute k (m + 1). \j = 0..m. f $ v ! j)" by (rule fps_power_nth_Suc) also have "natpermute k (m+1) = A \ B" unfolding A_def B_def by blast also have "(\v\\. \j = 0..m. f $ (v ! j)) = (\v\A. \j = 0..m. f $ (v ! j)) + (\v\B. \j = 0..m. f $ (v ! j))" by (intro sum.union_disjoint) simp_all also have "(\v\A. \j = 0..m. f $ (v ! j)) = of_nat (Suc m) * (f $ k * (f $ 0) ^ m)" by (simp add: A card_A) finally show ?thesis by (simp add: B_def) qed lemma fps_power_Suc_eqD: fixes f g :: "'a :: {idom,semiring_char_0} fps" assumes "f ^ Suc m = g ^ Suc m" "f $ 0 = g $ 0" "f $ 0 \ 0" shows "f = g" proof (rule fps_ext) fix k :: nat show "f $ k = g $ k" proof (induction k rule: less_induct) case (less k) show ?case proof (cases "k = 0") case False let ?h = "\f. (\v | v \ natpermute k (m + 1) \ k \ set v. \j = 0..m. f $ v ! j)" from False fps_power_Suc_nth[of k f m] fps_power_Suc_nth[of k g m] have "f $ k * (of_nat (Suc m) * (f $ 0) ^ m) + ?h f = g $ k * (of_nat (Suc m) * (f $ 0) ^ m) + ?h g" using assms by (simp add: mult_ac del: power_Suc of_nat_Suc) also have "v ! i < k" if "v \ {v\natpermute k (m+1). k \ set v}" "i \ m" for v i using that elem_le_sum_list[of i v] unfolding natpermute_def by (auto simp: set_conv_nth dest!: spec[of _ i]) hence "?h f = ?h g" by (intro sum.cong refl prod.cong less lessI) (simp add: natpermute_def) finally have "f $ k * (of_nat (Suc m) * (f $ 0) ^ m) = g $ k * (of_nat (Suc m) * (f $ 0) ^ m)" by simp with assms show "f $ k = g $ k" by (subst (asm) mult_right_cancel) (auto simp del: of_nat_Suc) qed (simp_all add: assms) qed qed lemma fps_power_Suc_eqD': fixes f g :: "'a :: {idom,semiring_char_0} fps" assumes "f ^ Suc m = g ^ Suc m" "f $ subdegree f = g $ subdegree g" shows "f = g" proof (cases "f = 0") case False have "Suc m * subdegree f = subdegree (f ^ Suc m)" by (rule subdegree_power [symmetric]) also have "f ^ Suc m = g ^ Suc m" by fact also have "subdegree \ = Suc m * subdegree g" by (rule subdegree_power) finally have [simp]: "subdegree f = subdegree g" by (subst (asm) Suc_mult_cancel1) have "fps_shift (subdegree f) f * fps_X ^ subdegree f = f" by (rule subdegree_decompose [symmetric]) also have "\ ^ Suc m = g ^ Suc m" by fact also have "g = fps_shift (subdegree g) g * fps_X ^ subdegree g" by (rule subdegree_decompose) also have "subdegree f = subdegree g" by fact finally have "fps_shift (subdegree g) f ^ Suc m = fps_shift (subdegree g) g ^ Suc m" by (simp add: algebra_simps power_mult_distrib del: power_Suc) hence "fps_shift (subdegree g) f = fps_shift (subdegree g) g" by (rule fps_power_Suc_eqD) (insert assms False, auto) with subdegree_decompose[of f] subdegree_decompose[of g] show ?thesis by simp qed (insert assms, simp_all) lemma fps_power_eqD': fixes f g :: "'a :: {idom,semiring_char_0} fps" assumes "f ^ m = g ^ m" "f $ subdegree f = g $ subdegree g" "m > 0" shows "f = g" using fps_power_Suc_eqD'[of f "m-1" g] assms by simp lemma fps_power_eqD: fixes f g :: "'a :: {idom,semiring_char_0} fps" assumes "f ^ m = g ^ m" "f $ 0 = g $ 0" "f $ 0 \ 0" "m > 0" shows "f = g" by (rule fps_power_eqD'[of f m g]) (insert assms, simp_all) lemma fps_compose_inj_right: assumes a0: "a$0 = (0::'a::idom)" and a1: "a$1 \ 0" shows "(b oo a = c oo a) \ b = c" (is "?lhs \?rhs") proof show ?lhs if ?rhs using that by simp show ?rhs if ?lhs proof - have "b$n = c$n" for n proof (induct n rule: nat_less_induct) fix n assume H: "\m?lhs\ have "(b oo a)$n = (c oo a)$n" by simp then show ?thesis using 0 by (simp add: fps_compose_nth) next case (Suc n1) have f: "finite {0 .. n1}" "finite {n}" by simp_all have eq: "{0 .. n1} \ {n} = {0 .. n}" using Suc by auto have d: "{0 .. n1} \ {n} = {}" using Suc by auto have seq: "(\i = 0..n1. b $ i * a ^ i $ n) = (\i = 0..n1. c $ i * a ^ i $ n)" apply (rule sum.cong) using H Suc apply auto done have th0: "(b oo a) $n = (\i = 0..n1. c $ i * a ^ i $ n) + b$n * (a$1)^n" unfolding fps_compose_nth sum.union_disjoint[OF f d, unfolded eq] seq using startsby_zero_power_nth_same[OF a0] by simp have th1: "(c oo a) $n = (\i = 0..n1. c $ i * a ^ i $ n) + c$n * (a$1)^n" unfolding fps_compose_nth sum.union_disjoint[OF f d, unfolded eq] using startsby_zero_power_nth_same[OF a0] by simp from \?lhs\[unfolded fps_eq_iff, rule_format, of n] th0 th1 a1 show ?thesis by auto qed qed then show ?rhs by (simp add: fps_eq_iff) qed qed subsection \Radicals\ declare prod.cong [fundef_cong] function radical :: "(nat \ 'a \ 'a) \ nat \ 'a::field fps \ nat \ 'a" where "radical r 0 a 0 = 1" | "radical r 0 a (Suc n) = 0" | "radical r (Suc k) a 0 = r (Suc k) (a$0)" | "radical r (Suc k) a (Suc n) = (a$ Suc n - sum (\xs. prod (\j. radical r (Suc k) a (xs ! j)) {0..k}) {xs. xs \ natpermute (Suc n) (Suc k) \ Suc n \ set xs}) / (of_nat (Suc k) * (radical r (Suc k) a 0)^k)" by pat_completeness auto termination radical proof let ?R = "measure (\(r, k, a, n). n)" { show "wf ?R" by auto next fix r :: "nat \ 'a \ 'a" and a :: "'a fps" and k n xs i assume xs: "xs \ {xs \ natpermute (Suc n) (Suc k). Suc n \ set xs}" and i: "i \ {0..k}" have False if c: "Suc n \ xs ! i" proof - from xs i have "xs !i \ Suc n" by (simp add: in_set_conv_nth natpermute_def) with c have c': "Suc n < xs!i" by arith have fths: "finite {0 ..< i}" "finite {i}" "finite {i+1.. ({i} \ {i+1 ..< Suc k}) = {}" "{i} \ {i+1..< Suc k} = {}" by auto have eqs: "{0.. ({i} \ {i+1 ..< Suc k})" using i by auto from xs have "Suc n = sum_list xs" by (simp add: natpermute_def) also have "\ = sum (nth xs) {0.. = xs!i + sum (nth xs) {0.. ?R" apply auto apply (metis not_less) done next fix r :: "nat \ 'a \ 'a" and a :: "'a fps" and k n show "((r, Suc k, a, 0), r, Suc k, a, Suc n) \ ?R" by simp } qed definition "fps_radical r n a = Abs_fps (radical r n a)" lemma fps_radical0[simp]: "fps_radical r 0 a = 1" apply (auto simp add: fps_eq_iff fps_radical_def) apply (case_tac n) apply auto done lemma fps_radical_nth_0[simp]: "fps_radical r n a $ 0 = (if n = 0 then 1 else r n (a$0))" by (cases n) (simp_all add: fps_radical_def) lemma fps_radical_power_nth[simp]: assumes r: "(r k (a$0)) ^ k = a$0" shows "fps_radical r k a ^ k $ 0 = (if k = 0 then 1 else a$0)" proof (cases k) case 0 then show ?thesis by simp next case (Suc h) have eq1: "fps_radical r k a ^ k $ 0 = (\j\{0..h}. fps_radical r k a $ (replicate k 0) ! j)" unfolding fps_power_nth Suc by simp also have "\ = (\j\{0..h}. r k (a$0))" apply (rule prod.cong) apply simp using Suc apply (subgoal_tac "replicate k 0 ! x = 0") apply (auto intro: nth_replicate simp del: replicate.simps) done also have "\ = a$0" using r Suc by (simp add: prod_constant) finally show ?thesis using Suc by simp qed lemma power_radical: fixes a:: "'a::field_char_0 fps" assumes a0: "a$0 \ 0" shows "(r (Suc k) (a$0)) ^ Suc k = a$0 \ (fps_radical r (Suc k) a) ^ (Suc k) = a" (is "?lhs \ ?rhs") proof let ?r = "fps_radical r (Suc k) a" show ?rhs if r0: ?lhs proof - from a0 r0 have r00: "r (Suc k) (a$0) \ 0" by auto have "?r ^ Suc k $ z = a$z" for z proof (induct z rule: nat_less_induct) fix n assume H: "\m 0" by simp let ?Pnk = "natpermute n (k + 1)" let ?Pnkn = "{xs \ ?Pnk. n \ set xs}" let ?Pnknn = "{xs \ ?Pnk. n \ set xs}" have eq: "?Pnkn \ ?Pnknn = ?Pnk" by blast have d: "?Pnkn \ ?Pnknn = {}" by blast have f: "finite ?Pnkn" "finite ?Pnknn" using finite_Un[of ?Pnkn ?Pnknn, unfolded eq] by (metis natpermute_finite)+ let ?f = "\v. \j\{0..k}. ?r $ v ! j" have "sum ?f ?Pnkn = sum (\v. ?r $ n * r (Suc k) (a $ 0) ^ k) ?Pnkn" proof (rule sum.cong) fix v assume v: "v \ {xs \ natpermute n (k + 1). n \ set xs}" let ?ths = "(\j\{0..k}. fps_radical r (Suc k) a $ v ! j) = fps_radical r (Suc k) a $ n * r (Suc k) (a $ 0) ^ k" from v obtain i where i: "i \ {0..k}" "v = (replicate (k+1) 0) [i:= n]" unfolding natpermute_contain_maximal by auto have "(\j\{0..k}. fps_radical r (Suc k) a $ v ! j) = (\j\{0..k}. if j = i then fps_radical r (Suc k) a $ n else r (Suc k) (a$0))" apply (rule prod.cong, simp) using i r0 apply (simp del: replicate.simps) done also have "\ = (fps_radical r (Suc k) a $ n) * r (Suc k) (a$0) ^ k" using i r0 by (simp add: prod_gen_delta) finally show ?ths . qed rule then have "sum ?f ?Pnkn = of_nat (k+1) * ?r $ n * r (Suc k) (a $ 0) ^ k" by (simp add: natpermute_max_card[OF \n \ 0\, simplified]) also have "\ = a$n - sum ?f ?Pnknn" unfolding Suc using r00 a0 by (simp add: field_simps fps_radical_def del: of_nat_Suc) finally have fn: "sum ?f ?Pnkn = a$n - sum ?f ?Pnknn" . have "(?r ^ Suc k)$n = sum ?f ?Pnkn + sum ?f ?Pnknn" unfolding fps_power_nth_Suc sum.union_disjoint[OF f d, unfolded eq] .. also have "\ = a$n" unfolding fn by simp finally show ?thesis . qed qed then show ?thesis using r0 by (simp add: fps_eq_iff) qed show ?lhs if ?rhs proof - from that have "((fps_radical r (Suc k) a) ^ (Suc k))$0 = a$0" by simp then show ?thesis unfolding fps_power_nth_Suc by (simp add: prod_constant del: replicate.simps) qed qed (* lemma power_radical: fixes a:: "'a::field_char_0 fps" assumes r0: "(r (Suc k) (a$0)) ^ Suc k = a$0" and a0: "a$0 \ 0" shows "(fps_radical r (Suc k) a) ^ (Suc k) = a" proof- let ?r = "fps_radical r (Suc k) a" from a0 r0 have r00: "r (Suc k) (a$0) \ 0" by auto {fix z have "?r ^ Suc k $ z = a$z" proof(induct z rule: nat_less_induct) fix n assume H: "\m 0" using n1 by arith let ?Pnk = "natpermute n (k + 1)" let ?Pnkn = "{xs \ ?Pnk. n \ set xs}" let ?Pnknn = "{xs \ ?Pnk. n \ set xs}" have eq: "?Pnkn \ ?Pnknn = ?Pnk" by blast have d: "?Pnkn \ ?Pnknn = {}" by blast have f: "finite ?Pnkn" "finite ?Pnknn" using finite_Un[of ?Pnkn ?Pnknn, unfolded eq] by (metis natpermute_finite)+ let ?f = "\v. \j\{0..k}. ?r $ v ! j" have "sum ?f ?Pnkn = sum (\v. ?r $ n * r (Suc k) (a $ 0) ^ k) ?Pnkn" proof(rule sum.cong2) fix v assume v: "v \ {xs \ natpermute n (k + 1). n \ set xs}" let ?ths = "(\j\{0..k}. fps_radical r (Suc k) a $ v ! j) = fps_radical r (Suc k) a $ n * r (Suc k) (a $ 0) ^ k" from v obtain i where i: "i \ {0..k}" "v = replicate (k+1) 0 [i:= n]" unfolding natpermute_contain_maximal by auto have "(\j\{0..k}. fps_radical r (Suc k) a $ v ! j) = (\j\{0..k}. if j = i then fps_radical r (Suc k) a $ n else r (Suc k) (a$0))" apply (rule prod.cong, simp) using i r0 by (simp del: replicate.simps) also have "\ = (fps_radical r (Suc k) a $ n) * r (Suc k) (a$0) ^ k" unfolding prod_gen_delta[OF fK] using i r0 by simp finally show ?ths . qed then have "sum ?f ?Pnkn = of_nat (k+1) * ?r $ n * r (Suc k) (a $ 0) ^ k" by (simp add: natpermute_max_card[OF nz, simplified]) also have "\ = a$n - sum ?f ?Pnknn" unfolding n1 using r00 a0 by (simp add: field_simps fps_radical_def del: of_nat_Suc ) finally have fn: "sum ?f ?Pnkn = a$n - sum ?f ?Pnknn" . have "(?r ^ Suc k)$n = sum ?f ?Pnkn + sum ?f ?Pnknn" unfolding fps_power_nth_Suc sum.union_disjoint[OF f d, unfolded eq] .. also have "\ = a$n" unfolding fn by simp finally have "?r ^ Suc k $ n = a $n" .} ultimately show "?r ^ Suc k $ n = a $n" by (cases n, auto) qed } then show ?thesis by (simp add: fps_eq_iff) qed *) lemma eq_divide_imp': fixes c :: "'a::field" shows "c \ 0 \ a * c = b \ a = b / c" by (simp add: field_simps) lemma radical_unique: assumes r0: "(r (Suc k) (b$0)) ^ Suc k = b$0" and a0: "r (Suc k) (b$0 ::'a::field_char_0) = a$0" and b0: "b$0 \ 0" shows "a^(Suc k) = b \ a = fps_radical r (Suc k) b" (is "?lhs \ ?rhs" is "_ \ a = ?r") proof show ?lhs if ?rhs using that using power_radical[OF b0, of r k, unfolded r0] by simp show ?rhs if ?lhs proof - have r00: "r (Suc k) (b$0) \ 0" using b0 r0 by auto have ceq: "card {0..k} = Suc k" by simp from a0 have a0r0: "a$0 = ?r$0" by simp have "a $ n = ?r $ n" for n proof (induct n rule: nat_less_induct) fix n assume h: "\m 0" using Suc by simp let ?Pnk = "natpermute n (Suc k)" let ?Pnkn = "{xs \ ?Pnk. n \ set xs}" let ?Pnknn = "{xs \ ?Pnk. n \ set xs}" have eq: "?Pnkn \ ?Pnknn = ?Pnk" by blast have d: "?Pnkn \ ?Pnknn = {}" by blast have f: "finite ?Pnkn" "finite ?Pnknn" using finite_Un[of ?Pnkn ?Pnknn, unfolded eq] by (metis natpermute_finite)+ let ?f = "\v. \j\{0..k}. ?r $ v ! j" let ?g = "\v. \j\{0..k}. a $ v ! j" have "sum ?g ?Pnkn = sum (\v. a $ n * (?r$0)^k) ?Pnkn" proof (rule sum.cong) fix v assume v: "v \ {xs \ natpermute n (Suc k). n \ set xs}" let ?ths = "(\j\{0..k}. a $ v ! j) = a $ n * (?r$0)^k" from v obtain i where i: "i \ {0..k}" "v = (replicate (k+1) 0) [i:= n]" unfolding Suc_eq_plus1 natpermute_contain_maximal by (auto simp del: replicate.simps) have "(\j\{0..k}. a $ v ! j) = (\j\{0..k}. if j = i then a $ n else r (Suc k) (b$0))" apply (rule prod.cong, simp) using i a0 apply (simp del: replicate.simps) done also have "\ = a $ n * (?r $ 0)^k" using i by (simp add: prod_gen_delta) finally show ?ths . qed rule then have th0: "sum ?g ?Pnkn = of_nat (k+1) * a $ n * (?r $ 0)^k" by (simp add: natpermute_max_card[OF nz, simplified]) have th1: "sum ?g ?Pnknn = sum ?f ?Pnknn" proof (rule sum.cong, rule refl, rule prod.cong, simp) fix xs i assume xs: "xs \ ?Pnknn" and i: "i \ {0..k}" have False if c: "n \ xs ! i" proof - from xs i have "xs ! i \ n" by (simp add: in_set_conv_nth natpermute_def) with c have c': "n < xs!i" by arith have fths: "finite {0 ..< i}" "finite {i}" "finite {i+1.. ({i} \ {i+1 ..< Suc k}) = {}" "{i} \ {i+1..< Suc k} = {}" by auto have eqs: "{0.. ({i} \ {i+1 ..< Suc k})" using i by auto from xs have "n = sum_list xs" by (simp add: natpermute_def) also have "\ = sum (nth xs) {0.. = xs!i + sum (nth xs) {0..x::'a. of_nat (Suc k) * (x * inverse (of_nat (Suc k))) = x" by (simp add: field_simps del: of_nat_Suc) from \?lhs\ have "b$n = a^Suc k $ n" by (simp add: fps_eq_iff) also have "a ^ Suc k$n = sum ?g ?Pnkn + sum ?g ?Pnknn" unfolding fps_power_nth_Suc using sum.union_disjoint[OF f d, unfolded Suc_eq_plus1[symmetric], unfolded eq, of ?g] by simp also have "\ = of_nat (k+1) * a $ n * (?r $ 0)^k + sum ?f ?Pnknn" unfolding th0 th1 .. finally have "of_nat (k+1) * a $ n * (?r $ 0)^k = b$n - sum ?f ?Pnknn" by simp then have "a$n = (b$n - sum ?f ?Pnknn) / (of_nat (k+1) * (?r $ 0)^k)" apply - apply (rule eq_divide_imp') using r00 apply (simp del: of_nat_Suc) apply (simp add: ac_simps) done then show ?thesis apply (simp del: of_nat_Suc) unfolding fps_radical_def Suc apply (simp add: field_simps Suc th00 del: of_nat_Suc) done qed qed then show ?rhs by (simp add: fps_eq_iff) qed qed lemma radical_power: assumes r0: "r (Suc k) ((a$0) ^ Suc k) = a$0" and a0: "(a$0 :: 'a::field_char_0) \ 0" shows "(fps_radical r (Suc k) (a ^ Suc k)) = a" proof - let ?ak = "a^ Suc k" have ak0: "?ak $ 0 = (a$0) ^ Suc k" by (simp add: fps_nth_power_0 del: power_Suc) from r0 have th0: "r (Suc k) (a ^ Suc k $ 0) ^ Suc k = a ^ Suc k $ 0" using ak0 by auto from r0 ak0 have th1: "r (Suc k) (a ^ Suc k $ 0) = a $ 0" by auto from ak0 a0 have ak00: "?ak $ 0 \0 " by auto from radical_unique[of r k ?ak a, OF th0 th1 ak00] show ?thesis by metis qed lemma fps_deriv_radical': fixes a :: "'a::field_char_0 fps" assumes r0: "(r (Suc k) (a$0)) ^ Suc k = a$0" and a0: "a$0 \ 0" shows "fps_deriv (fps_radical r (Suc k) a) = fps_deriv a / ((of_nat (Suc k)) * (fps_radical r (Suc k) a) ^ k)" proof - let ?r = "fps_radical r (Suc k) a" let ?w = "(of_nat (Suc k)) * ?r ^ k" from a0 r0 have r0': "r (Suc k) (a$0) \ 0" by auto from r0' have w0: "?w $ 0 \ 0" by (simp del: of_nat_Suc) note th0 = inverse_mult_eq_1[OF w0] let ?iw = "inverse ?w" from iffD1[OF power_radical[of a r], OF a0 r0] have "fps_deriv (?r ^ Suc k) = fps_deriv a" by simp then have "fps_deriv ?r * ?w = fps_deriv a" by (simp add: fps_deriv_power' ac_simps del: power_Suc) then have "?iw * fps_deriv ?r * ?w = ?iw * fps_deriv a" by simp with a0 r0 have "fps_deriv ?r * (?iw * ?w) = fps_deriv a / ?w" by (subst fps_divide_unit) (auto simp del: of_nat_Suc) then show ?thesis unfolding th0 by simp qed lemma fps_deriv_radical: fixes a :: "'a::field_char_0 fps" assumes r0: "(r (Suc k) (a$0)) ^ Suc k = a$0" and a0: "a$0 \ 0" shows "fps_deriv (fps_radical r (Suc k) a) = fps_deriv a / (fps_const (of_nat (Suc k)) * (fps_radical r (Suc k) a) ^ k)" using fps_deriv_radical'[of r k a, OF r0 a0] by (simp add: fps_of_nat[symmetric]) lemma radical_mult_distrib: fixes a :: "'a::field_char_0 fps" assumes k: "k > 0" and ra0: "r k (a $ 0) ^ k = a $ 0" and rb0: "r k (b $ 0) ^ k = b $ 0" and a0: "a $ 0 \ 0" and b0: "b $ 0 \ 0" shows "r k ((a * b) $ 0) = r k (a $ 0) * r k (b $ 0) \ fps_radical r k (a * b) = fps_radical r k a * fps_radical r k b" (is "?lhs \ ?rhs") proof show ?rhs if r0': ?lhs proof - from r0' have r0: "(r k ((a * b) $ 0)) ^ k = (a * b) $ 0" by (simp add: fps_mult_nth ra0 rb0 power_mult_distrib) show ?thesis proof (cases k) case 0 then show ?thesis using r0' by simp next case (Suc h) let ?ra = "fps_radical r (Suc h) a" let ?rb = "fps_radical r (Suc h) b" have th0: "r (Suc h) ((a * b) $ 0) = (fps_radical r (Suc h) a * fps_radical r (Suc h) b) $ 0" using r0' Suc by (simp add: fps_mult_nth) have ab0: "(a*b) $ 0 \ 0" using a0 b0 by (simp add: fps_mult_nth) from radical_unique[of r h "a*b" "fps_radical r (Suc h) a * fps_radical r (Suc h) b", OF r0[unfolded Suc] th0 ab0, symmetric] iffD1[OF power_radical[of _ r], OF a0 ra0[unfolded Suc]] iffD1[OF power_radical[of _ r], OF b0 rb0[unfolded Suc]] Suc r0' show ?thesis by (auto simp add: power_mult_distrib simp del: power_Suc) qed qed show ?lhs if ?rhs proof - from that have "(fps_radical r k (a * b)) $ 0 = (fps_radical r k a * fps_radical r k b) $ 0" by simp then show ?thesis using k by (simp add: fps_mult_nth) qed qed (* lemma radical_mult_distrib: fixes a:: "'a::field_char_0 fps" assumes ra0: "r k (a $ 0) ^ k = a $ 0" and rb0: "r k (b $ 0) ^ k = b $ 0" and r0': "r k ((a * b) $ 0) = r k (a $ 0) * r k (b $ 0)" and a0: "a$0 \ 0" and b0: "b$0 \ 0" shows "fps_radical r (k) (a*b) = fps_radical r (k) a * fps_radical r (k) (b)" proof- from r0' have r0: "(r (k) ((a*b)$0)) ^ k = (a*b)$0" by (simp add: fps_mult_nth ra0 rb0 power_mult_distrib) {assume "k=0" then have ?thesis by simp} moreover {fix h assume k: "k = Suc h" let ?ra = "fps_radical r (Suc h) a" let ?rb = "fps_radical r (Suc h) b" have th0: "r (Suc h) ((a * b) $ 0) = (fps_radical r (Suc h) a * fps_radical r (Suc h) b) $ 0" using r0' k by (simp add: fps_mult_nth) have ab0: "(a*b) $ 0 \ 0" using a0 b0 by (simp add: fps_mult_nth) from radical_unique[of r h "a*b" "fps_radical r (Suc h) a * fps_radical r (Suc h) b", OF r0[unfolded k] th0 ab0, symmetric] power_radical[of r, OF ra0[unfolded k] a0] power_radical[of r, OF rb0[unfolded k] b0] k have ?thesis by (auto simp add: power_mult_distrib simp del: power_Suc)} ultimately show ?thesis by (cases k, auto) qed *) lemma radical_divide: fixes a :: "'a::field_char_0 fps" assumes kp: "k > 0" and ra0: "(r k (a $ 0)) ^ k = a $ 0" and rb0: "(r k (b $ 0)) ^ k = b $ 0" and a0: "a$0 \ 0" and b0: "b$0 \ 0" shows "r k ((a $ 0) / (b$0)) = r k (a$0) / r k (b $ 0) \ fps_radical r k (a/b) = fps_radical r k a / fps_radical r k b" (is "?lhs = ?rhs") proof let ?r = "fps_radical r k" from kp obtain h where k: "k = Suc h" by (cases k) auto have ra0': "r k (a$0) \ 0" using a0 ra0 k by auto have rb0': "r k (b$0) \ 0" using b0 rb0 k by auto show ?lhs if ?rhs proof - from that have "?r (a/b) $ 0 = (?r a / ?r b)$0" by simp then show ?thesis using k a0 b0 rb0' by (simp add: fps_divide_unit fps_mult_nth fps_inverse_def divide_inverse) qed show ?rhs if ?lhs proof - from a0 b0 have ab0[simp]: "(a/b)$0 = a$0 / b$0" by (simp add: fps_divide_def fps_mult_nth divide_inverse fps_inverse_def) have th0: "r k ((a/b)$0) ^ k = (a/b)$0" by (simp add: \?lhs\ power_divide ra0 rb0) from a0 b0 ra0' rb0' kp \?lhs\ have th1: "r k ((a / b) $ 0) = (fps_radical r k a / fps_radical r k b) $ 0" by (simp add: fps_divide_unit fps_mult_nth fps_inverse_def divide_inverse) from a0 b0 ra0' rb0' kp have ab0': "(a / b) $ 0 \ 0" by (simp add: fps_divide_unit fps_mult_nth fps_inverse_def nonzero_imp_inverse_nonzero) note tha[simp] = iffD1[OF power_radical[where r=r and k=h], OF a0 ra0[unfolded k], unfolded k[symmetric]] note thb[simp] = iffD1[OF power_radical[where r=r and k=h], OF b0 rb0[unfolded k], unfolded k[symmetric]] from b0 rb0' have th2: "(?r a / ?r b)^k = a/b" by (simp add: fps_divide_unit power_mult_distrib fps_inverse_power[symmetric]) from iffD1[OF radical_unique[where r=r and a="?r a / ?r b" and b="a/b" and k=h], symmetric, unfolded k[symmetric], OF th0 th1 ab0' th2] show ?thesis . qed qed lemma radical_inverse: fixes a :: "'a::field_char_0 fps" assumes k: "k > 0" and ra0: "r k (a $ 0) ^ k = a $ 0" and r1: "(r k 1)^k = 1" and a0: "a$0 \ 0" shows "r k (inverse (a $ 0)) = r k 1 / (r k (a $ 0)) \ fps_radical r k (inverse a) = fps_radical r k 1 / fps_radical r k a" using radical_divide[where k=k and r=r and a=1 and b=a, OF k ] ra0 r1 a0 by (simp add: divide_inverse fps_divide_def) subsection \Derivative of composition\ lemma fps_compose_deriv: fixes a :: "'a::idom fps" assumes b0: "b$0 = 0" shows "fps_deriv (a oo b) = ((fps_deriv a) oo b) * fps_deriv b" proof - have "(fps_deriv (a oo b))$n = (((fps_deriv a) oo b) * (fps_deriv b)) $n" for n proof - have "(fps_deriv (a oo b))$n = sum (\i. a $ i * (fps_deriv (b^i))$n) {0.. Suc n}" by (simp add: fps_compose_def field_simps sum_distrib_left del: of_nat_Suc) also have "\ = sum (\i. a$i * ((fps_const (of_nat i)) * (fps_deriv b * (b^(i - 1))))$n) {0.. Suc n}" by (simp add: field_simps fps_deriv_power del: fps_mult_left_const_nth of_nat_Suc) also have "\ = sum (\i. of_nat i * a$i * (((b^(i - 1)) * fps_deriv b))$n) {0.. Suc n}" unfolding fps_mult_left_const_nth by (simp add: field_simps) also have "\ = sum (\i. of_nat i * a$i * (sum (\j. (b^ (i - 1))$j * (fps_deriv b)$(n - j)) {0..n})) {0.. Suc n}" unfolding fps_mult_nth .. also have "\ = sum (\i. of_nat i * a$i * (sum (\j. (b^ (i - 1))$j * (fps_deriv b)$(n - j)) {0..n})) {1.. Suc n}" apply (rule sum.mono_neutral_right) apply (auto simp add: mult_delta_left not_le) done also have "\ = sum (\i. of_nat (i + 1) * a$(i+1) * (sum (\j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}" unfolding fps_deriv_nth by (rule sum.reindex_cong [of Suc]) (simp_all add: mult.assoc) finally have th0: "(fps_deriv (a oo b))$n = sum (\i. of_nat (i + 1) * a$(i+1) * (sum (\j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}" . have "(((fps_deriv a) oo b) * (fps_deriv b))$n = sum (\i. (fps_deriv b)$ (n - i) * ((fps_deriv a) oo b)$i) {0..n}" unfolding fps_mult_nth by (simp add: ac_simps) also have "\ = sum (\i. sum (\j. of_nat (n - i +1) * b$(n - i + 1) * of_nat (j + 1) * a$(j+1) * (b^j)$i) {0..n}) {0..n}" unfolding fps_deriv_nth fps_compose_nth sum_distrib_left mult.assoc apply (rule sum.cong) apply (rule refl) apply (rule sum.mono_neutral_left) apply (simp_all add: subset_eq) apply clarify apply (subgoal_tac "b^i$x = 0") apply simp apply (rule startsby_zero_power_prefix[OF b0, rule_format]) apply simp done also have "\ = sum (\i. of_nat (i + 1) * a$(i+1) * (sum (\j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}" unfolding sum_distrib_left apply (subst sum.swap) apply (rule sum.cong, rule refl)+ apply simp done finally show ?thesis unfolding th0 by simp qed then show ?thesis by (simp add: fps_eq_iff) qed subsection \Finite FPS (i.e. polynomials) and fps_X\ lemma fps_poly_sum_fps_X: assumes "\i > n. a$i = 0" shows "a = sum (\i. fps_const (a$i) * fps_X^i) {0..n}" (is "a = ?r") proof - have "a$i = ?r$i" for i unfolding fps_sum_nth fps_mult_left_const_nth fps_X_power_nth by (simp add: mult_delta_right assms) then show ?thesis unfolding fps_eq_iff by blast qed subsection \Compositional inverses\ fun compinv :: "'a fps \ nat \ 'a::field" where "compinv a 0 = fps_X$0" | "compinv a (Suc n) = (fps_X$ Suc n - sum (\i. (compinv a i) * (a^i)$Suc n) {0 .. n}) / (a$1) ^ Suc n" definition "fps_inv a = Abs_fps (compinv a)" lemma fps_inv: assumes a0: "a$0 = 0" and a1: "a$1 \ 0" shows "fps_inv a oo a = fps_X" proof - let ?i = "fps_inv a oo a" have "?i $n = fps_X$n" for n proof (induct n rule: nat_less_induct) fix n assume h: "\mi. (fps_inv a $ i) * (a^i)$n) {0 .. n1} + fps_inv a $ Suc n1 * (a $ 1)^ Suc n1" by (simp only: fps_compose_nth) (simp add: Suc startsby_zero_power_nth_same [OF a0] del: power_Suc) also have "\ = sum (\i. (fps_inv a $ i) * (a^i)$n) {0 .. n1} + (fps_X$ Suc n1 - sum (\i. (fps_inv a $ i) * (a^i)$n) {0 .. n1})" using a0 a1 Suc by (simp add: fps_inv_def) also have "\ = fps_X$n" using Suc by simp finally show ?thesis . qed qed then show ?thesis by (simp add: fps_eq_iff) qed fun gcompinv :: "'a fps \ 'a fps \ nat \ 'a::field" where "gcompinv b a 0 = b$0" | "gcompinv b a (Suc n) = (b$ Suc n - sum (\i. (gcompinv b a i) * (a^i)$Suc n) {0 .. n}) / (a$1) ^ Suc n" definition "fps_ginv b a = Abs_fps (gcompinv b a)" lemma fps_ginv: assumes a0: "a$0 = 0" and a1: "a$1 \ 0" shows "fps_ginv b a oo a = b" proof - let ?i = "fps_ginv b a oo a" have "?i $n = b$n" for n proof (induct n rule: nat_less_induct) fix n assume h: "\mi. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1} + fps_ginv b a $ Suc n1 * (a $ 1)^ Suc n1" by (simp only: fps_compose_nth) (simp add: Suc startsby_zero_power_nth_same [OF a0] del: power_Suc) also have "\ = sum (\i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1} + (b$ Suc n1 - sum (\i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1})" using a0 a1 Suc by (simp add: fps_ginv_def) also have "\ = b$n" using Suc by simp finally show ?thesis . qed qed then show ?thesis by (simp add: fps_eq_iff) qed lemma fps_inv_ginv: "fps_inv = fps_ginv fps_X" apply (auto simp add: fun_eq_iff fps_eq_iff fps_inv_def fps_ginv_def) apply (induct_tac n rule: nat_less_induct) apply auto apply (case_tac na) apply simp apply simp done lemma fps_compose_1[simp]: "1 oo a = 1" by (simp add: fps_eq_iff fps_compose_nth mult_delta_left) lemma fps_compose_0[simp]: "0 oo a = 0" by (simp add: fps_eq_iff fps_compose_nth) lemma fps_compose_0_right[simp]: "a oo 0 = fps_const (a $ 0)" by (simp add: fps_eq_iff fps_compose_nth power_0_left sum.neutral) lemma fps_compose_add_distrib: "(a + b) oo c = (a oo c) + (b oo c)" by (simp add: fps_eq_iff fps_compose_nth field_simps sum.distrib) lemma fps_compose_sum_distrib: "(sum f S) oo a = sum (\i. f i oo a) S" proof (cases "finite S") case True show ?thesis proof (rule finite_induct[OF True]) show "sum f {} oo a = (\i\{}. f i oo a)" by simp next fix x F assume fF: "finite F" and xF: "x \ F" and h: "sum f F oo a = sum (\i. f i oo a) F" show "sum f (insert x F) oo a = sum (\i. f i oo a) (insert x F)" using fF xF h by (simp add: fps_compose_add_distrib) qed next case False then show ?thesis by simp qed lemma convolution_eq: "sum (\i. a (i :: nat) * b (n - i)) {0 .. n} = sum (\(i,j). a i * b j) {(i,j). i \ n \ j \ n \ i + j = n}" by (rule sum.reindex_bij_witness[where i=fst and j="\i. (i, n - i)"]) auto lemma product_composition_lemma: assumes c0: "c$0 = (0::'a::idom)" and d0: "d$0 = 0" shows "((a oo c) * (b oo d))$n = sum (\(k,m). a$k * b$m * (c^k * d^m) $ n) {(k,m). k + m \ n}" (is "?l = ?r") proof - let ?S = "{(k::nat, m::nat). k + m \ n}" have s: "?S \ {0..n} \ {0..n}" by (simp add: subset_eq) have f: "finite {(k::nat, m::nat). k + m \ n}" apply (rule finite_subset[OF s]) apply auto done have "?r = sum (\i. sum (\(k,m). a$k * (c^k)$i * b$m * (d^m) $ (n - i)) {(k,m). k + m \ n}) {0..n}" apply (simp add: fps_mult_nth sum_distrib_left) apply (subst sum.swap) apply (rule sum.cong) apply (auto simp add: field_simps) done also have "\ = ?l" apply (simp add: fps_mult_nth fps_compose_nth sum_product) apply (rule sum.cong) apply (rule refl) apply (simp add: sum.cartesian_product mult.assoc) apply (rule sum.mono_neutral_right[OF f]) apply (simp add: subset_eq) apply presburger apply clarsimp apply (rule ccontr) apply (clarsimp simp add: not_le) apply (case_tac "x < aa") apply simp apply (frule_tac startsby_zero_power_prefix[rule_format, OF c0]) apply blast apply simp apply (frule_tac startsby_zero_power_prefix[rule_format, OF d0]) apply blast done finally show ?thesis by simp qed lemma product_composition_lemma': assumes c0: "c$0 = (0::'a::idom)" and d0: "d$0 = 0" shows "((a oo c) * (b oo d))$n = sum (\k. sum (\m. a$k * b$m * (c^k * d^m) $ n) {0..n}) {0..n}" (is "?l = ?r") unfolding product_composition_lemma[OF c0 d0] unfolding sum.cartesian_product apply (rule sum.mono_neutral_left) apply simp apply (clarsimp simp add: subset_eq) apply clarsimp apply (rule ccontr) apply (subgoal_tac "(c^aa * d^ba) $ n = 0") apply simp unfolding fps_mult_nth apply (rule sum.neutral) apply (clarsimp simp add: not_le) apply (case_tac "x < aa") apply (rule startsby_zero_power_prefix[OF c0, rule_format]) apply simp apply (subgoal_tac "n - x < ba") apply (frule_tac k = "ba" in startsby_zero_power_prefix[OF d0, rule_format]) apply simp apply arith done lemma sum_pair_less_iff: "sum (\((k::nat),m). a k * b m * c (k + m)) {(k,m). k + m \ n} = sum (\s. sum (\i. a i * b (s - i) * c s) {0..s}) {0..n}" (is "?l = ?r") proof - let ?KM = "{(k,m). k + m \ n}" let ?f = "\s. \i\{0..s}. {(i, s - i)}" have th0: "?KM = \ (?f ` {0..n})" by auto show "?l = ?r " unfolding th0 apply (subst sum.UNION_disjoint) apply auto apply (subst sum.UNION_disjoint) apply auto done qed lemma fps_compose_mult_distrib_lemma: assumes c0: "c$0 = (0::'a::idom)" shows "((a oo c) * (b oo c))$n = sum (\s. sum (\i. a$i * b$(s - i) * (c^s) $ n) {0..s}) {0..n}" unfolding product_composition_lemma[OF c0 c0] power_add[symmetric] unfolding sum_pair_less_iff[where a = "\k. a$k" and b="\m. b$m" and c="\s. (c ^ s)$n" and n = n] .. lemma fps_compose_mult_distrib: assumes c0: "c $ 0 = (0::'a::idom)" shows "(a * b) oo c = (a oo c) * (b oo c)" apply (simp add: fps_eq_iff fps_compose_mult_distrib_lemma [OF c0]) apply (simp add: fps_compose_nth fps_mult_nth sum_distrib_right) done lemma fps_compose_prod_distrib: assumes c0: "c$0 = (0::'a::idom)" shows "prod a S oo c = prod (\k. a k oo c) S" apply (cases "finite S") apply simp_all apply (induct S rule: finite_induct) apply simp apply (simp add: fps_compose_mult_distrib[OF c0]) done lemma fps_compose_divide: assumes [simp]: "g dvd f" "h $ 0 = 0" shows "fps_compose f h = fps_compose (f / g :: 'a :: field fps) h * fps_compose g h" proof - have "f = (f / g) * g" by simp also have "fps_compose \ h = fps_compose (f / g) h * fps_compose g h" by (subst fps_compose_mult_distrib) simp_all finally show ?thesis . qed lemma fps_compose_divide_distrib: assumes "g dvd f" "h $ 0 = 0" "fps_compose g h \ 0" shows "fps_compose (f / g :: 'a :: field fps) h = fps_compose f h / fps_compose g h" using fps_compose_divide[OF assms(1,2)] assms(3) by simp lemma fps_compose_power: assumes c0: "c$0 = (0::'a::idom)" shows "(a oo c)^n = a^n oo c" proof (cases n) case 0 then show ?thesis by simp next case (Suc m) have "(\n = 0..m. a) oo c = (\n = 0..m. a oo c)" using c0 fps_compose_prod_distrib by blast moreover have th0: "a^n = prod (\k. a) {0..m}" "(a oo c) ^ n = prod (\k. a oo c) {0..m}" by (simp_all add: prod_constant Suc) ultimately show ?thesis by presburger qed lemma fps_compose_uminus: "- (a::'a::ring_1 fps) oo c = - (a oo c)" by (simp add: fps_eq_iff fps_compose_nth field_simps sum_negf[symmetric]) lemma fps_compose_sub_distrib: "(a - b) oo (c::'a::ring_1 fps) = (a oo c) - (b oo c)" using fps_compose_add_distrib [of a "- b" c] by (simp add: fps_compose_uminus) lemma fps_X_fps_compose: "fps_X oo a = Abs_fps (\n. if n = 0 then (0::'a::comm_ring_1) else a$n)" by (simp add: fps_eq_iff fps_compose_nth mult_delta_left) lemma fps_inverse_compose: assumes b0: "(b$0 :: 'a::field) = 0" and a0: "a$0 \ 0" shows "inverse a oo b = inverse (a oo b)" proof - let ?ia = "inverse a" let ?ab = "a oo b" let ?iab = "inverse ?ab" from a0 have ia0: "?ia $ 0 \ 0" by simp from a0 have ab0: "?ab $ 0 \ 0" by (simp add: fps_compose_def) have "(?ia oo b) * (a oo b) = 1" unfolding fps_compose_mult_distrib[OF b0, symmetric] unfolding inverse_mult_eq_1[OF a0] fps_compose_1 .. then have "(?ia oo b) * (a oo b) * ?iab = 1 * ?iab" by simp then have "(?ia oo b) * (?iab * (a oo b)) = ?iab" by simp then show ?thesis unfolding inverse_mult_eq_1[OF ab0] by simp qed lemma fps_divide_compose: assumes c0: "(c$0 :: 'a::field) = 0" and b0: "b$0 \ 0" shows "(a/b) oo c = (a oo c) / (b oo c)" using b0 c0 by (simp add: fps_divide_unit fps_inverse_compose fps_compose_mult_distrib) lemma gp: assumes a0: "a$0 = (0::'a::field)" shows "(Abs_fps (\n. 1)) oo a = 1/(1 - a)" (is "?one oo a = _") proof - have o0: "?one $ 0 \ 0" by simp have th0: "(1 - fps_X) $ 0 \ (0::'a)" by simp from fps_inverse_gp[where ?'a = 'a] have "inverse ?one = 1 - fps_X" by (simp add: fps_eq_iff) then have "inverse (inverse ?one) = inverse (1 - fps_X)" by simp then have th: "?one = 1/(1 - fps_X)" unfolding fps_inverse_idempotent[OF o0] by (simp add: fps_divide_def) show ?thesis unfolding th unfolding fps_divide_compose[OF a0 th0] fps_compose_1 fps_compose_sub_distrib fps_X_fps_compose_startby0[OF a0] .. qed lemma fps_compose_radical: assumes b0: "b$0 = (0::'a::field_char_0)" and ra0: "r (Suc k) (a$0) ^ Suc k = a$0" and a0: "a$0 \ 0" shows "fps_radical r (Suc k) a oo b = fps_radical r (Suc k) (a oo b)" proof - let ?r = "fps_radical r (Suc k)" let ?ab = "a oo b" have ab0: "?ab $ 0 = a$0" by (simp add: fps_compose_def) from ab0 a0 ra0 have rab0: "?ab $ 0 \ 0" "r (Suc k) (?ab $ 0) ^ Suc k = ?ab $ 0" by simp_all have th00: "r (Suc k) ((a oo b) $ 0) = (fps_radical r (Suc k) a oo b) $ 0" by (simp add: ab0 fps_compose_def) have th0: "(?r a oo b) ^ (Suc k) = a oo b" unfolding fps_compose_power[OF b0] unfolding iffD1[OF power_radical[of a r k], OF a0 ra0] .. from iffD1[OF radical_unique[where r=r and k=k and b= ?ab and a = "?r a oo b", OF rab0(2) th00 rab0(1)], OF th0] show ?thesis . qed lemma fps_const_mult_apply_left: "fps_const c * (a oo b) = (fps_const c * a) oo b" by (simp add: fps_eq_iff fps_compose_nth sum_distrib_left mult.assoc) lemma fps_const_mult_apply_right: "(a oo b) * fps_const (c::'a::comm_semiring_1) = (fps_const c * a) oo b" by (simp add: fps_const_mult_apply_left mult.commute) lemma fps_compose_assoc: assumes c0: "c$0 = (0::'a::idom)" and b0: "b$0 = 0" shows "a oo (b oo c) = a oo b oo c" (is "?l = ?r") proof - have "?l$n = ?r$n" for n proof - have "?l$n = (sum (\i. (fps_const (a$i) * b^i) oo c) {0..n})$n" by (simp add: fps_compose_nth fps_compose_power[OF c0] fps_const_mult_apply_left sum_distrib_left mult.assoc fps_sum_nth) also have "\ = ((sum (\i. fps_const (a$i) * b^i) {0..n}) oo c)$n" by (simp add: fps_compose_sum_distrib) also have "\ = ?r$n" apply (simp add: fps_compose_nth fps_sum_nth sum_distrib_right mult.assoc) apply (rule sum.cong) apply (rule refl) apply (rule sum.mono_neutral_right) apply (auto simp add: not_le) apply (erule startsby_zero_power_prefix[OF b0, rule_format]) done finally show ?thesis . qed then show ?thesis by (simp add: fps_eq_iff) qed lemma fps_X_power_compose: assumes a0: "a$0=0" shows "fps_X^k oo a = (a::'a::idom fps)^k" (is "?l = ?r") proof (cases k) case 0 then show ?thesis by simp next case (Suc h) have "?l $ n = ?r $n" for n proof - consider "k > n" | "k \ n" by arith then show ?thesis proof cases case 1 then show ?thesis using a0 startsby_zero_power_prefix[OF a0] Suc by (simp add: fps_compose_nth del: power_Suc) next case 2 then show ?thesis by (simp add: fps_compose_nth mult_delta_left) qed qed then show ?thesis unfolding fps_eq_iff by blast qed lemma fps_inv_right: assumes a0: "a$0 = 0" and a1: "a$1 \ 0" shows "a oo fps_inv a = fps_X" proof - let ?ia = "fps_inv a" let ?iaa = "a oo fps_inv a" have th0: "?ia $ 0 = 0" by (simp add: fps_inv_def) have th1: "?iaa $ 0 = 0" using a0 a1 by (simp add: fps_inv_def fps_compose_nth) have th2: "fps_X$0 = 0" by simp from fps_inv[OF a0 a1] have "a oo (fps_inv a oo a) = a oo fps_X" by simp then have "(a oo fps_inv a) oo a = fps_X oo a" by (simp add: fps_compose_assoc[OF a0 th0] fps_X_fps_compose_startby0[OF a0]) with fps_compose_inj_right[OF a0 a1] show ?thesis by simp qed lemma fps_inv_deriv: assumes a0: "a$0 = (0::'a::field)" and a1: "a$1 \ 0" shows "fps_deriv (fps_inv a) = inverse (fps_deriv a oo fps_inv a)" proof - let ?ia = "fps_inv a" let ?d = "fps_deriv a oo ?ia" let ?dia = "fps_deriv ?ia" have ia0: "?ia$0 = 0" by (simp add: fps_inv_def) have th0: "?d$0 \ 0" using a1 by (simp add: fps_compose_nth) from fps_inv_right[OF a0 a1] have "?d * ?dia = 1" by (simp add: fps_compose_deriv[OF ia0, of a, symmetric] ) then have "inverse ?d * ?d * ?dia = inverse ?d * 1" by simp with inverse_mult_eq_1 [OF th0] show "?dia = inverse ?d" by simp qed lemma fps_inv_idempotent: assumes a0: "a$0 = 0" and a1: "a$1 \ 0" shows "fps_inv (fps_inv a) = a" proof - let ?r = "fps_inv" have ra0: "?r a $ 0 = 0" by (simp add: fps_inv_def) from a1 have ra1: "?r a $ 1 \ 0" by (simp add: fps_inv_def field_simps) have fps_X0: "fps_X$0 = 0" by simp from fps_inv[OF ra0 ra1] have "?r (?r a) oo ?r a = fps_X" . then have "?r (?r a) oo ?r a oo a = fps_X oo a" by simp then have "?r (?r a) oo (?r a oo a) = a" unfolding fps_X_fps_compose_startby0[OF a0] unfolding fps_compose_assoc[OF a0 ra0, symmetric] . then show ?thesis unfolding fps_inv[OF a0 a1] by simp qed lemma fps_ginv_ginv: assumes a0: "a$0 = 0" and a1: "a$1 \ 0" and c0: "c$0 = 0" and c1: "c$1 \ 0" shows "fps_ginv b (fps_ginv c a) = b oo a oo fps_inv c" proof - let ?r = "fps_ginv" from c0 have rca0: "?r c a $0 = 0" by (simp add: fps_ginv_def) from a1 c1 have rca1: "?r c a $ 1 \ 0" by (simp add: fps_ginv_def field_simps) from fps_ginv[OF rca0 rca1] have "?r b (?r c a) oo ?r c a = b" . then have "?r b (?r c a) oo ?r c a oo a = b oo a" by simp then have "?r b (?r c a) oo (?r c a oo a) = b oo a" apply (subst fps_compose_assoc) using a0 c0 apply (simp_all add: fps_ginv_def) done then have "?r b (?r c a) oo c = b oo a" unfolding fps_ginv[OF a0 a1] . then have "?r b (?r c a) oo c oo fps_inv c= b oo a oo fps_inv c" by simp then have "?r b (?r c a) oo (c oo fps_inv c) = b oo a oo fps_inv c" apply (subst fps_compose_assoc) using a0 c0 apply (simp_all add: fps_inv_def) done then show ?thesis unfolding fps_inv_right[OF c0 c1] by simp qed lemma fps_ginv_deriv: assumes a0:"a$0 = (0::'a::field)" and a1: "a$1 \ 0" shows "fps_deriv (fps_ginv b a) = (fps_deriv b / fps_deriv a) oo fps_ginv fps_X a" proof - let ?ia = "fps_ginv b a" let ?ifps_Xa = "fps_ginv fps_X a" let ?d = "fps_deriv" let ?dia = "?d ?ia" have ifps_Xa0: "?ifps_Xa $ 0 = 0" by (simp add: fps_ginv_def) have da0: "?d a $ 0 \ 0" using a1 by simp from fps_ginv[OF a0 a1, of b] have "?d (?ia oo a) = fps_deriv b" by simp then have "(?d ?ia oo a) * ?d a = ?d b" unfolding fps_compose_deriv[OF a0] . then have "(?d ?ia oo a) * ?d a * inverse (?d a) = ?d b * inverse (?d a)" by simp with a1 have "(?d ?ia oo a) * (inverse (?d a) * ?d a) = ?d b / ?d a" by (simp add: fps_divide_unit) then have "(?d ?ia oo a) oo ?ifps_Xa = (?d b / ?d a) oo ?ifps_Xa" unfolding inverse_mult_eq_1[OF da0] by simp then have "?d ?ia oo (a oo ?ifps_Xa) = (?d b / ?d a) oo ?ifps_Xa" unfolding fps_compose_assoc[OF ifps_Xa0 a0] . then show ?thesis unfolding fps_inv_ginv[symmetric] unfolding fps_inv_right[OF a0 a1] by simp qed lemma fps_compose_linear: "fps_compose (f :: 'a :: comm_ring_1 fps) (fps_const c * fps_X) = Abs_fps (\n. c^n * f $ n)" by (simp add: fps_eq_iff fps_compose_def power_mult_distrib if_distrib cong: if_cong) lemma fps_compose_uminus': "fps_compose f (-fps_X :: 'a :: comm_ring_1 fps) = Abs_fps (\n. (-1)^n * f $ n)" using fps_compose_linear[of f "-1"] by (simp only: fps_const_neg [symmetric] fps_const_1_eq_1) simp subsection \Elementary series\ subsubsection \Exponential series\ definition "fps_exp x = Abs_fps (\n. x^n / of_nat (fact n))" lemma fps_exp_deriv[simp]: "fps_deriv (fps_exp a) = fps_const (a::'a::field_char_0) * fps_exp a" (is "?l = ?r") proof - have "?l$n = ?r $ n" for n apply (auto simp add: fps_exp_def field_simps power_Suc[symmetric] simp del: fact_Suc of_nat_Suc power_Suc) apply (simp add: field_simps) done then show ?thesis by (simp add: fps_eq_iff) qed lemma fps_exp_unique_ODE: "fps_deriv a = fps_const c * a \ a = fps_const (a$0) * fps_exp (c::'a::field_char_0)" (is "?lhs \ ?rhs") proof show ?rhs if ?lhs proof - from that have th: "\n. a $ Suc n = c * a$n / of_nat (Suc n)" by (simp add: fps_deriv_def fps_eq_iff field_simps del: of_nat_Suc) have th': "a$n = a$0 * c ^ n/ (fact n)" for n proof (induct n) case 0 then show ?case by simp next case Suc then show ?case unfolding th using fact_gt_zero apply (simp add: field_simps del: of_nat_Suc fact_Suc) apply simp done qed show ?thesis by (auto simp add: fps_eq_iff fps_const_mult_left fps_exp_def intro: th') qed show ?lhs if ?rhs using that by (metis fps_exp_deriv fps_deriv_mult_const_left mult.left_commute) qed lemma fps_exp_add_mult: "fps_exp (a + b) = fps_exp (a::'a::field_char_0) * fps_exp b" (is "?l = ?r") proof - have "fps_deriv ?r = fps_const (a + b) * ?r" by (simp add: fps_const_add[symmetric] field_simps del: fps_const_add) then have "?r = ?l" by (simp only: fps_exp_unique_ODE) (simp add: fps_mult_nth fps_exp_def) then show ?thesis .. qed lemma fps_exp_nth[simp]: "fps_exp a $ n = a^n / of_nat (fact n)" by (simp add: fps_exp_def) lemma fps_exp_0[simp]: "fps_exp (0::'a::field) = 1" by (simp add: fps_eq_iff power_0_left) lemma fps_exp_neg: "fps_exp (- a) = inverse (fps_exp (a::'a::field_char_0))" proof - from fps_exp_add_mult[of a "- a"] have th0: "fps_exp a * fps_exp (- a) = 1" by simp from fps_inverse_unique[OF th0] show ?thesis by simp qed lemma fps_exp_nth_deriv[simp]: "fps_nth_deriv n (fps_exp (a::'a::field_char_0)) = (fps_const a)^n * (fps_exp a)" by (induct n) auto lemma fps_X_compose_fps_exp[simp]: "fps_X oo fps_exp (a::'a::field) = fps_exp a - 1" by (simp add: fps_eq_iff fps_X_fps_compose) lemma fps_inv_fps_exp_compose: assumes a: "a \ 0" shows "fps_inv (fps_exp a - 1) oo (fps_exp a - 1) = fps_X" and "(fps_exp a - 1) oo fps_inv (fps_exp a - 1) = fps_X" proof - let ?b = "fps_exp a - 1" have b0: "?b $ 0 = 0" by simp have b1: "?b $ 1 \ 0" by (simp add: a) from fps_inv[OF b0 b1] show "fps_inv (fps_exp a - 1) oo (fps_exp a - 1) = fps_X" . from fps_inv_right[OF b0 b1] show "(fps_exp a - 1) oo fps_inv (fps_exp a - 1) = fps_X" . qed lemma fps_exp_power_mult: "(fps_exp (c::'a::field_char_0))^n = fps_exp (of_nat n * c)" by (induct n) (simp_all add: field_simps fps_exp_add_mult) lemma radical_fps_exp: assumes r: "r (Suc k) 1 = 1" shows "fps_radical r (Suc k) (fps_exp (c::'a::field_char_0)) = fps_exp (c / of_nat (Suc k))" proof - let ?ck = "(c / of_nat (Suc k))" let ?r = "fps_radical r (Suc k)" have eq0[simp]: "?ck * of_nat (Suc k) = c" "of_nat (Suc k) * ?ck = c" by (simp_all del: of_nat_Suc) have th0: "fps_exp ?ck ^ (Suc k) = fps_exp c" unfolding fps_exp_power_mult eq0 .. have th: "r (Suc k) (fps_exp c $0) ^ Suc k = fps_exp c $ 0" "r (Suc k) (fps_exp c $ 0) = fps_exp ?ck $ 0" "fps_exp c $ 0 \ 0" using r by simp_all from th0 radical_unique[where r=r and k=k, OF th] show ?thesis by auto qed lemma fps_exp_compose_linear [simp]: "fps_exp (d::'a::field_char_0) oo (fps_const c * fps_X) = fps_exp (c * d)" by (simp add: fps_compose_linear fps_exp_def fps_eq_iff power_mult_distrib) lemma fps_fps_exp_compose_minus [simp]: "fps_compose (fps_exp c) (-fps_X) = fps_exp (-c :: 'a :: field_char_0)" using fps_exp_compose_linear[of c "-1 :: 'a"] unfolding fps_const_neg [symmetric] fps_const_1_eq_1 by simp lemma fps_exp_eq_iff [simp]: "fps_exp c = fps_exp d \ c = (d :: 'a :: field_char_0)" proof assume "fps_exp c = fps_exp d" from arg_cong[of _ _ "\F. F $ 1", OF this] show "c = d" by simp qed simp_all lemma fps_exp_eq_fps_const_iff [simp]: "fps_exp (c :: 'a :: field_char_0) = fps_const c' \ c = 0 \ c' = 1" proof assume "c = 0 \ c' = 1" thus "fps_exp c = fps_const c'" by (simp add: fps_eq_iff) next assume "fps_exp c = fps_const c'" from arg_cong[of _ _ "\F. F $ 1", OF this] arg_cong[of _ _ "\F. F $ 0", OF this] show "c = 0 \ c' = 1" by simp_all qed lemma fps_exp_neq_0 [simp]: "\fps_exp (c :: 'a :: field_char_0) = 0" unfolding fps_const_0_eq_0 [symmetric] fps_exp_eq_fps_const_iff by simp lemma fps_exp_eq_1_iff [simp]: "fps_exp (c :: 'a :: field_char_0) = 1 \ c = 0" unfolding fps_const_1_eq_1 [symmetric] fps_exp_eq_fps_const_iff by simp lemma fps_exp_neq_numeral_iff [simp]: "fps_exp (c :: 'a :: field_char_0) = numeral n \ c = 0 \ n = Num.One" unfolding numeral_fps_const fps_exp_eq_fps_const_iff by simp subsubsection \Logarithmic series\ lemma Abs_fps_if_0: "Abs_fps (\n. if n = 0 then (v::'a::ring_1) else f n) = fps_const v + fps_X * Abs_fps (\n. f (Suc n))" by (simp add: fps_eq_iff) definition fps_ln :: "'a::field_char_0 \ 'a fps" where "fps_ln c = fps_const (1/c) * Abs_fps (\n. if n = 0 then 0 else (- 1) ^ (n - 1) / of_nat n)" lemma fps_ln_deriv: "fps_deriv (fps_ln c) = fps_const (1/c) * inverse (1 + fps_X)" unfolding fps_inverse_fps_X_plus1 by (simp add: fps_ln_def fps_eq_iff del: of_nat_Suc) lemma fps_ln_nth: "fps_ln c $ n = (if n = 0 then 0 else 1/c * ((- 1) ^ (n - 1) / of_nat n))" by (simp add: fps_ln_def field_simps) lemma fps_ln_0 [simp]: "fps_ln c $ 0 = 0" by (simp add: fps_ln_def) lemma fps_ln_fps_exp_inv: fixes a :: "'a::field_char_0" assumes a: "a \ 0" shows "fps_ln a = fps_inv (fps_exp a - 1)" (is "?l = ?r") proof - let ?b = "fps_exp a - 1" have b0: "?b $ 0 = 0" by simp have b1: "?b $ 1 \ 0" by (simp add: a) have "fps_deriv (fps_exp a - 1) oo fps_inv (fps_exp a - 1) = (fps_const a * (fps_exp a - 1) + fps_const a) oo fps_inv (fps_exp a - 1)" by (simp add: field_simps) also have "\ = fps_const a * (fps_X + 1)" apply (simp add: fps_compose_add_distrib fps_const_mult_apply_left[symmetric] fps_inv_right[OF b0 b1]) apply (simp add: field_simps) done finally have eq: "fps_deriv (fps_exp a - 1) oo fps_inv (fps_exp a - 1) = fps_const a * (fps_X + 1)" . from fps_inv_deriv[OF b0 b1, unfolded eq] have "fps_deriv (fps_inv ?b) = fps_const (inverse a) / (fps_X + 1)" using a by (simp add: fps_const_inverse eq fps_divide_def fps_inverse_mult) then have "fps_deriv ?l = fps_deriv ?r" by (simp add: fps_ln_deriv add.commute fps_divide_def divide_inverse) then show ?thesis unfolding fps_deriv_eq_iff by (simp add: fps_ln_nth fps_inv_def) qed lemma fps_ln_mult_add: assumes c0: "c\0" and d0: "d\0" shows "fps_ln c + fps_ln d = fps_const (c+d) * fps_ln (c*d)" (is "?r = ?l") proof- from c0 d0 have eq: "1/c + 1/d = (c+d)/(c*d)" by (simp add: field_simps) have "fps_deriv ?r = fps_const (1/c + 1/d) * inverse (1 + fps_X)" by (simp add: fps_ln_deriv fps_const_add[symmetric] algebra_simps del: fps_const_add) also have "\ = fps_deriv ?l" apply (simp add: fps_ln_deriv) apply (simp add: fps_eq_iff eq) done finally show ?thesis unfolding fps_deriv_eq_iff by simp qed lemma fps_X_dvd_fps_ln [simp]: "fps_X dvd fps_ln c" proof - have "fps_ln c = fps_X * Abs_fps (\n. (-1) ^ n / (of_nat (Suc n) * c))" by (intro fps_ext) (simp add: fps_ln_def of_nat_diff) thus ?thesis by simp qed subsubsection \Binomial series\ definition "fps_binomial a = Abs_fps (\n. a gchoose n)" lemma fps_binomial_nth[simp]: "fps_binomial a $ n = a gchoose n" by (simp add: fps_binomial_def) lemma fps_binomial_ODE_unique: fixes c :: "'a::field_char_0" shows "fps_deriv a = (fps_const c * a) / (1 + fps_X) \ a = fps_const (a$0) * fps_binomial c" (is "?lhs \ ?rhs") proof let ?da = "fps_deriv a" let ?x1 = "(1 + fps_X):: 'a fps" let ?l = "?x1 * ?da" let ?r = "fps_const c * a" have eq: "?l = ?r \ ?lhs" proof - have x10: "?x1 $ 0 \ 0" by simp have "?l = ?r \ inverse ?x1 * ?l = inverse ?x1 * ?r" by simp also have "\ \ ?da = (fps_const c * a) / ?x1" apply (simp only: fps_divide_def mult.assoc[symmetric] inverse_mult_eq_1[OF x10]) apply (simp add: field_simps) done finally show ?thesis . qed show ?rhs if ?lhs proof - from eq that have h: "?l = ?r" .. have th0: "a$ Suc n = ((c - of_nat n) / of_nat (Suc n)) * a $n" for n proof - from h have "?l $ n = ?r $ n" by simp then show ?thesis apply (simp add: field_simps del: of_nat_Suc) apply (cases n) apply (simp_all add: field_simps del: of_nat_Suc) done qed have th1: "a $ n = (c gchoose n) * a $ 0" for n proof (induct n) case 0 then show ?case by simp next case (Suc m) then show ?case unfolding th0 apply (simp add: field_simps del: of_nat_Suc) unfolding mult.assoc[symmetric] gbinomial_mult_1 apply (simp add: field_simps) done qed show ?thesis apply (simp add: fps_eq_iff) apply (subst th1) apply (simp add: field_simps) done qed show ?lhs if ?rhs proof - have th00: "x * (a $ 0 * y) = a $ 0 * (x * y)" for x y by (simp add: mult.commute) have "?l = ?r" apply (subst \?rhs\) apply (subst (2) \?rhs\) apply (clarsimp simp add: fps_eq_iff field_simps) unfolding mult.assoc[symmetric] th00 gbinomial_mult_1 apply (simp add: field_simps gbinomial_mult_1) done with eq show ?thesis .. qed qed lemma fps_binomial_ODE_unique': "(fps_deriv a = fps_const c * a / (1 + fps_X) \ a $ 0 = 1) \ (a = fps_binomial c)" by (subst fps_binomial_ODE_unique) auto lemma fps_binomial_deriv: "fps_deriv (fps_binomial c) = fps_const c * fps_binomial c / (1 + fps_X)" proof - let ?a = "fps_binomial c" have th0: "?a = fps_const (?a$0) * ?a" by (simp) from iffD2[OF fps_binomial_ODE_unique, OF th0] show ?thesis . qed lemma fps_binomial_add_mult: "fps_binomial (c+d) = fps_binomial c * fps_binomial d" (is "?l = ?r") proof - let ?P = "?r - ?l" let ?b = "fps_binomial" let ?db = "\x. fps_deriv (?b x)" have "fps_deriv ?P = ?db c * ?b d + ?b c * ?db d - ?db (c + d)" by simp also have "\ = inverse (1 + fps_X) * (fps_const c * ?b c * ?b d + fps_const d * ?b c * ?b d - fps_const (c+d) * ?b (c + d))" unfolding fps_binomial_deriv by (simp add: fps_divide_def field_simps) also have "\ = (fps_const (c + d)/ (1 + fps_X)) * ?P" by (simp add: field_simps fps_divide_unit fps_const_add[symmetric] del: fps_const_add) finally have th0: "fps_deriv ?P = fps_const (c+d) * ?P / (1 + fps_X)" by (simp add: fps_divide_def) have "?P = fps_const (?P$0) * ?b (c + d)" unfolding fps_binomial_ODE_unique[symmetric] using th0 by simp then have "?P = 0" by (simp add: fps_mult_nth) then show ?thesis by simp qed lemma fps_binomial_minus_one: "fps_binomial (- 1) = inverse (1 + fps_X)" (is "?l = inverse ?r") proof- have th: "?r$0 \ 0" by simp have th': "fps_deriv (inverse ?r) = fps_const (- 1) * inverse ?r / (1 + fps_X)" by (simp add: fps_inverse_deriv[OF th] fps_divide_def power2_eq_square mult.commute fps_const_neg[symmetric] del: fps_const_neg) have eq: "inverse ?r $ 0 = 1" by (simp add: fps_inverse_def) from iffD1[OF fps_binomial_ODE_unique[of "inverse (1 + fps_X)" "- 1"] th'] eq show ?thesis by (simp add: fps_inverse_def) qed lemma fps_binomial_of_nat: "fps_binomial (of_nat n) = (1 + fps_X :: 'a :: field_char_0 fps) ^ n" proof (cases "n = 0") case [simp]: True have "fps_deriv ((1 + fps_X) ^ n :: 'a fps) = 0" by simp also have "\ = fps_const (of_nat n) * (1 + fps_X) ^ n / (1 + fps_X)" by (simp add: fps_binomial_def) finally show ?thesis by (subst sym, subst fps_binomial_ODE_unique' [symmetric]) simp_all next case False have "fps_deriv ((1 + fps_X) ^ n :: 'a fps) = fps_const (of_nat n) * (1 + fps_X) ^ (n - 1)" by (simp add: fps_deriv_power) also have "(1 + fps_X :: 'a fps) $ 0 \ 0" by simp hence "(1 + fps_X :: 'a fps) \ 0" by (intro notI) (simp only: , simp) with False have "(1 + fps_X :: 'a fps) ^ (n - 1) = (1 + fps_X) ^ n / (1 + fps_X)" by (cases n) (simp_all ) also have "fps_const (of_nat n :: 'a) * ((1 + fps_X) ^ n / (1 + fps_X)) = fps_const (of_nat n) * (1 + fps_X) ^ n / (1 + fps_X)" by (simp add: unit_div_mult_swap) finally show ?thesis by (subst sym, subst fps_binomial_ODE_unique' [symmetric]) (simp_all add: fps_power_nth) qed lemma fps_binomial_0 [simp]: "fps_binomial 0 = 1" using fps_binomial_of_nat[of 0] by simp lemma fps_binomial_power: "fps_binomial a ^ n = fps_binomial (of_nat n * a)" by (induction n) (simp_all add: fps_binomial_add_mult ring_distribs) lemma fps_binomial_1: "fps_binomial 1 = 1 + fps_X" using fps_binomial_of_nat[of 1] by simp lemma fps_binomial_minus_of_nat: "fps_binomial (- of_nat n) = inverse ((1 + fps_X :: 'a :: field_char_0 fps) ^ n)" by (rule sym, rule fps_inverse_unique) (simp add: fps_binomial_of_nat [symmetric] fps_binomial_add_mult [symmetric]) lemma one_minus_const_fps_X_power: "c \ 0 \ (1 - fps_const c * fps_X) ^ n = fps_compose (fps_binomial (of_nat n)) (-fps_const c * fps_X)" by (subst fps_binomial_of_nat) (simp add: fps_compose_power [symmetric] fps_compose_add_distrib fps_const_neg [symmetric] del: fps_const_neg) lemma one_minus_fps_X_const_neg_power: "inverse ((1 - fps_const c * fps_X) ^ n) = fps_compose (fps_binomial (-of_nat n)) (-fps_const c * fps_X)" proof (cases "c = 0") case False thus ?thesis by (subst fps_binomial_minus_of_nat) (simp add: fps_compose_power [symmetric] fps_inverse_compose fps_compose_add_distrib fps_const_neg [symmetric] del: fps_const_neg) qed simp lemma fps_X_plus_const_power: "c \ 0 \ (fps_X + fps_const c) ^ n = fps_const (c^n) * fps_compose (fps_binomial (of_nat n)) (fps_const (inverse c) * fps_X)" by (subst fps_binomial_of_nat) (simp add: fps_compose_power [symmetric] fps_binomial_of_nat fps_compose_add_distrib fps_const_power [symmetric] power_mult_distrib [symmetric] algebra_simps inverse_mult_eq_1' del: fps_const_power) lemma fps_X_plus_const_neg_power: "c \ 0 \ inverse ((fps_X + fps_const c) ^ n) = fps_const (inverse c^n) * fps_compose (fps_binomial (-of_nat n)) (fps_const (inverse c) * fps_X)" by (subst fps_binomial_minus_of_nat) (simp add: fps_compose_power [symmetric] fps_binomial_of_nat fps_compose_add_distrib fps_const_power [symmetric] power_mult_distrib [symmetric] fps_inverse_compose algebra_simps fps_const_inverse [symmetric] fps_inverse_mult [symmetric] fps_inverse_power [symmetric] inverse_mult_eq_1' del: fps_const_power) lemma one_minus_const_fps_X_neg_power': "n > 0 \ inverse ((1 - fps_const (c :: 'a :: field_char_0) * fps_X) ^ n) = Abs_fps (\k. of_nat ((n + k - 1) choose k) * c^k)" apply (rule fps_ext) apply (subst one_minus_fps_X_const_neg_power, subst fps_const_neg, subst fps_compose_linear) apply (simp add: power_mult_distrib [symmetric] mult.assoc [symmetric] gbinomial_minus binomial_gbinomial of_nat_diff) done text \Vandermonde's Identity as a consequence.\ lemma gbinomial_Vandermonde: "sum (\k. (a gchoose k) * (b gchoose (n - k))) {0..n} = (a + b) gchoose n" proof - let ?ba = "fps_binomial a" let ?bb = "fps_binomial b" let ?bab = "fps_binomial (a + b)" from fps_binomial_add_mult[of a b] have "?bab $ n = (?ba * ?bb)$n" by simp then show ?thesis by (simp add: fps_mult_nth) qed lemma binomial_Vandermonde: "sum (\k. (a choose k) * (b choose (n - k))) {0..n} = (a + b) choose n" using gbinomial_Vandermonde[of "(of_nat a)" "of_nat b" n] by (simp only: binomial_gbinomial[symmetric] of_nat_mult[symmetric] of_nat_sum[symmetric] of_nat_add[symmetric] of_nat_eq_iff) lemma binomial_Vandermonde_same: "sum (\k. (n choose k)\<^sup>2) {0..n} = (2 * n) choose n" using binomial_Vandermonde[of n n n, symmetric] unfolding mult_2 apply (simp add: power2_eq_square) apply (rule sum.cong) apply (auto intro: binomial_symmetric) done lemma Vandermonde_pochhammer_lemma: fixes a :: "'a::field_char_0" assumes b: "\j\{0 .. of_nat j" shows "sum (\k. (pochhammer (- a) k * pochhammer (- (of_nat n)) k) / (of_nat (fact k) * pochhammer (b - of_nat n + 1) k)) {0..n} = pochhammer (- (a + b)) n / pochhammer (- b) n" (is "?l = ?r") proof - let ?m1 = "\m. (- 1 :: 'a) ^ m" let ?f = "\m. of_nat (fact m)" let ?p = "\(x::'a). pochhammer (- x)" from b have bn0: "?p b n \ 0" unfolding pochhammer_eq_0_iff by simp have th00: "b gchoose (n - k) = (?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k)" (is ?gchoose) "pochhammer (1 + b - of_nat n) k \ 0" (is ?pochhammer) if kn: "k \ {0..n}" for k proof - from kn have "k \ n" by simp have nz: "pochhammer (1 + b - of_nat n) n \ 0" proof assume "pochhammer (1 + b - of_nat n) n = 0" then have c: "pochhammer (b - of_nat n + 1) n = 0" by (simp add: algebra_simps) then obtain j where j: "j < n" "b - of_nat n + 1 = - of_nat j" unfolding pochhammer_eq_0_iff by blast from j have "b = of_nat n - of_nat j - of_nat 1" by (simp add: algebra_simps) then have "b = of_nat (n - j - 1)" using j kn by (simp add: of_nat_diff) with b show False using j by auto qed from nz kn [simplified] have nz': "pochhammer (1 + b - of_nat n) k \ 0" by (rule pochhammer_neq_0_mono) consider "k = 0 \ n = 0" | "k \ 0" "n \ 0" by blast then have "b gchoose (n - k) = (?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k)" proof cases case 1 then show ?thesis using kn by (cases "k = 0") (simp_all add: gbinomial_pochhammer) next case neq: 2 then obtain m where m: "n = Suc m" by (cases n) auto from neq(1) obtain h where h: "k = Suc h" by (cases k) auto show ?thesis proof (cases "k = n") case True then show ?thesis using pochhammer_minus'[where k=k and b=b] apply (simp add: pochhammer_same) using bn0 apply (simp add: field_simps power_add[symmetric]) done next case False with kn have kn': "k < n" by simp have m1nk: "?m1 n = prod (\i. - 1) {..m}" "?m1 k = prod (\i. - 1) {0..h}" by (simp_all add: prod_constant m h) have bnz0: "pochhammer (b - of_nat n + 1) k \ 0" using bn0 kn unfolding pochhammer_eq_0_iff apply auto apply (erule_tac x= "n - ka - 1" in allE) apply (auto simp add: algebra_simps of_nat_diff) done have eq1: "prod (\k. (1::'a) + of_nat m - of_nat k) {..h} = prod of_nat {Suc (m - h) .. Suc m}" using kn' h m by (intro prod.reindex_bij_witness[where i="\k. Suc m - k" and j="\k. Suc m - k"]) (auto simp: of_nat_diff) have th1: "(?m1 k * ?p (of_nat n) k) / ?f n = 1 / of_nat(fact (n - k))" apply (simp add: pochhammer_minus field_simps) using \k \ n\ apply (simp add: fact_split [of k n]) apply (simp add: pochhammer_prod) using prod.atLeastLessThan_shift_bounds [where ?'a = 'a, of "\i. 1 + of_nat i" 0 "n - k" k] apply (auto simp add: of_nat_diff field_simps) done have th20: "?m1 n * ?p b n = prod (\i. b - of_nat i) {0..m}" apply (simp add: pochhammer_minus field_simps m) apply (auto simp add: pochhammer_prod_rev of_nat_diff prod.atLeast_Suc_atMost_Suc_shift simp del: prod.cl_ivl_Suc) done have th21:"pochhammer (b - of_nat n + 1) k = prod (\i. b - of_nat i) {n - k .. n - 1}" using kn apply (simp add: pochhammer_prod_rev m h prod.atLeast_Suc_atMost_Suc_shift del: prod.op_ivl_Suc del: prod.cl_ivl_Suc) using prod.atLeastAtMost_shift_0 [of "m - h" m, where ?'a = 'a] apply (auto simp add: of_nat_diff field_simps) done have "?m1 n * ?p b n = prod (\i. b - of_nat i) {0.. n - k - 1} * pochhammer (b - of_nat n + 1) k" using kn' m h unfolding th20 th21 apply simp apply (subst prod.union_disjoint [symmetric]) apply auto apply (rule prod.cong) apply auto done then have th2: "(?m1 n * ?p b n)/pochhammer (b - of_nat n + 1) k = prod (\i. b - of_nat i) {0.. n - k - 1}" using nz' by (simp add: field_simps) have "(?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k) = ((?m1 k * ?p (of_nat n) k) / ?f n) * ((?m1 n * ?p b n)/pochhammer (b - of_nat n + 1) k)" using bnz0 by (simp add: field_simps) also have "\ = b gchoose (n - k)" unfolding th1 th2 using kn' m h apply (simp add: field_simps gbinomial_mult_fact) apply (rule prod.cong) apply auto done finally show ?thesis by simp qed qed then show ?gchoose and ?pochhammer apply (cases "n = 0") using nz' apply auto done qed have "?r = ((a + b) gchoose n) * (of_nat (fact n) / (?m1 n * pochhammer (- b) n))" unfolding gbinomial_pochhammer using bn0 by (auto simp add: field_simps) also have "\ = ?l" unfolding gbinomial_Vandermonde[symmetric] apply (simp add: th00) unfolding gbinomial_pochhammer using bn0 apply (simp add: sum_distrib_right sum_distrib_left field_simps) done finally show ?thesis by simp qed lemma Vandermonde_pochhammer: fixes a :: "'a::field_char_0" assumes c: "\i \ {0..< n}. c \ - of_nat i" shows "sum (\k. (pochhammer a k * pochhammer (- (of_nat n)) k) / (of_nat (fact k) * pochhammer c k)) {0..n} = pochhammer (c - a) n / pochhammer c n" proof - let ?a = "- a" let ?b = "c + of_nat n - 1" have h: "\ j \{0..< n}. ?b \ of_nat j" using c apply (auto simp add: algebra_simps of_nat_diff) apply (erule_tac x = "n - j - 1" in ballE) apply (auto simp add: of_nat_diff algebra_simps) done have th0: "pochhammer (- (?a + ?b)) n = (- 1)^n * pochhammer (c - a) n" unfolding pochhammer_minus by (simp add: algebra_simps) have th1: "pochhammer (- ?b) n = (- 1)^n * pochhammer c n" unfolding pochhammer_minus by simp have nz: "pochhammer c n \ 0" using c by (simp add: pochhammer_eq_0_iff) from Vandermonde_pochhammer_lemma[where a = "?a" and b="?b" and n=n, OF h, unfolded th0 th1] show ?thesis using nz by (simp add: field_simps sum_distrib_left) qed subsubsection \Formal trigonometric functions\ definition "fps_sin (c::'a::field_char_0) = Abs_fps (\n. if even n then 0 else (- 1) ^((n - 1) div 2) * c^n /(of_nat (fact n)))" definition "fps_cos (c::'a::field_char_0) = Abs_fps (\n. if even n then (- 1) ^ (n div 2) * c^n / (of_nat (fact n)) else 0)" lemma fps_sin_0 [simp]: "fps_sin 0 = 0" by (intro fps_ext) (auto simp: fps_sin_def elim!: oddE) lemma fps_cos_0 [simp]: "fps_cos 0 = 1" by (intro fps_ext) (simp add: fps_cos_def) lemma fps_sin_deriv: "fps_deriv (fps_sin c) = fps_const c * fps_cos c" (is "?lhs = ?rhs") proof (rule fps_ext) fix n :: nat show "?lhs $ n = ?rhs $ n" proof (cases "even n") case True have "?lhs$n = of_nat (n+1) * (fps_sin c $ (n+1))" by simp also have "\ = of_nat (n+1) * ((- 1)^(n div 2) * c^Suc n / of_nat (fact (Suc n)))" using True by (simp add: fps_sin_def) also have "\ = (- 1)^(n div 2) * c^Suc n * (of_nat (n+1) / (of_nat (Suc n) * of_nat (fact n)))" unfolding fact_Suc of_nat_mult by (simp add: field_simps del: of_nat_add of_nat_Suc) also have "\ = (- 1)^(n div 2) * c^Suc n / of_nat (fact n)" by (simp add: field_simps del: of_nat_add of_nat_Suc) finally show ?thesis using True by (simp add: fps_cos_def field_simps) next case False then show ?thesis by (simp_all add: fps_deriv_def fps_sin_def fps_cos_def) qed qed lemma fps_cos_deriv: "fps_deriv (fps_cos c) = fps_const (- c)* (fps_sin c)" (is "?lhs = ?rhs") proof (rule fps_ext) have th0: "- ((- 1::'a) ^ n) = (- 1)^Suc n" for n by simp show "?lhs $ n = ?rhs $ n" for n proof (cases "even n") case False then have n0: "n \ 0" by presburger from False have th1: "Suc ((n - 1) div 2) = Suc n div 2" by (cases n) simp_all have "?lhs$n = of_nat (n+1) * (fps_cos c $ (n+1))" by simp also have "\ = of_nat (n+1) * ((- 1)^((n + 1) div 2) * c^Suc n / of_nat (fact (Suc n)))" using False by (simp add: fps_cos_def) also have "\ = (- 1)^((n + 1) div 2) * c^Suc n * (of_nat (n+1) / (of_nat (Suc n) * of_nat (fact n)))" unfolding fact_Suc of_nat_mult by (simp add: field_simps del: of_nat_add of_nat_Suc) also have "\ = (- 1)^((n + 1) div 2) * c^Suc n / of_nat (fact n)" by (simp add: field_simps del: of_nat_add of_nat_Suc) also have "\ = (- ((- 1)^((n - 1) div 2))) * c^Suc n / of_nat (fact n)" unfolding th0 unfolding th1 by simp finally show ?thesis using False by (simp add: fps_sin_def field_simps) next case True then show ?thesis by (simp_all add: fps_deriv_def fps_sin_def fps_cos_def) qed qed lemma fps_sin_cos_sum_of_squares: "(fps_cos c)\<^sup>2 + (fps_sin c)\<^sup>2 = 1" (is "?lhs = _") proof - have "fps_deriv ?lhs = 0" apply (simp add: fps_deriv_power fps_sin_deriv fps_cos_deriv) apply (simp add: field_simps fps_const_neg[symmetric] del: fps_const_neg) done then have "?lhs = fps_const (?lhs $ 0)" unfolding fps_deriv_eq_0_iff . also have "\ = 1" by (simp add: fps_eq_iff numeral_2_eq_2 fps_mult_nth fps_cos_def fps_sin_def) finally show ?thesis . qed lemma fps_sin_nth_0 [simp]: "fps_sin c $ 0 = 0" unfolding fps_sin_def by simp lemma fps_sin_nth_1 [simp]: "fps_sin c $ 1 = c" unfolding fps_sin_def by simp lemma fps_sin_nth_add_2: "fps_sin c $ (n + 2) = - (c * c * fps_sin c $ n / (of_nat (n + 1) * of_nat (n + 2)))" unfolding fps_sin_def apply (cases n) apply simp apply (simp add: nonzero_divide_eq_eq nonzero_eq_divide_eq del: of_nat_Suc fact_Suc) apply simp done lemma fps_cos_nth_0 [simp]: "fps_cos c $ 0 = 1" unfolding fps_cos_def by simp lemma fps_cos_nth_1 [simp]: "fps_cos c $ 1 = 0" unfolding fps_cos_def by simp lemma fps_cos_nth_add_2: "fps_cos c $ (n + 2) = - (c * c * fps_cos c $ n / (of_nat (n + 1) * of_nat (n + 2)))" unfolding fps_cos_def apply (simp add: nonzero_divide_eq_eq nonzero_eq_divide_eq del: of_nat_Suc fact_Suc) apply simp done lemma nat_add_1_add_1: "(n::nat) + 1 + 1 = n + 2" by simp lemma eq_fps_sin: assumes 0: "a $ 0 = 0" and 1: "a $ 1 = c" and 2: "fps_deriv (fps_deriv a) = - (fps_const c * fps_const c * a)" shows "a = fps_sin c" apply (rule fps_ext) apply (induct_tac n rule: nat_induct2) apply (simp add: 0) apply (simp add: 1 del: One_nat_def) apply (rename_tac m, cut_tac f="\a. a $ m" in arg_cong [OF 2]) apply (simp add: nat_add_1_add_1 fps_sin_nth_add_2 del: One_nat_def of_nat_Suc of_nat_add add_2_eq_Suc') apply (subst minus_divide_left) apply (subst nonzero_eq_divide_eq) apply (simp del: of_nat_add of_nat_Suc) apply (simp only: ac_simps) done lemma eq_fps_cos: assumes 0: "a $ 0 = 1" and 1: "a $ 1 = 0" and 2: "fps_deriv (fps_deriv a) = - (fps_const c * fps_const c * a)" shows "a = fps_cos c" apply (rule fps_ext) apply (induct_tac n rule: nat_induct2) apply (simp add: 0) apply (simp add: 1 del: One_nat_def) apply (rename_tac m, cut_tac f="\a. a $ m" in arg_cong [OF 2]) apply (simp add: nat_add_1_add_1 fps_cos_nth_add_2 del: One_nat_def of_nat_Suc of_nat_add add_2_eq_Suc') apply (subst minus_divide_left) apply (subst nonzero_eq_divide_eq) apply (simp del: of_nat_add of_nat_Suc) apply (simp only: ac_simps) done lemma fps_sin_add: "fps_sin (a + b) = fps_sin a * fps_cos b + fps_cos a * fps_sin b" apply (rule eq_fps_sin [symmetric], simp, simp del: One_nat_def) apply (simp del: fps_const_neg fps_const_add fps_const_mult add: fps_const_add [symmetric] fps_const_neg [symmetric] fps_sin_deriv fps_cos_deriv algebra_simps) done lemma fps_cos_add: "fps_cos (a + b) = fps_cos a * fps_cos b - fps_sin a * fps_sin b" apply (rule eq_fps_cos [symmetric], simp, simp del: One_nat_def) apply (simp del: fps_const_neg fps_const_add fps_const_mult add: fps_const_add [symmetric] fps_const_neg [symmetric] fps_sin_deriv fps_cos_deriv algebra_simps) done lemma fps_sin_even: "fps_sin (- c) = - fps_sin c" by (simp add: fps_eq_iff fps_sin_def) lemma fps_cos_odd: "fps_cos (- c) = fps_cos c" by (simp add: fps_eq_iff fps_cos_def) definition "fps_tan c = fps_sin c / fps_cos c" lemma fps_tan_0 [simp]: "fps_tan 0 = 0" by (simp add: fps_tan_def) lemma fps_tan_deriv: "fps_deriv (fps_tan c) = fps_const c / (fps_cos c)\<^sup>2" proof - have th0: "fps_cos c $ 0 \ 0" by (simp add: fps_cos_def) from this have "fps_cos c \ 0" by (intro notI) simp hence "fps_deriv (fps_tan c) = fps_const c * (fps_cos c^2 + fps_sin c^2) / (fps_cos c^2)" by (simp add: fps_tan_def fps_divide_deriv power2_eq_square algebra_simps fps_sin_deriv fps_cos_deriv fps_const_neg[symmetric] div_mult_swap del: fps_const_neg) also note fps_sin_cos_sum_of_squares finally show ?thesis by simp qed text \Connection to @{const "fps_exp"} over the complex numbers --- Euler and de Moivre.\ lemma fps_exp_ii_sin_cos: "fps_exp (\ * c) = fps_cos c + fps_const \ * fps_sin c" (is "?l = ?r") proof - have "?l $ n = ?r $ n" for n proof (cases "even n") case True then obtain m where m: "n = 2 * m" .. show ?thesis by (simp add: m fps_sin_def fps_cos_def power_mult_distrib power_mult power_minus [of "c ^ 2"]) next case False then obtain m where m: "n = 2 * m + 1" .. show ?thesis by (simp add: m fps_sin_def fps_cos_def power_mult_distrib power_mult power_minus [of "c ^ 2"]) qed then show ?thesis by (simp add: fps_eq_iff) qed lemma fps_exp_minus_ii_sin_cos: "fps_exp (- (\ * c)) = fps_cos c - fps_const \ * fps_sin c" unfolding minus_mult_right fps_exp_ii_sin_cos by (simp add: fps_sin_even fps_cos_odd) lemma fps_cos_fps_exp_ii: "fps_cos c = (fps_exp (\ * c) + fps_exp (- \ * c)) / fps_const 2" proof - have th: "fps_cos c + fps_cos c = fps_cos c * fps_const 2" by (simp add: numeral_fps_const) show ?thesis unfolding fps_exp_ii_sin_cos minus_mult_commute by (simp add: fps_sin_even fps_cos_odd numeral_fps_const fps_divide_unit fps_const_inverse th) qed lemma fps_sin_fps_exp_ii: "fps_sin c = (fps_exp (\ * c) - fps_exp (- \ * c)) / fps_const (2*\)" proof - have th: "fps_const \ * fps_sin c + fps_const \ * fps_sin c = fps_sin c * fps_const (2 * \)" by (simp add: fps_eq_iff numeral_fps_const) show ?thesis unfolding fps_exp_ii_sin_cos minus_mult_commute by (simp add: fps_sin_even fps_cos_odd fps_divide_unit fps_const_inverse th) qed lemma fps_tan_fps_exp_ii: "fps_tan c = (fps_exp (\ * c) - fps_exp (- \ * c)) / (fps_const \ * (fps_exp (\ * c) + fps_exp (- \ * c)))" unfolding fps_tan_def fps_sin_fps_exp_ii fps_cos_fps_exp_ii mult_minus_left fps_exp_neg apply (simp add: fps_divide_unit fps_inverse_mult fps_const_mult[symmetric] fps_const_inverse del: fps_const_mult) apply simp done lemma fps_demoivre: "(fps_cos a + fps_const \ * fps_sin a)^n = fps_cos (of_nat n * a) + fps_const \ * fps_sin (of_nat n * a)" unfolding fps_exp_ii_sin_cos[symmetric] fps_exp_power_mult by (simp add: ac_simps) subsection \Hypergeometric series\ definition "fps_hypergeo as bs (c::'a::field_char_0) = Abs_fps (\n. (foldl (\r a. r* pochhammer a n) 1 as * c^n) / (foldl (\r b. r * pochhammer b n) 1 bs * of_nat (fact n)))" lemma fps_hypergeo_nth[simp]: "fps_hypergeo as bs c $ n = (foldl (\r a. r* pochhammer a n) 1 as * c^n) / (foldl (\r b. r * pochhammer b n) 1 bs * of_nat (fact n))" by (simp add: fps_hypergeo_def) lemma foldl_mult_start: fixes v :: "'a::comm_ring_1" shows "foldl (\r x. r * f x) v as * x = foldl (\r x. r * f x) (v * x) as " by (induct as arbitrary: x v) (auto simp add: algebra_simps) lemma foldr_mult_foldl: fixes v :: "'a::comm_ring_1" shows "foldr (\x r. r * f x) as v = foldl (\r x. r * f x) v as" by (induct as arbitrary: v) (simp_all add: foldl_mult_start) lemma fps_hypergeo_nth_alt: "fps_hypergeo as bs c $ n = foldr (\a r. r * pochhammer a n) as (c ^ n) / foldr (\b r. r * pochhammer b n) bs (of_nat (fact n))" by (simp add: foldl_mult_start foldr_mult_foldl) lemma fps_hypergeo_fps_exp[simp]: "fps_hypergeo [] [] c = fps_exp c" by (simp add: fps_eq_iff) lemma fps_hypergeo_1_0[simp]: "fps_hypergeo [1] [] c = 1/(1 - fps_const c * fps_X)" proof - let ?a = "(Abs_fps (\n. 1)) oo (fps_const c * fps_X)" have th0: "(fps_const c * fps_X) $ 0 = 0" by simp show ?thesis unfolding gp[OF th0, symmetric] by (simp add: fps_eq_iff pochhammer_fact[symmetric] fps_compose_nth power_mult_distrib if_distrib cong del: if_weak_cong) qed lemma fps_hypergeo_B[simp]: "fps_hypergeo [-a] [] (- 1) = fps_binomial a" by (simp add: fps_eq_iff gbinomial_pochhammer algebra_simps) lemma fps_hypergeo_0[simp]: "fps_hypergeo as bs c $ 0 = 1" apply simp apply (subgoal_tac "\as. foldl (\(r::'a) (a::'a). r) 1 as = 1") apply auto apply (induct_tac as) apply auto done lemma foldl_prod_prod: "foldl (\(r::'b::comm_ring_1) (x::'a::comm_ring_1). r * f x) v as * foldl (\r x. r * g x) w as = foldl (\r x. r * f x * g x) (v * w) as" by (induct as arbitrary: v w) (simp_all add: algebra_simps) lemma fps_hypergeo_rec: "fps_hypergeo as bs c $ Suc n = ((foldl (\r a. r* (a + of_nat n)) c as) / (foldl (\r b. r * (b + of_nat n)) (of_nat (Suc n)) bs )) * fps_hypergeo as bs c $ n" apply (simp del: of_nat_Suc of_nat_add fact_Suc) apply (simp add: foldl_mult_start del: fact_Suc of_nat_Suc) unfolding foldl_prod_prod[unfolded foldl_mult_start] pochhammer_Suc apply (simp add: algebra_simps) done lemma fps_XD_nth[simp]: "fps_XD a $ n = of_nat n * a$n" by (simp add: fps_XD_def) lemma fps_XD_0th[simp]: "fps_XD a $ 0 = 0" by simp lemma fps_XD_Suc[simp]:" fps_XD a $ Suc n = of_nat (Suc n) * a $ Suc n" by simp definition "fps_XDp c a = fps_XD a + fps_const c * a" lemma fps_XDp_nth[simp]: "fps_XDp c a $ n = (c + of_nat n) * a$n" by (simp add: fps_XDp_def algebra_simps) lemma fps_XDp_commute: "fps_XDp b \ fps_XDp (c::'a::comm_ring_1) = fps_XDp c \ fps_XDp b" by (simp add: fps_XDp_def fun_eq_iff fps_eq_iff algebra_simps) lemma fps_XDp0 [simp]: "fps_XDp 0 = fps_XD" by (simp add: fun_eq_iff fps_eq_iff) lemma fps_XDp_fps_integral [simp]: fixes a :: "'a::{division_ring,ring_char_0} fps" shows "fps_XDp 0 (fps_integral a c) = fps_X * a" using fps_deriv_fps_integral[of a c] by (simp add: fps_XD_def) lemma fps_hypergeo_minus_nat: "fps_hypergeo [- of_nat n] [- of_nat (n + m)] (c::'a::field_char_0) $ k = (if k \ n then pochhammer (- of_nat n) k * c ^ k / (pochhammer (- of_nat (n + m)) k * of_nat (fact k)) else 0)" "fps_hypergeo [- of_nat m] [- of_nat (m + n)] (c::'a::field_char_0) $ k = (if k \ m then pochhammer (- of_nat m) k * c ^ k / (pochhammer (- of_nat (m + n)) k * of_nat (fact k)) else 0)" by (simp_all add: pochhammer_eq_0_iff) lemma sum_eq_if: "sum f {(n::nat) .. m} = (if m < n then 0 else f n + sum f {n+1 .. m})" apply simp apply (subst sum.insert[symmetric]) apply (auto simp add: not_less sum.atLeast_Suc_atMost) done lemma pochhammer_rec_if: "pochhammer a n = (if n = 0 then 1 else a * pochhammer (a + 1) (n - 1))" by (cases n) (simp_all add: pochhammer_rec) lemma fps_XDp_foldr_nth [simp]: "foldr (\c r. fps_XDp c \ r) cs (\c. fps_XDp c a) c0 $ n = foldr (\c r. (c + of_nat n) * r) cs (c0 + of_nat n) * a$n" by (induct cs arbitrary: c0) (simp_all add: algebra_simps) lemma genric_fps_XDp_foldr_nth: assumes f: "\n c a. f c a $ n = (of_nat n + k c) * a$n" shows "foldr (\c r. f c \ r) cs (\c. g c a) c0 $ n = foldr (\c r. (k c + of_nat n) * r) cs (g c0 a $ n)" by (induct cs arbitrary: c0) (simp_all add: algebra_simps f) lemma dist_less_imp_nth_equal: assumes "dist f g < inverse (2 ^ i)" and"j \ i" shows "f $ j = g $ j" proof (rule ccontr) assume "f $ j \ g $ j" hence "f \ g" by auto with assms have "i < subdegree (f - g)" by (simp add: if_split_asm dist_fps_def) also have "\ \ j" using \f $ j \ g $ j\ by (intro subdegree_leI) simp_all finally show False using \j \ i\ by simp qed lemma nth_equal_imp_dist_less: assumes "\j. j \ i \ f $ j = g $ j" shows "dist f g < inverse (2 ^ i)" proof (cases "f = g") case True then show ?thesis by simp next case False with assms have "dist f g = inverse (2 ^ subdegree (f - g))" by (simp add: if_split_asm dist_fps_def) moreover from assms and False have "i < subdegree (f - g)" by (intro subdegree_greaterI) simp_all ultimately show ?thesis by simp qed lemma dist_less_eq_nth_equal: "dist f g < inverse (2 ^ i) \ (\j \ i. f $ j = g $ j)" using dist_less_imp_nth_equal nth_equal_imp_dist_less by blast instance fps :: (comm_ring_1) complete_space proof fix fps_X :: "nat \ 'a fps" assume "Cauchy fps_X" obtain M where M: "\i. \m \ M i. \j \ i. fps_X (M i) $ j = fps_X m $ j" proof - have "\M. \m \ M. \j\i. fps_X M $ j = fps_X m $ j" for i proof - have "0 < inverse ((2::real)^i)" by simp from metric_CauchyD[OF \Cauchy fps_X\ this] dist_less_imp_nth_equal show ?thesis by blast qed then show ?thesis using that by metis qed show "convergent fps_X" proof (rule convergentI) show "fps_X \ Abs_fps (\i. fps_X (M i) $ i)" unfolding tendsto_iff proof safe fix e::real assume e: "0 < e" have "(\n. inverse (2 ^ n) :: real) \ 0" by (rule LIMSEQ_inverse_realpow_zero) simp_all from this and e have "eventually (\i. inverse (2 ^ i) < e) sequentially" by (rule order_tendstoD) then obtain i where "inverse (2 ^ i) < e" by (auto simp: eventually_sequentially) have "eventually (\x. M i \ x) sequentially" by (auto simp: eventually_sequentially) then show "eventually (\x. dist (fps_X x) (Abs_fps (\i. fps_X (M i) $ i)) < e) sequentially" proof eventually_elim fix x assume x: "M i \ x" have "fps_X (M i) $ j = fps_X (M j) $ j" if "j \ i" for j using M that by (metis nat_le_linear) with x have "dist (fps_X x) (Abs_fps (\j. fps_X (M j) $ j)) < inverse (2 ^ i)" using M by (force simp: dist_less_eq_nth_equal) also note \inverse (2 ^ i) < e\ finally show "dist (fps_X x) (Abs_fps (\j. fps_X (M j) $ j)) < e" . qed qed qed qed (* TODO: Figure out better notation for this thing *) no_notation fps_nth (infixl "$" 75) bundle fps_notation begin notation fps_nth (infixl "$" 75) end end diff --git a/src/HOL/Computational_Algebra/Normalized_Fraction.thy b/src/HOL/Computational_Algebra/Normalized_Fraction.thy --- a/src/HOL/Computational_Algebra/Normalized_Fraction.thy +++ b/src/HOL/Computational_Algebra/Normalized_Fraction.thy @@ -1,363 +1,364 @@ (* Title: HOL/Computational_Algebra/Normalized_Fraction.thy Author: Manuel Eberl *) theory Normalized_Fraction imports Main Euclidean_Algorithm Fraction_Field begin definition quot_to_fract :: "'a :: {idom} \ 'a \ 'a fract" where "quot_to_fract = (\(a,b). Fraction_Field.Fract a b)" -definition normalize_quot :: "'a :: {ring_gcd,idom_divide} \ 'a \ 'a \ 'a" where +definition normalize_quot :: "'a :: {ring_gcd,idom_divide,semiring_gcd_mult_normalize} \ 'a \ 'a \ 'a" where "normalize_quot = (\(a,b). if b = 0 then (0,1) else let d = gcd a b * unit_factor b in (a div d, b div d))" lemma normalize_quot_zero [simp]: "normalize_quot (a, 0) = (0, 1)" by (simp add: normalize_quot_def) lemma normalize_quot_proj: "fst (normalize_quot (a, b)) = a div (gcd a b * unit_factor b)" "snd (normalize_quot (a, b)) = normalize b div gcd a b" if "b \ 0" using that by (simp_all add: normalize_quot_def Let_def mult.commute [of _ "unit_factor b"] dvd_div_mult2_eq mult_unit_dvd_iff') definition normalized_fracts :: "('a :: {ring_gcd,idom_divide} \ 'a) set" where "normalized_fracts = {(a,b). coprime a b \ unit_factor b = 1}" lemma not_normalized_fracts_0_denom [simp]: "(a, 0) \ normalized_fracts" by (auto simp: normalized_fracts_def) lemma unit_factor_snd_normalize_quot [simp]: "unit_factor (snd (normalize_quot x)) = 1" by (simp add: normalize_quot_def case_prod_unfold Let_def dvd_unit_factor_div mult_unit_dvd_iff unit_factor_mult unit_factor_gcd) lemma snd_normalize_quot_nonzero [simp]: "snd (normalize_quot x) \ 0" using unit_factor_snd_normalize_quot[of x] by (auto simp del: unit_factor_snd_normalize_quot) lemma normalize_quot_aux: fixes a b assumes "b \ 0" defines "d \ gcd a b * unit_factor b" shows "a = fst (normalize_quot (a,b)) * d" "b = snd (normalize_quot (a,b)) * d" "d dvd a" "d dvd b" "d \ 0" proof - from assms show "d dvd a" "d dvd b" by (simp_all add: d_def mult_unit_dvd_iff) thus "a = fst (normalize_quot (a,b)) * d" "b = snd (normalize_quot (a,b)) * d" "d \ 0" by (auto simp: normalize_quot_def Let_def d_def \b \ 0\) qed lemma normalize_quotE: assumes "b \ 0" obtains d where "a = fst (normalize_quot (a,b)) * d" "b = snd (normalize_quot (a,b)) * d" "d dvd a" "d dvd b" "d \ 0" using that[OF normalize_quot_aux[OF assms]] . lemma normalize_quotE': assumes "snd x \ 0" obtains d where "fst x = fst (normalize_quot x) * d" "snd x = snd (normalize_quot x) * d" "d dvd fst x" "d dvd snd x" "d \ 0" proof - from normalize_quotE[OF assms, of "fst x"] guess d . from this show ?thesis unfolding prod.collapse by (intro that[of d]) qed lemma coprime_normalize_quot: "coprime (fst (normalize_quot x)) (snd (normalize_quot x))" by (simp add: normalize_quot_def case_prod_unfold div_mult_unit2) (metis coprime_mult_self_right_iff div_gcd_coprime unit_div_mult_self unit_factor_is_unit) lemma normalize_quot_in_normalized_fracts [simp]: "normalize_quot x \ normalized_fracts" by (simp add: normalized_fracts_def coprime_normalize_quot case_prod_unfold) lemma normalize_quot_eq_iff: assumes "b \ 0" "d \ 0" shows "normalize_quot (a,b) = normalize_quot (c,d) \ a * d = b * c" proof - define x y where "x = normalize_quot (a,b)" and "y = normalize_quot (c,d)" from normalize_quotE[OF assms(1), of a] normalize_quotE[OF assms(2), of c] obtain d1 d2 where "a = fst x * d1" "b = snd x * d1" "c = fst y * d2" "d = snd y * d2" "d1 \ 0" "d2 \ 0" unfolding x_def y_def by metis hence "a * d = b * c \ fst x * snd y = snd x * fst y" by simp also have "\ \ fst x = fst y \ snd x = snd y" by (intro coprime_crossproduct') (simp_all add: x_def y_def coprime_normalize_quot) also have "\ \ x = y" using prod_eqI by blast finally show "x = y \ a * d = b * c" .. qed lemma normalize_quot_eq_iff': assumes "snd x \ 0" "snd y \ 0" shows "normalize_quot x = normalize_quot y \ fst x * snd y = snd x * fst y" using assms by (cases x, cases y, hypsubst) (subst normalize_quot_eq_iff, simp_all) lemma normalize_quot_id: "x \ normalized_fracts \ normalize_quot x = x" by (auto simp: normalized_fracts_def normalize_quot_def case_prod_unfold) lemma normalize_quot_idem [simp]: "normalize_quot (normalize_quot x) = normalize_quot x" by (rule normalize_quot_id) simp_all lemma fractrel_iff_normalize_quot_eq: "fractrel x y \ normalize_quot x = normalize_quot y \ snd x \ 0 \ snd y \ 0" by (cases x, cases y) (auto simp: fractrel_def normalize_quot_eq_iff) lemma fractrel_normalize_quot_left: assumes "snd x \ 0" shows "fractrel (normalize_quot x) y \ fractrel x y" using assms by (subst (1 2) fractrel_iff_normalize_quot_eq) auto lemma fractrel_normalize_quot_right: assumes "snd x \ 0" shows "fractrel y (normalize_quot x) \ fractrel y x" using assms by (subst (1 2) fractrel_iff_normalize_quot_eq) auto -lift_definition quot_of_fract :: "'a :: {ring_gcd,idom_divide} fract \ 'a \ 'a" +lift_definition quot_of_fract :: + "'a :: {ring_gcd,idom_divide,semiring_gcd_mult_normalize} fract \ 'a \ 'a" is normalize_quot by (subst (asm) fractrel_iff_normalize_quot_eq) simp_all lemma quot_to_fract_quot_of_fract [simp]: "quot_to_fract (quot_of_fract x) = x" unfolding quot_to_fract_def proof transfer fix x :: "'a \ 'a" assume rel: "fractrel x x" define x' where "x' = normalize_quot x" obtain a b where [simp]: "x = (a, b)" by (cases x) from rel have "b \ 0" by simp from normalize_quotE[OF this, of a] guess d . hence "a = fst x' * d" "b = snd x' * d" "d \ 0" "snd x' \ 0" by (simp_all add: x'_def) thus "fractrel (case x' of (a, b) \ if b = 0 then (0, 1) else (a, b)) x" by (auto simp add: case_prod_unfold) qed lemma quot_of_fract_quot_to_fract: "quot_of_fract (quot_to_fract x) = normalize_quot x" proof (cases "snd x = 0") case True thus ?thesis unfolding quot_to_fract_def by transfer (simp add: case_prod_unfold normalize_quot_def) next case False thus ?thesis unfolding quot_to_fract_def by transfer (simp add: case_prod_unfold) qed lemma quot_of_fract_quot_to_fract': "x \ normalized_fracts \ quot_of_fract (quot_to_fract x) = x" unfolding quot_to_fract_def by transfer (auto simp: normalize_quot_id) lemma quot_of_fract_in_normalized_fracts [simp]: "quot_of_fract x \ normalized_fracts" by transfer simp lemma normalize_quotI: assumes "a * d = b * c" "b \ 0" "(c, d) \ normalized_fracts" shows "normalize_quot (a, b) = (c, d)" proof - from assms have "normalize_quot (a, b) = normalize_quot (c, d)" by (subst normalize_quot_eq_iff) auto also have "\ = (c, d)" by (intro normalize_quot_id) fact finally show ?thesis . qed lemma td_normalized_fract: "type_definition quot_of_fract quot_to_fract normalized_fracts" by standard (simp_all add: quot_of_fract_quot_to_fract') lemma quot_of_fract_add_aux: assumes "snd x \ 0" "snd y \ 0" shows "(fst x * snd y + fst y * snd x) * (snd (normalize_quot x) * snd (normalize_quot y)) = snd x * snd y * (fst (normalize_quot x) * snd (normalize_quot y) + snd (normalize_quot x) * fst (normalize_quot y))" proof - from normalize_quotE'[OF assms(1)] guess d . note d = this from normalize_quotE'[OF assms(2)] guess e . note e = this show ?thesis by (simp_all add: d e algebra_simps) qed locale fract_as_normalized_quot begin setup_lifting td_normalized_fract end lemma quot_of_fract_add: "quot_of_fract (x + y) = (let (a,b) = quot_of_fract x; (c,d) = quot_of_fract y in normalize_quot (a * d + b * c, b * d))" by transfer (insert quot_of_fract_add_aux, simp_all add: Let_def case_prod_unfold normalize_quot_eq_iff) lemma quot_of_fract_uminus: "quot_of_fract (-x) = (let (a,b) = quot_of_fract x in (-a, b))" by transfer (auto simp: case_prod_unfold Let_def normalize_quot_def dvd_neg_div mult_unit_dvd_iff) lemma quot_of_fract_diff: "quot_of_fract (x - y) = (let (a,b) = quot_of_fract x; (c,d) = quot_of_fract y in normalize_quot (a * d - b * c, b * d))" (is "_ = ?rhs") proof - have "x - y = x + -y" by simp also have "quot_of_fract \ = ?rhs" by (simp only: quot_of_fract_add quot_of_fract_uminus Let_def case_prod_unfold) simp_all finally show ?thesis . qed lemma normalize_quot_mult_coprime: assumes "coprime a b" "coprime c d" "unit_factor b = 1" "unit_factor d = 1" defines "e \ fst (normalize_quot (a, d))" and "f \ snd (normalize_quot (a, d))" and "g \ fst (normalize_quot (c, b))" and "h \ snd (normalize_quot (c, b))" shows "normalize_quot (a * c, b * d) = (e * g, f * h)" proof (rule normalize_quotI) from assms have "gcd a b = 1" "gcd c d = 1" by simp_all from assms have "b \ 0" "d \ 0" by auto with assms have "normalize b = b" "normalize d = d" by (auto intro: normalize_unit_factor_eqI) from normalize_quotE [OF \b \ 0\, of c] guess k . note k = this [folded \gcd a b = 1\ \gcd c d = 1\ assms(3) assms(4)] from normalize_quotE [OF \d \ 0\, of a] guess l . note l = this [folded \gcd a b = 1\ \gcd c d = 1\ assms(3) assms(4)] from k l show "a * c * (f * h) = b * d * (e * g)" by (metis e_def f_def g_def h_def mult.commute mult.left_commute) from assms have [simp]: "unit_factor f = 1" "unit_factor h = 1" by simp_all from assms have "coprime e f" "coprime g h" by (simp_all add: coprime_normalize_quot) with k l assms(1,2) \b \ 0\ \d \ 0\ \unit_factor b = 1\ \unit_factor d = 1\ \normalize b = b\ \normalize d = d\ show "(e * g, f * h) \ normalized_fracts" by (simp add: normalized_fracts_def unit_factor_mult e_def f_def g_def h_def coprime_normalize_quot dvd_unit_factor_div unit_factor_gcd) (metis coprime_mult_left_iff coprime_mult_right_iff) qed (insert assms(3,4), auto) lemma normalize_quot_mult: assumes "snd x \ 0" "snd y \ 0" shows "normalize_quot (fst x * fst y, snd x * snd y) = normalize_quot (fst (normalize_quot x) * fst (normalize_quot y), snd (normalize_quot x) * snd (normalize_quot y))" proof - from normalize_quotE'[OF assms(1)] guess d . note d = this from normalize_quotE'[OF assms(2)] guess e . note e = this show ?thesis by (simp_all add: d e algebra_simps normalize_quot_eq_iff) qed lemma quot_of_fract_mult: "quot_of_fract (x * y) = (let (a,b) = quot_of_fract x; (c,d) = quot_of_fract y; (e,f) = normalize_quot (a,d); (g,h) = normalize_quot (c,b) in (e*g, f*h))" by transfer (simp add: split_def Let_def coprime_normalize_quot normalize_quot_mult normalize_quot_mult_coprime) lemma normalize_quot_0 [simp]: "normalize_quot (0, x) = (0, 1)" "normalize_quot (x, 0) = (0, 1)" by (simp_all add: normalize_quot_def) lemma normalize_quot_eq_0_iff [simp]: "fst (normalize_quot x) = 0 \ fst x = 0 \ snd x = 0" by (auto simp: normalize_quot_def case_prod_unfold Let_def div_mult_unit2 dvd_div_eq_0_iff) lemma fst_quot_of_fract_0_imp: "fst (quot_of_fract x) = 0 \ snd (quot_of_fract x) = 1" by transfer auto lemma normalize_quot_swap: assumes "a \ 0" "b \ 0" defines "a' \ fst (normalize_quot (a, b))" and "b' \ snd (normalize_quot (a, b))" shows "normalize_quot (b, a) = (b' div unit_factor a', a' div unit_factor a')" proof (rule normalize_quotI) from normalize_quotE[OF assms(2), of a] guess d . note d = this [folded assms(3,4)] show "b * (a' div unit_factor a') = a * (b' div unit_factor a')" using assms(1,2) d by (simp add: div_unit_factor [symmetric] unit_div_mult_swap mult_ac del: div_unit_factor) have "coprime a' b'" by (simp add: a'_def b'_def coprime_normalize_quot) thus "(b' div unit_factor a', a' div unit_factor a') \ normalized_fracts" using assms(1,2) d by (auto simp add: normalized_fracts_def ac_simps dvd_div_unit_iff elim: coprime_imp_coprime) qed fact+ lemma quot_of_fract_inverse: "quot_of_fract (inverse x) = (let (a,b) = quot_of_fract x; d = unit_factor a in if d = 0 then (0, 1) else (b div d, a div d))" proof (transfer, goal_cases) case (1 x) from normalize_quot_swap[of "fst x" "snd x"] show ?case by (auto simp: Let_def case_prod_unfold) qed lemma normalize_quot_div_unit_left: fixes x y u assumes "is_unit u" defines "x' \ fst (normalize_quot (x, y))" and "y' \ snd (normalize_quot (x, y))" shows "normalize_quot (x div u, y) = (x' div u, y')" proof (cases "y = 0") case False define v where "v = 1 div u" with \is_unit u\ have "is_unit v" and u: "\a. a div u = a * v" by simp_all from \is_unit v\ have "coprime v = top" by (simp add: fun_eq_iff is_unit_left_imp_coprime) from normalize_quotE[OF False, of x] guess d . note d = this[folded assms(2,3)] from assms have "coprime x' y'" "unit_factor y' = 1" by (simp_all add: coprime_normalize_quot) with d \coprime v = top\ have "normalize_quot (x * v, y) = (x' * v, y')" by (auto simp: normalized_fracts_def intro: normalize_quotI) then show ?thesis by (simp add: u) qed (simp_all add: assms) lemma normalize_quot_div_unit_right: fixes x y u assumes "is_unit u" defines "x' \ fst (normalize_quot (x, y))" and "y' \ snd (normalize_quot (x, y))" shows "normalize_quot (x, y div u) = (x' * u, y')" proof (cases "y = 0") case False from normalize_quotE[OF this, of x] guess d . note d = this[folded assms(2,3)] from assms have "coprime x' y'" "unit_factor y' = 1" by (simp_all add: coprime_normalize_quot) with d \is_unit u\ show ?thesis by (auto simp add: normalized_fracts_def is_unit_left_imp_coprime unit_div_eq_0_iff intro: normalize_quotI) qed (simp_all add: assms) lemma normalize_quot_normalize_left: fixes x y u defines "x' \ fst (normalize_quot (x, y))" and "y' \ snd (normalize_quot (x, y))" shows "normalize_quot (normalize x, y) = (x' div unit_factor x, y')" using normalize_quot_div_unit_left[of "unit_factor x" x y] by (cases "x = 0") (simp_all add: assms) lemma normalize_quot_normalize_right: fixes x y u defines "x' \ fst (normalize_quot (x, y))" and "y' \ snd (normalize_quot (x, y))" shows "normalize_quot (x, normalize y) = (x' * unit_factor y, y')" using normalize_quot_div_unit_right[of "unit_factor y" x y] by (cases "y = 0") (simp_all add: assms) lemma quot_of_fract_0 [simp]: "quot_of_fract 0 = (0, 1)" by transfer auto lemma quot_of_fract_1 [simp]: "quot_of_fract 1 = (1, 1)" by transfer (rule normalize_quotI, simp_all add: normalized_fracts_def) lemma quot_of_fract_divide: "quot_of_fract (x / y) = (if y = 0 then (0, 1) else (let (a,b) = quot_of_fract x; (c,d) = quot_of_fract y; (e,f) = normalize_quot (a,c); (g,h) = normalize_quot (d,b) in (e * g, f * h)))" (is "_ = ?rhs") proof (cases "y = 0") case False hence A: "fst (quot_of_fract y) \ 0" by transfer auto have "x / y = x * inverse y" by (simp add: divide_inverse) also from False A have "quot_of_fract \ = ?rhs" by (simp only: quot_of_fract_mult quot_of_fract_inverse) (simp_all add: Let_def case_prod_unfold fst_quot_of_fract_0_imp normalize_quot_div_unit_left normalize_quot_div_unit_right normalize_quot_normalize_right normalize_quot_normalize_left) finally show ?thesis . qed simp_all end diff --git a/src/HOL/Computational_Algebra/Nth_Powers.thy b/src/HOL/Computational_Algebra/Nth_Powers.thy --- a/src/HOL/Computational_Algebra/Nth_Powers.thy +++ b/src/HOL/Computational_Algebra/Nth_Powers.thy @@ -1,308 +1,308 @@ (* File: HOL/Computational_Algebra/Nth_Powers.thy Author: Manuel Eberl n-th powers in general and n-th roots of natural numbers *) section \$n$-th powers and roots of naturals\ theory Nth_Powers imports Primes begin subsection \The set of $n$-th powers\ definition is_nth_power :: "nat \ 'a :: monoid_mult \ bool" where "is_nth_power n x \ (\y. x = y ^ n)" lemma is_nth_power_nth_power [simp, intro]: "is_nth_power n (x ^ n)" by (auto simp add: is_nth_power_def) lemma is_nth_powerI [intro?]: "x = y ^ n \ is_nth_power n x" by (auto simp: is_nth_power_def) lemma is_nth_powerE: "is_nth_power n x \ (\y. x = y ^ n \ P) \ P" by (auto simp: is_nth_power_def) abbreviation is_square where "is_square \ is_nth_power 2" lemma is_zeroth_power [simp]: "is_nth_power 0 x \ x = 1" by (simp add: is_nth_power_def) lemma is_first_power [simp]: "is_nth_power 1 x" by (simp add: is_nth_power_def) lemma is_first_power' [simp]: "is_nth_power (Suc 0) x" by (simp add: is_nth_power_def) lemma is_nth_power_0 [simp]: "n > 0 \ is_nth_power n (0 :: 'a :: semiring_1)" by (auto simp: is_nth_power_def power_0_left intro!: exI[of _ 0]) lemma is_nth_power_0_iff [simp]: "is_nth_power n (0 :: 'a :: semiring_1) \ n > 0" by (cases n) auto lemma is_nth_power_1 [simp]: "is_nth_power n 1" by (auto simp: is_nth_power_def intro!: exI[of _ 1]) lemma is_nth_power_Suc_0 [simp]: "is_nth_power n (Suc 0)" by (simp add: One_nat_def [symmetric] del: One_nat_def) lemma is_nth_power_conv_multiplicity: - fixes x :: "'a :: factorial_semiring" + fixes x :: "'a :: {factorial_semiring, normalization_semidom_multiplicative}" assumes "n > 0" shows "is_nth_power n (normalize x) \ (\p. prime p \ n dvd multiplicity p x)" proof (cases "x = 0") case False show ?thesis proof (safe intro!: is_nth_powerI elim!: is_nth_powerE) fix y p :: 'a assume *: "normalize x = y ^ n" "prime p" with assms and False have [simp]: "y \ 0" by (auto simp: power_0_left) have "multiplicity p x = multiplicity p (y ^ n)" by (subst *(1) [symmetric]) simp with False and * and assms show "n dvd multiplicity p x" by (auto simp: prime_elem_multiplicity_power_distrib) next assume *: "\p. prime p \ n dvd multiplicity p x" have "multiplicity p ((\p\prime_factors x. p ^ (multiplicity p x div n)) ^ n) = multiplicity p x" if "prime p" for p proof - from that and * have "n dvd multiplicity p x" by blast have "multiplicity p x = 0" if "p \ prime_factors x" using that and \prime p\ by (simp add: prime_factors_multiplicity) with that and * and assms show ?thesis unfolding prod_power_distrib power_mult [symmetric] by (subst multiplicity_prod_prime_powers) (auto simp: in_prime_factors_imp_prime elim: dvdE) qed with assms False have "normalize x = normalize ((\p\prime_factors x. p ^ (multiplicity p x div n)) ^ n)" by (intro multiplicity_eq_imp_eq) (auto simp: multiplicity_prod_prime_powers) thus "normalize x = normalize (\p\prime_factors x. p ^ (multiplicity p x div n)) ^ n" by (simp add: normalize_power) qed qed (insert assms, auto) lemma is_nth_power_conv_multiplicity_nat: assumes "n > 0" shows "is_nth_power n (x :: nat) \ (\p. prime p \ n dvd multiplicity p x)" using is_nth_power_conv_multiplicity[OF assms, of x] by simp lemma is_nth_power_mult: assumes "is_nth_power n a" "is_nth_power n b" shows "is_nth_power n (a * b :: 'a :: comm_monoid_mult)" proof - from assms obtain a' b' where "a = a' ^ n" "b = b' ^ n" by (auto elim!: is_nth_powerE) hence "a * b = (a' * b') ^ n" by (simp add: power_mult_distrib) thus ?thesis by (rule is_nth_powerI) qed lemma is_nth_power_mult_coprime_natD: fixes a b :: nat assumes "coprime a b" "is_nth_power n (a * b)" "a > 0" "b > 0" shows "is_nth_power n a" "is_nth_power n b" proof - have A: "is_nth_power n a" if "coprime a b" "is_nth_power n (a * b)" "a \ 0" "b \ 0" "n > 0" for a b :: nat unfolding is_nth_power_conv_multiplicity_nat[OF \n > 0\] proof safe fix p :: nat assume p: "prime p" from \coprime a b\ have "\(p dvd a \ p dvd b)" using coprime_common_divisor_nat[of a b p] p by auto moreover from that and p have "n dvd multiplicity p a + multiplicity p b" by (auto simp: is_nth_power_conv_multiplicity_nat prime_elem_multiplicity_mult_distrib) ultimately show "n dvd multiplicity p a" by (auto simp: not_dvd_imp_multiplicity_0) qed from A [of a b] assms show "is_nth_power n a" by (cases "n = 0") simp_all from A [of b a] assms show "is_nth_power n b" by (cases "n = 0") (simp_all add: ac_simps) qed lemma is_nth_power_mult_coprime_nat_iff: fixes a b :: nat assumes "coprime a b" shows "is_nth_power n (a * b) \ is_nth_power n a \is_nth_power n b" using assms by (cases "a = 0"; cases "b = 0") (auto intro: is_nth_power_mult dest: is_nth_power_mult_coprime_natD[of a b n] simp del: One_nat_def) lemma is_nth_power_prime_power_nat_iff: fixes p :: nat assumes "prime p" shows "is_nth_power n (p ^ k) \ n dvd k" using assms by (cases "n > 0") (auto simp: is_nth_power_conv_multiplicity_nat prime_elem_multiplicity_power_distrib) lemma is_nth_power_nth_power': assumes "n dvd n'" shows "is_nth_power n (m ^ n')" proof - from assms have "n' = n' div n * n" by simp also have "m ^ \ = (m ^ (n' div n)) ^ n" by (simp add: power_mult) also have "is_nth_power n \" by simp finally show ?thesis . qed definition is_nth_power_nat :: "nat \ nat \ bool" where [code_abbrev]: "is_nth_power_nat = is_nth_power" lemma is_nth_power_nat_code [code]: "is_nth_power_nat n m = (if n = 0 then m = 1 else if m = 0 then n > 0 else if n = 1 then True else (\k\{1..m}. k ^ n = m))" by (auto simp: is_nth_power_nat_def is_nth_power_def power_eq_iff_eq_base self_le_power) (* TODO: Harmonise with Discrete.sqrt *) subsection \The $n$-root of a natural number\ definition nth_root_nat :: "nat \ nat \ nat" where "nth_root_nat k n = (if k = 0 then 0 else Max {m. m ^ k \ n})" lemma zeroth_root_nat [simp]: "nth_root_nat 0 n = 0" by (simp add: nth_root_nat_def) lemma nth_root_nat_aux1: assumes "k > 0" shows "{m::nat. m ^ k \ n} \ {..n}" proof safe fix m assume "m ^ k \ n" show "m \ n" proof (cases "m = 0") case False with assms have "m ^ 1 \ m ^ k" by (intro power_increasing) simp_all also note \m ^ k \ n\ finally show ?thesis by simp qed simp_all qed lemma nth_root_nat_aux2: assumes "k > 0" shows "finite {m::nat. m ^ k \ n}" "{m::nat. m ^ k \ n} \ {}" proof - from assms have "{m. m ^ k \ n} \ {..n}" by (rule nth_root_nat_aux1) moreover have "finite {..n}" by simp ultimately show "finite {m::nat. m ^ k \ n}" by (rule finite_subset) next from assms show "{m::nat. m ^ k \ n} \ {}" by (auto intro!: exI[of _ 0] simp: power_0_left) qed lemma assumes "k > 0" shows nth_root_nat_power_le: "nth_root_nat k n ^ k \ n" and nth_root_nat_ge: "x ^ k \ n \ x \ nth_root_nat k n" using Max_in[OF nth_root_nat_aux2[OF assms], of n] Max_ge[OF nth_root_nat_aux2(1)[OF assms], of x n] assms by (auto simp: nth_root_nat_def) lemma nth_root_nat_less: assumes "k > 0" "x ^ k > n" shows "nth_root_nat k n < x" proof - from \k > 0\ have "nth_root_nat k n ^ k \ n" by (rule nth_root_nat_power_le) also have "n < x ^ k" by fact finally show ?thesis by (rule power_less_imp_less_base) simp_all qed lemma nth_root_nat_unique: assumes "m ^ k \ n" "(m + 1) ^ k > n" shows "nth_root_nat k n = m" proof (cases "k > 0") case True from nth_root_nat_less[OF \k > 0\ assms(2)] have "nth_root_nat k n \ m" by simp moreover from \k > 0\ and assms(1) have "nth_root_nat k n \ m" by (intro nth_root_nat_ge) ultimately show ?thesis by (rule antisym) qed (insert assms, auto) lemma nth_root_nat_0 [simp]: "nth_root_nat k 0 = 0" by (simp add: nth_root_nat_def) lemma nth_root_nat_1 [simp]: "k > 0 \ nth_root_nat k 1 = 1" by (rule nth_root_nat_unique) (auto simp del: One_nat_def) lemma nth_root_nat_Suc_0 [simp]: "k > 0 \ nth_root_nat k (Suc 0) = Suc 0" using nth_root_nat_1 by (simp del: nth_root_nat_1) lemma first_root_nat [simp]: "nth_root_nat 1 n = n" by (intro nth_root_nat_unique) auto lemma first_root_nat' [simp]: "nth_root_nat (Suc 0) n = n" by (intro nth_root_nat_unique) auto lemma nth_root_nat_code_naive': "nth_root_nat k n = (if k = 0 then 0 else Max (Set.filter (\m. m ^ k \ n) {..n}))" proof (cases "k > 0") case True hence "{m. m ^ k \ n} \ {..n}" by (rule nth_root_nat_aux1) hence "Set.filter (\m. m ^ k \ n) {..n} = {m. m ^ k \ n}" by (auto simp: Set.filter_def) with True show ?thesis by (simp add: nth_root_nat_def Set.filter_def) qed simp function nth_root_nat_aux :: "nat \ nat \ nat \ nat \ nat" where "nth_root_nat_aux m k acc n = (let acc' = (k + 1) ^ m in if k \ n \ acc' > n then k else nth_root_nat_aux m (k+1) acc' n)" by auto termination by (relation "measure (\(_,k,_,n). n - k)", goal_cases) auto lemma nth_root_nat_aux_le: assumes "k ^ m \ n" "m > 0" shows "nth_root_nat_aux m k (k ^ m) n ^ m \ n" using assms by (induction m k "k ^ m" n rule: nth_root_nat_aux.induct) (auto simp: Let_def) lemma nth_root_nat_aux_gt: assumes "m > 0" shows "(nth_root_nat_aux m k (k ^ m) n + 1) ^ m > n" using assms proof (induction m k "k ^ m" n rule: nth_root_nat_aux.induct) case (1 m k n) have "n < Suc k ^ m" if "n \ k" proof - note that also have "k < Suc k ^ 1" by simp also from \m > 0\ have "\ \ Suc k ^ m" by (intro power_increasing) simp_all finally show ?thesis . qed with 1 show ?case by (auto simp: Let_def) qed lemma nth_root_nat_aux_correct: assumes "k ^ m \ n" "m > 0" shows "nth_root_nat_aux m k (k ^ m) n = nth_root_nat m n" by (rule sym, intro nth_root_nat_unique nth_root_nat_aux_le nth_root_nat_aux_gt assms) lemma nth_root_nat_naive_code [code]: "nth_root_nat m n = (if m = 0 \ n = 0 then 0 else if m = 1 \ n = 1 then n else nth_root_nat_aux m 1 1 n)" using nth_root_nat_aux_correct[of 1 m n] by (auto simp: ) lemma nth_root_nat_nth_power [simp]: "k > 0 \ nth_root_nat k (n ^ k) = n" by (intro nth_root_nat_unique order.refl power_strict_mono) simp_all lemma nth_root_nat_nth_power': assumes "k > 0" "k dvd m" shows "nth_root_nat k (n ^ m) = n ^ (m div k)" proof - from assms have "m = (m div k) * k" by simp also have "n ^ \ = (n ^ (m div k)) ^ k" by (simp add: power_mult) also from assms have "nth_root_nat k \ = n ^ (m div k)" by simp finally show ?thesis . qed lemma nth_root_nat_mono: assumes "m \ n" shows "nth_root_nat k m \ nth_root_nat k n" proof (cases "k = 0") case False with assms show ?thesis unfolding nth_root_nat_def using nth_root_nat_aux2[of k m] nth_root_nat_aux2[of k n] by (auto intro!: Max_mono) qed auto end diff --git a/src/HOL/Computational_Algebra/Polynomial.thy b/src/HOL/Computational_Algebra/Polynomial.thy --- a/src/HOL/Computational_Algebra/Polynomial.thy +++ b/src/HOL/Computational_Algebra/Polynomial.thy @@ -1,4699 +1,4717 @@ (* Title: HOL/Computational_Algebra/Polynomial.thy Author: Brian Huffman Author: Clemens Ballarin Author: Amine Chaieb Author: Florian Haftmann *) section \Polynomials as type over a ring structure\ theory Polynomial imports Complex_Main "HOL-Library.More_List" "HOL-Library.Infinite_Set" Factorial_Ring begin subsection \Auxiliary: operations for lists (later) representing coefficients\ definition cCons :: "'a::zero \ 'a list \ 'a list" (infixr "##" 65) where "x ## xs = (if xs = [] \ x = 0 then [] else x # xs)" lemma cCons_0_Nil_eq [simp]: "0 ## [] = []" by (simp add: cCons_def) lemma cCons_Cons_eq [simp]: "x ## y # ys = x # y # ys" by (simp add: cCons_def) lemma cCons_append_Cons_eq [simp]: "x ## xs @ y # ys = x # xs @ y # ys" by (simp add: cCons_def) lemma cCons_not_0_eq [simp]: "x \ 0 \ x ## xs = x # xs" by (simp add: cCons_def) lemma strip_while_not_0_Cons_eq [simp]: "strip_while (\x. x = 0) (x # xs) = x ## strip_while (\x. x = 0) xs" proof (cases "x = 0") case False then show ?thesis by simp next case True show ?thesis proof (induct xs rule: rev_induct) case Nil with True show ?case by simp next case (snoc y ys) then show ?case by (cases "y = 0") (simp_all add: append_Cons [symmetric] del: append_Cons) qed qed lemma tl_cCons [simp]: "tl (x ## xs) = xs" by (simp add: cCons_def) subsection \Definition of type \poly\\ typedef (overloaded) 'a poly = "{f :: nat \ 'a::zero. \\<^sub>\ n. f n = 0}" morphisms coeff Abs_poly by (auto intro!: ALL_MOST) setup_lifting type_definition_poly lemma poly_eq_iff: "p = q \ (\n. coeff p n = coeff q n)" by (simp add: coeff_inject [symmetric] fun_eq_iff) lemma poly_eqI: "(\n. coeff p n = coeff q n) \ p = q" by (simp add: poly_eq_iff) lemma MOST_coeff_eq_0: "\\<^sub>\ n. coeff p n = 0" using coeff [of p] by simp subsection \Degree of a polynomial\ definition degree :: "'a::zero poly \ nat" where "degree p = (LEAST n. \i>n. coeff p i = 0)" lemma coeff_eq_0: assumes "degree p < n" shows "coeff p n = 0" proof - have "\n. \i>n. coeff p i = 0" using MOST_coeff_eq_0 by (simp add: MOST_nat) then have "\i>degree p. coeff p i = 0" unfolding degree_def by (rule LeastI_ex) with assms show ?thesis by simp qed lemma le_degree: "coeff p n \ 0 \ n \ degree p" by (erule contrapos_np, rule coeff_eq_0, simp) lemma degree_le: "\i>n. coeff p i = 0 \ degree p \ n" unfolding degree_def by (erule Least_le) lemma less_degree_imp: "n < degree p \ \i>n. coeff p i \ 0" unfolding degree_def by (drule not_less_Least, simp) subsection \The zero polynomial\ instantiation poly :: (zero) zero begin lift_definition zero_poly :: "'a poly" is "\_. 0" by (rule MOST_I) simp instance .. end lemma coeff_0 [simp]: "coeff 0 n = 0" by transfer rule lemma degree_0 [simp]: "degree 0 = 0" by (rule order_antisym [OF degree_le le0]) simp lemma leading_coeff_neq_0: assumes "p \ 0" shows "coeff p (degree p) \ 0" proof (cases "degree p") case 0 from \p \ 0\ obtain n where "coeff p n \ 0" by (auto simp add: poly_eq_iff) then have "n \ degree p" by (rule le_degree) with \coeff p n \ 0\ and \degree p = 0\ show "coeff p (degree p) \ 0" by simp next case (Suc n) from \degree p = Suc n\ have "n < degree p" by simp then have "\i>n. coeff p i \ 0" by (rule less_degree_imp) then obtain i where "n < i" and "coeff p i \ 0" by blast from \degree p = Suc n\ and \n < i\ have "degree p \ i" by simp also from \coeff p i \ 0\ have "i \ degree p" by (rule le_degree) finally have "degree p = i" . with \coeff p i \ 0\ show "coeff p (degree p) \ 0" by simp qed lemma leading_coeff_0_iff [simp]: "coeff p (degree p) = 0 \ p = 0" by (cases "p = 0") (simp_all add: leading_coeff_neq_0) lemma eq_zero_or_degree_less: assumes "degree p \ n" and "coeff p n = 0" shows "p = 0 \ degree p < n" proof (cases n) case 0 with \degree p \ n\ and \coeff p n = 0\ have "coeff p (degree p) = 0" by simp then have "p = 0" by simp then show ?thesis .. next case (Suc m) from \degree p \ n\ have "\i>n. coeff p i = 0" by (simp add: coeff_eq_0) with \coeff p n = 0\ have "\i\n. coeff p i = 0" by (simp add: le_less) with \n = Suc m\ have "\i>m. coeff p i = 0" by (simp add: less_eq_Suc_le) then have "degree p \ m" by (rule degree_le) with \n = Suc m\ have "degree p < n" by (simp add: less_Suc_eq_le) then show ?thesis .. qed lemma coeff_0_degree_minus_1: "coeff rrr dr = 0 \ degree rrr \ dr \ degree rrr \ dr - 1" using eq_zero_or_degree_less by fastforce subsection \List-style constructor for polynomials\ lift_definition pCons :: "'a::zero \ 'a poly \ 'a poly" is "\a p. case_nat a (coeff p)" by (rule MOST_SucD) (simp add: MOST_coeff_eq_0) lemmas coeff_pCons = pCons.rep_eq lemma coeff_pCons_0 [simp]: "coeff (pCons a p) 0 = a" by transfer simp lemma coeff_pCons_Suc [simp]: "coeff (pCons a p) (Suc n) = coeff p n" by (simp add: coeff_pCons) lemma degree_pCons_le: "degree (pCons a p) \ Suc (degree p)" by (rule degree_le) (simp add: coeff_eq_0 coeff_pCons split: nat.split) lemma degree_pCons_eq: "p \ 0 \ degree (pCons a p) = Suc (degree p)" apply (rule order_antisym [OF degree_pCons_le]) apply (rule le_degree, simp) done lemma degree_pCons_0: "degree (pCons a 0) = 0" apply (rule order_antisym [OF _ le0]) apply (rule degree_le, simp add: coeff_pCons split: nat.split) done lemma degree_pCons_eq_if [simp]: "degree (pCons a p) = (if p = 0 then 0 else Suc (degree p))" apply (cases "p = 0", simp_all) apply (rule order_antisym [OF _ le0]) apply (rule degree_le, simp add: coeff_pCons split: nat.split) apply (rule order_antisym [OF degree_pCons_le]) apply (rule le_degree, simp) done lemma pCons_0_0 [simp]: "pCons 0 0 = 0" by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) lemma pCons_eq_iff [simp]: "pCons a p = pCons b q \ a = b \ p = q" proof safe assume "pCons a p = pCons b q" then have "coeff (pCons a p) 0 = coeff (pCons b q) 0" by simp then show "a = b" by simp next assume "pCons a p = pCons b q" then have "coeff (pCons a p) (Suc n) = coeff (pCons b q) (Suc n)" for n by simp then show "p = q" by (simp add: poly_eq_iff) qed lemma pCons_eq_0_iff [simp]: "pCons a p = 0 \ a = 0 \ p = 0" using pCons_eq_iff [of a p 0 0] by simp lemma pCons_cases [cases type: poly]: obtains (pCons) a q where "p = pCons a q" proof show "p = pCons (coeff p 0) (Abs_poly (\n. coeff p (Suc n)))" by transfer (simp_all add: MOST_inj[where f=Suc and P="\n. p n = 0" for p] fun_eq_iff Abs_poly_inverse split: nat.split) qed lemma pCons_induct [case_names 0 pCons, induct type: poly]: assumes zero: "P 0" assumes pCons: "\a p. a \ 0 \ p \ 0 \ P p \ P (pCons a p)" shows "P p" proof (induct p rule: measure_induct_rule [where f=degree]) case (less p) obtain a q where "p = pCons a q" by (rule pCons_cases) have "P q" proof (cases "q = 0") case True then show "P q" by (simp add: zero) next case False then have "degree (pCons a q) = Suc (degree q)" by (rule degree_pCons_eq) with \p = pCons a q\ have "degree q < degree p" by simp then show "P q" by (rule less.hyps) qed have "P (pCons a q)" proof (cases "a \ 0 \ q \ 0") case True with \P q\ show ?thesis by (auto intro: pCons) next case False with zero show ?thesis by simp qed with \p = pCons a q\ show ?case by simp qed lemma degree_eq_zeroE: fixes p :: "'a::zero poly" assumes "degree p = 0" obtains a where "p = pCons a 0" proof - obtain a q where p: "p = pCons a q" by (cases p) with assms have "q = 0" by (cases "q = 0") simp_all with p have "p = pCons a 0" by simp then show thesis .. qed subsection \Quickcheck generator for polynomials\ quickcheck_generator poly constructors: "0 :: _ poly", pCons subsection \List-style syntax for polynomials\ syntax "_poly" :: "args \ 'a poly" ("[:(_):]") translations "[:x, xs:]" \ "CONST pCons x [:xs:]" "[:x:]" \ "CONST pCons x 0" "[:x:]" \ "CONST pCons x (_constrain 0 t)" subsection \Representation of polynomials by lists of coefficients\ primrec Poly :: "'a::zero list \ 'a poly" where [code_post]: "Poly [] = 0" | [code_post]: "Poly (a # as) = pCons a (Poly as)" lemma Poly_replicate_0 [simp]: "Poly (replicate n 0) = 0" by (induct n) simp_all lemma Poly_eq_0: "Poly as = 0 \ (\n. as = replicate n 0)" by (induct as) (auto simp add: Cons_replicate_eq) lemma Poly_append_replicate_zero [simp]: "Poly (as @ replicate n 0) = Poly as" by (induct as) simp_all lemma Poly_snoc_zero [simp]: "Poly (as @ [0]) = Poly as" using Poly_append_replicate_zero [of as 1] by simp lemma Poly_cCons_eq_pCons_Poly [simp]: "Poly (a ## p) = pCons a (Poly p)" by (simp add: cCons_def) lemma Poly_on_rev_starting_with_0 [simp]: "hd as = 0 \ Poly (rev (tl as)) = Poly (rev as)" by (cases as) simp_all lemma degree_Poly: "degree (Poly xs) \ length xs" by (induct xs) simp_all lemma coeff_Poly_eq [simp]: "coeff (Poly xs) = nth_default 0 xs" by (induct xs) (simp_all add: fun_eq_iff coeff_pCons split: nat.splits) definition coeffs :: "'a poly \ 'a::zero list" where "coeffs p = (if p = 0 then [] else map (\i. coeff p i) [0 ..< Suc (degree p)])" lemma coeffs_eq_Nil [simp]: "coeffs p = [] \ p = 0" by (simp add: coeffs_def) lemma not_0_coeffs_not_Nil: "p \ 0 \ coeffs p \ []" by simp lemma coeffs_0_eq_Nil [simp]: "coeffs 0 = []" by simp lemma coeffs_pCons_eq_cCons [simp]: "coeffs (pCons a p) = a ## coeffs p" proof - have *: "\m\set ms. m > 0 \ map (case_nat x f) ms = map f (map (\n. n - 1) ms)" for ms :: "nat list" and f :: "nat \ 'a" and x :: "'a" by (induct ms) (auto split: nat.split) show ?thesis by (simp add: * coeffs_def upt_conv_Cons coeff_pCons map_decr_upt del: upt_Suc) qed lemma length_coeffs: "p \ 0 \ length (coeffs p) = degree p + 1" by (simp add: coeffs_def) lemma coeffs_nth: "p \ 0 \ n \ degree p \ coeffs p ! n = coeff p n" by (auto simp: coeffs_def simp del: upt_Suc) lemma coeff_in_coeffs: "p \ 0 \ n \ degree p \ coeff p n \ set (coeffs p)" using coeffs_nth [of p n, symmetric] by (simp add: length_coeffs) lemma not_0_cCons_eq [simp]: "p \ 0 \ a ## coeffs p = a # coeffs p" by (simp add: cCons_def) lemma Poly_coeffs [simp, code abstype]: "Poly (coeffs p) = p" by (induct p) auto lemma coeffs_Poly [simp]: "coeffs (Poly as) = strip_while (HOL.eq 0) as" proof (induct as) case Nil then show ?case by simp next case (Cons a as) from replicate_length_same [of as 0] have "(\n. as \ replicate n 0) \ (\a\set as. a \ 0)" by (auto dest: sym [of _ as]) with Cons show ?case by auto qed lemma no_trailing_coeffs [simp]: "no_trailing (HOL.eq 0) (coeffs p)" by (induct p) auto lemma strip_while_coeffs [simp]: "strip_while (HOL.eq 0) (coeffs p) = coeffs p" by simp lemma coeffs_eq_iff: "p = q \ coeffs p = coeffs q" (is "?P \ ?Q") proof assume ?P then show ?Q by simp next assume ?Q then have "Poly (coeffs p) = Poly (coeffs q)" by simp then show ?P by simp qed lemma nth_default_coeffs_eq: "nth_default 0 (coeffs p) = coeff p" by (simp add: fun_eq_iff coeff_Poly_eq [symmetric]) lemma [code]: "coeff p = nth_default 0 (coeffs p)" by (simp add: nth_default_coeffs_eq) lemma coeffs_eqI: assumes coeff: "\n. coeff p n = nth_default 0 xs n" assumes zero: "no_trailing (HOL.eq 0) xs" shows "coeffs p = xs" proof - from coeff have "p = Poly xs" by (simp add: poly_eq_iff) with zero show ?thesis by simp qed lemma degree_eq_length_coeffs [code]: "degree p = length (coeffs p) - 1" by (simp add: coeffs_def) lemma length_coeffs_degree: "p \ 0 \ length (coeffs p) = Suc (degree p)" by (induct p) (auto simp: cCons_def) lemma [code abstract]: "coeffs 0 = []" by (fact coeffs_0_eq_Nil) lemma [code abstract]: "coeffs (pCons a p) = a ## coeffs p" by (fact coeffs_pCons_eq_cCons) lemma set_coeffs_subset_singleton_0_iff [simp]: "set (coeffs p) \ {0} \ p = 0" by (auto simp add: coeffs_def intro: classical) lemma set_coeffs_not_only_0 [simp]: "set (coeffs p) \ {0}" by (auto simp add: set_eq_subset) lemma forall_coeffs_conv: "(\n. P (coeff p n)) \ (\c \ set (coeffs p). P c)" if "P 0" using that by (auto simp add: coeffs_def) (metis atLeastLessThan_iff coeff_eq_0 not_less_iff_gr_or_eq zero_le) instantiation poly :: ("{zero, equal}") equal begin definition [code]: "HOL.equal (p::'a poly) q \ HOL.equal (coeffs p) (coeffs q)" instance by standard (simp add: equal equal_poly_def coeffs_eq_iff) end lemma [code nbe]: "HOL.equal (p :: _ poly) p \ True" by (fact equal_refl) definition is_zero :: "'a::zero poly \ bool" where [code]: "is_zero p \ List.null (coeffs p)" lemma is_zero_null [code_abbrev]: "is_zero p \ p = 0" by (simp add: is_zero_def null_def) subsubsection \Reconstructing the polynomial from the list\ \ \contributed by Sebastiaan J.C. Joosten and René Thiemann\ definition poly_of_list :: "'a::comm_monoid_add list \ 'a poly" where [simp]: "poly_of_list = Poly" lemma poly_of_list_impl [code abstract]: "coeffs (poly_of_list as) = strip_while (HOL.eq 0) as" by simp subsection \Fold combinator for polynomials\ definition fold_coeffs :: "('a::zero \ 'b \ 'b) \ 'a poly \ 'b \ 'b" where "fold_coeffs f p = foldr f (coeffs p)" lemma fold_coeffs_0_eq [simp]: "fold_coeffs f 0 = id" by (simp add: fold_coeffs_def) lemma fold_coeffs_pCons_eq [simp]: "f 0 = id \ fold_coeffs f (pCons a p) = f a \ fold_coeffs f p" by (simp add: fold_coeffs_def cCons_def fun_eq_iff) lemma fold_coeffs_pCons_0_0_eq [simp]: "fold_coeffs f (pCons 0 0) = id" by (simp add: fold_coeffs_def) lemma fold_coeffs_pCons_coeff_not_0_eq [simp]: "a \ 0 \ fold_coeffs f (pCons a p) = f a \ fold_coeffs f p" by (simp add: fold_coeffs_def) lemma fold_coeffs_pCons_not_0_0_eq [simp]: "p \ 0 \ fold_coeffs f (pCons a p) = f a \ fold_coeffs f p" by (simp add: fold_coeffs_def) subsection \Canonical morphism on polynomials -- evaluation\ definition poly :: "'a::comm_semiring_0 poly \ 'a \ 'a" where "poly p = fold_coeffs (\a f x. a + x * f x) p (\x. 0)" \ \The Horner Schema\ lemma poly_0 [simp]: "poly 0 x = 0" by (simp add: poly_def) lemma poly_pCons [simp]: "poly (pCons a p) x = a + x * poly p x" by (cases "p = 0 \ a = 0") (auto simp add: poly_def) lemma poly_altdef: "poly p x = (\i\degree p. coeff p i * x ^ i)" for x :: "'a::{comm_semiring_0,semiring_1}" proof (induction p rule: pCons_induct) case 0 then show ?case by simp next case (pCons a p) show ?case proof (cases "p = 0") case True then show ?thesis by simp next case False let ?p' = "pCons a p" note poly_pCons[of a p x] also note pCons.IH also have "a + x * (\i\degree p. coeff p i * x ^ i) = coeff ?p' 0 * x^0 + (\i\degree p. coeff ?p' (Suc i) * x^Suc i)" by (simp add: field_simps sum_distrib_left coeff_pCons) also note sum.atMost_Suc_shift[symmetric] also note degree_pCons_eq[OF \p \ 0\, of a, symmetric] finally show ?thesis . qed qed lemma poly_0_coeff_0: "poly p 0 = coeff p 0" by (cases p) (auto simp: poly_altdef) subsection \Monomials\ lift_definition monom :: "'a \ nat \ 'a::zero poly" is "\a m n. if m = n then a else 0" by (simp add: MOST_iff_cofinite) lemma coeff_monom [simp]: "coeff (monom a m) n = (if m = n then a else 0)" by transfer rule lemma monom_0: "monom a 0 = pCons a 0" by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) lemma monom_Suc: "monom a (Suc n) = pCons 0 (monom a n)" by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) lemma monom_eq_0 [simp]: "monom 0 n = 0" by (rule poly_eqI) simp lemma monom_eq_0_iff [simp]: "monom a n = 0 \ a = 0" by (simp add: poly_eq_iff) lemma monom_eq_iff [simp]: "monom a n = monom b n \ a = b" by (simp add: poly_eq_iff) lemma degree_monom_le: "degree (monom a n) \ n" by (rule degree_le, simp) lemma degree_monom_eq: "a \ 0 \ degree (monom a n) = n" apply (rule order_antisym [OF degree_monom_le]) apply (rule le_degree) apply simp done lemma coeffs_monom [code abstract]: "coeffs (monom a n) = (if a = 0 then [] else replicate n 0 @ [a])" by (induct n) (simp_all add: monom_0 monom_Suc) lemma fold_coeffs_monom [simp]: "a \ 0 \ fold_coeffs f (monom a n) = f 0 ^^ n \ f a" by (simp add: fold_coeffs_def coeffs_monom fun_eq_iff) lemma poly_monom: "poly (monom a n) x = a * x ^ n" for a x :: "'a::comm_semiring_1" by (cases "a = 0", simp_all) (induct n, simp_all add: mult.left_commute poly_def) lemma monom_eq_iff': "monom c n = monom d m \ c = d \ (c = 0 \ n = m)" by (auto simp: poly_eq_iff) lemma monom_eq_const_iff: "monom c n = [:d:] \ c = d \ (c = 0 \ n = 0)" using monom_eq_iff'[of c n d 0] by (simp add: monom_0) subsection \Leading coefficient\ abbreviation lead_coeff:: "'a::zero poly \ 'a" where "lead_coeff p \ coeff p (degree p)" lemma lead_coeff_pCons[simp]: "p \ 0 \ lead_coeff (pCons a p) = lead_coeff p" "p = 0 \ lead_coeff (pCons a p) = a" by auto lemma lead_coeff_monom [simp]: "lead_coeff (monom c n) = c" by (cases "c = 0") (simp_all add: degree_monom_eq) lemma last_coeffs_eq_coeff_degree: "last (coeffs p) = lead_coeff p" if "p \ 0" using that by (simp add: coeffs_def) subsection \Addition and subtraction\ instantiation poly :: (comm_monoid_add) comm_monoid_add begin lift_definition plus_poly :: "'a poly \ 'a poly \ 'a poly" is "\p q n. coeff p n + coeff q n" proof - fix q p :: "'a poly" show "\\<^sub>\n. coeff p n + coeff q n = 0" using MOST_coeff_eq_0[of p] MOST_coeff_eq_0[of q] by eventually_elim simp qed lemma coeff_add [simp]: "coeff (p + q) n = coeff p n + coeff q n" by (simp add: plus_poly.rep_eq) instance proof fix p q r :: "'a poly" show "(p + q) + r = p + (q + r)" by (simp add: poly_eq_iff add.assoc) show "p + q = q + p" by (simp add: poly_eq_iff add.commute) show "0 + p = p" by (simp add: poly_eq_iff) qed end instantiation poly :: (cancel_comm_monoid_add) cancel_comm_monoid_add begin lift_definition minus_poly :: "'a poly \ 'a poly \ 'a poly" is "\p q n. coeff p n - coeff q n" proof - fix q p :: "'a poly" show "\\<^sub>\n. coeff p n - coeff q n = 0" using MOST_coeff_eq_0[of p] MOST_coeff_eq_0[of q] by eventually_elim simp qed lemma coeff_diff [simp]: "coeff (p - q) n = coeff p n - coeff q n" by (simp add: minus_poly.rep_eq) instance proof fix p q r :: "'a poly" show "p + q - p = q" by (simp add: poly_eq_iff) show "p - q - r = p - (q + r)" by (simp add: poly_eq_iff diff_diff_eq) qed end instantiation poly :: (ab_group_add) ab_group_add begin lift_definition uminus_poly :: "'a poly \ 'a poly" is "\p n. - coeff p n" proof - fix p :: "'a poly" show "\\<^sub>\n. - coeff p n = 0" using MOST_coeff_eq_0 by simp qed lemma coeff_minus [simp]: "coeff (- p) n = - coeff p n" by (simp add: uminus_poly.rep_eq) instance proof fix p q :: "'a poly" show "- p + p = 0" by (simp add: poly_eq_iff) show "p - q = p + - q" by (simp add: poly_eq_iff) qed end lemma add_pCons [simp]: "pCons a p + pCons b q = pCons (a + b) (p + q)" by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) lemma minus_pCons [simp]: "- pCons a p = pCons (- a) (- p)" by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) lemma diff_pCons [simp]: "pCons a p - pCons b q = pCons (a - b) (p - q)" by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) lemma degree_add_le_max: "degree (p + q) \ max (degree p) (degree q)" by (rule degree_le) (auto simp add: coeff_eq_0) lemma degree_add_le: "degree p \ n \ degree q \ n \ degree (p + q) \ n" by (auto intro: order_trans degree_add_le_max) lemma degree_add_less: "degree p < n \ degree q < n \ degree (p + q) < n" by (auto intro: le_less_trans degree_add_le_max) lemma degree_add_eq_right: "degree p < degree q \ degree (p + q) = degree q" apply (cases "q = 0") apply simp apply (rule order_antisym) apply (simp add: degree_add_le) apply (rule le_degree) apply (simp add: coeff_eq_0) done lemma degree_add_eq_left: "degree q < degree p \ degree (p + q) = degree p" using degree_add_eq_right [of q p] by (simp add: add.commute) lemma degree_minus [simp]: "degree (- p) = degree p" by (simp add: degree_def) lemma lead_coeff_add_le: "degree p < degree q \ lead_coeff (p + q) = lead_coeff q" by (metis coeff_add coeff_eq_0 monoid_add_class.add.left_neutral degree_add_eq_right) lemma lead_coeff_minus: "lead_coeff (- p) = - lead_coeff p" by (metis coeff_minus degree_minus) lemma degree_diff_le_max: "degree (p - q) \ max (degree p) (degree q)" for p q :: "'a::ab_group_add poly" using degree_add_le [where p=p and q="-q"] by simp lemma degree_diff_le: "degree p \ n \ degree q \ n \ degree (p - q) \ n" for p q :: "'a::ab_group_add poly" using degree_add_le [of p n "- q"] by simp lemma degree_diff_less: "degree p < n \ degree q < n \ degree (p - q) < n" for p q :: "'a::ab_group_add poly" using degree_add_less [of p n "- q"] by simp lemma add_monom: "monom a n + monom b n = monom (a + b) n" by (rule poly_eqI) simp lemma diff_monom: "monom a n - monom b n = monom (a - b) n" by (rule poly_eqI) simp lemma minus_monom: "- monom a n = monom (- a) n" by (rule poly_eqI) simp lemma coeff_sum: "coeff (\x\A. p x) i = (\x\A. coeff (p x) i)" by (induct A rule: infinite_finite_induct) simp_all lemma monom_sum: "monom (\x\A. a x) n = (\x\A. monom (a x) n)" by (rule poly_eqI) (simp add: coeff_sum) fun plus_coeffs :: "'a::comm_monoid_add list \ 'a list \ 'a list" where "plus_coeffs xs [] = xs" | "plus_coeffs [] ys = ys" | "plus_coeffs (x # xs) (y # ys) = (x + y) ## plus_coeffs xs ys" lemma coeffs_plus_eq_plus_coeffs [code abstract]: "coeffs (p + q) = plus_coeffs (coeffs p) (coeffs q)" proof - have *: "nth_default 0 (plus_coeffs xs ys) n = nth_default 0 xs n + nth_default 0 ys n" for xs ys :: "'a list" and n proof (induct xs ys arbitrary: n rule: plus_coeffs.induct) case (3 x xs y ys n) then show ?case by (cases n) (auto simp add: cCons_def) qed simp_all have **: "no_trailing (HOL.eq 0) (plus_coeffs xs ys)" if "no_trailing (HOL.eq 0) xs" and "no_trailing (HOL.eq 0) ys" for xs ys :: "'a list" using that by (induct xs ys rule: plus_coeffs.induct) (simp_all add: cCons_def) show ?thesis by (rule coeffs_eqI) (auto simp add: * nth_default_coeffs_eq intro: **) qed lemma coeffs_uminus [code abstract]: "coeffs (- p) = map uminus (coeffs p)" proof - have eq_0: "HOL.eq 0 \ uminus = HOL.eq (0::'a)" by (simp add: fun_eq_iff) show ?thesis by (rule coeffs_eqI) (simp_all add: nth_default_map_eq nth_default_coeffs_eq no_trailing_map eq_0) qed lemma [code]: "p - q = p + - q" for p q :: "'a::ab_group_add poly" by (fact diff_conv_add_uminus) lemma poly_add [simp]: "poly (p + q) x = poly p x + poly q x" apply (induct p arbitrary: q) apply simp apply (case_tac q, simp, simp add: algebra_simps) done lemma poly_minus [simp]: "poly (- p) x = - poly p x" for x :: "'a::comm_ring" by (induct p) simp_all lemma poly_diff [simp]: "poly (p - q) x = poly p x - poly q x" for x :: "'a::comm_ring" using poly_add [of p "- q" x] by simp lemma poly_sum: "poly (\k\A. p k) x = (\k\A. poly (p k) x)" by (induct A rule: infinite_finite_induct) simp_all lemma degree_sum_le: "finite S \ (\p. p \ S \ degree (f p) \ n) \ degree (sum f S) \ n" proof (induct S rule: finite_induct) case empty then show ?case by simp next case (insert p S) then have "degree (sum f S) \ n" "degree (f p) \ n" by auto then show ?case unfolding sum.insert[OF insert(1-2)] by (metis degree_add_le) qed lemma poly_as_sum_of_monoms': assumes "degree p \ n" shows "(\i\n. monom (coeff p i) i) = p" proof - have eq: "\i. {..n} \ {i} = (if i \ n then {i} else {})" by auto from assms show ?thesis by (simp add: poly_eq_iff coeff_sum coeff_eq_0 sum.If_cases eq if_distrib[where f="\x. x * a" for a]) qed lemma poly_as_sum_of_monoms: "(\i\degree p. monom (coeff p i) i) = p" by (intro poly_as_sum_of_monoms' order_refl) lemma Poly_snoc: "Poly (xs @ [x]) = Poly xs + monom x (length xs)" by (induct xs) (simp_all add: monom_0 monom_Suc) subsection \Multiplication by a constant, polynomial multiplication and the unit polynomial\ lift_definition smult :: "'a::comm_semiring_0 \ 'a poly \ 'a poly" is "\a p n. a * coeff p n" proof - fix a :: 'a and p :: "'a poly" show "\\<^sub>\ i. a * coeff p i = 0" using MOST_coeff_eq_0[of p] by eventually_elim simp qed lemma coeff_smult [simp]: "coeff (smult a p) n = a * coeff p n" by (simp add: smult.rep_eq) lemma degree_smult_le: "degree (smult a p) \ degree p" by (rule degree_le) (simp add: coeff_eq_0) lemma smult_smult [simp]: "smult a (smult b p) = smult (a * b) p" by (rule poly_eqI) (simp add: mult.assoc) lemma smult_0_right [simp]: "smult a 0 = 0" by (rule poly_eqI) simp lemma smult_0_left [simp]: "smult 0 p = 0" by (rule poly_eqI) simp lemma smult_1_left [simp]: "smult (1::'a::comm_semiring_1) p = p" by (rule poly_eqI) simp lemma smult_add_right: "smult a (p + q) = smult a p + smult a q" by (rule poly_eqI) (simp add: algebra_simps) lemma smult_add_left: "smult (a + b) p = smult a p + smult b p" by (rule poly_eqI) (simp add: algebra_simps) lemma smult_minus_right [simp]: "smult a (- p) = - smult a p" for a :: "'a::comm_ring" by (rule poly_eqI) simp lemma smult_minus_left [simp]: "smult (- a) p = - smult a p" for a :: "'a::comm_ring" by (rule poly_eqI) simp lemma smult_diff_right: "smult a (p - q) = smult a p - smult a q" for a :: "'a::comm_ring" by (rule poly_eqI) (simp add: algebra_simps) lemma smult_diff_left: "smult (a - b) p = smult a p - smult b p" for a b :: "'a::comm_ring" by (rule poly_eqI) (simp add: algebra_simps) lemmas smult_distribs = smult_add_left smult_add_right smult_diff_left smult_diff_right lemma smult_pCons [simp]: "smult a (pCons b p) = pCons (a * b) (smult a p)" by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) lemma smult_monom: "smult a (monom b n) = monom (a * b) n" by (induct n) (simp_all add: monom_0 monom_Suc) lemma smult_Poly: "smult c (Poly xs) = Poly (map ((*) c) xs)" by (auto simp: poly_eq_iff nth_default_def) lemma degree_smult_eq [simp]: "degree (smult a p) = (if a = 0 then 0 else degree p)" for a :: "'a::{comm_semiring_0,semiring_no_zero_divisors}" by (cases "a = 0") (simp_all add: degree_def) lemma smult_eq_0_iff [simp]: "smult a p = 0 \ a = 0 \ p = 0" for a :: "'a::{comm_semiring_0,semiring_no_zero_divisors}" by (simp add: poly_eq_iff) lemma coeffs_smult [code abstract]: "coeffs (smult a p) = (if a = 0 then [] else map (Groups.times a) (coeffs p))" for p :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" proof - have eq_0: "HOL.eq 0 \ times a = HOL.eq (0::'a)" if "a \ 0" using that by (simp add: fun_eq_iff) show ?thesis by (rule coeffs_eqI) (auto simp add: no_trailing_map nth_default_map_eq nth_default_coeffs_eq eq_0) qed lemma smult_eq_iff: fixes b :: "'a :: field" assumes "b \ 0" shows "smult a p = smult b q \ smult (a / b) p = q" (is "?lhs \ ?rhs") proof assume ?lhs also from assms have "smult (inverse b) \ = q" by simp finally show ?rhs by (simp add: field_simps) next assume ?rhs with assms show ?lhs by auto qed instantiation poly :: (comm_semiring_0) comm_semiring_0 begin definition "p * q = fold_coeffs (\a p. smult a q + pCons 0 p) p 0" lemma mult_poly_0_left: "(0::'a poly) * q = 0" by (simp add: times_poly_def) lemma mult_pCons_left [simp]: "pCons a p * q = smult a q + pCons 0 (p * q)" by (cases "p = 0 \ a = 0") (auto simp add: times_poly_def) lemma mult_poly_0_right: "p * (0::'a poly) = 0" by (induct p) (simp_all add: mult_poly_0_left) lemma mult_pCons_right [simp]: "p * pCons a q = smult a p + pCons 0 (p * q)" by (induct p) (simp_all add: mult_poly_0_left algebra_simps) lemmas mult_poly_0 = mult_poly_0_left mult_poly_0_right lemma mult_smult_left [simp]: "smult a p * q = smult a (p * q)" by (induct p) (simp_all add: mult_poly_0 smult_add_right) lemma mult_smult_right [simp]: "p * smult a q = smult a (p * q)" by (induct q) (simp_all add: mult_poly_0 smult_add_right) lemma mult_poly_add_left: "(p + q) * r = p * r + q * r" for p q r :: "'a poly" by (induct r) (simp_all add: mult_poly_0 smult_distribs algebra_simps) instance proof fix p q r :: "'a poly" show 0: "0 * p = 0" by (rule mult_poly_0_left) show "p * 0 = 0" by (rule mult_poly_0_right) show "(p + q) * r = p * r + q * r" by (rule mult_poly_add_left) show "(p * q) * r = p * (q * r)" by (induct p) (simp_all add: mult_poly_0 mult_poly_add_left) show "p * q = q * p" by (induct p) (simp_all add: mult_poly_0) qed end lemma coeff_mult_degree_sum: "coeff (p * q) (degree p + degree q) = coeff p (degree p) * coeff q (degree q)" by (induct p) (simp_all add: coeff_eq_0) instance poly :: ("{comm_semiring_0,semiring_no_zero_divisors}") semiring_no_zero_divisors proof fix p q :: "'a poly" assume "p \ 0" and "q \ 0" have "coeff (p * q) (degree p + degree q) = coeff p (degree p) * coeff q (degree q)" by (rule coeff_mult_degree_sum) also from \p \ 0\ \q \ 0\ have "coeff p (degree p) * coeff q (degree q) \ 0" by simp finally have "\n. coeff (p * q) n \ 0" .. then show "p * q \ 0" by (simp add: poly_eq_iff) qed instance poly :: (comm_semiring_0_cancel) comm_semiring_0_cancel .. lemma coeff_mult: "coeff (p * q) n = (\i\n. coeff p i * coeff q (n-i))" proof (induct p arbitrary: n) case 0 show ?case by simp next case (pCons a p n) then show ?case by (cases n) (simp_all add: sum.atMost_Suc_shift del: sum.atMost_Suc) qed lemma degree_mult_le: "degree (p * q) \ degree p + degree q" apply (rule degree_le) apply (induct p) apply simp apply (simp add: coeff_eq_0 coeff_pCons split: nat.split) done lemma mult_monom: "monom a m * monom b n = monom (a * b) (m + n)" by (induct m) (simp add: monom_0 smult_monom, simp add: monom_Suc) instantiation poly :: (comm_semiring_1) comm_semiring_1 begin lift_definition one_poly :: "'a poly" is "\n. of_bool (n = 0)" by (rule MOST_SucD) simp lemma coeff_1 [simp]: "coeff 1 n = of_bool (n = 0)" by (simp add: one_poly.rep_eq) lemma one_pCons: "1 = [:1:]" by (simp add: poly_eq_iff coeff_pCons split: nat.splits) lemma pCons_one: "[:1:] = 1" by (simp add: one_pCons) instance by standard (simp_all add: one_pCons) end lemma poly_1 [simp]: "poly 1 x = 1" by (simp add: one_pCons) lemma one_poly_eq_simps [simp]: "1 = [:1:] \ True" "[:1:] = 1 \ True" by (simp_all add: one_pCons) lemma degree_1 [simp]: "degree 1 = 0" by (simp add: one_pCons) lemma coeffs_1_eq [simp, code abstract]: "coeffs 1 = [1]" by (simp add: one_pCons) lemma smult_one [simp]: "smult c 1 = [:c:]" by (simp add: one_pCons) lemma monom_eq_1 [simp]: "monom 1 0 = 1" by (simp add: monom_0 one_pCons) lemma monom_eq_1_iff: "monom c n = 1 \ c = 1 \ n = 0" using monom_eq_const_iff [of c n 1] by auto lemma monom_altdef: "monom c n = smult c ([:0, 1:] ^ n)" by (induct n) (simp_all add: monom_0 monom_Suc) instance poly :: ("{comm_semiring_1,semiring_1_no_zero_divisors}") semiring_1_no_zero_divisors .. instance poly :: (comm_ring) comm_ring .. instance poly :: (comm_ring_1) comm_ring_1 .. instance poly :: (comm_ring_1) comm_semiring_1_cancel .. lemma degree_power_le: "degree (p ^ n) \ degree p * n" by (induct n) (auto intro: order_trans degree_mult_le) lemma coeff_0_power: "coeff (p ^ n) 0 = coeff p 0 ^ n" by (induct n) (simp_all add: coeff_mult) lemma poly_smult [simp]: "poly (smult a p) x = a * poly p x" by (induct p) (simp_all add: algebra_simps) lemma poly_mult [simp]: "poly (p * q) x = poly p x * poly q x" by (induct p) (simp_all add: algebra_simps) lemma poly_power [simp]: "poly (p ^ n) x = poly p x ^ n" for p :: "'a::comm_semiring_1 poly" by (induct n) simp_all lemma poly_prod: "poly (\k\A. p k) x = (\k\A. poly (p k) x)" by (induct A rule: infinite_finite_induct) simp_all lemma degree_prod_sum_le: "finite S \ degree (prod f S) \ sum (degree \ f) S" proof (induct S rule: finite_induct) case empty then show ?case by simp next case (insert a S) show ?case unfolding prod.insert[OF insert(1-2)] sum.insert[OF insert(1-2)] by (rule le_trans[OF degree_mult_le]) (use insert in auto) qed lemma coeff_0_prod_list: "coeff (prod_list xs) 0 = prod_list (map (\p. coeff p 0) xs)" by (induct xs) (simp_all add: coeff_mult) lemma coeff_monom_mult: "coeff (monom c n * p) k = (if k < n then 0 else c * coeff p (k - n))" proof - have "coeff (monom c n * p) k = (\i\k. (if n = i then c else 0) * coeff p (k - i))" by (simp add: coeff_mult) also have "\ = (\i\k. (if n = i then c * coeff p (k - i) else 0))" by (intro sum.cong) simp_all also have "\ = (if k < n then 0 else c * coeff p (k - n))" by simp finally show ?thesis . qed lemma monom_1_dvd_iff': "monom 1 n dvd p \ (\kkkk. coeff p (k + n))" have "\\<^sub>\k. coeff p (k + n) = 0" by (subst cofinite_eq_sequentially, subst eventually_sequentially_seg, subst cofinite_eq_sequentially [symmetric]) transfer then have coeff_r [simp]: "coeff r k = coeff p (k + n)" for k unfolding r_def by (subst poly.Abs_poly_inverse) simp_all have "p = monom 1 n * r" by (rule poly_eqI, subst coeff_monom_mult) (simp_all add: zero) then show "monom 1 n dvd p" by simp qed subsection \Mapping polynomials\ definition map_poly :: "('a :: zero \ 'b :: zero) \ 'a poly \ 'b poly" where "map_poly f p = Poly (map f (coeffs p))" lemma map_poly_0 [simp]: "map_poly f 0 = 0" by (simp add: map_poly_def) lemma map_poly_1: "map_poly f 1 = [:f 1:]" by (simp add: map_poly_def) lemma map_poly_1' [simp]: "f 1 = 1 \ map_poly f 1 = 1" by (simp add: map_poly_def one_pCons) lemma coeff_map_poly: assumes "f 0 = 0" shows "coeff (map_poly f p) n = f (coeff p n)" by (auto simp: assms map_poly_def nth_default_def coeffs_def not_less Suc_le_eq coeff_eq_0 simp del: upt_Suc) lemma coeffs_map_poly [code abstract]: "coeffs (map_poly f p) = strip_while ((=) 0) (map f (coeffs p))" by (simp add: map_poly_def) lemma coeffs_map_poly': assumes "\x. x \ 0 \ f x \ 0" shows "coeffs (map_poly f p) = map f (coeffs p)" using assms by (auto simp add: coeffs_map_poly strip_while_idem_iff last_coeffs_eq_coeff_degree no_trailing_unfold last_map) lemma set_coeffs_map_poly: "(\x. f x = 0 \ x = 0) \ set (coeffs (map_poly f p)) = f ` set (coeffs p)" by (simp add: coeffs_map_poly') lemma degree_map_poly: assumes "\x. x \ 0 \ f x \ 0" shows "degree (map_poly f p) = degree p" by (simp add: degree_eq_length_coeffs coeffs_map_poly' assms) lemma map_poly_eq_0_iff: assumes "f 0 = 0" "\x. x \ set (coeffs p) \ x \ 0 \ f x \ 0" shows "map_poly f p = 0 \ p = 0" proof - have "(coeff (map_poly f p) n = 0) = (coeff p n = 0)" for n proof - have "coeff (map_poly f p) n = f (coeff p n)" by (simp add: coeff_map_poly assms) also have "\ = 0 \ coeff p n = 0" proof (cases "n < length (coeffs p)") case True then have "coeff p n \ set (coeffs p)" by (auto simp: coeffs_def simp del: upt_Suc) with assms show "f (coeff p n) = 0 \ coeff p n = 0" by auto next case False then show ?thesis by (auto simp: assms length_coeffs nth_default_coeffs_eq [symmetric] nth_default_def) qed finally show ?thesis . qed then show ?thesis by (auto simp: poly_eq_iff) qed lemma map_poly_smult: assumes "f 0 = 0""\c x. f (c * x) = f c * f x" shows "map_poly f (smult c p) = smult (f c) (map_poly f p)" by (intro poly_eqI) (simp_all add: assms coeff_map_poly) lemma map_poly_pCons: assumes "f 0 = 0" shows "map_poly f (pCons c p) = pCons (f c) (map_poly f p)" by (intro poly_eqI) (simp_all add: assms coeff_map_poly coeff_pCons split: nat.splits) lemma map_poly_map_poly: assumes "f 0 = 0" "g 0 = 0" shows "map_poly f (map_poly g p) = map_poly (f \ g) p" by (intro poly_eqI) (simp add: coeff_map_poly assms) lemma map_poly_id [simp]: "map_poly id p = p" by (simp add: map_poly_def) lemma map_poly_id' [simp]: "map_poly (\x. x) p = p" by (simp add: map_poly_def) lemma map_poly_cong: assumes "(\x. x \ set (coeffs p) \ f x = g x)" shows "map_poly f p = map_poly g p" proof - from assms have "map f (coeffs p) = map g (coeffs p)" by (intro map_cong) simp_all then show ?thesis by (simp only: coeffs_eq_iff coeffs_map_poly) qed lemma map_poly_monom: "f 0 = 0 \ map_poly f (monom c n) = monom (f c) n" by (intro poly_eqI) (simp_all add: coeff_map_poly) lemma map_poly_idI: assumes "\x. x \ set (coeffs p) \ f x = x" shows "map_poly f p = p" using map_poly_cong[OF assms, of _ id] by simp lemma map_poly_idI': assumes "\x. x \ set (coeffs p) \ f x = x" shows "p = map_poly f p" using map_poly_cong[OF assms, of _ id] by simp lemma smult_conv_map_poly: "smult c p = map_poly (\x. c * x) p" by (intro poly_eqI) (simp_all add: coeff_map_poly) subsection \Conversions\ lemma of_nat_poly: "of_nat n = [:of_nat n:]" by (induct n) (simp_all add: one_pCons) lemma of_nat_monom: "of_nat n = monom (of_nat n) 0" by (simp add: of_nat_poly monom_0) lemma degree_of_nat [simp]: "degree (of_nat n) = 0" by (simp add: of_nat_poly) lemma lead_coeff_of_nat [simp]: "lead_coeff (of_nat n) = of_nat n" by (simp add: of_nat_poly) lemma of_int_poly: "of_int k = [:of_int k:]" by (simp only: of_int_of_nat of_nat_poly) simp lemma of_int_monom: "of_int k = monom (of_int k) 0" by (simp add: of_int_poly monom_0) lemma degree_of_int [simp]: "degree (of_int k) = 0" by (simp add: of_int_poly) lemma lead_coeff_of_int [simp]: "lead_coeff (of_int k) = of_int k" by (simp add: of_int_poly) lemma numeral_poly: "numeral n = [:numeral n:]" proof - have "numeral n = of_nat (numeral n)" by simp also have "\ = [:of_nat (numeral n):]" by (simp add: of_nat_poly) finally show ?thesis by simp qed lemma numeral_monom: "numeral n = monom (numeral n) 0" by (simp add: numeral_poly monom_0) lemma degree_numeral [simp]: "degree (numeral n) = 0" by (simp add: numeral_poly) lemma lead_coeff_numeral [simp]: "lead_coeff (numeral n) = numeral n" by (simp add: numeral_poly) subsection \Lemmas about divisibility\ lemma dvd_smult: assumes "p dvd q" shows "p dvd smult a q" proof - from assms obtain k where "q = p * k" .. then have "smult a q = p * smult a k" by simp then show "p dvd smult a q" .. qed lemma dvd_smult_cancel: "p dvd smult a q \ a \ 0 \ p dvd q" for a :: "'a::field" by (drule dvd_smult [where a="inverse a"]) simp lemma dvd_smult_iff: "a \ 0 \ p dvd smult a q \ p dvd q" for a :: "'a::field" by (safe elim!: dvd_smult dvd_smult_cancel) lemma smult_dvd_cancel: assumes "smult a p dvd q" shows "p dvd q" proof - from assms obtain k where "q = smult a p * k" .. then have "q = p * smult a k" by simp then show "p dvd q" .. qed lemma smult_dvd: "p dvd q \ a \ 0 \ smult a p dvd q" for a :: "'a::field" by (rule smult_dvd_cancel [where a="inverse a"]) simp lemma smult_dvd_iff: "smult a p dvd q \ (if a = 0 then q = 0 else p dvd q)" for a :: "'a::field" by (auto elim: smult_dvd smult_dvd_cancel) lemma is_unit_smult_iff: "smult c p dvd 1 \ c dvd 1 \ p dvd 1" proof - have "smult c p = [:c:] * p" by simp also have "\ dvd 1 \ c dvd 1 \ p dvd 1" proof safe assume *: "[:c:] * p dvd 1" then show "p dvd 1" by (rule dvd_mult_right) from * obtain q where q: "1 = [:c:] * p * q" by (rule dvdE) have "c dvd c * (coeff p 0 * coeff q 0)" by simp also have "\ = coeff ([:c:] * p * q) 0" by (simp add: mult.assoc coeff_mult) also note q [symmetric] finally have "c dvd coeff 1 0" . then show "c dvd 1" by simp next assume "c dvd 1" "p dvd 1" from this(1) obtain d where "1 = c * d" by (rule dvdE) then have "1 = [:c:] * [:d:]" by (simp add: one_pCons ac_simps) then have "[:c:] dvd 1" by (rule dvdI) from mult_dvd_mono[OF this \p dvd 1\] show "[:c:] * p dvd 1" by simp qed finally show ?thesis . qed subsection \Polynomials form an integral domain\ instance poly :: (idom) idom .. instance poly :: ("{ring_char_0, comm_ring_1}") ring_char_0 by standard (auto simp add: of_nat_poly intro: injI) lemma degree_mult_eq: "p \ 0 \ q \ 0 \ degree (p * q) = degree p + degree q" for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" by (rule order_antisym [OF degree_mult_le le_degree]) (simp add: coeff_mult_degree_sum) lemma degree_mult_eq_0: "degree (p * q) = 0 \ p = 0 \ q = 0 \ (p \ 0 \ q \ 0 \ degree p = 0 \ degree q = 0)" for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" by (auto simp: degree_mult_eq) lemma degree_power_eq: "p \ 0 \ degree ((p :: 'a :: idom poly) ^ n) = n * degree p" by (induction n) (simp_all add: degree_mult_eq) lemma degree_mult_right_le: fixes p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" assumes "q \ 0" shows "degree p \ degree (p * q)" using assms by (cases "p = 0") (simp_all add: degree_mult_eq) lemma coeff_degree_mult: "coeff (p * q) (degree (p * q)) = coeff q (degree q) * coeff p (degree p)" for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" by (cases "p = 0 \ q = 0") (auto simp: degree_mult_eq coeff_mult_degree_sum mult_ac) lemma dvd_imp_degree_le: "p dvd q \ q \ 0 \ degree p \ degree q" for p q :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" by (erule dvdE, hypsubst, subst degree_mult_eq) auto lemma divides_degree: fixes p q :: "'a ::{comm_semiring_1,semiring_no_zero_divisors} poly" assumes "p dvd q" shows "degree p \ degree q \ q = 0" by (metis dvd_imp_degree_le assms) lemma const_poly_dvd_iff: fixes c :: "'a::{comm_semiring_1,semiring_no_zero_divisors}" shows "[:c:] dvd p \ (\n. c dvd coeff p n)" proof (cases "c = 0 \ p = 0") case True then show ?thesis by (auto intro!: poly_eqI) next case False show ?thesis proof assume "[:c:] dvd p" then show "\n. c dvd coeff p n" by (auto elim!: dvdE simp: coeffs_def) next assume *: "\n. c dvd coeff p n" define mydiv where "mydiv x y = (SOME z. x = y * z)" for x y :: 'a have mydiv: "x = y * mydiv x y" if "y dvd x" for x y using that unfolding mydiv_def dvd_def by (rule someI_ex) define q where "q = Poly (map (\a. mydiv a c) (coeffs p))" from False * have "p = q * [:c:]" by (intro poly_eqI) (auto simp: q_def nth_default_def not_less length_coeffs_degree coeffs_nth intro!: coeff_eq_0 mydiv) then show "[:c:] dvd p" by (simp only: dvd_triv_right) qed qed lemma const_poly_dvd_const_poly_iff [simp]: "[:a:] dvd [:b:] \ a dvd b" for a b :: "'a::{comm_semiring_1,semiring_no_zero_divisors}" by (subst const_poly_dvd_iff) (auto simp: coeff_pCons split: nat.splits) lemma lead_coeff_mult: "lead_coeff (p * q) = lead_coeff p * lead_coeff q" for p q :: "'a::{comm_semiring_0, semiring_no_zero_divisors} poly" by (cases "p = 0 \ q = 0") (auto simp: coeff_mult_degree_sum degree_mult_eq) lemma lead_coeff_smult: "lead_coeff (smult c p) = c * lead_coeff p" for p :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" proof - have "smult c p = [:c:] * p" by simp also have "lead_coeff \ = c * lead_coeff p" by (subst lead_coeff_mult) simp_all finally show ?thesis . qed lemma lead_coeff_1 [simp]: "lead_coeff 1 = 1" by simp lemma lead_coeff_power: "lead_coeff (p ^ n) = lead_coeff p ^ n" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" by (induct n) (simp_all add: lead_coeff_mult) subsection \Polynomials form an ordered integral domain\ definition pos_poly :: "'a::linordered_semidom poly \ bool" where "pos_poly p \ 0 < coeff p (degree p)" lemma pos_poly_pCons: "pos_poly (pCons a p) \ pos_poly p \ (p = 0 \ 0 < a)" by (simp add: pos_poly_def) lemma not_pos_poly_0 [simp]: "\ pos_poly 0" by (simp add: pos_poly_def) lemma pos_poly_add: "pos_poly p \ pos_poly q \ pos_poly (p + q)" apply (induct p arbitrary: q) apply simp apply (case_tac q) apply (force simp add: pos_poly_pCons add_pos_pos) done lemma pos_poly_mult: "pos_poly p \ pos_poly q \ pos_poly (p * q)" unfolding pos_poly_def apply (subgoal_tac "p \ 0 \ q \ 0") apply (simp add: degree_mult_eq coeff_mult_degree_sum) apply auto done lemma pos_poly_total: "p = 0 \ pos_poly p \ pos_poly (- p)" for p :: "'a::linordered_idom poly" by (induct p) (auto simp: pos_poly_pCons) lemma pos_poly_coeffs [code]: "pos_poly p \ (let as = coeffs p in as \ [] \ last as > 0)" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (auto simp add: pos_poly_def last_coeffs_eq_coeff_degree) next assume ?lhs then have *: "0 < coeff p (degree p)" by (simp add: pos_poly_def) then have "p \ 0" by auto with * show ?rhs by (simp add: last_coeffs_eq_coeff_degree) qed instantiation poly :: (linordered_idom) linordered_idom begin definition "x < y \ pos_poly (y - x)" definition "x \ y \ x = y \ pos_poly (y - x)" definition "\x::'a poly\ = (if x < 0 then - x else x)" definition "sgn (x::'a poly) = (if x = 0 then 0 else if 0 < x then 1 else - 1)" instance proof fix x y z :: "'a poly" show "x < y \ x \ y \ \ y \ x" unfolding less_eq_poly_def less_poly_def apply safe apply simp apply (drule (1) pos_poly_add) apply simp done show "x \ x" by (simp add: less_eq_poly_def) show "x \ y \ y \ z \ x \ z" unfolding less_eq_poly_def apply safe apply (drule (1) pos_poly_add) apply (simp add: algebra_simps) done show "x \ y \ y \ x \ x = y" unfolding less_eq_poly_def apply safe apply (drule (1) pos_poly_add) apply simp done show "x \ y \ z + x \ z + y" unfolding less_eq_poly_def apply safe apply (simp add: algebra_simps) done show "x \ y \ y \ x" unfolding less_eq_poly_def using pos_poly_total [of "x - y"] by auto show "x < y \ 0 < z \ z * x < z * y" by (simp add: less_poly_def right_diff_distrib [symmetric] pos_poly_mult) show "\x\ = (if x < 0 then - x else x)" by (rule abs_poly_def) show "sgn x = (if x = 0 then 0 else if 0 < x then 1 else - 1)" by (rule sgn_poly_def) qed end text \TODO: Simplification rules for comparisons\ subsection \Synthetic division and polynomial roots\ subsubsection \Synthetic division\ text \Synthetic division is simply division by the linear polynomial \<^term>\x - c\.\ definition synthetic_divmod :: "'a::comm_semiring_0 poly \ 'a \ 'a poly \ 'a" where "synthetic_divmod p c = fold_coeffs (\a (q, r). (pCons r q, a + c * r)) p (0, 0)" definition synthetic_div :: "'a::comm_semiring_0 poly \ 'a \ 'a poly" where "synthetic_div p c = fst (synthetic_divmod p c)" lemma synthetic_divmod_0 [simp]: "synthetic_divmod 0 c = (0, 0)" by (simp add: synthetic_divmod_def) lemma synthetic_divmod_pCons [simp]: "synthetic_divmod (pCons a p) c = (\(q, r). (pCons r q, a + c * r)) (synthetic_divmod p c)" by (cases "p = 0 \ a = 0") (auto simp add: synthetic_divmod_def) lemma synthetic_div_0 [simp]: "synthetic_div 0 c = 0" by (simp add: synthetic_div_def) lemma synthetic_div_unique_lemma: "smult c p = pCons a p \ p = 0" by (induct p arbitrary: a) simp_all lemma snd_synthetic_divmod: "snd (synthetic_divmod p c) = poly p c" by (induct p) (simp_all add: split_def) lemma synthetic_div_pCons [simp]: "synthetic_div (pCons a p) c = pCons (poly p c) (synthetic_div p c)" by (simp add: synthetic_div_def split_def snd_synthetic_divmod) lemma synthetic_div_eq_0_iff: "synthetic_div p c = 0 \ degree p = 0" proof (induct p) case 0 then show ?case by simp next case (pCons a p) then show ?case by (cases p) simp qed lemma degree_synthetic_div: "degree (synthetic_div p c) = degree p - 1" by (induct p) (simp_all add: synthetic_div_eq_0_iff) lemma synthetic_div_correct: "p + smult c (synthetic_div p c) = pCons (poly p c) (synthetic_div p c)" by (induct p) simp_all lemma synthetic_div_unique: "p + smult c q = pCons r q \ r = poly p c \ q = synthetic_div p c" apply (induct p arbitrary: q r) apply simp apply (frule synthetic_div_unique_lemma) apply simp apply (case_tac q, force) done lemma synthetic_div_correct': "[:-c, 1:] * synthetic_div p c + [:poly p c:] = p" for c :: "'a::comm_ring_1" using synthetic_div_correct [of p c] by (simp add: algebra_simps) subsubsection \Polynomial roots\ lemma poly_eq_0_iff_dvd: "poly p c = 0 \ [:- c, 1:] dvd p" (is "?lhs \ ?rhs") for c :: "'a::comm_ring_1" proof assume ?lhs with synthetic_div_correct' [of c p] have "p = [:-c, 1:] * synthetic_div p c" by simp then show ?rhs .. next assume ?rhs then obtain k where "p = [:-c, 1:] * k" by (rule dvdE) then show ?lhs by simp qed lemma dvd_iff_poly_eq_0: "[:c, 1:] dvd p \ poly p (- c) = 0" for c :: "'a::comm_ring_1" by (simp add: poly_eq_0_iff_dvd) lemma poly_roots_finite: "p \ 0 \ finite {x. poly p x = 0}" for p :: "'a::{comm_ring_1,ring_no_zero_divisors} poly" proof (induct n \ "degree p" arbitrary: p) case 0 then obtain a where "a \ 0" and "p = [:a:]" by (cases p) (simp split: if_splits) then show "finite {x. poly p x = 0}" by simp next case (Suc n) show "finite {x. poly p x = 0}" proof (cases "\x. poly p x = 0") case False then show "finite {x. poly p x = 0}" by simp next case True then obtain a where "poly p a = 0" .. then have "[:-a, 1:] dvd p" by (simp only: poly_eq_0_iff_dvd) then obtain k where k: "p = [:-a, 1:] * k" .. with \p \ 0\ have "k \ 0" by auto with k have "degree p = Suc (degree k)" by (simp add: degree_mult_eq del: mult_pCons_left) with \Suc n = degree p\ have "n = degree k" by simp from this \k \ 0\ have "finite {x. poly k x = 0}" by (rule Suc.hyps) then have "finite (insert a {x. poly k x = 0})" by simp then show "finite {x. poly p x = 0}" by (simp add: k Collect_disj_eq del: mult_pCons_left) qed qed lemma poly_eq_poly_eq_iff: "poly p = poly q \ p = q" (is "?lhs \ ?rhs") for p q :: "'a::{comm_ring_1,ring_no_zero_divisors,ring_char_0} poly" proof assume ?rhs then show ?lhs by simp next assume ?lhs have "poly p = poly 0 \ p = 0" for p :: "'a poly" apply (cases "p = 0") apply simp_all apply (drule poly_roots_finite) apply (auto simp add: infinite_UNIV_char_0) done from \?lhs\ and this [of "p - q"] show ?rhs by auto qed lemma poly_all_0_iff_0: "(\x. poly p x = 0) \ p = 0" for p :: "'a::{ring_char_0,comm_ring_1,ring_no_zero_divisors} poly" by (auto simp add: poly_eq_poly_eq_iff [symmetric]) subsubsection \Order of polynomial roots\ definition order :: "'a::idom \ 'a poly \ nat" where "order a p = (LEAST n. \ [:-a, 1:] ^ Suc n dvd p)" lemma coeff_linear_power: "coeff ([:a, 1:] ^ n) n = 1" for a :: "'a::comm_semiring_1" apply (induct n) apply simp_all apply (subst coeff_eq_0) apply (auto intro: le_less_trans degree_power_le) done lemma degree_linear_power: "degree ([:a, 1:] ^ n) = n" for a :: "'a::comm_semiring_1" apply (rule order_antisym) apply (rule ord_le_eq_trans [OF degree_power_le]) apply simp apply (rule le_degree) apply (simp add: coeff_linear_power) done lemma order_1: "[:-a, 1:] ^ order a p dvd p" apply (cases "p = 0") apply simp apply (cases "order a p") apply simp apply (subgoal_tac "nat < (LEAST n. \ [:-a, 1:] ^ Suc n dvd p)") apply (drule not_less_Least) apply simp apply (fold order_def) apply simp done lemma order_2: "p \ 0 \ \ [:-a, 1:] ^ Suc (order a p) dvd p" unfolding order_def apply (rule LeastI_ex) apply (rule_tac x="degree p" in exI) apply (rule notI) apply (drule (1) dvd_imp_degree_le) apply (simp only: degree_linear_power) done lemma order: "p \ 0 \ [:-a, 1:] ^ order a p dvd p \ \ [:-a, 1:] ^ Suc (order a p) dvd p" by (rule conjI [OF order_1 order_2]) lemma order_degree: assumes p: "p \ 0" shows "order a p \ degree p" proof - have "order a p = degree ([:-a, 1:] ^ order a p)" by (simp only: degree_linear_power) also from order_1 p have "\ \ degree p" by (rule dvd_imp_degree_le) finally show ?thesis . qed lemma order_root: "poly p a = 0 \ p = 0 \ order a p \ 0" apply (cases "p = 0") apply simp_all apply (rule iffI) apply (metis order_2 not_gr0 poly_eq_0_iff_dvd power_0 power_Suc_0 power_one_right) unfolding poly_eq_0_iff_dvd apply (metis dvd_power dvd_trans order_1) done lemma order_0I: "poly p a \ 0 \ order a p = 0" by (subst (asm) order_root) auto lemma order_unique_lemma: fixes p :: "'a::idom poly" assumes "[:-a, 1:] ^ n dvd p" "\ [:-a, 1:] ^ Suc n dvd p" shows "n = order a p" unfolding Polynomial.order_def apply (rule Least_equality [symmetric]) apply (fact assms) apply (rule classical) apply (erule notE) unfolding not_less_eq_eq using assms(1) apply (rule power_le_dvd) apply assumption done lemma order_mult: "p * q \ 0 \ order a (p * q) = order a p + order a q" proof - define i where "i = order a p" define j where "j = order a q" define t where "t = [:-a, 1:]" have t_dvd_iff: "\u. t dvd u \ poly u a = 0" by (simp add: t_def dvd_iff_poly_eq_0) assume "p * q \ 0" then show "order a (p * q) = i + j" apply clarsimp apply (drule order [where a=a and p=p, folded i_def t_def]) apply (drule order [where a=a and p=q, folded j_def t_def]) apply clarify apply (erule dvdE)+ apply (rule order_unique_lemma [symmetric], fold t_def) apply (simp_all add: power_add t_dvd_iff) done qed lemma order_smult: assumes "c \ 0" shows "order x (smult c p) = order x p" proof (cases "p = 0") case True then show ?thesis by simp next case False have "smult c p = [:c:] * p" by simp also from assms False have "order x \ = order x [:c:] + order x p" by (subst order_mult) simp_all also have "order x [:c:] = 0" by (rule order_0I) (use assms in auto) finally show ?thesis by simp qed (* Next three lemmas contributed by Wenda Li *) lemma order_1_eq_0 [simp]:"order x 1 = 0" by (metis order_root poly_1 zero_neq_one) lemma order_uminus[simp]: "order x (-p) = order x p" by (metis neg_equal_0_iff_equal order_smult smult_1_left smult_minus_left) lemma order_power_n_n: "order a ([:-a,1:]^n)=n" proof (induct n) (*might be proved more concisely using nat_less_induct*) case 0 then show ?case by (metis order_root poly_1 power_0 zero_neq_one) next case (Suc n) have "order a ([:- a, 1:] ^ Suc n) = order a ([:- a, 1:] ^ n) + order a [:-a,1:]" by (metis (no_types, hide_lams) One_nat_def add_Suc_right monoid_add_class.add.right_neutral one_neq_zero order_mult pCons_eq_0_iff power_add power_eq_0_iff power_one_right) moreover have "order a [:-a,1:] = 1" unfolding order_def proof (rule Least_equality, rule notI) assume "[:- a, 1:] ^ Suc 1 dvd [:- a, 1:]" then have "degree ([:- a, 1:] ^ Suc 1) \ degree ([:- a, 1:])" by (rule dvd_imp_degree_le) auto then show False by auto next fix y assume *: "\ [:- a, 1:] ^ Suc y dvd [:- a, 1:]" show "1 \ y" proof (rule ccontr) assume "\ 1 \ y" then have "y = 0" by auto then have "[:- a, 1:] ^ Suc y dvd [:- a, 1:]" by auto with * show False by auto qed qed ultimately show ?case using Suc by auto qed lemma order_0_monom [simp]: "c \ 0 \ order 0 (monom c n) = n" using order_power_n_n[of 0 n] by (simp add: monom_altdef order_smult) lemma dvd_imp_order_le: "q \ 0 \ p dvd q \ Polynomial.order a p \ Polynomial.order a q" by (auto simp: order_mult elim: dvdE) text \Now justify the standard squarefree decomposition, i.e. \f / gcd f f'\.\ lemma order_divides: "[:-a, 1:] ^ n dvd p \ p = 0 \ n \ order a p" apply (cases "p = 0") apply auto apply (drule order_2 [where a=a and p=p]) apply (metis not_less_eq_eq power_le_dvd) apply (erule power_le_dvd [OF order_1]) done lemma order_decomp: assumes "p \ 0" shows "\q. p = [:- a, 1:] ^ order a p * q \ \ [:- a, 1:] dvd q" proof - from assms have *: "[:- a, 1:] ^ order a p dvd p" and **: "\ [:- a, 1:] ^ Suc (order a p) dvd p" by (auto dest: order) from * obtain q where q: "p = [:- a, 1:] ^ order a p * q" .. with ** have "\ [:- a, 1:] ^ Suc (order a p) dvd [:- a, 1:] ^ order a p * q" by simp then have "\ [:- a, 1:] ^ order a p * [:- a, 1:] dvd [:- a, 1:] ^ order a p * q" by simp with idom_class.dvd_mult_cancel_left [of "[:- a, 1:] ^ order a p" "[:- a, 1:]" q] have "\ [:- a, 1:] dvd q" by auto with q show ?thesis by blast qed lemma monom_1_dvd_iff: "p \ 0 \ monom 1 n dvd p \ n \ order 0 p" using order_divides[of 0 n p] by (simp add: monom_altdef) subsection \Additional induction rules on polynomials\ text \ An induction rule for induction over the roots of a polynomial with a certain property. (e.g. all positive roots) \ lemma poly_root_induct [case_names 0 no_roots root]: fixes p :: "'a :: idom poly" assumes "Q 0" and "\p. (\a. P a \ poly p a \ 0) \ Q p" and "\a p. P a \ Q p \ Q ([:a, -1:] * p)" shows "Q p" proof (induction "degree p" arbitrary: p rule: less_induct) case (less p) show ?case proof (cases "p = 0") case True with assms(1) show ?thesis by simp next case False show ?thesis proof (cases "\a. P a \ poly p a = 0") case False then show ?thesis by (intro assms(2)) blast next case True then obtain a where a: "P a" "poly p a = 0" by blast then have "-[:-a, 1:] dvd p" by (subst minus_dvd_iff) (simp add: poly_eq_0_iff_dvd) then obtain q where q: "p = [:a, -1:] * q" by (elim dvdE) simp with False have "q \ 0" by auto have "degree p = Suc (degree q)" by (subst q, subst degree_mult_eq) (simp_all add: \q \ 0\) then have "Q q" by (intro less) simp with a(1) have "Q ([:a, -1:] * q)" by (rule assms(3)) with q show ?thesis by simp qed qed qed lemma dropWhile_replicate_append: "dropWhile ((=) a) (replicate n a @ ys) = dropWhile ((=) a) ys" by (induct n) simp_all lemma Poly_append_replicate_0: "Poly (xs @ replicate n 0) = Poly xs" by (subst coeffs_eq_iff) (simp_all add: strip_while_def dropWhile_replicate_append) text \ An induction rule for simultaneous induction over two polynomials, prepending one coefficient in each step. \ lemma poly_induct2 [case_names 0 pCons]: assumes "P 0 0" "\a p b q. P p q \ P (pCons a p) (pCons b q)" shows "P p q" proof - define n where "n = max (length (coeffs p)) (length (coeffs q))" define xs where "xs = coeffs p @ (replicate (n - length (coeffs p)) 0)" define ys where "ys = coeffs q @ (replicate (n - length (coeffs q)) 0)" have "length xs = length ys" by (simp add: xs_def ys_def n_def) then have "P (Poly xs) (Poly ys)" by (induct rule: list_induct2) (simp_all add: assms) also have "Poly xs = p" by (simp add: xs_def Poly_append_replicate_0) also have "Poly ys = q" by (simp add: ys_def Poly_append_replicate_0) finally show ?thesis . qed subsection \Composition of polynomials\ (* Several lemmas contributed by René Thiemann and Akihisa Yamada *) definition pcompose :: "'a::comm_semiring_0 poly \ 'a poly \ 'a poly" where "pcompose p q = fold_coeffs (\a c. [:a:] + q * c) p 0" notation pcompose (infixl "\\<^sub>p" 71) lemma pcompose_0 [simp]: "pcompose 0 q = 0" by (simp add: pcompose_def) lemma pcompose_pCons: "pcompose (pCons a p) q = [:a:] + q * pcompose p q" by (cases "p = 0 \ a = 0") (auto simp add: pcompose_def) lemma pcompose_1: "pcompose 1 p = 1" for p :: "'a::comm_semiring_1 poly" by (auto simp: one_pCons pcompose_pCons) lemma poly_pcompose: "poly (pcompose p q) x = poly p (poly q x)" by (induct p) (simp_all add: pcompose_pCons) lemma degree_pcompose_le: "degree (pcompose p q) \ degree p * degree q" apply (induct p) apply simp apply (simp add: pcompose_pCons) apply clarify apply (rule degree_add_le) apply simp apply (rule order_trans [OF degree_mult_le]) apply simp done lemma pcompose_add: "pcompose (p + q) r = pcompose p r + pcompose q r" for p q r :: "'a::{comm_semiring_0, ab_semigroup_add} poly" proof (induction p q rule: poly_induct2) case 0 then show ?case by simp next case (pCons a p b q) have "pcompose (pCons a p + pCons b q) r = [:a + b:] + r * pcompose p r + r * pcompose q r" by (simp_all add: pcompose_pCons pCons.IH algebra_simps) also have "[:a + b:] = [:a:] + [:b:]" by simp also have "\ + r * pcompose p r + r * pcompose q r = pcompose (pCons a p) r + pcompose (pCons b q) r" by (simp only: pcompose_pCons add_ac) finally show ?case . qed lemma pcompose_uminus: "pcompose (-p) r = -pcompose p r" for p r :: "'a::comm_ring poly" by (induct p) (simp_all add: pcompose_pCons) lemma pcompose_diff: "pcompose (p - q) r = pcompose p r - pcompose q r" for p q r :: "'a::comm_ring poly" using pcompose_add[of p "-q"] by (simp add: pcompose_uminus) lemma pcompose_smult: "pcompose (smult a p) r = smult a (pcompose p r)" for p r :: "'a::comm_semiring_0 poly" by (induct p) (simp_all add: pcompose_pCons pcompose_add smult_add_right) lemma pcompose_mult: "pcompose (p * q) r = pcompose p r * pcompose q r" for p q r :: "'a::comm_semiring_0 poly" by (induct p arbitrary: q) (simp_all add: pcompose_add pcompose_smult pcompose_pCons algebra_simps) lemma pcompose_assoc: "pcompose p (pcompose q r) = pcompose (pcompose p q) r" for p q r :: "'a::comm_semiring_0 poly" by (induct p arbitrary: q) (simp_all add: pcompose_pCons pcompose_add pcompose_mult) lemma pcompose_idR[simp]: "pcompose p [: 0, 1 :] = p" for p :: "'a::comm_semiring_1 poly" by (induct p) (simp_all add: pcompose_pCons) lemma pcompose_sum: "pcompose (sum f A) p = sum (\i. pcompose (f i) p) A" by (induct A rule: infinite_finite_induct) (simp_all add: pcompose_1 pcompose_add) lemma pcompose_prod: "pcompose (prod f A) p = prod (\i. pcompose (f i) p) A" by (induct A rule: infinite_finite_induct) (simp_all add: pcompose_1 pcompose_mult) lemma pcompose_const [simp]: "pcompose [:a:] q = [:a:]" by (subst pcompose_pCons) simp lemma pcompose_0': "pcompose p 0 = [:coeff p 0:]" by (induct p) (auto simp add: pcompose_pCons) lemma degree_pcompose: "degree (pcompose p q) = degree p * degree q" for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" proof (induct p) case 0 then show ?case by auto next case (pCons a p) consider "degree (q * pcompose p q) = 0" | "degree (q * pcompose p q) > 0" by blast then show ?case proof cases case prems: 1 show ?thesis proof (cases "p = 0") case True then show ?thesis by auto next case False from prems have "degree q = 0 \ pcompose p q = 0" by (auto simp add: degree_mult_eq_0) moreover have False if "pcompose p q = 0" "degree q \ 0" proof - from pCons.hyps(2) that have "degree p = 0" by auto then obtain a1 where "p = [:a1:]" by (metis degree_pCons_eq_if old.nat.distinct(2) pCons_cases) with \pcompose p q = 0\ \p \ 0\ show False by auto qed ultimately have "degree (pCons a p) * degree q = 0" by auto moreover have "degree (pcompose (pCons a p) q) = 0" proof - from prems have "0 = max (degree [:a:]) (degree (q * pcompose p q))" by simp also have "\ \ degree ([:a:] + q * pcompose p q)" by (rule degree_add_le_max) finally show ?thesis by (auto simp add: pcompose_pCons) qed ultimately show ?thesis by simp qed next case prems: 2 then have "p \ 0" "q \ 0" "pcompose p q \ 0" by auto from prems degree_add_eq_right [of "[:a:]"] have "degree (pcompose (pCons a p) q) = degree (q * pcompose p q)" by (auto simp: pcompose_pCons) with pCons.hyps(2) degree_mult_eq[OF \q\0\ \pcompose p q\0\] show ?thesis by auto qed qed lemma pcompose_eq_0: fixes p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" assumes "pcompose p q = 0" "degree q > 0" shows "p = 0" proof - from assms degree_pcompose [of p q] have "degree p = 0" by auto then obtain a where "p = [:a:]" by (metis degree_pCons_eq_if gr0_conv_Suc neq0_conv pCons_cases) with assms(1) have "a = 0" by auto with \p = [:a:]\ show ?thesis by simp qed lemma lead_coeff_comp: fixes p q :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" assumes "degree q > 0" shows "lead_coeff (pcompose p q) = lead_coeff p * lead_coeff q ^ (degree p)" proof (induct p) case 0 then show ?case by auto next case (pCons a p) consider "degree (q * pcompose p q) = 0" | "degree (q * pcompose p q) > 0" by blast then show ?case proof cases case prems: 1 then have "pcompose p q = 0" by (metis assms degree_0 degree_mult_eq_0 neq0_conv) with pcompose_eq_0[OF _ \degree q > 0\] have "p = 0" by simp then show ?thesis by auto next case prems: 2 then have "degree [:a:] < degree (q * pcompose p q)" by simp then have "lead_coeff ([:a:] + q * p \\<^sub>p q) = lead_coeff (q * p \\<^sub>p q)" by (rule lead_coeff_add_le) then have "lead_coeff (pcompose (pCons a p) q) = lead_coeff (q * pcompose p q)" by (simp add: pcompose_pCons) also have "\ = lead_coeff q * (lead_coeff p * lead_coeff q ^ degree p)" using pCons.hyps(2) lead_coeff_mult[of q "pcompose p q"] by simp also have "\ = lead_coeff p * lead_coeff q ^ (degree p + 1)" by (auto simp: mult_ac) finally show ?thesis by auto qed qed subsection \Shifting polynomials\ definition poly_shift :: "nat \ 'a::zero poly \ 'a poly" where "poly_shift n p = Abs_poly (\i. coeff p (i + n))" lemma nth_default_drop: "nth_default x (drop n xs) m = nth_default x xs (m + n)" by (auto simp add: nth_default_def add_ac) lemma nth_default_take: "nth_default x (take n xs) m = (if m < n then nth_default x xs m else x)" by (auto simp add: nth_default_def add_ac) lemma coeff_poly_shift: "coeff (poly_shift n p) i = coeff p (i + n)" proof - from MOST_coeff_eq_0[of p] obtain m where "\k>m. coeff p k = 0" by (auto simp: MOST_nat) then have "\k>m. coeff p (k + n) = 0" by auto then have "\\<^sub>\k. coeff p (k + n) = 0" by (auto simp: MOST_nat) then show ?thesis by (simp add: poly_shift_def poly.Abs_poly_inverse) qed lemma poly_shift_id [simp]: "poly_shift 0 = (\x. x)" by (simp add: poly_eq_iff fun_eq_iff coeff_poly_shift) lemma poly_shift_0 [simp]: "poly_shift n 0 = 0" by (simp add: poly_eq_iff coeff_poly_shift) lemma poly_shift_1: "poly_shift n 1 = (if n = 0 then 1 else 0)" by (simp add: poly_eq_iff coeff_poly_shift) lemma poly_shift_monom: "poly_shift n (monom c m) = (if m \ n then monom c (m - n) else 0)" by (auto simp add: poly_eq_iff coeff_poly_shift) lemma coeffs_shift_poly [code abstract]: "coeffs (poly_shift n p) = drop n (coeffs p)" proof (cases "p = 0") case True then show ?thesis by simp next case False then show ?thesis by (intro coeffs_eqI) (simp_all add: coeff_poly_shift nth_default_drop nth_default_coeffs_eq) qed subsection \Truncating polynomials\ definition poly_cutoff where "poly_cutoff n p = Abs_poly (\k. if k < n then coeff p k else 0)" lemma coeff_poly_cutoff: "coeff (poly_cutoff n p) k = (if k < n then coeff p k else 0)" unfolding poly_cutoff_def by (subst poly.Abs_poly_inverse) (auto simp: MOST_nat intro: exI[of _ n]) lemma poly_cutoff_0 [simp]: "poly_cutoff n 0 = 0" by (simp add: poly_eq_iff coeff_poly_cutoff) lemma poly_cutoff_1 [simp]: "poly_cutoff n 1 = (if n = 0 then 0 else 1)" by (simp add: poly_eq_iff coeff_poly_cutoff) lemma coeffs_poly_cutoff [code abstract]: "coeffs (poly_cutoff n p) = strip_while ((=) 0) (take n (coeffs p))" proof (cases "strip_while ((=) 0) (take n (coeffs p)) = []") case True then have "coeff (poly_cutoff n p) k = 0" for k unfolding coeff_poly_cutoff by (auto simp: nth_default_coeffs_eq [symmetric] nth_default_def set_conv_nth) then have "poly_cutoff n p = 0" by (simp add: poly_eq_iff) then show ?thesis by (subst True) simp_all next case False have "no_trailing ((=) 0) (strip_while ((=) 0) (take n (coeffs p)))" by simp with False have "last (strip_while ((=) 0) (take n (coeffs p))) \ 0" unfolding no_trailing_unfold by auto then show ?thesis by (intro coeffs_eqI) (simp_all add: coeff_poly_cutoff nth_default_take nth_default_coeffs_eq) qed subsection \Reflecting polynomials\ definition reflect_poly :: "'a::zero poly \ 'a poly" where "reflect_poly p = Poly (rev (coeffs p))" lemma coeffs_reflect_poly [code abstract]: "coeffs (reflect_poly p) = rev (dropWhile ((=) 0) (coeffs p))" by (simp add: reflect_poly_def) lemma reflect_poly_0 [simp]: "reflect_poly 0 = 0" by (simp add: reflect_poly_def) lemma reflect_poly_1 [simp]: "reflect_poly 1 = 1" by (simp add: reflect_poly_def one_pCons) lemma coeff_reflect_poly: "coeff (reflect_poly p) n = (if n > degree p then 0 else coeff p (degree p - n))" by (cases "p = 0") (auto simp add: reflect_poly_def nth_default_def rev_nth degree_eq_length_coeffs coeffs_nth not_less dest: le_imp_less_Suc) lemma coeff_0_reflect_poly_0_iff [simp]: "coeff (reflect_poly p) 0 = 0 \ p = 0" by (simp add: coeff_reflect_poly) lemma reflect_poly_at_0_eq_0_iff [simp]: "poly (reflect_poly p) 0 = 0 \ p = 0" by (simp add: coeff_reflect_poly poly_0_coeff_0) lemma reflect_poly_pCons': "p \ 0 \ reflect_poly (pCons c p) = reflect_poly p + monom c (Suc (degree p))" by (intro poly_eqI) (auto simp: coeff_reflect_poly coeff_pCons not_less Suc_diff_le split: nat.split) lemma reflect_poly_const [simp]: "reflect_poly [:a:] = [:a:]" by (cases "a = 0") (simp_all add: reflect_poly_def) lemma poly_reflect_poly_nz: "x \ 0 \ poly (reflect_poly p) x = x ^ degree p * poly p (inverse x)" for x :: "'a::field" by (induct rule: pCons_induct) (simp_all add: field_simps reflect_poly_pCons' poly_monom) lemma coeff_0_reflect_poly [simp]: "coeff (reflect_poly p) 0 = lead_coeff p" by (simp add: coeff_reflect_poly) lemma poly_reflect_poly_0 [simp]: "poly (reflect_poly p) 0 = lead_coeff p" by (simp add: poly_0_coeff_0) lemma reflect_poly_reflect_poly [simp]: "coeff p 0 \ 0 \ reflect_poly (reflect_poly p) = p" by (cases p rule: pCons_cases) (simp add: reflect_poly_def ) lemma degree_reflect_poly_le: "degree (reflect_poly p) \ degree p" by (simp add: degree_eq_length_coeffs coeffs_reflect_poly length_dropWhile_le diff_le_mono) lemma reflect_poly_pCons: "a \ 0 \ reflect_poly (pCons a p) = Poly (rev (a # coeffs p))" by (subst coeffs_eq_iff) (simp add: coeffs_reflect_poly) lemma degree_reflect_poly_eq [simp]: "coeff p 0 \ 0 \ degree (reflect_poly p) = degree p" by (cases p rule: pCons_cases) (simp add: reflect_poly_pCons degree_eq_length_coeffs) (* TODO: does this work with zero divisors as well? Probably not. *) lemma reflect_poly_mult: "reflect_poly (p * q) = reflect_poly p * reflect_poly q" for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" proof (cases "p = 0 \ q = 0") case False then have [simp]: "p \ 0" "q \ 0" by auto show ?thesis proof (rule poly_eqI) show "coeff (reflect_poly (p * q)) i = coeff (reflect_poly p * reflect_poly q) i" for i proof (cases "i \ degree (p * q)") case True define A where "A = {..i} \ {i - degree q..degree p}" define B where "B = {..degree p} \ {degree p - i..degree (p*q) - i}" let ?f = "\j. degree p - j" from True have "coeff (reflect_poly (p * q)) i = coeff (p * q) (degree (p * q) - i)" by (simp add: coeff_reflect_poly) also have "\ = (\j\degree (p * q) - i. coeff p j * coeff q (degree (p * q) - i - j))" by (simp add: coeff_mult) also have "\ = (\j\B. coeff p j * coeff q (degree (p * q) - i - j))" by (intro sum.mono_neutral_right) (auto simp: B_def degree_mult_eq not_le coeff_eq_0) also from True have "\ = (\j\A. coeff p (degree p - j) * coeff q (degree q - (i - j)))" by (intro sum.reindex_bij_witness[of _ ?f ?f]) (auto simp: A_def B_def degree_mult_eq add_ac) also have "\ = (\j\i. if j \ {i - degree q..degree p} then coeff p (degree p - j) * coeff q (degree q - (i - j)) else 0)" by (subst sum.inter_restrict [symmetric]) (simp_all add: A_def) also have "\ = coeff (reflect_poly p * reflect_poly q) i" by (fastforce simp: coeff_mult coeff_reflect_poly intro!: sum.cong) finally show ?thesis . qed (auto simp: coeff_mult coeff_reflect_poly coeff_eq_0 degree_mult_eq intro!: sum.neutral) qed qed auto lemma reflect_poly_smult: "reflect_poly (smult c p) = smult c (reflect_poly p)" for p :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" using reflect_poly_mult[of "[:c:]" p] by simp lemma reflect_poly_power: "reflect_poly (p ^ n) = reflect_poly p ^ n" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" by (induct n) (simp_all add: reflect_poly_mult) lemma reflect_poly_prod: "reflect_poly (prod f A) = prod (\x. reflect_poly (f x)) A" for f :: "_ \ _::{comm_semiring_0,semiring_no_zero_divisors} poly" by (induct A rule: infinite_finite_induct) (simp_all add: reflect_poly_mult) lemma reflect_poly_prod_list: "reflect_poly (prod_list xs) = prod_list (map reflect_poly xs)" for xs :: "_::{comm_semiring_0,semiring_no_zero_divisors} poly list" by (induct xs) (simp_all add: reflect_poly_mult) lemma reflect_poly_Poly_nz: "no_trailing (HOL.eq 0) xs \ reflect_poly (Poly xs) = Poly (rev xs)" by (simp add: reflect_poly_def) lemmas reflect_poly_simps = reflect_poly_0 reflect_poly_1 reflect_poly_const reflect_poly_smult reflect_poly_mult reflect_poly_power reflect_poly_prod reflect_poly_prod_list subsection \Derivatives\ function pderiv :: "('a :: {comm_semiring_1,semiring_no_zero_divisors}) poly \ 'a poly" where "pderiv (pCons a p) = (if p = 0 then 0 else p + pCons 0 (pderiv p))" by (auto intro: pCons_cases) termination pderiv by (relation "measure degree") simp_all declare pderiv.simps[simp del] lemma pderiv_0 [simp]: "pderiv 0 = 0" using pderiv.simps [of 0 0] by simp lemma pderiv_pCons: "pderiv (pCons a p) = p + pCons 0 (pderiv p)" by (simp add: pderiv.simps) lemma pderiv_1 [simp]: "pderiv 1 = 0" by (simp add: one_pCons pderiv_pCons) lemma pderiv_of_nat [simp]: "pderiv (of_nat n) = 0" and pderiv_numeral [simp]: "pderiv (numeral m) = 0" by (simp_all add: of_nat_poly numeral_poly pderiv_pCons) lemma coeff_pderiv: "coeff (pderiv p) n = of_nat (Suc n) * coeff p (Suc n)" by (induct p arbitrary: n) (auto simp add: pderiv_pCons coeff_pCons algebra_simps split: nat.split) fun pderiv_coeffs_code :: "'a::{comm_semiring_1,semiring_no_zero_divisors} \ 'a list \ 'a list" where "pderiv_coeffs_code f (x # xs) = cCons (f * x) (pderiv_coeffs_code (f+1) xs)" | "pderiv_coeffs_code f [] = []" definition pderiv_coeffs :: "'a::{comm_semiring_1,semiring_no_zero_divisors} list \ 'a list" where "pderiv_coeffs xs = pderiv_coeffs_code 1 (tl xs)" (* Efficient code for pderiv contributed by René Thiemann and Akihisa Yamada *) lemma pderiv_coeffs_code: "nth_default 0 (pderiv_coeffs_code f xs) n = (f + of_nat n) * nth_default 0 xs n" proof (induct xs arbitrary: f n) case Nil then show ?case by simp next case (Cons x xs) show ?case proof (cases n) case 0 then show ?thesis by (cases "pderiv_coeffs_code (f + 1) xs = [] \ f * x = 0") (auto simp: cCons_def) next case n: (Suc m) show ?thesis proof (cases "pderiv_coeffs_code (f + 1) xs = [] \ f * x = 0") case False then have "nth_default 0 (pderiv_coeffs_code f (x # xs)) n = nth_default 0 (pderiv_coeffs_code (f + 1) xs) m" by (auto simp: cCons_def n) also have "\ = (f + of_nat n) * nth_default 0 xs m" by (simp add: Cons n add_ac) finally show ?thesis by (simp add: n) next case True have empty: "pderiv_coeffs_code g xs = [] \ g + of_nat m = 0 \ nth_default 0 xs m = 0" for g proof (induct xs arbitrary: g m) case Nil then show ?case by simp next case (Cons x xs) from Cons(2) have empty: "pderiv_coeffs_code (g + 1) xs = []" and g: "g = 0 \ x = 0" by (auto simp: cCons_def split: if_splits) note IH = Cons(1)[OF empty] from IH[of m] IH[of "m - 1"] g show ?case by (cases m) (auto simp: field_simps) qed from True have "nth_default 0 (pderiv_coeffs_code f (x # xs)) n = 0" by (auto simp: cCons_def n) moreover from True have "(f + of_nat n) * nth_default 0 (x # xs) n = 0" by (simp add: n) (use empty[of "f+1"] in \auto simp: field_simps\) ultimately show ?thesis by simp qed qed qed lemma coeffs_pderiv_code [code abstract]: "coeffs (pderiv p) = pderiv_coeffs (coeffs p)" unfolding pderiv_coeffs_def proof (rule coeffs_eqI, unfold pderiv_coeffs_code coeff_pderiv, goal_cases) case (1 n) have id: "coeff p (Suc n) = nth_default 0 (map (\i. coeff p (Suc i)) [0.. degree p = 0" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly" apply (rule iffI) apply (cases p) apply simp apply (simp add: poly_eq_iff coeff_pderiv del: of_nat_Suc) apply (simp add: poly_eq_iff coeff_pderiv coeff_eq_0) done lemma degree_pderiv: "degree (pderiv p) = degree p - 1" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly" apply (rule order_antisym [OF degree_le]) apply (simp add: coeff_pderiv coeff_eq_0) apply (cases "degree p") apply simp apply (rule le_degree) apply (simp add: coeff_pderiv del: of_nat_Suc) apply (metis degree_0 leading_coeff_0_iff nat.distinct(1)) done lemma not_dvd_pderiv: fixes p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly" assumes "degree p \ 0" shows "\ p dvd pderiv p" proof assume dvd: "p dvd pderiv p" then obtain q where p: "pderiv p = p * q" unfolding dvd_def by auto from dvd have le: "degree p \ degree (pderiv p)" by (simp add: assms dvd_imp_degree_le pderiv_eq_0_iff) from assms and this [unfolded degree_pderiv] show False by auto qed lemma dvd_pderiv_iff [simp]: "p dvd pderiv p \ degree p = 0" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly" using not_dvd_pderiv[of p] by (auto simp: pderiv_eq_0_iff [symmetric]) lemma pderiv_singleton [simp]: "pderiv [:a:] = 0" by (simp add: pderiv_pCons) lemma pderiv_add: "pderiv (p + q) = pderiv p + pderiv q" by (rule poly_eqI) (simp add: coeff_pderiv algebra_simps) lemma pderiv_minus: "pderiv (- p :: 'a :: idom poly) = - pderiv p" by (rule poly_eqI) (simp add: coeff_pderiv algebra_simps) lemma pderiv_diff: "pderiv ((p :: _ :: idom poly) - q) = pderiv p - pderiv q" by (rule poly_eqI) (simp add: coeff_pderiv algebra_simps) lemma pderiv_smult: "pderiv (smult a p) = smult a (pderiv p)" by (rule poly_eqI) (simp add: coeff_pderiv algebra_simps) lemma pderiv_mult: "pderiv (p * q) = p * pderiv q + q * pderiv p" by (induct p) (auto simp: pderiv_add pderiv_smult pderiv_pCons algebra_simps) lemma pderiv_power_Suc: "pderiv (p ^ Suc n) = smult (of_nat (Suc n)) (p ^ n) * pderiv p" apply (induct n) apply simp apply (subst power_Suc) apply (subst pderiv_mult) apply (erule ssubst) apply (simp only: of_nat_Suc smult_add_left smult_1_left) apply (simp add: algebra_simps) done lemma pderiv_pcompose: "pderiv (pcompose p q) = pcompose (pderiv p) q * pderiv q" by (induction p rule: pCons_induct) (auto simp: pcompose_pCons pderiv_add pderiv_mult pderiv_pCons pcompose_add algebra_simps) lemma pderiv_prod: "pderiv (prod f (as)) = (\a\as. prod f (as - {a}) * pderiv (f a))" proof (induct as rule: infinite_finite_induct) case (insert a as) then have id: "prod f (insert a as) = f a * prod f as" "\g. sum g (insert a as) = g a + sum g as" "insert a as - {a} = as" by auto have "prod f (insert a as - {b}) = f a * prod f (as - {b})" if "b \ as" for b proof - from \a \ as\ that have *: "insert a as - {b} = insert a (as - {b})" by auto show ?thesis unfolding * by (subst prod.insert) (use insert in auto) qed then show ?case unfolding id pderiv_mult insert(3) sum_distrib_left by (auto simp add: ac_simps intro!: sum.cong) qed auto lemma DERIV_pow2: "DERIV (\x. x ^ Suc n) x :> real (Suc n) * (x ^ n)" by (rule DERIV_cong, rule DERIV_pow) simp declare DERIV_pow2 [simp] DERIV_pow [simp] lemma DERIV_add_const: "DERIV f x :> D \ DERIV (\x. a + f x :: 'a::real_normed_field) x :> D" by (rule DERIV_cong, rule DERIV_add) auto lemma poly_DERIV [simp]: "DERIV (\x. poly p x) x :> poly (pderiv p) x" by (induct p) (auto intro!: derivative_eq_intros simp add: pderiv_pCons) lemma poly_isCont[simp]: fixes x::"'a::real_normed_field" shows "isCont (\x. poly p x) x" by (rule poly_DERIV [THEN DERIV_isCont]) lemma tendsto_poly [tendsto_intros]: "(f \ a) F \ ((\x. poly p (f x)) \ poly p a) F" for f :: "_ \ 'a::real_normed_field" by (rule isCont_tendsto_compose [OF poly_isCont]) lemma continuous_within_poly: "continuous (at z within s) (poly p)" for z :: "'a::{real_normed_field}" by (simp add: continuous_within tendsto_poly) lemma continuous_poly [continuous_intros]: "continuous F f \ continuous F (\x. poly p (f x))" for f :: "_ \ 'a::real_normed_field" unfolding continuous_def by (rule tendsto_poly) lemma continuous_on_poly [continuous_intros]: fixes p :: "'a :: {real_normed_field} poly" assumes "continuous_on A f" shows "continuous_on A (\x. poly p (f x))" by (metis DERIV_continuous_on assms continuous_on_compose2 poly_DERIV subset_UNIV) text \Consequences of the derivative theorem above.\ lemma poly_differentiable[simp]: "(\x. poly p x) differentiable (at x)" for x :: real by (simp add: real_differentiable_def) (blast intro: poly_DERIV) lemma poly_IVT_pos: "a < b \ poly p a < 0 \ 0 < poly p b \ \x. a < x \ x < b \ poly p x = 0" for a b :: real using IVT_objl [of "poly p" a 0 b] by (auto simp add: order_le_less) lemma poly_IVT_neg: "a < b \ 0 < poly p a \ poly p b < 0 \ \x. a < x \ x < b \ poly p x = 0" for a b :: real using poly_IVT_pos [where p = "- p"] by simp lemma poly_IVT: "a < b \ poly p a * poly p b < 0 \ \x>a. x < b \ poly p x = 0" for p :: "real poly" by (metis less_not_sym mult_less_0_iff poly_IVT_neg poly_IVT_pos) lemma poly_MVT: "a < b \ \x. a < x \ x < b \ poly p b - poly p a = (b - a) * poly (pderiv p) x" for a b :: real using MVT [of a b "poly p"] apply simp by (metis (full_types) DERIV_continuous_on DERIV_unique has_field_derivative_at_within poly_DERIV) lemma poly_MVT': fixes a b :: real assumes "{min a b..max a b} \ A" shows "\x\A. poly p b - poly p a = (b - a) * poly (pderiv p) x" proof (cases a b rule: linorder_cases) case less from poly_MVT[OF less, of p] guess x by (elim exE conjE) then show ?thesis by (intro bexI[of _ x]) (auto intro!: subsetD[OF assms]) next case greater from poly_MVT[OF greater, of p] guess x by (elim exE conjE) then show ?thesis by (intro bexI[of _ x]) (auto simp: algebra_simps intro!: subsetD[OF assms]) qed (use assms in auto) lemma poly_pinfty_gt_lc: fixes p :: "real poly" assumes "lead_coeff p > 0" shows "\n. \ x \ n. poly p x \ lead_coeff p" using assms proof (induct p) case 0 then show ?case by auto next case (pCons a p) from this(1) consider "a \ 0" "p = 0" | "p \ 0" by auto then show ?case proof cases case 1 then show ?thesis by auto next case 2 with pCons obtain n1 where gte_lcoeff: "\x\n1. lead_coeff p \ poly p x" by auto from pCons(3) \p \ 0\ have gt_0: "lead_coeff p > 0" by auto define n where "n = max n1 (1 + \a\ / lead_coeff p)" have "lead_coeff (pCons a p) \ poly (pCons a p) x" if "n \ x" for x proof - from gte_lcoeff that have "lead_coeff p \ poly p x" by (auto simp: n_def) with gt_0 have "\a\ / lead_coeff p \ \a\ / poly p x" and "poly p x > 0" by (auto intro: frac_le) with \n \ x\[unfolded n_def] have "x \ 1 + \a\ / poly p x" by auto with \lead_coeff p \ poly p x\ \poly p x > 0\ \p \ 0\ show "lead_coeff (pCons a p) \ poly (pCons a p) x" by (auto simp: field_simps) qed then show ?thesis by blast qed qed lemma lemma_order_pderiv1: "pderiv ([:- a, 1:] ^ Suc n * q) = [:- a, 1:] ^ Suc n * pderiv q + smult (of_nat (Suc n)) (q * [:- a, 1:] ^ n)" by (simp only: pderiv_mult pderiv_power_Suc) (simp del: power_Suc of_nat_Suc add: pderiv_pCons) lemma lemma_order_pderiv: fixes p :: "'a :: field_char_0 poly" assumes n: "0 < n" and pd: "pderiv p \ 0" and pe: "p = [:- a, 1:] ^ n * q" and nd: "\ [:- a, 1:] dvd q" shows "n = Suc (order a (pderiv p))" proof - from assms have "pderiv ([:- a, 1:] ^ n * q) \ 0" by auto from assms obtain n' where "n = Suc n'" "0 < Suc n'" "pderiv ([:- a, 1:] ^ Suc n' * q) \ 0" by (cases n) auto have *: "k dvd k * pderiv q + smult (of_nat (Suc n')) l \ k dvd l" for k l by (auto simp del: of_nat_Suc simp: dvd_add_right_iff dvd_smult_iff) have "n' = order a (pderiv ([:- a, 1:] ^ Suc n' * q))" proof (rule order_unique_lemma) show "[:- a, 1:] ^ n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)" apply (subst lemma_order_pderiv1) apply (rule dvd_add) apply (metis dvdI dvd_mult2 power_Suc2) apply (metis dvd_smult dvd_triv_right) done show "\ [:- a, 1:] ^ Suc n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)" apply (subst lemma_order_pderiv1) apply (metis * nd dvd_mult_cancel_right power_not_zero pCons_eq_0_iff power_Suc zero_neq_one) done qed then show ?thesis by (metis \n = Suc n'\ pe) qed lemma order_pderiv: "pderiv p \ 0 \ order a p \ 0 \ order a p = Suc (order a (pderiv p))" for p :: "'a::field_char_0 poly" apply (cases "p = 0") apply simp apply (drule_tac a = a and p = p in order_decomp) using neq0_conv apply (blast intro: lemma_order_pderiv) done lemma poly_squarefree_decomp_order: fixes p :: "'a::field_char_0 poly" assumes "pderiv p \ 0" and p: "p = q * d" and p': "pderiv p = e * d" and d: "d = r * p + s * pderiv p" shows "order a q = (if order a p = 0 then 0 else 1)" proof (rule classical) assume 1: "\ ?thesis" from \pderiv p \ 0\ have "p \ 0" by auto with p have "order a p = order a q + order a d" by (simp add: order_mult) with 1 have "order a p \ 0" by (auto split: if_splits) from \pderiv p \ 0\ \pderiv p = e * d\ have "order a (pderiv p) = order a e + order a d" by (simp add: order_mult) from \pderiv p \ 0\ \order a p \ 0\ have "order a p = Suc (order a (pderiv p))" by (rule order_pderiv) from \p \ 0\ \p = q * d\ have "d \ 0" by simp have "([:-a, 1:] ^ (order a (pderiv p))) dvd d" apply (simp add: d) apply (rule dvd_add) apply (rule dvd_mult) apply (simp add: order_divides \p \ 0\ \order a p = Suc (order a (pderiv p))\) apply (rule dvd_mult) apply (simp add: order_divides) done with \d \ 0\ have "order a (pderiv p) \ order a d" by (simp add: order_divides) show ?thesis using \order a p = order a q + order a d\ and \order a (pderiv p) = order a e + order a d\ and \order a p = Suc (order a (pderiv p))\ and \order a (pderiv p) \ order a d\ by auto qed lemma poly_squarefree_decomp_order2: "pderiv p \ 0 \ p = q * d \ pderiv p = e * d \ d = r * p + s * pderiv p \ \a. order a q = (if order a p = 0 then 0 else 1)" for p :: "'a::field_char_0 poly" by (blast intro: poly_squarefree_decomp_order) lemma order_pderiv2: "pderiv p \ 0 \ order a p \ 0 \ order a (pderiv p) = n \ order a p = Suc n" for p :: "'a::field_char_0 poly" by (auto dest: order_pderiv) definition rsquarefree :: "'a::idom poly \ bool" where "rsquarefree p \ p \ 0 \ (\a. order a p = 0 \ order a p = 1)" lemma pderiv_iszero: "pderiv p = 0 \ \h. p = [:h:]" for p :: "'a::{semidom,semiring_char_0} poly" by (cases p) (auto simp: pderiv_eq_0_iff split: if_splits) lemma rsquarefree_roots: "rsquarefree p \ (\a. \ (poly p a = 0 \ poly (pderiv p) a = 0))" for p :: "'a::field_char_0 poly" apply (simp add: rsquarefree_def) apply (case_tac "p = 0") apply simp apply simp apply (case_tac "pderiv p = 0") apply simp apply (drule pderiv_iszero, clarsimp) apply (metis coeff_0 coeff_pCons_0 degree_pCons_0 le0 le_antisym order_degree) apply (force simp add: order_root order_pderiv2) done lemma poly_squarefree_decomp: fixes p :: "'a::field_char_0 poly" assumes "pderiv p \ 0" and "p = q * d" and "pderiv p = e * d" and "d = r * p + s * pderiv p" shows "rsquarefree q \ (\a. poly q a = 0 \ poly p a = 0)" proof - from \pderiv p \ 0\ have "p \ 0" by auto with \p = q * d\ have "q \ 0" by simp from assms have "\a. order a q = (if order a p = 0 then 0 else 1)" by (rule poly_squarefree_decomp_order2) with \p \ 0\ \q \ 0\ show ?thesis by (simp add: rsquarefree_def order_root) qed subsection \Algebraic numbers\ text \ Algebraic numbers can be defined in two equivalent ways: all real numbers that are roots of rational polynomials or of integer polynomials. The Algebraic-Numbers AFP entry uses the rational definition, but we need the integer definition. The equivalence is obvious since any rational polynomial can be multiplied with the LCM of its coefficients, yielding an integer polynomial with the same roots. \ definition algebraic :: "'a :: field_char_0 \ bool" where "algebraic x \ (\p. (\i. coeff p i \ \) \ p \ 0 \ poly p x = 0)" lemma algebraicI: "(\i. coeff p i \ \) \ p \ 0 \ poly p x = 0 \ algebraic x" unfolding algebraic_def by blast lemma algebraicE: assumes "algebraic x" obtains p where "\i. coeff p i \ \" "p \ 0" "poly p x = 0" using assms unfolding algebraic_def by blast lemma algebraic_altdef: "algebraic x \ (\p. (\i. coeff p i \ \) \ p \ 0 \ poly p x = 0)" for p :: "'a::field_char_0 poly" proof safe fix p assume rat: "\i. coeff p i \ \" and root: "poly p x = 0" and nz: "p \ 0" define cs where "cs = coeffs p" from rat have "\c\range (coeff p). \c'. c = of_rat c'" unfolding Rats_def by blast then obtain f where f: "coeff p i = of_rat (f (coeff p i))" for i by (subst (asm) bchoice_iff) blast define cs' where "cs' = map (quotient_of \ f) (coeffs p)" define d where "d = Lcm (set (map snd cs'))" define p' where "p' = smult (of_int d) p" have "coeff p' n \ \" for n proof (cases "n \ degree p") case True define c where "c = coeff p n" define a where "a = fst (quotient_of (f (coeff p n)))" define b where "b = snd (quotient_of (f (coeff p n)))" have b_pos: "b > 0" unfolding b_def using quotient_of_denom_pos' by simp have "coeff p' n = of_int d * coeff p n" by (simp add: p'_def) also have "coeff p n = of_rat (of_int a / of_int b)" unfolding a_def b_def by (subst quotient_of_div [of "f (coeff p n)", symmetric]) (simp_all add: f [symmetric]) also have "of_int d * \ = of_rat (of_int (a*d) / of_int b)" by (simp add: of_rat_mult of_rat_divide) also from nz True have "b \ snd ` set cs'" by (force simp: cs'_def o_def b_def coeffs_def simp del: upt_Suc) then have "b dvd (a * d)" by (simp add: d_def) then have "of_int (a * d) / of_int b \ (\ :: rat set)" by (rule of_int_divide_in_Ints) then have "of_rat (of_int (a * d) / of_int b) \ \" by (elim Ints_cases) auto finally show ?thesis . next case False then show ?thesis by (auto simp: p'_def not_le coeff_eq_0) qed moreover have "set (map snd cs') \ {0<..}" unfolding cs'_def using quotient_of_denom_pos' by (auto simp: coeffs_def simp del: upt_Suc) then have "d \ 0" unfolding d_def by (induct cs') simp_all with nz have "p' \ 0" by (simp add: p'_def) moreover from root have "poly p' x = 0" by (simp add: p'_def) ultimately show "algebraic x" unfolding algebraic_def by blast next assume "algebraic x" then obtain p where p: "coeff p i \ \" "poly p x = 0" "p \ 0" for i by (force simp: algebraic_def) moreover have "coeff p i \ \ \ coeff p i \ \" for i by (elim Ints_cases) simp ultimately show "\p. (\i. coeff p i \ \) \ p \ 0 \ poly p x = 0" by auto qed subsection \Division of polynomials\ subsubsection \Division in general\ instantiation poly :: (idom_divide) idom_divide begin fun divide_poly_main :: "'a \ 'a poly \ 'a poly \ 'a poly \ nat \ nat \ 'a poly" where "divide_poly_main lc q r d dr (Suc n) = (let cr = coeff r dr; a = cr div lc; mon = monom a n in if False \ a * lc = cr then \ \\False \\ is only because of problem in function-package\ divide_poly_main lc (q + mon) (r - mon * d) d (dr - 1) n else 0)" | "divide_poly_main lc q r d dr 0 = q" definition divide_poly :: "'a poly \ 'a poly \ 'a poly" where "divide_poly f g = (if g = 0 then 0 else divide_poly_main (coeff g (degree g)) 0 f g (degree f) (1 + length (coeffs f) - length (coeffs g)))" lemma divide_poly_main: assumes d: "d \ 0" "lc = coeff d (degree d)" and "degree (d * r) \ dr" "divide_poly_main lc q (d * r) d dr n = q'" and "n = 1 + dr - degree d \ dr = 0 \ n = 0 \ d * r = 0" shows "q' = q + r" using assms(3-) proof (induct n arbitrary: q r dr) case (Suc n) let ?rr = "d * r" let ?a = "coeff ?rr dr" let ?qq = "?a div lc" define b where [simp]: "b = monom ?qq n" let ?rrr = "d * (r - b)" let ?qqq = "q + b" note res = Suc(3) from Suc(4) have dr: "dr = n + degree d" by auto from d have lc: "lc \ 0" by auto have "coeff (b * d) dr = coeff b n * coeff d (degree d)" proof (cases "?qq = 0") case True then show ?thesis by simp next case False then have n: "n = degree b" by (simp add: degree_monom_eq) show ?thesis unfolding n dr by (simp add: coeff_mult_degree_sum) qed also have "\ = lc * coeff b n" by (simp add: d) finally have c2: "coeff (b * d) dr = lc * coeff b n" . have rrr: "?rrr = ?rr - b * d" by (simp add: field_simps) have c1: "coeff (d * r) dr = lc * coeff r n" proof (cases "degree r = n") case True with Suc(2) show ?thesis unfolding dr using coeff_mult_degree_sum[of d r] d by (auto simp: ac_simps) next case False from dr Suc(2) have "degree r \ n" by auto (metis add.commute add_le_cancel_left d(1) degree_0 degree_mult_eq diff_is_0_eq diff_zero le_cases) with False have r_n: "degree r < n" by auto then have right: "lc * coeff r n = 0" by (simp add: coeff_eq_0) have "coeff (d * r) dr = coeff (d * r) (degree d + n)" by (simp add: dr ac_simps) also from r_n have "\ = 0" by (metis False Suc.prems(1) add.commute add_left_imp_eq coeff_degree_mult coeff_eq_0 coeff_mult_degree_sum degree_mult_le dr le_eq_less_or_eq) finally show ?thesis by (simp only: right) qed have c0: "coeff ?rrr dr = 0" and id: "lc * (coeff (d * r) dr div lc) = coeff (d * r) dr" unfolding rrr coeff_diff c2 unfolding b_def coeff_monom coeff_smult c1 using lc by auto from res[unfolded divide_poly_main.simps[of lc q] Let_def] id have res: "divide_poly_main lc ?qqq ?rrr d (dr - 1) n = q'" by (simp del: divide_poly_main.simps add: field_simps) note IH = Suc(1)[OF _ res] from Suc(4) have dr: "dr = n + degree d" by auto from Suc(2) have deg_rr: "degree ?rr \ dr" by auto have deg_bd: "degree (b * d) \ dr" unfolding dr b_def by (rule order.trans[OF degree_mult_le]) (auto simp: degree_monom_le) have "degree ?rrr \ dr" unfolding rrr by (rule degree_diff_le[OF deg_rr deg_bd]) with c0 have deg_rrr: "degree ?rrr \ (dr - 1)" by (rule coeff_0_degree_minus_1) have "n = 1 + (dr - 1) - degree d \ dr - 1 = 0 \ n = 0 \ ?rrr = 0" proof (cases dr) case 0 with Suc(4) have 0: "dr = 0" "n = 0" "degree d = 0" by auto with deg_rrr have "degree ?rrr = 0" by simp from degree_eq_zeroE[OF this] obtain a where rrr: "?rrr = [:a:]" by metis show ?thesis unfolding 0 using c0 unfolding rrr 0 by simp next case _: Suc with Suc(4) show ?thesis by auto qed from IH[OF deg_rrr this] show ?case by simp next case 0 show ?case proof (cases "r = 0") case True with 0 show ?thesis by auto next case False from d False have "degree (d * r) = degree d + degree r" by (subst degree_mult_eq) auto with 0 d show ?thesis by auto qed qed lemma divide_poly_main_0: "divide_poly_main 0 0 r d dr n = 0" proof (induct n arbitrary: r d dr) case 0 then show ?case by simp next case Suc show ?case unfolding divide_poly_main.simps[of _ _ r] Let_def by (simp add: Suc del: divide_poly_main.simps) qed lemma divide_poly: assumes g: "g \ 0" shows "(f * g) div g = (f :: 'a poly)" proof - have len: "length (coeffs f) = Suc (degree f)" if "f \ 0" for f :: "'a poly" using that unfolding degree_eq_length_coeffs by auto have "divide_poly_main (coeff g (degree g)) 0 (g * f) g (degree (g * f)) (1 + length (coeffs (g * f)) - length (coeffs g)) = (f * g) div g" by (simp add: divide_poly_def Let_def ac_simps) note main = divide_poly_main[OF g refl le_refl this] have "(f * g) div g = 0 + f" proof (rule main, goal_cases) case 1 show ?case proof (cases "f = 0") case True with g show ?thesis by (auto simp: degree_eq_length_coeffs) next case False with g have fg: "g * f \ 0" by auto show ?thesis unfolding len[OF fg] len[OF g] by auto qed qed then show ?thesis by simp qed lemma divide_poly_0: "f div 0 = 0" for f :: "'a poly" by (simp add: divide_poly_def Let_def divide_poly_main_0) instance by standard (auto simp: divide_poly divide_poly_0) end instance poly :: (idom_divide) algebraic_semidom .. lemma div_const_poly_conv_map_poly: assumes "[:c:] dvd p" shows "p div [:c:] = map_poly (\x. x div c) p" proof (cases "c = 0") case True then show ?thesis by (auto intro!: poly_eqI simp: coeff_map_poly) next case False from assms obtain q where p: "p = [:c:] * q" by (rule dvdE) moreover { have "smult c q = [:c:] * q" by simp also have "\ div [:c:] = q" by (rule nonzero_mult_div_cancel_left) (use False in auto) finally have "smult c q div [:c:] = q" . } ultimately show ?thesis by (intro poly_eqI) (auto simp: coeff_map_poly False) qed lemma is_unit_monom_0: fixes a :: "'a::field" assumes "a \ 0" shows "is_unit (monom a 0)" proof from assms show "1 = monom a 0 * monom (inverse a) 0" by (simp add: mult_monom) qed lemma is_unit_triv: "a \ 0 \ is_unit [:a:]" for a :: "'a::field" by (simp add: is_unit_monom_0 monom_0 [symmetric]) lemma is_unit_iff_degree: fixes p :: "'a::field poly" assumes "p \ 0" shows "is_unit p \ degree p = 0" (is "?lhs \ ?rhs") proof assume ?rhs then obtain a where "p = [:a:]" by (rule degree_eq_zeroE) with assms show ?lhs by (simp add: is_unit_triv) next assume ?lhs then obtain q where "q \ 0" "p * q = 1" .. then have "degree (p * q) = degree 1" by simp with \p \ 0\ \q \ 0\ have "degree p + degree q = 0" by (simp add: degree_mult_eq) then show ?rhs by simp qed lemma is_unit_pCons_iff: "is_unit (pCons a p) \ p = 0 \ a \ 0" for p :: "'a::field poly" by (cases "p = 0") (auto simp: is_unit_triv is_unit_iff_degree) lemma is_unit_monom_trival: "is_unit p \ monom (coeff p (degree p)) 0 = p" for p :: "'a::field poly" by (cases p) (simp_all add: monom_0 is_unit_pCons_iff) lemma is_unit_const_poly_iff: "[:c:] dvd 1 \ c dvd 1" for c :: "'a::{comm_semiring_1,semiring_no_zero_divisors}" by (auto simp: one_pCons) lemma is_unit_polyE: fixes p :: "'a :: {comm_semiring_1,semiring_no_zero_divisors} poly" assumes "p dvd 1" obtains c where "p = [:c:]" "c dvd 1" proof - from assms obtain q where "1 = p * q" by (rule dvdE) then have "p \ 0" and "q \ 0" by auto from \1 = p * q\ have "degree 1 = degree (p * q)" by simp also from \p \ 0\ and \q \ 0\ have "\ = degree p + degree q" by (simp add: degree_mult_eq) finally have "degree p = 0" by simp with degree_eq_zeroE obtain c where c: "p = [:c:]" . with \p dvd 1\ have "c dvd 1" by (simp add: is_unit_const_poly_iff) with c show thesis .. qed lemma is_unit_polyE': fixes p :: "'a::field poly" assumes "is_unit p" obtains a where "p = monom a 0" and "a \ 0" proof - obtain a q where "p = pCons a q" by (cases p) with assms have "p = [:a:]" and "a \ 0" by (simp_all add: is_unit_pCons_iff) with that show thesis by (simp add: monom_0) qed lemma is_unit_poly_iff: "p dvd 1 \ (\c. p = [:c:] \ c dvd 1)" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" by (auto elim: is_unit_polyE simp add: is_unit_const_poly_iff) subsubsection \Pseudo-Division\ text \This part is by René Thiemann and Akihisa Yamada.\ fun pseudo_divmod_main :: "'a :: comm_ring_1 \ 'a poly \ 'a poly \ 'a poly \ nat \ nat \ 'a poly \ 'a poly" where "pseudo_divmod_main lc q r d dr (Suc n) = (let rr = smult lc r; qq = coeff r dr; rrr = rr - monom qq n * d; qqq = smult lc q + monom qq n in pseudo_divmod_main lc qqq rrr d (dr - 1) n)" | "pseudo_divmod_main lc q r d dr 0 = (q,r)" definition pseudo_divmod :: "'a :: comm_ring_1 poly \ 'a poly \ 'a poly \ 'a poly" where "pseudo_divmod p q \ if q = 0 then (0, p) else pseudo_divmod_main (coeff q (degree q)) 0 p q (degree p) (1 + length (coeffs p) - length (coeffs q))" lemma pseudo_divmod_main: assumes d: "d \ 0" "lc = coeff d (degree d)" and "degree r \ dr" "pseudo_divmod_main lc q r d dr n = (q',r')" and "n = 1 + dr - degree d \ dr = 0 \ n = 0 \ r = 0" shows "(r' = 0 \ degree r' < degree d) \ smult (lc^n) (d * q + r) = d * q' + r'" using assms(3-) proof (induct n arbitrary: q r dr) case 0 then show ?case by auto next case (Suc n) let ?rr = "smult lc r" let ?qq = "coeff r dr" define b where [simp]: "b = monom ?qq n" let ?rrr = "?rr - b * d" let ?qqq = "smult lc q + b" note res = Suc(3) from res[unfolded pseudo_divmod_main.simps[of lc q] Let_def] have res: "pseudo_divmod_main lc ?qqq ?rrr d (dr - 1) n = (q',r')" by (simp del: pseudo_divmod_main.simps) from Suc(4) have dr: "dr = n + degree d" by auto have "coeff (b * d) dr = coeff b n * coeff d (degree d)" proof (cases "?qq = 0") case True then show ?thesis by auto next case False then have n: "n = degree b" by (simp add: degree_monom_eq) show ?thesis unfolding n dr by (simp add: coeff_mult_degree_sum) qed also have "\ = lc * coeff b n" by (simp add: d) finally have "coeff (b * d) dr = lc * coeff b n" . moreover have "coeff ?rr dr = lc * coeff r dr" by simp ultimately have c0: "coeff ?rrr dr = 0" by auto from Suc(4) have dr: "dr = n + degree d" by auto have deg_rr: "degree ?rr \ dr" using Suc(2) degree_smult_le dual_order.trans by blast have deg_bd: "degree (b * d) \ dr" unfolding dr by (rule order.trans[OF degree_mult_le]) (auto simp: degree_monom_le) have "degree ?rrr \ dr" using degree_diff_le[OF deg_rr deg_bd] by auto with c0 have deg_rrr: "degree ?rrr \ (dr - 1)" by (rule coeff_0_degree_minus_1) have "n = 1 + (dr - 1) - degree d \ dr - 1 = 0 \ n = 0 \ ?rrr = 0" proof (cases dr) case 0 with Suc(4) have 0: "dr = 0" "n = 0" "degree d = 0" by auto with deg_rrr have "degree ?rrr = 0" by simp then have "\a. ?rrr = [:a:]" by (metis degree_pCons_eq_if old.nat.distinct(2) pCons_cases) from this obtain a where rrr: "?rrr = [:a:]" by auto show ?thesis unfolding 0 using c0 unfolding rrr 0 by simp next case _: Suc with Suc(4) show ?thesis by auto qed note IH = Suc(1)[OF deg_rrr res this] show ?case proof (intro conjI) from IH show "r' = 0 \ degree r' < degree d" by blast show "smult (lc ^ Suc n) (d * q + r) = d * q' + r'" unfolding IH[THEN conjunct2,symmetric] by (simp add: field_simps smult_add_right) qed qed lemma pseudo_divmod: assumes g: "g \ 0" and *: "pseudo_divmod f g = (q,r)" shows "smult (coeff g (degree g) ^ (Suc (degree f) - degree g)) f = g * q + r" (is ?A) and "r = 0 \ degree r < degree g" (is ?B) proof - from *[unfolded pseudo_divmod_def Let_def] have "pseudo_divmod_main (coeff g (degree g)) 0 f g (degree f) (1 + length (coeffs f) - length (coeffs g)) = (q, r)" by (auto simp: g) note main = pseudo_divmod_main[OF _ _ _ this, OF g refl le_refl] from g have "1 + length (coeffs f) - length (coeffs g) = 1 + degree f - degree g \ degree f = 0 \ 1 + length (coeffs f) - length (coeffs g) = 0 \ f = 0" by (cases "f = 0"; cases "coeffs g") (auto simp: degree_eq_length_coeffs) note main' = main[OF this] then show "r = 0 \ degree r < degree g" by auto show "smult (coeff g (degree g) ^ (Suc (degree f) - degree g)) f = g * q + r" by (subst main'[THEN conjunct2, symmetric], simp add: degree_eq_length_coeffs, cases "f = 0"; cases "coeffs g", use g in auto) qed definition "pseudo_mod_main lc r d dr n = snd (pseudo_divmod_main lc 0 r d dr n)" lemma snd_pseudo_divmod_main: "snd (pseudo_divmod_main lc q r d dr n) = snd (pseudo_divmod_main lc q' r d dr n)" by (induct n arbitrary: q q' lc r d dr) (simp_all add: Let_def) definition pseudo_mod :: "'a::{comm_ring_1,semiring_1_no_zero_divisors} poly \ 'a poly \ 'a poly" where "pseudo_mod f g = snd (pseudo_divmod f g)" lemma pseudo_mod: fixes f g :: "'a::{comm_ring_1,semiring_1_no_zero_divisors} poly" defines "r \ pseudo_mod f g" assumes g: "g \ 0" shows "\a q. a \ 0 \ smult a f = g * q + r" "r = 0 \ degree r < degree g" proof - let ?cg = "coeff g (degree g)" let ?cge = "?cg ^ (Suc (degree f) - degree g)" define a where "a = ?cge" from r_def[unfolded pseudo_mod_def] obtain q where pdm: "pseudo_divmod f g = (q, r)" by (cases "pseudo_divmod f g") auto from pseudo_divmod[OF g pdm] have id: "smult a f = g * q + r" and "r = 0 \ degree r < degree g" by (auto simp: a_def) show "r = 0 \ degree r < degree g" by fact from g have "a \ 0" by (auto simp: a_def) with id show "\a q. a \ 0 \ smult a f = g * q + r" by auto qed lemma fst_pseudo_divmod_main_as_divide_poly_main: assumes d: "d \ 0" defines lc: "lc \ coeff d (degree d)" shows "fst (pseudo_divmod_main lc q r d dr n) = divide_poly_main lc (smult (lc^n) q) (smult (lc^n) r) d dr n" proof (induct n arbitrary: q r dr) case 0 then show ?case by simp next case (Suc n) note lc0 = leading_coeff_neq_0[OF d, folded lc] then have "pseudo_divmod_main lc q r d dr (Suc n) = pseudo_divmod_main lc (smult lc q + monom (coeff r dr) n) (smult lc r - monom (coeff r dr) n * d) d (dr - 1) n" by (simp add: Let_def ac_simps) also have "fst \ = divide_poly_main lc (smult (lc^n) (smult lc q + monom (coeff r dr) n)) (smult (lc^n) (smult lc r - monom (coeff r dr) n * d)) d (dr - 1) n" by (simp only: Suc[unfolded divide_poly_main.simps Let_def]) also have "\ = divide_poly_main lc (smult (lc ^ Suc n) q) (smult (lc ^ Suc n) r) d dr (Suc n)" unfolding smult_monom smult_distribs mult_smult_left[symmetric] using lc0 by (simp add: Let_def ac_simps) finally show ?case . qed subsubsection \Division in polynomials over fields\ lemma pseudo_divmod_field: fixes g :: "'a::field poly" assumes g: "g \ 0" and *: "pseudo_divmod f g = (q,r)" defines "c \ coeff g (degree g) ^ (Suc (degree f) - degree g)" shows "f = g * smult (1/c) q + smult (1/c) r" proof - from leading_coeff_neq_0[OF g] have c0: "c \ 0" by (auto simp: c_def) from pseudo_divmod(1)[OF g *, folded c_def] have "smult c f = g * q + r" by auto also have "smult (1 / c) \ = g * smult (1 / c) q + smult (1 / c) r" by (simp add: smult_add_right) finally show ?thesis using c0 by auto qed lemma divide_poly_main_field: fixes d :: "'a::field poly" assumes d: "d \ 0" defines lc: "lc \ coeff d (degree d)" shows "divide_poly_main lc q r d dr n = fst (pseudo_divmod_main lc (smult ((1 / lc)^n) q) (smult ((1 / lc)^n) r) d dr n)" unfolding lc by (subst fst_pseudo_divmod_main_as_divide_poly_main) (auto simp: d power_one_over) lemma divide_poly_field: fixes f g :: "'a::field poly" defines "f' \ smult ((1 / coeff g (degree g)) ^ (Suc (degree f) - degree g)) f" shows "f div g = fst (pseudo_divmod f' g)" proof (cases "g = 0") case True show ?thesis unfolding divide_poly_def pseudo_divmod_def Let_def f'_def True by (simp add: divide_poly_main_0) next case False from leading_coeff_neq_0[OF False] have "degree f' = degree f" by (auto simp: f'_def) then show ?thesis using length_coeffs_degree[of f'] length_coeffs_degree[of f] unfolding divide_poly_def pseudo_divmod_def Let_def divide_poly_main_field[OF False] length_coeffs_degree[OF False] f'_def by force qed instantiation poly :: ("{semidom_divide_unit_factor,idom_divide}") normalization_semidom begin definition unit_factor_poly :: "'a poly \ 'a poly" where "unit_factor_poly p = [:unit_factor (lead_coeff p):]" definition normalize_poly :: "'a poly \ 'a poly" where "normalize p = p div [:unit_factor (lead_coeff p):]" instance proof fix p :: "'a poly" show "unit_factor p * normalize p = p" proof (cases "p = 0") case True then show ?thesis by (simp add: unit_factor_poly_def normalize_poly_def) next case False then have "lead_coeff p \ 0" by simp then have *: "unit_factor (lead_coeff p) \ 0" using unit_factor_is_unit [of "lead_coeff p"] by auto then have "unit_factor (lead_coeff p) dvd 1" by (auto intro: unit_factor_is_unit) then have **: "unit_factor (lead_coeff p) dvd c" for c by (rule dvd_trans) simp have ***: "unit_factor (lead_coeff p) * (c div unit_factor (lead_coeff p)) = c" for c proof - from ** obtain b where "c = unit_factor (lead_coeff p) * b" .. with False * show ?thesis by simp qed have "p div [:unit_factor (lead_coeff p):] = map_poly (\c. c div unit_factor (lead_coeff p)) p" by (simp add: const_poly_dvd_iff div_const_poly_conv_map_poly **) then show ?thesis by (simp add: normalize_poly_def unit_factor_poly_def smult_conv_map_poly map_poly_map_poly o_def ***) qed next fix p :: "'a poly" assume "is_unit p" then obtain c where p: "p = [:c:]" "c dvd 1" by (auto simp: is_unit_poly_iff) then show "unit_factor p = p" by (simp add: unit_factor_poly_def monom_0 is_unit_unit_factor) next fix p :: "'a poly" assume "p \ 0" then show "is_unit (unit_factor p)" by (simp add: unit_factor_poly_def monom_0 is_unit_poly_iff unit_factor_is_unit) +next + fix a b :: "'a poly" assume "is_unit a" + thus "unit_factor (a * b) = a * unit_factor b" + by (auto simp: unit_factor_poly_def lead_coeff_mult unit_factor_mult elim!: is_unit_polyE) qed (simp_all add: normalize_poly_def unit_factor_poly_def monom_0 lead_coeff_mult unit_factor_mult) end +instance poly :: ("{semidom_divide_unit_factor,idom_divide,normalization_semidom_multiplicative}") + normalization_semidom_multiplicative + by intro_classes (auto simp: unit_factor_poly_def lead_coeff_mult unit_factor_mult) + lemma normalize_poly_eq_map_poly: "normalize p = map_poly (\x. x div unit_factor (lead_coeff p)) p" proof - have "[:unit_factor (lead_coeff p):] dvd p" by (metis unit_factor_poly_def unit_factor_self) then show ?thesis by (simp add: normalize_poly_def div_const_poly_conv_map_poly) qed lemma coeff_normalize [simp]: "coeff (normalize p) n = coeff p n div unit_factor (lead_coeff p)" by (simp add: normalize_poly_eq_map_poly coeff_map_poly) class field_unit_factor = field + unit_factor + assumes unit_factor_field [simp]: "unit_factor = id" begin subclass semidom_divide_unit_factor proof fix a assume "a \ 0" then have "1 = a * inverse a" by simp then have "a dvd 1" .. then show "unit_factor a dvd 1" by simp qed simp_all end lemma unit_factor_pCons: "unit_factor (pCons a p) = (if p = 0 then [:unit_factor a:] else unit_factor p)" by (simp add: unit_factor_poly_def) lemma normalize_monom [simp]: "normalize (monom a n) = monom (normalize a) n" by (cases "a = 0") (simp_all add: map_poly_monom normalize_poly_eq_map_poly degree_monom_eq) lemma unit_factor_monom [simp]: "unit_factor (monom a n) = [:unit_factor a:]" by (cases "a = 0") (simp_all add: unit_factor_poly_def degree_monom_eq) lemma normalize_const_poly: "normalize [:c:] = [:normalize c:]" by (simp add: normalize_poly_eq_map_poly map_poly_pCons) -lemma normalize_smult: "normalize (smult c p) = smult (normalize c) (normalize p)" +lemma normalize_smult: + fixes c :: "'a :: {normalization_semidom_multiplicative, idom_divide}" + shows "normalize (smult c p) = smult (normalize c) (normalize p)" proof - have "smult c p = [:c:] * p" by simp also have "normalize \ = smult (normalize c) (normalize p)" by (subst normalize_mult) (simp add: normalize_const_poly) finally show ?thesis . qed inductive eucl_rel_poly :: "'a::field poly \ 'a poly \ 'a poly \ 'a poly \ bool" where eucl_rel_poly_by0: "eucl_rel_poly x 0 (0, x)" | eucl_rel_poly_dividesI: "y \ 0 \ x = q * y \ eucl_rel_poly x y (q, 0)" | eucl_rel_poly_remainderI: "y \ 0 \ degree r < degree y \ x = q * y + r \ eucl_rel_poly x y (q, r)" lemma eucl_rel_poly_iff: "eucl_rel_poly x y (q, r) \ x = q * y + r \ (if y = 0 then q = 0 else r = 0 \ degree r < degree y)" by (auto elim: eucl_rel_poly.cases intro: eucl_rel_poly_by0 eucl_rel_poly_dividesI eucl_rel_poly_remainderI) lemma eucl_rel_poly_0: "eucl_rel_poly 0 y (0, 0)" by (simp add: eucl_rel_poly_iff) lemma eucl_rel_poly_by_0: "eucl_rel_poly x 0 (0, x)" by (simp add: eucl_rel_poly_iff) lemma eucl_rel_poly_pCons: assumes rel: "eucl_rel_poly x y (q, r)" assumes y: "y \ 0" assumes b: "b = coeff (pCons a r) (degree y) / coeff y (degree y)" shows "eucl_rel_poly (pCons a x) y (pCons b q, pCons a r - smult b y)" (is "eucl_rel_poly ?x y (?q, ?r)") proof - from assms have x: "x = q * y + r" and r: "r = 0 \ degree r < degree y" by (simp_all add: eucl_rel_poly_iff) from b x have "?x = ?q * y + ?r" by simp moreover have "?r = 0 \ degree ?r < degree y" proof (rule eq_zero_or_degree_less) show "degree ?r \ degree y" proof (rule degree_diff_le) from r show "degree (pCons a r) \ degree y" by auto show "degree (smult b y) \ degree y" by (rule degree_smult_le) qed from \y \ 0\ show "coeff ?r (degree y) = 0" by (simp add: b) qed ultimately show ?thesis unfolding eucl_rel_poly_iff using \y \ 0\ by simp qed lemma eucl_rel_poly_exists: "\q r. eucl_rel_poly x y (q, r)" apply (cases "y = 0") apply (fast intro!: eucl_rel_poly_by_0) apply (induct x) apply (fast intro!: eucl_rel_poly_0) apply (fast intro!: eucl_rel_poly_pCons) done lemma eucl_rel_poly_unique: assumes 1: "eucl_rel_poly x y (q1, r1)" assumes 2: "eucl_rel_poly x y (q2, r2)" shows "q1 = q2 \ r1 = r2" proof (cases "y = 0") assume "y = 0" with assms show ?thesis by (simp add: eucl_rel_poly_iff) next assume [simp]: "y \ 0" from 1 have q1: "x = q1 * y + r1" and r1: "r1 = 0 \ degree r1 < degree y" unfolding eucl_rel_poly_iff by simp_all from 2 have q2: "x = q2 * y + r2" and r2: "r2 = 0 \ degree r2 < degree y" unfolding eucl_rel_poly_iff by simp_all from q1 q2 have q3: "(q1 - q2) * y = r2 - r1" by (simp add: algebra_simps) from r1 r2 have r3: "(r2 - r1) = 0 \ degree (r2 - r1) < degree y" by (auto intro: degree_diff_less) show "q1 = q2 \ r1 = r2" proof (rule classical) assume "\ ?thesis" with q3 have "q1 \ q2" and "r1 \ r2" by auto with r3 have "degree (r2 - r1) < degree y" by simp also have "degree y \ degree (q1 - q2) + degree y" by simp also from \q1 \ q2\ have "\ = degree ((q1 - q2) * y)" by (simp add: degree_mult_eq) also from q3 have "\ = degree (r2 - r1)" by simp finally have "degree (r2 - r1) < degree (r2 - r1)" . then show ?thesis by simp qed qed lemma eucl_rel_poly_0_iff: "eucl_rel_poly 0 y (q, r) \ q = 0 \ r = 0" by (auto dest: eucl_rel_poly_unique intro: eucl_rel_poly_0) lemma eucl_rel_poly_by_0_iff: "eucl_rel_poly x 0 (q, r) \ q = 0 \ r = x" by (auto dest: eucl_rel_poly_unique intro: eucl_rel_poly_by_0) lemmas eucl_rel_poly_unique_div = eucl_rel_poly_unique [THEN conjunct1] lemmas eucl_rel_poly_unique_mod = eucl_rel_poly_unique [THEN conjunct2] instantiation poly :: (field) semidom_modulo begin definition modulo_poly :: "'a poly \ 'a poly \ 'a poly" where mod_poly_def: "f mod g = (if g = 0 then f else pseudo_mod (smult ((1 / lead_coeff g) ^ (Suc (degree f) - degree g)) f) g)" instance proof fix x y :: "'a poly" show "x div y * y + x mod y = x" proof (cases "y = 0") case True then show ?thesis by (simp add: divide_poly_0 mod_poly_def) next case False then have "pseudo_divmod (smult ((1 / lead_coeff y) ^ (Suc (degree x) - degree y)) x) y = (x div y, x mod y)" by (simp add: divide_poly_field mod_poly_def pseudo_mod_def) with False pseudo_divmod [OF False this] show ?thesis by (simp add: power_mult_distrib [symmetric] ac_simps) qed qed end lemma eucl_rel_poly: "eucl_rel_poly x y (x div y, x mod y)" unfolding eucl_rel_poly_iff proof show "x = x div y * y + x mod y" by (simp add: div_mult_mod_eq) show "if y = 0 then x div y = 0 else x mod y = 0 \ degree (x mod y) < degree y" proof (cases "y = 0") case True then show ?thesis by auto next case False with pseudo_mod[OF this] show ?thesis by (simp add: mod_poly_def) qed qed lemma div_poly_eq: "eucl_rel_poly x y (q, r) \ x div y = q" for x :: "'a::field poly" by (rule eucl_rel_poly_unique_div [OF eucl_rel_poly]) lemma mod_poly_eq: "eucl_rel_poly x y (q, r) \ x mod y = r" for x :: "'a::field poly" by (rule eucl_rel_poly_unique_mod [OF eucl_rel_poly]) instance poly :: (field) idom_modulo .. lemma div_pCons_eq: "pCons a p div q = (if q = 0 then 0 else pCons (coeff (pCons a (p mod q)) (degree q) / lead_coeff q) (p div q))" using eucl_rel_poly_pCons [OF eucl_rel_poly _ refl, of q a p] by (auto intro: div_poly_eq) lemma mod_pCons_eq: "pCons a p mod q = (if q = 0 then pCons a p else pCons a (p mod q) - smult (coeff (pCons a (p mod q)) (degree q) / lead_coeff q) q)" using eucl_rel_poly_pCons [OF eucl_rel_poly _ refl, of q a p] by (auto intro: mod_poly_eq) lemma div_mod_fold_coeffs: "(p div q, p mod q) = (if q = 0 then (0, p) else fold_coeffs (\a (s, r). let b = coeff (pCons a r) (degree q) / coeff q (degree q) in (pCons b s, pCons a r - smult b q)) p (0, 0))" by (rule sym, induct p) (auto simp: div_pCons_eq mod_pCons_eq Let_def) lemma degree_mod_less: "y \ 0 \ x mod y = 0 \ degree (x mod y) < degree y" using eucl_rel_poly [of x y] unfolding eucl_rel_poly_iff by simp lemma degree_mod_less': "b \ 0 \ a mod b \ 0 \ degree (a mod b) < degree b" using degree_mod_less[of b a] by auto lemma div_poly_less: fixes x :: "'a::field poly" assumes "degree x < degree y" shows "x div y = 0" proof - from assms have "eucl_rel_poly x y (0, x)" by (simp add: eucl_rel_poly_iff) then show "x div y = 0" by (rule div_poly_eq) qed lemma mod_poly_less: assumes "degree x < degree y" shows "x mod y = x" proof - from assms have "eucl_rel_poly x y (0, x)" by (simp add: eucl_rel_poly_iff) then show "x mod y = x" by (rule mod_poly_eq) qed lemma eucl_rel_poly_smult_left: "eucl_rel_poly x y (q, r) \ eucl_rel_poly (smult a x) y (smult a q, smult a r)" by (simp add: eucl_rel_poly_iff smult_add_right) lemma div_smult_left: "(smult a x) div y = smult a (x div y)" for x y :: "'a::field poly" by (rule div_poly_eq, rule eucl_rel_poly_smult_left, rule eucl_rel_poly) lemma mod_smult_left: "(smult a x) mod y = smult a (x mod y)" by (rule mod_poly_eq, rule eucl_rel_poly_smult_left, rule eucl_rel_poly) lemma poly_div_minus_left [simp]: "(- x) div y = - (x div y)" for x y :: "'a::field poly" using div_smult_left [of "- 1::'a"] by simp lemma poly_mod_minus_left [simp]: "(- x) mod y = - (x mod y)" for x y :: "'a::field poly" using mod_smult_left [of "- 1::'a"] by simp lemma eucl_rel_poly_add_left: assumes "eucl_rel_poly x y (q, r)" assumes "eucl_rel_poly x' y (q', r')" shows "eucl_rel_poly (x + x') y (q + q', r + r')" using assms unfolding eucl_rel_poly_iff by (auto simp: algebra_simps degree_add_less) lemma poly_div_add_left: "(x + y) div z = x div z + y div z" for x y z :: "'a::field poly" using eucl_rel_poly_add_left [OF eucl_rel_poly eucl_rel_poly] by (rule div_poly_eq) lemma poly_mod_add_left: "(x + y) mod z = x mod z + y mod z" for x y z :: "'a::field poly" using eucl_rel_poly_add_left [OF eucl_rel_poly eucl_rel_poly] by (rule mod_poly_eq) lemma poly_div_diff_left: "(x - y) div z = x div z - y div z" for x y z :: "'a::field poly" by (simp only: diff_conv_add_uminus poly_div_add_left poly_div_minus_left) lemma poly_mod_diff_left: "(x - y) mod z = x mod z - y mod z" for x y z :: "'a::field poly" by (simp only: diff_conv_add_uminus poly_mod_add_left poly_mod_minus_left) lemma eucl_rel_poly_smult_right: "a \ 0 \ eucl_rel_poly x y (q, r) \ eucl_rel_poly x (smult a y) (smult (inverse a) q, r)" by (simp add: eucl_rel_poly_iff) lemma div_smult_right: "a \ 0 \ x div (smult a y) = smult (inverse a) (x div y)" for x y :: "'a::field poly" by (rule div_poly_eq, erule eucl_rel_poly_smult_right, rule eucl_rel_poly) lemma mod_smult_right: "a \ 0 \ x mod (smult a y) = x mod y" by (rule mod_poly_eq, erule eucl_rel_poly_smult_right, rule eucl_rel_poly) lemma poly_div_minus_right [simp]: "x div (- y) = - (x div y)" for x y :: "'a::field poly" using div_smult_right [of "- 1::'a"] by (simp add: nonzero_inverse_minus_eq) lemma poly_mod_minus_right [simp]: "x mod (- y) = x mod y" for x y :: "'a::field poly" using mod_smult_right [of "- 1::'a"] by simp lemma eucl_rel_poly_mult: "eucl_rel_poly x y (q, r) \ eucl_rel_poly q z (q', r') \ eucl_rel_poly x (y * z) (q', y * r' + r)" apply (cases "z = 0", simp add: eucl_rel_poly_iff) apply (cases "y = 0", simp add: eucl_rel_poly_by_0_iff eucl_rel_poly_0_iff) apply (cases "r = 0") apply (cases "r' = 0") apply (simp add: eucl_rel_poly_iff) apply (simp add: eucl_rel_poly_iff field_simps degree_mult_eq) apply (cases "r' = 0") apply (simp add: eucl_rel_poly_iff degree_mult_eq) apply (simp add: eucl_rel_poly_iff field_simps) apply (simp add: degree_mult_eq degree_add_less) done lemma poly_div_mult_right: "x div (y * z) = (x div y) div z" for x y z :: "'a::field poly" by (rule div_poly_eq, rule eucl_rel_poly_mult, (rule eucl_rel_poly)+) lemma poly_mod_mult_right: "x mod (y * z) = y * (x div y mod z) + x mod y" for x y z :: "'a::field poly" by (rule mod_poly_eq, rule eucl_rel_poly_mult, (rule eucl_rel_poly)+) lemma mod_pCons: fixes a :: "'a::field" and x y :: "'a::field poly" assumes y: "y \ 0" defines "b \ coeff (pCons a (x mod y)) (degree y) / coeff y (degree y)" shows "(pCons a x) mod y = pCons a (x mod y) - smult b y" unfolding b_def by (rule mod_poly_eq, rule eucl_rel_poly_pCons [OF eucl_rel_poly y refl]) subsubsection \List-based versions for fast implementation\ (* Subsection by: Sebastiaan Joosten René Thiemann Akihisa Yamada *) fun minus_poly_rev_list :: "'a :: group_add list \ 'a list \ 'a list" where "minus_poly_rev_list (x # xs) (y # ys) = (x - y) # (minus_poly_rev_list xs ys)" | "minus_poly_rev_list xs [] = xs" | "minus_poly_rev_list [] (y # ys) = []" fun pseudo_divmod_main_list :: "'a::comm_ring_1 \ 'a list \ 'a list \ 'a list \ nat \ 'a list \ 'a list" where "pseudo_divmod_main_list lc q r d (Suc n) = (let rr = map ((*) lc) r; a = hd r; qqq = cCons a (map ((*) lc) q); rrr = tl (if a = 0 then rr else minus_poly_rev_list rr (map ((*) a) d)) in pseudo_divmod_main_list lc qqq rrr d n)" | "pseudo_divmod_main_list lc q r d 0 = (q, r)" fun pseudo_mod_main_list :: "'a::comm_ring_1 \ 'a list \ 'a list \ nat \ 'a list" where "pseudo_mod_main_list lc r d (Suc n) = (let rr = map ((*) lc) r; a = hd r; rrr = tl (if a = 0 then rr else minus_poly_rev_list rr (map ((*) a) d)) in pseudo_mod_main_list lc rrr d n)" | "pseudo_mod_main_list lc r d 0 = r" fun divmod_poly_one_main_list :: "'a::comm_ring_1 list \ 'a list \ 'a list \ nat \ 'a list \ 'a list" where "divmod_poly_one_main_list q r d (Suc n) = (let a = hd r; qqq = cCons a q; rr = tl (if a = 0 then r else minus_poly_rev_list r (map ((*) a) d)) in divmod_poly_one_main_list qqq rr d n)" | "divmod_poly_one_main_list q r d 0 = (q, r)" fun mod_poly_one_main_list :: "'a::comm_ring_1 list \ 'a list \ nat \ 'a list" where "mod_poly_one_main_list r d (Suc n) = (let a = hd r; rr = tl (if a = 0 then r else minus_poly_rev_list r (map ((*) a) d)) in mod_poly_one_main_list rr d n)" | "mod_poly_one_main_list r d 0 = r" definition pseudo_divmod_list :: "'a::comm_ring_1 list \ 'a list \ 'a list \ 'a list" where "pseudo_divmod_list p q = (if q = [] then ([], p) else (let rq = rev q; (qu,re) = pseudo_divmod_main_list (hd rq) [] (rev p) rq (1 + length p - length q) in (qu, rev re)))" definition pseudo_mod_list :: "'a::comm_ring_1 list \ 'a list \ 'a list" where "pseudo_mod_list p q = (if q = [] then p else (let rq = rev q; re = pseudo_mod_main_list (hd rq) (rev p) rq (1 + length p - length q) in rev re))" lemma minus_zero_does_nothing: "minus_poly_rev_list x (map ((*) 0) y) = x" for x :: "'a::ring list" by (induct x y rule: minus_poly_rev_list.induct) auto lemma length_minus_poly_rev_list [simp]: "length (minus_poly_rev_list xs ys) = length xs" by (induct xs ys rule: minus_poly_rev_list.induct) auto lemma if_0_minus_poly_rev_list: "(if a = 0 then x else minus_poly_rev_list x (map ((*) a) y)) = minus_poly_rev_list x (map ((*) a) y)" for a :: "'a::ring" by(cases "a = 0") (simp_all add: minus_zero_does_nothing) lemma Poly_append: "Poly (a @ b) = Poly a + monom 1 (length a) * Poly b" for a :: "'a::comm_semiring_1 list" by (induct a) (auto simp: monom_0 monom_Suc) lemma minus_poly_rev_list: "length p \ length q \ Poly (rev (minus_poly_rev_list (rev p) (rev q))) = Poly p - monom 1 (length p - length q) * Poly q" for p q :: "'a :: comm_ring_1 list" proof (induct "rev p" "rev q" arbitrary: p q rule: minus_poly_rev_list.induct) case (1 x xs y ys) then have "length (rev q) \ length (rev p)" by simp from this[folded 1(2,3)] have ys_xs: "length ys \ length xs" by simp then have *: "Poly (rev (minus_poly_rev_list xs ys)) = Poly (rev xs) - monom 1 (length xs - length ys) * Poly (rev ys)" by (subst "1.hyps"(1)[of "rev xs" "rev ys", unfolded rev_rev_ident length_rev]) auto have "Poly p - monom 1 (length p - length q) * Poly q = Poly (rev (rev p)) - monom 1 (length (rev (rev p)) - length (rev (rev q))) * Poly (rev (rev q))" by simp also have "\ = Poly (rev (x # xs)) - monom 1 (length (x # xs) - length (y # ys)) * Poly (rev (y # ys))" unfolding 1(2,3) by simp also from ys_xs have "\ = Poly (rev xs) + monom x (length xs) - (monom 1 (length xs - length ys) * Poly (rev ys) + monom y (length xs))" by (simp add: Poly_append distrib_left mult_monom smult_monom) also have "\ = Poly (rev (minus_poly_rev_list xs ys)) + monom (x - y) (length xs)" unfolding * diff_monom[symmetric] by simp finally show ?case by (simp add: 1(2,3)[symmetric] smult_monom Poly_append) qed auto lemma smult_monom_mult: "smult a (monom b n * f) = monom (a * b) n * f" using smult_monom [of a _ n] by (metis mult_smult_left) lemma head_minus_poly_rev_list: "length d \ length r \ d \ [] \ hd (minus_poly_rev_list (map ((*) (last d)) r) (map ((*) (hd r)) (rev d))) = 0" for d r :: "'a::comm_ring list" proof (induct r) case Nil then show ?case by simp next case (Cons a rs) then show ?case by (cases "rev d") (simp_all add: ac_simps) qed lemma Poly_map: "Poly (map ((*) a) p) = smult a (Poly p)" proof (induct p) case Nil then show ?case by simp next case (Cons x xs) then show ?case by (cases "Poly xs = 0") auto qed lemma last_coeff_is_hd: "xs \ [] \ coeff (Poly xs) (length xs - 1) = hd (rev xs)" by (simp_all add: hd_conv_nth rev_nth nth_default_nth nth_append) lemma pseudo_divmod_main_list_invar: assumes leading_nonzero: "last d \ 0" and lc: "last d = lc" and "d \ []" and "pseudo_divmod_main_list lc q (rev r) (rev d) n = (q', rev r')" and "n = 1 + length r - length d" shows "pseudo_divmod_main lc (monom 1 n * Poly q) (Poly r) (Poly d) (length r - 1) n = (Poly q', Poly r')" using assms(4-) proof (induct n arbitrary: r q) case (Suc n) from Suc.prems have *: "\ Suc (length r) \ length d" by simp with \d \ []\ have "r \ []" using Suc_leI length_greater_0_conv list.size(3) by fastforce let ?a = "(hd (rev r))" let ?rr = "map ((*) lc) (rev r)" let ?rrr = "rev (tl (minus_poly_rev_list ?rr (map ((*) ?a) (rev d))))" let ?qq = "cCons ?a (map ((*) lc) q)" from * Suc(3) have n: "n = (1 + length r - length d - 1)" by simp from * have rr_val:"(length ?rrr) = (length r - 1)" by auto with \r \ []\ * have rr_smaller: "(1 + length r - length d - 1) = (1 + length ?rrr - length d)" by auto from * have id: "Suc (length r) - length d = Suc (length r - length d)" by auto from Suc.prems * have "pseudo_divmod_main_list lc ?qq (rev ?rrr) (rev d) (1 + length r - length d - 1) = (q', rev r')" by (simp add: Let_def if_0_minus_poly_rev_list id) with n have v: "pseudo_divmod_main_list lc ?qq (rev ?rrr) (rev d) n = (q', rev r')" by auto from * have sucrr:"Suc (length r) - length d = Suc (length r - length d)" using Suc_diff_le not_less_eq_eq by blast from Suc(3) \r \ []\ have n_ok : "n = 1 + (length ?rrr) - length d" by simp have cong: "\x1 x2 x3 x4 y1 y2 y3 y4. x1 = y1 \ x2 = y2 \ x3 = y3 \ x4 = y4 \ pseudo_divmod_main lc x1 x2 x3 x4 n = pseudo_divmod_main lc y1 y2 y3 y4 n" by simp have hd_rev: "coeff (Poly r) (length r - Suc 0) = hd (rev r)" using last_coeff_is_hd[OF \r \ []\] by simp show ?case unfolding Suc.hyps(1)[OF v n_ok, symmetric] pseudo_divmod_main.simps Let_def proof (rule cong[OF _ _ refl], goal_cases) case 1 show ?case by (simp add: monom_Suc hd_rev[symmetric] smult_monom Poly_map) next case 2 show ?case proof (subst Poly_on_rev_starting_with_0, goal_cases) show "hd (minus_poly_rev_list (map ((*) lc) (rev r)) (map ((*) (hd (rev r))) (rev d))) = 0" by (fold lc, subst head_minus_poly_rev_list, insert * \d \ []\, auto) from * have "length d \ length r" by simp then show "smult lc (Poly r) - monom (coeff (Poly r) (length r - 1)) n * Poly d = Poly (rev (minus_poly_rev_list (map ((*) lc) (rev r)) (map ((*) (hd (rev r))) (rev d))))" by (fold rev_map) (auto simp add: n smult_monom_mult Poly_map hd_rev [symmetric] minus_poly_rev_list) qed qed simp qed simp lemma pseudo_divmod_impl [code]: "pseudo_divmod f g = map_prod poly_of_list poly_of_list (pseudo_divmod_list (coeffs f) (coeffs g))" for f g :: "'a::comm_ring_1 poly" proof (cases "g = 0") case False then have "last (coeffs g) \ 0" and "last (coeffs g) = lead_coeff g" and "coeffs g \ []" by (simp_all add: last_coeffs_eq_coeff_degree) moreover obtain q r where qr: "pseudo_divmod_main_list (last (coeffs g)) (rev []) (rev (coeffs f)) (rev (coeffs g)) (1 + length (coeffs f) - length (coeffs g)) = (q, rev (rev r))" by force ultimately have "(Poly q, Poly (rev r)) = pseudo_divmod_main (lead_coeff g) 0 f g (length (coeffs f) - Suc 0) (Suc (length (coeffs f)) - length (coeffs g))" by (subst pseudo_divmod_main_list_invar [symmetric]) auto moreover have "pseudo_divmod_main_list (hd (rev (coeffs g))) [] (rev (coeffs f)) (rev (coeffs g)) (1 + length (coeffs f) - length (coeffs g)) = (q, r)" using qr hd_rev [OF \coeffs g \ []\] by simp ultimately show ?thesis by (auto simp: degree_eq_length_coeffs pseudo_divmod_def pseudo_divmod_list_def Let_def) next case True then show ?thesis by (auto simp add: pseudo_divmod_def pseudo_divmod_list_def) qed lemma pseudo_mod_main_list: "snd (pseudo_divmod_main_list l q xs ys n) = pseudo_mod_main_list l xs ys n" by (induct n arbitrary: l q xs ys) (auto simp: Let_def) lemma pseudo_mod_impl[code]: "pseudo_mod f g = poly_of_list (pseudo_mod_list (coeffs f) (coeffs g))" proof - have snd_case: "\f g p. snd ((\(x,y). (f x, g y)) p) = g (snd p)" by auto show ?thesis unfolding pseudo_mod_def pseudo_divmod_impl pseudo_divmod_list_def pseudo_mod_list_def Let_def by (simp add: snd_case pseudo_mod_main_list) qed subsubsection \Improved Code-Equations for Polynomial (Pseudo) Division\ lemma pdivmod_pdivmodrel: "eucl_rel_poly p q (r, s) \ (p div q, p mod q) = (r, s)" by (metis eucl_rel_poly eucl_rel_poly_unique) lemma pdivmod_via_pseudo_divmod: "(f div g, f mod g) = (if g = 0 then (0, f) else let ilc = inverse (coeff g (degree g)); h = smult ilc g; (q,r) = pseudo_divmod f h in (smult ilc q, r))" (is "?l = ?r") proof (cases "g = 0") case True then show ?thesis by simp next case False define lc where "lc = inverse (coeff g (degree g))" define h where "h = smult lc g" from False have h1: "coeff h (degree h) = 1" and lc: "lc \ 0" by (auto simp: h_def lc_def) then have h0: "h \ 0" by auto obtain q r where p: "pseudo_divmod f h = (q, r)" by force from False have id: "?r = (smult lc q, r)" by (auto simp: Let_def h_def[symmetric] lc_def[symmetric] p) from pseudo_divmod[OF h0 p, unfolded h1] have f: "f = h * q + r" and r: "r = 0 \ degree r < degree h" by auto from f r h0 have "eucl_rel_poly f h (q, r)" by (auto simp: eucl_rel_poly_iff) then have "(f div h, f mod h) = (q, r)" by (simp add: pdivmod_pdivmodrel) with lc have "(f div g, f mod g) = (smult lc q, r)" by (auto simp: h_def div_smult_right[OF lc] mod_smult_right[OF lc]) with id show ?thesis by auto qed lemma pdivmod_via_pseudo_divmod_list: "(f div g, f mod g) = (let cg = coeffs g in if cg = [] then (0, f) else let cf = coeffs f; ilc = inverse (last cg); ch = map ((*) ilc) cg; (q, r) = pseudo_divmod_main_list 1 [] (rev cf) (rev ch) (1 + length cf - length cg) in (poly_of_list (map ((*) ilc) q), poly_of_list (rev r)))" proof - note d = pdivmod_via_pseudo_divmod pseudo_divmod_impl pseudo_divmod_list_def show ?thesis proof (cases "g = 0") case True with d show ?thesis by auto next case False define ilc where "ilc = inverse (coeff g (degree g))" from False have ilc: "ilc \ 0" by (auto simp: ilc_def) with False have id: "g = 0 \ False" "coeffs g = [] \ False" "last (coeffs g) = coeff g (degree g)" "coeffs (smult ilc g) = [] \ False" by (auto simp: last_coeffs_eq_coeff_degree) have id2: "hd (rev (coeffs (smult ilc g))) = 1" by (subst hd_rev, insert id ilc, auto simp: coeffs_smult, subst last_map, auto simp: id ilc_def) have id3: "length (coeffs (smult ilc g)) = length (coeffs g)" "rev (coeffs (smult ilc g)) = rev (map ((*) ilc) (coeffs g))" unfolding coeffs_smult using ilc by auto obtain q r where pair: "pseudo_divmod_main_list 1 [] (rev (coeffs f)) (rev (map ((*) ilc) (coeffs g))) (1 + length (coeffs f) - length (coeffs g)) = (q, r)" by force show ?thesis unfolding d Let_def id if_False ilc_def[symmetric] map_prod_def[symmetric] id2 unfolding id3 pair map_prod_def split by (auto simp: Poly_map) qed qed lemma pseudo_divmod_main_list_1: "pseudo_divmod_main_list 1 = divmod_poly_one_main_list" proof (intro ext, goal_cases) case (1 q r d n) have *: "map ((*) 1) xs = xs" for xs :: "'a list" by (induct xs) auto show ?case by (induct n arbitrary: q r d) (auto simp: * Let_def) qed fun divide_poly_main_list :: "'a::idom_divide \ 'a list \ 'a list \ 'a list \ nat \ 'a list" where "divide_poly_main_list lc q r d (Suc n) = (let cr = hd r in if cr = 0 then divide_poly_main_list lc (cCons cr q) (tl r) d n else let a = cr div lc; qq = cCons a q; rr = minus_poly_rev_list r (map ((*) a) d) in if hd rr = 0 then divide_poly_main_list lc qq (tl rr) d n else [])" | "divide_poly_main_list lc q r d 0 = q" lemma divide_poly_main_list_simp [simp]: "divide_poly_main_list lc q r d (Suc n) = (let cr = hd r; a = cr div lc; qq = cCons a q; rr = minus_poly_rev_list r (map ((*) a) d) in if hd rr = 0 then divide_poly_main_list lc qq (tl rr) d n else [])" by (simp add: Let_def minus_zero_does_nothing) declare divide_poly_main_list.simps(1)[simp del] definition divide_poly_list :: "'a::idom_divide poly \ 'a poly \ 'a poly" where "divide_poly_list f g = (let cg = coeffs g in if cg = [] then g else let cf = coeffs f; cgr = rev cg in poly_of_list (divide_poly_main_list (hd cgr) [] (rev cf) cgr (1 + length cf - length cg)))" lemmas pdivmod_via_divmod_list = pdivmod_via_pseudo_divmod_list[unfolded pseudo_divmod_main_list_1] lemma mod_poly_one_main_list: "snd (divmod_poly_one_main_list q r d n) = mod_poly_one_main_list r d n" by (induct n arbitrary: q r d) (auto simp: Let_def) lemma mod_poly_code [code]: "f mod g = (let cg = coeffs g in if cg = [] then f else let cf = coeffs f; ilc = inverse (last cg); ch = map ((*) ilc) cg; r = mod_poly_one_main_list (rev cf) (rev ch) (1 + length cf - length cg) in poly_of_list (rev r))" (is "_ = ?rhs") proof - have "snd (f div g, f mod g) = ?rhs" unfolding pdivmod_via_divmod_list Let_def mod_poly_one_main_list [symmetric, of _ _ _ Nil] by (auto split: prod.splits) then show ?thesis by simp qed definition div_field_poly_impl :: "'a :: field poly \ 'a poly \ 'a poly" where "div_field_poly_impl f g = (let cg = coeffs g in if cg = [] then 0 else let cf = coeffs f; ilc = inverse (last cg); ch = map ((*) ilc) cg; q = fst (divmod_poly_one_main_list [] (rev cf) (rev ch) (1 + length cf - length cg)) in poly_of_list ((map ((*) ilc) q)))" text \We do not declare the following lemma as code equation, since then polynomial division on non-fields will no longer be executable. However, a code-unfold is possible, since \div_field_poly_impl\ is a bit more efficient than the generic polynomial division.\ lemma div_field_poly_impl[code_unfold]: "(div) = div_field_poly_impl" proof (intro ext) fix f g :: "'a poly" have "fst (f div g, f mod g) = div_field_poly_impl f g" unfolding div_field_poly_impl_def pdivmod_via_divmod_list Let_def by (auto split: prod.splits) then show "f div g = div_field_poly_impl f g" by simp qed lemma divide_poly_main_list: assumes lc0: "lc \ 0" and lc: "last d = lc" and d: "d \ []" and "n = (1 + length r - length d)" shows "Poly (divide_poly_main_list lc q (rev r) (rev d) n) = divide_poly_main lc (monom 1 n * Poly q) (Poly r) (Poly d) (length r - 1) n" using assms(4-) proof (induct "n" arbitrary: r q) case (Suc n) from Suc.prems have ifCond: "\ Suc (length r) \ length d" by simp with d have r: "r \ []" using Suc_leI length_greater_0_conv list.size(3) by fastforce then obtain rr lcr where r: "r = rr @ [lcr]" by (cases r rule: rev_cases) auto from d lc obtain dd where d: "d = dd @ [lc]" by (cases d rule: rev_cases) auto from Suc(2) ifCond have n: "n = 1 + length rr - length d" by (auto simp: r) from ifCond have len: "length dd \ length rr" by (simp add: r d) show ?case proof (cases "lcr div lc * lc = lcr") case False with r d show ?thesis unfolding Suc(2)[symmetric] by (auto simp add: Let_def nth_default_append) next case True with r d have id: "?thesis \ Poly (divide_poly_main_list lc (cCons (lcr div lc) q) (rev (rev (minus_poly_rev_list (rev rr) (rev (map ((*) (lcr div lc)) dd))))) (rev d) n) = divide_poly_main lc (monom 1 (Suc n) * Poly q + monom (lcr div lc) n) (Poly r - monom (lcr div lc) n * Poly d) (Poly d) (length rr - 1) n" by (cases r rule: rev_cases; cases "d" rule: rev_cases) (auto simp add: Let_def rev_map nth_default_append) have cong: "\x1 x2 x3 x4 y1 y2 y3 y4. x1 = y1 \ x2 = y2 \ x3 = y3 \ x4 = y4 \ divide_poly_main lc x1 x2 x3 x4 n = divide_poly_main lc y1 y2 y3 y4 n" by simp show ?thesis unfolding id proof (subst Suc(1), simp add: n, subst minus_poly_rev_list, force simp: len, rule cong[OF _ _ refl], goal_cases) case 2 have "monom lcr (length rr) = monom (lcr div lc) (length rr - length dd) * monom lc (length dd)" by (simp add: mult_monom len True) then show ?case unfolding r d Poly_append n ring_distribs by (auto simp: Poly_map smult_monom smult_monom_mult) qed (auto simp: len monom_Suc smult_monom) qed qed simp lemma divide_poly_list[code]: "f div g = divide_poly_list f g" proof - note d = divide_poly_def divide_poly_list_def show ?thesis proof (cases "g = 0") case True show ?thesis by (auto simp: d True) next case False then obtain cg lcg where cg: "coeffs g = cg @ [lcg]" by (cases "coeffs g" rule: rev_cases) auto with False have id: "(g = 0) = False" "(cg @ [lcg] = []) = False" by auto from cg False have lcg: "coeff g (degree g) = lcg" using last_coeffs_eq_coeff_degree last_snoc by force with False have "lcg \ 0" by auto from cg Poly_coeffs [of g] have ltp: "Poly (cg @ [lcg]) = g" by auto show ?thesis unfolding d cg Let_def id if_False poly_of_list_def by (subst divide_poly_main_list, insert False cg \lcg \ 0\) (auto simp: lcg ltp, simp add: degree_eq_length_coeffs) qed qed subsection \Primality and irreducibility in polynomial rings\ lemma prod_mset_const_poly: "(\x\#A. [:f x:]) = [:prod_mset (image_mset f A):]" by (induct A) (simp_all add: ac_simps) lemma irreducible_const_poly_iff: fixes c :: "'a :: {comm_semiring_1,semiring_no_zero_divisors}" shows "irreducible [:c:] \ irreducible c" proof assume A: "irreducible c" show "irreducible [:c:]" proof (rule irreducibleI) fix a b assume ab: "[:c:] = a * b" hence "degree [:c:] = degree (a * b)" by (simp only: ) also from A ab have "a \ 0" "b \ 0" by auto hence "degree (a * b) = degree a + degree b" by (simp add: degree_mult_eq) finally have "degree a = 0" "degree b = 0" by auto then obtain a' b' where ab': "a = [:a':]" "b = [:b':]" by (auto elim!: degree_eq_zeroE) from ab have "coeff [:c:] 0 = coeff (a * b) 0" by (simp only: ) hence "c = a' * b'" by (simp add: ab' mult_ac) from A and this have "a' dvd 1 \ b' dvd 1" by (rule irreducibleD) with ab' show "a dvd 1 \ b dvd 1" by (auto simp add: is_unit_const_poly_iff) qed (insert A, auto simp: irreducible_def is_unit_poly_iff) next assume A: "irreducible [:c:]" then have "c \ 0" and "\ c dvd 1" by (auto simp add: irreducible_def is_unit_const_poly_iff) then show "irreducible c" proof (rule irreducibleI) fix a b assume ab: "c = a * b" hence "[:c:] = [:a:] * [:b:]" by (simp add: mult_ac) from A and this have "[:a:] dvd 1 \ [:b:] dvd 1" by (rule irreducibleD) then show "a dvd 1 \ b dvd 1" by (auto simp add: is_unit_const_poly_iff) qed qed lemma lift_prime_elem_poly: assumes "prime_elem (c :: 'a :: semidom)" shows "prime_elem [:c:]" proof (rule prime_elemI) fix a b assume *: "[:c:] dvd a * b" from * have dvd: "c dvd coeff (a * b) n" for n by (subst (asm) const_poly_dvd_iff) blast { define m where "m = (GREATEST m. \c dvd coeff b m)" assume "\[:c:] dvd b" hence A: "\i. \c dvd coeff b i" by (subst (asm) const_poly_dvd_iff) blast have B: "\i. \c dvd coeff b i \ i \ degree b" by (auto intro: le_degree) have coeff_m: "\c dvd coeff b m" unfolding m_def by (rule GreatestI_ex_nat[OF A B]) have "i \ m" if "\c dvd coeff b i" for i unfolding m_def by (rule Greatest_le_nat[OF that B]) hence dvd_b: "c dvd coeff b i" if "i > m" for i using that by force have "c dvd coeff a i" for i proof (induction i rule: nat_descend_induct[of "degree a"]) case (base i) thus ?case by (simp add: coeff_eq_0) next case (descend i) let ?A = "{..i+m} - {i}" have "c dvd coeff (a * b) (i + m)" by (rule dvd) also have "coeff (a * b) (i + m) = (\k\i + m. coeff a k * coeff b (i + m - k))" by (simp add: coeff_mult) also have "{..i+m} = insert i ?A" by auto also have "(\k\\. coeff a k * coeff b (i + m - k)) = coeff a i * coeff b m + (\k\?A. coeff a k * coeff b (i + m - k))" (is "_ = _ + ?S") by (subst sum.insert) simp_all finally have eq: "c dvd coeff a i * coeff b m + ?S" . moreover have "c dvd ?S" proof (rule dvd_sum) fix k assume k: "k \ {..i+m} - {i}" show "c dvd coeff a k * coeff b (i + m - k)" proof (cases "k < i") case False with k have "c dvd coeff a k" by (intro descend.IH) simp thus ?thesis by simp next case True hence "c dvd coeff b (i + m - k)" by (intro dvd_b) simp thus ?thesis by simp qed qed ultimately have "c dvd coeff a i * coeff b m" by (simp add: dvd_add_left_iff) with assms coeff_m show "c dvd coeff a i" by (simp add: prime_elem_dvd_mult_iff) qed hence "[:c:] dvd a" by (subst const_poly_dvd_iff) blast } then show "[:c:] dvd a \ [:c:] dvd b" by blast next from assms show "[:c:] \ 0" and "\ [:c:] dvd 1" by (simp_all add: prime_elem_def is_unit_const_poly_iff) qed lemma prime_elem_const_poly_iff: fixes c :: "'a :: semidom" shows "prime_elem [:c:] \ prime_elem c" proof assume A: "prime_elem [:c:]" show "prime_elem c" proof (rule prime_elemI) fix a b assume "c dvd a * b" hence "[:c:] dvd [:a:] * [:b:]" by (simp add: mult_ac) from A and this have "[:c:] dvd [:a:] \ [:c:] dvd [:b:]" by (rule prime_elem_dvd_multD) thus "c dvd a \ c dvd b" by simp qed (insert A, auto simp: prime_elem_def is_unit_poly_iff) qed (auto intro: lift_prime_elem_poly) subsection \Content and primitive part of a polynomial\ definition content :: "'a::semiring_gcd poly \ 'a" where "content p = gcd_list (coeffs p)" lemma content_eq_fold_coeffs [code]: "content p = fold_coeffs gcd p 0" by (simp add: content_def Gcd_fin.set_eq_fold fold_coeffs_def foldr_fold fun_eq_iff ac_simps) lemma content_0 [simp]: "content 0 = 0" by (simp add: content_def) lemma content_1 [simp]: "content 1 = 1" by (simp add: content_def) lemma content_const [simp]: "content [:c:] = normalize c" by (simp add: content_def cCons_def) lemma const_poly_dvd_iff_dvd_content: "[:c:] dvd p \ c dvd content p" for c :: "'a::semiring_gcd" proof (cases "p = 0") case True then show ?thesis by simp next case False have "[:c:] dvd p \ (\n. c dvd coeff p n)" by (rule const_poly_dvd_iff) also have "\ \ (\a\set (coeffs p). c dvd a)" proof safe fix n :: nat assume "\a\set (coeffs p). c dvd a" then show "c dvd coeff p n" by (cases "n \ degree p") (auto simp: coeff_eq_0 coeffs_def split: if_splits) qed (auto simp: coeffs_def simp del: upt_Suc split: if_splits) also have "\ \ c dvd content p" by (simp add: content_def dvd_Gcd_fin_iff dvd_mult_unit_iff) finally show ?thesis . qed lemma content_dvd [simp]: "[:content p:] dvd p" by (subst const_poly_dvd_iff_dvd_content) simp_all lemma content_dvd_coeff [simp]: "content p dvd coeff p n" proof (cases "p = 0") case True then show ?thesis by simp next case False then show ?thesis by (cases "n \ degree p") (auto simp add: content_def not_le coeff_eq_0 coeff_in_coeffs intro: Gcd_fin_dvd) qed lemma content_dvd_coeffs: "c \ set (coeffs p) \ content p dvd c" by (simp add: content_def Gcd_fin_dvd) lemma normalize_content [simp]: "normalize (content p) = content p" by (simp add: content_def) lemma is_unit_content_iff [simp]: "is_unit (content p) \ content p = 1" proof assume "is_unit (content p)" then have "normalize (content p) = 1" by (simp add: is_unit_normalize del: normalize_content) then show "content p = 1" by simp qed auto -lemma content_smult [simp]: "content (smult c p) = normalize c * content p" - by (simp add: content_def coeffs_smult Gcd_fin_mult) +lemma content_smult [simp]: + fixes c :: "'a :: {normalization_semidom_multiplicative, semiring_gcd}" + shows "content (smult c p) = normalize c * content p" + by (simp add: content_def coeffs_smult Gcd_fin_mult normalize_mult) lemma content_eq_zero_iff [simp]: "content p = 0 \ p = 0" by (auto simp: content_def simp: poly_eq_iff coeffs_def) definition primitive_part :: "'a :: semiring_gcd poly \ 'a poly" where "primitive_part p = map_poly (\x. x div content p) p" lemma primitive_part_0 [simp]: "primitive_part 0 = 0" by (simp add: primitive_part_def) lemma content_times_primitive_part [simp]: "smult (content p) (primitive_part p) = p" for p :: "'a :: semiring_gcd poly" proof (cases "p = 0") case True then show ?thesis by simp next case False then show ?thesis unfolding primitive_part_def by (auto simp: smult_conv_map_poly map_poly_map_poly o_def content_dvd_coeffs intro: map_poly_idI) qed lemma primitive_part_eq_0_iff [simp]: "primitive_part p = 0 \ p = 0" proof (cases "p = 0") case True then show ?thesis by simp next case False then have "primitive_part p = map_poly (\x. x div content p) p" by (simp add: primitive_part_def) also from False have "\ = 0 \ p = 0" by (intro map_poly_eq_0_iff) (auto simp: dvd_div_eq_0_iff content_dvd_coeffs) finally show ?thesis using False by simp qed lemma content_primitive_part [simp]: + fixes p :: "'a :: {normalization_semidom_multiplicative, semiring_gcd} poly" assumes "p \ 0" shows "content (primitive_part p) = 1" proof - have "p = smult (content p) (primitive_part p)" by simp also have "content \ = content (primitive_part p) * content p" by (simp del: content_times_primitive_part add: ac_simps) finally have "1 * content p = content (primitive_part p) * content p" by simp then have "1 * content p div content p = content (primitive_part p) * content p div content p" by simp with assms show ?thesis by simp qed lemma content_decompose: - obtains p' :: "'a::semiring_gcd poly" + obtains p' :: "'a :: {normalization_semidom_multiplicative, semiring_gcd} poly" where "p = smult (content p) p'" "content p' = 1" proof (cases "p = 0") case True then have "p = smult (content p) 1" "content 1 = 1" by simp_all then show ?thesis .. next case False then have "p = smult (content p) (primitive_part p)" "content (primitive_part p) = 1" by simp_all then show ?thesis .. qed lemma content_dvd_contentI [intro]: "p dvd q \ content p dvd content q" using const_poly_dvd_iff_dvd_content content_dvd dvd_trans by blast lemma primitive_part_const_poly [simp]: "primitive_part [:x:] = [:unit_factor x:]" by (simp add: primitive_part_def map_poly_pCons) lemma primitive_part_prim: "content p = 1 \ primitive_part p = p" by (auto simp: primitive_part_def) lemma degree_primitive_part [simp]: "degree (primitive_part p) = degree p" proof (cases "p = 0") case True then show ?thesis by simp next case False have "p = smult (content p) (primitive_part p)" by simp also from False have "degree \ = degree (primitive_part p)" by (subst degree_smult_eq) simp_all finally show ?thesis .. qed lemma smult_content_normalize_primitive_part [simp]: - "smult (content p) (normalize (primitive_part p)) = normalize p" + fixes p :: "'a :: {normalization_semidom_multiplicative, semiring_gcd, idom_divide} poly" + shows "smult (content p) (normalize (primitive_part p)) = normalize p" proof - have "smult (content p) (normalize (primitive_part p)) = normalize ([:content p:] * primitive_part p)" by (subst normalize_mult) (simp_all add: normalize_const_poly) also have "[:content p:] * primitive_part p = p" by simp finally show ?thesis . qed context begin private lemma content_1_mult: fixes f g :: "'a :: {semiring_gcd, factorial_semiring} poly" assumes "content f = 1" "content g = 1" shows "content (f * g) = 1" proof (cases "f * g = 0") case False from assms have "f \ 0" "g \ 0" by auto hence "f * g \ 0" by auto { assume "\is_unit (content (f * g))" with False have "\p. p dvd content (f * g) \ prime p" by (intro prime_divisor_exists) simp_all then obtain p where "p dvd content (f * g)" "prime p" by blast from \p dvd content (f * g)\ have "[:p:] dvd f * g" by (simp add: const_poly_dvd_iff_dvd_content) moreover from \prime p\ have "prime_elem [:p:]" by (simp add: lift_prime_elem_poly) ultimately have "[:p:] dvd f \ [:p:] dvd g" by (simp add: prime_elem_dvd_mult_iff) with assms have "is_unit p" by (simp add: const_poly_dvd_iff_dvd_content) with \prime p\ have False by simp } hence "is_unit (content (f * g))" by blast hence "normalize (content (f * g)) = 1" by (simp add: is_unit_normalize del: normalize_content) thus ?thesis by simp qed (insert assms, auto) lemma content_mult: - fixes p q :: "'a :: {factorial_semiring, semiring_gcd} poly" + fixes p q :: "'a :: {factorial_semiring, semiring_gcd, normalization_semidom_multiplicative} poly" shows "content (p * q) = content p * content q" proof (cases "p * q = 0") case False then have "p \ 0" and "q \ 0" by simp_all then have *: "content (primitive_part p * primitive_part q) = 1" by (auto intro: content_1_mult) have "p * q = smult (content p) (primitive_part p) * smult (content q) (primitive_part q)" by simp also have "\ = smult (content p * content q) (primitive_part p * primitive_part q)" by (metis mult.commute mult_smult_right smult_smult) with * show ?thesis by (simp add: normalize_mult) next case True then show ?thesis by auto qed end lemma primitive_part_mult: - fixes p q :: "'a :: {factorial_semiring, semiring_Gcd, ring_gcd, idom_divide} poly" + fixes p q :: "'a :: {factorial_semiring, semiring_Gcd, ring_gcd, idom_divide, + normalization_semidom_multiplicative} poly" shows "primitive_part (p * q) = primitive_part p * primitive_part q" proof - have "primitive_part (p * q) = p * q div [:content (p * q):]" by (simp add: primitive_part_def div_const_poly_conv_map_poly) also have "\ = (p div [:content p:]) * (q div [:content q:])" by (subst div_mult_div_if_dvd) (simp_all add: content_mult mult_ac) also have "\ = primitive_part p * primitive_part q" by (simp add: primitive_part_def div_const_poly_conv_map_poly) finally show ?thesis . qed lemma primitive_part_smult: - fixes p :: "'a :: {factorial_semiring, semiring_Gcd, ring_gcd, idom_divide} poly" + fixes p :: "'a :: {factorial_semiring, semiring_Gcd, ring_gcd, idom_divide, + normalization_semidom_multiplicative} poly" shows "primitive_part (smult a p) = smult (unit_factor a) (primitive_part p)" proof - have "smult a p = [:a:] * p" by simp also have "primitive_part \ = smult (unit_factor a) (primitive_part p)" by (subst primitive_part_mult) simp_all finally show ?thesis . qed lemma primitive_part_dvd_primitive_partI [intro]: - fixes p q :: "'a :: {factorial_semiring, semiring_Gcd, ring_gcd, idom_divide} poly" + fixes p q :: "'a :: {factorial_semiring, semiring_Gcd, ring_gcd, idom_divide, + normalization_semidom_multiplicative} poly" shows "p dvd q \ primitive_part p dvd primitive_part q" by (auto elim!: dvdE simp: primitive_part_mult) lemma content_prod_mset: - fixes A :: "'a :: {factorial_semiring, semiring_Gcd} poly multiset" + fixes A :: "'a :: {factorial_semiring, semiring_Gcd, normalization_semidom_multiplicative} + poly multiset" shows "content (prod_mset A) = prod_mset (image_mset content A)" by (induction A) (simp_all add: content_mult mult_ac) lemma content_prod_eq_1_iff: - fixes p q :: "'a :: {factorial_semiring, semiring_Gcd} poly" + fixes p q :: "'a :: {factorial_semiring, semiring_Gcd, normalization_semidom_multiplicative} poly" shows "content (p * q) = 1 \ content p = 1 \ content q = 1" proof safe assume A: "content (p * q) = 1" { fix p q :: "'a poly" assume "content p * content q = 1" hence "1 = content p * content q" by simp hence "content p dvd 1" by (rule dvdI) hence "content p = 1" by simp } note B = this from A B[of p q] B [of q p] show "content p = 1" "content q = 1" by (simp_all add: content_mult mult_ac) qed (auto simp: content_mult) no_notation cCons (infixr "##" 65) end diff --git a/src/HOL/Computational_Algebra/Polynomial_Factorial.thy b/src/HOL/Computational_Algebra/Polynomial_Factorial.thy --- a/src/HOL/Computational_Algebra/Polynomial_Factorial.thy +++ b/src/HOL/Computational_Algebra/Polynomial_Factorial.thy @@ -1,848 +1,869 @@ (* Title: HOL/Computational_Algebra/Polynomial_Factorial.thy Author: Manuel Eberl *) section \Polynomials, fractions and rings\ theory Polynomial_Factorial imports Complex_Main Polynomial Normalized_Fraction begin subsection \Lifting elements into the field of fractions\ definition to_fract :: "'a :: idom \ 'a fract" where "to_fract x = Fract x 1" \ \FIXME: more idiomatic name, abbreviation\ lemma to_fract_0 [simp]: "to_fract 0 = 0" by (simp add: to_fract_def eq_fract Zero_fract_def) lemma to_fract_1 [simp]: "to_fract 1 = 1" by (simp add: to_fract_def eq_fract One_fract_def) lemma to_fract_add [simp]: "to_fract (x + y) = to_fract x + to_fract y" by (simp add: to_fract_def) lemma to_fract_diff [simp]: "to_fract (x - y) = to_fract x - to_fract y" by (simp add: to_fract_def) lemma to_fract_uminus [simp]: "to_fract (-x) = -to_fract x" by (simp add: to_fract_def) lemma to_fract_mult [simp]: "to_fract (x * y) = to_fract x * to_fract y" by (simp add: to_fract_def) lemma to_fract_eq_iff [simp]: "to_fract x = to_fract y \ x = y" by (simp add: to_fract_def eq_fract) lemma to_fract_eq_0_iff [simp]: "to_fract x = 0 \ x = 0" by (simp add: to_fract_def Zero_fract_def eq_fract) lemma snd_quot_of_fract_nonzero [simp]: "snd (quot_of_fract x) \ 0" by transfer simp lemma Fract_quot_of_fract [simp]: "Fract (fst (quot_of_fract x)) (snd (quot_of_fract x)) = x" by transfer (simp del: fractrel_iff, subst fractrel_normalize_quot_left, simp) lemma to_fract_quot_of_fract: assumes "snd (quot_of_fract x) = 1" shows "to_fract (fst (quot_of_fract x)) = x" proof - have "x = Fract (fst (quot_of_fract x)) (snd (quot_of_fract x))" by simp also note assms finally show ?thesis by (simp add: to_fract_def) qed lemma snd_quot_of_fract_Fract_whole: assumes "y dvd x" shows "snd (quot_of_fract (Fract x y)) = 1" using assms by transfer (auto simp: normalize_quot_def Let_def gcd_proj2_if_dvd) lemma Fract_conv_to_fract: "Fract a b = to_fract a / to_fract b" by (simp add: to_fract_def) lemma quot_of_fract_to_fract [simp]: "quot_of_fract (to_fract x) = (x, 1)" unfolding to_fract_def by transfer (simp add: normalize_quot_def) lemma fst_quot_of_fract_eq_0_iff [simp]: "fst (quot_of_fract x) = 0 \ x = 0" by transfer simp lemma snd_quot_of_fract_to_fract [simp]: "snd (quot_of_fract (to_fract x)) = 1" unfolding to_fract_def by (rule snd_quot_of_fract_Fract_whole) simp_all lemma coprime_quot_of_fract: "coprime (fst (quot_of_fract x)) (snd (quot_of_fract x))" by transfer (simp add: coprime_normalize_quot) lemma unit_factor_snd_quot_of_fract: "unit_factor (snd (quot_of_fract x)) = 1" using quot_of_fract_in_normalized_fracts[of x] by (simp add: normalized_fracts_def case_prod_unfold) lemma unit_factor_1_imp_normalized: "unit_factor x = 1 \ normalize x = x" by (subst (2) normalize_mult_unit_factor [symmetric, of x]) (simp del: normalize_mult_unit_factor) lemma normalize_snd_quot_of_fract: "normalize (snd (quot_of_fract x)) = snd (quot_of_fract x)" by (intro unit_factor_1_imp_normalized unit_factor_snd_quot_of_fract) subsection \Lifting polynomial coefficients to the field of fractions\ abbreviation (input) fract_poly where "fract_poly \ map_poly to_fract" abbreviation (input) unfract_poly where "unfract_poly \ map_poly (fst \ quot_of_fract)" lemma fract_poly_smult [simp]: "fract_poly (smult c p) = smult (to_fract c) (fract_poly p)" by (simp add: smult_conv_map_poly map_poly_map_poly o_def) lemma fract_poly_0 [simp]: "fract_poly 0 = 0" by (simp add: poly_eqI coeff_map_poly) lemma fract_poly_1 [simp]: "fract_poly 1 = 1" by (simp add: map_poly_pCons) lemma fract_poly_add [simp]: "fract_poly (p + q) = fract_poly p + fract_poly q" by (intro poly_eqI) (simp_all add: coeff_map_poly) lemma fract_poly_diff [simp]: "fract_poly (p - q) = fract_poly p - fract_poly q" by (intro poly_eqI) (simp_all add: coeff_map_poly) lemma to_fract_sum [simp]: "to_fract (sum f A) = sum (\x. to_fract (f x)) A" by (cases "finite A", induction A rule: finite_induct) simp_all lemma fract_poly_mult [simp]: "fract_poly (p * q) = fract_poly p * fract_poly q" by (intro poly_eqI) (simp_all add: coeff_map_poly coeff_mult) lemma fract_poly_eq_iff [simp]: "fract_poly p = fract_poly q \ p = q" by (auto simp: poly_eq_iff coeff_map_poly) lemma fract_poly_eq_0_iff [simp]: "fract_poly p = 0 \ p = 0" using fract_poly_eq_iff[of p 0] by (simp del: fract_poly_eq_iff) lemma fract_poly_dvd: "p dvd q \ fract_poly p dvd fract_poly q" by (auto elim!: dvdE) lemma prod_mset_fract_poly: "(\x\#A. map_poly to_fract (f x)) = fract_poly (prod_mset (image_mset f A))" by (induct A) (simp_all add: ac_simps) lemma is_unit_fract_poly_iff: "p dvd 1 \ fract_poly p dvd 1 \ content p = 1" proof safe assume A: "p dvd 1" with fract_poly_dvd [of p 1] show "is_unit (fract_poly p)" by simp from A show "content p = 1" by (auto simp: is_unit_poly_iff normalize_1_iff) next assume A: "fract_poly p dvd 1" and B: "content p = 1" from A obtain c where c: "fract_poly p = [:c:]" by (auto simp: is_unit_poly_iff) { fix n :: nat assume "n > 0" have "to_fract (coeff p n) = coeff (fract_poly p) n" by (simp add: coeff_map_poly) also note c also from \n > 0\ have "coeff [:c:] n = 0" by (simp add: coeff_pCons split: nat.splits) finally have "coeff p n = 0" by simp } hence "degree p \ 0" by (intro degree_le) simp_all with B show "p dvd 1" by (auto simp: is_unit_poly_iff normalize_1_iff elim!: degree_eq_zeroE) qed lemma fract_poly_is_unit: "p dvd 1 \ fract_poly p dvd 1" using fract_poly_dvd[of p 1] by simp lemma fract_poly_smult_eqE: - fixes c :: "'a :: {idom_divide,ring_gcd} fract" + fixes c :: "'a :: {idom_divide,ring_gcd,semiring_gcd_mult_normalize} fract" assumes "fract_poly p = smult c (fract_poly q)" obtains a b where "c = to_fract b / to_fract a" "smult a p = smult b q" "coprime a b" "normalize a = a" proof - define a b where "a = fst (quot_of_fract c)" and "b = snd (quot_of_fract c)" have "smult (to_fract a) (fract_poly q) = smult (to_fract b) (fract_poly p)" by (subst smult_eq_iff) (simp_all add: a_def b_def Fract_conv_to_fract [symmetric] assms) hence "fract_poly (smult a q) = fract_poly (smult b p)" by (simp del: fract_poly_eq_iff) hence "smult b p = smult a q" by (simp only: fract_poly_eq_iff) moreover have "c = to_fract a / to_fract b" "coprime b a" "normalize b = b" by (simp_all add: a_def b_def coprime_quot_of_fract [of c] ac_simps normalize_snd_quot_of_fract Fract_conv_to_fract [symmetric]) ultimately show ?thesis by (intro that[of a b]) qed subsection \Fractional content\ abbreviation (input) Lcm_coeff_denoms - :: "'a :: {semiring_Gcd,idom_divide,ring_gcd} fract poly \ 'a" + :: "'a :: {semiring_Gcd,idom_divide,ring_gcd,semiring_gcd_mult_normalize} fract poly \ 'a" where "Lcm_coeff_denoms p \ Lcm (snd ` quot_of_fract ` set (coeffs p))" definition fract_content :: - "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} fract poly \ 'a fract" where + "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} fract poly \ 'a fract" where "fract_content p = (let d = Lcm_coeff_denoms p in Fract (content (unfract_poly (smult (to_fract d) p))) d)" definition primitive_part_fract :: - "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} fract poly \ 'a poly" where + "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} fract poly \ 'a poly" where "primitive_part_fract p = primitive_part (unfract_poly (smult (to_fract (Lcm_coeff_denoms p)) p))" lemma primitive_part_fract_0 [simp]: "primitive_part_fract 0 = 0" by (simp add: primitive_part_fract_def) lemma fract_content_eq_0_iff [simp]: "fract_content p = 0 \ p = 0" unfolding fract_content_def Let_def Zero_fract_def by (subst eq_fract) (auto simp: Lcm_0_iff map_poly_eq_0_iff) -lemma content_primitive_part_fract [simp]: "p \ 0 \ content (primitive_part_fract p) = 1" +lemma content_primitive_part_fract [simp]: + fixes p :: "'a :: {semiring_gcd_mult_normalize, + factorial_semiring, ring_gcd, semiring_Gcd,idom_divide} fract poly" + shows "p \ 0 \ content (primitive_part_fract p) = 1" unfolding primitive_part_fract_def by (rule content_primitive_part) (auto simp: primitive_part_fract_def map_poly_eq_0_iff Lcm_0_iff) lemma content_times_primitive_part_fract: "smult (fract_content p) (fract_poly (primitive_part_fract p)) = p" proof - define p' where "p' = unfract_poly (smult (to_fract (Lcm_coeff_denoms p)) p)" have "fract_poly p' = map_poly (to_fract \ fst \ quot_of_fract) (smult (to_fract (Lcm_coeff_denoms p)) p)" unfolding primitive_part_fract_def p'_def by (subst map_poly_map_poly) (simp_all add: o_assoc) also have "\ = smult (to_fract (Lcm_coeff_denoms p)) p" proof (intro map_poly_idI, unfold o_apply) fix c assume "c \ set (coeffs (smult (to_fract (Lcm_coeff_denoms p)) p))" then obtain c' where c: "c' \ set (coeffs p)" "c = to_fract (Lcm_coeff_denoms p) * c'" by (auto simp add: Lcm_0_iff coeffs_smult split: if_splits) note c(2) also have "c' = Fract (fst (quot_of_fract c')) (snd (quot_of_fract c'))" by simp also have "to_fract (Lcm_coeff_denoms p) * \ = Fract (Lcm_coeff_denoms p * fst (quot_of_fract c')) (snd (quot_of_fract c'))" unfolding to_fract_def by (subst mult_fract) simp_all also have "snd (quot_of_fract \) = 1" by (intro snd_quot_of_fract_Fract_whole dvd_mult2 dvd_Lcm) (insert c(1), auto) finally show "to_fract (fst (quot_of_fract c)) = c" by (rule to_fract_quot_of_fract) qed also have "p' = smult (content p') (primitive_part p')" by (rule content_times_primitive_part [symmetric]) also have "primitive_part p' = primitive_part_fract p" by (simp add: primitive_part_fract_def p'_def) also have "fract_poly (smult (content p') (primitive_part_fract p)) = smult (to_fract (content p')) (fract_poly (primitive_part_fract p))" by simp finally have "smult (to_fract (content p')) (fract_poly (primitive_part_fract p)) = smult (to_fract (Lcm_coeff_denoms p)) p" . thus ?thesis by (subst (asm) smult_eq_iff) (auto simp add: Let_def p'_def Fract_conv_to_fract field_simps Lcm_0_iff fract_content_def) qed lemma fract_content_fract_poly [simp]: "fract_content (fract_poly p) = to_fract (content p)" proof - have "Lcm_coeff_denoms (fract_poly p) = 1" by (auto simp: set_coeffs_map_poly) hence "fract_content (fract_poly p) = to_fract (content (map_poly (fst \ quot_of_fract \ to_fract) p))" by (simp add: fract_content_def to_fract_def fract_collapse map_poly_map_poly del: Lcm_1_iff) also have "map_poly (fst \ quot_of_fract \ to_fract) p = p" by (intro map_poly_idI) simp_all finally show ?thesis . qed lemma content_decompose_fract: - fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} fract poly" + fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide, + semiring_gcd_mult_normalize} fract poly" obtains c p' where "p = smult c (map_poly to_fract p')" "content p' = 1" proof (cases "p = 0") case True hence "p = smult 0 (map_poly to_fract 1)" "content 1 = 1" by simp_all thus ?thesis .. next case False thus ?thesis by (rule that[OF content_times_primitive_part_fract [symmetric] content_primitive_part_fract]) qed subsection \More properties of content and primitive part\ lemma lift_prime_elem_poly: assumes "prime_elem (c :: 'a :: semidom)" shows "prime_elem [:c:]" proof (rule prime_elemI) fix a b assume *: "[:c:] dvd a * b" from * have dvd: "c dvd coeff (a * b) n" for n by (subst (asm) const_poly_dvd_iff) blast { define m where "m = (GREATEST m. \c dvd coeff b m)" assume "\[:c:] dvd b" hence A: "\i. \c dvd coeff b i" by (subst (asm) const_poly_dvd_iff) blast have B: "\i. \c dvd coeff b i \ i \ degree b" by (auto intro: le_degree) have coeff_m: "\c dvd coeff b m" unfolding m_def by (rule GreatestI_ex_nat[OF A B]) have "i \ m" if "\c dvd coeff b i" for i unfolding m_def by (rule Greatest_le_nat[OF that B]) hence dvd_b: "c dvd coeff b i" if "i > m" for i using that by force have "c dvd coeff a i" for i proof (induction i rule: nat_descend_induct[of "degree a"]) case (base i) thus ?case by (simp add: coeff_eq_0) next case (descend i) let ?A = "{..i+m} - {i}" have "c dvd coeff (a * b) (i + m)" by (rule dvd) also have "coeff (a * b) (i + m) = (\k\i + m. coeff a k * coeff b (i + m - k))" by (simp add: coeff_mult) also have "{..i+m} = insert i ?A" by auto also have "(\k\\. coeff a k * coeff b (i + m - k)) = coeff a i * coeff b m + (\k\?A. coeff a k * coeff b (i + m - k))" (is "_ = _ + ?S") by (subst sum.insert) simp_all finally have eq: "c dvd coeff a i * coeff b m + ?S" . moreover have "c dvd ?S" proof (rule dvd_sum) fix k assume k: "k \ {..i+m} - {i}" show "c dvd coeff a k * coeff b (i + m - k)" proof (cases "k < i") case False with k have "c dvd coeff a k" by (intro descend.IH) simp thus ?thesis by simp next case True hence "c dvd coeff b (i + m - k)" by (intro dvd_b) simp thus ?thesis by simp qed qed ultimately have "c dvd coeff a i * coeff b m" by (simp add: dvd_add_left_iff) with assms coeff_m show "c dvd coeff a i" by (simp add: prime_elem_dvd_mult_iff) qed hence "[:c:] dvd a" by (subst const_poly_dvd_iff) blast } then show "[:c:] dvd a \ [:c:] dvd b" by blast next from assms show "[:c:] \ 0" and "\ [:c:] dvd 1" by (simp_all add: prime_elem_def is_unit_const_poly_iff) qed lemma prime_elem_const_poly_iff: fixes c :: "'a :: semidom" shows "prime_elem [:c:] \ prime_elem c" proof assume A: "prime_elem [:c:]" show "prime_elem c" proof (rule prime_elemI) fix a b assume "c dvd a * b" hence "[:c:] dvd [:a:] * [:b:]" by (simp add: mult_ac) from A and this have "[:c:] dvd [:a:] \ [:c:] dvd [:b:]" by (rule prime_elem_dvd_multD) thus "c dvd a \ c dvd b" by simp qed (insert A, auto simp: prime_elem_def is_unit_poly_iff) qed (auto intro: lift_prime_elem_poly) lemma fract_poly_dvdD: - fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} poly" + fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide, + semiring_gcd_mult_normalize} poly" assumes "fract_poly p dvd fract_poly q" "content p = 1" shows "p dvd q" proof - from assms(1) obtain r where r: "fract_poly q = fract_poly p * r" by (erule dvdE) from content_decompose_fract[of r] guess c r' . note r' = this from r r' have eq: "fract_poly q = smult c (fract_poly (p * r'))" by simp from fract_poly_smult_eqE[OF this] guess a b . note ab = this have "content (smult a q) = content (smult b (p * r'))" by (simp only: ab(2)) hence eq': "normalize b = a * content q" by (simp add: assms content_mult r' ab(4)) have "1 = gcd a (normalize b)" by (simp add: ab) also note eq' also have "gcd a (a * content q) = a" by (simp add: gcd_proj1_if_dvd ab(4)) finally have [simp]: "a = 1" by simp from eq ab have "q = p * ([:b:] * r')" by simp thus ?thesis by (rule dvdI) qed subsection \Polynomials over a field are a Euclidean ring\ context begin interpretation field_poly: - normalization_euclidean_semiring where zero = "0 :: 'a :: field poly" + normalization_euclidean_semiring_multiplicative where zero = "0 :: 'a :: field poly" and one = 1 and plus = plus and minus = minus and times = times and normalize = "\p. smult (inverse (lead_coeff p)) p" and unit_factor = "\p. [:lead_coeff p:]" and euclidean_size = "\p. if p = 0 then 0 else 2 ^ degree p" and divide = divide and modulo = modulo rewrites "dvd.dvd (times :: 'a poly \ _) = Rings.dvd" and "comm_monoid_mult.prod_mset times 1 = prod_mset" and "comm_semiring_1.irreducible times 1 0 = irreducible" and "comm_semiring_1.prime_elem times 1 0 = prime_elem" proof - show "dvd.dvd (times :: 'a poly \ _) = Rings.dvd" by (simp add: dvd_dict) show "comm_monoid_mult.prod_mset times 1 = prod_mset" by (simp add: prod_mset_dict) show "comm_semiring_1.irreducible times 1 0 = irreducible" by (simp add: irreducible_dict) show "comm_semiring_1.prime_elem times 1 0 = prime_elem" by (simp add: prime_elem_dict) - show "class.normalization_euclidean_semiring divide plus minus (0 :: 'a poly) times 1 + show "class.normalization_euclidean_semiring_multiplicative divide plus minus (0 :: 'a poly) times 1 modulo (\p. if p = 0 then 0 else 2 ^ degree p) (\p. [:lead_coeff p:]) (\p. smult (inverse (lead_coeff p)) p)" proof (standard, fold dvd_dict) fix p :: "'a poly" show "[:lead_coeff p:] * smult (inverse (lead_coeff p)) p = p" by (cases "p = 0") simp_all next fix p :: "'a poly" assume "is_unit p" then show "[:lead_coeff p:] = p" by (elim is_unit_polyE) (auto simp: monom_0 one_poly_def field_simps) next fix p :: "'a poly" assume "p \ 0" then show "is_unit [:lead_coeff p:]" by (simp add: is_unit_pCons_iff) + next + fix a b :: "'a poly" assume "is_unit a" + thus "[:lead_coeff (a * b):] = a * [:lead_coeff b:]" + by (auto elim!: is_unit_polyE) qed (auto simp: lead_coeff_mult Rings.div_mult_mod_eq intro!: degree_mod_less' degree_mult_right_le) qed lemma field_poly_irreducible_imp_prime: "prime_elem p" if "irreducible p" for p :: "'a :: field poly" using that by (fact field_poly.irreducible_imp_prime_elem) - +find_theorems name:prod_mset_prime_factorization lemma field_poly_prod_mset_prime_factorization: "prod_mset (field_poly.prime_factorization p) = smult (inverse (lead_coeff p)) p" if "p \ 0" for p :: "'a :: field poly" using that by (fact field_poly.prod_mset_prime_factorization) lemma field_poly_in_prime_factorization_imp_prime: "prime_elem p" if "p \# field_poly.prime_factorization x" for p :: "'a :: field poly" by (rule field_poly.prime_imp_prime_elem, rule field_poly.in_prime_factors_imp_prime) (fact that) subsection \Primality and irreducibility in polynomial rings\ lemma nonconst_poly_irreducible_iff: - fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} poly" + fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} poly" assumes "degree p \ 0" shows "irreducible p \ irreducible (fract_poly p) \ content p = 1" proof safe assume p: "irreducible p" from content_decompose[of p] guess p' . note p' = this hence "p = [:content p:] * p'" by simp from p this have "[:content p:] dvd 1 \ p' dvd 1" by (rule irreducibleD) moreover have "\p' dvd 1" proof assume "p' dvd 1" hence "degree p = 0" by (subst p') (auto simp: is_unit_poly_iff) with assms show False by contradiction qed ultimately show [simp]: "content p = 1" by (simp add: is_unit_const_poly_iff) show "irreducible (map_poly to_fract p)" proof (rule irreducibleI) have "fract_poly p = 0 \ p = 0" by (intro map_poly_eq_0_iff) auto with assms show "map_poly to_fract p \ 0" by auto next show "\is_unit (fract_poly p)" proof assume "is_unit (map_poly to_fract p)" hence "degree (map_poly to_fract p) = 0" by (auto simp: is_unit_poly_iff) hence "degree p = 0" by (simp add: degree_map_poly) with assms show False by contradiction qed next fix q r assume qr: "fract_poly p = q * r" from content_decompose_fract[of q] guess cg q' . note q = this from content_decompose_fract[of r] guess cr r' . note r = this from qr q r p have nz: "cg \ 0" "cr \ 0" by auto from qr have eq: "fract_poly p = smult (cr * cg) (fract_poly (q' * r'))" by (simp add: q r) from fract_poly_smult_eqE[OF this] guess a b . note ab = this hence "content (smult a p) = content (smult b (q' * r'))" by (simp only:) with ab(4) have a: "a = normalize b" by (simp add: content_mult q r) then have "normalize b = gcd a b" by simp with \coprime a b\ have "normalize b = 1" by simp then have "a = 1" "is_unit b" by (simp_all add: a normalize_1_iff) note eq also from ab(1) \a = 1\ have "cr * cg = to_fract b" by simp also have "smult \ (fract_poly (q' * r')) = fract_poly (smult b (q' * r'))" by simp finally have "p = ([:b:] * q') * r'" by (simp del: fract_poly_smult) from p and this have "([:b:] * q') dvd 1 \ r' dvd 1" by (rule irreducibleD) hence "q' dvd 1 \ r' dvd 1" by (auto dest: dvd_mult_right simp del: mult_pCons_left) hence "fract_poly q' dvd 1 \ fract_poly r' dvd 1" by (auto simp: fract_poly_is_unit) with q r show "is_unit q \ is_unit r" by (auto simp add: is_unit_smult_iff dvd_field_iff nz) qed next assume irred: "irreducible (fract_poly p)" and primitive: "content p = 1" show "irreducible p" proof (rule irreducibleI) from irred show "p \ 0" by auto next from irred show "\p dvd 1" by (auto simp: irreducible_def dest: fract_poly_is_unit) next fix q r assume qr: "p = q * r" hence "fract_poly p = fract_poly q * fract_poly r" by simp from irred and this have "fract_poly q dvd 1 \ fract_poly r dvd 1" by (rule irreducibleD) with primitive qr show "q dvd 1 \ r dvd 1" by (auto simp: content_prod_eq_1_iff is_unit_fract_poly_iff) qed qed private lemma irreducible_imp_prime_poly: - fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} poly" + fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} poly" assumes "irreducible p" shows "prime_elem p" proof (cases "degree p = 0") case True with assms show ?thesis by (auto simp: prime_elem_const_poly_iff irreducible_const_poly_iff intro!: irreducible_imp_prime_elem elim!: degree_eq_zeroE) next case False from assms False have irred: "irreducible (fract_poly p)" and primitive: "content p = 1" by (simp_all add: nonconst_poly_irreducible_iff) from irred have prime: "prime_elem (fract_poly p)" by (rule field_poly_irreducible_imp_prime) show ?thesis proof (rule prime_elemI) fix q r assume "p dvd q * r" hence "fract_poly p dvd fract_poly (q * r)" by (rule fract_poly_dvd) hence "fract_poly p dvd fract_poly q * fract_poly r" by simp from prime and this have "fract_poly p dvd fract_poly q \ fract_poly p dvd fract_poly r" by (rule prime_elem_dvd_multD) with primitive show "p dvd q \ p dvd r" by (auto dest: fract_poly_dvdD) qed (insert assms, auto simp: irreducible_def) qed lemma degree_primitive_part_fract [simp]: "degree (primitive_part_fract p) = degree p" proof - have "p = smult (fract_content p) (fract_poly (primitive_part_fract p))" by (simp add: content_times_primitive_part_fract) also have "degree \ = degree (primitive_part_fract p)" by (auto simp: degree_map_poly) finally show ?thesis .. qed lemma irreducible_primitive_part_fract: - fixes p :: "'a :: {idom_divide, ring_gcd, factorial_semiring, semiring_Gcd} fract poly" + fixes p :: "'a :: {idom_divide, ring_gcd, factorial_semiring, semiring_Gcd,semiring_gcd_mult_normalize} fract poly" assumes "irreducible p" shows "irreducible (primitive_part_fract p)" proof - from assms have deg: "degree (primitive_part_fract p) \ 0" by (intro notI) (auto elim!: degree_eq_zeroE simp: irreducible_def is_unit_poly_iff dvd_field_iff) hence [simp]: "p \ 0" by auto note \irreducible p\ also have "p = [:fract_content p:] * fract_poly (primitive_part_fract p)" by (simp add: content_times_primitive_part_fract) also have "irreducible \ \ irreducible (fract_poly (primitive_part_fract p))" by (intro irreducible_mult_unit_left) (simp_all add: is_unit_poly_iff dvd_field_iff) finally show ?thesis using deg by (simp add: nonconst_poly_irreducible_iff) qed lemma prime_elem_primitive_part_fract: - fixes p :: "'a :: {idom_divide, ring_gcd, factorial_semiring, semiring_Gcd} fract poly" + fixes p :: "'a :: {idom_divide, ring_gcd, factorial_semiring, semiring_Gcd,semiring_gcd_mult_normalize} fract poly" shows "irreducible p \ prime_elem (primitive_part_fract p)" by (intro irreducible_imp_prime_poly irreducible_primitive_part_fract) lemma irreducible_linear_field_poly: fixes a b :: "'a::field" assumes "b \ 0" shows "irreducible [:a,b:]" proof (rule irreducibleI) fix p q assume pq: "[:a,b:] = p * q" also from pq assms have "degree \ = degree p + degree q" by (intro degree_mult_eq) auto finally have "degree p = 0 \ degree q = 0" using assms by auto with assms pq show "is_unit p \ is_unit q" by (auto simp: is_unit_const_poly_iff dvd_field_iff elim!: degree_eq_zeroE) qed (insert assms, auto simp: is_unit_poly_iff) lemma prime_elem_linear_field_poly: "(b :: 'a :: field) \ 0 \ prime_elem [:a,b:]" by (rule field_poly_irreducible_imp_prime, rule irreducible_linear_field_poly) lemma irreducible_linear_poly: - fixes a b :: "'a::{idom_divide,ring_gcd,factorial_semiring,semiring_Gcd}" + fixes a b :: "'a::{idom_divide,ring_gcd,factorial_semiring,semiring_Gcd,semiring_gcd_mult_normalize}" shows "b \ 0 \ coprime a b \ irreducible [:a,b:]" by (auto intro!: irreducible_linear_field_poly simp: nonconst_poly_irreducible_iff content_def map_poly_pCons) lemma prime_elem_linear_poly: - fixes a b :: "'a::{idom_divide,ring_gcd,factorial_semiring,semiring_Gcd}" + fixes a b :: "'a::{idom_divide,ring_gcd,factorial_semiring,semiring_Gcd,semiring_gcd_mult_normalize}" shows "b \ 0 \ coprime a b \ prime_elem [:a,b:]" by (rule irreducible_imp_prime_poly, rule irreducible_linear_poly) subsection \Prime factorisation of polynomials\ private lemma poly_prime_factorization_exists_content_1: - fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} poly" + fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} poly" assumes "p \ 0" "content p = 1" shows "\A. (\p. p \# A \ prime_elem p) \ prod_mset A = normalize p" proof - let ?P = "field_poly.prime_factorization (fract_poly p)" define c where "c = prod_mset (image_mset fract_content ?P)" define c' where "c' = c * to_fract (lead_coeff p)" define e where "e = prod_mset (image_mset primitive_part_fract ?P)" define A where "A = image_mset (normalize \ primitive_part_fract) ?P" have "content e = (\x\#field_poly.prime_factorization (map_poly to_fract p). content (primitive_part_fract x))" by (simp add: e_def content_prod_mset multiset.map_comp o_def) also have "image_mset (\x. content (primitive_part_fract x)) ?P = image_mset (\_. 1) ?P" by (intro image_mset_cong content_primitive_part_fract) auto finally have content_e: "content e = 1" by simp from \p \ 0\ have "fract_poly p = [:lead_coeff (fract_poly p):] * smult (inverse (lead_coeff (fract_poly p))) (fract_poly p)" by simp also have "[:lead_coeff (fract_poly p):] = [:to_fract (lead_coeff p):]" by (simp add: monom_0 degree_map_poly coeff_map_poly) also from assms have "smult (inverse (lead_coeff (fract_poly p))) (fract_poly p) = prod_mset ?P" by (subst field_poly_prod_mset_prime_factorization) simp_all also have "\ = prod_mset (image_mset id ?P)" by simp also have "image_mset id ?P = image_mset (\x. [:fract_content x:] * fract_poly (primitive_part_fract x)) ?P" by (intro image_mset_cong) (auto simp: content_times_primitive_part_fract) also have "prod_mset \ = smult c (fract_poly e)" by (subst prod_mset.distrib) (simp_all add: prod_mset_fract_poly prod_mset_const_poly c_def e_def) also have "[:to_fract (lead_coeff p):] * \ = smult c' (fract_poly e)" by (simp add: c'_def) finally have eq: "fract_poly p = smult c' (fract_poly e)" . also obtain b where b: "c' = to_fract b" "is_unit b" proof - from fract_poly_smult_eqE[OF eq] guess a b . note ab = this from ab(2) have "content (smult a p) = content (smult b e)" by (simp only: ) with assms content_e have "a = normalize b" by (simp add: ab(4)) with ab have ab': "a = 1" "is_unit b" by (simp_all add: normalize_1_iff) with ab ab' have "c' = to_fract b" by auto from this and \is_unit b\ show ?thesis by (rule that) qed hence "smult c' (fract_poly e) = fract_poly (smult b e)" by simp finally have "p = smult b e" by (simp only: fract_poly_eq_iff) hence "p = [:b:] * e" by simp with b have "normalize p = normalize e" by (simp only: normalize_mult) (simp add: is_unit_normalize is_unit_poly_iff) also have "normalize e = prod_mset A" by (simp add: multiset.map_comp e_def A_def normalize_prod_mset) finally have "prod_mset A = normalize p" .. have "prime_elem p" if "p \# A" for p using that by (auto simp: A_def prime_elem_primitive_part_fract prime_elem_imp_irreducible dest!: field_poly_in_prime_factorization_imp_prime ) from this and \prod_mset A = normalize p\ show ?thesis by (intro exI[of _ A]) blast qed lemma poly_prime_factorization_exists: - fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} poly" + fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} poly" assumes "p \ 0" - shows "\A. (\p. p \# A \ prime_elem p) \ prod_mset A = normalize p" + shows "\A. (\p. p \# A \ prime_elem p) \ normalize (prod_mset A) = normalize p" proof - define B where "B = image_mset (\x. [:x:]) (prime_factorization (content p))" have "\A. (\p. p \# A \ prime_elem p) \ prod_mset A = normalize (primitive_part p)" by (rule poly_prime_factorization_exists_content_1) (insert assms, simp_all) then guess A by (elim exE conjE) note A = this - moreover from assms have "prod_mset B = [:content p:]" - by (simp add: B_def prod_mset_const_poly prod_mset_prime_factorization) + have "normalize (prod_mset (A + B)) = normalize (prod_mset A * normalize (prod_mset B))" + by simp + also from assms have "normalize (prod_mset B) = normalize [:content p:]" + by (simp add: prod_mset_const_poly normalize_const_poly prod_mset_prime_factorization_weak B_def) + also have "prod_mset A = normalize (primitive_part p)" + using A by simp + finally have "normalize (prod_mset (A + B)) = normalize (primitive_part p * [:content p:])" + by simp moreover have "\p. p \# B \ prime_elem p" by (auto simp: B_def intro!: lift_prime_elem_poly dest: in_prime_factors_imp_prime) - ultimately show ?thesis by (intro exI[of _ "B + A"]) auto + ultimately show ?thesis using A by (intro exI[of _ "A + B"]) (auto) qed end subsection \Typeclass instances\ -instance poly :: (factorial_ring_gcd) factorial_semiring +instance poly :: ("{factorial_ring_gcd,semiring_gcd_mult_normalize}") factorial_semiring by standard (rule poly_prime_factorization_exists) -instantiation poly :: (factorial_ring_gcd) factorial_ring_gcd +instantiation poly :: ("{factorial_ring_gcd, semiring_gcd_mult_normalize}") factorial_ring_gcd begin definition gcd_poly :: "'a poly \ 'a poly \ 'a poly" where [code del]: "gcd_poly = gcd_factorial" definition lcm_poly :: "'a poly \ 'a poly \ 'a poly" where [code del]: "lcm_poly = lcm_factorial" definition Gcd_poly :: "'a poly set \ 'a poly" where [code del]: "Gcd_poly = Gcd_factorial" definition Lcm_poly :: "'a poly set \ 'a poly" where [code del]: "Lcm_poly = Lcm_factorial" instance by standard (simp_all add: gcd_poly_def lcm_poly_def Gcd_poly_def Lcm_poly_def) end -instantiation poly :: ("{field,factorial_ring_gcd}") "{unique_euclidean_ring, normalization_euclidean_semiring}" +instance poly :: ("{factorial_ring_gcd, semiring_gcd_mult_normalize}") semiring_gcd_mult_normalize .. + +instantiation poly :: ("{field,factorial_ring_gcd,semiring_gcd_mult_normalize}") + "{unique_euclidean_ring, normalization_euclidean_semiring}" begin definition euclidean_size_poly :: "'a poly \ nat" where "euclidean_size_poly p = (if p = 0 then 0 else 2 ^ degree p)" definition division_segment_poly :: "'a poly \ 'a poly" where [simp]: "division_segment_poly p = 1" instance proof show "(q * p + r) div p = q" if "p \ 0" and "euclidean_size r < euclidean_size p" for q p r :: "'a poly" proof - from \p \ 0\ eucl_rel_poly [of r p] have "eucl_rel_poly (r + q * p) p (q + r div p, r mod p)" by (simp add: eucl_rel_poly_iff distrib_right) then have "(r + q * p) div p = q + r div p" by (rule div_poly_eq) with that show ?thesis by (cases "r = 0") (simp_all add: euclidean_size_poly_def div_poly_less ac_simps) qed qed (auto simp: euclidean_size_poly_def Rings.div_mult_mod_eq div_poly_less degree_mult_eq power_add intro!: degree_mod_less' split: if_splits) end -instance poly :: ("{field, normalization_euclidean_semiring, factorial_ring_gcd}") euclidean_ring_gcd +instance poly :: ("{field, normalization_euclidean_semiring, factorial_ring_gcd, + semiring_gcd_mult_normalize}") euclidean_ring_gcd by (rule euclidean_ring_gcd_class.intro, rule factorial_euclidean_semiring_gcdI) standard subsection \Polynomial GCD\ lemma gcd_poly_decompose: - fixes p q :: "'a :: factorial_ring_gcd poly" + fixes p q :: "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly" shows "gcd p q = smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q))" proof (rule sym, rule gcdI) have "[:gcd (content p) (content q):] * gcd (primitive_part p) (primitive_part q) dvd [:content p:] * primitive_part p" by (intro mult_dvd_mono) simp_all thus "smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q)) dvd p" by simp next have "[:gcd (content p) (content q):] * gcd (primitive_part p) (primitive_part q) dvd [:content q:] * primitive_part q" by (intro mult_dvd_mono) simp_all thus "smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q)) dvd q" by simp next fix d assume "d dvd p" "d dvd q" hence "[:content d:] * primitive_part d dvd [:gcd (content p) (content q):] * gcd (primitive_part p) (primitive_part q)" by (intro mult_dvd_mono) auto thus "d dvd smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q))" by simp qed (auto simp: normalize_smult) lemma gcd_poly_pseudo_mod: - fixes p q :: "'a :: factorial_ring_gcd poly" + fixes p q :: "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly" assumes nz: "q \ 0" and prim: "content p = 1" "content q = 1" shows "gcd p q = gcd q (primitive_part (pseudo_mod p q))" proof - define r s where "r = fst (pseudo_divmod p q)" and "s = snd (pseudo_divmod p q)" define a where "a = [:coeff q (degree q) ^ (Suc (degree p) - degree q):]" have [simp]: "primitive_part a = unit_factor a" by (simp add: a_def unit_factor_poly_def unit_factor_power monom_0) from nz have [simp]: "a \ 0" by (auto simp: a_def) have rs: "pseudo_divmod p q = (r, s)" by (simp add: r_def s_def) have "gcd (q * r + s) q = gcd q s" using gcd_add_mult[of q r s] by (simp add: gcd.commute add_ac mult_ac) with pseudo_divmod(1)[OF nz rs] have "gcd (p * a) q = gcd q s" by (simp add: a_def) also from prim have "gcd (p * a) q = gcd p q" by (subst gcd_poly_decompose) (auto simp: primitive_part_mult gcd_mult_unit1 primitive_part_prim simp del: mult_pCons_right ) also from prim have "gcd q s = gcd q (primitive_part s)" by (subst gcd_poly_decompose) (simp_all add: primitive_part_prim) also have "s = pseudo_mod p q" by (simp add: s_def pseudo_mod_def) finally show ?thesis . qed lemma degree_pseudo_mod_less: assumes "q \ 0" "pseudo_mod p q \ 0" shows "degree (pseudo_mod p q) < degree q" using pseudo_mod(2)[of q p] assms by auto function gcd_poly_code_aux :: "'a :: factorial_ring_gcd poly \ 'a poly \ 'a poly" where "gcd_poly_code_aux p q = (if q = 0 then normalize p else gcd_poly_code_aux q (primitive_part (pseudo_mod p q)))" by auto termination by (relation "measure ((\p. if p = 0 then 0 else Suc (degree p)) \ snd)") (auto simp: degree_pseudo_mod_less) declare gcd_poly_code_aux.simps [simp del] lemma gcd_poly_code_aux_correct: assumes "content p = 1" "q = 0 \ content q = 1" shows "gcd_poly_code_aux p q = gcd p q" using assms proof (induction p q rule: gcd_poly_code_aux.induct) case (1 p q) show ?case proof (cases "q = 0") case True thus ?thesis by (subst gcd_poly_code_aux.simps) auto next case False hence "gcd_poly_code_aux p q = gcd_poly_code_aux q (primitive_part (pseudo_mod p q))" by (subst gcd_poly_code_aux.simps) simp_all also from "1.prems" False have "primitive_part (pseudo_mod p q) = 0 \ content (primitive_part (pseudo_mod p q)) = 1" by (cases "pseudo_mod p q = 0") auto with "1.prems" False have "gcd_poly_code_aux q (primitive_part (pseudo_mod p q)) = gcd q (primitive_part (pseudo_mod p q))" by (intro 1) simp_all also from "1.prems" False have "\ = gcd p q" by (intro gcd_poly_pseudo_mod [symmetric]) auto finally show ?thesis . qed qed definition gcd_poly_code :: "'a :: factorial_ring_gcd poly \ 'a poly \ 'a poly" where "gcd_poly_code p q = (if p = 0 then normalize q else if q = 0 then normalize p else smult (gcd (content p) (content q)) (gcd_poly_code_aux (primitive_part p) (primitive_part q)))" lemma gcd_poly_code [code]: "gcd p q = gcd_poly_code p q" by (simp add: gcd_poly_code_def gcd_poly_code_aux_correct gcd_poly_decompose [symmetric]) lemma lcm_poly_code [code]: - fixes p q :: "'a :: factorial_ring_gcd poly" - shows "lcm p q = normalize (p * q) div gcd p q" + fixes p q :: "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly" + shows "lcm p q = normalize (p * q div gcd p q)" by (fact lcm_gcd) -lemmas Gcd_poly_set_eq_fold [code] = Gcd_set_eq_fold [where ?'a = "'a :: factorial_ring_gcd poly"] -lemmas Lcm_poly_set_eq_fold [code] = Lcm_set_eq_fold [where ?'a = "'a :: factorial_ring_gcd poly"] +lemmas Gcd_poly_set_eq_fold [code] = + Gcd_set_eq_fold [where ?'a = "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly"] +lemmas Lcm_poly_set_eq_fold [code] = + Lcm_set_eq_fold [where ?'a = "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly"] text \Example: @{lemma "Lcm {[:1, 2, 3:], [:2, 3, 4:]} = [:[:2:], [:7:], [:16:], [:17:], [:12 :: int:]:]" by eval} \ end diff --git a/src/HOL/GCD.thy b/src/HOL/GCD.thy --- a/src/HOL/GCD.thy +++ b/src/HOL/GCD.thy @@ -1,2803 +1,2787 @@ (* Title: HOL/GCD.thy Author: Christophe Tabacznyj Author: Lawrence C. Paulson Author: Amine Chaieb Author: Thomas M. Rasmussen Author: Jeremy Avigad Author: Tobias Nipkow This file deals with the functions gcd and lcm. Definitions and lemmas are proved uniformly for the natural numbers and integers. This file combines and revises a number of prior developments. The original theories "GCD" and "Primes" were by Christophe Tabacznyj and Lawrence C. Paulson, based on @{cite davenport92}. They introduced gcd, lcm, and prime for the natural numbers. The original theory "IntPrimes" was by Thomas M. Rasmussen, and extended gcd, lcm, primes to the integers. Amine Chaieb provided another extension of the notions to the integers, and added a number of results to "Primes" and "GCD". IntPrimes also defined and developed the congruence relations on the integers. The notion was extended to the natural numbers by Chaieb. Jeremy Avigad combined all of these, made everything uniform for the natural numbers and the integers, and added a number of new theorems. Tobias Nipkow cleaned up a lot. *) section \Greatest common divisor and least common multiple\ theory GCD imports Groups_List begin subsection \Abstract bounded quasi semilattices as common foundation\ locale bounded_quasi_semilattice = abel_semigroup + fixes top :: 'a ("\<^bold>\") and bot :: 'a ("\<^bold>\") and normalize :: "'a \ 'a" assumes idem_normalize [simp]: "a \<^bold>* a = normalize a" and normalize_left_idem [simp]: "normalize a \<^bold>* b = a \<^bold>* b" and normalize_idem [simp]: "normalize (a \<^bold>* b) = a \<^bold>* b" and normalize_top [simp]: "normalize \<^bold>\ = \<^bold>\" and normalize_bottom [simp]: "normalize \<^bold>\ = \<^bold>\" and top_left_normalize [simp]: "\<^bold>\ \<^bold>* a = normalize a" and bottom_left_bottom [simp]: "\<^bold>\ \<^bold>* a = \<^bold>\" begin lemma left_idem [simp]: "a \<^bold>* (a \<^bold>* b) = a \<^bold>* b" using assoc [of a a b, symmetric] by simp lemma right_idem [simp]: "(a \<^bold>* b) \<^bold>* b = a \<^bold>* b" using left_idem [of b a] by (simp add: ac_simps) lemma comp_fun_idem: "comp_fun_idem f" by standard (simp_all add: fun_eq_iff ac_simps) interpretation comp_fun_idem f by (fact comp_fun_idem) lemma top_right_normalize [simp]: "a \<^bold>* \<^bold>\ = normalize a" using top_left_normalize [of a] by (simp add: ac_simps) lemma bottom_right_bottom [simp]: "a \<^bold>* \<^bold>\ = \<^bold>\" using bottom_left_bottom [of a] by (simp add: ac_simps) lemma normalize_right_idem [simp]: "a \<^bold>* normalize b = a \<^bold>* b" using normalize_left_idem [of b a] by (simp add: ac_simps) end locale bounded_quasi_semilattice_set = bounded_quasi_semilattice begin interpretation comp_fun_idem f by (fact comp_fun_idem) definition F :: "'a set \ 'a" where eq_fold: "F A = (if finite A then Finite_Set.fold f \<^bold>\ A else \<^bold>\)" lemma infinite [simp]: "infinite A \ F A = \<^bold>\" by (simp add: eq_fold) lemma set_eq_fold [code]: "F (set xs) = fold f xs \<^bold>\" by (simp add: eq_fold fold_set_fold) lemma empty [simp]: "F {} = \<^bold>\" by (simp add: eq_fold) lemma insert [simp]: "F (insert a A) = a \<^bold>* F A" by (cases "finite A") (simp_all add: eq_fold) lemma normalize [simp]: "normalize (F A) = F A" by (induct A rule: infinite_finite_induct) simp_all lemma in_idem: assumes "a \ A" shows "a \<^bold>* F A = F A" using assms by (induct A rule: infinite_finite_induct) (auto simp: left_commute [of a]) lemma union: "F (A \ B) = F A \<^bold>* F B" by (induct A rule: infinite_finite_induct) (simp_all add: ac_simps) lemma remove: assumes "a \ A" shows "F A = a \<^bold>* F (A - {a})" proof - from assms obtain B where "A = insert a B" and "a \ B" by (blast dest: mk_disjoint_insert) with assms show ?thesis by simp qed lemma insert_remove: "F (insert a A) = a \<^bold>* F (A - {a})" by (cases "a \ A") (simp_all add: insert_absorb remove) lemma subset: assumes "B \ A" shows "F B \<^bold>* F A = F A" using assms by (simp add: union [symmetric] Un_absorb1) end subsection \Abstract GCD and LCM\ class gcd = zero + one + dvd + fixes gcd :: "'a \ 'a \ 'a" and lcm :: "'a \ 'a \ 'a" class Gcd = gcd + fixes Gcd :: "'a set \ 'a" and Lcm :: "'a set \ 'a" syntax "_GCD1" :: "pttrns \ 'b \ 'b" ("(3GCD _./ _)" [0, 10] 10) "_GCD" :: "pttrn \ 'a set \ 'b \ 'b" ("(3GCD _\_./ _)" [0, 0, 10] 10) "_LCM1" :: "pttrns \ 'b \ 'b" ("(3LCM _./ _)" [0, 10] 10) "_LCM" :: "pttrn \ 'a set \ 'b \ 'b" ("(3LCM _\_./ _)" [0, 0, 10] 10) translations "GCD x y. f" \ "GCD x. GCD y. f" "GCD x. f" \ "CONST Gcd (CONST range (\x. f))" "GCD x\A. f" \ "CONST Gcd ((\x. f) ` A)" "LCM x y. f" \ "LCM x. LCM y. f" "LCM x. f" \ "CONST Lcm (CONST range (\x. f))" "LCM x\A. f" \ "CONST Lcm ((\x. f) ` A)" class semiring_gcd = normalization_semidom + gcd + assumes gcd_dvd1 [iff]: "gcd a b dvd a" and gcd_dvd2 [iff]: "gcd a b dvd b" and gcd_greatest: "c dvd a \ c dvd b \ c dvd gcd a b" and normalize_gcd [simp]: "normalize (gcd a b) = gcd a b" - and lcm_gcd: "lcm a b = normalize (a * b) div gcd a b" + and lcm_gcd: "lcm a b = normalize (a * b div gcd a b)" begin lemma gcd_greatest_iff [simp]: "a dvd gcd b c \ a dvd b \ a dvd c" by (blast intro!: gcd_greatest intro: dvd_trans) lemma gcd_dvdI1: "a dvd c \ gcd a b dvd c" by (rule dvd_trans) (rule gcd_dvd1) lemma gcd_dvdI2: "b dvd c \ gcd a b dvd c" by (rule dvd_trans) (rule gcd_dvd2) lemma dvd_gcdD1: "a dvd gcd b c \ a dvd b" using gcd_dvd1 [of b c] by (blast intro: dvd_trans) lemma dvd_gcdD2: "a dvd gcd b c \ a dvd c" using gcd_dvd2 [of b c] by (blast intro: dvd_trans) lemma gcd_0_left [simp]: "gcd 0 a = normalize a" by (rule associated_eqI) simp_all lemma gcd_0_right [simp]: "gcd a 0 = normalize a" by (rule associated_eqI) simp_all lemma gcd_eq_0_iff [simp]: "gcd a b = 0 \ a = 0 \ b = 0" (is "?P \ ?Q") proof assume ?P then have "0 dvd gcd a b" by simp then have "0 dvd a" and "0 dvd b" by (blast intro: dvd_trans)+ then show ?Q by simp next assume ?Q then show ?P by simp qed lemma unit_factor_gcd: "unit_factor (gcd a b) = (if a = 0 \ b = 0 then 0 else 1)" proof (cases "gcd a b = 0") case True then show ?thesis by simp next case False have "unit_factor (gcd a b) * normalize (gcd a b) = gcd a b" by (rule unit_factor_mult_normalize) then have "unit_factor (gcd a b) * gcd a b = gcd a b" by simp then have "unit_factor (gcd a b) * gcd a b div gcd a b = gcd a b div gcd a b" by simp with False show ?thesis by simp qed lemma is_unit_gcd_iff [simp]: "is_unit (gcd a b) \ gcd a b = 1" by (cases "a = 0 \ b = 0") (auto simp: unit_factor_gcd dest: is_unit_unit_factor) sublocale gcd: abel_semigroup gcd proof fix a b c show "gcd a b = gcd b a" by (rule associated_eqI) simp_all from gcd_dvd1 have "gcd (gcd a b) c dvd a" by (rule dvd_trans) simp moreover from gcd_dvd1 have "gcd (gcd a b) c dvd b" by (rule dvd_trans) simp ultimately have P1: "gcd (gcd a b) c dvd gcd a (gcd b c)" by (auto intro!: gcd_greatest) from gcd_dvd2 have "gcd a (gcd b c) dvd b" by (rule dvd_trans) simp moreover from gcd_dvd2 have "gcd a (gcd b c) dvd c" by (rule dvd_trans) simp ultimately have P2: "gcd a (gcd b c) dvd gcd (gcd a b) c" by (auto intro!: gcd_greatest) from P1 P2 show "gcd (gcd a b) c = gcd a (gcd b c)" by (rule associated_eqI) simp_all qed sublocale gcd: bounded_quasi_semilattice gcd 0 1 normalize proof show "gcd a a = normalize a" for a proof - have "a dvd gcd a a" by (rule gcd_greatest) simp_all then show ?thesis by (auto intro: associated_eqI) qed show "gcd (normalize a) b = gcd a b" for a b using gcd_dvd1 [of "normalize a" b] by (auto intro: associated_eqI) show "gcd 1 a = 1" for a by (rule associated_eqI) simp_all qed simp_all lemma gcd_self: "gcd a a = normalize a" by (fact gcd.idem_normalize) lemma gcd_left_idem: "gcd a (gcd a b) = gcd a b" by (fact gcd.left_idem) lemma gcd_right_idem: "gcd (gcd a b) b = gcd a b" by (fact gcd.right_idem) -lemma gcd_mult_left: "gcd (c * a) (c * b) = normalize c * gcd a b" -proof (cases "c = 0") - case True - then show ?thesis by simp -next - case False - then have *: "c * gcd a b dvd gcd (c * a) (c * b)" - by (auto intro: gcd_greatest) - moreover from False * have "gcd (c * a) (c * b) dvd c * gcd a b" - by (metis div_dvd_iff_mult dvd_mult_left gcd_dvd1 gcd_dvd2 gcd_greatest mult_commute) - ultimately have "normalize (gcd (c * a) (c * b)) = normalize (c * gcd a b)" - by (auto intro: associated_eqI) - then show ?thesis - by (simp add: normalize_mult) -qed - -lemma gcd_mult_right: "gcd (a * c) (b * c) = gcd b a * normalize c" - using gcd_mult_left [of c a b] by (simp add: ac_simps) - -lemma mult_gcd_left: "c * gcd a b = unit_factor c * gcd (c * a) (c * b)" - by (simp add: gcd_mult_left mult.assoc [symmetric]) - -lemma mult_gcd_right: "gcd a b * c = gcd (a * c) (b * c) * unit_factor c" - using mult_gcd_left [of c a b] by (simp add: ac_simps) - -lemma dvd_lcm1 [iff]: "a dvd lcm a b" -proof - - have "normalize (a * b) div gcd a b = normalize a * (normalize b div gcd a b)" - by (simp add: lcm_gcd normalize_mult div_mult_swap) - then show ?thesis - by (simp add: lcm_gcd) -qed - -lemma dvd_lcm2 [iff]: "b dvd lcm a b" -proof - - have "normalize (a * b) div gcd a b = normalize b * (normalize a div gcd a b)" - by (simp add: lcm_gcd normalize_mult div_mult_swap ac_simps) - then show ?thesis - by (simp add: lcm_gcd) -qed - -lemma dvd_lcmI1: "a dvd b \ a dvd lcm b c" - by (rule dvd_trans) (assumption, blast) - -lemma dvd_lcmI2: "a dvd c \ a dvd lcm b c" - by (rule dvd_trans) (assumption, blast) - -lemma lcm_dvdD1: "lcm a b dvd c \ a dvd c" - using dvd_lcm1 [of a b] by (blast intro: dvd_trans) - -lemma lcm_dvdD2: "lcm a b dvd c \ b dvd c" - using dvd_lcm2 [of a b] by (blast intro: dvd_trans) - -lemma lcm_least: - assumes "a dvd c" and "b dvd c" - shows "lcm a b dvd c" -proof (cases "c = 0") - case True - then show ?thesis by simp -next - case False - then have *: "is_unit (unit_factor c)" - by simp - show ?thesis - proof (cases "gcd a b = 0") - case True - with assms show ?thesis by simp - next - case False - then have "a \ 0 \ b \ 0" - by simp - with \c \ 0\ assms have "a * b dvd a * c" "a * b dvd c * b" - by (simp_all add: mult_dvd_mono) - then have "normalize (a * b) dvd gcd (a * c) (b * c)" - by (auto intro: gcd_greatest simp add: ac_simps) - then have "normalize (a * b) dvd gcd (a * c) (b * c) * unit_factor c" - using * by (simp add: dvd_mult_unit_iff) - then have "normalize (a * b) dvd gcd a b * c" - by (simp add: mult_gcd_right [of a b c]) - then have "normalize (a * b) div gcd a b dvd c" - using False by (simp add: div_dvd_iff_mult ac_simps) - then show ?thesis - by (simp add: lcm_gcd) - qed -qed - -lemma lcm_least_iff [simp]: "lcm a b dvd c \ a dvd c \ b dvd c" - by (blast intro!: lcm_least intro: dvd_trans) - -lemma normalize_lcm [simp]: "normalize (lcm a b) = lcm a b" - by (simp add: lcm_gcd dvd_normalize_div) - -lemma lcm_0_left [simp]: "lcm 0 a = 0" - by (simp add: lcm_gcd) - -lemma lcm_0_right [simp]: "lcm a 0 = 0" - by (simp add: lcm_gcd) - -lemma lcm_eq_0_iff: "lcm a b = 0 \ a = 0 \ b = 0" - (is "?P \ ?Q") -proof - assume ?P - then have "0 dvd lcm a b" - by simp - then have "0 dvd normalize (a * b) div gcd a b" - by (simp add: lcm_gcd) - then have "0 * gcd a b dvd normalize (a * b)" - using dvd_div_iff_mult [of "gcd a b" _ 0] by (cases "gcd a b = 0") simp_all - then have "normalize (a * b) = 0" - by simp - then show ?Q - by simp -next - assume ?Q - then show ?P - by auto -qed - -lemma lcm_eq_1_iff [simp]: "lcm a b = 1 \ is_unit a \ is_unit b" - by (auto intro: associated_eqI) - -lemma unit_factor_lcm: "unit_factor (lcm a b) = (if a = 0 \ b = 0 then 0 else 1)" - by (simp add: unit_factor_gcd dvd_unit_factor_div lcm_gcd) - -sublocale lcm: abel_semigroup lcm -proof - fix a b c - show "lcm a b = lcm b a" - by (simp add: lcm_gcd ac_simps normalize_mult dvd_normalize_div) - have "lcm (lcm a b) c dvd lcm a (lcm b c)" - and "lcm a (lcm b c) dvd lcm (lcm a b) c" - by (auto intro: lcm_least - dvd_trans [of b "lcm b c" "lcm a (lcm b c)"] - dvd_trans [of c "lcm b c" "lcm a (lcm b c)"] - dvd_trans [of a "lcm a b" "lcm (lcm a b) c"] - dvd_trans [of b "lcm a b" "lcm (lcm a b) c"]) - then show "lcm (lcm a b) c = lcm a (lcm b c)" - by (rule associated_eqI) simp_all -qed - -sublocale lcm: bounded_quasi_semilattice lcm 1 0 normalize -proof - show "lcm a a = normalize a" for a - proof - - have "lcm a a dvd a" - by (rule lcm_least) simp_all - then show ?thesis - by (auto intro: associated_eqI) - qed - show "lcm (normalize a) b = lcm a b" for a b - using dvd_lcm1 [of "normalize a" b] unfolding normalize_dvd_iff - by (auto intro: associated_eqI) - show "lcm 1 a = normalize a" for a - by (rule associated_eqI) simp_all -qed simp_all - -lemma lcm_self: "lcm a a = normalize a" - by (fact lcm.idem_normalize) - -lemma lcm_left_idem: "lcm a (lcm a b) = lcm a b" - by (fact lcm.left_idem) - -lemma lcm_right_idem: "lcm (lcm a b) b = lcm a b" - by (fact lcm.right_idem) - -lemma gcd_mult_lcm [simp]: "gcd a b * lcm a b = normalize a * normalize b" - by (simp add: lcm_gcd normalize_mult) - -lemma lcm_mult_gcd [simp]: "lcm a b * gcd a b = normalize a * normalize b" - using gcd_mult_lcm [of a b] by (simp add: ac_simps) - -lemma gcd_lcm: - assumes "a \ 0" and "b \ 0" - shows "gcd a b = normalize (a * b) div lcm a b" -proof - - from assms have "lcm a b \ 0" - by (simp add: lcm_eq_0_iff) - have "gcd a b * lcm a b = normalize a * normalize b" - by simp - then have "gcd a b * lcm a b div lcm a b = normalize (a * b) div lcm a b" - by (simp_all add: normalize_mult) - with \lcm a b \ 0\ show ?thesis - using nonzero_mult_div_cancel_right [of "lcm a b" "gcd a b"] by simp -qed - -lemma lcm_1_left: "lcm 1 a = normalize a" - by (fact lcm.top_left_normalize) - -lemma lcm_1_right: "lcm a 1 = normalize a" - by (fact lcm.top_right_normalize) - -lemma lcm_mult_left: "lcm (c * a) (c * b) = normalize c * lcm a b" - by (cases "c = 0") - (simp_all add: gcd_mult_right lcm_gcd div_mult_swap normalize_mult ac_simps, - simp add: dvd_div_mult2_eq mult.left_commute [of "normalize c", symmetric]) - -lemma lcm_mult_right: "lcm (a * c) (b * c) = lcm b a * normalize c" - using lcm_mult_left [of c a b] by (simp add: ac_simps) - -lemma mult_lcm_left: "c * lcm a b = unit_factor c * lcm (c * a) (c * b)" - by (simp add: lcm_mult_left mult.assoc [symmetric]) - -lemma mult_lcm_right: "lcm a b * c = lcm (a * c) (b * c) * unit_factor c" - using mult_lcm_left [of c a b] by (simp add: ac_simps) - lemma gcdI: assumes "c dvd a" and "c dvd b" and greatest: "\d. d dvd a \ d dvd b \ d dvd c" and "normalize c = c" shows "c = gcd a b" by (rule associated_eqI) (auto simp: assms intro: gcd_greatest) lemma gcd_unique: "d dvd a \ d dvd b \ normalize d = d \ (\e. e dvd a \ e dvd b \ e dvd d) \ d = gcd a b" by rule (auto intro: gcdI simp: gcd_greatest) lemma gcd_dvd_prod: "gcd a b dvd k * b" using mult_dvd_mono [of 1] by auto lemma gcd_proj2_if_dvd: "b dvd a \ gcd a b = normalize b" by (rule gcdI [symmetric]) simp_all lemma gcd_proj1_if_dvd: "a dvd b \ gcd a b = normalize a" by (rule gcdI [symmetric]) simp_all lemma gcd_proj1_iff: "gcd m n = normalize m \ m dvd n" proof assume *: "gcd m n = normalize m" show "m dvd n" proof (cases "m = 0") case True with * show ?thesis by simp next case [simp]: False from * have **: "m = gcd m n * unit_factor m" by (simp add: unit_eq_div2) show ?thesis by (subst **) (simp add: mult_unit_dvd_iff) qed next assume "m dvd n" then show "gcd m n = normalize m" by (rule gcd_proj1_if_dvd) qed lemma gcd_proj2_iff: "gcd m n = normalize n \ n dvd m" using gcd_proj1_iff [of n m] by (simp add: ac_simps) -lemma gcd_mult_distrib': "normalize c * gcd a b = gcd (c * a) (c * b)" - by (rule gcdI) (auto simp: normalize_mult gcd_greatest mult_dvd_mono gcd_mult_left[symmetric]) - -lemma gcd_mult_distrib: "k * gcd a b = gcd (k * a) (k * b) * unit_factor k" -proof- - have "normalize k * gcd a b = gcd (k * a) (k * b)" - by (simp add: gcd_mult_distrib') - then have "normalize k * gcd a b * unit_factor k = gcd (k * a) (k * b) * unit_factor k" +lemma gcd_mult_left: "gcd (c * a) (c * b) = normalize (c * gcd a b)" +proof (cases "c = 0") + case True + then show ?thesis by simp +next + case False + then have *: "c * gcd a b dvd gcd (c * a) (c * b)" + by (auto intro: gcd_greatest) + moreover from False * have "gcd (c * a) (c * b) dvd c * gcd a b" + by (metis div_dvd_iff_mult dvd_mult_left gcd_dvd1 gcd_dvd2 gcd_greatest mult_commute) + ultimately have "normalize (gcd (c * a) (c * b)) = normalize (c * gcd a b)" + by (auto intro: associated_eqI) + then show ?thesis + by (simp add: normalize_mult) +qed + +lemma gcd_mult_right: "gcd (a * c) (b * c) = normalize (gcd b a * c)" + using gcd_mult_left [of c a b] by (simp add: ac_simps) + +lemma dvd_lcm1 [iff]: "a dvd lcm a b" + by (metis div_mult_swap dvd_mult2 dvd_normalize_iff dvd_refl gcd_dvd2 lcm_gcd) + +lemma dvd_lcm2 [iff]: "b dvd lcm a b" + by (metis dvd_div_mult dvd_mult dvd_normalize_iff dvd_refl gcd_dvd1 lcm_gcd) + +lemma dvd_lcmI1: "a dvd b \ a dvd lcm b c" + by (rule dvd_trans) (assumption, blast) + +lemma dvd_lcmI2: "a dvd c \ a dvd lcm b c" + by (rule dvd_trans) (assumption, blast) + +lemma lcm_dvdD1: "lcm a b dvd c \ a dvd c" + using dvd_lcm1 [of a b] by (blast intro: dvd_trans) + +lemma lcm_dvdD2: "lcm a b dvd c \ b dvd c" + using dvd_lcm2 [of a b] by (blast intro: dvd_trans) + +lemma lcm_least: + assumes "a dvd c" and "b dvd c" + shows "lcm a b dvd c" +proof (cases "c = 0") + case True + then show ?thesis by simp +next + case False + then have *: "is_unit (unit_factor c)" by simp - then have "normalize k * unit_factor k * gcd a b = gcd (k * a) (k * b) * unit_factor k" - by (simp only: ac_simps) + show ?thesis + proof (cases "gcd a b = 0") + case True + with assms show ?thesis by simp + next + case False + have "a * b dvd normalize (c * gcd a b)" + using assms by (subst gcd_mult_left [symmetric]) (auto intro!: gcd_greatest simp: mult_ac) + with False have "(a * b div gcd a b) dvd c" + by (subst div_dvd_iff_mult) auto + thus ?thesis by (simp add: lcm_gcd) + qed +qed + +lemma lcm_least_iff [simp]: "lcm a b dvd c \ a dvd c \ b dvd c" + by (blast intro!: lcm_least intro: dvd_trans) + +lemma normalize_lcm [simp]: "normalize (lcm a b) = lcm a b" + by (simp add: lcm_gcd dvd_normalize_div) + +lemma lcm_0_left [simp]: "lcm 0 a = 0" + by (simp add: lcm_gcd) + +lemma lcm_0_right [simp]: "lcm a 0 = 0" + by (simp add: lcm_gcd) + +lemma lcm_eq_0_iff: "lcm a b = 0 \ a = 0 \ b = 0" + (is "?P \ ?Q") +proof + assume ?P + then have "0 dvd lcm a b" + by simp + also have "lcm a b dvd (a * b)" + by simp + finally show ?Q + by auto +next + assume ?Q + then show ?P + by auto +qed + +lemma zero_eq_lcm_iff: "0 = lcm a b \ a = 0 \ b = 0" + using lcm_eq_0_iff[of a b] by auto + +lemma lcm_eq_1_iff [simp]: "lcm a b = 1 \ is_unit a \ is_unit b" + by (auto intro: associated_eqI) + +lemma unit_factor_lcm: "unit_factor (lcm a b) = (if a = 0 \ b = 0 then 0 else 1)" + using lcm_eq_0_iff[of a b] by (cases "lcm a b = 0") (auto simp: lcm_gcd) + +sublocale lcm: abel_semigroup lcm +proof + fix a b c + show "lcm a b = lcm b a" + by (simp add: lcm_gcd ac_simps normalize_mult dvd_normalize_div) + have "lcm (lcm a b) c dvd lcm a (lcm b c)" + and "lcm a (lcm b c) dvd lcm (lcm a b) c" + by (auto intro: lcm_least + dvd_trans [of b "lcm b c" "lcm a (lcm b c)"] + dvd_trans [of c "lcm b c" "lcm a (lcm b c)"] + dvd_trans [of a "lcm a b" "lcm (lcm a b) c"] + dvd_trans [of b "lcm a b" "lcm (lcm a b) c"]) + then show "lcm (lcm a b) c = lcm a (lcm b c)" + by (rule associated_eqI) simp_all +qed + +sublocale lcm: bounded_quasi_semilattice lcm 1 0 normalize +proof + show "lcm a a = normalize a" for a + proof - + have "lcm a a dvd a" + by (rule lcm_least) simp_all + then show ?thesis + by (auto intro: associated_eqI) + qed + show "lcm (normalize a) b = lcm a b" for a b + using dvd_lcm1 [of "normalize a" b] unfolding normalize_dvd_iff + by (auto intro: associated_eqI) + show "lcm 1 a = normalize a" for a + by (rule associated_eqI) simp_all +qed simp_all + +lemma lcm_self: "lcm a a = normalize a" + by (fact lcm.idem_normalize) + +lemma lcm_left_idem: "lcm a (lcm a b) = lcm a b" + by (fact lcm.left_idem) + +lemma lcm_right_idem: "lcm (lcm a b) b = lcm a b" + by (fact lcm.right_idem) + +lemma gcd_lcm: + assumes "a \ 0" and "b \ 0" + shows "gcd a b = normalize (a * b div lcm a b)" +proof - + from assms have [simp]: "a * b div gcd a b \ 0" + by (subst dvd_div_eq_0_iff) auto + let ?u = "unit_factor (a * b div gcd a b)" + have "gcd a b * normalize (a * b div gcd a b) = + gcd a b * ((a * b div gcd a b) * (1 div ?u))" + by simp + also have "\ = a * b div ?u" + by (subst mult.assoc [symmetric]) auto + also have "\ dvd a * b" + by (subst div_unit_dvd_iff) auto + finally have "gcd a b dvd ((a * b) div lcm a b)" + by (intro dvd_mult_imp_div) (auto simp: lcm_gcd) + moreover have "a * b div lcm a b dvd a" and "a * b div lcm a b dvd b" + using assms by (subst div_dvd_iff_mult; simp add: lcm_eq_0_iff mult.commute[of b "lcm a b"])+ + ultimately have "normalize (gcd a b) = normalize (a * b div lcm a b)" + apply - + apply (rule associated_eqI) + using assms + apply (auto simp: div_dvd_iff_mult zero_eq_lcm_iff) + done + thus ?thesis by simp +qed + +lemma lcm_1_left: "lcm 1 a = normalize a" + by (fact lcm.top_left_normalize) + +lemma lcm_1_right: "lcm a 1 = normalize a" + by (fact lcm.top_right_normalize) + +lemma lcm_mult_left: "lcm (c * a) (c * b) = normalize (c * lcm a b)" +proof (cases "c = 0") + case True + then show ?thesis by simp +next + case False + then have *: "lcm (c * a) (c * b) dvd c * lcm a b" + by (auto intro: lcm_least) + moreover have "lcm a b dvd lcm (c * a) (c * b) div c" + by (intro lcm_least) (auto intro!: dvd_mult_imp_div simp: mult_ac) + hence "c * lcm a b dvd lcm (c * a) (c * b)" + using False by (subst (asm) dvd_div_iff_mult) (auto simp: mult_ac intro: dvd_lcmI1) + ultimately have "normalize (lcm (c * a) (c * b)) = normalize (c * lcm a b)" + by (auto intro: associated_eqI) then show ?thesis - by simp + by (simp add: normalize_mult) qed +lemma lcm_mult_right: "lcm (a * c) (b * c) = normalize (lcm b a * c)" + using lcm_mult_left [of c a b] by (simp add: ac_simps) + lemma lcm_mult_unit1: "is_unit a \ lcm (b * a) c = lcm b c" by (rule associated_eqI) (simp_all add: mult_unit_dvd_iff dvd_lcmI1) lemma lcm_mult_unit2: "is_unit a \ lcm b (c * a) = lcm b c" using lcm_mult_unit1 [of a c b] by (simp add: ac_simps) lemma lcm_div_unit1: "is_unit a \ lcm (b div a) c = lcm b c" by (erule is_unitE [of _ b]) (simp add: lcm_mult_unit1) lemma lcm_div_unit2: "is_unit a \ lcm b (c div a) = lcm b c" by (erule is_unitE [of _ c]) (simp add: lcm_mult_unit2) lemma normalize_lcm_left: "lcm (normalize a) b = lcm a b" by (fact lcm.normalize_left_idem) lemma normalize_lcm_right: "lcm a (normalize b) = lcm a b" by (fact lcm.normalize_right_idem) -lemma gcd_mult_unit1: - assumes "is_unit a" shows "gcd (b * a) c = gcd b c" -proof - - have "gcd (b * a) c dvd b" - using assms local.dvd_mult_unit_iff by blast - then show ?thesis - by (rule gcdI) simp_all -qed - -lemma gcd_mult_unit2: "is_unit a \ gcd b (c * a) = gcd b c" - using gcd.commute gcd_mult_unit1 by auto - -lemma gcd_div_unit1: "is_unit a \ gcd (b div a) c = gcd b c" - by (erule is_unitE [of _ b]) (simp add: gcd_mult_unit1) - -lemma gcd_div_unit2: "is_unit a \ gcd b (c div a) = gcd b c" - by (erule is_unitE [of _ c]) (simp add: gcd_mult_unit2) - -lemma normalize_gcd_left: "gcd (normalize a) b = gcd a b" - by (fact gcd.normalize_left_idem) - -lemma normalize_gcd_right: "gcd a (normalize b) = gcd a b" - by (fact gcd.normalize_right_idem) - lemma comp_fun_idem_gcd: "comp_fun_idem gcd" by standard (simp_all add: fun_eq_iff ac_simps) lemma comp_fun_idem_lcm: "comp_fun_idem lcm" by standard (simp_all add: fun_eq_iff ac_simps) lemma gcd_dvd_antisym: "gcd a b dvd gcd c d \ gcd c d dvd gcd a b \ gcd a b = gcd c d" proof (rule gcdI) assume *: "gcd a b dvd gcd c d" and **: "gcd c d dvd gcd a b" have "gcd c d dvd c" by simp with * show "gcd a b dvd c" by (rule dvd_trans) have "gcd c d dvd d" by simp with * show "gcd a b dvd d" by (rule dvd_trans) show "normalize (gcd a b) = gcd a b" by simp fix l assume "l dvd c" and "l dvd d" then have "l dvd gcd c d" by (rule gcd_greatest) from this and ** show "l dvd gcd a b" by (rule dvd_trans) qed -lemma gcd_add1 [simp]: "gcd (m + n) n = gcd m n" - by (rule gcdI [symmetric]) (simp_all add: dvd_add_left_iff) - -lemma gcd_add2 [simp]: "gcd m (m + n) = gcd m n" - using gcd_add1 [of n m] by (simp add: ac_simps) - -lemma gcd_add_mult: "gcd m (k * m + n) = gcd m n" - by (rule gcdI [symmetric]) (simp_all add: dvd_add_right_iff) - -lemma lcm_gcd_prod: "lcm a b * gcd a b = normalize (a * b)" - by (simp add: lcm_gcd) - declare unit_factor_lcm [simp] lemma lcmI: assumes "a dvd c" and "b dvd c" and "\d. a dvd d \ b dvd d \ c dvd d" and "normalize c = c" shows "c = lcm a b" by (rule associated_eqI) (auto simp: assms intro: lcm_least) lemma gcd_dvd_lcm [simp]: "gcd a b dvd lcm a b" using gcd_dvd2 by (rule dvd_lcmI2) lemmas lcm_0 = lcm_0_right lemma lcm_unique: "a dvd d \ b dvd d \ normalize d = d \ (\e. a dvd e \ b dvd e \ d dvd e) \ d = lcm a b" by rule (auto intro: lcmI simp: lcm_least lcm_eq_0_iff) lemma lcm_proj1_if_dvd: assumes "b dvd a" shows "lcm a b = normalize a" -proof (cases "a = 0") - case False - then show ?thesis - using assms gcd_proj2_if_dvd lcm_mult_gcd local.lcm_gcd by auto -qed auto +proof - + have "normalize (lcm a b) = normalize a" + by (rule associatedI) (use assms in auto) + thus ?thesis by simp +qed lemma lcm_proj2_if_dvd: "a dvd b \ lcm a b = normalize b" using lcm_proj1_if_dvd [of a b] by (simp add: ac_simps) lemma lcm_proj1_iff: "lcm m n = normalize m \ n dvd m" proof assume *: "lcm m n = normalize m" show "n dvd m" proof (cases "m = 0") case True then show ?thesis by simp next case [simp]: False from * have **: "m = lcm m n * unit_factor m" by (simp add: unit_eq_div2) show ?thesis by (subst **) simp qed next assume "n dvd m" then show "lcm m n = normalize m" by (rule lcm_proj1_if_dvd) qed lemma lcm_proj2_iff: "lcm m n = normalize n \ m dvd n" using lcm_proj1_iff [of n m] by (simp add: ac_simps) -lemma lcm_mult_distrib': "normalize c * lcm a b = lcm (c * a) (c * b)" - by (rule lcmI) (auto simp: normalize_mult lcm_least mult_dvd_mono lcm_mult_left [symmetric]) - -lemma lcm_mult_distrib: "k * lcm a b = lcm (k * a) (k * b) * unit_factor k" -proof- - have "normalize k * lcm a b = lcm (k * a) (k * b)" - by (simp add: lcm_mult_distrib') - then have "normalize k * lcm a b * unit_factor k = lcm (k * a) (k * b) * unit_factor k" - by simp - then have "normalize k * unit_factor k * lcm a b = lcm (k * a) (k * b) * unit_factor k" - by (simp only: ac_simps) - then show ?thesis - by simp -qed - lemma gcd_mono: "a dvd c \ b dvd d \ gcd a b dvd gcd c d" by (simp add: gcd_dvdI1 gcd_dvdI2) lemma lcm_mono: "a dvd c \ b dvd d \ lcm a b dvd lcm c d" by (simp add: dvd_lcmI1 dvd_lcmI2) lemma dvd_productE: assumes "p dvd a * b" obtains x y where "p = x * y" "x dvd a" "y dvd b" proof (cases "a = 0") case True thus ?thesis by (intro that[of p 1]) simp_all next case False define x y where "x = gcd a p" and "y = p div x" have "p = x * y" by (simp add: x_def y_def) moreover have "x dvd a" by (simp add: x_def) moreover from assms have "p dvd gcd (b * a) (b * p)" by (intro gcd_greatest) (simp_all add: mult.commute) - hence "p dvd b * gcd a p" by (simp add: gcd_mult_distrib) + hence "p dvd b * gcd a p" by (subst (asm) gcd_mult_left) auto with False have "y dvd b" by (simp add: x_def y_def div_dvd_iff_mult assms) ultimately show ?thesis by (rule that) qed +lemma gcd_mult_unit1: + assumes "is_unit a" shows "gcd (b * a) c = gcd b c" +proof - + have "gcd (b * a) c dvd b" + using assms dvd_mult_unit_iff by blast + then show ?thesis + by (rule gcdI) simp_all +qed + +lemma gcd_mult_unit2: "is_unit a \ gcd b (c * a) = gcd b c" + using gcd.commute gcd_mult_unit1 by auto + +lemma gcd_div_unit1: "is_unit a \ gcd (b div a) c = gcd b c" + by (erule is_unitE [of _ b]) (simp add: gcd_mult_unit1) + +lemma gcd_div_unit2: "is_unit a \ gcd b (c div a) = gcd b c" + by (erule is_unitE [of _ c]) (simp add: gcd_mult_unit2) + +lemma normalize_gcd_left: "gcd (normalize a) b = gcd a b" + by (fact gcd.normalize_left_idem) + +lemma normalize_gcd_right: "gcd a (normalize b) = gcd a b" + by (fact gcd.normalize_right_idem) + +lemma gcd_add1 [simp]: "gcd (m + n) n = gcd m n" + by (rule gcdI [symmetric]) (simp_all add: dvd_add_left_iff) + +lemma gcd_add2 [simp]: "gcd m (m + n) = gcd m n" + using gcd_add1 [of n m] by (simp add: ac_simps) + +lemma gcd_add_mult: "gcd m (k * m + n) = gcd m n" + by (rule gcdI [symmetric]) (simp_all add: dvd_add_right_iff) + end class ring_gcd = comm_ring_1 + semiring_gcd begin lemma gcd_neg1 [simp]: "gcd (-a) b = gcd a b" by (rule sym, rule gcdI) (simp_all add: gcd_greatest) lemma gcd_neg2 [simp]: "gcd a (-b) = gcd a b" by (rule sym, rule gcdI) (simp_all add: gcd_greatest) lemma gcd_neg_numeral_1 [simp]: "gcd (- numeral n) a = gcd (numeral n) a" by (fact gcd_neg1) lemma gcd_neg_numeral_2 [simp]: "gcd a (- numeral n) = gcd a (numeral n)" by (fact gcd_neg2) lemma gcd_diff1: "gcd (m - n) n = gcd m n" by (subst diff_conv_add_uminus, subst gcd_neg2[symmetric], subst gcd_add1, simp) lemma gcd_diff2: "gcd (n - m) n = gcd m n" by (subst gcd_neg1[symmetric]) (simp only: minus_diff_eq gcd_diff1) lemma lcm_neg1 [simp]: "lcm (-a) b = lcm a b" by (rule sym, rule lcmI) (simp_all add: lcm_least lcm_eq_0_iff) lemma lcm_neg2 [simp]: "lcm a (-b) = lcm a b" by (rule sym, rule lcmI) (simp_all add: lcm_least lcm_eq_0_iff) lemma lcm_neg_numeral_1 [simp]: "lcm (- numeral n) a = lcm (numeral n) a" by (fact lcm_neg1) lemma lcm_neg_numeral_2 [simp]: "lcm a (- numeral n) = lcm a (numeral n)" by (fact lcm_neg2) end class semiring_Gcd = semiring_gcd + Gcd + assumes Gcd_dvd: "a \ A \ Gcd A dvd a" and Gcd_greatest: "(\b. b \ A \ a dvd b) \ a dvd Gcd A" and normalize_Gcd [simp]: "normalize (Gcd A) = Gcd A" assumes dvd_Lcm: "a \ A \ a dvd Lcm A" and Lcm_least: "(\b. b \ A \ b dvd a) \ Lcm A dvd a" and normalize_Lcm [simp]: "normalize (Lcm A) = Lcm A" begin lemma Lcm_Gcd: "Lcm A = Gcd {b. \a\A. a dvd b}" by (rule associated_eqI) (auto intro: Gcd_dvd dvd_Lcm Gcd_greatest Lcm_least) lemma Gcd_Lcm: "Gcd A = Lcm {b. \a\A. b dvd a}" by (rule associated_eqI) (auto intro: Gcd_dvd dvd_Lcm Gcd_greatest Lcm_least) lemma Gcd_empty [simp]: "Gcd {} = 0" by (rule dvd_0_left, rule Gcd_greatest) simp lemma Lcm_empty [simp]: "Lcm {} = 1" by (auto intro: associated_eqI Lcm_least) lemma Gcd_insert [simp]: "Gcd (insert a A) = gcd a (Gcd A)" proof - have "Gcd (insert a A) dvd gcd a (Gcd A)" by (auto intro: Gcd_dvd Gcd_greatest) moreover have "gcd a (Gcd A) dvd Gcd (insert a A)" proof (rule Gcd_greatest) fix b assume "b \ insert a A" then show "gcd a (Gcd A) dvd b" proof assume "b = a" then show ?thesis by simp next assume "b \ A" then have "Gcd A dvd b" by (rule Gcd_dvd) moreover have "gcd a (Gcd A) dvd Gcd A" by simp ultimately show ?thesis by (blast intro: dvd_trans) qed qed ultimately show ?thesis by (auto intro: associated_eqI) qed lemma Lcm_insert [simp]: "Lcm (insert a A) = lcm a (Lcm A)" proof (rule sym) have "lcm a (Lcm A) dvd Lcm (insert a A)" by (auto intro: dvd_Lcm Lcm_least) moreover have "Lcm (insert a A) dvd lcm a (Lcm A)" proof (rule Lcm_least) fix b assume "b \ insert a A" then show "b dvd lcm a (Lcm A)" proof assume "b = a" then show ?thesis by simp next assume "b \ A" then have "b dvd Lcm A" by (rule dvd_Lcm) moreover have "Lcm A dvd lcm a (Lcm A)" by simp ultimately show ?thesis by (blast intro: dvd_trans) qed qed ultimately show "lcm a (Lcm A) = Lcm (insert a A)" by (rule associated_eqI) (simp_all add: lcm_eq_0_iff) qed lemma LcmI: assumes "\a. a \ A \ a dvd b" and "\c. (\a. a \ A \ a dvd c) \ b dvd c" and "normalize b = b" shows "b = Lcm A" by (rule associated_eqI) (auto simp: assms dvd_Lcm intro: Lcm_least) lemma Lcm_subset: "A \ B \ Lcm A dvd Lcm B" by (blast intro: Lcm_least dvd_Lcm) lemma Lcm_Un: "Lcm (A \ B) = lcm (Lcm A) (Lcm B)" proof - have "\d. \Lcm A dvd d; Lcm B dvd d\ \ Lcm (A \ B) dvd d" - by (meson UnE local.Lcm_least local.dvd_Lcm local.dvd_trans) + by (meson UnE Lcm_least dvd_Lcm dvd_trans) then show ?thesis - by (meson Lcm_subset local.lcm_unique local.normalize_Lcm sup.cobounded1 sup.cobounded2) + by (meson Lcm_subset lcm_unique normalize_Lcm sup.cobounded1 sup.cobounded2) qed lemma Gcd_0_iff [simp]: "Gcd A = 0 \ A \ {0}" (is "?P \ ?Q") proof assume ?P show ?Q proof fix a assume "a \ A" then have "Gcd A dvd a" by (rule Gcd_dvd) with \?P\ have "a = 0" by simp then show "a \ {0}" by simp qed next assume ?Q have "0 dvd Gcd A" proof (rule Gcd_greatest) fix a assume "a \ A" with \?Q\ have "a = 0" by auto then show "0 dvd a" by simp qed then show ?P by simp qed lemma Lcm_1_iff [simp]: "Lcm A = 1 \ (\a\A. is_unit a)" (is "?P \ ?Q") proof assume ?P show ?Q proof fix a assume "a \ A" then have "a dvd Lcm A" by (rule dvd_Lcm) with \?P\ show "is_unit a" by simp qed next assume ?Q then have "is_unit (Lcm A)" by (blast intro: Lcm_least) then have "normalize (Lcm A) = 1" by (rule is_unit_normalize) then show ?P by simp qed lemma unit_factor_Lcm: "unit_factor (Lcm A) = (if Lcm A = 0 then 0 else 1)" proof (cases "Lcm A = 0") case True then show ?thesis by simp next case False with unit_factor_normalize have "unit_factor (normalize (Lcm A)) = 1" by blast with False show ?thesis by simp qed lemma unit_factor_Gcd: "unit_factor (Gcd A) = (if Gcd A = 0 then 0 else 1)" by (simp add: Gcd_Lcm unit_factor_Lcm) lemma GcdI: assumes "\a. a \ A \ b dvd a" and "\c. (\a. a \ A \ c dvd a) \ c dvd b" and "normalize b = b" shows "b = Gcd A" by (rule associated_eqI) (auto simp: assms Gcd_dvd intro: Gcd_greatest) lemma Gcd_eq_1_I: assumes "is_unit a" and "a \ A" shows "Gcd A = 1" proof - from assms have "is_unit (Gcd A)" by (blast intro: Gcd_dvd dvd_unit_imp_unit) then have "normalize (Gcd A) = 1" by (rule is_unit_normalize) then show ?thesis by simp qed lemma Lcm_eq_0_I: assumes "0 \ A" shows "Lcm A = 0" proof - from assms have "0 dvd Lcm A" by (rule dvd_Lcm) then show ?thesis by simp qed lemma Gcd_UNIV [simp]: "Gcd UNIV = 1" using dvd_refl by (rule Gcd_eq_1_I) simp lemma Lcm_UNIV [simp]: "Lcm UNIV = 0" by (rule Lcm_eq_0_I) simp lemma Lcm_0_iff: assumes "finite A" shows "Lcm A = 0 \ 0 \ A" proof (cases "A = {}") case True then show ?thesis by simp next case False with assms show ?thesis by (induct A rule: finite_ne_induct) (auto simp: lcm_eq_0_iff) qed lemma Gcd_image_normalize [simp]: "Gcd (normalize ` A) = Gcd A" proof - have "Gcd (normalize ` A) dvd a" if "a \ A" for a proof - from that obtain B where "A = insert a B" by blast moreover have "gcd (normalize a) (Gcd (normalize ` B)) dvd normalize a" by (rule gcd_dvd1) ultimately show "Gcd (normalize ` A) dvd a" by simp qed then have "Gcd (normalize ` A) dvd Gcd A" and "Gcd A dvd Gcd (normalize ` A)" by (auto intro!: Gcd_greatest intro: Gcd_dvd) then show ?thesis by (auto intro: associated_eqI) qed lemma Gcd_eqI: assumes "normalize a = a" assumes "\b. b \ A \ a dvd b" and "\c. (\b. b \ A \ c dvd b) \ c dvd a" shows "Gcd A = a" using assms by (blast intro: associated_eqI Gcd_greatest Gcd_dvd normalize_Gcd) lemma dvd_GcdD: "x dvd Gcd A \ y \ A \ x dvd y" using Gcd_dvd dvd_trans by blast lemma dvd_Gcd_iff: "x dvd Gcd A \ (\y\A. x dvd y)" by (blast dest: dvd_GcdD intro: Gcd_greatest) -lemma Gcd_mult: "Gcd ((*) c ` A) = normalize c * Gcd A" +lemma Gcd_mult: "Gcd ((*) c ` A) = normalize (c * Gcd A)" proof (cases "c = 0") case True then show ?thesis by auto next case [simp]: False have "Gcd ((*) c ` A) div c dvd Gcd A" by (intro Gcd_greatest, subst div_dvd_iff_mult) (auto intro!: Gcd_greatest Gcd_dvd simp: mult.commute[of _ c]) then have "Gcd ((*) c ` A) dvd c * Gcd A" by (subst (asm) div_dvd_iff_mult) (auto intro: Gcd_greatest simp: mult_ac) - also have "c * Gcd A = (normalize c * Gcd A) * unit_factor c" - by (subst unit_factor_mult_normalize [symmetric]) (simp only: mult_ac) - also have "Gcd ((*) c ` A) dvd \ \ Gcd ((*) c ` A) dvd normalize c * Gcd A" - by (simp add: dvd_mult_unit_iff) - finally have "Gcd ((*) c ` A) dvd normalize c * Gcd A" . - moreover have "normalize c * Gcd A dvd Gcd ((*) c ` A)" + moreover have "c * Gcd A dvd Gcd ((*) c ` A)" by (intro Gcd_greatest) (auto intro: mult_dvd_mono Gcd_dvd) - ultimately have "normalize (Gcd ((*) c ` A)) = normalize (normalize c * Gcd A)" + ultimately have "normalize (Gcd ((*) c ` A)) = normalize (c * Gcd A)" by (rule associatedI) - then show ?thesis - by (simp add: normalize_mult) + then show ?thesis by simp qed lemma Lcm_eqI: assumes "normalize a = a" and "\b. b \ A \ b dvd a" and "\c. (\b. b \ A \ b dvd c) \ a dvd c" shows "Lcm A = a" using assms by (blast intro: associated_eqI Lcm_least dvd_Lcm normalize_Lcm) lemma Lcm_dvdD: "Lcm A dvd x \ y \ A \ y dvd x" using dvd_Lcm dvd_trans by blast lemma Lcm_dvd_iff: "Lcm A dvd x \ (\y\A. y dvd x)" by (blast dest: Lcm_dvdD intro: Lcm_least) lemma Lcm_mult: assumes "A \ {}" - shows "Lcm ((*) c ` A) = normalize c * Lcm A" + shows "Lcm ((*) c ` A) = normalize (c * Lcm A)" proof (cases "c = 0") case True with assms have "(*) c ` A = {0}" by auto with True show ?thesis by auto next case [simp]: False from assms obtain x where x: "x \ A" by blast have "c dvd c * x" by simp also from x have "c * x dvd Lcm ((*) c ` A)" by (intro dvd_Lcm) auto finally have dvd: "c dvd Lcm ((*) c ` A)" . moreover have "Lcm A dvd Lcm ((*) c ` A) div c" by (intro Lcm_least dvd_mult_imp_div) (auto intro!: Lcm_least dvd_Lcm simp: mult.commute[of _ c]) ultimately have "c * Lcm A dvd Lcm ((*) c ` A)" by auto - also have "c * Lcm A = (normalize c * Lcm A) * unit_factor c" - by (subst unit_factor_mult_normalize [symmetric]) (simp only: mult_ac) - also have "\ dvd Lcm ((*) c ` A) \ normalize c * Lcm A dvd Lcm ((*) c ` A)" - by (simp add: mult_unit_dvd_iff) - finally have "normalize c * Lcm A dvd Lcm ((*) c ` A)" . - moreover have "Lcm ((*) c ` A) dvd normalize c * Lcm A" + moreover have "Lcm ((*) c ` A) dvd c * Lcm A" by (intro Lcm_least) (auto intro: mult_dvd_mono dvd_Lcm) - ultimately have "normalize (normalize c * Lcm A) = normalize (Lcm ((*) c ` A))" + ultimately have "normalize (c * Lcm A) = normalize (Lcm ((*) c ` A))" by (rule associatedI) - then show ?thesis - by (simp add: normalize_mult) + then show ?thesis by simp qed lemma Lcm_no_units: "Lcm A = Lcm (A - {a. is_unit a})" proof - have "(A - {a. is_unit a}) \ {a\A. is_unit a} = A" by blast then have "Lcm A = lcm (Lcm (A - {a. is_unit a})) (Lcm {a\A. is_unit a})" by (simp add: Lcm_Un [symmetric]) also have "Lcm {a\A. is_unit a} = 1" by simp finally show ?thesis by simp qed lemma Lcm_0_iff': "Lcm A = 0 \ (\l. l \ 0 \ (\a\A. a dvd l))" by (metis Lcm_least dvd_0_left dvd_Lcm) lemma Lcm_no_multiple: "(\m. m \ 0 \ (\a\A. \ a dvd m)) \ Lcm A = 0" by (auto simp: Lcm_0_iff') lemma Lcm_singleton [simp]: "Lcm {a} = normalize a" by simp lemma Lcm_2 [simp]: "Lcm {a, b} = lcm a b" by simp lemma Gcd_1: "1 \ A \ Gcd A = 1" by (auto intro!: Gcd_eq_1_I) lemma Gcd_singleton [simp]: "Gcd {a} = normalize a" by simp lemma Gcd_2 [simp]: "Gcd {a, b} = gcd a b" by simp lemma Gcd_mono: assumes "\x. x \ A \ f x dvd g x" shows "(GCD x\A. f x) dvd (GCD x\A. g x)" proof (intro Gcd_greatest, safe) fix x assume "x \ A" hence "(GCD x\A. f x) dvd f x" by (intro Gcd_dvd) auto also have "f x dvd g x" using \x \ A\ assms by blast finally show "(GCD x\A. f x) dvd \" . qed lemma Lcm_mono: assumes "\x. x \ A \ f x dvd g x" shows "(LCM x\A. f x) dvd (LCM x\A. g x)" proof (intro Lcm_least, safe) fix x assume "x \ A" hence "f x dvd g x" by (rule assms) also have "g x dvd (LCM x\A. g x)" using \x \ A\ by (intro dvd_Lcm) auto finally show "f x dvd \" . qed end subsection \An aside: GCD and LCM on finite sets for incomplete gcd rings\ context semiring_gcd begin sublocale Gcd_fin: bounded_quasi_semilattice_set gcd 0 1 normalize defines Gcd_fin ("Gcd\<^sub>f\<^sub>i\<^sub>n") = "Gcd_fin.F :: 'a set \ 'a" .. abbreviation gcd_list :: "'a list \ 'a" where "gcd_list xs \ Gcd\<^sub>f\<^sub>i\<^sub>n (set xs)" sublocale Lcm_fin: bounded_quasi_semilattice_set lcm 1 0 normalize defines Lcm_fin ("Lcm\<^sub>f\<^sub>i\<^sub>n") = Lcm_fin.F .. abbreviation lcm_list :: "'a list \ 'a" where "lcm_list xs \ Lcm\<^sub>f\<^sub>i\<^sub>n (set xs)" lemma Gcd_fin_dvd: "a \ A \ Gcd\<^sub>f\<^sub>i\<^sub>n A dvd a" by (induct A rule: infinite_finite_induct) (auto intro: dvd_trans) lemma dvd_Lcm_fin: "a \ A \ a dvd Lcm\<^sub>f\<^sub>i\<^sub>n A" by (induct A rule: infinite_finite_induct) (auto intro: dvd_trans) lemma Gcd_fin_greatest: "a dvd Gcd\<^sub>f\<^sub>i\<^sub>n A" if "finite A" and "\b. b \ A \ a dvd b" using that by (induct A) simp_all lemma Lcm_fin_least: "Lcm\<^sub>f\<^sub>i\<^sub>n A dvd a" if "finite A" and "\b. b \ A \ b dvd a" using that by (induct A) simp_all lemma gcd_list_greatest: "a dvd gcd_list bs" if "\b. b \ set bs \ a dvd b" by (rule Gcd_fin_greatest) (simp_all add: that) lemma lcm_list_least: "lcm_list bs dvd a" if "\b. b \ set bs \ b dvd a" by (rule Lcm_fin_least) (simp_all add: that) lemma dvd_Gcd_fin_iff: "b dvd Gcd\<^sub>f\<^sub>i\<^sub>n A \ (\a\A. b dvd a)" if "finite A" using that by (auto intro: Gcd_fin_greatest Gcd_fin_dvd dvd_trans [of b "Gcd\<^sub>f\<^sub>i\<^sub>n A"]) lemma dvd_gcd_list_iff: "b dvd gcd_list xs \ (\a\set xs. b dvd a)" by (simp add: dvd_Gcd_fin_iff) lemma Lcm_fin_dvd_iff: "Lcm\<^sub>f\<^sub>i\<^sub>n A dvd b \ (\a\A. a dvd b)" if "finite A" using that by (auto intro: Lcm_fin_least dvd_Lcm_fin dvd_trans [of _ "Lcm\<^sub>f\<^sub>i\<^sub>n A" b]) lemma lcm_list_dvd_iff: "lcm_list xs dvd b \ (\a\set xs. a dvd b)" by (simp add: Lcm_fin_dvd_iff) lemma Gcd_fin_mult: - "Gcd\<^sub>f\<^sub>i\<^sub>n (image (times b) A) = normalize b * Gcd\<^sub>f\<^sub>i\<^sub>n A" if "finite A" -using that proof induct - case empty - then show ?case - by simp -next - case (insert a A) - have "gcd (b * a) (b * Gcd\<^sub>f\<^sub>i\<^sub>n A) = gcd (b * a) (normalize (b * Gcd\<^sub>f\<^sub>i\<^sub>n A))" - by simp - also have "\ = gcd (b * a) (normalize b * Gcd\<^sub>f\<^sub>i\<^sub>n A)" - by (simp add: normalize_mult) - finally show ?case - using insert by (simp add: gcd_mult_distrib') -qed + "Gcd\<^sub>f\<^sub>i\<^sub>n (image (times b) A) = normalize (b * Gcd\<^sub>f\<^sub>i\<^sub>n A)" if "finite A" + using that by induction (auto simp: gcd_mult_left) lemma Lcm_fin_mult: - "Lcm\<^sub>f\<^sub>i\<^sub>n (image (times b) A) = normalize b * Lcm\<^sub>f\<^sub>i\<^sub>n A" if "A \ {}" + "Lcm\<^sub>f\<^sub>i\<^sub>n (image (times b) A) = normalize (b * Lcm\<^sub>f\<^sub>i\<^sub>n A)" if "A \ {}" proof (cases "b = 0") case True moreover from that have "times 0 ` A = {0}" by auto ultimately show ?thesis by simp next case False show ?thesis proof (cases "finite A") case False moreover have "inj_on (times b) A" using \b \ 0\ by (rule inj_on_mult) ultimately have "infinite (times b ` A)" by (simp add: finite_image_iff) with False show ?thesis by simp next case True - then show ?thesis using that proof (induct A rule: finite_ne_induct) - case (singleton a) - then show ?case - by (simp add: normalize_mult) - next - case (insert a A) - have "lcm (b * a) (b * Lcm\<^sub>f\<^sub>i\<^sub>n A) = lcm (b * a) (normalize (b * Lcm\<^sub>f\<^sub>i\<^sub>n A))" - by simp - also have "\ = lcm (b * a) (normalize b * Lcm\<^sub>f\<^sub>i\<^sub>n A)" - by (simp add: normalize_mult) - finally show ?case - using insert by (simp add: lcm_mult_distrib') - qed + then show ?thesis using that + by (induct A rule: finite_ne_induct) (auto simp: lcm_mult_left) qed qed lemma unit_factor_Gcd_fin: "unit_factor (Gcd\<^sub>f\<^sub>i\<^sub>n A) = of_bool (Gcd\<^sub>f\<^sub>i\<^sub>n A \ 0)" by (rule normalize_idem_imp_unit_factor_eq) simp lemma unit_factor_Lcm_fin: "unit_factor (Lcm\<^sub>f\<^sub>i\<^sub>n A) = of_bool (Lcm\<^sub>f\<^sub>i\<^sub>n A \ 0)" by (rule normalize_idem_imp_unit_factor_eq) simp lemma is_unit_Gcd_fin_iff [simp]: "is_unit (Gcd\<^sub>f\<^sub>i\<^sub>n A) \ Gcd\<^sub>f\<^sub>i\<^sub>n A = 1" by (rule normalize_idem_imp_is_unit_iff) simp lemma is_unit_Lcm_fin_iff [simp]: "is_unit (Lcm\<^sub>f\<^sub>i\<^sub>n A) \ Lcm\<^sub>f\<^sub>i\<^sub>n A = 1" by (rule normalize_idem_imp_is_unit_iff) simp lemma Gcd_fin_0_iff: "Gcd\<^sub>f\<^sub>i\<^sub>n A = 0 \ A \ {0} \ finite A" by (induct A rule: infinite_finite_induct) simp_all lemma Lcm_fin_0_iff: "Lcm\<^sub>f\<^sub>i\<^sub>n A = 0 \ 0 \ A" if "finite A" using that by (induct A) (auto simp: lcm_eq_0_iff) lemma Lcm_fin_1_iff: "Lcm\<^sub>f\<^sub>i\<^sub>n A = 1 \ (\a\A. is_unit a) \ finite A" by (induct A rule: infinite_finite_induct) simp_all end context semiring_Gcd begin lemma Gcd_fin_eq_Gcd [simp]: "Gcd\<^sub>f\<^sub>i\<^sub>n A = Gcd A" if "finite A" for A :: "'a set" using that by induct simp_all lemma Gcd_set_eq_fold [code_unfold]: "Gcd (set xs) = fold gcd xs 0" by (simp add: Gcd_fin.set_eq_fold [symmetric]) lemma Lcm_fin_eq_Lcm [simp]: "Lcm\<^sub>f\<^sub>i\<^sub>n A = Lcm A" if "finite A" for A :: "'a set" using that by induct simp_all lemma Lcm_set_eq_fold [code_unfold]: "Lcm (set xs) = fold lcm xs 1" by (simp add: Lcm_fin.set_eq_fold [symmetric]) end subsection \Coprimality\ context semiring_gcd begin lemma coprime_imp_gcd_eq_1 [simp]: "gcd a b = 1" if "coprime a b" proof - define t r s where "t = gcd a b" and "r = a div t" and "s = b div t" then have "a = t * r" and "b = t * s" by simp_all with that have "coprime (t * r) (t * s)" by simp then show ?thesis by (simp add: t_def) qed lemma gcd_eq_1_imp_coprime [dest!]: "coprime a b" if "gcd a b = 1" proof (rule coprimeI) fix c assume "c dvd a" and "c dvd b" then have "c dvd gcd a b" by (rule gcd_greatest) with that show "is_unit c" by simp qed lemma coprime_iff_gcd_eq_1 [presburger, code]: "coprime a b \ gcd a b = 1" by rule (simp_all add: gcd_eq_1_imp_coprime) lemma is_unit_gcd [simp]: "is_unit (gcd a b) \ coprime a b" by (simp add: coprime_iff_gcd_eq_1) lemma coprime_add_one_left [simp]: "coprime (a + 1) a" by (simp add: gcd_eq_1_imp_coprime ac_simps) lemma coprime_add_one_right [simp]: "coprime a (a + 1)" using coprime_add_one_left [of a] by (simp add: ac_simps) lemma coprime_mult_left_iff [simp]: "coprime (a * b) c \ coprime a c \ coprime b c" proof assume "coprime (a * b) c" with coprime_common_divisor [of "a * b" c] have *: "is_unit d" if "d dvd a * b" and "d dvd c" for d using that by blast have "coprime a c" by (rule coprimeI, rule *) simp_all moreover have "coprime b c" by (rule coprimeI, rule *) simp_all ultimately show "coprime a c \ coprime b c" .. next assume "coprime a c \ coprime b c" then have "coprime a c" "coprime b c" by simp_all show "coprime (a * b) c" proof (rule coprimeI) fix d assume "d dvd a * b" then obtain r s where d: "d = r * s" "r dvd a" "s dvd b" by (rule dvd_productE) assume "d dvd c" with d have "r * s dvd c" by simp then have "r dvd c" "s dvd c" by (auto intro: dvd_mult_left dvd_mult_right) from \coprime a c\ \r dvd a\ \r dvd c\ have "is_unit r" by (rule coprime_common_divisor) moreover from \coprime b c\ \s dvd b\ \s dvd c\ have "is_unit s" by (rule coprime_common_divisor) ultimately show "is_unit d" by (simp add: d is_unit_mult_iff) qed qed lemma coprime_mult_right_iff [simp]: "coprime c (a * b) \ coprime c a \ coprime c b" using coprime_mult_left_iff [of a b c] by (simp add: ac_simps) lemma coprime_power_left_iff [simp]: "coprime (a ^ n) b \ coprime a b \ n = 0" proof (cases "n = 0") case True then show ?thesis by simp next case False then have "n > 0" by simp then show ?thesis by (induction n rule: nat_induct_non_zero) simp_all qed lemma coprime_power_right_iff [simp]: "coprime a (b ^ n) \ coprime a b \ n = 0" using coprime_power_left_iff [of b n a] by (simp add: ac_simps) lemma prod_coprime_left: "coprime (\i\A. f i) a" if "\i. i \ A \ coprime (f i) a" using that by (induct A rule: infinite_finite_induct) simp_all lemma prod_coprime_right: "coprime a (\i\A. f i)" if "\i. i \ A \ coprime a (f i)" using that prod_coprime_left [of A f a] by (simp add: ac_simps) lemma prod_list_coprime_left: "coprime (prod_list xs) a" if "\x. x \ set xs \ coprime x a" using that by (induct xs) simp_all lemma prod_list_coprime_right: "coprime a (prod_list xs)" if "\x. x \ set xs \ coprime a x" using that prod_list_coprime_left [of xs a] by (simp add: ac_simps) lemma coprime_dvd_mult_left_iff: "a dvd b * c \ a dvd b" if "coprime a c" proof assume "a dvd b" then show "a dvd b * c" by simp next assume "a dvd b * c" show "a dvd b" proof (cases "b = 0") case True then show ?thesis by simp next case False then have unit: "is_unit (unit_factor b)" by simp - from \coprime a c\ mult_gcd_left [of b a c] + from \coprime a c\ have "gcd (b * a) (b * c) * unit_factor b = b" - by (simp add: ac_simps) + by (subst gcd_mult_left) (simp add: ac_simps) moreover from \a dvd b * c\ have "a dvd gcd (b * a) (b * c) * unit_factor b" by (simp add: dvd_mult_unit_iff unit) ultimately show ?thesis by simp qed qed lemma coprime_dvd_mult_right_iff: "a dvd c * b \ a dvd b" if "coprime a c" using that coprime_dvd_mult_left_iff [of a c b] by (simp add: ac_simps) lemma divides_mult: "a * b dvd c" if "a dvd c" and "b dvd c" and "coprime a b" proof - from \b dvd c\ obtain b' where "c = b * b'" .. with \a dvd c\ have "a dvd b' * b" by (simp add: ac_simps) with \coprime a b\ have "a dvd b'" by (simp add: coprime_dvd_mult_left_iff) then obtain a' where "b' = a * a'" .. with \c = b * b'\ have "c = (a * b) * a'" by (simp add: ac_simps) then show ?thesis .. qed lemma div_gcd_coprime: assumes "a \ 0 \ b \ 0" shows "coprime (a div gcd a b) (b div gcd a b)" proof - let ?g = "gcd a b" let ?a' = "a div ?g" let ?b' = "b div ?g" let ?g' = "gcd ?a' ?b'" have dvdg: "?g dvd a" "?g dvd b" by simp_all have dvdg': "?g' dvd ?a'" "?g' dvd ?b'" by simp_all from dvdg dvdg' obtain ka kb ka' kb' where kab: "a = ?g * ka" "b = ?g * kb" "?a' = ?g' * ka'" "?b' = ?g' * kb'" unfolding dvd_def by blast from this [symmetric] have "?g * ?a' = (?g * ?g') * ka'" "?g * ?b' = (?g * ?g') * kb'" by (simp_all add: mult.assoc mult.left_commute [of "gcd a b"]) then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b" by (auto simp: dvd_mult_div_cancel [OF dvdg(1)] dvd_mult_div_cancel [OF dvdg(2)] dvd_def) have "?g \ 0" using assms by simp moreover from gcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" . ultimately show ?thesis using dvd_times_left_cancel_iff [of "gcd a b" _ 1] by simp (simp only: coprime_iff_gcd_eq_1) qed lemma gcd_coprime: assumes c: "gcd a b \ 0" and a: "a = a' * gcd a b" and b: "b = b' * gcd a b" shows "coprime a' b'" proof - from c have "a \ 0 \ b \ 0" by simp with div_gcd_coprime have "coprime (a div gcd a b) (b div gcd a b)" . also from assms have "a div gcd a b = a'" - using dvd_div_eq_mult local.gcd_dvd1 by blast + using dvd_div_eq_mult gcd_dvd1 by blast also from assms have "b div gcd a b = b'" - using dvd_div_eq_mult local.gcd_dvd1 by blast + using dvd_div_eq_mult gcd_dvd1 by blast finally show ?thesis . qed lemma gcd_coprime_exists: assumes "gcd a b \ 0" shows "\a' b'. a = a' * gcd a b \ b = b' * gcd a b \ coprime a' b'" proof - have "coprime (a div gcd a b) (b div gcd a b)" using assms div_gcd_coprime by auto then show ?thesis by force qed lemma pow_divides_pow_iff [simp]: "a ^ n dvd b ^ n \ a dvd b" if "n > 0" proof (cases "gcd a b = 0") case True then show ?thesis by simp next case False show ?thesis proof let ?d = "gcd a b" from \n > 0\ obtain m where m: "n = Suc m" by (cases n) simp_all from False have zn: "?d ^ n \ 0" by (rule power_not_zero) from gcd_coprime_exists [OF False] obtain a' b' where ab': "a = a' * ?d" "b = b' * ?d" "coprime a' b'" by blast assume "a ^ n dvd b ^ n" then have "(a' * ?d) ^ n dvd (b' * ?d) ^ n" by (simp add: ab'(1,2)[symmetric]) then have "?d^n * a'^n dvd ?d^n * b'^n" by (simp only: power_mult_distrib ac_simps) with zn have "a' ^ n dvd b' ^ n" by simp then have "a' dvd b' ^ n" using dvd_trans[of a' "a'^n" "b'^n"] by (simp add: m) then have "a' dvd b' ^ m * b'" by (simp add: m ac_simps) moreover have "coprime a' (b' ^ n)" using \coprime a' b'\ by simp then have "a' dvd b'" using \a' dvd b' ^ n\ coprime_dvd_mult_left_iff dvd_mult by blast then have "a' * ?d dvd b' * ?d" by (rule mult_dvd_mono) simp with ab'(1,2) show "a dvd b" by simp next assume "a dvd b" with \n > 0\ show "a ^ n dvd b ^ n" by (induction rule: nat_induct_non_zero) (simp_all add: mult_dvd_mono) qed qed lemma coprime_crossproduct: fixes a b c d :: 'a assumes "coprime a d" and "coprime b c" shows "normalize a * normalize c = normalize b * normalize d \ normalize a = normalize b \ normalize c = normalize d" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by simp next assume ?lhs from \?lhs\ have "normalize a dvd normalize b * normalize d" by (auto intro: dvdI dest: sym) with \coprime a d\ have "a dvd b" by (simp add: coprime_dvd_mult_left_iff normalize_mult [symmetric]) from \?lhs\ have "normalize b dvd normalize a * normalize c" by (auto intro: dvdI dest: sym) with \coprime b c\ have "b dvd a" by (simp add: coprime_dvd_mult_left_iff normalize_mult [symmetric]) from \?lhs\ have "normalize c dvd normalize d * normalize b" by (auto intro: dvdI dest: sym simp add: mult.commute) with \coprime b c\ have "c dvd d" by (simp add: coprime_dvd_mult_left_iff coprime_commute normalize_mult [symmetric]) from \?lhs\ have "normalize d dvd normalize c * normalize a" by (auto intro: dvdI dest: sym simp add: mult.commute) with \coprime a d\ have "d dvd c" by (simp add: coprime_dvd_mult_left_iff coprime_commute normalize_mult [symmetric]) from \a dvd b\ \b dvd a\ have "normalize a = normalize b" by (rule associatedI) moreover from \c dvd d\ \d dvd c\ have "normalize c = normalize d" by (rule associatedI) ultimately show ?rhs .. qed -lemma coprime_crossproduct': - fixes a b c d - assumes "b \ 0" - assumes unit_factors: "unit_factor b = unit_factor d" - assumes coprime: "coprime a b" "coprime c d" - shows "a * d = b * c \ a = c \ b = d" -proof safe - assume eq: "a * d = b * c" - hence "normalize a * normalize d = normalize c * normalize b" - by (simp only: normalize_mult [symmetric] mult_ac) - with coprime have "normalize b = normalize d" - by (subst (asm) coprime_crossproduct) simp_all - from this and unit_factors show "b = d" - by (rule normalize_unit_factor_eqI) - from eq have "a * d = c * d" by (simp only: \b = d\ mult_ac) - with \b \ 0\ \b = d\ show "a = c" by simp -qed (simp_all add: mult_ac) - lemma gcd_mult_left_left_cancel: "gcd (c * a) b = gcd a b" if "coprime b c" proof - have "coprime (gcd b (a * c)) c" by (rule coprimeI) (auto intro: that coprime_common_divisor) then have "gcd b (a * c) dvd a" using coprime_dvd_mult_left_iff [of "gcd b (a * c)" c a] by simp then show ?thesis by (auto intro: associated_eqI simp add: ac_simps) qed lemma gcd_mult_left_right_cancel: "gcd (a * c) b = gcd a b" if "coprime b c" using that gcd_mult_left_left_cancel [of b c a] by (simp add: ac_simps) lemma gcd_mult_right_left_cancel: "gcd a (c * b) = gcd a b" if "coprime a c" using that gcd_mult_left_left_cancel [of a c b] by (simp add: ac_simps) lemma gcd_mult_right_right_cancel: "gcd a (b * c) = gcd a b" if "coprime a c" using that gcd_mult_right_left_cancel [of a c b] by (simp add: ac_simps) -lemma gcd_exp [simp]: - "gcd (a ^ n) (b ^ n) = gcd a b ^ n" +lemma gcd_exp_weak: + "gcd (a ^ n) (b ^ n) = normalize (gcd a b ^ n)" proof (cases "a = 0 \ b = 0 \ n = 0") case True then show ?thesis by (cases n) simp_all next case False then have "coprime (a div gcd a b) (b div gcd a b)" and "n > 0" by (auto intro: div_gcd_coprime) then have "coprime ((a div gcd a b) ^ n) ((b div gcd a b) ^ n)" by simp then have "1 = gcd ((a div gcd a b) ^ n) ((b div gcd a b) ^ n)" by simp - then have "gcd a b ^ n = gcd a b ^ n * \" + then have "normalize (gcd a b ^ n) = normalize (gcd a b ^ n * \)" by simp - also note gcd_mult_distrib - also have "unit_factor (gcd a b ^ n) = 1" - using False by (auto simp: unit_factor_power unit_factor_gcd) + also have "\ = gcd (gcd a b ^ n * (a div gcd a b) ^ n) (gcd a b ^ n * (b div gcd a b) ^ n)" + by (rule gcd_mult_left [symmetric]) also have "(gcd a b) ^ n * (a div gcd a b) ^ n = a ^ n" by (simp add: ac_simps div_power dvd_power_same) also have "(gcd a b) ^ n * (b div gcd a b) ^ n = b ^ n" by (simp add: ac_simps div_power dvd_power_same) finally show ?thesis by simp qed lemma division_decomp: assumes "a dvd b * c" shows "\b' c'. a = b' * c' \ b' dvd b \ c' dvd c" proof (cases "gcd a b = 0") case True then have "a = 0 \ b = 0" by simp then have "a = 0 * c \ 0 dvd b \ c dvd c" by simp then show ?thesis by blast next case False let ?d = "gcd a b" from gcd_coprime_exists [OF False] obtain a' b' where ab': "a = a' * ?d" "b = b' * ?d" "coprime a' b'" by blast from ab'(1) have "a' dvd a" .. with assms have "a' dvd b * c" using dvd_trans [of a' a "b * c"] by simp from assms ab'(1,2) have "a' * ?d dvd (b' * ?d) * c" by simp then have "?d * a' dvd ?d * (b' * c)" by (simp add: mult_ac) with \?d \ 0\ have "a' dvd b' * c" by simp then have "a' dvd c" using \coprime a' b'\ by (simp add: coprime_dvd_mult_right_iff) with ab'(1) have "a = ?d * a' \ ?d dvd b \ a' dvd c" by (simp add: ac_simps) then show ?thesis by blast qed lemma lcm_coprime: "coprime a b \ lcm a b = normalize (a * b)" by (subst lcm_gcd) simp end context ring_gcd begin lemma coprime_minus_left_iff [simp]: "coprime (- a) b \ coprime a b" by (rule; rule coprimeI) (auto intro: coprime_common_divisor) lemma coprime_minus_right_iff [simp]: "coprime a (- b) \ coprime a b" using coprime_minus_left_iff [of b a] by (simp add: ac_simps) lemma coprime_diff_one_left [simp]: "coprime (a - 1) a" using coprime_add_one_right [of "a - 1"] by simp lemma coprime_doff_one_right [simp]: "coprime a (a - 1)" using coprime_diff_one_left [of a] by (simp add: ac_simps) end context semiring_Gcd begin lemma Lcm_coprime: assumes "finite A" and "A \ {}" and "\a b. a \ A \ b \ A \ a \ b \ coprime a b" shows "Lcm A = normalize (\A)" using assms proof (induct rule: finite_ne_induct) case singleton then show ?case by simp next case (insert a A) have "Lcm (insert a A) = lcm a (Lcm A)" by simp also from insert have "Lcm A = normalize (\A)" by blast also have "lcm a \ = lcm a (\A)" by (cases "\A = 0") (simp_all add: lcm_div_unit2) also from insert have "coprime a (\A)" by (subst coprime_commute, intro prod_coprime_left) auto with insert have "lcm a (\A) = normalize (\(insert a A))" by (simp add: lcm_coprime) finally show ?case . qed lemma Lcm_coprime': "card A \ 0 \ (\a b. a \ A \ b \ A \ a \ b \ coprime a b) \ Lcm A = normalize (\A)" by (rule Lcm_coprime) (simp_all add: card_eq_0_iff) end +subsection \GCD and LCM for multiplicative normalisation functions\ + +class semiring_gcd_mult_normalize = semiring_gcd + normalization_semidom_multiplicative +begin + +lemma mult_gcd_left: "c * gcd a b = unit_factor c * gcd (c * a) (c * b)" + by (simp add: gcd_mult_left normalize_mult mult.assoc [symmetric]) + +lemma mult_gcd_right: "gcd a b * c = gcd (a * c) (b * c) * unit_factor c" + using mult_gcd_left [of c a b] by (simp add: ac_simps) + +lemma gcd_mult_distrib': "normalize c * gcd a b = gcd (c * a) (c * b)" + by (subst gcd_mult_left) (simp_all add: normalize_mult) + +lemma gcd_mult_distrib: "k * gcd a b = gcd (k * a) (k * b) * unit_factor k" +proof- + have "normalize k * gcd a b = gcd (k * a) (k * b)" + by (simp add: gcd_mult_distrib') + then have "normalize k * gcd a b * unit_factor k = gcd (k * a) (k * b) * unit_factor k" + by simp + then have "normalize k * unit_factor k * gcd a b = gcd (k * a) (k * b) * unit_factor k" + by (simp only: ac_simps) + then show ?thesis + by simp +qed + +lemma gcd_mult_lcm [simp]: "gcd a b * lcm a b = normalize a * normalize b" + by (simp add: lcm_gcd normalize_mult dvd_normalize_div) + +lemma lcm_mult_gcd [simp]: "lcm a b * gcd a b = normalize a * normalize b" + using gcd_mult_lcm [of a b] by (simp add: ac_simps) + +lemma mult_lcm_left: "c * lcm a b = unit_factor c * lcm (c * a) (c * b)" + by (simp add: lcm_mult_left mult.assoc [symmetric] normalize_mult) + +lemma mult_lcm_right: "lcm a b * c = lcm (a * c) (b * c) * unit_factor c" + using mult_lcm_left [of c a b] by (simp add: ac_simps) + +lemma lcm_gcd_prod: "lcm a b * gcd a b = normalize (a * b)" + by (simp add: lcm_gcd dvd_normalize_div normalize_mult) + +lemma lcm_mult_distrib': "normalize c * lcm a b = lcm (c * a) (c * b)" + by (subst lcm_mult_left) (simp add: normalize_mult) + +lemma lcm_mult_distrib: "k * lcm a b = lcm (k * a) (k * b) * unit_factor k" +proof- + have "normalize k * lcm a b = lcm (k * a) (k * b)" + by (simp add: lcm_mult_distrib') + then have "normalize k * lcm a b * unit_factor k = lcm (k * a) (k * b) * unit_factor k" + by simp + then have "normalize k * unit_factor k * lcm a b = lcm (k * a) (k * b) * unit_factor k" + by (simp only: ac_simps) + then show ?thesis + by simp +qed + +lemma coprime_crossproduct': + fixes a b c d + assumes "b \ 0" + assumes unit_factors: "unit_factor b = unit_factor d" + assumes coprime: "coprime a b" "coprime c d" + shows "a * d = b * c \ a = c \ b = d" +proof safe + assume eq: "a * d = b * c" + hence "normalize a * normalize d = normalize c * normalize b" + by (simp only: normalize_mult [symmetric] mult_ac) + with coprime have "normalize b = normalize d" + by (subst (asm) coprime_crossproduct) simp_all + from this and unit_factors show "b = d" + by (rule normalize_unit_factor_eqI) + from eq have "a * d = c * d" by (simp only: \b = d\ mult_ac) + with \b \ 0\ \b = d\ show "a = c" by simp +qed (simp_all add: mult_ac) + +lemma gcd_exp [simp]: + "gcd (a ^ n) (b ^ n) = gcd a b ^ n" + using gcd_exp_weak[of a n b] by (simp add: normalize_power) + +end + + subsection \GCD and LCM on \<^typ>\nat\ and \<^typ>\int\\ instantiation nat :: gcd begin fun gcd_nat :: "nat \ nat \ nat" where "gcd_nat x y = (if y = 0 then x else gcd y (x mod y))" definition lcm_nat :: "nat \ nat \ nat" where "lcm_nat x y = x * y div (gcd x y)" instance .. end instantiation int :: gcd begin definition gcd_int :: "int \ int \ int" where "gcd_int x y = int (gcd (nat \x\) (nat \y\))" definition lcm_int :: "int \ int \ int" where "lcm_int x y = int (lcm (nat \x\) (nat \y\))" instance .. end lemma gcd_int_int_eq [simp]: "gcd (int m) (int n) = int (gcd m n)" by (simp add: gcd_int_def) lemma gcd_nat_abs_left_eq [simp]: "gcd (nat \k\) n = nat (gcd k (int n))" by (simp add: gcd_int_def) lemma gcd_nat_abs_right_eq [simp]: "gcd n (nat \k\) = nat (gcd (int n) k)" by (simp add: gcd_int_def) lemma abs_gcd_int [simp]: "\gcd x y\ = gcd x y" for x y :: int by (simp only: gcd_int_def) lemma gcd_abs1_int [simp]: "gcd \x\ y = gcd x y" for x y :: int by (simp only: gcd_int_def) simp lemma gcd_abs2_int [simp]: "gcd x \y\ = gcd x y" for x y :: int by (simp only: gcd_int_def) simp lemma lcm_int_int_eq [simp]: "lcm (int m) (int n) = int (lcm m n)" by (simp add: lcm_int_def) lemma lcm_nat_abs_left_eq [simp]: "lcm (nat \k\) n = nat (lcm k (int n))" by (simp add: lcm_int_def) lemma lcm_nat_abs_right_eq [simp]: "lcm n (nat \k\) = nat (lcm (int n) k)" by (simp add: lcm_int_def) lemma lcm_abs1_int [simp]: "lcm \x\ y = lcm x y" for x y :: int by (simp only: lcm_int_def) simp lemma lcm_abs2_int [simp]: "lcm x \y\ = lcm x y" for x y :: int by (simp only: lcm_int_def) simp lemma abs_lcm_int [simp]: "\lcm i j\ = lcm i j" for i j :: int by (simp only: lcm_int_def) lemma gcd_nat_induct [case_names base step]: fixes m n :: nat assumes "\m. P m 0" and "\m n. 0 < n \ P n (m mod n) \ P m n" shows "P m n" proof (induction m n rule: gcd_nat.induct) case (1 x y) then show ?case using assms neq0_conv by blast qed lemma gcd_neg1_int [simp]: "gcd (- x) y = gcd x y" for x y :: int by (simp only: gcd_int_def) simp lemma gcd_neg2_int [simp]: "gcd x (- y) = gcd x y" for x y :: int by (simp only: gcd_int_def) simp lemma gcd_cases_int: fixes x y :: int assumes "x \ 0 \ y \ 0 \ P (gcd x y)" and "x \ 0 \ y \ 0 \ P (gcd x (- y))" and "x \ 0 \ y \ 0 \ P (gcd (- x) y)" and "x \ 0 \ y \ 0 \ P (gcd (- x) (- y))" shows "P (gcd x y)" using assms by auto arith lemma gcd_ge_0_int [simp]: "gcd (x::int) y >= 0" for x y :: int by (simp add: gcd_int_def) lemma lcm_neg1_int: "lcm (- x) y = lcm x y" for x y :: int by (simp only: lcm_int_def) simp lemma lcm_neg2_int: "lcm x (- y) = lcm x y" for x y :: int by (simp only: lcm_int_def) simp lemma lcm_cases_int: fixes x y :: int assumes "x \ 0 \ y \ 0 \ P (lcm x y)" and "x \ 0 \ y \ 0 \ P (lcm x (- y))" and "x \ 0 \ y \ 0 \ P (lcm (- x) y)" and "x \ 0 \ y \ 0 \ P (lcm (- x) (- y))" shows "P (lcm x y)" using assms by (auto simp: lcm_neg1_int lcm_neg2_int) arith lemma lcm_ge_0_int [simp]: "lcm x y \ 0" for x y :: int by (simp only: lcm_int_def) lemma gcd_0_nat: "gcd x 0 = x" for x :: nat by simp lemma gcd_0_int [simp]: "gcd x 0 = \x\" for x :: int by (auto simp: gcd_int_def) lemma gcd_0_left_nat: "gcd 0 x = x" for x :: nat by simp lemma gcd_0_left_int [simp]: "gcd 0 x = \x\" for x :: int by (auto simp: gcd_int_def) lemma gcd_red_nat: "gcd x y = gcd y (x mod y)" for x y :: nat by (cases "y = 0") auto text \Weaker, but useful for the simplifier.\ lemma gcd_non_0_nat: "y \ 0 \ gcd x y = gcd y (x mod y)" for x y :: nat by simp lemma gcd_1_nat [simp]: "gcd m 1 = 1" for m :: nat by simp lemma gcd_Suc_0 [simp]: "gcd m (Suc 0) = Suc 0" for m :: nat by simp lemma gcd_1_int [simp]: "gcd m 1 = 1" for m :: int by (simp add: gcd_int_def) lemma gcd_idem_nat: "gcd x x = x" for x :: nat by simp lemma gcd_idem_int: "gcd x x = \x\" for x :: int by (auto simp: gcd_int_def) declare gcd_nat.simps [simp del] text \ \<^medskip> \<^term>\gcd m n\ divides \m\ and \n\. The conjunctions don't seem provable separately. \ instance nat :: semiring_gcd proof fix m n :: nat show "gcd m n dvd m" and "gcd m n dvd n" proof (induct m n rule: gcd_nat_induct) case (step m n) then have "gcd n (m mod n) dvd m" by (metis dvd_mod_imp_dvd) with step show "gcd m n dvd m" by (simp add: gcd_non_0_nat) qed (simp_all add: gcd_0_nat gcd_non_0_nat) next fix m n k :: nat assume "k dvd m" and "k dvd n" then show "k dvd gcd m n" by (induct m n rule: gcd_nat_induct) (simp_all add: gcd_non_0_nat dvd_mod gcd_0_nat) qed (simp_all add: lcm_nat_def) instance int :: ring_gcd proof fix k l r :: int - show "gcd k l dvd k" "gcd k l dvd l" + show [simp]: "gcd k l dvd k" "gcd k l dvd l" using gcd_dvd1 [of "nat \k\" "nat \l\"] gcd_dvd2 [of "nat \k\" "nat \l\"] by simp_all - show "lcm k l = normalize (k * l) div gcd k l" + show "lcm k l = normalize (k * l div gcd k l)" using lcm_gcd [of "nat \k\" "nat \l\"] - by (simp add: nat_eq_iff of_nat_div abs_mult) + by (simp add: nat_eq_iff of_nat_div abs_mult abs_div) assume "r dvd k" "r dvd l" then show "r dvd gcd k l" using gcd_greatest [of "nat \r\" "nat \k\" "nat \l\"] by simp qed simp lemma gcd_le1_nat [simp]: "a \ 0 \ gcd a b \ a" for a b :: nat by (rule dvd_imp_le) auto lemma gcd_le2_nat [simp]: "b \ 0 \ gcd a b \ b" for a b :: nat by (rule dvd_imp_le) auto lemma gcd_le1_int [simp]: "a > 0 \ gcd a b \ a" for a b :: int by (rule zdvd_imp_le) auto lemma gcd_le2_int [simp]: "b > 0 \ gcd a b \ b" for a b :: int by (rule zdvd_imp_le) auto lemma gcd_pos_nat [simp]: "gcd m n > 0 \ m \ 0 \ n \ 0" for m n :: nat using gcd_eq_0_iff [of m n] by arith lemma gcd_pos_int [simp]: "gcd m n > 0 \ m \ 0 \ n \ 0" for m n :: int using gcd_eq_0_iff [of m n] gcd_ge_0_int [of m n] by arith lemma gcd_unique_nat: "d dvd a \ d dvd b \ (\e. e dvd a \ e dvd b \ e dvd d) \ d = gcd a b" for d a :: nat using gcd_unique by fastforce lemma gcd_unique_int: "d \ 0 \ d dvd a \ d dvd b \ (\e. e dvd a \ e dvd b \ e dvd d) \ d = gcd a b" for d a :: int using zdvd_antisym_nonneg by auto interpretation gcd_nat: semilattice_neutr_order gcd "0::nat" Rings.dvd "\m n. m dvd n \ m \ n" by standard (auto simp: gcd_unique_nat [symmetric] intro: dvd_antisym dvd_trans) lemma gcd_proj1_if_dvd_int [simp]: "x dvd y \ gcd x y = \x\" for x y :: int by (metis abs_dvd_iff gcd_0_left_int gcd_unique_int) lemma gcd_proj2_if_dvd_int [simp]: "y dvd x \ gcd x y = \y\" for x y :: int by (metis gcd_proj1_if_dvd_int gcd.commute) text \\<^medskip> Multiplication laws.\ lemma gcd_mult_distrib_nat: "k * gcd m n = gcd (k * m) (k * n)" for k m n :: nat \ \@{cite \page 27\ davenport92}\ -proof (induct m n rule: gcd_nat_induct) - case (step m n) - then show ?case - by (metis gcd_mult_distrib' gcd_red_nat) -qed auto + by (simp add: gcd_mult_left) lemma gcd_mult_distrib_int: "\k\ * gcd m n = gcd (k * m) (k * n)" for k m n :: int - using gcd_mult_distrib' [of k m n] by simp + by (simp add: gcd_mult_left abs_mult) text \\medskip Addition laws.\ (* TODO: add the other variations? *) lemma gcd_diff1_nat: "m \ n \ gcd (m - n) n = gcd m n" for m n :: nat by (subst gcd_add1 [symmetric]) auto lemma gcd_diff2_nat: "n \ m \ gcd (n - m) n = gcd m n" for m n :: nat by (metis gcd.commute gcd_add2 gcd_diff1_nat le_add_diff_inverse2) lemma gcd_non_0_int: fixes x y :: int assumes "y > 0" shows "gcd x y = gcd y (x mod y)" proof (cases "x mod y = 0") case False then have neg: "x mod y = y - (- x) mod y" by (simp add: zmod_zminus1_eq_if) have xy: "0 \ x mod y" by (simp add: assms) show ?thesis proof (cases "x < 0") case True have "nat (- x mod y) \ nat y" by (simp add: assms dual_order.order_iff_strict) moreover have "gcd (nat (- x)) (nat y) = gcd (nat (- x mod y)) (nat y)" using True assms gcd_non_0_nat nat_mod_distrib by auto ultimately have "gcd (nat (- x)) (nat y) = gcd (nat y) (nat (x mod y))" using assms by (simp add: neg nat_diff_distrib') (metis gcd.commute gcd_diff2_nat) with assms \0 \ x mod y\ show ?thesis by (simp add: True dual_order.order_iff_strict gcd_int_def) next case False with assms xy have "gcd (nat x) (nat y) = gcd (nat y) (nat x mod nat y)" using gcd_red_nat by blast with False assms show ?thesis by (simp add: gcd_int_def nat_mod_distrib) qed qed (use assms in auto) lemma gcd_red_int: "gcd x y = gcd y (x mod y)" for x y :: int proof (cases y "0::int" rule: linorder_cases) case less with gcd_non_0_int [of "- y" "- x"] show ?thesis by auto next case greater with gcd_non_0_int [of y x] show ?thesis by auto qed auto (* TODO: differences, and all variations of addition rules as simplification rules for nat and int *) (* TODO: add the three variations of these, and for ints? *) lemma finite_divisors_nat [simp]: (* FIXME move *) fixes m :: nat assumes "m > 0" shows "finite {d. d dvd m}" proof- from assms have "{d. d dvd m} \ {d. d \ m}" by (auto dest: dvd_imp_le) then show ?thesis using finite_Collect_le_nat by (rule finite_subset) qed lemma finite_divisors_int [simp]: fixes i :: int assumes "i \ 0" shows "finite {d. d dvd i}" proof - have "{d. \d\ \ \i\} = {- \i\..\i\}" by (auto simp: abs_if) then have "finite {d. \d\ \ \i\}" by simp from finite_subset [OF _ this] show ?thesis using assms by (simp add: dvd_imp_le_int subset_iff) qed lemma Max_divisors_self_nat [simp]: "n \ 0 \ Max {d::nat. d dvd n} = n" by (fastforce intro: antisym Max_le_iff[THEN iffD2] simp: dvd_imp_le) lemma Max_divisors_self_int [simp]: assumes "n \ 0" shows "Max {d::int. d dvd n} = \n\" proof (rule antisym) show "Max {d. d dvd n} \ \n\" using assms by (auto intro: abs_le_D1 dvd_imp_le_int intro!: Max_le_iff [THEN iffD2]) qed (simp add: assms) lemma gcd_is_Max_divisors_nat: fixes m n :: nat assumes "n > 0" shows "gcd m n = Max {d. d dvd m \ d dvd n}" proof (rule Max_eqI[THEN sym], simp_all) show "finite {d. d dvd m \ d dvd n}" by (simp add: \n > 0\) show "\y. y dvd m \ y dvd n \ y \ gcd m n" by (simp add: \n > 0\ dvd_imp_le) qed lemma gcd_is_Max_divisors_int: fixes m n :: int assumes "n \ 0" shows "gcd m n = Max {d. d dvd m \ d dvd n}" proof (rule Max_eqI[THEN sym], simp_all) show "finite {d. d dvd m \ d dvd n}" by (simp add: \n \ 0\) show "\y. y dvd m \ y dvd n \ y \ gcd m n" by (simp add: \n \ 0\ zdvd_imp_le) qed lemma gcd_code_int [code]: "gcd k l = \if l = 0 then k else gcd l (\k\ mod \l\)\" for k l :: int using gcd_red_int [of "\k\" "\l\"] by simp lemma coprime_Suc_left_nat [simp]: "coprime (Suc n) n" using coprime_add_one_left [of n] by simp lemma coprime_Suc_right_nat [simp]: "coprime n (Suc n)" using coprime_Suc_left_nat [of n] by (simp add: ac_simps) lemma coprime_diff_one_left_nat [simp]: "coprime (n - 1) n" if "n > 0" for n :: nat using that coprime_Suc_right_nat [of "n - 1"] by simp lemma coprime_diff_one_right_nat [simp]: "coprime n (n - 1)" if "n > 0" for n :: nat using that coprime_diff_one_left_nat [of n] by (simp add: ac_simps) lemma coprime_crossproduct_nat: fixes a b c d :: nat assumes "coprime a d" and "coprime b c" shows "a * c = b * d \ a = b \ c = d" using assms coprime_crossproduct [of a d b c] by simp lemma coprime_crossproduct_int: fixes a b c d :: int assumes "coprime a d" and "coprime b c" shows "\a\ * \c\ = \b\ * \d\ \ \a\ = \b\ \ \c\ = \d\" using assms coprime_crossproduct [of a d b c] by simp subsection \Bezout's theorem\ text \ Function \bezw\ returns a pair of witnesses to Bezout's theorem -- see the theorems that follow the definition. \ fun bezw :: "nat \ nat \ int * int" where "bezw x y = (if y = 0 then (1, 0) else (snd (bezw y (x mod y)), fst (bezw y (x mod y)) - snd (bezw y (x mod y)) * int(x div y)))" lemma bezw_0 [simp]: "bezw x 0 = (1, 0)" by simp lemma bezw_non_0: "y > 0 \ bezw x y = (snd (bezw y (x mod y)), fst (bezw y (x mod y)) - snd (bezw y (x mod y)) * int(x div y))" by simp declare bezw.simps [simp del] lemma bezw_aux: "int (gcd x y) = fst (bezw x y) * int x + snd (bezw x y) * int y" proof (induct x y rule: gcd_nat_induct) case (step m n) then have "fst (bezw m n) * int m + snd (bezw m n) * int n - int (gcd m n) = int m * snd (bezw n (m mod n)) - (int (m mod n) * snd (bezw n (m mod n)) + int n * (int (m div n) * snd (bezw n (m mod n))))" by (simp add: bezw_non_0 gcd_non_0_nat field_simps) also have "\ = int m * snd (bezw n (m mod n)) - (int (m mod n) + int (n * (m div n))) * snd (bezw n (m mod n))" by (simp add: distrib_right) also have "\ = 0" by (metis cancel_comm_monoid_add_class.diff_cancel mod_mult_div_eq of_nat_add) finally show ?case by simp qed auto lemma bezout_int: "\u v. u * x + v * y = gcd x y" for x y :: int proof - have aux: "x \ 0 \ y \ 0 \ \u v. u * x + v * y = gcd x y" for x y :: int apply (rule_tac x = "fst (bezw (nat x) (nat y))" in exI) apply (rule_tac x = "snd (bezw (nat x) (nat y))" in exI) by (simp add: bezw_aux gcd_int_def) consider "x \ 0" "y \ 0" | "x \ 0" "y \ 0" | "x \ 0" "y \ 0" | "x \ 0" "y \ 0" using linear by blast then show ?thesis proof cases case 1 then show ?thesis by (rule aux) next case 2 then show ?thesis using aux [of x "-y"] by (metis gcd_neg2_int mult.commute mult_minus_right neg_0_le_iff_le) next case 3 then show ?thesis using aux [of "-x" y] by (metis gcd.commute gcd_neg2_int mult.commute mult_minus_right neg_0_le_iff_le) next case 4 then show ?thesis using aux [of "-x" "-y"] by (metis diff_0 diff_ge_0_iff_ge gcd_neg1_int gcd_neg2_int mult.commute mult_minus_right) qed qed text \Versions of Bezout for \nat\, by Amine Chaieb.\ lemma Euclid_induct [case_names swap zero add]: fixes P :: "nat \ nat \ bool" assumes c: "\a b. P a b \ P b a" and z: "\a. P a 0" and add: "\a b. P a b \ P a (a + b)" shows "P a b" proof (induct "a + b" arbitrary: a b rule: less_induct) case less consider (eq) "a = b" | (lt) "a < b" "a + b - a < a + b" | "b = 0" | "b + a - b < a + b" by arith show ?case proof (cases a b rule: linorder_cases) case equal with add [rule_format, OF z [rule_format, of a]] show ?thesis by simp next case lt: less then consider "a = 0" | "a + b - a < a + b" by arith then show ?thesis proof cases case 1 with z c show ?thesis by blast next case 2 also have *: "a + b - a = a + (b - a)" using lt by arith finally have "a + (b - a) < a + b" . then have "P a (a + (b - a))" by (rule add [rule_format, OF less]) then show ?thesis by (simp add: *[symmetric]) qed next case gt: greater then consider "b = 0" | "b + a - b < a + b" by arith then show ?thesis proof cases case 1 with z c show ?thesis by blast next case 2 also have *: "b + a - b = b + (a - b)" using gt by arith finally have "b + (a - b) < a + b" . then have "P b (b + (a - b))" by (rule add [rule_format, OF less]) then have "P b a" by (simp add: *[symmetric]) with c show ?thesis by blast qed qed qed lemma bezout_lemma_nat: fixes d::nat shows "\d dvd a; d dvd b; a * x = b * y + d \ b * x = a * y + d\ \ \x y. d dvd a \ d dvd a + b \ (a * x = (a + b) * y + d \ (a + b) * x = a * y + d)" apply auto apply (metis add_mult_distrib2 left_add_mult_distrib) apply (rule_tac x=x in exI) by (metis add_mult_distrib2 mult.commute add.assoc) lemma bezout_add_nat: "\(d::nat) x y. d dvd a \ d dvd b \ (a * x = b * y + d \ b * x = a * y + d)" proof (induct a b rule: Euclid_induct) case (swap a b) then show ?case by blast next case (zero a) then show ?case by fastforce next case (add a b) then show ?case by (meson bezout_lemma_nat) qed lemma bezout1_nat: "\(d::nat) x y. d dvd a \ d dvd b \ (a * x - b * y = d \ b * x - a * y = d)" using bezout_add_nat[of a b] by (metis add_diff_cancel_left') lemma bezout_add_strong_nat: fixes a b :: nat assumes a: "a \ 0" shows "\d x y. d dvd a \ d dvd b \ a * x = b * y + d" proof - consider d x y where "d dvd a" "d dvd b" "a * x = b * y + d" | d x y where "d dvd a" "d dvd b" "b * x = a * y + d" using bezout_add_nat [of a b] by blast then show ?thesis proof cases case 1 then show ?thesis by blast next case H: 2 show ?thesis proof (cases "b = 0") case True with H show ?thesis by simp next case False then have bp: "b > 0" by simp with dvd_imp_le [OF H(2)] consider "d = b" | "d < b" by atomize_elim auto then show ?thesis proof cases case 1 with a H show ?thesis by (metis Suc_pred add.commute mult.commute mult_Suc_right neq0_conv) next case 2 show ?thesis proof (cases "x = 0") case True with a H show ?thesis by simp next case x0: False then have xp: "x > 0" by simp from \d < b\ have "d \ b - 1" by simp then have "d * b \ b * (b - 1)" by simp with xp mult_mono[of "1" "x" "d * b" "b * (b - 1)"] have dble: "d * b \ x * b * (b - 1)" using bp by simp from H(3) have "d + (b - 1) * (b * x) = d + (b - 1) * (a * y + d)" by simp then have "d + (b - 1) * a * y + (b - 1) * d = d + (b - 1) * b * x" by (simp only: mult.assoc distrib_left) then have "a * ((b - 1) * y) + d * (b - 1 + 1) = d + x * b * (b - 1)" by algebra then have "a * ((b - 1) * y) = d + x * b * (b - 1) - d * b" using bp by simp then have "a * ((b - 1) * y) = d + (x * b * (b - 1) - d * b)" by (simp only: diff_add_assoc[OF dble, of d, symmetric]) then have "a * ((b - 1) * y) = b * (x * (b - 1) - d) + d" by (simp only: diff_mult_distrib2 ac_simps) with H(1,2) show ?thesis by blast qed qed qed qed qed lemma bezout_nat: fixes a :: nat assumes a: "a \ 0" shows "\x y. a * x = b * y + gcd a b" proof - obtain d x y where d: "d dvd a" "d dvd b" and eq: "a * x = b * y + d" using bezout_add_strong_nat [OF a, of b] by blast from d have "d dvd gcd a b" by simp then obtain k where k: "gcd a b = d * k" unfolding dvd_def by blast from eq have "a * x * k = (b * y + d) * k" by auto then have "a * (x * k) = b * (y * k) + gcd a b" by (algebra add: k) then show ?thesis by blast qed subsection \LCM properties on \<^typ>\nat\ and \<^typ>\int\\ lemma lcm_altdef_int [code]: "lcm a b = \a\ * \b\ div gcd a b" for a b :: int - by (simp add: abs_mult lcm_gcd) + by (simp add: abs_mult lcm_gcd abs_div) lemma prod_gcd_lcm_nat: "m * n = gcd m n * lcm m n" for m n :: nat - by simp + by (simp add: lcm_gcd) lemma prod_gcd_lcm_int: "\m\ * \n\ = gcd m n * lcm m n" for m n :: int - by simp + by (simp add: lcm_gcd abs_div abs_mult) lemma lcm_pos_nat: "m > 0 \ n > 0 \ lcm m n > 0" for m n :: nat using lcm_eq_0_iff [of m n] by auto lemma lcm_pos_int: "m \ 0 \ n \ 0 \ lcm m n > 0" for m n :: int by (simp add: less_le lcm_eq_0_iff) lemma dvd_pos_nat: "n > 0 \ m dvd n \ m > 0" (* FIXME move *) for m n :: nat by auto lemma lcm_unique_nat: "a dvd d \ b dvd d \ (\e. a dvd e \ b dvd e \ d dvd e) \ d = lcm a b" for a b d :: nat by (auto intro: dvd_antisym lcm_least) lemma lcm_unique_int: "d \ 0 \ a dvd d \ b dvd d \ (\e. a dvd e \ b dvd e \ d dvd e) \ d = lcm a b" for a b d :: int using lcm_least zdvd_antisym_nonneg by auto lemma lcm_proj2_if_dvd_nat [simp]: "x dvd y \ lcm x y = y" for x y :: nat by (simp add: lcm_proj2_if_dvd) lemma lcm_proj2_if_dvd_int [simp]: "x dvd y \ lcm x y = \y\" for x y :: int by (simp add: lcm_proj2_if_dvd) lemma lcm_proj1_if_dvd_nat [simp]: "x dvd y \ lcm y x = y" for x y :: nat by (subst lcm.commute) (erule lcm_proj2_if_dvd_nat) lemma lcm_proj1_if_dvd_int [simp]: "x dvd y \ lcm y x = \y\" for x y :: int by (subst lcm.commute) (erule lcm_proj2_if_dvd_int) lemma lcm_proj1_iff_nat [simp]: "lcm m n = m \ n dvd m" for m n :: nat by (metis lcm_proj1_if_dvd_nat lcm_unique_nat) lemma lcm_proj2_iff_nat [simp]: "lcm m n = n \ m dvd n" for m n :: nat by (metis lcm_proj2_if_dvd_nat lcm_unique_nat) lemma lcm_proj1_iff_int [simp]: "lcm m n = \m\ \ n dvd m" for m n :: int by (metis dvd_abs_iff lcm_proj1_if_dvd_int lcm_unique_int) lemma lcm_proj2_iff_int [simp]: "lcm m n = \n\ \ m dvd n" for m n :: int by (metis dvd_abs_iff lcm_proj2_if_dvd_int lcm_unique_int) lemma lcm_1_iff_nat [simp]: "lcm m n = Suc 0 \ m = Suc 0 \ n = Suc 0" for m n :: nat using lcm_eq_1_iff [of m n] by simp lemma lcm_1_iff_int [simp]: "lcm m n = 1 \ (m = 1 \ m = -1) \ (n = 1 \ n = -1)" for m n :: int by auto subsection \The complete divisibility lattice on \<^typ>\nat\ and \<^typ>\int\\ text \ Lifting \gcd\ and \lcm\ to sets (\Gcd\ / \Lcm\). \Gcd\ is defined via \Lcm\ to facilitate the proof that we have a complete lattice. \ instantiation nat :: semiring_Gcd begin interpretation semilattice_neutr_set lcm "1::nat" by standard simp_all definition "Lcm M = (if finite M then F M else 0)" for M :: "nat set" lemma Lcm_nat_empty: "Lcm {} = (1::nat)" by (simp add: Lcm_nat_def del: One_nat_def) lemma Lcm_nat_insert: "Lcm (insert n M) = lcm n (Lcm M)" for n :: nat by (cases "finite M") (auto simp: Lcm_nat_def simp del: One_nat_def) lemma Lcm_nat_infinite: "infinite M \ Lcm M = 0" for M :: "nat set" by (simp add: Lcm_nat_def) lemma dvd_Lcm_nat [simp]: fixes M :: "nat set" assumes "m \ M" shows "m dvd Lcm M" proof - from assms have "insert m M = M" by auto moreover have "m dvd Lcm (insert m M)" by (simp add: Lcm_nat_insert) ultimately show ?thesis by simp qed lemma Lcm_dvd_nat [simp]: fixes M :: "nat set" assumes "\m\M. m dvd n" shows "Lcm M dvd n" proof (cases "n > 0") case False then show ?thesis by simp next case True then have "finite {d. d dvd n}" by (rule finite_divisors_nat) moreover have "M \ {d. d dvd n}" using assms by fast ultimately have "finite M" by (rule rev_finite_subset) then show ?thesis using assms by (induct M) (simp_all add: Lcm_nat_empty Lcm_nat_insert) qed definition "Gcd M = Lcm {d. \m\M. d dvd m}" for M :: "nat set" instance proof fix N :: "nat set" fix n :: nat show "Gcd N dvd n" if "n \ N" using that by (induct N rule: infinite_finite_induct) (auto simp: Gcd_nat_def) show "n dvd Gcd N" if "\m. m \ N \ n dvd m" using that by (induct N rule: infinite_finite_induct) (auto simp: Gcd_nat_def) show "n dvd Lcm N" if "n \ N" using that by (induct N rule: infinite_finite_induct) auto show "Lcm N dvd n" if "\m. m \ N \ m dvd n" using that by (induct N rule: infinite_finite_induct) auto show "normalize (Gcd N) = Gcd N" and "normalize (Lcm N) = Lcm N" by simp_all qed end lemma Gcd_nat_eq_one: "1 \ N \ Gcd N = 1" for N :: "nat set" by (rule Gcd_eq_1_I) auto +instance nat :: semiring_gcd_mult_normalize + by intro_classes (auto simp: unit_factor_nat_def) + text \Alternative characterizations of Gcd:\ lemma Gcd_eq_Max: fixes M :: "nat set" assumes "finite (M::nat set)" and "M \ {}" and "0 \ M" shows "Gcd M = Max (\m\M. {d. d dvd m})" proof (rule antisym) from assms obtain m where "m \ M" and "m > 0" by auto from \m > 0\ have "finite {d. d dvd m}" by (blast intro: finite_divisors_nat) with \m \ M\ have fin: "finite (\m\M. {d. d dvd m})" by blast from fin show "Gcd M \ Max (\m\M. {d. d dvd m})" by (auto intro: Max_ge Gcd_dvd) from fin show "Max (\m\M. {d. d dvd m}) \ Gcd M" proof (rule Max.boundedI, simp_all) show "(\m\M. {d. d dvd m}) \ {}" by auto show "\a. \x\M. a dvd x \ a \ Gcd M" by (meson Gcd_dvd Gcd_greatest \0 < m\ \m \ M\ dvd_imp_le dvd_pos_nat) qed qed lemma Gcd_remove0_nat: "finite M \ Gcd M = Gcd (M - {0})" for M :: "nat set" proof (induct pred: finite) case (insert x M) then show ?case by (simp add: insert_Diff_if) qed auto lemma Lcm_in_lcm_closed_set_nat: fixes M :: "nat set" assumes "finite M" "M \ {}" "\m n. \m \ M; n \ M\ \ lcm m n \ M" shows "Lcm M \ M" using assms proof (induction M rule: finite_linorder_min_induct) case (insert x M) then have "\m n. m \ M \ n \ M \ lcm m n \ M" by (metis dvd_lcm1 gr0I insert_iff lcm_pos_nat nat_dvd_not_less) with insert show ?case by simp (metis Lcm_nat_empty One_nat_def dvd_1_left dvd_lcm2) qed auto lemma Lcm_eq_Max_nat: fixes M :: "nat set" assumes M: "finite M" "M \ {}" "0 \ M" and lcm: "\m n. \m \ M; n \ M\ \ lcm m n \ M" shows "Lcm M = Max M" proof (rule antisym) show "Lcm M \ Max M" by (simp add: Lcm_in_lcm_closed_set_nat \finite M\ \M \ {}\ lcm) show "Max M \ Lcm M" by (meson Lcm_0_iff Max_in M dvd_Lcm dvd_imp_le le_0_eq not_le) qed lemma mult_inj_if_coprime_nat: "inj_on f A \ inj_on g B \ (\a b. \a\A; b\B\ \ coprime (f a) (g b)) \ inj_on (\(a, b). f a * g b) (A \ B)" for f :: "'a \ nat" and g :: "'b \ nat" by (auto simp: inj_on_def coprime_crossproduct_nat simp del: One_nat_def) subsubsection \Setwise GCD and LCM for integers\ instantiation int :: Gcd begin definition Gcd_int :: "int set \ int" where "Gcd K = int (GCD k\K. (nat \ abs) k)" definition Lcm_int :: "int set \ int" where "Lcm K = int (LCM k\K. (nat \ abs) k)" instance .. end lemma Gcd_int_eq [simp]: "(GCD n\N. int n) = int (Gcd N)" by (simp add: Gcd_int_def image_image) lemma Gcd_nat_abs_eq [simp]: "(GCD k\K. nat \k\) = nat (Gcd K)" by (simp add: Gcd_int_def) lemma abs_Gcd_eq [simp]: "\Gcd K\ = Gcd K" for K :: "int set" by (simp only: Gcd_int_def) lemma Gcd_int_greater_eq_0 [simp]: "Gcd K \ 0" for K :: "int set" using abs_ge_zero [of "Gcd K"] by simp lemma Gcd_abs_eq [simp]: "(GCD k\K. \k\) = Gcd K" for K :: "int set" by (simp only: Gcd_int_def image_image) simp lemma Lcm_int_eq [simp]: "(LCM n\N. int n) = int (Lcm N)" by (simp add: Lcm_int_def image_image) lemma Lcm_nat_abs_eq [simp]: "(LCM k\K. nat \k\) = nat (Lcm K)" by (simp add: Lcm_int_def) lemma abs_Lcm_eq [simp]: "\Lcm K\ = Lcm K" for K :: "int set" by (simp only: Lcm_int_def) lemma Lcm_int_greater_eq_0 [simp]: "Lcm K \ 0" for K :: "int set" using abs_ge_zero [of "Lcm K"] by simp lemma Lcm_abs_eq [simp]: "(LCM k\K. \k\) = Lcm K" for K :: "int set" by (simp only: Lcm_int_def image_image) simp instance int :: semiring_Gcd proof fix K :: "int set" and k :: int show "Gcd K dvd k" and "k dvd Lcm K" if "k \ K" using that Gcd_dvd [of "nat \k\" "(nat \ abs) ` K"] dvd_Lcm [of "nat \k\" "(nat \ abs) ` K"] by (simp_all add: comp_def) show "k dvd Gcd K" if "\l. l \ K \ k dvd l" proof - have "nat \k\ dvd (GCD k\K. nat \k\)" by (rule Gcd_greatest) (use that in auto) then show ?thesis by simp qed show "Lcm K dvd k" if "\l. l \ K \ l dvd k" proof - have "(LCM k\K. nat \k\) dvd nat \k\" by (rule Lcm_least) (use that in auto) then show ?thesis by simp qed -qed simp_all +qed (simp_all add: sgn_mult) + +instance int :: semiring_gcd_mult_normalize + by intro_classes (auto simp: sgn_mult) subsection \GCD and LCM on \<^typ>\integer\\ instantiation integer :: gcd begin context includes integer.lifting begin lift_definition gcd_integer :: "integer \ integer \ integer" is gcd . lift_definition lcm_integer :: "integer \ integer \ integer" is lcm . end instance .. end lifting_update integer.lifting lifting_forget integer.lifting context includes integer.lifting begin lemma gcd_code_integer [code]: "gcd k l = \if l = (0::integer) then k else gcd l (\k\ mod \l\)\" by transfer (fact gcd_code_int) lemma lcm_code_integer [code]: "lcm a b = \a\ * \b\ div gcd a b" for a b :: integer by transfer (fact lcm_altdef_int) end code_printing constant "gcd :: integer \ _" \ (OCaml) "!(fun k l -> if Z.equal k Z.zero then/ Z.abs l else if Z.equal/ l Z.zero then Z.abs k else Z.gcd k l)" and (Haskell) "Prelude.gcd" and (Scala) "_.gcd'((_)')" \ \There is no gcd operation in the SML standard library, so no code setup for SML\ text \Some code equations\ lemmas Gcd_nat_set_eq_fold [code] = Gcd_set_eq_fold [where ?'a = nat] lemmas Lcm_nat_set_eq_fold [code] = Lcm_set_eq_fold [where ?'a = nat] lemmas Gcd_int_set_eq_fold [code] = Gcd_set_eq_fold [where ?'a = int] lemmas Lcm_int_set_eq_fold [code] = Lcm_set_eq_fold [where ?'a = int] text \Fact aliases.\ lemma lcm_0_iff_nat [simp]: "lcm m n = 0 \ m = 0 \ n = 0" for m n :: nat by (fact lcm_eq_0_iff) lemma lcm_0_iff_int [simp]: "lcm m n = 0 \ m = 0 \ n = 0" for m n :: int by (fact lcm_eq_0_iff) lemma dvd_lcm_I1_nat [simp]: "k dvd m \ k dvd lcm m n" for k m n :: nat by (fact dvd_lcmI1) lemma dvd_lcm_I2_nat [simp]: "k dvd n \ k dvd lcm m n" for k m n :: nat by (fact dvd_lcmI2) lemma dvd_lcm_I1_int [simp]: "i dvd m \ i dvd lcm m n" for i m n :: int by (fact dvd_lcmI1) lemma dvd_lcm_I2_int [simp]: "i dvd n \ i dvd lcm m n" for i m n :: int by (fact dvd_lcmI2) lemmas Gcd_dvd_nat [simp] = Gcd_dvd [where ?'a = nat] lemmas Gcd_dvd_int [simp] = Gcd_dvd [where ?'a = int] lemmas Gcd_greatest_nat [simp] = Gcd_greatest [where ?'a = nat] lemmas Gcd_greatest_int [simp] = Gcd_greatest [where ?'a = int] lemma dvd_Lcm_int [simp]: "m \ M \ m dvd Lcm M" for M :: "int set" by (fact dvd_Lcm) lemma gcd_neg_numeral_1_int [simp]: "gcd (- numeral n :: int) x = gcd (numeral n) x" by (fact gcd_neg1_int) lemma gcd_neg_numeral_2_int [simp]: "gcd x (- numeral n :: int) = gcd x (numeral n)" by (fact gcd_neg2_int) lemma gcd_proj1_if_dvd_nat [simp]: "x dvd y \ gcd x y = x" for x y :: nat by (fact gcd_nat.absorb1) lemma gcd_proj2_if_dvd_nat [simp]: "y dvd x \ gcd x y = y" for x y :: nat by (fact gcd_nat.absorb2) lemmas Lcm_eq_0_I_nat [simp] = Lcm_eq_0_I [where ?'a = nat] lemmas Lcm_0_iff_nat [simp] = Lcm_0_iff [where ?'a = nat] lemmas Lcm_least_int [simp] = Lcm_least [where ?'a = int] end diff --git a/src/HOL/Library/Multiset.thy b/src/HOL/Library/Multiset.thy --- a/src/HOL/Library/Multiset.thy +++ b/src/HOL/Library/Multiset.thy @@ -1,3927 +1,3938 @@ (* Title: HOL/Library/Multiset.thy Author: Tobias Nipkow, Markus Wenzel, Lawrence C Paulson, Norbert Voelker Author: Andrei Popescu, TU Muenchen Author: Jasmin Blanchette, Inria, LORIA, MPII Author: Dmitriy Traytel, TU Muenchen Author: Mathias Fleury, MPII *) section \(Finite) Multisets\ theory Multiset imports Cancellation begin subsection \The type of multisets\ definition "multiset = {f :: 'a \ nat. finite {x. f x > 0}}" typedef 'a multiset = "multiset :: ('a \ nat) set" morphisms count Abs_multiset unfolding multiset_def proof show "(\x. 0::nat) \ {f. finite {x. f x > 0}}" by simp qed setup_lifting type_definition_multiset lemma multiset_eq_iff: "M = N \ (\a. count M a = count N a)" by (simp only: count_inject [symmetric] fun_eq_iff) lemma multiset_eqI: "(\x. count A x = count B x) \ A = B" using multiset_eq_iff by auto text \Preservation of the representing set \<^term>\multiset\.\ lemma const0_in_multiset: "(\a. 0) \ multiset" by (simp add: multiset_def) lemma only1_in_multiset: "(\b. if b = a then n else 0) \ multiset" by (simp add: multiset_def) lemma union_preserves_multiset: "M \ multiset \ N \ multiset \ (\a. M a + N a) \ multiset" by (simp add: multiset_def) lemma diff_preserves_multiset: assumes "M \ multiset" shows "(\a. M a - N a) \ multiset" proof - have "{x. N x < M x} \ {x. 0 < M x}" by auto with assms show ?thesis by (auto simp add: multiset_def intro: finite_subset) qed lemma filter_preserves_multiset: assumes "M \ multiset" shows "(\x. if P x then M x else 0) \ multiset" proof - have "{x. (P x \ 0 < M x) \ P x} \ {x. 0 < M x}" by auto with assms show ?thesis by (auto simp add: multiset_def intro: finite_subset) qed lemmas in_multiset = const0_in_multiset only1_in_multiset union_preserves_multiset diff_preserves_multiset filter_preserves_multiset subsection \Representing multisets\ text \Multiset enumeration\ instantiation multiset :: (type) cancel_comm_monoid_add begin lift_definition zero_multiset :: "'a multiset" is "\a. 0" by (rule const0_in_multiset) abbreviation Mempty :: "'a multiset" ("{#}") where "Mempty \ 0" lift_definition plus_multiset :: "'a multiset \ 'a multiset \ 'a multiset" is "\M N. (\a. M a + N a)" by (rule union_preserves_multiset) lift_definition minus_multiset :: "'a multiset \ 'a multiset \ 'a multiset" is "\ M N. \a. M a - N a" by (rule diff_preserves_multiset) instance by (standard; transfer; simp add: fun_eq_iff) end context begin qualified definition is_empty :: "'a multiset \ bool" where [code_abbrev]: "is_empty A \ A = {#}" end lemma add_mset_in_multiset: assumes M: \M \ multiset\ shows \(\b. if b = a then Suc (M b) else M b) \ multiset\ using assms by (simp add: multiset_def flip: insert_Collect) lift_definition add_mset :: "'a \ 'a multiset \ 'a multiset" is "\a M b. if b = a then Suc (M b) else M b" by (rule add_mset_in_multiset) syntax "_multiset" :: "args \ 'a multiset" ("{#(_)#}") translations "{#x, xs#}" == "CONST add_mset x {#xs#}" "{#x#}" == "CONST add_mset x {#}" lemma count_empty [simp]: "count {#} a = 0" by (simp add: zero_multiset.rep_eq) lemma count_add_mset [simp]: "count (add_mset b A) a = (if b = a then Suc (count A a) else count A a)" by (simp add: add_mset.rep_eq) lemma count_single: "count {#b#} a = (if b = a then 1 else 0)" by simp lemma add_mset_not_empty [simp]: \add_mset a A \ {#}\ and empty_not_add_mset [simp]: "{#} \ add_mset a A" by (auto simp: multiset_eq_iff) lemma add_mset_add_mset_same_iff [simp]: "add_mset a A = add_mset a B \ A = B" by (auto simp: multiset_eq_iff) lemma add_mset_commute: "add_mset x (add_mset y M) = add_mset y (add_mset x M)" by (auto simp: multiset_eq_iff) subsection \Basic operations\ subsubsection \Conversion to set and membership\ definition set_mset :: "'a multiset \ 'a set" where "set_mset M = {x. count M x > 0}" abbreviation Melem :: "'a \ 'a multiset \ bool" where "Melem a M \ a \ set_mset M" notation Melem ("'(\#')") and Melem ("(_/ \# _)" [51, 51] 50) notation (ASCII) Melem ("'(:#')") and Melem ("(_/ :# _)" [51, 51] 50) abbreviation not_Melem :: "'a \ 'a multiset \ bool" where "not_Melem a M \ a \ set_mset M" notation not_Melem ("'(\#')") and not_Melem ("(_/ \# _)" [51, 51] 50) notation (ASCII) not_Melem ("'(~:#')") and not_Melem ("(_/ ~:# _)" [51, 51] 50) context begin qualified abbreviation Ball :: "'a multiset \ ('a \ bool) \ bool" where "Ball M \ Set.Ball (set_mset M)" qualified abbreviation Bex :: "'a multiset \ ('a \ bool) \ bool" where "Bex M \ Set.Bex (set_mset M)" end syntax "_MBall" :: "pttrn \ 'a set \ bool \ bool" ("(3\_\#_./ _)" [0, 0, 10] 10) "_MBex" :: "pttrn \ 'a set \ bool \ bool" ("(3\_\#_./ _)" [0, 0, 10] 10) syntax (ASCII) "_MBall" :: "pttrn \ 'a set \ bool \ bool" ("(3\_:#_./ _)" [0, 0, 10] 10) "_MBex" :: "pttrn \ 'a set \ bool \ bool" ("(3\_:#_./ _)" [0, 0, 10] 10) translations "\x\#A. P" \ "CONST Multiset.Ball A (\x. P)" "\x\#A. P" \ "CONST Multiset.Bex A (\x. P)" lemma count_eq_zero_iff: "count M x = 0 \ x \# M" by (auto simp add: set_mset_def) lemma not_in_iff: "x \# M \ count M x = 0" by (auto simp add: count_eq_zero_iff) lemma count_greater_zero_iff [simp]: "count M x > 0 \ x \# M" by (auto simp add: set_mset_def) lemma count_inI: assumes "count M x = 0 \ False" shows "x \# M" proof (rule ccontr) assume "x \# M" with assms show False by (simp add: not_in_iff) qed lemma in_countE: assumes "x \# M" obtains n where "count M x = Suc n" proof - from assms have "count M x > 0" by simp then obtain n where "count M x = Suc n" using gr0_conv_Suc by blast with that show thesis . qed lemma count_greater_eq_Suc_zero_iff [simp]: "count M x \ Suc 0 \ x \# M" by (simp add: Suc_le_eq) lemma count_greater_eq_one_iff [simp]: "count M x \ 1 \ x \# M" by simp lemma set_mset_empty [simp]: "set_mset {#} = {}" by (simp add: set_mset_def) lemma set_mset_single: "set_mset {#b#} = {b}" by (simp add: set_mset_def) lemma set_mset_eq_empty_iff [simp]: "set_mset M = {} \ M = {#}" by (auto simp add: multiset_eq_iff count_eq_zero_iff) lemma finite_set_mset [iff]: "finite (set_mset M)" using count [of M] by (simp add: multiset_def) lemma set_mset_add_mset_insert [simp]: \set_mset (add_mset a A) = insert a (set_mset A)\ by (auto simp flip: count_greater_eq_Suc_zero_iff split: if_splits) lemma multiset_nonemptyE [elim]: assumes "A \ {#}" obtains x where "x \# A" proof - have "\x. x \# A" by (rule ccontr) (insert assms, auto) with that show ?thesis by blast qed subsubsection \Union\ lemma count_union [simp]: "count (M + N) a = count M a + count N a" by (simp add: plus_multiset.rep_eq) lemma set_mset_union [simp]: "set_mset (M + N) = set_mset M \ set_mset N" by (simp only: set_eq_iff count_greater_zero_iff [symmetric] count_union) simp lemma union_mset_add_mset_left [simp]: "add_mset a A + B = add_mset a (A + B)" by (auto simp: multiset_eq_iff) lemma union_mset_add_mset_right [simp]: "A + add_mset a B = add_mset a (A + B)" by (auto simp: multiset_eq_iff) lemma add_mset_add_single: \add_mset a A = A + {#a#}\ by (subst union_mset_add_mset_right, subst add.comm_neutral) standard subsubsection \Difference\ instance multiset :: (type) comm_monoid_diff by standard (transfer; simp add: fun_eq_iff) lemma count_diff [simp]: "count (M - N) a = count M a - count N a" by (simp add: minus_multiset.rep_eq) lemma add_mset_diff_bothsides: \add_mset a M - add_mset a A = M - A\ by (auto simp: multiset_eq_iff) lemma in_diff_count: "a \# M - N \ count N a < count M a" by (simp add: set_mset_def) lemma count_in_diffI: assumes "\n. count N x = n + count M x \ False" shows "x \# M - N" proof (rule ccontr) assume "x \# M - N" then have "count N x = (count N x - count M x) + count M x" by (simp add: in_diff_count not_less) with assms show False by auto qed lemma in_diff_countE: assumes "x \# M - N" obtains n where "count M x = Suc n + count N x" proof - from assms have "count M x - count N x > 0" by (simp add: in_diff_count) then have "count M x > count N x" by simp then obtain n where "count M x = Suc n + count N x" using less_iff_Suc_add by auto with that show thesis . qed lemma in_diffD: assumes "a \# M - N" shows "a \# M" proof - have "0 \ count N a" by simp also from assms have "count N a < count M a" by (simp add: in_diff_count) finally show ?thesis by simp qed lemma set_mset_diff: "set_mset (M - N) = {a. count N a < count M a}" by (simp add: set_mset_def) lemma diff_empty [simp]: "M - {#} = M \ {#} - M = {#}" by rule (fact Groups.diff_zero, fact Groups.zero_diff) lemma diff_cancel: "A - A = {#}" by (fact Groups.diff_cancel) lemma diff_union_cancelR: "M + N - N = (M::'a multiset)" by (fact add_diff_cancel_right') lemma diff_union_cancelL: "N + M - N = (M::'a multiset)" by (fact add_diff_cancel_left') lemma diff_right_commute: fixes M N Q :: "'a multiset" shows "M - N - Q = M - Q - N" by (fact diff_right_commute) lemma diff_add: fixes M N Q :: "'a multiset" shows "M - (N + Q) = M - N - Q" by (rule sym) (fact diff_diff_add) lemma insert_DiffM [simp]: "x \# M \ add_mset x (M - {#x#}) = M" by (clarsimp simp: multiset_eq_iff) lemma insert_DiffM2: "x \# M \ (M - {#x#}) + {#x#} = M" by simp lemma diff_union_swap: "a \ b \ add_mset b (M - {#a#}) = add_mset b M - {#a#}" by (auto simp add: multiset_eq_iff) lemma diff_add_mset_swap [simp]: "b \# A \ add_mset b M - A = add_mset b (M - A)" by (auto simp add: multiset_eq_iff simp: not_in_iff) lemma diff_union_swap2 [simp]: "y \# M \ add_mset x M - {#y#} = add_mset x (M - {#y#})" by (metis add_mset_diff_bothsides diff_union_swap diff_zero insert_DiffM) lemma diff_diff_add_mset [simp]: "(M::'a multiset) - N - P = M - (N + P)" by (rule diff_diff_add) lemma diff_union_single_conv: "a \# J \ I + J - {#a#} = I + (J - {#a#})" by (simp add: multiset_eq_iff Suc_le_eq) lemma mset_add [elim?]: assumes "a \# A" obtains B where "A = add_mset a B" proof - from assms have "A = add_mset a (A - {#a#})" by simp with that show thesis . qed lemma union_iff: "a \# A + B \ a \# A \ a \# B" by auto subsubsection \Min and Max\ abbreviation Min_mset :: "'a::linorder multiset \ 'a" where "Min_mset m \ Min (set_mset m)" abbreviation Max_mset :: "'a::linorder multiset \ 'a" where "Max_mset m \ Max (set_mset m)" subsubsection \Equality of multisets\ lemma single_eq_single [simp]: "{#a#} = {#b#} \ a = b" by (auto simp add: multiset_eq_iff) lemma union_eq_empty [iff]: "M + N = {#} \ M = {#} \ N = {#}" by (auto simp add: multiset_eq_iff) lemma empty_eq_union [iff]: "{#} = M + N \ M = {#} \ N = {#}" by (auto simp add: multiset_eq_iff) lemma multi_self_add_other_not_self [simp]: "M = add_mset x M \ False" by (auto simp add: multiset_eq_iff) lemma add_mset_remove_trivial [simp]: \add_mset x M - {#x#} = M\ by (auto simp: multiset_eq_iff) lemma diff_single_trivial: "\ x \# M \ M - {#x#} = M" by (auto simp add: multiset_eq_iff not_in_iff) lemma diff_single_eq_union: "x \# M \ M - {#x#} = N \ M = add_mset x N" by auto lemma union_single_eq_diff: "add_mset x M = N \ M = N - {#x#}" unfolding add_mset_add_single[of _ M] by (fact add_implies_diff) lemma union_single_eq_member: "add_mset x M = N \ x \# N" by auto lemma add_mset_remove_trivial_If: "add_mset a (N - {#a#}) = (if a \# N then N else add_mset a N)" by (simp add: diff_single_trivial) lemma add_mset_remove_trivial_eq: \N = add_mset a (N - {#a#}) \ a \# N\ by (auto simp: add_mset_remove_trivial_If) lemma union_is_single: "M + N = {#a#} \ M = {#a#} \ N = {#} \ M = {#} \ N = {#a#}" (is "?lhs = ?rhs") proof show ?lhs if ?rhs using that by auto show ?rhs if ?lhs by (metis Multiset.diff_cancel add.commute add_diff_cancel_left' diff_add_zero diff_single_trivial insert_DiffM that) qed lemma single_is_union: "{#a#} = M + N \ {#a#} = M \ N = {#} \ M = {#} \ {#a#} = N" by (auto simp add: eq_commute [of "{#a#}" "M + N"] union_is_single) lemma add_eq_conv_diff: "add_mset a M = add_mset b N \ M = N \ a = b \ M = add_mset b (N - {#a#}) \ N = add_mset a (M - {#b#})" (is "?lhs \ ?rhs") (* shorter: by (simp add: multiset_eq_iff) fastforce *) proof show ?lhs if ?rhs using that by (auto simp add: add_mset_commute[of a b]) show ?rhs if ?lhs proof (cases "a = b") case True with \?lhs\ show ?thesis by simp next case False from \?lhs\ have "a \# add_mset b N" by (rule union_single_eq_member) with False have "a \# N" by auto moreover from \?lhs\ have "M = add_mset b N - {#a#}" by (rule union_single_eq_diff) moreover note False ultimately show ?thesis by (auto simp add: diff_right_commute [of _ "{#a#}"]) qed qed lemma add_mset_eq_single [iff]: "add_mset b M = {#a#} \ b = a \ M = {#}" by (auto simp: add_eq_conv_diff) lemma single_eq_add_mset [iff]: "{#a#} = add_mset b M \ b = a \ M = {#}" by (auto simp: add_eq_conv_diff) lemma insert_noteq_member: assumes BC: "add_mset b B = add_mset c C" and bnotc: "b \ c" shows "c \# B" proof - have "c \# add_mset c C" by simp have nc: "\ c \# {#b#}" using bnotc by simp then have "c \# add_mset b B" using BC by simp then show "c \# B" using nc by simp qed lemma add_eq_conv_ex: "(add_mset a M = add_mset b N) = (M = N \ a = b \ (\K. M = add_mset b K \ N = add_mset a K))" by (auto simp add: add_eq_conv_diff) lemma multi_member_split: "x \# M \ \A. M = add_mset x A" by (rule exI [where x = "M - {#x#}"]) simp lemma multiset_add_sub_el_shuffle: assumes "c \# B" and "b \ c" shows "add_mset b (B - {#c#}) = add_mset b B - {#c#}" proof - from \c \# B\ obtain A where B: "B = add_mset c A" by (blast dest: multi_member_split) have "add_mset b A = add_mset c (add_mset b A) - {#c#}" by simp then have "add_mset b A = add_mset b (add_mset c A) - {#c#}" by (simp add: \b \ c\) then show ?thesis using B by simp qed lemma add_mset_eq_singleton_iff[iff]: "add_mset x M = {#y#} \ M = {#} \ x = y" by auto subsubsection \Pointwise ordering induced by count\ definition subseteq_mset :: "'a multiset \ 'a multiset \ bool" (infix "\#" 50) where "A \# B \ (\a. count A a \ count B a)" definition subset_mset :: "'a multiset \ 'a multiset \ bool" (infix "\#" 50) where "A \# B \ A \# B \ A \ B" abbreviation (input) supseteq_mset :: "'a multiset \ 'a multiset \ bool" (infix "\#" 50) where "supseteq_mset A B \ B \# A" abbreviation (input) supset_mset :: "'a multiset \ 'a multiset \ bool" (infix "\#" 50) where "supset_mset A B \ B \# A" notation (input) subseteq_mset (infix "\#" 50) and supseteq_mset (infix "\#" 50) notation (ASCII) subseteq_mset (infix "<=#" 50) and subset_mset (infix "<#" 50) and supseteq_mset (infix ">=#" 50) and supset_mset (infix ">#" 50) interpretation subset_mset: ordered_ab_semigroup_add_imp_le "(+)" "(-)" "(\#)" "(\#)" by standard (auto simp add: subset_mset_def subseteq_mset_def multiset_eq_iff intro: order_trans antisym) \ \FIXME: avoid junk stemming from type class interpretation\ interpretation subset_mset: ordered_ab_semigroup_monoid_add_imp_le "(+)" 0 "(-)" "(\#)" "(\#)" by standard \ \FIXME: avoid junk stemming from type class interpretation\ lemma mset_subset_eqI: "(\a. count A a \ count B a) \ A \# B" by (simp add: subseteq_mset_def) lemma mset_subset_eq_count: "A \# B \ count A a \ count B a" by (simp add: subseteq_mset_def) lemma mset_subset_eq_exists_conv: "(A::'a multiset) \# B \ (\C. B = A + C)" unfolding subseteq_mset_def apply (rule iffI) apply (rule exI [where x = "B - A"]) apply (auto intro: multiset_eq_iff [THEN iffD2]) done interpretation subset_mset: ordered_cancel_comm_monoid_diff "(+)" 0 "(\#)" "(\#)" "(-)" by standard (simp, fact mset_subset_eq_exists_conv) \ \FIXME: avoid junk stemming from type class interpretation\ declare subset_mset.add_diff_assoc[simp] subset_mset.add_diff_assoc2[simp] lemma mset_subset_eq_mono_add_right_cancel: "(A::'a multiset) + C \# B + C \ A \# B" by (fact subset_mset.add_le_cancel_right) lemma mset_subset_eq_mono_add_left_cancel: "C + (A::'a multiset) \# C + B \ A \# B" by (fact subset_mset.add_le_cancel_left) lemma mset_subset_eq_mono_add: "(A::'a multiset) \# B \ C \# D \ A + C \# B + D" by (fact subset_mset.add_mono) lemma mset_subset_eq_add_left: "(A::'a multiset) \# A + B" by simp lemma mset_subset_eq_add_right: "B \# (A::'a multiset) + B" by simp lemma single_subset_iff [simp]: "{#a#} \# M \ a \# M" by (auto simp add: subseteq_mset_def Suc_le_eq) lemma mset_subset_eq_single: "a \# B \ {#a#} \# B" by simp lemma mset_subset_eq_add_mset_cancel: \add_mset a A \# add_mset a B \ A \# B\ unfolding add_mset_add_single[of _ A] add_mset_add_single[of _ B] by (rule mset_subset_eq_mono_add_right_cancel) lemma multiset_diff_union_assoc: fixes A B C D :: "'a multiset" shows "C \# B \ A + B - C = A + (B - C)" by (fact subset_mset.diff_add_assoc) lemma mset_subset_eq_multiset_union_diff_commute: fixes A B C D :: "'a multiset" shows "B \# A \ A - B + C = A + C - B" by (fact subset_mset.add_diff_assoc2) lemma diff_subset_eq_self[simp]: "(M::'a multiset) - N \# M" by (simp add: subseteq_mset_def) lemma mset_subset_eqD: assumes "A \# B" and "x \# A" shows "x \# B" proof - from \x \# A\ have "count A x > 0" by simp also from \A \# B\ have "count A x \ count B x" by (simp add: subseteq_mset_def) finally show ?thesis by simp qed lemma mset_subsetD: "A \# B \ x \# A \ x \# B" by (auto intro: mset_subset_eqD [of A]) lemma set_mset_mono: "A \# B \ set_mset A \ set_mset B" by (metis mset_subset_eqD subsetI) lemma mset_subset_eq_insertD: "add_mset x A \# B \ x \# B \ A \# B" apply (rule conjI) apply (simp add: mset_subset_eqD) apply (clarsimp simp: subset_mset_def subseteq_mset_def) apply safe apply (erule_tac x = a in allE) apply (auto split: if_split_asm) done lemma mset_subset_insertD: "add_mset x A \# B \ x \# B \ A \# B" by (rule mset_subset_eq_insertD) simp lemma mset_subset_of_empty[simp]: "A \# {#} \ False" by (simp only: subset_mset.not_less_zero) lemma empty_subset_add_mset[simp]: "{#} \# add_mset x M" by (auto intro: subset_mset.gr_zeroI) lemma empty_le: "{#} \# A" by (fact subset_mset.zero_le) lemma insert_subset_eq_iff: "add_mset a A \# B \ a \# B \ A \# B - {#a#}" using le_diff_conv2 [of "Suc 0" "count B a" "count A a"] apply (auto simp add: subseteq_mset_def not_in_iff Suc_le_eq) apply (rule ccontr) apply (auto simp add: not_in_iff) done lemma insert_union_subset_iff: "add_mset a A \# B \ a \# B \ A \# B - {#a#}" by (auto simp add: insert_subset_eq_iff subset_mset_def) lemma subset_eq_diff_conv: "A - C \# B \ A \# B + C" by (simp add: subseteq_mset_def le_diff_conv) lemma multi_psub_of_add_self [simp]: "A \# add_mset x A" by (auto simp: subset_mset_def subseteq_mset_def) lemma multi_psub_self: "A \# A = False" by simp lemma mset_subset_add_mset [simp]: "add_mset x N \# add_mset x M \ N \# M" unfolding add_mset_add_single[of _ N] add_mset_add_single[of _ M] by (fact subset_mset.add_less_cancel_right) lemma mset_subset_diff_self: "c \# B \ B - {#c#} \# B" by (auto simp: subset_mset_def elim: mset_add) lemma Diff_eq_empty_iff_mset: "A - B = {#} \ A \# B" by (auto simp: multiset_eq_iff subseteq_mset_def) lemma add_mset_subseteq_single_iff[iff]: "add_mset a M \# {#b#} \ M = {#} \ a = b" proof assume A: "add_mset a M \# {#b#}" then have \a = b\ by (auto dest: mset_subset_eq_insertD) then show "M={#} \ a=b" using A by (simp add: mset_subset_eq_add_mset_cancel) qed simp subsubsection \Intersection and bounded union\ definition inf_subset_mset :: "'a multiset \ 'a multiset \ 'a multiset" (infixl "\#" 70) where multiset_inter_def: "inf_subset_mset A B = A - (A - B)" interpretation subset_mset: semilattice_inf inf_subset_mset "(\#)" "(\#)" proof - have [simp]: "m \ n \ m \ q \ m \ n - (n - q)" for m n q :: nat by arith show "class.semilattice_inf (\#) (\#) (\#)" by standard (auto simp add: multiset_inter_def subseteq_mset_def) qed \ \FIXME: avoid junk stemming from type class interpretation\ definition sup_subset_mset :: "'a multiset \ 'a multiset \ 'a multiset"(infixl "\#" 70) where "sup_subset_mset A B = A + (B - A)" \ \FIXME irregular fact name\ interpretation subset_mset: semilattice_sup sup_subset_mset "(\#)" "(\#)" proof - have [simp]: "m \ n \ q \ n \ m + (q - m) \ n" for m n q :: nat by arith show "class.semilattice_sup (\#) (\#) (\#)" by standard (auto simp add: sup_subset_mset_def subseteq_mset_def) qed \ \FIXME: avoid junk stemming from type class interpretation\ interpretation subset_mset: bounded_lattice_bot "(\#)" "(\#)" "(\#)" "(\#)" "{#}" by standard auto \ \FIXME: avoid junk stemming from type class interpretation\ subsubsection \Additional intersection facts\ lemma multiset_inter_count [simp]: fixes A B :: "'a multiset" shows "count (A \# B) x = min (count A x) (count B x)" by (simp add: multiset_inter_def) lemma set_mset_inter [simp]: "set_mset (A \# B) = set_mset A \ set_mset B" by (simp only: set_eq_iff count_greater_zero_iff [symmetric] multiset_inter_count) simp lemma diff_intersect_left_idem [simp]: "M - M \# N = M - N" by (simp add: multiset_eq_iff min_def) lemma diff_intersect_right_idem [simp]: "M - N \# M = M - N" by (simp add: multiset_eq_iff min_def) lemma multiset_inter_single[simp]: "a \ b \ {#a#} \# {#b#} = {#}" by (rule multiset_eqI) auto lemma multiset_union_diff_commute: assumes "B \# C = {#}" shows "A + B - C = A - C + B" proof (rule multiset_eqI) fix x from assms have "min (count B x) (count C x) = 0" by (auto simp add: multiset_eq_iff) then have "count B x = 0 \ count C x = 0" unfolding min_def by (auto split: if_splits) then show "count (A + B - C) x = count (A - C + B) x" by auto qed lemma disjunct_not_in: "A \# B = {#} \ (\a. a \# A \ a \# B)" (is "?P \ ?Q") proof assume ?P show ?Q proof fix a from \?P\ have "min (count A a) (count B a) = 0" by (simp add: multiset_eq_iff) then have "count A a = 0 \ count B a = 0" by (cases "count A a \ count B a") (simp_all add: min_def) then show "a \# A \ a \# B" by (simp add: not_in_iff) qed next assume ?Q show ?P proof (rule multiset_eqI) fix a from \?Q\ have "count A a = 0 \ count B a = 0" by (auto simp add: not_in_iff) then show "count (A \# B) a = count {#} a" by auto qed qed lemma inter_mset_empty_distrib_right: "A \# (B + C) = {#} \ A \# B = {#} \ A \# C = {#}" by (meson disjunct_not_in union_iff) lemma inter_mset_empty_distrib_left: "(A + B) \# C = {#} \ A \# C = {#} \ B \# C = {#}" by (meson disjunct_not_in union_iff) lemma add_mset_inter_add_mset[simp]: "add_mset a A \# add_mset a B = add_mset a (A \# B)" by (metis add_mset_add_single add_mset_diff_bothsides diff_subset_eq_self multiset_inter_def subset_mset.diff_add_assoc2) lemma add_mset_disjoint [simp]: "add_mset a A \# B = {#} \ a \# B \ A \# B = {#}" "{#} = add_mset a A \# B \ a \# B \ {#} = A \# B" by (auto simp: disjunct_not_in) lemma disjoint_add_mset [simp]: "B \# add_mset a A = {#} \ a \# B \ B \# A = {#}" "{#} = A \# add_mset b B \ b \# A \ {#} = A \# B" by (auto simp: disjunct_not_in) lemma inter_add_left1: "\ x \# N \ (add_mset x M) \# N = M \# N" by (simp add: multiset_eq_iff not_in_iff) lemma inter_add_left2: "x \# N \ (add_mset x M) \# N = add_mset x (M \# (N - {#x#}))" by (auto simp add: multiset_eq_iff elim: mset_add) lemma inter_add_right1: "\ x \# N \ N \# (add_mset x M) = N \# M" by (simp add: multiset_eq_iff not_in_iff) lemma inter_add_right2: "x \# N \ N \# (add_mset x M) = add_mset x ((N - {#x#}) \# M)" by (auto simp add: multiset_eq_iff elim: mset_add) lemma disjunct_set_mset_diff: assumes "M \# N = {#}" shows "set_mset (M - N) = set_mset M" proof (rule set_eqI) fix a from assms have "a \# M \ a \# N" by (simp add: disjunct_not_in) then show "a \# M - N \ a \# M" by (auto dest: in_diffD) (simp add: in_diff_count not_in_iff) qed lemma at_most_one_mset_mset_diff: assumes "a \# M - {#a#}" shows "set_mset (M - {#a#}) = set_mset M - {a}" using assms by (auto simp add: not_in_iff in_diff_count set_eq_iff) lemma more_than_one_mset_mset_diff: assumes "a \# M - {#a#}" shows "set_mset (M - {#a#}) = set_mset M" proof (rule set_eqI) fix b have "Suc 0 < count M b \ count M b > 0" by arith then show "b \# M - {#a#} \ b \# M" using assms by (auto simp add: in_diff_count) qed lemma inter_iff: "a \# A \# B \ a \# A \ a \# B" by simp lemma inter_union_distrib_left: "A \# B + C = (A + C) \# (B + C)" by (simp add: multiset_eq_iff min_add_distrib_left) lemma inter_union_distrib_right: "C + A \# B = (C + A) \# (C + B)" using inter_union_distrib_left [of A B C] by (simp add: ac_simps) lemma inter_subset_eq_union: "A \# B \# A + B" by (auto simp add: subseteq_mset_def) subsubsection \Additional bounded union facts\ lemma sup_subset_mset_count [simp]: \ \FIXME irregular fact name\ "count (A \# B) x = max (count A x) (count B x)" by (simp add: sup_subset_mset_def) lemma set_mset_sup [simp]: "set_mset (A \# B) = set_mset A \ set_mset B" by (simp only: set_eq_iff count_greater_zero_iff [symmetric] sup_subset_mset_count) (auto simp add: not_in_iff elim: mset_add) lemma sup_union_left1 [simp]: "\ x \# N \ (add_mset x M) \# N = add_mset x (M \# N)" by (simp add: multiset_eq_iff not_in_iff) lemma sup_union_left2: "x \# N \ (add_mset x M) \# N = add_mset x (M \# (N - {#x#}))" by (simp add: multiset_eq_iff) lemma sup_union_right1 [simp]: "\ x \# N \ N \# (add_mset x M) = add_mset x (N \# M)" by (simp add: multiset_eq_iff not_in_iff) lemma sup_union_right2: "x \# N \ N \# (add_mset x M) = add_mset x ((N - {#x#}) \# M)" by (simp add: multiset_eq_iff) lemma sup_union_distrib_left: "A \# B + C = (A + C) \# (B + C)" by (simp add: multiset_eq_iff max_add_distrib_left) lemma union_sup_distrib_right: "C + A \# B = (C + A) \# (C + B)" using sup_union_distrib_left [of A B C] by (simp add: ac_simps) lemma union_diff_inter_eq_sup: "A + B - A \# B = A \# B" by (auto simp add: multiset_eq_iff) lemma union_diff_sup_eq_inter: "A + B - A \# B = A \# B" by (auto simp add: multiset_eq_iff) lemma add_mset_union: \add_mset a A \# add_mset a B = add_mset a (A \# B)\ by (auto simp: multiset_eq_iff max_def) subsection \Replicate and repeat operations\ definition replicate_mset :: "nat \ 'a \ 'a multiset" where "replicate_mset n x = (add_mset x ^^ n) {#}" lemma replicate_mset_0[simp]: "replicate_mset 0 x = {#}" unfolding replicate_mset_def by simp lemma replicate_mset_Suc [simp]: "replicate_mset (Suc n) x = add_mset x (replicate_mset n x)" unfolding replicate_mset_def by (induct n) (auto intro: add.commute) lemma count_replicate_mset[simp]: "count (replicate_mset n x) y = (if y = x then n else 0)" unfolding replicate_mset_def by (induct n) auto fun repeat_mset :: "nat \ 'a multiset \ 'a multiset" where "repeat_mset 0 _ = {#}" | "repeat_mset (Suc n) A = A + repeat_mset n A" lemma count_repeat_mset [simp]: "count (repeat_mset i A) a = i * count A a" by (induction i) auto lemma repeat_mset_right [simp]: "repeat_mset a (repeat_mset b A) = repeat_mset (a * b) A" by (auto simp: multiset_eq_iff left_diff_distrib') lemma left_diff_repeat_mset_distrib': \repeat_mset (i - j) u = repeat_mset i u - repeat_mset j u\ by (auto simp: multiset_eq_iff left_diff_distrib') lemma left_add_mult_distrib_mset: "repeat_mset i u + (repeat_mset j u + k) = repeat_mset (i+j) u + k" by (auto simp: multiset_eq_iff add_mult_distrib) lemma repeat_mset_distrib: "repeat_mset (m + n) A = repeat_mset m A + repeat_mset n A" by (auto simp: multiset_eq_iff Nat.add_mult_distrib) lemma repeat_mset_distrib2[simp]: "repeat_mset n (A + B) = repeat_mset n A + repeat_mset n B" by (auto simp: multiset_eq_iff add_mult_distrib2) lemma repeat_mset_replicate_mset[simp]: "repeat_mset n {#a#} = replicate_mset n a" by (auto simp: multiset_eq_iff) lemma repeat_mset_distrib_add_mset[simp]: "repeat_mset n (add_mset a A) = replicate_mset n a + repeat_mset n A" by (auto simp: multiset_eq_iff) lemma repeat_mset_empty[simp]: "repeat_mset n {#} = {#}" by (induction n) simp_all subsubsection \Simprocs\ lemma repeat_mset_iterate_add: \repeat_mset n M = iterate_add n M\ unfolding iterate_add_def by (induction n) auto lemma mset_subseteq_add_iff1: "j \ (i::nat) \ (repeat_mset i u + m \# repeat_mset j u + n) = (repeat_mset (i-j) u + m \# n)" by (auto simp add: subseteq_mset_def nat_le_add_iff1) lemma mset_subseteq_add_iff2: "i \ (j::nat) \ (repeat_mset i u + m \# repeat_mset j u + n) = (m \# repeat_mset (j-i) u + n)" by (auto simp add: subseteq_mset_def nat_le_add_iff2) lemma mset_subset_add_iff1: "j \ (i::nat) \ (repeat_mset i u + m \# repeat_mset j u + n) = (repeat_mset (i-j) u + m \# n)" unfolding subset_mset_def repeat_mset_iterate_add by (simp add: iterate_add_eq_add_iff1 mset_subseteq_add_iff1[unfolded repeat_mset_iterate_add]) lemma mset_subset_add_iff2: "i \ (j::nat) \ (repeat_mset i u + m \# repeat_mset j u + n) = (m \# repeat_mset (j-i) u + n)" unfolding subset_mset_def repeat_mset_iterate_add by (simp add: iterate_add_eq_add_iff2 mset_subseteq_add_iff2[unfolded repeat_mset_iterate_add]) ML_file \multiset_simprocs.ML\ lemma add_mset_replicate_mset_safe[cancelation_simproc_pre]: \NO_MATCH {#} M \ add_mset a M = {#a#} + M\ by simp declare repeat_mset_iterate_add[cancelation_simproc_pre] declare iterate_add_distrib[cancelation_simproc_pre] declare repeat_mset_iterate_add[symmetric, cancelation_simproc_post] declare add_mset_not_empty[cancelation_simproc_eq_elim] empty_not_add_mset[cancelation_simproc_eq_elim] subset_mset.le_zero_eq[cancelation_simproc_eq_elim] empty_not_add_mset[cancelation_simproc_eq_elim] add_mset_not_empty[cancelation_simproc_eq_elim] subset_mset.le_zero_eq[cancelation_simproc_eq_elim] le_zero_eq[cancelation_simproc_eq_elim] simproc_setup mseteq_cancel ("(l::'a multiset) + m = n" | "(l::'a multiset) = m + n" | "add_mset a m = n" | "m = add_mset a n" | "replicate_mset p a = n" | "m = replicate_mset p a" | "repeat_mset p m = n" | "m = repeat_mset p m") = \fn phi => Cancel_Simprocs.eq_cancel\ simproc_setup msetsubset_cancel ("(l::'a multiset) + m \# n" | "(l::'a multiset) \# m + n" | "add_mset a m \# n" | "m \# add_mset a n" | "replicate_mset p r \# n" | "m \# replicate_mset p r" | "repeat_mset p m \# n" | "m \# repeat_mset p m") = \fn phi => Multiset_Simprocs.subset_cancel_msets\ simproc_setup msetsubset_eq_cancel ("(l::'a multiset) + m \# n" | "(l::'a multiset) \# m + n" | "add_mset a m \# n" | "m \# add_mset a n" | "replicate_mset p r \# n" | "m \# replicate_mset p r" | "repeat_mset p m \# n" | "m \# repeat_mset p m") = \fn phi => Multiset_Simprocs.subseteq_cancel_msets\ simproc_setup msetdiff_cancel ("((l::'a multiset) + m) - n" | "(l::'a multiset) - (m + n)" | "add_mset a m - n" | "m - add_mset a n" | "replicate_mset p r - n" | "m - replicate_mset p r" | "repeat_mset p m - n" | "m - repeat_mset p m") = \fn phi => Cancel_Simprocs.diff_cancel\ subsubsection \Conditionally complete lattice\ instantiation multiset :: (type) Inf begin lift_definition Inf_multiset :: "'a multiset set \ 'a multiset" is "\A i. if A = {} then 0 else Inf ((\f. f i) ` A)" proof - fix A :: "('a \ nat) set" assume *: "\x. x \ A \ x \ multiset" have "finite {i. (if A = {} then 0 else Inf ((\f. f i) ` A)) > 0}" unfolding multiset_def proof (cases "A = {}") case False then obtain f where "f \ A" by blast hence "{i. Inf ((\f. f i) ` A) > 0} \ {i. f i > 0}" by (auto intro: less_le_trans[OF _ cInf_lower]) moreover from \f \ A\ * have "finite \" by (simp add: multiset_def) ultimately have "finite {i. Inf ((\f. f i) ` A) > 0}" by (rule finite_subset) with False show ?thesis by simp qed simp_all thus "(\i. if A = {} then 0 else INF f\A. f i) \ multiset" by (simp add: multiset_def) qed instance .. end lemma Inf_multiset_empty: "Inf {} = {#}" by transfer simp_all lemma count_Inf_multiset_nonempty: "A \ {} \ count (Inf A) x = Inf ((\X. count X x) ` A)" by transfer simp_all instantiation multiset :: (type) Sup begin definition Sup_multiset :: "'a multiset set \ 'a multiset" where "Sup_multiset A = (if A \ {} \ subset_mset.bdd_above A then Abs_multiset (\i. Sup ((\X. count X i) ` A)) else {#})" lemma Sup_multiset_empty: "Sup {} = {#}" by (simp add: Sup_multiset_def) lemma Sup_multiset_unbounded: "\subset_mset.bdd_above A \ Sup A = {#}" by (simp add: Sup_multiset_def) instance .. end lemma bdd_above_multiset_imp_bdd_above_count: assumes "subset_mset.bdd_above (A :: 'a multiset set)" shows "bdd_above ((\X. count X x) ` A)" proof - from assms obtain Y where Y: "\X\A. X \# Y" by (auto simp: subset_mset.bdd_above_def) hence "count X x \ count Y x" if "X \ A" for X using that by (auto intro: mset_subset_eq_count) thus ?thesis by (intro bdd_aboveI[of _ "count Y x"]) auto qed lemma bdd_above_multiset_imp_finite_support: assumes "A \ {}" "subset_mset.bdd_above (A :: 'a multiset set)" shows "finite (\X\A. {x. count X x > 0})" proof - from assms obtain Y where Y: "\X\A. X \# Y" by (auto simp: subset_mset.bdd_above_def) hence "count X x \ count Y x" if "X \ A" for X x using that by (auto intro: mset_subset_eq_count) hence "(\X\A. {x. count X x > 0}) \ {x. count Y x > 0}" by safe (erule less_le_trans) moreover have "finite \" by simp ultimately show ?thesis by (rule finite_subset) qed lemma Sup_multiset_in_multiset: assumes "A \ {}" "subset_mset.bdd_above A" shows "(\i. SUP X\A. count X i) \ multiset" unfolding multiset_def proof have "{i. Sup ((\X. count X i) ` A) > 0} \ (\X\A. {i. 0 < count X i})" proof safe fix i assume pos: "(SUP X\A. count X i) > 0" show "i \ (\X\A. {i. 0 < count X i})" proof (rule ccontr) assume "i \ (\X\A. {i. 0 < count X i})" hence "\X\A. count X i \ 0" by (auto simp: count_eq_zero_iff) with assms have "(SUP X\A. count X i) \ 0" by (intro cSup_least bdd_above_multiset_imp_bdd_above_count) auto with pos show False by simp qed qed moreover from assms have "finite \" by (rule bdd_above_multiset_imp_finite_support) ultimately show "finite {i. Sup ((\X. count X i) ` A) > 0}" by (rule finite_subset) qed lemma count_Sup_multiset_nonempty: assumes "A \ {}" "subset_mset.bdd_above A" shows "count (Sup A) x = (SUP X\A. count X x)" using assms by (simp add: Sup_multiset_def Abs_multiset_inverse Sup_multiset_in_multiset) interpretation subset_mset: conditionally_complete_lattice Inf Sup "(\#)" "(\#)" "(\#)" "(\#)" proof fix X :: "'a multiset" and A assume "X \ A" show "Inf A \# X" proof (rule mset_subset_eqI) fix x from \X \ A\ have "A \ {}" by auto hence "count (Inf A) x = (INF X\A. count X x)" by (simp add: count_Inf_multiset_nonempty) also from \X \ A\ have "\ \ count X x" by (intro cInf_lower) simp_all finally show "count (Inf A) x \ count X x" . qed next fix X :: "'a multiset" and A assume nonempty: "A \ {}" and le: "\Y. Y \ A \ X \# Y" show "X \# Inf A" proof (rule mset_subset_eqI) fix x from nonempty have "count X x \ (INF X\A. count X x)" by (intro cInf_greatest) (auto intro: mset_subset_eq_count le) also from nonempty have "\ = count (Inf A) x" by (simp add: count_Inf_multiset_nonempty) finally show "count X x \ count (Inf A) x" . qed next fix X :: "'a multiset" and A assume X: "X \ A" and bdd: "subset_mset.bdd_above A" show "X \# Sup A" proof (rule mset_subset_eqI) fix x from X have "A \ {}" by auto have "count X x \ (SUP X\A. count X x)" by (intro cSUP_upper X bdd_above_multiset_imp_bdd_above_count bdd) also from count_Sup_multiset_nonempty[OF \A \ {}\ bdd] have "(SUP X\A. count X x) = count (Sup A) x" by simp finally show "count X x \ count (Sup A) x" . qed next fix X :: "'a multiset" and A assume nonempty: "A \ {}" and ge: "\Y. Y \ A \ Y \# X" from ge have bdd: "subset_mset.bdd_above A" by (rule subset_mset.bdd_aboveI[of _ X]) show "Sup A \# X" proof (rule mset_subset_eqI) fix x from count_Sup_multiset_nonempty[OF \A \ {}\ bdd] have "count (Sup A) x = (SUP X\A. count X x)" . also from nonempty have "\ \ count X x" by (intro cSup_least) (auto intro: mset_subset_eq_count ge) finally show "count (Sup A) x \ count X x" . qed qed \ \FIXME: avoid junk stemming from type class interpretation\ lemma set_mset_Inf: assumes "A \ {}" shows "set_mset (Inf A) = (\X\A. set_mset X)" proof safe fix x X assume "x \# Inf A" "X \ A" hence nonempty: "A \ {}" by (auto simp: Inf_multiset_empty) from \x \# Inf A\ have "{#x#} \# Inf A" by auto also from \X \ A\ have "\ \# X" by (rule subset_mset.cInf_lower) simp_all finally show "x \# X" by simp next fix x assume x: "x \ (\X\A. set_mset X)" hence "{#x#} \# X" if "X \ A" for X using that by auto from assms and this have "{#x#} \# Inf A" by (rule subset_mset.cInf_greatest) thus "x \# Inf A" by simp qed lemma in_Inf_multiset_iff: assumes "A \ {}" shows "x \# Inf A \ (\X\A. x \# X)" proof - from assms have "set_mset (Inf A) = (\X\A. set_mset X)" by (rule set_mset_Inf) also have "x \ \ \ (\X\A. x \# X)" by simp finally show ?thesis . qed lemma in_Inf_multisetD: "x \# Inf A \ X \ A \ x \# X" by (subst (asm) in_Inf_multiset_iff) auto lemma set_mset_Sup: assumes "subset_mset.bdd_above A" shows "set_mset (Sup A) = (\X\A. set_mset X)" proof safe fix x assume "x \# Sup A" hence nonempty: "A \ {}" by (auto simp: Sup_multiset_empty) show "x \ (\X\A. set_mset X)" proof (rule ccontr) assume x: "x \ (\X\A. set_mset X)" have "count X x \ count (Sup A) x" if "X \ A" for X x using that by (intro mset_subset_eq_count subset_mset.cSup_upper assms) with x have "X \# Sup A - {#x#}" if "X \ A" for X using that by (auto simp: subseteq_mset_def algebra_simps not_in_iff) hence "Sup A \# Sup A - {#x#}" by (intro subset_mset.cSup_least nonempty) with \x \# Sup A\ show False by (auto simp: subseteq_mset_def simp flip: count_greater_zero_iff dest!: spec[of _ x]) qed next fix x X assume "x \ set_mset X" "X \ A" hence "{#x#} \# X" by auto also have "X \# Sup A" by (intro subset_mset.cSup_upper \X \ A\ assms) finally show "x \ set_mset (Sup A)" by simp qed lemma in_Sup_multiset_iff: assumes "subset_mset.bdd_above A" shows "x \# Sup A \ (\X\A. x \# X)" proof - from assms have "set_mset (Sup A) = (\X\A. set_mset X)" by (rule set_mset_Sup) also have "x \ \ \ (\X\A. x \# X)" by simp finally show ?thesis . qed lemma in_Sup_multisetD: assumes "x \# Sup A" shows "\X\A. x \# X" proof - have "subset_mset.bdd_above A" by (rule ccontr) (insert assms, simp_all add: Sup_multiset_unbounded) with assms show ?thesis by (simp add: in_Sup_multiset_iff) qed interpretation subset_mset: distrib_lattice "(\#)" "(\#)" "(\#)" "(\#)" proof fix A B C :: "'a multiset" show "A \# (B \# C) = A \# B \# (A \# C)" by (intro multiset_eqI) simp_all qed \ \FIXME: avoid junk stemming from type class interpretation\ subsubsection \Filter (with comprehension syntax)\ text \Multiset comprehension\ lift_definition filter_mset :: "('a \ bool) \ 'a multiset \ 'a multiset" is "\P M. \x. if P x then M x else 0" by (rule filter_preserves_multiset) syntax (ASCII) "_MCollect" :: "pttrn \ 'a multiset \ bool \ 'a multiset" ("(1{#_ :# _./ _#})") syntax "_MCollect" :: "pttrn \ 'a multiset \ bool \ 'a multiset" ("(1{#_ \# _./ _#})") translations "{#x \# M. P#}" == "CONST filter_mset (\x. P) M" lemma count_filter_mset [simp]: "count (filter_mset P M) a = (if P a then count M a else 0)" by (simp add: filter_mset.rep_eq) lemma set_mset_filter [simp]: "set_mset (filter_mset P M) = {a \ set_mset M. P a}" by (simp only: set_eq_iff count_greater_zero_iff [symmetric] count_filter_mset) simp lemma filter_empty_mset [simp]: "filter_mset P {#} = {#}" by (rule multiset_eqI) simp lemma filter_single_mset: "filter_mset P {#x#} = (if P x then {#x#} else {#})" by (rule multiset_eqI) simp lemma filter_union_mset [simp]: "filter_mset P (M + N) = filter_mset P M + filter_mset P N" by (rule multiset_eqI) simp lemma filter_diff_mset [simp]: "filter_mset P (M - N) = filter_mset P M - filter_mset P N" by (rule multiset_eqI) simp lemma filter_inter_mset [simp]: "filter_mset P (M \# N) = filter_mset P M \# filter_mset P N" by (rule multiset_eqI) simp lemma filter_sup_mset[simp]: "filter_mset P (A \# B) = filter_mset P A \# filter_mset P B" by (rule multiset_eqI) simp lemma filter_mset_add_mset [simp]: "filter_mset P (add_mset x A) = (if P x then add_mset x (filter_mset P A) else filter_mset P A)" by (auto simp: multiset_eq_iff) lemma multiset_filter_subset[simp]: "filter_mset f M \# M" by (simp add: mset_subset_eqI) lemma multiset_filter_mono: assumes "A \# B" shows "filter_mset f A \# filter_mset f B" proof - from assms[unfolded mset_subset_eq_exists_conv] obtain C where B: "B = A + C" by auto show ?thesis unfolding B by auto qed lemma filter_mset_eq_conv: "filter_mset P M = N \ N \# M \ (\b\#N. P b) \ (\a\#M - N. \ P a)" (is "?P \ ?Q") proof assume ?P then show ?Q by auto (simp add: multiset_eq_iff in_diff_count) next assume ?Q then obtain Q where M: "M = N + Q" by (auto simp add: mset_subset_eq_exists_conv) then have MN: "M - N = Q" by simp show ?P proof (rule multiset_eqI) fix a from \?Q\ MN have *: "\ P a \ a \# N" "P a \ a \# Q" by auto show "count (filter_mset P M) a = count N a" proof (cases "a \# M") case True with * show ?thesis by (simp add: not_in_iff M) next case False then have "count M a = 0" by (simp add: not_in_iff) with M show ?thesis by simp qed qed qed lemma filter_filter_mset: "filter_mset P (filter_mset Q M) = {#x \# M. Q x \ P x#}" by (auto simp: multiset_eq_iff) lemma filter_mset_True[simp]: "{#y \# M. True#} = M" and filter_mset_False[simp]: "{#y \# M. False#} = {#}" by (auto simp: multiset_eq_iff) subsubsection \Size\ definition wcount where "wcount f M = (\x. count M x * Suc (f x))" lemma wcount_union: "wcount f (M + N) a = wcount f M a + wcount f N a" by (auto simp: wcount_def add_mult_distrib) lemma wcount_add_mset: "wcount f (add_mset x M) a = (if x = a then Suc (f a) else 0) + wcount f M a" unfolding add_mset_add_single[of _ M] wcount_union by (auto simp: wcount_def) definition size_multiset :: "('a \ nat) \ 'a multiset \ nat" where "size_multiset f M = sum (wcount f M) (set_mset M)" lemmas size_multiset_eq = size_multiset_def[unfolded wcount_def] instantiation multiset :: (type) size begin definition size_multiset where size_multiset_overloaded_def: "size_multiset = Multiset.size_multiset (\_. 0)" instance .. end lemmas size_multiset_overloaded_eq = size_multiset_overloaded_def[THEN fun_cong, unfolded size_multiset_eq, simplified] lemma size_multiset_empty [simp]: "size_multiset f {#} = 0" by (simp add: size_multiset_def) lemma size_empty [simp]: "size {#} = 0" by (simp add: size_multiset_overloaded_def) lemma size_multiset_single : "size_multiset f {#b#} = Suc (f b)" by (simp add: size_multiset_eq) lemma size_single: "size {#b#} = 1" by (simp add: size_multiset_overloaded_def size_multiset_single) lemma sum_wcount_Int: "finite A \ sum (wcount f N) (A \ set_mset N) = sum (wcount f N) A" by (induct rule: finite_induct) (simp_all add: Int_insert_left wcount_def count_eq_zero_iff) lemma size_multiset_union [simp]: "size_multiset f (M + N::'a multiset) = size_multiset f M + size_multiset f N" apply (simp add: size_multiset_def sum_Un_nat sum.distrib sum_wcount_Int wcount_union) apply (subst Int_commute) apply (simp add: sum_wcount_Int) done lemma size_multiset_add_mset [simp]: "size_multiset f (add_mset a M) = Suc (f a) + size_multiset f M" unfolding add_mset_add_single[of _ M] size_multiset_union by (auto simp: size_multiset_single) lemma size_add_mset [simp]: "size (add_mset a A) = Suc (size A)" by (simp add: size_multiset_overloaded_def wcount_add_mset) lemma size_union [simp]: "size (M + N::'a multiset) = size M + size N" by (auto simp add: size_multiset_overloaded_def) lemma size_multiset_eq_0_iff_empty [iff]: "size_multiset f M = 0 \ M = {#}" by (auto simp add: size_multiset_eq count_eq_zero_iff) lemma size_eq_0_iff_empty [iff]: "(size M = 0) = (M = {#})" by (auto simp add: size_multiset_overloaded_def) lemma nonempty_has_size: "(S \ {#}) = (0 < size S)" by (metis gr0I gr_implies_not0 size_empty size_eq_0_iff_empty) lemma size_eq_Suc_imp_elem: "size M = Suc n \ \a. a \# M" apply (unfold size_multiset_overloaded_eq) apply (drule sum_SucD) apply auto done lemma size_eq_Suc_imp_eq_union: assumes "size M = Suc n" shows "\a N. M = add_mset a N" proof - from assms obtain a where "a \# M" by (erule size_eq_Suc_imp_elem [THEN exE]) then have "M = add_mset a (M - {#a#})" by simp then show ?thesis by blast qed lemma size_mset_mono: fixes A B :: "'a multiset" assumes "A \# B" shows "size A \ size B" proof - from assms[unfolded mset_subset_eq_exists_conv] obtain C where B: "B = A + C" by auto show ?thesis unfolding B by (induct C) auto qed lemma size_filter_mset_lesseq[simp]: "size (filter_mset f M) \ size M" by (rule size_mset_mono[OF multiset_filter_subset]) lemma size_Diff_submset: "M \# M' \ size (M' - M) = size M' - size(M::'a multiset)" by (metis add_diff_cancel_left' size_union mset_subset_eq_exists_conv) subsection \Induction and case splits\ theorem multiset_induct [case_names empty add, induct type: multiset]: assumes empty: "P {#}" assumes add: "\x M. P M \ P (add_mset x M)" shows "P M" proof (induct "size M" arbitrary: M) case 0 thus "P M" by (simp add: empty) next case (Suc k) obtain N x where "M = add_mset x N" using \Suc k = size M\ [symmetric] using size_eq_Suc_imp_eq_union by fast with Suc add show "P M" by simp qed lemma multiset_induct_min[case_names empty add]: fixes M :: "'a::linorder multiset" assumes empty: "P {#}" and add: "\x M. P M \ (\y \# M. y \ x) \ P (add_mset x M)" shows "P M" proof (induct "size M" arbitrary: M) case (Suc k) note ih = this(1) and Sk_eq_sz_M = this(2) let ?y = "Min_mset M" let ?N = "M - {#?y#}" have M: "M = add_mset ?y ?N" by (metis Min_in Sk_eq_sz_M finite_set_mset insert_DiffM lessI not_less_zero set_mset_eq_empty_iff size_empty) show ?case by (subst M, rule add, rule ih, metis M Sk_eq_sz_M nat.inject size_add_mset, meson Min_le finite_set_mset in_diffD) qed (simp add: empty) lemma multiset_induct_max[case_names empty add]: fixes M :: "'a::linorder multiset" assumes empty: "P {#}" and add: "\x M. P M \ (\y \# M. y \ x) \ P (add_mset x M)" shows "P M" proof (induct "size M" arbitrary: M) case (Suc k) note ih = this(1) and Sk_eq_sz_M = this(2) let ?y = "Max_mset M" let ?N = "M - {#?y#}" have M: "M = add_mset ?y ?N" by (metis Max_in Sk_eq_sz_M finite_set_mset insert_DiffM lessI not_less_zero set_mset_eq_empty_iff size_empty) show ?case by (subst M, rule add, rule ih, metis M Sk_eq_sz_M nat.inject size_add_mset, meson Max_ge finite_set_mset in_diffD) qed (simp add: empty) lemma multi_nonempty_split: "M \ {#} \ \A a. M = add_mset a A" by (induct M) auto lemma multiset_cases [cases type]: obtains (empty) "M = {#}" | (add) x N where "M = add_mset x N" by (induct M) simp_all lemma multi_drop_mem_not_eq: "c \# B \ B - {#c#} \ B" by (cases "B = {#}") (auto dest: multi_member_split) lemma union_filter_mset_complement[simp]: "\x. P x = (\ Q x) \ filter_mset P M + filter_mset Q M = M" by (subst multiset_eq_iff) auto lemma multiset_partition: "M = {#x \# M. P x#} + {#x \# M. \ P x#}" by simp lemma mset_subset_size: "A \# B \ size A < size B" proof (induct A arbitrary: B) case empty then show ?case using nonempty_has_size by auto next case (add x A) have "add_mset x A \# B" by (meson add.prems subset_mset_def) then show ?case by (metis (no_types) add.prems add.right_neutral add_diff_cancel_left' leD nat_neq_iff size_Diff_submset size_eq_0_iff_empty size_mset_mono subset_mset.le_iff_add subset_mset_def) qed lemma size_1_singleton_mset: "size M = 1 \ \a. M = {#a#}" by (cases M) auto subsubsection \Strong induction and subset induction for multisets\ text \Well-foundedness of strict subset relation\ lemma wf_subset_mset_rel: "wf {(M, N :: 'a multiset). M \# N}" apply (rule wf_measure [THEN wf_subset, where f1=size]) apply (clarsimp simp: measure_def inv_image_def mset_subset_size) done lemma full_multiset_induct [case_names less]: assumes ih: "\B. \(A::'a multiset). A \# B \ P A \ P B" shows "P B" apply (rule wf_subset_mset_rel [THEN wf_induct]) apply (rule ih, auto) done lemma multi_subset_induct [consumes 2, case_names empty add]: assumes "F \# A" and empty: "P {#}" and insert: "\a F. a \# A \ P F \ P (add_mset a F)" shows "P F" proof - from \F \# A\ show ?thesis proof (induct F) show "P {#}" by fact next fix x F assume P: "F \# A \ P F" and i: "add_mset x F \# A" show "P (add_mset x F)" proof (rule insert) from i show "x \# A" by (auto dest: mset_subset_eq_insertD) from i have "F \# A" by (auto dest: mset_subset_eq_insertD) with P show "P F" . qed qed qed subsection \The fold combinator\ definition fold_mset :: "('a \ 'b \ 'b) \ 'b \ 'a multiset \ 'b" where "fold_mset f s M = Finite_Set.fold (\x. f x ^^ count M x) s (set_mset M)" lemma fold_mset_empty [simp]: "fold_mset f s {#} = s" by (simp add: fold_mset_def) context comp_fun_commute begin lemma fold_mset_add_mset [simp]: "fold_mset f s (add_mset x M) = f x (fold_mset f s M)" proof - interpret mset: comp_fun_commute "\y. f y ^^ count M y" by (fact comp_fun_commute_funpow) interpret mset_union: comp_fun_commute "\y. f y ^^ count (add_mset x M) y" by (fact comp_fun_commute_funpow) show ?thesis proof (cases "x \ set_mset M") case False then have *: "count (add_mset x M) x = 1" by (simp add: not_in_iff) from False have "Finite_Set.fold (\y. f y ^^ count (add_mset x M) y) s (set_mset M) = Finite_Set.fold (\y. f y ^^ count M y) s (set_mset M)" by (auto intro!: Finite_Set.fold_cong comp_fun_commute_funpow) with False * show ?thesis by (simp add: fold_mset_def del: count_add_mset) next case True define N where "N = set_mset M - {x}" from N_def True have *: "set_mset M = insert x N" "x \ N" "finite N" by auto then have "Finite_Set.fold (\y. f y ^^ count (add_mset x M) y) s N = Finite_Set.fold (\y. f y ^^ count M y) s N" by (auto intro!: Finite_Set.fold_cong comp_fun_commute_funpow) with * show ?thesis by (simp add: fold_mset_def del: count_add_mset) simp qed qed corollary fold_mset_single: "fold_mset f s {#x#} = f x s" by simp lemma fold_mset_fun_left_comm: "f x (fold_mset f s M) = fold_mset f (f x s) M" by (induct M) (simp_all add: fun_left_comm) lemma fold_mset_union [simp]: "fold_mset f s (M + N) = fold_mset f (fold_mset f s M) N" by (induct M) (simp_all add: fold_mset_fun_left_comm) lemma fold_mset_fusion: assumes "comp_fun_commute g" and *: "\x y. h (g x y) = f x (h y)" shows "h (fold_mset g w A) = fold_mset f (h w) A" proof - interpret comp_fun_commute g by (fact assms) from * show ?thesis by (induct A) auto qed end lemma union_fold_mset_add_mset: "A + B = fold_mset add_mset A B" proof - interpret comp_fun_commute add_mset by standard auto show ?thesis by (induction B) auto qed text \ A note on code generation: When defining some function containing a subterm \<^term>\fold_mset F\, code generation is not automatic. When interpreting locale \left_commutative\ with \F\, the would be code thms for \<^const>\fold_mset\ become thms like \<^term>\fold_mset F z {#} = z\ where \F\ is not a pattern but contains defined symbols, i.e.\ is not a code thm. Hence a separate constant with its own code thms needs to be introduced for \F\. See the image operator below. \ subsection \Image\ definition image_mset :: "('a \ 'b) \ 'a multiset \ 'b multiset" where "image_mset f = fold_mset (add_mset \ f) {#}" lemma comp_fun_commute_mset_image: "comp_fun_commute (add_mset \ f)" by unfold_locales (simp add: fun_eq_iff) lemma image_mset_empty [simp]: "image_mset f {#} = {#}" by (simp add: image_mset_def) lemma image_mset_single: "image_mset f {#x#} = {#f x#}" by (simp add: comp_fun_commute.fold_mset_add_mset comp_fun_commute_mset_image image_mset_def) lemma image_mset_union [simp]: "image_mset f (M + N) = image_mset f M + image_mset f N" proof - interpret comp_fun_commute "add_mset \ f" by (fact comp_fun_commute_mset_image) show ?thesis by (induct N) (simp_all add: image_mset_def) qed corollary image_mset_add_mset [simp]: "image_mset f (add_mset a M) = add_mset (f a) (image_mset f M)" unfolding image_mset_union add_mset_add_single[of a M] by (simp add: image_mset_single) lemma set_image_mset [simp]: "set_mset (image_mset f M) = image f (set_mset M)" by (induct M) simp_all lemma size_image_mset [simp]: "size (image_mset f M) = size M" by (induct M) simp_all lemma image_mset_is_empty_iff [simp]: "image_mset f M = {#} \ M = {#}" by (cases M) auto lemma image_mset_If: "image_mset (\x. if P x then f x else g x) A = image_mset f (filter_mset P A) + image_mset g (filter_mset (\x. \P x) A)" by (induction A) auto lemma image_mset_Diff: assumes "B \# A" shows "image_mset f (A - B) = image_mset f A - image_mset f B" proof - have "image_mset f (A - B + B) = image_mset f (A - B) + image_mset f B" by simp also from assms have "A - B + B = A" by (simp add: subset_mset.diff_add) finally show ?thesis by simp qed lemma count_image_mset: "count (image_mset f A) x = (\y\f -` {x} \ set_mset A. count A y)" proof (induction A) case empty then show ?case by simp next case (add x A) moreover have *: "(if x = y then Suc n else n) = n + (if x = y then 1 else 0)" for n y by simp ultimately show ?case by (auto simp: sum.distrib intro!: sum.mono_neutral_left) qed lemma image_mset_subseteq_mono: "A \# B \ image_mset f A \# image_mset f B" by (metis image_mset_union subset_mset.le_iff_add) lemma image_mset_subset_mono: "M \# N \ image_mset f M \# image_mset f N" by (metis (no_types) Diff_eq_empty_iff_mset image_mset_Diff image_mset_is_empty_iff image_mset_subseteq_mono subset_mset.less_le_not_le) syntax (ASCII) "_comprehension_mset" :: "'a \ 'b \ 'b multiset \ 'a multiset" ("({#_/. _ :# _#})") syntax "_comprehension_mset" :: "'a \ 'b \ 'b multiset \ 'a multiset" ("({#_/. _ \# _#})") translations "{#e. x \# M#}" \ "CONST image_mset (\x. e) M" syntax (ASCII) "_comprehension_mset'" :: "'a \ 'b \ 'b multiset \ bool \ 'a multiset" ("({#_/ | _ :# _./ _#})") syntax "_comprehension_mset'" :: "'a \ 'b \ 'b multiset \ bool \ 'a multiset" ("({#_/ | _ \# _./ _#})") translations "{#e | x\#M. P#}" \ "{#e. x \# {# x\#M. P#}#}" text \ This allows to write not just filters like \<^term>\{#x\#M. x but also images like \<^term>\{#x+x. x\#M #}\ and @{term [source] "{#x+x|x\#M. x\{#x+x|x\#M. x. \ lemma in_image_mset: "y \# {#f x. x \# M#} \ y \ f ` set_mset M" by simp functor image_mset: image_mset proof - fix f g show "image_mset f \ image_mset g = image_mset (f \ g)" proof fix A show "(image_mset f \ image_mset g) A = image_mset (f \ g) A" by (induct A) simp_all qed show "image_mset id = id" proof fix A show "image_mset id A = id A" by (induct A) simp_all qed qed declare image_mset.id [simp] image_mset.identity [simp] lemma image_mset_id[simp]: "image_mset id x = x" unfolding id_def by auto lemma image_mset_cong: "(\x. x \# M \ f x = g x) \ {#f x. x \# M#} = {#g x. x \# M#}" by (induct M) auto lemma image_mset_cong_pair: "(\x y. (x, y) \# M \ f x y = g x y) \ {#f x y. (x, y) \# M#} = {#g x y. (x, y) \# M#}" by (metis image_mset_cong split_cong) lemma image_mset_const_eq: "{#c. a \# M#} = replicate_mset (size M) c" by (induct M) simp_all subsection \Further conversions\ primrec mset :: "'a list \ 'a multiset" where "mset [] = {#}" | "mset (a # x) = add_mset a (mset x)" lemma in_multiset_in_set: "x \# mset xs \ x \ set xs" by (induct xs) simp_all lemma count_mset: "count (mset xs) x = length (filter (\y. x = y) xs)" by (induct xs) simp_all lemma mset_zero_iff[simp]: "(mset x = {#}) = (x = [])" by (induct x) auto lemma mset_zero_iff_right[simp]: "({#} = mset x) = (x = [])" by (induct x) auto lemma count_mset_gt_0: "x \ set xs \ count (mset xs) x > 0" by (induction xs) auto lemma count_mset_0_iff [simp]: "count (mset xs) x = 0 \ x \ set xs" by (induction xs) auto lemma mset_single_iff[iff]: "mset xs = {#x#} \ xs = [x]" by (cases xs) auto lemma mset_single_iff_right[iff]: "{#x#} = mset xs \ xs = [x]" by (cases xs) auto lemma set_mset_mset[simp]: "set_mset (mset xs) = set xs" by (induct xs) auto lemma set_mset_comp_mset [simp]: "set_mset \ mset = set" by (simp add: fun_eq_iff) lemma size_mset [simp]: "size (mset xs) = length xs" by (induct xs) simp_all lemma mset_append [simp]: "mset (xs @ ys) = mset xs + mset ys" by (induct xs arbitrary: ys) auto lemma mset_filter[simp]: "mset (filter P xs) = {#x \# mset xs. P x #}" by (induct xs) simp_all lemma mset_rev [simp]: "mset (rev xs) = mset xs" by (induct xs) simp_all lemma surj_mset: "surj mset" apply (unfold surj_def) apply (rule allI) apply (rule_tac M = y in multiset_induct) apply auto apply (rule_tac x = "x # xa" in exI) apply auto done lemma distinct_count_atmost_1: "distinct x = (\a. count (mset x) a = (if a \ set x then 1 else 0))" proof (induct x) case Nil then show ?case by simp next case (Cons x xs) show ?case (is "?lhs \ ?rhs") proof assume ?lhs then show ?rhs using Cons by simp next assume ?rhs then have "x \ set xs" by (simp split: if_splits) moreover from \?rhs\ have "(\a. count (mset xs) a = (if a \ set xs then 1 else 0))" by (auto split: if_splits simp add: count_eq_zero_iff) ultimately show ?lhs using Cons by simp qed qed lemma mset_eq_setD: assumes "mset xs = mset ys" shows "set xs = set ys" proof - from assms have "set_mset (mset xs) = set_mset (mset ys)" by simp then show ?thesis by simp qed lemma set_eq_iff_mset_eq_distinct: "distinct x \ distinct y \ (set x = set y) = (mset x = mset y)" by (auto simp: multiset_eq_iff distinct_count_atmost_1) lemma set_eq_iff_mset_remdups_eq: "(set x = set y) = (mset (remdups x) = mset (remdups y))" apply (rule iffI) apply (simp add: set_eq_iff_mset_eq_distinct[THEN iffD1]) apply (drule distinct_remdups [THEN distinct_remdups [THEN set_eq_iff_mset_eq_distinct [THEN iffD2]]]) apply simp done lemma nth_mem_mset: "i < length ls \ (ls ! i) \# mset ls" proof (induct ls arbitrary: i) case Nil then show ?case by simp next case Cons then show ?case by (cases i) auto qed lemma mset_remove1[simp]: "mset (remove1 a xs) = mset xs - {#a#}" by (induct xs) (auto simp add: multiset_eq_iff) lemma mset_eq_length: assumes "mset xs = mset ys" shows "length xs = length ys" using assms by (metis size_mset) lemma mset_eq_length_filter: assumes "mset xs = mset ys" shows "length (filter (\x. z = x) xs) = length (filter (\y. z = y) ys)" using assms by (metis count_mset) lemma fold_multiset_equiv: assumes f: "\x y. x \ set xs \ y \ set xs \ f x \ f y = f y \ f x" and equiv: "mset xs = mset ys" shows "List.fold f xs = List.fold f ys" using f equiv [symmetric] proof (induct xs arbitrary: ys) case Nil then show ?case by simp next case (Cons x xs) then have *: "set ys = set (x # xs)" by (blast dest: mset_eq_setD) have "\x y. x \ set ys \ y \ set ys \ f x \ f y = f y \ f x" by (rule Cons.prems(1)) (simp_all add: *) moreover from * have "x \ set ys" by simp ultimately have "List.fold f ys = List.fold f (remove1 x ys) \ f x" by (fact fold_remove1_split) moreover from Cons.prems have "List.fold f xs = List.fold f (remove1 x ys)" by (auto intro: Cons.hyps) ultimately show ?case by simp qed lemma mset_shuffles: "zs \ shuffles xs ys \ mset zs = mset xs + mset ys" by (induction xs ys arbitrary: zs rule: shuffles.induct) auto lemma mset_insort [simp]: "mset (insort x xs) = add_mset x (mset xs)" by (induct xs) simp_all lemma mset_map[simp]: "mset (map f xs) = image_mset f (mset xs)" by (induct xs) simp_all global_interpretation mset_set: folding add_mset "{#}" defines mset_set = "folding.F add_mset {#}" by standard (simp add: fun_eq_iff) lemma sum_multiset_singleton [simp]: "sum (\n. {#n#}) A = mset_set A" by (induction A rule: infinite_finite_induct) auto lemma count_mset_set [simp]: "finite A \ x \ A \ count (mset_set A) x = 1" (is "PROP ?P") "\ finite A \ count (mset_set A) x = 0" (is "PROP ?Q") "x \ A \ count (mset_set A) x = 0" (is "PROP ?R") proof - have *: "count (mset_set A) x = 0" if "x \ A" for A proof (cases "finite A") case False then show ?thesis by simp next case True from True \x \ A\ show ?thesis by (induct A) auto qed then show "PROP ?P" "PROP ?Q" "PROP ?R" by (auto elim!: Set.set_insert) qed \ \TODO: maybe define \<^const>\mset_set\ also in terms of \<^const>\Abs_multiset\\ lemma elem_mset_set[simp, intro]: "finite A \ x \# mset_set A \ x \ A" by (induct A rule: finite_induct) simp_all lemma mset_set_Union: "finite A \ finite B \ A \ B = {} \ mset_set (A \ B) = mset_set A + mset_set B" by (induction A rule: finite_induct) auto lemma filter_mset_mset_set [simp]: "finite A \ filter_mset P (mset_set A) = mset_set {x\A. P x}" proof (induction A rule: finite_induct) case (insert x A) from insert.hyps have "filter_mset P (mset_set (insert x A)) = filter_mset P (mset_set A) + mset_set (if P x then {x} else {})" by simp also have "filter_mset P (mset_set A) = mset_set {x\A. P x}" by (rule insert.IH) also from insert.hyps have "\ + mset_set (if P x then {x} else {}) = mset_set ({x \ A. P x} \ (if P x then {x} else {}))" (is "_ = mset_set ?A") by (intro mset_set_Union [symmetric]) simp_all also from insert.hyps have "?A = {y\insert x A. P y}" by auto finally show ?case . qed simp_all lemma mset_set_Diff: assumes "finite A" "B \ A" shows "mset_set (A - B) = mset_set A - mset_set B" proof - from assms have "mset_set ((A - B) \ B) = mset_set (A - B) + mset_set B" by (intro mset_set_Union) (auto dest: finite_subset) also from assms have "A - B \ B = A" by blast finally show ?thesis by simp qed lemma mset_set_set: "distinct xs \ mset_set (set xs) = mset xs" by (induction xs) simp_all lemma count_mset_set': "count (mset_set A) x = (if finite A \ x \ A then 1 else 0)" by auto lemma subset_imp_msubset_mset_set: assumes "A \ B" "finite B" shows "mset_set A \# mset_set B" proof (rule mset_subset_eqI) fix x :: 'a from assms have "finite A" by (rule finite_subset) with assms show "count (mset_set A) x \ count (mset_set B) x" by (cases "x \ A"; cases "x \ B") auto qed lemma mset_set_set_mset_msubset: "mset_set (set_mset A) \# A" proof (rule mset_subset_eqI) fix x show "count (mset_set (set_mset A)) x \ count A x" by (cases "x \# A") simp_all qed context linorder begin definition sorted_list_of_multiset :: "'a multiset \ 'a list" where "sorted_list_of_multiset M = fold_mset insort [] M" lemma sorted_list_of_multiset_empty [simp]: "sorted_list_of_multiset {#} = []" by (simp add: sorted_list_of_multiset_def) lemma sorted_list_of_multiset_singleton [simp]: "sorted_list_of_multiset {#x#} = [x]" proof - interpret comp_fun_commute insort by (fact comp_fun_commute_insort) show ?thesis by (simp add: sorted_list_of_multiset_def) qed lemma sorted_list_of_multiset_insert [simp]: "sorted_list_of_multiset (add_mset x M) = List.insort x (sorted_list_of_multiset M)" proof - interpret comp_fun_commute insort by (fact comp_fun_commute_insort) show ?thesis by (simp add: sorted_list_of_multiset_def) qed end lemma mset_sorted_list_of_multiset[simp]: "mset (sorted_list_of_multiset M) = M" by (induct M) simp_all lemma sorted_list_of_multiset_mset[simp]: "sorted_list_of_multiset (mset xs) = sort xs" by (induct xs) simp_all lemma finite_set_mset_mset_set[simp]: "finite A \ set_mset (mset_set A) = A" by auto lemma mset_set_empty_iff: "mset_set A = {#} \ A = {} \ infinite A" using finite_set_mset_mset_set by fastforce lemma infinite_set_mset_mset_set: "\ finite A \ set_mset (mset_set A) = {}" by simp lemma set_sorted_list_of_multiset [simp]: "set (sorted_list_of_multiset M) = set_mset M" by (induct M) (simp_all add: set_insort_key) lemma sorted_list_of_mset_set [simp]: "sorted_list_of_multiset (mset_set A) = sorted_list_of_set A" by (cases "finite A") (induct A rule: finite_induct, simp_all) lemma mset_upt [simp]: "mset [m.. {#the (map_of xs i). i \# mset (map fst xs)#} = mset (map snd xs)" proof (induction xs) case (Cons x xs) have "{#the (map_of (x # xs) i). i \# mset (map fst (x # xs))#} = add_mset (snd x) {#the (if i = fst x then Some (snd x) else map_of xs i). i \# mset (map fst xs)#}" (is "_ = add_mset _ ?A") by simp also from Cons.prems have "?A = {#the (map_of xs i). i :# mset (map fst xs)#}" by (cases x, intro image_mset_cong) (auto simp: in_multiset_in_set) also from Cons.prems have "\ = mset (map snd xs)" by (intro Cons.IH) simp_all finally show ?case by simp qed simp_all lemma msubset_mset_set_iff[simp]: assumes "finite A" "finite B" shows "mset_set A \# mset_set B \ A \ B" using assms set_mset_mono subset_imp_msubset_mset_set by fastforce lemma mset_set_eq_iff[simp]: assumes "finite A" "finite B" shows "mset_set A = mset_set B \ A = B" using assms by (fastforce dest: finite_set_mset_mset_set) lemma image_mset_mset_set: \<^marker>\contributor \Lukas Bulwahn\\ assumes "inj_on f A" shows "image_mset f (mset_set A) = mset_set (f ` A)" proof cases assume "finite A" from this \inj_on f A\ show ?thesis by (induct A) auto next assume "infinite A" from this \inj_on f A\ have "infinite (f ` A)" using finite_imageD by blast from \infinite A\ \infinite (f ` A)\ show ?thesis by simp qed subsection \More properties of the replicate and repeat operations\ lemma in_replicate_mset[simp]: "x \# replicate_mset n y \ n > 0 \ x = y" unfolding replicate_mset_def by (induct n) auto lemma set_mset_replicate_mset_subset[simp]: "set_mset (replicate_mset n x) = (if n = 0 then {} else {x})" by (auto split: if_splits) lemma size_replicate_mset[simp]: "size (replicate_mset n M) = n" by (induct n, simp_all) lemma count_le_replicate_mset_subset_eq: "n \ count M x \ replicate_mset n x \# M" by (auto simp add: mset_subset_eqI) (metis count_replicate_mset subseteq_mset_def) lemma filter_eq_replicate_mset: "{#y \# D. y = x#} = replicate_mset (count D x) x" by (induct D) simp_all lemma replicate_count_mset_eq_filter_eq: "replicate (count (mset xs) k) k = filter (HOL.eq k) xs" by (induct xs) auto lemma replicate_mset_eq_empty_iff [simp]: "replicate_mset n a = {#} \ n = 0" by (induct n) simp_all lemma replicate_mset_eq_iff: "replicate_mset m a = replicate_mset n b \ m = 0 \ n = 0 \ m = n \ a = b" by (auto simp add: multiset_eq_iff) lemma repeat_mset_cancel1: "repeat_mset a A = repeat_mset a B \ A = B \ a = 0" by (auto simp: multiset_eq_iff) lemma repeat_mset_cancel2: "repeat_mset a A = repeat_mset b A \ a = b \ A = {#}" by (auto simp: multiset_eq_iff) lemma repeat_mset_eq_empty_iff: "repeat_mset n A = {#} \ n = 0 \ A = {#}" by (cases n) auto lemma image_replicate_mset [simp]: "image_mset f (replicate_mset n a) = replicate_mset n (f a)" by (induct n) simp_all lemma replicate_mset_msubseteq_iff: "replicate_mset m a \# replicate_mset n b \ m = 0 \ a = b \ m \ n" by (cases m) (auto simp: insert_subset_eq_iff simp flip: count_le_replicate_mset_subset_eq) lemma msubseteq_replicate_msetE: assumes "A \# replicate_mset n a" obtains m where "m \ n" and "A = replicate_mset m a" proof (cases "n = 0") case True with assms that show thesis by simp next case False from assms have "set_mset A \ set_mset (replicate_mset n a)" by (rule set_mset_mono) with False have "set_mset A \ {a}" by simp then have "\m. A = replicate_mset m a" proof (induction A) case empty then show ?case by simp next case (add b A) then obtain m where "A = replicate_mset m a" by auto with add.prems show ?case by (auto intro: exI [of _ "Suc m"]) qed then obtain m where A: "A = replicate_mset m a" .. with assms have "m \ n" by (auto simp add: replicate_mset_msubseteq_iff) then show thesis using A .. qed subsection \Big operators\ locale comm_monoid_mset = comm_monoid begin interpretation comp_fun_commute f by standard (simp add: fun_eq_iff left_commute) interpretation comp?: comp_fun_commute "f \ g" by (fact comp_comp_fun_commute) context begin definition F :: "'a multiset \ 'a" where eq_fold: "F M = fold_mset f \<^bold>1 M" lemma empty [simp]: "F {#} = \<^bold>1" by (simp add: eq_fold) lemma singleton [simp]: "F {#x#} = x" proof - interpret comp_fun_commute by standard (simp add: fun_eq_iff left_commute) show ?thesis by (simp add: eq_fold) qed lemma union [simp]: "F (M + N) = F M \<^bold>* F N" proof - interpret comp_fun_commute f by standard (simp add: fun_eq_iff left_commute) show ?thesis by (induct N) (simp_all add: left_commute eq_fold) qed lemma add_mset [simp]: "F (add_mset x N) = x \<^bold>* F N" unfolding add_mset_add_single[of x N] union by (simp add: ac_simps) lemma insert [simp]: shows "F (image_mset g (add_mset x A)) = g x \<^bold>* F (image_mset g A)" by (simp add: eq_fold) lemma remove: assumes "x \# A" shows "F A = x \<^bold>* F (A - {#x#})" using multi_member_split[OF assms] by auto lemma neutral: "\x\#A. x = \<^bold>1 \ F A = \<^bold>1" by (induct A) simp_all lemma neutral_const [simp]: "F (image_mset (\_. \<^bold>1) A) = \<^bold>1" by (simp add: neutral) private lemma F_image_mset_product: "F {#g x j \<^bold>* F {#g i j. i \# A#}. j \# B#} = F (image_mset (g x) B) \<^bold>* F {#F {#g i j. i \# A#}. j \# B#}" by (induction B) (simp_all add: left_commute semigroup.assoc semigroup_axioms) lemma swap: "F (image_mset (\i. F (image_mset (g i) B)) A) = F (image_mset (\j. F (image_mset (\i. g i j) A)) B)" apply (induction A, simp) apply (induction B, auto simp add: F_image_mset_product ac_simps) done lemma distrib: "F (image_mset (\x. g x \<^bold>* h x) A) = F (image_mset g A) \<^bold>* F (image_mset h A)" by (induction A) (auto simp: ac_simps) lemma union_disjoint: "A \# B = {#} \ F (image_mset g (A \# B)) = F (image_mset g A) \<^bold>* F (image_mset g B)" by (induction A) (auto simp: ac_simps) end end lemma comp_fun_commute_plus_mset[simp]: "comp_fun_commute ((+) :: 'a multiset \ _ \ _)" by standard (simp add: add_ac comp_def) declare comp_fun_commute.fold_mset_add_mset[OF comp_fun_commute_plus_mset, simp] lemma in_mset_fold_plus_iff[iff]: "x \# fold_mset (+) M NN \ x \# M \ (\N. N \# NN \ x \# N)" by (induct NN) auto context comm_monoid_add begin sublocale sum_mset: comm_monoid_mset plus 0 defines sum_mset = sum_mset.F .. lemma sum_unfold_sum_mset: "sum f A = sum_mset (image_mset f (mset_set A))" by (cases "finite A") (induct A rule: finite_induct, simp_all) end syntax (ASCII) "_sum_mset_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_add" ("(3SUM _:#_. _)" [0, 51, 10] 10) syntax "_sum_mset_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_add" ("(3\_\#_. _)" [0, 51, 10] 10) translations "\i \# A. b" \ "CONST sum_mset (CONST image_mset (\i. b) A)" context comm_monoid_add begin lemma sum_mset_sum_list: "sum_mset (mset xs) = sum_list xs" by (induction xs) auto end context canonically_ordered_monoid_add begin lemma sum_mset_0_iff [simp]: "sum_mset M = 0 \ (\x \ set_mset M. x = 0)" by (induction M) auto end context ordered_comm_monoid_add begin lemma sum_mset_mono: "sum_mset (image_mset f K) \ sum_mset (image_mset g K)" if "\i. i \# K \ f i \ g i" using that by (induction K) (simp_all add: add_mono) end context ordered_cancel_comm_monoid_diff begin lemma sum_mset_diff: "sum_mset (M - N) = sum_mset M - sum_mset N" if "N \# M" for M N :: "'a multiset" using that by (auto simp add: subset_mset.le_iff_add) end context semiring_0 begin lemma sum_mset_distrib_left: "c * (\x \# M. f x) = (\x \# M. c * f(x))" by (induction M) (simp_all add: algebra_simps) lemma sum_mset_distrib_right: "(\x \# M. f x) * c = (\x \# M. f x * c)" by (induction M) (simp_all add: algebra_simps) end lemma sum_mset_product: fixes f :: "'a::{comm_monoid_add,times} \ 'b::semiring_0" shows "(\i \# A. f i) * (\i \# B. g i) = (\i\#A. \j\#B. f i * g j)" by (subst sum_mset.swap) (simp add: sum_mset_distrib_left sum_mset_distrib_right) context semiring_1 begin lemma sum_mset_replicate_mset [simp]: "sum_mset (replicate_mset n a) = of_nat n * a" by (induction n) (simp_all add: algebra_simps) lemma sum_mset_delta: "sum_mset (image_mset (\x. if x = y then c else 0) A) = c * of_nat (count A y)" by (induction A) (simp_all add: algebra_simps) lemma sum_mset_delta': "sum_mset (image_mset (\x. if y = x then c else 0) A) = c * of_nat (count A y)" by (induction A) (simp_all add: algebra_simps) end lemma of_nat_sum_mset [simp]: "of_nat (sum_mset A) = sum_mset (image_mset of_nat A)" by (induction A) auto lemma size_eq_sum_mset: "size M = (\a\#M. 1)" using image_mset_const_eq [of "1::nat" M] by simp lemma size_mset_set [simp]: "size (mset_set A) = card A" by (simp only: size_eq_sum_mset card_eq_sum sum_unfold_sum_mset) lemma sum_mset_constant [simp]: fixes y :: "'b::semiring_1" shows \(\x\#A. y) = of_nat (size A) * y\ by (induction A) (auto simp: algebra_simps) abbreviation Union_mset :: "'a multiset multiset \ 'a multiset" ("\#") where "\# MM \ sum_mset MM" \ \FIXME ambiguous notation -- could likewise refer to \\#\\ lemma set_mset_Union_mset[simp]: "set_mset (\# MM) = (\M \ set_mset MM. set_mset M)" by (induct MM) auto lemma in_Union_mset_iff[iff]: "x \# \# MM \ (\M. M \# MM \ x \# M)" by (induct MM) auto lemma count_sum: "count (sum f A) x = sum (\a. count (f a) x) A" by (induct A rule: infinite_finite_induct) simp_all lemma sum_eq_empty_iff: assumes "finite A" shows "sum f A = {#} \ (\a\A. f a = {#})" using assms by induct simp_all lemma Union_mset_empty_conv[simp]: "\# M = {#} \ (\i\#M. i = {#})" by (induction M) auto lemma Union_image_single_mset[simp]: "\# (image_mset (\x. {#x#}) m) = m" by(induction m) auto context comm_monoid_mult begin sublocale prod_mset: comm_monoid_mset times 1 defines prod_mset = prod_mset.F .. lemma prod_mset_empty: "prod_mset {#} = 1" by (fact prod_mset.empty) lemma prod_mset_singleton: "prod_mset {#x#} = x" by (fact prod_mset.singleton) lemma prod_mset_Un: "prod_mset (A + B) = prod_mset A * prod_mset B" by (fact prod_mset.union) lemma prod_mset_prod_list: "prod_mset (mset xs) = prod_list xs" by (induct xs) auto lemma prod_mset_replicate_mset [simp]: "prod_mset (replicate_mset n a) = a ^ n" by (induct n) simp_all lemma prod_unfold_prod_mset: "prod f A = prod_mset (image_mset f (mset_set A))" by (cases "finite A") (induct A rule: finite_induct, simp_all) lemma prod_mset_multiplicity: "prod_mset M = prod (\x. x ^ count M x) (set_mset M)" by (simp add: fold_mset_def prod.eq_fold prod_mset.eq_fold funpow_times_power comp_def) lemma prod_mset_delta: "prod_mset (image_mset (\x. if x = y then c else 1) A) = c ^ count A y" by (induction A) simp_all lemma prod_mset_delta': "prod_mset (image_mset (\x. if y = x then c else 1) A) = c ^ count A y" by (induction A) simp_all lemma prod_mset_subset_imp_dvd: assumes "A \# B" shows "prod_mset A dvd prod_mset B" proof - from assms have "B = (B - A) + A" by (simp add: subset_mset.diff_add) also have "prod_mset \ = prod_mset (B - A) * prod_mset A" by simp also have "prod_mset A dvd \" by simp finally show ?thesis . qed lemma dvd_prod_mset: assumes "x \# A" shows "x dvd prod_mset A" using assms prod_mset_subset_imp_dvd [of "{#x#}" A] by simp end syntax (ASCII) "_prod_mset_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_mult" ("(3PROD _:#_. _)" [0, 51, 10] 10) syntax "_prod_mset_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_mult" ("(3\_\#_. _)" [0, 51, 10] 10) translations "\i \# A. b" \ "CONST prod_mset (CONST image_mset (\i. b) A)" lemma prod_mset_constant [simp]: "(\_\#A. c) = c ^ size A" by (simp add: image_mset_const_eq) lemma (in semidom) prod_mset_zero_iff [iff]: "prod_mset A = 0 \ 0 \# A" by (induct A) auto lemma (in semidom_divide) prod_mset_diff: assumes "B \# A" and "0 \# B" shows "prod_mset (A - B) = prod_mset A div prod_mset B" proof - from assms obtain C where "A = B + C" by (metis subset_mset.add_diff_inverse) with assms show ?thesis by simp qed lemma (in semidom_divide) prod_mset_minus: assumes "a \# A" and "a \ 0" shows "prod_mset (A - {#a#}) = prod_mset A div a" using assms prod_mset_diff [of "{#a#}" A] by auto +lemma (in normalization_semidom) normalize_prod_mset_normalize: + "normalize (prod_mset (image_mset normalize A)) = normalize (prod_mset A)" +proof (induction A) + case (add x A) + have "normalize (prod_mset (image_mset normalize (add_mset x A))) = + normalize (x * normalize (prod_mset (image_mset normalize A)))" + by simp + also note add.IH + finally show ?case by simp +qed auto + lemma (in algebraic_semidom) is_unit_prod_mset_iff: "is_unit (prod_mset A) \ (\x \# A. is_unit x)" by (induct A) (auto simp: is_unit_mult_iff) -lemma (in normalization_semidom) normalize_prod_mset: +lemma (in normalization_semidom_multiplicative) normalize_prod_mset: "normalize (prod_mset A) = prod_mset (image_mset normalize A)" by (induct A) (simp_all add: normalize_mult) -lemma (in normalization_semidom) normalized_prod_msetI: +lemma (in normalization_semidom_multiplicative) normalized_prod_msetI: assumes "\a. a \# A \ normalize a = a" shows "normalize (prod_mset A) = prod_mset A" proof - from assms have "image_mset normalize A = A" by (induct A) simp_all then show ?thesis by (simp add: normalize_prod_mset) qed subsection \Alternative representations\ subsubsection \Lists\ context linorder begin lemma mset_insort [simp]: "mset (insort_key k x xs) = add_mset x (mset xs)" by (induct xs) simp_all lemma mset_sort [simp]: "mset (sort_key k xs) = mset xs" by (induct xs) simp_all text \ This lemma shows which properties suffice to show that a function \f\ with \f xs = ys\ behaves like sort. \ lemma properties_for_sort_key: assumes "mset ys = mset xs" and "\k. k \ set ys \ filter (\x. f k = f x) ys = filter (\x. f k = f x) xs" and "sorted (map f ys)" shows "sort_key f xs = ys" using assms proof (induct xs arbitrary: ys) case Nil then show ?case by simp next case (Cons x xs) from Cons.prems(2) have "\k \ set ys. filter (\x. f k = f x) (remove1 x ys) = filter (\x. f k = f x) xs" by (simp add: filter_remove1) with Cons.prems have "sort_key f xs = remove1 x ys" by (auto intro!: Cons.hyps simp add: sorted_map_remove1) moreover from Cons.prems have "x \# mset ys" by auto then have "x \ set ys" by simp ultimately show ?case using Cons.prems by (simp add: insort_key_remove1) qed lemma properties_for_sort: assumes multiset: "mset ys = mset xs" and "sorted ys" shows "sort xs = ys" proof (rule properties_for_sort_key) from multiset show "mset ys = mset xs" . from \sorted ys\ show "sorted (map (\x. x) ys)" by simp from multiset have "length (filter (\y. k = y) ys) = length (filter (\x. k = x) xs)" for k by (rule mset_eq_length_filter) then have "replicate (length (filter (\y. k = y) ys)) k = replicate (length (filter (\x. k = x) xs)) k" for k by simp then show "k \ set ys \ filter (\y. k = y) ys = filter (\x. k = x) xs" for k by (simp add: replicate_length_filter) qed lemma sort_key_inj_key_eq: assumes mset_equal: "mset xs = mset ys" and "inj_on f (set xs)" and "sorted (map f ys)" shows "sort_key f xs = ys" proof (rule properties_for_sort_key) from mset_equal show "mset ys = mset xs" by simp from \sorted (map f ys)\ show "sorted (map f ys)" . show "[x\ys . f k = f x] = [x\xs . f k = f x]" if "k \ set ys" for k proof - from mset_equal have set_equal: "set xs = set ys" by (rule mset_eq_setD) with that have "insert k (set ys) = set ys" by auto with \inj_on f (set xs)\ have inj: "inj_on f (insert k (set ys))" by (simp add: set_equal) from inj have "[x\ys . f k = f x] = filter (HOL.eq k) ys" by (auto intro!: inj_on_filter_key_eq) also have "\ = replicate (count (mset ys) k) k" by (simp add: replicate_count_mset_eq_filter_eq) also have "\ = replicate (count (mset xs) k) k" using mset_equal by simp also have "\ = filter (HOL.eq k) xs" by (simp add: replicate_count_mset_eq_filter_eq) also have "\ = [x\xs . f k = f x]" using inj by (auto intro!: inj_on_filter_key_eq [symmetric] simp add: set_equal) finally show ?thesis . qed qed lemma sort_key_eq_sort_key: assumes "mset xs = mset ys" and "inj_on f (set xs)" shows "sort_key f xs = sort_key f ys" by (rule sort_key_inj_key_eq) (simp_all add: assms) lemma sort_key_by_quicksort: "sort_key f xs = sort_key f [x\xs. f x < f (xs ! (length xs div 2))] @ [x\xs. f x = f (xs ! (length xs div 2))] @ sort_key f [x\xs. f x > f (xs ! (length xs div 2))]" (is "sort_key f ?lhs = ?rhs") proof (rule properties_for_sort_key) show "mset ?rhs = mset ?lhs" by (rule multiset_eqI) auto show "sorted (map f ?rhs)" by (auto simp add: sorted_append intro: sorted_map_same) next fix l assume "l \ set ?rhs" let ?pivot = "f (xs ! (length xs div 2))" have *: "\x. f l = f x \ f x = f l" by auto have "[x \ sort_key f xs . f x = f l] = [x \ xs. f x = f l]" unfolding filter_sort by (rule properties_for_sort_key) (auto intro: sorted_map_same) with * have **: "[x \ sort_key f xs . f l = f x] = [x \ xs. f l = f x]" by simp have "\x P. P (f x) ?pivot \ f l = f x \ P (f l) ?pivot \ f l = f x" by auto then have "\P. [x \ sort_key f xs . P (f x) ?pivot \ f l = f x] = [x \ sort_key f xs. P (f l) ?pivot \ f l = f x]" by simp note *** = this [of "(<)"] this [of "(>)"] this [of "(=)"] show "[x \ ?rhs. f l = f x] = [x \ ?lhs. f l = f x]" proof (cases "f l" ?pivot rule: linorder_cases) case less then have "f l \ ?pivot" and "\ f l > ?pivot" by auto with less show ?thesis by (simp add: filter_sort [symmetric] ** ***) next case equal then show ?thesis by (simp add: * less_le) next case greater then have "f l \ ?pivot" and "\ f l < ?pivot" by auto with greater show ?thesis by (simp add: filter_sort [symmetric] ** ***) qed qed lemma sort_by_quicksort: "sort xs = sort [x\xs. x < xs ! (length xs div 2)] @ [x\xs. x = xs ! (length xs div 2)] @ sort [x\xs. x > xs ! (length xs div 2)]" (is "sort ?lhs = ?rhs") using sort_key_by_quicksort [of "\x. x", symmetric] by simp text \A stable parameterized quicksort\ definition part :: "('b \ 'a) \ 'a \ 'b list \ 'b list \ 'b list \ 'b list" where "part f pivot xs = ([x \ xs. f x < pivot], [x \ xs. f x = pivot], [x \ xs. pivot < f x])" lemma part_code [code]: "part f pivot [] = ([], [], [])" "part f pivot (x # xs) = (let (lts, eqs, gts) = part f pivot xs; x' = f x in if x' < pivot then (x # lts, eqs, gts) else if x' > pivot then (lts, eqs, x # gts) else (lts, x # eqs, gts))" by (auto simp add: part_def Let_def split_def) lemma sort_key_by_quicksort_code [code]: "sort_key f xs = (case xs of [] \ [] | [x] \ xs | [x, y] \ (if f x \ f y then xs else [y, x]) | _ \ let (lts, eqs, gts) = part f (f (xs ! (length xs div 2))) xs in sort_key f lts @ eqs @ sort_key f gts)" proof (cases xs) case Nil then show ?thesis by simp next case (Cons _ ys) note hyps = Cons show ?thesis proof (cases ys) case Nil with hyps show ?thesis by simp next case (Cons _ zs) note hyps = hyps Cons show ?thesis proof (cases zs) case Nil with hyps show ?thesis by auto next case Cons from sort_key_by_quicksort [of f xs] have "sort_key f xs = (let (lts, eqs, gts) = part f (f (xs ! (length xs div 2))) xs in sort_key f lts @ eqs @ sort_key f gts)" by (simp only: split_def Let_def part_def fst_conv snd_conv) with hyps Cons show ?thesis by (simp only: list.cases) qed qed qed end hide_const (open) part lemma mset_remdups_subset_eq: "mset (remdups xs) \# mset xs" by (induct xs) (auto intro: subset_mset.order_trans) lemma mset_update: "i < length ls \ mset (ls[i := v]) = add_mset v (mset ls - {#ls ! i#})" proof (induct ls arbitrary: i) case Nil then show ?case by simp next case (Cons x xs) show ?case proof (cases i) case 0 then show ?thesis by simp next case (Suc i') with Cons show ?thesis by (cases \x = xs ! i'\) auto qed qed lemma mset_swap: "i < length ls \ j < length ls \ mset (ls[j := ls ! i, i := ls ! j]) = mset ls" by (cases "i = j") (simp_all add: mset_update nth_mem_mset) subsection \The multiset order\ subsubsection \Well-foundedness\ definition mult1 :: "('a \ 'a) set \ ('a multiset \ 'a multiset) set" where "mult1 r = {(N, M). \a M0 K. M = add_mset a M0 \ N = M0 + K \ (\b. b \# K \ (b, a) \ r)}" definition mult :: "('a \ 'a) set \ ('a multiset \ 'a multiset) set" where "mult r = (mult1 r)\<^sup>+" lemma mult1I: assumes "M = add_mset a M0" and "N = M0 + K" and "\b. b \# K \ (b, a) \ r" shows "(N, M) \ mult1 r" using assms unfolding mult1_def by blast lemma mult1E: assumes "(N, M) \ mult1 r" obtains a M0 K where "M = add_mset a M0" "N = M0 + K" "\b. b \# K \ (b, a) \ r" using assms unfolding mult1_def by blast lemma mono_mult1: assumes "r \ r'" shows "mult1 r \ mult1 r'" unfolding mult1_def using assms by blast lemma mono_mult: assumes "r \ r'" shows "mult r \ mult r'" unfolding mult_def using mono_mult1[OF assms] trancl_mono by blast lemma not_less_empty [iff]: "(M, {#}) \ mult1 r" by (simp add: mult1_def) lemma less_add: assumes mult1: "(N, add_mset a M0) \ mult1 r" shows "(\M. (M, M0) \ mult1 r \ N = add_mset a M) \ (\K. (\b. b \# K \ (b, a) \ r) \ N = M0 + K)" proof - let ?r = "\K a. \b. b \# K \ (b, a) \ r" let ?R = "\N M. \a M0 K. M = add_mset a M0 \ N = M0 + K \ ?r K a" obtain a' M0' K where M0: "add_mset a M0 = add_mset a' M0'" and N: "N = M0' + K" and r: "?r K a'" using mult1 unfolding mult1_def by auto show ?thesis (is "?case1 \ ?case2") proof - from M0 consider "M0 = M0'" "a = a'" | K' where "M0 = add_mset a' K'" "M0' = add_mset a K'" by atomize_elim (simp only: add_eq_conv_ex) then show ?thesis proof cases case 1 with N r have "?r K a \ N = M0 + K" by simp then have ?case2 .. then show ?thesis .. next case 2 from N 2(2) have n: "N = add_mset a (K' + K)" by simp with r 2(1) have "?R (K' + K) M0" by blast with n have ?case1 by (simp add: mult1_def) then show ?thesis .. qed qed qed lemma all_accessible: assumes "wf r" shows "\M. M \ Wellfounded.acc (mult1 r)" proof let ?R = "mult1 r" let ?W = "Wellfounded.acc ?R" { fix M M0 a assume M0: "M0 \ ?W" and wf_hyp: "\b. (b, a) \ r \ (\M \ ?W. add_mset b M \ ?W)" and acc_hyp: "\M. (M, M0) \ ?R \ add_mset a M \ ?W" have "add_mset a M0 \ ?W" proof (rule accI [of "add_mset a M0"]) fix N assume "(N, add_mset a M0) \ ?R" then consider M where "(M, M0) \ ?R" "N = add_mset a M" | K where "\b. b \# K \ (b, a) \ r" "N = M0 + K" by atomize_elim (rule less_add) then show "N \ ?W" proof cases case 1 from acc_hyp have "(M, M0) \ ?R \ add_mset a M \ ?W" .. from this and \(M, M0) \ ?R\ have "add_mset a M \ ?W" .. then show "N \ ?W" by (simp only: \N = add_mset a M\) next case 2 from this(1) have "M0 + K \ ?W" proof (induct K) case empty from M0 show "M0 + {#} \ ?W" by simp next case (add x K) from add.prems have "(x, a) \ r" by simp with wf_hyp have "\M \ ?W. add_mset x M \ ?W" by blast moreover from add have "M0 + K \ ?W" by simp ultimately have "add_mset x (M0 + K) \ ?W" .. then show "M0 + (add_mset x K) \ ?W" by simp qed then show "N \ ?W" by (simp only: 2(2)) qed qed } note tedious_reasoning = this show "M \ ?W" for M proof (induct M) show "{#} \ ?W" proof (rule accI) fix b assume "(b, {#}) \ ?R" with not_less_empty show "b \ ?W" by contradiction qed fix M a assume "M \ ?W" from \wf r\ have "\M \ ?W. add_mset a M \ ?W" proof induct fix a assume r: "\b. (b, a) \ r \ (\M \ ?W. add_mset b M \ ?W)" show "\M \ ?W. add_mset a M \ ?W" proof fix M assume "M \ ?W" then show "add_mset a M \ ?W" by (rule acc_induct) (rule tedious_reasoning [OF _ r]) qed qed from this and \M \ ?W\ show "add_mset a M \ ?W" .. qed qed theorem wf_mult1: "wf r \ wf (mult1 r)" by (rule acc_wfI) (rule all_accessible) theorem wf_mult: "wf r \ wf (mult r)" unfolding mult_def by (rule wf_trancl) (rule wf_mult1) subsubsection \Closure-free presentation\ text \One direction.\ lemma mult_implies_one_step: assumes trans: "trans r" and MN: "(M, N) \ mult r" shows "\I J K. N = I + J \ M = I + K \ J \ {#} \ (\k \ set_mset K. \j \ set_mset J. (k, j) \ r)" using MN unfolding mult_def mult1_def proof (induction rule: converse_trancl_induct) case (base y) then show ?case by force next case (step y z) note yz = this(1) and zN = this(2) and N_decomp = this(3) obtain I J K where N: "N = I + J" "z = I + K" "J \ {#}" "\k\#K. \j\#J. (k, j) \ r" using N_decomp by blast obtain a M0 K' where z: "z = add_mset a M0" and y: "y = M0 + K'" and K: "\b. b \# K' \ (b, a) \ r" using yz by blast show ?case proof (cases "a \# K") case True moreover have "\j\#J. (k, j) \ r" if "k \# K'" for k using K N trans True by (meson that transE) ultimately show ?thesis by (rule_tac x = I in exI, rule_tac x = J in exI, rule_tac x = "(K - {#a#}) + K'" in exI) (use z y N in \auto simp del: subset_mset.add_diff_assoc2 dest: in_diffD\) next case False then have "a \# I" by (metis N(2) union_iff union_single_eq_member z) moreover have "M0 = I + K - {#a#}" using N(2) z by force ultimately show ?thesis by (rule_tac x = "I - {#a#}" in exI, rule_tac x = "add_mset a J" in exI, rule_tac x = "K + K'" in exI) (use z y N False K in \auto simp: add.assoc\) qed qed lemma one_step_implies_mult: assumes "J \ {#}" and "\k \ set_mset K. \j \ set_mset J. (k, j) \ r" shows "(I + K, I + J) \ mult r" using assms proof (induction "size J" arbitrary: I J K) case 0 then show ?case by auto next case (Suc n) note IH = this(1) and size_J = this(2)[THEN sym] obtain J' a where J: "J = add_mset a J'" using size_J by (blast dest: size_eq_Suc_imp_eq_union) show ?case proof (cases "J' = {#}") case True then show ?thesis using J Suc by (fastforce simp add: mult_def mult1_def) next case [simp]: False have K: "K = {#x \# K. (x, a) \ r#} + {#x \# K. (x, a) \ r#}" by simp have "(I + K, (I + {# x \# K. (x, a) \ r #}) + J') \ mult r" using IH[of J' "{# x \# K. (x, a) \ r#}" "I + {# x \# K. (x, a) \ r#}"] J Suc.prems K size_J by (auto simp: ac_simps) moreover have "(I + {#x \# K. (x, a) \ r#} + J', I + J) \ mult r" by (fastforce simp: J mult1_def mult_def) ultimately show ?thesis unfolding mult_def by simp qed qed lemma subset_implies_mult: assumes sub: "A \# B" shows "(A, B) \ mult r" proof - have ApBmA: "A + (B - A) = B" using sub by simp have BmA: "B - A \ {#}" using sub by (simp add: Diff_eq_empty_iff_mset subset_mset.less_le_not_le) thus ?thesis by (rule one_step_implies_mult[of "B - A" "{#}" _ A, unfolded ApBmA, simplified]) qed subsection \The multiset extension is cancellative for multiset union\ lemma mult_cancel: assumes "trans s" and "irrefl s" shows "(X + Z, Y + Z) \ mult s \ (X, Y) \ mult s" (is "?L \ ?R") proof assume ?L thus ?R proof (induct Z) case (add z Z) obtain X' Y' Z' where *: "add_mset z X + Z = Z' + X'" "add_mset z Y + Z = Z' + Y'" "Y' \ {#}" "\x \ set_mset X'. \y \ set_mset Y'. (x, y) \ s" using mult_implies_one_step[OF \trans s\ add(2)] by auto consider Z2 where "Z' = add_mset z Z2" | X2 Y2 where "X' = add_mset z X2" "Y' = add_mset z Y2" using *(1,2) by (metis add_mset_remove_trivial_If insert_iff set_mset_add_mset_insert union_iff) thus ?case proof (cases) case 1 thus ?thesis using * one_step_implies_mult[of Y' X' s Z2] by (auto simp: add.commute[of _ "{#_#}"] add.assoc intro: add(1)) next case 2 then obtain y where "y \ set_mset Y2" "(z, y) \ s" using *(4) \irrefl s\ by (auto simp: irrefl_def) moreover from this transD[OF \trans s\ _ this(2)] have "x' \ set_mset X2 \ \y \ set_mset Y2. (x', y) \ s" for x' using 2 *(4)[rule_format, of x'] by auto ultimately show ?thesis using * one_step_implies_mult[of Y2 X2 s Z'] 2 by (force simp: add.commute[of "{#_#}"] add.assoc[symmetric] intro: add(1)) qed qed auto next assume ?R then obtain I J K where "Y = I + J" "X = I + K" "J \ {#}" "\k \ set_mset K. \j \ set_mset J. (k, j) \ s" using mult_implies_one_step[OF \trans s\] by blast thus ?L using one_step_implies_mult[of J K s "I + Z"] by (auto simp: ac_simps) qed lemmas mult_cancel_add_mset = mult_cancel[of _ _ "{#_#}", unfolded union_mset_add_mset_right add.comm_neutral] lemma mult_cancel_max: assumes "trans s" and "irrefl s" shows "(X, Y) \ mult s \ (X - X \# Y, Y - X \# Y) \ mult s" (is "?L \ ?R") proof - have "X - X \# Y + X \# Y = X" "Y - X \# Y + X \# Y = Y" by (auto simp flip: count_inject) thus ?thesis using mult_cancel[OF assms, of "X - X \# Y" "X \# Y" "Y - X \# Y"] by auto qed subsection \Quasi-executable version of the multiset extension\ text \ Predicate variants of \mult\ and the reflexive closure of \mult\, which are executable whenever the given predicate \P\ is. Together with the standard code equations for \(\#\) and \(-\) this should yield quadratic (with respect to calls to \P\) implementations of \multp\ and \multeqp\. \ definition multp :: "('a \ 'a \ bool) \ 'a multiset \ 'a multiset \ bool" where "multp P N M = (let Z = M \# N; X = M - Z in X \ {#} \ (let Y = N - Z in (\y \ set_mset Y. \x \ set_mset X. P y x)))" definition multeqp :: "('a \ 'a \ bool) \ 'a multiset \ 'a multiset \ bool" where "multeqp P N M = (let Z = M \# N; X = M - Z; Y = N - Z in (\y \ set_mset Y. \x \ set_mset X. P y x))" lemma multp_iff: assumes "irrefl R" and "trans R" and [simp]: "\x y. P x y \ (x, y) \ R" shows "multp P N M \ (N, M) \ mult R" (is "?L \ ?R") proof - have *: "M \# N + (N - M \# N) = N" "M \# N + (M - M \# N) = M" "(M - M \# N) \# (N - M \# N) = {#}" by (auto simp flip: count_inject) show ?thesis proof assume ?L thus ?R using one_step_implies_mult[of "M - M \# N" "N - M \# N" R "M \# N"] * by (auto simp: multp_def Let_def) next { fix I J K :: "'a multiset" assume "(I + J) \# (I + K) = {#}" then have "I = {#}" by (metis inter_union_distrib_right union_eq_empty) } note [dest!] = this assume ?R thus ?L using mult_implies_one_step[OF assms(2), of "N - M \# N" "M - M \# N"] mult_cancel_max[OF assms(2,1), of "N" "M"] * by (auto simp: multp_def) qed qed lemma multeqp_iff: assumes "irrefl R" and "trans R" and "\x y. P x y \ (x, y) \ R" shows "multeqp P N M \ (N, M) \ (mult R)\<^sup>=" proof - { assume "N \ M" "M - M \# N = {#}" then obtain y where "count N y \ count M y" by (auto simp flip: count_inject) then have "\y. count M y < count N y" using \M - M \# N = {#}\ by (auto simp flip: count_inject dest!: le_neq_implies_less fun_cong[of _ _ y]) } then have "multeqp P N M \ multp P N M \ N = M" by (auto simp: multeqp_def multp_def Let_def in_diff_count) thus ?thesis using multp_iff[OF assms] by simp qed subsubsection \Partial-order properties\ lemma (in preorder) mult1_lessE: assumes "(N, M) \ mult1 {(a, b). a < b}" obtains a M0 K where "M = add_mset a M0" "N = M0 + K" "a \# K" "\b. b \# K \ b < a" proof - from assms obtain a M0 K where "M = add_mset a M0" "N = M0 + K" and *: "b \# K \ b < a" for b by (blast elim: mult1E) moreover from * [of a] have "a \# K" by auto ultimately show thesis by (auto intro: that) qed instantiation multiset :: (preorder) order begin definition less_multiset :: "'a multiset \ 'a multiset \ bool" where "M' < M \ (M', M) \ mult {(x', x). x' < x}" definition less_eq_multiset :: "'a multiset \ 'a multiset \ bool" where "less_eq_multiset M' M \ M' < M \ M' = M" instance proof - have irrefl: "\ M < M" for M :: "'a multiset" proof assume "M < M" then have MM: "(M, M) \ mult {(x, y). x < y}" by (simp add: less_multiset_def) have "trans {(x'::'a, x). x' < x}" by (metis (mono_tags, lifting) case_prodD case_prodI less_trans mem_Collect_eq transI) moreover note MM ultimately have "\I J K. M = I + J \ M = I + K \ J \ {#} \ (\k\set_mset K. \j\set_mset J. (k, j) \ {(x, y). x < y})" by (rule mult_implies_one_step) then obtain I J K where "M = I + J" and "M = I + K" and "J \ {#}" and "(\k\set_mset K. \j\set_mset J. (k, j) \ {(x, y). x < y})" by blast then have *: "K \ {#}" and **: "\k\set_mset K. \j\set_mset K. k < j" by auto have "finite (set_mset K)" by simp moreover note ** ultimately have "set_mset K = {}" by (induct rule: finite_induct) (auto intro: order_less_trans) with * show False by simp qed have trans: "K < M \ M < N \ K < N" for K M N :: "'a multiset" unfolding less_multiset_def mult_def by (blast intro: trancl_trans) show "OFCLASS('a multiset, order_class)" by standard (auto simp add: less_eq_multiset_def irrefl dest: trans) qed end \ \FIXME avoid junk stemming from type class interpretation\ lemma mset_le_irrefl [elim!]: fixes M :: "'a::preorder multiset" shows "M < M \ R" by simp subsubsection \Monotonicity of multiset union\ lemma mult1_union: "(B, D) \ mult1 r \ (C + B, C + D) \ mult1 r" by (force simp: mult1_def) lemma union_le_mono2: "B < D \ C + B < C + (D::'a::preorder multiset)" apply (unfold less_multiset_def mult_def) apply (erule trancl_induct) apply (blast intro: mult1_union) apply (blast intro: mult1_union trancl_trans) done lemma union_le_mono1: "B < D \ B + C < D + (C::'a::preorder multiset)" apply (subst add.commute [of B C]) apply (subst add.commute [of D C]) apply (erule union_le_mono2) done lemma union_less_mono: fixes A B C D :: "'a::preorder multiset" shows "A < C \ B < D \ A + B < C + D" by (blast intro!: union_le_mono1 union_le_mono2 less_trans) instantiation multiset :: (preorder) ordered_ab_semigroup_add begin instance by standard (auto simp add: less_eq_multiset_def intro: union_le_mono2) end subsubsection \Termination proofs with multiset orders\ lemma multi_member_skip: "x \# XS \ x \# {# y #} + XS" and multi_member_this: "x \# {# x #} + XS" and multi_member_last: "x \# {# x #}" by auto definition "ms_strict = mult pair_less" definition "ms_weak = ms_strict \ Id" lemma ms_reduction_pair: "reduction_pair (ms_strict, ms_weak)" unfolding reduction_pair_def ms_strict_def ms_weak_def pair_less_def by (auto intro: wf_mult1 wf_trancl simp: mult_def) lemma smsI: "(set_mset A, set_mset B) \ max_strict \ (Z + A, Z + B) \ ms_strict" unfolding ms_strict_def by (rule one_step_implies_mult) (auto simp add: max_strict_def pair_less_def elim!:max_ext.cases) lemma wmsI: "(set_mset A, set_mset B) \ max_strict \ A = {#} \ B = {#} \ (Z + A, Z + B) \ ms_weak" unfolding ms_weak_def ms_strict_def by (auto simp add: pair_less_def max_strict_def elim!:max_ext.cases intro: one_step_implies_mult) inductive pw_leq where pw_leq_empty: "pw_leq {#} {#}" | pw_leq_step: "\(x,y) \ pair_leq; pw_leq X Y \ \ pw_leq ({#x#} + X) ({#y#} + Y)" lemma pw_leq_lstep: "(x, y) \ pair_leq \ pw_leq {#x#} {#y#}" by (drule pw_leq_step) (rule pw_leq_empty, simp) lemma pw_leq_split: assumes "pw_leq X Y" shows "\A B Z. X = A + Z \ Y = B + Z \ ((set_mset A, set_mset B) \ max_strict \ (B = {#} \ A = {#}))" using assms proof induct case pw_leq_empty thus ?case by auto next case (pw_leq_step x y X Y) then obtain A B Z where [simp]: "X = A + Z" "Y = B + Z" and 1[simp]: "(set_mset A, set_mset B) \ max_strict \ (B = {#} \ A = {#})" by auto from pw_leq_step consider "x = y" | "(x, y) \ pair_less" unfolding pair_leq_def by auto thus ?case proof cases case [simp]: 1 have "{#x#} + X = A + ({#y#}+Z) \ {#y#} + Y = B + ({#y#}+Z) \ ((set_mset A, set_mset B) \ max_strict \ (B = {#} \ A = {#}))" by auto thus ?thesis by blast next case 2 let ?A' = "{#x#} + A" and ?B' = "{#y#} + B" have "{#x#} + X = ?A' + Z" "{#y#} + Y = ?B' + Z" by auto moreover have "(set_mset ?A', set_mset ?B') \ max_strict" using 1 2 unfolding max_strict_def by (auto elim!: max_ext.cases) ultimately show ?thesis by blast qed qed lemma assumes pwleq: "pw_leq Z Z'" shows ms_strictI: "(set_mset A, set_mset B) \ max_strict \ (Z + A, Z' + B) \ ms_strict" and ms_weakI1: "(set_mset A, set_mset B) \ max_strict \ (Z + A, Z' + B) \ ms_weak" and ms_weakI2: "(Z + {#}, Z' + {#}) \ ms_weak" proof - from pw_leq_split[OF pwleq] obtain A' B' Z'' where [simp]: "Z = A' + Z''" "Z' = B' + Z''" and mx_or_empty: "(set_mset A', set_mset B') \ max_strict \ (A' = {#} \ B' = {#})" by blast { assume max: "(set_mset A, set_mset B) \ max_strict" from mx_or_empty have "(Z'' + (A + A'), Z'' + (B + B')) \ ms_strict" proof assume max': "(set_mset A', set_mset B') \ max_strict" with max have "(set_mset (A + A'), set_mset (B + B')) \ max_strict" by (auto simp: max_strict_def intro: max_ext_additive) thus ?thesis by (rule smsI) next assume [simp]: "A' = {#} \ B' = {#}" show ?thesis by (rule smsI) (auto intro: max) qed thus "(Z + A, Z' + B) \ ms_strict" by (simp add: ac_simps) thus "(Z + A, Z' + B) \ ms_weak" by (simp add: ms_weak_def) } from mx_or_empty have "(Z'' + A', Z'' + B') \ ms_weak" by (rule wmsI) thus "(Z + {#}, Z' + {#}) \ ms_weak" by (simp add: ac_simps) qed lemma empty_neutral: "{#} + x = x" "x + {#} = x" and nonempty_plus: "{# x #} + rs \ {#}" and nonempty_single: "{# x #} \ {#}" by auto setup \ let fun msetT T = Type (\<^type_name>\multiset\, [T]); fun mk_mset T [] = Const (\<^const_abbrev>\Mempty\, msetT T) | mk_mset T [x] = Const (\<^const_name>\add_mset\, T --> msetT T --> msetT T) $ x $ Const (\<^const_abbrev>\Mempty\, msetT T) | mk_mset T (x :: xs) = Const (\<^const_name>\plus\, msetT T --> msetT T --> msetT T) $ mk_mset T [x] $ mk_mset T xs fun mset_member_tac ctxt m i = if m <= 0 then resolve_tac ctxt @{thms multi_member_this} i ORELSE resolve_tac ctxt @{thms multi_member_last} i else resolve_tac ctxt @{thms multi_member_skip} i THEN mset_member_tac ctxt (m - 1) i fun mset_nonempty_tac ctxt = resolve_tac ctxt @{thms nonempty_plus} ORELSE' resolve_tac ctxt @{thms nonempty_single} fun regroup_munion_conv ctxt = Function_Lib.regroup_conv ctxt \<^const_abbrev>\Mempty\ \<^const_name>\plus\ (map (fn t => t RS eq_reflection) (@{thms ac_simps} @ @{thms empty_neutral})) fun unfold_pwleq_tac ctxt i = (resolve_tac ctxt @{thms pw_leq_step} i THEN (fn st => unfold_pwleq_tac ctxt (i + 1) st)) ORELSE (resolve_tac ctxt @{thms pw_leq_lstep} i) ORELSE (resolve_tac ctxt @{thms pw_leq_empty} i) val set_mset_simps = [@{thm set_mset_empty}, @{thm set_mset_single}, @{thm set_mset_union}, @{thm Un_insert_left}, @{thm Un_empty_left}] in ScnpReconstruct.multiset_setup (ScnpReconstruct.Multiset { msetT=msetT, mk_mset=mk_mset, mset_regroup_conv=regroup_munion_conv, mset_member_tac=mset_member_tac, mset_nonempty_tac=mset_nonempty_tac, mset_pwleq_tac=unfold_pwleq_tac, set_of_simps=set_mset_simps, smsI'= @{thm ms_strictI}, wmsI2''= @{thm ms_weakI2}, wmsI1= @{thm ms_weakI1}, reduction_pair = @{thm ms_reduction_pair} }) end \ subsection \Legacy theorem bindings\ lemmas multi_count_eq = multiset_eq_iff [symmetric] lemma union_commute: "M + N = N + (M::'a multiset)" by (fact add.commute) lemma union_assoc: "(M + N) + K = M + (N + (K::'a multiset))" by (fact add.assoc) lemma union_lcomm: "M + (N + K) = N + (M + (K::'a multiset))" by (fact add.left_commute) lemmas union_ac = union_assoc union_commute union_lcomm add_mset_commute lemma union_right_cancel: "M + K = N + K \ M = (N::'a multiset)" by (fact add_right_cancel) lemma union_left_cancel: "K + M = K + N \ M = (N::'a multiset)" by (fact add_left_cancel) lemma multi_union_self_other_eq: "(A::'a multiset) + X = A + Y \ X = Y" by (fact add_left_imp_eq) lemma mset_subset_trans: "(M::'a multiset) \# K \ K \# N \ M \# N" by (fact subset_mset.less_trans) lemma multiset_inter_commute: "A \# B = B \# A" by (fact subset_mset.inf.commute) lemma multiset_inter_assoc: "A \# (B \# C) = A \# B \# C" by (fact subset_mset.inf.assoc [symmetric]) lemma multiset_inter_left_commute: "A \# (B \# C) = B \# (A \# C)" by (fact subset_mset.inf.left_commute) lemmas multiset_inter_ac = multiset_inter_commute multiset_inter_assoc multiset_inter_left_commute lemma mset_le_not_refl: "\ M < (M::'a::preorder multiset)" by (fact less_irrefl) lemma mset_le_trans: "K < M \ M < N \ K < (N::'a::preorder multiset)" by (fact less_trans) lemma mset_le_not_sym: "M < N \ \ N < (M::'a::preorder multiset)" by (fact less_not_sym) lemma mset_le_asym: "M < N \ (\ P \ N < (M::'a::preorder multiset)) \ P" by (fact less_asym) declaration \ let fun multiset_postproc _ maybe_name all_values (T as Type (_, [elem_T])) (Const _ $ t') = let val (maybe_opt, ps) = Nitpick_Model.dest_plain_fun t' ||> (~~) ||> map (apsnd (snd o HOLogic.dest_number)) fun elems_for t = (case AList.lookup (=) ps t of SOME n => replicate n t | NONE => [Const (maybe_name, elem_T --> elem_T) $ t]) in (case maps elems_for (all_values elem_T) @ (if maybe_opt then [Const (Nitpick_Model.unrep_mixfix (), elem_T)] else []) of [] => Const (\<^const_name>\zero_class.zero\, T) | ts => foldl1 (fn (s, t) => Const (\<^const_name>\add_mset\, elem_T --> T --> T) $ s $ t) ts) end | multiset_postproc _ _ _ _ t = t in Nitpick_Model.register_term_postprocessor \<^typ>\'a multiset\ multiset_postproc end \ subsection \Naive implementation using lists\ code_datatype mset lemma [code]: "{#} = mset []" by simp lemma [code]: "add_mset x (mset xs) = mset (x # xs)" by simp lemma [code]: "Multiset.is_empty (mset xs) \ List.null xs" by (simp add: Multiset.is_empty_def List.null_def) lemma union_code [code]: "mset xs + mset ys = mset (xs @ ys)" by simp lemma [code]: "image_mset f (mset xs) = mset (map f xs)" by simp lemma [code]: "filter_mset f (mset xs) = mset (filter f xs)" by simp lemma [code]: "mset xs - mset ys = mset (fold remove1 ys xs)" by (rule sym, induct ys arbitrary: xs) (simp_all add: diff_add diff_right_commute diff_diff_add) lemma [code]: "mset xs \# mset ys = mset (snd (fold (\x (ys, zs). if x \ set ys then (remove1 x ys, x # zs) else (ys, zs)) xs (ys, [])))" proof - have "\zs. mset (snd (fold (\x (ys, zs). if x \ set ys then (remove1 x ys, x # zs) else (ys, zs)) xs (ys, zs))) = (mset xs \# mset ys) + mset zs" by (induct xs arbitrary: ys) (auto simp add: inter_add_right1 inter_add_right2 ac_simps) then show ?thesis by simp qed lemma [code]: "mset xs \# mset ys = mset (case_prod append (fold (\x (ys, zs). (remove1 x ys, x # zs)) xs (ys, [])))" proof - have "\zs. mset (case_prod append (fold (\x (ys, zs). (remove1 x ys, x # zs)) xs (ys, zs))) = (mset xs \# mset ys) + mset zs" by (induct xs arbitrary: ys) (simp_all add: multiset_eq_iff) then show ?thesis by simp qed declare in_multiset_in_set [code_unfold] lemma [code]: "count (mset xs) x = fold (\y. if x = y then Suc else id) xs 0" proof - have "\n. fold (\y. if x = y then Suc else id) xs n = count (mset xs) x + n" by (induct xs) simp_all then show ?thesis by simp qed declare set_mset_mset [code] declare sorted_list_of_multiset_mset [code] lemma [code]: \ \not very efficient, but representation-ignorant!\ "mset_set A = mset (sorted_list_of_set A)" apply (cases "finite A") apply simp_all apply (induct A rule: finite_induct) apply simp_all done declare size_mset [code] fun subset_eq_mset_impl :: "'a list \ 'a list \ bool option" where "subset_eq_mset_impl [] ys = Some (ys \ [])" | "subset_eq_mset_impl (Cons x xs) ys = (case List.extract ((=) x) ys of None \ None | Some (ys1,_,ys2) \ subset_eq_mset_impl xs (ys1 @ ys2))" lemma subset_eq_mset_impl: "(subset_eq_mset_impl xs ys = None \ \ mset xs \# mset ys) \ (subset_eq_mset_impl xs ys = Some True \ mset xs \# mset ys) \ (subset_eq_mset_impl xs ys = Some False \ mset xs = mset ys)" proof (induct xs arbitrary: ys) case (Nil ys) show ?case by (auto simp: subset_mset.zero_less_iff_neq_zero) next case (Cons x xs ys) show ?case proof (cases "List.extract ((=) x) ys") case None hence x: "x \ set ys" by (simp add: extract_None_iff) { assume "mset (x # xs) \# mset ys" from set_mset_mono[OF this] x have False by simp } note nle = this moreover { assume "mset (x # xs) \# mset ys" hence "mset (x # xs) \# mset ys" by auto from nle[OF this] have False . } ultimately show ?thesis using None by auto next case (Some res) obtain ys1 y ys2 where res: "res = (ys1,y,ys2)" by (cases res, auto) note Some = Some[unfolded res] from extract_SomeE[OF Some] have "ys = ys1 @ x # ys2" by simp hence id: "mset ys = add_mset x (mset (ys1 @ ys2))" by auto show ?thesis unfolding subset_eq_mset_impl.simps unfolding Some option.simps split unfolding id using Cons[of "ys1 @ ys2"] unfolding subset_mset_def subseteq_mset_def by auto qed qed lemma [code]: "mset xs \# mset ys \ subset_eq_mset_impl xs ys \ None" using subset_eq_mset_impl[of xs ys] by (cases "subset_eq_mset_impl xs ys", auto) lemma [code]: "mset xs \# mset ys \ subset_eq_mset_impl xs ys = Some True" using subset_eq_mset_impl[of xs ys] by (cases "subset_eq_mset_impl xs ys", auto) instantiation multiset :: (equal) equal begin definition [code del]: "HOL.equal A (B :: 'a multiset) \ A = B" lemma [code]: "HOL.equal (mset xs) (mset ys) \ subset_eq_mset_impl xs ys = Some False" unfolding equal_multiset_def using subset_eq_mset_impl[of xs ys] by (cases "subset_eq_mset_impl xs ys", auto) instance by standard (simp add: equal_multiset_def) end declare sum_mset_sum_list [code] lemma [code]: "prod_mset (mset xs) = fold times xs 1" proof - have "\x. fold times xs x = prod_mset (mset xs) * x" by (induct xs) (simp_all add: ac_simps) then show ?thesis by simp qed text \ Exercise for the casual reader: add implementations for \<^term>\(\)\ and \<^term>\(<)\ (multiset order). \ text \Quickcheck generators\ definition (in term_syntax) msetify :: "'a::typerep list \ (unit \ Code_Evaluation.term) \ 'a multiset \ (unit \ Code_Evaluation.term)" where [code_unfold]: "msetify xs = Code_Evaluation.valtermify mset {\} xs" notation fcomp (infixl "\>" 60) notation scomp (infixl "\\" 60) instantiation multiset :: (random) random begin definition "Quickcheck_Random.random i = Quickcheck_Random.random i \\ (\xs. Pair (msetify xs))" instance .. end no_notation fcomp (infixl "\>" 60) no_notation scomp (infixl "\\" 60) instantiation multiset :: (full_exhaustive) full_exhaustive begin definition full_exhaustive_multiset :: "('a multiset \ (unit \ term) \ (bool \ term list) option) \ natural \ (bool \ term list) option" where "full_exhaustive_multiset f i = Quickcheck_Exhaustive.full_exhaustive (\xs. f (msetify xs)) i" instance .. end hide_const (open) msetify subsection \BNF setup\ definition rel_mset where "rel_mset R X Y \ (\xs ys. mset xs = X \ mset ys = Y \ list_all2 R xs ys)" lemma mset_zip_take_Cons_drop_twice: assumes "length xs = length ys" "j \ length xs" shows "mset (zip (take j xs @ x # drop j xs) (take j ys @ y # drop j ys)) = add_mset (x,y) (mset (zip xs ys))" using assms proof (induct xs ys arbitrary: x y j rule: list_induct2) case Nil thus ?case by simp next case (Cons x xs y ys) thus ?case proof (cases "j = 0") case True thus ?thesis by simp next case False then obtain k where k: "j = Suc k" by (cases j) simp hence "k \ length xs" using Cons.prems by auto hence "mset (zip (take k xs @ x # drop k xs) (take k ys @ y # drop k ys)) = add_mset (x,y) (mset (zip xs ys))" by (rule Cons.hyps(2)) thus ?thesis unfolding k by auto qed qed lemma ex_mset_zip_left: assumes "length xs = length ys" "mset xs' = mset xs" shows "\ys'. length ys' = length xs' \ mset (zip xs' ys') = mset (zip xs ys)" using assms proof (induct xs ys arbitrary: xs' rule: list_induct2) case Nil thus ?case by auto next case (Cons x xs y ys xs') obtain j where j_len: "j < length xs'" and nth_j: "xs' ! j = x" by (metis Cons.prems in_set_conv_nth list.set_intros(1) mset_eq_setD) define xsa where "xsa = take j xs' @ drop (Suc j) xs'" have "mset xs' = {#x#} + mset xsa" unfolding xsa_def using j_len nth_j by (metis Cons_nth_drop_Suc union_mset_add_mset_right add_mset_remove_trivial add_diff_cancel_left' append_take_drop_id mset.simps(2) mset_append) hence ms_x: "mset xsa = mset xs" by (simp add: Cons.prems) then obtain ysa where len_a: "length ysa = length xsa" and ms_a: "mset (zip xsa ysa) = mset (zip xs ys)" using Cons.hyps(2) by blast define ys' where "ys' = take j ysa @ y # drop j ysa" have xs': "xs' = take j xsa @ x # drop j xsa" using ms_x j_len nth_j Cons.prems xsa_def by (metis append_eq_append_conv append_take_drop_id diff_Suc_Suc Cons_nth_drop_Suc length_Cons length_drop size_mset) have j_len': "j \ length xsa" using j_len xs' xsa_def by (metis add_Suc_right append_take_drop_id length_Cons length_append less_eq_Suc_le not_less) have "length ys' = length xs'" unfolding ys'_def using Cons.prems len_a ms_x by (metis add_Suc_right append_take_drop_id length_Cons length_append mset_eq_length) moreover have "mset (zip xs' ys') = mset (zip (x # xs) (y # ys))" unfolding xs' ys'_def by (rule trans[OF mset_zip_take_Cons_drop_twice]) (auto simp: len_a ms_a j_len') ultimately show ?case by blast qed lemma list_all2_reorder_left_invariance: assumes rel: "list_all2 R xs ys" and ms_x: "mset xs' = mset xs" shows "\ys'. list_all2 R xs' ys' \ mset ys' = mset ys" proof - have len: "length xs = length ys" using rel list_all2_conv_all_nth by auto obtain ys' where len': "length xs' = length ys'" and ms_xy: "mset (zip xs' ys') = mset (zip xs ys)" using len ms_x by (metis ex_mset_zip_left) have "list_all2 R xs' ys'" using assms(1) len' ms_xy unfolding list_all2_iff by (blast dest: mset_eq_setD) moreover have "mset ys' = mset ys" using len len' ms_xy map_snd_zip mset_map by metis ultimately show ?thesis by blast qed lemma ex_mset: "\xs. mset xs = X" by (induct X) (simp, metis mset.simps(2)) inductive pred_mset :: "('a \ bool) \ 'a multiset \ bool" where "pred_mset P {#}" | "\P a; pred_mset P M\ \ pred_mset P (add_mset a M)" bnf "'a multiset" map: image_mset sets: set_mset bd: natLeq wits: "{#}" rel: rel_mset pred: pred_mset proof - show "image_mset id = id" by (rule image_mset.id) show "image_mset (g \ f) = image_mset g \ image_mset f" for f g unfolding comp_def by (rule ext) (simp add: comp_def image_mset.compositionality) show "(\z. z \ set_mset X \ f z = g z) \ image_mset f X = image_mset g X" for f g X by (induct X) simp_all show "set_mset \ image_mset f = (`) f \ set_mset" for f by auto show "card_order natLeq" by (rule natLeq_card_order) show "BNF_Cardinal_Arithmetic.cinfinite natLeq" by (rule natLeq_cinfinite) show "ordLeq3 (card_of (set_mset X)) natLeq" for X by transfer (auto intro!: ordLess_imp_ordLeq simp: finite_iff_ordLess_natLeq[symmetric] multiset_def) show "rel_mset R OO rel_mset S \ rel_mset (R OO S)" for R S unfolding rel_mset_def[abs_def] OO_def apply clarify subgoal for X Z Y xs ys' ys zs apply (drule list_all2_reorder_left_invariance [where xs = ys' and ys = zs and xs' = ys]) apply (auto intro: list_all2_trans) done done show "rel_mset R = (\x y. \z. set_mset z \ {(x, y). R x y} \ image_mset fst z = x \ image_mset snd z = y)" for R unfolding rel_mset_def[abs_def] apply (rule ext)+ apply safe apply (rule_tac x = "mset (zip xs ys)" in exI; auto simp: in_set_zip list_all2_iff simp flip: mset_map) apply (rename_tac XY) apply (cut_tac X = XY in ex_mset) apply (erule exE) apply (rename_tac xys) apply (rule_tac x = "map fst xys" in exI) apply (auto simp: mset_map) apply (rule_tac x = "map snd xys" in exI) apply (auto simp: mset_map list_all2I subset_eq zip_map_fst_snd) done show "z \ set_mset {#} \ False" for z by auto show "pred_mset P = (\x. Ball (set_mset x) P)" for P proof (intro ext iffI) fix x assume "pred_mset P x" then show "Ball (set_mset x) P" by (induct pred: pred_mset; simp) next fix x assume "Ball (set_mset x) P" then show "pred_mset P x" by (induct x; auto intro: pred_mset.intros) qed qed inductive rel_mset' where Zero[intro]: "rel_mset' R {#} {#}" | Plus[intro]: "\R a b; rel_mset' R M N\ \ rel_mset' R (add_mset a M) (add_mset b N)" lemma rel_mset_Zero: "rel_mset R {#} {#}" unfolding rel_mset_def Grp_def by auto declare multiset.count[simp] declare Abs_multiset_inverse[simp] declare multiset.count_inverse[simp] declare union_preserves_multiset[simp] lemma rel_mset_Plus: assumes ab: "R a b" and MN: "rel_mset R M N" shows "rel_mset R (add_mset a M) (add_mset b N)" proof - have "\ya. add_mset a (image_mset fst y) = image_mset fst ya \ add_mset b (image_mset snd y) = image_mset snd ya \ set_mset ya \ {(x, y). R x y}" if "R a b" and "set_mset y \ {(x, y). R x y}" for y using that by (intro exI[of _ "add_mset (a,b) y"]) auto thus ?thesis using assms unfolding multiset.rel_compp_Grp Grp_def by blast qed lemma rel_mset'_imp_rel_mset: "rel_mset' R M N \ rel_mset R M N" by (induct rule: rel_mset'.induct) (auto simp: rel_mset_Zero rel_mset_Plus) lemma rel_mset_size: "rel_mset R M N \ size M = size N" unfolding multiset.rel_compp_Grp Grp_def by auto lemma multiset_induct2[case_names empty addL addR]: assumes empty: "P {#} {#}" and addL: "\a M N. P M N \ P (add_mset a M) N" and addR: "\a M N. P M N \ P M (add_mset a N)" shows "P M N" apply(induct N rule: multiset_induct) apply(induct M rule: multiset_induct, rule empty, erule addL) apply(induct M rule: multiset_induct, erule addR, erule addR) done lemma multiset_induct2_size[consumes 1, case_names empty add]: assumes c: "size M = size N" and empty: "P {#} {#}" and add: "\a b M N a b. P M N \ P (add_mset a M) (add_mset b N)" shows "P M N" using c proof (induct M arbitrary: N rule: measure_induct_rule[of size]) case (less M) show ?case proof(cases "M = {#}") case True hence "N = {#}" using less.prems by auto thus ?thesis using True empty by auto next case False then obtain M1 a where M: "M = add_mset a M1" by (metis multi_nonempty_split) have "N \ {#}" using False less.prems by auto then obtain N1 b where N: "N = add_mset b N1" by (metis multi_nonempty_split) have "size M1 = size N1" using less.prems unfolding M N by auto thus ?thesis using M N less.hyps add by auto qed qed lemma msed_map_invL: assumes "image_mset f (add_mset a M) = N" shows "\N1. N = add_mset (f a) N1 \ image_mset f M = N1" proof - have "f a \# N" using assms multiset.set_map[of f "add_mset a M"] by auto then obtain N1 where N: "N = add_mset (f a) N1" using multi_member_split by metis have "image_mset f M = N1" using assms unfolding N by simp thus ?thesis using N by blast qed lemma msed_map_invR: assumes "image_mset f M = add_mset b N" shows "\M1 a. M = add_mset a M1 \ f a = b \ image_mset f M1 = N" proof - obtain a where a: "a \# M" and fa: "f a = b" using multiset.set_map[of f M] unfolding assms by (metis image_iff union_single_eq_member) then obtain M1 where M: "M = add_mset a M1" using multi_member_split by metis have "image_mset f M1 = N" using assms unfolding M fa[symmetric] by simp thus ?thesis using M fa by blast qed lemma msed_rel_invL: assumes "rel_mset R (add_mset a M) N" shows "\N1 b. N = add_mset b N1 \ R a b \ rel_mset R M N1" proof - obtain K where KM: "image_mset fst K = add_mset a M" and KN: "image_mset snd K = N" and sK: "set_mset K \ {(a, b). R a b}" using assms unfolding multiset.rel_compp_Grp Grp_def by auto obtain K1 ab where K: "K = add_mset ab K1" and a: "fst ab = a" and K1M: "image_mset fst K1 = M" using msed_map_invR[OF KM] by auto obtain N1 where N: "N = add_mset (snd ab) N1" and K1N1: "image_mset snd K1 = N1" using msed_map_invL[OF KN[unfolded K]] by auto have Rab: "R a (snd ab)" using sK a unfolding K by auto have "rel_mset R M N1" using sK K1M K1N1 unfolding K multiset.rel_compp_Grp Grp_def by auto thus ?thesis using N Rab by auto qed lemma msed_rel_invR: assumes "rel_mset R M (add_mset b N)" shows "\M1 a. M = add_mset a M1 \ R a b \ rel_mset R M1 N" proof - obtain K where KN: "image_mset snd K = add_mset b N" and KM: "image_mset fst K = M" and sK: "set_mset K \ {(a, b). R a b}" using assms unfolding multiset.rel_compp_Grp Grp_def by auto obtain K1 ab where K: "K = add_mset ab K1" and b: "snd ab = b" and K1N: "image_mset snd K1 = N" using msed_map_invR[OF KN] by auto obtain M1 where M: "M = add_mset (fst ab) M1" and K1M1: "image_mset fst K1 = M1" using msed_map_invL[OF KM[unfolded K]] by auto have Rab: "R (fst ab) b" using sK b unfolding K by auto have "rel_mset R M1 N" using sK K1N K1M1 unfolding K multiset.rel_compp_Grp Grp_def by auto thus ?thesis using M Rab by auto qed lemma rel_mset_imp_rel_mset': assumes "rel_mset R M N" shows "rel_mset' R M N" using assms proof(induct M arbitrary: N rule: measure_induct_rule[of size]) case (less M) have c: "size M = size N" using rel_mset_size[OF less.prems] . show ?case proof(cases "M = {#}") case True hence "N = {#}" using c by simp thus ?thesis using True rel_mset'.Zero by auto next case False then obtain M1 a where M: "M = add_mset a M1" by (metis multi_nonempty_split) obtain N1 b where N: "N = add_mset b N1" and R: "R a b" and ms: "rel_mset R M1 N1" using msed_rel_invL[OF less.prems[unfolded M]] by auto have "rel_mset' R M1 N1" using less.hyps[of M1 N1] ms unfolding M by simp thus ?thesis using rel_mset'.Plus[of R a b, OF R] unfolding M N by simp qed qed lemma rel_mset_rel_mset': "rel_mset R M N = rel_mset' R M N" using rel_mset_imp_rel_mset' rel_mset'_imp_rel_mset by auto text \The main end product for \<^const>\rel_mset\: inductive characterization:\ lemmas rel_mset_induct[case_names empty add, induct pred: rel_mset] = rel_mset'.induct[unfolded rel_mset_rel_mset'[symmetric]] subsection \Size setup\ lemma size_multiset_o_map: "size_multiset g \ image_mset f = size_multiset (g \ f)" apply (rule ext) subgoal for x by (induct x) auto done setup \ BNF_LFP_Size.register_size_global \<^type_name>\multiset\ \<^const_name>\size_multiset\ @{thm size_multiset_overloaded_def} @{thms size_multiset_empty size_multiset_single size_multiset_union size_empty size_single size_union} @{thms size_multiset_o_map} \ subsection \Lemmas about Size\ lemma size_mset_SucE: "size A = Suc n \ (\a B. A = {#a#} + B \ size B = n \ P) \ P" by (cases A) (auto simp add: ac_simps) lemma size_Suc_Diff1: "x \# M \ Suc (size (M - {#x#})) = size M" using arg_cong[OF insert_DiffM, of _ _ size] by simp lemma size_Diff_singleton: "x \# M \ size (M - {#x#}) = size M - 1" by (simp flip: size_Suc_Diff1) lemma size_Diff_singleton_if: "size (A - {#x#}) = (if x \# A then size A - 1 else size A)" by (simp add: diff_single_trivial size_Diff_singleton) lemma size_Un_Int: "size A + size B = size (A \# B) + size (A \# B)" by (metis inter_subset_eq_union size_union subset_mset.diff_add union_diff_inter_eq_sup) lemma size_Un_disjoint: "A \# B = {#} \ size (A \# B) = size A + size B" using size_Un_Int[of A B] by simp lemma size_Diff_subset_Int: "size (M - M') = size M - size (M \# M')" by (metis diff_intersect_left_idem size_Diff_submset subset_mset.inf_le1) lemma diff_size_le_size_Diff: "size (M :: _ multiset) - size M' \ size (M - M')" by (simp add: diff_le_mono2 size_Diff_subset_Int size_mset_mono) lemma size_Diff1_less: "x\# M \ size (M - {#x#}) < size M" by (rule Suc_less_SucD) (simp add: size_Suc_Diff1) lemma size_Diff2_less: "x\# M \ y\# M \ size (M - {#x#} - {#y#}) < size M" by (metis less_imp_diff_less size_Diff1_less size_Diff_subset_Int) lemma size_Diff1_le: "size (M - {#x#}) \ size M" by (cases "x \# M") (simp_all add: size_Diff1_less less_imp_le diff_single_trivial) lemma size_psubset: "M \# M' \ size M < size M' \ M \# M'" using less_irrefl subset_mset_def by blast hide_const (open) wcount end diff --git a/src/HOL/Number_Theory/Prime_Powers.thy b/src/HOL/Number_Theory/Prime_Powers.thy --- a/src/HOL/Number_Theory/Prime_Powers.thy +++ b/src/HOL/Number_Theory/Prime_Powers.thy @@ -1,513 +1,518 @@ (* File: HOL/Number_Theory/Prime_Powers.thy Author: Manuel Eberl Prime powers and the Mangoldt function *) section \Prime powers\ theory Prime_Powers imports Complex_Main "HOL-Computational_Algebra.Primes" "HOL-Library.FuncSet" begin definition aprimedivisor :: "'a :: normalization_semidom \ 'a" where "aprimedivisor q = (SOME p. prime p \ p dvd q)" definition primepow :: "'a :: normalization_semidom \ bool" where "primepow n \ (\p k. prime p \ k > 0 \ n = p ^ k)" definition primepow_factors :: "'a :: normalization_semidom \ 'a set" where "primepow_factors n = {x. primepow x \ x dvd n}" lemma primepow_gt_Suc_0: "primepow n \ n > Suc 0" using one_less_power[of "p::nat" for p] by (auto simp: primepow_def prime_nat_iff) lemma assumes "prime p" "p dvd n" shows prime_aprimedivisor: "prime (aprimedivisor n)" and aprimedivisor_dvd: "aprimedivisor n dvd n" proof - from assms have "\p. prime p \ p dvd n" by auto from someI_ex[OF this] show "prime (aprimedivisor n)" "aprimedivisor n dvd n" unfolding aprimedivisor_def by (simp_all add: conj_commute) qed lemma assumes "n \ 0" "\is_unit (n :: 'a :: factorial_semiring)" shows prime_aprimedivisor': "prime (aprimedivisor n)" and aprimedivisor_dvd': "aprimedivisor n dvd n" proof - from someI_ex[OF prime_divisor_exists[OF assms]] show "prime (aprimedivisor n)" "aprimedivisor n dvd n" unfolding aprimedivisor_def by (simp_all add: conj_commute) qed lemma aprimedivisor_of_prime [simp]: assumes "prime p" shows "aprimedivisor p = p" proof - from assms have "\q. prime q \ q dvd p" by auto from someI_ex[OF this, folded aprimedivisor_def] assms show ?thesis by (auto intro: primes_dvd_imp_eq) qed lemma aprimedivisor_pos_nat: "(n::nat) > 1 \ aprimedivisor n > 0" using aprimedivisor_dvd'[of n] by (auto elim: dvdE intro!: Nat.gr0I) lemma aprimedivisor_primepow_power: assumes "primepow n" "k > 0" shows "aprimedivisor (n ^ k) = aprimedivisor n" proof - from assms obtain p l where l: "prime p" "l > 0" "n = p ^ l" by (auto simp: primepow_def) from l assms have *: "prime (aprimedivisor (n ^ k))" "aprimedivisor (n ^ k) dvd n ^ k" by (intro prime_aprimedivisor[of p] aprimedivisor_dvd[of p] dvd_power; simp add: power_mult [symmetric])+ from * l have "aprimedivisor (n ^ k) dvd p ^ (l * k)" by (simp add: power_mult) with assms * l have "aprimedivisor (n ^ k) dvd p" by (subst (asm) prime_dvd_power_iff) simp_all with l assms have "aprimedivisor (n ^ k) = p" by (intro primes_dvd_imp_eq prime_aprimedivisor l) (auto simp: power_mult [symmetric]) moreover from l have "aprimedivisor n dvd p ^ l" by (auto intro: aprimedivisor_dvd simp: prime_gt_0_nat) with assms l have "aprimedivisor n dvd p" by (subst (asm) prime_dvd_power_iff) (auto intro!: prime_aprimedivisor simp: prime_gt_0_nat) with l assms have "aprimedivisor n = p" by (intro primes_dvd_imp_eq prime_aprimedivisor l) auto ultimately show ?thesis by simp qed lemma aprimedivisor_prime_power: assumes "prime p" "k > 0" shows "aprimedivisor (p ^ k) = p" proof - from assms have *: "prime (aprimedivisor (p ^ k))" "aprimedivisor (p ^ k) dvd p ^ k" by (intro prime_aprimedivisor[of p] aprimedivisor_dvd[of p]; simp add: prime_nat_iff)+ from assms * have "aprimedivisor (p ^ k) dvd p" by (subst (asm) prime_dvd_power_iff) simp_all with assms * show "aprimedivisor (p ^ k) = p" by (intro primes_dvd_imp_eq) qed lemma prime_factorization_primepow: assumes "primepow n" shows "prime_factorization n = replicate_mset (multiplicity (aprimedivisor n) n) (aprimedivisor n)" using assms by (auto simp: primepow_def aprimedivisor_prime_power prime_factorization_prime_power) lemma primepow_decompose: + fixes n :: "'a :: factorial_semiring_multiplicative" assumes "primepow n" shows "aprimedivisor n ^ multiplicity (aprimedivisor n) n = n" proof - from assms have "n \ 0" by (intro notI) (auto simp: primepow_def) hence "n = unit_factor n * prod_mset (prime_factorization n)" by (subst prod_mset_prime_factorization) simp_all also from assms have "unit_factor n = 1" by (auto simp: primepow_def unit_factor_power) also have "prime_factorization n = replicate_mset (multiplicity (aprimedivisor n) n) (aprimedivisor n)" by (intro prime_factorization_primepow assms) also have "prod_mset \ = aprimedivisor n ^ multiplicity (aprimedivisor n) n" by simp finally show ?thesis by simp qed lemma prime_power_not_one: assumes "prime p" "k > 0" shows "p ^ k \ 1" proof assume "p ^ k = 1" hence "is_unit (p ^ k)" by simp thus False using assms by (simp add: is_unit_power_iff) qed lemma zero_not_primepow [simp]: "\primepow 0" by (auto simp: primepow_def) lemma one_not_primepow [simp]: "\primepow 1" by (auto simp: primepow_def prime_power_not_one) lemma primepow_not_unit [simp]: "primepow p \ \is_unit p" by (auto simp: primepow_def is_unit_power_iff) lemma not_primepow_Suc_0_nat [simp]: "\primepow (Suc 0)" using primepow_gt_Suc_0[of "Suc 0"] by auto lemma primepow_gt_0_nat: "primepow n \ n > (0::nat)" using primepow_gt_Suc_0[of n] by simp -lemma unit_factor_primepow: "primepow p \ unit_factor p = 1" +lemma unit_factor_primepow: + fixes p :: "'a :: factorial_semiring_multiplicative" + shows "primepow p \ unit_factor p = 1" by (auto simp: primepow_def unit_factor_power) lemma aprimedivisor_primepow: - assumes "prime p" "p dvd n" "primepow (n :: 'a :: factorial_semiring)" + assumes "prime p" "p dvd n" "primepow (n :: 'a :: factorial_semiring_multiplicative)" shows "aprimedivisor (p * n) = p" "aprimedivisor n = p" proof - from assms have [simp]: "n \ 0" by auto define q where "q = aprimedivisor n" with assms have q: "prime q" by (auto simp: q_def intro!: prime_aprimedivisor) from \primepow n\ have n: "n = q ^ multiplicity q n" by (simp add: primepow_decompose q_def) have nz: "multiplicity q n \ 0" proof assume "multiplicity q n = 0" with n have n': "n = unit_factor n" by simp have "is_unit n" by (subst n', rule unit_factor_is_unit) (insert assms, auto) with assms show False by auto qed with \prime p\ \p dvd n\ q have "p dvd q" by (subst (asm) n) (auto intro: prime_dvd_power) with \prime p\ q have "p = q" by (intro primes_dvd_imp_eq) thus "aprimedivisor n = p" by (simp add: q_def) define r where "r = aprimedivisor (p * n)" with assms have r: "r dvd (p * n)" "prime r" unfolding r_def by (intro aprimedivisor_dvd[of p] prime_aprimedivisor[of p]; simp)+ hence "r dvd q ^ Suc (multiplicity q n)" by (subst (asm) n) (auto simp: \p = q\ dest: dvd_unit_imp_unit) with r have "r dvd q" by (auto intro: prime_dvd_power_nat simp: prime_dvd_mult_iff dest: prime_dvd_power) with r q have "r = q" by (intro primes_dvd_imp_eq) thus "aprimedivisor (p * n) = p" by (simp add: r_def \p = q\) qed lemma power_eq_prime_powerD: fixes p :: "'a :: factorial_semiring" assumes "prime p" "n > 0" "x ^ n = p ^ k" shows "\i. normalize x = normalize (p ^ i)" proof - have "normalize x = normalize (p ^ multiplicity p x)" proof (rule multiplicity_eq_imp_eq) fix q :: 'a assume "prime q" from assms have "multiplicity q (x ^ n) = multiplicity q (p ^ k)" by simp with \prime q\ and assms have "n * multiplicity q x = k * multiplicity q p" by (subst (asm) (1 2) prime_elem_multiplicity_power_distrib) (auto simp: power_0_left) with assms and \prime q\ show "multiplicity q x = multiplicity q (p ^ multiplicity p x)" by (cases "p = q") (auto simp: multiplicity_distinct_prime_power prime_multiplicity_other) qed (insert assms, auto simp: power_0_left) thus ?thesis by auto qed lemma primepow_power_iff: + fixes p :: "'a :: factorial_semiring_multiplicative" assumes "unit_factor p = 1" - shows "primepow (p ^ n) \ primepow (p :: 'a :: factorial_semiring) \ n > 0" + shows "primepow (p ^ n) \ primepow p \ n > 0" proof safe assume "primepow (p ^ n)" hence n: "n \ 0" by (auto intro!: Nat.gr0I) thus "n > 0" by simp from assms have [simp]: "normalize p = p" using normalize_mult_unit_factor[of p] by (simp only: mult.right_neutral) from \primepow (p ^ n)\ obtain q k where *: "k > 0" "prime q" "p ^ n = q ^ k" by (auto simp: primepow_def) with power_eq_prime_powerD[of q n p k] n obtain i where eq: "normalize p = normalize (q ^ i)" by auto with primepow_not_unit[OF \primepow (p ^ n)\] have "i \ 0" by (intro notI) (simp add: normalize_1_iff is_unit_power_iff del: primepow_not_unit) with \normalize p = normalize (q ^ i)\ \prime q\ show "primepow p" by (auto simp: normalize_power primepow_def intro!: exI[of _ q] exI[of _ i]) next assume "primepow p" "n > 0" then obtain q k where *: "k > 0" "prime q" "p = q ^ k" by (auto simp: primepow_def) with \n > 0\ show "primepow (p ^ n)" by (auto simp: primepow_def power_mult intro!: exI[of _ q] exI[of _ "k * n"]) qed lemma primepow_power_iff_nat: "p > 0 \ primepow (p ^ n) \ primepow (p :: nat) \ n > 0" by (rule primepow_power_iff) (simp_all add: unit_factor_nat_def) lemma primepow_prime [simp]: "prime n \ primepow n" by (auto simp: primepow_def intro!: exI[of _ n] exI[of _ "1::nat"]) lemma primepow_prime_power [simp]: - "prime (p :: 'a :: factorial_semiring) \ primepow (p ^ n) \ n > 0" + "prime (p :: 'a :: factorial_semiring_multiplicative) \ primepow (p ^ n) \ n > 0" by (subst primepow_power_iff) auto lemma aprimedivisor_vimage: - assumes "prime (p :: 'a :: factorial_semiring)" + assumes "prime (p :: 'a :: factorial_semiring_multiplicative)" shows "aprimedivisor -` {p} \ primepow_factors n = {p ^ k |k. k > 0 \ p ^ k dvd n}" proof safe fix q assume q: "q \ primepow_factors n" hence q': "q \ 0" "q \ 1" by (auto simp: primepow_def primepow_factors_def prime_power_not_one) let ?n = "multiplicity (aprimedivisor q) q" from q q' have "q = aprimedivisor q ^ ?n \ ?n > 0 \ aprimedivisor q ^ ?n dvd n" by (auto simp: primepow_decompose primepow_factors_def prime_multiplicity_gt_zero_iff prime_aprimedivisor' prime_imp_prime_elem aprimedivisor_dvd') thus "\k. q = aprimedivisor q ^ k \ k > 0 \ aprimedivisor q ^ k dvd n" .. next fix k :: nat assume k: "p ^ k dvd n" "k > 0" with assms show "p ^ k \ aprimedivisor -` {p}" by (auto simp: aprimedivisor_prime_power) with assms k show "p ^ k \ primepow_factors n" by (auto simp: primepow_factors_def primepow_def aprimedivisor_prime_power intro: Suc_leI) qed lemma aprimedivisor_nat: assumes "n \ (Suc 0::nat)" shows "prime (aprimedivisor n)" "aprimedivisor n dvd n" proof - from assms have "\p. prime p \ p dvd n" by (intro prime_factor_nat) auto from someI_ex[OF this, folded aprimedivisor_def] show "prime (aprimedivisor n)" "aprimedivisor n dvd n" by blast+ qed lemma aprimedivisor_gt_Suc_0: assumes "n \ Suc 0" shows "aprimedivisor n > Suc 0" proof - from assms have "prime (aprimedivisor n)" by (rule aprimedivisor_nat) thus "aprimedivisor n > Suc 0" by (simp add: prime_nat_iff) qed lemma aprimedivisor_le_nat: assumes "n > Suc 0" shows "aprimedivisor n \ n" proof - from assms have "aprimedivisor n dvd n" by (intro aprimedivisor_nat) simp_all with assms show "aprimedivisor n \ n" by (intro dvd_imp_le) simp_all qed lemma bij_betw_primepows: - "bij_betw (\(p,k). p ^ Suc k :: 'a :: factorial_semiring) + "bij_betw (\(p,k). p ^ Suc k :: 'a :: factorial_semiring_multiplicative) (Collect prime \ UNIV) (Collect primepow)" proof (rule bij_betwI [where ?g = "(\n. (aprimedivisor n, multiplicity (aprimedivisor n) n - 1))"], goal_cases) case 1 show "(\(p, k). p ^ Suc k :: 'a) \ Collect prime \ UNIV \ Collect primepow" by (auto intro!: primepow_prime_power simp del: power_Suc ) next case 2 show ?case by (auto simp: primepow_def prime_aprimedivisor) next case (3 n) thus ?case by (auto simp: aprimedivisor_prime_power simp del: power_Suc) next case (4 n) hence *: "0 < multiplicity (aprimedivisor n) n" by (subst prime_multiplicity_gt_zero_iff) (auto intro!: prime_imp_prime_elem aprimedivisor_dvd simp: primepow_def prime_aprimedivisor) have "aprimedivisor n * aprimedivisor n ^ (multiplicity (aprimedivisor n) n - Suc 0) = aprimedivisor n ^ Suc (multiplicity (aprimedivisor n) n - Suc 0)" by simp also from * have "Suc (multiplicity (aprimedivisor n) n - Suc 0) = multiplicity (aprimedivisor n) n" by (subst Suc_diff_Suc) (auto simp: prime_multiplicity_gt_zero_iff) also have "aprimedivisor n ^ \ = n" using 4 by (subst primepow_decompose) auto finally show ?case by auto qed (* TODO Generalise *) lemma primepow_multD: assumes "primepow (a * b :: nat)" shows "a = 1 \ primepow a" "b = 1 \ primepow b" proof - from assms obtain p k where k: "k > 0" "a * b = p ^ k" "prime p" unfolding primepow_def by auto then obtain i j where "a = p ^ i" "b = p ^ j" using prime_power_mult_nat[of p a b] by blast with \prime p\ show "a = 1 \ primepow a" "b = 1 \ primepow b" by auto qed lemma primepow_mult_aprimedivisorI: - assumes "primepow (n :: 'a :: factorial_semiring)" + assumes "primepow (n :: 'a :: factorial_semiring_multiplicative)" shows "primepow (aprimedivisor n * n)" by (subst (2) primepow_decompose[OF assms, symmetric], subst power_Suc [symmetric], subst primepow_prime_power) (insert assms, auto intro!: prime_aprimedivisor' dest: primepow_gt_Suc_0) lemma primepow_factors_altdef: - fixes x :: "'a :: factorial_semiring" + fixes x :: "'a :: factorial_semiring_multiplicative" assumes "x \ 0" shows "primepow_factors x = {p ^ k |p k. p \ prime_factors x \ k \ {0<.. multiplicity p x}}" proof (intro equalityI subsetI) fix q assume "q \ primepow_factors x" then obtain p k where pk: "prime p" "k > 0" "q = p ^ k" "q dvd x" unfolding primepow_factors_def primepow_def by blast moreover have "k \ multiplicity p x" using pk assms by (intro multiplicity_geI) auto ultimately show "q \ {p ^ k |p k. p \ prime_factors x \ k \ {0<.. multiplicity p x}}" by (auto simp: prime_factors_multiplicity intro!: exI[of _ p] exI[of _ k]) qed (auto simp: primepow_factors_def prime_factors_multiplicity multiplicity_dvd') lemma finite_primepow_factors: - assumes "x \ (0 :: 'a :: factorial_semiring)" + assumes "x \ (0 :: 'a :: factorial_semiring_multiplicative)" shows "finite (primepow_factors x)" proof - have "finite (SIGMA p:prime_factors x. {0<..multiplicity p x})" by (intro finite_SigmaI) simp_all hence "finite ((\(p,k). p ^ k) ` \)" (is "finite ?A") by (rule finite_imageI) also have "?A = primepow_factors x" using assms by (subst primepow_factors_altdef) fast+ finally show ?thesis . qed lemma aprimedivisor_primepow_factors_conv_prime_factorization: - assumes [simp]: "n \ (0 :: 'a :: factorial_semiring)" + assumes [simp]: "n \ (0 :: 'a :: factorial_semiring_multiplicative)" shows "image_mset aprimedivisor (mset_set (primepow_factors n)) = prime_factorization n" (is "?lhs = ?rhs") proof (intro multiset_eqI) fix p :: 'a show "count ?lhs p = count ?rhs p" proof (cases "prime p") case False have "p \# image_mset aprimedivisor (mset_set (primepow_factors n))" proof assume "p \# image_mset aprimedivisor (mset_set (primepow_factors n))" then obtain q where "p = aprimedivisor q" "q \ primepow_factors n" by (auto simp: finite_primepow_factors) with False prime_aprimedivisor'[of q] have "q = 0 \ is_unit q" by auto with \q \ primepow_factors n\ show False by (auto simp: primepow_factors_def primepow_def) qed hence "count ?lhs p = 0" by (simp only: Multiset.not_in_iff) with False show ?thesis by (simp add: count_prime_factorization) next case True hence p: "p \ 0" "\is_unit p" by auto have "count ?lhs p = card (aprimedivisor -` {p} \ primepow_factors n)" by (simp add: count_image_mset finite_primepow_factors) also have "aprimedivisor -` {p} \ primepow_factors n = {p^k |k. k > 0 \ p ^ k dvd n}" using True by (rule aprimedivisor_vimage) also from True have "\ = (\k. p ^ k) ` {0<..multiplicity p n}" by (subst power_dvd_iff_le_multiplicity) auto also from p True have "card \ = multiplicity p n" by (subst card_image) (auto intro!: inj_onI dest: prime_power_inj) also from True have "\ = count (prime_factorization n) p" by (simp add: count_prime_factorization) finally show ?thesis . qed qed lemma prime_elem_aprimedivisor_nat: "d > Suc 0 \ prime_elem (aprimedivisor d)" using prime_aprimedivisor'[of d] by simp lemma aprimedivisor_gt_0_nat [simp]: "d > Suc 0 \ aprimedivisor d > 0" using prime_aprimedivisor'[of d] by (simp add: prime_gt_0_nat) lemma aprimedivisor_gt_Suc_0_nat [simp]: "d > Suc 0 \ aprimedivisor d > Suc 0" using prime_aprimedivisor'[of d] by (simp add: prime_gt_Suc_0_nat) lemma aprimedivisor_not_Suc_0_nat [simp]: "d > Suc 0 \ aprimedivisor d \ Suc 0" using aprimedivisor_gt_Suc_0[of d] by (intro notI) auto lemma multiplicity_aprimedivisor_gt_0_nat [simp]: "d > Suc 0 \ multiplicity (aprimedivisor d) d > 0" by (subst multiplicity_gt_zero_iff) (auto intro: aprimedivisor_dvd') lemma primepowI: "prime p \ k > 0 \ p ^ k = n \ primepow n \ aprimedivisor n = p" unfolding primepow_def by (auto simp: aprimedivisor_prime_power) lemma not_primepowI: assumes "prime p" "prime q" "p \ q" "p dvd n" "q dvd n" shows "\primepow n" using assms by (auto simp: primepow_def dest!: prime_dvd_power[rotated] dest: primes_dvd_imp_eq) lemma sum_prime_factorization_conv_sum_primepow_factors: + fixes n :: "'a :: factorial_semiring_multiplicative" assumes "n \ 0" shows "(\q\primepow_factors n. f (aprimedivisor q)) = (\p\#prime_factorization n. f p)" proof - from assms have "prime_factorization n = image_mset aprimedivisor (mset_set (primepow_factors n))" by (rule aprimedivisor_primepow_factors_conv_prime_factorization [symmetric]) also have "(\p\#\. f p) = (\q\primepow_factors n. f (aprimedivisor q))" by (simp add: image_mset.compositionality sum_unfold_sum_mset o_def) finally show ?thesis .. qed lemma multiplicity_aprimedivisor_Suc_0_iff: - assumes "primepow (n :: 'a :: factorial_semiring)" + assumes "primepow (n :: 'a :: factorial_semiring_multiplicative)" shows "multiplicity (aprimedivisor n) n = Suc 0 \ prime n" by (subst (3) primepow_decompose [OF assms, symmetric]) (insert assms, auto simp add: prime_power_iff intro!: prime_aprimedivisor') definition mangoldt :: "nat \ 'a :: real_algebra_1" where "mangoldt n = (if primepow n then of_real (ln (real (aprimedivisor n))) else 0)" lemma mangoldt_0 [simp]: "mangoldt 0 = 0" by (simp add: mangoldt_def) lemma mangoldt_Suc_0 [simp]: "mangoldt (Suc 0) = 0" by (simp add: mangoldt_def) lemma of_real_mangoldt [simp]: "of_real (mangoldt n) = mangoldt n" by (simp add: mangoldt_def) lemma mangoldt_sum: assumes "n \ 0" shows "(\d | d dvd n. mangoldt d :: 'a :: real_algebra_1) = of_real (ln (real n))" proof - have "(\d | d dvd n. mangoldt d :: 'a) = of_real (\d | d dvd n. mangoldt d)" by simp also have "(\d | d dvd n. mangoldt d) = (\d \ primepow_factors n. ln (real (aprimedivisor d)))" using assms by (intro sum.mono_neutral_cong_right) (auto simp: primepow_factors_def mangoldt_def) also have "\ = ln (real (\d \ primepow_factors n. aprimedivisor d))" using assms finite_primepow_factors[of n] by (subst ln_prod [symmetric]) (auto simp: primepow_factors_def intro!: aprimedivisor_pos_nat intro: Nat.gr0I primepow_gt_Suc_0) also have "primepow_factors n = (\(p,k). p ^ k) ` (SIGMA p:prime_factors n. {0<..multiplicity p n})" (is "_ = _ ` ?A") by (subst primepow_factors_altdef[OF assms]) fast+ also have "prod aprimedivisor \ = (\(p,k)\?A. aprimedivisor (p ^ k))" by (subst prod.reindex) (auto simp: inj_on_def prime_power_inj'' prime_factors_multiplicity prod.Sigma [symmetric] case_prod_unfold) also have "\ = (\(p,k)\?A. p)" by (intro prod.cong refl) (auto simp: aprimedivisor_prime_power prime_factors_multiplicity) also have "\ = (\x\prime_factors n. \k\{0<..multiplicity x n}. x)" by (rule prod.Sigma [symmetric]) auto also have "\ = (\x\prime_factors n. x ^ multiplicity x n)" by (intro prod.cong refl) (simp add: prod_constant) also have "\ = n" using assms by (intro prime_factorization_nat [symmetric]) simp finally show ?thesis . qed lemma mangoldt_primepow: "prime p \ mangoldt (p ^ k) = (if k > 0 then of_real (ln (real p)) else 0)" by (simp add: mangoldt_def aprimedivisor_prime_power) lemma mangoldt_primepow' [simp]: "prime p \ k > 0 \ mangoldt (p ^ k) = of_real (ln (real p))" by (subst mangoldt_primepow) auto lemma mangoldt_prime [simp]: "prime p \ mangoldt p = of_real (ln (real p))" using mangoldt_primepow[of p 1] by simp lemma mangoldt_nonneg: "0 \ (mangoldt d :: real)" using aprimedivisor_gt_Suc_0_nat[of d] by (auto simp: mangoldt_def of_nat_le_iff[of 1 x for x, unfolded of_nat_1] Suc_le_eq intro!: ln_ge_zero dest: primepow_gt_Suc_0) lemma norm_mangoldt [simp]: "norm (mangoldt n :: 'a :: real_normed_algebra_1) = mangoldt n" proof (cases "primepow n") case True hence "prime (aprimedivisor n)" by (intro prime_aprimedivisor') (auto simp: primepow_def prime_gt_0_nat) hence "aprimedivisor n > 1" by (simp add: prime_gt_Suc_0_nat) with True show ?thesis by (auto simp: mangoldt_def abs_if) qed (auto simp: mangoldt_def) lemma Re_mangoldt [simp]: "Re (mangoldt n) = mangoldt n" and Im_mangoldt [simp]: "Im (mangoldt n) = 0" by (simp_all add: mangoldt_def) lemma abs_mangoldt [simp]: "abs (mangoldt n :: real) = mangoldt n" using norm_mangoldt[of n, where ?'a = real, unfolded real_norm_def] . lemma mangoldt_le: assumes "n > 0" shows "mangoldt n \ ln n" proof (cases "primepow n") case True from True have "prime (aprimedivisor n)" by (intro prime_aprimedivisor') (auto simp: primepow_def prime_gt_0_nat) hence gt_1: "aprimedivisor n > 1" by (simp add: prime_gt_Suc_0_nat) from True have "mangoldt n = ln (aprimedivisor n)" by (simp add: mangoldt_def) also have "\ \ ln n" using True gt_1 by (subst ln_le_cancel_iff) (auto intro!: Nat.gr0I dvd_imp_le aprimedivisor_dvd') finally show ?thesis . qed (insert assms, auto simp: mangoldt_def) end diff --git a/src/HOL/Power.thy b/src/HOL/Power.thy --- a/src/HOL/Power.thy +++ b/src/HOL/Power.thy @@ -1,1015 +1,1015 @@ (* Title: HOL/Power.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1997 University of Cambridge *) section \Exponentiation\ theory Power imports Num begin subsection \Powers for Arbitrary Monoids\ class power = one + times begin primrec power :: "'a \ nat \ 'a" (infixr "^" 80) where power_0: "a ^ 0 = 1" | power_Suc: "a ^ Suc n = a * a ^ n" notation (latex output) power ("(_\<^bsup>_\<^esup>)" [1000] 1000) text \Special syntax for squares.\ abbreviation power2 :: "'a \ 'a" ("(_\<^sup>2)" [1000] 999) where "x\<^sup>2 \ x ^ 2" end context includes lifting_syntax begin lemma power_transfer [transfer_rule]: \(R ===> (=) ===> R) (^) (^)\ if [transfer_rule]: \R 1 1\ \(R ===> R ===> R) (*) (*)\ for R :: \'a::power \ 'b::power \ bool\ by (simp only: power_def [abs_def]) transfer_prover end context monoid_mult begin subclass power . lemma power_one [simp]: "1 ^ n = 1" by (induct n) simp_all lemma power_one_right [simp]: "a ^ 1 = a" by simp lemma power_Suc0_right [simp]: "a ^ Suc 0 = a" by simp lemma power_commutes: "a ^ n * a = a * a ^ n" by (induct n) (simp_all add: mult.assoc) lemma power_Suc2: "a ^ Suc n = a ^ n * a" by (simp add: power_commutes) lemma power_add: "a ^ (m + n) = a ^ m * a ^ n" by (induct m) (simp_all add: algebra_simps) lemma power_mult: "a ^ (m * n) = (a ^ m) ^ n" by (induct n) (simp_all add: power_add) lemma power_even_eq: "a ^ (2 * n) = (a ^ n)\<^sup>2" by (subst mult.commute) (simp add: power_mult) lemma power_odd_eq: "a ^ Suc (2*n) = a * (a ^ n)\<^sup>2" by (simp add: power_even_eq) lemma power_numeral_even: "z ^ numeral (Num.Bit0 w) = (let w = z ^ (numeral w) in w * w)" by (simp only: numeral_Bit0 power_add Let_def) lemma power_numeral_odd: "z ^ numeral (Num.Bit1 w) = (let w = z ^ (numeral w) in z * w * w)" by (simp only: numeral_Bit1 One_nat_def add_Suc_right add_0_right power_Suc power_add Let_def mult.assoc) lemma power2_eq_square: "a\<^sup>2 = a * a" by (simp add: numeral_2_eq_2) lemma power3_eq_cube: "a ^ 3 = a * a * a" by (simp add: numeral_3_eq_3 mult.assoc) lemma power4_eq_xxxx: "x^4 = x * x * x * x" by (simp add: mult.assoc power_numeral_even) lemma funpow_times_power: "(times x ^^ f x) = times (x ^ f x)" proof (induct "f x" arbitrary: f) case 0 then show ?case by (simp add: fun_eq_iff) next case (Suc n) define g where "g x = f x - 1" for x with Suc have "n = g x" by simp with Suc have "times x ^^ g x = times (x ^ g x)" by simp moreover from Suc g_def have "f x = g x + 1" by simp ultimately show ?case by (simp add: power_add funpow_add fun_eq_iff mult.assoc) qed lemma power_commuting_commutes: assumes "x * y = y * x" shows "x ^ n * y = y * x ^n" proof (induct n) case 0 then show ?case by simp next case (Suc n) have "x ^ Suc n * y = x ^ n * y * x" by (subst power_Suc2) (simp add: assms ac_simps) also have "\ = y * x ^ Suc n" by (simp only: Suc power_Suc2) (simp add: ac_simps) finally show ?case . qed lemma power_minus_mult: "0 < n \ a ^ (n - 1) * a = a ^ n" by (simp add: power_commutes split: nat_diff_split) lemma left_right_inverse_power: assumes "x * y = 1" shows "x ^ n * y ^ n = 1" proof (induct n) case (Suc n) moreover have "x ^ Suc n * y ^ Suc n = x^n * (x * y) * y^n" by (simp add: power_Suc2[symmetric] mult.assoc[symmetric]) ultimately show ?case by (simp add: assms) qed simp end context comm_monoid_mult begin lemma power_mult_distrib [algebra_simps, algebra_split_simps, field_simps, field_split_simps, divide_simps]: "(a * b) ^ n = (a ^ n) * (b ^ n)" by (induction n) (simp_all add: ac_simps) end text \Extract constant factors from powers.\ declare power_mult_distrib [where a = "numeral w" for w, simp] declare power_mult_distrib [where b = "numeral w" for w, simp] lemma power_add_numeral [simp]: "a^numeral m * a^numeral n = a^numeral (m + n)" for a :: "'a::monoid_mult" by (simp add: power_add [symmetric]) lemma power_add_numeral2 [simp]: "a^numeral m * (a^numeral n * b) = a^numeral (m + n) * b" for a :: "'a::monoid_mult" by (simp add: mult.assoc [symmetric]) lemma power_mult_numeral [simp]: "(a^numeral m)^numeral n = a^numeral (m * n)" for a :: "'a::monoid_mult" by (simp only: numeral_mult power_mult) context semiring_numeral begin lemma numeral_sqr: "numeral (Num.sqr k) = numeral k * numeral k" by (simp only: sqr_conv_mult numeral_mult) lemma numeral_pow: "numeral (Num.pow k l) = numeral k ^ numeral l" by (induct l) (simp_all only: numeral_class.numeral.simps pow.simps numeral_sqr numeral_mult power_add power_one_right) lemma power_numeral [simp]: "numeral k ^ numeral l = numeral (Num.pow k l)" by (rule numeral_pow [symmetric]) end context semiring_1 begin lemma of_nat_power [simp]: "of_nat (m ^ n) = of_nat m ^ n" by (induct n) simp_all lemma zero_power: "0 < n \ 0 ^ n = 0" by (cases n) simp_all lemma power_zero_numeral [simp]: "0 ^ numeral k = 0" by (simp add: numeral_eq_Suc) lemma zero_power2: "0\<^sup>2 = 0" (* delete? *) by (rule power_zero_numeral) lemma one_power2: "1\<^sup>2 = 1" (* delete? *) by (rule power_one) lemma power_0_Suc [simp]: "0 ^ Suc n = 0" by simp text \It looks plausible as a simprule, but its effect can be strange.\ lemma power_0_left: "0 ^ n = (if n = 0 then 1 else 0)" by (cases n) simp_all end context semiring_char_0 begin lemma numeral_power_eq_of_nat_cancel_iff [simp]: "numeral x ^ n = of_nat y \ numeral x ^ n = y" using of_nat_eq_iff by fastforce lemma real_of_nat_eq_numeral_power_cancel_iff [simp]: "of_nat y = numeral x ^ n \ y = numeral x ^ n" using numeral_power_eq_of_nat_cancel_iff [of x n y] by (metis (mono_tags)) lemma of_nat_eq_of_nat_power_cancel_iff[simp]: "(of_nat b) ^ w = of_nat x \ b ^ w = x" by (metis of_nat_power of_nat_eq_iff) lemma of_nat_power_eq_of_nat_cancel_iff[simp]: "of_nat x = (of_nat b) ^ w \ x = b ^ w" by (metis of_nat_eq_of_nat_power_cancel_iff) end context comm_semiring_1 begin text \The divides relation.\ lemma le_imp_power_dvd: assumes "m \ n" shows "a ^ m dvd a ^ n" proof from assms have "a ^ n = a ^ (m + (n - m))" by simp also have "\ = a ^ m * a ^ (n - m)" by (rule power_add) finally show "a ^ n = a ^ m * a ^ (n - m)" . qed lemma power_le_dvd: "a ^ n dvd b \ m \ n \ a ^ m dvd b" by (rule dvd_trans [OF le_imp_power_dvd]) lemma dvd_power_same: "x dvd y \ x ^ n dvd y ^ n" by (induct n) (auto simp add: mult_dvd_mono) lemma dvd_power_le: "x dvd y \ m \ n \ x ^ n dvd y ^ m" by (rule power_le_dvd [OF dvd_power_same]) lemma dvd_power [simp]: fixes n :: nat assumes "n > 0 \ x = 1" shows "x dvd (x ^ n)" using assms proof assume "0 < n" then have "x ^ n = x ^ Suc (n - 1)" by simp then show "x dvd (x ^ n)" by simp next assume "x = 1" then show "x dvd (x ^ n)" by simp qed end context semiring_1_no_zero_divisors begin subclass power . lemma power_eq_0_iff [simp]: "a ^ n = 0 \ a = 0 \ n > 0" by (induct n) auto lemma power_not_zero: "a \ 0 \ a ^ n \ 0" by (induct n) auto lemma zero_eq_power2 [simp]: "a\<^sup>2 = 0 \ a = 0" unfolding power2_eq_square by simp end context ring_1 begin lemma power_minus: "(- a) ^ n = (- 1) ^ n * a ^ n" proof (induct n) case 0 show ?case by simp next case (Suc n) then show ?case by (simp del: power_Suc add: power_Suc2 mult.assoc) qed lemma power_minus': "NO_MATCH 1 x \ (-x) ^ n = (-1)^n * x ^ n" by (rule power_minus) lemma power_minus_Bit0: "(- x) ^ numeral (Num.Bit0 k) = x ^ numeral (Num.Bit0 k)" by (induct k, simp_all only: numeral_class.numeral.simps power_add power_one_right mult_minus_left mult_minus_right minus_minus) lemma power_minus_Bit1: "(- x) ^ numeral (Num.Bit1 k) = - (x ^ numeral (Num.Bit1 k))" by (simp only: eval_nat_numeral(3) power_Suc power_minus_Bit0 mult_minus_left) lemma power2_minus [simp]: "(- a)\<^sup>2 = a\<^sup>2" by (fact power_minus_Bit0) lemma power_minus1_even [simp]: "(- 1) ^ (2*n) = 1" proof (induct n) case 0 show ?case by simp next case (Suc n) then show ?case by (simp add: power_add power2_eq_square) qed lemma power_minus1_odd: "(- 1) ^ Suc (2*n) = -1" by simp lemma power_minus_even [simp]: "(-a) ^ (2*n) = a ^ (2*n)" by (simp add: power_minus [of a]) end context ring_1_no_zero_divisors begin lemma power2_eq_1_iff: "a\<^sup>2 = 1 \ a = 1 \ a = - 1" using square_eq_1_iff [of a] by (simp add: power2_eq_square) end context idom begin lemma power2_eq_iff: "x\<^sup>2 = y\<^sup>2 \ x = y \ x = - y" unfolding power2_eq_square by (rule square_eq_iff) end context semidom_divide begin lemma power_diff: "a ^ (m - n) = (a ^ m) div (a ^ n)" if "a \ 0" and "n \ m" proof - define q where "q = m - n" with \n \ m\ have "m = q + n" by simp with \a \ 0\ q_def show ?thesis by (simp add: power_add) qed end context algebraic_semidom begin lemma div_power: "b dvd a \ (a div b) ^ n = a ^ n div b ^ n" by (induct n) (simp_all add: div_mult_div_if_dvd dvd_power_same) lemma is_unit_power_iff: "is_unit (a ^ n) \ is_unit a \ n = 0" by (induct n) (auto simp add: is_unit_mult_iff) lemma dvd_power_iff: assumes "x \ 0" shows "x ^ m dvd x ^ n \ is_unit x \ m \ n" proof assume *: "x ^ m dvd x ^ n" { assume "m > n" note * also have "x ^ n = x ^ n * 1" by simp also from \m > n\ have "m = n + (m - n)" by simp also have "x ^ \ = x ^ n * x ^ (m - n)" by (rule power_add) finally have "x ^ (m - n) dvd 1" by (subst (asm) dvd_times_left_cancel_iff) (insert assms, simp_all) with \m > n\ have "is_unit x" by (simp add: is_unit_power_iff) } thus "is_unit x \ m \ n" by force qed (auto intro: unit_imp_dvd simp: is_unit_power_iff le_imp_power_dvd) end -context normalization_semidom +context normalization_semidom_multiplicative begin lemma normalize_power: "normalize (a ^ n) = normalize a ^ n" by (induct n) (simp_all add: normalize_mult) lemma unit_factor_power: "unit_factor (a ^ n) = unit_factor a ^ n" by (induct n) (simp_all add: unit_factor_mult) end context division_ring begin text \Perhaps these should be simprules.\ lemma power_inverse [field_simps, field_split_simps, divide_simps]: "inverse a ^ n = inverse (a ^ n)" proof (cases "a = 0") case True then show ?thesis by (simp add: power_0_left) next case False then have "inverse (a ^ n) = inverse a ^ n" by (induct n) (simp_all add: nonzero_inverse_mult_distrib power_commutes) then show ?thesis by simp qed lemma power_one_over [field_simps, field_split_simps, divide_simps]: "(1 / a) ^ n = 1 / a ^ n" using power_inverse [of a] by (simp add: divide_inverse) end context field begin lemma power_divide [field_simps, field_split_simps, divide_simps]: "(a / b) ^ n = a ^ n / b ^ n" by (induct n) simp_all end subsection \Exponentiation on ordered types\ context linordered_semidom begin lemma zero_less_power [simp]: "0 < a \ 0 < a ^ n" by (induct n) simp_all lemma zero_le_power [simp]: "0 \ a \ 0 \ a ^ n" by (induct n) simp_all lemma power_mono: "a \ b \ 0 \ a \ a ^ n \ b ^ n" by (induct n) (auto intro: mult_mono order_trans [of 0 a b]) lemma one_le_power [simp]: "1 \ a \ 1 \ a ^ n" using power_mono [of 1 a n] by simp lemma power_le_one: "0 \ a \ a \ 1 \ a ^ n \ 1" using power_mono [of a 1 n] by simp lemma power_gt1_lemma: assumes gt1: "1 < a" shows "1 < a * a ^ n" proof - from gt1 have "0 \ a" by (fact order_trans [OF zero_le_one less_imp_le]) from gt1 have "1 * 1 < a * 1" by simp also from gt1 have "\ \ a * a ^ n" by (simp only: mult_mono \0 \ a\ one_le_power order_less_imp_le zero_le_one order_refl) finally show ?thesis by simp qed lemma power_gt1: "1 < a \ 1 < a ^ Suc n" by (simp add: power_gt1_lemma) lemma one_less_power [simp]: "1 < a \ 0 < n \ 1 < a ^ n" by (cases n) (simp_all add: power_gt1_lemma) lemma power_le_imp_le_exp: assumes gt1: "1 < a" shows "a ^ m \ a ^ n \ m \ n" proof (induct m arbitrary: n) case 0 show ?case by simp next case (Suc m) show ?case proof (cases n) case 0 with Suc have "a * a ^ m \ 1" by simp with gt1 show ?thesis by (force simp only: power_gt1_lemma not_less [symmetric]) next case (Suc n) with Suc.prems Suc.hyps show ?thesis by (force dest: mult_left_le_imp_le simp add: less_trans [OF zero_less_one gt1]) qed qed lemma of_nat_zero_less_power_iff [simp]: "of_nat x ^ n > 0 \ x > 0 \ n = 0" by (induct n) auto text \Surely we can strengthen this? It holds for \0 too.\ lemma power_inject_exp [simp]: "1 < a \ a ^ m = a ^ n \ m = n" by (force simp add: order_antisym power_le_imp_le_exp) text \ Can relax the first premise to \<^term>\0 in the case of the natural numbers. \ lemma power_less_imp_less_exp: "1 < a \ a ^ m < a ^ n \ m < n" by (simp add: order_less_le [of m n] less_le [of "a^m" "a^n"] power_le_imp_le_exp) lemma power_strict_mono [rule_format]: "a < b \ 0 \ a \ 0 < n \ a ^ n < b ^ n" by (induct n) (auto simp: mult_strict_mono le_less_trans [of 0 a b]) lemma power_mono_iff [simp]: shows "\a \ 0; b \ 0; n>0\ \ a ^ n \ b ^ n \ a \ b" using power_mono [of a b] power_strict_mono [of b a] not_le by auto text\Lemma for \power_strict_decreasing\\ lemma power_Suc_less: "0 < a \ a < 1 \ a * a ^ n < a ^ n" by (induct n) (auto simp: mult_strict_left_mono) lemma power_strict_decreasing [rule_format]: "n < N \ 0 < a \ a < 1 \ a ^ N < a ^ n" proof (induct N) case 0 then show ?case by simp next case (Suc N) then show ?case apply (auto simp add: power_Suc_less less_Suc_eq) apply (subgoal_tac "a * a^N < 1 * a^n") apply simp apply (rule mult_strict_mono) apply auto done qed text \Proof resembles that of \power_strict_decreasing\.\ lemma power_decreasing: "n \ N \ 0 \ a \ a \ 1 \ a ^ N \ a ^ n" proof (induct N) case 0 then show ?case by simp next case (Suc N) then show ?case apply (auto simp add: le_Suc_eq) apply (subgoal_tac "a * a^N \ 1 * a^n") apply simp apply (rule mult_mono) apply auto done qed lemma power_decreasing_iff [simp]: "\0 < b; b < 1\ \ b ^ m \ b ^ n \ n \ m" using power_strict_decreasing [of m n b] by (auto intro: power_decreasing ccontr) lemma power_strict_decreasing_iff [simp]: "\0 < b; b < 1\ \ b ^ m < b ^ n \ n < m" using power_decreasing_iff [of b m n] unfolding le_less by (auto dest: power_strict_decreasing le_neq_implies_less) lemma power_Suc_less_one: "0 < a \ a < 1 \ a ^ Suc n < 1" using power_strict_decreasing [of 0 "Suc n" a] by simp text \Proof again resembles that of \power_strict_decreasing\.\ lemma power_increasing: "n \ N \ 1 \ a \ a ^ n \ a ^ N" proof (induct N) case 0 then show ?case by simp next case (Suc N) then show ?case apply (auto simp add: le_Suc_eq) apply (subgoal_tac "1 * a^n \ a * a^N") apply simp apply (rule mult_mono) apply (auto simp add: order_trans [OF zero_le_one]) done qed text \Lemma for \power_strict_increasing\.\ lemma power_less_power_Suc: "1 < a \ a ^ n < a * a ^ n" by (induct n) (auto simp: mult_strict_left_mono less_trans [OF zero_less_one]) lemma power_strict_increasing: "n < N \ 1 < a \ a ^ n < a ^ N" proof (induct N) case 0 then show ?case by simp next case (Suc N) then show ?case apply (auto simp add: power_less_power_Suc less_Suc_eq) apply (subgoal_tac "1 * a^n < a * a^N") apply simp apply (rule mult_strict_mono) apply (auto simp add: less_trans [OF zero_less_one] less_imp_le) done qed lemma power_increasing_iff [simp]: "1 < b \ b ^ x \ b ^ y \ x \ y" by (blast intro: power_le_imp_le_exp power_increasing less_imp_le) lemma power_strict_increasing_iff [simp]: "1 < b \ b ^ x < b ^ y \ x < y" by (blast intro: power_less_imp_less_exp power_strict_increasing) lemma power_le_imp_le_base: assumes le: "a ^ Suc n \ b ^ Suc n" and "0 \ b" shows "a \ b" proof (rule ccontr) assume "\ ?thesis" then have "b < a" by (simp only: linorder_not_le) then have "b ^ Suc n < a ^ Suc n" by (simp only: assms(2) power_strict_mono) with le show False by (simp add: linorder_not_less [symmetric]) qed lemma power_less_imp_less_base: assumes less: "a ^ n < b ^ n" assumes nonneg: "0 \ b" shows "a < b" proof (rule contrapos_pp [OF less]) assume "\ ?thesis" then have "b \ a" by (simp only: linorder_not_less) from this nonneg have "b ^ n \ a ^ n" by (rule power_mono) then show "\ a ^ n < b ^ n" by (simp only: linorder_not_less) qed lemma power_inject_base: "a ^ Suc n = b ^ Suc n \ 0 \ a \ 0 \ b \ a = b" by (blast intro: power_le_imp_le_base antisym eq_refl sym) lemma power_eq_imp_eq_base: "a ^ n = b ^ n \ 0 \ a \ 0 \ b \ 0 < n \ a = b" by (cases n) (simp_all del: power_Suc, rule power_inject_base) lemma power_eq_iff_eq_base: "0 < n \ 0 \ a \ 0 \ b \ a ^ n = b ^ n \ a = b" using power_eq_imp_eq_base [of a n b] by auto lemma power2_le_imp_le: "x\<^sup>2 \ y\<^sup>2 \ 0 \ y \ x \ y" unfolding numeral_2_eq_2 by (rule power_le_imp_le_base) lemma power2_less_imp_less: "x\<^sup>2 < y\<^sup>2 \ 0 \ y \ x < y" by (rule power_less_imp_less_base) lemma power2_eq_imp_eq: "x\<^sup>2 = y\<^sup>2 \ 0 \ x \ 0 \ y \ x = y" unfolding numeral_2_eq_2 by (erule (2) power_eq_imp_eq_base) simp lemma power_Suc_le_self: "0 \ a \ a \ 1 \ a ^ Suc n \ a" using power_decreasing [of 1 "Suc n" a] by simp lemma power2_eq_iff_nonneg [simp]: assumes "0 \ x" "0 \ y" shows "(x ^ 2 = y ^ 2) \ x = y" using assms power2_eq_imp_eq by blast lemma of_nat_less_numeral_power_cancel_iff[simp]: "of_nat x < numeral i ^ n \ x < numeral i ^ n" using of_nat_less_iff[of x "numeral i ^ n", unfolded of_nat_numeral of_nat_power] . lemma of_nat_le_numeral_power_cancel_iff[simp]: "of_nat x \ numeral i ^ n \ x \ numeral i ^ n" using of_nat_le_iff[of x "numeral i ^ n", unfolded of_nat_numeral of_nat_power] . lemma numeral_power_less_of_nat_cancel_iff[simp]: "numeral i ^ n < of_nat x \ numeral i ^ n < x" using of_nat_less_iff[of "numeral i ^ n" x, unfolded of_nat_numeral of_nat_power] . lemma numeral_power_le_of_nat_cancel_iff[simp]: "numeral i ^ n \ of_nat x \ numeral i ^ n \ x" using of_nat_le_iff[of "numeral i ^ n" x, unfolded of_nat_numeral of_nat_power] . lemma of_nat_le_of_nat_power_cancel_iff[simp]: "(of_nat b) ^ w \ of_nat x \ b ^ w \ x" by (metis of_nat_le_iff of_nat_power) lemma of_nat_power_le_of_nat_cancel_iff[simp]: "of_nat x \ (of_nat b) ^ w \ x \ b ^ w" by (metis of_nat_le_iff of_nat_power) lemma of_nat_less_of_nat_power_cancel_iff[simp]: "(of_nat b) ^ w < of_nat x \ b ^ w < x" by (metis of_nat_less_iff of_nat_power) lemma of_nat_power_less_of_nat_cancel_iff[simp]: "of_nat x < (of_nat b) ^ w \ x < b ^ w" by (metis of_nat_less_iff of_nat_power) end text \Some @{typ nat}-specific lemmas:\ lemma mono_ge2_power_minus_self: assumes "k \ 2" shows "mono (\m. k ^ m - m)" unfolding mono_iff_le_Suc proof fix n have "k ^ n < k ^ Suc n" using power_strict_increasing_iff[of k "n" "Suc n"] assms by linarith thus "k ^ n - n \ k ^ Suc n - Suc n" by linarith qed lemma self_le_ge2_pow[simp]: assumes "k \ 2" shows "m \ k ^ m" proof (induction m) case 0 show ?case by simp next case (Suc m) hence "Suc m \ Suc (k ^ m)" by simp also have "... \ k^m + k^m" using one_le_power[of k m] assms by linarith also have "... \ k * k^m" by (metis mult_2 mult_le_mono1[OF assms]) finally show ?case by simp qed lemma diff_le_diff_pow[simp]: assumes "k \ 2" shows "m - n \ k ^ m - k ^ n" proof (cases "n \ m") case True thus ?thesis using monoD[OF mono_ge2_power_minus_self[OF assms] True] self_le_ge2_pow[OF assms, of m] by (simp add: le_diff_conv le_diff_conv2) qed auto context linordered_ring_strict begin lemma sum_squares_eq_zero_iff: "x * x + y * y = 0 \ x = 0 \ y = 0" by (simp add: add_nonneg_eq_0_iff) lemma sum_squares_le_zero_iff: "x * x + y * y \ 0 \ x = 0 \ y = 0" by (simp add: le_less not_sum_squares_lt_zero sum_squares_eq_zero_iff) lemma sum_squares_gt_zero_iff: "0 < x * x + y * y \ x \ 0 \ y \ 0" by (simp add: not_le [symmetric] sum_squares_le_zero_iff) end context linordered_idom begin lemma zero_le_power2 [simp]: "0 \ a\<^sup>2" by (simp add: power2_eq_square) lemma zero_less_power2 [simp]: "0 < a\<^sup>2 \ a \ 0" by (force simp add: power2_eq_square zero_less_mult_iff linorder_neq_iff) lemma power2_less_0 [simp]: "\ a\<^sup>2 < 0" by (force simp add: power2_eq_square mult_less_0_iff) lemma power_abs: "\a ^ n\ = \a\ ^ n" \ \FIXME simp?\ by (induct n) (simp_all add: abs_mult) lemma power_sgn [simp]: "sgn (a ^ n) = sgn a ^ n" by (induct n) (simp_all add: sgn_mult) lemma abs_power_minus [simp]: "\(- a) ^ n\ = \a ^ n\" by (simp add: power_abs) lemma zero_less_power_abs_iff [simp]: "0 < \a\ ^ n \ a \ 0 \ n = 0" proof (induct n) case 0 show ?case by simp next case Suc then show ?case by (auto simp: zero_less_mult_iff) qed lemma zero_le_power_abs [simp]: "0 \ \a\ ^ n" by (rule zero_le_power [OF abs_ge_zero]) lemma power2_less_eq_zero_iff [simp]: "a\<^sup>2 \ 0 \ a = 0" by (simp add: le_less) lemma abs_power2 [simp]: "\a\<^sup>2\ = a\<^sup>2" by (simp add: power2_eq_square) lemma power2_abs [simp]: "\a\\<^sup>2 = a\<^sup>2" by (simp add: power2_eq_square) lemma odd_power_less_zero: "a < 0 \ a ^ Suc (2 * n) < 0" proof (induct n) case 0 then show ?case by simp next case (Suc n) have "a ^ Suc (2 * Suc n) = (a*a) * a ^ Suc(2*n)" by (simp add: ac_simps power_add power2_eq_square) then show ?case by (simp del: power_Suc add: Suc mult_less_0_iff mult_neg_neg) qed lemma odd_0_le_power_imp_0_le: "0 \ a ^ Suc (2 * n) \ 0 \ a" using odd_power_less_zero [of a n] by (force simp add: linorder_not_less [symmetric]) lemma zero_le_even_power'[simp]: "0 \ a ^ (2 * n)" proof (induct n) case 0 show ?case by simp next case (Suc n) have "a ^ (2 * Suc n) = (a*a) * a ^ (2*n)" by (simp add: ac_simps power_add power2_eq_square) then show ?case by (simp add: Suc zero_le_mult_iff) qed lemma sum_power2_ge_zero: "0 \ x\<^sup>2 + y\<^sup>2" by (intro add_nonneg_nonneg zero_le_power2) lemma not_sum_power2_lt_zero: "\ x\<^sup>2 + y\<^sup>2 < 0" unfolding not_less by (rule sum_power2_ge_zero) lemma sum_power2_eq_zero_iff: "x\<^sup>2 + y\<^sup>2 = 0 \ x = 0 \ y = 0" unfolding power2_eq_square by (simp add: add_nonneg_eq_0_iff) lemma sum_power2_le_zero_iff: "x\<^sup>2 + y\<^sup>2 \ 0 \ x = 0 \ y = 0" by (simp add: le_less sum_power2_eq_zero_iff not_sum_power2_lt_zero) lemma sum_power2_gt_zero_iff: "0 < x\<^sup>2 + y\<^sup>2 \ x \ 0 \ y \ 0" unfolding not_le [symmetric] by (simp add: sum_power2_le_zero_iff) lemma abs_le_square_iff: "\x\ \ \y\ \ x\<^sup>2 \ y\<^sup>2" (is "?lhs \ ?rhs") proof assume ?lhs then have "\x\\<^sup>2 \ \y\\<^sup>2" by (rule power_mono) simp then show ?rhs by simp next assume ?rhs then show ?lhs by (auto intro!: power2_le_imp_le [OF _ abs_ge_zero]) qed lemma abs_square_le_1:"x\<^sup>2 \ 1 \ \x\ \ 1" using abs_le_square_iff [of x 1] by simp lemma abs_square_eq_1: "x\<^sup>2 = 1 \ \x\ = 1" by (auto simp add: abs_if power2_eq_1_iff) lemma abs_square_less_1: "x\<^sup>2 < 1 \ \x\ < 1" using abs_square_eq_1 [of x] abs_square_le_1 [of x] by (auto simp add: le_less) lemma square_le_1: assumes "- 1 \ x" "x \ 1" shows "x\<^sup>2 \ 1" using assms by (metis add.inverse_inverse linear mult_le_one neg_equal_0_iff_equal neg_le_iff_le power2_eq_square power_minus_Bit0) end subsection \Miscellaneous rules\ lemma (in linordered_semidom) self_le_power: "1 \ a \ 0 < n \ a \ a ^ n" using power_increasing [of 1 n a] power_one_right [of a] by auto lemma (in power) power_eq_if: "p ^ m = (if m=0 then 1 else p * (p ^ (m - 1)))" unfolding One_nat_def by (cases m) simp_all lemma (in comm_semiring_1) power2_sum: "(x + y)\<^sup>2 = x\<^sup>2 + y\<^sup>2 + 2 * x * y" by (simp add: algebra_simps power2_eq_square mult_2_right) context comm_ring_1 begin lemma power2_diff: "(x - y)\<^sup>2 = x\<^sup>2 + y\<^sup>2 - 2 * x * y" by (simp add: algebra_simps power2_eq_square mult_2_right) lemma power2_commute: "(x - y)\<^sup>2 = (y - x)\<^sup>2" by (simp add: algebra_simps power2_eq_square) lemma minus_power_mult_self: "(- a) ^ n * (- a) ^ n = a ^ (2 * n)" by (simp add: power_mult_distrib [symmetric]) (simp add: power2_eq_square [symmetric] power_mult [symmetric]) lemma minus_one_mult_self [simp]: "(- 1) ^ n * (- 1) ^ n = 1" using minus_power_mult_self [of 1 n] by simp lemma left_minus_one_mult_self [simp]: "(- 1) ^ n * ((- 1) ^ n * a) = a" by (simp add: mult.assoc [symmetric]) end text \Simprules for comparisons where common factors can be cancelled.\ lemmas zero_compare_simps = add_strict_increasing add_strict_increasing2 add_increasing zero_le_mult_iff zero_le_divide_iff zero_less_mult_iff zero_less_divide_iff mult_le_0_iff divide_le_0_iff mult_less_0_iff divide_less_0_iff zero_le_power2 power2_less_0 subsection \Exponentiation for the Natural Numbers\ lemma nat_one_le_power [simp]: "Suc 0 \ i \ Suc 0 \ i ^ n" by (rule one_le_power [of i n, unfolded One_nat_def]) lemma nat_zero_less_power_iff [simp]: "x ^ n > 0 \ x > 0 \ n = 0" for x :: nat by (induct n) auto lemma nat_power_eq_Suc_0_iff [simp]: "x ^ m = Suc 0 \ m = 0 \ x = Suc 0" by (induct m) auto lemma power_Suc_0 [simp]: "Suc 0 ^ n = Suc 0" by simp text \ Valid for the naturals, but what if \0 < i < 1\? Premises cannot be weakened: consider the case where \i = 0\, \m = 1\ and \n = 0\. \ lemma nat_power_less_imp_less: fixes i :: nat assumes nonneg: "0 < i" assumes less: "i ^ m < i ^ n" shows "m < n" proof (cases "i = 1") case True with less power_one [where 'a = nat] show ?thesis by simp next case False with nonneg have "1 < i" by auto from power_strict_increasing_iff [OF this] less show ?thesis .. qed lemma power_dvd_imp_le: "i ^ m dvd i ^ n \ 1 < i \ m \ n" for i m n :: nat apply (rule power_le_imp_le_exp) apply assumption apply (erule dvd_imp_le) apply simp done lemma dvd_power_iff_le: fixes k::nat shows "2 \ k \ ((k ^ m) dvd (k ^ n) \ m \ n)" using le_imp_power_dvd power_dvd_imp_le by force lemma power2_nat_le_eq_le: "m\<^sup>2 \ n\<^sup>2 \ m \ n" for m n :: nat by (auto intro: power2_le_imp_le power_mono) lemma power2_nat_le_imp_le: fixes m n :: nat assumes "m\<^sup>2 \ n" shows "m \ n" proof (cases m) case 0 then show ?thesis by simp next case (Suc k) show ?thesis proof (rule ccontr) assume "\ ?thesis" then have "n < m" by simp with assms Suc show False by (simp add: power2_eq_square) qed qed lemma ex_power_ivl1: fixes b k :: nat assumes "b \ 2" shows "k \ 1 \ \n. b^n \ k \ k < b^(n+1)" (is "_ \ \n. ?P k n") proof(induction k) case 0 thus ?case by simp next case (Suc k) show ?case proof cases assume "k=0" hence "?P (Suc k) 0" using assms by simp thus ?case .. next assume "k\0" with Suc obtain n where IH: "?P k n" by auto show ?case proof (cases "k = b^(n+1) - 1") case True hence "?P (Suc k) (n+1)" using assms by (simp add: power_less_power_Suc) thus ?thesis .. next case False hence "?P (Suc k) n" using IH by auto thus ?thesis .. qed qed qed lemma ex_power_ivl2: fixes b k :: nat assumes "b \ 2" "k \ 2" shows "\n. b^n < k \ k \ b^(n+1)" proof - have "1 \ k - 1" using assms(2) by arith from ex_power_ivl1[OF assms(1) this] obtain n where "b ^ n \ k - 1 \ k - 1 < b ^ (n + 1)" .. hence "b^n < k \ k \ b^(n+1)" using assms by auto thus ?thesis .. qed subsubsection \Cardinality of the Powerset\ lemma card_UNIV_bool [simp]: "card (UNIV :: bool set) = 2" unfolding UNIV_bool by simp lemma card_Pow: "finite A \ card (Pow A) = 2 ^ card A" proof (induct rule: finite_induct) case empty show ?case by simp next case (insert x A) from \x \ A\ have disjoint: "Pow A \ insert x ` Pow A = {}" by blast from \x \ A\ have inj_on: "inj_on (insert x) (Pow A)" unfolding inj_on_def by auto have "card (Pow (insert x A)) = card (Pow A \ insert x ` Pow A)" by (simp only: Pow_insert) also have "\ = card (Pow A) + card (insert x ` Pow A)" by (rule card_Un_disjoint) (use \finite A\ disjoint in simp_all) also from inj_on have "card (insert x ` Pow A) = card (Pow A)" by (rule card_image) also have "\ + \ = 2 * \" by (simp add: mult_2) also from insert(3) have "\ = 2 ^ Suc (card A)" by simp also from insert(1,2) have "Suc (card A) = card (insert x A)" by (rule card_insert_disjoint [symmetric]) finally show ?case . qed subsection \Code generator tweak\ code_identifier code_module Power \ (SML) Arith and (OCaml) Arith and (Haskell) Arith end diff --git a/src/HOL/Rings.thy b/src/HOL/Rings.thy --- a/src/HOL/Rings.thy +++ b/src/HOL/Rings.thy @@ -1,2631 +1,2670 @@ (* Title: HOL/Rings.thy Author: Gertrud Bauer Author: Steven Obua Author: Tobias Nipkow Author: Lawrence C Paulson Author: Markus Wenzel Author: Jeremy Avigad *) section \Rings\ theory Rings imports Groups Set Fun begin subsection \Semirings and rings\ class semiring = ab_semigroup_add + semigroup_mult + assumes distrib_right [algebra_simps, algebra_split_simps]: "(a + b) * c = a * c + b * c" assumes distrib_left [algebra_simps, algebra_split_simps]: "a * (b + c) = a * b + a * c" begin text \For the \combine_numerals\ simproc\ lemma combine_common_factor: "a * e + (b * e + c) = (a + b) * e + c" by (simp add: distrib_right ac_simps) end class mult_zero = times + zero + assumes mult_zero_left [simp]: "0 * a = 0" assumes mult_zero_right [simp]: "a * 0 = 0" begin lemma mult_not_zero: "a * b \ 0 \ a \ 0 \ b \ 0" by auto end class semiring_0 = semiring + comm_monoid_add + mult_zero class semiring_0_cancel = semiring + cancel_comm_monoid_add begin subclass semiring_0 proof fix a :: 'a have "0 * a + 0 * a = 0 * a + 0" by (simp add: distrib_right [symmetric]) then show "0 * a = 0" by (simp only: add_left_cancel) have "a * 0 + a * 0 = a * 0 + 0" by (simp add: distrib_left [symmetric]) then show "a * 0 = 0" by (simp only: add_left_cancel) qed end class comm_semiring = ab_semigroup_add + ab_semigroup_mult + assumes distrib: "(a + b) * c = a * c + b * c" begin subclass semiring proof fix a b c :: 'a show "(a + b) * c = a * c + b * c" by (simp add: distrib) have "a * (b + c) = (b + c) * a" by (simp add: ac_simps) also have "\ = b * a + c * a" by (simp only: distrib) also have "\ = a * b + a * c" by (simp add: ac_simps) finally show "a * (b + c) = a * b + a * c" by blast qed end class comm_semiring_0 = comm_semiring + comm_monoid_add + mult_zero begin subclass semiring_0 .. end class comm_semiring_0_cancel = comm_semiring + cancel_comm_monoid_add begin subclass semiring_0_cancel .. subclass comm_semiring_0 .. end class zero_neq_one = zero + one + assumes zero_neq_one [simp]: "0 \ 1" begin lemma one_neq_zero [simp]: "1 \ 0" by (rule not_sym) (rule zero_neq_one) definition of_bool :: "bool \ 'a" where "of_bool p = (if p then 1 else 0)" lemma of_bool_eq [simp, code]: "of_bool False = 0" "of_bool True = 1" by (simp_all add: of_bool_def) lemma of_bool_eq_iff: "of_bool p = of_bool q \ p = q" by (simp add: of_bool_def) lemma split_of_bool [split]: "P (of_bool p) \ (p \ P 1) \ (\ p \ P 0)" by (cases p) simp_all lemma split_of_bool_asm: "P (of_bool p) \ \ (p \ \ P 1 \ \ p \ \ P 0)" by (cases p) simp_all end class semiring_1 = zero_neq_one + semiring_0 + monoid_mult begin lemma of_bool_conj: "of_bool (P \ Q) = of_bool P * of_bool Q" by auto end lemma lambda_zero: "(\h::'a::mult_zero. 0) = (*) 0" by auto lemma lambda_one: "(\x::'a::monoid_mult. x) = (*) 1" by auto subsection \Abstract divisibility\ class dvd = times begin definition dvd :: "'a \ 'a \ bool" (infix "dvd" 50) where "b dvd a \ (\k. a = b * k)" lemma dvdI [intro?]: "a = b * k \ b dvd a" unfolding dvd_def .. lemma dvdE [elim]: "b dvd a \ (\k. a = b * k \ P) \ P" unfolding dvd_def by blast end context comm_monoid_mult begin subclass dvd . lemma dvd_refl [simp]: "a dvd a" proof show "a = a * 1" by simp qed lemma dvd_trans [trans]: assumes "a dvd b" and "b dvd c" shows "a dvd c" proof - from assms obtain v where "b = a * v" by auto moreover from assms obtain w where "c = b * w" by auto ultimately have "c = a * (v * w)" by (simp add: mult.assoc) then show ?thesis .. qed lemma subset_divisors_dvd: "{c. c dvd a} \ {c. c dvd b} \ a dvd b" by (auto simp add: subset_iff intro: dvd_trans) lemma strict_subset_divisors_dvd: "{c. c dvd a} \ {c. c dvd b} \ a dvd b \ \ b dvd a" by (auto simp add: subset_iff intro: dvd_trans) lemma one_dvd [simp]: "1 dvd a" by (auto intro: dvdI) lemma dvd_mult [simp]: "a dvd (b * c)" if "a dvd c" using that by rule (auto intro: mult.left_commute dvdI) lemma dvd_mult2 [simp]: "a dvd (b * c)" if "a dvd b" using that dvd_mult [of a b c] by (simp add: ac_simps) lemma dvd_triv_right [simp]: "a dvd b * a" by (rule dvd_mult) (rule dvd_refl) lemma dvd_triv_left [simp]: "a dvd a * b" by (rule dvd_mult2) (rule dvd_refl) lemma mult_dvd_mono: assumes "a dvd b" and "c dvd d" shows "a * c dvd b * d" proof - from \a dvd b\ obtain b' where "b = a * b'" .. moreover from \c dvd d\ obtain d' where "d = c * d'" .. ultimately have "b * d = (a * c) * (b' * d')" by (simp add: ac_simps) then show ?thesis .. qed lemma dvd_mult_left: "a * b dvd c \ a dvd c" by (simp add: dvd_def mult.assoc) blast lemma dvd_mult_right: "a * b dvd c \ b dvd c" using dvd_mult_left [of b a c] by (simp add: ac_simps) end class comm_semiring_1 = zero_neq_one + comm_semiring_0 + comm_monoid_mult begin subclass semiring_1 .. lemma dvd_0_left_iff [simp]: "0 dvd a \ a = 0" by auto lemma dvd_0_right [iff]: "a dvd 0" proof show "0 = a * 0" by simp qed lemma dvd_0_left: "0 dvd a \ a = 0" by simp lemma dvd_add [simp]: assumes "a dvd b" and "a dvd c" shows "a dvd (b + c)" proof - from \a dvd b\ obtain b' where "b = a * b'" .. moreover from \a dvd c\ obtain c' where "c = a * c'" .. ultimately have "b + c = a * (b' + c')" by (simp add: distrib_left) then show ?thesis .. qed end class semiring_1_cancel = semiring + cancel_comm_monoid_add + zero_neq_one + monoid_mult begin subclass semiring_0_cancel .. subclass semiring_1 .. end class comm_semiring_1_cancel = comm_semiring + cancel_comm_monoid_add + zero_neq_one + comm_monoid_mult + assumes right_diff_distrib' [algebra_simps, algebra_split_simps]: "a * (b - c) = a * b - a * c" begin subclass semiring_1_cancel .. subclass comm_semiring_0_cancel .. subclass comm_semiring_1 .. lemma left_diff_distrib' [algebra_simps, algebra_split_simps]: "(b - c) * a = b * a - c * a" by (simp add: algebra_simps) lemma dvd_add_times_triv_left_iff [simp]: "a dvd c * a + b \ a dvd b" proof - have "a dvd a * c + b \ a dvd b" (is "?P \ ?Q") proof assume ?Q then show ?P by simp next assume ?P then obtain d where "a * c + b = a * d" .. then have "a * c + b - a * c = a * d - a * c" by simp then have "b = a * d - a * c" by simp then have "b = a * (d - c)" by (simp add: algebra_simps) then show ?Q .. qed then show "a dvd c * a + b \ a dvd b" by (simp add: ac_simps) qed lemma dvd_add_times_triv_right_iff [simp]: "a dvd b + c * a \ a dvd b" using dvd_add_times_triv_left_iff [of a c b] by (simp add: ac_simps) lemma dvd_add_triv_left_iff [simp]: "a dvd a + b \ a dvd b" using dvd_add_times_triv_left_iff [of a 1 b] by simp lemma dvd_add_triv_right_iff [simp]: "a dvd b + a \ a dvd b" using dvd_add_times_triv_right_iff [of a b 1] by simp lemma dvd_add_right_iff: assumes "a dvd b" shows "a dvd b + c \ a dvd c" (is "?P \ ?Q") proof assume ?P then obtain d where "b + c = a * d" .. moreover from \a dvd b\ obtain e where "b = a * e" .. ultimately have "a * e + c = a * d" by simp then have "a * e + c - a * e = a * d - a * e" by simp then have "c = a * d - a * e" by simp then have "c = a * (d - e)" by (simp add: algebra_simps) then show ?Q .. next assume ?Q with assms show ?P by simp qed lemma dvd_add_left_iff: "a dvd c \ a dvd b + c \ a dvd b" using dvd_add_right_iff [of a c b] by (simp add: ac_simps) end class ring = semiring + ab_group_add begin subclass semiring_0_cancel .. text \Distribution rules\ lemma minus_mult_left: "- (a * b) = - a * b" by (rule minus_unique) (simp add: distrib_right [symmetric]) lemma minus_mult_right: "- (a * b) = a * - b" by (rule minus_unique) (simp add: distrib_left [symmetric]) text \Extract signs from products\ lemmas mult_minus_left [simp] = minus_mult_left [symmetric] lemmas mult_minus_right [simp] = minus_mult_right [symmetric] lemma minus_mult_minus [simp]: "- a * - b = a * b" by simp lemma minus_mult_commute: "- a * b = a * - b" by simp lemma right_diff_distrib [algebra_simps, algebra_split_simps]: "a * (b - c) = a * b - a * c" using distrib_left [of a b "-c "] by simp lemma left_diff_distrib [algebra_simps, algebra_split_simps]: "(a - b) * c = a * c - b * c" using distrib_right [of a "- b" c] by simp lemmas ring_distribs = distrib_left distrib_right left_diff_distrib right_diff_distrib lemma eq_add_iff1: "a * e + c = b * e + d \ (a - b) * e + c = d" by (simp add: algebra_simps) lemma eq_add_iff2: "a * e + c = b * e + d \ c = (b - a) * e + d" by (simp add: algebra_simps) end lemmas ring_distribs = distrib_left distrib_right left_diff_distrib right_diff_distrib class comm_ring = comm_semiring + ab_group_add begin subclass ring .. subclass comm_semiring_0_cancel .. lemma square_diff_square_factored: "x * x - y * y = (x + y) * (x - y)" by (simp add: algebra_simps) end class ring_1 = ring + zero_neq_one + monoid_mult begin subclass semiring_1_cancel .. lemma square_diff_one_factored: "x * x - 1 = (x + 1) * (x - 1)" by (simp add: algebra_simps) end class comm_ring_1 = comm_ring + zero_neq_one + comm_monoid_mult begin subclass ring_1 .. subclass comm_semiring_1_cancel by standard (simp add: algebra_simps) lemma dvd_minus_iff [simp]: "x dvd - y \ x dvd y" proof assume "x dvd - y" then have "x dvd - 1 * - y" by (rule dvd_mult) then show "x dvd y" by simp next assume "x dvd y" then have "x dvd - 1 * y" by (rule dvd_mult) then show "x dvd - y" by simp qed lemma minus_dvd_iff [simp]: "- x dvd y \ x dvd y" proof assume "- x dvd y" then obtain k where "y = - x * k" .. then have "y = x * - k" by simp then show "x dvd y" .. next assume "x dvd y" then obtain k where "y = x * k" .. then have "y = - x * - k" by simp then show "- x dvd y" .. qed lemma dvd_diff [simp]: "x dvd y \ x dvd z \ x dvd (y - z)" using dvd_add [of x y "- z"] by simp end subsection \Towards integral domains\ class semiring_no_zero_divisors = semiring_0 + assumes no_zero_divisors: "a \ 0 \ b \ 0 \ a * b \ 0" begin lemma divisors_zero: assumes "a * b = 0" shows "a = 0 \ b = 0" proof (rule classical) assume "\ ?thesis" then have "a \ 0" and "b \ 0" by auto with no_zero_divisors have "a * b \ 0" by blast with assms show ?thesis by simp qed lemma mult_eq_0_iff [simp]: "a * b = 0 \ a = 0 \ b = 0" proof (cases "a = 0 \ b = 0") case False then have "a \ 0" and "b \ 0" by auto then show ?thesis using no_zero_divisors by simp next case True then show ?thesis by auto qed end class semiring_1_no_zero_divisors = semiring_1 + semiring_no_zero_divisors class semiring_no_zero_divisors_cancel = semiring_no_zero_divisors + assumes mult_cancel_right [simp]: "a * c = b * c \ c = 0 \ a = b" and mult_cancel_left [simp]: "c * a = c * b \ c = 0 \ a = b" begin lemma mult_left_cancel: "c \ 0 \ c * a = c * b \ a = b" by simp lemma mult_right_cancel: "c \ 0 \ a * c = b * c \ a = b" by simp end class ring_no_zero_divisors = ring + semiring_no_zero_divisors begin subclass semiring_no_zero_divisors_cancel proof fix a b c have "a * c = b * c \ (a - b) * c = 0" by (simp add: algebra_simps) also have "\ \ c = 0 \ a = b" by auto finally show "a * c = b * c \ c = 0 \ a = b" . have "c * a = c * b \ c * (a - b) = 0" by (simp add: algebra_simps) also have "\ \ c = 0 \ a = b" by auto finally show "c * a = c * b \ c = 0 \ a = b" . qed end class ring_1_no_zero_divisors = ring_1 + ring_no_zero_divisors begin subclass semiring_1_no_zero_divisors .. lemma square_eq_1_iff: "x * x = 1 \ x = 1 \ x = - 1" proof - have "(x - 1) * (x + 1) = x * x - 1" by (simp add: algebra_simps) then have "x * x = 1 \ (x - 1) * (x + 1) = 0" by simp then show ?thesis by (simp add: eq_neg_iff_add_eq_0) qed lemma mult_cancel_right1 [simp]: "c = b * c \ c = 0 \ b = 1" using mult_cancel_right [of 1 c b] by auto lemma mult_cancel_right2 [simp]: "a * c = c \ c = 0 \ a = 1" using mult_cancel_right [of a c 1] by simp lemma mult_cancel_left1 [simp]: "c = c * b \ c = 0 \ b = 1" using mult_cancel_left [of c 1 b] by force lemma mult_cancel_left2 [simp]: "c * a = c \ c = 0 \ a = 1" using mult_cancel_left [of c a 1] by simp end class semidom = comm_semiring_1_cancel + semiring_no_zero_divisors begin subclass semiring_1_no_zero_divisors .. end class idom = comm_ring_1 + semiring_no_zero_divisors begin subclass semidom .. subclass ring_1_no_zero_divisors .. lemma dvd_mult_cancel_right [simp]: "a * c dvd b * c \ c = 0 \ a dvd b" proof - have "a * c dvd b * c \ (\k. b * c = (a * k) * c)" by (auto simp add: ac_simps) also have "(\k. b * c = (a * k) * c) \ c = 0 \ a dvd b" by auto finally show ?thesis . qed lemma dvd_mult_cancel_left [simp]: "c * a dvd c * b \ c = 0 \ a dvd b" using dvd_mult_cancel_right [of a c b] by (simp add: ac_simps) lemma square_eq_iff: "a * a = b * b \ a = b \ a = - b" proof assume "a * a = b * b" then have "(a - b) * (a + b) = 0" by (simp add: algebra_simps) then show "a = b \ a = - b" by (simp add: eq_neg_iff_add_eq_0) next assume "a = b \ a = - b" then show "a * a = b * b" by auto qed end class idom_abs_sgn = idom + abs + sgn + assumes sgn_mult_abs: "sgn a * \a\ = a" and sgn_sgn [simp]: "sgn (sgn a) = sgn a" and abs_abs [simp]: "\\a\\ = \a\" and abs_0 [simp]: "\0\ = 0" and sgn_0 [simp]: "sgn 0 = 0" and sgn_1 [simp]: "sgn 1 = 1" and sgn_minus_1: "sgn (- 1) = - 1" and sgn_mult: "sgn (a * b) = sgn a * sgn b" begin lemma sgn_eq_0_iff: "sgn a = 0 \ a = 0" proof - { assume "sgn a = 0" then have "sgn a * \a\ = 0" by simp then have "a = 0" by (simp add: sgn_mult_abs) } then show ?thesis by auto qed lemma abs_eq_0_iff: "\a\ = 0 \ a = 0" proof - { assume "\a\ = 0" then have "sgn a * \a\ = 0" by simp then have "a = 0" by (simp add: sgn_mult_abs) } then show ?thesis by auto qed lemma abs_mult_sgn: "\a\ * sgn a = a" using sgn_mult_abs [of a] by (simp add: ac_simps) lemma abs_1 [simp]: "\1\ = 1" using sgn_mult_abs [of 1] by simp lemma sgn_abs [simp]: "\sgn a\ = of_bool (a \ 0)" using sgn_mult_abs [of "sgn a"] mult_cancel_left [of "sgn a" "\sgn a\" 1] by (auto simp add: sgn_eq_0_iff) lemma abs_sgn [simp]: "sgn \a\ = of_bool (a \ 0)" using sgn_mult_abs [of "\a\"] mult_cancel_right [of "sgn \a\" "\a\" 1] by (auto simp add: abs_eq_0_iff) lemma abs_mult: "\a * b\ = \a\ * \b\" proof (cases "a = 0 \ b = 0") case True then show ?thesis by auto next case False then have *: "sgn (a * b) \ 0" by (simp add: sgn_eq_0_iff) from abs_mult_sgn [of "a * b"] abs_mult_sgn [of a] abs_mult_sgn [of b] have "\a * b\ * sgn (a * b) = \a\ * sgn a * \b\ * sgn b" by (simp add: ac_simps) then have "\a * b\ * sgn (a * b) = \a\ * \b\ * sgn (a * b)" by (simp add: sgn_mult ac_simps) with * show ?thesis by simp qed lemma sgn_minus [simp]: "sgn (- a) = - sgn a" proof - from sgn_minus_1 have "sgn (- 1 * a) = - 1 * sgn a" by (simp only: sgn_mult) then show ?thesis by simp qed lemma abs_minus [simp]: "\- a\ = \a\" proof - have [simp]: "\- 1\ = 1" using sgn_mult_abs [of "- 1"] by simp then have "\- 1 * a\ = 1 * \a\" by (simp only: abs_mult) then show ?thesis by simp qed end subsection \(Partial) Division\ class divide = fixes divide :: "'a \ 'a \ 'a" (infixl "div" 70) setup \Sign.add_const_constraint (\<^const_name>\divide\, SOME \<^typ>\'a \ 'a \ 'a\)\ context semiring begin lemma [field_simps, field_split_simps]: shows distrib_left_NO_MATCH: "NO_MATCH (x div y) a \ a * (b + c) = a * b + a * c" and distrib_right_NO_MATCH: "NO_MATCH (x div y) c \ (a + b) * c = a * c + b * c" by (rule distrib_left distrib_right)+ end context ring begin lemma [field_simps, field_split_simps]: shows left_diff_distrib_NO_MATCH: "NO_MATCH (x div y) c \ (a - b) * c = a * c - b * c" and right_diff_distrib_NO_MATCH: "NO_MATCH (x div y) a \ a * (b - c) = a * b - a * c" by (rule left_diff_distrib right_diff_distrib)+ end setup \Sign.add_const_constraint (\<^const_name>\divide\, SOME \<^typ>\'a::divide \ 'a \ 'a\)\ text \Algebraic classes with division\ class semidom_divide = semidom + divide + assumes nonzero_mult_div_cancel_right [simp]: "b \ 0 \ (a * b) div b = a" assumes div_by_0 [simp]: "a div 0 = 0" begin lemma nonzero_mult_div_cancel_left [simp]: "a \ 0 \ (a * b) div a = b" using nonzero_mult_div_cancel_right [of a b] by (simp add: ac_simps) subclass semiring_no_zero_divisors_cancel proof show *: "a * c = b * c \ c = 0 \ a = b" for a b c proof (cases "c = 0") case True then show ?thesis by simp next case False have "a = b" if "a * c = b * c" proof - from that have "a * c div c = b * c div c" by simp with False show ?thesis by simp qed then show ?thesis by auto qed show "c * a = c * b \ c = 0 \ a = b" for a b c using * [of a c b] by (simp add: ac_simps) qed lemma div_self [simp]: "a \ 0 \ a div a = 1" using nonzero_mult_div_cancel_left [of a 1] by simp lemma div_0 [simp]: "0 div a = 0" proof (cases "a = 0") case True then show ?thesis by simp next case False then have "a * 0 div a = 0" by (rule nonzero_mult_div_cancel_left) then show ?thesis by simp qed lemma div_by_1 [simp]: "a div 1 = a" using nonzero_mult_div_cancel_left [of 1 a] by simp lemma dvd_div_eq_0_iff: assumes "b dvd a" shows "a div b = 0 \ a = 0" using assms by (elim dvdE, cases "b = 0") simp_all lemma dvd_div_eq_cancel: "a div c = b div c \ c dvd a \ c dvd b \ a = b" by (elim dvdE, cases "c = 0") simp_all lemma dvd_div_eq_iff: "c dvd a \ c dvd b \ a div c = b div c \ a = b" by (elim dvdE, cases "c = 0") simp_all lemma inj_on_mult: "inj_on ((*) a) A" if "a \ 0" proof (rule inj_onI) fix b c assume "a * b = a * c" then have "a * b div a = a * c div a" by (simp only:) with that show "b = c" by simp qed end class idom_divide = idom + semidom_divide begin lemma dvd_neg_div: assumes "b dvd a" shows "- a div b = - (a div b)" proof (cases "b = 0") case True then show ?thesis by simp next case False from assms obtain c where "a = b * c" .. then have "- a div b = (b * - c) div b" by simp from False also have "\ = - c" by (rule nonzero_mult_div_cancel_left) with False \a = b * c\ show ?thesis by simp qed lemma dvd_div_neg: assumes "b dvd a" shows "a div - b = - (a div b)" proof (cases "b = 0") case True then show ?thesis by simp next case False then have "- b \ 0" by simp from assms obtain c where "a = b * c" .. then have "a div - b = (- b * - c) div - b" by simp from \- b \ 0\ also have "\ = - c" by (rule nonzero_mult_div_cancel_left) with False \a = b * c\ show ?thesis by simp qed end class algebraic_semidom = semidom_divide begin text \ Class \<^class>\algebraic_semidom\ enriches a integral domain by notions from algebra, like units in a ring. It is a separate class to avoid spoiling fields with notions which are degenerated there. \ lemma dvd_times_left_cancel_iff [simp]: assumes "a \ 0" shows "a * b dvd a * c \ b dvd c" (is "?lhs \ ?rhs") proof assume ?lhs then obtain d where "a * c = a * b * d" .. with assms have "c = b * d" by (simp add: ac_simps) then show ?rhs .. next assume ?rhs then obtain d where "c = b * d" .. then have "a * c = a * b * d" by (simp add: ac_simps) then show ?lhs .. qed lemma dvd_times_right_cancel_iff [simp]: assumes "a \ 0" shows "b * a dvd c * a \ b dvd c" using dvd_times_left_cancel_iff [of a b c] assms by (simp add: ac_simps) lemma div_dvd_iff_mult: assumes "b \ 0" and "b dvd a" shows "a div b dvd c \ a dvd c * b" proof - from \b dvd a\ obtain d where "a = b * d" .. with \b \ 0\ show ?thesis by (simp add: ac_simps) qed lemma dvd_div_iff_mult: assumes "c \ 0" and "c dvd b" shows "a dvd b div c \ a * c dvd b" proof - from \c dvd b\ obtain d where "b = c * d" .. with \c \ 0\ show ?thesis by (simp add: mult.commute [of a]) qed lemma div_dvd_div [simp]: assumes "a dvd b" and "a dvd c" shows "b div a dvd c div a \ b dvd c" proof (cases "a = 0") case True with assms show ?thesis by simp next case False moreover from assms obtain k l where "b = a * k" and "c = a * l" by blast ultimately show ?thesis by simp qed lemma div_add [simp]: assumes "c dvd a" and "c dvd b" shows "(a + b) div c = a div c + b div c" proof (cases "c = 0") case True then show ?thesis by simp next case False moreover from assms obtain k l where "a = c * k" and "b = c * l" by blast moreover have "c * k + c * l = c * (k + l)" by (simp add: algebra_simps) ultimately show ?thesis by simp qed lemma div_mult_div_if_dvd: assumes "b dvd a" and "d dvd c" shows "(a div b) * (c div d) = (a * c) div (b * d)" proof (cases "b = 0 \ c = 0") case True with assms show ?thesis by auto next case False moreover from assms obtain k l where "a = b * k" and "c = d * l" by blast moreover have "b * k * (d * l) div (b * d) = (b * d) * (k * l) div (b * d)" by (simp add: ac_simps) ultimately show ?thesis by simp qed lemma dvd_div_eq_mult: assumes "a \ 0" and "a dvd b" shows "b div a = c \ b = c * a" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (simp add: assms) next assume ?lhs then have "b div a * a = c * a" by simp moreover from assms have "b div a * a = b" by (auto simp add: ac_simps) ultimately show ?rhs by simp qed lemma dvd_div_mult_self [simp]: "a dvd b \ b div a * a = b" by (cases "a = 0") (auto simp add: ac_simps) lemma dvd_mult_div_cancel [simp]: "a dvd b \ a * (b div a) = b" using dvd_div_mult_self [of a b] by (simp add: ac_simps) lemma div_mult_swap: assumes "c dvd b" shows "a * (b div c) = (a * b) div c" proof (cases "c = 0") case True then show ?thesis by simp next case False from assms obtain d where "b = c * d" .. moreover from False have "a * divide (d * c) c = ((a * d) * c) div c" by simp ultimately show ?thesis by (simp add: ac_simps) qed lemma dvd_div_mult: "c dvd b \ b div c * a = (b * a) div c" using div_mult_swap [of c b a] by (simp add: ac_simps) lemma dvd_div_mult2_eq: assumes "b * c dvd a" shows "a div (b * c) = a div b div c" proof - from assms obtain k where "a = b * c * k" .. then show ?thesis by (cases "b = 0 \ c = 0") (auto, simp add: ac_simps) qed lemma dvd_div_div_eq_mult: assumes "a \ 0" "c \ 0" and "a dvd b" "c dvd d" shows "b div a = d div c \ b * c = a * d" (is "?lhs \ ?rhs") proof - from assms have "a * c \ 0" by simp then have "?lhs \ b div a * (a * c) = d div c * (a * c)" by simp also have "\ \ (a * (b div a)) * c = (c * (d div c)) * a" by (simp add: ac_simps) also have "\ \ (a * b div a) * c = (c * d div c) * a" using assms by (simp add: div_mult_swap) also have "\ \ ?rhs" using assms by (simp add: ac_simps) finally show ?thesis . qed lemma dvd_mult_imp_div: assumes "a * c dvd b" shows "a dvd b div c" proof (cases "c = 0") case True then show ?thesis by simp next case False from \a * c dvd b\ obtain d where "b = a * c * d" .. with False show ?thesis by (simp add: mult.commute [of a] mult.assoc) qed lemma div_div_eq_right: assumes "c dvd b" "b dvd a" shows "a div (b div c) = a div b * c" proof (cases "c = 0 \ b = 0") case True then show ?thesis by auto next case False from assms obtain r s where "b = c * r" and "a = c * r * s" by blast moreover with False have "r \ 0" by auto ultimately show ?thesis using False by simp (simp add: mult.commute [of _ r] mult.assoc mult.commute [of c]) qed lemma div_div_div_same: assumes "d dvd b" "b dvd a" shows "(a div d) div (b div d) = a div b" proof (cases "b = 0 \ d = 0") case True with assms show ?thesis by auto next case False from assms obtain r s where "a = d * r * s" and "b = d * r" by blast with False show ?thesis by simp (simp add: ac_simps) qed text \Units: invertible elements in a ring\ abbreviation is_unit :: "'a \ bool" where "is_unit a \ a dvd 1" lemma not_is_unit_0 [simp]: "\ is_unit 0" by simp lemma unit_imp_dvd [dest]: "is_unit b \ b dvd a" by (rule dvd_trans [of _ 1]) simp_all lemma unit_dvdE: assumes "is_unit a" obtains c where "a \ 0" and "b = a * c" proof - from assms have "a dvd b" by auto then obtain c where "b = a * c" .. moreover from assms have "a \ 0" by auto ultimately show thesis using that by blast qed lemma dvd_unit_imp_unit: "a dvd b \ is_unit b \ is_unit a" by (rule dvd_trans) lemma unit_div_1_unit [simp, intro]: assumes "is_unit a" shows "is_unit (1 div a)" proof - from assms have "1 = 1 div a * a" by simp then show "is_unit (1 div a)" by (rule dvdI) qed lemma is_unitE [elim?]: assumes "is_unit a" obtains b where "a \ 0" and "b \ 0" and "is_unit b" and "1 div a = b" and "1 div b = a" and "a * b = 1" and "c div a = c * b" proof (rule that) define b where "b = 1 div a" then show "1 div a = b" by simp from assms b_def show "is_unit b" by simp with assms show "a \ 0" and "b \ 0" by auto from assms b_def show "a * b = 1" by simp then have "1 = a * b" .. with b_def \b \ 0\ show "1 div b = a" by simp from assms have "a dvd c" .. then obtain d where "c = a * d" .. with \a \ 0\ \a * b = 1\ show "c div a = c * b" by (simp add: mult.assoc mult.left_commute [of a]) qed lemma unit_prod [intro]: "is_unit a \ is_unit b \ is_unit (a * b)" by (subst mult_1_left [of 1, symmetric]) (rule mult_dvd_mono) lemma is_unit_mult_iff: "is_unit (a * b) \ is_unit a \ is_unit b" by (auto dest: dvd_mult_left dvd_mult_right) lemma unit_div [intro]: "is_unit a \ is_unit b \ is_unit (a div b)" by (erule is_unitE [of b a]) (simp add: ac_simps unit_prod) lemma mult_unit_dvd_iff: assumes "is_unit b" shows "a * b dvd c \ a dvd c" proof assume "a * b dvd c" with assms show "a dvd c" by (simp add: dvd_mult_left) next assume "a dvd c" then obtain k where "c = a * k" .. with assms have "c = (a * b) * (1 div b * k)" by (simp add: mult_ac) then show "a * b dvd c" by (rule dvdI) qed lemma mult_unit_dvd_iff': "is_unit a \ (a * b) dvd c \ b dvd c" using mult_unit_dvd_iff [of a b c] by (simp add: ac_simps) lemma dvd_mult_unit_iff: assumes "is_unit b" shows "a dvd c * b \ a dvd c" proof assume "a dvd c * b" with assms have "c * b dvd c * (b * (1 div b))" by (subst mult_assoc [symmetric]) simp also from assms have "b * (1 div b) = 1" by (rule is_unitE) simp finally have "c * b dvd c" by simp with \a dvd c * b\ show "a dvd c" by (rule dvd_trans) next assume "a dvd c" then show "a dvd c * b" by simp qed lemma dvd_mult_unit_iff': "is_unit b \ a dvd b * c \ a dvd c" using dvd_mult_unit_iff [of b a c] by (simp add: ac_simps) lemma div_unit_dvd_iff: "is_unit b \ a div b dvd c \ a dvd c" by (erule is_unitE [of _ a]) (auto simp add: mult_unit_dvd_iff) lemma dvd_div_unit_iff: "is_unit b \ a dvd c div b \ a dvd c" by (erule is_unitE [of _ c]) (simp add: dvd_mult_unit_iff) lemmas unit_dvd_iff = mult_unit_dvd_iff mult_unit_dvd_iff' dvd_mult_unit_iff dvd_mult_unit_iff' div_unit_dvd_iff dvd_div_unit_iff (* FIXME consider named_theorems *) lemma unit_mult_div_div [simp]: "is_unit a \ b * (1 div a) = b div a" by (erule is_unitE [of _ b]) simp lemma unit_div_mult_self [simp]: "is_unit a \ b div a * a = b" by (rule dvd_div_mult_self) auto lemma unit_div_1_div_1 [simp]: "is_unit a \ 1 div (1 div a) = a" by (erule is_unitE) simp lemma unit_div_mult_swap: "is_unit c \ a * (b div c) = (a * b) div c" by (erule unit_dvdE [of _ b]) (simp add: mult.left_commute [of _ c]) lemma unit_div_commute: "is_unit b \ (a div b) * c = (a * c) div b" using unit_div_mult_swap [of b c a] by (simp add: ac_simps) lemma unit_eq_div1: "is_unit b \ a div b = c \ a = c * b" by (auto elim: is_unitE) lemma unit_eq_div2: "is_unit b \ a = c div b \ a * b = c" using unit_eq_div1 [of b c a] by auto lemma unit_mult_left_cancel: "is_unit a \ a * b = a * c \ b = c" using mult_cancel_left [of a b c] by auto lemma unit_mult_right_cancel: "is_unit a \ b * a = c * a \ b = c" using unit_mult_left_cancel [of a b c] by (auto simp add: ac_simps) lemma unit_div_cancel: assumes "is_unit a" shows "b div a = c div a \ b = c" proof - from assms have "is_unit (1 div a)" by simp then have "b * (1 div a) = c * (1 div a) \ b = c" by (rule unit_mult_right_cancel) with assms show ?thesis by simp qed lemma is_unit_div_mult2_eq: assumes "is_unit b" and "is_unit c" shows "a div (b * c) = a div b div c" proof - from assms have "is_unit (b * c)" by (simp add: unit_prod) then have "b * c dvd a" by (rule unit_imp_dvd) then show ?thesis by (rule dvd_div_mult2_eq) qed lemma is_unit_div_mult_cancel_left: assumes "a \ 0" and "is_unit b" shows "a div (a * b) = 1 div b" proof - from assms have "a div (a * b) = a div a div b" by (simp add: mult_unit_dvd_iff dvd_div_mult2_eq) with assms show ?thesis by simp qed lemma is_unit_div_mult_cancel_right: assumes "a \ 0" and "is_unit b" shows "a div (b * a) = 1 div b" using assms is_unit_div_mult_cancel_left [of a b] by (simp add: ac_simps) lemma unit_div_eq_0_iff: assumes "is_unit b" shows "a div b = 0 \ a = 0" by (rule dvd_div_eq_0_iff) (insert assms, auto) lemma div_mult_unit2: "is_unit c \ b dvd a \ a div (b * c) = a div b div c" by (rule dvd_div_mult2_eq) (simp_all add: mult_unit_dvd_iff) text \Coprimality\ definition coprime :: "'a \ 'a \ bool" where "coprime a b \ (\c. c dvd a \ c dvd b \ is_unit c)" lemma coprimeI: assumes "\c. c dvd a \ c dvd b \ is_unit c" shows "coprime a b" using assms by (auto simp: coprime_def) lemma not_coprimeI: assumes "c dvd a" and "c dvd b" and "\ is_unit c" shows "\ coprime a b" using assms by (auto simp: coprime_def) lemma coprime_common_divisor: "is_unit c" if "coprime a b" and "c dvd a" and "c dvd b" using that by (auto simp: coprime_def) lemma not_coprimeE: assumes "\ coprime a b" obtains c where "c dvd a" and "c dvd b" and "\ is_unit c" using assms by (auto simp: coprime_def) lemma coprime_imp_coprime: "coprime a b" if "coprime c d" and "\e. \ is_unit e \ e dvd a \ e dvd b \ e dvd c" and "\e. \ is_unit e \ e dvd a \ e dvd b \ e dvd d" proof (rule coprimeI) fix e assume "e dvd a" and "e dvd b" with that have "e dvd c" and "e dvd d" by (auto intro: dvd_trans) with \coprime c d\ show "is_unit e" by (rule coprime_common_divisor) qed lemma coprime_divisors: "coprime a b" if "a dvd c" "b dvd d" and "coprime c d" using \coprime c d\ proof (rule coprime_imp_coprime) fix e assume "e dvd a" then show "e dvd c" using \a dvd c\ by (rule dvd_trans) assume "e dvd b" then show "e dvd d" using \b dvd d\ by (rule dvd_trans) qed lemma coprime_self [simp]: "coprime a a \ is_unit a" (is "?P \ ?Q") proof assume ?P then show ?Q by (rule coprime_common_divisor) simp_all next assume ?Q show ?P by (rule coprimeI) (erule dvd_unit_imp_unit, rule \?Q\) qed lemma coprime_commute [ac_simps]: "coprime b a \ coprime a b" unfolding coprime_def by auto lemma is_unit_left_imp_coprime: "coprime a b" if "is_unit a" proof (rule coprimeI) fix c assume "c dvd a" with that show "is_unit c" by (auto intro: dvd_unit_imp_unit) qed lemma is_unit_right_imp_coprime: "coprime a b" if "is_unit b" using that is_unit_left_imp_coprime [of b a] by (simp add: ac_simps) lemma coprime_1_left [simp]: "coprime 1 a" by (rule coprimeI) lemma coprime_1_right [simp]: "coprime a 1" by (rule coprimeI) lemma coprime_0_left_iff [simp]: "coprime 0 a \ is_unit a" by (auto intro: coprimeI dvd_unit_imp_unit coprime_common_divisor [of 0 a a]) lemma coprime_0_right_iff [simp]: "coprime a 0 \ is_unit a" using coprime_0_left_iff [of a] by (simp add: ac_simps) lemma coprime_mult_self_left_iff [simp]: "coprime (c * a) (c * b) \ is_unit c \ coprime a b" by (auto intro: coprime_common_divisor) (rule coprimeI, auto intro: coprime_common_divisor simp add: dvd_mult_unit_iff')+ lemma coprime_mult_self_right_iff [simp]: "coprime (a * c) (b * c) \ is_unit c \ coprime a b" using coprime_mult_self_left_iff [of c a b] by (simp add: ac_simps) lemma coprime_absorb_left: assumes "x dvd y" shows "coprime x y \ is_unit x" using assms coprime_common_divisor is_unit_left_imp_coprime by auto lemma coprime_absorb_right: assumes "y dvd x" shows "coprime x y \ is_unit y" using assms coprime_common_divisor is_unit_right_imp_coprime by auto end class unit_factor = fixes unit_factor :: "'a \ 'a" class semidom_divide_unit_factor = semidom_divide + unit_factor + assumes unit_factor_0 [simp]: "unit_factor 0 = 0" and is_unit_unit_factor: "a dvd 1 \ unit_factor a = a" and unit_factor_is_unit: "a \ 0 \ unit_factor a dvd 1" - and unit_factor_mult: "unit_factor (a * b) = unit_factor a * unit_factor b" + and unit_factor_mult_unit_left: "a dvd 1 \ unit_factor (a * b) = a * unit_factor b" \ \This fine-grained hierarchy will later on allow lean normalization of polynomials\ - +begin + +lemma unit_factor_mult_unit_right: "a dvd 1 \ unit_factor (b * a) = unit_factor b * a" + using unit_factor_mult_unit_left[of a b] by (simp add: mult_ac) + +lemmas [simp] = unit_factor_mult_unit_left unit_factor_mult_unit_right + +end + class normalization_semidom = algebraic_semidom + semidom_divide_unit_factor + fixes normalize :: "'a \ 'a" assumes unit_factor_mult_normalize [simp]: "unit_factor a * normalize a = a" and normalize_0 [simp]: "normalize 0 = 0" begin text \ Class \<^class>\normalization_semidom\ cultivates the idea that each integral domain can be split into equivalence classes whose representants are associated, i.e. divide each other. \<^const>\normalize\ specifies a canonical representant for each equivalence class. The rationale behind this is that it is easier to reason about equality than equivalences, hence we prefer to think about equality of normalized values rather than associated elements. \ declare unit_factor_is_unit [iff] lemma unit_factor_dvd [simp]: "a \ 0 \ unit_factor a dvd b" by (rule unit_imp_dvd) simp lemma unit_factor_self [simp]: "unit_factor a dvd a" by (cases "a = 0") simp_all lemma normalize_mult_unit_factor [simp]: "normalize a * unit_factor a = a" using unit_factor_mult_normalize [of a] by (simp add: ac_simps) lemma normalize_eq_0_iff [simp]: "normalize a = 0 \ a = 0" (is "?lhs \ ?rhs") proof assume ?lhs moreover have "unit_factor a * normalize a = a" by simp ultimately show ?rhs by simp next assume ?rhs then show ?lhs by simp qed lemma unit_factor_eq_0_iff [simp]: "unit_factor a = 0 \ a = 0" (is "?lhs \ ?rhs") proof assume ?lhs moreover have "unit_factor a * normalize a = a" by simp ultimately show ?rhs by simp next assume ?rhs then show ?lhs by simp qed lemma div_unit_factor [simp]: "a div unit_factor a = normalize a" proof (cases "a = 0") case True then show ?thesis by simp next case False then have "unit_factor a \ 0" by simp with nonzero_mult_div_cancel_left have "unit_factor a * normalize a div unit_factor a = normalize a" by blast then show ?thesis by simp qed lemma normalize_div [simp]: "normalize a div a = 1 div unit_factor a" proof (cases "a = 0") case True then show ?thesis by simp next case False have "normalize a div a = normalize a div (unit_factor a * normalize a)" by simp also have "\ = 1 div unit_factor a" using False by (subst is_unit_div_mult_cancel_right) simp_all finally show ?thesis . qed lemma is_unit_normalize: assumes "is_unit a" shows "normalize a = 1" proof - from assms have "unit_factor a = a" by (rule is_unit_unit_factor) moreover from assms have "a \ 0" by auto moreover have "normalize a = a div unit_factor a" by simp ultimately show ?thesis by simp qed lemma unit_factor_1 [simp]: "unit_factor 1 = 1" by (rule is_unit_unit_factor) simp lemma normalize_1 [simp]: "normalize 1 = 1" by (rule is_unit_normalize) simp lemma normalize_1_iff: "normalize a = 1 \ is_unit a" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (rule is_unit_normalize) next assume ?lhs then have "unit_factor a * normalize a = unit_factor a * 1" by simp then have "unit_factor a = a" by simp moreover from \?lhs\ have "a \ 0" by auto then have "is_unit (unit_factor a)" by simp ultimately show ?rhs by simp qed lemma div_normalize [simp]: "a div normalize a = unit_factor a" proof (cases "a = 0") case True then show ?thesis by simp next case False then have "normalize a \ 0" by simp with nonzero_mult_div_cancel_right have "unit_factor a * normalize a div normalize a = unit_factor a" by blast then show ?thesis by simp qed lemma mult_one_div_unit_factor [simp]: "a * (1 div unit_factor b) = a div unit_factor b" by (cases "b = 0") simp_all lemma inv_unit_factor_eq_0_iff [simp]: "1 div unit_factor a = 0 \ a = 0" (is "?lhs \ ?rhs") proof assume ?lhs then have "a * (1 div unit_factor a) = a * 0" by simp then show ?rhs by simp next assume ?rhs then show ?lhs by simp qed -lemma normalize_mult: "normalize (a * b) = normalize a * normalize b" -proof (cases "a = 0 \ b = 0") - case True - then show ?thesis by auto -next - case False - have "unit_factor (a * b) * normalize (a * b) = a * b" - by (rule unit_factor_mult_normalize) - then have "normalize (a * b) = a * b div unit_factor (a * b)" - by simp - also have "\ = a * b div unit_factor (b * a)" - by (simp add: ac_simps) - also have "\ = a * b div unit_factor b div unit_factor a" - using False by (simp add: unit_factor_mult is_unit_div_mult2_eq [symmetric]) - also have "\ = a * (b div unit_factor b) div unit_factor a" - using False by (subst unit_div_mult_swap) simp_all - also have "\ = normalize a * normalize b" - using False - by (simp add: mult.commute [of a] mult.commute [of "normalize a"] unit_div_mult_swap [symmetric]) - finally show ?thesis . -qed - lemma unit_factor_idem [simp]: "unit_factor (unit_factor a) = unit_factor a" by (cases "a = 0") (auto intro: is_unit_unit_factor) lemma normalize_unit_factor [simp]: "a \ 0 \ normalize (unit_factor a) = 1" by (rule is_unit_normalize) simp +lemma normalize_mult_unit_left [simp]: + assumes "a dvd 1" + shows "normalize (a * b) = normalize b" +proof (cases "b = 0") + case False + have "a * unit_factor b * normalize (a * b) = unit_factor (a * b) * normalize (a * b)" + using assms by (subst unit_factor_mult_unit_left) auto + also have "\ = a * b" by simp + also have "b = unit_factor b * normalize b" by simp + hence "a * b = a * unit_factor b * normalize b" + by (simp only: mult_ac) + finally show ?thesis + using assms False by auto +qed auto + +lemma normalize_mult_unit_right [simp]: + assumes "b dvd 1" + shows "normalize (a * b) = normalize a" + using assms by (subst mult.commute) auto + lemma normalize_idem [simp]: "normalize (normalize a) = normalize a" proof (cases "a = 0") - case True - then show ?thesis by simp -next case False have "normalize a = normalize (unit_factor a * normalize a)" by simp - also have "\ = normalize (unit_factor a) * normalize (normalize a)" - by (simp only: normalize_mult) - finally show ?thesis - using False by simp_all -qed + also from False have "\ = normalize (normalize a)" + by (subst normalize_mult_unit_left) auto + finally show ?thesis .. +qed auto lemma unit_factor_normalize [simp]: assumes "a \ 0" shows "unit_factor (normalize a) = 1" proof - from assms have *: "normalize a \ 0" by simp have "unit_factor (normalize a) * normalize (normalize a) = normalize a" by (simp only: unit_factor_mult_normalize) then have "unit_factor (normalize a) * normalize a = normalize a" by simp with * have "unit_factor (normalize a) * normalize a div normalize a = normalize a div normalize a" by simp with * show ?thesis by simp qed -lemma dvd_unit_factor_div: - assumes "b dvd a" - shows "unit_factor (a div b) = unit_factor a div unit_factor b" -proof - - from assms have "a = a div b * b" - by simp - then have "unit_factor a = unit_factor (a div b * b)" - by simp - then show ?thesis - by (cases "b = 0") (simp_all add: unit_factor_mult) -qed - -lemma dvd_normalize_div: - assumes "b dvd a" - shows "normalize (a div b) = normalize a div normalize b" -proof - - from assms have "a = a div b * b" - by simp - then have "normalize a = normalize (a div b * b)" - by simp - then show ?thesis - by (cases "b = 0") (simp_all add: normalize_mult) -qed - lemma normalize_dvd_iff [simp]: "normalize a dvd b \ a dvd b" proof - have "normalize a dvd b \ unit_factor a * normalize a dvd b" using mult_unit_dvd_iff [of "unit_factor a" "normalize a" b] by (cases "a = 0") simp_all then show ?thesis by simp qed lemma dvd_normalize_iff [simp]: "a dvd normalize b \ a dvd b" proof - have "a dvd normalize b \ a dvd normalize b * unit_factor b" using dvd_mult_unit_iff [of "unit_factor b" a "normalize b"] by (cases "b = 0") simp_all then show ?thesis by simp qed lemma normalize_idem_imp_unit_factor_eq: assumes "normalize a = a" shows "unit_factor a = of_bool (a \ 0)" proof (cases "a = 0") case True then show ?thesis by simp next case False then show ?thesis using assms unit_factor_normalize [of a] by simp qed lemma normalize_idem_imp_is_unit_iff: assumes "normalize a = a" shows "is_unit a \ a = 1" using assms by (cases "a = 0") (auto dest: is_unit_normalize) lemma coprime_normalize_left_iff [simp]: "coprime (normalize a) b \ coprime a b" by (rule; rule coprimeI) (auto intro: coprime_common_divisor) lemma coprime_normalize_right_iff [simp]: "coprime a (normalize b) \ coprime a b" using coprime_normalize_left_iff [of b a] by (simp add: ac_simps) text \ We avoid an explicit definition of associated elements but prefer explicit normalisation instead. In theory we could define an abbreviation like \<^prop>\associated a b \ normalize a = normalize b\ but this is counterproductive without suggestive infix syntax, which we do not want to sacrifice for this purpose here. \ lemma associatedI: assumes "a dvd b" and "b dvd a" shows "normalize a = normalize b" proof (cases "a = 0 \ b = 0") case True with assms show ?thesis by auto next case False from \a dvd b\ obtain c where b: "b = a * c" .. moreover from \b dvd a\ obtain d where a: "a = b * d" .. ultimately have "b * 1 = b * (c * d)" by (simp add: ac_simps) with False have "1 = c * d" unfolding mult_cancel_left by simp then have "is_unit c" and "is_unit d" by auto with a b show ?thesis - by (simp add: normalize_mult is_unit_normalize) + by (simp add: is_unit_normalize) qed lemma associatedD1: "normalize a = normalize b \ a dvd b" using dvd_normalize_iff [of _ b, symmetric] normalize_dvd_iff [of a _, symmetric] by simp lemma associatedD2: "normalize a = normalize b \ b dvd a" using dvd_normalize_iff [of _ a, symmetric] normalize_dvd_iff [of b _, symmetric] by simp lemma associated_unit: "normalize a = normalize b \ is_unit a \ is_unit b" using dvd_unit_imp_unit by (auto dest!: associatedD1 associatedD2) lemma associated_iff_dvd: "normalize a = normalize b \ a dvd b \ b dvd a" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (auto intro!: associatedI) next assume ?lhs then have "unit_factor a * normalize a = unit_factor a * normalize b" by simp then have *: "normalize b * unit_factor a = a" by (simp add: ac_simps) show ?rhs proof (cases "a = 0 \ b = 0") case True with \?lhs\ show ?thesis by auto next case False then have "b dvd normalize b * unit_factor a" and "normalize b * unit_factor a dvd b" by (simp_all add: mult_unit_dvd_iff dvd_mult_unit_iff) with * show ?thesis by simp qed qed lemma associated_eqI: assumes "a dvd b" and "b dvd a" assumes "normalize a = a" and "normalize b = b" shows "a = b" proof - from assms have "normalize a = normalize b" unfolding associated_iff_dvd by simp with \normalize a = a\ have "a = normalize b" by simp with \normalize b = b\ show "a = b" by simp qed lemma normalize_unit_factor_eqI: assumes "normalize a = normalize b" and "unit_factor a = unit_factor b" shows "a = b" proof - from assms have "unit_factor a * normalize a = unit_factor b * normalize b" by simp then show ?thesis by simp qed +lemma normalize_mult_normalize_left [simp]: "normalize (normalize a * b) = normalize (a * b)" + by (rule associated_eqI) (auto intro!: mult_dvd_mono) + +lemma normalize_mult_normalize_right [simp]: "normalize (a * normalize b) = normalize (a * b)" + by (rule associated_eqI) (auto intro!: mult_dvd_mono) + end +class normalization_semidom_multiplicative = normalization_semidom + + assumes unit_factor_mult: "unit_factor (a * b) = unit_factor a * unit_factor b" +begin + +lemma normalize_mult: "normalize (a * b) = normalize a * normalize b" +proof (cases "a = 0 \ b = 0") + case True + then show ?thesis by auto +next + case False + have "unit_factor (a * b) * normalize (a * b) = a * b" + by (rule unit_factor_mult_normalize) + then have "normalize (a * b) = a * b div unit_factor (a * b)" + by simp + also have "\ = a * b div unit_factor (b * a)" + by (simp add: ac_simps) + also have "\ = a * b div unit_factor b div unit_factor a" + using False by (simp add: unit_factor_mult is_unit_div_mult2_eq [symmetric]) + also have "\ = a * (b div unit_factor b) div unit_factor a" + using False by (subst unit_div_mult_swap) simp_all + also have "\ = normalize a * normalize b" + using False + by (simp add: mult.commute [of a] mult.commute [of "normalize a"] unit_div_mult_swap [symmetric]) + finally show ?thesis . +qed + +lemma dvd_unit_factor_div: + assumes "b dvd a" + shows "unit_factor (a div b) = unit_factor a div unit_factor b" +proof - + from assms have "a = a div b * b" + by simp + then have "unit_factor a = unit_factor (a div b * b)" + by simp + then show ?thesis + by (cases "b = 0") (simp_all add: unit_factor_mult) +qed + +lemma dvd_normalize_div: + assumes "b dvd a" + shows "normalize (a div b) = normalize a div normalize b" +proof - + from assms have "a = a div b * b" + by simp + then have "normalize a = normalize (a div b * b)" + by simp + then show ?thesis + by (cases "b = 0") (simp_all add: normalize_mult) +qed + +end + + + + text \Syntactic division remainder operator\ class modulo = dvd + divide + fixes modulo :: "'a \ 'a \ 'a" (infixl "mod" 70) text \Arbitrary quotient and remainder partitions\ class semiring_modulo = comm_semiring_1_cancel + divide + modulo + assumes div_mult_mod_eq: "a div b * b + a mod b = a" begin lemma mod_div_decomp: fixes a b obtains q r where "q = a div b" and "r = a mod b" and "a = q * b + r" proof - from div_mult_mod_eq have "a = a div b * b + a mod b" by simp moreover have "a div b = a div b" .. moreover have "a mod b = a mod b" .. note that ultimately show thesis by blast qed lemma mult_div_mod_eq: "b * (a div b) + a mod b = a" using div_mult_mod_eq [of a b] by (simp add: ac_simps) lemma mod_div_mult_eq: "a mod b + a div b * b = a" using div_mult_mod_eq [of a b] by (simp add: ac_simps) lemma mod_mult_div_eq: "a mod b + b * (a div b) = a" using div_mult_mod_eq [of a b] by (simp add: ac_simps) lemma minus_div_mult_eq_mod: "a - a div b * b = a mod b" by (rule add_implies_diff [symmetric]) (fact mod_div_mult_eq) lemma minus_mult_div_eq_mod: "a - b * (a div b) = a mod b" by (rule add_implies_diff [symmetric]) (fact mod_mult_div_eq) lemma minus_mod_eq_div_mult: "a - a mod b = a div b * b" by (rule add_implies_diff [symmetric]) (fact div_mult_mod_eq) lemma minus_mod_eq_mult_div: "a - a mod b = b * (a div b)" by (rule add_implies_diff [symmetric]) (fact mult_div_mod_eq) lemma mod_0_imp_dvd [dest!]: "b dvd a" if "a mod b = 0" proof - have "b dvd (a div b) * b" by simp also have "(a div b) * b = a" using div_mult_mod_eq [of a b] by (simp add: that) finally show ?thesis . qed lemma [nitpick_unfold]: "a mod b = a - a div b * b" by (fact minus_div_mult_eq_mod [symmetric]) end subsection \Quotient and remainder in integral domains\ class semidom_modulo = algebraic_semidom + semiring_modulo begin lemma mod_0 [simp]: "0 mod a = 0" using div_mult_mod_eq [of 0 a] by simp lemma mod_by_0 [simp]: "a mod 0 = a" using div_mult_mod_eq [of a 0] by simp lemma mod_by_1 [simp]: "a mod 1 = 0" proof - from div_mult_mod_eq [of a one] div_by_1 have "a + a mod 1 = a" by simp then have "a + a mod 1 = a + 0" by simp then show ?thesis by (rule add_left_imp_eq) qed lemma mod_self [simp]: "a mod a = 0" using div_mult_mod_eq [of a a] by simp lemma dvd_imp_mod_0 [simp]: "b mod a = 0" if "a dvd b" using that minus_div_mult_eq_mod [of b a] by simp lemma mod_eq_0_iff_dvd: "a mod b = 0 \ b dvd a" by (auto intro: mod_0_imp_dvd) lemma dvd_eq_mod_eq_0 [nitpick_unfold, code]: "a dvd b \ b mod a = 0" by (simp add: mod_eq_0_iff_dvd) lemma dvd_mod_iff: assumes "c dvd b" shows "c dvd a mod b \ c dvd a" proof - from assms have "(c dvd a mod b) \ (c dvd ((a div b) * b + a mod b))" by (simp add: dvd_add_right_iff) also have "(a div b) * b + a mod b = a" using div_mult_mod_eq [of a b] by simp finally show ?thesis . qed lemma dvd_mod_imp_dvd: assumes "c dvd a mod b" and "c dvd b" shows "c dvd a" using assms dvd_mod_iff [of c b a] by simp lemma dvd_minus_mod [simp]: "b dvd a - a mod b" by (simp add: minus_mod_eq_div_mult) lemma cancel_div_mod_rules: "((a div b) * b + a mod b) + c = a + c" "(b * (a div b) + a mod b) + c = a + c" by (simp_all add: div_mult_mod_eq mult_div_mod_eq) end class idom_modulo = idom + semidom_modulo begin subclass idom_divide .. lemma div_diff [simp]: "c dvd a \ c dvd b \ (a - b) div c = a div c - b div c" using div_add [of _ _ "- b"] by (simp add: dvd_neg_div) end subsection \Interlude: basic tool support for algebraic and arithmetic calculations\ named_theorems arith "arith facts -- only ground formulas" ML_file \Tools/arith_data.ML\ ML_file \~~/src/Provers/Arith/cancel_div_mod.ML\ ML \ structure Cancel_Div_Mod_Ring = Cancel_Div_Mod ( val div_name = \<^const_name>\divide\; val mod_name = \<^const_name>\modulo\; val mk_binop = HOLogic.mk_binop; val mk_sum = Arith_Data.mk_sum; val dest_sum = Arith_Data.dest_sum; val div_mod_eqs = map mk_meta_eq @{thms cancel_div_mod_rules}; val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac @{thms diff_conv_add_uminus add_0_left add_0_right ac_simps}) ) \ simproc_setup cancel_div_mod_int ("(a::'a::semidom_modulo) + b") = \K Cancel_Div_Mod_Ring.proc\ subsection \Ordered semirings and rings\ text \ The theory of partially ordered rings is taken from the books: \<^item> \<^emph>\Lattice Theory\ by Garret Birkhoff, American Mathematical Society, 1979 \<^item> \<^emph>\Partially Ordered Algebraic Systems\, Pergamon Press, 1963 Most of the used notions can also be looked up in \<^item> \<^url>\http://www.mathworld.com\ by Eric Weisstein et. al. \<^item> \<^emph>\Algebra I\ by van der Waerden, Springer \ class ordered_semiring = semiring + ordered_comm_monoid_add + assumes mult_left_mono: "a \ b \ 0 \ c \ c * a \ c * b" assumes mult_right_mono: "a \ b \ 0 \ c \ a * c \ b * c" begin lemma mult_mono: "a \ b \ c \ d \ 0 \ b \ 0 \ c \ a * c \ b * d" apply (erule (1) mult_right_mono [THEN order_trans]) apply (erule (1) mult_left_mono) done lemma mult_mono': "a \ b \ c \ d \ 0 \ a \ 0 \ c \ a * c \ b * d" by (rule mult_mono) (fast intro: order_trans)+ end class ordered_semiring_0 = semiring_0 + ordered_semiring begin lemma mult_nonneg_nonneg [simp]: "0 \ a \ 0 \ b \ 0 \ a * b" using mult_left_mono [of 0 b a] by simp lemma mult_nonneg_nonpos: "0 \ a \ b \ 0 \ a * b \ 0" using mult_left_mono [of b 0 a] by simp lemma mult_nonpos_nonneg: "a \ 0 \ 0 \ b \ a * b \ 0" using mult_right_mono [of a 0 b] by simp text \Legacy -- use @{thm [source] mult_nonpos_nonneg}.\ lemma mult_nonneg_nonpos2: "0 \ a \ b \ 0 \ b * a \ 0" by (drule mult_right_mono [of b 0]) auto lemma split_mult_neg_le: "(0 \ a \ b \ 0) \ (a \ 0 \ 0 \ b) \ a * b \ 0" by (auto simp add: mult_nonneg_nonpos mult_nonneg_nonpos2) end class ordered_cancel_semiring = ordered_semiring + cancel_comm_monoid_add begin subclass semiring_0_cancel .. subclass ordered_semiring_0 .. end class linordered_semiring = ordered_semiring + linordered_cancel_ab_semigroup_add begin subclass ordered_cancel_semiring .. subclass ordered_cancel_comm_monoid_add .. subclass ordered_ab_semigroup_monoid_add_imp_le .. lemma mult_left_less_imp_less: "c * a < c * b \ 0 \ c \ a < b" by (force simp add: mult_left_mono not_le [symmetric]) lemma mult_right_less_imp_less: "a * c < b * c \ 0 \ c \ a < b" by (force simp add: mult_right_mono not_le [symmetric]) end class zero_less_one = order + zero + one + assumes zero_less_one [simp]: "0 < 1" class linordered_semiring_1 = linordered_semiring + semiring_1 + zero_less_one begin lemma convex_bound_le: assumes "x \ a" "y \ a" "0 \ u" "0 \ v" "u + v = 1" shows "u * x + v * y \ a" proof- from assms have "u * x + v * y \ u * a + v * a" by (simp add: add_mono mult_left_mono) with assms show ?thesis unfolding distrib_right[symmetric] by simp qed end class linordered_semiring_strict = semiring + comm_monoid_add + linordered_cancel_ab_semigroup_add + assumes mult_strict_left_mono: "a < b \ 0 < c \ c * a < c * b" assumes mult_strict_right_mono: "a < b \ 0 < c \ a * c < b * c" begin subclass semiring_0_cancel .. subclass linordered_semiring proof fix a b c :: 'a assume *: "a \ b" "0 \ c" then show "c * a \ c * b" unfolding le_less using mult_strict_left_mono by (cases "c = 0") auto from * show "a * c \ b * c" unfolding le_less using mult_strict_right_mono by (cases "c = 0") auto qed lemma mult_left_le_imp_le: "c * a \ c * b \ 0 < c \ a \ b" by (auto simp add: mult_strict_left_mono _not_less [symmetric]) lemma mult_right_le_imp_le: "a * c \ b * c \ 0 < c \ a \ b" by (auto simp add: mult_strict_right_mono not_less [symmetric]) lemma mult_pos_pos[simp]: "0 < a \ 0 < b \ 0 < a * b" using mult_strict_left_mono [of 0 b a] by simp lemma mult_pos_neg: "0 < a \ b < 0 \ a * b < 0" using mult_strict_left_mono [of b 0 a] by simp lemma mult_neg_pos: "a < 0 \ 0 < b \ a * b < 0" using mult_strict_right_mono [of a 0 b] by simp text \Legacy -- use @{thm [source] mult_neg_pos}.\ lemma mult_pos_neg2: "0 < a \ b < 0 \ b * a < 0" by (drule mult_strict_right_mono [of b 0]) auto lemma zero_less_mult_pos: "0 < a * b \ 0 < a \ 0 < b" apply (cases "b \ 0") apply (auto simp add: le_less not_less) apply (drule_tac mult_pos_neg [of a b]) apply (auto dest: less_not_sym) done lemma zero_less_mult_pos2: "0 < b * a \ 0 < a \ 0 < b" apply (cases "b \ 0") apply (auto simp add: le_less not_less) apply (drule_tac mult_pos_neg2 [of a b]) apply (auto dest: less_not_sym) done text \Strict monotonicity in both arguments\ lemma mult_strict_mono: assumes "a < b" and "c < d" and "0 < b" and "0 \ c" shows "a * c < b * d" using assms apply (cases "c = 0") apply simp apply (erule mult_strict_right_mono [THEN less_trans]) apply (auto simp add: le_less) apply (erule (1) mult_strict_left_mono) done text \This weaker variant has more natural premises\ lemma mult_strict_mono': assumes "a < b" and "c < d" and "0 \ a" and "0 \ c" shows "a * c < b * d" by (rule mult_strict_mono) (insert assms, auto) lemma mult_less_le_imp_less: assumes "a < b" and "c \ d" and "0 \ a" and "0 < c" shows "a * c < b * d" using assms apply (subgoal_tac "a * c < b * c") apply (erule less_le_trans) apply (erule mult_left_mono) apply simp apply (erule (1) mult_strict_right_mono) done lemma mult_le_less_imp_less: assumes "a \ b" and "c < d" and "0 < a" and "0 \ c" shows "a * c < b * d" using assms apply (subgoal_tac "a * c \ b * c") apply (erule le_less_trans) apply (erule mult_strict_left_mono) apply simp apply (erule (1) mult_right_mono) done end class linordered_semiring_1_strict = linordered_semiring_strict + semiring_1 + zero_less_one begin subclass linordered_semiring_1 .. lemma convex_bound_lt: assumes "x < a" "y < a" "0 \ u" "0 \ v" "u + v = 1" shows "u * x + v * y < a" proof - from assms have "u * x + v * y < u * a + v * a" by (cases "u = 0") (auto intro!: add_less_le_mono mult_strict_left_mono mult_left_mono) with assms show ?thesis unfolding distrib_right[symmetric] by simp qed end class ordered_comm_semiring = comm_semiring_0 + ordered_ab_semigroup_add + assumes comm_mult_left_mono: "a \ b \ 0 \ c \ c * a \ c * b" begin subclass ordered_semiring proof fix a b c :: 'a assume "a \ b" "0 \ c" then show "c * a \ c * b" by (rule comm_mult_left_mono) then show "a * c \ b * c" by (simp only: mult.commute) qed end class ordered_cancel_comm_semiring = ordered_comm_semiring + cancel_comm_monoid_add begin subclass comm_semiring_0_cancel .. subclass ordered_comm_semiring .. subclass ordered_cancel_semiring .. end class linordered_comm_semiring_strict = comm_semiring_0 + linordered_cancel_ab_semigroup_add + assumes comm_mult_strict_left_mono: "a < b \ 0 < c \ c * a < c * b" begin subclass linordered_semiring_strict proof fix a b c :: 'a assume "a < b" "0 < c" then show "c * a < c * b" by (rule comm_mult_strict_left_mono) then show "a * c < b * c" by (simp only: mult.commute) qed subclass ordered_cancel_comm_semiring proof fix a b c :: 'a assume "a \ b" "0 \ c" then show "c * a \ c * b" unfolding le_less using mult_strict_left_mono by (cases "c = 0") auto qed end class ordered_ring = ring + ordered_cancel_semiring begin subclass ordered_ab_group_add .. lemma less_add_iff1: "a * e + c < b * e + d \ (a - b) * e + c < d" by (simp add: algebra_simps) lemma less_add_iff2: "a * e + c < b * e + d \ c < (b - a) * e + d" by (simp add: algebra_simps) lemma le_add_iff1: "a * e + c \ b * e + d \ (a - b) * e + c \ d" by (simp add: algebra_simps) lemma le_add_iff2: "a * e + c \ b * e + d \ c \ (b - a) * e + d" by (simp add: algebra_simps) lemma mult_left_mono_neg: "b \ a \ c \ 0 \ c * a \ c * b" apply (drule mult_left_mono [of _ _ "- c"]) apply simp_all done lemma mult_right_mono_neg: "b \ a \ c \ 0 \ a * c \ b * c" apply (drule mult_right_mono [of _ _ "- c"]) apply simp_all done lemma mult_nonpos_nonpos: "a \ 0 \ b \ 0 \ 0 \ a * b" using mult_right_mono_neg [of a 0 b] by simp lemma split_mult_pos_le: "(0 \ a \ 0 \ b) \ (a \ 0 \ b \ 0) \ 0 \ a * b" by (auto simp add: mult_nonpos_nonpos) end class abs_if = minus + uminus + ord + zero + abs + assumes abs_if: "\a\ = (if a < 0 then - a else a)" class linordered_ring = ring + linordered_semiring + linordered_ab_group_add + abs_if begin subclass ordered_ring .. subclass ordered_ab_group_add_abs proof fix a b show "\a + b\ \ \a\ + \b\" by (auto simp add: abs_if not_le not_less algebra_simps simp del: add.commute dest: add_neg_neg add_nonneg_nonneg) qed (auto simp: abs_if) lemma zero_le_square [simp]: "0 \ a * a" using linear [of 0 a] by (auto simp add: mult_nonpos_nonpos) lemma not_square_less_zero [simp]: "\ (a * a < 0)" by (simp add: not_less) proposition abs_eq_iff: "\x\ = \y\ \ x = y \ x = -y" by (auto simp add: abs_if split: if_split_asm) lemma abs_eq_iff': "\a\ = b \ b \ 0 \ (a = b \ a = - b)" by (cases "a \ 0") auto lemma eq_abs_iff': "a = \b\ \ a \ 0 \ (b = a \ b = - a)" using abs_eq_iff' [of b a] by auto lemma sum_squares_ge_zero: "0 \ x * x + y * y" by (intro add_nonneg_nonneg zero_le_square) lemma not_sum_squares_lt_zero: "\ x * x + y * y < 0" by (simp add: not_less sum_squares_ge_zero) end class linordered_ring_strict = ring + linordered_semiring_strict + ordered_ab_group_add + abs_if begin subclass linordered_ring .. lemma mult_strict_left_mono_neg: "b < a \ c < 0 \ c * a < c * b" using mult_strict_left_mono [of b a "- c"] by simp lemma mult_strict_right_mono_neg: "b < a \ c < 0 \ a * c < b * c" using mult_strict_right_mono [of b a "- c"] by simp lemma mult_neg_neg: "a < 0 \ b < 0 \ 0 < a * b" using mult_strict_right_mono_neg [of a 0 b] by simp subclass ring_no_zero_divisors proof fix a b assume "a \ 0" then have a: "a < 0 \ 0 < a" by (simp add: neq_iff) assume "b \ 0" then have b: "b < 0 \ 0 < b" by (simp add: neq_iff) have "a * b < 0 \ 0 < a * b" proof (cases "a < 0") case True show ?thesis proof (cases "b < 0") case True with \a < 0\ show ?thesis by (auto dest: mult_neg_neg) next case False with b have "0 < b" by auto with \a < 0\ show ?thesis by (auto dest: mult_strict_right_mono) qed next case False with a have "0 < a" by auto show ?thesis proof (cases "b < 0") case True with \0 < a\ show ?thesis by (auto dest: mult_strict_right_mono_neg) next case False with b have "0 < b" by auto with \0 < a\ show ?thesis by auto qed qed then show "a * b \ 0" by (simp add: neq_iff) qed lemma zero_less_mult_iff [algebra_split_simps, field_split_simps]: "0 < a * b \ 0 < a \ 0 < b \ a < 0 \ b < 0" by (cases a 0 b 0 rule: linorder_cases[case_product linorder_cases]) (auto simp add: mult_neg_neg not_less le_less dest: zero_less_mult_pos zero_less_mult_pos2) lemma zero_le_mult_iff [algebra_split_simps, field_split_simps]: "0 \ a * b \ 0 \ a \ 0 \ b \ a \ 0 \ b \ 0" by (auto simp add: eq_commute [of 0] le_less not_less zero_less_mult_iff) lemma mult_less_0_iff [algebra_split_simps, field_split_simps]: "a * b < 0 \ 0 < a \ b < 0 \ a < 0 \ 0 < b" using zero_less_mult_iff [of "- a" b] by auto lemma mult_le_0_iff [algebra_split_simps, field_split_simps]: "a * b \ 0 \ 0 \ a \ b \ 0 \ a \ 0 \ 0 \ b" using zero_le_mult_iff [of "- a" b] by auto text \ Cancellation laws for \<^term>\c * a < c * b\ and \<^term>\a * c < b * c\, also with the relations \\\ and equality. \ text \ These ``disjunction'' versions produce two cases when the comparison is an assumption, but effectively four when the comparison is a goal. \ lemma mult_less_cancel_right_disj: "a * c < b * c \ 0 < c \ a < b \ c < 0 \ b < a" apply (cases "c = 0") apply (auto simp add: neq_iff mult_strict_right_mono mult_strict_right_mono_neg) apply (auto simp add: not_less not_le [symmetric, of "a*c"] not_le [symmetric, of a]) apply (erule_tac [!] notE) apply (auto simp add: less_imp_le mult_right_mono mult_right_mono_neg) done lemma mult_less_cancel_left_disj: "c * a < c * b \ 0 < c \ a < b \ c < 0 \ b < a" apply (cases "c = 0") apply (auto simp add: neq_iff mult_strict_left_mono mult_strict_left_mono_neg) apply (auto simp add: not_less not_le [symmetric, of "c * a"] not_le [symmetric, of a]) apply (erule_tac [!] notE) apply (auto simp add: less_imp_le mult_left_mono mult_left_mono_neg) done text \ The ``conjunction of implication'' lemmas produce two cases when the comparison is a goal, but give four when the comparison is an assumption. \ lemma mult_less_cancel_right: "a * c < b * c \ (0 \ c \ a < b) \ (c \ 0 \ b < a)" using mult_less_cancel_right_disj [of a c b] by auto lemma mult_less_cancel_left: "c * a < c * b \ (0 \ c \ a < b) \ (c \ 0 \ b < a)" using mult_less_cancel_left_disj [of c a b] by auto lemma mult_le_cancel_right: "a * c \ b * c \ (0 < c \ a \ b) \ (c < 0 \ b \ a)" by (simp add: not_less [symmetric] mult_less_cancel_right_disj) lemma mult_le_cancel_left: "c * a \ c * b \ (0 < c \ a \ b) \ (c < 0 \ b \ a)" by (simp add: not_less [symmetric] mult_less_cancel_left_disj) lemma mult_le_cancel_left_pos: "0 < c \ c * a \ c * b \ a \ b" by (auto simp: mult_le_cancel_left) lemma mult_le_cancel_left_neg: "c < 0 \ c * a \ c * b \ b \ a" by (auto simp: mult_le_cancel_left) lemma mult_less_cancel_left_pos: "0 < c \ c * a < c * b \ a < b" by (auto simp: mult_less_cancel_left) lemma mult_less_cancel_left_neg: "c < 0 \ c * a < c * b \ b < a" by (auto simp: mult_less_cancel_left) end lemmas mult_sign_intros = mult_nonneg_nonneg mult_nonneg_nonpos mult_nonpos_nonneg mult_nonpos_nonpos mult_pos_pos mult_pos_neg mult_neg_pos mult_neg_neg class ordered_comm_ring = comm_ring + ordered_comm_semiring begin subclass ordered_ring .. subclass ordered_cancel_comm_semiring .. end class linordered_nonzero_semiring = ordered_comm_semiring + monoid_mult + linorder + zero_less_one + assumes add_mono1: "a < b \ a + 1 < b + 1" begin subclass zero_neq_one by standard (insert zero_less_one, blast) subclass comm_semiring_1 by standard (rule mult_1_left) lemma zero_le_one [simp]: "0 \ 1" by (rule zero_less_one [THEN less_imp_le]) lemma not_one_le_zero [simp]: "\ 1 \ 0" by (simp add: not_le) lemma not_one_less_zero [simp]: "\ 1 < 0" by (simp add: not_less) lemma mult_left_le: "c \ 1 \ 0 \ a \ a * c \ a" using mult_left_mono[of c 1 a] by simp lemma mult_le_one: "a \ 1 \ 0 \ b \ b \ 1 \ a * b \ 1" using mult_mono[of a 1 b 1] by simp lemma zero_less_two: "0 < 1 + 1" using add_pos_pos[OF zero_less_one zero_less_one] . end class linordered_semidom = semidom + linordered_comm_semiring_strict + zero_less_one + assumes le_add_diff_inverse2 [simp]: "b \ a \ a - b + b = a" begin subclass linordered_nonzero_semiring proof show "a + 1 < b + 1" if "a < b" for a b proof (rule ccontr, simp add: not_less) assume "b \ a" with that show False by (simp add: ) qed qed text \Addition is the inverse of subtraction.\ lemma le_add_diff_inverse [simp]: "b \ a \ b + (a - b) = a" by (frule le_add_diff_inverse2) (simp add: add.commute) lemma add_diff_inverse: "\ a < b \ b + (a - b) = a" by simp lemma add_le_imp_le_diff: "i + k \ n \ i \ n - k" apply (subst add_le_cancel_right [where c=k, symmetric]) apply (frule le_add_diff_inverse2) apply (simp only: add.assoc [symmetric]) using add_implies_diff apply fastforce done lemma add_le_add_imp_diff_le: assumes 1: "i + k \ n" and 2: "n \ j + k" shows "i + k \ n \ n \ j + k \ n - k \ j" proof - have "n - (i + k) + (i + k) = n" using 1 by simp moreover have "n - k = n - k - i + i" using 1 by (simp add: add_le_imp_le_diff) ultimately show ?thesis using 2 apply (simp add: add.assoc [symmetric]) apply (rule add_le_imp_le_diff [of _ k "j + k", simplified add_diff_cancel_right']) apply (simp add: add.commute diff_diff_add) done qed lemma less_1_mult: "1 < m \ 1 < n \ 1 < m * n" using mult_strict_mono [of 1 m 1 n] by (simp add: less_trans [OF zero_less_one]) end class linordered_idom = comm_ring_1 + linordered_comm_semiring_strict + ordered_ab_group_add + abs_if + sgn + assumes sgn_if: "sgn x = (if x = 0 then 0 else if 0 < x then 1 else - 1)" begin subclass linordered_ring_strict .. subclass linordered_semiring_1_strict proof have "0 \ 1 * 1" by (fact zero_le_square) then show "0 < 1" by (simp add: le_less) qed subclass ordered_comm_ring .. subclass idom .. subclass linordered_semidom by standard simp subclass idom_abs_sgn by standard (auto simp add: sgn_if abs_if zero_less_mult_iff) lemma linorder_neqE_linordered_idom: assumes "x \ y" obtains "x < y" | "y < x" using assms by (rule neqE) text \These cancellation simp rules also produce two cases when the comparison is a goal.\ lemma mult_le_cancel_right1: "c \ b * c \ (0 < c \ 1 \ b) \ (c < 0 \ b \ 1)" using mult_le_cancel_right [of 1 c b] by simp lemma mult_le_cancel_right2: "a * c \ c \ (0 < c \ a \ 1) \ (c < 0 \ 1 \ a)" using mult_le_cancel_right [of a c 1] by simp lemma mult_le_cancel_left1: "c \ c * b \ (0 < c \ 1 \ b) \ (c < 0 \ b \ 1)" using mult_le_cancel_left [of c 1 b] by simp lemma mult_le_cancel_left2: "c * a \ c \ (0 < c \ a \ 1) \ (c < 0 \ 1 \ a)" using mult_le_cancel_left [of c a 1] by simp lemma mult_less_cancel_right1: "c < b * c \ (0 \ c \ 1 < b) \ (c \ 0 \ b < 1)" using mult_less_cancel_right [of 1 c b] by simp lemma mult_less_cancel_right2: "a * c < c \ (0 \ c \ a < 1) \ (c \ 0 \ 1 < a)" using mult_less_cancel_right [of a c 1] by simp lemma mult_less_cancel_left1: "c < c * b \ (0 \ c \ 1 < b) \ (c \ 0 \ b < 1)" using mult_less_cancel_left [of c 1 b] by simp lemma mult_less_cancel_left2: "c * a < c \ (0 \ c \ a < 1) \ (c \ 0 \ 1 < a)" using mult_less_cancel_left [of c a 1] by simp lemma sgn_0_0: "sgn a = 0 \ a = 0" by (fact sgn_eq_0_iff) lemma sgn_1_pos: "sgn a = 1 \ a > 0" unfolding sgn_if by simp lemma sgn_1_neg: "sgn a = - 1 \ a < 0" unfolding sgn_if by auto lemma sgn_pos [simp]: "0 < a \ sgn a = 1" by (simp only: sgn_1_pos) lemma sgn_neg [simp]: "a < 0 \ sgn a = - 1" by (simp only: sgn_1_neg) lemma abs_sgn: "\k\ = k * sgn k" unfolding sgn_if abs_if by auto lemma sgn_greater [simp]: "0 < sgn a \ 0 < a" unfolding sgn_if by auto lemma sgn_less [simp]: "sgn a < 0 \ a < 0" unfolding sgn_if by auto lemma abs_sgn_eq_1 [simp]: "a \ 0 \ \sgn a\ = 1" by simp lemma abs_sgn_eq: "\sgn a\ = (if a = 0 then 0 else 1)" by (simp add: sgn_if) lemma sgn_mult_self_eq [simp]: "sgn a * sgn a = of_bool (a \ 0)" by (cases "a > 0") simp_all lemma abs_mult_self_eq [simp]: "\a\ * \a\ = a * a" by (cases "a > 0") simp_all lemma same_sgn_sgn_add: "sgn (a + b) = sgn a" if "sgn b = sgn a" proof (cases a 0 rule: linorder_cases) case equal with that show ?thesis by simp next case less with that have "b < 0" by (simp add: sgn_1_neg) with \a < 0\ have "a + b < 0" by (rule add_neg_neg) with \a < 0\ show ?thesis by simp next case greater with that have "b > 0" by (simp add: sgn_1_pos) with \a > 0\ have "a + b > 0" by (rule add_pos_pos) with \a > 0\ show ?thesis by simp qed lemma same_sgn_abs_add: "\a + b\ = \a\ + \b\" if "sgn b = sgn a" proof - have "a + b = sgn a * \a\ + sgn b * \b\" by (simp add: sgn_mult_abs) also have "\ = sgn a * (\a\ + \b\)" using that by (simp add: algebra_simps) finally show ?thesis by (auto simp add: abs_mult) qed lemma sgn_not_eq_imp: "sgn a = - sgn b" if "sgn b \ sgn a" and "sgn a \ 0" and "sgn b \ 0" using that by (cases "a < 0") (auto simp add: sgn_0_0 sgn_1_pos sgn_1_neg) lemma abs_dvd_iff [simp]: "\m\ dvd k \ m dvd k" by (simp add: abs_if) lemma dvd_abs_iff [simp]: "m dvd \k\ \ m dvd k" by (simp add: abs_if) lemma dvd_if_abs_eq: "\l\ = \k\ \ l dvd k" by (subst abs_dvd_iff [symmetric]) simp text \ The following lemmas can be proven in more general structures, but are dangerous as simp rules in absence of @{thm neg_equal_zero}, @{thm neg_less_pos}, @{thm neg_less_eq_nonneg}. \ lemma equation_minus_iff_1 [simp, no_atp]: "1 = - a \ a = - 1" by (fact equation_minus_iff) lemma minus_equation_iff_1 [simp, no_atp]: "- a = 1 \ a = - 1" by (subst minus_equation_iff, auto) lemma le_minus_iff_1 [simp, no_atp]: "1 \ - b \ b \ - 1" by (fact le_minus_iff) lemma minus_le_iff_1 [simp, no_atp]: "- a \ 1 \ - 1 \ a" by (fact minus_le_iff) lemma less_minus_iff_1 [simp, no_atp]: "1 < - b \ b < - 1" by (fact less_minus_iff) lemma minus_less_iff_1 [simp, no_atp]: "- a < 1 \ - 1 < a" by (fact minus_less_iff) lemma add_less_zeroD: shows "x+y < 0 \ x<0 \ y<0" by (auto simp: not_less intro: le_less_trans [of _ "x+y"]) end text \Reasoning about inequalities with division\ context linordered_semidom begin lemma less_add_one: "a < a + 1" proof - have "a + 0 < a + 1" by (blast intro: zero_less_one add_strict_left_mono) then show ?thesis by simp qed end context linordered_idom begin lemma mult_right_le_one_le: "0 \ x \ 0 \ y \ y \ 1 \ x * y \ x" by (rule mult_left_le) lemma mult_left_le_one_le: "0 \ x \ 0 \ y \ y \ 1 \ y * x \ x" by (auto simp add: mult_le_cancel_right2) end text \Absolute Value\ context linordered_idom begin lemma mult_sgn_abs: "sgn x * \x\ = x" by (fact sgn_mult_abs) lemma abs_one: "\1\ = 1" by (fact abs_1) end class ordered_ring_abs = ordered_ring + ordered_ab_group_add_abs + assumes abs_eq_mult: "(0 \ a \ a \ 0) \ (0 \ b \ b \ 0) \ \a * b\ = \a\ * \b\" context linordered_idom begin subclass ordered_ring_abs by standard (auto simp: abs_if not_less mult_less_0_iff) lemma abs_mult_self: "\a\ * \a\ = a * a" by (fact abs_mult_self_eq) lemma abs_mult_less: assumes ac: "\a\ < c" and bd: "\b\ < d" shows "\a\ * \b\ < c * d" proof - from ac have "0 < c" by (blast intro: le_less_trans abs_ge_zero) with bd show ?thesis by (simp add: ac mult_strict_mono) qed lemma abs_less_iff: "\a\ < b \ a < b \ - a < b" by (simp add: less_le abs_le_iff) (auto simp add: abs_if) lemma abs_mult_pos: "0 \ x \ \y\ * x = \y * x\" by (simp add: abs_mult) lemma abs_diff_less_iff: "\x - a\ < r \ a - r < x \ x < a + r" by (auto simp add: diff_less_eq ac_simps abs_less_iff) lemma abs_diff_le_iff: "\x - a\ \ r \ a - r \ x \ x \ a + r" by (auto simp add: diff_le_eq ac_simps abs_le_iff) lemma abs_add_one_gt_zero: "0 < 1 + \x\" by (auto simp: abs_if not_less intro: zero_less_one add_strict_increasing less_trans) end subsection \Dioids\ text \ Dioids are the alternative extensions of semirings, a semiring can either be a ring or a dioid but never both. \ class dioid = semiring_1 + canonically_ordered_monoid_add begin subclass ordered_semiring by standard (auto simp: le_iff_add distrib_left distrib_right) end hide_fact (open) comm_mult_left_mono comm_mult_strict_left_mono distrib code_identifier code_module Rings \ (SML) Arith and (OCaml) Arith and (Haskell) Arith end