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,6465 +1,6244 @@ (* 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" +lemma fps_mult_nth_1: "(f * g) $ 1 = f$0 * g$1 + f$1 * g$0" + by (simp add: fps_mult_nth) + +lemma fps_mult_nth_1' [simp]: "(f * g) $ Suc 0 = f$0 * g$Suc 0 + f$Suc 0 * 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_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" 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 + proof (cases "f = 0") + case True + then show ?thesis + by (simp add: fps_euclidean_size_def) + next + case False + then show ?thesis + using le_less_linear[of "subdegree g" "subdegree f"] + by (force simp add: fps_mod_eq_zero fps_euclidean_size_def subdegree_mod) + qed 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 +proof (intro strip sum.mono_neutral_right) + show "\i. i \ {0..n} - {0..k} \ a ^ i $ k = 0" + by (simp add: a0 startsby_zero_power_prefix) +qed (use kn in auto) 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 + by (simp add: not_less le_less_Suc_eq) 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 +proof - + have "\length xs = 1; n = sum_list xs\ \ xs = [sum_list xs]" for xs + by (cases xs) auto + then show ?thesis + by (auto simp add: natpermute_def) +qed + +lemma natlist_trivial_Suc0 [simp]: "natpermute n (Suc 0) = {[n]}" + using natlist_trivial_1 by force 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 + using assms by (force simp add: natpermute_def) 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 + have "sum_list (take h l) \ sum_list l" + using l_take_drop ls m by presburger + with xs ys ls l show "l \ ?R" + by simp (metis append_take_drop_id m) 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 + by (auto simp add: set_replicate_conv_if natpermute_def replicate_length_same) 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 + by (simp add: natpermute_0) 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 + then show ?case + using natpermute_split [of k "Suc k"] finite_UN_I by simp 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 + by (auto simp add: natpermute_def atLeastLessThanSuc_atLeastAtMost sum_list_sum_nth) 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 + unfolding xs using set_update_memI i + by (metis Suc_eq_plus1 atLeast0AtMost atMost_iff le_simps(2) length_replicate) 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 + by simp 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 "\ = (\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] sum_distrib_right .. + also have "... = (\i = 0..n. + \v\(\l1. l1 @ [n - i]) ` natpermute i (Suc k). + (\j = 0..k. a j $ v ! j) * a (Suc k) $ v ! Suc k)" + by (intro sum.cong [OF refl sym] sum.reindex_cong) (auto simp: inj_on_def natpermute_def nth_append Suc) + also have "... = (\v\(\x\{0..n}. {l1 @ [n - x] |l1. l1 \ natpermute x (Suc k)}). + (\j = 0..k. a j $ v ! j) * a (Suc k) $ v ! Suc k)" + by (subst sum.UNION_disjoint) (auto simp add: natpermute_finite setcompr_eq_image) 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 + using natpermute_split[of m "m + 1"] by (simp add: Suc) 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 + using H Suc by auto 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 + using not_less by auto 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 radical_0 [simp]: "\n. 0 < n \ radical r 0 a n = 0" + using radical.elims by blast + 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 + by (auto simp add: fps_eq_iff fps_radical_def) 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 + proof (rule prod.cong [OF refl]) + show "fps_radical r k a $ replicate k 0 ! j = r k (a $ 0)" if "j \ {0..h}" for j + proof - + have "j < Suc h" + using that by presburger + then show ?thesis + by (metis Suc fps_radical_nth_0 nth_replicate old.nat.distinct(2)) + qed + qed also have "\ = a$0" - using r Suc by (simp add: prod_constant) + using r Suc by simp 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 + (\j\{0..k}. if j = i then fps_radical r (Suc k) a $ n else r (Suc k) (a$0))" + using i r0 by (auto simp del: replicate.simps intro: prod.cong) 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 + using i a0 by (auto simp del: replicate.simps intro: prod.cong) 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" + 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 + have "a$n = (b$n - sum ?f ?Pnknn) / (of_nat (k+1) * (?r $ 0)^k)" + apply (rule eq_divide_imp) + using r00 \
by (simp_all add: ac_simps del: of_nat_Suc) 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 + by (simp del: of_nat_Suc) 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 + by (intro sum.mono_neutral_right) (auto simp add: mult_delta_left not_le) 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 + by (auto simp: subset_eq b0 startsby_zero_power_prefix sum.mono_neutral_left intro: sum.cong) 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 + by (subst sum.swap) (force intro: sum.cong) 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 +proof - + have "compinv x n = gcompinv fps_X x n" for n and x :: "'a fps" + proof (induction n rule: nat_less_induct) + case (1 n) + then show ?case + by (cases n) auto + qed + then show ?thesis + by (auto simp add: fun_eq_iff fps_eq_iff fps_inv_def fps_ginv_def) +qed 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) + by (auto intro: finite_subset[OF s]) + have "?r = (\(k, m) \ {(k, m). k + m \ n}. \j = 0..n. a $ k * b $ m * (c ^ k $ j * d ^ m $ (n - j)))" + by (simp add: fps_mult_nth sum_distrib_left) + also have "\ = (\i = 0..n. \(k,m)\{(k,m). k+m \ n}. a $ k * c ^ k $ i * b $ m * d ^ m $ (n-i))" + unfolding sum.swap [where A = "{0..n}"] by (auto simp add: field_simps intro: sum.cong) + also have "... = (\i = 0..n. + \q = 0..i. \j = 0..n - i. a $ q * c ^ q $ i * (b $ j * d ^ j $ (n - i)))" + apply (rule sum.cong [OF 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 + apply (rule sum.mono_neutral_right[OF f], force) + by clarsimp (meson c0 d0 leI startsby_zero_power_prefix) + also have "\ = ?l" + by (simp add: fps_mult_nth fps_compose_nth sum_product) 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})" + have th0: "{(k, m). k + m \ n} = (\s\{0..n}. \i\{0..s}. {(i, s - i)})" by auto - show "?l = ?r " + show "?l = ?r" unfolding th0 - apply (subst sum.UNION_disjoint) - apply auto - apply (subst sum.UNION_disjoint) - apply auto - done + by (simp add: sum.UNION_disjoint eq_diff_iff disjoint_iff) 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 +proof (clarsimp simp add: fps_eq_iff fps_compose_mult_distrib_lemma [OF c0]) + show "(a * b oo c) $ n = (\s = 0..n. \i = 0..s. a $ i * b $ (s - i) * c ^ s $ n)" for n + by (simp add: fps_compose_nth fps_mult_nth sum_distrib_right) +qed + 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 +proof (induct S rule: infinite_finite_induct) +next + case (insert) + then show ?case + by (simp add: fps_compose_mult_distrib[OF c0]) +qed auto 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 "... = (\i = 0..n. \j = 0..n. a $ j * (b ^ j $ i * c ^ i $ n))" + by (simp add: fps_compose_nth fps_sum_nth sum_distrib_right mult.assoc) + also have "... = (\i = 0..n. \j = 0..i. a $ j * (b ^ j $ i * c ^ i $ n))" + by (intro sum.cong [OF refl] sum.mono_neutral_right; simp add: b0 startsby_zero_power_prefix) 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 + by (simp add: fps_compose_nth sum_distrib_right mult.assoc) 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 + by (simp add: a0 fps_compose_assoc rca0) 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 + by (metis c0 c1 fps_compose_assoc fps_compose_nth_0 fps_inv fps_inv_right) 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 + using of_nat_neq_0 by (auto simp add: fps_exp_def divide_simps) 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 + by (simp add: th divide_simps) 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 + by (simp add: fps_compose_add_distrib fps_inv_right[OF b0 b1] distrib_left flip: fps_const_mult_apply_left) 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 + by (simp add: eq fps_ln_deriv) 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 + unfolding fps_divide_def mult.assoc[symmetric] inverse_mult_eq_1[OF x10] + by (simp add: field_simps) 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 + by (simp add: field_simps del: of_nat_Suc split: if_split_asm) 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 + have "(c - of_nat m) * (c gchoose m) = (c gchoose Suc m) * of_nat (Suc m)" + by (metis gbinomial_absorb_comp gbinomial_absorption mult.commute) + with Suc 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 + by (simp add: divide_simps del: of_nat_Suc) qed show ?thesis - apply (simp add: fps_eq_iff) - apply (subst th1) - apply (simp add: field_simps) - done + by (metis expand_fps_eq fps_binomial_nth fps_mult_right_const_nth mult.commute th1) 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 + have "?l = (1 + fps_X) * fps_deriv (fps_const (a $ 0) * fps_binomial c)" + using that by auto + also have "... = fps_const c * (fps_const (a $ 0) * fps_binomial c)" + proof (clarsimp simp add: fps_eq_iff algebra_simps) + show "a $ 0 * (c gchoose Suc n) + (of_nat n * ((c gchoose n) * a $ 0) + of_nat n * (a $ 0 * (c gchoose Suc n))) + = c * ((c gchoose n) * a $ 0)" for n + unfolding mult.assoc[symmetric] + by (simp add: field_simps gbinomial_mult_1) + qed + also have "... = ?r" + using that by auto + finally have "?l = ?r" . 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 + fixes c :: "'a :: field_char_0" + assumes "n > 0" + shows "inverse ((1 - fps_const c * fps_X) ^ n) = Abs_fps (\k. of_nat ((n + k - 1) choose k) * c^k)" +proof - + have \
: "\j. Abs_fps (\na. (- c) ^ na * fps_binomial (- of_nat n) $ na) $ j = + Abs_fps (\k. of_nat (n + k - 1 choose k) * c ^ k) $ j" + using assms + by (simp add: gbinomial_minus binomial_gbinomial of_nat_diff flip: power_mult_distrib mult.assoc) + show ?thesis + apply (rule fps_ext) + using \
+ by (metis (no_types, lifting) one_minus_fps_X_const_neg_power fps_const_neg fps_compose_linear fps_nth_Abs_fps) +qed 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 + by (metis atMost_atLeast0 choose_square_sum mult_2) lemma Vandermonde_pochhammer_lemma: fixes a :: "'a::field_char_0" - assumes b: "\j\{0 .. of_nat j" + assumes b: "\j. j b \ 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 + then show False + using \j < n\ j b 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 + with pochhammer_minus'[where k=k and b=b] bn0 show ?thesis + by (simp add: pochhammer_same) next case False with kn have kn': "k < n" by simp + have "h \ m" + using \k \ n\ h m by blast have m1nk: "?m1 n = prod (\i. - 1) {..m}" "?m1 k = prod (\i. - 1) {0..h}" - by (simp_all add: prod_constant m h) + by (simp_all add: 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 + by (metis add.commute add_diff_eq nz' pochhammer_eq_0_iff) 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) + have "(\i = 0..x = n - k..k \ n\ 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 + by (auto simp add: of_nat_diff field_simps) + then have "fact (n - k) * pochhammer ((1::'a) + of_nat n - of_nat k) k = fact n" + using \k \ n\ + by (auto simp add: fact_split [of k n] pochhammer_prod field_simps) + then have th1: "(?m1 k * ?p (of_nat n) k) / ?f n = 1 / of_nat(fact (n - k))" + by (simp add: pochhammer_minus field_simps) + have "?m1 n * ?p b n = pochhammer (b - of_nat m) (Suc m)" + by (simp add: pochhammer_minus field_simps m) + also have "... = (\i = 0..m. b - of_nat i)" + by (auto simp add: pochhammer_prod_rev of_nat_diff prod.atLeast_Suc_atMost_Suc_shift simp del: prod.cl_ivl_Suc) + finally have th20: "?m1 n * ?p b n = prod (\i. b - of_nat i) {0..m}" . + have "(\x = 0..h. b - of_nat m + of_nat (h - x)) = (\i = m - h..m. b - of_nat i)" + using \h \ m\ prod.atLeastAtMost_shift_0 [of "m - h" m, where ?'a = 'a] + by (auto simp add: of_nat_diff field_simps) + then have th21:"pochhammer (b - of_nat n + 1) k = prod (\i. b - of_nat i) {n - k .. n - 1}" + using kn by (simp add: pochhammer_prod_rev m h prod.atLeast_Suc_atMost_Suc_shift del: prod.op_ivl_Suc del: prod.cl_ivl_Suc) 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 + using kn' m h unfolding th20 th21 + by (auto simp flip: prod.union_disjoint intro: prod.cong) 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 + by (auto simp: field_simps gbinomial_mult_fact intro: prod.cong) finally show ?thesis by simp qed qed then show ?gchoose and ?pochhammer - apply (cases "n = 0") - using nz' - apply auto - done + using nz' by force+ 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" + using bn0 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 + by (simp add: gbinomial_pochhammer sum_distrib_right sum_distrib_left field_simps) 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 h: "?b \ of_nat j" if "j < n" for j + proof - + have "c \ - of_nat (n - j - 1)" + using c that by auto + with that show ?thesis + by (auto simp add: algebra_simps of_nat_diff) + qed 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 + by (simp add: fps_deriv_power fps_sin_deriv fps_cos_deriv field_simps flip: fps_const_neg) 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" +lemma fps_sin_nth_1 [simp]: "fps_sin c $ Suc 0 = 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 +proof (cases n) + case (Suc n') + then show ?thesis + unfolding fps_sin_def by (simp add: field_simps) +qed (auto simp: fps_sin_def) + 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" +lemma fps_cos_nth_1 [simp]: "fps_cos c $ Suc 0 = 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 +proof (cases n) + case (Suc n') + then show ?thesis + unfolding fps_cos_def by (simp add: field_simps) +qed (auto simp: fps_cos_def) 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 + assumes a0: "a $ 0 = 0" + and a1: "a $ 1 = c" + and a2: "fps_deriv (fps_deriv a) = - (fps_const c * fps_const c * a)" + shows "fps_sin c = a" +proof (rule fps_ext) + fix n + show "fps_sin c $ n = a $ n" + proof (induction n rule: nat_induct2) + case (step n) + then have "of_nat (n + 1) * (of_nat (n + 2) * a $ (n + 2)) = + - (c * c * fps_sin c $ n)" + using a2 + by (metis fps_const_mult fps_deriv_nth fps_mult_left_const_nth fps_neg_nth nat_add_1_add_1) + with step show ?case + by (metis (no_types, lifting) a0 add.commute add.inverse_inverse fps_sin_nth_0 fps_sin_nth_add_2 mult_divide_mult_cancel_left_if mult_minus_right nonzero_mult_div_cancel_left not_less_zero of_nat_eq_0_iff plus_1_eq_Suc zero_less_Suc) + qed (use assms in auto) +qed 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 + assumes a0: "a $ 0 = 1" + and a1: "a $ 1 = 0" + and a2: "fps_deriv (fps_deriv a) = - (fps_const c * fps_const c * a)" + shows "fps_cos c = a" +proof (rule fps_ext) + fix n + show "fps_cos c $ n = a $ n" + proof (induction n rule: nat_induct2) + case (step n) + then have "of_nat (n + 1) * (of_nat (n + 2) * a $ (n + 2)) = + - (c * c * fps_cos c $ n)" + using a2 + by (metis fps_const_mult fps_deriv_nth fps_mult_left_const_nth fps_neg_nth nat_add_1_add_1) + with step show ?case + by (metis (no_types, lifting) a0 add.commute add.inverse_inverse fps_cos_nth_0 fps_cos_nth_add_2 mult_divide_mult_cancel_left_if mult_minus_right nonzero_mult_div_cancel_left not_less_zero of_nat_eq_0_iff plus_1_eq_Suc zero_less_Suc) + qed (use assms in auto) +qed 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 +proof - + have "fps_deriv (fps_deriv (fps_sin a * fps_cos b + fps_cos a * fps_sin b)) = + - (fps_const (a + b) * fps_const (a + b) * (fps_sin a * fps_cos b + fps_cos a * fps_sin b))" + by (simp flip: fps_const_neg fps_const_add fps_const_mult + add: fps_sin_deriv fps_cos_deriv algebra_simps) + then show ?thesis + by (auto intro: eq_fps_sin) +qed + 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 +proof - + have "fps_deriv + (fps_deriv (fps_cos a * fps_cos b - fps_sin a * fps_sin b)) = + - (fps_const (a + b) * fps_const (a + b) * + (fps_cos a * fps_cos b - fps_sin a * fps_sin b))" + by (simp flip: fps_const_neg fps_const_add fps_const_mult + add: fps_sin_deriv fps_cos_deriv algebra_simps) + then show ?thesis + by (auto intro: eq_fps_cos) +qed 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 + unfolding fps_tan_def fps_sin_fps_exp_ii fps_cos_fps_exp_ii + by (simp add: fps_divide_unit fps_inverse_mult fps_const_inverse) 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 +proof - + have "foldl (\(r::'a) (a::'a). r) 1 as = 1" for as + by (induction as) auto + then show ?thesis + by auto +qed 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) + apply (simp add: foldl_mult_start del: of_nat_Suc of_nat_add fact_Suc) unfolding foldl_prod_prod[unfolded foldl_mult_start] pochhammer_Suc - apply (simp add: algebra_simps) - done + by (simp add: algebra_simps) 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/Set_Interval.thy b/src/HOL/Set_Interval.thy --- a/src/HOL/Set_Interval.thy +++ b/src/HOL/Set_Interval.thy @@ -1,2517 +1,2522 @@ (* Title: HOL/Set_Interval.thy Author: Tobias Nipkow, Clemens Ballarin, Jeremy Avigad lessThan, greaterThan, atLeast, atMost and two-sided intervals Modern convention: Ixy stands for an interval where x and y describe the lower and upper bound and x,y : {c,o,i} where c = closed, o = open, i = infinite. Examples: Ico = {_ ..< _} and Ici = {_ ..} *) section \Set intervals\ theory Set_Interval imports Divides begin (* Belongs in Finite_Set but 2 is not available there *) lemma card_2_iff: "card S = 2 \ (\x y. S = {x,y} \ x \ y)" by (auto simp: card_Suc_eq numeral_eq_Suc) lemma card_2_iff': "card S = 2 \ (\x\S. \y\S. x \ y \ (\z\S. z = x \ z = y))" by (auto simp: card_Suc_eq numeral_eq_Suc) context ord begin definition lessThan :: "'a => 'a set" ("(1{..<_})") where "{.. 'a set" ("(1{.._})") where "{..u} == {x. x \ u}" definition greaterThan :: "'a => 'a set" ("(1{_<..})") where "{l<..} == {x. l 'a set" ("(1{_..})") where "{l..} == {x. l\x}" definition greaterThanLessThan :: "'a => 'a => 'a set" ("(1{_<..<_})") where "{l<.. 'a => 'a set" ("(1{_..<_})") where "{l.. 'a => 'a set" ("(1{_<.._})") where "{l<..u} == {l<..} Int {..u}" definition atLeastAtMost :: "'a => 'a => 'a set" ("(1{_.._})") where "{l..u} == {l..} Int {..u}" end text\A note of warning when using \<^term>\{.. on type \<^typ>\nat\: it is equivalent to \<^term>\{0::nat.. but some lemmas involving \<^term>\{m.. may not exist in \<^term>\{..-form as well.\ syntax (ASCII) "_UNION_le" :: "'a => 'a => 'b set => 'b set" ("(3UN _<=_./ _)" [0, 0, 10] 10) "_UNION_less" :: "'a => 'a => 'b set => 'b set" ("(3UN _<_./ _)" [0, 0, 10] 10) "_INTER_le" :: "'a => 'a => 'b set => 'b set" ("(3INT _<=_./ _)" [0, 0, 10] 10) "_INTER_less" :: "'a => 'a => 'b set => 'b set" ("(3INT _<_./ _)" [0, 0, 10] 10) syntax (latex output) "_UNION_le" :: "'a \ 'a => 'b set => 'b set" ("(3\(\unbreakable\_ \ _)/ _)" [0, 0, 10] 10) "_UNION_less" :: "'a \ 'a => 'b set => 'b set" ("(3\(\unbreakable\_ < _)/ _)" [0, 0, 10] 10) "_INTER_le" :: "'a \ 'a => 'b set => 'b set" ("(3\(\unbreakable\_ \ _)/ _)" [0, 0, 10] 10) "_INTER_less" :: "'a \ 'a => 'b set => 'b set" ("(3\(\unbreakable\_ < _)/ _)" [0, 0, 10] 10) syntax "_UNION_le" :: "'a => 'a => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) "_UNION_less" :: "'a => 'a => 'b set => 'b set" ("(3\_<_./ _)" [0, 0, 10] 10) "_INTER_le" :: "'a => 'a => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) "_INTER_less" :: "'a => 'a => 'b set => 'b set" ("(3\_<_./ _)" [0, 0, 10] 10) translations "\i\n. A" \ "\i\{..n}. A" "\i "\i\{..i\n. A" \ "\i\{..n}. A" "\i "\i\{..Various equivalences\ lemma (in ord) lessThan_iff [iff]: "(i \ lessThan k) = (i greaterThan k) = (k atLeast k) = (k<=i)" by (simp add: atLeast_def) lemma Compl_atLeast [simp]: "!!k:: 'a::linorder. -atLeast k = lessThan k" by (auto simp add: lessThan_def atLeast_def) lemma (in ord) atMost_iff [iff]: "(i \ atMost k) = (i<=k)" by (simp add: atMost_def) lemma atMost_Int_atLeast: "!!n:: 'a::order. atMost n Int atLeast n = {n}" by (blast intro: order_antisym) lemma (in linorder) lessThan_Int_lessThan: "{ a <..} \ { b <..} = { max a b <..}" by auto lemma (in linorder) greaterThan_Int_greaterThan: "{..< a} \ {..< b} = {..< min a b}" by auto subsection \Logical Equivalences for Set Inclusion and Equality\ lemma atLeast_empty_triv [simp]: "{{}..} = UNIV" by auto lemma atMost_UNIV_triv [simp]: "{..UNIV} = UNIV" by auto lemma atLeast_subset_iff [iff]: "(atLeast x \ atLeast y) = (y \ (x::'a::preorder))" by (blast intro: order_trans) lemma atLeast_eq_iff [iff]: "(atLeast x = atLeast y) = (x = (y::'a::order))" by (blast intro: order_antisym order_trans) lemma greaterThan_subset_iff [iff]: "(greaterThan x \ greaterThan y) = (y \ (x::'a::linorder))" unfolding greaterThan_def by (auto simp: linorder_not_less [symmetric]) lemma greaterThan_eq_iff [iff]: "(greaterThan x = greaterThan y) = (x = (y::'a::linorder))" by (auto simp: elim!: equalityE) lemma atMost_subset_iff [iff]: "(atMost x \ atMost y) = (x \ (y::'a::preorder))" by (blast intro: order_trans) lemma atMost_eq_iff [iff]: "(atMost x = atMost y) = (x = (y::'a::order))" by (blast intro: order_antisym order_trans) lemma lessThan_subset_iff [iff]: "(lessThan x \ lessThan y) = (x \ (y::'a::linorder))" unfolding lessThan_def by (auto simp: linorder_not_less [symmetric]) lemma lessThan_eq_iff [iff]: "(lessThan x = lessThan y) = (x = (y::'a::linorder))" by (auto simp: elim!: equalityE) lemma lessThan_strict_subset_iff: fixes m n :: "'a::linorder" shows "{.. m < n" by (metis leD lessThan_subset_iff linorder_linear not_less_iff_gr_or_eq psubset_eq) lemma (in linorder) Ici_subset_Ioi_iff: "{a ..} \ {b <..} \ b < a" by auto lemma (in linorder) Iic_subset_Iio_iff: "{.. a} \ {..< b} \ a < b" by auto lemma (in preorder) Ioi_le_Ico: "{a <..} \ {a ..}" by (auto intro: less_imp_le) subsection \Two-sided intervals\ context ord begin lemma greaterThanLessThan_iff [simp]: "(i \ {l<.. i < u)" by (simp add: greaterThanLessThan_def) lemma atLeastLessThan_iff [simp]: "(i \ {l.. i \ i < u)" by (simp add: atLeastLessThan_def) lemma greaterThanAtMost_iff [simp]: "(i \ {l<..u}) = (l < i \ i \ u)" by (simp add: greaterThanAtMost_def) lemma atLeastAtMost_iff [simp]: "(i \ {l..u}) = (l \ i \ i \ u)" by (simp add: atLeastAtMost_def) text \The above four lemmas could be declared as iffs. Unfortunately this breaks many proofs. Since it only helps blast, it is better to leave them alone.\ lemma greaterThanLessThan_eq: "{ a <..< b} = { a <..} \ {..< b }" by auto lemma (in order) atLeastLessThan_eq_atLeastAtMost_diff: "{a..Emptyness, singletons, subset\ context preorder begin lemma atLeastatMost_empty_iff[simp]: "{a..b} = {} \ (\ a \ b)" by auto (blast intro: order_trans) lemma atLeastatMost_empty_iff2[simp]: "{} = {a..b} \ (\ a \ b)" by auto (blast intro: order_trans) lemma atLeastLessThan_empty_iff[simp]: "{a.. (\ a < b)" by auto (blast intro: le_less_trans) lemma atLeastLessThan_empty_iff2[simp]: "{} = {a.. (\ a < b)" by auto (blast intro: le_less_trans) lemma greaterThanAtMost_empty_iff[simp]: "{k<..l} = {} \ \ k < l" by auto (blast intro: less_le_trans) lemma greaterThanAtMost_empty_iff2[simp]: "{} = {k<..l} \ \ k < l" by auto (blast intro: less_le_trans) lemma atLeastatMost_subset_iff[simp]: "{a..b} \ {c..d} \ (\ a \ b) \ c \ a \ b \ d" unfolding atLeastAtMost_def atLeast_def atMost_def by (blast intro: order_trans) lemma atLeastatMost_psubset_iff: "{a..b} < {c..d} \ ((\ a \ b) \ c \ a \ b \ d \ (c < a \ b < d)) \ c \ d" by(simp add: psubset_eq set_eq_iff less_le_not_le)(blast intro: order_trans) lemma atLeastAtMost_subseteq_atLeastLessThan_iff: "{a..b} \ {c ..< d} \ (a \ b \ c \ a \ b < d)" by auto (blast intro: local.order_trans local.le_less_trans elim: )+ lemma Icc_subset_Ici_iff[simp]: "{l..h} \ {l'..} = (\ l\h \ l\l')" by(auto simp: subset_eq intro: order_trans) lemma Icc_subset_Iic_iff[simp]: "{l..h} \ {..h'} = (\ l\h \ h\h')" by(auto simp: subset_eq intro: order_trans) lemma not_Ici_eq_empty[simp]: "{l..} \ {}" by(auto simp: set_eq_iff) lemma not_Iic_eq_empty[simp]: "{..h} \ {}" by(auto simp: set_eq_iff) lemmas not_empty_eq_Ici_eq_empty[simp] = not_Ici_eq_empty[symmetric] lemmas not_empty_eq_Iic_eq_empty[simp] = not_Iic_eq_empty[symmetric] end context order begin lemma atLeastatMost_empty[simp]: "b < a \ {a..b} = {}" by(auto simp: atLeastAtMost_def atLeast_def atMost_def) lemma atLeastLessThan_empty[simp]: "b \ a \ {a.. k ==> {k<..l} = {}" by(auto simp:greaterThanAtMost_def greaterThan_def atMost_def) lemma greaterThanLessThan_empty[simp]:"l \ k ==> {k<.. {a .. b} = {a}" by simp lemma Icc_eq_Icc[simp]: "{l..h} = {l'..h'} = (l=l' \ h=h' \ \ l\h \ \ l'\h')" by(simp add: order_class.eq_iff)(auto intro: order_trans) lemma atLeastAtMost_singleton_iff[simp]: "{a .. b} = {c} \ a = b \ b = c" proof assume "{a..b} = {c}" hence *: "\ (\ a \ b)" unfolding atLeastatMost_empty_iff[symmetric] by simp with \{a..b} = {c}\ have "c \ a \ b \ c" by auto with * show "a = b \ b = c" by auto qed simp end context no_top begin (* also holds for no_bot but no_top should suffice *) lemma not_UNIV_le_Icc[simp]: "\ UNIV \ {l..h}" using gt_ex[of h] by(auto simp: subset_eq less_le_not_le) lemma not_UNIV_le_Iic[simp]: "\ UNIV \ {..h}" using gt_ex[of h] by(auto simp: subset_eq less_le_not_le) lemma not_Ici_le_Icc[simp]: "\ {l..} \ {l'..h'}" using gt_ex[of h'] by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans) lemma not_Ici_le_Iic[simp]: "\ {l..} \ {..h'}" using gt_ex[of h'] by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans) end context no_bot begin lemma not_UNIV_le_Ici[simp]: "\ UNIV \ {l..}" using lt_ex[of l] by(auto simp: subset_eq less_le_not_le) lemma not_Iic_le_Icc[simp]: "\ {..h} \ {l'..h'}" using lt_ex[of l'] by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans) lemma not_Iic_le_Ici[simp]: "\ {..h} \ {l'..}" using lt_ex[of l'] by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans) end context no_top begin (* also holds for no_bot but no_top should suffice *) lemma not_UNIV_eq_Icc[simp]: "\ UNIV = {l'..h'}" using gt_ex[of h'] by(auto simp: set_eq_iff less_le_not_le) lemmas not_Icc_eq_UNIV[simp] = not_UNIV_eq_Icc[symmetric] lemma not_UNIV_eq_Iic[simp]: "\ UNIV = {..h'}" using gt_ex[of h'] by(auto simp: set_eq_iff less_le_not_le) lemmas not_Iic_eq_UNIV[simp] = not_UNIV_eq_Iic[symmetric] lemma not_Icc_eq_Ici[simp]: "\ {l..h} = {l'..}" unfolding atLeastAtMost_def using not_Ici_le_Iic[of l'] by blast lemmas not_Ici_eq_Icc[simp] = not_Icc_eq_Ici[symmetric] (* also holds for no_bot but no_top should suffice *) lemma not_Iic_eq_Ici[simp]: "\ {..h} = {l'..}" using not_Ici_le_Iic[of l' h] by blast lemmas not_Ici_eq_Iic[simp] = not_Iic_eq_Ici[symmetric] end context no_bot begin lemma not_UNIV_eq_Ici[simp]: "\ UNIV = {l'..}" using lt_ex[of l'] by(auto simp: set_eq_iff less_le_not_le) lemmas not_Ici_eq_UNIV[simp] = not_UNIV_eq_Ici[symmetric] lemma not_Icc_eq_Iic[simp]: "\ {l..h} = {..h'}" unfolding atLeastAtMost_def using not_Iic_le_Ici[of h'] by blast lemmas not_Iic_eq_Icc[simp] = not_Icc_eq_Iic[symmetric] end context dense_linorder begin lemma greaterThanLessThan_empty_iff[simp]: "{ a <..< b } = {} \ b \ a" using dense[of a b] by (cases "a < b") auto lemma greaterThanLessThan_empty_iff2[simp]: "{} = { a <..< b } \ b \ a" using dense[of a b] by (cases "a < b") auto lemma atLeastLessThan_subseteq_atLeastAtMost_iff: "{a ..< b} \ { c .. d } \ (a < b \ c \ a \ b \ d)" using dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanAtMost_subseteq_atLeastAtMost_iff: "{a <.. b} \ { c .. d } \ (a < b \ c \ a \ b \ d)" using dense[of "a" "min c b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanLessThan_subseteq_atLeastAtMost_iff: "{a <..< b} \ { c .. d } \ (a < b \ c \ a \ b \ d)" using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanLessThan_subseteq_greaterThanLessThan: "{a <..< b} \ {c <..< d} \ (a < b \ a \ c \ b \ d)" using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanAtMost_subseteq_atLeastLessThan_iff: "{a <.. b} \ { c ..< d } \ (a < b \ c \ a \ b < d)" using dense[of "a" "min c b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanLessThan_subseteq_atLeastLessThan_iff: "{a <..< b} \ { c ..< d } \ (a < b \ c \ a \ b \ d)" using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanLessThan_subseteq_greaterThanAtMost_iff: "{a <..< b} \ { c <.. d } \ (a < b \ c \ a \ b \ d)" using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) end context no_top begin lemma greaterThan_non_empty[simp]: "{x <..} \ {}" using gt_ex[of x] by auto end context no_bot begin lemma lessThan_non_empty[simp]: "{..< x} \ {}" using lt_ex[of x] by auto end lemma (in linorder) atLeastLessThan_subset_iff: "{a.. {c.. b \ a \ c\a \ b\d" apply (auto simp:subset_eq Ball_def not_le) apply(frule_tac x=a in spec) apply(erule_tac x=d in allE) apply (auto simp: ) done lemma atLeastLessThan_inj: fixes a b c d :: "'a::linorder" assumes eq: "{a ..< b} = {c ..< d}" and "a < b" "c < d" shows "a = c" "b = d" using assms by (metis atLeastLessThan_subset_iff eq less_le_not_le antisym_conv2 subset_refl)+ lemma atLeastLessThan_eq_iff: fixes a b c d :: "'a::linorder" assumes "a < b" "c < d" shows "{a ..< b} = {c ..< d} \ a = c \ b = d" using atLeastLessThan_inj assms by auto lemma (in linorder) Ioc_inj: "{a <.. b} = {c <.. d} \ (b \ a \ d \ c) \ a = c \ b = d" by (metis eq_iff greaterThanAtMost_empty_iff2 greaterThanAtMost_iff le_cases not_le) lemma (in order) Iio_Int_singleton: "{.. {x} = (if x < k then {x} else {})" by auto lemma (in linorder) Ioc_subset_iff: "{a<..b} \ {c<..d} \ (b \ a \ c \ a \ b \ d)" by (auto simp: subset_eq Ball_def) (metis less_le not_less) lemma (in order_bot) atLeast_eq_UNIV_iff: "{x..} = UNIV \ x = bot" by (auto simp: set_eq_iff intro: le_bot) lemma (in order_top) atMost_eq_UNIV_iff: "{..x} = UNIV \ x = top" by (auto simp: set_eq_iff intro: top_le) lemma (in bounded_lattice) atLeastAtMost_eq_UNIV_iff: "{x..y} = UNIV \ (x = bot \ y = top)" by (auto simp: set_eq_iff intro: top_le le_bot) lemma Iio_eq_empty_iff: "{..< n::'a::{linorder, order_bot}} = {} \ n = bot" by (auto simp: set_eq_iff not_less le_bot) lemma lessThan_empty_iff: "{..< n::nat} = {} \ n = 0" by (simp add: Iio_eq_empty_iff bot_nat_def) lemma mono_image_least: assumes f_mono: "mono f" and f_img: "f ` {m ..< n} = {m' ..< n'}" "m < n" shows "f m = m'" proof - from f_img have "{m' ..< n'} \ {}" by (metis atLeastLessThan_empty_iff image_is_empty) with f_img have "m' \ f ` {m ..< n}" by auto then obtain k where "f k = m'" "m \ k" by auto moreover have "m' \ f m" using f_img by auto ultimately show "f m = m'" using f_mono by (auto elim: monoE[where x=m and y=k]) qed subsection \Infinite intervals\ context dense_linorder begin lemma infinite_Ioo: assumes "a < b" shows "\ finite {a<.. {}" using \a < b\ by auto ultimately have "a < Max {a <..< b}" "Max {a <..< b} < b" using Max_in[of "{a <..< b}"] by auto then obtain x where "Max {a <..< b} < x" "x < b" using dense[of "Max {a<.. {a <..< b}" using \a < Max {a <..< b}\ by auto then have "x \ Max {a <..< b}" using fin by auto with \Max {a <..< b} < x\ show False by auto qed lemma infinite_Icc: "a < b \ \ finite {a .. b}" using greaterThanLessThan_subseteq_atLeastAtMost_iff[of a b a b] infinite_Ioo[of a b] by (auto dest: finite_subset) lemma infinite_Ico: "a < b \ \ finite {a ..< b}" using greaterThanLessThan_subseteq_atLeastLessThan_iff[of a b a b] infinite_Ioo[of a b] by (auto dest: finite_subset) lemma infinite_Ioc: "a < b \ \ finite {a <.. b}" using greaterThanLessThan_subseteq_greaterThanAtMost_iff[of a b a b] infinite_Ioo[of a b] by (auto dest: finite_subset) lemma infinite_Ioo_iff [simp]: "infinite {a<.. a < b" using not_less_iff_gr_or_eq by (fastforce simp: infinite_Ioo) lemma infinite_Icc_iff [simp]: "infinite {a .. b} \ a < b" using not_less_iff_gr_or_eq by (fastforce simp: infinite_Icc) lemma infinite_Ico_iff [simp]: "infinite {a.. a < b" using not_less_iff_gr_or_eq by (fastforce simp: infinite_Ico) lemma infinite_Ioc_iff [simp]: "infinite {a<..b} \ a < b" using not_less_iff_gr_or_eq by (fastforce simp: infinite_Ioc) end lemma infinite_Iio: "\ finite {..< a :: 'a :: {no_bot, linorder}}" proof assume "finite {..< a}" then have *: "\x. x < a \ Min {..< a} \ x" by auto obtain x where "x < a" using lt_ex by auto obtain y where "y < Min {..< a}" using lt_ex by auto also have "Min {..< a} \ x" using \x < a\ by fact also note \x < a\ finally have "Min {..< a} \ y" by fact with \y < Min {..< a}\ show False by auto qed lemma infinite_Iic: "\ finite {.. a :: 'a :: {no_bot, linorder}}" using infinite_Iio[of a] finite_subset[of "{..< a}" "{.. a}"] by (auto simp: subset_eq less_imp_le) lemma infinite_Ioi: "\ finite {a :: 'a :: {no_top, linorder} <..}" proof assume "finite {a <..}" then have *: "\x. a < x \ x \ Max {a <..}" by auto obtain y where "Max {a <..} < y" using gt_ex by auto obtain x where x: "a < x" using gt_ex by auto also from x have "x \ Max {a <..}" by fact also note \Max {a <..} < y\ finally have "y \ Max { a <..}" by fact with \Max {a <..} < y\ show False by auto qed lemma infinite_Ici: "\ finite {a :: 'a :: {no_top, linorder} ..}" using infinite_Ioi[of a] finite_subset[of "{a <..}" "{a ..}"] by (auto simp: subset_eq less_imp_le) subsubsection \Intersection\ context linorder begin lemma Int_atLeastAtMost[simp]: "{a..b} Int {c..d} = {max a c .. min b d}" by auto lemma Int_atLeastAtMostR1[simp]: "{..b} Int {c..d} = {c .. min b d}" by auto lemma Int_atLeastAtMostR2[simp]: "{a..} Int {c..d} = {max a c .. d}" by auto lemma Int_atLeastAtMostL1[simp]: "{a..b} Int {..d} = {a .. min b d}" by auto lemma Int_atLeastAtMostL2[simp]: "{a..b} Int {c..} = {max a c .. b}" by auto lemma Int_atLeastLessThan[simp]: "{a.. {..b} = {.. min a b}" by (auto simp: min_def) lemma Ioc_disjoint: "{a<..b} \ {c<..d} = {} \ b \ a \ d \ c \ b \ c \ d \ a" by auto end context complete_lattice begin lemma shows Sup_atLeast[simp]: "Sup {x ..} = top" and Sup_greaterThanAtLeast[simp]: "x < top \ Sup {x <..} = top" and Sup_atMost[simp]: "Sup {.. y} = y" and Sup_atLeastAtMost[simp]: "x \ y \ Sup { x .. y} = y" and Sup_greaterThanAtMost[simp]: "x < y \ Sup { x <.. y} = y" by (auto intro!: Sup_eqI) lemma shows Inf_atMost[simp]: "Inf {.. x} = bot" and Inf_atMostLessThan[simp]: "top < x \ Inf {..< x} = bot" and Inf_atLeast[simp]: "Inf {x ..} = x" and Inf_atLeastAtMost[simp]: "x \ y \ Inf { x .. y} = x" and Inf_atLeastLessThan[simp]: "x < y \ Inf { x ..< y} = x" by (auto intro!: Inf_eqI) end lemma fixes x y :: "'a :: {complete_lattice, dense_linorder}" shows Sup_lessThan[simp]: "Sup {..< y} = y" and Sup_atLeastLessThan[simp]: "x < y \ Sup { x ..< y} = y" and Sup_greaterThanLessThan[simp]: "x < y \ Sup { x <..< y} = y" and Inf_greaterThan[simp]: "Inf {x <..} = x" and Inf_greaterThanAtMost[simp]: "x < y \ Inf { x <.. y} = x" and Inf_greaterThanLessThan[simp]: "x < y \ Inf { x <..< y} = x" by (auto intro!: Inf_eqI Sup_eqI intro: dense_le dense_le_bounded dense_ge dense_ge_bounded) subsection \Intervals of natural numbers\ subsubsection \The Constant \<^term>\lessThan\\ lemma lessThan_0 [simp]: "lessThan (0::nat) = {}" by (simp add: lessThan_def) lemma lessThan_Suc: "lessThan (Suc k) = insert k (lessThan k)" by (simp add: lessThan_def less_Suc_eq, blast) text \The following proof is convenient in induction proofs where new elements get indices at the beginning. So it is used to transform \<^term>\{.. to \<^term>\0::nat\ and \<^term>\{..< n}\.\ lemma zero_notin_Suc_image [simp]: "0 \ Suc ` A" by auto lemma lessThan_Suc_eq_insert_0: "{..m::nat. lessThan m) = UNIV" by blast subsubsection \The Constant \<^term>\greaterThan\\ lemma greaterThan_0: "greaterThan 0 = range Suc" unfolding greaterThan_def by (blast dest: gr0_conv_Suc [THEN iffD1]) lemma greaterThan_Suc: "greaterThan (Suc k) = greaterThan k - {Suc k}" unfolding greaterThan_def by (auto elim: linorder_neqE) lemma INT_greaterThan_UNIV: "(\m::nat. greaterThan m) = {}" by blast subsubsection \The Constant \<^term>\atLeast\\ lemma atLeast_0 [simp]: "atLeast (0::nat) = UNIV" by (unfold atLeast_def UNIV_def, simp) lemma atLeast_Suc: "atLeast (Suc k) = atLeast k - {k}" unfolding atLeast_def by (auto simp: order_le_less Suc_le_eq) lemma atLeast_Suc_greaterThan: "atLeast (Suc k) = greaterThan k" by (auto simp add: greaterThan_def atLeast_def less_Suc_eq_le) lemma UN_atLeast_UNIV: "(\m::nat. atLeast m) = UNIV" by blast subsubsection \The Constant \<^term>\atMost\\ lemma atMost_0 [simp]: "atMost (0::nat) = {0}" by (simp add: atMost_def) lemma atMost_Suc: "atMost (Suc k) = insert (Suc k) (atMost k)" unfolding atMost_def by (auto simp add: less_Suc_eq order_le_less) lemma UN_atMost_UNIV: "(\m::nat. atMost m) = UNIV" by blast subsubsection \The Constant \<^term>\atLeastLessThan\\ text\The orientation of the following 2 rules is tricky. The lhs is defined in terms of the rhs. Hence the chosen orientation makes sense in this theory --- the reverse orientation complicates proofs (eg nontermination). But outside, when the definition of the lhs is rarely used, the opposite orientation seems preferable because it reduces a specific concept to a more general one.\ lemma atLeast0LessThan [code_abbrev]: "{0::nat..The Constant \<^term>\atLeastAtMost\\ lemma Icc_eq_insert_lb_nat: "m \ n \ {m..n} = insert m {Suc m..n}" by auto lemma atLeast0_atMost_Suc: "{0..Suc n} = insert (Suc n) {0..n}" by (simp add: atLeast0AtMost atMost_Suc) lemma atLeast0_atMost_Suc_eq_insert_0: "{0..Suc n} = insert 0 (Suc ` {0..n})" by (simp add: atLeast0AtMost atMost_Suc_eq_insert_0) subsubsection \Intervals of nats with \<^term>\Suc\\ text\Not a simprule because the RHS is too messy.\ lemma atLeastLessThanSuc: "{m.. n then insert n {m.. Suc n \ {m..Suc n} = insert (Suc n) {m..n}" by auto lemma atLeastAtMost_insertL: "m \ n \ insert m {Suc m..n} = {m ..n}" by auto text \The analogous result is useful on \<^typ>\int\:\ (* here, because we don't have an own int section *) lemma atLeastAtMostPlus1_int_conv: "m \ 1+n \ {m..1+n} = insert (1+n) {m..n::int}" by (auto intro: set_eqI) lemma atLeastLessThan_add_Un: "i \ j \ {i.. {j..Intervals and numerals\ lemma lessThan_nat_numeral: \ \Evaluation for specific numerals\ "lessThan (numeral k :: nat) = insert (pred_numeral k) (lessThan (pred_numeral k))" by (simp add: numeral_eq_Suc lessThan_Suc) lemma atMost_nat_numeral: \ \Evaluation for specific numerals\ "atMost (numeral k :: nat) = insert (numeral k) (atMost (pred_numeral k))" by (simp add: numeral_eq_Suc atMost_Suc) lemma atLeastLessThan_nat_numeral: \ \Evaluation for specific numerals\ "atLeastLessThan m (numeral k :: nat) = (if m \ (pred_numeral k) then insert (pred_numeral k) (atLeastLessThan m (pred_numeral k)) else {})" by (simp add: numeral_eq_Suc atLeastLessThanSuc) subsubsection \Image\ context linordered_semidom begin lemma image_add_atLeast[simp]: "plus k ` {i..} = {k + i..}" proof - have "n = k + (n - k)" if "i + k \ n" for n proof - have "n = (n - (k + i)) + (k + i)" using that by (metis add_commute le_add_diff_inverse) then show "n = k + (n - k)" by (metis local.add_diff_cancel_left' add_assoc add_commute) qed then show ?thesis by (fastforce simp: add_le_imp_le_diff add.commute) qed lemma image_add_atLeastAtMost [simp]: "plus k ` {i..j} = {i + k..j + k}" (is "?A = ?B") proof show "?A \ ?B" by (auto simp add: ac_simps) next show "?B \ ?A" proof fix n assume "n \ ?B" then have "i \ n - k" by (simp add: add_le_imp_le_diff) have "n = n - k + k" proof - from \n \ ?B\ have "n = n - (i + k) + (i + k)" by simp also have "\ = n - k - i + i + k" by (simp add: algebra_simps) also have "\ = n - k + k" using \i \ n - k\ by simp finally show ?thesis . qed moreover have "n - k \ {i..j}" using \n \ ?B\ by (auto simp: add_le_imp_le_diff add_le_add_imp_diff_le) ultimately show "n \ ?A" by (simp add: ac_simps) qed qed lemma image_add_atLeastAtMost' [simp]: "(\n. n + k) ` {i..j} = {i + k..j + k}" by (simp add: add.commute [of _ k]) lemma image_add_atLeastLessThan [simp]: "plus k ` {i..n. n + k) ` {i.. uminus ` {x<..}" by (rule imageI) (simp add: *) thus "y \ uminus ` {x<..}" by simp next fix y assume "y \ -x" have "- (-y) \ uminus ` {x..}" by (rule imageI) (insert \y \ -x\[THEN le_imp_neg_le], simp) thus "y \ uminus ` {x..}" by simp qed simp_all lemma fixes x :: 'a shows image_uminus_lessThan[simp]: "uminus ` {.. = {c - b<..c - a}" by simp finally show ?thesis by simp qed lemma image_minus_const_greaterThanAtMost[simp]: fixes a b c::"'a::linordered_idom" shows "(-) c ` {a<..b} = {c - b.. = {c - b.. = {..c - a}" by simp finally show ?thesis by simp qed lemma image_minus_const_AtMost[simp]: fixes b c::"'a::linordered_idom" shows "(-) c ` {..b} = {c - b..}" proof - have "(-) c ` {..b} = (+) c ` uminus ` {..b}" unfolding image_image by simp also have "\ = {c - b..}" by simp finally show ?thesis by simp qed lemma image_minus_const_atLeastAtMost' [simp]: "(\t. t-d)`{a..b} = {a-d..b-d}" for d::"'a::linordered_idom" by (metis (no_types, lifting) diff_conv_add_uminus image_add_atLeastAtMost' image_cong) context linordered_field begin lemma image_mult_atLeastAtMost [simp]: "((*) d ` {a..b}) = {d*a..d*b}" if "d>0" using that by (auto simp: field_simps mult_le_cancel_right intro: rev_image_eqI [where x="x/d" for x]) lemma image_divide_atLeastAtMost [simp]: "((\c. c / d) ` {a..b}) = {a/d..b/d}" if "d>0" proof - from that have "inverse d > 0" by simp with image_mult_atLeastAtMost [of "inverse d" a b] have "(*) (inverse d) ` {a..b} = {inverse d * a..inverse d * b}" by blast moreover have "(*) (inverse d) = (\c. c / d)" by (simp add: fun_eq_iff field_simps) ultimately show ?thesis by simp qed lemma image_mult_atLeastAtMost_if: "(*) c ` {x .. y} = (if c > 0 then {c * x .. c * y} else if x \ y then {c * y .. c * x} else {})" proof (cases "c = 0 \ x > y") case True then show ?thesis by auto next case False then have "x \ y" by auto from False consider "c < 0"| "c > 0" by (auto simp add: neq_iff) then show ?thesis proof cases case 1 have "(*) c ` {x..y} = {c * y..c * x}" proof (rule set_eqI) fix d from 1 have "inj (\z. z / c)" by (auto intro: injI) then have "d \ (*) c ` {x..y} \ d / c \ (\z. z div c) ` (*) c ` {x..y}" by (subst inj_image_mem_iff) simp_all also have "\ \ d / c \ {x..y}" using 1 by (simp add: image_image) also have "\ \ d \ {c * y..c * x}" by (auto simp add: field_simps 1) finally show "d \ (*) c ` {x..y} \ d \ {c * y..c * x}" . qed with \x \ y\ show ?thesis by auto qed (simp add: mult_left_mono_neg) qed lemma image_mult_atLeastAtMost_if': "(\x. x * c) ` {x..y} = (if x \ y then if c > 0 then {x * c .. y * c} else {y * c .. x * c} else {})" using image_mult_atLeastAtMost_if [of c x y] by (auto simp add: ac_simps) lemma image_affinity_atLeastAtMost: "((\x. m * x + c) ` {a..b}) = (if {a..b} = {} then {} else if 0 \ m then {m * a + c .. m * b + c} else {m * b + c .. m * a + c})" proof - have *: "(\x. m * x + c) = ((\x. x + c) \ (*) m)" by (simp add: fun_eq_iff) show ?thesis by (simp only: * image_comp [symmetric] image_mult_atLeastAtMost_if) (auto simp add: mult_le_cancel_left) qed lemma image_affinity_atLeastAtMost_diff: "((\x. m*x - c) ` {a..b}) = (if {a..b}={} then {} else if 0 \ m then {m*a - c .. m*b - c} else {m*b - c .. m*a - c})" using image_affinity_atLeastAtMost [of m "-c" a b] by simp lemma image_affinity_atLeastAtMost_div: "((\x. x/m + c) ` {a..b}) = (if {a..b}={} then {} else if 0 \ m then {a/m + c .. b/m + c} else {b/m + c .. a/m + c})" using image_affinity_atLeastAtMost [of "inverse m" c a b] by (simp add: field_class.field_divide_inverse algebra_simps inverse_eq_divide) lemma image_affinity_atLeastAtMost_div_diff: "((\x. x/m - c) ` {a..b}) = (if {a..b}={} then {} else if 0 \ m then {a/m - c .. b/m - c} else {b/m - c .. a/m - c})" using image_affinity_atLeastAtMost_diff [of "inverse m" c a b] by (simp add: field_class.field_divide_inverse algebra_simps inverse_eq_divide) end lemma atLeast1_lessThan_eq_remove0: "{Suc 0..x. x + (l::int)) ` {0..i. i - c) ` {x ..< y} = (if c < y then {x - c ..< y - c} else if x < y then {0} else {})" (is "_ = ?right") proof safe fix a assume a: "a \ ?right" show "a \ (\i. i - c) ` {x ..< y}" proof cases assume "c < y" with a show ?thesis by (auto intro!: image_eqI[of _ _ "a + c"]) next assume "\ c < y" with a show ?thesis by (auto intro!: image_eqI[of _ _ x] split: if_split_asm) qed qed auto lemma image_int_atLeastLessThan: "int ` {a..Finiteness\ lemma finite_lessThan [iff]: fixes k :: nat shows "finite {..A bounded set of natural numbers is finite.\ lemma bounded_nat_set_is_finite: "(\i\N. i < (n::nat)) \ finite N" by (rule finite_subset [OF _ finite_lessThan]) auto text \A set of natural numbers is finite iff it is bounded.\ lemma finite_nat_set_iff_bounded: "finite(N::nat set) = (\m. \n\N. n?F\, simplified less_Suc_eq_le[symmetric]] by blast next assume ?B show ?F using \?B\ by(blast intro:bounded_nat_set_is_finite) qed lemma finite_nat_set_iff_bounded_le: "finite(N::nat set) = (\m. \n\N. n\m)" unfolding finite_nat_set_iff_bounded by (blast dest:less_imp_le_nat le_imp_less_Suc) lemma finite_less_ub: "!!f::nat=>nat. (!!n. n \ f n) ==> finite {n. f n \ u}" by (rule_tac B="{..u}" in finite_subset, auto intro: order_trans) lemma bounded_Max_nat: fixes P :: "nat \ bool" assumes x: "P x" and M: "\x. P x \ x \ M" obtains m where "P m" "\x. P x \ x \ m" proof - have "finite {x. P x}" using M finite_nat_set_iff_bounded_le by auto then have "Max {x. P x} \ {x. P x}" using Max_in x by auto then show ?thesis by (simp add: \finite {x. P x}\ that) qed text\Any subset of an interval of natural numbers the size of the subset is exactly that interval.\ lemma subset_card_intvl_is_intvl: assumes "A \ {k.. A" by auto with insert have "A \ {k..Proving Inclusions and Equalities between Unions\ lemma UN_le_eq_Un0: "(\i\n::nat. M i) = (\i\{1..n}. M i) \ M 0" (is "?A = ?B") proof show "?A \ ?B" proof fix x assume "x \ ?A" then obtain i where i: "i\n" "x \ M i" by auto show "x \ ?B" proof(cases i) case 0 with i show ?thesis by simp next case (Suc j) with i show ?thesis by auto qed qed next show "?B \ ?A" by fastforce qed lemma UN_le_add_shift: "(\i\n::nat. M(i+k)) = (\i\{k..n+k}. M i)" (is "?A = ?B") proof show "?A \ ?B" by fastforce next show "?B \ ?A" proof fix x assume "x \ ?B" then obtain i where i: "i \ {k..n+k}" "x \ M(i)" by auto hence "i-k\n \ x \ M((i-k)+k)" by auto thus "x \ ?A" by blast qed qed lemma UN_le_add_shift_strict: "(\ii\{k.. ?A" proof fix x assume "x \ ?B" then obtain i where i: "i \ {k.. M(i)" by auto then have "i - k < n \ x \ M((i-k) + k)" by auto then show "x \ ?A" using UN_le_add_shift by blast qed qed (fastforce) lemma UN_UN_finite_eq: "(\n::nat. \i\{0..n. A n)" by (auto simp add: atLeast0LessThan) lemma UN_finite_subset: "(\n::nat. (\i\{0.. C) \ (\n. A n) \ C" by (subst UN_UN_finite_eq [symmetric]) blast lemma UN_finite2_subset: assumes "\n::nat. (\i\{0.. (\i\{0..n. A n) \ (\n. B n)" proof (rule UN_finite_subset, rule) fix n and a from assms have "(\i\{0.. (\i\{0.. (\i\{0.. (\i\{0.. (\i. B i)" by (auto simp add: UN_UN_finite_eq) qed lemma UN_finite2_eq: "(\n::nat. (\i\{0..i\{0.. (\n. A n) = (\n. B n)" apply (rule subset_antisym [OF UN_finite_subset UN_finite2_subset]) apply auto apply (force simp add: atLeastLessThan_add_Un [of 0])+ done subsubsection \Cardinality\ lemma card_lessThan [simp]: "card {..x. x + l) ` {.. {0.. {0..n}" shows "finite N" using assms finite_atLeastAtMost by (rule finite_subset) lemma ex_bij_betw_nat_finite: "finite M \ \h. bij_betw h {0.. \h. bij_betw h M {0.. finite B \ card A = card B \ \h. bij_betw h A B" apply(drule ex_bij_betw_finite_nat) apply(drule ex_bij_betw_nat_finite) apply(auto intro!:bij_betw_trans) done lemma ex_bij_betw_nat_finite_1: "finite M \ \h. bij_betw h {1 .. card M} M" by (rule finite_same_card_bij) auto lemma bij_betw_iff_card: assumes "finite A" "finite B" shows "(\f. bij_betw f A B) \ (card A = card B)" proof assume "card A = card B" moreover obtain f where "bij_betw f A {0 ..< card A}" using assms ex_bij_betw_finite_nat by blast moreover obtain g where "bij_betw g {0 ..< card B} B" using assms ex_bij_betw_nat_finite by blast ultimately have "bij_betw (g \ f) A B" by (auto simp: bij_betw_trans) thus "(\f. bij_betw f A B)" by blast qed (auto simp: bij_betw_same_card) lemma subset_eq_atLeast0_lessThan_card: fixes n :: nat assumes "N \ {0.. n" proof - from assms finite_lessThan have "card N \ card {0..Relational version of @{thm [source] card_inj_on_le}:\ lemma card_le_if_inj_on_rel: assumes "finite B" "\a. a \ A \ \b. b\B \ r a b" "\a1 a2 b. \ a1 \ A; a2 \ A; b \ B; r a1 b; r a2 b \ \ a1 = a2" shows "card A \ card B" proof - let ?P = "\a b. b \ B \ r a b" let ?f = "\a. SOME b. ?P a b" have 1: "?f ` A \ B" by (auto intro: someI2_ex[OF assms(2)]) have "inj_on ?f A" proof (auto simp: inj_on_def) fix a1 a2 assume asms: "a1 \ A" "a2 \ A" "?f a1 = ?f a2" have 0: "?f a1 \ B" using "1" \a1 \ A\ by blast have 1: "r a1 (?f a1)" using someI_ex[OF assms(2)[OF \a1 \ A\]] by blast have 2: "r a2 (?f a1)" using someI_ex[OF assms(2)[OF \a2 \ A\]] asms(3) by auto show "a1 = a2" using assms(3)[OF asms(1,2) 0 1 2] . qed with 1 show ?thesis using card_inj_on_le[of ?f A B] assms(1) by simp qed subsection \Intervals of integers\ lemma atLeastLessThanPlusOne_atLeastAtMost_int: "{l..Finiteness\ lemma image_atLeastZeroLessThan_int: "0 \ u ==> {(0::int).. u") case True then show ?thesis by (auto simp: image_atLeastZeroLessThan_int) qed auto lemma finite_atLeastLessThan_int [iff]: "finite {l..Cardinality\ lemma card_atLeastZeroLessThan_int: "card {(0::int).. u") case True then show ?thesis by (auto simp: image_atLeastZeroLessThan_int card_image inj_on_def) qed auto lemma card_atLeastLessThan_int [simp]: "card {l.. k < (i::nat)}" proof - have "{k. P k \ k < i} \ {.. M" shows "card {k \ M. k < Suc i} \ 0" proof - from zero_in_M have "{k \ M. k < Suc i} \ {}" by auto with finite_M_bounded_by_nat show ?thesis by (auto simp add: card_eq_0_iff) qed lemma card_less_Suc2: assumes "0 \ M" shows "card {k. Suc k \ M \ k < i} = card {k \ M. k < Suc i}" proof - have *: "\j \ M; j < Suc i\ \ j - Suc 0 < i \ Suc (j - Suc 0) \ M \ Suc 0 \ j" for j by (cases j) (use assms in auto) show ?thesis proof (rule card_bij_eq) show "inj_on Suc {k. Suc k \ M \ k < i}" by force show "inj_on (\x. x - Suc 0) {k \ M. k < Suc i}" by (rule inj_on_diff_nat) (use * in blast) qed (use * in auto) qed lemma card_less_Suc: assumes "0 \ M" shows "Suc (card {k. Suc k \ M \ k < i}) = card {k \ M. k < Suc i}" proof - have "Suc (card {k. Suc k \ M \ k < i}) = Suc (card {k. Suc k \ M - {0} \ k < i})" by simp also have "\ = Suc (card {k \ M - {0}. k < Suc i})" apply (subst card_less_Suc2) using assms by auto also have "\ = Suc (card ({k \ M. k < Suc i} - {0}))" by (force intro: arg_cong [where f=card]) also have "\ = card (insert 0 ({k \ M. k < Suc i} - {0}))" by (simp add: card.insert_remove) also have "... = card {k \ M. k < Suc i}" using assms by (force simp add: intro: arg_cong [where f=card]) finally show ?thesis. qed subsection \Lemmas useful with the summation operator sum\ text \For examples, see Algebra/poly/UnivPoly2.thy\ subsubsection \Disjoint Unions\ text \Singletons and open intervals\ lemma ivl_disj_un_singleton: "{l::'a::linorder} Un {l<..} = {l..}" "{.. {l} Un {l<.. {l<.. u ==> {l} Un {l<..u} = {l..u}" "(l::'a::linorder) \ u ==> {l..One- and two-sided intervals\ lemma ivl_disj_un_one: "(l::'a::linorder) < u ==> {..l} Un {l<.. u ==> {.. u ==> {..l} Un {l<..u} = {..u}" "(l::'a::linorder) \ u ==> {.. u ==> {l<..u} Un {u<..} = {l<..}" "(l::'a::linorder) < u ==> {l<.. u ==> {l..u} Un {u<..} = {l..}" "(l::'a::linorder) \ u ==> {l..Two- and two-sided intervals\ lemma ivl_disj_un_two: "[| (l::'a::linorder) < m; m \ u |] ==> {l<.. m; m < u |] ==> {l<..m} Un {m<.. m; m \ u |] ==> {l.. m; m < u |] ==> {l..m} Un {m<.. u |] ==> {l<.. m; m \ u |] ==> {l<..m} Un {m<..u} = {l<..u}" "[| (l::'a::linorder) \ m; m \ u |] ==> {l.. m; m \ u |] ==> {l..m} Un {m<..u} = {l..u}" by auto lemma ivl_disj_un_two_touch: "[| (l::'a::linorder) < m; m < u |] ==> {l<..m} Un {m.. m; m < u |] ==> {l..m} Un {m.. u |] ==> {l<..m} Un {m..u} = {l<..u}" "[| (l::'a::linorder) \ m; m \ u |] ==> {l..m} Un {m..u} = {l..u}" by auto lemmas ivl_disj_un = ivl_disj_un_singleton ivl_disj_un_one ivl_disj_un_two ivl_disj_un_two_touch subsubsection \Disjoint Intersections\ text \One- and two-sided intervals\ lemma ivl_disj_int_one: "{..l::'a::order} Int {l<..Two- and two-sided intervals\ lemma ivl_disj_int_two: "{l::'a::order<..Some Differences\ lemma ivl_diff[simp]: "i \ n \ {i..Some Subset Conditions\ lemma ivl_subset [simp]: "({i.. {m.. i \ m \ i \ j \ (n::'a::linorder))" using linorder_class.le_less_linear[of i n] apply (auto simp: linorder_not_le) apply (force intro: leI)+ done lemma obtain_subset_with_card_n: assumes "n \ card S" obtains T where "T \ S" "card T = n" "finite T" proof - obtain n' where "card S = n + n'" by (metis assms le_add_diff_inverse) with that show thesis proof (induct n' arbitrary: S) case 0 then show ?case by (cases "finite S") auto next case Suc then show ?case by (simp add: card_Suc_eq) (metis subset_insertI2) qed qed subsection \Generic big monoid operation over intervals\ context semiring_char_0 begin lemma inj_on_of_nat [simp]: "inj_on of_nat N" by rule simp lemma bij_betw_of_nat [simp]: "bij_betw of_nat N A \ of_nat ` N = A" by (simp add: bij_betw_def) end context comm_monoid_set begin lemma atLeastLessThan_reindex: "F g {h m.. h) {m.. h) {m..n}" if "bij_betw h {m..n} {h m..h n}" for m n ::nat proof - from that have "inj_on h {m..n}" and "h ` {m..n} = {h m..h n}" by (simp_all add: bij_betw_def) then show ?thesis using reindex [of h "{m..n}" g] by simp qed lemma atLeastLessThan_shift_bounds: "F g {m + k.. plus k) {m.. plus k) {m..n}" for m n k :: nat using atLeastAtMost_reindex [of "plus k" m n g] by (simp add: ac_simps) lemma atLeast_Suc_lessThan_Suc_shift: "F g {Suc m.. Suc) {m.. Suc) {m..n}" using atLeastAtMost_shift_bounds [of _ _ 1] by (simp add: plus_1_eq_Suc) lemma atLeast_int_lessThan_int_shift: "F g {int m.. int) {m.. int) {m..n}" by (rule atLeastAtMost_reindex) (simp add: image_int_atLeastAtMost) lemma atLeast0_lessThan_Suc: "F g {0..* g n" by (simp add: atLeast0_lessThan_Suc ac_simps) lemma atLeast0_atMost_Suc: "F g {0..Suc n} = F g {0..n} \<^bold>* g (Suc n)" by (simp add: atLeast0_atMost_Suc ac_simps) lemma atLeast0_lessThan_Suc_shift: "F g {0..* F (g \ Suc) {0..* F (g \ Suc) {0..n}" by (simp add: atLeast0_atMost_Suc_eq_insert_0 atLeast_Suc_atMost_Suc_shift) lemma atLeast_Suc_lessThan: "F g {m..* F g {Suc m..* F g {Suc m..n}" if "m \ n" proof - from that have "{m..n} = insert m {Suc m..n}" by auto then show ?thesis by simp qed lemma ivl_cong: "a = c \ b = d \ (\x. c \ x \ x < d \ g x = h x) \ F g {a.. plus m) {0.. n") simp_all lemma atLeastAtMost_shift_0: fixes m n p :: nat assumes "m \ n" shows "F g {m..n} = F (g \ plus m) {0..n - m}" using assms atLeastAtMost_shift_bounds [of g 0 m "n - m"] by simp lemma atLeastLessThan_concat: fixes m n p :: nat shows "m \ n \ n \ p \ F g {m..* F g {n..i. g (m + n - Suc i)) {n..i. g (m + n - i)) {n..m}" by (rule reindex_bij_witness [where i="\i. m + n - i" and j="\i. m + n - i"]) auto lemma atLeastLessThan_rev_at_least_Suc_atMost: "F g {n..i. g (m + n - i)) {Suc n..m}" unfolding atLeastLessThan_rev [of g n m] by (cases m) (simp_all add: atLeast_Suc_atMost_Suc_shift atLeastLessThanSuc_atLeastAtMost) end subsection \Summation indexed over intervals\ syntax (ASCII) "_from_to_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(SUM _ = _.._./ _)" [0,0,0,10] 10) "_from_upto_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(SUM _ = _..<_./ _)" [0,0,0,10] 10) "_upt_sum" :: "idt \ 'a \ 'b \ 'b" ("(SUM _<_./ _)" [0,0,10] 10) "_upto_sum" :: "idt \ 'a \ 'b \ 'b" ("(SUM _<=_./ _)" [0,0,10] 10) syntax (latex_sum output) "_from_to_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\sum_{\_ = _\<^latex>\}^{\_\<^latex>\}$\ _)" [0,0,0,10] 10) "_from_upto_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\sum_{\_ = _\<^latex>\}^{<\_\<^latex>\}$\ _)" [0,0,0,10] 10) "_upt_sum" :: "idt \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\sum_{\_ < _\<^latex>\}$\ _)" [0,0,10] 10) "_upto_sum" :: "idt \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\sum_{\_ \ _\<^latex>\}$\ _)" [0,0,10] 10) syntax "_from_to_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _.._./ _)" [0,0,0,10] 10) "_from_upto_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _..<_./ _)" [0,0,0,10] 10) "_upt_sum" :: "idt \ 'a \ 'b \ 'b" ("(3\_<_./ _)" [0,0,10] 10) "_upto_sum" :: "idt \ 'a \ 'b \ 'b" ("(3\_\_./ _)" [0,0,10] 10) translations "\x=a..b. t" == "CONST sum (\x. t) {a..b}" "\x=a..x. t) {a..i\n. t" == "CONST sum (\i. t) {..n}" "\ii. t) {..The above introduces some pretty alternative syntaxes for summation over intervals: \begin{center} \begin{tabular}{lll} Old & New & \LaTeX\\ @{term[source]"\x\{a..b}. e"} & \<^term>\\x=a..b. e\ & @{term[mode=latex_sum]"\x=a..b. e"}\\ @{term[source]"\x\{a..\\x=a.. & @{term[mode=latex_sum]"\x=a..x\{..b}. e"} & \<^term>\\x\b. e\ & @{term[mode=latex_sum]"\x\b. e"}\\ @{term[source]"\x\{..\\x & @{term[mode=latex_sum]"\xlatex_sum\ (e.g.\ via \mode = latex_sum\ in antiquotations). It is not the default \LaTeX\ output because it only works well with italic-style formulae, not tt-style. Note that for uniformity on \<^typ>\nat\ it is better to use \<^term>\\x::nat=0.. rather than \\x: \sum\ may not provide all lemmas available for \<^term>\{m.. also in the special form for \<^term>\{...\ text\This congruence rule should be used for sums over intervals as the standard theorem @{text[source]sum.cong} does not work well with the simplifier who adds the unsimplified premise \<^term>\x\B\ to the context.\ context comm_monoid_set begin lemma zero_middle: assumes "1 \ p" "k \ p" shows "F (\j. if j < k then g j else if j = k then \<^bold>1 else h (j - Suc 0)) {..p} = F (\j. if j < k then g j else h j) {..p - Suc 0}" (is "?lhs = ?rhs") proof - have [simp]: "{..p - Suc 0} \ {j. j < k} = {.. - {j. j < k} = {k..p - Suc 0}" using assms by auto have "?lhs = F g {..* F (\j. if j = k then \<^bold>1 else h (j - Suc 0)) {k..p}" using union_disjoint [of "{.. = F g {..* F (\j. h (j - Suc 0)) {Suc k..p}" by (simp add: atLeast_Suc_atMost [of k p] assms) also have "\ = F g {..* F h {k .. p - Suc 0}" using reindex [of Suc "{k..p - Suc 0}"] assms by simp also have "\ = ?rhs" by (simp add: If_cases) finally show ?thesis . qed lemma atMost_Suc [simp]: "F g {..Suc n} = F g {..n} \<^bold>* g (Suc n)" by (simp add: atMost_Suc ac_simps) lemma lessThan_Suc [simp]: "F g {..* g n" by (simp add: lessThan_Suc ac_simps) lemma cl_ivl_Suc [simp]: "F g {m..Suc n} = (if Suc n < m then \<^bold>1 else F g {m..n} \<^bold>* g(Suc n))" by (auto simp: ac_simps atLeastAtMostSuc_conv) lemma op_ivl_Suc [simp]: "F g {m..1 else F g {m..* g(n))" by (auto simp: ac_simps atLeastLessThanSuc) lemma head: fixes n :: nat assumes mn: "m \ n" shows "F g {m..n} = g m \<^bold>* F g {m<..n}" (is "?lhs = ?rhs") proof - from mn have "{m..n} = {m} \ {m<..n}" by (auto intro: ivl_disj_un_singleton) hence "?lhs = F g ({m} \ {m<..n})" by (simp add: atLeast0LessThan) also have "\ = ?rhs" by simp finally show ?thesis . qed +lemma last_plus: + fixes n::nat shows "m \ n \ F g {m..n} = g n \<^bold>* F g {m..1 else F g {m..* g(n))" + by (simp add: commute last_plus) + lemma ub_add_nat: assumes "(m::nat) \ n + 1" shows "F g {m..n + p} = F g {m..n} \<^bold>* F g {n + 1..n + p}" proof- have "{m .. n+p} = {m..n} \ {n+1..n+p}" using \m \ n+1\ by auto thus ?thesis by (auto simp: ivl_disj_int union_disjoint atLeastSucAtMost_greaterThanAtMost) qed lemma nat_group: fixes k::nat shows "F (\m. F g {m * k ..< m*k + k}) {.. 0" by auto then show ?thesis by (induct n) (simp_all add: atLeastLessThan_concat add.commute atLeast0LessThan[symmetric]) qed auto lemma triangle_reindex: fixes n :: nat shows "F (\(i,j). g i j) {(i,j). i+j < n} = F (\k. F (\i. g i (k - i)) {..k}) {..(i,j). g i j) {(i,j). i+j \ n} = F (\k. F (\i. g i (k - i)) {..k}) {..n}" using triangle_reindex [of g "Suc n"] by (simp only: Nat.less_Suc_eq_le lessThan_Suc_atMost) lemma nat_diff_reindex: "F (\i. g (n - Suc i)) {..i. g(i + k)){m..i. g(i + k)){m..n::nat}" by (rule reindex_bij_witness[where i="\i. i + k" and j="\i. i - k"]) auto corollary shift_bounds_cl_Suc_ivl: "F g {Suc m..Suc n} = F (\i. g(Suc i)){m..n}" by (simp add: shift_bounds_cl_nat_ivl[where k="Suc 0", simplified]) corollary Suc_reindex_ivl: "m \ n \ F g {m..n} \<^bold>* g (Suc n) = g m \<^bold>* F (\i. g (Suc i)) {m..n}" by (simp add: assoc atLeast_Suc_atMost flip: shift_bounds_cl_Suc_ivl) corollary shift_bounds_Suc_ivl: "F g {Suc m..i. g(Suc i)){m..* F (\i. g (Suc i)) {..n}" proof (induct n) case 0 show ?case by simp next case (Suc n) note IH = this have "F g {..Suc (Suc n)} = F g {..Suc n} \<^bold>* g (Suc (Suc n))" by (rule atMost_Suc) also have "F g {..Suc n} = g 0 \<^bold>* F (\i. g (Suc i)) {..n}" by (rule IH) also have "g 0 \<^bold>* F (\i. g (Suc i)) {..n} \<^bold>* g (Suc (Suc n)) = g 0 \<^bold>* (F (\i. g (Suc i)) {..n} \<^bold>* g (Suc (Suc n)))" by (rule assoc) also have "F (\i. g (Suc i)) {..n} \<^bold>* g (Suc (Suc n)) = F (\i. g (Suc i)) {..Suc n}" by (rule atMost_Suc [symmetric]) finally show ?case . qed lemma lessThan_Suc_shift: "F g {..* F (\i. g (Suc i)) {..* F (\i. g (Suc i)) {.. n \ F g {m..n} = g n \<^bold>* F g {m..i. F (\j. a i j) {0..j. F (\i. a i j) {Suc j..n}) {0..i. F (\j. a i j) {..j. F (\i. a i j) {Suc j..n}) {..k. g (Suc k)) {.. = F (\k. g (Suc k)) {.. b \ F g {a..* g b" by (simp add: atLeastLessThanSuc commute) lemma nat_ivl_Suc': assumes "m \ Suc n" shows "F g {m..Suc n} = g (Suc n) \<^bold>* F g {m..n}" proof - from assms have "{m..Suc n} = insert (Suc n) {m..n}" by auto also have "F g \ = g (Suc n) \<^bold>* F g {m..n}" by simp finally show ?thesis . qed lemma in_pairs: "F g {2*m..Suc(2*n)} = F (\i. g(2*i) \<^bold>* g(Suc(2*i))) {m..n}" proof (induction n) case 0 show ?case by (cases "m=0") auto next case (Suc n) then show ?case by (auto simp: assoc split: if_split_asm) qed lemma in_pairs_0: "F g {..Suc(2*n)} = F (\i. g(2*i) \<^bold>* g(Suc(2*i))) {..n}" using in_pairs [of _ 0 n] by (simp add: atLeast0AtMost) end lemma sum_natinterval_diff: fixes f:: "nat \ ('a::ab_group_add)" shows "sum (\k. f k - f(k + 1)) {(m::nat) .. n} = (if m \ n then f m - f(n + 1) else 0)" by (induct n, auto simp add: algebra_simps not_le le_Suc_eq) lemma sum_diff_nat_ivl: fixes f :: "nat \ 'a::ab_group_add" shows "\ m \ n; n \ p \ \ sum f {m..x. Q x \ P x \ (\xxxk = 0..k = 0..k = Suc 0..k = Suc 0..k = 0..Shifting bounds\ context comm_monoid_add begin context fixes f :: "nat \ 'a" assumes "f 0 = 0" begin lemma sum_shift_lb_Suc0_0_upt: "sum f {Suc 0..f 0 = 0\ by simp qed lemma sum_shift_lb_Suc0_0: "sum f {Suc 0..k} = sum f {0..k}" proof (cases k) case 0 with \f 0 = 0\ show ?thesis by simp next case (Suc k) moreover have "{0..Suc k} = insert 0 {Suc 0..Suc k}" by auto ultimately show ?thesis using \f 0 = 0\ by simp qed end end lemma sum_Suc_diff: fixes f :: "nat \ 'a::ab_group_add" assumes "m \ Suc n" shows "(\i = m..n. f(Suc i) - f i) = f (Suc n) - f m" using assms by (induct n) (auto simp: le_Suc_eq) lemma sum_Suc_diff': fixes f :: "nat \ 'a::ab_group_add" assumes "m \ n" shows "(\i = m..Telescoping\ lemma sum_telescope: fixes f::"nat \ 'a::ab_group_add" shows "sum (\i. f i - f (Suc i)) {.. i} = f 0 - f (Suc i)" by (induct i) simp_all lemma sum_telescope'': assumes "m \ n" shows "(\k\{Suc m..n}. f k - f (k - 1)) = f n - (f m :: 'a :: ab_group_add)" by (rule dec_induct[OF assms]) (simp_all add: algebra_simps) lemma sum_lessThan_telescope: "(\nnThe formula for geometric sums\ lemma sum_power2: "(\i=0.. 1" shows "(\i 0" by simp_all moreover have "(\iy \ 0\) ultimately show ?thesis by simp qed lemma diff_power_eq_sum: fixes y :: "'a::{comm_ring,monoid_mult}" shows "x ^ (Suc n) - y ^ (Suc n) = (x - y) * (\pppp \\COMPLEX_POLYFUN\ in HOL Light\ fixes x :: "'a::{comm_ring,monoid_mult}" shows "x^n - y^n = (x - y) * (\iiiii\n. x^i) = 1 - x^Suc n" by (simp only: one_diff_power_eq lessThan_Suc_atMost) lemma sum_power_shift: fixes x :: "'a::{comm_ring,monoid_mult}" assumes "m \ n" shows "(\i=m..n. x^i) = x^m * (\i\n-m. x^i)" proof - have "(\i=m..n. x^i) = x^m * (\i=m..n. x^(i-m))" by (simp add: sum_distrib_left power_add [symmetric]) also have "(\i=m..n. x^(i-m)) = (\i\n-m. x^i)" using \m \ n\ by (intro sum.reindex_bij_witness[where j="\i. i - m" and i="\i. i + m"]) auto finally show ?thesis . qed lemma sum_gp_multiplied: fixes x :: "'a::{comm_ring,monoid_mult}" assumes "m \ n" shows "(1 - x) * (\i=m..n. x^i) = x^m - x^Suc n" proof - have "(1 - x) * (\i=m..n. x^i) = x^m * (1 - x) * (\i\n-m. x^i)" by (metis mult.assoc mult.commute assms sum_power_shift) also have "... =x^m * (1 - x^Suc(n-m))" by (metis mult.assoc sum_gp_basic) also have "... = x^m - x^Suc n" using assms by (simp add: algebra_simps) (metis le_add_diff_inverse power_add) finally show ?thesis . qed lemma sum_gp: fixes x :: "'a::{comm_ring,division_ring}" shows "(\i=m..n. x^i) = (if n < m then 0 else if x = 1 then of_nat((n + 1) - m) else (x^m - x^Suc n) / (1 - x))" using sum_gp_multiplied [of m n x] apply auto by (metis eq_iff_diff_eq_0 mult.commute nonzero_divide_eq_eq) subsubsection\Geometric progressions\ lemma sum_gp0: fixes x :: "'a::{comm_ring,division_ring}" shows "(\i\n. x^i) = (if x = 1 then of_nat(n + 1) else (1 - x^Suc n) / (1 - x))" using sum_gp_basic[of x n] by (simp add: mult.commute field_split_simps) lemma sum_power_add: fixes x :: "'a::{comm_ring,monoid_mult}" shows "(\i\I. x^(m+i)) = x^m * (\i\I. x^i)" by (simp add: sum_distrib_left power_add) lemma sum_gp_offset: fixes x :: "'a::{comm_ring,division_ring}" shows "(\i=m..m+n. x^i) = (if x = 1 then of_nat n + 1 else x^m * (1 - x^Suc n) / (1 - x))" using sum_gp [of x m "m+n"] by (auto simp: power_add algebra_simps) lemma sum_gp_strict: fixes x :: "'a::{comm_ring,division_ring}" shows "(\iThe formulae for arithmetic sums\ context comm_semiring_1 begin lemma double_gauss_sum: "2 * (\i = 0..n. of_nat i) = of_nat n * (of_nat n + 1)" by (induct n) (simp_all add: sum.atLeast0_atMost_Suc algebra_simps left_add_twice) lemma double_gauss_sum_from_Suc_0: "2 * (\i = Suc 0..n. of_nat i) = of_nat n * (of_nat n + 1)" proof - have "sum of_nat {Suc 0..n} = sum of_nat (insert 0 {Suc 0..n})" by simp also have "\ = sum of_nat {0..n}" by (cases n) (simp_all add: atLeast0_atMost_Suc_eq_insert_0) finally show ?thesis by (simp add: double_gauss_sum) qed lemma double_arith_series: "2 * (\i = 0..n. a + of_nat i * d) = (of_nat n + 1) * (2 * a + of_nat n * d)" proof - have "(\i = 0..n. a + of_nat i * d) = ((\i = 0..n. a) + (\i = 0..n. of_nat i * d))" by (rule sum.distrib) also have "\ = (of_nat (Suc n) * a + d * (\i = 0..n. of_nat i))" by (simp add: sum_distrib_left algebra_simps) finally show ?thesis by (simp add: algebra_simps double_gauss_sum left_add_twice) qed end context unique_euclidean_semiring_with_nat begin lemma gauss_sum: "(\i = 0..n. of_nat i) = of_nat n * (of_nat n + 1) div 2" using double_gauss_sum [of n, symmetric] by simp lemma gauss_sum_from_Suc_0: "(\i = Suc 0..n. of_nat i) = of_nat n * (of_nat n + 1) div 2" using double_gauss_sum_from_Suc_0 [of n, symmetric] by simp lemma arith_series: "(\i = 0..n. a + of_nat i * d) = (of_nat n + 1) * (2 * a + of_nat n * d) div 2" using double_arith_series [of a d n, symmetric] by simp end lemma gauss_sum_nat: "\{0..n} = (n * Suc n) div 2" using gauss_sum [of n, where ?'a = nat] by simp lemma arith_series_nat: "(\i = 0..n. a + i * d) = Suc n * (2 * a + n * d) div 2" using arith_series [of a d n] by simp lemma Sum_Icc_int: "\{m..n} = (n * (n + 1) - m * (m - 1)) div 2" if "m \ n" for m n :: int using that proof (induct i \ "nat (n - m)" arbitrary: m n) case 0 then have "m = n" by arith then show ?case by (simp add: algebra_simps mult_2 [symmetric]) next case (Suc i) have 0: "i = nat((n-1) - m)" "m \ n-1" using Suc(2,3) by arith+ have "\ {m..n} = \ {m..1+(n-1)}" by simp also have "\ = \ {m..n-1} + n" using \m \ n\ by(subst atLeastAtMostPlus1_int_conv) simp_all also have "\ = ((n-1)*(n-1+1) - m*(m-1)) div 2 + n" by(simp add: Suc(1)[OF 0]) also have "\ = ((n-1)*(n-1+1) - m*(m-1) + 2*n) div 2" by simp also have "\ = (n*(n+1) - m*(m-1)) div 2" by (simp add: algebra_simps mult_2_right) finally show ?case . qed lemma Sum_Icc_nat: "\{m..n} = (n * (n + 1) - m * (m - 1)) div 2" for m n :: nat proof (cases "m \ n") case True then have *: "m * (m - 1) \ n * (n + 1)" by (meson diff_le_self order_trans le_add1 mult_le_mono) have "int (\{m..n}) = (\{int m..int n})" by (simp add: sum.atLeast_int_atMost_int_shift) also have "\ = (int n * (int n + 1) - int m * (int m - 1)) div 2" using \m \ n\ by (simp add: Sum_Icc_int) also have "\ = int ((n * (n + 1) - m * (m - 1)) div 2)" using le_square * by (simp add: algebra_simps of_nat_div of_nat_diff) finally show ?thesis by (simp only: of_nat_eq_iff) next case False then show ?thesis by (auto dest: less_imp_Suc_add simp add: not_le algebra_simps) qed lemma Sum_Ico_nat: "\{m..Division remainder\ lemma range_mod: fixes n :: nat assumes "n > 0" shows "range (\m. m mod n) = {0.. ?A \ m \ ?B" proof assume "m \ ?A" with assms show "m \ ?B" by auto next assume "m \ ?B" moreover have "m mod n \ ?A" by (rule rangeI) ultimately show "m \ ?A" by simp qed qed subsection \Products indexed over intervals\ syntax (ASCII) "_from_to_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(PROD _ = _.._./ _)" [0,0,0,10] 10) "_from_upto_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(PROD _ = _..<_./ _)" [0,0,0,10] 10) "_upt_prod" :: "idt \ 'a \ 'b \ 'b" ("(PROD _<_./ _)" [0,0,10] 10) "_upto_prod" :: "idt \ 'a \ 'b \ 'b" ("(PROD _<=_./ _)" [0,0,10] 10) syntax (latex_prod output) "_from_to_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\prod_{\_ = _\<^latex>\}^{\_\<^latex>\}$\ _)" [0,0,0,10] 10) "_from_upto_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\prod_{\_ = _\<^latex>\}^{<\_\<^latex>\}$\ _)" [0,0,0,10] 10) "_upt_prod" :: "idt \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\prod_{\_ < _\<^latex>\}$\ _)" [0,0,10] 10) "_upto_prod" :: "idt \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\prod_{\_ \ _\<^latex>\}$\ _)" [0,0,10] 10) syntax "_from_to_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _.._./ _)" [0,0,0,10] 10) "_from_upto_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _..<_./ _)" [0,0,0,10] 10) "_upt_prod" :: "idt \ 'a \ 'b \ 'b" ("(3\_<_./ _)" [0,0,10] 10) "_upto_prod" :: "idt \ 'a \ 'b \ 'b" ("(3\_\_./ _)" [0,0,10] 10) translations "\x=a..b. t" \ "CONST prod (\x. t) {a..b}" "\x=a.. "CONST prod (\x. t) {a..i\n. t" \ "CONST prod (\i. t) {..n}" "\i "CONST prod (\i. t) {..{int i..int (i+j)}" by (induct j) (auto simp add: atLeastAtMostSuc_conv atLeastAtMostPlus1_int_conv) lemma prod_int_eq: "prod int {i..j} = \{int i..int j}" proof (cases "i \ j") case True then show ?thesis by (metis le_iff_add prod_int_plus_eq) next case False then show ?thesis by auto qed subsection \Efficient folding over intervals\ function fold_atLeastAtMost_nat where [simp del]: "fold_atLeastAtMost_nat f a (b::nat) acc = (if a > b then acc else fold_atLeastAtMost_nat f (a+1) b (f a acc))" by pat_completeness auto termination by (relation "measure (\(_,a,b,_). Suc b - a)") auto lemma fold_atLeastAtMost_nat: assumes "comp_fun_commute f" shows "fold_atLeastAtMost_nat f a b acc = Finite_Set.fold f acc {a..b}" using assms proof (induction f a b acc rule: fold_atLeastAtMost_nat.induct, goal_cases) case (1 f a b acc) interpret comp_fun_commute f by fact show ?case proof (cases "a > b") case True thus ?thesis by (subst fold_atLeastAtMost_nat.simps) auto next case False with 1 show ?thesis by (subst fold_atLeastAtMost_nat.simps) (auto simp: atLeastAtMost_insertL[symmetric] fold_fun_left_comm) qed qed lemma sum_atLeastAtMost_code: "sum f {a..b} = fold_atLeastAtMost_nat (\a acc. f a + acc) a b 0" proof - have "comp_fun_commute (\a. (+) (f a))" by unfold_locales (auto simp: o_def add_ac) thus ?thesis by (simp add: sum.eq_fold fold_atLeastAtMost_nat o_def) qed lemma prod_atLeastAtMost_code: "prod f {a..b} = fold_atLeastAtMost_nat (\a acc. f a * acc) a b 1" proof - have "comp_fun_commute (\a. (*) (f a))" by unfold_locales (auto simp: o_def mult_ac) thus ?thesis by (simp add: prod.eq_fold fold_atLeastAtMost_nat o_def) qed (* TODO: Add support for folding over more kinds of intervals here *) end