# HG changeset patch # User paulson # Date 1528035750 -3600 # Node ID 20375f232f3bfb7ebcea2e7efa2517cb2905f30d # Parent 0f19c98fa7be2a4b76fc5430621cc27cdccdb64b infinite product material diff -r 0f19c98fa7be -r 20375f232f3b src/HOL/Analysis/Brouwer_Fixpoint.thy --- a/src/HOL/Analysis/Brouwer_Fixpoint.thy Sat Jun 02 22:57:44 2018 +0100 +++ b/src/HOL/Analysis/Brouwer_Fixpoint.thy Sun Jun 03 15:22:30 2018 +0100 @@ -40,9 +40,7 @@ lemmas swap_apply1 = swap_apply(1) lemmas swap_apply2 = swap_apply(2) -lemmas lessThan_empty_iff = Iio_eq_empty_iff_nat lemmas Zero_notin_Suc = zero_notin_Suc_image -lemmas atMost_Suc_eq_insert_0 = Iic_Suc_eq_insert_0 lemma pointwise_minimal_pointwise_maximal: fixes s :: "(nat \ nat) set" diff -r 0f19c98fa7be -r 20375f232f3b src/HOL/Analysis/Henstock_Kurzweil_Integration.thy --- a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Sat Jun 02 22:57:44 2018 +0100 +++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Sun Jun 03 15:22:30 2018 +0100 @@ -3925,55 +3925,42 @@ subsection \Stronger form with finite number of exceptional points\ lemma fundamental_theorem_of_calculus_interior_strong: - fixes f :: "real \ 'a::banach" - assumes "finite s" - and "a \ b" - and "continuous_on {a..b} f" - and "\x\{a <..< b} - s. (f has_vector_derivative f'(x)) (at x)" - shows "(f' has_integral (f b - f a)) {a..b}" + fixes f :: "real \ 'a::banach" + assumes "finite S" + and "a \ b" "\x. x \ {a <..< b} - S \ (f has_vector_derivative f'(x)) (at x)" + and "continuous_on {a .. b} f" + shows "(f' has_integral (f b - f a)) {a .. b}" using assms -proof (induct "card s" arbitrary: s a b) - case 0 +proof (induction arbitrary: a b) +case empty then show ?case - by (force simp add: intro: fundamental_theorem_of_calculus_interior) + using fundamental_theorem_of_calculus_interior by force next - case (Suc n) - then obtain c s' where cs: "s = insert c s'" and n: "n = card s'" - by (metis card_eq_SucD) - then have "finite s'" - using \finite s\ by force +case (insert x S) show ?case - proof (cases "c \ box a b") - case False - with \finite s'\ show ?thesis - using cs n Suc - by (metis Diff_iff box_real(1) insert_iff) + proof (cases "x \ {a<.. c" "c \ b" - by (auto simp: mem_box) - moreover have "?P a c" "?P c b" - using Suc.prems(4) True \a \ c\ cs(1) by auto - moreover have "continuous_on {a..c} f" "continuous_on {c..b} f" - using \continuous_on {a..b} f\ \a \ c\ \c \ b\ continuous_on_subset by fastforce+ - ultimately have "(f' has_integral f c - f a + (f b - f c)) {a..b}" - using Suc.hyps(1) \finite s'\ \n = card s'\ by (blast intro: has_integral_combine) - then show ?thesis - by auto + case True then have "a < x" "x < b" + by auto + have "(f' has_integral f x - f a) {a..x}" "(f' has_integral f b - f x) {x..b}" + using \continuous_on {a..b} f\ \a < x\ \x < b\ continuous_on_subset by (force simp: intro!: insert)+ + then have "(f' has_integral f x - f a + (f b - f x)) {a..b}" + using \a < x\ \x < b\ has_integral_combine less_imp_le by blast + then show ?thesis + by simp qed qed corollary fundamental_theorem_of_calculus_strong: fixes f :: "real \ 'a::banach" - assumes "finite s" + assumes "finite S" and "a \ b" + and vec: "\x. x \ {a..b} - S \ (f has_vector_derivative f'(x)) (at x)" and "continuous_on {a..b} f" - and vec: "\x\{a..b} - s. (f has_vector_derivative f'(x)) (at x)" shows "(f' has_integral (f b - f a)) {a..b}" - apply (rule fundamental_theorem_of_calculus_interior_strong[OF assms(1-3), of f']) - using vec apply (auto simp: mem_box) - done + by (rule fundamental_theorem_of_calculus_interior_strong [OF \finite S\]) (force simp: assms)+ proposition indefinite_integral_continuous_left: fixes f:: "real \ 'a::banach" @@ -7242,7 +7229,6 @@ define f where "f = (\k x. if x \ {inverse (of_nat (Suc k))..c} then x powr a else 0)" define F where "F = (\k. if inverse (of_nat (Suc k)) \ c then c powr (a+1)/(a+1) - inverse (real (Suc k)) powr (a+1)/(a+1) else 0)" - { fix k :: nat have "(f k has_integral F k) {0..c}" @@ -7448,8 +7434,8 @@ lemma integrable_on_subinterval: fixes f :: "'n::ordered_euclidean_space \ 'a::banach" - assumes "f integrable_on s" "{a..b} \ s" + assumes "f integrable_on S" "{a..b} \ S" shows "f integrable_on {a..b}" - using integrable_on_subcbox[of f s a b] assms by (simp add: cbox_interval) + using integrable_on_subcbox[of f S a b] assms by (simp add: cbox_interval) end diff -r 0f19c98fa7be -r 20375f232f3b src/HOL/Analysis/Infinite_Products.thy --- a/src/HOL/Analysis/Infinite_Products.thy Sat Jun 02 22:57:44 2018 +0100 +++ b/src/HOL/Analysis/Infinite_Products.thy Sun Jun 03 15:22:30 2018 +0100 @@ -50,21 +50,21 @@ by (rule tendsto_eq_intros refl | simp)+ qed auto -definition gen_has_prod :: "[nat \ 'a::{t2_space, comm_semiring_1}, nat, 'a] \ bool" - where "gen_has_prod f M p \ (\n. \i\n. f (i+M)) \ p \ p \ 0" +definition raw_has_prod :: "[nat \ 'a::{t2_space, comm_semiring_1}, nat, 'a] \ bool" + where "raw_has_prod f M p \ (\n. \i\n. f (i+M)) \ p \ p \ 0" text\The nonzero and zero cases, as in \emph{Complex Analysis} by Joseph Bak and Donald J.Newman, page 241\ definition has_prod :: "(nat \ 'a::{t2_space, comm_semiring_1}) \ 'a \ bool" (infixr "has'_prod" 80) - where "f has_prod p \ gen_has_prod f 0 p \ (\i q. p = 0 \ f i = 0 \ gen_has_prod f (Suc i) q)" + where "f has_prod p \ raw_has_prod f 0 p \ (\i q. p = 0 \ f i = 0 \ raw_has_prod f (Suc i) q)" definition convergent_prod :: "(nat \ 'a :: {t2_space,comm_semiring_1}) \ bool" where - "convergent_prod f \ \M p. gen_has_prod f M p" + "convergent_prod f \ \M p. raw_has_prod f M p" definition prodinf :: "(nat \ 'a::{t2_space, comm_semiring_1}) \ 'a" (binder "\" 10) where "prodinf f = (THE p. f has_prod p)" -lemmas prod_defs = gen_has_prod_def has_prod_def convergent_prod_def prodinf_def +lemmas prod_defs = raw_has_prod_def has_prod_def convergent_prod_def prodinf_def lemma has_prod_subst[trans]: "f = g \ g has_prod z \ f has_prod z" by simp @@ -72,34 +72,39 @@ lemma has_prod_cong: "(\n. f n = g n) \ f has_prod c \ g has_prod c" by presburger -lemma gen_has_prod_nonzero [simp]: "\ gen_has_prod f M 0" - by (simp add: gen_has_prod_def) +lemma raw_has_prod_nonzero [simp]: "\ raw_has_prod f M 0" + by (simp add: raw_has_prod_def) -lemma gen_has_prod_eq_0: - fixes f :: "nat \ 'a::{idom,t2_space}" - assumes p: "gen_has_prod f m p" and i: "f i = 0" "i \ m" +lemma raw_has_prod_eq_0: + fixes f :: "nat \ 'a::{semidom,t2_space}" + assumes p: "raw_has_prod f m p" and i: "f i = 0" "i \ m" shows "p = 0" proof - have eq0: "(\k\n. f (k+m)) = 0" if "i - m \ n" for n - by (metis i that atMost_atLeast0 atMost_iff diff_add finite_atLeastAtMost prod_zero_iff) + proof - + have "\k\n. f (k + m) = 0" + using i that by auto + then show ?thesis + by auto + qed have "(\n. \i\n. f (i + m)) \ 0" by (rule LIMSEQ_offset [where k = "i-m"]) (simp add: eq0) with p show ?thesis - unfolding gen_has_prod_def + unfolding raw_has_prod_def using LIMSEQ_unique by blast qed -lemma has_prod_0_iff: "f has_prod 0 \ (\i. f i = 0 \ (\p. gen_has_prod f (Suc i) p))" +lemma has_prod_0_iff: "f has_prod 0 \ (\i. f i = 0 \ (\p. raw_has_prod f (Suc i) p))" by (simp add: has_prod_def) lemma has_prod_unique2: - fixes f :: "nat \ 'a::{idom,t2_space}" + fixes f :: "nat \ 'a::{semidom,t2_space}" assumes "f has_prod a" "f has_prod b" shows "a = b" using assms - by (auto simp: has_prod_def gen_has_prod_eq_0) (meson gen_has_prod_def sequentially_bot tendsto_unique) + by (auto simp: has_prod_def raw_has_prod_eq_0) (meson raw_has_prod_def sequentially_bot tendsto_unique) lemma has_prod_unique: - fixes f :: "nat \ 'a :: {idom,t2_space}" + fixes f :: "nat \ 'a :: {semidom,t2_space}" shows "f has_prod s \ s = prodinf f" by (simp add: has_prod_unique2 prodinf_def the_equality) @@ -378,42 +383,55 @@ shows "abs_convergent_prod f" using assms unfolding abs_convergent_prod_def by (rule convergent_prod_offset) -lemma convergent_prod_ignore_initial_segment: - fixes f :: "nat \ 'a :: {real_normed_field}" - assumes "convergent_prod f" - shows "convergent_prod (\n. f (n + m))" +lemma raw_has_prod_ignore_initial_segment: + fixes f :: "nat \ 'a :: real_normed_field" + assumes "raw_has_prod f M p" "N \ M" + obtains q where "raw_has_prod f N q" proof - - from assms obtain M L - where L: "(\n. \k\n. f (k + M)) \ L" "L \ 0" and nz: "\n. n \ M \ f n \ 0" - by (auto simp: convergent_prod_altdef) - define C where "C = (\kn. \k\n. f (k + M)) \ p" and "p \ 0" + using assms by (auto simp: raw_has_prod_def) + then have nz: "\n. n \ M \ f n \ 0" + using assms by (auto simp: raw_has_prod_eq_0) + define C where "C = (\k 0" by (auto simp: C_def) - from L(1) have "(\n. \k\n+m. f (k + M)) \ L" + from p have "(\i. \k\i + (N-M). f (k + M)) \ p" by (rule LIMSEQ_ignore_initial_segment) - also have "(\n. \k\n+m. f (k + M)) = (\n. C * (\k\n. f (k + M + m)))" + also have "(\i. \k\i + (N-M). f (k + M)) = (\n. C * (\k\n. f (k + N)))" proof (rule ext, goal_cases) case (1 n) - have "{..n+m} = {.. {m..n+m}" by auto - also have "(\k\\. f (k + M)) = C * (\k=m..n+m. f (k + M))" + have "{..n+(N-M)} = {..<(N-M)} \ {(N-M)..n+(N-M)}" by auto + also have "(\k\\. f (k + M)) = C * (\k=(N-M)..n+(N-M). f (k + M))" unfolding C_def by (rule prod.union_disjoint) auto - also have "(\k=m..n+m. f (k + M)) = (\k\n. f (k + m + M))" - by (intro ext prod.reindex_bij_witness[of _ "\k. k + m" "\k. k - m"]) auto - finally show ?case by (simp add: add_ac) + also have "(\k=(N-M)..n+(N-M). f (k + M)) = (\k\n. f (k + (N-M) + M))" + by (intro ext prod.reindex_bij_witness[of _ "\k. k + (N-M)" "\k. k - (N-M)"]) auto + finally show ?case + using \N \ M\ by (simp add: add_ac) qed - finally have "(\n. C * (\k\n. f (k + M + m)) / C) \ L / C" + finally have "(\n. C * (\k\n. f (k + N)) / C) \ p / C" by (intro tendsto_divide tendsto_const) auto - hence "(\n. \k\n. f (k + M + m)) \ L / C" by simp - moreover from \L \ 0\ have "L / C \ 0" by simp - ultimately show ?thesis - unfolding prod_defs by blast + hence "(\n. \k\n. f (k + N)) \ p / C" by simp + moreover from \p \ 0\ have "p / C \ 0" by simp + ultimately show ?thesis + using raw_has_prod_def that by blast qed +corollary convergent_prod_ignore_initial_segment: + fixes f :: "nat \ 'a :: real_normed_field" + assumes "convergent_prod f" + shows "convergent_prod (\n. f (n + m))" + using assms + unfolding convergent_prod_def + apply clarify + apply (erule_tac N="M+m" in raw_has_prod_ignore_initial_segment) + apply (auto simp add: raw_has_prod_def add_ac) + done + corollary convergent_prod_ignore_nonzero_segment: fixes f :: "nat \ 'a :: real_normed_field" assumes f: "convergent_prod f" and nz: "\i. i \ M \ f i \ 0" - shows "\p. gen_has_prod f M p" + shows "\p. raw_has_prod f M p" using convergent_prod_ignore_initial_segment [OF f] by (metis convergent_LIMSEQ_iff convergent_prod_iff_convergent le_add_same_cancel2 nz prod_defs(1) zero_order(1)) @@ -551,13 +569,13 @@ with L show ?thesis by (auto simp: prod_defs) qed -lemma gen_has_prod_cases: +lemma raw_has_prod_cases: fixes f :: "nat \ 'a :: {idom,topological_semigroup_mult,t2_space}" - assumes "gen_has_prod f M p" - obtains i where "in. \i\n. f (i + M)) \ p" "p \ 0" - using assms unfolding gen_has_prod_def by blast+ + using assms unfolding raw_has_prod_def by blast+ then have "(\n. prod f {..i\n. f (i + M))) \ prod f {..i\n. f (i + M)) = prod f {..n+M}" for n @@ -572,8 +590,8 @@ qed ultimately have "(\n. prod f {..n}) \ prod f {..i 0" - using \p \ 0\ assms that by (auto simp: gen_has_prod_def) + then have "raw_has_prod f 0 (prod f {..i 0" + using \p \ 0\ assms that by (auto simp: raw_has_prod_def) then show thesis using that by blast qed @@ -581,8 +599,8 @@ corollary convergent_prod_offset_0: fixes f :: "nat \ 'a :: {idom,topological_semigroup_mult,t2_space}" assumes "convergent_prod f" "\i. f i \ 0" - shows "\p. gen_has_prod f 0 p" - using assms convergent_prod_def gen_has_prod_cases by blast + shows "\p. raw_has_prod f 0 p" + using assms convergent_prod_def raw_has_prod_cases by blast lemma prodinf_eq_lim: fixes f :: "nat \ 'a :: {idom,topological_semigroup_mult,t2_space}" @@ -648,7 +666,7 @@ qed lemma has_prod_finite: - fixes f :: "nat \ 'a::{idom,t2_space}" + fixes f :: "nat \ 'a::{semidom,t2_space}" assumes [simp]: "finite N" and f: "\n. n \ N \ f n = 1" shows "f has_prod (\n\N. f n)" @@ -668,7 +686,7 @@ moreover have "(\n. prod f {..n}) \ prod f N" by (rule LIMSEQ_offset[of _ "Suc (Max N)"]) (simp add: eq atLeast0LessThan del: add_Suc_right) ultimately show ?thesis - by (simp add: gen_has_prod_def has_prod_def) + by (simp add: raw_has_prod_def has_prod_def) next case False then obtain k where "k \ N" "f k = 0" @@ -692,8 +710,8 @@ (use Max.coboundedI [OF \finite N\] f in \force+\) finally show ?thesis . qed - have q: "gen_has_prod f (Suc (Max ?Z)) ?q" - proof (simp add: gen_has_prod_def) + have q: "raw_has_prod f (Suc (Max ?Z)) ?q" + proof (simp add: raw_has_prod_def) show "(\n. \i\n. f (Suc (i + Max ?Z))) \ ?q" by (rule LIMSEQ_offset[of _ "(Max N)"]) (simp add: eq) qed @@ -709,17 +727,20 @@ qed corollary has_prod_0: - fixes f :: "nat \ 'a::{idom,t2_space}" + fixes f :: "nat \ 'a::{semidom,t2_space}" assumes "\n. f n = 1" shows "f has_prod 1" by (simp add: assms has_prod_cong) +lemma prodinf_zero[simp]: "prodinf (\n. 1::'a::real_normed_field) = 1" + using has_prod_unique by force + lemma convergent_prod_finite: fixes f :: "nat \ 'a::{idom,t2_space}" assumes "finite N" "\n. n \ N \ f n = 1" shows "convergent_prod f" proof - - have "\n p. gen_has_prod f n p" + have "\n p. raw_has_prod f n p" using assms has_prod_def has_prod_finite by blast then show ?thesis by (simp add: convergent_prod_def) @@ -759,12 +780,12 @@ assumes "convergent_prod f" shows "\p. f has_prod p" proof - - obtain M p where p: "gen_has_prod f M p" + obtain M p where p: "raw_has_prod f M p" using assms convergent_prod_def by blast then have "p \ 0" - using gen_has_prod_nonzero by blast + using raw_has_prod_nonzero by blast with p have fnz: "f i \ 0" if "i \ M" for i - using gen_has_prod_eq_0 that by blast + using raw_has_prod_eq_0 that by blast define C where "C = (\nn\M. f n \ 0") @@ -781,7 +802,7 @@ by (metis (mono_tags, lifting) GreatestI_ex_nat nat_le_linear) have "f i \ 0" if "i > ?N" for i by (metis (mono_tags, lifting) Greatest_le_nat fnz leD linear that) - then have "\p. gen_has_prod f (Suc ?N) p" + then have "\p. raw_has_prod f (Suc ?N) p" using assms by (auto simp: intro!: convergent_prod_ignore_nonzero_segment) then show ?thesis unfolding has_prod_def using 0 by blast @@ -796,7 +817,7 @@ lemma convergent_prod_LIMSEQ: shows "convergent_prod f \ (\n. \i\n. f i) \ prodinf f" by (metis convergent_LIMSEQ_iff convergent_prod_has_prod convergent_prod_imp_convergent - convergent_prod_to_zero_iff gen_has_prod_eq_0 has_prod_def prodinf_eq_lim zero_le) + convergent_prod_to_zero_iff raw_has_prod_eq_0 has_prod_def prodinf_eq_lim zero_le) lemma has_prod_iff: "f has_prod x \ convergent_prod f \ prodinf f = x" proof @@ -818,4 +839,494 @@ end +subsection \Infinite products on ordered, topological monoids\ + +lemma LIMSEQ_prod_0: + fixes f :: "nat \ 'a::{semidom,topological_space}" + assumes "f i = 0" + shows "(\n. prod f {..n}) \ 0" +proof (subst tendsto_cong) + show "\\<^sub>F n in sequentially. prod f {..n} = 0" + proof + show "prod f {..n} = 0" if "n \ i" for n + using that assms by auto + qed +qed auto + +lemma LIMSEQ_prod_nonneg: + fixes f :: "nat \ 'a::{linordered_semidom,linorder_topology}" + assumes 0: "\n. 0 \ f n" and a: "(\n. prod f {..n}) \ a" + shows "a \ 0" + by (simp add: "0" prod_nonneg LIMSEQ_le_const [OF a]) + + +context + fixes f :: "nat \ 'a::{linordered_semidom,linorder_topology}" +begin + +lemma has_prod_le: + assumes f: "f has_prod a" and g: "g has_prod b" and le: "\n. 0 \ f n \ f n \ g n" + shows "a \ b" +proof (cases "a=0 \ b=0") + case True + then show ?thesis + proof + assume [simp]: "a=0" + have "b \ 0" + proof (rule LIMSEQ_prod_nonneg) + show "(\n. prod g {..n}) \ b" + using g by (auto simp: has_prod_def raw_has_prod_def LIMSEQ_prod_0) + qed (use le order_trans in auto) + then show ?thesis + by auto + next + assume [simp]: "b=0" + then obtain i where "g i = 0" + using g by (auto simp: prod_defs) + then have "f i = 0" + using antisym le by force + then have "a=0" + using f by (auto simp: prod_defs LIMSEQ_prod_0 LIMSEQ_unique) + then show ?thesis + by auto + qed +next + case False + then show ?thesis + using assms + unfolding has_prod_def raw_has_prod_def + by (force simp: LIMSEQ_prod_0 intro!: LIMSEQ_le prod_mono) +qed + +lemma prodinf_le: + assumes f: "f has_prod a" and g: "g has_prod b" and le: "\n. 0 \ f n \ f n \ g n" + shows "prodinf f \ prodinf g" + using has_prod_le [OF assms] has_prod_unique f g by blast + end + + +lemma prod_le_prodinf: + fixes f :: "nat \ 'a::{linordered_idom,linorder_topology}" + assumes "f has_prod a" "\i. 0 \ f i" "\i. i\n \ 1 \ f i" + shows "prod f {.. prodinf f" + by(rule has_prod_le[OF has_prod_If_finite_set]) (use assms has_prod_unique in auto) + +lemma prodinf_nonneg: + fixes f :: "nat \ 'a::{linordered_idom,linorder_topology}" + assumes "f has_prod a" "\i. 1 \ f i" + shows "1 \ prodinf f" + using prod_le_prodinf[of f a 0] assms + by (metis order_trans prod_ge_1 zero_le_one) + +lemma prodinf_le_const: + fixes f :: "nat \ real" + assumes "convergent_prod f" "\n. prod f {.. x" + shows "prodinf f \ x" + by (metis lessThan_Suc_atMost assms convergent_prod_LIMSEQ LIMSEQ_le_const2) + +lemma prodinf_eq_one_iff: + fixes f :: "nat \ real" + assumes f: "convergent_prod f" and ge1: "\n. 1 \ f n" + shows "prodinf f = 1 \ (\n. f n = 1)" +proof + assume "prodinf f = 1" + then have "(\n. \i 1" + using convergent_prod_LIMSEQ[of f] assms by (simp add: LIMSEQ_lessThan_iff_atMost) + then have "\i. (\n\{i}. f n) \ 1" + proof (rule LIMSEQ_le_const) + have "1 \ prod f n" for n + by (simp add: ge1 prod_ge_1) + have "prod f {..\n. 1 \ prod f n\ \prodinf f = 1\ antisym f convergent_prod_has_prod ge1 order_trans prod_le_prodinf zero_le_one) + then have "(\n\{i}. f n) \ prod f {.. Suc i" for i n + by (metis mult.left_neutral order_refl prod.cong prod.neutral_const prod_lessThan_Suc) + then show "\N. \n\N. (\n\{i}. f n) \ prod f {..n. f n = 1" + by (auto intro!: antisym) +qed (metis prodinf_zero fun_eq_iff) + +lemma prodinf_pos_iff: + fixes f :: "nat \ real" + assumes "convergent_prod f" "\n. 1 \ f n" + shows "1 < prodinf f \ (\i. 1 < f i)" + using prod_le_prodinf[of f 1] prodinf_eq_one_iff + by (metis convergent_prod_has_prod assms less_le prodinf_nonneg) + +lemma less_1_prodinf2: + fixes f :: "nat \ real" + assumes "convergent_prod f" "\n. 1 \ f n" "1 < f i" + shows "1 < prodinf f" +proof - + have "1 < (\n \ prodinf f" + by (intro prod_le_prodinf) (use assms order_trans zero_le_one in \blast+\) + finally show ?thesis . +qed + +lemma less_1_prodinf: + fixes f :: "nat \ real" + shows "\convergent_prod f; \n. 1 < f n\ \ 1 < prodinf f" + by (intro less_1_prodinf2[where i=1]) (auto intro: less_imp_le) + +lemma prodinf_nonzero: + fixes f :: "nat \ 'a :: {idom,topological_semigroup_mult,t2_space}" + assumes "convergent_prod f" "\i. f i \ 0" + shows "prodinf f \ 0" + by (metis assms convergent_prod_offset_0 has_prod_unique raw_has_prod_def has_prod_def) + +lemma less_0_prodinf: + fixes f :: "nat \ real" + assumes f: "convergent_prod f" and 0: "\i. f i > 0" + shows "0 < prodinf f" +proof - + have "prodinf f \ 0" + by (metis assms less_irrefl prodinf_nonzero) + moreover have "0 < (\n 0" + using convergent_prod_LIMSEQ [OF f] LIMSEQ_prod_nonneg 0 less_le by blast + ultimately show ?thesis + by auto +qed + +lemma prod_less_prodinf2: + fixes f :: "nat \ real" + assumes f: "convergent_prod f" and 1: "\m. m\n \ 1 \ f m" and 0: "\m. 0 < f m" and i: "n \ i" "1 < f i" + shows "prod f {.. prod f {.. prodinf f" + using prod_le_prodinf[of f _ "Suc i"] + by (meson "0" "1" Suc_leD convergent_prod_has_prod f \n \ i\ le_trans less_eq_real_def) + ultimately show ?thesis + by (metis le_less_trans mult.commute not_le prod_lessThan_Suc) +qed + +lemma prod_less_prodinf: + fixes f :: "nat \ real" + assumes f: "convergent_prod f" and 1: "\m. m\n \ 1 < f m" and 0: "\m. 0 < f m" + shows "prod f {.. real" + assumes pos: "\n. 1 \ f n" + and le: "\n. (\i x" + shows "\p. raw_has_prod f 0 p" + unfolding raw_has_prod_def add_0_right +proof (rule exI LIMSEQ_incseq_SUP conjI)+ + show "bdd_above (range (\n. prod f {..n}))" + by (metis bdd_aboveI2 le lessThan_Suc_atMost) + then have "(SUP i. prod f {..i}) > 0" + by (metis UNIV_I cSUP_upper less_le_trans pos prod_pos zero_less_one) + then show "(SUP i. prod f {..i}) \ 0" + by auto + show "incseq (\n. prod f {..n})" + using pos order_trans [OF zero_le_one] by (auto simp: mono_def intro!: prod_mono2) +qed + +lemma convergent_prodI_nonneg_bounded: + fixes f :: "nat \ real" + assumes "\n. 1 \ f n" "\n. (\i x" + shows "convergent_prod f" + using convergent_prod_def raw_has_prodI_bounded [OF assms] by blast + + +subsection \Infinite products on topological monoids\ + +context + fixes f g :: "nat \ 'a::{t2_space,topological_semigroup_mult,idom}" +begin + +lemma raw_has_prod_mult: "\raw_has_prod f M a; raw_has_prod g M b\ \ raw_has_prod (\n. f n * g n) M (a * b)" + by (force simp add: prod.distrib tendsto_mult raw_has_prod_def) + +lemma has_prod_mult_nz: "\f has_prod a; g has_prod b; a \ 0; b \ 0\ \ (\n. f n * g n) has_prod (a * b)" + by (simp add: raw_has_prod_mult has_prod_def) + +end + + +context + fixes f g :: "nat \ 'a::real_normed_field" +begin + +lemma has_prod_mult: + assumes f: "f has_prod a" and g: "g has_prod b" + shows "(\n. f n * g n) has_prod (a * b)" + using f [unfolded has_prod_def] +proof (elim disjE exE conjE) + assume f0: "raw_has_prod f 0 a" + show ?thesis + using g [unfolded has_prod_def] + proof (elim disjE exE conjE) + assume g0: "raw_has_prod g 0 b" + with f0 show ?thesis + by (force simp add: has_prod_def prod.distrib tendsto_mult raw_has_prod_def) + next + fix j q + assume "b = 0" and "g j = 0" and q: "raw_has_prod g (Suc j) q" + obtain p where p: "raw_has_prod f (Suc j) p" + using f0 raw_has_prod_ignore_initial_segment by blast + then have "Ex (raw_has_prod (\n. f n * g n) (Suc j))" + using q raw_has_prod_mult by blast + then show ?thesis + using \b = 0\ \g j = 0\ has_prod_0_iff by fastforce + qed +next + fix i p + assume "a = 0" and "f i = 0" and p: "raw_has_prod f (Suc i) p" + show ?thesis + using g [unfolded has_prod_def] + proof (elim disjE exE conjE) + assume g0: "raw_has_prod g 0 b" + obtain q where q: "raw_has_prod g (Suc i) q" + using g0 raw_has_prod_ignore_initial_segment by blast + then have "Ex (raw_has_prod (\n. f n * g n) (Suc i))" + using raw_has_prod_mult p by blast + then show ?thesis + using \a = 0\ \f i = 0\ has_prod_0_iff by fastforce + next + fix j q + assume "b = 0" and "g j = 0" and q: "raw_has_prod g (Suc j) q" + obtain p' where p': "raw_has_prod f (Suc (max i j)) p'" + by (metis raw_has_prod_ignore_initial_segment max_Suc_Suc max_def p) + moreover + obtain q' where q': "raw_has_prod g (Suc (max i j)) q'" + by (metis raw_has_prod_ignore_initial_segment max.cobounded2 max_Suc_Suc q) + ultimately show ?thesis + using \b = 0\ by (simp add: has_prod_def) (metis \f i = 0\ \g j = 0\ raw_has_prod_mult max_def) + qed +qed + +lemma convergent_prod_mult: + assumes f: "convergent_prod f" and g: "convergent_prod g" + shows "convergent_prod (\n. f n * g n)" + unfolding convergent_prod_def +proof - + obtain M p N q where p: "raw_has_prod f M p" and q: "raw_has_prod g N q" + using convergent_prod_def f g by blast+ + then obtain p' q' where p': "raw_has_prod f (max M N) p'" and q': "raw_has_prod g (max M N) q'" + by (meson raw_has_prod_ignore_initial_segment max.cobounded1 max.cobounded2) + then show "\M p. raw_has_prod (\n. f n * g n) M p" + using raw_has_prod_mult by blast +qed + +lemma prodinf_mult: "convergent_prod f \ convergent_prod g \ prodinf f * prodinf g = (\n. f n * g n)" + by (intro has_prod_unique has_prod_mult convergent_prod_has_prod) + +end + +context + fixes f :: "'i \ nat \ 'a::real_normed_field" + and I :: "'i set" +begin + +lemma has_prod_prod: "(\i. i \ I \ (f i) has_prod (x i)) \ (\n. \i\I. f i n) has_prod (\i\I. x i)" + by (induct I rule: infinite_finite_induct) (auto intro!: has_prod_mult) + +lemma prodinf_prod: "(\i. i \ I \ convergent_prod (f i)) \ (\n. \i\I. f i n) = (\i\I. \n. f i n)" + using has_prod_unique[OF has_prod_prod, OF convergent_prod_has_prod] by simp + +lemma convergent_prod_prod: "(\i. i \ I \ convergent_prod (f i)) \ convergent_prod (\n. \i\I. f i n)" + using convergent_prod_has_prod_iff has_prod_prod prodinf_prod by force + +end + +subsection \Infinite summability on real normed vector spaces\ + +context + fixes f :: "nat \ 'a::real_normed_field" +begin + +lemma raw_has_prod_Suc_iff: "raw_has_prod f M (a * f M) \ raw_has_prod (\n. f (Suc n)) M a \ f M \ 0" +proof - + have "raw_has_prod f M (a * f M) \ (\i. \j\Suc i. f (j+M)) \ a * f M \ a * f M \ 0" + by (subst LIMSEQ_Suc_iff) (simp add: raw_has_prod_def) + also have "\ \ (\i. (\j\i. f (Suc j + M)) * f M) \ a * f M \ a * f M \ 0" + by (simp add: ac_simps atMost_Suc_eq_insert_0 image_Suc_atMost prod_atLeast1_atMost_eq lessThan_Suc_atMost) + also have "\ \ raw_has_prod (\n. f (Suc n)) M a \ f M \ 0" + proof safe + assume tends: "(\i. (\j\i. f (Suc j + M)) * f M) \ a * f M" and 0: "a * f M \ 0" + with tendsto_divide[OF tends tendsto_const, of "f M"] + show "raw_has_prod (\n. f (Suc n)) M a" + by (simp add: raw_has_prod_def) + qed (auto intro: tendsto_mult_right simp: raw_has_prod_def) + finally show ?thesis . +qed + +lemma has_prod_Suc_iff: + assumes "f 0 \ 0" shows "(\n. f (Suc n)) has_prod a \ f has_prod (a * f 0)" +proof (cases "a = 0") + case True + then show ?thesis + proof (simp add: has_prod_def, safe) + fix i x + assume "f (Suc i) = 0" and "raw_has_prod (\n. f (Suc n)) (Suc i) x" + then obtain y where "raw_has_prod f (Suc (Suc i)) y" + by (metis (no_types) raw_has_prod_eq_0 Suc_n_not_le_n raw_has_prod_Suc_iff raw_has_prod_ignore_initial_segment raw_has_prod_nonzero linear) + then show "\i. f i = 0 \ Ex (raw_has_prod f (Suc i))" + using \f (Suc i) = 0\ by blast + next + fix i x + assume "f i = 0" and x: "raw_has_prod f (Suc i) x" + then obtain j where j: "i = Suc j" + by (metis assms not0_implies_Suc) + moreover have "\ y. raw_has_prod (\n. f (Suc n)) i y" + using x by (auto simp: raw_has_prod_def) + then show "\i. f (Suc i) = 0 \ Ex (raw_has_prod (\n. f (Suc n)) (Suc i))" + using \f i = 0\ j by blast + qed +next + case False + then show ?thesis + by (auto simp: has_prod_def raw_has_prod_Suc_iff assms) +qed + +lemma convergent_prod_Suc_iff: + assumes "\k. f k \ 0" shows "convergent_prod (\n. f (Suc n)) = convergent_prod f" +proof + assume "convergent_prod f" + then have "f has_prod prodinf f" + by (rule convergent_prod_has_prod) + moreover have "prodinf f \ 0" + by (simp add: \convergent_prod f\ assms prodinf_nonzero) + ultimately have "(\n. f (Suc n)) has_prod (prodinf f * inverse (f 0))" + by (simp add: has_prod_Suc_iff inverse_eq_divide assms) + then show "convergent_prod (\n. f (Suc n))" + using has_prod_iff by blast +next + assume "convergent_prod (\n. f (Suc n))" + then show "convergent_prod f" + using assms convergent_prod_def raw_has_prod_Suc_iff by blast +qed + +lemma raw_has_prod_inverse: + assumes "raw_has_prod f M a" shows "raw_has_prod (\n. inverse (f n)) M (inverse a)" + using assms unfolding raw_has_prod_def by (auto dest: tendsto_inverse simp: prod_inversef [symmetric]) + +lemma has_prod_inverse: + assumes "f has_prod a" shows "(\n. inverse (f n)) has_prod (inverse a)" +using assms raw_has_prod_inverse unfolding has_prod_def by auto + +lemma convergent_prod_inverse: + assumes "convergent_prod f" + shows "convergent_prod (\n. inverse (f n))" + using assms unfolding convergent_prod_def by (blast intro: raw_has_prod_inverse elim: ) + +end + +context (* Separate contexts are necessary to allow general use of the results above, here. *) + fixes f :: "nat \ 'a::real_normed_field" +begin + +lemma raw_has_prod_Suc_iff': "raw_has_prod f M a \ raw_has_prod (\n. f (Suc n)) M (a / f M) \ f M \ 0" + by (metis raw_has_prod_eq_0 add.commute add.left_neutral raw_has_prod_Suc_iff raw_has_prod_nonzero le_add1 nonzero_mult_div_cancel_right times_divide_eq_left) + +lemma has_prod_divide: "f has_prod a \ g has_prod b \ (\n. f n / g n) has_prod (a / b)" + unfolding divide_inverse by (intro has_prod_inverse has_prod_mult) + +lemma convergent_prod_divide: + assumes f: "convergent_prod f" and g: "convergent_prod g" + shows "convergent_prod (\n. f n / g n)" + using f g has_prod_divide has_prod_iff by blast + +lemma prodinf_divide: "convergent_prod f \ convergent_prod g \ prodinf f / prodinf g = (\n. f n / g n)" + by (intro has_prod_unique has_prod_divide convergent_prod_has_prod) + +lemma prodinf_inverse: "convergent_prod f \ (\n. inverse (f n)) = inverse (\n. f n)" + by (intro has_prod_unique [symmetric] has_prod_inverse convergent_prod_has_prod) + +lemma has_prod_iff_shift: + assumes "\i. i < n \ f i \ 0" + shows "(\i. f (i + n)) has_prod a \ f has_prod (a * (\ii. f (Suc i + n)) has_prod a \ (\i. f (i + n)) has_prod (a * f n)" + by (subst has_prod_Suc_iff) auto + with Suc show ?case + by (simp add: ac_simps) +qed + +corollary has_prod_iff_shift': + assumes "\i. i < n \ f i \ 0" + shows "(\i. f (i + n)) has_prod (a / (\i f has_prod a" + by (simp add: assms has_prod_iff_shift) + +lemma has_prod_one_iff_shift: + assumes "\i. i < n \ f i = 1" + shows "(\i. f (i+n)) has_prod a \ (\i. f i) has_prod a" + by (simp add: assms has_prod_iff_shift) + +lemma convergent_prod_iff_shift: + shows "convergent_prod (\i. f (i + n)) \ convergent_prod f" + apply safe + using convergent_prod_offset apply blast + using convergent_prod_ignore_initial_segment convergent_prod_def by blast + +lemma has_prod_split_initial_segment: + assumes "f has_prod a" "\i. i < n \ f i \ 0" + shows "(\i. f (i + n)) has_prod (a / (\ii. i < n \ f i \ 0" + shows "(\i. f (i + n)) = (\i. f i) / (\ii. i < n \ f i \ 0" + shows "prodinf f = (\i. f (i + n)) * (\i 0" + shows "(\n. f (Suc n)) = prodinf f / f 0" + using prodinf_split_initial_segment[of 1] assms by simp + +end + +context (* Separate contexts are necessary to allow general use of the results above, here. *) + fixes f :: "nat \ 'a::real_normed_field" +begin + +lemma convergent_prod_inverse_iff: "convergent_prod (\n. inverse (f n)) \ convergent_prod f" + by (auto dest: convergent_prod_inverse) + +lemma convergent_prod_const_iff: + fixes c :: "'a :: {real_normed_field}" + shows "convergent_prod (\_. c) \ c = 1" +proof + assume "convergent_prod (\_. c)" + then show "c = 1" + using convergent_prod_imp_LIMSEQ LIMSEQ_unique by blast +next + assume "c = 1" + then show "convergent_prod (\_. c)" + by auto +qed + +lemma has_prod_power: "f has_prod a \ (\i. f i ^ n) has_prod (a ^ n)" + by (induction n) (auto simp: has_prod_mult) + +lemma convergent_prod_power: "convergent_prod f \ convergent_prod (\i. f i ^ n)" + by (induction n) (auto simp: convergent_prod_mult) + +lemma prodinf_power: "convergent_prod f \ prodinf (\i. f i ^ n) = prodinf f ^ n" + by (metis has_prod_unique convergent_prod_imp_has_prod has_prod_power) + +end + +end diff -r 0f19c98fa7be -r 20375f232f3b src/HOL/Groups_Big.thy --- a/src/HOL/Groups_Big.thy Sat Jun 02 22:57:44 2018 +0100 +++ b/src/HOL/Groups_Big.thy Sun Jun 03 15:22:30 2018 +0100 @@ -1302,17 +1302,20 @@ by simp qed -lemma (in linordered_semidom) prod_nonneg: "(\a\A. 0 \ f a) \ 0 \ prod f A" +context linordered_semidom +begin + +lemma prod_nonneg: "(\a\A. 0 \ f a) \ 0 \ prod f A" by (induct A rule: infinite_finite_induct) simp_all -lemma (in linordered_semidom) prod_pos: "(\a\A. 0 < f a) \ 0 < prod f A" +lemma prod_pos: "(\a\A. 0 < f a) \ 0 < prod f A" by (induct A rule: infinite_finite_induct) simp_all -lemma (in linordered_semidom) prod_mono: +lemma prod_mono: "(\i. i \ A \ 0 \ f i \ f i \ g i) \ prod f A \ prod g A" by (induct A rule: infinite_finite_induct) (force intro!: prod_nonneg mult_mono)+ -lemma (in linordered_semidom) prod_mono_strict: +lemma prod_mono_strict: assumes "finite A" "\i. i \ A \ 0 \ f i \ f i < g i" "A \ {}" shows "prod f A < prod g A" using assms @@ -1324,6 +1327,42 @@ then show ?case by (force intro: mult_strict_mono' prod_nonneg) qed +end + +lemma prod_mono2: + fixes f :: "'a \ 'b :: linordered_idom" + assumes fin: "finite B" + and sub: "A \ B" + and nn: "\b. b \ B-A \ 1 \ f b" + and A: "\a. a \ A \ 0 \ f a" + shows "prod f A \ prod f B" +proof - + have "prod f A \ prod f A * prod f (B-A)" + by (metis prod_ge_1 A mult_le_cancel_left1 nn not_less prod_nonneg) + also from fin finite_subset[OF sub fin] have "\ = prod f (A \ (B-A))" + by (simp add: prod.union_disjoint del: Un_Diff_cancel) + also from sub have "A \ (B-A) = B" by blast + finally show ?thesis . +qed + +lemma less_1_prod: + fixes f :: "'a \ 'b::linordered_idom" + shows "finite I \ I \ {} \ (\i. i \ I \ 1 < f i) \ 1 < prod f I" + by (induct I rule: finite_ne_induct) (auto intro: less_1_mult) + +lemma less_1_prod2: + fixes f :: "'a \ 'b::linordered_idom" + assumes I: "finite I" "i \ I" "1 < f i" "\i. i \ I \ 1 \ f i" + shows "1 < prod f I" +proof - + have "1 < f i * prod f (I - {i})" + using assms + by (meson DiffD1 leI less_1_mult less_le_trans mult_le_cancel_left1 prod_ge_1) + also have "\ = prod f I" + using assms by (simp add: prod.remove) + finally show ?thesis . +qed + lemma (in linordered_field) abs_prod: "\prod f A\ = (\x\A. \f x\)" by (induct A rule: infinite_finite_induct) (simp_all add: abs_mult) diff -r 0f19c98fa7be -r 20375f232f3b src/HOL/Set_Interval.thy --- a/src/HOL/Set_Interval.thy Sat Jun 02 22:57:44 2018 +0100 +++ b/src/HOL/Set_Interval.thy Sun Jun 03 15:22:30 2018 +0100 @@ -507,7 +507,7 @@ lemma Iio_eq_empty_iff: "{..< n::'a::{linorder, order_bot}} = {} \ n = bot" by (auto simp: set_eq_iff not_less le_bot) -lemma Iio_eq_empty_iff_nat: "{..< n::nat} = {} \ n = 0" +lemma lessThan_empty_iff: "{..< n::nat} = {} \ n = 0" by (simp add: Iio_eq_empty_iff bot_nat_def) lemma mono_image_least: @@ -709,7 +709,7 @@ lemma lessThan_Suc_atMost: "lessThan (Suc k) = atMost k" by (simp add: lessThan_def atMost_def less_Suc_eq_le) -lemma Iic_Suc_eq_insert_0: "{.. Suc n} = insert 0 (Suc ` {.. n})" +lemma atMost_Suc_eq_insert_0: "{.. Suc n} = insert 0 (Suc ` {.. n})" unfolding lessThan_Suc_atMost[symmetric] lessThan_Suc_eq_insert_0[of "Suc n"] .. lemma UN_lessThan_UNIV: "(UN m::nat. lessThan m) = UNIV" @@ -803,7 +803,7 @@ lemma atLeast0_atMost_Suc_eq_insert_0: "{0..Suc n} = insert 0 (Suc ` {0..n})" - by (simp add: atLeast0AtMost Iic_Suc_eq_insert_0) + by (simp add: atLeast0AtMost atMost_Suc_eq_insert_0) subsubsection \Intervals of nats with @{term Suc}\ @@ -2408,6 +2408,16 @@ "\i\n. t" \ "CONST prod (\i. t) {..n}" "\i "CONST prod (\i. t) {..k = (\k{int i..int (i+j)}" by (induct j) (auto simp add: atLeastAtMostSuc_conv atLeastAtMostPlus1_int_conv) @@ -2441,7 +2451,7 @@ "prod f {Suc m..iin. f (Suc n)) \ l = f \ l" by (blast intro: LIMSEQ_imp_Suc LIMSEQ_Suc) +lemma LIMSEQ_lessThan_iff_atMost: + shows "(\n. f {.. x \ (\n. f {..n}) \ x" + apply (subst LIMSEQ_Suc_iff [symmetric]) + apply (simp only: lessThan_Suc_atMost) + done + lemma LIMSEQ_unique: "X \ a \ X \ b \ a = b" for a b :: "'a::t2_space" using trivial_limit_sequentially by (rule tendsto_unique)