# HG changeset patch # User hoelzl # Date 1470332191 -7200 # Node ID bd218a9320b54363495253e9ce2706b08b8821ec # Parent bbcb05504fdc90533386fc0d4d59fe4b5f5a742f HOL-Multivariate_Analysis: rename theories for more descriptive names diff -r bbcb05504fdc -r bd218a9320b5 src/HOL/Multivariate_Analysis/Bounded_Continuous_Function.thy --- a/src/HOL/Multivariate_Analysis/Bounded_Continuous_Function.thy Thu Aug 04 18:45:28 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Bounded_Continuous_Function.thy Thu Aug 04 19:36:31 2016 +0200 @@ -1,7 +1,7 @@ section \Bounded Continuous Functions\ theory Bounded_Continuous_Function -imports Integration +imports Henstock_Kurzweil_Integration begin subsection \Definition\ diff -r bbcb05504fdc -r bd218a9320b5 src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Thu Aug 04 18:45:28 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Thu Aug 04 19:36:31 2016 +0200 @@ -1,7 +1,7 @@ section \Instantiates the finite Cartesian product of Euclidean spaces as a Euclidean space.\ theory Cartesian_Euclidean_Space -imports Finite_Cartesian_Product Integration +imports Finite_Cartesian_Product Henstock_Kurzweil_Integration begin lemma subspace_special_hyperplane: "subspace {x. x $ k = 0}" diff -r bbcb05504fdc -r bd218a9320b5 src/HOL/Multivariate_Analysis/Cauchy_Integral_Theorem.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Multivariate_Analysis/Cauchy_Integral_Theorem.thy Thu Aug 04 19:36:31 2016 +0200 @@ -0,0 +1,7539 @@ +section \Complex path integrals and Cauchy's integral theorem\ + +text\By John Harrison et al. Ported from HOL Light by L C Paulson (2015)\ + +theory Cauchy_Integral_Theorem +imports Complex_Transcendental Weierstrass_Theorems Ordered_Euclidean_Space +begin + +subsection\Homeomorphisms of arc images\ + +lemma homeomorphism_arc: + fixes g :: "real \ 'a::t2_space" + assumes "arc g" + obtains h where "homeomorphism {0..1} (path_image g) g h" +using assms by (force simp add: arc_def homeomorphism_compact path_def path_image_def) + +lemma homeomorphic_arc_image_interval: + fixes g :: "real \ 'a::t2_space" and a::real + assumes "arc g" "a < b" + shows "(path_image g) homeomorphic {a..b}" +proof - + have "(path_image g) homeomorphic {0..1::real}" + by (meson assms(1) homeomorphic_def homeomorphic_sym homeomorphism_arc) + also have "... homeomorphic {a..b}" + using assms by (force intro: homeomorphic_closed_intervals_real) + finally show ?thesis . +qed + +lemma homeomorphic_arc_images: + fixes g :: "real \ 'a::t2_space" and h :: "real \ 'b::t2_space" + assumes "arc g" "arc h" + shows "(path_image g) homeomorphic (path_image h)" +proof - + have "(path_image g) homeomorphic {0..1::real}" + by (meson assms homeomorphic_def homeomorphic_sym homeomorphism_arc) + also have "... homeomorphic (path_image h)" + by (meson assms homeomorphic_def homeomorphism_arc) + finally show ?thesis . +qed + +subsection \Piecewise differentiable functions\ + +definition piecewise_differentiable_on + (infixr "piecewise'_differentiable'_on" 50) + where "f piecewise_differentiable_on i \ + continuous_on i f \ + (\s. finite s \ (\x \ i - s. f differentiable (at x within i)))" + +lemma piecewise_differentiable_on_imp_continuous_on: + "f piecewise_differentiable_on s \ continuous_on s f" +by (simp add: piecewise_differentiable_on_def) + +lemma piecewise_differentiable_on_subset: + "f piecewise_differentiable_on s \ t \ s \ f piecewise_differentiable_on t" + using continuous_on_subset + unfolding piecewise_differentiable_on_def + apply safe + apply (blast intro: elim: continuous_on_subset) + by (meson Diff_iff differentiable_within_subset subsetCE) + +lemma differentiable_on_imp_piecewise_differentiable: + fixes a:: "'a::{linorder_topology,real_normed_vector}" + shows "f differentiable_on {a..b} \ f piecewise_differentiable_on {a..b}" + apply (simp add: piecewise_differentiable_on_def differentiable_imp_continuous_on) + apply (rule_tac x="{a,b}" in exI, simp add: differentiable_on_def) + done + +lemma differentiable_imp_piecewise_differentiable: + "(\x. x \ s \ f differentiable (at x within s)) + \ f piecewise_differentiable_on s" +by (auto simp: piecewise_differentiable_on_def differentiable_imp_continuous_on differentiable_on_def + intro: differentiable_within_subset) + +lemma piecewise_differentiable_const [iff]: "(\x. z) piecewise_differentiable_on s" + by (simp add: differentiable_imp_piecewise_differentiable) + +lemma piecewise_differentiable_compose: + "\f piecewise_differentiable_on s; g piecewise_differentiable_on (f ` s); + \x. finite (s \ f-`{x})\ + \ (g o f) piecewise_differentiable_on s" + apply (simp add: piecewise_differentiable_on_def, safe) + apply (blast intro: continuous_on_compose2) + apply (rename_tac A B) + apply (rule_tac x="A \ (\x\B. s \ f-`{x})" in exI) + apply (blast intro: differentiable_chain_within) + done + +lemma piecewise_differentiable_affine: + fixes m::real + assumes "f piecewise_differentiable_on ((\x. m *\<^sub>R x + c) ` s)" + shows "(f o (\x. m *\<^sub>R x + c)) piecewise_differentiable_on s" +proof (cases "m = 0") + case True + then show ?thesis + unfolding o_def + by (force intro: differentiable_imp_piecewise_differentiable differentiable_const) +next + case False + show ?thesis + apply (rule piecewise_differentiable_compose [OF differentiable_imp_piecewise_differentiable]) + apply (rule assms derivative_intros | simp add: False vimage_def real_vector_affinity_eq)+ + done +qed + +lemma piecewise_differentiable_cases: + fixes c::real + assumes "f piecewise_differentiable_on {a..c}" + "g piecewise_differentiable_on {c..b}" + "a \ c" "c \ b" "f c = g c" + shows "(\x. if x \ c then f x else g x) piecewise_differentiable_on {a..b}" +proof - + obtain s t where st: "finite s" "finite t" + "\x\{a..c} - s. f differentiable at x within {a..c}" + "\x\{c..b} - t. g differentiable at x within {c..b}" + using assms + by (auto simp: piecewise_differentiable_on_def) + have finabc: "finite ({a,b,c} \ (s \ t))" + by (metis \finite s\ \finite t\ finite_Un finite_insert finite.emptyI) + have "continuous_on {a..c} f" "continuous_on {c..b} g" + using assms piecewise_differentiable_on_def by auto + then have "continuous_on {a..b} (\x. if x \ c then f x else g x)" + using continuous_on_cases [OF closed_real_atLeastAtMost [of a c], + OF closed_real_atLeastAtMost [of c b], + of f g "\x. x\c"] assms + by (force simp: ivl_disj_un_two_touch) + moreover + { fix x + assume x: "x \ {a..b} - ({a,b,c} \ (s \ t))" + have "(\x. if x \ c then f x else g x) differentiable at x within {a..b}" (is "?diff_fg") + proof (cases x c rule: le_cases) + case le show ?diff_fg + apply (rule differentiable_transform_within [where d = "dist x c" and f = f]) + using x le st + apply (simp_all add: dist_real_def) + apply (rule differentiable_at_withinI) + apply (rule differentiable_within_open [where s = "{a<..s. finite s \ + (\x\{a..b} - s. (\x. if x \ c then f x else g x) differentiable at x within {a..b})" + by (meson finabc) + ultimately show ?thesis + by (simp add: piecewise_differentiable_on_def) +qed + +lemma piecewise_differentiable_neg: + "f piecewise_differentiable_on s \ (\x. -(f x)) piecewise_differentiable_on s" + by (auto simp: piecewise_differentiable_on_def continuous_on_minus) + +lemma piecewise_differentiable_add: + assumes "f piecewise_differentiable_on i" + "g piecewise_differentiable_on i" + shows "(\x. f x + g x) piecewise_differentiable_on i" +proof - + obtain s t where st: "finite s" "finite t" + "\x\i - s. f differentiable at x within i" + "\x\i - t. g differentiable at x within i" + using assms by (auto simp: piecewise_differentiable_on_def) + then have "finite (s \ t) \ (\x\i - (s \ t). (\x. f x + g x) differentiable at x within i)" + by auto + moreover have "continuous_on i f" "continuous_on i g" + using assms piecewise_differentiable_on_def by auto + ultimately show ?thesis + by (auto simp: piecewise_differentiable_on_def continuous_on_add) +qed + +lemma piecewise_differentiable_diff: + "\f piecewise_differentiable_on s; g piecewise_differentiable_on s\ + \ (\x. f x - g x) piecewise_differentiable_on s" + unfolding diff_conv_add_uminus + by (metis piecewise_differentiable_add piecewise_differentiable_neg) + +lemma continuous_on_joinpaths_D1: + "continuous_on {0..1} (g1 +++ g2) \ continuous_on {0..1} g1" + apply (rule continuous_on_eq [of _ "(g1 +++ g2) o (op*(inverse 2))"]) + apply (rule continuous_intros | simp)+ + apply (auto elim!: continuous_on_subset simp: joinpaths_def) + done + +lemma continuous_on_joinpaths_D2: + "\continuous_on {0..1} (g1 +++ g2); pathfinish g1 = pathstart g2\ \ continuous_on {0..1} g2" + apply (rule continuous_on_eq [of _ "(g1 +++ g2) o (\x. inverse 2*x + 1/2)"]) + apply (rule continuous_intros | simp)+ + apply (auto elim!: continuous_on_subset simp add: joinpaths_def pathfinish_def pathstart_def Ball_def) + done + +lemma piecewise_differentiable_D1: + "(g1 +++ g2) piecewise_differentiable_on {0..1} \ g1 piecewise_differentiable_on {0..1}" + apply (clarsimp simp add: piecewise_differentiable_on_def dest!: continuous_on_joinpaths_D1) + apply (rule_tac x="insert 1 ((op*2)`s)" in exI) + apply simp + apply (intro ballI) + apply (rule_tac d="dist (x/2) (1/2)" and f = "(g1 +++ g2) o (op*(inverse 2))" + in differentiable_transform_within) + apply (auto simp: dist_real_def joinpaths_def) + apply (rule differentiable_chain_within derivative_intros | simp)+ + apply (rule differentiable_subset) + apply (force simp:)+ + done + +lemma piecewise_differentiable_D2: + "\(g1 +++ g2) piecewise_differentiable_on {0..1}; pathfinish g1 = pathstart g2\ + \ g2 piecewise_differentiable_on {0..1}" + apply (clarsimp simp add: piecewise_differentiable_on_def dest!: continuous_on_joinpaths_D2) + apply (rule_tac x="insert 0 ((\x. 2*x-1)`s)" in exI) + apply simp + apply (intro ballI) + apply (rule_tac d="dist ((x+1)/2) (1/2)" and f = "(g1 +++ g2) o (\x. (x+1)/2)" + in differentiable_transform_within) + apply (auto simp: dist_real_def joinpaths_def abs_if field_simps split: if_split_asm) + apply (rule differentiable_chain_within derivative_intros | simp)+ + apply (rule differentiable_subset) + apply (force simp: divide_simps)+ + done + + +subsubsection\The concept of continuously differentiable\ + +text \ +John Harrison writes as follows: + +``The usual assumption in complex analysis texts is that a path \\\ should be piecewise +continuously differentiable, which ensures that the path integral exists at least for any continuous +f, since all piecewise continuous functions are integrable. However, our notion of validity is +weaker, just piecewise differentiability... [namely] continuity plus differentiability except on a +finite set ... [Our] underlying theory of integration is the Kurzweil-Henstock theory. In contrast to +the Riemann or Lebesgue theory (but in common with a simple notion based on antiderivatives), this +can integrate all derivatives.'' + +"Formalizing basic complex analysis." From Insight to Proof: Festschrift in Honour of Andrzej Trybulec. +Studies in Logic, Grammar and Rhetoric 10.23 (2007): 151-165. + +And indeed he does not assume that his derivatives are continuous, but the penalty is unreasonably +difficult proofs concerning winding numbers. We need a self-contained and straightforward theorem +asserting that all derivatives can be integrated before we can adopt Harrison's choice.\ + +definition C1_differentiable_on :: "(real \ 'a::real_normed_vector) \ real set \ bool" + (infix "C1'_differentiable'_on" 50) + where + "f C1_differentiable_on s \ + (\D. (\x \ s. (f has_vector_derivative (D x)) (at x)) \ continuous_on s D)" + +lemma C1_differentiable_on_eq: + "f C1_differentiable_on s \ + (\x \ s. f differentiable at x) \ continuous_on s (\x. vector_derivative f (at x))" + unfolding C1_differentiable_on_def + apply safe + using differentiable_def has_vector_derivative_def apply blast + apply (erule continuous_on_eq) + using vector_derivative_at apply fastforce + using vector_derivative_works apply fastforce + done + +lemma C1_differentiable_on_subset: + "f C1_differentiable_on t \ s \ t \ f C1_differentiable_on s" + unfolding C1_differentiable_on_def continuous_on_eq_continuous_within + by (blast intro: continuous_within_subset) + +lemma C1_differentiable_compose: + "\f C1_differentiable_on s; g C1_differentiable_on (f ` s); + \x. finite (s \ f-`{x})\ + \ (g o f) C1_differentiable_on s" + apply (simp add: C1_differentiable_on_eq, safe) + using differentiable_chain_at apply blast + apply (rule continuous_on_eq [of _ "\x. vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x))"]) + apply (rule Limits.continuous_on_scaleR, assumption) + apply (metis (mono_tags, lifting) continuous_on_eq continuous_at_imp_continuous_on continuous_on_compose differentiable_imp_continuous_within o_def) + by (simp add: vector_derivative_chain_at) + +lemma C1_diff_imp_diff: "f C1_differentiable_on s \ f differentiable_on s" + by (simp add: C1_differentiable_on_eq differentiable_at_imp_differentiable_on) + +lemma C1_differentiable_on_ident [simp, derivative_intros]: "(\x. x) C1_differentiable_on s" + by (auto simp: C1_differentiable_on_eq continuous_on_const) + +lemma C1_differentiable_on_const [simp, derivative_intros]: "(\z. a) C1_differentiable_on s" + by (auto simp: C1_differentiable_on_eq continuous_on_const) + +lemma C1_differentiable_on_add [simp, derivative_intros]: + "f C1_differentiable_on s \ g C1_differentiable_on s \ (\x. f x + g x) C1_differentiable_on s" + unfolding C1_differentiable_on_eq by (auto intro: continuous_intros) + +lemma C1_differentiable_on_minus [simp, derivative_intros]: + "f C1_differentiable_on s \ (\x. - f x) C1_differentiable_on s" + unfolding C1_differentiable_on_eq by (auto intro: continuous_intros) + +lemma C1_differentiable_on_diff [simp, derivative_intros]: + "f C1_differentiable_on s \ g C1_differentiable_on s \ (\x. f x - g x) C1_differentiable_on s" + unfolding C1_differentiable_on_eq by (auto intro: continuous_intros) + +lemma C1_differentiable_on_mult [simp, derivative_intros]: + fixes f g :: "real \ 'a :: real_normed_algebra" + shows "f C1_differentiable_on s \ g C1_differentiable_on s \ (\x. f x * g x) C1_differentiable_on s" + unfolding C1_differentiable_on_eq + by (auto simp: continuous_on_add continuous_on_mult continuous_at_imp_continuous_on differentiable_imp_continuous_within) + +lemma C1_differentiable_on_scaleR [simp, derivative_intros]: + "f C1_differentiable_on s \ g C1_differentiable_on s \ (\x. f x *\<^sub>R g x) C1_differentiable_on s" + unfolding C1_differentiable_on_eq + by (rule continuous_intros | simp add: continuous_at_imp_continuous_on differentiable_imp_continuous_within)+ + + +definition piecewise_C1_differentiable_on + (infixr "piecewise'_C1'_differentiable'_on" 50) + where "f piecewise_C1_differentiable_on i \ + continuous_on i f \ + (\s. finite s \ (f C1_differentiable_on (i - s)))" + +lemma C1_differentiable_imp_piecewise: + "f C1_differentiable_on s \ f piecewise_C1_differentiable_on s" + by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq continuous_at_imp_continuous_on differentiable_imp_continuous_within) + +lemma piecewise_C1_imp_differentiable: + "f piecewise_C1_differentiable_on i \ f piecewise_differentiable_on i" + by (auto simp: piecewise_C1_differentiable_on_def piecewise_differentiable_on_def + C1_differentiable_on_def differentiable_def has_vector_derivative_def + intro: has_derivative_at_within) + +lemma piecewise_C1_differentiable_compose: + "\f piecewise_C1_differentiable_on s; g piecewise_C1_differentiable_on (f ` s); + \x. finite (s \ f-`{x})\ + \ (g o f) piecewise_C1_differentiable_on s" + apply (simp add: piecewise_C1_differentiable_on_def, safe) + apply (blast intro: continuous_on_compose2) + apply (rename_tac A B) + apply (rule_tac x="A \ (\x\B. s \ f-`{x})" in exI) + apply (rule conjI, blast) + apply (rule C1_differentiable_compose) + apply (blast intro: C1_differentiable_on_subset) + apply (blast intro: C1_differentiable_on_subset) + by (simp add: Diff_Int_distrib2) + +lemma piecewise_C1_differentiable_on_subset: + "f piecewise_C1_differentiable_on s \ t \ s \ f piecewise_C1_differentiable_on t" + by (auto simp: piecewise_C1_differentiable_on_def elim!: continuous_on_subset C1_differentiable_on_subset) + +lemma C1_differentiable_imp_continuous_on: + "f C1_differentiable_on s \ continuous_on s f" + unfolding C1_differentiable_on_eq continuous_on_eq_continuous_within + using differentiable_at_withinI differentiable_imp_continuous_within by blast + +lemma C1_differentiable_on_empty [iff]: "f C1_differentiable_on {}" + unfolding C1_differentiable_on_def + by auto + +lemma piecewise_C1_differentiable_affine: + fixes m::real + assumes "f piecewise_C1_differentiable_on ((\x. m * x + c) ` s)" + shows "(f o (\x. m *\<^sub>R x + c)) piecewise_C1_differentiable_on s" +proof (cases "m = 0") + case True + then show ?thesis + unfolding o_def by (auto simp: piecewise_C1_differentiable_on_def continuous_on_const) +next + case False + show ?thesis + apply (rule piecewise_C1_differentiable_compose [OF C1_differentiable_imp_piecewise]) + apply (rule assms derivative_intros | simp add: False vimage_def)+ + using real_vector_affinity_eq [OF False, where c=c, unfolded scaleR_conv_of_real] + apply simp + done +qed + +lemma piecewise_C1_differentiable_cases: + fixes c::real + assumes "f piecewise_C1_differentiable_on {a..c}" + "g piecewise_C1_differentiable_on {c..b}" + "a \ c" "c \ b" "f c = g c" + shows "(\x. if x \ c then f x else g x) piecewise_C1_differentiable_on {a..b}" +proof - + obtain s t where st: "f C1_differentiable_on ({a..c} - s)" + "g C1_differentiable_on ({c..b} - t)" + "finite s" "finite t" + using assms + by (force simp: piecewise_C1_differentiable_on_def) + then have f_diff: "f differentiable_on {a..x. if x \ c then f x else g x)" + using continuous_on_cases [OF closed_real_atLeastAtMost [of a c], + OF closed_real_atLeastAtMost [of c b], + of f g "\x. x\c"] assms + by (force simp: ivl_disj_un_two_touch) + { fix x + assume x: "x \ {a..b} - insert c (s \ t)" + have "(\x. if x \ c then f x else g x) differentiable at x" (is "?diff_fg") + proof (cases x c rule: le_cases) + case le show ?diff_fg + apply (rule differentiable_transform_within [where f=f and d = "dist x c"]) + using x dist_real_def le st by (auto simp: C1_differentiable_on_eq) + next + case ge show ?diff_fg + apply (rule differentiable_transform_within [where f=g and d = "dist x c"]) + using dist_nz x dist_real_def ge st x by (auto simp: C1_differentiable_on_eq) + qed + } + then have "(\x \ {a..b} - insert c (s \ t). (\x. if x \ c then f x else g x) differentiable at x)" + by auto + moreover + { assume fcon: "continuous_on ({a<..x. vector_derivative f (at x))" + and gcon: "continuous_on ({c<..x. vector_derivative g (at x))" + have "open ({a<..x. vector_derivative (\x. if x \ c then f x else g x) (at x))" + apply (rule continuous_on_eq [OF fcon]) + apply (simp add:) + apply (rule vector_derivative_at [symmetric]) + apply (rule_tac f=f and d="dist x c" in has_vector_derivative_transform_within) + apply (simp_all add: dist_norm vector_derivative_works [symmetric]) + apply (metis (full_types) C1_differentiable_on_eq Diff_iff Groups.add_ac(2) add_mono_thms_linordered_field(5) atLeastAtMost_iff linorder_not_le order_less_irrefl st(1)) + apply auto + done + moreover have "continuous_on ({c<..x. vector_derivative (\x. if x \ c then f x else g x) (at x))" + apply (rule continuous_on_eq [OF gcon]) + apply (simp add:) + apply (rule vector_derivative_at [symmetric]) + apply (rule_tac f=g and d="dist x c" in has_vector_derivative_transform_within) + apply (simp_all add: dist_norm vector_derivative_works [symmetric]) + apply (metis (full_types) C1_differentiable_on_eq Diff_iff Groups.add_ac(2) add_mono_thms_linordered_field(5) atLeastAtMost_iff less_irrefl not_le st(2)) + apply auto + done + ultimately have "continuous_on ({a<.. t)) + (\x. vector_derivative (\x. if x \ c then f x else g x) (at x))" + apply (rule continuous_on_subset [OF continuous_on_open_Un], auto) + done + } note * = this + have "continuous_on ({a<.. t)) (\x. vector_derivative (\x. if x \ c then f x else g x) (at x))" + using st + by (auto simp: C1_differentiable_on_eq elim!: continuous_on_subset intro: *) + ultimately have "\s. finite s \ ((\x. if x \ c then f x else g x) C1_differentiable_on {a..b} - s)" + apply (rule_tac x="{a,b,c} \ s \ t" in exI) + using st by (auto simp: C1_differentiable_on_eq elim!: continuous_on_subset) + with cab show ?thesis + by (simp add: piecewise_C1_differentiable_on_def) +qed + +lemma piecewise_C1_differentiable_neg: + "f piecewise_C1_differentiable_on s \ (\x. -(f x)) piecewise_C1_differentiable_on s" + unfolding piecewise_C1_differentiable_on_def + by (auto intro!: continuous_on_minus C1_differentiable_on_minus) + +lemma piecewise_C1_differentiable_add: + assumes "f piecewise_C1_differentiable_on i" + "g piecewise_C1_differentiable_on i" + shows "(\x. f x + g x) piecewise_C1_differentiable_on i" +proof - + obtain s t where st: "finite s" "finite t" + "f C1_differentiable_on (i-s)" + "g C1_differentiable_on (i-t)" + using assms by (auto simp: piecewise_C1_differentiable_on_def) + then have "finite (s \ t) \ (\x. f x + g x) C1_differentiable_on i - (s \ t)" + by (auto intro: C1_differentiable_on_add elim!: C1_differentiable_on_subset) + moreover have "continuous_on i f" "continuous_on i g" + using assms piecewise_C1_differentiable_on_def by auto + ultimately show ?thesis + by (auto simp: piecewise_C1_differentiable_on_def continuous_on_add) +qed + +lemma piecewise_C1_differentiable_diff: + "\f piecewise_C1_differentiable_on s; g piecewise_C1_differentiable_on s\ + \ (\x. f x - g x) piecewise_C1_differentiable_on s" + unfolding diff_conv_add_uminus + by (metis piecewise_C1_differentiable_add piecewise_C1_differentiable_neg) + +lemma piecewise_C1_differentiable_D1: + fixes g1 :: "real \ 'a::real_normed_field" + assumes "(g1 +++ g2) piecewise_C1_differentiable_on {0..1}" + shows "g1 piecewise_C1_differentiable_on {0..1}" +proof - + obtain s where "finite s" + and co12: "continuous_on ({0..1} - s) (\x. vector_derivative (g1 +++ g2) (at x))" + and g12D: "\x\{0..1} - s. g1 +++ g2 differentiable at x" + using assms by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + then have g1D: "g1 differentiable at x" if "x \ {0..1} - insert 1 (op * 2 ` s)" for x + apply (rule_tac d="dist (x/2) (1/2)" and f = "(g1 +++ g2) o (op*(inverse 2))" in differentiable_transform_within) + using that + apply (simp_all add: dist_real_def joinpaths_def) + apply (rule differentiable_chain_at derivative_intros | force)+ + done + have [simp]: "vector_derivative (g1 \ op * 2) (at (x/2)) = 2 *\<^sub>R vector_derivative g1 (at x)" + if "x \ {0..1} - insert 1 (op * 2 ` s)" for x + apply (subst vector_derivative_chain_at) + using that + apply (rule derivative_eq_intros g1D | simp)+ + done + have "continuous_on ({0..1/2} - insert (1/2) s) (\x. vector_derivative (g1 +++ g2) (at x))" + using co12 by (rule continuous_on_subset) force + then have coDhalf: "continuous_on ({0..1/2} - insert (1/2) s) (\x. vector_derivative (g1 o op*2) (at x))" + apply (rule continuous_on_eq [OF _ vector_derivative_at]) + apply (rule_tac f="g1 o op*2" and d="dist x (1/2)" in has_vector_derivative_transform_within) + apply (simp_all add: dist_norm joinpaths_def vector_derivative_works [symmetric]) + apply (force intro: g1D differentiable_chain_at) + apply auto + done + have "continuous_on ({0..1} - insert 1 (op * 2 ` s)) + ((\x. 1/2 * vector_derivative (g1 o op*2) (at x)) o op*(1/2))" + apply (rule continuous_intros)+ + using coDhalf + apply (simp add: scaleR_conv_of_real image_set_diff image_image) + done + then have con_g1: "continuous_on ({0..1} - insert 1 (op * 2 ` s)) (\x. vector_derivative g1 (at x))" + by (rule continuous_on_eq) (simp add: scaleR_conv_of_real) + have "continuous_on {0..1} g1" + using continuous_on_joinpaths_D1 assms piecewise_C1_differentiable_on_def by blast + with \finite s\ show ?thesis + apply (clarsimp simp add: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + apply (rule_tac x="insert 1 ((op*2)`s)" in exI) + apply (simp add: g1D con_g1) + done +qed + +lemma piecewise_C1_differentiable_D2: + fixes g2 :: "real \ 'a::real_normed_field" + assumes "(g1 +++ g2) piecewise_C1_differentiable_on {0..1}" "pathfinish g1 = pathstart g2" + shows "g2 piecewise_C1_differentiable_on {0..1}" +proof - + obtain s where "finite s" + and co12: "continuous_on ({0..1} - s) (\x. vector_derivative (g1 +++ g2) (at x))" + and g12D: "\x\{0..1} - s. g1 +++ g2 differentiable at x" + using assms by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + then have g2D: "g2 differentiable at x" if "x \ {0..1} - insert 0 ((\x. 2*x-1) ` s)" for x + apply (rule_tac d="dist ((x+1)/2) (1/2)" and f = "(g1 +++ g2) o (\x. (x+1)/2)" in differentiable_transform_within) + using that + apply (simp_all add: dist_real_def joinpaths_def) + apply (auto simp: dist_real_def joinpaths_def field_simps) + apply (rule differentiable_chain_at derivative_intros | force)+ + apply (drule_tac x= "(x + 1) / 2" in bspec, force simp: divide_simps) + apply assumption + done + have [simp]: "vector_derivative (g2 \ (\x. 2*x-1)) (at ((x+1)/2)) = 2 *\<^sub>R vector_derivative g2 (at x)" + if "x \ {0..1} - insert 0 ((\x. 2*x-1) ` s)" for x + using that by (auto simp: vector_derivative_chain_at divide_simps g2D) + have "continuous_on ({1/2..1} - insert (1/2) s) (\x. vector_derivative (g1 +++ g2) (at x))" + using co12 by (rule continuous_on_subset) force + then have coDhalf: "continuous_on ({1/2..1} - insert (1/2) s) (\x. vector_derivative (g2 o (\x. 2*x-1)) (at x))" + apply (rule continuous_on_eq [OF _ vector_derivative_at]) + apply (rule_tac f="g2 o (\x. 2*x-1)" and d="dist (3/4) ((x+1)/2)" in has_vector_derivative_transform_within) + apply (auto simp: dist_real_def field_simps joinpaths_def vector_derivative_works [symmetric] + intro!: g2D differentiable_chain_at) + done + have [simp]: "((\x. (x + 1) / 2) ` ({0..1} - insert 0 ((\x. 2 * x - 1) ` s))) = ({1/2..1} - insert (1/2) s)" + apply (simp add: image_set_diff inj_on_def image_image) + apply (auto simp: image_affinity_atLeastAtMost_div add_divide_distrib) + done + have "continuous_on ({0..1} - insert 0 ((\x. 2*x-1) ` s)) + ((\x. 1/2 * vector_derivative (g2 \ (\x. 2*x-1)) (at x)) o (\x. (x+1)/2))" + by (rule continuous_intros | simp add: coDhalf)+ + then have con_g2: "continuous_on ({0..1} - insert 0 ((\x. 2*x-1) ` s)) (\x. vector_derivative g2 (at x))" + by (rule continuous_on_eq) (simp add: scaleR_conv_of_real) + have "continuous_on {0..1} g2" + using continuous_on_joinpaths_D2 assms piecewise_C1_differentiable_on_def by blast + with \finite s\ show ?thesis + apply (clarsimp simp add: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + apply (rule_tac x="insert 0 ((\x. 2 * x - 1) ` s)" in exI) + apply (simp add: g2D con_g2) + done +qed + +subsection \Valid paths, and their start and finish\ + +lemma Diff_Un_eq: "A - (B \ C) = A - B - C" + by blast + +definition valid_path :: "(real \ 'a :: real_normed_vector) \ bool" + where "valid_path f \ f piecewise_C1_differentiable_on {0..1::real}" + +definition closed_path :: "(real \ 'a :: real_normed_vector) \ bool" + where "closed_path g \ g 0 = g 1" + +subsubsection\In particular, all results for paths apply\ + +lemma valid_path_imp_path: "valid_path g \ path g" +by (simp add: path_def piecewise_C1_differentiable_on_def valid_path_def) + +lemma connected_valid_path_image: "valid_path g \ connected(path_image g)" + by (metis connected_path_image valid_path_imp_path) + +lemma compact_valid_path_image: "valid_path g \ compact(path_image g)" + by (metis compact_path_image valid_path_imp_path) + +lemma bounded_valid_path_image: "valid_path g \ bounded(path_image g)" + by (metis bounded_path_image valid_path_imp_path) + +lemma closed_valid_path_image: "valid_path g \ closed(path_image g)" + by (metis closed_path_image valid_path_imp_path) + +proposition valid_path_compose: + assumes "valid_path g" + and der: "\x. x \ path_image g \ \f'. (f has_field_derivative f') (at x)" + and con: "continuous_on (path_image g) (deriv f)" + shows "valid_path (f o g)" +proof - + obtain s where "finite s" and g_diff: "g C1_differentiable_on {0..1} - s" + using \valid_path g\ unfolding valid_path_def piecewise_C1_differentiable_on_def by auto + have "f \ g differentiable at t" when "t\{0..1} - s" for t + proof (rule differentiable_chain_at) + show "g differentiable at t" using \valid_path g\ + by (meson C1_differentiable_on_eq \g C1_differentiable_on {0..1} - s\ that) + next + have "g t\path_image g" using that DiffD1 image_eqI path_image_def by metis + then obtain f' where "(f has_field_derivative f') (at (g t))" + using der by auto + then have " (f has_derivative op * f') (at (g t))" + using has_field_derivative_imp_has_derivative[of f f' "at (g t)"] by auto + then show "f differentiable at (g t)" using differentiableI by auto + qed + moreover have "continuous_on ({0..1} - s) (\x. vector_derivative (f \ g) (at x))" + proof (rule continuous_on_eq [where f = "\x. vector_derivative g (at x) * deriv f (g x)"], + rule continuous_intros) + show "continuous_on ({0..1} - s) (\x. vector_derivative g (at x))" + using g_diff C1_differentiable_on_eq by auto + next + have "continuous_on {0..1} (\x. deriv f (g x))" + using continuous_on_compose[OF _ con[unfolded path_image_def],unfolded comp_def] + \valid_path g\ piecewise_C1_differentiable_on_def valid_path_def + by blast + then show "continuous_on ({0..1} - s) (\x. deriv f (g x))" + using continuous_on_subset by blast + next + show "vector_derivative g (at t) * deriv f (g t) = vector_derivative (f \ g) (at t)" + when "t \ {0..1} - s" for t + proof (rule vector_derivative_chain_at_general[symmetric]) + show "g differentiable at t" by (meson C1_differentiable_on_eq g_diff that) + next + have "g t\path_image g" using that DiffD1 image_eqI path_image_def by metis + then obtain f' where "(f has_field_derivative f') (at (g t))" + using der by auto + then show "\g'. (f has_field_derivative g') (at (g t))" by auto + qed + qed + ultimately have "f o g C1_differentiable_on {0..1} - s" + using C1_differentiable_on_eq by blast + moreover have "path (f o g)" + proof - + have "isCont f x" when "x\path_image g" for x + proof - + obtain f' where "(f has_field_derivative f') (at x)" + using der[rule_format] \x\path_image g\ by auto + thus ?thesis using DERIV_isCont by auto + qed + then have "continuous_on (path_image g) f" using continuous_at_imp_continuous_on by auto + then show ?thesis using path_continuous_image \valid_path g\ valid_path_imp_path by auto + qed + ultimately show ?thesis unfolding valid_path_def piecewise_C1_differentiable_on_def path_def + using \finite s\ by auto +qed + + +subsection\Contour Integrals along a path\ + +text\This definition is for complex numbers only, and does not generalise to line integrals in a vector field\ + +text\piecewise differentiable function on [0,1]\ + +definition has_contour_integral :: "(complex \ complex) \ complex \ (real \ complex) \ bool" + (infixr "has'_contour'_integral" 50) + where "(f has_contour_integral i) g \ + ((\x. f(g x) * vector_derivative g (at x within {0..1})) + has_integral i) {0..1}" + +definition contour_integrable_on + (infixr "contour'_integrable'_on" 50) + where "f contour_integrable_on g \ \i. (f has_contour_integral i) g" + +definition contour_integral + where "contour_integral g f \ @i. (f has_contour_integral i) g \ ~ f contour_integrable_on g \ i=0" + +lemma not_integrable_contour_integral: "~ f contour_integrable_on g \ contour_integral g f = 0" + unfolding contour_integrable_on_def contour_integral_def by blast + +lemma contour_integral_unique: "(f has_contour_integral i) g \ contour_integral g f = i" + apply (simp add: contour_integral_def has_contour_integral_def contour_integrable_on_def) + using has_integral_unique by blast + +corollary has_contour_integral_eqpath: + "\(f has_contour_integral y) p; f contour_integrable_on \; + contour_integral p f = contour_integral \ f\ + \ (f has_contour_integral y) \" +using contour_integrable_on_def contour_integral_unique by auto + +lemma has_contour_integral_integral: + "f contour_integrable_on i \ (f has_contour_integral (contour_integral i f)) i" + by (metis contour_integral_unique contour_integrable_on_def) + +lemma has_contour_integral_unique: + "(f has_contour_integral i) g \ (f has_contour_integral j) g \ i = j" + using has_integral_unique + by (auto simp: has_contour_integral_def) + +lemma has_contour_integral_integrable: "(f has_contour_integral i) g \ f contour_integrable_on g" + using contour_integrable_on_def by blast + +(* Show that we can forget about the localized derivative.*) + +lemma vector_derivative_within_interior: + "\x \ interior s; NO_MATCH UNIV s\ + \ vector_derivative f (at x within s) = vector_derivative f (at x)" + apply (simp add: vector_derivative_def has_vector_derivative_def has_derivative_def netlimit_within_interior) + apply (subst lim_within_interior, auto) + done + +lemma has_integral_localized_vector_derivative: + "((\x. f (g x) * vector_derivative g (at x within {a..b})) has_integral i) {a..b} \ + ((\x. f (g x) * vector_derivative g (at x)) has_integral i) {a..b}" +proof - + have "{a..b} - {a,b} = interior {a..b}" + by (simp add: atLeastAtMost_diff_ends) + show ?thesis + apply (rule has_integral_spike_eq [of "{a,b}"]) + apply (auto simp: vector_derivative_within_interior) + done +qed + +lemma integrable_on_localized_vector_derivative: + "(\x. f (g x) * vector_derivative g (at x within {a..b})) integrable_on {a..b} \ + (\x. f (g x) * vector_derivative g (at x)) integrable_on {a..b}" + by (simp add: integrable_on_def has_integral_localized_vector_derivative) + +lemma has_contour_integral: + "(f has_contour_integral i) g \ + ((\x. f (g x) * vector_derivative g (at x)) has_integral i) {0..1}" + by (simp add: has_integral_localized_vector_derivative has_contour_integral_def) + +lemma contour_integrable_on: + "f contour_integrable_on g \ + (\t. f(g t) * vector_derivative g (at t)) integrable_on {0..1}" + by (simp add: has_contour_integral integrable_on_def contour_integrable_on_def) + +subsection\Reversing a path\ + +lemma valid_path_imp_reverse: + assumes "valid_path g" + shows "valid_path(reversepath g)" +proof - + obtain s where "finite s" "g C1_differentiable_on ({0..1} - s)" + using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) + then have "finite (op - 1 ` s)" "(reversepath g C1_differentiable_on ({0..1} - op - 1 ` s))" + apply (auto simp: reversepath_def) + apply (rule C1_differentiable_compose [of "\x::real. 1-x" _ g, unfolded o_def]) + apply (auto simp: C1_differentiable_on_eq) + apply (rule continuous_intros, force) + apply (force elim!: continuous_on_subset) + apply (simp add: finite_vimageI inj_on_def) + done + then show ?thesis using assms + by (auto simp: valid_path_def piecewise_C1_differentiable_on_def path_def [symmetric]) +qed + +lemma valid_path_reversepath [simp]: "valid_path(reversepath g) \ valid_path g" + using valid_path_imp_reverse by force + +lemma has_contour_integral_reversepath: + assumes "valid_path g" "(f has_contour_integral i) g" + shows "(f has_contour_integral (-i)) (reversepath g)" +proof - + { fix s x + assume xs: "g C1_differentiable_on ({0..1} - s)" "x \ op - 1 ` s" "0 \ x" "x \ 1" + have "vector_derivative (\x. g (1 - x)) (at x within {0..1}) = + - vector_derivative g (at (1 - x) within {0..1})" + proof - + obtain f' where f': "(g has_vector_derivative f') (at (1 - x))" + using xs + by (force simp: has_vector_derivative_def C1_differentiable_on_def) + have "(g o (\x. 1 - x) has_vector_derivative -1 *\<^sub>R f') (at x)" + apply (rule vector_diff_chain_within) + apply (intro vector_diff_chain_within derivative_eq_intros | simp)+ + apply (rule has_vector_derivative_at_within [OF f']) + done + then have mf': "((\x. g (1 - x)) has_vector_derivative -f') (at x)" + by (simp add: o_def) + show ?thesis + using xs + by (auto simp: vector_derivative_at_within_ivl [OF mf'] vector_derivative_at_within_ivl [OF f']) + qed + } note * = this + have 01: "{0..1::real} = cbox 0 1" + by simp + show ?thesis using assms + apply (auto simp: has_contour_integral_def) + apply (drule has_integral_affinity01 [where m= "-1" and c=1]) + apply (auto simp: reversepath_def valid_path_def piecewise_C1_differentiable_on_def) + apply (drule has_integral_neg) + apply (rule_tac s = "(\x. 1 - x) ` s" in has_integral_spike_finite) + apply (auto simp: *) + done +qed + +lemma contour_integrable_reversepath: + "valid_path g \ f contour_integrable_on g \ f contour_integrable_on (reversepath g)" + using has_contour_integral_reversepath contour_integrable_on_def by blast + +lemma contour_integrable_reversepath_eq: + "valid_path g \ (f contour_integrable_on (reversepath g) \ f contour_integrable_on g)" + using contour_integrable_reversepath valid_path_reversepath by fastforce + +lemma contour_integral_reversepath: + assumes "valid_path g" + shows "contour_integral (reversepath g) f = - (contour_integral g f)" +proof (cases "f contour_integrable_on g") + case True then show ?thesis + by (simp add: assms contour_integral_unique has_contour_integral_integral has_contour_integral_reversepath) +next + case False then have "~ f contour_integrable_on (reversepath g)" + by (simp add: assms contour_integrable_reversepath_eq) + with False show ?thesis by (simp add: not_integrable_contour_integral) +qed + + +subsection\Joining two paths together\ + +lemma valid_path_join: + assumes "valid_path g1" "valid_path g2" "pathfinish g1 = pathstart g2" + shows "valid_path(g1 +++ g2)" +proof - + have "g1 1 = g2 0" + using assms by (auto simp: pathfinish_def pathstart_def) + moreover have "(g1 o (\x. 2*x)) piecewise_C1_differentiable_on {0..1/2}" + apply (rule piecewise_C1_differentiable_compose) + using assms + apply (auto simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_joinpaths) + apply (rule continuous_intros | simp)+ + apply (force intro: finite_vimageI [where h = "op*2"] inj_onI) + done + moreover have "(g2 o (\x. 2*x-1)) piecewise_C1_differentiable_on {1/2..1}" + apply (rule piecewise_C1_differentiable_compose) + using assms unfolding valid_path_def piecewise_C1_differentiable_on_def + by (auto intro!: continuous_intros finite_vimageI [where h = "(\x. 2*x - 1)"] inj_onI + simp: image_affinity_atLeastAtMost_diff continuous_on_joinpaths) + ultimately show ?thesis + apply (simp only: valid_path_def continuous_on_joinpaths joinpaths_def) + apply (rule piecewise_C1_differentiable_cases) + apply (auto simp: o_def) + done +qed + +lemma valid_path_join_D1: + fixes g1 :: "real \ 'a::real_normed_field" + shows "valid_path (g1 +++ g2) \ valid_path g1" + unfolding valid_path_def + by (rule piecewise_C1_differentiable_D1) + +lemma valid_path_join_D2: + fixes g2 :: "real \ 'a::real_normed_field" + shows "\valid_path (g1 +++ g2); pathfinish g1 = pathstart g2\ \ valid_path g2" + unfolding valid_path_def + by (rule piecewise_C1_differentiable_D2) + +lemma valid_path_join_eq [simp]: + fixes g2 :: "real \ 'a::real_normed_field" + shows "pathfinish g1 = pathstart g2 \ (valid_path(g1 +++ g2) \ valid_path g1 \ valid_path g2)" + using valid_path_join_D1 valid_path_join_D2 valid_path_join by blast + +lemma has_contour_integral_join: + assumes "(f has_contour_integral i1) g1" "(f has_contour_integral i2) g2" + "valid_path g1" "valid_path g2" + shows "(f has_contour_integral (i1 + i2)) (g1 +++ g2)" +proof - + obtain s1 s2 + where s1: "finite s1" "\x\{0..1} - s1. g1 differentiable at x" + and s2: "finite s2" "\x\{0..1} - s2. g2 differentiable at x" + using assms + by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + have 1: "((\x. f (g1 x) * vector_derivative g1 (at x)) has_integral i1) {0..1}" + and 2: "((\x. f (g2 x) * vector_derivative g2 (at x)) has_integral i2) {0..1}" + using assms + by (auto simp: has_contour_integral) + have i1: "((\x. (2*f (g1 (2*x))) * vector_derivative g1 (at (2*x))) has_integral i1) {0..1/2}" + and i2: "((\x. (2*f (g2 (2*x - 1))) * vector_derivative g2 (at (2*x - 1))) has_integral i2) {1/2..1}" + using has_integral_affinity01 [OF 1, where m= 2 and c=0, THEN has_integral_cmul [where c=2]] + has_integral_affinity01 [OF 2, where m= 2 and c="-1", THEN has_integral_cmul [where c=2]] + by (simp_all only: image_affinity_atLeastAtMost_div_diff, simp_all add: scaleR_conv_of_real mult_ac) + have g1: "\0 \ z; z*2 < 1; z*2 \ s1\ \ + vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at z) = + 2 *\<^sub>R vector_derivative g1 (at (z*2))" for z + apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g1(2*x))" and d = "\z - 1/2\"]]) + apply (simp_all add: dist_real_def abs_if split: if_split_asm) + apply (rule vector_diff_chain_at [of "\x. 2*x" 2 _ g1, simplified o_def]) + apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) + using s1 + apply (auto simp: algebra_simps vector_derivative_works) + done + have g2: "\1 < z*2; z \ 1; z*2 - 1 \ s2\ \ + vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at z) = + 2 *\<^sub>R vector_derivative g2 (at (z*2 - 1))" for z + apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g2 (2*x - 1))" and d = "\z - 1/2\"]]) + apply (simp_all add: dist_real_def abs_if split: if_split_asm) + apply (rule vector_diff_chain_at [of "\x. 2*x - 1" 2 _ g2, simplified o_def]) + apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) + using s2 + apply (auto simp: algebra_simps vector_derivative_works) + done + have "((\x. f ((g1 +++ g2) x) * vector_derivative (g1 +++ g2) (at x)) has_integral i1) {0..1/2}" + apply (rule has_integral_spike_finite [OF _ _ i1, of "insert (1/2) (op*2 -` s1)"]) + using s1 + apply (force intro: finite_vimageI [where h = "op*2"] inj_onI) + apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g1) + done + moreover have "((\x. f ((g1 +++ g2) x) * vector_derivative (g1 +++ g2) (at x)) has_integral i2) {1/2..1}" + apply (rule has_integral_spike_finite [OF _ _ i2, of "insert (1/2) ((\x. 2*x-1) -` s2)"]) + using s2 + apply (force intro: finite_vimageI [where h = "\x. 2*x-1"] inj_onI) + apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g2) + done + ultimately + show ?thesis + apply (simp add: has_contour_integral) + apply (rule has_integral_combine [where c = "1/2"], auto) + done +qed + +lemma contour_integrable_joinI: + assumes "f contour_integrable_on g1" "f contour_integrable_on g2" + "valid_path g1" "valid_path g2" + shows "f contour_integrable_on (g1 +++ g2)" + using assms + by (meson has_contour_integral_join contour_integrable_on_def) + +lemma contour_integrable_joinD1: + assumes "f contour_integrable_on (g1 +++ g2)" "valid_path g1" + shows "f contour_integrable_on g1" +proof - + obtain s1 + where s1: "finite s1" "\x\{0..1} - s1. g1 differentiable at x" + using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + have "(\x. f ((g1 +++ g2) (x/2)) * vector_derivative (g1 +++ g2) (at (x/2))) integrable_on {0..1}" + using assms + apply (auto simp: contour_integrable_on) + apply (drule integrable_on_subcbox [where a=0 and b="1/2"]) + apply (auto intro: integrable_affinity [of _ 0 "1/2::real" "1/2" 0, simplified]) + done + then have *: "(\x. (f ((g1 +++ g2) (x/2))/2) * vector_derivative (g1 +++ g2) (at (x/2))) integrable_on {0..1}" + by (auto dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real) + have g1: "\0 < z; z < 1; z \ s1\ \ + vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at (z/2)) = + 2 *\<^sub>R vector_derivative g1 (at z)" for z + apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g1(2*x))" and d = "\(z-1)/2\"]]) + apply (simp_all add: field_simps dist_real_def abs_if split: if_split_asm) + apply (rule vector_diff_chain_at [of "\x. x*2" 2 _ g1, simplified o_def]) + using s1 + apply (auto simp: vector_derivative_works has_vector_derivative_def has_derivative_def bounded_linear_mult_left) + done + show ?thesis + using s1 + apply (auto simp: contour_integrable_on) + apply (rule integrable_spike_finite [of "{0,1} \ s1", OF _ _ *]) + apply (auto simp: joinpaths_def scaleR_conv_of_real g1) + done +qed + +lemma contour_integrable_joinD2: + assumes "f contour_integrable_on (g1 +++ g2)" "valid_path g2" + shows "f contour_integrable_on g2" +proof - + obtain s2 + where s2: "finite s2" "\x\{0..1} - s2. g2 differentiable at x" + using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + have "(\x. f ((g1 +++ g2) (x/2 + 1/2)) * vector_derivative (g1 +++ g2) (at (x/2 + 1/2))) integrable_on {0..1}" + using assms + apply (auto simp: contour_integrable_on) + apply (drule integrable_on_subcbox [where a="1/2" and b=1], auto) + apply (drule integrable_affinity [of _ "1/2::real" 1 "1/2" "1/2", simplified]) + apply (simp add: image_affinity_atLeastAtMost_diff) + done + then have *: "(\x. (f ((g1 +++ g2) (x/2 + 1/2))/2) * vector_derivative (g1 +++ g2) (at (x/2 + 1/2))) + integrable_on {0..1}" + by (auto dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real) + have g2: "\0 < z; z < 1; z \ s2\ \ + vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at (z/2+1/2)) = + 2 *\<^sub>R vector_derivative g2 (at z)" for z + apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g2(2*x-1))" and d = "\z/2\"]]) + apply (simp_all add: field_simps dist_real_def abs_if split: if_split_asm) + apply (rule vector_diff_chain_at [of "\x. x*2-1" 2 _ g2, simplified o_def]) + using s2 + apply (auto simp: has_vector_derivative_def has_derivative_def bounded_linear_mult_left + vector_derivative_works add_divide_distrib) + done + show ?thesis + using s2 + apply (auto simp: contour_integrable_on) + apply (rule integrable_spike_finite [of "{0,1} \ s2", OF _ _ *]) + apply (auto simp: joinpaths_def scaleR_conv_of_real g2) + done +qed + +lemma contour_integrable_join [simp]: + shows + "\valid_path g1; valid_path g2\ + \ f contour_integrable_on (g1 +++ g2) \ f contour_integrable_on g1 \ f contour_integrable_on g2" +using contour_integrable_joinD1 contour_integrable_joinD2 contour_integrable_joinI by blast + +lemma contour_integral_join [simp]: + shows + "\f contour_integrable_on g1; f contour_integrable_on g2; valid_path g1; valid_path g2\ + \ contour_integral (g1 +++ g2) f = contour_integral g1 f + contour_integral g2 f" + by (simp add: has_contour_integral_integral has_contour_integral_join contour_integral_unique) + + +subsection\Shifting the starting point of a (closed) path\ + +lemma shiftpath_alt_def: "shiftpath a f = (\x. if x \ 1-a then f (a + x) else f (a + x - 1))" + by (auto simp: shiftpath_def) + +lemma valid_path_shiftpath [intro]: + assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" + shows "valid_path(shiftpath a g)" + using assms + apply (auto simp: valid_path_def shiftpath_alt_def) + apply (rule piecewise_C1_differentiable_cases) + apply (auto simp: algebra_simps) + apply (rule piecewise_C1_differentiable_affine [of g 1 a, simplified o_def scaleR_one]) + apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset) + apply (rule piecewise_C1_differentiable_affine [of g 1 "a-1", simplified o_def scaleR_one algebra_simps]) + apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset) + done + +lemma has_contour_integral_shiftpath: + assumes f: "(f has_contour_integral i) g" "valid_path g" + and a: "a \ {0..1}" + shows "(f has_contour_integral i) (shiftpath a g)" +proof - + obtain s + where s: "finite s" and g: "\x\{0..1} - s. g differentiable at x" + using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + have *: "((\x. f (g x) * vector_derivative g (at x)) has_integral i) {0..1}" + using assms by (auto simp: has_contour_integral) + then have i: "i = integral {a..1} (\x. f (g x) * vector_derivative g (at x)) + + integral {0..a} (\x. f (g x) * vector_derivative g (at x))" + apply (rule has_integral_unique) + apply (subst add.commute) + apply (subst integral_combine) + using assms * integral_unique by auto + { fix x + have "0 \ x \ x + a < 1 \ x \ (\x. x - a) ` s \ + vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a))" + unfolding shiftpath_def + apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g(a+x))" and d = "dist(1-a) x"]]) + apply (auto simp: field_simps dist_real_def abs_if split: if_split_asm) + apply (rule vector_diff_chain_at [of "\x. x+a" 1 _ g, simplified o_def scaleR_one]) + apply (intro derivative_eq_intros | simp)+ + using g + apply (drule_tac x="x+a" in bspec) + using a apply (auto simp: has_vector_derivative_def vector_derivative_works image_def add.commute) + done + } note vd1 = this + { fix x + have "1 < x + a \ x \ 1 \ x \ (\x. x - a + 1) ` s \ + vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a - 1))" + unfolding shiftpath_def + apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g(a+x-1))" and d = "dist (1-a) x"]]) + apply (auto simp: field_simps dist_real_def abs_if split: if_split_asm) + apply (rule vector_diff_chain_at [of "\x. x+a-1" 1 _ g, simplified o_def scaleR_one]) + apply (intro derivative_eq_intros | simp)+ + using g + apply (drule_tac x="x+a-1" in bspec) + using a apply (auto simp: has_vector_derivative_def vector_derivative_works image_def add.commute) + done + } note vd2 = this + have va1: "(\x. f (g x) * vector_derivative g (at x)) integrable_on ({a..1})" + using * a by (fastforce intro: integrable_subinterval_real) + have v0a: "(\x. f (g x) * vector_derivative g (at x)) integrable_on ({0..a})" + apply (rule integrable_subinterval_real) + using * a by auto + have "((\x. f (shiftpath a g x) * vector_derivative (shiftpath a g) (at x)) + has_integral integral {a..1} (\x. f (g x) * vector_derivative g (at x))) {0..1 - a}" + apply (rule has_integral_spike_finite + [where s = "{1-a} \ (\x. x-a) ` s" and f = "\x. f(g(a+x)) * vector_derivative g (at(a+x))"]) + using s apply blast + using a apply (auto simp: algebra_simps vd1) + apply (force simp: shiftpath_def add.commute) + using has_integral_affinity [where m=1 and c=a, simplified, OF integrable_integral [OF va1]] + apply (simp add: image_affinity_atLeastAtMost_diff [where m=1 and c=a, simplified] add.commute) + done + moreover + have "((\x. f (shiftpath a g x) * vector_derivative (shiftpath a g) (at x)) + has_integral integral {0..a} (\x. f (g x) * vector_derivative g (at x))) {1 - a..1}" + apply (rule has_integral_spike_finite + [where s = "{1-a} \ (\x. x-a+1) ` s" and f = "\x. f(g(a+x-1)) * vector_derivative g (at(a+x-1))"]) + using s apply blast + using a apply (auto simp: algebra_simps vd2) + apply (force simp: shiftpath_def add.commute) + using has_integral_affinity [where m=1 and c="a-1", simplified, OF integrable_integral [OF v0a]] + apply (simp add: image_affinity_atLeastAtMost [where m=1 and c="1-a", simplified]) + apply (simp add: algebra_simps) + done + ultimately show ?thesis + using a + by (auto simp: i has_contour_integral intro: has_integral_combine [where c = "1-a"]) +qed + +lemma has_contour_integral_shiftpath_D: + assumes "(f has_contour_integral i) (shiftpath a g)" + "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" + shows "(f has_contour_integral i) g" +proof - + obtain s + where s: "finite s" and g: "\x\{0..1} - s. g differentiable at x" + using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + { fix x + assume x: "0 < x" "x < 1" "x \ s" + then have gx: "g differentiable at x" + using g by auto + have "vector_derivative g (at x within {0..1}) = + vector_derivative (shiftpath (1 - a) (shiftpath a g)) (at x within {0..1})" + apply (rule vector_derivative_at_within_ivl + [OF has_vector_derivative_transform_within_open + [where f = "(shiftpath (1 - a) (shiftpath a g))" and s = "{0<..<1}-s"]]) + using s g assms x + apply (auto simp: finite_imp_closed open_Diff shiftpath_shiftpath + vector_derivative_within_interior vector_derivative_works [symmetric]) + apply (rule differentiable_transform_within [OF gx, of "min x (1-x)"]) + apply (auto simp: dist_real_def shiftpath_shiftpath abs_if split: if_split_asm) + done + } note vd = this + have fi: "(f has_contour_integral i) (shiftpath (1 - a) (shiftpath a g))" + using assms by (auto intro!: has_contour_integral_shiftpath) + show ?thesis + apply (simp add: has_contour_integral_def) + apply (rule has_integral_spike_finite [of "{0,1} \ s", OF _ _ fi [unfolded has_contour_integral_def]]) + using s assms vd + apply (auto simp: Path_Connected.shiftpath_shiftpath) + done +qed + +lemma has_contour_integral_shiftpath_eq: + assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" + shows "(f has_contour_integral i) (shiftpath a g) \ (f has_contour_integral i) g" + using assms has_contour_integral_shiftpath has_contour_integral_shiftpath_D by blast + +lemma contour_integrable_on_shiftpath_eq: + assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" + shows "f contour_integrable_on (shiftpath a g) \ f contour_integrable_on g" +using assms contour_integrable_on_def has_contour_integral_shiftpath_eq by auto + +lemma contour_integral_shiftpath: + assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" + shows "contour_integral (shiftpath a g) f = contour_integral g f" + using assms + by (simp add: contour_integral_def contour_integrable_on_def has_contour_integral_shiftpath_eq) + + +subsection\More about straight-line paths\ + +lemma has_vector_derivative_linepath_within: + "(linepath a b has_vector_derivative (b - a)) (at x within s)" +apply (simp add: linepath_def has_vector_derivative_def algebra_simps) +apply (rule derivative_eq_intros | simp)+ +done + +lemma vector_derivative_linepath_within: + "x \ {0..1} \ vector_derivative (linepath a b) (at x within {0..1}) = b - a" + apply (rule vector_derivative_within_closed_interval [of 0 "1::real", simplified]) + apply (auto simp: has_vector_derivative_linepath_within) + done + +lemma vector_derivative_linepath_at [simp]: "vector_derivative (linepath a b) (at x) = b - a" + by (simp add: has_vector_derivative_linepath_within vector_derivative_at) + +lemma valid_path_linepath [iff]: "valid_path (linepath a b)" + apply (simp add: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq continuous_on_linepath) + apply (rule_tac x="{}" in exI) + apply (simp add: differentiable_on_def differentiable_def) + using has_vector_derivative_def has_vector_derivative_linepath_within + apply (fastforce simp add: continuous_on_eq_continuous_within) + done + +lemma has_contour_integral_linepath: + shows "(f has_contour_integral i) (linepath a b) \ + ((\x. f(linepath a b x) * (b - a)) has_integral i) {0..1}" + by (simp add: has_contour_integral vector_derivative_linepath_at) + +lemma linepath_in_path: + shows "x \ {0..1} \ linepath a b x \ closed_segment a b" + by (auto simp: segment linepath_def) + +lemma linepath_image_01: "linepath a b ` {0..1} = closed_segment a b" + by (auto simp: segment linepath_def) + +lemma linepath_in_convex_hull: + fixes x::real + assumes a: "a \ convex hull s" + and b: "b \ convex hull s" + and x: "0\x" "x\1" + shows "linepath a b x \ convex hull s" + apply (rule closed_segment_subset_convex_hull [OF a b, THEN subsetD]) + using x + apply (auto simp: linepath_image_01 [symmetric]) + done + +lemma Re_linepath: "Re(linepath (of_real a) (of_real b) x) = (1 - x)*a + x*b" + by (simp add: linepath_def) + +lemma Im_linepath: "Im(linepath (of_real a) (of_real b) x) = 0" + by (simp add: linepath_def) + +lemma has_contour_integral_trivial [iff]: "(f has_contour_integral 0) (linepath a a)" + by (simp add: has_contour_integral_linepath) + +lemma contour_integral_trivial [simp]: "contour_integral (linepath a a) f = 0" + using has_contour_integral_trivial contour_integral_unique by blast + + +subsection\Relation to subpath construction\ + +lemma valid_path_subpath: + fixes g :: "real \ 'a :: real_normed_vector" + assumes "valid_path g" "u \ {0..1}" "v \ {0..1}" + shows "valid_path(subpath u v g)" +proof (cases "v=u") + case True + then show ?thesis + unfolding valid_path_def subpath_def + by (force intro: C1_differentiable_on_const C1_differentiable_imp_piecewise) +next + case False + have "(g o (\x. ((v-u) * x + u))) piecewise_C1_differentiable_on {0..1}" + apply (rule piecewise_C1_differentiable_compose) + apply (simp add: C1_differentiable_imp_piecewise) + apply (simp add: image_affinity_atLeastAtMost) + using assms False + apply (auto simp: algebra_simps valid_path_def piecewise_C1_differentiable_on_subset) + apply (subst Int_commute) + apply (auto simp: inj_on_def algebra_simps crossproduct_eq finite_vimage_IntI) + done + then show ?thesis + by (auto simp: o_def valid_path_def subpath_def) +qed + +lemma has_contour_integral_subpath_refl [iff]: "(f has_contour_integral 0) (subpath u u g)" + by (simp add: has_contour_integral subpath_def) + +lemma contour_integrable_subpath_refl [iff]: "f contour_integrable_on (subpath u u g)" + using has_contour_integral_subpath_refl contour_integrable_on_def by blast + +lemma contour_integral_subpath_refl [simp]: "contour_integral (subpath u u g) f = 0" + by (simp add: has_contour_integral_subpath_refl contour_integral_unique) + +lemma has_contour_integral_subpath: + assumes f: "f contour_integrable_on g" and g: "valid_path g" + and uv: "u \ {0..1}" "v \ {0..1}" "u \ v" + shows "(f has_contour_integral integral {u..v} (\x. f(g x) * vector_derivative g (at x))) + (subpath u v g)" +proof (cases "v=u") + case True + then show ?thesis + using f by (simp add: contour_integrable_on_def subpath_def has_contour_integral) +next + case False + obtain s where s: "\x. x \ {0..1} - s \ g differentiable at x" and fs: "finite s" + using g unfolding piecewise_C1_differentiable_on_def C1_differentiable_on_eq valid_path_def by blast + have *: "((\x. f (g ((v - u) * x + u)) * vector_derivative g (at ((v - u) * x + u))) + has_integral (1 / (v - u)) * integral {u..v} (\t. f (g t) * vector_derivative g (at t))) + {0..1}" + using f uv + apply (simp add: contour_integrable_on subpath_def has_contour_integral) + apply (drule integrable_on_subcbox [where a=u and b=v, simplified]) + apply (simp_all add: has_integral_integral) + apply (drule has_integral_affinity [where m="v-u" and c=u, simplified]) + apply (simp_all add: False image_affinity_atLeastAtMost_div_diff scaleR_conv_of_real) + apply (simp add: divide_simps False) + done + { fix x + have "x \ {0..1} \ + x \ (\t. (v-u) *\<^sub>R t + u) -` s \ + vector_derivative (\x. g ((v-u) * x + u)) (at x) = (v-u) *\<^sub>R vector_derivative g (at ((v-u) * x + u))" + apply (rule vector_derivative_at [OF vector_diff_chain_at [simplified o_def]]) + apply (intro derivative_eq_intros | simp)+ + apply (cut_tac s [of "(v - u) * x + u"]) + using uv mult_left_le [of x "v-u"] + apply (auto simp: vector_derivative_works) + done + } note vd = this + show ?thesis + apply (cut_tac has_integral_cmul [OF *, where c = "v-u"]) + using fs assms + apply (simp add: False subpath_def has_contour_integral) + apply (rule_tac s = "(\t. ((v-u) *\<^sub>R t + u)) -` s" in has_integral_spike_finite) + apply (auto simp: inj_on_def False finite_vimageI vd scaleR_conv_of_real) + done +qed + +lemma contour_integrable_subpath: + assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" + shows "f contour_integrable_on (subpath u v g)" + apply (cases u v rule: linorder_class.le_cases) + apply (metis contour_integrable_on_def has_contour_integral_subpath [OF assms]) + apply (subst reversepath_subpath [symmetric]) + apply (rule contour_integrable_reversepath) + using assms apply (blast intro: valid_path_subpath) + apply (simp add: contour_integrable_on_def) + using assms apply (blast intro: has_contour_integral_subpath) + done + +lemma has_integral_integrable_integral: "(f has_integral i) s \ f integrable_on s \ integral s f = i" + by blast + +lemma has_integral_contour_integral_subpath: + assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "u \ v" + shows "(((\x. f(g x) * vector_derivative g (at x))) + has_integral contour_integral (subpath u v g) f) {u..v}" + using assms + apply (auto simp: has_integral_integrable_integral) + apply (rule integrable_on_subcbox [where a=u and b=v and s = "{0..1}", simplified]) + apply (auto simp: contour_integral_unique [OF has_contour_integral_subpath] contour_integrable_on) + done + +lemma contour_integral_subcontour_integral: + assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "u \ v" + shows "contour_integral (subpath u v g) f = + integral {u..v} (\x. f(g x) * vector_derivative g (at x))" + using assms has_contour_integral_subpath contour_integral_unique by blast + +lemma contour_integral_subpath_combine_less: + assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "w \ {0..1}" + "u {0..1}" "v \ {0..1}" "w \ {0..1}" + shows "contour_integral (subpath u v g) f + contour_integral (subpath v w g) f = + contour_integral (subpath u w g) f" +proof (cases "u\v \ v\w \ u\w") + case True + have *: "subpath v u g = reversepath(subpath u v g) \ + subpath w u g = reversepath(subpath u w g) \ + subpath w v g = reversepath(subpath v w g)" + by (auto simp: reversepath_subpath) + have "u < v \ v < w \ + u < w \ w < v \ + v < u \ u < w \ + v < w \ w < u \ + w < u \ u < v \ + w < v \ v < u" + using True assms by linarith + with assms show ?thesis + using contour_integral_subpath_combine_less [of f g u v w] + contour_integral_subpath_combine_less [of f g u w v] + contour_integral_subpath_combine_less [of f g v u w] + contour_integral_subpath_combine_less [of f g v w u] + contour_integral_subpath_combine_less [of f g w u v] + contour_integral_subpath_combine_less [of f g w v u] + apply simp + apply (elim disjE) + apply (auto simp: * contour_integral_reversepath contour_integrable_subpath + valid_path_reversepath valid_path_subpath algebra_simps) + done +next + case False + then show ?thesis + apply (auto simp: contour_integral_subpath_refl) + using assms + by (metis eq_neg_iff_add_eq_0 contour_integrable_subpath contour_integral_reversepath reversepath_subpath valid_path_subpath) +qed + +lemma contour_integral_integral: + "contour_integral g f = integral {0..1} (\x. f (g x) * vector_derivative g (at x))" + by (simp add: contour_integral_def integral_def has_contour_integral contour_integrable_on) + + +text\Cauchy's theorem where there's a primitive\ + +lemma contour_integral_primitive_lemma: + fixes f :: "complex \ complex" and g :: "real \ complex" + assumes "a \ b" + and "\x. x \ s \ (f has_field_derivative f' x) (at x within s)" + and "g piecewise_differentiable_on {a..b}" "\x. x \ {a..b} \ g x \ s" + shows "((\x. f'(g x) * vector_derivative g (at x within {a..b})) + has_integral (f(g b) - f(g a))) {a..b}" +proof - + obtain k where k: "finite k" "\x\{a..b} - k. g differentiable (at x within {a..b})" and cg: "continuous_on {a..b} g" + using assms by (auto simp: piecewise_differentiable_on_def) + have cfg: "continuous_on {a..b} (\x. f (g x))" + apply (rule continuous_on_compose [OF cg, unfolded o_def]) + using assms + apply (metis field_differentiable_def field_differentiable_imp_continuous_at continuous_on_eq_continuous_within continuous_on_subset image_subset_iff) + done + { fix x::real + assume a: "a < x" and b: "x < b" and xk: "x \ k" + then have "g differentiable at x within {a..b}" + using k by (simp add: differentiable_at_withinI) + then have "(g has_vector_derivative vector_derivative g (at x within {a..b})) (at x within {a..b})" + by (simp add: vector_derivative_works has_field_derivative_def scaleR_conv_of_real) + then have gdiff: "(g has_derivative (\u. u * vector_derivative g (at x within {a..b}))) (at x within {a..b})" + by (simp add: has_vector_derivative_def scaleR_conv_of_real) + have "(f has_field_derivative (f' (g x))) (at (g x) within g ` {a..b})" + using assms by (metis a atLeastAtMost_iff b DERIV_subset image_subset_iff less_eq_real_def) + then have fdiff: "(f has_derivative op * (f' (g x))) (at (g x) within g ` {a..b})" + by (simp add: has_field_derivative_def) + have "((\x. f (g x)) has_vector_derivative f' (g x) * vector_derivative g (at x within {a..b})) (at x within {a..b})" + using diff_chain_within [OF gdiff fdiff] + by (simp add: has_vector_derivative_def scaleR_conv_of_real o_def mult_ac) + } note * = this + show ?thesis + apply (rule fundamental_theorem_of_calculus_interior_strong) + using k assms cfg * + apply (auto simp: at_within_closed_interval) + done +qed + +lemma contour_integral_primitive: + assumes "\x. x \ s \ (f has_field_derivative f' x) (at x within s)" + and "valid_path g" "path_image g \ s" + shows "(f' has_contour_integral (f(pathfinish g) - f(pathstart g))) g" + using assms + apply (simp add: valid_path_def path_image_def pathfinish_def pathstart_def has_contour_integral_def) + apply (auto intro!: piecewise_C1_imp_differentiable contour_integral_primitive_lemma [of 0 1 s]) + done + +corollary Cauchy_theorem_primitive: + assumes "\x. x \ s \ (f has_field_derivative f' x) (at x within s)" + and "valid_path g" "path_image g \ s" "pathfinish g = pathstart g" + shows "(f' has_contour_integral 0) g" + using assms + by (metis diff_self contour_integral_primitive) + + +text\Existence of path integral for continuous function\ +lemma contour_integrable_continuous_linepath: + assumes "continuous_on (closed_segment a b) f" + shows "f contour_integrable_on (linepath a b)" +proof - + have "continuous_on {0..1} ((\x. f x * (b - a)) o linepath a b)" + apply (rule continuous_on_compose [OF continuous_on_linepath], simp add: linepath_image_01) + apply (rule continuous_intros | simp add: assms)+ + done + then show ?thesis + apply (simp add: contour_integrable_on_def has_contour_integral_def integrable_on_def [symmetric]) + apply (rule integrable_continuous [of 0 "1::real", simplified]) + apply (rule continuous_on_eq [where f = "\x. f(linepath a b x)*(b - a)"]) + apply (auto simp: vector_derivative_linepath_within) + done +qed + +lemma has_field_der_id: "((\x. x\<^sup>2 / 2) has_field_derivative x) (at x)" + by (rule has_derivative_imp_has_field_derivative) + (rule derivative_intros | simp)+ + +lemma contour_integral_id [simp]: "contour_integral (linepath a b) (\y. y) = (b^2 - a^2)/2" + apply (rule contour_integral_unique) + using contour_integral_primitive [of UNIV "\x. x^2/2" "\x. x" "linepath a b"] + apply (auto simp: field_simps has_field_der_id) + done + +lemma contour_integrable_on_const [iff]: "(\x. c) contour_integrable_on (linepath a b)" + by (simp add: continuous_on_const contour_integrable_continuous_linepath) + +lemma contour_integrable_on_id [iff]: "(\x. x) contour_integrable_on (linepath a b)" + by (simp add: continuous_on_id contour_integrable_continuous_linepath) + + +subsection\Arithmetical combining theorems\ + +lemma has_contour_integral_neg: + "(f has_contour_integral i) g \ ((\x. -(f x)) has_contour_integral (-i)) g" + by (simp add: has_integral_neg has_contour_integral_def) + +lemma has_contour_integral_add: + "\(f1 has_contour_integral i1) g; (f2 has_contour_integral i2) g\ + \ ((\x. f1 x + f2 x) has_contour_integral (i1 + i2)) g" + by (simp add: has_integral_add has_contour_integral_def algebra_simps) + +lemma has_contour_integral_diff: + "\(f1 has_contour_integral i1) g; (f2 has_contour_integral i2) g\ + \ ((\x. f1 x - f2 x) has_contour_integral (i1 - i2)) g" + by (simp add: has_integral_sub has_contour_integral_def algebra_simps) + +lemma has_contour_integral_lmul: + "(f has_contour_integral i) g \ ((\x. c * (f x)) has_contour_integral (c*i)) g" +apply (simp add: has_contour_integral_def) +apply (drule has_integral_mult_right) +apply (simp add: algebra_simps) +done + +lemma has_contour_integral_rmul: + "(f has_contour_integral i) g \ ((\x. (f x) * c) has_contour_integral (i*c)) g" +apply (drule has_contour_integral_lmul) +apply (simp add: mult.commute) +done + +lemma has_contour_integral_div: + "(f has_contour_integral i) g \ ((\x. f x/c) has_contour_integral (i/c)) g" + by (simp add: field_class.field_divide_inverse) (metis has_contour_integral_rmul) + +lemma has_contour_integral_eq: + "\(f has_contour_integral y) p; \x. x \ path_image p \ f x = g x\ \ (g has_contour_integral y) p" +apply (simp add: path_image_def has_contour_integral_def) +by (metis (no_types, lifting) image_eqI has_integral_eq) + +lemma has_contour_integral_bound_linepath: + assumes "(f has_contour_integral i) (linepath a b)" + "0 \ B" "\x. x \ closed_segment a b \ norm(f x) \ B" + shows "norm i \ B * norm(b - a)" +proof - + { fix x::real + assume x: "0 \ x" "x \ 1" + have "norm (f (linepath a b x)) * + norm (vector_derivative (linepath a b) (at x within {0..1})) \ B * norm (b - a)" + by (auto intro: mult_mono simp: assms linepath_in_path of_real_linepath vector_derivative_linepath_within x) + } note * = this + have "norm i \ (B * norm (b - a)) * content (cbox 0 (1::real))" + apply (rule has_integral_bound + [of _ "\x. f (linepath a b x) * vector_derivative (linepath a b) (at x within {0..1})"]) + using assms * unfolding has_contour_integral_def + apply (auto simp: norm_mult) + done + then show ?thesis + by (auto simp: content_real) +qed + +(*UNUSED +lemma has_contour_integral_bound_linepath_strong: + fixes a :: real and f :: "complex \ real" + assumes "(f has_contour_integral i) (linepath a b)" + "finite k" + "0 \ B" "\x::real. x \ closed_segment a b - k \ norm(f x) \ B" + shows "norm i \ B*norm(b - a)" +*) + +lemma has_contour_integral_const_linepath: "((\x. c) has_contour_integral c*(b - a))(linepath a b)" + unfolding has_contour_integral_linepath + by (metis content_real diff_0_right has_integral_const_real lambda_one of_real_1 scaleR_conv_of_real zero_le_one) + +lemma has_contour_integral_0: "((\x. 0) has_contour_integral 0) g" + by (simp add: has_contour_integral_def) + +lemma has_contour_integral_is_0: + "(\z. z \ path_image g \ f z = 0) \ (f has_contour_integral 0) g" + by (rule has_contour_integral_eq [OF has_contour_integral_0]) auto + +lemma has_contour_integral_setsum: + "\finite s; \a. a \ s \ (f a has_contour_integral i a) p\ + \ ((\x. setsum (\a. f a x) s) has_contour_integral setsum i s) p" + by (induction s rule: finite_induct) (auto simp: has_contour_integral_0 has_contour_integral_add) + + +subsection \Operations on path integrals\ + +lemma contour_integral_const_linepath [simp]: "contour_integral (linepath a b) (\x. c) = c*(b - a)" + by (rule contour_integral_unique [OF has_contour_integral_const_linepath]) + +lemma contour_integral_neg: + "f contour_integrable_on g \ contour_integral g (\x. -(f x)) = -(contour_integral g f)" + by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_neg) + +lemma contour_integral_add: + "f1 contour_integrable_on g \ f2 contour_integrable_on g \ contour_integral g (\x. f1 x + f2 x) = + contour_integral g f1 + contour_integral g f2" + by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_add) + +lemma contour_integral_diff: + "f1 contour_integrable_on g \ f2 contour_integrable_on g \ contour_integral g (\x. f1 x - f2 x) = + contour_integral g f1 - contour_integral g f2" + by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_diff) + +lemma contour_integral_lmul: + shows "f contour_integrable_on g + \ contour_integral g (\x. c * f x) = c*contour_integral g f" + by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_lmul) + +lemma contour_integral_rmul: + shows "f contour_integrable_on g + \ contour_integral g (\x. f x * c) = contour_integral g f * c" + by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_rmul) + +lemma contour_integral_div: + shows "f contour_integrable_on g + \ contour_integral g (\x. f x / c) = contour_integral g f / c" + by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_div) + +lemma contour_integral_eq: + "(\x. x \ path_image p \ f x = g x) \ contour_integral p f = contour_integral p g" + apply (simp add: contour_integral_def) + using has_contour_integral_eq + by (metis contour_integral_unique has_contour_integral_integrable has_contour_integral_integral) + +lemma contour_integral_eq_0: + "(\z. z \ path_image g \ f z = 0) \ contour_integral g f = 0" + by (simp add: has_contour_integral_is_0 contour_integral_unique) + +lemma contour_integral_bound_linepath: + shows + "\f contour_integrable_on (linepath a b); + 0 \ B; \x. x \ closed_segment a b \ norm(f x) \ B\ + \ norm(contour_integral (linepath a b) f) \ B*norm(b - a)" + apply (rule has_contour_integral_bound_linepath [of f]) + apply (auto simp: has_contour_integral_integral) + done + +lemma contour_integral_0 [simp]: "contour_integral g (\x. 0) = 0" + by (simp add: contour_integral_unique has_contour_integral_0) + +lemma contour_integral_setsum: + "\finite s; \a. a \ s \ (f a) contour_integrable_on p\ + \ contour_integral p (\x. setsum (\a. f a x) s) = setsum (\a. contour_integral p (f a)) s" + by (auto simp: contour_integral_unique has_contour_integral_setsum has_contour_integral_integral) + +lemma contour_integrable_eq: + "\f contour_integrable_on p; \x. x \ path_image p \ f x = g x\ \ g contour_integrable_on p" + unfolding contour_integrable_on_def + by (metis has_contour_integral_eq) + + +subsection \Arithmetic theorems for path integrability\ + +lemma contour_integrable_neg: + "f contour_integrable_on g \ (\x. -(f x)) contour_integrable_on g" + using has_contour_integral_neg contour_integrable_on_def by blast + +lemma contour_integrable_add: + "\f1 contour_integrable_on g; f2 contour_integrable_on g\ \ (\x. f1 x + f2 x) contour_integrable_on g" + using has_contour_integral_add contour_integrable_on_def + by fastforce + +lemma contour_integrable_diff: + "\f1 contour_integrable_on g; f2 contour_integrable_on g\ \ (\x. f1 x - f2 x) contour_integrable_on g" + using has_contour_integral_diff contour_integrable_on_def + by fastforce + +lemma contour_integrable_lmul: + "f contour_integrable_on g \ (\x. c * f x) contour_integrable_on g" + using has_contour_integral_lmul contour_integrable_on_def + by fastforce + +lemma contour_integrable_rmul: + "f contour_integrable_on g \ (\x. f x * c) contour_integrable_on g" + using has_contour_integral_rmul contour_integrable_on_def + by fastforce + +lemma contour_integrable_div: + "f contour_integrable_on g \ (\x. f x / c) contour_integrable_on g" + using has_contour_integral_div contour_integrable_on_def + by fastforce + +lemma contour_integrable_setsum: + "\finite s; \a. a \ s \ (f a) contour_integrable_on p\ + \ (\x. setsum (\a. f a x) s) contour_integrable_on p" + unfolding contour_integrable_on_def + by (metis has_contour_integral_setsum) + + +subsection\Reversing a path integral\ + +lemma has_contour_integral_reverse_linepath: + "(f has_contour_integral i) (linepath a b) + \ (f has_contour_integral (-i)) (linepath b a)" + using has_contour_integral_reversepath valid_path_linepath by fastforce + +lemma contour_integral_reverse_linepath: + "continuous_on (closed_segment a b) f + \ contour_integral (linepath a b) f = - (contour_integral(linepath b a) f)" +apply (rule contour_integral_unique) +apply (rule has_contour_integral_reverse_linepath) +by (simp add: closed_segment_commute contour_integrable_continuous_linepath has_contour_integral_integral) + + +(* Splitting a path integral in a flat way.*) + +lemma has_contour_integral_split: + assumes f: "(f has_contour_integral i) (linepath a c)" "(f has_contour_integral j) (linepath c b)" + and k: "0 \ k" "k \ 1" + and c: "c - a = k *\<^sub>R (b - a)" + shows "(f has_contour_integral (i + j)) (linepath a b)" +proof (cases "k = 0 \ k = 1") + case True + then show ?thesis + using assms + apply auto + apply (metis add.left_neutral has_contour_integral_trivial has_contour_integral_unique) + apply (metis add.right_neutral has_contour_integral_trivial has_contour_integral_unique) + done +next + case False + then have k: "0 < k" "k < 1" "complex_of_real k \ 1" + using assms apply auto + using of_real_eq_iff by fastforce + have c': "c = k *\<^sub>R (b - a) + a" + by (metis diff_add_cancel c) + have bc: "(b - c) = (1 - k) *\<^sub>R (b - a)" + by (simp add: algebra_simps c') + { assume *: "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R c) * (c - a)) has_integral i) {0..1}" + have **: "\x. ((k - x) / k) *\<^sub>R a + (x / k) *\<^sub>R c = (1 - x) *\<^sub>R a + x *\<^sub>R b" + using False + apply (simp add: c' algebra_simps) + apply (simp add: real_vector.scale_left_distrib [symmetric] divide_simps) + done + have "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral i) {0..k}" + using * k + apply - + apply (drule has_integral_affinity [of _ _ 0 "1::real" "inverse k" "0", simplified]) + apply (simp_all add: divide_simps mult.commute [of _ "k"] image_affinity_atLeastAtMost ** c) + apply (drule has_integral_cmul [where c = "inverse k"]) + apply (simp add: has_integral_cmul) + done + } note fi = this + { assume *: "((\x. f ((1 - x) *\<^sub>R c + x *\<^sub>R b) * (b - c)) has_integral j) {0..1}" + have **: "\x. (((1 - x) / (1 - k)) *\<^sub>R c + ((x - k) / (1 - k)) *\<^sub>R b) = ((1 - x) *\<^sub>R a + x *\<^sub>R b)" + using k + apply (simp add: c' field_simps) + apply (simp add: scaleR_conv_of_real divide_simps) + apply (simp add: field_simps) + done + have "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral j) {k..1}" + using * k + apply - + apply (drule has_integral_affinity [of _ _ 0 "1::real" "inverse(1 - k)" "-(k/(1 - k))", simplified]) + apply (simp_all add: divide_simps mult.commute [of _ "1-k"] image_affinity_atLeastAtMost ** bc) + apply (drule has_integral_cmul [where k = "(1 - k) *\<^sub>R j" and c = "inverse (1 - k)"]) + apply (simp add: has_integral_cmul) + done + } note fj = this + show ?thesis + using f k + apply (simp add: has_contour_integral_linepath) + apply (simp add: linepath_def) + apply (rule has_integral_combine [OF _ _ fi fj], simp_all) + done +qed + +lemma continuous_on_closed_segment_transform: + assumes f: "continuous_on (closed_segment a b) f" + and k: "0 \ k" "k \ 1" + and c: "c - a = k *\<^sub>R (b - a)" + shows "continuous_on (closed_segment a c) f" +proof - + have c': "c = (1 - k) *\<^sub>R a + k *\<^sub>R b" + using c by (simp add: algebra_simps) + show "continuous_on (closed_segment a c) f" + apply (rule continuous_on_subset [OF f]) + apply (simp add: segment_convex_hull) + apply (rule convex_hull_subset) + using assms + apply (auto simp: hull_inc c' Convex.convexD_alt) + done +qed + +lemma contour_integral_split: + assumes f: "continuous_on (closed_segment a b) f" + and k: "0 \ k" "k \ 1" + and c: "c - a = k *\<^sub>R (b - a)" + shows "contour_integral(linepath a b) f = contour_integral(linepath a c) f + contour_integral(linepath c b) f" +proof - + have c': "c = (1 - k) *\<^sub>R a + k *\<^sub>R b" + using c by (simp add: algebra_simps) + have *: "continuous_on (closed_segment a c) f" "continuous_on (closed_segment c b) f" + apply (rule_tac [!] continuous_on_subset [OF f]) + apply (simp_all add: segment_convex_hull) + apply (rule_tac [!] convex_hull_subset) + using assms + apply (auto simp: hull_inc c' Convex.convexD_alt) + done + show ?thesis + apply (rule contour_integral_unique) + apply (rule has_contour_integral_split [OF has_contour_integral_integral has_contour_integral_integral k c]) + apply (rule contour_integrable_continuous_linepath *)+ + done +qed + +lemma contour_integral_split_linepath: + assumes f: "continuous_on (closed_segment a b) f" + and c: "c \ closed_segment a b" + shows "contour_integral(linepath a b) f = contour_integral(linepath a c) f + contour_integral(linepath c b) f" + using c + by (auto simp: closed_segment_def algebra_simps intro!: contour_integral_split [OF f]) + +(* The special case of midpoints used in the main quadrisection.*) + +lemma has_contour_integral_midpoint: + assumes "(f has_contour_integral i) (linepath a (midpoint a b))" + "(f has_contour_integral j) (linepath (midpoint a b) b)" + shows "(f has_contour_integral (i + j)) (linepath a b)" + apply (rule has_contour_integral_split [where c = "midpoint a b" and k = "1/2"]) + using assms + apply (auto simp: midpoint_def algebra_simps scaleR_conv_of_real) + done + +lemma contour_integral_midpoint: + "continuous_on (closed_segment a b) f + \ contour_integral (linepath a b) f = + contour_integral (linepath a (midpoint a b)) f + contour_integral (linepath (midpoint a b) b) f" + apply (rule contour_integral_split [where c = "midpoint a b" and k = "1/2"]) + apply (auto simp: midpoint_def algebra_simps scaleR_conv_of_real) + done + + +text\A couple of special case lemmas that are useful below\ + +lemma triangle_linear_has_chain_integral: + "((\x. m*x + d) has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" + apply (rule Cauchy_theorem_primitive [of UNIV "\x. m/2 * x^2 + d*x"]) + apply (auto intro!: derivative_eq_intros) + done + +lemma has_chain_integral_chain_integral3: + "(f has_contour_integral i) (linepath a b +++ linepath b c +++ linepath c d) + \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c d) f = i" + apply (subst contour_integral_unique [symmetric], assumption) + apply (drule has_contour_integral_integrable) + apply (simp add: valid_path_join) + done + +lemma has_chain_integral_chain_integral4: + "(f has_contour_integral i) (linepath a b +++ linepath b c +++ linepath c d +++ linepath d e) + \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c d) f + contour_integral (linepath d e) f = i" + apply (subst contour_integral_unique [symmetric], assumption) + apply (drule has_contour_integral_integrable) + apply (simp add: valid_path_join) + done + +subsection\Reversing the order in a double path integral\ + +text\The condition is stronger than needed but it's often true in typical situations\ + +lemma fst_im_cbox [simp]: "cbox c d \ {} \ (fst ` cbox (a,c) (b,d)) = cbox a b" + by (auto simp: cbox_Pair_eq) + +lemma snd_im_cbox [simp]: "cbox a b \ {} \ (snd ` cbox (a,c) (b,d)) = cbox c d" + by (auto simp: cbox_Pair_eq) + +lemma contour_integral_swap: + assumes fcon: "continuous_on (path_image g \ path_image h) (\(y1,y2). f y1 y2)" + and vp: "valid_path g" "valid_path h" + and gvcon: "continuous_on {0..1} (\t. vector_derivative g (at t))" + and hvcon: "continuous_on {0..1} (\t. vector_derivative h (at t))" + shows "contour_integral g (\w. contour_integral h (f w)) = + contour_integral h (\z. contour_integral g (\w. f w z))" +proof - + have gcon: "continuous_on {0..1} g" and hcon: "continuous_on {0..1} h" + using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) + have fgh1: "\x. (\t. f (g x) (h t)) = (\(y1,y2). f y1 y2) o (\t. (g x, h t))" + by (rule ext) simp + have fgh2: "\x. (\t. f (g t) (h x)) = (\(y1,y2). f y1 y2) o (\t. (g t, h x))" + by (rule ext) simp + have fcon_im1: "\x. 0 \ x \ x \ 1 \ continuous_on ((\t. (g x, h t)) ` {0..1}) (\(x, y). f x y)" + by (rule continuous_on_subset [OF fcon]) (auto simp: path_image_def) + have fcon_im2: "\x. 0 \ x \ x \ 1 \ continuous_on ((\t. (g t, h x)) ` {0..1}) (\(x, y). f x y)" + by (rule continuous_on_subset [OF fcon]) (auto simp: path_image_def) + have vdg: "\y. y \ {0..1} \ (\x. f (g x) (h y) * vector_derivative g (at x)) integrable_on {0..1}" + apply (rule integrable_continuous_real) + apply (rule continuous_on_mult [OF _ gvcon]) + apply (subst fgh2) + apply (rule fcon_im2 gcon continuous_intros | simp)+ + done + have "(\z. vector_derivative g (at (fst z))) = (\x. vector_derivative g (at x)) o fst" + by auto + then have gvcon': "continuous_on (cbox (0, 0) (1, 1::real)) (\x. vector_derivative g (at (fst x)))" + apply (rule ssubst) + apply (rule continuous_intros | simp add: gvcon)+ + done + have "(\z. vector_derivative h (at (snd z))) = (\x. vector_derivative h (at x)) o snd" + by auto + then have hvcon': "continuous_on (cbox (0, 0) (1::real, 1)) (\x. vector_derivative h (at (snd x)))" + apply (rule ssubst) + apply (rule continuous_intros | simp add: hvcon)+ + done + have "(\x. f (g (fst x)) (h (snd x))) = (\(y1,y2). f y1 y2) o (\w. ((g o fst) w, (h o snd) w))" + by auto + then have fgh: "continuous_on (cbox (0, 0) (1, 1)) (\x. f (g (fst x)) (h (snd x)))" + apply (rule ssubst) + apply (rule gcon hcon continuous_intros | simp)+ + apply (auto simp: path_image_def intro: continuous_on_subset [OF fcon]) + done + have "integral {0..1} (\x. contour_integral h (f (g x)) * vector_derivative g (at x)) = + integral {0..1} (\x. contour_integral h (\y. f (g x) y * vector_derivative g (at x)))" + apply (rule integral_cong [OF contour_integral_rmul [symmetric]]) + apply (clarsimp simp: contour_integrable_on) + apply (rule integrable_continuous_real) + apply (rule continuous_on_mult [OF _ hvcon]) + apply (subst fgh1) + apply (rule fcon_im1 hcon continuous_intros | simp)+ + done + also have "... = integral {0..1} + (\y. contour_integral g (\x. f x (h y) * vector_derivative h (at y)))" + apply (simp only: contour_integral_integral) + apply (subst integral_swap_continuous [where 'a = real and 'b = real, of 0 0 1 1, simplified]) + apply (rule fgh gvcon' hvcon' continuous_intros | simp add: split_def)+ + unfolding integral_mult_left [symmetric] + apply (simp only: mult_ac) + done + also have "... = contour_integral h (\z. contour_integral g (\w. f w z))" + apply (simp add: contour_integral_integral) + apply (rule integral_cong) + unfolding integral_mult_left [symmetric] + apply (simp add: algebra_simps) + done + finally show ?thesis + by (simp add: contour_integral_integral) +qed + + +subsection\The key quadrisection step\ + +lemma norm_sum_half: + assumes "norm(a + b) >= e" + shows "norm a >= e/2 \ norm b >= e/2" +proof - + have "e \ norm (- a - b)" + by (simp add: add.commute assms norm_minus_commute) + thus ?thesis + using norm_triangle_ineq4 order_trans by fastforce +qed + +lemma norm_sum_lemma: + assumes "e \ norm (a + b + c + d)" + shows "e / 4 \ norm a \ e / 4 \ norm b \ e / 4 \ norm c \ e / 4 \ norm d" +proof - + have "e \ norm ((a + b) + (c + d))" using assms + by (simp add: algebra_simps) + then show ?thesis + by (auto dest!: norm_sum_half) +qed + +lemma Cauchy_theorem_quadrisection: + assumes f: "continuous_on (convex hull {a,b,c}) f" + and dist: "dist a b \ K" "dist b c \ K" "dist c a \ K" + and e: "e * K^2 \ + norm (contour_integral(linepath a b) f + contour_integral(linepath b c) f + contour_integral(linepath c a) f)" + shows "\a' b' c'. + a' \ convex hull {a,b,c} \ b' \ convex hull {a,b,c} \ c' \ convex hull {a,b,c} \ + dist a' b' \ K/2 \ dist b' c' \ K/2 \ dist c' a' \ K/2 \ + e * (K/2)^2 \ norm(contour_integral(linepath a' b') f + contour_integral(linepath b' c') f + contour_integral(linepath c' a') f)" +proof - + note divide_le_eq_numeral1 [simp del] + define a' where "a' = midpoint b c" + define b' where "b' = midpoint c a" + define c' where "c' = midpoint a b" + have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" + using f continuous_on_subset segments_subset_convex_hull by metis+ + have fcont': "continuous_on (closed_segment c' b') f" + "continuous_on (closed_segment a' c') f" + "continuous_on (closed_segment b' a') f" + unfolding a'_def b'_def c'_def + apply (rule continuous_on_subset [OF f], + metis midpoints_in_convex_hull convex_hull_subset hull_subset insert_subset segment_convex_hull)+ + done + let ?pathint = "\x y. contour_integral(linepath x y) f" + have *: "?pathint a b + ?pathint b c + ?pathint c a = + (?pathint a c' + ?pathint c' b' + ?pathint b' a) + + (?pathint a' c' + ?pathint c' b + ?pathint b a') + + (?pathint a' c + ?pathint c b' + ?pathint b' a') + + (?pathint a' b' + ?pathint b' c' + ?pathint c' a')" + apply (simp add: fcont' contour_integral_reverse_linepath) + apply (simp add: a'_def b'_def c'_def contour_integral_midpoint fabc) + done + have [simp]: "\x y. cmod (x * 2 - y * 2) = cmod (x - y) * 2" + by (metis left_diff_distrib mult.commute norm_mult_numeral1) + have [simp]: "\x y. cmod (x - y) = cmod (y - x)" + by (simp add: norm_minus_commute) + consider "e * K\<^sup>2 / 4 \ cmod (?pathint a c' + ?pathint c' b' + ?pathint b' a)" | + "e * K\<^sup>2 / 4 \ cmod (?pathint a' c' + ?pathint c' b + ?pathint b a')" | + "e * K\<^sup>2 / 4 \ cmod (?pathint a' c + ?pathint c b' + ?pathint b' a')" | + "e * K\<^sup>2 / 4 \ cmod (?pathint a' b' + ?pathint b' c' + ?pathint c' a')" + using assms + apply (simp only: *) + apply (blast intro: that dest!: norm_sum_lemma) + done + then show ?thesis + proof cases + case 1 then show ?thesis + apply (rule_tac x=a in exI) + apply (rule exI [where x=c']) + apply (rule exI [where x=b']) + using assms + apply (auto simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) + apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real divide_simps) + done + next + case 2 then show ?thesis + apply (rule_tac x=a' in exI) + apply (rule exI [where x=c']) + apply (rule exI [where x=b]) + using assms + apply (auto simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) + apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real divide_simps) + done + next + case 3 then show ?thesis + apply (rule_tac x=a' in exI) + apply (rule exI [where x=c]) + apply (rule exI [where x=b']) + using assms + apply (auto simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) + apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real divide_simps) + done + next + case 4 then show ?thesis + apply (rule_tac x=a' in exI) + apply (rule exI [where x=b']) + apply (rule exI [where x=c']) + using assms + apply (auto simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) + apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real divide_simps) + done + qed +qed + +subsection\Cauchy's theorem for triangles\ + +lemma triangle_points_closer: + fixes a::complex + shows "\x \ convex hull {a,b,c}; y \ convex hull {a,b,c}\ + \ norm(x - y) \ norm(a - b) \ + norm(x - y) \ norm(b - c) \ + norm(x - y) \ norm(c - a)" + using simplex_extremal_le [of "{a,b,c}"] + by (auto simp: norm_minus_commute) + +lemma holomorphic_point_small_triangle: + assumes x: "x \ s" + and f: "continuous_on s f" + and cd: "f field_differentiable (at x within s)" + and e: "0 < e" + shows "\k>0. \a b c. dist a b \ k \ dist b c \ k \ dist c a \ k \ + x \ convex hull {a,b,c} \ convex hull {a,b,c} \ s + \ norm(contour_integral(linepath a b) f + contour_integral(linepath b c) f + + contour_integral(linepath c a) f) + \ e*(dist a b + dist b c + dist c a)^2" + (is "\k>0. \a b c. _ \ ?normle a b c") +proof - + have le_of_3: "\a x y z. \0 \ x*y; 0 \ x*z; 0 \ y*z; a \ (e*(x + y + z))*x + (e*(x + y + z))*y + (e*(x + y + z))*z\ + \ a \ e*(x + y + z)^2" + by (simp add: algebra_simps power2_eq_square) + have disj_le: "\x \ a \ x \ b \ x \ c; 0 \ a; 0 \ b; 0 \ c\ \ x \ a + b + c" + for x::real and a b c + by linarith + have fabc: "f contour_integrable_on linepath a b" "f contour_integrable_on linepath b c" "f contour_integrable_on linepath c a" + if "convex hull {a, b, c} \ s" for a b c + using segments_subset_convex_hull that + by (metis continuous_on_subset f contour_integrable_continuous_linepath)+ + note path_bound = has_contour_integral_bound_linepath [simplified norm_minus_commute, OF has_contour_integral_integral] + { fix f' a b c d + assume d: "0 < d" + and f': "\y. \cmod (y - x) \ d; y \ s\ \ cmod (f y - f x - f' * (y - x)) \ e * cmod (y - x)" + and le: "cmod (a - b) \ d" "cmod (b - c) \ d" "cmod (c - a) \ d" + and xc: "x \ convex hull {a, b, c}" + and s: "convex hull {a, b, c} \ s" + have pa: "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = + contour_integral (linepath a b) (\y. f y - f x - f'*(y - x)) + + contour_integral (linepath b c) (\y. f y - f x - f'*(y - x)) + + contour_integral (linepath c a) (\y. f y - f x - f'*(y - x))" + apply (simp add: contour_integral_diff contour_integral_lmul contour_integrable_lmul contour_integrable_diff fabc [OF s]) + apply (simp add: field_simps) + done + { fix y + assume yc: "y \ convex hull {a,b,c}" + have "cmod (f y - f x - f' * (y - x)) \ e*norm(y - x)" + apply (rule f') + apply (metis triangle_points_closer [OF xc yc] le norm_minus_commute order_trans) + using s yc by blast + also have "... \ e * (cmod (a - b) + cmod (b - c) + cmod (c - a))" + by (simp add: yc e xc disj_le [OF triangle_points_closer]) + finally have "cmod (f y - f x - f' * (y - x)) \ e * (cmod (a - b) + cmod (b - c) + cmod (c - a))" . + } note cm_le = this + have "?normle a b c" + apply (simp add: dist_norm pa) + apply (rule le_of_3) + using f' xc s e + apply simp_all + apply (intro norm_triangle_le add_mono path_bound) + apply (simp_all add: contour_integral_diff contour_integral_lmul contour_integrable_lmul contour_integrable_diff fabc) + apply (blast intro: cm_le elim: dest: segments_subset_convex_hull [THEN subsetD])+ + done + } note * = this + show ?thesis + using cd e + apply (simp add: field_differentiable_def has_field_derivative_def has_derivative_within_alt approachable_lt_le2 Ball_def) + apply (clarify dest!: spec mp) + using * + apply (simp add: dist_norm, blast) + done +qed + + +(* Hence the most basic theorem for a triangle.*) +locale Chain = + fixes x0 At Follows + assumes At0: "At x0 0" + and AtSuc: "\x n. At x n \ \x'. At x' (Suc n) \ Follows x' x" +begin + primrec f where + "f 0 = x0" + | "f (Suc n) = (SOME x. At x (Suc n) \ Follows x (f n))" + + lemma At: "At (f n) n" + proof (induct n) + case 0 show ?case + by (simp add: At0) + next + case (Suc n) show ?case + by (metis (no_types, lifting) AtSuc [OF Suc] f.simps(2) someI_ex) + qed + + lemma Follows: "Follows (f(Suc n)) (f n)" + by (metis (no_types, lifting) AtSuc [OF At [of n]] f.simps(2) someI_ex) + + declare f.simps(2) [simp del] +end + +lemma Chain3: + assumes At0: "At x0 y0 z0 0" + and AtSuc: "\x y z n. At x y z n \ \x' y' z'. At x' y' z' (Suc n) \ Follows x' y' z' x y z" + obtains f g h where + "f 0 = x0" "g 0 = y0" "h 0 = z0" + "\n. At (f n) (g n) (h n) n" + "\n. Follows (f(Suc n)) (g(Suc n)) (h(Suc n)) (f n) (g n) (h n)" +proof - + interpret three: Chain "(x0,y0,z0)" "\(x,y,z). At x y z" "\(x',y',z'). \(x,y,z). Follows x' y' z' x y z" + apply unfold_locales + using At0 AtSuc by auto + show ?thesis + apply (rule that [of "\n. fst (three.f n)" "\n. fst (snd (three.f n))" "\n. snd (snd (three.f n))"]) + apply simp_all + using three.At three.Follows + apply (simp_all add: split_beta') + done +qed + +lemma Cauchy_theorem_triangle: + assumes "f holomorphic_on (convex hull {a,b,c})" + shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" +proof - + have contf: "continuous_on (convex hull {a,b,c}) f" + by (metis assms holomorphic_on_imp_continuous_on) + let ?pathint = "\x y. contour_integral(linepath x y) f" + { fix y::complex + assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)" + and ynz: "y \ 0" + define K where "K = 1 + max (dist a b) (max (dist b c) (dist c a))" + define e where "e = norm y / K^2" + have K1: "K \ 1" by (simp add: K_def max.coboundedI1) + then have K: "K > 0" by linarith + have [iff]: "dist a b \ K" "dist b c \ K" "dist c a \ K" + by (simp_all add: K_def) + have e: "e > 0" + unfolding e_def using ynz K1 by simp + define At where "At x y z n \ + convex hull {x,y,z} \ convex hull {a,b,c} \ + dist x y \ K/2^n \ dist y z \ K/2^n \ dist z x \ K/2^n \ + norm(?pathint x y + ?pathint y z + ?pathint z x) \ e*(K/2^n)^2" + for x y z n + have At0: "At a b c 0" + using fy + by (simp add: At_def e_def has_chain_integral_chain_integral3) + { fix x y z n + assume At: "At x y z n" + then have contf': "continuous_on (convex hull {x,y,z}) f" + using contf At_def continuous_on_subset by blast + have "\x' y' z'. At x' y' z' (Suc n) \ convex hull {x',y',z'} \ convex hull {x,y,z}" + using At + apply (simp add: At_def) + using Cauchy_theorem_quadrisection [OF contf', of "K/2^n" e] + apply clarsimp + apply (rule_tac x="a'" in exI) + apply (rule_tac x="b'" in exI) + apply (rule_tac x="c'" in exI) + apply (simp add: algebra_simps) + apply (meson convex_hull_subset empty_subsetI insert_subset subsetCE) + done + } note AtSuc = this + obtain fa fb fc + where f0 [simp]: "fa 0 = a" "fb 0 = b" "fc 0 = c" + and cosb: "\n. convex hull {fa n, fb n, fc n} \ convex hull {a,b,c}" + and dist: "\n. dist (fa n) (fb n) \ K/2^n" + "\n. dist (fb n) (fc n) \ K/2^n" + "\n. dist (fc n) (fa n) \ K/2^n" + and no: "\n. norm(?pathint (fa n) (fb n) + + ?pathint (fb n) (fc n) + + ?pathint (fc n) (fa n)) \ e * (K/2^n)^2" + and conv_le: "\n. convex hull {fa(Suc n), fb(Suc n), fc(Suc n)} \ convex hull {fa n, fb n, fc n}" + apply (rule Chain3 [of At, OF At0 AtSuc]) + apply (auto simp: At_def) + done + have "\x. \n. x \ convex hull {fa n, fb n, fc n}" + apply (rule bounded_closed_nest) + apply (simp_all add: compact_imp_closed finite_imp_compact_convex_hull finite_imp_bounded_convex_hull) + apply (rule allI) + apply (rule transitive_stepwise_le) + apply (auto simp: conv_le) + done + then obtain x where x: "\n. x \ convex hull {fa n, fb n, fc n}" by auto + then have xin: "x \ convex hull {a,b,c}" + using assms f0 by blast + then have fx: "f field_differentiable at x within (convex hull {a,b,c})" + using assms holomorphic_on_def by blast + { fix k n + assume k: "0 < k" + and le: + "\x' y' z'. + \dist x' y' \ k; dist y' z' \ k; dist z' x' \ k; + x \ convex hull {x',y',z'}; + convex hull {x',y',z'} \ convex hull {a,b,c}\ + \ + cmod (?pathint x' y' + ?pathint y' z' + ?pathint z' x') * 10 + \ e * (dist x' y' + dist y' z' + dist z' x')\<^sup>2" + and Kk: "K / k < 2 ^ n" + have "K / 2 ^ n < k" using Kk k + by (auto simp: field_simps) + then have DD: "dist (fa n) (fb n) \ k" "dist (fb n) (fc n) \ k" "dist (fc n) (fa n) \ k" + using dist [of n] k + by linarith+ + have dle: "(dist (fa n) (fb n) + dist (fb n) (fc n) + dist (fc n) (fa n))\<^sup>2 + \ (3 * K / 2 ^ n)\<^sup>2" + using dist [of n] e K + by (simp add: abs_le_square_iff [symmetric]) + have less10: "\x y::real. 0 < x \ y \ 9*x \ y < x*10" + by linarith + have "e * (dist (fa n) (fb n) + dist (fb n) (fc n) + dist (fc n) (fa n))\<^sup>2 \ e * (3 * K / 2 ^ n)\<^sup>2" + using ynz dle e mult_le_cancel_left_pos by blast + also have "... < + cmod (?pathint (fa n) (fb n) + ?pathint (fb n) (fc n) + ?pathint (fc n) (fa n)) * 10" + using no [of n] e K + apply (simp add: e_def field_simps) + apply (simp only: zero_less_norm_iff [symmetric]) + done + finally have False + using le [OF DD x cosb] by auto + } then + have ?thesis + using holomorphic_point_small_triangle [OF xin contf fx, of "e/10"] e + apply clarsimp + apply (rule_tac y1="K/k" in exE [OF real_arch_pow[of 2]]) + apply force+ + done + } + moreover have "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)" + by simp (meson contf continuous_on_subset contour_integrable_continuous_linepath segments_subset_convex_hull(1) + segments_subset_convex_hull(3) segments_subset_convex_hull(5)) + ultimately show ?thesis + using has_contour_integral_integral by fastforce +qed + + +subsection\Version needing function holomorphic in interior only\ + +lemma Cauchy_theorem_flat_lemma: + assumes f: "continuous_on (convex hull {a,b,c}) f" + and c: "c - a = k *\<^sub>R (b - a)" + and k: "0 \ k" + shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0" +proof - + have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" + using f continuous_on_subset segments_subset_convex_hull by metis+ + show ?thesis + proof (cases "k \ 1") + case True show ?thesis + by (simp add: contour_integral_split [OF fabc(1) k True c] contour_integral_reverse_linepath fabc) + next + case False then show ?thesis + using fabc c + apply (subst contour_integral_split [of a c f "1/k" b, symmetric]) + apply (metis closed_segment_commute fabc(3)) + apply (auto simp: k contour_integral_reverse_linepath) + done + qed +qed + +lemma Cauchy_theorem_flat: + assumes f: "continuous_on (convex hull {a,b,c}) f" + and c: "c - a = k *\<^sub>R (b - a)" + shows "contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0" +proof (cases "0 \ k") + case True with assms show ?thesis + by (blast intro: Cauchy_theorem_flat_lemma) +next + case False + have "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" + using f continuous_on_subset segments_subset_convex_hull by metis+ + moreover have "contour_integral (linepath b a) f + contour_integral (linepath a c) f + + contour_integral (linepath c b) f = 0" + apply (rule Cauchy_theorem_flat_lemma [of b a c f "1-k"]) + using False c + apply (auto simp: f insert_commute scaleR_conv_of_real algebra_simps) + done + ultimately show ?thesis + apply (auto simp: contour_integral_reverse_linepath) + using add_eq_0_iff by force +qed + + +lemma Cauchy_theorem_triangle_interior: + assumes contf: "continuous_on (convex hull {a,b,c}) f" + and holf: "f holomorphic_on interior (convex hull {a,b,c})" + shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" +proof - + have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" + using contf continuous_on_subset segments_subset_convex_hull by metis+ + have "bounded (f ` (convex hull {a,b,c}))" + by (simp add: compact_continuous_image compact_convex_hull compact_imp_bounded contf) + then obtain B where "0 < B" and Bnf: "\x. x \ convex hull {a,b,c} \ norm (f x) \ B" + by (auto simp: dest!: bounded_pos [THEN iffD1]) + have "bounded (convex hull {a,b,c})" + by (simp add: bounded_convex_hull) + then obtain C where C: "0 < C" and Cno: "\y. y \ convex hull {a,b,c} \ norm y < C" + using bounded_pos_less by blast + then have diff_2C: "norm(x - y) \ 2*C" + if x: "x \ convex hull {a, b, c}" and y: "y \ convex hull {a, b, c}" for x y + proof - + have "cmod x \ C" + using x by (meson Cno not_le not_less_iff_gr_or_eq) + hence "cmod (x - y) \ C + C" + using y by (meson Cno add_mono_thms_linordered_field(4) less_eq_real_def norm_triangle_ineq4 order_trans) + thus "cmod (x - y) \ 2 * C" + by (metis mult_2) + qed + have contf': "continuous_on (convex hull {b,a,c}) f" + using contf by (simp add: insert_commute) + { fix y::complex + assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)" + and ynz: "y \ 0" + have pi_eq_y: "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = y" + by (rule has_chain_integral_chain_integral3 [OF fy]) + have ?thesis + proof (cases "c=a \ a=b \ b=c") + case True then show ?thesis + using Cauchy_theorem_flat [OF contf, of 0] + using has_chain_integral_chain_integral3 [OF fy] ynz + by (force simp: fabc contour_integral_reverse_linepath) + next + case False + then have car3: "card {a, b, c} = Suc (DIM(complex))" + by auto + { assume "interior(convex hull {a,b,c}) = {}" + then have "collinear{a,b,c}" + using interior_convex_hull_eq_empty [OF car3] + by (simp add: collinear_3_eq_affine_dependent) + then have "False" + using False + apply (clarsimp simp add: collinear_3 collinear_lemma) + apply (drule Cauchy_theorem_flat [OF contf']) + using pi_eq_y ynz + apply (simp add: fabc add_eq_0_iff contour_integral_reverse_linepath) + done + } + then obtain d where d: "d \ interior (convex hull {a, b, c})" + by blast + { fix d1 + assume d1_pos: "0 < d1" + and d1: "\x x'. \x\convex hull {a, b, c}; x'\convex hull {a, b, c}; cmod (x' - x) < d1\ + \ cmod (f x' - f x) < cmod y / (24 * C)" + define e where "e = min 1 (min (d1/(4*C)) ((norm y / 24 / C) / B))" + define shrink where "shrink x = x - e *\<^sub>R (x - d)" for x + let ?pathint = "\x y. contour_integral(linepath x y) f" + have e: "0 < e" "e \ 1" "e \ d1 / (4 * C)" "e \ cmod y / 24 / C / B" + using d1_pos \C>0\ \B>0\ ynz by (simp_all add: e_def) + then have eCB: "24 * e * C * B \ cmod y" + using \C>0\ \B>0\ by (simp add: field_simps) + have e_le_d1: "e * (4 * C) \ d1" + using e \C>0\ by (simp add: field_simps) + have "shrink a \ interior(convex hull {a,b,c})" + "shrink b \ interior(convex hull {a,b,c})" + "shrink c \ interior(convex hull {a,b,c})" + using d e by (auto simp: hull_inc mem_interior_convex_shrink shrink_def) + then have fhp0: "(f has_contour_integral 0) + (linepath (shrink a) (shrink b) +++ linepath (shrink b) (shrink c) +++ linepath (shrink c) (shrink a))" + by (simp add: Cauchy_theorem_triangle holomorphic_on_subset [OF holf] hull_minimal convex_interior) + then have f_0_shrink: "?pathint (shrink a) (shrink b) + ?pathint (shrink b) (shrink c) + ?pathint (shrink c) (shrink a) = 0" + by (simp add: has_chain_integral_chain_integral3) + have fpi_abc: "f contour_integrable_on linepath (shrink a) (shrink b)" + "f contour_integrable_on linepath (shrink b) (shrink c)" + "f contour_integrable_on linepath (shrink c) (shrink a)" + using fhp0 by (auto simp: valid_path_join dest: has_contour_integral_integrable) + have cmod_shr: "\x y. cmod (shrink y - shrink x - (y - x)) = e * cmod (x - y)" + using e by (simp add: shrink_def real_vector.scale_right_diff_distrib [symmetric]) + have sh_eq: "\a b d::complex. (b - e *\<^sub>R (b - d)) - (a - e *\<^sub>R (a - d)) - (b - a) = e *\<^sub>R (a - b)" + by (simp add: algebra_simps) + have "cmod y / (24 * C) \ cmod y / cmod (b - a) / 12" + using False \C>0\ diff_2C [of b a] ynz + by (auto simp: divide_simps hull_inc) + have less_C: "\u \ convex hull {a, b, c}; 0 \ x; x \ 1\ \ x * cmod u < C" for x u + apply (cases "x=0", simp add: \0) + using Cno [of u] mult_left_le_one_le [of "cmod u" x] le_less_trans norm_ge_zero by blast + { fix u v + assume uv: "u \ convex hull {a, b, c}" "v \ convex hull {a, b, c}" "u\v" + and fpi_uv: "f contour_integrable_on linepath (shrink u) (shrink v)" + have shr_uv: "shrink u \ interior(convex hull {a,b,c})" + "shrink v \ interior(convex hull {a,b,c})" + using d e uv + by (auto simp: hull_inc mem_interior_convex_shrink shrink_def) + have cmod_fuv: "\x. 0\x \ x\1 \ cmod (f (linepath (shrink u) (shrink v) x)) \ B" + using shr_uv by (blast intro: Bnf linepath_in_convex_hull interior_subset [THEN subsetD]) + have By_uv: "B * (12 * (e * cmod (u - v))) \ cmod y" + apply (rule order_trans [OF _ eCB]) + using e \B>0\ diff_2C [of u v] uv + by (auto simp: field_simps) + { fix x::real assume x: "0\x" "x\1" + have cmod_less_4C: "cmod ((1 - x) *\<^sub>R u - (1 - x) *\<^sub>R d) + cmod (x *\<^sub>R v - x *\<^sub>R d) < (C+C) + (C+C)" + apply (rule add_strict_mono; rule norm_triangle_half_l [of _ 0]) + using uv x d interior_subset + apply (auto simp: hull_inc intro!: less_C) + done + have ll: "linepath (shrink u) (shrink v) x - linepath u v x = -e * ((1 - x) *\<^sub>R (u - d) + x *\<^sub>R (v - d))" + by (simp add: linepath_def shrink_def algebra_simps scaleR_conv_of_real) + have cmod_less_dt: "cmod (linepath (shrink u) (shrink v) x - linepath u v x) < d1" + using \e>0\ + apply (simp add: ll norm_mult scaleR_diff_right) + apply (rule less_le_trans [OF _ e_le_d1]) + using cmod_less_4C + apply (force intro: norm_triangle_lt) + done + have "cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) < cmod y / (24 * C)" + using x uv shr_uv cmod_less_dt + by (auto simp: hull_inc intro: d1 interior_subset [THEN subsetD] linepath_in_convex_hull) + also have "... \ cmod y / cmod (v - u) / 12" + using False uv \C>0\ diff_2C [of v u] ynz + by (auto simp: divide_simps hull_inc) + finally have "cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) \ cmod y / cmod (v - u) / 12" + by simp + then have cmod_12_le: "cmod (v - u) * cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) * 12 \ cmod y" + using uv False by (auto simp: field_simps) + have "cmod (f (linepath (shrink u) (shrink v) x)) * cmod (shrink v - shrink u - (v - u)) + + cmod (v - u) * cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) + \ cmod y / 6" + apply (rule order_trans [of _ "B*((norm y / 24 / C / B)*2*C) + (2*C)*(norm y /24 / C)"]) + apply (rule add_mono [OF mult_mono]) + using By_uv e \0 < B\ \0 < C\ x ynz + apply (simp_all add: cmod_fuv cmod_shr cmod_12_le hull_inc) + apply (simp add: field_simps) + done + } note cmod_diff_le = this + have f_uv: "continuous_on (closed_segment u v) f" + by (blast intro: uv continuous_on_subset [OF contf closed_segment_subset_convex_hull]) + have **: "\f' x' f x::complex. f'*x' - f*x = f'*(x' - x) + x*(f' - f)" + by (simp add: algebra_simps) + have "norm (?pathint (shrink u) (shrink v) - ?pathint u v) \ norm y / 6" + apply (rule order_trans) + apply (rule has_integral_bound + [of "B*(norm y /24/C/B)*2*C + (2*C)*(norm y/24/C)" + "\x. f(linepath (shrink u) (shrink v) x) * (shrink v - shrink u) - f(linepath u v x)*(v - u)" + _ 0 1 ]) + using ynz \0 < B\ \0 < C\ + apply (simp_all del: le_divide_eq_numeral1) + apply (simp add: has_integral_sub has_contour_integral_linepath [symmetric] has_contour_integral_integral + fpi_uv f_uv contour_integrable_continuous_linepath, clarify) + apply (simp only: **) + apply (simp add: norm_triangle_le norm_mult cmod_diff_le del: le_divide_eq_numeral1) + done + } note * = this + have "norm (?pathint (shrink a) (shrink b) - ?pathint a b) \ norm y / 6" + using False fpi_abc by (rule_tac *) (auto simp: hull_inc) + moreover + have "norm (?pathint (shrink b) (shrink c) - ?pathint b c) \ norm y / 6" + using False fpi_abc by (rule_tac *) (auto simp: hull_inc) + moreover + have "norm (?pathint (shrink c) (shrink a) - ?pathint c a) \ norm y / 6" + using False fpi_abc by (rule_tac *) (auto simp: hull_inc) + ultimately + have "norm((?pathint (shrink a) (shrink b) - ?pathint a b) + + (?pathint (shrink b) (shrink c) - ?pathint b c) + (?pathint (shrink c) (shrink a) - ?pathint c a)) + \ norm y / 6 + norm y / 6 + norm y / 6" + by (metis norm_triangle_le add_mono) + also have "... = norm y / 2" + by simp + finally have "norm((?pathint (shrink a) (shrink b) + ?pathint (shrink b) (shrink c) + ?pathint (shrink c) (shrink a)) - + (?pathint a b + ?pathint b c + ?pathint c a)) + \ norm y / 2" + by (simp add: algebra_simps) + then + have "norm(?pathint a b + ?pathint b c + ?pathint c a) \ norm y / 2" + by (simp add: f_0_shrink) (metis (mono_tags) add.commute minus_add_distrib norm_minus_cancel uminus_add_conv_diff) + then have "False" + using pi_eq_y ynz by auto + } + moreover have "uniformly_continuous_on (convex hull {a,b,c}) f" + by (simp add: contf compact_convex_hull compact_uniformly_continuous) + ultimately have "False" + unfolding uniformly_continuous_on_def + by (force simp: ynz \0 < C\ dist_norm) + then show ?thesis .. + qed + } + moreover have "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)" + using fabc contour_integrable_continuous_linepath by auto + ultimately show ?thesis + using has_contour_integral_integral by fastforce +qed + + +subsection\Version allowing finite number of exceptional points\ + +lemma Cauchy_theorem_triangle_cofinite: + assumes "continuous_on (convex hull {a,b,c}) f" + and "finite s" + and "(\x. x \ interior(convex hull {a,b,c}) - s \ f field_differentiable (at x))" + shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" +using assms +proof (induction "card s" arbitrary: a b c s rule: less_induct) + case (less s a b c) + show ?case + proof (cases "s={}") + case True with less show ?thesis + by (fastforce simp: holomorphic_on_def field_differentiable_at_within + Cauchy_theorem_triangle_interior) + next + case False + then obtain d s' where d: "s = insert d s'" "d \ s'" + by (meson Set.set_insert all_not_in_conv) + then show ?thesis + proof (cases "d \ convex hull {a,b,c}") + case False + show "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" + apply (rule less.hyps [of "s'"]) + using False d \finite s\ interior_subset + apply (auto intro!: less.prems) + done + next + case True + have *: "convex hull {a, b, d} \ convex hull {a, b, c}" + by (meson True hull_subset insert_subset convex_hull_subset) + have abd: "(f has_contour_integral 0) (linepath a b +++ linepath b d +++ linepath d a)" + apply (rule less.hyps [of "s'"]) + using True d \finite s\ not_in_interior_convex_hull_3 + apply (auto intro!: less.prems continuous_on_subset [OF _ *]) + apply (metis * insert_absorb insert_subset interior_mono) + done + have *: "convex hull {b, c, d} \ convex hull {a, b, c}" + by (meson True hull_subset insert_subset convex_hull_subset) + have bcd: "(f has_contour_integral 0) (linepath b c +++ linepath c d +++ linepath d b)" + apply (rule less.hyps [of "s'"]) + using True d \finite s\ not_in_interior_convex_hull_3 + apply (auto intro!: less.prems continuous_on_subset [OF _ *]) + apply (metis * insert_absorb insert_subset interior_mono) + done + have *: "convex hull {c, a, d} \ convex hull {a, b, c}" + by (meson True hull_subset insert_subset convex_hull_subset) + have cad: "(f has_contour_integral 0) (linepath c a +++ linepath a d +++ linepath d c)" + apply (rule less.hyps [of "s'"]) + using True d \finite s\ not_in_interior_convex_hull_3 + apply (auto intro!: less.prems continuous_on_subset [OF _ *]) + apply (metis * insert_absorb insert_subset interior_mono) + done + have "f contour_integrable_on linepath a b" + using less.prems + by (metis continuous_on_subset insert_commute contour_integrable_continuous_linepath segments_subset_convex_hull(3)) + moreover have "f contour_integrable_on linepath b c" + using less.prems + by (metis continuous_on_subset contour_integrable_continuous_linepath segments_subset_convex_hull(3)) + moreover have "f contour_integrable_on linepath c a" + using less.prems + by (metis continuous_on_subset insert_commute contour_integrable_continuous_linepath segments_subset_convex_hull(3)) + ultimately have fpi: "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)" + by auto + { fix y::complex + assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)" + and ynz: "y \ 0" + have cont_ad: "continuous_on (closed_segment a d) f" + by (meson "*" continuous_on_subset less.prems(1) segments_subset_convex_hull(3)) + have cont_bd: "continuous_on (closed_segment b d) f" + by (meson True closed_segment_subset_convex_hull continuous_on_subset hull_subset insert_subset less.prems(1)) + have cont_cd: "continuous_on (closed_segment c d) f" + by (meson "*" continuous_on_subset less.prems(1) segments_subset_convex_hull(2)) + have "contour_integral (linepath a b) f = - (contour_integral (linepath b d) f + (contour_integral (linepath d a) f))" + "contour_integral (linepath b c) f = - (contour_integral (linepath c d) f + (contour_integral (linepath d b) f))" + "contour_integral (linepath c a) f = - (contour_integral (linepath a d) f + contour_integral (linepath d c) f)" + using has_chain_integral_chain_integral3 [OF abd] + has_chain_integral_chain_integral3 [OF bcd] + has_chain_integral_chain_integral3 [OF cad] + by (simp_all add: algebra_simps add_eq_0_iff) + then have ?thesis + using cont_ad cont_bd cont_cd fy has_chain_integral_chain_integral3 contour_integral_reverse_linepath by fastforce + } + then show ?thesis + using fpi contour_integrable_on_def by blast + qed + qed +qed + + +subsection\Cauchy's theorem for an open starlike set\ + +lemma starlike_convex_subset: + assumes s: "a \ s" "closed_segment b c \ s" and subs: "\x. x \ s \ closed_segment a x \ s" + shows "convex hull {a,b,c} \ s" + using s + apply (clarsimp simp add: convex_hull_insert [of "{b,c}" a] segment_convex_hull) + apply (meson subs convexD convex_closed_segment ends_in_segment(1) ends_in_segment(2) subsetCE) + done + +lemma triangle_contour_integrals_starlike_primitive: + assumes contf: "continuous_on s f" + and s: "a \ s" "open s" + and x: "x \ s" + and subs: "\y. y \ s \ closed_segment a y \ s" + and zer: "\b c. closed_segment b c \ s + \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0" + shows "((\x. contour_integral(linepath a x) f) has_field_derivative f x) (at x)" +proof - + let ?pathint = "\x y. contour_integral(linepath x y) f" + { fix e y + assume e: "0 < e" and bxe: "ball x e \ s" and close: "cmod (y - x) < e" + have y: "y \ s" + using bxe close by (force simp: dist_norm norm_minus_commute) + have cont_ayf: "continuous_on (closed_segment a y) f" + using contf continuous_on_subset subs y by blast + have xys: "closed_segment x y \ s" + apply (rule order_trans [OF _ bxe]) + using close + by (auto simp: dist_norm ball_def norm_minus_commute dest: segment_bound) + have "?pathint a y - ?pathint a x = ?pathint x y" + using zer [OF xys] contour_integral_reverse_linepath [OF cont_ayf] add_eq_0_iff by force + } note [simp] = this + { fix e::real + assume e: "0 < e" + have cont_atx: "continuous (at x) f" + using x s contf continuous_on_eq_continuous_at by blast + then obtain d1 where d1: "d1>0" and d1_less: "\y. cmod (y - x) < d1 \ cmod (f y - f x) < e/2" + unfolding continuous_at Lim_at dist_norm using e + by (drule_tac x="e/2" in spec) force + obtain d2 where d2: "d2>0" "ball x d2 \ s" using \open s\ x + by (auto simp: open_contains_ball) + have dpos: "min d1 d2 > 0" using d1 d2 by simp + { fix y + assume yx: "y \ x" and close: "cmod (y - x) < min d1 d2" + have y: "y \ s" + using d2 close by (force simp: dist_norm norm_minus_commute) + have fxy: "f contour_integrable_on linepath x y" + apply (rule contour_integrable_continuous_linepath) + apply (rule continuous_on_subset [OF contf]) + using close d2 + apply (auto simp: dist_norm norm_minus_commute dest!: segment_bound(1)) + done + then obtain i where i: "(f has_contour_integral i) (linepath x y)" + by (auto simp: contour_integrable_on_def) + then have "((\w. f w - f x) has_contour_integral (i - f x * (y - x))) (linepath x y)" + by (rule has_contour_integral_diff [OF _ has_contour_integral_const_linepath]) + then have "cmod (i - f x * (y - x)) \ e / 2 * cmod (y - x)" + apply (rule has_contour_integral_bound_linepath [where B = "e/2"]) + using e apply simp + apply (rule d1_less [THEN less_imp_le]) + using close segment_bound + apply force + done + also have "... < e * cmod (y - x)" + by (simp add: e yx) + finally have "cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e" + using i yx by (simp add: contour_integral_unique divide_less_eq) + } + then have "\d>0. \y. y \ x \ cmod (y-x) < d \ cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e" + using dpos by blast + } + then have *: "(\y. (?pathint x y - f x * (y - x)) /\<^sub>R cmod (y - x)) \x\ 0" + by (simp add: Lim_at dist_norm inverse_eq_divide) + show ?thesis + apply (simp add: has_field_derivative_def has_derivative_at bounded_linear_mult_right) + apply (rule Lim_transform [OF * Lim_eventually]) + apply (simp add: inverse_eq_divide [symmetric] eventually_at) + using \open s\ x + apply (force simp: dist_norm open_contains_ball) + done +qed + +(** Existence of a primitive.*) + +lemma holomorphic_starlike_primitive: + fixes f :: "complex \ complex" + assumes contf: "continuous_on s f" + and s: "starlike s" and os: "open s" + and k: "finite k" + and fcd: "\x. x \ s - k \ f field_differentiable at x" + shows "\g. \x \ s. (g has_field_derivative f x) (at x)" +proof - + obtain a where a: "a\s" and a_cs: "\x. x\s \ closed_segment a x \ s" + using s by (auto simp: starlike_def) + { fix x b c + assume "x \ s" "closed_segment b c \ s" + then have abcs: "convex hull {a, b, c} \ s" + by (simp add: a a_cs starlike_convex_subset) + then have *: "continuous_on (convex hull {a, b, c}) f" + by (simp add: continuous_on_subset [OF contf]) + have "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" + apply (rule Cauchy_theorem_triangle_cofinite [OF _ k]) + using abcs apply (simp add: continuous_on_subset [OF contf]) + using * abcs interior_subset apply (auto intro: fcd) + done + } note 0 = this + show ?thesis + apply (intro exI ballI) + apply (rule triangle_contour_integrals_starlike_primitive [OF contf a os], assumption) + apply (metis a_cs) + apply (metis has_chain_integral_chain_integral3 0) + done +qed + +lemma Cauchy_theorem_starlike: + "\open s; starlike s; finite k; continuous_on s f; + \x. x \ s - k \ f field_differentiable at x; + valid_path g; path_image g \ s; pathfinish g = pathstart g\ + \ (f has_contour_integral 0) g" + by (metis holomorphic_starlike_primitive Cauchy_theorem_primitive at_within_open) + +lemma Cauchy_theorem_starlike_simple: + "\open s; starlike s; f holomorphic_on s; valid_path g; path_image g \ s; pathfinish g = pathstart g\ + \ (f has_contour_integral 0) g" +apply (rule Cauchy_theorem_starlike [OF _ _ finite.emptyI]) +apply (simp_all add: holomorphic_on_imp_continuous_on) +apply (metis at_within_open holomorphic_on_def) +done + + +subsection\Cauchy's theorem for a convex set\ + +text\For a convex set we can avoid assuming openness and boundary analyticity\ + +lemma triangle_contour_integrals_convex_primitive: + assumes contf: "continuous_on s f" + and s: "a \ s" "convex s" + and x: "x \ s" + and zer: "\b c. \b \ s; c \ s\ + \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0" + shows "((\x. contour_integral(linepath a x) f) has_field_derivative f x) (at x within s)" +proof - + let ?pathint = "\x y. contour_integral(linepath x y) f" + { fix y + assume y: "y \ s" + have cont_ayf: "continuous_on (closed_segment a y) f" + using s y by (meson contf continuous_on_subset convex_contains_segment) + have xys: "closed_segment x y \ s" (*?*) + using convex_contains_segment s x y by auto + have "?pathint a y - ?pathint a x = ?pathint x y" + using zer [OF x y] contour_integral_reverse_linepath [OF cont_ayf] add_eq_0_iff by force + } note [simp] = this + { fix e::real + assume e: "0 < e" + have cont_atx: "continuous (at x within s) f" + using x s contf by (simp add: continuous_on_eq_continuous_within) + then obtain d1 where d1: "d1>0" and d1_less: "\y. \y \ s; cmod (y - x) < d1\ \ cmod (f y - f x) < e/2" + unfolding continuous_within Lim_within dist_norm using e + by (drule_tac x="e/2" in spec) force + { fix y + assume yx: "y \ x" and close: "cmod (y - x) < d1" and y: "y \ s" + have fxy: "f contour_integrable_on linepath x y" + using convex_contains_segment s x y + by (blast intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf]) + then obtain i where i: "(f has_contour_integral i) (linepath x y)" + by (auto simp: contour_integrable_on_def) + then have "((\w. f w - f x) has_contour_integral (i - f x * (y - x))) (linepath x y)" + by (rule has_contour_integral_diff [OF _ has_contour_integral_const_linepath]) + then have "cmod (i - f x * (y - x)) \ e / 2 * cmod (y - x)" + apply (rule has_contour_integral_bound_linepath [where B = "e/2"]) + using e apply simp + apply (rule d1_less [THEN less_imp_le]) + using convex_contains_segment s(2) x y apply blast + using close segment_bound(1) apply fastforce + done + also have "... < e * cmod (y - x)" + by (simp add: e yx) + finally have "cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e" + using i yx by (simp add: contour_integral_unique divide_less_eq) + } + then have "\d>0. \y\s. y \ x \ cmod (y-x) < d \ cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e" + using d1 by blast + } + then have *: "((\y. (contour_integral (linepath x y) f - f x * (y - x)) /\<^sub>R cmod (y - x)) \ 0) (at x within s)" + by (simp add: Lim_within dist_norm inverse_eq_divide) + show ?thesis + apply (simp add: has_field_derivative_def has_derivative_within bounded_linear_mult_right) + apply (rule Lim_transform [OF * Lim_eventually]) + using linordered_field_no_ub + apply (force simp: inverse_eq_divide [symmetric] eventually_at) + done +qed + +lemma contour_integral_convex_primitive: + "\convex s; continuous_on s f; + \a b c. \a \ s; b \ s; c \ s\ \ (f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)\ + \ \g. \x \ s. (g has_field_derivative f x) (at x within s)" + apply (cases "s={}") + apply (simp_all add: ex_in_conv [symmetric]) + apply (blast intro: triangle_contour_integrals_convex_primitive has_chain_integral_chain_integral3) + done + +lemma holomorphic_convex_primitive: + fixes f :: "complex \ complex" + shows + "\convex s; finite k; continuous_on s f; + \x. x \ interior s - k \ f field_differentiable at x\ + \ \g. \x \ s. (g has_field_derivative f x) (at x within s)" +apply (rule contour_integral_convex_primitive [OF _ _ Cauchy_theorem_triangle_cofinite]) +prefer 3 +apply (erule continuous_on_subset) +apply (simp add: subset_hull continuous_on_subset, assumption+) +by (metis Diff_iff convex_contains_segment insert_absorb insert_subset interior_mono segment_convex_hull subset_hull) + +lemma Cauchy_theorem_convex: + "\continuous_on s f; convex s; finite k; + \x. x \ interior s - k \ f field_differentiable at x; + valid_path g; path_image g \ s; + pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" + by (metis holomorphic_convex_primitive Cauchy_theorem_primitive) + +lemma Cauchy_theorem_convex_simple: + "\f holomorphic_on s; convex s; + valid_path g; path_image g \ s; + pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" + apply (rule Cauchy_theorem_convex) + apply (simp_all add: holomorphic_on_imp_continuous_on) + apply (rule finite.emptyI) + using at_within_interior holomorphic_on_def interior_subset by fastforce + + +text\In particular for a disc\ +lemma Cauchy_theorem_disc: + "\finite k; continuous_on (cball a e) f; + \x. x \ ball a e - k \ f field_differentiable at x; + valid_path g; path_image g \ cball a e; + pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" + apply (rule Cauchy_theorem_convex) + apply (auto simp: convex_cball interior_cball) + done + +lemma Cauchy_theorem_disc_simple: + "\f holomorphic_on (ball a e); valid_path g; path_image g \ ball a e; + pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" +by (simp add: Cauchy_theorem_convex_simple) + + +subsection\Generalize integrability to local primitives\ + +lemma contour_integral_local_primitive_lemma: + fixes f :: "complex\complex" + shows + "\g piecewise_differentiable_on {a..b}; + \x. x \ s \ (f has_field_derivative f' x) (at x within s); + \x. x \ {a..b} \ g x \ s\ + \ (\x. f' (g x) * vector_derivative g (at x within {a..b})) + integrable_on {a..b}" + apply (cases "cbox a b = {}", force) + apply (simp add: integrable_on_def) + apply (rule exI) + apply (rule contour_integral_primitive_lemma, assumption+) + using atLeastAtMost_iff by blast + +lemma contour_integral_local_primitive_any: + fixes f :: "complex \ complex" + assumes gpd: "g piecewise_differentiable_on {a..b}" + and dh: "\x. x \ s + \ \d h. 0 < d \ + (\y. norm(y - x) < d \ (h has_field_derivative f y) (at y within s))" + and gs: "\x. x \ {a..b} \ g x \ s" + shows "(\x. f(g x) * vector_derivative g (at x)) integrable_on {a..b}" +proof - + { fix x + assume x: "a \ x" "x \ b" + obtain d h where d: "0 < d" + and h: "(\y. norm(y - g x) < d \ (h has_field_derivative f y) (at y within s))" + using x gs dh by (metis atLeastAtMost_iff) + have "continuous_on {a..b} g" using gpd piecewise_differentiable_on_def by blast + then obtain e where e: "e>0" and lessd: "\x'. x' \ {a..b} \ \x' - x\ < e \ cmod (g x' - g x) < d" + using x d + apply (auto simp: dist_norm continuous_on_iff) + apply (drule_tac x=x in bspec) + using x apply simp + apply (drule_tac x=d in spec, auto) + done + have "\d>0. \u v. u \ x \ x \ v \ {u..v} \ ball x d \ (u \ v \ a \ u \ v \ b) \ + (\x. f (g x) * vector_derivative g (at x)) integrable_on {u..v}" + apply (rule_tac x=e in exI) + using e + apply (simp add: integrable_on_localized_vector_derivative [symmetric], clarify) + apply (rule_tac f = h and s = "g ` {u..v}" in contour_integral_local_primitive_lemma) + apply (meson atLeastatMost_subset_iff gpd piecewise_differentiable_on_subset) + apply (force simp: ball_def dist_norm intro: lessd gs DERIV_subset [OF h], force) + done + } then + show ?thesis + by (force simp: intro!: integrable_on_little_subintervals [of a b, simplified]) +qed + +lemma contour_integral_local_primitive: + fixes f :: "complex \ complex" + assumes g: "valid_path g" "path_image g \ s" + and dh: "\x. x \ s + \ \d h. 0 < d \ + (\y. norm(y - x) < d \ (h has_field_derivative f y) (at y within s))" + shows "f contour_integrable_on g" + using g + apply (simp add: valid_path_def path_image_def contour_integrable_on_def has_contour_integral_def + has_integral_localized_vector_derivative integrable_on_def [symmetric]) + using contour_integral_local_primitive_any [OF _ dh] + by (meson image_subset_iff piecewise_C1_imp_differentiable) + + +text\In particular if a function is holomorphic\ + +lemma contour_integrable_holomorphic: + assumes contf: "continuous_on s f" + and os: "open s" + and k: "finite k" + and g: "valid_path g" "path_image g \ s" + and fcd: "\x. x \ s - k \ f field_differentiable at x" + shows "f contour_integrable_on g" +proof - + { fix z + assume z: "z \ s" + obtain d where d: "d>0" "ball z d \ s" using \open s\ z + by (auto simp: open_contains_ball) + then have contfb: "continuous_on (ball z d) f" + using contf continuous_on_subset by blast + obtain h where "\y\ball z d. (h has_field_derivative f y) (at y within ball z d)" + using holomorphic_convex_primitive [OF convex_ball k contfb fcd] d + interior_subset by force + then have "\y\ball z d. (h has_field_derivative f y) (at y within s)" + by (metis Topology_Euclidean_Space.open_ball at_within_open d(2) os subsetCE) + then have "\h. (\y. cmod (y - z) < d \ (h has_field_derivative f y) (at y within s))" + by (force simp: dist_norm norm_minus_commute) + then have "\d h. 0 < d \ (\y. cmod (y - z) < d \ (h has_field_derivative f y) (at y within s))" + using d by blast + } + then show ?thesis + by (rule contour_integral_local_primitive [OF g]) +qed + +lemma contour_integrable_holomorphic_simple: + assumes fh: "f holomorphic_on s" + and os: "open s" + and g: "valid_path g" "path_image g \ s" + shows "f contour_integrable_on g" + apply (rule contour_integrable_holomorphic [OF _ os Finite_Set.finite.emptyI g]) + apply (simp add: fh holomorphic_on_imp_continuous_on) + using fh by (simp add: field_differentiable_def holomorphic_on_open os) + +lemma continuous_on_inversediff: + fixes z:: "'a::real_normed_field" shows "z \ s \ continuous_on s (\w. 1 / (w - z))" + by (rule continuous_intros | force)+ + +corollary contour_integrable_inversediff: + "\valid_path g; z \ path_image g\ \ (\w. 1 / (w-z)) contour_integrable_on g" +apply (rule contour_integrable_holomorphic_simple [of _ "UNIV-{z}"]) +apply (auto simp: holomorphic_on_open open_delete intro!: derivative_eq_intros) +done + +text\Key fact that path integral is the same for a "nearby" path. This is the + main lemma for the homotopy form of Cauchy's theorem and is also useful + if we want "without loss of generality" to assume some nice properties of a + path (e.g. smoothness). It can also be used to define the integrals of + analytic functions over arbitrary continuous paths. This is just done for + winding numbers now. +\ + +text\A technical definition to avoid duplication of similar proofs, + for paths joined at the ends versus looping paths\ +definition linked_paths :: "bool \ (real \ 'a) \ (real \ 'a::topological_space) \ bool" + where "linked_paths atends g h == + (if atends then pathstart h = pathstart g \ pathfinish h = pathfinish g + else pathfinish g = pathstart g \ pathfinish h = pathstart h)" + +text\This formulation covers two cases: @{term g} and @{term h} share their + start and end points; @{term g} and @{term h} both loop upon themselves.\ +lemma contour_integral_nearby: + assumes os: "open s" and p: "path p" "path_image p \ s" + shows + "\d. 0 < d \ + (\g h. valid_path g \ valid_path h \ + (\t \ {0..1}. norm(g t - p t) < d \ norm(h t - p t) < d) \ + linked_paths atends g h + \ path_image g \ s \ path_image h \ s \ + (\f. f holomorphic_on s \ contour_integral h f = contour_integral g f))" +proof - + have "\z. \e. z \ path_image p \ 0 < e \ ball z e \ s" + using open_contains_ball os p(2) by blast + then obtain ee where ee: "\z. z \ path_image p \ 0 < ee z \ ball z (ee z) \ s" + by metis + define cover where "cover = (\z. ball z (ee z/3)) ` (path_image p)" + have "compact (path_image p)" + by (metis p(1) compact_path_image) + moreover have "path_image p \ (\c\path_image p. ball c (ee c / 3))" + using ee by auto + ultimately have "\D \ cover. finite D \ path_image p \ \D" + by (simp add: compact_eq_heine_borel cover_def) + then obtain D where D: "D \ cover" "finite D" "path_image p \ \D" + by blast + then obtain k where k: "k \ {0..1}" "finite k" and D_eq: "D = ((\z. ball z (ee z / 3)) \ p) ` k" + apply (simp add: cover_def path_image_def image_comp) + apply (blast dest!: finite_subset_image [OF \finite D\]) + done + then have kne: "k \ {}" + using D by auto + have pi: "\i. i \ k \ p i \ path_image p" + using k by (auto simp: path_image_def) + then have eepi: "\i. i \ k \ 0 < ee((p i))" + by (metis ee) + define e where "e = Min((ee o p) ` k)" + have fin_eep: "finite ((ee o p) ` k)" + using k by blast + have enz: "0 < e" + using ee k by (simp add: kne e_def Min_gr_iff [OF fin_eep] eepi) + have "uniformly_continuous_on {0..1} p" + using p by (simp add: path_def compact_uniformly_continuous) + then obtain d::real where d: "d>0" + and de: "\x x'. \x' - x\ < d \ x\{0..1} \ x'\{0..1} \ cmod (p x' - p x) < e/3" + unfolding uniformly_continuous_on_def dist_norm real_norm_def + by (metis divide_pos_pos enz zero_less_numeral) + then obtain N::nat where N: "N>0" "inverse N < d" + using real_arch_inverse [of d] by auto + { fix g h + assume g: "valid_path g" and gp: "\t\{0..1}. cmod (g t - p t) < e / 3" + and h: "valid_path h" and hp: "\t\{0..1}. cmod (h t - p t) < e / 3" + and joins: "linked_paths atends g h" + { fix t::real + assume t: "0 \ t" "t \ 1" + then obtain u where u: "u \ k" and ptu: "p t \ ball(p u) (ee(p u) / 3)" + using \path_image p \ \D\ D_eq by (force simp: path_image_def) + then have ele: "e \ ee (p u)" using fin_eep + by (simp add: e_def) + have "cmod (g t - p t) < e / 3" "cmod (h t - p t) < e / 3" + using gp hp t by auto + with ele have "cmod (g t - p t) < ee (p u) / 3" + "cmod (h t - p t) < ee (p u) / 3" + by linarith+ + then have "g t \ ball(p u) (ee(p u))" "h t \ ball(p u) (ee(p u))" + using norm_diff_triangle_ineq [of "g t" "p t" "p t" "p u"] + norm_diff_triangle_ineq [of "h t" "p t" "p t" "p u"] ptu eepi u + by (force simp: dist_norm ball_def norm_minus_commute)+ + then have "g t \ s" "h t \ s" using ee u k + by (auto simp: path_image_def ball_def) + } + then have ghs: "path_image g \ s" "path_image h \ s" + by (auto simp: path_image_def) + moreover + { fix f + assume fhols: "f holomorphic_on s" + then have fpa: "f contour_integrable_on g" "f contour_integrable_on h" + using g ghs h holomorphic_on_imp_continuous_on os contour_integrable_holomorphic_simple + by blast+ + have contf: "continuous_on s f" + by (simp add: fhols holomorphic_on_imp_continuous_on) + { fix z + assume z: "z \ path_image p" + have "f holomorphic_on ball z (ee z)" + using fhols ee z holomorphic_on_subset by blast + then have "\ff. (\w \ ball z (ee z). (ff has_field_derivative f w) (at w))" + using holomorphic_convex_primitive [of "ball z (ee z)" "{}" f, simplified] + by (metis open_ball at_within_open holomorphic_on_def holomorphic_on_imp_continuous_on mem_ball) + } + then obtain ff where ff: + "\z w. \z \ path_image p; w \ ball z (ee z)\ \ (ff z has_field_derivative f w) (at w)" + by metis + { fix n + assume n: "n \ N" + then have "contour_integral(subpath 0 (n/N) h) f - contour_integral(subpath 0 (n/N) g) f = + contour_integral(linepath (g(n/N)) (h(n/N))) f - contour_integral(linepath (g 0) (h 0)) f" + proof (induct n) + case 0 show ?case by simp + next + case (Suc n) + obtain t where t: "t \ k" and "p (n/N) \ ball(p t) (ee(p t) / 3)" + using \path_image p \ \D\ [THEN subsetD, where c="p (n/N)"] D_eq N Suc.prems + by (force simp: path_image_def) + then have ptu: "cmod (p t - p (n/N)) < ee (p t) / 3" + by (simp add: dist_norm) + have e3le: "e/3 \ ee (p t) / 3" using fin_eep t + by (simp add: e_def) + { fix x + assume x: "n/N \ x" "x \ (1 + n)/N" + then have nN01: "0 \ n/N" "(1 + n)/N \ 1" + using Suc.prems by auto + then have x01: "0 \ x" "x \ 1" + using x by linarith+ + have "cmod (p t - p x) < ee (p t) / 3 + e/3" + apply (rule norm_diff_triangle_less [OF ptu de]) + using x N x01 Suc.prems + apply (auto simp: field_simps) + done + then have ptx: "cmod (p t - p x) < 2*ee (p t)/3" + using e3le eepi [OF t] by simp + have "cmod (p t - g x) < 2*ee (p t)/3 + e/3 " + apply (rule norm_diff_triangle_less [OF ptx]) + using gp x01 by (simp add: norm_minus_commute) + also have "... \ ee (p t)" + using e3le eepi [OF t] by simp + finally have gg: "cmod (p t - g x) < ee (p t)" . + have "cmod (p t - h x) < 2*ee (p t)/3 + e/3 " + apply (rule norm_diff_triangle_less [OF ptx]) + using hp x01 by (simp add: norm_minus_commute) + also have "... \ ee (p t)" + using e3le eepi [OF t] by simp + finally have "cmod (p t - g x) < ee (p t)" + "cmod (p t - h x) < ee (p t)" + using gg by auto + } note ptgh_ee = this + have pi_hgn: "path_image (linepath (h (n/N)) (g (n/N))) \ ball (p t) (ee (p t))" + using ptgh_ee [of "n/N"] Suc.prems + by (auto simp: field_simps dist_norm dest: segment_furthest_le [where y="p t"]) + then have gh_ns: "closed_segment (g (n/N)) (h (n/N)) \ s" + using \N>0\ Suc.prems + apply (simp add: path_image_join field_simps closed_segment_commute) + apply (erule order_trans) + apply (simp add: ee pi t) + done + have pi_ghn': "path_image (linepath (g ((1 + n) / N)) (h ((1 + n) / N))) + \ ball (p t) (ee (p t))" + using ptgh_ee [of "(1+n)/N"] Suc.prems + by (auto simp: field_simps dist_norm dest: segment_furthest_le [where y="p t"]) + then have gh_n's: "closed_segment (g ((1 + n) / N)) (h ((1 + n) / N)) \ s" + using \N>0\ Suc.prems ee pi t + by (auto simp: Path_Connected.path_image_join field_simps) + have pi_subset_ball: + "path_image (subpath (n/N) ((1+n) / N) g +++ linepath (g ((1+n) / N)) (h ((1+n) / N)) +++ + subpath ((1+n) / N) (n/N) h +++ linepath (h (n/N)) (g (n/N))) + \ ball (p t) (ee (p t))" + apply (intro subset_path_image_join pi_hgn pi_ghn') + using \N>0\ Suc.prems + apply (auto simp: path_image_subpath dist_norm field_simps closed_segment_eq_real_ivl ptgh_ee) + done + have pi0: "(f has_contour_integral 0) + (subpath (n/ N) ((Suc n)/N) g +++ linepath(g ((Suc n) / N)) (h((Suc n) / N)) +++ + subpath ((Suc n) / N) (n/N) h +++ linepath(h (n/N)) (g (n/N)))" + apply (rule Cauchy_theorem_primitive [of "ball(p t) (ee(p t))" "ff (p t)" "f"]) + apply (metis ff open_ball at_within_open pi t) + apply (intro valid_path_join) + using Suc.prems pi_subset_ball apply (simp_all add: valid_path_subpath g h) + done + have fpa1: "f contour_integrable_on subpath (real n / real N) (real (Suc n) / real N) g" + using Suc.prems by (simp add: contour_integrable_subpath g fpa) + have fpa2: "f contour_integrable_on linepath (g (real (Suc n) / real N)) (h (real (Suc n) / real N))" + using gh_n's + by (auto intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf]) + have fpa3: "f contour_integrable_on linepath (h (real n / real N)) (g (real n / real N))" + using gh_ns + by (auto simp: closed_segment_commute intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf]) + have eq0: "contour_integral (subpath (n/N) ((Suc n) / real N) g) f + + contour_integral (linepath (g ((Suc n) / N)) (h ((Suc n) / N))) f + + contour_integral (subpath ((Suc n) / N) (n/N) h) f + + contour_integral (linepath (h (n/N)) (g (n/N))) f = 0" + using contour_integral_unique [OF pi0] Suc.prems + by (simp add: g h fpa valid_path_subpath contour_integrable_subpath + fpa1 fpa2 fpa3 algebra_simps del: of_nat_Suc) + have *: "\hn he hn' gn gd gn' hgn ghn gh0 ghn'. + \hn - gn = ghn - gh0; + gd + ghn' + he + hgn = (0::complex); + hn - he = hn'; gn + gd = gn'; hgn = -ghn\ \ hn' - gn' = ghn' - gh0" + by (auto simp: algebra_simps) + have "contour_integral (subpath 0 (n/N) h) f - contour_integral (subpath ((Suc n) / N) (n/N) h) f = + contour_integral (subpath 0 (n/N) h) f + contour_integral (subpath (n/N) ((Suc n) / N) h) f" + unfolding reversepath_subpath [symmetric, of "((Suc n) / N)"] + using Suc.prems by (simp add: h fpa contour_integral_reversepath valid_path_subpath contour_integrable_subpath) + also have "... = contour_integral (subpath 0 ((Suc n) / N) h) f" + using Suc.prems by (simp add: contour_integral_subpath_combine h fpa) + finally have pi0_eq: + "contour_integral (subpath 0 (n/N) h) f - contour_integral (subpath ((Suc n) / N) (n/N) h) f = + contour_integral (subpath 0 ((Suc n) / N) h) f" . + show ?case + apply (rule * [OF Suc.hyps eq0 pi0_eq]) + using Suc.prems + apply (simp_all add: g h fpa contour_integral_subpath_combine + contour_integral_reversepath [symmetric] contour_integrable_continuous_linepath + continuous_on_subset [OF contf gh_ns]) + done + qed + } note ind = this + have "contour_integral h f = contour_integral g f" + using ind [OF order_refl] N joins + by (simp add: linked_paths_def pathstart_def pathfinish_def split: if_split_asm) + } + ultimately + have "path_image g \ s \ path_image h \ s \ (\f. f holomorphic_on s \ contour_integral h f = contour_integral g f)" + by metis + } note * = this + show ?thesis + apply (rule_tac x="e/3" in exI) + apply (rule conjI) + using enz apply simp + apply (clarsimp simp only: ball_conj_distrib) + apply (rule *; assumption) + done +qed + + +lemma + assumes "open s" "path p" "path_image p \ s" + shows contour_integral_nearby_ends: + "\d. 0 < d \ + (\g h. valid_path g \ valid_path h \ + (\t \ {0..1}. norm(g t - p t) < d \ norm(h t - p t) < d) \ + pathstart h = pathstart g \ pathfinish h = pathfinish g + \ path_image g \ s \ + path_image h \ s \ + (\f. f holomorphic_on s + \ contour_integral h f = contour_integral g f))" + and contour_integral_nearby_loops: + "\d. 0 < d \ + (\g h. valid_path g \ valid_path h \ + (\t \ {0..1}. norm(g t - p t) < d \ norm(h t - p t) < d) \ + pathfinish g = pathstart g \ pathfinish h = pathstart h + \ path_image g \ s \ + path_image h \ s \ + (\f. f holomorphic_on s + \ contour_integral h f = contour_integral g f))" + using contour_integral_nearby [OF assms, where atends=True] + using contour_integral_nearby [OF assms, where atends=False] + unfolding linked_paths_def by simp_all + +corollary differentiable_polynomial_function: + fixes p :: "real \ 'a::euclidean_space" + shows "polynomial_function p \ p differentiable_on s" +by (meson has_vector_derivative_polynomial_function differentiable_at_imp_differentiable_on differentiable_def has_vector_derivative_def) + +lemma C1_differentiable_polynomial_function: + fixes p :: "real \ 'a::euclidean_space" + shows "polynomial_function p \ p C1_differentiable_on s" + by (metis continuous_on_polymonial_function C1_differentiable_on_def has_vector_derivative_polynomial_function) + +lemma valid_path_polynomial_function: + fixes p :: "real \ 'a::euclidean_space" + shows "polynomial_function p \ valid_path p" +by (force simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_polymonial_function C1_differentiable_polynomial_function) + +lemma valid_path_subpath_trivial [simp]: + fixes g :: "real \ 'a::euclidean_space" + shows "z \ g x \ valid_path (subpath x x g)" + by (simp add: subpath_def valid_path_polynomial_function) + +lemma contour_integral_bound_exists: +assumes s: "open s" + and g: "valid_path g" + and pag: "path_image g \ s" + shows "\L. 0 < L \ + (\f B. f holomorphic_on s \ (\z \ s. norm(f z) \ B) + \ norm(contour_integral g f) \ L*B)" +proof - +have "path g" using g + by (simp add: valid_path_imp_path) +then obtain d::real and p + where d: "0 < d" + and p: "polynomial_function p" "path_image p \ s" + and pi: "\f. f holomorphic_on s \ contour_integral g f = contour_integral p f" + using contour_integral_nearby_ends [OF s \path g\ pag] + apply clarify + apply (drule_tac x=g in spec) + apply (simp only: assms) + apply (force simp: valid_path_polynomial_function dest: path_approx_polynomial_function) + done +then obtain p' where p': "polynomial_function p'" + "\x. (p has_vector_derivative (p' x)) (at x)" + using has_vector_derivative_polynomial_function by force +then have "bounded(p' ` {0..1})" + using continuous_on_polymonial_function + by (force simp: intro!: compact_imp_bounded compact_continuous_image) +then obtain L where L: "L>0" and nop': "\x. x \ {0..1} \ norm (p' x) \ L" + by (force simp: bounded_pos) +{ fix f B + assume f: "f holomorphic_on s" + and B: "\z. z\s \ cmod (f z) \ B" + then have "f contour_integrable_on p \ valid_path p" + using p s + by (blast intro: valid_path_polynomial_function contour_integrable_holomorphic_simple holomorphic_on_imp_continuous_on) + moreover have "\x. x \ {0..1} \ cmod (vector_derivative p (at x)) * cmod (f (p x)) \ L * B" + apply (rule mult_mono) + apply (subst Derivative.vector_derivative_at; force intro: p' nop') + using L B p + apply (auto simp: path_image_def image_subset_iff) + done + ultimately have "cmod (contour_integral g f) \ L * B" + apply (simp add: pi [OF f]) + apply (simp add: contour_integral_integral) + apply (rule order_trans [OF integral_norm_bound_integral]) + apply (auto simp: mult.commute integral_norm_bound_integral contour_integrable_on [symmetric] norm_mult) + done +} then +show ?thesis + by (force simp: L contour_integral_integral) +qed + +subsection\Constancy of a function from a connected set into a finite, disconnected or discrete set\ + +text\Still missing: versions for a set that is smaller than R, or countable.\ + +lemma continuous_disconnected_range_constant: + assumes s: "connected s" + and conf: "continuous_on s f" + and fim: "f ` s \ t" + and cct: "\y. y \ t \ connected_component_set t y = {y}" + shows "\a. \x \ s. f x = a" +proof (cases "s = {}") + case True then show ?thesis by force +next + case False + { fix x assume "x \ s" + then have "f ` s \ {f x}" + by (metis connected_continuous_image conf connected_component_maximal fim image_subset_iff rev_image_eqI s cct) + } + with False show ?thesis + by blast +qed + +lemma discrete_subset_disconnected: + fixes s :: "'a::topological_space set" + fixes t :: "'b::real_normed_vector set" + assumes conf: "continuous_on s f" + and no: "\x. x \ s \ \e>0. \y. y \ s \ f y \ f x \ e \ norm (f y - f x)" + shows "f ` s \ {y. connected_component_set (f ` s) y = {y}}" +proof - + { fix x assume x: "x \ s" + then obtain e where "e>0" and ele: "\y. \y \ s; f y \ f x\ \ e \ norm (f y - f x)" + using conf no [OF x] by auto + then have e2: "0 \ e / 2" + by simp + have "f y = f x" if "y \ s" and ccs: "f y \ connected_component_set (f ` s) (f x)" for y + apply (rule ccontr) + using connected_closed [of "connected_component_set (f ` s) (f x)"] \e>0\ + apply (simp add: del: ex_simps) + apply (drule spec [where x="cball (f x) (e / 2)"]) + apply (drule spec [where x="- ball(f x) e"]) + apply (auto simp: dist_norm open_closed [symmetric] simp del: le_divide_eq_numeral1 dest!: connected_component_in) + apply (metis diff_self e2 ele norm_minus_commute norm_zero not_less) + using centre_in_cball connected_component_refl_eq e2 x apply blast + using ccs + apply (force simp: cball_def dist_norm norm_minus_commute dest: ele [OF \y \ s\]) + done + moreover have "connected_component_set (f ` s) (f x) \ f ` s" + by (auto simp: connected_component_in) + ultimately have "connected_component_set (f ` s) (f x) = {f x}" + by (auto simp: x) + } + with assms show ?thesis + by blast +qed + +lemma finite_implies_discrete: + fixes s :: "'a::topological_space set" + assumes "finite (f ` s)" + shows "(\x \ s. \e>0. \y. y \ s \ f y \ f x \ e \ norm (f y - f x))" +proof - + have "\e>0. \y. y \ s \ f y \ f x \ e \ norm (f y - f x)" if "x \ s" for x + proof (cases "f ` s - {f x} = {}") + case True + with zero_less_numeral show ?thesis + by (fastforce simp add: Set.image_subset_iff cong: conj_cong) + next + case False + then obtain z where z: "z \ s" "f z \ f x" + by blast + have finn: "finite {norm (z - f x) |z. z \ f ` s - {f x}}" + using assms by simp + then have *: "0 < Inf{norm(z - f x) | z. z \ f ` s - {f x}}" + apply (rule finite_imp_less_Inf) + using z apply force+ + done + show ?thesis + by (force intro!: * cInf_le_finite [OF finn]) + qed + with assms show ?thesis + by blast +qed + +text\This proof requires the existence of two separate values of the range type.\ +lemma finite_range_constant_imp_connected: + assumes "\f::'a::topological_space \ 'b::real_normed_algebra_1. + \continuous_on s f; finite(f ` s)\ \ \a. \x \ s. f x = a" + shows "connected s" +proof - + { fix t u + assume clt: "closedin (subtopology euclidean s) t" + and clu: "closedin (subtopology euclidean s) u" + and tue: "t \ u = {}" and tus: "t \ u = s" + have conif: "continuous_on s (\x. if x \ t then 0 else 1)" + apply (subst tus [symmetric]) + apply (rule continuous_on_cases_local) + using clt clu tue + apply (auto simp: tus continuous_on_const) + done + have fi: "finite ((\x. if x \ t then 0 else 1) ` s)" + by (rule finite_subset [of _ "{0,1}"]) auto + have "t = {} \ u = {}" + using assms [OF conif fi] tus [symmetric] + by (auto simp: Ball_def) (metis IntI empty_iff one_neq_zero tue) + } + then show ?thesis + by (simp add: connected_closedin_eq) +qed + +lemma continuous_disconnected_range_constant_eq: + "(connected s \ + (\f::'a::topological_space \ 'b::real_normed_algebra_1. + \t. continuous_on s f \ f ` s \ t \ (\y \ t. connected_component_set t y = {y}) + \ (\a::'b. \x \ s. f x = a)))" (is ?thesis1) + and continuous_discrete_range_constant_eq: + "(connected s \ + (\f::'a::topological_space \ 'b::real_normed_algebra_1. + continuous_on s f \ + (\x \ s. \e. 0 < e \ (\y. y \ s \ (f y \ f x) \ e \ norm(f y - f x))) + \ (\a::'b. \x \ s. f x = a)))" (is ?thesis2) + and continuous_finite_range_constant_eq: + "(connected s \ + (\f::'a::topological_space \ 'b::real_normed_algebra_1. + continuous_on s f \ finite (f ` s) + \ (\a::'b. \x \ s. f x = a)))" (is ?thesis3) +proof - + have *: "\s t u v. \s \ t; t \ u; u \ v; v \ s\ + \ (s \ t) \ (s \ u) \ (s \ v)" + by blast + have "?thesis1 \ ?thesis2 \ ?thesis3" + apply (rule *) + using continuous_disconnected_range_constant apply metis + apply clarify + apply (frule discrete_subset_disconnected; blast) + apply (blast dest: finite_implies_discrete) + apply (blast intro!: finite_range_constant_imp_connected) + done + then show ?thesis1 ?thesis2 ?thesis3 + by blast+ +qed + +lemma continuous_discrete_range_constant: + fixes f :: "'a::topological_space \ 'b::real_normed_algebra_1" + assumes s: "connected s" + and "continuous_on s f" + and "\x. x \ s \ \e>0. \y. y \ s \ f y \ f x \ e \ norm (f y - f x)" + shows "\a. \x \ s. f x = a" + using continuous_discrete_range_constant_eq [THEN iffD1, OF s] assms + by blast + +lemma continuous_finite_range_constant: + fixes f :: "'a::topological_space \ 'b::real_normed_algebra_1" + assumes "connected s" + and "continuous_on s f" + and "finite (f ` s)" + shows "\a. \x \ s. f x = a" + using assms continuous_finite_range_constant_eq + by blast + + +text\We can treat even non-rectifiable paths as having a "length" for bounds on analytic functions in open sets.\ + +subsection\Winding Numbers\ + +definition winding_number:: "[real \ complex, complex] \ complex" where + "winding_number \ z \ + @n. \e > 0. \p. valid_path p \ z \ path_image p \ + pathstart p = pathstart \ \ + pathfinish p = pathfinish \ \ + (\t \ {0..1}. norm(\ t - p t) < e) \ + contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" + +lemma winding_number: + assumes "path \" "z \ path_image \" "0 < e" + shows "\p. valid_path p \ z \ path_image p \ + pathstart p = pathstart \ \ + pathfinish p = pathfinish \ \ + (\t \ {0..1}. norm (\ t - p t) < e) \ + contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * winding_number \ z" +proof - + have "path_image \ \ UNIV - {z}" + using assms by blast + then obtain d + where d: "d>0" + and pi_eq: "\h1 h2. valid_path h1 \ valid_path h2 \ + (\t\{0..1}. cmod (h1 t - \ t) < d \ cmod (h2 t - \ t) < d) \ + pathstart h2 = pathstart h1 \ pathfinish h2 = pathfinish h1 \ + path_image h1 \ UNIV - {z} \ path_image h2 \ UNIV - {z} \ + (\f. f holomorphic_on UNIV - {z} \ contour_integral h2 f = contour_integral h1 f)" + using contour_integral_nearby_ends [of "UNIV - {z}" \] assms by (auto simp: open_delete) + then obtain h where h: "polynomial_function h \ pathstart h = pathstart \ \ pathfinish h = pathfinish \ \ + (\t \ {0..1}. norm(h t - \ t) < d/2)" + using path_approx_polynomial_function [OF \path \\, of "d/2"] d by auto + define nn where "nn = 1/(2* pi*\) * contour_integral h (\w. 1/(w - z))" + have "\n. \e > 0. \p. valid_path p \ z \ path_image p \ + pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ + (\t \ {0..1}. norm(\ t - p t) < e) \ + contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" + (is "\n. \e > 0. ?PP e n") + proof (rule_tac x=nn in exI, clarify) + fix e::real + assume e: "e>0" + obtain p where p: "polynomial_function p \ + pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ (\t\{0..1}. cmod (p t - \ t) < min e (d / 2))" + using path_approx_polynomial_function [OF \path \\, of "min e (d/2)"] d \0 by auto + have "(\w. 1 / (w - z)) holomorphic_on UNIV - {z}" + by (auto simp: intro!: holomorphic_intros) + then show "?PP e nn" + apply (rule_tac x=p in exI) + using pi_eq [of h p] h p d + apply (auto simp: valid_path_polynomial_function norm_minus_commute nn_def) + done + qed + then show ?thesis + unfolding winding_number_def + apply (rule someI2_ex) + apply (blast intro: \0) + done +qed + +lemma winding_number_unique: + assumes \: "path \" "z \ path_image \" + and pi: + "\e. e>0 \ \p. valid_path p \ z \ path_image p \ + pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ + (\t \ {0..1}. norm (\ t - p t) < e) \ + contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" + shows "winding_number \ z = n" +proof - + have "path_image \ \ UNIV - {z}" + using assms by blast + then obtain e + where e: "e>0" + and pi_eq: "\h1 h2 f. \valid_path h1; valid_path h2; + (\t\{0..1}. cmod (h1 t - \ t) < e \ cmod (h2 t - \ t) < e); + pathstart h2 = pathstart h1; pathfinish h2 = pathfinish h1; f holomorphic_on UNIV - {z}\ \ + contour_integral h2 f = contour_integral h1 f" + using contour_integral_nearby_ends [of "UNIV - {z}" \] assms by (auto simp: open_delete) + obtain p where p: + "valid_path p \ z \ path_image p \ + pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ + (\t \ {0..1}. norm (\ t - p t) < e) \ + contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" + using pi [OF e] by blast + obtain q where q: + "valid_path q \ z \ path_image q \ + pathstart q = pathstart \ \ pathfinish q = pathfinish \ \ + (\t\{0..1}. cmod (\ t - q t) < e) \ contour_integral q (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number \ z" + using winding_number [OF \ e] by blast + have "2 * complex_of_real pi * \ * n = contour_integral p (\w. 1 / (w - z))" + using p by auto + also have "... = contour_integral q (\w. 1 / (w - z))" + apply (rule pi_eq) + using p q + by (auto simp: valid_path_polynomial_function norm_minus_commute intro!: holomorphic_intros) + also have "... = 2 * complex_of_real pi * \ * winding_number \ z" + using q by auto + finally have "2 * complex_of_real pi * \ * n = 2 * complex_of_real pi * \ * winding_number \ z" . + then show ?thesis + by simp +qed + +lemma winding_number_unique_loop: + assumes \: "path \" "z \ path_image \" + and loop: "pathfinish \ = pathstart \" + and pi: + "\e. e>0 \ \p. valid_path p \ z \ path_image p \ + pathfinish p = pathstart p \ + (\t \ {0..1}. norm (\ t - p t) < e) \ + contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" + shows "winding_number \ z = n" +proof - + have "path_image \ \ UNIV - {z}" + using assms by blast + then obtain e + where e: "e>0" + and pi_eq: "\h1 h2 f. \valid_path h1; valid_path h2; + (\t\{0..1}. cmod (h1 t - \ t) < e \ cmod (h2 t - \ t) < e); + pathfinish h1 = pathstart h1; pathfinish h2 = pathstart h2; f holomorphic_on UNIV - {z}\ \ + contour_integral h2 f = contour_integral h1 f" + using contour_integral_nearby_loops [of "UNIV - {z}" \] assms by (auto simp: open_delete) + obtain p where p: + "valid_path p \ z \ path_image p \ + pathfinish p = pathstart p \ + (\t \ {0..1}. norm (\ t - p t) < e) \ + contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" + using pi [OF e] by blast + obtain q where q: + "valid_path q \ z \ path_image q \ + pathstart q = pathstart \ \ pathfinish q = pathfinish \ \ + (\t\{0..1}. cmod (\ t - q t) < e) \ contour_integral q (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number \ z" + using winding_number [OF \ e] by blast + have "2 * complex_of_real pi * \ * n = contour_integral p (\w. 1 / (w - z))" + using p by auto + also have "... = contour_integral q (\w. 1 / (w - z))" + apply (rule pi_eq) + using p q loop + by (auto simp: valid_path_polynomial_function norm_minus_commute intro!: holomorphic_intros) + also have "... = 2 * complex_of_real pi * \ * winding_number \ z" + using q by auto + finally have "2 * complex_of_real pi * \ * n = 2 * complex_of_real pi * \ * winding_number \ z" . + then show ?thesis + by simp +qed + +lemma winding_number_valid_path: + assumes "valid_path \" "z \ path_image \" + shows "winding_number \ z = 1/(2*pi*\) * contour_integral \ (\w. 1/(w - z))" + using assms by (auto simp: valid_path_imp_path intro!: winding_number_unique) + +lemma has_contour_integral_winding_number: + assumes \: "valid_path \" "z \ path_image \" + shows "((\w. 1/(w - z)) has_contour_integral (2*pi*\*winding_number \ z)) \" +by (simp add: winding_number_valid_path has_contour_integral_integral contour_integrable_inversediff assms) + +lemma winding_number_trivial [simp]: "z \ a \ winding_number(linepath a a) z = 0" + by (simp add: winding_number_valid_path) + +lemma winding_number_subpath_trivial [simp]: "z \ g x \ winding_number (subpath x x g) z = 0" + by (simp add: path_image_subpath winding_number_valid_path) + +lemma winding_number_join: + assumes g1: "path g1" "z \ path_image g1" + and g2: "path g2" "z \ path_image g2" + and "pathfinish g1 = pathstart g2" + shows "winding_number(g1 +++ g2) z = winding_number g1 z + winding_number g2 z" + apply (rule winding_number_unique) + using assms apply (simp_all add: not_in_path_image_join) + apply (frule winding_number [OF g2]) + apply (frule winding_number [OF g1], clarify) + apply (rename_tac p2 p1) + apply (rule_tac x="p1+++p2" in exI) + apply (simp add: not_in_path_image_join contour_integrable_inversediff algebra_simps) + apply (auto simp: joinpaths_def) + done + +lemma winding_number_reversepath: + assumes "path \" "z \ path_image \" + shows "winding_number(reversepath \) z = - (winding_number \ z)" + apply (rule winding_number_unique) + using assms + apply simp_all + apply (frule winding_number [OF assms], clarify) + apply (rule_tac x="reversepath p" in exI) + apply (simp add: contour_integral_reversepath contour_integrable_inversediff valid_path_imp_reverse) + apply (auto simp: reversepath_def) + done + +lemma winding_number_shiftpath: + assumes \: "path \" "z \ path_image \" + and "pathfinish \ = pathstart \" "a \ {0..1}" + shows "winding_number(shiftpath a \) z = winding_number \ z" + apply (rule winding_number_unique_loop) + using assms + apply (simp_all add: path_shiftpath path_image_shiftpath pathfinish_shiftpath pathstart_shiftpath) + apply (frule winding_number [OF \], clarify) + apply (rule_tac x="shiftpath a p" in exI) + apply (simp add: contour_integral_shiftpath path_image_shiftpath pathfinish_shiftpath pathstart_shiftpath valid_path_shiftpath) + apply (auto simp: shiftpath_def) + done + +lemma winding_number_split_linepath: + assumes "c \ closed_segment a b" "z \ closed_segment a b" + shows "winding_number(linepath a b) z = winding_number(linepath a c) z + winding_number(linepath c b) z" +proof - + have "z \ closed_segment a c" "z \ closed_segment c b" + using assms apply (meson convex_contains_segment convex_segment ends_in_segment(1) subsetCE) + using assms by (meson convex_contains_segment convex_segment ends_in_segment(2) subsetCE) + then show ?thesis + using assms + by (simp add: winding_number_valid_path contour_integral_split_linepath [symmetric] continuous_on_inversediff field_simps) +qed + +lemma winding_number_cong: + "(\t. \0 \ t; t \ 1\ \ p t = q t) \ winding_number p z = winding_number q z" + by (simp add: winding_number_def pathstart_def pathfinish_def) + +lemma winding_number_offset: "winding_number p z = winding_number (\w. p w - z) 0" + apply (simp add: winding_number_def contour_integral_integral path_image_def valid_path_def pathstart_def pathfinish_def) + apply (intro ext arg_cong [where f = Eps] arg_cong [where f = All] imp_cong refl, safe) + apply (rename_tac g) + apply (rule_tac x="\t. g t - z" in exI) + apply (force simp: vector_derivative_def has_vector_derivative_diff_const piecewise_C1_differentiable_diff C1_differentiable_imp_piecewise) + apply (rename_tac g) + apply (rule_tac x="\t. g t + z" in exI) + apply (simp add: piecewise_C1_differentiable_add vector_derivative_def has_vector_derivative_add_const C1_differentiable_imp_piecewise) + apply (force simp: algebra_simps) + done + +(* A combined theorem deducing several things piecewise.*) +lemma winding_number_join_pos_combined: + "\valid_path \1; z \ path_image \1; 0 < Re(winding_number \1 z); + valid_path \2; z \ path_image \2; 0 < Re(winding_number \2 z); pathfinish \1 = pathstart \2\ + \ valid_path(\1 +++ \2) \ z \ path_image(\1 +++ \2) \ 0 < Re(winding_number(\1 +++ \2) z)" + by (simp add: valid_path_join path_image_join winding_number_join valid_path_imp_path) + + +(* Useful sufficient conditions for the winding number to be positive etc.*) + +lemma Re_winding_number: + "\valid_path \; z \ path_image \\ + \ Re(winding_number \ z) = Im(contour_integral \ (\w. 1/(w - z))) / (2*pi)" +by (simp add: winding_number_valid_path field_simps Re_divide power2_eq_square) + +lemma winding_number_pos_le: + assumes \: "valid_path \" "z \ path_image \" + and ge: "\x. \0 < x; x < 1\ \ 0 \ Im (vector_derivative \ (at x) * cnj(\ x - z))" + shows "0 \ Re(winding_number \ z)" +proof - + have *: "0 \ Im (vector_derivative \ (at x) / (\ x - z))" if x: "0 < x" "x < 1" for x + using ge by (simp add: Complex.Im_divide algebra_simps x) + show ?thesis + apply (simp add: Re_winding_number [OF \] field_simps) + apply (rule has_integral_component_nonneg + [of \ "\x. if x \ {0<..<1} + then 1/(\ x - z) * vector_derivative \ (at x) else 0", simplified]) + prefer 3 apply (force simp: *) + apply (simp add: Basis_complex_def) + apply (rule has_integral_spike_interior [of 0 1 _ "\x. 1/(\ x - z) * vector_derivative \ (at x)"]) + apply simp + apply (simp only: box_real) + apply (subst has_contour_integral [symmetric]) + using \ + apply (simp add: contour_integrable_inversediff has_contour_integral_integral) + done +qed + +lemma winding_number_pos_lt_lemma: + assumes \: "valid_path \" "z \ path_image \" + and e: "0 < e" + and ge: "\x. \0 < x; x < 1\ \ e \ Im (vector_derivative \ (at x) / (\ x - z))" + shows "0 < Re(winding_number \ z)" +proof - + have "e \ Im (contour_integral \ (\w. 1 / (w - z)))" + apply (rule has_integral_component_le + [of \ "\x. \*e" "\*e" "{0..1}" + "\x. if x \ {0<..<1} then 1/(\ x - z) * vector_derivative \ (at x) else \*e" + "contour_integral \ (\w. 1/(w - z))", simplified]) + using e + apply (simp_all add: Basis_complex_def) + using has_integral_const_real [of _ 0 1] apply force + apply (rule has_integral_spike_interior [of 0 1 _ "\x. 1/(\ x - z) * vector_derivative \ (at x)", simplified box_real]) + apply simp + apply (subst has_contour_integral [symmetric]) + using \ + apply (simp_all add: contour_integrable_inversediff has_contour_integral_integral ge) + done + with e show ?thesis + by (simp add: Re_winding_number [OF \] field_simps) +qed + +lemma winding_number_pos_lt: + assumes \: "valid_path \" "z \ path_image \" + and e: "0 < e" + and ge: "\x. \0 < x; x < 1\ \ e \ Im (vector_derivative \ (at x) * cnj(\ x - z))" + shows "0 < Re (winding_number \ z)" +proof - + have bm: "bounded ((\w. w - z) ` (path_image \))" + using bounded_translation [of _ "-z"] \ by (simp add: bounded_valid_path_image) + then obtain B where B: "B > 0" and Bno: "\x. x \ (\w. w - z) ` (path_image \) \ norm x \ B" + using bounded_pos [THEN iffD1, OF bm] by blast + { fix x::real assume x: "0 < x" "x < 1" + then have B2: "cmod (\ x - z)^2 \ B^2" using Bno [of "\ x - z"] + by (simp add: path_image_def power2_eq_square mult_mono') + with x have "\ x \ z" using \ + using path_image_def by fastforce + then have "e / B\<^sup>2 \ Im (vector_derivative \ (at x) * cnj (\ x - z)) / (cmod (\ x - z))\<^sup>2" + using B ge [OF x] B2 e + apply (rule_tac y="e / (cmod (\ x - z))\<^sup>2" in order_trans) + apply (auto simp: divide_left_mono divide_right_mono) + done + then have "e / B\<^sup>2 \ Im (vector_derivative \ (at x) / (\ x - z))" + by (simp add: Im_divide_Reals complex_div_cnj [of _ "\ x - z" for x] del: complex_cnj_diff times_complex.sel) + } note * = this + show ?thesis + using e B by (simp add: * winding_number_pos_lt_lemma [OF \, of "e/B^2"]) +qed + +subsection\The winding number is an integer\ + +text\Proof from the book Complex Analysis by Lars V. Ahlfors, Chapter 4, section 2.1, + Also on page 134 of Serge Lang's book with the name title, etc.\ + +lemma exp_fg: + fixes z::complex + assumes g: "(g has_vector_derivative g') (at x within s)" + and f: "(f has_vector_derivative (g' / (g x - z))) (at x within s)" + and z: "g x \ z" + shows "((\x. exp(-f x) * (g x - z)) has_vector_derivative 0) (at x within s)" +proof - + have *: "(exp o (\x. (- f x)) has_vector_derivative - (g' / (g x - z)) * exp (- f x)) (at x within s)" + using assms unfolding has_vector_derivative_def scaleR_conv_of_real + by (auto intro!: derivative_eq_intros) + show ?thesis + apply (rule has_vector_derivative_eq_rhs) + apply (rule bounded_bilinear.has_vector_derivative [OF bounded_bilinear_mult]) + using z + apply (auto simp: intro!: derivative_eq_intros * [unfolded o_def] g) + done +qed + +lemma winding_number_exp_integral: + fixes z::complex + assumes \: "\ piecewise_C1_differentiable_on {a..b}" + and ab: "a \ b" + and z: "z \ \ ` {a..b}" + shows "(\x. vector_derivative \ (at x) / (\ x - z)) integrable_on {a..b}" + (is "?thesis1") + "exp (- (integral {a..b} (\x. vector_derivative \ (at x) / (\ x - z)))) * (\ b - z) = \ a - z" + (is "?thesis2") +proof - + let ?D\ = "\x. vector_derivative \ (at x)" + have [simp]: "\x. a \ x \ x \ b \ \ x \ z" + using z by force + have cong: "continuous_on {a..b} \" + using \ by (simp add: piecewise_C1_differentiable_on_def) + obtain k where fink: "finite k" and g_C1_diff: "\ C1_differentiable_on ({a..b} - k)" + using \ by (force simp: piecewise_C1_differentiable_on_def) + have o: "open ({a<..finite k\ by (simp add: finite_imp_closed open_Diff) + moreover have "{a<.. {a..b} - k" + by force + ultimately have g_diff_at: "\x. \x \ k; x \ {a<.. \ \ differentiable at x" + by (metis Diff_iff differentiable_on_subset C1_diff_imp_diff [OF g_C1_diff] differentiable_on_def differentiable_within_open) + { fix w + assume "w \ z" + have "continuous_on (ball w (cmod (w - z))) (\w. 1 / (w - z))" + by (auto simp: dist_norm intro!: continuous_intros) + moreover have "\x. cmod (w - x) < cmod (w - z) \ \f'. ((\w. 1 / (w - z)) has_field_derivative f') (at x)" + by (auto simp: intro!: derivative_eq_intros) + ultimately have "\h. \y. norm(y - w) < norm(w - z) \ (h has_field_derivative 1/(y - z)) (at y)" + using holomorphic_convex_primitive [of "ball w (norm(w - z))" "{}" "\w. 1/(w - z)"] + by (simp add: field_differentiable_def Ball_def dist_norm at_within_open_NO_MATCH norm_minus_commute) + } + then obtain h where h: "\w y. w \ z \ norm(y - w) < norm(w - z) \ (h w has_field_derivative 1/(y - z)) (at y)" + by meson + have exy: "\y. ((\x. inverse (\ x - z) * ?D\ x) has_integral y) {a..b}" + unfolding integrable_on_def [symmetric] + apply (rule contour_integral_local_primitive_any [OF piecewise_C1_imp_differentiable [OF \], of "-{z}"]) + apply (rename_tac w) + apply (rule_tac x="norm(w - z)" in exI) + apply (simp_all add: inverse_eq_divide) + apply (metis has_field_derivative_at_within h) + done + have vg_int: "(\x. ?D\ x / (\ x - z)) integrable_on {a..b}" + unfolding box_real [symmetric] divide_inverse_commute + by (auto intro!: exy integrable_subinterval simp add: integrable_on_def ab) + with ab show ?thesis1 + by (simp add: divide_inverse_commute integral_def integrable_on_def) + { fix t + assume t: "t \ {a..b}" + have cball: "continuous_on (ball (\ t) (dist (\ t) z)) (\x. inverse (x - z))" + using z by (auto intro!: continuous_intros simp: dist_norm) + have icd: "\x. cmod (\ t - x) < cmod (\ t - z) \ (\w. inverse (w - z)) field_differentiable at x" + unfolding field_differentiable_def by (force simp: intro!: derivative_eq_intros) + obtain h where h: "\x. cmod (\ t - x) < cmod (\ t - z) \ + (h has_field_derivative inverse (x - z)) (at x within {y. cmod (\ t - y) < cmod (\ t - z)})" + using holomorphic_convex_primitive [where f = "\w. inverse(w - z)", OF convex_ball finite.emptyI cball icd] + by simp (auto simp: ball_def dist_norm that) + { fix x D + assume x: "x \ k" "a < x" "x < b" + then have "x \ interior ({a..b} - k)" + using open_subset_interior [OF o] by fastforce + then have con: "isCont (\x. ?D\ x) x" + using g_C1_diff x by (auto simp: C1_differentiable_on_eq intro: continuous_on_interior) + then have con_vd: "continuous (at x within {a..b}) (\x. ?D\ x)" + by (rule continuous_at_imp_continuous_within) + have gdx: "\ differentiable at x" + using x by (simp add: g_diff_at) + have "((\c. exp (- integral {a..c} (\x. vector_derivative \ (at x) / (\ x - z))) * (\ c - z)) has_derivative (\h. 0)) + (at x within {a..b})" + using x gdx t + apply (clarsimp simp add: differentiable_iff_scaleR) + apply (rule exp_fg [unfolded has_vector_derivative_def, simplified], blast intro: has_derivative_at_within) + apply (simp_all add: has_vector_derivative_def [symmetric]) + apply (rule has_vector_derivative_eq_rhs [OF integral_has_vector_derivative_continuous_at]) + apply (rule con_vd continuous_intros cong vg_int | simp add: continuous_at_imp_continuous_within has_vector_derivative_continuous vector_derivative_at)+ + done + } note * = this + have "exp (- (integral {a..t} (\x. ?D\ x / (\ x - z)))) * (\ t - z) =\ a - z" + apply (rule has_derivative_zero_unique_strong_interval [of "{a,b} \ k" a b]) + using t + apply (auto intro!: * continuous_intros fink cong indefinite_integral_continuous [OF vg_int] simp add: ab)+ + done + } + with ab show ?thesis2 + by (simp add: divide_inverse_commute integral_def) +qed + +corollary winding_number_exp_2pi: + "\path p; z \ path_image p\ + \ pathfinish p - z = exp (2 * pi * \ * winding_number p z) * (pathstart p - z)" +using winding_number [of p z 1] unfolding valid_path_def path_image_def pathstart_def pathfinish_def + by (force dest: winding_number_exp_integral(2) [of _ 0 1 z] simp: field_simps contour_integral_integral exp_minus) + + +subsection\The version with complex integers and equality\ + +lemma integer_winding_number_eq: + assumes \: "path \" and z: "z \ path_image \" + shows "winding_number \ z \ \ \ pathfinish \ = pathstart \" +proof - + have *: "\i::complex. \g0 g1. \i \ 0; g0 \ z; (g1 - z) / i = g0 - z\ \ (i = 1 \ g1 = g0)" + by (simp add: field_simps) algebra + obtain p where p: "valid_path p" "z \ path_image p" + "pathstart p = pathstart \" "pathfinish p = pathfinish \" + "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number \ z" + using winding_number [OF assms, of 1] by auto + have [simp]: "(winding_number \ z \ \) = (exp (contour_integral p (\w. 1 / (w - z))) = 1)" + using p by (simp add: exp_eq_1 complex_is_Int_iff) + have "winding_number p z \ \ \ pathfinish p = pathstart p" + using p z + apply (simp add: winding_number_valid_path valid_path_def path_image_def pathstart_def pathfinish_def) + using winding_number_exp_integral(2) [of p 0 1 z] + apply (simp add: field_simps contour_integral_integral exp_minus) + apply (rule *) + apply (auto simp: path_image_def field_simps) + done + then show ?thesis using p + by (auto simp: winding_number_valid_path) +qed + +theorem integer_winding_number: + "\path \; pathfinish \ = pathstart \; z \ path_image \\ \ winding_number \ z \ \" +by (metis integer_winding_number_eq) + + +text\If the winding number's magnitude is at least one, then the path must contain points in every direction.*) + We can thus bound the winding number of a path that doesn't intersect a given ray. \ + +lemma winding_number_pos_meets: + fixes z::complex + assumes \: "valid_path \" and z: "z \ path_image \" and 1: "Re (winding_number \ z) \ 1" + and w: "w \ z" + shows "\a::real. 0 < a \ z + a*(w - z) \ path_image \" +proof - + have [simp]: "\x. 0 \ x \ x \ 1 \ \ x \ z" + using z by (auto simp: path_image_def) + have [simp]: "z \ \ ` {0..1}" + using path_image_def z by auto + have gpd: "\ piecewise_C1_differentiable_on {0..1}" + using \ valid_path_def by blast + define r where "r = (w - z) / (\ 0 - z)" + have [simp]: "r \ 0" + using w z by (auto simp: r_def) + have "Arg r \ 2*pi" + by (simp add: Arg less_eq_real_def) + also have "... \ Im (integral {0..1} (\x. vector_derivative \ (at x) / (\ x - z)))" + using 1 + apply (simp add: winding_number_valid_path [OF \ z] contour_integral_integral) + apply (simp add: Complex.Re_divide field_simps power2_eq_square) + done + finally have "Arg r \ Im (integral {0..1} (\x. vector_derivative \ (at x) / (\ x - z)))" . + then have "\t. t \ {0..1} \ Im(integral {0..t} (\x. vector_derivative \ (at x)/(\ x - z))) = Arg r" + apply (simp add:) + apply (rule IVT') + apply (simp_all add: Arg_ge_0) + apply (intro continuous_intros indefinite_integral_continuous winding_number_exp_integral [OF gpd]; simp) + done + then obtain t where t: "t \ {0..1}" + and eqArg: "Im (integral {0..t} (\x. vector_derivative \ (at x)/(\ x - z))) = Arg r" + by blast + define i where "i = integral {0..t} (\x. vector_derivative \ (at x) / (\ x - z))" + have iArg: "Arg r = Im i" + using eqArg by (simp add: i_def) + have gpdt: "\ piecewise_C1_differentiable_on {0..t}" + by (metis atLeastAtMost_iff atLeastatMost_subset_iff order_refl piecewise_C1_differentiable_on_subset gpd t) + have "exp (- i) * (\ t - z) = \ 0 - z" + unfolding i_def + apply (rule winding_number_exp_integral [OF gpdt]) + using t z unfolding path_image_def + apply force+ + done + then have *: "\ t - z = exp i * (\ 0 - z)" + by (simp add: exp_minus field_simps) + then have "(w - z) = r * (\ 0 - z)" + by (simp add: r_def) + then have "z + complex_of_real (exp (Re i)) * (w - z) / complex_of_real (cmod r) = \ t" + apply (simp add:) + apply (subst Complex_Transcendental.Arg_eq [of r]) + apply (simp add: iArg) + using * + apply (simp add: exp_eq_polar field_simps) + done + with t show ?thesis + by (rule_tac x="exp(Re i) / norm r" in exI) (auto simp: path_image_def) +qed + +lemma winding_number_big_meets: + fixes z::complex + assumes \: "valid_path \" and z: "z \ path_image \" and "\Re (winding_number \ z)\ \ 1" + and w: "w \ z" + shows "\a::real. 0 < a \ z + a*(w - z) \ path_image \" +proof - + { assume "Re (winding_number \ z) \ - 1" + then have "Re (winding_number (reversepath \) z) \ 1" + by (simp add: \ valid_path_imp_path winding_number_reversepath z) + moreover have "valid_path (reversepath \)" + using \ valid_path_imp_reverse by auto + moreover have "z \ path_image (reversepath \)" + by (simp add: z) + ultimately have "\a::real. 0 < a \ z + a*(w - z) \ path_image (reversepath \)" + using winding_number_pos_meets w by blast + then have ?thesis + by simp + } + then show ?thesis + using assms + by (simp add: abs_if winding_number_pos_meets split: if_split_asm) +qed + +lemma winding_number_less_1: + fixes z::complex + shows + "\valid_path \; z \ path_image \; w \ z; + \a::real. 0 < a \ z + a*(w - z) \ path_image \\ + \ \Re(winding_number \ z)\ < 1" + by (auto simp: not_less dest: winding_number_big_meets) + +text\One way of proving that WN=1 for a loop.\ +lemma winding_number_eq_1: + fixes z::complex + assumes \: "valid_path \" and z: "z \ path_image \" and loop: "pathfinish \ = pathstart \" + and 0: "0 < Re(winding_number \ z)" and 2: "Re(winding_number \ z) < 2" + shows "winding_number \ z = 1" +proof - + have "winding_number \ z \ Ints" + by (simp add: \ integer_winding_number loop valid_path_imp_path z) + then show ?thesis + using 0 2 by (auto simp: Ints_def) +qed + + +subsection\Continuity of winding number and invariance on connected sets.\ + +lemma continuous_at_winding_number: + fixes z::complex + assumes \: "path \" and z: "z \ path_image \" + shows "continuous (at z) (winding_number \)" +proof - + obtain e where "e>0" and cbg: "cball z e \ - path_image \" + using open_contains_cball [of "- path_image \"] z + by (force simp: closed_def [symmetric] closed_path_image [OF \]) + then have ppag: "path_image \ \ - cball z (e/2)" + by (force simp: cball_def dist_norm) + have oc: "open (- cball z (e / 2))" + by (simp add: closed_def [symmetric]) + obtain d where "d>0" and pi_eq: + "\h1 h2. \valid_path h1; valid_path h2; + (\t\{0..1}. cmod (h1 t - \ t) < d \ cmod (h2 t - \ t) < d); + pathstart h2 = pathstart h1; pathfinish h2 = pathfinish h1\ + \ + path_image h1 \ - cball z (e / 2) \ + path_image h2 \ - cball z (e / 2) \ + (\f. f holomorphic_on - cball z (e / 2) \ contour_integral h2 f = contour_integral h1 f)" + using contour_integral_nearby_ends [OF oc \ ppag] by metis + obtain p where p: "valid_path p" "z \ path_image p" + "pathstart p = pathstart \ \ pathfinish p = pathfinish \" + and pg: "\t. t\{0..1} \ cmod (\ t - p t) < min d e / 2" + and pi: "contour_integral p (\x. 1 / (x - z)) = complex_of_real (2 * pi) * \ * winding_number \ z" + using winding_number [OF \ z, of "min d e / 2"] \d>0\ \e>0\ by auto + { fix w + assume d2: "cmod (w - z) < d/2" and e2: "cmod (w - z) < e/2" + then have wnotp: "w \ path_image p" + using cbg \d>0\ \e>0\ + apply (simp add: path_image_def cball_def dist_norm, clarify) + apply (frule pg) + apply (drule_tac c="\ x" in subsetD) + apply (auto simp: less_eq_real_def norm_minus_commute norm_triangle_half_l) + done + have wnotg: "w \ path_image \" + using cbg e2 \e>0\ by (force simp: dist_norm norm_minus_commute) + { fix k::real + assume k: "k>0" + then obtain q where q: "valid_path q" "w \ path_image q" + "pathstart q = pathstart \ \ pathfinish q = pathfinish \" + and qg: "\t. t \ {0..1} \ cmod (\ t - q t) < min k (min d e) / 2" + and qi: "contour_integral q (\u. 1 / (u - w)) = complex_of_real (2 * pi) * \ * winding_number \ w" + using winding_number [OF \ wnotg, of "min k (min d e) / 2"] \d>0\ \e>0\ k + by (force simp: min_divide_distrib_right) + have "contour_integral p (\u. 1 / (u - w)) = contour_integral q (\u. 1 / (u - w))" + apply (rule pi_eq [OF \valid_path q\ \valid_path p\, THEN conjunct2, THEN conjunct2, rule_format]) + apply (frule pg) + apply (frule qg) + using p q \d>0\ e2 + apply (auto simp: dist_norm norm_minus_commute intro!: holomorphic_intros) + done + then have "contour_integral p (\x. 1 / (x - w)) = complex_of_real (2 * pi) * \ * winding_number \ w" + by (simp add: pi qi) + } note pip = this + have "path p" + using p by (simp add: valid_path_imp_path) + then have "winding_number p w = winding_number \ w" + apply (rule winding_number_unique [OF _ wnotp]) + apply (rule_tac x=p in exI) + apply (simp add: p wnotp min_divide_distrib_right pip) + done + } note wnwn = this + obtain pe where "pe>0" and cbp: "cball z (3 / 4 * pe) \ - path_image p" + using p open_contains_cball [of "- path_image p"] + by (force simp: closed_def [symmetric] closed_path_image [OF valid_path_imp_path]) + obtain L + where "L>0" + and L: "\f B. \f holomorphic_on - cball z (3 / 4 * pe); + \z \ - cball z (3 / 4 * pe). cmod (f z) \ B\ \ + cmod (contour_integral p f) \ L * B" + using contour_integral_bound_exists [of "- cball z (3/4*pe)" p] cbp \valid_path p\ by force + { fix e::real and w::complex + assume e: "0 < e" and w: "cmod (w - z) < pe/4" "cmod (w - z) < e * pe\<^sup>2 / (8 * L)" + then have [simp]: "w \ path_image p" + using cbp p(2) \0 < pe\ + by (force simp: dist_norm norm_minus_commute path_image_def cball_def) + have [simp]: "contour_integral p (\x. 1/(x - w)) - contour_integral p (\x. 1/(x - z)) = + contour_integral p (\x. 1/(x - w) - 1/(x - z))" + by (simp add: p contour_integrable_inversediff contour_integral_diff) + { fix x + assume pe: "3/4 * pe < cmod (z - x)" + have "cmod (w - x) < pe/4 + cmod (z - x)" + by (meson add_less_cancel_right norm_diff_triangle_le order_refl order_trans_rules(21) w(1)) + then have wx: "cmod (w - x) < 4/3 * cmod (z - x)" using pe by simp + have "cmod (z - x) \ cmod (z - w) + cmod (w - x)" + using norm_diff_triangle_le by blast + also have "... < pe/4 + cmod (w - x)" + using w by (simp add: norm_minus_commute) + finally have "pe/2 < cmod (w - x)" + using pe by auto + then have "(pe/2)^2 < cmod (w - x) ^ 2" + apply (rule power_strict_mono) + using \pe>0\ by auto + then have pe2: "pe^2 < 4 * cmod (w - x) ^ 2" + by (simp add: power_divide) + have "8 * L * cmod (w - z) < e * pe\<^sup>2" + using w \L>0\ by (simp add: field_simps) + also have "... < e * 4 * cmod (w - x) * cmod (w - x)" + using pe2 \e>0\ by (simp add: power2_eq_square) + also have "... < e * 4 * cmod (w - x) * (4/3 * cmod (z - x))" + using wx + apply (rule mult_strict_left_mono) + using pe2 e not_less_iff_gr_or_eq by fastforce + finally have "L * cmod (w - z) < 2/3 * e * cmod (w - x) * cmod (z - x)" + by simp + also have "... \ e * cmod (w - x) * cmod (z - x)" + using e by simp + finally have Lwz: "L * cmod (w - z) < e * cmod (w - x) * cmod (z - x)" . + have "L * cmod (1 / (x - w) - 1 / (x - z)) \ e" + apply (cases "x=z \ x=w") + using pe \pe>0\ w \L>0\ + apply (force simp: norm_minus_commute) + using wx w(2) \L>0\ pe pe2 Lwz + apply (auto simp: divide_simps mult_less_0_iff norm_minus_commute norm_divide norm_mult power2_eq_square) + done + } note L_cmod_le = this + have *: "cmod (contour_integral p (\x. 1 / (x - w) - 1 / (x - z))) \ L * (e * pe\<^sup>2 / L / 4 * (inverse (pe / 2))\<^sup>2)" + apply (rule L) + using \pe>0\ w + apply (force simp: dist_norm norm_minus_commute intro!: holomorphic_intros) + using \pe>0\ w \L>0\ + apply (auto simp: cball_def dist_norm field_simps L_cmod_le simp del: less_divide_eq_numeral1 le_divide_eq_numeral1) + done + have "cmod (contour_integral p (\x. 1 / (x - w)) - contour_integral p (\x. 1 / (x - z))) < 2*e" + apply (simp add:) + apply (rule le_less_trans [OF *]) + using \L>0\ e + apply (force simp: field_simps) + done + then have "cmod (winding_number p w - winding_number p z) < e" + using pi_ge_two e + by (force simp: winding_number_valid_path p field_simps norm_divide norm_mult intro: less_le_trans) + } note cmod_wn_diff = this + then have "isCont (winding_number p) z" + apply (simp add: continuous_at_eps_delta, clarify) + apply (rule_tac x="min (pe/4) (e/2*pe^2/L/4)" in exI) + using \pe>0\ \L>0\ + apply (simp add: dist_norm cmod_wn_diff) + done + then show ?thesis + apply (rule continuous_transform_within [where d = "min d e / 2"]) + apply (auto simp: \d>0\ \e>0\ dist_norm wnwn) + done +qed + +corollary continuous_on_winding_number: + "path \ \ continuous_on (- path_image \) (\w. winding_number \ w)" + by (simp add: continuous_at_imp_continuous_on continuous_at_winding_number) + + +subsection\The winding number is constant on a connected region\ + +lemma winding_number_constant: + assumes \: "path \" and loop: "pathfinish \ = pathstart \" and cs: "connected s" and sg: "s \ path_image \ = {}" + shows "\k. \z \ s. winding_number \ z = k" +proof - + { fix y z + assume ne: "winding_number \ y \ winding_number \ z" + assume "y \ s" "z \ s" + then have "winding_number \ y \ \" "winding_number \ z \ \" + using integer_winding_number [OF \ loop] sg \y \ s\ by auto + with ne have "1 \ cmod (winding_number \ y - winding_number \ z)" + by (auto simp: Ints_def of_int_diff [symmetric] simp del: of_int_diff) + } note * = this + show ?thesis + apply (rule continuous_discrete_range_constant [OF cs]) + using continuous_on_winding_number [OF \] sg + apply (metis Diff_Compl Diff_eq_empty_iff continuous_on_subset) + apply (rule_tac x=1 in exI) + apply (auto simp: *) + done +qed + +lemma winding_number_eq: + "\path \; pathfinish \ = pathstart \; w \ s; z \ s; connected s; s \ path_image \ = {}\ + \ winding_number \ w = winding_number \ z" +using winding_number_constant by fastforce + +lemma open_winding_number_levelsets: + assumes \: "path \" and loop: "pathfinish \ = pathstart \" + shows "open {z. z \ path_image \ \ winding_number \ z = k}" +proof - + have op: "open (- path_image \)" + by (simp add: closed_path_image \ open_Compl) + { fix z assume z: "z \ path_image \" and k: "k = winding_number \ z" + obtain e where e: "e>0" "ball z e \ - path_image \" + using open_contains_ball [of "- path_image \"] op z + by blast + have "\e>0. \y. dist y z < e \ y \ path_image \ \ winding_number \ y = winding_number \ z" + apply (rule_tac x=e in exI) + using e apply (simp add: dist_norm ball_def norm_minus_commute) + apply (auto simp: dist_norm norm_minus_commute intro!: winding_number_eq [OF assms, where s = "ball z e"]) + done + } then + show ?thesis + by (auto simp: open_dist) +qed + +subsection\Winding number is zero "outside" a curve, in various senses\ + +lemma winding_number_zero_in_outside: + assumes \: "path \" and loop: "pathfinish \ = pathstart \" and z: "z \ outside (path_image \)" + shows "winding_number \ z = 0" +proof - + obtain B::real where "0 < B" and B: "path_image \ \ ball 0 B" + using bounded_subset_ballD [OF bounded_path_image [OF \]] by auto + obtain w::complex where w: "w \ ball 0 (B + 1)" + by (metis abs_of_nonneg le_less less_irrefl mem_ball_0 norm_of_real) + have "- ball 0 (B + 1) \ outside (path_image \)" + apply (rule outside_subset_convex) + using B subset_ball by auto + then have wout: "w \ outside (path_image \)" + using w by blast + moreover obtain k where "\z. z \ outside (path_image \) \ winding_number \ z = k" + using winding_number_constant [OF \ loop, of "outside(path_image \)"] connected_outside + by (metis DIM_complex bounded_path_image dual_order.refl \ outside_no_overlap) + ultimately have "winding_number \ z = winding_number \ w" + using z by blast + also have "... = 0" + proof - + have wnot: "w \ path_image \" using wout by (simp add: outside_def) + { fix e::real assume "0" "pathfinish p = pathfinish \" + and pg1: "(\t. \0 \ t; t \ 1\ \ cmod (p t - \ t) < 1)" + and pge: "(\t. \0 \ t; t \ 1\ \ cmod (p t - \ t) < e)" + using path_approx_polynomial_function [OF \, of "min 1 e"] \e>0\ by force + have pip: "path_image p \ ball 0 (B + 1)" + using B + apply (clarsimp simp add: path_image_def dist_norm ball_def) + apply (frule (1) pg1) + apply (fastforce dest: norm_add_less) + done + then have "w \ path_image p" using w by blast + then have "\p. valid_path p \ w \ path_image p \ + pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ + (\t\{0..1}. cmod (\ t - p t) < e) \ contour_integral p (\wa. 1 / (wa - w)) = 0" + apply (rule_tac x=p in exI) + apply (simp add: p valid_path_polynomial_function) + apply (intro conjI) + using pge apply (simp add: norm_minus_commute) + apply (rule contour_integral_unique [OF Cauchy_theorem_convex_simple [OF _ convex_ball [of 0 "B+1"]]]) + apply (rule holomorphic_intros | simp add: dist_norm)+ + using mem_ball_0 w apply blast + using p apply (simp_all add: valid_path_polynomial_function loop pip) + done + } + then show ?thesis + by (auto intro: winding_number_unique [OF \] simp add: wnot) + qed + finally show ?thesis . +qed + +lemma winding_number_zero_outside: + "\path \; convex s; pathfinish \ = pathstart \; z \ s; path_image \ \ s\ \ winding_number \ z = 0" + by (meson convex_in_outside outside_mono subsetCE winding_number_zero_in_outside) + +lemma winding_number_zero_at_infinity: + assumes \: "path \" and loop: "pathfinish \ = pathstart \" + shows "\B. \z. B \ norm z \ winding_number \ z = 0" +proof - + obtain B::real where "0 < B" and B: "path_image \ \ ball 0 B" + using bounded_subset_ballD [OF bounded_path_image [OF \]] by auto + then show ?thesis + apply (rule_tac x="B+1" in exI, clarify) + apply (rule winding_number_zero_outside [OF \ convex_cball [of 0 B] loop]) + apply (meson less_add_one mem_cball_0 not_le order_trans) + using ball_subset_cball by blast +qed + +lemma winding_number_zero_point: + "\path \; convex s; pathfinish \ = pathstart \; open s; path_image \ \ s\ + \ \z. z \ s \ winding_number \ z = 0" + using outside_compact_in_open [of "path_image \" s] path_image_nonempty winding_number_zero_in_outside + by (fastforce simp add: compact_path_image) + + +text\If a path winds round a set, it winds rounds its inside.\ +lemma winding_number_around_inside: + assumes \: "path \" and loop: "pathfinish \ = pathstart \" + and cls: "closed s" and cos: "connected s" and s_disj: "s \ path_image \ = {}" + and z: "z \ s" and wn_nz: "winding_number \ z \ 0" and w: "w \ s \ inside s" + shows "winding_number \ w = winding_number \ z" +proof - + have ssb: "s \ inside(path_image \)" + proof + fix x :: complex + assume "x \ s" + hence "x \ path_image \" + by (meson disjoint_iff_not_equal s_disj) + thus "x \ inside (path_image \)" + using \x \ s\ by (metis (no_types) ComplI UnE cos \ loop s_disj union_with_outside winding_number_eq winding_number_zero_in_outside wn_nz z) +qed + show ?thesis + apply (rule winding_number_eq [OF \ loop w]) + using z apply blast + apply (simp add: cls connected_with_inside cos) + apply (simp add: Int_Un_distrib2 s_disj, safe) + by (meson ssb inside_inside_compact_connected [OF cls, of "path_image \"] compact_path_image connected_path_image contra_subsetD disjoint_iff_not_equal \ inside_no_overlap) + qed + + +text\Bounding a WN by 1/2 for a path and point in opposite halfspaces.\ +lemma winding_number_subpath_continuous: + assumes \: "valid_path \" and z: "z \ path_image \" + shows "continuous_on {0..1} (\x. winding_number(subpath 0 x \) z)" +proof - + have *: "integral {0..x} (\t. vector_derivative \ (at t) / (\ t - z)) / (2 * of_real pi * \) = + winding_number (subpath 0 x \) z" + if x: "0 \ x" "x \ 1" for x + proof - + have "integral {0..x} (\t. vector_derivative \ (at t) / (\ t - z)) / (2 * of_real pi * \) = + 1 / (2*pi*\) * contour_integral (subpath 0 x \) (\w. 1/(w - z))" + using assms x + apply (simp add: contour_integral_subcontour_integral [OF contour_integrable_inversediff]) + done + also have "... = winding_number (subpath 0 x \) z" + apply (subst winding_number_valid_path) + using assms x + apply (simp_all add: path_image_subpath valid_path_subpath) + by (force simp: path_image_def) + finally show ?thesis . + qed + show ?thesis + apply (rule continuous_on_eq + [where f = "\x. 1 / (2*pi*\) * + integral {0..x} (\t. 1/(\ t - z) * vector_derivative \ (at t))"]) + apply (rule continuous_intros)+ + apply (rule indefinite_integral_continuous) + apply (rule contour_integrable_inversediff [OF assms, unfolded contour_integrable_on]) + using assms + apply (simp add: *) + done +qed + +lemma winding_number_ivt_pos: + assumes \: "valid_path \" and z: "z \ path_image \" and "0 \ w" "w \ Re(winding_number \ z)" + shows "\t \ {0..1}. Re(winding_number(subpath 0 t \) z) = w" + apply (rule ivt_increasing_component_on_1 [of 0 1, where ?k = "1::complex", simplified complex_inner_1_right]) + apply (simp add:) + apply (rule winding_number_subpath_continuous [OF \ z]) + using assms + apply (auto simp: path_image_def image_def) + done + +lemma winding_number_ivt_neg: + assumes \: "valid_path \" and z: "z \ path_image \" and "Re(winding_number \ z) \ w" "w \ 0" + shows "\t \ {0..1}. Re(winding_number(subpath 0 t \) z) = w" + apply (rule ivt_decreasing_component_on_1 [of 0 1, where ?k = "1::complex", simplified complex_inner_1_right]) + apply (simp add:) + apply (rule winding_number_subpath_continuous [OF \ z]) + using assms + apply (auto simp: path_image_def image_def) + done + +lemma winding_number_ivt_abs: + assumes \: "valid_path \" and z: "z \ path_image \" and "0 \ w" "w \ \Re(winding_number \ z)\" + shows "\t \ {0..1}. \Re (winding_number (subpath 0 t \) z)\ = w" + using assms winding_number_ivt_pos [of \ z w] winding_number_ivt_neg [of \ z "-w"] + by force + +lemma winding_number_lt_half_lemma: + assumes \: "valid_path \" and z: "z \ path_image \" and az: "a \ z \ b" and pag: "path_image \ \ {w. a \ w > b}" + shows "Re(winding_number \ z) < 1/2" +proof - + { assume "Re(winding_number \ z) \ 1/2" + then obtain t::real where t: "0 \ t" "t \ 1" and sub12: "Re (winding_number (subpath 0 t \) z) = 1/2" + using winding_number_ivt_pos [OF \ z, of "1/2"] by auto + have gt: "\ t - z = - (of_real (exp (- (2 * pi * Im (winding_number (subpath 0 t \) z)))) * (\ 0 - z))" + using winding_number_exp_2pi [of "subpath 0 t \" z] + apply (simp add: t \ valid_path_imp_path) + using closed_segment_eq_real_ivl path_image_def t z by (fastforce simp: path_image_subpath Euler sub12) + have "b < a \ \ 0" + proof - + have "\ 0 \ {c. b < a \ c}" + by (metis (no_types) pag atLeastAtMost_iff image_subset_iff order_refl path_image_def zero_le_one) + thus ?thesis + by blast + qed + moreover have "b < a \ \ t" + proof - + have "\ t \ {c. b < a \ c}" + by (metis (no_types) pag atLeastAtMost_iff image_subset_iff path_image_def t) + thus ?thesis + by blast + qed + ultimately have "0 < a \ (\ 0 - z)" "0 < a \ (\ t - z)" using az + by (simp add: inner_diff_right)+ + then have False + by (simp add: gt inner_mult_right mult_less_0_iff) + } + then show ?thesis by force +qed + +lemma winding_number_lt_half: + assumes "valid_path \" "a \ z \ b" "path_image \ \ {w. a \ w > b}" + shows "\Re (winding_number \ z)\ < 1/2" +proof - + have "z \ path_image \" using assms by auto + with assms show ?thesis + apply (simp add: winding_number_lt_half_lemma abs_if del: less_divide_eq_numeral1) + apply (metis complex_inner_1_right winding_number_lt_half_lemma [OF valid_path_imp_reverse, of \ z a b] + winding_number_reversepath valid_path_imp_path inner_minus_left path_image_reversepath) + done +qed + +lemma winding_number_le_half: + assumes \: "valid_path \" and z: "z \ path_image \" + and anz: "a \ 0" and azb: "a \ z \ b" and pag: "path_image \ \ {w. a \ w \ b}" + shows "\Re (winding_number \ z)\ \ 1/2" +proof - + { assume wnz_12: "\Re (winding_number \ z)\ > 1/2" + have "isCont (winding_number \) z" + by (metis continuous_at_winding_number valid_path_imp_path \ z) + then obtain d where "d>0" and d: "\x'. dist x' z < d \ dist (winding_number \ x') (winding_number \ z) < \Re(winding_number \ z)\ - 1/2" + using continuous_at_eps_delta wnz_12 diff_gt_0_iff_gt by blast + define z' where "z' = z - (d / (2 * cmod a)) *\<^sub>R a" + have *: "a \ z' \ b - d / 3 * cmod a" + unfolding z'_def inner_mult_right' divide_inverse + apply (simp add: divide_simps algebra_simps dot_square_norm power2_eq_square anz) + apply (metis \0 < d\ add_increasing azb less_eq_real_def mult_nonneg_nonneg mult_right_mono norm_ge_zero norm_numeral) + done + have "cmod (winding_number \ z' - winding_number \ z) < \Re (winding_number \ z)\ - 1/2" + using d [of z'] anz \d>0\ by (simp add: dist_norm z'_def) + then have "1/2 < \Re (winding_number \ z)\ - cmod (winding_number \ z' - winding_number \ z)" + by simp + then have "1/2 < \Re (winding_number \ z)\ - \Re (winding_number \ z') - Re (winding_number \ z)\" + using abs_Re_le_cmod [of "winding_number \ z' - winding_number \ z"] by simp + then have wnz_12': "\Re (winding_number \ z')\ > 1/2" + by linarith + moreover have "\Re (winding_number \ z')\ < 1/2" + apply (rule winding_number_lt_half [OF \ *]) + using azb \d>0\ pag + apply (auto simp: add_strict_increasing anz divide_simps algebra_simps dest!: subsetD) + done + ultimately have False + by simp + } + then show ?thesis by force +qed + +lemma winding_number_lt_half_linepath: "z \ closed_segment a b \ \Re (winding_number (linepath a b) z)\ < 1/2" + using separating_hyperplane_closed_point [of "closed_segment a b" z] + apply auto + apply (simp add: closed_segment_def) + apply (drule less_imp_le) + apply (frule winding_number_lt_half [OF valid_path_linepath [of a b]]) + apply (auto simp: segment) + done + + +text\ Positivity of WN for a linepath.\ +lemma winding_number_linepath_pos_lt: + assumes "0 < Im ((b - a) * cnj (b - z))" + shows "0 < Re(winding_number(linepath a b) z)" +proof - + have z: "z \ path_image (linepath a b)" + using assms + by (simp add: closed_segment_def) (force simp: algebra_simps) + show ?thesis + apply (rule winding_number_pos_lt [OF valid_path_linepath z assms]) + apply (simp add: linepath_def algebra_simps) + done +qed + + +subsection\Cauchy's integral formula, again for a convex enclosing set.\ + +lemma Cauchy_integral_formula_weak: + assumes s: "convex s" and "finite k" and conf: "continuous_on s f" + and fcd: "(\x. x \ interior s - k \ f field_differentiable at x)" + and z: "z \ interior s - k" and vpg: "valid_path \" + and pasz: "path_image \ \ s - {z}" and loop: "pathfinish \ = pathstart \" + shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" +proof - + obtain f' where f': "(f has_field_derivative f') (at z)" + using fcd [OF z] by (auto simp: field_differentiable_def) + have pas: "path_image \ \ s" and znotin: "z \ path_image \" using pasz by blast+ + have c: "continuous (at x within s) (\w. if w = z then f' else (f w - f z) / (w - z))" if "x \ s" for x + proof (cases "x = z") + case True then show ?thesis + apply (simp add: continuous_within) + apply (rule Lim_transform_away_within [of _ "z+1" _ "\w::complex. (f w - f z)/(w - z)"]) + using has_field_derivative_at_within DERIV_within_iff f' + apply (fastforce simp add:)+ + done + next + case False + then have dxz: "dist x z > 0" by auto + have cf: "continuous (at x within s) f" + using conf continuous_on_eq_continuous_within that by blast + have "continuous (at x within s) (\w. (f w - f z) / (w - z))" + by (rule cf continuous_intros | simp add: False)+ + then show ?thesis + apply (rule continuous_transform_within [OF _ dxz that, of "\w::complex. (f w - f z)/(w - z)"]) + apply (force simp: dist_commute) + done + qed + have fink': "finite (insert z k)" using \finite k\ by blast + have *: "((\w. if w = z then f' else (f w - f z) / (w - z)) has_contour_integral 0) \" + apply (rule Cauchy_theorem_convex [OF _ s fink' _ vpg pas loop]) + using c apply (force simp: continuous_on_eq_continuous_within) + apply (rename_tac w) + apply (rule_tac d="dist w z" and f = "\w. (f w - f z)/(w - z)" in field_differentiable_transform_within) + apply (simp_all add: dist_pos_lt dist_commute) + apply (metis less_irrefl) + apply (rule derivative_intros fcd | simp)+ + done + show ?thesis + apply (rule has_contour_integral_eq) + using znotin has_contour_integral_add [OF has_contour_integral_lmul [OF has_contour_integral_winding_number [OF vpg znotin], of "f z"] *] + apply (auto simp: mult_ac divide_simps) + done +qed + +theorem Cauchy_integral_formula_convex_simple: + "\convex s; f holomorphic_on s; z \ interior s; valid_path \; path_image \ \ s - {z}; + pathfinish \ = pathstart \\ + \ ((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" + apply (rule Cauchy_integral_formula_weak [where k = "{}"]) + using holomorphic_on_imp_continuous_on + by auto (metis at_within_interior holomorphic_on_def interiorE subsetCE) + + +subsection\Homotopy forms of Cauchy's theorem\ + +proposition Cauchy_theorem_homotopic: + assumes hom: "if atends then homotopic_paths s g h else homotopic_loops s g h" + and "open s" and f: "f holomorphic_on s" + and vpg: "valid_path g" and vph: "valid_path h" + shows "contour_integral g f = contour_integral h f" +proof - + have pathsf: "linked_paths atends g h" + using hom by (auto simp: linked_paths_def homotopic_paths_imp_pathstart homotopic_paths_imp_pathfinish homotopic_loops_imp_loop) + obtain k :: "real \ real \ complex" + where contk: "continuous_on ({0..1} \ {0..1}) k" + and ks: "k ` ({0..1} \ {0..1}) \ s" + and k [simp]: "\x. k (0, x) = g x" "\x. k (1, x) = h x" + and ksf: "\t\{0..1}. linked_paths atends g (\x. k (t, x))" + using hom pathsf by (auto simp: linked_paths_def homotopic_paths_def homotopic_loops_def homotopic_with_def split: if_split_asm) + have ucontk: "uniformly_continuous_on ({0..1} \ {0..1}) k" + by (blast intro: compact_Times compact_uniformly_continuous [OF contk]) + { fix t::real assume t: "t \ {0..1}" + have pak: "path (k o (\u. (t, u)))" + unfolding path_def + apply (rule continuous_intros continuous_on_subset [OF contk])+ + using t by force + have pik: "path_image (k \ Pair t) \ s" + using ks t by (auto simp: path_image_def) + obtain e where "e>0" and e: + "\g h. \valid_path g; valid_path h; + \u\{0..1}. cmod (g u - (k \ Pair t) u) < e \ cmod (h u - (k \ Pair t) u) < e; + linked_paths atends g h\ + \ contour_integral h f = contour_integral g f" + using contour_integral_nearby [OF \open s\ pak pik, of atends] f by metis + obtain d where "d>0" and d: + "\x x'. \x \ {0..1} \ {0..1}; x' \ {0..1} \ {0..1}; norm (x'-x) < d\ \ norm (k x' - k x) < e/4" + by (rule uniformly_continuous_onE [OF ucontk, of "e/4"]) (auto simp: dist_norm \e>0\) + { fix t1 t2 + assume t1: "0 \ t1" "t1 \ 1" and t2: "0 \ t2" "t2 \ 1" and ltd: "\t1 - t\ < d" "\t2 - t\ < d" + have no2: "\g1 k1 kt. \norm(g1 - k1) < e/4; norm(k1 - kt) < e/4\ \ norm(g1 - kt) < e" + using \e > 0\ + apply (rule_tac y = k1 in norm_triangle_half_l) + apply (auto simp: norm_minus_commute intro: order_less_trans) + done + have "\d>0. \g1 g2. valid_path g1 \ valid_path g2 \ + (\u\{0..1}. cmod (g1 u - k (t1, u)) < d \ cmod (g2 u - k (t2, u)) < d) \ + linked_paths atends g1 g2 \ + contour_integral g2 f = contour_integral g1 f" + apply (rule_tac x="e/4" in exI) + using t t1 t2 ltd \e > 0\ + apply (auto intro!: e simp: d no2 simp del: less_divide_eq_numeral1) + done + } + then have "\e. 0 < e \ + (\t1 t2. t1 \ {0..1} \ t2 \ {0..1} \ \t1 - t\ < e \ \t2 - t\ < e + \ (\d. 0 < d \ + (\g1 g2. valid_path g1 \ valid_path g2 \ + (\u \ {0..1}. + norm(g1 u - k((t1,u))) < d \ norm(g2 u - k((t2,u))) < d) \ + linked_paths atends g1 g2 + \ contour_integral g2 f = contour_integral g1 f)))" + by (rule_tac x=d in exI) (simp add: \d > 0\) + } + then obtain ee where ee: + "\t. t \ {0..1} \ ee t > 0 \ + (\t1 t2. t1 \ {0..1} \ t2 \ {0..1} \ \t1 - t\ < ee t \ \t2 - t\ < ee t + \ (\d. 0 < d \ + (\g1 g2. valid_path g1 \ valid_path g2 \ + (\u \ {0..1}. + norm(g1 u - k((t1,u))) < d \ norm(g2 u - k((t2,u))) < d) \ + linked_paths atends g1 g2 + \ contour_integral g2 f = contour_integral g1 f)))" + by metis + note ee_rule = ee [THEN conjunct2, rule_format] + define C where "C = (\t. ball t (ee t / 3)) ` {0..1}" + have "\t \ C. open t" by (simp add: C_def) + moreover have "{0..1} \ \C" + using ee [THEN conjunct1] by (auto simp: C_def dist_norm) + ultimately obtain C' where C': "C' \ C" "finite C'" and C'01: "{0..1} \ \C'" + by (rule compactE [OF compact_interval]) + define kk where "kk = {t \ {0..1}. ball t (ee t / 3) \ C'}" + have kk01: "kk \ {0..1}" by (auto simp: kk_def) + define e where "e = Min (ee ` kk)" + have C'_eq: "C' = (\t. ball t (ee t / 3)) ` kk" + using C' by (auto simp: kk_def C_def) + have ee_pos[simp]: "\t. t \ {0..1} \ ee t > 0" + by (simp add: kk_def ee) + moreover have "finite kk" + using \finite C'\ kk01 by (force simp: C'_eq inj_on_def ball_eq_ball_iff dest: ee_pos finite_imageD) + moreover have "kk \ {}" using \{0..1} \ \C'\ C'_eq by force + ultimately have "e > 0" + using finite_less_Inf_iff [of "ee ` kk" 0] kk01 by (force simp: e_def) + then obtain N::nat where "N > 0" and N: "1/N < e/3" + by (meson divide_pos_pos nat_approx_posE zero_less_Suc zero_less_numeral) + have e_le_ee: "\i. i \ kk \ e \ ee i" + using \finite kk\ by (simp add: e_def Min_le_iff [of "ee ` kk"]) + have plus: "\t \ kk. x \ ball t (ee t / 3)" if "x \ {0..1}" for x + using C' subsetD [OF C'01 that] unfolding C'_eq by blast + have [OF order_refl]: + "\d. 0 < d \ (\j. valid_path j \ (\u \ {0..1}. norm(j u - k (n/N, u)) < d) \ linked_paths atends g j + \ contour_integral j f = contour_integral g f)" + if "n \ N" for n + using that + proof (induct n) + case 0 show ?case using ee_rule [of 0 0 0] + apply clarsimp + apply (rule_tac x=d in exI, safe) + by (metis diff_self vpg norm_zero) + next + case (Suc n) + then have N01: "n/N \ {0..1}" "(Suc n)/N \ {0..1}" by auto + then obtain t where t: "t \ kk" "n/N \ ball t (ee t / 3)" + using plus [of "n/N"] by blast + then have nN_less: "\n/N - t\ < ee t" + by (simp add: dist_norm del: less_divide_eq_numeral1) + have n'N_less: "\real (Suc n) / real N - t\ < ee t" + using t N \N > 0\ e_le_ee [of t] + by (simp add: dist_norm add_divide_distrib abs_diff_less_iff del: less_divide_eq_numeral1) (simp add: field_simps) + have t01: "t \ {0..1}" using \kk \ {0..1}\ \t \ kk\ by blast + obtain d1 where "d1 > 0" and d1: + "\g1 g2. \valid_path g1; valid_path g2; + \u\{0..1}. cmod (g1 u - k (n/N, u)) < d1 \ cmod (g2 u - k ((Suc n) / N, u)) < d1; + linked_paths atends g1 g2\ + \ contour_integral g2 f = contour_integral g1 f" + using ee [THEN conjunct2, rule_format, OF t01 N01 nN_less n'N_less] by fastforce + have "n \ N" using Suc.prems by auto + with Suc.hyps + obtain d2 where "d2 > 0" + and d2: "\j. \valid_path j; \u\{0..1}. cmod (j u - k (n/N, u)) < d2; linked_paths atends g j\ + \ contour_integral j f = contour_integral g f" + by auto + have "continuous_on {0..1} (k o (\u. (n/N, u)))" + apply (rule continuous_intros continuous_on_subset [OF contk])+ + using N01 by auto + then have pkn: "path (\u. k (n/N, u))" + by (simp add: path_def) + have min12: "min d1 d2 > 0" by (simp add: \0 < d1\ \0 < d2\) + obtain p where "polynomial_function p" + and psf: "pathstart p = pathstart (\u. k (n/N, u))" + "pathfinish p = pathfinish (\u. k (n/N, u))" + and pk_le: "\t. t\{0..1} \ cmod (p t - k (n/N, t)) < min d1 d2" + using path_approx_polynomial_function [OF pkn min12] by blast + then have vpp: "valid_path p" using valid_path_polynomial_function by blast + have lpa: "linked_paths atends g p" + by (metis (mono_tags, lifting) N01(1) ksf linked_paths_def pathfinish_def pathstart_def psf) + show ?case + apply (rule_tac x="min d1 d2" in exI) + apply (simp add: \0 < d1\ \0 < d2\, clarify) + apply (rule_tac s="contour_integral p f" in trans) + using pk_le N01(1) ksf pathfinish_def pathstart_def + apply (force intro!: vpp d1 simp add: linked_paths_def psf ksf) + using pk_le N01 apply (force intro!: vpp d2 lpa simp add: linked_paths_def psf ksf) + done + qed + then obtain d where "0 < d" + "\j. valid_path j \ (\u \ {0..1}. norm(j u - k (1,u)) < d) \ + linked_paths atends g j + \ contour_integral j f = contour_integral g f" + using \N>0\ by auto + then have "linked_paths atends g h \ contour_integral h f = contour_integral g f" + using \N>0\ vph by fastforce + then show ?thesis + by (simp add: pathsf) +qed + +proposition Cauchy_theorem_homotopic_paths: + assumes hom: "homotopic_paths s g h" + and "open s" and f: "f holomorphic_on s" + and vpg: "valid_path g" and vph: "valid_path h" + shows "contour_integral g f = contour_integral h f" + using Cauchy_theorem_homotopic [of True s g h] assms by simp + +proposition Cauchy_theorem_homotopic_loops: + assumes hom: "homotopic_loops s g h" + and "open s" and f: "f holomorphic_on s" + and vpg: "valid_path g" and vph: "valid_path h" + shows "contour_integral g f = contour_integral h f" + using Cauchy_theorem_homotopic [of False s g h] assms by simp + +lemma has_contour_integral_newpath: + "\(f has_contour_integral y) h; f contour_integrable_on g; contour_integral g f = contour_integral h f\ + \ (f has_contour_integral y) g" + using has_contour_integral_integral contour_integral_unique by auto + +lemma Cauchy_theorem_null_homotopic: + "\f holomorphic_on s; open s; valid_path g; homotopic_loops s g (linepath a a)\ \ (f has_contour_integral 0) g" + apply (rule has_contour_integral_newpath [where h = "linepath a a"], simp) + using contour_integrable_holomorphic_simple + apply (blast dest: holomorphic_on_imp_continuous_on homotopic_loops_imp_subset) + by (simp add: Cauchy_theorem_homotopic_loops) + + + +subsection\More winding number properties\ + +text\including the fact that it's +-1 inside a simple closed curve.\ + +lemma winding_number_homotopic_paths: + assumes "homotopic_paths (-{z}) g h" + shows "winding_number g z = winding_number h z" +proof - + have "path g" "path h" using homotopic_paths_imp_path [OF assms] by auto + moreover have pag: "z \ path_image g" and pah: "z \ path_image h" + using homotopic_paths_imp_subset [OF assms] by auto + ultimately obtain d e where "d > 0" "e > 0" + and d: "\p. \path p; pathstart p = pathstart g; pathfinish p = pathfinish g; \t\{0..1}. norm (p t - g t) < d\ + \ homotopic_paths (-{z}) g p" + and e: "\q. \path q; pathstart q = pathstart h; pathfinish q = pathfinish h; \t\{0..1}. norm (q t - h t) < e\ + \ homotopic_paths (-{z}) h q" + using homotopic_nearby_paths [of g "-{z}"] homotopic_nearby_paths [of h "-{z}"] by force + obtain p where p: + "valid_path p" "z \ path_image p" + "pathstart p = pathstart g" "pathfinish p = pathfinish g" + and gp_less:"\t\{0..1}. cmod (g t - p t) < d" + and pap: "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number g z" + using winding_number [OF \path g\ pag \0 < d\] by blast + obtain q where q: + "valid_path q" "z \ path_image q" + "pathstart q = pathstart h" "pathfinish q = pathfinish h" + and hq_less: "\t\{0..1}. cmod (h t - q t) < e" + and paq: "contour_integral q (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number h z" + using winding_number [OF \path h\ pah \0 < e\] by blast + have gp: "homotopic_paths (- {z}) g p" + by (simp add: d p valid_path_imp_path norm_minus_commute gp_less) + have hq: "homotopic_paths (- {z}) h q" + by (simp add: e q valid_path_imp_path norm_minus_commute hq_less) + have "contour_integral p (\w. 1/(w - z)) = contour_integral q (\w. 1/(w - z))" + apply (rule Cauchy_theorem_homotopic_paths [of "-{z}"]) + apply (blast intro: homotopic_paths_trans homotopic_paths_sym gp hq assms) + apply (auto intro!: holomorphic_intros simp: p q) + done + then show ?thesis + by (simp add: pap paq) +qed + +lemma winding_number_homotopic_loops: + assumes "homotopic_loops (-{z}) g h" + shows "winding_number g z = winding_number h z" +proof - + have "path g" "path h" using homotopic_loops_imp_path [OF assms] by auto + moreover have pag: "z \ path_image g" and pah: "z \ path_image h" + using homotopic_loops_imp_subset [OF assms] by auto + moreover have gloop: "pathfinish g = pathstart g" and hloop: "pathfinish h = pathstart h" + using homotopic_loops_imp_loop [OF assms] by auto + ultimately obtain d e where "d > 0" "e > 0" + and d: "\p. \path p; pathfinish p = pathstart p; \t\{0..1}. norm (p t - g t) < d\ + \ homotopic_loops (-{z}) g p" + and e: "\q. \path q; pathfinish q = pathstart q; \t\{0..1}. norm (q t - h t) < e\ + \ homotopic_loops (-{z}) h q" + using homotopic_nearby_loops [of g "-{z}"] homotopic_nearby_loops [of h "-{z}"] by force + obtain p where p: + "valid_path p" "z \ path_image p" + "pathstart p = pathstart g" "pathfinish p = pathfinish g" + and gp_less:"\t\{0..1}. cmod (g t - p t) < d" + and pap: "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number g z" + using winding_number [OF \path g\ pag \0 < d\] by blast + obtain q where q: + "valid_path q" "z \ path_image q" + "pathstart q = pathstart h" "pathfinish q = pathfinish h" + and hq_less: "\t\{0..1}. cmod (h t - q t) < e" + and paq: "contour_integral q (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number h z" + using winding_number [OF \path h\ pah \0 < e\] by blast + have gp: "homotopic_loops (- {z}) g p" + by (simp add: gloop d gp_less norm_minus_commute p valid_path_imp_path) + have hq: "homotopic_loops (- {z}) h q" + by (simp add: e hloop hq_less norm_minus_commute q valid_path_imp_path) + have "contour_integral p (\w. 1/(w - z)) = contour_integral q (\w. 1/(w - z))" + apply (rule Cauchy_theorem_homotopic_loops [of "-{z}"]) + apply (blast intro: homotopic_loops_trans homotopic_loops_sym gp hq assms) + apply (auto intro!: holomorphic_intros simp: p q) + done + then show ?thesis + by (simp add: pap paq) +qed + +lemma winding_number_paths_linear_eq: + "\path g; path h; pathstart h = pathstart g; pathfinish h = pathfinish g; + \t. t \ {0..1} \ z \ closed_segment (g t) (h t)\ + \ winding_number h z = winding_number g z" + by (blast intro: sym homotopic_paths_linear winding_number_homotopic_paths elim: ) + +lemma winding_number_loops_linear_eq: + "\path g; path h; pathfinish g = pathstart g; pathfinish h = pathstart h; + \t. t \ {0..1} \ z \ closed_segment (g t) (h t)\ + \ winding_number h z = winding_number g z" + by (blast intro: sym homotopic_loops_linear winding_number_homotopic_loops elim: ) + +lemma winding_number_nearby_paths_eq: + "\path g; path h; + pathstart h = pathstart g; pathfinish h = pathfinish g; + \t. t \ {0..1} \ norm(h t - g t) < norm(g t - z)\ + \ winding_number h z = winding_number g z" + by (metis segment_bound(2) norm_minus_commute not_le winding_number_paths_linear_eq) + +lemma winding_number_nearby_loops_eq: + "\path g; path h; + pathfinish g = pathstart g; + pathfinish h = pathstart h; + \t. t \ {0..1} \ norm(h t - g t) < norm(g t - z)\ + \ winding_number h z = winding_number g z" + by (metis segment_bound(2) norm_minus_commute not_le winding_number_loops_linear_eq) + + +proposition winding_number_subpath_combine: + "\path g; z \ path_image g; + u \ {0..1}; v \ {0..1}; w \ {0..1}\ + \ winding_number (subpath u v g) z + winding_number (subpath v w g) z = + winding_number (subpath u w g) z" +apply (rule trans [OF winding_number_join [THEN sym] + winding_number_homotopic_paths [OF homotopic_join_subpaths]]) +apply (auto dest: path_image_subpath_subset) +done + + +subsection\Partial circle path\ + +definition part_circlepath :: "[complex, real, real, real, real] \ complex" + where "part_circlepath z r s t \ \x. z + of_real r * exp (\ * of_real (linepath s t x))" + +lemma pathstart_part_circlepath [simp]: + "pathstart(part_circlepath z r s t) = z + r*exp(\ * s)" +by (metis part_circlepath_def pathstart_def pathstart_linepath) + +lemma pathfinish_part_circlepath [simp]: + "pathfinish(part_circlepath z r s t) = z + r*exp(\*t)" +by (metis part_circlepath_def pathfinish_def pathfinish_linepath) + +proposition has_vector_derivative_part_circlepath [derivative_intros]: + "((part_circlepath z r s t) has_vector_derivative + (\ * r * (of_real t - of_real s) * exp(\ * linepath s t x))) + (at x within X)" + apply (simp add: part_circlepath_def linepath_def scaleR_conv_of_real) + apply (rule has_vector_derivative_real_complex) + apply (rule derivative_eq_intros | simp)+ + done + +corollary vector_derivative_part_circlepath: + "vector_derivative (part_circlepath z r s t) (at x) = + \ * r * (of_real t - of_real s) * exp(\ * linepath s t x)" + using has_vector_derivative_part_circlepath vector_derivative_at by blast + +corollary vector_derivative_part_circlepath01: + "\0 \ x; x \ 1\ + \ vector_derivative (part_circlepath z r s t) (at x within {0..1}) = + \ * r * (of_real t - of_real s) * exp(\ * linepath s t x)" + using has_vector_derivative_part_circlepath + by (auto simp: vector_derivative_at_within_ivl) + +lemma valid_path_part_circlepath [simp]: "valid_path (part_circlepath z r s t)" + apply (simp add: valid_path_def) + apply (rule C1_differentiable_imp_piecewise) + apply (auto simp: C1_differentiable_on_eq vector_derivative_works vector_derivative_part_circlepath has_vector_derivative_part_circlepath + intro!: continuous_intros) + done + +lemma path_part_circlepath [simp]: "path (part_circlepath z r s t)" + by (simp add: valid_path_imp_path) + +proposition path_image_part_circlepath: + assumes "s \ t" + shows "path_image (part_circlepath z r s t) = {z + r * exp(\ * of_real x) | x. s \ x \ x \ t}" +proof - + { fix z::real + assume "0 \ z" "z \ 1" + with \s \ t\ have "\x. (exp (\ * linepath s t z) = exp (\ * of_real x)) \ s \ x \ x \ t" + apply (rule_tac x="(1 - z) * s + z * t" in exI) + apply (simp add: linepath_def scaleR_conv_of_real algebra_simps) + apply (rule conjI) + using mult_right_mono apply blast + using affine_ineq by (metis "mult.commute") + } + moreover + { fix z + assume "s \ z" "z \ t" + then have "z + of_real r * exp (\ * of_real z) \ (\x. z + of_real r * exp (\ * linepath s t x)) ` {0..1}" + apply (rule_tac x="(z - s)/(t - s)" in image_eqI) + apply (simp add: linepath_def scaleR_conv_of_real divide_simps exp_eq) + apply (auto simp: algebra_simps divide_simps) + done + } + ultimately show ?thesis + by (fastforce simp add: path_image_def part_circlepath_def) +qed + +corollary path_image_part_circlepath_subset: + "\s \ t; 0 \ r\ \ path_image(part_circlepath z r s t) \ sphere z r" +by (auto simp: path_image_part_circlepath sphere_def dist_norm algebra_simps norm_mult) + +proposition in_path_image_part_circlepath: + assumes "w \ path_image(part_circlepath z r s t)" "s \ t" "0 \ r" + shows "norm(w - z) = r" +proof - + have "w \ {c. dist z c = r}" + by (metis (no_types) path_image_part_circlepath_subset sphere_def subset_eq assms) + thus ?thesis + by (simp add: dist_norm norm_minus_commute) +qed + +proposition finite_bounded_log: "finite {z::complex. norm z \ b \ exp z = w}" +proof (cases "w = 0") + case True then show ?thesis by auto +next + case False + have *: "finite {x. cmod (complex_of_real (2 * real_of_int x * pi) * \) \ b + cmod (Ln w)}" + apply (simp add: norm_mult finite_int_iff_bounded_le) + apply (rule_tac x="\(b + cmod (Ln w)) / (2*pi)\" in exI) + apply (auto simp: divide_simps le_floor_iff) + done + have [simp]: "\P f. {z. P z \ (\n. z = f n)} = f ` {n. P (f n)}" + by blast + show ?thesis + apply (subst exp_Ln [OF False, symmetric]) + apply (simp add: exp_eq) + using norm_add_leD apply (fastforce intro: finite_subset [OF _ *]) + done +qed + +lemma finite_bounded_log2: + fixes a::complex + assumes "a \ 0" + shows "finite {z. norm z \ b \ exp(a*z) = w}" +proof - + have *: "finite ((\z. z / a) ` {z. cmod z \ b * cmod a \ exp z = w})" + by (rule finite_imageI [OF finite_bounded_log]) + show ?thesis + by (rule finite_subset [OF _ *]) (force simp: assms norm_mult) +qed + +proposition has_contour_integral_bound_part_circlepath_strong: + assumes fi: "(f has_contour_integral i) (part_circlepath z r s t)" + and "finite k" and le: "0 \ B" "0 < r" "s \ t" + and B: "\x. x \ path_image(part_circlepath z r s t) - k \ norm(f x) \ B" + shows "cmod i \ B * r * (t - s)" +proof - + consider "s = t" | "s < t" using \s \ t\ by linarith + then show ?thesis + proof cases + case 1 with fi [unfolded has_contour_integral] + have "i = 0" by (simp add: vector_derivative_part_circlepath) + with assms show ?thesis by simp + next + case 2 + have [simp]: "\r\ = r" using \r > 0\ by linarith + have [simp]: "cmod (complex_of_real t - complex_of_real s) = t-s" + by (metis "2" abs_of_pos diff_gt_0_iff_gt norm_of_real of_real_diff) + have "finite (part_circlepath z r s t -` {y} \ {0..1})" if "y \ k" for y + proof - + define w where "w = (y - z)/of_real r / exp(\ * of_real s)" + have fin: "finite (of_real -` {z. cmod z \ 1 \ exp (\ * complex_of_real (t - s) * z) = w})" + apply (rule finite_vimageI [OF finite_bounded_log2]) + using \s < t\ apply (auto simp: inj_of_real) + done + show ?thesis + apply (simp add: part_circlepath_def linepath_def vimage_def) + apply (rule finite_subset [OF _ fin]) + using le + apply (auto simp: w_def algebra_simps scaleR_conv_of_real exp_add exp_diff) + done + qed + then have fin01: "finite ((part_circlepath z r s t) -` k \ {0..1})" + by (rule finite_finite_vimage_IntI [OF \finite k\]) + have **: "((\x. if (part_circlepath z r s t x) \ k then 0 + else f(part_circlepath z r s t x) * + vector_derivative (part_circlepath z r s t) (at x)) has_integral i) {0..1}" + apply (rule has_integral_spike + [where f = "\x. f(part_circlepath z r s t x) * vector_derivative (part_circlepath z r s t) (at x)"]) + apply (rule negligible_finite [OF fin01]) + using fi has_contour_integral + apply auto + done + have *: "\x. \0 \ x; x \ 1; part_circlepath z r s t x \ k\ \ cmod (f (part_circlepath z r s t x)) \ B" + by (auto intro!: B [unfolded path_image_def image_def, simplified]) + show ?thesis + apply (rule has_integral_bound [where 'a=real, simplified, OF _ **, simplified]) + using assms apply force + apply (simp add: norm_mult vector_derivative_part_circlepath) + using le * "2" \r > 0\ by auto + qed +qed + +corollary has_contour_integral_bound_part_circlepath: + "\(f has_contour_integral i) (part_circlepath z r s t); + 0 \ B; 0 < r; s \ t; + \x. x \ path_image(part_circlepath z r s t) \ norm(f x) \ B\ + \ norm i \ B*r*(t - s)" + by (auto intro: has_contour_integral_bound_part_circlepath_strong) + +proposition contour_integrable_continuous_part_circlepath: + "continuous_on (path_image (part_circlepath z r s t)) f + \ f contour_integrable_on (part_circlepath z r s t)" + apply (simp add: contour_integrable_on has_contour_integral_def vector_derivative_part_circlepath path_image_def) + apply (rule integrable_continuous_real) + apply (fast intro: path_part_circlepath [unfolded path_def] continuous_intros continuous_on_compose2 [where g=f, OF _ _ order_refl]) + done + +proposition winding_number_part_circlepath_pos_less: + assumes "s < t" and no: "norm(w - z) < r" + shows "0 < Re (winding_number(part_circlepath z r s t) w)" +proof - + have "0 < r" by (meson no norm_not_less_zero not_le order.strict_trans2) + note valid_path_part_circlepath + moreover have " w \ path_image (part_circlepath z r s t)" + using assms by (auto simp: path_image_def image_def part_circlepath_def norm_mult linepath_def) + moreover have "0 < r * (t - s) * (r - cmod (w - z))" + using assms by (metis \0 < r\ diff_gt_0_iff_gt mult_pos_pos) + ultimately show ?thesis + apply (rule winding_number_pos_lt [where e = "r*(t - s)*(r - norm(w - z))"]) + apply (simp add: vector_derivative_part_circlepath right_diff_distrib [symmetric] mult_ac) + apply (rule mult_left_mono)+ + using Re_Im_le_cmod [of "w-z" "linepath s t x" for x] + apply (simp add: exp_Euler cos_of_real sin_of_real part_circlepath_def algebra_simps cos_squared_eq [unfolded power2_eq_square]) + using assms \0 < r\ by auto +qed + +proposition simple_path_part_circlepath: + "simple_path(part_circlepath z r s t) \ (r \ 0 \ s \ t \ \s - t\ \ 2*pi)" +proof (cases "r = 0 \ s = t") + case True + then show ?thesis + apply (rule disjE) + apply (force simp: part_circlepath_def simple_path_def intro: bexI [where x = "1/4"] bexI [where x = "1/3"])+ + done +next + case False then have "r \ 0" "s \ t" by auto + have *: "\x y z s t. \*((1 - x) * s + x * t) = \*(((1 - y) * s + y * t)) + z \ \*(x - y) * (t - s) = z" + by (simp add: algebra_simps) + have abs01: "\x y::real. 0 \ x \ x \ 1 \ 0 \ y \ y \ 1 + \ (x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0 \ \x - y\ \ {0,1})" + by auto + have abs_away: "\P. (\x\{0..1}. \y\{0..1}. P \x - y\) \ (\x::real. 0 \ x \ x \ 1 \ P x)" + by force + have **: "\x y. (\n. (complex_of_real x - of_real y) * (of_real t - of_real s) = 2 * (of_int n * of_real pi)) \ + (\n. \x - y\ * (t - s) = 2 * (of_int n * pi))" + by (force simp: algebra_simps abs_if dest: arg_cong [where f=Re] arg_cong [where f=complex_of_real] + intro: exI [where x = "-n" for n]) + have 1: "\x. 0 \ x \ x \ 1 \ (\n. x * (t - s) = 2 * (real_of_int n * pi)) \ x = 0 \ x = 1 \ \s - t\ \ 2 * pi" + apply (rule ccontr) + apply (drule_tac x="2*pi / \t - s\" in spec) + using False + apply (simp add: abs_minus_commute divide_simps) + apply (frule_tac x=1 in spec) + apply (drule_tac x="-1" in spec) + apply (simp add:) + done + have 2: "\s - t\ = \2 * (real_of_int n * pi) / x\" if "x \ 0" "x * (t - s) = 2 * (real_of_int n * pi)" for x n + proof - + have "t-s = 2 * (real_of_int n * pi)/x" + using that by (simp add: field_simps) + then show ?thesis by (metis abs_minus_commute) + qed + show ?thesis using False + apply (simp add: simple_path_def path_part_circlepath) + apply (simp add: part_circlepath_def linepath_def exp_eq * ** abs01 del: Set.insert_iff) + apply (subst abs_away) + apply (auto simp: 1) + apply (rule ccontr) + apply (auto simp: 2 divide_simps abs_mult dest: of_int_leD) + done +qed + +proposition arc_part_circlepath: + assumes "r \ 0" "s \ t" "\s - t\ < 2*pi" + shows "arc (part_circlepath z r s t)" +proof - + have *: "x = y" if eq: "\ * (linepath s t x) = \ * (linepath s t y) + 2 * of_int n * complex_of_real pi * \" + and x: "x \ {0..1}" and y: "y \ {0..1}" for x y n + proof - + have "(linepath s t x) = (linepath s t y) + 2 * of_int n * complex_of_real pi" + by (metis add_divide_eq_iff complex_i_not_zero mult.commute nonzero_mult_divide_cancel_left eq) + then have "s*y + t*x = s*x + (t*y + of_int n * (pi * 2))" + by (force simp: algebra_simps linepath_def dest: arg_cong [where f=Re]) + then have st: "x \ y \ (s-t) = (of_int n * (pi * 2) / (y-x))" + by (force simp: field_simps) + show ?thesis + apply (rule ccontr) + using assms x y + apply (simp add: st abs_mult field_simps) + using st + apply (auto simp: dest: of_int_lessD) + done + qed + show ?thesis + using assms + apply (simp add: arc_def) + apply (simp add: part_circlepath_def inj_on_def exp_eq) + apply (blast intro: *) + done +qed + + +subsection\Special case of one complete circle\ + +definition circlepath :: "[complex, real, real] \ complex" + where "circlepath z r \ part_circlepath z r 0 (2*pi)" + +lemma circlepath: "circlepath z r = (\x. z + r * exp(2 * of_real pi * \ * of_real x))" + by (simp add: circlepath_def part_circlepath_def linepath_def algebra_simps) + +lemma pathstart_circlepath [simp]: "pathstart (circlepath z r) = z + r" + by (simp add: circlepath_def) + +lemma pathfinish_circlepath [simp]: "pathfinish (circlepath z r) = z + r" + by (simp add: circlepath_def) (metis exp_two_pi_i mult.commute) + +lemma circlepath_minus: "circlepath z (-r) x = circlepath z r (x + 1/2)" +proof - + have "z + of_real r * exp (2 * pi * \ * (x + 1 / 2)) = + z + of_real r * exp (2 * pi * \ * x + pi * \)" + by (simp add: divide_simps) (simp add: algebra_simps) + also have "... = z - r * exp (2 * pi * \ * x)" + by (simp add: exp_add) + finally show ?thesis + by (simp add: circlepath path_image_def sphere_def dist_norm) +qed + +lemma circlepath_add1: "circlepath z r (x+1) = circlepath z r x" + using circlepath_minus [of z r "x+1/2"] circlepath_minus [of z "-r" x] + by (simp add: add.commute) + +lemma circlepath_add_half: "circlepath z r (x + 1/2) = circlepath z r (x - 1/2)" + using circlepath_add1 [of z r "x-1/2"] + by (simp add: add.commute) + +lemma path_image_circlepath_minus_subset: + "path_image (circlepath z (-r)) \ path_image (circlepath z r)" + apply (simp add: path_image_def image_def circlepath_minus, clarify) + apply (case_tac "xa \ 1/2", force) + apply (force simp add: circlepath_add_half)+ + done + +lemma path_image_circlepath_minus: "path_image (circlepath z (-r)) = path_image (circlepath z r)" + using path_image_circlepath_minus_subset by fastforce + +proposition has_vector_derivative_circlepath [derivative_intros]: + "((circlepath z r) has_vector_derivative (2 * pi * \ * r * exp (2 * of_real pi * \ * of_real x))) + (at x within X)" + apply (simp add: circlepath_def scaleR_conv_of_real) + apply (rule derivative_eq_intros) + apply (simp add: algebra_simps) + done + +corollary vector_derivative_circlepath: + "vector_derivative (circlepath z r) (at x) = + 2 * pi * \ * r * exp(2 * of_real pi * \ * x)" +using has_vector_derivative_circlepath vector_derivative_at by blast + +corollary vector_derivative_circlepath01: + "\0 \ x; x \ 1\ + \ vector_derivative (circlepath z r) (at x within {0..1}) = + 2 * pi * \ * r * exp(2 * of_real pi * \ * x)" + using has_vector_derivative_circlepath + by (auto simp: vector_derivative_at_within_ivl) + +lemma valid_path_circlepath [simp]: "valid_path (circlepath z r)" + by (simp add: circlepath_def) + +lemma path_circlepath [simp]: "path (circlepath z r)" + by (simp add: valid_path_imp_path) + +lemma path_image_circlepath_nonneg: + assumes "0 \ r" shows "path_image (circlepath z r) = sphere z r" +proof - + have *: "x \ (\u. z + (cmod (x - z)) * exp (\ * (of_real u * (of_real pi * 2)))) ` {0..1}" for x + proof (cases "x = z") + case True then show ?thesis by force + next + case False + define w where "w = x - z" + then have "w \ 0" by (simp add: False) + have **: "\t. \Re w = cos t * cmod w; Im w = sin t * cmod w\ \ w = of_real (cmod w) * exp (\ * t)" + using cis_conv_exp complex_eq_iff by auto + show ?thesis + apply (rule sincos_total_2pi [of "Re(w/of_real(norm w))" "Im(w/of_real(norm w))"]) + apply (simp add: divide_simps \w \ 0\ cmod_power2 [symmetric]) + apply (rule_tac x="t / (2*pi)" in image_eqI) + apply (simp add: divide_simps \w \ 0\) + using False ** + apply (auto simp: w_def) + done + qed + show ?thesis + unfolding circlepath path_image_def sphere_def dist_norm + by (force simp: assms algebra_simps norm_mult norm_minus_commute intro: *) +qed + +proposition path_image_circlepath [simp]: + "path_image (circlepath z r) = sphere z \r\" + using path_image_circlepath_minus + by (force simp add: path_image_circlepath_nonneg abs_if) + +lemma has_contour_integral_bound_circlepath_strong: + "\(f has_contour_integral i) (circlepath z r); + finite k; 0 \ B; 0 < r; + \x. \norm(x - z) = r; x \ k\ \ norm(f x) \ B\ + \ norm i \ B*(2*pi*r)" + unfolding circlepath_def + by (auto simp: algebra_simps in_path_image_part_circlepath dest!: has_contour_integral_bound_part_circlepath_strong) + +corollary has_contour_integral_bound_circlepath: + "\(f has_contour_integral i) (circlepath z r); + 0 \ B; 0 < r; \x. norm(x - z) = r \ norm(f x) \ B\ + \ norm i \ B*(2*pi*r)" + by (auto intro: has_contour_integral_bound_circlepath_strong) + +proposition contour_integrable_continuous_circlepath: + "continuous_on (path_image (circlepath z r)) f + \ f contour_integrable_on (circlepath z r)" + by (simp add: circlepath_def contour_integrable_continuous_part_circlepath) + +lemma simple_path_circlepath: "simple_path(circlepath z r) \ (r \ 0)" + by (simp add: circlepath_def simple_path_part_circlepath) + +lemma notin_path_image_circlepath [simp]: "cmod (w - z) < r \ w \ path_image (circlepath z r)" + by (simp add: sphere_def dist_norm norm_minus_commute) + +proposition contour_integral_circlepath: + "0 < r \ contour_integral (circlepath z r) (\w. 1 / (w - z)) = 2 * complex_of_real pi * \" + apply (rule contour_integral_unique) + apply (simp add: has_contour_integral_def) + apply (subst has_integral_cong) + apply (simp add: vector_derivative_circlepath01) + using has_integral_const_real [of _ 0 1] + apply (force simp: circlepath) + done + +lemma winding_number_circlepath_centre: "0 < r \ winding_number (circlepath z r) z = 1" + apply (rule winding_number_unique_loop) + apply (simp_all add: sphere_def valid_path_imp_path) + apply (rule_tac x="circlepath z r" in exI) + apply (simp add: sphere_def contour_integral_circlepath) + done + +proposition winding_number_circlepath: + assumes "norm(w - z) < r" shows "winding_number(circlepath z r) w = 1" +proof (cases "w = z") + case True then show ?thesis + using assms winding_number_circlepath_centre by auto +next + case False + have [simp]: "r > 0" + using assms le_less_trans norm_ge_zero by blast + define r' where "r' = norm(w - z)" + have "r' < r" + by (simp add: assms r'_def) + have disjo: "cball z r' \ sphere z r = {}" + using \r' < r\ by (force simp: cball_def sphere_def) + have "winding_number(circlepath z r) w = winding_number(circlepath z r) z" + apply (rule winding_number_around_inside [where s = "cball z r'"]) + apply (simp_all add: disjo order.strict_implies_order winding_number_circlepath_centre) + apply (simp_all add: False r'_def dist_norm norm_minus_commute) + done + also have "... = 1" + by (simp add: winding_number_circlepath_centre) + finally show ?thesis . +qed + + +text\ Hence the Cauchy formula for points inside a circle.\ + +theorem Cauchy_integral_circlepath: + assumes "continuous_on (cball z r) f" "f holomorphic_on (ball z r)" "norm(w - z) < r" + shows "((\u. f u/(u - w)) has_contour_integral (2 * of_real pi * \ * f w)) + (circlepath z r)" +proof - + have "r > 0" + using assms le_less_trans norm_ge_zero by blast + have "((\u. f u / (u - w)) has_contour_integral (2 * pi) * \ * winding_number (circlepath z r) w * f w) + (circlepath z r)" + apply (rule Cauchy_integral_formula_weak [where s = "cball z r" and k = "{}"]) + using assms \r > 0\ + apply (simp_all add: dist_norm norm_minus_commute) + apply (metis at_within_interior dist_norm holomorphic_on_def interior_ball mem_ball norm_minus_commute) + apply (simp add: cball_def sphere_def dist_norm, clarify) + apply (simp add:) + by (metis dist_commute dist_norm less_irrefl) + then show ?thesis + by (simp add: winding_number_circlepath assms) +qed + +corollary Cauchy_integral_circlepath_simple: + assumes "f holomorphic_on cball z r" "norm(w - z) < r" + shows "((\u. f u/(u - w)) has_contour_integral (2 * of_real pi * \ * f w)) + (circlepath z r)" +using assms by (force simp: holomorphic_on_imp_continuous_on holomorphic_on_subset Cauchy_integral_circlepath) + + +lemma no_bounded_connected_component_imp_winding_number_zero: + assumes g: "path g" "path_image g \ s" "pathfinish g = pathstart g" "z \ s" + and nb: "\z. bounded (connected_component_set (- s) z) \ z \ s" + shows "winding_number g z = 0" +apply (rule winding_number_zero_in_outside) +apply (simp_all add: assms) +by (metis nb [of z] \path_image g \ s\ \z \ s\ contra_subsetD mem_Collect_eq outside outside_mono) + +lemma no_bounded_path_component_imp_winding_number_zero: + assumes g: "path g" "path_image g \ s" "pathfinish g = pathstart g" "z \ s" + and nb: "\z. bounded (path_component_set (- s) z) \ z \ s" + shows "winding_number g z = 0" +apply (rule no_bounded_connected_component_imp_winding_number_zero [OF g]) +by (simp add: bounded_subset nb path_component_subset_connected_component) + + +subsection\ Uniform convergence of path integral\ + +text\Uniform convergence when the derivative of the path is bounded, and in particular for the special case of a circle.\ + +proposition contour_integral_uniform_limit: + assumes ev_fint: "eventually (\n::'a. (f n) contour_integrable_on \) F" + and ev_no: "\e. 0 < e \ eventually (\n. \x \ path_image \. norm(f n x - l x) < e) F" + and noleB: "\t. t \ {0..1} \ norm (vector_derivative \ (at t)) \ B" + and \: "valid_path \" + and [simp]: "~ (trivial_limit F)" + shows "l contour_integrable_on \" "((\n. contour_integral \ (f n)) \ contour_integral \ l) F" +proof - + have "0 \ B" by (meson noleB [of 0] atLeastAtMost_iff norm_ge_zero order_refl order_trans zero_le_one) + { fix e::real + assume "0 < e" + then have eB: "0 < e / (\B\ + 1)" by simp + obtain a where fga: "\x. x \ {0..1} \ cmod (f a (\ x) - l (\ x)) < e / (\B\ + 1)" + and inta: "(\t. f a (\ t) * vector_derivative \ (at t)) integrable_on {0..1}" + using eventually_happens [OF eventually_conj [OF ev_no [OF eB] ev_fint]] + by (fastforce simp: contour_integrable_on path_image_def) + have Ble: "B * e / (\B\ + 1) \ e" + using \0 \ B\ \0 < e\ by (simp add: divide_simps) + have "\h. (\x\{0..1}. cmod (l (\ x) * vector_derivative \ (at x) - h x) \ e) \ h integrable_on {0..1}" + apply (rule_tac x="\x. f (a::'a) (\ x) * vector_derivative \ (at x)" in exI) + apply (intro inta conjI ballI) + apply (rule order_trans [OF _ Ble]) + apply (frule noleB) + apply (frule fga) + using \0 \ B\ \0 < e\ + apply (simp add: norm_mult left_diff_distrib [symmetric] norm_minus_commute divide_simps) + apply (drule (1) mult_mono [OF less_imp_le]) + apply (simp_all add: mult_ac) + done + } + then show lintg: "l contour_integrable_on \" + apply (simp add: contour_integrable_on) + apply (blast intro: integrable_uniform_limit_real) + done + { fix e::real + define B' where "B' = B + 1" + have B': "B' > 0" "B' > B" using \0 \ B\ by (auto simp: B'_def) + assume "0 < e" + then have ev_no': "\\<^sub>F n in F. \x\path_image \. 2 * cmod (f n x - l x) < e / B'" + using ev_no [of "e / B' / 2"] B' by (simp add: field_simps) + have ie: "integral {0..1::real} (\x. e / 2) < e" using \0 < e\ by simp + have *: "cmod (f x (\ t) * vector_derivative \ (at t) - l (\ t) * vector_derivative \ (at t)) \ e / 2" + if t: "t\{0..1}" and leB': "2 * cmod (f x (\ t) - l (\ t)) < e / B'" for x t + proof - + have "2 * cmod (f x (\ t) - l (\ t)) * cmod (vector_derivative \ (at t)) \ e * (B/ B')" + using mult_mono [OF less_imp_le [OF leB'] noleB] B' \0 < e\ t by auto + also have "... < e" + by (simp add: B' \0 < e\ mult_imp_div_pos_less) + finally have "2 * cmod (f x (\ t) - l (\ t)) * cmod (vector_derivative \ (at t)) < e" . + then show ?thesis + by (simp add: left_diff_distrib [symmetric] norm_mult) + qed + have "\\<^sub>F x in F. dist (contour_integral \ (f x)) (contour_integral \ l) < e" + apply (rule eventually_mono [OF eventually_conj [OF ev_no' ev_fint]]) + apply (simp add: dist_norm contour_integrable_on path_image_def contour_integral_integral) + apply (simp add: lintg integral_diff [symmetric] contour_integrable_on [symmetric], clarify) + apply (rule le_less_trans [OF integral_norm_bound_integral ie]) + apply (simp add: lintg integrable_diff contour_integrable_on [symmetric]) + apply (blast intro: *)+ + done + } + then show "((\n. contour_integral \ (f n)) \ contour_integral \ l) F" + by (rule tendstoI) +qed + +proposition contour_integral_uniform_limit_circlepath: + assumes ev_fint: "eventually (\n::'a. (f n) contour_integrable_on (circlepath z r)) F" + and ev_no: "\e. 0 < e \ eventually (\n. \x \ path_image (circlepath z r). norm(f n x - l x) < e) F" + and [simp]: "~ (trivial_limit F)" "0 < r" + shows "l contour_integrable_on (circlepath z r)" "((\n. contour_integral (circlepath z r) (f n)) \ contour_integral (circlepath z r) l) F" +by (auto simp: vector_derivative_circlepath norm_mult intro: contour_integral_uniform_limit assms) + + +subsection\ General stepping result for derivative formulas.\ + +lemma sum_sqs_eq: + fixes x::"'a::idom" shows "x * x + y * y = x * (y * 2) \ y = x" + by algebra + +proposition Cauchy_next_derivative: + assumes "continuous_on (path_image \) f'" + and leB: "\t. t \ {0..1} \ norm (vector_derivative \ (at t)) \ B" + and int: "\w. w \ s - path_image \ \ ((\u. f' u / (u - w)^k) has_contour_integral f w) \" + and k: "k \ 0" + and "open s" + and \: "valid_path \" + and w: "w \ s - path_image \" + shows "(\u. f' u / (u - w)^(Suc k)) contour_integrable_on \" + and "(f has_field_derivative (k * contour_integral \ (\u. f' u/(u - w)^(Suc k)))) + (at w)" (is "?thes2") +proof - + have "open (s - path_image \)" using \open s\ closed_valid_path_image \ by blast + then obtain d where "d>0" and d: "ball w d \ s - path_image \" using w + using open_contains_ball by blast + have [simp]: "\n. cmod (1 + of_nat n) = 1 + of_nat n" + by (metis norm_of_nat of_nat_Suc) + have 1: "\\<^sub>F n in at w. (\x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k) + contour_integrable_on \" + apply (simp add: eventually_at) + apply (rule_tac x=d in exI) + apply (simp add: \d > 0\ dist_norm field_simps, clarify) + apply (rule contour_integrable_div [OF contour_integrable_diff]) + using int w d + apply (force simp: dist_norm norm_minus_commute intro!: has_contour_integral_integrable)+ + done + have bim_g: "bounded (image f' (path_image \))" + by (simp add: compact_imp_bounded compact_continuous_image compact_valid_path_image assms) + then obtain C where "C > 0" and C: "\x. \0 \ x; x \ 1\ \ cmod (f' (\ x)) \ C" + by (force simp: bounded_pos path_image_def) + have twom: "\\<^sub>F n in at w. + \x\path_image \. + cmod ((inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / k - inverse (x - w) ^ Suc k) < e" + if "0 < e" for e + proof - + have *: "cmod ((inverse (x - u) ^ k - inverse (x - w) ^ k) / ((u - w) * k) - inverse (x - w) ^ Suc k) < e" + if x: "x \ path_image \" and "u \ w" and uwd: "cmod (u - w) < d/2" + and uw_less: "cmod (u - w) < e * (d / 2) ^ (k+2) / (1 + real k)" + for u x + proof - + define ff where [abs_def]: + "ff n w = + (if n = 0 then inverse(x - w)^k + else if n = 1 then k / (x - w)^(Suc k) + else (k * of_real(Suc k)) / (x - w)^(k + 2))" for n :: nat and w + have km1: "\z::complex. z \ 0 \ z ^ (k - Suc 0) = z ^ k / z" + by (simp add: field_simps) (metis Suc_pred \k \ 0\ neq0_conv power_Suc) + have ff1: "(ff i has_field_derivative ff (Suc i) z) (at z within ball w (d / 2))" + if "z \ ball w (d / 2)" "i \ 1" for i z + proof - + have "z \ path_image \" + using \x \ path_image \\ d that ball_divide_subset_numeral by blast + then have xz[simp]: "x \ z" using \x \ path_image \\ by blast + then have neq: "x * x + z * z \ x * (z * 2)" + by (blast intro: dest!: sum_sqs_eq) + with xz have "\v. v \ 0 \ (x * x + z * z) * v \ (x * (z * 2) * v)" by auto + then have neqq: "\v. v \ 0 \ x * (x * v) + z * (z * v) \ x * (z * (2 * v))" + by (simp add: algebra_simps) + show ?thesis using \i \ 1\ + apply (simp add: ff_def dist_norm Nat.le_Suc_eq km1, safe) + apply (rule derivative_eq_intros | simp add: km1 | simp add: field_simps neq neqq)+ + done + qed + { fix a::real and b::real assume ab: "a > 0" "b > 0" + then have "k * (1 + real k) * (1 / a) \ k * (1 + real k) * (4 / b) \ b \ 4 * a" + apply (subst mult_le_cancel_left_pos) + using \k \ 0\ + apply (auto simp: divide_simps) + done + with ab have "real k * (1 + real k) / a \ (real k * 4 + real k * real k * 4) / b \ b \ 4 * a" + by (simp add: field_simps) + } note canc = this + have ff2: "cmod (ff (Suc 1) v) \ real (k * (k + 1)) / (d / 2) ^ (k + 2)" + if "v \ ball w (d / 2)" for v + proof - + have "d/2 \ cmod (x - v)" using d x that + apply (simp add: dist_norm path_image_def ball_def not_less [symmetric] del: divide_const_simps, clarify) + apply (drule subsetD) + prefer 2 apply blast + apply (metis norm_minus_commute norm_triangle_half_r CollectI) + done + then have "d \ cmod (x - v) * 2" + by (simp add: divide_simps) + then have dpow_le: "d ^ (k+2) \ (cmod (x - v) * 2) ^ (k+2)" + using \0 < d\ order_less_imp_le power_mono by blast + have "x \ v" using that + using \x \ path_image \\ ball_divide_subset_numeral d by fastforce + then show ?thesis + using \d > 0\ + apply (simp add: ff_def norm_mult norm_divide norm_power dist_norm canc) + using dpow_le + apply (simp add: algebra_simps divide_simps mult_less_0_iff) + done + qed + have ub: "u \ ball w (d / 2)" + using uwd by (simp add: dist_commute dist_norm) + have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) + \ (real k * 4 + real k * real k * 4) * (cmod (u - w) * cmod (u - w)) / (d * (d * (d / 2) ^ k))" + using complex_taylor [OF _ ff1 ff2 _ ub, of w, simplified] + by (simp add: ff_def \0 < d\) + then have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) + \ (cmod (u - w) * real k) * (1 + real k) * cmod (u - w) / (d / 2) ^ (k+2)" + by (simp add: field_simps) + then have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) + / (cmod (u - w) * real k) + \ (1 + real k) * cmod (u - w) / (d / 2) ^ (k+2)" + using \k \ 0\ \u \ w\ by (simp add: mult_ac zero_less_mult_iff pos_divide_le_eq) + also have "... < e" + using uw_less \0 < d\ by (simp add: mult_ac divide_simps) + finally have e: "cmod (inverse (x-u)^k - (inverse (x-w)^k + of_nat k * (u-w) / ((x-w) * (x-w)^k))) + / cmod ((u - w) * real k) < e" + by (simp add: norm_mult) + have "x \ u" + using uwd \0 < d\ x d by (force simp: dist_norm ball_def norm_minus_commute) + show ?thesis + apply (rule le_less_trans [OF _ e]) + using \k \ 0\ \x \ u\ \u \ w\ + apply (simp add: field_simps norm_divide [symmetric]) + done + qed + show ?thesis + unfolding eventually_at + apply (rule_tac x = "min (d/2) ((e*(d/2)^(k + 2))/(Suc k))" in exI) + apply (force simp: \d > 0\ dist_norm that simp del: power_Suc intro: *) + done + qed + have 2: "\\<^sub>F n in at w. + \x\path_image \. + cmod (f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k - f' x / (x - w) ^ Suc k) < e" + if "0 < e" for e + proof - + have *: "cmod (f' (\ x) * (inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - + f' (\ x) / ((\ x - w) * (\ x - w) ^ k)) < e" + if ec: "cmod ((inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - + inverse (\ x - w) * inverse (\ x - w) ^ k) < e / C" + and x: "0 \ x" "x \ 1" + for u x + proof (cases "(f' (\ x)) = 0") + case True then show ?thesis by (simp add: \0 < e\) + next + case False + have "cmod (f' (\ x) * (inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - + f' (\ x) / ((\ x - w) * (\ x - w) ^ k)) = + cmod (f' (\ x) * ((inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - + inverse (\ x - w) * inverse (\ x - w) ^ k))" + by (simp add: field_simps) + also have "... = cmod (f' (\ x)) * + cmod ((inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - + inverse (\ x - w) * inverse (\ x - w) ^ k)" + by (simp add: norm_mult) + also have "... < cmod (f' (\ x)) * (e/C)" + apply (rule mult_strict_left_mono [OF ec]) + using False by simp + also have "... \ e" using C + by (metis False \0 < e\ frac_le less_eq_real_def mult.commute pos_le_divide_eq x zero_less_norm_iff) + finally show ?thesis . + qed + show ?thesis + using twom [OF divide_pos_pos [OF that \C > 0\]] unfolding path_image_def + by (force intro: * elim: eventually_mono) + qed + show "(\u. f' u / (u - w) ^ (Suc k)) contour_integrable_on \" + by (rule contour_integral_uniform_limit [OF 1 2 leB \]) auto + have *: "(\n. contour_integral \ (\x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / k)) + \w\ contour_integral \ (\u. f' u / (u - w) ^ (Suc k))" + by (rule contour_integral_uniform_limit [OF 1 2 leB \]) auto + have **: "contour_integral \ (\x. f' x * (inverse (x - u) ^ k - inverse (x - w) ^ k) / ((u - w) * k)) = + (f u - f w) / (u - w) / k" + if "dist u w < d" for u + apply (rule contour_integral_unique) + apply (simp add: diff_divide_distrib algebra_simps) + apply (rule has_contour_integral_diff; rule has_contour_integral_div; simp add: field_simps; rule int) + apply (metis contra_subsetD d dist_commute mem_ball that) + apply (rule w) + done + show ?thes2 + apply (simp add: DERIV_within_iff del: power_Suc) + apply (rule Lim_transform_within [OF tendsto_mult_left [OF *] \0 < d\ ]) + apply (simp add: \k \ 0\ **) + done +qed + +corollary Cauchy_next_derivative_circlepath: + assumes contf: "continuous_on (path_image (circlepath z r)) f" + and int: "\w. w \ ball z r \ ((\u. f u / (u - w)^k) has_contour_integral g w) (circlepath z r)" + and k: "k \ 0" + and w: "w \ ball z r" + shows "(\u. f u / (u - w)^(Suc k)) contour_integrable_on (circlepath z r)" + (is "?thes1") + and "(g has_field_derivative (k * contour_integral (circlepath z r) (\u. f u/(u - w)^(Suc k)))) (at w)" + (is "?thes2") +proof - + have "r > 0" using w + using ball_eq_empty by fastforce + have wim: "w \ ball z r - path_image (circlepath z r)" + using w by (auto simp: dist_norm) + show ?thes1 ?thes2 + by (rule Cauchy_next_derivative [OF contf _ int k open_ball valid_path_circlepath wim, where B = "2 * pi * \r\"]; + auto simp: vector_derivative_circlepath norm_mult)+ +qed + + +text\ In particular, the first derivative formula.\ + +proposition Cauchy_derivative_integral_circlepath: + assumes contf: "continuous_on (cball z r) f" + and holf: "f holomorphic_on ball z r" + and w: "w \ ball z r" + shows "(\u. f u/(u - w)^2) contour_integrable_on (circlepath z r)" + (is "?thes1") + and "(f has_field_derivative (1 / (2 * of_real pi * \) * contour_integral(circlepath z r) (\u. f u / (u - w)^2))) (at w)" + (is "?thes2") +proof - + have [simp]: "r \ 0" using w + using ball_eq_empty by fastforce + have f: "continuous_on (path_image (circlepath z r)) f" + by (rule continuous_on_subset [OF contf]) (force simp add: cball_def sphere_def) + have int: "\w. dist z w < r \ + ((\u. f u / (u - w)) has_contour_integral (\x. 2 * of_real pi * \ * f x) w) (circlepath z r)" + by (rule Cauchy_integral_circlepath [OF contf holf]) (simp add: dist_norm norm_minus_commute) + show ?thes1 + apply (simp add: power2_eq_square) + apply (rule Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1, simplified]) + apply (blast intro: int) + done + have "((\x. 2 * of_real pi * \ * f x) has_field_derivative contour_integral (circlepath z r) (\u. f u / (u - w)^2)) (at w)" + apply (simp add: power2_eq_square) + apply (rule Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1 and g = "\x. 2 * of_real pi * \ * f x", simplified]) + apply (blast intro: int) + done + then have fder: "(f has_field_derivative contour_integral (circlepath z r) (\u. f u / (u - w)^2) / (2 * of_real pi * \)) (at w)" + by (rule DERIV_cdivide [where f = "\x. 2 * of_real pi * \ * f x" and c = "2 * of_real pi * \", simplified]) + show ?thes2 + by simp (rule fder) +qed + +subsection\ Existence of all higher derivatives.\ + +proposition derivative_is_holomorphic: + assumes "open s" + and fder: "\z. z \ s \ (f has_field_derivative f' z) (at z)" + shows "f' holomorphic_on s" +proof - + have *: "\h. (f' has_field_derivative h) (at z)" if "z \ s" for z + proof - + obtain r where "r > 0" and r: "cball z r \ s" + using open_contains_cball \z \ s\ \open s\ by blast + then have holf_cball: "f holomorphic_on cball z r" + apply (simp add: holomorphic_on_def) + using field_differentiable_at_within field_differentiable_def fder by blast + then have "continuous_on (path_image (circlepath z r)) f" + using \r > 0\ by (force elim: holomorphic_on_subset [THEN holomorphic_on_imp_continuous_on]) + then have contfpi: "continuous_on (path_image (circlepath z r)) (\x. 1/(2 * of_real pi*\) * f x)" + by (auto intro: continuous_intros)+ + have contf_cball: "continuous_on (cball z r) f" using holf_cball + by (simp add: holomorphic_on_imp_continuous_on holomorphic_on_subset) + have holf_ball: "f holomorphic_on ball z r" using holf_cball + using ball_subset_cball holomorphic_on_subset by blast + { fix w assume w: "w \ ball z r" + have intf: "(\u. f u / (u - w)\<^sup>2) contour_integrable_on circlepath z r" + by (blast intro: w Cauchy_derivative_integral_circlepath [OF contf_cball holf_ball]) + have fder': "(f has_field_derivative 1 / (2 * of_real pi * \) * contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2)) + (at w)" + by (blast intro: w Cauchy_derivative_integral_circlepath [OF contf_cball holf_ball]) + have f'_eq: "f' w = contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \)" + using fder' ball_subset_cball r w by (force intro: DERIV_unique [OF fder]) + have "((\u. f u / (u - w)\<^sup>2 / (2 * of_real pi * \)) has_contour_integral + contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \)) + (circlepath z r)" + by (rule has_contour_integral_div [OF has_contour_integral_integral [OF intf]]) + then have "((\u. f u / (2 * of_real pi * \ * (u - w)\<^sup>2)) has_contour_integral + contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \)) + (circlepath z r)" + by (simp add: algebra_simps) + then have "((\u. f u / (2 * of_real pi * \ * (u - w)\<^sup>2)) has_contour_integral f' w) (circlepath z r)" + by (simp add: f'_eq) + } note * = this + show ?thesis + apply (rule exI) + apply (rule Cauchy_next_derivative_circlepath [OF contfpi, of 2 f', simplified]) + apply (simp_all add: \0 < r\ * dist_norm) + done + qed + show ?thesis + by (simp add: holomorphic_on_open [OF \open s\] *) +qed + +lemma holomorphic_deriv [holomorphic_intros]: + "\f holomorphic_on s; open s\ \ (deriv f) holomorphic_on s" +by (metis DERIV_deriv_iff_field_differentiable at_within_open derivative_is_holomorphic holomorphic_on_def) + +lemma analytic_deriv: "f analytic_on s \ (deriv f) analytic_on s" + using analytic_on_holomorphic holomorphic_deriv by auto + +lemma holomorphic_higher_deriv [holomorphic_intros]: "\f holomorphic_on s; open s\ \ (deriv ^^ n) f holomorphic_on s" + by (induction n) (auto simp: holomorphic_deriv) + +lemma analytic_higher_deriv: "f analytic_on s \ (deriv ^^ n) f analytic_on s" + unfolding analytic_on_def using holomorphic_higher_deriv by blast + +lemma has_field_derivative_higher_deriv: + "\f holomorphic_on s; open s; x \ s\ + \ ((deriv ^^ n) f has_field_derivative (deriv ^^ (Suc n)) f x) (at x)" +by (metis (no_types, hide_lams) DERIV_deriv_iff_field_differentiable at_within_open comp_apply + funpow.simps(2) holomorphic_higher_deriv holomorphic_on_def) + +lemma valid_path_compose_holomorphic: + assumes "valid_path g" and holo:"f holomorphic_on s" and "open s" "path_image g \ s" + shows "valid_path (f o g)" +proof (rule valid_path_compose[OF \valid_path g\]) + fix x assume "x \ path_image g" + then show "\f'. (f has_field_derivative f') (at x)" + using holo holomorphic_on_open[OF \open s\] \path_image g \ s\ by auto +next + have "deriv f holomorphic_on s" + using holomorphic_deriv holo \open s\ by auto + then show "continuous_on (path_image g) (deriv f)" + using assms(4) holomorphic_on_imp_continuous_on holomorphic_on_subset by auto +qed + + +subsection\ Morera's theorem.\ + +lemma Morera_local_triangle_ball: + assumes "\z. z \ s + \ \e a. 0 < e \ z \ ball a e \ continuous_on (ball a e) f \ + (\b c. closed_segment b c \ ball a e + \ contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0)" + shows "f analytic_on s" +proof - + { fix z assume "z \ s" + with assms obtain e a where + "0 < e" and z: "z \ ball a e" and contf: "continuous_on (ball a e) f" + and 0: "\b c. closed_segment b c \ ball a e + \ contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0" + by fastforce + have az: "dist a z < e" using mem_ball z by blast + have sb_ball: "ball z (e - dist a z) \ ball a e" + by (simp add: dist_commute ball_subset_ball_iff) + have "\e>0. f holomorphic_on ball z e" + apply (rule_tac x="e - dist a z" in exI) + apply (simp add: az) + apply (rule holomorphic_on_subset [OF _ sb_ball]) + apply (rule derivative_is_holomorphic[OF open_ball]) + apply (rule triangle_contour_integrals_starlike_primitive [OF contf _ open_ball, of a]) + apply (simp_all add: 0 \0 < e\) + apply (meson \0 < e\ centre_in_ball convex_ball convex_contains_segment mem_ball) + done + } + then show ?thesis + by (simp add: analytic_on_def) +qed + +lemma Morera_local_triangle: + assumes "\z. z \ s + \ \t. open t \ z \ t \ continuous_on t f \ + (\a b c. convex hull {a,b,c} \ t + \ contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0)" + shows "f analytic_on s" +proof - + { fix z assume "z \ s" + with assms obtain t where + "open t" and z: "z \ t" and contf: "continuous_on t f" + and 0: "\a b c. convex hull {a,b,c} \ t + \ contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0" + by force + then obtain e where "e>0" and e: "ball z e \ t" + using open_contains_ball by blast + have [simp]: "continuous_on (ball z e) f" using contf + using continuous_on_subset e by blast + have "\e a. 0 < e \ + z \ ball a e \ + continuous_on (ball a e) f \ + (\b c. closed_segment b c \ ball a e \ + contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0)" + apply (rule_tac x=e in exI) + apply (rule_tac x=z in exI) + apply (simp add: \e > 0\, clarify) + apply (rule 0) + apply (meson z \0 < e\ centre_in_ball closed_segment_subset convex_ball dual_order.trans e starlike_convex_subset) + done + } + then show ?thesis + by (simp add: Morera_local_triangle_ball) +qed + +proposition Morera_triangle: + "\continuous_on s f; open s; + \a b c. convex hull {a,b,c} \ s + \ contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0\ + \ f analytic_on s" + using Morera_local_triangle by blast + + + +subsection\ Combining theorems for higher derivatives including Leibniz rule.\ + +lemma higher_deriv_linear [simp]: + "(deriv ^^ n) (\w. c*w) = (\z. if n = 0 then c*z else if n = 1 then c else 0)" + by (induction n) (auto simp: deriv_const deriv_linear) + +lemma higher_deriv_const [simp]: "(deriv ^^ n) (\w. c) = (\w. if n=0 then c else 0)" + by (induction n) (auto simp: deriv_const) + +lemma higher_deriv_ident [simp]: + "(deriv ^^ n) (\w. w) z = (if n = 0 then z else if n = 1 then 1 else 0)" + apply (induction n, simp) + apply (metis higher_deriv_linear lambda_one) + done + +corollary higher_deriv_id [simp]: + "(deriv ^^ n) id z = (if n = 0 then z else if n = 1 then 1 else 0)" + by (simp add: id_def) + +lemma has_complex_derivative_funpow_1: + "\(f has_field_derivative 1) (at z); f z = z\ \ (f^^n has_field_derivative 1) (at z)" + apply (induction n) + apply auto + apply (metis DERIV_ident DERIV_transform_at id_apply zero_less_one) + by (metis DERIV_chain comp_funpow comp_id funpow_swap1 mult.right_neutral) + +proposition higher_deriv_uminus: + assumes "f holomorphic_on s" "open s" and z: "z \ s" + shows "(deriv ^^ n) (\w. -(f w)) z = - ((deriv ^^ n) f z)" +using z +proof (induction n arbitrary: z) + case 0 then show ?case by simp +next + case (Suc n z) + have *: "((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)" + using Suc.prems assms has_field_derivative_higher_deriv by auto + show ?case + apply simp + apply (rule DERIV_imp_deriv) + apply (rule DERIV_transform_within_open [of "\w. -((deriv ^^ n) f w)"]) + apply (rule derivative_eq_intros | rule * refl assms Suc)+ + apply (simp add: Suc) + done +qed + +proposition higher_deriv_add: + fixes z::complex + assumes "f holomorphic_on s" "g holomorphic_on s" "open s" and z: "z \ s" + shows "(deriv ^^ n) (\w. f w + g w) z = (deriv ^^ n) f z + (deriv ^^ n) g z" +using z +proof (induction n arbitrary: z) + case 0 then show ?case by simp +next + case (Suc n z) + have *: "((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)" + "((deriv ^^ n) g has_field_derivative deriv ((deriv ^^ n) g) z) (at z)" + using Suc.prems assms has_field_derivative_higher_deriv by auto + show ?case + apply simp + apply (rule DERIV_imp_deriv) + apply (rule DERIV_transform_within_open [of "\w. (deriv ^^ n) f w + (deriv ^^ n) g w"]) + apply (rule derivative_eq_intros | rule * refl assms Suc)+ + apply (simp add: Suc) + done +qed + +corollary higher_deriv_diff: + fixes z::complex + assumes "f holomorphic_on s" "g holomorphic_on s" "open s" and z: "z \ s" + shows "(deriv ^^ n) (\w. f w - g w) z = (deriv ^^ n) f z - (deriv ^^ n) g z" + apply (simp only: Groups.group_add_class.diff_conv_add_uminus higher_deriv_add) + apply (subst higher_deriv_add) + using assms holomorphic_on_minus apply (auto simp: higher_deriv_uminus) + done + + +lemma bb: "Suc n choose k = (n choose k) + (if k = 0 then 0 else (n choose (k - 1)))" + by (cases k) simp_all + +proposition higher_deriv_mult: + fixes z::complex + assumes "f holomorphic_on s" "g holomorphic_on s" "open s" and z: "z \ s" + shows "(deriv ^^ n) (\w. f w * g w) z = + (\i = 0..n. of_nat (n choose i) * (deriv ^^ i) f z * (deriv ^^ (n - i)) g z)" +using z +proof (induction n arbitrary: z) + case 0 then show ?case by simp +next + case (Suc n z) + have *: "\n. ((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)" + "\n. ((deriv ^^ n) g has_field_derivative deriv ((deriv ^^ n) g) z) (at z)" + using Suc.prems assms has_field_derivative_higher_deriv by auto + have sumeq: "(\i = 0..n. + of_nat (n choose i) * (deriv ((deriv ^^ i) f) z * (deriv ^^ (n - i)) g z + deriv ((deriv ^^ (n - i)) g) z * (deriv ^^ i) f z)) = + g z * deriv ((deriv ^^ n) f) z + (\i = 0..n. (deriv ^^ i) f z * (of_nat (Suc n choose i) * (deriv ^^ (Suc n - i)) g z))" + apply (simp add: bb algebra_simps setsum.distrib) + apply (subst (4) setsum_Suc_reindex) + apply (auto simp: algebra_simps Suc_diff_le intro: setsum.cong) + done + show ?case + apply (simp only: funpow.simps o_apply) + apply (rule DERIV_imp_deriv) + apply (rule DERIV_transform_within_open + [of "\w. (\i = 0..n. of_nat (n choose i) * (deriv ^^ i) f w * (deriv ^^ (n - i)) g w)"]) + apply (simp add: algebra_simps) + apply (rule DERIV_cong [OF DERIV_setsum]) + apply (rule DERIV_cmult) + apply (auto simp: intro: DERIV_mult * sumeq \open s\ Suc.prems Suc.IH [symmetric]) + done +qed + + +proposition higher_deriv_transform_within_open: + fixes z::complex + assumes "f holomorphic_on s" "g holomorphic_on s" "open s" and z: "z \ s" + and fg: "\w. w \ s \ f w = g w" + shows "(deriv ^^ i) f z = (deriv ^^ i) g z" +using z +by (induction i arbitrary: z) + (auto simp: fg intro: complex_derivative_transform_within_open holomorphic_higher_deriv assms) + +proposition higher_deriv_compose_linear: + fixes z::complex + assumes f: "f holomorphic_on t" and s: "open s" and t: "open t" and z: "z \ s" + and fg: "\w. w \ s \ u * w \ t" + shows "(deriv ^^ n) (\w. f (u * w)) z = u^n * (deriv ^^ n) f (u * z)" +using z +proof (induction n arbitrary: z) + case 0 then show ?case by simp +next + case (Suc n z) + have holo0: "f holomorphic_on op * u ` s" + by (meson fg f holomorphic_on_subset image_subset_iff) + have holo1: "(\w. f (u * w)) holomorphic_on s" + apply (rule holomorphic_on_compose [where g=f, unfolded o_def]) + apply (rule holo0 holomorphic_intros)+ + done + have holo2: "(\z. u ^ n * (deriv ^^ n) f (u * z)) holomorphic_on s" + apply (rule holomorphic_intros)+ + apply (rule holomorphic_on_compose [where g="(deriv ^^ n) f", unfolded o_def]) + apply (rule holomorphic_intros) + apply (rule holomorphic_on_subset [where s=t]) + apply (rule holomorphic_intros assms)+ + apply (blast intro: fg) + done + have "deriv ((deriv ^^ n) (\w. f (u * w))) z = deriv (\z. u^n * (deriv ^^ n) f (u*z)) z" + apply (rule complex_derivative_transform_within_open [OF _ holo2 s Suc.prems]) + apply (rule holomorphic_higher_deriv [OF holo1 s]) + apply (simp add: Suc.IH) + done + also have "... = u^n * deriv (\z. (deriv ^^ n) f (u * z)) z" + apply (rule deriv_cmult) + apply (rule analytic_on_imp_differentiable_at [OF _ Suc.prems]) + apply (rule analytic_on_compose_gen [where g="(deriv ^^ n) f" and t=t, unfolded o_def]) + apply (simp add: analytic_on_linear) + apply (simp add: analytic_on_open f holomorphic_higher_deriv t) + apply (blast intro: fg) + done + also have "... = u * u ^ n * deriv ((deriv ^^ n) f) (u * z)" + apply (subst complex_derivative_chain [where g = "(deriv ^^ n) f" and f = "op*u", unfolded o_def]) + apply (rule derivative_intros) + using Suc.prems field_differentiable_def f fg has_field_derivative_higher_deriv t apply blast + apply (simp add: deriv_linear) + done + finally show ?case + by simp +qed + +lemma higher_deriv_add_at: + assumes "f analytic_on {z}" "g analytic_on {z}" + shows "(deriv ^^ n) (\w. f w + g w) z = (deriv ^^ n) f z + (deriv ^^ n) g z" +proof - + have "f analytic_on {z} \ g analytic_on {z}" + using assms by blast + with higher_deriv_add show ?thesis + by (auto simp: analytic_at_two) +qed + +lemma higher_deriv_diff_at: + assumes "f analytic_on {z}" "g analytic_on {z}" + shows "(deriv ^^ n) (\w. f w - g w) z = (deriv ^^ n) f z - (deriv ^^ n) g z" +proof - + have "f analytic_on {z} \ g analytic_on {z}" + using assms by blast + with higher_deriv_diff show ?thesis + by (auto simp: analytic_at_two) +qed + +lemma higher_deriv_uminus_at: + "f analytic_on {z} \ (deriv ^^ n) (\w. -(f w)) z = - ((deriv ^^ n) f z)" + using higher_deriv_uminus + by (auto simp: analytic_at) + +lemma higher_deriv_mult_at: + assumes "f analytic_on {z}" "g analytic_on {z}" + shows "(deriv ^^ n) (\w. f w * g w) z = + (\i = 0..n. of_nat (n choose i) * (deriv ^^ i) f z * (deriv ^^ (n - i)) g z)" +proof - + have "f analytic_on {z} \ g analytic_on {z}" + using assms by blast + with higher_deriv_mult show ?thesis + by (auto simp: analytic_at_two) +qed + + +text\ Nonexistence of isolated singularities and a stronger integral formula.\ + +proposition no_isolated_singularity: + fixes z::complex + assumes f: "continuous_on s f" and holf: "f holomorphic_on (s - k)" and s: "open s" and k: "finite k" + shows "f holomorphic_on s" +proof - + { fix z + assume "z \ s" and cdf: "\x. x\s - k \ f field_differentiable at x" + have "f field_differentiable at z" + proof (cases "z \ k") + case False then show ?thesis by (blast intro: cdf \z \ s\) + next + case True + with finite_set_avoid [OF k, of z] + obtain d where "d>0" and d: "\x. \x\k; x \ z\ \ d \ dist z x" + by blast + obtain e where "e>0" and e: "ball z e \ s" + using s \z \ s\ by (force simp add: open_contains_ball) + have fde: "continuous_on (ball z (min d e)) f" + by (metis Int_iff ball_min_Int continuous_on_subset e f subsetI) + have "\g. \w \ ball z (min d e). (g has_field_derivative f w) (at w within ball z (min d e))" + apply (rule contour_integral_convex_primitive [OF convex_ball fde]) + apply (rule Cauchy_theorem_triangle_cofinite [OF _ k]) + apply (metis continuous_on_subset [OF fde] closed_segment_subset convex_ball starlike_convex_subset) + apply (rule cdf) + apply (metis Diff_iff Int_iff ball_min_Int bot_least contra_subsetD convex_ball e insert_subset + interior_mono interior_subset subset_hull) + done + then have "f holomorphic_on ball z (min d e)" + by (metis open_ball at_within_open derivative_is_holomorphic) + then show ?thesis + unfolding holomorphic_on_def + by (metis open_ball \0 < d\ \0 < e\ at_within_open centre_in_ball min_less_iff_conj) + qed + } + with holf s k show ?thesis + by (simp add: holomorphic_on_open open_Diff finite_imp_closed field_differentiable_def [symmetric]) +qed + +proposition Cauchy_integral_formula_convex: + assumes s: "convex s" and k: "finite k" and contf: "continuous_on s f" + and fcd: "(\x. x \ interior s - k \ f field_differentiable at x)" + and z: "z \ interior s" and vpg: "valid_path \" + and pasz: "path_image \ \ s - {z}" and loop: "pathfinish \ = pathstart \" + shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" + apply (rule Cauchy_integral_formula_weak [OF s finite.emptyI contf]) + apply (simp add: holomorphic_on_open [symmetric] field_differentiable_def) + using no_isolated_singularity [where s = "interior s"] + apply (metis k contf fcd holomorphic_on_open field_differentiable_def continuous_on_subset + has_field_derivative_at_within holomorphic_on_def interior_subset open_interior) + using assms + apply auto + done + + +text\ Formula for higher derivatives.\ + +proposition Cauchy_has_contour_integral_higher_derivative_circlepath: + assumes contf: "continuous_on (cball z r) f" + and holf: "f holomorphic_on ball z r" + and w: "w \ ball z r" + shows "((\u. f u / (u - w) ^ (Suc k)) has_contour_integral ((2 * pi * \) / (fact k) * (deriv ^^ k) f w)) + (circlepath z r)" +using w +proof (induction k arbitrary: w) + case 0 then show ?case + using assms by (auto simp: Cauchy_integral_circlepath dist_commute dist_norm) +next + case (Suc k) + have [simp]: "r > 0" using w + using ball_eq_empty by fastforce + have f: "continuous_on (path_image (circlepath z r)) f" + by (rule continuous_on_subset [OF contf]) (force simp add: cball_def sphere_def less_imp_le) + obtain X where X: "((\u. f u / (u - w) ^ Suc (Suc k)) has_contour_integral X) (circlepath z r)" + using Cauchy_next_derivative_circlepath(1) [OF f Suc.IH _ Suc.prems] + by (auto simp: contour_integrable_on_def) + then have con: "contour_integral (circlepath z r) ((\u. f u / (u - w) ^ Suc (Suc k))) = X" + by (rule contour_integral_unique) + have "\n. ((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) w) (at w)" + using Suc.prems assms has_field_derivative_higher_deriv by auto + then have dnf_diff: "\n. (deriv ^^ n) f field_differentiable (at w)" + by (force simp add: field_differentiable_def) + have "deriv (\w. complex_of_real (2 * pi) * \ / (fact k) * (deriv ^^ k) f w) w = + of_nat (Suc k) * contour_integral (circlepath z r) (\u. f u / (u - w) ^ Suc (Suc k))" + by (force intro!: DERIV_imp_deriv Cauchy_next_derivative_circlepath [OF f Suc.IH _ Suc.prems]) + also have "... = of_nat (Suc k) * X" + by (simp only: con) + finally have "deriv (\w. ((2 * pi) * \ / (fact k)) * (deriv ^^ k) f w) w = of_nat (Suc k) * X" . + then have "((2 * pi) * \ / (fact k)) * deriv (\w. (deriv ^^ k) f w) w = of_nat (Suc k) * X" + by (metis deriv_cmult dnf_diff) + then have "deriv (\w. (deriv ^^ k) f w) w = of_nat (Suc k) * X / ((2 * pi) * \ / (fact k))" + by (simp add: field_simps) + then show ?case + using of_nat_eq_0_iff X by fastforce +qed + +proposition Cauchy_higher_derivative_integral_circlepath: + assumes contf: "continuous_on (cball z r) f" + and holf: "f holomorphic_on ball z r" + and w: "w \ ball z r" + shows "(\u. f u / (u - w)^(Suc k)) contour_integrable_on (circlepath z r)" + (is "?thes1") + and "(deriv ^^ k) f w = (fact k) / (2 * pi * \) * contour_integral(circlepath z r) (\u. f u/(u - w)^(Suc k))" + (is "?thes2") +proof - + have *: "((\u. f u / (u - w) ^ Suc k) has_contour_integral (2 * pi) * \ / (fact k) * (deriv ^^ k) f w) + (circlepath z r)" + using Cauchy_has_contour_integral_higher_derivative_circlepath [OF assms] + by simp + show ?thes1 using * + using contour_integrable_on_def by blast + show ?thes2 + unfolding contour_integral_unique [OF *] by (simp add: divide_simps) +qed + +corollary Cauchy_contour_integral_circlepath: + assumes "continuous_on (cball z r) f" "f holomorphic_on ball z r" "w \ ball z r" + shows "contour_integral(circlepath z r) (\u. f u/(u - w)^(Suc k)) = (2 * pi * \) * (deriv ^^ k) f w / (fact k)" +by (simp add: Cauchy_higher_derivative_integral_circlepath [OF assms]) + +corollary Cauchy_contour_integral_circlepath_2: + assumes "continuous_on (cball z r) f" "f holomorphic_on ball z r" "w \ ball z r" + shows "contour_integral(circlepath z r) (\u. f u/(u - w)^2) = (2 * pi * \) * deriv f w" + using Cauchy_contour_integral_circlepath [OF assms, of 1] + by (simp add: power2_eq_square) + + +subsection\A holomorphic function is analytic, i.e. has local power series.\ + +theorem holomorphic_power_series: + assumes holf: "f holomorphic_on ball z r" + and w: "w \ ball z r" + shows "((\n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)" +proof - + have fh': "f holomorphic_on cball z ((r + dist w z) / 2)" + apply (rule holomorphic_on_subset [OF holf]) + apply (clarsimp simp del: divide_const_simps) + apply (metis add.commute dist_commute le_less_trans mem_ball real_gt_half_sum w) + done + \\Replacing @{term r} and the original (weak) premises\ + obtain r where "0 < r" and holfc: "f holomorphic_on cball z r" and w: "w \ ball z r" + apply (rule that [of "(r + dist w z) / 2"]) + apply (simp_all add: fh') + apply (metis add_0_iff ball_eq_empty dist_nz dist_self empty_iff not_less pos_add_strict w) + apply (metis add_less_cancel_right dist_commute mem_ball mult_2_right w) + done + then have holf: "f holomorphic_on ball z r" and contf: "continuous_on (cball z r) f" + using ball_subset_cball holomorphic_on_subset apply blast + by (simp add: holfc holomorphic_on_imp_continuous_on) + have cint: "\k. (\u. f u / (u - z) ^ Suc k) contour_integrable_on circlepath z r" + apply (rule Cauchy_higher_derivative_integral_circlepath [OF contf holf]) + apply (simp add: \0 < r\) + done + obtain B where "0 < B" and B: "\u. u \ cball z r \ norm(f u) \ B" + by (metis (no_types) bounded_pos compact_cball compact_continuous_image compact_imp_bounded contf image_eqI) + obtain k where k: "0 < k" "k \ r" and wz_eq: "norm(w - z) = r - k" + and kle: "\u. norm(u - z) = r \ k \ norm(u - w)" + apply (rule_tac k = "r - dist z w" in that) + using w + apply (auto simp: dist_norm norm_minus_commute) + by (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute) + have *: "\\<^sub>F n in sequentially. \x\path_image (circlepath z r). + norm ((\k (r - k) / r" "(r - k) / r < 1" using k by auto + obtain n where n: "((r - k) / r) ^ n < e / B * k" + using real_arch_pow_inv [of "e/B*k" "(r - k)/r"] \0 < e\ \0 < B\ k by force + have "norm ((\k N" and r: "r = dist z u" for N u + proof - + have N: "((r - k) / r) ^ N < e / B * k" + apply (rule le_less_trans [OF power_decreasing n]) + using \n \ N\ k by auto + have u [simp]: "(u \ z) \ (u \ w)" + using \0 < r\ r w by auto + have wzu_not1: "(w - z) / (u - z) \ 1" + by (metis (no_types) dist_norm divide_eq_1_iff less_irrefl mem_ball norm_minus_commute r w) + have "norm ((\kk0 < B\ + apply (auto simp: geometric_sum [OF wzu_not1]) + apply (simp add: field_simps norm_mult [symmetric]) + done + also have "... = norm ((u-z) ^ N * (w - u) - ((w - z) ^ N - (u-z) ^ N) * (u-w)) / (r ^ N * norm (u-w)) * norm (f u)" + using \0 < r\ r by (simp add: divide_simps norm_mult norm_divide norm_power dist_norm norm_minus_commute) + also have "... = norm ((w - z) ^ N * (w - u)) / (r ^ N * norm (u - w)) * norm (f u)" + by (simp add: algebra_simps) + also have "... = norm (w - z) ^ N * norm (f u) / r ^ N" + by (simp add: norm_mult norm_power norm_minus_commute) + also have "... \ (((r - k)/r)^N) * B" + using \0 < r\ w k + apply (simp add: divide_simps) + apply (rule mult_mono [OF power_mono]) + apply (auto simp: norm_divide wz_eq norm_power dist_norm norm_minus_commute B r) + done + also have "... < e * k" + using \0 < B\ N by (simp add: divide_simps) + also have "... \ e * norm (u - w)" + using r kle \0 < e\ by (simp add: dist_commute dist_norm) + finally show ?thesis + by (simp add: divide_simps norm_divide del: power_Suc) + qed + with \0 < r\ show ?thesis + by (auto simp: mult_ac less_imp_le eventually_sequentially Ball_def) + qed + have eq: "\\<^sub>F x in sequentially. + contour_integral (circlepath z r) (\u. \kku. f u / (u - z) ^ Suc k) * (w - z) ^ k)" + apply (rule eventuallyI) + apply (subst contour_integral_setsum, simp) + using contour_integrable_lmul [OF cint, of "(w - z) ^ a" for a] apply (simp add: field_simps) + apply (simp only: contour_integral_lmul cint algebra_simps) + done + have "(\k. contour_integral (circlepath z r) (\u. f u/(u - z)^(Suc k)) * (w - z)^k) + sums contour_integral (circlepath z r) (\u. f u/(u - w))" + unfolding sums_def + apply (rule Lim_transform_eventually [OF eq]) + apply (rule contour_integral_uniform_limit_circlepath [OF eventuallyI *]) + apply (rule contour_integrable_setsum, simp) + apply (rule contour_integrable_lmul) + apply (rule Cauchy_higher_derivative_integral_circlepath [OF contf holf]) + using \0 < r\ + apply auto + done + then have "(\k. contour_integral (circlepath z r) (\u. f u/(u - z)^(Suc k)) * (w - z)^k) + sums (2 * of_real pi * \ * f w)" + using w by (auto simp: dist_commute dist_norm contour_integral_unique [OF Cauchy_integral_circlepath_simple [OF holfc]]) + then have "(\k. contour_integral (circlepath z r) (\u. f u / (u - z) ^ Suc k) * (w - z)^k / (\ * (of_real pi * 2))) + sums ((2 * of_real pi * \ * f w) / (\ * (complex_of_real pi * 2)))" + by (rule sums_divide) + then have "(\n. (w - z) ^ n * contour_integral (circlepath z r) (\u. f u / (u - z) ^ Suc n) / (\ * (of_real pi * 2))) + sums f w" + by (simp add: field_simps) + then show ?thesis + by (simp add: field_simps \0 < r\ Cauchy_higher_derivative_integral_circlepath [OF contf holf]) +qed + + +subsection\The Liouville theorem and the Fundamental Theorem of Algebra.\ + +text\ These weak Liouville versions don't even need the derivative formula.\ + +lemma Liouville_weak_0: + assumes holf: "f holomorphic_on UNIV" and inf: "(f \ 0) at_infinity" + shows "f z = 0" +proof (rule ccontr) + assume fz: "f z \ 0" + with inf [unfolded Lim_at_infinity, rule_format, of "norm(f z)/2"] + obtain B where B: "\x. B \ cmod x \ norm (f x) * 2 < cmod (f z)" + by (auto simp: dist_norm) + define R where "R = 1 + \B\ + norm z" + have "R > 0" unfolding R_def + proof - + have "0 \ cmod z + \B\" + by (metis (full_types) add_nonneg_nonneg norm_ge_zero real_norm_def) + then show "0 < 1 + \B\ + cmod z" + by linarith + qed + have *: "((\u. f u / (u - z)) has_contour_integral 2 * complex_of_real pi * \ * f z) (circlepath z R)" + apply (rule Cauchy_integral_circlepath) + using \R > 0\ apply (auto intro: holomorphic_on_subset [OF holf] holomorphic_on_imp_continuous_on)+ + done + have "cmod (x - z) = R \ cmod (f x) * 2 \ cmod (f z)" for x + apply (simp add: R_def) + apply (rule less_imp_le) + apply (rule B) + using norm_triangle_ineq4 [of x z] + apply (auto simp:) + done + with \R > 0\ fz show False + using has_contour_integral_bound_circlepath [OF *, of "norm(f z)/2/R"] + by (auto simp: norm_mult norm_divide divide_simps) +qed + +proposition Liouville_weak: + assumes "f holomorphic_on UNIV" and "(f \ l) at_infinity" + shows "f z = l" + using Liouville_weak_0 [of "\z. f z - l"] + by (simp add: assms holomorphic_on_const holomorphic_on_diff LIM_zero) + + +proposition Liouville_weak_inverse: + assumes "f holomorphic_on UNIV" and unbounded: "\B. eventually (\x. norm (f x) \ B) at_infinity" + obtains z where "f z = 0" +proof - + { assume f: "\z. f z \ 0" + have 1: "(\x. 1 / f x) holomorphic_on UNIV" + by (simp add: holomorphic_on_divide holomorphic_on_const assms f) + have 2: "((\x. 1 / f x) \ 0) at_infinity" + apply (rule tendstoI [OF eventually_mono]) + apply (rule_tac B="2/e" in unbounded) + apply (simp add: dist_norm norm_divide divide_simps mult_ac) + done + have False + using Liouville_weak_0 [OF 1 2] f by simp + } + then show ?thesis + using that by blast +qed + + +text\ In particular we get the Fundamental Theorem of Algebra.\ + +theorem fundamental_theorem_of_algebra: + fixes a :: "nat \ complex" + assumes "a 0 = 0 \ (\i \ {1..n}. a i \ 0)" + obtains z where "(\i\n. a i * z^i) = 0" +using assms +proof (elim disjE bexE) + assume "a 0 = 0" then show ?thesis + by (auto simp: that [of 0]) +next + fix i + assume i: "i \ {1..n}" and nz: "a i \ 0" + have 1: "(\z. \i\n. a i * z^i) holomorphic_on UNIV" + by (rule holomorphic_intros)+ + show ?thesis + apply (rule Liouville_weak_inverse [OF 1]) + apply (rule polyfun_extremal) + apply (rule nz) + using i that + apply (auto simp:) + done +qed + + +subsection\ Weierstrass convergence theorem.\ + +proposition holomorphic_uniform_limit: + assumes cont: "eventually (\n. continuous_on (cball z r) (f n) \ (f n) holomorphic_on ball z r) F" + and lim: "\e. 0 < e \ eventually (\n. \x \ cball z r. norm(f n x - g x) < e) F" + and F: "~ trivial_limit F" + obtains "continuous_on (cball z r) g" "g holomorphic_on ball z r" +proof (cases r "0::real" rule: linorder_cases) + case less then show ?thesis by (force simp add: ball_empty less_imp_le continuous_on_def holomorphic_on_def intro: that) +next + case equal then show ?thesis + by (force simp add: holomorphic_on_def continuous_on_sing intro: that) +next + case greater + have contg: "continuous_on (cball z r) g" + using cont + by (fastforce simp: eventually_conj_iff dist_norm intro: eventually_mono [OF lim] continuous_uniform_limit [OF F]) + have 1: "continuous_on (path_image (circlepath z r)) (\x. 1 / (2 * complex_of_real pi * \) * g x)" + apply (rule continuous_intros continuous_on_subset [OF contg])+ + using \0 < r\ by auto + have 2: "((\u. 1 / (2 * of_real pi * \) * g u / (u - w) ^ 1) has_contour_integral g w) (circlepath z r)" + if w: "w \ ball z r" for w + proof - + define d where "d = (r - norm(w - z))" + have "0 < d" "d \ r" using w by (auto simp: norm_minus_commute d_def dist_norm) + have dle: "\u. cmod (z - u) = r \ d \ cmod (u - w)" + unfolding d_def by (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute) + have ev_int: "\\<^sub>F n in F. (\u. f n u / (u - w)) contour_integrable_on circlepath z r" + apply (rule eventually_mono [OF cont]) + using w + apply (auto intro: Cauchy_higher_derivative_integral_circlepath [where k=0, simplified]) + done + have ev_less: "\\<^sub>F n in F. \x\path_image (circlepath z r). cmod (f n x / (x - w) - g x / (x - w)) < e" + if "e > 0" for e + using greater \0 < d\ \0 < e\ + apply (simp add: norm_divide diff_divide_distrib [symmetric] divide_simps) + apply (rule_tac e1="e * d" in eventually_mono [OF lim]) + apply (force simp: dist_norm intro: dle mult_left_mono less_le_trans)+ + done + have g_cint: "(\u. g u/(u - w)) contour_integrable_on circlepath z r" + by (rule contour_integral_uniform_limit_circlepath [OF ev_int ev_less F \0 < r\]) + have cif_tends_cig: "((\n. contour_integral(circlepath z r) (\u. f n u / (u - w))) \ contour_integral(circlepath z r) (\u. g u/(u - w))) F" + by (rule contour_integral_uniform_limit_circlepath [OF ev_int ev_less F \0 < r\]) + have f_tends_cig: "((\n. 2 * of_real pi * \ * f n w) \ contour_integral (circlepath z r) (\u. g u / (u - w))) F" + apply (rule Lim_transform_eventually [where f = "\n. contour_integral (circlepath z r) (\u. f n u/(u - w))"]) + apply (rule eventually_mono [OF cont]) + apply (rule contour_integral_unique [OF Cauchy_integral_circlepath]) + using w + apply (auto simp: norm_minus_commute dist_norm cif_tends_cig) + done + have "((\n. 2 * of_real pi * \ * f n w) \ 2 * of_real pi * \ * g w) F" + apply (rule tendsto_mult_left [OF tendstoI]) + apply (rule eventually_mono [OF lim], assumption) + using w + apply (force simp add: dist_norm) + done + then have "((\u. g u / (u - w)) has_contour_integral 2 * of_real pi * \ * g w) (circlepath z r)" + using has_contour_integral_integral [OF g_cint] tendsto_unique [OF F f_tends_cig] w + by (force simp add: dist_norm) + then have "((\u. g u / (2 * of_real pi * \ * (u - w))) has_contour_integral g w) (circlepath z r)" + using has_contour_integral_div [where c = "2 * of_real pi * \"] + by (force simp add: field_simps) + then show ?thesis + by (simp add: dist_norm) + qed + show ?thesis + using Cauchy_next_derivative_circlepath(2) [OF 1 2, simplified] + by (fastforce simp add: holomorphic_on_open contg intro: that) +qed + + +text\ Version showing that the limit is the limit of the derivatives.\ + +proposition has_complex_derivative_uniform_limit: + fixes z::complex + assumes cont: "eventually (\n. continuous_on (cball z r) (f n) \ + (\w \ ball z r. ((f n) has_field_derivative (f' n w)) (at w))) F" + and lim: "\e. 0 < e \ eventually (\n. \x \ cball z r. norm(f n x - g x) < e) F" + and F: "~ trivial_limit F" and "0 < r" + obtains g' where + "continuous_on (cball z r) g" + "\w. w \ ball z r \ (g has_field_derivative (g' w)) (at w) \ ((\n. f' n w) \ g' w) F" +proof - + let ?conint = "contour_integral (circlepath z r)" + have g: "continuous_on (cball z r) g" "g holomorphic_on ball z r" + by (rule holomorphic_uniform_limit [OF eventually_mono [OF cont] lim F]; + auto simp: holomorphic_on_open field_differentiable_def)+ + then obtain g' where g': "\x. x \ ball z r \ (g has_field_derivative g' x) (at x)" + using DERIV_deriv_iff_has_field_derivative + by (fastforce simp add: holomorphic_on_open) + then have derg: "\x. x \ ball z r \ deriv g x = g' x" + by (simp add: DERIV_imp_deriv) + have tends_f'n_g': "((\n. f' n w) \ g' w) F" if w: "w \ ball z r" for w + proof - + have eq_f': "?conint (\x. f n x / (x - w)\<^sup>2) - ?conint (\x. g x / (x - w)\<^sup>2) = (f' n w - g' w) * (2 * of_real pi * \)" + if cont_fn: "continuous_on (cball z r) (f n)" + and fnd: "\w. w \ ball z r \ (f n has_field_derivative f' n w) (at w)" for n + proof - + have hol_fn: "f n holomorphic_on ball z r" + using fnd by (force simp add: holomorphic_on_open) + have "(f n has_field_derivative 1 / (2 * of_real pi * \) * ?conint (\u. f n u / (u - w)\<^sup>2)) (at w)" + by (rule Cauchy_derivative_integral_circlepath [OF cont_fn hol_fn w]) + then have f': "f' n w = 1 / (2 * of_real pi * \) * ?conint (\u. f n u / (u - w)\<^sup>2)" + using DERIV_unique [OF fnd] w by blast + show ?thesis + by (simp add: f' Cauchy_contour_integral_circlepath_2 [OF g w] derg [OF w] divide_simps) + qed + define d where "d = (r - norm(w - z))^2" + have "d > 0" + using w by (simp add: dist_commute dist_norm d_def) + have dle: "\y. r = cmod (z - y) \ d \ cmod ((y - w)\<^sup>2)" + apply (simp add: d_def norm_power) + apply (rule power_mono) + apply (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute) + apply (metis diff_ge_0_iff_ge dist_commute dist_norm less_eq_real_def mem_ball w) + done + have 1: "\\<^sub>F n in F. (\x. f n x / (x - w)\<^sup>2) contour_integrable_on circlepath z r" + by (force simp add: holomorphic_on_open intro: w Cauchy_derivative_integral_circlepath eventually_mono [OF cont]) + have 2: "0 < e \ \\<^sub>F n in F. \x \ path_image (circlepath z r). cmod (f n x / (x - w)\<^sup>2 - g x / (x - w)\<^sup>2) < e" for e + using \r > 0\ + apply (simp add: diff_divide_distrib [symmetric] norm_divide divide_simps sphere_def) + apply (rule eventually_mono [OF lim, of "e*d"]) + apply (simp add: \0 < d\) + apply (force simp add: dist_norm dle intro: less_le_trans) + done + have "((\n. contour_integral (circlepath z r) (\x. f n x / (x - w)\<^sup>2)) + \ contour_integral (circlepath z r) ((\x. g x / (x - w)\<^sup>2))) F" + by (rule contour_integral_uniform_limit_circlepath [OF 1 2 F \0 < r\]) + then have tendsto_0: "((\n. 1 / (2 * of_real pi * \) * (?conint (\x. f n x / (x - w)\<^sup>2) - ?conint (\x. g x / (x - w)\<^sup>2))) \ 0) F" + using Lim_null by (force intro!: tendsto_mult_right_zero) + have "((\n. f' n w - g' w) \ 0) F" + apply (rule Lim_transform_eventually [OF _ tendsto_0]) + apply (force simp add: divide_simps intro: eq_f' eventually_mono [OF cont]) + done + then show ?thesis using Lim_null by blast + qed + obtain g' where "\w. w \ ball z r \ (g has_field_derivative (g' w)) (at w) \ ((\n. f' n w) \ g' w) F" + by (blast intro: tends_f'n_g' g' ) + then show ?thesis using g + using that by blast +qed + + +subsection\Some more simple/convenient versions for applications.\ + +lemma holomorphic_uniform_sequence: + assumes s: "open s" + and hol_fn: "\n. (f n) holomorphic_on s" + and to_g: "\x. x \ s + \ \d. 0 < d \ cball x d \ s \ + (\e. 0 < e \ eventually (\n. \y \ cball x d. norm(f n y - g y) < e) sequentially)" + shows "g holomorphic_on s" +proof - + have "\f'. (g has_field_derivative f') (at z)" if "z \ s" for z + proof - + obtain r where "0 < r" and r: "cball z r \ s" + and fg: "\e. 0 < e \ eventually (\n. \y \ cball z r. norm(f n y - g y) < e) sequentially" + using to_g [OF \z \ s\] by blast + have *: "\\<^sub>F n in sequentially. continuous_on (cball z r) (f n) \ f n holomorphic_on ball z r" + apply (intro eventuallyI conjI) + using hol_fn holomorphic_on_imp_continuous_on holomorphic_on_subset r apply blast + apply (metis hol_fn holomorphic_on_subset interior_cball interior_subset r) + done + show ?thesis + apply (rule holomorphic_uniform_limit [OF *]) + using \0 < r\ centre_in_ball fg + apply (auto simp: holomorphic_on_open) + done + qed + with s show ?thesis + by (simp add: holomorphic_on_open) +qed + +lemma has_complex_derivative_uniform_sequence: + fixes s :: "complex set" + assumes s: "open s" + and hfd: "\n x. x \ s \ ((f n) has_field_derivative f' n x) (at x)" + and to_g: "\x. x \ s + \ \d. 0 < d \ cball x d \ s \ + (\e. 0 < e \ eventually (\n. \y \ cball x d. norm(f n y - g y) < e) sequentially)" + shows "\g'. \x \ s. (g has_field_derivative g' x) (at x) \ ((\n. f' n x) \ g' x) sequentially" +proof - + have y: "\y. (g has_field_derivative y) (at z) \ (\n. f' n z) \ y" if "z \ s" for z + proof - + obtain r where "0 < r" and r: "cball z r \ s" + and fg: "\e. 0 < e \ eventually (\n. \y \ cball z r. norm(f n y - g y) < e) sequentially" + using to_g [OF \z \ s\] by blast + have *: "\\<^sub>F n in sequentially. continuous_on (cball z r) (f n) \ + (\w \ ball z r. ((f n) has_field_derivative (f' n w)) (at w))" + apply (intro eventuallyI conjI) + apply (meson hfd holomorphic_on_imp_continuous_on holomorphic_on_open holomorphic_on_subset r s) + using ball_subset_cball hfd r apply blast + done + show ?thesis + apply (rule has_complex_derivative_uniform_limit [OF *, of g]) + using \0 < r\ centre_in_ball fg + apply force+ + done + qed + show ?thesis + by (rule bchoice) (blast intro: y) +qed + + +subsection\On analytic functions defined by a series.\ + +lemma series_and_derivative_comparison: + fixes s :: "complex set" + assumes s: "open s" + and h: "summable h" + and hfd: "\n x. x \ s \ (f n has_field_derivative f' n x) (at x)" + and to_g: "\n x. \N \ n; x \ s\ \ norm(f n x) \ h n" + obtains g g' where "\x \ s. ((\n. f n x) sums g x) \ ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" +proof - + obtain g where g: "\e. e>0 \ \N. \n x. N \ n \ x \ s \ dist (\nd>0. cball x d \ s \ (\e>0. \\<^sub>F n in sequentially. \y\cball x d. cmod ((\a s" for x + proof - + obtain d where "d>0" and d: "cball x d \ s" + using open_contains_cball [of "s"] \x \ s\ s by blast + then show ?thesis + apply (rule_tac x=d in exI) + apply (auto simp: dist_norm eventually_sequentially) + apply (metis g contra_subsetD dist_norm) + done + qed + have "(\x\s. (\n. \i g x)" + using g by (force simp add: lim_sequentially) + moreover have "\g'. \x\s. (g has_field_derivative g' x) (at x) \ (\n. \i g' x" + by (rule has_complex_derivative_uniform_sequence [OF s]) (auto intro: * hfd DERIV_setsum)+ + ultimately show ?thesis + by (force simp add: sums_def conj_commute intro: that) +qed + +text\A version where we only have local uniform/comparative convergence.\ + +lemma series_and_derivative_comparison_local: + fixes s :: "complex set" + assumes s: "open s" + and hfd: "\n x. x \ s \ (f n has_field_derivative f' n x) (at x)" + and to_g: "\x. x \ s \ + \d h N. 0 < d \ summable h \ (\n y. N \ n \ y \ ball x d \ norm(f n y) \ h n)" + shows "\g g'. \x \ s. ((\n. f n x) sums g x) \ ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" +proof - + have "\y. (\n. f n z) sums (\n. f n z) \ (\n. f' n z) sums y \ ((\x. \n. f n x) has_field_derivative y) (at z)" + if "z \ s" for z + proof - + obtain d h N where "0 < d" "summable h" and le_h: "\n y. \N \ n; y \ ball z d\ \ norm(f n y) \ h n" + using to_g \z \ s\ by meson + then obtain r where "r>0" and r: "ball z r \ ball z d \ s" using \z \ s\ s + by (metis Int_iff open_ball centre_in_ball open_Int open_contains_ball_eq) + have 1: "open (ball z d \ s)" + by (simp add: open_Int s) + have 2: "\n x. x \ ball z d \ s \ (f n has_field_derivative f' n x) (at x)" + by (auto simp: hfd) + obtain g g' where gg': "\x \ ball z d \ s. ((\n. f n x) sums g x) \ + ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" + by (auto intro: le_h series_and_derivative_comparison [OF 1 \summable h\ hfd]) + then have "(\n. f' n z) sums g' z" + by (meson \0 < r\ centre_in_ball contra_subsetD r) + moreover have "(\n. f n z) sums (\n. f n z)" + by (metis summable_comparison_test' summable_sums centre_in_ball \0 < d\ \summable h\ le_h) + moreover have "((\x. \n. f n x) has_field_derivative g' z) (at z)" + apply (rule_tac f=g in DERIV_transform_at [OF _ \0 < r\]) + apply (simp add: gg' \z \ s\ \0 < d\) + apply (metis (full_types) contra_subsetD dist_commute gg' mem_ball r sums_unique) + done + ultimately show ?thesis by auto + qed + then show ?thesis + by (rule_tac x="\x. suminf (\n. f n x)" in exI) meson +qed + + +text\Sometimes convenient to compare with a complex series of positive reals. (?)\ + +lemma series_and_derivative_comparison_complex: + fixes s :: "complex set" + assumes s: "open s" + and hfd: "\n x. x \ s \ (f n has_field_derivative f' n x) (at x)" + and to_g: "\x. x \ s \ + \d h N. 0 < d \ summable h \ range h \ nonneg_Reals \ (\n y. N \ n \ y \ ball x d \ cmod(f n y) \ cmod (h n))" + shows "\g g'. \x \ s. ((\n. f n x) sums g x) \ ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" +apply (rule series_and_derivative_comparison_local [OF s hfd], assumption) +apply (frule to_g) +apply (erule ex_forward) +apply (erule exE) +apply (rule_tac x="Re o h" in exI) +apply (erule ex_forward) +apply (simp add: summable_Re o_def ) +apply (elim conjE all_forward) +apply (simp add: nonneg_Reals_cmod_eq_Re image_subset_iff) +done + + +text\In particular, a power series is analytic inside circle of convergence.\ + +lemma power_series_and_derivative_0: + fixes a :: "nat \ complex" and r::real + assumes "summable (\n. a n * r^n)" + shows "\g g'. \z. cmod z < r \ + ((\n. a n * z^n) sums g z) \ ((\n. of_nat n * a n * z^(n - 1)) sums g' z) \ (g has_field_derivative g' z) (at z)" +proof (cases "0 < r") + case True + have der: "\n z. ((\x. a n * x ^ n) has_field_derivative of_nat n * a n * z ^ (n - 1)) (at z)" + by (rule derivative_eq_intros | simp)+ + have y_le: "\cmod (z - y) * 2 < r - cmod z\ \ cmod y \ cmod (of_real r + of_real (cmod z)) / 2" for z y + using \r > 0\ + apply (auto simp: algebra_simps norm_mult norm_divide norm_power of_real_add [symmetric] simp del: of_real_add) + using norm_triangle_ineq2 [of y z] + apply (simp only: diff_le_eq norm_minus_commute mult_2) + done + have "summable (\n. a n * complex_of_real r ^ n)" + using assms \r > 0\ by simp + moreover have "\z. cmod z < r \ cmod ((of_real r + of_real (cmod z)) / 2) < cmod (of_real r)" + using \r > 0\ + by (simp add: of_real_add [symmetric] del: of_real_add) + ultimately have sum: "\z. cmod z < r \ summable (\n. of_real (cmod (a n)) * ((of_real r + complex_of_real (cmod z)) / 2) ^ n)" + by (rule power_series_conv_imp_absconv_weak) + have "\g g'. \z \ ball 0 r. (\n. (a n) * z ^ n) sums g z \ + (\n. of_nat n * (a n) * z ^ (n - 1)) sums g' z \ (g has_field_derivative g' z) (at z)" + apply (rule series_and_derivative_comparison_complex [OF open_ball der]) + apply (rule_tac x="(r - norm z)/2" in exI) + apply (simp add: dist_norm) + apply (rule_tac x="\n. of_real(norm(a n)*((r + norm z)/2)^n)" in exI) + using \r > 0\ + apply (auto simp: sum nonneg_Reals_divide_I) + apply (rule_tac x=0 in exI) + apply (force simp: norm_mult norm_divide norm_power intro!: mult_left_mono power_mono y_le) + done + then show ?thesis + by (simp add: dist_0_norm ball_def) +next + case False then show ?thesis + apply (simp add: not_less) + using less_le_trans norm_not_less_zero by blast +qed + +proposition power_series_and_derivative: + fixes a :: "nat \ complex" and r::real + assumes "summable (\n. a n * r^n)" + obtains g g' where "\z \ ball w r. + ((\n. a n * (z - w) ^ n) sums g z) \ ((\n. of_nat n * a n * (z - w) ^ (n - 1)) sums g' z) \ + (g has_field_derivative g' z) (at z)" + using power_series_and_derivative_0 [OF assms] + apply clarify + apply (rule_tac g="(\z. g(z - w))" in that) + using DERIV_shift [where z="-w"] + apply (auto simp: norm_minus_commute Ball_def dist_norm) + done + +proposition power_series_holomorphic: + assumes "\w. w \ ball z r \ ((\n. a n*(w - z)^n) sums f w)" + shows "f holomorphic_on ball z r" +proof - + have "\f'. (f has_field_derivative f') (at w)" if w: "dist z w < r" for w + proof - + have inb: "z + complex_of_real ((dist z w + r) / 2) \ ball z r" + proof - + have wz: "cmod (w - z) < r" using w + by (auto simp: divide_simps dist_norm norm_minus_commute) + then have "0 \ r" + by (meson less_eq_real_def norm_ge_zero order_trans) + show ?thesis + using w by (simp add: dist_norm \0\r\ of_real_add [symmetric] del: of_real_add) + qed + have sum: "summable (\n. a n * of_real (((cmod (z - w) + r) / 2) ^ n))" + using assms [OF inb] by (force simp add: summable_def dist_norm) + obtain g g' where gg': "\u. u \ ball z ((cmod (z - w) + r) / 2) \ + (\n. a n * (u - z) ^ n) sums g u \ + (\n. of_nat n * a n * (u - z) ^ (n - 1)) sums g' u \ (g has_field_derivative g' u) (at u)" + by (rule power_series_and_derivative [OF sum, of z]) fastforce + have [simp]: "g u = f u" if "cmod (u - w) < (r - cmod (z - w)) / 2" for u + proof - + have less: "cmod (z - u) * 2 < cmod (z - w) + r" + using that dist_triangle2 [of z u w] + by (simp add: dist_norm [symmetric] algebra_simps) + show ?thesis + apply (rule sums_unique2 [of "\n. a n*(u - z)^n"]) + using gg' [of u] less w + apply (auto simp: assms dist_norm) + done + qed + show ?thesis + apply (rule_tac x="g' w" in exI) + apply (rule DERIV_transform_at [where f=g and d="(r - norm(z - w))/2"]) + using w gg' [of w] + apply (auto simp: dist_norm) + done + qed + then show ?thesis by (simp add: holomorphic_on_open) +qed + +corollary holomorphic_iff_power_series: + "f holomorphic_on ball z r \ + (\w \ ball z r. (\n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)" + apply (intro iffI ballI) + using holomorphic_power_series apply force + apply (rule power_series_holomorphic [where a = "\n. (deriv ^^ n) f z / (fact n)"]) + apply force + done + +corollary power_series_analytic: + "(\w. w \ ball z r \ (\n. a n*(w - z)^n) sums f w) \ f analytic_on ball z r" + by (force simp add: analytic_on_open intro!: power_series_holomorphic) + +corollary analytic_iff_power_series: + "f analytic_on ball z r \ + (\w \ ball z r. (\n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)" + by (simp add: analytic_on_open holomorphic_iff_power_series) + + +subsection\Equality between holomorphic functions, on open ball then connected set.\ + +lemma holomorphic_fun_eq_on_ball: + "\f holomorphic_on ball z r; g holomorphic_on ball z r; + w \ ball z r; + \n. (deriv ^^ n) f z = (deriv ^^ n) g z\ + \ f w = g w" + apply (rule sums_unique2 [of "\n. (deriv ^^ n) f z / (fact n) * (w - z)^n"]) + apply (auto simp: holomorphic_iff_power_series) + done + +lemma holomorphic_fun_eq_0_on_ball: + "\f holomorphic_on ball z r; w \ ball z r; + \n. (deriv ^^ n) f z = 0\ + \ f w = 0" + apply (rule sums_unique2 [of "\n. (deriv ^^ n) f z / (fact n) * (w - z)^n"]) + apply (auto simp: holomorphic_iff_power_series) + done + +lemma holomorphic_fun_eq_0_on_connected: + assumes holf: "f holomorphic_on s" and "open s" + and cons: "connected s" + and der: "\n. (deriv ^^ n) f z = 0" + and "z \ s" "w \ s" + shows "f w = 0" +proof - + have *: "\x e. \ \xa. (deriv ^^ xa) f x = 0; ball x e \ s\ + \ ball x e \ (\n. {w \ s. (deriv ^^ n) f w = 0})" + apply auto + apply (rule holomorphic_fun_eq_0_on_ball [OF holomorphic_higher_deriv]) + apply (rule holomorphic_on_subset [OF holf], simp_all) + by (metis funpow_add o_apply) + have 1: "openin (subtopology euclidean s) (\n. {w \ s. (deriv ^^ n) f w = 0})" + apply (rule open_subset, force) + using \open s\ + apply (simp add: open_contains_ball Ball_def) + apply (erule all_forward) + using "*" by auto blast+ + have 2: "closedin (subtopology euclidean s) (\n. {w \ s. (deriv ^^ n) f w = 0})" + using assms + by (auto intro: continuous_closedin_preimage_constant holomorphic_on_imp_continuous_on holomorphic_higher_deriv) + obtain e where "e>0" and e: "ball w e \ s" using openE [OF \open s\ \w \ s\] . + then have holfb: "f holomorphic_on ball w e" + using holf holomorphic_on_subset by blast + have 3: "(\n. {w \ s. (deriv ^^ n) f w = 0}) = s \ f w = 0" + using \e>0\ e by (force intro: holomorphic_fun_eq_0_on_ball [OF holfb]) + show ?thesis + using cons der \z \ s\ + apply (simp add: connected_clopen) + apply (drule_tac x="\n. {w \ s. (deriv ^^ n) f w = 0}" in spec) + apply (auto simp: 1 2 3) + done +qed + +lemma holomorphic_fun_eq_on_connected: + assumes "f holomorphic_on s" "g holomorphic_on s" and "open s" "connected s" + and "\n. (deriv ^^ n) f z = (deriv ^^ n) g z" + and "z \ s" "w \ s" + shows "f w = g w" + apply (rule holomorphic_fun_eq_0_on_connected [of "\x. f x - g x" s z, simplified]) + apply (intro assms holomorphic_intros) + using assms apply simp_all + apply (subst higher_deriv_diff, auto) + done + +lemma holomorphic_fun_eq_const_on_connected: + assumes holf: "f holomorphic_on s" and "open s" + and cons: "connected s" + and der: "\n. 0 < n \ (deriv ^^ n) f z = 0" + and "z \ s" "w \ s" + shows "f w = f z" + apply (rule holomorphic_fun_eq_0_on_connected [of "\w. f w - f z" s z, simplified]) + apply (intro assms holomorphic_intros) + using assms apply simp_all + apply (subst higher_deriv_diff) + apply (intro holomorphic_intros | simp)+ + done + + +subsection\Some basic lemmas about poles/singularities.\ + +lemma pole_lemma: + assumes holf: "f holomorphic_on s" and a: "a \ interior s" + shows "(\z. if z = a then deriv f a + else (f z - f a) / (z - a)) holomorphic_on s" (is "?F holomorphic_on s") +proof - + have F1: "?F field_differentiable (at u within s)" if "u \ s" "u \ a" for u + proof - + have fcd: "f field_differentiable at u within s" + using holf holomorphic_on_def by (simp add: \u \ s\) + have cd: "(\z. (f z - f a) / (z - a)) field_differentiable at u within s" + by (rule fcd derivative_intros | simp add: that)+ + have "0 < dist a u" using that dist_nz by blast + then show ?thesis + by (rule field_differentiable_transform_within [OF _ _ _ cd]) (auto simp: \u \ s\) + qed + have F2: "?F field_differentiable at a" if "0 < e" "ball a e \ s" for e + proof - + have holfb: "f holomorphic_on ball a e" + by (rule holomorphic_on_subset [OF holf \ball a e \ s\]) + have 2: "?F holomorphic_on ball a e - {a}" + apply (rule holomorphic_on_subset [where s = "s - {a}"]) + apply (simp add: holomorphic_on_def field_differentiable_def [symmetric]) + using mem_ball that + apply (auto intro: F1 field_differentiable_within_subset) + done + have "isCont (\z. if z = a then deriv f a else (f z - f a) / (z - a)) x" + if "dist a x < e" for x + proof (cases "x=a") + case True then show ?thesis + using holfb \0 < e\ + apply (simp add: holomorphic_on_open field_differentiable_def [symmetric]) + apply (drule_tac x=a in bspec) + apply (auto simp: DERIV_deriv_iff_field_differentiable [symmetric] continuous_at DERIV_iff2 + elim: rev_iffD1 [OF _ LIM_equal]) + done + next + case False with 2 that show ?thesis + by (force simp: holomorphic_on_open open_Diff field_differentiable_def [symmetric] field_differentiable_imp_continuous_at) + qed + then have 1: "continuous_on (ball a e) ?F" + by (clarsimp simp: continuous_on_eq_continuous_at) + have "?F holomorphic_on ball a e" + by (auto intro: no_isolated_singularity [OF 1 2]) + with that show ?thesis + by (simp add: holomorphic_on_open field_differentiable_def [symmetric] + field_differentiable_at_within) + qed + show ?thesis + proof + fix x assume "x \ s" show "?F field_differentiable at x within s" + proof (cases "x=a") + case True then show ?thesis + using a by (auto simp: mem_interior intro: field_differentiable_at_within F2) + next + case False with F1 \x \ s\ + show ?thesis by blast + qed + qed +qed + +proposition pole_theorem: + assumes holg: "g holomorphic_on s" and a: "a \ interior s" + and eq: "\z. z \ s - {a} \ g z = (z - a) * f z" + shows "(\z. if z = a then deriv g a + else f z - g a/(z - a)) holomorphic_on s" + using pole_lemma [OF holg a] + by (rule holomorphic_transform) (simp add: eq divide_simps) + +lemma pole_lemma_open: + assumes "f holomorphic_on s" "open s" + shows "(\z. if z = a then deriv f a else (f z - f a)/(z - a)) holomorphic_on s" +proof (cases "a \ s") + case True with assms interior_eq pole_lemma + show ?thesis by fastforce +next + case False with assms show ?thesis + apply (simp add: holomorphic_on_def field_differentiable_def [symmetric], clarify) + apply (rule field_differentiable_transform_within [where f = "\z. (f z - f a)/(z - a)" and d = 1]) + apply (rule derivative_intros | force)+ + done +qed + +proposition pole_theorem_open: + assumes holg: "g holomorphic_on s" and s: "open s" + and eq: "\z. z \ s - {a} \ g z = (z - a) * f z" + shows "(\z. if z = a then deriv g a + else f z - g a/(z - a)) holomorphic_on s" + using pole_lemma_open [OF holg s] + by (rule holomorphic_transform) (auto simp: eq divide_simps) + +proposition pole_theorem_0: + assumes holg: "g holomorphic_on s" and a: "a \ interior s" + and eq: "\z. z \ s - {a} \ g z = (z - a) * f z" + and [simp]: "f a = deriv g a" "g a = 0" + shows "f holomorphic_on s" + using pole_theorem [OF holg a eq] + by (rule holomorphic_transform) (auto simp: eq divide_simps) + +proposition pole_theorem_open_0: + assumes holg: "g holomorphic_on s" and s: "open s" + and eq: "\z. z \ s - {a} \ g z = (z - a) * f z" + and [simp]: "f a = deriv g a" "g a = 0" + shows "f holomorphic_on s" + using pole_theorem_open [OF holg s eq] + by (rule holomorphic_transform) (auto simp: eq divide_simps) + +lemma pole_theorem_analytic: + assumes g: "g analytic_on s" + and eq: "\z. z \ s + \ \d. 0 < d \ (\w \ ball z d - {a}. g w = (w - a) * f w)" + shows "(\z. if z = a then deriv g a + else f z - g a/(z - a)) analytic_on s" +using g +apply (simp add: analytic_on_def Ball_def) +apply (safe elim!: all_forward dest!: eq) +apply (rule_tac x="min d e" in exI, simp) +apply (rule pole_theorem_open) +apply (auto simp: holomorphic_on_subset subset_ball) +done + +lemma pole_theorem_analytic_0: + assumes g: "g analytic_on s" + and eq: "\z. z \ s \ \d. 0 < d \ (\w \ ball z d - {a}. g w = (w - a) * f w)" + and [simp]: "f a = deriv g a" "g a = 0" + shows "f analytic_on s" +proof - + have [simp]: "(\z. if z = a then deriv g a else f z - g a / (z - a)) = f" + by auto + show ?thesis + using pole_theorem_analytic [OF g eq] by simp +qed + +lemma pole_theorem_analytic_open_superset: + assumes g: "g analytic_on s" and "s \ t" "open t" + and eq: "\z. z \ t - {a} \ g z = (z - a) * f z" + shows "(\z. if z = a then deriv g a + else f z - g a/(z - a)) analytic_on s" + apply (rule pole_theorem_analytic [OF g]) + apply (rule openE [OF \open t\]) + using assms eq by auto + +lemma pole_theorem_analytic_open_superset_0: + assumes g: "g analytic_on s" "s \ t" "open t" "\z. z \ t - {a} \ g z = (z - a) * f z" + and [simp]: "f a = deriv g a" "g a = 0" + shows "f analytic_on s" +proof - + have [simp]: "(\z. if z = a then deriv g a else f z - g a / (z - a)) = f" + by auto + have "(\z. if z = a then deriv g a else f z - g a/(z - a)) analytic_on s" + by (rule pole_theorem_analytic_open_superset [OF g]) + then show ?thesis by simp +qed + + + +subsection\General, homology form of Cauchy's theorem.\ + +text\Proof is based on Dixon's, as presented in Lang's "Complex Analysis" book (page 147).\ + +lemma contour_integral_continuous_on_linepath_2D: + assumes "open u" and cont_dw: "\w. w \ u \ F w contour_integrable_on (linepath a b)" + and cond_uu: "continuous_on (u \ u) (\(x,y). F x y)" + and abu: "closed_segment a b \ u" + shows "continuous_on u (\w. contour_integral (linepath a b) (F w))" +proof - + have *: "\d>0. \x'\u. dist x' w < d \ + dist (contour_integral (linepath a b) (F x')) + (contour_integral (linepath a b) (F w)) \ \" + if "w \ u" "0 < \" "a \ b" for w \ + proof - + obtain \ where "\>0" and \: "cball w \ \ u" using open_contains_cball \open u\ \w \ u\ by force + let ?TZ = "{(t,z) |t z. t \ cball w \ \ z \ closed_segment a b}" + have "uniformly_continuous_on ?TZ (\(x,y). F x y)" + apply (rule compact_uniformly_continuous) + apply (rule continuous_on_subset[OF cond_uu]) + using abu \ + apply (auto simp: compact_Times simp del: mem_cball) + done + then obtain \ where "\>0" + and \: "\x x'. \x\?TZ; x'\?TZ; dist x' x < \\ \ + dist ((\(x,y). F x y) x') ((\(x,y). F x y) x) < \/norm(b - a)" + apply (rule uniformly_continuous_onE [where e = "\/norm(b - a)"]) + using \0 < \\ \a \ b\ by auto + have \: "\norm (w - x1) \ \; x2 \ closed_segment a b; + norm (w - x1') \ \; x2' \ closed_segment a b; norm ((x1', x2') - (x1, x2)) < \\ + \ norm (F x1' x2' - F x1 x2) \ \ / cmod (b - a)" + for x1 x2 x1' x2' + using \ [of "(x1,x2)" "(x1',x2')"] by (force simp add: dist_norm) + have le_ee: "cmod (contour_integral (linepath a b) (\x. F x' x - F w x)) \ \" + if "x' \ u" "cmod (x' - w) < \" "cmod (x' - w) < \" for x' + proof - + have "cmod (contour_integral (linepath a b) (\x. F x' x - F w x)) \ \/norm(b - a) * norm(b - a)" + apply (rule has_contour_integral_bound_linepath [OF has_contour_integral_integral _ \]) + apply (rule contour_integrable_diff [OF cont_dw cont_dw]) + using \0 < \\ \a \ b\ \0 < \\ \w \ u\ that + apply (auto simp: norm_minus_commute) + done + also have "... = \" using \a \ b\ by simp + finally show ?thesis . + qed + show ?thesis + apply (rule_tac x="min \ \" in exI) + using \0 < \\ \0 < \\ + apply (auto simp: dist_norm contour_integral_diff [OF cont_dw cont_dw, symmetric] \w \ u\ intro: le_ee) + done + qed + show ?thesis + apply (rule continuous_onI) + apply (cases "a=b") + apply (auto intro: *) + done +qed + +text\This version has @{term"polynomial_function \"} as an additional assumption.\ +lemma Cauchy_integral_formula_global_weak: + assumes u: "open u" and holf: "f holomorphic_on u" + and z: "z \ u" and \: "polynomial_function \" + and pasz: "path_image \ \ u - {z}" and loop: "pathfinish \ = pathstart \" + and zero: "\w. w \ u \ winding_number \ w = 0" + shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" +proof - + obtain \' where pf\': "polynomial_function \'" and \': "\x. (\ has_vector_derivative (\' x)) (at x)" + using has_vector_derivative_polynomial_function [OF \] by blast + then have "bounded(path_image \')" + by (simp add: path_image_def compact_imp_bounded compact_continuous_image continuous_on_polymonial_function) + then obtain B where "B>0" and B: "\x. x \ path_image \' \ norm x \ B" + using bounded_pos by force + define d where [abs_def]: "d z w = (if w = z then deriv f z else (f w - f z)/(w - z))" for z w + define v where "v = {w. w \ path_image \ \ winding_number \ w = 0}" + have "path \" "valid_path \" using \ + by (auto simp: path_polynomial_function valid_path_polynomial_function) + then have ov: "open v" + by (simp add: v_def open_winding_number_levelsets loop) + have uv_Un: "u \ v = UNIV" + using pasz zero by (auto simp: v_def) + have conf: "continuous_on u f" + by (metis holf holomorphic_on_imp_continuous_on) + have hol_d: "(d y) holomorphic_on u" if "y \ u" for y + proof - + have *: "(\c. if c = y then deriv f y else (f c - f y) / (c - y)) holomorphic_on u" + by (simp add: holf pole_lemma_open u) + then have "isCont (\x. if x = y then deriv f y else (f x - f y) / (x - y)) y" + using at_within_open field_differentiable_imp_continuous_at holomorphic_on_def that u by fastforce + then have "continuous_on u (d y)" + apply (simp add: d_def continuous_on_eq_continuous_at u, clarify) + using * holomorphic_on_def + by (meson field_differentiable_within_open field_differentiable_imp_continuous_at u) + moreover have "d y holomorphic_on u - {y}" + apply (simp add: d_def holomorphic_on_open u open_delete field_differentiable_def [symmetric]) + apply (intro ballI) + apply (rename_tac w) + apply (rule_tac d="dist w y" and f = "\w. (f w - f y)/(w - y)" in field_differentiable_transform_within) + apply (auto simp: dist_pos_lt dist_commute intro!: derivative_intros) + using analytic_on_imp_differentiable_at analytic_on_open holf u apply blast + done + ultimately show ?thesis + by (rule no_isolated_singularity) (auto simp: u) + qed + have cint_fxy: "(\x. (f x - f y) / (x - y)) contour_integrable_on \" if "y \ path_image \" for y + apply (rule contour_integrable_holomorphic_simple [where s = "u-{y}"]) + using \valid_path \\ pasz + apply (auto simp: u open_delete) + apply (rule continuous_intros holomorphic_intros continuous_on_subset [OF conf] holomorphic_on_subset [OF holf] | + force simp add: that)+ + done + define h where + "h z = (if z \ u then contour_integral \ (d z) else contour_integral \ (\w. f w/(w - z)))" for z + have U: "\z. z \ u \ ((d z) has_contour_integral h z) \" + apply (simp add: h_def) + apply (rule has_contour_integral_integral [OF contour_integrable_holomorphic_simple [where s=u]]) + using u pasz \valid_path \\ + apply (auto intro: holomorphic_on_imp_continuous_on hol_d) + done + have V: "((\w. f w / (w - z)) has_contour_integral h z) \" if z: "z \ v" for z + proof - + have 0: "0 = (f z) * 2 * of_real (2 * pi) * \ * winding_number \ z" + using v_def z by auto + then have "((\x. 1 / (x - z)) has_contour_integral 0) \" + using z v_def has_contour_integral_winding_number [OF \valid_path \\] by fastforce + then have "((\x. f z * (1 / (x - z))) has_contour_integral 0) \" + using has_contour_integral_lmul by fastforce + then have "((\x. f z / (x - z)) has_contour_integral 0) \" + by (simp add: divide_simps) + moreover have "((\x. (f x - f z) / (x - z)) has_contour_integral contour_integral \ (d z)) \" + using z + apply (auto simp: v_def) + apply (metis (no_types, lifting) contour_integrable_eq d_def has_contour_integral_eq has_contour_integral_integral cint_fxy) + done + ultimately have *: "((\x. f z / (x - z) + (f x - f z) / (x - z)) has_contour_integral (0 + contour_integral \ (d z))) \" + by (rule has_contour_integral_add) + have "((\w. f w / (w - z)) has_contour_integral contour_integral \ (d z)) \" + if "z \ u" + using * by (auto simp: divide_simps has_contour_integral_eq) + moreover have "((\w. f w / (w - z)) has_contour_integral contour_integral \ (\w. f w / (w - z))) \" + if "z \ u" + apply (rule has_contour_integral_integral [OF contour_integrable_holomorphic_simple [where s=u]]) + using u pasz \valid_path \\ that + apply (auto intro: holomorphic_on_imp_continuous_on hol_d) + apply (rule continuous_intros conf holomorphic_intros holf | force)+ + done + ultimately show ?thesis + using z by (simp add: h_def) + qed + have znot: "z \ path_image \" + using pasz by blast + obtain d0 where "d0>0" and d0: "\x y. x \ path_image \ \ y \ - u \ d0 \ dist x y" + using separate_compact_closed [of "path_image \" "-u"] pasz u + by (fastforce simp add: \path \\ compact_path_image) + obtain dd where "0 < dd" and dd: "{y + k | y k. y \ path_image \ \ k \ ball 0 dd} \ u" + apply (rule that [of "d0/2"]) + using \0 < d0\ + apply (auto simp: dist_norm dest: d0) + done + have "\x x'. \x \ path_image \; dist x x' * 2 < dd\ \ \y k. x' = y + k \ y \ path_image \ \ dist 0 k * 2 \ dd" + apply (rule_tac x=x in exI) + apply (rule_tac x="x'-x" in exI) + apply (force simp add: dist_norm) + done + then have 1: "path_image \ \ interior {y + k |y k. y \ path_image \ \ k \ cball 0 (dd / 2)}" + apply (clarsimp simp add: mem_interior) + using \0 < dd\ + apply (rule_tac x="dd/2" in exI, auto) + done + obtain t where "compact t" and subt: "path_image \ \ interior t" and t: "t \ u" + apply (rule that [OF _ 1]) + apply (fastforce simp add: \valid_path \\ compact_valid_path_image intro!: compact_sums) + apply (rule order_trans [OF _ dd]) + using \0 < dd\ by fastforce + obtain L where "L>0" + and L: "\f B. \f holomorphic_on interior t; \z. z\interior t \ cmod (f z) \ B\ \ + cmod (contour_integral \ f) \ L * B" + using contour_integral_bound_exists [OF open_interior \valid_path \\ subt] + by blast + have "bounded(f ` t)" + by (meson \compact t\ compact_continuous_image compact_imp_bounded conf continuous_on_subset t) + then obtain D where "D>0" and D: "\x. x \ t \ norm (f x) \ D" + by (auto simp: bounded_pos) + obtain C where "C>0" and C: "\x. x \ t \ norm x \ C" + using \compact t\ bounded_pos compact_imp_bounded by force + have "dist (h y) 0 \ e" if "0 < e" and le: "D * L / e + C \ cmod y" for e y + proof - + have "D * L / e > 0" using \D>0\ \L>0\ \e>0\ by simp + with le have ybig: "norm y > C" by force + with C have "y \ t" by force + then have ynot: "y \ path_image \" + using subt interior_subset by blast + have [simp]: "winding_number \ y = 0" + apply (rule winding_number_zero_outside [of _ "cball 0 C"]) + using ybig interior_subset subt + apply (force simp add: loop \path \\ dist_norm intro!: C)+ + done + have [simp]: "h y = contour_integral \ (\w. f w/(w - y))" + by (rule contour_integral_unique [symmetric]) (simp add: v_def ynot V) + have holint: "(\w. f w / (w - y)) holomorphic_on interior t" + apply (rule holomorphic_on_divide) + using holf holomorphic_on_subset interior_subset t apply blast + apply (rule holomorphic_intros)+ + using \y \ t\ interior_subset by auto + have leD: "cmod (f z / (z - y)) \ D * (e / L / D)" if z: "z \ interior t" for z + proof - + have "D * L / e + cmod z \ cmod y" + using le C [of z] z using interior_subset by force + then have DL2: "D * L / e \ cmod (z - y)" + using norm_triangle_ineq2 [of y z] by (simp add: norm_minus_commute) + have "cmod (f z / (z - y)) = cmod (f z) * inverse (cmod (z - y))" + by (simp add: norm_mult norm_inverse Fields.field_class.field_divide_inverse) + also have "... \ D * (e / L / D)" + apply (rule mult_mono) + using that D interior_subset apply blast + using \L>0\ \e>0\ \D>0\ DL2 + apply (auto simp: norm_divide divide_simps algebra_simps) + done + finally show ?thesis . + qed + have "dist (h y) 0 = cmod (contour_integral \ (\w. f w / (w - y)))" + by (simp add: dist_norm) + also have "... \ L * (D * (e / L / D))" + by (rule L [OF holint leD]) + also have "... = e" + using \L>0\ \0 < D\ by auto + finally show ?thesis . + qed + then have "(h \ 0) at_infinity" + by (meson Lim_at_infinityI) + moreover have "h holomorphic_on UNIV" + proof - + have con_ff: "continuous (at (x,z)) (\(x,y). (f y - f x) / (y - x))" + if "x \ u" "z \ u" "x \ z" for x z + using that conf + apply (simp add: split_def continuous_on_eq_continuous_at u) + apply (simp | rule continuous_intros continuous_within_compose2 [where g=f])+ + done + have con_fstsnd: "continuous_on UNIV (\x. (fst x - snd x) ::complex)" + by (rule continuous_intros)+ + have open_uu_Id: "open (u \ u - Id)" + apply (rule open_Diff) + apply (simp add: open_Times u) + using continuous_closed_preimage_constant [OF con_fstsnd closed_UNIV, of 0] + apply (auto simp: Id_fstsnd_eq algebra_simps) + done + have con_derf: "continuous (at z) (deriv f)" if "z \ u" for z + apply (rule continuous_on_interior [of u]) + apply (simp add: holf holomorphic_deriv holomorphic_on_imp_continuous_on u) + by (simp add: interior_open that u) + have tendsto_f': "((\(x,y). if y = x then deriv f (x) + else (f (y) - f (x)) / (y - x)) \ deriv f x) + (at (x, x) within u \ u)" if "x \ u" for x + proof (rule Lim_withinI) + fix e::real assume "0 < e" + obtain k1 where "k1>0" and k1: "\x'. norm (x' - x) \ k1 \ norm (deriv f x' - deriv f x) < e" + using \0 < e\ continuous_within_E [OF con_derf [OF \x \ u\]] + by (metis UNIV_I dist_norm) + obtain k2 where "k2>0" and k2: "ball x k2 \ u" by (blast intro: openE [OF u] \x \ u\) + have neq: "norm ((f z' - f x') / (z' - x') - deriv f x) \ e" + if "z' \ x'" and less_k1: "norm (x'-x, z'-x) < k1" and less_k2: "norm (x'-x, z'-x) < k2" + for x' z' + proof - + have cs_less: "w \ closed_segment x' z' \ cmod (w - x) \ norm (x'-x, z'-x)" for w + apply (drule segment_furthest_le [where y=x]) + by (metis (no_types) dist_commute dist_norm norm_fst_le norm_snd_le order_trans) + have derf_le: "w \ closed_segment x' z' \ z' \ x' \ cmod (deriv f w - deriv f x) \ e" for w + by (blast intro: cs_less less_k1 k1 [unfolded divide_const_simps dist_norm] less_imp_le le_less_trans) + have f_has_der: "\x. x \ u \ (f has_field_derivative deriv f x) (at x within u)" + by (metis DERIV_deriv_iff_field_differentiable at_within_open holf holomorphic_on_def u) + have "closed_segment x' z' \ u" + by (rule order_trans [OF _ k2]) (simp add: cs_less le_less_trans [OF _ less_k2] dist_complex_def norm_minus_commute subset_iff) + then have cint_derf: "(deriv f has_contour_integral f z' - f x') (linepath x' z')" + using contour_integral_primitive [OF f_has_der valid_path_linepath] pasz by simp + then have *: "((\x. deriv f x / (z' - x')) has_contour_integral (f z' - f x') / (z' - x')) (linepath x' z')" + by (rule has_contour_integral_div) + have "norm ((f z' - f x') / (z' - x') - deriv f x) \ e/norm(z' - x') * norm(z' - x')" + apply (rule has_contour_integral_bound_linepath [OF has_contour_integral_diff [OF *]]) + using has_contour_integral_div [where c = "z' - x'", OF has_contour_integral_const_linepath [of "deriv f x" z' x']] + \e > 0\ \z' \ x'\ + apply (auto simp: norm_divide divide_simps derf_le) + done + also have "... \ e" using \0 < e\ by simp + finally show ?thesis . + qed + show "\d>0. \xa\u \ u. + 0 < dist xa (x, x) \ dist xa (x, x) < d \ + dist (case xa of (x, y) \ if y = x then deriv f x else (f y - f x) / (y - x)) (deriv f x) \ e" + apply (rule_tac x="min k1 k2" in exI) + using \k1>0\ \k2>0\ \e>0\ + apply (force simp: dist_norm neq intro: dual_order.strict_trans2 k1 less_imp_le norm_fst_le) + done + qed + have con_pa_f: "continuous_on (path_image \) f" + by (meson holf holomorphic_on_imp_continuous_on holomorphic_on_subset interior_subset subt t) + have le_B: "\t. t \ {0..1} \ cmod (vector_derivative \ (at t)) \ B" + apply (rule B) + using \' using path_image_def vector_derivative_at by fastforce + have f_has_cint: "\w. w \ v - path_image \ \ ((\u. f u / (u - w) ^ 1) has_contour_integral h w) \" + by (simp add: V) + have cond_uu: "continuous_on (u \ u) (\(x,y). d x y)" + apply (simp add: continuous_on_eq_continuous_within d_def continuous_within tendsto_f') + apply (simp add: tendsto_within_open_NO_MATCH open_Times u, clarify) + apply (rule Lim_transform_within_open [OF _ open_uu_Id, where f = "(\(x,y). (f y - f x) / (y - x))"]) + using con_ff + apply (auto simp: continuous_within) + done + have hol_dw: "(\z. d z w) holomorphic_on u" if "w \ u" for w + proof - + have "continuous_on u ((\(x,y). d x y) o (\z. (w,z)))" + by (rule continuous_on_compose continuous_intros continuous_on_subset [OF cond_uu] | force intro: that)+ + then have *: "continuous_on u (\z. if w = z then deriv f z else (f w - f z) / (w - z))" + by (rule rev_iffD1 [OF _ continuous_on_cong [OF refl]]) (simp add: d_def field_simps) + have **: "\x. \x \ u; x \ w\ \ (\z. if w = z then deriv f z else (f w - f z) / (w - z)) field_differentiable at x" + apply (rule_tac f = "\x. (f w - f x)/(w - x)" and d = "dist x w" in field_differentiable_transform_within) + apply (rule u derivative_intros holomorphic_on_imp_differentiable_at [OF holf] | force simp add: dist_commute)+ + done + show ?thesis + unfolding d_def + apply (rule no_isolated_singularity [OF * _ u, where k = "{w}"]) + apply (auto simp: field_differentiable_def [symmetric] holomorphic_on_open open_Diff u **) + done + qed + { fix a b + assume abu: "closed_segment a b \ u" + then have "\w. w \ u \ (\z. d z w) contour_integrable_on (linepath a b)" + by (metis hol_dw continuous_on_subset contour_integrable_continuous_linepath holomorphic_on_imp_continuous_on) + then have cont_cint_d: "continuous_on u (\w. contour_integral (linepath a b) (\z. d z w))" + apply (rule contour_integral_continuous_on_linepath_2D [OF \open u\ _ _ abu]) + apply (auto simp: intro: continuous_on_swap_args cond_uu) + done + have cont_cint_d\: "continuous_on {0..1} ((\w. contour_integral (linepath a b) (\z. d z w)) o \)" + apply (rule continuous_on_compose) + using \path \\ path_def pasz + apply (auto intro!: continuous_on_subset [OF cont_cint_d]) + apply (force simp add: path_image_def) + done + have cint_cint: "(\w. contour_integral (linepath a b) (\z. d z w)) contour_integrable_on \" + apply (simp add: contour_integrable_on) + apply (rule integrable_continuous_real) + apply (rule continuous_on_mult [OF cont_cint_d\ [unfolded o_def]]) + using pf\' + by (simp add: continuous_on_polymonial_function vector_derivative_at [OF \']) + have "contour_integral (linepath a b) h = contour_integral (linepath a b) (\z. contour_integral \ (d z))" + using abu by (force simp add: h_def intro: contour_integral_eq) + also have "... = contour_integral \ (\w. contour_integral (linepath a b) (\z. d z w))" + apply (rule contour_integral_swap) + apply (rule continuous_on_subset [OF cond_uu]) + using abu pasz \valid_path \\ + apply (auto intro!: continuous_intros) + by (metis \' continuous_on_eq path_def path_polynomial_function pf\' vector_derivative_at) + finally have cint_h_eq: + "contour_integral (linepath a b) h = + contour_integral \ (\w. contour_integral (linepath a b) (\z. d z w))" . + note cint_cint cint_h_eq + } note cint_h = this + have conthu: "continuous_on u h" + proof (simp add: continuous_on_sequentially, clarify) + fix a x + assume x: "x \ u" and au: "\n. a n \ u" and ax: "a \ x" + then have A1: "\\<^sub>F n in sequentially. d (a n) contour_integrable_on \" + by (meson U contour_integrable_on_def eventuallyI) + obtain dd where "dd>0" and dd: "cball x dd \ u" using open_contains_cball u x by force + have A2: "\\<^sub>F n in sequentially. \xa\path_image \. cmod (d (a n) xa - d x xa) < ee" if "0 < ee" for ee + proof - + let ?ddpa = "{(w,z) |w z. w \ cball x dd \ z \ path_image \}" + have "uniformly_continuous_on ?ddpa (\(x,y). d x y)" + apply (rule compact_uniformly_continuous [OF continuous_on_subset[OF cond_uu]]) + using dd pasz \valid_path \\ + apply (auto simp: compact_Times compact_valid_path_image simp del: mem_cball) + done + then obtain kk where "kk>0" + and kk: "\x x'. \x\?ddpa; x'\?ddpa; dist x' x < kk\ \ + dist ((\(x,y). d x y) x') ((\(x,y). d x y) x) < ee" + apply (rule uniformly_continuous_onE [where e = ee]) + using \0 < ee\ by auto + have kk: "\norm (w - x) \ dd; z \ path_image \; norm ((w, z) - (x, z)) < kk\ \ norm (d w z - d x z) < ee" + for w z + using \dd>0\ kk [of "(x,z)" "(w,z)"] by (force simp add: norm_minus_commute dist_norm) + show ?thesis + using ax unfolding lim_sequentially eventually_sequentially + apply (drule_tac x="min dd kk" in spec) + using \dd > 0\ \kk > 0\ + apply (fastforce simp: kk dist_norm) + done + qed + have tendsto_hx: "(\n. contour_integral \ (d (a n))) \ h x" + apply (simp add: contour_integral_unique [OF U, symmetric] x) + apply (rule contour_integral_uniform_limit [OF A1 A2 le_B]) + apply (auto simp: \valid_path \\) + done + then show "(h \ a) \ h x" + by (simp add: h_def x au o_def) + qed + show ?thesis + proof (simp add: holomorphic_on_open field_differentiable_def [symmetric], clarify) + fix z0 + consider "z0 \ v" | "z0 \ u" using uv_Un by blast + then show "h field_differentiable at z0" + proof cases + assume "z0 \ v" then show ?thesis + using Cauchy_next_derivative [OF con_pa_f le_B f_has_cint _ ov] V f_has_cint \valid_path \\ + by (auto simp: field_differentiable_def v_def) + next + assume "z0 \ u" then + obtain e where "e>0" and e: "ball z0 e \ u" by (blast intro: openE [OF u]) + have *: "contour_integral (linepath a b) h + contour_integral (linepath b c) h + contour_integral (linepath c a) h = 0" + if abc_subset: "convex hull {a, b, c} \ ball z0 e" for a b c + proof - + have *: "\x1 x2 z. z \ u \ closed_segment x1 x2 \ u \ (\w. d w z) contour_integrable_on linepath x1 x2" + using hol_dw holomorphic_on_imp_continuous_on u + by (auto intro!: contour_integrable_holomorphic_simple) + have abc: "closed_segment a b \ u" "closed_segment b c \ u" "closed_segment c a \ u" + using that e segments_subset_convex_hull by fastforce+ + have eq0: "\w. w \ u \ contour_integral (linepath a b +++ linepath b c +++ linepath c a) (\z. d z w) = 0" + apply (rule contour_integral_unique [OF Cauchy_theorem_triangle]) + apply (rule holomorphic_on_subset [OF hol_dw]) + using e abc_subset by auto + have "contour_integral \ + (\x. contour_integral (linepath a b) (\z. d z x) + + (contour_integral (linepath b c) (\z. d z x) + + contour_integral (linepath c a) (\z. d z x))) = 0" + apply (rule contour_integral_eq_0) + using abc pasz u + apply (subst contour_integral_join [symmetric], auto intro: eq0 *)+ + done + then show ?thesis + by (simp add: cint_h abc contour_integrable_add contour_integral_add [symmetric] add_ac) + qed + show ?thesis + using e \e > 0\ + by (auto intro!: holomorphic_on_imp_differentiable_at [OF _ open_ball] analytic_imp_holomorphic + Morera_triangle continuous_on_subset [OF conthu] *) + qed + qed + qed + ultimately have [simp]: "h z = 0" for z + by (meson Liouville_weak) + have "((\w. 1 / (w - z)) has_contour_integral complex_of_real (2 * pi) * \ * winding_number \ z) \" + by (rule has_contour_integral_winding_number [OF \valid_path \\ znot]) + then have "((\w. f z * (1 / (w - z))) has_contour_integral complex_of_real (2 * pi) * \ * winding_number \ z * f z) \" + by (metis mult.commute has_contour_integral_lmul) + then have 1: "((\w. f z / (w - z)) has_contour_integral complex_of_real (2 * pi) * \ * winding_number \ z * f z) \" + by (simp add: divide_simps) + moreover have 2: "((\w. (f w - f z) / (w - z)) has_contour_integral 0) \" + using U [OF z] pasz d_def by (force elim: has_contour_integral_eq [where g = "\w. (f w - f z)/(w - z)"]) + show ?thesis + using has_contour_integral_add [OF 1 2] by (simp add: diff_divide_distrib) +qed + + +theorem Cauchy_integral_formula_global: + assumes s: "open s" and holf: "f holomorphic_on s" + and z: "z \ s" and vpg: "valid_path \" + and pasz: "path_image \ \ s - {z}" and loop: "pathfinish \ = pathstart \" + and zero: "\w. w \ s \ winding_number \ w = 0" + shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" +proof - + have "path \" using vpg by (blast intro: valid_path_imp_path) + have hols: "(\w. f w / (w - z)) holomorphic_on s - {z}" "(\w. 1 / (w - z)) holomorphic_on s - {z}" + by (rule holomorphic_intros holomorphic_on_subset [OF holf] | force)+ + then have cint_fw: "(\w. f w / (w - z)) contour_integrable_on \" + by (meson contour_integrable_holomorphic_simple holomorphic_on_imp_continuous_on open_delete s vpg pasz) + obtain d where "d>0" + and d: "\g h. \valid_path g; valid_path h; \t\{0..1}. cmod (g t - \ t) < d \ cmod (h t - \ t) < d; + pathstart h = pathstart g \ pathfinish h = pathfinish g\ + \ path_image h \ s - {z} \ (\f. f holomorphic_on s - {z} \ contour_integral h f = contour_integral g f)" + using contour_integral_nearby_ends [OF _ \path \\ pasz] s by (simp add: open_Diff) metis + obtain p where polyp: "polynomial_function p" + and ps: "pathstart p = pathstart \" and pf: "pathfinish p = pathfinish \" and led: "\t\{0..1}. cmod (p t - \ t) < d" + using path_approx_polynomial_function [OF \path \\ \d > 0\] by blast + then have ploop: "pathfinish p = pathstart p" using loop by auto + have vpp: "valid_path p" using polyp valid_path_polynomial_function by blast + have [simp]: "z \ path_image \" using pasz by blast + have paps: "path_image p \ s - {z}" and cint_eq: "(\f. f holomorphic_on s - {z} \ contour_integral p f = contour_integral \ f)" + using pf ps led d [OF vpg vpp] \d > 0\ by auto + have wn_eq: "winding_number p z = winding_number \ z" + using vpp paps + by (simp add: subset_Diff_insert vpg valid_path_polynomial_function winding_number_valid_path cint_eq hols) + have "winding_number p w = winding_number \ w" if "w \ s" for w + proof - + have hol: "(\v. 1 / (v - w)) holomorphic_on s - {z}" + using that by (force intro: holomorphic_intros holomorphic_on_subset [OF holf]) + have "w \ path_image p" "w \ path_image \" using paps pasz that by auto + then show ?thesis + using vpp vpg by (simp add: subset_Diff_insert valid_path_polynomial_function winding_number_valid_path cint_eq [OF hol]) + qed + then have wn0: "\w. w \ s \ winding_number p w = 0" + by (simp add: zero) + show ?thesis + using Cauchy_integral_formula_global_weak [OF s holf z polyp paps ploop wn0] hols + by (metis wn_eq cint_eq has_contour_integral_eqpath cint_fw cint_eq) +qed + +theorem Cauchy_theorem_global: + assumes s: "open s" and holf: "f holomorphic_on s" + and vpg: "valid_path \" and loop: "pathfinish \ = pathstart \" + and pas: "path_image \ \ s" + and zero: "\w. w \ s \ winding_number \ w = 0" + shows "(f has_contour_integral 0) \" +proof - + obtain z where "z \ s" and znot: "z \ path_image \" + proof - + have "compact (path_image \)" + using compact_valid_path_image vpg by blast + then have "path_image \ \ s" + by (metis (no_types) compact_open path_image_nonempty s) + with pas show ?thesis by (blast intro: that) + qed + then have pasz: "path_image \ \ s - {z}" using pas by blast + have hol: "(\w. (w - z) * f w) holomorphic_on s" + by (rule holomorphic_intros holf)+ + show ?thesis + using Cauchy_integral_formula_global [OF s hol \z \ s\ vpg pasz loop zero] + by (auto simp: znot elim!: has_contour_integral_eq) +qed + +corollary Cauchy_theorem_global_outside: + assumes "open s" "f holomorphic_on s" "valid_path \" "pathfinish \ = pathstart \" "path_image \ \ s" + "\w. w \ s \ w \ outside(path_image \)" + shows "(f has_contour_integral 0) \" +by (metis Cauchy_theorem_global assms winding_number_zero_in_outside valid_path_imp_path) + + +end diff -r bbcb05504fdc -r bd218a9320b5 src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy --- a/src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy Thu Aug 04 18:45:28 2016 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,7539 +0,0 @@ -section \Complex path integrals and Cauchy's integral theorem\ - -text\By John Harrison et al. Ported from HOL Light by L C Paulson (2015)\ - -theory Cauchy_Integral_Thm -imports Complex_Transcendental Weierstrass Ordered_Euclidean_Space -begin - -subsection\Homeomorphisms of arc images\ - -lemma homeomorphism_arc: - fixes g :: "real \ 'a::t2_space" - assumes "arc g" - obtains h where "homeomorphism {0..1} (path_image g) g h" -using assms by (force simp add: arc_def homeomorphism_compact path_def path_image_def) - -lemma homeomorphic_arc_image_interval: - fixes g :: "real \ 'a::t2_space" and a::real - assumes "arc g" "a < b" - shows "(path_image g) homeomorphic {a..b}" -proof - - have "(path_image g) homeomorphic {0..1::real}" - by (meson assms(1) homeomorphic_def homeomorphic_sym homeomorphism_arc) - also have "... homeomorphic {a..b}" - using assms by (force intro: homeomorphic_closed_intervals_real) - finally show ?thesis . -qed - -lemma homeomorphic_arc_images: - fixes g :: "real \ 'a::t2_space" and h :: "real \ 'b::t2_space" - assumes "arc g" "arc h" - shows "(path_image g) homeomorphic (path_image h)" -proof - - have "(path_image g) homeomorphic {0..1::real}" - by (meson assms homeomorphic_def homeomorphic_sym homeomorphism_arc) - also have "... homeomorphic (path_image h)" - by (meson assms homeomorphic_def homeomorphism_arc) - finally show ?thesis . -qed - -subsection \Piecewise differentiable functions\ - -definition piecewise_differentiable_on - (infixr "piecewise'_differentiable'_on" 50) - where "f piecewise_differentiable_on i \ - continuous_on i f \ - (\s. finite s \ (\x \ i - s. f differentiable (at x within i)))" - -lemma piecewise_differentiable_on_imp_continuous_on: - "f piecewise_differentiable_on s \ continuous_on s f" -by (simp add: piecewise_differentiable_on_def) - -lemma piecewise_differentiable_on_subset: - "f piecewise_differentiable_on s \ t \ s \ f piecewise_differentiable_on t" - using continuous_on_subset - unfolding piecewise_differentiable_on_def - apply safe - apply (blast intro: elim: continuous_on_subset) - by (meson Diff_iff differentiable_within_subset subsetCE) - -lemma differentiable_on_imp_piecewise_differentiable: - fixes a:: "'a::{linorder_topology,real_normed_vector}" - shows "f differentiable_on {a..b} \ f piecewise_differentiable_on {a..b}" - apply (simp add: piecewise_differentiable_on_def differentiable_imp_continuous_on) - apply (rule_tac x="{a,b}" in exI, simp add: differentiable_on_def) - done - -lemma differentiable_imp_piecewise_differentiable: - "(\x. x \ s \ f differentiable (at x within s)) - \ f piecewise_differentiable_on s" -by (auto simp: piecewise_differentiable_on_def differentiable_imp_continuous_on differentiable_on_def - intro: differentiable_within_subset) - -lemma piecewise_differentiable_const [iff]: "(\x. z) piecewise_differentiable_on s" - by (simp add: differentiable_imp_piecewise_differentiable) - -lemma piecewise_differentiable_compose: - "\f piecewise_differentiable_on s; g piecewise_differentiable_on (f ` s); - \x. finite (s \ f-`{x})\ - \ (g o f) piecewise_differentiable_on s" - apply (simp add: piecewise_differentiable_on_def, safe) - apply (blast intro: continuous_on_compose2) - apply (rename_tac A B) - apply (rule_tac x="A \ (\x\B. s \ f-`{x})" in exI) - apply (blast intro: differentiable_chain_within) - done - -lemma piecewise_differentiable_affine: - fixes m::real - assumes "f piecewise_differentiable_on ((\x. m *\<^sub>R x + c) ` s)" - shows "(f o (\x. m *\<^sub>R x + c)) piecewise_differentiable_on s" -proof (cases "m = 0") - case True - then show ?thesis - unfolding o_def - by (force intro: differentiable_imp_piecewise_differentiable differentiable_const) -next - case False - show ?thesis - apply (rule piecewise_differentiable_compose [OF differentiable_imp_piecewise_differentiable]) - apply (rule assms derivative_intros | simp add: False vimage_def real_vector_affinity_eq)+ - done -qed - -lemma piecewise_differentiable_cases: - fixes c::real - assumes "f piecewise_differentiable_on {a..c}" - "g piecewise_differentiable_on {c..b}" - "a \ c" "c \ b" "f c = g c" - shows "(\x. if x \ c then f x else g x) piecewise_differentiable_on {a..b}" -proof - - obtain s t where st: "finite s" "finite t" - "\x\{a..c} - s. f differentiable at x within {a..c}" - "\x\{c..b} - t. g differentiable at x within {c..b}" - using assms - by (auto simp: piecewise_differentiable_on_def) - have finabc: "finite ({a,b,c} \ (s \ t))" - by (metis \finite s\ \finite t\ finite_Un finite_insert finite.emptyI) - have "continuous_on {a..c} f" "continuous_on {c..b} g" - using assms piecewise_differentiable_on_def by auto - then have "continuous_on {a..b} (\x. if x \ c then f x else g x)" - using continuous_on_cases [OF closed_real_atLeastAtMost [of a c], - OF closed_real_atLeastAtMost [of c b], - of f g "\x. x\c"] assms - by (force simp: ivl_disj_un_two_touch) - moreover - { fix x - assume x: "x \ {a..b} - ({a,b,c} \ (s \ t))" - have "(\x. if x \ c then f x else g x) differentiable at x within {a..b}" (is "?diff_fg") - proof (cases x c rule: le_cases) - case le show ?diff_fg - apply (rule differentiable_transform_within [where d = "dist x c" and f = f]) - using x le st - apply (simp_all add: dist_real_def) - apply (rule differentiable_at_withinI) - apply (rule differentiable_within_open [where s = "{a<..s. finite s \ - (\x\{a..b} - s. (\x. if x \ c then f x else g x) differentiable at x within {a..b})" - by (meson finabc) - ultimately show ?thesis - by (simp add: piecewise_differentiable_on_def) -qed - -lemma piecewise_differentiable_neg: - "f piecewise_differentiable_on s \ (\x. -(f x)) piecewise_differentiable_on s" - by (auto simp: piecewise_differentiable_on_def continuous_on_minus) - -lemma piecewise_differentiable_add: - assumes "f piecewise_differentiable_on i" - "g piecewise_differentiable_on i" - shows "(\x. f x + g x) piecewise_differentiable_on i" -proof - - obtain s t where st: "finite s" "finite t" - "\x\i - s. f differentiable at x within i" - "\x\i - t. g differentiable at x within i" - using assms by (auto simp: piecewise_differentiable_on_def) - then have "finite (s \ t) \ (\x\i - (s \ t). (\x. f x + g x) differentiable at x within i)" - by auto - moreover have "continuous_on i f" "continuous_on i g" - using assms piecewise_differentiable_on_def by auto - ultimately show ?thesis - by (auto simp: piecewise_differentiable_on_def continuous_on_add) -qed - -lemma piecewise_differentiable_diff: - "\f piecewise_differentiable_on s; g piecewise_differentiable_on s\ - \ (\x. f x - g x) piecewise_differentiable_on s" - unfolding diff_conv_add_uminus - by (metis piecewise_differentiable_add piecewise_differentiable_neg) - -lemma continuous_on_joinpaths_D1: - "continuous_on {0..1} (g1 +++ g2) \ continuous_on {0..1} g1" - apply (rule continuous_on_eq [of _ "(g1 +++ g2) o (op*(inverse 2))"]) - apply (rule continuous_intros | simp)+ - apply (auto elim!: continuous_on_subset simp: joinpaths_def) - done - -lemma continuous_on_joinpaths_D2: - "\continuous_on {0..1} (g1 +++ g2); pathfinish g1 = pathstart g2\ \ continuous_on {0..1} g2" - apply (rule continuous_on_eq [of _ "(g1 +++ g2) o (\x. inverse 2*x + 1/2)"]) - apply (rule continuous_intros | simp)+ - apply (auto elim!: continuous_on_subset simp add: joinpaths_def pathfinish_def pathstart_def Ball_def) - done - -lemma piecewise_differentiable_D1: - "(g1 +++ g2) piecewise_differentiable_on {0..1} \ g1 piecewise_differentiable_on {0..1}" - apply (clarsimp simp add: piecewise_differentiable_on_def dest!: continuous_on_joinpaths_D1) - apply (rule_tac x="insert 1 ((op*2)`s)" in exI) - apply simp - apply (intro ballI) - apply (rule_tac d="dist (x/2) (1/2)" and f = "(g1 +++ g2) o (op*(inverse 2))" - in differentiable_transform_within) - apply (auto simp: dist_real_def joinpaths_def) - apply (rule differentiable_chain_within derivative_intros | simp)+ - apply (rule differentiable_subset) - apply (force simp:)+ - done - -lemma piecewise_differentiable_D2: - "\(g1 +++ g2) piecewise_differentiable_on {0..1}; pathfinish g1 = pathstart g2\ - \ g2 piecewise_differentiable_on {0..1}" - apply (clarsimp simp add: piecewise_differentiable_on_def dest!: continuous_on_joinpaths_D2) - apply (rule_tac x="insert 0 ((\x. 2*x-1)`s)" in exI) - apply simp - apply (intro ballI) - apply (rule_tac d="dist ((x+1)/2) (1/2)" and f = "(g1 +++ g2) o (\x. (x+1)/2)" - in differentiable_transform_within) - apply (auto simp: dist_real_def joinpaths_def abs_if field_simps split: if_split_asm) - apply (rule differentiable_chain_within derivative_intros | simp)+ - apply (rule differentiable_subset) - apply (force simp: divide_simps)+ - done - - -subsubsection\The concept of continuously differentiable\ - -text \ -John Harrison writes as follows: - -``The usual assumption in complex analysis texts is that a path \\\ should be piecewise -continuously differentiable, which ensures that the path integral exists at least for any continuous -f, since all piecewise continuous functions are integrable. However, our notion of validity is -weaker, just piecewise differentiability... [namely] continuity plus differentiability except on a -finite set ... [Our] underlying theory of integration is the Kurzweil-Henstock theory. In contrast to -the Riemann or Lebesgue theory (but in common with a simple notion based on antiderivatives), this -can integrate all derivatives.'' - -"Formalizing basic complex analysis." From Insight to Proof: Festschrift in Honour of Andrzej Trybulec. -Studies in Logic, Grammar and Rhetoric 10.23 (2007): 151-165. - -And indeed he does not assume that his derivatives are continuous, but the penalty is unreasonably -difficult proofs concerning winding numbers. We need a self-contained and straightforward theorem -asserting that all derivatives can be integrated before we can adopt Harrison's choice.\ - -definition C1_differentiable_on :: "(real \ 'a::real_normed_vector) \ real set \ bool" - (infix "C1'_differentiable'_on" 50) - where - "f C1_differentiable_on s \ - (\D. (\x \ s. (f has_vector_derivative (D x)) (at x)) \ continuous_on s D)" - -lemma C1_differentiable_on_eq: - "f C1_differentiable_on s \ - (\x \ s. f differentiable at x) \ continuous_on s (\x. vector_derivative f (at x))" - unfolding C1_differentiable_on_def - apply safe - using differentiable_def has_vector_derivative_def apply blast - apply (erule continuous_on_eq) - using vector_derivative_at apply fastforce - using vector_derivative_works apply fastforce - done - -lemma C1_differentiable_on_subset: - "f C1_differentiable_on t \ s \ t \ f C1_differentiable_on s" - unfolding C1_differentiable_on_def continuous_on_eq_continuous_within - by (blast intro: continuous_within_subset) - -lemma C1_differentiable_compose: - "\f C1_differentiable_on s; g C1_differentiable_on (f ` s); - \x. finite (s \ f-`{x})\ - \ (g o f) C1_differentiable_on s" - apply (simp add: C1_differentiable_on_eq, safe) - using differentiable_chain_at apply blast - apply (rule continuous_on_eq [of _ "\x. vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x))"]) - apply (rule Limits.continuous_on_scaleR, assumption) - apply (metis (mono_tags, lifting) continuous_on_eq continuous_at_imp_continuous_on continuous_on_compose differentiable_imp_continuous_within o_def) - by (simp add: vector_derivative_chain_at) - -lemma C1_diff_imp_diff: "f C1_differentiable_on s \ f differentiable_on s" - by (simp add: C1_differentiable_on_eq differentiable_at_imp_differentiable_on) - -lemma C1_differentiable_on_ident [simp, derivative_intros]: "(\x. x) C1_differentiable_on s" - by (auto simp: C1_differentiable_on_eq continuous_on_const) - -lemma C1_differentiable_on_const [simp, derivative_intros]: "(\z. a) C1_differentiable_on s" - by (auto simp: C1_differentiable_on_eq continuous_on_const) - -lemma C1_differentiable_on_add [simp, derivative_intros]: - "f C1_differentiable_on s \ g C1_differentiable_on s \ (\x. f x + g x) C1_differentiable_on s" - unfolding C1_differentiable_on_eq by (auto intro: continuous_intros) - -lemma C1_differentiable_on_minus [simp, derivative_intros]: - "f C1_differentiable_on s \ (\x. - f x) C1_differentiable_on s" - unfolding C1_differentiable_on_eq by (auto intro: continuous_intros) - -lemma C1_differentiable_on_diff [simp, derivative_intros]: - "f C1_differentiable_on s \ g C1_differentiable_on s \ (\x. f x - g x) C1_differentiable_on s" - unfolding C1_differentiable_on_eq by (auto intro: continuous_intros) - -lemma C1_differentiable_on_mult [simp, derivative_intros]: - fixes f g :: "real \ 'a :: real_normed_algebra" - shows "f C1_differentiable_on s \ g C1_differentiable_on s \ (\x. f x * g x) C1_differentiable_on s" - unfolding C1_differentiable_on_eq - by (auto simp: continuous_on_add continuous_on_mult continuous_at_imp_continuous_on differentiable_imp_continuous_within) - -lemma C1_differentiable_on_scaleR [simp, derivative_intros]: - "f C1_differentiable_on s \ g C1_differentiable_on s \ (\x. f x *\<^sub>R g x) C1_differentiable_on s" - unfolding C1_differentiable_on_eq - by (rule continuous_intros | simp add: continuous_at_imp_continuous_on differentiable_imp_continuous_within)+ - - -definition piecewise_C1_differentiable_on - (infixr "piecewise'_C1'_differentiable'_on" 50) - where "f piecewise_C1_differentiable_on i \ - continuous_on i f \ - (\s. finite s \ (f C1_differentiable_on (i - s)))" - -lemma C1_differentiable_imp_piecewise: - "f C1_differentiable_on s \ f piecewise_C1_differentiable_on s" - by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq continuous_at_imp_continuous_on differentiable_imp_continuous_within) - -lemma piecewise_C1_imp_differentiable: - "f piecewise_C1_differentiable_on i \ f piecewise_differentiable_on i" - by (auto simp: piecewise_C1_differentiable_on_def piecewise_differentiable_on_def - C1_differentiable_on_def differentiable_def has_vector_derivative_def - intro: has_derivative_at_within) - -lemma piecewise_C1_differentiable_compose: - "\f piecewise_C1_differentiable_on s; g piecewise_C1_differentiable_on (f ` s); - \x. finite (s \ f-`{x})\ - \ (g o f) piecewise_C1_differentiable_on s" - apply (simp add: piecewise_C1_differentiable_on_def, safe) - apply (blast intro: continuous_on_compose2) - apply (rename_tac A B) - apply (rule_tac x="A \ (\x\B. s \ f-`{x})" in exI) - apply (rule conjI, blast) - apply (rule C1_differentiable_compose) - apply (blast intro: C1_differentiable_on_subset) - apply (blast intro: C1_differentiable_on_subset) - by (simp add: Diff_Int_distrib2) - -lemma piecewise_C1_differentiable_on_subset: - "f piecewise_C1_differentiable_on s \ t \ s \ f piecewise_C1_differentiable_on t" - by (auto simp: piecewise_C1_differentiable_on_def elim!: continuous_on_subset C1_differentiable_on_subset) - -lemma C1_differentiable_imp_continuous_on: - "f C1_differentiable_on s \ continuous_on s f" - unfolding C1_differentiable_on_eq continuous_on_eq_continuous_within - using differentiable_at_withinI differentiable_imp_continuous_within by blast - -lemma C1_differentiable_on_empty [iff]: "f C1_differentiable_on {}" - unfolding C1_differentiable_on_def - by auto - -lemma piecewise_C1_differentiable_affine: - fixes m::real - assumes "f piecewise_C1_differentiable_on ((\x. m * x + c) ` s)" - shows "(f o (\x. m *\<^sub>R x + c)) piecewise_C1_differentiable_on s" -proof (cases "m = 0") - case True - then show ?thesis - unfolding o_def by (auto simp: piecewise_C1_differentiable_on_def continuous_on_const) -next - case False - show ?thesis - apply (rule piecewise_C1_differentiable_compose [OF C1_differentiable_imp_piecewise]) - apply (rule assms derivative_intros | simp add: False vimage_def)+ - using real_vector_affinity_eq [OF False, where c=c, unfolded scaleR_conv_of_real] - apply simp - done -qed - -lemma piecewise_C1_differentiable_cases: - fixes c::real - assumes "f piecewise_C1_differentiable_on {a..c}" - "g piecewise_C1_differentiable_on {c..b}" - "a \ c" "c \ b" "f c = g c" - shows "(\x. if x \ c then f x else g x) piecewise_C1_differentiable_on {a..b}" -proof - - obtain s t where st: "f C1_differentiable_on ({a..c} - s)" - "g C1_differentiable_on ({c..b} - t)" - "finite s" "finite t" - using assms - by (force simp: piecewise_C1_differentiable_on_def) - then have f_diff: "f differentiable_on {a..x. if x \ c then f x else g x)" - using continuous_on_cases [OF closed_real_atLeastAtMost [of a c], - OF closed_real_atLeastAtMost [of c b], - of f g "\x. x\c"] assms - by (force simp: ivl_disj_un_two_touch) - { fix x - assume x: "x \ {a..b} - insert c (s \ t)" - have "(\x. if x \ c then f x else g x) differentiable at x" (is "?diff_fg") - proof (cases x c rule: le_cases) - case le show ?diff_fg - apply (rule differentiable_transform_within [where f=f and d = "dist x c"]) - using x dist_real_def le st by (auto simp: C1_differentiable_on_eq) - next - case ge show ?diff_fg - apply (rule differentiable_transform_within [where f=g and d = "dist x c"]) - using dist_nz x dist_real_def ge st x by (auto simp: C1_differentiable_on_eq) - qed - } - then have "(\x \ {a..b} - insert c (s \ t). (\x. if x \ c then f x else g x) differentiable at x)" - by auto - moreover - { assume fcon: "continuous_on ({a<..x. vector_derivative f (at x))" - and gcon: "continuous_on ({c<..x. vector_derivative g (at x))" - have "open ({a<..x. vector_derivative (\x. if x \ c then f x else g x) (at x))" - apply (rule continuous_on_eq [OF fcon]) - apply (simp add:) - apply (rule vector_derivative_at [symmetric]) - apply (rule_tac f=f and d="dist x c" in has_vector_derivative_transform_within) - apply (simp_all add: dist_norm vector_derivative_works [symmetric]) - apply (metis (full_types) C1_differentiable_on_eq Diff_iff Groups.add_ac(2) add_mono_thms_linordered_field(5) atLeastAtMost_iff linorder_not_le order_less_irrefl st(1)) - apply auto - done - moreover have "continuous_on ({c<..x. vector_derivative (\x. if x \ c then f x else g x) (at x))" - apply (rule continuous_on_eq [OF gcon]) - apply (simp add:) - apply (rule vector_derivative_at [symmetric]) - apply (rule_tac f=g and d="dist x c" in has_vector_derivative_transform_within) - apply (simp_all add: dist_norm vector_derivative_works [symmetric]) - apply (metis (full_types) C1_differentiable_on_eq Diff_iff Groups.add_ac(2) add_mono_thms_linordered_field(5) atLeastAtMost_iff less_irrefl not_le st(2)) - apply auto - done - ultimately have "continuous_on ({a<.. t)) - (\x. vector_derivative (\x. if x \ c then f x else g x) (at x))" - apply (rule continuous_on_subset [OF continuous_on_open_Un], auto) - done - } note * = this - have "continuous_on ({a<.. t)) (\x. vector_derivative (\x. if x \ c then f x else g x) (at x))" - using st - by (auto simp: C1_differentiable_on_eq elim!: continuous_on_subset intro: *) - ultimately have "\s. finite s \ ((\x. if x \ c then f x else g x) C1_differentiable_on {a..b} - s)" - apply (rule_tac x="{a,b,c} \ s \ t" in exI) - using st by (auto simp: C1_differentiable_on_eq elim!: continuous_on_subset) - with cab show ?thesis - by (simp add: piecewise_C1_differentiable_on_def) -qed - -lemma piecewise_C1_differentiable_neg: - "f piecewise_C1_differentiable_on s \ (\x. -(f x)) piecewise_C1_differentiable_on s" - unfolding piecewise_C1_differentiable_on_def - by (auto intro!: continuous_on_minus C1_differentiable_on_minus) - -lemma piecewise_C1_differentiable_add: - assumes "f piecewise_C1_differentiable_on i" - "g piecewise_C1_differentiable_on i" - shows "(\x. f x + g x) piecewise_C1_differentiable_on i" -proof - - obtain s t where st: "finite s" "finite t" - "f C1_differentiable_on (i-s)" - "g C1_differentiable_on (i-t)" - using assms by (auto simp: piecewise_C1_differentiable_on_def) - then have "finite (s \ t) \ (\x. f x + g x) C1_differentiable_on i - (s \ t)" - by (auto intro: C1_differentiable_on_add elim!: C1_differentiable_on_subset) - moreover have "continuous_on i f" "continuous_on i g" - using assms piecewise_C1_differentiable_on_def by auto - ultimately show ?thesis - by (auto simp: piecewise_C1_differentiable_on_def continuous_on_add) -qed - -lemma piecewise_C1_differentiable_diff: - "\f piecewise_C1_differentiable_on s; g piecewise_C1_differentiable_on s\ - \ (\x. f x - g x) piecewise_C1_differentiable_on s" - unfolding diff_conv_add_uminus - by (metis piecewise_C1_differentiable_add piecewise_C1_differentiable_neg) - -lemma piecewise_C1_differentiable_D1: - fixes g1 :: "real \ 'a::real_normed_field" - assumes "(g1 +++ g2) piecewise_C1_differentiable_on {0..1}" - shows "g1 piecewise_C1_differentiable_on {0..1}" -proof - - obtain s where "finite s" - and co12: "continuous_on ({0..1} - s) (\x. vector_derivative (g1 +++ g2) (at x))" - and g12D: "\x\{0..1} - s. g1 +++ g2 differentiable at x" - using assms by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - then have g1D: "g1 differentiable at x" if "x \ {0..1} - insert 1 (op * 2 ` s)" for x - apply (rule_tac d="dist (x/2) (1/2)" and f = "(g1 +++ g2) o (op*(inverse 2))" in differentiable_transform_within) - using that - apply (simp_all add: dist_real_def joinpaths_def) - apply (rule differentiable_chain_at derivative_intros | force)+ - done - have [simp]: "vector_derivative (g1 \ op * 2) (at (x/2)) = 2 *\<^sub>R vector_derivative g1 (at x)" - if "x \ {0..1} - insert 1 (op * 2 ` s)" for x - apply (subst vector_derivative_chain_at) - using that - apply (rule derivative_eq_intros g1D | simp)+ - done - have "continuous_on ({0..1/2} - insert (1/2) s) (\x. vector_derivative (g1 +++ g2) (at x))" - using co12 by (rule continuous_on_subset) force - then have coDhalf: "continuous_on ({0..1/2} - insert (1/2) s) (\x. vector_derivative (g1 o op*2) (at x))" - apply (rule continuous_on_eq [OF _ vector_derivative_at]) - apply (rule_tac f="g1 o op*2" and d="dist x (1/2)" in has_vector_derivative_transform_within) - apply (simp_all add: dist_norm joinpaths_def vector_derivative_works [symmetric]) - apply (force intro: g1D differentiable_chain_at) - apply auto - done - have "continuous_on ({0..1} - insert 1 (op * 2 ` s)) - ((\x. 1/2 * vector_derivative (g1 o op*2) (at x)) o op*(1/2))" - apply (rule continuous_intros)+ - using coDhalf - apply (simp add: scaleR_conv_of_real image_set_diff image_image) - done - then have con_g1: "continuous_on ({0..1} - insert 1 (op * 2 ` s)) (\x. vector_derivative g1 (at x))" - by (rule continuous_on_eq) (simp add: scaleR_conv_of_real) - have "continuous_on {0..1} g1" - using continuous_on_joinpaths_D1 assms piecewise_C1_differentiable_on_def by blast - with \finite s\ show ?thesis - apply (clarsimp simp add: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - apply (rule_tac x="insert 1 ((op*2)`s)" in exI) - apply (simp add: g1D con_g1) - done -qed - -lemma piecewise_C1_differentiable_D2: - fixes g2 :: "real \ 'a::real_normed_field" - assumes "(g1 +++ g2) piecewise_C1_differentiable_on {0..1}" "pathfinish g1 = pathstart g2" - shows "g2 piecewise_C1_differentiable_on {0..1}" -proof - - obtain s where "finite s" - and co12: "continuous_on ({0..1} - s) (\x. vector_derivative (g1 +++ g2) (at x))" - and g12D: "\x\{0..1} - s. g1 +++ g2 differentiable at x" - using assms by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - then have g2D: "g2 differentiable at x" if "x \ {0..1} - insert 0 ((\x. 2*x-1) ` s)" for x - apply (rule_tac d="dist ((x+1)/2) (1/2)" and f = "(g1 +++ g2) o (\x. (x+1)/2)" in differentiable_transform_within) - using that - apply (simp_all add: dist_real_def joinpaths_def) - apply (auto simp: dist_real_def joinpaths_def field_simps) - apply (rule differentiable_chain_at derivative_intros | force)+ - apply (drule_tac x= "(x + 1) / 2" in bspec, force simp: divide_simps) - apply assumption - done - have [simp]: "vector_derivative (g2 \ (\x. 2*x-1)) (at ((x+1)/2)) = 2 *\<^sub>R vector_derivative g2 (at x)" - if "x \ {0..1} - insert 0 ((\x. 2*x-1) ` s)" for x - using that by (auto simp: vector_derivative_chain_at divide_simps g2D) - have "continuous_on ({1/2..1} - insert (1/2) s) (\x. vector_derivative (g1 +++ g2) (at x))" - using co12 by (rule continuous_on_subset) force - then have coDhalf: "continuous_on ({1/2..1} - insert (1/2) s) (\x. vector_derivative (g2 o (\x. 2*x-1)) (at x))" - apply (rule continuous_on_eq [OF _ vector_derivative_at]) - apply (rule_tac f="g2 o (\x. 2*x-1)" and d="dist (3/4) ((x+1)/2)" in has_vector_derivative_transform_within) - apply (auto simp: dist_real_def field_simps joinpaths_def vector_derivative_works [symmetric] - intro!: g2D differentiable_chain_at) - done - have [simp]: "((\x. (x + 1) / 2) ` ({0..1} - insert 0 ((\x. 2 * x - 1) ` s))) = ({1/2..1} - insert (1/2) s)" - apply (simp add: image_set_diff inj_on_def image_image) - apply (auto simp: image_affinity_atLeastAtMost_div add_divide_distrib) - done - have "continuous_on ({0..1} - insert 0 ((\x. 2*x-1) ` s)) - ((\x. 1/2 * vector_derivative (g2 \ (\x. 2*x-1)) (at x)) o (\x. (x+1)/2))" - by (rule continuous_intros | simp add: coDhalf)+ - then have con_g2: "continuous_on ({0..1} - insert 0 ((\x. 2*x-1) ` s)) (\x. vector_derivative g2 (at x))" - by (rule continuous_on_eq) (simp add: scaleR_conv_of_real) - have "continuous_on {0..1} g2" - using continuous_on_joinpaths_D2 assms piecewise_C1_differentiable_on_def by blast - with \finite s\ show ?thesis - apply (clarsimp simp add: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - apply (rule_tac x="insert 0 ((\x. 2 * x - 1) ` s)" in exI) - apply (simp add: g2D con_g2) - done -qed - -subsection \Valid paths, and their start and finish\ - -lemma Diff_Un_eq: "A - (B \ C) = A - B - C" - by blast - -definition valid_path :: "(real \ 'a :: real_normed_vector) \ bool" - where "valid_path f \ f piecewise_C1_differentiable_on {0..1::real}" - -definition closed_path :: "(real \ 'a :: real_normed_vector) \ bool" - where "closed_path g \ g 0 = g 1" - -subsubsection\In particular, all results for paths apply\ - -lemma valid_path_imp_path: "valid_path g \ path g" -by (simp add: path_def piecewise_C1_differentiable_on_def valid_path_def) - -lemma connected_valid_path_image: "valid_path g \ connected(path_image g)" - by (metis connected_path_image valid_path_imp_path) - -lemma compact_valid_path_image: "valid_path g \ compact(path_image g)" - by (metis compact_path_image valid_path_imp_path) - -lemma bounded_valid_path_image: "valid_path g \ bounded(path_image g)" - by (metis bounded_path_image valid_path_imp_path) - -lemma closed_valid_path_image: "valid_path g \ closed(path_image g)" - by (metis closed_path_image valid_path_imp_path) - -proposition valid_path_compose: - assumes "valid_path g" - and der: "\x. x \ path_image g \ \f'. (f has_field_derivative f') (at x)" - and con: "continuous_on (path_image g) (deriv f)" - shows "valid_path (f o g)" -proof - - obtain s where "finite s" and g_diff: "g C1_differentiable_on {0..1} - s" - using \valid_path g\ unfolding valid_path_def piecewise_C1_differentiable_on_def by auto - have "f \ g differentiable at t" when "t\{0..1} - s" for t - proof (rule differentiable_chain_at) - show "g differentiable at t" using \valid_path g\ - by (meson C1_differentiable_on_eq \g C1_differentiable_on {0..1} - s\ that) - next - have "g t\path_image g" using that DiffD1 image_eqI path_image_def by metis - then obtain f' where "(f has_field_derivative f') (at (g t))" - using der by auto - then have " (f has_derivative op * f') (at (g t))" - using has_field_derivative_imp_has_derivative[of f f' "at (g t)"] by auto - then show "f differentiable at (g t)" using differentiableI by auto - qed - moreover have "continuous_on ({0..1} - s) (\x. vector_derivative (f \ g) (at x))" - proof (rule continuous_on_eq [where f = "\x. vector_derivative g (at x) * deriv f (g x)"], - rule continuous_intros) - show "continuous_on ({0..1} - s) (\x. vector_derivative g (at x))" - using g_diff C1_differentiable_on_eq by auto - next - have "continuous_on {0..1} (\x. deriv f (g x))" - using continuous_on_compose[OF _ con[unfolded path_image_def],unfolded comp_def] - \valid_path g\ piecewise_C1_differentiable_on_def valid_path_def - by blast - then show "continuous_on ({0..1} - s) (\x. deriv f (g x))" - using continuous_on_subset by blast - next - show "vector_derivative g (at t) * deriv f (g t) = vector_derivative (f \ g) (at t)" - when "t \ {0..1} - s" for t - proof (rule vector_derivative_chain_at_general[symmetric]) - show "g differentiable at t" by (meson C1_differentiable_on_eq g_diff that) - next - have "g t\path_image g" using that DiffD1 image_eqI path_image_def by metis - then obtain f' where "(f has_field_derivative f') (at (g t))" - using der by auto - then show "\g'. (f has_field_derivative g') (at (g t))" by auto - qed - qed - ultimately have "f o g C1_differentiable_on {0..1} - s" - using C1_differentiable_on_eq by blast - moreover have "path (f o g)" - proof - - have "isCont f x" when "x\path_image g" for x - proof - - obtain f' where "(f has_field_derivative f') (at x)" - using der[rule_format] \x\path_image g\ by auto - thus ?thesis using DERIV_isCont by auto - qed - then have "continuous_on (path_image g) f" using continuous_at_imp_continuous_on by auto - then show ?thesis using path_continuous_image \valid_path g\ valid_path_imp_path by auto - qed - ultimately show ?thesis unfolding valid_path_def piecewise_C1_differentiable_on_def path_def - using \finite s\ by auto -qed - - -subsection\Contour Integrals along a path\ - -text\This definition is for complex numbers only, and does not generalise to line integrals in a vector field\ - -text\piecewise differentiable function on [0,1]\ - -definition has_contour_integral :: "(complex \ complex) \ complex \ (real \ complex) \ bool" - (infixr "has'_contour'_integral" 50) - where "(f has_contour_integral i) g \ - ((\x. f(g x) * vector_derivative g (at x within {0..1})) - has_integral i) {0..1}" - -definition contour_integrable_on - (infixr "contour'_integrable'_on" 50) - where "f contour_integrable_on g \ \i. (f has_contour_integral i) g" - -definition contour_integral - where "contour_integral g f \ @i. (f has_contour_integral i) g \ ~ f contour_integrable_on g \ i=0" - -lemma not_integrable_contour_integral: "~ f contour_integrable_on g \ contour_integral g f = 0" - unfolding contour_integrable_on_def contour_integral_def by blast - -lemma contour_integral_unique: "(f has_contour_integral i) g \ contour_integral g f = i" - apply (simp add: contour_integral_def has_contour_integral_def contour_integrable_on_def) - using has_integral_unique by blast - -corollary has_contour_integral_eqpath: - "\(f has_contour_integral y) p; f contour_integrable_on \; - contour_integral p f = contour_integral \ f\ - \ (f has_contour_integral y) \" -using contour_integrable_on_def contour_integral_unique by auto - -lemma has_contour_integral_integral: - "f contour_integrable_on i \ (f has_contour_integral (contour_integral i f)) i" - by (metis contour_integral_unique contour_integrable_on_def) - -lemma has_contour_integral_unique: - "(f has_contour_integral i) g \ (f has_contour_integral j) g \ i = j" - using has_integral_unique - by (auto simp: has_contour_integral_def) - -lemma has_contour_integral_integrable: "(f has_contour_integral i) g \ f contour_integrable_on g" - using contour_integrable_on_def by blast - -(* Show that we can forget about the localized derivative.*) - -lemma vector_derivative_within_interior: - "\x \ interior s; NO_MATCH UNIV s\ - \ vector_derivative f (at x within s) = vector_derivative f (at x)" - apply (simp add: vector_derivative_def has_vector_derivative_def has_derivative_def netlimit_within_interior) - apply (subst lim_within_interior, auto) - done - -lemma has_integral_localized_vector_derivative: - "((\x. f (g x) * vector_derivative g (at x within {a..b})) has_integral i) {a..b} \ - ((\x. f (g x) * vector_derivative g (at x)) has_integral i) {a..b}" -proof - - have "{a..b} - {a,b} = interior {a..b}" - by (simp add: atLeastAtMost_diff_ends) - show ?thesis - apply (rule has_integral_spike_eq [of "{a,b}"]) - apply (auto simp: vector_derivative_within_interior) - done -qed - -lemma integrable_on_localized_vector_derivative: - "(\x. f (g x) * vector_derivative g (at x within {a..b})) integrable_on {a..b} \ - (\x. f (g x) * vector_derivative g (at x)) integrable_on {a..b}" - by (simp add: integrable_on_def has_integral_localized_vector_derivative) - -lemma has_contour_integral: - "(f has_contour_integral i) g \ - ((\x. f (g x) * vector_derivative g (at x)) has_integral i) {0..1}" - by (simp add: has_integral_localized_vector_derivative has_contour_integral_def) - -lemma contour_integrable_on: - "f contour_integrable_on g \ - (\t. f(g t) * vector_derivative g (at t)) integrable_on {0..1}" - by (simp add: has_contour_integral integrable_on_def contour_integrable_on_def) - -subsection\Reversing a path\ - -lemma valid_path_imp_reverse: - assumes "valid_path g" - shows "valid_path(reversepath g)" -proof - - obtain s where "finite s" "g C1_differentiable_on ({0..1} - s)" - using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) - then have "finite (op - 1 ` s)" "(reversepath g C1_differentiable_on ({0..1} - op - 1 ` s))" - apply (auto simp: reversepath_def) - apply (rule C1_differentiable_compose [of "\x::real. 1-x" _ g, unfolded o_def]) - apply (auto simp: C1_differentiable_on_eq) - apply (rule continuous_intros, force) - apply (force elim!: continuous_on_subset) - apply (simp add: finite_vimageI inj_on_def) - done - then show ?thesis using assms - by (auto simp: valid_path_def piecewise_C1_differentiable_on_def path_def [symmetric]) -qed - -lemma valid_path_reversepath [simp]: "valid_path(reversepath g) \ valid_path g" - using valid_path_imp_reverse by force - -lemma has_contour_integral_reversepath: - assumes "valid_path g" "(f has_contour_integral i) g" - shows "(f has_contour_integral (-i)) (reversepath g)" -proof - - { fix s x - assume xs: "g C1_differentiable_on ({0..1} - s)" "x \ op - 1 ` s" "0 \ x" "x \ 1" - have "vector_derivative (\x. g (1 - x)) (at x within {0..1}) = - - vector_derivative g (at (1 - x) within {0..1})" - proof - - obtain f' where f': "(g has_vector_derivative f') (at (1 - x))" - using xs - by (force simp: has_vector_derivative_def C1_differentiable_on_def) - have "(g o (\x. 1 - x) has_vector_derivative -1 *\<^sub>R f') (at x)" - apply (rule vector_diff_chain_within) - apply (intro vector_diff_chain_within derivative_eq_intros | simp)+ - apply (rule has_vector_derivative_at_within [OF f']) - done - then have mf': "((\x. g (1 - x)) has_vector_derivative -f') (at x)" - by (simp add: o_def) - show ?thesis - using xs - by (auto simp: vector_derivative_at_within_ivl [OF mf'] vector_derivative_at_within_ivl [OF f']) - qed - } note * = this - have 01: "{0..1::real} = cbox 0 1" - by simp - show ?thesis using assms - apply (auto simp: has_contour_integral_def) - apply (drule has_integral_affinity01 [where m= "-1" and c=1]) - apply (auto simp: reversepath_def valid_path_def piecewise_C1_differentiable_on_def) - apply (drule has_integral_neg) - apply (rule_tac s = "(\x. 1 - x) ` s" in has_integral_spike_finite) - apply (auto simp: *) - done -qed - -lemma contour_integrable_reversepath: - "valid_path g \ f contour_integrable_on g \ f contour_integrable_on (reversepath g)" - using has_contour_integral_reversepath contour_integrable_on_def by blast - -lemma contour_integrable_reversepath_eq: - "valid_path g \ (f contour_integrable_on (reversepath g) \ f contour_integrable_on g)" - using contour_integrable_reversepath valid_path_reversepath by fastforce - -lemma contour_integral_reversepath: - assumes "valid_path g" - shows "contour_integral (reversepath g) f = - (contour_integral g f)" -proof (cases "f contour_integrable_on g") - case True then show ?thesis - by (simp add: assms contour_integral_unique has_contour_integral_integral has_contour_integral_reversepath) -next - case False then have "~ f contour_integrable_on (reversepath g)" - by (simp add: assms contour_integrable_reversepath_eq) - with False show ?thesis by (simp add: not_integrable_contour_integral) -qed - - -subsection\Joining two paths together\ - -lemma valid_path_join: - assumes "valid_path g1" "valid_path g2" "pathfinish g1 = pathstart g2" - shows "valid_path(g1 +++ g2)" -proof - - have "g1 1 = g2 0" - using assms by (auto simp: pathfinish_def pathstart_def) - moreover have "(g1 o (\x. 2*x)) piecewise_C1_differentiable_on {0..1/2}" - apply (rule piecewise_C1_differentiable_compose) - using assms - apply (auto simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_joinpaths) - apply (rule continuous_intros | simp)+ - apply (force intro: finite_vimageI [where h = "op*2"] inj_onI) - done - moreover have "(g2 o (\x. 2*x-1)) piecewise_C1_differentiable_on {1/2..1}" - apply (rule piecewise_C1_differentiable_compose) - using assms unfolding valid_path_def piecewise_C1_differentiable_on_def - by (auto intro!: continuous_intros finite_vimageI [where h = "(\x. 2*x - 1)"] inj_onI - simp: image_affinity_atLeastAtMost_diff continuous_on_joinpaths) - ultimately show ?thesis - apply (simp only: valid_path_def continuous_on_joinpaths joinpaths_def) - apply (rule piecewise_C1_differentiable_cases) - apply (auto simp: o_def) - done -qed - -lemma valid_path_join_D1: - fixes g1 :: "real \ 'a::real_normed_field" - shows "valid_path (g1 +++ g2) \ valid_path g1" - unfolding valid_path_def - by (rule piecewise_C1_differentiable_D1) - -lemma valid_path_join_D2: - fixes g2 :: "real \ 'a::real_normed_field" - shows "\valid_path (g1 +++ g2); pathfinish g1 = pathstart g2\ \ valid_path g2" - unfolding valid_path_def - by (rule piecewise_C1_differentiable_D2) - -lemma valid_path_join_eq [simp]: - fixes g2 :: "real \ 'a::real_normed_field" - shows "pathfinish g1 = pathstart g2 \ (valid_path(g1 +++ g2) \ valid_path g1 \ valid_path g2)" - using valid_path_join_D1 valid_path_join_D2 valid_path_join by blast - -lemma has_contour_integral_join: - assumes "(f has_contour_integral i1) g1" "(f has_contour_integral i2) g2" - "valid_path g1" "valid_path g2" - shows "(f has_contour_integral (i1 + i2)) (g1 +++ g2)" -proof - - obtain s1 s2 - where s1: "finite s1" "\x\{0..1} - s1. g1 differentiable at x" - and s2: "finite s2" "\x\{0..1} - s2. g2 differentiable at x" - using assms - by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - have 1: "((\x. f (g1 x) * vector_derivative g1 (at x)) has_integral i1) {0..1}" - and 2: "((\x. f (g2 x) * vector_derivative g2 (at x)) has_integral i2) {0..1}" - using assms - by (auto simp: has_contour_integral) - have i1: "((\x. (2*f (g1 (2*x))) * vector_derivative g1 (at (2*x))) has_integral i1) {0..1/2}" - and i2: "((\x. (2*f (g2 (2*x - 1))) * vector_derivative g2 (at (2*x - 1))) has_integral i2) {1/2..1}" - using has_integral_affinity01 [OF 1, where m= 2 and c=0, THEN has_integral_cmul [where c=2]] - has_integral_affinity01 [OF 2, where m= 2 and c="-1", THEN has_integral_cmul [where c=2]] - by (simp_all only: image_affinity_atLeastAtMost_div_diff, simp_all add: scaleR_conv_of_real mult_ac) - have g1: "\0 \ z; z*2 < 1; z*2 \ s1\ \ - vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at z) = - 2 *\<^sub>R vector_derivative g1 (at (z*2))" for z - apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g1(2*x))" and d = "\z - 1/2\"]]) - apply (simp_all add: dist_real_def abs_if split: if_split_asm) - apply (rule vector_diff_chain_at [of "\x. 2*x" 2 _ g1, simplified o_def]) - apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) - using s1 - apply (auto simp: algebra_simps vector_derivative_works) - done - have g2: "\1 < z*2; z \ 1; z*2 - 1 \ s2\ \ - vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at z) = - 2 *\<^sub>R vector_derivative g2 (at (z*2 - 1))" for z - apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g2 (2*x - 1))" and d = "\z - 1/2\"]]) - apply (simp_all add: dist_real_def abs_if split: if_split_asm) - apply (rule vector_diff_chain_at [of "\x. 2*x - 1" 2 _ g2, simplified o_def]) - apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) - using s2 - apply (auto simp: algebra_simps vector_derivative_works) - done - have "((\x. f ((g1 +++ g2) x) * vector_derivative (g1 +++ g2) (at x)) has_integral i1) {0..1/2}" - apply (rule has_integral_spike_finite [OF _ _ i1, of "insert (1/2) (op*2 -` s1)"]) - using s1 - apply (force intro: finite_vimageI [where h = "op*2"] inj_onI) - apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g1) - done - moreover have "((\x. f ((g1 +++ g2) x) * vector_derivative (g1 +++ g2) (at x)) has_integral i2) {1/2..1}" - apply (rule has_integral_spike_finite [OF _ _ i2, of "insert (1/2) ((\x. 2*x-1) -` s2)"]) - using s2 - apply (force intro: finite_vimageI [where h = "\x. 2*x-1"] inj_onI) - apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g2) - done - ultimately - show ?thesis - apply (simp add: has_contour_integral) - apply (rule has_integral_combine [where c = "1/2"], auto) - done -qed - -lemma contour_integrable_joinI: - assumes "f contour_integrable_on g1" "f contour_integrable_on g2" - "valid_path g1" "valid_path g2" - shows "f contour_integrable_on (g1 +++ g2)" - using assms - by (meson has_contour_integral_join contour_integrable_on_def) - -lemma contour_integrable_joinD1: - assumes "f contour_integrable_on (g1 +++ g2)" "valid_path g1" - shows "f contour_integrable_on g1" -proof - - obtain s1 - where s1: "finite s1" "\x\{0..1} - s1. g1 differentiable at x" - using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - have "(\x. f ((g1 +++ g2) (x/2)) * vector_derivative (g1 +++ g2) (at (x/2))) integrable_on {0..1}" - using assms - apply (auto simp: contour_integrable_on) - apply (drule integrable_on_subcbox [where a=0 and b="1/2"]) - apply (auto intro: integrable_affinity [of _ 0 "1/2::real" "1/2" 0, simplified]) - done - then have *: "(\x. (f ((g1 +++ g2) (x/2))/2) * vector_derivative (g1 +++ g2) (at (x/2))) integrable_on {0..1}" - by (auto dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real) - have g1: "\0 < z; z < 1; z \ s1\ \ - vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at (z/2)) = - 2 *\<^sub>R vector_derivative g1 (at z)" for z - apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g1(2*x))" and d = "\(z-1)/2\"]]) - apply (simp_all add: field_simps dist_real_def abs_if split: if_split_asm) - apply (rule vector_diff_chain_at [of "\x. x*2" 2 _ g1, simplified o_def]) - using s1 - apply (auto simp: vector_derivative_works has_vector_derivative_def has_derivative_def bounded_linear_mult_left) - done - show ?thesis - using s1 - apply (auto simp: contour_integrable_on) - apply (rule integrable_spike_finite [of "{0,1} \ s1", OF _ _ *]) - apply (auto simp: joinpaths_def scaleR_conv_of_real g1) - done -qed - -lemma contour_integrable_joinD2: - assumes "f contour_integrable_on (g1 +++ g2)" "valid_path g2" - shows "f contour_integrable_on g2" -proof - - obtain s2 - where s2: "finite s2" "\x\{0..1} - s2. g2 differentiable at x" - using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - have "(\x. f ((g1 +++ g2) (x/2 + 1/2)) * vector_derivative (g1 +++ g2) (at (x/2 + 1/2))) integrable_on {0..1}" - using assms - apply (auto simp: contour_integrable_on) - apply (drule integrable_on_subcbox [where a="1/2" and b=1], auto) - apply (drule integrable_affinity [of _ "1/2::real" 1 "1/2" "1/2", simplified]) - apply (simp add: image_affinity_atLeastAtMost_diff) - done - then have *: "(\x. (f ((g1 +++ g2) (x/2 + 1/2))/2) * vector_derivative (g1 +++ g2) (at (x/2 + 1/2))) - integrable_on {0..1}" - by (auto dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real) - have g2: "\0 < z; z < 1; z \ s2\ \ - vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at (z/2+1/2)) = - 2 *\<^sub>R vector_derivative g2 (at z)" for z - apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g2(2*x-1))" and d = "\z/2\"]]) - apply (simp_all add: field_simps dist_real_def abs_if split: if_split_asm) - apply (rule vector_diff_chain_at [of "\x. x*2-1" 2 _ g2, simplified o_def]) - using s2 - apply (auto simp: has_vector_derivative_def has_derivative_def bounded_linear_mult_left - vector_derivative_works add_divide_distrib) - done - show ?thesis - using s2 - apply (auto simp: contour_integrable_on) - apply (rule integrable_spike_finite [of "{0,1} \ s2", OF _ _ *]) - apply (auto simp: joinpaths_def scaleR_conv_of_real g2) - done -qed - -lemma contour_integrable_join [simp]: - shows - "\valid_path g1; valid_path g2\ - \ f contour_integrable_on (g1 +++ g2) \ f contour_integrable_on g1 \ f contour_integrable_on g2" -using contour_integrable_joinD1 contour_integrable_joinD2 contour_integrable_joinI by blast - -lemma contour_integral_join [simp]: - shows - "\f contour_integrable_on g1; f contour_integrable_on g2; valid_path g1; valid_path g2\ - \ contour_integral (g1 +++ g2) f = contour_integral g1 f + contour_integral g2 f" - by (simp add: has_contour_integral_integral has_contour_integral_join contour_integral_unique) - - -subsection\Shifting the starting point of a (closed) path\ - -lemma shiftpath_alt_def: "shiftpath a f = (\x. if x \ 1-a then f (a + x) else f (a + x - 1))" - by (auto simp: shiftpath_def) - -lemma valid_path_shiftpath [intro]: - assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" - shows "valid_path(shiftpath a g)" - using assms - apply (auto simp: valid_path_def shiftpath_alt_def) - apply (rule piecewise_C1_differentiable_cases) - apply (auto simp: algebra_simps) - apply (rule piecewise_C1_differentiable_affine [of g 1 a, simplified o_def scaleR_one]) - apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset) - apply (rule piecewise_C1_differentiable_affine [of g 1 "a-1", simplified o_def scaleR_one algebra_simps]) - apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset) - done - -lemma has_contour_integral_shiftpath: - assumes f: "(f has_contour_integral i) g" "valid_path g" - and a: "a \ {0..1}" - shows "(f has_contour_integral i) (shiftpath a g)" -proof - - obtain s - where s: "finite s" and g: "\x\{0..1} - s. g differentiable at x" - using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - have *: "((\x. f (g x) * vector_derivative g (at x)) has_integral i) {0..1}" - using assms by (auto simp: has_contour_integral) - then have i: "i = integral {a..1} (\x. f (g x) * vector_derivative g (at x)) + - integral {0..a} (\x. f (g x) * vector_derivative g (at x))" - apply (rule has_integral_unique) - apply (subst add.commute) - apply (subst Integration.integral_combine) - using assms * integral_unique by auto - { fix x - have "0 \ x \ x + a < 1 \ x \ (\x. x - a) ` s \ - vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a))" - unfolding shiftpath_def - apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g(a+x))" and d = "dist(1-a) x"]]) - apply (auto simp: field_simps dist_real_def abs_if split: if_split_asm) - apply (rule vector_diff_chain_at [of "\x. x+a" 1 _ g, simplified o_def scaleR_one]) - apply (intro derivative_eq_intros | simp)+ - using g - apply (drule_tac x="x+a" in bspec) - using a apply (auto simp: has_vector_derivative_def vector_derivative_works image_def add.commute) - done - } note vd1 = this - { fix x - have "1 < x + a \ x \ 1 \ x \ (\x. x - a + 1) ` s \ - vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a - 1))" - unfolding shiftpath_def - apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g(a+x-1))" and d = "dist (1-a) x"]]) - apply (auto simp: field_simps dist_real_def abs_if split: if_split_asm) - apply (rule vector_diff_chain_at [of "\x. x+a-1" 1 _ g, simplified o_def scaleR_one]) - apply (intro derivative_eq_intros | simp)+ - using g - apply (drule_tac x="x+a-1" in bspec) - using a apply (auto simp: has_vector_derivative_def vector_derivative_works image_def add.commute) - done - } note vd2 = this - have va1: "(\x. f (g x) * vector_derivative g (at x)) integrable_on ({a..1})" - using * a by (fastforce intro: integrable_subinterval_real) - have v0a: "(\x. f (g x) * vector_derivative g (at x)) integrable_on ({0..a})" - apply (rule integrable_subinterval_real) - using * a by auto - have "((\x. f (shiftpath a g x) * vector_derivative (shiftpath a g) (at x)) - has_integral integral {a..1} (\x. f (g x) * vector_derivative g (at x))) {0..1 - a}" - apply (rule has_integral_spike_finite - [where s = "{1-a} \ (\x. x-a) ` s" and f = "\x. f(g(a+x)) * vector_derivative g (at(a+x))"]) - using s apply blast - using a apply (auto simp: algebra_simps vd1) - apply (force simp: shiftpath_def add.commute) - using has_integral_affinity [where m=1 and c=a, simplified, OF integrable_integral [OF va1]] - apply (simp add: image_affinity_atLeastAtMost_diff [where m=1 and c=a, simplified] add.commute) - done - moreover - have "((\x. f (shiftpath a g x) * vector_derivative (shiftpath a g) (at x)) - has_integral integral {0..a} (\x. f (g x) * vector_derivative g (at x))) {1 - a..1}" - apply (rule has_integral_spike_finite - [where s = "{1-a} \ (\x. x-a+1) ` s" and f = "\x. f(g(a+x-1)) * vector_derivative g (at(a+x-1))"]) - using s apply blast - using a apply (auto simp: algebra_simps vd2) - apply (force simp: shiftpath_def add.commute) - using has_integral_affinity [where m=1 and c="a-1", simplified, OF integrable_integral [OF v0a]] - apply (simp add: image_affinity_atLeastAtMost [where m=1 and c="1-a", simplified]) - apply (simp add: algebra_simps) - done - ultimately show ?thesis - using a - by (auto simp: i has_contour_integral intro: has_integral_combine [where c = "1-a"]) -qed - -lemma has_contour_integral_shiftpath_D: - assumes "(f has_contour_integral i) (shiftpath a g)" - "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" - shows "(f has_contour_integral i) g" -proof - - obtain s - where s: "finite s" and g: "\x\{0..1} - s. g differentiable at x" - using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - { fix x - assume x: "0 < x" "x < 1" "x \ s" - then have gx: "g differentiable at x" - using g by auto - have "vector_derivative g (at x within {0..1}) = - vector_derivative (shiftpath (1 - a) (shiftpath a g)) (at x within {0..1})" - apply (rule vector_derivative_at_within_ivl - [OF has_vector_derivative_transform_within_open - [where f = "(shiftpath (1 - a) (shiftpath a g))" and s = "{0<..<1}-s"]]) - using s g assms x - apply (auto simp: finite_imp_closed open_Diff shiftpath_shiftpath - vector_derivative_within_interior vector_derivative_works [symmetric]) - apply (rule differentiable_transform_within [OF gx, of "min x (1-x)"]) - apply (auto simp: dist_real_def shiftpath_shiftpath abs_if split: if_split_asm) - done - } note vd = this - have fi: "(f has_contour_integral i) (shiftpath (1 - a) (shiftpath a g))" - using assms by (auto intro!: has_contour_integral_shiftpath) - show ?thesis - apply (simp add: has_contour_integral_def) - apply (rule has_integral_spike_finite [of "{0,1} \ s", OF _ _ fi [unfolded has_contour_integral_def]]) - using s assms vd - apply (auto simp: Path_Connected.shiftpath_shiftpath) - done -qed - -lemma has_contour_integral_shiftpath_eq: - assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" - shows "(f has_contour_integral i) (shiftpath a g) \ (f has_contour_integral i) g" - using assms has_contour_integral_shiftpath has_contour_integral_shiftpath_D by blast - -lemma contour_integrable_on_shiftpath_eq: - assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" - shows "f contour_integrable_on (shiftpath a g) \ f contour_integrable_on g" -using assms contour_integrable_on_def has_contour_integral_shiftpath_eq by auto - -lemma contour_integral_shiftpath: - assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" - shows "contour_integral (shiftpath a g) f = contour_integral g f" - using assms - by (simp add: contour_integral_def contour_integrable_on_def has_contour_integral_shiftpath_eq) - - -subsection\More about straight-line paths\ - -lemma has_vector_derivative_linepath_within: - "(linepath a b has_vector_derivative (b - a)) (at x within s)" -apply (simp add: linepath_def has_vector_derivative_def algebra_simps) -apply (rule derivative_eq_intros | simp)+ -done - -lemma vector_derivative_linepath_within: - "x \ {0..1} \ vector_derivative (linepath a b) (at x within {0..1}) = b - a" - apply (rule vector_derivative_within_closed_interval [of 0 "1::real", simplified]) - apply (auto simp: has_vector_derivative_linepath_within) - done - -lemma vector_derivative_linepath_at [simp]: "vector_derivative (linepath a b) (at x) = b - a" - by (simp add: has_vector_derivative_linepath_within vector_derivative_at) - -lemma valid_path_linepath [iff]: "valid_path (linepath a b)" - apply (simp add: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq continuous_on_linepath) - apply (rule_tac x="{}" in exI) - apply (simp add: differentiable_on_def differentiable_def) - using has_vector_derivative_def has_vector_derivative_linepath_within - apply (fastforce simp add: continuous_on_eq_continuous_within) - done - -lemma has_contour_integral_linepath: - shows "(f has_contour_integral i) (linepath a b) \ - ((\x. f(linepath a b x) * (b - a)) has_integral i) {0..1}" - by (simp add: has_contour_integral vector_derivative_linepath_at) - -lemma linepath_in_path: - shows "x \ {0..1} \ linepath a b x \ closed_segment a b" - by (auto simp: segment linepath_def) - -lemma linepath_image_01: "linepath a b ` {0..1} = closed_segment a b" - by (auto simp: segment linepath_def) - -lemma linepath_in_convex_hull: - fixes x::real - assumes a: "a \ convex hull s" - and b: "b \ convex hull s" - and x: "0\x" "x\1" - shows "linepath a b x \ convex hull s" - apply (rule closed_segment_subset_convex_hull [OF a b, THEN subsetD]) - using x - apply (auto simp: linepath_image_01 [symmetric]) - done - -lemma Re_linepath: "Re(linepath (of_real a) (of_real b) x) = (1 - x)*a + x*b" - by (simp add: linepath_def) - -lemma Im_linepath: "Im(linepath (of_real a) (of_real b) x) = 0" - by (simp add: linepath_def) - -lemma has_contour_integral_trivial [iff]: "(f has_contour_integral 0) (linepath a a)" - by (simp add: has_contour_integral_linepath) - -lemma contour_integral_trivial [simp]: "contour_integral (linepath a a) f = 0" - using has_contour_integral_trivial contour_integral_unique by blast - - -subsection\Relation to subpath construction\ - -lemma valid_path_subpath: - fixes g :: "real \ 'a :: real_normed_vector" - assumes "valid_path g" "u \ {0..1}" "v \ {0..1}" - shows "valid_path(subpath u v g)" -proof (cases "v=u") - case True - then show ?thesis - unfolding valid_path_def subpath_def - by (force intro: C1_differentiable_on_const C1_differentiable_imp_piecewise) -next - case False - have "(g o (\x. ((v-u) * x + u))) piecewise_C1_differentiable_on {0..1}" - apply (rule piecewise_C1_differentiable_compose) - apply (simp add: C1_differentiable_imp_piecewise) - apply (simp add: image_affinity_atLeastAtMost) - using assms False - apply (auto simp: algebra_simps valid_path_def piecewise_C1_differentiable_on_subset) - apply (subst Int_commute) - apply (auto simp: inj_on_def algebra_simps crossproduct_eq finite_vimage_IntI) - done - then show ?thesis - by (auto simp: o_def valid_path_def subpath_def) -qed - -lemma has_contour_integral_subpath_refl [iff]: "(f has_contour_integral 0) (subpath u u g)" - by (simp add: has_contour_integral subpath_def) - -lemma contour_integrable_subpath_refl [iff]: "f contour_integrable_on (subpath u u g)" - using has_contour_integral_subpath_refl contour_integrable_on_def by blast - -lemma contour_integral_subpath_refl [simp]: "contour_integral (subpath u u g) f = 0" - by (simp add: has_contour_integral_subpath_refl contour_integral_unique) - -lemma has_contour_integral_subpath: - assumes f: "f contour_integrable_on g" and g: "valid_path g" - and uv: "u \ {0..1}" "v \ {0..1}" "u \ v" - shows "(f has_contour_integral integral {u..v} (\x. f(g x) * vector_derivative g (at x))) - (subpath u v g)" -proof (cases "v=u") - case True - then show ?thesis - using f by (simp add: contour_integrable_on_def subpath_def has_contour_integral) -next - case False - obtain s where s: "\x. x \ {0..1} - s \ g differentiable at x" and fs: "finite s" - using g unfolding piecewise_C1_differentiable_on_def C1_differentiable_on_eq valid_path_def by blast - have *: "((\x. f (g ((v - u) * x + u)) * vector_derivative g (at ((v - u) * x + u))) - has_integral (1 / (v - u)) * integral {u..v} (\t. f (g t) * vector_derivative g (at t))) - {0..1}" - using f uv - apply (simp add: contour_integrable_on subpath_def has_contour_integral) - apply (drule integrable_on_subcbox [where a=u and b=v, simplified]) - apply (simp_all add: has_integral_integral) - apply (drule has_integral_affinity [where m="v-u" and c=u, simplified]) - apply (simp_all add: False image_affinity_atLeastAtMost_div_diff scaleR_conv_of_real) - apply (simp add: divide_simps False) - done - { fix x - have "x \ {0..1} \ - x \ (\t. (v-u) *\<^sub>R t + u) -` s \ - vector_derivative (\x. g ((v-u) * x + u)) (at x) = (v-u) *\<^sub>R vector_derivative g (at ((v-u) * x + u))" - apply (rule vector_derivative_at [OF vector_diff_chain_at [simplified o_def]]) - apply (intro derivative_eq_intros | simp)+ - apply (cut_tac s [of "(v - u) * x + u"]) - using uv mult_left_le [of x "v-u"] - apply (auto simp: vector_derivative_works) - done - } note vd = this - show ?thesis - apply (cut_tac has_integral_cmul [OF *, where c = "v-u"]) - using fs assms - apply (simp add: False subpath_def has_contour_integral) - apply (rule_tac s = "(\t. ((v-u) *\<^sub>R t + u)) -` s" in has_integral_spike_finite) - apply (auto simp: inj_on_def False finite_vimageI vd scaleR_conv_of_real) - done -qed - -lemma contour_integrable_subpath: - assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" - shows "f contour_integrable_on (subpath u v g)" - apply (cases u v rule: linorder_class.le_cases) - apply (metis contour_integrable_on_def has_contour_integral_subpath [OF assms]) - apply (subst reversepath_subpath [symmetric]) - apply (rule contour_integrable_reversepath) - using assms apply (blast intro: valid_path_subpath) - apply (simp add: contour_integrable_on_def) - using assms apply (blast intro: has_contour_integral_subpath) - done - -lemma has_integral_integrable_integral: "(f has_integral i) s \ f integrable_on s \ integral s f = i" - by blast - -lemma has_integral_contour_integral_subpath: - assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "u \ v" - shows "(((\x. f(g x) * vector_derivative g (at x))) - has_integral contour_integral (subpath u v g) f) {u..v}" - using assms - apply (auto simp: has_integral_integrable_integral) - apply (rule integrable_on_subcbox [where a=u and b=v and s = "{0..1}", simplified]) - apply (auto simp: contour_integral_unique [OF has_contour_integral_subpath] contour_integrable_on) - done - -lemma contour_integral_subcontour_integral: - assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "u \ v" - shows "contour_integral (subpath u v g) f = - integral {u..v} (\x. f(g x) * vector_derivative g (at x))" - using assms has_contour_integral_subpath contour_integral_unique by blast - -lemma contour_integral_subpath_combine_less: - assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "w \ {0..1}" - "u {0..1}" "v \ {0..1}" "w \ {0..1}" - shows "contour_integral (subpath u v g) f + contour_integral (subpath v w g) f = - contour_integral (subpath u w g) f" -proof (cases "u\v \ v\w \ u\w") - case True - have *: "subpath v u g = reversepath(subpath u v g) \ - subpath w u g = reversepath(subpath u w g) \ - subpath w v g = reversepath(subpath v w g)" - by (auto simp: reversepath_subpath) - have "u < v \ v < w \ - u < w \ w < v \ - v < u \ u < w \ - v < w \ w < u \ - w < u \ u < v \ - w < v \ v < u" - using True assms by linarith - with assms show ?thesis - using contour_integral_subpath_combine_less [of f g u v w] - contour_integral_subpath_combine_less [of f g u w v] - contour_integral_subpath_combine_less [of f g v u w] - contour_integral_subpath_combine_less [of f g v w u] - contour_integral_subpath_combine_less [of f g w u v] - contour_integral_subpath_combine_less [of f g w v u] - apply simp - apply (elim disjE) - apply (auto simp: * contour_integral_reversepath contour_integrable_subpath - valid_path_reversepath valid_path_subpath algebra_simps) - done -next - case False - then show ?thesis - apply (auto simp: contour_integral_subpath_refl) - using assms - by (metis eq_neg_iff_add_eq_0 contour_integrable_subpath contour_integral_reversepath reversepath_subpath valid_path_subpath) -qed - -lemma contour_integral_integral: - "contour_integral g f = integral {0..1} (\x. f (g x) * vector_derivative g (at x))" - by (simp add: contour_integral_def integral_def has_contour_integral contour_integrable_on) - - -text\Cauchy's theorem where there's a primitive\ - -lemma contour_integral_primitive_lemma: - fixes f :: "complex \ complex" and g :: "real \ complex" - assumes "a \ b" - and "\x. x \ s \ (f has_field_derivative f' x) (at x within s)" - and "g piecewise_differentiable_on {a..b}" "\x. x \ {a..b} \ g x \ s" - shows "((\x. f'(g x) * vector_derivative g (at x within {a..b})) - has_integral (f(g b) - f(g a))) {a..b}" -proof - - obtain k where k: "finite k" "\x\{a..b} - k. g differentiable (at x within {a..b})" and cg: "continuous_on {a..b} g" - using assms by (auto simp: piecewise_differentiable_on_def) - have cfg: "continuous_on {a..b} (\x. f (g x))" - apply (rule continuous_on_compose [OF cg, unfolded o_def]) - using assms - apply (metis field_differentiable_def field_differentiable_imp_continuous_at continuous_on_eq_continuous_within continuous_on_subset image_subset_iff) - done - { fix x::real - assume a: "a < x" and b: "x < b" and xk: "x \ k" - then have "g differentiable at x within {a..b}" - using k by (simp add: differentiable_at_withinI) - then have "(g has_vector_derivative vector_derivative g (at x within {a..b})) (at x within {a..b})" - by (simp add: vector_derivative_works has_field_derivative_def scaleR_conv_of_real) - then have gdiff: "(g has_derivative (\u. u * vector_derivative g (at x within {a..b}))) (at x within {a..b})" - by (simp add: has_vector_derivative_def scaleR_conv_of_real) - have "(f has_field_derivative (f' (g x))) (at (g x) within g ` {a..b})" - using assms by (metis a atLeastAtMost_iff b DERIV_subset image_subset_iff less_eq_real_def) - then have fdiff: "(f has_derivative op * (f' (g x))) (at (g x) within g ` {a..b})" - by (simp add: has_field_derivative_def) - have "((\x. f (g x)) has_vector_derivative f' (g x) * vector_derivative g (at x within {a..b})) (at x within {a..b})" - using diff_chain_within [OF gdiff fdiff] - by (simp add: has_vector_derivative_def scaleR_conv_of_real o_def mult_ac) - } note * = this - show ?thesis - apply (rule fundamental_theorem_of_calculus_interior_strong) - using k assms cfg * - apply (auto simp: at_within_closed_interval) - done -qed - -lemma contour_integral_primitive: - assumes "\x. x \ s \ (f has_field_derivative f' x) (at x within s)" - and "valid_path g" "path_image g \ s" - shows "(f' has_contour_integral (f(pathfinish g) - f(pathstart g))) g" - using assms - apply (simp add: valid_path_def path_image_def pathfinish_def pathstart_def has_contour_integral_def) - apply (auto intro!: piecewise_C1_imp_differentiable contour_integral_primitive_lemma [of 0 1 s]) - done - -corollary Cauchy_theorem_primitive: - assumes "\x. x \ s \ (f has_field_derivative f' x) (at x within s)" - and "valid_path g" "path_image g \ s" "pathfinish g = pathstart g" - shows "(f' has_contour_integral 0) g" - using assms - by (metis diff_self contour_integral_primitive) - - -text\Existence of path integral for continuous function\ -lemma contour_integrable_continuous_linepath: - assumes "continuous_on (closed_segment a b) f" - shows "f contour_integrable_on (linepath a b)" -proof - - have "continuous_on {0..1} ((\x. f x * (b - a)) o linepath a b)" - apply (rule continuous_on_compose [OF continuous_on_linepath], simp add: linepath_image_01) - apply (rule continuous_intros | simp add: assms)+ - done - then show ?thesis - apply (simp add: contour_integrable_on_def has_contour_integral_def integrable_on_def [symmetric]) - apply (rule integrable_continuous [of 0 "1::real", simplified]) - apply (rule continuous_on_eq [where f = "\x. f(linepath a b x)*(b - a)"]) - apply (auto simp: vector_derivative_linepath_within) - done -qed - -lemma has_field_der_id: "((\x. x\<^sup>2 / 2) has_field_derivative x) (at x)" - by (rule has_derivative_imp_has_field_derivative) - (rule derivative_intros | simp)+ - -lemma contour_integral_id [simp]: "contour_integral (linepath a b) (\y. y) = (b^2 - a^2)/2" - apply (rule contour_integral_unique) - using contour_integral_primitive [of UNIV "\x. x^2/2" "\x. x" "linepath a b"] - apply (auto simp: field_simps has_field_der_id) - done - -lemma contour_integrable_on_const [iff]: "(\x. c) contour_integrable_on (linepath a b)" - by (simp add: continuous_on_const contour_integrable_continuous_linepath) - -lemma contour_integrable_on_id [iff]: "(\x. x) contour_integrable_on (linepath a b)" - by (simp add: continuous_on_id contour_integrable_continuous_linepath) - - -subsection\Arithmetical combining theorems\ - -lemma has_contour_integral_neg: - "(f has_contour_integral i) g \ ((\x. -(f x)) has_contour_integral (-i)) g" - by (simp add: has_integral_neg has_contour_integral_def) - -lemma has_contour_integral_add: - "\(f1 has_contour_integral i1) g; (f2 has_contour_integral i2) g\ - \ ((\x. f1 x + f2 x) has_contour_integral (i1 + i2)) g" - by (simp add: has_integral_add has_contour_integral_def algebra_simps) - -lemma has_contour_integral_diff: - "\(f1 has_contour_integral i1) g; (f2 has_contour_integral i2) g\ - \ ((\x. f1 x - f2 x) has_contour_integral (i1 - i2)) g" - by (simp add: has_integral_sub has_contour_integral_def algebra_simps) - -lemma has_contour_integral_lmul: - "(f has_contour_integral i) g \ ((\x. c * (f x)) has_contour_integral (c*i)) g" -apply (simp add: has_contour_integral_def) -apply (drule has_integral_mult_right) -apply (simp add: algebra_simps) -done - -lemma has_contour_integral_rmul: - "(f has_contour_integral i) g \ ((\x. (f x) * c) has_contour_integral (i*c)) g" -apply (drule has_contour_integral_lmul) -apply (simp add: mult.commute) -done - -lemma has_contour_integral_div: - "(f has_contour_integral i) g \ ((\x. f x/c) has_contour_integral (i/c)) g" - by (simp add: field_class.field_divide_inverse) (metis has_contour_integral_rmul) - -lemma has_contour_integral_eq: - "\(f has_contour_integral y) p; \x. x \ path_image p \ f x = g x\ \ (g has_contour_integral y) p" -apply (simp add: path_image_def has_contour_integral_def) -by (metis (no_types, lifting) image_eqI has_integral_eq) - -lemma has_contour_integral_bound_linepath: - assumes "(f has_contour_integral i) (linepath a b)" - "0 \ B" "\x. x \ closed_segment a b \ norm(f x) \ B" - shows "norm i \ B * norm(b - a)" -proof - - { fix x::real - assume x: "0 \ x" "x \ 1" - have "norm (f (linepath a b x)) * - norm (vector_derivative (linepath a b) (at x within {0..1})) \ B * norm (b - a)" - by (auto intro: mult_mono simp: assms linepath_in_path of_real_linepath vector_derivative_linepath_within x) - } note * = this - have "norm i \ (B * norm (b - a)) * content (cbox 0 (1::real))" - apply (rule has_integral_bound - [of _ "\x. f (linepath a b x) * vector_derivative (linepath a b) (at x within {0..1})"]) - using assms * unfolding has_contour_integral_def - apply (auto simp: norm_mult) - done - then show ?thesis - by (auto simp: content_real) -qed - -(*UNUSED -lemma has_contour_integral_bound_linepath_strong: - fixes a :: real and f :: "complex \ real" - assumes "(f has_contour_integral i) (linepath a b)" - "finite k" - "0 \ B" "\x::real. x \ closed_segment a b - k \ norm(f x) \ B" - shows "norm i \ B*norm(b - a)" -*) - -lemma has_contour_integral_const_linepath: "((\x. c) has_contour_integral c*(b - a))(linepath a b)" - unfolding has_contour_integral_linepath - by (metis content_real diff_0_right has_integral_const_real lambda_one of_real_1 scaleR_conv_of_real zero_le_one) - -lemma has_contour_integral_0: "((\x. 0) has_contour_integral 0) g" - by (simp add: has_contour_integral_def) - -lemma has_contour_integral_is_0: - "(\z. z \ path_image g \ f z = 0) \ (f has_contour_integral 0) g" - by (rule has_contour_integral_eq [OF has_contour_integral_0]) auto - -lemma has_contour_integral_setsum: - "\finite s; \a. a \ s \ (f a has_contour_integral i a) p\ - \ ((\x. setsum (\a. f a x) s) has_contour_integral setsum i s) p" - by (induction s rule: finite_induct) (auto simp: has_contour_integral_0 has_contour_integral_add) - - -subsection \Operations on path integrals\ - -lemma contour_integral_const_linepath [simp]: "contour_integral (linepath a b) (\x. c) = c*(b - a)" - by (rule contour_integral_unique [OF has_contour_integral_const_linepath]) - -lemma contour_integral_neg: - "f contour_integrable_on g \ contour_integral g (\x. -(f x)) = -(contour_integral g f)" - by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_neg) - -lemma contour_integral_add: - "f1 contour_integrable_on g \ f2 contour_integrable_on g \ contour_integral g (\x. f1 x + f2 x) = - contour_integral g f1 + contour_integral g f2" - by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_add) - -lemma contour_integral_diff: - "f1 contour_integrable_on g \ f2 contour_integrable_on g \ contour_integral g (\x. f1 x - f2 x) = - contour_integral g f1 - contour_integral g f2" - by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_diff) - -lemma contour_integral_lmul: - shows "f contour_integrable_on g - \ contour_integral g (\x. c * f x) = c*contour_integral g f" - by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_lmul) - -lemma contour_integral_rmul: - shows "f contour_integrable_on g - \ contour_integral g (\x. f x * c) = contour_integral g f * c" - by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_rmul) - -lemma contour_integral_div: - shows "f contour_integrable_on g - \ contour_integral g (\x. f x / c) = contour_integral g f / c" - by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_div) - -lemma contour_integral_eq: - "(\x. x \ path_image p \ f x = g x) \ contour_integral p f = contour_integral p g" - apply (simp add: contour_integral_def) - using has_contour_integral_eq - by (metis contour_integral_unique has_contour_integral_integrable has_contour_integral_integral) - -lemma contour_integral_eq_0: - "(\z. z \ path_image g \ f z = 0) \ contour_integral g f = 0" - by (simp add: has_contour_integral_is_0 contour_integral_unique) - -lemma contour_integral_bound_linepath: - shows - "\f contour_integrable_on (linepath a b); - 0 \ B; \x. x \ closed_segment a b \ norm(f x) \ B\ - \ norm(contour_integral (linepath a b) f) \ B*norm(b - a)" - apply (rule has_contour_integral_bound_linepath [of f]) - apply (auto simp: has_contour_integral_integral) - done - -lemma contour_integral_0 [simp]: "contour_integral g (\x. 0) = 0" - by (simp add: contour_integral_unique has_contour_integral_0) - -lemma contour_integral_setsum: - "\finite s; \a. a \ s \ (f a) contour_integrable_on p\ - \ contour_integral p (\x. setsum (\a. f a x) s) = setsum (\a. contour_integral p (f a)) s" - by (auto simp: contour_integral_unique has_contour_integral_setsum has_contour_integral_integral) - -lemma contour_integrable_eq: - "\f contour_integrable_on p; \x. x \ path_image p \ f x = g x\ \ g contour_integrable_on p" - unfolding contour_integrable_on_def - by (metis has_contour_integral_eq) - - -subsection \Arithmetic theorems for path integrability\ - -lemma contour_integrable_neg: - "f contour_integrable_on g \ (\x. -(f x)) contour_integrable_on g" - using has_contour_integral_neg contour_integrable_on_def by blast - -lemma contour_integrable_add: - "\f1 contour_integrable_on g; f2 contour_integrable_on g\ \ (\x. f1 x + f2 x) contour_integrable_on g" - using has_contour_integral_add contour_integrable_on_def - by fastforce - -lemma contour_integrable_diff: - "\f1 contour_integrable_on g; f2 contour_integrable_on g\ \ (\x. f1 x - f2 x) contour_integrable_on g" - using has_contour_integral_diff contour_integrable_on_def - by fastforce - -lemma contour_integrable_lmul: - "f contour_integrable_on g \ (\x. c * f x) contour_integrable_on g" - using has_contour_integral_lmul contour_integrable_on_def - by fastforce - -lemma contour_integrable_rmul: - "f contour_integrable_on g \ (\x. f x * c) contour_integrable_on g" - using has_contour_integral_rmul contour_integrable_on_def - by fastforce - -lemma contour_integrable_div: - "f contour_integrable_on g \ (\x. f x / c) contour_integrable_on g" - using has_contour_integral_div contour_integrable_on_def - by fastforce - -lemma contour_integrable_setsum: - "\finite s; \a. a \ s \ (f a) contour_integrable_on p\ - \ (\x. setsum (\a. f a x) s) contour_integrable_on p" - unfolding contour_integrable_on_def - by (metis has_contour_integral_setsum) - - -subsection\Reversing a path integral\ - -lemma has_contour_integral_reverse_linepath: - "(f has_contour_integral i) (linepath a b) - \ (f has_contour_integral (-i)) (linepath b a)" - using has_contour_integral_reversepath valid_path_linepath by fastforce - -lemma contour_integral_reverse_linepath: - "continuous_on (closed_segment a b) f - \ contour_integral (linepath a b) f = - (contour_integral(linepath b a) f)" -apply (rule contour_integral_unique) -apply (rule has_contour_integral_reverse_linepath) -by (simp add: closed_segment_commute contour_integrable_continuous_linepath has_contour_integral_integral) - - -(* Splitting a path integral in a flat way.*) - -lemma has_contour_integral_split: - assumes f: "(f has_contour_integral i) (linepath a c)" "(f has_contour_integral j) (linepath c b)" - and k: "0 \ k" "k \ 1" - and c: "c - a = k *\<^sub>R (b - a)" - shows "(f has_contour_integral (i + j)) (linepath a b)" -proof (cases "k = 0 \ k = 1") - case True - then show ?thesis - using assms - apply auto - apply (metis add.left_neutral has_contour_integral_trivial has_contour_integral_unique) - apply (metis add.right_neutral has_contour_integral_trivial has_contour_integral_unique) - done -next - case False - then have k: "0 < k" "k < 1" "complex_of_real k \ 1" - using assms apply auto - using of_real_eq_iff by fastforce - have c': "c = k *\<^sub>R (b - a) + a" - by (metis diff_add_cancel c) - have bc: "(b - c) = (1 - k) *\<^sub>R (b - a)" - by (simp add: algebra_simps c') - { assume *: "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R c) * (c - a)) has_integral i) {0..1}" - have **: "\x. ((k - x) / k) *\<^sub>R a + (x / k) *\<^sub>R c = (1 - x) *\<^sub>R a + x *\<^sub>R b" - using False - apply (simp add: c' algebra_simps) - apply (simp add: real_vector.scale_left_distrib [symmetric] divide_simps) - done - have "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral i) {0..k}" - using * k - apply - - apply (drule has_integral_affinity [of _ _ 0 "1::real" "inverse k" "0", simplified]) - apply (simp_all add: divide_simps mult.commute [of _ "k"] image_affinity_atLeastAtMost ** c) - apply (drule Integration.has_integral_cmul [where c = "inverse k"]) - apply (simp add: Integration.has_integral_cmul) - done - } note fi = this - { assume *: "((\x. f ((1 - x) *\<^sub>R c + x *\<^sub>R b) * (b - c)) has_integral j) {0..1}" - have **: "\x. (((1 - x) / (1 - k)) *\<^sub>R c + ((x - k) / (1 - k)) *\<^sub>R b) = ((1 - x) *\<^sub>R a + x *\<^sub>R b)" - using k - apply (simp add: c' field_simps) - apply (simp add: scaleR_conv_of_real divide_simps) - apply (simp add: field_simps) - done - have "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral j) {k..1}" - using * k - apply - - apply (drule has_integral_affinity [of _ _ 0 "1::real" "inverse(1 - k)" "-(k/(1 - k))", simplified]) - apply (simp_all add: divide_simps mult.commute [of _ "1-k"] image_affinity_atLeastAtMost ** bc) - apply (drule Integration.has_integral_cmul [where k = "(1 - k) *\<^sub>R j" and c = "inverse (1 - k)"]) - apply (simp add: Integration.has_integral_cmul) - done - } note fj = this - show ?thesis - using f k - apply (simp add: has_contour_integral_linepath) - apply (simp add: linepath_def) - apply (rule has_integral_combine [OF _ _ fi fj], simp_all) - done -qed - -lemma continuous_on_closed_segment_transform: - assumes f: "continuous_on (closed_segment a b) f" - and k: "0 \ k" "k \ 1" - and c: "c - a = k *\<^sub>R (b - a)" - shows "continuous_on (closed_segment a c) f" -proof - - have c': "c = (1 - k) *\<^sub>R a + k *\<^sub>R b" - using c by (simp add: algebra_simps) - show "continuous_on (closed_segment a c) f" - apply (rule continuous_on_subset [OF f]) - apply (simp add: segment_convex_hull) - apply (rule convex_hull_subset) - using assms - apply (auto simp: hull_inc c' Convex.convexD_alt) - done -qed - -lemma contour_integral_split: - assumes f: "continuous_on (closed_segment a b) f" - and k: "0 \ k" "k \ 1" - and c: "c - a = k *\<^sub>R (b - a)" - shows "contour_integral(linepath a b) f = contour_integral(linepath a c) f + contour_integral(linepath c b) f" -proof - - have c': "c = (1 - k) *\<^sub>R a + k *\<^sub>R b" - using c by (simp add: algebra_simps) - have *: "continuous_on (closed_segment a c) f" "continuous_on (closed_segment c b) f" - apply (rule_tac [!] continuous_on_subset [OF f]) - apply (simp_all add: segment_convex_hull) - apply (rule_tac [!] convex_hull_subset) - using assms - apply (auto simp: hull_inc c' Convex.convexD_alt) - done - show ?thesis - apply (rule contour_integral_unique) - apply (rule has_contour_integral_split [OF has_contour_integral_integral has_contour_integral_integral k c]) - apply (rule contour_integrable_continuous_linepath *)+ - done -qed - -lemma contour_integral_split_linepath: - assumes f: "continuous_on (closed_segment a b) f" - and c: "c \ closed_segment a b" - shows "contour_integral(linepath a b) f = contour_integral(linepath a c) f + contour_integral(linepath c b) f" - using c - by (auto simp: closed_segment_def algebra_simps intro!: contour_integral_split [OF f]) - -(* The special case of midpoints used in the main quadrisection.*) - -lemma has_contour_integral_midpoint: - assumes "(f has_contour_integral i) (linepath a (midpoint a b))" - "(f has_contour_integral j) (linepath (midpoint a b) b)" - shows "(f has_contour_integral (i + j)) (linepath a b)" - apply (rule has_contour_integral_split [where c = "midpoint a b" and k = "1/2"]) - using assms - apply (auto simp: midpoint_def algebra_simps scaleR_conv_of_real) - done - -lemma contour_integral_midpoint: - "continuous_on (closed_segment a b) f - \ contour_integral (linepath a b) f = - contour_integral (linepath a (midpoint a b)) f + contour_integral (linepath (midpoint a b) b) f" - apply (rule contour_integral_split [where c = "midpoint a b" and k = "1/2"]) - apply (auto simp: midpoint_def algebra_simps scaleR_conv_of_real) - done - - -text\A couple of special case lemmas that are useful below\ - -lemma triangle_linear_has_chain_integral: - "((\x. m*x + d) has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" - apply (rule Cauchy_theorem_primitive [of UNIV "\x. m/2 * x^2 + d*x"]) - apply (auto intro!: derivative_eq_intros) - done - -lemma has_chain_integral_chain_integral3: - "(f has_contour_integral i) (linepath a b +++ linepath b c +++ linepath c d) - \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c d) f = i" - apply (subst contour_integral_unique [symmetric], assumption) - apply (drule has_contour_integral_integrable) - apply (simp add: valid_path_join) - done - -lemma has_chain_integral_chain_integral4: - "(f has_contour_integral i) (linepath a b +++ linepath b c +++ linepath c d +++ linepath d e) - \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c d) f + contour_integral (linepath d e) f = i" - apply (subst contour_integral_unique [symmetric], assumption) - apply (drule has_contour_integral_integrable) - apply (simp add: valid_path_join) - done - -subsection\Reversing the order in a double path integral\ - -text\The condition is stronger than needed but it's often true in typical situations\ - -lemma fst_im_cbox [simp]: "cbox c d \ {} \ (fst ` cbox (a,c) (b,d)) = cbox a b" - by (auto simp: cbox_Pair_eq) - -lemma snd_im_cbox [simp]: "cbox a b \ {} \ (snd ` cbox (a,c) (b,d)) = cbox c d" - by (auto simp: cbox_Pair_eq) - -lemma contour_integral_swap: - assumes fcon: "continuous_on (path_image g \ path_image h) (\(y1,y2). f y1 y2)" - and vp: "valid_path g" "valid_path h" - and gvcon: "continuous_on {0..1} (\t. vector_derivative g (at t))" - and hvcon: "continuous_on {0..1} (\t. vector_derivative h (at t))" - shows "contour_integral g (\w. contour_integral h (f w)) = - contour_integral h (\z. contour_integral g (\w. f w z))" -proof - - have gcon: "continuous_on {0..1} g" and hcon: "continuous_on {0..1} h" - using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) - have fgh1: "\x. (\t. f (g x) (h t)) = (\(y1,y2). f y1 y2) o (\t. (g x, h t))" - by (rule ext) simp - have fgh2: "\x. (\t. f (g t) (h x)) = (\(y1,y2). f y1 y2) o (\t. (g t, h x))" - by (rule ext) simp - have fcon_im1: "\x. 0 \ x \ x \ 1 \ continuous_on ((\t. (g x, h t)) ` {0..1}) (\(x, y). f x y)" - by (rule continuous_on_subset [OF fcon]) (auto simp: path_image_def) - have fcon_im2: "\x. 0 \ x \ x \ 1 \ continuous_on ((\t. (g t, h x)) ` {0..1}) (\(x, y). f x y)" - by (rule continuous_on_subset [OF fcon]) (auto simp: path_image_def) - have vdg: "\y. y \ {0..1} \ (\x. f (g x) (h y) * vector_derivative g (at x)) integrable_on {0..1}" - apply (rule integrable_continuous_real) - apply (rule continuous_on_mult [OF _ gvcon]) - apply (subst fgh2) - apply (rule fcon_im2 gcon continuous_intros | simp)+ - done - have "(\z. vector_derivative g (at (fst z))) = (\x. vector_derivative g (at x)) o fst" - by auto - then have gvcon': "continuous_on (cbox (0, 0) (1, 1::real)) (\x. vector_derivative g (at (fst x)))" - apply (rule ssubst) - apply (rule continuous_intros | simp add: gvcon)+ - done - have "(\z. vector_derivative h (at (snd z))) = (\x. vector_derivative h (at x)) o snd" - by auto - then have hvcon': "continuous_on (cbox (0, 0) (1::real, 1)) (\x. vector_derivative h (at (snd x)))" - apply (rule ssubst) - apply (rule continuous_intros | simp add: hvcon)+ - done - have "(\x. f (g (fst x)) (h (snd x))) = (\(y1,y2). f y1 y2) o (\w. ((g o fst) w, (h o snd) w))" - by auto - then have fgh: "continuous_on (cbox (0, 0) (1, 1)) (\x. f (g (fst x)) (h (snd x)))" - apply (rule ssubst) - apply (rule gcon hcon continuous_intros | simp)+ - apply (auto simp: path_image_def intro: continuous_on_subset [OF fcon]) - done - have "integral {0..1} (\x. contour_integral h (f (g x)) * vector_derivative g (at x)) = - integral {0..1} (\x. contour_integral h (\y. f (g x) y * vector_derivative g (at x)))" - apply (rule integral_cong [OF contour_integral_rmul [symmetric]]) - apply (clarsimp simp: contour_integrable_on) - apply (rule integrable_continuous_real) - apply (rule continuous_on_mult [OF _ hvcon]) - apply (subst fgh1) - apply (rule fcon_im1 hcon continuous_intros | simp)+ - done - also have "... = integral {0..1} - (\y. contour_integral g (\x. f x (h y) * vector_derivative h (at y)))" - apply (simp only: contour_integral_integral) - apply (subst integral_swap_continuous [where 'a = real and 'b = real, of 0 0 1 1, simplified]) - apply (rule fgh gvcon' hvcon' continuous_intros | simp add: split_def)+ - unfolding integral_mult_left [symmetric] - apply (simp only: mult_ac) - done - also have "... = contour_integral h (\z. contour_integral g (\w. f w z))" - apply (simp add: contour_integral_integral) - apply (rule integral_cong) - unfolding integral_mult_left [symmetric] - apply (simp add: algebra_simps) - done - finally show ?thesis - by (simp add: contour_integral_integral) -qed - - -subsection\The key quadrisection step\ - -lemma norm_sum_half: - assumes "norm(a + b) >= e" - shows "norm a >= e/2 \ norm b >= e/2" -proof - - have "e \ norm (- a - b)" - by (simp add: add.commute assms norm_minus_commute) - thus ?thesis - using norm_triangle_ineq4 order_trans by fastforce -qed - -lemma norm_sum_lemma: - assumes "e \ norm (a + b + c + d)" - shows "e / 4 \ norm a \ e / 4 \ norm b \ e / 4 \ norm c \ e / 4 \ norm d" -proof - - have "e \ norm ((a + b) + (c + d))" using assms - by (simp add: algebra_simps) - then show ?thesis - by (auto dest!: norm_sum_half) -qed - -lemma Cauchy_theorem_quadrisection: - assumes f: "continuous_on (convex hull {a,b,c}) f" - and dist: "dist a b \ K" "dist b c \ K" "dist c a \ K" - and e: "e * K^2 \ - norm (contour_integral(linepath a b) f + contour_integral(linepath b c) f + contour_integral(linepath c a) f)" - shows "\a' b' c'. - a' \ convex hull {a,b,c} \ b' \ convex hull {a,b,c} \ c' \ convex hull {a,b,c} \ - dist a' b' \ K/2 \ dist b' c' \ K/2 \ dist c' a' \ K/2 \ - e * (K/2)^2 \ norm(contour_integral(linepath a' b') f + contour_integral(linepath b' c') f + contour_integral(linepath c' a') f)" -proof - - note divide_le_eq_numeral1 [simp del] - define a' where "a' = midpoint b c" - define b' where "b' = midpoint c a" - define c' where "c' = midpoint a b" - have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" - using f continuous_on_subset segments_subset_convex_hull by metis+ - have fcont': "continuous_on (closed_segment c' b') f" - "continuous_on (closed_segment a' c') f" - "continuous_on (closed_segment b' a') f" - unfolding a'_def b'_def c'_def - apply (rule continuous_on_subset [OF f], - metis midpoints_in_convex_hull convex_hull_subset hull_subset insert_subset segment_convex_hull)+ - done - let ?pathint = "\x y. contour_integral(linepath x y) f" - have *: "?pathint a b + ?pathint b c + ?pathint c a = - (?pathint a c' + ?pathint c' b' + ?pathint b' a) + - (?pathint a' c' + ?pathint c' b + ?pathint b a') + - (?pathint a' c + ?pathint c b' + ?pathint b' a') + - (?pathint a' b' + ?pathint b' c' + ?pathint c' a')" - apply (simp add: fcont' contour_integral_reverse_linepath) - apply (simp add: a'_def b'_def c'_def contour_integral_midpoint fabc) - done - have [simp]: "\x y. cmod (x * 2 - y * 2) = cmod (x - y) * 2" - by (metis left_diff_distrib mult.commute norm_mult_numeral1) - have [simp]: "\x y. cmod (x - y) = cmod (y - x)" - by (simp add: norm_minus_commute) - consider "e * K\<^sup>2 / 4 \ cmod (?pathint a c' + ?pathint c' b' + ?pathint b' a)" | - "e * K\<^sup>2 / 4 \ cmod (?pathint a' c' + ?pathint c' b + ?pathint b a')" | - "e * K\<^sup>2 / 4 \ cmod (?pathint a' c + ?pathint c b' + ?pathint b' a')" | - "e * K\<^sup>2 / 4 \ cmod (?pathint a' b' + ?pathint b' c' + ?pathint c' a')" - using assms - apply (simp only: *) - apply (blast intro: that dest!: norm_sum_lemma) - done - then show ?thesis - proof cases - case 1 then show ?thesis - apply (rule_tac x=a in exI) - apply (rule exI [where x=c']) - apply (rule exI [where x=b']) - using assms - apply (auto simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) - apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real divide_simps) - done - next - case 2 then show ?thesis - apply (rule_tac x=a' in exI) - apply (rule exI [where x=c']) - apply (rule exI [where x=b]) - using assms - apply (auto simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) - apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real divide_simps) - done - next - case 3 then show ?thesis - apply (rule_tac x=a' in exI) - apply (rule exI [where x=c]) - apply (rule exI [where x=b']) - using assms - apply (auto simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) - apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real divide_simps) - done - next - case 4 then show ?thesis - apply (rule_tac x=a' in exI) - apply (rule exI [where x=b']) - apply (rule exI [where x=c']) - using assms - apply (auto simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) - apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real divide_simps) - done - qed -qed - -subsection\Cauchy's theorem for triangles\ - -lemma triangle_points_closer: - fixes a::complex - shows "\x \ convex hull {a,b,c}; y \ convex hull {a,b,c}\ - \ norm(x - y) \ norm(a - b) \ - norm(x - y) \ norm(b - c) \ - norm(x - y) \ norm(c - a)" - using simplex_extremal_le [of "{a,b,c}"] - by (auto simp: norm_minus_commute) - -lemma holomorphic_point_small_triangle: - assumes x: "x \ s" - and f: "continuous_on s f" - and cd: "f field_differentiable (at x within s)" - and e: "0 < e" - shows "\k>0. \a b c. dist a b \ k \ dist b c \ k \ dist c a \ k \ - x \ convex hull {a,b,c} \ convex hull {a,b,c} \ s - \ norm(contour_integral(linepath a b) f + contour_integral(linepath b c) f + - contour_integral(linepath c a) f) - \ e*(dist a b + dist b c + dist c a)^2" - (is "\k>0. \a b c. _ \ ?normle a b c") -proof - - have le_of_3: "\a x y z. \0 \ x*y; 0 \ x*z; 0 \ y*z; a \ (e*(x + y + z))*x + (e*(x + y + z))*y + (e*(x + y + z))*z\ - \ a \ e*(x + y + z)^2" - by (simp add: algebra_simps power2_eq_square) - have disj_le: "\x \ a \ x \ b \ x \ c; 0 \ a; 0 \ b; 0 \ c\ \ x \ a + b + c" - for x::real and a b c - by linarith - have fabc: "f contour_integrable_on linepath a b" "f contour_integrable_on linepath b c" "f contour_integrable_on linepath c a" - if "convex hull {a, b, c} \ s" for a b c - using segments_subset_convex_hull that - by (metis continuous_on_subset f contour_integrable_continuous_linepath)+ - note path_bound = has_contour_integral_bound_linepath [simplified norm_minus_commute, OF has_contour_integral_integral] - { fix f' a b c d - assume d: "0 < d" - and f': "\y. \cmod (y - x) \ d; y \ s\ \ cmod (f y - f x - f' * (y - x)) \ e * cmod (y - x)" - and le: "cmod (a - b) \ d" "cmod (b - c) \ d" "cmod (c - a) \ d" - and xc: "x \ convex hull {a, b, c}" - and s: "convex hull {a, b, c} \ s" - have pa: "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = - contour_integral (linepath a b) (\y. f y - f x - f'*(y - x)) + - contour_integral (linepath b c) (\y. f y - f x - f'*(y - x)) + - contour_integral (linepath c a) (\y. f y - f x - f'*(y - x))" - apply (simp add: contour_integral_diff contour_integral_lmul contour_integrable_lmul contour_integrable_diff fabc [OF s]) - apply (simp add: field_simps) - done - { fix y - assume yc: "y \ convex hull {a,b,c}" - have "cmod (f y - f x - f' * (y - x)) \ e*norm(y - x)" - apply (rule f') - apply (metis triangle_points_closer [OF xc yc] le norm_minus_commute order_trans) - using s yc by blast - also have "... \ e * (cmod (a - b) + cmod (b - c) + cmod (c - a))" - by (simp add: yc e xc disj_le [OF triangle_points_closer]) - finally have "cmod (f y - f x - f' * (y - x)) \ e * (cmod (a - b) + cmod (b - c) + cmod (c - a))" . - } note cm_le = this - have "?normle a b c" - apply (simp add: dist_norm pa) - apply (rule le_of_3) - using f' xc s e - apply simp_all - apply (intro norm_triangle_le add_mono path_bound) - apply (simp_all add: contour_integral_diff contour_integral_lmul contour_integrable_lmul contour_integrable_diff fabc) - apply (blast intro: cm_le elim: dest: segments_subset_convex_hull [THEN subsetD])+ - done - } note * = this - show ?thesis - using cd e - apply (simp add: field_differentiable_def has_field_derivative_def has_derivative_within_alt approachable_lt_le2 Ball_def) - apply (clarify dest!: spec mp) - using * - apply (simp add: dist_norm, blast) - done -qed - - -(* Hence the most basic theorem for a triangle.*) -locale Chain = - fixes x0 At Follows - assumes At0: "At x0 0" - and AtSuc: "\x n. At x n \ \x'. At x' (Suc n) \ Follows x' x" -begin - primrec f where - "f 0 = x0" - | "f (Suc n) = (SOME x. At x (Suc n) \ Follows x (f n))" - - lemma At: "At (f n) n" - proof (induct n) - case 0 show ?case - by (simp add: At0) - next - case (Suc n) show ?case - by (metis (no_types, lifting) AtSuc [OF Suc] f.simps(2) someI_ex) - qed - - lemma Follows: "Follows (f(Suc n)) (f n)" - by (metis (no_types, lifting) AtSuc [OF At [of n]] f.simps(2) someI_ex) - - declare f.simps(2) [simp del] -end - -lemma Chain3: - assumes At0: "At x0 y0 z0 0" - and AtSuc: "\x y z n. At x y z n \ \x' y' z'. At x' y' z' (Suc n) \ Follows x' y' z' x y z" - obtains f g h where - "f 0 = x0" "g 0 = y0" "h 0 = z0" - "\n. At (f n) (g n) (h n) n" - "\n. Follows (f(Suc n)) (g(Suc n)) (h(Suc n)) (f n) (g n) (h n)" -proof - - interpret three: Chain "(x0,y0,z0)" "\(x,y,z). At x y z" "\(x',y',z'). \(x,y,z). Follows x' y' z' x y z" - apply unfold_locales - using At0 AtSuc by auto - show ?thesis - apply (rule that [of "\n. fst (three.f n)" "\n. fst (snd (three.f n))" "\n. snd (snd (three.f n))"]) - apply simp_all - using three.At three.Follows - apply (simp_all add: split_beta') - done -qed - -lemma Cauchy_theorem_triangle: - assumes "f holomorphic_on (convex hull {a,b,c})" - shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" -proof - - have contf: "continuous_on (convex hull {a,b,c}) f" - by (metis assms holomorphic_on_imp_continuous_on) - let ?pathint = "\x y. contour_integral(linepath x y) f" - { fix y::complex - assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)" - and ynz: "y \ 0" - define K where "K = 1 + max (dist a b) (max (dist b c) (dist c a))" - define e where "e = norm y / K^2" - have K1: "K \ 1" by (simp add: K_def max.coboundedI1) - then have K: "K > 0" by linarith - have [iff]: "dist a b \ K" "dist b c \ K" "dist c a \ K" - by (simp_all add: K_def) - have e: "e > 0" - unfolding e_def using ynz K1 by simp - define At where "At x y z n \ - convex hull {x,y,z} \ convex hull {a,b,c} \ - dist x y \ K/2^n \ dist y z \ K/2^n \ dist z x \ K/2^n \ - norm(?pathint x y + ?pathint y z + ?pathint z x) \ e*(K/2^n)^2" - for x y z n - have At0: "At a b c 0" - using fy - by (simp add: At_def e_def has_chain_integral_chain_integral3) - { fix x y z n - assume At: "At x y z n" - then have contf': "continuous_on (convex hull {x,y,z}) f" - using contf At_def continuous_on_subset by blast - have "\x' y' z'. At x' y' z' (Suc n) \ convex hull {x',y',z'} \ convex hull {x,y,z}" - using At - apply (simp add: At_def) - using Cauchy_theorem_quadrisection [OF contf', of "K/2^n" e] - apply clarsimp - apply (rule_tac x="a'" in exI) - apply (rule_tac x="b'" in exI) - apply (rule_tac x="c'" in exI) - apply (simp add: algebra_simps) - apply (meson convex_hull_subset empty_subsetI insert_subset subsetCE) - done - } note AtSuc = this - obtain fa fb fc - where f0 [simp]: "fa 0 = a" "fb 0 = b" "fc 0 = c" - and cosb: "\n. convex hull {fa n, fb n, fc n} \ convex hull {a,b,c}" - and dist: "\n. dist (fa n) (fb n) \ K/2^n" - "\n. dist (fb n) (fc n) \ K/2^n" - "\n. dist (fc n) (fa n) \ K/2^n" - and no: "\n. norm(?pathint (fa n) (fb n) + - ?pathint (fb n) (fc n) + - ?pathint (fc n) (fa n)) \ e * (K/2^n)^2" - and conv_le: "\n. convex hull {fa(Suc n), fb(Suc n), fc(Suc n)} \ convex hull {fa n, fb n, fc n}" - apply (rule Chain3 [of At, OF At0 AtSuc]) - apply (auto simp: At_def) - done - have "\x. \n. x \ convex hull {fa n, fb n, fc n}" - apply (rule bounded_closed_nest) - apply (simp_all add: compact_imp_closed finite_imp_compact_convex_hull finite_imp_bounded_convex_hull) - apply (rule allI) - apply (rule transitive_stepwise_le) - apply (auto simp: conv_le) - done - then obtain x where x: "\n. x \ convex hull {fa n, fb n, fc n}" by auto - then have xin: "x \ convex hull {a,b,c}" - using assms f0 by blast - then have fx: "f field_differentiable at x within (convex hull {a,b,c})" - using assms holomorphic_on_def by blast - { fix k n - assume k: "0 < k" - and le: - "\x' y' z'. - \dist x' y' \ k; dist y' z' \ k; dist z' x' \ k; - x \ convex hull {x',y',z'}; - convex hull {x',y',z'} \ convex hull {a,b,c}\ - \ - cmod (?pathint x' y' + ?pathint y' z' + ?pathint z' x') * 10 - \ e * (dist x' y' + dist y' z' + dist z' x')\<^sup>2" - and Kk: "K / k < 2 ^ n" - have "K / 2 ^ n < k" using Kk k - by (auto simp: field_simps) - then have DD: "dist (fa n) (fb n) \ k" "dist (fb n) (fc n) \ k" "dist (fc n) (fa n) \ k" - using dist [of n] k - by linarith+ - have dle: "(dist (fa n) (fb n) + dist (fb n) (fc n) + dist (fc n) (fa n))\<^sup>2 - \ (3 * K / 2 ^ n)\<^sup>2" - using dist [of n] e K - by (simp add: abs_le_square_iff [symmetric]) - have less10: "\x y::real. 0 < x \ y \ 9*x \ y < x*10" - by linarith - have "e * (dist (fa n) (fb n) + dist (fb n) (fc n) + dist (fc n) (fa n))\<^sup>2 \ e * (3 * K / 2 ^ n)\<^sup>2" - using ynz dle e mult_le_cancel_left_pos by blast - also have "... < - cmod (?pathint (fa n) (fb n) + ?pathint (fb n) (fc n) + ?pathint (fc n) (fa n)) * 10" - using no [of n] e K - apply (simp add: e_def field_simps) - apply (simp only: zero_less_norm_iff [symmetric]) - done - finally have False - using le [OF DD x cosb] by auto - } then - have ?thesis - using holomorphic_point_small_triangle [OF xin contf fx, of "e/10"] e - apply clarsimp - apply (rule_tac y1="K/k" in exE [OF real_arch_pow[of 2]]) - apply force+ - done - } - moreover have "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)" - by simp (meson contf continuous_on_subset contour_integrable_continuous_linepath segments_subset_convex_hull(1) - segments_subset_convex_hull(3) segments_subset_convex_hull(5)) - ultimately show ?thesis - using has_contour_integral_integral by fastforce -qed - - -subsection\Version needing function holomorphic in interior only\ - -lemma Cauchy_theorem_flat_lemma: - assumes f: "continuous_on (convex hull {a,b,c}) f" - and c: "c - a = k *\<^sub>R (b - a)" - and k: "0 \ k" - shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f + - contour_integral (linepath c a) f = 0" -proof - - have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" - using f continuous_on_subset segments_subset_convex_hull by metis+ - show ?thesis - proof (cases "k \ 1") - case True show ?thesis - by (simp add: contour_integral_split [OF fabc(1) k True c] contour_integral_reverse_linepath fabc) - next - case False then show ?thesis - using fabc c - apply (subst contour_integral_split [of a c f "1/k" b, symmetric]) - apply (metis closed_segment_commute fabc(3)) - apply (auto simp: k contour_integral_reverse_linepath) - done - qed -qed - -lemma Cauchy_theorem_flat: - assumes f: "continuous_on (convex hull {a,b,c}) f" - and c: "c - a = k *\<^sub>R (b - a)" - shows "contour_integral (linepath a b) f + - contour_integral (linepath b c) f + - contour_integral (linepath c a) f = 0" -proof (cases "0 \ k") - case True with assms show ?thesis - by (blast intro: Cauchy_theorem_flat_lemma) -next - case False - have "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" - using f continuous_on_subset segments_subset_convex_hull by metis+ - moreover have "contour_integral (linepath b a) f + contour_integral (linepath a c) f + - contour_integral (linepath c b) f = 0" - apply (rule Cauchy_theorem_flat_lemma [of b a c f "1-k"]) - using False c - apply (auto simp: f insert_commute scaleR_conv_of_real algebra_simps) - done - ultimately show ?thesis - apply (auto simp: contour_integral_reverse_linepath) - using add_eq_0_iff by force -qed - - -lemma Cauchy_theorem_triangle_interior: - assumes contf: "continuous_on (convex hull {a,b,c}) f" - and holf: "f holomorphic_on interior (convex hull {a,b,c})" - shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" -proof - - have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" - using contf continuous_on_subset segments_subset_convex_hull by metis+ - have "bounded (f ` (convex hull {a,b,c}))" - by (simp add: compact_continuous_image compact_convex_hull compact_imp_bounded contf) - then obtain B where "0 < B" and Bnf: "\x. x \ convex hull {a,b,c} \ norm (f x) \ B" - by (auto simp: dest!: bounded_pos [THEN iffD1]) - have "bounded (convex hull {a,b,c})" - by (simp add: bounded_convex_hull) - then obtain C where C: "0 < C" and Cno: "\y. y \ convex hull {a,b,c} \ norm y < C" - using bounded_pos_less by blast - then have diff_2C: "norm(x - y) \ 2*C" - if x: "x \ convex hull {a, b, c}" and y: "y \ convex hull {a, b, c}" for x y - proof - - have "cmod x \ C" - using x by (meson Cno not_le not_less_iff_gr_or_eq) - hence "cmod (x - y) \ C + C" - using y by (meson Cno add_mono_thms_linordered_field(4) less_eq_real_def norm_triangle_ineq4 order_trans) - thus "cmod (x - y) \ 2 * C" - by (metis mult_2) - qed - have contf': "continuous_on (convex hull {b,a,c}) f" - using contf by (simp add: insert_commute) - { fix y::complex - assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)" - and ynz: "y \ 0" - have pi_eq_y: "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = y" - by (rule has_chain_integral_chain_integral3 [OF fy]) - have ?thesis - proof (cases "c=a \ a=b \ b=c") - case True then show ?thesis - using Cauchy_theorem_flat [OF contf, of 0] - using has_chain_integral_chain_integral3 [OF fy] ynz - by (force simp: fabc contour_integral_reverse_linepath) - next - case False - then have car3: "card {a, b, c} = Suc (DIM(complex))" - by auto - { assume "interior(convex hull {a,b,c}) = {}" - then have "collinear{a,b,c}" - using interior_convex_hull_eq_empty [OF car3] - by (simp add: collinear_3_eq_affine_dependent) - then have "False" - using False - apply (clarsimp simp add: collinear_3 collinear_lemma) - apply (drule Cauchy_theorem_flat [OF contf']) - using pi_eq_y ynz - apply (simp add: fabc add_eq_0_iff contour_integral_reverse_linepath) - done - } - then obtain d where d: "d \ interior (convex hull {a, b, c})" - by blast - { fix d1 - assume d1_pos: "0 < d1" - and d1: "\x x'. \x\convex hull {a, b, c}; x'\convex hull {a, b, c}; cmod (x' - x) < d1\ - \ cmod (f x' - f x) < cmod y / (24 * C)" - define e where "e = min 1 (min (d1/(4*C)) ((norm y / 24 / C) / B))" - define shrink where "shrink x = x - e *\<^sub>R (x - d)" for x - let ?pathint = "\x y. contour_integral(linepath x y) f" - have e: "0 < e" "e \ 1" "e \ d1 / (4 * C)" "e \ cmod y / 24 / C / B" - using d1_pos \C>0\ \B>0\ ynz by (simp_all add: e_def) - then have eCB: "24 * e * C * B \ cmod y" - using \C>0\ \B>0\ by (simp add: field_simps) - have e_le_d1: "e * (4 * C) \ d1" - using e \C>0\ by (simp add: field_simps) - have "shrink a \ interior(convex hull {a,b,c})" - "shrink b \ interior(convex hull {a,b,c})" - "shrink c \ interior(convex hull {a,b,c})" - using d e by (auto simp: hull_inc mem_interior_convex_shrink shrink_def) - then have fhp0: "(f has_contour_integral 0) - (linepath (shrink a) (shrink b) +++ linepath (shrink b) (shrink c) +++ linepath (shrink c) (shrink a))" - by (simp add: Cauchy_theorem_triangle holomorphic_on_subset [OF holf] hull_minimal convex_interior) - then have f_0_shrink: "?pathint (shrink a) (shrink b) + ?pathint (shrink b) (shrink c) + ?pathint (shrink c) (shrink a) = 0" - by (simp add: has_chain_integral_chain_integral3) - have fpi_abc: "f contour_integrable_on linepath (shrink a) (shrink b)" - "f contour_integrable_on linepath (shrink b) (shrink c)" - "f contour_integrable_on linepath (shrink c) (shrink a)" - using fhp0 by (auto simp: valid_path_join dest: has_contour_integral_integrable) - have cmod_shr: "\x y. cmod (shrink y - shrink x - (y - x)) = e * cmod (x - y)" - using e by (simp add: shrink_def real_vector.scale_right_diff_distrib [symmetric]) - have sh_eq: "\a b d::complex. (b - e *\<^sub>R (b - d)) - (a - e *\<^sub>R (a - d)) - (b - a) = e *\<^sub>R (a - b)" - by (simp add: algebra_simps) - have "cmod y / (24 * C) \ cmod y / cmod (b - a) / 12" - using False \C>0\ diff_2C [of b a] ynz - by (auto simp: divide_simps hull_inc) - have less_C: "\u \ convex hull {a, b, c}; 0 \ x; x \ 1\ \ x * cmod u < C" for x u - apply (cases "x=0", simp add: \0) - using Cno [of u] mult_left_le_one_le [of "cmod u" x] le_less_trans norm_ge_zero by blast - { fix u v - assume uv: "u \ convex hull {a, b, c}" "v \ convex hull {a, b, c}" "u\v" - and fpi_uv: "f contour_integrable_on linepath (shrink u) (shrink v)" - have shr_uv: "shrink u \ interior(convex hull {a,b,c})" - "shrink v \ interior(convex hull {a,b,c})" - using d e uv - by (auto simp: hull_inc mem_interior_convex_shrink shrink_def) - have cmod_fuv: "\x. 0\x \ x\1 \ cmod (f (linepath (shrink u) (shrink v) x)) \ B" - using shr_uv by (blast intro: Bnf linepath_in_convex_hull interior_subset [THEN subsetD]) - have By_uv: "B * (12 * (e * cmod (u - v))) \ cmod y" - apply (rule order_trans [OF _ eCB]) - using e \B>0\ diff_2C [of u v] uv - by (auto simp: field_simps) - { fix x::real assume x: "0\x" "x\1" - have cmod_less_4C: "cmod ((1 - x) *\<^sub>R u - (1 - x) *\<^sub>R d) + cmod (x *\<^sub>R v - x *\<^sub>R d) < (C+C) + (C+C)" - apply (rule add_strict_mono; rule norm_triangle_half_l [of _ 0]) - using uv x d interior_subset - apply (auto simp: hull_inc intro!: less_C) - done - have ll: "linepath (shrink u) (shrink v) x - linepath u v x = -e * ((1 - x) *\<^sub>R (u - d) + x *\<^sub>R (v - d))" - by (simp add: linepath_def shrink_def algebra_simps scaleR_conv_of_real) - have cmod_less_dt: "cmod (linepath (shrink u) (shrink v) x - linepath u v x) < d1" - using \e>0\ - apply (simp add: ll norm_mult scaleR_diff_right) - apply (rule less_le_trans [OF _ e_le_d1]) - using cmod_less_4C - apply (force intro: norm_triangle_lt) - done - have "cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) < cmod y / (24 * C)" - using x uv shr_uv cmod_less_dt - by (auto simp: hull_inc intro: d1 interior_subset [THEN subsetD] linepath_in_convex_hull) - also have "... \ cmod y / cmod (v - u) / 12" - using False uv \C>0\ diff_2C [of v u] ynz - by (auto simp: divide_simps hull_inc) - finally have "cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) \ cmod y / cmod (v - u) / 12" - by simp - then have cmod_12_le: "cmod (v - u) * cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) * 12 \ cmod y" - using uv False by (auto simp: field_simps) - have "cmod (f (linepath (shrink u) (shrink v) x)) * cmod (shrink v - shrink u - (v - u)) + - cmod (v - u) * cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) - \ cmod y / 6" - apply (rule order_trans [of _ "B*((norm y / 24 / C / B)*2*C) + (2*C)*(norm y /24 / C)"]) - apply (rule add_mono [OF mult_mono]) - using By_uv e \0 < B\ \0 < C\ x ynz - apply (simp_all add: cmod_fuv cmod_shr cmod_12_le hull_inc) - apply (simp add: field_simps) - done - } note cmod_diff_le = this - have f_uv: "continuous_on (closed_segment u v) f" - by (blast intro: uv continuous_on_subset [OF contf closed_segment_subset_convex_hull]) - have **: "\f' x' f x::complex. f'*x' - f*x = f'*(x' - x) + x*(f' - f)" - by (simp add: algebra_simps) - have "norm (?pathint (shrink u) (shrink v) - ?pathint u v) \ norm y / 6" - apply (rule order_trans) - apply (rule has_integral_bound - [of "B*(norm y /24/C/B)*2*C + (2*C)*(norm y/24/C)" - "\x. f(linepath (shrink u) (shrink v) x) * (shrink v - shrink u) - f(linepath u v x)*(v - u)" - _ 0 1 ]) - using ynz \0 < B\ \0 < C\ - apply (simp_all del: le_divide_eq_numeral1) - apply (simp add: has_integral_sub has_contour_integral_linepath [symmetric] has_contour_integral_integral - fpi_uv f_uv contour_integrable_continuous_linepath, clarify) - apply (simp only: **) - apply (simp add: norm_triangle_le norm_mult cmod_diff_le del: le_divide_eq_numeral1) - done - } note * = this - have "norm (?pathint (shrink a) (shrink b) - ?pathint a b) \ norm y / 6" - using False fpi_abc by (rule_tac *) (auto simp: hull_inc) - moreover - have "norm (?pathint (shrink b) (shrink c) - ?pathint b c) \ norm y / 6" - using False fpi_abc by (rule_tac *) (auto simp: hull_inc) - moreover - have "norm (?pathint (shrink c) (shrink a) - ?pathint c a) \ norm y / 6" - using False fpi_abc by (rule_tac *) (auto simp: hull_inc) - ultimately - have "norm((?pathint (shrink a) (shrink b) - ?pathint a b) + - (?pathint (shrink b) (shrink c) - ?pathint b c) + (?pathint (shrink c) (shrink a) - ?pathint c a)) - \ norm y / 6 + norm y / 6 + norm y / 6" - by (metis norm_triangle_le add_mono) - also have "... = norm y / 2" - by simp - finally have "norm((?pathint (shrink a) (shrink b) + ?pathint (shrink b) (shrink c) + ?pathint (shrink c) (shrink a)) - - (?pathint a b + ?pathint b c + ?pathint c a)) - \ norm y / 2" - by (simp add: algebra_simps) - then - have "norm(?pathint a b + ?pathint b c + ?pathint c a) \ norm y / 2" - by (simp add: f_0_shrink) (metis (mono_tags) add.commute minus_add_distrib norm_minus_cancel uminus_add_conv_diff) - then have "False" - using pi_eq_y ynz by auto - } - moreover have "uniformly_continuous_on (convex hull {a,b,c}) f" - by (simp add: contf compact_convex_hull compact_uniformly_continuous) - ultimately have "False" - unfolding uniformly_continuous_on_def - by (force simp: ynz \0 < C\ dist_norm) - then show ?thesis .. - qed - } - moreover have "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)" - using fabc contour_integrable_continuous_linepath by auto - ultimately show ?thesis - using has_contour_integral_integral by fastforce -qed - - -subsection\Version allowing finite number of exceptional points\ - -lemma Cauchy_theorem_triangle_cofinite: - assumes "continuous_on (convex hull {a,b,c}) f" - and "finite s" - and "(\x. x \ interior(convex hull {a,b,c}) - s \ f field_differentiable (at x))" - shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" -using assms -proof (induction "card s" arbitrary: a b c s rule: less_induct) - case (less s a b c) - show ?case - proof (cases "s={}") - case True with less show ?thesis - by (fastforce simp: holomorphic_on_def field_differentiable_at_within - Cauchy_theorem_triangle_interior) - next - case False - then obtain d s' where d: "s = insert d s'" "d \ s'" - by (meson Set.set_insert all_not_in_conv) - then show ?thesis - proof (cases "d \ convex hull {a,b,c}") - case False - show "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" - apply (rule less.hyps [of "s'"]) - using False d \finite s\ interior_subset - apply (auto intro!: less.prems) - done - next - case True - have *: "convex hull {a, b, d} \ convex hull {a, b, c}" - by (meson True hull_subset insert_subset convex_hull_subset) - have abd: "(f has_contour_integral 0) (linepath a b +++ linepath b d +++ linepath d a)" - apply (rule less.hyps [of "s'"]) - using True d \finite s\ not_in_interior_convex_hull_3 - apply (auto intro!: less.prems continuous_on_subset [OF _ *]) - apply (metis * insert_absorb insert_subset interior_mono) - done - have *: "convex hull {b, c, d} \ convex hull {a, b, c}" - by (meson True hull_subset insert_subset convex_hull_subset) - have bcd: "(f has_contour_integral 0) (linepath b c +++ linepath c d +++ linepath d b)" - apply (rule less.hyps [of "s'"]) - using True d \finite s\ not_in_interior_convex_hull_3 - apply (auto intro!: less.prems continuous_on_subset [OF _ *]) - apply (metis * insert_absorb insert_subset interior_mono) - done - have *: "convex hull {c, a, d} \ convex hull {a, b, c}" - by (meson True hull_subset insert_subset convex_hull_subset) - have cad: "(f has_contour_integral 0) (linepath c a +++ linepath a d +++ linepath d c)" - apply (rule less.hyps [of "s'"]) - using True d \finite s\ not_in_interior_convex_hull_3 - apply (auto intro!: less.prems continuous_on_subset [OF _ *]) - apply (metis * insert_absorb insert_subset interior_mono) - done - have "f contour_integrable_on linepath a b" - using less.prems - by (metis continuous_on_subset insert_commute contour_integrable_continuous_linepath segments_subset_convex_hull(3)) - moreover have "f contour_integrable_on linepath b c" - using less.prems - by (metis continuous_on_subset contour_integrable_continuous_linepath segments_subset_convex_hull(3)) - moreover have "f contour_integrable_on linepath c a" - using less.prems - by (metis continuous_on_subset insert_commute contour_integrable_continuous_linepath segments_subset_convex_hull(3)) - ultimately have fpi: "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)" - by auto - { fix y::complex - assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)" - and ynz: "y \ 0" - have cont_ad: "continuous_on (closed_segment a d) f" - by (meson "*" continuous_on_subset less.prems(1) segments_subset_convex_hull(3)) - have cont_bd: "continuous_on (closed_segment b d) f" - by (meson True closed_segment_subset_convex_hull continuous_on_subset hull_subset insert_subset less.prems(1)) - have cont_cd: "continuous_on (closed_segment c d) f" - by (meson "*" continuous_on_subset less.prems(1) segments_subset_convex_hull(2)) - have "contour_integral (linepath a b) f = - (contour_integral (linepath b d) f + (contour_integral (linepath d a) f))" - "contour_integral (linepath b c) f = - (contour_integral (linepath c d) f + (contour_integral (linepath d b) f))" - "contour_integral (linepath c a) f = - (contour_integral (linepath a d) f + contour_integral (linepath d c) f)" - using has_chain_integral_chain_integral3 [OF abd] - has_chain_integral_chain_integral3 [OF bcd] - has_chain_integral_chain_integral3 [OF cad] - by (simp_all add: algebra_simps add_eq_0_iff) - then have ?thesis - using cont_ad cont_bd cont_cd fy has_chain_integral_chain_integral3 contour_integral_reverse_linepath by fastforce - } - then show ?thesis - using fpi contour_integrable_on_def by blast - qed - qed -qed - - -subsection\Cauchy's theorem for an open starlike set\ - -lemma starlike_convex_subset: - assumes s: "a \ s" "closed_segment b c \ s" and subs: "\x. x \ s \ closed_segment a x \ s" - shows "convex hull {a,b,c} \ s" - using s - apply (clarsimp simp add: convex_hull_insert [of "{b,c}" a] segment_convex_hull) - apply (meson subs convexD convex_closed_segment ends_in_segment(1) ends_in_segment(2) subsetCE) - done - -lemma triangle_contour_integrals_starlike_primitive: - assumes contf: "continuous_on s f" - and s: "a \ s" "open s" - and x: "x \ s" - and subs: "\y. y \ s \ closed_segment a y \ s" - and zer: "\b c. closed_segment b c \ s - \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + - contour_integral (linepath c a) f = 0" - shows "((\x. contour_integral(linepath a x) f) has_field_derivative f x) (at x)" -proof - - let ?pathint = "\x y. contour_integral(linepath x y) f" - { fix e y - assume e: "0 < e" and bxe: "ball x e \ s" and close: "cmod (y - x) < e" - have y: "y \ s" - using bxe close by (force simp: dist_norm norm_minus_commute) - have cont_ayf: "continuous_on (closed_segment a y) f" - using contf continuous_on_subset subs y by blast - have xys: "closed_segment x y \ s" - apply (rule order_trans [OF _ bxe]) - using close - by (auto simp: dist_norm ball_def norm_minus_commute dest: segment_bound) - have "?pathint a y - ?pathint a x = ?pathint x y" - using zer [OF xys] contour_integral_reverse_linepath [OF cont_ayf] add_eq_0_iff by force - } note [simp] = this - { fix e::real - assume e: "0 < e" - have cont_atx: "continuous (at x) f" - using x s contf continuous_on_eq_continuous_at by blast - then obtain d1 where d1: "d1>0" and d1_less: "\y. cmod (y - x) < d1 \ cmod (f y - f x) < e/2" - unfolding continuous_at Lim_at dist_norm using e - by (drule_tac x="e/2" in spec) force - obtain d2 where d2: "d2>0" "ball x d2 \ s" using \open s\ x - by (auto simp: open_contains_ball) - have dpos: "min d1 d2 > 0" using d1 d2 by simp - { fix y - assume yx: "y \ x" and close: "cmod (y - x) < min d1 d2" - have y: "y \ s" - using d2 close by (force simp: dist_norm norm_minus_commute) - have fxy: "f contour_integrable_on linepath x y" - apply (rule contour_integrable_continuous_linepath) - apply (rule continuous_on_subset [OF contf]) - using close d2 - apply (auto simp: dist_norm norm_minus_commute dest!: segment_bound(1)) - done - then obtain i where i: "(f has_contour_integral i) (linepath x y)" - by (auto simp: contour_integrable_on_def) - then have "((\w. f w - f x) has_contour_integral (i - f x * (y - x))) (linepath x y)" - by (rule has_contour_integral_diff [OF _ has_contour_integral_const_linepath]) - then have "cmod (i - f x * (y - x)) \ e / 2 * cmod (y - x)" - apply (rule has_contour_integral_bound_linepath [where B = "e/2"]) - using e apply simp - apply (rule d1_less [THEN less_imp_le]) - using close segment_bound - apply force - done - also have "... < e * cmod (y - x)" - by (simp add: e yx) - finally have "cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e" - using i yx by (simp add: contour_integral_unique divide_less_eq) - } - then have "\d>0. \y. y \ x \ cmod (y-x) < d \ cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e" - using dpos by blast - } - then have *: "(\y. (?pathint x y - f x * (y - x)) /\<^sub>R cmod (y - x)) \x\ 0" - by (simp add: Lim_at dist_norm inverse_eq_divide) - show ?thesis - apply (simp add: has_field_derivative_def has_derivative_at bounded_linear_mult_right) - apply (rule Lim_transform [OF * Lim_eventually]) - apply (simp add: inverse_eq_divide [symmetric] eventually_at) - using \open s\ x - apply (force simp: dist_norm open_contains_ball) - done -qed - -(** Existence of a primitive.*) - -lemma holomorphic_starlike_primitive: - fixes f :: "complex \ complex" - assumes contf: "continuous_on s f" - and s: "starlike s" and os: "open s" - and k: "finite k" - and fcd: "\x. x \ s - k \ f field_differentiable at x" - shows "\g. \x \ s. (g has_field_derivative f x) (at x)" -proof - - obtain a where a: "a\s" and a_cs: "\x. x\s \ closed_segment a x \ s" - using s by (auto simp: starlike_def) - { fix x b c - assume "x \ s" "closed_segment b c \ s" - then have abcs: "convex hull {a, b, c} \ s" - by (simp add: a a_cs starlike_convex_subset) - then have *: "continuous_on (convex hull {a, b, c}) f" - by (simp add: continuous_on_subset [OF contf]) - have "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" - apply (rule Cauchy_theorem_triangle_cofinite [OF _ k]) - using abcs apply (simp add: continuous_on_subset [OF contf]) - using * abcs interior_subset apply (auto intro: fcd) - done - } note 0 = this - show ?thesis - apply (intro exI ballI) - apply (rule triangle_contour_integrals_starlike_primitive [OF contf a os], assumption) - apply (metis a_cs) - apply (metis has_chain_integral_chain_integral3 0) - done -qed - -lemma Cauchy_theorem_starlike: - "\open s; starlike s; finite k; continuous_on s f; - \x. x \ s - k \ f field_differentiable at x; - valid_path g; path_image g \ s; pathfinish g = pathstart g\ - \ (f has_contour_integral 0) g" - by (metis holomorphic_starlike_primitive Cauchy_theorem_primitive at_within_open) - -lemma Cauchy_theorem_starlike_simple: - "\open s; starlike s; f holomorphic_on s; valid_path g; path_image g \ s; pathfinish g = pathstart g\ - \ (f has_contour_integral 0) g" -apply (rule Cauchy_theorem_starlike [OF _ _ finite.emptyI]) -apply (simp_all add: holomorphic_on_imp_continuous_on) -apply (metis at_within_open holomorphic_on_def) -done - - -subsection\Cauchy's theorem for a convex set\ - -text\For a convex set we can avoid assuming openness and boundary analyticity\ - -lemma triangle_contour_integrals_convex_primitive: - assumes contf: "continuous_on s f" - and s: "a \ s" "convex s" - and x: "x \ s" - and zer: "\b c. \b \ s; c \ s\ - \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + - contour_integral (linepath c a) f = 0" - shows "((\x. contour_integral(linepath a x) f) has_field_derivative f x) (at x within s)" -proof - - let ?pathint = "\x y. contour_integral(linepath x y) f" - { fix y - assume y: "y \ s" - have cont_ayf: "continuous_on (closed_segment a y) f" - using s y by (meson contf continuous_on_subset convex_contains_segment) - have xys: "closed_segment x y \ s" (*?*) - using convex_contains_segment s x y by auto - have "?pathint a y - ?pathint a x = ?pathint x y" - using zer [OF x y] contour_integral_reverse_linepath [OF cont_ayf] add_eq_0_iff by force - } note [simp] = this - { fix e::real - assume e: "0 < e" - have cont_atx: "continuous (at x within s) f" - using x s contf by (simp add: continuous_on_eq_continuous_within) - then obtain d1 where d1: "d1>0" and d1_less: "\y. \y \ s; cmod (y - x) < d1\ \ cmod (f y - f x) < e/2" - unfolding continuous_within Lim_within dist_norm using e - by (drule_tac x="e/2" in spec) force - { fix y - assume yx: "y \ x" and close: "cmod (y - x) < d1" and y: "y \ s" - have fxy: "f contour_integrable_on linepath x y" - using convex_contains_segment s x y - by (blast intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf]) - then obtain i where i: "(f has_contour_integral i) (linepath x y)" - by (auto simp: contour_integrable_on_def) - then have "((\w. f w - f x) has_contour_integral (i - f x * (y - x))) (linepath x y)" - by (rule has_contour_integral_diff [OF _ has_contour_integral_const_linepath]) - then have "cmod (i - f x * (y - x)) \ e / 2 * cmod (y - x)" - apply (rule has_contour_integral_bound_linepath [where B = "e/2"]) - using e apply simp - apply (rule d1_less [THEN less_imp_le]) - using convex_contains_segment s(2) x y apply blast - using close segment_bound(1) apply fastforce - done - also have "... < e * cmod (y - x)" - by (simp add: e yx) - finally have "cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e" - using i yx by (simp add: contour_integral_unique divide_less_eq) - } - then have "\d>0. \y\s. y \ x \ cmod (y-x) < d \ cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e" - using d1 by blast - } - then have *: "((\y. (contour_integral (linepath x y) f - f x * (y - x)) /\<^sub>R cmod (y - x)) \ 0) (at x within s)" - by (simp add: Lim_within dist_norm inverse_eq_divide) - show ?thesis - apply (simp add: has_field_derivative_def has_derivative_within bounded_linear_mult_right) - apply (rule Lim_transform [OF * Lim_eventually]) - using linordered_field_no_ub - apply (force simp: inverse_eq_divide [symmetric] eventually_at) - done -qed - -lemma contour_integral_convex_primitive: - "\convex s; continuous_on s f; - \a b c. \a \ s; b \ s; c \ s\ \ (f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)\ - \ \g. \x \ s. (g has_field_derivative f x) (at x within s)" - apply (cases "s={}") - apply (simp_all add: ex_in_conv [symmetric]) - apply (blast intro: triangle_contour_integrals_convex_primitive has_chain_integral_chain_integral3) - done - -lemma holomorphic_convex_primitive: - fixes f :: "complex \ complex" - shows - "\convex s; finite k; continuous_on s f; - \x. x \ interior s - k \ f field_differentiable at x\ - \ \g. \x \ s. (g has_field_derivative f x) (at x within s)" -apply (rule contour_integral_convex_primitive [OF _ _ Cauchy_theorem_triangle_cofinite]) -prefer 3 -apply (erule continuous_on_subset) -apply (simp add: subset_hull continuous_on_subset, assumption+) -by (metis Diff_iff convex_contains_segment insert_absorb insert_subset interior_mono segment_convex_hull subset_hull) - -lemma Cauchy_theorem_convex: - "\continuous_on s f; convex s; finite k; - \x. x \ interior s - k \ f field_differentiable at x; - valid_path g; path_image g \ s; - pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" - by (metis holomorphic_convex_primitive Cauchy_theorem_primitive) - -lemma Cauchy_theorem_convex_simple: - "\f holomorphic_on s; convex s; - valid_path g; path_image g \ s; - pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" - apply (rule Cauchy_theorem_convex) - apply (simp_all add: holomorphic_on_imp_continuous_on) - apply (rule finite.emptyI) - using at_within_interior holomorphic_on_def interior_subset by fastforce - - -text\In particular for a disc\ -lemma Cauchy_theorem_disc: - "\finite k; continuous_on (cball a e) f; - \x. x \ ball a e - k \ f field_differentiable at x; - valid_path g; path_image g \ cball a e; - pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" - apply (rule Cauchy_theorem_convex) - apply (auto simp: convex_cball interior_cball) - done - -lemma Cauchy_theorem_disc_simple: - "\f holomorphic_on (ball a e); valid_path g; path_image g \ ball a e; - pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" -by (simp add: Cauchy_theorem_convex_simple) - - -subsection\Generalize integrability to local primitives\ - -lemma contour_integral_local_primitive_lemma: - fixes f :: "complex\complex" - shows - "\g piecewise_differentiable_on {a..b}; - \x. x \ s \ (f has_field_derivative f' x) (at x within s); - \x. x \ {a..b} \ g x \ s\ - \ (\x. f' (g x) * vector_derivative g (at x within {a..b})) - integrable_on {a..b}" - apply (cases "cbox a b = {}", force) - apply (simp add: integrable_on_def) - apply (rule exI) - apply (rule contour_integral_primitive_lemma, assumption+) - using atLeastAtMost_iff by blast - -lemma contour_integral_local_primitive_any: - fixes f :: "complex \ complex" - assumes gpd: "g piecewise_differentiable_on {a..b}" - and dh: "\x. x \ s - \ \d h. 0 < d \ - (\y. norm(y - x) < d \ (h has_field_derivative f y) (at y within s))" - and gs: "\x. x \ {a..b} \ g x \ s" - shows "(\x. f(g x) * vector_derivative g (at x)) integrable_on {a..b}" -proof - - { fix x - assume x: "a \ x" "x \ b" - obtain d h where d: "0 < d" - and h: "(\y. norm(y - g x) < d \ (h has_field_derivative f y) (at y within s))" - using x gs dh by (metis atLeastAtMost_iff) - have "continuous_on {a..b} g" using gpd piecewise_differentiable_on_def by blast - then obtain e where e: "e>0" and lessd: "\x'. x' \ {a..b} \ \x' - x\ < e \ cmod (g x' - g x) < d" - using x d - apply (auto simp: dist_norm continuous_on_iff) - apply (drule_tac x=x in bspec) - using x apply simp - apply (drule_tac x=d in spec, auto) - done - have "\d>0. \u v. u \ x \ x \ v \ {u..v} \ ball x d \ (u \ v \ a \ u \ v \ b) \ - (\x. f (g x) * vector_derivative g (at x)) integrable_on {u..v}" - apply (rule_tac x=e in exI) - using e - apply (simp add: integrable_on_localized_vector_derivative [symmetric], clarify) - apply (rule_tac f = h and s = "g ` {u..v}" in contour_integral_local_primitive_lemma) - apply (meson atLeastatMost_subset_iff gpd piecewise_differentiable_on_subset) - apply (force simp: ball_def dist_norm intro: lessd gs DERIV_subset [OF h], force) - done - } then - show ?thesis - by (force simp: intro!: integrable_on_little_subintervals [of a b, simplified]) -qed - -lemma contour_integral_local_primitive: - fixes f :: "complex \ complex" - assumes g: "valid_path g" "path_image g \ s" - and dh: "\x. x \ s - \ \d h. 0 < d \ - (\y. norm(y - x) < d \ (h has_field_derivative f y) (at y within s))" - shows "f contour_integrable_on g" - using g - apply (simp add: valid_path_def path_image_def contour_integrable_on_def has_contour_integral_def - has_integral_localized_vector_derivative integrable_on_def [symmetric]) - using contour_integral_local_primitive_any [OF _ dh] - by (meson image_subset_iff piecewise_C1_imp_differentiable) - - -text\In particular if a function is holomorphic\ - -lemma contour_integrable_holomorphic: - assumes contf: "continuous_on s f" - and os: "open s" - and k: "finite k" - and g: "valid_path g" "path_image g \ s" - and fcd: "\x. x \ s - k \ f field_differentiable at x" - shows "f contour_integrable_on g" -proof - - { fix z - assume z: "z \ s" - obtain d where d: "d>0" "ball z d \ s" using \open s\ z - by (auto simp: open_contains_ball) - then have contfb: "continuous_on (ball z d) f" - using contf continuous_on_subset by blast - obtain h where "\y\ball z d. (h has_field_derivative f y) (at y within ball z d)" - using holomorphic_convex_primitive [OF convex_ball k contfb fcd] d - interior_subset by force - then have "\y\ball z d. (h has_field_derivative f y) (at y within s)" - by (metis Topology_Euclidean_Space.open_ball at_within_open d(2) os subsetCE) - then have "\h. (\y. cmod (y - z) < d \ (h has_field_derivative f y) (at y within s))" - by (force simp: dist_norm norm_minus_commute) - then have "\d h. 0 < d \ (\y. cmod (y - z) < d \ (h has_field_derivative f y) (at y within s))" - using d by blast - } - then show ?thesis - by (rule contour_integral_local_primitive [OF g]) -qed - -lemma contour_integrable_holomorphic_simple: - assumes fh: "f holomorphic_on s" - and os: "open s" - and g: "valid_path g" "path_image g \ s" - shows "f contour_integrable_on g" - apply (rule contour_integrable_holomorphic [OF _ os Finite_Set.finite.emptyI g]) - apply (simp add: fh holomorphic_on_imp_continuous_on) - using fh by (simp add: field_differentiable_def holomorphic_on_open os) - -lemma continuous_on_inversediff: - fixes z:: "'a::real_normed_field" shows "z \ s \ continuous_on s (\w. 1 / (w - z))" - by (rule continuous_intros | force)+ - -corollary contour_integrable_inversediff: - "\valid_path g; z \ path_image g\ \ (\w. 1 / (w-z)) contour_integrable_on g" -apply (rule contour_integrable_holomorphic_simple [of _ "UNIV-{z}"]) -apply (auto simp: holomorphic_on_open open_delete intro!: derivative_eq_intros) -done - -text\Key fact that path integral is the same for a "nearby" path. This is the - main lemma for the homotopy form of Cauchy's theorem and is also useful - if we want "without loss of generality" to assume some nice properties of a - path (e.g. smoothness). It can also be used to define the integrals of - analytic functions over arbitrary continuous paths. This is just done for - winding numbers now. -\ - -text\A technical definition to avoid duplication of similar proofs, - for paths joined at the ends versus looping paths\ -definition linked_paths :: "bool \ (real \ 'a) \ (real \ 'a::topological_space) \ bool" - where "linked_paths atends g h == - (if atends then pathstart h = pathstart g \ pathfinish h = pathfinish g - else pathfinish g = pathstart g \ pathfinish h = pathstart h)" - -text\This formulation covers two cases: @{term g} and @{term h} share their - start and end points; @{term g} and @{term h} both loop upon themselves.\ -lemma contour_integral_nearby: - assumes os: "open s" and p: "path p" "path_image p \ s" - shows - "\d. 0 < d \ - (\g h. valid_path g \ valid_path h \ - (\t \ {0..1}. norm(g t - p t) < d \ norm(h t - p t) < d) \ - linked_paths atends g h - \ path_image g \ s \ path_image h \ s \ - (\f. f holomorphic_on s \ contour_integral h f = contour_integral g f))" -proof - - have "\z. \e. z \ path_image p \ 0 < e \ ball z e \ s" - using open_contains_ball os p(2) by blast - then obtain ee where ee: "\z. z \ path_image p \ 0 < ee z \ ball z (ee z) \ s" - by metis - define cover where "cover = (\z. ball z (ee z/3)) ` (path_image p)" - have "compact (path_image p)" - by (metis p(1) compact_path_image) - moreover have "path_image p \ (\c\path_image p. ball c (ee c / 3))" - using ee by auto - ultimately have "\D \ cover. finite D \ path_image p \ \D" - by (simp add: compact_eq_heine_borel cover_def) - then obtain D where D: "D \ cover" "finite D" "path_image p \ \D" - by blast - then obtain k where k: "k \ {0..1}" "finite k" and D_eq: "D = ((\z. ball z (ee z / 3)) \ p) ` k" - apply (simp add: cover_def path_image_def image_comp) - apply (blast dest!: finite_subset_image [OF \finite D\]) - done - then have kne: "k \ {}" - using D by auto - have pi: "\i. i \ k \ p i \ path_image p" - using k by (auto simp: path_image_def) - then have eepi: "\i. i \ k \ 0 < ee((p i))" - by (metis ee) - define e where "e = Min((ee o p) ` k)" - have fin_eep: "finite ((ee o p) ` k)" - using k by blast - have enz: "0 < e" - using ee k by (simp add: kne e_def Min_gr_iff [OF fin_eep] eepi) - have "uniformly_continuous_on {0..1} p" - using p by (simp add: path_def compact_uniformly_continuous) - then obtain d::real where d: "d>0" - and de: "\x x'. \x' - x\ < d \ x\{0..1} \ x'\{0..1} \ cmod (p x' - p x) < e/3" - unfolding uniformly_continuous_on_def dist_norm real_norm_def - by (metis divide_pos_pos enz zero_less_numeral) - then obtain N::nat where N: "N>0" "inverse N < d" - using real_arch_inverse [of d] by auto - { fix g h - assume g: "valid_path g" and gp: "\t\{0..1}. cmod (g t - p t) < e / 3" - and h: "valid_path h" and hp: "\t\{0..1}. cmod (h t - p t) < e / 3" - and joins: "linked_paths atends g h" - { fix t::real - assume t: "0 \ t" "t \ 1" - then obtain u where u: "u \ k" and ptu: "p t \ ball(p u) (ee(p u) / 3)" - using \path_image p \ \D\ D_eq by (force simp: path_image_def) - then have ele: "e \ ee (p u)" using fin_eep - by (simp add: e_def) - have "cmod (g t - p t) < e / 3" "cmod (h t - p t) < e / 3" - using gp hp t by auto - with ele have "cmod (g t - p t) < ee (p u) / 3" - "cmod (h t - p t) < ee (p u) / 3" - by linarith+ - then have "g t \ ball(p u) (ee(p u))" "h t \ ball(p u) (ee(p u))" - using norm_diff_triangle_ineq [of "g t" "p t" "p t" "p u"] - norm_diff_triangle_ineq [of "h t" "p t" "p t" "p u"] ptu eepi u - by (force simp: dist_norm ball_def norm_minus_commute)+ - then have "g t \ s" "h t \ s" using ee u k - by (auto simp: path_image_def ball_def) - } - then have ghs: "path_image g \ s" "path_image h \ s" - by (auto simp: path_image_def) - moreover - { fix f - assume fhols: "f holomorphic_on s" - then have fpa: "f contour_integrable_on g" "f contour_integrable_on h" - using g ghs h holomorphic_on_imp_continuous_on os contour_integrable_holomorphic_simple - by blast+ - have contf: "continuous_on s f" - by (simp add: fhols holomorphic_on_imp_continuous_on) - { fix z - assume z: "z \ path_image p" - have "f holomorphic_on ball z (ee z)" - using fhols ee z holomorphic_on_subset by blast - then have "\ff. (\w \ ball z (ee z). (ff has_field_derivative f w) (at w))" - using holomorphic_convex_primitive [of "ball z (ee z)" "{}" f, simplified] - by (metis open_ball at_within_open holomorphic_on_def holomorphic_on_imp_continuous_on mem_ball) - } - then obtain ff where ff: - "\z w. \z \ path_image p; w \ ball z (ee z)\ \ (ff z has_field_derivative f w) (at w)" - by metis - { fix n - assume n: "n \ N" - then have "contour_integral(subpath 0 (n/N) h) f - contour_integral(subpath 0 (n/N) g) f = - contour_integral(linepath (g(n/N)) (h(n/N))) f - contour_integral(linepath (g 0) (h 0)) f" - proof (induct n) - case 0 show ?case by simp - next - case (Suc n) - obtain t where t: "t \ k" and "p (n/N) \ ball(p t) (ee(p t) / 3)" - using \path_image p \ \D\ [THEN subsetD, where c="p (n/N)"] D_eq N Suc.prems - by (force simp: path_image_def) - then have ptu: "cmod (p t - p (n/N)) < ee (p t) / 3" - by (simp add: dist_norm) - have e3le: "e/3 \ ee (p t) / 3" using fin_eep t - by (simp add: e_def) - { fix x - assume x: "n/N \ x" "x \ (1 + n)/N" - then have nN01: "0 \ n/N" "(1 + n)/N \ 1" - using Suc.prems by auto - then have x01: "0 \ x" "x \ 1" - using x by linarith+ - have "cmod (p t - p x) < ee (p t) / 3 + e/3" - apply (rule norm_diff_triangle_less [OF ptu de]) - using x N x01 Suc.prems - apply (auto simp: field_simps) - done - then have ptx: "cmod (p t - p x) < 2*ee (p t)/3" - using e3le eepi [OF t] by simp - have "cmod (p t - g x) < 2*ee (p t)/3 + e/3 " - apply (rule norm_diff_triangle_less [OF ptx]) - using gp x01 by (simp add: norm_minus_commute) - also have "... \ ee (p t)" - using e3le eepi [OF t] by simp - finally have gg: "cmod (p t - g x) < ee (p t)" . - have "cmod (p t - h x) < 2*ee (p t)/3 + e/3 " - apply (rule norm_diff_triangle_less [OF ptx]) - using hp x01 by (simp add: norm_minus_commute) - also have "... \ ee (p t)" - using e3le eepi [OF t] by simp - finally have "cmod (p t - g x) < ee (p t)" - "cmod (p t - h x) < ee (p t)" - using gg by auto - } note ptgh_ee = this - have pi_hgn: "path_image (linepath (h (n/N)) (g (n/N))) \ ball (p t) (ee (p t))" - using ptgh_ee [of "n/N"] Suc.prems - by (auto simp: field_simps dist_norm dest: segment_furthest_le [where y="p t"]) - then have gh_ns: "closed_segment (g (n/N)) (h (n/N)) \ s" - using \N>0\ Suc.prems - apply (simp add: path_image_join field_simps closed_segment_commute) - apply (erule order_trans) - apply (simp add: ee pi t) - done - have pi_ghn': "path_image (linepath (g ((1 + n) / N)) (h ((1 + n) / N))) - \ ball (p t) (ee (p t))" - using ptgh_ee [of "(1+n)/N"] Suc.prems - by (auto simp: field_simps dist_norm dest: segment_furthest_le [where y="p t"]) - then have gh_n's: "closed_segment (g ((1 + n) / N)) (h ((1 + n) / N)) \ s" - using \N>0\ Suc.prems ee pi t - by (auto simp: Path_Connected.path_image_join field_simps) - have pi_subset_ball: - "path_image (subpath (n/N) ((1+n) / N) g +++ linepath (g ((1+n) / N)) (h ((1+n) / N)) +++ - subpath ((1+n) / N) (n/N) h +++ linepath (h (n/N)) (g (n/N))) - \ ball (p t) (ee (p t))" - apply (intro subset_path_image_join pi_hgn pi_ghn') - using \N>0\ Suc.prems - apply (auto simp: path_image_subpath dist_norm field_simps closed_segment_eq_real_ivl ptgh_ee) - done - have pi0: "(f has_contour_integral 0) - (subpath (n/ N) ((Suc n)/N) g +++ linepath(g ((Suc n) / N)) (h((Suc n) / N)) +++ - subpath ((Suc n) / N) (n/N) h +++ linepath(h (n/N)) (g (n/N)))" - apply (rule Cauchy_theorem_primitive [of "ball(p t) (ee(p t))" "ff (p t)" "f"]) - apply (metis ff open_ball at_within_open pi t) - apply (intro valid_path_join) - using Suc.prems pi_subset_ball apply (simp_all add: valid_path_subpath g h) - done - have fpa1: "f contour_integrable_on subpath (real n / real N) (real (Suc n) / real N) g" - using Suc.prems by (simp add: contour_integrable_subpath g fpa) - have fpa2: "f contour_integrable_on linepath (g (real (Suc n) / real N)) (h (real (Suc n) / real N))" - using gh_n's - by (auto intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf]) - have fpa3: "f contour_integrable_on linepath (h (real n / real N)) (g (real n / real N))" - using gh_ns - by (auto simp: closed_segment_commute intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf]) - have eq0: "contour_integral (subpath (n/N) ((Suc n) / real N) g) f + - contour_integral (linepath (g ((Suc n) / N)) (h ((Suc n) / N))) f + - contour_integral (subpath ((Suc n) / N) (n/N) h) f + - contour_integral (linepath (h (n/N)) (g (n/N))) f = 0" - using contour_integral_unique [OF pi0] Suc.prems - by (simp add: g h fpa valid_path_subpath contour_integrable_subpath - fpa1 fpa2 fpa3 algebra_simps del: of_nat_Suc) - have *: "\hn he hn' gn gd gn' hgn ghn gh0 ghn'. - \hn - gn = ghn - gh0; - gd + ghn' + he + hgn = (0::complex); - hn - he = hn'; gn + gd = gn'; hgn = -ghn\ \ hn' - gn' = ghn' - gh0" - by (auto simp: algebra_simps) - have "contour_integral (subpath 0 (n/N) h) f - contour_integral (subpath ((Suc n) / N) (n/N) h) f = - contour_integral (subpath 0 (n/N) h) f + contour_integral (subpath (n/N) ((Suc n) / N) h) f" - unfolding reversepath_subpath [symmetric, of "((Suc n) / N)"] - using Suc.prems by (simp add: h fpa contour_integral_reversepath valid_path_subpath contour_integrable_subpath) - also have "... = contour_integral (subpath 0 ((Suc n) / N) h) f" - using Suc.prems by (simp add: contour_integral_subpath_combine h fpa) - finally have pi0_eq: - "contour_integral (subpath 0 (n/N) h) f - contour_integral (subpath ((Suc n) / N) (n/N) h) f = - contour_integral (subpath 0 ((Suc n) / N) h) f" . - show ?case - apply (rule * [OF Suc.hyps eq0 pi0_eq]) - using Suc.prems - apply (simp_all add: g h fpa contour_integral_subpath_combine - contour_integral_reversepath [symmetric] contour_integrable_continuous_linepath - continuous_on_subset [OF contf gh_ns]) - done - qed - } note ind = this - have "contour_integral h f = contour_integral g f" - using ind [OF order_refl] N joins - by (simp add: linked_paths_def pathstart_def pathfinish_def split: if_split_asm) - } - ultimately - have "path_image g \ s \ path_image h \ s \ (\f. f holomorphic_on s \ contour_integral h f = contour_integral g f)" - by metis - } note * = this - show ?thesis - apply (rule_tac x="e/3" in exI) - apply (rule conjI) - using enz apply simp - apply (clarsimp simp only: ball_conj_distrib) - apply (rule *; assumption) - done -qed - - -lemma - assumes "open s" "path p" "path_image p \ s" - shows contour_integral_nearby_ends: - "\d. 0 < d \ - (\g h. valid_path g \ valid_path h \ - (\t \ {0..1}. norm(g t - p t) < d \ norm(h t - p t) < d) \ - pathstart h = pathstart g \ pathfinish h = pathfinish g - \ path_image g \ s \ - path_image h \ s \ - (\f. f holomorphic_on s - \ contour_integral h f = contour_integral g f))" - and contour_integral_nearby_loops: - "\d. 0 < d \ - (\g h. valid_path g \ valid_path h \ - (\t \ {0..1}. norm(g t - p t) < d \ norm(h t - p t) < d) \ - pathfinish g = pathstart g \ pathfinish h = pathstart h - \ path_image g \ s \ - path_image h \ s \ - (\f. f holomorphic_on s - \ contour_integral h f = contour_integral g f))" - using contour_integral_nearby [OF assms, where atends=True] - using contour_integral_nearby [OF assms, where atends=False] - unfolding linked_paths_def by simp_all - -corollary differentiable_polynomial_function: - fixes p :: "real \ 'a::euclidean_space" - shows "polynomial_function p \ p differentiable_on s" -by (meson has_vector_derivative_polynomial_function differentiable_at_imp_differentiable_on differentiable_def has_vector_derivative_def) - -lemma C1_differentiable_polynomial_function: - fixes p :: "real \ 'a::euclidean_space" - shows "polynomial_function p \ p C1_differentiable_on s" - by (metis continuous_on_polymonial_function C1_differentiable_on_def has_vector_derivative_polynomial_function) - -lemma valid_path_polynomial_function: - fixes p :: "real \ 'a::euclidean_space" - shows "polynomial_function p \ valid_path p" -by (force simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_polymonial_function C1_differentiable_polynomial_function) - -lemma valid_path_subpath_trivial [simp]: - fixes g :: "real \ 'a::euclidean_space" - shows "z \ g x \ valid_path (subpath x x g)" - by (simp add: subpath_def valid_path_polynomial_function) - -lemma contour_integral_bound_exists: -assumes s: "open s" - and g: "valid_path g" - and pag: "path_image g \ s" - shows "\L. 0 < L \ - (\f B. f holomorphic_on s \ (\z \ s. norm(f z) \ B) - \ norm(contour_integral g f) \ L*B)" -proof - -have "path g" using g - by (simp add: valid_path_imp_path) -then obtain d::real and p - where d: "0 < d" - and p: "polynomial_function p" "path_image p \ s" - and pi: "\f. f holomorphic_on s \ contour_integral g f = contour_integral p f" - using contour_integral_nearby_ends [OF s \path g\ pag] - apply clarify - apply (drule_tac x=g in spec) - apply (simp only: assms) - apply (force simp: valid_path_polynomial_function dest: path_approx_polynomial_function) - done -then obtain p' where p': "polynomial_function p'" - "\x. (p has_vector_derivative (p' x)) (at x)" - using has_vector_derivative_polynomial_function by force -then have "bounded(p' ` {0..1})" - using continuous_on_polymonial_function - by (force simp: intro!: compact_imp_bounded compact_continuous_image) -then obtain L where L: "L>0" and nop': "\x. x \ {0..1} \ norm (p' x) \ L" - by (force simp: bounded_pos) -{ fix f B - assume f: "f holomorphic_on s" - and B: "\z. z\s \ cmod (f z) \ B" - then have "f contour_integrable_on p \ valid_path p" - using p s - by (blast intro: valid_path_polynomial_function contour_integrable_holomorphic_simple holomorphic_on_imp_continuous_on) - moreover have "\x. x \ {0..1} \ cmod (vector_derivative p (at x)) * cmod (f (p x)) \ L * B" - apply (rule mult_mono) - apply (subst Derivative.vector_derivative_at; force intro: p' nop') - using L B p - apply (auto simp: path_image_def image_subset_iff) - done - ultimately have "cmod (contour_integral g f) \ L * B" - apply (simp add: pi [OF f]) - apply (simp add: contour_integral_integral) - apply (rule order_trans [OF integral_norm_bound_integral]) - apply (auto simp: mult.commute integral_norm_bound_integral contour_integrable_on [symmetric] norm_mult) - done -} then -show ?thesis - by (force simp: L contour_integral_integral) -qed - -subsection\Constancy of a function from a connected set into a finite, disconnected or discrete set\ - -text\Still missing: versions for a set that is smaller than R, or countable.\ - -lemma continuous_disconnected_range_constant: - assumes s: "connected s" - and conf: "continuous_on s f" - and fim: "f ` s \ t" - and cct: "\y. y \ t \ connected_component_set t y = {y}" - shows "\a. \x \ s. f x = a" -proof (cases "s = {}") - case True then show ?thesis by force -next - case False - { fix x assume "x \ s" - then have "f ` s \ {f x}" - by (metis connected_continuous_image conf connected_component_maximal fim image_subset_iff rev_image_eqI s cct) - } - with False show ?thesis - by blast -qed - -lemma discrete_subset_disconnected: - fixes s :: "'a::topological_space set" - fixes t :: "'b::real_normed_vector set" - assumes conf: "continuous_on s f" - and no: "\x. x \ s \ \e>0. \y. y \ s \ f y \ f x \ e \ norm (f y - f x)" - shows "f ` s \ {y. connected_component_set (f ` s) y = {y}}" -proof - - { fix x assume x: "x \ s" - then obtain e where "e>0" and ele: "\y. \y \ s; f y \ f x\ \ e \ norm (f y - f x)" - using conf no [OF x] by auto - then have e2: "0 \ e / 2" - by simp - have "f y = f x" if "y \ s" and ccs: "f y \ connected_component_set (f ` s) (f x)" for y - apply (rule ccontr) - using connected_closed [of "connected_component_set (f ` s) (f x)"] \e>0\ - apply (simp add: del: ex_simps) - apply (drule spec [where x="cball (f x) (e / 2)"]) - apply (drule spec [where x="- ball(f x) e"]) - apply (auto simp: dist_norm open_closed [symmetric] simp del: le_divide_eq_numeral1 dest!: connected_component_in) - apply (metis diff_self e2 ele norm_minus_commute norm_zero not_less) - using centre_in_cball connected_component_refl_eq e2 x apply blast - using ccs - apply (force simp: cball_def dist_norm norm_minus_commute dest: ele [OF \y \ s\]) - done - moreover have "connected_component_set (f ` s) (f x) \ f ` s" - by (auto simp: connected_component_in) - ultimately have "connected_component_set (f ` s) (f x) = {f x}" - by (auto simp: x) - } - with assms show ?thesis - by blast -qed - -lemma finite_implies_discrete: - fixes s :: "'a::topological_space set" - assumes "finite (f ` s)" - shows "(\x \ s. \e>0. \y. y \ s \ f y \ f x \ e \ norm (f y - f x))" -proof - - have "\e>0. \y. y \ s \ f y \ f x \ e \ norm (f y - f x)" if "x \ s" for x - proof (cases "f ` s - {f x} = {}") - case True - with zero_less_numeral show ?thesis - by (fastforce simp add: Set.image_subset_iff cong: conj_cong) - next - case False - then obtain z where z: "z \ s" "f z \ f x" - by blast - have finn: "finite {norm (z - f x) |z. z \ f ` s - {f x}}" - using assms by simp - then have *: "0 < Inf{norm(z - f x) | z. z \ f ` s - {f x}}" - apply (rule finite_imp_less_Inf) - using z apply force+ - done - show ?thesis - by (force intro!: * cInf_le_finite [OF finn]) - qed - with assms show ?thesis - by blast -qed - -text\This proof requires the existence of two separate values of the range type.\ -lemma finite_range_constant_imp_connected: - assumes "\f::'a::topological_space \ 'b::real_normed_algebra_1. - \continuous_on s f; finite(f ` s)\ \ \a. \x \ s. f x = a" - shows "connected s" -proof - - { fix t u - assume clt: "closedin (subtopology euclidean s) t" - and clu: "closedin (subtopology euclidean s) u" - and tue: "t \ u = {}" and tus: "t \ u = s" - have conif: "continuous_on s (\x. if x \ t then 0 else 1)" - apply (subst tus [symmetric]) - apply (rule continuous_on_cases_local) - using clt clu tue - apply (auto simp: tus continuous_on_const) - done - have fi: "finite ((\x. if x \ t then 0 else 1) ` s)" - by (rule finite_subset [of _ "{0,1}"]) auto - have "t = {} \ u = {}" - using assms [OF conif fi] tus [symmetric] - by (auto simp: Ball_def) (metis IntI empty_iff one_neq_zero tue) - } - then show ?thesis - by (simp add: connected_closedin_eq) -qed - -lemma continuous_disconnected_range_constant_eq: - "(connected s \ - (\f::'a::topological_space \ 'b::real_normed_algebra_1. - \t. continuous_on s f \ f ` s \ t \ (\y \ t. connected_component_set t y = {y}) - \ (\a::'b. \x \ s. f x = a)))" (is ?thesis1) - and continuous_discrete_range_constant_eq: - "(connected s \ - (\f::'a::topological_space \ 'b::real_normed_algebra_1. - continuous_on s f \ - (\x \ s. \e. 0 < e \ (\y. y \ s \ (f y \ f x) \ e \ norm(f y - f x))) - \ (\a::'b. \x \ s. f x = a)))" (is ?thesis2) - and continuous_finite_range_constant_eq: - "(connected s \ - (\f::'a::topological_space \ 'b::real_normed_algebra_1. - continuous_on s f \ finite (f ` s) - \ (\a::'b. \x \ s. f x = a)))" (is ?thesis3) -proof - - have *: "\s t u v. \s \ t; t \ u; u \ v; v \ s\ - \ (s \ t) \ (s \ u) \ (s \ v)" - by blast - have "?thesis1 \ ?thesis2 \ ?thesis3" - apply (rule *) - using continuous_disconnected_range_constant apply metis - apply clarify - apply (frule discrete_subset_disconnected; blast) - apply (blast dest: finite_implies_discrete) - apply (blast intro!: finite_range_constant_imp_connected) - done - then show ?thesis1 ?thesis2 ?thesis3 - by blast+ -qed - -lemma continuous_discrete_range_constant: - fixes f :: "'a::topological_space \ 'b::real_normed_algebra_1" - assumes s: "connected s" - and "continuous_on s f" - and "\x. x \ s \ \e>0. \y. y \ s \ f y \ f x \ e \ norm (f y - f x)" - shows "\a. \x \ s. f x = a" - using continuous_discrete_range_constant_eq [THEN iffD1, OF s] assms - by blast - -lemma continuous_finite_range_constant: - fixes f :: "'a::topological_space \ 'b::real_normed_algebra_1" - assumes "connected s" - and "continuous_on s f" - and "finite (f ` s)" - shows "\a. \x \ s. f x = a" - using assms continuous_finite_range_constant_eq - by blast - - -text\We can treat even non-rectifiable paths as having a "length" for bounds on analytic functions in open sets.\ - -subsection\Winding Numbers\ - -definition winding_number:: "[real \ complex, complex] \ complex" where - "winding_number \ z \ - @n. \e > 0. \p. valid_path p \ z \ path_image p \ - pathstart p = pathstart \ \ - pathfinish p = pathfinish \ \ - (\t \ {0..1}. norm(\ t - p t) < e) \ - contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" - -lemma winding_number: - assumes "path \" "z \ path_image \" "0 < e" - shows "\p. valid_path p \ z \ path_image p \ - pathstart p = pathstart \ \ - pathfinish p = pathfinish \ \ - (\t \ {0..1}. norm (\ t - p t) < e) \ - contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * winding_number \ z" -proof - - have "path_image \ \ UNIV - {z}" - using assms by blast - then obtain d - where d: "d>0" - and pi_eq: "\h1 h2. valid_path h1 \ valid_path h2 \ - (\t\{0..1}. cmod (h1 t - \ t) < d \ cmod (h2 t - \ t) < d) \ - pathstart h2 = pathstart h1 \ pathfinish h2 = pathfinish h1 \ - path_image h1 \ UNIV - {z} \ path_image h2 \ UNIV - {z} \ - (\f. f holomorphic_on UNIV - {z} \ contour_integral h2 f = contour_integral h1 f)" - using contour_integral_nearby_ends [of "UNIV - {z}" \] assms by (auto simp: open_delete) - then obtain h where h: "polynomial_function h \ pathstart h = pathstart \ \ pathfinish h = pathfinish \ \ - (\t \ {0..1}. norm(h t - \ t) < d/2)" - using path_approx_polynomial_function [OF \path \\, of "d/2"] d by auto - define nn where "nn = 1/(2* pi*\) * contour_integral h (\w. 1/(w - z))" - have "\n. \e > 0. \p. valid_path p \ z \ path_image p \ - pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ - (\t \ {0..1}. norm(\ t - p t) < e) \ - contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" - (is "\n. \e > 0. ?PP e n") - proof (rule_tac x=nn in exI, clarify) - fix e::real - assume e: "e>0" - obtain p where p: "polynomial_function p \ - pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ (\t\{0..1}. cmod (p t - \ t) < min e (d / 2))" - using path_approx_polynomial_function [OF \path \\, of "min e (d/2)"] d \0 by auto - have "(\w. 1 / (w - z)) holomorphic_on UNIV - {z}" - by (auto simp: intro!: holomorphic_intros) - then show "?PP e nn" - apply (rule_tac x=p in exI) - using pi_eq [of h p] h p d - apply (auto simp: valid_path_polynomial_function norm_minus_commute nn_def) - done - qed - then show ?thesis - unfolding winding_number_def - apply (rule someI2_ex) - apply (blast intro: \0) - done -qed - -lemma winding_number_unique: - assumes \: "path \" "z \ path_image \" - and pi: - "\e. e>0 \ \p. valid_path p \ z \ path_image p \ - pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ - (\t \ {0..1}. norm (\ t - p t) < e) \ - contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" - shows "winding_number \ z = n" -proof - - have "path_image \ \ UNIV - {z}" - using assms by blast - then obtain e - where e: "e>0" - and pi_eq: "\h1 h2 f. \valid_path h1; valid_path h2; - (\t\{0..1}. cmod (h1 t - \ t) < e \ cmod (h2 t - \ t) < e); - pathstart h2 = pathstart h1; pathfinish h2 = pathfinish h1; f holomorphic_on UNIV - {z}\ \ - contour_integral h2 f = contour_integral h1 f" - using contour_integral_nearby_ends [of "UNIV - {z}" \] assms by (auto simp: open_delete) - obtain p where p: - "valid_path p \ z \ path_image p \ - pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ - (\t \ {0..1}. norm (\ t - p t) < e) \ - contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" - using pi [OF e] by blast - obtain q where q: - "valid_path q \ z \ path_image q \ - pathstart q = pathstart \ \ pathfinish q = pathfinish \ \ - (\t\{0..1}. cmod (\ t - q t) < e) \ contour_integral q (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number \ z" - using winding_number [OF \ e] by blast - have "2 * complex_of_real pi * \ * n = contour_integral p (\w. 1 / (w - z))" - using p by auto - also have "... = contour_integral q (\w. 1 / (w - z))" - apply (rule pi_eq) - using p q - by (auto simp: valid_path_polynomial_function norm_minus_commute intro!: holomorphic_intros) - also have "... = 2 * complex_of_real pi * \ * winding_number \ z" - using q by auto - finally have "2 * complex_of_real pi * \ * n = 2 * complex_of_real pi * \ * winding_number \ z" . - then show ?thesis - by simp -qed - -lemma winding_number_unique_loop: - assumes \: "path \" "z \ path_image \" - and loop: "pathfinish \ = pathstart \" - and pi: - "\e. e>0 \ \p. valid_path p \ z \ path_image p \ - pathfinish p = pathstart p \ - (\t \ {0..1}. norm (\ t - p t) < e) \ - contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" - shows "winding_number \ z = n" -proof - - have "path_image \ \ UNIV - {z}" - using assms by blast - then obtain e - where e: "e>0" - and pi_eq: "\h1 h2 f. \valid_path h1; valid_path h2; - (\t\{0..1}. cmod (h1 t - \ t) < e \ cmod (h2 t - \ t) < e); - pathfinish h1 = pathstart h1; pathfinish h2 = pathstart h2; f holomorphic_on UNIV - {z}\ \ - contour_integral h2 f = contour_integral h1 f" - using contour_integral_nearby_loops [of "UNIV - {z}" \] assms by (auto simp: open_delete) - obtain p where p: - "valid_path p \ z \ path_image p \ - pathfinish p = pathstart p \ - (\t \ {0..1}. norm (\ t - p t) < e) \ - contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" - using pi [OF e] by blast - obtain q where q: - "valid_path q \ z \ path_image q \ - pathstart q = pathstart \ \ pathfinish q = pathfinish \ \ - (\t\{0..1}. cmod (\ t - q t) < e) \ contour_integral q (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number \ z" - using winding_number [OF \ e] by blast - have "2 * complex_of_real pi * \ * n = contour_integral p (\w. 1 / (w - z))" - using p by auto - also have "... = contour_integral q (\w. 1 / (w - z))" - apply (rule pi_eq) - using p q loop - by (auto simp: valid_path_polynomial_function norm_minus_commute intro!: holomorphic_intros) - also have "... = 2 * complex_of_real pi * \ * winding_number \ z" - using q by auto - finally have "2 * complex_of_real pi * \ * n = 2 * complex_of_real pi * \ * winding_number \ z" . - then show ?thesis - by simp -qed - -lemma winding_number_valid_path: - assumes "valid_path \" "z \ path_image \" - shows "winding_number \ z = 1/(2*pi*\) * contour_integral \ (\w. 1/(w - z))" - using assms by (auto simp: valid_path_imp_path intro!: winding_number_unique) - -lemma has_contour_integral_winding_number: - assumes \: "valid_path \" "z \ path_image \" - shows "((\w. 1/(w - z)) has_contour_integral (2*pi*\*winding_number \ z)) \" -by (simp add: winding_number_valid_path has_contour_integral_integral contour_integrable_inversediff assms) - -lemma winding_number_trivial [simp]: "z \ a \ winding_number(linepath a a) z = 0" - by (simp add: winding_number_valid_path) - -lemma winding_number_subpath_trivial [simp]: "z \ g x \ winding_number (subpath x x g) z = 0" - by (simp add: path_image_subpath winding_number_valid_path) - -lemma winding_number_join: - assumes g1: "path g1" "z \ path_image g1" - and g2: "path g2" "z \ path_image g2" - and "pathfinish g1 = pathstart g2" - shows "winding_number(g1 +++ g2) z = winding_number g1 z + winding_number g2 z" - apply (rule winding_number_unique) - using assms apply (simp_all add: not_in_path_image_join) - apply (frule winding_number [OF g2]) - apply (frule winding_number [OF g1], clarify) - apply (rename_tac p2 p1) - apply (rule_tac x="p1+++p2" in exI) - apply (simp add: not_in_path_image_join contour_integrable_inversediff algebra_simps) - apply (auto simp: joinpaths_def) - done - -lemma winding_number_reversepath: - assumes "path \" "z \ path_image \" - shows "winding_number(reversepath \) z = - (winding_number \ z)" - apply (rule winding_number_unique) - using assms - apply simp_all - apply (frule winding_number [OF assms], clarify) - apply (rule_tac x="reversepath p" in exI) - apply (simp add: contour_integral_reversepath contour_integrable_inversediff valid_path_imp_reverse) - apply (auto simp: reversepath_def) - done - -lemma winding_number_shiftpath: - assumes \: "path \" "z \ path_image \" - and "pathfinish \ = pathstart \" "a \ {0..1}" - shows "winding_number(shiftpath a \) z = winding_number \ z" - apply (rule winding_number_unique_loop) - using assms - apply (simp_all add: path_shiftpath path_image_shiftpath pathfinish_shiftpath pathstart_shiftpath) - apply (frule winding_number [OF \], clarify) - apply (rule_tac x="shiftpath a p" in exI) - apply (simp add: contour_integral_shiftpath path_image_shiftpath pathfinish_shiftpath pathstart_shiftpath valid_path_shiftpath) - apply (auto simp: shiftpath_def) - done - -lemma winding_number_split_linepath: - assumes "c \ closed_segment a b" "z \ closed_segment a b" - shows "winding_number(linepath a b) z = winding_number(linepath a c) z + winding_number(linepath c b) z" -proof - - have "z \ closed_segment a c" "z \ closed_segment c b" - using assms apply (meson convex_contains_segment convex_segment ends_in_segment(1) subsetCE) - using assms by (meson convex_contains_segment convex_segment ends_in_segment(2) subsetCE) - then show ?thesis - using assms - by (simp add: winding_number_valid_path contour_integral_split_linepath [symmetric] continuous_on_inversediff field_simps) -qed - -lemma winding_number_cong: - "(\t. \0 \ t; t \ 1\ \ p t = q t) \ winding_number p z = winding_number q z" - by (simp add: winding_number_def pathstart_def pathfinish_def) - -lemma winding_number_offset: "winding_number p z = winding_number (\w. p w - z) 0" - apply (simp add: winding_number_def contour_integral_integral path_image_def valid_path_def pathstart_def pathfinish_def) - apply (intro ext arg_cong [where f = Eps] arg_cong [where f = All] imp_cong refl, safe) - apply (rename_tac g) - apply (rule_tac x="\t. g t - z" in exI) - apply (force simp: vector_derivative_def has_vector_derivative_diff_const piecewise_C1_differentiable_diff C1_differentiable_imp_piecewise) - apply (rename_tac g) - apply (rule_tac x="\t. g t + z" in exI) - apply (simp add: piecewise_C1_differentiable_add vector_derivative_def has_vector_derivative_add_const C1_differentiable_imp_piecewise) - apply (force simp: algebra_simps) - done - -(* A combined theorem deducing several things piecewise.*) -lemma winding_number_join_pos_combined: - "\valid_path \1; z \ path_image \1; 0 < Re(winding_number \1 z); - valid_path \2; z \ path_image \2; 0 < Re(winding_number \2 z); pathfinish \1 = pathstart \2\ - \ valid_path(\1 +++ \2) \ z \ path_image(\1 +++ \2) \ 0 < Re(winding_number(\1 +++ \2) z)" - by (simp add: valid_path_join path_image_join winding_number_join valid_path_imp_path) - - -(* Useful sufficient conditions for the winding number to be positive etc.*) - -lemma Re_winding_number: - "\valid_path \; z \ path_image \\ - \ Re(winding_number \ z) = Im(contour_integral \ (\w. 1/(w - z))) / (2*pi)" -by (simp add: winding_number_valid_path field_simps Re_divide power2_eq_square) - -lemma winding_number_pos_le: - assumes \: "valid_path \" "z \ path_image \" - and ge: "\x. \0 < x; x < 1\ \ 0 \ Im (vector_derivative \ (at x) * cnj(\ x - z))" - shows "0 \ Re(winding_number \ z)" -proof - - have *: "0 \ Im (vector_derivative \ (at x) / (\ x - z))" if x: "0 < x" "x < 1" for x - using ge by (simp add: Complex.Im_divide algebra_simps x) - show ?thesis - apply (simp add: Re_winding_number [OF \] field_simps) - apply (rule has_integral_component_nonneg - [of \ "\x. if x \ {0<..<1} - then 1/(\ x - z) * vector_derivative \ (at x) else 0", simplified]) - prefer 3 apply (force simp: *) - apply (simp add: Basis_complex_def) - apply (rule has_integral_spike_interior [of 0 1 _ "\x. 1/(\ x - z) * vector_derivative \ (at x)"]) - apply simp - apply (simp only: box_real) - apply (subst has_contour_integral [symmetric]) - using \ - apply (simp add: contour_integrable_inversediff has_contour_integral_integral) - done -qed - -lemma winding_number_pos_lt_lemma: - assumes \: "valid_path \" "z \ path_image \" - and e: "0 < e" - and ge: "\x. \0 < x; x < 1\ \ e \ Im (vector_derivative \ (at x) / (\ x - z))" - shows "0 < Re(winding_number \ z)" -proof - - have "e \ Im (contour_integral \ (\w. 1 / (w - z)))" - apply (rule has_integral_component_le - [of \ "\x. \*e" "\*e" "{0..1}" - "\x. if x \ {0<..<1} then 1/(\ x - z) * vector_derivative \ (at x) else \*e" - "contour_integral \ (\w. 1/(w - z))", simplified]) - using e - apply (simp_all add: Basis_complex_def) - using has_integral_const_real [of _ 0 1] apply force - apply (rule has_integral_spike_interior [of 0 1 _ "\x. 1/(\ x - z) * vector_derivative \ (at x)", simplified box_real]) - apply simp - apply (subst has_contour_integral [symmetric]) - using \ - apply (simp_all add: contour_integrable_inversediff has_contour_integral_integral ge) - done - with e show ?thesis - by (simp add: Re_winding_number [OF \] field_simps) -qed - -lemma winding_number_pos_lt: - assumes \: "valid_path \" "z \ path_image \" - and e: "0 < e" - and ge: "\x. \0 < x; x < 1\ \ e \ Im (vector_derivative \ (at x) * cnj(\ x - z))" - shows "0 < Re (winding_number \ z)" -proof - - have bm: "bounded ((\w. w - z) ` (path_image \))" - using bounded_translation [of _ "-z"] \ by (simp add: bounded_valid_path_image) - then obtain B where B: "B > 0" and Bno: "\x. x \ (\w. w - z) ` (path_image \) \ norm x \ B" - using bounded_pos [THEN iffD1, OF bm] by blast - { fix x::real assume x: "0 < x" "x < 1" - then have B2: "cmod (\ x - z)^2 \ B^2" using Bno [of "\ x - z"] - by (simp add: path_image_def power2_eq_square mult_mono') - with x have "\ x \ z" using \ - using path_image_def by fastforce - then have "e / B\<^sup>2 \ Im (vector_derivative \ (at x) * cnj (\ x - z)) / (cmod (\ x - z))\<^sup>2" - using B ge [OF x] B2 e - apply (rule_tac y="e / (cmod (\ x - z))\<^sup>2" in order_trans) - apply (auto simp: divide_left_mono divide_right_mono) - done - then have "e / B\<^sup>2 \ Im (vector_derivative \ (at x) / (\ x - z))" - by (simp add: Im_divide_Reals complex_div_cnj [of _ "\ x - z" for x] del: complex_cnj_diff times_complex.sel) - } note * = this - show ?thesis - using e B by (simp add: * winding_number_pos_lt_lemma [OF \, of "e/B^2"]) -qed - -subsection\The winding number is an integer\ - -text\Proof from the book Complex Analysis by Lars V. Ahlfors, Chapter 4, section 2.1, - Also on page 134 of Serge Lang's book with the name title, etc.\ - -lemma exp_fg: - fixes z::complex - assumes g: "(g has_vector_derivative g') (at x within s)" - and f: "(f has_vector_derivative (g' / (g x - z))) (at x within s)" - and z: "g x \ z" - shows "((\x. exp(-f x) * (g x - z)) has_vector_derivative 0) (at x within s)" -proof - - have *: "(exp o (\x. (- f x)) has_vector_derivative - (g' / (g x - z)) * exp (- f x)) (at x within s)" - using assms unfolding has_vector_derivative_def scaleR_conv_of_real - by (auto intro!: derivative_eq_intros) - show ?thesis - apply (rule has_vector_derivative_eq_rhs) - apply (rule bounded_bilinear.has_vector_derivative [OF bounded_bilinear_mult]) - using z - apply (auto simp: intro!: derivative_eq_intros * [unfolded o_def] g) - done -qed - -lemma winding_number_exp_integral: - fixes z::complex - assumes \: "\ piecewise_C1_differentiable_on {a..b}" - and ab: "a \ b" - and z: "z \ \ ` {a..b}" - shows "(\x. vector_derivative \ (at x) / (\ x - z)) integrable_on {a..b}" - (is "?thesis1") - "exp (- (integral {a..b} (\x. vector_derivative \ (at x) / (\ x - z)))) * (\ b - z) = \ a - z" - (is "?thesis2") -proof - - let ?D\ = "\x. vector_derivative \ (at x)" - have [simp]: "\x. a \ x \ x \ b \ \ x \ z" - using z by force - have cong: "continuous_on {a..b} \" - using \ by (simp add: piecewise_C1_differentiable_on_def) - obtain k where fink: "finite k" and g_C1_diff: "\ C1_differentiable_on ({a..b} - k)" - using \ by (force simp: piecewise_C1_differentiable_on_def) - have o: "open ({a<..finite k\ by (simp add: finite_imp_closed open_Diff) - moreover have "{a<.. {a..b} - k" - by force - ultimately have g_diff_at: "\x. \x \ k; x \ {a<.. \ \ differentiable at x" - by (metis Diff_iff differentiable_on_subset C1_diff_imp_diff [OF g_C1_diff] differentiable_on_def differentiable_within_open) - { fix w - assume "w \ z" - have "continuous_on (ball w (cmod (w - z))) (\w. 1 / (w - z))" - by (auto simp: dist_norm intro!: continuous_intros) - moreover have "\x. cmod (w - x) < cmod (w - z) \ \f'. ((\w. 1 / (w - z)) has_field_derivative f') (at x)" - by (auto simp: intro!: derivative_eq_intros) - ultimately have "\h. \y. norm(y - w) < norm(w - z) \ (h has_field_derivative 1/(y - z)) (at y)" - using holomorphic_convex_primitive [of "ball w (norm(w - z))" "{}" "\w. 1/(w - z)"] - by (simp add: field_differentiable_def Ball_def dist_norm at_within_open_NO_MATCH norm_minus_commute) - } - then obtain h where h: "\w y. w \ z \ norm(y - w) < norm(w - z) \ (h w has_field_derivative 1/(y - z)) (at y)" - by meson - have exy: "\y. ((\x. inverse (\ x - z) * ?D\ x) has_integral y) {a..b}" - unfolding integrable_on_def [symmetric] - apply (rule contour_integral_local_primitive_any [OF piecewise_C1_imp_differentiable [OF \], of "-{z}"]) - apply (rename_tac w) - apply (rule_tac x="norm(w - z)" in exI) - apply (simp_all add: inverse_eq_divide) - apply (metis has_field_derivative_at_within h) - done - have vg_int: "(\x. ?D\ x / (\ x - z)) integrable_on {a..b}" - unfolding box_real [symmetric] divide_inverse_commute - by (auto intro!: exy integrable_subinterval simp add: integrable_on_def ab) - with ab show ?thesis1 - by (simp add: divide_inverse_commute integral_def integrable_on_def) - { fix t - assume t: "t \ {a..b}" - have cball: "continuous_on (ball (\ t) (dist (\ t) z)) (\x. inverse (x - z))" - using z by (auto intro!: continuous_intros simp: dist_norm) - have icd: "\x. cmod (\ t - x) < cmod (\ t - z) \ (\w. inverse (w - z)) field_differentiable at x" - unfolding field_differentiable_def by (force simp: intro!: derivative_eq_intros) - obtain h where h: "\x. cmod (\ t - x) < cmod (\ t - z) \ - (h has_field_derivative inverse (x - z)) (at x within {y. cmod (\ t - y) < cmod (\ t - z)})" - using holomorphic_convex_primitive [where f = "\w. inverse(w - z)", OF convex_ball finite.emptyI cball icd] - by simp (auto simp: ball_def dist_norm that) - { fix x D - assume x: "x \ k" "a < x" "x < b" - then have "x \ interior ({a..b} - k)" - using open_subset_interior [OF o] by fastforce - then have con: "isCont (\x. ?D\ x) x" - using g_C1_diff x by (auto simp: C1_differentiable_on_eq intro: continuous_on_interior) - then have con_vd: "continuous (at x within {a..b}) (\x. ?D\ x)" - by (rule continuous_at_imp_continuous_within) - have gdx: "\ differentiable at x" - using x by (simp add: g_diff_at) - have "((\c. exp (- integral {a..c} (\x. vector_derivative \ (at x) / (\ x - z))) * (\ c - z)) has_derivative (\h. 0)) - (at x within {a..b})" - using x gdx t - apply (clarsimp simp add: differentiable_iff_scaleR) - apply (rule exp_fg [unfolded has_vector_derivative_def, simplified], blast intro: has_derivative_at_within) - apply (simp_all add: has_vector_derivative_def [symmetric]) - apply (rule has_vector_derivative_eq_rhs [OF integral_has_vector_derivative_continuous_at]) - apply (rule con_vd continuous_intros cong vg_int | simp add: continuous_at_imp_continuous_within has_vector_derivative_continuous vector_derivative_at)+ - done - } note * = this - have "exp (- (integral {a..t} (\x. ?D\ x / (\ x - z)))) * (\ t - z) =\ a - z" - apply (rule has_derivative_zero_unique_strong_interval [of "{a,b} \ k" a b]) - using t - apply (auto intro!: * continuous_intros fink cong indefinite_integral_continuous [OF vg_int] simp add: ab)+ - done - } - with ab show ?thesis2 - by (simp add: divide_inverse_commute integral_def) -qed - -corollary winding_number_exp_2pi: - "\path p; z \ path_image p\ - \ pathfinish p - z = exp (2 * pi * \ * winding_number p z) * (pathstart p - z)" -using winding_number [of p z 1] unfolding valid_path_def path_image_def pathstart_def pathfinish_def - by (force dest: winding_number_exp_integral(2) [of _ 0 1 z] simp: field_simps contour_integral_integral exp_minus) - - -subsection\The version with complex integers and equality\ - -lemma integer_winding_number_eq: - assumes \: "path \" and z: "z \ path_image \" - shows "winding_number \ z \ \ \ pathfinish \ = pathstart \" -proof - - have *: "\i::complex. \g0 g1. \i \ 0; g0 \ z; (g1 - z) / i = g0 - z\ \ (i = 1 \ g1 = g0)" - by (simp add: field_simps) algebra - obtain p where p: "valid_path p" "z \ path_image p" - "pathstart p = pathstart \" "pathfinish p = pathfinish \" - "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number \ z" - using winding_number [OF assms, of 1] by auto - have [simp]: "(winding_number \ z \ \) = (exp (contour_integral p (\w. 1 / (w - z))) = 1)" - using p by (simp add: exp_eq_1 complex_is_Int_iff) - have "winding_number p z \ \ \ pathfinish p = pathstart p" - using p z - apply (simp add: winding_number_valid_path valid_path_def path_image_def pathstart_def pathfinish_def) - using winding_number_exp_integral(2) [of p 0 1 z] - apply (simp add: field_simps contour_integral_integral exp_minus) - apply (rule *) - apply (auto simp: path_image_def field_simps) - done - then show ?thesis using p - by (auto simp: winding_number_valid_path) -qed - -theorem integer_winding_number: - "\path \; pathfinish \ = pathstart \; z \ path_image \\ \ winding_number \ z \ \" -by (metis integer_winding_number_eq) - - -text\If the winding number's magnitude is at least one, then the path must contain points in every direction.*) - We can thus bound the winding number of a path that doesn't intersect a given ray. \ - -lemma winding_number_pos_meets: - fixes z::complex - assumes \: "valid_path \" and z: "z \ path_image \" and 1: "Re (winding_number \ z) \ 1" - and w: "w \ z" - shows "\a::real. 0 < a \ z + a*(w - z) \ path_image \" -proof - - have [simp]: "\x. 0 \ x \ x \ 1 \ \ x \ z" - using z by (auto simp: path_image_def) - have [simp]: "z \ \ ` {0..1}" - using path_image_def z by auto - have gpd: "\ piecewise_C1_differentiable_on {0..1}" - using \ valid_path_def by blast - define r where "r = (w - z) / (\ 0 - z)" - have [simp]: "r \ 0" - using w z by (auto simp: r_def) - have "Arg r \ 2*pi" - by (simp add: Arg less_eq_real_def) - also have "... \ Im (integral {0..1} (\x. vector_derivative \ (at x) / (\ x - z)))" - using 1 - apply (simp add: winding_number_valid_path [OF \ z] Cauchy_Integral_Thm.contour_integral_integral) - apply (simp add: Complex.Re_divide field_simps power2_eq_square) - done - finally have "Arg r \ Im (integral {0..1} (\x. vector_derivative \ (at x) / (\ x - z)))" . - then have "\t. t \ {0..1} \ Im(integral {0..t} (\x. vector_derivative \ (at x)/(\ x - z))) = Arg r" - apply (simp add:) - apply (rule Topological_Spaces.IVT') - apply (simp_all add: Complex_Transcendental.Arg_ge_0) - apply (intro continuous_intros indefinite_integral_continuous winding_number_exp_integral [OF gpd]; simp) - done - then obtain t where t: "t \ {0..1}" - and eqArg: "Im (integral {0..t} (\x. vector_derivative \ (at x)/(\ x - z))) = Arg r" - by blast - define i where "i = integral {0..t} (\x. vector_derivative \ (at x) / (\ x - z))" - have iArg: "Arg r = Im i" - using eqArg by (simp add: i_def) - have gpdt: "\ piecewise_C1_differentiable_on {0..t}" - by (metis atLeastAtMost_iff atLeastatMost_subset_iff order_refl piecewise_C1_differentiable_on_subset gpd t) - have "exp (- i) * (\ t - z) = \ 0 - z" - unfolding i_def - apply (rule winding_number_exp_integral [OF gpdt]) - using t z unfolding path_image_def - apply force+ - done - then have *: "\ t - z = exp i * (\ 0 - z)" - by (simp add: exp_minus field_simps) - then have "(w - z) = r * (\ 0 - z)" - by (simp add: r_def) - then have "z + complex_of_real (exp (Re i)) * (w - z) / complex_of_real (cmod r) = \ t" - apply (simp add:) - apply (subst Complex_Transcendental.Arg_eq [of r]) - apply (simp add: iArg) - using * - apply (simp add: exp_eq_polar field_simps) - done - with t show ?thesis - by (rule_tac x="exp(Re i) / norm r" in exI) (auto simp: path_image_def) -qed - -lemma winding_number_big_meets: - fixes z::complex - assumes \: "valid_path \" and z: "z \ path_image \" and "\Re (winding_number \ z)\ \ 1" - and w: "w \ z" - shows "\a::real. 0 < a \ z + a*(w - z) \ path_image \" -proof - - { assume "Re (winding_number \ z) \ - 1" - then have "Re (winding_number (reversepath \) z) \ 1" - by (simp add: \ valid_path_imp_path winding_number_reversepath z) - moreover have "valid_path (reversepath \)" - using \ valid_path_imp_reverse by auto - moreover have "z \ path_image (reversepath \)" - by (simp add: z) - ultimately have "\a::real. 0 < a \ z + a*(w - z) \ path_image (reversepath \)" - using winding_number_pos_meets w by blast - then have ?thesis - by simp - } - then show ?thesis - using assms - by (simp add: Groups.abs_if_class.abs_if winding_number_pos_meets split: if_split_asm) -qed - -lemma winding_number_less_1: - fixes z::complex - shows - "\valid_path \; z \ path_image \; w \ z; - \a::real. 0 < a \ z + a*(w - z) \ path_image \\ - \ \Re(winding_number \ z)\ < 1" - by (auto simp: not_less dest: winding_number_big_meets) - -text\One way of proving that WN=1 for a loop.\ -lemma winding_number_eq_1: - fixes z::complex - assumes \: "valid_path \" and z: "z \ path_image \" and loop: "pathfinish \ = pathstart \" - and 0: "0 < Re(winding_number \ z)" and 2: "Re(winding_number \ z) < 2" - shows "winding_number \ z = 1" -proof - - have "winding_number \ z \ Ints" - by (simp add: \ integer_winding_number loop valid_path_imp_path z) - then show ?thesis - using 0 2 by (auto simp: Ints_def) -qed - - -subsection\Continuity of winding number and invariance on connected sets.\ - -lemma continuous_at_winding_number: - fixes z::complex - assumes \: "path \" and z: "z \ path_image \" - shows "continuous (at z) (winding_number \)" -proof - - obtain e where "e>0" and cbg: "cball z e \ - path_image \" - using open_contains_cball [of "- path_image \"] z - by (force simp: closed_def [symmetric] closed_path_image [OF \]) - then have ppag: "path_image \ \ - cball z (e/2)" - by (force simp: cball_def dist_norm) - have oc: "open (- cball z (e / 2))" - by (simp add: closed_def [symmetric]) - obtain d where "d>0" and pi_eq: - "\h1 h2. \valid_path h1; valid_path h2; - (\t\{0..1}. cmod (h1 t - \ t) < d \ cmod (h2 t - \ t) < d); - pathstart h2 = pathstart h1; pathfinish h2 = pathfinish h1\ - \ - path_image h1 \ - cball z (e / 2) \ - path_image h2 \ - cball z (e / 2) \ - (\f. f holomorphic_on - cball z (e / 2) \ contour_integral h2 f = contour_integral h1 f)" - using contour_integral_nearby_ends [OF oc \ ppag] by metis - obtain p where p: "valid_path p" "z \ path_image p" - "pathstart p = pathstart \ \ pathfinish p = pathfinish \" - and pg: "\t. t\{0..1} \ cmod (\ t - p t) < min d e / 2" - and pi: "contour_integral p (\x. 1 / (x - z)) = complex_of_real (2 * pi) * \ * winding_number \ z" - using winding_number [OF \ z, of "min d e / 2"] \d>0\ \e>0\ by auto - { fix w - assume d2: "cmod (w - z) < d/2" and e2: "cmod (w - z) < e/2" - then have wnotp: "w \ path_image p" - using cbg \d>0\ \e>0\ - apply (simp add: path_image_def cball_def dist_norm, clarify) - apply (frule pg) - apply (drule_tac c="\ x" in subsetD) - apply (auto simp: less_eq_real_def norm_minus_commute norm_triangle_half_l) - done - have wnotg: "w \ path_image \" - using cbg e2 \e>0\ by (force simp: dist_norm norm_minus_commute) - { fix k::real - assume k: "k>0" - then obtain q where q: "valid_path q" "w \ path_image q" - "pathstart q = pathstart \ \ pathfinish q = pathfinish \" - and qg: "\t. t \ {0..1} \ cmod (\ t - q t) < min k (min d e) / 2" - and qi: "contour_integral q (\u. 1 / (u - w)) = complex_of_real (2 * pi) * \ * winding_number \ w" - using winding_number [OF \ wnotg, of "min k (min d e) / 2"] \d>0\ \e>0\ k - by (force simp: min_divide_distrib_right) - have "contour_integral p (\u. 1 / (u - w)) = contour_integral q (\u. 1 / (u - w))" - apply (rule pi_eq [OF \valid_path q\ \valid_path p\, THEN conjunct2, THEN conjunct2, rule_format]) - apply (frule pg) - apply (frule qg) - using p q \d>0\ e2 - apply (auto simp: dist_norm norm_minus_commute intro!: holomorphic_intros) - done - then have "contour_integral p (\x. 1 / (x - w)) = complex_of_real (2 * pi) * \ * winding_number \ w" - by (simp add: pi qi) - } note pip = this - have "path p" - using p by (simp add: valid_path_imp_path) - then have "winding_number p w = winding_number \ w" - apply (rule winding_number_unique [OF _ wnotp]) - apply (rule_tac x=p in exI) - apply (simp add: p wnotp min_divide_distrib_right pip) - done - } note wnwn = this - obtain pe where "pe>0" and cbp: "cball z (3 / 4 * pe) \ - path_image p" - using p open_contains_cball [of "- path_image p"] - by (force simp: closed_def [symmetric] closed_path_image [OF valid_path_imp_path]) - obtain L - where "L>0" - and L: "\f B. \f holomorphic_on - cball z (3 / 4 * pe); - \z \ - cball z (3 / 4 * pe). cmod (f z) \ B\ \ - cmod (contour_integral p f) \ L * B" - using contour_integral_bound_exists [of "- cball z (3/4*pe)" p] cbp \valid_path p\ by force - { fix e::real and w::complex - assume e: "0 < e" and w: "cmod (w - z) < pe/4" "cmod (w - z) < e * pe\<^sup>2 / (8 * L)" - then have [simp]: "w \ path_image p" - using cbp p(2) \0 < pe\ - by (force simp: dist_norm norm_minus_commute path_image_def cball_def) - have [simp]: "contour_integral p (\x. 1/(x - w)) - contour_integral p (\x. 1/(x - z)) = - contour_integral p (\x. 1/(x - w) - 1/(x - z))" - by (simp add: p contour_integrable_inversediff contour_integral_diff) - { fix x - assume pe: "3/4 * pe < cmod (z - x)" - have "cmod (w - x) < pe/4 + cmod (z - x)" - by (meson add_less_cancel_right norm_diff_triangle_le order_refl order_trans_rules(21) w(1)) - then have wx: "cmod (w - x) < 4/3 * cmod (z - x)" using pe by simp - have "cmod (z - x) \ cmod (z - w) + cmod (w - x)" - using norm_diff_triangle_le by blast - also have "... < pe/4 + cmod (w - x)" - using w by (simp add: norm_minus_commute) - finally have "pe/2 < cmod (w - x)" - using pe by auto - then have "(pe/2)^2 < cmod (w - x) ^ 2" - apply (rule power_strict_mono) - using \pe>0\ by auto - then have pe2: "pe^2 < 4 * cmod (w - x) ^ 2" - by (simp add: power_divide) - have "8 * L * cmod (w - z) < e * pe\<^sup>2" - using w \L>0\ by (simp add: field_simps) - also have "... < e * 4 * cmod (w - x) * cmod (w - x)" - using pe2 \e>0\ by (simp add: power2_eq_square) - also have "... < e * 4 * cmod (w - x) * (4/3 * cmod (z - x))" - using wx - apply (rule mult_strict_left_mono) - using pe2 e not_less_iff_gr_or_eq by fastforce - finally have "L * cmod (w - z) < 2/3 * e * cmod (w - x) * cmod (z - x)" - by simp - also have "... \ e * cmod (w - x) * cmod (z - x)" - using e by simp - finally have Lwz: "L * cmod (w - z) < e * cmod (w - x) * cmod (z - x)" . - have "L * cmod (1 / (x - w) - 1 / (x - z)) \ e" - apply (cases "x=z \ x=w") - using pe \pe>0\ w \L>0\ - apply (force simp: norm_minus_commute) - using wx w(2) \L>0\ pe pe2 Lwz - apply (auto simp: divide_simps mult_less_0_iff norm_minus_commute norm_divide norm_mult power2_eq_square) - done - } note L_cmod_le = this - have *: "cmod (contour_integral p (\x. 1 / (x - w) - 1 / (x - z))) \ L * (e * pe\<^sup>2 / L / 4 * (inverse (pe / 2))\<^sup>2)" - apply (rule L) - using \pe>0\ w - apply (force simp: dist_norm norm_minus_commute intro!: holomorphic_intros) - using \pe>0\ w \L>0\ - apply (auto simp: cball_def dist_norm field_simps L_cmod_le simp del: less_divide_eq_numeral1 le_divide_eq_numeral1) - done - have "cmod (contour_integral p (\x. 1 / (x - w)) - contour_integral p (\x. 1 / (x - z))) < 2*e" - apply (simp add:) - apply (rule le_less_trans [OF *]) - using \L>0\ e - apply (force simp: field_simps) - done - then have "cmod (winding_number p w - winding_number p z) < e" - using pi_ge_two e - by (force simp: winding_number_valid_path p field_simps norm_divide norm_mult intro: less_le_trans) - } note cmod_wn_diff = this - then have "isCont (winding_number p) z" - apply (simp add: continuous_at_eps_delta, clarify) - apply (rule_tac x="min (pe/4) (e/2*pe^2/L/4)" in exI) - using \pe>0\ \L>0\ - apply (simp add: dist_norm cmod_wn_diff) - done - then show ?thesis - apply (rule continuous_transform_within [where d = "min d e / 2"]) - apply (auto simp: \d>0\ \e>0\ dist_norm wnwn) - done -qed - -corollary continuous_on_winding_number: - "path \ \ continuous_on (- path_image \) (\w. winding_number \ w)" - by (simp add: continuous_at_imp_continuous_on continuous_at_winding_number) - - -subsection\The winding number is constant on a connected region\ - -lemma winding_number_constant: - assumes \: "path \" and loop: "pathfinish \ = pathstart \" and cs: "connected s" and sg: "s \ path_image \ = {}" - shows "\k. \z \ s. winding_number \ z = k" -proof - - { fix y z - assume ne: "winding_number \ y \ winding_number \ z" - assume "y \ s" "z \ s" - then have "winding_number \ y \ \" "winding_number \ z \ \" - using integer_winding_number [OF \ loop] sg \y \ s\ by auto - with ne have "1 \ cmod (winding_number \ y - winding_number \ z)" - by (auto simp: Ints_def of_int_diff [symmetric] simp del: of_int_diff) - } note * = this - show ?thesis - apply (rule continuous_discrete_range_constant [OF cs]) - using continuous_on_winding_number [OF \] sg - apply (metis Diff_Compl Diff_eq_empty_iff continuous_on_subset) - apply (rule_tac x=1 in exI) - apply (auto simp: *) - done -qed - -lemma winding_number_eq: - "\path \; pathfinish \ = pathstart \; w \ s; z \ s; connected s; s \ path_image \ = {}\ - \ winding_number \ w = winding_number \ z" -using winding_number_constant by fastforce - -lemma open_winding_number_levelsets: - assumes \: "path \" and loop: "pathfinish \ = pathstart \" - shows "open {z. z \ path_image \ \ winding_number \ z = k}" -proof - - have op: "open (- path_image \)" - by (simp add: closed_path_image \ open_Compl) - { fix z assume z: "z \ path_image \" and k: "k = winding_number \ z" - obtain e where e: "e>0" "ball z e \ - path_image \" - using open_contains_ball [of "- path_image \"] op z - by blast - have "\e>0. \y. dist y z < e \ y \ path_image \ \ winding_number \ y = winding_number \ z" - apply (rule_tac x=e in exI) - using e apply (simp add: dist_norm ball_def norm_minus_commute) - apply (auto simp: dist_norm norm_minus_commute intro!: winding_number_eq [OF assms, where s = "ball z e"]) - done - } then - show ?thesis - by (auto simp: open_dist) -qed - -subsection\Winding number is zero "outside" a curve, in various senses\ - -lemma winding_number_zero_in_outside: - assumes \: "path \" and loop: "pathfinish \ = pathstart \" and z: "z \ outside (path_image \)" - shows "winding_number \ z = 0" -proof - - obtain B::real where "0 < B" and B: "path_image \ \ ball 0 B" - using bounded_subset_ballD [OF bounded_path_image [OF \]] by auto - obtain w::complex where w: "w \ ball 0 (B + 1)" - by (metis abs_of_nonneg le_less less_irrefl mem_ball_0 norm_of_real) - have "- ball 0 (B + 1) \ outside (path_image \)" - apply (rule outside_subset_convex) - using B subset_ball by auto - then have wout: "w \ outside (path_image \)" - using w by blast - moreover obtain k where "\z. z \ outside (path_image \) \ winding_number \ z = k" - using winding_number_constant [OF \ loop, of "outside(path_image \)"] connected_outside - by (metis DIM_complex bounded_path_image dual_order.refl \ outside_no_overlap) - ultimately have "winding_number \ z = winding_number \ w" - using z by blast - also have "... = 0" - proof - - have wnot: "w \ path_image \" using wout by (simp add: outside_def) - { fix e::real assume "0" "pathfinish p = pathfinish \" - and pg1: "(\t. \0 \ t; t \ 1\ \ cmod (p t - \ t) < 1)" - and pge: "(\t. \0 \ t; t \ 1\ \ cmod (p t - \ t) < e)" - using path_approx_polynomial_function [OF \, of "min 1 e"] \e>0\ by force - have pip: "path_image p \ ball 0 (B + 1)" - using B - apply (clarsimp simp add: path_image_def dist_norm ball_def) - apply (frule (1) pg1) - apply (fastforce dest: norm_add_less) - done - then have "w \ path_image p" using w by blast - then have "\p. valid_path p \ w \ path_image p \ - pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ - (\t\{0..1}. cmod (\ t - p t) < e) \ contour_integral p (\wa. 1 / (wa - w)) = 0" - apply (rule_tac x=p in exI) - apply (simp add: p valid_path_polynomial_function) - apply (intro conjI) - using pge apply (simp add: norm_minus_commute) - apply (rule contour_integral_unique [OF Cauchy_theorem_convex_simple [OF _ convex_ball [of 0 "B+1"]]]) - apply (rule holomorphic_intros | simp add: dist_norm)+ - using mem_ball_0 w apply blast - using p apply (simp_all add: valid_path_polynomial_function loop pip) - done - } - then show ?thesis - by (auto intro: winding_number_unique [OF \] simp add: wnot) - qed - finally show ?thesis . -qed - -lemma winding_number_zero_outside: - "\path \; convex s; pathfinish \ = pathstart \; z \ s; path_image \ \ s\ \ winding_number \ z = 0" - by (meson convex_in_outside outside_mono subsetCE winding_number_zero_in_outside) - -lemma winding_number_zero_at_infinity: - assumes \: "path \" and loop: "pathfinish \ = pathstart \" - shows "\B. \z. B \ norm z \ winding_number \ z = 0" -proof - - obtain B::real where "0 < B" and B: "path_image \ \ ball 0 B" - using bounded_subset_ballD [OF bounded_path_image [OF \]] by auto - then show ?thesis - apply (rule_tac x="B+1" in exI, clarify) - apply (rule winding_number_zero_outside [OF \ convex_cball [of 0 B] loop]) - apply (meson less_add_one mem_cball_0 not_le order_trans) - using ball_subset_cball by blast -qed - -lemma winding_number_zero_point: - "\path \; convex s; pathfinish \ = pathstart \; open s; path_image \ \ s\ - \ \z. z \ s \ winding_number \ z = 0" - using outside_compact_in_open [of "path_image \" s] path_image_nonempty winding_number_zero_in_outside - by (fastforce simp add: compact_path_image) - - -text\If a path winds round a set, it winds rounds its inside.\ -lemma winding_number_around_inside: - assumes \: "path \" and loop: "pathfinish \ = pathstart \" - and cls: "closed s" and cos: "connected s" and s_disj: "s \ path_image \ = {}" - and z: "z \ s" and wn_nz: "winding_number \ z \ 0" and w: "w \ s \ inside s" - shows "winding_number \ w = winding_number \ z" -proof - - have ssb: "s \ inside(path_image \)" - proof - fix x :: complex - assume "x \ s" - hence "x \ path_image \" - by (meson disjoint_iff_not_equal s_disj) - thus "x \ inside (path_image \)" - using \x \ s\ by (metis (no_types) ComplI UnE cos \ loop s_disj union_with_outside winding_number_eq winding_number_zero_in_outside wn_nz z) -qed - show ?thesis - apply (rule winding_number_eq [OF \ loop w]) - using z apply blast - apply (simp add: cls connected_with_inside cos) - apply (simp add: Int_Un_distrib2 s_disj, safe) - by (meson ssb inside_inside_compact_connected [OF cls, of "path_image \"] compact_path_image connected_path_image contra_subsetD disjoint_iff_not_equal \ inside_no_overlap) - qed - - -text\Bounding a WN by 1/2 for a path and point in opposite halfspaces.\ -lemma winding_number_subpath_continuous: - assumes \: "valid_path \" and z: "z \ path_image \" - shows "continuous_on {0..1} (\x. winding_number(subpath 0 x \) z)" -proof - - have *: "integral {0..x} (\t. vector_derivative \ (at t) / (\ t - z)) / (2 * of_real pi * \) = - winding_number (subpath 0 x \) z" - if x: "0 \ x" "x \ 1" for x - proof - - have "integral {0..x} (\t. vector_derivative \ (at t) / (\ t - z)) / (2 * of_real pi * \) = - 1 / (2*pi*\) * contour_integral (subpath 0 x \) (\w. 1/(w - z))" - using assms x - apply (simp add: contour_integral_subcontour_integral [OF contour_integrable_inversediff]) - done - also have "... = winding_number (subpath 0 x \) z" - apply (subst winding_number_valid_path) - using assms x - apply (simp_all add: path_image_subpath valid_path_subpath) - by (force simp: path_image_def) - finally show ?thesis . - qed - show ?thesis - apply (rule continuous_on_eq - [where f = "\x. 1 / (2*pi*\) * - integral {0..x} (\t. 1/(\ t - z) * vector_derivative \ (at t))"]) - apply (rule continuous_intros)+ - apply (rule indefinite_integral_continuous) - apply (rule contour_integrable_inversediff [OF assms, unfolded contour_integrable_on]) - using assms - apply (simp add: *) - done -qed - -lemma winding_number_ivt_pos: - assumes \: "valid_path \" and z: "z \ path_image \" and "0 \ w" "w \ Re(winding_number \ z)" - shows "\t \ {0..1}. Re(winding_number(subpath 0 t \) z) = w" - apply (rule ivt_increasing_component_on_1 [of 0 1, where ?k = "1::complex", simplified complex_inner_1_right]) - apply (simp add:) - apply (rule winding_number_subpath_continuous [OF \ z]) - using assms - apply (auto simp: path_image_def image_def) - done - -lemma winding_number_ivt_neg: - assumes \: "valid_path \" and z: "z \ path_image \" and "Re(winding_number \ z) \ w" "w \ 0" - shows "\t \ {0..1}. Re(winding_number(subpath 0 t \) z) = w" - apply (rule ivt_decreasing_component_on_1 [of 0 1, where ?k = "1::complex", simplified complex_inner_1_right]) - apply (simp add:) - apply (rule winding_number_subpath_continuous [OF \ z]) - using assms - apply (auto simp: path_image_def image_def) - done - -lemma winding_number_ivt_abs: - assumes \: "valid_path \" and z: "z \ path_image \" and "0 \ w" "w \ \Re(winding_number \ z)\" - shows "\t \ {0..1}. \Re (winding_number (subpath 0 t \) z)\ = w" - using assms winding_number_ivt_pos [of \ z w] winding_number_ivt_neg [of \ z "-w"] - by force - -lemma winding_number_lt_half_lemma: - assumes \: "valid_path \" and z: "z \ path_image \" and az: "a \ z \ b" and pag: "path_image \ \ {w. a \ w > b}" - shows "Re(winding_number \ z) < 1/2" -proof - - { assume "Re(winding_number \ z) \ 1/2" - then obtain t::real where t: "0 \ t" "t \ 1" and sub12: "Re (winding_number (subpath 0 t \) z) = 1/2" - using winding_number_ivt_pos [OF \ z, of "1/2"] by auto - have gt: "\ t - z = - (of_real (exp (- (2 * pi * Im (winding_number (subpath 0 t \) z)))) * (\ 0 - z))" - using winding_number_exp_2pi [of "subpath 0 t \" z] - apply (simp add: t \ valid_path_imp_path) - using closed_segment_eq_real_ivl path_image_def t z by (fastforce simp: path_image_subpath Euler sub12) - have "b < a \ \ 0" - proof - - have "\ 0 \ {c. b < a \ c}" - by (metis (no_types) pag atLeastAtMost_iff image_subset_iff order_refl path_image_def zero_le_one) - thus ?thesis - by blast - qed - moreover have "b < a \ \ t" - proof - - have "\ t \ {c. b < a \ c}" - by (metis (no_types) pag atLeastAtMost_iff image_subset_iff path_image_def t) - thus ?thesis - by blast - qed - ultimately have "0 < a \ (\ 0 - z)" "0 < a \ (\ t - z)" using az - by (simp add: inner_diff_right)+ - then have False - by (simp add: gt inner_mult_right mult_less_0_iff) - } - then show ?thesis by force -qed - -lemma winding_number_lt_half: - assumes "valid_path \" "a \ z \ b" "path_image \ \ {w. a \ w > b}" - shows "\Re (winding_number \ z)\ < 1/2" -proof - - have "z \ path_image \" using assms by auto - with assms show ?thesis - apply (simp add: winding_number_lt_half_lemma abs_if del: less_divide_eq_numeral1) - apply (metis complex_inner_1_right winding_number_lt_half_lemma [OF valid_path_imp_reverse, of \ z a b] - winding_number_reversepath valid_path_imp_path inner_minus_left path_image_reversepath) - done -qed - -lemma winding_number_le_half: - assumes \: "valid_path \" and z: "z \ path_image \" - and anz: "a \ 0" and azb: "a \ z \ b" and pag: "path_image \ \ {w. a \ w \ b}" - shows "\Re (winding_number \ z)\ \ 1/2" -proof - - { assume wnz_12: "\Re (winding_number \ z)\ > 1/2" - have "isCont (winding_number \) z" - by (metis continuous_at_winding_number valid_path_imp_path \ z) - then obtain d where "d>0" and d: "\x'. dist x' z < d \ dist (winding_number \ x') (winding_number \ z) < \Re(winding_number \ z)\ - 1/2" - using continuous_at_eps_delta wnz_12 diff_gt_0_iff_gt by blast - define z' where "z' = z - (d / (2 * cmod a)) *\<^sub>R a" - have *: "a \ z' \ b - d / 3 * cmod a" - unfolding z'_def inner_mult_right' divide_inverse - apply (simp add: divide_simps algebra_simps dot_square_norm power2_eq_square anz) - apply (metis \0 < d\ add_increasing azb less_eq_real_def mult_nonneg_nonneg mult_right_mono norm_ge_zero norm_numeral) - done - have "cmod (winding_number \ z' - winding_number \ z) < \Re (winding_number \ z)\ - 1/2" - using d [of z'] anz \d>0\ by (simp add: dist_norm z'_def) - then have "1/2 < \Re (winding_number \ z)\ - cmod (winding_number \ z' - winding_number \ z)" - by simp - then have "1/2 < \Re (winding_number \ z)\ - \Re (winding_number \ z') - Re (winding_number \ z)\" - using abs_Re_le_cmod [of "winding_number \ z' - winding_number \ z"] by simp - then have wnz_12': "\Re (winding_number \ z')\ > 1/2" - by linarith - moreover have "\Re (winding_number \ z')\ < 1/2" - apply (rule winding_number_lt_half [OF \ *]) - using azb \d>0\ pag - apply (auto simp: add_strict_increasing anz divide_simps algebra_simps dest!: subsetD) - done - ultimately have False - by simp - } - then show ?thesis by force -qed - -lemma winding_number_lt_half_linepath: "z \ closed_segment a b \ \Re (winding_number (linepath a b) z)\ < 1/2" - using separating_hyperplane_closed_point [of "closed_segment a b" z] - apply auto - apply (simp add: closed_segment_def) - apply (drule less_imp_le) - apply (frule winding_number_lt_half [OF valid_path_linepath [of a b]]) - apply (auto simp: segment) - done - - -text\ Positivity of WN for a linepath.\ -lemma winding_number_linepath_pos_lt: - assumes "0 < Im ((b - a) * cnj (b - z))" - shows "0 < Re(winding_number(linepath a b) z)" -proof - - have z: "z \ path_image (linepath a b)" - using assms - by (simp add: closed_segment_def) (force simp: algebra_simps) - show ?thesis - apply (rule winding_number_pos_lt [OF valid_path_linepath z assms]) - apply (simp add: linepath_def algebra_simps) - done -qed - - -subsection\Cauchy's integral formula, again for a convex enclosing set.\ - -lemma Cauchy_integral_formula_weak: - assumes s: "convex s" and "finite k" and conf: "continuous_on s f" - and fcd: "(\x. x \ interior s - k \ f field_differentiable at x)" - and z: "z \ interior s - k" and vpg: "valid_path \" - and pasz: "path_image \ \ s - {z}" and loop: "pathfinish \ = pathstart \" - shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" -proof - - obtain f' where f': "(f has_field_derivative f') (at z)" - using fcd [OF z] by (auto simp: field_differentiable_def) - have pas: "path_image \ \ s" and znotin: "z \ path_image \" using pasz by blast+ - have c: "continuous (at x within s) (\w. if w = z then f' else (f w - f z) / (w - z))" if "x \ s" for x - proof (cases "x = z") - case True then show ?thesis - apply (simp add: continuous_within) - apply (rule Lim_transform_away_within [of _ "z+1" _ "\w::complex. (f w - f z)/(w - z)"]) - using has_field_derivative_at_within DERIV_within_iff f' - apply (fastforce simp add:)+ - done - next - case False - then have dxz: "dist x z > 0" by auto - have cf: "continuous (at x within s) f" - using conf continuous_on_eq_continuous_within that by blast - have "continuous (at x within s) (\w. (f w - f z) / (w - z))" - by (rule cf continuous_intros | simp add: False)+ - then show ?thesis - apply (rule continuous_transform_within [OF _ dxz that, of "\w::complex. (f w - f z)/(w - z)"]) - apply (force simp: dist_commute) - done - qed - have fink': "finite (insert z k)" using \finite k\ by blast - have *: "((\w. if w = z then f' else (f w - f z) / (w - z)) has_contour_integral 0) \" - apply (rule Cauchy_theorem_convex [OF _ s fink' _ vpg pas loop]) - using c apply (force simp: continuous_on_eq_continuous_within) - apply (rename_tac w) - apply (rule_tac d="dist w z" and f = "\w. (f w - f z)/(w - z)" in field_differentiable_transform_within) - apply (simp_all add: dist_pos_lt dist_commute) - apply (metis less_irrefl) - apply (rule derivative_intros fcd | simp)+ - done - show ?thesis - apply (rule has_contour_integral_eq) - using znotin has_contour_integral_add [OF has_contour_integral_lmul [OF has_contour_integral_winding_number [OF vpg znotin], of "f z"] *] - apply (auto simp: mult_ac divide_simps) - done -qed - -theorem Cauchy_integral_formula_convex_simple: - "\convex s; f holomorphic_on s; z \ interior s; valid_path \; path_image \ \ s - {z}; - pathfinish \ = pathstart \\ - \ ((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" - apply (rule Cauchy_integral_formula_weak [where k = "{}"]) - using holomorphic_on_imp_continuous_on - by auto (metis at_within_interior holomorphic_on_def interiorE subsetCE) - - -subsection\Homotopy forms of Cauchy's theorem\ - -proposition Cauchy_theorem_homotopic: - assumes hom: "if atends then homotopic_paths s g h else homotopic_loops s g h" - and "open s" and f: "f holomorphic_on s" - and vpg: "valid_path g" and vph: "valid_path h" - shows "contour_integral g f = contour_integral h f" -proof - - have pathsf: "linked_paths atends g h" - using hom by (auto simp: linked_paths_def homotopic_paths_imp_pathstart homotopic_paths_imp_pathfinish homotopic_loops_imp_loop) - obtain k :: "real \ real \ complex" - where contk: "continuous_on ({0..1} \ {0..1}) k" - and ks: "k ` ({0..1} \ {0..1}) \ s" - and k [simp]: "\x. k (0, x) = g x" "\x. k (1, x) = h x" - and ksf: "\t\{0..1}. linked_paths atends g (\x. k (t, x))" - using hom pathsf by (auto simp: linked_paths_def homotopic_paths_def homotopic_loops_def homotopic_with_def split: if_split_asm) - have ucontk: "uniformly_continuous_on ({0..1} \ {0..1}) k" - by (blast intro: compact_Times compact_uniformly_continuous [OF contk]) - { fix t::real assume t: "t \ {0..1}" - have pak: "path (k o (\u. (t, u)))" - unfolding path_def - apply (rule continuous_intros continuous_on_subset [OF contk])+ - using t by force - have pik: "path_image (k \ Pair t) \ s" - using ks t by (auto simp: path_image_def) - obtain e where "e>0" and e: - "\g h. \valid_path g; valid_path h; - \u\{0..1}. cmod (g u - (k \ Pair t) u) < e \ cmod (h u - (k \ Pair t) u) < e; - linked_paths atends g h\ - \ contour_integral h f = contour_integral g f" - using contour_integral_nearby [OF \open s\ pak pik, of atends] f by metis - obtain d where "d>0" and d: - "\x x'. \x \ {0..1} \ {0..1}; x' \ {0..1} \ {0..1}; norm (x'-x) < d\ \ norm (k x' - k x) < e/4" - by (rule uniformly_continuous_onE [OF ucontk, of "e/4"]) (auto simp: dist_norm \e>0\) - { fix t1 t2 - assume t1: "0 \ t1" "t1 \ 1" and t2: "0 \ t2" "t2 \ 1" and ltd: "\t1 - t\ < d" "\t2 - t\ < d" - have no2: "\g1 k1 kt. \norm(g1 - k1) < e/4; norm(k1 - kt) < e/4\ \ norm(g1 - kt) < e" - using \e > 0\ - apply (rule_tac y = k1 in norm_triangle_half_l) - apply (auto simp: norm_minus_commute intro: order_less_trans) - done - have "\d>0. \g1 g2. valid_path g1 \ valid_path g2 \ - (\u\{0..1}. cmod (g1 u - k (t1, u)) < d \ cmod (g2 u - k (t2, u)) < d) \ - linked_paths atends g1 g2 \ - contour_integral g2 f = contour_integral g1 f" - apply (rule_tac x="e/4" in exI) - using t t1 t2 ltd \e > 0\ - apply (auto intro!: e simp: d no2 simp del: less_divide_eq_numeral1) - done - } - then have "\e. 0 < e \ - (\t1 t2. t1 \ {0..1} \ t2 \ {0..1} \ \t1 - t\ < e \ \t2 - t\ < e - \ (\d. 0 < d \ - (\g1 g2. valid_path g1 \ valid_path g2 \ - (\u \ {0..1}. - norm(g1 u - k((t1,u))) < d \ norm(g2 u - k((t2,u))) < d) \ - linked_paths atends g1 g2 - \ contour_integral g2 f = contour_integral g1 f)))" - by (rule_tac x=d in exI) (simp add: \d > 0\) - } - then obtain ee where ee: - "\t. t \ {0..1} \ ee t > 0 \ - (\t1 t2. t1 \ {0..1} \ t2 \ {0..1} \ \t1 - t\ < ee t \ \t2 - t\ < ee t - \ (\d. 0 < d \ - (\g1 g2. valid_path g1 \ valid_path g2 \ - (\u \ {0..1}. - norm(g1 u - k((t1,u))) < d \ norm(g2 u - k((t2,u))) < d) \ - linked_paths atends g1 g2 - \ contour_integral g2 f = contour_integral g1 f)))" - by metis - note ee_rule = ee [THEN conjunct2, rule_format] - define C where "C = (\t. ball t (ee t / 3)) ` {0..1}" - have "\t \ C. open t" by (simp add: C_def) - moreover have "{0..1} \ \C" - using ee [THEN conjunct1] by (auto simp: C_def dist_norm) - ultimately obtain C' where C': "C' \ C" "finite C'" and C'01: "{0..1} \ \C'" - by (rule compactE [OF compact_interval]) - define kk where "kk = {t \ {0..1}. ball t (ee t / 3) \ C'}" - have kk01: "kk \ {0..1}" by (auto simp: kk_def) - define e where "e = Min (ee ` kk)" - have C'_eq: "C' = (\t. ball t (ee t / 3)) ` kk" - using C' by (auto simp: kk_def C_def) - have ee_pos[simp]: "\t. t \ {0..1} \ ee t > 0" - by (simp add: kk_def ee) - moreover have "finite kk" - using \finite C'\ kk01 by (force simp: C'_eq inj_on_def ball_eq_ball_iff dest: ee_pos finite_imageD) - moreover have "kk \ {}" using \{0..1} \ \C'\ C'_eq by force - ultimately have "e > 0" - using finite_less_Inf_iff [of "ee ` kk" 0] kk01 by (force simp: e_def) - then obtain N::nat where "N > 0" and N: "1/N < e/3" - by (meson divide_pos_pos nat_approx_posE zero_less_Suc zero_less_numeral) - have e_le_ee: "\i. i \ kk \ e \ ee i" - using \finite kk\ by (simp add: e_def Min_le_iff [of "ee ` kk"]) - have plus: "\t \ kk. x \ ball t (ee t / 3)" if "x \ {0..1}" for x - using C' subsetD [OF C'01 that] unfolding C'_eq by blast - have [OF order_refl]: - "\d. 0 < d \ (\j. valid_path j \ (\u \ {0..1}. norm(j u - k (n/N, u)) < d) \ linked_paths atends g j - \ contour_integral j f = contour_integral g f)" - if "n \ N" for n - using that - proof (induct n) - case 0 show ?case using ee_rule [of 0 0 0] - apply clarsimp - apply (rule_tac x=d in exI, safe) - by (metis diff_self vpg norm_zero) - next - case (Suc n) - then have N01: "n/N \ {0..1}" "(Suc n)/N \ {0..1}" by auto - then obtain t where t: "t \ kk" "n/N \ ball t (ee t / 3)" - using plus [of "n/N"] by blast - then have nN_less: "\n/N - t\ < ee t" - by (simp add: dist_norm del: less_divide_eq_numeral1) - have n'N_less: "\real (Suc n) / real N - t\ < ee t" - using t N \N > 0\ e_le_ee [of t] - by (simp add: dist_norm add_divide_distrib abs_diff_less_iff del: less_divide_eq_numeral1) (simp add: field_simps) - have t01: "t \ {0..1}" using \kk \ {0..1}\ \t \ kk\ by blast - obtain d1 where "d1 > 0" and d1: - "\g1 g2. \valid_path g1; valid_path g2; - \u\{0..1}. cmod (g1 u - k (n/N, u)) < d1 \ cmod (g2 u - k ((Suc n) / N, u)) < d1; - linked_paths atends g1 g2\ - \ contour_integral g2 f = contour_integral g1 f" - using ee [THEN conjunct2, rule_format, OF t01 N01 nN_less n'N_less] by fastforce - have "n \ N" using Suc.prems by auto - with Suc.hyps - obtain d2 where "d2 > 0" - and d2: "\j. \valid_path j; \u\{0..1}. cmod (j u - k (n/N, u)) < d2; linked_paths atends g j\ - \ contour_integral j f = contour_integral g f" - by auto - have "continuous_on {0..1} (k o (\u. (n/N, u)))" - apply (rule continuous_intros continuous_on_subset [OF contk])+ - using N01 by auto - then have pkn: "path (\u. k (n/N, u))" - by (simp add: path_def) - have min12: "min d1 d2 > 0" by (simp add: \0 < d1\ \0 < d2\) - obtain p where "polynomial_function p" - and psf: "pathstart p = pathstart (\u. k (n/N, u))" - "pathfinish p = pathfinish (\u. k (n/N, u))" - and pk_le: "\t. t\{0..1} \ cmod (p t - k (n/N, t)) < min d1 d2" - using path_approx_polynomial_function [OF pkn min12] by blast - then have vpp: "valid_path p" using valid_path_polynomial_function by blast - have lpa: "linked_paths atends g p" - by (metis (mono_tags, lifting) N01(1) ksf linked_paths_def pathfinish_def pathstart_def psf) - show ?case - apply (rule_tac x="min d1 d2" in exI) - apply (simp add: \0 < d1\ \0 < d2\, clarify) - apply (rule_tac s="contour_integral p f" in trans) - using pk_le N01(1) ksf pathfinish_def pathstart_def - apply (force intro!: vpp d1 simp add: linked_paths_def psf ksf) - using pk_le N01 apply (force intro!: vpp d2 lpa simp add: linked_paths_def psf ksf) - done - qed - then obtain d where "0 < d" - "\j. valid_path j \ (\u \ {0..1}. norm(j u - k (1,u)) < d) \ - linked_paths atends g j - \ contour_integral j f = contour_integral g f" - using \N>0\ by auto - then have "linked_paths atends g h \ contour_integral h f = contour_integral g f" - using \N>0\ vph by fastforce - then show ?thesis - by (simp add: pathsf) -qed - -proposition Cauchy_theorem_homotopic_paths: - assumes hom: "homotopic_paths s g h" - and "open s" and f: "f holomorphic_on s" - and vpg: "valid_path g" and vph: "valid_path h" - shows "contour_integral g f = contour_integral h f" - using Cauchy_theorem_homotopic [of True s g h] assms by simp - -proposition Cauchy_theorem_homotopic_loops: - assumes hom: "homotopic_loops s g h" - and "open s" and f: "f holomorphic_on s" - and vpg: "valid_path g" and vph: "valid_path h" - shows "contour_integral g f = contour_integral h f" - using Cauchy_theorem_homotopic [of False s g h] assms by simp - -lemma has_contour_integral_newpath: - "\(f has_contour_integral y) h; f contour_integrable_on g; contour_integral g f = contour_integral h f\ - \ (f has_contour_integral y) g" - using has_contour_integral_integral contour_integral_unique by auto - -lemma Cauchy_theorem_null_homotopic: - "\f holomorphic_on s; open s; valid_path g; homotopic_loops s g (linepath a a)\ \ (f has_contour_integral 0) g" - apply (rule has_contour_integral_newpath [where h = "linepath a a"], simp) - using contour_integrable_holomorphic_simple - apply (blast dest: holomorphic_on_imp_continuous_on homotopic_loops_imp_subset) - by (simp add: Cauchy_theorem_homotopic_loops) - - - -subsection\More winding number properties\ - -text\including the fact that it's +-1 inside a simple closed curve.\ - -lemma winding_number_homotopic_paths: - assumes "homotopic_paths (-{z}) g h" - shows "winding_number g z = winding_number h z" -proof - - have "path g" "path h" using homotopic_paths_imp_path [OF assms] by auto - moreover have pag: "z \ path_image g" and pah: "z \ path_image h" - using homotopic_paths_imp_subset [OF assms] by auto - ultimately obtain d e where "d > 0" "e > 0" - and d: "\p. \path p; pathstart p = pathstart g; pathfinish p = pathfinish g; \t\{0..1}. norm (p t - g t) < d\ - \ homotopic_paths (-{z}) g p" - and e: "\q. \path q; pathstart q = pathstart h; pathfinish q = pathfinish h; \t\{0..1}. norm (q t - h t) < e\ - \ homotopic_paths (-{z}) h q" - using homotopic_nearby_paths [of g "-{z}"] homotopic_nearby_paths [of h "-{z}"] by force - obtain p where p: - "valid_path p" "z \ path_image p" - "pathstart p = pathstart g" "pathfinish p = pathfinish g" - and gp_less:"\t\{0..1}. cmod (g t - p t) < d" - and pap: "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number g z" - using winding_number [OF \path g\ pag \0 < d\] by blast - obtain q where q: - "valid_path q" "z \ path_image q" - "pathstart q = pathstart h" "pathfinish q = pathfinish h" - and hq_less: "\t\{0..1}. cmod (h t - q t) < e" - and paq: "contour_integral q (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number h z" - using winding_number [OF \path h\ pah \0 < e\] by blast - have gp: "homotopic_paths (- {z}) g p" - by (simp add: d p valid_path_imp_path norm_minus_commute gp_less) - have hq: "homotopic_paths (- {z}) h q" - by (simp add: e q valid_path_imp_path norm_minus_commute hq_less) - have "contour_integral p (\w. 1/(w - z)) = contour_integral q (\w. 1/(w - z))" - apply (rule Cauchy_theorem_homotopic_paths [of "-{z}"]) - apply (blast intro: homotopic_paths_trans homotopic_paths_sym gp hq assms) - apply (auto intro!: holomorphic_intros simp: p q) - done - then show ?thesis - by (simp add: pap paq) -qed - -lemma winding_number_homotopic_loops: - assumes "homotopic_loops (-{z}) g h" - shows "winding_number g z = winding_number h z" -proof - - have "path g" "path h" using homotopic_loops_imp_path [OF assms] by auto - moreover have pag: "z \ path_image g" and pah: "z \ path_image h" - using homotopic_loops_imp_subset [OF assms] by auto - moreover have gloop: "pathfinish g = pathstart g" and hloop: "pathfinish h = pathstart h" - using homotopic_loops_imp_loop [OF assms] by auto - ultimately obtain d e where "d > 0" "e > 0" - and d: "\p. \path p; pathfinish p = pathstart p; \t\{0..1}. norm (p t - g t) < d\ - \ homotopic_loops (-{z}) g p" - and e: "\q. \path q; pathfinish q = pathstart q; \t\{0..1}. norm (q t - h t) < e\ - \ homotopic_loops (-{z}) h q" - using homotopic_nearby_loops [of g "-{z}"] homotopic_nearby_loops [of h "-{z}"] by force - obtain p where p: - "valid_path p" "z \ path_image p" - "pathstart p = pathstart g" "pathfinish p = pathfinish g" - and gp_less:"\t\{0..1}. cmod (g t - p t) < d" - and pap: "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number g z" - using winding_number [OF \path g\ pag \0 < d\] by blast - obtain q where q: - "valid_path q" "z \ path_image q" - "pathstart q = pathstart h" "pathfinish q = pathfinish h" - and hq_less: "\t\{0..1}. cmod (h t - q t) < e" - and paq: "contour_integral q (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number h z" - using winding_number [OF \path h\ pah \0 < e\] by blast - have gp: "homotopic_loops (- {z}) g p" - by (simp add: gloop d gp_less norm_minus_commute p valid_path_imp_path) - have hq: "homotopic_loops (- {z}) h q" - by (simp add: e hloop hq_less norm_minus_commute q valid_path_imp_path) - have "contour_integral p (\w. 1/(w - z)) = contour_integral q (\w. 1/(w - z))" - apply (rule Cauchy_theorem_homotopic_loops [of "-{z}"]) - apply (blast intro: homotopic_loops_trans homotopic_loops_sym gp hq assms) - apply (auto intro!: holomorphic_intros simp: p q) - done - then show ?thesis - by (simp add: pap paq) -qed - -lemma winding_number_paths_linear_eq: - "\path g; path h; pathstart h = pathstart g; pathfinish h = pathfinish g; - \t. t \ {0..1} \ z \ closed_segment (g t) (h t)\ - \ winding_number h z = winding_number g z" - by (blast intro: sym homotopic_paths_linear winding_number_homotopic_paths elim: ) - -lemma winding_number_loops_linear_eq: - "\path g; path h; pathfinish g = pathstart g; pathfinish h = pathstart h; - \t. t \ {0..1} \ z \ closed_segment (g t) (h t)\ - \ winding_number h z = winding_number g z" - by (blast intro: sym homotopic_loops_linear winding_number_homotopic_loops elim: ) - -lemma winding_number_nearby_paths_eq: - "\path g; path h; - pathstart h = pathstart g; pathfinish h = pathfinish g; - \t. t \ {0..1} \ norm(h t - g t) < norm(g t - z)\ - \ winding_number h z = winding_number g z" - by (metis segment_bound(2) norm_minus_commute not_le winding_number_paths_linear_eq) - -lemma winding_number_nearby_loops_eq: - "\path g; path h; - pathfinish g = pathstart g; - pathfinish h = pathstart h; - \t. t \ {0..1} \ norm(h t - g t) < norm(g t - z)\ - \ winding_number h z = winding_number g z" - by (metis segment_bound(2) norm_minus_commute not_le winding_number_loops_linear_eq) - - -proposition winding_number_subpath_combine: - "\path g; z \ path_image g; - u \ {0..1}; v \ {0..1}; w \ {0..1}\ - \ winding_number (subpath u v g) z + winding_number (subpath v w g) z = - winding_number (subpath u w g) z" -apply (rule trans [OF winding_number_join [THEN sym] - winding_number_homotopic_paths [OF homotopic_join_subpaths]]) -apply (auto dest: path_image_subpath_subset) -done - - -subsection\Partial circle path\ - -definition part_circlepath :: "[complex, real, real, real, real] \ complex" - where "part_circlepath z r s t \ \x. z + of_real r * exp (\ * of_real (linepath s t x))" - -lemma pathstart_part_circlepath [simp]: - "pathstart(part_circlepath z r s t) = z + r*exp(\ * s)" -by (metis part_circlepath_def pathstart_def pathstart_linepath) - -lemma pathfinish_part_circlepath [simp]: - "pathfinish(part_circlepath z r s t) = z + r*exp(\*t)" -by (metis part_circlepath_def pathfinish_def pathfinish_linepath) - -proposition has_vector_derivative_part_circlepath [derivative_intros]: - "((part_circlepath z r s t) has_vector_derivative - (\ * r * (of_real t - of_real s) * exp(\ * linepath s t x))) - (at x within X)" - apply (simp add: part_circlepath_def linepath_def scaleR_conv_of_real) - apply (rule has_vector_derivative_real_complex) - apply (rule derivative_eq_intros | simp)+ - done - -corollary vector_derivative_part_circlepath: - "vector_derivative (part_circlepath z r s t) (at x) = - \ * r * (of_real t - of_real s) * exp(\ * linepath s t x)" - using has_vector_derivative_part_circlepath vector_derivative_at by blast - -corollary vector_derivative_part_circlepath01: - "\0 \ x; x \ 1\ - \ vector_derivative (part_circlepath z r s t) (at x within {0..1}) = - \ * r * (of_real t - of_real s) * exp(\ * linepath s t x)" - using has_vector_derivative_part_circlepath - by (auto simp: vector_derivative_at_within_ivl) - -lemma valid_path_part_circlepath [simp]: "valid_path (part_circlepath z r s t)" - apply (simp add: valid_path_def) - apply (rule C1_differentiable_imp_piecewise) - apply (auto simp: C1_differentiable_on_eq vector_derivative_works vector_derivative_part_circlepath has_vector_derivative_part_circlepath - intro!: continuous_intros) - done - -lemma path_part_circlepath [simp]: "path (part_circlepath z r s t)" - by (simp add: valid_path_imp_path) - -proposition path_image_part_circlepath: - assumes "s \ t" - shows "path_image (part_circlepath z r s t) = {z + r * exp(\ * of_real x) | x. s \ x \ x \ t}" -proof - - { fix z::real - assume "0 \ z" "z \ 1" - with \s \ t\ have "\x. (exp (\ * linepath s t z) = exp (\ * of_real x)) \ s \ x \ x \ t" - apply (rule_tac x="(1 - z) * s + z * t" in exI) - apply (simp add: linepath_def scaleR_conv_of_real algebra_simps) - apply (rule conjI) - using mult_right_mono apply blast - using affine_ineq by (metis "mult.commute") - } - moreover - { fix z - assume "s \ z" "z \ t" - then have "z + of_real r * exp (\ * of_real z) \ (\x. z + of_real r * exp (\ * linepath s t x)) ` {0..1}" - apply (rule_tac x="(z - s)/(t - s)" in image_eqI) - apply (simp add: linepath_def scaleR_conv_of_real divide_simps exp_eq) - apply (auto simp: algebra_simps divide_simps) - done - } - ultimately show ?thesis - by (fastforce simp add: path_image_def part_circlepath_def) -qed - -corollary path_image_part_circlepath_subset: - "\s \ t; 0 \ r\ \ path_image(part_circlepath z r s t) \ sphere z r" -by (auto simp: path_image_part_circlepath sphere_def dist_norm algebra_simps norm_mult) - -proposition in_path_image_part_circlepath: - assumes "w \ path_image(part_circlepath z r s t)" "s \ t" "0 \ r" - shows "norm(w - z) = r" -proof - - have "w \ {c. dist z c = r}" - by (metis (no_types) path_image_part_circlepath_subset sphere_def subset_eq assms) - thus ?thesis - by (simp add: dist_norm norm_minus_commute) -qed - -proposition finite_bounded_log: "finite {z::complex. norm z \ b \ exp z = w}" -proof (cases "w = 0") - case True then show ?thesis by auto -next - case False - have *: "finite {x. cmod (complex_of_real (2 * real_of_int x * pi) * \) \ b + cmod (Ln w)}" - apply (simp add: norm_mult finite_int_iff_bounded_le) - apply (rule_tac x="\(b + cmod (Ln w)) / (2*pi)\" in exI) - apply (auto simp: divide_simps le_floor_iff) - done - have [simp]: "\P f. {z. P z \ (\n. z = f n)} = f ` {n. P (f n)}" - by blast - show ?thesis - apply (subst exp_Ln [OF False, symmetric]) - apply (simp add: exp_eq) - using norm_add_leD apply (fastforce intro: finite_subset [OF _ *]) - done -qed - -lemma finite_bounded_log2: - fixes a::complex - assumes "a \ 0" - shows "finite {z. norm z \ b \ exp(a*z) = w}" -proof - - have *: "finite ((\z. z / a) ` {z. cmod z \ b * cmod a \ exp z = w})" - by (rule finite_imageI [OF finite_bounded_log]) - show ?thesis - by (rule finite_subset [OF _ *]) (force simp: assms norm_mult) -qed - -proposition has_contour_integral_bound_part_circlepath_strong: - assumes fi: "(f has_contour_integral i) (part_circlepath z r s t)" - and "finite k" and le: "0 \ B" "0 < r" "s \ t" - and B: "\x. x \ path_image(part_circlepath z r s t) - k \ norm(f x) \ B" - shows "cmod i \ B * r * (t - s)" -proof - - consider "s = t" | "s < t" using \s \ t\ by linarith - then show ?thesis - proof cases - case 1 with fi [unfolded has_contour_integral] - have "i = 0" by (simp add: vector_derivative_part_circlepath) - with assms show ?thesis by simp - next - case 2 - have [simp]: "\r\ = r" using \r > 0\ by linarith - have [simp]: "cmod (complex_of_real t - complex_of_real s) = t-s" - by (metis "2" abs_of_pos diff_gt_0_iff_gt norm_of_real of_real_diff) - have "finite (part_circlepath z r s t -` {y} \ {0..1})" if "y \ k" for y - proof - - define w where "w = (y - z)/of_real r / exp(\ * of_real s)" - have fin: "finite (of_real -` {z. cmod z \ 1 \ exp (\ * complex_of_real (t - s) * z) = w})" - apply (rule finite_vimageI [OF finite_bounded_log2]) - using \s < t\ apply (auto simp: inj_of_real) - done - show ?thesis - apply (simp add: part_circlepath_def linepath_def vimage_def) - apply (rule finite_subset [OF _ fin]) - using le - apply (auto simp: w_def algebra_simps scaleR_conv_of_real exp_add exp_diff) - done - qed - then have fin01: "finite ((part_circlepath z r s t) -` k \ {0..1})" - by (rule finite_finite_vimage_IntI [OF \finite k\]) - have **: "((\x. if (part_circlepath z r s t x) \ k then 0 - else f(part_circlepath z r s t x) * - vector_derivative (part_circlepath z r s t) (at x)) has_integral i) {0..1}" - apply (rule has_integral_spike - [where f = "\x. f(part_circlepath z r s t x) * vector_derivative (part_circlepath z r s t) (at x)"]) - apply (rule negligible_finite [OF fin01]) - using fi has_contour_integral - apply auto - done - have *: "\x. \0 \ x; x \ 1; part_circlepath z r s t x \ k\ \ cmod (f (part_circlepath z r s t x)) \ B" - by (auto intro!: B [unfolded path_image_def image_def, simplified]) - show ?thesis - apply (rule has_integral_bound [where 'a=real, simplified, OF _ **, simplified]) - using assms apply force - apply (simp add: norm_mult vector_derivative_part_circlepath) - using le * "2" \r > 0\ by auto - qed -qed - -corollary has_contour_integral_bound_part_circlepath: - "\(f has_contour_integral i) (part_circlepath z r s t); - 0 \ B; 0 < r; s \ t; - \x. x \ path_image(part_circlepath z r s t) \ norm(f x) \ B\ - \ norm i \ B*r*(t - s)" - by (auto intro: has_contour_integral_bound_part_circlepath_strong) - -proposition contour_integrable_continuous_part_circlepath: - "continuous_on (path_image (part_circlepath z r s t)) f - \ f contour_integrable_on (part_circlepath z r s t)" - apply (simp add: contour_integrable_on has_contour_integral_def vector_derivative_part_circlepath path_image_def) - apply (rule integrable_continuous_real) - apply (fast intro: path_part_circlepath [unfolded path_def] continuous_intros continuous_on_compose2 [where g=f, OF _ _ order_refl]) - done - -proposition winding_number_part_circlepath_pos_less: - assumes "s < t" and no: "norm(w - z) < r" - shows "0 < Re (winding_number(part_circlepath z r s t) w)" -proof - - have "0 < r" by (meson no norm_not_less_zero not_le order.strict_trans2) - note valid_path_part_circlepath - moreover have " w \ path_image (part_circlepath z r s t)" - using assms by (auto simp: path_image_def image_def part_circlepath_def norm_mult linepath_def) - moreover have "0 < r * (t - s) * (r - cmod (w - z))" - using assms by (metis \0 < r\ diff_gt_0_iff_gt mult_pos_pos) - ultimately show ?thesis - apply (rule winding_number_pos_lt [where e = "r*(t - s)*(r - norm(w - z))"]) - apply (simp add: vector_derivative_part_circlepath right_diff_distrib [symmetric] mult_ac) - apply (rule mult_left_mono)+ - using Re_Im_le_cmod [of "w-z" "linepath s t x" for x] - apply (simp add: exp_Euler cos_of_real sin_of_real part_circlepath_def algebra_simps cos_squared_eq [unfolded power2_eq_square]) - using assms \0 < r\ by auto -qed - -proposition simple_path_part_circlepath: - "simple_path(part_circlepath z r s t) \ (r \ 0 \ s \ t \ \s - t\ \ 2*pi)" -proof (cases "r = 0 \ s = t") - case True - then show ?thesis - apply (rule disjE) - apply (force simp: part_circlepath_def simple_path_def intro: bexI [where x = "1/4"] bexI [where x = "1/3"])+ - done -next - case False then have "r \ 0" "s \ t" by auto - have *: "\x y z s t. \*((1 - x) * s + x * t) = \*(((1 - y) * s + y * t)) + z \ \*(x - y) * (t - s) = z" - by (simp add: algebra_simps) - have abs01: "\x y::real. 0 \ x \ x \ 1 \ 0 \ y \ y \ 1 - \ (x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0 \ \x - y\ \ {0,1})" - by auto - have abs_away: "\P. (\x\{0..1}. \y\{0..1}. P \x - y\) \ (\x::real. 0 \ x \ x \ 1 \ P x)" - by force - have **: "\x y. (\n. (complex_of_real x - of_real y) * (of_real t - of_real s) = 2 * (of_int n * of_real pi)) \ - (\n. \x - y\ * (t - s) = 2 * (of_int n * pi))" - by (force simp: algebra_simps abs_if dest: arg_cong [where f=Re] arg_cong [where f=complex_of_real] - intro: exI [where x = "-n" for n]) - have 1: "\x. 0 \ x \ x \ 1 \ (\n. x * (t - s) = 2 * (real_of_int n * pi)) \ x = 0 \ x = 1 \ \s - t\ \ 2 * pi" - apply (rule ccontr) - apply (drule_tac x="2*pi / \t - s\" in spec) - using False - apply (simp add: abs_minus_commute divide_simps) - apply (frule_tac x=1 in spec) - apply (drule_tac x="-1" in spec) - apply (simp add:) - done - have 2: "\s - t\ = \2 * (real_of_int n * pi) / x\" if "x \ 0" "x * (t - s) = 2 * (real_of_int n * pi)" for x n - proof - - have "t-s = 2 * (real_of_int n * pi)/x" - using that by (simp add: field_simps) - then show ?thesis by (metis abs_minus_commute) - qed - show ?thesis using False - apply (simp add: simple_path_def path_part_circlepath) - apply (simp add: part_circlepath_def linepath_def exp_eq * ** abs01 del: Set.insert_iff) - apply (subst abs_away) - apply (auto simp: 1) - apply (rule ccontr) - apply (auto simp: 2 divide_simps abs_mult dest: of_int_leD) - done -qed - -proposition arc_part_circlepath: - assumes "r \ 0" "s \ t" "\s - t\ < 2*pi" - shows "arc (part_circlepath z r s t)" -proof - - have *: "x = y" if eq: "\ * (linepath s t x) = \ * (linepath s t y) + 2 * of_int n * complex_of_real pi * \" - and x: "x \ {0..1}" and y: "y \ {0..1}" for x y n - proof - - have "(linepath s t x) = (linepath s t y) + 2 * of_int n * complex_of_real pi" - by (metis add_divide_eq_iff complex_i_not_zero mult.commute nonzero_mult_divide_cancel_left eq) - then have "s*y + t*x = s*x + (t*y + of_int n * (pi * 2))" - by (force simp: algebra_simps linepath_def dest: arg_cong [where f=Re]) - then have st: "x \ y \ (s-t) = (of_int n * (pi * 2) / (y-x))" - by (force simp: field_simps) - show ?thesis - apply (rule ccontr) - using assms x y - apply (simp add: st abs_mult field_simps) - using st - apply (auto simp: dest: of_int_lessD) - done - qed - show ?thesis - using assms - apply (simp add: arc_def) - apply (simp add: part_circlepath_def inj_on_def exp_eq) - apply (blast intro: *) - done -qed - - -subsection\Special case of one complete circle\ - -definition circlepath :: "[complex, real, real] \ complex" - where "circlepath z r \ part_circlepath z r 0 (2*pi)" - -lemma circlepath: "circlepath z r = (\x. z + r * exp(2 * of_real pi * \ * of_real x))" - by (simp add: circlepath_def part_circlepath_def linepath_def algebra_simps) - -lemma pathstart_circlepath [simp]: "pathstart (circlepath z r) = z + r" - by (simp add: circlepath_def) - -lemma pathfinish_circlepath [simp]: "pathfinish (circlepath z r) = z + r" - by (simp add: circlepath_def) (metis exp_two_pi_i mult.commute) - -lemma circlepath_minus: "circlepath z (-r) x = circlepath z r (x + 1/2)" -proof - - have "z + of_real r * exp (2 * pi * \ * (x + 1 / 2)) = - z + of_real r * exp (2 * pi * \ * x + pi * \)" - by (simp add: divide_simps) (simp add: algebra_simps) - also have "... = z - r * exp (2 * pi * \ * x)" - by (simp add: exp_add) - finally show ?thesis - by (simp add: circlepath path_image_def sphere_def dist_norm) -qed - -lemma circlepath_add1: "circlepath z r (x+1) = circlepath z r x" - using circlepath_minus [of z r "x+1/2"] circlepath_minus [of z "-r" x] - by (simp add: add.commute) - -lemma circlepath_add_half: "circlepath z r (x + 1/2) = circlepath z r (x - 1/2)" - using circlepath_add1 [of z r "x-1/2"] - by (simp add: add.commute) - -lemma path_image_circlepath_minus_subset: - "path_image (circlepath z (-r)) \ path_image (circlepath z r)" - apply (simp add: path_image_def image_def circlepath_minus, clarify) - apply (case_tac "xa \ 1/2", force) - apply (force simp add: circlepath_add_half)+ - done - -lemma path_image_circlepath_minus: "path_image (circlepath z (-r)) = path_image (circlepath z r)" - using path_image_circlepath_minus_subset by fastforce - -proposition has_vector_derivative_circlepath [derivative_intros]: - "((circlepath z r) has_vector_derivative (2 * pi * \ * r * exp (2 * of_real pi * \ * of_real x))) - (at x within X)" - apply (simp add: circlepath_def scaleR_conv_of_real) - apply (rule derivative_eq_intros) - apply (simp add: algebra_simps) - done - -corollary vector_derivative_circlepath: - "vector_derivative (circlepath z r) (at x) = - 2 * pi * \ * r * exp(2 * of_real pi * \ * x)" -using has_vector_derivative_circlepath vector_derivative_at by blast - -corollary vector_derivative_circlepath01: - "\0 \ x; x \ 1\ - \ vector_derivative (circlepath z r) (at x within {0..1}) = - 2 * pi * \ * r * exp(2 * of_real pi * \ * x)" - using has_vector_derivative_circlepath - by (auto simp: vector_derivative_at_within_ivl) - -lemma valid_path_circlepath [simp]: "valid_path (circlepath z r)" - by (simp add: circlepath_def) - -lemma path_circlepath [simp]: "path (circlepath z r)" - by (simp add: valid_path_imp_path) - -lemma path_image_circlepath_nonneg: - assumes "0 \ r" shows "path_image (circlepath z r) = sphere z r" -proof - - have *: "x \ (\u. z + (cmod (x - z)) * exp (\ * (of_real u * (of_real pi * 2)))) ` {0..1}" for x - proof (cases "x = z") - case True then show ?thesis by force - next - case False - define w where "w = x - z" - then have "w \ 0" by (simp add: False) - have **: "\t. \Re w = cos t * cmod w; Im w = sin t * cmod w\ \ w = of_real (cmod w) * exp (\ * t)" - using cis_conv_exp complex_eq_iff by auto - show ?thesis - apply (rule sincos_total_2pi [of "Re(w/of_real(norm w))" "Im(w/of_real(norm w))"]) - apply (simp add: divide_simps \w \ 0\ cmod_power2 [symmetric]) - apply (rule_tac x="t / (2*pi)" in image_eqI) - apply (simp add: divide_simps \w \ 0\) - using False ** - apply (auto simp: w_def) - done - qed - show ?thesis - unfolding circlepath path_image_def sphere_def dist_norm - by (force simp: assms algebra_simps norm_mult norm_minus_commute intro: *) -qed - -proposition path_image_circlepath [simp]: - "path_image (circlepath z r) = sphere z \r\" - using path_image_circlepath_minus - by (force simp add: path_image_circlepath_nonneg abs_if) - -lemma has_contour_integral_bound_circlepath_strong: - "\(f has_contour_integral i) (circlepath z r); - finite k; 0 \ B; 0 < r; - \x. \norm(x - z) = r; x \ k\ \ norm(f x) \ B\ - \ norm i \ B*(2*pi*r)" - unfolding circlepath_def - by (auto simp: algebra_simps in_path_image_part_circlepath dest!: has_contour_integral_bound_part_circlepath_strong) - -corollary has_contour_integral_bound_circlepath: - "\(f has_contour_integral i) (circlepath z r); - 0 \ B; 0 < r; \x. norm(x - z) = r \ norm(f x) \ B\ - \ norm i \ B*(2*pi*r)" - by (auto intro: has_contour_integral_bound_circlepath_strong) - -proposition contour_integrable_continuous_circlepath: - "continuous_on (path_image (circlepath z r)) f - \ f contour_integrable_on (circlepath z r)" - by (simp add: circlepath_def contour_integrable_continuous_part_circlepath) - -lemma simple_path_circlepath: "simple_path(circlepath z r) \ (r \ 0)" - by (simp add: circlepath_def simple_path_part_circlepath) - -lemma notin_path_image_circlepath [simp]: "cmod (w - z) < r \ w \ path_image (circlepath z r)" - by (simp add: sphere_def dist_norm norm_minus_commute) - -proposition contour_integral_circlepath: - "0 < r \ contour_integral (circlepath z r) (\w. 1 / (w - z)) = 2 * complex_of_real pi * \" - apply (rule contour_integral_unique) - apply (simp add: has_contour_integral_def) - apply (subst has_integral_cong) - apply (simp add: vector_derivative_circlepath01) - using has_integral_const_real [of _ 0 1] - apply (force simp: circlepath) - done - -lemma winding_number_circlepath_centre: "0 < r \ winding_number (circlepath z r) z = 1" - apply (rule winding_number_unique_loop) - apply (simp_all add: sphere_def valid_path_imp_path) - apply (rule_tac x="circlepath z r" in exI) - apply (simp add: sphere_def contour_integral_circlepath) - done - -proposition winding_number_circlepath: - assumes "norm(w - z) < r" shows "winding_number(circlepath z r) w = 1" -proof (cases "w = z") - case True then show ?thesis - using assms winding_number_circlepath_centre by auto -next - case False - have [simp]: "r > 0" - using assms le_less_trans norm_ge_zero by blast - define r' where "r' = norm(w - z)" - have "r' < r" - by (simp add: assms r'_def) - have disjo: "cball z r' \ sphere z r = {}" - using \r' < r\ by (force simp: cball_def sphere_def) - have "winding_number(circlepath z r) w = winding_number(circlepath z r) z" - apply (rule winding_number_around_inside [where s = "cball z r'"]) - apply (simp_all add: disjo order.strict_implies_order winding_number_circlepath_centre) - apply (simp_all add: False r'_def dist_norm norm_minus_commute) - done - also have "... = 1" - by (simp add: winding_number_circlepath_centre) - finally show ?thesis . -qed - - -text\ Hence the Cauchy formula for points inside a circle.\ - -theorem Cauchy_integral_circlepath: - assumes "continuous_on (cball z r) f" "f holomorphic_on (ball z r)" "norm(w - z) < r" - shows "((\u. f u/(u - w)) has_contour_integral (2 * of_real pi * \ * f w)) - (circlepath z r)" -proof - - have "r > 0" - using assms le_less_trans norm_ge_zero by blast - have "((\u. f u / (u - w)) has_contour_integral (2 * pi) * \ * winding_number (circlepath z r) w * f w) - (circlepath z r)" - apply (rule Cauchy_integral_formula_weak [where s = "cball z r" and k = "{}"]) - using assms \r > 0\ - apply (simp_all add: dist_norm norm_minus_commute) - apply (metis at_within_interior dist_norm holomorphic_on_def interior_ball mem_ball norm_minus_commute) - apply (simp add: cball_def sphere_def dist_norm, clarify) - apply (simp add:) - by (metis dist_commute dist_norm less_irrefl) - then show ?thesis - by (simp add: winding_number_circlepath assms) -qed - -corollary Cauchy_integral_circlepath_simple: - assumes "f holomorphic_on cball z r" "norm(w - z) < r" - shows "((\u. f u/(u - w)) has_contour_integral (2 * of_real pi * \ * f w)) - (circlepath z r)" -using assms by (force simp: holomorphic_on_imp_continuous_on holomorphic_on_subset Cauchy_integral_circlepath) - - -lemma no_bounded_connected_component_imp_winding_number_zero: - assumes g: "path g" "path_image g \ s" "pathfinish g = pathstart g" "z \ s" - and nb: "\z. bounded (connected_component_set (- s) z) \ z \ s" - shows "winding_number g z = 0" -apply (rule winding_number_zero_in_outside) -apply (simp_all add: assms) -by (metis nb [of z] \path_image g \ s\ \z \ s\ contra_subsetD mem_Collect_eq outside outside_mono) - -lemma no_bounded_path_component_imp_winding_number_zero: - assumes g: "path g" "path_image g \ s" "pathfinish g = pathstart g" "z \ s" - and nb: "\z. bounded (path_component_set (- s) z) \ z \ s" - shows "winding_number g z = 0" -apply (rule no_bounded_connected_component_imp_winding_number_zero [OF g]) -by (simp add: bounded_subset nb path_component_subset_connected_component) - - -subsection\ Uniform convergence of path integral\ - -text\Uniform convergence when the derivative of the path is bounded, and in particular for the special case of a circle.\ - -proposition contour_integral_uniform_limit: - assumes ev_fint: "eventually (\n::'a. (f n) contour_integrable_on \) F" - and ev_no: "\e. 0 < e \ eventually (\n. \x \ path_image \. norm(f n x - l x) < e) F" - and noleB: "\t. t \ {0..1} \ norm (vector_derivative \ (at t)) \ B" - and \: "valid_path \" - and [simp]: "~ (trivial_limit F)" - shows "l contour_integrable_on \" "((\n. contour_integral \ (f n)) \ contour_integral \ l) F" -proof - - have "0 \ B" by (meson noleB [of 0] atLeastAtMost_iff norm_ge_zero order_refl order_trans zero_le_one) - { fix e::real - assume "0 < e" - then have eB: "0 < e / (\B\ + 1)" by simp - obtain a where fga: "\x. x \ {0..1} \ cmod (f a (\ x) - l (\ x)) < e / (\B\ + 1)" - and inta: "(\t. f a (\ t) * vector_derivative \ (at t)) integrable_on {0..1}" - using eventually_happens [OF eventually_conj [OF ev_no [OF eB] ev_fint]] - by (fastforce simp: contour_integrable_on path_image_def) - have Ble: "B * e / (\B\ + 1) \ e" - using \0 \ B\ \0 < e\ by (simp add: divide_simps) - have "\h. (\x\{0..1}. cmod (l (\ x) * vector_derivative \ (at x) - h x) \ e) \ h integrable_on {0..1}" - apply (rule_tac x="\x. f (a::'a) (\ x) * vector_derivative \ (at x)" in exI) - apply (intro inta conjI ballI) - apply (rule order_trans [OF _ Ble]) - apply (frule noleB) - apply (frule fga) - using \0 \ B\ \0 < e\ - apply (simp add: norm_mult left_diff_distrib [symmetric] norm_minus_commute divide_simps) - apply (drule (1) mult_mono [OF less_imp_le]) - apply (simp_all add: mult_ac) - done - } - then show lintg: "l contour_integrable_on \" - apply (simp add: contour_integrable_on) - apply (blast intro: integrable_uniform_limit_real) - done - { fix e::real - define B' where "B' = B + 1" - have B': "B' > 0" "B' > B" using \0 \ B\ by (auto simp: B'_def) - assume "0 < e" - then have ev_no': "\\<^sub>F n in F. \x\path_image \. 2 * cmod (f n x - l x) < e / B'" - using ev_no [of "e / B' / 2"] B' by (simp add: field_simps) - have ie: "integral {0..1::real} (\x. e / 2) < e" using \0 < e\ by simp - have *: "cmod (f x (\ t) * vector_derivative \ (at t) - l (\ t) * vector_derivative \ (at t)) \ e / 2" - if t: "t\{0..1}" and leB': "2 * cmod (f x (\ t) - l (\ t)) < e / B'" for x t - proof - - have "2 * cmod (f x (\ t) - l (\ t)) * cmod (vector_derivative \ (at t)) \ e * (B/ B')" - using mult_mono [OF less_imp_le [OF leB'] noleB] B' \0 < e\ t by auto - also have "... < e" - by (simp add: B' \0 < e\ mult_imp_div_pos_less) - finally have "2 * cmod (f x (\ t) - l (\ t)) * cmod (vector_derivative \ (at t)) < e" . - then show ?thesis - by (simp add: left_diff_distrib [symmetric] norm_mult) - qed - have "\\<^sub>F x in F. dist (contour_integral \ (f x)) (contour_integral \ l) < e" - apply (rule eventually_mono [OF eventually_conj [OF ev_no' ev_fint]]) - apply (simp add: dist_norm contour_integrable_on path_image_def contour_integral_integral) - apply (simp add: lintg integral_diff [symmetric] contour_integrable_on [symmetric], clarify) - apply (rule le_less_trans [OF integral_norm_bound_integral ie]) - apply (simp add: lintg integrable_diff contour_integrable_on [symmetric]) - apply (blast intro: *)+ - done - } - then show "((\n. contour_integral \ (f n)) \ contour_integral \ l) F" - by (rule tendstoI) -qed - -proposition contour_integral_uniform_limit_circlepath: - assumes ev_fint: "eventually (\n::'a. (f n) contour_integrable_on (circlepath z r)) F" - and ev_no: "\e. 0 < e \ eventually (\n. \x \ path_image (circlepath z r). norm(f n x - l x) < e) F" - and [simp]: "~ (trivial_limit F)" "0 < r" - shows "l contour_integrable_on (circlepath z r)" "((\n. contour_integral (circlepath z r) (f n)) \ contour_integral (circlepath z r) l) F" -by (auto simp: vector_derivative_circlepath norm_mult intro: contour_integral_uniform_limit assms) - - -subsection\ General stepping result for derivative formulas.\ - -lemma sum_sqs_eq: - fixes x::"'a::idom" shows "x * x + y * y = x * (y * 2) \ y = x" - by algebra - -proposition Cauchy_next_derivative: - assumes "continuous_on (path_image \) f'" - and leB: "\t. t \ {0..1} \ norm (vector_derivative \ (at t)) \ B" - and int: "\w. w \ s - path_image \ \ ((\u. f' u / (u - w)^k) has_contour_integral f w) \" - and k: "k \ 0" - and "open s" - and \: "valid_path \" - and w: "w \ s - path_image \" - shows "(\u. f' u / (u - w)^(Suc k)) contour_integrable_on \" - and "(f has_field_derivative (k * contour_integral \ (\u. f' u/(u - w)^(Suc k)))) - (at w)" (is "?thes2") -proof - - have "open (s - path_image \)" using \open s\ closed_valid_path_image \ by blast - then obtain d where "d>0" and d: "ball w d \ s - path_image \" using w - using open_contains_ball by blast - have [simp]: "\n. cmod (1 + of_nat n) = 1 + of_nat n" - by (metis norm_of_nat of_nat_Suc) - have 1: "\\<^sub>F n in at w. (\x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k) - contour_integrable_on \" - apply (simp add: eventually_at) - apply (rule_tac x=d in exI) - apply (simp add: \d > 0\ dist_norm field_simps, clarify) - apply (rule contour_integrable_div [OF contour_integrable_diff]) - using int w d - apply (force simp: dist_norm norm_minus_commute intro!: has_contour_integral_integrable)+ - done - have bim_g: "bounded (image f' (path_image \))" - by (simp add: compact_imp_bounded compact_continuous_image compact_valid_path_image assms) - then obtain C where "C > 0" and C: "\x. \0 \ x; x \ 1\ \ cmod (f' (\ x)) \ C" - by (force simp: bounded_pos path_image_def) - have twom: "\\<^sub>F n in at w. - \x\path_image \. - cmod ((inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / k - inverse (x - w) ^ Suc k) < e" - if "0 < e" for e - proof - - have *: "cmod ((inverse (x - u) ^ k - inverse (x - w) ^ k) / ((u - w) * k) - inverse (x - w) ^ Suc k) < e" - if x: "x \ path_image \" and "u \ w" and uwd: "cmod (u - w) < d/2" - and uw_less: "cmod (u - w) < e * (d / 2) ^ (k+2) / (1 + real k)" - for u x - proof - - define ff where [abs_def]: - "ff n w = - (if n = 0 then inverse(x - w)^k - else if n = 1 then k / (x - w)^(Suc k) - else (k * of_real(Suc k)) / (x - w)^(k + 2))" for n :: nat and w - have km1: "\z::complex. z \ 0 \ z ^ (k - Suc 0) = z ^ k / z" - by (simp add: field_simps) (metis Suc_pred \k \ 0\ neq0_conv power_Suc) - have ff1: "(ff i has_field_derivative ff (Suc i) z) (at z within ball w (d / 2))" - if "z \ ball w (d / 2)" "i \ 1" for i z - proof - - have "z \ path_image \" - using \x \ path_image \\ d that ball_divide_subset_numeral by blast - then have xz[simp]: "x \ z" using \x \ path_image \\ by blast - then have neq: "x * x + z * z \ x * (z * 2)" - by (blast intro: dest!: sum_sqs_eq) - with xz have "\v. v \ 0 \ (x * x + z * z) * v \ (x * (z * 2) * v)" by auto - then have neqq: "\v. v \ 0 \ x * (x * v) + z * (z * v) \ x * (z * (2 * v))" - by (simp add: algebra_simps) - show ?thesis using \i \ 1\ - apply (simp add: ff_def dist_norm Nat.le_Suc_eq km1, safe) - apply (rule derivative_eq_intros | simp add: km1 | simp add: field_simps neq neqq)+ - done - qed - { fix a::real and b::real assume ab: "a > 0" "b > 0" - then have "k * (1 + real k) * (1 / a) \ k * (1 + real k) * (4 / b) \ b \ 4 * a" - apply (subst mult_le_cancel_left_pos) - using \k \ 0\ - apply (auto simp: divide_simps) - done - with ab have "real k * (1 + real k) / a \ (real k * 4 + real k * real k * 4) / b \ b \ 4 * a" - by (simp add: field_simps) - } note canc = this - have ff2: "cmod (ff (Suc 1) v) \ real (k * (k + 1)) / (d / 2) ^ (k + 2)" - if "v \ ball w (d / 2)" for v - proof - - have "d/2 \ cmod (x - v)" using d x that - apply (simp add: dist_norm path_image_def ball_def not_less [symmetric] del: divide_const_simps, clarify) - apply (drule subsetD) - prefer 2 apply blast - apply (metis norm_minus_commute norm_triangle_half_r CollectI) - done - then have "d \ cmod (x - v) * 2" - by (simp add: divide_simps) - then have dpow_le: "d ^ (k+2) \ (cmod (x - v) * 2) ^ (k+2)" - using \0 < d\ order_less_imp_le power_mono by blast - have "x \ v" using that - using \x \ path_image \\ ball_divide_subset_numeral d by fastforce - then show ?thesis - using \d > 0\ - apply (simp add: ff_def norm_mult norm_divide norm_power dist_norm canc) - using dpow_le - apply (simp add: algebra_simps divide_simps mult_less_0_iff) - done - qed - have ub: "u \ ball w (d / 2)" - using uwd by (simp add: dist_commute dist_norm) - have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) - \ (real k * 4 + real k * real k * 4) * (cmod (u - w) * cmod (u - w)) / (d * (d * (d / 2) ^ k))" - using complex_taylor [OF _ ff1 ff2 _ ub, of w, simplified] - by (simp add: ff_def \0 < d\) - then have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) - \ (cmod (u - w) * real k) * (1 + real k) * cmod (u - w) / (d / 2) ^ (k+2)" - by (simp add: field_simps) - then have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) - / (cmod (u - w) * real k) - \ (1 + real k) * cmod (u - w) / (d / 2) ^ (k+2)" - using \k \ 0\ \u \ w\ by (simp add: mult_ac zero_less_mult_iff pos_divide_le_eq) - also have "... < e" - using uw_less \0 < d\ by (simp add: mult_ac divide_simps) - finally have e: "cmod (inverse (x-u)^k - (inverse (x-w)^k + of_nat k * (u-w) / ((x-w) * (x-w)^k))) - / cmod ((u - w) * real k) < e" - by (simp add: norm_mult) - have "x \ u" - using uwd \0 < d\ x d by (force simp: dist_norm ball_def norm_minus_commute) - show ?thesis - apply (rule le_less_trans [OF _ e]) - using \k \ 0\ \x \ u\ \u \ w\ - apply (simp add: field_simps norm_divide [symmetric]) - done - qed - show ?thesis - unfolding eventually_at - apply (rule_tac x = "min (d/2) ((e*(d/2)^(k + 2))/(Suc k))" in exI) - apply (force simp: \d > 0\ dist_norm that simp del: power_Suc intro: *) - done - qed - have 2: "\\<^sub>F n in at w. - \x\path_image \. - cmod (f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k - f' x / (x - w) ^ Suc k) < e" - if "0 < e" for e - proof - - have *: "cmod (f' (\ x) * (inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - - f' (\ x) / ((\ x - w) * (\ x - w) ^ k)) < e" - if ec: "cmod ((inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - - inverse (\ x - w) * inverse (\ x - w) ^ k) < e / C" - and x: "0 \ x" "x \ 1" - for u x - proof (cases "(f' (\ x)) = 0") - case True then show ?thesis by (simp add: \0 < e\) - next - case False - have "cmod (f' (\ x) * (inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - - f' (\ x) / ((\ x - w) * (\ x - w) ^ k)) = - cmod (f' (\ x) * ((inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - - inverse (\ x - w) * inverse (\ x - w) ^ k))" - by (simp add: field_simps) - also have "... = cmod (f' (\ x)) * - cmod ((inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - - inverse (\ x - w) * inverse (\ x - w) ^ k)" - by (simp add: norm_mult) - also have "... < cmod (f' (\ x)) * (e/C)" - apply (rule mult_strict_left_mono [OF ec]) - using False by simp - also have "... \ e" using C - by (metis False \0 < e\ frac_le less_eq_real_def mult.commute pos_le_divide_eq x zero_less_norm_iff) - finally show ?thesis . - qed - show ?thesis - using twom [OF divide_pos_pos [OF that \C > 0\]] unfolding path_image_def - by (force intro: * elim: eventually_mono) - qed - show "(\u. f' u / (u - w) ^ (Suc k)) contour_integrable_on \" - by (rule contour_integral_uniform_limit [OF 1 2 leB \]) auto - have *: "(\n. contour_integral \ (\x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / k)) - \w\ contour_integral \ (\u. f' u / (u - w) ^ (Suc k))" - by (rule contour_integral_uniform_limit [OF 1 2 leB \]) auto - have **: "contour_integral \ (\x. f' x * (inverse (x - u) ^ k - inverse (x - w) ^ k) / ((u - w) * k)) = - (f u - f w) / (u - w) / k" - if "dist u w < d" for u - apply (rule contour_integral_unique) - apply (simp add: diff_divide_distrib algebra_simps) - apply (rule has_contour_integral_diff; rule has_contour_integral_div; simp add: field_simps; rule int) - apply (metis contra_subsetD d dist_commute mem_ball that) - apply (rule w) - done - show ?thes2 - apply (simp add: DERIV_within_iff del: power_Suc) - apply (rule Lim_transform_within [OF tendsto_mult_left [OF *] \0 < d\ ]) - apply (simp add: \k \ 0\ **) - done -qed - -corollary Cauchy_next_derivative_circlepath: - assumes contf: "continuous_on (path_image (circlepath z r)) f" - and int: "\w. w \ ball z r \ ((\u. f u / (u - w)^k) has_contour_integral g w) (circlepath z r)" - and k: "k \ 0" - and w: "w \ ball z r" - shows "(\u. f u / (u - w)^(Suc k)) contour_integrable_on (circlepath z r)" - (is "?thes1") - and "(g has_field_derivative (k * contour_integral (circlepath z r) (\u. f u/(u - w)^(Suc k)))) (at w)" - (is "?thes2") -proof - - have "r > 0" using w - using ball_eq_empty by fastforce - have wim: "w \ ball z r - path_image (circlepath z r)" - using w by (auto simp: dist_norm) - show ?thes1 ?thes2 - by (rule Cauchy_next_derivative [OF contf _ int k open_ball valid_path_circlepath wim, where B = "2 * pi * \r\"]; - auto simp: vector_derivative_circlepath norm_mult)+ -qed - - -text\ In particular, the first derivative formula.\ - -proposition Cauchy_derivative_integral_circlepath: - assumes contf: "continuous_on (cball z r) f" - and holf: "f holomorphic_on ball z r" - and w: "w \ ball z r" - shows "(\u. f u/(u - w)^2) contour_integrable_on (circlepath z r)" - (is "?thes1") - and "(f has_field_derivative (1 / (2 * of_real pi * \) * contour_integral(circlepath z r) (\u. f u / (u - w)^2))) (at w)" - (is "?thes2") -proof - - have [simp]: "r \ 0" using w - using ball_eq_empty by fastforce - have f: "continuous_on (path_image (circlepath z r)) f" - by (rule continuous_on_subset [OF contf]) (force simp add: cball_def sphere_def) - have int: "\w. dist z w < r \ - ((\u. f u / (u - w)) has_contour_integral (\x. 2 * of_real pi * \ * f x) w) (circlepath z r)" - by (rule Cauchy_integral_circlepath [OF contf holf]) (simp add: dist_norm norm_minus_commute) - show ?thes1 - apply (simp add: power2_eq_square) - apply (rule Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1, simplified]) - apply (blast intro: int) - done - have "((\x. 2 * of_real pi * \ * f x) has_field_derivative contour_integral (circlepath z r) (\u. f u / (u - w)^2)) (at w)" - apply (simp add: power2_eq_square) - apply (rule Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1 and g = "\x. 2 * of_real pi * \ * f x", simplified]) - apply (blast intro: int) - done - then have fder: "(f has_field_derivative contour_integral (circlepath z r) (\u. f u / (u - w)^2) / (2 * of_real pi * \)) (at w)" - by (rule DERIV_cdivide [where f = "\x. 2 * of_real pi * \ * f x" and c = "2 * of_real pi * \", simplified]) - show ?thes2 - by simp (rule fder) -qed - -subsection\ Existence of all higher derivatives.\ - -proposition derivative_is_holomorphic: - assumes "open s" - and fder: "\z. z \ s \ (f has_field_derivative f' z) (at z)" - shows "f' holomorphic_on s" -proof - - have *: "\h. (f' has_field_derivative h) (at z)" if "z \ s" for z - proof - - obtain r where "r > 0" and r: "cball z r \ s" - using open_contains_cball \z \ s\ \open s\ by blast - then have holf_cball: "f holomorphic_on cball z r" - apply (simp add: holomorphic_on_def) - using field_differentiable_at_within field_differentiable_def fder by blast - then have "continuous_on (path_image (circlepath z r)) f" - using \r > 0\ by (force elim: holomorphic_on_subset [THEN holomorphic_on_imp_continuous_on]) - then have contfpi: "continuous_on (path_image (circlepath z r)) (\x. 1/(2 * of_real pi*\) * f x)" - by (auto intro: continuous_intros)+ - have contf_cball: "continuous_on (cball z r) f" using holf_cball - by (simp add: holomorphic_on_imp_continuous_on holomorphic_on_subset) - have holf_ball: "f holomorphic_on ball z r" using holf_cball - using ball_subset_cball holomorphic_on_subset by blast - { fix w assume w: "w \ ball z r" - have intf: "(\u. f u / (u - w)\<^sup>2) contour_integrable_on circlepath z r" - by (blast intro: w Cauchy_derivative_integral_circlepath [OF contf_cball holf_ball]) - have fder': "(f has_field_derivative 1 / (2 * of_real pi * \) * contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2)) - (at w)" - by (blast intro: w Cauchy_derivative_integral_circlepath [OF contf_cball holf_ball]) - have f'_eq: "f' w = contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \)" - using fder' ball_subset_cball r w by (force intro: DERIV_unique [OF fder]) - have "((\u. f u / (u - w)\<^sup>2 / (2 * of_real pi * \)) has_contour_integral - contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \)) - (circlepath z r)" - by (rule Cauchy_Integral_Thm.has_contour_integral_div [OF has_contour_integral_integral [OF intf]]) - then have "((\u. f u / (2 * of_real pi * \ * (u - w)\<^sup>2)) has_contour_integral - contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \)) - (circlepath z r)" - by (simp add: algebra_simps) - then have "((\u. f u / (2 * of_real pi * \ * (u - w)\<^sup>2)) has_contour_integral f' w) (circlepath z r)" - by (simp add: f'_eq) - } note * = this - show ?thesis - apply (rule exI) - apply (rule Cauchy_next_derivative_circlepath [OF contfpi, of 2 f', simplified]) - apply (simp_all add: \0 < r\ * dist_norm) - done - qed - show ?thesis - by (simp add: holomorphic_on_open [OF \open s\] *) -qed - -lemma holomorphic_deriv [holomorphic_intros]: - "\f holomorphic_on s; open s\ \ (deriv f) holomorphic_on s" -by (metis DERIV_deriv_iff_field_differentiable at_within_open derivative_is_holomorphic holomorphic_on_def) - -lemma analytic_deriv: "f analytic_on s \ (deriv f) analytic_on s" - using analytic_on_holomorphic holomorphic_deriv by auto - -lemma holomorphic_higher_deriv [holomorphic_intros]: "\f holomorphic_on s; open s\ \ (deriv ^^ n) f holomorphic_on s" - by (induction n) (auto simp: holomorphic_deriv) - -lemma analytic_higher_deriv: "f analytic_on s \ (deriv ^^ n) f analytic_on s" - unfolding analytic_on_def using holomorphic_higher_deriv by blast - -lemma has_field_derivative_higher_deriv: - "\f holomorphic_on s; open s; x \ s\ - \ ((deriv ^^ n) f has_field_derivative (deriv ^^ (Suc n)) f x) (at x)" -by (metis (no_types, hide_lams) DERIV_deriv_iff_field_differentiable at_within_open comp_apply - funpow.simps(2) holomorphic_higher_deriv holomorphic_on_def) - -lemma valid_path_compose_holomorphic: - assumes "valid_path g" and holo:"f holomorphic_on s" and "open s" "path_image g \ s" - shows "valid_path (f o g)" -proof (rule valid_path_compose[OF \valid_path g\]) - fix x assume "x \ path_image g" - then show "\f'. (f has_field_derivative f') (at x)" - using holo holomorphic_on_open[OF \open s\] \path_image g \ s\ by auto -next - have "deriv f holomorphic_on s" - using holomorphic_deriv holo \open s\ by auto - then show "continuous_on (path_image g) (deriv f)" - using assms(4) holomorphic_on_imp_continuous_on holomorphic_on_subset by auto -qed - - -subsection\ Morera's theorem.\ - -lemma Morera_local_triangle_ball: - assumes "\z. z \ s - \ \e a. 0 < e \ z \ ball a e \ continuous_on (ball a e) f \ - (\b c. closed_segment b c \ ball a e - \ contour_integral (linepath a b) f + - contour_integral (linepath b c) f + - contour_integral (linepath c a) f = 0)" - shows "f analytic_on s" -proof - - { fix z assume "z \ s" - with assms obtain e a where - "0 < e" and z: "z \ ball a e" and contf: "continuous_on (ball a e) f" - and 0: "\b c. closed_segment b c \ ball a e - \ contour_integral (linepath a b) f + - contour_integral (linepath b c) f + - contour_integral (linepath c a) f = 0" - by fastforce - have az: "dist a z < e" using mem_ball z by blast - have sb_ball: "ball z (e - dist a z) \ ball a e" - by (simp add: dist_commute ball_subset_ball_iff) - have "\e>0. f holomorphic_on ball z e" - apply (rule_tac x="e - dist a z" in exI) - apply (simp add: az) - apply (rule holomorphic_on_subset [OF _ sb_ball]) - apply (rule derivative_is_holomorphic[OF open_ball]) - apply (rule triangle_contour_integrals_starlike_primitive [OF contf _ open_ball, of a]) - apply (simp_all add: 0 \0 < e\) - apply (meson \0 < e\ centre_in_ball convex_ball convex_contains_segment mem_ball) - done - } - then show ?thesis - by (simp add: analytic_on_def) -qed - -lemma Morera_local_triangle: - assumes "\z. z \ s - \ \t. open t \ z \ t \ continuous_on t f \ - (\a b c. convex hull {a,b,c} \ t - \ contour_integral (linepath a b) f + - contour_integral (linepath b c) f + - contour_integral (linepath c a) f = 0)" - shows "f analytic_on s" -proof - - { fix z assume "z \ s" - with assms obtain t where - "open t" and z: "z \ t" and contf: "continuous_on t f" - and 0: "\a b c. convex hull {a,b,c} \ t - \ contour_integral (linepath a b) f + - contour_integral (linepath b c) f + - contour_integral (linepath c a) f = 0" - by force - then obtain e where "e>0" and e: "ball z e \ t" - using open_contains_ball by blast - have [simp]: "continuous_on (ball z e) f" using contf - using continuous_on_subset e by blast - have "\e a. 0 < e \ - z \ ball a e \ - continuous_on (ball a e) f \ - (\b c. closed_segment b c \ ball a e \ - contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0)" - apply (rule_tac x=e in exI) - apply (rule_tac x=z in exI) - apply (simp add: \e > 0\, clarify) - apply (rule 0) - apply (meson z \0 < e\ centre_in_ball closed_segment_subset convex_ball dual_order.trans e starlike_convex_subset) - done - } - then show ?thesis - by (simp add: Morera_local_triangle_ball) -qed - -proposition Morera_triangle: - "\continuous_on s f; open s; - \a b c. convex hull {a,b,c} \ s - \ contour_integral (linepath a b) f + - contour_integral (linepath b c) f + - contour_integral (linepath c a) f = 0\ - \ f analytic_on s" - using Morera_local_triangle by blast - - - -subsection\ Combining theorems for higher derivatives including Leibniz rule.\ - -lemma higher_deriv_linear [simp]: - "(deriv ^^ n) (\w. c*w) = (\z. if n = 0 then c*z else if n = 1 then c else 0)" - by (induction n) (auto simp: deriv_const deriv_linear) - -lemma higher_deriv_const [simp]: "(deriv ^^ n) (\w. c) = (\w. if n=0 then c else 0)" - by (induction n) (auto simp: deriv_const) - -lemma higher_deriv_ident [simp]: - "(deriv ^^ n) (\w. w) z = (if n = 0 then z else if n = 1 then 1 else 0)" - apply (induction n, simp) - apply (metis higher_deriv_linear lambda_one) - done - -corollary higher_deriv_id [simp]: - "(deriv ^^ n) id z = (if n = 0 then z else if n = 1 then 1 else 0)" - by (simp add: id_def) - -lemma has_complex_derivative_funpow_1: - "\(f has_field_derivative 1) (at z); f z = z\ \ (f^^n has_field_derivative 1) (at z)" - apply (induction n) - apply auto - apply (metis DERIV_ident DERIV_transform_at id_apply zero_less_one) - by (metis DERIV_chain comp_funpow comp_id funpow_swap1 mult.right_neutral) - -proposition higher_deriv_uminus: - assumes "f holomorphic_on s" "open s" and z: "z \ s" - shows "(deriv ^^ n) (\w. -(f w)) z = - ((deriv ^^ n) f z)" -using z -proof (induction n arbitrary: z) - case 0 then show ?case by simp -next - case (Suc n z) - have *: "((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)" - using Suc.prems assms has_field_derivative_higher_deriv by auto - show ?case - apply simp - apply (rule DERIV_imp_deriv) - apply (rule DERIV_transform_within_open [of "\w. -((deriv ^^ n) f w)"]) - apply (rule derivative_eq_intros | rule * refl assms Suc)+ - apply (simp add: Suc) - done -qed - -proposition higher_deriv_add: - fixes z::complex - assumes "f holomorphic_on s" "g holomorphic_on s" "open s" and z: "z \ s" - shows "(deriv ^^ n) (\w. f w + g w) z = (deriv ^^ n) f z + (deriv ^^ n) g z" -using z -proof (induction n arbitrary: z) - case 0 then show ?case by simp -next - case (Suc n z) - have *: "((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)" - "((deriv ^^ n) g has_field_derivative deriv ((deriv ^^ n) g) z) (at z)" - using Suc.prems assms has_field_derivative_higher_deriv by auto - show ?case - apply simp - apply (rule DERIV_imp_deriv) - apply (rule DERIV_transform_within_open [of "\w. (deriv ^^ n) f w + (deriv ^^ n) g w"]) - apply (rule derivative_eq_intros | rule * refl assms Suc)+ - apply (simp add: Suc) - done -qed - -corollary higher_deriv_diff: - fixes z::complex - assumes "f holomorphic_on s" "g holomorphic_on s" "open s" and z: "z \ s" - shows "(deriv ^^ n) (\w. f w - g w) z = (deriv ^^ n) f z - (deriv ^^ n) g z" - apply (simp only: Groups.group_add_class.diff_conv_add_uminus higher_deriv_add) - apply (subst higher_deriv_add) - using assms holomorphic_on_minus apply (auto simp: higher_deriv_uminus) - done - - -lemma bb: "Suc n choose k = (n choose k) + (if k = 0 then 0 else (n choose (k - 1)))" - by (cases k) simp_all - -proposition higher_deriv_mult: - fixes z::complex - assumes "f holomorphic_on s" "g holomorphic_on s" "open s" and z: "z \ s" - shows "(deriv ^^ n) (\w. f w * g w) z = - (\i = 0..n. of_nat (n choose i) * (deriv ^^ i) f z * (deriv ^^ (n - i)) g z)" -using z -proof (induction n arbitrary: z) - case 0 then show ?case by simp -next - case (Suc n z) - have *: "\n. ((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)" - "\n. ((deriv ^^ n) g has_field_derivative deriv ((deriv ^^ n) g) z) (at z)" - using Suc.prems assms has_field_derivative_higher_deriv by auto - have sumeq: "(\i = 0..n. - of_nat (n choose i) * (deriv ((deriv ^^ i) f) z * (deriv ^^ (n - i)) g z + deriv ((deriv ^^ (n - i)) g) z * (deriv ^^ i) f z)) = - g z * deriv ((deriv ^^ n) f) z + (\i = 0..n. (deriv ^^ i) f z * (of_nat (Suc n choose i) * (deriv ^^ (Suc n - i)) g z))" - apply (simp add: bb algebra_simps setsum.distrib) - apply (subst (4) setsum_Suc_reindex) - apply (auto simp: algebra_simps Suc_diff_le intro: setsum.cong) - done - show ?case - apply (simp only: funpow.simps o_apply) - apply (rule DERIV_imp_deriv) - apply (rule DERIV_transform_within_open - [of "\w. (\i = 0..n. of_nat (n choose i) * (deriv ^^ i) f w * (deriv ^^ (n - i)) g w)"]) - apply (simp add: algebra_simps) - apply (rule DERIV_cong [OF DERIV_setsum]) - apply (rule DERIV_cmult) - apply (auto simp: intro: DERIV_mult * sumeq \open s\ Suc.prems Suc.IH [symmetric]) - done -qed - - -proposition higher_deriv_transform_within_open: - fixes z::complex - assumes "f holomorphic_on s" "g holomorphic_on s" "open s" and z: "z \ s" - and fg: "\w. w \ s \ f w = g w" - shows "(deriv ^^ i) f z = (deriv ^^ i) g z" -using z -by (induction i arbitrary: z) - (auto simp: fg intro: complex_derivative_transform_within_open holomorphic_higher_deriv assms) - -proposition higher_deriv_compose_linear: - fixes z::complex - assumes f: "f holomorphic_on t" and s: "open s" and t: "open t" and z: "z \ s" - and fg: "\w. w \ s \ u * w \ t" - shows "(deriv ^^ n) (\w. f (u * w)) z = u^n * (deriv ^^ n) f (u * z)" -using z -proof (induction n arbitrary: z) - case 0 then show ?case by simp -next - case (Suc n z) - have holo0: "f holomorphic_on op * u ` s" - by (meson fg f holomorphic_on_subset image_subset_iff) - have holo1: "(\w. f (u * w)) holomorphic_on s" - apply (rule holomorphic_on_compose [where g=f, unfolded o_def]) - apply (rule holo0 holomorphic_intros)+ - done - have holo2: "(\z. u ^ n * (deriv ^^ n) f (u * z)) holomorphic_on s" - apply (rule holomorphic_intros)+ - apply (rule holomorphic_on_compose [where g="(deriv ^^ n) f", unfolded o_def]) - apply (rule holomorphic_intros) - apply (rule holomorphic_on_subset [where s=t]) - apply (rule holomorphic_intros assms)+ - apply (blast intro: fg) - done - have "deriv ((deriv ^^ n) (\w. f (u * w))) z = deriv (\z. u^n * (deriv ^^ n) f (u*z)) z" - apply (rule complex_derivative_transform_within_open [OF _ holo2 s Suc.prems]) - apply (rule holomorphic_higher_deriv [OF holo1 s]) - apply (simp add: Suc.IH) - done - also have "... = u^n * deriv (\z. (deriv ^^ n) f (u * z)) z" - apply (rule deriv_cmult) - apply (rule analytic_on_imp_differentiable_at [OF _ Suc.prems]) - apply (rule analytic_on_compose_gen [where g="(deriv ^^ n) f" and t=t, unfolded o_def]) - apply (simp add: analytic_on_linear) - apply (simp add: analytic_on_open f holomorphic_higher_deriv t) - apply (blast intro: fg) - done - also have "... = u * u ^ n * deriv ((deriv ^^ n) f) (u * z)" - apply (subst complex_derivative_chain [where g = "(deriv ^^ n) f" and f = "op*u", unfolded o_def]) - apply (rule derivative_intros) - using Suc.prems field_differentiable_def f fg has_field_derivative_higher_deriv t apply blast - apply (simp add: deriv_linear) - done - finally show ?case - by simp -qed - -lemma higher_deriv_add_at: - assumes "f analytic_on {z}" "g analytic_on {z}" - shows "(deriv ^^ n) (\w. f w + g w) z = (deriv ^^ n) f z + (deriv ^^ n) g z" -proof - - have "f analytic_on {z} \ g analytic_on {z}" - using assms by blast - with higher_deriv_add show ?thesis - by (auto simp: analytic_at_two) -qed - -lemma higher_deriv_diff_at: - assumes "f analytic_on {z}" "g analytic_on {z}" - shows "(deriv ^^ n) (\w. f w - g w) z = (deriv ^^ n) f z - (deriv ^^ n) g z" -proof - - have "f analytic_on {z} \ g analytic_on {z}" - using assms by blast - with higher_deriv_diff show ?thesis - by (auto simp: analytic_at_two) -qed - -lemma higher_deriv_uminus_at: - "f analytic_on {z} \ (deriv ^^ n) (\w. -(f w)) z = - ((deriv ^^ n) f z)" - using higher_deriv_uminus - by (auto simp: analytic_at) - -lemma higher_deriv_mult_at: - assumes "f analytic_on {z}" "g analytic_on {z}" - shows "(deriv ^^ n) (\w. f w * g w) z = - (\i = 0..n. of_nat (n choose i) * (deriv ^^ i) f z * (deriv ^^ (n - i)) g z)" -proof - - have "f analytic_on {z} \ g analytic_on {z}" - using assms by blast - with higher_deriv_mult show ?thesis - by (auto simp: analytic_at_two) -qed - - -text\ Nonexistence of isolated singularities and a stronger integral formula.\ - -proposition no_isolated_singularity: - fixes z::complex - assumes f: "continuous_on s f" and holf: "f holomorphic_on (s - k)" and s: "open s" and k: "finite k" - shows "f holomorphic_on s" -proof - - { fix z - assume "z \ s" and cdf: "\x. x\s - k \ f field_differentiable at x" - have "f field_differentiable at z" - proof (cases "z \ k") - case False then show ?thesis by (blast intro: cdf \z \ s\) - next - case True - with finite_set_avoid [OF k, of z] - obtain d where "d>0" and d: "\x. \x\k; x \ z\ \ d \ dist z x" - by blast - obtain e where "e>0" and e: "ball z e \ s" - using s \z \ s\ by (force simp add: open_contains_ball) - have fde: "continuous_on (ball z (min d e)) f" - by (metis Int_iff ball_min_Int continuous_on_subset e f subsetI) - have "\g. \w \ ball z (min d e). (g has_field_derivative f w) (at w within ball z (min d e))" - apply (rule contour_integral_convex_primitive [OF convex_ball fde]) - apply (rule Cauchy_theorem_triangle_cofinite [OF _ k]) - apply (metis continuous_on_subset [OF fde] closed_segment_subset convex_ball starlike_convex_subset) - apply (rule cdf) - apply (metis Diff_iff Int_iff ball_min_Int bot_least contra_subsetD convex_ball e insert_subset - interior_mono interior_subset subset_hull) - done - then have "f holomorphic_on ball z (min d e)" - by (metis open_ball at_within_open derivative_is_holomorphic) - then show ?thesis - unfolding holomorphic_on_def - by (metis open_ball \0 < d\ \0 < e\ at_within_open centre_in_ball min_less_iff_conj) - qed - } - with holf s k show ?thesis - by (simp add: holomorphic_on_open open_Diff finite_imp_closed field_differentiable_def [symmetric]) -qed - -proposition Cauchy_integral_formula_convex: - assumes s: "convex s" and k: "finite k" and contf: "continuous_on s f" - and fcd: "(\x. x \ interior s - k \ f field_differentiable at x)" - and z: "z \ interior s" and vpg: "valid_path \" - and pasz: "path_image \ \ s - {z}" and loop: "pathfinish \ = pathstart \" - shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" - apply (rule Cauchy_integral_formula_weak [OF s finite.emptyI contf]) - apply (simp add: holomorphic_on_open [symmetric] field_differentiable_def) - using no_isolated_singularity [where s = "interior s"] - apply (metis k contf fcd holomorphic_on_open field_differentiable_def continuous_on_subset - has_field_derivative_at_within holomorphic_on_def interior_subset open_interior) - using assms - apply auto - done - - -text\ Formula for higher derivatives.\ - -proposition Cauchy_has_contour_integral_higher_derivative_circlepath: - assumes contf: "continuous_on (cball z r) f" - and holf: "f holomorphic_on ball z r" - and w: "w \ ball z r" - shows "((\u. f u / (u - w) ^ (Suc k)) has_contour_integral ((2 * pi * \) / (fact k) * (deriv ^^ k) f w)) - (circlepath z r)" -using w -proof (induction k arbitrary: w) - case 0 then show ?case - using assms by (auto simp: Cauchy_integral_circlepath dist_commute dist_norm) -next - case (Suc k) - have [simp]: "r > 0" using w - using ball_eq_empty by fastforce - have f: "continuous_on (path_image (circlepath z r)) f" - by (rule continuous_on_subset [OF contf]) (force simp add: cball_def sphere_def less_imp_le) - obtain X where X: "((\u. f u / (u - w) ^ Suc (Suc k)) has_contour_integral X) (circlepath z r)" - using Cauchy_next_derivative_circlepath(1) [OF f Suc.IH _ Suc.prems] - by (auto simp: contour_integrable_on_def) - then have con: "contour_integral (circlepath z r) ((\u. f u / (u - w) ^ Suc (Suc k))) = X" - by (rule contour_integral_unique) - have "\n. ((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) w) (at w)" - using Suc.prems assms has_field_derivative_higher_deriv by auto - then have dnf_diff: "\n. (deriv ^^ n) f field_differentiable (at w)" - by (force simp add: field_differentiable_def) - have "deriv (\w. complex_of_real (2 * pi) * \ / (fact k) * (deriv ^^ k) f w) w = - of_nat (Suc k) * contour_integral (circlepath z r) (\u. f u / (u - w) ^ Suc (Suc k))" - by (force intro!: DERIV_imp_deriv Cauchy_next_derivative_circlepath [OF f Suc.IH _ Suc.prems]) - also have "... = of_nat (Suc k) * X" - by (simp only: con) - finally have "deriv (\w. ((2 * pi) * \ / (fact k)) * (deriv ^^ k) f w) w = of_nat (Suc k) * X" . - then have "((2 * pi) * \ / (fact k)) * deriv (\w. (deriv ^^ k) f w) w = of_nat (Suc k) * X" - by (metis deriv_cmult dnf_diff) - then have "deriv (\w. (deriv ^^ k) f w) w = of_nat (Suc k) * X / ((2 * pi) * \ / (fact k))" - by (simp add: field_simps) - then show ?case - using of_nat_eq_0_iff X by fastforce -qed - -proposition Cauchy_higher_derivative_integral_circlepath: - assumes contf: "continuous_on (cball z r) f" - and holf: "f holomorphic_on ball z r" - and w: "w \ ball z r" - shows "(\u. f u / (u - w)^(Suc k)) contour_integrable_on (circlepath z r)" - (is "?thes1") - and "(deriv ^^ k) f w = (fact k) / (2 * pi * \) * contour_integral(circlepath z r) (\u. f u/(u - w)^(Suc k))" - (is "?thes2") -proof - - have *: "((\u. f u / (u - w) ^ Suc k) has_contour_integral (2 * pi) * \ / (fact k) * (deriv ^^ k) f w) - (circlepath z r)" - using Cauchy_has_contour_integral_higher_derivative_circlepath [OF assms] - by simp - show ?thes1 using * - using contour_integrable_on_def by blast - show ?thes2 - unfolding contour_integral_unique [OF *] by (simp add: divide_simps) -qed - -corollary Cauchy_contour_integral_circlepath: - assumes "continuous_on (cball z r) f" "f holomorphic_on ball z r" "w \ ball z r" - shows "contour_integral(circlepath z r) (\u. f u/(u - w)^(Suc k)) = (2 * pi * \) * (deriv ^^ k) f w / (fact k)" -by (simp add: Cauchy_higher_derivative_integral_circlepath [OF assms]) - -corollary Cauchy_contour_integral_circlepath_2: - assumes "continuous_on (cball z r) f" "f holomorphic_on ball z r" "w \ ball z r" - shows "contour_integral(circlepath z r) (\u. f u/(u - w)^2) = (2 * pi * \) * deriv f w" - using Cauchy_contour_integral_circlepath [OF assms, of 1] - by (simp add: power2_eq_square) - - -subsection\A holomorphic function is analytic, i.e. has local power series.\ - -theorem holomorphic_power_series: - assumes holf: "f holomorphic_on ball z r" - and w: "w \ ball z r" - shows "((\n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)" -proof - - have fh': "f holomorphic_on cball z ((r + dist w z) / 2)" - apply (rule holomorphic_on_subset [OF holf]) - apply (clarsimp simp del: divide_const_simps) - apply (metis add.commute dist_commute le_less_trans mem_ball real_gt_half_sum w) - done - \\Replacing @{term r} and the original (weak) premises\ - obtain r where "0 < r" and holfc: "f holomorphic_on cball z r" and w: "w \ ball z r" - apply (rule that [of "(r + dist w z) / 2"]) - apply (simp_all add: fh') - apply (metis add_0_iff ball_eq_empty dist_nz dist_self empty_iff not_less pos_add_strict w) - apply (metis add_less_cancel_right dist_commute mem_ball mult_2_right w) - done - then have holf: "f holomorphic_on ball z r" and contf: "continuous_on (cball z r) f" - using ball_subset_cball holomorphic_on_subset apply blast - by (simp add: holfc holomorphic_on_imp_continuous_on) - have cint: "\k. (\u. f u / (u - z) ^ Suc k) contour_integrable_on circlepath z r" - apply (rule Cauchy_higher_derivative_integral_circlepath [OF contf holf]) - apply (simp add: \0 < r\) - done - obtain B where "0 < B" and B: "\u. u \ cball z r \ norm(f u) \ B" - by (metis (no_types) bounded_pos compact_cball compact_continuous_image compact_imp_bounded contf image_eqI) - obtain k where k: "0 < k" "k \ r" and wz_eq: "norm(w - z) = r - k" - and kle: "\u. norm(u - z) = r \ k \ norm(u - w)" - apply (rule_tac k = "r - dist z w" in that) - using w - apply (auto simp: dist_norm norm_minus_commute) - by (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute) - have *: "\\<^sub>F n in sequentially. \x\path_image (circlepath z r). - norm ((\k (r - k) / r" "(r - k) / r < 1" using k by auto - obtain n where n: "((r - k) / r) ^ n < e / B * k" - using real_arch_pow_inv [of "e/B*k" "(r - k)/r"] \0 < e\ \0 < B\ k by force - have "norm ((\k N" and r: "r = dist z u" for N u - proof - - have N: "((r - k) / r) ^ N < e / B * k" - apply (rule le_less_trans [OF power_decreasing n]) - using \n \ N\ k by auto - have u [simp]: "(u \ z) \ (u \ w)" - using \0 < r\ r w by auto - have wzu_not1: "(w - z) / (u - z) \ 1" - by (metis (no_types) dist_norm divide_eq_1_iff less_irrefl mem_ball norm_minus_commute r w) - have "norm ((\kk0 < B\ - apply (auto simp: geometric_sum [OF wzu_not1]) - apply (simp add: field_simps norm_mult [symmetric]) - done - also have "... = norm ((u-z) ^ N * (w - u) - ((w - z) ^ N - (u-z) ^ N) * (u-w)) / (r ^ N * norm (u-w)) * norm (f u)" - using \0 < r\ r by (simp add: divide_simps norm_mult norm_divide norm_power dist_norm norm_minus_commute) - also have "... = norm ((w - z) ^ N * (w - u)) / (r ^ N * norm (u - w)) * norm (f u)" - by (simp add: algebra_simps) - also have "... = norm (w - z) ^ N * norm (f u) / r ^ N" - by (simp add: norm_mult norm_power norm_minus_commute) - also have "... \ (((r - k)/r)^N) * B" - using \0 < r\ w k - apply (simp add: divide_simps) - apply (rule mult_mono [OF power_mono]) - apply (auto simp: norm_divide wz_eq norm_power dist_norm norm_minus_commute B r) - done - also have "... < e * k" - using \0 < B\ N by (simp add: divide_simps) - also have "... \ e * norm (u - w)" - using r kle \0 < e\ by (simp add: dist_commute dist_norm) - finally show ?thesis - by (simp add: divide_simps norm_divide del: power_Suc) - qed - with \0 < r\ show ?thesis - by (auto simp: mult_ac less_imp_le eventually_sequentially Ball_def) - qed - have eq: "\\<^sub>F x in sequentially. - contour_integral (circlepath z r) (\u. \kku. f u / (u - z) ^ Suc k) * (w - z) ^ k)" - apply (rule eventuallyI) - apply (subst contour_integral_setsum, simp) - using contour_integrable_lmul [OF cint, of "(w - z) ^ a" for a] apply (simp add: field_simps) - apply (simp only: contour_integral_lmul cint algebra_simps) - done - have "(\k. contour_integral (circlepath z r) (\u. f u/(u - z)^(Suc k)) * (w - z)^k) - sums contour_integral (circlepath z r) (\u. f u/(u - w))" - unfolding sums_def - apply (rule Lim_transform_eventually [OF eq]) - apply (rule contour_integral_uniform_limit_circlepath [OF eventuallyI *]) - apply (rule contour_integrable_setsum, simp) - apply (rule contour_integrable_lmul) - apply (rule Cauchy_higher_derivative_integral_circlepath [OF contf holf]) - using \0 < r\ - apply auto - done - then have "(\k. contour_integral (circlepath z r) (\u. f u/(u - z)^(Suc k)) * (w - z)^k) - sums (2 * of_real pi * \ * f w)" - using w by (auto simp: dist_commute dist_norm contour_integral_unique [OF Cauchy_integral_circlepath_simple [OF holfc]]) - then have "(\k. contour_integral (circlepath z r) (\u. f u / (u - z) ^ Suc k) * (w - z)^k / (\ * (of_real pi * 2))) - sums ((2 * of_real pi * \ * f w) / (\ * (complex_of_real pi * 2)))" - by (rule sums_divide) - then have "(\n. (w - z) ^ n * contour_integral (circlepath z r) (\u. f u / (u - z) ^ Suc n) / (\ * (of_real pi * 2))) - sums f w" - by (simp add: field_simps) - then show ?thesis - by (simp add: field_simps \0 < r\ Cauchy_higher_derivative_integral_circlepath [OF contf holf]) -qed - - -subsection\The Liouville theorem and the Fundamental Theorem of Algebra.\ - -text\ These weak Liouville versions don't even need the derivative formula.\ - -lemma Liouville_weak_0: - assumes holf: "f holomorphic_on UNIV" and inf: "(f \ 0) at_infinity" - shows "f z = 0" -proof (rule ccontr) - assume fz: "f z \ 0" - with inf [unfolded Lim_at_infinity, rule_format, of "norm(f z)/2"] - obtain B where B: "\x. B \ cmod x \ norm (f x) * 2 < cmod (f z)" - by (auto simp: dist_norm) - define R where "R = 1 + \B\ + norm z" - have "R > 0" unfolding R_def - proof - - have "0 \ cmod z + \B\" - by (metis (full_types) add_nonneg_nonneg norm_ge_zero real_norm_def) - then show "0 < 1 + \B\ + cmod z" - by linarith - qed - have *: "((\u. f u / (u - z)) has_contour_integral 2 * complex_of_real pi * \ * f z) (circlepath z R)" - apply (rule Cauchy_integral_circlepath) - using \R > 0\ apply (auto intro: holomorphic_on_subset [OF holf] holomorphic_on_imp_continuous_on)+ - done - have "cmod (x - z) = R \ cmod (f x) * 2 \ cmod (f z)" for x - apply (simp add: R_def) - apply (rule less_imp_le) - apply (rule B) - using norm_triangle_ineq4 [of x z] - apply (auto simp:) - done - with \R > 0\ fz show False - using has_contour_integral_bound_circlepath [OF *, of "norm(f z)/2/R"] - by (auto simp: norm_mult norm_divide divide_simps) -qed - -proposition Liouville_weak: - assumes "f holomorphic_on UNIV" and "(f \ l) at_infinity" - shows "f z = l" - using Liouville_weak_0 [of "\z. f z - l"] - by (simp add: assms holomorphic_on_const holomorphic_on_diff LIM_zero) - - -proposition Liouville_weak_inverse: - assumes "f holomorphic_on UNIV" and unbounded: "\B. eventually (\x. norm (f x) \ B) at_infinity" - obtains z where "f z = 0" -proof - - { assume f: "\z. f z \ 0" - have 1: "(\x. 1 / f x) holomorphic_on UNIV" - by (simp add: holomorphic_on_divide holomorphic_on_const assms f) - have 2: "((\x. 1 / f x) \ 0) at_infinity" - apply (rule tendstoI [OF eventually_mono]) - apply (rule_tac B="2/e" in unbounded) - apply (simp add: dist_norm norm_divide divide_simps mult_ac) - done - have False - using Liouville_weak_0 [OF 1 2] f by simp - } - then show ?thesis - using that by blast -qed - - -text\ In particular we get the Fundamental Theorem of Algebra.\ - -theorem fundamental_theorem_of_algebra: - fixes a :: "nat \ complex" - assumes "a 0 = 0 \ (\i \ {1..n}. a i \ 0)" - obtains z where "(\i\n. a i * z^i) = 0" -using assms -proof (elim disjE bexE) - assume "a 0 = 0" then show ?thesis - by (auto simp: that [of 0]) -next - fix i - assume i: "i \ {1..n}" and nz: "a i \ 0" - have 1: "(\z. \i\n. a i * z^i) holomorphic_on UNIV" - by (rule holomorphic_intros)+ - show ?thesis - apply (rule Liouville_weak_inverse [OF 1]) - apply (rule polyfun_extremal) - apply (rule nz) - using i that - apply (auto simp:) - done -qed - - -subsection\ Weierstrass convergence theorem.\ - -proposition holomorphic_uniform_limit: - assumes cont: "eventually (\n. continuous_on (cball z r) (f n) \ (f n) holomorphic_on ball z r) F" - and lim: "\e. 0 < e \ eventually (\n. \x \ cball z r. norm(f n x - g x) < e) F" - and F: "~ trivial_limit F" - obtains "continuous_on (cball z r) g" "g holomorphic_on ball z r" -proof (cases r "0::real" rule: linorder_cases) - case less then show ?thesis by (force simp add: ball_empty less_imp_le continuous_on_def holomorphic_on_def intro: that) -next - case equal then show ?thesis - by (force simp add: holomorphic_on_def continuous_on_sing intro: that) -next - case greater - have contg: "continuous_on (cball z r) g" - using cont - by (fastforce simp: eventually_conj_iff dist_norm intro: eventually_mono [OF lim] continuous_uniform_limit [OF F]) - have 1: "continuous_on (path_image (circlepath z r)) (\x. 1 / (2 * complex_of_real pi * \) * g x)" - apply (rule continuous_intros continuous_on_subset [OF contg])+ - using \0 < r\ by auto - have 2: "((\u. 1 / (2 * of_real pi * \) * g u / (u - w) ^ 1) has_contour_integral g w) (circlepath z r)" - if w: "w \ ball z r" for w - proof - - define d where "d = (r - norm(w - z))" - have "0 < d" "d \ r" using w by (auto simp: norm_minus_commute d_def dist_norm) - have dle: "\u. cmod (z - u) = r \ d \ cmod (u - w)" - unfolding d_def by (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute) - have ev_int: "\\<^sub>F n in F. (\u. f n u / (u - w)) contour_integrable_on circlepath z r" - apply (rule eventually_mono [OF cont]) - using w - apply (auto intro: Cauchy_higher_derivative_integral_circlepath [where k=0, simplified]) - done - have ev_less: "\\<^sub>F n in F. \x\path_image (circlepath z r). cmod (f n x / (x - w) - g x / (x - w)) < e" - if "e > 0" for e - using greater \0 < d\ \0 < e\ - apply (simp add: norm_divide diff_divide_distrib [symmetric] divide_simps) - apply (rule_tac e1="e * d" in eventually_mono [OF lim]) - apply (force simp: dist_norm intro: dle mult_left_mono less_le_trans)+ - done - have g_cint: "(\u. g u/(u - w)) contour_integrable_on circlepath z r" - by (rule contour_integral_uniform_limit_circlepath [OF ev_int ev_less F \0 < r\]) - have cif_tends_cig: "((\n. contour_integral(circlepath z r) (\u. f n u / (u - w))) \ contour_integral(circlepath z r) (\u. g u/(u - w))) F" - by (rule contour_integral_uniform_limit_circlepath [OF ev_int ev_less F \0 < r\]) - have f_tends_cig: "((\n. 2 * of_real pi * \ * f n w) \ contour_integral (circlepath z r) (\u. g u / (u - w))) F" - apply (rule Lim_transform_eventually [where f = "\n. contour_integral (circlepath z r) (\u. f n u/(u - w))"]) - apply (rule eventually_mono [OF cont]) - apply (rule contour_integral_unique [OF Cauchy_integral_circlepath]) - using w - apply (auto simp: norm_minus_commute dist_norm cif_tends_cig) - done - have "((\n. 2 * of_real pi * \ * f n w) \ 2 * of_real pi * \ * g w) F" - apply (rule tendsto_mult_left [OF tendstoI]) - apply (rule eventually_mono [OF lim], assumption) - using w - apply (force simp add: dist_norm) - done - then have "((\u. g u / (u - w)) has_contour_integral 2 * of_real pi * \ * g w) (circlepath z r)" - using has_contour_integral_integral [OF g_cint] tendsto_unique [OF F f_tends_cig] w - by (force simp add: dist_norm) - then have "((\u. g u / (2 * of_real pi * \ * (u - w))) has_contour_integral g w) (circlepath z r)" - using has_contour_integral_div [where c = "2 * of_real pi * \"] - by (force simp add: field_simps) - then show ?thesis - by (simp add: dist_norm) - qed - show ?thesis - using Cauchy_next_derivative_circlepath(2) [OF 1 2, simplified] - by (fastforce simp add: holomorphic_on_open contg intro: that) -qed - - -text\ Version showing that the limit is the limit of the derivatives.\ - -proposition has_complex_derivative_uniform_limit: - fixes z::complex - assumes cont: "eventually (\n. continuous_on (cball z r) (f n) \ - (\w \ ball z r. ((f n) has_field_derivative (f' n w)) (at w))) F" - and lim: "\e. 0 < e \ eventually (\n. \x \ cball z r. norm(f n x - g x) < e) F" - and F: "~ trivial_limit F" and "0 < r" - obtains g' where - "continuous_on (cball z r) g" - "\w. w \ ball z r \ (g has_field_derivative (g' w)) (at w) \ ((\n. f' n w) \ g' w) F" -proof - - let ?conint = "contour_integral (circlepath z r)" - have g: "continuous_on (cball z r) g" "g holomorphic_on ball z r" - by (rule holomorphic_uniform_limit [OF eventually_mono [OF cont] lim F]; - auto simp: holomorphic_on_open field_differentiable_def)+ - then obtain g' where g': "\x. x \ ball z r \ (g has_field_derivative g' x) (at x)" - using DERIV_deriv_iff_has_field_derivative - by (fastforce simp add: holomorphic_on_open) - then have derg: "\x. x \ ball z r \ deriv g x = g' x" - by (simp add: DERIV_imp_deriv) - have tends_f'n_g': "((\n. f' n w) \ g' w) F" if w: "w \ ball z r" for w - proof - - have eq_f': "?conint (\x. f n x / (x - w)\<^sup>2) - ?conint (\x. g x / (x - w)\<^sup>2) = (f' n w - g' w) * (2 * of_real pi * \)" - if cont_fn: "continuous_on (cball z r) (f n)" - and fnd: "\w. w \ ball z r \ (f n has_field_derivative f' n w) (at w)" for n - proof - - have hol_fn: "f n holomorphic_on ball z r" - using fnd by (force simp add: holomorphic_on_open) - have "(f n has_field_derivative 1 / (2 * of_real pi * \) * ?conint (\u. f n u / (u - w)\<^sup>2)) (at w)" - by (rule Cauchy_derivative_integral_circlepath [OF cont_fn hol_fn w]) - then have f': "f' n w = 1 / (2 * of_real pi * \) * ?conint (\u. f n u / (u - w)\<^sup>2)" - using DERIV_unique [OF fnd] w by blast - show ?thesis - by (simp add: f' Cauchy_contour_integral_circlepath_2 [OF g w] derg [OF w] divide_simps) - qed - define d where "d = (r - norm(w - z))^2" - have "d > 0" - using w by (simp add: dist_commute dist_norm d_def) - have dle: "\y. r = cmod (z - y) \ d \ cmod ((y - w)\<^sup>2)" - apply (simp add: d_def norm_power) - apply (rule power_mono) - apply (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute) - apply (metis diff_ge_0_iff_ge dist_commute dist_norm less_eq_real_def mem_ball w) - done - have 1: "\\<^sub>F n in F. (\x. f n x / (x - w)\<^sup>2) contour_integrable_on circlepath z r" - by (force simp add: holomorphic_on_open intro: w Cauchy_derivative_integral_circlepath eventually_mono [OF cont]) - have 2: "0 < e \ \\<^sub>F n in F. \x \ path_image (circlepath z r). cmod (f n x / (x - w)\<^sup>2 - g x / (x - w)\<^sup>2) < e" for e - using \r > 0\ - apply (simp add: diff_divide_distrib [symmetric] norm_divide divide_simps sphere_def) - apply (rule eventually_mono [OF lim, of "e*d"]) - apply (simp add: \0 < d\) - apply (force simp add: dist_norm dle intro: less_le_trans) - done - have "((\n. contour_integral (circlepath z r) (\x. f n x / (x - w)\<^sup>2)) - \ contour_integral (circlepath z r) ((\x. g x / (x - w)\<^sup>2))) F" - by (rule Cauchy_Integral_Thm.contour_integral_uniform_limit_circlepath [OF 1 2 F \0 < r\]) - then have tendsto_0: "((\n. 1 / (2 * of_real pi * \) * (?conint (\x. f n x / (x - w)\<^sup>2) - ?conint (\x. g x / (x - w)\<^sup>2))) \ 0) F" - using Lim_null by (force intro!: tendsto_mult_right_zero) - have "((\n. f' n w - g' w) \ 0) F" - apply (rule Lim_transform_eventually [OF _ tendsto_0]) - apply (force simp add: divide_simps intro: eq_f' eventually_mono [OF cont]) - done - then show ?thesis using Lim_null by blast - qed - obtain g' where "\w. w \ ball z r \ (g has_field_derivative (g' w)) (at w) \ ((\n. f' n w) \ g' w) F" - by (blast intro: tends_f'n_g' g' ) - then show ?thesis using g - using that by blast -qed - - -subsection\Some more simple/convenient versions for applications.\ - -lemma holomorphic_uniform_sequence: - assumes s: "open s" - and hol_fn: "\n. (f n) holomorphic_on s" - and to_g: "\x. x \ s - \ \d. 0 < d \ cball x d \ s \ - (\e. 0 < e \ eventually (\n. \y \ cball x d. norm(f n y - g y) < e) sequentially)" - shows "g holomorphic_on s" -proof - - have "\f'. (g has_field_derivative f') (at z)" if "z \ s" for z - proof - - obtain r where "0 < r" and r: "cball z r \ s" - and fg: "\e. 0 < e \ eventually (\n. \y \ cball z r. norm(f n y - g y) < e) sequentially" - using to_g [OF \z \ s\] by blast - have *: "\\<^sub>F n in sequentially. continuous_on (cball z r) (f n) \ f n holomorphic_on ball z r" - apply (intro eventuallyI conjI) - using hol_fn holomorphic_on_imp_continuous_on holomorphic_on_subset r apply blast - apply (metis hol_fn holomorphic_on_subset interior_cball interior_subset r) - done - show ?thesis - apply (rule holomorphic_uniform_limit [OF *]) - using \0 < r\ centre_in_ball fg - apply (auto simp: holomorphic_on_open) - done - qed - with s show ?thesis - by (simp add: holomorphic_on_open) -qed - -lemma has_complex_derivative_uniform_sequence: - fixes s :: "complex set" - assumes s: "open s" - and hfd: "\n x. x \ s \ ((f n) has_field_derivative f' n x) (at x)" - and to_g: "\x. x \ s - \ \d. 0 < d \ cball x d \ s \ - (\e. 0 < e \ eventually (\n. \y \ cball x d. norm(f n y - g y) < e) sequentially)" - shows "\g'. \x \ s. (g has_field_derivative g' x) (at x) \ ((\n. f' n x) \ g' x) sequentially" -proof - - have y: "\y. (g has_field_derivative y) (at z) \ (\n. f' n z) \ y" if "z \ s" for z - proof - - obtain r where "0 < r" and r: "cball z r \ s" - and fg: "\e. 0 < e \ eventually (\n. \y \ cball z r. norm(f n y - g y) < e) sequentially" - using to_g [OF \z \ s\] by blast - have *: "\\<^sub>F n in sequentially. continuous_on (cball z r) (f n) \ - (\w \ ball z r. ((f n) has_field_derivative (f' n w)) (at w))" - apply (intro eventuallyI conjI) - apply (meson hfd holomorphic_on_imp_continuous_on holomorphic_on_open holomorphic_on_subset r s) - using ball_subset_cball hfd r apply blast - done - show ?thesis - apply (rule has_complex_derivative_uniform_limit [OF *, of g]) - using \0 < r\ centre_in_ball fg - apply force+ - done - qed - show ?thesis - by (rule bchoice) (blast intro: y) -qed - - -subsection\On analytic functions defined by a series.\ - -lemma series_and_derivative_comparison: - fixes s :: "complex set" - assumes s: "open s" - and h: "summable h" - and hfd: "\n x. x \ s \ (f n has_field_derivative f' n x) (at x)" - and to_g: "\n x. \N \ n; x \ s\ \ norm(f n x) \ h n" - obtains g g' where "\x \ s. ((\n. f n x) sums g x) \ ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" -proof - - obtain g where g: "\e. e>0 \ \N. \n x. N \ n \ x \ s \ dist (\nd>0. cball x d \ s \ (\e>0. \\<^sub>F n in sequentially. \y\cball x d. cmod ((\a s" for x - proof - - obtain d where "d>0" and d: "cball x d \ s" - using open_contains_cball [of "s"] \x \ s\ s by blast - then show ?thesis - apply (rule_tac x=d in exI) - apply (auto simp: dist_norm eventually_sequentially) - apply (metis g contra_subsetD dist_norm) - done - qed - have "(\x\s. (\n. \i g x)" - using g by (force simp add: lim_sequentially) - moreover have "\g'. \x\s. (g has_field_derivative g' x) (at x) \ (\n. \i g' x" - by (rule has_complex_derivative_uniform_sequence [OF s]) (auto intro: * hfd DERIV_setsum)+ - ultimately show ?thesis - by (force simp add: sums_def conj_commute intro: that) -qed - -text\A version where we only have local uniform/comparative convergence.\ - -lemma series_and_derivative_comparison_local: - fixes s :: "complex set" - assumes s: "open s" - and hfd: "\n x. x \ s \ (f n has_field_derivative f' n x) (at x)" - and to_g: "\x. x \ s \ - \d h N. 0 < d \ summable h \ (\n y. N \ n \ y \ ball x d \ norm(f n y) \ h n)" - shows "\g g'. \x \ s. ((\n. f n x) sums g x) \ ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" -proof - - have "\y. (\n. f n z) sums (\n. f n z) \ (\n. f' n z) sums y \ ((\x. \n. f n x) has_field_derivative y) (at z)" - if "z \ s" for z - proof - - obtain d h N where "0 < d" "summable h" and le_h: "\n y. \N \ n; y \ ball z d\ \ norm(f n y) \ h n" - using to_g \z \ s\ by meson - then obtain r where "r>0" and r: "ball z r \ ball z d \ s" using \z \ s\ s - by (metis Int_iff open_ball centre_in_ball open_Int open_contains_ball_eq) - have 1: "open (ball z d \ s)" - by (simp add: open_Int s) - have 2: "\n x. x \ ball z d \ s \ (f n has_field_derivative f' n x) (at x)" - by (auto simp: hfd) - obtain g g' where gg': "\x \ ball z d \ s. ((\n. f n x) sums g x) \ - ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" - by (auto intro: le_h series_and_derivative_comparison [OF 1 \summable h\ hfd]) - then have "(\n. f' n z) sums g' z" - by (meson \0 < r\ centre_in_ball contra_subsetD r) - moreover have "(\n. f n z) sums (\n. f n z)" - by (metis summable_comparison_test' summable_sums centre_in_ball \0 < d\ \summable h\ le_h) - moreover have "((\x. \n. f n x) has_field_derivative g' z) (at z)" - apply (rule_tac f=g in DERIV_transform_at [OF _ \0 < r\]) - apply (simp add: gg' \z \ s\ \0 < d\) - apply (metis (full_types) contra_subsetD dist_commute gg' mem_ball r sums_unique) - done - ultimately show ?thesis by auto - qed - then show ?thesis - by (rule_tac x="\x. suminf (\n. f n x)" in exI) meson -qed - - -text\Sometimes convenient to compare with a complex series of positive reals. (?)\ - -lemma series_and_derivative_comparison_complex: - fixes s :: "complex set" - assumes s: "open s" - and hfd: "\n x. x \ s \ (f n has_field_derivative f' n x) (at x)" - and to_g: "\x. x \ s \ - \d h N. 0 < d \ summable h \ range h \ nonneg_Reals \ (\n y. N \ n \ y \ ball x d \ cmod(f n y) \ cmod (h n))" - shows "\g g'. \x \ s. ((\n. f n x) sums g x) \ ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" -apply (rule series_and_derivative_comparison_local [OF s hfd], assumption) -apply (frule to_g) -apply (erule ex_forward) -apply (erule exE) -apply (rule_tac x="Re o h" in exI) -apply (erule ex_forward) -apply (simp add: summable_Re o_def ) -apply (elim conjE all_forward) -apply (simp add: nonneg_Reals_cmod_eq_Re image_subset_iff) -done - - -text\In particular, a power series is analytic inside circle of convergence.\ - -lemma power_series_and_derivative_0: - fixes a :: "nat \ complex" and r::real - assumes "summable (\n. a n * r^n)" - shows "\g g'. \z. cmod z < r \ - ((\n. a n * z^n) sums g z) \ ((\n. of_nat n * a n * z^(n - 1)) sums g' z) \ (g has_field_derivative g' z) (at z)" -proof (cases "0 < r") - case True - have der: "\n z. ((\x. a n * x ^ n) has_field_derivative of_nat n * a n * z ^ (n - 1)) (at z)" - by (rule derivative_eq_intros | simp)+ - have y_le: "\cmod (z - y) * 2 < r - cmod z\ \ cmod y \ cmod (of_real r + of_real (cmod z)) / 2" for z y - using \r > 0\ - apply (auto simp: algebra_simps norm_mult norm_divide norm_power of_real_add [symmetric] simp del: of_real_add) - using norm_triangle_ineq2 [of y z] - apply (simp only: diff_le_eq norm_minus_commute mult_2) - done - have "summable (\n. a n * complex_of_real r ^ n)" - using assms \r > 0\ by simp - moreover have "\z. cmod z < r \ cmod ((of_real r + of_real (cmod z)) / 2) < cmod (of_real r)" - using \r > 0\ - by (simp add: of_real_add [symmetric] del: of_real_add) - ultimately have sum: "\z. cmod z < r \ summable (\n. of_real (cmod (a n)) * ((of_real r + complex_of_real (cmod z)) / 2) ^ n)" - by (rule power_series_conv_imp_absconv_weak) - have "\g g'. \z \ ball 0 r. (\n. (a n) * z ^ n) sums g z \ - (\n. of_nat n * (a n) * z ^ (n - 1)) sums g' z \ (g has_field_derivative g' z) (at z)" - apply (rule series_and_derivative_comparison_complex [OF open_ball der]) - apply (rule_tac x="(r - norm z)/2" in exI) - apply (simp add: dist_norm) - apply (rule_tac x="\n. of_real(norm(a n)*((r + norm z)/2)^n)" in exI) - using \r > 0\ - apply (auto simp: sum nonneg_Reals_divide_I) - apply (rule_tac x=0 in exI) - apply (force simp: norm_mult norm_divide norm_power intro!: mult_left_mono power_mono y_le) - done - then show ?thesis - by (simp add: dist_0_norm ball_def) -next - case False then show ?thesis - apply (simp add: not_less) - using less_le_trans norm_not_less_zero by blast -qed - -proposition power_series_and_derivative: - fixes a :: "nat \ complex" and r::real - assumes "summable (\n. a n * r^n)" - obtains g g' where "\z \ ball w r. - ((\n. a n * (z - w) ^ n) sums g z) \ ((\n. of_nat n * a n * (z - w) ^ (n - 1)) sums g' z) \ - (g has_field_derivative g' z) (at z)" - using power_series_and_derivative_0 [OF assms] - apply clarify - apply (rule_tac g="(\z. g(z - w))" in that) - using DERIV_shift [where z="-w"] - apply (auto simp: norm_minus_commute Ball_def dist_norm) - done - -proposition power_series_holomorphic: - assumes "\w. w \ ball z r \ ((\n. a n*(w - z)^n) sums f w)" - shows "f holomorphic_on ball z r" -proof - - have "\f'. (f has_field_derivative f') (at w)" if w: "dist z w < r" for w - proof - - have inb: "z + complex_of_real ((dist z w + r) / 2) \ ball z r" - proof - - have wz: "cmod (w - z) < r" using w - by (auto simp: divide_simps dist_norm norm_minus_commute) - then have "0 \ r" - by (meson less_eq_real_def norm_ge_zero order_trans) - show ?thesis - using w by (simp add: dist_norm \0\r\ of_real_add [symmetric] del: of_real_add) - qed - have sum: "summable (\n. a n * of_real (((cmod (z - w) + r) / 2) ^ n))" - using assms [OF inb] by (force simp add: summable_def dist_norm) - obtain g g' where gg': "\u. u \ ball z ((cmod (z - w) + r) / 2) \ - (\n. a n * (u - z) ^ n) sums g u \ - (\n. of_nat n * a n * (u - z) ^ (n - 1)) sums g' u \ (g has_field_derivative g' u) (at u)" - by (rule power_series_and_derivative [OF sum, of z]) fastforce - have [simp]: "g u = f u" if "cmod (u - w) < (r - cmod (z - w)) / 2" for u - proof - - have less: "cmod (z - u) * 2 < cmod (z - w) + r" - using that dist_triangle2 [of z u w] - by (simp add: dist_norm [symmetric] algebra_simps) - show ?thesis - apply (rule sums_unique2 [of "\n. a n*(u - z)^n"]) - using gg' [of u] less w - apply (auto simp: assms dist_norm) - done - qed - show ?thesis - apply (rule_tac x="g' w" in exI) - apply (rule DERIV_transform_at [where f=g and d="(r - norm(z - w))/2"]) - using w gg' [of w] - apply (auto simp: dist_norm) - done - qed - then show ?thesis by (simp add: holomorphic_on_open) -qed - -corollary holomorphic_iff_power_series: - "f holomorphic_on ball z r \ - (\w \ ball z r. (\n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)" - apply (intro iffI ballI) - using holomorphic_power_series apply force - apply (rule power_series_holomorphic [where a = "\n. (deriv ^^ n) f z / (fact n)"]) - apply force - done - -corollary power_series_analytic: - "(\w. w \ ball z r \ (\n. a n*(w - z)^n) sums f w) \ f analytic_on ball z r" - by (force simp add: analytic_on_open intro!: power_series_holomorphic) - -corollary analytic_iff_power_series: - "f analytic_on ball z r \ - (\w \ ball z r. (\n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)" - by (simp add: analytic_on_open holomorphic_iff_power_series) - - -subsection\Equality between holomorphic functions, on open ball then connected set.\ - -lemma holomorphic_fun_eq_on_ball: - "\f holomorphic_on ball z r; g holomorphic_on ball z r; - w \ ball z r; - \n. (deriv ^^ n) f z = (deriv ^^ n) g z\ - \ f w = g w" - apply (rule sums_unique2 [of "\n. (deriv ^^ n) f z / (fact n) * (w - z)^n"]) - apply (auto simp: holomorphic_iff_power_series) - done - -lemma holomorphic_fun_eq_0_on_ball: - "\f holomorphic_on ball z r; w \ ball z r; - \n. (deriv ^^ n) f z = 0\ - \ f w = 0" - apply (rule sums_unique2 [of "\n. (deriv ^^ n) f z / (fact n) * (w - z)^n"]) - apply (auto simp: holomorphic_iff_power_series) - done - -lemma holomorphic_fun_eq_0_on_connected: - assumes holf: "f holomorphic_on s" and "open s" - and cons: "connected s" - and der: "\n. (deriv ^^ n) f z = 0" - and "z \ s" "w \ s" - shows "f w = 0" -proof - - have *: "\x e. \ \xa. (deriv ^^ xa) f x = 0; ball x e \ s\ - \ ball x e \ (\n. {w \ s. (deriv ^^ n) f w = 0})" - apply auto - apply (rule holomorphic_fun_eq_0_on_ball [OF holomorphic_higher_deriv]) - apply (rule holomorphic_on_subset [OF holf], simp_all) - by (metis funpow_add o_apply) - have 1: "openin (subtopology euclidean s) (\n. {w \ s. (deriv ^^ n) f w = 0})" - apply (rule open_subset, force) - using \open s\ - apply (simp add: open_contains_ball Ball_def) - apply (erule all_forward) - using "*" by auto blast+ - have 2: "closedin (subtopology euclidean s) (\n. {w \ s. (deriv ^^ n) f w = 0})" - using assms - by (auto intro: continuous_closedin_preimage_constant holomorphic_on_imp_continuous_on holomorphic_higher_deriv) - obtain e where "e>0" and e: "ball w e \ s" using openE [OF \open s\ \w \ s\] . - then have holfb: "f holomorphic_on ball w e" - using holf holomorphic_on_subset by blast - have 3: "(\n. {w \ s. (deriv ^^ n) f w = 0}) = s \ f w = 0" - using \e>0\ e by (force intro: holomorphic_fun_eq_0_on_ball [OF holfb]) - show ?thesis - using cons der \z \ s\ - apply (simp add: connected_clopen) - apply (drule_tac x="\n. {w \ s. (deriv ^^ n) f w = 0}" in spec) - apply (auto simp: 1 2 3) - done -qed - -lemma holomorphic_fun_eq_on_connected: - assumes "f holomorphic_on s" "g holomorphic_on s" and "open s" "connected s" - and "\n. (deriv ^^ n) f z = (deriv ^^ n) g z" - and "z \ s" "w \ s" - shows "f w = g w" - apply (rule holomorphic_fun_eq_0_on_connected [of "\x. f x - g x" s z, simplified]) - apply (intro assms holomorphic_intros) - using assms apply simp_all - apply (subst higher_deriv_diff, auto) - done - -lemma holomorphic_fun_eq_const_on_connected: - assumes holf: "f holomorphic_on s" and "open s" - and cons: "connected s" - and der: "\n. 0 < n \ (deriv ^^ n) f z = 0" - and "z \ s" "w \ s" - shows "f w = f z" - apply (rule holomorphic_fun_eq_0_on_connected [of "\w. f w - f z" s z, simplified]) - apply (intro assms holomorphic_intros) - using assms apply simp_all - apply (subst higher_deriv_diff) - apply (intro holomorphic_intros | simp)+ - done - - -subsection\Some basic lemmas about poles/singularities.\ - -lemma pole_lemma: - assumes holf: "f holomorphic_on s" and a: "a \ interior s" - shows "(\z. if z = a then deriv f a - else (f z - f a) / (z - a)) holomorphic_on s" (is "?F holomorphic_on s") -proof - - have F1: "?F field_differentiable (at u within s)" if "u \ s" "u \ a" for u - proof - - have fcd: "f field_differentiable at u within s" - using holf holomorphic_on_def by (simp add: \u \ s\) - have cd: "(\z. (f z - f a) / (z - a)) field_differentiable at u within s" - by (rule fcd derivative_intros | simp add: that)+ - have "0 < dist a u" using that dist_nz by blast - then show ?thesis - by (rule field_differentiable_transform_within [OF _ _ _ cd]) (auto simp: \u \ s\) - qed - have F2: "?F field_differentiable at a" if "0 < e" "ball a e \ s" for e - proof - - have holfb: "f holomorphic_on ball a e" - by (rule holomorphic_on_subset [OF holf \ball a e \ s\]) - have 2: "?F holomorphic_on ball a e - {a}" - apply (rule holomorphic_on_subset [where s = "s - {a}"]) - apply (simp add: holomorphic_on_def field_differentiable_def [symmetric]) - using mem_ball that - apply (auto intro: F1 field_differentiable_within_subset) - done - have "isCont (\z. if z = a then deriv f a else (f z - f a) / (z - a)) x" - if "dist a x < e" for x - proof (cases "x=a") - case True then show ?thesis - using holfb \0 < e\ - apply (simp add: holomorphic_on_open field_differentiable_def [symmetric]) - apply (drule_tac x=a in bspec) - apply (auto simp: DERIV_deriv_iff_field_differentiable [symmetric] continuous_at DERIV_iff2 - elim: rev_iffD1 [OF _ LIM_equal]) - done - next - case False with 2 that show ?thesis - by (force simp: holomorphic_on_open open_Diff field_differentiable_def [symmetric] field_differentiable_imp_continuous_at) - qed - then have 1: "continuous_on (ball a e) ?F" - by (clarsimp simp: continuous_on_eq_continuous_at) - have "?F holomorphic_on ball a e" - by (auto intro: no_isolated_singularity [OF 1 2]) - with that show ?thesis - by (simp add: holomorphic_on_open field_differentiable_def [symmetric] - field_differentiable_at_within) - qed - show ?thesis - proof - fix x assume "x \ s" show "?F field_differentiable at x within s" - proof (cases "x=a") - case True then show ?thesis - using a by (auto simp: mem_interior intro: field_differentiable_at_within F2) - next - case False with F1 \x \ s\ - show ?thesis by blast - qed - qed -qed - -proposition pole_theorem: - assumes holg: "g holomorphic_on s" and a: "a \ interior s" - and eq: "\z. z \ s - {a} \ g z = (z - a) * f z" - shows "(\z. if z = a then deriv g a - else f z - g a/(z - a)) holomorphic_on s" - using pole_lemma [OF holg a] - by (rule holomorphic_transform) (simp add: eq divide_simps) - -lemma pole_lemma_open: - assumes "f holomorphic_on s" "open s" - shows "(\z. if z = a then deriv f a else (f z - f a)/(z - a)) holomorphic_on s" -proof (cases "a \ s") - case True with assms interior_eq pole_lemma - show ?thesis by fastforce -next - case False with assms show ?thesis - apply (simp add: holomorphic_on_def field_differentiable_def [symmetric], clarify) - apply (rule field_differentiable_transform_within [where f = "\z. (f z - f a)/(z - a)" and d = 1]) - apply (rule derivative_intros | force)+ - done -qed - -proposition pole_theorem_open: - assumes holg: "g holomorphic_on s" and s: "open s" - and eq: "\z. z \ s - {a} \ g z = (z - a) * f z" - shows "(\z. if z = a then deriv g a - else f z - g a/(z - a)) holomorphic_on s" - using pole_lemma_open [OF holg s] - by (rule holomorphic_transform) (auto simp: eq divide_simps) - -proposition pole_theorem_0: - assumes holg: "g holomorphic_on s" and a: "a \ interior s" - and eq: "\z. z \ s - {a} \ g z = (z - a) * f z" - and [simp]: "f a = deriv g a" "g a = 0" - shows "f holomorphic_on s" - using pole_theorem [OF holg a eq] - by (rule holomorphic_transform) (auto simp: eq divide_simps) - -proposition pole_theorem_open_0: - assumes holg: "g holomorphic_on s" and s: "open s" - and eq: "\z. z \ s - {a} \ g z = (z - a) * f z" - and [simp]: "f a = deriv g a" "g a = 0" - shows "f holomorphic_on s" - using pole_theorem_open [OF holg s eq] - by (rule holomorphic_transform) (auto simp: eq divide_simps) - -lemma pole_theorem_analytic: - assumes g: "g analytic_on s" - and eq: "\z. z \ s - \ \d. 0 < d \ (\w \ ball z d - {a}. g w = (w - a) * f w)" - shows "(\z. if z = a then deriv g a - else f z - g a/(z - a)) analytic_on s" -using g -apply (simp add: analytic_on_def Ball_def) -apply (safe elim!: all_forward dest!: eq) -apply (rule_tac x="min d e" in exI, simp) -apply (rule pole_theorem_open) -apply (auto simp: holomorphic_on_subset subset_ball) -done - -lemma pole_theorem_analytic_0: - assumes g: "g analytic_on s" - and eq: "\z. z \ s \ \d. 0 < d \ (\w \ ball z d - {a}. g w = (w - a) * f w)" - and [simp]: "f a = deriv g a" "g a = 0" - shows "f analytic_on s" -proof - - have [simp]: "(\z. if z = a then deriv g a else f z - g a / (z - a)) = f" - by auto - show ?thesis - using pole_theorem_analytic [OF g eq] by simp -qed - -lemma pole_theorem_analytic_open_superset: - assumes g: "g analytic_on s" and "s \ t" "open t" - and eq: "\z. z \ t - {a} \ g z = (z - a) * f z" - shows "(\z. if z = a then deriv g a - else f z - g a/(z - a)) analytic_on s" - apply (rule pole_theorem_analytic [OF g]) - apply (rule openE [OF \open t\]) - using assms eq by auto - -lemma pole_theorem_analytic_open_superset_0: - assumes g: "g analytic_on s" "s \ t" "open t" "\z. z \ t - {a} \ g z = (z - a) * f z" - and [simp]: "f a = deriv g a" "g a = 0" - shows "f analytic_on s" -proof - - have [simp]: "(\z. if z = a then deriv g a else f z - g a / (z - a)) = f" - by auto - have "(\z. if z = a then deriv g a else f z - g a/(z - a)) analytic_on s" - by (rule pole_theorem_analytic_open_superset [OF g]) - then show ?thesis by simp -qed - - - -subsection\General, homology form of Cauchy's theorem.\ - -text\Proof is based on Dixon's, as presented in Lang's "Complex Analysis" book (page 147).\ - -lemma contour_integral_continuous_on_linepath_2D: - assumes "open u" and cont_dw: "\w. w \ u \ F w contour_integrable_on (linepath a b)" - and cond_uu: "continuous_on (u \ u) (\(x,y). F x y)" - and abu: "closed_segment a b \ u" - shows "continuous_on u (\w. contour_integral (linepath a b) (F w))" -proof - - have *: "\d>0. \x'\u. dist x' w < d \ - dist (contour_integral (linepath a b) (F x')) - (contour_integral (linepath a b) (F w)) \ \" - if "w \ u" "0 < \" "a \ b" for w \ - proof - - obtain \ where "\>0" and \: "cball w \ \ u" using open_contains_cball \open u\ \w \ u\ by force - let ?TZ = "{(t,z) |t z. t \ cball w \ \ z \ closed_segment a b}" - have "uniformly_continuous_on ?TZ (\(x,y). F x y)" - apply (rule compact_uniformly_continuous) - apply (rule continuous_on_subset[OF cond_uu]) - using abu \ - apply (auto simp: compact_Times simp del: mem_cball) - done - then obtain \ where "\>0" - and \: "\x x'. \x\?TZ; x'\?TZ; dist x' x < \\ \ - dist ((\(x,y). F x y) x') ((\(x,y). F x y) x) < \/norm(b - a)" - apply (rule uniformly_continuous_onE [where e = "\/norm(b - a)"]) - using \0 < \\ \a \ b\ by auto - have \: "\norm (w - x1) \ \; x2 \ closed_segment a b; - norm (w - x1') \ \; x2' \ closed_segment a b; norm ((x1', x2') - (x1, x2)) < \\ - \ norm (F x1' x2' - F x1 x2) \ \ / cmod (b - a)" - for x1 x2 x1' x2' - using \ [of "(x1,x2)" "(x1',x2')"] by (force simp add: dist_norm) - have le_ee: "cmod (contour_integral (linepath a b) (\x. F x' x - F w x)) \ \" - if "x' \ u" "cmod (x' - w) < \" "cmod (x' - w) < \" for x' - proof - - have "cmod (contour_integral (linepath a b) (\x. F x' x - F w x)) \ \/norm(b - a) * norm(b - a)" - apply (rule has_contour_integral_bound_linepath [OF has_contour_integral_integral _ \]) - apply (rule contour_integrable_diff [OF cont_dw cont_dw]) - using \0 < \\ \a \ b\ \0 < \\ \w \ u\ that - apply (auto simp: norm_minus_commute) - done - also have "... = \" using \a \ b\ by simp - finally show ?thesis . - qed - show ?thesis - apply (rule_tac x="min \ \" in exI) - using \0 < \\ \0 < \\ - apply (auto simp: dist_norm contour_integral_diff [OF cont_dw cont_dw, symmetric] \w \ u\ intro: le_ee) - done - qed - show ?thesis - apply (rule continuous_onI) - apply (cases "a=b") - apply (auto intro: *) - done -qed - -text\This version has @{term"polynomial_function \"} as an additional assumption.\ -lemma Cauchy_integral_formula_global_weak: - assumes u: "open u" and holf: "f holomorphic_on u" - and z: "z \ u" and \: "polynomial_function \" - and pasz: "path_image \ \ u - {z}" and loop: "pathfinish \ = pathstart \" - and zero: "\w. w \ u \ winding_number \ w = 0" - shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" -proof - - obtain \' where pf\': "polynomial_function \'" and \': "\x. (\ has_vector_derivative (\' x)) (at x)" - using has_vector_derivative_polynomial_function [OF \] by blast - then have "bounded(path_image \')" - by (simp add: path_image_def compact_imp_bounded compact_continuous_image continuous_on_polymonial_function) - then obtain B where "B>0" and B: "\x. x \ path_image \' \ norm x \ B" - using bounded_pos by force - define d where [abs_def]: "d z w = (if w = z then deriv f z else (f w - f z)/(w - z))" for z w - define v where "v = {w. w \ path_image \ \ winding_number \ w = 0}" - have "path \" "valid_path \" using \ - by (auto simp: path_polynomial_function valid_path_polynomial_function) - then have ov: "open v" - by (simp add: v_def open_winding_number_levelsets loop) - have uv_Un: "u \ v = UNIV" - using pasz zero by (auto simp: v_def) - have conf: "continuous_on u f" - by (metis holf holomorphic_on_imp_continuous_on) - have hol_d: "(d y) holomorphic_on u" if "y \ u" for y - proof - - have *: "(\c. if c = y then deriv f y else (f c - f y) / (c - y)) holomorphic_on u" - by (simp add: holf pole_lemma_open u) - then have "isCont (\x. if x = y then deriv f y else (f x - f y) / (x - y)) y" - using at_within_open field_differentiable_imp_continuous_at holomorphic_on_def that u by fastforce - then have "continuous_on u (d y)" - apply (simp add: d_def continuous_on_eq_continuous_at u, clarify) - using * holomorphic_on_def - by (meson field_differentiable_within_open field_differentiable_imp_continuous_at u) - moreover have "d y holomorphic_on u - {y}" - apply (simp add: d_def holomorphic_on_open u open_delete field_differentiable_def [symmetric]) - apply (intro ballI) - apply (rename_tac w) - apply (rule_tac d="dist w y" and f = "\w. (f w - f y)/(w - y)" in field_differentiable_transform_within) - apply (auto simp: dist_pos_lt dist_commute intro!: derivative_intros) - using analytic_on_imp_differentiable_at analytic_on_open holf u apply blast - done - ultimately show ?thesis - by (rule no_isolated_singularity) (auto simp: u) - qed - have cint_fxy: "(\x. (f x - f y) / (x - y)) contour_integrable_on \" if "y \ path_image \" for y - apply (rule contour_integrable_holomorphic_simple [where s = "u-{y}"]) - using \valid_path \\ pasz - apply (auto simp: u open_delete) - apply (rule continuous_intros holomorphic_intros continuous_on_subset [OF conf] holomorphic_on_subset [OF holf] | - force simp add: that)+ - done - define h where - "h z = (if z \ u then contour_integral \ (d z) else contour_integral \ (\w. f w/(w - z)))" for z - have U: "\z. z \ u \ ((d z) has_contour_integral h z) \" - apply (simp add: h_def) - apply (rule has_contour_integral_integral [OF contour_integrable_holomorphic_simple [where s=u]]) - using u pasz \valid_path \\ - apply (auto intro: holomorphic_on_imp_continuous_on hol_d) - done - have V: "((\w. f w / (w - z)) has_contour_integral h z) \" if z: "z \ v" for z - proof - - have 0: "0 = (f z) * 2 * of_real (2 * pi) * \ * winding_number \ z" - using v_def z by auto - then have "((\x. 1 / (x - z)) has_contour_integral 0) \" - using z v_def has_contour_integral_winding_number [OF \valid_path \\] by fastforce - then have "((\x. f z * (1 / (x - z))) has_contour_integral 0) \" - using has_contour_integral_lmul by fastforce - then have "((\x. f z / (x - z)) has_contour_integral 0) \" - by (simp add: divide_simps) - moreover have "((\x. (f x - f z) / (x - z)) has_contour_integral contour_integral \ (d z)) \" - using z - apply (auto simp: v_def) - apply (metis (no_types, lifting) contour_integrable_eq d_def has_contour_integral_eq has_contour_integral_integral cint_fxy) - done - ultimately have *: "((\x. f z / (x - z) + (f x - f z) / (x - z)) has_contour_integral (0 + contour_integral \ (d z))) \" - by (rule has_contour_integral_add) - have "((\w. f w / (w - z)) has_contour_integral contour_integral \ (d z)) \" - if "z \ u" - using * by (auto simp: divide_simps has_contour_integral_eq) - moreover have "((\w. f w / (w - z)) has_contour_integral contour_integral \ (\w. f w / (w - z))) \" - if "z \ u" - apply (rule has_contour_integral_integral [OF contour_integrable_holomorphic_simple [where s=u]]) - using u pasz \valid_path \\ that - apply (auto intro: holomorphic_on_imp_continuous_on hol_d) - apply (rule continuous_intros conf holomorphic_intros holf | force)+ - done - ultimately show ?thesis - using z by (simp add: h_def) - qed - have znot: "z \ path_image \" - using pasz by blast - obtain d0 where "d0>0" and d0: "\x y. x \ path_image \ \ y \ - u \ d0 \ dist x y" - using separate_compact_closed [of "path_image \" "-u"] pasz u - by (fastforce simp add: \path \\ compact_path_image) - obtain dd where "0 < dd" and dd: "{y + k | y k. y \ path_image \ \ k \ ball 0 dd} \ u" - apply (rule that [of "d0/2"]) - using \0 < d0\ - apply (auto simp: dist_norm dest: d0) - done - have "\x x'. \x \ path_image \; dist x x' * 2 < dd\ \ \y k. x' = y + k \ y \ path_image \ \ dist 0 k * 2 \ dd" - apply (rule_tac x=x in exI) - apply (rule_tac x="x'-x" in exI) - apply (force simp add: dist_norm) - done - then have 1: "path_image \ \ interior {y + k |y k. y \ path_image \ \ k \ cball 0 (dd / 2)}" - apply (clarsimp simp add: mem_interior) - using \0 < dd\ - apply (rule_tac x="dd/2" in exI, auto) - done - obtain t where "compact t" and subt: "path_image \ \ interior t" and t: "t \ u" - apply (rule that [OF _ 1]) - apply (fastforce simp add: \valid_path \\ compact_valid_path_image intro!: compact_sums) - apply (rule order_trans [OF _ dd]) - using \0 < dd\ by fastforce - obtain L where "L>0" - and L: "\f B. \f holomorphic_on interior t; \z. z\interior t \ cmod (f z) \ B\ \ - cmod (contour_integral \ f) \ L * B" - using contour_integral_bound_exists [OF open_interior \valid_path \\ subt] - by blast - have "bounded(f ` t)" - by (meson \compact t\ compact_continuous_image compact_imp_bounded conf continuous_on_subset t) - then obtain D where "D>0" and D: "\x. x \ t \ norm (f x) \ D" - by (auto simp: bounded_pos) - obtain C where "C>0" and C: "\x. x \ t \ norm x \ C" - using \compact t\ bounded_pos compact_imp_bounded by force - have "dist (h y) 0 \ e" if "0 < e" and le: "D * L / e + C \ cmod y" for e y - proof - - have "D * L / e > 0" using \D>0\ \L>0\ \e>0\ by simp - with le have ybig: "norm y > C" by force - with C have "y \ t" by force - then have ynot: "y \ path_image \" - using subt interior_subset by blast - have [simp]: "winding_number \ y = 0" - apply (rule winding_number_zero_outside [of _ "cball 0 C"]) - using ybig interior_subset subt - apply (force simp add: loop \path \\ dist_norm intro!: C)+ - done - have [simp]: "h y = contour_integral \ (\w. f w/(w - y))" - by (rule contour_integral_unique [symmetric]) (simp add: v_def ynot V) - have holint: "(\w. f w / (w - y)) holomorphic_on interior t" - apply (rule holomorphic_on_divide) - using holf holomorphic_on_subset interior_subset t apply blast - apply (rule holomorphic_intros)+ - using \y \ t\ interior_subset by auto - have leD: "cmod (f z / (z - y)) \ D * (e / L / D)" if z: "z \ interior t" for z - proof - - have "D * L / e + cmod z \ cmod y" - using le C [of z] z using interior_subset by force - then have DL2: "D * L / e \ cmod (z - y)" - using norm_triangle_ineq2 [of y z] by (simp add: norm_minus_commute) - have "cmod (f z / (z - y)) = cmod (f z) * inverse (cmod (z - y))" - by (simp add: norm_mult norm_inverse Fields.field_class.field_divide_inverse) - also have "... \ D * (e / L / D)" - apply (rule mult_mono) - using that D interior_subset apply blast - using \L>0\ \e>0\ \D>0\ DL2 - apply (auto simp: norm_divide divide_simps algebra_simps) - done - finally show ?thesis . - qed - have "dist (h y) 0 = cmod (contour_integral \ (\w. f w / (w - y)))" - by (simp add: dist_norm) - also have "... \ L * (D * (e / L / D))" - by (rule L [OF holint leD]) - also have "... = e" - using \L>0\ \0 < D\ by auto - finally show ?thesis . - qed - then have "(h \ 0) at_infinity" - by (meson Lim_at_infinityI) - moreover have "h holomorphic_on UNIV" - proof - - have con_ff: "continuous (at (x,z)) (\(x,y). (f y - f x) / (y - x))" - if "x \ u" "z \ u" "x \ z" for x z - using that conf - apply (simp add: split_def continuous_on_eq_continuous_at u) - apply (simp | rule continuous_intros continuous_within_compose2 [where g=f])+ - done - have con_fstsnd: "continuous_on UNIV (\x. (fst x - snd x) ::complex)" - by (rule continuous_intros)+ - have open_uu_Id: "open (u \ u - Id)" - apply (rule open_Diff) - apply (simp add: open_Times u) - using continuous_closed_preimage_constant [OF con_fstsnd closed_UNIV, of 0] - apply (auto simp: Id_fstsnd_eq algebra_simps) - done - have con_derf: "continuous (at z) (deriv f)" if "z \ u" for z - apply (rule continuous_on_interior [of u]) - apply (simp add: holf holomorphic_deriv holomorphic_on_imp_continuous_on u) - by (simp add: interior_open that u) - have tendsto_f': "((\(x,y). if y = x then deriv f (x) - else (f (y) - f (x)) / (y - x)) \ deriv f x) - (at (x, x) within u \ u)" if "x \ u" for x - proof (rule Lim_withinI) - fix e::real assume "0 < e" - obtain k1 where "k1>0" and k1: "\x'. norm (x' - x) \ k1 \ norm (deriv f x' - deriv f x) < e" - using \0 < e\ continuous_within_E [OF con_derf [OF \x \ u\]] - by (metis UNIV_I dist_norm) - obtain k2 where "k2>0" and k2: "ball x k2 \ u" by (blast intro: openE [OF u] \x \ u\) - have neq: "norm ((f z' - f x') / (z' - x') - deriv f x) \ e" - if "z' \ x'" and less_k1: "norm (x'-x, z'-x) < k1" and less_k2: "norm (x'-x, z'-x) < k2" - for x' z' - proof - - have cs_less: "w \ closed_segment x' z' \ cmod (w - x) \ norm (x'-x, z'-x)" for w - apply (drule segment_furthest_le [where y=x]) - by (metis (no_types) dist_commute dist_norm norm_fst_le norm_snd_le order_trans) - have derf_le: "w \ closed_segment x' z' \ z' \ x' \ cmod (deriv f w - deriv f x) \ e" for w - by (blast intro: cs_less less_k1 k1 [unfolded divide_const_simps dist_norm] less_imp_le le_less_trans) - have f_has_der: "\x. x \ u \ (f has_field_derivative deriv f x) (at x within u)" - by (metis DERIV_deriv_iff_field_differentiable at_within_open holf holomorphic_on_def u) - have "closed_segment x' z' \ u" - by (rule order_trans [OF _ k2]) (simp add: cs_less le_less_trans [OF _ less_k2] dist_complex_def norm_minus_commute subset_iff) - then have cint_derf: "(deriv f has_contour_integral f z' - f x') (linepath x' z')" - using contour_integral_primitive [OF f_has_der valid_path_linepath] pasz by simp - then have *: "((\x. deriv f x / (z' - x')) has_contour_integral (f z' - f x') / (z' - x')) (linepath x' z')" - by (rule has_contour_integral_div) - have "norm ((f z' - f x') / (z' - x') - deriv f x) \ e/norm(z' - x') * norm(z' - x')" - apply (rule has_contour_integral_bound_linepath [OF has_contour_integral_diff [OF *]]) - using has_contour_integral_div [where c = "z' - x'", OF has_contour_integral_const_linepath [of "deriv f x" z' x']] - \e > 0\ \z' \ x'\ - apply (auto simp: norm_divide divide_simps derf_le) - done - also have "... \ e" using \0 < e\ by simp - finally show ?thesis . - qed - show "\d>0. \xa\u \ u. - 0 < dist xa (x, x) \ dist xa (x, x) < d \ - dist (case xa of (x, y) \ if y = x then deriv f x else (f y - f x) / (y - x)) (deriv f x) \ e" - apply (rule_tac x="min k1 k2" in exI) - using \k1>0\ \k2>0\ \e>0\ - apply (force simp: dist_norm neq intro: dual_order.strict_trans2 k1 less_imp_le norm_fst_le) - done - qed - have con_pa_f: "continuous_on (path_image \) f" - by (meson holf holomorphic_on_imp_continuous_on holomorphic_on_subset interior_subset subt t) - have le_B: "\t. t \ {0..1} \ cmod (vector_derivative \ (at t)) \ B" - apply (rule B) - using \' using path_image_def vector_derivative_at by fastforce - have f_has_cint: "\w. w \ v - path_image \ \ ((\u. f u / (u - w) ^ 1) has_contour_integral h w) \" - by (simp add: V) - have cond_uu: "continuous_on (u \ u) (\(x,y). d x y)" - apply (simp add: continuous_on_eq_continuous_within d_def continuous_within tendsto_f') - apply (simp add: tendsto_within_open_NO_MATCH open_Times u, clarify) - apply (rule Lim_transform_within_open [OF _ open_uu_Id, where f = "(\(x,y). (f y - f x) / (y - x))"]) - using con_ff - apply (auto simp: continuous_within) - done - have hol_dw: "(\z. d z w) holomorphic_on u" if "w \ u" for w - proof - - have "continuous_on u ((\(x,y). d x y) o (\z. (w,z)))" - by (rule continuous_on_compose continuous_intros continuous_on_subset [OF cond_uu] | force intro: that)+ - then have *: "continuous_on u (\z. if w = z then deriv f z else (f w - f z) / (w - z))" - by (rule rev_iffD1 [OF _ continuous_on_cong [OF refl]]) (simp add: d_def field_simps) - have **: "\x. \x \ u; x \ w\ \ (\z. if w = z then deriv f z else (f w - f z) / (w - z)) field_differentiable at x" - apply (rule_tac f = "\x. (f w - f x)/(w - x)" and d = "dist x w" in field_differentiable_transform_within) - apply (rule u derivative_intros holomorphic_on_imp_differentiable_at [OF holf] | force simp add: dist_commute)+ - done - show ?thesis - unfolding d_def - apply (rule no_isolated_singularity [OF * _ u, where k = "{w}"]) - apply (auto simp: field_differentiable_def [symmetric] holomorphic_on_open open_Diff u **) - done - qed - { fix a b - assume abu: "closed_segment a b \ u" - then have "\w. w \ u \ (\z. d z w) contour_integrable_on (linepath a b)" - by (metis hol_dw continuous_on_subset contour_integrable_continuous_linepath holomorphic_on_imp_continuous_on) - then have cont_cint_d: "continuous_on u (\w. contour_integral (linepath a b) (\z. d z w))" - apply (rule contour_integral_continuous_on_linepath_2D [OF \open u\ _ _ abu]) - apply (auto simp: intro: continuous_on_swap_args cond_uu) - done - have cont_cint_d\: "continuous_on {0..1} ((\w. contour_integral (linepath a b) (\z. d z w)) o \)" - apply (rule continuous_on_compose) - using \path \\ path_def pasz - apply (auto intro!: continuous_on_subset [OF cont_cint_d]) - apply (force simp add: path_image_def) - done - have cint_cint: "(\w. contour_integral (linepath a b) (\z. d z w)) contour_integrable_on \" - apply (simp add: contour_integrable_on) - apply (rule integrable_continuous_real) - apply (rule continuous_on_mult [OF cont_cint_d\ [unfolded o_def]]) - using pf\' - by (simp add: continuous_on_polymonial_function vector_derivative_at [OF \']) - have "contour_integral (linepath a b) h = contour_integral (linepath a b) (\z. contour_integral \ (d z))" - using abu by (force simp add: h_def intro: contour_integral_eq) - also have "... = contour_integral \ (\w. contour_integral (linepath a b) (\z. d z w))" - apply (rule contour_integral_swap) - apply (rule continuous_on_subset [OF cond_uu]) - using abu pasz \valid_path \\ - apply (auto intro!: continuous_intros) - by (metis \' continuous_on_eq path_def path_polynomial_function pf\' vector_derivative_at) - finally have cint_h_eq: - "contour_integral (linepath a b) h = - contour_integral \ (\w. contour_integral (linepath a b) (\z. d z w))" . - note cint_cint cint_h_eq - } note cint_h = this - have conthu: "continuous_on u h" - proof (simp add: continuous_on_sequentially, clarify) - fix a x - assume x: "x \ u" and au: "\n. a n \ u" and ax: "a \ x" - then have A1: "\\<^sub>F n in sequentially. d (a n) contour_integrable_on \" - by (meson U contour_integrable_on_def eventuallyI) - obtain dd where "dd>0" and dd: "cball x dd \ u" using open_contains_cball u x by force - have A2: "\\<^sub>F n in sequentially. \xa\path_image \. cmod (d (a n) xa - d x xa) < ee" if "0 < ee" for ee - proof - - let ?ddpa = "{(w,z) |w z. w \ cball x dd \ z \ path_image \}" - have "uniformly_continuous_on ?ddpa (\(x,y). d x y)" - apply (rule compact_uniformly_continuous [OF continuous_on_subset[OF cond_uu]]) - using dd pasz \valid_path \\ - apply (auto simp: compact_Times compact_valid_path_image simp del: mem_cball) - done - then obtain kk where "kk>0" - and kk: "\x x'. \x\?ddpa; x'\?ddpa; dist x' x < kk\ \ - dist ((\(x,y). d x y) x') ((\(x,y). d x y) x) < ee" - apply (rule uniformly_continuous_onE [where e = ee]) - using \0 < ee\ by auto - have kk: "\norm (w - x) \ dd; z \ path_image \; norm ((w, z) - (x, z)) < kk\ \ norm (d w z - d x z) < ee" - for w z - using \dd>0\ kk [of "(x,z)" "(w,z)"] by (force simp add: norm_minus_commute dist_norm) - show ?thesis - using ax unfolding lim_sequentially eventually_sequentially - apply (drule_tac x="min dd kk" in spec) - using \dd > 0\ \kk > 0\ - apply (fastforce simp: kk dist_norm) - done - qed - have tendsto_hx: "(\n. contour_integral \ (d (a n))) \ h x" - apply (simp add: contour_integral_unique [OF U, symmetric] x) - apply (rule contour_integral_uniform_limit [OF A1 A2 le_B]) - apply (auto simp: \valid_path \\) - done - then show "(h \ a) \ h x" - by (simp add: h_def x au o_def) - qed - show ?thesis - proof (simp add: holomorphic_on_open field_differentiable_def [symmetric], clarify) - fix z0 - consider "z0 \ v" | "z0 \ u" using uv_Un by blast - then show "h field_differentiable at z0" - proof cases - assume "z0 \ v" then show ?thesis - using Cauchy_next_derivative [OF con_pa_f le_B f_has_cint _ ov] V f_has_cint \valid_path \\ - by (auto simp: field_differentiable_def v_def) - next - assume "z0 \ u" then - obtain e where "e>0" and e: "ball z0 e \ u" by (blast intro: openE [OF u]) - have *: "contour_integral (linepath a b) h + contour_integral (linepath b c) h + contour_integral (linepath c a) h = 0" - if abc_subset: "convex hull {a, b, c} \ ball z0 e" for a b c - proof - - have *: "\x1 x2 z. z \ u \ closed_segment x1 x2 \ u \ (\w. d w z) contour_integrable_on linepath x1 x2" - using hol_dw holomorphic_on_imp_continuous_on u - by (auto intro!: contour_integrable_holomorphic_simple) - have abc: "closed_segment a b \ u" "closed_segment b c \ u" "closed_segment c a \ u" - using that e segments_subset_convex_hull by fastforce+ - have eq0: "\w. w \ u \ contour_integral (linepath a b +++ linepath b c +++ linepath c a) (\z. d z w) = 0" - apply (rule contour_integral_unique [OF Cauchy_theorem_triangle]) - apply (rule holomorphic_on_subset [OF hol_dw]) - using e abc_subset by auto - have "contour_integral \ - (\x. contour_integral (linepath a b) (\z. d z x) + - (contour_integral (linepath b c) (\z. d z x) + - contour_integral (linepath c a) (\z. d z x))) = 0" - apply (rule contour_integral_eq_0) - using abc pasz u - apply (subst contour_integral_join [symmetric], auto intro: eq0 *)+ - done - then show ?thesis - by (simp add: cint_h abc contour_integrable_add contour_integral_add [symmetric] add_ac) - qed - show ?thesis - using e \e > 0\ - by (auto intro!: holomorphic_on_imp_differentiable_at [OF _ open_ball] analytic_imp_holomorphic - Morera_triangle continuous_on_subset [OF conthu] *) - qed - qed - qed - ultimately have [simp]: "h z = 0" for z - by (meson Liouville_weak) - have "((\w. 1 / (w - z)) has_contour_integral complex_of_real (2 * pi) * \ * winding_number \ z) \" - by (rule has_contour_integral_winding_number [OF \valid_path \\ znot]) - then have "((\w. f z * (1 / (w - z))) has_contour_integral complex_of_real (2 * pi) * \ * winding_number \ z * f z) \" - by (metis mult.commute has_contour_integral_lmul) - then have 1: "((\w. f z / (w - z)) has_contour_integral complex_of_real (2 * pi) * \ * winding_number \ z * f z) \" - by (simp add: divide_simps) - moreover have 2: "((\w. (f w - f z) / (w - z)) has_contour_integral 0) \" - using U [OF z] pasz d_def by (force elim: has_contour_integral_eq [where g = "\w. (f w - f z)/(w - z)"]) - show ?thesis - using has_contour_integral_add [OF 1 2] by (simp add: diff_divide_distrib) -qed - - -theorem Cauchy_integral_formula_global: - assumes s: "open s" and holf: "f holomorphic_on s" - and z: "z \ s" and vpg: "valid_path \" - and pasz: "path_image \ \ s - {z}" and loop: "pathfinish \ = pathstart \" - and zero: "\w. w \ s \ winding_number \ w = 0" - shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" -proof - - have "path \" using vpg by (blast intro: valid_path_imp_path) - have hols: "(\w. f w / (w - z)) holomorphic_on s - {z}" "(\w. 1 / (w - z)) holomorphic_on s - {z}" - by (rule holomorphic_intros holomorphic_on_subset [OF holf] | force)+ - then have cint_fw: "(\w. f w / (w - z)) contour_integrable_on \" - by (meson contour_integrable_holomorphic_simple holomorphic_on_imp_continuous_on open_delete s vpg pasz) - obtain d where "d>0" - and d: "\g h. \valid_path g; valid_path h; \t\{0..1}. cmod (g t - \ t) < d \ cmod (h t - \ t) < d; - pathstart h = pathstart g \ pathfinish h = pathfinish g\ - \ path_image h \ s - {z} \ (\f. f holomorphic_on s - {z} \ contour_integral h f = contour_integral g f)" - using contour_integral_nearby_ends [OF _ \path \\ pasz] s by (simp add: open_Diff) metis - obtain p where polyp: "polynomial_function p" - and ps: "pathstart p = pathstart \" and pf: "pathfinish p = pathfinish \" and led: "\t\{0..1}. cmod (p t - \ t) < d" - using path_approx_polynomial_function [OF \path \\ \d > 0\] by blast - then have ploop: "pathfinish p = pathstart p" using loop by auto - have vpp: "valid_path p" using polyp valid_path_polynomial_function by blast - have [simp]: "z \ path_image \" using pasz by blast - have paps: "path_image p \ s - {z}" and cint_eq: "(\f. f holomorphic_on s - {z} \ contour_integral p f = contour_integral \ f)" - using pf ps led d [OF vpg vpp] \d > 0\ by auto - have wn_eq: "winding_number p z = winding_number \ z" - using vpp paps - by (simp add: subset_Diff_insert vpg valid_path_polynomial_function winding_number_valid_path cint_eq hols) - have "winding_number p w = winding_number \ w" if "w \ s" for w - proof - - have hol: "(\v. 1 / (v - w)) holomorphic_on s - {z}" - using that by (force intro: holomorphic_intros holomorphic_on_subset [OF holf]) - have "w \ path_image p" "w \ path_image \" using paps pasz that by auto - then show ?thesis - using vpp vpg by (simp add: subset_Diff_insert valid_path_polynomial_function winding_number_valid_path cint_eq [OF hol]) - qed - then have wn0: "\w. w \ s \ winding_number p w = 0" - by (simp add: zero) - show ?thesis - using Cauchy_integral_formula_global_weak [OF s holf z polyp paps ploop wn0] hols - by (metis wn_eq cint_eq has_contour_integral_eqpath cint_fw cint_eq) -qed - -theorem Cauchy_theorem_global: - assumes s: "open s" and holf: "f holomorphic_on s" - and vpg: "valid_path \" and loop: "pathfinish \ = pathstart \" - and pas: "path_image \ \ s" - and zero: "\w. w \ s \ winding_number \ w = 0" - shows "(f has_contour_integral 0) \" -proof - - obtain z where "z \ s" and znot: "z \ path_image \" - proof - - have "compact (path_image \)" - using compact_valid_path_image vpg by blast - then have "path_image \ \ s" - by (metis (no_types) compact_open path_image_nonempty s) - with pas show ?thesis by (blast intro: that) - qed - then have pasz: "path_image \ \ s - {z}" using pas by blast - have hol: "(\w. (w - z) * f w) holomorphic_on s" - by (rule holomorphic_intros holf)+ - show ?thesis - using Cauchy_integral_formula_global [OF s hol \z \ s\ vpg pasz loop zero] - by (auto simp: znot elim!: has_contour_integral_eq) -qed - -corollary Cauchy_theorem_global_outside: - assumes "open s" "f holomorphic_on s" "valid_path \" "pathfinish \ = pathstart \" "path_image \ \ s" - "\w. w \ s \ w \ outside(path_image \)" - shows "(f has_contour_integral 0) \" -by (metis Cauchy_theorem_global assms winding_number_zero_in_outside valid_path_imp_path) - - -end diff -r bbcb05504fdc -r bd218a9320b5 src/HOL/Multivariate_Analysis/Complex_Transcendental.thy --- a/src/HOL/Multivariate_Analysis/Complex_Transcendental.thy Thu Aug 04 18:45:28 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Complex_Transcendental.thy Thu Aug 04 19:36:31 2016 +0200 @@ -5,7 +5,7 @@ theory Complex_Transcendental imports Complex_Analysis_Basics - Summation + Summation_Tests begin (* TODO: Figure out what to do with Möbius transformations *) @@ -1647,7 +1647,7 @@ lemma norm_powr_real_powr: "w \ \ \ 0 \ Re w \ cmod (w powr z) = Re w powr Re z" - by (cases "w = 0") (auto simp add: norm_powr_real powr_def Im_Ln_eq_0 + by (cases "w = 0") (auto simp add: norm_powr_real powr_def Im_Ln_eq_0 complex_is_Real_iff in_Reals_norm complex_eq_iff) lemma tendsto_ln_complex [tendsto_intros]: diff -r bbcb05504fdc -r bd218a9320b5 src/HOL/Multivariate_Analysis/Conformal_Mappings.thy --- a/src/HOL/Multivariate_Analysis/Conformal_Mappings.thy Thu Aug 04 18:45:28 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Conformal_Mappings.thy Thu Aug 04 19:36:31 2016 +0200 @@ -5,7 +5,7 @@ text\Also Cauchy's residue theorem by Wenda Li (2016)\ theory Conformal_Mappings -imports "~~/src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm" +imports "~~/src/HOL/Multivariate_Analysis/Cauchy_Integral_Theorem" begin @@ -1897,7 +1897,7 @@ qed qed -lemma Bloch_lemma: +lemma Bloch_lemma: assumes holf: "f holomorphic_on cball a r" and "0 < r" and le: "\z. z \ ball a r \ norm(deriv f z) \ 2 * norm(deriv f a)" shows "ball (f a) ((3 - 2 * sqrt 2) * r * norm(deriv f a)) \ f ` ball a r" @@ -1928,7 +1928,7 @@ done qed -proposition Bloch_unit: +proposition Bloch_unit: assumes holf: "f holomorphic_on ball a 1" and [simp]: "deriv f a = 1" obtains b r where "1/12 < r" "ball b r \ f ` (ball a 1)" proof - @@ -1947,23 +1947,23 @@ by (rule compact_continuous_image [OF _ compact_cball]) have 2: "g ` cball a r \ {}" using \r > 0\ by auto - obtain p where pr: "p \ cball a r" + obtain p where pr: "p \ cball a r" and pge: "\y. y \ cball a r \ norm (g y) \ norm (g p)" using distance_attains_sup [OF 1 2, of 0] by force define t where "t = (r - norm(p - a)) / 2" have "norm (p - a) \ r" using pge [of a] \r > 0\ by (auto simp: g_def norm_mult) - then have "norm (p - a) < r" using pr + then have "norm (p - a) < r" using pr by (simp add: norm_minus_commute dist_norm) - then have "0 < t" + then have "0 < t" by (simp add: t_def) have cpt: "cball p t \ ball a r" using \0 < t\ by (simp add: cball_subset_ball_iff dist_norm t_def field_simps) - have gen_le_dfp: "norm (deriv f y) * (r - norm (y - a)) / (r - norm (p - a)) \ norm (deriv f p)" + have gen_le_dfp: "norm (deriv f y) * (r - norm (y - a)) / (r - norm (p - a)) \ norm (deriv f p)" if "y \ cball a r" for y proof - have [simp]: "norm (y - a) \ r" - using that by (simp add: dist_norm norm_minus_commute) + using that by (simp add: dist_norm norm_minus_commute) have "norm (g y) \ norm (g p)" using pge [OF that] by simp then have "norm (deriv f y) * abs (r - norm (y - a)) \ norm (deriv f p) * abs (r - norm (p - a))" @@ -1982,14 +1982,14 @@ by (meson ball_subset_cball subsetD cpt that) then have "norm(z - a) < r" by (metis ball_subset_cball contra_subsetD cpt dist_norm mem_ball norm_minus_commute that) - have "norm (deriv f z) * (r - norm (z - a)) / (r - norm (p - a)) \ norm (deriv f p)" + have "norm (deriv f z) * (r - norm (z - a)) / (r - norm (p - a)) \ norm (deriv f p)" using gen_le_dfp [OF z] by simp - with \norm (z - a) < r\ \norm (p - a) < r\ + with \norm (z - a) < r\ \norm (p - a) < r\ have "norm (deriv f z) \ (r - norm (p - a)) / (r - norm (z - a)) * norm (deriv f p)" by (simp add: field_simps) also have "... \ 2 * norm (deriv f p)" apply (rule mult_right_mono) - using that \norm (p - a) < r\ \norm(z - a) < r\ + using that \norm (p - a) < r\ \norm(z - a) < r\ apply (simp_all add: field_simps t_def dist_norm [symmetric]) using dist_triangle3 [of z a p] by linarith finally show ?thesis . @@ -1998,10 +1998,10 @@ by (rule real_less_lsqrt) (auto simp: power2_eq_square) then have sq3: "0 < 3 - 2 * sqrt 2" by simp have "1 / 12 / ((3 - 2 * sqrt 2) / 2) < r" - using sq3 sqrt2 by (auto simp: field_simps r_def) + using sq3 sqrt2 by (auto simp: field_simps r_def) also have "... \ cmod (deriv f p) * (r - cmod (p - a))" - using \norm (p - a) < r\ le_norm_dfp by (simp add: pos_divide_le_eq) - finally have "1 / 12 < cmod (deriv f p) * (r - cmod (p - a)) * ((3 - 2 * sqrt 2) / 2)" + using \norm (p - a) < r\ le_norm_dfp by (simp add: pos_divide_le_eq) + finally have "1 / 12 < cmod (deriv f p) * (r - cmod (p - a)) * ((3 - 2 * sqrt 2) / 2)" using pos_divide_less_eq half_gt_zero_iff sq3 by blast then have **: "1 / 12 < (3 - 2 * sqrt 2) * t * norm (deriv f p)" using sq3 by (simp add: mult.commute t_def) @@ -2020,7 +2020,7 @@ theorem Bloch: - assumes holf: "f holomorphic_on ball a r" and "0 < r" + assumes holf: "f holomorphic_on ball a r" and "0 < r" and r': "r' \ r * norm (deriv f a) / 12" obtains b where "ball b r' \ f ` (ball a r)" proof (cases "deriv f a = 0") @@ -2043,7 +2043,7 @@ apply (rule holomorphic_intros holomorphic_on_compose holf' | simp add: fo)+ using \0 < r\ by (simp add: C_def False) have "((\z. f (a + of_real r * z) / (C * of_real r)) has_field_derivative - (deriv f (a + of_real r * z) / C)) (at z)" + (deriv f (a + of_real r * z) / C)) (at z)" if "norm z < 1" for z proof - have *: "((\x. f (a + of_real r * x)) has_field_derivative @@ -2065,23 +2065,23 @@ apply (rule DERIV_chain [OF field_differentiable_derivI]) apply (simp add: dfa) apply (rule derivative_eq_intros | simp add: C_def False fo)+ - using \0 < r\ + using \0 < r\ apply (simp add: C_def False fo) apply (simp add: derivative_intros dfa complex_derivative_chain) done - have sb1: "op * (C * r) ` (\z. f (a + of_real r * z) / (C * r)) ` ball 0 1 + have sb1: "op * (C * r) ` (\z. f (a + of_real r * z) / (C * r)) ` ball 0 1 \ f ` ball a r" using \0 < r\ by (auto simp: dist_norm norm_mult C_def False) - have sb2: "ball (C * r * b) r' \ op * (C * r) ` ball b t" + have sb2: "ball (C * r * b) r' \ op * (C * r) ` ball b t" if "1 / 12 < t" for b t proof - have *: "r * cmod (deriv f a) / 12 \ r * (t * cmod (deriv f a))" - using that \0 < r\ less_eq_real_def mult.commute mult.right_neutral mult_left_mono norm_ge_zero times_divide_eq_right + using that \0 < r\ less_eq_real_def mult.commute mult.right_neutral mult_left_mono norm_ge_zero times_divide_eq_right by auto show ?thesis apply clarify apply (rule_tac x="x / (C * r)" in image_eqI) - using \0 < r\ + using \0 < r\ apply (simp_all add: dist_norm norm_mult norm_divide C_def False field_simps) apply (erule less_le_trans) apply (rule order_trans [OF r' *]) @@ -2098,7 +2098,7 @@ qed corollary Bloch_general: - assumes holf: "f holomorphic_on s" and "a \ s" + assumes holf: "f holomorphic_on s" and "a \ s" and tle: "\z. z \ frontier s \ t \ dist a z" and rle: "r \ t * norm(deriv f a) / 12" obtains b where "ball b r \ f ` s" @@ -2140,11 +2140,11 @@ Interactive Theorem Proving\ definition residue :: "(complex \ complex) \ complex \ complex" where - "residue f z = (SOME int. \e>0. \\>0. \e>0. \\>0. \ (f has_contour_integral 2*pi* \ *int) (circlepath z \))" lemma contour_integral_circlepath_eq: - assumes "open s" and f_holo:"f holomorphic_on (s-{z})" and "0e2" + assumes "open s" and f_holo:"f holomorphic_on (s-{z})" and "0e2" and e2_cball:"cball z e2 \ s" shows "f contour_integrable_on circlepath z e1" @@ -2155,24 +2155,24 @@ have [simp]:"valid_path l" "pathstart l=z+e2" "pathfinish l=z+e1" unfolding l_def by auto have "e2>0" using \e1>0\ \e1\e2\ by auto have zl_img:"z\path_image l" - proof + proof assume "z \ path_image l" then have "e2 \ cmod (e2 - e1)" using segment_furthest_le[of z "z+e2" "z+e1" "z+e2",simplified] \e1>0\ \e2>0\ unfolding l_def by (auto simp add:closed_segment_commute) - thus False using \e2>0\ \e1>0\ \e1\e2\ + thus False using \e2>0\ \e1>0\ \e1\e2\ apply (subst (asm) norm_of_real) by auto qed define g where "g \ circlepath z e2 +++ l +++ reversepath (circlepath z e1) +++ reversepath l" - show [simp]: "f contour_integrable_on circlepath z e2" "f contour_integrable_on (circlepath z e1)" + show [simp]: "f contour_integrable_on circlepath z e2" "f contour_integrable_on (circlepath z e1)" proof - - show "f contour_integrable_on circlepath z e2" - apply (intro contour_integrable_continuous_circlepath[OF + show "f contour_integrable_on circlepath z e2" + apply (intro contour_integrable_continuous_circlepath[OF continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) using \e2>0\ e2_cball by auto - show "f contour_integrable_on (circlepath z e1)" - apply (intro contour_integrable_continuous_circlepath[OF + show "f contour_integrable_on (circlepath z e1)" + apply (intro contour_integrable_continuous_circlepath[OF continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) using \e1>0\ \e1\e2\ e2_cball by auto qed @@ -2183,7 +2183,7 @@ hence "closed_segment (z + e2) (z + e1) \ s - {z}" using zl_img e2_cball unfolding l_def by auto then show "f contour_integrable_on l" unfolding l_def - apply (intro contour_integrable_continuous_linepath[OF + apply (intro contour_integrable_continuous_linepath[OF continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) by auto qed @@ -2197,14 +2197,14 @@ have path_img:"path_image g \ cball z e2" proof - have "closed_segment (z + e2) (z + e1) \ cball z e2" using \e2>0\ \e1>0\ \e1\e2\ - by (intro closed_segment_subset,auto simp add:dist_norm) + by (intro closed_segment_subset,auto simp add:dist_norm) moreover have "sphere z \e1\ \ cball z e2" using \e2>0\ \e1\e2\ \e1>0\ by auto - ultimately show ?thesis unfolding g_def l_def using \e2>0\ + ultimately show ?thesis unfolding g_def l_def using \e2>0\ by (simp add: path_image_join closed_segment_commute) qed - show "path_image g \ s - {z}" + show "path_image g \ s - {z}" proof - - have "z\path_image g" using zl_img + have "z\path_image g" using zl_img unfolding g_def l_def by (auto simp add: path_image_join closed_segment_commute) moreover note \cball z e2 \ s\ and path_img ultimately show ?thesis by auto @@ -2214,12 +2214,12 @@ have "winding_number g w = 0" when "w\s" using that e2_cball apply (intro winding_number_zero_outside[OF _ _ _ _ path_img]) by (auto simp add:g_def l_def) - moreover have "winding_number g z=0" + moreover have "winding_number g z=0" proof - let ?Wz="\g. winding_number g z" - have "?Wz g = ?Wz (circlepath z e2) + ?Wz l + ?Wz (reversepath (circlepath z e1)) + have "?Wz g = ?Wz (circlepath z e2) + ?Wz l + ?Wz (reversepath (circlepath z e1)) + ?Wz (reversepath l)" - using \e2>0\ \e1>0\ zl_img unfolding g_def l_def + using \e2>0\ \e1>0\ zl_img unfolding g_def l_def by (subst winding_number_join,auto simp add:path_image_join closed_segment_commute)+ also have "... = ?Wz (circlepath z e2) + ?Wz (reversepath (circlepath z e1))" using zl_img @@ -2237,43 +2237,43 @@ finally show ?thesis . qed ultimately show ?thesis using that by auto - qed - qed + qed + qed then have "0 = ?ig g" using contour_integral_unique by simp - also have "... = ?ig (circlepath z e2) + ?ig l + ?ig (reversepath (circlepath z e1)) + also have "... = ?ig (circlepath z e2) + ?ig l + ?ig (reversepath (circlepath z e1)) + ?ig (reversepath l)" - unfolding g_def + unfolding g_def by (auto simp add:contour_integrable_reversepath_eq) also have "... = ?ig (circlepath z e2) - ?ig (circlepath z e1)" by (auto simp add:contour_integral_reversepath) - finally show "contour_integral (circlepath z e2) f = contour_integral (circlepath z e1) f" + finally show "contour_integral (circlepath z e2) f = contour_integral (circlepath z e1) f" by simp qed lemma base_residue: - assumes "open s" "z\s" "r>0" and f_holo:"f holomorphic_on (s - {z})" + assumes "open s" "z\s" "r>0" and f_holo:"f holomorphic_on (s - {z})" and r_cball:"cball z r \ s" shows "(f has_contour_integral 2 * pi * \ * (residue f z)) (circlepath z r)" proof - - obtain e where "e>0" and e_cball:"cball z e \ s" + obtain e where "e>0" and e_cball:"cball z e \ s" using open_contains_cball[of s] \open s\ \z\s\ by auto define c where "c \ 2 * pi * \" define i where "i \ contour_integral (circlepath z e) f / c" - have "(f has_contour_integral c*i) (circlepath z \)" when "\>0" "\ + have "(f has_contour_integral c*i) (circlepath z \)" when "\>0" "\ proof - have "contour_integral (circlepath z e) f = contour_integral (circlepath z \) f" "f contour_integrable_on circlepath z \" "f contour_integrable_on circlepath z e" using \\ - by (intro contour_integral_circlepath_eq[OF \open s\ f_holo \\>0\ _ e_cball],auto)+ + by (intro contour_integral_circlepath_eq[OF \open s\ f_holo \\>0\ _ e_cball],auto)+ then show ?thesis unfolding i_def c_def by (auto intro:has_contour_integral_integral) qed - then have "\e>0. \\>0. \ (f has_contour_integral c * (residue f z)) (circlepath z \)" + then have "\e>0. \\>0. \ (f has_contour_integral c * (residue f z)) (circlepath z \)" unfolding residue_def c_def apply (rule_tac someI[of _ i],intro exI[where x=e]) by (auto simp add:\e>0\ c_def) - then obtain e' where "e'>0" + then obtain e' where "e'>0" and e'_def:"\\>0. \ (f has_contour_integral c * (residue f z)) (circlepath z \)" by auto let ?int="\e. contour_integral (circlepath z e) f" @@ -2281,8 +2281,8 @@ have "\>0" "\\r" "\r>0\ \e'>0\ unfolding \_def by auto have "(f has_contour_integral c * (residue f z)) (circlepath z \)" using e'_def[rule_format,OF \\>0\ \\] . - then show ?thesis unfolding c_def - using contour_integral_circlepath_eq[OF \open s\ f_holo \\>0\ \\\r\ r_cball] + then show ?thesis unfolding c_def + using contour_integral_circlepath_eq[OF \open s\ f_holo \\>0\ \\\r\ r_cball] by (auto elim: has_contour_integral_eqpath[of _ _ "circlepath z \" "circlepath z r"]) qed @@ -2292,17 +2292,17 @@ shows "residue f z = 0" proof - define c where "c \ 2 * pi * \" - obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ + obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ using open_contains_cball_eq by blast have "(f has_contour_integral c*residue f z) (circlepath z e)" using f_holo - by (auto intro: base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) - moreover have "(f has_contour_integral 0) (circlepath z e)" + by (auto intro: base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) + moreover have "(f has_contour_integral 0) (circlepath z e)" using f_holo e_cball \e>0\ by (auto intro: Cauchy_theorem_convex_simple[of _ "cball z e"]) ultimately have "c*residue f z =0" using has_contour_integral_unique by blast - thus ?thesis unfolding c_def by auto + thus ?thesis unfolding c_def by auto qed @@ -2311,39 +2311,39 @@ lemma residue_add: - assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" + assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" and g_holo:"g holomorphic_on s - {z}" shows "residue (\z. f z + g z) z= residue f z + residue g z" proof - define c where "c \ 2 * pi * \" define fg where "fg \ (\z. f z+g z)" - obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ + obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ using open_contains_cball_eq by blast - have "(fg has_contour_integral c * residue fg z) (circlepath z e)" + have "(fg has_contour_integral c * residue fg z) (circlepath z e)" unfolding fg_def using f_holo g_holo apply (intro base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) by (auto intro:holomorphic_intros) moreover have "(fg has_contour_integral c*residue f z + c* residue g z) (circlepath z e)" unfolding fg_def using f_holo g_holo - by (auto intro: has_contour_integral_add base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) + by (auto intro: has_contour_integral_add base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) ultimately have "c*(residue f z + residue g z) = c * residue fg z" using has_contour_integral_unique by (auto simp add:distrib_left) - thus ?thesis unfolding fg_def + thus ?thesis unfolding fg_def by (auto simp add:c_def) qed - + lemma residue_lmul: - assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" + assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" shows "residue (\z. c * (f z)) z= c * residue f z" proof (cases "c=0") case True thus ?thesis using residue_const by auto -next +next case False def c'\"2 * pi * \" def f'\"(\z. c * (f z))" - obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ + obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ using open_contains_cball_eq by blast have "(f' has_contour_integral c' * residue f' z) (circlepath z e)" unfolding f'_def using f_holo @@ -2351,34 +2351,34 @@ by (auto intro:holomorphic_intros) moreover have "(f' has_contour_integral c * (c' * residue f z)) (circlepath z e)" unfolding f'_def using f_holo - by (auto intro: has_contour_integral_lmul - base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c'_def]) + by (auto intro: has_contour_integral_lmul + base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c'_def]) ultimately have "c' * residue f' z = c * (c' * residue f z)" - using has_contour_integral_unique by auto + using has_contour_integral_unique by auto thus ?thesis unfolding f'_def c'_def using False by (auto simp add:field_simps) qed lemma residue_rmul: - assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" + assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" shows "residue (\z. (f z) * c) z= residue f z * c" using residue_lmul[OF assms,of c] by (auto simp add:algebra_simps) lemma residue_div: - assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" + assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" shows "residue (\z. (f z) / c) z= residue f z / c " using residue_lmul[OF assms,of "1/c"] by (auto simp add:algebra_simps) lemma residue_neg: - assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" + assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" shows "residue (\z. - (f z)) z= - residue f z" using residue_lmul[OF assms,of "-1"] by auto lemma residue_diff: - assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" + assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" and g_holo:"g holomorphic_on s - {z}" shows "residue (\z. f z - g z) z= residue f z - residue g z" -using residue_add[OF assms(1,2,3),of "\z. - g z"] residue_neg[OF assms(1,2,4)] +using residue_add[OF assms(1,2,3),of "\z. - g z"] residue_neg[OF assms(1,2,4)] by (auto intro:holomorphic_intros g_holo) lemma residue_simple: @@ -2387,7 +2387,7 @@ proof - define c where "c \ 2 * pi * \" def f'\"\w. f w / (w - z)" - obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ + obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ using open_contains_cball_eq by blast have "(f' has_contour_integral c * f z) (circlepath z e)" unfolding f'_def c_def using \e>0\ f_holo e_cball @@ -2407,52 +2407,52 @@ lemma get_integrable_path: assumes "open s" "connected (s-pts)" "finite pts" "f holomorphic_on (s-pts) " "a\s-pts" "b\s-pts" - obtains g where "valid_path g" "pathstart g = a" "pathfinish g = b" + obtains g where "valid_path g" "pathstart g = a" "pathfinish g = b" "path_image g \ s-pts" "f contour_integrable_on g" using assms -proof (induct arbitrary:s thesis a rule:finite_induct[OF \finite pts\]) +proof (induct arbitrary:s thesis a rule:finite_induct[OF \finite pts\]) case 1 obtain g where "valid_path g" "path_image g \ s" "pathstart g = a" "pathfinish g = b" using connected_open_polynomial_connected[OF \open s\,of a b ] \connected (s - {})\ valid_path_polynomial_function "1.prems"(6) "1.prems"(7) by auto - moreover have "f contour_integrable_on g" + moreover have "f contour_integrable_on g" using contour_integrable_holomorphic_simple[OF _ \open s\ \valid_path g\ \path_image g \ s\,of f] \f holomorphic_on s - {}\ by auto ultimately show ?case using "1"(1)[of g] by auto next - case idt:(2 p pts) + case idt:(2 p pts) obtain e where "e>0" and e:"\w\ball a e. w \ s \ (w \ a \ w \ insert p pts)" - using finite_ball_avoid[OF \open s\ \finite (insert p pts)\, of a] + using finite_ball_avoid[OF \open s\ \finite (insert p pts)\, of a] \a \ s - insert p pts\ by auto define a' where "a' \ a+e/2" - have "a'\s-{p} -pts" using e[rule_format,of "a+e/2"] \e>0\ + have "a'\s-{p} -pts" using e[rule_format,of "a+e/2"] \e>0\ by (auto simp add:dist_complex_def a'_def) - then obtain g' where g'[simp]:"valid_path g'" "pathstart g' = a'" "pathfinish g' = b" - "path_image g' \ s - {p} - pts" "f contour_integrable_on g'" - using idt.hyps(3)[of a' "s-{p}"] idt.prems idt.hyps(1) + then obtain g' where g'[simp]:"valid_path g'" "pathstart g' = a'" "pathfinish g' = b" + "path_image g' \ s - {p} - pts" "f contour_integrable_on g'" + using idt.hyps(3)[of a' "s-{p}"] idt.prems idt.hyps(1) by (metis Diff_insert2 open_delete) define g where "g \ linepath a a' +++ g'" have "valid_path g" unfolding g_def by (auto intro: valid_path_join) - moreover have "pathstart g = a" and "pathfinish g = b" unfolding g_def by auto - moreover have "path_image g \ s - insert p pts" unfolding g_def + moreover have "pathstart g = a" and "pathfinish g = b" unfolding g_def by auto + moreover have "path_image g \ s - insert p pts" unfolding g_def proof (rule subset_path_image_join) - have "closed_segment a a' \ ball a e" using \e>0\ + have "closed_segment a a' \ ball a e" using \e>0\ by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute) then show "path_image (linepath a a') \ s - insert p pts" using e idt(9) by auto next show "path_image g' \ s - insert p pts" using g'(4) by blast qed - moreover have "f contour_integrable_on g" + moreover have "f contour_integrable_on g" proof - - have "closed_segment a a' \ ball a e" using \e>0\ + have "closed_segment a a' \ ball a e" using \e>0\ by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute) - then have "continuous_on (closed_segment a a') f" + then have "continuous_on (closed_segment a a') f" using e idt.prems(6) holomorphic_on_imp_continuous_on[OF idt.prems(5)] apply (elim continuous_on_subset) by auto - then have "f contour_integrable_on linepath a a'" + then have "f contour_integrable_on linepath a a'" using contour_integrable_continuous_linepath by auto then show ?thesis unfolding g_def apply (rule contour_integrable_joinI) @@ -2463,16 +2463,16 @@ lemma Cauchy_theorem_aux: assumes "open s" "connected (s-pts)" "finite pts" "pts \ s" "f holomorphic_on s-pts" - "valid_path g" "pathfinish g = pathstart g" "path_image g \ s-pts" - "\z. (z \ s) \ winding_number g z = 0" + "valid_path g" "pathfinish g = pathstart g" "path_image g \ s-pts" + "\z. (z \ s) \ winding_number g z = 0" "\p\s. h p>0 \ (\w\cball p (h p). w\s \ (w\p \ w \ pts))" shows "contour_integral g f = (\p\pts. winding_number g p * contour_integral (circlepath p (h p)) f)" using assms -proof (induct arbitrary:s g rule:finite_induct[OF \finite pts\]) +proof (induct arbitrary:s g rule:finite_induct[OF \finite pts\]) case 1 then show ?case by (simp add: Cauchy_theorem_global contour_integral_unique) next - case (2 p pts) + case (2 p pts) note fin[simp] = \finite (insert p pts)\ and connected = \connected (s - insert p pts)\ and valid[simp] = \valid_path g\ @@ -2484,12 +2484,12 @@ have "h p>0" and "p\s" and h_p: "\w\cball p (h p). w \ s \ (w \ p \ w \ insert p pts)" using h \insert p pts \ s\ by auto - obtain pg where pg[simp]: "valid_path pg" "pathstart pg = pathstart g" "pathfinish pg=p+h p" - "path_image pg \ s-insert p pts" "f contour_integrable_on pg" + obtain pg where pg[simp]: "valid_path pg" "pathstart pg = pathstart g" "pathfinish pg=p+h p" + "path_image pg \ s-insert p pts" "f contour_integrable_on pg" proof - - have "p + h p\cball p (h p)" using h[rule_format,of p] + have "p + h p\cball p (h p)" using h[rule_format,of p] by (simp add: \p \ s\ dist_norm) - then have "p + h p \ s - insert p pts" using h[rule_format,of p] \insert p pts \ s\ + then have "p + h p \ s - insert p pts" using h[rule_format,of p] \insert p pts \ s\ by fastforce moreover have "pathstart g \ s - insert p pts " using path_img by auto ultimately show ?thesis @@ -2497,32 +2497,32 @@ by blast qed obtain n::int where "n=winding_number g p" - using integer_winding_number[OF _ g_loop,of p] valid path_img + using integer_winding_number[OF _ g_loop,of p] valid path_img by (metis DiffD2 Ints_cases insertI1 subset_eq valid_path_imp_path) - define p_circ where "p_circ \ circlepath p (h p)" + define p_circ where "p_circ \ circlepath p (h p)" define p_circ_pt where "p_circ_pt \ linepath (p+h p) (p+h p)" define n_circ where "n_circ \ \n. (op +++ p_circ ^^ n) p_circ_pt" define cp where "cp \ if n\0 then reversepath (n_circ (nat n)) else n_circ (nat (- n))" - have n_circ:"valid_path (n_circ k)" + have n_circ:"valid_path (n_circ k)" "winding_number (n_circ k) p = k" - "pathstart (n_circ k) = p + h p" "pathfinish (n_circ k) = p + h p" + "pathstart (n_circ k) = p + h p" "pathfinish (n_circ k) = p + h p" "path_image (n_circ k) = (if k=0 then {p + h p} else sphere p (h p))" "p \ path_image (n_circ k)" "\p'. p'\s - pts \ winding_number (n_circ k) p'=0 \ p'\path_image (n_circ k)" "f contour_integrable_on (n_circ k)" "contour_integral (n_circ k) f = k * contour_integral p_circ f" - for k - proof (induct k) - case 0 - show "valid_path (n_circ 0)" + for k + proof (induct k) + case 0 + show "valid_path (n_circ 0)" and "path_image (n_circ 0) = (if 0=0 then {p + h p} else sphere p (h p))" - and "winding_number (n_circ 0) p = of_nat 0" + and "winding_number (n_circ 0) p = of_nat 0" and "pathstart (n_circ 0) = p + h p" and "pathfinish (n_circ 0) = p + h p" and "p \ path_image (n_circ 0)" unfolding n_circ_def p_circ_pt_def using \h p > 0\ by (auto simp add: dist_norm) - show "winding_number (n_circ 0) p'=0 \ p'\path_image (n_circ 0)" when "p'\s- pts" for p' + show "winding_number (n_circ 0) p'=0 \ p'\path_image (n_circ 0)" when "p'\s- pts" for p' unfolding n_circ_def p_circ_pt_def apply (auto intro!:winding_number_trivial) by (metis Diff_iff pathfinish_in_path_image pg(3) pg(4) subsetCE subset_insertI that)+ @@ -2534,51 +2534,51 @@ next case (Suc k) have n_Suc:"n_circ (Suc k) = p_circ +++ n_circ k" unfolding n_circ_def by auto - have pcirc:"p \ path_image p_circ" "valid_path p_circ" "pathfinish p_circ = pathstart (n_circ k)" + have pcirc:"p \ path_image p_circ" "valid_path p_circ" "pathfinish p_circ = pathstart (n_circ k)" using Suc(3) unfolding p_circ_def using \h p > 0\ by (auto simp add: p_circ_def) - have pcirc_image:"path_image p_circ \ s - insert p pts" + have pcirc_image:"path_image p_circ \ s - insert p pts" proof - have "path_image p_circ \ cball p (h p)" using \0 < h p\ p_circ_def by auto then show ?thesis using h_p pcirc(1) by auto qed have pcirc_integrable:"f contour_integrable_on p_circ" - by (auto simp add:p_circ_def intro!: pcirc_image[unfolded p_circ_def] - contour_integrable_continuous_circlepath holomorphic_on_imp_continuous_on + by (auto simp add:p_circ_def intro!: pcirc_image[unfolded p_circ_def] + contour_integrable_continuous_circlepath holomorphic_on_imp_continuous_on holomorphic_on_subset[OF holo]) - show "valid_path (n_circ (Suc k))" + show "valid_path (n_circ (Suc k))" using valid_path_join[OF pcirc(2) Suc(1) pcirc(3)] unfolding n_circ_def by auto - show "path_image (n_circ (Suc k)) - = (if Suc k = 0 then {p + complex_of_real (h p)} else sphere p (h p))" + show "path_image (n_circ (Suc k)) + = (if Suc k = 0 then {p + complex_of_real (h p)} else sphere p (h p))" proof - - have "path_image p_circ = sphere p (h p)" + have "path_image p_circ = sphere p (h p)" unfolding p_circ_def using \0 < h p\ by auto then show ?thesis unfolding n_Suc using Suc.hyps(5) \h p>0\ by (auto simp add: path_image_join[OF pcirc(3)] dist_norm) qed then show "p \ path_image (n_circ (Suc k))" using \h p>0\ by auto - show "winding_number (n_circ (Suc k)) p = of_nat (Suc k)" + show "winding_number (n_circ (Suc k)) p = of_nat (Suc k)" proof - - have "winding_number p_circ p = 1" + have "winding_number p_circ p = 1" by (simp add: \h p > 0\ p_circ_def winding_number_circlepath_centre) moreover have "p \ path_image (n_circ k)" using Suc(5) \h p>0\ by auto - then have "winding_number (p_circ +++ n_circ k) p + then have "winding_number (p_circ +++ n_circ k) p = winding_number p_circ p + winding_number (n_circ k) p" - using valid_path_imp_path Suc.hyps(1) Suc.hyps(2) pcirc + using valid_path_imp_path Suc.hyps(1) Suc.hyps(2) pcirc apply (intro winding_number_join) by auto - ultimately show ?thesis using Suc(2) unfolding n_circ_def + ultimately show ?thesis using Suc(2) unfolding n_circ_def by auto - qed + qed show "pathstart (n_circ (Suc k)) = p + h p" by (simp add: n_circ_def p_circ_def) show "pathfinish (n_circ (Suc k)) = p + h p" - using Suc(4) unfolding n_circ_def by auto + using Suc(4) unfolding n_circ_def by auto show "winding_number (n_circ (Suc k)) p'=0 \ p'\path_image (n_circ (Suc k))" when "p'\s-pts" for p' proof - have " p' \ path_image p_circ" using \p \ s\ h p_circ_def that using pcirc_image by blast - moreover have "p' \ path_image (n_circ k)" + moreover have "p' \ path_image (n_circ k)" using Suc.hyps(7) that by blast - moreover have "winding_number p_circ p' = 0" + moreover have "winding_number p_circ p' = 0" proof - have "path_image p_circ \ cball p (h p)" using h unfolding p_circ_def using \p \ s\ by fastforce @@ -2597,28 +2597,28 @@ by (rule contour_integrable_joinI[OF pcirc_integrable Suc(8) pcirc(2) Suc(1)]) show "contour_integral (n_circ (Suc k)) f = (Suc k) * contour_integral p_circ f" unfolding n_Suc - by (auto simp add:contour_integral_join[OF pcirc_integrable Suc(8) pcirc(2) Suc(1)] + by (auto simp add:contour_integral_join[OF pcirc_integrable Suc(8) pcirc(2) Suc(1)] Suc(9) algebra_simps) qed have cp[simp]:"pathstart cp = p + h p" "pathfinish cp = p + h p" - "valid_path cp" "path_image cp \ s - insert p pts" + "valid_path cp" "path_image cp \ s - insert p pts" "winding_number cp p = - n" "\p'. p'\s - pts \ winding_number cp p'=0 \ p' \ path_image cp" "f contour_integrable_on cp" "contour_integral cp f = - n * contour_integral p_circ f" - proof - + proof - show "pathstart cp = p + h p" and "pathfinish cp = p + h p" and "valid_path cp" using n_circ unfolding cp_def by auto next - have "sphere p (h p) \ s - insert p pts" + have "sphere p (h p) \ s - insert p pts" using h[rule_format,of p] \insert p pts \ s\ by force - moreover have "p + complex_of_real (h p) \ s - insert p pts" + moreover have "p + complex_of_real (h p) \ s - insert p pts" using pg(3) pg(4) by (metis pathfinish_in_path_image subsetCE) ultimately show "path_image cp \ s - insert p pts" unfolding cp_def using n_circ(5) by auto next - show "winding_number cp p = - n" - unfolding cp_def using winding_number_reversepath n_circ \h p>0\ + show "winding_number cp p = - n" + unfolding cp_def using winding_number_reversepath n_circ \h p>0\ by (auto simp: valid_path_imp_path) next show "winding_number cp p'=0 \ p' \ path_image cp" when "p'\s - pts" for p' @@ -2636,23 +2636,23 @@ qed def g'\"g +++ pg +++ cp +++ (reversepath pg)" have "contour_integral g' f = (\p\pts. winding_number g' p * contour_integral (circlepath p (h p)) f)" - proof (rule "2.hyps"(3)[of "s-{p}" "g'",OF _ _ \finite pts\ ]) + proof (rule "2.hyps"(3)[of "s-{p}" "g'",OF _ _ \finite pts\ ]) show "connected (s - {p} - pts)" using connected by (metis Diff_insert2) show "open (s - {p})" using \open s\ by auto - show " pts \ s - {p}" using \insert p pts \ s\ \ p \ pts\ by blast + show " pts \ s - {p}" using \insert p pts \ s\ \ p \ pts\ by blast show "f holomorphic_on s - {p} - pts" using holo \p \ pts\ by (metis Diff_insert2) - show "valid_path g'" + show "valid_path g'" unfolding g'_def cp_def using n_circ valid pg g_loop by (auto intro!:valid_path_join ) show "pathfinish g' = pathstart g'" unfolding g'_def cp_def using pg(2) by simp - show "path_image g' \ s - {p} - pts" + show "path_image g' \ s - {p} - pts" proof - def s'\"s - {p} - pts" have s':"s' = s-insert p pts " unfolding s'_def by auto - then show ?thesis using path_img pg(4) cp(4) - unfolding g'_def - apply (fold s'_def s') + then show ?thesis using path_img pg(4) cp(4) + unfolding g'_def + apply (fold s'_def s') apply (intro subset_path_image_join) by auto qed @@ -2660,12 +2660,12 @@ show "\z. z \ s - {p} \ winding_number g' z = 0" proof clarify fix z assume z:"z\s - {p}" - have "winding_number (g +++ pg +++ cp +++ reversepath pg) z = winding_number g z + have "winding_number (g +++ pg +++ cp +++ reversepath pg) z = winding_number g z + winding_number (pg +++ cp +++ (reversepath pg)) z" proof (rule winding_number_join) show "path g" using \valid_path g\ by (simp add: valid_path_imp_path) show "z \ path_image g" using z path_img by auto - show "path (pg +++ cp +++ reversepath pg)" using pg(3) cp + show "path (pg +++ cp +++ reversepath pg)" using pg(3) cp by (simp add: valid_path_imp_path) next have "path_image (pg +++ cp +++ reversepath pg) \ s - insert p pts" @@ -2674,52 +2674,52 @@ next show "pathfinish g = pathstart (pg +++ cp +++ reversepath pg)" using g_loop by auto qed - also have "... = winding_number g z + (winding_number pg z + also have "... = winding_number g z + (winding_number pg z + winding_number (cp +++ (reversepath pg)) z)" proof (subst add_left_cancel,rule winding_number_join) - show "path pg" and "path (cp +++ reversepath pg)" - and "pathfinish pg = pathstart (cp +++ reversepath pg)" + show "path pg" and "path (cp +++ reversepath pg)" + and "pathfinish pg = pathstart (cp +++ reversepath pg)" by (auto simp add: valid_path_imp_path) show "z \ path_image pg" using pg(4) z by blast - show "z \ path_image (cp +++ reversepath pg)" using z - by (metis Diff_iff \z \ path_image pg\ contra_subsetD cp(4) insertI1 + show "z \ path_image (cp +++ reversepath pg)" using z + by (metis Diff_iff \z \ path_image pg\ contra_subsetD cp(4) insertI1 not_in_path_image_join path_image_reversepath singletonD) qed - also have "... = winding_number g z + (winding_number pg z - + (winding_number cp z + winding_number (reversepath pg) z))" + also have "... = winding_number g z + (winding_number pg z + + (winding_number cp z + winding_number (reversepath pg) z))" apply (auto intro!:winding_number_join simp: valid_path_imp_path) apply (metis Diff_iff contra_subsetD cp(4) insertI1 singletonD z) - by (metis Diff_insert2 Diff_subset contra_subsetD pg(4) z) + by (metis Diff_insert2 Diff_subset contra_subsetD pg(4) z) also have "... = winding_number g z + winding_number cp z" apply (subst winding_number_reversepath) apply (auto simp: valid_path_imp_path) by (metis Diff_iff contra_subsetD insertI1 pg(4) singletonD z) - finally have "winding_number g' z = winding_number g z + winding_number cp z" + finally have "winding_number g' z = winding_number g z + winding_number cp z" unfolding g'_def . - moreover have "winding_number g z + winding_number cp z = 0" + moreover have "winding_number g z + winding_number cp z = 0" using winding z \n=winding_number g p\ by auto ultimately show "winding_number g' z = 0" unfolding g'_def by auto qed show "\pa\s - {p}. 0 < h pa \ (\w\cball pa (h pa). w \ s - {p} \ (w \ pa \ w \ pts))" using h by fastforce qed - moreover have "contour_integral g' f = contour_integral g f - - winding_number g p * contour_integral p_circ f" + moreover have "contour_integral g' f = contour_integral g f + - winding_number g p * contour_integral p_circ f" proof - - have "contour_integral g' f = contour_integral g f + have "contour_integral g' f = contour_integral g f + contour_integral (pg +++ cp +++ reversepath pg) f" unfolding g'_def - apply (subst Cauchy_Integral_Thm.contour_integral_join) + apply (subst contour_integral_join) by (auto simp add:open_Diff[OF \open s\,OF finite_imp_closed[OF fin]] - intro!: contour_integrable_holomorphic_simple[OF holo _ _ path_img] + intro!: contour_integrable_holomorphic_simple[OF holo _ _ path_img] contour_integrable_reversepath) - also have "... = contour_integral g f + contour_integral pg f + also have "... = contour_integral g f + contour_integral pg f + contour_integral (cp +++ reversepath pg) f" - apply (subst Cauchy_Integral_Thm.contour_integral_join) + apply (subst contour_integral_join) by (auto simp add:contour_integrable_reversepath) - also have "... = contour_integral g f + contour_integral pg f + also have "... = contour_integral g f + contour_integral pg f + contour_integral cp f + contour_integral (reversepath pg) f" - apply (subst Cauchy_Integral_Thm.contour_integral_join) + apply (subst contour_integral_join) by (auto simp add:contour_integrable_reversepath) also have "... = contour_integral g f + contour_integral cp f" using contour_integral_reversepath @@ -2728,21 +2728,21 @@ using \n=winding_number g p\ by auto finally show ?thesis . qed - moreover have "winding_number g' p' = winding_number g p'" when "p'\pts" for p' + moreover have "winding_number g' p' = winding_number g p'" when "p'\pts" for p' proof - - have [simp]: "p' \ path_image g" "p' \ path_image pg" "p'\path_image cp" - using "2.prems"(8) that + have [simp]: "p' \ path_image g" "p' \ path_image pg" "p'\path_image cp" + using "2.prems"(8) that apply blast apply (metis Diff_iff Diff_insert2 contra_subsetD pg(4) that) by (meson DiffD2 cp(4) set_rev_mp subset_insertI that) - have "winding_number g' p' = winding_number g p' - + winding_number (pg +++ cp +++ reversepath pg) p'" unfolding g'_def + have "winding_number g' p' = winding_number g p' + + winding_number (pg +++ cp +++ reversepath pg) p'" unfolding g'_def apply (subst winding_number_join) apply (simp_all add: valid_path_imp_path) apply (intro not_in_path_image_join) by auto also have "... = winding_number g p' + winding_number pg p' - + winding_number (cp +++ reversepath pg) p'" + + winding_number (cp +++ reversepath pg) p'" apply (subst winding_number_join) apply (simp_all add: valid_path_imp_path) apply (intro not_in_path_image_join) @@ -2757,16 +2757,16 @@ also have "... = winding_number g p'" using that by auto finally show ?thesis . qed - ultimately show ?case unfolding p_circ_def - apply (subst (asm) setsum.cong[OF refl, + ultimately show ?case unfolding p_circ_def + apply (subst (asm) setsum.cong[OF refl, of pts _ "\p. winding_number g p * contour_integral (circlepath p (h p)) f"]) - by (auto simp add:setsum.insert[OF \finite pts\ \p\pts\] algebra_simps) -qed + by (auto simp add:setsum.insert[OF \finite pts\ \p\pts\] algebra_simps) +qed lemma Cauchy_theorem_singularities: - assumes "open s" "connected s" "finite pts" and + assumes "open s" "connected s" "finite pts" and holo:"f holomorphic_on s-pts" and - "valid_path g" and + "valid_path g" and loop:"pathfinish g = pathstart g" and "path_image g \ s-pts" and homo:"\z. (z \ s) \ winding_number g z = 0" and @@ -2775,32 +2775,32 @@ (is "?L=?R") proof - define circ where "circ \ \p. winding_number g p * contour_integral (circlepath p (h p)) f" - define pts1 where "pts1 \ pts \ s" + define pts1 where "pts1 \ pts \ s" define pts2 where "pts2 \ pts - pts1" - have "pts=pts1 \ pts2" "pts1 \ pts2 = {}" "pts2 \ s={}" "pts1\s" + have "pts=pts1 \ pts2" "pts1 \ pts2 = {}" "pts2 \ s={}" "pts1\s" unfolding pts1_def pts2_def by auto have "contour_integral g f = (\p\pts1. circ p)" unfolding circ_def proof (rule Cauchy_theorem_aux[OF \open s\ _ _ \pts1\s\ _ \valid_path g\ loop _ homo]) have "finite pts1" unfolding pts1_def using \finite pts\ by auto - then show "connected (s - pts1)" + then show "connected (s - pts1)" using \open s\ \connected s\ connected_open_delete_finite[of s] by auto next show "finite pts1" using \pts = pts1 \ pts2\ assms(3) by auto show "f holomorphic_on s - pts1" by (metis Diff_Int2 Int_absorb holo pts1_def) - show "path_image g \ s - pts1" using assms(7) pts1_def by auto + show "path_image g \ s - pts1" using assms(7) pts1_def by auto show "\p\s. 0 < h p \ (\w\cball p (h p). w \ s \ (w \ p \ w \ pts1))" by (simp add: avoid pts1_def) qed moreover have "setsum circ pts2=0" proof - - have "winding_number g p=0" when "p\pts2" for p + have "winding_number g p=0" when "p\pts2" for p using \pts2 \ s={}\ that homo[rule_format,of p] by auto thus ?thesis unfolding circ_def apply (intro setsum.neutral) by auto qed moreover have "?R=setsum circ pts1 + setsum circ pts2" - unfolding circ_def + unfolding circ_def using setsum.union_disjoint[OF _ _ \pts1 \ pts2 = {}\] \finite pts\ \pts=pts1 \ pts2\ by blast ultimately show ?thesis @@ -2811,34 +2811,34 @@ lemma Residue_theorem: fixes s pts::"complex set" and f::"complex \ complex" and g::"real \ complex" - assumes "open s" "connected s" "finite pts" and + assumes "open s" "connected s" "finite pts" and holo:"f holomorphic_on s-pts" and - "valid_path g" and + "valid_path g" and loop:"pathfinish g = pathstart g" and "path_image g \ s-pts" and - homo:"\z. (z \ s) \ winding_number g z = 0" + homo:"\z. (z \ s) \ winding_number g z = 0" shows "contour_integral g f = 2 * pi * \ *(\p\pts. winding_number g p * residue f p)" proof - define c where "c \ 2 * pi * \" - obtain h where avoid:"\p\s. h p>0 \ (\w\cball p (h p). w\s \ (w\p \ w \ pts))" + obtain h where avoid:"\p\s. h p>0 \ (\w\cball p (h p). w\s \ (w\p \ w \ pts))" using finite_cball_avoid[OF \open s\ \finite pts\] by metis - have "contour_integral g f + have "contour_integral g f = (\p\pts. winding_number g p * contour_integral (circlepath p (h p)) f)" using Cauchy_theorem_singularities[OF assms avoid] . also have "... = (\p\pts. c * winding_number g p * residue f p)" proof (intro setsum.cong) show "pts = pts" by simp next - fix x assume "x \ pts" - show "winding_number g x * contour_integral (circlepath x (h x)) f - = c * winding_number g x * residue f x" + fix x assume "x \ pts" + show "winding_number g x * contour_integral (circlepath x (h x)) f + = c * winding_number g x * residue f x" proof (cases "x\s") case False then have "winding_number g x=0" using homo by auto thus ?thesis by auto next case True - have "contour_integral (circlepath x (h x)) f = c* residue f x" + have "contour_integral (circlepath x (h x)) f = c* residue f x" using \x\pts\ \finite pts\ avoid[rule_format,OF True] apply (intro base_residue[of "s-(pts-{x})",THEN contour_integral_unique,folded c_def]) by (auto intro:holomorphic_on_subset[OF holo] open_Diff[OF \open s\ finite_imp_closed]) @@ -2846,9 +2846,9 @@ qed qed also have "... = c * (\p\pts. winding_number g p * residue f p)" - by (simp add: setsum_right_distrib algebra_simps) + by (simp add: setsum_right_distrib algebra_simps) finally show ?thesis unfolding c_def . -qed +qed subsection \The argument principle\ @@ -2858,13 +2858,13 @@ lemma is_pole_tendsto: fixes f::"('a::topological_space \ 'b::real_normed_div_algebra)" shows "is_pole f x \ ((inverse o f) \ 0) (at x)" -unfolding is_pole_def +unfolding is_pole_def by (auto simp add:filterlim_inverse_at_iff[symmetric] comp_def filterlim_at) lemma is_pole_inverse_holomorphic: assumes "open s" - and f_holo:"f holomorphic_on (s-{z})" - and pole:"is_pole f z" + and f_holo:"f holomorphic_on (s-{z})" + and pole:"is_pole f z" and non_z:"\x\s-{z}. f x\0" shows "(\x. if x=z then 0 else inverse (f x)) holomorphic_on s" proof - @@ -2875,7 +2875,7 @@ moreover have "continuous_on (s-{z}) f" using f_holo holomorphic_on_imp_continuous_on by auto hence "continuous_on (s-{z}) (inverse o f)" unfolding comp_def by (auto elim!:continuous_on_inverse simp add:non_z) - hence "continuous_on (s-{z}) g" unfolding g_def + hence "continuous_on (s-{z}) g" unfolding g_def apply (subst continuous_on_cong[where t="s-{z}" and g="inverse o f"]) by auto ultimately have "continuous_on s g" using open_delete[OF \open s\] \open s\ @@ -2887,17 +2887,17 @@ apply (subst holomorphic_cong[where t="s-{z}" and g="inverse o f"]) by (auto simp add:g_def) ultimately show ?thesis unfolding g_def using \open s\ - by (auto elim!: no_isolated_singularity) + by (auto elim!: no_isolated_singularity) qed (*order of the zero of f at z*) definition zorder::"(complex \ complex) \ complex \ nat" where - "zorder f z = (THE n. n>0 \ (\h r. r>0 \ h holomorphic_on cball z r + "zorder f z = (THE n. n>0 \ (\h r. r>0 \ h holomorphic_on cball z r \ (\w\cball z r. f w = h w * (w-z)^n \ h w \0)))" definition zer_poly::"[complex \ complex,complex]\complex \ complex" where - "zer_poly f z = (SOME h. \r . r>0 \ h holomorphic_on cball z r + "zer_poly f z = (SOME h. \r . r>0 \ h holomorphic_on cball z r \ (\w\cball z r. f w = h w * (w-z)^(zorder f z) \ h w \0))" (*order of the pole of f at z*) @@ -2905,25 +2905,25 @@ "porder f z = (let f'=(\x. if x=z then 0 else inverse (f x)) in zorder f' z)" definition pol_poly::"[complex \ complex,complex]\complex \ complex" where - "pol_poly f z = (let f'=(\ x. if x=z then 0 else inverse (f x)) + "pol_poly f z = (let f'=(\ x. if x=z then 0 else inverse (f x)) in inverse o zer_poly f' z)" -lemma holomorphic_factor_zero_unique: - fixes f::"complex \ complex" and z::complex and r::real +lemma holomorphic_factor_zero_unique: + fixes f::"complex \ complex" and z::complex and r::real assumes "r>0" and asm:"\w\ball z r. f w = (w-z)^n * g w \ g w\0 \ f w = (w - z)^m * h w \ h w\0" and g_holo:"g holomorphic_on ball z r" and h_holo:"h holomorphic_on ball z r" - shows "n=m" + shows "n=m" proof - have "n>m \ False" proof - assume "n>m" - have "(h \ 0) (at z within ball z r)" - proof (rule Lim_transform_within[OF _ \r>0\, where f="\w. (w - z) ^ (n - m) * g w"]) + have "(h \ 0) (at z within ball z r)" + proof (rule Lim_transform_within[OF _ \r>0\, where f="\w. (w - z) ^ (n - m) * g w"]) have "\w\ball z r. w\z \ h w = (w-z)^(n-m) * g w" using \n>m\ asm by (auto simp add:field_simps power_diff) - then show "\x' \ ball z r; 0 < dist x' z;dist x' z < r\ + then show "\x' \ ball z r; 0 < dist x' z;dist x' z < r\ \ (x' - z) ^ (n - m) * g x' = h x'" for x' by auto next define F where "F \ at z within ball z r" @@ -2932,16 +2932,16 @@ moreover have "continuous F f'" unfolding f'_def F_def by (intro continuous_intros) ultimately have "(f' \ 0) F" unfolding F_def - by (simp add: continuous_within) + by (simp add: continuous_within) moreover have "(g \ g z) F" using holomorphic_on_imp_continuous_on[OF g_holo,unfolded continuous_on_def] \r>0\ unfolding F_def by auto ultimately show " ((\w. f' w * g w) \ 0) F" using tendsto_mult by fastforce qed - moreover have "(h \ h z) (at z within ball z r)" - using holomorphic_on_imp_continuous_on[OF h_holo] - by (auto simp add:continuous_on_def \r>0\) - moreover have "at z within ball z r \ bot" using \r>0\ + moreover have "(h \ h z) (at z within ball z r)" + using holomorphic_on_imp_continuous_on[OF h_holo] + by (auto simp add:continuous_on_def \r>0\) + moreover have "at z within ball z r \ bot" using \r>0\ by (auto simp add:trivial_limit_within islimpt_ball) ultimately have "h z=0" by (auto intro: tendsto_unique) thus False using asm \r>0\ by auto @@ -2949,11 +2949,11 @@ moreover have "m>n \ False" proof - assume "m>n" - have "(g \ 0) (at z within ball z r)" - proof (rule Lim_transform_within[OF _ \r>0\, where f="\w. (w - z) ^ (m - n) * h w"]) + have "(g \ 0) (at z within ball z r)" + proof (rule Lim_transform_within[OF _ \r>0\, where f="\w. (w - z) ^ (m - n) * h w"]) have "\w\ball z r. w\z \ g w = (w-z)^(m-n) * h w" using \m>n\ asm by (auto simp add:field_simps power_diff) - then show "\x' \ ball z r; 0 < dist x' z;dist x' z < r\ + then show "\x' \ ball z r; 0 < dist x' z;dist x' z < r\ \ (x' - z) ^ (m - n) * h x' = g x'" for x' by auto next define F where "F \ at z within ball z r" @@ -2962,16 +2962,16 @@ moreover have "continuous F f'" unfolding f'_def F_def by (intro continuous_intros) ultimately have "(f' \ 0) F" unfolding F_def - by (simp add: continuous_within) + by (simp add: continuous_within) moreover have "(h \ h z) F" using holomorphic_on_imp_continuous_on[OF h_holo,unfolded continuous_on_def] \r>0\ unfolding F_def by auto ultimately show " ((\w. f' w * h w) \ 0) F" using tendsto_mult by fastforce qed - moreover have "(g \ g z) (at z within ball z r)" - using holomorphic_on_imp_continuous_on[OF g_holo] - by (auto simp add:continuous_on_def \r>0\) - moreover have "at z within ball z r \ bot" using \r>0\ + moreover have "(g \ g z) (at z within ball z r)" + using holomorphic_on_imp_continuous_on[OF g_holo] + by (auto simp add:continuous_on_def \r>0\) + moreover have "at z within ball z r \ bot" using \r>0\ by (auto simp add:trivial_limit_within islimpt_ball) ultimately have "g z=0" by (auto intro: tendsto_unique) thus False using asm \r>0\ by auto @@ -2982,62 +2982,62 @@ lemma holomorphic_factor_zero_Ex1: assumes "open s" "connected s" "z \ s" and - holo:"f holomorphic_on s" + holo:"f holomorphic_on s" and "f z = 0" and "\w\s. f w \ 0" shows "\!n. \g r. 0 < n \ 0 < r \ - g holomorphic_on cball z r + g holomorphic_on cball z r \ (\w\cball z r. f w = (w-z)^n * g w \ g w\0)" proof (rule ex_ex1I) - obtain g r n where "0 < n" "0 < r" "ball z r \ s" and + obtain g r n where "0 < n" "0 < r" "ball z r \ s" and g:"g holomorphic_on ball z r" - "\w. w \ ball z r \ f w = (w - z) ^ n * g w" + "\w. w \ ball z r \ f w = (w - z) ^ n * g w" "\w. w \ ball z r \ g w \ 0" using holomorphic_factor_zero_nonconstant[OF holo \open s\ \connected s\ \z\s\ \f z=0\] by (metis assms(3) assms(5) assms(6)) def r'\"r/2" have "cball z r' \ ball z r" unfolding r'_def by (simp add: \0 < r\ cball_subset_ball_iff) - hence "cball z r' \ s" "g holomorphic_on cball z r'" + hence "cball z r' \ s" "g holomorphic_on cball z r'" "(\w\cball z r'. f w = (w - z) ^ n * g w \ g w \ 0)" using g \ball z r \ s\ by auto moreover have "r'>0" unfolding r'_def using \0 by auto - ultimately show "\n g r. 0 < n \ 0 < r \ g holomorphic_on cball z r + ultimately show "\n g r. 0 < n \ 0 < r \ g holomorphic_on cball z r \ (\w\cball z r. f w = (w - z) ^ n * g w \ g w \ 0)" apply (intro exI[of _ n] exI[of _ g] exI[of _ r']) by (simp add:\0 < n\) next - fix m n - define fac where "fac \ \n g r. \w\cball z r. f w = (w - z) ^ n * g w \ g w \ 0" + fix m n + define fac where "fac \ \n g r. \w\cball z r. f w = (w - z) ^ n * g w \ g w \ 0" assume n_asm:"\g r1. 0 < n \ 0 < r1 \ g holomorphic_on cball z r1 \ fac n g r1" and m_asm:"\h r2. 0 < m \ 0 < r2 \ h holomorphic_on cball z r2 \ fac m h r2" - obtain g r1 where "0 < n" "0 < r1" and g_holo: "g holomorphic_on cball z r1" + obtain g r1 where "0 < n" "0 < r1" and g_holo: "g holomorphic_on cball z r1" and "fac n g r1" using n_asm by auto - obtain h r2 where "0 < m" "0 < r2" and h_holo: "h holomorphic_on cball z r2" + obtain h r2 where "0 < m" "0 < r2" and h_holo: "h holomorphic_on cball z r2" and "fac m h r2" using m_asm by auto define r where "r \ min r1 r2" have "r>0" using \r1>0\ \r2>0\ unfolding r_def by auto - moreover have "\w\ball z r. f w = (w-z)^n * g w \ g w\0 \ f w = (w - z)^m * h w \ h w\0" + moreover have "\w\ball z r. f w = (w-z)^n * g w \ g w\0 \ f w = (w - z)^m * h w \ h w\0" using \fac m h r2\ \fac n g r1\ unfolding fac_def r_def by fastforce - ultimately show "m=n" using g_holo h_holo + ultimately show "m=n" using g_holo h_holo apply (elim holomorphic_factor_zero_unique[of r z f n g m h,symmetric,rotated]) by (auto simp add:r_def) qed -lemma zorder_exist: +lemma zorder_exist: fixes f::"complex \ complex" and z::complex defines "n\zorder f z" and "h\zer_poly f z" - assumes "open s" "connected s" "z\s" + assumes "open s" "connected s" "z\s" and holo: "f holomorphic_on s" and "f z=0" "\w\s. f w\0" - shows "\r. n>0 \ r>0 \ cball z r \ s \ h holomorphic_on cball z r + shows "\r. n>0 \ r>0 \ cball z r \ s \ h holomorphic_on cball z r \ (\w\cball z r. f w = h w * (w-z)^n \ h w \0) " proof - - define P where "P \ \h r n. r>0 \ h holomorphic_on cball z r + define P where "P \ \h r n. r>0 \ h holomorphic_on cball z r \ (\w\cball z r. ( f w = h w * (w-z)^n) \ h w \0)" have "(\!n. n>0 \ (\ h r. P h r n))" proof - have "\!n. \h r. n>0 \ P h r n" - using holomorphic_factor_zero_Ex1[OF \open s\ \connected s\ \z\s\ holo \f z=0\ + using holomorphic_factor_zero_Ex1[OF \open s\ \connected s\ \z\s\ holo \f z=0\ \\w\s. f w\0\] unfolding P_def apply (subst mult.commute) by auto @@ -3045,77 +3045,77 @@ qed moreover have n:"n=(THE n. n>0 \ (\h r. P h r n))" unfolding n_def zorder_def P_def by simp - ultimately have "n>0 \ (\h r. P h r n)" + ultimately have "n>0 \ (\h r. P h r n)" apply (drule_tac theI') by simp then have "n>0" and "\h r. P h r n" by auto moreover have "h=(SOME h. \r. P h r n)" unfolding h_def P_def zer_poly_def[of f z,folded n_def P_def] by simp - ultimately have "\r. P h r n" + ultimately have "\r. P h r n" apply (drule_tac someI_ex) by simp then obtain r1 where "P h r1 n" by auto - obtain r2 where "r2>0" "cball z r2 \ s" + obtain r2 where "r2>0" "cball z r2 \ s" using assms(3) assms(5) open_contains_cball_eq by blast define r3 where "r3 \ min r1 r2" have "P h r3 n" using \P h r1 n\ \r2>0\ unfolding P_def r3_def by auto moreover have "cball z r3 \ s" using \cball z r2 \ s\ unfolding r3_def by auto ultimately show ?thesis using \n>0\ unfolding P_def by auto -qed +qed -lemma porder_exist: +lemma porder_exist: fixes f::"complex \ complex" and z::complex defines "n \ porder f z" and "h \ pol_poly f z" - assumes "open s" "z \ s" + assumes "open s" "z \ s" and holo:"f holomorphic_on s-{z}" and "is_pole f z" - shows "\r. n>0 \ r>0 \ cball z r \ s \ h holomorphic_on cball z r + shows "\r. n>0 \ r>0 \ cball z r \ s \ h holomorphic_on cball z r \ (\w\cball z r. (w\z \ f w = h w / (w-z)^n) \ h w \0)" proof - obtain e where "e>0" and e_ball:"ball z e \ s"and e_def: "\x\ball z e-{z}. f x\0" - proof - - have "\\<^sub>F z in at z. f z \ 0" - using \is_pole f z\ filterlim_at_infinity_imp_eventually_ne unfolding is_pole_def + proof - + have "\\<^sub>F z in at z. f z \ 0" + using \is_pole f z\ filterlim_at_infinity_imp_eventually_ne unfolding is_pole_def by auto then obtain e1 where "e1>0" and e1_def: "\x. x \ z \ dist x z < e1 \ f x \ 0" using eventually_at[of "\x. f x\0" z,simplified] by auto obtain e2 where "e2>0" and "ball z e2 \s" using \open s\ \z\s\ openE by auto define e where "e \ min e1 e2" have "e>0" using \e1>0\ \e2>0\ unfolding e_def by auto - moreover have "ball z e \ s" unfolding e_def using \ball z e2 \ s\ by auto - moreover have "\x\ball z e-{z}. f x\0" using e1_def \e1>0\ \e2>0\ unfolding e_def + moreover have "ball z e \ s" unfolding e_def using \ball z e2 \ s\ by auto + moreover have "\x\ball z e-{z}. f x\0" using e1_def \e1>0\ \e2>0\ unfolding e_def by (simp add: DiffD1 DiffD2 dist_commute singletonI) ultimately show ?thesis using that by auto qed define g where "g \ \x. if x=z then 0 else inverse (f x)" - define zo where "zo \ zorder g z" - define zp where "zp \ zer_poly g z" - have "\w\ball z e. g w \ 0" + define zo where "zo \ zorder g z" + define zp where "zp \ zer_poly g z" + have "\w\ball z e. g w \ 0" proof - obtain w where w:"w\ball z e-{z}" using \0 < e\ - by (metis open_ball all_not_in_conv centre_in_ball insert_Diff_single + by (metis open_ball all_not_in_conv centre_in_ball insert_Diff_single insert_absorb not_open_singleton) hence "w\z" "f w\0" using e_def[rule_format,of w] mem_ball by (auto simp add:dist_commute) then show ?thesis unfolding g_def using w by auto qed - moreover have "g holomorphic_on ball z e" + moreover have "g holomorphic_on ball z e" apply (intro is_pole_inverse_holomorphic[of "ball z e",OF _ _ \is_pole f z\ e_def,folded g_def]) using holo e_ball by auto moreover have "g z=0" unfolding g_def by auto - ultimately obtain r where "0 < zo" "0 < r" "cball z r \ ball z e" + ultimately obtain r where "0 < zo" "0 < r" "cball z r \ ball z e" and zp_holo: "zp holomorphic_on cball z r" and zp_fac: "\w\cball z r. g w = zp w * (w - z) ^ zo \ zp w \ 0" - using zorder_exist[of "ball z e" z g,simplified,folded zo_def zp_def] \e>0\ + using zorder_exist[of "ball z e" z g,simplified,folded zo_def zp_def] \e>0\ by auto - have n:"n=zo" and h:"h=inverse o zp" + have n:"n=zo" and h:"h=inverse o zp" unfolding n_def zo_def porder_def h_def zp_def pol_poly_def g_def by simp_all have "h holomorphic_on cball z r" using zp_holo zp_fac holomorphic_on_inverse unfolding h comp_def by blast moreover have "\w\cball z r. (w\z \ f w = h w / (w-z)^n) \ h w \0" using zp_fac unfolding h n comp_def g_def - by (metis divide_inverse_commute field_class.field_inverse_zero inverse_inverse_eq + by (metis divide_inverse_commute field_class.field_inverse_zero inverse_inverse_eq inverse_mult_distrib mult.commute) moreover have "0 < n" unfolding n using \zo>0\ by simp ultimately show ?thesis using \0 < r\ \cball z r \ ball z e\ e_ball by auto @@ -3124,13 +3124,13 @@ lemma residue_porder: fixes f::"complex \ complex" and z::complex defines "n \ porder f z" and "h \ pol_poly f z" - assumes "open s" "z \ s" + assumes "open s" "z \ s" and holo:"f holomorphic_on s - {z}" and pole:"is_pole f z" shows "residue f z = ((deriv ^^ (n - 1)) h z / fact (n-1))" proof - define g where "g \ \x. if x=z then 0 else inverse (f x)" - obtain r where "0 < n" "0 < r" and r_cball:"cball z r \ s" and h_holo: "h holomorphic_on cball z r" + obtain r where "0 < n" "0 < r" and r_cball:"cball z r \ s" and h_holo: "h holomorphic_on cball z r" and h_divide:"(\w\cball z r. (w\z \ f w = h w / (w - z) ^ n) \ h w \ 0)" using porder_exist[OF \open s\ \z \ s\ holo pole, folded n_def h_def] by blast have r_nonzero:"\w. w \ ball z r - {z} \ f w \ 0" @@ -3138,7 +3138,7 @@ define c where "c \ 2 * pi * \" define der_f where "der_f \ ((deriv ^^ (n - 1)) h z / fact (n-1))" def h'\"\u. h u / (u - z) ^ n" - have "(h' has_contour_integral c / fact (n - 1) * (deriv ^^ (n - 1)) h z) (circlepath z r)" + have "(h' has_contour_integral c / fact (n - 1) * (deriv ^^ (n - 1)) h z) (circlepath z r)" unfolding h'_def proof (rule Cauchy_has_contour_integral_higher_derivative_circlepath[of z r h z "n-1", folded c_def Suc_pred'[OF \n>0\]]) @@ -3147,9 +3147,9 @@ show " z \ ball z r" using \r>0\ by auto qed then have "(h' has_contour_integral c * der_f) (circlepath z r)" unfolding der_f_def by auto - then have "(f has_contour_integral c * der_f) (circlepath z r)" + then have "(f has_contour_integral c * der_f) (circlepath z r)" proof (elim has_contour_integral_eq) - fix x assume "x \ path_image (circlepath z r)" + fix x assume "x \ path_image (circlepath z r)" hence "x\cball z r - {z}" using \r>0\ by auto then show "h' x = f x" using h_divide unfolding h'_def by auto qed @@ -3163,18 +3163,18 @@ theorem argument_principle: fixes f::"complex \ complex" and poles s:: "complex set" defines "zeros\{p. f p=0} - poles" - assumes "open s" and - "connected s" and - f_holo:"f holomorphic_on s-poles" and + assumes "open s" and + "connected s" and + f_holo:"f holomorphic_on s-poles" and h_holo:"h holomorphic_on s" and - "valid_path g" and + "valid_path g" and loop:"pathfinish g = pathstart g" and path_img:"path_image g \ s - (zeros \ poles)" and homo:"\z. (z \ s) \ winding_number g z = 0" and finite:"finite (zeros \ poles)" and poles:"\p\poles. is_pole f p" shows "contour_integral g (\x. deriv f x * h x / f x) = 2 * pi * \ * - ((\p\zeros. winding_number g p * h p * zorder f p) + ((\p\zeros. winding_number g p * h p * zorder f p) - (\p\poles. winding_number g p * h p * porder f p))" (is "?L=?R") proof - @@ -3183,40 +3183,40 @@ define cont_pole where "cont_pole \ \ff p e. (ff has_contour_integral - c * porder f p * h p) (circlepath p e)" define cont_zero where "cont_zero \ \ff p e. (ff has_contour_integral c * zorder f p * h p ) (circlepath p e)" define avoid where "avoid \ \p e. \w\cball p e. w \ s \ (w \ p \ w \ zeros \ poles)" - have "\e>0. avoid p e \ (p\poles \ cont_pole ff p e) \ (p\zeros \ cont_zero ff p e)" - when "p\s" for p + have "\e>0. avoid p e \ (p\poles \ cont_pole ff p e) \ (p\zeros \ cont_zero ff p e)" + when "p\s" for p proof - obtain e1 where "e1>0" and e1_avoid:"avoid p e1" using finite_cball_avoid[OF \open s\ finite] \p\s\ unfolding avoid_def by auto have "\e2>0. cball p e2 \ ball p e1 \ cont_pole ff p e2" - when "p\poles" - proof - + when "p\poles" + proof - define po where "po \ porder f p" define pp where "pp \ pol_poly f p" def f'\"\w. pp w / (w - p) ^ po" - def ff'\"(\x. deriv f' x * h x / f' x)" - have "f holomorphic_on ball p e1 - {p}" + def ff'\"(\x. deriv f' x * h x / f' x)" + have "f holomorphic_on ball p e1 - {p}" apply (intro holomorphic_on_subset[OF f_holo]) using e1_avoid \p\poles\ unfolding avoid_def by auto - then obtain r where + then obtain r where "0 < po" "r>0" "cball p r \ ball p e1" and pp_holo:"pp holomorphic_on cball p r" and pp_po:"(\w\cball p r. (w\p \ f w = pp w / (w - p) ^ po) \ pp w \ 0)" using porder_exist[of "ball p e1" p f,simplified,OF \e1>0\] poles \p\poles\ - unfolding po_def pp_def - by auto + unfolding po_def pp_def + by auto define e2 where "e2 \ r/2" have "e2>0" using \r>0\ unfolding e2_def by auto - define anal where "anal \ \w. deriv pp w * h w / pp w" + define anal where "anal \ \w. deriv pp w * h w / pp w" define prin where "prin \ \w. - of_nat po * h w / (w - p)" have "((\w. prin w + anal w) has_contour_integral - c * po * h p) (circlepath p e2)" proof (rule has_contour_integral_add[of _ _ _ _ 0,simplified]) - have "ball p r \ s" + have "ball p r \ s" using \cball p r \ ball p e1\ avoid_def ball_subset_cball e1_avoid by blast then have "cball p e2 \ s" - using \r>0\ unfolding e2_def by auto - then have "(\w. - of_nat po * h w) holomorphic_on cball p e2" + using \r>0\ unfolding e2_def by auto + then have "(\w. - of_nat po * h w) holomorphic_on cball p e2" using h_holo by (auto intro!: holomorphic_intros) then show "(prin has_contour_integral - c * of_nat po * h p ) (circlepath p e2)" @@ -3224,7 +3224,7 @@ \e2>0\ unfolding prin_def by (auto simp add: mult.assoc) - have "anal holomorphic_on ball p r" unfolding anal_def + have "anal holomorphic_on ball p r" unfolding anal_def using pp_holo h_holo pp_po \ball p r \ s\ by (auto intro!: holomorphic_intros) then show "(anal has_contour_integral 0) (circlepath p e2)" @@ -3233,19 +3233,19 @@ qed then have "cont_pole ff' p e2" unfolding cont_pole_def po_def proof (elim has_contour_integral_eq) - fix w assume "w \ path_image (circlepath p e2)" + fix w assume "w \ path_image (circlepath p e2)" then have "w\ball p r" and "w\p" unfolding e2_def using \r>0\ by auto define wp where "wp \ w-p" - have "wp\0" and "pp w \0" - unfolding wp_def using \w\p\ \w\ball p r\ pp_po by auto - moreover have der_f':"deriv f' w = - po * pp w / (w-p)^(po+1) + deriv pp w / (w-p)^po" + have "wp\0" and "pp w \0" + unfolding wp_def using \w\p\ \w\ball p r\ pp_po by auto + moreover have der_f':"deriv f' w = - po * pp w / (w-p)^(po+1) + deriv pp w / (w-p)^po" proof (rule DERIV_imp_deriv) define der where "der \ - po * pp w / (w-p)^(po+1) + deriv pp w / (w-p)^po" have po:"po = Suc (po - Suc 0) " using \po>0\ by auto - have "(pp has_field_derivative (deriv pp w)) (at w)" - using DERIV_deriv_iff_has_field_derivative pp_holo \w\p\ + have "(pp has_field_derivative (deriv pp w)) (at w)" + using DERIV_deriv_iff_has_field_derivative pp_holo \w\p\ by (meson open_ball \w \ ball p r\ ball_subset_cball holomorphic_derivI holomorphic_on_subset) - then show "(f' has_field_derivative der) (at w)" + then show "(f' has_field_derivative der) (at w)" using \w\p\ \po>0\ unfolding der_def f'_def apply (auto intro!: derivative_eq_intros simp add:field_simps) apply (subst (4) po) @@ -3259,16 +3259,16 @@ apply (fold wp_def) by (auto simp add:field_simps) qed - then have "cont_pole ff p e2" unfolding cont_pole_def + then have "cont_pole ff p e2" unfolding cont_pole_def proof (elim has_contour_integral_eq) fix w assume "w \ path_image (circlepath p e2)" then have "w\ball p r" and "w\p" unfolding e2_def using \r>0\ by auto - have "deriv f' w = deriv f w" + have "deriv f' w = deriv f w" proof (rule complex_derivative_transform_within_open[where s="ball p r - {p}"]) show "f' holomorphic_on ball p r - {p}" unfolding f'_def using pp_holo by (auto intro!: holomorphic_intros) - next - have "ball p e1 - {p} \ s - poles" + next + have "ball p e1 - {p} \ s - poles" using avoid_def ball_subset_cball e1_avoid by auto then have "ball p r - {p} \ s - poles" using \cball p r \ ball p e1\ @@ -3281,16 +3281,16 @@ show "w \ ball p r - {p}" using \w\ball p r\ \w\p\ by auto next fix x assume "x \ ball p r - {p}" - then show "f' x = f x" + then show "f' x = f x" using pp_po unfolding f'_def by auto qed - moreover have " f' w = f w " - using \w \ ball p r\ ball_subset_cball subset_iff pp_po \w\p\ + moreover have " f' w = f w " + using \w \ ball p r\ ball_subset_cball subset_iff pp_po \w\p\ unfolding f'_def by auto ultimately show "ff' w = ff w" unfolding ff'_def ff_def by simp qed - moreover have "cball p e2 \ ball p e1" + moreover have "cball p e2 \ ball p e1" using \0 < r\ \cball p r \ ball p e1\ e2_def by auto ultimately show ?thesis using \e2>0\ by auto qed @@ -3303,15 +3303,15 @@ define zp where "zp \ zer_poly f p" def f'\"\w. zp w * (w - p) ^ zo" def ff'\"(\x. deriv f' x * h x / f' x)" - have "f holomorphic_on ball p e1" + have "f holomorphic_on ball p e1" proof - - have "ball p e1 \ s - poles" + have "ball p e1 \ s - poles" using avoid_def ball_subset_cball e1_avoid that zeros_def by fastforce thus ?thesis using f_holo by auto qed - moreover have "f p = 0" using \p\zeros\ + moreover have "f p = 0" using \p\zeros\ using DiffD1 mem_Collect_eq zeros_def by blast - moreover have "\w\ball p e1. f w \ 0" + moreover have "\w\ball p e1. f w \ 0" proof - def p'\"p+e1/2" have "p'\ball p e1" and "p'\p" using \e1>0\ unfolding p'_def by (auto simp add:dist_norm) @@ -3319,24 +3319,24 @@ apply (rule_tac x=p' in bexI) by (auto simp add:zeros_def) qed - ultimately obtain r where + ultimately obtain r where "0 < zo" "r>0" "cball p r \ ball p e1" and pp_holo:"zp holomorphic_on cball p r" and pp_po:"(\w\cball p r. f w = zp w * (w - p) ^ zo \ zp w \ 0)" using zorder_exist[of "ball p e1" p f,simplified,OF \e1>0\] unfolding zo_def zp_def - by auto + by auto define e2 where "e2 \ r/2" have "e2>0" using \r>0\ unfolding e2_def by auto - define anal where "anal \ \w. deriv zp w * h w / zp w" + define anal where "anal \ \w. deriv zp w * h w / zp w" define prin where "prin \ \w. of_nat zo * h w / (w - p)" have "((\w. prin w + anal w) has_contour_integral c * zo * h p) (circlepath p e2)" proof (rule has_contour_integral_add[of _ _ _ _ 0,simplified]) - have "ball p r \ s" + have "ball p r \ s" using \cball p r \ ball p e1\ avoid_def ball_subset_cball e1_avoid by blast then have "cball p e2 \ s" - using \r>0\ unfolding e2_def by auto - then have "(\w. of_nat zo * h w) holomorphic_on cball p e2" + using \r>0\ unfolding e2_def by auto + then have "(\w. of_nat zo * h w) holomorphic_on cball p e2" using h_holo by (auto intro!: holomorphic_intros) then show "(prin has_contour_integral c * of_nat zo * h p ) (circlepath p e2)" @@ -3344,7 +3344,7 @@ \e2>0\ unfolding prin_def by (auto simp add: mult.assoc) - have "anal holomorphic_on ball p r" unfolding anal_def + have "anal holomorphic_on ball p r" unfolding anal_def using pp_holo h_holo pp_po \ball p r \ s\ by (auto intro!: holomorphic_intros) then show "(anal has_contour_integral 0) (circlepath p e2)" @@ -3353,19 +3353,19 @@ qed then have "cont_zero ff' p e2" unfolding cont_zero_def zo_def proof (elim has_contour_integral_eq) - fix w assume "w \ path_image (circlepath p e2)" + fix w assume "w \ path_image (circlepath p e2)" then have "w\ball p r" and "w\p" unfolding e2_def using \r>0\ by auto define wp where "wp \ w-p" - have "wp\0" and "zp w \0" - unfolding wp_def using \w\p\ \w\ball p r\ pp_po by auto - moreover have der_f':"deriv f' w = zo * zp w * (w-p)^(zo-1) + deriv zp w * (w-p)^zo" + have "wp\0" and "zp w \0" + unfolding wp_def using \w\p\ \w\ball p r\ pp_po by auto + moreover have der_f':"deriv f' w = zo * zp w * (w-p)^(zo-1) + deriv zp w * (w-p)^zo" proof (rule DERIV_imp_deriv) define der where "der \ zo * zp w * (w-p)^(zo-1) + deriv zp w * (w-p)^zo" have po:"zo = Suc (zo - Suc 0) " using \zo>0\ by auto - have "(zp has_field_derivative (deriv zp w)) (at w)" + have "(zp has_field_derivative (deriv zp w)) (at w)" using DERIV_deriv_iff_has_field_derivative pp_holo by (meson Topology_Euclidean_Space.open_ball \w \ ball p r\ ball_subset_cball holomorphic_derivI holomorphic_on_subset) - then show "(f' has_field_derivative der) (at w)" + then show "(f' has_field_derivative der) (at w)" using \w\p\ \zo>0\ unfolding der_def f'_def by (auto intro!: derivative_eq_intros simp add:field_simps) qed @@ -3377,16 +3377,16 @@ apply (auto simp add:field_simps) by (metis Suc_diff_Suc minus_nat.diff_0 power_Suc) qed - then have "cont_zero ff p e2" unfolding cont_zero_def + then have "cont_zero ff p e2" unfolding cont_zero_def proof (elim has_contour_integral_eq) fix w assume "w \ path_image (circlepath p e2)" then have "w\ball p r" and "w\p" unfolding e2_def using \r>0\ by auto - have "deriv f' w = deriv f w" + have "deriv f' w = deriv f w" proof (rule complex_derivative_transform_within_open[where s="ball p r - {p}"]) show "f' holomorphic_on ball p r - {p}" unfolding f'_def using pp_holo by (auto intro!: holomorphic_intros) - next - have "ball p e1 - {p} \ s - poles" + next + have "ball p e1 - {p} \ s - poles" using avoid_def ball_subset_cball e1_avoid by auto then have "ball p r - {p} \ s - poles" using \cball p r \ ball p e1\ using ball_subset_cball by blast @@ -3398,37 +3398,37 @@ show "w \ ball p r - {p}" using \w\ball p r\ \w\p\ by auto next fix x assume "x \ ball p r - {p}" - then show "f' x = f x" + then show "f' x = f x" using pp_po unfolding f'_def by auto qed - moreover have " f' w = f w " + moreover have " f' w = f w " using \w \ ball p r\ ball_subset_cball subset_iff pp_po unfolding f'_def by auto ultimately show "ff' w = ff w" unfolding ff'_def ff_def by simp qed - moreover have "cball p e2 \ ball p e1" + moreover have "cball p e2 \ ball p e1" using \0 < r\ \cball p r \ ball p e1\ e2_def by auto ultimately show ?thesis using \e2>0\ by auto qed then obtain e3 where e3:"p\zeros \ e3>0 \ cball p e3 \ ball p e1 \ cont_zero ff p e3" - by auto + by auto define e4 where "e4 \ if p\poles then e2 else if p\zeros then e3 else e1" have "e4>0" using e2 e3 \e1>0\ unfolding e4_def by auto moreover have "avoid p e4" using e2 e3 \e1>0\ e1_avoid unfolding e4_def avoid_def by auto - moreover have "p\poles \ cont_pole ff p e4" and "p\zeros \ cont_zero ff p e4" + moreover have "p\poles \ cont_pole ff p e4" and "p\zeros \ cont_zero ff p e4" by (auto simp add: e2 e3 e4_def zeros_def) ultimately show ?thesis by auto qed - then obtain get_e where get_e:"\p\s. get_e p>0 \ avoid p (get_e p) + then obtain get_e where get_e:"\p\s. get_e p>0 \ avoid p (get_e p) \ (p\poles \ cont_pole ff p (get_e p)) \ (p\zeros \ cont_zero ff p (get_e p))" by metis define cont where "cont \ \p. contour_integral (circlepath p (get_e p)) ff" define w where "w \ \p. winding_number g p" have "contour_integral g ff = (\p\zeros \ poles. w p * cont p)" unfolding cont_def w_def - proof (rule Cauchy_theorem_singularities[OF \open s\ \connected s\ finite _ \valid_path g\ loop + proof (rule Cauchy_theorem_singularities[OF \open s\ \connected s\ finite _ \valid_path g\ loop path_img homo]) - have "open (s - (zeros \ poles))" using open_Diff[OF _ finite_imp_closed[OF finite]] \open s\ by auto + have "open (s - (zeros \ poles))" using open_Diff[OF _ finite_imp_closed[OF finite]] \open s\ by auto then show "ff holomorphic_on s - (zeros \ poles)" unfolding ff_def using f_holo h_holo by (auto intro!: holomorphic_intros simp add:zeros_def) next @@ -3436,12 +3436,12 @@ using get_e using avoid_def by blast qed also have "... = (\p\zeros. w p * cont p) + (\p\poles. w p * cont p)" - using finite + using finite apply (subst setsum.union_disjoint) by (auto simp add:zeros_def) also have "... = c * ((\p\zeros. w p * h p * zorder f p) - (\p\poles. w p * h p * porder f p))" proof - - have "(\p\zeros. w p * cont p) = (\p\zeros. c * w p * h p * zorder f p)" + have "(\p\zeros. w p * cont p) = (\p\zeros. c * w p * h p * zorder f p)" proof (rule setsum.cong[of zeros zeros,simplified]) fix p assume "p \ zeros" show "w p * cont p = c * w p * h p * (zorder f p)" @@ -3449,7 +3449,7 @@ assume "p \ s" have "cont p = c * h p * (zorder f p)" unfolding cont_def apply (rule contour_integral_unique) - using get_e \p\s\ \p\zeros\ unfolding cont_zero_def + using get_e \p\s\ \p\zeros\ unfolding cont_zero_def by (metis mult.assoc mult.commute) thus ?thesis by auto next @@ -3458,10 +3458,10 @@ then show ?thesis by auto qed qed - then have "(\p\zeros. w p * cont p) = c * (\p\zeros. w p * h p * zorder f p)" + then have "(\p\zeros. w p * cont p) = c * (\p\zeros. w p * h p * zorder f p)" apply (subst setsum_right_distrib) by (simp add:algebra_simps) - moreover have "(\p\poles. w p * cont p) = (\p\poles. - c * w p * h p * porder f p)" + moreover have "(\p\poles. w p * cont p) = (\p\poles. - c * w p * h p * porder f p)" proof (rule setsum.cong[of poles poles,simplified]) fix p assume "p \ poles" show "w p * cont p = - c * w p * h p * (porder f p)" @@ -3469,7 +3469,7 @@ assume "p \ s" have "cont p = - c * h p * (porder f p)" unfolding cont_def apply (rule contour_integral_unique) - using get_e \p\s\ \p\poles\ unfolding cont_pole_def + using get_e \p\s\ \p\poles\ unfolding cont_pole_def by (metis mult.assoc mult.commute) thus ?thesis by auto next @@ -3478,11 +3478,11 @@ then show ?thesis by auto qed qed - then have "(\p\poles. w p * cont p) = - c * (\p\poles. w p * h p * porder f p)" + then have "(\p\poles. w p * cont p) = - c * (\p\poles. w p * h p * porder f p)" apply (subst setsum_right_distrib) by (simp add:algebra_simps) ultimately show ?thesis by (simp add: right_diff_distrib) - qed + qed finally show ?thesis unfolding w_def ff_def c_def by auto qed @@ -3493,59 +3493,59 @@ defines "fg\(\p. f p+ g p)" defines "zeros_fg\{p. fg p =0}" and "zeros_f\{p. f p=0}" assumes - "open s" and "connected s" and - "finite zeros_fg" and + "open s" and "connected s" and + "finite zeros_fg" and "finite zeros_f" and - f_holo:"f holomorphic_on s" and + f_holo:"f holomorphic_on s" and g_holo:"g holomorphic_on s" and - "valid_path \" and + "valid_path \" and loop:"pathfinish \ = pathstart \" and path_img:"path_image \ \ s " and - path_less:"\z\path_image \. cmod(f z) > cmod(g z)" and + path_less:"\z\path_image \. cmod(f z) > cmod(g z)" and homo:"\z. (z \ s) \ winding_number \ z = 0" - shows "(\p\zeros_fg. winding_number \ p * zorder fg p) + shows "(\p\zeros_fg. winding_number \ p * zorder fg p) = (\p\zeros_f. winding_number \ p * zorder f p)" proof - have path_fg:"path_image \ \ s - zeros_fg" proof - - have False when "z\path_image \" and "f z + g z=0" for z + have False when "z\path_image \" and "f z + g z=0" for z proof - have "cmod (f z) > cmod (g z)" using \z\path_image \\ path_less by auto moreover have "f z = - g z" using \f z + g z =0\ by (simp add: eq_neg_iff_add_eq_0) then have "cmod (f z) = cmod (g z)" by auto ultimately show False by auto qed - then show ?thesis unfolding zeros_fg_def fg_def using path_img by auto + then show ?thesis unfolding zeros_fg_def fg_def using path_img by auto qed have path_f:"path_image \ \ s - zeros_f" proof - - have False when "z\path_image \" and "f z =0" for z + have False when "z\path_image \" and "f z =0" for z proof - have "cmod (g z) < cmod (f z) " using \z\path_image \\ path_less by auto - then have "cmod (g z) < 0" using \f z=0\ by auto + then have "cmod (g z) < 0" using \f z=0\ by auto then show False by auto qed - then show ?thesis unfolding zeros_f_def using path_img by auto + then show ?thesis unfolding zeros_f_def using path_img by auto qed define w where "w \ \p. winding_number \ p" define c where "c \ 2 * complex_of_real pi * \" - define h where "h \ \p. g p / f p + 1" + define h where "h \ \p. g p / f p + 1" obtain spikes where "finite spikes" and spikes: "\x\{0..1} - spikes. \ differentiable at x" using \valid_path \\ by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - have h_contour:"((\x. deriv h x / h x) has_contour_integral 0) \" + have h_contour:"((\x. deriv h x / h x) has_contour_integral 0) \" proof - - have outside_img:"0 \ outside (path_image (h o \))" + have outside_img:"0 \ outside (path_image (h o \))" proof - have "h p \ ball 1 1" when "p\path_image \" for p proof - - have "cmod (g p/f p) <1" using path_less[rule_format,OF that] + have "cmod (g p/f p) <1" using path_less[rule_format,OF that] apply (cases "cmod (f p) = 0") by (auto simp add: norm_divide) then show ?thesis unfolding h_def by (auto simp add:dist_complex_def) qed - then have "path_image (h o \) \ ball 1 1" + then have "path_image (h o \) \ ball 1 1" by (simp add: image_subset_iff path_image_compose) moreover have " (0::complex) \ ball 1 1" by (simp add: dist_norm) ultimately show "?thesis" @@ -3553,67 +3553,67 @@ qed have valid_h:"valid_path (h \ \)" proof (rule valid_path_compose_holomorphic[OF \valid_path \\ _ _ path_f]) - show "h holomorphic_on s - zeros_f" + show "h holomorphic_on s - zeros_f" unfolding h_def using f_holo g_holo by (auto intro!: holomorphic_intros simp add:zeros_f_def) next show "open (s - zeros_f)" using \finite zeros_f\ \open s\ finite_imp_closed by auto qed - have "((\z. 1/z) has_contour_integral 0) (h \ \)" + have "((\z. 1/z) has_contour_integral 0) (h \ \)" proof - have "0 \ path_image (h \ \)" using outside_img by (simp add: outside_def) then have "((\z. 1/z) has_contour_integral c * winding_number (h \ \) 0) (h \ \)" - using has_contour_integral_winding_number[of "h o \" 0,simplified] valid_h + using has_contour_integral_winding_number[of "h o \" 0,simplified] valid_h unfolding c_def by auto - moreover have "winding_number (h o \) 0 = 0" + moreover have "winding_number (h o \) 0 = 0" proof - have "0 \ outside (path_image (h \ \))" using outside_img . - moreover have "path (h o \)" + moreover have "path (h o \)" using valid_h by (simp add: valid_path_imp_path) - moreover have "pathfinish (h o \) = pathstart (h o \)" + moreover have "pathfinish (h o \) = pathstart (h o \)" by (simp add: loop pathfinish_compose pathstart_compose) ultimately show ?thesis using winding_number_zero_in_outside by auto qed ultimately show ?thesis by auto qed moreover have "vector_derivative (h \ \) (at x) = vector_derivative \ (at x) * deriv h (\ x)" - when "x\{0..1} - spikes" for x + when "x\{0..1} - spikes" for x proof (rule vector_derivative_chain_at_general) show "\ differentiable at x" using that \valid_path \\ spikes by auto next define der where "der \ \p. (deriv g p * f p - g p * deriv f p)/(f p * f p)" define t where "t \ \ x" - have "f t\0" unfolding zeros_f_def t_def + have "f t\0" unfolding zeros_f_def t_def by (metis DiffD1 image_eqI norm_not_less_zero norm_zero path_defs(4) path_less that) - moreover have "t\s" + moreover have "t\s" using contra_subsetD path_image_def path_fg t_def that by fastforce ultimately have "(h has_field_derivative der t) (at t)" unfolding h_def der_def using g_holo f_holo \open s\ by (auto intro!: holomorphic_derivI derivative_eq_intros ) then show "\g'. (h has_field_derivative g') (at (\ x))" unfolding t_def by auto qed - then have " (op / 1 has_contour_integral 0) (h \ \) - = ((\x. deriv h x / h x) has_contour_integral 0) \" + then have " (op / 1 has_contour_integral 0) (h \ \) + = ((\x. deriv h x / h x) has_contour_integral 0) \" unfolding has_contour_integral apply (intro has_integral_spike_eq[OF negligible_finite, OF \finite spikes\]) by auto ultimately show ?thesis by auto qed - then have "contour_integral \ (\x. deriv h x / h x) = 0" + then have "contour_integral \ (\x. deriv h x / h x) = 0" using contour_integral_unique by simp - moreover have "contour_integral \ (\x. deriv fg x / fg x) = contour_integral \ (\x. deriv f x / f x) - + contour_integral \ (\p. deriv h p / h p)" + moreover have "contour_integral \ (\x. deriv fg x / fg x) = contour_integral \ (\x. deriv f x / f x) + + contour_integral \ (\p. deriv h p / h p)" proof - - have "(\p. deriv f p / f p) contour_integrable_on \" + have "(\p. deriv f p / f p) contour_integrable_on \" proof (rule contour_integrable_holomorphic_simple[OF _ _ \valid_path \\ path_f]) - show "open (s - zeros_f)" using finite_imp_closed[OF \finite zeros_f\] \open s\ + show "open (s - zeros_f)" using finite_imp_closed[OF \finite zeros_f\] \open s\ by auto then show "(\p. deriv f p / f p) holomorphic_on s - zeros_f" using f_holo by (auto intro!: holomorphic_intros simp add:zeros_f_def) qed - moreover have "(\p. deriv h p / h p) contour_integrable_on \" + moreover have "(\p. deriv h p / h p) contour_integrable_on \" using h_contour by (simp add: has_contour_integral_integrable) ultimately have "contour_integral \ (\x. deriv f x / f x + deriv h x / h x) = @@ -3621,16 +3621,16 @@ using contour_integral_add[of "(\p. deriv f p / f p)" \ "(\p. deriv h p / h p)" ] by auto moreover have "deriv fg p / fg p = deriv f p / f p + deriv h p / h p" - when "p\ path_image \" for p + when "p\ path_image \" for p proof - have "fg p\0" and "f p\0" using path_f path_fg that unfolding zeros_f_def zeros_fg_def by auto - have "h p\0" + have "h p\0" proof (rule ccontr) assume "\ h p \ 0" then have "g p / f p= -1" unfolding h_def by (simp add: add_eq_0_iff2) then have "cmod (g p/f p) = 1" by auto - moreover have "cmod (g p/f p) <1" using path_less[rule_format,OF that] + moreover have "cmod (g p/f p) <1" using path_less[rule_format,OF that] apply (cases "cmod (f p) = 0") by (auto simp add: norm_divide) ultimately show False by auto @@ -3640,7 +3640,7 @@ by auto have der_h:"deriv h p = (deriv g p * f p - g p * deriv f p)/(f p * f p)" proof - - define der where "der \ \p. (deriv g p * f p - g p * deriv f p)/(f p * f p)" + define der where "der \ \p. (deriv g p * f p - g p * deriv f p)/(f p * f p)" have "p\s" using path_img that by auto then have "(h has_field_derivative der p) (at p)" unfolding h_def der_def using g_holo f_holo \open s\ \f p\0\ @@ -3648,17 +3648,17 @@ then show ?thesis unfolding der_def using DERIV_imp_deriv by auto qed show ?thesis - apply (simp only:der_fg der_h) + apply (simp only:der_fg der_h) apply (auto simp add:field_simps \h p\0\ \f p\0\ \fg p\0\) - by (auto simp add:field_simps h_def \f p\0\ fg_def) + by (auto simp add:field_simps h_def \f p\0\ fg_def) qed - then have "contour_integral \ (\p. deriv fg p / fg p) + then have "contour_integral \ (\p. deriv fg p / fg p) = contour_integral \ (\p. deriv f p / f p + deriv h p / h p)" - by (elim contour_integral_eq) + by (elim contour_integral_eq) ultimately show ?thesis by auto qed moreover have "contour_integral \ (\x. deriv fg x / fg x) = c * (\p\zeros_fg. w p * zorder fg p)" - unfolding c_def zeros_fg_def w_def + unfolding c_def zeros_fg_def w_def proof (rule argument_principle[OF \open s\ \connected s\ _ _ \valid_path \\ loop _ homo , of _ "{}" "\_. 1",simplified]) show "fg holomorphic_on s" unfolding fg_def using f_holo g_holo holomorphic_on_add by auto @@ -3666,13 +3666,13 @@ show " finite {p. fg p = 0}" using \finite zeros_fg\ unfolding zeros_fg_def . qed moreover have "contour_integral \ (\x. deriv f x / f x) = c * (\p\zeros_f. w p * zorder f p)" - unfolding c_def zeros_f_def w_def + unfolding c_def zeros_f_def w_def proof (rule argument_principle[OF \open s\ \connected s\ _ _ \valid_path \\ loop _ homo , of _ "{}" "\_. 1",simplified]) show "f holomorphic_on s" using f_holo g_holo holomorphic_on_add by auto show "path_image \ \ s - {p. f p = 0}" using path_f unfolding zeros_f_def . show " finite {p. f p = 0}" using \finite zeros_f\ unfolding zeros_f_def . - qed + qed ultimately have " c* (\p\zeros_fg. w p * (zorder fg p)) = c* (\p\zeros_f. w p * (zorder f p))" by auto then show ?thesis unfolding c_def using w_def by auto diff -r bbcb05504fdc -r bd218a9320b5 src/HOL/Multivariate_Analysis/Continuous_Extension.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Multivariate_Analysis/Continuous_Extension.thy Thu Aug 04 19:36:31 2016 +0200 @@ -0,0 +1,547 @@ +(* Title: HOL/Multivariate_Analysis/Extension.thy + Authors: LC Paulson, based on material from HOL Light +*) + +section \Continuous extensions of functions: Urysohn's lemma, Dugundji extension theorem, Tietze\ + +theory Continuous_Extension +imports Convex_Euclidean_Space +begin + +subsection\Partitions of unity subordinate to locally finite open coverings\ + +text\A difference from HOL Light: all summations over infinite sets equal zero, + so the "support" must be made explicit in the summation below!\ + +proposition subordinate_partition_of_unity: + fixes S :: "'a :: euclidean_space set" + assumes "S \ \\" and opC: "\T. T \ \ \ open T" + and fin: "\x. x \ S \ \V. open V \ x \ V \ finite {U \ \. U \ V \ {}}" + obtains F :: "['a set, 'a] \ real" + where "\U. U \ \ \ continuous_on S (F U) \ (\x \ S. 0 \ F U x)" + and "\x U. \U \ \; x \ S; x \ U\ \ F U x = 0" + and "\x. x \ S \ supp_setsum (\W. F W x) \ = 1" + and "\x. x \ S \ \V. open V \ x \ V \ finite {U \ \. \x\V. F U x \ 0}" +proof (cases "\W. W \ \ \ S \ W") + case True + then obtain W where "W \ \" "S \ W" by metis + then show ?thesis + apply (rule_tac F = "\V x. if V = W then 1 else 0" in that) + apply (auto simp: continuous_on_const supp_setsum_def support_on_def) + done +next + case False + have nonneg: "0 \ supp_setsum (\V. setdist {x} (S - V)) \" for x + by (simp add: supp_setsum_def setsum_nonneg) + have sd_pos: "0 < setdist {x} (S - V)" if "V \ \" "x \ S" "x \ V" for V x + proof - + have "closedin (subtopology euclidean S) (S - V)" + by (simp add: Diff_Diff_Int Diff_subset closedin_def opC openin_open_Int \V \ \\) + with that False setdist_eq_0_closedin [of S "S-V" x] setdist_pos_le [of "{x}" "S - V"] + show ?thesis + by (simp add: order_class.order.order_iff_strict) + qed + have ss_pos: "0 < supp_setsum (\V. setdist {x} (S - V)) \" if "x \ S" for x + proof - + obtain U where "U \ \" "x \ U" using \x \ S\ \S \ \\\ + by blast + obtain V where "open V" "x \ V" "finite {U \ \. U \ V \ {}}" + using \x \ S\ fin by blast + then have *: "finite {A \ \. \ S \ A \ x \ closure (S - A)}" + using closure_def that by (blast intro: rev_finite_subset) + have "x \ closure (S - U)" + by (metis \U \ \\ \x \ U\ less_irrefl sd_pos setdist_eq_0_sing_1 that) + then show ?thesis + apply (simp add: setdist_eq_0_sing_1 supp_setsum_def support_on_def) + apply (rule ordered_comm_monoid_add_class.setsum_pos2 [OF *, of U]) + using \U \ \\ \x \ U\ False + apply (auto simp: setdist_pos_le sd_pos that) + done + qed + define F where + "F \ \W x. if x \ S then setdist {x} (S - W) / supp_setsum (\V. setdist {x} (S - V)) \ + else 0" + show ?thesis + proof (rule_tac F = F in that) + have "continuous_on S (F U)" if "U \ \" for U + proof - + have *: "continuous_on S (\x. supp_setsum (\V. setdist {x} (S - V)) \)" + proof (clarsimp simp add: continuous_on_eq_continuous_within) + fix x assume "x \ S" + then obtain X where "open X" and x: "x \ S \ X" and finX: "finite {U \ \. U \ X \ {}}" + using assms by blast + then have OSX: "openin (subtopology euclidean S) (S \ X)" by blast + have sumeq: "\x. x \ S \ X \ + (\V | V \ \ \ V \ X \ {}. setdist {x} (S - V)) + = supp_setsum (\V. setdist {x} (S - V)) \" + apply (simp add: supp_setsum_def) + apply (rule setsum.mono_neutral_right [OF finX]) + apply (auto simp: setdist_eq_0_sing_1 support_on_def subset_iff) + apply (meson DiffI closure_subset disjoint_iff_not_equal subsetCE) + done + show "continuous (at x within S) (\x. supp_setsum (\V. setdist {x} (S - V)) \)" + apply (rule continuous_transform_within_openin + [where f = "\x. (setsum (\V. setdist {x} (S - V)) {V \ \. V \ X \ {}})" + and S ="S \ X"]) + apply (rule continuous_intros continuous_at_setdist continuous_at_imp_continuous_at_within OSX x)+ + apply (simp add: sumeq) + done + qed + show ?thesis + apply (simp add: F_def) + apply (rule continuous_intros *)+ + using ss_pos apply force + done + qed + moreover have "\U \ \; x \ S\ \ 0 \ F U x" for U x + using nonneg [of x] by (simp add: F_def divide_simps setdist_pos_le) + ultimately show "\U. U \ \ \ continuous_on S (F U) \ (\x\S. 0 \ F U x)" + by metis + next + show "\x U. \U \ \; x \ S; x \ U\ \ F U x = 0" + by (simp add: setdist_eq_0_sing_1 closure_def F_def) + next + show "supp_setsum (\W. F W x) \ = 1" if "x \ S" for x + using that ss_pos [OF that] + by (simp add: F_def divide_simps supp_setsum_divide_distrib [symmetric]) + next + show "\V. open V \ x \ V \ finite {U \ \. \x\V. F U x \ 0}" if "x \ S" for x + using fin [OF that] that + by (fastforce simp: setdist_eq_0_sing_1 closure_def F_def elim!: rev_finite_subset) + qed +qed + + +subsection\Urysohn's lemma (for Euclidean spaces, where the proof is easy using distances)\ + +lemma Urysohn_both_ne: + assumes US: "closedin (subtopology euclidean U) S" + and UT: "closedin (subtopology euclidean U) T" + and "S \ T = {}" "S \ {}" "T \ {}" "a \ b" + obtains f :: "'a::euclidean_space \ 'b::real_normed_vector" + where "continuous_on U f" + "\x. x \ U \ f x \ closed_segment a b" + "\x. x \ U \ (f x = a \ x \ S)" + "\x. x \ U \ (f x = b \ x \ T)" +proof - + have S0: "\x. x \ U \ setdist {x} S = 0 \ x \ S" + using \S \ {}\ US setdist_eq_0_closedin by auto + have T0: "\x. x \ U \ setdist {x} T = 0 \ x \ T" + using \T \ {}\ UT setdist_eq_0_closedin by auto + have sdpos: "0 < setdist {x} S + setdist {x} T" if "x \ U" for x + proof - + have "~ (setdist {x} S = 0 \ setdist {x} T = 0)" + using assms by (metis IntI empty_iff setdist_eq_0_closedin that) + then show ?thesis + by (metis add.left_neutral add.right_neutral add_pos_pos linorder_neqE_linordered_idom not_le setdist_pos_le) + qed + define f where "f \ \x. a + (setdist {x} S / (setdist {x} S + setdist {x} T)) *\<^sub>R (b - a)" + show ?thesis + proof (rule_tac f = f in that) + show "continuous_on U f" + using sdpos unfolding f_def + by (intro continuous_intros | force)+ + show "f x \ closed_segment a b" if "x \ U" for x + unfolding f_def + apply (simp add: closed_segment_def) + apply (rule_tac x="(setdist {x} S / (setdist {x} S + setdist {x} T))" in exI) + using sdpos that apply (simp add: algebra_simps) + done + show "\x. x \ U \ (f x = a \ x \ S)" + using S0 \a \ b\ f_def sdpos by force + show "(f x = b \ x \ T)" if "x \ U" for x + proof - + have "f x = b \ (setdist {x} S / (setdist {x} S + setdist {x} T)) = 1" + unfolding f_def + apply (rule iffI) + apply (metis \a \ b\ add_diff_cancel_left' eq_iff_diff_eq_0 pth_1 real_vector.scale_right_imp_eq, force) + done + also have "... \ setdist {x} T = 0 \ setdist {x} S \ 0" + using sdpos that + by (simp add: divide_simps) linarith + also have "... \ x \ T" + using \S \ {}\ \T \ {}\ \S \ T = {}\ that + by (force simp: S0 T0) + finally show ?thesis . + qed + qed +qed + +proposition Urysohn_local_strong: + assumes US: "closedin (subtopology euclidean U) S" + and UT: "closedin (subtopology euclidean U) T" + and "S \ T = {}" "a \ b" + obtains f :: "'a::euclidean_space \ 'b::euclidean_space" + where "continuous_on U f" + "\x. x \ U \ f x \ closed_segment a b" + "\x. x \ U \ (f x = a \ x \ S)" + "\x. x \ U \ (f x = b \ x \ T)" +proof (cases "S = {}") + case True show ?thesis + proof (cases "T = {}") + case True show ?thesis + proof (rule_tac f = "\x. midpoint a b" in that) + show "continuous_on U (\x. midpoint a b)" + by (intro continuous_intros) + show "midpoint a b \ closed_segment a b" + using csegment_midpoint_subset by blast + show "(midpoint a b = a) = (x \ S)" for x + using \S = {}\ \a \ b\ by simp + show "(midpoint a b = b) = (x \ T)" for x + using \T = {}\ \a \ b\ by simp + qed + next + case False + show ?thesis + proof (cases "T = U") + case True with \S = {}\ \a \ b\ show ?thesis + by (rule_tac f = "\x. b" in that) (auto simp: continuous_on_const) + next + case False + with UT closedin_subset obtain c where c: "c \ U" "c \ T" + by fastforce + obtain f where f: "continuous_on U f" + "\x. x \ U \ f x \ closed_segment (midpoint a b) b" + "\x. x \ U \ (f x = midpoint a b \ x = c)" + "\x. x \ U \ (f x = b \ x \ T)" + apply (rule Urysohn_both_ne [of U "{c}" T "midpoint a b" "b"]) + using c \T \ {}\ assms apply simp_all + done + show ?thesis + apply (rule_tac f=f in that) + using \S = {}\ \T \ {}\ f csegment_midpoint_subset notin_segment_midpoint [OF \a \ b\] + apply force+ + done + qed + qed +next + case False + show ?thesis + proof (cases "T = {}") + case True show ?thesis + proof (cases "S = U") + case True with \T = {}\ \a \ b\ show ?thesis + by (rule_tac f = "\x. a" in that) (auto simp: continuous_on_const) + next + case False + with US closedin_subset obtain c where c: "c \ U" "c \ S" + by fastforce + obtain f where f: "continuous_on U f" + "\x. x \ U \ f x \ closed_segment a (midpoint a b)" + "\x. x \ U \ (f x = midpoint a b \ x = c)" + "\x. x \ U \ (f x = a \ x \ S)" + apply (rule Urysohn_both_ne [of U S "{c}" a "midpoint a b"]) + using c \S \ {}\ assms apply simp_all + apply (metis midpoint_eq_endpoint) + done + show ?thesis + apply (rule_tac f=f in that) + using \S \ {}\ \T = {}\ f \a \ b\ + apply simp_all + apply (metis (no_types) closed_segment_commute csegment_midpoint_subset midpoint_sym subset_iff) + apply (metis closed_segment_commute midpoint_sym notin_segment_midpoint) + done + qed + next + case False + show ?thesis + using Urysohn_both_ne [OF US UT \S \ T = {}\ \S \ {}\ \T \ {}\ \a \ b\] that + by blast + qed +qed + +lemma Urysohn_local: + assumes US: "closedin (subtopology euclidean U) S" + and UT: "closedin (subtopology euclidean U) T" + and "S \ T = {}" + obtains f :: "'a::euclidean_space \ 'b::euclidean_space" + where "continuous_on U f" + "\x. x \ U \ f x \ closed_segment a b" + "\x. x \ S \ f x = a" + "\x. x \ T \ f x = b" +proof (cases "a = b") + case True then show ?thesis + by (rule_tac f = "\x. b" in that) (auto simp: continuous_on_const) +next + case False + then show ?thesis + apply (rule Urysohn_local_strong [OF assms]) + apply (erule that, assumption) + apply (meson US closedin_singleton closedin_trans) + apply (meson UT closedin_singleton closedin_trans) + done +qed + +lemma Urysohn_strong: + assumes US: "closed S" + and UT: "closed T" + and "S \ T = {}" "a \ b" + obtains f :: "'a::euclidean_space \ 'b::euclidean_space" + where "continuous_on UNIV f" + "\x. f x \ closed_segment a b" + "\x. f x = a \ x \ S" + "\x. f x = b \ x \ T" +apply (rule Urysohn_local_strong [of UNIV S T]) +using assms +apply (auto simp: closed_closedin) +done + +proposition Urysohn: + assumes US: "closed S" + and UT: "closed T" + and "S \ T = {}" + obtains f :: "'a::euclidean_space \ 'b::euclidean_space" + where "continuous_on UNIV f" + "\x. f x \ closed_segment a b" + "\x. x \ S \ f x = a" + "\x. x \ T \ f x = b" +apply (rule Urysohn_local [of UNIV S T a b]) +using assms +apply (auto simp: closed_closedin) +done + + +subsection\ The Dugundji extension theorem, and Tietze variants as corollaries.\ + +text\J. Dugundji. An extension of Tietze's theorem. Pacific J. Math. Volume 1, Number 3 (1951), 353-367. +http://projecteuclid.org/euclid.pjm/1103052106\ + +theorem Dugundji: + fixes f :: "'a::euclidean_space \ 'b::real_inner" + assumes "convex C" "C \ {}" + and cloin: "closedin (subtopology euclidean U) S" + and contf: "continuous_on S f" and "f ` S \ C" + obtains g where "continuous_on U g" "g ` U \ C" + "\x. x \ S \ g x = f x" +proof (cases "S = {}") + case True then show thesis + apply (rule_tac g="\x. @y. y \ C" in that) + apply (rule continuous_intros) + apply (meson all_not_in_conv \C \ {}\ image_subsetI someI_ex, simp) + done +next + case False + then have sd_pos: "\x. \x \ U; x \ S\ \ 0 < setdist {x} S" + using setdist_eq_0_closedin [OF cloin] le_less setdist_pos_le by fastforce + define \ where "\ = {ball x (setdist {x} S / 2) |x. x \ U - S}" + have [simp]: "\T. T \ \ \ open T" + by (auto simp: \_def) + have USS: "U - S \ \\" + by (auto simp: sd_pos \_def) + obtain \ where USsub: "U - S \ \\" + and nbrhd: "\U. U \ \ \ open U \ (\T. T \ \ \ U \ T)" + and fin: "\x. x \ U - S + \ \V. open V \ x \ V \ finite {U. U \ \ \ U \ V \ {}}" + using paracompact [OF USS] by auto + have "\v a. v \ U \ v \ S \ a \ S \ + T \ ball v (setdist {v} S / 2) \ + dist v a \ 2 * setdist {v} S" if "T \ \" for T + proof - + obtain v where v: "T \ ball v (setdist {v} S / 2)" "v \ U" "v \ S" + using \T \ \\ nbrhd by (force simp: \_def) + then obtain a where "a \ S" "dist v a < 2 * setdist {v} S" + using setdist_ltE [of "{v}" S "2 * setdist {v} S"] + using False sd_pos by force + with v show ?thesis + apply (rule_tac x=v in exI) + apply (rule_tac x=a in exI, auto) + done + qed + then obtain \ \ where + VA: "\T. T \ \ \ \ T \ U \ \ T \ S \ \ T \ S \ + T \ ball (\ T) (setdist {\ T} S / 2) \ + dist (\ T) (\ T) \ 2 * setdist {\ T} S" + by metis + have sdle: "setdist {\ T} S \ 2 * setdist {v} S" if "T \ \" "v \ T" for T v + using setdist_Lipschitz [of "\ T" S v] VA [OF \T \ \\] \v \ T\ by auto + have d6: "dist a (\ T) \ 6 * dist a v" if "T \ \" "v \ T" "a \ S" for T v a + proof - + have "dist (\ T) v < setdist {\ T} S / 2" + using that VA mem_ball by blast + also have "... \ setdist {v} S" + using sdle [OF \T \ \\ \v \ T\] by simp + also have vS: "setdist {v} S \ dist a v" + by (simp add: setdist_le_dist setdist_sym \a \ S\) + finally have VTV: "dist (\ T) v < dist a v" . + have VTS: "setdist {\ T} S \ 2 * dist a v" + using sdle that vS by force + have "dist a (\ T) \ dist a v + dist v (\ T) + dist (\ T) (\ T)" + by (metis add.commute add_le_cancel_left dist_commute dist_triangle2 dist_triangle_le) + also have "... \ dist a v + dist a v + dist (\ T) (\ T)" + using VTV by (simp add: dist_commute) + also have "... \ 2 * dist a v + 2 * setdist {\ T} S" + using VA [OF \T \ \\] by auto + finally show ?thesis + using VTS by linarith + qed + obtain H :: "['a set, 'a] \ real" + where Hcont: "\Z. Z \ \ \ continuous_on (U-S) (H Z)" + and Hge0: "\Z x. \Z \ \; x \ U-S\ \ 0 \ H Z x" + and Heq0: "\x Z. \Z \ \; x \ U-S; x \ Z\ \ H Z x = 0" + and H1: "\x. x \ U-S \ supp_setsum (\W. H W x) \ = 1" + and Hfin: "\x. x \ U-S \ \V. open V \ x \ V \ finite {U \ \. \x\V. H U x \ 0}" + apply (rule subordinate_partition_of_unity [OF USsub _ fin]) + using nbrhd by auto + define g where "g \ \x. if x \ S then f x else supp_setsum (\T. H T x *\<^sub>R f(\ T)) \" + show ?thesis + proof (rule that) + show "continuous_on U g" + proof (clarsimp simp: continuous_on_eq_continuous_within) + fix a assume "a \ U" + show "continuous (at a within U) g" + proof (cases "a \ S") + case True show ?thesis + proof (clarsimp simp add: continuous_within_topological) + fix W + assume "open W" "g a \ W" + then obtain e where "0 < e" and e: "ball (f a) e \ W" + using openE True g_def by auto + have "continuous (at a within S) f" + using True contf continuous_on_eq_continuous_within by blast + then obtain d where "0 < d" + and d: "\x. \x \ S; dist x a < d\ \ dist (f x) (f a) < e" + using continuous_within_eps_delta \0 < e\ by force + have "g y \ ball (f a) e" if "y \ U" and y: "y \ ball a (d / 6)" for y + proof (cases "y \ S") + case True + then have "dist (f a) (f y) < e" + by (metis ball_divide_subset_numeral dist_commute in_mono mem_ball y d) + then show ?thesis + by (simp add: True g_def) + next + case False + have *: "dist (f (\ T)) (f a) < e" if "T \ \" "H T y \ 0" for T + proof - + have "y \ T" + using Heq0 that False \y \ U\ by blast + have "dist (\ T) a < d" + using d6 [OF \T \ \\ \y \ T\ \a \ S\] y + by (simp add: dist_commute mult.commute) + then show ?thesis + using VA [OF \T \ \\] by (auto simp: d) + qed + have "supp_setsum (\T. H T y *\<^sub>R f (\ T)) \ \ ball (f a) e" + apply (rule convex_supp_setsum [OF convex_ball]) + apply (simp_all add: False H1 Hge0 \y \ U\) + by (metis dist_commute *) + then show ?thesis + by (simp add: False g_def) + qed + then show "\A. open A \ a \ A \ (\y\U. y \ A \ g y \ W)" + apply (rule_tac x = "ball a (d / 6)" in exI) + using e \0 < d\ by fastforce + qed + next + case False + obtain N where N: "open N" "a \ N" + and finN: "finite {U \ \. \a\N. H U a \ 0}" + using Hfin False \a \ U\ by auto + have oUS: "openin (subtopology euclidean U) (U - S)" + using cloin by (simp add: openin_diff) + have HcontU: "continuous (at a within U) (H T)" if "T \ \" for T + using Hcont [OF \T \ \\] False \a \ U\ \T \ \\ + apply (simp add: continuous_on_eq_continuous_within continuous_within) + apply (rule Lim_transform_within_set) + using oUS + apply (force simp: eventually_at openin_contains_ball dist_commute dest!: bspec)+ + done + show ?thesis + proof (rule continuous_transform_within_openin [OF _ oUS]) + show "continuous (at a within U) (\x. supp_setsum (\T. H T x *\<^sub>R f (\ T)) \)" + proof (rule continuous_transform_within_openin) + show "continuous (at a within U) + (\x. \T\{U \ \. \x\N. H U x \ 0}. H T x *\<^sub>R f (\ T))" + by (force intro: continuous_intros HcontU)+ + next + show "openin (subtopology euclidean U) ((U - S) \ N)" + using N oUS openin_trans by blast + next + show "a \ (U - S) \ N" using False \a \ U\ N by blast + next + show "\x. x \ (U - S) \ N \ + (\T \ {U \ \. \x\N. H U x \ 0}. H T x *\<^sub>R f (\ T)) + = supp_setsum (\T. H T x *\<^sub>R f (\ T)) \" + by (auto simp: supp_setsum_def support_on_def + intro: setsum.mono_neutral_right [OF finN]) + qed + next + show "a \ U - S" using False \a \ U\ by blast + next + show "\x. x \ U - S \ supp_setsum (\T. H T x *\<^sub>R f (\ T)) \ = g x" + by (simp add: g_def) + qed + qed + qed + show "g ` U \ C" + using \f ` S \ C\ VA + by (fastforce simp: g_def Hge0 intro!: convex_supp_setsum [OF \convex C\] H1) + show "\x. x \ S \ g x = f x" + by (simp add: g_def) + qed +qed + + +corollary Tietze: + fixes f :: "'a::euclidean_space \ 'b::real_inner" + assumes "continuous_on S f" + and "closedin (subtopology euclidean U) S" + and "0 \ B" + and "\x. x \ S \ norm(f x) \ B" + obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" + "\x. x \ U \ norm(g x) \ B" +using assms +by (auto simp: image_subset_iff intro: Dugundji [of "cball 0 B" U S f]) + +corollary Tietze_closed_interval: + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes "continuous_on S f" + and "closedin (subtopology euclidean U) S" + and "cbox a b \ {}" + and "\x. x \ S \ f x \ cbox a b" + obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" + "\x. x \ U \ g x \ cbox a b" +apply (rule Dugundji [of "cbox a b" U S f]) +using assms by auto + +corollary Tietze_closed_interval_1: + fixes f :: "'a::euclidean_space \ real" + assumes "continuous_on S f" + and "closedin (subtopology euclidean U) S" + and "a \ b" + and "\x. x \ S \ f x \ cbox a b" + obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" + "\x. x \ U \ g x \ cbox a b" +apply (rule Dugundji [of "cbox a b" U S f]) +using assms by (auto simp: image_subset_iff) + +corollary Tietze_open_interval: + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes "continuous_on S f" + and "closedin (subtopology euclidean U) S" + and "box a b \ {}" + and "\x. x \ S \ f x \ box a b" + obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" + "\x. x \ U \ g x \ box a b" +apply (rule Dugundji [of "box a b" U S f]) +using assms by auto + +corollary Tietze_open_interval_1: + fixes f :: "'a::euclidean_space \ real" + assumes "continuous_on S f" + and "closedin (subtopology euclidean U) S" + and "a < b" + and no: "\x. x \ S \ f x \ box a b" + obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" + "\x. x \ U \ g x \ box a b" +apply (rule Dugundji [of "box a b" U S f]) +using assms by (auto simp: image_subset_iff) + +corollary Tietze_unbounded: + fixes f :: "'a::euclidean_space \ 'b::real_inner" + assumes "continuous_on S f" + and "closedin (subtopology euclidean U) S" + obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" +apply (rule Dugundji [of UNIV U S f]) +using assms by auto + +end diff -r bbcb05504fdc -r bd218a9320b5 src/HOL/Multivariate_Analysis/Extension.thy --- a/src/HOL/Multivariate_Analysis/Extension.thy Thu Aug 04 18:45:28 2016 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,547 +0,0 @@ -(* Title: HOL/Multivariate_Analysis/Extension.thy - Authors: LC Paulson, based on material from HOL Light -*) - -section \Continuous extensions of functions: Urysohn's lemma, Dugundji extension theorem, Tietze\ - -theory Extension -imports Convex_Euclidean_Space -begin - -subsection\Partitions of unity subordinate to locally finite open coverings\ - -text\A difference from HOL Light: all summations over infinite sets equal zero, - so the "support" must be made explicit in the summation below!\ - -proposition subordinate_partition_of_unity: - fixes S :: "'a :: euclidean_space set" - assumes "S \ \\" and opC: "\T. T \ \ \ open T" - and fin: "\x. x \ S \ \V. open V \ x \ V \ finite {U \ \. U \ V \ {}}" - obtains F :: "['a set, 'a] \ real" - where "\U. U \ \ \ continuous_on S (F U) \ (\x \ S. 0 \ F U x)" - and "\x U. \U \ \; x \ S; x \ U\ \ F U x = 0" - and "\x. x \ S \ supp_setsum (\W. F W x) \ = 1" - and "\x. x \ S \ \V. open V \ x \ V \ finite {U \ \. \x\V. F U x \ 0}" -proof (cases "\W. W \ \ \ S \ W") - case True - then obtain W where "W \ \" "S \ W" by metis - then show ?thesis - apply (rule_tac F = "\V x. if V = W then 1 else 0" in that) - apply (auto simp: continuous_on_const supp_setsum_def support_on_def) - done -next - case False - have nonneg: "0 \ supp_setsum (\V. setdist {x} (S - V)) \" for x - by (simp add: supp_setsum_def setsum_nonneg) - have sd_pos: "0 < setdist {x} (S - V)" if "V \ \" "x \ S" "x \ V" for V x - proof - - have "closedin (subtopology euclidean S) (S - V)" - by (simp add: Diff_Diff_Int Diff_subset closedin_def opC openin_open_Int \V \ \\) - with that False setdist_eq_0_closedin [of S "S-V" x] setdist_pos_le [of "{x}" "S - V"] - show ?thesis - by (simp add: order_class.order.order_iff_strict) - qed - have ss_pos: "0 < supp_setsum (\V. setdist {x} (S - V)) \" if "x \ S" for x - proof - - obtain U where "U \ \" "x \ U" using \x \ S\ \S \ \\\ - by blast - obtain V where "open V" "x \ V" "finite {U \ \. U \ V \ {}}" - using \x \ S\ fin by blast - then have *: "finite {A \ \. \ S \ A \ x \ closure (S - A)}" - using closure_def that by (blast intro: rev_finite_subset) - have "x \ closure (S - U)" - by (metis \U \ \\ \x \ U\ less_irrefl sd_pos setdist_eq_0_sing_1 that) - then show ?thesis - apply (simp add: setdist_eq_0_sing_1 supp_setsum_def support_on_def) - apply (rule ordered_comm_monoid_add_class.setsum_pos2 [OF *, of U]) - using \U \ \\ \x \ U\ False - apply (auto simp: setdist_pos_le sd_pos that) - done - qed - define F where - "F \ \W x. if x \ S then setdist {x} (S - W) / supp_setsum (\V. setdist {x} (S - V)) \ - else 0" - show ?thesis - proof (rule_tac F = F in that) - have "continuous_on S (F U)" if "U \ \" for U - proof - - have *: "continuous_on S (\x. supp_setsum (\V. setdist {x} (S - V)) \)" - proof (clarsimp simp add: continuous_on_eq_continuous_within) - fix x assume "x \ S" - then obtain X where "open X" and x: "x \ S \ X" and finX: "finite {U \ \. U \ X \ {}}" - using assms by blast - then have OSX: "openin (subtopology euclidean S) (S \ X)" by blast - have sumeq: "\x. x \ S \ X \ - (\V | V \ \ \ V \ X \ {}. setdist {x} (S - V)) - = supp_setsum (\V. setdist {x} (S - V)) \" - apply (simp add: supp_setsum_def) - apply (rule setsum.mono_neutral_right [OF finX]) - apply (auto simp: setdist_eq_0_sing_1 support_on_def subset_iff) - apply (meson DiffI closure_subset disjoint_iff_not_equal subsetCE) - done - show "continuous (at x within S) (\x. supp_setsum (\V. setdist {x} (S - V)) \)" - apply (rule continuous_transform_within_openin - [where f = "\x. (setsum (\V. setdist {x} (S - V)) {V \ \. V \ X \ {}})" - and S ="S \ X"]) - apply (rule continuous_intros continuous_at_setdist continuous_at_imp_continuous_at_within OSX x)+ - apply (simp add: sumeq) - done - qed - show ?thesis - apply (simp add: F_def) - apply (rule continuous_intros *)+ - using ss_pos apply force - done - qed - moreover have "\U \ \; x \ S\ \ 0 \ F U x" for U x - using nonneg [of x] by (simp add: F_def divide_simps setdist_pos_le) - ultimately show "\U. U \ \ \ continuous_on S (F U) \ (\x\S. 0 \ F U x)" - by metis - next - show "\x U. \U \ \; x \ S; x \ U\ \ F U x = 0" - by (simp add: setdist_eq_0_sing_1 closure_def F_def) - next - show "supp_setsum (\W. F W x) \ = 1" if "x \ S" for x - using that ss_pos [OF that] - by (simp add: F_def divide_simps supp_setsum_divide_distrib [symmetric]) - next - show "\V. open V \ x \ V \ finite {U \ \. \x\V. F U x \ 0}" if "x \ S" for x - using fin [OF that] that - by (fastforce simp: setdist_eq_0_sing_1 closure_def F_def elim!: rev_finite_subset) - qed -qed - - -subsection\Urysohn's lemma (for Euclidean spaces, where the proof is easy using distances)\ - -lemma Urysohn_both_ne: - assumes US: "closedin (subtopology euclidean U) S" - and UT: "closedin (subtopology euclidean U) T" - and "S \ T = {}" "S \ {}" "T \ {}" "a \ b" - obtains f :: "'a::euclidean_space \ 'b::real_normed_vector" - where "continuous_on U f" - "\x. x \ U \ f x \ closed_segment a b" - "\x. x \ U \ (f x = a \ x \ S)" - "\x. x \ U \ (f x = b \ x \ T)" -proof - - have S0: "\x. x \ U \ setdist {x} S = 0 \ x \ S" - using \S \ {}\ US setdist_eq_0_closedin by auto - have T0: "\x. x \ U \ setdist {x} T = 0 \ x \ T" - using \T \ {}\ UT setdist_eq_0_closedin by auto - have sdpos: "0 < setdist {x} S + setdist {x} T" if "x \ U" for x - proof - - have "~ (setdist {x} S = 0 \ setdist {x} T = 0)" - using assms by (metis IntI empty_iff setdist_eq_0_closedin that) - then show ?thesis - by (metis add.left_neutral add.right_neutral add_pos_pos linorder_neqE_linordered_idom not_le setdist_pos_le) - qed - define f where "f \ \x. a + (setdist {x} S / (setdist {x} S + setdist {x} T)) *\<^sub>R (b - a)" - show ?thesis - proof (rule_tac f = f in that) - show "continuous_on U f" - using sdpos unfolding f_def - by (intro continuous_intros | force)+ - show "f x \ closed_segment a b" if "x \ U" for x - unfolding f_def - apply (simp add: closed_segment_def) - apply (rule_tac x="(setdist {x} S / (setdist {x} S + setdist {x} T))" in exI) - using sdpos that apply (simp add: algebra_simps) - done - show "\x. x \ U \ (f x = a \ x \ S)" - using S0 \a \ b\ f_def sdpos by force - show "(f x = b \ x \ T)" if "x \ U" for x - proof - - have "f x = b \ (setdist {x} S / (setdist {x} S + setdist {x} T)) = 1" - unfolding f_def - apply (rule iffI) - apply (metis \a \ b\ add_diff_cancel_left' eq_iff_diff_eq_0 pth_1 real_vector.scale_right_imp_eq, force) - done - also have "... \ setdist {x} T = 0 \ setdist {x} S \ 0" - using sdpos that - by (simp add: divide_simps) linarith - also have "... \ x \ T" - using \S \ {}\ \T \ {}\ \S \ T = {}\ that - by (force simp: S0 T0) - finally show ?thesis . - qed - qed -qed - -proposition Urysohn_local_strong: - assumes US: "closedin (subtopology euclidean U) S" - and UT: "closedin (subtopology euclidean U) T" - and "S \ T = {}" "a \ b" - obtains f :: "'a::euclidean_space \ 'b::euclidean_space" - where "continuous_on U f" - "\x. x \ U \ f x \ closed_segment a b" - "\x. x \ U \ (f x = a \ x \ S)" - "\x. x \ U \ (f x = b \ x \ T)" -proof (cases "S = {}") - case True show ?thesis - proof (cases "T = {}") - case True show ?thesis - proof (rule_tac f = "\x. midpoint a b" in that) - show "continuous_on U (\x. midpoint a b)" - by (intro continuous_intros) - show "midpoint a b \ closed_segment a b" - using csegment_midpoint_subset by blast - show "(midpoint a b = a) = (x \ S)" for x - using \S = {}\ \a \ b\ by simp - show "(midpoint a b = b) = (x \ T)" for x - using \T = {}\ \a \ b\ by simp - qed - next - case False - show ?thesis - proof (cases "T = U") - case True with \S = {}\ \a \ b\ show ?thesis - by (rule_tac f = "\x. b" in that) (auto simp: continuous_on_const) - next - case False - with UT closedin_subset obtain c where c: "c \ U" "c \ T" - by fastforce - obtain f where f: "continuous_on U f" - "\x. x \ U \ f x \ closed_segment (midpoint a b) b" - "\x. x \ U \ (f x = midpoint a b \ x = c)" - "\x. x \ U \ (f x = b \ x \ T)" - apply (rule Urysohn_both_ne [of U "{c}" T "midpoint a b" "b"]) - using c \T \ {}\ assms apply simp_all - done - show ?thesis - apply (rule_tac f=f in that) - using \S = {}\ \T \ {}\ f csegment_midpoint_subset notin_segment_midpoint [OF \a \ b\] - apply force+ - done - qed - qed -next - case False - show ?thesis - proof (cases "T = {}") - case True show ?thesis - proof (cases "S = U") - case True with \T = {}\ \a \ b\ show ?thesis - by (rule_tac f = "\x. a" in that) (auto simp: continuous_on_const) - next - case False - with US closedin_subset obtain c where c: "c \ U" "c \ S" - by fastforce - obtain f where f: "continuous_on U f" - "\x. x \ U \ f x \ closed_segment a (midpoint a b)" - "\x. x \ U \ (f x = midpoint a b \ x = c)" - "\x. x \ U \ (f x = a \ x \ S)" - apply (rule Urysohn_both_ne [of U S "{c}" a "midpoint a b"]) - using c \S \ {}\ assms apply simp_all - apply (metis midpoint_eq_endpoint) - done - show ?thesis - apply (rule_tac f=f in that) - using \S \ {}\ \T = {}\ f \a \ b\ - apply simp_all - apply (metis (no_types) closed_segment_commute csegment_midpoint_subset midpoint_sym subset_iff) - apply (metis closed_segment_commute midpoint_sym notin_segment_midpoint) - done - qed - next - case False - show ?thesis - using Urysohn_both_ne [OF US UT \S \ T = {}\ \S \ {}\ \T \ {}\ \a \ b\] that - by blast - qed -qed - -lemma Urysohn_local: - assumes US: "closedin (subtopology euclidean U) S" - and UT: "closedin (subtopology euclidean U) T" - and "S \ T = {}" - obtains f :: "'a::euclidean_space \ 'b::euclidean_space" - where "continuous_on U f" - "\x. x \ U \ f x \ closed_segment a b" - "\x. x \ S \ f x = a" - "\x. x \ T \ f x = b" -proof (cases "a = b") - case True then show ?thesis - by (rule_tac f = "\x. b" in that) (auto simp: continuous_on_const) -next - case False - then show ?thesis - apply (rule Urysohn_local_strong [OF assms]) - apply (erule that, assumption) - apply (meson US closedin_singleton closedin_trans) - apply (meson UT closedin_singleton closedin_trans) - done -qed - -lemma Urysohn_strong: - assumes US: "closed S" - and UT: "closed T" - and "S \ T = {}" "a \ b" - obtains f :: "'a::euclidean_space \ 'b::euclidean_space" - where "continuous_on UNIV f" - "\x. f x \ closed_segment a b" - "\x. f x = a \ x \ S" - "\x. f x = b \ x \ T" -apply (rule Urysohn_local_strong [of UNIV S T]) -using assms -apply (auto simp: closed_closedin) -done - -proposition Urysohn: - assumes US: "closed S" - and UT: "closed T" - and "S \ T = {}" - obtains f :: "'a::euclidean_space \ 'b::euclidean_space" - where "continuous_on UNIV f" - "\x. f x \ closed_segment a b" - "\x. x \ S \ f x = a" - "\x. x \ T \ f x = b" -apply (rule Urysohn_local [of UNIV S T a b]) -using assms -apply (auto simp: closed_closedin) -done - - -subsection\ The Dugundji extension theorem, and Tietze variants as corollaries.\ - -text\J. Dugundji. An extension of Tietze's theorem. Pacific J. Math. Volume 1, Number 3 (1951), 353-367. -http://projecteuclid.org/euclid.pjm/1103052106\ - -theorem Dugundji: - fixes f :: "'a::euclidean_space \ 'b::real_inner" - assumes "convex C" "C \ {}" - and cloin: "closedin (subtopology euclidean U) S" - and contf: "continuous_on S f" and "f ` S \ C" - obtains g where "continuous_on U g" "g ` U \ C" - "\x. x \ S \ g x = f x" -proof (cases "S = {}") - case True then show thesis - apply (rule_tac g="\x. @y. y \ C" in that) - apply (rule continuous_intros) - apply (meson all_not_in_conv \C \ {}\ image_subsetI someI_ex, simp) - done -next - case False - then have sd_pos: "\x. \x \ U; x \ S\ \ 0 < setdist {x} S" - using setdist_eq_0_closedin [OF cloin] le_less setdist_pos_le by fastforce - define \ where "\ = {ball x (setdist {x} S / 2) |x. x \ U - S}" - have [simp]: "\T. T \ \ \ open T" - by (auto simp: \_def) - have USS: "U - S \ \\" - by (auto simp: sd_pos \_def) - obtain \ where USsub: "U - S \ \\" - and nbrhd: "\U. U \ \ \ open U \ (\T. T \ \ \ U \ T)" - and fin: "\x. x \ U - S - \ \V. open V \ x \ V \ finite {U. U \ \ \ U \ V \ {}}" - using paracompact [OF USS] by auto - have "\v a. v \ U \ v \ S \ a \ S \ - T \ ball v (setdist {v} S / 2) \ - dist v a \ 2 * setdist {v} S" if "T \ \" for T - proof - - obtain v where v: "T \ ball v (setdist {v} S / 2)" "v \ U" "v \ S" - using \T \ \\ nbrhd by (force simp: \_def) - then obtain a where "a \ S" "dist v a < 2 * setdist {v} S" - using setdist_ltE [of "{v}" S "2 * setdist {v} S"] - using False sd_pos by force - with v show ?thesis - apply (rule_tac x=v in exI) - apply (rule_tac x=a in exI, auto) - done - qed - then obtain \ \ where - VA: "\T. T \ \ \ \ T \ U \ \ T \ S \ \ T \ S \ - T \ ball (\ T) (setdist {\ T} S / 2) \ - dist (\ T) (\ T) \ 2 * setdist {\ T} S" - by metis - have sdle: "setdist {\ T} S \ 2 * setdist {v} S" if "T \ \" "v \ T" for T v - using setdist_Lipschitz [of "\ T" S v] VA [OF \T \ \\] \v \ T\ by auto - have d6: "dist a (\ T) \ 6 * dist a v" if "T \ \" "v \ T" "a \ S" for T v a - proof - - have "dist (\ T) v < setdist {\ T} S / 2" - using that VA mem_ball by blast - also have "... \ setdist {v} S" - using sdle [OF \T \ \\ \v \ T\] by simp - also have vS: "setdist {v} S \ dist a v" - by (simp add: setdist_le_dist setdist_sym \a \ S\) - finally have VTV: "dist (\ T) v < dist a v" . - have VTS: "setdist {\ T} S \ 2 * dist a v" - using sdle that vS by force - have "dist a (\ T) \ dist a v + dist v (\ T) + dist (\ T) (\ T)" - by (metis add.commute add_le_cancel_left dist_commute dist_triangle2 dist_triangle_le) - also have "... \ dist a v + dist a v + dist (\ T) (\ T)" - using VTV by (simp add: dist_commute) - also have "... \ 2 * dist a v + 2 * setdist {\ T} S" - using VA [OF \T \ \\] by auto - finally show ?thesis - using VTS by linarith - qed - obtain H :: "['a set, 'a] \ real" - where Hcont: "\Z. Z \ \ \ continuous_on (U-S) (H Z)" - and Hge0: "\Z x. \Z \ \; x \ U-S\ \ 0 \ H Z x" - and Heq0: "\x Z. \Z \ \; x \ U-S; x \ Z\ \ H Z x = 0" - and H1: "\x. x \ U-S \ supp_setsum (\W. H W x) \ = 1" - and Hfin: "\x. x \ U-S \ \V. open V \ x \ V \ finite {U \ \. \x\V. H U x \ 0}" - apply (rule subordinate_partition_of_unity [OF USsub _ fin]) - using nbrhd by auto - define g where "g \ \x. if x \ S then f x else supp_setsum (\T. H T x *\<^sub>R f(\ T)) \" - show ?thesis - proof (rule that) - show "continuous_on U g" - proof (clarsimp simp: continuous_on_eq_continuous_within) - fix a assume "a \ U" - show "continuous (at a within U) g" - proof (cases "a \ S") - case True show ?thesis - proof (clarsimp simp add: continuous_within_topological) - fix W - assume "open W" "g a \ W" - then obtain e where "0 < e" and e: "ball (f a) e \ W" - using openE True g_def by auto - have "continuous (at a within S) f" - using True contf continuous_on_eq_continuous_within by blast - then obtain d where "0 < d" - and d: "\x. \x \ S; dist x a < d\ \ dist (f x) (f a) < e" - using continuous_within_eps_delta \0 < e\ by force - have "g y \ ball (f a) e" if "y \ U" and y: "y \ ball a (d / 6)" for y - proof (cases "y \ S") - case True - then have "dist (f a) (f y) < e" - by (metis ball_divide_subset_numeral dist_commute in_mono mem_ball y d) - then show ?thesis - by (simp add: True g_def) - next - case False - have *: "dist (f (\ T)) (f a) < e" if "T \ \" "H T y \ 0" for T - proof - - have "y \ T" - using Heq0 that False \y \ U\ by blast - have "dist (\ T) a < d" - using d6 [OF \T \ \\ \y \ T\ \a \ S\] y - by (simp add: dist_commute mult.commute) - then show ?thesis - using VA [OF \T \ \\] by (auto simp: d) - qed - have "supp_setsum (\T. H T y *\<^sub>R f (\ T)) \ \ ball (f a) e" - apply (rule convex_supp_setsum [OF convex_ball]) - apply (simp_all add: False H1 Hge0 \y \ U\) - by (metis dist_commute *) - then show ?thesis - by (simp add: False g_def) - qed - then show "\A. open A \ a \ A \ (\y\U. y \ A \ g y \ W)" - apply (rule_tac x = "ball a (d / 6)" in exI) - using e \0 < d\ by fastforce - qed - next - case False - obtain N where N: "open N" "a \ N" - and finN: "finite {U \ \. \a\N. H U a \ 0}" - using Hfin False \a \ U\ by auto - have oUS: "openin (subtopology euclidean U) (U - S)" - using cloin by (simp add: openin_diff) - have HcontU: "continuous (at a within U) (H T)" if "T \ \" for T - using Hcont [OF \T \ \\] False \a \ U\ \T \ \\ - apply (simp add: continuous_on_eq_continuous_within continuous_within) - apply (rule Lim_transform_within_set) - using oUS - apply (force simp: eventually_at openin_contains_ball dist_commute dest!: bspec)+ - done - show ?thesis - proof (rule continuous_transform_within_openin [OF _ oUS]) - show "continuous (at a within U) (\x. supp_setsum (\T. H T x *\<^sub>R f (\ T)) \)" - proof (rule continuous_transform_within_openin) - show "continuous (at a within U) - (\x. \T\{U \ \. \x\N. H U x \ 0}. H T x *\<^sub>R f (\ T))" - by (force intro: continuous_intros HcontU)+ - next - show "openin (subtopology euclidean U) ((U - S) \ N)" - using N oUS openin_trans by blast - next - show "a \ (U - S) \ N" using False \a \ U\ N by blast - next - show "\x. x \ (U - S) \ N \ - (\T \ {U \ \. \x\N. H U x \ 0}. H T x *\<^sub>R f (\ T)) - = supp_setsum (\T. H T x *\<^sub>R f (\ T)) \" - by (auto simp: supp_setsum_def support_on_def - intro: setsum.mono_neutral_right [OF finN]) - qed - next - show "a \ U - S" using False \a \ U\ by blast - next - show "\x. x \ U - S \ supp_setsum (\T. H T x *\<^sub>R f (\ T)) \ = g x" - by (simp add: g_def) - qed - qed - qed - show "g ` U \ C" - using \f ` S \ C\ VA - by (fastforce simp: g_def Hge0 intro!: convex_supp_setsum [OF \convex C\] H1) - show "\x. x \ S \ g x = f x" - by (simp add: g_def) - qed -qed - - -corollary Tietze: - fixes f :: "'a::euclidean_space \ 'b::real_inner" - assumes "continuous_on S f" - and "closedin (subtopology euclidean U) S" - and "0 \ B" - and "\x. x \ S \ norm(f x) \ B" - obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" - "\x. x \ U \ norm(g x) \ B" -using assms -by (auto simp: image_subset_iff intro: Dugundji [of "cball 0 B" U S f]) - -corollary Tietze_closed_interval: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "continuous_on S f" - and "closedin (subtopology euclidean U) S" - and "cbox a b \ {}" - and "\x. x \ S \ f x \ cbox a b" - obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" - "\x. x \ U \ g x \ cbox a b" -apply (rule Dugundji [of "cbox a b" U S f]) -using assms by auto - -corollary Tietze_closed_interval_1: - fixes f :: "'a::euclidean_space \ real" - assumes "continuous_on S f" - and "closedin (subtopology euclidean U) S" - and "a \ b" - and "\x. x \ S \ f x \ cbox a b" - obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" - "\x. x \ U \ g x \ cbox a b" -apply (rule Dugundji [of "cbox a b" U S f]) -using assms by (auto simp: image_subset_iff) - -corollary Tietze_open_interval: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "continuous_on S f" - and "closedin (subtopology euclidean U) S" - and "box a b \ {}" - and "\x. x \ S \ f x \ box a b" - obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" - "\x. x \ U \ g x \ box a b" -apply (rule Dugundji [of "box a b" U S f]) -using assms by auto - -corollary Tietze_open_interval_1: - fixes f :: "'a::euclidean_space \ real" - assumes "continuous_on S f" - and "closedin (subtopology euclidean U) S" - and "a < b" - and no: "\x. x \ S \ f x \ box a b" - obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" - "\x. x \ U \ g x \ box a b" -apply (rule Dugundji [of "box a b" U S f]) -using assms by (auto simp: image_subset_iff) - -corollary Tietze_unbounded: - fixes f :: "'a::euclidean_space \ 'b::real_inner" - assumes "continuous_on S f" - and "closedin (subtopology euclidean U) S" - obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" -apply (rule Dugundji [of UNIV U S f]) -using assms by auto - -end diff -r bbcb05504fdc -r bd218a9320b5 src/HOL/Multivariate_Analysis/Fashoda.thy --- a/src/HOL/Multivariate_Analysis/Fashoda.thy Thu Aug 04 18:45:28 2016 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1136 +0,0 @@ -(* Author: John Harrison - Author: Robert Himmelmann, TU Muenchen (translation from HOL light) -*) - -section \Fashoda meet theorem\ - -theory Fashoda -imports Brouwer_Fixpoint Path_Connected Cartesian_Euclidean_Space -begin - -subsection \Bijections between intervals.\ - -definition interval_bij :: "'a \ 'a \ 'a \ 'a \ 'a \ 'a::euclidean_space" - where "interval_bij = - (\(a, b) (u, v) x. (\i\Basis. (u\i + (x\i - a\i) / (b\i - a\i) * (v\i - u\i)) *\<^sub>R i))" - -lemma interval_bij_affine: - "interval_bij (a,b) (u,v) = (\x. (\i\Basis. ((v\i - u\i) / (b\i - a\i) * (x\i)) *\<^sub>R i) + - (\i\Basis. (u\i - (v\i - u\i) / (b\i - a\i) * (a\i)) *\<^sub>R i))" - by (auto simp: setsum.distrib[symmetric] scaleR_add_left[symmetric] interval_bij_def fun_eq_iff - field_simps inner_simps add_divide_distrib[symmetric] intro!: setsum.cong) - -lemma continuous_interval_bij: - fixes a b :: "'a::euclidean_space" - shows "continuous (at x) (interval_bij (a, b) (u, v))" - by (auto simp add: divide_inverse interval_bij_def intro!: continuous_setsum continuous_intros) - -lemma continuous_on_interval_bij: "continuous_on s (interval_bij (a, b) (u, v))" - apply(rule continuous_at_imp_continuous_on) - apply (rule, rule continuous_interval_bij) - done - -lemma in_interval_interval_bij: - fixes a b u v x :: "'a::euclidean_space" - assumes "x \ cbox a b" - and "cbox u v \ {}" - shows "interval_bij (a, b) (u, v) x \ cbox u v" - apply (simp only: interval_bij_def split_conv mem_box inner_setsum_left_Basis cong: ball_cong) - apply safe -proof - - fix i :: 'a - assume i: "i \ Basis" - have "cbox a b \ {}" - using assms by auto - with i have *: "a\i \ b\i" "u\i \ v\i" - using assms(2) by (auto simp add: box_eq_empty) - have x: "a\i\x\i" "x\i\b\i" - using assms(1)[unfolded mem_box] using i by auto - have "0 \ (x \ i - a \ i) / (b \ i - a \ i) * (v \ i - u \ i)" - using * x by auto - then show "u \ i \ u \ i + (x \ i - a \ i) / (b \ i - a \ i) * (v \ i - u \ i)" - using * by auto - have "((x \ i - a \ i) / (b \ i - a \ i)) * (v \ i - u \ i) \ 1 * (v \ i - u \ i)" - apply (rule mult_right_mono) - unfolding divide_le_eq_1 - using * x - apply auto - done - then show "u \ i + (x \ i - a \ i) / (b \ i - a \ i) * (v \ i - u \ i) \ v \ i" - using * by auto -qed - -lemma interval_bij_bij: - "\(i::'a::euclidean_space)\Basis. a\i < b\i \ u\i < v\i \ - interval_bij (a, b) (u, v) (interval_bij (u, v) (a, b) x) = x" - by (auto simp: interval_bij_def euclidean_eq_iff[where 'a='a]) - -lemma interval_bij_bij_cart: fixes x::"real^'n" assumes "\i. a$i < b$i \ u$i < v$i" - shows "interval_bij (a,b) (u,v) (interval_bij (u,v) (a,b) x) = x" - using assms by (intro interval_bij_bij) (auto simp: Basis_vec_def inner_axis) - - -subsection \Fashoda meet theorem\ - -lemma infnorm_2: - fixes x :: "real^2" - shows "infnorm x = max \x$1\ \x$2\" - unfolding infnorm_cart UNIV_2 by (rule cSup_eq) auto - -lemma infnorm_eq_1_2: - fixes x :: "real^2" - shows "infnorm x = 1 \ - \x$1\ \ 1 \ \x$2\ \ 1 \ (x$1 = -1 \ x$1 = 1 \ x$2 = -1 \ x$2 = 1)" - unfolding infnorm_2 by auto - -lemma infnorm_eq_1_imp: - fixes x :: "real^2" - assumes "infnorm x = 1" - shows "\x$1\ \ 1" and "\x$2\ \ 1" - using assms unfolding infnorm_eq_1_2 by auto - -lemma fashoda_unit: - fixes f g :: "real \ real^2" - assumes "f ` {-1 .. 1} \ cbox (-1) 1" - and "g ` {-1 .. 1} \ cbox (-1) 1" - and "continuous_on {-1 .. 1} f" - and "continuous_on {-1 .. 1} g" - and "f (- 1)$1 = - 1" - and "f 1$1 = 1" "g (- 1) $2 = -1" - and "g 1 $2 = 1" - shows "\s\{-1 .. 1}. \t\{-1 .. 1}. f s = g t" -proof (rule ccontr) - assume "\ ?thesis" - note as = this[unfolded bex_simps,rule_format] - define sqprojection - where [abs_def]: "sqprojection z = (inverse (infnorm z)) *\<^sub>R z" for z :: "real^2" - define negatex :: "real^2 \ real^2" - where "negatex x = (vector [-(x$1), x$2])" for x - have lem1: "\z::real^2. infnorm (negatex z) = infnorm z" - unfolding negatex_def infnorm_2 vector_2 by auto - have lem2: "\z. z \ 0 \ infnorm (sqprojection z) = 1" - unfolding sqprojection_def - unfolding infnorm_mul[unfolded scalar_mult_eq_scaleR] - unfolding abs_inverse real_abs_infnorm - apply (subst infnorm_eq_0[symmetric]) - apply auto - done - let ?F = "\w::real^2. (f \ (\x. x$1)) w - (g \ (\x. x$2)) w" - have *: "\i. (\x::real^2. x $ i) ` cbox (- 1) 1 = {-1 .. 1}" - apply (rule set_eqI) - unfolding image_iff Bex_def mem_interval_cart interval_cbox_cart - apply rule - defer - apply (rule_tac x="vec x" in exI) - apply auto - done - { - fix x - assume "x \ (\w. (f \ (\x. x $ 1)) w - (g \ (\x. x $ 2)) w) ` (cbox (- 1) (1::real^2))" - then obtain w :: "real^2" where w: - "w \ cbox (- 1) 1" - "x = (f \ (\x. x $ 1)) w - (g \ (\x. x $ 2)) w" - unfolding image_iff .. - then have "x \ 0" - using as[of "w$1" "w$2"] - unfolding mem_interval_cart atLeastAtMost_iff - by auto - } note x0 = this - have 21: "\i::2. i \ 1 \ i = 2" - using UNIV_2 by auto - have 1: "box (- 1) (1::real^2) \ {}" - unfolding interval_eq_empty_cart by auto - have 2: "continuous_on (cbox (- 1) 1) (negatex \ sqprojection \ ?F)" - apply (intro continuous_intros continuous_on_component) - unfolding * - apply (rule assms)+ - apply (subst sqprojection_def) - apply (intro continuous_intros) - apply (simp add: infnorm_eq_0 x0) - apply (rule linear_continuous_on) - proof - - show "bounded_linear negatex" - apply (rule bounded_linearI') - unfolding vec_eq_iff - proof (rule_tac[!] allI) - fix i :: 2 - fix x y :: "real^2" - fix c :: real - show "negatex (x + y) $ i = - (negatex x + negatex y) $ i" "negatex (c *\<^sub>R x) $ i = (c *\<^sub>R negatex x) $ i" - apply - - apply (case_tac[!] "i\1") - prefer 3 - apply (drule_tac[1-2] 21) - unfolding negatex_def - apply (auto simp add:vector_2) - done - qed - qed - have 3: "(negatex \ sqprojection \ ?F) ` cbox (-1) 1 \ cbox (-1) 1" - unfolding subset_eq - proof (rule, goal_cases) - case (1 x) - then obtain y :: "real^2" where y: - "y \ cbox (- 1) 1" - "x = (negatex \ sqprojection \ (\w. (f \ (\x. x $ 1)) w - (g \ (\x. x $ 2)) w)) y" - unfolding image_iff .. - have "?F y \ 0" - apply (rule x0) - using y(1) - apply auto - done - then have *: "infnorm (sqprojection (?F y)) = 1" - unfolding y o_def - by - (rule lem2[rule_format]) - have "infnorm x = 1" - unfolding *[symmetric] y o_def - by (rule lem1[rule_format]) - then show "x \ cbox (-1) 1" - unfolding mem_interval_cart interval_cbox_cart infnorm_2 - apply - - apply rule - proof - - fix i - assume "max \x $ 1\ \x $ 2\ = 1" - then show "(- 1) $ i \ x $ i \ x $ i \ 1 $ i" - apply (cases "i = 1") - defer - apply (drule 21) - apply auto - done - qed - qed - obtain x :: "real^2" where x: - "x \ cbox (- 1) 1" - "(negatex \ sqprojection \ (\w. (f \ (\x. x $ 1)) w - (g \ (\x. x $ 2)) w)) x = x" - apply (rule brouwer_weak[of "cbox (- 1) (1::real^2)" "negatex \ sqprojection \ ?F"]) - apply (rule compact_cbox convex_box)+ - unfolding interior_cbox - apply (rule 1 2 3)+ - apply blast - done - have "?F x \ 0" - apply (rule x0) - using x(1) - apply auto - done - then have *: "infnorm (sqprojection (?F x)) = 1" - unfolding o_def - by (rule lem2[rule_format]) - have nx: "infnorm x = 1" - apply (subst x(2)[symmetric]) - unfolding *[symmetric] o_def - apply (rule lem1[rule_format]) - done - have "\x i. x \ 0 \ (0 < (sqprojection x)$i \ 0 < x$i)" - and "\x i. x \ 0 \ ((sqprojection x)$i < 0 \ x$i < 0)" - apply - - apply (rule_tac[!] allI impI)+ - proof - - fix x :: "real^2" - fix i :: 2 - assume x: "x \ 0" - have "inverse (infnorm x) > 0" - using x[unfolded infnorm_pos_lt[symmetric]] by auto - then show "(0 < sqprojection x $ i) = (0 < x $ i)" - and "(sqprojection x $ i < 0) = (x $ i < 0)" - unfolding sqprojection_def vector_component_simps vector_scaleR_component real_scaleR_def - unfolding zero_less_mult_iff mult_less_0_iff - by (auto simp add: field_simps) - qed - note lem3 = this[rule_format] - have x1: "x $ 1 \ {- 1..1::real}" "x $ 2 \ {- 1..1::real}" - using x(1) unfolding mem_interval_cart by auto - then have nz: "f (x $ 1) - g (x $ 2) \ 0" - unfolding right_minus_eq - apply - - apply (rule as) - apply auto - done - have "x $ 1 = -1 \ x $ 1 = 1 \ x $ 2 = -1 \ x $ 2 = 1" - using nx unfolding infnorm_eq_1_2 by auto - then show False - proof - - fix P Q R S - presume "P \ Q \ R \ S" - and "P \ False" - and "Q \ False" - and "R \ False" - and "S \ False" - then show False by auto - next - assume as: "x$1 = 1" - then have *: "f (x $ 1) $ 1 = 1" - using assms(6) by auto - have "sqprojection (f (x$1) - g (x$2)) $ 1 < 0" - using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=1]] - unfolding as negatex_def vector_2 - by auto - moreover - from x1 have "g (x $ 2) \ cbox (-1) 1" - apply - - apply (rule assms(2)[unfolded subset_eq,rule_format]) - apply auto - done - ultimately show False - unfolding lem3[OF nz] vector_component_simps * mem_interval_cart - apply (erule_tac x=1 in allE) - apply auto - done - next - assume as: "x$1 = -1" - then have *: "f (x $ 1) $ 1 = - 1" - using assms(5) by auto - have "sqprojection (f (x$1) - g (x$2)) $ 1 > 0" - using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=1]] - unfolding as negatex_def vector_2 - by auto - moreover - from x1 have "g (x $ 2) \ cbox (-1) 1" - apply - - apply (rule assms(2)[unfolded subset_eq,rule_format]) - apply auto - done - ultimately show False - unfolding lem3[OF nz] vector_component_simps * mem_interval_cart - apply (erule_tac x=1 in allE) - apply auto - done - next - assume as: "x$2 = 1" - then have *: "g (x $ 2) $ 2 = 1" - using assms(8) by auto - have "sqprojection (f (x$1) - g (x$2)) $ 2 > 0" - using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=2]] - unfolding as negatex_def vector_2 - by auto - moreover - from x1 have "f (x $ 1) \ cbox (-1) 1" - apply - - apply (rule assms(1)[unfolded subset_eq,rule_format]) - apply auto - done - ultimately show False - unfolding lem3[OF nz] vector_component_simps * mem_interval_cart - apply (erule_tac x=2 in allE) - apply auto - done - next - assume as: "x$2 = -1" - then have *: "g (x $ 2) $ 2 = - 1" - using assms(7) by auto - have "sqprojection (f (x$1) - g (x$2)) $ 2 < 0" - using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=2]] - unfolding as negatex_def vector_2 - by auto - moreover - from x1 have "f (x $ 1) \ cbox (-1) 1" - apply - - apply (rule assms(1)[unfolded subset_eq,rule_format]) - apply auto - done - ultimately show False - unfolding lem3[OF nz] vector_component_simps * mem_interval_cart - apply (erule_tac x=2 in allE) - apply auto - done - qed auto -qed - -lemma fashoda_unit_path: - fixes f g :: "real \ real^2" - assumes "path f" - and "path g" - and "path_image f \ cbox (-1) 1" - and "path_image g \ cbox (-1) 1" - and "(pathstart f)$1 = -1" - and "(pathfinish f)$1 = 1" - and "(pathstart g)$2 = -1" - and "(pathfinish g)$2 = 1" - obtains z where "z \ path_image f" and "z \ path_image g" -proof - - note assms=assms[unfolded path_def pathstart_def pathfinish_def path_image_def] - define iscale where [abs_def]: "iscale z = inverse 2 *\<^sub>R (z + 1)" for z :: real - have isc: "iscale ` {- 1..1} \ {0..1}" - unfolding iscale_def by auto - have "\s\{- 1..1}. \t\{- 1..1}. (f \ iscale) s = (g \ iscale) t" - proof (rule fashoda_unit) - show "(f \ iscale) ` {- 1..1} \ cbox (- 1) 1" "(g \ iscale) ` {- 1..1} \ cbox (- 1) 1" - using isc and assms(3-4) by (auto simp add: image_comp [symmetric]) - have *: "continuous_on {- 1..1} iscale" - unfolding iscale_def by (rule continuous_intros)+ - show "continuous_on {- 1..1} (f \ iscale)" "continuous_on {- 1..1} (g \ iscale)" - apply - - apply (rule_tac[!] continuous_on_compose[OF *]) - apply (rule_tac[!] continuous_on_subset[OF _ isc]) - apply (rule assms)+ - done - have *: "(1 / 2) *\<^sub>R (1 + (1::real^1)) = 1" - unfolding vec_eq_iff by auto - show "(f \ iscale) (- 1) $ 1 = - 1" - and "(f \ iscale) 1 $ 1 = 1" - and "(g \ iscale) (- 1) $ 2 = -1" - and "(g \ iscale) 1 $ 2 = 1" - unfolding o_def iscale_def - using assms - by (auto simp add: *) - qed - then obtain s t where st: - "s \ {- 1..1}" - "t \ {- 1..1}" - "(f \ iscale) s = (g \ iscale) t" - by auto - show thesis - apply (rule_tac z = "f (iscale s)" in that) - using st - unfolding o_def path_image_def image_iff - apply - - apply (rule_tac x="iscale s" in bexI) - prefer 3 - apply (rule_tac x="iscale t" in bexI) - using isc[unfolded subset_eq, rule_format] - apply auto - done -qed - -lemma fashoda: - fixes b :: "real^2" - assumes "path f" - and "path g" - and "path_image f \ cbox a b" - and "path_image g \ cbox a b" - and "(pathstart f)$1 = a$1" - and "(pathfinish f)$1 = b$1" - and "(pathstart g)$2 = a$2" - and "(pathfinish g)$2 = b$2" - obtains z where "z \ path_image f" and "z \ path_image g" -proof - - fix P Q S - presume "P \ Q \ S" "P \ thesis" and "Q \ thesis" and "S \ thesis" - then show thesis - by auto -next - have "cbox a b \ {}" - using assms(3) using path_image_nonempty[of f] by auto - then have "a \ b" - unfolding interval_eq_empty_cart less_eq_vec_def by (auto simp add: not_less) - then show "a$1 = b$1 \ a$2 = b$2 \ (a$1 < b$1 \ a$2 < b$2)" - unfolding less_eq_vec_def forall_2 by auto -next - assume as: "a$1 = b$1" - have "\z\path_image g. z$2 = (pathstart f)$2" - apply (rule connected_ivt_component_cart) - apply (rule connected_path_image assms)+ - apply (rule pathstart_in_path_image) - apply (rule pathfinish_in_path_image) - unfolding assms using assms(3)[unfolded path_image_def subset_eq,rule_format,of "f 0"] - unfolding pathstart_def - apply (auto simp add: less_eq_vec_def mem_interval_cart) - done - then obtain z :: "real^2" where z: "z \ path_image g" "z $ 2 = pathstart f $ 2" .. - have "z \ cbox a b" - using z(1) assms(4) - unfolding path_image_def - by blast - then have "z = f 0" - unfolding vec_eq_iff forall_2 - unfolding z(2) pathstart_def - using assms(3)[unfolded path_image_def subset_eq mem_interval_cart,rule_format,of "f 0" 1] - unfolding mem_interval_cart - apply (erule_tac x=1 in allE) - using as - apply auto - done - then show thesis - apply - - apply (rule that[OF _ z(1)]) - unfolding path_image_def - apply auto - done -next - assume as: "a$2 = b$2" - have "\z\path_image f. z$1 = (pathstart g)$1" - apply (rule connected_ivt_component_cart) - apply (rule connected_path_image assms)+ - apply (rule pathstart_in_path_image) - apply (rule pathfinish_in_path_image) - unfolding assms - using assms(4)[unfolded path_image_def subset_eq,rule_format,of "g 0"] - unfolding pathstart_def - apply (auto simp add: less_eq_vec_def mem_interval_cart) - done - then obtain z where z: "z \ path_image f" "z $ 1 = pathstart g $ 1" .. - have "z \ cbox a b" - using z(1) assms(3) - unfolding path_image_def - by blast - then have "z = g 0" - unfolding vec_eq_iff forall_2 - unfolding z(2) pathstart_def - using assms(4)[unfolded path_image_def subset_eq mem_interval_cart,rule_format,of "g 0" 2] - unfolding mem_interval_cart - apply (erule_tac x=2 in allE) - using as - apply auto - done - then show thesis - apply - - apply (rule that[OF z(1)]) - unfolding path_image_def - apply auto - done -next - assume as: "a $ 1 < b $ 1 \ a $ 2 < b $ 2" - have int_nem: "cbox (-1) (1::real^2) \ {}" - unfolding interval_eq_empty_cart by auto - obtain z :: "real^2" where z: - "z \ (interval_bij (a, b) (- 1, 1) \ f) ` {0..1}" - "z \ (interval_bij (a, b) (- 1, 1) \ g) ` {0..1}" - apply (rule fashoda_unit_path[of "interval_bij (a,b) (- 1,1) \ f" "interval_bij (a,b) (- 1,1) \ g"]) - unfolding path_def path_image_def pathstart_def pathfinish_def - apply (rule_tac[1-2] continuous_on_compose) - apply (rule assms[unfolded path_def] continuous_on_interval_bij)+ - unfolding subset_eq - apply(rule_tac[1-2] ballI) - proof - - fix x - assume "x \ (interval_bij (a, b) (- 1, 1) \ f) ` {0..1}" - then obtain y where y: - "y \ {0..1}" - "x = (interval_bij (a, b) (- 1, 1) \ f) y" - unfolding image_iff .. - show "x \ cbox (- 1) 1" - unfolding y o_def - apply (rule in_interval_interval_bij) - using y(1) - using assms(3)[unfolded path_image_def subset_eq] int_nem - apply auto - done - next - fix x - assume "x \ (interval_bij (a, b) (- 1, 1) \ g) ` {0..1}" - then obtain y where y: - "y \ {0..1}" - "x = (interval_bij (a, b) (- 1, 1) \ g) y" - unfolding image_iff .. - show "x \ cbox (- 1) 1" - unfolding y o_def - apply (rule in_interval_interval_bij) - using y(1) - using assms(4)[unfolded path_image_def subset_eq] int_nem - apply auto - done - next - show "(interval_bij (a, b) (- 1, 1) \ f) 0 $ 1 = -1" - and "(interval_bij (a, b) (- 1, 1) \ f) 1 $ 1 = 1" - and "(interval_bij (a, b) (- 1, 1) \ g) 0 $ 2 = -1" - and "(interval_bij (a, b) (- 1, 1) \ g) 1 $ 2 = 1" - using assms as - by (simp_all add: axis_in_Basis cart_eq_inner_axis pathstart_def pathfinish_def interval_bij_def) - (simp_all add: inner_axis) - qed - from z(1) obtain zf where zf: - "zf \ {0..1}" - "z = (interval_bij (a, b) (- 1, 1) \ f) zf" - unfolding image_iff .. - from z(2) obtain zg where zg: - "zg \ {0..1}" - "z = (interval_bij (a, b) (- 1, 1) \ g) zg" - unfolding image_iff .. - have *: "\i. (- 1) $ i < (1::real^2) $ i \ a $ i < b $ i" - unfolding forall_2 - using as - by auto - show thesis - apply (rule_tac z="interval_bij (- 1,1) (a,b) z" in that) - apply (subst zf) - defer - apply (subst zg) - unfolding o_def interval_bij_bij_cart[OF *] path_image_def - using zf(1) zg(1) - apply auto - done -qed - - -subsection \Some slightly ad hoc lemmas I use below\ - -lemma segment_vertical: - fixes a :: "real^2" - assumes "a$1 = b$1" - shows "x \ closed_segment a b \ - x$1 = a$1 \ x$1 = b$1 \ (a$2 \ x$2 \ x$2 \ b$2 \ b$2 \ x$2 \ x$2 \ a$2)" - (is "_ = ?R") -proof - - let ?L = "\u. (x $ 1 = (1 - u) * a $ 1 + u * b $ 1 \ x $ 2 = (1 - u) * a $ 2 + u * b $ 2) \ 0 \ u \ u \ 1" - { - presume "?L \ ?R" and "?R \ ?L" - then show ?thesis - unfolding closed_segment_def mem_Collect_eq - unfolding vec_eq_iff forall_2 scalar_mult_eq_scaleR[symmetric] vector_component_simps - by blast - } - { - assume ?L - then obtain u where u: - "x $ 1 = (1 - u) * a $ 1 + u * b $ 1" - "x $ 2 = (1 - u) * a $ 2 + u * b $ 2" - "0 \ u" - "u \ 1" - by blast - { fix b a - assume "b + u * a > a + u * b" - then have "(1 - u) * b > (1 - u) * a" - by (auto simp add:field_simps) - then have "b \ a" - apply (drule_tac mult_left_less_imp_less) - using u - apply auto - done - then have "u * a \ u * b" - apply - - apply (rule mult_left_mono[OF _ u(3)]) - using u(3-4) - apply (auto simp add: field_simps) - done - } note * = this - { - fix a b - assume "u * b > u * a" - then have "(1 - u) * a \ (1 - u) * b" - apply - - apply (rule mult_left_mono) - apply (drule mult_left_less_imp_less) - using u - apply auto - done - then have "a + u * b \ b + u * a" - by (auto simp add: field_simps) - } note ** = this - then show ?R - unfolding u assms - using u - by (auto simp add:field_simps not_le intro: * **) - } - { - assume ?R - then show ?L - proof (cases "x$2 = b$2") - case True - then show ?L - apply (rule_tac x="(x$2 - a$2) / (b$2 - a$2)" in exI) - unfolding assms True - using \?R\ - apply (auto simp add: field_simps) - done - next - case False - then show ?L - apply (rule_tac x="1 - (x$2 - b$2) / (a$2 - b$2)" in exI) - unfolding assms - using \?R\ - apply (auto simp add: field_simps) - done - qed - } -qed - -lemma segment_horizontal: - fixes a :: "real^2" - assumes "a$2 = b$2" - shows "x \ closed_segment a b \ - x$2 = a$2 \ x$2 = b$2 \ (a$1 \ x$1 \ x$1 \ b$1 \ b$1 \ x$1 \ x$1 \ a$1)" - (is "_ = ?R") -proof - - let ?L = "\u. (x $ 1 = (1 - u) * a $ 1 + u * b $ 1 \ x $ 2 = (1 - u) * a $ 2 + u * b $ 2) \ 0 \ u \ u \ 1" - { - presume "?L \ ?R" and "?R \ ?L" - then show ?thesis - unfolding closed_segment_def mem_Collect_eq - unfolding vec_eq_iff forall_2 scalar_mult_eq_scaleR[symmetric] vector_component_simps - by blast - } - { - assume ?L - then obtain u where u: - "x $ 1 = (1 - u) * a $ 1 + u * b $ 1" - "x $ 2 = (1 - u) * a $ 2 + u * b $ 2" - "0 \ u" - "u \ 1" - by blast - { - fix b a - assume "b + u * a > a + u * b" - then have "(1 - u) * b > (1 - u) * a" - by (auto simp add: field_simps) - then have "b \ a" - apply (drule_tac mult_left_less_imp_less) - using u - apply auto - done - then have "u * a \ u * b" - apply - - apply (rule mult_left_mono[OF _ u(3)]) - using u(3-4) - apply (auto simp add: field_simps) - done - } note * = this - { - fix a b - assume "u * b > u * a" - then have "(1 - u) * a \ (1 - u) * b" - apply - - apply (rule mult_left_mono) - apply (drule mult_left_less_imp_less) - using u - apply auto - done - then have "a + u * b \ b + u * a" - by (auto simp add: field_simps) - } note ** = this - then show ?R - unfolding u assms - using u - by (auto simp add: field_simps not_le intro: * **) - } - { - assume ?R - then show ?L - proof (cases "x$1 = b$1") - case True - then show ?L - apply (rule_tac x="(x$1 - a$1) / (b$1 - a$1)" in exI) - unfolding assms True - using \?R\ - apply (auto simp add: field_simps) - done - next - case False - then show ?L - apply (rule_tac x="1 - (x$1 - b$1) / (a$1 - b$1)" in exI) - unfolding assms - using \?R\ - apply (auto simp add: field_simps) - done - qed - } -qed - - -subsection \Useful Fashoda corollary pointed out to me by Tom Hales\ - -lemma fashoda_interlace: - fixes a :: "real^2" - assumes "path f" - and "path g" - and "path_image f \ cbox a b" - and "path_image g \ cbox a b" - and "(pathstart f)$2 = a$2" - and "(pathfinish f)$2 = a$2" - and "(pathstart g)$2 = a$2" - and "(pathfinish g)$2 = a$2" - and "(pathstart f)$1 < (pathstart g)$1" - and "(pathstart g)$1 < (pathfinish f)$1" - and "(pathfinish f)$1 < (pathfinish g)$1" - obtains z where "z \ path_image f" and "z \ path_image g" -proof - - have "cbox a b \ {}" - using path_image_nonempty[of f] using assms(3) by auto - note ab=this[unfolded interval_eq_empty_cart not_ex forall_2 not_less] - have "pathstart f \ cbox a b" - and "pathfinish f \ cbox a b" - and "pathstart g \ cbox a b" - and "pathfinish g \ cbox a b" - using pathstart_in_path_image pathfinish_in_path_image - using assms(3-4) - by auto - note startfin = this[unfolded mem_interval_cart forall_2] - let ?P1 = "linepath (vector[a$1 - 2, a$2 - 2]) (vector[(pathstart f)$1,a$2 - 2]) +++ - linepath(vector[(pathstart f)$1,a$2 - 2])(pathstart f) +++ f +++ - linepath(pathfinish f)(vector[(pathfinish f)$1,a$2 - 2]) +++ - linepath(vector[(pathfinish f)$1,a$2 - 2])(vector[b$1 + 2,a$2 - 2])" - let ?P2 = "linepath(vector[(pathstart g)$1, (pathstart g)$2 - 3])(pathstart g) +++ g +++ - linepath(pathfinish g)(vector[(pathfinish g)$1,a$2 - 1]) +++ - linepath(vector[(pathfinish g)$1,a$2 - 1])(vector[b$1 + 1,a$2 - 1]) +++ - linepath(vector[b$1 + 1,a$2 - 1])(vector[b$1 + 1,b$2 + 3])" - let ?a = "vector[a$1 - 2, a$2 - 3]" - let ?b = "vector[b$1 + 2, b$2 + 3]" - have P1P2: "path_image ?P1 = path_image (linepath (vector[a$1 - 2, a$2 - 2]) (vector[(pathstart f)$1,a$2 - 2])) \ - path_image (linepath(vector[(pathstart f)$1,a$2 - 2])(pathstart f)) \ path_image f \ - path_image (linepath(pathfinish f)(vector[(pathfinish f)$1,a$2 - 2])) \ - path_image (linepath(vector[(pathfinish f)$1,a$2 - 2])(vector[b$1 + 2,a$2 - 2]))" - "path_image ?P2 = path_image(linepath(vector[(pathstart g)$1, (pathstart g)$2 - 3])(pathstart g)) \ path_image g \ - path_image(linepath(pathfinish g)(vector[(pathfinish g)$1,a$2 - 1])) \ - path_image(linepath(vector[(pathfinish g)$1,a$2 - 1])(vector[b$1 + 1,a$2 - 1])) \ - path_image(linepath(vector[b$1 + 1,a$2 - 1])(vector[b$1 + 1,b$2 + 3]))" using assms(1-2) - by(auto simp add: path_image_join path_linepath) - have abab: "cbox a b \ cbox ?a ?b" - unfolding interval_cbox_cart[symmetric] - by (auto simp add:less_eq_vec_def forall_2 vector_2) - obtain z where - "z \ path_image - (linepath (vector [a $ 1 - 2, a $ 2 - 2]) (vector [pathstart f $ 1, a $ 2 - 2]) +++ - linepath (vector [pathstart f $ 1, a $ 2 - 2]) (pathstart f) +++ - f +++ - linepath (pathfinish f) (vector [pathfinish f $ 1, a $ 2 - 2]) +++ - linepath (vector [pathfinish f $ 1, a $ 2 - 2]) (vector [b $ 1 + 2, a $ 2 - 2]))" - "z \ path_image - (linepath (vector [pathstart g $ 1, pathstart g $ 2 - 3]) (pathstart g) +++ - g +++ - linepath (pathfinish g) (vector [pathfinish g $ 1, a $ 2 - 1]) +++ - linepath (vector [pathfinish g $ 1, a $ 2 - 1]) (vector [b $ 1 + 1, a $ 2 - 1]) +++ - linepath (vector [b $ 1 + 1, a $ 2 - 1]) (vector [b $ 1 + 1, b $ 2 + 3]))" - apply (rule fashoda[of ?P1 ?P2 ?a ?b]) - unfolding pathstart_join pathfinish_join pathstart_linepath pathfinish_linepath vector_2 - proof - - show "path ?P1" and "path ?P2" - using assms by auto - have "path_image ?P1 \ cbox ?a ?b" - unfolding P1P2 path_image_linepath - apply (rule Un_least)+ - defer 3 - apply (rule_tac[1-4] convex_box(1)[unfolded convex_contains_segment,rule_format]) - unfolding mem_interval_cart forall_2 vector_2 - using ab startfin abab assms(3) - using assms(9-) - unfolding assms - apply (auto simp add: field_simps box_def) - done - then show "path_image ?P1 \ cbox ?a ?b" . - have "path_image ?P2 \ cbox ?a ?b" - unfolding P1P2 path_image_linepath - apply (rule Un_least)+ - defer 2 - apply (rule_tac[1-4] convex_box(1)[unfolded convex_contains_segment,rule_format]) - unfolding mem_interval_cart forall_2 vector_2 - using ab startfin abab assms(4) - using assms(9-) - unfolding assms - apply (auto simp add: field_simps box_def) - done - then show "path_image ?P2 \ cbox ?a ?b" . - show "a $ 1 - 2 = a $ 1 - 2" - and "b $ 1 + 2 = b $ 1 + 2" - and "pathstart g $ 2 - 3 = a $ 2 - 3" - and "b $ 2 + 3 = b $ 2 + 3" - by (auto simp add: assms) - qed - note z=this[unfolded P1P2 path_image_linepath] - show thesis - apply (rule that[of z]) - proof - - have "(z \ closed_segment (vector [a $ 1 - 2, a $ 2 - 2]) (vector [pathstart f $ 1, a $ 2 - 2]) \ - z \ closed_segment (vector [pathstart f $ 1, a $ 2 - 2]) (pathstart f)) \ - z \ closed_segment (pathfinish f) (vector [pathfinish f $ 1, a $ 2 - 2]) \ - z \ closed_segment (vector [pathfinish f $ 1, a $ 2 - 2]) (vector [b $ 1 + 2, a $ 2 - 2]) \ - (((z \ closed_segment (vector [pathstart g $ 1, pathstart g $ 2 - 3]) (pathstart g)) \ - z \ closed_segment (pathfinish g) (vector [pathfinish g $ 1, a $ 2 - 1])) \ - z \ closed_segment (vector [pathfinish g $ 1, a $ 2 - 1]) (vector [b $ 1 + 1, a $ 2 - 1])) \ - z \ closed_segment (vector [b $ 1 + 1, a $ 2 - 1]) (vector [b $ 1 + 1, b $ 2 + 3]) \ False" - proof (simp only: segment_vertical segment_horizontal vector_2, goal_cases) - case prems: 1 - have "pathfinish f \ cbox a b" - using assms(3) pathfinish_in_path_image[of f] by auto - then have "1 + b $ 1 \ pathfinish f $ 1 \ False" - unfolding mem_interval_cart forall_2 by auto - then have "z$1 \ pathfinish f$1" - using prems(2) - using assms ab - by (auto simp add: field_simps) - moreover have "pathstart f \ cbox a b" - using assms(3) pathstart_in_path_image[of f] - by auto - then have "1 + b $ 1 \ pathstart f $ 1 \ False" - unfolding mem_interval_cart forall_2 - by auto - then have "z$1 \ pathstart f$1" - using prems(2) using assms ab - by (auto simp add: field_simps) - ultimately have *: "z$2 = a$2 - 2" - using prems(1) - by auto - have "z$1 \ pathfinish g$1" - using prems(2) - using assms ab - by (auto simp add: field_simps *) - moreover have "pathstart g \ cbox a b" - using assms(4) pathstart_in_path_image[of g] - by auto - note this[unfolded mem_interval_cart forall_2] - then have "z$1 \ pathstart g$1" - using prems(1) - using assms ab - by (auto simp add: field_simps *) - ultimately have "a $ 2 - 1 \ z $ 2 \ z $ 2 \ b $ 2 + 3 \ b $ 2 + 3 \ z $ 2 \ z $ 2 \ a $ 2 - 1" - using prems(2) - unfolding * assms - by (auto simp add: field_simps) - then show False - unfolding * using ab by auto - qed - then have "z \ path_image f \ z \ path_image g" - using z unfolding Un_iff by blast - then have z': "z \ cbox a b" - using assms(3-4) - by auto - have "a $ 2 = z $ 2 \ (z $ 1 = pathstart f $ 1 \ z $ 1 = pathfinish f $ 1) \ - z = pathstart f \ z = pathfinish f" - unfolding vec_eq_iff forall_2 assms - by auto - with z' show "z \ path_image f" - using z(1) - unfolding Un_iff mem_interval_cart forall_2 - apply - - apply (simp only: segment_vertical segment_horizontal vector_2) - unfolding assms - apply auto - done - have "a $ 2 = z $ 2 \ (z $ 1 = pathstart g $ 1 \ z $ 1 = pathfinish g $ 1) \ - z = pathstart g \ z = pathfinish g" - unfolding vec_eq_iff forall_2 assms - by auto - with z' show "z \ path_image g" - using z(2) - unfolding Un_iff mem_interval_cart forall_2 - apply (simp only: segment_vertical segment_horizontal vector_2) - unfolding assms - apply auto - done - qed -qed - -(** The Following still needs to be translated. Maybe I will do that later. - -(* ------------------------------------------------------------------------- *) -(* Complement in dimension N >= 2 of set homeomorphic to any interval in *) -(* any dimension is (path-)connected. This naively generalizes the argument *) -(* in Ryuji Maehara's paper "The Jordan curve theorem via the Brouwer *) -(* fixed point theorem", American Mathematical Monthly 1984. *) -(* ------------------------------------------------------------------------- *) - -let RETRACTION_INJECTIVE_IMAGE_INTERVAL = prove - (`!p:real^M->real^N a b. - ~(interval[a,b] = {}) /\ - p continuous_on interval[a,b] /\ - (!x y. x IN interval[a,b] /\ y IN interval[a,b] /\ p x = p y ==> x = y) - ==> ?f. f continuous_on (:real^N) /\ - IMAGE f (:real^N) SUBSET (IMAGE p (interval[a,b])) /\ - (!x. x IN (IMAGE p (interval[a,b])) ==> f x = x)`, - REPEAT STRIP_TAC THEN - FIRST_ASSUM(MP_TAC o GEN_REWRITE_RULE I [INJECTIVE_ON_LEFT_INVERSE]) THEN - DISCH_THEN(X_CHOOSE_TAC `q:real^N->real^M`) THEN - SUBGOAL_THEN `(q:real^N->real^M) continuous_on - (IMAGE p (interval[a:real^M,b]))` - ASSUME_TAC THENL - [MATCH_MP_TAC CONTINUOUS_ON_INVERSE THEN ASM_REWRITE_TAC[COMPACT_INTERVAL]; - ALL_TAC] THEN - MP_TAC(ISPECL [`q:real^N->real^M`; - `IMAGE (p:real^M->real^N) - (interval[a,b])`; - `a:real^M`; `b:real^M`] - TIETZE_CLOSED_INTERVAL) THEN - ASM_SIMP_TAC[COMPACT_INTERVAL; COMPACT_CONTINUOUS_IMAGE; - COMPACT_IMP_CLOSED] THEN - ANTS_TAC THENL [ASM SET_TAC[]; ALL_TAC] THEN - DISCH_THEN(X_CHOOSE_THEN `r:real^N->real^M` STRIP_ASSUME_TAC) THEN - EXISTS_TAC `(p:real^M->real^N) o (r:real^N->real^M)` THEN - REWRITE_TAC[SUBSET; FORALL_IN_IMAGE; o_THM; IN_UNIV] THEN - CONJ_TAC THENL [ALL_TAC; ASM SET_TAC[]] THEN - MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN ASM_REWRITE_TAC[] THEN - FIRST_X_ASSUM(MATCH_MP_TAC o MATCH_MP(REWRITE_RULE[IMP_CONJ] - CONTINUOUS_ON_SUBSET)) THEN ASM SET_TAC[]);; - -let UNBOUNDED_PATH_COMPONENTS_COMPLEMENT_HOMEOMORPHIC_INTERVAL = prove - (`!s:real^N->bool a b:real^M. - s homeomorphic (interval[a,b]) - ==> !x. ~(x IN s) ==> ~bounded(path_component((:real^N) DIFF s) x)`, - REPEAT GEN_TAC THEN REWRITE_TAC[homeomorphic; homeomorphism] THEN - REWRITE_TAC[LEFT_IMP_EXISTS_THM] THEN - MAP_EVERY X_GEN_TAC [`p':real^N->real^M`; `p:real^M->real^N`] THEN - DISCH_TAC THEN - SUBGOAL_THEN - `!x y. x IN interval[a,b] /\ y IN interval[a,b] /\ - (p:real^M->real^N) x = p y ==> x = y` - ASSUME_TAC THENL [ASM_MESON_TAC[]; ALL_TAC] THEN - FIRST_X_ASSUM(MP_TAC o funpow 4 CONJUNCT2) THEN - DISCH_THEN(CONJUNCTS_THEN2 (SUBST1_TAC o SYM) ASSUME_TAC) THEN - ASM_CASES_TAC `interval[a:real^M,b] = {}` THEN - ASM_REWRITE_TAC[IMAGE_CLAUSES; DIFF_EMPTY; PATH_COMPONENT_UNIV; - NOT_BOUNDED_UNIV] THEN - ABBREV_TAC `s = (:real^N) DIFF (IMAGE p (interval[a:real^M,b]))` THEN - X_GEN_TAC `c:real^N` THEN REPEAT STRIP_TAC THEN - SUBGOAL_THEN `(c:real^N) IN s` ASSUME_TAC THENL [ASM SET_TAC[]; ALL_TAC] THEN - SUBGOAL_THEN `bounded((path_component s c) UNION - (IMAGE (p:real^M->real^N) (interval[a,b])))` - MP_TAC THENL - [ASM_SIMP_TAC[BOUNDED_UNION; COMPACT_IMP_BOUNDED; COMPACT_IMP_BOUNDED; - COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL]; - ALL_TAC] THEN - DISCH_THEN(MP_TAC o SPEC `c:real^N` o MATCH_MP BOUNDED_SUBSET_BALL) THEN - REWRITE_TAC[UNION_SUBSET] THEN - DISCH_THEN(X_CHOOSE_THEN `B:real` STRIP_ASSUME_TAC) THEN - MP_TAC(ISPECL [`p:real^M->real^N`; `a:real^M`; `b:real^M`] - RETRACTION_INJECTIVE_IMAGE_INTERVAL) THEN - ASM_REWRITE_TAC[SUBSET; IN_UNIV] THEN - DISCH_THEN(X_CHOOSE_THEN `r:real^N->real^N` MP_TAC) THEN - DISCH_THEN(CONJUNCTS_THEN2 ASSUME_TAC - (CONJUNCTS_THEN2 MP_TAC ASSUME_TAC)) THEN - REWRITE_TAC[FORALL_IN_IMAGE; IN_UNIV] THEN DISCH_TAC THEN - ABBREV_TAC `q = \z:real^N. if z IN path_component s c then r(z) else z` THEN - SUBGOAL_THEN - `(q:real^N->real^N) continuous_on - (closure(path_component s c) UNION ((:real^N) DIFF (path_component s c)))` - MP_TAC THENL - [EXPAND_TAC "q" THEN MATCH_MP_TAC CONTINUOUS_ON_CASES THEN - REWRITE_TAC[CLOSED_CLOSURE; CONTINUOUS_ON_ID; GSYM OPEN_CLOSED] THEN - REPEAT CONJ_TAC THENL - [MATCH_MP_TAC OPEN_PATH_COMPONENT THEN EXPAND_TAC "s" THEN - ASM_SIMP_TAC[GSYM CLOSED_OPEN; COMPACT_IMP_CLOSED; - COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL]; - ASM_MESON_TAC[CONTINUOUS_ON_SUBSET; SUBSET_UNIV]; - ALL_TAC] THEN - X_GEN_TAC `z:real^N` THEN - REWRITE_TAC[SET_RULE `~(z IN (s DIFF t) /\ z IN t)`] THEN - STRIP_TAC THEN FIRST_X_ASSUM MATCH_MP_TAC THEN - MP_TAC(ISPECL - [`path_component s (z:real^N)`; `path_component s (c:real^N)`] - OPEN_INTER_CLOSURE_EQ_EMPTY) THEN - ASM_REWRITE_TAC[GSYM DISJOINT; PATH_COMPONENT_DISJOINT] THEN ANTS_TAC THENL - [MATCH_MP_TAC OPEN_PATH_COMPONENT THEN EXPAND_TAC "s" THEN - ASM_SIMP_TAC[GSYM CLOSED_OPEN; COMPACT_IMP_CLOSED; - COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL]; - REWRITE_TAC[DISJOINT; EXTENSION; IN_INTER; NOT_IN_EMPTY] THEN - DISCH_THEN(MP_TAC o SPEC `z:real^N`) THEN ASM_REWRITE_TAC[] THEN - GEN_REWRITE_TAC (LAND_CONV o ONCE_DEPTH_CONV) [IN] THEN - REWRITE_TAC[PATH_COMPONENT_REFL_EQ] THEN ASM SET_TAC[]]; - ALL_TAC] THEN - SUBGOAL_THEN - `closure(path_component s c) UNION ((:real^N) DIFF (path_component s c)) = - (:real^N)` - SUBST1_TAC THENL - [MATCH_MP_TAC(SET_RULE `s SUBSET t ==> t UNION (UNIV DIFF s) = UNIV`) THEN - REWRITE_TAC[CLOSURE_SUBSET]; - DISCH_TAC] THEN - MP_TAC(ISPECL - [`(\x. &2 % c - x) o - (\x. c + B / norm(x - c) % (x - c)) o (q:real^N->real^N)`; - `cball(c:real^N,B)`] - BROUWER) THEN - REWRITE_TAC[NOT_IMP; GSYM CONJ_ASSOC; COMPACT_CBALL; CONVEX_CBALL] THEN - ASM_SIMP_TAC[CBALL_EQ_EMPTY; REAL_LT_IMP_LE; REAL_NOT_LT] THEN - SUBGOAL_THEN `!x. ~((q:real^N->real^N) x = c)` ASSUME_TAC THENL - [X_GEN_TAC `x:real^N` THEN EXPAND_TAC "q" THEN - REWRITE_TAC[NORM_EQ_0; VECTOR_SUB_EQ] THEN COND_CASES_TAC THEN - ASM SET_TAC[PATH_COMPONENT_REFL_EQ]; - ALL_TAC] THEN - REPEAT CONJ_TAC THENL - [MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN - SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST] THEN - MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN CONJ_TAC THENL - [ASM_MESON_TAC[CONTINUOUS_ON_SUBSET; SUBSET_UNIV]; ALL_TAC] THEN - MATCH_MP_TAC CONTINUOUS_ON_ADD THEN REWRITE_TAC[CONTINUOUS_ON_CONST] THEN - MATCH_MP_TAC CONTINUOUS_ON_MUL THEN - SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST] THEN - REWRITE_TAC[o_DEF; real_div; LIFT_CMUL] THEN - MATCH_MP_TAC CONTINUOUS_ON_CMUL THEN - MATCH_MP_TAC(REWRITE_RULE[o_DEF] CONTINUOUS_ON_INV) THEN - ASM_REWRITE_TAC[FORALL_IN_IMAGE; NORM_EQ_0; VECTOR_SUB_EQ] THEN - SUBGOAL_THEN - `(\x:real^N. lift(norm(x - c))) = (lift o norm) o (\x. x - c)` - SUBST1_TAC THENL [REWRITE_TAC[o_DEF]; ALL_TAC] THEN - MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN - ASM_SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST; - CONTINUOUS_ON_LIFT_NORM]; - REWRITE_TAC[SUBSET; FORALL_IN_IMAGE; IN_CBALL; o_THM; dist] THEN - X_GEN_TAC `x:real^N` THEN DISCH_TAC THEN - REWRITE_TAC[VECTOR_ARITH `c - (&2 % c - (c + x)) = x`] THEN - REWRITE_TAC[NORM_MUL; REAL_ABS_DIV; REAL_ABS_NORM] THEN - ASM_SIMP_TAC[REAL_DIV_RMUL; NORM_EQ_0; VECTOR_SUB_EQ] THEN - ASM_REAL_ARITH_TAC; - REWRITE_TAC[NOT_EXISTS_THM; TAUT `~(c /\ b) <=> c ==> ~b`] THEN - REWRITE_TAC[IN_CBALL; o_THM; dist] THEN - X_GEN_TAC `x:real^N` THEN DISCH_TAC THEN - REWRITE_TAC[VECTOR_ARITH `&2 % c - (c + x') = x <=> --x' = x - c`] THEN - ASM_CASES_TAC `(x:real^N) IN path_component s c` THENL - [MATCH_MP_TAC(NORM_ARITH `norm(y) < B /\ norm(x) = B ==> ~(--x = y)`) THEN - REWRITE_TAC[NORM_MUL; REAL_ABS_DIV; REAL_ABS_NORM] THEN - ASM_SIMP_TAC[REAL_DIV_RMUL; NORM_EQ_0; VECTOR_SUB_EQ] THEN - ASM_SIMP_TAC[REAL_ARITH `&0 < B ==> abs B = B`] THEN - UNDISCH_TAC `path_component s c SUBSET ball(c:real^N,B)` THEN - REWRITE_TAC[SUBSET; IN_BALL] THEN ASM_MESON_TAC[dist; NORM_SUB]; - EXPAND_TAC "q" THEN REWRITE_TAC[] THEN ASM_REWRITE_TAC[] THEN - REWRITE_TAC[VECTOR_ARITH `--(c % x) = x <=> (&1 + c) % x = vec 0`] THEN - ASM_REWRITE_TAC[DE_MORGAN_THM; VECTOR_SUB_EQ; VECTOR_MUL_EQ_0] THEN - SUBGOAL_THEN `~(x:real^N = c)` ASSUME_TAC THENL - [ASM_MESON_TAC[PATH_COMPONENT_REFL; IN]; ALL_TAC] THEN - ASM_REWRITE_TAC[] THEN - MATCH_MP_TAC(REAL_ARITH `&0 < x ==> ~(&1 + x = &0)`) THEN - ASM_SIMP_TAC[REAL_LT_DIV; NORM_POS_LT; VECTOR_SUB_EQ]]]);; - -let PATH_CONNECTED_COMPLEMENT_HOMEOMORPHIC_INTERVAL = prove - (`!s:real^N->bool a b:real^M. - 2 <= dimindex(:N) /\ s homeomorphic interval[a,b] - ==> path_connected((:real^N) DIFF s)`, - REPEAT STRIP_TAC THEN REWRITE_TAC[PATH_CONNECTED_IFF_PATH_COMPONENT] THEN - FIRST_ASSUM(MP_TAC o MATCH_MP - UNBOUNDED_PATH_COMPONENTS_COMPLEMENT_HOMEOMORPHIC_INTERVAL) THEN - ASM_REWRITE_TAC[SET_RULE `~(x IN s) <=> x IN (UNIV DIFF s)`] THEN - ABBREV_TAC `t = (:real^N) DIFF s` THEN - DISCH_TAC THEN MAP_EVERY X_GEN_TAC [`x:real^N`; `y:real^N`] THEN - STRIP_TAC THEN FIRST_ASSUM(MP_TAC o MATCH_MP HOMEOMORPHIC_COMPACTNESS) THEN - REWRITE_TAC[COMPACT_INTERVAL] THEN - DISCH_THEN(MP_TAC o MATCH_MP COMPACT_IMP_BOUNDED) THEN - REWRITE_TAC[BOUNDED_POS; LEFT_IMP_EXISTS_THM] THEN - X_GEN_TAC `B:real` THEN STRIP_TAC THEN - SUBGOAL_THEN `(?u:real^N. u IN path_component t x /\ B < norm(u)) /\ - (?v:real^N. v IN path_component t y /\ B < norm(v))` - STRIP_ASSUME_TAC THENL - [ASM_MESON_TAC[BOUNDED_POS; REAL_NOT_LE]; ALL_TAC] THEN - MATCH_MP_TAC PATH_COMPONENT_TRANS THEN EXISTS_TAC `u:real^N` THEN - CONJ_TAC THENL [ASM_MESON_TAC[IN]; ALL_TAC] THEN - MATCH_MP_TAC PATH_COMPONENT_SYM THEN - MATCH_MP_TAC PATH_COMPONENT_TRANS THEN EXISTS_TAC `v:real^N` THEN - CONJ_TAC THENL [ASM_MESON_TAC[IN]; ALL_TAC] THEN - MATCH_MP_TAC PATH_COMPONENT_OF_SUBSET THEN - EXISTS_TAC `(:real^N) DIFF cball(vec 0,B)` THEN CONJ_TAC THENL - [EXPAND_TAC "t" THEN MATCH_MP_TAC(SET_RULE - `s SUBSET t ==> (u DIFF t) SUBSET (u DIFF s)`) THEN - ASM_REWRITE_TAC[SUBSET; IN_CBALL_0]; - MP_TAC(ISPEC `cball(vec 0:real^N,B)` - PATH_CONNECTED_COMPLEMENT_BOUNDED_CONVEX) THEN - ASM_REWRITE_TAC[BOUNDED_CBALL; CONVEX_CBALL] THEN - REWRITE_TAC[PATH_CONNECTED_IFF_PATH_COMPONENT] THEN - DISCH_THEN MATCH_MP_TAC THEN - ASM_REWRITE_TAC[IN_DIFF; IN_UNIV; IN_CBALL_0; REAL_NOT_LE]]);; - -(* ------------------------------------------------------------------------- *) -(* In particular, apply all these to the special case of an arc. *) -(* ------------------------------------------------------------------------- *) - -let RETRACTION_ARC = prove - (`!p. arc p - ==> ?f. f continuous_on (:real^N) /\ - IMAGE f (:real^N) SUBSET path_image p /\ - (!x. x IN path_image p ==> f x = x)`, - REWRITE_TAC[arc; path; path_image] THEN REPEAT STRIP_TAC THEN - MATCH_MP_TAC RETRACTION_INJECTIVE_IMAGE_INTERVAL THEN - ASM_REWRITE_TAC[INTERVAL_EQ_EMPTY_CART_1; DROP_VEC; REAL_NOT_LT; REAL_POS]);; - -let PATH_CONNECTED_ARC_COMPLEMENT = prove - (`!p. 2 <= dimindex(:N) /\ arc p - ==> path_connected((:real^N) DIFF path_image p)`, - REWRITE_TAC[arc; path] THEN REPEAT STRIP_TAC THEN SIMP_TAC[path_image] THEN - MP_TAC(ISPECL [`path_image p:real^N->bool`; `vec 0:real^1`; `vec 1:real^1`] - PATH_CONNECTED_COMPLEMENT_HOMEOMORPHIC_INTERVAL) THEN - ASM_REWRITE_TAC[path_image] THEN DISCH_THEN MATCH_MP_TAC THEN - ONCE_REWRITE_TAC[HOMEOMORPHIC_SYM] THEN - MATCH_MP_TAC HOMEOMORPHIC_COMPACT THEN - EXISTS_TAC `p:real^1->real^N` THEN ASM_REWRITE_TAC[COMPACT_INTERVAL]);; - -let CONNECTED_ARC_COMPLEMENT = prove - (`!p. 2 <= dimindex(:N) /\ arc p - ==> connected((:real^N) DIFF path_image p)`, - SIMP_TAC[PATH_CONNECTED_ARC_COMPLEMENT; PATH_CONNECTED_IMP_CONNECTED]);; *) - -end diff -r bbcb05504fdc -r bd218a9320b5 src/HOL/Multivariate_Analysis/Fashoda_Theorem.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Multivariate_Analysis/Fashoda_Theorem.thy Thu Aug 04 19:36:31 2016 +0200 @@ -0,0 +1,1136 @@ +(* Author: John Harrison + Author: Robert Himmelmann, TU Muenchen (translation from HOL light) +*) + +section \Fashoda meet theorem\ + +theory Fashoda_Theorem +imports Brouwer_Fixpoint Path_Connected Cartesian_Euclidean_Space +begin + +subsection \Bijections between intervals.\ + +definition interval_bij :: "'a \ 'a \ 'a \ 'a \ 'a \ 'a::euclidean_space" + where "interval_bij = + (\(a, b) (u, v) x. (\i\Basis. (u\i + (x\i - a\i) / (b\i - a\i) * (v\i - u\i)) *\<^sub>R i))" + +lemma interval_bij_affine: + "interval_bij (a,b) (u,v) = (\x. (\i\Basis. ((v\i - u\i) / (b\i - a\i) * (x\i)) *\<^sub>R i) + + (\i\Basis. (u\i - (v\i - u\i) / (b\i - a\i) * (a\i)) *\<^sub>R i))" + by (auto simp: setsum.distrib[symmetric] scaleR_add_left[symmetric] interval_bij_def fun_eq_iff + field_simps inner_simps add_divide_distrib[symmetric] intro!: setsum.cong) + +lemma continuous_interval_bij: + fixes a b :: "'a::euclidean_space" + shows "continuous (at x) (interval_bij (a, b) (u, v))" + by (auto simp add: divide_inverse interval_bij_def intro!: continuous_setsum continuous_intros) + +lemma continuous_on_interval_bij: "continuous_on s (interval_bij (a, b) (u, v))" + apply(rule continuous_at_imp_continuous_on) + apply (rule, rule continuous_interval_bij) + done + +lemma in_interval_interval_bij: + fixes a b u v x :: "'a::euclidean_space" + assumes "x \ cbox a b" + and "cbox u v \ {}" + shows "interval_bij (a, b) (u, v) x \ cbox u v" + apply (simp only: interval_bij_def split_conv mem_box inner_setsum_left_Basis cong: ball_cong) + apply safe +proof - + fix i :: 'a + assume i: "i \ Basis" + have "cbox a b \ {}" + using assms by auto + with i have *: "a\i \ b\i" "u\i \ v\i" + using assms(2) by (auto simp add: box_eq_empty) + have x: "a\i\x\i" "x\i\b\i" + using assms(1)[unfolded mem_box] using i by auto + have "0 \ (x \ i - a \ i) / (b \ i - a \ i) * (v \ i - u \ i)" + using * x by auto + then show "u \ i \ u \ i + (x \ i - a \ i) / (b \ i - a \ i) * (v \ i - u \ i)" + using * by auto + have "((x \ i - a \ i) / (b \ i - a \ i)) * (v \ i - u \ i) \ 1 * (v \ i - u \ i)" + apply (rule mult_right_mono) + unfolding divide_le_eq_1 + using * x + apply auto + done + then show "u \ i + (x \ i - a \ i) / (b \ i - a \ i) * (v \ i - u \ i) \ v \ i" + using * by auto +qed + +lemma interval_bij_bij: + "\(i::'a::euclidean_space)\Basis. a\i < b\i \ u\i < v\i \ + interval_bij (a, b) (u, v) (interval_bij (u, v) (a, b) x) = x" + by (auto simp: interval_bij_def euclidean_eq_iff[where 'a='a]) + +lemma interval_bij_bij_cart: fixes x::"real^'n" assumes "\i. a$i < b$i \ u$i < v$i" + shows "interval_bij (a,b) (u,v) (interval_bij (u,v) (a,b) x) = x" + using assms by (intro interval_bij_bij) (auto simp: Basis_vec_def inner_axis) + + +subsection \Fashoda meet theorem\ + +lemma infnorm_2: + fixes x :: "real^2" + shows "infnorm x = max \x$1\ \x$2\" + unfolding infnorm_cart UNIV_2 by (rule cSup_eq) auto + +lemma infnorm_eq_1_2: + fixes x :: "real^2" + shows "infnorm x = 1 \ + \x$1\ \ 1 \ \x$2\ \ 1 \ (x$1 = -1 \ x$1 = 1 \ x$2 = -1 \ x$2 = 1)" + unfolding infnorm_2 by auto + +lemma infnorm_eq_1_imp: + fixes x :: "real^2" + assumes "infnorm x = 1" + shows "\x$1\ \ 1" and "\x$2\ \ 1" + using assms unfolding infnorm_eq_1_2 by auto + +lemma fashoda_unit: + fixes f g :: "real \ real^2" + assumes "f ` {-1 .. 1} \ cbox (-1) 1" + and "g ` {-1 .. 1} \ cbox (-1) 1" + and "continuous_on {-1 .. 1} f" + and "continuous_on {-1 .. 1} g" + and "f (- 1)$1 = - 1" + and "f 1$1 = 1" "g (- 1) $2 = -1" + and "g 1 $2 = 1" + shows "\s\{-1 .. 1}. \t\{-1 .. 1}. f s = g t" +proof (rule ccontr) + assume "\ ?thesis" + note as = this[unfolded bex_simps,rule_format] + define sqprojection + where [abs_def]: "sqprojection z = (inverse (infnorm z)) *\<^sub>R z" for z :: "real^2" + define negatex :: "real^2 \ real^2" + where "negatex x = (vector [-(x$1), x$2])" for x + have lem1: "\z::real^2. infnorm (negatex z) = infnorm z" + unfolding negatex_def infnorm_2 vector_2 by auto + have lem2: "\z. z \ 0 \ infnorm (sqprojection z) = 1" + unfolding sqprojection_def + unfolding infnorm_mul[unfolded scalar_mult_eq_scaleR] + unfolding abs_inverse real_abs_infnorm + apply (subst infnorm_eq_0[symmetric]) + apply auto + done + let ?F = "\w::real^2. (f \ (\x. x$1)) w - (g \ (\x. x$2)) w" + have *: "\i. (\x::real^2. x $ i) ` cbox (- 1) 1 = {-1 .. 1}" + apply (rule set_eqI) + unfolding image_iff Bex_def mem_interval_cart interval_cbox_cart + apply rule + defer + apply (rule_tac x="vec x" in exI) + apply auto + done + { + fix x + assume "x \ (\w. (f \ (\x. x $ 1)) w - (g \ (\x. x $ 2)) w) ` (cbox (- 1) (1::real^2))" + then obtain w :: "real^2" where w: + "w \ cbox (- 1) 1" + "x = (f \ (\x. x $ 1)) w - (g \ (\x. x $ 2)) w" + unfolding image_iff .. + then have "x \ 0" + using as[of "w$1" "w$2"] + unfolding mem_interval_cart atLeastAtMost_iff + by auto + } note x0 = this + have 21: "\i::2. i \ 1 \ i = 2" + using UNIV_2 by auto + have 1: "box (- 1) (1::real^2) \ {}" + unfolding interval_eq_empty_cart by auto + have 2: "continuous_on (cbox (- 1) 1) (negatex \ sqprojection \ ?F)" + apply (intro continuous_intros continuous_on_component) + unfolding * + apply (rule assms)+ + apply (subst sqprojection_def) + apply (intro continuous_intros) + apply (simp add: infnorm_eq_0 x0) + apply (rule linear_continuous_on) + proof - + show "bounded_linear negatex" + apply (rule bounded_linearI') + unfolding vec_eq_iff + proof (rule_tac[!] allI) + fix i :: 2 + fix x y :: "real^2" + fix c :: real + show "negatex (x + y) $ i = + (negatex x + negatex y) $ i" "negatex (c *\<^sub>R x) $ i = (c *\<^sub>R negatex x) $ i" + apply - + apply (case_tac[!] "i\1") + prefer 3 + apply (drule_tac[1-2] 21) + unfolding negatex_def + apply (auto simp add:vector_2) + done + qed + qed + have 3: "(negatex \ sqprojection \ ?F) ` cbox (-1) 1 \ cbox (-1) 1" + unfolding subset_eq + proof (rule, goal_cases) + case (1 x) + then obtain y :: "real^2" where y: + "y \ cbox (- 1) 1" + "x = (negatex \ sqprojection \ (\w. (f \ (\x. x $ 1)) w - (g \ (\x. x $ 2)) w)) y" + unfolding image_iff .. + have "?F y \ 0" + apply (rule x0) + using y(1) + apply auto + done + then have *: "infnorm (sqprojection (?F y)) = 1" + unfolding y o_def + by - (rule lem2[rule_format]) + have "infnorm x = 1" + unfolding *[symmetric] y o_def + by (rule lem1[rule_format]) + then show "x \ cbox (-1) 1" + unfolding mem_interval_cart interval_cbox_cart infnorm_2 + apply - + apply rule + proof - + fix i + assume "max \x $ 1\ \x $ 2\ = 1" + then show "(- 1) $ i \ x $ i \ x $ i \ 1 $ i" + apply (cases "i = 1") + defer + apply (drule 21) + apply auto + done + qed + qed + obtain x :: "real^2" where x: + "x \ cbox (- 1) 1" + "(negatex \ sqprojection \ (\w. (f \ (\x. x $ 1)) w - (g \ (\x. x $ 2)) w)) x = x" + apply (rule brouwer_weak[of "cbox (- 1) (1::real^2)" "negatex \ sqprojection \ ?F"]) + apply (rule compact_cbox convex_box)+ + unfolding interior_cbox + apply (rule 1 2 3)+ + apply blast + done + have "?F x \ 0" + apply (rule x0) + using x(1) + apply auto + done + then have *: "infnorm (sqprojection (?F x)) = 1" + unfolding o_def + by (rule lem2[rule_format]) + have nx: "infnorm x = 1" + apply (subst x(2)[symmetric]) + unfolding *[symmetric] o_def + apply (rule lem1[rule_format]) + done + have "\x i. x \ 0 \ (0 < (sqprojection x)$i \ 0 < x$i)" + and "\x i. x \ 0 \ ((sqprojection x)$i < 0 \ x$i < 0)" + apply - + apply (rule_tac[!] allI impI)+ + proof - + fix x :: "real^2" + fix i :: 2 + assume x: "x \ 0" + have "inverse (infnorm x) > 0" + using x[unfolded infnorm_pos_lt[symmetric]] by auto + then show "(0 < sqprojection x $ i) = (0 < x $ i)" + and "(sqprojection x $ i < 0) = (x $ i < 0)" + unfolding sqprojection_def vector_component_simps vector_scaleR_component real_scaleR_def + unfolding zero_less_mult_iff mult_less_0_iff + by (auto simp add: field_simps) + qed + note lem3 = this[rule_format] + have x1: "x $ 1 \ {- 1..1::real}" "x $ 2 \ {- 1..1::real}" + using x(1) unfolding mem_interval_cart by auto + then have nz: "f (x $ 1) - g (x $ 2) \ 0" + unfolding right_minus_eq + apply - + apply (rule as) + apply auto + done + have "x $ 1 = -1 \ x $ 1 = 1 \ x $ 2 = -1 \ x $ 2 = 1" + using nx unfolding infnorm_eq_1_2 by auto + then show False + proof - + fix P Q R S + presume "P \ Q \ R \ S" + and "P \ False" + and "Q \ False" + and "R \ False" + and "S \ False" + then show False by auto + next + assume as: "x$1 = 1" + then have *: "f (x $ 1) $ 1 = 1" + using assms(6) by auto + have "sqprojection (f (x$1) - g (x$2)) $ 1 < 0" + using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=1]] + unfolding as negatex_def vector_2 + by auto + moreover + from x1 have "g (x $ 2) \ cbox (-1) 1" + apply - + apply (rule assms(2)[unfolded subset_eq,rule_format]) + apply auto + done + ultimately show False + unfolding lem3[OF nz] vector_component_simps * mem_interval_cart + apply (erule_tac x=1 in allE) + apply auto + done + next + assume as: "x$1 = -1" + then have *: "f (x $ 1) $ 1 = - 1" + using assms(5) by auto + have "sqprojection (f (x$1) - g (x$2)) $ 1 > 0" + using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=1]] + unfolding as negatex_def vector_2 + by auto + moreover + from x1 have "g (x $ 2) \ cbox (-1) 1" + apply - + apply (rule assms(2)[unfolded subset_eq,rule_format]) + apply auto + done + ultimately show False + unfolding lem3[OF nz] vector_component_simps * mem_interval_cart + apply (erule_tac x=1 in allE) + apply auto + done + next + assume as: "x$2 = 1" + then have *: "g (x $ 2) $ 2 = 1" + using assms(8) by auto + have "sqprojection (f (x$1) - g (x$2)) $ 2 > 0" + using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=2]] + unfolding as negatex_def vector_2 + by auto + moreover + from x1 have "f (x $ 1) \ cbox (-1) 1" + apply - + apply (rule assms(1)[unfolded subset_eq,rule_format]) + apply auto + done + ultimately show False + unfolding lem3[OF nz] vector_component_simps * mem_interval_cart + apply (erule_tac x=2 in allE) + apply auto + done + next + assume as: "x$2 = -1" + then have *: "g (x $ 2) $ 2 = - 1" + using assms(7) by auto + have "sqprojection (f (x$1) - g (x$2)) $ 2 < 0" + using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=2]] + unfolding as negatex_def vector_2 + by auto + moreover + from x1 have "f (x $ 1) \ cbox (-1) 1" + apply - + apply (rule assms(1)[unfolded subset_eq,rule_format]) + apply auto + done + ultimately show False + unfolding lem3[OF nz] vector_component_simps * mem_interval_cart + apply (erule_tac x=2 in allE) + apply auto + done + qed auto +qed + +lemma fashoda_unit_path: + fixes f g :: "real \ real^2" + assumes "path f" + and "path g" + and "path_image f \ cbox (-1) 1" + and "path_image g \ cbox (-1) 1" + and "(pathstart f)$1 = -1" + and "(pathfinish f)$1 = 1" + and "(pathstart g)$2 = -1" + and "(pathfinish g)$2 = 1" + obtains z where "z \ path_image f" and "z \ path_image g" +proof - + note assms=assms[unfolded path_def pathstart_def pathfinish_def path_image_def] + define iscale where [abs_def]: "iscale z = inverse 2 *\<^sub>R (z + 1)" for z :: real + have isc: "iscale ` {- 1..1} \ {0..1}" + unfolding iscale_def by auto + have "\s\{- 1..1}. \t\{- 1..1}. (f \ iscale) s = (g \ iscale) t" + proof (rule fashoda_unit) + show "(f \ iscale) ` {- 1..1} \ cbox (- 1) 1" "(g \ iscale) ` {- 1..1} \ cbox (- 1) 1" + using isc and assms(3-4) by (auto simp add: image_comp [symmetric]) + have *: "continuous_on {- 1..1} iscale" + unfolding iscale_def by (rule continuous_intros)+ + show "continuous_on {- 1..1} (f \ iscale)" "continuous_on {- 1..1} (g \ iscale)" + apply - + apply (rule_tac[!] continuous_on_compose[OF *]) + apply (rule_tac[!] continuous_on_subset[OF _ isc]) + apply (rule assms)+ + done + have *: "(1 / 2) *\<^sub>R (1 + (1::real^1)) = 1" + unfolding vec_eq_iff by auto + show "(f \ iscale) (- 1) $ 1 = - 1" + and "(f \ iscale) 1 $ 1 = 1" + and "(g \ iscale) (- 1) $ 2 = -1" + and "(g \ iscale) 1 $ 2 = 1" + unfolding o_def iscale_def + using assms + by (auto simp add: *) + qed + then obtain s t where st: + "s \ {- 1..1}" + "t \ {- 1..1}" + "(f \ iscale) s = (g \ iscale) t" + by auto + show thesis + apply (rule_tac z = "f (iscale s)" in that) + using st + unfolding o_def path_image_def image_iff + apply - + apply (rule_tac x="iscale s" in bexI) + prefer 3 + apply (rule_tac x="iscale t" in bexI) + using isc[unfolded subset_eq, rule_format] + apply auto + done +qed + +lemma fashoda: + fixes b :: "real^2" + assumes "path f" + and "path g" + and "path_image f \ cbox a b" + and "path_image g \ cbox a b" + and "(pathstart f)$1 = a$1" + and "(pathfinish f)$1 = b$1" + and "(pathstart g)$2 = a$2" + and "(pathfinish g)$2 = b$2" + obtains z where "z \ path_image f" and "z \ path_image g" +proof - + fix P Q S + presume "P \ Q \ S" "P \ thesis" and "Q \ thesis" and "S \ thesis" + then show thesis + by auto +next + have "cbox a b \ {}" + using assms(3) using path_image_nonempty[of f] by auto + then have "a \ b" + unfolding interval_eq_empty_cart less_eq_vec_def by (auto simp add: not_less) + then show "a$1 = b$1 \ a$2 = b$2 \ (a$1 < b$1 \ a$2 < b$2)" + unfolding less_eq_vec_def forall_2 by auto +next + assume as: "a$1 = b$1" + have "\z\path_image g. z$2 = (pathstart f)$2" + apply (rule connected_ivt_component_cart) + apply (rule connected_path_image assms)+ + apply (rule pathstart_in_path_image) + apply (rule pathfinish_in_path_image) + unfolding assms using assms(3)[unfolded path_image_def subset_eq,rule_format,of "f 0"] + unfolding pathstart_def + apply (auto simp add: less_eq_vec_def mem_interval_cart) + done + then obtain z :: "real^2" where z: "z \ path_image g" "z $ 2 = pathstart f $ 2" .. + have "z \ cbox a b" + using z(1) assms(4) + unfolding path_image_def + by blast + then have "z = f 0" + unfolding vec_eq_iff forall_2 + unfolding z(2) pathstart_def + using assms(3)[unfolded path_image_def subset_eq mem_interval_cart,rule_format,of "f 0" 1] + unfolding mem_interval_cart + apply (erule_tac x=1 in allE) + using as + apply auto + done + then show thesis + apply - + apply (rule that[OF _ z(1)]) + unfolding path_image_def + apply auto + done +next + assume as: "a$2 = b$2" + have "\z\path_image f. z$1 = (pathstart g)$1" + apply (rule connected_ivt_component_cart) + apply (rule connected_path_image assms)+ + apply (rule pathstart_in_path_image) + apply (rule pathfinish_in_path_image) + unfolding assms + using assms(4)[unfolded path_image_def subset_eq,rule_format,of "g 0"] + unfolding pathstart_def + apply (auto simp add: less_eq_vec_def mem_interval_cart) + done + then obtain z where z: "z \ path_image f" "z $ 1 = pathstart g $ 1" .. + have "z \ cbox a b" + using z(1) assms(3) + unfolding path_image_def + by blast + then have "z = g 0" + unfolding vec_eq_iff forall_2 + unfolding z(2) pathstart_def + using assms(4)[unfolded path_image_def subset_eq mem_interval_cart,rule_format,of "g 0" 2] + unfolding mem_interval_cart + apply (erule_tac x=2 in allE) + using as + apply auto + done + then show thesis + apply - + apply (rule that[OF z(1)]) + unfolding path_image_def + apply auto + done +next + assume as: "a $ 1 < b $ 1 \ a $ 2 < b $ 2" + have int_nem: "cbox (-1) (1::real^2) \ {}" + unfolding interval_eq_empty_cart by auto + obtain z :: "real^2" where z: + "z \ (interval_bij (a, b) (- 1, 1) \ f) ` {0..1}" + "z \ (interval_bij (a, b) (- 1, 1) \ g) ` {0..1}" + apply (rule fashoda_unit_path[of "interval_bij (a,b) (- 1,1) \ f" "interval_bij (a,b) (- 1,1) \ g"]) + unfolding path_def path_image_def pathstart_def pathfinish_def + apply (rule_tac[1-2] continuous_on_compose) + apply (rule assms[unfolded path_def] continuous_on_interval_bij)+ + unfolding subset_eq + apply(rule_tac[1-2] ballI) + proof - + fix x + assume "x \ (interval_bij (a, b) (- 1, 1) \ f) ` {0..1}" + then obtain y where y: + "y \ {0..1}" + "x = (interval_bij (a, b) (- 1, 1) \ f) y" + unfolding image_iff .. + show "x \ cbox (- 1) 1" + unfolding y o_def + apply (rule in_interval_interval_bij) + using y(1) + using assms(3)[unfolded path_image_def subset_eq] int_nem + apply auto + done + next + fix x + assume "x \ (interval_bij (a, b) (- 1, 1) \ g) ` {0..1}" + then obtain y where y: + "y \ {0..1}" + "x = (interval_bij (a, b) (- 1, 1) \ g) y" + unfolding image_iff .. + show "x \ cbox (- 1) 1" + unfolding y o_def + apply (rule in_interval_interval_bij) + using y(1) + using assms(4)[unfolded path_image_def subset_eq] int_nem + apply auto + done + next + show "(interval_bij (a, b) (- 1, 1) \ f) 0 $ 1 = -1" + and "(interval_bij (a, b) (- 1, 1) \ f) 1 $ 1 = 1" + and "(interval_bij (a, b) (- 1, 1) \ g) 0 $ 2 = -1" + and "(interval_bij (a, b) (- 1, 1) \ g) 1 $ 2 = 1" + using assms as + by (simp_all add: axis_in_Basis cart_eq_inner_axis pathstart_def pathfinish_def interval_bij_def) + (simp_all add: inner_axis) + qed + from z(1) obtain zf where zf: + "zf \ {0..1}" + "z = (interval_bij (a, b) (- 1, 1) \ f) zf" + unfolding image_iff .. + from z(2) obtain zg where zg: + "zg \ {0..1}" + "z = (interval_bij (a, b) (- 1, 1) \ g) zg" + unfolding image_iff .. + have *: "\i. (- 1) $ i < (1::real^2) $ i \ a $ i < b $ i" + unfolding forall_2 + using as + by auto + show thesis + apply (rule_tac z="interval_bij (- 1,1) (a,b) z" in that) + apply (subst zf) + defer + apply (subst zg) + unfolding o_def interval_bij_bij_cart[OF *] path_image_def + using zf(1) zg(1) + apply auto + done +qed + + +subsection \Some slightly ad hoc lemmas I use below\ + +lemma segment_vertical: + fixes a :: "real^2" + assumes "a$1 = b$1" + shows "x \ closed_segment a b \ + x$1 = a$1 \ x$1 = b$1 \ (a$2 \ x$2 \ x$2 \ b$2 \ b$2 \ x$2 \ x$2 \ a$2)" + (is "_ = ?R") +proof - + let ?L = "\u. (x $ 1 = (1 - u) * a $ 1 + u * b $ 1 \ x $ 2 = (1 - u) * a $ 2 + u * b $ 2) \ 0 \ u \ u \ 1" + { + presume "?L \ ?R" and "?R \ ?L" + then show ?thesis + unfolding closed_segment_def mem_Collect_eq + unfolding vec_eq_iff forall_2 scalar_mult_eq_scaleR[symmetric] vector_component_simps + by blast + } + { + assume ?L + then obtain u where u: + "x $ 1 = (1 - u) * a $ 1 + u * b $ 1" + "x $ 2 = (1 - u) * a $ 2 + u * b $ 2" + "0 \ u" + "u \ 1" + by blast + { fix b a + assume "b + u * a > a + u * b" + then have "(1 - u) * b > (1 - u) * a" + by (auto simp add:field_simps) + then have "b \ a" + apply (drule_tac mult_left_less_imp_less) + using u + apply auto + done + then have "u * a \ u * b" + apply - + apply (rule mult_left_mono[OF _ u(3)]) + using u(3-4) + apply (auto simp add: field_simps) + done + } note * = this + { + fix a b + assume "u * b > u * a" + then have "(1 - u) * a \ (1 - u) * b" + apply - + apply (rule mult_left_mono) + apply (drule mult_left_less_imp_less) + using u + apply auto + done + then have "a + u * b \ b + u * a" + by (auto simp add: field_simps) + } note ** = this + then show ?R + unfolding u assms + using u + by (auto simp add:field_simps not_le intro: * **) + } + { + assume ?R + then show ?L + proof (cases "x$2 = b$2") + case True + then show ?L + apply (rule_tac x="(x$2 - a$2) / (b$2 - a$2)" in exI) + unfolding assms True + using \?R\ + apply (auto simp add: field_simps) + done + next + case False + then show ?L + apply (rule_tac x="1 - (x$2 - b$2) / (a$2 - b$2)" in exI) + unfolding assms + using \?R\ + apply (auto simp add: field_simps) + done + qed + } +qed + +lemma segment_horizontal: + fixes a :: "real^2" + assumes "a$2 = b$2" + shows "x \ closed_segment a b \ + x$2 = a$2 \ x$2 = b$2 \ (a$1 \ x$1 \ x$1 \ b$1 \ b$1 \ x$1 \ x$1 \ a$1)" + (is "_ = ?R") +proof - + let ?L = "\u. (x $ 1 = (1 - u) * a $ 1 + u * b $ 1 \ x $ 2 = (1 - u) * a $ 2 + u * b $ 2) \ 0 \ u \ u \ 1" + { + presume "?L \ ?R" and "?R \ ?L" + then show ?thesis + unfolding closed_segment_def mem_Collect_eq + unfolding vec_eq_iff forall_2 scalar_mult_eq_scaleR[symmetric] vector_component_simps + by blast + } + { + assume ?L + then obtain u where u: + "x $ 1 = (1 - u) * a $ 1 + u * b $ 1" + "x $ 2 = (1 - u) * a $ 2 + u * b $ 2" + "0 \ u" + "u \ 1" + by blast + { + fix b a + assume "b + u * a > a + u * b" + then have "(1 - u) * b > (1 - u) * a" + by (auto simp add: field_simps) + then have "b \ a" + apply (drule_tac mult_left_less_imp_less) + using u + apply auto + done + then have "u * a \ u * b" + apply - + apply (rule mult_left_mono[OF _ u(3)]) + using u(3-4) + apply (auto simp add: field_simps) + done + } note * = this + { + fix a b + assume "u * b > u * a" + then have "(1 - u) * a \ (1 - u) * b" + apply - + apply (rule mult_left_mono) + apply (drule mult_left_less_imp_less) + using u + apply auto + done + then have "a + u * b \ b + u * a" + by (auto simp add: field_simps) + } note ** = this + then show ?R + unfolding u assms + using u + by (auto simp add: field_simps not_le intro: * **) + } + { + assume ?R + then show ?L + proof (cases "x$1 = b$1") + case True + then show ?L + apply (rule_tac x="(x$1 - a$1) / (b$1 - a$1)" in exI) + unfolding assms True + using \?R\ + apply (auto simp add: field_simps) + done + next + case False + then show ?L + apply (rule_tac x="1 - (x$1 - b$1) / (a$1 - b$1)" in exI) + unfolding assms + using \?R\ + apply (auto simp add: field_simps) + done + qed + } +qed + + +subsection \Useful Fashoda corollary pointed out to me by Tom Hales\ + +lemma fashoda_interlace: + fixes a :: "real^2" + assumes "path f" + and "path g" + and "path_image f \ cbox a b" + and "path_image g \ cbox a b" + and "(pathstart f)$2 = a$2" + and "(pathfinish f)$2 = a$2" + and "(pathstart g)$2 = a$2" + and "(pathfinish g)$2 = a$2" + and "(pathstart f)$1 < (pathstart g)$1" + and "(pathstart g)$1 < (pathfinish f)$1" + and "(pathfinish f)$1 < (pathfinish g)$1" + obtains z where "z \ path_image f" and "z \ path_image g" +proof - + have "cbox a b \ {}" + using path_image_nonempty[of f] using assms(3) by auto + note ab=this[unfolded interval_eq_empty_cart not_ex forall_2 not_less] + have "pathstart f \ cbox a b" + and "pathfinish f \ cbox a b" + and "pathstart g \ cbox a b" + and "pathfinish g \ cbox a b" + using pathstart_in_path_image pathfinish_in_path_image + using assms(3-4) + by auto + note startfin = this[unfolded mem_interval_cart forall_2] + let ?P1 = "linepath (vector[a$1 - 2, a$2 - 2]) (vector[(pathstart f)$1,a$2 - 2]) +++ + linepath(vector[(pathstart f)$1,a$2 - 2])(pathstart f) +++ f +++ + linepath(pathfinish f)(vector[(pathfinish f)$1,a$2 - 2]) +++ + linepath(vector[(pathfinish f)$1,a$2 - 2])(vector[b$1 + 2,a$2 - 2])" + let ?P2 = "linepath(vector[(pathstart g)$1, (pathstart g)$2 - 3])(pathstart g) +++ g +++ + linepath(pathfinish g)(vector[(pathfinish g)$1,a$2 - 1]) +++ + linepath(vector[(pathfinish g)$1,a$2 - 1])(vector[b$1 + 1,a$2 - 1]) +++ + linepath(vector[b$1 + 1,a$2 - 1])(vector[b$1 + 1,b$2 + 3])" + let ?a = "vector[a$1 - 2, a$2 - 3]" + let ?b = "vector[b$1 + 2, b$2 + 3]" + have P1P2: "path_image ?P1 = path_image (linepath (vector[a$1 - 2, a$2 - 2]) (vector[(pathstart f)$1,a$2 - 2])) \ + path_image (linepath(vector[(pathstart f)$1,a$2 - 2])(pathstart f)) \ path_image f \ + path_image (linepath(pathfinish f)(vector[(pathfinish f)$1,a$2 - 2])) \ + path_image (linepath(vector[(pathfinish f)$1,a$2 - 2])(vector[b$1 + 2,a$2 - 2]))" + "path_image ?P2 = path_image(linepath(vector[(pathstart g)$1, (pathstart g)$2 - 3])(pathstart g)) \ path_image g \ + path_image(linepath(pathfinish g)(vector[(pathfinish g)$1,a$2 - 1])) \ + path_image(linepath(vector[(pathfinish g)$1,a$2 - 1])(vector[b$1 + 1,a$2 - 1])) \ + path_image(linepath(vector[b$1 + 1,a$2 - 1])(vector[b$1 + 1,b$2 + 3]))" using assms(1-2) + by(auto simp add: path_image_join path_linepath) + have abab: "cbox a b \ cbox ?a ?b" + unfolding interval_cbox_cart[symmetric] + by (auto simp add:less_eq_vec_def forall_2 vector_2) + obtain z where + "z \ path_image + (linepath (vector [a $ 1 - 2, a $ 2 - 2]) (vector [pathstart f $ 1, a $ 2 - 2]) +++ + linepath (vector [pathstart f $ 1, a $ 2 - 2]) (pathstart f) +++ + f +++ + linepath (pathfinish f) (vector [pathfinish f $ 1, a $ 2 - 2]) +++ + linepath (vector [pathfinish f $ 1, a $ 2 - 2]) (vector [b $ 1 + 2, a $ 2 - 2]))" + "z \ path_image + (linepath (vector [pathstart g $ 1, pathstart g $ 2 - 3]) (pathstart g) +++ + g +++ + linepath (pathfinish g) (vector [pathfinish g $ 1, a $ 2 - 1]) +++ + linepath (vector [pathfinish g $ 1, a $ 2 - 1]) (vector [b $ 1 + 1, a $ 2 - 1]) +++ + linepath (vector [b $ 1 + 1, a $ 2 - 1]) (vector [b $ 1 + 1, b $ 2 + 3]))" + apply (rule fashoda[of ?P1 ?P2 ?a ?b]) + unfolding pathstart_join pathfinish_join pathstart_linepath pathfinish_linepath vector_2 + proof - + show "path ?P1" and "path ?P2" + using assms by auto + have "path_image ?P1 \ cbox ?a ?b" + unfolding P1P2 path_image_linepath + apply (rule Un_least)+ + defer 3 + apply (rule_tac[1-4] convex_box(1)[unfolded convex_contains_segment,rule_format]) + unfolding mem_interval_cart forall_2 vector_2 + using ab startfin abab assms(3) + using assms(9-) + unfolding assms + apply (auto simp add: field_simps box_def) + done + then show "path_image ?P1 \ cbox ?a ?b" . + have "path_image ?P2 \ cbox ?a ?b" + unfolding P1P2 path_image_linepath + apply (rule Un_least)+ + defer 2 + apply (rule_tac[1-4] convex_box(1)[unfolded convex_contains_segment,rule_format]) + unfolding mem_interval_cart forall_2 vector_2 + using ab startfin abab assms(4) + using assms(9-) + unfolding assms + apply (auto simp add: field_simps box_def) + done + then show "path_image ?P2 \ cbox ?a ?b" . + show "a $ 1 - 2 = a $ 1 - 2" + and "b $ 1 + 2 = b $ 1 + 2" + and "pathstart g $ 2 - 3 = a $ 2 - 3" + and "b $ 2 + 3 = b $ 2 + 3" + by (auto simp add: assms) + qed + note z=this[unfolded P1P2 path_image_linepath] + show thesis + apply (rule that[of z]) + proof - + have "(z \ closed_segment (vector [a $ 1 - 2, a $ 2 - 2]) (vector [pathstart f $ 1, a $ 2 - 2]) \ + z \ closed_segment (vector [pathstart f $ 1, a $ 2 - 2]) (pathstart f)) \ + z \ closed_segment (pathfinish f) (vector [pathfinish f $ 1, a $ 2 - 2]) \ + z \ closed_segment (vector [pathfinish f $ 1, a $ 2 - 2]) (vector [b $ 1 + 2, a $ 2 - 2]) \ + (((z \ closed_segment (vector [pathstart g $ 1, pathstart g $ 2 - 3]) (pathstart g)) \ + z \ closed_segment (pathfinish g) (vector [pathfinish g $ 1, a $ 2 - 1])) \ + z \ closed_segment (vector [pathfinish g $ 1, a $ 2 - 1]) (vector [b $ 1 + 1, a $ 2 - 1])) \ + z \ closed_segment (vector [b $ 1 + 1, a $ 2 - 1]) (vector [b $ 1 + 1, b $ 2 + 3]) \ False" + proof (simp only: segment_vertical segment_horizontal vector_2, goal_cases) + case prems: 1 + have "pathfinish f \ cbox a b" + using assms(3) pathfinish_in_path_image[of f] by auto + then have "1 + b $ 1 \ pathfinish f $ 1 \ False" + unfolding mem_interval_cart forall_2 by auto + then have "z$1 \ pathfinish f$1" + using prems(2) + using assms ab + by (auto simp add: field_simps) + moreover have "pathstart f \ cbox a b" + using assms(3) pathstart_in_path_image[of f] + by auto + then have "1 + b $ 1 \ pathstart f $ 1 \ False" + unfolding mem_interval_cart forall_2 + by auto + then have "z$1 \ pathstart f$1" + using prems(2) using assms ab + by (auto simp add: field_simps) + ultimately have *: "z$2 = a$2 - 2" + using prems(1) + by auto + have "z$1 \ pathfinish g$1" + using prems(2) + using assms ab + by (auto simp add: field_simps *) + moreover have "pathstart g \ cbox a b" + using assms(4) pathstart_in_path_image[of g] + by auto + note this[unfolded mem_interval_cart forall_2] + then have "z$1 \ pathstart g$1" + using prems(1) + using assms ab + by (auto simp add: field_simps *) + ultimately have "a $ 2 - 1 \ z $ 2 \ z $ 2 \ b $ 2 + 3 \ b $ 2 + 3 \ z $ 2 \ z $ 2 \ a $ 2 - 1" + using prems(2) + unfolding * assms + by (auto simp add: field_simps) + then show False + unfolding * using ab by auto + qed + then have "z \ path_image f \ z \ path_image g" + using z unfolding Un_iff by blast + then have z': "z \ cbox a b" + using assms(3-4) + by auto + have "a $ 2 = z $ 2 \ (z $ 1 = pathstart f $ 1 \ z $ 1 = pathfinish f $ 1) \ + z = pathstart f \ z = pathfinish f" + unfolding vec_eq_iff forall_2 assms + by auto + with z' show "z \ path_image f" + using z(1) + unfolding Un_iff mem_interval_cart forall_2 + apply - + apply (simp only: segment_vertical segment_horizontal vector_2) + unfolding assms + apply auto + done + have "a $ 2 = z $ 2 \ (z $ 1 = pathstart g $ 1 \ z $ 1 = pathfinish g $ 1) \ + z = pathstart g \ z = pathfinish g" + unfolding vec_eq_iff forall_2 assms + by auto + with z' show "z \ path_image g" + using z(2) + unfolding Un_iff mem_interval_cart forall_2 + apply (simp only: segment_vertical segment_horizontal vector_2) + unfolding assms + apply auto + done + qed +qed + +(** The Following still needs to be translated. Maybe I will do that later. + +(* ------------------------------------------------------------------------- *) +(* Complement in dimension N >= 2 of set homeomorphic to any interval in *) +(* any dimension is (path-)connected. This naively generalizes the argument *) +(* in Ryuji Maehara's paper "The Jordan curve theorem via the Brouwer *) +(* fixed point theorem", American Mathematical Monthly 1984. *) +(* ------------------------------------------------------------------------- *) + +let RETRACTION_INJECTIVE_IMAGE_INTERVAL = prove + (`!p:real^M->real^N a b. + ~(interval[a,b] = {}) /\ + p continuous_on interval[a,b] /\ + (!x y. x IN interval[a,b] /\ y IN interval[a,b] /\ p x = p y ==> x = y) + ==> ?f. f continuous_on (:real^N) /\ + IMAGE f (:real^N) SUBSET (IMAGE p (interval[a,b])) /\ + (!x. x IN (IMAGE p (interval[a,b])) ==> f x = x)`, + REPEAT STRIP_TAC THEN + FIRST_ASSUM(MP_TAC o GEN_REWRITE_RULE I [INJECTIVE_ON_LEFT_INVERSE]) THEN + DISCH_THEN(X_CHOOSE_TAC `q:real^N->real^M`) THEN + SUBGOAL_THEN `(q:real^N->real^M) continuous_on + (IMAGE p (interval[a:real^M,b]))` + ASSUME_TAC THENL + [MATCH_MP_TAC CONTINUOUS_ON_INVERSE THEN ASM_REWRITE_TAC[COMPACT_INTERVAL]; + ALL_TAC] THEN + MP_TAC(ISPECL [`q:real^N->real^M`; + `IMAGE (p:real^M->real^N) + (interval[a,b])`; + `a:real^M`; `b:real^M`] + TIETZE_CLOSED_INTERVAL) THEN + ASM_SIMP_TAC[COMPACT_INTERVAL; COMPACT_CONTINUOUS_IMAGE; + COMPACT_IMP_CLOSED] THEN + ANTS_TAC THENL [ASM SET_TAC[]; ALL_TAC] THEN + DISCH_THEN(X_CHOOSE_THEN `r:real^N->real^M` STRIP_ASSUME_TAC) THEN + EXISTS_TAC `(p:real^M->real^N) o (r:real^N->real^M)` THEN + REWRITE_TAC[SUBSET; FORALL_IN_IMAGE; o_THM; IN_UNIV] THEN + CONJ_TAC THENL [ALL_TAC; ASM SET_TAC[]] THEN + MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN ASM_REWRITE_TAC[] THEN + FIRST_X_ASSUM(MATCH_MP_TAC o MATCH_MP(REWRITE_RULE[IMP_CONJ] + CONTINUOUS_ON_SUBSET)) THEN ASM SET_TAC[]);; + +let UNBOUNDED_PATH_COMPONENTS_COMPLEMENT_HOMEOMORPHIC_INTERVAL = prove + (`!s:real^N->bool a b:real^M. + s homeomorphic (interval[a,b]) + ==> !x. ~(x IN s) ==> ~bounded(path_component((:real^N) DIFF s) x)`, + REPEAT GEN_TAC THEN REWRITE_TAC[homeomorphic; homeomorphism] THEN + REWRITE_TAC[LEFT_IMP_EXISTS_THM] THEN + MAP_EVERY X_GEN_TAC [`p':real^N->real^M`; `p:real^M->real^N`] THEN + DISCH_TAC THEN + SUBGOAL_THEN + `!x y. x IN interval[a,b] /\ y IN interval[a,b] /\ + (p:real^M->real^N) x = p y ==> x = y` + ASSUME_TAC THENL [ASM_MESON_TAC[]; ALL_TAC] THEN + FIRST_X_ASSUM(MP_TAC o funpow 4 CONJUNCT2) THEN + DISCH_THEN(CONJUNCTS_THEN2 (SUBST1_TAC o SYM) ASSUME_TAC) THEN + ASM_CASES_TAC `interval[a:real^M,b] = {}` THEN + ASM_REWRITE_TAC[IMAGE_CLAUSES; DIFF_EMPTY; PATH_COMPONENT_UNIV; + NOT_BOUNDED_UNIV] THEN + ABBREV_TAC `s = (:real^N) DIFF (IMAGE p (interval[a:real^M,b]))` THEN + X_GEN_TAC `c:real^N` THEN REPEAT STRIP_TAC THEN + SUBGOAL_THEN `(c:real^N) IN s` ASSUME_TAC THENL [ASM SET_TAC[]; ALL_TAC] THEN + SUBGOAL_THEN `bounded((path_component s c) UNION + (IMAGE (p:real^M->real^N) (interval[a,b])))` + MP_TAC THENL + [ASM_SIMP_TAC[BOUNDED_UNION; COMPACT_IMP_BOUNDED; COMPACT_IMP_BOUNDED; + COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL]; + ALL_TAC] THEN + DISCH_THEN(MP_TAC o SPEC `c:real^N` o MATCH_MP BOUNDED_SUBSET_BALL) THEN + REWRITE_TAC[UNION_SUBSET] THEN + DISCH_THEN(X_CHOOSE_THEN `B:real` STRIP_ASSUME_TAC) THEN + MP_TAC(ISPECL [`p:real^M->real^N`; `a:real^M`; `b:real^M`] + RETRACTION_INJECTIVE_IMAGE_INTERVAL) THEN + ASM_REWRITE_TAC[SUBSET; IN_UNIV] THEN + DISCH_THEN(X_CHOOSE_THEN `r:real^N->real^N` MP_TAC) THEN + DISCH_THEN(CONJUNCTS_THEN2 ASSUME_TAC + (CONJUNCTS_THEN2 MP_TAC ASSUME_TAC)) THEN + REWRITE_TAC[FORALL_IN_IMAGE; IN_UNIV] THEN DISCH_TAC THEN + ABBREV_TAC `q = \z:real^N. if z IN path_component s c then r(z) else z` THEN + SUBGOAL_THEN + `(q:real^N->real^N) continuous_on + (closure(path_component s c) UNION ((:real^N) DIFF (path_component s c)))` + MP_TAC THENL + [EXPAND_TAC "q" THEN MATCH_MP_TAC CONTINUOUS_ON_CASES THEN + REWRITE_TAC[CLOSED_CLOSURE; CONTINUOUS_ON_ID; GSYM OPEN_CLOSED] THEN + REPEAT CONJ_TAC THENL + [MATCH_MP_TAC OPEN_PATH_COMPONENT THEN EXPAND_TAC "s" THEN + ASM_SIMP_TAC[GSYM CLOSED_OPEN; COMPACT_IMP_CLOSED; + COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL]; + ASM_MESON_TAC[CONTINUOUS_ON_SUBSET; SUBSET_UNIV]; + ALL_TAC] THEN + X_GEN_TAC `z:real^N` THEN + REWRITE_TAC[SET_RULE `~(z IN (s DIFF t) /\ z IN t)`] THEN + STRIP_TAC THEN FIRST_X_ASSUM MATCH_MP_TAC THEN + MP_TAC(ISPECL + [`path_component s (z:real^N)`; `path_component s (c:real^N)`] + OPEN_INTER_CLOSURE_EQ_EMPTY) THEN + ASM_REWRITE_TAC[GSYM DISJOINT; PATH_COMPONENT_DISJOINT] THEN ANTS_TAC THENL + [MATCH_MP_TAC OPEN_PATH_COMPONENT THEN EXPAND_TAC "s" THEN + ASM_SIMP_TAC[GSYM CLOSED_OPEN; COMPACT_IMP_CLOSED; + COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL]; + REWRITE_TAC[DISJOINT; EXTENSION; IN_INTER; NOT_IN_EMPTY] THEN + DISCH_THEN(MP_TAC o SPEC `z:real^N`) THEN ASM_REWRITE_TAC[] THEN + GEN_REWRITE_TAC (LAND_CONV o ONCE_DEPTH_CONV) [IN] THEN + REWRITE_TAC[PATH_COMPONENT_REFL_EQ] THEN ASM SET_TAC[]]; + ALL_TAC] THEN + SUBGOAL_THEN + `closure(path_component s c) UNION ((:real^N) DIFF (path_component s c)) = + (:real^N)` + SUBST1_TAC THENL + [MATCH_MP_TAC(SET_RULE `s SUBSET t ==> t UNION (UNIV DIFF s) = UNIV`) THEN + REWRITE_TAC[CLOSURE_SUBSET]; + DISCH_TAC] THEN + MP_TAC(ISPECL + [`(\x. &2 % c - x) o + (\x. c + B / norm(x - c) % (x - c)) o (q:real^N->real^N)`; + `cball(c:real^N,B)`] + BROUWER) THEN + REWRITE_TAC[NOT_IMP; GSYM CONJ_ASSOC; COMPACT_CBALL; CONVEX_CBALL] THEN + ASM_SIMP_TAC[CBALL_EQ_EMPTY; REAL_LT_IMP_LE; REAL_NOT_LT] THEN + SUBGOAL_THEN `!x. ~((q:real^N->real^N) x = c)` ASSUME_TAC THENL + [X_GEN_TAC `x:real^N` THEN EXPAND_TAC "q" THEN + REWRITE_TAC[NORM_EQ_0; VECTOR_SUB_EQ] THEN COND_CASES_TAC THEN + ASM SET_TAC[PATH_COMPONENT_REFL_EQ]; + ALL_TAC] THEN + REPEAT CONJ_TAC THENL + [MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN + SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST] THEN + MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN CONJ_TAC THENL + [ASM_MESON_TAC[CONTINUOUS_ON_SUBSET; SUBSET_UNIV]; ALL_TAC] THEN + MATCH_MP_TAC CONTINUOUS_ON_ADD THEN REWRITE_TAC[CONTINUOUS_ON_CONST] THEN + MATCH_MP_TAC CONTINUOUS_ON_MUL THEN + SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST] THEN + REWRITE_TAC[o_DEF; real_div; LIFT_CMUL] THEN + MATCH_MP_TAC CONTINUOUS_ON_CMUL THEN + MATCH_MP_TAC(REWRITE_RULE[o_DEF] CONTINUOUS_ON_INV) THEN + ASM_REWRITE_TAC[FORALL_IN_IMAGE; NORM_EQ_0; VECTOR_SUB_EQ] THEN + SUBGOAL_THEN + `(\x:real^N. lift(norm(x - c))) = (lift o norm) o (\x. x - c)` + SUBST1_TAC THENL [REWRITE_TAC[o_DEF]; ALL_TAC] THEN + MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN + ASM_SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST; + CONTINUOUS_ON_LIFT_NORM]; + REWRITE_TAC[SUBSET; FORALL_IN_IMAGE; IN_CBALL; o_THM; dist] THEN + X_GEN_TAC `x:real^N` THEN DISCH_TAC THEN + REWRITE_TAC[VECTOR_ARITH `c - (&2 % c - (c + x)) = x`] THEN + REWRITE_TAC[NORM_MUL; REAL_ABS_DIV; REAL_ABS_NORM] THEN + ASM_SIMP_TAC[REAL_DIV_RMUL; NORM_EQ_0; VECTOR_SUB_EQ] THEN + ASM_REAL_ARITH_TAC; + REWRITE_TAC[NOT_EXISTS_THM; TAUT `~(c /\ b) <=> c ==> ~b`] THEN + REWRITE_TAC[IN_CBALL; o_THM; dist] THEN + X_GEN_TAC `x:real^N` THEN DISCH_TAC THEN + REWRITE_TAC[VECTOR_ARITH `&2 % c - (c + x') = x <=> --x' = x - c`] THEN + ASM_CASES_TAC `(x:real^N) IN path_component s c` THENL + [MATCH_MP_TAC(NORM_ARITH `norm(y) < B /\ norm(x) = B ==> ~(--x = y)`) THEN + REWRITE_TAC[NORM_MUL; REAL_ABS_DIV; REAL_ABS_NORM] THEN + ASM_SIMP_TAC[REAL_DIV_RMUL; NORM_EQ_0; VECTOR_SUB_EQ] THEN + ASM_SIMP_TAC[REAL_ARITH `&0 < B ==> abs B = B`] THEN + UNDISCH_TAC `path_component s c SUBSET ball(c:real^N,B)` THEN + REWRITE_TAC[SUBSET; IN_BALL] THEN ASM_MESON_TAC[dist; NORM_SUB]; + EXPAND_TAC "q" THEN REWRITE_TAC[] THEN ASM_REWRITE_TAC[] THEN + REWRITE_TAC[VECTOR_ARITH `--(c % x) = x <=> (&1 + c) % x = vec 0`] THEN + ASM_REWRITE_TAC[DE_MORGAN_THM; VECTOR_SUB_EQ; VECTOR_MUL_EQ_0] THEN + SUBGOAL_THEN `~(x:real^N = c)` ASSUME_TAC THENL + [ASM_MESON_TAC[PATH_COMPONENT_REFL; IN]; ALL_TAC] THEN + ASM_REWRITE_TAC[] THEN + MATCH_MP_TAC(REAL_ARITH `&0 < x ==> ~(&1 + x = &0)`) THEN + ASM_SIMP_TAC[REAL_LT_DIV; NORM_POS_LT; VECTOR_SUB_EQ]]]);; + +let PATH_CONNECTED_COMPLEMENT_HOMEOMORPHIC_INTERVAL = prove + (`!s:real^N->bool a b:real^M. + 2 <= dimindex(:N) /\ s homeomorphic interval[a,b] + ==> path_connected((:real^N) DIFF s)`, + REPEAT STRIP_TAC THEN REWRITE_TAC[PATH_CONNECTED_IFF_PATH_COMPONENT] THEN + FIRST_ASSUM(MP_TAC o MATCH_MP + UNBOUNDED_PATH_COMPONENTS_COMPLEMENT_HOMEOMORPHIC_INTERVAL) THEN + ASM_REWRITE_TAC[SET_RULE `~(x IN s) <=> x IN (UNIV DIFF s)`] THEN + ABBREV_TAC `t = (:real^N) DIFF s` THEN + DISCH_TAC THEN MAP_EVERY X_GEN_TAC [`x:real^N`; `y:real^N`] THEN + STRIP_TAC THEN FIRST_ASSUM(MP_TAC o MATCH_MP HOMEOMORPHIC_COMPACTNESS) THEN + REWRITE_TAC[COMPACT_INTERVAL] THEN + DISCH_THEN(MP_TAC o MATCH_MP COMPACT_IMP_BOUNDED) THEN + REWRITE_TAC[BOUNDED_POS; LEFT_IMP_EXISTS_THM] THEN + X_GEN_TAC `B:real` THEN STRIP_TAC THEN + SUBGOAL_THEN `(?u:real^N. u IN path_component t x /\ B < norm(u)) /\ + (?v:real^N. v IN path_component t y /\ B < norm(v))` + STRIP_ASSUME_TAC THENL + [ASM_MESON_TAC[BOUNDED_POS; REAL_NOT_LE]; ALL_TAC] THEN + MATCH_MP_TAC PATH_COMPONENT_TRANS THEN EXISTS_TAC `u:real^N` THEN + CONJ_TAC THENL [ASM_MESON_TAC[IN]; ALL_TAC] THEN + MATCH_MP_TAC PATH_COMPONENT_SYM THEN + MATCH_MP_TAC PATH_COMPONENT_TRANS THEN EXISTS_TAC `v:real^N` THEN + CONJ_TAC THENL [ASM_MESON_TAC[IN]; ALL_TAC] THEN + MATCH_MP_TAC PATH_COMPONENT_OF_SUBSET THEN + EXISTS_TAC `(:real^N) DIFF cball(vec 0,B)` THEN CONJ_TAC THENL + [EXPAND_TAC "t" THEN MATCH_MP_TAC(SET_RULE + `s SUBSET t ==> (u DIFF t) SUBSET (u DIFF s)`) THEN + ASM_REWRITE_TAC[SUBSET; IN_CBALL_0]; + MP_TAC(ISPEC `cball(vec 0:real^N,B)` + PATH_CONNECTED_COMPLEMENT_BOUNDED_CONVEX) THEN + ASM_REWRITE_TAC[BOUNDED_CBALL; CONVEX_CBALL] THEN + REWRITE_TAC[PATH_CONNECTED_IFF_PATH_COMPONENT] THEN + DISCH_THEN MATCH_MP_TAC THEN + ASM_REWRITE_TAC[IN_DIFF; IN_UNIV; IN_CBALL_0; REAL_NOT_LE]]);; + +(* ------------------------------------------------------------------------- *) +(* In particular, apply all these to the special case of an arc. *) +(* ------------------------------------------------------------------------- *) + +let RETRACTION_ARC = prove + (`!p. arc p + ==> ?f. f continuous_on (:real^N) /\ + IMAGE f (:real^N) SUBSET path_image p /\ + (!x. x IN path_image p ==> f x = x)`, + REWRITE_TAC[arc; path; path_image] THEN REPEAT STRIP_TAC THEN + MATCH_MP_TAC RETRACTION_INJECTIVE_IMAGE_INTERVAL THEN + ASM_REWRITE_TAC[INTERVAL_EQ_EMPTY_CART_1; DROP_VEC; REAL_NOT_LT; REAL_POS]);; + +let PATH_CONNECTED_ARC_COMPLEMENT = prove + (`!p. 2 <= dimindex(:N) /\ arc p + ==> path_connected((:real^N) DIFF path_image p)`, + REWRITE_TAC[arc; path] THEN REPEAT STRIP_TAC THEN SIMP_TAC[path_image] THEN + MP_TAC(ISPECL [`path_image p:real^N->bool`; `vec 0:real^1`; `vec 1:real^1`] + PATH_CONNECTED_COMPLEMENT_HOMEOMORPHIC_INTERVAL) THEN + ASM_REWRITE_TAC[path_image] THEN DISCH_THEN MATCH_MP_TAC THEN + ONCE_REWRITE_TAC[HOMEOMORPHIC_SYM] THEN + MATCH_MP_TAC HOMEOMORPHIC_COMPACT THEN + EXISTS_TAC `p:real^1->real^N` THEN ASM_REWRITE_TAC[COMPACT_INTERVAL]);; + +let CONNECTED_ARC_COMPLEMENT = prove + (`!p. 2 <= dimindex(:N) /\ arc p + ==> connected((:real^N) DIFF path_image p)`, + SIMP_TAC[PATH_CONNECTED_ARC_COMPLEMENT; PATH_CONNECTED_IMP_CONNECTED]);; *) + +end diff -r bbcb05504fdc -r bd218a9320b5 src/HOL/Multivariate_Analysis/Gamma.thy --- a/src/HOL/Multivariate_Analysis/Gamma.thy Thu Aug 04 18:45:28 2016 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,2969 +0,0 @@ -(* Title: HOL/Multivariate_Analysis/Gamma.thy - Author: Manuel Eberl, TU München -*) - -section \The Gamma Function\ - -theory Gamma -imports - Complex_Transcendental - Summation - Harmonic_Numbers - "~~/src/HOL/Library/Nonpos_Ints" - "~~/src/HOL/Library/Periodic_Fun" -begin - -text \ - Several equivalent definitions of the Gamma function and its - most important properties. Also contains the definition and some properties - of the log-Gamma function and the Digamma function and the other Polygamma functions. - - Based on the Gamma function, we also prove the Weierstraß product form of the - sin function and, based on this, the solution of the Basel problem (the - sum over all @{term "1 / (n::nat)^2"}. -\ - -lemma pochhammer_eq_0_imp_nonpos_Int: - "pochhammer (x::'a::field_char_0) n = 0 \ x \ \\<^sub>\\<^sub>0" - by (auto simp: pochhammer_eq_0_iff) - -lemma closed_nonpos_Ints [simp]: "closed (\\<^sub>\\<^sub>0 :: 'a :: real_normed_algebra_1 set)" -proof - - have "\\<^sub>\\<^sub>0 = (of_int ` {n. n \ 0} :: 'a set)" - by (auto elim!: nonpos_Ints_cases intro!: nonpos_Ints_of_int) - also have "closed \" by (rule closed_of_int_image) - finally show ?thesis . -qed - -lemma plus_one_in_nonpos_Ints_imp: "z + 1 \ \\<^sub>\\<^sub>0 \ z \ \\<^sub>\\<^sub>0" - using nonpos_Ints_diff_Nats[of "z+1" "1"] by simp_all - -lemma of_int_in_nonpos_Ints_iff: - "(of_int n :: 'a :: ring_char_0) \ \\<^sub>\\<^sub>0 \ n \ 0" - by (auto simp: nonpos_Ints_def) - -lemma one_plus_of_int_in_nonpos_Ints_iff: - "(1 + of_int n :: 'a :: ring_char_0) \ \\<^sub>\\<^sub>0 \ n \ -1" -proof - - have "1 + of_int n = (of_int (n + 1) :: 'a)" by simp - also have "\ \ \\<^sub>\\<^sub>0 \ n + 1 \ 0" by (subst of_int_in_nonpos_Ints_iff) simp_all - also have "\ \ n \ -1" by presburger - finally show ?thesis . -qed - -lemma one_minus_of_nat_in_nonpos_Ints_iff: - "(1 - of_nat n :: 'a :: ring_char_0) \ \\<^sub>\\<^sub>0 \ n > 0" -proof - - have "(1 - of_nat n :: 'a) = of_int (1 - int n)" by simp - also have "\ \ \\<^sub>\\<^sub>0 \ n > 0" by (subst of_int_in_nonpos_Ints_iff) presburger - finally show ?thesis . -qed - -lemma fraction_not_in_ints: - assumes "\(n dvd m)" "n \ 0" - shows "of_int m / of_int n \ (\ :: 'a :: {division_ring,ring_char_0} set)" -proof - assume "of_int m / (of_int n :: 'a) \ \" - then obtain k where "of_int m / of_int n = (of_int k :: 'a)" by (elim Ints_cases) - with assms have "of_int m = (of_int (k * n) :: 'a)" by (auto simp add: divide_simps) - hence "m = k * n" by (subst (asm) of_int_eq_iff) - hence "n dvd m" by simp - with assms(1) show False by contradiction -qed - -lemma fraction_not_in_nats: - assumes "\n dvd m" "n \ 0" - shows "of_int m / of_int n \ (\ :: 'a :: {division_ring,ring_char_0} set)" -proof - assume "of_int m / of_int n \ (\ :: 'a set)" - also note Nats_subset_Ints - finally have "of_int m / of_int n \ (\ :: 'a set)" . - moreover have "of_int m / of_int n \ (\ :: 'a set)" - using assms by (intro fraction_not_in_ints) - ultimately show False by contradiction -qed - -lemma not_in_Ints_imp_not_in_nonpos_Ints: "z \ \ \ z \ \\<^sub>\\<^sub>0" - by (auto simp: Ints_def nonpos_Ints_def) - -lemma double_in_nonpos_Ints_imp: - assumes "2 * (z :: 'a :: field_char_0) \ \\<^sub>\\<^sub>0" - shows "z \ \\<^sub>\\<^sub>0 \ z + 1/2 \ \\<^sub>\\<^sub>0" -proof- - from assms obtain k where k: "2 * z = - of_nat k" by (elim nonpos_Ints_cases') - thus ?thesis by (cases "even k") (auto elim!: evenE oddE simp: field_simps) -qed - - -lemma sin_series: "(\n. ((-1)^n / fact (2*n+1)) *\<^sub>R z^(2*n+1)) sums sin z" -proof - - from sin_converges[of z] have "(\n. sin_coeff n *\<^sub>R z^n) sums sin z" . - also have "(\n. sin_coeff n *\<^sub>R z^n) sums sin z \ - (\n. ((-1)^n / fact (2*n+1)) *\<^sub>R z^(2*n+1)) sums sin z" - by (subst sums_mono_reindex[of "\n. 2*n+1", symmetric]) - (auto simp: sin_coeff_def subseq_def ac_simps elim!: oddE) - finally show ?thesis . -qed - -lemma cos_series: "(\n. ((-1)^n / fact (2*n)) *\<^sub>R z^(2*n)) sums cos z" -proof - - from cos_converges[of z] have "(\n. cos_coeff n *\<^sub>R z^n) sums cos z" . - also have "(\n. cos_coeff n *\<^sub>R z^n) sums cos z \ - (\n. ((-1)^n / fact (2*n)) *\<^sub>R z^(2*n)) sums cos z" - by (subst sums_mono_reindex[of "\n. 2*n", symmetric]) - (auto simp: cos_coeff_def subseq_def ac_simps elim!: evenE) - finally show ?thesis . -qed - -lemma sin_z_over_z_series: - fixes z :: "'a :: {real_normed_field,banach}" - assumes "z \ 0" - shows "(\n. (-1)^n / fact (2*n+1) * z^(2*n)) sums (sin z / z)" -proof - - from sin_series[of z] have "(\n. z * ((-1)^n / fact (2*n+1)) * z^(2*n)) sums sin z" - by (simp add: field_simps scaleR_conv_of_real) - from sums_mult[OF this, of "inverse z"] and assms show ?thesis - by (simp add: field_simps) -qed - -lemma sin_z_over_z_series': - fixes z :: "'a :: {real_normed_field,banach}" - assumes "z \ 0" - shows "(\n. sin_coeff (n+1) *\<^sub>R z^n) sums (sin z / z)" -proof - - from sums_split_initial_segment[OF sin_converges[of z], of 1] - have "(\n. z * (sin_coeff (n+1) *\<^sub>R z ^ n)) sums sin z" by simp - from sums_mult[OF this, of "inverse z"] assms show ?thesis by (simp add: field_simps) -qed - -lemma has_field_derivative_sin_z_over_z: - fixes A :: "'a :: {real_normed_field,banach} set" - shows "((\z. if z = 0 then 1 else sin z / z) has_field_derivative 0) (at 0 within A)" - (is "(?f has_field_derivative ?f') _") -proof (rule has_field_derivative_at_within) - have "((\z::'a. \n. of_real (sin_coeff (n+1)) * z^n) - has_field_derivative (\n. diffs (\n. of_real (sin_coeff (n+1))) n * 0^n)) (at 0)" - proof (rule termdiffs_strong) - from summable_ignore_initial_segment[OF sums_summable[OF sin_converges[of "1::'a"]], of 1] - show "summable (\n. of_real (sin_coeff (n+1)) * (1::'a)^n)" by (simp add: of_real_def) - qed simp - also have "(\z::'a. \n. of_real (sin_coeff (n+1)) * z^n) = ?f" - proof - fix z - show "(\n. of_real (sin_coeff (n+1)) * z^n) = ?f z" - by (cases "z = 0") (insert sin_z_over_z_series'[of z], - simp_all add: scaleR_conv_of_real sums_iff powser_zero sin_coeff_def) - qed - also have "(\n. diffs (\n. of_real (sin_coeff (n + 1))) n * (0::'a) ^ n) = - diffs (\n. of_real (sin_coeff (Suc n))) 0" by (simp add: powser_zero) - also have "\ = 0" by (simp add: sin_coeff_def diffs_def) - finally show "((\z::'a. if z = 0 then 1 else sin z / z) has_field_derivative 0) (at 0)" . -qed - -lemma round_Re_minimises_norm: - "norm ((z::complex) - of_int m) \ norm (z - of_int (round (Re z)))" -proof - - let ?n = "round (Re z)" - have "norm (z - of_int ?n) = sqrt ((Re z - of_int ?n)\<^sup>2 + (Im z)\<^sup>2)" - by (simp add: cmod_def) - also have "\Re z - of_int ?n\ \ \Re z - of_int m\" by (rule round_diff_minimal) - hence "sqrt ((Re z - of_int ?n)\<^sup>2 + (Im z)\<^sup>2) \ sqrt ((Re z - of_int m)\<^sup>2 + (Im z)\<^sup>2)" - by (intro real_sqrt_le_mono add_mono) (simp_all add: abs_le_square_iff) - also have "\ = norm (z - of_int m)" by (simp add: cmod_def) - finally show ?thesis . -qed - -lemma Re_pos_in_ball: - assumes "Re z > 0" "t \ ball z (Re z/2)" - shows "Re t > 0" -proof - - have "Re (z - t) \ norm (z - t)" by (rule complex_Re_le_cmod) - also from assms have "\ < Re z / 2" by (simp add: dist_complex_def) - finally show "Re t > 0" using assms by simp -qed - -lemma no_nonpos_Int_in_ball_complex: - assumes "Re z > 0" "t \ ball z (Re z/2)" - shows "t \ \\<^sub>\\<^sub>0" - using Re_pos_in_ball[OF assms] by (force elim!: nonpos_Ints_cases) - -lemma no_nonpos_Int_in_ball: - assumes "t \ ball z (dist z (round (Re z)))" - shows "t \ \\<^sub>\\<^sub>0" -proof - assume "t \ \\<^sub>\\<^sub>0" - then obtain n where "t = of_int n" by (auto elim!: nonpos_Ints_cases) - have "dist z (of_int n) \ dist z t + dist t (of_int n)" by (rule dist_triangle) - also from assms have "dist z t < dist z (round (Re z))" by simp - also have "\ \ dist z (of_int n)" - using round_Re_minimises_norm[of z] by (simp add: dist_complex_def) - finally have "dist t (of_int n) > 0" by simp - with \t = of_int n\ show False by simp -qed - -lemma no_nonpos_Int_in_ball': - assumes "(z :: 'a :: {euclidean_space,real_normed_algebra_1}) \ \\<^sub>\\<^sub>0" - obtains d where "d > 0" "\t. t \ ball z d \ t \ \\<^sub>\\<^sub>0" -proof (rule that) - from assms show "setdist {z} \\<^sub>\\<^sub>0 > 0" by (subst setdist_gt_0_compact_closed) auto -next - fix t assume "t \ ball z (setdist {z} \\<^sub>\\<^sub>0)" - thus "t \ \\<^sub>\\<^sub>0" using setdist_le_dist[of z "{z}" t "\\<^sub>\\<^sub>0"] by force -qed - -lemma no_nonpos_Real_in_ball: - assumes z: "z \ \\<^sub>\\<^sub>0" and t: "t \ ball z (if Im z = 0 then Re z / 2 else abs (Im z) / 2)" - shows "t \ \\<^sub>\\<^sub>0" -using z -proof (cases "Im z = 0") - assume A: "Im z = 0" - with z have "Re z > 0" by (force simp add: complex_nonpos_Reals_iff) - with t A Re_pos_in_ball[of z t] show ?thesis by (force simp add: complex_nonpos_Reals_iff) -next - assume A: "Im z \ 0" - have "abs (Im z) - abs (Im t) \ abs (Im z - Im t)" by linarith - also have "\ = abs (Im (z - t))" by simp - also have "\ \ norm (z - t)" by (rule abs_Im_le_cmod) - also from A t have "\ \ abs (Im z) / 2" by (simp add: dist_complex_def) - finally have "abs (Im t) > 0" using A by simp - thus ?thesis by (force simp add: complex_nonpos_Reals_iff) -qed - - -subsection \Definitions\ - -text \ - We define the Gamma function by first defining its multiplicative inverse @{term "Gamma_inv"}. - This is more convenient because @{term "Gamma_inv"} is entire, which makes proofs of its - properties more convenient because one does not have to watch out for discontinuities. - (e.g. @{term "Gamma_inv"} fulfils @{term "rGamma z = z * rGamma (z + 1)"} everywhere, - whereas @{term "Gamma"} does not fulfil the analogous equation on the non-positive integers) - - We define the Gamma function (resp. its inverse) in the Euler form. This form has the advantage - that it is a relatively simple limit that converges everywhere. The limit at the poles is 0 - (due to division by 0). The functional equation @{term "Gamma (z + 1) = z * Gamma z"} follows - immediately from the definition. -\ - -definition Gamma_series :: "('a :: {banach,real_normed_field}) \ nat \ 'a" where - "Gamma_series z n = fact n * exp (z * of_real (ln (of_nat n))) / pochhammer z (n+1)" - -definition Gamma_series' :: "('a :: {banach,real_normed_field}) \ nat \ 'a" where - "Gamma_series' z n = fact (n - 1) * exp (z * of_real (ln (of_nat n))) / pochhammer z n" - -definition rGamma_series :: "('a :: {banach,real_normed_field}) \ nat \ 'a" where - "rGamma_series z n = pochhammer z (n+1) / (fact n * exp (z * of_real (ln (of_nat n))))" - -lemma Gamma_series_altdef: "Gamma_series z n = inverse (rGamma_series z n)" - and rGamma_series_altdef: "rGamma_series z n = inverse (Gamma_series z n)" - unfolding Gamma_series_def rGamma_series_def by simp_all - -lemma rGamma_series_minus_of_nat: - "eventually (\n. rGamma_series (- of_nat k) n = 0) sequentially" - using eventually_ge_at_top[of k] - by eventually_elim (auto simp: rGamma_series_def pochhammer_of_nat_eq_0_iff) - -lemma Gamma_series_minus_of_nat: - "eventually (\n. Gamma_series (- of_nat k) n = 0) sequentially" - using eventually_ge_at_top[of k] - by eventually_elim (auto simp: Gamma_series_def pochhammer_of_nat_eq_0_iff) - -lemma Gamma_series'_minus_of_nat: - "eventually (\n. Gamma_series' (- of_nat k) n = 0) sequentially" - using eventually_gt_at_top[of k] - by eventually_elim (auto simp: Gamma_series'_def pochhammer_of_nat_eq_0_iff) - -lemma rGamma_series_nonpos_Ints_LIMSEQ: "z \ \\<^sub>\\<^sub>0 \ rGamma_series z \ 0" - by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule rGamma_series_minus_of_nat, simp) - -lemma Gamma_series_nonpos_Ints_LIMSEQ: "z \ \\<^sub>\\<^sub>0 \ Gamma_series z \ 0" - by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule Gamma_series_minus_of_nat, simp) - -lemma Gamma_series'_nonpos_Ints_LIMSEQ: "z \ \\<^sub>\\<^sub>0 \ Gamma_series' z \ 0" - by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule Gamma_series'_minus_of_nat, simp) - -lemma Gamma_series_Gamma_series': - assumes z: "z \ \\<^sub>\\<^sub>0" - shows "(\n. Gamma_series' z n / Gamma_series z n) \ 1" -proof (rule Lim_transform_eventually) - from eventually_gt_at_top[of "0::nat"] - show "eventually (\n. z / of_nat n + 1 = Gamma_series' z n / Gamma_series z n) sequentially" - proof eventually_elim - fix n :: nat assume n: "n > 0" - from n z have "Gamma_series' z n / Gamma_series z n = (z + of_nat n) / of_nat n" - by (cases n, simp) - (auto simp add: Gamma_series_def Gamma_series'_def pochhammer_rec' - dest: pochhammer_eq_0_imp_nonpos_Int plus_of_nat_eq_0_imp) - also from n have "\ = z / of_nat n + 1" by (simp add: divide_simps) - finally show "z / of_nat n + 1 = Gamma_series' z n / Gamma_series z n" .. - qed - have "(\x. z / of_nat x) \ 0" - by (rule tendsto_norm_zero_cancel) - (insert tendsto_mult[OF tendsto_const[of "norm z"] lim_inverse_n], - simp add: norm_divide inverse_eq_divide) - from tendsto_add[OF this tendsto_const[of 1]] show "(\n. z / of_nat n + 1) \ 1" by simp -qed - - -subsection \Convergence of the Euler series form\ - -text \ - We now show that the series that defines the Gamma function in the Euler form converges - and that the function defined by it is continuous on the complex halfspace with positive - real part. - - We do this by showing that the logarithm of the Euler series is continuous and converges - locally uniformly, which means that the log-Gamma function defined by its limit is also - continuous. - - This will later allow us to lift holomorphicity and continuity from the log-Gamma - function to the inverse of the Gamma function, and from that to the Gamma function itself. -\ - -definition ln_Gamma_series :: "('a :: {banach,real_normed_field,ln}) \ nat \ 'a" where - "ln_Gamma_series z n = z * ln (of_nat n) - ln z - (\k=1..n. ln (z / of_nat k + 1))" - -definition ln_Gamma_series' :: "('a :: {banach,real_normed_field,ln}) \ nat \ 'a" where - "ln_Gamma_series' z n = - - euler_mascheroni*z - ln z + (\k=1..n. z / of_nat n - ln (z / of_nat k + 1))" - -definition ln_Gamma :: "('a :: {banach,real_normed_field,ln}) \ 'a" where - "ln_Gamma z = lim (ln_Gamma_series z)" - - -text \ - We now show that the log-Gamma series converges locally uniformly for all complex numbers except - the non-positive integers. We do this by proving that the series is locally Cauchy, adapting this - proof: - http://math.stackexchange.com/questions/887158/convergence-of-gammaz-lim-n-to-infty-fracnzn-prod-m-0nzm -\ - -context -begin - -private lemma ln_Gamma_series_complex_converges_aux: - fixes z :: complex and k :: nat - assumes z: "z \ 0" and k: "of_nat k \ 2*norm z" "k \ 2" - shows "norm (z * ln (1 - 1/of_nat k) + ln (z/of_nat k + 1)) \ 2*(norm z + norm z^2) / of_nat k^2" -proof - - let ?k = "of_nat k :: complex" and ?z = "norm z" - have "z *ln (1 - 1/?k) + ln (z/?k+1) = z*(ln (1 - 1/?k :: complex) + 1/?k) + (ln (1+z/?k) - z/?k)" - by (simp add: algebra_simps) - also have "norm ... \ ?z * norm (ln (1-1/?k) + 1/?k :: complex) + norm (ln (1+z/?k) - z/?k)" - by (subst norm_mult [symmetric], rule norm_triangle_ineq) - also have "norm (Ln (1 + -1/?k) - (-1/?k)) \ (norm (-1/?k))\<^sup>2 / (1 - norm(-1/?k))" - using k by (intro Ln_approx_linear) (simp add: norm_divide) - hence "?z * norm (ln (1-1/?k) + 1/?k) \ ?z * ((norm (1/?k))^2 / (1 - norm (1/?k)))" - by (intro mult_left_mono) simp_all - also have "... \ (?z * (of_nat k / (of_nat k - 1))) / of_nat k^2" using k - by (simp add: field_simps power2_eq_square norm_divide) - also have "... \ (?z * 2) / of_nat k^2" using k - by (intro divide_right_mono mult_left_mono) (simp_all add: field_simps) - also have "norm (ln (1+z/?k) - z/?k) \ norm (z/?k)^2 / (1 - norm (z/?k))" using k - by (intro Ln_approx_linear) (simp add: norm_divide) - hence "norm (ln (1+z/?k) - z/?k) \ ?z^2 / of_nat k^2 / (1 - ?z / of_nat k)" - by (simp add: field_simps norm_divide) - also have "... \ (?z^2 * (of_nat k / (of_nat k - ?z))) / of_nat k^2" using k - by (simp add: field_simps power2_eq_square) - also have "... \ (?z^2 * 2) / of_nat k^2" using k - by (intro divide_right_mono mult_left_mono) (simp_all add: field_simps) - also note add_divide_distrib [symmetric] - finally show ?thesis by (simp only: distrib_left mult.commute) -qed - -lemma ln_Gamma_series_complex_converges: - assumes z: "z \ \\<^sub>\\<^sub>0" - assumes d: "d > 0" "\n. n \ \\<^sub>\\<^sub>0 \ norm (z - of_int n) > d" - shows "uniformly_convergent_on (ball z d) (\n z. ln_Gamma_series z n :: complex)" -proof (intro Cauchy_uniformly_convergent uniformly_Cauchy_onI') - fix e :: real assume e: "e > 0" - define e'' where "e'' = (SUP t:ball z d. norm t + norm t^2)" - define e' where "e' = e / (2*e'')" - have "bounded ((\t. norm t + norm t^2) ` cball z d)" - by (intro compact_imp_bounded compact_continuous_image) (auto intro!: continuous_intros) - hence "bounded ((\t. norm t + norm t^2) ` ball z d)" by (rule bounded_subset) auto - hence bdd: "bdd_above ((\t. norm t + norm t^2) ` ball z d)" by (rule bounded_imp_bdd_above) - - with z d(1) d(2)[of "-1"] have e''_pos: "e'' > 0" unfolding e''_def - by (subst less_cSUP_iff) (auto intro!: add_pos_nonneg bexI[of _ z]) - have e'': "norm t + norm t^2 \ e''" if "t \ ball z d" for t unfolding e''_def using that - by (rule cSUP_upper[OF _ bdd]) - from e z e''_pos have e': "e' > 0" unfolding e'_def - by (intro divide_pos_pos mult_pos_pos add_pos_pos) (simp_all add: field_simps) - - have "summable (\k. inverse ((real_of_nat k)^2))" - by (rule inverse_power_summable) simp - from summable_partial_sum_bound[OF this e'] guess M . note M = this - - define N where "N = max 2 (max (nat \2 * (norm z + d)\) M)" - { - from d have "\2 * (cmod z + d)\ \ \0::real\" - by (intro ceiling_mono mult_nonneg_nonneg add_nonneg_nonneg) simp_all - hence "2 * (norm z + d) \ of_nat (nat \2 * (norm z + d)\)" unfolding N_def - by (simp_all add: le_of_int_ceiling) - also have "... \ of_nat N" unfolding N_def - by (subst of_nat_le_iff) (rule max.coboundedI2, rule max.cobounded1) - finally have "of_nat N \ 2 * (norm z + d)" . - moreover have "N \ 2" "N \ M" unfolding N_def by simp_all - moreover have "(\k=m..n. 1/(of_nat k)\<^sup>2) < e'" if "m \ N" for m n - using M[OF order.trans[OF \N \ M\ that]] unfolding real_norm_def - by (subst (asm) abs_of_nonneg) (auto intro: setsum_nonneg simp: divide_simps) - moreover note calculation - } note N = this - - show "\M. \t\ball z d. \m\M. \n>m. dist (ln_Gamma_series t m) (ln_Gamma_series t n) < e" - unfolding dist_complex_def - proof (intro exI[of _ N] ballI allI impI) - fix t m n assume t: "t \ ball z d" and mn: "m \ N" "n > m" - from d(2)[of 0] t have "0 < dist z 0 - dist z t" by (simp add: field_simps dist_complex_def) - also have "dist z 0 - dist z t \ dist 0 t" using dist_triangle[of 0 z t] - by (simp add: dist_commute) - finally have t_nz: "t \ 0" by auto - - have "norm t \ norm z + norm (t - z)" by (rule norm_triangle_sub) - also from t have "norm (t - z) < d" by (simp add: dist_complex_def norm_minus_commute) - also have "2 * (norm z + d) \ of_nat N" by (rule N) - also have "N \ m" by (rule mn) - finally have norm_t: "2 * norm t < of_nat m" by simp - - have "ln_Gamma_series t m - ln_Gamma_series t n = - (-(t * Ln (of_nat n)) - (-(t * Ln (of_nat m)))) + - ((\k=1..n. Ln (t / of_nat k + 1)) - (\k=1..m. Ln (t / of_nat k + 1)))" - by (simp add: ln_Gamma_series_def algebra_simps) - also have "(\k=1..n. Ln (t / of_nat k + 1)) - (\k=1..m. Ln (t / of_nat k + 1)) = - (\k\{1..n}-{1..m}. Ln (t / of_nat k + 1))" using mn - by (simp add: setsum_diff) - also from mn have "{1..n}-{1..m} = {Suc m..n}" by fastforce - also have "-(t * Ln (of_nat n)) - (-(t * Ln (of_nat m))) = - (\k = Suc m..n. t * Ln (of_nat (k - 1)) - t * Ln (of_nat k))" using mn - by (subst setsum_telescope'' [symmetric]) simp_all - also have "... = (\k = Suc m..n. t * Ln (of_nat (k - 1) / of_nat k))" using mn N - by (intro setsum_cong_Suc) - (simp_all del: of_nat_Suc add: field_simps Ln_of_nat Ln_of_nat_over_of_nat) - also have "of_nat (k - 1) / of_nat k = 1 - 1 / (of_nat k :: complex)" if "k \ {Suc m..n}" for k - using that of_nat_eq_0_iff[of "Suc i" for i] by (cases k) (simp_all add: divide_simps) - hence "(\k = Suc m..n. t * Ln (of_nat (k - 1) / of_nat k)) = - (\k = Suc m..n. t * Ln (1 - 1 / of_nat k))" using mn N - by (intro setsum.cong) simp_all - also note setsum.distrib [symmetric] - also have "norm (\k=Suc m..n. t * Ln (1 - 1/of_nat k) + Ln (t/of_nat k + 1)) \ - (\k=Suc m..n. 2 * (norm t + (norm t)\<^sup>2) / (real_of_nat k)\<^sup>2)" using t_nz N(2) mn norm_t - by (intro order.trans[OF norm_setsum setsum_mono[OF ln_Gamma_series_complex_converges_aux]]) simp_all - also have "... \ 2 * (norm t + norm t^2) * (\k=Suc m..n. 1 / (of_nat k)\<^sup>2)" - by (simp add: setsum_right_distrib) - also have "... < 2 * (norm t + norm t^2) * e'" using mn z t_nz - by (intro mult_strict_left_mono N mult_pos_pos add_pos_pos) simp_all - also from e''_pos have "... = e * ((cmod t + (cmod t)\<^sup>2) / e'')" - by (simp add: e'_def field_simps power2_eq_square) - also from e''[OF t] e''_pos e - have "\ \ e * 1" by (intro mult_left_mono) (simp_all add: field_simps) - finally show "norm (ln_Gamma_series t m - ln_Gamma_series t n) < e" by simp - qed -qed - -end - -lemma ln_Gamma_series_complex_converges': - assumes z: "(z :: complex) \ \\<^sub>\\<^sub>0" - shows "\d>0. uniformly_convergent_on (ball z d) (\n z. ln_Gamma_series z n)" -proof - - define d' where "d' = Re z" - define d where "d = (if d' > 0 then d' / 2 else norm (z - of_int (round d')) / 2)" - have "of_int (round d') \ \\<^sub>\\<^sub>0" if "d' \ 0" using that - by (intro nonpos_Ints_of_int) (simp_all add: round_def) - with assms have d_pos: "d > 0" unfolding d_def by (force simp: not_less) - - have "d < cmod (z - of_int n)" if "n \ \\<^sub>\\<^sub>0" for n - proof (cases "Re z > 0") - case True - from nonpos_Ints_nonpos[OF that] have n: "n \ 0" by simp - from True have "d = Re z/2" by (simp add: d_def d'_def) - also from n True have "\ < Re (z - of_int n)" by simp - also have "\ \ norm (z - of_int n)" by (rule complex_Re_le_cmod) - finally show ?thesis . - next - case False - with assms nonpos_Ints_of_int[of "round (Re z)"] - have "z \ of_int (round d')" by (auto simp: not_less) - with False have "d < norm (z - of_int (round d'))" by (simp add: d_def d'_def) - also have "\ \ norm (z - of_int n)" unfolding d'_def by (rule round_Re_minimises_norm) - finally show ?thesis . - qed - hence conv: "uniformly_convergent_on (ball z d) (\n z. ln_Gamma_series z n)" - by (intro ln_Gamma_series_complex_converges d_pos z) simp_all - from d_pos conv show ?thesis by blast -qed - -lemma ln_Gamma_series_complex_converges'': "(z :: complex) \ \\<^sub>\\<^sub>0 \ convergent (ln_Gamma_series z)" - by (drule ln_Gamma_series_complex_converges') (auto intro: uniformly_convergent_imp_convergent) - -lemma ln_Gamma_complex_LIMSEQ: "(z :: complex) \ \\<^sub>\\<^sub>0 \ ln_Gamma_series z \ ln_Gamma z" - using ln_Gamma_series_complex_converges'' by (simp add: convergent_LIMSEQ_iff ln_Gamma_def) - -lemma exp_ln_Gamma_series_complex: - assumes "n > 0" "z \ \\<^sub>\\<^sub>0" - shows "exp (ln_Gamma_series z n :: complex) = Gamma_series z n" -proof - - from assms obtain m where m: "n = Suc m" by (cases n) blast - from assms have "z \ 0" by (intro notI) auto - with assms have "exp (ln_Gamma_series z n) = - (of_nat n) powr z / (z * (\k=1..n. exp (Ln (z / of_nat k + 1))))" - unfolding ln_Gamma_series_def powr_def by (simp add: exp_diff exp_setsum) - also from assms have "(\k=1..n. exp (Ln (z / of_nat k + 1))) = (\k=1..n. z / of_nat k + 1)" - by (intro setprod.cong[OF refl], subst exp_Ln) (auto simp: field_simps plus_of_nat_eq_0_imp) - also have "... = (\k=1..n. z + k) / fact n" - by (simp add: fact_setprod) - (subst setprod_dividef [symmetric], simp_all add: field_simps) - also from m have "z * ... = (\k=0..n. z + k) / fact n" - by (simp add: setprod.atLeast0_atMost_Suc_shift setprod.atLeast_Suc_atMost_Suc_shift) - also have "(\k=0..n. z + k) = pochhammer z (Suc n)" - unfolding pochhammer_setprod - by (simp add: setprod.atLeast0_atMost_Suc atLeastLessThanSuc_atLeastAtMost) - also have "of_nat n powr z / (pochhammer z (Suc n) / fact n) = Gamma_series z n" - unfolding Gamma_series_def using assms by (simp add: divide_simps powr_def Ln_of_nat) - finally show ?thesis . -qed - - -lemma ln_Gamma_series'_aux: - assumes "(z::complex) \ \\<^sub>\\<^sub>0" - shows "(\k. z / of_nat (Suc k) - ln (1 + z / of_nat (Suc k))) sums - (ln_Gamma z + euler_mascheroni * z + ln z)" (is "?f sums ?s") -unfolding sums_def -proof (rule Lim_transform) - show "(\n. ln_Gamma_series z n + of_real (harm n - ln (of_nat n)) * z + ln z) \ ?s" - (is "?g \ _") - by (intro tendsto_intros ln_Gamma_complex_LIMSEQ euler_mascheroni_LIMSEQ_of_real assms) - - have A: "eventually (\n. (\k 0" - have "(\kk=1..n. z / of_nat k - ln (1 + z / of_nat k))" - by (subst atLeast0LessThan [symmetric], subst setsum_shift_bounds_Suc_ivl [symmetric], - subst atLeastLessThanSuc_atLeastAtMost) simp_all - also have "\ = z * of_real (harm n) - (\k=1..n. ln (1 + z / of_nat k))" - by (simp add: harm_def setsum_subtractf setsum_right_distrib divide_inverse) - also from n have "\ - ?g n = 0" - by (simp add: ln_Gamma_series_def setsum_subtractf algebra_simps Ln_of_nat) - finally show "(\kn. (\k 0" by (subst tendsto_cong[OF A]) simp_all -qed - - -lemma uniformly_summable_deriv_ln_Gamma: - assumes z: "(z :: 'a :: {real_normed_field,banach}) \ 0" and d: "d > 0" "d \ norm z/2" - shows "uniformly_convergent_on (ball z d) - (\k z. \ik z. \i ball z d" - have "norm z = norm (t + (z - t))" by simp - have "norm (t + (z - t)) \ norm t + norm (z - t)" by (rule norm_triangle_ineq) - also from t d have "norm (z - t) < norm z / 2" by (simp add: dist_norm) - finally have A: "norm t > norm z / 2" using z by (simp add: field_simps) - - have "norm t = norm (z + (t - z))" by simp - also have "\ \ norm z + norm (t - z)" by (rule norm_triangle_ineq) - also from t d have "norm (t - z) \ norm z / 2" by (simp add: dist_norm norm_minus_commute) - also from z have "\ < norm z" by simp - finally have B: "norm t < 2 * norm z" by simp - note A B - } note ball = this - - show "eventually (\n. \t\ball z d. norm (?f n t) \ 4 * norm z * inverse (of_nat (Suc n)^2)) sequentially" - using eventually_gt_at_top apply eventually_elim - proof safe - fix t :: 'a assume t: "t \ ball z d" - from z ball[OF t] have t_nz: "t \ 0" by auto - fix n :: nat assume n: "n > nat \4 * norm z\" - from ball[OF t] t_nz have "4 * norm z > 2 * norm t" by simp - also from n have "\ < of_nat n" by linarith - finally have n: "of_nat n > 2 * norm t" . - hence "of_nat n > norm t" by simp - hence t': "t \ -of_nat (Suc n)" by (intro notI) (simp del: of_nat_Suc) - - with t_nz have "?f n t = 1 / (of_nat (Suc n) * (1 + of_nat (Suc n)/t))" - by (simp add: divide_simps eq_neg_iff_add_eq_0 del: of_nat_Suc) - also have "norm \ = inverse (of_nat (Suc n)) * inverse (norm (of_nat (Suc n)/t + 1))" - by (simp add: norm_divide norm_mult divide_simps add_ac del: of_nat_Suc) - also { - from z t_nz ball[OF t] have "of_nat (Suc n) / (4 * norm z) \ of_nat (Suc n) / (2 * norm t)" - by (intro divide_left_mono mult_pos_pos) simp_all - also have "\ < norm (of_nat (Suc n) / t) - norm (1 :: 'a)" - using t_nz n by (simp add: field_simps norm_divide del: of_nat_Suc) - also have "\ \ norm (of_nat (Suc n)/t + 1)" by (rule norm_diff_ineq) - finally have "inverse (norm (of_nat (Suc n)/t + 1)) \ 4 * norm z / of_nat (Suc n)" - using z by (simp add: divide_simps norm_divide mult_ac del: of_nat_Suc) - } - also have "inverse (real_of_nat (Suc n)) * (4 * norm z / real_of_nat (Suc n)) = - 4 * norm z * inverse (of_nat (Suc n)^2)" - by (simp add: divide_simps power2_eq_square del: of_nat_Suc) - finally show "norm (?f n t) \ 4 * norm z * inverse (of_nat (Suc n)^2)" - by (simp del: of_nat_Suc) - qed -next - show "summable (\n. 4 * norm z * inverse ((of_nat (Suc n))^2))" - by (subst summable_Suc_iff) (simp add: summable_mult inverse_power_summable) -qed - -lemma summable_deriv_ln_Gamma: - "z \ (0 :: 'a :: {real_normed_field,banach}) \ - summable (\n. inverse (of_nat (Suc n)) - inverse (z + of_nat (Suc n)))" - unfolding summable_iff_convergent - by (rule uniformly_convergent_imp_convergent, - rule uniformly_summable_deriv_ln_Gamma[of z "norm z/2"]) simp_all - - -definition Polygamma :: "nat \ ('a :: {real_normed_field,banach}) \ 'a" where - "Polygamma n z = (if n = 0 then - (\k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) - euler_mascheroni else - (-1)^Suc n * fact n * (\k. inverse ((z + of_nat k)^Suc n)))" - -abbreviation Digamma :: "('a :: {real_normed_field,banach}) \ 'a" where - "Digamma \ Polygamma 0" - -lemma Digamma_def: - "Digamma z = (\k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) - euler_mascheroni" - by (simp add: Polygamma_def) - - -lemma summable_Digamma: - assumes "(z :: 'a :: {real_normed_field,banach}) \ 0" - shows "summable (\n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))" -proof - - have sums: "(\n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n)) sums - (0 - inverse (z + of_nat 0))" - by (intro telescope_sums filterlim_compose[OF tendsto_inverse_0] - tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat) - from summable_add[OF summable_deriv_ln_Gamma[OF assms] sums_summable[OF sums]] - show "summable (\n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))" by simp -qed - -lemma summable_offset: - assumes "summable (\n. f (n + k) :: 'a :: real_normed_vector)" - shows "summable f" -proof - - from assms have "convergent (\m. \nm. (\nnm. (\nnm. \n {k..n\\. f n) = (\nn=k..n=k..n=0..n. n + k"]) - (simp, subst image_add_atLeastLessThan, auto) - finally show "(\nnna. setsum f {.. lim (\m. setsum f {.. 0" and n: "n \ 2" - shows "uniformly_convergent_on (ball z d) (\k z. \inorm z * e\" - { - fix t assume t: "t \ ball z d" - have "norm t = norm (z + (t - z))" by simp - also have "\ \ norm z + norm (t - z)" by (rule norm_triangle_ineq) - also from t have "norm (t - z) < d" by (simp add: dist_norm norm_minus_commute) - finally have "norm t < norm z * e" using z by (simp add: divide_simps e_def) - } note ball = this - - show "eventually (\k. \t\ball z d. norm (inverse ((t + of_nat k)^n)) \ - inverse (of_nat (k - m)^n)) sequentially" - using eventually_gt_at_top[of m] apply eventually_elim - proof (intro ballI) - fix k :: nat and t :: 'a assume k: "k > m" and t: "t \ ball z d" - from k have "real_of_nat (k - m) = of_nat k - of_nat m" by (simp add: of_nat_diff) - also have "\ \ norm (of_nat k :: 'a) - norm z * e" - unfolding m_def by (subst norm_of_nat) linarith - also from ball[OF t] have "\ \ norm (of_nat k :: 'a) - norm t" by simp - also have "\ \ norm (of_nat k + t)" by (rule norm_diff_ineq) - finally have "inverse ((norm (t + of_nat k))^n) \ inverse (real_of_nat (k - m)^n)" using k n - by (intro le_imp_inverse_le power_mono) (simp_all add: add_ac del: of_nat_Suc) - thus "norm (inverse ((t + of_nat k)^n)) \ inverse (of_nat (k - m)^n)" - by (simp add: norm_inverse norm_power power_inverse) - qed - - have "summable (\k. inverse ((real_of_nat k)^n))" - using inverse_power_summable[of n] n by simp - hence "summable (\k. inverse ((real_of_nat (k + m - m))^n))" by simp - thus "summable (\k. inverse ((real_of_nat (k - m))^n))" by (rule summable_offset) -qed - -lemma Polygamma_converges': - fixes z :: "'a :: {real_normed_field,banach}" - assumes z: "z \ 0" and n: "n \ 2" - shows "summable (\k. inverse ((z + of_nat k)^n))" - using uniformly_convergent_imp_convergent[OF Polygamma_converges[OF assms, of 1], of z] - by (simp add: summable_iff_convergent) - -lemma has_field_derivative_ln_Gamma_complex [derivative_intros]: - fixes z :: complex - assumes z: "z \ \\<^sub>\\<^sub>0" - shows "(ln_Gamma has_field_derivative Digamma z) (at z)" -proof - - have not_nonpos_Int [simp]: "t \ \\<^sub>\\<^sub>0" if "Re t > 0" for t - using that by (auto elim!: nonpos_Ints_cases') - from z have z': "z \ \\<^sub>\\<^sub>0" and z'': "z \ 0" using nonpos_Ints_subset_nonpos_Reals nonpos_Reals_zero_I - by blast+ - let ?f' = "\z k. inverse (of_nat (Suc k)) - inverse (z + of_nat (Suc k))" - let ?f = "\z k. z / of_nat (Suc k) - ln (1 + z / of_nat (Suc k))" and ?F' = "\z. \n. ?f' z n" - define d where "d = min (norm z/2) (if Im z = 0 then Re z / 2 else abs (Im z) / 2)" - from z have d: "d > 0" "norm z/2 \ d" by (auto simp add: complex_nonpos_Reals_iff d_def) - have ball: "Im t = 0 \ Re t > 0" if "dist z t < d" for t - using no_nonpos_Real_in_ball[OF z, of t] that unfolding d_def by (force simp add: complex_nonpos_Reals_iff) - have sums: "(\n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n)) sums - (0 - inverse (z + of_nat 0))" - by (intro telescope_sums filterlim_compose[OF tendsto_inverse_0] - tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat) - - have "((\z. \n. ?f z n) has_field_derivative ?F' z) (at z)" - using d z ln_Gamma_series'_aux[OF z'] - apply (intro has_field_derivative_series'(2)[of "ball z d" _ _ z] uniformly_summable_deriv_ln_Gamma) - apply (auto intro!: derivative_eq_intros add_pos_pos mult_pos_pos dest!: ball - simp: field_simps sums_iff nonpos_Reals_divide_of_nat_iff - simp del: of_nat_Suc) - apply (auto simp add: complex_nonpos_Reals_iff) - done - with z have "((\z. (\k. ?f z k) - euler_mascheroni * z - Ln z) has_field_derivative - ?F' z - euler_mascheroni - inverse z) (at z)" - by (force intro!: derivative_eq_intros simp: Digamma_def) - also have "?F' z - euler_mascheroni - inverse z = (?F' z + -inverse z) - euler_mascheroni" by simp - also from sums have "-inverse z = (\n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n))" - by (simp add: sums_iff) - also from sums summable_deriv_ln_Gamma[OF z''] - have "?F' z + \ = (\n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))" - by (subst suminf_add) (simp_all add: add_ac sums_iff) - also have "\ - euler_mascheroni = Digamma z" by (simp add: Digamma_def) - finally have "((\z. (\k. ?f z k) - euler_mascheroni * z - Ln z) - has_field_derivative Digamma z) (at z)" . - moreover from eventually_nhds_ball[OF d(1), of z] - have "eventually (\z. ln_Gamma z = (\k. ?f z k) - euler_mascheroni * z - Ln z) (nhds z)" - proof eventually_elim - fix t assume "t \ ball z d" - hence "t \ \\<^sub>\\<^sub>0" by (auto dest!: ball elim!: nonpos_Ints_cases) - from ln_Gamma_series'_aux[OF this] - show "ln_Gamma t = (\k. ?f t k) - euler_mascheroni * t - Ln t" by (simp add: sums_iff) - qed - ultimately show ?thesis by (subst DERIV_cong_ev[OF refl _ refl]) -qed - -declare has_field_derivative_ln_Gamma_complex[THEN DERIV_chain2, derivative_intros] - - -lemma Digamma_1 [simp]: "Digamma (1 :: 'a :: {real_normed_field,banach}) = - euler_mascheroni" - by (simp add: Digamma_def) - -lemma Digamma_plus1: - assumes "z \ 0" - shows "Digamma (z+1) = Digamma z + 1/z" -proof - - have sums: "(\k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k))) - sums (inverse (z + of_nat 0) - 0)" - by (intro telescope_sums'[OF filterlim_compose[OF tendsto_inverse_0]] - tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat) - have "Digamma (z+1) = (\k. inverse (of_nat (Suc k)) - inverse (z + of_nat (Suc k))) - - euler_mascheroni" (is "_ = suminf ?f - _") by (simp add: Digamma_def add_ac) - also have "suminf ?f = (\k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) + - (\k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k)))" - using summable_Digamma[OF assms] sums by (subst suminf_add) (simp_all add: add_ac sums_iff) - also have "(\k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k))) = 1/z" - using sums by (simp add: sums_iff inverse_eq_divide) - finally show ?thesis by (simp add: Digamma_def[of z]) -qed - -lemma Polygamma_plus1: - assumes "z \ 0" - shows "Polygamma n (z + 1) = Polygamma n z + (-1)^n * fact n / (z ^ Suc n)" -proof (cases "n = 0") - assume n: "n \ 0" - let ?f = "\k. inverse ((z + of_nat k) ^ Suc n)" - have "Polygamma n (z + 1) = (-1) ^ Suc n * fact n * (\k. ?f (k+1))" - using n by (simp add: Polygamma_def add_ac) - also have "(\k. ?f (k+1)) + (\k<1. ?f k) = (\k. ?f k)" - using Polygamma_converges'[OF assms, of "Suc n"] n - by (subst suminf_split_initial_segment [symmetric]) simp_all - hence "(\k. ?f (k+1)) = (\k. ?f k) - inverse (z ^ Suc n)" by (simp add: algebra_simps) - also have "(-1) ^ Suc n * fact n * ((\k. ?f k) - inverse (z ^ Suc n)) = - Polygamma n z + (-1)^n * fact n / (z ^ Suc n)" using n - by (simp add: inverse_eq_divide algebra_simps Polygamma_def) - finally show ?thesis . -qed (insert assms, simp add: Digamma_plus1 inverse_eq_divide) - -lemma Digamma_of_nat: - "Digamma (of_nat (Suc n) :: 'a :: {real_normed_field,banach}) = harm n - euler_mascheroni" -proof (induction n) - case (Suc n) - have "Digamma (of_nat (Suc (Suc n)) :: 'a) = Digamma (of_nat (Suc n) + 1)" by simp - also have "\ = Digamma (of_nat (Suc n)) + inverse (of_nat (Suc n))" - by (subst Digamma_plus1) (simp_all add: inverse_eq_divide del: of_nat_Suc) - also have "Digamma (of_nat (Suc n) :: 'a) = harm n - euler_mascheroni " by (rule Suc) - also have "\ + inverse (of_nat (Suc n)) = harm (Suc n) - euler_mascheroni" - by (simp add: harm_Suc) - finally show ?case . -qed (simp add: harm_def) - -lemma Digamma_numeral: "Digamma (numeral n) = harm (pred_numeral n) - euler_mascheroni" - by (subst of_nat_numeral[symmetric], subst numeral_eq_Suc, subst Digamma_of_nat) (rule refl) - -lemma Polygamma_of_real: "x \ 0 \ Polygamma n (of_real x) = of_real (Polygamma n x)" - unfolding Polygamma_def using summable_Digamma[of x] Polygamma_converges'[of x "Suc n"] - by (simp_all add: suminf_of_real) - -lemma Polygamma_Real: "z \ \ \ z \ 0 \ Polygamma n z \ \" - by (elim Reals_cases, hypsubst, subst Polygamma_of_real) simp_all - -lemma Digamma_half_integer: - "Digamma (of_nat n + 1/2 :: 'a :: {real_normed_field,banach}) = - (\kk. inverse (of_nat (Suc k)) - inverse (of_nat k + 1/2)) - euler_mascheroni" - by (simp add: Digamma_def add_ac) - also have "(\k. inverse (of_nat (Suc k) :: real) - inverse (of_nat k + 1/2)) = - (\k. inverse (1/2) * (inverse (2 * of_nat (Suc k)) - inverse (2 * of_nat k + 1)))" - by (simp_all add: add_ac inverse_mult_distrib[symmetric] ring_distribs del: inverse_divide) - also have "\ = - 2 * ln 2" using sums_minus[OF alternating_harmonic_series_sums'] - by (subst suminf_mult) (simp_all add: algebra_simps sums_iff) - finally show ?case by simp -next - case (Suc n) - have nz: "2 * of_nat n + (1:: 'a) \ 0" - using of_nat_neq_0[of "2*n"] by (simp only: of_nat_Suc) (simp add: add_ac) - hence nz': "of_nat n + (1/2::'a) \ 0" by (simp add: field_simps) - have "Digamma (of_nat (Suc n) + 1/2 :: 'a) = Digamma (of_nat n + 1/2 + 1)" by simp - also from nz' have "\ = Digamma (of_nat n + 1 / 2) + 1 / (of_nat n + 1 / 2)" - by (rule Digamma_plus1) - also from nz nz' have "1 / (of_nat n + 1 / 2 :: 'a) = 2 / (2 * of_nat n + 1)" - by (subst divide_eq_eq) simp_all - also note Suc - finally show ?case by (simp add: add_ac) -qed - -lemma Digamma_one_half: "Digamma (1/2) = - euler_mascheroni - of_real (2 * ln 2)" - using Digamma_half_integer[of 0] by simp - -lemma Digamma_real_three_halves_pos: "Digamma (3/2 :: real) > 0" -proof - - have "-Digamma (3/2 :: real) = -Digamma (of_nat 1 + 1/2)" by simp - also have "\ = 2 * ln 2 + euler_mascheroni - 2" by (subst Digamma_half_integer) simp - also note euler_mascheroni_less_13_over_22 - also note ln2_le_25_over_36 - finally show ?thesis by simp -qed - - -lemma has_field_derivative_Polygamma [derivative_intros]: - fixes z :: "'a :: {real_normed_field,euclidean_space}" - assumes z: "z \ \\<^sub>\\<^sub>0" - shows "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z within A)" -proof (rule has_field_derivative_at_within, cases "n = 0") - assume n: "n = 0" - let ?f = "\k z. inverse (of_nat (Suc k)) - inverse (z + of_nat k)" - let ?F = "\z. \k. ?f k z" and ?f' = "\k z. inverse ((z + of_nat k)\<^sup>2)" - from no_nonpos_Int_in_ball'[OF z] guess d . note d = this - from z have summable: "summable (\k. inverse (of_nat (Suc k)) - inverse (z + of_nat k))" - by (intro summable_Digamma) force - from z have conv: "uniformly_convergent_on (ball z d) (\k z. \i2))" - by (intro Polygamma_converges) auto - with d have "summable (\k. inverse ((z + of_nat k)\<^sup>2))" unfolding summable_iff_convergent - by (auto dest!: uniformly_convergent_imp_convergent simp: summable_iff_convergent ) - - have "(?F has_field_derivative (\k. ?f' k z)) (at z)" - proof (rule has_field_derivative_series'[of "ball z d" _ _ z]) - fix k :: nat and t :: 'a assume t: "t \ ball z d" - from t d(2)[of t] show "((\z. ?f k z) has_field_derivative ?f' k t) (at t within ball z d)" - by (auto intro!: derivative_eq_intros simp: power2_eq_square simp del: of_nat_Suc - dest!: plus_of_nat_eq_0_imp elim!: nonpos_Ints_cases) - qed (insert d(1) summable conv, (assumption|simp)+) - with z show "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z)" - unfolding Digamma_def [abs_def] Polygamma_def [abs_def] using n - by (force simp: power2_eq_square intro!: derivative_eq_intros) -next - assume n: "n \ 0" - from z have z': "z \ 0" by auto - from no_nonpos_Int_in_ball'[OF z] guess d . note d = this - define n' where "n' = Suc n" - from n have n': "n' \ 2" by (simp add: n'_def) - have "((\z. \k. inverse ((z + of_nat k) ^ n')) has_field_derivative - (\k. - of_nat n' * inverse ((z + of_nat k) ^ (n'+1)))) (at z)" - proof (rule has_field_derivative_series'[of "ball z d" _ _ z]) - fix k :: nat and t :: 'a assume t: "t \ ball z d" - with d have t': "t \ \\<^sub>\\<^sub>0" "t \ 0" by auto - show "((\a. inverse ((a + of_nat k) ^ n')) has_field_derivative - - of_nat n' * inverse ((t + of_nat k) ^ (n'+1))) (at t within ball z d)" using t' - by (fastforce intro!: derivative_eq_intros simp: divide_simps power_diff dest: plus_of_nat_eq_0_imp) - next - have "uniformly_convergent_on (ball z d) - (\k z. (- of_nat n' :: 'a) * (\ik z. \ik. - of_nat n' * inverse ((z + of_nat k) ^ (n' + 1))) = - (- of_nat n') * (\k. inverse ((z + of_nat k) ^ (n' + 1)))" - using Polygamma_converges'[OF z', of "n'+1"] n' by (subst suminf_mult) simp_all - finally have "((\z. \k. inverse ((z + of_nat k) ^ n')) has_field_derivative - - of_nat n' * (\k. inverse ((z + of_nat k) ^ (n' + 1)))) (at z)" . - from DERIV_cmult[OF this, of "(-1)^Suc n * fact n :: 'a"] - show "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z)" - unfolding n'_def Polygamma_def[abs_def] using n by (simp add: algebra_simps) -qed - -declare has_field_derivative_Polygamma[THEN DERIV_chain2, derivative_intros] - -lemma isCont_Polygamma [continuous_intros]: - fixes f :: "_ \ 'a :: {real_normed_field,euclidean_space}" - shows "isCont f z \ f z \ \\<^sub>\\<^sub>0 \ isCont (\x. Polygamma n (f x)) z" - by (rule isCont_o2[OF _ DERIV_isCont[OF has_field_derivative_Polygamma]]) - -lemma continuous_on_Polygamma: - "A \ \\<^sub>\\<^sub>0 = {} \ continuous_on A (Polygamma n :: _ \ 'a :: {real_normed_field,euclidean_space})" - by (intro continuous_at_imp_continuous_on isCont_Polygamma[OF continuous_ident] ballI) blast - -lemma isCont_ln_Gamma_complex [continuous_intros]: - fixes f :: "'a::t2_space \ complex" - shows "isCont f z \ f z \ \\<^sub>\\<^sub>0 \ isCont (\z. ln_Gamma (f z)) z" - by (rule isCont_o2[OF _ DERIV_isCont[OF has_field_derivative_ln_Gamma_complex]]) - -lemma continuous_on_ln_Gamma_complex [continuous_intros]: - fixes A :: "complex set" - shows "A \ \\<^sub>\\<^sub>0 = {} \ continuous_on A ln_Gamma" - by (intro continuous_at_imp_continuous_on ballI isCont_ln_Gamma_complex[OF continuous_ident]) - fastforce - -text \ - We define a type class that captures all the fundamental properties of the inverse of the Gamma function - and defines the Gamma function upon that. This allows us to instantiate the type class both for - the reals and for the complex numbers with a minimal amount of proof duplication. -\ - -class Gamma = real_normed_field + complete_space + - fixes rGamma :: "'a \ 'a" - assumes rGamma_eq_zero_iff_aux: "rGamma z = 0 \ (\n. z = - of_nat n)" - assumes differentiable_rGamma_aux1: - "(\n. z \ - of_nat n) \ - let d = (THE d. (\n. \k d) - scaleR euler_mascheroni 1 - in filterlim (\y. (rGamma y - rGamma z + rGamma z * d * (y - z)) /\<^sub>R - norm (y - z)) (nhds 0) (at z)" - assumes differentiable_rGamma_aux2: - "let z = - of_nat n - in filterlim (\y. (rGamma y - rGamma z - (-1)^n * (setprod of_nat {1..n}) * (y - z)) /\<^sub>R - norm (y - z)) (nhds 0) (at z)" - assumes rGamma_series_aux: "(\n. z \ - of_nat n) \ - let fact' = (\n. setprod of_nat {1..n}); - exp = (\x. THE e. (\n. \kR fact k) \ e); - pochhammer' = (\a n. (\n = 0..n. a + of_nat n)) - in filterlim (\n. pochhammer' z n / (fact' n * exp (z * (ln (of_nat n) *\<^sub>R 1)))) - (nhds (rGamma z)) sequentially" -begin -subclass banach .. -end - -definition "Gamma z = inverse (rGamma z)" - - -subsection \Basic properties\ - -lemma Gamma_nonpos_Int: "z \ \\<^sub>\\<^sub>0 \ Gamma z = 0" - and rGamma_nonpos_Int: "z \ \\<^sub>\\<^sub>0 \ rGamma z = 0" - using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases') - -lemma Gamma_nonzero: "z \ \\<^sub>\\<^sub>0 \ Gamma z \ 0" - and rGamma_nonzero: "z \ \\<^sub>\\<^sub>0 \ rGamma z \ 0" - using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases') - -lemma Gamma_eq_zero_iff: "Gamma z = 0 \ z \ \\<^sub>\\<^sub>0" - and rGamma_eq_zero_iff: "rGamma z = 0 \ z \ \\<^sub>\\<^sub>0" - using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases') - -lemma rGamma_inverse_Gamma: "rGamma z = inverse (Gamma z)" - unfolding Gamma_def by simp - -lemma rGamma_series_LIMSEQ [tendsto_intros]: - "rGamma_series z \ rGamma z" -proof (cases "z \ \\<^sub>\\<^sub>0") - case False - hence "z \ - of_nat n" for n by auto - from rGamma_series_aux[OF this] show ?thesis - by (simp add: rGamma_series_def[abs_def] fact_setprod pochhammer_Suc_setprod - exp_def of_real_def[symmetric] suminf_def sums_def[abs_def] atLeast0AtMost) -qed (insert rGamma_eq_zero_iff[of z], simp_all add: rGamma_series_nonpos_Ints_LIMSEQ) - -lemma Gamma_series_LIMSEQ [tendsto_intros]: - "Gamma_series z \ Gamma z" -proof (cases "z \ \\<^sub>\\<^sub>0") - case False - hence "(\n. inverse (rGamma_series z n)) \ inverse (rGamma z)" - by (intro tendsto_intros) (simp_all add: rGamma_eq_zero_iff) - also have "(\n. inverse (rGamma_series z n)) = Gamma_series z" - by (simp add: rGamma_series_def Gamma_series_def[abs_def]) - finally show ?thesis by (simp add: Gamma_def) -qed (insert Gamma_eq_zero_iff[of z], simp_all add: Gamma_series_nonpos_Ints_LIMSEQ) - -lemma Gamma_altdef: "Gamma z = lim (Gamma_series z)" - using Gamma_series_LIMSEQ[of z] by (simp add: limI) - -lemma rGamma_1 [simp]: "rGamma 1 = 1" -proof - - have A: "eventually (\n. rGamma_series 1 n = of_nat (Suc n) / of_nat n) sequentially" - using eventually_gt_at_top[of "0::nat"] - by (force elim!: eventually_mono simp: rGamma_series_def exp_of_real pochhammer_fact - divide_simps pochhammer_rec' dest!: pochhammer_eq_0_imp_nonpos_Int) - have "rGamma_series 1 \ 1" by (subst tendsto_cong[OF A]) (rule LIMSEQ_Suc_n_over_n) - moreover have "rGamma_series 1 \ rGamma 1" by (rule tendsto_intros) - ultimately show ?thesis by (intro LIMSEQ_unique) -qed - -lemma rGamma_plus1: "z * rGamma (z + 1) = rGamma z" -proof - - let ?f = "\n. (z + 1) * inverse (of_nat n) + 1" - have "eventually (\n. ?f n * rGamma_series z n = z * rGamma_series (z + 1) n) sequentially" - using eventually_gt_at_top[of "0::nat"] - proof eventually_elim - fix n :: nat assume n: "n > 0" - hence "z * rGamma_series (z + 1) n = inverse (of_nat n) * - pochhammer z (Suc (Suc n)) / (fact n * exp (z * of_real (ln (of_nat n))))" - by (subst pochhammer_rec) (simp add: rGamma_series_def field_simps exp_add exp_of_real) - also from n have "\ = ?f n * rGamma_series z n" - by (subst pochhammer_rec') (simp_all add: divide_simps rGamma_series_def add_ac) - finally show "?f n * rGamma_series z n = z * rGamma_series (z + 1) n" .. - qed - moreover have "(\n. ?f n * rGamma_series z n) \ ((z+1) * 0 + 1) * rGamma z" - by (intro tendsto_intros lim_inverse_n) - hence "(\n. ?f n * rGamma_series z n) \ rGamma z" by simp - ultimately have "(\n. z * rGamma_series (z + 1) n) \ rGamma z" - by (rule Lim_transform_eventually) - moreover have "(\n. z * rGamma_series (z + 1) n) \ z * rGamma (z + 1)" - by (intro tendsto_intros) - ultimately show "z * rGamma (z + 1) = rGamma z" using LIMSEQ_unique by blast -qed - - -lemma pochhammer_rGamma: "rGamma z = pochhammer z n * rGamma (z + of_nat n)" -proof (induction n arbitrary: z) - case (Suc n z) - have "rGamma z = pochhammer z n * rGamma (z + of_nat n)" by (rule Suc.IH) - also note rGamma_plus1 [symmetric] - finally show ?case by (simp add: add_ac pochhammer_rec') -qed simp_all - -lemma Gamma_plus1: "z \ \\<^sub>\\<^sub>0 \ Gamma (z + 1) = z * Gamma z" - using rGamma_plus1[of z] by (simp add: rGamma_inverse_Gamma field_simps Gamma_eq_zero_iff) - -lemma pochhammer_Gamma: "z \ \\<^sub>\\<^sub>0 \ pochhammer z n = Gamma (z + of_nat n) / Gamma z" - using pochhammer_rGamma[of z] - by (simp add: rGamma_inverse_Gamma Gamma_eq_zero_iff field_simps) - -lemma Gamma_0 [simp]: "Gamma 0 = 0" - and rGamma_0 [simp]: "rGamma 0 = 0" - and Gamma_neg_1 [simp]: "Gamma (- 1) = 0" - and rGamma_neg_1 [simp]: "rGamma (- 1) = 0" - and Gamma_neg_numeral [simp]: "Gamma (- numeral n) = 0" - and rGamma_neg_numeral [simp]: "rGamma (- numeral n) = 0" - and Gamma_neg_of_nat [simp]: "Gamma (- of_nat m) = 0" - and rGamma_neg_of_nat [simp]: "rGamma (- of_nat m) = 0" - by (simp_all add: rGamma_eq_zero_iff Gamma_eq_zero_iff) - -lemma Gamma_1 [simp]: "Gamma 1 = 1" unfolding Gamma_def by simp - -lemma Gamma_fact: "Gamma (1 + of_nat n) = fact n" - by (simp add: pochhammer_fact pochhammer_Gamma of_nat_in_nonpos_Ints_iff - of_nat_Suc [symmetric] del: of_nat_Suc) - -lemma Gamma_numeral: "Gamma (numeral n) = fact (pred_numeral n)" - by (subst of_nat_numeral[symmetric], subst numeral_eq_Suc, - subst of_nat_Suc, subst Gamma_fact) (rule refl) - -lemma Gamma_of_int: "Gamma (of_int n) = (if n > 0 then fact (nat (n - 1)) else 0)" -proof (cases "n > 0") - case True - hence "Gamma (of_int n) = Gamma (of_nat (Suc (nat (n - 1))))" by (subst of_nat_Suc) simp_all - with True show ?thesis by (subst (asm) of_nat_Suc, subst (asm) Gamma_fact) simp -qed (simp_all add: Gamma_eq_zero_iff nonpos_Ints_of_int) - -lemma rGamma_of_int: "rGamma (of_int n) = (if n > 0 then inverse (fact (nat (n - 1))) else 0)" - by (simp add: Gamma_of_int rGamma_inverse_Gamma) - -lemma Gamma_seriesI: - assumes "(\n. g n / Gamma_series z n) \ 1" - shows "g \ Gamma z" -proof (rule Lim_transform_eventually) - have "1/2 > (0::real)" by simp - from tendstoD[OF assms, OF this] - show "eventually (\n. g n / Gamma_series z n * Gamma_series z n = g n) sequentially" - by (force elim!: eventually_mono simp: dist_real_def dist_0_norm) - from assms have "(\n. g n / Gamma_series z n * Gamma_series z n) \ 1 * Gamma z" - by (intro tendsto_intros) - thus "(\n. g n / Gamma_series z n * Gamma_series z n) \ Gamma z" by simp -qed - -lemma Gamma_seriesI': - assumes "f \ rGamma z" - assumes "(\n. g n * f n) \ 1" - assumes "z \ \\<^sub>\\<^sub>0" - shows "g \ Gamma z" -proof (rule Lim_transform_eventually) - have "1/2 > (0::real)" by simp - from tendstoD[OF assms(2), OF this] show "eventually (\n. g n * f n / f n = g n) sequentially" - by (force elim!: eventually_mono simp: dist_real_def dist_0_norm) - from assms have "(\n. g n * f n / f n) \ 1 / rGamma z" - by (intro tendsto_divide assms) (simp_all add: rGamma_eq_zero_iff) - thus "(\n. g n * f n / f n) \ Gamma z" by (simp add: Gamma_def divide_inverse) -qed - -lemma Gamma_series'_LIMSEQ: "Gamma_series' z \ Gamma z" - by (cases "z \ \\<^sub>\\<^sub>0") (simp_all add: Gamma_nonpos_Int Gamma_seriesI[OF Gamma_series_Gamma_series'] - Gamma_series'_nonpos_Ints_LIMSEQ[of z]) - - -subsection \Differentiability\ - -lemma has_field_derivative_rGamma_no_nonpos_int: - assumes "z \ \\<^sub>\\<^sub>0" - shows "(rGamma has_field_derivative -rGamma z * Digamma z) (at z within A)" -proof (rule has_field_derivative_at_within) - from assms have "z \ - of_nat n" for n by auto - from differentiable_rGamma_aux1[OF this] - show "(rGamma has_field_derivative -rGamma z * Digamma z) (at z)" - unfolding Digamma_def suminf_def sums_def[abs_def] - has_field_derivative_def has_derivative_def netlimit_at - by (simp add: Let_def bounded_linear_mult_right mult_ac of_real_def [symmetric]) -qed - -lemma has_field_derivative_rGamma_nonpos_int: - "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n) within A)" - apply (rule has_field_derivative_at_within) - using differentiable_rGamma_aux2[of n] - unfolding Let_def has_field_derivative_def has_derivative_def netlimit_at - by (simp only: bounded_linear_mult_right mult_ac of_real_def [symmetric] fact_setprod) simp - -lemma has_field_derivative_rGamma [derivative_intros]: - "(rGamma has_field_derivative (if z \ \\<^sub>\\<^sub>0 then (-1)^(nat \norm z\) * fact (nat \norm z\) - else -rGamma z * Digamma z)) (at z within A)" -using has_field_derivative_rGamma_no_nonpos_int[of z A] - has_field_derivative_rGamma_nonpos_int[of "nat \norm z\" A] - by (auto elim!: nonpos_Ints_cases') - -declare has_field_derivative_rGamma_no_nonpos_int [THEN DERIV_chain2, derivative_intros] -declare has_field_derivative_rGamma [THEN DERIV_chain2, derivative_intros] -declare has_field_derivative_rGamma_nonpos_int [derivative_intros] -declare has_field_derivative_rGamma_no_nonpos_int [derivative_intros] -declare has_field_derivative_rGamma [derivative_intros] - - -lemma has_field_derivative_Gamma [derivative_intros]: - "z \ \\<^sub>\\<^sub>0 \ (Gamma has_field_derivative Gamma z * Digamma z) (at z within A)" - unfolding Gamma_def [abs_def] - by (fastforce intro!: derivative_eq_intros simp: rGamma_eq_zero_iff) - -declare has_field_derivative_Gamma[THEN DERIV_chain2, derivative_intros] - -(* TODO: Hide ugly facts properly *) -hide_fact rGamma_eq_zero_iff_aux differentiable_rGamma_aux1 differentiable_rGamma_aux2 - differentiable_rGamma_aux2 rGamma_series_aux Gamma_class.rGamma_eq_zero_iff_aux - - - -(* TODO: differentiable etc. *) - - -subsection \Continuity\ - -lemma continuous_on_rGamma [continuous_intros]: "continuous_on A rGamma" - by (rule DERIV_continuous_on has_field_derivative_rGamma)+ - -lemma continuous_on_Gamma [continuous_intros]: "A \ \\<^sub>\\<^sub>0 = {} \ continuous_on A Gamma" - by (rule DERIV_continuous_on has_field_derivative_Gamma)+ blast - -lemma isCont_rGamma [continuous_intros]: - "isCont f z \ isCont (\x. rGamma (f x)) z" - by (rule isCont_o2[OF _ DERIV_isCont[OF has_field_derivative_rGamma]]) - -lemma isCont_Gamma [continuous_intros]: - "isCont f z \ f z \ \\<^sub>\\<^sub>0 \ isCont (\x. Gamma (f x)) z" - by (rule isCont_o2[OF _ DERIV_isCont[OF has_field_derivative_Gamma]]) - - - -text \The complex Gamma function\ - -instantiation complex :: Gamma -begin - -definition rGamma_complex :: "complex \ complex" where - "rGamma_complex z = lim (rGamma_series z)" - -lemma rGamma_series_complex_converges: - "convergent (rGamma_series (z :: complex))" (is "?thesis1") - and rGamma_complex_altdef: - "rGamma z = (if z \ \\<^sub>\\<^sub>0 then 0 else exp (-ln_Gamma z))" (is "?thesis2") -proof - - have "?thesis1 \ ?thesis2" - proof (cases "z \ \\<^sub>\\<^sub>0") - case False - have "rGamma_series z \ exp (- ln_Gamma z)" - proof (rule Lim_transform_eventually) - from ln_Gamma_series_complex_converges'[OF False] guess d by (elim exE conjE) - from this(1) uniformly_convergent_imp_convergent[OF this(2), of z] - have "ln_Gamma_series z \ lim (ln_Gamma_series z)" by (simp add: convergent_LIMSEQ_iff) - thus "(\n. exp (-ln_Gamma_series z n)) \ exp (- ln_Gamma z)" - unfolding convergent_def ln_Gamma_def by (intro tendsto_exp tendsto_minus) - from eventually_gt_at_top[of "0::nat"] exp_ln_Gamma_series_complex False - show "eventually (\n. exp (-ln_Gamma_series z n) = rGamma_series z n) sequentially" - by (force elim!: eventually_mono simp: exp_minus Gamma_series_def rGamma_series_def) - qed - with False show ?thesis - by (auto simp: convergent_def rGamma_complex_def intro!: limI) - next - case True - then obtain k where "z = - of_nat k" by (erule nonpos_Ints_cases') - also have "rGamma_series \ \ 0" - by (subst tendsto_cong[OF rGamma_series_minus_of_nat]) (simp_all add: convergent_const) - finally show ?thesis using True - by (auto simp: rGamma_complex_def convergent_def intro!: limI) - qed - thus "?thesis1" "?thesis2" by blast+ -qed - -context -begin - -(* TODO: duplication *) -private lemma rGamma_complex_plus1: "z * rGamma (z + 1) = rGamma (z :: complex)" -proof - - let ?f = "\n. (z + 1) * inverse (of_nat n) + 1" - have "eventually (\n. ?f n * rGamma_series z n = z * rGamma_series (z + 1) n) sequentially" - using eventually_gt_at_top[of "0::nat"] - proof eventually_elim - fix n :: nat assume n: "n > 0" - hence "z * rGamma_series (z + 1) n = inverse (of_nat n) * - pochhammer z (Suc (Suc n)) / (fact n * exp (z * of_real (ln (of_nat n))))" - by (subst pochhammer_rec) (simp add: rGamma_series_def field_simps exp_add exp_of_real) - also from n have "\ = ?f n * rGamma_series z n" - by (subst pochhammer_rec') (simp_all add: divide_simps rGamma_series_def add_ac) - finally show "?f n * rGamma_series z n = z * rGamma_series (z + 1) n" .. - qed - moreover have "(\n. ?f n * rGamma_series z n) \ ((z+1) * 0 + 1) * rGamma z" - using rGamma_series_complex_converges - by (intro tendsto_intros lim_inverse_n) - (simp_all add: convergent_LIMSEQ_iff rGamma_complex_def) - hence "(\n. ?f n * rGamma_series z n) \ rGamma z" by simp - ultimately have "(\n. z * rGamma_series (z + 1) n) \ rGamma z" - by (rule Lim_transform_eventually) - moreover have "(\n. z * rGamma_series (z + 1) n) \ z * rGamma (z + 1)" - using rGamma_series_complex_converges - by (auto intro!: tendsto_mult simp: rGamma_complex_def convergent_LIMSEQ_iff) - ultimately show "z * rGamma (z + 1) = rGamma z" using LIMSEQ_unique by blast -qed - -private lemma has_field_derivative_rGamma_complex_no_nonpos_Int: - assumes "(z :: complex) \ \\<^sub>\\<^sub>0" - shows "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)" -proof - - have diff: "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)" if "Re z > 0" for z - proof (subst DERIV_cong_ev[OF refl _ refl]) - from that have "eventually (\t. t \ ball z (Re z/2)) (nhds z)" - by (intro eventually_nhds_in_nhd) simp_all - thus "eventually (\t. rGamma t = exp (- ln_Gamma t)) (nhds z)" - using no_nonpos_Int_in_ball_complex[OF that] - by (auto elim!: eventually_mono simp: rGamma_complex_altdef) - next - have "z \ \\<^sub>\\<^sub>0" using that by (simp add: complex_nonpos_Reals_iff) - with that show "((\t. exp (- ln_Gamma t)) has_field_derivative (-rGamma z * Digamma z)) (at z)" - by (force elim!: nonpos_Ints_cases intro!: derivative_eq_intros simp: rGamma_complex_altdef) - qed - - from assms show "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)" - proof (induction "nat \1 - Re z\" arbitrary: z) - case (Suc n z) - from Suc.prems have z: "z \ 0" by auto - from Suc.hyps have "n = nat \- Re z\" by linarith - hence A: "n = nat \1 - Re (z + 1)\" by simp - from Suc.prems have B: "z + 1 \ \\<^sub>\\<^sub>0" by (force dest: plus_one_in_nonpos_Ints_imp) - - have "((\z. z * (rGamma \ (\z. z + 1)) z) has_field_derivative - -rGamma (z + 1) * (Digamma (z + 1) * z - 1)) (at z)" - by (rule derivative_eq_intros DERIV_chain Suc refl A B)+ (simp add: algebra_simps) - also have "(\z. z * (rGamma \ (\z. z + 1 :: complex)) z) = rGamma" - by (simp add: rGamma_complex_plus1) - also from z have "Digamma (z + 1) * z - 1 = z * Digamma z" - by (subst Digamma_plus1) (simp_all add: field_simps) - also have "-rGamma (z + 1) * (z * Digamma z) = -rGamma z * Digamma z" - by (simp add: rGamma_complex_plus1[of z, symmetric]) - finally show ?case . - qed (intro diff, simp) -qed - -private lemma rGamma_complex_1: "rGamma (1 :: complex) = 1" -proof - - have A: "eventually (\n. rGamma_series 1 n = of_nat (Suc n) / of_nat n) sequentially" - using eventually_gt_at_top[of "0::nat"] - by (force elim!: eventually_mono simp: rGamma_series_def exp_of_real pochhammer_fact - divide_simps pochhammer_rec' dest!: pochhammer_eq_0_imp_nonpos_Int) - have "rGamma_series 1 \ 1" by (subst tendsto_cong[OF A]) (rule LIMSEQ_Suc_n_over_n) - thus "rGamma 1 = (1 :: complex)" unfolding rGamma_complex_def by (rule limI) -qed - -private lemma has_field_derivative_rGamma_complex_nonpos_Int: - "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n :: complex))" -proof (induction n) - case 0 - have A: "(0::complex) + 1 \ \\<^sub>\\<^sub>0" by simp - have "((\z. z * (rGamma \ (\z. z + 1 :: complex)) z) has_field_derivative 1) (at 0)" - by (rule derivative_eq_intros DERIV_chain refl - has_field_derivative_rGamma_complex_no_nonpos_Int A)+ (simp add: rGamma_complex_1) - thus ?case by (simp add: rGamma_complex_plus1) -next - case (Suc n) - hence A: "(rGamma has_field_derivative (-1)^n * fact n) - (at (- of_nat (Suc n) + 1 :: complex))" by simp - have "((\z. z * (rGamma \ (\z. z + 1 :: complex)) z) has_field_derivative - (- 1) ^ Suc n * fact (Suc n)) (at (- of_nat (Suc n)))" - by (rule derivative_eq_intros refl A DERIV_chain)+ - (simp add: algebra_simps rGamma_complex_altdef) - thus ?case by (simp add: rGamma_complex_plus1) -qed - -instance proof - fix z :: complex show "(rGamma z = 0) \ (\n. z = - of_nat n)" - by (auto simp: rGamma_complex_altdef elim!: nonpos_Ints_cases') -next - fix z :: complex assume "\n. z \ - of_nat n" - hence "z \ \\<^sub>\\<^sub>0" by (auto elim!: nonpos_Ints_cases') - from has_field_derivative_rGamma_complex_no_nonpos_Int[OF this] - show "let d = (THE d. (\n. \k d) - euler_mascheroni *\<^sub>R 1 in (\y. (rGamma y - rGamma z + - rGamma z * d * (y - z)) /\<^sub>R cmod (y - z)) \z\ 0" - by (simp add: has_field_derivative_def has_derivative_def Digamma_def sums_def [abs_def] - netlimit_at of_real_def[symmetric] suminf_def) -next - fix n :: nat - from has_field_derivative_rGamma_complex_nonpos_Int[of n] - show "let z = - of_nat n in (\y. (rGamma y - rGamma z - (- 1) ^ n * setprod of_nat {1..n} * - (y - z)) /\<^sub>R cmod (y - z)) \z\ 0" - by (simp add: has_field_derivative_def has_derivative_def fact_setprod netlimit_at Let_def) -next - fix z :: complex - from rGamma_series_complex_converges[of z] have "rGamma_series z \ rGamma z" - by (simp add: convergent_LIMSEQ_iff rGamma_complex_def) - thus "let fact' = \n. setprod of_nat {1..n}; - exp = \x. THE e. (\n. \kR fact k) \ e; - pochhammer' = \a n. \n = 0..n. a + of_nat n - in (\n. pochhammer' z n / (fact' n * exp (z * ln (real_of_nat n) *\<^sub>R 1))) \ rGamma z" - by (simp add: fact_setprod pochhammer_Suc_setprod rGamma_series_def [abs_def] exp_def - of_real_def [symmetric] suminf_def sums_def [abs_def] atLeast0AtMost) -qed - -end -end - - -lemma Gamma_complex_altdef: - "Gamma z = (if z \ \\<^sub>\\<^sub>0 then 0 else exp (ln_Gamma (z :: complex)))" - unfolding Gamma_def rGamma_complex_altdef by (simp add: exp_minus) - -lemma cnj_rGamma: "cnj (rGamma z) = rGamma (cnj z)" -proof - - have "rGamma_series (cnj z) = (\n. cnj (rGamma_series z n))" - by (intro ext) (simp_all add: rGamma_series_def exp_cnj) - also have "... \ cnj (rGamma z)" by (intro tendsto_cnj tendsto_intros) - finally show ?thesis unfolding rGamma_complex_def by (intro sym[OF limI]) -qed - -lemma cnj_Gamma: "cnj (Gamma z) = Gamma (cnj z)" - unfolding Gamma_def by (simp add: cnj_rGamma) - -lemma Gamma_complex_real: - "z \ \ \ Gamma z \ (\ :: complex set)" and rGamma_complex_real: "z \ \ \ rGamma z \ \" - by (simp_all add: Reals_cnj_iff cnj_Gamma cnj_rGamma) - -lemma field_differentiable_rGamma: "rGamma field_differentiable (at z within A)" - using has_field_derivative_rGamma[of z] unfolding field_differentiable_def by blast - -lemma holomorphic_on_rGamma: "rGamma holomorphic_on A" - unfolding holomorphic_on_def by (auto intro!: field_differentiable_rGamma) - -lemma analytic_on_rGamma: "rGamma analytic_on A" - unfolding analytic_on_def by (auto intro!: exI[of _ 1] holomorphic_on_rGamma) - - -lemma field_differentiable_Gamma: "z \ \\<^sub>\\<^sub>0 \ Gamma field_differentiable (at z within A)" - using has_field_derivative_Gamma[of z] unfolding field_differentiable_def by auto - -lemma holomorphic_on_Gamma: "A \ \\<^sub>\\<^sub>0 = {} \ Gamma holomorphic_on A" - unfolding holomorphic_on_def by (auto intro!: field_differentiable_Gamma) - -lemma analytic_on_Gamma: "A \ \\<^sub>\\<^sub>0 = {} \ Gamma analytic_on A" - by (rule analytic_on_subset[of _ "UNIV - \\<^sub>\\<^sub>0"], subst analytic_on_open) - (auto intro!: holomorphic_on_Gamma) - -lemma has_field_derivative_rGamma_complex' [derivative_intros]: - "(rGamma has_field_derivative (if z \ \\<^sub>\\<^sub>0 then (-1)^(nat \-Re z\) * fact (nat \-Re z\) else - -rGamma z * Digamma z)) (at z within A)" - using has_field_derivative_rGamma[of z] by (auto elim!: nonpos_Ints_cases') - -declare has_field_derivative_rGamma_complex'[THEN DERIV_chain2, derivative_intros] - - -lemma field_differentiable_Polygamma: - fixes z::complex - shows - "z \ \\<^sub>\\<^sub>0 \ Polygamma n field_differentiable (at z within A)" - using has_field_derivative_Polygamma[of z n] unfolding field_differentiable_def by auto - -lemma holomorphic_on_Polygamma: "A \ \\<^sub>\\<^sub>0 = {} \ Polygamma n holomorphic_on A" - unfolding holomorphic_on_def by (auto intro!: field_differentiable_Polygamma) - -lemma analytic_on_Polygamma: "A \ \\<^sub>\\<^sub>0 = {} \ Polygamma n analytic_on A" - by (rule analytic_on_subset[of _ "UNIV - \\<^sub>\\<^sub>0"], subst analytic_on_open) - (auto intro!: holomorphic_on_Polygamma) - - - -text \The real Gamma function\ - -lemma rGamma_series_real: - "eventually (\n. rGamma_series x n = Re (rGamma_series (of_real x) n)) sequentially" - using eventually_gt_at_top[of "0 :: nat"] -proof eventually_elim - fix n :: nat assume n: "n > 0" - have "Re (rGamma_series (of_real x) n) = - Re (of_real (pochhammer x (Suc n)) / (fact n * exp (of_real (x * ln (real_of_nat n)))))" - using n by (simp add: rGamma_series_def powr_def Ln_of_nat pochhammer_of_real) - also from n have "\ = Re (of_real ((pochhammer x (Suc n)) / - (fact n * (exp (x * ln (real_of_nat n))))))" - by (subst exp_of_real) simp - also from n have "\ = rGamma_series x n" - by (subst Re_complex_of_real) (simp add: rGamma_series_def powr_def) - finally show "rGamma_series x n = Re (rGamma_series (of_real x) n)" .. -qed - -instantiation real :: Gamma -begin - -definition "rGamma_real x = Re (rGamma (of_real x :: complex))" - -instance proof - fix x :: real - have "rGamma x = Re (rGamma (of_real x))" by (simp add: rGamma_real_def) - also have "of_real \ = rGamma (of_real x :: complex)" - by (intro of_real_Re rGamma_complex_real) simp_all - also have "\ = 0 \ x \ \\<^sub>\\<^sub>0" by (simp add: rGamma_eq_zero_iff of_real_in_nonpos_Ints_iff) - also have "\ \ (\n. x = - of_nat n)" by (auto elim!: nonpos_Ints_cases') - finally show "(rGamma x) = 0 \ (\n. x = - real_of_nat n)" by simp -next - fix x :: real assume "\n. x \ - of_nat n" - hence x: "complex_of_real x \ \\<^sub>\\<^sub>0" - by (subst of_real_in_nonpos_Ints_iff) (auto elim!: nonpos_Ints_cases') - then have "x \ 0" by auto - with x have "(rGamma has_field_derivative - rGamma x * Digamma x) (at x)" - by (fastforce intro!: derivative_eq_intros has_vector_derivative_real_complex - simp: Polygamma_of_real rGamma_real_def [abs_def]) - thus "let d = (THE d. (\n. \k d) - euler_mascheroni *\<^sub>R 1 in (\y. (rGamma y - rGamma x + - rGamma x * d * (y - x)) /\<^sub>R norm (y - x)) \x\ 0" - by (simp add: has_field_derivative_def has_derivative_def Digamma_def sums_def [abs_def] - netlimit_at of_real_def[symmetric] suminf_def) -next - fix n :: nat - have "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n :: real))" - by (fastforce intro!: derivative_eq_intros has_vector_derivative_real_complex - simp: Polygamma_of_real rGamma_real_def [abs_def]) - thus "let x = - of_nat n in (\y. (rGamma y - rGamma x - (- 1) ^ n * setprod of_nat {1..n} * - (y - x)) /\<^sub>R norm (y - x)) \x::real\ 0" - by (simp add: has_field_derivative_def has_derivative_def fact_setprod netlimit_at Let_def) -next - fix x :: real - have "rGamma_series x \ rGamma x" - proof (rule Lim_transform_eventually) - show "(\n. Re (rGamma_series (of_real x) n)) \ rGamma x" unfolding rGamma_real_def - by (intro tendsto_intros) - qed (insert rGamma_series_real, simp add: eq_commute) - thus "let fact' = \n. setprod of_nat {1..n}; - exp = \x. THE e. (\n. \kR fact k) \ e; - pochhammer' = \a n. \n = 0..n. a + of_nat n - in (\n. pochhammer' x n / (fact' n * exp (x * ln (real_of_nat n) *\<^sub>R 1))) \ rGamma x" - by (simp add: fact_setprod pochhammer_Suc_setprod rGamma_series_def [abs_def] exp_def - of_real_def [symmetric] suminf_def sums_def [abs_def] atLeast0AtMost) -qed - -end - - -lemma rGamma_complex_of_real: "rGamma (complex_of_real x) = complex_of_real (rGamma x)" - unfolding rGamma_real_def using rGamma_complex_real by simp - -lemma Gamma_complex_of_real: "Gamma (complex_of_real x) = complex_of_real (Gamma x)" - unfolding Gamma_def by (simp add: rGamma_complex_of_real) - -lemma rGamma_real_altdef: "rGamma x = lim (rGamma_series (x :: real))" - by (rule sym, rule limI, rule tendsto_intros) - -lemma Gamma_real_altdef1: "Gamma x = lim (Gamma_series (x :: real))" - by (rule sym, rule limI, rule tendsto_intros) - -lemma Gamma_real_altdef2: "Gamma x = Re (Gamma (of_real x))" - using rGamma_complex_real[OF Reals_of_real[of x]] - by (elim Reals_cases) - (simp only: Gamma_def rGamma_real_def of_real_inverse[symmetric] Re_complex_of_real) - -lemma ln_Gamma_series_complex_of_real: - "x > 0 \ n > 0 \ ln_Gamma_series (complex_of_real x) n = of_real (ln_Gamma_series x n)" -proof - - assume xn: "x > 0" "n > 0" - have "Ln (complex_of_real x / of_nat k + 1) = of_real (ln (x / of_nat k + 1))" if "k \ 1" for k - using that xn by (subst Ln_of_real [symmetric]) (auto intro!: add_nonneg_pos simp: field_simps) - with xn show ?thesis by (simp add: ln_Gamma_series_def Ln_of_nat Ln_of_real) -qed - -lemma ln_Gamma_real_converges: - assumes "(x::real) > 0" - shows "convergent (ln_Gamma_series x)" -proof - - have "(\n. ln_Gamma_series (complex_of_real x) n) \ ln_Gamma (of_real x)" using assms - by (intro ln_Gamma_complex_LIMSEQ) (auto simp: of_real_in_nonpos_Ints_iff) - moreover from eventually_gt_at_top[of "0::nat"] - have "eventually (\n. complex_of_real (ln_Gamma_series x n) = - ln_Gamma_series (complex_of_real x) n) sequentially" - by eventually_elim (simp add: ln_Gamma_series_complex_of_real assms) - ultimately have "(\n. complex_of_real (ln_Gamma_series x n)) \ ln_Gamma (of_real x)" - by (subst tendsto_cong) assumption+ - from tendsto_Re[OF this] show ?thesis by (auto simp: convergent_def) -qed - -lemma ln_Gamma_real_LIMSEQ: "(x::real) > 0 \ ln_Gamma_series x \ ln_Gamma x" - using ln_Gamma_real_converges[of x] unfolding ln_Gamma_def by (simp add: convergent_LIMSEQ_iff) - -lemma ln_Gamma_complex_of_real: "x > 0 \ ln_Gamma (complex_of_real x) = of_real (ln_Gamma x)" -proof (unfold ln_Gamma_def, rule limI, rule Lim_transform_eventually) - assume x: "x > 0" - show "eventually (\n. of_real (ln_Gamma_series x n) = - ln_Gamma_series (complex_of_real x) n) sequentially" - using eventually_gt_at_top[of "0::nat"] - by eventually_elim (simp add: ln_Gamma_series_complex_of_real x) -qed (intro tendsto_of_real, insert ln_Gamma_real_LIMSEQ[of x], simp add: ln_Gamma_def) - -lemma Gamma_real_pos_exp: "x > (0 :: real) \ Gamma x = exp (ln_Gamma x)" - by (auto simp: Gamma_real_altdef2 Gamma_complex_altdef of_real_in_nonpos_Ints_iff - ln_Gamma_complex_of_real exp_of_real) - -lemma ln_Gamma_real_pos: "x > 0 \ ln_Gamma x = ln (Gamma x :: real)" - unfolding Gamma_real_pos_exp by simp - -lemma Gamma_real_pos: "x > (0::real) \ Gamma x > 0" - by (simp add: Gamma_real_pos_exp) - -lemma has_field_derivative_ln_Gamma_real [derivative_intros]: - assumes x: "x > (0::real)" - shows "(ln_Gamma has_field_derivative Digamma x) (at x)" -proof (subst DERIV_cong_ev[OF refl _ refl]) - from assms show "((Re \ ln_Gamma \ complex_of_real) has_field_derivative Digamma x) (at x)" - by (auto intro!: derivative_eq_intros has_vector_derivative_real_complex - simp: Polygamma_of_real o_def) - from eventually_nhds_in_nhd[of x "{0<..}"] assms - show "eventually (\y. ln_Gamma y = (Re \ ln_Gamma \ of_real) y) (nhds x)" - by (auto elim!: eventually_mono simp: ln_Gamma_complex_of_real interior_open) -qed - -declare has_field_derivative_ln_Gamma_real[THEN DERIV_chain2, derivative_intros] - - -lemma has_field_derivative_rGamma_real' [derivative_intros]: - "(rGamma has_field_derivative (if x \ \\<^sub>\\<^sub>0 then (-1)^(nat \-x\) * fact (nat \-x\) else - -rGamma x * Digamma x)) (at x within A)" - using has_field_derivative_rGamma[of x] by (force elim!: nonpos_Ints_cases') - -declare has_field_derivative_rGamma_real'[THEN DERIV_chain2, derivative_intros] - -lemma Polygamma_real_odd_pos: - assumes "(x::real) \ \\<^sub>\\<^sub>0" "odd n" - shows "Polygamma n x > 0" -proof - - from assms have "x \ 0" by auto - with assms show ?thesis - unfolding Polygamma_def using Polygamma_converges'[of x "Suc n"] - by (auto simp: zero_less_power_eq simp del: power_Suc - dest: plus_of_nat_eq_0_imp intro!: mult_pos_pos suminf_pos) -qed - -lemma Polygamma_real_even_neg: - assumes "(x::real) > 0" "n > 0" "even n" - shows "Polygamma n x < 0" - using assms unfolding Polygamma_def using Polygamma_converges'[of x "Suc n"] - by (auto intro!: mult_pos_pos suminf_pos) - -lemma Polygamma_real_strict_mono: - assumes "x > 0" "x < (y::real)" "even n" - shows "Polygamma n x < Polygamma n y" -proof - - have "\\. x < \ \ \ < y \ Polygamma n y - Polygamma n x = (y - x) * Polygamma (Suc n) \" - using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases) - then guess \ by (elim exE conjE) note \ = this - note \(3) - also from \(1,2) assms have "(y - x) * Polygamma (Suc n) \ > 0" - by (intro mult_pos_pos Polygamma_real_odd_pos) (auto elim!: nonpos_Ints_cases) - finally show ?thesis by simp -qed - -lemma Polygamma_real_strict_antimono: - assumes "x > 0" "x < (y::real)" "odd n" - shows "Polygamma n x > Polygamma n y" -proof - - have "\\. x < \ \ \ < y \ Polygamma n y - Polygamma n x = (y - x) * Polygamma (Suc n) \" - using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases) - then guess \ by (elim exE conjE) note \ = this - note \(3) - also from \(1,2) assms have "(y - x) * Polygamma (Suc n) \ < 0" - by (intro mult_pos_neg Polygamma_real_even_neg) simp_all - finally show ?thesis by simp -qed - -lemma Polygamma_real_mono: - assumes "x > 0" "x \ (y::real)" "even n" - shows "Polygamma n x \ Polygamma n y" - using Polygamma_real_strict_mono[OF assms(1) _ assms(3), of y] assms(2) - by (cases "x = y") simp_all - -lemma Digamma_real_ge_three_halves_pos: - assumes "x \ 3/2" - shows "Digamma (x :: real) > 0" -proof - - have "0 < Digamma (3/2 :: real)" by (fact Digamma_real_three_halves_pos) - also from assms have "\ \ Digamma x" by (intro Polygamma_real_mono) simp_all - finally show ?thesis . -qed - -lemma ln_Gamma_real_strict_mono: - assumes "x \ 3/2" "x < y" - shows "ln_Gamma (x :: real) < ln_Gamma y" -proof - - have "\\. x < \ \ \ < y \ ln_Gamma y - ln_Gamma x = (y - x) * Digamma \" - using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases) - then guess \ by (elim exE conjE) note \ = this - note \(3) - also from \(1,2) assms have "(y - x) * Digamma \ > 0" - by (intro mult_pos_pos Digamma_real_ge_three_halves_pos) simp_all - finally show ?thesis by simp -qed - -lemma Gamma_real_strict_mono: - assumes "x \ 3/2" "x < y" - shows "Gamma (x :: real) < Gamma y" -proof - - from Gamma_real_pos_exp[of x] assms have "Gamma x = exp (ln_Gamma x)" by simp - also have "\ < exp (ln_Gamma y)" by (intro exp_less_mono ln_Gamma_real_strict_mono assms) - also from Gamma_real_pos_exp[of y] assms have "\ = Gamma y" by simp - finally show ?thesis . -qed - -lemma log_convex_Gamma_real: "convex_on {0<..} (ln \ Gamma :: real \ real)" - by (rule convex_on_realI[of _ _ Digamma]) - (auto intro!: derivative_eq_intros Polygamma_real_mono Gamma_real_pos - simp: o_def Gamma_eq_zero_iff elim!: nonpos_Ints_cases') - - -subsection \Beta function\ - -definition Beta where "Beta a b = Gamma a * Gamma b / Gamma (a + b)" - -lemma Beta_altdef: "Beta a b = Gamma a * Gamma b * rGamma (a + b)" - by (simp add: inverse_eq_divide Beta_def Gamma_def) - -lemma Beta_commute: "Beta a b = Beta b a" - unfolding Beta_def by (simp add: ac_simps) - -lemma has_field_derivative_Beta1 [derivative_intros]: - assumes "x \ \\<^sub>\\<^sub>0" "x + y \ \\<^sub>\\<^sub>0" - shows "((\x. Beta x y) has_field_derivative (Beta x y * (Digamma x - Digamma (x + y)))) - (at x within A)" unfolding Beta_altdef - by (rule DERIV_cong, (rule derivative_intros assms)+) (simp add: algebra_simps) - -lemma Beta_pole1: "x \ \\<^sub>\\<^sub>0 \ Beta x y = 0" - by (auto simp add: Beta_def elim!: nonpos_Ints_cases') - -lemma Beta_pole2: "y \ \\<^sub>\\<^sub>0 \ Beta x y = 0" - by (auto simp add: Beta_def elim!: nonpos_Ints_cases') - -lemma Beta_zero: "x + y \ \\<^sub>\\<^sub>0 \ Beta x y = 0" - by (auto simp add: Beta_def elim!: nonpos_Ints_cases') - -lemma has_field_derivative_Beta2 [derivative_intros]: - assumes "y \ \\<^sub>\\<^sub>0" "x + y \ \\<^sub>\\<^sub>0" - shows "((\y. Beta x y) has_field_derivative (Beta x y * (Digamma y - Digamma (x + y)))) - (at y within A)" - using has_field_derivative_Beta1[of y x A] assms by (simp add: Beta_commute add_ac) - -lemma Beta_plus1_plus1: - assumes "x \ \\<^sub>\\<^sub>0" "y \ \\<^sub>\\<^sub>0" - shows "Beta (x + 1) y + Beta x (y + 1) = Beta x y" -proof - - have "Beta (x + 1) y + Beta x (y + 1) = - (Gamma (x + 1) * Gamma y + Gamma x * Gamma (y + 1)) * rGamma ((x + y) + 1)" - by (simp add: Beta_altdef add_divide_distrib algebra_simps) - also have "\ = (Gamma x * Gamma y) * ((x + y) * rGamma ((x + y) + 1))" - by (subst assms[THEN Gamma_plus1])+ (simp add: algebra_simps) - also from assms have "\ = Beta x y" unfolding Beta_altdef by (subst rGamma_plus1) simp - finally show ?thesis . -qed - -lemma Beta_plus1_left: - assumes "x \ \\<^sub>\\<^sub>0" - shows "(x + y) * Beta (x + 1) y = x * Beta x y" -proof - - have "(x + y) * Beta (x + 1) y = Gamma (x + 1) * Gamma y * ((x + y) * rGamma ((x + y) + 1))" - unfolding Beta_altdef by (simp only: ac_simps) - also have "\ = x * Beta x y" unfolding Beta_altdef - by (subst assms[THEN Gamma_plus1] rGamma_plus1)+ (simp only: ac_simps) - finally show ?thesis . -qed - -lemma Beta_plus1_right: - assumes "y \ \\<^sub>\\<^sub>0" - shows "(x + y) * Beta x (y + 1) = y * Beta x y" - using Beta_plus1_left[of y x] assms by (simp_all add: Beta_commute add.commute) - -lemma Gamma_Gamma_Beta: - assumes "x + y \ \\<^sub>\\<^sub>0" - shows "Gamma x * Gamma y = Beta x y * Gamma (x + y)" - unfolding Beta_altdef using assms Gamma_eq_zero_iff[of "x+y"] - by (simp add: rGamma_inverse_Gamma) - - - -subsection \Legendre duplication theorem\ - -context -begin - -private lemma Gamma_legendre_duplication_aux: - fixes z :: "'a :: Gamma" - assumes "z \ \\<^sub>\\<^sub>0" "z + 1/2 \ \\<^sub>\\<^sub>0" - shows "Gamma z * Gamma (z + 1/2) = exp ((1 - 2*z) * of_real (ln 2)) * Gamma (1/2) * Gamma (2*z)" -proof - - let ?powr = "\b a. exp (a * of_real (ln (of_nat b)))" - let ?h = "\n. (fact (n-1))\<^sup>2 / fact (2*n-1) * of_nat (2^(2*n)) * - exp (1/2 * of_real (ln (real_of_nat n)))" - { - fix z :: 'a assume z: "z \ \\<^sub>\\<^sub>0" "z + 1/2 \ \\<^sub>\\<^sub>0" - let ?g = "\n. ?powr 2 (2*z) * Gamma_series' z n * Gamma_series' (z + 1/2) n / - Gamma_series' (2*z) (2*n)" - have "eventually (\n. ?g n = ?h n) sequentially" using eventually_gt_at_top - proof eventually_elim - fix n :: nat assume n: "n > 0" - let ?f = "fact (n - 1) :: 'a" and ?f' = "fact (2*n - 1) :: 'a" - have A: "exp t * exp t = exp (2*t :: 'a)" for t by (subst exp_add [symmetric]) simp - have A: "Gamma_series' z n * Gamma_series' (z + 1/2) n = ?f^2 * ?powr n (2*z + 1/2) / - (pochhammer z n * pochhammer (z + 1/2) n)" - by (simp add: Gamma_series'_def exp_add ring_distribs power2_eq_square A mult_ac) - have B: "Gamma_series' (2*z) (2*n) = - ?f' * ?powr 2 (2*z) * ?powr n (2*z) / - (of_nat (2^(2*n)) * pochhammer z n * pochhammer (z+1/2) n)" using n - by (simp add: Gamma_series'_def ln_mult exp_add ring_distribs pochhammer_double) - from z have "pochhammer z n \ 0" by (auto dest: pochhammer_eq_0_imp_nonpos_Int) - moreover from z have "pochhammer (z + 1/2) n \ 0" by (auto dest: pochhammer_eq_0_imp_nonpos_Int) - ultimately have "?powr 2 (2*z) * (Gamma_series' z n * Gamma_series' (z + 1/2) n) / Gamma_series' (2*z) (2*n) = - ?f^2 / ?f' * of_nat (2^(2*n)) * (?powr n ((4*z + 1)/2) * ?powr n (-2*z))" - using n unfolding A B by (simp add: divide_simps exp_minus) - also have "?powr n ((4*z + 1)/2) * ?powr n (-2*z) = ?powr n (1/2)" - by (simp add: algebra_simps exp_add[symmetric] add_divide_distrib) - finally show "?g n = ?h n" by (simp only: mult_ac) - qed - - moreover from z double_in_nonpos_Ints_imp[of z] have "2 * z \ \\<^sub>\\<^sub>0" by auto - hence "?g \ ?powr 2 (2*z) * Gamma z * Gamma (z+1/2) / Gamma (2*z)" - using LIMSEQ_subseq_LIMSEQ[OF Gamma_series'_LIMSEQ, of "op*2" "2*z"] - by (intro tendsto_intros Gamma_series'_LIMSEQ) - (simp_all add: o_def subseq_def Gamma_eq_zero_iff) - ultimately have "?h \ ?powr 2 (2*z) * Gamma z * Gamma (z+1/2) / Gamma (2*z)" - by (rule Lim_transform_eventually) - } note lim = this - - from assms double_in_nonpos_Ints_imp[of z] have z': "2 * z \ \\<^sub>\\<^sub>0" by auto - from fraction_not_in_ints[of 2 1] have "(1/2 :: 'a) \ \\<^sub>\\<^sub>0" - by (intro not_in_Ints_imp_not_in_nonpos_Ints) simp_all - with lim[of "1/2 :: 'a"] have "?h \ 2 * Gamma (1 / 2 :: 'a)" by (simp add: exp_of_real) - from LIMSEQ_unique[OF this lim[OF assms]] z' show ?thesis - by (simp add: divide_simps Gamma_eq_zero_iff ring_distribs exp_diff exp_of_real ac_simps) -qed - -(* TODO: perhaps this is unnecessary once we have the fact that a holomorphic function is - infinitely differentiable *) -private lemma Gamma_reflection_aux: - defines "h \ \z::complex. if z \ \ then 0 else - (of_real pi * cot (of_real pi*z) + Digamma z - Digamma (1 - z))" - defines "a \ complex_of_real pi" - obtains h' where "continuous_on UNIV h'" "\z. (h has_field_derivative (h' z)) (at z)" -proof - - define f where "f n = a * of_real (cos_coeff (n+1) - sin_coeff (n+2))" for n - define F where "F z = (if z = 0 then 0 else (cos (a*z) - sin (a*z)/(a*z)) / z)" for z - define g where "g n = complex_of_real (sin_coeff (n+1))" for n - define G where "G z = (if z = 0 then 1 else sin (a*z)/(a*z))" for z - have a_nz: "a \ 0" unfolding a_def by simp - - have "(\n. f n * (a*z)^n) sums (F z) \ (\n. g n * (a*z)^n) sums (G z)" - if "abs (Re z) < 1" for z - proof (cases "z = 0"; rule conjI) - assume "z \ 0" - note z = this that - - from z have sin_nz: "sin (a*z) \ 0" unfolding a_def by (auto simp: sin_eq_0) - have "(\n. of_real (sin_coeff n) * (a*z)^n) sums (sin (a*z))" using sin_converges[of "a*z"] - by (simp add: scaleR_conv_of_real) - from sums_split_initial_segment[OF this, of 1] - have "(\n. (a*z) * of_real (sin_coeff (n+1)) * (a*z)^n) sums (sin (a*z))" by (simp add: mult_ac) - from sums_mult[OF this, of "inverse (a*z)"] z a_nz - have A: "(\n. g n * (a*z)^n) sums (sin (a*z)/(a*z))" - by (simp add: field_simps g_def) - with z show "(\n. g n * (a*z)^n) sums (G z)" by (simp add: G_def) - from A z a_nz sin_nz have g_nz: "(\n. g n * (a*z)^n) \ 0" by (simp add: sums_iff g_def) - - have [simp]: "sin_coeff (Suc 0) = 1" by (simp add: sin_coeff_def) - from sums_split_initial_segment[OF sums_diff[OF cos_converges[of "a*z"] A], of 1] - have "(\n. z * f n * (a*z)^n) sums (cos (a*z) - sin (a*z) / (a*z))" - by (simp add: mult_ac scaleR_conv_of_real ring_distribs f_def g_def) - from sums_mult[OF this, of "inverse z"] z assms - show "(\n. f n * (a*z)^n) sums (F z)" by (simp add: divide_simps mult_ac f_def F_def) - next - assume z: "z = 0" - have "(\n. f n * (a * z) ^ n) sums f 0" using powser_sums_zero[of f] z by simp - with z show "(\n. f n * (a * z) ^ n) sums (F z)" - by (simp add: f_def F_def sin_coeff_def cos_coeff_def) - have "(\n. g n * (a * z) ^ n) sums g 0" using powser_sums_zero[of g] z by simp - with z show "(\n. g n * (a * z) ^ n) sums (G z)" - by (simp add: g_def G_def sin_coeff_def cos_coeff_def) - qed - note sums = conjunct1[OF this] conjunct2[OF this] - - define h2 where [abs_def]: - "h2 z = (\n. f n * (a*z)^n) / (\n. g n * (a*z)^n) + Digamma (1 + z) - Digamma (1 - z)" for z - define POWSER where [abs_def]: "POWSER f z = (\n. f n * (z^n :: complex))" for f z - define POWSER' where [abs_def]: "POWSER' f z = (\n. diffs f n * (z^n))" for f and z :: complex - define h2' where [abs_def]: - "h2' z = a * (POWSER g (a*z) * POWSER' f (a*z) - POWSER f (a*z) * POWSER' g (a*z)) / - (POWSER g (a*z))^2 + Polygamma 1 (1 + z) + Polygamma 1 (1 - z)" for z - - have h_eq: "h t = h2 t" if "abs (Re t) < 1" for t - proof - - from that have t: "t \ \ \ t = 0" by (auto elim!: Ints_cases simp: dist_0_norm) - hence "h t = a*cot (a*t) - 1/t + Digamma (1 + t) - Digamma (1 - t)" - unfolding h_def using Digamma_plus1[of t] by (force simp: field_simps a_def) - also have "a*cot (a*t) - 1/t = (F t) / (G t)" - using t by (auto simp add: divide_simps sin_eq_0 cot_def a_def F_def G_def) - also have "\ = (\n. f n * (a*t)^n) / (\n. g n * (a*t)^n)" - using sums[of t] that by (simp add: sums_iff dist_0_norm) - finally show "h t = h2 t" by (simp only: h2_def) - qed - - let ?A = "{z. abs (Re z) < 1}" - have "open ({z. Re z < 1} \ {z. Re z > -1})" - using open_halfspace_Re_gt open_halfspace_Re_lt by auto - also have "({z. Re z < 1} \ {z. Re z > -1}) = {z. abs (Re z) < 1}" by auto - finally have open_A: "open ?A" . - hence [simp]: "interior ?A = ?A" by (simp add: interior_open) - - have summable_f: "summable (\n. f n * z^n)" for z - by (rule powser_inside, rule sums_summable, rule sums[of "\ * of_real (norm z + 1) / a"]) - (simp_all add: norm_mult a_def del: of_real_add) - have summable_g: "summable (\n. g n * z^n)" for z - by (rule powser_inside, rule sums_summable, rule sums[of "\ * of_real (norm z + 1) / a"]) - (simp_all add: norm_mult a_def del: of_real_add) - have summable_fg': "summable (\n. diffs f n * z^n)" "summable (\n. diffs g n * z^n)" for z - by (intro termdiff_converges_all summable_f summable_g)+ - have "(POWSER f has_field_derivative (POWSER' f z)) (at z)" - "(POWSER g has_field_derivative (POWSER' g z)) (at z)" for z - unfolding POWSER_def POWSER'_def - by (intro termdiffs_strong_converges_everywhere summable_f summable_g)+ - note derivs = this[THEN DERIV_chain2[OF _ DERIV_cmult[OF DERIV_ident]], unfolded POWSER_def] - have "isCont (POWSER f) z" "isCont (POWSER g) z" "isCont (POWSER' f) z" "isCont (POWSER' g) z" - for z unfolding POWSER_def POWSER'_def - by (intro isCont_powser_converges_everywhere summable_f summable_g summable_fg')+ - note cont = this[THEN isCont_o2[rotated], unfolded POWSER_def POWSER'_def] - - { - fix z :: complex assume z: "abs (Re z) < 1" - define d where "d = \ * of_real (norm z + 1)" - have d: "abs (Re d) < 1" "norm z < norm d" by (simp_all add: d_def norm_mult del: of_real_add) - have "eventually (\z. h z = h2 z) (nhds z)" - using eventually_nhds_in_nhd[of z ?A] using h_eq z - by (auto elim!: eventually_mono simp: dist_0_norm) - - moreover from sums(2)[OF z] z have nz: "(\n. g n * (a * z) ^ n) \ 0" - unfolding G_def by (auto simp: sums_iff sin_eq_0 a_def) - have A: "z \ \ \ z = 0" using z by (auto elim!: Ints_cases) - have no_int: "1 + z \ \ \ z = 0" using z Ints_diff[of "1+z" 1] A - by (auto elim!: nonpos_Ints_cases) - have no_int': "1 - z \ \ \ z = 0" using z Ints_diff[of 1 "1-z"] A - by (auto elim!: nonpos_Ints_cases) - from no_int no_int' have no_int: "1 - z \ \\<^sub>\\<^sub>0" "1 + z \ \\<^sub>\\<^sub>0" by auto - have "(h2 has_field_derivative h2' z) (at z)" unfolding h2_def - by (rule DERIV_cong, (rule derivative_intros refl derivs[unfolded POWSER_def] nz no_int)+) - (auto simp: h2'_def POWSER_def field_simps power2_eq_square) - ultimately have deriv: "(h has_field_derivative h2' z) (at z)" - by (subst DERIV_cong_ev[OF refl _ refl]) - - from sums(2)[OF z] z have "(\n. g n * (a * z) ^ n) \ 0" - unfolding G_def by (auto simp: sums_iff a_def sin_eq_0) - hence "isCont h2' z" using no_int unfolding h2'_def[abs_def] POWSER_def POWSER'_def - by (intro continuous_intros cont - continuous_on_compose2[OF _ continuous_on_Polygamma[of "{z. Re z > 0}"]]) auto - note deriv and this - } note A = this - - interpret h: periodic_fun_simple' h - proof - fix z :: complex - show "h (z + 1) = h z" - proof (cases "z \ \") - assume z: "z \ \" - hence A: "z + 1 \ \" "z \ 0" using Ints_diff[of "z+1" 1] by auto - hence "Digamma (z + 1) - Digamma (-z) = Digamma z - Digamma (-z + 1)" - by (subst (1 2) Digamma_plus1) simp_all - with A z show "h (z + 1) = h z" - by (simp add: h_def sin_plus_pi cos_plus_pi ring_distribs cot_def) - qed (simp add: h_def) - qed - - have h2'_eq: "h2' (z - 1) = h2' z" if z: "Re z > 0" "Re z < 1" for z - proof - - have "((\z. h (z - 1)) has_field_derivative h2' (z - 1)) (at z)" - by (rule DERIV_cong, rule DERIV_chain'[OF _ A(1)]) - (insert z, auto intro!: derivative_eq_intros) - hence "(h has_field_derivative h2' (z - 1)) (at z)" by (subst (asm) h.minus_1) - moreover from z have "(h has_field_derivative h2' z) (at z)" by (intro A) simp_all - ultimately show "h2' (z - 1) = h2' z" by (rule DERIV_unique) - qed - - define h2'' where "h2'' z = h2' (z - of_int \Re z\)" for z - have deriv: "(h has_field_derivative h2'' z) (at z)" for z - proof - - fix z :: complex - have B: "\Re z - real_of_int \Re z\\ < 1" by linarith - have "((\t. h (t - of_int \Re z\)) has_field_derivative h2'' z) (at z)" - unfolding h2''_def by (rule DERIV_cong, rule DERIV_chain'[OF _ A(1)]) - (insert B, auto intro!: derivative_intros) - thus "(h has_field_derivative h2'' z) (at z)" by (simp add: h.minus_of_int) - qed - - have cont: "continuous_on UNIV h2''" - proof (intro continuous_at_imp_continuous_on ballI) - fix z :: complex - define r where "r = \Re z\" - define A where "A = {t. of_int r - 1 < Re t \ Re t < of_int r + 1}" - have "continuous_on A (\t. h2' (t - of_int r))" unfolding A_def - by (intro continuous_at_imp_continuous_on isCont_o2[OF _ A(2)] ballI continuous_intros) - (simp_all add: abs_real_def) - moreover have "h2'' t = h2' (t - of_int r)" if t: "t \ A" for t - proof (cases "Re t \ of_int r") - case True - from t have "of_int r - 1 < Re t" "Re t < of_int r + 1" by (simp_all add: A_def) - with True have "\Re t\ = \Re z\" unfolding r_def by linarith - thus ?thesis by (auto simp: r_def h2''_def) - next - case False - from t have t: "of_int r - 1 < Re t" "Re t < of_int r + 1" by (simp_all add: A_def) - with False have t': "\Re t\ = \Re z\ - 1" unfolding r_def by linarith - moreover from t False have "h2' (t - of_int r + 1 - 1) = h2' (t - of_int r + 1)" - by (intro h2'_eq) simp_all - ultimately show ?thesis by (auto simp: r_def h2''_def algebra_simps t') - qed - ultimately have "continuous_on A h2''" by (subst continuous_on_cong[OF refl]) - moreover { - have "open ({t. of_int r - 1 < Re t} \ {t. of_int r + 1 > Re t})" - by (intro open_Int open_halfspace_Re_gt open_halfspace_Re_lt) - also have "{t. of_int r - 1 < Re t} \ {t. of_int r + 1 > Re t} = A" - unfolding A_def by blast - finally have "open A" . - } - ultimately have C: "isCont h2'' t" if "t \ A" for t using that - by (subst (asm) continuous_on_eq_continuous_at) auto - have "of_int r - 1 < Re z" "Re z < of_int r + 1" unfolding r_def by linarith+ - thus "isCont h2'' z" by (intro C) (simp_all add: A_def) - qed - - from that[OF cont deriv] show ?thesis . -qed - -lemma Gamma_reflection_complex: - fixes z :: complex - shows "Gamma z * Gamma (1 - z) = of_real pi / sin (of_real pi * z)" -proof - - let ?g = "\z::complex. Gamma z * Gamma (1 - z) * sin (of_real pi * z)" - define g where [abs_def]: "g z = (if z \ \ then of_real pi else ?g z)" for z :: complex - let ?h = "\z::complex. (of_real pi * cot (of_real pi*z) + Digamma z - Digamma (1 - z))" - define h where [abs_def]: "h z = (if z \ \ then 0 else ?h z)" for z :: complex - - \ \@{term g} is periodic with period 1.\ - interpret g: periodic_fun_simple' g - proof - fix z :: complex - show "g (z + 1) = g z" - proof (cases "z \ \") - case False - hence "z * g z = z * Beta z (- z + 1) * sin (of_real pi * z)" by (simp add: g_def Beta_def) - also have "z * Beta z (- z + 1) = (z + 1 + -z) * Beta (z + 1) (- z + 1)" - using False Ints_diff[of 1 "1 - z"] nonpos_Ints_subset_Ints - by (subst Beta_plus1_left [symmetric]) auto - also have "\ * sin (of_real pi * z) = z * (Beta (z + 1) (-z) * sin (of_real pi * (z + 1)))" - using False Ints_diff[of "z+1" 1] Ints_minus[of "-z"] nonpos_Ints_subset_Ints - by (subst Beta_plus1_right) (auto simp: ring_distribs sin_plus_pi) - also from False have "Beta (z + 1) (-z) * sin (of_real pi * (z + 1)) = g (z + 1)" - using Ints_diff[of "z+1" 1] by (auto simp: g_def Beta_def) - finally show "g (z + 1) = g z" using False by (subst (asm) mult_left_cancel) auto - qed (simp add: g_def) - qed - - \ \@{term g} is entire.\ - have g_g': "(g has_field_derivative (h z * g z)) (at z)" for z :: complex - proof (cases "z \ \") - let ?h' = "\z. Beta z (1 - z) * ((Digamma z - Digamma (1 - z)) * sin (z * of_real pi) + - of_real pi * cos (z * of_real pi))" - case False - from False have "eventually (\t. t \ UNIV - \) (nhds z)" - by (intro eventually_nhds_in_open) (auto simp: open_Diff) - hence "eventually (\t. g t = ?g t) (nhds z)" by eventually_elim (simp add: g_def) - moreover { - from False Ints_diff[of 1 "1-z"] have "1 - z \ \" by auto - hence "(?g has_field_derivative ?h' z) (at z)" using nonpos_Ints_subset_Ints - by (auto intro!: derivative_eq_intros simp: algebra_simps Beta_def) - also from False have "sin (of_real pi * z) \ 0" by (subst sin_eq_0) auto - hence "?h' z = h z * g z" - using False unfolding g_def h_def cot_def by (simp add: field_simps Beta_def) - finally have "(?g has_field_derivative (h z * g z)) (at z)" . - } - ultimately show ?thesis by (subst DERIV_cong_ev[OF refl _ refl]) - next - case True - then obtain n where z: "z = of_int n" by (auto elim!: Ints_cases) - let ?t = "(\z::complex. if z = 0 then 1 else sin z / z) \ (\z. of_real pi * z)" - have deriv_0: "(g has_field_derivative 0) (at 0)" - proof (subst DERIV_cong_ev[OF refl _ refl]) - show "eventually (\z. g z = of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z) (nhds 0)" - using eventually_nhds_ball[OF zero_less_one, of "0::complex"] - proof eventually_elim - fix z :: complex assume z: "z \ ball 0 1" - show "g z = of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z" - proof (cases "z = 0") - assume z': "z \ 0" - with z have z'': "z \ \\<^sub>\\<^sub>0" "z \ \" by (auto elim!: Ints_cases simp: dist_0_norm) - from Gamma_plus1[OF this(1)] have "Gamma z = Gamma (z + 1) / z" by simp - with z'' z' show ?thesis by (simp add: g_def ac_simps) - qed (simp add: g_def) - qed - have "(?t has_field_derivative (0 * of_real pi)) (at 0)" - using has_field_derivative_sin_z_over_z[of "UNIV :: complex set"] - by (intro DERIV_chain) simp_all - thus "((\z. of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z) has_field_derivative 0) (at 0)" - by (auto intro!: derivative_eq_intros simp: o_def) - qed - - have "((g \ (\x. x - of_int n)) has_field_derivative 0 * 1) (at (of_int n))" - using deriv_0 by (intro DERIV_chain) (auto intro!: derivative_eq_intros) - also have "g \ (\x. x - of_int n) = g" by (intro ext) (simp add: g.minus_of_int) - finally show "(g has_field_derivative (h z * g z)) (at z)" by (simp add: z h_def) - qed - - have g_eq: "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * g z" if "Re z > -1" "Re z < 2" for z - proof (cases "z \ \") - case True - with that have "z = 0 \ z = 1" by (force elim!: Ints_cases) - moreover have "g 0 * g (1/2) = Gamma (1/2)^2 * g 0" - using fraction_not_in_ints[where 'a = complex, of 2 1] by (simp add: g_def power2_eq_square) - moreover have "g (1/2) * g 1 = Gamma (1/2)^2 * g 1" - using fraction_not_in_ints[where 'a = complex, of 2 1] - by (simp add: g_def power2_eq_square Beta_def algebra_simps) - ultimately show ?thesis by force - next - case False - hence z: "z/2 \ \" "(z+1)/2 \ \" using Ints_diff[of "z+1" 1] by (auto elim!: Ints_cases) - hence z': "z/2 \ \\<^sub>\\<^sub>0" "(z+1)/2 \ \\<^sub>\\<^sub>0" by (auto elim!: nonpos_Ints_cases) - from z have "1-z/2 \ \" "1-((z+1)/2) \ \" - using Ints_diff[of 1 "1-z/2"] Ints_diff[of 1 "1-((z+1)/2)"] by auto - hence z'': "1-z/2 \ \\<^sub>\\<^sub>0" "1-((z+1)/2) \ \\<^sub>\\<^sub>0" by (auto elim!: nonpos_Ints_cases) - from z have "g (z/2) * g ((z+1)/2) = - (Gamma (z/2) * Gamma ((z+1)/2)) * (Gamma (1-z/2) * Gamma (1-((z+1)/2))) * - (sin (of_real pi * z/2) * sin (of_real pi * (z+1)/2))" - by (simp add: g_def) - also from z' Gamma_legendre_duplication_aux[of "z/2"] - have "Gamma (z/2) * Gamma ((z+1)/2) = exp ((1-z) * of_real (ln 2)) * Gamma (1/2) * Gamma z" - by (simp add: add_divide_distrib) - also from z'' Gamma_legendre_duplication_aux[of "1-(z+1)/2"] - have "Gamma (1-z/2) * Gamma (1-(z+1)/2) = - Gamma (1-z) * Gamma (1/2) * exp (z * of_real (ln 2))" - by (simp add: add_divide_distrib ac_simps) - finally have "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * (Gamma z * Gamma (1-z) * - (2 * (sin (of_real pi*z/2) * sin (of_real pi*(z+1)/2))))" - by (simp add: add_ac power2_eq_square exp_add ring_distribs exp_diff exp_of_real) - also have "sin (of_real pi*(z+1)/2) = cos (of_real pi*z/2)" - using cos_sin_eq[of "- of_real pi * z/2", symmetric] - by (simp add: ring_distribs add_divide_distrib ac_simps) - also have "2 * (sin (of_real pi*z/2) * cos (of_real pi*z/2)) = sin (of_real pi * z)" - by (subst sin_times_cos) (simp add: field_simps) - also have "Gamma z * Gamma (1 - z) * sin (complex_of_real pi * z) = g z" - using \z \ \\ by (simp add: g_def) - finally show ?thesis . - qed - have g_eq: "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * g z" for z - proof - - define r where "r = \Re z / 2\" - have "Gamma (1/2)^2 * g z = Gamma (1/2)^2 * g (z - of_int (2*r))" by (simp only: g.minus_of_int) - also have "of_int (2*r) = 2 * of_int r" by simp - also have "Re z - 2 * of_int r > -1" "Re z - 2 * of_int r < 2" unfolding r_def by linarith+ - hence "Gamma (1/2)^2 * g (z - 2 * of_int r) = - g ((z - 2 * of_int r)/2) * g ((z - 2 * of_int r + 1)/2)" - unfolding r_def by (intro g_eq[symmetric]) simp_all - also have "(z - 2 * of_int r) / 2 = z/2 - of_int r" by simp - also have "g \ = g (z/2)" by (rule g.minus_of_int) - also have "(z - 2 * of_int r + 1) / 2 = (z + 1)/2 - of_int r" by simp - also have "g \ = g ((z+1)/2)" by (rule g.minus_of_int) - finally show ?thesis .. - qed - - have g_nz [simp]: "g z \ 0" for z :: complex - unfolding g_def using Ints_diff[of 1 "1 - z"] - by (auto simp: Gamma_eq_zero_iff sin_eq_0 dest!: nonpos_Ints_Int) - - have h_eq: "h z = (h (z/2) + h ((z+1)/2)) / 2" for z - proof - - have "((\t. g (t/2) * g ((t+1)/2)) has_field_derivative - (g (z/2) * g ((z+1)/2)) * ((h (z/2) + h ((z+1)/2)) / 2)) (at z)" - by (auto intro!: derivative_eq_intros g_g'[THEN DERIV_chain2] simp: field_simps) - hence "((\t. Gamma (1/2)^2 * g t) has_field_derivative - Gamma (1/2)^2 * g z * ((h (z/2) + h ((z+1)/2)) / 2)) (at z)" - by (subst (1 2) g_eq[symmetric]) simp - from DERIV_cmult[OF this, of "inverse ((Gamma (1/2))^2)"] - have "(g has_field_derivative (g z * ((h (z/2) + h ((z+1)/2))/2))) (at z)" - using fraction_not_in_ints[where 'a = complex, of 2 1] - by (simp add: divide_simps Gamma_eq_zero_iff not_in_Ints_imp_not_in_nonpos_Ints) - moreover have "(g has_field_derivative (g z * h z)) (at z)" - using g_g'[of z] by (simp add: ac_simps) - ultimately have "g z * h z = g z * ((h (z/2) + h ((z+1)/2))/2)" - by (intro DERIV_unique) - thus "h z = (h (z/2) + h ((z+1)/2)) / 2" by simp - qed - - obtain h' where h'_cont: "continuous_on UNIV h'" and - h_h': "\z. (h has_field_derivative h' z) (at z)" - unfolding h_def by (erule Gamma_reflection_aux) - - have h'_eq: "h' z = (h' (z/2) + h' ((z+1)/2)) / 4" for z - proof - - have "((\t. (h (t/2) + h ((t+1)/2)) / 2) has_field_derivative - ((h' (z/2) + h' ((z+1)/2)) / 4)) (at z)" - by (fastforce intro!: derivative_eq_intros h_h'[THEN DERIV_chain2]) - hence "(h has_field_derivative ((h' (z/2) + h' ((z+1)/2))/4)) (at z)" - by (subst (asm) h_eq[symmetric]) - from h_h' and this show "h' z = (h' (z/2) + h' ((z+1)/2)) / 4" by (rule DERIV_unique) - qed - - have h'_zero: "h' z = 0" for z - proof - - define m where "m = max 1 \Re z\" - define B where "B = {t. abs (Re t) \ m \ abs (Im t) \ abs (Im z)}" - have "closed ({t. Re t \ -m} \ {t. Re t \ m} \ - {t. Im t \ -\Im z\} \ {t. Im t \ \Im z\})" - (is "closed ?B") by (intro closed_Int closed_halfspace_Re_ge closed_halfspace_Re_le - closed_halfspace_Im_ge closed_halfspace_Im_le) - also have "?B = B" unfolding B_def by fastforce - finally have "closed B" . - moreover have "bounded B" unfolding bounded_iff - proof (intro ballI exI) - fix t assume t: "t \ B" - have "norm t \ \Re t\ + \Im t\" by (rule cmod_le) - also from t have "\Re t\ \ m" unfolding B_def by blast - also from t have "\Im t\ \ \Im z\" unfolding B_def by blast - finally show "norm t \ m + \Im z\" by - simp - qed - ultimately have compact: "compact B" by (subst compact_eq_bounded_closed) blast - - define M where "M = (SUP z:B. norm (h' z))" - have "compact (h' ` B)" - by (intro compact_continuous_image continuous_on_subset[OF h'_cont] compact) blast+ - hence bdd: "bdd_above ((\z. norm (h' z)) ` B)" - using bdd_above_norm[of "h' ` B"] by (simp add: image_comp o_def compact_imp_bounded) - have "norm (h' z) \ M" unfolding M_def by (intro cSUP_upper bdd) (simp_all add: B_def m_def) - also have "M \ M/2" - proof (subst M_def, subst cSUP_le_iff) - have "z \ B" unfolding B_def m_def by simp - thus "B \ {}" by auto - next - show "\z\B. norm (h' z) \ M/2" - proof - fix t :: complex assume t: "t \ B" - from h'_eq[of t] t have "h' t = (h' (t/2) + h' ((t+1)/2)) / 4" by (simp add: dist_0_norm) - also have "norm \ = norm (h' (t/2) + h' ((t+1)/2)) / 4" by simp - also have "norm (h' (t/2) + h' ((t+1)/2)) \ norm (h' (t/2)) + norm (h' ((t+1)/2))" - by (rule norm_triangle_ineq) - also from t have "abs (Re ((t + 1)/2)) \ m" unfolding m_def B_def by auto - with t have "t/2 \ B" "(t+1)/2 \ B" unfolding B_def by auto - hence "norm (h' (t/2)) + norm (h' ((t+1)/2)) \ M + M" unfolding M_def - by (intro add_mono cSUP_upper bdd) (auto simp: B_def) - also have "(M + M) / 4 = M / 2" by simp - finally show "norm (h' t) \ M/2" by - simp_all - qed - qed (insert bdd, auto simp: cball_eq_empty) - hence "M \ 0" by simp - finally show "h' z = 0" by simp - qed - have h_h'_2: "(h has_field_derivative 0) (at z)" for z - using h_h'[of z] h'_zero[of z] by simp - - have g_real: "g z \ \" if "z \ \" for z - unfolding g_def using that by (auto intro!: Reals_mult Gamma_complex_real) - have h_real: "h z \ \" if "z \ \" for z - unfolding h_def using that by (auto intro!: Reals_mult Reals_add Reals_diff Polygamma_Real) - have g_nz: "g z \ 0" for z unfolding g_def using Ints_diff[of 1 "1-z"] - by (auto simp: Gamma_eq_zero_iff sin_eq_0) - - from h'_zero h_h'_2 have "\c. \z\UNIV. h z = c" - by (intro has_field_derivative_zero_constant) (simp_all add: dist_0_norm) - then obtain c where c: "\z. h z = c" by auto - have "\u. u \ closed_segment 0 1 \ Re (g 1) - Re (g 0) = Re (h u * g u * (1 - 0))" - by (intro complex_mvt_line g_g') - find_theorems name:deriv Reals - then guess u by (elim exE conjE) note u = this - from u(1) have u': "u \ \" unfolding closed_segment_def - by (auto simp: scaleR_conv_of_real) - from u' g_real[of u] g_nz[of u] have "Re (g u) \ 0" by (auto elim!: Reals_cases) - with u(2) c[of u] g_real[of u] g_nz[of u] u' - have "Re c = 0" by (simp add: complex_is_Real_iff g.of_1) - with h_real[of 0] c[of 0] have "c = 0" by (auto elim!: Reals_cases) - with c have A: "h z * g z = 0" for z by simp - hence "(g has_field_derivative 0) (at z)" for z using g_g'[of z] by simp - hence "\c'. \z\UNIV. g z = c'" by (intro has_field_derivative_zero_constant) simp_all - then obtain c' where c: "\z. g z = c'" by (force simp: dist_0_norm) - from this[of 0] have "c' = pi" unfolding g_def by simp - with c have "g z = pi" by simp - - show ?thesis - proof (cases "z \ \") - case False - with \g z = pi\ show ?thesis by (auto simp: g_def divide_simps) - next - case True - then obtain n where n: "z = of_int n" by (elim Ints_cases) - with sin_eq_0[of "of_real pi * z"] have "sin (of_real pi * z) = 0" by force - moreover have "of_int (1 - n) \ \\<^sub>\\<^sub>0" if "n > 0" using that by (intro nonpos_Ints_of_int) simp - ultimately show ?thesis using n - by (cases "n \ 0") (auto simp: Gamma_eq_zero_iff nonpos_Ints_of_int) - qed -qed - -lemma rGamma_reflection_complex: - "rGamma z * rGamma (1 - z :: complex) = sin (of_real pi * z) / of_real pi" - using Gamma_reflection_complex[of z] - by (simp add: Gamma_def divide_simps split: if_split_asm) - -lemma rGamma_reflection_complex': - "rGamma z * rGamma (- z :: complex) = -z * sin (of_real pi * z) / of_real pi" -proof - - have "rGamma z * rGamma (-z) = -z * (rGamma z * rGamma (1 - z))" - using rGamma_plus1[of "-z", symmetric] by simp - also have "rGamma z * rGamma (1 - z) = sin (of_real pi * z) / of_real pi" - by (rule rGamma_reflection_complex) - finally show ?thesis by simp -qed - -lemma Gamma_reflection_complex': - "Gamma z * Gamma (- z :: complex) = - of_real pi / (z * sin (of_real pi * z))" - using rGamma_reflection_complex'[of z] by (force simp add: Gamma_def divide_simps mult_ac) - - - -lemma Gamma_one_half_real: "Gamma (1/2 :: real) = sqrt pi" -proof - - from Gamma_reflection_complex[of "1/2"] fraction_not_in_ints[where 'a = complex, of 2 1] - have "Gamma (1/2 :: complex)^2 = of_real pi" by (simp add: power2_eq_square) - hence "of_real pi = Gamma (complex_of_real (1/2))^2" by simp - also have "\ = of_real ((Gamma (1/2))^2)" by (subst Gamma_complex_of_real) simp_all - finally have "Gamma (1/2)^2 = pi" by (subst (asm) of_real_eq_iff) simp_all - moreover have "Gamma (1/2 :: real) \ 0" using Gamma_real_pos[of "1/2"] by simp - ultimately show ?thesis by (rule real_sqrt_unique [symmetric]) -qed - -lemma Gamma_one_half_complex: "Gamma (1/2 :: complex) = of_real (sqrt pi)" -proof - - have "Gamma (1/2 :: complex) = Gamma (of_real (1/2))" by simp - also have "\ = of_real (sqrt pi)" by (simp only: Gamma_complex_of_real Gamma_one_half_real) - finally show ?thesis . -qed - -lemma Gamma_legendre_duplication: - fixes z :: complex - assumes "z \ \\<^sub>\\<^sub>0" "z + 1/2 \ \\<^sub>\\<^sub>0" - shows "Gamma z * Gamma (z + 1/2) = - exp ((1 - 2*z) * of_real (ln 2)) * of_real (sqrt pi) * Gamma (2*z)" - using Gamma_legendre_duplication_aux[OF assms] by (simp add: Gamma_one_half_complex) - -end - - -subsection \Limits and residues\ - -text \ - The inverse of the Gamma function has simple zeros: -\ - -lemma rGamma_zeros: - "(\z. rGamma z / (z + of_nat n)) \ (- of_nat n) \ ((-1)^n * fact n :: 'a :: Gamma)" -proof (subst tendsto_cong) - let ?f = "\z. pochhammer z n * rGamma (z + of_nat (Suc n)) :: 'a" - from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV] - show "eventually (\z. rGamma z / (z + of_nat n) = ?f z) (at (- of_nat n))" - by (subst pochhammer_rGamma[of _ "Suc n"]) - (auto elim!: eventually_mono simp: divide_simps pochhammer_rec' eq_neg_iff_add_eq_0) - have "isCont ?f (- of_nat n)" by (intro continuous_intros) - thus "?f \ (- of_nat n) \ (- 1) ^ n * fact n" unfolding isCont_def - by (simp add: pochhammer_same) -qed - - -text \ - The simple zeros of the inverse of the Gamma function correspond to simple poles of the Gamma function, - and their residues can easily be computed from the limit we have just proven: -\ - -lemma Gamma_poles: "filterlim Gamma at_infinity (at (- of_nat n :: 'a :: Gamma))" -proof - - from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV] - have "eventually (\z. rGamma z \ (0 :: 'a)) (at (- of_nat n))" - by (auto elim!: eventually_mono nonpos_Ints_cases' - simp: rGamma_eq_zero_iff dist_of_nat dist_minus) - with isCont_rGamma[of "- of_nat n :: 'a", OF continuous_ident] - have "filterlim (\z. inverse (rGamma z) :: 'a) at_infinity (at (- of_nat n))" - unfolding isCont_def by (intro filterlim_compose[OF filterlim_inverse_at_infinity]) - (simp_all add: filterlim_at) - moreover have "(\z. inverse (rGamma z) :: 'a) = Gamma" - by (intro ext) (simp add: rGamma_inverse_Gamma) - ultimately show ?thesis by (simp only: ) -qed - -lemma Gamma_residues: - "(\z. Gamma z * (z + of_nat n)) \ (- of_nat n) \ ((-1)^n / fact n :: 'a :: Gamma)" -proof (subst tendsto_cong) - let ?c = "(- 1) ^ n / fact n :: 'a" - from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV] - show "eventually (\z. Gamma z * (z + of_nat n) = inverse (rGamma z / (z + of_nat n))) - (at (- of_nat n))" - by (auto elim!: eventually_mono simp: divide_simps rGamma_inverse_Gamma) - have "(\z. inverse (rGamma z / (z + of_nat n))) \ (- of_nat n) \ - inverse ((- 1) ^ n * fact n :: 'a)" - by (intro tendsto_intros rGamma_zeros) simp_all - also have "inverse ((- 1) ^ n * fact n) = ?c" - by (simp_all add: field_simps power_mult_distrib [symmetric] del: power_mult_distrib) - finally show "(\z. inverse (rGamma z / (z + of_nat n))) \ (- of_nat n) \ ?c" . -qed - - - -subsection \Alternative definitions\ - - -subsubsection \Variant of the Euler form\ - - -definition Gamma_series_euler' where - "Gamma_series_euler' z n = - inverse z * (\k=1..n. exp (z * of_real (ln (1 + inverse (of_nat k)))) / (1 + z / of_nat k))" - -context -begin -private lemma Gamma_euler'_aux1: - fixes z :: "'a :: {real_normed_field,banach}" - assumes n: "n > 0" - shows "exp (z * of_real (ln (of_nat n + 1))) = (\k=1..n. exp (z * of_real (ln (1 + 1 / of_nat k))))" -proof - - have "(\k=1..n. exp (z * of_real (ln (1 + 1 / of_nat k)))) = - exp (z * of_real (\k = 1..n. ln (1 + 1 / real_of_nat k)))" - by (subst exp_setsum [symmetric]) (simp_all add: setsum_right_distrib) - also have "(\k=1..n. ln (1 + 1 / of_nat k) :: real) = ln (\k=1..n. 1 + 1 / real_of_nat k)" - by (subst ln_setprod [symmetric]) (auto intro!: add_pos_nonneg) - also have "(\k=1..n. 1 + 1 / of_nat k :: real) = (\k=1..n. (of_nat k + 1) / of_nat k)" - by (intro setprod.cong) (simp_all add: divide_simps) - also have "(\k=1..n. (of_nat k + 1) / of_nat k :: real) = of_nat n + 1" - by (induction n) (simp_all add: setprod_nat_ivl_Suc' divide_simps) - finally show ?thesis .. -qed - -lemma Gamma_series_euler': - assumes z: "(z :: 'a :: Gamma) \ \\<^sub>\\<^sub>0" - shows "(\n. Gamma_series_euler' z n) \ Gamma z" -proof (rule Gamma_seriesI, rule Lim_transform_eventually) - let ?f = "\n. fact n * exp (z * of_real (ln (of_nat n + 1))) / pochhammer z (n + 1)" - let ?r = "\n. ?f n / Gamma_series z n" - let ?r' = "\n. exp (z * of_real (ln (of_nat (Suc n) / of_nat n)))" - from z have z': "z \ 0" by auto - - have "eventually (\n. ?r' n = ?r n) sequentially" using eventually_gt_at_top[of "0::nat"] - using z by (auto simp: divide_simps Gamma_series_def ring_distribs exp_diff ln_div add_ac - elim!: eventually_mono dest: pochhammer_eq_0_imp_nonpos_Int) - moreover have "?r' \ exp (z * of_real (ln 1))" - by (intro tendsto_intros LIMSEQ_Suc_n_over_n) simp_all - ultimately show "?r \ 1" by (force dest!: Lim_transform_eventually) - - from eventually_gt_at_top[of "0::nat"] - show "eventually (\n. ?r n = Gamma_series_euler' z n / Gamma_series z n) sequentially" - proof eventually_elim - fix n :: nat assume n: "n > 0" - from n z' have "Gamma_series_euler' z n = - exp (z * of_real (ln (of_nat n + 1))) / (z * (\k=1..n. (1 + z / of_nat k)))" - by (subst Gamma_euler'_aux1) - (simp_all add: Gamma_series_euler'_def setprod.distrib - setprod_inversef[symmetric] divide_inverse) - also have "(\k=1..n. (1 + z / of_nat k)) = pochhammer (z + 1) n / fact n" - by (cases n) (simp_all add: pochhammer_setprod fact_setprod atLeastLessThanSuc_atLeastAtMost - setprod_dividef [symmetric] field_simps setprod.atLeast_Suc_atMost_Suc_shift) - also have "z * \ = pochhammer z (Suc n) / fact n" by (simp add: pochhammer_rec) - finally show "?r n = Gamma_series_euler' z n / Gamma_series z n" by simp - qed -qed - -end - - - -subsubsection \Weierstrass form\ - -definition Gamma_series_weierstrass :: "'a :: {banach,real_normed_field} \ nat \ 'a" where - "Gamma_series_weierstrass z n = - exp (-euler_mascheroni * z) / z * (\k=1..n. exp (z / of_nat k) / (1 + z / of_nat k))" - -definition rGamma_series_weierstrass :: "'a :: {banach,real_normed_field} \ nat \ 'a" where - "rGamma_series_weierstrass z n = - exp (euler_mascheroni * z) * z * (\k=1..n. (1 + z / of_nat k) * exp (-z / of_nat k))" - -lemma Gamma_series_weierstrass_nonpos_Ints: - "eventually (\k. Gamma_series_weierstrass (- of_nat n) k = 0) sequentially" - using eventually_ge_at_top[of n] by eventually_elim (auto simp: Gamma_series_weierstrass_def) - -lemma rGamma_series_weierstrass_nonpos_Ints: - "eventually (\k. rGamma_series_weierstrass (- of_nat n) k = 0) sequentially" - using eventually_ge_at_top[of n] by eventually_elim (auto simp: rGamma_series_weierstrass_def) - -lemma Gamma_weierstrass_complex: "Gamma_series_weierstrass z \ Gamma (z :: complex)" -proof (cases "z \ \\<^sub>\\<^sub>0") - case True - then obtain n where "z = - of_nat n" by (elim nonpos_Ints_cases') - also from True have "Gamma_series_weierstrass \ \ Gamma z" - by (simp add: tendsto_cong[OF Gamma_series_weierstrass_nonpos_Ints] Gamma_nonpos_Int) - finally show ?thesis . -next - case False - hence z: "z \ 0" by auto - let ?f = "(\x. \x = Suc 0..x. exp (z / of_nat x) / (1 + z / of_nat x))" - have A: "exp (ln (1 + z / of_nat n)) = (1 + z / of_nat n)" if "n \ 1" for n :: nat - using False that by (subst exp_Ln) (auto simp: field_simps dest!: plus_of_nat_eq_0_imp) - have "(\n. \k=1..n. z / of_nat k - ln (1 + z / of_nat k)) \ ln_Gamma z + euler_mascheroni * z + ln z" - using ln_Gamma_series'_aux[OF False] - by (simp only: atLeastLessThanSuc_atLeastAtMost [symmetric] One_nat_def - setsum_shift_bounds_Suc_ivl sums_def atLeast0LessThan) - from tendsto_exp[OF this] False z have "?f \ z * exp (euler_mascheroni * z) * Gamma z" - by (simp add: exp_add exp_setsum exp_diff mult_ac Gamma_complex_altdef A) - from tendsto_mult[OF tendsto_const[of "exp (-euler_mascheroni * z) / z"] this] z - show "Gamma_series_weierstrass z \ Gamma z" - by (simp add: exp_minus divide_simps Gamma_series_weierstrass_def [abs_def]) -qed - -lemma tendsto_complex_of_real_iff: "((\x. complex_of_real (f x)) \ of_real c) F = (f \ c) F" - by (rule tendsto_of_real_iff) - -lemma Gamma_weierstrass_real: "Gamma_series_weierstrass x \ Gamma (x :: real)" - using Gamma_weierstrass_complex[of "of_real x"] unfolding Gamma_series_weierstrass_def[abs_def] - by (subst tendsto_complex_of_real_iff [symmetric]) - (simp_all add: exp_of_real[symmetric] Gamma_complex_of_real) - -lemma rGamma_weierstrass_complex: "rGamma_series_weierstrass z \ rGamma (z :: complex)" -proof (cases "z \ \\<^sub>\\<^sub>0") - case True - then obtain n where "z = - of_nat n" by (elim nonpos_Ints_cases') - also from True have "rGamma_series_weierstrass \ \ rGamma z" - by (simp add: tendsto_cong[OF rGamma_series_weierstrass_nonpos_Ints] rGamma_nonpos_Int) - finally show ?thesis . -next - case False - have "rGamma_series_weierstrass z = (\n. inverse (Gamma_series_weierstrass z n))" - by (simp add: rGamma_series_weierstrass_def[abs_def] Gamma_series_weierstrass_def - exp_minus divide_inverse setprod_inversef[symmetric] mult_ac) - also from False have "\ \ inverse (Gamma z)" - by (intro tendsto_intros Gamma_weierstrass_complex) (simp add: Gamma_eq_zero_iff) - finally show ?thesis by (simp add: Gamma_def) -qed - -subsubsection \Binomial coefficient form\ - -lemma Gamma_gbinomial: - "(\n. ((z + of_nat n) gchoose n) * exp (-z * of_real (ln (of_nat n)))) \ rGamma (z+1)" -proof (cases "z = 0") - case False - show ?thesis - proof (rule Lim_transform_eventually) - let ?powr = "\a b. exp (b * of_real (ln (of_nat a)))" - show "eventually (\n. rGamma_series z n / z = - ((z + of_nat n) gchoose n) * ?powr n (-z)) sequentially" - proof (intro always_eventually allI) - fix n :: nat - from False have "((z + of_nat n) gchoose n) = pochhammer z (Suc n) / z / fact n" - by (simp add: gbinomial_pochhammer' pochhammer_rec) - also have "pochhammer z (Suc n) / z / fact n * ?powr n (-z) = rGamma_series z n / z" - by (simp add: rGamma_series_def divide_simps exp_minus) - finally show "rGamma_series z n / z = ((z + of_nat n) gchoose n) * ?powr n (-z)" .. - qed - - from False have "(\n. rGamma_series z n / z) \ rGamma z / z" by (intro tendsto_intros) - also from False have "rGamma z / z = rGamma (z + 1)" using rGamma_plus1[of z] - by (simp add: field_simps) - finally show "(\n. rGamma_series z n / z) \ rGamma (z+1)" . - qed -qed (simp_all add: binomial_gbinomial [symmetric]) - -lemma gbinomial_minus': "(a + of_nat b) gchoose b = (- 1) ^ b * (- (a + 1) gchoose b)" - by (subst gbinomial_minus) (simp add: power_mult_distrib [symmetric]) - -lemma gbinomial_asymptotic: - fixes z :: "'a :: Gamma" - shows "(\n. (z gchoose n) / ((-1)^n / exp ((z+1) * of_real (ln (real n))))) \ - inverse (Gamma (- z))" - unfolding rGamma_inverse_Gamma [symmetric] using Gamma_gbinomial[of "-z-1"] - by (subst (asm) gbinomial_minus') - (simp add: add_ac mult_ac divide_inverse power_inverse [symmetric]) - -lemma fact_binomial_limit: - "(\n. of_nat ((k + n) choose n) / of_nat (n ^ k) :: 'a :: Gamma) \ 1 / fact k" -proof (rule Lim_transform_eventually) - have "(\n. of_nat ((k + n) choose n) / of_real (exp (of_nat k * ln (real_of_nat n)))) - \ 1 / Gamma (of_nat (Suc k) :: 'a)" (is "?f \ _") - using Gamma_gbinomial[of "of_nat k :: 'a"] - by (simp add: binomial_gbinomial add_ac Gamma_def divide_simps exp_of_real [symmetric] exp_minus) - also have "Gamma (of_nat (Suc k)) = fact k" by (simp add: Gamma_fact) - finally show "?f \ 1 / fact k" . - - show "eventually (\n. ?f n = of_nat ((k + n) choose n) / of_nat (n ^ k)) sequentially" - using eventually_gt_at_top[of "0::nat"] - proof eventually_elim - fix n :: nat assume n: "n > 0" - from n have "exp (real_of_nat k * ln (real_of_nat n)) = real_of_nat (n^k)" - by (simp add: exp_of_nat_mult) - thus "?f n = of_nat ((k + n) choose n) / of_nat (n ^ k)" by simp - qed -qed - -lemma binomial_asymptotic': - "(\n. of_nat ((k + n) choose n) / (of_nat (n ^ k) / fact k) :: 'a :: Gamma) \ 1" - using tendsto_mult[OF fact_binomial_limit[of k] tendsto_const[of "fact k :: 'a"]] by simp - -lemma gbinomial_Beta: - assumes "z + 1 \ \\<^sub>\\<^sub>0" - shows "((z::'a::Gamma) gchoose n) = inverse ((z + 1) * Beta (z - of_nat n + 1) (of_nat n + 1))" -using assms -proof (induction n arbitrary: z) - case 0 - hence "z + 2 \ \\<^sub>\\<^sub>0" - using plus_one_in_nonpos_Ints_imp[of "z+1"] by (auto simp: add.commute) - with 0 show ?case - by (auto simp: Beta_def Gamma_eq_zero_iff Gamma_plus1 [symmetric] add.commute) -next - case (Suc n z) - show ?case - proof (cases "z \ \\<^sub>\\<^sub>0") - case True - with Suc.prems have "z = 0" - by (auto elim!: nonpos_Ints_cases simp: algebra_simps one_plus_of_int_in_nonpos_Ints_iff) - show ?thesis - proof (cases "n = 0") - case True - with \z = 0\ show ?thesis - by (simp add: Beta_def Gamma_eq_zero_iff Gamma_plus1 [symmetric]) - next - case False - with \z = 0\ show ?thesis - by (simp_all add: Beta_pole1 one_minus_of_nat_in_nonpos_Ints_iff gbinomial_1) - qed - next - case False - have "(z gchoose (Suc n)) = ((z - 1 + 1) gchoose (Suc n))" by simp - also have "\ = (z - 1 gchoose n) * ((z - 1) + 1) / of_nat (Suc n)" - by (subst gbinomial_factors) (simp add: field_simps) - also from False have "\ = inverse (of_nat (Suc n) * Beta (z - of_nat n) (of_nat (Suc n)))" - (is "_ = inverse ?x") by (subst Suc.IH) (simp_all add: field_simps Beta_pole1) - also have "of_nat (Suc n) \ (\\<^sub>\\<^sub>0 :: 'a set)" by (subst of_nat_in_nonpos_Ints_iff) simp_all - hence "?x = (z + 1) * Beta (z - of_nat (Suc n) + 1) (of_nat (Suc n) + 1)" - by (subst Beta_plus1_right [symmetric]) simp_all - finally show ?thesis . - qed -qed - -lemma gbinomial_Gamma: - assumes "z + 1 \ \\<^sub>\\<^sub>0" - shows "(z gchoose n) = Gamma (z + 1) / (fact n * Gamma (z - of_nat n + 1))" -proof - - have "(z gchoose n) = Gamma (z + 2) / (z + 1) / (fact n * Gamma (z - of_nat n + 1))" - by (subst gbinomial_Beta[OF assms]) (simp_all add: Beta_def Gamma_fact [symmetric] add_ac) - also from assms have "Gamma (z + 2) / (z + 1) = Gamma (z + 1)" - using Gamma_plus1[of "z+1"] by (auto simp add: divide_simps mult_ac add_ac) - finally show ?thesis . -qed - - -subsubsection \Integral form\ - -lemma integrable_Gamma_integral_bound: - fixes a c :: real - assumes a: "a > -1" and c: "c \ 0" - defines "f \ \x. if x \ {0..c} then x powr a else exp (-x/2)" - shows "f integrable_on {0..}" -proof - - have "f integrable_on {0..c}" - by (rule integrable_spike_finite[of "{}", OF _ _ integrable_on_powr_from_0[of a c]]) - (insert a c, simp_all add: f_def) - moreover have A: "(\x. exp (-x/2)) integrable_on {c..}" - using integrable_on_exp_minus_to_infinity[of "1/2"] by simp - have "f integrable_on {c..}" - by (rule integrable_spike_finite[of "{c}", OF _ _ A]) (simp_all add: f_def) - ultimately show "f integrable_on {0..}" - by (rule integrable_union') (insert c, auto simp: max_def) -qed - -lemma Gamma_integral_complex: - assumes z: "Re z > 0" - shows "((\t. of_real t powr (z - 1) / of_real (exp t)) has_integral Gamma z) {0..}" -proof - - have A: "((\t. (of_real t) powr (z - 1) * of_real ((1 - t) ^ n)) - has_integral (fact n / pochhammer z (n+1))) {0..1}" - if "Re z > 0" for n z using that - proof (induction n arbitrary: z) - case 0 - have "((\t. complex_of_real t powr (z - 1)) has_integral - (of_real 1 powr z / z - of_real 0 powr z / z)) {0..1}" using 0 - by (intro fundamental_theorem_of_calculus_interior) - (auto intro!: continuous_intros derivative_eq_intros has_vector_derivative_real_complex) - thus ?case by simp - next - case (Suc n) - let ?f = "\t. complex_of_real t powr z / z" - let ?f' = "\t. complex_of_real t powr (z - 1)" - let ?g = "\t. (1 - complex_of_real t) ^ Suc n" - let ?g' = "\t. - ((1 - complex_of_real t) ^ n) * of_nat (Suc n)" - have "((\t. ?f' t * ?g t) has_integral - (of_nat (Suc n)) * fact n / pochhammer z (n+2)) {0..1}" - (is "(_ has_integral ?I) _") - proof (rule integration_by_parts_interior[where f' = ?f' and g = ?g]) - from Suc.prems show "continuous_on {0..1} ?f" "continuous_on {0..1} ?g" - by (auto intro!: continuous_intros) - next - fix t :: real assume t: "t \ {0<..<1}" - show "(?f has_vector_derivative ?f' t) (at t)" using t Suc.prems - by (auto intro!: derivative_eq_intros has_vector_derivative_real_complex) - show "(?g has_vector_derivative ?g' t) (at t)" - by (rule has_vector_derivative_real_complex derivative_eq_intros refl)+ simp_all - next - from Suc.prems have [simp]: "z \ 0" by auto - from Suc.prems have A: "Re (z + of_nat n) > 0" for n by simp - have [simp]: "z + of_nat n \ 0" "z + 1 + of_nat n \ 0" for n - using A[of n] A[of "Suc n"] by (auto simp add: add.assoc simp del: plus_complex.sel) - have "((\x. of_real x powr z * of_real ((1 - x) ^ n) * (- of_nat (Suc n) / z)) has_integral - fact n / pochhammer (z+1) (n+1) * (- of_nat (Suc n) / z)) {0..1}" - (is "(?A has_integral ?B) _") - using Suc.IH[of "z+1"] Suc.prems by (intro has_integral_mult_left) (simp_all add: add_ac pochhammer_rec) - also have "?A = (\t. ?f t * ?g' t)" by (intro ext) (simp_all add: field_simps) - also have "?B = - (of_nat (Suc n) * fact n / pochhammer z (n+2))" - by (simp add: divide_simps pochhammer_rec - setprod_shift_bounds_cl_Suc_ivl del: of_nat_Suc) - finally show "((\t. ?f t * ?g' t) has_integral (?f 1 * ?g 1 - ?f 0 * ?g 0 - ?I)) {0..1}" - by simp - qed (simp_all add: bounded_bilinear_mult) - thus ?case by simp - qed - - have B: "((\t. if t \ {0..of_nat n} then - of_real t powr (z - 1) * (1 - of_real t / of_nat n) ^ n else 0) - has_integral (of_nat n powr z * fact n / pochhammer z (n+1))) {0..}" for n - proof (cases "n > 0") - case [simp]: True - hence [simp]: "n \ 0" by auto - with has_integral_affinity01[OF A[OF z, of n], of "inverse (of_nat n)" 0] - have "((\x. (of_nat n - of_real x) ^ n * (of_real x / of_nat n) powr (z - 1) / of_nat n ^ n) - has_integral fact n * of_nat n / pochhammer z (n+1)) ((\x. real n * x)`{0..1})" - (is "(?f has_integral ?I) ?ivl") by (simp add: field_simps scaleR_conv_of_real) - also from True have "((\x. real n*x)`{0..1}) = {0..real n}" - by (subst image_mult_atLeastAtMost) simp_all - also have "?f = (\x. (of_real x / of_nat n) powr (z - 1) * (1 - of_real x / of_nat n) ^ n)" - using True by (intro ext) (simp add: field_simps) - finally have "((\x. (of_real x / of_nat n) powr (z - 1) * (1 - of_real x / of_nat n) ^ n) - has_integral ?I) {0..real n}" (is ?P) . - also have "?P \ ((\x. exp ((z - 1) * of_real (ln (x / of_nat n))) * (1 - of_real x / of_nat n) ^ n) - has_integral ?I) {0..real n}" - by (intro has_integral_spike_finite_eq[of "{0}"]) (auto simp: powr_def Ln_of_real [symmetric]) - also have "\ \ ((\x. exp ((z - 1) * of_real (ln x - ln (of_nat n))) * (1 - of_real x / of_nat n) ^ n) - has_integral ?I) {0..real n}" - by (intro has_integral_spike_finite_eq[of "{0}"]) (simp_all add: ln_div) - finally have \ . - note B = has_integral_mult_right[OF this, of "exp ((z - 1) * ln (of_nat n))"] - have "((\x. exp ((z - 1) * of_real (ln x)) * (1 - of_real x / of_nat n) ^ n) - has_integral (?I * exp ((z - 1) * ln (of_nat n)))) {0..real n}" (is ?P) - by (insert B, subst (asm) mult.assoc [symmetric], subst (asm) exp_add [symmetric]) - (simp add: Ln_of_nat algebra_simps) - also have "?P \ ((\x. of_real x powr (z - 1) * (1 - of_real x / of_nat n) ^ n) - has_integral (?I * exp ((z - 1) * ln (of_nat n)))) {0..real n}" - by (intro has_integral_spike_finite_eq[of "{0}"]) (simp_all add: powr_def Ln_of_real) - also have "fact n * of_nat n / pochhammer z (n+1) * exp ((z - 1) * Ln (of_nat n)) = - (of_nat n powr z * fact n / pochhammer z (n+1))" - by (auto simp add: powr_def algebra_simps exp_diff) - finally show ?thesis by (subst has_integral_restrict) simp_all - next - case False - thus ?thesis by (subst has_integral_restrict) (simp_all add: has_integral_refl) - qed - - have "eventually (\n. Gamma_series z n = - of_nat n powr z * fact n / pochhammer z (n+1)) sequentially" - using eventually_gt_at_top[of "0::nat"] - by eventually_elim (simp add: powr_def algebra_simps Ln_of_nat Gamma_series_def) - from this and Gamma_series_LIMSEQ[of z] - have C: "(\k. of_nat k powr z * fact k / pochhammer z (k+1)) \ Gamma z" - by (rule Lim_transform_eventually) - - { - fix x :: real assume x: "x \ 0" - have lim_exp: "(\k. (1 - x / real k) ^ k) \ exp (-x)" - using tendsto_exp_limit_sequentially[of "-x"] by simp - have "(\k. of_real x powr (z - 1) * of_real ((1 - x / of_nat k) ^ k)) - \ of_real x powr (z - 1) * of_real (exp (-x))" (is ?P) - by (intro tendsto_intros lim_exp) - also from eventually_gt_at_top[of "nat \x\"] - have "eventually (\k. of_nat k > x) sequentially" by eventually_elim linarith - hence "?P \ (\k. if x \ of_nat k then - of_real x powr (z - 1) * of_real ((1 - x / of_nat k) ^ k) else 0) - \ of_real x powr (z - 1) * of_real (exp (-x))" - by (intro tendsto_cong) (auto elim!: eventually_mono) - finally have \ . - } - hence D: "\x\{0..}. (\k. if x \ {0..real k} then - of_real x powr (z - 1) * (1 - of_real x / of_nat k) ^ k else 0) - \ of_real x powr (z - 1) / of_real (exp x)" - by (simp add: exp_minus field_simps cong: if_cong) - - have "((\x. (Re z - 1) * (ln x / x)) \ (Re z - 1) * 0) at_top" - by (intro tendsto_intros ln_x_over_x_tendsto_0) - hence "((\x. ((Re z - 1) * ln x) / x) \ 0) at_top" by simp - from order_tendstoD(2)[OF this, of "1/2"] - have "eventually (\x. (Re z - 1) * ln x / x < 1/2) at_top" by simp - from eventually_conj[OF this eventually_gt_at_top[of 0]] - obtain x0 where "\x\x0. (Re z - 1) * ln x / x < 1/2 \ x > 0" - by (auto simp: eventually_at_top_linorder) - hence x0: "x0 > 0" "\x. x \ x0 \ (Re z - 1) * ln x < x / 2" by auto - - define h where "h = (\x. if x \ {0..x0} then x powr (Re z - 1) else exp (-x/2))" - have le_h: "x powr (Re z - 1) * exp (-x) \ h x" if x: "x \ 0" for x - proof (cases "x > x0") - case True - from True x0(1) have "x powr (Re z - 1) * exp (-x) = exp ((Re z - 1) * ln x - x)" - by (simp add: powr_def exp_diff exp_minus field_simps exp_add) - also from x0(2)[of x] True have "\ < exp (-x/2)" - by (simp add: field_simps) - finally show ?thesis using True by (auto simp add: h_def) - next - case False - from x have "x powr (Re z - 1) * exp (- x) \ x powr (Re z - 1) * 1" - by (intro mult_left_mono) simp_all - with False show ?thesis by (auto simp add: h_def) - qed - - have E: "\x\{0..}. cmod (if x \ {0..real k} then of_real x powr (z - 1) * - (1 - complex_of_real x / of_nat k) ^ k else 0) \ h x" - (is "\x\_. ?f x \ _") for k - proof safe - fix x :: real assume x: "x \ 0" - { - fix x :: real and n :: nat assume x: "x \ of_nat n" - have "(1 - complex_of_real x / of_nat n) = complex_of_real ((1 - x / of_nat n))" by simp - also have "norm \ = \(1 - x / real n)\" by (subst norm_of_real) (rule refl) - also from x have "\ = (1 - x / real n)" by (intro abs_of_nonneg) (simp_all add: divide_simps) - finally have "cmod (1 - complex_of_real x / of_nat n) = 1 - x / real n" . - } note D = this - from D[of x k] x - have "?f x \ (if of_nat k \ x \ k > 0 then x powr (Re z - 1) * (1 - x / real k) ^ k else 0)" - by (auto simp: norm_mult norm_powr_real_powr norm_power intro!: mult_nonneg_nonneg) - also have "\ \ x powr (Re z - 1) * exp (-x)" - by (auto intro!: mult_left_mono exp_ge_one_minus_x_over_n_power_n) - also from x have "\ \ h x" by (rule le_h) - finally show "?f x \ h x" . - qed - - have F: "h integrable_on {0..}" unfolding h_def - by (rule integrable_Gamma_integral_bound) (insert assms x0(1), simp_all) - show ?thesis - by (rule has_integral_dominated_convergence[OF B F E D C]) -qed - -lemma Gamma_integral_real: - assumes x: "x > (0 :: real)" - shows "((\t. t powr (x - 1) / exp t) has_integral Gamma x) {0..}" -proof - - have A: "((\t. complex_of_real t powr (complex_of_real x - 1) / - complex_of_real (exp t)) has_integral complex_of_real (Gamma x)) {0..}" - using Gamma_integral_complex[of x] assms by (simp_all add: Gamma_complex_of_real powr_of_real) - have "((\t. complex_of_real (t powr (x - 1) / exp t)) has_integral of_real (Gamma x)) {0..}" - by (rule has_integral_eq[OF _ A]) (simp_all add: powr_of_real [symmetric]) - from has_integral_linear[OF this bounded_linear_Re] show ?thesis by (simp add: o_def) -qed - - - -subsection \The Weierstraß product formula for the sine\ - -lemma sin_product_formula_complex: - fixes z :: complex - shows "(\n. of_real pi * z * (\k=1..n. 1 - z^2 / of_nat k^2)) \ sin (of_real pi * z)" -proof - - let ?f = "rGamma_series_weierstrass" - have "(\n. (- of_real pi * inverse z) * (?f z n * ?f (- z) n)) - \ (- of_real pi * inverse z) * (rGamma z * rGamma (- z))" - by (intro tendsto_intros rGamma_weierstrass_complex) - also have "(\n. (- of_real pi * inverse z) * (?f z n * ?f (-z) n)) = - (\n. of_real pi * z * (\k=1..n. 1 - z^2 / of_nat k ^ 2))" - proof - fix n :: nat - have "(- of_real pi * inverse z) * (?f z n * ?f (-z) n) = - of_real pi * z * (\k=1..n. (of_nat k - z) * (of_nat k + z) / of_nat k ^ 2)" - by (simp add: rGamma_series_weierstrass_def mult_ac exp_minus - divide_simps setprod.distrib[symmetric] power2_eq_square) - also have "(\k=1..n. (of_nat k - z) * (of_nat k + z) / of_nat k ^ 2) = - (\k=1..n. 1 - z^2 / of_nat k ^ 2)" - by (intro setprod.cong) (simp_all add: power2_eq_square field_simps) - finally show "(- of_real pi * inverse z) * (?f z n * ?f (-z) n) = of_real pi * z * \" - by (simp add: divide_simps) - qed - also have "(- of_real pi * inverse z) * (rGamma z * rGamma (- z)) = sin (of_real pi * z)" - by (subst rGamma_reflection_complex') (simp add: divide_simps) - finally show ?thesis . -qed - -lemma sin_product_formula_real: - "(\n. pi * (x::real) * (\k=1..n. 1 - x^2 / of_nat k^2)) \ sin (pi * x)" -proof - - from sin_product_formula_complex[of "of_real x"] - have "(\n. of_real pi * of_real x * (\k=1..n. 1 - (of_real x)^2 / (of_nat k)^2)) - \ sin (of_real pi * of_real x :: complex)" (is "?f \ ?y") . - also have "?f = (\n. of_real (pi * x * (\k=1..n. 1 - x^2 / (of_nat k^2))))" by simp - also have "?y = of_real (sin (pi * x))" by (simp only: sin_of_real [symmetric] of_real_mult) - finally show ?thesis by (subst (asm) tendsto_of_real_iff) -qed - -lemma sin_product_formula_real': - assumes "x \ (0::real)" - shows "(\n. (\k=1..n. 1 - x^2 / of_nat k^2)) \ sin (pi * x) / (pi * x)" - using tendsto_divide[OF sin_product_formula_real[of x] tendsto_const[of "pi * x"]] assms - by simp - - -subsection \The Solution to the Basel problem\ - -theorem inverse_squares_sums: "(\n. 1 / (n + 1)\<^sup>2) sums (pi\<^sup>2 / 6)" -proof - - define P where "P x n = (\k=1..n. 1 - x^2 / of_nat k^2)" for x :: real and n - define K where "K = (\n. inverse (real_of_nat (Suc n))^2)" - define f where [abs_def]: "f x = (\n. P x n / of_nat (Suc n)^2)" for x - define g where [abs_def]: "g x = (1 - sin (pi * x) / (pi * x))" for x - - have sums: "(\n. P x n / of_nat (Suc n)^2) sums (if x = 0 then K else g x / x^2)" for x - proof (cases "x = 0") - assume x: "x = 0" - have "summable (\n. inverse ((real_of_nat (Suc n))\<^sup>2))" - using inverse_power_summable[of 2] by (subst summable_Suc_iff) simp - thus ?thesis by (simp add: x g_def P_def K_def inverse_eq_divide power_divide summable_sums) - next - assume x: "x \ 0" - have "(\n. P x n - P x (Suc n)) sums (P x 0 - sin (pi * x) / (pi * x))" - unfolding P_def using x by (intro telescope_sums' sin_product_formula_real') - also have "(\n. P x n - P x (Suc n)) = (\n. (x^2 / of_nat (Suc n)^2) * P x n)" - unfolding P_def by (simp add: setprod_nat_ivl_Suc' algebra_simps) - also have "P x 0 = 1" by (simp add: P_def) - finally have "(\n. x\<^sup>2 / (of_nat (Suc n))\<^sup>2 * P x n) sums (1 - sin (pi * x) / (pi * x))" . - from sums_divide[OF this, of "x^2"] x show ?thesis unfolding g_def by simp - qed - - have "continuous_on (ball 0 1) f" - proof (rule uniform_limit_theorem; (intro always_eventually allI)?) - show "uniform_limit (ball 0 1) (\n x. \k ball 0 1" - { - fix k :: nat assume k: "k \ 1" - from x have "x^2 < 1" by (auto simp: dist_0_norm abs_square_less_1) - also from k have "\ \ of_nat k^2" by simp - finally have "(1 - x^2 / of_nat k^2) \ {0..1}" using k - by (simp_all add: field_simps del: of_nat_Suc) - } - hence "(\k=1..n. abs (1 - x^2 / of_nat k^2)) \ (\k=1..n. 1)" by (intro setprod_mono) simp - thus "norm (P x n / (of_nat (Suc n)^2)) \ 1 / of_nat (Suc n)^2" - unfolding P_def by (simp add: field_simps abs_setprod del: of_nat_Suc) - qed (subst summable_Suc_iff, insert inverse_power_summable[of 2], simp add: inverse_eq_divide) - qed (auto simp: P_def intro!: continuous_intros) - hence "isCont f 0" by (subst (asm) continuous_on_eq_continuous_at) simp_all - hence "(f \ 0 \ f 0)" by (simp add: isCont_def) - also have "f 0 = K" unfolding f_def P_def K_def by (simp add: inverse_eq_divide power_divide) - finally have "f \ 0 \ K" . - - moreover have "f \ 0 \ pi^2 / 6" - proof (rule Lim_transform_eventually) - define f' where [abs_def]: "f' x = (\n. - sin_coeff (n+3) * pi ^ (n+2) * x^n)" for x - have "eventually (\x. x \ (0::real)) (at 0)" - by (auto simp add: eventually_at intro!: exI[of _ 1]) - thus "eventually (\x. f' x = f x) (at 0)" - proof eventually_elim - fix x :: real assume x: "x \ 0" - have "sin_coeff 1 = (1 :: real)" "sin_coeff 2 = (0::real)" by (simp_all add: sin_coeff_def) - with sums_split_initial_segment[OF sums_minus[OF sin_converges], of 3 "pi*x"] - have "(\n. - (sin_coeff (n+3) * (pi*x)^(n+3))) sums (pi * x - sin (pi*x))" - by (simp add: eval_nat_numeral) - from sums_divide[OF this, of "x^3 * pi"] x - have "(\n. - (sin_coeff (n+3) * pi^(n+2) * x^n)) sums ((1 - sin (pi*x) / (pi*x)) / x^2)" - by (simp add: divide_simps eval_nat_numeral power_mult_distrib mult_ac) - with x have "(\n. - (sin_coeff (n+3) * pi^(n+2) * x^n)) sums (g x / x^2)" - by (simp add: g_def) - hence "f' x = g x / x^2" by (simp add: sums_iff f'_def) - also have "\ = f x" using sums[of x] x by (simp add: sums_iff g_def f_def) - finally show "f' x = f x" . - qed - - have "isCont f' 0" unfolding f'_def - proof (intro isCont_powser_converges_everywhere) - fix x :: real show "summable (\n. -sin_coeff (n+3) * pi^(n+2) * x^n)" - proof (cases "x = 0") - assume x: "x \ 0" - from summable_divide[OF sums_summable[OF sums_split_initial_segment[OF - sin_converges[of "pi*x"]], of 3], of "-pi*x^3"] x - show ?thesis by (simp add: mult_ac power_mult_distrib divide_simps eval_nat_numeral) - qed (simp only: summable_0_powser) - qed - hence "f' \ 0 \ f' 0" by (simp add: isCont_def) - also have "f' 0 = pi * pi / fact 3" unfolding f'_def - by (subst powser_zero) (simp add: sin_coeff_def) - finally show "f' \ 0 \ pi^2 / 6" by (simp add: eval_nat_numeral) - qed - - ultimately have "K = pi^2 / 6" by (rule LIM_unique) - moreover from inverse_power_summable[of 2] - have "summable (\n. (inverse (real_of_nat (Suc n)))\<^sup>2)" - by (subst summable_Suc_iff) (simp add: power_inverse) - ultimately show ?thesis unfolding K_def - by (auto simp add: sums_iff power_divide inverse_eq_divide) -qed - -end diff -r bbcb05504fdc -r bd218a9320b5 src/HOL/Multivariate_Analysis/Gamma_Function.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Multivariate_Analysis/Gamma_Function.thy Thu Aug 04 19:36:31 2016 +0200 @@ -0,0 +1,2969 @@ +(* Title: HOL/Multivariate_Analysis/Gamma.thy + Author: Manuel Eberl, TU München +*) + +section \The Gamma Function\ + +theory Gamma_Function +imports + Complex_Transcendental + Summation_Tests + Harmonic_Numbers + "~~/src/HOL/Library/Nonpos_Ints" + "~~/src/HOL/Library/Periodic_Fun" +begin + +text \ + Several equivalent definitions of the Gamma function and its + most important properties. Also contains the definition and some properties + of the log-Gamma function and the Digamma function and the other Polygamma functions. + + Based on the Gamma function, we also prove the Weierstraß product form of the + sin function and, based on this, the solution of the Basel problem (the + sum over all @{term "1 / (n::nat)^2"}. +\ + +lemma pochhammer_eq_0_imp_nonpos_Int: + "pochhammer (x::'a::field_char_0) n = 0 \ x \ \\<^sub>\\<^sub>0" + by (auto simp: pochhammer_eq_0_iff) + +lemma closed_nonpos_Ints [simp]: "closed (\\<^sub>\\<^sub>0 :: 'a :: real_normed_algebra_1 set)" +proof - + have "\\<^sub>\\<^sub>0 = (of_int ` {n. n \ 0} :: 'a set)" + by (auto elim!: nonpos_Ints_cases intro!: nonpos_Ints_of_int) + also have "closed \" by (rule closed_of_int_image) + finally show ?thesis . +qed + +lemma plus_one_in_nonpos_Ints_imp: "z + 1 \ \\<^sub>\\<^sub>0 \ z \ \\<^sub>\\<^sub>0" + using nonpos_Ints_diff_Nats[of "z+1" "1"] by simp_all + +lemma of_int_in_nonpos_Ints_iff: + "(of_int n :: 'a :: ring_char_0) \ \\<^sub>\\<^sub>0 \ n \ 0" + by (auto simp: nonpos_Ints_def) + +lemma one_plus_of_int_in_nonpos_Ints_iff: + "(1 + of_int n :: 'a :: ring_char_0) \ \\<^sub>\\<^sub>0 \ n \ -1" +proof - + have "1 + of_int n = (of_int (n + 1) :: 'a)" by simp + also have "\ \ \\<^sub>\\<^sub>0 \ n + 1 \ 0" by (subst of_int_in_nonpos_Ints_iff) simp_all + also have "\ \ n \ -1" by presburger + finally show ?thesis . +qed + +lemma one_minus_of_nat_in_nonpos_Ints_iff: + "(1 - of_nat n :: 'a :: ring_char_0) \ \\<^sub>\\<^sub>0 \ n > 0" +proof - + have "(1 - of_nat n :: 'a) = of_int (1 - int n)" by simp + also have "\ \ \\<^sub>\\<^sub>0 \ n > 0" by (subst of_int_in_nonpos_Ints_iff) presburger + finally show ?thesis . +qed + +lemma fraction_not_in_ints: + assumes "\(n dvd m)" "n \ 0" + shows "of_int m / of_int n \ (\ :: 'a :: {division_ring,ring_char_0} set)" +proof + assume "of_int m / (of_int n :: 'a) \ \" + then obtain k where "of_int m / of_int n = (of_int k :: 'a)" by (elim Ints_cases) + with assms have "of_int m = (of_int (k * n) :: 'a)" by (auto simp add: divide_simps) + hence "m = k * n" by (subst (asm) of_int_eq_iff) + hence "n dvd m" by simp + with assms(1) show False by contradiction +qed + +lemma fraction_not_in_nats: + assumes "\n dvd m" "n \ 0" + shows "of_int m / of_int n \ (\ :: 'a :: {division_ring,ring_char_0} set)" +proof + assume "of_int m / of_int n \ (\ :: 'a set)" + also note Nats_subset_Ints + finally have "of_int m / of_int n \ (\ :: 'a set)" . + moreover have "of_int m / of_int n \ (\ :: 'a set)" + using assms by (intro fraction_not_in_ints) + ultimately show False by contradiction +qed + +lemma not_in_Ints_imp_not_in_nonpos_Ints: "z \ \ \ z \ \\<^sub>\\<^sub>0" + by (auto simp: Ints_def nonpos_Ints_def) + +lemma double_in_nonpos_Ints_imp: + assumes "2 * (z :: 'a :: field_char_0) \ \\<^sub>\\<^sub>0" + shows "z \ \\<^sub>\\<^sub>0 \ z + 1/2 \ \\<^sub>\\<^sub>0" +proof- + from assms obtain k where k: "2 * z = - of_nat k" by (elim nonpos_Ints_cases') + thus ?thesis by (cases "even k") (auto elim!: evenE oddE simp: field_simps) +qed + + +lemma sin_series: "(\n. ((-1)^n / fact (2*n+1)) *\<^sub>R z^(2*n+1)) sums sin z" +proof - + from sin_converges[of z] have "(\n. sin_coeff n *\<^sub>R z^n) sums sin z" . + also have "(\n. sin_coeff n *\<^sub>R z^n) sums sin z \ + (\n. ((-1)^n / fact (2*n+1)) *\<^sub>R z^(2*n+1)) sums sin z" + by (subst sums_mono_reindex[of "\n. 2*n+1", symmetric]) + (auto simp: sin_coeff_def subseq_def ac_simps elim!: oddE) + finally show ?thesis . +qed + +lemma cos_series: "(\n. ((-1)^n / fact (2*n)) *\<^sub>R z^(2*n)) sums cos z" +proof - + from cos_converges[of z] have "(\n. cos_coeff n *\<^sub>R z^n) sums cos z" . + also have "(\n. cos_coeff n *\<^sub>R z^n) sums cos z \ + (\n. ((-1)^n / fact (2*n)) *\<^sub>R z^(2*n)) sums cos z" + by (subst sums_mono_reindex[of "\n. 2*n", symmetric]) + (auto simp: cos_coeff_def subseq_def ac_simps elim!: evenE) + finally show ?thesis . +qed + +lemma sin_z_over_z_series: + fixes z :: "'a :: {real_normed_field,banach}" + assumes "z \ 0" + shows "(\n. (-1)^n / fact (2*n+1) * z^(2*n)) sums (sin z / z)" +proof - + from sin_series[of z] have "(\n. z * ((-1)^n / fact (2*n+1)) * z^(2*n)) sums sin z" + by (simp add: field_simps scaleR_conv_of_real) + from sums_mult[OF this, of "inverse z"] and assms show ?thesis + by (simp add: field_simps) +qed + +lemma sin_z_over_z_series': + fixes z :: "'a :: {real_normed_field,banach}" + assumes "z \ 0" + shows "(\n. sin_coeff (n+1) *\<^sub>R z^n) sums (sin z / z)" +proof - + from sums_split_initial_segment[OF sin_converges[of z], of 1] + have "(\n. z * (sin_coeff (n+1) *\<^sub>R z ^ n)) sums sin z" by simp + from sums_mult[OF this, of "inverse z"] assms show ?thesis by (simp add: field_simps) +qed + +lemma has_field_derivative_sin_z_over_z: + fixes A :: "'a :: {real_normed_field,banach} set" + shows "((\z. if z = 0 then 1 else sin z / z) has_field_derivative 0) (at 0 within A)" + (is "(?f has_field_derivative ?f') _") +proof (rule has_field_derivative_at_within) + have "((\z::'a. \n. of_real (sin_coeff (n+1)) * z^n) + has_field_derivative (\n. diffs (\n. of_real (sin_coeff (n+1))) n * 0^n)) (at 0)" + proof (rule termdiffs_strong) + from summable_ignore_initial_segment[OF sums_summable[OF sin_converges[of "1::'a"]], of 1] + show "summable (\n. of_real (sin_coeff (n+1)) * (1::'a)^n)" by (simp add: of_real_def) + qed simp + also have "(\z::'a. \n. of_real (sin_coeff (n+1)) * z^n) = ?f" + proof + fix z + show "(\n. of_real (sin_coeff (n+1)) * z^n) = ?f z" + by (cases "z = 0") (insert sin_z_over_z_series'[of z], + simp_all add: scaleR_conv_of_real sums_iff powser_zero sin_coeff_def) + qed + also have "(\n. diffs (\n. of_real (sin_coeff (n + 1))) n * (0::'a) ^ n) = + diffs (\n. of_real (sin_coeff (Suc n))) 0" by (simp add: powser_zero) + also have "\ = 0" by (simp add: sin_coeff_def diffs_def) + finally show "((\z::'a. if z = 0 then 1 else sin z / z) has_field_derivative 0) (at 0)" . +qed + +lemma round_Re_minimises_norm: + "norm ((z::complex) - of_int m) \ norm (z - of_int (round (Re z)))" +proof - + let ?n = "round (Re z)" + have "norm (z - of_int ?n) = sqrt ((Re z - of_int ?n)\<^sup>2 + (Im z)\<^sup>2)" + by (simp add: cmod_def) + also have "\Re z - of_int ?n\ \ \Re z - of_int m\" by (rule round_diff_minimal) + hence "sqrt ((Re z - of_int ?n)\<^sup>2 + (Im z)\<^sup>2) \ sqrt ((Re z - of_int m)\<^sup>2 + (Im z)\<^sup>2)" + by (intro real_sqrt_le_mono add_mono) (simp_all add: abs_le_square_iff) + also have "\ = norm (z - of_int m)" by (simp add: cmod_def) + finally show ?thesis . +qed + +lemma Re_pos_in_ball: + assumes "Re z > 0" "t \ ball z (Re z/2)" + shows "Re t > 0" +proof - + have "Re (z - t) \ norm (z - t)" by (rule complex_Re_le_cmod) + also from assms have "\ < Re z / 2" by (simp add: dist_complex_def) + finally show "Re t > 0" using assms by simp +qed + +lemma no_nonpos_Int_in_ball_complex: + assumes "Re z > 0" "t \ ball z (Re z/2)" + shows "t \ \\<^sub>\\<^sub>0" + using Re_pos_in_ball[OF assms] by (force elim!: nonpos_Ints_cases) + +lemma no_nonpos_Int_in_ball: + assumes "t \ ball z (dist z (round (Re z)))" + shows "t \ \\<^sub>\\<^sub>0" +proof + assume "t \ \\<^sub>\\<^sub>0" + then obtain n where "t = of_int n" by (auto elim!: nonpos_Ints_cases) + have "dist z (of_int n) \ dist z t + dist t (of_int n)" by (rule dist_triangle) + also from assms have "dist z t < dist z (round (Re z))" by simp + also have "\ \ dist z (of_int n)" + using round_Re_minimises_norm[of z] by (simp add: dist_complex_def) + finally have "dist t (of_int n) > 0" by simp + with \t = of_int n\ show False by simp +qed + +lemma no_nonpos_Int_in_ball': + assumes "(z :: 'a :: {euclidean_space,real_normed_algebra_1}) \ \\<^sub>\\<^sub>0" + obtains d where "d > 0" "\t. t \ ball z d \ t \ \\<^sub>\\<^sub>0" +proof (rule that) + from assms show "setdist {z} \\<^sub>\\<^sub>0 > 0" by (subst setdist_gt_0_compact_closed) auto +next + fix t assume "t \ ball z (setdist {z} \\<^sub>\\<^sub>0)" + thus "t \ \\<^sub>\\<^sub>0" using setdist_le_dist[of z "{z}" t "\\<^sub>\\<^sub>0"] by force +qed + +lemma no_nonpos_Real_in_ball: + assumes z: "z \ \\<^sub>\\<^sub>0" and t: "t \ ball z (if Im z = 0 then Re z / 2 else abs (Im z) / 2)" + shows "t \ \\<^sub>\\<^sub>0" +using z +proof (cases "Im z = 0") + assume A: "Im z = 0" + with z have "Re z > 0" by (force simp add: complex_nonpos_Reals_iff) + with t A Re_pos_in_ball[of z t] show ?thesis by (force simp add: complex_nonpos_Reals_iff) +next + assume A: "Im z \ 0" + have "abs (Im z) - abs (Im t) \ abs (Im z - Im t)" by linarith + also have "\ = abs (Im (z - t))" by simp + also have "\ \ norm (z - t)" by (rule abs_Im_le_cmod) + also from A t have "\ \ abs (Im z) / 2" by (simp add: dist_complex_def) + finally have "abs (Im t) > 0" using A by simp + thus ?thesis by (force simp add: complex_nonpos_Reals_iff) +qed + + +subsection \Definitions\ + +text \ + We define the Gamma function by first defining its multiplicative inverse @{term "Gamma_inv"}. + This is more convenient because @{term "Gamma_inv"} is entire, which makes proofs of its + properties more convenient because one does not have to watch out for discontinuities. + (e.g. @{term "Gamma_inv"} fulfils @{term "rGamma z = z * rGamma (z + 1)"} everywhere, + whereas @{term "Gamma"} does not fulfil the analogous equation on the non-positive integers) + + We define the Gamma function (resp. its inverse) in the Euler form. This form has the advantage + that it is a relatively simple limit that converges everywhere. The limit at the poles is 0 + (due to division by 0). The functional equation @{term "Gamma (z + 1) = z * Gamma z"} follows + immediately from the definition. +\ + +definition Gamma_series :: "('a :: {banach,real_normed_field}) \ nat \ 'a" where + "Gamma_series z n = fact n * exp (z * of_real (ln (of_nat n))) / pochhammer z (n+1)" + +definition Gamma_series' :: "('a :: {banach,real_normed_field}) \ nat \ 'a" where + "Gamma_series' z n = fact (n - 1) * exp (z * of_real (ln (of_nat n))) / pochhammer z n" + +definition rGamma_series :: "('a :: {banach,real_normed_field}) \ nat \ 'a" where + "rGamma_series z n = pochhammer z (n+1) / (fact n * exp (z * of_real (ln (of_nat n))))" + +lemma Gamma_series_altdef: "Gamma_series z n = inverse (rGamma_series z n)" + and rGamma_series_altdef: "rGamma_series z n = inverse (Gamma_series z n)" + unfolding Gamma_series_def rGamma_series_def by simp_all + +lemma rGamma_series_minus_of_nat: + "eventually (\n. rGamma_series (- of_nat k) n = 0) sequentially" + using eventually_ge_at_top[of k] + by eventually_elim (auto simp: rGamma_series_def pochhammer_of_nat_eq_0_iff) + +lemma Gamma_series_minus_of_nat: + "eventually (\n. Gamma_series (- of_nat k) n = 0) sequentially" + using eventually_ge_at_top[of k] + by eventually_elim (auto simp: Gamma_series_def pochhammer_of_nat_eq_0_iff) + +lemma Gamma_series'_minus_of_nat: + "eventually (\n. Gamma_series' (- of_nat k) n = 0) sequentially" + using eventually_gt_at_top[of k] + by eventually_elim (auto simp: Gamma_series'_def pochhammer_of_nat_eq_0_iff) + +lemma rGamma_series_nonpos_Ints_LIMSEQ: "z \ \\<^sub>\\<^sub>0 \ rGamma_series z \ 0" + by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule rGamma_series_minus_of_nat, simp) + +lemma Gamma_series_nonpos_Ints_LIMSEQ: "z \ \\<^sub>\\<^sub>0 \ Gamma_series z \ 0" + by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule Gamma_series_minus_of_nat, simp) + +lemma Gamma_series'_nonpos_Ints_LIMSEQ: "z \ \\<^sub>\\<^sub>0 \ Gamma_series' z \ 0" + by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule Gamma_series'_minus_of_nat, simp) + +lemma Gamma_series_Gamma_series': + assumes z: "z \ \\<^sub>\\<^sub>0" + shows "(\n. Gamma_series' z n / Gamma_series z n) \ 1" +proof (rule Lim_transform_eventually) + from eventually_gt_at_top[of "0::nat"] + show "eventually (\n. z / of_nat n + 1 = Gamma_series' z n / Gamma_series z n) sequentially" + proof eventually_elim + fix n :: nat assume n: "n > 0" + from n z have "Gamma_series' z n / Gamma_series z n = (z + of_nat n) / of_nat n" + by (cases n, simp) + (auto simp add: Gamma_series_def Gamma_series'_def pochhammer_rec' + dest: pochhammer_eq_0_imp_nonpos_Int plus_of_nat_eq_0_imp) + also from n have "\ = z / of_nat n + 1" by (simp add: divide_simps) + finally show "z / of_nat n + 1 = Gamma_series' z n / Gamma_series z n" .. + qed + have "(\x. z / of_nat x) \ 0" + by (rule tendsto_norm_zero_cancel) + (insert tendsto_mult[OF tendsto_const[of "norm z"] lim_inverse_n], + simp add: norm_divide inverse_eq_divide) + from tendsto_add[OF this tendsto_const[of 1]] show "(\n. z / of_nat n + 1) \ 1" by simp +qed + + +subsection \Convergence of the Euler series form\ + +text \ + We now show that the series that defines the Gamma function in the Euler form converges + and that the function defined by it is continuous on the complex halfspace with positive + real part. + + We do this by showing that the logarithm of the Euler series is continuous and converges + locally uniformly, which means that the log-Gamma function defined by its limit is also + continuous. + + This will later allow us to lift holomorphicity and continuity from the log-Gamma + function to the inverse of the Gamma function, and from that to the Gamma function itself. +\ + +definition ln_Gamma_series :: "('a :: {banach,real_normed_field,ln}) \ nat \ 'a" where + "ln_Gamma_series z n = z * ln (of_nat n) - ln z - (\k=1..n. ln (z / of_nat k + 1))" + +definition ln_Gamma_series' :: "('a :: {banach,real_normed_field,ln}) \ nat \ 'a" where + "ln_Gamma_series' z n = + - euler_mascheroni*z - ln z + (\k=1..n. z / of_nat n - ln (z / of_nat k + 1))" + +definition ln_Gamma :: "('a :: {banach,real_normed_field,ln}) \ 'a" where + "ln_Gamma z = lim (ln_Gamma_series z)" + + +text \ + We now show that the log-Gamma series converges locally uniformly for all complex numbers except + the non-positive integers. We do this by proving that the series is locally Cauchy, adapting this + proof: + http://math.stackexchange.com/questions/887158/convergence-of-gammaz-lim-n-to-infty-fracnzn-prod-m-0nzm +\ + +context +begin + +private lemma ln_Gamma_series_complex_converges_aux: + fixes z :: complex and k :: nat + assumes z: "z \ 0" and k: "of_nat k \ 2*norm z" "k \ 2" + shows "norm (z * ln (1 - 1/of_nat k) + ln (z/of_nat k + 1)) \ 2*(norm z + norm z^2) / of_nat k^2" +proof - + let ?k = "of_nat k :: complex" and ?z = "norm z" + have "z *ln (1 - 1/?k) + ln (z/?k+1) = z*(ln (1 - 1/?k :: complex) + 1/?k) + (ln (1+z/?k) - z/?k)" + by (simp add: algebra_simps) + also have "norm ... \ ?z * norm (ln (1-1/?k) + 1/?k :: complex) + norm (ln (1+z/?k) - z/?k)" + by (subst norm_mult [symmetric], rule norm_triangle_ineq) + also have "norm (Ln (1 + -1/?k) - (-1/?k)) \ (norm (-1/?k))\<^sup>2 / (1 - norm(-1/?k))" + using k by (intro Ln_approx_linear) (simp add: norm_divide) + hence "?z * norm (ln (1-1/?k) + 1/?k) \ ?z * ((norm (1/?k))^2 / (1 - norm (1/?k)))" + by (intro mult_left_mono) simp_all + also have "... \ (?z * (of_nat k / (of_nat k - 1))) / of_nat k^2" using k + by (simp add: field_simps power2_eq_square norm_divide) + also have "... \ (?z * 2) / of_nat k^2" using k + by (intro divide_right_mono mult_left_mono) (simp_all add: field_simps) + also have "norm (ln (1+z/?k) - z/?k) \ norm (z/?k)^2 / (1 - norm (z/?k))" using k + by (intro Ln_approx_linear) (simp add: norm_divide) + hence "norm (ln (1+z/?k) - z/?k) \ ?z^2 / of_nat k^2 / (1 - ?z / of_nat k)" + by (simp add: field_simps norm_divide) + also have "... \ (?z^2 * (of_nat k / (of_nat k - ?z))) / of_nat k^2" using k + by (simp add: field_simps power2_eq_square) + also have "... \ (?z^2 * 2) / of_nat k^2" using k + by (intro divide_right_mono mult_left_mono) (simp_all add: field_simps) + also note add_divide_distrib [symmetric] + finally show ?thesis by (simp only: distrib_left mult.commute) +qed + +lemma ln_Gamma_series_complex_converges: + assumes z: "z \ \\<^sub>\\<^sub>0" + assumes d: "d > 0" "\n. n \ \\<^sub>\\<^sub>0 \ norm (z - of_int n) > d" + shows "uniformly_convergent_on (ball z d) (\n z. ln_Gamma_series z n :: complex)" +proof (intro Cauchy_uniformly_convergent uniformly_Cauchy_onI') + fix e :: real assume e: "e > 0" + define e'' where "e'' = (SUP t:ball z d. norm t + norm t^2)" + define e' where "e' = e / (2*e'')" + have "bounded ((\t. norm t + norm t^2) ` cball z d)" + by (intro compact_imp_bounded compact_continuous_image) (auto intro!: continuous_intros) + hence "bounded ((\t. norm t + norm t^2) ` ball z d)" by (rule bounded_subset) auto + hence bdd: "bdd_above ((\t. norm t + norm t^2) ` ball z d)" by (rule bounded_imp_bdd_above) + + with z d(1) d(2)[of "-1"] have e''_pos: "e'' > 0" unfolding e''_def + by (subst less_cSUP_iff) (auto intro!: add_pos_nonneg bexI[of _ z]) + have e'': "norm t + norm t^2 \ e''" if "t \ ball z d" for t unfolding e''_def using that + by (rule cSUP_upper[OF _ bdd]) + from e z e''_pos have e': "e' > 0" unfolding e'_def + by (intro divide_pos_pos mult_pos_pos add_pos_pos) (simp_all add: field_simps) + + have "summable (\k. inverse ((real_of_nat k)^2))" + by (rule inverse_power_summable) simp + from summable_partial_sum_bound[OF this e'] guess M . note M = this + + define N where "N = max 2 (max (nat \2 * (norm z + d)\) M)" + { + from d have "\2 * (cmod z + d)\ \ \0::real\" + by (intro ceiling_mono mult_nonneg_nonneg add_nonneg_nonneg) simp_all + hence "2 * (norm z + d) \ of_nat (nat \2 * (norm z + d)\)" unfolding N_def + by (simp_all add: le_of_int_ceiling) + also have "... \ of_nat N" unfolding N_def + by (subst of_nat_le_iff) (rule max.coboundedI2, rule max.cobounded1) + finally have "of_nat N \ 2 * (norm z + d)" . + moreover have "N \ 2" "N \ M" unfolding N_def by simp_all + moreover have "(\k=m..n. 1/(of_nat k)\<^sup>2) < e'" if "m \ N" for m n + using M[OF order.trans[OF \N \ M\ that]] unfolding real_norm_def + by (subst (asm) abs_of_nonneg) (auto intro: setsum_nonneg simp: divide_simps) + moreover note calculation + } note N = this + + show "\M. \t\ball z d. \m\M. \n>m. dist (ln_Gamma_series t m) (ln_Gamma_series t n) < e" + unfolding dist_complex_def + proof (intro exI[of _ N] ballI allI impI) + fix t m n assume t: "t \ ball z d" and mn: "m \ N" "n > m" + from d(2)[of 0] t have "0 < dist z 0 - dist z t" by (simp add: field_simps dist_complex_def) + also have "dist z 0 - dist z t \ dist 0 t" using dist_triangle[of 0 z t] + by (simp add: dist_commute) + finally have t_nz: "t \ 0" by auto + + have "norm t \ norm z + norm (t - z)" by (rule norm_triangle_sub) + also from t have "norm (t - z) < d" by (simp add: dist_complex_def norm_minus_commute) + also have "2 * (norm z + d) \ of_nat N" by (rule N) + also have "N \ m" by (rule mn) + finally have norm_t: "2 * norm t < of_nat m" by simp + + have "ln_Gamma_series t m - ln_Gamma_series t n = + (-(t * Ln (of_nat n)) - (-(t * Ln (of_nat m)))) + + ((\k=1..n. Ln (t / of_nat k + 1)) - (\k=1..m. Ln (t / of_nat k + 1)))" + by (simp add: ln_Gamma_series_def algebra_simps) + also have "(\k=1..n. Ln (t / of_nat k + 1)) - (\k=1..m. Ln (t / of_nat k + 1)) = + (\k\{1..n}-{1..m}. Ln (t / of_nat k + 1))" using mn + by (simp add: setsum_diff) + also from mn have "{1..n}-{1..m} = {Suc m..n}" by fastforce + also have "-(t * Ln (of_nat n)) - (-(t * Ln (of_nat m))) = + (\k = Suc m..n. t * Ln (of_nat (k - 1)) - t * Ln (of_nat k))" using mn + by (subst setsum_telescope'' [symmetric]) simp_all + also have "... = (\k = Suc m..n. t * Ln (of_nat (k - 1) / of_nat k))" using mn N + by (intro setsum_cong_Suc) + (simp_all del: of_nat_Suc add: field_simps Ln_of_nat Ln_of_nat_over_of_nat) + also have "of_nat (k - 1) / of_nat k = 1 - 1 / (of_nat k :: complex)" if "k \ {Suc m..n}" for k + using that of_nat_eq_0_iff[of "Suc i" for i] by (cases k) (simp_all add: divide_simps) + hence "(\k = Suc m..n. t * Ln (of_nat (k - 1) / of_nat k)) = + (\k = Suc m..n. t * Ln (1 - 1 / of_nat k))" using mn N + by (intro setsum.cong) simp_all + also note setsum.distrib [symmetric] + also have "norm (\k=Suc m..n. t * Ln (1 - 1/of_nat k) + Ln (t/of_nat k + 1)) \ + (\k=Suc m..n. 2 * (norm t + (norm t)\<^sup>2) / (real_of_nat k)\<^sup>2)" using t_nz N(2) mn norm_t + by (intro order.trans[OF norm_setsum setsum_mono[OF ln_Gamma_series_complex_converges_aux]]) simp_all + also have "... \ 2 * (norm t + norm t^2) * (\k=Suc m..n. 1 / (of_nat k)\<^sup>2)" + by (simp add: setsum_right_distrib) + also have "... < 2 * (norm t + norm t^2) * e'" using mn z t_nz + by (intro mult_strict_left_mono N mult_pos_pos add_pos_pos) simp_all + also from e''_pos have "... = e * ((cmod t + (cmod t)\<^sup>2) / e'')" + by (simp add: e'_def field_simps power2_eq_square) + also from e''[OF t] e''_pos e + have "\ \ e * 1" by (intro mult_left_mono) (simp_all add: field_simps) + finally show "norm (ln_Gamma_series t m - ln_Gamma_series t n) < e" by simp + qed +qed + +end + +lemma ln_Gamma_series_complex_converges': + assumes z: "(z :: complex) \ \\<^sub>\\<^sub>0" + shows "\d>0. uniformly_convergent_on (ball z d) (\n z. ln_Gamma_series z n)" +proof - + define d' where "d' = Re z" + define d where "d = (if d' > 0 then d' / 2 else norm (z - of_int (round d')) / 2)" + have "of_int (round d') \ \\<^sub>\\<^sub>0" if "d' \ 0" using that + by (intro nonpos_Ints_of_int) (simp_all add: round_def) + with assms have d_pos: "d > 0" unfolding d_def by (force simp: not_less) + + have "d < cmod (z - of_int n)" if "n \ \\<^sub>\\<^sub>0" for n + proof (cases "Re z > 0") + case True + from nonpos_Ints_nonpos[OF that] have n: "n \ 0" by simp + from True have "d = Re z/2" by (simp add: d_def d'_def) + also from n True have "\ < Re (z - of_int n)" by simp + also have "\ \ norm (z - of_int n)" by (rule complex_Re_le_cmod) + finally show ?thesis . + next + case False + with assms nonpos_Ints_of_int[of "round (Re z)"] + have "z \ of_int (round d')" by (auto simp: not_less) + with False have "d < norm (z - of_int (round d'))" by (simp add: d_def d'_def) + also have "\ \ norm (z - of_int n)" unfolding d'_def by (rule round_Re_minimises_norm) + finally show ?thesis . + qed + hence conv: "uniformly_convergent_on (ball z d) (\n z. ln_Gamma_series z n)" + by (intro ln_Gamma_series_complex_converges d_pos z) simp_all + from d_pos conv show ?thesis by blast +qed + +lemma ln_Gamma_series_complex_converges'': "(z :: complex) \ \\<^sub>\\<^sub>0 \ convergent (ln_Gamma_series z)" + by (drule ln_Gamma_series_complex_converges') (auto intro: uniformly_convergent_imp_convergent) + +lemma ln_Gamma_complex_LIMSEQ: "(z :: complex) \ \\<^sub>\\<^sub>0 \ ln_Gamma_series z \ ln_Gamma z" + using ln_Gamma_series_complex_converges'' by (simp add: convergent_LIMSEQ_iff ln_Gamma_def) + +lemma exp_ln_Gamma_series_complex: + assumes "n > 0" "z \ \\<^sub>\\<^sub>0" + shows "exp (ln_Gamma_series z n :: complex) = Gamma_series z n" +proof - + from assms obtain m where m: "n = Suc m" by (cases n) blast + from assms have "z \ 0" by (intro notI) auto + with assms have "exp (ln_Gamma_series z n) = + (of_nat n) powr z / (z * (\k=1..n. exp (Ln (z / of_nat k + 1))))" + unfolding ln_Gamma_series_def powr_def by (simp add: exp_diff exp_setsum) + also from assms have "(\k=1..n. exp (Ln (z / of_nat k + 1))) = (\k=1..n. z / of_nat k + 1)" + by (intro setprod.cong[OF refl], subst exp_Ln) (auto simp: field_simps plus_of_nat_eq_0_imp) + also have "... = (\k=1..n. z + k) / fact n" + by (simp add: fact_setprod) + (subst setprod_dividef [symmetric], simp_all add: field_simps) + also from m have "z * ... = (\k=0..n. z + k) / fact n" + by (simp add: setprod.atLeast0_atMost_Suc_shift setprod.atLeast_Suc_atMost_Suc_shift) + also have "(\k=0..n. z + k) = pochhammer z (Suc n)" + unfolding pochhammer_setprod + by (simp add: setprod.atLeast0_atMost_Suc atLeastLessThanSuc_atLeastAtMost) + also have "of_nat n powr z / (pochhammer z (Suc n) / fact n) = Gamma_series z n" + unfolding Gamma_series_def using assms by (simp add: divide_simps powr_def Ln_of_nat) + finally show ?thesis . +qed + + +lemma ln_Gamma_series'_aux: + assumes "(z::complex) \ \\<^sub>\\<^sub>0" + shows "(\k. z / of_nat (Suc k) - ln (1 + z / of_nat (Suc k))) sums + (ln_Gamma z + euler_mascheroni * z + ln z)" (is "?f sums ?s") +unfolding sums_def +proof (rule Lim_transform) + show "(\n. ln_Gamma_series z n + of_real (harm n - ln (of_nat n)) * z + ln z) \ ?s" + (is "?g \ _") + by (intro tendsto_intros ln_Gamma_complex_LIMSEQ euler_mascheroni_LIMSEQ_of_real assms) + + have A: "eventually (\n. (\k 0" + have "(\kk=1..n. z / of_nat k - ln (1 + z / of_nat k))" + by (subst atLeast0LessThan [symmetric], subst setsum_shift_bounds_Suc_ivl [symmetric], + subst atLeastLessThanSuc_atLeastAtMost) simp_all + also have "\ = z * of_real (harm n) - (\k=1..n. ln (1 + z / of_nat k))" + by (simp add: harm_def setsum_subtractf setsum_right_distrib divide_inverse) + also from n have "\ - ?g n = 0" + by (simp add: ln_Gamma_series_def setsum_subtractf algebra_simps Ln_of_nat) + finally show "(\kn. (\k 0" by (subst tendsto_cong[OF A]) simp_all +qed + + +lemma uniformly_summable_deriv_ln_Gamma: + assumes z: "(z :: 'a :: {real_normed_field,banach}) \ 0" and d: "d > 0" "d \ norm z/2" + shows "uniformly_convergent_on (ball z d) + (\k z. \ik z. \i ball z d" + have "norm z = norm (t + (z - t))" by simp + have "norm (t + (z - t)) \ norm t + norm (z - t)" by (rule norm_triangle_ineq) + also from t d have "norm (z - t) < norm z / 2" by (simp add: dist_norm) + finally have A: "norm t > norm z / 2" using z by (simp add: field_simps) + + have "norm t = norm (z + (t - z))" by simp + also have "\ \ norm z + norm (t - z)" by (rule norm_triangle_ineq) + also from t d have "norm (t - z) \ norm z / 2" by (simp add: dist_norm norm_minus_commute) + also from z have "\ < norm z" by simp + finally have B: "norm t < 2 * norm z" by simp + note A B + } note ball = this + + show "eventually (\n. \t\ball z d. norm (?f n t) \ 4 * norm z * inverse (of_nat (Suc n)^2)) sequentially" + using eventually_gt_at_top apply eventually_elim + proof safe + fix t :: 'a assume t: "t \ ball z d" + from z ball[OF t] have t_nz: "t \ 0" by auto + fix n :: nat assume n: "n > nat \4 * norm z\" + from ball[OF t] t_nz have "4 * norm z > 2 * norm t" by simp + also from n have "\ < of_nat n" by linarith + finally have n: "of_nat n > 2 * norm t" . + hence "of_nat n > norm t" by simp + hence t': "t \ -of_nat (Suc n)" by (intro notI) (simp del: of_nat_Suc) + + with t_nz have "?f n t = 1 / (of_nat (Suc n) * (1 + of_nat (Suc n)/t))" + by (simp add: divide_simps eq_neg_iff_add_eq_0 del: of_nat_Suc) + also have "norm \ = inverse (of_nat (Suc n)) * inverse (norm (of_nat (Suc n)/t + 1))" + by (simp add: norm_divide norm_mult divide_simps add_ac del: of_nat_Suc) + also { + from z t_nz ball[OF t] have "of_nat (Suc n) / (4 * norm z) \ of_nat (Suc n) / (2 * norm t)" + by (intro divide_left_mono mult_pos_pos) simp_all + also have "\ < norm (of_nat (Suc n) / t) - norm (1 :: 'a)" + using t_nz n by (simp add: field_simps norm_divide del: of_nat_Suc) + also have "\ \ norm (of_nat (Suc n)/t + 1)" by (rule norm_diff_ineq) + finally have "inverse (norm (of_nat (Suc n)/t + 1)) \ 4 * norm z / of_nat (Suc n)" + using z by (simp add: divide_simps norm_divide mult_ac del: of_nat_Suc) + } + also have "inverse (real_of_nat (Suc n)) * (4 * norm z / real_of_nat (Suc n)) = + 4 * norm z * inverse (of_nat (Suc n)^2)" + by (simp add: divide_simps power2_eq_square del: of_nat_Suc) + finally show "norm (?f n t) \ 4 * norm z * inverse (of_nat (Suc n)^2)" + by (simp del: of_nat_Suc) + qed +next + show "summable (\n. 4 * norm z * inverse ((of_nat (Suc n))^2))" + by (subst summable_Suc_iff) (simp add: summable_mult inverse_power_summable) +qed + +lemma summable_deriv_ln_Gamma: + "z \ (0 :: 'a :: {real_normed_field,banach}) \ + summable (\n. inverse (of_nat (Suc n)) - inverse (z + of_nat (Suc n)))" + unfolding summable_iff_convergent + by (rule uniformly_convergent_imp_convergent, + rule uniformly_summable_deriv_ln_Gamma[of z "norm z/2"]) simp_all + + +definition Polygamma :: "nat \ ('a :: {real_normed_field,banach}) \ 'a" where + "Polygamma n z = (if n = 0 then + (\k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) - euler_mascheroni else + (-1)^Suc n * fact n * (\k. inverse ((z + of_nat k)^Suc n)))" + +abbreviation Digamma :: "('a :: {real_normed_field,banach}) \ 'a" where + "Digamma \ Polygamma 0" + +lemma Digamma_def: + "Digamma z = (\k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) - euler_mascheroni" + by (simp add: Polygamma_def) + + +lemma summable_Digamma: + assumes "(z :: 'a :: {real_normed_field,banach}) \ 0" + shows "summable (\n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))" +proof - + have sums: "(\n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n)) sums + (0 - inverse (z + of_nat 0))" + by (intro telescope_sums filterlim_compose[OF tendsto_inverse_0] + tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat) + from summable_add[OF summable_deriv_ln_Gamma[OF assms] sums_summable[OF sums]] + show "summable (\n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))" by simp +qed + +lemma summable_offset: + assumes "summable (\n. f (n + k) :: 'a :: real_normed_vector)" + shows "summable f" +proof - + from assms have "convergent (\m. \nm. (\nnm. (\nnm. \n {k..n\\. f n) = (\nn=k..n=k..n=0..n. n + k"]) + (simp, subst image_add_atLeastLessThan, auto) + finally show "(\nnna. setsum f {.. lim (\m. setsum f {.. 0" and n: "n \ 2" + shows "uniformly_convergent_on (ball z d) (\k z. \inorm z * e\" + { + fix t assume t: "t \ ball z d" + have "norm t = norm (z + (t - z))" by simp + also have "\ \ norm z + norm (t - z)" by (rule norm_triangle_ineq) + also from t have "norm (t - z) < d" by (simp add: dist_norm norm_minus_commute) + finally have "norm t < norm z * e" using z by (simp add: divide_simps e_def) + } note ball = this + + show "eventually (\k. \t\ball z d. norm (inverse ((t + of_nat k)^n)) \ + inverse (of_nat (k - m)^n)) sequentially" + using eventually_gt_at_top[of m] apply eventually_elim + proof (intro ballI) + fix k :: nat and t :: 'a assume k: "k > m" and t: "t \ ball z d" + from k have "real_of_nat (k - m) = of_nat k - of_nat m" by (simp add: of_nat_diff) + also have "\ \ norm (of_nat k :: 'a) - norm z * e" + unfolding m_def by (subst norm_of_nat) linarith + also from ball[OF t] have "\ \ norm (of_nat k :: 'a) - norm t" by simp + also have "\ \ norm (of_nat k + t)" by (rule norm_diff_ineq) + finally have "inverse ((norm (t + of_nat k))^n) \ inverse (real_of_nat (k - m)^n)" using k n + by (intro le_imp_inverse_le power_mono) (simp_all add: add_ac del: of_nat_Suc) + thus "norm (inverse ((t + of_nat k)^n)) \ inverse (of_nat (k - m)^n)" + by (simp add: norm_inverse norm_power power_inverse) + qed + + have "summable (\k. inverse ((real_of_nat k)^n))" + using inverse_power_summable[of n] n by simp + hence "summable (\k. inverse ((real_of_nat (k + m - m))^n))" by simp + thus "summable (\k. inverse ((real_of_nat (k - m))^n))" by (rule summable_offset) +qed + +lemma Polygamma_converges': + fixes z :: "'a :: {real_normed_field,banach}" + assumes z: "z \ 0" and n: "n \ 2" + shows "summable (\k. inverse ((z + of_nat k)^n))" + using uniformly_convergent_imp_convergent[OF Polygamma_converges[OF assms, of 1], of z] + by (simp add: summable_iff_convergent) + +lemma has_field_derivative_ln_Gamma_complex [derivative_intros]: + fixes z :: complex + assumes z: "z \ \\<^sub>\\<^sub>0" + shows "(ln_Gamma has_field_derivative Digamma z) (at z)" +proof - + have not_nonpos_Int [simp]: "t \ \\<^sub>\\<^sub>0" if "Re t > 0" for t + using that by (auto elim!: nonpos_Ints_cases') + from z have z': "z \ \\<^sub>\\<^sub>0" and z'': "z \ 0" using nonpos_Ints_subset_nonpos_Reals nonpos_Reals_zero_I + by blast+ + let ?f' = "\z k. inverse (of_nat (Suc k)) - inverse (z + of_nat (Suc k))" + let ?f = "\z k. z / of_nat (Suc k) - ln (1 + z / of_nat (Suc k))" and ?F' = "\z. \n. ?f' z n" + define d where "d = min (norm z/2) (if Im z = 0 then Re z / 2 else abs (Im z) / 2)" + from z have d: "d > 0" "norm z/2 \ d" by (auto simp add: complex_nonpos_Reals_iff d_def) + have ball: "Im t = 0 \ Re t > 0" if "dist z t < d" for t + using no_nonpos_Real_in_ball[OF z, of t] that unfolding d_def by (force simp add: complex_nonpos_Reals_iff) + have sums: "(\n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n)) sums + (0 - inverse (z + of_nat 0))" + by (intro telescope_sums filterlim_compose[OF tendsto_inverse_0] + tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat) + + have "((\z. \n. ?f z n) has_field_derivative ?F' z) (at z)" + using d z ln_Gamma_series'_aux[OF z'] + apply (intro has_field_derivative_series'(2)[of "ball z d" _ _ z] uniformly_summable_deriv_ln_Gamma) + apply (auto intro!: derivative_eq_intros add_pos_pos mult_pos_pos dest!: ball + simp: field_simps sums_iff nonpos_Reals_divide_of_nat_iff + simp del: of_nat_Suc) + apply (auto simp add: complex_nonpos_Reals_iff) + done + with z have "((\z. (\k. ?f z k) - euler_mascheroni * z - Ln z) has_field_derivative + ?F' z - euler_mascheroni - inverse z) (at z)" + by (force intro!: derivative_eq_intros simp: Digamma_def) + also have "?F' z - euler_mascheroni - inverse z = (?F' z + -inverse z) - euler_mascheroni" by simp + also from sums have "-inverse z = (\n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n))" + by (simp add: sums_iff) + also from sums summable_deriv_ln_Gamma[OF z''] + have "?F' z + \ = (\n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))" + by (subst suminf_add) (simp_all add: add_ac sums_iff) + also have "\ - euler_mascheroni = Digamma z" by (simp add: Digamma_def) + finally have "((\z. (\k. ?f z k) - euler_mascheroni * z - Ln z) + has_field_derivative Digamma z) (at z)" . + moreover from eventually_nhds_ball[OF d(1), of z] + have "eventually (\z. ln_Gamma z = (\k. ?f z k) - euler_mascheroni * z - Ln z) (nhds z)" + proof eventually_elim + fix t assume "t \ ball z d" + hence "t \ \\<^sub>\\<^sub>0" by (auto dest!: ball elim!: nonpos_Ints_cases) + from ln_Gamma_series'_aux[OF this] + show "ln_Gamma t = (\k. ?f t k) - euler_mascheroni * t - Ln t" by (simp add: sums_iff) + qed + ultimately show ?thesis by (subst DERIV_cong_ev[OF refl _ refl]) +qed + +declare has_field_derivative_ln_Gamma_complex[THEN DERIV_chain2, derivative_intros] + + +lemma Digamma_1 [simp]: "Digamma (1 :: 'a :: {real_normed_field,banach}) = - euler_mascheroni" + by (simp add: Digamma_def) + +lemma Digamma_plus1: + assumes "z \ 0" + shows "Digamma (z+1) = Digamma z + 1/z" +proof - + have sums: "(\k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k))) + sums (inverse (z + of_nat 0) - 0)" + by (intro telescope_sums'[OF filterlim_compose[OF tendsto_inverse_0]] + tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat) + have "Digamma (z+1) = (\k. inverse (of_nat (Suc k)) - inverse (z + of_nat (Suc k))) - + euler_mascheroni" (is "_ = suminf ?f - _") by (simp add: Digamma_def add_ac) + also have "suminf ?f = (\k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) + + (\k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k)))" + using summable_Digamma[OF assms] sums by (subst suminf_add) (simp_all add: add_ac sums_iff) + also have "(\k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k))) = 1/z" + using sums by (simp add: sums_iff inverse_eq_divide) + finally show ?thesis by (simp add: Digamma_def[of z]) +qed + +lemma Polygamma_plus1: + assumes "z \ 0" + shows "Polygamma n (z + 1) = Polygamma n z + (-1)^n * fact n / (z ^ Suc n)" +proof (cases "n = 0") + assume n: "n \ 0" + let ?f = "\k. inverse ((z + of_nat k) ^ Suc n)" + have "Polygamma n (z + 1) = (-1) ^ Suc n * fact n * (\k. ?f (k+1))" + using n by (simp add: Polygamma_def add_ac) + also have "(\k. ?f (k+1)) + (\k<1. ?f k) = (\k. ?f k)" + using Polygamma_converges'[OF assms, of "Suc n"] n + by (subst suminf_split_initial_segment [symmetric]) simp_all + hence "(\k. ?f (k+1)) = (\k. ?f k) - inverse (z ^ Suc n)" by (simp add: algebra_simps) + also have "(-1) ^ Suc n * fact n * ((\k. ?f k) - inverse (z ^ Suc n)) = + Polygamma n z + (-1)^n * fact n / (z ^ Suc n)" using n + by (simp add: inverse_eq_divide algebra_simps Polygamma_def) + finally show ?thesis . +qed (insert assms, simp add: Digamma_plus1 inverse_eq_divide) + +lemma Digamma_of_nat: + "Digamma (of_nat (Suc n) :: 'a :: {real_normed_field,banach}) = harm n - euler_mascheroni" +proof (induction n) + case (Suc n) + have "Digamma (of_nat (Suc (Suc n)) :: 'a) = Digamma (of_nat (Suc n) + 1)" by simp + also have "\ = Digamma (of_nat (Suc n)) + inverse (of_nat (Suc n))" + by (subst Digamma_plus1) (simp_all add: inverse_eq_divide del: of_nat_Suc) + also have "Digamma (of_nat (Suc n) :: 'a) = harm n - euler_mascheroni " by (rule Suc) + also have "\ + inverse (of_nat (Suc n)) = harm (Suc n) - euler_mascheroni" + by (simp add: harm_Suc) + finally show ?case . +qed (simp add: harm_def) + +lemma Digamma_numeral: "Digamma (numeral n) = harm (pred_numeral n) - euler_mascheroni" + by (subst of_nat_numeral[symmetric], subst numeral_eq_Suc, subst Digamma_of_nat) (rule refl) + +lemma Polygamma_of_real: "x \ 0 \ Polygamma n (of_real x) = of_real (Polygamma n x)" + unfolding Polygamma_def using summable_Digamma[of x] Polygamma_converges'[of x "Suc n"] + by (simp_all add: suminf_of_real) + +lemma Polygamma_Real: "z \ \ \ z \ 0 \ Polygamma n z \ \" + by (elim Reals_cases, hypsubst, subst Polygamma_of_real) simp_all + +lemma Digamma_half_integer: + "Digamma (of_nat n + 1/2 :: 'a :: {real_normed_field,banach}) = + (\kk. inverse (of_nat (Suc k)) - inverse (of_nat k + 1/2)) - euler_mascheroni" + by (simp add: Digamma_def add_ac) + also have "(\k. inverse (of_nat (Suc k) :: real) - inverse (of_nat k + 1/2)) = + (\k. inverse (1/2) * (inverse (2 * of_nat (Suc k)) - inverse (2 * of_nat k + 1)))" + by (simp_all add: add_ac inverse_mult_distrib[symmetric] ring_distribs del: inverse_divide) + also have "\ = - 2 * ln 2" using sums_minus[OF alternating_harmonic_series_sums'] + by (subst suminf_mult) (simp_all add: algebra_simps sums_iff) + finally show ?case by simp +next + case (Suc n) + have nz: "2 * of_nat n + (1:: 'a) \ 0" + using of_nat_neq_0[of "2*n"] by (simp only: of_nat_Suc) (simp add: add_ac) + hence nz': "of_nat n + (1/2::'a) \ 0" by (simp add: field_simps) + have "Digamma (of_nat (Suc n) + 1/2 :: 'a) = Digamma (of_nat n + 1/2 + 1)" by simp + also from nz' have "\ = Digamma (of_nat n + 1 / 2) + 1 / (of_nat n + 1 / 2)" + by (rule Digamma_plus1) + also from nz nz' have "1 / (of_nat n + 1 / 2 :: 'a) = 2 / (2 * of_nat n + 1)" + by (subst divide_eq_eq) simp_all + also note Suc + finally show ?case by (simp add: add_ac) +qed + +lemma Digamma_one_half: "Digamma (1/2) = - euler_mascheroni - of_real (2 * ln 2)" + using Digamma_half_integer[of 0] by simp + +lemma Digamma_real_three_halves_pos: "Digamma (3/2 :: real) > 0" +proof - + have "-Digamma (3/2 :: real) = -Digamma (of_nat 1 + 1/2)" by simp + also have "\ = 2 * ln 2 + euler_mascheroni - 2" by (subst Digamma_half_integer) simp + also note euler_mascheroni_less_13_over_22 + also note ln2_le_25_over_36 + finally show ?thesis by simp +qed + + +lemma has_field_derivative_Polygamma [derivative_intros]: + fixes z :: "'a :: {real_normed_field,euclidean_space}" + assumes z: "z \ \\<^sub>\\<^sub>0" + shows "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z within A)" +proof (rule has_field_derivative_at_within, cases "n = 0") + assume n: "n = 0" + let ?f = "\k z. inverse (of_nat (Suc k)) - inverse (z + of_nat k)" + let ?F = "\z. \k. ?f k z" and ?f' = "\k z. inverse ((z + of_nat k)\<^sup>2)" + from no_nonpos_Int_in_ball'[OF z] guess d . note d = this + from z have summable: "summable (\k. inverse (of_nat (Suc k)) - inverse (z + of_nat k))" + by (intro summable_Digamma) force + from z have conv: "uniformly_convergent_on (ball z d) (\k z. \i2))" + by (intro Polygamma_converges) auto + with d have "summable (\k. inverse ((z + of_nat k)\<^sup>2))" unfolding summable_iff_convergent + by (auto dest!: uniformly_convergent_imp_convergent simp: summable_iff_convergent ) + + have "(?F has_field_derivative (\k. ?f' k z)) (at z)" + proof (rule has_field_derivative_series'[of "ball z d" _ _ z]) + fix k :: nat and t :: 'a assume t: "t \ ball z d" + from t d(2)[of t] show "((\z. ?f k z) has_field_derivative ?f' k t) (at t within ball z d)" + by (auto intro!: derivative_eq_intros simp: power2_eq_square simp del: of_nat_Suc + dest!: plus_of_nat_eq_0_imp elim!: nonpos_Ints_cases) + qed (insert d(1) summable conv, (assumption|simp)+) + with z show "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z)" + unfolding Digamma_def [abs_def] Polygamma_def [abs_def] using n + by (force simp: power2_eq_square intro!: derivative_eq_intros) +next + assume n: "n \ 0" + from z have z': "z \ 0" by auto + from no_nonpos_Int_in_ball'[OF z] guess d . note d = this + define n' where "n' = Suc n" + from n have n': "n' \ 2" by (simp add: n'_def) + have "((\z. \k. inverse ((z + of_nat k) ^ n')) has_field_derivative + (\k. - of_nat n' * inverse ((z + of_nat k) ^ (n'+1)))) (at z)" + proof (rule has_field_derivative_series'[of "ball z d" _ _ z]) + fix k :: nat and t :: 'a assume t: "t \ ball z d" + with d have t': "t \ \\<^sub>\\<^sub>0" "t \ 0" by auto + show "((\a. inverse ((a + of_nat k) ^ n')) has_field_derivative + - of_nat n' * inverse ((t + of_nat k) ^ (n'+1))) (at t within ball z d)" using t' + by (fastforce intro!: derivative_eq_intros simp: divide_simps power_diff dest: plus_of_nat_eq_0_imp) + next + have "uniformly_convergent_on (ball z d) + (\k z. (- of_nat n' :: 'a) * (\ik z. \ik. - of_nat n' * inverse ((z + of_nat k) ^ (n' + 1))) = + (- of_nat n') * (\k. inverse ((z + of_nat k) ^ (n' + 1)))" + using Polygamma_converges'[OF z', of "n'+1"] n' by (subst suminf_mult) simp_all + finally have "((\z. \k. inverse ((z + of_nat k) ^ n')) has_field_derivative + - of_nat n' * (\k. inverse ((z + of_nat k) ^ (n' + 1)))) (at z)" . + from DERIV_cmult[OF this, of "(-1)^Suc n * fact n :: 'a"] + show "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z)" + unfolding n'_def Polygamma_def[abs_def] using n by (simp add: algebra_simps) +qed + +declare has_field_derivative_Polygamma[THEN DERIV_chain2, derivative_intros] + +lemma isCont_Polygamma [continuous_intros]: + fixes f :: "_ \ 'a :: {real_normed_field,euclidean_space}" + shows "isCont f z \ f z \ \\<^sub>\\<^sub>0 \ isCont (\x. Polygamma n (f x)) z" + by (rule isCont_o2[OF _ DERIV_isCont[OF has_field_derivative_Polygamma]]) + +lemma continuous_on_Polygamma: + "A \ \\<^sub>\\<^sub>0 = {} \ continuous_on A (Polygamma n :: _ \ 'a :: {real_normed_field,euclidean_space})" + by (intro continuous_at_imp_continuous_on isCont_Polygamma[OF continuous_ident] ballI) blast + +lemma isCont_ln_Gamma_complex [continuous_intros]: + fixes f :: "'a::t2_space \ complex" + shows "isCont f z \ f z \ \\<^sub>\\<^sub>0 \ isCont (\z. ln_Gamma (f z)) z" + by (rule isCont_o2[OF _ DERIV_isCont[OF has_field_derivative_ln_Gamma_complex]]) + +lemma continuous_on_ln_Gamma_complex [continuous_intros]: + fixes A :: "complex set" + shows "A \ \\<^sub>\\<^sub>0 = {} \ continuous_on A ln_Gamma" + by (intro continuous_at_imp_continuous_on ballI isCont_ln_Gamma_complex[OF continuous_ident]) + fastforce + +text \ + We define a type class that captures all the fundamental properties of the inverse of the Gamma function + and defines the Gamma function upon that. This allows us to instantiate the type class both for + the reals and for the complex numbers with a minimal amount of proof duplication. +\ + +class Gamma = real_normed_field + complete_space + + fixes rGamma :: "'a \ 'a" + assumes rGamma_eq_zero_iff_aux: "rGamma z = 0 \ (\n. z = - of_nat n)" + assumes differentiable_rGamma_aux1: + "(\n. z \ - of_nat n) \ + let d = (THE d. (\n. \k d) - scaleR euler_mascheroni 1 + in filterlim (\y. (rGamma y - rGamma z + rGamma z * d * (y - z)) /\<^sub>R + norm (y - z)) (nhds 0) (at z)" + assumes differentiable_rGamma_aux2: + "let z = - of_nat n + in filterlim (\y. (rGamma y - rGamma z - (-1)^n * (setprod of_nat {1..n}) * (y - z)) /\<^sub>R + norm (y - z)) (nhds 0) (at z)" + assumes rGamma_series_aux: "(\n. z \ - of_nat n) \ + let fact' = (\n. setprod of_nat {1..n}); + exp = (\x. THE e. (\n. \kR fact k) \ e); + pochhammer' = (\a n. (\n = 0..n. a + of_nat n)) + in filterlim (\n. pochhammer' z n / (fact' n * exp (z * (ln (of_nat n) *\<^sub>R 1)))) + (nhds (rGamma z)) sequentially" +begin +subclass banach .. +end + +definition "Gamma z = inverse (rGamma z)" + + +subsection \Basic properties\ + +lemma Gamma_nonpos_Int: "z \ \\<^sub>\\<^sub>0 \ Gamma z = 0" + and rGamma_nonpos_Int: "z \ \\<^sub>\\<^sub>0 \ rGamma z = 0" + using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases') + +lemma Gamma_nonzero: "z \ \\<^sub>\\<^sub>0 \ Gamma z \ 0" + and rGamma_nonzero: "z \ \\<^sub>\\<^sub>0 \ rGamma z \ 0" + using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases') + +lemma Gamma_eq_zero_iff: "Gamma z = 0 \ z \ \\<^sub>\\<^sub>0" + and rGamma_eq_zero_iff: "rGamma z = 0 \ z \ \\<^sub>\\<^sub>0" + using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases') + +lemma rGamma_inverse_Gamma: "rGamma z = inverse (Gamma z)" + unfolding Gamma_def by simp + +lemma rGamma_series_LIMSEQ [tendsto_intros]: + "rGamma_series z \ rGamma z" +proof (cases "z \ \\<^sub>\\<^sub>0") + case False + hence "z \ - of_nat n" for n by auto + from rGamma_series_aux[OF this] show ?thesis + by (simp add: rGamma_series_def[abs_def] fact_setprod pochhammer_Suc_setprod + exp_def of_real_def[symmetric] suminf_def sums_def[abs_def] atLeast0AtMost) +qed (insert rGamma_eq_zero_iff[of z], simp_all add: rGamma_series_nonpos_Ints_LIMSEQ) + +lemma Gamma_series_LIMSEQ [tendsto_intros]: + "Gamma_series z \ Gamma z" +proof (cases "z \ \\<^sub>\\<^sub>0") + case False + hence "(\n. inverse (rGamma_series z n)) \ inverse (rGamma z)" + by (intro tendsto_intros) (simp_all add: rGamma_eq_zero_iff) + also have "(\n. inverse (rGamma_series z n)) = Gamma_series z" + by (simp add: rGamma_series_def Gamma_series_def[abs_def]) + finally show ?thesis by (simp add: Gamma_def) +qed (insert Gamma_eq_zero_iff[of z], simp_all add: Gamma_series_nonpos_Ints_LIMSEQ) + +lemma Gamma_altdef: "Gamma z = lim (Gamma_series z)" + using Gamma_series_LIMSEQ[of z] by (simp add: limI) + +lemma rGamma_1 [simp]: "rGamma 1 = 1" +proof - + have A: "eventually (\n. rGamma_series 1 n = of_nat (Suc n) / of_nat n) sequentially" + using eventually_gt_at_top[of "0::nat"] + by (force elim!: eventually_mono simp: rGamma_series_def exp_of_real pochhammer_fact + divide_simps pochhammer_rec' dest!: pochhammer_eq_0_imp_nonpos_Int) + have "rGamma_series 1 \ 1" by (subst tendsto_cong[OF A]) (rule LIMSEQ_Suc_n_over_n) + moreover have "rGamma_series 1 \ rGamma 1" by (rule tendsto_intros) + ultimately show ?thesis by (intro LIMSEQ_unique) +qed + +lemma rGamma_plus1: "z * rGamma (z + 1) = rGamma z" +proof - + let ?f = "\n. (z + 1) * inverse (of_nat n) + 1" + have "eventually (\n. ?f n * rGamma_series z n = z * rGamma_series (z + 1) n) sequentially" + using eventually_gt_at_top[of "0::nat"] + proof eventually_elim + fix n :: nat assume n: "n > 0" + hence "z * rGamma_series (z + 1) n = inverse (of_nat n) * + pochhammer z (Suc (Suc n)) / (fact n * exp (z * of_real (ln (of_nat n))))" + by (subst pochhammer_rec) (simp add: rGamma_series_def field_simps exp_add exp_of_real) + also from n have "\ = ?f n * rGamma_series z n" + by (subst pochhammer_rec') (simp_all add: divide_simps rGamma_series_def add_ac) + finally show "?f n * rGamma_series z n = z * rGamma_series (z + 1) n" .. + qed + moreover have "(\n. ?f n * rGamma_series z n) \ ((z+1) * 0 + 1) * rGamma z" + by (intro tendsto_intros lim_inverse_n) + hence "(\n. ?f n * rGamma_series z n) \ rGamma z" by simp + ultimately have "(\n. z * rGamma_series (z + 1) n) \ rGamma z" + by (rule Lim_transform_eventually) + moreover have "(\n. z * rGamma_series (z + 1) n) \ z * rGamma (z + 1)" + by (intro tendsto_intros) + ultimately show "z * rGamma (z + 1) = rGamma z" using LIMSEQ_unique by blast +qed + + +lemma pochhammer_rGamma: "rGamma z = pochhammer z n * rGamma (z + of_nat n)" +proof (induction n arbitrary: z) + case (Suc n z) + have "rGamma z = pochhammer z n * rGamma (z + of_nat n)" by (rule Suc.IH) + also note rGamma_plus1 [symmetric] + finally show ?case by (simp add: add_ac pochhammer_rec') +qed simp_all + +lemma Gamma_plus1: "z \ \\<^sub>\\<^sub>0 \ Gamma (z + 1) = z * Gamma z" + using rGamma_plus1[of z] by (simp add: rGamma_inverse_Gamma field_simps Gamma_eq_zero_iff) + +lemma pochhammer_Gamma: "z \ \\<^sub>\\<^sub>0 \ pochhammer z n = Gamma (z + of_nat n) / Gamma z" + using pochhammer_rGamma[of z] + by (simp add: rGamma_inverse_Gamma Gamma_eq_zero_iff field_simps) + +lemma Gamma_0 [simp]: "Gamma 0 = 0" + and rGamma_0 [simp]: "rGamma 0 = 0" + and Gamma_neg_1 [simp]: "Gamma (- 1) = 0" + and rGamma_neg_1 [simp]: "rGamma (- 1) = 0" + and Gamma_neg_numeral [simp]: "Gamma (- numeral n) = 0" + and rGamma_neg_numeral [simp]: "rGamma (- numeral n) = 0" + and Gamma_neg_of_nat [simp]: "Gamma (- of_nat m) = 0" + and rGamma_neg_of_nat [simp]: "rGamma (- of_nat m) = 0" + by (simp_all add: rGamma_eq_zero_iff Gamma_eq_zero_iff) + +lemma Gamma_1 [simp]: "Gamma 1 = 1" unfolding Gamma_def by simp + +lemma Gamma_fact: "Gamma (1 + of_nat n) = fact n" + by (simp add: pochhammer_fact pochhammer_Gamma of_nat_in_nonpos_Ints_iff + of_nat_Suc [symmetric] del: of_nat_Suc) + +lemma Gamma_numeral: "Gamma (numeral n) = fact (pred_numeral n)" + by (subst of_nat_numeral[symmetric], subst numeral_eq_Suc, + subst of_nat_Suc, subst Gamma_fact) (rule refl) + +lemma Gamma_of_int: "Gamma (of_int n) = (if n > 0 then fact (nat (n - 1)) else 0)" +proof (cases "n > 0") + case True + hence "Gamma (of_int n) = Gamma (of_nat (Suc (nat (n - 1))))" by (subst of_nat_Suc) simp_all + with True show ?thesis by (subst (asm) of_nat_Suc, subst (asm) Gamma_fact) simp +qed (simp_all add: Gamma_eq_zero_iff nonpos_Ints_of_int) + +lemma rGamma_of_int: "rGamma (of_int n) = (if n > 0 then inverse (fact (nat (n - 1))) else 0)" + by (simp add: Gamma_of_int rGamma_inverse_Gamma) + +lemma Gamma_seriesI: + assumes "(\n. g n / Gamma_series z n) \ 1" + shows "g \ Gamma z" +proof (rule Lim_transform_eventually) + have "1/2 > (0::real)" by simp + from tendstoD[OF assms, OF this] + show "eventually (\n. g n / Gamma_series z n * Gamma_series z n = g n) sequentially" + by (force elim!: eventually_mono simp: dist_real_def dist_0_norm) + from assms have "(\n. g n / Gamma_series z n * Gamma_series z n) \ 1 * Gamma z" + by (intro tendsto_intros) + thus "(\n. g n / Gamma_series z n * Gamma_series z n) \ Gamma z" by simp +qed + +lemma Gamma_seriesI': + assumes "f \ rGamma z" + assumes "(\n. g n * f n) \ 1" + assumes "z \ \\<^sub>\\<^sub>0" + shows "g \ Gamma z" +proof (rule Lim_transform_eventually) + have "1/2 > (0::real)" by simp + from tendstoD[OF assms(2), OF this] show "eventually (\n. g n * f n / f n = g n) sequentially" + by (force elim!: eventually_mono simp: dist_real_def dist_0_norm) + from assms have "(\n. g n * f n / f n) \ 1 / rGamma z" + by (intro tendsto_divide assms) (simp_all add: rGamma_eq_zero_iff) + thus "(\n. g n * f n / f n) \ Gamma z" by (simp add: Gamma_def divide_inverse) +qed + +lemma Gamma_series'_LIMSEQ: "Gamma_series' z \ Gamma z" + by (cases "z \ \\<^sub>\\<^sub>0") (simp_all add: Gamma_nonpos_Int Gamma_seriesI[OF Gamma_series_Gamma_series'] + Gamma_series'_nonpos_Ints_LIMSEQ[of z]) + + +subsection \Differentiability\ + +lemma has_field_derivative_rGamma_no_nonpos_int: + assumes "z \ \\<^sub>\\<^sub>0" + shows "(rGamma has_field_derivative -rGamma z * Digamma z) (at z within A)" +proof (rule has_field_derivative_at_within) + from assms have "z \ - of_nat n" for n by auto + from differentiable_rGamma_aux1[OF this] + show "(rGamma has_field_derivative -rGamma z * Digamma z) (at z)" + unfolding Digamma_def suminf_def sums_def[abs_def] + has_field_derivative_def has_derivative_def netlimit_at + by (simp add: Let_def bounded_linear_mult_right mult_ac of_real_def [symmetric]) +qed + +lemma has_field_derivative_rGamma_nonpos_int: + "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n) within A)" + apply (rule has_field_derivative_at_within) + using differentiable_rGamma_aux2[of n] + unfolding Let_def has_field_derivative_def has_derivative_def netlimit_at + by (simp only: bounded_linear_mult_right mult_ac of_real_def [symmetric] fact_setprod) simp + +lemma has_field_derivative_rGamma [derivative_intros]: + "(rGamma has_field_derivative (if z \ \\<^sub>\\<^sub>0 then (-1)^(nat \norm z\) * fact (nat \norm z\) + else -rGamma z * Digamma z)) (at z within A)" +using has_field_derivative_rGamma_no_nonpos_int[of z A] + has_field_derivative_rGamma_nonpos_int[of "nat \norm z\" A] + by (auto elim!: nonpos_Ints_cases') + +declare has_field_derivative_rGamma_no_nonpos_int [THEN DERIV_chain2, derivative_intros] +declare has_field_derivative_rGamma [THEN DERIV_chain2, derivative_intros] +declare has_field_derivative_rGamma_nonpos_int [derivative_intros] +declare has_field_derivative_rGamma_no_nonpos_int [derivative_intros] +declare has_field_derivative_rGamma [derivative_intros] + + +lemma has_field_derivative_Gamma [derivative_intros]: + "z \ \\<^sub>\\<^sub>0 \ (Gamma has_field_derivative Gamma z * Digamma z) (at z within A)" + unfolding Gamma_def [abs_def] + by (fastforce intro!: derivative_eq_intros simp: rGamma_eq_zero_iff) + +declare has_field_derivative_Gamma[THEN DERIV_chain2, derivative_intros] + +(* TODO: Hide ugly facts properly *) +hide_fact rGamma_eq_zero_iff_aux differentiable_rGamma_aux1 differentiable_rGamma_aux2 + differentiable_rGamma_aux2 rGamma_series_aux Gamma_class.rGamma_eq_zero_iff_aux + + + +(* TODO: differentiable etc. *) + + +subsection \Continuity\ + +lemma continuous_on_rGamma [continuous_intros]: "continuous_on A rGamma" + by (rule DERIV_continuous_on has_field_derivative_rGamma)+ + +lemma continuous_on_Gamma [continuous_intros]: "A \ \\<^sub>\\<^sub>0 = {} \ continuous_on A Gamma" + by (rule DERIV_continuous_on has_field_derivative_Gamma)+ blast + +lemma isCont_rGamma [continuous_intros]: + "isCont f z \ isCont (\x. rGamma (f x)) z" + by (rule isCont_o2[OF _ DERIV_isCont[OF has_field_derivative_rGamma]]) + +lemma isCont_Gamma [continuous_intros]: + "isCont f z \ f z \ \\<^sub>\\<^sub>0 \ isCont (\x. Gamma (f x)) z" + by (rule isCont_o2[OF _ DERIV_isCont[OF has_field_derivative_Gamma]]) + + + +text \The complex Gamma function\ + +instantiation complex :: Gamma +begin + +definition rGamma_complex :: "complex \ complex" where + "rGamma_complex z = lim (rGamma_series z)" + +lemma rGamma_series_complex_converges: + "convergent (rGamma_series (z :: complex))" (is "?thesis1") + and rGamma_complex_altdef: + "rGamma z = (if z \ \\<^sub>\\<^sub>0 then 0 else exp (-ln_Gamma z))" (is "?thesis2") +proof - + have "?thesis1 \ ?thesis2" + proof (cases "z \ \\<^sub>\\<^sub>0") + case False + have "rGamma_series z \ exp (- ln_Gamma z)" + proof (rule Lim_transform_eventually) + from ln_Gamma_series_complex_converges'[OF False] guess d by (elim exE conjE) + from this(1) uniformly_convergent_imp_convergent[OF this(2), of z] + have "ln_Gamma_series z \ lim (ln_Gamma_series z)" by (simp add: convergent_LIMSEQ_iff) + thus "(\n. exp (-ln_Gamma_series z n)) \ exp (- ln_Gamma z)" + unfolding convergent_def ln_Gamma_def by (intro tendsto_exp tendsto_minus) + from eventually_gt_at_top[of "0::nat"] exp_ln_Gamma_series_complex False + show "eventually (\n. exp (-ln_Gamma_series z n) = rGamma_series z n) sequentially" + by (force elim!: eventually_mono simp: exp_minus Gamma_series_def rGamma_series_def) + qed + with False show ?thesis + by (auto simp: convergent_def rGamma_complex_def intro!: limI) + next + case True + then obtain k where "z = - of_nat k" by (erule nonpos_Ints_cases') + also have "rGamma_series \ \ 0" + by (subst tendsto_cong[OF rGamma_series_minus_of_nat]) (simp_all add: convergent_const) + finally show ?thesis using True + by (auto simp: rGamma_complex_def convergent_def intro!: limI) + qed + thus "?thesis1" "?thesis2" by blast+ +qed + +context +begin + +(* TODO: duplication *) +private lemma rGamma_complex_plus1: "z * rGamma (z + 1) = rGamma (z :: complex)" +proof - + let ?f = "\n. (z + 1) * inverse (of_nat n) + 1" + have "eventually (\n. ?f n * rGamma_series z n = z * rGamma_series (z + 1) n) sequentially" + using eventually_gt_at_top[of "0::nat"] + proof eventually_elim + fix n :: nat assume n: "n > 0" + hence "z * rGamma_series (z + 1) n = inverse (of_nat n) * + pochhammer z (Suc (Suc n)) / (fact n * exp (z * of_real (ln (of_nat n))))" + by (subst pochhammer_rec) (simp add: rGamma_series_def field_simps exp_add exp_of_real) + also from n have "\ = ?f n * rGamma_series z n" + by (subst pochhammer_rec') (simp_all add: divide_simps rGamma_series_def add_ac) + finally show "?f n * rGamma_series z n = z * rGamma_series (z + 1) n" .. + qed + moreover have "(\n. ?f n * rGamma_series z n) \ ((z+1) * 0 + 1) * rGamma z" + using rGamma_series_complex_converges + by (intro tendsto_intros lim_inverse_n) + (simp_all add: convergent_LIMSEQ_iff rGamma_complex_def) + hence "(\n. ?f n * rGamma_series z n) \ rGamma z" by simp + ultimately have "(\n. z * rGamma_series (z + 1) n) \ rGamma z" + by (rule Lim_transform_eventually) + moreover have "(\n. z * rGamma_series (z + 1) n) \ z * rGamma (z + 1)" + using rGamma_series_complex_converges + by (auto intro!: tendsto_mult simp: rGamma_complex_def convergent_LIMSEQ_iff) + ultimately show "z * rGamma (z + 1) = rGamma z" using LIMSEQ_unique by blast +qed + +private lemma has_field_derivative_rGamma_complex_no_nonpos_Int: + assumes "(z :: complex) \ \\<^sub>\\<^sub>0" + shows "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)" +proof - + have diff: "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)" if "Re z > 0" for z + proof (subst DERIV_cong_ev[OF refl _ refl]) + from that have "eventually (\t. t \ ball z (Re z/2)) (nhds z)" + by (intro eventually_nhds_in_nhd) simp_all + thus "eventually (\t. rGamma t = exp (- ln_Gamma t)) (nhds z)" + using no_nonpos_Int_in_ball_complex[OF that] + by (auto elim!: eventually_mono simp: rGamma_complex_altdef) + next + have "z \ \\<^sub>\\<^sub>0" using that by (simp add: complex_nonpos_Reals_iff) + with that show "((\t. exp (- ln_Gamma t)) has_field_derivative (-rGamma z * Digamma z)) (at z)" + by (force elim!: nonpos_Ints_cases intro!: derivative_eq_intros simp: rGamma_complex_altdef) + qed + + from assms show "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)" + proof (induction "nat \1 - Re z\" arbitrary: z) + case (Suc n z) + from Suc.prems have z: "z \ 0" by auto + from Suc.hyps have "n = nat \- Re z\" by linarith + hence A: "n = nat \1 - Re (z + 1)\" by simp + from Suc.prems have B: "z + 1 \ \\<^sub>\\<^sub>0" by (force dest: plus_one_in_nonpos_Ints_imp) + + have "((\z. z * (rGamma \ (\z. z + 1)) z) has_field_derivative + -rGamma (z + 1) * (Digamma (z + 1) * z - 1)) (at z)" + by (rule derivative_eq_intros DERIV_chain Suc refl A B)+ (simp add: algebra_simps) + also have "(\z. z * (rGamma \ (\z. z + 1 :: complex)) z) = rGamma" + by (simp add: rGamma_complex_plus1) + also from z have "Digamma (z + 1) * z - 1 = z * Digamma z" + by (subst Digamma_plus1) (simp_all add: field_simps) + also have "-rGamma (z + 1) * (z * Digamma z) = -rGamma z * Digamma z" + by (simp add: rGamma_complex_plus1[of z, symmetric]) + finally show ?case . + qed (intro diff, simp) +qed + +private lemma rGamma_complex_1: "rGamma (1 :: complex) = 1" +proof - + have A: "eventually (\n. rGamma_series 1 n = of_nat (Suc n) / of_nat n) sequentially" + using eventually_gt_at_top[of "0::nat"] + by (force elim!: eventually_mono simp: rGamma_series_def exp_of_real pochhammer_fact + divide_simps pochhammer_rec' dest!: pochhammer_eq_0_imp_nonpos_Int) + have "rGamma_series 1 \ 1" by (subst tendsto_cong[OF A]) (rule LIMSEQ_Suc_n_over_n) + thus "rGamma 1 = (1 :: complex)" unfolding rGamma_complex_def by (rule limI) +qed + +private lemma has_field_derivative_rGamma_complex_nonpos_Int: + "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n :: complex))" +proof (induction n) + case 0 + have A: "(0::complex) + 1 \ \\<^sub>\\<^sub>0" by simp + have "((\z. z * (rGamma \ (\z. z + 1 :: complex)) z) has_field_derivative 1) (at 0)" + by (rule derivative_eq_intros DERIV_chain refl + has_field_derivative_rGamma_complex_no_nonpos_Int A)+ (simp add: rGamma_complex_1) + thus ?case by (simp add: rGamma_complex_plus1) +next + case (Suc n) + hence A: "(rGamma has_field_derivative (-1)^n * fact n) + (at (- of_nat (Suc n) + 1 :: complex))" by simp + have "((\z. z * (rGamma \ (\z. z + 1 :: complex)) z) has_field_derivative + (- 1) ^ Suc n * fact (Suc n)) (at (- of_nat (Suc n)))" + by (rule derivative_eq_intros refl A DERIV_chain)+ + (simp add: algebra_simps rGamma_complex_altdef) + thus ?case by (simp add: rGamma_complex_plus1) +qed + +instance proof + fix z :: complex show "(rGamma z = 0) \ (\n. z = - of_nat n)" + by (auto simp: rGamma_complex_altdef elim!: nonpos_Ints_cases') +next + fix z :: complex assume "\n. z \ - of_nat n" + hence "z \ \\<^sub>\\<^sub>0" by (auto elim!: nonpos_Ints_cases') + from has_field_derivative_rGamma_complex_no_nonpos_Int[OF this] + show "let d = (THE d. (\n. \k d) - euler_mascheroni *\<^sub>R 1 in (\y. (rGamma y - rGamma z + + rGamma z * d * (y - z)) /\<^sub>R cmod (y - z)) \z\ 0" + by (simp add: has_field_derivative_def has_derivative_def Digamma_def sums_def [abs_def] + netlimit_at of_real_def[symmetric] suminf_def) +next + fix n :: nat + from has_field_derivative_rGamma_complex_nonpos_Int[of n] + show "let z = - of_nat n in (\y. (rGamma y - rGamma z - (- 1) ^ n * setprod of_nat {1..n} * + (y - z)) /\<^sub>R cmod (y - z)) \z\ 0" + by (simp add: has_field_derivative_def has_derivative_def fact_setprod netlimit_at Let_def) +next + fix z :: complex + from rGamma_series_complex_converges[of z] have "rGamma_series z \ rGamma z" + by (simp add: convergent_LIMSEQ_iff rGamma_complex_def) + thus "let fact' = \n. setprod of_nat {1..n}; + exp = \x. THE e. (\n. \kR fact k) \ e; + pochhammer' = \a n. \n = 0..n. a + of_nat n + in (\n. pochhammer' z n / (fact' n * exp (z * ln (real_of_nat n) *\<^sub>R 1))) \ rGamma z" + by (simp add: fact_setprod pochhammer_Suc_setprod rGamma_series_def [abs_def] exp_def + of_real_def [symmetric] suminf_def sums_def [abs_def] atLeast0AtMost) +qed + +end +end + + +lemma Gamma_complex_altdef: + "Gamma z = (if z \ \\<^sub>\\<^sub>0 then 0 else exp (ln_Gamma (z :: complex)))" + unfolding Gamma_def rGamma_complex_altdef by (simp add: exp_minus) + +lemma cnj_rGamma: "cnj (rGamma z) = rGamma (cnj z)" +proof - + have "rGamma_series (cnj z) = (\n. cnj (rGamma_series z n))" + by (intro ext) (simp_all add: rGamma_series_def exp_cnj) + also have "... \ cnj (rGamma z)" by (intro tendsto_cnj tendsto_intros) + finally show ?thesis unfolding rGamma_complex_def by (intro sym[OF limI]) +qed + +lemma cnj_Gamma: "cnj (Gamma z) = Gamma (cnj z)" + unfolding Gamma_def by (simp add: cnj_rGamma) + +lemma Gamma_complex_real: + "z \ \ \ Gamma z \ (\ :: complex set)" and rGamma_complex_real: "z \ \ \ rGamma z \ \" + by (simp_all add: Reals_cnj_iff cnj_Gamma cnj_rGamma) + +lemma field_differentiable_rGamma: "rGamma field_differentiable (at z within A)" + using has_field_derivative_rGamma[of z] unfolding field_differentiable_def by blast + +lemma holomorphic_on_rGamma: "rGamma holomorphic_on A" + unfolding holomorphic_on_def by (auto intro!: field_differentiable_rGamma) + +lemma analytic_on_rGamma: "rGamma analytic_on A" + unfolding analytic_on_def by (auto intro!: exI[of _ 1] holomorphic_on_rGamma) + + +lemma field_differentiable_Gamma: "z \ \\<^sub>\\<^sub>0 \ Gamma field_differentiable (at z within A)" + using has_field_derivative_Gamma[of z] unfolding field_differentiable_def by auto + +lemma holomorphic_on_Gamma: "A \ \\<^sub>\\<^sub>0 = {} \ Gamma holomorphic_on A" + unfolding holomorphic_on_def by (auto intro!: field_differentiable_Gamma) + +lemma analytic_on_Gamma: "A \ \\<^sub>\\<^sub>0 = {} \ Gamma analytic_on A" + by (rule analytic_on_subset[of _ "UNIV - \\<^sub>\\<^sub>0"], subst analytic_on_open) + (auto intro!: holomorphic_on_Gamma) + +lemma has_field_derivative_rGamma_complex' [derivative_intros]: + "(rGamma has_field_derivative (if z \ \\<^sub>\\<^sub>0 then (-1)^(nat \-Re z\) * fact (nat \-Re z\) else + -rGamma z * Digamma z)) (at z within A)" + using has_field_derivative_rGamma[of z] by (auto elim!: nonpos_Ints_cases') + +declare has_field_derivative_rGamma_complex'[THEN DERIV_chain2, derivative_intros] + + +lemma field_differentiable_Polygamma: + fixes z::complex + shows + "z \ \\<^sub>\\<^sub>0 \ Polygamma n field_differentiable (at z within A)" + using has_field_derivative_Polygamma[of z n] unfolding field_differentiable_def by auto + +lemma holomorphic_on_Polygamma: "A \ \\<^sub>\\<^sub>0 = {} \ Polygamma n holomorphic_on A" + unfolding holomorphic_on_def by (auto intro!: field_differentiable_Polygamma) + +lemma analytic_on_Polygamma: "A \ \\<^sub>\\<^sub>0 = {} \ Polygamma n analytic_on A" + by (rule analytic_on_subset[of _ "UNIV - \\<^sub>\\<^sub>0"], subst analytic_on_open) + (auto intro!: holomorphic_on_Polygamma) + + + +text \The real Gamma function\ + +lemma rGamma_series_real: + "eventually (\n. rGamma_series x n = Re (rGamma_series (of_real x) n)) sequentially" + using eventually_gt_at_top[of "0 :: nat"] +proof eventually_elim + fix n :: nat assume n: "n > 0" + have "Re (rGamma_series (of_real x) n) = + Re (of_real (pochhammer x (Suc n)) / (fact n * exp (of_real (x * ln (real_of_nat n)))))" + using n by (simp add: rGamma_series_def powr_def Ln_of_nat pochhammer_of_real) + also from n have "\ = Re (of_real ((pochhammer x (Suc n)) / + (fact n * (exp (x * ln (real_of_nat n))))))" + by (subst exp_of_real) simp + also from n have "\ = rGamma_series x n" + by (subst Re_complex_of_real) (simp add: rGamma_series_def powr_def) + finally show "rGamma_series x n = Re (rGamma_series (of_real x) n)" .. +qed + +instantiation real :: Gamma +begin + +definition "rGamma_real x = Re (rGamma (of_real x :: complex))" + +instance proof + fix x :: real + have "rGamma x = Re (rGamma (of_real x))" by (simp add: rGamma_real_def) + also have "of_real \ = rGamma (of_real x :: complex)" + by (intro of_real_Re rGamma_complex_real) simp_all + also have "\ = 0 \ x \ \\<^sub>\\<^sub>0" by (simp add: rGamma_eq_zero_iff of_real_in_nonpos_Ints_iff) + also have "\ \ (\n. x = - of_nat n)" by (auto elim!: nonpos_Ints_cases') + finally show "(rGamma x) = 0 \ (\n. x = - real_of_nat n)" by simp +next + fix x :: real assume "\n. x \ - of_nat n" + hence x: "complex_of_real x \ \\<^sub>\\<^sub>0" + by (subst of_real_in_nonpos_Ints_iff) (auto elim!: nonpos_Ints_cases') + then have "x \ 0" by auto + with x have "(rGamma has_field_derivative - rGamma x * Digamma x) (at x)" + by (fastforce intro!: derivative_eq_intros has_vector_derivative_real_complex + simp: Polygamma_of_real rGamma_real_def [abs_def]) + thus "let d = (THE d. (\n. \k d) - euler_mascheroni *\<^sub>R 1 in (\y. (rGamma y - rGamma x + + rGamma x * d * (y - x)) /\<^sub>R norm (y - x)) \x\ 0" + by (simp add: has_field_derivative_def has_derivative_def Digamma_def sums_def [abs_def] + netlimit_at of_real_def[symmetric] suminf_def) +next + fix n :: nat + have "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n :: real))" + by (fastforce intro!: derivative_eq_intros has_vector_derivative_real_complex + simp: Polygamma_of_real rGamma_real_def [abs_def]) + thus "let x = - of_nat n in (\y. (rGamma y - rGamma x - (- 1) ^ n * setprod of_nat {1..n} * + (y - x)) /\<^sub>R norm (y - x)) \x::real\ 0" + by (simp add: has_field_derivative_def has_derivative_def fact_setprod netlimit_at Let_def) +next + fix x :: real + have "rGamma_series x \ rGamma x" + proof (rule Lim_transform_eventually) + show "(\n. Re (rGamma_series (of_real x) n)) \ rGamma x" unfolding rGamma_real_def + by (intro tendsto_intros) + qed (insert rGamma_series_real, simp add: eq_commute) + thus "let fact' = \n. setprod of_nat {1..n}; + exp = \x. THE e. (\n. \kR fact k) \ e; + pochhammer' = \a n. \n = 0..n. a + of_nat n + in (\n. pochhammer' x n / (fact' n * exp (x * ln (real_of_nat n) *\<^sub>R 1))) \ rGamma x" + by (simp add: fact_setprod pochhammer_Suc_setprod rGamma_series_def [abs_def] exp_def + of_real_def [symmetric] suminf_def sums_def [abs_def] atLeast0AtMost) +qed + +end + + +lemma rGamma_complex_of_real: "rGamma (complex_of_real x) = complex_of_real (rGamma x)" + unfolding rGamma_real_def using rGamma_complex_real by simp + +lemma Gamma_complex_of_real: "Gamma (complex_of_real x) = complex_of_real (Gamma x)" + unfolding Gamma_def by (simp add: rGamma_complex_of_real) + +lemma rGamma_real_altdef: "rGamma x = lim (rGamma_series (x :: real))" + by (rule sym, rule limI, rule tendsto_intros) + +lemma Gamma_real_altdef1: "Gamma x = lim (Gamma_series (x :: real))" + by (rule sym, rule limI, rule tendsto_intros) + +lemma Gamma_real_altdef2: "Gamma x = Re (Gamma (of_real x))" + using rGamma_complex_real[OF Reals_of_real[of x]] + by (elim Reals_cases) + (simp only: Gamma_def rGamma_real_def of_real_inverse[symmetric] Re_complex_of_real) + +lemma ln_Gamma_series_complex_of_real: + "x > 0 \ n > 0 \ ln_Gamma_series (complex_of_real x) n = of_real (ln_Gamma_series x n)" +proof - + assume xn: "x > 0" "n > 0" + have "Ln (complex_of_real x / of_nat k + 1) = of_real (ln (x / of_nat k + 1))" if "k \ 1" for k + using that xn by (subst Ln_of_real [symmetric]) (auto intro!: add_nonneg_pos simp: field_simps) + with xn show ?thesis by (simp add: ln_Gamma_series_def Ln_of_nat Ln_of_real) +qed + +lemma ln_Gamma_real_converges: + assumes "(x::real) > 0" + shows "convergent (ln_Gamma_series x)" +proof - + have "(\n. ln_Gamma_series (complex_of_real x) n) \ ln_Gamma (of_real x)" using assms + by (intro ln_Gamma_complex_LIMSEQ) (auto simp: of_real_in_nonpos_Ints_iff) + moreover from eventually_gt_at_top[of "0::nat"] + have "eventually (\n. complex_of_real (ln_Gamma_series x n) = + ln_Gamma_series (complex_of_real x) n) sequentially" + by eventually_elim (simp add: ln_Gamma_series_complex_of_real assms) + ultimately have "(\n. complex_of_real (ln_Gamma_series x n)) \ ln_Gamma (of_real x)" + by (subst tendsto_cong) assumption+ + from tendsto_Re[OF this] show ?thesis by (auto simp: convergent_def) +qed + +lemma ln_Gamma_real_LIMSEQ: "(x::real) > 0 \ ln_Gamma_series x \ ln_Gamma x" + using ln_Gamma_real_converges[of x] unfolding ln_Gamma_def by (simp add: convergent_LIMSEQ_iff) + +lemma ln_Gamma_complex_of_real: "x > 0 \ ln_Gamma (complex_of_real x) = of_real (ln_Gamma x)" +proof (unfold ln_Gamma_def, rule limI, rule Lim_transform_eventually) + assume x: "x > 0" + show "eventually (\n. of_real (ln_Gamma_series x n) = + ln_Gamma_series (complex_of_real x) n) sequentially" + using eventually_gt_at_top[of "0::nat"] + by eventually_elim (simp add: ln_Gamma_series_complex_of_real x) +qed (intro tendsto_of_real, insert ln_Gamma_real_LIMSEQ[of x], simp add: ln_Gamma_def) + +lemma Gamma_real_pos_exp: "x > (0 :: real) \ Gamma x = exp (ln_Gamma x)" + by (auto simp: Gamma_real_altdef2 Gamma_complex_altdef of_real_in_nonpos_Ints_iff + ln_Gamma_complex_of_real exp_of_real) + +lemma ln_Gamma_real_pos: "x > 0 \ ln_Gamma x = ln (Gamma x :: real)" + unfolding Gamma_real_pos_exp by simp + +lemma Gamma_real_pos: "x > (0::real) \ Gamma x > 0" + by (simp add: Gamma_real_pos_exp) + +lemma has_field_derivative_ln_Gamma_real [derivative_intros]: + assumes x: "x > (0::real)" + shows "(ln_Gamma has_field_derivative Digamma x) (at x)" +proof (subst DERIV_cong_ev[OF refl _ refl]) + from assms show "((Re \ ln_Gamma \ complex_of_real) has_field_derivative Digamma x) (at x)" + by (auto intro!: derivative_eq_intros has_vector_derivative_real_complex + simp: Polygamma_of_real o_def) + from eventually_nhds_in_nhd[of x "{0<..}"] assms + show "eventually (\y. ln_Gamma y = (Re \ ln_Gamma \ of_real) y) (nhds x)" + by (auto elim!: eventually_mono simp: ln_Gamma_complex_of_real interior_open) +qed + +declare has_field_derivative_ln_Gamma_real[THEN DERIV_chain2, derivative_intros] + + +lemma has_field_derivative_rGamma_real' [derivative_intros]: + "(rGamma has_field_derivative (if x \ \\<^sub>\\<^sub>0 then (-1)^(nat \-x\) * fact (nat \-x\) else + -rGamma x * Digamma x)) (at x within A)" + using has_field_derivative_rGamma[of x] by (force elim!: nonpos_Ints_cases') + +declare has_field_derivative_rGamma_real'[THEN DERIV_chain2, derivative_intros] + +lemma Polygamma_real_odd_pos: + assumes "(x::real) \ \\<^sub>\\<^sub>0" "odd n" + shows "Polygamma n x > 0" +proof - + from assms have "x \ 0" by auto + with assms show ?thesis + unfolding Polygamma_def using Polygamma_converges'[of x "Suc n"] + by (auto simp: zero_less_power_eq simp del: power_Suc + dest: plus_of_nat_eq_0_imp intro!: mult_pos_pos suminf_pos) +qed + +lemma Polygamma_real_even_neg: + assumes "(x::real) > 0" "n > 0" "even n" + shows "Polygamma n x < 0" + using assms unfolding Polygamma_def using Polygamma_converges'[of x "Suc n"] + by (auto intro!: mult_pos_pos suminf_pos) + +lemma Polygamma_real_strict_mono: + assumes "x > 0" "x < (y::real)" "even n" + shows "Polygamma n x < Polygamma n y" +proof - + have "\\. x < \ \ \ < y \ Polygamma n y - Polygamma n x = (y - x) * Polygamma (Suc n) \" + using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases) + then guess \ by (elim exE conjE) note \ = this + note \(3) + also from \(1,2) assms have "(y - x) * Polygamma (Suc n) \ > 0" + by (intro mult_pos_pos Polygamma_real_odd_pos) (auto elim!: nonpos_Ints_cases) + finally show ?thesis by simp +qed + +lemma Polygamma_real_strict_antimono: + assumes "x > 0" "x < (y::real)" "odd n" + shows "Polygamma n x > Polygamma n y" +proof - + have "\\. x < \ \ \ < y \ Polygamma n y - Polygamma n x = (y - x) * Polygamma (Suc n) \" + using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases) + then guess \ by (elim exE conjE) note \ = this + note \(3) + also from \(1,2) assms have "(y - x) * Polygamma (Suc n) \ < 0" + by (intro mult_pos_neg Polygamma_real_even_neg) simp_all + finally show ?thesis by simp +qed + +lemma Polygamma_real_mono: + assumes "x > 0" "x \ (y::real)" "even n" + shows "Polygamma n x \ Polygamma n y" + using Polygamma_real_strict_mono[OF assms(1) _ assms(3), of y] assms(2) + by (cases "x = y") simp_all + +lemma Digamma_real_ge_three_halves_pos: + assumes "x \ 3/2" + shows "Digamma (x :: real) > 0" +proof - + have "0 < Digamma (3/2 :: real)" by (fact Digamma_real_three_halves_pos) + also from assms have "\ \ Digamma x" by (intro Polygamma_real_mono) simp_all + finally show ?thesis . +qed + +lemma ln_Gamma_real_strict_mono: + assumes "x \ 3/2" "x < y" + shows "ln_Gamma (x :: real) < ln_Gamma y" +proof - + have "\\. x < \ \ \ < y \ ln_Gamma y - ln_Gamma x = (y - x) * Digamma \" + using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases) + then guess \ by (elim exE conjE) note \ = this + note \(3) + also from \(1,2) assms have "(y - x) * Digamma \ > 0" + by (intro mult_pos_pos Digamma_real_ge_three_halves_pos) simp_all + finally show ?thesis by simp +qed + +lemma Gamma_real_strict_mono: + assumes "x \ 3/2" "x < y" + shows "Gamma (x :: real) < Gamma y" +proof - + from Gamma_real_pos_exp[of x] assms have "Gamma x = exp (ln_Gamma x)" by simp + also have "\ < exp (ln_Gamma y)" by (intro exp_less_mono ln_Gamma_real_strict_mono assms) + also from Gamma_real_pos_exp[of y] assms have "\ = Gamma y" by simp + finally show ?thesis . +qed + +lemma log_convex_Gamma_real: "convex_on {0<..} (ln \ Gamma :: real \ real)" + by (rule convex_on_realI[of _ _ Digamma]) + (auto intro!: derivative_eq_intros Polygamma_real_mono Gamma_real_pos + simp: o_def Gamma_eq_zero_iff elim!: nonpos_Ints_cases') + + +subsection \Beta function\ + +definition Beta where "Beta a b = Gamma a * Gamma b / Gamma (a + b)" + +lemma Beta_altdef: "Beta a b = Gamma a * Gamma b * rGamma (a + b)" + by (simp add: inverse_eq_divide Beta_def Gamma_def) + +lemma Beta_commute: "Beta a b = Beta b a" + unfolding Beta_def by (simp add: ac_simps) + +lemma has_field_derivative_Beta1 [derivative_intros]: + assumes "x \ \\<^sub>\\<^sub>0" "x + y \ \\<^sub>\\<^sub>0" + shows "((\x. Beta x y) has_field_derivative (Beta x y * (Digamma x - Digamma (x + y)))) + (at x within A)" unfolding Beta_altdef + by (rule DERIV_cong, (rule derivative_intros assms)+) (simp add: algebra_simps) + +lemma Beta_pole1: "x \ \\<^sub>\\<^sub>0 \ Beta x y = 0" + by (auto simp add: Beta_def elim!: nonpos_Ints_cases') + +lemma Beta_pole2: "y \ \\<^sub>\\<^sub>0 \ Beta x y = 0" + by (auto simp add: Beta_def elim!: nonpos_Ints_cases') + +lemma Beta_zero: "x + y \ \\<^sub>\\<^sub>0 \ Beta x y = 0" + by (auto simp add: Beta_def elim!: nonpos_Ints_cases') + +lemma has_field_derivative_Beta2 [derivative_intros]: + assumes "y \ \\<^sub>\\<^sub>0" "x + y \ \\<^sub>\\<^sub>0" + shows "((\y. Beta x y) has_field_derivative (Beta x y * (Digamma y - Digamma (x + y)))) + (at y within A)" + using has_field_derivative_Beta1[of y x A] assms by (simp add: Beta_commute add_ac) + +lemma Beta_plus1_plus1: + assumes "x \ \\<^sub>\\<^sub>0" "y \ \\<^sub>\\<^sub>0" + shows "Beta (x + 1) y + Beta x (y + 1) = Beta x y" +proof - + have "Beta (x + 1) y + Beta x (y + 1) = + (Gamma (x + 1) * Gamma y + Gamma x * Gamma (y + 1)) * rGamma ((x + y) + 1)" + by (simp add: Beta_altdef add_divide_distrib algebra_simps) + also have "\ = (Gamma x * Gamma y) * ((x + y) * rGamma ((x + y) + 1))" + by (subst assms[THEN Gamma_plus1])+ (simp add: algebra_simps) + also from assms have "\ = Beta x y" unfolding Beta_altdef by (subst rGamma_plus1) simp + finally show ?thesis . +qed + +lemma Beta_plus1_left: + assumes "x \ \\<^sub>\\<^sub>0" + shows "(x + y) * Beta (x + 1) y = x * Beta x y" +proof - + have "(x + y) * Beta (x + 1) y = Gamma (x + 1) * Gamma y * ((x + y) * rGamma ((x + y) + 1))" + unfolding Beta_altdef by (simp only: ac_simps) + also have "\ = x * Beta x y" unfolding Beta_altdef + by (subst assms[THEN Gamma_plus1] rGamma_plus1)+ (simp only: ac_simps) + finally show ?thesis . +qed + +lemma Beta_plus1_right: + assumes "y \ \\<^sub>\\<^sub>0" + shows "(x + y) * Beta x (y + 1) = y * Beta x y" + using Beta_plus1_left[of y x] assms by (simp_all add: Beta_commute add.commute) + +lemma Gamma_Gamma_Beta: + assumes "x + y \ \\<^sub>\\<^sub>0" + shows "Gamma x * Gamma y = Beta x y * Gamma (x + y)" + unfolding Beta_altdef using assms Gamma_eq_zero_iff[of "x+y"] + by (simp add: rGamma_inverse_Gamma) + + + +subsection \Legendre duplication theorem\ + +context +begin + +private lemma Gamma_legendre_duplication_aux: + fixes z :: "'a :: Gamma" + assumes "z \ \\<^sub>\\<^sub>0" "z + 1/2 \ \\<^sub>\\<^sub>0" + shows "Gamma z * Gamma (z + 1/2) = exp ((1 - 2*z) * of_real (ln 2)) * Gamma (1/2) * Gamma (2*z)" +proof - + let ?powr = "\b a. exp (a * of_real (ln (of_nat b)))" + let ?h = "\n. (fact (n-1))\<^sup>2 / fact (2*n-1) * of_nat (2^(2*n)) * + exp (1/2 * of_real (ln (real_of_nat n)))" + { + fix z :: 'a assume z: "z \ \\<^sub>\\<^sub>0" "z + 1/2 \ \\<^sub>\\<^sub>0" + let ?g = "\n. ?powr 2 (2*z) * Gamma_series' z n * Gamma_series' (z + 1/2) n / + Gamma_series' (2*z) (2*n)" + have "eventually (\n. ?g n = ?h n) sequentially" using eventually_gt_at_top + proof eventually_elim + fix n :: nat assume n: "n > 0" + let ?f = "fact (n - 1) :: 'a" and ?f' = "fact (2*n - 1) :: 'a" + have A: "exp t * exp t = exp (2*t :: 'a)" for t by (subst exp_add [symmetric]) simp + have A: "Gamma_series' z n * Gamma_series' (z + 1/2) n = ?f^2 * ?powr n (2*z + 1/2) / + (pochhammer z n * pochhammer (z + 1/2) n)" + by (simp add: Gamma_series'_def exp_add ring_distribs power2_eq_square A mult_ac) + have B: "Gamma_series' (2*z) (2*n) = + ?f' * ?powr 2 (2*z) * ?powr n (2*z) / + (of_nat (2^(2*n)) * pochhammer z n * pochhammer (z+1/2) n)" using n + by (simp add: Gamma_series'_def ln_mult exp_add ring_distribs pochhammer_double) + from z have "pochhammer z n \ 0" by (auto dest: pochhammer_eq_0_imp_nonpos_Int) + moreover from z have "pochhammer (z + 1/2) n \ 0" by (auto dest: pochhammer_eq_0_imp_nonpos_Int) + ultimately have "?powr 2 (2*z) * (Gamma_series' z n * Gamma_series' (z + 1/2) n) / Gamma_series' (2*z) (2*n) = + ?f^2 / ?f' * of_nat (2^(2*n)) * (?powr n ((4*z + 1)/2) * ?powr n (-2*z))" + using n unfolding A B by (simp add: divide_simps exp_minus) + also have "?powr n ((4*z + 1)/2) * ?powr n (-2*z) = ?powr n (1/2)" + by (simp add: algebra_simps exp_add[symmetric] add_divide_distrib) + finally show "?g n = ?h n" by (simp only: mult_ac) + qed + + moreover from z double_in_nonpos_Ints_imp[of z] have "2 * z \ \\<^sub>\\<^sub>0" by auto + hence "?g \ ?powr 2 (2*z) * Gamma z * Gamma (z+1/2) / Gamma (2*z)" + using LIMSEQ_subseq_LIMSEQ[OF Gamma_series'_LIMSEQ, of "op*2" "2*z"] + by (intro tendsto_intros Gamma_series'_LIMSEQ) + (simp_all add: o_def subseq_def Gamma_eq_zero_iff) + ultimately have "?h \ ?powr 2 (2*z) * Gamma z * Gamma (z+1/2) / Gamma (2*z)" + by (rule Lim_transform_eventually) + } note lim = this + + from assms double_in_nonpos_Ints_imp[of z] have z': "2 * z \ \\<^sub>\\<^sub>0" by auto + from fraction_not_in_ints[of 2 1] have "(1/2 :: 'a) \ \\<^sub>\\<^sub>0" + by (intro not_in_Ints_imp_not_in_nonpos_Ints) simp_all + with lim[of "1/2 :: 'a"] have "?h \ 2 * Gamma (1 / 2 :: 'a)" by (simp add: exp_of_real) + from LIMSEQ_unique[OF this lim[OF assms]] z' show ?thesis + by (simp add: divide_simps Gamma_eq_zero_iff ring_distribs exp_diff exp_of_real ac_simps) +qed + +(* TODO: perhaps this is unnecessary once we have the fact that a holomorphic function is + infinitely differentiable *) +private lemma Gamma_reflection_aux: + defines "h \ \z::complex. if z \ \ then 0 else + (of_real pi * cot (of_real pi*z) + Digamma z - Digamma (1 - z))" + defines "a \ complex_of_real pi" + obtains h' where "continuous_on UNIV h'" "\z. (h has_field_derivative (h' z)) (at z)" +proof - + define f where "f n = a * of_real (cos_coeff (n+1) - sin_coeff (n+2))" for n + define F where "F z = (if z = 0 then 0 else (cos (a*z) - sin (a*z)/(a*z)) / z)" for z + define g where "g n = complex_of_real (sin_coeff (n+1))" for n + define G where "G z = (if z = 0 then 1 else sin (a*z)/(a*z))" for z + have a_nz: "a \ 0" unfolding a_def by simp + + have "(\n. f n * (a*z)^n) sums (F z) \ (\n. g n * (a*z)^n) sums (G z)" + if "abs (Re z) < 1" for z + proof (cases "z = 0"; rule conjI) + assume "z \ 0" + note z = this that + + from z have sin_nz: "sin (a*z) \ 0" unfolding a_def by (auto simp: sin_eq_0) + have "(\n. of_real (sin_coeff n) * (a*z)^n) sums (sin (a*z))" using sin_converges[of "a*z"] + by (simp add: scaleR_conv_of_real) + from sums_split_initial_segment[OF this, of 1] + have "(\n. (a*z) * of_real (sin_coeff (n+1)) * (a*z)^n) sums (sin (a*z))" by (simp add: mult_ac) + from sums_mult[OF this, of "inverse (a*z)"] z a_nz + have A: "(\n. g n * (a*z)^n) sums (sin (a*z)/(a*z))" + by (simp add: field_simps g_def) + with z show "(\n. g n * (a*z)^n) sums (G z)" by (simp add: G_def) + from A z a_nz sin_nz have g_nz: "(\n. g n * (a*z)^n) \ 0" by (simp add: sums_iff g_def) + + have [simp]: "sin_coeff (Suc 0) = 1" by (simp add: sin_coeff_def) + from sums_split_initial_segment[OF sums_diff[OF cos_converges[of "a*z"] A], of 1] + have "(\n. z * f n * (a*z)^n) sums (cos (a*z) - sin (a*z) / (a*z))" + by (simp add: mult_ac scaleR_conv_of_real ring_distribs f_def g_def) + from sums_mult[OF this, of "inverse z"] z assms + show "(\n. f n * (a*z)^n) sums (F z)" by (simp add: divide_simps mult_ac f_def F_def) + next + assume z: "z = 0" + have "(\n. f n * (a * z) ^ n) sums f 0" using powser_sums_zero[of f] z by simp + with z show "(\n. f n * (a * z) ^ n) sums (F z)" + by (simp add: f_def F_def sin_coeff_def cos_coeff_def) + have "(\n. g n * (a * z) ^ n) sums g 0" using powser_sums_zero[of g] z by simp + with z show "(\n. g n * (a * z) ^ n) sums (G z)" + by (simp add: g_def G_def sin_coeff_def cos_coeff_def) + qed + note sums = conjunct1[OF this] conjunct2[OF this] + + define h2 where [abs_def]: + "h2 z = (\n. f n * (a*z)^n) / (\n. g n * (a*z)^n) + Digamma (1 + z) - Digamma (1 - z)" for z + define POWSER where [abs_def]: "POWSER f z = (\n. f n * (z^n :: complex))" for f z + define POWSER' where [abs_def]: "POWSER' f z = (\n. diffs f n * (z^n))" for f and z :: complex + define h2' where [abs_def]: + "h2' z = a * (POWSER g (a*z) * POWSER' f (a*z) - POWSER f (a*z) * POWSER' g (a*z)) / + (POWSER g (a*z))^2 + Polygamma 1 (1 + z) + Polygamma 1 (1 - z)" for z + + have h_eq: "h t = h2 t" if "abs (Re t) < 1" for t + proof - + from that have t: "t \ \ \ t = 0" by (auto elim!: Ints_cases simp: dist_0_norm) + hence "h t = a*cot (a*t) - 1/t + Digamma (1 + t) - Digamma (1 - t)" + unfolding h_def using Digamma_plus1[of t] by (force simp: field_simps a_def) + also have "a*cot (a*t) - 1/t = (F t) / (G t)" + using t by (auto simp add: divide_simps sin_eq_0 cot_def a_def F_def G_def) + also have "\ = (\n. f n * (a*t)^n) / (\n. g n * (a*t)^n)" + using sums[of t] that by (simp add: sums_iff dist_0_norm) + finally show "h t = h2 t" by (simp only: h2_def) + qed + + let ?A = "{z. abs (Re z) < 1}" + have "open ({z. Re z < 1} \ {z. Re z > -1})" + using open_halfspace_Re_gt open_halfspace_Re_lt by auto + also have "({z. Re z < 1} \ {z. Re z > -1}) = {z. abs (Re z) < 1}" by auto + finally have open_A: "open ?A" . + hence [simp]: "interior ?A = ?A" by (simp add: interior_open) + + have summable_f: "summable (\n. f n * z^n)" for z + by (rule powser_inside, rule sums_summable, rule sums[of "\ * of_real (norm z + 1) / a"]) + (simp_all add: norm_mult a_def del: of_real_add) + have summable_g: "summable (\n. g n * z^n)" for z + by (rule powser_inside, rule sums_summable, rule sums[of "\ * of_real (norm z + 1) / a"]) + (simp_all add: norm_mult a_def del: of_real_add) + have summable_fg': "summable (\n. diffs f n * z^n)" "summable (\n. diffs g n * z^n)" for z + by (intro termdiff_converges_all summable_f summable_g)+ + have "(POWSER f has_field_derivative (POWSER' f z)) (at z)" + "(POWSER g has_field_derivative (POWSER' g z)) (at z)" for z + unfolding POWSER_def POWSER'_def + by (intro termdiffs_strong_converges_everywhere summable_f summable_g)+ + note derivs = this[THEN DERIV_chain2[OF _ DERIV_cmult[OF DERIV_ident]], unfolded POWSER_def] + have "isCont (POWSER f) z" "isCont (POWSER g) z" "isCont (POWSER' f) z" "isCont (POWSER' g) z" + for z unfolding POWSER_def POWSER'_def + by (intro isCont_powser_converges_everywhere summable_f summable_g summable_fg')+ + note cont = this[THEN isCont_o2[rotated], unfolded POWSER_def POWSER'_def] + + { + fix z :: complex assume z: "abs (Re z) < 1" + define d where "d = \ * of_real (norm z + 1)" + have d: "abs (Re d) < 1" "norm z < norm d" by (simp_all add: d_def norm_mult del: of_real_add) + have "eventually (\z. h z = h2 z) (nhds z)" + using eventually_nhds_in_nhd[of z ?A] using h_eq z + by (auto elim!: eventually_mono simp: dist_0_norm) + + moreover from sums(2)[OF z] z have nz: "(\n. g n * (a * z) ^ n) \ 0" + unfolding G_def by (auto simp: sums_iff sin_eq_0 a_def) + have A: "z \ \ \ z = 0" using z by (auto elim!: Ints_cases) + have no_int: "1 + z \ \ \ z = 0" using z Ints_diff[of "1+z" 1] A + by (auto elim!: nonpos_Ints_cases) + have no_int': "1 - z \ \ \ z = 0" using z Ints_diff[of 1 "1-z"] A + by (auto elim!: nonpos_Ints_cases) + from no_int no_int' have no_int: "1 - z \ \\<^sub>\\<^sub>0" "1 + z \ \\<^sub>\\<^sub>0" by auto + have "(h2 has_field_derivative h2' z) (at z)" unfolding h2_def + by (rule DERIV_cong, (rule derivative_intros refl derivs[unfolded POWSER_def] nz no_int)+) + (auto simp: h2'_def POWSER_def field_simps power2_eq_square) + ultimately have deriv: "(h has_field_derivative h2' z) (at z)" + by (subst DERIV_cong_ev[OF refl _ refl]) + + from sums(2)[OF z] z have "(\n. g n * (a * z) ^ n) \ 0" + unfolding G_def by (auto simp: sums_iff a_def sin_eq_0) + hence "isCont h2' z" using no_int unfolding h2'_def[abs_def] POWSER_def POWSER'_def + by (intro continuous_intros cont + continuous_on_compose2[OF _ continuous_on_Polygamma[of "{z. Re z > 0}"]]) auto + note deriv and this + } note A = this + + interpret h: periodic_fun_simple' h + proof + fix z :: complex + show "h (z + 1) = h z" + proof (cases "z \ \") + assume z: "z \ \" + hence A: "z + 1 \ \" "z \ 0" using Ints_diff[of "z+1" 1] by auto + hence "Digamma (z + 1) - Digamma (-z) = Digamma z - Digamma (-z + 1)" + by (subst (1 2) Digamma_plus1) simp_all + with A z show "h (z + 1) = h z" + by (simp add: h_def sin_plus_pi cos_plus_pi ring_distribs cot_def) + qed (simp add: h_def) + qed + + have h2'_eq: "h2' (z - 1) = h2' z" if z: "Re z > 0" "Re z < 1" for z + proof - + have "((\z. h (z - 1)) has_field_derivative h2' (z - 1)) (at z)" + by (rule DERIV_cong, rule DERIV_chain'[OF _ A(1)]) + (insert z, auto intro!: derivative_eq_intros) + hence "(h has_field_derivative h2' (z - 1)) (at z)" by (subst (asm) h.minus_1) + moreover from z have "(h has_field_derivative h2' z) (at z)" by (intro A) simp_all + ultimately show "h2' (z - 1) = h2' z" by (rule DERIV_unique) + qed + + define h2'' where "h2'' z = h2' (z - of_int \Re z\)" for z + have deriv: "(h has_field_derivative h2'' z) (at z)" for z + proof - + fix z :: complex + have B: "\Re z - real_of_int \Re z\\ < 1" by linarith + have "((\t. h (t - of_int \Re z\)) has_field_derivative h2'' z) (at z)" + unfolding h2''_def by (rule DERIV_cong, rule DERIV_chain'[OF _ A(1)]) + (insert B, auto intro!: derivative_intros) + thus "(h has_field_derivative h2'' z) (at z)" by (simp add: h.minus_of_int) + qed + + have cont: "continuous_on UNIV h2''" + proof (intro continuous_at_imp_continuous_on ballI) + fix z :: complex + define r where "r = \Re z\" + define A where "A = {t. of_int r - 1 < Re t \ Re t < of_int r + 1}" + have "continuous_on A (\t. h2' (t - of_int r))" unfolding A_def + by (intro continuous_at_imp_continuous_on isCont_o2[OF _ A(2)] ballI continuous_intros) + (simp_all add: abs_real_def) + moreover have "h2'' t = h2' (t - of_int r)" if t: "t \ A" for t + proof (cases "Re t \ of_int r") + case True + from t have "of_int r - 1 < Re t" "Re t < of_int r + 1" by (simp_all add: A_def) + with True have "\Re t\ = \Re z\" unfolding r_def by linarith + thus ?thesis by (auto simp: r_def h2''_def) + next + case False + from t have t: "of_int r - 1 < Re t" "Re t < of_int r + 1" by (simp_all add: A_def) + with False have t': "\Re t\ = \Re z\ - 1" unfolding r_def by linarith + moreover from t False have "h2' (t - of_int r + 1 - 1) = h2' (t - of_int r + 1)" + by (intro h2'_eq) simp_all + ultimately show ?thesis by (auto simp: r_def h2''_def algebra_simps t') + qed + ultimately have "continuous_on A h2''" by (subst continuous_on_cong[OF refl]) + moreover { + have "open ({t. of_int r - 1 < Re t} \ {t. of_int r + 1 > Re t})" + by (intro open_Int open_halfspace_Re_gt open_halfspace_Re_lt) + also have "{t. of_int r - 1 < Re t} \ {t. of_int r + 1 > Re t} = A" + unfolding A_def by blast + finally have "open A" . + } + ultimately have C: "isCont h2'' t" if "t \ A" for t using that + by (subst (asm) continuous_on_eq_continuous_at) auto + have "of_int r - 1 < Re z" "Re z < of_int r + 1" unfolding r_def by linarith+ + thus "isCont h2'' z" by (intro C) (simp_all add: A_def) + qed + + from that[OF cont deriv] show ?thesis . +qed + +lemma Gamma_reflection_complex: + fixes z :: complex + shows "Gamma z * Gamma (1 - z) = of_real pi / sin (of_real pi * z)" +proof - + let ?g = "\z::complex. Gamma z * Gamma (1 - z) * sin (of_real pi * z)" + define g where [abs_def]: "g z = (if z \ \ then of_real pi else ?g z)" for z :: complex + let ?h = "\z::complex. (of_real pi * cot (of_real pi*z) + Digamma z - Digamma (1 - z))" + define h where [abs_def]: "h z = (if z \ \ then 0 else ?h z)" for z :: complex + + \ \@{term g} is periodic with period 1.\ + interpret g: periodic_fun_simple' g + proof + fix z :: complex + show "g (z + 1) = g z" + proof (cases "z \ \") + case False + hence "z * g z = z * Beta z (- z + 1) * sin (of_real pi * z)" by (simp add: g_def Beta_def) + also have "z * Beta z (- z + 1) = (z + 1 + -z) * Beta (z + 1) (- z + 1)" + using False Ints_diff[of 1 "1 - z"] nonpos_Ints_subset_Ints + by (subst Beta_plus1_left [symmetric]) auto + also have "\ * sin (of_real pi * z) = z * (Beta (z + 1) (-z) * sin (of_real pi * (z + 1)))" + using False Ints_diff[of "z+1" 1] Ints_minus[of "-z"] nonpos_Ints_subset_Ints + by (subst Beta_plus1_right) (auto simp: ring_distribs sin_plus_pi) + also from False have "Beta (z + 1) (-z) * sin (of_real pi * (z + 1)) = g (z + 1)" + using Ints_diff[of "z+1" 1] by (auto simp: g_def Beta_def) + finally show "g (z + 1) = g z" using False by (subst (asm) mult_left_cancel) auto + qed (simp add: g_def) + qed + + \ \@{term g} is entire.\ + have g_g': "(g has_field_derivative (h z * g z)) (at z)" for z :: complex + proof (cases "z \ \") + let ?h' = "\z. Beta z (1 - z) * ((Digamma z - Digamma (1 - z)) * sin (z * of_real pi) + + of_real pi * cos (z * of_real pi))" + case False + from False have "eventually (\t. t \ UNIV - \) (nhds z)" + by (intro eventually_nhds_in_open) (auto simp: open_Diff) + hence "eventually (\t. g t = ?g t) (nhds z)" by eventually_elim (simp add: g_def) + moreover { + from False Ints_diff[of 1 "1-z"] have "1 - z \ \" by auto + hence "(?g has_field_derivative ?h' z) (at z)" using nonpos_Ints_subset_Ints + by (auto intro!: derivative_eq_intros simp: algebra_simps Beta_def) + also from False have "sin (of_real pi * z) \ 0" by (subst sin_eq_0) auto + hence "?h' z = h z * g z" + using False unfolding g_def h_def cot_def by (simp add: field_simps Beta_def) + finally have "(?g has_field_derivative (h z * g z)) (at z)" . + } + ultimately show ?thesis by (subst DERIV_cong_ev[OF refl _ refl]) + next + case True + then obtain n where z: "z = of_int n" by (auto elim!: Ints_cases) + let ?t = "(\z::complex. if z = 0 then 1 else sin z / z) \ (\z. of_real pi * z)" + have deriv_0: "(g has_field_derivative 0) (at 0)" + proof (subst DERIV_cong_ev[OF refl _ refl]) + show "eventually (\z. g z = of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z) (nhds 0)" + using eventually_nhds_ball[OF zero_less_one, of "0::complex"] + proof eventually_elim + fix z :: complex assume z: "z \ ball 0 1" + show "g z = of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z" + proof (cases "z = 0") + assume z': "z \ 0" + with z have z'': "z \ \\<^sub>\\<^sub>0" "z \ \" by (auto elim!: Ints_cases simp: dist_0_norm) + from Gamma_plus1[OF this(1)] have "Gamma z = Gamma (z + 1) / z" by simp + with z'' z' show ?thesis by (simp add: g_def ac_simps) + qed (simp add: g_def) + qed + have "(?t has_field_derivative (0 * of_real pi)) (at 0)" + using has_field_derivative_sin_z_over_z[of "UNIV :: complex set"] + by (intro DERIV_chain) simp_all + thus "((\z. of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z) has_field_derivative 0) (at 0)" + by (auto intro!: derivative_eq_intros simp: o_def) + qed + + have "((g \ (\x. x - of_int n)) has_field_derivative 0 * 1) (at (of_int n))" + using deriv_0 by (intro DERIV_chain) (auto intro!: derivative_eq_intros) + also have "g \ (\x. x - of_int n) = g" by (intro ext) (simp add: g.minus_of_int) + finally show "(g has_field_derivative (h z * g z)) (at z)" by (simp add: z h_def) + qed + + have g_eq: "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * g z" if "Re z > -1" "Re z < 2" for z + proof (cases "z \ \") + case True + with that have "z = 0 \ z = 1" by (force elim!: Ints_cases) + moreover have "g 0 * g (1/2) = Gamma (1/2)^2 * g 0" + using fraction_not_in_ints[where 'a = complex, of 2 1] by (simp add: g_def power2_eq_square) + moreover have "g (1/2) * g 1 = Gamma (1/2)^2 * g 1" + using fraction_not_in_ints[where 'a = complex, of 2 1] + by (simp add: g_def power2_eq_square Beta_def algebra_simps) + ultimately show ?thesis by force + next + case False + hence z: "z/2 \ \" "(z+1)/2 \ \" using Ints_diff[of "z+1" 1] by (auto elim!: Ints_cases) + hence z': "z/2 \ \\<^sub>\\<^sub>0" "(z+1)/2 \ \\<^sub>\\<^sub>0" by (auto elim!: nonpos_Ints_cases) + from z have "1-z/2 \ \" "1-((z+1)/2) \ \" + using Ints_diff[of 1 "1-z/2"] Ints_diff[of 1 "1-((z+1)/2)"] by auto + hence z'': "1-z/2 \ \\<^sub>\\<^sub>0" "1-((z+1)/2) \ \\<^sub>\\<^sub>0" by (auto elim!: nonpos_Ints_cases) + from z have "g (z/2) * g ((z+1)/2) = + (Gamma (z/2) * Gamma ((z+1)/2)) * (Gamma (1-z/2) * Gamma (1-((z+1)/2))) * + (sin (of_real pi * z/2) * sin (of_real pi * (z+1)/2))" + by (simp add: g_def) + also from z' Gamma_legendre_duplication_aux[of "z/2"] + have "Gamma (z/2) * Gamma ((z+1)/2) = exp ((1-z) * of_real (ln 2)) * Gamma (1/2) * Gamma z" + by (simp add: add_divide_distrib) + also from z'' Gamma_legendre_duplication_aux[of "1-(z+1)/2"] + have "Gamma (1-z/2) * Gamma (1-(z+1)/2) = + Gamma (1-z) * Gamma (1/2) * exp (z * of_real (ln 2))" + by (simp add: add_divide_distrib ac_simps) + finally have "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * (Gamma z * Gamma (1-z) * + (2 * (sin (of_real pi*z/2) * sin (of_real pi*(z+1)/2))))" + by (simp add: add_ac power2_eq_square exp_add ring_distribs exp_diff exp_of_real) + also have "sin (of_real pi*(z+1)/2) = cos (of_real pi*z/2)" + using cos_sin_eq[of "- of_real pi * z/2", symmetric] + by (simp add: ring_distribs add_divide_distrib ac_simps) + also have "2 * (sin (of_real pi*z/2) * cos (of_real pi*z/2)) = sin (of_real pi * z)" + by (subst sin_times_cos) (simp add: field_simps) + also have "Gamma z * Gamma (1 - z) * sin (complex_of_real pi * z) = g z" + using \z \ \\ by (simp add: g_def) + finally show ?thesis . + qed + have g_eq: "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * g z" for z + proof - + define r where "r = \Re z / 2\" + have "Gamma (1/2)^2 * g z = Gamma (1/2)^2 * g (z - of_int (2*r))" by (simp only: g.minus_of_int) + also have "of_int (2*r) = 2 * of_int r" by simp + also have "Re z - 2 * of_int r > -1" "Re z - 2 * of_int r < 2" unfolding r_def by linarith+ + hence "Gamma (1/2)^2 * g (z - 2 * of_int r) = + g ((z - 2 * of_int r)/2) * g ((z - 2 * of_int r + 1)/2)" + unfolding r_def by (intro g_eq[symmetric]) simp_all + also have "(z - 2 * of_int r) / 2 = z/2 - of_int r" by simp + also have "g \ = g (z/2)" by (rule g.minus_of_int) + also have "(z - 2 * of_int r + 1) / 2 = (z + 1)/2 - of_int r" by simp + also have "g \ = g ((z+1)/2)" by (rule g.minus_of_int) + finally show ?thesis .. + qed + + have g_nz [simp]: "g z \ 0" for z :: complex + unfolding g_def using Ints_diff[of 1 "1 - z"] + by (auto simp: Gamma_eq_zero_iff sin_eq_0 dest!: nonpos_Ints_Int) + + have h_eq: "h z = (h (z/2) + h ((z+1)/2)) / 2" for z + proof - + have "((\t. g (t/2) * g ((t+1)/2)) has_field_derivative + (g (z/2) * g ((z+1)/2)) * ((h (z/2) + h ((z+1)/2)) / 2)) (at z)" + by (auto intro!: derivative_eq_intros g_g'[THEN DERIV_chain2] simp: field_simps) + hence "((\t. Gamma (1/2)^2 * g t) has_field_derivative + Gamma (1/2)^2 * g z * ((h (z/2) + h ((z+1)/2)) / 2)) (at z)" + by (subst (1 2) g_eq[symmetric]) simp + from DERIV_cmult[OF this, of "inverse ((Gamma (1/2))^2)"] + have "(g has_field_derivative (g z * ((h (z/2) + h ((z+1)/2))/2))) (at z)" + using fraction_not_in_ints[where 'a = complex, of 2 1] + by (simp add: divide_simps Gamma_eq_zero_iff not_in_Ints_imp_not_in_nonpos_Ints) + moreover have "(g has_field_derivative (g z * h z)) (at z)" + using g_g'[of z] by (simp add: ac_simps) + ultimately have "g z * h z = g z * ((h (z/2) + h ((z+1)/2))/2)" + by (intro DERIV_unique) + thus "h z = (h (z/2) + h ((z+1)/2)) / 2" by simp + qed + + obtain h' where h'_cont: "continuous_on UNIV h'" and + h_h': "\z. (h has_field_derivative h' z) (at z)" + unfolding h_def by (erule Gamma_reflection_aux) + + have h'_eq: "h' z = (h' (z/2) + h' ((z+1)/2)) / 4" for z + proof - + have "((\t. (h (t/2) + h ((t+1)/2)) / 2) has_field_derivative + ((h' (z/2) + h' ((z+1)/2)) / 4)) (at z)" + by (fastforce intro!: derivative_eq_intros h_h'[THEN DERIV_chain2]) + hence "(h has_field_derivative ((h' (z/2) + h' ((z+1)/2))/4)) (at z)" + by (subst (asm) h_eq[symmetric]) + from h_h' and this show "h' z = (h' (z/2) + h' ((z+1)/2)) / 4" by (rule DERIV_unique) + qed + + have h'_zero: "h' z = 0" for z + proof - + define m where "m = max 1 \Re z\" + define B where "B = {t. abs (Re t) \ m \ abs (Im t) \ abs (Im z)}" + have "closed ({t. Re t \ -m} \ {t. Re t \ m} \ + {t. Im t \ -\Im z\} \ {t. Im t \ \Im z\})" + (is "closed ?B") by (intro closed_Int closed_halfspace_Re_ge closed_halfspace_Re_le + closed_halfspace_Im_ge closed_halfspace_Im_le) + also have "?B = B" unfolding B_def by fastforce + finally have "closed B" . + moreover have "bounded B" unfolding bounded_iff + proof (intro ballI exI) + fix t assume t: "t \ B" + have "norm t \ \Re t\ + \Im t\" by (rule cmod_le) + also from t have "\Re t\ \ m" unfolding B_def by blast + also from t have "\Im t\ \ \Im z\" unfolding B_def by blast + finally show "norm t \ m + \Im z\" by - simp + qed + ultimately have compact: "compact B" by (subst compact_eq_bounded_closed) blast + + define M where "M = (SUP z:B. norm (h' z))" + have "compact (h' ` B)" + by (intro compact_continuous_image continuous_on_subset[OF h'_cont] compact) blast+ + hence bdd: "bdd_above ((\z. norm (h' z)) ` B)" + using bdd_above_norm[of "h' ` B"] by (simp add: image_comp o_def compact_imp_bounded) + have "norm (h' z) \ M" unfolding M_def by (intro cSUP_upper bdd) (simp_all add: B_def m_def) + also have "M \ M/2" + proof (subst M_def, subst cSUP_le_iff) + have "z \ B" unfolding B_def m_def by simp + thus "B \ {}" by auto + next + show "\z\B. norm (h' z) \ M/2" + proof + fix t :: complex assume t: "t \ B" + from h'_eq[of t] t have "h' t = (h' (t/2) + h' ((t+1)/2)) / 4" by (simp add: dist_0_norm) + also have "norm \ = norm (h' (t/2) + h' ((t+1)/2)) / 4" by simp + also have "norm (h' (t/2) + h' ((t+1)/2)) \ norm (h' (t/2)) + norm (h' ((t+1)/2))" + by (rule norm_triangle_ineq) + also from t have "abs (Re ((t + 1)/2)) \ m" unfolding m_def B_def by auto + with t have "t/2 \ B" "(t+1)/2 \ B" unfolding B_def by auto + hence "norm (h' (t/2)) + norm (h' ((t+1)/2)) \ M + M" unfolding M_def + by (intro add_mono cSUP_upper bdd) (auto simp: B_def) + also have "(M + M) / 4 = M / 2" by simp + finally show "norm (h' t) \ M/2" by - simp_all + qed + qed (insert bdd, auto simp: cball_eq_empty) + hence "M \ 0" by simp + finally show "h' z = 0" by simp + qed + have h_h'_2: "(h has_field_derivative 0) (at z)" for z + using h_h'[of z] h'_zero[of z] by simp + + have g_real: "g z \ \" if "z \ \" for z + unfolding g_def using that by (auto intro!: Reals_mult Gamma_complex_real) + have h_real: "h z \ \" if "z \ \" for z + unfolding h_def using that by (auto intro!: Reals_mult Reals_add Reals_diff Polygamma_Real) + have g_nz: "g z \ 0" for z unfolding g_def using Ints_diff[of 1 "1-z"] + by (auto simp: Gamma_eq_zero_iff sin_eq_0) + + from h'_zero h_h'_2 have "\c. \z\UNIV. h z = c" + by (intro has_field_derivative_zero_constant) (simp_all add: dist_0_norm) + then obtain c where c: "\z. h z = c" by auto + have "\u. u \ closed_segment 0 1 \ Re (g 1) - Re (g 0) = Re (h u * g u * (1 - 0))" + by (intro complex_mvt_line g_g') + find_theorems name:deriv Reals + then guess u by (elim exE conjE) note u = this + from u(1) have u': "u \ \" unfolding closed_segment_def + by (auto simp: scaleR_conv_of_real) + from u' g_real[of u] g_nz[of u] have "Re (g u) \ 0" by (auto elim!: Reals_cases) + with u(2) c[of u] g_real[of u] g_nz[of u] u' + have "Re c = 0" by (simp add: complex_is_Real_iff g.of_1) + with h_real[of 0] c[of 0] have "c = 0" by (auto elim!: Reals_cases) + with c have A: "h z * g z = 0" for z by simp + hence "(g has_field_derivative 0) (at z)" for z using g_g'[of z] by simp + hence "\c'. \z\UNIV. g z = c'" by (intro has_field_derivative_zero_constant) simp_all + then obtain c' where c: "\z. g z = c'" by (force simp: dist_0_norm) + from this[of 0] have "c' = pi" unfolding g_def by simp + with c have "g z = pi" by simp + + show ?thesis + proof (cases "z \ \") + case False + with \g z = pi\ show ?thesis by (auto simp: g_def divide_simps) + next + case True + then obtain n where n: "z = of_int n" by (elim Ints_cases) + with sin_eq_0[of "of_real pi * z"] have "sin (of_real pi * z) = 0" by force + moreover have "of_int (1 - n) \ \\<^sub>\\<^sub>0" if "n > 0" using that by (intro nonpos_Ints_of_int) simp + ultimately show ?thesis using n + by (cases "n \ 0") (auto simp: Gamma_eq_zero_iff nonpos_Ints_of_int) + qed +qed + +lemma rGamma_reflection_complex: + "rGamma z * rGamma (1 - z :: complex) = sin (of_real pi * z) / of_real pi" + using Gamma_reflection_complex[of z] + by (simp add: Gamma_def divide_simps split: if_split_asm) + +lemma rGamma_reflection_complex': + "rGamma z * rGamma (- z :: complex) = -z * sin (of_real pi * z) / of_real pi" +proof - + have "rGamma z * rGamma (-z) = -z * (rGamma z * rGamma (1 - z))" + using rGamma_plus1[of "-z", symmetric] by simp + also have "rGamma z * rGamma (1 - z) = sin (of_real pi * z) / of_real pi" + by (rule rGamma_reflection_complex) + finally show ?thesis by simp +qed + +lemma Gamma_reflection_complex': + "Gamma z * Gamma (- z :: complex) = - of_real pi / (z * sin (of_real pi * z))" + using rGamma_reflection_complex'[of z] by (force simp add: Gamma_def divide_simps mult_ac) + + + +lemma Gamma_one_half_real: "Gamma (1/2 :: real) = sqrt pi" +proof - + from Gamma_reflection_complex[of "1/2"] fraction_not_in_ints[where 'a = complex, of 2 1] + have "Gamma (1/2 :: complex)^2 = of_real pi" by (simp add: power2_eq_square) + hence "of_real pi = Gamma (complex_of_real (1/2))^2" by simp + also have "\ = of_real ((Gamma (1/2))^2)" by (subst Gamma_complex_of_real) simp_all + finally have "Gamma (1/2)^2 = pi" by (subst (asm) of_real_eq_iff) simp_all + moreover have "Gamma (1/2 :: real) \ 0" using Gamma_real_pos[of "1/2"] by simp + ultimately show ?thesis by (rule real_sqrt_unique [symmetric]) +qed + +lemma Gamma_one_half_complex: "Gamma (1/2 :: complex) = of_real (sqrt pi)" +proof - + have "Gamma (1/2 :: complex) = Gamma (of_real (1/2))" by simp + also have "\ = of_real (sqrt pi)" by (simp only: Gamma_complex_of_real Gamma_one_half_real) + finally show ?thesis . +qed + +lemma Gamma_legendre_duplication: + fixes z :: complex + assumes "z \ \\<^sub>\\<^sub>0" "z + 1/2 \ \\<^sub>\\<^sub>0" + shows "Gamma z * Gamma (z + 1/2) = + exp ((1 - 2*z) * of_real (ln 2)) * of_real (sqrt pi) * Gamma (2*z)" + using Gamma_legendre_duplication_aux[OF assms] by (simp add: Gamma_one_half_complex) + +end + + +subsection \Limits and residues\ + +text \ + The inverse of the Gamma function has simple zeros: +\ + +lemma rGamma_zeros: + "(\z. rGamma z / (z + of_nat n)) \ (- of_nat n) \ ((-1)^n * fact n :: 'a :: Gamma)" +proof (subst tendsto_cong) + let ?f = "\z. pochhammer z n * rGamma (z + of_nat (Suc n)) :: 'a" + from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV] + show "eventually (\z. rGamma z / (z + of_nat n) = ?f z) (at (- of_nat n))" + by (subst pochhammer_rGamma[of _ "Suc n"]) + (auto elim!: eventually_mono simp: divide_simps pochhammer_rec' eq_neg_iff_add_eq_0) + have "isCont ?f (- of_nat n)" by (intro continuous_intros) + thus "?f \ (- of_nat n) \ (- 1) ^ n * fact n" unfolding isCont_def + by (simp add: pochhammer_same) +qed + + +text \ + The simple zeros of the inverse of the Gamma function correspond to simple poles of the Gamma function, + and their residues can easily be computed from the limit we have just proven: +\ + +lemma Gamma_poles: "filterlim Gamma at_infinity (at (- of_nat n :: 'a :: Gamma))" +proof - + from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV] + have "eventually (\z. rGamma z \ (0 :: 'a)) (at (- of_nat n))" + by (auto elim!: eventually_mono nonpos_Ints_cases' + simp: rGamma_eq_zero_iff dist_of_nat dist_minus) + with isCont_rGamma[of "- of_nat n :: 'a", OF continuous_ident] + have "filterlim (\z. inverse (rGamma z) :: 'a) at_infinity (at (- of_nat n))" + unfolding isCont_def by (intro filterlim_compose[OF filterlim_inverse_at_infinity]) + (simp_all add: filterlim_at) + moreover have "(\z. inverse (rGamma z) :: 'a) = Gamma" + by (intro ext) (simp add: rGamma_inverse_Gamma) + ultimately show ?thesis by (simp only: ) +qed + +lemma Gamma_residues: + "(\z. Gamma z * (z + of_nat n)) \ (- of_nat n) \ ((-1)^n / fact n :: 'a :: Gamma)" +proof (subst tendsto_cong) + let ?c = "(- 1) ^ n / fact n :: 'a" + from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV] + show "eventually (\z. Gamma z * (z + of_nat n) = inverse (rGamma z / (z + of_nat n))) + (at (- of_nat n))" + by (auto elim!: eventually_mono simp: divide_simps rGamma_inverse_Gamma) + have "(\z. inverse (rGamma z / (z + of_nat n))) \ (- of_nat n) \ + inverse ((- 1) ^ n * fact n :: 'a)" + by (intro tendsto_intros rGamma_zeros) simp_all + also have "inverse ((- 1) ^ n * fact n) = ?c" + by (simp_all add: field_simps power_mult_distrib [symmetric] del: power_mult_distrib) + finally show "(\z. inverse (rGamma z / (z + of_nat n))) \ (- of_nat n) \ ?c" . +qed + + + +subsection \Alternative definitions\ + + +subsubsection \Variant of the Euler form\ + + +definition Gamma_series_euler' where + "Gamma_series_euler' z n = + inverse z * (\k=1..n. exp (z * of_real (ln (1 + inverse (of_nat k)))) / (1 + z / of_nat k))" + +context +begin +private lemma Gamma_euler'_aux1: + fixes z :: "'a :: {real_normed_field,banach}" + assumes n: "n > 0" + shows "exp (z * of_real (ln (of_nat n + 1))) = (\k=1..n. exp (z * of_real (ln (1 + 1 / of_nat k))))" +proof - + have "(\k=1..n. exp (z * of_real (ln (1 + 1 / of_nat k)))) = + exp (z * of_real (\k = 1..n. ln (1 + 1 / real_of_nat k)))" + by (subst exp_setsum [symmetric]) (simp_all add: setsum_right_distrib) + also have "(\k=1..n. ln (1 + 1 / of_nat k) :: real) = ln (\k=1..n. 1 + 1 / real_of_nat k)" + by (subst ln_setprod [symmetric]) (auto intro!: add_pos_nonneg) + also have "(\k=1..n. 1 + 1 / of_nat k :: real) = (\k=1..n. (of_nat k + 1) / of_nat k)" + by (intro setprod.cong) (simp_all add: divide_simps) + also have "(\k=1..n. (of_nat k + 1) / of_nat k :: real) = of_nat n + 1" + by (induction n) (simp_all add: setprod_nat_ivl_Suc' divide_simps) + finally show ?thesis .. +qed + +lemma Gamma_series_euler': + assumes z: "(z :: 'a :: Gamma) \ \\<^sub>\\<^sub>0" + shows "(\n. Gamma_series_euler' z n) \ Gamma z" +proof (rule Gamma_seriesI, rule Lim_transform_eventually) + let ?f = "\n. fact n * exp (z * of_real (ln (of_nat n + 1))) / pochhammer z (n + 1)" + let ?r = "\n. ?f n / Gamma_series z n" + let ?r' = "\n. exp (z * of_real (ln (of_nat (Suc n) / of_nat n)))" + from z have z': "z \ 0" by auto + + have "eventually (\n. ?r' n = ?r n) sequentially" using eventually_gt_at_top[of "0::nat"] + using z by (auto simp: divide_simps Gamma_series_def ring_distribs exp_diff ln_div add_ac + elim!: eventually_mono dest: pochhammer_eq_0_imp_nonpos_Int) + moreover have "?r' \ exp (z * of_real (ln 1))" + by (intro tendsto_intros LIMSEQ_Suc_n_over_n) simp_all + ultimately show "?r \ 1" by (force dest!: Lim_transform_eventually) + + from eventually_gt_at_top[of "0::nat"] + show "eventually (\n. ?r n = Gamma_series_euler' z n / Gamma_series z n) sequentially" + proof eventually_elim + fix n :: nat assume n: "n > 0" + from n z' have "Gamma_series_euler' z n = + exp (z * of_real (ln (of_nat n + 1))) / (z * (\k=1..n. (1 + z / of_nat k)))" + by (subst Gamma_euler'_aux1) + (simp_all add: Gamma_series_euler'_def setprod.distrib + setprod_inversef[symmetric] divide_inverse) + also have "(\k=1..n. (1 + z / of_nat k)) = pochhammer (z + 1) n / fact n" + by (cases n) (simp_all add: pochhammer_setprod fact_setprod atLeastLessThanSuc_atLeastAtMost + setprod_dividef [symmetric] field_simps setprod.atLeast_Suc_atMost_Suc_shift) + also have "z * \ = pochhammer z (Suc n) / fact n" by (simp add: pochhammer_rec) + finally show "?r n = Gamma_series_euler' z n / Gamma_series z n" by simp + qed +qed + +end + + + +subsubsection \Weierstrass form\ + +definition Gamma_series_weierstrass :: "'a :: {banach,real_normed_field} \ nat \ 'a" where + "Gamma_series_weierstrass z n = + exp (-euler_mascheroni * z) / z * (\k=1..n. exp (z / of_nat k) / (1 + z / of_nat k))" + +definition rGamma_series_weierstrass :: "'a :: {banach,real_normed_field} \ nat \ 'a" where + "rGamma_series_weierstrass z n = + exp (euler_mascheroni * z) * z * (\k=1..n. (1 + z / of_nat k) * exp (-z / of_nat k))" + +lemma Gamma_series_weierstrass_nonpos_Ints: + "eventually (\k. Gamma_series_weierstrass (- of_nat n) k = 0) sequentially" + using eventually_ge_at_top[of n] by eventually_elim (auto simp: Gamma_series_weierstrass_def) + +lemma rGamma_series_weierstrass_nonpos_Ints: + "eventually (\k. rGamma_series_weierstrass (- of_nat n) k = 0) sequentially" + using eventually_ge_at_top[of n] by eventually_elim (auto simp: rGamma_series_weierstrass_def) + +lemma Gamma_weierstrass_complex: "Gamma_series_weierstrass z \ Gamma (z :: complex)" +proof (cases "z \ \\<^sub>\\<^sub>0") + case True + then obtain n where "z = - of_nat n" by (elim nonpos_Ints_cases') + also from True have "Gamma_series_weierstrass \ \ Gamma z" + by (simp add: tendsto_cong[OF Gamma_series_weierstrass_nonpos_Ints] Gamma_nonpos_Int) + finally show ?thesis . +next + case False + hence z: "z \ 0" by auto + let ?f = "(\x. \x = Suc 0..x. exp (z / of_nat x) / (1 + z / of_nat x))" + have A: "exp (ln (1 + z / of_nat n)) = (1 + z / of_nat n)" if "n \ 1" for n :: nat + using False that by (subst exp_Ln) (auto simp: field_simps dest!: plus_of_nat_eq_0_imp) + have "(\n. \k=1..n. z / of_nat k - ln (1 + z / of_nat k)) \ ln_Gamma z + euler_mascheroni * z + ln z" + using ln_Gamma_series'_aux[OF False] + by (simp only: atLeastLessThanSuc_atLeastAtMost [symmetric] One_nat_def + setsum_shift_bounds_Suc_ivl sums_def atLeast0LessThan) + from tendsto_exp[OF this] False z have "?f \ z * exp (euler_mascheroni * z) * Gamma z" + by (simp add: exp_add exp_setsum exp_diff mult_ac Gamma_complex_altdef A) + from tendsto_mult[OF tendsto_const[of "exp (-euler_mascheroni * z) / z"] this] z + show "Gamma_series_weierstrass z \ Gamma z" + by (simp add: exp_minus divide_simps Gamma_series_weierstrass_def [abs_def]) +qed + +lemma tendsto_complex_of_real_iff: "((\x. complex_of_real (f x)) \ of_real c) F = (f \ c) F" + by (rule tendsto_of_real_iff) + +lemma Gamma_weierstrass_real: "Gamma_series_weierstrass x \ Gamma (x :: real)" + using Gamma_weierstrass_complex[of "of_real x"] unfolding Gamma_series_weierstrass_def[abs_def] + by (subst tendsto_complex_of_real_iff [symmetric]) + (simp_all add: exp_of_real[symmetric] Gamma_complex_of_real) + +lemma rGamma_weierstrass_complex: "rGamma_series_weierstrass z \ rGamma (z :: complex)" +proof (cases "z \ \\<^sub>\\<^sub>0") + case True + then obtain n where "z = - of_nat n" by (elim nonpos_Ints_cases') + also from True have "rGamma_series_weierstrass \ \ rGamma z" + by (simp add: tendsto_cong[OF rGamma_series_weierstrass_nonpos_Ints] rGamma_nonpos_Int) + finally show ?thesis . +next + case False + have "rGamma_series_weierstrass z = (\n. inverse (Gamma_series_weierstrass z n))" + by (simp add: rGamma_series_weierstrass_def[abs_def] Gamma_series_weierstrass_def + exp_minus divide_inverse setprod_inversef[symmetric] mult_ac) + also from False have "\ \ inverse (Gamma z)" + by (intro tendsto_intros Gamma_weierstrass_complex) (simp add: Gamma_eq_zero_iff) + finally show ?thesis by (simp add: Gamma_def) +qed + +subsubsection \Binomial coefficient form\ + +lemma Gamma_gbinomial: + "(\n. ((z + of_nat n) gchoose n) * exp (-z * of_real (ln (of_nat n)))) \ rGamma (z+1)" +proof (cases "z = 0") + case False + show ?thesis + proof (rule Lim_transform_eventually) + let ?powr = "\a b. exp (b * of_real (ln (of_nat a)))" + show "eventually (\n. rGamma_series z n / z = + ((z + of_nat n) gchoose n) * ?powr n (-z)) sequentially" + proof (intro always_eventually allI) + fix n :: nat + from False have "((z + of_nat n) gchoose n) = pochhammer z (Suc n) / z / fact n" + by (simp add: gbinomial_pochhammer' pochhammer_rec) + also have "pochhammer z (Suc n) / z / fact n * ?powr n (-z) = rGamma_series z n / z" + by (simp add: rGamma_series_def divide_simps exp_minus) + finally show "rGamma_series z n / z = ((z + of_nat n) gchoose n) * ?powr n (-z)" .. + qed + + from False have "(\n. rGamma_series z n / z) \ rGamma z / z" by (intro tendsto_intros) + also from False have "rGamma z / z = rGamma (z + 1)" using rGamma_plus1[of z] + by (simp add: field_simps) + finally show "(\n. rGamma_series z n / z) \ rGamma (z+1)" . + qed +qed (simp_all add: binomial_gbinomial [symmetric]) + +lemma gbinomial_minus': "(a + of_nat b) gchoose b = (- 1) ^ b * (- (a + 1) gchoose b)" + by (subst gbinomial_minus) (simp add: power_mult_distrib [symmetric]) + +lemma gbinomial_asymptotic: + fixes z :: "'a :: Gamma" + shows "(\n. (z gchoose n) / ((-1)^n / exp ((z+1) * of_real (ln (real n))))) \ + inverse (Gamma (- z))" + unfolding rGamma_inverse_Gamma [symmetric] using Gamma_gbinomial[of "-z-1"] + by (subst (asm) gbinomial_minus') + (simp add: add_ac mult_ac divide_inverse power_inverse [symmetric]) + +lemma fact_binomial_limit: + "(\n. of_nat ((k + n) choose n) / of_nat (n ^ k) :: 'a :: Gamma) \ 1 / fact k" +proof (rule Lim_transform_eventually) + have "(\n. of_nat ((k + n) choose n) / of_real (exp (of_nat k * ln (real_of_nat n)))) + \ 1 / Gamma (of_nat (Suc k) :: 'a)" (is "?f \ _") + using Gamma_gbinomial[of "of_nat k :: 'a"] + by (simp add: binomial_gbinomial add_ac Gamma_def divide_simps exp_of_real [symmetric] exp_minus) + also have "Gamma (of_nat (Suc k)) = fact k" by (simp add: Gamma_fact) + finally show "?f \ 1 / fact k" . + + show "eventually (\n. ?f n = of_nat ((k + n) choose n) / of_nat (n ^ k)) sequentially" + using eventually_gt_at_top[of "0::nat"] + proof eventually_elim + fix n :: nat assume n: "n > 0" + from n have "exp (real_of_nat k * ln (real_of_nat n)) = real_of_nat (n^k)" + by (simp add: exp_of_nat_mult) + thus "?f n = of_nat ((k + n) choose n) / of_nat (n ^ k)" by simp + qed +qed + +lemma binomial_asymptotic': + "(\n. of_nat ((k + n) choose n) / (of_nat (n ^ k) / fact k) :: 'a :: Gamma) \ 1" + using tendsto_mult[OF fact_binomial_limit[of k] tendsto_const[of "fact k :: 'a"]] by simp + +lemma gbinomial_Beta: + assumes "z + 1 \ \\<^sub>\\<^sub>0" + shows "((z::'a::Gamma) gchoose n) = inverse ((z + 1) * Beta (z - of_nat n + 1) (of_nat n + 1))" +using assms +proof (induction n arbitrary: z) + case 0 + hence "z + 2 \ \\<^sub>\\<^sub>0" + using plus_one_in_nonpos_Ints_imp[of "z+1"] by (auto simp: add.commute) + with 0 show ?case + by (auto simp: Beta_def Gamma_eq_zero_iff Gamma_plus1 [symmetric] add.commute) +next + case (Suc n z) + show ?case + proof (cases "z \ \\<^sub>\\<^sub>0") + case True + with Suc.prems have "z = 0" + by (auto elim!: nonpos_Ints_cases simp: algebra_simps one_plus_of_int_in_nonpos_Ints_iff) + show ?thesis + proof (cases "n = 0") + case True + with \z = 0\ show ?thesis + by (simp add: Beta_def Gamma_eq_zero_iff Gamma_plus1 [symmetric]) + next + case False + with \z = 0\ show ?thesis + by (simp_all add: Beta_pole1 one_minus_of_nat_in_nonpos_Ints_iff gbinomial_1) + qed + next + case False + have "(z gchoose (Suc n)) = ((z - 1 + 1) gchoose (Suc n))" by simp + also have "\ = (z - 1 gchoose n) * ((z - 1) + 1) / of_nat (Suc n)" + by (subst gbinomial_factors) (simp add: field_simps) + also from False have "\ = inverse (of_nat (Suc n) * Beta (z - of_nat n) (of_nat (Suc n)))" + (is "_ = inverse ?x") by (subst Suc.IH) (simp_all add: field_simps Beta_pole1) + also have "of_nat (Suc n) \ (\\<^sub>\\<^sub>0 :: 'a set)" by (subst of_nat_in_nonpos_Ints_iff) simp_all + hence "?x = (z + 1) * Beta (z - of_nat (Suc n) + 1) (of_nat (Suc n) + 1)" + by (subst Beta_plus1_right [symmetric]) simp_all + finally show ?thesis . + qed +qed + +lemma gbinomial_Gamma: + assumes "z + 1 \ \\<^sub>\\<^sub>0" + shows "(z gchoose n) = Gamma (z + 1) / (fact n * Gamma (z - of_nat n + 1))" +proof - + have "(z gchoose n) = Gamma (z + 2) / (z + 1) / (fact n * Gamma (z - of_nat n + 1))" + by (subst gbinomial_Beta[OF assms]) (simp_all add: Beta_def Gamma_fact [symmetric] add_ac) + also from assms have "Gamma (z + 2) / (z + 1) = Gamma (z + 1)" + using Gamma_plus1[of "z+1"] by (auto simp add: divide_simps mult_ac add_ac) + finally show ?thesis . +qed + + +subsubsection \Integral form\ + +lemma integrable_Gamma_integral_bound: + fixes a c :: real + assumes a: "a > -1" and c: "c \ 0" + defines "f \ \x. if x \ {0..c} then x powr a else exp (-x/2)" + shows "f integrable_on {0..}" +proof - + have "f integrable_on {0..c}" + by (rule integrable_spike_finite[of "{}", OF _ _ integrable_on_powr_from_0[of a c]]) + (insert a c, simp_all add: f_def) + moreover have A: "(\x. exp (-x/2)) integrable_on {c..}" + using integrable_on_exp_minus_to_infinity[of "1/2"] by simp + have "f integrable_on {c..}" + by (rule integrable_spike_finite[of "{c}", OF _ _ A]) (simp_all add: f_def) + ultimately show "f integrable_on {0..}" + by (rule integrable_union') (insert c, auto simp: max_def) +qed + +lemma Gamma_integral_complex: + assumes z: "Re z > 0" + shows "((\t. of_real t powr (z - 1) / of_real (exp t)) has_integral Gamma z) {0..}" +proof - + have A: "((\t. (of_real t) powr (z - 1) * of_real ((1 - t) ^ n)) + has_integral (fact n / pochhammer z (n+1))) {0..1}" + if "Re z > 0" for n z using that + proof (induction n arbitrary: z) + case 0 + have "((\t. complex_of_real t powr (z - 1)) has_integral + (of_real 1 powr z / z - of_real 0 powr z / z)) {0..1}" using 0 + by (intro fundamental_theorem_of_calculus_interior) + (auto intro!: continuous_intros derivative_eq_intros has_vector_derivative_real_complex) + thus ?case by simp + next + case (Suc n) + let ?f = "\t. complex_of_real t powr z / z" + let ?f' = "\t. complex_of_real t powr (z - 1)" + let ?g = "\t. (1 - complex_of_real t) ^ Suc n" + let ?g' = "\t. - ((1 - complex_of_real t) ^ n) * of_nat (Suc n)" + have "((\t. ?f' t * ?g t) has_integral + (of_nat (Suc n)) * fact n / pochhammer z (n+2)) {0..1}" + (is "(_ has_integral ?I) _") + proof (rule integration_by_parts_interior[where f' = ?f' and g = ?g]) + from Suc.prems show "continuous_on {0..1} ?f" "continuous_on {0..1} ?g" + by (auto intro!: continuous_intros) + next + fix t :: real assume t: "t \ {0<..<1}" + show "(?f has_vector_derivative ?f' t) (at t)" using t Suc.prems + by (auto intro!: derivative_eq_intros has_vector_derivative_real_complex) + show "(?g has_vector_derivative ?g' t) (at t)" + by (rule has_vector_derivative_real_complex derivative_eq_intros refl)+ simp_all + next + from Suc.prems have [simp]: "z \ 0" by auto + from Suc.prems have A: "Re (z + of_nat n) > 0" for n by simp + have [simp]: "z + of_nat n \ 0" "z + 1 + of_nat n \ 0" for n + using A[of n] A[of "Suc n"] by (auto simp add: add.assoc simp del: plus_complex.sel) + have "((\x. of_real x powr z * of_real ((1 - x) ^ n) * (- of_nat (Suc n) / z)) has_integral + fact n / pochhammer (z+1) (n+1) * (- of_nat (Suc n) / z)) {0..1}" + (is "(?A has_integral ?B) _") + using Suc.IH[of "z+1"] Suc.prems by (intro has_integral_mult_left) (simp_all add: add_ac pochhammer_rec) + also have "?A = (\t. ?f t * ?g' t)" by (intro ext) (simp_all add: field_simps) + also have "?B = - (of_nat (Suc n) * fact n / pochhammer z (n+2))" + by (simp add: divide_simps pochhammer_rec + setprod_shift_bounds_cl_Suc_ivl del: of_nat_Suc) + finally show "((\t. ?f t * ?g' t) has_integral (?f 1 * ?g 1 - ?f 0 * ?g 0 - ?I)) {0..1}" + by simp + qed (simp_all add: bounded_bilinear_mult) + thus ?case by simp + qed + + have B: "((\t. if t \ {0..of_nat n} then + of_real t powr (z - 1) * (1 - of_real t / of_nat n) ^ n else 0) + has_integral (of_nat n powr z * fact n / pochhammer z (n+1))) {0..}" for n + proof (cases "n > 0") + case [simp]: True + hence [simp]: "n \ 0" by auto + with has_integral_affinity01[OF A[OF z, of n], of "inverse (of_nat n)" 0] + have "((\x. (of_nat n - of_real x) ^ n * (of_real x / of_nat n) powr (z - 1) / of_nat n ^ n) + has_integral fact n * of_nat n / pochhammer z (n+1)) ((\x. real n * x)`{0..1})" + (is "(?f has_integral ?I) ?ivl") by (simp add: field_simps scaleR_conv_of_real) + also from True have "((\x. real n*x)`{0..1}) = {0..real n}" + by (subst image_mult_atLeastAtMost) simp_all + also have "?f = (\x. (of_real x / of_nat n) powr (z - 1) * (1 - of_real x / of_nat n) ^ n)" + using True by (intro ext) (simp add: field_simps) + finally have "((\x. (of_real x / of_nat n) powr (z - 1) * (1 - of_real x / of_nat n) ^ n) + has_integral ?I) {0..real n}" (is ?P) . + also have "?P \ ((\x. exp ((z - 1) * of_real (ln (x / of_nat n))) * (1 - of_real x / of_nat n) ^ n) + has_integral ?I) {0..real n}" + by (intro has_integral_spike_finite_eq[of "{0}"]) (auto simp: powr_def Ln_of_real [symmetric]) + also have "\ \ ((\x. exp ((z - 1) * of_real (ln x - ln (of_nat n))) * (1 - of_real x / of_nat n) ^ n) + has_integral ?I) {0..real n}" + by (intro has_integral_spike_finite_eq[of "{0}"]) (simp_all add: ln_div) + finally have \ . + note B = has_integral_mult_right[OF this, of "exp ((z - 1) * ln (of_nat n))"] + have "((\x. exp ((z - 1) * of_real (ln x)) * (1 - of_real x / of_nat n) ^ n) + has_integral (?I * exp ((z - 1) * ln (of_nat n)))) {0..real n}" (is ?P) + by (insert B, subst (asm) mult.assoc [symmetric], subst (asm) exp_add [symmetric]) + (simp add: Ln_of_nat algebra_simps) + also have "?P \ ((\x. of_real x powr (z - 1) * (1 - of_real x / of_nat n) ^ n) + has_integral (?I * exp ((z - 1) * ln (of_nat n)))) {0..real n}" + by (intro has_integral_spike_finite_eq[of "{0}"]) (simp_all add: powr_def Ln_of_real) + also have "fact n * of_nat n / pochhammer z (n+1) * exp ((z - 1) * Ln (of_nat n)) = + (of_nat n powr z * fact n / pochhammer z (n+1))" + by (auto simp add: powr_def algebra_simps exp_diff) + finally show ?thesis by (subst has_integral_restrict) simp_all + next + case False + thus ?thesis by (subst has_integral_restrict) (simp_all add: has_integral_refl) + qed + + have "eventually (\n. Gamma_series z n = + of_nat n powr z * fact n / pochhammer z (n+1)) sequentially" + using eventually_gt_at_top[of "0::nat"] + by eventually_elim (simp add: powr_def algebra_simps Ln_of_nat Gamma_series_def) + from this and Gamma_series_LIMSEQ[of z] + have C: "(\k. of_nat k powr z * fact k / pochhammer z (k+1)) \ Gamma z" + by (rule Lim_transform_eventually) + + { + fix x :: real assume x: "x \ 0" + have lim_exp: "(\k. (1 - x / real k) ^ k) \ exp (-x)" + using tendsto_exp_limit_sequentially[of "-x"] by simp + have "(\k. of_real x powr (z - 1) * of_real ((1 - x / of_nat k) ^ k)) + \ of_real x powr (z - 1) * of_real (exp (-x))" (is ?P) + by (intro tendsto_intros lim_exp) + also from eventually_gt_at_top[of "nat \x\"] + have "eventually (\k. of_nat k > x) sequentially" by eventually_elim linarith + hence "?P \ (\k. if x \ of_nat k then + of_real x powr (z - 1) * of_real ((1 - x / of_nat k) ^ k) else 0) + \ of_real x powr (z - 1) * of_real (exp (-x))" + by (intro tendsto_cong) (auto elim!: eventually_mono) + finally have \ . + } + hence D: "\x\{0..}. (\k. if x \ {0..real k} then + of_real x powr (z - 1) * (1 - of_real x / of_nat k) ^ k else 0) + \ of_real x powr (z - 1) / of_real (exp x)" + by (simp add: exp_minus field_simps cong: if_cong) + + have "((\x. (Re z - 1) * (ln x / x)) \ (Re z - 1) * 0) at_top" + by (intro tendsto_intros ln_x_over_x_tendsto_0) + hence "((\x. ((Re z - 1) * ln x) / x) \ 0) at_top" by simp + from order_tendstoD(2)[OF this, of "1/2"] + have "eventually (\x. (Re z - 1) * ln x / x < 1/2) at_top" by simp + from eventually_conj[OF this eventually_gt_at_top[of 0]] + obtain x0 where "\x\x0. (Re z - 1) * ln x / x < 1/2 \ x > 0" + by (auto simp: eventually_at_top_linorder) + hence x0: "x0 > 0" "\x. x \ x0 \ (Re z - 1) * ln x < x / 2" by auto + + define h where "h = (\x. if x \ {0..x0} then x powr (Re z - 1) else exp (-x/2))" + have le_h: "x powr (Re z - 1) * exp (-x) \ h x" if x: "x \ 0" for x + proof (cases "x > x0") + case True + from True x0(1) have "x powr (Re z - 1) * exp (-x) = exp ((Re z - 1) * ln x - x)" + by (simp add: powr_def exp_diff exp_minus field_simps exp_add) + also from x0(2)[of x] True have "\ < exp (-x/2)" + by (simp add: field_simps) + finally show ?thesis using True by (auto simp add: h_def) + next + case False + from x have "x powr (Re z - 1) * exp (- x) \ x powr (Re z - 1) * 1" + by (intro mult_left_mono) simp_all + with False show ?thesis by (auto simp add: h_def) + qed + + have E: "\x\{0..}. cmod (if x \ {0..real k} then of_real x powr (z - 1) * + (1 - complex_of_real x / of_nat k) ^ k else 0) \ h x" + (is "\x\_. ?f x \ _") for k + proof safe + fix x :: real assume x: "x \ 0" + { + fix x :: real and n :: nat assume x: "x \ of_nat n" + have "(1 - complex_of_real x / of_nat n) = complex_of_real ((1 - x / of_nat n))" by simp + also have "norm \ = \(1 - x / real n)\" by (subst norm_of_real) (rule refl) + also from x have "\ = (1 - x / real n)" by (intro abs_of_nonneg) (simp_all add: divide_simps) + finally have "cmod (1 - complex_of_real x / of_nat n) = 1 - x / real n" . + } note D = this + from D[of x k] x + have "?f x \ (if of_nat k \ x \ k > 0 then x powr (Re z - 1) * (1 - x / real k) ^ k else 0)" + by (auto simp: norm_mult norm_powr_real_powr norm_power intro!: mult_nonneg_nonneg) + also have "\ \ x powr (Re z - 1) * exp (-x)" + by (auto intro!: mult_left_mono exp_ge_one_minus_x_over_n_power_n) + also from x have "\ \ h x" by (rule le_h) + finally show "?f x \ h x" . + qed + + have F: "h integrable_on {0..}" unfolding h_def + by (rule integrable_Gamma_integral_bound) (insert assms x0(1), simp_all) + show ?thesis + by (rule has_integral_dominated_convergence[OF B F E D C]) +qed + +lemma Gamma_integral_real: + assumes x: "x > (0 :: real)" + shows "((\t. t powr (x - 1) / exp t) has_integral Gamma x) {0..}" +proof - + have A: "((\t. complex_of_real t powr (complex_of_real x - 1) / + complex_of_real (exp t)) has_integral complex_of_real (Gamma x)) {0..}" + using Gamma_integral_complex[of x] assms by (simp_all add: Gamma_complex_of_real powr_of_real) + have "((\t. complex_of_real (t powr (x - 1) / exp t)) has_integral of_real (Gamma x)) {0..}" + by (rule has_integral_eq[OF _ A]) (simp_all add: powr_of_real [symmetric]) + from has_integral_linear[OF this bounded_linear_Re] show ?thesis by (simp add: o_def) +qed + + + +subsection \The Weierstraß product formula for the sine\ + +lemma sin_product_formula_complex: + fixes z :: complex + shows "(\n. of_real pi * z * (\k=1..n. 1 - z^2 / of_nat k^2)) \ sin (of_real pi * z)" +proof - + let ?f = "rGamma_series_weierstrass" + have "(\n. (- of_real pi * inverse z) * (?f z n * ?f (- z) n)) + \ (- of_real pi * inverse z) * (rGamma z * rGamma (- z))" + by (intro tendsto_intros rGamma_weierstrass_complex) + also have "(\n. (- of_real pi * inverse z) * (?f z n * ?f (-z) n)) = + (\n. of_real pi * z * (\k=1..n. 1 - z^2 / of_nat k ^ 2))" + proof + fix n :: nat + have "(- of_real pi * inverse z) * (?f z n * ?f (-z) n) = + of_real pi * z * (\k=1..n. (of_nat k - z) * (of_nat k + z) / of_nat k ^ 2)" + by (simp add: rGamma_series_weierstrass_def mult_ac exp_minus + divide_simps setprod.distrib[symmetric] power2_eq_square) + also have "(\k=1..n. (of_nat k - z) * (of_nat k + z) / of_nat k ^ 2) = + (\k=1..n. 1 - z^2 / of_nat k ^ 2)" + by (intro setprod.cong) (simp_all add: power2_eq_square field_simps) + finally show "(- of_real pi * inverse z) * (?f z n * ?f (-z) n) = of_real pi * z * \" + by (simp add: divide_simps) + qed + also have "(- of_real pi * inverse z) * (rGamma z * rGamma (- z)) = sin (of_real pi * z)" + by (subst rGamma_reflection_complex') (simp add: divide_simps) + finally show ?thesis . +qed + +lemma sin_product_formula_real: + "(\n. pi * (x::real) * (\k=1..n. 1 - x^2 / of_nat k^2)) \ sin (pi * x)" +proof - + from sin_product_formula_complex[of "of_real x"] + have "(\n. of_real pi * of_real x * (\k=1..n. 1 - (of_real x)^2 / (of_nat k)^2)) + \ sin (of_real pi * of_real x :: complex)" (is "?f \ ?y") . + also have "?f = (\n. of_real (pi * x * (\k=1..n. 1 - x^2 / (of_nat k^2))))" by simp + also have "?y = of_real (sin (pi * x))" by (simp only: sin_of_real [symmetric] of_real_mult) + finally show ?thesis by (subst (asm) tendsto_of_real_iff) +qed + +lemma sin_product_formula_real': + assumes "x \ (0::real)" + shows "(\n. (\k=1..n. 1 - x^2 / of_nat k^2)) \ sin (pi * x) / (pi * x)" + using tendsto_divide[OF sin_product_formula_real[of x] tendsto_const[of "pi * x"]] assms + by simp + + +subsection \The Solution to the Basel problem\ + +theorem inverse_squares_sums: "(\n. 1 / (n + 1)\<^sup>2) sums (pi\<^sup>2 / 6)" +proof - + define P where "P x n = (\k=1..n. 1 - x^2 / of_nat k^2)" for x :: real and n + define K where "K = (\n. inverse (real_of_nat (Suc n))^2)" + define f where [abs_def]: "f x = (\n. P x n / of_nat (Suc n)^2)" for x + define g where [abs_def]: "g x = (1 - sin (pi * x) / (pi * x))" for x + + have sums: "(\n. P x n / of_nat (Suc n)^2) sums (if x = 0 then K else g x / x^2)" for x + proof (cases "x = 0") + assume x: "x = 0" + have "summable (\n. inverse ((real_of_nat (Suc n))\<^sup>2))" + using inverse_power_summable[of 2] by (subst summable_Suc_iff) simp + thus ?thesis by (simp add: x g_def P_def K_def inverse_eq_divide power_divide summable_sums) + next + assume x: "x \ 0" + have "(\n. P x n - P x (Suc n)) sums (P x 0 - sin (pi * x) / (pi * x))" + unfolding P_def using x by (intro telescope_sums' sin_product_formula_real') + also have "(\n. P x n - P x (Suc n)) = (\n. (x^2 / of_nat (Suc n)^2) * P x n)" + unfolding P_def by (simp add: setprod_nat_ivl_Suc' algebra_simps) + also have "P x 0 = 1" by (simp add: P_def) + finally have "(\n. x\<^sup>2 / (of_nat (Suc n))\<^sup>2 * P x n) sums (1 - sin (pi * x) / (pi * x))" . + from sums_divide[OF this, of "x^2"] x show ?thesis unfolding g_def by simp + qed + + have "continuous_on (ball 0 1) f" + proof (rule uniform_limit_theorem; (intro always_eventually allI)?) + show "uniform_limit (ball 0 1) (\n x. \k ball 0 1" + { + fix k :: nat assume k: "k \ 1" + from x have "x^2 < 1" by (auto simp: dist_0_norm abs_square_less_1) + also from k have "\ \ of_nat k^2" by simp + finally have "(1 - x^2 / of_nat k^2) \ {0..1}" using k + by (simp_all add: field_simps del: of_nat_Suc) + } + hence "(\k=1..n. abs (1 - x^2 / of_nat k^2)) \ (\k=1..n. 1)" by (intro setprod_mono) simp + thus "norm (P x n / (of_nat (Suc n)^2)) \ 1 / of_nat (Suc n)^2" + unfolding P_def by (simp add: field_simps abs_setprod del: of_nat_Suc) + qed (subst summable_Suc_iff, insert inverse_power_summable[of 2], simp add: inverse_eq_divide) + qed (auto simp: P_def intro!: continuous_intros) + hence "isCont f 0" by (subst (asm) continuous_on_eq_continuous_at) simp_all + hence "(f \ 0 \ f 0)" by (simp add: isCont_def) + also have "f 0 = K" unfolding f_def P_def K_def by (simp add: inverse_eq_divide power_divide) + finally have "f \ 0 \ K" . + + moreover have "f \ 0 \ pi^2 / 6" + proof (rule Lim_transform_eventually) + define f' where [abs_def]: "f' x = (\n. - sin_coeff (n+3) * pi ^ (n+2) * x^n)" for x + have "eventually (\x. x \ (0::real)) (at 0)" + by (auto simp add: eventually_at intro!: exI[of _ 1]) + thus "eventually (\x. f' x = f x) (at 0)" + proof eventually_elim + fix x :: real assume x: "x \ 0" + have "sin_coeff 1 = (1 :: real)" "sin_coeff 2 = (0::real)" by (simp_all add: sin_coeff_def) + with sums_split_initial_segment[OF sums_minus[OF sin_converges], of 3 "pi*x"] + have "(\n. - (sin_coeff (n+3) * (pi*x)^(n+3))) sums (pi * x - sin (pi*x))" + by (simp add: eval_nat_numeral) + from sums_divide[OF this, of "x^3 * pi"] x + have "(\n. - (sin_coeff (n+3) * pi^(n+2) * x^n)) sums ((1 - sin (pi*x) / (pi*x)) / x^2)" + by (simp add: divide_simps eval_nat_numeral power_mult_distrib mult_ac) + with x have "(\n. - (sin_coeff (n+3) * pi^(n+2) * x^n)) sums (g x / x^2)" + by (simp add: g_def) + hence "f' x = g x / x^2" by (simp add: sums_iff f'_def) + also have "\ = f x" using sums[of x] x by (simp add: sums_iff g_def f_def) + finally show "f' x = f x" . + qed + + have "isCont f' 0" unfolding f'_def + proof (intro isCont_powser_converges_everywhere) + fix x :: real show "summable (\n. -sin_coeff (n+3) * pi^(n+2) * x^n)" + proof (cases "x = 0") + assume x: "x \ 0" + from summable_divide[OF sums_summable[OF sums_split_initial_segment[OF + sin_converges[of "pi*x"]], of 3], of "-pi*x^3"] x + show ?thesis by (simp add: mult_ac power_mult_distrib divide_simps eval_nat_numeral) + qed (simp only: summable_0_powser) + qed + hence "f' \ 0 \ f' 0" by (simp add: isCont_def) + also have "f' 0 = pi * pi / fact 3" unfolding f'_def + by (subst powser_zero) (simp add: sin_coeff_def) + finally show "f' \ 0 \ pi^2 / 6" by (simp add: eval_nat_numeral) + qed + + ultimately have "K = pi^2 / 6" by (rule LIM_unique) + moreover from inverse_power_summable[of 2] + have "summable (\n. (inverse (real_of_nat (Suc n)))\<^sup>2)" + by (subst summable_Suc_iff) (simp add: power_inverse) + ultimately show ?thesis unfolding K_def + by (auto simp add: sums_iff power_divide inverse_eq_divide) +qed + +end diff -r bbcb05504fdc -r bd218a9320b5 src/HOL/Multivariate_Analysis/Generalised_Binomial_Theorem.thy --- a/src/HOL/Multivariate_Analysis/Generalised_Binomial_Theorem.thy Thu Aug 04 18:45:28 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Generalised_Binomial_Theorem.thy Thu Aug 04 19:36:31 2016 +0200 @@ -11,10 +11,10 @@ \ theory Generalised_Binomial_Theorem -imports - Complex_Main +imports + Complex_Main Complex_Transcendental - Summation + Summation_Tests begin lemma gbinomial_ratio_limit: @@ -36,14 +36,14 @@ by (simp add: setprod.atLeast0_atMost_Suc atLeastLessThanSuc_atLeastAtMost) also have "?P / \ = (?P / ?P) / (a - of_nat n)" by (rule divide_divide_eq_left[symmetric]) also from assms have "?P / ?P = 1" by auto - also have "of_nat (Suc n) * (1 / (a - of_nat n)) = + also have "of_nat (Suc n) * (1 / (a - of_nat n)) = inverse (inverse (of_nat (Suc n)) * (a - of_nat n))" by (simp add: field_simps) also have "inverse (of_nat (Suc n)) * (a - of_nat n) = a / of_nat (Suc n) - of_nat n / of_nat (Suc n)" by (simp add: field_simps del: of_nat_Suc) finally show "?f n = (a gchoose n) / (a gchoose Suc n)" by simp qed - have "(\n. norm a / (of_nat (Suc n))) \ 0" + have "(\n. norm a / (of_nat (Suc n))) \ 0" unfolding divide_inverse by (intro tendsto_mult_right_zero LIMSEQ_inverse_real_of_nat) hence "(\n. a / of_nat (Suc n)) \ 0" @@ -84,26 +84,26 @@ with K have summable: "summable (\n. ?f n * z ^ n)" if "norm z < K" for z using that by auto hence summable': "summable (\n. ?f' n * z ^ n)" if "norm z < K" for z using that by (intro termdiff_converges[of _ K]) simp_all - + define f f' where [abs_def]: "f z = (\n. ?f n * z ^ n)" "f' z = (\n. ?f' n * z ^ n)" for z { fix z :: complex assume z: "norm z < K" from summable_mult2[OF summable'[OF z], of z] have summable1: "summable (\n. ?f' n * z ^ Suc n)" by (simp add: mult_ac) - hence summable2: "summable (\n. of_nat n * ?f n * z^n)" + hence summable2: "summable (\n. of_nat n * ?f n * z^n)" unfolding diffs_def by (subst (asm) summable_Suc_iff) have "(1 + z) * f' z = (\n. ?f' n * z^n) + (\n. ?f' n * z^Suc n)" unfolding f_f'_def using summable' z by (simp add: algebra_simps suminf_mult) also have "(\n. ?f' n * z^n) = (\n. of_nat (Suc n) * ?f (Suc n) * z^n)" by (intro suminf_cong) (simp add: diffs_def) - also have "(\n. ?f' n * z^Suc n) = (\n. of_nat n * ?f n * z ^ n)" + also have "(\n. ?f' n * z^Suc n) = (\n. of_nat n * ?f n * z ^ n)" using summable1 suminf_split_initial_segment[OF summable1] unfolding diffs_def by (subst suminf_split_head, subst (asm) summable_Suc_iff) simp_all also have "(\n. of_nat (Suc n) * ?f (Suc n) * z^n) + (\n. of_nat n * ?f n * z^n) = (\n. a * ?f n * z^n)" by (subst gbinomial_mult_1, subst suminf_add) - (insert summable'[OF z] summable2, + (insert summable'[OF z] summable2, simp_all add: summable_powser_split_head algebra_simps diffs_def) also have "\ = a * f z" unfolding f_f'_def by (subst suminf_mult[symmetric]) (simp_all add: summable[OF z] mult_ac) @@ -124,16 +124,16 @@ with K have nz: "1 + z \ 0" by (auto dest!: minus_unique) from z K have "norm z < 1" by simp hence "(1 + z) \ \\<^sub>\\<^sub>0" by (cases z) (auto simp: complex_nonpos_Reals_iff) - hence "((\z. f z * (1 + z) powr (-a)) has_field_derivative + hence "((\z. f z * (1 + z) powr (-a)) has_field_derivative f' z * (1 + z) powr (-a) - a * f z * (1 + z) powr (-a-1)) (at z)" using z by (auto intro!: derivative_eq_intros) also from z have "a * f z = (1 + z) * f' z" by (rule deriv) - finally show "((\z. f z * (1 + z) powr (-a)) has_field_derivative 0) (at z within ball 0 K)" + finally show "((\z. f z * (1 + z) powr (-a)) has_field_derivative 0) (at z within ball 0 K)" using nz by (simp add: field_simps powr_diff_complex at_within_open[OF z']) qed simp_all then obtain c where c: "\z. z \ ball 0 K \ f z * (1 + z) powr (-a) = c" by blast from c[of 0] and K have "c = 1" by simp - with c[of z] have "f z = (1 + z) powr a" using K + with c[of z] have "f z = (1 + z) powr a" using K by (simp add: powr_minus_complex field_simps dist_complex_def) with summable K show ?thesis unfolding f_f'_def by (simp add: sums_iff) qed @@ -141,7 +141,7 @@ lemma gen_binomial_complex': fixes x y :: real and a :: complex assumes "\x\ < \y\" - shows "(\n. (a gchoose n) * of_real x^n * of_real y powr (a - of_nat n)) sums + shows "(\n. (a gchoose n) * of_real x^n * of_real y powr (a - of_nat n)) sums of_real (x + y) powr a" (is "?P x y") proof - { @@ -150,7 +150,7 @@ note xy = xy this from xy have "(\n. (a gchoose n) * of_real (x / y) ^ n) sums (1 + of_real (x / y)) powr a" by (intro gen_binomial_complex) (simp add: norm_divide) - hence "(\n. (a gchoose n) * of_real (x / y) ^ n * y powr a) sums + hence "(\n. (a gchoose n) * of_real (x / y) ^ n * y powr a) sums ((1 + of_real (x / y)) powr a * y powr a)" by (rule sums_mult2) also have "(1 + complex_of_real (x / y)) = complex_of_real (1 + x/y)" by simp @@ -172,7 +172,7 @@ by (subst powr_neg_real_complex) (simp add: abs_real_def split: if_split_asm) also { fix n :: nat - from y have "(a gchoose n) * of_real (-x) ^ n * of_real (-y) powr (a - of_nat n) = + from y have "(a gchoose n) * of_real (-x) ^ n * of_real (-y) powr (a - of_nat n) = (a gchoose n) * (-of_real x / -of_real y) ^ n * (- of_real y) powr a" by (subst power_divide) (simp add: powr_diff_complex powr_nat) also from y have "(- of_real y) powr a = (-1) powr -a * of_real y powr a" @@ -180,7 +180,7 @@ also have "-complex_of_real x / -complex_of_real y = complex_of_real x / complex_of_real y" by simp also have "... ^ n = of_real x ^ n / of_real y ^ n" by (simp add: power_divide) - also have "(a gchoose n) * ... * ((-1) powr -a * of_real y powr a) = + also have "(a gchoose n) * ... * ((-1) powr -a * of_real y powr a) = (-1) powr -a * ((a gchoose n) * of_real x ^ n * of_real y powr (a - n))" by (simp add: algebra_simps powr_diff_complex powr_nat) finally have "(a gchoose n) * of_real (- x) ^ n * of_real (- y) powr (a - of_nat n) = @@ -194,7 +194,7 @@ lemma gen_binomial_complex'': fixes x y :: real and a :: complex assumes "\y\ < \x\" - shows "(\n. (a gchoose n) * of_real x powr (a - of_nat n) * of_real y ^ n) sums + shows "(\n. (a gchoose n) * of_real x powr (a - of_nat n) * of_real y ^ n) sums of_real (x + y) powr a" using gen_binomial_complex'[OF assms] by (simp add: mult_ac add.commute) @@ -209,12 +209,12 @@ (of_real (1 + z)) powr (of_real a)" by simp also have "(of_real (1 + z) :: complex) powr (of_real a) = of_real ((1 + z) powr a)" using assms by (subst powr_of_real) simp_all - also have "(of_real a gchoose n :: complex) = of_real (a gchoose n)" for n + also have "(of_real a gchoose n :: complex) = of_real (a gchoose n)" for n by (simp add: gbinomial_setprod_rev) hence "(\n. (of_real a gchoose n :: complex) * of_real z ^ n) = (\n. of_real ((a gchoose n) * z ^ n))" by (intro ext) simp finally show ?thesis by (simp only: sums_of_real_iff) -qed +qed lemma gen_binomial_real': fixes x y a :: real @@ -228,7 +228,7 @@ by (rule gen_binomial_real) hence "(\n. (a gchoose n) * (x / y) ^ n * y powr a) sums ((1 + x / y) powr a * y powr a)" by (rule sums_mult2) - with xy show ?thesis + with xy show ?thesis by (simp add: field_simps powr_divide powr_divide2[symmetric] powr_realpow) qed @@ -245,7 +245,7 @@ using gen_binomial_real'[OF assms] by (simp add: mult_ac add.commute) lemma sqrt_series': - "\z\ < a \ (\n. ((1/2) gchoose n) * a powr (1/2 - real_of_nat n) * z ^ n) sums + "\z\ < a \ (\n. ((1/2) gchoose n) * a powr (1/2 - real_of_nat n) * z ^ n) sums sqrt (a + z :: real)" using gen_binomial_real''[of z a "1/2"] by (simp add: powr_half_sqrt) diff -r bbcb05504fdc -r bd218a9320b5 src/HOL/Multivariate_Analysis/Harmonic_Numbers.thy --- a/src/HOL/Multivariate_Analysis/Harmonic_Numbers.thy Thu Aug 04 18:45:28 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Harmonic_Numbers.thy Thu Aug 04 19:36:31 2016 +0200 @@ -5,15 +5,15 @@ section \Harmonic Numbers\ theory Harmonic_Numbers -imports +imports Complex_Transcendental - Summation + Summation_Tests Integral_Test begin text \ The definition of the Harmonic Numbers and the Euler-Mascheroni constant. - Also provides a reasonably accurate approximation of @{term "ln 2 :: real"} + Also provides a reasonably accurate approximation of @{term "ln 2 :: real"} and the Euler-Mascheroni constant. \ @@ -51,11 +51,11 @@ lemma of_real_harm: "of_real (harm n) = harm n" unfolding harm_def by simp - + lemma norm_harm: "norm (harm n) = harm n" by (subst of_real_harm [symmetric]) (simp add: harm_nonneg) -lemma harm_expand: +lemma harm_expand: "harm 0 = 0" "harm (Suc 0) = 1" "harm (numeral n) = harm (pred_numeral n) + inverse (numeral n)" @@ -134,7 +134,7 @@ has_field_derivative_iff_has_vector_derivative[symmetric]) hence "integral {0..of_nat n} (\x. inverse (x + 1) :: real) = ln (of_nat (Suc n))" by (auto dest!: integral_unique) - ultimately show ?thesis + ultimately show ?thesis by (simp add: euler_mascheroni.sum_integral_diff_series_def atLeast0AtMost) qed @@ -151,7 +151,7 @@ lemma euler_mascheroni_convergent: "convergent (\n. harm n - ln (of_nat n) :: real)" proof - - have A: "(\n. harm (Suc n) - ln (of_nat (Suc n))) = + have A: "(\n. harm (Suc n) - ln (of_nat (Suc n))) = euler_mascheroni.sum_integral_diff_series" by (subst euler_mascheroni_sum_integral_diff_series [symmetric]) (rule refl) have "convergent (\n. harm (Suc n) - ln (of_nat (Suc n) :: real))" @@ -159,13 +159,13 @@ thus ?thesis by (subst (asm) convergent_Suc_iff) qed -lemma euler_mascheroni_LIMSEQ: +lemma euler_mascheroni_LIMSEQ: "(\n. harm n - ln (of_nat n) :: real) \ euler_mascheroni" unfolding euler_mascheroni_def by (simp add: convergent_LIMSEQ_iff [symmetric] euler_mascheroni_convergent) -lemma euler_mascheroni_LIMSEQ_of_real: - "(\n. of_real (harm n - ln (of_nat n))) \ +lemma euler_mascheroni_LIMSEQ_of_real: + "(\n. of_real (harm n - ln (of_nat n))) \ (euler_mascheroni :: 'a :: {real_normed_algebra_1, topological_space})" proof - have "(\n. of_real (harm n - ln (of_nat n))) \ (of_real (euler_mascheroni) :: 'a)" @@ -175,7 +175,7 @@ lemma euler_mascheroni_sum: "(\n. inverse (of_nat (n+1)) + ln (of_nat (n+1)) - ln (of_nat (n+2)) :: real) - sums euler_mascheroni" + sums euler_mascheroni" using sums_add[OF telescope_sums[OF LIMSEQ_Suc[OF euler_mascheroni_LIMSEQ]] telescope_sums'[OF LIMSEQ_inverse_real_of_nat]] by (simp_all add: harm_def algebra_simps) @@ -198,21 +198,21 @@ by (intro setsum.mono_neutral_right) auto also have "\ = (\k|k<2*n \ odd k. 2 / (real_of_nat (Suc k)))" by (intro setsum.cong) auto - also have "(\k|k<2*n \ odd k. 2 / (real_of_nat (Suc k))) = harm n" + also have "(\k|k<2*n \ odd k. 2 / (real_of_nat (Suc k))) = harm n" unfolding harm_altdef by (intro setsum.reindex_cong[of "\n. 2*n+1"]) (auto simp: inj_on_def field_simps elim!: oddE) also have "harm (2*n) - harm n = ?em (2*n) - ?em n + ln 2" using n by (simp_all add: algebra_simps ln_mult) finally show "?em (2*n) - ?em n + ln 2 = (\k<2*n. (-1)^k / real_of_nat (Suc k))" .. qed - moreover have "(\n. ?em (2*n) - ?em n + ln (2::real)) + moreover have "(\n. ?em (2*n) - ?em n + ln (2::real)) \ euler_mascheroni - euler_mascheroni + ln 2" by (intro tendsto_intros euler_mascheroni_LIMSEQ filterlim_compose[OF euler_mascheroni_LIMSEQ] filterlim_subseq) (auto simp: subseq_def) hence "(\n. ?em (2*n) - ?em n + ln (2::real)) \ ln 2" by simp ultimately have "(\n. (\k<2*n. (-1)^k / real_of_nat (Suc k))) \ ln 2" by (rule Lim_transform_eventually) - + moreover have "summable (\k. (-1)^k * inverse (real_of_nat (Suc k)))" using LIMSEQ_inverse_real_of_nat by (intro summable_Leibniz(1) decseq_imp_monoseq decseq_SucI) simp_all @@ -225,12 +225,12 @@ with A show ?thesis by (simp add: sums_def) qed -lemma alternating_harmonic_series_sums': +lemma alternating_harmonic_series_sums': "(\k. inverse (real_of_nat (2*k+1)) - inverse (real_of_nat (2*k+2))) sums ln 2" unfolding sums_def proof (rule Lim_transform_eventually) show "(\n. \k<2*n. (-1)^k / (real_of_nat (Suc k))) \ ln 2" - using alternating_harmonic_series_sums unfolding sums_def + using alternating_harmonic_series_sums unfolding sums_def by (rule filterlim_compose) (rule mult_nat_left_at_top, simp) show "eventually (\n. (\k<2*n. (-1)^k / (real_of_nat (Suc k))) = (\kkBounds on the Euler--Mascheroni constant\ @@ -254,16 +254,16 @@ have f'_nonpos: "f' \ 0" using assms by (simp add: f'_def divide_simps) let ?f = "\t. (t - x) * f' + inverse x" let ?F = "\t. (t - x)^2 * f' / 2 + t * inverse x" - have diff: "\t\{x..x+a}. (?F has_vector_derivative ?f t) + have diff: "\t\{x..x+a}. (?F has_vector_derivative ?f t) (at t within {x..x+a})" using assms - by (auto intro!: derivative_eq_intros + by (auto intro!: derivative_eq_intros simp: has_field_derivative_iff_has_vector_derivative[symmetric]) from assms have "(?f has_integral (?F (x+a) - ?F x)) {x..x+a}" by (intro fundamental_theorem_of_calculus[OF _ diff]) (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric] field_simps intro!: derivative_eq_intros) also have "?F (x+a) - ?F x = (a*2 + f'*a\<^sup>2*x) / (2*x)" using assms by (simp add: field_simps) - also have "f'*a^2 = - (a^2) / (x*(x + a))" using assms + also have "f'*a^2 = - (a^2) / (x*(x + a))" using assms by (simp add: divide_simps f'_def power2_eq_square) also have "(a*2 + - a\<^sup>2/(x*(x+a))*x) / (2*x) = ?A" using assms by (simp add: divide_simps power2_eq_square) (simp add: algebra_simps) @@ -281,7 +281,7 @@ have "inverse t = inverse ((1 - (t - x) / a) *\<^sub>R x + ((t - x) / a) *\<^sub>R (x + a))" (is "_ = ?A") using assms t' by (simp add: field_simps) also from assms have "convex_on {x..x+a} inverse" by (intro convex_on_inverse) auto - from convex_onD_Icc[OF this _ t] assms + from convex_onD_Icc[OF this _ t] assms have "?A \ (1 - (t - x) / a) * inverse x + (t - x) / a * inverse (x + a)" by simp also have "\ = (t - x) * f' + inverse x" using assms by (simp add: f'_def divide_simps) (simp add: f'_def field_simps) @@ -305,7 +305,7 @@ by (intro fundamental_theorem_of_calculus) (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric] divide_simps intro!: derivative_eq_intros) - also from m have "?F y - ?F x = ((y - m)^2 - (x - m)^2) * f' / 2 + (y - x) / m" + also from m have "?F y - ?F x = ((y - m)^2 - (x - m)^2) * f' / 2 + (y - x) / m" by (simp add: field_simps) also have "((y - m)^2 - (x - m)^2) = 0" by (simp add: m_def power2_eq_square field_simps) also have "0 * f' / 2 + (y - x) / m = ?A" by (simp add: m_def) @@ -332,7 +332,7 @@ qed -lemma euler_mascheroni_lower: +lemma euler_mascheroni_lower: "euler_mascheroni \ harm (Suc n) - ln (real_of_nat (n + 2)) + 1/real_of_nat (2 * (n + 2))" and euler_mascheroni_upper: "euler_mascheroni \ harm (Suc n) - ln (real_of_nat (n + 2)) + 1/real_of_nat (2 * (n + 1))" @@ -370,11 +370,11 @@ also have "\ = D k" unfolding D_def inv_def .. finally show "D (k' + Suc n) \ (inv (k' + Suc n + 1) - inv (k' + Suc n + 2)) / 2" by (simp add: k_def) - from sums_summable[OF sums] + from sums_summable[OF sums] show "summable (\k. (inv (k + Suc n + 1) - inv (k + Suc n + 2))/2)" by simp qed also from sums have "\ = -inv (n+2) / 2" by (simp add: sums_iff) - finally have "euler_mascheroni \ (\k\n. D k) + 1 / (of_nat (2 * (n+2)))" + finally have "euler_mascheroni \ (\k\n. D k) + 1 / (of_nat (2 * (n+2)))" by (simp add: inv_def field_simps) also have "(\k\n. D k) = harm (Suc n) - (\k\n. ln (real_of_nat (Suc k+1)) - ln (of_nat (k+1)))" unfolding harm_altdef D_def by (subst lessThan_Suc_atMost) (simp add: setsum.distrib setsum_subtractf) @@ -382,7 +382,7 @@ by (subst atLeast0AtMost [symmetric], subst setsum_Suc_diff) simp_all finally show "euler_mascheroni \ harm (Suc n) - ln (real_of_nat (n + 2)) + 1/real_of_nat (2 * (n + 2))" by simp - + note sum also have "-(\k. D (k + Suc n)) \ -(\k. (inv (Suc (k + n)) - inv (Suc (Suc k + n)))/2)" proof (intro le_imp_neg_le suminf_le allI summable_ignore_initial_segment[OF summable]) @@ -393,22 +393,22 @@ from ln_inverse_approx_ge[of "of_nat k + 1" "of_nat k + 2"] have "2 / (2 * real_of_nat k + 3) \ ln (of_nat (k+2)) - ln (real_of_nat (k+1))" by (simp add: add_ac) - hence "D k \ 1 / real_of_nat (k+1) - 2 / (2 * real_of_nat k + 3)" + hence "D k \ 1 / real_of_nat (k+1) - 2 / (2 * real_of_nat k + 3)" by (simp add: D_def inverse_eq_divide inv_def) also have "\ = inv ((k+1)*(2*k+3))" unfolding inv_def by (simp add: field_simps) also have "\ \ inv (2*k*(k+1))" unfolding inv_def using k - by (intro le_imp_inverse_le) + by (intro le_imp_inverse_le) (simp add: algebra_simps, simp del: of_nat_add) also have "\ = (inv k - inv (k+1))/2" unfolding inv_def using k by (simp add: divide_simps del: of_nat_mult) (simp add: algebra_simps) finally show "D k \ (inv (Suc (k' + n)) - inv (Suc (Suc k' + n)))/2" unfolding k_def by simp next - from sums_summable[OF sums'] + from sums_summable[OF sums'] show "summable (\k. (inv (Suc (k + n)) - inv (Suc (Suc k + n)))/2)" by simp qed also from sums' have "(\k. (inv (Suc (k + n)) - inv (Suc (Suc k + n)))/2) = inv (n+1)/2" by (simp add: sums_iff) - finally have "euler_mascheroni \ (\k\n. D k) + 1 / of_nat (2 * (n+1))" + finally have "euler_mascheroni \ (\k\n. D k) + 1 / of_nat (2 * (n+1))" by (simp add: inv_def field_simps) also have "(\k\n. D k) = harm (Suc n) - (\k\n. ln (real_of_nat (Suc k+1)) - ln (of_nat (k+1)))" unfolding harm_altdef D_def by (subst lessThan_Suc_atMost) (simp add: setsum.distrib setsum_subtractf) @@ -428,7 +428,7 @@ fixes n :: nat and x :: real defines "y \ (x-1)/(x+1)" assumes x: "x > 0" "x \ 1" - shows "inverse (2*y^(2*n+1)) * (ln x - (\k + shows "inverse (2*y^(2*n+1)) * (ln x - (\k {0..(1 / (1 - y^2) / of_nat (2*n+1))}" proof - from x have norm_y: "norm y < 1" unfolding y_def by simp @@ -468,21 +468,21 @@ and ln_approx_abs: "abs (ln x - (approx + d)) \ d" proof - define c where "c = 2*y^(2*n+1)" - from x have c_pos: "c > 0" unfolding c_def y_def + from x have c_pos: "c > 0" unfolding c_def y_def by (intro mult_pos_pos zero_less_power) simp_all have A: "inverse c * (ln x - (\k {0.. (1 / (1 - y^2) / of_nat (2*n+1))}" using assms unfolding y_def c_def by (intro ln_approx_aux) simp_all hence "inverse c * (ln x - (\k (1 / (1-y^2) / of_nat (2*n+1))" by simp - hence "(ln x - (\k (1 / (1 - y^2) / of_nat (2*n+1))" + hence "(ln x - (\k (1 / (1 - y^2) / of_nat (2*n+1))" by (auto simp add: divide_simps) with c_pos have "ln x \ c / (1 - y^2) / of_nat (2*n+1) + approx" by (subst (asm) pos_divide_le_eq) (simp_all add: mult_ac approx_def) moreover { from A c_pos have "0 \ c * (inverse c * (ln x - (\k = (c * inverse c) * (ln x - (\k = (c * inverse c) * (ln x - (\k approx" by (simp add: approx_def) @@ -501,7 +501,7 @@ lemma euler_mascheroni_bounds': fixes n :: nat assumes "n \ 1" "ln (real_of_nat (Suc n)) \ {l<.. + shows "euler_mascheroni \ {harm n - u + inverse (of_nat (2*(n+1)))<.. -lemma ln2_ge_two_thirds: "2/3 \ ln (2::real)" +lemma ln2_ge_two_thirds: "2/3 \ ln (2::real)" and ln2_le_25_over_36: "ln (2::real) \ 25/36" using ln_approx_bounds[of 2 1, simplified, simplified eval_nat_numeral, simplified] by simp_all text \ - Approximation of the Euler--Mascheroni constant. The lower bound is accurate to about 0.0015; + Approximation of the Euler--Mascheroni constant. The lower bound is accurate to about 0.0015; the upper bound is accurate to about 0.015. \ lemma euler_mascheroni_gt_19_over_33: "(euler_mascheroni :: real) > 19/33" (is ?th1) diff -r bbcb05504fdc -r bd218a9320b5 src/HOL/Multivariate_Analysis/Henstock_Kurzweil_Integration.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Multivariate_Analysis/Henstock_Kurzweil_Integration.thy Thu Aug 04 19:36:31 2016 +0200 @@ -0,0 +1,12645 @@ +(* Author: John Harrison + Author: Robert Himmelmann, TU Muenchen (Translation from HOL light); proofs reworked by LCP +*) + +section \Henstock-Kurzweil gauge integration in many dimensions.\ + +theory Henstock_Kurzweil_Integration +imports + Derivative + Uniform_Limit + "~~/src/HOL/Library/Indicator_Function" +begin + +lemmas scaleR_simps = scaleR_zero_left scaleR_minus_left scaleR_left_diff_distrib + scaleR_zero_right scaleR_minus_right scaleR_right_diff_distrib scaleR_eq_0_iff + scaleR_cancel_left scaleR_cancel_right scaleR_add_right scaleR_add_left real_vector_class.scaleR_one + + +subsection \Sundries\ + +lemma conjunctD2: assumes "a \ b" shows a b using assms by auto +lemma conjunctD3: assumes "a \ b \ c" shows a b c using assms by auto +lemma conjunctD4: assumes "a \ b \ c \ d" shows a b c d using assms by auto + +declare norm_triangle_ineq4[intro] + +lemma simple_image: "{f x |x . x \ s} = f ` s" + by blast + +lemma linear_simps: + assumes "bounded_linear f" + shows + "f (a + b) = f a + f b" + "f (a - b) = f a - f b" + "f 0 = 0" + "f (- a) = - f a" + "f (s *\<^sub>R v) = s *\<^sub>R (f v)" +proof - + interpret f: bounded_linear f by fact + show "f (a + b) = f a + f b" by (rule f.add) + show "f (a - b) = f a - f b" by (rule f.diff) + show "f 0 = 0" by (rule f.zero) + show "f (- a) = - f a" by (rule f.minus) + show "f (s *\<^sub>R v) = s *\<^sub>R (f v)" by (rule f.scaleR) +qed + +lemma bounded_linearI: + assumes "\x y. f (x + y) = f x + f y" + and "\r x. f (r *\<^sub>R x) = r *\<^sub>R f x" + and "\x. norm (f x) \ norm x * K" + shows "bounded_linear f" + using assms by (rule bounded_linear_intro) (* FIXME: duplicate *) + +lemma bounded_linear_component [intro]: "bounded_linear (\x::'a::euclidean_space. x \ k)" + by (rule bounded_linear_inner_left) + +lemma transitive_stepwise_lt_eq: + assumes "(\x y z::nat. R x y \ R y z \ R x z)" + shows "((\m. \n>m. R m n) \ (\n. R n (Suc n)))" + (is "?l = ?r") +proof safe + assume ?r + fix n m :: nat + assume "m < n" + then show "R m n" + proof (induct n arbitrary: m) + case 0 + then show ?case by auto + next + case (Suc n) + show ?case + proof (cases "m < n") + case True + show ?thesis + apply (rule assms[OF Suc(1)[OF True]]) + using \?r\ + apply auto + done + next + case False + then have "m = n" + using Suc(2) by auto + then show ?thesis + using \?r\ by auto + qed + qed +qed auto + +lemma transitive_stepwise_gt: + assumes "\x y z. R x y \ R y z \ R x z" "\n. R n (Suc n)" + shows "\n>m. R m n" +proof - + have "\m. \n>m. R m n" + apply (subst transitive_stepwise_lt_eq) + apply (blast intro: assms)+ + done + then show ?thesis by auto +qed + +lemma transitive_stepwise_le_eq: + assumes "\x. R x x" "\x y z. R x y \ R y z \ R x z" + shows "(\m. \n\m. R m n) \ (\n. R n (Suc n))" + (is "?l = ?r") +proof safe + assume ?r + fix m n :: nat + assume "m \ n" + then show "R m n" + proof (induct n arbitrary: m) + case 0 + with assms show ?case by auto + next + case (Suc n) + show ?case + proof (cases "m \ n") + case True + with Suc.hyps \\n. R n (Suc n)\ assms show ?thesis + by blast + next + case False + then have "m = Suc n" + using Suc(2) by auto + then show ?thesis + using assms(1) by auto + qed + qed +qed auto + +lemma transitive_stepwise_le: + assumes "\x. R x x" "\x y z. R x y \ R y z \ R x z" + and "\n. R n (Suc n)" + shows "\n\m. R m n" +proof - + have "\m. \n\m. R m n" + apply (subst transitive_stepwise_le_eq) + apply (blast intro: assms)+ + done + then show ?thesis by auto +qed + + +subsection \Some useful lemmas about intervals.\ + +lemma empty_as_interval: "{} = cbox One (0::'a::euclidean_space)" + using nonempty_Basis + by (fastforce simp add: set_eq_iff mem_box) + +lemma interior_subset_union_intervals: + assumes "i = cbox a b" + and "j = cbox c d" + and "interior j \ {}" + and "i \ j \ s" + and "interior i \ interior j = {}" + shows "interior i \ interior s" +proof - + have "box a b \ cbox c d = {}" + using inter_interval_mixed_eq_empty[of c d a b] and assms(3,5) + unfolding assms(1,2) interior_cbox by auto + moreover + have "box a b \ cbox c d \ s" + apply (rule order_trans,rule box_subset_cbox) + using assms(4) unfolding assms(1,2) + apply auto + done + ultimately + show ?thesis + unfolding assms interior_cbox + by auto (metis IntI UnE empty_iff interior_maximal open_box subsetCE subsetI) +qed + +lemma inter_interior_unions_intervals: + fixes f::"('a::euclidean_space) set set" + assumes "finite f" + and "open s" + and "\t\f. \a b. t = cbox a b" + and "\t\f. s \ (interior t) = {}" + shows "s \ interior (\f) = {}" +proof (clarsimp simp only: all_not_in_conv [symmetric]) + fix x + assume x: "x \ s" "x \ interior (\f)" + have lem1: "\x e s U. ball x e \ s \ interior U \ ball x e \ s \ U" + using interior_subset + by auto (meson Topology_Euclidean_Space.open_ball contra_subsetD interior_maximal mem_ball) + have "\t\f. \x. \e>0. ball x e \ s \ t" + if "finite f" and "\t\f. \a b. t = cbox a b" and "\x. x \ s \ interior (\f)" for f + using that + proof (induct rule: finite_induct) + case empty + obtain x where "x \ s \ interior (\{})" + using empty(2) .. + then have False + unfolding Union_empty interior_empty by auto + then show ?case by auto + next + case (insert i f) + obtain x where x: "x \ s \ interior (\insert i f)" + using insert(5) .. + then obtain e where e: "0 < e \ ball x e \ s \ interior (\insert i f)" + unfolding open_contains_ball_eq[OF open_Int[OF assms(2) open_interior], rule_format] .. + obtain a where "\b. i = cbox a b" + using insert(4)[rule_format,OF insertI1] .. + then obtain b where ab: "i = cbox a b" .. + show ?case + proof (cases "x \ i") + case False + then have "x \ UNIV - cbox a b" + unfolding ab by auto + then obtain d where "0 < d \ ball x d \ UNIV - cbox a b" + unfolding open_contains_ball_eq[OF open_Diff[OF open_UNIV closed_cbox],rule_format] .. + then have "0 < d" "ball x (min d e) \ UNIV - i" + unfolding ab ball_min_Int by auto + then have "ball x (min d e) \ s \ interior (\f)" + using e unfolding lem1 unfolding ball_min_Int by auto + then have "x \ s \ interior (\f)" using \d>0\ e by auto + then have "\t\f. \x e. 0 < e \ ball x e \ s \ t" + using insert.hyps(3) insert.prems(1) by blast + then show ?thesis by auto + next + case True show ?thesis + proof (cases "x\box a b") + case True + then obtain d where "0 < d \ ball x d \ box a b" + unfolding open_contains_ball_eq[OF open_box,rule_format] .. + then show ?thesis + apply (rule_tac x=i in bexI, rule_tac x=x in exI, rule_tac x="min d e" in exI) + unfolding ab + using box_subset_cbox[of a b] and e + apply fastforce+ + done + next + case False + then obtain k where "x\k \ a\k \ x\k \ b\k" and k: "k \ Basis" + unfolding mem_box by (auto simp add: not_less) + then have "x\k = a\k \ x\k = b\k" + using True unfolding ab and mem_box + apply (erule_tac x = k in ballE) + apply auto + done + then have "\x. ball x (e/2) \ s \ (\f)" + proof (rule disjE) + let ?z = "x - (e/2) *\<^sub>R k" + assume as: "x\k = a\k" + have "ball ?z (e / 2) \ i = {}" + proof (clarsimp simp only: all_not_in_conv [symmetric]) + fix y + assume "y \ ball ?z (e / 2)" and yi: "y \ i" + then have "dist ?z y < e/2" by auto + then have "\(?z - y) \ k\ < e/2" + using Basis_le_norm[OF k, of "?z - y"] unfolding dist_norm by auto + then have "y\k < a\k" + using e k + by (auto simp add: field_simps abs_less_iff as inner_simps) + then have "y \ i" + unfolding ab mem_box by (auto intro!: bexI[OF _ k]) + then show False using yi by auto + qed + moreover + have "ball ?z (e/2) \ s \ (\insert i f)" + apply (rule order_trans[OF _ e[THEN conjunct2, unfolded lem1]]) + proof + fix y + assume as: "y \ ball ?z (e/2)" + have "norm (x - y) \ \e\ / 2 + norm (x - y - (e / 2) *\<^sub>R k)" + apply (rule order_trans,rule norm_triangle_sub[of "x - y" "(e/2) *\<^sub>R k"]) + unfolding norm_scaleR norm_Basis[OF k] + apply auto + done + also have "\ < \e\ / 2 + \e\ / 2" + apply (rule add_strict_left_mono) + using as e + apply (auto simp add: field_simps dist_norm) + done + finally show "y \ ball x e" + unfolding mem_ball dist_norm using e by (auto simp add:field_simps) + qed + ultimately show ?thesis + apply (rule_tac x="?z" in exI) + unfolding Union_insert + apply auto + done + next + let ?z = "x + (e/2) *\<^sub>R k" + assume as: "x\k = b\k" + have "ball ?z (e / 2) \ i = {}" + proof (clarsimp simp only: all_not_in_conv [symmetric]) + fix y + assume "y \ ball ?z (e / 2)" and yi: "y \ i" + then have "dist ?z y < e/2" + by auto + then have "\(?z - y) \ k\ < e/2" + using Basis_le_norm[OF k, of "?z - y"] + unfolding dist_norm by auto + then have "y\k > b\k" + using e k + by (auto simp add:field_simps inner_simps inner_Basis as) + then have "y \ i" + unfolding ab mem_box by (auto intro!: bexI[OF _ k]) + then show False using yi by auto + qed + moreover + have "ball ?z (e/2) \ s \ (\insert i f)" + apply (rule order_trans[OF _ e[THEN conjunct2, unfolded lem1]]) + proof + fix y + assume as: "y\ ball ?z (e/2)" + have "norm (x - y) \ \e\ / 2 + norm (x - y + (e / 2) *\<^sub>R k)" + apply (rule order_trans,rule norm_triangle_sub[of "x - y" "- (e/2) *\<^sub>R k"]) + unfolding norm_scaleR + apply (auto simp: k) + done + also have "\ < \e\ / 2 + \e\ / 2" + apply (rule add_strict_left_mono) + using as unfolding mem_ball dist_norm + using e apply (auto simp add: field_simps) + done + finally show "y \ ball x e" + unfolding mem_ball dist_norm using e by (auto simp add:field_simps) + qed + ultimately show ?thesis + apply (rule_tac x="?z" in exI) + unfolding Union_insert + apply auto + done + qed + then obtain x where "ball x (e / 2) \ s \ \f" .. + then have "x \ s \ interior (\f)" + unfolding lem1[where U="\f", symmetric] + using centre_in_ball e by auto + then show ?thesis + using insert.hyps(3) insert.prems(1) by blast + qed + qed + qed + from this[OF assms(1,3)] x + obtain t x e where "t \ f" "0 < e" "ball x e \ s \ t" + by blast + then have "x \ s" "x \ interior t" + using open_subset_interior[OF open_ball, of x e t] + by auto + then show False + using \t \ f\ assms(4) by auto +qed + +subsection \Bounds on intervals where they exist.\ + +definition interval_upperbound :: "('a::euclidean_space) set \ 'a" + where "interval_upperbound s = (\i\Basis. (SUP x:s. x\i) *\<^sub>R i)" + +definition interval_lowerbound :: "('a::euclidean_space) set \ 'a" + where "interval_lowerbound s = (\i\Basis. (INF x:s. x\i) *\<^sub>R i)" + +lemma interval_upperbound[simp]: + "\i\Basis. a\i \ b\i \ + interval_upperbound (cbox a b) = (b::'a::euclidean_space)" + unfolding interval_upperbound_def euclidean_representation_setsum cbox_def + by (safe intro!: cSup_eq) auto + +lemma interval_lowerbound[simp]: + "\i\Basis. a\i \ b\i \ + interval_lowerbound (cbox a b) = (a::'a::euclidean_space)" + unfolding interval_lowerbound_def euclidean_representation_setsum cbox_def + by (safe intro!: cInf_eq) auto + +lemmas interval_bounds = interval_upperbound interval_lowerbound + +lemma + fixes X::"real set" + shows interval_upperbound_real[simp]: "interval_upperbound X = Sup X" + and interval_lowerbound_real[simp]: "interval_lowerbound X = Inf X" + by (auto simp: interval_upperbound_def interval_lowerbound_def) + +lemma interval_bounds'[simp]: + assumes "cbox a b \ {}" + shows "interval_upperbound (cbox a b) = b" + and "interval_lowerbound (cbox a b) = a" + using assms unfolding box_ne_empty by auto + + +lemma interval_upperbound_Times: + assumes "A \ {}" and "B \ {}" + shows "interval_upperbound (A \ B) = (interval_upperbound A, interval_upperbound B)" +proof- + from assms have fst_image_times': "A = fst ` (A \ B)" by simp + have "(\i\Basis. (SUP x:A \ B. x \ (i, 0)) *\<^sub>R i) = (\i\Basis. (SUP x:A. x \ i) *\<^sub>R i)" + by (subst (2) fst_image_times') (simp del: fst_image_times add: o_def inner_Pair_0) + moreover from assms have snd_image_times': "B = snd ` (A \ B)" by simp + have "(\i\Basis. (SUP x:A \ B. x \ (0, i)) *\<^sub>R i) = (\i\Basis. (SUP x:B. x \ i) *\<^sub>R i)" + by (subst (2) snd_image_times') (simp del: snd_image_times add: o_def inner_Pair_0) + ultimately show ?thesis unfolding interval_upperbound_def + by (subst setsum_Basis_prod_eq) (auto simp add: setsum_prod) +qed + +lemma interval_lowerbound_Times: + assumes "A \ {}" and "B \ {}" + shows "interval_lowerbound (A \ B) = (interval_lowerbound A, interval_lowerbound B)" +proof- + from assms have fst_image_times': "A = fst ` (A \ B)" by simp + have "(\i\Basis. (INF x:A \ B. x \ (i, 0)) *\<^sub>R i) = (\i\Basis. (INF x:A. x \ i) *\<^sub>R i)" + by (subst (2) fst_image_times') (simp del: fst_image_times add: o_def inner_Pair_0) + moreover from assms have snd_image_times': "B = snd ` (A \ B)" by simp + have "(\i\Basis. (INF x:A \ B. x \ (0, i)) *\<^sub>R i) = (\i\Basis. (INF x:B. x \ i) *\<^sub>R i)" + by (subst (2) snd_image_times') (simp del: snd_image_times add: o_def inner_Pair_0) + ultimately show ?thesis unfolding interval_lowerbound_def + by (subst setsum_Basis_prod_eq) (auto simp add: setsum_prod) +qed + +subsection \Content (length, area, volume...) of an interval.\ + +definition "content (s::('a::euclidean_space) set) = + (if s = {} then 0 else (\i\Basis. (interval_upperbound s)\i - (interval_lowerbound s)\i))" + +lemma interval_not_empty: "\i\Basis. a\i \ b\i \ cbox a b \ {}" + unfolding box_eq_empty unfolding not_ex not_less by auto + +lemma content_cbox: + fixes a :: "'a::euclidean_space" + assumes "\i\Basis. a\i \ b\i" + shows "content (cbox a b) = (\i\Basis. b\i - a\i)" + using interval_not_empty[OF assms] + unfolding content_def + by auto + +lemma content_cbox': + fixes a :: "'a::euclidean_space" + assumes "cbox a b \ {}" + shows "content (cbox a b) = (\i\Basis. b\i - a\i)" + using assms box_ne_empty(1) content_cbox by blast + +lemma content_real: "a \ b \ content {a..b} = b - a" + by (auto simp: interval_upperbound_def interval_lowerbound_def content_def) + +lemma abs_eq_content: "\y - x\ = (if x\y then content {x .. y} else content {y..x})" + by (auto simp: content_real) + +lemma content_singleton[simp]: "content {a} = 0" +proof - + have "content (cbox a a) = 0" + by (subst content_cbox) (auto simp: ex_in_conv) + then show ?thesis by (simp add: cbox_sing) +qed + +lemma content_unit[iff]: "content(cbox 0 (One::'a::euclidean_space)) = 1" + proof - + have *: "\i\Basis. (0::'a)\i \ (One::'a)\i" + by auto + have "0 \ cbox 0 (One::'a)" + unfolding mem_box by auto + then show ?thesis + unfolding content_def interval_bounds[OF *] using setprod.neutral_const by auto + qed + +lemma content_pos_le[intro]: + fixes a::"'a::euclidean_space" + shows "0 \ content (cbox a b)" +proof (cases "cbox a b = {}") + case False + then have *: "\i\Basis. a \ i \ b \ i" + unfolding box_ne_empty . + have "0 \ (\i\Basis. interval_upperbound (cbox a b) \ i - interval_lowerbound (cbox a b) \ i)" + apply (rule setprod_nonneg) + unfolding interval_bounds[OF *] + using * + apply auto + done + also have "\ = content (cbox a b)" using False by (simp add: content_def) + finally show ?thesis . +qed (simp add: content_def) + +corollary content_nonneg [simp]: + fixes a::"'a::euclidean_space" + shows "~ content (cbox a b) < 0" +using not_le by blast + +lemma content_pos_lt: + fixes a :: "'a::euclidean_space" + assumes "\i\Basis. a\i < b\i" + shows "0 < content (cbox a b)" + using assms + by (auto simp: content_def box_eq_empty intro!: setprod_pos) + +lemma content_eq_0: + "content (cbox a b) = 0 \ (\i\Basis. b\i \ a\i)" + by (auto simp: content_def box_eq_empty intro!: setprod_pos bexI) + +lemma cond_cases: "(P \ Q x) \ (\ P \ Q y) \ Q (if P then x else y)" + by auto + +lemma content_cbox_cases: + "content (cbox a (b::'a::euclidean_space)) = + (if \i\Basis. a\i \ b\i then setprod (\i. b\i - a\i) Basis else 0)" + by (auto simp: not_le content_eq_0 intro: less_imp_le content_cbox) + +lemma content_eq_0_interior: "content (cbox a b) = 0 \ interior(cbox a b) = {}" + unfolding content_eq_0 interior_cbox box_eq_empty + by auto + +lemma content_pos_lt_eq: + "0 < content (cbox a (b::'a::euclidean_space)) \ (\i\Basis. a\i < b\i)" +proof (rule iffI) + assume "0 < content (cbox a b)" + then have "content (cbox a b) \ 0" by auto + then show "\i\Basis. a\i < b\i" + unfolding content_eq_0 not_ex not_le by fastforce +next + assume "\i\Basis. a \ i < b \ i" + then show "0 < content (cbox a b)" + by (metis content_pos_lt) +qed + +lemma content_empty [simp]: "content {} = 0" + unfolding content_def by auto + +lemma content_real_if [simp]: "content {a..b} = (if a \ b then b - a else 0)" + by (simp add: content_real) + +lemma content_subset: + assumes "cbox a b \ cbox c d" + shows "content (cbox a b) \ content (cbox c d)" +proof (cases "cbox a b = {}") + case True + then show ?thesis + using content_pos_le[of c d] by auto +next + case False + then have ab_ne: "\i\Basis. a \ i \ b \ i" + unfolding box_ne_empty by auto + then have ab_ab: "a\cbox a b" "b\cbox a b" + unfolding mem_box by auto + have "cbox c d \ {}" using assms False by auto + then have cd_ne: "\i\Basis. c \ i \ d \ i" + using assms unfolding box_ne_empty by auto + have "\i. i \ Basis \ 0 \ b \ i - a \ i" + using ab_ne by auto + moreover + have "\i. i \ Basis \ b \ i - a \ i \ d \ i - c \ i" + using assms[unfolded subset_eq mem_box,rule_format,OF ab_ab(2)] + assms[unfolded subset_eq mem_box,rule_format,OF ab_ab(1)] + by (metis diff_mono) + ultimately show ?thesis + unfolding content_def interval_bounds[OF ab_ne] interval_bounds[OF cd_ne] + by (simp add: setprod_mono if_not_P[OF False] if_not_P[OF \cbox c d \ {}\]) +qed + +lemma content_lt_nz: "0 < content (cbox a b) \ content (cbox a b) \ 0" + unfolding content_pos_lt_eq content_eq_0 unfolding not_ex not_le by fastforce + +lemma content_times[simp]: "content (A \ B) = content A * content B" +proof (cases "A \ B = {}") + let ?ub1 = "interval_upperbound" and ?lb1 = "interval_lowerbound" + let ?ub2 = "interval_upperbound" and ?lb2 = "interval_lowerbound" + assume nonempty: "A \ B \ {}" + hence "content (A \ B) = (\i\Basis. (?ub1 A, ?ub2 B) \ i - (?lb1 A, ?lb2 B) \ i)" + unfolding content_def by (simp add: interval_upperbound_Times interval_lowerbound_Times) + also have "... = content A * content B" unfolding content_def using nonempty + apply (subst Basis_prod_def, subst setprod.union_disjoint, force, force, force, simp) + apply (subst (1 2) setprod.reindex, auto intro: inj_onI) + done + finally show ?thesis . +qed (auto simp: content_def) + +lemma content_Pair: "content (cbox (a,c) (b,d)) = content (cbox a b) * content (cbox c d)" + by (simp add: cbox_Pair_eq) + +lemma content_cbox_pair_eq0_D: + "content (cbox (a,c) (b,d)) = 0 \ content (cbox a b) = 0 \ content (cbox c d) = 0" + by (simp add: content_Pair) + +lemma content_eq_0_gen: + fixes s :: "'a::euclidean_space set" + assumes "bounded s" + shows "content s = 0 \ (\i\Basis. \v. \x \ s. x \ i = v)" (is "_ = ?rhs") +proof safe + assume "content s = 0" then show ?rhs + apply (clarsimp simp: ex_in_conv content_def split: if_split_asm) + apply (rule_tac x=a in bexI) + apply (rule_tac x="interval_lowerbound s \ a" in exI) + apply (clarsimp simp: interval_upperbound_def interval_lowerbound_def) + apply (drule cSUP_eq_cINF_D) + apply (auto simp: bounded_inner_imp_bdd_above [OF assms] bounded_inner_imp_bdd_below [OF assms]) + done +next + fix i a + assume "i \ Basis" "\x\s. x \ i = a" + then show "content s = 0" + apply (clarsimp simp: content_def) + apply (rule_tac x=i in bexI) + apply (auto simp: interval_upperbound_def interval_lowerbound_def) + done +qed + +lemma content_0_subset_gen: + fixes a :: "'a::euclidean_space" + assumes "content t = 0" "s \ t" "bounded t" shows "content s = 0" +proof - + have "bounded s" + using assms by (metis bounded_subset) + then show ?thesis + using assms + by (auto simp: content_eq_0_gen) +qed + +lemma content_0_subset: "\content(cbox a b) = 0; s \ cbox a b\ \ content s = 0" + by (simp add: content_0_subset_gen bounded_cbox) + + +lemma interval_split: + fixes a :: "'a::euclidean_space" + assumes "k \ Basis" + shows + "cbox a b \ {x. x\k \ c} = cbox a (\i\Basis. (if i = k then min (b\k) c else b\i) *\<^sub>R i)" + "cbox a b \ {x. x\k \ c} = cbox (\i\Basis. (if i = k then max (a\k) c else a\i) *\<^sub>R i) b" + apply (rule_tac[!] set_eqI) + unfolding Int_iff mem_box mem_Collect_eq + using assms + apply auto + done + +lemma content_split: + fixes a :: "'a::euclidean_space" + assumes "k \ Basis" + shows "content (cbox a b) = content(cbox a b \ {x. x\k \ c}) + content(cbox a b \ {x. x\k \ c})" +proof cases + note simps = interval_split[OF assms] content_cbox_cases + have *: "Basis = insert k (Basis - {k})" "\x. finite (Basis-{x})" "\x. x\Basis-{x}" + using assms by auto + have *: "\X Y Z. (\i\Basis. Z i (if i = k then X else Y i)) = Z k X * (\i\Basis-{k}. Z i (Y i))" + "(\i\Basis. b\i - a\i) = (\i\Basis-{k}. b\i - a\i) * (b\k - a\k)" + apply (subst *(1)) + defer + apply (subst *(1)) + unfolding setprod.insert[OF *(2-)] + apply auto + done + assume as: "\i\Basis. a \ i \ b \ i" + moreover + have "\x. min (b \ k) c = max (a \ k) c \ + x * (b\k - a\k) = x * (max (a \ k) c - a \ k) + x * (b \ k - max (a \ k) c)" + by (auto simp add: field_simps) + moreover + have **: "(\i\Basis. ((\i\Basis. (if i = k then min (b \ k) c else b \ i) *\<^sub>R i) \ i - a \ i)) = + (\i\Basis. (if i = k then min (b \ k) c else b \ i) - a \ i)" + "(\i\Basis. b \ i - ((\i\Basis. (if i = k then max (a \ k) c else a \ i) *\<^sub>R i) \ i)) = + (\i\Basis. b \ i - (if i = k then max (a \ k) c else a \ i))" + by (auto intro!: setprod.cong) + have "\ a \ k \ c \ \ c \ b \ k \ False" + unfolding not_le + using as[unfolded ,rule_format,of k] assms + by auto + ultimately show ?thesis + using assms + unfolding simps ** + unfolding *(1)[of "\i x. b\i - x"] *(1)[of "\i x. x - a\i"] + unfolding *(2) + by auto +next + assume "\ (\i\Basis. a \ i \ b \ i)" + then have "cbox a b = {}" + unfolding box_eq_empty by (auto simp: not_le) + then show ?thesis + by (auto simp: not_le) +qed + +subsection \The notion of a gauge --- simply an open set containing the point.\ + +definition "gauge d \ (\x. x \ d x \ open (d x))" + +lemma gaugeI: + assumes "\x. x \ g x" + and "\x. open (g x)" + shows "gauge g" + using assms unfolding gauge_def by auto + +lemma gaugeD[dest]: + assumes "gauge d" + shows "x \ d x" + and "open (d x)" + using assms unfolding gauge_def by auto + +lemma gauge_ball_dependent: "\x. 0 < e x \ gauge (\x. ball x (e x))" + unfolding gauge_def by auto + +lemma gauge_ball[intro]: "0 < e \ gauge (\x. ball x e)" + unfolding gauge_def by auto + +lemma gauge_trivial[intro!]: "gauge (\x. ball x 1)" + by (rule gauge_ball) auto + +lemma gauge_inter[intro]: "gauge d1 \ gauge d2 \ gauge (\x. d1 x \ d2 x)" + unfolding gauge_def by auto + +lemma gauge_inters: + assumes "finite s" + and "\d\s. gauge (f d)" + shows "gauge (\x. \{f d x | d. d \ s})" +proof - + have *: "\x. {f d x |d. d \ s} = (\d. f d x) ` s" + by auto + show ?thesis + unfolding gauge_def unfolding * + using assms unfolding Ball_def Inter_iff mem_Collect_eq gauge_def by auto +qed + +lemma gauge_existence_lemma: + "(\x. \d :: real. p x \ 0 < d \ q d x) \ (\x. \d>0. p x \ q d x)" + by (metis zero_less_one) + + +subsection \Divisions.\ + +definition division_of (infixl "division'_of" 40) +where + "s division_of i \ + finite s \ + (\k\s. k \ i \ k \ {} \ (\a b. k = cbox a b)) \ + (\k1\s. \k2\s. k1 \ k2 \ interior(k1) \ interior(k2) = {}) \ + (\s = i)" + +lemma division_ofD[dest]: + assumes "s division_of i" + shows "finite s" + and "\k. k \ s \ k \ i" + and "\k. k \ s \ k \ {}" + and "\k. k \ s \ \a b. k = cbox a b" + and "\k1 k2. k1 \ s \ k2 \ s \ k1 \ k2 \ interior(k1) \ interior(k2) = {}" + and "\s = i" + using assms unfolding division_of_def by auto + +lemma division_ofI: + assumes "finite s" + and "\k. k \ s \ k \ i" + and "\k. k \ s \ k \ {}" + and "\k. k \ s \ \a b. k = cbox a b" + and "\k1 k2. k1 \ s \ k2 \ s \ k1 \ k2 \ interior k1 \ interior k2 = {}" + and "\s = i" + shows "s division_of i" + using assms unfolding division_of_def by auto + +lemma division_of_finite: "s division_of i \ finite s" + unfolding division_of_def by auto + +lemma division_of_self[intro]: "cbox a b \ {} \ {cbox a b} division_of (cbox a b)" + unfolding division_of_def by auto + +lemma division_of_trivial[simp]: "s division_of {} \ s = {}" + unfolding division_of_def by auto + +lemma division_of_sing[simp]: + "s division_of cbox a (a::'a::euclidean_space) \ s = {cbox a a}" + (is "?l = ?r") +proof + assume ?r + moreover + { fix k + assume "s = {{a}}" "k\s" + then have "\x y. k = cbox x y" + apply (rule_tac x=a in exI)+ + apply (force simp: cbox_sing) + done + } + ultimately show ?l + unfolding division_of_def cbox_sing by auto +next + assume ?l + note * = conjunctD4[OF this[unfolded division_of_def cbox_sing]] + { + fix x + assume x: "x \ s" have "x = {a}" + using *(2)[rule_format,OF x] by auto + } + moreover have "s \ {}" + using *(4) by auto + ultimately show ?r + unfolding cbox_sing by auto +qed + +lemma elementary_empty: obtains p where "p division_of {}" + unfolding division_of_trivial by auto + +lemma elementary_interval: obtains p where "p division_of (cbox a b)" + by (metis division_of_trivial division_of_self) + +lemma division_contains: "s division_of i \ \x\i. \k\s. x \ k" + unfolding division_of_def by auto + +lemma forall_in_division: + "d division_of i \ (\x\d. P x) \ (\a b. cbox a b \ d \ P (cbox a b))" + unfolding division_of_def by fastforce + +lemma division_of_subset: + assumes "p division_of (\p)" + and "q \ p" + shows "q division_of (\q)" +proof (rule division_ofI) + note * = division_ofD[OF assms(1)] + show "finite q" + using "*"(1) assms(2) infinite_super by auto + { + fix k + assume "k \ q" + then have kp: "k \ p" + using assms(2) by auto + show "k \ \q" + using \k \ q\ by auto + show "\a b. k = cbox a b" + using *(4)[OF kp] by auto + show "k \ {}" + using *(3)[OF kp] by auto + } + fix k1 k2 + assume "k1 \ q" "k2 \ q" "k1 \ k2" + then have **: "k1 \ p" "k2 \ p" "k1 \ k2" + using assms(2) by auto + show "interior k1 \ interior k2 = {}" + using *(5)[OF **] by auto +qed auto + +lemma division_of_union_self[intro]: "p division_of s \ p division_of (\p)" + unfolding division_of_def by auto + +lemma division_of_content_0: + assumes "content (cbox a b) = 0" "d division_of (cbox a b)" + shows "\k\d. content k = 0" + unfolding forall_in_division[OF assms(2)] + by (metis antisym_conv assms content_pos_le content_subset division_ofD(2)) + +lemma division_inter: + fixes s1 s2 :: "'a::euclidean_space set" + assumes "p1 division_of s1" + and "p2 division_of s2" + shows "{k1 \ k2 | k1 k2 .k1 \ p1 \ k2 \ p2 \ k1 \ k2 \ {}} division_of (s1 \ s2)" + (is "?A' division_of _") +proof - + let ?A = "{s. s \ (\(k1,k2). k1 \ k2) ` (p1 \ p2) \ s \ {}}" + have *: "?A' = ?A" by auto + show ?thesis + unfolding * + proof (rule division_ofI) + have "?A \ (\(x, y). x \ y) ` (p1 \ p2)" + by auto + moreover have "finite (p1 \ p2)" + using assms unfolding division_of_def by auto + ultimately show "finite ?A" by auto + have *: "\s. \{x\s. x \ {}} = \s" + by auto + show "\?A = s1 \ s2" + apply (rule set_eqI) + unfolding * and UN_iff + using division_ofD(6)[OF assms(1)] and division_ofD(6)[OF assms(2)] + apply auto + done + { + fix k + assume "k \ ?A" + then obtain k1 k2 where k: "k = k1 \ k2" "k1 \ p1" "k2 \ p2" "k \ {}" + by auto + then show "k \ {}" + by auto + show "k \ s1 \ s2" + using division_ofD(2)[OF assms(1) k(2)] and division_ofD(2)[OF assms(2) k(3)] + unfolding k by auto + obtain a1 b1 where k1: "k1 = cbox a1 b1" + using division_ofD(4)[OF assms(1) k(2)] by blast + obtain a2 b2 where k2: "k2 = cbox a2 b2" + using division_ofD(4)[OF assms(2) k(3)] by blast + show "\a b. k = cbox a b" + unfolding k k1 k2 unfolding inter_interval by auto + } + fix k1 k2 + assume "k1 \ ?A" + then obtain x1 y1 where k1: "k1 = x1 \ y1" "x1 \ p1" "y1 \ p2" "k1 \ {}" + by auto + assume "k2 \ ?A" + then obtain x2 y2 where k2: "k2 = x2 \ y2" "x2 \ p1" "y2 \ p2" "k2 \ {}" + by auto + assume "k1 \ k2" + then have th: "x1 \ x2 \ y1 \ y2" + unfolding k1 k2 by auto + have *: "interior x1 \ interior x2 = {} \ interior y1 \ interior y2 = {} \ + interior (x1 \ y1) \ interior x1 \ interior (x1 \ y1) \ interior y1 \ + interior (x2 \ y2) \ interior x2 \ interior (x2 \ y2) \ interior y2 \ + interior (x1 \ y1) \ interior (x2 \ y2) = {}" by auto + show "interior k1 \ interior k2 = {}" + unfolding k1 k2 + apply (rule *) + using assms division_ofD(5) k1 k2(2) k2(3) th apply auto + done + qed +qed + +lemma division_inter_1: + assumes "d division_of i" + and "cbox a (b::'a::euclidean_space) \ i" + shows "{cbox a b \ k | k. k \ d \ cbox a b \ k \ {}} division_of (cbox a b)" +proof (cases "cbox a b = {}") + case True + show ?thesis + unfolding True and division_of_trivial by auto +next + case False + have *: "cbox a b \ i = cbox a b" using assms(2) by auto + show ?thesis + using division_inter[OF division_of_self[OF False] assms(1)] + unfolding * by auto +qed + +lemma elementary_inter: + fixes s t :: "'a::euclidean_space set" + assumes "p1 division_of s" + and "p2 division_of t" + shows "\p. p division_of (s \ t)" +using assms division_inter by blast + +lemma elementary_inters: + assumes "finite f" + and "f \ {}" + and "\s\f. \p. p division_of (s::('a::euclidean_space) set)" + shows "\p. p division_of (\f)" + using assms +proof (induct f rule: finite_induct) + case (insert x f) + show ?case + proof (cases "f = {}") + case True + then show ?thesis + unfolding True using insert by auto + next + case False + obtain p where "p division_of \f" + using insert(3)[OF False insert(5)[unfolded ball_simps,THEN conjunct2]] .. + moreover obtain px where "px division_of x" + using insert(5)[rule_format,OF insertI1] .. + ultimately show ?thesis + by (simp add: elementary_inter Inter_insert) + qed +qed auto + +lemma division_disjoint_union: + assumes "p1 division_of s1" + and "p2 division_of s2" + and "interior s1 \ interior s2 = {}" + shows "(p1 \ p2) division_of (s1 \ s2)" +proof (rule division_ofI) + note d1 = division_ofD[OF assms(1)] + note d2 = division_ofD[OF assms(2)] + show "finite (p1 \ p2)" + using d1(1) d2(1) by auto + show "\(p1 \ p2) = s1 \ s2" + using d1(6) d2(6) by auto + { + fix k1 k2 + assume as: "k1 \ p1 \ p2" "k2 \ p1 \ p2" "k1 \ k2" + moreover + let ?g="interior k1 \ interior k2 = {}" + { + assume as: "k1\p1" "k2\p2" + have ?g + using interior_mono[OF d1(2)[OF as(1)]] interior_mono[OF d2(2)[OF as(2)]] + using assms(3) by blast + } + moreover + { + assume as: "k1\p2" "k2\p1" + have ?g + using interior_mono[OF d1(2)[OF as(2)]] interior_mono[OF d2(2)[OF as(1)]] + using assms(3) by blast + } + ultimately show ?g + using d1(5)[OF _ _ as(3)] and d2(5)[OF _ _ as(3)] by auto + } + fix k + assume k: "k \ p1 \ p2" + show "k \ s1 \ s2" + using k d1(2) d2(2) by auto + show "k \ {}" + using k d1(3) d2(3) by auto + show "\a b. k = cbox a b" + using k d1(4) d2(4) by auto +qed + +lemma partial_division_extend_1: + fixes a b c d :: "'a::euclidean_space" + assumes incl: "cbox c d \ cbox a b" + and nonempty: "cbox c d \ {}" + obtains p where "p division_of (cbox a b)" "cbox c d \ p" +proof + let ?B = "\f::'a\'a \ 'a. + cbox (\i\Basis. (fst (f i) \ i) *\<^sub>R i) (\i\Basis. (snd (f i) \ i) *\<^sub>R i)" + define p where "p = ?B ` (Basis \\<^sub>E {(a, c), (c, d), (d, b)})" + + show "cbox c d \ p" + unfolding p_def + by (auto simp add: box_eq_empty cbox_def intro!: image_eqI[where x="\(i::'a)\Basis. (c, d)"]) + { + fix i :: 'a + assume "i \ Basis" + with incl nonempty have "a \ i \ c \ i" "c \ i \ d \ i" "d \ i \ b \ i" + unfolding box_eq_empty subset_box by (auto simp: not_le) + } + note ord = this + + show "p division_of (cbox a b)" + proof (rule division_ofI) + show "finite p" + unfolding p_def by (auto intro!: finite_PiE) + { + fix k + assume "k \ p" + then obtain f where f: "f \ Basis \\<^sub>E {(a, c), (c, d), (d, b)}" and k: "k = ?B f" + by (auto simp: p_def) + then show "\a b. k = cbox a b" + by auto + have "k \ cbox a b \ k \ {}" + proof (simp add: k box_eq_empty subset_box not_less, safe) + fix i :: 'a + assume i: "i \ Basis" + with f have "f i = (a, c) \ f i = (c, d) \ f i = (d, b)" + by (auto simp: PiE_iff) + with i ord[of i] + show "a \ i \ fst (f i) \ i" "snd (f i) \ i \ b \ i" "fst (f i) \ i \ snd (f i) \ i" + by auto + qed + then show "k \ {}" "k \ cbox a b" + by auto + { + fix l + assume "l \ p" + then obtain g where g: "g \ Basis \\<^sub>E {(a, c), (c, d), (d, b)}" and l: "l = ?B g" + by (auto simp: p_def) + assume "l \ k" + have "\i\Basis. f i \ g i" + proof (rule ccontr) + assume "\ ?thesis" + with f g have "f = g" + by (auto simp: PiE_iff extensional_def intro!: ext) + with \l \ k\ show False + by (simp add: l k) + qed + then obtain i where *: "i \ Basis" "f i \ g i" .. + then have "f i = (a, c) \ f i = (c, d) \ f i = (d, b)" + "g i = (a, c) \ g i = (c, d) \ g i = (d, b)" + using f g by (auto simp: PiE_iff) + with * ord[of i] show "interior l \ interior k = {}" + by (auto simp add: l k interior_cbox disjoint_interval intro!: bexI[of _ i]) + } + note \k \ cbox a b\ + } + moreover + { + fix x assume x: "x \ cbox a b" + have "\i\Basis. \l. x \ i \ {fst l \ i .. snd l \ i} \ l \ {(a, c), (c, d), (d, b)}" + proof + fix i :: 'a + assume "i \ Basis" + with x ord[of i] + have "(a \ i \ x \ i \ x \ i \ c \ i) \ (c \ i \ x \ i \ x \ i \ d \ i) \ + (d \ i \ x \ i \ x \ i \ b \ i)" + by (auto simp: cbox_def) + then show "\l. x \ i \ {fst l \ i .. snd l \ i} \ l \ {(a, c), (c, d), (d, b)}" + by auto + qed + then obtain f where + f: "\i\Basis. x \ i \ {fst (f i) \ i..snd (f i) \ i} \ f i \ {(a, c), (c, d), (d, b)}" + unfolding bchoice_iff .. + moreover from f have "restrict f Basis \ Basis \\<^sub>E {(a, c), (c, d), (d, b)}" + by auto + moreover from f have "x \ ?B (restrict f Basis)" + by (auto simp: mem_box) + ultimately have "\k\p. x \ k" + unfolding p_def by blast + } + ultimately show "\p = cbox a b" + by auto + qed +qed + +lemma partial_division_extend_interval: + assumes "p division_of (\p)" "(\p) \ cbox a b" + obtains q where "p \ q" "q division_of cbox a (b::'a::euclidean_space)" +proof (cases "p = {}") + case True + obtain q where "q division_of (cbox a b)" + by (rule elementary_interval) + then show ?thesis + using True that by blast +next + case False + note p = division_ofD[OF assms(1)] + have div_cbox: "\k\p. \q. q division_of cbox a b \ k \ q" + proof + fix k + assume kp: "k \ p" + obtain c d where k: "k = cbox c d" + using p(4)[OF kp] by blast + have *: "cbox c d \ cbox a b" "cbox c d \ {}" + using p(2,3)[OF kp, unfolded k] using assms(2) + by (blast intro: order.trans)+ + obtain q where "q division_of cbox a b" "cbox c d \ q" + by (rule partial_division_extend_1[OF *]) + then show "\q. q division_of cbox a b \ k \ q" + unfolding k by auto + qed + obtain q where q: "\x. x \ p \ q x division_of cbox a b" "\x. x \ p \ x \ q x" + using bchoice[OF div_cbox] by blast + { fix x + assume x: "x \ p" + have "q x division_of \q x" + apply (rule division_ofI) + using division_ofD[OF q(1)[OF x]] + apply auto + done } + then have "\x. x \ p \ \d. d division_of \(q x - {x})" + by (meson Diff_subset division_of_subset) + then have "\d. d division_of \((\i. \(q i - {i})) ` p)" + apply - + apply (rule elementary_inters [OF finite_imageI[OF p(1)]]) + apply (auto simp: False elementary_inters [OF finite_imageI[OF p(1)]]) + done + then obtain d where d: "d division_of \((\i. \(q i - {i})) ` p)" .. + have "d \ p division_of cbox a b" + proof - + have te: "\s f t. s \ {} \ \i\s. f i \ i = t \ t = \(f ` s) \ \s" by auto + have cbox_eq: "cbox a b = \((\i. \(q i - {i})) ` p) \ \p" + proof (rule te[OF False], clarify) + fix i + assume i: "i \ p" + show "\(q i - {i}) \ i = cbox a b" + using division_ofD(6)[OF q(1)[OF i]] using q(2)[OF i] by auto + qed + { fix k + assume k: "k \ p" + have *: "\u t s. t \ s = {} \ u \ s \ u \ t = {}" + by auto + have "interior (\i\p. \(q i - {i})) \ interior k = {}" + proof (rule *[OF inter_interior_unions_intervals]) + note qk=division_ofD[OF q(1)[OF k]] + show "finite (q k - {k})" "open (interior k)" "\t\q k - {k}. \a b. t = cbox a b" + using qk by auto + show "\t\q k - {k}. interior k \ interior t = {}" + using qk(5) using q(2)[OF k] by auto + show "interior (\i\p. \(q i - {i})) \ interior (\(q k - {k}))" + apply (rule interior_mono)+ + using k + apply auto + done + qed } note [simp] = this + show "d \ p division_of (cbox a b)" + unfolding cbox_eq + apply (rule division_disjoint_union[OF d assms(1)]) + apply (rule inter_interior_unions_intervals) + apply (rule p open_interior ballI)+ + apply simp_all + done + qed + then show ?thesis + by (meson Un_upper2 that) +qed + +lemma elementary_bounded[dest]: + fixes s :: "'a::euclidean_space set" + shows "p division_of s \ bounded s" + unfolding division_of_def by (metis bounded_Union bounded_cbox) + +lemma elementary_subset_cbox: + "p division_of s \ \a b. s \ cbox a (b::'a::euclidean_space)" + by (meson elementary_bounded bounded_subset_cbox) + +lemma division_union_intervals_exists: + fixes a b :: "'a::euclidean_space" + assumes "cbox a b \ {}" + obtains p where "(insert (cbox a b) p) division_of (cbox a b \ cbox c d)" +proof (cases "cbox c d = {}") + case True + show ?thesis + apply (rule that[of "{}"]) + unfolding True + using assms + apply auto + done +next + case False + show ?thesis + proof (cases "cbox a b \ cbox c d = {}") + case True + then show ?thesis + by (metis that False assms division_disjoint_union division_of_self insert_is_Un interior_Int interior_empty) + next + case False + obtain u v where uv: "cbox a b \ cbox c d = cbox u v" + unfolding inter_interval by auto + have uv_sub: "cbox u v \ cbox c d" using uv by auto + obtain p where "p division_of cbox c d" "cbox u v \ p" + by (rule partial_division_extend_1[OF uv_sub False[unfolded uv]]) + note p = this division_ofD[OF this(1)] + have "interior (cbox a b \ \(p - {cbox u v})) = interior(cbox u v \ \(p - {cbox u v}))" + apply (rule arg_cong[of _ _ interior]) + using p(8) uv by auto + also have "\ = {}" + unfolding interior_Int + apply (rule inter_interior_unions_intervals) + using p(6) p(7)[OF p(2)] p(3) + apply auto + done + finally have [simp]: "interior (cbox a b) \ interior (\(p - {cbox u v})) = {}" by simp + have cbe: "cbox a b \ cbox c d = cbox a b \ \(p - {cbox u v})" + using p(8) unfolding uv[symmetric] by auto + have "insert (cbox a b) (p - {cbox u v}) division_of cbox a b \ \(p - {cbox u v})" + proof - + have "{cbox a b} division_of cbox a b" + by (simp add: assms division_of_self) + then show "insert (cbox a b) (p - {cbox u v}) division_of cbox a b \ \(p - {cbox u v})" + by (metis (no_types) Diff_subset \interior (cbox a b) \ interior (\(p - {cbox u v})) = {}\ division_disjoint_union division_of_subset insert_is_Un p(1) p(8)) + qed + with that[of "p - {cbox u v}"] show ?thesis by (simp add: cbe) + qed +qed + +lemma division_of_unions: + assumes "finite f" + and "\p. p \ f \ p division_of (\p)" + and "\k1 k2. k1 \ \f \ k2 \ \f \ k1 \ k2 \ interior k1 \ interior k2 = {}" + shows "\f division_of \\f" + using assms + by (auto intro!: division_ofI) + +lemma elementary_union_interval: + fixes a b :: "'a::euclidean_space" + assumes "p division_of \p" + obtains q where "q division_of (cbox a b \ \p)" +proof - + note assm = division_ofD[OF assms] + have lem1: "\f s. \\(f ` s) = \((\x. \(f x)) ` s)" + by auto + have lem2: "\f s. f \ {} \ \{s \ t |t. t \ f} = s \ \f" + by auto + { + presume "p = {} \ thesis" + "cbox a b = {} \ thesis" + "cbox a b \ {} \ interior (cbox a b) = {} \ thesis" + "p \ {} \ interior (cbox a b)\{} \ cbox a b \ {} \ thesis" + then show thesis by auto + next + assume as: "p = {}" + obtain p where "p division_of (cbox a b)" + by (rule elementary_interval) + then show thesis + using as that by auto + next + assume as: "cbox a b = {}" + show thesis + using as assms that by auto + next + assume as: "interior (cbox a b) = {}" "cbox a b \ {}" + show thesis + apply (rule that[of "insert (cbox a b) p"],rule division_ofI) + unfolding finite_insert + apply (rule assm(1)) unfolding Union_insert + using assm(2-4) as + apply - + apply (fast dest: assm(5))+ + done + next + assume as: "p \ {}" "interior (cbox a b) \ {}" "cbox a b \ {}" + have "\k\p. \q. (insert (cbox a b) q) division_of (cbox a b \ k)" + proof + fix k + assume kp: "k \ p" + from assm(4)[OF kp] obtain c d where "k = cbox c d" by blast + then show "\q. (insert (cbox a b) q) division_of (cbox a b \ k)" + by (meson as(3) division_union_intervals_exists) + qed + from bchoice[OF this] obtain q where "\x\p. insert (cbox a b) (q x) division_of (cbox a b) \ x" .. + note q = division_ofD[OF this[rule_format]] + let ?D = "\{insert (cbox a b) (q k) | k. k \ p}" + show thesis + proof (rule that[OF division_ofI]) + have *: "{insert (cbox a b) (q k) |k. k \ p} = (\k. insert (cbox a b) (q k)) ` p" + by auto + show "finite ?D" + using "*" assm(1) q(1) by auto + show "\?D = cbox a b \ \p" + unfolding * lem1 + unfolding lem2[OF as(1), of "cbox a b", symmetric] + using q(6) + by auto + fix k + assume k: "k \ ?D" + then show "k \ cbox a b \ \p" + using q(2) by auto + show "k \ {}" + using q(3) k by auto + show "\a b. k = cbox a b" + using q(4) k by auto + fix k' + assume k': "k' \ ?D" "k \ k'" + obtain x where x: "k \ insert (cbox a b) (q x)" "x\p" + using k by auto + obtain x' where x': "k'\insert (cbox a b) (q x')" "x'\p" + using k' by auto + show "interior k \ interior k' = {}" + proof (cases "x = x'") + case True + show ?thesis + using True k' q(5) x' x by auto + next + case False + { + presume "k = cbox a b \ ?thesis" + and "k' = cbox a b \ ?thesis" + and "k \ cbox a b \ k' \ cbox a b \ ?thesis" + then show ?thesis by linarith + next + assume as': "k = cbox a b" + show ?thesis + using as' k' q(5) x' by blast + next + assume as': "k' = cbox a b" + show ?thesis + using as' k'(2) q(5) x by blast + } + assume as': "k \ cbox a b" "k' \ cbox a b" + obtain c d where k: "k = cbox c d" + using q(4)[OF x(2,1)] by blast + have "interior k \ interior (cbox a b) = {}" + using as' k'(2) q(5) x by blast + then have "interior k \ interior x" + using interior_subset_union_intervals + by (metis as(2) k q(2) x interior_subset_union_intervals) + moreover + obtain c d where c_d: "k' = cbox c d" + using q(4)[OF x'(2,1)] by blast + have "interior k' \ interior (cbox a b) = {}" + using as'(2) q(5) x' by blast + then have "interior k' \ interior x'" + by (metis as(2) c_d interior_subset_union_intervals q(2) x'(1) x'(2)) + ultimately show ?thesis + using assm(5)[OF x(2) x'(2) False] by auto + qed + qed + } +qed + +lemma elementary_unions_intervals: + assumes fin: "finite f" + and "\s. s \ f \ \a b. s = cbox a (b::'a::euclidean_space)" + obtains p where "p division_of (\f)" +proof - + have "\p. p division_of (\f)" + proof (induct_tac f rule:finite_subset_induct) + show "\p. p division_of \{}" using elementary_empty by auto + next + fix x F + assume as: "finite F" "x \ F" "\p. p division_of \F" "x\f" + from this(3) obtain p where p: "p division_of \F" .. + from assms(2)[OF as(4)] obtain a b where x: "x = cbox a b" by blast + have *: "\F = \p" + using division_ofD[OF p] by auto + show "\p. p division_of \insert x F" + using elementary_union_interval[OF p[unfolded *], of a b] + unfolding Union_insert x * by metis + qed (insert assms, auto) + then show ?thesis + using that by auto +qed + +lemma elementary_union: + fixes s t :: "'a::euclidean_space set" + assumes "ps division_of s" "pt division_of t" + obtains p where "p division_of (s \ t)" +proof - + have *: "s \ t = \ps \ \pt" + using assms unfolding division_of_def by auto + show ?thesis + apply (rule elementary_unions_intervals[of "ps \ pt"]) + using assms apply auto + by (simp add: * that) +qed + +lemma partial_division_extend: + fixes t :: "'a::euclidean_space set" + assumes "p division_of s" + and "q division_of t" + and "s \ t" + obtains r where "p \ r" and "r division_of t" +proof - + note divp = division_ofD[OF assms(1)] and divq = division_ofD[OF assms(2)] + obtain a b where ab: "t \ cbox a b" + using elementary_subset_cbox[OF assms(2)] by auto + obtain r1 where "p \ r1" "r1 division_of (cbox a b)" + using assms + by (metis ab dual_order.trans partial_division_extend_interval divp(6)) + note r1 = this division_ofD[OF this(2)] + obtain p' where "p' division_of \(r1 - p)" + apply (rule elementary_unions_intervals[of "r1 - p"]) + using r1(3,6) + apply auto + done + then obtain r2 where r2: "r2 division_of (\(r1 - p)) \ (\q)" + by (metis assms(2) divq(6) elementary_inter) + { + fix x + assume x: "x \ t" "x \ s" + then have "x\\r1" + unfolding r1 using ab by auto + then obtain r where r: "r \ r1" "x \ r" + unfolding Union_iff .. + moreover + have "r \ p" + proof + assume "r \ p" + then have "x \ s" using divp(2) r by auto + then show False using x by auto + qed + ultimately have "x\\(r1 - p)" by auto + } + then have *: "t = \p \ (\(r1 - p) \ \q)" + unfolding divp divq using assms(3) by auto + show ?thesis + apply (rule that[of "p \ r2"]) + unfolding * + defer + apply (rule division_disjoint_union) + unfolding divp(6) + apply(rule assms r2)+ + proof - + have "interior s \ interior (\(r1-p)) = {}" + proof (rule inter_interior_unions_intervals) + show "finite (r1 - p)" and "open (interior s)" and "\t\r1-p. \a b. t = cbox a b" + using r1 by auto + have *: "\s. (\x. x \ s \ False) \ s = {}" + by auto + show "\t\r1-p. interior s \ interior t = {}" + proof + fix m x + assume as: "m \ r1 - p" + have "interior m \ interior (\p) = {}" + proof (rule inter_interior_unions_intervals) + show "finite p" and "open (interior m)" and "\t\p. \a b. t = cbox a b" + using divp by auto + show "\t\p. interior m \ interior t = {}" + by (metis DiffD1 DiffD2 as r1(1) r1(7) set_rev_mp) + qed + then show "interior s \ interior m = {}" + unfolding divp by auto + qed + qed + then show "interior s \ interior (\(r1-p) \ (\q)) = {}" + using interior_subset by auto + qed auto +qed + +lemma division_split_left_inj: + fixes type :: "'a::euclidean_space" + assumes "d division_of i" + and "k1 \ d" + and "k2 \ d" + and "k1 \ k2" + and "k1 \ {x::'a. x\k \ c} = k2 \ {x. x\k \ c}" + and k: "k\Basis" + shows "content(k1 \ {x. x\k \ c}) = 0" +proof - + note d=division_ofD[OF assms(1)] + have *: "\(a::'a) b c. content (cbox a b \ {x. x\k \ c}) = 0 \ + interior(cbox a b \ {x. x\k \ c}) = {}" + unfolding interval_split[OF k] content_eq_0_interior by auto + guess u1 v1 using d(4)[OF assms(2)] by (elim exE) note uv1=this + guess u2 v2 using d(4)[OF assms(3)] by (elim exE) note uv2=this + have **: "\s t u. s \ t = {} \ u \ s \ u \ t \ u = {}" + by auto + show ?thesis + unfolding uv1 uv2 * + apply (rule **[OF d(5)[OF assms(2-4)]]) + apply (simp add: uv1) + using assms(5) uv1 by auto +qed + +lemma division_split_right_inj: + fixes type :: "'a::euclidean_space" + assumes "d division_of i" + and "k1 \ d" + and "k2 \ d" + and "k1 \ k2" + and "k1 \ {x::'a. x\k \ c} = k2 \ {x. x\k \ c}" + and k: "k \ Basis" + shows "content (k1 \ {x. x\k \ c}) = 0" +proof - + note d=division_ofD[OF assms(1)] + have *: "\a b::'a. \c. content(cbox a b \ {x. x\k \ c}) = 0 \ + interior(cbox a b \ {x. x\k \ c}) = {}" + unfolding interval_split[OF k] content_eq_0_interior by auto + guess u1 v1 using d(4)[OF assms(2)] by (elim exE) note uv1=this + guess u2 v2 using d(4)[OF assms(3)] by (elim exE) note uv2=this + have **: "\s t u. s \ t = {} \ u \ s \ u \ t \ u = {}" + by auto + show ?thesis + unfolding uv1 uv2 * + apply (rule **[OF d(5)[OF assms(2-4)]]) + apply (simp add: uv1) + using assms(5) uv1 by auto +qed + + +lemma division_split: + fixes a :: "'a::euclidean_space" + assumes "p division_of (cbox a b)" + and k: "k\Basis" + shows "{l \ {x. x\k \ c} | l. l \ p \ l \ {x. x\k \ c} \ {}} division_of(cbox a b \ {x. x\k \ c})" + (is "?p1 division_of ?I1") + and "{l \ {x. x\k \ c} | l. l \ p \ l \ {x. x\k \ c} \ {}} division_of (cbox a b \ {x. x\k \ c})" + (is "?p2 division_of ?I2") +proof (rule_tac[!] division_ofI) + note p = division_ofD[OF assms(1)] + show "finite ?p1" "finite ?p2" + using p(1) by auto + show "\?p1 = ?I1" "\?p2 = ?I2" + unfolding p(6)[symmetric] by auto + { + fix k + assume "k \ ?p1" + then guess l unfolding mem_Collect_eq by (elim exE conjE) note l=this + guess u v using p(4)[OF l(2)] by (elim exE) note uv=this + show "k \ ?I1" + using l p(2) uv by force + show "k \ {}" + by (simp add: l) + show "\a b. k = cbox a b" + apply (simp add: l uv p(2-3)[OF l(2)]) + apply (subst interval_split[OF k]) + apply (auto intro: order.trans) + done + fix k' + assume "k' \ ?p1" + then guess l' unfolding mem_Collect_eq by (elim exE conjE) note l'=this + assume "k \ k'" + then show "interior k \ interior k' = {}" + unfolding l l' using p(5)[OF l(2) l'(2)] by auto + } + { + fix k + assume "k \ ?p2" + then guess l unfolding mem_Collect_eq by (elim exE conjE) note l=this + guess u v using p(4)[OF l(2)] by (elim exE) note uv=this + show "k \ ?I2" + using l p(2) uv by force + show "k \ {}" + by (simp add: l) + show "\a b. k = cbox a b" + apply (simp add: l uv p(2-3)[OF l(2)]) + apply (subst interval_split[OF k]) + apply (auto intro: order.trans) + done + fix k' + assume "k' \ ?p2" + then guess l' unfolding mem_Collect_eq by (elim exE conjE) note l'=this + assume "k \ k'" + then show "interior k \ interior k' = {}" + unfolding l l' using p(5)[OF l(2) l'(2)] by auto + } +qed + +subsection \Tagged (partial) divisions.\ + +definition tagged_partial_division_of (infixr "tagged'_partial'_division'_of" 40) + where "s tagged_partial_division_of i \ + finite s \ + (\x k. (x, k) \ s \ x \ k \ k \ i \ (\a b. k = cbox a b)) \ + (\x1 k1 x2 k2. (x1, k1) \ s \ (x2, k2) \ s \ (x1, k1) \ (x2, k2) \ + interior k1 \ interior k2 = {})" + +lemma tagged_partial_division_ofD[dest]: + assumes "s tagged_partial_division_of i" + shows "finite s" + and "\x k. (x,k) \ s \ x \ k" + and "\x k. (x,k) \ s \ k \ i" + and "\x k. (x,k) \ s \ \a b. k = cbox a b" + and "\x1 k1 x2 k2. (x1,k1) \ s \ + (x2, k2) \ s \ (x1, k1) \ (x2, k2) \ interior k1 \ interior k2 = {}" + using assms unfolding tagged_partial_division_of_def by blast+ + +definition tagged_division_of (infixr "tagged'_division'_of" 40) + where "s tagged_division_of i \ s tagged_partial_division_of i \ (\{k. \x. (x,k) \ s} = i)" + +lemma tagged_division_of_finite: "s tagged_division_of i \ finite s" + unfolding tagged_division_of_def tagged_partial_division_of_def by auto + +lemma tagged_division_of: + "s tagged_division_of i \ + finite s \ + (\x k. (x, k) \ s \ x \ k \ k \ i \ (\a b. k = cbox a b)) \ + (\x1 k1 x2 k2. (x1, k1) \ s \ (x2, k2) \ s \ (x1, k1) \ (x2, k2) \ + interior k1 \ interior k2 = {}) \ + (\{k. \x. (x,k) \ s} = i)" + unfolding tagged_division_of_def tagged_partial_division_of_def by auto + +lemma tagged_division_ofI: + assumes "finite s" + and "\x k. (x,k) \ s \ x \ k" + and "\x k. (x,k) \ s \ k \ i" + and "\x k. (x,k) \ s \ \a b. k = cbox a b" + and "\x1 k1 x2 k2. (x1,k1) \ s \ (x2, k2) \ s \ (x1, k1) \ (x2, k2) \ + interior k1 \ interior k2 = {}" + and "(\{k. \x. (x,k) \ s} = i)" + shows "s tagged_division_of i" + unfolding tagged_division_of + using assms + apply auto + apply fastforce+ + done + +lemma tagged_division_ofD[dest]: (*FIXME USE A LOCALE*) + assumes "s tagged_division_of i" + shows "finite s" + and "\x k. (x,k) \ s \ x \ k" + and "\x k. (x,k) \ s \ k \ i" + and "\x k. (x,k) \ s \ \a b. k = cbox a b" + and "\x1 k1 x2 k2. (x1, k1) \ s \ (x2, k2) \ s \ (x1, k1) \ (x2, k2) \ + interior k1 \ interior k2 = {}" + and "(\{k. \x. (x,k) \ s} = i)" + using assms unfolding tagged_division_of by blast+ + +lemma division_of_tagged_division: + assumes "s tagged_division_of i" + shows "(snd ` s) division_of i" +proof (rule division_ofI) + note assm = tagged_division_ofD[OF assms] + show "\(snd ` s) = i" "finite (snd ` s)" + using assm by auto + fix k + assume k: "k \ snd ` s" + then obtain xk where xk: "(xk, k) \ s" + by auto + then show "k \ i" "k \ {}" "\a b. k = cbox a b" + using assm by fastforce+ + fix k' + assume k': "k' \ snd ` s" "k \ k'" + from this(1) obtain xk' where xk': "(xk', k') \ s" + by auto + then show "interior k \ interior k' = {}" + using assm(5) k'(2) xk by blast +qed + +lemma partial_division_of_tagged_division: + assumes "s tagged_partial_division_of i" + shows "(snd ` s) division_of \(snd ` s)" +proof (rule division_ofI) + note assm = tagged_partial_division_ofD[OF assms] + show "finite (snd ` s)" "\(snd ` s) = \(snd ` s)" + using assm by auto + fix k + assume k: "k \ snd ` s" + then obtain xk where xk: "(xk, k) \ s" + by auto + then show "k \ {}" "\a b. k = cbox a b" "k \ \(snd ` s)" + using assm by auto + fix k' + assume k': "k' \ snd ` s" "k \ k'" + from this(1) obtain xk' where xk': "(xk', k') \ s" + by auto + then show "interior k \ interior k' = {}" + using assm(5) k'(2) xk by auto +qed + +lemma tagged_partial_division_subset: + assumes "s tagged_partial_division_of i" + and "t \ s" + shows "t tagged_partial_division_of i" + using assms + unfolding tagged_partial_division_of_def + using finite_subset[OF assms(2)] + by blast + +lemma (in comm_monoid_set) over_tagged_division_lemma: + assumes "p tagged_division_of i" + and "\u v. cbox u v \ {} \ content (cbox u v) = 0 \ d (cbox u v) = \<^bold>1" + shows "F (\(x,k). d k) p = F d (snd ` p)" +proof - + have *: "(\(x,k). d k) = d \ snd" + unfolding o_def by (rule ext) auto + note assm = tagged_division_ofD[OF assms(1)] + show ?thesis + unfolding * + proof (rule reindex_nontrivial[symmetric]) + show "finite p" + using assm by auto + fix x y + assume "x\p" "y\p" "x\y" "snd x = snd y" + obtain a b where ab: "snd x = cbox a b" + using assm(4)[of "fst x" "snd x"] \x\p\ by auto + have "(fst x, snd y) \ p" "(fst x, snd y) \ y" + by (metis prod.collapse \x\p\ \snd x = snd y\ \x \ y\)+ + with \x\p\ \y\p\ have "interior (snd x) \ interior (snd y) = {}" + by (intro assm(5)[of "fst x" _ "fst y"]) auto + then have "content (cbox a b) = 0" + unfolding \snd x = snd y\[symmetric] ab content_eq_0_interior by auto + then have "d (cbox a b) = \<^bold>1" + using assm(2)[of "fst x" "snd x"] \x\p\ ab[symmetric] by (intro assms(2)) auto + then show "d (snd x) = \<^bold>1" + unfolding ab by auto + qed +qed + +lemma tag_in_interval: "p tagged_division_of i \ (x, k) \ p \ x \ i" + by auto + +lemma tagged_division_of_empty: "{} tagged_division_of {}" + unfolding tagged_division_of by auto + +lemma tagged_partial_division_of_trivial[simp]: "p tagged_partial_division_of {} \ p = {}" + unfolding tagged_partial_division_of_def by auto + +lemma tagged_division_of_trivial[simp]: "p tagged_division_of {} \ p = {}" + unfolding tagged_division_of by auto + +lemma tagged_division_of_self: "x \ cbox a b \ {(x,cbox a b)} tagged_division_of (cbox a b)" + by (rule tagged_division_ofI) auto + +lemma tagged_division_of_self_real: "x \ {a .. b::real} \ {(x,{a .. b})} tagged_division_of {a .. b}" + unfolding box_real[symmetric] + by (rule tagged_division_of_self) + +lemma tagged_division_union: + assumes "p1 tagged_division_of s1" + and "p2 tagged_division_of s2" + and "interior s1 \ interior s2 = {}" + shows "(p1 \ p2) tagged_division_of (s1 \ s2)" +proof (rule tagged_division_ofI) + note p1 = tagged_division_ofD[OF assms(1)] + note p2 = tagged_division_ofD[OF assms(2)] + show "finite (p1 \ p2)" + using p1(1) p2(1) by auto + show "\{k. \x. (x, k) \ p1 \ p2} = s1 \ s2" + using p1(6) p2(6) by blast + fix x k + assume xk: "(x, k) \ p1 \ p2" + show "x \ k" "\a b. k = cbox a b" + using xk p1(2,4) p2(2,4) by auto + show "k \ s1 \ s2" + using xk p1(3) p2(3) by blast + fix x' k' + assume xk': "(x', k') \ p1 \ p2" "(x, k) \ (x', k')" + have *: "\a b. a \ s1 \ b \ s2 \ interior a \ interior b = {}" + using assms(3) interior_mono by blast + show "interior k \ interior k' = {}" + apply (cases "(x, k) \ p1") + apply (meson "*" UnE assms(1) assms(2) p1(5) tagged_division_ofD(3) xk'(1) xk'(2)) + by (metis "*" UnE assms(1) assms(2) inf_sup_aci(1) p2(5) tagged_division_ofD(3) xk xk'(1) xk'(2)) +qed + +lemma tagged_division_unions: + assumes "finite iset" + and "\i\iset. pfn i tagged_division_of i" + and "\i1\iset. \i2\iset. i1 \ i2 \ interior(i1) \ interior(i2) = {}" + shows "\(pfn ` iset) tagged_division_of (\iset)" +proof (rule tagged_division_ofI) + note assm = tagged_division_ofD[OF assms(2)[rule_format]] + show "finite (\(pfn ` iset))" + apply (rule finite_Union) + using assms + apply auto + done + have "\{k. \x. (x, k) \ \(pfn ` iset)} = \((\i. \{k. \x. (x, k) \ pfn i}) ` iset)" + by blast + also have "\ = \iset" + using assm(6) by auto + finally show "\{k. \x. (x, k) \ \(pfn ` iset)} = \iset" . + fix x k + assume xk: "(x, k) \ \(pfn ` iset)" + then obtain i where i: "i \ iset" "(x, k) \ pfn i" + by auto + show "x \ k" "\a b. k = cbox a b" "k \ \iset" + using assm(2-4)[OF i] using i(1) by auto + fix x' k' + assume xk': "(x', k') \ \(pfn ` iset)" "(x, k) \ (x', k')" + then obtain i' where i': "i' \ iset" "(x', k') \ pfn i'" + by auto + have *: "\a b. i \ i' \ a \ i \ b \ i' \ interior a \ interior b = {}" + using i(1) i'(1) + using assms(3)[rule_format] interior_mono + by blast + show "interior k \ interior k' = {}" + apply (cases "i = i'") + using assm(5) i' i(2) xk'(2) apply blast + using "*" assm(3) i' i by auto +qed + +lemma tagged_partial_division_of_union_self: + assumes "p tagged_partial_division_of s" + shows "p tagged_division_of (\(snd ` p))" + apply (rule tagged_division_ofI) + using tagged_partial_division_ofD[OF assms] + apply auto + done + +lemma tagged_division_of_union_self: + assumes "p tagged_division_of s" + shows "p tagged_division_of (\(snd ` p))" + apply (rule tagged_division_ofI) + using tagged_division_ofD[OF assms] + apply auto + done + +subsection \Functions closed on boxes: morphisms from boxes to monoids\ + +text \This auxiliary structure is used to sum up over the elements of a division. Main theorem is + @{text operative_division}. Instances for the monoid are @{typ "'a option"}, @{typ real}, and + @{typ bool}.\ + +lemma property_empty_interval: "\a b. content (cbox a b) = 0 \ P (cbox a b) \ P {}" + using content_empty unfolding empty_as_interval by auto + +paragraph \Using additivity of lifted function to encode definedness.\ + +definition lift_option :: "('a \ 'b \ 'c) \ 'a option \ 'b option \ 'c option" +where + "lift_option f a' b' = Option.bind a' (\a. Option.bind b' (\b. Some (f a b)))" + +lemma lift_option_simps[simp]: + "lift_option f (Some a) (Some b) = Some (f a b)" + "lift_option f None b' = None" + "lift_option f a' None = None" + by (auto simp: lift_option_def) + +lemma (in comm_monoid) comm_monoid_lift_option: "comm_monoid (lift_option op \<^bold>*) (Some \<^bold>1)" + proof qed (auto simp: lift_option_def ac_simps split: bind_split) + +lemma (in comm_monoid) comm_monoid_set_lift_option: "comm_monoid_set (lift_option op \<^bold>*) (Some \<^bold>1)" + proof qed (auto simp: lift_option_def ac_simps split: bind_split) + +lemma comm_monoid_and: "comm_monoid op \ True" + proof qed auto + +lemma comm_monoid_set_and: "comm_monoid_set op \ True" + proof qed auto + +paragraph \Operative\ + +definition (in comm_monoid) operative :: "('b::euclidean_space set \ 'a) \ bool" + where "operative g \ + (\a b. content (cbox a b) = 0 \ g (cbox a b) = \<^bold>1) \ + (\a b c. \k\Basis. g (cbox a b) = g (cbox a b \ {x. x\k \ c}) \<^bold>* g (cbox a b \ {x. x\k \ c}))" + +lemma (in comm_monoid) operativeD[dest]: + assumes "operative g" + shows "\a b. content (cbox a b) = 0 \ g (cbox a b) = \<^bold>1" + and "\a b c k. k \ Basis \ g (cbox a b) = g (cbox a b \ {x. x\k \ c}) \<^bold>* g (cbox a b \ {x. x\k \ c})" + using assms unfolding operative_def by auto + +lemma (in comm_monoid) operative_empty: "operative g \ g {} = \<^bold>1" + unfolding operative_def by (rule property_empty_interval) auto + +lemma operative_content[intro]: "add.operative content" + by (force simp add: add.operative_def content_split[symmetric]) + +definition "division_points (k::('a::euclidean_space) set) d = + {(j,x). j \ Basis \ (interval_lowerbound k)\j < x \ x < (interval_upperbound k)\j \ + (\i\d. (interval_lowerbound i)\j = x \ (interval_upperbound i)\j = x)}" + +lemma division_points_finite: + fixes i :: "'a::euclidean_space set" + assumes "d division_of i" + shows "finite (division_points i d)" +proof - + note assm = division_ofD[OF assms] + let ?M = "\j. {(j,x)|x. (interval_lowerbound i)\j < x \ x < (interval_upperbound i)\j \ + (\i\d. (interval_lowerbound i)\j = x \ (interval_upperbound i)\j = x)}" + have *: "division_points i d = \(?M ` Basis)" + unfolding division_points_def by auto + show ?thesis + unfolding * using assm by auto +qed + +lemma division_points_subset: + fixes a :: "'a::euclidean_space" + assumes "d division_of (cbox a b)" + and "\i\Basis. a\i < b\i" "a\k < c" "c < b\k" + and k: "k \ Basis" + shows "division_points (cbox a b \ {x. x\k \ c}) {l \ {x. x\k \ c} | l . l \ d \ l \ {x. x\k \ c} \ {}} \ + division_points (cbox a b) d" (is ?t1) + and "division_points (cbox a b \ {x. x\k \ c}) {l \ {x. x\k \ c} | l . l \ d \ ~(l \ {x. x\k \ c} = {})} \ + division_points (cbox a b) d" (is ?t2) +proof - + note assm = division_ofD[OF assms(1)] + have *: "\i\Basis. a\i \ b\i" + "\i\Basis. a\i \ (\i\Basis. (if i = k then min (b \ k) c else b \ i) *\<^sub>R i) \ i" + "\i\Basis. (\i\Basis. (if i = k then max (a \ k) c else a \ i) *\<^sub>R i) \ i \ b\i" + "min (b \ k) c = c" "max (a \ k) c = c" + using assms using less_imp_le by auto + show ?t1 (*FIXME a horrible mess*) + unfolding division_points_def interval_split[OF k, of a b] + unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)] + unfolding * + apply (rule subsetI) + unfolding mem_Collect_eq split_beta + apply (erule bexE conjE)+ + apply (simp add: ) + apply (erule exE conjE)+ + proof + fix i l x + assume as: + "a \ fst x < snd x" "snd x < (if fst x = k then c else b \ fst x)" + "interval_lowerbound i \ fst x = snd x \ interval_upperbound i \ fst x = snd x" + "i = l \ {x. x \ k \ c}" "l \ d" "l \ {x. x \ k \ c} \ {}" + and fstx: "fst x \ Basis" + from assm(4)[OF this(5)] guess u v apply-by(erule exE)+ note l=this + have *: "\i\Basis. u \ i \ (\i\Basis. (if i = k then min (v \ k) c else v \ i) *\<^sub>R i) \ i" + using as(6) unfolding l interval_split[OF k] box_ne_empty as . + have **: "\i\Basis. u\i \ v\i" + using l using as(6) unfolding box_ne_empty[symmetric] by auto + show "\i\d. interval_lowerbound i \ fst x = snd x \ interval_upperbound i \ fst x = snd x" + apply (rule bexI[OF _ \l \ d\]) + using as(1-3,5) fstx + unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] as + apply (auto split: if_split_asm) + done + show "snd x < b \ fst x" + using as(2) \c < b\k\ by (auto split: if_split_asm) + qed + show ?t2 + unfolding division_points_def interval_split[OF k, of a b] + unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)] + unfolding * + unfolding subset_eq + apply rule + unfolding mem_Collect_eq split_beta + apply (erule bexE conjE)+ + apply (simp only: mem_Collect_eq inner_setsum_left_Basis simp_thms) + apply (erule exE conjE)+ + proof + fix i l x + assume as: + "(if fst x = k then c else a \ fst x) < snd x" "snd x < b \ fst x" + "interval_lowerbound i \ fst x = snd x \ interval_upperbound i \ fst x = snd x" + "i = l \ {x. c \ x \ k}" "l \ d" "l \ {x. c \ x \ k} \ {}" + and fstx: "fst x \ Basis" + from assm(4)[OF this(5)] guess u v by (elim exE) note l=this + have *: "\i\Basis. (\i\Basis. (if i = k then max (u \ k) c else u \ i) *\<^sub>R i) \ i \ v \ i" + using as(6) unfolding l interval_split[OF k] box_ne_empty as . + have **: "\i\Basis. u\i \ v\i" + using l using as(6) unfolding box_ne_empty[symmetric] by auto + show "\i\d. interval_lowerbound i \ fst x = snd x \ interval_upperbound i \ fst x = snd x" + apply (rule bexI[OF _ \l \ d\]) + using as(1-3,5) fstx + unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] as + apply (auto split: if_split_asm) + done + show "a \ fst x < snd x" + using as(1) \a\k < c\ by (auto split: if_split_asm) + qed +qed + +lemma division_points_psubset: + fixes a :: "'a::euclidean_space" + assumes "d division_of (cbox a b)" + and "\i\Basis. a\i < b\i" "a\k < c" "c < b\k" + and "l \ d" + and "interval_lowerbound l\k = c \ interval_upperbound l\k = c" + and k: "k \ Basis" + shows "division_points (cbox a b \ {x. x\k \ c}) {l \ {x. x\k \ c} | l. l\d \ l \ {x. x\k \ c} \ {}} \ + division_points (cbox a b) d" (is "?D1 \ ?D") + and "division_points (cbox a b \ {x. x\k \ c}) {l \ {x. x\k \ c} | l. l\d \ l \ {x. x\k \ c} \ {}} \ + division_points (cbox a b) d" (is "?D2 \ ?D") +proof - + have ab: "\i\Basis. a\i \ b\i" + using assms(2) by (auto intro!:less_imp_le) + guess u v using division_ofD(4)[OF assms(1,5)] by (elim exE) note l=this + have uv: "\i\Basis. u\i \ v\i" "\i\Basis. a\i \ u\i \ v\i \ b\i" + using division_ofD(2,2,3)[OF assms(1,5)] unfolding l box_ne_empty + using subset_box(1) + apply auto + apply blast+ + done + have *: "interval_upperbound (cbox a b \ {x. x \ k \ interval_upperbound l \ k}) \ k = interval_upperbound l \ k" + "interval_upperbound (cbox a b \ {x. x \ k \ interval_lowerbound l \ k}) \ k = interval_lowerbound l \ k" + unfolding l interval_split[OF k] interval_bounds[OF uv(1)] + using uv[rule_format, of k] ab k + by auto + have "\x. x \ ?D - ?D1" + using assms(3-) + unfolding division_points_def interval_bounds[OF ab] + apply - + apply (erule disjE) + apply (rule_tac x="(k,(interval_lowerbound l)\k)" in exI, force simp add: *) + apply (rule_tac x="(k,(interval_upperbound l)\k)" in exI, force simp add: *) + done + moreover have "?D1 \ ?D" + by (auto simp add: assms division_points_subset) + ultimately show "?D1 \ ?D" + by blast + have *: "interval_lowerbound (cbox a b \ {x. x \ k \ interval_lowerbound l \ k}) \ k = interval_lowerbound l \ k" + "interval_lowerbound (cbox a b \ {x. x \ k \ interval_upperbound l \ k}) \ k = interval_upperbound l \ k" + unfolding l interval_split[OF k] interval_bounds[OF uv(1)] + using uv[rule_format, of k] ab k + by auto + have "\x. x \ ?D - ?D2" + using assms(3-) + unfolding division_points_def interval_bounds[OF ab] + apply - + apply (erule disjE) + apply (rule_tac x="(k,(interval_lowerbound l)\k)" in exI, force simp add: *) + apply (rule_tac x="(k,(interval_upperbound l)\k)" in exI, force simp add: *) + done + moreover have "?D2 \ ?D" + by (auto simp add: assms division_points_subset) + ultimately show "?D2 \ ?D" + by blast +qed + +lemma (in comm_monoid_set) operative_division: + fixes g :: "'b::euclidean_space set \ 'a" + assumes g: "operative g" and d: "d division_of (cbox a b)" shows "F g d = g (cbox a b)" +proof - + define C where [abs_def]: "C = card (division_points (cbox a b) d)" + then show ?thesis + using d + proof (induction C arbitrary: a b d rule: less_induct) + case (less a b d) + show ?case + proof cases + show "content (cbox a b) = 0 \ F g d = g (cbox a b)" + using division_of_content_0[OF _ less.prems] operativeD(1)[OF g] division_ofD(4)[OF less.prems] + by (fastforce intro!: neutral) + next + assume "content (cbox a b) \ 0" + note ab = this[unfolded content_lt_nz[symmetric] content_pos_lt_eq] + then have ab': "\i\Basis. a\i \ b\i" + by (auto intro!: less_imp_le) + show "F g d = g (cbox a b)" + proof (cases "division_points (cbox a b) d = {}") + case True + { fix u v and j :: 'b + assume j: "j \ Basis" and as: "cbox u v \ d" + then have "cbox u v \ {}" + using less.prems by blast + then have uv: "\i\Basis. u\i \ v\i" "u\j \ v\j" + using j unfolding box_ne_empty by auto + have *: "\p r Q. \ j\Basis \ p \ r \ (\x\d. Q x) \ p \ r \ Q (cbox u v)" + using as j by auto + have "(j, u\j) \ division_points (cbox a b) d" + "(j, v\j) \ division_points (cbox a b) d" using True by auto + note this[unfolded de_Morgan_conj division_points_def mem_Collect_eq split_conv interval_bounds[OF ab'] bex_simps] + note *[OF this(1)] *[OF this(2)] note this[unfolded interval_bounds[OF uv(1)]] + moreover + have "a\j \ u\j" "v\j \ b\j" + using division_ofD(2,2,3)[OF \d division_of cbox a b\ as] + apply (metis j subset_box(1) uv(1)) + by (metis \cbox u v \ cbox a b\ j subset_box(1) uv(1)) + ultimately have "u\j = a\j \ v\j = a\j \ u\j = b\j \ v\j = b\j \ u\j = a\j \ v\j = b\j" + unfolding not_less de_Morgan_disj using ab[rule_format,of j] uv(2) j by force } + then have d': "\i\d. \u v. i = cbox u v \ + (\j\Basis. u\j = a\j \ v\j = a\j \ u\j = b\j \ v\j = b\j \ u\j = a\j \ v\j = b\j)" + unfolding forall_in_division[OF less.prems] by blast + have "(1/2) *\<^sub>R (a+b) \ cbox a b" + unfolding mem_box using ab by(auto intro!: less_imp_le simp: inner_simps) + note this[unfolded division_ofD(6)[OF \d division_of cbox a b\,symmetric] Union_iff] + then guess i .. note i=this + guess u v using d'[rule_format,OF i(1)] by (elim exE conjE) note uv=this + have "cbox a b \ d" + proof - + have "u = a" "v = b" + unfolding euclidean_eq_iff[where 'a='b] + proof safe + fix j :: 'b + assume j: "j \ Basis" + note i(2)[unfolded uv mem_box,rule_format,of j] + then show "u \ j = a \ j" and "v \ j = b \ j" + using uv(2)[rule_format,of j] j by (auto simp: inner_simps) + qed + then have "i = cbox a b" using uv by auto + then show ?thesis using i by auto + qed + then have deq: "d = insert (cbox a b) (d - {cbox a b})" + by auto + have "F g (d - {cbox a b}) = \<^bold>1" + proof (intro neutral ballI) + fix x + assume x: "x \ d - {cbox a b}" + then have "x\d" + by auto note d'[rule_format,OF this] + then guess u v by (elim exE conjE) note uv=this + have "u \ a \ v \ b" + using x[unfolded uv] by auto + then obtain j where "u\j \ a\j \ v\j \ b\j" and j: "j \ Basis" + unfolding euclidean_eq_iff[where 'a='b] by auto + then have "u\j = v\j" + using uv(2)[rule_format,OF j] by auto + then have "content (cbox u v) = 0" + unfolding content_eq_0 using j + by force + then show "g x = \<^bold>1" + unfolding uv(1) by (rule operativeD(1)[OF g]) + qed + then show "F g d = g (cbox a b)" + using division_ofD[OF less.prems] + apply (subst deq) + apply (subst insert) + apply auto + done + next + case False + then have "\x. x \ division_points (cbox a b) d" + by auto + then guess k c + unfolding split_paired_Ex division_points_def mem_Collect_eq split_conv + apply (elim exE conjE) + done + note this(2-4,1) note kc=this[unfolded interval_bounds[OF ab']] + from this(3) guess j .. note j=this + define d1 where "d1 = {l \ {x. x\k \ c} | l. l \ d \ l \ {x. x\k \ c} \ {}}" + define d2 where "d2 = {l \ {x. x\k \ c} | l. l \ d \ l \ {x. x\k \ c} \ {}}" + define cb where "cb = (\i\Basis. (if i = k then c else b\i) *\<^sub>R i)" + define ca where "ca = (\i\Basis. (if i = k then c else a\i) *\<^sub>R i)" + note division_points_psubset[OF \d division_of cbox a b\ ab kc(1-2) j] + note psubset_card_mono[OF _ this(1)] psubset_card_mono[OF _ this(2)] + then have *: "F g d1 = g (cbox a b \ {x. x\k \ c})" "F g d2 = g (cbox a b \ {x. x\k \ c})" + unfolding interval_split[OF kc(4)] + apply (rule_tac[!] "less.hyps"[rule_format]) + using division_split[OF \d division_of cbox a b\, where k=k and c=c] + apply (simp_all add: interval_split kc d1_def d2_def division_points_finite[OF \d division_of cbox a b\]) + done + { fix l y + assume as: "l \ d" "y \ d" "l \ {x. x \ k \ c} = y \ {x. x \ k \ c}" "l \ y" + from division_ofD(4)[OF \d division_of cbox a b\ this(1)] guess u v by (elim exE) note leq=this + have "g (l \ {x. x \ k \ c}) = \<^bold>1" + unfolding leq interval_split[OF kc(4)] + apply (rule operativeD[OF g]) + unfolding interval_split[symmetric, OF kc(4)] + using division_split_left_inj less as kc leq by blast + } note fxk_le = this + { fix l y + assume as: "l \ d" "y \ d" "l \ {x. c \ x \ k} = y \ {x. c \ x \ k}" "l \ y" + from division_ofD(4)[OF \d division_of cbox a b\ this(1)] guess u v by (elim exE) note leq=this + have "g (l \ {x. x \ k \ c}) = \<^bold>1" + unfolding leq interval_split[OF kc(4)] + apply (rule operativeD(1)[OF g]) + unfolding interval_split[symmetric,OF kc(4)] + using division_split_right_inj less leq as kc by blast + } note fxk_ge = this + have d1_alt: "d1 = (\l. l \ {x. x\k \ c}) ` {l \ d. l \ {x. x\k \ c} \ {}}" + using d1_def by auto + have d2_alt: "d2 = (\l. l \ {x. x\k \ c}) ` {l \ d. l \ {x. x\k \ c} \ {}}" + using d2_def by auto + have "g (cbox a b) = F g d1 \<^bold>* F g d2" (is "_ = ?prev") + unfolding * using g kc(4) by blast + also have "F g d1 = F (\l. g (l \ {x. x\k \ c})) d" + unfolding d1_alt using division_of_finite[OF less.prems] fxk_le + by (subst reindex_nontrivial) (auto intro!: mono_neutral_cong_left simp: operative_empty[OF g]) + also have "F g d2 = F (\l. g (l \ {x. x\k \ c})) d" + unfolding d2_alt using division_of_finite[OF less.prems] fxk_ge + by (subst reindex_nontrivial) (auto intro!: mono_neutral_cong_left simp: operative_empty[OF g]) + also have *: "\x\d. g x = g (x \ {x. x \ k \ c}) \<^bold>* g (x \ {x. c \ x \ k})" + unfolding forall_in_division[OF \d division_of cbox a b\] + using g kc(4) by blast + have "F (\l. g (l \ {x. x\k \ c})) d \<^bold>* F (\l. g (l \ {x. x\k \ c})) d = F g d" + using * by (simp add: distrib) + finally show ?thesis by auto + qed + qed + qed +qed + +lemma (in comm_monoid_set) operative_tagged_division: + assumes f: "operative g" and d: "d tagged_division_of (cbox a b)" + shows "F (\(x, l). g l) d = g (cbox a b)" + unfolding d[THEN division_of_tagged_division, THEN operative_division[OF f], symmetric] + by (simp add: f[THEN operativeD(1)] over_tagged_division_lemma[OF d]) + +lemma additive_content_division: "d division_of (cbox a b) \ setsum content d = content (cbox a b)" + by (metis operative_content setsum.operative_division) + +lemma additive_content_tagged_division: + "d tagged_division_of (cbox a b) \ setsum (\(x,l). content l) d = content (cbox a b)" + unfolding setsum.operative_tagged_division[OF operative_content, symmetric] by blast + +lemma + shows real_inner_1_left: "inner 1 x = x" + and real_inner_1_right: "inner x 1 = x" + by simp_all + +lemma content_real_eq_0: "content {a .. b::real} = 0 \ a \ b" + by (metis atLeastatMost_empty_iff2 content_empty content_real diff_self eq_iff le_cases le_iff_diff_le_0) + +lemma interval_real_split: + "{a .. b::real} \ {x. x \ c} = {a .. min b c}" + "{a .. b} \ {x. c \ x} = {max a c .. b}" + apply (metis Int_atLeastAtMostL1 atMost_def) + apply (metis Int_atLeastAtMostL2 atLeast_def) + done + +lemma (in comm_monoid) operative_1_lt: + "operative (g :: real set \ 'a) \ + ((\a b. b \ a \ g {a .. b} = \<^bold>1) \ (\a b c. a < c \ c < b \ g {a .. c} \<^bold>* g {c .. b} = g {a .. b}))" + apply (simp add: operative_def content_real_eq_0 atMost_def[symmetric] atLeast_def[symmetric] + del: content_real_if) +proof safe + fix a b c :: real + assume *: "\a b c. g {a..b} = g {a..min b c} \<^bold>* g {max a c..b}" + assume "a < c" "c < b" + with *[rule_format, of a b c] show "g {a..c} \<^bold>* g {c..b} = g {a..b}" + by (simp add: less_imp_le min.absorb2 max.absorb2) +next + fix a b c :: real + assume as: "\a b. b \ a \ g {a..b} = \<^bold>1" + "\a b c. a < c \ c < b \ g {a..c} \<^bold>* g {c..b} = g {a..b}" + from as(1)[rule_format, of 0 1] as(1)[rule_format, of a a for a] as(2) + have [simp]: "g {} = \<^bold>1" "\a. g {a} = \<^bold>1" + "\a b c. a < c \ c < b \ g {a..c} \<^bold>* g {c..b} = g {a..b}" + by auto + show "g {a..b} = g {a..min b c} \<^bold>* g {max a c..b}" + by (auto simp: min_def max_def le_less) +qed + +lemma (in comm_monoid) operative_1_le: + "operative (g :: real set \ 'a) \ + ((\a b. b \ a \ g {a..b} = \<^bold>1) \ (\a b c. a \ c \ c \ b \ g {a .. c} \<^bold>* g {c .. b} = g {a .. b}))" + unfolding operative_1_lt +proof safe + fix a b c :: real + assume as: "\a b c. a \ c \ c \ b \ g {a..c} \<^bold>* g {c..b} = g {a..b}" "a < c" "c < b" + show "g {a..c} \<^bold>* g {c..b} = g {a..b}" + apply (rule as(1)[rule_format]) + using as(2-) + apply auto + done +next + fix a b c :: real + assume "\a b. b \ a \ g {a .. b} = \<^bold>1" + and "\a b c. a < c \ c < b \ g {a..c} \<^bold>* g {c..b} = g {a..b}" + and "a \ c" + and "c \ b" + note as = this[rule_format] + show "g {a..c} \<^bold>* g {c..b} = g {a..b}" + proof (cases "c = a \ c = b") + case False + then show ?thesis + apply - + apply (subst as(2)) + using as(3-) + apply auto + done + next + case True + then show ?thesis + proof + assume *: "c = a" + then have "g {a .. c} = \<^bold>1" + apply - + apply (rule as(1)[rule_format]) + apply auto + done + then show ?thesis + unfolding * by auto + next + assume *: "c = b" + then have "g {c .. b} = \<^bold>1" + apply - + apply (rule as(1)[rule_format]) + apply auto + done + then show ?thesis + unfolding * by auto + qed + qed +qed + +subsection \Fine-ness of a partition w.r.t. a gauge.\ + +definition fine (infixr "fine" 46) + where "d fine s \ (\(x,k) \ s. k \ d x)" + +lemma fineI: + assumes "\x k. (x, k) \ s \ k \ d x" + shows "d fine s" + using assms unfolding fine_def by auto + +lemma fineD[dest]: + assumes "d fine s" + shows "\x k. (x,k) \ s \ k \ d x" + using assms unfolding fine_def by auto + +lemma fine_inter: "(\x. d1 x \ d2 x) fine p \ d1 fine p \ d2 fine p" + unfolding fine_def by auto + +lemma fine_inters: + "(\x. \{f d x | d. d \ s}) fine p \ (\d\s. (f d) fine p)" + unfolding fine_def by blast + +lemma fine_union: "d fine p1 \ d fine p2 \ d fine (p1 \ p2)" + unfolding fine_def by blast + +lemma fine_unions: "(\p. p \ ps \ d fine p) \ d fine (\ps)" + unfolding fine_def by auto + +lemma fine_subset: "p \ q \ d fine q \ d fine p" + unfolding fine_def by blast + + +subsection \Gauge integral. Define on compact intervals first, then use a limit.\ + +definition has_integral_compact_interval (infixr "has'_integral'_compact'_interval" 46) + where "(f has_integral_compact_interval y) i \ + (\e>0. \d. gauge d \ + (\p. p tagged_division_of i \ d fine p \ + norm (setsum (\(x,k). content k *\<^sub>R f x) p - y) < e))" + +definition has_integral :: + "('n::euclidean_space \ 'b::real_normed_vector) \ 'b \ 'n set \ bool" + (infixr "has'_integral" 46) + where "(f has_integral y) i \ + (if \a b. i = cbox a b + then (f has_integral_compact_interval y) i + else (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ + (\z. ((\x. if x \ i then f x else 0) has_integral_compact_interval z) (cbox a b) \ + norm (z - y) < e)))" + +lemma has_integral: + "(f has_integral y) (cbox a b) \ + (\e>0. \d. gauge d \ + (\p. p tagged_division_of (cbox a b) \ d fine p \ + norm (setsum (\(x,k). content(k) *\<^sub>R f x) p - y) < e))" + unfolding has_integral_def has_integral_compact_interval_def + by auto + +lemma has_integral_real: + "(f has_integral y) {a .. b::real} \ + (\e>0. \d. gauge d \ + (\p. p tagged_division_of {a .. b} \ d fine p \ + norm (setsum (\(x,k). content(k) *\<^sub>R f x) p - y) < e))" + unfolding box_real[symmetric] + by (rule has_integral) + +lemma has_integralD[dest]: + assumes "(f has_integral y) (cbox a b)" + and "e > 0" + obtains d where "gauge d" + and "\p. p tagged_division_of (cbox a b) \ d fine p \ + norm (setsum (\(x,k). content(k) *\<^sub>R f(x)) p - y) < e" + using assms unfolding has_integral by auto + +lemma has_integral_alt: + "(f has_integral y) i \ + (if \a b. i = cbox a b + then (f has_integral y) i + else (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ + (\z. ((\x. if x \ i then f(x) else 0) has_integral z) (cbox a b) \ norm (z - y) < e)))" + unfolding has_integral + unfolding has_integral_compact_interval_def has_integral_def + by auto + +lemma has_integral_altD: + assumes "(f has_integral y) i" + and "\ (\a b. i = cbox a b)" + and "e>0" + obtains B where "B > 0" + and "\a b. ball 0 B \ cbox a b \ + (\z. ((\x. if x \ i then f(x) else 0) has_integral z) (cbox a b) \ norm(z - y) < e)" + using assms + unfolding has_integral + unfolding has_integral_compact_interval_def has_integral_def + by auto + +definition integrable_on (infixr "integrable'_on" 46) + where "f integrable_on i \ (\y. (f has_integral y) i)" + +definition "integral i f = (SOME y. (f has_integral y) i \ ~ f integrable_on i \ y=0)" + +lemma integrable_integral[dest]: "f integrable_on i \ (f has_integral (integral i f)) i" + unfolding integrable_on_def integral_def by (metis (mono_tags, lifting) someI_ex) + +lemma not_integrable_integral: "~ f integrable_on i \ integral i f = 0" + unfolding integrable_on_def integral_def by blast + +lemma has_integral_integrable[intro]: "(f has_integral i) s \ f integrable_on s" + unfolding integrable_on_def by auto + +lemma has_integral_integral: "f integrable_on s \ (f has_integral (integral s f)) s" + by auto + +lemma setsum_content_null: + assumes "content (cbox a b) = 0" + and "p tagged_division_of (cbox a b)" + shows "setsum (\(x,k). content k *\<^sub>R f x) p = (0::'a::real_normed_vector)" +proof (rule setsum.neutral, rule) + fix y + assume y: "y \ p" + obtain x k where xk: "y = (x, k)" + using surj_pair[of y] by blast + note assm = tagged_division_ofD(3-4)[OF assms(2) y[unfolded xk]] + from this(2) obtain c d where k: "k = cbox c d" by blast + have "(\(x, k). content k *\<^sub>R f x) y = content k *\<^sub>R f x" + unfolding xk by auto + also have "\ = 0" + using content_subset[OF assm(1)[unfolded k]] content_pos_le[of c d] + unfolding assms(1) k + by auto + finally show "(\(x, k). content k *\<^sub>R f x) y = 0" . +qed + + +subsection \Some basic combining lemmas.\ + +lemma tagged_division_unions_exists: + assumes "finite iset" + and "\i\iset. \p. p tagged_division_of i \ d fine p" + and "\i1\iset. \i2\iset. i1 \ i2 \ interior i1 \ interior i2 = {}" + and "\iset = i" + obtains p where "p tagged_division_of i" and "d fine p" +proof - + obtain pfn where pfn: + "\x. x \ iset \ pfn x tagged_division_of x" + "\x. x \ iset \ d fine pfn x" + using bchoice[OF assms(2)] by auto + show thesis + apply (rule_tac p="\(pfn ` iset)" in that) + using assms(1) assms(3) assms(4) pfn(1) tagged_division_unions apply force + by (metis (mono_tags, lifting) fine_unions imageE pfn(2)) +qed + + +subsection \The set we're concerned with must be closed.\ + +lemma division_of_closed: + fixes i :: "'n::euclidean_space set" + shows "s division_of i \ closed i" + unfolding division_of_def by fastforce + +subsection \General bisection principle for intervals; might be useful elsewhere.\ + +lemma interval_bisection_step: + fixes type :: "'a::euclidean_space" + assumes "P {}" + and "\s t. P s \ P t \ interior(s) \ interior(t) = {} \ P (s \ t)" + and "\ P (cbox a (b::'a))" + obtains c d where "\ P (cbox c d)" + and "\i\Basis. a\i \ c\i \ c\i \ d\i \ d\i \ b\i \ 2 * (d\i - c\i) \ b\i - a\i" +proof - + have "cbox a b \ {}" + using assms(1,3) by metis + then have ab: "\i. i\Basis \ a \ i \ b \ i" + by (force simp: mem_box) + { fix f + have "\finite f; + \s. s\f \ P s; + \s. s\f \ \a b. s = cbox a b; + \s t. s\f \ t\f \ s \ t \ interior s \ interior t = {}\ \ P (\f)" + proof (induct f rule: finite_induct) + case empty + show ?case + using assms(1) by auto + next + case (insert x f) + show ?case + unfolding Union_insert + apply (rule assms(2)[rule_format]) + using inter_interior_unions_intervals [of f "interior x"] + apply (auto simp: insert) + by (metis IntI empty_iff insert.hyps(2) insert.prems(3) insert_iff) + qed + } note UN_cases = this + let ?A = "{cbox c d | c d::'a. \i\Basis. (c\i = a\i) \ (d\i = (a\i + b\i) / 2) \ + (c\i = (a\i + b\i) / 2) \ (d\i = b\i)}" + let ?PP = "\c d. \i\Basis. a\i \ c\i \ c\i \ d\i \ d\i \ b\i \ 2 * (d\i - c\i) \ b\i - a\i" + { + presume "\c d. ?PP c d \ P (cbox c d) \ False" + then show thesis + unfolding atomize_not not_all + by (blast intro: that) + } + assume as: "\c d. ?PP c d \ P (cbox c d)" + have "P (\?A)" + proof (rule UN_cases) + let ?B = "(\s. cbox (\i\Basis. (if i \ s then a\i else (a\i + b\i) / 2) *\<^sub>R i::'a) + (\i\Basis. (if i \ s then (a\i + b\i) / 2 else b\i) *\<^sub>R i)) ` {s. s \ Basis}" + have "?A \ ?B" + proof + fix x + assume "x \ ?A" + then obtain c d + where x: "x = cbox c d" + "\i. i \ Basis \ + c \ i = a \ i \ d \ i = (a \ i + b \ i) / 2 \ + c \ i = (a \ i + b \ i) / 2 \ d \ i = b \ i" by blast + show "x \ ?B" + unfolding image_iff x + apply (rule_tac x="{i. i\Basis \ c\i = a\i}" in bexI) + apply (rule arg_cong2 [where f = cbox]) + using x(2) ab + apply (auto simp add: euclidean_eq_iff[where 'a='a]) + by fastforce + qed + then show "finite ?A" + by (rule finite_subset) auto + next + fix s + assume "s \ ?A" + then obtain c d + where s: "s = cbox c d" + "\i. i \ Basis \ + c \ i = a \ i \ d \ i = (a \ i + b \ i) / 2 \ + c \ i = (a \ i + b \ i) / 2 \ d \ i = b \ i" + by blast + show "P s" + unfolding s + apply (rule as[rule_format]) + using ab s(2) by force + show "\a b. s = cbox a b" + unfolding s by auto + fix t + assume "t \ ?A" + then obtain e f where t: + "t = cbox e f" + "\i. i \ Basis \ + e \ i = a \ i \ f \ i = (a \ i + b \ i) / 2 \ + e \ i = (a \ i + b \ i) / 2 \ f \ i = b \ i" + by blast + assume "s \ t" + then have "\ (c = e \ d = f)" + unfolding s t by auto + then obtain i where "c\i \ e\i \ d\i \ f\i" and i': "i \ Basis" + unfolding euclidean_eq_iff[where 'a='a] by auto + then have i: "c\i \ e\i" "d\i \ f\i" + using s(2) t(2) apply fastforce + using t(2)[OF i'] \c \ i \ e \ i \ d \ i \ f \ i\ i' s(2) t(2) by fastforce + have *: "\s t. (\a. a \ s \ a \ t \ False) \ s \ t = {}" + by auto + show "interior s \ interior t = {}" + unfolding s t interior_cbox + proof (rule *) + fix x + assume "x \ box c d" "x \ box e f" + then have x: "c\i < d\i" "e\i < f\i" "c\i < f\i" "e\i < d\i" + unfolding mem_box using i' + by force+ + show False using s(2)[OF i'] + proof safe + assume as: "c \ i = a \ i" "d \ i = (a \ i + b \ i) / 2" + show False + using t(2)[OF i'] and i x unfolding as by (fastforce simp add:field_simps) + next + assume as: "c \ i = (a \ i + b \ i) / 2" "d \ i = b \ i" + show False + using t(2)[OF i'] and i x unfolding as by(fastforce simp add:field_simps) + qed + qed + qed + also have "\?A = cbox a b" + proof (rule set_eqI,rule) + fix x + assume "x \ \?A" + then obtain c d where x: + "x \ cbox c d" + "\i. i \ Basis \ + c \ i = a \ i \ d \ i = (a \ i + b \ i) / 2 \ + c \ i = (a \ i + b \ i) / 2 \ d \ i = b \ i" + by blast + show "x\cbox a b" + unfolding mem_box + proof safe + fix i :: 'a + assume i: "i \ Basis" + then show "a \ i \ x \ i" "x \ i \ b \ i" + using x(2)[OF i] x(1)[unfolded mem_box,THEN bspec, OF i] by auto + qed + next + fix x + assume x: "x \ cbox a b" + have "\i\Basis. + \c d. (c = a\i \ d = (a\i + b\i) / 2 \ c = (a\i + b\i) / 2 \ d = b\i) \ c\x\i \ x\i \ d" + (is "\i\Basis. \c d. ?P i c d") + unfolding mem_box + proof + fix i :: 'a + assume i: "i \ Basis" + have "?P i (a\i) ((a \ i + b \ i) / 2) \ ?P i ((a \ i + b \ i) / 2) (b\i)" + using x[unfolded mem_box,THEN bspec, OF i] by auto + then show "\c d. ?P i c d" + by blast + qed + then show "x\\?A" + unfolding Union_iff Bex_def mem_Collect_eq choice_Basis_iff + apply auto + apply (rule_tac x="cbox xa xaa" in exI) + unfolding mem_box + apply auto + done + qed + finally show False + using assms by auto +qed + +lemma interval_bisection: + fixes type :: "'a::euclidean_space" + assumes "P {}" + and "(\s t. P s \ P t \ interior(s) \ interior(t) = {} \ P(s \ t))" + and "\ P (cbox a (b::'a))" + obtains x where "x \ cbox a b" + and "\e>0. \c d. x \ cbox c d \ cbox c d \ ball x e \ cbox c d \ cbox a b \ \ P (cbox c d)" +proof - + have "\x. \y. \ P (cbox (fst x) (snd x)) \ (\ P (cbox (fst y) (snd y)) \ + (\i\Basis. fst x\i \ fst y\i \ fst y\i \ snd y\i \ snd y\i \ snd x\i \ + 2 * (snd y\i - fst y\i) \ snd x\i - fst x\i))" (is "\x. ?P x") + proof + show "?P x" for x + proof (cases "P (cbox (fst x) (snd x))") + case True + then show ?thesis by auto + next + case as: False + obtain c d where "\ P (cbox c d)" + "\i\Basis. + fst x \ i \ c \ i \ + c \ i \ d \ i \ + d \ i \ snd x \ i \ + 2 * (d \ i - c \ i) \ snd x \ i - fst x \ i" + by (rule interval_bisection_step[of P, OF assms(1-2) as]) + then show ?thesis + apply - + apply (rule_tac x="(c,d)" in exI) + apply auto + done + qed + qed + then obtain f where f: + "\x. + \ P (cbox (fst x) (snd x)) \ + \ P (cbox (fst (f x)) (snd (f x))) \ + (\i\Basis. + fst x \ i \ fst (f x) \ i \ + fst (f x) \ i \ snd (f x) \ i \ + snd (f x) \ i \ snd x \ i \ + 2 * (snd (f x) \ i - fst (f x) \ i) \ snd x \ i - fst x \ i)" + apply - + apply (drule choice) + apply blast + done + define AB A B where ab_def: "AB n = (f ^^ n) (a,b)" "A n = fst(AB n)" "B n = snd(AB n)" for n + have "A 0 = a" "B 0 = b" "\n. \ P (cbox (A(Suc n)) (B(Suc n))) \ + (\i\Basis. A(n)\i \ A(Suc n)\i \ A(Suc n)\i \ B(Suc n)\i \ B(Suc n)\i \ B(n)\i \ + 2 * (B(Suc n)\i - A(Suc n)\i) \ B(n)\i - A(n)\i)" (is "\n. ?P n") + proof - + show "A 0 = a" "B 0 = b" + unfolding ab_def by auto + note S = ab_def funpow.simps o_def id_apply + show "?P n" for n + proof (induct n) + case 0 + then show ?case + unfolding S + apply (rule f[rule_format]) using assms(3) + apply auto + done + next + case (Suc n) + show ?case + unfolding S + apply (rule f[rule_format]) + using Suc + unfolding S + apply auto + done + qed + qed + note AB = this(1-2) conjunctD2[OF this(3),rule_format] + + have interv: "\n. \x\cbox (A n) (B n). \y\cbox (A n) (B n). dist x y < e" + if e: "0 < e" for e + proof - + obtain n where n: "(\i\Basis. b \ i - a \ i) / e < 2 ^ n" + using real_arch_pow[of 2 "(setsum (\i. b\i - a\i) Basis) / e"] by auto + show ?thesis + proof (rule exI [where x=n], clarify) + fix x y + assume xy: "x\cbox (A n) (B n)" "y\cbox (A n) (B n)" + have "dist x y \ setsum (\i. \(x - y)\i\) Basis" + unfolding dist_norm by(rule norm_le_l1) + also have "\ \ setsum (\i. B n\i - A n\i) Basis" + proof (rule setsum_mono) + fix i :: 'a + assume i: "i \ Basis" + show "\(x - y) \ i\ \ B n \ i - A n \ i" + using xy[unfolded mem_box,THEN bspec, OF i] + by (auto simp: inner_diff_left) + qed + also have "\ \ setsum (\i. b\i - a\i) Basis / 2^n" + unfolding setsum_divide_distrib + proof (rule setsum_mono) + show "B n \ i - A n \ i \ (b \ i - a \ i) / 2 ^ n" if i: "i \ Basis" for i + proof (induct n) + case 0 + then show ?case + unfolding AB by auto + next + case (Suc n) + have "B (Suc n) \ i - A (Suc n) \ i \ (B n \ i - A n \ i) / 2" + using AB(4)[of i n] using i by auto + also have "\ \ (b \ i - a \ i) / 2 ^ Suc n" + using Suc by (auto simp add: field_simps) + finally show ?case . + qed + qed + also have "\ < e" + using n using e by (auto simp add: field_simps) + finally show "dist x y < e" . + qed + qed + { + fix n m :: nat + assume "m \ n" then have "cbox (A n) (B n) \ cbox (A m) (B m)" + proof (induction rule: inc_induct) + case (step i) + show ?case + using AB(4) by (intro order_trans[OF step.IH] subset_box_imp) auto + qed simp + } note ABsubset = this + have "\a. \n. a\ cbox (A n) (B n)" + by (rule decreasing_closed_nest[rule_format,OF closed_cbox _ ABsubset interv]) + (metis nat.exhaust AB(1-3) assms(1,3)) + then obtain x0 where x0: "\n. x0 \ cbox (A n) (B n)" + by blast + show thesis + proof (rule that[rule_format, of x0]) + show "x0\cbox a b" + using x0[of 0] unfolding AB . + fix e :: real + assume "e > 0" + from interv[OF this] obtain n + where n: "\x\cbox (A n) (B n). \y\cbox (A n) (B n). dist x y < e" .. + have "\ P (cbox (A n) (B n))" + apply (cases "0 < n") + using AB(3)[of "n - 1"] assms(3) AB(1-2) + apply auto + done + moreover have "cbox (A n) (B n) \ ball x0 e" + using n using x0[of n] by auto + moreover have "cbox (A n) (B n) \ cbox a b" + unfolding AB(1-2)[symmetric] by (rule ABsubset) auto + ultimately show "\c d. x0 \ cbox c d \ cbox c d \ ball x0 e \ cbox c d \ cbox a b \ \ P (cbox c d)" + apply (rule_tac x="A n" in exI) + apply (rule_tac x="B n" in exI) + apply (auto simp: x0) + done + qed +qed + + +subsection \Cousin's lemma.\ + +lemma fine_division_exists: + fixes a b :: "'a::euclidean_space" + assumes "gauge g" + obtains p where "p tagged_division_of (cbox a b)" "g fine p" +proof - + presume "\ (\p. p tagged_division_of (cbox a b) \ g fine p) \ False" + then obtain p where "p tagged_division_of (cbox a b)" "g fine p" + by blast + then show thesis .. +next + assume as: "\ (\p. p tagged_division_of (cbox a b) \ g fine p)" + obtain x where x: + "x \ (cbox a b)" + "\e. 0 < e \ + \c d. + x \ cbox c d \ + cbox c d \ ball x e \ + cbox c d \ (cbox a b) \ + \ (\p. p tagged_division_of cbox c d \ g fine p)" + apply (rule interval_bisection[of "\s. \p. p tagged_division_of s \ g fine p", OF _ _ as]) + apply (simp add: fine_def) + apply (metis tagged_division_union fine_union) + apply (auto simp: ) + done + obtain e where e: "e > 0" "ball x e \ g x" + using gaugeD[OF assms, of x] unfolding open_contains_ball by auto + from x(2)[OF e(1)] + obtain c d where c_d: "x \ cbox c d" + "cbox c d \ ball x e" + "cbox c d \ cbox a b" + "\ (\p. p tagged_division_of cbox c d \ g fine p)" + by blast + have "g fine {(x, cbox c d)}" + unfolding fine_def using e using c_d(2) by auto + then show False + using tagged_division_of_self[OF c_d(1)] using c_d by auto +qed + +lemma fine_division_exists_real: + fixes a b :: real + assumes "gauge g" + obtains p where "p tagged_division_of {a .. b}" "g fine p" + by (metis assms box_real(2) fine_division_exists) + +subsection \Basic theorems about integrals.\ + +lemma has_integral_unique: + fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" + assumes "(f has_integral k1) i" + and "(f has_integral k2) i" + shows "k1 = k2" +proof (rule ccontr) + let ?e = "norm (k1 - k2) / 2" + assume as: "k1 \ k2" + then have e: "?e > 0" + by auto + have lem: False + if f_k1: "(f has_integral k1) (cbox a b)" + and f_k2: "(f has_integral k2) (cbox a b)" + and "k1 \ k2" + for f :: "'n \ 'a" and a b k1 k2 + proof - + let ?e = "norm (k1 - k2) / 2" + from \k1 \ k2\ have e: "?e > 0" by auto + obtain d1 where d1: + "gauge d1" + "\p. p tagged_division_of cbox a b \ + d1 fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x) - k1) < norm (k1 - k2) / 2" + by (rule has_integralD[OF f_k1 e]) blast + obtain d2 where d2: + "gauge d2" + "\p. p tagged_division_of cbox a b \ + d2 fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x) - k2) < norm (k1 - k2) / 2" + by (rule has_integralD[OF f_k2 e]) blast + obtain p where p: + "p tagged_division_of cbox a b" + "(\x. d1 x \ d2 x) fine p" + by (rule fine_division_exists[OF gauge_inter[OF d1(1) d2(1)]]) + let ?c = "(\(x, k)\p. content k *\<^sub>R f x)" + have "norm (k1 - k2) \ norm (?c - k2) + norm (?c - k1)" + using norm_triangle_ineq4[of "k1 - ?c" "k2 - ?c"] + by (auto simp add:algebra_simps norm_minus_commute) + also have "\ < norm (k1 - k2) / 2 + norm (k1 - k2) / 2" + apply (rule add_strict_mono) + apply (rule_tac[!] d2(2) d1(2)) + using p unfolding fine_def + apply auto + done + finally show False by auto + qed + { + presume "\ (\a b. i = cbox a b) \ False" + then show False + using as assms lem by blast + } + assume as: "\ (\a b. i = cbox a b)" + obtain B1 where B1: + "0 < B1" + "\a b. ball 0 B1 \ cbox a b \ + \z. ((\x. if x \ i then f x else 0) has_integral z) (cbox a b) \ + norm (z - k1) < norm (k1 - k2) / 2" + by (rule has_integral_altD[OF assms(1) as,OF e]) blast + obtain B2 where B2: + "0 < B2" + "\a b. ball 0 B2 \ cbox a b \ + \z. ((\x. if x \ i then f x else 0) has_integral z) (cbox a b) \ + norm (z - k2) < norm (k1 - k2) / 2" + by (rule has_integral_altD[OF assms(2) as,OF e]) blast + have "\a b::'n. ball 0 B1 \ ball 0 B2 \ cbox a b" + apply (rule bounded_subset_cbox) + using bounded_Un bounded_ball + apply auto + done + then obtain a b :: 'n where ab: "ball 0 B1 \ cbox a b" "ball 0 B2 \ cbox a b" + by blast + obtain w where w: + "((\x. if x \ i then f x else 0) has_integral w) (cbox a b)" + "norm (w - k1) < norm (k1 - k2) / 2" + using B1(2)[OF ab(1)] by blast + obtain z where z: + "((\x. if x \ i then f x else 0) has_integral z) (cbox a b)" + "norm (z - k2) < norm (k1 - k2) / 2" + using B2(2)[OF ab(2)] by blast + have "z = w" + using lem[OF w(1) z(1)] by auto + then have "norm (k1 - k2) \ norm (z - k2) + norm (w - k1)" + using norm_triangle_ineq4 [of "k1 - w" "k2 - z"] + by (auto simp add: norm_minus_commute) + also have "\ < norm (k1 - k2) / 2 + norm (k1 - k2) / 2" + apply (rule add_strict_mono) + apply (rule_tac[!] z(2) w(2)) + done + finally show False by auto +qed + +lemma integral_unique [intro]: "(f has_integral y) k \ integral k f = y" + unfolding integral_def + by (rule some_equality) (auto intro: has_integral_unique) + +lemma eq_integralD: "integral k f = y \ (f has_integral y) k \ ~ f integrable_on k \ y=0" + unfolding integral_def integrable_on_def + apply (erule subst) + apply (rule someI_ex) + by blast + +lemma has_integral_is_0: + fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" + assumes "\x\s. f x = 0" + shows "(f has_integral 0) s" +proof - + have lem: "\a b. \f::'n \ 'a. + (\x\cbox a b. f(x) = 0) \ (f has_integral 0) (cbox a b)" + unfolding has_integral + proof clarify + fix a b e + fix f :: "'n \ 'a" + assume as: "\x\cbox a b. f x = 0" "0 < (e::real)" + have "norm ((\(x, k)\p. content k *\<^sub>R f x) - 0) < e" + if p: "p tagged_division_of cbox a b" for p + proof - + have "(\(x, k)\p. content k *\<^sub>R f x) = 0" + proof (rule setsum.neutral, rule) + fix x + assume x: "x \ p" + have "f (fst x) = 0" + using tagged_division_ofD(2-3)[OF p, of "fst x" "snd x"] using as x by auto + then show "(\(x, k). content k *\<^sub>R f x) x = 0" + apply (subst surjective_pairing[of x]) + unfolding split_conv + apply auto + done + qed + then show ?thesis + using as by auto + qed + then show "\d. gauge d \ + (\p. p tagged_division_of (cbox a b) \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x) - 0) < e)" + by auto + qed + { + presume "\ (\a b. s = cbox a b) \ ?thesis" + with assms lem show ?thesis + by blast + } + have *: "(\x. if x \ s then f x else 0) = (\x. 0)" + apply (rule ext) + using assms + apply auto + done + assume "\ (\a b. s = cbox a b)" + then show ?thesis + using lem + by (subst has_integral_alt) (force simp add: *) +qed + +lemma has_integral_0[simp]: "((\x::'n::euclidean_space. 0) has_integral 0) s" + by (rule has_integral_is_0) auto + +lemma has_integral_0_eq[simp]: "((\x. 0) has_integral i) s \ i = 0" + using has_integral_unique[OF has_integral_0] by auto + +lemma has_integral_linear: + fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" + assumes "(f has_integral y) s" + and "bounded_linear h" + shows "((h \ f) has_integral ((h y))) s" +proof - + interpret bounded_linear h + using assms(2) . + from pos_bounded obtain B where B: "0 < B" "\x. norm (h x) \ norm x * B" + by blast + have lem: "\(f :: 'n \ 'a) y a b. + (f has_integral y) (cbox a b) \ ((h \ f) has_integral h y) (cbox a b)" + unfolding has_integral + proof (clarify, goal_cases) + case prems: (1 f y a b e) + from pos_bounded + obtain B where B: "0 < B" "\x. norm (h x) \ norm x * B" + by blast + have "e / B > 0" using prems(2) B by simp + then obtain g + where g: "gauge g" + "\p. p tagged_division_of (cbox a b) \ g fine p \ + norm ((\(x, k)\p. content k *\<^sub>R f x) - y) < e / B" + using prems(1) by auto + { + fix p + assume as: "p tagged_division_of (cbox a b)" "g fine p" + have hc: "\x k. h ((\(x, k). content k *\<^sub>R f x) x) = (\(x, k). h (content k *\<^sub>R f x)) x" + by auto + then have "(\(x, k)\p. content k *\<^sub>R (h \ f) x) = setsum (h \ (\(x, k). content k *\<^sub>R f x)) p" + unfolding o_def unfolding scaleR[symmetric] hc by simp + also have "\ = h (\(x, k)\p. content k *\<^sub>R f x)" + using setsum[of "\(x,k). content k *\<^sub>R f x" p] using as by auto + finally have "(\(x, k)\p. content k *\<^sub>R (h \ f) x) = h (\(x, k)\p. content k *\<^sub>R f x)" . + then have "norm ((\(x, k)\p. content k *\<^sub>R (h \ f) x) - h y) < e" + apply (simp add: diff[symmetric]) + apply (rule le_less_trans[OF B(2)]) + using g(2)[OF as] B(1) + apply (auto simp add: field_simps) + done + } + with g show ?case + by (rule_tac x=g in exI) auto + qed + { + presume "\ (\a b. s = cbox a b) \ ?thesis" + then show ?thesis + using assms(1) lem by blast + } + assume as: "\ (\a b. s = cbox a b)" + then show ?thesis + proof (subst has_integral_alt, clarsimp) + fix e :: real + assume e: "e > 0" + have *: "0 < e/B" using e B(1) by simp + obtain M where M: + "M > 0" + "\a b. ball 0 M \ cbox a b \ + \z. ((\x. if x \ s then f x else 0) has_integral z) (cbox a b) \ norm (z - y) < e / B" + using has_integral_altD[OF assms(1) as *] by blast + show "\B>0. \a b. ball 0 B \ cbox a b \ + (\z. ((\x. if x \ s then (h \ f) x else 0) has_integral z) (cbox a b) \ norm (z - h y) < e)" + proof (rule_tac x=M in exI, clarsimp simp add: M, goal_cases) + case prems: (1 a b) + obtain z where z: + "((\x. if x \ s then f x else 0) has_integral z) (cbox a b)" + "norm (z - y) < e / B" + using M(2)[OF prems(1)] by blast + have *: "(\x. if x \ s then (h \ f) x else 0) = h \ (\x. if x \ s then f x else 0)" + using zero by auto + show ?case + apply (rule_tac x="h z" in exI) + apply (simp add: * lem z(1)) + apply (metis B diff le_less_trans pos_less_divide_eq z(2)) + done + qed + qed +qed + +lemma has_integral_scaleR_left: + "(f has_integral y) s \ ((\x. f x *\<^sub>R c) has_integral (y *\<^sub>R c)) s" + using has_integral_linear[OF _ bounded_linear_scaleR_left] by (simp add: comp_def) + +lemma has_integral_mult_left: + fixes c :: "_ :: real_normed_algebra" + shows "(f has_integral y) s \ ((\x. f x * c) has_integral (y * c)) s" + using has_integral_linear[OF _ bounded_linear_mult_left] by (simp add: comp_def) + +text\The case analysis eliminates the condition @{term "f integrable_on s"} at the cost + of the type class constraint \division_ring\\ +corollary integral_mult_left [simp]: + fixes c:: "'a::{real_normed_algebra,division_ring}" + shows "integral s (\x. f x * c) = integral s f * c" +proof (cases "f integrable_on s \ c = 0") + case True then show ?thesis + by (force intro: has_integral_mult_left) +next + case False then have "~ (\x. f x * c) integrable_on s" + using has_integral_mult_left [of "(\x. f x * c)" _ s "inverse c"] + by (force simp add: mult.assoc) + with False show ?thesis by (simp add: not_integrable_integral) +qed + +corollary integral_mult_right [simp]: + fixes c:: "'a::{real_normed_field}" + shows "integral s (\x. c * f x) = c * integral s f" +by (simp add: mult.commute [of c]) + +corollary integral_divide [simp]: + fixes z :: "'a::real_normed_field" + shows "integral S (\x. f x / z) = integral S (\x. f x) / z" +using integral_mult_left [of S f "inverse z"] + by (simp add: divide_inverse_commute) + +lemma has_integral_mult_right: + fixes c :: "'a :: real_normed_algebra" + shows "(f has_integral y) i \ ((\x. c * f x) has_integral (c * y)) i" + using has_integral_linear[OF _ bounded_linear_mult_right] by (simp add: comp_def) + +lemma has_integral_cmul: "(f has_integral k) s \ ((\x. c *\<^sub>R f x) has_integral (c *\<^sub>R k)) s" + unfolding o_def[symmetric] + by (metis has_integral_linear bounded_linear_scaleR_right) + +lemma has_integral_cmult_real: + fixes c :: real + assumes "c \ 0 \ (f has_integral x) A" + shows "((\x. c * f x) has_integral c * x) A" +proof (cases "c = 0") + case True + then show ?thesis by simp +next + case False + from has_integral_cmul[OF assms[OF this], of c] show ?thesis + unfolding real_scaleR_def . +qed + +lemma has_integral_neg: "(f has_integral k) s \ ((\x. -(f x)) has_integral -k) s" + by (drule_tac c="-1" in has_integral_cmul) auto + +lemma has_integral_add: + fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" + assumes "(f has_integral k) s" + and "(g has_integral l) s" + shows "((\x. f x + g x) has_integral (k + l)) s" +proof - + have lem: "((\x. f x + g x) has_integral (k + l)) (cbox a b)" + if f_k: "(f has_integral k) (cbox a b)" + and g_l: "(g has_integral l) (cbox a b)" + for f :: "'n \ 'a" and g a b k l + unfolding has_integral + proof clarify + fix e :: real + assume e: "e > 0" + then have *: "e / 2 > 0" + by auto + obtain d1 where d1: + "gauge d1" + "\p. p tagged_division_of (cbox a b) \ d1 fine p \ + norm ((\(x, k)\p. content k *\<^sub>R f x) - k) < e / 2" + using has_integralD[OF f_k *] by blast + obtain d2 where d2: + "gauge d2" + "\p. p tagged_division_of (cbox a b) \ d2 fine p \ + norm ((\(x, k)\p. content k *\<^sub>R g x) - l) < e / 2" + using has_integralD[OF g_l *] by blast + show "\d. gauge d \ (\p. p tagged_division_of (cbox a b) \ d fine p \ + norm ((\(x, k)\p. content k *\<^sub>R (f x + g x)) - (k + l)) < e)" + proof (rule exI [where x="\x. (d1 x) \ (d2 x)"], clarsimp simp add: gauge_inter[OF d1(1) d2(1)]) + fix p + assume as: "p tagged_division_of (cbox a b)" "(\x. d1 x \ d2 x) fine p" + have *: "(\(x, k)\p. content k *\<^sub>R (f x + g x)) = + (\(x, k)\p. content k *\<^sub>R f x) + (\(x, k)\p. content k *\<^sub>R g x)" + unfolding scaleR_right_distrib setsum.distrib[of "\(x,k). content k *\<^sub>R f x" "\(x,k). content k *\<^sub>R g x" p,symmetric] + by (rule setsum.cong) auto + from as have fine: "d1 fine p" "d2 fine p" + unfolding fine_inter by auto + have "norm ((\(x, k)\p. content k *\<^sub>R (f x + g x)) - (k + l)) = + norm (((\(x, k)\p. content k *\<^sub>R f x) - k) + ((\(x, k)\p. content k *\<^sub>R g x) - l))" + unfolding * by (auto simp add: algebra_simps) + also have "\ < e/2 + e/2" + apply (rule le_less_trans[OF norm_triangle_ineq]) + using as d1 d2 fine + apply (blast intro: add_strict_mono) + done + finally show "norm ((\(x, k)\p. content k *\<^sub>R (f x + g x)) - (k + l)) < e" + by auto + qed + qed + { + presume "\ (\a b. s = cbox a b) \ ?thesis" + then show ?thesis + using assms lem by force + } + assume as: "\ (\a b. s = cbox a b)" + then show ?thesis + proof (subst has_integral_alt, clarsimp, goal_cases) + case (1 e) + then have *: "e / 2 > 0" + by auto + from has_integral_altD[OF assms(1) as *] + obtain B1 where B1: + "0 < B1" + "\a b. ball 0 B1 \ cbox a b \ + \z. ((\x. if x \ s then f x else 0) has_integral z) (cbox a b) \ norm (z - k) < e / 2" + by blast + from has_integral_altD[OF assms(2) as *] + obtain B2 where B2: + "0 < B2" + "\a b. ball 0 B2 \ (cbox a b) \ + \z. ((\x. if x \ s then g x else 0) has_integral z) (cbox a b) \ norm (z - l) < e / 2" + by blast + show ?case + proof (rule_tac x="max B1 B2" in exI, clarsimp simp add: max.strict_coboundedI1 B1) + fix a b + assume "ball 0 (max B1 B2) \ cbox a (b::'n)" + then have *: "ball 0 B1 \ cbox a (b::'n)" "ball 0 B2 \ cbox a (b::'n)" + by auto + obtain w where w: + "((\x. if x \ s then f x else 0) has_integral w) (cbox a b)" + "norm (w - k) < e / 2" + using B1(2)[OF *(1)] by blast + obtain z where z: + "((\x. if x \ s then g x else 0) has_integral z) (cbox a b)" + "norm (z - l) < e / 2" + using B2(2)[OF *(2)] by blast + have *: "\x. (if x \ s then f x + g x else 0) = + (if x \ s then f x else 0) + (if x \ s then g x else 0)" + by auto + show "\z. ((\x. if x \ s then f x + g x else 0) has_integral z) (cbox a b) \ norm (z - (k + l)) < e" + apply (rule_tac x="w + z" in exI) + apply (simp add: lem[OF w(1) z(1), unfolded *[symmetric]]) + using norm_triangle_ineq[of "w - k" "z - l"] w(2) z(2) + apply (auto simp add: field_simps) + done + qed + qed +qed + +lemma has_integral_sub: + "(f has_integral k) s \ (g has_integral l) s \ + ((\x. f x - g x) has_integral (k - l)) s" + using has_integral_add[OF _ has_integral_neg, of f k s g l] + by (auto simp: algebra_simps) + +lemma integral_0 [simp]: + "integral s (\x::'n::euclidean_space. 0::'m::real_normed_vector) = 0" + by (rule integral_unique has_integral_0)+ + +lemma integral_add: "f integrable_on s \ g integrable_on s \ + integral s (\x. f x + g x) = integral s f + integral s g" + by (rule integral_unique) (metis integrable_integral has_integral_add) + +lemma integral_cmul [simp]: "integral s (\x. c *\<^sub>R f x) = c *\<^sub>R integral s f" +proof (cases "f integrable_on s \ c = 0") + case True with has_integral_cmul show ?thesis by force +next + case False then have "~ (\x. c *\<^sub>R f x) integrable_on s" + using has_integral_cmul [of "(\x. c *\<^sub>R f x)" _ s "inverse c"] + by force + with False show ?thesis by (simp add: not_integrable_integral) +qed + +lemma integral_neg [simp]: "integral s (\x. - f x) = - integral s f" +proof (cases "f integrable_on s") + case True then show ?thesis + by (simp add: has_integral_neg integrable_integral integral_unique) +next + case False then have "~ (\x. - f x) integrable_on s" + using has_integral_neg [of "(\x. - f x)" _ s ] + by force + with False show ?thesis by (simp add: not_integrable_integral) +qed + +lemma integral_diff: "f integrable_on s \ g integrable_on s \ + integral s (\x. f x - g x) = integral s f - integral s g" + by (rule integral_unique) (metis integrable_integral has_integral_sub) + +lemma integrable_0: "(\x. 0) integrable_on s" + unfolding integrable_on_def using has_integral_0 by auto + +lemma integrable_add: "f integrable_on s \ g integrable_on s \ (\x. f x + g x) integrable_on s" + unfolding integrable_on_def by(auto intro: has_integral_add) + +lemma integrable_cmul: "f integrable_on s \ (\x. c *\<^sub>R f(x)) integrable_on s" + unfolding integrable_on_def by(auto intro: has_integral_cmul) + +lemma integrable_on_cmult_iff: + fixes c :: real + assumes "c \ 0" + shows "(\x. c * f x) integrable_on s \ f integrable_on s" + using integrable_cmul[of "\x. c * f x" s "1 / c"] integrable_cmul[of f s c] \c \ 0\ + by auto + +lemma integrable_on_cmult_left: + assumes "f integrable_on s" + shows "(\x. of_real c * f x) integrable_on s" + using integrable_cmul[of f s "of_real c"] assms + by (simp add: scaleR_conv_of_real) + +lemma integrable_neg: "f integrable_on s \ (\x. -f(x)) integrable_on s" + unfolding integrable_on_def by(auto intro: has_integral_neg) + +lemma integrable_diff: + "f integrable_on s \ g integrable_on s \ (\x. f x - g x) integrable_on s" + unfolding integrable_on_def by(auto intro: has_integral_sub) + +lemma integrable_linear: + "f integrable_on s \ bounded_linear h \ (h \ f) integrable_on s" + unfolding integrable_on_def by(auto intro: has_integral_linear) + +lemma integral_linear: + "f integrable_on s \ bounded_linear h \ integral s (h \ f) = h (integral s f)" + apply (rule has_integral_unique [where i=s and f = "h \ f"]) + apply (simp_all add: integrable_integral integrable_linear has_integral_linear ) + done + +lemma integral_component_eq[simp]: + fixes f :: "'n::euclidean_space \ 'm::euclidean_space" + assumes "f integrable_on s" + shows "integral s (\x. f x \ k) = integral s f \ k" + unfolding integral_linear[OF assms(1) bounded_linear_component,unfolded o_def] .. + +lemma has_integral_setsum: + assumes "finite t" + and "\a\t. ((f a) has_integral (i a)) s" + shows "((\x. setsum (\a. f a x) t) has_integral (setsum i t)) s" + using assms(1) subset_refl[of t] +proof (induct rule: finite_subset_induct) + case empty + then show ?case by auto +next + case (insert x F) + with assms show ?case + by (simp add: has_integral_add) +qed + +lemma integral_setsum: + "\finite t; \a\t. (f a) integrable_on s\ \ + integral s (\x. setsum (\a. f a x) t) = setsum (\a. integral s (f a)) t" + by (auto intro: has_integral_setsum integrable_integral) + +lemma integrable_setsum: + "\finite t; \a\t. (f a) integrable_on s\ \ (\x. setsum (\a. f a x) t) integrable_on s" + unfolding integrable_on_def + apply (drule bchoice) + using has_integral_setsum[of t] + apply auto + done + +lemma has_integral_eq: + assumes "\x. x \ s \ f x = g x" + and "(f has_integral k) s" + shows "(g has_integral k) s" + using has_integral_sub[OF assms(2), of "\x. f x - g x" 0] + using has_integral_is_0[of s "\x. f x - g x"] + using assms(1) + by auto + +lemma integrable_eq: "(\x. x \ s \ f x = g x) \ f integrable_on s \ g integrable_on s" + unfolding integrable_on_def + using has_integral_eq[of s f g] has_integral_eq by blast + +lemma has_integral_cong: + assumes "\x. x \ s \ f x = g x" + shows "(f has_integral i) s = (g has_integral i) s" + using has_integral_eq[of s f g] has_integral_eq[of s g f] assms + by auto + +lemma integral_cong: + assumes "\x. x \ s \ f x = g x" + shows "integral s f = integral s g" + unfolding integral_def +by (metis (full_types, hide_lams) assms has_integral_cong integrable_eq) + +lemma integrable_on_cmult_left_iff [simp]: + assumes "c \ 0" + shows "(\x. of_real c * f x) integrable_on s \ f integrable_on s" + (is "?lhs = ?rhs") +proof + assume ?lhs + then have "(\x. of_real (1 / c) * (of_real c * f x)) integrable_on s" + using integrable_cmul[of "\x. of_real c * f x" s "1 / of_real c"] + by (simp add: scaleR_conv_of_real) + then have "(\x. (of_real (1 / c) * of_real c * f x)) integrable_on s" + by (simp add: algebra_simps) + with \c \ 0\ show ?rhs + by (metis (no_types, lifting) integrable_eq mult.left_neutral nonzero_divide_eq_eq of_real_1 of_real_mult) +qed (blast intro: integrable_on_cmult_left) + +lemma integrable_on_cmult_right: + fixes f :: "_ \ 'b :: {comm_ring,real_algebra_1,real_normed_vector}" + assumes "f integrable_on s" + shows "(\x. f x * of_real c) integrable_on s" +using integrable_on_cmult_left [OF assms] by (simp add: mult.commute) + +lemma integrable_on_cmult_right_iff [simp]: + fixes f :: "_ \ 'b :: {comm_ring,real_algebra_1,real_normed_vector}" + assumes "c \ 0" + shows "(\x. f x * of_real c) integrable_on s \ f integrable_on s" +using integrable_on_cmult_left_iff [OF assms] by (simp add: mult.commute) + +lemma integrable_on_cdivide: + fixes f :: "_ \ 'b :: real_normed_field" + assumes "f integrable_on s" + shows "(\x. f x / of_real c) integrable_on s" +by (simp add: integrable_on_cmult_right divide_inverse assms of_real_inverse [symmetric] del: of_real_inverse) + +lemma integrable_on_cdivide_iff [simp]: + fixes f :: "_ \ 'b :: real_normed_field" + assumes "c \ 0" + shows "(\x. f x / of_real c) integrable_on s \ f integrable_on s" +by (simp add: divide_inverse assms of_real_inverse [symmetric] del: of_real_inverse) + +lemma has_integral_null [intro]: + assumes "content(cbox a b) = 0" + shows "(f has_integral 0) (cbox a b)" +proof - + have "gauge (\x. ball x 1)" + by auto + moreover + { + fix e :: real + fix p + assume e: "e > 0" + assume p: "p tagged_division_of (cbox a b)" + have "norm ((\(x, k)\p. content k *\<^sub>R f x) - 0) = 0" + unfolding norm_eq_zero diff_0_right + using setsum_content_null[OF assms(1) p, of f] . + then have "norm ((\(x, k)\p. content k *\<^sub>R f x) - 0) < e" + using e by auto + } + ultimately show ?thesis + by (auto simp: has_integral) +qed + +lemma has_integral_null_real [intro]: + assumes "content {a .. b::real} = 0" + shows "(f has_integral 0) {a .. b}" + by (metis assms box_real(2) has_integral_null) + +lemma has_integral_null_eq[simp]: "content (cbox a b) = 0 \ (f has_integral i) (cbox a b) \ i = 0" + by (auto simp add: has_integral_null dest!: integral_unique) + +lemma integral_null [simp]: "content (cbox a b) = 0 \ integral (cbox a b) f = 0" + by (metis has_integral_null integral_unique) + +lemma integrable_on_null [intro]: "content (cbox a b) = 0 \ f integrable_on (cbox a b)" + by (simp add: has_integral_integrable) + +lemma has_integral_empty[intro]: "(f has_integral 0) {}" + by (simp add: has_integral_is_0) + +lemma has_integral_empty_eq[simp]: "(f has_integral i) {} \ i = 0" + by (auto simp add: has_integral_empty has_integral_unique) + +lemma integrable_on_empty[intro]: "f integrable_on {}" + unfolding integrable_on_def by auto + +lemma integral_empty[simp]: "integral {} f = 0" + by (rule integral_unique) (rule has_integral_empty) + +lemma has_integral_refl[intro]: + fixes a :: "'a::euclidean_space" + shows "(f has_integral 0) (cbox a a)" + and "(f has_integral 0) {a}" +proof - + have *: "{a} = cbox a a" + apply (rule set_eqI) + unfolding mem_box singleton_iff euclidean_eq_iff[where 'a='a] + apply safe + prefer 3 + apply (erule_tac x=b in ballE) + apply (auto simp add: field_simps) + done + show "(f has_integral 0) (cbox a a)" "(f has_integral 0) {a}" + unfolding * + apply (rule_tac[!] has_integral_null) + unfolding content_eq_0_interior + unfolding interior_cbox + using box_sing + apply auto + done +qed + +lemma integrable_on_refl[intro]: "f integrable_on cbox a a" + unfolding integrable_on_def by auto + +lemma integral_refl [simp]: "integral (cbox a a) f = 0" + by (rule integral_unique) auto + +lemma integral_singleton [simp]: "integral {a} f = 0" + by auto + +lemma integral_blinfun_apply: + assumes "f integrable_on s" + shows "integral s (\x. blinfun_apply h (f x)) = blinfun_apply h (integral s f)" + by (subst integral_linear[symmetric, OF assms blinfun.bounded_linear_right]) (simp add: o_def) + +lemma blinfun_apply_integral: + assumes "f integrable_on s" + shows "blinfun_apply (integral s f) x = integral s (\y. blinfun_apply (f y) x)" + by (metis (no_types, lifting) assms blinfun.prod_left.rep_eq integral_blinfun_apply integral_cong) + +lemma has_integral_componentwise_iff: + fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" + shows "(f has_integral y) A \ (\b\Basis. ((\x. f x \ b) has_integral (y \ b)) A)" +proof safe + fix b :: 'b assume "(f has_integral y) A" + from has_integral_linear[OF this(1) bounded_linear_component, of b] + show "((\x. f x \ b) has_integral (y \ b)) A" by (simp add: o_def) +next + assume "(\b\Basis. ((\x. f x \ b) has_integral (y \ b)) A)" + hence "\b\Basis. (((\x. x *\<^sub>R b) \ (\x. f x \ b)) has_integral ((y \ b) *\<^sub>R b)) A" + by (intro ballI has_integral_linear) (simp_all add: bounded_linear_scaleR_left) + hence "((\x. \b\Basis. (f x \ b) *\<^sub>R b) has_integral (\b\Basis. (y \ b) *\<^sub>R b)) A" + by (intro has_integral_setsum) (simp_all add: o_def) + thus "(f has_integral y) A" by (simp add: euclidean_representation) +qed + +lemma has_integral_componentwise: + fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" + shows "(\b. b \ Basis \ ((\x. f x \ b) has_integral (y \ b)) A) \ (f has_integral y) A" + by (subst has_integral_componentwise_iff) blast + +lemma integrable_componentwise_iff: + fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" + shows "f integrable_on A \ (\b\Basis. (\x. f x \ b) integrable_on A)" +proof + assume "f integrable_on A" + then obtain y where "(f has_integral y) A" by (auto simp: integrable_on_def) + hence "(\b\Basis. ((\x. f x \ b) has_integral (y \ b)) A)" + by (subst (asm) has_integral_componentwise_iff) + thus "(\b\Basis. (\x. f x \ b) integrable_on A)" by (auto simp: integrable_on_def) +next + assume "(\b\Basis. (\x. f x \ b) integrable_on A)" + then obtain y where "\b\Basis. ((\x. f x \ b) has_integral y b) A" + unfolding integrable_on_def by (subst (asm) bchoice_iff) blast + hence "\b\Basis. (((\x. x *\<^sub>R b) \ (\x. f x \ b)) has_integral (y b *\<^sub>R b)) A" + by (intro ballI has_integral_linear) (simp_all add: bounded_linear_scaleR_left) + hence "((\x. \b\Basis. (f x \ b) *\<^sub>R b) has_integral (\b\Basis. y b *\<^sub>R b)) A" + by (intro has_integral_setsum) (simp_all add: o_def) + thus "f integrable_on A" by (auto simp: integrable_on_def o_def euclidean_representation) +qed + +lemma integrable_componentwise: + fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" + shows "(\b. b \ Basis \ (\x. f x \ b) integrable_on A) \ f integrable_on A" + by (subst integrable_componentwise_iff) blast + +lemma integral_componentwise: + fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" + assumes "f integrable_on A" + shows "integral A f = (\b\Basis. integral A (\x. (f x \ b) *\<^sub>R b))" +proof - + from assms have integrable: "\b\Basis. (\x. x *\<^sub>R b) \ (\x. (f x \ b)) integrable_on A" + by (subst (asm) integrable_componentwise_iff, intro integrable_linear ballI) + (simp_all add: bounded_linear_scaleR_left) + have "integral A f = integral A (\x. \b\Basis. (f x \ b) *\<^sub>R b)" + by (simp add: euclidean_representation) + also from integrable have "\ = (\a\Basis. integral A (\x. (f x \ a) *\<^sub>R a))" + by (subst integral_setsum) (simp_all add: o_def) + finally show ?thesis . +qed + +lemma integrable_component: + "f integrable_on A \ (\x. f x \ (y :: 'b :: euclidean_space)) integrable_on A" + by (drule integrable_linear[OF _ bounded_linear_component[of y]]) (simp add: o_def) + + + +subsection \Cauchy-type criterion for integrability.\ + +(* XXXXXXX *) +lemma integrable_cauchy: + fixes f :: "'n::euclidean_space \ 'a::{real_normed_vector,complete_space}" + shows "f integrable_on cbox a b \ + (\e>0.\d. gauge d \ + (\p1 p2. p1 tagged_division_of (cbox a b) \ d fine p1 \ + p2 tagged_division_of (cbox a b) \ d fine p2 \ + norm (setsum (\(x,k). content k *\<^sub>R f x) p1 - + setsum (\(x,k). content k *\<^sub>R f x) p2) < e))" + (is "?l = (\e>0. \d. ?P e d)") +proof + assume ?l + then guess y unfolding integrable_on_def has_integral .. note y=this + show "\e>0. \d. ?P e d" + proof (clarify, goal_cases) + case (1 e) + then have "e/2 > 0" by auto + then guess d + apply - + apply (drule y[rule_format]) + apply (elim exE conjE) + done + note d=this[rule_format] + show ?case + proof (rule_tac x=d in exI, clarsimp simp: d) + fix p1 p2 + assume as: "p1 tagged_division_of (cbox a b)" "d fine p1" + "p2 tagged_division_of (cbox a b)" "d fine p2" + show "norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e" + apply (rule dist_triangle_half_l[where y=y,unfolded dist_norm]) + using d(2)[OF conjI[OF as(1-2)]] d(2)[OF conjI[OF as(3-4)]] . + qed + qed +next + assume "\e>0. \d. ?P e d" + then have "\n::nat. \d. ?P (inverse(of_nat (n + 1))) d" + by auto + from choice[OF this] guess d .. note d=conjunctD2[OF this[rule_format],rule_format] + have "\n. gauge (\x. \{d i x |i. i \ {0..n}})" + apply (rule gauge_inters) + using d(1) + apply auto + done + then have "\n. \p. p tagged_division_of (cbox a b) \ (\x. \{d i x |i. i \ {0..n}}) fine p" + by (meson fine_division_exists) + from choice[OF this] guess p .. note p = conjunctD2[OF this[rule_format]] + have dp: "\i n. i\n \ d i fine p n" + using p(2) unfolding fine_inters by auto + have "Cauchy (\n. setsum (\(x,k). content k *\<^sub>R (f x)) (p n))" + proof (rule CauchyI, goal_cases) + case (1 e) + then guess N unfolding real_arch_inverse[of e] .. note N=this + show ?case + apply (rule_tac x=N in exI) + proof clarify + fix m n + assume mn: "N \ m" "N \ n" + have *: "N = (N - 1) + 1" using N by auto + show "norm ((\(x, k)\p m. content k *\<^sub>R f x) - (\(x, k)\p n. content k *\<^sub>R f x)) < e" + apply (rule less_trans[OF _ N[THEN conjunct2,THEN conjunct2]]) + apply(subst *) + using dp p(1) mn d(2) by auto + qed + qed + then guess y unfolding convergent_eq_cauchy[symmetric] .. note y=this[THEN LIMSEQ_D] + show ?l + unfolding integrable_on_def has_integral + proof (rule_tac x=y in exI, clarify) + fix e :: real + assume "e>0" + then have *:"e/2 > 0" by auto + then guess N1 unfolding real_arch_inverse[of "e/2"] .. note N1=this + then have N1': "N1 = N1 - 1 + 1" + by auto + guess N2 using y[OF *] .. note N2=this + have "gauge (d (N1 + N2))" + using d by auto + moreover + { + fix q + assume as: "q tagged_division_of (cbox a b)" "d (N1 + N2) fine q" + have *: "inverse (of_nat (N1 + N2 + 1)) < e / 2" + apply (rule less_trans) + using N1 + apply auto + done + have "norm ((\(x, k)\q. content k *\<^sub>R f x) - y) < e" + apply (rule norm_triangle_half_r) + apply (rule less_trans[OF _ *]) + apply (subst N1', rule d(2)[of "p (N1+N2)"]) + using N1' as(1) as(2) dp + apply (simp add: \\x. p x tagged_division_of cbox a b \ (\xa. \{d i xa |i. i \ {0..x}}) fine p x\) + using N2 le_add2 by blast + } + ultimately show "\d. gauge d \ + (\p. p tagged_division_of (cbox a b) \ d fine p \ + norm ((\(x, k)\p. content k *\<^sub>R f x) - y) < e)" + by (rule_tac x="d (N1 + N2)" in exI) auto + qed +qed + + +subsection \Additivity of integral on abutting intervals.\ + +lemma tagged_division_split_left_inj: + fixes x1 :: "'a::euclidean_space" + assumes d: "d tagged_division_of i" + and k12: "(x1, k1) \ d" + "(x2, k2) \ d" + "k1 \ k2" + "k1 \ {x. x\k \ c} = k2 \ {x. x\k \ c}" + "k \ Basis" + shows "content (k1 \ {x. x\k \ c}) = 0" +proof - + have *: "\a b c. (a,b) \ c \ b \ snd ` c" + by force + show ?thesis + using k12 + by (fastforce intro!: division_split_left_inj[OF division_of_tagged_division[OF d]] *) +qed + +lemma tagged_division_split_right_inj: + fixes x1 :: "'a::euclidean_space" + assumes d: "d tagged_division_of i" + and k12: "(x1, k1) \ d" + "(x2, k2) \ d" + "k1 \ k2" + "k1 \ {x. x\k \ c} = k2 \ {x. x\k \ c}" + "k \ Basis" + shows "content (k1 \ {x. x\k \ c}) = 0" +proof - + have *: "\a b c. (a,b) \ c \ b \ snd ` c" + by force + show ?thesis + using k12 + by (fastforce intro!: division_split_right_inj[OF division_of_tagged_division[OF d]] *) +qed + +lemma has_integral_split: + fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" + assumes fi: "(f has_integral i) (cbox a b \ {x. x\k \ c})" + and fj: "(f has_integral j) (cbox a b \ {x. x\k \ c})" + and k: "k \ Basis" + shows "(f has_integral (i + j)) (cbox a b)" +proof (unfold has_integral, rule, rule, goal_cases) + case (1 e) + then have e: "e/2 > 0" + by auto + obtain d1 + where d1: "gauge d1" + and d1norm: + "\p. \p tagged_division_of cbox a b \ {x. x \ k \ c}; + d1 fine p\ \ norm ((\(x, k) \ p. content k *\<^sub>R f x) - i) < e / 2" + apply (rule has_integralD[OF fi[unfolded interval_split[OF k]] e]) + apply (simp add: interval_split[symmetric] k) + done + obtain d2 + where d2: "gauge d2" + and d2norm: + "\p. \p tagged_division_of cbox a b \ {x. c \ x \ k}; + d2 fine p\ \ norm ((\(x, k) \ p. content k *\<^sub>R f x) - j) < e / 2" + apply (rule has_integralD[OF fj[unfolded interval_split[OF k]] e]) + apply (simp add: interval_split[symmetric] k) + done + let ?d = "\x. if x\k = c then (d1 x \ d2 x) else ball x \x\k - c\ \ d1 x \ d2 x" + have "gauge ?d" + using d1 d2 unfolding gauge_def by auto + then show ?case + proof (rule_tac x="?d" in exI, safe) + fix p + assume "p tagged_division_of (cbox a b)" "?d fine p" + note p = this tagged_division_ofD[OF this(1)] + have xk_le_c: "\x kk. (x, kk) \ p \ kk \ {x. x\k \ c} \ {} \ x\k \ c" + proof - + fix x kk + assume as: "(x, kk) \ p" and kk: "kk \ {x. x\k \ c} \ {}" + show "x\k \ c" + proof (rule ccontr) + assume **: "\ ?thesis" + from this[unfolded not_le] + have "kk \ ball x \x \ k - c\" + using p(2)[unfolded fine_def, rule_format,OF as] by auto + with kk obtain y where y: "y \ ball x \x \ k - c\" "y\k \ c" + by blast + then have "\x \ k - y \ k\ < \x \ k - c\" + using Basis_le_norm[OF k, of "x - y"] + by (auto simp add: dist_norm inner_diff_left intro: le_less_trans) + with y show False + using ** by (auto simp add: field_simps) + qed + qed + have xk_ge_c: "\x kk. (x, kk) \ p \ kk \ {x. x\k \ c} \ {} \ x\k \ c" + proof - + fix x kk + assume as: "(x, kk) \ p" and kk: "kk \ {x. x\k \ c} \ {}" + show "x\k \ c" + proof (rule ccontr) + assume **: "\ ?thesis" + from this[unfolded not_le] have "kk \ ball x \x \ k - c\" + using p(2)[unfolded fine_def,rule_format,OF as,unfolded split_conv] by auto + with kk obtain y where y: "y \ ball x \x \ k - c\" "y\k \ c" + by blast + then have "\x \ k - y \ k\ < \x \ k - c\" + using Basis_le_norm[OF k, of "x - y"] + by (auto simp add: dist_norm inner_diff_left intro: le_less_trans) + with y show False + using ** by (auto simp add: field_simps) + qed + qed + + have lem1: "\f P Q. (\x k. (x, k) \ {(x, f k) | x k. P x k} \ Q x k) \ + (\x k. P x k \ Q x (f k))" + by auto + have fin_finite: "finite {(x,f k) | x k. (x,k) \ s \ P x k}" if "finite s" for f s P + proof - + from that have "finite ((\(x, k). (x, f k)) ` s)" + by auto + then show ?thesis + by (rule rev_finite_subset) auto + qed + { fix g :: "'a set \ 'a set" + fix i :: "'a \ 'a set" + assume "i \ (\(x, k). (x, g k)) ` p - {(x, g k) |x k. (x, k) \ p \ g k \ {}}" + then obtain x k where xk: + "i = (x, g k)" "(x, k) \ p" + "(x, g k) \ {(x, g k) |x k. (x, k) \ p \ g k \ {}}" + by auto + have "content (g k) = 0" + using xk using content_empty by auto + then have "(\(x, k). content k *\<^sub>R f x) i = 0" + unfolding xk split_conv by auto + } note [simp] = this + have lem3: "\g :: 'a set \ 'a set. finite p \ + setsum (\(x, k). content k *\<^sub>R f x) {(x,g k) |x k. (x,k) \ p \ g k \ {}} = + setsum (\(x, k). content k *\<^sub>R f x) ((\(x, k). (x, g k)) ` p)" + by (rule setsum.mono_neutral_left) auto + let ?M1 = "{(x, kk \ {x. x\k \ c}) |x kk. (x, kk) \ p \ kk \ {x. x\k \ c} \ {}}" + have d1_fine: "d1 fine ?M1" + by (force intro: fineI dest: fineD[OF p(2)] simp add: split: if_split_asm) + have "norm ((\(x, k)\?M1. content k *\<^sub>R f x) - i) < e/2" + proof (rule d1norm [OF tagged_division_ofI d1_fine]) + show "finite ?M1" + by (rule fin_finite p(3))+ + show "\{k. \x. (x, k) \ ?M1} = cbox a b \ {x. x\k \ c}" + unfolding p(8)[symmetric] by auto + fix x l + assume xl: "(x, l) \ ?M1" + then guess x' l' unfolding mem_Collect_eq unfolding prod.inject by (elim exE conjE) note xl'=this + show "x \ l" "l \ cbox a b \ {x. x \ k \ c}" + unfolding xl' + using p(4-6)[OF xl'(3)] using xl'(4) + using xk_le_c[OF xl'(3-4)] by auto + show "\a b. l = cbox a b" + unfolding xl' + using p(6)[OF xl'(3)] + by (fastforce simp add: interval_split[OF k,where c=c]) + fix y r + let ?goal = "interior l \ interior r = {}" + assume yr: "(y, r) \ ?M1" + then guess y' r' unfolding mem_Collect_eq unfolding prod.inject by (elim exE conjE) note yr'=this + assume as: "(x, l) \ (y, r)" + show "interior l \ interior r = {}" + proof (cases "l' = r' \ x' = y'") + case False + then show ?thesis + using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto + next + case True + then have "l' \ r'" + using as unfolding xl' yr' by auto + then show ?thesis + using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto + qed + qed + moreover + let ?M2 = "{(x,kk \ {x. x\k \ c}) |x kk. (x,kk) \ p \ kk \ {x. x\k \ c} \ {}}" + have d2_fine: "d2 fine ?M2" + by (force intro: fineI dest: fineD[OF p(2)] simp add: split: if_split_asm) + have "norm ((\(x, k)\?M2. content k *\<^sub>R f x) - j) < e/2" + proof (rule d2norm [OF tagged_division_ofI d2_fine]) + show "finite ?M2" + by (rule fin_finite p(3))+ + show "\{k. \x. (x, k) \ ?M2} = cbox a b \ {x. x\k \ c}" + unfolding p(8)[symmetric] by auto + fix x l + assume xl: "(x, l) \ ?M2" + then guess x' l' unfolding mem_Collect_eq unfolding prod.inject by (elim exE conjE) note xl'=this + show "x \ l" "l \ cbox a b \ {x. x \ k \ c}" + unfolding xl' + using p(4-6)[OF xl'(3)] xl'(4) xk_ge_c[OF xl'(3-4)] + by auto + show "\a b. l = cbox a b" + unfolding xl' + using p(6)[OF xl'(3)] + by (fastforce simp add: interval_split[OF k, where c=c]) + fix y r + let ?goal = "interior l \ interior r = {}" + assume yr: "(y, r) \ ?M2" + then guess y' r' unfolding mem_Collect_eq unfolding prod.inject by (elim exE conjE) note yr'=this + assume as: "(x, l) \ (y, r)" + show "interior l \ interior r = {}" + proof (cases "l' = r' \ x' = y'") + case False + then show ?thesis + using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto + next + case True + then have "l' \ r'" + using as unfolding xl' yr' by auto + then show ?thesis + using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto + qed + qed + ultimately + have "norm (((\(x, k)\?M1. content k *\<^sub>R f x) - i) + ((\(x, k)\?M2. content k *\<^sub>R f x) - j)) < e/2 + e/2" + using norm_add_less by blast + also { + have eq0: "\x y. x = (0::real) \ x *\<^sub>R (y::'b) = 0" + using scaleR_zero_left by auto + have cont_eq: "\g. (\(x,l). content l *\<^sub>R f x) \ (\(x,l). (x,g l)) = (\(x,l). content (g l) *\<^sub>R f x)" + by auto + have "((\(x, k)\?M1. content k *\<^sub>R f x) - i) + ((\(x, k)\?M2. content k *\<^sub>R f x) - j) = + (\(x, k)\?M1. content k *\<^sub>R f x) + (\(x, k)\?M2. content k *\<^sub>R f x) - (i + j)" + by auto + also have "\ = (\(x, ka)\p. content (ka \ {x. x \ k \ c}) *\<^sub>R f x) + + (\(x, ka)\p. content (ka \ {x. c \ x \ k}) *\<^sub>R f x) - (i + j)" + unfolding lem3[OF p(3)] + by (subst setsum.reindex_nontrivial[OF p(3)], auto intro!: k eq0 tagged_division_split_left_inj[OF p(1)] tagged_division_split_right_inj[OF p(1)] + simp: cont_eq)+ + also note setsum.distrib[symmetric] + also have "\x. x \ p \ + (\(x,ka). content (ka \ {x. x \ k \ c}) *\<^sub>R f x) x + + (\(x,ka). content (ka \ {x. c \ x \ k}) *\<^sub>R f x) x = + (\(x,ka). content ka *\<^sub>R f x) x" + proof clarify + fix a b + assume "(a, b) \ p" + from p(6)[OF this] guess u v by (elim exE) note uv=this + then show "content (b \ {x. x \ k \ c}) *\<^sub>R f a + content (b \ {x. c \ x \ k}) *\<^sub>R f a = + content b *\<^sub>R f a" + unfolding scaleR_left_distrib[symmetric] + unfolding uv content_split[OF k,of u v c] + by auto + qed + note setsum.cong [OF _ this] + finally have "(\(x, k)\{(x, kk \ {x. x \ k \ c}) |x kk. (x, kk) \ p \ kk \ {x. x \ k \ c} \ {}}. content k *\<^sub>R f x) - i + + ((\(x, k)\{(x, kk \ {x. c \ x \ k}) |x kk. (x, kk) \ p \ kk \ {x. c \ x \ k} \ {}}. content k *\<^sub>R f x) - j) = + (\(x, ka)\p. content ka *\<^sub>R f x) - (i + j)" + by auto + } + finally show "norm ((\(x, k)\p. content k *\<^sub>R f x) - (i + j)) < e" + by auto + qed +qed + + +subsection \A sort of converse, integrability on subintervals.\ + +lemma tagged_division_union_interval: + fixes a :: "'a::euclidean_space" + assumes "p1 tagged_division_of (cbox a b \ {x. x\k \ (c::real)})" + and "p2 tagged_division_of (cbox a b \ {x. x\k \ c})" + and k: "k \ Basis" + shows "(p1 \ p2) tagged_division_of (cbox a b)" +proof - + have *: "cbox a b = (cbox a b \ {x. x\k \ c}) \ (cbox a b \ {x. x\k \ c})" + by auto + show ?thesis + apply (subst *) + apply (rule tagged_division_union[OF assms(1-2)]) + unfolding interval_split[OF k] interior_cbox + using k + apply (auto simp add: box_def elim!: ballE[where x=k]) + done +qed + +lemma tagged_division_union_interval_real: + fixes a :: real + assumes "p1 tagged_division_of ({a .. b} \ {x. x\k \ (c::real)})" + and "p2 tagged_division_of ({a .. b} \ {x. x\k \ c})" + and k: "k \ Basis" + shows "(p1 \ p2) tagged_division_of {a .. b}" + using assms + unfolding box_real[symmetric] + by (rule tagged_division_union_interval) + +lemma has_integral_separate_sides: + fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" + assumes "(f has_integral i) (cbox a b)" + and "e > 0" + and k: "k \ Basis" + obtains d where "gauge d" + "\p1 p2. p1 tagged_division_of (cbox a b \ {x. x\k \ c}) \ d fine p1 \ + p2 tagged_division_of (cbox a b \ {x. x\k \ c}) \ d fine p2 \ + norm ((setsum (\(x,k). content k *\<^sub>R f x) p1 + setsum (\(x,k). content k *\<^sub>R f x) p2) - i) < e" +proof - + guess d using has_integralD[OF assms(1-2)] . note d=this + { fix p1 p2 + assume "p1 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p1" + note p1=tagged_division_ofD[OF this(1)] this + assume "p2 tagged_division_of (cbox a b) \ {x. c \ x \ k}" "d fine p2" + note p2=tagged_division_ofD[OF this(1)] this + note tagged_division_union_interval[OF p1(7) p2(7)] note p12 = tagged_division_ofD[OF this] this + { fix a b + assume ab: "(a, b) \ p1 \ p2" + have "(a, b) \ p1" + using ab by auto + with p1 obtain u v where uv: "b = cbox u v" by auto + have "b \ {x. x\k = c}" + using ab p1(3)[of a b] p2(3)[of a b] by fastforce + moreover + have "interior {x::'a. x \ k = c} = {}" + proof (rule ccontr) + assume "\ ?thesis" + then obtain x where x: "x \ interior {x::'a. x\k = c}" + by auto + then guess e unfolding mem_interior .. note e=this + have x: "x\k = c" + using x interior_subset by fastforce + have *: "\i. i \ Basis \ \(x - (x + (e / 2) *\<^sub>R k)) \ i\ = (if i = k then e/2 else 0)" + using e k by (auto simp: inner_simps inner_not_same_Basis) + have "(\i\Basis. \(x - (x + (e / 2 ) *\<^sub>R k)) \ i\) = + (\i\Basis. (if i = k then e / 2 else 0))" + using "*" by (blast intro: setsum.cong) + also have "\ < e" + apply (subst setsum.delta) + using e + apply auto + done + finally have "x + (e/2) *\<^sub>R k \ ball x e" + unfolding mem_ball dist_norm by(rule le_less_trans[OF norm_le_l1]) + then have "x + (e/2) *\<^sub>R k \ {x. x\k = c}" + using e by auto + then show False + unfolding mem_Collect_eq using e x k by (auto simp: inner_simps) + qed + ultimately have "content b = 0" + unfolding uv content_eq_0_interior + using interior_mono by blast + then have "content b *\<^sub>R f a = 0" + by auto + } + then have "norm ((\(x, k)\p1. content k *\<^sub>R f x) + (\(x, k)\p2. content k *\<^sub>R f x) - i) = + norm ((\(x, k)\p1 \ p2. content k *\<^sub>R f x) - i)" + by (subst setsum.union_inter_neutral) (auto simp: p1 p2) + also have "\ < e" + by (rule k d(2) p12 fine_union p1 p2)+ + finally have "norm ((\(x, k)\p1. content k *\<^sub>R f x) + (\(x, k)\p2. content k *\<^sub>R f x) - i) < e" . + } + then show ?thesis + by (auto intro: that[of d] d elim: ) +qed + +lemma integrable_split[intro]: + fixes f :: "'a::euclidean_space \ 'b::{real_normed_vector,complete_space}" + assumes "f integrable_on cbox a b" + and k: "k \ Basis" + shows "f integrable_on (cbox a b \ {x. x\k \ c})" (is ?t1) + and "f integrable_on (cbox a b \ {x. x\k \ c})" (is ?t2) +proof - + guess y using assms(1) unfolding integrable_on_def .. note y=this + define b' where "b' = (\i\Basis. (if i = k then min (b\k) c else b\i)*\<^sub>R i)" + define a' where "a' = (\i\Basis. (if i = k then max (a\k) c else a\i)*\<^sub>R i)" + show ?t1 ?t2 + unfolding interval_split[OF k] integrable_cauchy + unfolding interval_split[symmetric,OF k] + proof (rule_tac[!] allI impI)+ + fix e :: real + assume "e > 0" + then have "e/2>0" + by auto + from has_integral_separate_sides[OF y this k,of c] guess d . note d=this[rule_format] + let ?P = "\A. \d. gauge d \ (\p1 p2. p1 tagged_division_of (cbox a b) \ A \ d fine p1 \ + p2 tagged_division_of (cbox a b) \ A \ d fine p2 \ + norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e)" + show "?P {x. x \ k \ c}" + proof (rule_tac x=d in exI, clarsimp simp add: d) + fix p1 p2 + assume as: "p1 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p1" + "p2 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p2" + show "norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e" + proof (rule fine_division_exists[OF d(1), of a' b] ) + fix p + assume "p tagged_division_of cbox a' b" "d fine p" + then show ?thesis + using as norm_triangle_half_l[OF d(2)[of p1 p] d(2)[of p2 p]] + unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric] + by (auto simp add: algebra_simps) + qed + qed + show "?P {x. x \ k \ c}" + proof (rule_tac x=d in exI, clarsimp simp add: d) + fix p1 p2 + assume as: "p1 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p1" + "p2 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p2" + show "norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e" + proof (rule fine_division_exists[OF d(1), of a b'] ) + fix p + assume "p tagged_division_of cbox a b'" "d fine p" + then show ?thesis + using as norm_triangle_half_l[OF d(2)[of p p1] d(2)[of p p2]] + unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric] + by (auto simp add: algebra_simps) + qed + qed + qed +qed + +lemma operative_integral: + fixes f :: "'a::euclidean_space \ 'b::banach" + shows "comm_monoid.operative (lift_option op +) (Some 0) (\i. if f integrable_on i then Some (integral i f) else None)" + unfolding comm_monoid.operative_def[OF add.comm_monoid_lift_option] +proof safe + fix a b c + fix k :: 'a + assume k: "k \ Basis" + show "(if f integrable_on cbox a b then Some (integral (cbox a b) f) else None) = + lift_option op + (if f integrable_on cbox a b \ {x. x \ k \ c} then Some (integral (cbox a b \ {x. x \ k \ c}) f) else None) + (if f integrable_on cbox a b \ {x. c \ x \ k} then Some (integral (cbox a b \ {x. c \ x \ k}) f) else None)" + proof (cases "f integrable_on cbox a b") + case True + with k show ?thesis + apply (simp add: integrable_split) + apply (rule integral_unique [OF has_integral_split[OF _ _ k]]) + apply (auto intro: integrable_integral) + done + next + case False + have "\ (f integrable_on cbox a b \ {x. x \ k \ c}) \ \ ( f integrable_on cbox a b \ {x. c \ x \ k})" + proof (rule ccontr) + assume "\ ?thesis" + then have "f integrable_on cbox a b" + unfolding integrable_on_def + apply (rule_tac x="integral (cbox a b \ {x. x \ k \ c}) f + integral (cbox a b \ {x. x \ k \ c}) f" in exI) + apply (rule has_integral_split[OF _ _ k]) + apply (auto intro: integrable_integral) + done + then show False + using False by auto + qed + then show ?thesis + using False by auto + qed +next + fix a b :: 'a + assume "content (cbox a b) = 0" + then show "(if f integrable_on cbox a b then Some (integral (cbox a b) f) else None) = Some 0" + using has_integral_null_eq + by (auto simp: integrable_on_null) +qed + +subsection \Finally, the integral of a constant\ + +lemma has_integral_const [intro]: + fixes a b :: "'a::euclidean_space" + shows "((\x. c) has_integral (content (cbox a b) *\<^sub>R c)) (cbox a b)" + apply (auto intro!: exI [where x="\x. ball x 1"] simp: split_def has_integral) + apply (subst scaleR_left.setsum[symmetric, unfolded o_def]) + apply (subst additive_content_tagged_division[unfolded split_def]) + apply auto + done + +lemma has_integral_const_real [intro]: + fixes a b :: real + shows "((\x. c) has_integral (content {a .. b} *\<^sub>R c)) {a .. b}" + by (metis box_real(2) has_integral_const) + +lemma integral_const [simp]: + fixes a b :: "'a::euclidean_space" + shows "integral (cbox a b) (\x. c) = content (cbox a b) *\<^sub>R c" + by (rule integral_unique) (rule has_integral_const) + +lemma integral_const_real [simp]: + fixes a b :: real + shows "integral {a .. b} (\x. c) = content {a .. b} *\<^sub>R c" + by (metis box_real(2) integral_const) + + +subsection \Bounds on the norm of Riemann sums and the integral itself.\ + +lemma dsum_bound: + assumes "p division_of (cbox a b)" + and "norm c \ e" + shows "norm (setsum (\l. content l *\<^sub>R c) p) \ e * content(cbox a b)" +proof - + have sumeq: "(\i\p. \content i\) = setsum content p" + apply (rule setsum.cong) + using assms + apply simp + apply (metis abs_of_nonneg assms(1) content_pos_le division_ofD(4)) + done + have e: "0 \ e" + using assms(2) norm_ge_zero order_trans by blast + have "norm (setsum (\l. content l *\<^sub>R c) p) \ (\i\p. norm (content i *\<^sub>R c))" + using norm_setsum by blast + also have "... \ e * (\i\p. \content i\)" + apply (simp add: setsum_right_distrib[symmetric] mult.commute) + using assms(2) mult_right_mono by blast + also have "... \ e * content (cbox a b)" + apply (rule mult_left_mono [OF _ e]) + apply (simp add: sumeq) + using additive_content_division assms(1) eq_iff apply blast + done + finally show ?thesis . +qed + +lemma rsum_bound: + assumes p: "p tagged_division_of (cbox a b)" + and "\x\cbox a b. norm (f x) \ e" + shows "norm (setsum (\(x,k). content k *\<^sub>R f x) p) \ e * content (cbox a b)" +proof (cases "cbox a b = {}") + case True show ?thesis + using p unfolding True tagged_division_of_trivial by auto +next + case False + then have e: "e \ 0" + by (meson ex_in_conv assms(2) norm_ge_zero order_trans) + have setsum_le: "setsum (content \ snd) p \ content (cbox a b)" + unfolding additive_content_tagged_division[OF p, symmetric] split_def + by (auto intro: eq_refl) + have con: "\xk. xk \ p \ 0 \ content (snd xk)" + using tagged_division_ofD(4) [OF p] content_pos_le + by force + have norm: "\xk. xk \ p \ norm (f (fst xk)) \ e" + unfolding fst_conv using tagged_division_ofD(2,3)[OF p] assms + by (metis prod.collapse subset_eq) + have "norm (setsum (\(x,k). content k *\<^sub>R f x) p) \ (\i\p. norm (case i of (x, k) \ content k *\<^sub>R f x))" + by (rule norm_setsum) + also have "... \ e * content (cbox a b)" + unfolding split_def norm_scaleR + apply (rule order_trans[OF setsum_mono]) + apply (rule mult_left_mono[OF _ abs_ge_zero, of _ e]) + apply (metis norm) + unfolding setsum_left_distrib[symmetric] + using con setsum_le + apply (auto simp: mult.commute intro: mult_left_mono [OF _ e]) + done + finally show ?thesis . +qed + +lemma rsum_diff_bound: + assumes "p tagged_division_of (cbox a b)" + and "\x\cbox a b. norm (f x - g x) \ e" + shows "norm (setsum (\(x,k). content k *\<^sub>R f x) p - setsum (\(x,k). content k *\<^sub>R g x) p) \ + e * content (cbox a b)" + apply (rule order_trans[OF _ rsum_bound[OF assms]]) + apply (simp add: split_def scaleR_diff_right setsum_subtractf eq_refl) + done + +lemma has_integral_bound: + fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" + assumes "0 \ B" + and "(f has_integral i) (cbox a b)" + and "\x\cbox a b. norm (f x) \ B" + shows "norm i \ B * content (cbox a b)" +proof (rule ccontr) + assume "\ ?thesis" + then have *: "norm i - B * content (cbox a b) > 0" + by auto + from assms(2)[unfolded has_integral,rule_format,OF *] + guess d by (elim exE conjE) note d=this[rule_format] + from fine_division_exists[OF this(1), of a b] guess p . note p=this + have *: "\s B. norm s \ B \ \ norm (s - i) < norm i - B" + unfolding not_less + by (metis norm_triangle_sub[of i] add.commute le_less_trans less_diff_eq linorder_not_le norm_minus_commute) + show False + using d(2)[OF conjI[OF p]] *[OF rsum_bound[OF p(1) assms(3)]] by auto +qed + +corollary has_integral_bound_real: + fixes f :: "real \ 'b::real_normed_vector" + assumes "0 \ B" + and "(f has_integral i) {a .. b}" + and "\x\{a .. b}. norm (f x) \ B" + shows "norm i \ B * content {a .. b}" + by (metis assms box_real(2) has_integral_bound) + +corollary integrable_bound: + fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" + assumes "0 \ B" + and "f integrable_on (cbox a b)" + and "\x. x\cbox a b \ norm (f x) \ B" + shows "norm (integral (cbox a b) f) \ B * content (cbox a b)" +by (metis integrable_integral has_integral_bound assms) + + +subsection \Similar theorems about relationship among components.\ + +lemma rsum_component_le: + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes "p tagged_division_of (cbox a b)" + and "\x\cbox a b. (f x)\i \ (g x)\i" + shows "(setsum (\(x,k). content k *\<^sub>R f x) p)\i \ (setsum (\(x,k). content k *\<^sub>R g x) p)\i" +unfolding inner_setsum_left +proof (rule setsum_mono, clarify) + fix a b + assume ab: "(a, b) \ p" + note tagged = tagged_division_ofD(2-4)[OF assms(1) ab] + from this(3) guess u v by (elim exE) note b=this + show "(content b *\<^sub>R f a) \ i \ (content b *\<^sub>R g a) \ i" + unfolding b inner_simps real_scaleR_def + apply (rule mult_left_mono) + using assms(2) tagged + by (auto simp add: content_pos_le) +qed + +lemma has_integral_component_le: + fixes f g :: "'a::euclidean_space \ 'b::euclidean_space" + assumes k: "k \ Basis" + assumes "(f has_integral i) s" "(g has_integral j) s" + and "\x\s. (f x)\k \ (g x)\k" + shows "i\k \ j\k" +proof - + have lem: "i\k \ j\k" + if f_i: "(f has_integral i) (cbox a b)" + and g_j: "(g has_integral j) (cbox a b)" + and le: "\x\cbox a b. (f x)\k \ (g x)\k" + for a b i and j :: 'b and f g :: "'a \ 'b" + proof (rule ccontr) + assume "\ ?thesis" + then have *: "0 < (i\k - j\k) / 3" + by auto + guess d1 using f_i[unfolded has_integral,rule_format,OF *] by (elim exE conjE) note d1=this[rule_format] + guess d2 using g_j[unfolded has_integral,rule_format,OF *] by (elim exE conjE) note d2=this[rule_format] + obtain p where p: "p tagged_division_of cbox a b" "d1 fine p" "d2 fine p" + using fine_division_exists[OF gauge_inter[OF d1(1) d2(1)], of a b] unfolding fine_inter + by metis + note le_less_trans[OF Basis_le_norm[OF k]] + then have "\((\(x, k)\p. content k *\<^sub>R f x) - i) \ k\ < (i \ k - j \ k) / 3" + "\((\(x, k)\p. content k *\<^sub>R g x) - j) \ k\ < (i \ k - j \ k) / 3" + using k norm_bound_Basis_lt d1 d2 p + by blast+ + then show False + unfolding inner_simps + using rsum_component_le[OF p(1) le] + by (simp add: abs_real_def split: if_split_asm) + qed + show ?thesis + proof (cases "\a b. s = cbox a b") + case True + with lem assms show ?thesis + by auto + next + case False + show ?thesis + proof (rule ccontr) + assume "\ i\k \ j\k" + then have ij: "(i\k - j\k) / 3 > 0" + by auto + note has_integral_altD[OF _ False this] + from this[OF assms(2)] this[OF assms(3)] guess B1 B2 . note B=this[rule_format] + have "bounded (ball 0 B1 \ ball (0::'a) B2)" + unfolding bounded_Un by(rule conjI bounded_ball)+ + from bounded_subset_cbox[OF this] guess a b by (elim exE) + note ab = conjunctD2[OF this[unfolded Un_subset_iff]] + guess w1 using B(2)[OF ab(1)] .. note w1=conjunctD2[OF this] + guess w2 using B(4)[OF ab(2)] .. note w2=conjunctD2[OF this] + have *: "\w1 w2 j i::real .\w1 - i\ < (i - j) / 3 \ \w2 - j\ < (i - j) / 3 \ w1 \ w2 \ False" + by (simp add: abs_real_def split: if_split_asm) + note le_less_trans[OF Basis_le_norm[OF k]] + note this[OF w1(2)] this[OF w2(2)] + moreover + have "w1\k \ w2\k" + by (rule lem[OF w1(1) w2(1)]) (simp add: assms(4)) + ultimately show False + unfolding inner_simps by(rule *) + qed + qed +qed + +lemma integral_component_le: + fixes g f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes "k \ Basis" + and "f integrable_on s" "g integrable_on s" + and "\x\s. (f x)\k \ (g x)\k" + shows "(integral s f)\k \ (integral s g)\k" + apply (rule has_integral_component_le) + using integrable_integral assms + apply auto + done + +lemma has_integral_component_nonneg: + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes "k \ Basis" + and "(f has_integral i) s" + and "\x\s. 0 \ (f x)\k" + shows "0 \ i\k" + using has_integral_component_le[OF assms(1) has_integral_0 assms(2)] + using assms(3-) + by auto + +lemma integral_component_nonneg: + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes "k \ Basis" + and "\x\s. 0 \ (f x)\k" + shows "0 \ (integral s f)\k" +proof (cases "f integrable_on s") + case True show ?thesis + apply (rule has_integral_component_nonneg) + using assms True + apply auto + done +next + case False then show ?thesis by (simp add: not_integrable_integral) +qed + +lemma has_integral_component_neg: + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes "k \ Basis" + and "(f has_integral i) s" + and "\x\s. (f x)\k \ 0" + shows "i\k \ 0" + using has_integral_component_le[OF assms(1,2) has_integral_0] assms(2-) + by auto + +lemma has_integral_component_lbound: + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes "(f has_integral i) (cbox a b)" + and "\x\cbox a b. B \ f(x)\k" + and "k \ Basis" + shows "B * content (cbox a b) \ i\k" + using has_integral_component_le[OF assms(3) has_integral_const assms(1),of "(\i\Basis. B *\<^sub>R i)::'b"] assms(2-) + by (auto simp add: field_simps) + +lemma has_integral_component_ubound: + fixes f::"'a::euclidean_space => 'b::euclidean_space" + assumes "(f has_integral i) (cbox a b)" + and "\x\cbox a b. f x\k \ B" + and "k \ Basis" + shows "i\k \ B * content (cbox a b)" + using has_integral_component_le[OF assms(3,1) has_integral_const, of "\i\Basis. B *\<^sub>R i"] assms(2-) + by (auto simp add: field_simps) + +lemma integral_component_lbound: + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes "f integrable_on cbox a b" + and "\x\cbox a b. B \ f(x)\k" + and "k \ Basis" + shows "B * content (cbox a b) \ (integral(cbox a b) f)\k" + apply (rule has_integral_component_lbound) + using assms + unfolding has_integral_integral + apply auto + done + +lemma integral_component_lbound_real: + assumes "f integrable_on {a ::real .. b}" + and "\x\{a .. b}. B \ f(x)\k" + and "k \ Basis" + shows "B * content {a .. b} \ (integral {a .. b} f)\k" + using assms + by (metis box_real(2) integral_component_lbound) + +lemma integral_component_ubound: + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes "f integrable_on cbox a b" + and "\x\cbox a b. f x\k \ B" + and "k \ Basis" + shows "(integral (cbox a b) f)\k \ B * content (cbox a b)" + apply (rule has_integral_component_ubound) + using assms + unfolding has_integral_integral + apply auto + done + +lemma integral_component_ubound_real: + fixes f :: "real \ 'a::euclidean_space" + assumes "f integrable_on {a .. b}" + and "\x\{a .. b}. f x\k \ B" + and "k \ Basis" + shows "(integral {a .. b} f)\k \ B * content {a .. b}" + using assms + by (metis box_real(2) integral_component_ubound) + +subsection \Uniform limit of integrable functions is integrable.\ + +lemma real_arch_invD: + "0 < (e::real) \ (\n::nat. n \ 0 \ 0 < inverse (real n) \ inverse (real n) < e)" + by (subst(asm) real_arch_inverse) + +lemma integrable_uniform_limit: + fixes f :: "'a::euclidean_space \ 'b::banach" + assumes "\e>0. \g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" + shows "f integrable_on cbox a b" +proof (cases "content (cbox a b) > 0") + case False then show ?thesis + using has_integral_null + by (simp add: content_lt_nz integrable_on_def) +next + case True + have *: "\P. \e>(0::real). P e \ \n::nat. P (inverse (real n + 1))" + by auto + from choice[OF *[OF assms]] guess g .. note g=conjunctD2[OF this[rule_format],rule_format] + from choice[OF allI[OF g(2)[unfolded integrable_on_def], of "\x. x"]] + obtain i where i: "\x. (g x has_integral i x) (cbox a b)" + by auto + have "Cauchy i" + unfolding Cauchy_def + proof clarify + fix e :: real + assume "e>0" + then have "e / 4 / content (cbox a b) > 0" + using True by (auto simp add: field_simps) + then obtain M :: nat + where M: "M \ 0" "0 < inverse (real_of_nat M)" "inverse (of_nat M) < e / 4 / content (cbox a b)" + by (subst (asm) real_arch_inverse) auto + show "\M. \m\M. \n\M. dist (i m) (i n) < e" + proof (rule exI [where x=M], clarify) + fix m n + assume m: "M \ m" and n: "M \ n" + have "e/4>0" using \e>0\ by auto + note * = i[unfolded has_integral,rule_format,OF this] + from *[of m] guess gm by (elim conjE exE) note gm=this[rule_format] + from *[of n] guess gn by (elim conjE exE) note gn=this[rule_format] + from fine_division_exists[OF gauge_inter[OF gm(1) gn(1)], of a b] + obtain p where p: "p tagged_division_of cbox a b" "(\x. gm x \ gn x) fine p" + by auto + { fix s1 s2 i1 and i2::'b + assume no: "norm(s2 - s1) \ e/2" "norm (s1 - i1) < e/4" "norm (s2 - i2) < e/4" + have "norm (i1 - i2) \ norm (i1 - s1) + norm (s1 - s2) + norm (s2 - i2)" + using norm_triangle_ineq[of "i1 - s1" "s1 - i2"] + using norm_triangle_ineq[of "s1 - s2" "s2 - i2"] + by (auto simp add: algebra_simps) + also have "\ < e" + using no + unfolding norm_minus_commute + by (auto simp add: algebra_simps) + finally have "norm (i1 - i2) < e" . + } note triangle3 = this + have finep: "gm fine p" "gn fine p" + using fine_inter p by auto + { fix x + assume x: "x \ cbox a b" + have "norm (f x - g n x) + norm (f x - g m x) \ inverse (real n + 1) + inverse (real m + 1)" + using g(1)[OF x, of n] g(1)[OF x, of m] by auto + also have "\ \ inverse (real M) + inverse (real M)" + apply (rule add_mono) + using M(2) m n by auto + also have "\ = 2 / real M" + unfolding divide_inverse by auto + finally have "norm (g n x - g m x) \ 2 / real M" + using norm_triangle_le[of "g n x - f x" "f x - g m x" "2 / real M"] + by (auto simp add: algebra_simps simp add: norm_minus_commute) + } note norm_le = this + have le_e2: "norm ((\(x, k)\p. content k *\<^sub>R g n x) - (\(x, k)\p. content k *\<^sub>R g m x)) \ e / 2" + apply (rule order_trans [OF rsum_diff_bound[OF p(1), where e="2 / real M"]]) + apply (blast intro: norm_le) + using M True + by (auto simp add: field_simps) + then show "dist (i m) (i n) < e" + unfolding dist_norm + using gm gn p finep + by (auto intro!: triangle3) + qed + qed + then obtain s where s: "i \ s" + using convergent_eq_cauchy[symmetric] by blast + show ?thesis + unfolding integrable_on_def has_integral + proof (rule_tac x=s in exI, clarify) + fix e::real + assume e: "0 < e" + then have *: "e/3 > 0" by auto + then obtain N1 where N1: "\n\N1. norm (i n - s) < e / 3" + using LIMSEQ_D [OF s] by metis + from e True have "e / 3 / content (cbox a b) > 0" + by (auto simp add: field_simps) + from real_arch_invD[OF this] guess N2 by (elim exE conjE) note N2=this + from i[of "N1 + N2",unfolded has_integral,rule_format,OF *] guess g' .. note g'=conjunctD2[OF this,rule_format] + { fix sf sg i + assume no: "norm (sf - sg) \ e / 3" + "norm(i - s) < e / 3" + "norm (sg - i) < e / 3" + have "norm (sf - s) \ norm (sf - sg) + norm (sg - i) + norm (i - s)" + using norm_triangle_ineq[of "sf - sg" "sg - s"] + using norm_triangle_ineq[of "sg - i" " i - s"] + by (auto simp add: algebra_simps) + also have "\ < e" + using no + unfolding norm_minus_commute + by (auto simp add: algebra_simps) + finally have "norm (sf - s) < e" . + } note lem = this + { fix p + assume p: "p tagged_division_of (cbox a b) \ g' fine p" + then have norm_less: "norm ((\(x, k)\p. content k *\<^sub>R g (N1 + N2) x) - i (N1 + N2)) < e / 3" + using g' by blast + have "content (cbox a b) < e / 3 * (of_nat N2)" + using N2 unfolding inverse_eq_divide using True by (auto simp add: field_simps) + moreover have "e / 3 * of_nat N2 \ e / 3 * (of_nat (N1 + N2) + 1)" + using \e>0\ by auto + ultimately have "content (cbox a b) < e / 3 * (of_nat (N1 + N2) + 1)" + by linarith + then have le_e3: "inverse (real (N1 + N2) + 1) * content (cbox a b) \ e / 3" + unfolding inverse_eq_divide + by (auto simp add: field_simps) + have ne3: "norm (i (N1 + N2) - s) < e / 3" + using N1 by auto + have "norm ((\(x, k)\p. content k *\<^sub>R f x) - s) < e" + apply (rule lem[OF order_trans [OF _ le_e3] ne3 norm_less]) + apply (rule rsum_diff_bound[OF p[THEN conjunct1]]) + apply (blast intro: g) + done } + then show "\d. gauge d \ + (\p. p tagged_division_of cbox a b \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x) - s) < e)" + by (blast intro: g') + qed +qed + +lemmas integrable_uniform_limit_real = integrable_uniform_limit [where 'a=real, simplified] + + +subsection \Negligible sets.\ + +definition "negligible (s:: 'a::euclidean_space set) \ + (\a b. ((indicator s :: 'a\real) has_integral 0) (cbox a b))" + + +subsection \Negligibility of hyperplane.\ + +lemma interval_doublesplit: + fixes a :: "'a::euclidean_space" + assumes "k \ Basis" + shows "cbox a b \ {x . \x\k - c\ \ (e::real)} = + cbox (\i\Basis. (if i = k then max (a\k) (c - e) else a\i) *\<^sub>R i) + (\i\Basis. (if i = k then min (b\k) (c + e) else b\i) *\<^sub>R i)" +proof - + have *: "\x c e::real. \x - c\ \ e \ x \ c - e \ x \ c + e" + by auto + have **: "\s P Q. s \ {x. P x \ Q x} = (s \ {x. Q x}) \ {x. P x}" + by blast + show ?thesis + unfolding * ** interval_split[OF assms] by (rule refl) +qed + +lemma division_doublesplit: + fixes a :: "'a::euclidean_space" + assumes "p division_of (cbox a b)" + and k: "k \ Basis" + shows "(\l. l \ {x. \x\k - c\ \ e}) ` {l\p. l \ {x. \x\k - c\ \ e} \ {}} + division_of (cbox a b \ {x. \x\k - c\ \ e})" +proof - + have *: "\x c. \x - c\ \ e \ x \ c - e \ x \ c + e" + by auto + have **: "\p q p' q'. p division_of q \ p = p' \ q = q' \ p' division_of q'" + by auto + note division_split(1)[OF assms, where c="c+e",unfolded interval_split[OF k]] + note division_split(2)[OF this, where c="c-e" and k=k,OF k] + then show ?thesis + apply (rule **) + subgoal + apply (simp add: abs_diff_le_iff field_simps Collect_conj_eq setcompr_eq_image[symmetric]) + apply (rule equalityI) + apply blast + apply clarsimp + apply (rule_tac x="l \ {x. c + e \ x \ k}" in exI) + apply auto + done + by (simp add: interval_split k interval_doublesplit) +qed + +lemma content_doublesplit: + fixes a :: "'a::euclidean_space" + assumes "0 < e" + and k: "k \ Basis" + obtains d where "0 < d" and "content (cbox a b \ {x. \x\k - c\ \ d}) < e" +proof (cases "content (cbox a b) = 0") + case True + then have ce: "content (cbox a b) < e" + by (metis \0 < e\) + show ?thesis + apply (rule that[of 1]) + apply simp + unfolding interval_doublesplit[OF k] + apply (rule le_less_trans[OF content_subset ce]) + apply (auto simp: interval_doublesplit[symmetric] k) + done +next + case False + define d where "d = e / 3 / setprod (\i. b\i - a\i) (Basis - {k})" + note False[unfolded content_eq_0 not_ex not_le, rule_format] + then have "\x. x \ Basis \ b\x > a\x" + by (auto simp add:not_le) + then have prod0: "0 < setprod (\i. b\i - a\i) (Basis - {k})" + by (force simp add: setprod_pos field_simps) + then have "d > 0" + using assms + by (auto simp add: d_def field_simps) + then show ?thesis + proof (rule that[of d]) + have *: "Basis = insert k (Basis - {k})" + using k by auto + have less_e: "(min (b \ k) (c + d) - max (a \ k) (c - d)) * (\i\Basis - {k}. b \ i - a \ i) < e" + proof - + have "(min (b \ k) (c + d) - max (a \ k) (c - d)) \ 2 * d" + by auto + also have "\ < e / (\i\Basis - {k}. b \ i - a \ i)" + unfolding d_def + using assms prod0 + by (auto simp add: field_simps) + finally show "(min (b \ k) (c + d) - max (a \ k) (c - d)) * (\i\Basis - {k}. b \ i - a \ i) < e" + unfolding pos_less_divide_eq[OF prod0] . + qed + show "content (cbox a b \ {x. \x \ k - c\ \ d}) < e" + proof (cases "cbox a b \ {x. \x \ k - c\ \ d} = {}") + case True + then show ?thesis + using assms by simp + next + case False + then have + "(\i\Basis - {k}. interval_upperbound (cbox a b \ {x. \x \ k - c\ \ d}) \ i - + interval_lowerbound (cbox a b \ {x. \x \ k - c\ \ d}) \ i) + = (\i\Basis - {k}. b\i - a\i)" + by (simp add: box_eq_empty interval_doublesplit[OF k]) + then show "content (cbox a b \ {x. \x \ k - c\ \ d}) < e" + unfolding content_def + using assms False + apply (subst *) + apply (subst setprod.insert) + apply (simp_all add: interval_doublesplit[OF k] box_eq_empty not_less less_e) + done + qed + qed +qed + +lemma negligible_standard_hyperplane[intro]: + fixes k :: "'a::euclidean_space" + assumes k: "k \ Basis" + shows "negligible {x. x\k = c}" + unfolding negligible_def has_integral +proof (clarify, goal_cases) + case (1 a b e) + from this and k obtain d where d: "0 < d" "content (cbox a b \ {x. \x \ k - c\ \ d}) < e" + by (rule content_doublesplit) + let ?i = "indicator {x::'a. x\k = c} :: 'a\real" + show ?case + apply (rule_tac x="\x. ball x d" in exI) + apply rule + apply (rule gauge_ball) + apply (rule d) + proof (rule, rule) + fix p + assume p: "p tagged_division_of (cbox a b) \ (\x. ball x d) fine p" + have *: "(\(x, ka)\p. content ka *\<^sub>R ?i x) = + (\(x, ka)\p. content (ka \ {x. \x\k - c\ \ d}) *\<^sub>R ?i x)" + apply (rule setsum.cong) + apply (rule refl) + unfolding split_paired_all real_scaleR_def mult_cancel_right split_conv + apply cases + apply (rule disjI1) + apply assumption + apply (rule disjI2) + proof - + fix x l + assume as: "(x, l) \ p" "?i x \ 0" + then have xk: "x\k = c" + unfolding indicator_def + apply - + apply (rule ccontr) + apply auto + done + show "content l = content (l \ {x. \x \ k - c\ \ d})" + apply (rule arg_cong[where f=content]) + apply (rule set_eqI) + apply rule + apply rule + unfolding mem_Collect_eq + proof - + fix y + assume y: "y \ l" + note p[THEN conjunct2,unfolded fine_def,rule_format,OF as(1),unfolded split_conv] + note this[unfolded subset_eq mem_ball dist_norm,rule_format,OF y] + note le_less_trans[OF Basis_le_norm[OF k] this] + then show "\y \ k - c\ \ d" + unfolding inner_simps xk by auto + qed auto + qed + note p'= tagged_division_ofD[OF p[THEN conjunct1]] and p''=division_of_tagged_division[OF p[THEN conjunct1]] + show "norm ((\(x, ka)\p. content ka *\<^sub>R ?i x) - 0) < e" + unfolding diff_0_right * + unfolding real_scaleR_def real_norm_def + apply (subst abs_of_nonneg) + apply (rule setsum_nonneg) + apply rule + unfolding split_paired_all split_conv + apply (rule mult_nonneg_nonneg) + apply (drule p'(4)) + apply (erule exE)+ + apply(rule_tac b=b in back_subst) + prefer 2 + apply (subst(asm) eq_commute) + apply assumption + apply (subst interval_doublesplit[OF k]) + apply (rule content_pos_le) + apply (rule indicator_pos_le) + proof - + have "(\(x, ka)\p. content (ka \ {x. \x \ k - c\ \ d}) * ?i x) \ + (\(x, ka)\p. content (ka \ {x. \x \ k - c\ \ d}))" + apply (rule setsum_mono) + unfolding split_paired_all split_conv + apply (rule mult_right_le_one_le) + apply (drule p'(4)) + apply (auto simp add:interval_doublesplit[OF k]) + done + also have "\ < e" + proof (subst setsum.over_tagged_division_lemma[OF p[THEN conjunct1]], goal_cases) + case prems: (1 u v) + have "content (cbox u v \ {x. \x \ k - c\ \ d}) \ content (cbox u v)" + unfolding interval_doublesplit[OF k] + apply (rule content_subset) + unfolding interval_doublesplit[symmetric,OF k] + apply auto + done + then show ?case + unfolding prems interval_doublesplit[OF k] + by (blast intro: antisym) + next + have "(\l\snd ` p. content (l \ {x. \x \ k - c\ \ d})) = + setsum content ((\l. l \ {x. \x \ k - c\ \ d})`{l\snd ` p. l \ {x. \x \ k - c\ \ d} \ {}})" + proof (subst (2) setsum.reindex_nontrivial) + fix x y assume "x \ {l \ snd ` p. l \ {x. \x \ k - c\ \ d} \ {}}" "y \ {l \ snd ` p. l \ {x. \x \ k - c\ \ d} \ {}}" + "x \ y" and eq: "x \ {x. \x \ k - c\ \ d} = y \ {x. \x \ k - c\ \ d}" + then obtain x' y' where "(x', x) \ p" "x \ {x. \x \ k - c\ \ d} \ {}" "(y', y) \ p" "y \ {x. \x \ k - c\ \ d} \ {}" + by (auto) + from p'(5)[OF \(x', x) \ p\ \(y', y) \ p\] \x \ y\ have "interior (x \ y) = {}" + by auto + moreover have "interior ((x \ {x. \x \ k - c\ \ d}) \ (y \ {x. \x \ k - c\ \ d})) \ interior (x \ y)" + by (auto intro: interior_mono) + ultimately have "interior (x \ {x. \x \ k - c\ \ d}) = {}" + by (auto simp: eq) + then show "content (x \ {x. \x \ k - c\ \ d}) = 0" + using p'(4)[OF \(x', x) \ p\] by (auto simp: interval_doublesplit[OF k] content_eq_0_interior simp del: interior_Int) + qed (insert p'(1), auto intro!: setsum.mono_neutral_right) + also have "\ \ norm (\l\(\l. l \ {x. \x \ k - c\ \ d})`{l\snd ` p. l \ {x. \x \ k - c\ \ d} \ {}}. content l *\<^sub>R 1::real)" + by simp + also have "\ \ 1 * content (cbox a b \ {x. \x \ k - c\ \ d})" + using division_doublesplit[OF p'' k, unfolded interval_doublesplit[OF k]] + unfolding interval_doublesplit[OF k] by (intro dsum_bound) auto + also have "\ < e" + using d(2) by simp + finally show "(\ka\snd ` p. content (ka \ {x. \x \ k - c\ \ d})) < e" . + qed + finally show "(\(x, ka)\p. content (ka \ {x. \x \ k - c\ \ d}) * ?i x) < e" . + qed + qed +qed + + +subsection \A technical lemma about "refinement" of division.\ + +lemma tagged_division_finer: + fixes p :: "('a::euclidean_space \ ('a::euclidean_space set)) set" + assumes "p tagged_division_of (cbox a b)" + and "gauge d" + obtains q where "q tagged_division_of (cbox a b)" + and "d fine q" + and "\(x,k) \ p. k \ d(x) \ (x,k) \ q" +proof - + let ?P = "\p. p tagged_partial_division_of (cbox a b) \ gauge d \ + (\q. q tagged_division_of (\{k. \x. (x,k) \ p}) \ d fine q \ + (\(x,k) \ p. k \ d(x) \ (x,k) \ q))" + { + have *: "finite p" "p tagged_partial_division_of (cbox a b)" + using assms(1) + unfolding tagged_division_of_def + by auto + presume "\p. finite p \ ?P p" + from this[rule_format,OF * assms(2)] guess q .. note q=this + then show ?thesis + apply - + apply (rule that[of q]) + unfolding tagged_division_ofD[OF assms(1)] + apply auto + done + } + fix p :: "('a::euclidean_space \ ('a::euclidean_space set)) set" + assume as: "finite p" + show "?P p" + apply rule + apply rule + using as + proof (induct p) + case empty + show ?case + apply (rule_tac x="{}" in exI) + unfolding fine_def + apply auto + done + next + case (insert xk p) + guess x k using surj_pair[of xk] by (elim exE) note xk=this + note tagged_partial_division_subset[OF insert(4) subset_insertI] + from insert(3)[OF this insert(5)] guess q1 .. note q1 = conjunctD3[OF this] + have *: "\{l. \y. (y,l) \ insert xk p} = k \ \{l. \y. (y,l) \ p}" + unfolding xk by auto + note p = tagged_partial_division_ofD[OF insert(4)] + from p(4)[unfolded xk, OF insertI1] guess u v by (elim exE) note uv=this + + have "finite {k. \x. (x, k) \ p}" + apply (rule finite_subset[of _ "snd ` p"]) + using p + apply safe + apply (metis image_iff snd_conv) + apply auto + done + then have int: "interior (cbox u v) \ interior (\{k. \x. (x, k) \ p}) = {}" + apply (rule inter_interior_unions_intervals) + apply (rule open_interior) + apply (rule_tac[!] ballI) + unfolding mem_Collect_eq + apply (erule_tac[!] exE) + apply (drule p(4)[OF insertI2]) + apply assumption + apply (rule p(5)) + unfolding uv xk + apply (rule insertI1) + apply (rule insertI2) + apply assumption + using insert(2) + unfolding uv xk + apply auto + done + show ?case + proof (cases "cbox u v \ d x") + case True + then show ?thesis + apply (rule_tac x="{(x,cbox u v)} \ q1" in exI) + apply rule + unfolding * uv + apply (rule tagged_division_union) + apply (rule tagged_division_of_self) + apply (rule p[unfolded xk uv] insertI1)+ + apply (rule q1) + apply (rule int) + apply rule + apply (rule fine_union) + apply (subst fine_def) + defer + apply (rule q1) + unfolding Ball_def split_paired_All split_conv + apply rule + apply rule + apply rule + apply rule + apply (erule insertE) + apply (simp add: uv xk) + apply (rule UnI2) + apply (drule q1(3)[rule_format]) + unfolding xk uv + apply auto + done + next + case False + from fine_division_exists[OF assms(2), of u v] guess q2 . note q2=this + show ?thesis + apply (rule_tac x="q2 \ q1" in exI) + apply rule + unfolding * uv + apply (rule tagged_division_union q2 q1 int fine_union)+ + unfolding Ball_def split_paired_All split_conv + apply rule + apply (rule fine_union) + apply (rule q1 q2)+ + apply rule + apply rule + apply rule + apply rule + apply (erule insertE) + apply (rule UnI2) + apply (simp add: False uv xk) + apply (drule q1(3)[rule_format]) + using False + unfolding xk uv + apply auto + done + qed + qed +qed + + +subsection \Hence the main theorem about negligible sets.\ + +lemma finite_product_dependent: + assumes "finite s" + and "\x. x \ s \ finite (t x)" + shows "finite {(i, j) |i j. i \ s \ j \ t i}" + using assms +proof induct + case (insert x s) + have *: "{(i, j) |i j. i \ insert x s \ j \ t i} = + (\y. (x,y)) ` (t x) \ {(i, j) |i j. i \ s \ j \ t i}" by auto + show ?case + unfolding * + apply (rule finite_UnI) + using insert + apply auto + done +qed auto + +lemma sum_sum_product: + assumes "finite s" + and "\i\s. finite (t i)" + shows "setsum (\i. setsum (x i) (t i)::real) s = + setsum (\(i,j). x i j) {(i,j) | i j. i \ s \ j \ t i}" + using assms +proof induct + case (insert a s) + have *: "{(i, j) |i j. i \ insert a s \ j \ t i} = + (\y. (a,y)) ` (t a) \ {(i, j) |i j. i \ s \ j \ t i}" by auto + show ?case + unfolding * + apply (subst setsum.union_disjoint) + unfolding setsum.insert[OF insert(1-2)] + prefer 4 + apply (subst insert(3)) + unfolding add_right_cancel + proof - + show "setsum (x a) (t a) = (\(xa, y)\ Pair a ` t a. x xa y)" + apply (subst setsum.reindex) + unfolding inj_on_def + apply auto + done + show "finite {(i, j) |i j. i \ s \ j \ t i}" + apply (rule finite_product_dependent) + using insert + apply auto + done + qed (insert insert, auto) +qed auto + +lemma has_integral_negligible: + fixes f :: "'b::euclidean_space \ 'a::real_normed_vector" + assumes "negligible s" + and "\x\(t - s). f x = 0" + shows "(f has_integral 0) t" +proof - + presume P: "\f::'b::euclidean_space \ 'a. + \a b. \x. x \ s \ f x = 0 \ (f has_integral 0) (cbox a b)" + let ?f = "(\x. if x \ t then f x else 0)" + show ?thesis + apply (rule_tac f="?f" in has_integral_eq) + unfolding if_P + apply (rule refl) + apply (subst has_integral_alt) + apply cases + apply (subst if_P, assumption) + unfolding if_not_P + proof - + assume "\a b. t = cbox a b" + then guess a b apply - by (erule exE)+ note t = this + show "(?f has_integral 0) t" + unfolding t + apply (rule P) + using assms(2) + unfolding t + apply auto + done + next + show "\e>0. \B>0. \a b. ball 0 B \ cbox a b \ + (\z. ((\x. if x \ t then ?f x else 0) has_integral z) (cbox a b) \ norm (z - 0) < e)" + apply safe + apply (rule_tac x=1 in exI) + apply rule + apply (rule zero_less_one) + apply safe + apply (rule_tac x=0 in exI) + apply rule + apply (rule P) + using assms(2) + apply auto + done + qed +next + fix f :: "'b \ 'a" + fix a b :: 'b + assume assm: "\x. x \ s \ f x = 0" + show "(f has_integral 0) (cbox a b)" + unfolding has_integral + proof (safe, goal_cases) + case prems: (1 e) + then have "\n. e / 2 / ((real n+1) * (2 ^ n)) > 0" + apply - + apply (rule divide_pos_pos) + defer + apply (rule mult_pos_pos) + apply (auto simp add:field_simps) + done + note assms(1)[unfolded negligible_def has_integral,rule_format,OF this,of a b] + note allI[OF this,of "\x. x"] + from choice[OF this] guess d .. note d=conjunctD2[OF this[rule_format]] + show ?case + apply (rule_tac x="\x. d (nat \norm (f x)\) x" in exI) + proof safe + show "gauge (\x. d (nat \norm (f x)\) x)" + using d(1) unfolding gauge_def by auto + fix p + assume as: "p tagged_division_of (cbox a b)" "(\x. d (nat \norm (f x)\) x) fine p" + let ?goal = "norm ((\(x, k)\p. content k *\<^sub>R f x) - 0) < e" + { + presume "p \ {} \ ?goal" + then show ?goal + apply (cases "p = {}") + using prems + apply auto + done + } + assume as': "p \ {}" + from real_arch_simple[of "Max((\(x,k). norm(f x)) ` p)"] guess N .. + then have N: "\x\(\(x, k). norm (f x)) ` p. x \ real N" + by (meson Max_ge as(1) dual_order.trans finite_imageI tagged_division_of_finite) + have "\i. \q. q tagged_division_of (cbox a b) \ (d i) fine q \ (\(x, k)\p. k \ (d i) x \ (x, k) \ q)" + by (auto intro: tagged_division_finer[OF as(1) d(1)]) + from choice[OF this] guess q .. note q=conjunctD3[OF this[rule_format]] + have *: "\i. (\(x, k)\q i. content k *\<^sub>R indicator s x) \ (0::real)" + apply (rule setsum_nonneg) + apply safe + unfolding real_scaleR_def + apply (drule tagged_division_ofD(4)[OF q(1)]) + apply (auto intro: mult_nonneg_nonneg) + done + have **: "finite s \ finite t \ (\(x,y) \ t. (0::real) \ g(x,y)) \ + (\y\s. \x. (x,y) \ t \ f(y) \ g(x,y)) \ setsum f s \ setsum g t" for f g s t + apply (rule setsum_le_included[of s t g snd f]) + prefer 4 + apply safe + apply (erule_tac x=x in ballE) + apply (erule exE) + apply (rule_tac x="(xa,x)" in bexI) + apply auto + done + have "norm ((\(x, k)\p. content k *\<^sub>R f x) - 0) \ setsum (\i. (real i + 1) * + norm (setsum (\(x,k). content k *\<^sub>R indicator s x :: real) (q i))) {..N+1}" + unfolding real_norm_def setsum_right_distrib abs_of_nonneg[OF *] diff_0_right + apply (rule order_trans) + apply (rule norm_setsum) + apply (subst sum_sum_product) + prefer 3 + proof (rule **, safe) + show "finite {(i, j) |i j. i \ {..N + 1} \ j \ q i}" + apply (rule finite_product_dependent) + using q + apply auto + done + fix i a b + assume as'': "(a, b) \ q i" + show "0 \ (real i + 1) * (content b *\<^sub>R indicator s a)" + unfolding real_scaleR_def + using tagged_division_ofD(4)[OF q(1) as''] + by (auto intro!: mult_nonneg_nonneg) + next + fix i :: nat + show "finite (q i)" + using q by auto + next + fix x k + assume xk: "(x, k) \ p" + define n where "n = nat \norm (f x)\" + have *: "norm (f x) \ (\(x, k). norm (f x)) ` p" + using xk by auto + have nfx: "real n \ norm (f x)" "norm (f x) \ real n + 1" + unfolding n_def by auto + then have "n \ {0..N + 1}" + using N[rule_format,OF *] by auto + moreover + note as(2)[unfolded fine_def,rule_format,OF xk,unfolded split_conv] + note q(3)[rule_format,OF xk,unfolded split_conv,rule_format,OF this] + note this[unfolded n_def[symmetric]] + moreover + have "norm (content k *\<^sub>R f x) \ (real n + 1) * (content k * indicator s x)" + proof (cases "x \ s") + case False + then show ?thesis + using assm by auto + next + case True + have *: "content k \ 0" + using tagged_division_ofD(4)[OF as(1) xk] by auto + moreover + have "content k * norm (f x) \ content k * (real n + 1)" + apply (rule mult_mono) + using nfx * + apply auto + done + ultimately + show ?thesis + unfolding abs_mult + using nfx True + by (auto simp add: field_simps) + qed + ultimately show "\y. (y, x, k) \ {(i, j) |i j. i \ {..N + 1} \ j \ q i} \ norm (content k *\<^sub>R f x) \ + (real y + 1) * (content k *\<^sub>R indicator s x)" + apply (rule_tac x=n in exI) + apply safe + apply (rule_tac x=n in exI) + apply (rule_tac x="(x,k)" in exI) + apply safe + apply auto + done + qed (insert as, auto) + also have "\ \ setsum (\i. e / 2 / 2 ^ i) {..N+1}" + proof (rule setsum_mono, goal_cases) + case (1 i) + then show ?case + apply (subst mult.commute, subst pos_le_divide_eq[symmetric]) + using d(2)[rule_format, of "q i" i] + using q[rule_format] + apply (auto simp add: field_simps) + done + qed + also have "\ < e * inverse 2 * 2" + unfolding divide_inverse setsum_right_distrib[symmetric] + apply (rule mult_strict_left_mono) + unfolding power_inverse [symmetric] lessThan_Suc_atMost[symmetric] + apply (subst geometric_sum) + using prems + apply auto + done + finally show "?goal" by auto + qed + qed +qed + +lemma has_integral_spike: + fixes f :: "'b::euclidean_space \ 'a::real_normed_vector" + assumes "negligible s" + and "(\x\(t - s). g x = f x)" + and "(f has_integral y) t" + shows "(g has_integral y) t" +proof - + { + fix a b :: 'b + fix f g :: "'b \ 'a" + fix y :: 'a + assume as: "\x \ cbox a b - s. g x = f x" "(f has_integral y) (cbox a b)" + have "((\x. f x + (g x - f x)) has_integral (y + 0)) (cbox a b)" + apply (rule has_integral_add[OF as(2)]) + apply (rule has_integral_negligible[OF assms(1)]) + using as + apply auto + done + then have "(g has_integral y) (cbox a b)" + by auto + } note * = this + show ?thesis + apply (subst has_integral_alt) + using assms(2-) + apply - + apply (rule cond_cases) + apply safe + apply (rule *) + apply assumption+ + apply (subst(asm) has_integral_alt) + unfolding if_not_P + apply (erule_tac x=e in allE) + apply safe + apply (rule_tac x=B in exI) + apply safe + apply (erule_tac x=a in allE) + apply (erule_tac x=b in allE) + apply safe + apply (rule_tac x=z in exI) + apply safe + apply (rule *[where fa2="\x. if x\t then f x else 0"]) + apply auto + done +qed + +lemma has_integral_spike_eq: + assumes "negligible s" + and "\x\(t - s). g x = f x" + shows "((f has_integral y) t \ (g has_integral y) t)" + apply rule + apply (rule_tac[!] has_integral_spike[OF assms(1)]) + using assms(2) + apply auto + done + +lemma integrable_spike: + assumes "negligible s" + and "\x\(t - s). g x = f x" + and "f integrable_on t" + shows "g integrable_on t" + using assms + unfolding integrable_on_def + apply - + apply (erule exE) + apply rule + apply (rule has_integral_spike) + apply fastforce+ + done + +lemma integral_spike: + assumes "negligible s" + and "\x\(t - s). g x = f x" + shows "integral t f = integral t g" + using has_integral_spike_eq[OF assms] by (simp add: integral_def integrable_on_def) + + +subsection \Some other trivialities about negligible sets.\ + +lemma negligible_subset[intro]: + assumes "negligible s" + and "t \ s" + shows "negligible t" + unfolding negligible_def +proof (safe, goal_cases) + case (1 a b) + show ?case + using assms(1)[unfolded negligible_def,rule_format,of a b] + apply - + apply (rule has_integral_spike[OF assms(1)]) + defer + apply assumption + using assms(2) + unfolding indicator_def + apply auto + done +qed + +lemma negligible_diff[intro?]: + assumes "negligible s" + shows "negligible (s - t)" + using assms by auto + +lemma negligible_Int: + assumes "negligible s \ negligible t" + shows "negligible (s \ t)" + using assms by auto + +lemma negligible_Un: + assumes "negligible s" + and "negligible t" + shows "negligible (s \ t)" + unfolding negligible_def +proof (safe, goal_cases) + case (1 a b) + note assm = assms[unfolded negligible_def,rule_format,of a b] + then show ?case + apply (subst has_integral_spike_eq[OF assms(2)]) + defer + apply assumption + unfolding indicator_def + apply auto + done +qed + +lemma negligible_Un_eq[simp]: "negligible (s \ t) \ negligible s \ negligible t" + using negligible_Un by auto + +lemma negligible_sing[intro]: "negligible {a::'a::euclidean_space}" + using negligible_standard_hyperplane[OF SOME_Basis, of "a \ (SOME i. i \ Basis)"] by auto + +lemma negligible_insert[simp]: "negligible (insert a s) \ negligible s" + apply (subst insert_is_Un) + unfolding negligible_Un_eq + apply auto + done + +lemma negligible_empty[iff]: "negligible {}" + by auto + +lemma negligible_finite[intro]: + assumes "finite s" + shows "negligible s" + using assms by (induct s) auto + +lemma negligible_Union[intro]: + assumes "finite s" + and "\t\s. negligible t" + shows "negligible(\s)" + using assms by induct auto + +lemma negligible: + "negligible s \ (\t::('a::euclidean_space) set. ((indicator s::'a\real) has_integral 0) t)" + apply safe + defer + apply (subst negligible_def) +proof - + fix t :: "'a set" + assume as: "negligible s" + have *: "(\x. if x \ s \ t then 1 else 0) = (\x. if x\t then if x\s then 1 else 0 else 0)" + by auto + show "((indicator s::'a\real) has_integral 0) t" + apply (subst has_integral_alt) + apply cases + apply (subst if_P,assumption) + unfolding if_not_P + apply safe + apply (rule as[unfolded negligible_def,rule_format]) + apply (rule_tac x=1 in exI) + apply safe + apply (rule zero_less_one) + apply (rule_tac x=0 in exI) + using negligible_subset[OF as,of "s \ t"] + unfolding negligible_def indicator_def [abs_def] + unfolding * + apply auto + done +qed auto + + +subsection \Finite case of the spike theorem is quite commonly needed.\ + +lemma has_integral_spike_finite: + assumes "finite s" + and "\x\t-s. g x = f x" + and "(f has_integral y) t" + shows "(g has_integral y) t" + apply (rule has_integral_spike) + using assms + apply auto + done + +lemma has_integral_spike_finite_eq: + assumes "finite s" + and "\x\t-s. g x = f x" + shows "((f has_integral y) t \ (g has_integral y) t)" + apply rule + apply (rule_tac[!] has_integral_spike_finite) + using assms + apply auto + done + +lemma integrable_spike_finite: + assumes "finite s" + and "\x\t-s. g x = f x" + and "f integrable_on t" + shows "g integrable_on t" + using assms + unfolding integrable_on_def + apply safe + apply (rule_tac x=y in exI) + apply (rule has_integral_spike_finite) + apply auto + done + + +subsection \In particular, the boundary of an interval is negligible.\ + +lemma negligible_frontier_interval: "negligible(cbox (a::'a::euclidean_space) b - box a b)" +proof - + let ?A = "\((\k. {x. x\k = a\k} \ {x::'a. x\k = b\k}) ` Basis)" + have "cbox a b - box a b \ ?A" + apply rule unfolding Diff_iff mem_box + apply simp + apply(erule conjE bexE)+ + apply(rule_tac x=i in bexI) + apply auto + done + then show ?thesis + apply - + apply (rule negligible_subset[of ?A]) + apply (rule negligible_Union[OF finite_imageI]) + apply auto + done +qed + +lemma has_integral_spike_interior: + assumes "\x\box a b. g x = f x" + and "(f has_integral y) (cbox a b)" + shows "(g has_integral y) (cbox a b)" + apply (rule has_integral_spike[OF negligible_frontier_interval _ assms(2)]) + using assms(1) + apply auto + done + +lemma has_integral_spike_interior_eq: + assumes "\x\box a b. g x = f x" + shows "(f has_integral y) (cbox a b) \ (g has_integral y) (cbox a b)" + apply rule + apply (rule_tac[!] has_integral_spike_interior) + using assms + apply auto + done + +lemma integrable_spike_interior: + assumes "\x\box a b. g x = f x" + and "f integrable_on cbox a b" + shows "g integrable_on cbox a b" + using assms + unfolding integrable_on_def + using has_integral_spike_interior[OF assms(1)] + by auto + + +subsection \Integrability of continuous functions.\ + +lemma operative_approximable: + fixes f :: "'b::euclidean_space \ 'a::banach" + assumes "0 \ e" + shows "comm_monoid.operative op \ True (\i. \g. (\x\i. norm (f x - g (x::'b)) \ e) \ g integrable_on i)" + unfolding comm_monoid.operative_def[OF comm_monoid_and] +proof safe + fix a b :: 'b + show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" + if "content (cbox a b) = 0" + apply (rule_tac x=f in exI) + using assms that + apply (auto intro!: integrable_on_null) + done + { + fix c g + fix k :: 'b + assume as: "\x\cbox a b. norm (f x - g x) \ e" "g integrable_on cbox a b" + assume k: "k \ Basis" + show "\g. (\x\cbox a b \ {x. x \ k \ c}. norm (f x - g x) \ e) \ g integrable_on cbox a b \ {x. x \ k \ c}" + "\g. (\x\cbox a b \ {x. c \ x \ k}. norm (f x - g x) \ e) \ g integrable_on cbox a b \ {x. c \ x \ k}" + apply (rule_tac[!] x=g in exI) + using as(1) integrable_split[OF as(2) k] + apply auto + done + } + fix c k g1 g2 + assume as: "\x\cbox a b \ {x. x \ k \ c}. norm (f x - g1 x) \ e" "g1 integrable_on cbox a b \ {x. x \ k \ c}" + "\x\cbox a b \ {x. c \ x \ k}. norm (f x - g2 x) \ e" "g2 integrable_on cbox a b \ {x. c \ x \ k}" + assume k: "k \ Basis" + let ?g = "\x. if x\k = c then f x else if x\k \ c then g1 x else g2 x" + show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" + apply (rule_tac x="?g" in exI) + apply safe + proof goal_cases + case (1 x) + then show ?case + apply - + apply (cases "x\k=c") + apply (case_tac "x\k < c") + using as assms + apply auto + done + next + case 2 + presume "?g integrable_on cbox a b \ {x. x \ k \ c}" + and "?g integrable_on cbox a b \ {x. x \ k \ c}" + then guess h1 h2 unfolding integrable_on_def by auto + from has_integral_split[OF this k] show ?case + unfolding integrable_on_def by auto + next + show "?g integrable_on cbox a b \ {x. x \ k \ c}" "?g integrable_on cbox a b \ {x. x \ k \ c}" + apply(rule_tac[!] integrable_spike[OF negligible_standard_hyperplane[of k c]]) + using k as(2,4) + apply auto + done + qed +qed + +lemma comm_monoid_set_F_and: "comm_monoid_set.F op \ True f s \ (finite s \ (\x\s. f x))" +proof - + interpret bool: comm_monoid_set "op \" True + proof qed auto + show ?thesis + by (induction s rule: infinite_finite_induct) auto +qed + +lemma approximable_on_division: + fixes f :: "'b::euclidean_space \ 'a::banach" + assumes "0 \ e" + and "d division_of (cbox a b)" + and "\i\d. \g. (\x\i. norm (f x - g x) \ e) \ g integrable_on i" + obtains g where "\x\cbox a b. norm (f x - g x) \ e" "g integrable_on cbox a b" +proof - + note * = comm_monoid_set.operative_division[OF comm_monoid_set_and operative_approximable[OF assms(1)] assms(2)] + from assms(3) this[unfolded comm_monoid_set_F_and, of f] division_of_finite[OF assms(2)] + guess g by auto + then show thesis + apply - + apply (rule that[of g]) + apply auto + done +qed + +lemma integrable_continuous: + fixes f :: "'b::euclidean_space \ 'a::banach" + assumes "continuous_on (cbox a b) f" + shows "f integrable_on cbox a b" +proof (rule integrable_uniform_limit, safe) + fix e :: real + assume e: "e > 0" + from compact_uniformly_continuous[OF assms compact_cbox,unfolded uniformly_continuous_on_def,rule_format,OF e] guess d .. + note d=conjunctD2[OF this,rule_format] + from fine_division_exists[OF gauge_ball[OF d(1)], of a b] guess p . note p=this + note p' = tagged_division_ofD[OF p(1)] + have *: "\i\snd ` p. \g. (\x\i. norm (f x - g x) \ e) \ g integrable_on i" + proof (safe, unfold snd_conv) + fix x l + assume as: "(x, l) \ p" + from p'(4)[OF this] guess a b by (elim exE) note l=this + show "\g. (\x\l. norm (f x - g x) \ e) \ g integrable_on l" + apply (rule_tac x="\y. f x" in exI) + proof safe + show "(\y. f x) integrable_on l" + unfolding integrable_on_def l + apply rule + apply (rule has_integral_const) + done + fix y + assume y: "y \ l" + note fineD[OF p(2) as,unfolded subset_eq,rule_format,OF this] + note d(2)[OF _ _ this[unfolded mem_ball]] + then show "norm (f y - f x) \ e" + using y p'(2-3)[OF as] unfolding dist_norm l norm_minus_commute by fastforce + qed + qed + from e have "e \ 0" + by auto + from approximable_on_division[OF this division_of_tagged_division[OF p(1)] *] guess g . + then show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" + by auto +qed + +lemma integrable_continuous_real: + fixes f :: "real \ 'a::banach" + assumes "continuous_on {a .. b} f" + shows "f integrable_on {a .. b}" + by (metis assms box_real(2) integrable_continuous) + +subsection \Specialization of additivity to one dimension.\ + +subsection \Special case of additivity we need for the FTC.\ + +lemma additive_tagged_division_1: + fixes f :: "real \ 'a::real_normed_vector" + assumes "a \ b" + and "p tagged_division_of {a..b}" + shows "setsum (\(x,k). f(Sup k) - f(Inf k)) p = f b - f a" +proof - + let ?f = "(\k::(real) set. if k = {} then 0 else f(interval_upperbound k) - f(interval_lowerbound k))" + have ***: "\i\Basis. a \ i \ b \ i" + using assms by auto + have *: "add.operative ?f" + unfolding add.operative_1_lt box_eq_empty + by auto + have **: "cbox a b \ {}" + using assms(1) by auto + note setsum.operative_tagged_division[OF * assms(2)[simplified box_real[symmetric]]] + note * = this[unfolded if_not_P[OF **] interval_bounds[OF ***],symmetric] + show ?thesis + unfolding * + apply (rule setsum.cong) + unfolding split_paired_all split_conv + using assms(2) + apply auto + done +qed + + +subsection \A useful lemma allowing us to factor out the content size.\ + +lemma has_integral_factor_content: + "(f has_integral i) (cbox a b) \ + (\e>0. \d. gauge d \ (\p. p tagged_division_of (cbox a b) \ d fine p \ + norm (setsum (\(x,k). content k *\<^sub>R f x) p - i) \ e * content (cbox a b)))" +proof (cases "content (cbox a b) = 0") + case True + show ?thesis + unfolding has_integral_null_eq[OF True] + apply safe + apply (rule, rule, rule gauge_trivial, safe) + unfolding setsum_content_null[OF True] True + defer + apply (erule_tac x=1 in allE) + apply safe + defer + apply (rule fine_division_exists[of _ a b]) + apply assumption + apply (erule_tac x=p in allE) + unfolding setsum_content_null[OF True] + apply auto + done +next + case False + note F = this[unfolded content_lt_nz[symmetric]] + let ?P = "\e opp. \d. gauge d \ + (\p. p tagged_division_of (cbox a b) \ d fine p \ opp (norm ((\(x, k)\p. content k *\<^sub>R f x) - i)) e)" + show ?thesis + apply (subst has_integral) + proof safe + fix e :: real + assume e: "e > 0" + { + assume "\e>0. ?P e op <" + then show "?P (e * content (cbox a b)) op \" + apply (erule_tac x="e * content (cbox a b)" in allE) + apply (erule impE) + defer + apply (erule exE,rule_tac x=d in exI) + using F e + apply (auto simp add:field_simps) + done + } + { + assume "\e>0. ?P (e * content (cbox a b)) op \" + then show "?P e op <" + apply (erule_tac x="e / 2 / content (cbox a b)" in allE) + apply (erule impE) + defer + apply (erule exE,rule_tac x=d in exI) + using F e + apply (auto simp add: field_simps) + done + } + qed +qed + +lemma has_integral_factor_content_real: + "(f has_integral i) {a .. b::real} \ + (\e>0. \d. gauge d \ (\p. p tagged_division_of {a .. b} \ d fine p \ + norm (setsum (\(x,k). content k *\<^sub>R f x) p - i) \ e * content {a .. b} ))" + unfolding box_real[symmetric] + by (rule has_integral_factor_content) + + +subsection \Fundamental theorem of calculus.\ + +lemma interval_bounds_real: + fixes q b :: real + assumes "a \ b" + shows "Sup {a..b} = b" + and "Inf {a..b} = a" + using assms by auto + +lemma fundamental_theorem_of_calculus: + fixes f :: "real \ 'a::banach" + assumes "a \ b" + and "\x\{a .. b}. (f has_vector_derivative f' x) (at x within {a .. b})" + shows "(f' has_integral (f b - f a)) {a .. b}" + unfolding has_integral_factor_content box_real[symmetric] +proof safe + fix e :: real + assume e: "e > 0" + note assm = assms(2)[unfolded has_vector_derivative_def has_derivative_within_alt] + have *: "\P Q. \x\{a .. b}. P x \ (\e>0. \d>0. Q x e d) \ \x. \(d::real)>0. x\{a .. b} \ Q x e d" + using e by blast + note this[OF assm,unfolded gauge_existence_lemma] + from choice[OF this,unfolded Ball_def[symmetric]] guess d .. + note d=conjunctD2[OF this[rule_format],rule_format] + show "\d. gauge d \ (\p. p tagged_division_of (cbox a b) \ d fine p \ + norm ((\(x, k)\p. content k *\<^sub>R f' x) - (f b - f a)) \ e * content (cbox a b))" + apply (rule_tac x="\x. ball x (d x)" in exI) + apply safe + apply (rule gauge_ball_dependent) + apply rule + apply (rule d(1)) + proof - + fix p + assume as: "p tagged_division_of cbox a b" "(\x. ball x (d x)) fine p" + show "norm ((\(x, k)\p. content k *\<^sub>R f' x) - (f b - f a)) \ e * content (cbox a b)" + unfolding content_real[OF assms(1), simplified box_real[symmetric]] additive_tagged_division_1[OF assms(1) as(1)[simplified box_real],of f,symmetric] + unfolding additive_tagged_division_1[OF assms(1) as(1)[simplified box_real],of "\x. x",symmetric] + unfolding setsum_right_distrib + defer + unfolding setsum_subtractf[symmetric] + proof (rule setsum_norm_le,safe) + fix x k + assume "(x, k) \ p" + note xk = tagged_division_ofD(2-4)[OF as(1) this] + from this(3) guess u v by (elim exE) note k=this + have *: "u \ v" + using xk unfolding k by auto + have ball: "\xa\k. xa \ ball x (d x)" + using as(2)[unfolded fine_def,rule_format,OF \(x,k)\p\,unfolded split_conv subset_eq] . + have "norm ((v - u) *\<^sub>R f' x - (f v - f u)) \ + norm (f u - f x - (u - x) *\<^sub>R f' x) + norm (f v - f x - (v - x) *\<^sub>R f' x)" + apply (rule order_trans[OF _ norm_triangle_ineq4]) + apply (rule eq_refl) + apply (rule arg_cong[where f=norm]) + unfolding scaleR_diff_left + apply (auto simp add:algebra_simps) + done + also have "\ \ e * norm (u - x) + e * norm (v - x)" + apply (rule add_mono) + apply (rule d(2)[of "x" "u",unfolded o_def]) + prefer 4 + apply (rule d(2)[of "x" "v",unfolded o_def]) + using ball[rule_format,of u] ball[rule_format,of v] + using xk(1-2) + unfolding k subset_eq + apply (auto simp add:dist_real_def) + done + also have "\ \ e * (Sup k - Inf k)" + unfolding k interval_bounds_real[OF *] + using xk(1) + unfolding k + by (auto simp add: dist_real_def field_simps) + finally show "norm (content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) \ + e * (Sup k - Inf k)" + unfolding box_real k interval_bounds_real[OF *] content_real[OF *] + interval_upperbound_real interval_lowerbound_real + . + qed + qed +qed + +lemma ident_has_integral: + fixes a::real + assumes "a \ b" + shows "((\x. x) has_integral (b\<^sup>2 - a\<^sup>2) / 2) {a..b}" +proof - + have "((\x. x) has_integral inverse 2 * b\<^sup>2 - inverse 2 * a\<^sup>2) {a..b}" + apply (rule fundamental_theorem_of_calculus [OF assms], clarify) + unfolding power2_eq_square + by (rule derivative_eq_intros | simp)+ + then show ?thesis + by (simp add: field_simps) +qed + +lemma integral_ident [simp]: + fixes a::real + assumes "a \ b" + shows "integral {a..b} (\x. x) = (if a \ b then (b\<^sup>2 - a\<^sup>2) / 2 else 0)" +using ident_has_integral integral_unique by fastforce + +lemma ident_integrable_on: + fixes a::real + shows "(\x. x) integrable_on {a..b}" +by (metis atLeastatMost_empty_iff integrable_on_def has_integral_empty ident_has_integral) + + +subsection \Taylor series expansion\ + +lemma (in bounded_bilinear) setsum_prod_derivatives_has_vector_derivative: + assumes "p>0" + and f0: "Df 0 = f" + and Df: "\m t. m < p \ a \ t \ t \ b \ + (Df m has_vector_derivative Df (Suc m) t) (at t within {a .. b})" + and g0: "Dg 0 = g" + and Dg: "\m t. m < p \ a \ t \ t \ b \ + (Dg m has_vector_derivative Dg (Suc m) t) (at t within {a .. b})" + and ivl: "a \ t" "t \ b" + shows "((\t. \iR prod (Df i t) (Dg (p - Suc i) t)) + has_vector_derivative + prod (f t) (Dg p t) - (-1)^p *\<^sub>R prod (Df p t) (g t)) + (at t within {a .. b})" + using assms +proof cases + assume p: "p \ 1" + define p' where "p' = p - 2" + from assms p have p': "{..i. i \ p' \ Suc (Suc p' - i) = (Suc (Suc p') - i)" + by auto + let ?f = "\i. (-1) ^ i *\<^sub>R (prod (Df i t) (Dg ((p - i)) t))" + have "(\iR (prod (Df i t) (Dg (Suc (p - Suc i)) t) + + prod (Df (Suc i) t) (Dg (p - Suc i) t))) = + (\i\(Suc p'). ?f i - ?f (Suc i))" + by (auto simp: algebra_simps p'(2) numeral_2_eq_2 * lessThan_Suc_atMost) + also note setsum_telescope + finally + have "(\iR (prod (Df i t) (Dg (Suc (p - Suc i)) t) + + prod (Df (Suc i) t) (Dg (p - Suc i) t))) + = prod (f t) (Dg p t) - (- 1) ^ p *\<^sub>R prod (Df p t) (g t)" + unfolding p'[symmetric] + by (simp add: assms) + thus ?thesis + using assms + by (auto intro!: derivative_eq_intros has_vector_derivative) +qed (auto intro!: derivative_eq_intros has_vector_derivative) + +lemma + fixes f::"real\'a::banach" + assumes "p>0" + and f0: "Df 0 = f" + and Df: "\m t. m < p \ a \ t \ t \ b \ + (Df m has_vector_derivative Df (Suc m) t) (at t within {a .. b})" + and ivl: "a \ b" + defines "i \ \x. ((b - x) ^ (p - 1) / fact (p - 1)) *\<^sub>R Df p x" + shows taylor_has_integral: + "(i has_integral f b - (\iR Df i a)) {a..b}" + and taylor_integral: + "f b = (\iR Df i a) + integral {a..b} i" + and taylor_integrable: + "i integrable_on {a .. b}" +proof goal_cases + case 1 + interpret bounded_bilinear "scaleR::real\'a\'a" + by (rule bounded_bilinear_scaleR) + define g where "g s = (b - s)^(p - 1)/fact (p - 1)" for s + define Dg where [abs_def]: + "Dg n s = (if n < p then (-1)^n * (b - s)^(p - 1 - n) / fact (p - 1 - n) else 0)" for n s + have g0: "Dg 0 = g" + using \p > 0\ + by (auto simp add: Dg_def divide_simps g_def split: if_split_asm) + { + fix m + assume "p > Suc m" + hence "p - Suc m = Suc (p - Suc (Suc m))" + by auto + hence "real (p - Suc m) * fact (p - Suc (Suc m)) = fact (p - Suc m)" + by auto + } note fact_eq = this + have Dg: "\m t. m < p \ a \ t \ t \ b \ + (Dg m has_vector_derivative Dg (Suc m) t) (at t within {a .. b})" + unfolding Dg_def + by (auto intro!: derivative_eq_intros simp: has_vector_derivative_def fact_eq divide_simps) + let ?sum = "\t. \iR Dg i t *\<^sub>R Df (p - Suc i) t" + from setsum_prod_derivatives_has_vector_derivative[of _ Dg _ _ _ Df, + OF \p > 0\ g0 Dg f0 Df] + have deriv: "\t. a \ t \ t \ b \ + (?sum has_vector_derivative + g t *\<^sub>R Df p t - (- 1) ^ p *\<^sub>R Dg p t *\<^sub>R f t) (at t within {a..b})" + by auto + from fundamental_theorem_of_calculus[rule_format, OF \a \ b\ deriv] + have "(i has_integral ?sum b - ?sum a) {a .. b}" + by (simp add: i_def g_def Dg_def) + also + have one: "(- 1) ^ p' * (- 1) ^ p' = (1::real)" + and "{.. {i. p = Suc i} = {p - 1}" + for p' + using \p > 0\ + by (auto simp: power_mult_distrib[symmetric]) + then have "?sum b = f b" + using Suc_pred'[OF \p > 0\] + by (simp add: diff_eq_eq Dg_def power_0_left le_Suc_eq if_distrib + cond_application_beta setsum.If_cases f0) + also + have "{..x. p - x - 1) ` {.. (\x. p - x - 1) ` {..iR Df i a)" + by (rule setsum.reindex_cong) (auto simp add: inj_on_def Dg_def one) + finally show c: ?case . + case 2 show ?case using c integral_unique by force + case 3 show ?case using c by force +qed + + +subsection \Attempt a systematic general set of "offset" results for components.\ + +lemma gauge_modify: + assumes "(\s. open s \ open {x. f(x) \ s})" "gauge d" + shows "gauge (\x. {y. f y \ d (f x)})" + using assms + unfolding gauge_def + apply safe + defer + apply (erule_tac x="f x" in allE) + apply (erule_tac x="d (f x)" in allE) + apply auto + done + + +subsection \Only need trivial subintervals if the interval itself is trivial.\ + +lemma division_of_nontrivial: + fixes s :: "'a::euclidean_space set set" + assumes "s division_of (cbox a b)" + and "content (cbox a b) \ 0" + shows "{k. k \ s \ content k \ 0} division_of (cbox a b)" + using assms(1) + apply - +proof (induct "card s" arbitrary: s rule: nat_less_induct) + fix s::"'a set set" + assume assm: "s division_of (cbox a b)" + "\mx. m = card x \ + x division_of (cbox a b) \ {k \ x. content k \ 0} division_of (cbox a b)" + note s = division_ofD[OF assm(1)] + let ?thesis = "{k \ s. content k \ 0} division_of (cbox a b)" + { + presume *: "{k \ s. content k \ 0} \ s \ ?thesis" + show ?thesis + apply cases + defer + apply (rule *) + apply assumption + using assm(1) + apply auto + done + } + assume noteq: "{k \ s. content k \ 0} \ s" + then obtain k where k: "k \ s" "content k = 0" + by auto + from s(4)[OF k(1)] guess c d by (elim exE) note k=k this + from k have "card s > 0" + unfolding card_gt_0_iff using assm(1) by auto + then have card: "card (s - {k}) < card s" + using assm(1) k(1) + apply (subst card_Diff_singleton_if) + apply auto + done + have *: "closed (\(s - {k}))" + apply (rule closed_Union) + defer + apply rule + apply (drule DiffD1,drule s(4)) + using assm(1) + apply auto + done + have "k \ \(s - {k})" + apply safe + apply (rule *[unfolded closed_limpt,rule_format]) + unfolding islimpt_approachable + proof safe + fix x + fix e :: real + assume as: "x \ k" "e > 0" + from k(2)[unfolded k content_eq_0] guess i .. + then have i:"c\i = d\i" "i\Basis" + using s(3)[OF k(1),unfolded k] unfolding box_ne_empty by auto + then have xi: "x\i = d\i" + using as unfolding k mem_box by (metis antisym) + define y where "y = (\j\Basis. (if j = i then if c\i \ (a\i + b\i) / 2 then c\i + + min e (b\i - c\i) / 2 else c\i - min e (c\i - a\i) / 2 else x\j) *\<^sub>R j)" + show "\x'\\(s - {k}). x' \ x \ dist x' x < e" + apply (rule_tac x=y in bexI) + proof + have "d \ cbox c d" + using s(3)[OF k(1)] + unfolding k box_eq_empty mem_box + by (fastforce simp add: not_less) + then have "d \ cbox a b" + using s(2)[OF k(1)] + unfolding k + by auto + note di = this[unfolded mem_box,THEN bspec[where x=i]] + then have xyi: "y\i \ x\i" + unfolding y_def i xi + using as(2) assms(2)[unfolded content_eq_0] i(2) + by (auto elim!: ballE[of _ _ i]) + then show "y \ x" + unfolding euclidean_eq_iff[where 'a='a] using i by auto + have *: "Basis = insert i (Basis - {i})" + using i by auto + have "norm (y - x) < e + setsum (\i. 0) Basis" + apply (rule le_less_trans[OF norm_le_l1]) + apply (subst *) + apply (subst setsum.insert) + prefer 3 + apply (rule add_less_le_mono) + proof - + show "\(y - x) \ i\ < e" + using di as(2) y_def i xi by (auto simp: inner_simps) + show "(\i\Basis - {i}. \(y - x) \ i\) \ (\i\Basis. 0)" + unfolding y_def by (auto simp: inner_simps) + qed auto + then show "dist y x < e" + unfolding dist_norm by auto + have "y \ k" + unfolding k mem_box + apply rule + apply (erule_tac x=i in ballE) + using xyi k i xi + apply auto + done + moreover + have "y \ \s" + using set_rev_mp[OF as(1) s(2)[OF k(1)]] as(2) di i + unfolding s mem_box y_def + by (auto simp: field_simps elim!: ballE[of _ _ i]) + ultimately + show "y \ \(s - {k})" by auto + qed + qed + then have "\(s - {k}) = cbox a b" + unfolding s(6)[symmetric] by auto + then have "{ka \ s - {k}. content ka \ 0} division_of (cbox a b)" + apply - + apply (rule assm(2)[rule_format,OF card refl]) + apply (rule division_ofI) + defer + apply (rule_tac[1-4] s) + using assm(1) + apply auto + done + moreover + have "{ka \ s - {k}. content ka \ 0} = {k \ s. content k \ 0}" + using k by auto + ultimately show ?thesis by auto +qed + + +subsection \Integrability on subintervals.\ + +lemma operative_integrable: + fixes f :: "'b::euclidean_space \ 'a::banach" + shows "comm_monoid.operative op \ True (\i. f integrable_on i)" + unfolding comm_monoid.operative_def[OF comm_monoid_and] + apply safe + apply (subst integrable_on_def) + unfolding has_integral_null_eq + apply (rule, rule refl) + apply (rule, assumption, assumption)+ + unfolding integrable_on_def + by (auto intro!: has_integral_split) + +lemma integrable_subinterval: + fixes f :: "'b::euclidean_space \ 'a::banach" + assumes "f integrable_on cbox a b" + and "cbox c d \ cbox a b" + shows "f integrable_on cbox c d" + apply (cases "cbox c d = {}") + defer + apply (rule partial_division_extend_1[OF assms(2)],assumption) + using comm_monoid_set.operative_division[OF comm_monoid_set_and operative_integrable,symmetric,of _ _ _ f] assms(1) + apply (auto simp: comm_monoid_set_F_and) + done + +lemma integrable_subinterval_real: + fixes f :: "real \ 'a::banach" + assumes "f integrable_on {a .. b}" + and "{c .. d} \ {a .. b}" + shows "f integrable_on {c .. d}" + by (metis assms(1) assms(2) box_real(2) integrable_subinterval) + + +subsection \Combining adjacent intervals in 1 dimension.\ + +lemma has_integral_combine: + fixes a b c :: real + assumes "a \ c" + and "c \ b" + and "(f has_integral i) {a .. c}" + and "(f has_integral (j::'a::banach)) {c .. b}" + shows "(f has_integral (i + j)) {a .. b}" +proof - + note operative_integral[of f, unfolded comm_monoid.operative_1_le[OF add.comm_monoid_lift_option]] + note conjunctD2[OF this,rule_format] + note * = this(2)[OF conjI[OF assms(1-2)],unfolded if_P[OF assms(3)]] + then have "f integrable_on cbox a b" + apply - + apply (rule ccontr) + apply (subst(asm) if_P) + defer + apply (subst(asm) if_P) + using assms(3-) + apply auto + done + with * + show ?thesis + apply - + apply (subst(asm) if_P) + defer + apply (subst(asm) if_P) + defer + apply (subst(asm) if_P) + using assms(3-) + apply (auto simp add: integrable_on_def integral_unique) + done +qed + +lemma integral_combine: + fixes f :: "real \ 'a::banach" + assumes "a \ c" + and "c \ b" + and "f integrable_on {a .. b}" + shows "integral {a .. c} f + integral {c .. b} f = integral {a .. b} f" + apply (rule integral_unique[symmetric]) + apply (rule has_integral_combine[OF assms(1-2)]) + apply (metis assms(2) assms(3) atLeastatMost_subset_iff box_real(2) content_pos_le content_real_eq_0 integrable_integral integrable_subinterval le_add_same_cancel2 monoid_add_class.add.left_neutral) + by (metis assms(1) assms(3) atLeastatMost_subset_iff box_real(2) content_pos_le content_real_eq_0 integrable_integral integrable_subinterval le_add_same_cancel1 monoid_add_class.add.right_neutral) + +lemma integrable_combine: + fixes f :: "real \ 'a::banach" + assumes "a \ c" + and "c \ b" + and "f integrable_on {a .. c}" + and "f integrable_on {c .. b}" + shows "f integrable_on {a .. b}" + using assms + unfolding integrable_on_def + by (fastforce intro!:has_integral_combine) + + +subsection \Reduce integrability to "local" integrability.\ + +lemma integrable_on_little_subintervals: + fixes f :: "'b::euclidean_space \ 'a::banach" + assumes "\x\cbox a b. \d>0. \u v. x \ cbox u v \ cbox u v \ ball x d \ cbox u v \ cbox a b \ + f integrable_on cbox u v" + shows "f integrable_on cbox a b" +proof - + have "\x. \d. x\cbox a b \ d>0 \ (\u v. x \ cbox u v \ cbox u v \ ball x d \ cbox u v \ cbox a b \ + f integrable_on cbox u v)" + using assms by auto + note this[unfolded gauge_existence_lemma] + from choice[OF this] guess d .. note d=this[rule_format] + guess p + apply (rule fine_division_exists[OF gauge_ball_dependent,of d a b]) + using d + by auto + note p=this(1-2) + note division_of_tagged_division[OF this(1)] + note * = comm_monoid_set.operative_division[OF comm_monoid_set_and operative_integrable, OF this, symmetric, of f] + show ?thesis + unfolding * comm_monoid_set_F_and + apply safe + unfolding snd_conv + proof - + fix x k + assume "(x, k) \ p" + note tagged_division_ofD(2-4)[OF p(1) this] fineD[OF p(2) this] + then show "f integrable_on k" + apply safe + apply (rule d[THEN conjunct2,rule_format,of x]) + apply (auto intro: order.trans) + done + qed +qed + + +subsection \Second FTC or existence of antiderivative.\ + +lemma integrable_const[intro]: "(\x. c) integrable_on cbox a b" + unfolding integrable_on_def + apply rule + apply (rule has_integral_const) + done + +lemma integral_has_vector_derivative_continuous_at: + fixes f :: "real \ 'a::banach" + assumes f: "f integrable_on {a..b}" + and x: "x \ {a..b}" + and fx: "continuous (at x within {a..b}) f" + shows "((\u. integral {a..u} f) has_vector_derivative f x) (at x within {a..b})" +proof - + let ?I = "\a b. integral {a..b} f" + { fix e::real + assume "e > 0" + obtain d where "d>0" and d: "\x'. \x' \ {a..b}; \x' - x\ < d\ \ norm(f x' - f x) \ e" + using \e>0\ fx by (auto simp: continuous_within_eps_delta dist_norm less_imp_le) + have "norm (integral {a..y} f - integral {a..x} f - (y - x) *\<^sub>R f x) \ e * \y - x\" + if y: "y \ {a..b}" and yx: "\y - x\ < d" for y + proof (cases "y < x") + case False + have "f integrable_on {a..y}" + using f y by (simp add: integrable_subinterval_real) + then have Idiff: "?I a y - ?I a x = ?I x y" + using False x by (simp add: algebra_simps integral_combine) + have fux_int: "((\u. f u - f x) has_integral integral {x..y} f - (y - x) *\<^sub>R f x) {x..y}" + apply (rule has_integral_sub) + using x y apply (force intro: integrable_integral [OF integrable_subinterval_real [OF f]]) + using has_integral_const_real [of "f x" x y] False + apply (simp add: ) + done + show ?thesis + using False + apply (simp add: abs_eq_content del: content_real_if) + apply (rule has_integral_bound_real[where f="(\u. f u - f x)"]) + using yx False d x y \e>0\ apply (auto simp add: Idiff fux_int) + done + next + case True + have "f integrable_on {a..x}" + using f x by (simp add: integrable_subinterval_real) + then have Idiff: "?I a x - ?I a y = ?I y x" + using True x y by (simp add: algebra_simps integral_combine) + have fux_int: "((\u. f u - f x) has_integral integral {y..x} f - (x - y) *\<^sub>R f x) {y..x}" + apply (rule has_integral_sub) + using x y apply (force intro: integrable_integral [OF integrable_subinterval_real [OF f]]) + using has_integral_const_real [of "f x" y x] True + apply (simp add: ) + done + have "norm (integral {a..x} f - integral {a..y} f - (x - y) *\<^sub>R f x) \ e * \y - x\" + using True + apply (simp add: abs_eq_content del: content_real_if) + apply (rule has_integral_bound_real[where f="(\u. f u - f x)"]) + using yx True d x y \e>0\ apply (auto simp add: Idiff fux_int) + done + then show ?thesis + by (simp add: algebra_simps norm_minus_commute) + qed + then have "\d>0. \y\{a..b}. \y - x\ < d \ norm (integral {a..y} f - integral {a..x} f - (y - x) *\<^sub>R f x) \ e * \y - x\" + using \d>0\ by blast + } + then show ?thesis + by (simp add: has_vector_derivative_def has_derivative_within_alt bounded_linear_scaleR_left) +qed + +lemma integral_has_vector_derivative: + fixes f :: "real \ 'a::banach" + assumes "continuous_on {a .. b} f" + and "x \ {a .. b}" + shows "((\u. integral {a .. u} f) has_vector_derivative f(x)) (at x within {a .. b})" +apply (rule integral_has_vector_derivative_continuous_at [OF integrable_continuous_real]) +using assms +apply (auto simp: continuous_on_eq_continuous_within) +done + +lemma antiderivative_continuous: + fixes q b :: real + assumes "continuous_on {a .. b} f" + obtains g where "\x\{a .. b}. (g has_vector_derivative (f x::_::banach)) (at x within {a .. b})" + apply (rule that) + apply rule + using integral_has_vector_derivative[OF assms] + apply auto + done + + +subsection \Combined fundamental theorem of calculus.\ + +lemma antiderivative_integral_continuous: + fixes f :: "real \ 'a::banach" + assumes "continuous_on {a .. b} f" + obtains g where "\u\{a .. b}. \v \ {a .. b}. u \ v \ (f has_integral (g v - g u)) {u .. v}" +proof - + from antiderivative_continuous[OF assms] guess g . note g=this + show ?thesis + apply (rule that[of g]) + apply safe + proof goal_cases + case prems: (1 u v) + have "\x\cbox u v. (g has_vector_derivative f x) (at x within cbox u v)" + apply rule + apply (rule has_vector_derivative_within_subset) + apply (rule g[rule_format]) + using prems(1,2) + apply auto + done + then show ?case + using fundamental_theorem_of_calculus[OF prems(3), of g f] by auto + qed +qed + + +subsection \General "twiddling" for interval-to-interval function image.\ + +lemma has_integral_twiddle: + assumes "0 < r" + and "\x. h(g x) = x" + and "\x. g(h x) = x" + and "\x. continuous (at x) g" + and "\u v. \w z. g ` cbox u v = cbox w z" + and "\u v. \w z. h ` cbox u v = cbox w z" + and "\u v. content(g ` cbox u v) = r * content (cbox u v)" + and "(f has_integral i) (cbox a b)" + shows "((\x. f(g x)) has_integral (1 / r) *\<^sub>R i) (h ` cbox a b)" +proof - + show ?thesis when *: "cbox a b \ {} \ ?thesis" + apply cases + defer + apply (rule *) + apply assumption + proof goal_cases + case prems: 1 + then show ?thesis + unfolding prems assms(8)[unfolded prems has_integral_empty_eq] by auto + qed + assume "cbox a b \ {}" + from assms(6)[rule_format,of a b] guess w z by (elim exE) note wz=this + have inj: "inj g" "inj h" + unfolding inj_on_def + apply safe + apply(rule_tac[!] ccontr) + using assms(2) + apply(erule_tac x=x in allE) + using assms(2) + apply(erule_tac x=y in allE) + defer + using assms(3) + apply (erule_tac x=x in allE) + using assms(3) + apply(erule_tac x=y in allE) + apply auto + done + show ?thesis + unfolding has_integral_def has_integral_compact_interval_def + apply (subst if_P) + apply rule + apply rule + apply (rule wz) + proof safe + fix e :: real + assume e: "e > 0" + with assms(1) have "e * r > 0" by simp + from assms(8)[unfolded has_integral,rule_format,OF this] guess d by (elim exE conjE) note d=this[rule_format] + define d' where "d' x = {y. g y \ d (g x)}" for x + have d': "\x. d' x = {y. g y \ (d (g x))}" + unfolding d'_def .. + show "\d. gauge d \ (\p. p tagged_division_of h ` cbox a b \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e)" + proof (rule_tac x=d' in exI, safe) + show "gauge d'" + using d(1) + unfolding gauge_def d' + using continuous_open_preimage_univ[OF assms(4)] + by auto + fix p + assume as: "p tagged_division_of h ` cbox a b" "d' fine p" + note p = tagged_division_ofD[OF as(1)] + have "(\(x, k). (g x, g ` k)) ` p tagged_division_of (cbox a b) \ d fine (\(x, k). (g x, g ` k)) ` p" + unfolding tagged_division_of + proof safe + show "finite ((\(x, k). (g x, g ` k)) ` p)" + using as by auto + show "d fine (\(x, k). (g x, g ` k)) ` p" + using as(2) unfolding fine_def d' by auto + fix x k + assume xk[intro]: "(x, k) \ p" + show "g x \ g ` k" + using p(2)[OF xk] by auto + show "\u v. g ` k = cbox u v" + using p(4)[OF xk] using assms(5-6) by auto + { + fix y + assume "y \ k" + then show "g y \ cbox a b" "g y \ cbox a b" + using p(3)[OF xk,unfolded subset_eq,rule_format,of "h (g y)"] + using assms(2)[rule_format,of y] + unfolding inj_image_mem_iff[OF inj(2)] + by auto + } + fix x' k' + assume xk': "(x', k') \ p" + fix z + assume z: "z \ interior (g ` k)" "z \ interior (g ` k')" + have same: "(x, k) = (x', k')" + apply - + apply (rule ccontr) + apply (drule p(5)[OF xk xk']) + proof - + assume as: "interior k \ interior k' = {}" + have "z \ g ` (interior k \ interior k')" + using interior_image_subset[OF assms(4) inj(1)] z + unfolding image_Int[OF inj(1)] by blast + then show False + using as by blast + qed + then show "g x = g x'" + by auto + { + fix z + assume "z \ k" + then show "g z \ g ` k'" + using same by auto + } + { + fix z + assume "z \ k'" + then show "g z \ g ` k" + using same by auto + } + next + fix x + assume "x \ cbox a b" + then have "h x \ \{k. \x. (x, k) \ p}" + using p(6) by auto + then guess X unfolding Union_iff .. note X=this + from this(1) guess y unfolding mem_Collect_eq .. + then show "x \ \{k. \x. (x, k) \ (\(x, k). (g x, g ` k)) ` p}" + apply - + apply (rule_tac X="g ` X" in UnionI) + defer + apply (rule_tac x="h x" in image_eqI) + using X(2) assms(3)[rule_format,of x] + apply auto + done + qed + note ** = d(2)[OF this] + have *: "inj_on (\(x, k). (g x, g ` k)) p" + using inj(1) unfolding inj_on_def by fastforce + have "(\(x, k)\(\(x, k). (g x, g ` k)) ` p. content k *\<^sub>R f x) - i = r *\<^sub>R (\(x, k)\p. content k *\<^sub>R f (g x)) - i" (is "?l = _") + using assms(7) + apply (simp only: algebra_simps add_left_cancel scaleR_right.setsum) + apply (subst setsum.reindex_bij_betw[symmetric, where h="\(x, k). (g x, g ` k)" and S=p]) + apply (auto intro!: * setsum.cong simp: bij_betw_def dest!: p(4)) + done + also have "\ = r *\<^sub>R ((\(x, k)\p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i)" (is "_ = ?r") + unfolding scaleR_diff_right scaleR_scaleR + using assms(1) + by auto + finally have *: "?l = ?r" . + show "norm ((\(x, k)\p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e" + using ** + unfolding * + unfolding norm_scaleR + using assms(1) + by (auto simp add:field_simps) + qed + qed +qed + + +subsection \Special case of a basic affine transformation.\ + +lemma interval_image_affinity_interval: + "\u v. (\x. m *\<^sub>R (x::'a::euclidean_space) + c) ` cbox a b = cbox u v" + unfolding image_affinity_cbox + by auto + +lemma content_image_affinity_cbox: + "content((\x::'a::euclidean_space. m *\<^sub>R x + c) ` cbox a b) = + \m\ ^ DIM('a) * content (cbox a b)" (is "?l = ?r") +proof (cases "cbox a b = {}") + case True then show ?thesis by simp +next + case False + show ?thesis + proof (cases "m \ 0") + case True + with \cbox a b \ {}\ have "cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c) \ {}" + unfolding box_ne_empty + apply (intro ballI) + apply (erule_tac x=i in ballE) + apply (auto simp: inner_simps mult_left_mono) + done + moreover from True have *: "\i. (m *\<^sub>R b + c) \ i - (m *\<^sub>R a + c) \ i = m *\<^sub>R (b - a) \ i" + by (simp add: inner_simps field_simps) + ultimately show ?thesis + by (simp add: image_affinity_cbox True content_cbox' + setprod.distrib setprod_constant inner_diff_left) + next + case False + with \cbox a b \ {}\ have "cbox (m *\<^sub>R b + c) (m *\<^sub>R a + c) \ {}" + unfolding box_ne_empty + apply (intro ballI) + apply (erule_tac x=i in ballE) + apply (auto simp: inner_simps mult_left_mono) + done + moreover from False have *: "\i. (m *\<^sub>R a + c) \ i - (m *\<^sub>R b + c) \ i = (-m) *\<^sub>R (b - a) \ i" + by (simp add: inner_simps field_simps) + ultimately show ?thesis using False + by (simp add: image_affinity_cbox content_cbox' + setprod.distrib[symmetric] setprod_constant[symmetric] inner_diff_left) + qed +qed + +lemma has_integral_affinity: + fixes a :: "'a::euclidean_space" + assumes "(f has_integral i) (cbox a b)" + and "m \ 0" + shows "((\x. f(m *\<^sub>R x + c)) has_integral ((1 / (\m\ ^ DIM('a))) *\<^sub>R i)) ((\x. (1 / m) *\<^sub>R x + -((1 / m) *\<^sub>R c)) ` cbox a b)" + apply (rule has_integral_twiddle) + using assms + apply (safe intro!: interval_image_affinity_interval content_image_affinity_cbox) + apply (rule zero_less_power) + unfolding scaleR_right_distrib + apply auto + done + +lemma integrable_affinity: + assumes "f integrable_on cbox a b" + and "m \ 0" + shows "(\x. f(m *\<^sub>R x + c)) integrable_on ((\x. (1 / m) *\<^sub>R x + -((1/m) *\<^sub>R c)) ` cbox a b)" + using assms + unfolding integrable_on_def + apply safe + apply (drule has_integral_affinity) + apply auto + done + +lemmas has_integral_affinity01 = has_integral_affinity [of _ _ 0 "1::real", simplified] + +subsection \Special case of stretching coordinate axes separately.\ + +lemma content_image_stretch_interval: + "content ((\x::'a::euclidean_space. (\k\Basis. (m k * (x\k))*\<^sub>R k)::'a) ` cbox a b) = + \setprod m Basis\ * content (cbox a b)" +proof (cases "cbox a b = {}") + case True + then show ?thesis + unfolding content_def image_is_empty image_stretch_interval if_P[OF True] by auto +next + case False + then have "(\x. (\k\Basis. (m k * (x\k))*\<^sub>R k)) ` cbox a b \ {}" + by auto + then show ?thesis + using False + unfolding content_def image_stretch_interval + apply - + unfolding interval_bounds' if_not_P + unfolding abs_setprod setprod.distrib[symmetric] + apply (rule setprod.cong) + apply (rule refl) + unfolding lessThan_iff + apply (simp only: inner_setsum_left_Basis) + proof - + fix i :: 'a + assume i: "i \ Basis" + have "(m i < 0 \ m i > 0) \ m i = 0" + by auto + then show "max (m i * (a \ i)) (m i * (b \ i)) - min (m i * (a \ i)) (m i * (b \ i)) = + \m i\ * (b \ i - a \ i)" + apply - + apply (erule disjE)+ + unfolding min_def max_def + using False[unfolded box_ne_empty,rule_format,of i] i + apply (auto simp add:field_simps not_le mult_le_cancel_left_neg mult_le_cancel_left_pos) + done + qed +qed + +lemma has_integral_stretch: + fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" + assumes "(f has_integral i) (cbox a b)" + and "\k\Basis. m k \ 0" + shows "((\x. f (\k\Basis. (m k * (x\k))*\<^sub>R k)) has_integral + ((1/ \setprod m Basis\) *\<^sub>R i)) ((\x. (\k\Basis. (1 / m k * (x\k))*\<^sub>R k)) ` cbox a b)" + apply (rule has_integral_twiddle[where f=f]) + unfolding zero_less_abs_iff content_image_stretch_interval + unfolding image_stretch_interval empty_as_interval euclidean_eq_iff[where 'a='a] + using assms +proof - + show "\y::'a. continuous (at y) (\x. (\k\Basis. (m k * (x\k))*\<^sub>R k))" + apply rule + apply (rule linear_continuous_at) + unfolding linear_linear + unfolding linear_iff inner_simps euclidean_eq_iff[where 'a='a] + apply (auto simp add: field_simps) + done +qed auto + +lemma integrable_stretch: + fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" + assumes "f integrable_on cbox a b" + and "\k\Basis. m k \ 0" + shows "(\x::'a. f (\k\Basis. (m k * (x\k))*\<^sub>R k)) integrable_on + ((\x. \k\Basis. (1 / m k * (x\k))*\<^sub>R k) ` cbox a b)" + using assms + unfolding integrable_on_def + apply - + apply (erule exE) + apply (drule has_integral_stretch) + apply assumption + apply auto + done + +subsection \even more special cases.\ + +lemma uminus_interval_vector[simp]: + fixes a b :: "'a::euclidean_space" + shows "uminus ` cbox a b = cbox (-b) (-a)" + apply (rule set_eqI) + apply rule + defer + unfolding image_iff + apply (rule_tac x="-x" in bexI) + apply (auto simp add:minus_le_iff le_minus_iff mem_box) + done + +lemma has_integral_reflect_lemma[intro]: + assumes "(f has_integral i) (cbox a b)" + shows "((\x. f(-x)) has_integral i) (cbox (-b) (-a))" + using has_integral_affinity[OF assms, of "-1" 0] + by auto + +lemma has_integral_reflect_lemma_real[intro]: + assumes "(f has_integral i) {a .. b::real}" + shows "((\x. f(-x)) has_integral i) {-b .. -a}" + using assms + unfolding box_real[symmetric] + by (rule has_integral_reflect_lemma) + +lemma has_integral_reflect[simp]: + "((\x. f (-x)) has_integral i) (cbox (-b) (-a)) \ (f has_integral i) (cbox a b)" + apply rule + apply (drule_tac[!] has_integral_reflect_lemma) + apply auto + done + +lemma integrable_reflect[simp]: "(\x. f(-x)) integrable_on cbox (-b) (-a) \ f integrable_on cbox a b" + unfolding integrable_on_def by auto + +lemma integrable_reflect_real[simp]: "(\x. f(-x)) integrable_on {-b .. -a} \ f integrable_on {a .. b::real}" + unfolding box_real[symmetric] + by (rule integrable_reflect) + +lemma integral_reflect[simp]: "integral (cbox (-b) (-a)) (\x. f (-x)) = integral (cbox a b) f" + unfolding integral_def by auto + +lemma integral_reflect_real[simp]: "integral {-b .. -a} (\x. f (-x)) = integral {a .. b::real} f" + unfolding box_real[symmetric] + by (rule integral_reflect) + + +subsection \Stronger form of FCT; quite a tedious proof.\ + +lemma bgauge_existence_lemma: "(\x\s. \d::real. 0 < d \ q d x) \ (\x. \d>0. x\s \ q d x)" + by (meson zero_less_one) + +lemma additive_tagged_division_1': + fixes f :: "real \ 'a::real_normed_vector" + assumes "a \ b" + and "p tagged_division_of {a..b}" + shows "setsum (\(x,k). f (Sup k) - f(Inf k)) p = f b - f a" + using additive_tagged_division_1[OF _ assms(2), of f] + using assms(1) + by auto + +lemma split_minus[simp]: "(\(x, k). f x k) x - (\(x, k). g x k) x = (\(x, k). f x k - g x k) x" + by (simp add: split_def) + +lemma norm_triangle_le_sub: "norm x + norm y \ e \ norm (x - y) \ e" + apply (subst(asm)(2) norm_minus_cancel[symmetric]) + apply (drule norm_triangle_le) + apply (auto simp add: algebra_simps) + done + +lemma fundamental_theorem_of_calculus_interior: + fixes f :: "real \ 'a::real_normed_vector" + assumes "a \ b" + and "continuous_on {a .. b} f" + and "\x\{a <..< b}. (f has_vector_derivative f'(x)) (at x)" + shows "(f' has_integral (f b - f a)) {a .. b}" +proof - + { + presume *: "a < b \ ?thesis" + show ?thesis + proof (cases "a < b") + case True + then show ?thesis by (rule *) + next + case False + then have "a = b" + using assms(1) by auto + then have *: "cbox a b = {b}" "f b - f a = 0" + by (auto simp add: order_antisym) + show ?thesis + unfolding *(2) + unfolding content_eq_0 + using * \a = b\ + by (auto simp: ex_in_conv) + qed + } + assume ab: "a < b" + let ?P = "\e. \d. gauge d \ (\p. p tagged_division_of {a .. b} \ d fine p \ + norm ((\(x, k)\p. content k *\<^sub>R f' x) - (f b - f a)) \ e * content {a .. b})" + { presume "\e. e > 0 \ ?P e" then show ?thesis unfolding has_integral_factor_content_real by auto } + fix e :: real + assume e: "e > 0" + note assms(3)[unfolded has_vector_derivative_def has_derivative_at_alt ball_conj_distrib] + note conjunctD2[OF this] + note bounded=this(1) and this(2) + from this(2) have "\x\box a b. \d>0. \y. norm (y - x) < d \ + norm (f y - f x - (y - x) *\<^sub>R f' x) \ e/2 * norm (y - x)" + apply - + apply safe + apply (erule_tac x=x in ballE) + apply (erule_tac x="e/2" in allE) + using e + apply auto + done + note this[unfolded bgauge_existence_lemma] + from choice[OF this] guess d .. + note conjunctD2[OF this[rule_format]] + note d = this[rule_format] + have "bounded (f ` cbox a b)" + apply (rule compact_imp_bounded compact_continuous_image)+ + using compact_cbox assms + apply auto + done + from this[unfolded bounded_pos] guess B .. note B = this[rule_format] + + have "\da. 0 < da \ (\c. a \ c \ {a .. c} \ {a .. b} \ {a .. c} \ ball a da \ + norm (content {a .. c} *\<^sub>R f' a - (f c - f a)) \ (e * (b - a)) / 4)" + proof - + have "a \ {a .. b}" + using ab by auto + note assms(2)[unfolded continuous_on_eq_continuous_within,rule_format,OF this] + note * = this[unfolded continuous_within Lim_within,rule_format] + have "(e * (b - a)) / 8 > 0" + using e ab by (auto simp add: field_simps) + from *[OF this] guess k .. note k = conjunctD2[OF this,rule_format] + have "\l. 0 < l \ norm(l *\<^sub>R f' a) \ (e * (b - a)) / 8" + proof (cases "f' a = 0") + case True + thus ?thesis using ab e by auto + next + case False + then show ?thesis + apply (rule_tac x="(e * (b - a)) / 8 / norm (f' a)" in exI) + using ab e + apply (auto simp add: field_simps) + done + qed + then guess l .. note l = conjunctD2[OF this] + show ?thesis + apply (rule_tac x="min k l" in exI) + apply safe + unfolding min_less_iff_conj + apply rule + apply (rule l k)+ + proof - + fix c + assume as: "a \ c" "{a .. c} \ {a .. b}" "{a .. c} \ ball a (min k l)" + note as' = this[unfolded subset_eq Ball_def mem_ball dist_real_def mem_box] + have "norm ((c - a) *\<^sub>R f' a - (f c - f a)) \ norm ((c - a) *\<^sub>R f' a) + norm (f c - f a)" + by (rule norm_triangle_ineq4) + also have "\ \ e * (b - a) / 8 + e * (b - a) / 8" + proof (rule add_mono) + have "\c - a\ \ \l\" + using as' by auto + then show "norm ((c - a) *\<^sub>R f' a) \ e * (b - a) / 8" + apply - + apply (rule order_trans[OF _ l(2)]) + unfolding norm_scaleR + apply (rule mult_right_mono) + apply auto + done + next + show "norm (f c - f a) \ e * (b - a) / 8" + apply (rule less_imp_le) + apply (cases "a = c") + defer + apply (rule k(2)[unfolded dist_norm]) + using as' e ab + apply (auto simp add: field_simps) + done + qed + finally show "norm (content {a .. c} *\<^sub>R f' a - (f c - f a)) \ e * (b - a) / 4" + unfolding content_real[OF as(1)] by auto + qed + qed + then guess da .. note da=conjunctD2[OF this,rule_format] + + have "\db>0. \c\b. {c .. b} \ {a .. b} \ {c .. b} \ ball b db \ + norm (content {c .. b} *\<^sub>R f' b - (f b - f c)) \ (e * (b - a)) / 4" + proof - + have "b \ {a .. b}" + using ab by auto + note assms(2)[unfolded continuous_on_eq_continuous_within,rule_format,OF this] + note * = this[unfolded continuous_within Lim_within,rule_format] have "(e * (b - a)) / 8 > 0" + using e ab by (auto simp add: field_simps) + from *[OF this] guess k .. note k = conjunctD2[OF this,rule_format] + have "\l. 0 < l \ norm (l *\<^sub>R f' b) \ (e * (b - a)) / 8" + proof (cases "f' b = 0") + case True + thus ?thesis using ab e by auto + next + case False + then show ?thesis + apply (rule_tac x="(e * (b - a)) / 8 / norm (f' b)" in exI) + using ab e + apply (auto simp add: field_simps) + done + qed + then guess l .. note l = conjunctD2[OF this] + show ?thesis + apply (rule_tac x="min k l" in exI) + apply safe + unfolding min_less_iff_conj + apply rule + apply (rule l k)+ + proof - + fix c + assume as: "c \ b" "{c..b} \ {a..b}" "{c..b} \ ball b (min k l)" + note as' = this[unfolded subset_eq Ball_def mem_ball dist_real_def mem_box] + have "norm ((b - c) *\<^sub>R f' b - (f b - f c)) \ norm ((b - c) *\<^sub>R f' b) + norm (f b - f c)" + by (rule norm_triangle_ineq4) + also have "\ \ e * (b - a) / 8 + e * (b - a) / 8" + proof (rule add_mono) + have "\c - b\ \ \l\" + using as' by auto + then show "norm ((b - c) *\<^sub>R f' b) \ e * (b - a) / 8" + apply - + apply (rule order_trans[OF _ l(2)]) + unfolding norm_scaleR + apply (rule mult_right_mono) + apply auto + done + next + show "norm (f b - f c) \ e * (b - a) / 8" + apply (rule less_imp_le) + apply (cases "b = c") + defer + apply (subst norm_minus_commute) + apply (rule k(2)[unfolded dist_norm]) + using as' e ab + apply (auto simp add: field_simps) + done + qed + finally show "norm (content {c .. b} *\<^sub>R f' b - (f b - f c)) \ e * (b - a) / 4" + unfolding content_real[OF as(1)] by auto + qed + qed + then guess db .. note db=conjunctD2[OF this,rule_format] + + let ?d = "(\x. ball x (if x=a then da else if x=b then db else d x))" + show "?P e" + apply (rule_tac x="?d" in exI) + proof (safe, goal_cases) + case 1 + show ?case + apply (rule gauge_ball_dependent) + using ab db(1) da(1) d(1) + apply auto + done + next + case as: (2 p) + let ?A = "{t. fst t \ {a, b}}" + note p = tagged_division_ofD[OF as(1)] + have pA: "p = (p \ ?A) \ (p - ?A)" "finite (p \ ?A)" "finite (p - ?A)" "(p \ ?A) \ (p - ?A) = {}" + using as by auto + note * = additive_tagged_division_1'[OF assms(1) as(1), symmetric] + have **: "\n1 s1 n2 s2::real. n2 \ s2 / 2 \ n1 - s1 \ s2 / 2 \ n1 + n2 \ s1 + s2" + by arith + show ?case + unfolding content_real[OF assms(1)] and *[of "\x. x"] *[of f] setsum_subtractf[symmetric] split_minus + unfolding setsum_right_distrib + apply (subst(2) pA) + apply (subst pA) + unfolding setsum.union_disjoint[OF pA(2-)] + proof (rule norm_triangle_le, rule **, goal_cases) + case 1 + show ?case + apply (rule order_trans) + apply (rule setsum_norm_le) + defer + apply (subst setsum_divide_distrib) + apply (rule order_refl) + apply safe + apply (unfold not_le o_def split_conv fst_conv) + proof (rule ccontr) + fix x k + assume xk: "(x, k) \ p" + "e * (Sup k - Inf k) / 2 < + norm (content k *\<^sub>R f' x - (f (Sup k) - f (Inf k)))" + from p(4)[OF this(1)] guess u v by (elim exE) note k=this + then have "u \ v" and uv: "{u, v} \ cbox u v" + using p(2)[OF xk(1)] by auto + note result = xk(2)[unfolded k box_real interval_bounds_real[OF this(1)] content_real[OF this(1)]] + + assume as': "x \ a" "x \ b" + then have "x \ box a b" + using p(2-3)[OF xk(1)] by (auto simp: mem_box) + note * = d(2)[OF this] + have "norm ((v - u) *\<^sub>R f' (x) - (f (v) - f (u))) = + norm ((f (u) - f (x) - (u - x) *\<^sub>R f' (x)) - (f (v) - f (x) - (v - x) *\<^sub>R f' (x)))" + apply (rule arg_cong[of _ _ norm]) + unfolding scaleR_left.diff + apply auto + done + also have "\ \ e / 2 * norm (u - x) + e / 2 * norm (v - x)" + apply (rule norm_triangle_le_sub) + apply (rule add_mono) + apply (rule_tac[!] *) + using fineD[OF as(2) xk(1)] as' + unfolding k subset_eq + apply - + apply (erule_tac x=u in ballE) + apply (erule_tac[3] x=v in ballE) + using uv + apply (auto simp:dist_real_def) + done + also have "\ \ e / 2 * norm (v - u)" + using p(2)[OF xk(1)] + unfolding k + by (auto simp add: field_simps) + finally have "e * (v - u) / 2 < e * (v - u) / 2" + apply - + apply (rule less_le_trans[OF result]) + using uv + apply auto + done + then show False by auto + qed + next + have *: "\x s1 s2::real. 0 \ s1 \ x \ (s1 + s2) / 2 \ x - s1 \ s2 / 2" + by auto + case 2 + show ?case + apply (rule *) + apply (rule setsum_nonneg) + apply rule + apply (unfold split_paired_all split_conv) + defer + unfolding setsum.union_disjoint[OF pA(2-),symmetric] pA(1)[symmetric] + unfolding setsum_right_distrib[symmetric] + apply (subst additive_tagged_division_1[OF _ as(1)]) + apply (rule assms) + proof - + fix x k + assume "(x, k) \ p \ {t. fst t \ {a, b}}" + note xk=IntD1[OF this] + from p(4)[OF this] guess u v by (elim exE) note uv=this + with p(2)[OF xk] have "cbox u v \ {}" + by auto + then show "0 \ e * ((Sup k) - (Inf k))" + unfolding uv using e by (auto simp add: field_simps) + next + have *: "\s f t e. setsum f s = setsum f t \ norm (setsum f t) \ e \ norm (setsum f s) \ e" + by auto + show "norm (\(x, k)\p \ ?A. content k *\<^sub>R f' x - + (f ((Sup k)) - f ((Inf k)))) \ e * (b - a) / 2" + apply (rule *[where t1="p \ {t. fst t \ {a, b} \ content(snd t) \ 0}"]) + apply (rule setsum.mono_neutral_right[OF pA(2)]) + defer + apply rule + unfolding split_paired_all split_conv o_def + proof goal_cases + fix x k + assume "(x, k) \ p \ {t. fst t \ {a, b}} - p \ {t. fst t \ {a, b} \ content (snd t) \ 0}" + then have xk: "(x, k) \ p" "content k = 0" + by auto + from p(4)[OF xk(1)] guess u v by (elim exE) note uv=this + have "k \ {}" + using p(2)[OF xk(1)] by auto + then have *: "u = v" + using xk + unfolding uv content_eq_0 box_eq_empty + by auto + then show "content k *\<^sub>R (f' (x)) - (f ((Sup k)) - f ((Inf k))) = 0" + using xk unfolding uv by auto + next + have *: "p \ {t. fst t \ {a, b} \ content(snd t) \ 0} = + {t. t\p \ fst t = a \ content(snd t) \ 0} \ {t. t\p \ fst t = b \ content(snd t) \ 0}" + by blast + have **: "norm (setsum f s) \ e" + if "\x y. x \ s \ y \ s \ x = y" + and "\x. x \ s \ norm (f x) \ e" + and "e > 0" + for s f and e :: real + proof (cases "s = {}") + case True + with that show ?thesis by auto + next + case False + then obtain x where "x \ s" + by auto + then have *: "s = {x}" + using that(1) by auto + then show ?thesis + using \x \ s\ that(2) by auto + qed + case 2 + show ?case + apply (subst *) + apply (subst setsum.union_disjoint) + prefer 4 + apply (rule order_trans[of _ "e * (b - a)/4 + e * (b - a)/4"]) + apply (rule norm_triangle_le,rule add_mono) + apply (rule_tac[1-2] **) + proof - + let ?B = "\x. {t \ p. fst t = x \ content (snd t) \ 0}" + have pa: "\v. k = cbox a v \ a \ v" if "(a, k) \ p" for k + proof - + guess u v using p(4)[OF that] by (elim exE) note uv=this + have *: "u \ v" + using p(2)[OF that] unfolding uv by auto + have u: "u = a" + proof (rule ccontr) + have "u \ cbox u v" + using p(2-3)[OF that(1)] unfolding uv by auto + have "u \ a" + using p(2-3)[OF that(1)] unfolding uv subset_eq by auto + moreover assume "\ ?thesis" + ultimately have "u > a" by auto + then show False + using p(2)[OF that(1)] unfolding uv by (auto simp add:) + qed + then show ?thesis + apply (rule_tac x=v in exI) + unfolding uv + using * + apply auto + done + qed + have pb: "\v. k = cbox v b \ b \ v" if "(b, k) \ p" for k + proof - + guess u v using p(4)[OF that] by (elim exE) note uv=this + have *: "u \ v" + using p(2)[OF that] unfolding uv by auto + have u: "v = b" + proof (rule ccontr) + have "u \ cbox u v" + using p(2-3)[OF that(1)] unfolding uv by auto + have "v \ b" + using p(2-3)[OF that(1)] unfolding uv subset_eq by auto + moreover assume "\ ?thesis" + ultimately have "v < b" by auto + then show False + using p(2)[OF that(1)] unfolding uv by (auto simp add:) + qed + then show ?thesis + apply (rule_tac x=u in exI) + unfolding uv + using * + apply auto + done + qed + show "\x y. x \ ?B a \ y \ ?B a \ x = y" + apply (rule,rule,rule,unfold split_paired_all) + unfolding mem_Collect_eq fst_conv snd_conv + apply safe + proof - + fix x k k' + assume k: "(a, k) \ p" "(a, k') \ p" "content k \ 0" "content k' \ 0" + guess v using pa[OF k(1)] .. note v = conjunctD2[OF this] + guess v' using pa[OF k(2)] .. note v' = conjunctD2[OF this] let ?v = "min v v'" + have "box a ?v \ k \ k'" + unfolding v v' by (auto simp add: mem_box) + note interior_mono[OF this,unfolded interior_Int] + moreover have "(a + ?v)/2 \ box a ?v" + using k(3-) + unfolding v v' content_eq_0 not_le + by (auto simp add: mem_box) + ultimately have "(a + ?v)/2 \ interior k \ interior k'" + unfolding interior_open[OF open_box] by auto + then have *: "k = k'" + apply - + apply (rule ccontr) + using p(5)[OF k(1-2)] + apply auto + done + { assume "x \ k" then show "x \ k'" unfolding * . } + { assume "x \ k'" then show "x \ k" unfolding * . } + qed + show "\x y. x \ ?B b \ y \ ?B b \ x = y" + apply rule + apply rule + apply rule + apply (unfold split_paired_all) + unfolding mem_Collect_eq fst_conv snd_conv + apply safe + proof - + fix x k k' + assume k: "(b, k) \ p" "(b, k') \ p" "content k \ 0" "content k' \ 0" + guess v using pb[OF k(1)] .. note v = conjunctD2[OF this] + guess v' using pb[OF k(2)] .. note v' = conjunctD2[OF this] + let ?v = "max v v'" + have "box ?v b \ k \ k'" + unfolding v v' by (auto simp: mem_box) + note interior_mono[OF this,unfolded interior_Int] + moreover have " ((b + ?v)/2) \ box ?v b" + using k(3-) unfolding v v' content_eq_0 not_le by (auto simp: mem_box) + ultimately have " ((b + ?v)/2) \ interior k \ interior k'" + unfolding interior_open[OF open_box] by auto + then have *: "k = k'" + apply - + apply (rule ccontr) + using p(5)[OF k(1-2)] + apply auto + done + { assume "x \ k" then show "x \ k'" unfolding * . } + { assume "x \ k'" then show "x\k" unfolding * . } + qed + + let ?a = a and ?b = b (* a is something else while proofing the next theorem. *) + show "\x. x \ ?B a \ norm ((\(x, k). content k *\<^sub>R f' x - (f (Sup k) - + f (Inf k))) x) \ e * (b - a) / 4" + apply rule + apply rule + unfolding mem_Collect_eq + unfolding split_paired_all fst_conv snd_conv + proof (safe, goal_cases) + case prems: 1 + guess v using pa[OF prems(1)] .. note v = conjunctD2[OF this] + have "?a \ {?a..v}" + using v(2) by auto + then have "v \ ?b" + using p(3)[OF prems(1)] unfolding subset_eq v by auto + moreover have "{?a..v} \ ball ?a da" + using fineD[OF as(2) prems(1)] + apply - + apply (subst(asm) if_P) + apply (rule refl) + unfolding subset_eq + apply safe + apply (erule_tac x=" x" in ballE) + apply (auto simp add:subset_eq dist_real_def v) + done + ultimately show ?case + unfolding v interval_bounds_real[OF v(2)] box_real + apply - + apply(rule da(2)[of "v"]) + using prems fineD[OF as(2) prems(1)] + unfolding v content_eq_0 + apply auto + done + qed + show "\x. x \ ?B b \ norm ((\(x, k). content k *\<^sub>R f' x - + (f (Sup k) - f (Inf k))) x) \ e * (b - a) / 4" + apply rule + apply rule + unfolding mem_Collect_eq + unfolding split_paired_all fst_conv snd_conv + proof (safe, goal_cases) + case prems: 1 + guess v using pb[OF prems(1)] .. note v = conjunctD2[OF this] + have "?b \ {v.. ?b}" + using v(2) by auto + then have "v \ ?a" using p(3)[OF prems(1)] + unfolding subset_eq v by auto + moreover have "{v..?b} \ ball ?b db" + using fineD[OF as(2) prems(1)] + apply - + apply (subst(asm) if_P, rule refl) + unfolding subset_eq + apply safe + apply (erule_tac x=" x" in ballE) + using ab + apply (auto simp add:subset_eq v dist_real_def) + done + ultimately show ?case + unfolding v + unfolding interval_bounds_real[OF v(2)] box_real + apply - + apply(rule db(2)[of "v"]) + using prems fineD[OF as(2) prems(1)] + unfolding v content_eq_0 + apply auto + done + qed + qed (insert p(1) ab e, auto simp add: field_simps) + qed auto + qed + qed + qed +qed + + +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}" + using assms +proof (induct "card s" arbitrary: s a b) + case 0 + show ?case + apply (rule fundamental_theorem_of_calculus_interior) + using 0 + apply auto + done +next + case (Suc n) + from this(2) guess c s' + apply - + apply (subst(asm) eq_commute) + unfolding card_Suc_eq + apply (subst(asm)(2) eq_commute) + apply (elim exE conjE) + done + note cs = this[rule_format] + show ?case + proof (cases "c \ box a b") + case False + then show ?thesis + apply - + apply (rule Suc(1)[OF cs(3) _ Suc(4,5)]) + apply safe + defer + apply (rule Suc(6)[rule_format]) + using Suc(3) + unfolding cs + apply auto + done + next + have *: "f b - f a = (f c - f a) + (f b - f c)" + by auto + case True + then have "a \ c" "c \ b" + by (auto simp: mem_box) + then show ?thesis + apply (subst *) + apply (rule has_integral_combine) + apply assumption+ + apply (rule_tac[!] Suc(1)[OF cs(3)]) + using Suc(3) + unfolding cs + proof - + show "continuous_on {a .. c} f" "continuous_on {c .. b} f" + apply (rule_tac[!] continuous_on_subset[OF Suc(5)]) + using True + apply (auto simp: mem_box) + done + let ?P = "\i j. \x\{i <..< j} - s'. (f has_vector_derivative f' x) (at x)" + show "?P a c" "?P c b" + apply safe + apply (rule_tac[!] Suc(6)[rule_format]) + using True + unfolding cs + apply (auto simp: mem_box) + done + qed auto + qed +qed + +lemma fundamental_theorem_of_calculus_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}" + apply (rule fundamental_theorem_of_calculus_interior_strong[OF assms(1-3), of f']) + using assms(4) + apply (auto simp: mem_box) + done + +lemma indefinite_integral_continuous_left: + fixes f:: "real \ 'a::banach" + assumes "f integrable_on {a .. b}" + and "a < c" + and "c \ b" + and "e > 0" + obtains d where "d > 0" + and "\t. c - d < t \ t \ c \ norm (integral {a .. c} f - integral {a .. t} f) < e" +proof - + have "\w>0. \t. c - w < t \ t < c \ norm (f c) * norm(c - t) < e / 3" + proof (cases "f c = 0") + case False + hence "0 < e / 3 / norm (f c)" using \e>0\ by simp + then show ?thesis + apply - + apply rule + apply rule + apply assumption + apply safe + proof - + fix t + assume as: "t < c" and "c - e / 3 / norm (f c) < t" + then have "c - t < e / 3 / norm (f c)" + by auto + then have "norm (c - t) < e / 3 / norm (f c)" + using as by auto + then show "norm (f c) * norm (c - t) < e / 3" + using False + apply - + apply (subst mult.commute) + apply (subst pos_less_divide_eq[symmetric]) + apply auto + done + qed + next + case True + show ?thesis + apply (rule_tac x=1 in exI) + unfolding True + using \e > 0\ + apply auto + done + qed + then guess w .. note w = conjunctD2[OF this,rule_format] + + have *: "e / 3 > 0" + using assms by auto + have "f integrable_on {a .. c}" + apply (rule integrable_subinterval_real[OF assms(1)]) + using assms(2-3) + apply auto + done + from integrable_integral[OF this,unfolded has_integral_real,rule_format,OF *] guess d1 .. + note d1 = conjunctD2[OF this,rule_format] + define d where [abs_def]: "d x = ball x w \ d1 x" for x + have "gauge d" + unfolding d_def using w(1) d1 by auto + note this[unfolded gauge_def,rule_format,of c] + note conjunctD2[OF this] + from this(2)[unfolded open_contains_ball,rule_format,OF this(1)] guess k .. + note k=conjunctD2[OF this] + + let ?d = "min k (c - a) / 2" + show ?thesis + apply (rule that[of ?d]) + apply safe + proof - + show "?d > 0" + using k(1) using assms(2) by auto + fix t + assume as: "c - ?d < t" "t \ c" + let ?thesis = "norm (integral ({a .. c}) f - integral ({a .. t}) f) < e" + { + presume *: "t < c \ ?thesis" + show ?thesis + apply (cases "t = c") + defer + apply (rule *) + apply (subst less_le) + using \e > 0\ as(2) + apply auto + done + } + assume "t < c" + + have "f integrable_on {a .. t}" + apply (rule integrable_subinterval_real[OF assms(1)]) + using assms(2-3) as(2) + apply auto + done + from integrable_integral[OF this,unfolded has_integral_real,rule_format,OF *] guess d2 .. + note d2 = conjunctD2[OF this,rule_format] + define d3 where "d3 x = (if x \ t then d1 x \ d2 x else d1 x)" for x + have "gauge d3" + using d2(1) d1(1) unfolding d3_def gauge_def by auto + from fine_division_exists_real[OF this, of a t] guess p . note p=this + note p'=tagged_division_ofD[OF this(1)] + have pt: "\(x,k)\p. x \ t" + proof (safe, goal_cases) + case prems: 1 + from p'(2,3)[OF prems] show ?case + by auto + qed + with p(2) have "d2 fine p" + unfolding fine_def d3_def + apply safe + apply (erule_tac x="(a,b)" in ballE)+ + apply auto + done + note d2_fin = d2(2)[OF conjI[OF p(1) this]] + + have *: "{a .. c} \ {x. x \ 1 \ t} = {a .. t}" "{a .. c} \ {x. x \ 1 \ t} = {t .. c}" + using assms(2-3) as by (auto simp add: field_simps) + have "p \ {(c, {t .. c})} tagged_division_of {a .. c} \ d1 fine p \ {(c, {t .. c})}" + apply rule + apply (rule tagged_division_union_interval_real[of _ _ _ 1 "t"]) + unfolding * + apply (rule p) + apply (rule tagged_division_of_self_real) + unfolding fine_def + apply safe + proof - + fix x k y + assume "(x,k) \ p" and "y \ k" + then show "y \ d1 x" + using p(2) pt + unfolding fine_def d3_def + apply - + apply (erule_tac x="(x,k)" in ballE)+ + apply auto + done + next + fix x assume "x \ {t..c}" + then have "dist c x < k" + unfolding dist_real_def + using as(1) + by (auto simp add: field_simps) + then show "x \ d1 c" + using k(2) + unfolding d_def + by auto + qed (insert as(2), auto) note d1_fin = d1(2)[OF this] + + have *: "integral {a .. c} f - integral {a .. t} f = -(((c - t) *\<^sub>R f c + (\(x, k)\p. content k *\<^sub>R f x)) - + integral {a .. c} f) + ((\(x, k)\p. content k *\<^sub>R f x) - integral {a .. t} f) + (c - t) *\<^sub>R f c" + "e = (e/3 + e/3) + e/3" + by auto + have **: "(\(x, k)\p \ {(c, {t .. c})}. content k *\<^sub>R f x) = + (c - t) *\<^sub>R f c + (\(x, k)\p. content k *\<^sub>R f x)" + proof - + have **: "\x F. F \ {x} = insert x F" + by auto + have "(c, cbox t c) \ p" + proof (safe, goal_cases) + case prems: 1 + from p'(2-3)[OF prems] have "c \ cbox a t" + by auto + then show False using \t < c\ + by auto + qed + then show ?thesis + unfolding ** box_real + apply - + apply (subst setsum.insert) + apply (rule p') + unfolding split_conv + defer + apply (subst content_real) + using as(2) + apply auto + done + qed + have ***: "c - w < t \ t < c" + proof - + have "c - k < t" + using \k>0\ as(1) by (auto simp add: field_simps) + moreover have "k \ w" + apply (rule ccontr) + using k(2) + unfolding subset_eq + apply (erule_tac x="c + ((k + w)/2)" in ballE) + unfolding d_def + using \k > 0\ \w > 0\ + apply (auto simp add: field_simps not_le not_less dist_real_def) + done + ultimately show ?thesis using \t < c\ + by (auto simp add: field_simps) + qed + show ?thesis + unfolding *(1) + apply (subst *(2)) + apply (rule norm_triangle_lt add_strict_mono)+ + unfolding norm_minus_cancel + apply (rule d1_fin[unfolded **]) + apply (rule d2_fin) + using w(2)[OF ***] + unfolding norm_scaleR + apply (auto simp add: field_simps) + done + qed +qed + +lemma indefinite_integral_continuous_right: + fixes f :: "real \ 'a::banach" + assumes "f integrable_on {a .. b}" + and "a \ c" + and "c < b" + and "e > 0" + obtains d where "0 < d" + and "\t. c \ t \ t < c + d \ norm (integral {a .. c} f - integral {a .. t} f) < e" +proof - + have *: "(\x. f (- x)) integrable_on {-b .. -a}" "- b < - c" "- c \ - a" + using assms by auto + from indefinite_integral_continuous_left[OF * \e>0\] guess d . note d = this + let ?d = "min d (b - c)" + show ?thesis + apply (rule that[of "?d"]) + apply safe + proof - + show "0 < ?d" + using d(1) assms(3) by auto + fix t :: real + assume as: "c \ t" "t < c + ?d" + have *: "integral {a .. c} f = integral {a .. b} f - integral {c .. b} f" + "integral {a .. t} f = integral {a .. b} f - integral {t .. b} f" + apply (simp_all only: algebra_simps) + apply (rule_tac[!] integral_combine) + using assms as + apply auto + done + have "(- c) - d < (- t) \ - t \ - c" + using as by auto note d(2)[rule_format,OF this] + then show "norm (integral {a .. c} f - integral {a .. t} f) < e" + unfolding * + unfolding integral_reflect + apply (subst norm_minus_commute) + apply (auto simp add: algebra_simps) + done + qed +qed + +lemma indefinite_integral_continuous: + fixes f :: "real \ 'a::banach" + assumes "f integrable_on {a .. b}" + shows "continuous_on {a .. b} (\x. integral {a .. x} f)" +proof (unfold continuous_on_iff, safe) + fix x e :: real + assume as: "x \ {a .. b}" "e > 0" + let ?thesis = "\d>0. \x'\{a .. b}. dist x' x < d \ dist (integral {a .. x'} f) (integral {a .. x} f) < e" + { + presume *: "a < b \ ?thesis" + show ?thesis + apply cases + apply (rule *) + apply assumption + proof goal_cases + case 1 + then have "cbox a b = {x}" + using as(1) + apply - + apply (rule set_eqI) + apply auto + done + then show ?case using \e > 0\ by auto + qed + } + assume "a < b" + have "(x = a \ x = b) \ (a < x \ x < b)" + using as(1) by auto + then show ?thesis + apply (elim disjE) + proof - + assume "x = a" + have "a \ a" .. + from indefinite_integral_continuous_right[OF assms(1) this \a \e>0\] guess d . note d=this + show ?thesis + apply rule + apply rule + apply (rule d) + apply safe + apply (subst dist_commute) + unfolding \x = a\ dist_norm + apply (rule d(2)[rule_format]) + apply auto + done + next + assume "x = b" + have "b \ b" .. + from indefinite_integral_continuous_left[OF assms(1) \a this \e>0\] guess d . note d=this + show ?thesis + apply rule + apply rule + apply (rule d) + apply safe + apply (subst dist_commute) + unfolding \x = b\ dist_norm + apply (rule d(2)[rule_format]) + apply auto + done + next + assume "a < x \ x < b" + then have xl: "a < x" "x \ b" and xr: "a \ x" "x < b" + by auto + from indefinite_integral_continuous_left [OF assms(1) xl \e>0\] guess d1 . note d1=this + from indefinite_integral_continuous_right[OF assms(1) xr \e>0\] guess d2 . note d2=this + show ?thesis + apply (rule_tac x="min d1 d2" in exI) + proof safe + show "0 < min d1 d2" + using d1 d2 by auto + fix y + assume "y \ {a .. b}" and "dist y x < min d1 d2" + then show "dist (integral {a .. y} f) (integral {a .. x} f) < e" + apply (subst dist_commute) + apply (cases "y < x") + unfolding dist_norm + apply (rule d1(2)[rule_format]) + defer + apply (rule d2(2)[rule_format]) + unfolding not_less + apply (auto simp add: field_simps) + done + qed + qed +qed + + +subsection \This doesn't directly involve integration, but that gives an easy proof.\ + +lemma has_derivative_zero_unique_strong_interval: + fixes f :: "real \ 'a::banach" + assumes "finite k" + and "continuous_on {a .. b} f" + and "f a = y" + and "\x\({a .. b} - k). (f has_derivative (\h. 0)) (at x within {a .. b})" "x \ {a .. b}" + shows "f x = y" +proof - + have ab: "a \ b" + using assms by auto + have *: "a \ x" + using assms(5) by auto + have "((\x. 0::'a) has_integral f x - f a) {a .. x}" + apply (rule fundamental_theorem_of_calculus_interior_strong[OF assms(1) *]) + apply (rule continuous_on_subset[OF assms(2)]) + defer + apply safe + unfolding has_vector_derivative_def + apply (subst has_derivative_within_open[symmetric]) + apply assumption + apply (rule open_greaterThanLessThan) + apply (rule has_derivative_within_subset[where s="{a .. b}"]) + using assms(4) assms(5) + apply (auto simp: mem_box) + done + note this[unfolded *] + note has_integral_unique[OF has_integral_0 this] + then show ?thesis + unfolding assms by auto +qed + + +subsection \Generalize a bit to any convex set.\ + +lemma has_derivative_zero_unique_strong_convex: + fixes f :: "'a::euclidean_space \ 'b::banach" + assumes "convex s" + and "finite k" + and "continuous_on s f" + and "c \ s" + and "f c = y" + and "\x\(s - k). (f has_derivative (\h. 0)) (at x within s)" + and "x \ s" + shows "f x = y" +proof - + { + presume *: "x \ c \ ?thesis" + show ?thesis + apply cases + apply (rule *) + apply assumption + unfolding assms(5)[symmetric] + apply auto + done + } + assume "x \ c" + note conv = assms(1)[unfolded convex_alt,rule_format] + have as1: "continuous_on {0 ..1} (f \ (\t. (1 - t) *\<^sub>R c + t *\<^sub>R x))" + apply (rule continuous_intros)+ + apply (rule continuous_on_subset[OF assms(3)]) + apply safe + apply (rule conv) + using assms(4,7) + apply auto + done + have *: "t = xa" if "(1 - t) *\<^sub>R c + t *\<^sub>R x = (1 - xa) *\<^sub>R c + xa *\<^sub>R x" for t xa + proof - + from that have "(t - xa) *\<^sub>R x = (t - xa) *\<^sub>R c" + unfolding scaleR_simps by (auto simp add: algebra_simps) + then show ?thesis + using \x \ c\ by auto + qed + have as2: "finite {t. ((1 - t) *\<^sub>R c + t *\<^sub>R x) \ k}" + using assms(2) + apply (rule finite_surj[where f="\z. SOME t. (1-t) *\<^sub>R c + t *\<^sub>R x = z"]) + apply safe + unfolding image_iff + apply rule + defer + apply assumption + apply (rule sym) + apply (rule some_equality) + defer + apply (drule *) + apply auto + done + have "(f \ (\t. (1 - t) *\<^sub>R c + t *\<^sub>R x)) 1 = y" + apply (rule has_derivative_zero_unique_strong_interval[OF as2 as1, of ]) + unfolding o_def + using assms(5) + defer + apply - + apply rule + proof - + fix t + assume as: "t \ {0 .. 1} - {t. (1 - t) *\<^sub>R c + t *\<^sub>R x \ k}" + have *: "c - t *\<^sub>R c + t *\<^sub>R x \ s - k" + apply safe + apply (rule conv[unfolded scaleR_simps]) + using \x \ s\ \c \ s\ as + by (auto simp add: algebra_simps) + have "(f \ (\t. (1 - t) *\<^sub>R c + t *\<^sub>R x) has_derivative (\x. 0) \ (\z. (0 - z *\<^sub>R c) + z *\<^sub>R x)) + (at t within {0 .. 1})" + apply (intro derivative_eq_intros) + apply simp_all + apply (simp add: field_simps) + unfolding scaleR_simps + apply (rule has_derivative_within_subset,rule assms(6)[rule_format]) + apply (rule *) + apply safe + apply (rule conv[unfolded scaleR_simps]) + using \x \ s\ \c \ s\ + apply auto + done + then show "((\xa. f ((1 - xa) *\<^sub>R c + xa *\<^sub>R x)) has_derivative (\h. 0)) (at t within {0 .. 1})" + unfolding o_def . + qed auto + then show ?thesis + by auto +qed + + +text \Also to any open connected set with finite set of exceptions. Could + generalize to locally convex set with limpt-free set of exceptions.\ + +lemma has_derivative_zero_unique_strong_connected: + fixes f :: "'a::euclidean_space \ 'b::banach" + assumes "connected s" + and "open s" + and "finite k" + and "continuous_on s f" + and "c \ s" + and "f c = y" + and "\x\(s - k). (f has_derivative (\h. 0)) (at x within s)" + and "x\s" + shows "f x = y" +proof - + have "{x \ s. f x \ {y}} = {} \ {x \ s. f x \ {y}} = s" + apply (rule assms(1)[unfolded connected_clopen,rule_format]) + apply rule + defer + apply (rule continuous_closedin_preimage[OF assms(4) closed_singleton]) + apply (rule open_openin_trans[OF assms(2)]) + unfolding open_contains_ball + proof safe + fix x + assume "x \ s" + from assms(2)[unfolded open_contains_ball,rule_format,OF this] guess e .. note e=conjunctD2[OF this] + show "\e>0. ball x e \ {xa \ s. f xa \ {f x}}" + apply rule + apply rule + apply (rule e) + proof safe + fix y + assume y: "y \ ball x e" + then show "y \ s" + using e by auto + show "f y = f x" + apply (rule has_derivative_zero_unique_strong_convex[OF convex_ball]) + apply (rule assms) + apply (rule continuous_on_subset) + apply (rule assms) + apply (rule e)+ + apply (subst centre_in_ball) + apply (rule e) + apply rule + apply safe + apply (rule has_derivative_within_subset) + apply (rule assms(7)[rule_format]) + using y e + apply auto + done + qed + qed + then show ?thesis + using \x \ s\ \f c = y\ \c \ s\ by auto +qed + +lemma has_derivative_zero_connected_constant: + fixes f :: "'a::euclidean_space \ 'b::banach" + assumes "connected s" + and "open s" + and "finite k" + and "continuous_on s f" + and "\x\(s - k). (f has_derivative (\h. 0)) (at x within s)" + obtains c where "\x. x \ s \ f(x) = c" +proof (cases "s = {}") + case True + then show ?thesis +by (metis empty_iff that) +next + case False + then obtain c where "c \ s" + by (metis equals0I) + then show ?thesis + by (metis has_derivative_zero_unique_strong_connected assms that) +qed + + +subsection \Integrating characteristic function of an interval\ + +lemma has_integral_restrict_open_subinterval: + fixes f :: "'a::euclidean_space \ 'b::banach" + assumes "(f has_integral i) (cbox c d)" + and "cbox c d \ cbox a b" + shows "((\x. if x \ box c d then f x else 0) has_integral i) (cbox a b)" +proof - + define g where [abs_def]: "g x = (if x \box c d then f x else 0)" for x + { + presume *: "cbox c d \ {} \ ?thesis" + show ?thesis + apply cases + apply (rule *) + apply assumption + proof goal_cases + case prems: 1 + then have *: "box c d = {}" + by (metis bot.extremum_uniqueI box_subset_cbox) + show ?thesis + using assms(1) + unfolding * + using prems + by auto + qed + } + assume "cbox c d \ {}" + from partial_division_extend_1[OF assms(2) this] guess p . note p=this + note operat = comm_monoid_set.operative_division[OF add.comm_monoid_set_lift_option operative_integral p(1), symmetric] + let ?P = "(if g integrable_on cbox a b then Some (integral (cbox a b) g) else None) = Some i" + { + presume "?P" + then have "g integrable_on cbox a b \ integral (cbox a b) g = i" + apply - + apply cases + apply (subst(asm) if_P) + apply assumption + apply auto + done + then show ?thesis + using integrable_integral + unfolding g_def + by auto + } + let ?F = "comm_monoid_set.F (lift_option op +) (Some 0)" + have iterate:"?F (\i. if g integrable_on i then Some (integral i g) else None) (p - {cbox c d}) = Some 0" + proof (intro comm_monoid_set.neutral[OF add.comm_monoid_set_lift_option] ballI) + fix x + assume x: "x \ p - {cbox c d}" + then have "x \ p" + by auto + note div = division_ofD(2-5)[OF p(1) this] + from div(3) guess u v by (elim exE) note uv=this + have "interior x \ interior (cbox c d) = {}" + using div(4)[OF p(2)] x by auto + then have "(g has_integral 0) x" + unfolding uv + apply - + apply (rule has_integral_spike_interior[where f="\x. 0"]) + unfolding g_def interior_cbox + apply auto + done + then show "(if g integrable_on x then Some (integral x g) else None) = Some 0" + by auto + qed + + have *: "p = insert (cbox c d) (p - {cbox c d})" + using p by auto + have **: "g integrable_on cbox c d" + apply (rule integrable_spike_interior[where f=f]) + unfolding g_def using assms(1) + apply auto + done + moreover + have "integral (cbox c d) g = i" + apply (rule has_integral_unique[OF _ assms(1)]) + apply (rule has_integral_spike_interior[where f=g]) + defer + apply (rule integrable_integral[OF **]) + unfolding g_def + apply auto + done + ultimately show ?P + unfolding operat + using p + apply (subst *) + apply (subst comm_monoid_set.insert[OF add.comm_monoid_set_lift_option]) + apply (simp_all add: division_of_finite iterate) + done +qed + +lemma has_integral_restrict_closed_subinterval: + fixes f :: "'a::euclidean_space \ 'b::banach" + assumes "(f has_integral i) (cbox c d)" + and "cbox c d \ cbox a b" + shows "((\x. if x \ cbox c d then f x else 0) has_integral i) (cbox a b)" +proof - + note has_integral_restrict_open_subinterval[OF assms] + note * = has_integral_spike[OF negligible_frontier_interval _ this] + show ?thesis + apply (rule *[of c d]) + using box_subset_cbox[of c d] + apply auto + done +qed + +lemma has_integral_restrict_closed_subintervals_eq: + fixes f :: "'a::euclidean_space \ 'b::banach" + assumes "cbox c d \ cbox a b" + shows "((\x. if x \ cbox c d then f x else 0) has_integral i) (cbox a b) \ (f has_integral i) (cbox c d)" + (is "?l = ?r") +proof (cases "cbox c d = {}") + case False + let ?g = "\x. if x \ cbox c d then f x else 0" + show ?thesis + apply rule + defer + apply (rule has_integral_restrict_closed_subinterval[OF _ assms]) + apply assumption + proof - + assume ?l + then have "?g integrable_on cbox c d" + using assms has_integral_integrable integrable_subinterval by blast + then have *: "f integrable_on cbox c d" + apply - + apply (rule integrable_eq) + apply auto + done + then have "i = integral (cbox c d) f" + apply - + apply (rule has_integral_unique) + apply (rule \?l\) + apply (rule has_integral_restrict_closed_subinterval[OF _ assms]) + apply auto + done + then show ?r + using * by auto + qed +qed auto + + +text \Hence we can apply the limit process uniformly to all integrals.\ + +lemma has_integral': + fixes f :: "'n::euclidean_space \ 'a::banach" + shows "(f has_integral i) s \ + (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ + (\z. ((\x. if x \ s then f(x) else 0) has_integral z) (cbox a b) \ norm(z - i) < e))" + (is "?l \ (\e>0. ?r e)") +proof - + { + presume *: "\a b. s = cbox a b \ ?thesis" + show ?thesis + apply cases + apply (rule *) + apply assumption + apply (subst has_integral_alt) + apply auto + done + } + assume "\a b. s = cbox a b" + then guess a b by (elim exE) note s=this + from bounded_cbox[of a b, unfolded bounded_pos] guess B .. + note B = conjunctD2[OF this,rule_format] show ?thesis + apply safe + proof - + fix e :: real + assume ?l and "e > 0" + show "?r e" + apply (rule_tac x="B+1" in exI) + apply safe + defer + apply (rule_tac x=i in exI) + proof + fix c d :: 'n + assume as: "ball 0 (B+1) \ cbox c d" + then show "((\x. if x \ s then f x else 0) has_integral i) (cbox c d)" + unfolding s + apply - + apply (rule has_integral_restrict_closed_subinterval) + apply (rule \?l\[unfolded s]) + apply safe + apply (drule B(2)[rule_format]) + unfolding subset_eq + apply (erule_tac x=x in ballE) + apply (auto simp add: dist_norm) + done + qed (insert B \e>0\, auto) + next + assume as: "\e>0. ?r e" + from this[rule_format,OF zero_less_one] guess C .. note C=conjunctD2[OF this,rule_format] + define c :: 'n where "c = (\i\Basis. (- max B C) *\<^sub>R i)" + define d :: 'n where "d = (\i\Basis. max B C *\<^sub>R i)" + have c_d: "cbox a b \ cbox c d" + apply safe + apply (drule B(2)) + unfolding mem_box + proof + fix x i + show "c \ i \ x \ i \ x \ i \ d \ i" if "norm x \ B" and "i \ Basis" + using that and Basis_le_norm[OF \i\Basis\, of x] + unfolding c_def d_def + by (auto simp add: field_simps setsum_negf) + qed + have "ball 0 C \ cbox c d" + apply (rule subsetI) + unfolding mem_box mem_ball dist_norm + proof + fix x i :: 'n + assume x: "norm (0 - x) < C" and i: "i \ Basis" + show "c \ i \ x \ i \ x \ i \ d \ i" + using Basis_le_norm[OF i, of x] and x i + unfolding c_def d_def + by (auto simp: setsum_negf) + qed + from C(2)[OF this] have "\y. (f has_integral y) (cbox a b)" + unfolding has_integral_restrict_closed_subintervals_eq[OF c_d,symmetric] + unfolding s + by auto + then guess y .. note y=this + + have "y = i" + proof (rule ccontr) + assume "\ ?thesis" + then have "0 < norm (y - i)" + by auto + from as[rule_format,OF this] guess C .. note C=conjunctD2[OF this,rule_format] + define c :: 'n where "c = (\i\Basis. (- max B C) *\<^sub>R i)" + define d :: 'n where "d = (\i\Basis. max B C *\<^sub>R i)" + have c_d: "cbox a b \ cbox c d" + apply safe + apply (drule B(2)) + unfolding mem_box + proof + fix x i :: 'n + assume "norm x \ B" and "i \ Basis" + then show "c \ i \ x \ i \ x \ i \ d \ i" + using Basis_le_norm[of i x] + unfolding c_def d_def + by (auto simp add: field_simps setsum_negf) + qed + have "ball 0 C \ cbox c d" + apply (rule subsetI) + unfolding mem_box mem_ball dist_norm + proof + fix x i :: 'n + assume "norm (0 - x) < C" and "i \ Basis" + then show "c \ i \ x \ i \ x \ i \ d \ i" + using Basis_le_norm[of i x] + unfolding c_def d_def + by (auto simp: setsum_negf) + qed + note C(2)[OF this] then guess z .. note z = conjunctD2[OF this, unfolded s] + note this[unfolded has_integral_restrict_closed_subintervals_eq[OF c_d]] + then have "z = y" and "norm (z - i) < norm (y - i)" + apply - + apply (rule has_integral_unique[OF _ y(1)]) + apply assumption + apply assumption + done + then show False + by auto + qed + then show ?l + using y + unfolding s + by auto + qed +qed + +lemma has_integral_le: + fixes f :: "'n::euclidean_space \ real" + assumes "(f has_integral i) s" + and "(g has_integral j) s" + and "\x\s. f x \ g x" + shows "i \ j" + using has_integral_component_le[OF _ assms(1-2), of 1] + using assms(3) + by auto + +lemma integral_le: + fixes f :: "'n::euclidean_space \ real" + assumes "f integrable_on s" + and "g integrable_on s" + and "\x\s. f x \ g x" + shows "integral s f \ integral s g" + by (rule has_integral_le[OF assms(1,2)[unfolded has_integral_integral] assms(3)]) + +lemma has_integral_nonneg: + fixes f :: "'n::euclidean_space \ real" + assumes "(f has_integral i) s" + and "\x\s. 0 \ f x" + shows "0 \ i" + using has_integral_component_nonneg[of 1 f i s] + unfolding o_def + using assms + by auto + +lemma integral_nonneg: + fixes f :: "'n::euclidean_space \ real" + assumes "f integrable_on s" + and "\x\s. 0 \ f x" + shows "0 \ integral s f" + by (rule has_integral_nonneg[OF assms(1)[unfolded has_integral_integral] assms(2)]) + + +text \Hence a general restriction property.\ + +lemma has_integral_restrict[simp]: + assumes "s \ t" + shows "((\x. if x \ s then f x else (0::'a::banach)) has_integral i) t \ (f has_integral i) s" +proof - + have *: "\x. (if x \ t then if x \ s then f x else 0 else 0) = (if x\s then f x else 0)" + using assms by auto + show ?thesis + apply (subst(2) has_integral') + apply (subst has_integral') + unfolding * + apply rule + done +qed + +lemma has_integral_restrict_univ: + fixes f :: "'n::euclidean_space \ 'a::banach" + shows "((\x. if x \ s then f x else 0) has_integral i) UNIV \ (f has_integral i) s" + by auto + +lemma has_integral_on_superset: + fixes f :: "'n::euclidean_space \ 'a::banach" + assumes "\x. x \ s \ f x = 0" + and "s \ t" + and "(f has_integral i) s" + shows "(f has_integral i) t" +proof - + have "(\x. if x \ s then f x else 0) = (\x. if x \ t then f x else 0)" + apply rule + using assms(1-2) + apply auto + done + then show ?thesis + using assms(3) + apply (subst has_integral_restrict_univ[symmetric]) + apply (subst(asm) has_integral_restrict_univ[symmetric]) + apply auto + done +qed + +lemma integrable_on_superset: + fixes f :: "'n::euclidean_space \ 'a::banach" + assumes "\x. x \ s \ f x = 0" + and "s \ t" + and "f integrable_on s" + shows "f integrable_on t" + using assms + unfolding integrable_on_def + by (auto intro:has_integral_on_superset) + +lemma integral_restrict_univ[intro]: + fixes f :: "'n::euclidean_space \ 'a::banach" + shows "f integrable_on s \ integral UNIV (\x. if x \ s then f x else 0) = integral s f" + apply (rule integral_unique) + unfolding has_integral_restrict_univ + apply auto + done + +lemma integrable_restrict_univ: + fixes f :: "'n::euclidean_space \ 'a::banach" + shows "(\x. if x \ s then f x else 0) integrable_on UNIV \ f integrable_on s" + unfolding integrable_on_def + by auto + +lemma negligible_on_intervals: "negligible s \ (\a b. negligible(s \ cbox a b))" (is "?l \ ?r") +proof + assume ?r + show ?l + unfolding negligible_def + proof safe + fix a b + show "(indicator s has_integral 0) (cbox a b)" + apply (rule has_integral_negligible[OF \?r\[rule_format,of a b]]) + unfolding indicator_def + apply auto + done + qed +qed auto + +lemma negligible_translation: + assumes "negligible S" + shows "negligible (op + c ` S)" +proof - + have inj: "inj (op + c)" + by simp + show ?thesis + using assms + proof (clarsimp simp: negligible_def) + fix a b + assume "\x y. (indicator S has_integral 0) (cbox x y)" + then have *: "(indicator S has_integral 0) (cbox (a-c) (b-c))" + by (meson Diff_iff assms has_integral_negligible indicator_simps(2)) + have eq: "indicator (op + c ` S) = (\x. indicator S (x - c))" + by (force simp add: indicator_def) + show "(indicator (op + c ` S) has_integral 0) (cbox a b)" + using has_integral_affinity [OF *, of 1 "-c"] + cbox_translation [of "c" "-c+a" "-c+b"] + by (simp add: eq add.commute) + qed +qed + +lemma negligible_translation_rev: + assumes "negligible (op + c ` S)" + shows "negligible S" +by (metis negligible_translation [OF assms, of "-c"] translation_galois) + +lemma has_integral_spike_set_eq: + fixes f :: "'n::euclidean_space \ 'a::banach" + assumes "negligible ((s - t) \ (t - s))" + shows "(f has_integral y) s \ (f has_integral y) t" + unfolding has_integral_restrict_univ[symmetric,of f] + apply (rule has_integral_spike_eq[OF assms]) + by (auto split: if_split_asm) + +lemma has_integral_spike_set[dest]: + fixes f :: "'n::euclidean_space \ 'a::banach" + assumes "negligible ((s - t) \ (t - s))" + and "(f has_integral y) s" + shows "(f has_integral y) t" + using assms has_integral_spike_set_eq + by auto + +lemma integrable_spike_set[dest]: + fixes f :: "'n::euclidean_space \ 'a::banach" + assumes "negligible ((s - t) \ (t - s))" + and "f integrable_on s" + shows "f integrable_on t" + using assms(2) + unfolding integrable_on_def + unfolding has_integral_spike_set_eq[OF assms(1)] . + +lemma integrable_spike_set_eq: + fixes f :: "'n::euclidean_space \ 'a::banach" + assumes "negligible ((s - t) \ (t - s))" + shows "f integrable_on s \ f integrable_on t" + apply rule + apply (rule_tac[!] integrable_spike_set) + using assms + apply auto + done + +(*lemma integral_spike_set: + "\f:real^M->real^N g s t. + negligible(s DIFF t \ t DIFF s) + \ integral s f = integral t f" +qed REPEAT STRIP_TAC THEN REWRITE_TAC[integral] THEN + AP_TERM_TAC THEN ABS_TAC THEN MATCH_MP_TAC HAS_INTEGRAL_SPIKE_SET_EQ THEN + ASM_MESON_TAC[]);; + +lemma has_integral_interior: + "\f:real^M->real^N y s. + negligible(frontier s) + \ ((f has_integral y) (interior s) \ (f has_integral y) s)" +qed REPEAT STRIP_TAC THEN MATCH_MP_TAC HAS_INTEGRAL_SPIKE_SET_EQ THEN + FIRST_X_ASSUM(MATCH_MP_TAC o MATCH_MP (REWRITE_RULE[IMP_CONJ] + NEGLIGIBLE_SUBSET)) THEN + REWRITE_TAC[frontier] THEN + MP_TAC(ISPEC `s:real^M->bool` INTERIOR_SUBSET) THEN + MP_TAC(ISPEC `s:real^M->bool` CLOSURE_SUBSET) THEN + SET_TAC[]);; + +lemma has_integral_closure: + "\f:real^M->real^N y s. + negligible(frontier s) + \ ((f has_integral y) (closure s) \ (f has_integral y) s)" +qed REPEAT STRIP_TAC THEN MATCH_MP_TAC HAS_INTEGRAL_SPIKE_SET_EQ THEN + FIRST_X_ASSUM(MATCH_MP_TAC o MATCH_MP (REWRITE_RULE[IMP_CONJ] + NEGLIGIBLE_SUBSET)) THEN + REWRITE_TAC[frontier] THEN + MP_TAC(ISPEC `s:real^M->bool` INTERIOR_SUBSET) THEN + MP_TAC(ISPEC `s:real^M->bool` CLOSURE_SUBSET) THEN + SET_TAC[]);;*) + + +subsection \More lemmas that are useful later\ + +lemma has_integral_subset_component_le: + fixes f :: "'n::euclidean_space \ 'm::euclidean_space" + assumes k: "k \ Basis" + and as: "s \ t" "(f has_integral i) s" "(f has_integral j) t" "\x\t. 0 \ f(x)\k" + shows "i\k \ j\k" +proof - + note has_integral_restrict_univ[symmetric, of f] + note as(2-3)[unfolded this] note * = has_integral_component_le[OF k this] + show ?thesis + apply (rule *) + using as(1,4) + apply auto + done +qed + +lemma has_integral_subset_le: + fixes f :: "'n::euclidean_space \ real" + assumes "s \ t" + and "(f has_integral i) s" + and "(f has_integral j) t" + and "\x\t. 0 \ f x" + shows "i \ j" + using has_integral_subset_component_le[OF _ assms(1), of 1 f i j] + using assms + by auto + +lemma integral_subset_component_le: + fixes f :: "'n::euclidean_space \ 'm::euclidean_space" + assumes "k \ Basis" + and "s \ t" + and "f integrable_on s" + and "f integrable_on t" + and "\x \ t. 0 \ f x \ k" + shows "(integral s f)\k \ (integral t f)\k" + apply (rule has_integral_subset_component_le) + using assms + apply auto + done + +lemma integral_subset_le: + fixes f :: "'n::euclidean_space \ real" + assumes "s \ t" + and "f integrable_on s" + and "f integrable_on t" + and "\x \ t. 0 \ f x" + shows "integral s f \ integral t f" + apply (rule has_integral_subset_le) + using assms + apply auto + done + +lemma has_integral_alt': + fixes f :: "'n::euclidean_space \ 'a::banach" + shows "(f has_integral i) s \ (\a b. (\x. if x \ s then f x else 0) integrable_on cbox a b) \ + (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ + norm (integral (cbox a b) (\x. if x \ s then f x else 0) - i) < e)" + (is "?l = ?r") +proof + assume ?r + show ?l + apply (subst has_integral') + apply safe + proof goal_cases + case (1 e) + from \?r\[THEN conjunct2,rule_format,OF this] guess B .. note B=conjunctD2[OF this] + show ?case + apply rule + apply rule + apply (rule B) + apply safe + apply (rule_tac x="integral (cbox a b) (\x. if x \ s then f x else 0)" in exI) + apply (drule B(2)[rule_format]) + using integrable_integral[OF \?r\[THEN conjunct1,rule_format]] + apply auto + done + qed +next + assume ?l note as = this[unfolded has_integral'[of f],rule_format] + let ?f = "\x. if x \ s then f x else 0" + show ?r + proof safe + fix a b :: 'n + from as[OF zero_less_one] guess B .. note B=conjunctD2[OF this,rule_format] + let ?a = "\i\Basis. min (a\i) (-B) *\<^sub>R i::'n" + let ?b = "\i\Basis. max (b\i) B *\<^sub>R i::'n" + show "?f integrable_on cbox a b" + proof (rule integrable_subinterval[of _ ?a ?b]) + have "ball 0 B \ cbox ?a ?b" + apply (rule subsetI) + unfolding mem_ball mem_box dist_norm + proof (rule, goal_cases) + case (1 x i) + then show ?case using Basis_le_norm[of i x] + by (auto simp add:field_simps) + qed + from B(2)[OF this] guess z .. note conjunct1[OF this] + then show "?f integrable_on cbox ?a ?b" + unfolding integrable_on_def by auto + show "cbox a b \ cbox ?a ?b" + apply safe + unfolding mem_box + apply rule + apply (erule_tac x=i in ballE) + apply auto + done + qed + + fix e :: real + assume "e > 0" + from as[OF this] guess B .. note B=conjunctD2[OF this,rule_format] + show "\B>0. \a b. ball 0 B \ cbox a b \ + norm (integral (cbox a b) (\x. if x \ s then f x else 0) - i) < e" + apply rule + apply rule + apply (rule B) + apply safe + proof goal_cases + case 1 + from B(2)[OF this] guess z .. note z=conjunctD2[OF this] + from integral_unique[OF this(1)] show ?case + using z(2) by auto + qed + qed +qed + + +subsection \Continuity of the integral (for a 1-dimensional interval).\ + +lemma integrable_alt: + fixes f :: "'n::euclidean_space \ 'a::banach" + shows "f integrable_on s \ + (\a b. (\x. if x \ s then f x else 0) integrable_on cbox a b) \ + (\e>0. \B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ + norm (integral (cbox a b) (\x. if x \ s then f x else 0) - + integral (cbox c d) (\x. if x \ s then f x else 0)) < e)" + (is "?l = ?r") +proof + assume ?l + then guess y unfolding integrable_on_def .. note this[unfolded has_integral_alt'[of f]] + note y=conjunctD2[OF this,rule_format] + show ?r + apply safe + apply (rule y) + proof goal_cases + case (1 e) + then have "e/2 > 0" + by auto + from y(2)[OF this] guess B .. note B=conjunctD2[OF this,rule_format] + show ?case + apply rule + apply rule + apply (rule B) + apply safe + proof goal_cases + case prems: (1 a b c d) + show ?case + apply (rule norm_triangle_half_l) + using B(2)[OF prems(1)] B(2)[OF prems(2)] + apply auto + done + qed + qed +next + assume ?r + note as = conjunctD2[OF this,rule_format] + let ?cube = "\n. cbox (\i\Basis. - real n *\<^sub>R i::'n) (\i\Basis. real n *\<^sub>R i)" + have "Cauchy (\n. integral (?cube n) (\x. if x \ s then f x else 0))" + proof (unfold Cauchy_def, safe, goal_cases) + case (1 e) + from as(2)[OF this] guess B .. note B = conjunctD2[OF this,rule_format] + from real_arch_simple[of B] guess N .. note N = this + { + fix n + assume n: "n \ N" + have "ball 0 B \ ?cube n" + apply (rule subsetI) + unfolding mem_ball mem_box dist_norm + proof (rule, goal_cases) + case (1 x i) + then show ?case + using Basis_le_norm[of i x] \i\Basis\ + using n N + by (auto simp add: field_simps setsum_negf) + qed + } + then show ?case + apply - + apply (rule_tac x=N in exI) + apply safe + unfolding dist_norm + apply (rule B(2)) + apply auto + done + qed + from this[unfolded convergent_eq_cauchy[symmetric]] guess i .. + note i = this[THEN LIMSEQ_D] + + show ?l unfolding integrable_on_def has_integral_alt'[of f] + apply (rule_tac x=i in exI) + apply safe + apply (rule as(1)[unfolded integrable_on_def]) + proof goal_cases + case (1 e) + then have *: "e/2 > 0" by auto + from i[OF this] guess N .. note N =this[rule_format] + from as(2)[OF *] guess B .. note B=conjunctD2[OF this,rule_format] + let ?B = "max (real N) B" + show ?case + apply (rule_tac x="?B" in exI) + proof safe + show "0 < ?B" + using B(1) by auto + fix a b :: 'n + assume ab: "ball 0 ?B \ cbox a b" + from real_arch_simple[of ?B] guess n .. note n=this + show "norm (integral (cbox a b) (\x. if x \ s then f x else 0) - i) < e" + apply (rule norm_triangle_half_l) + apply (rule B(2)) + defer + apply (subst norm_minus_commute) + apply (rule N[of n]) + proof safe + show "N \ n" + using n by auto + fix x :: 'n + assume x: "x \ ball 0 B" + then have "x \ ball 0 ?B" + by auto + then show "x \ cbox a b" + using ab by blast + show "x \ ?cube n" + using x + unfolding mem_box mem_ball dist_norm + apply - + proof (rule, goal_cases) + case (1 i) + then show ?case + using Basis_le_norm[of i x] \i \ Basis\ + using n + by (auto simp add: field_simps setsum_negf) + qed + qed + qed + qed +qed + +lemma integrable_altD: + fixes f :: "'n::euclidean_space \ 'a::banach" + assumes "f integrable_on s" + shows "\a b. (\x. if x \ s then f x else 0) integrable_on cbox a b" + and "\e. e > 0 \ \B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ + norm (integral (cbox a b) (\x. if x \ s then f x else 0) - integral (cbox c d) (\x. if x \ s then f x else 0)) < e" + using assms[unfolded integrable_alt[of f]] by auto + +lemma integrable_on_subcbox: + fixes f :: "'n::euclidean_space \ 'a::banach" + assumes "f integrable_on s" + and "cbox a b \ s" + shows "f integrable_on cbox a b" + apply (rule integrable_eq) + defer + apply (rule integrable_altD(1)[OF assms(1)]) + using assms(2) + apply auto + done + + +subsection \A straddling criterion for integrability\ + +lemma integrable_straddle_interval: + fixes f :: "'n::euclidean_space \ real" + assumes "\e>0. \g h i j. (g has_integral i) (cbox a b) \ (h has_integral j) (cbox a b) \ + norm (i - j) < e \ (\x\cbox a b. (g x) \ f x \ f x \ h x)" + shows "f integrable_on cbox a b" +proof (subst integrable_cauchy, safe, goal_cases) + case (1 e) + then have e: "e/3 > 0" + by auto + note assms[rule_format,OF this] + then guess g h i j by (elim exE conjE) note obt = this + from obt(1)[unfolded has_integral[of g], rule_format, OF e] guess d1 .. note d1=conjunctD2[OF this,rule_format] + from obt(2)[unfolded has_integral[of h], rule_format, OF e] guess d2 .. note d2=conjunctD2[OF this,rule_format] + show ?case + apply (rule_tac x="\x. d1 x \ d2 x" in exI) + apply (rule conjI gauge_inter d1 d2)+ + unfolding fine_inter + proof (safe, goal_cases) + have **: "\i j g1 g2 h1 h2 f1 f2. g1 - h2 \ f1 - f2 \ f1 - f2 \ h1 - g2 \ + \i - j\ < e / 3 \ \g2 - i\ < e / 3 \ \g1 - i\ < e / 3 \ + \h2 - j\ < e / 3 \ \h1 - j\ < e / 3 \ \f1 - f2\ < e" + using \e > 0\ by arith + case prems: (1 p1 p2) + note tagged_division_ofD(2-4) note * = this[OF prems(1)] this[OF prems(4)] + + have "(\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p1. content k *\<^sub>R g x) \ 0" + and "0 \ (\(x, k)\p2. content k *\<^sub>R h x) - (\(x, k)\p2. content k *\<^sub>R f x)" + and "(\(x, k)\p2. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R g x) \ 0" + and "0 \ (\(x, k)\p1. content k *\<^sub>R h x) - (\(x, k)\p1. content k *\<^sub>R f x)" + unfolding setsum_subtractf[symmetric] + apply - + apply (rule_tac[!] setsum_nonneg) + apply safe + unfolding real_scaleR_def right_diff_distrib[symmetric] + apply (rule_tac[!] mult_nonneg_nonneg) + proof - + fix a b + assume ab: "(a, b) \ p1" + show "0 \ content b" + using *(3)[OF ab] + apply safe + apply (rule content_pos_le) + done + then show "0 \ content b" . + show "0 \ f a - g a" "0 \ h a - f a" + using *(1-2)[OF ab] + using obt(4)[rule_format,of a] + by auto + next + fix a b + assume ab: "(a, b) \ p2" + show "0 \ content b" + using *(6)[OF ab] + apply safe + apply (rule content_pos_le) + done + then show "0 \ content b" . + show "0 \ f a - g a" and "0 \ h a - f a" + using *(4-5)[OF ab] using obt(4)[rule_format,of a] by auto + qed + then show ?case + apply - + unfolding real_norm_def + apply (rule **) + defer + defer + unfolding real_norm_def[symmetric] + apply (rule obt(3)) + apply (rule d1(2)[OF conjI[OF prems(4,5)]]) + apply (rule d1(2)[OF conjI[OF prems(1,2)]]) + apply (rule d2(2)[OF conjI[OF prems(4,6)]]) + apply (rule d2(2)[OF conjI[OF prems(1,3)]]) + apply auto + done + qed +qed + +lemma integrable_straddle: + fixes f :: "'n::euclidean_space \ real" + assumes "\e>0. \g h i j. (g has_integral i) s \ (h has_integral j) s \ + norm (i - j) < e \ (\x\s. g x \ f x \ f x \ h x)" + shows "f integrable_on s" +proof - + have "\a b. (\x. if x \ s then f x else 0) integrable_on cbox a b" + proof (rule integrable_straddle_interval, safe, goal_cases) + case (1 a b e) + then have *: "e/4 > 0" + by auto + from assms[rule_format,OF this] guess g h i j by (elim exE conjE) note obt=this + note obt(1)[unfolded has_integral_alt'[of g]] + note conjunctD2[OF this, rule_format] + note g = this(1) and this(2)[OF *] + from this(2) guess B1 .. note B1 = conjunctD2[OF this,rule_format] + note obt(2)[unfolded has_integral_alt'[of h]] + note conjunctD2[OF this, rule_format] + note h = this(1) and this(2)[OF *] + from this(2) guess B2 .. note B2 = conjunctD2[OF this,rule_format] + define c :: 'n where "c = (\i\Basis. min (a\i) (- (max B1 B2)) *\<^sub>R i)" + define d :: 'n where "d = (\i\Basis. max (b\i) (max B1 B2) *\<^sub>R i)" + have *: "ball 0 B1 \ cbox c d" "ball 0 B2 \ cbox c d" + apply safe + unfolding mem_ball mem_box dist_norm + apply (rule_tac[!] ballI) + proof goal_cases + case (1 x i) + then show ?case using Basis_le_norm[of i x] + unfolding c_def d_def by auto + next + case (2 x i) + then show ?case using Basis_le_norm[of i x] + unfolding c_def d_def by auto + qed + have **: "\ch cg ag ah::real. norm (ah - ag) \ norm (ch - cg) \ norm (cg - i) < e / 4 \ + norm (ch - j) < e / 4 \ norm (ag - ah) < e" + using obt(3) + unfolding real_norm_def + by arith + show ?case + apply (rule_tac x="\x. if x \ s then g x else 0" in exI) + apply (rule_tac x="\x. if x \ s then h x else 0" in exI) + apply (rule_tac x="integral (cbox a b) (\x. if x \ s then g x else 0)" in exI) + apply (rule_tac x="integral (cbox a b) (\x. if x \ s then h x else 0)" in exI) + apply safe + apply (rule_tac[1-2] integrable_integral,rule g) + apply (rule h) + apply (rule **[OF _ B1(2)[OF *(1)] B2(2)[OF *(2)]]) + proof - + have *: "\x f g. (if x \ s then f x else 0) - (if x \ s then g x else 0) = + (if x \ s then f x - g x else (0::real))" + by auto + note ** = abs_of_nonneg[OF integral_nonneg[OF integrable_diff, OF h g]] + show "norm (integral (cbox a b) (\x. if x \ s then h x else 0) - + integral (cbox a b) (\x. if x \ s then g x else 0)) \ + norm (integral (cbox c d) (\x. if x \ s then h x else 0) - + integral (cbox c d) (\x. if x \ s then g x else 0))" + unfolding integral_diff[OF h g,symmetric] real_norm_def + apply (subst **) + defer + apply (subst **) + defer + apply (rule has_integral_subset_le) + defer + apply (rule integrable_integral integrable_diff h g)+ + proof safe + fix x + assume "x \ cbox a b" + then show "x \ cbox c d" + unfolding mem_box c_def d_def + apply - + apply rule + apply (erule_tac x=i in ballE) + apply auto + done + qed (insert obt(4), auto) + qed (insert obt(4), auto) + qed + note interv = this + + show ?thesis + unfolding integrable_alt[of f] + apply safe + apply (rule interv) + proof goal_cases + case (1 e) + then have *: "e/3 > 0" + by auto + from assms[rule_format,OF this] guess g h i j by (elim exE conjE) note obt=this + note obt(1)[unfolded has_integral_alt'[of g]] + note conjunctD2[OF this, rule_format] + note g = this(1) and this(2)[OF *] + from this(2) guess B1 .. note B1 = conjunctD2[OF this,rule_format] + note obt(2)[unfolded has_integral_alt'[of h]] + note conjunctD2[OF this, rule_format] + note h = this(1) and this(2)[OF *] + from this(2) guess B2 .. note B2 = conjunctD2[OF this,rule_format] + show ?case + apply (rule_tac x="max B1 B2" in exI) + apply safe + apply (rule max.strict_coboundedI1) + apply (rule B1) + proof - + fix a b c d :: 'n + assume as: "ball 0 (max B1 B2) \ cbox a b" "ball 0 (max B1 B2) \ cbox c d" + have **: "ball 0 B1 \ ball (0::'n) (max B1 B2)" "ball 0 B2 \ ball (0::'n) (max B1 B2)" + by auto + have *: "\ga gc ha hc fa fc::real. + \ga - i\ < e / 3 \ \gc - i\ < e / 3 \ \ha - j\ < e / 3 \ + \hc - j\ < e / 3 \ \i - j\ < e / 3 \ ga \ fa \ fa \ ha \ gc \ fc \ fc \ hc \ + \fa - fc\ < e" + by (simp add: abs_real_def split: if_split_asm) + show "norm (integral (cbox a b) (\x. if x \ s then f x else 0) - integral (cbox c d) + (\x. if x \ s then f x else 0)) < e" + unfolding real_norm_def + apply (rule *) + apply safe + unfolding real_norm_def[symmetric] + apply (rule B1(2)) + apply (rule order_trans) + apply (rule **) + apply (rule as(1)) + apply (rule B1(2)) + apply (rule order_trans) + apply (rule **) + apply (rule as(2)) + apply (rule B2(2)) + apply (rule order_trans) + apply (rule **) + apply (rule as(1)) + apply (rule B2(2)) + apply (rule order_trans) + apply (rule **) + apply (rule as(2)) + apply (rule obt) + apply (rule_tac[!] integral_le) + using obt + apply (auto intro!: h g interv) + done + qed + qed +qed + + +subsection \Adding integrals over several sets\ + +lemma has_integral_union: + fixes f :: "'n::euclidean_space \ 'a::banach" + assumes "(f has_integral i) s" + and "(f has_integral j) t" + and "negligible (s \ t)" + shows "(f has_integral (i + j)) (s \ t)" +proof - + note * = has_integral_restrict_univ[symmetric, of f] + show ?thesis + unfolding * + apply (rule has_integral_spike[OF assms(3)]) + defer + apply (rule has_integral_add[OF assms(1-2)[unfolded *]]) + apply auto + done +qed + +lemma integrable_union: + fixes f :: "'a::euclidean_space \ 'b :: banach" + assumes "negligible (A \ B)" "f integrable_on A" "f integrable_on B" + shows "f integrable_on (A \ B)" +proof - + from assms obtain y z where "(f has_integral y) A" "(f has_integral z) B" + by (auto simp: integrable_on_def) + from has_integral_union[OF this assms(1)] show ?thesis by (auto simp: integrable_on_def) +qed + +lemma integrable_union': + fixes f :: "'a::euclidean_space \ 'b :: banach" + assumes "f integrable_on A" "f integrable_on B" "negligible (A \ B)" "C = A \ B" + shows "f integrable_on C" + using integrable_union[of A B f] assms by simp + +lemma has_integral_unions: + fixes f :: "'n::euclidean_space \ 'a::banach" + assumes "finite t" + and "\s\t. (f has_integral (i s)) s" + and "\s\t. \s'\t. s \ s' \ negligible (s \ s')" + shows "(f has_integral (setsum i t)) (\t)" +proof - + note * = has_integral_restrict_univ[symmetric, of f] + have **: "negligible (\((\(a,b). a \ b) ` {(a,b). a \ t \ b \ {y. y \ t \ a \ y}}))" + apply (rule negligible_Union) + apply (rule finite_imageI) + apply (rule finite_subset[of _ "t \ t"]) + defer + apply (rule finite_cartesian_product[OF assms(1,1)]) + using assms(3) + apply auto + done + note assms(2)[unfolded *] + note has_integral_setsum[OF assms(1) this] + then show ?thesis + unfolding * + apply - + apply (rule has_integral_spike[OF **]) + defer + apply assumption + apply safe + proof goal_cases + case prems: (1 x) + then show ?case + proof (cases "x \ \t") + case True + then guess s unfolding Union_iff .. note s=this + then have *: "\b\t. x \ b \ b = s" + using prems(3) by blast + show ?thesis + unfolding if_P[OF True] + apply (rule trans) + defer + apply (rule setsum.cong) + apply (rule refl) + apply (subst *) + apply assumption + apply (rule refl) + unfolding setsum.delta[OF assms(1)] + using s + apply auto + done + qed auto + qed +qed + + +text \In particular adding integrals over a division, maybe not of an interval.\ + +lemma has_integral_combine_division: + fixes f :: "'n::euclidean_space \ 'a::banach" + assumes "d division_of s" + and "\k\d. (f has_integral (i k)) k" + shows "(f has_integral (setsum i d)) s" +proof - + note d = division_ofD[OF assms(1)] + show ?thesis + unfolding d(6)[symmetric] + apply (rule has_integral_unions) + apply (rule d assms)+ + apply rule + apply rule + apply rule + proof goal_cases + case prems: (1 s s') + from d(4)[OF this(1)] d(4)[OF this(2)] guess a c b d by (elim exE) note obt=this + from d(5)[OF prems] show ?case + unfolding obt interior_cbox + apply - + apply (rule negligible_subset[of "(cbox a b-box a b) \ (cbox c d-box c d)"]) + apply (rule negligible_Un negligible_frontier_interval)+ + apply auto + done + qed +qed + +lemma integral_combine_division_bottomup: + fixes f :: "'n::euclidean_space \ 'a::banach" + assumes "d division_of s" + and "\k\d. f integrable_on k" + shows "integral s f = setsum (\i. integral i f) d" + apply (rule integral_unique) + apply (rule has_integral_combine_division[OF assms(1)]) + using assms(2) + unfolding has_integral_integral + apply assumption + done + +lemma has_integral_combine_division_topdown: + fixes f :: "'n::euclidean_space \ 'a::banach" + assumes "f integrable_on s" + and "d division_of k" + and "k \ s" + shows "(f has_integral (setsum (\i. integral i f) d)) k" + apply (rule has_integral_combine_division[OF assms(2)]) + apply safe + unfolding has_integral_integral[symmetric] +proof goal_cases + case (1 k) + from division_ofD(2,4)[OF assms(2) this] + show ?case + apply safe + apply (rule integrable_on_subcbox) + apply (rule assms) + using assms(3) + apply auto + done +qed + +lemma integral_combine_division_topdown: + fixes f :: "'n::euclidean_space \ 'a::banach" + assumes "f integrable_on s" + and "d division_of s" + shows "integral s f = setsum (\i. integral i f) d" + apply (rule integral_unique) + apply (rule has_integral_combine_division_topdown) + using assms + apply auto + done + +lemma integrable_combine_division: + fixes f :: "'n::euclidean_space \ 'a::banach" + assumes "d division_of s" + and "\i\d. f integrable_on i" + shows "f integrable_on s" + using assms(2) + unfolding integrable_on_def + by (metis has_integral_combine_division[OF assms(1)]) + +lemma integrable_on_subdivision: + fixes f :: "'n::euclidean_space \ 'a::banach" + assumes "d division_of i" + and "f integrable_on s" + and "i \ s" + shows "f integrable_on i" + apply (rule integrable_combine_division assms)+ + apply safe +proof goal_cases + case 1 + note division_ofD(2,4)[OF assms(1) this] + then show ?case + apply safe + apply (rule integrable_on_subcbox[OF assms(2)]) + using assms(3) + apply auto + done +qed + + +subsection \Also tagged divisions\ + +lemma has_integral_combine_tagged_division: + fixes f :: "'n::euclidean_space \ 'a::banach" + assumes "p tagged_division_of s" + and "\(x,k) \ p. (f has_integral (i k)) k" + shows "(f has_integral (setsum (\(x,k). i k) p)) s" +proof - + have *: "(f has_integral (setsum (\k. integral k f) (snd ` p))) s" + apply (rule has_integral_combine_division) + apply (rule division_of_tagged_division[OF assms(1)]) + using assms(2) + unfolding has_integral_integral[symmetric] + apply safe + apply auto + done + then show ?thesis + apply - + apply (rule subst[where P="\i. (f has_integral i) s"]) + defer + apply assumption + apply (rule trans[of _ "setsum (\(x,k). integral k f) p"]) + apply (subst eq_commute) + apply (rule setsum.over_tagged_division_lemma[OF assms(1)]) + apply (rule integral_null) + apply assumption + apply (rule setsum.cong) + using assms(2) + apply auto + done +qed + +lemma integral_combine_tagged_division_bottomup: + fixes f :: "'n::euclidean_space \ 'a::banach" + assumes "p tagged_division_of (cbox a b)" + and "\(x,k)\p. f integrable_on k" + shows "integral (cbox a b) f = setsum (\(x,k). integral k f) p" + apply (rule integral_unique) + apply (rule has_integral_combine_tagged_division[OF assms(1)]) + using assms(2) + apply auto + done + +lemma has_integral_combine_tagged_division_topdown: + fixes f :: "'n::euclidean_space \ 'a::banach" + assumes "f integrable_on cbox a b" + and "p tagged_division_of (cbox a b)" + shows "(f has_integral (setsum (\(x,k). integral k f) p)) (cbox a b)" + apply (rule has_integral_combine_tagged_division[OF assms(2)]) + apply safe +proof goal_cases + case 1 + note tagged_division_ofD(3-4)[OF assms(2) this] + then show ?case + using integrable_subinterval[OF assms(1)] by blast +qed + +lemma integral_combine_tagged_division_topdown: + fixes f :: "'n::euclidean_space \ 'a::banach" + assumes "f integrable_on cbox a b" + and "p tagged_division_of (cbox a b)" + shows "integral (cbox a b) f = setsum (\(x,k). integral k f) p" + apply (rule integral_unique) + apply (rule has_integral_combine_tagged_division_topdown) + using assms + apply auto + done + + +subsection \Henstock's lemma\ + +lemma henstock_lemma_part1: + fixes f :: "'n::euclidean_space \ 'a::banach" + assumes "f integrable_on cbox a b" + and "e > 0" + and "gauge d" + and "(\p. p tagged_division_of (cbox a b) \ d fine p \ + norm (setsum (\(x,k). content k *\<^sub>R f x) p - integral(cbox a b) f) < e)" + and p: "p tagged_partial_division_of (cbox a b)" "d fine p" + shows "norm (setsum (\(x,k). content k *\<^sub>R f x - integral k f) p) \ e" + (is "?x \ e") +proof - + { presume "\k. 0 ?x \ e + k" then show ?thesis by (blast intro: field_le_epsilon) } + fix k :: real + assume k: "k > 0" + note p' = tagged_partial_division_ofD[OF p(1)] + have "\(snd ` p) \ cbox a b" + using p'(3) by fastforce + note partial_division_of_tagged_division[OF p(1)] this + from partial_division_extend_interval[OF this] guess q . note q=this and q' = division_ofD[OF this(2)] + define r where "r = q - snd ` p" + have "snd ` p \ r = {}" + unfolding r_def by auto + have r: "finite r" + using q' unfolding r_def by auto + + have "\i\r. \p. p tagged_division_of i \ d fine p \ + norm (setsum (\(x,j). content j *\<^sub>R f x) p - integral i f) < k / (real (card r) + 1)" + apply safe + proof goal_cases + case (1 i) + then have i: "i \ q" + unfolding r_def by auto + from q'(4)[OF this] guess u v by (elim exE) note uv=this + have *: "k / (real (card r) + 1) > 0" using k by simp + have "f integrable_on cbox u v" + apply (rule integrable_subinterval[OF assms(1)]) + using q'(2)[OF i] + unfolding uv + apply auto + done + note integrable_integral[OF this, unfolded has_integral[of f]] + from this[rule_format,OF *] guess dd .. note dd=conjunctD2[OF this,rule_format] + note gauge_inter[OF \gauge d\ dd(1)] + from fine_division_exists[OF this,of u v] guess qq . + then show ?case + apply (rule_tac x=qq in exI) + using dd(2)[of qq] + unfolding fine_inter uv + apply auto + done + qed + from bchoice[OF this] guess qq .. note qq=this[rule_format] + + let ?p = "p \ \(qq ` r)" + have "norm ((\(x, k)\?p. content k *\<^sub>R f x) - integral (cbox a b) f) < e" + apply (rule assms(4)[rule_format]) + proof + show "d fine ?p" + apply (rule fine_union) + apply (rule p) + apply (rule fine_unions) + using qq + apply auto + done + note * = tagged_partial_division_of_union_self[OF p(1)] + have "p \ \(qq ` r) tagged_division_of \(snd ` p) \ \r" + using r + proof (rule tagged_division_union[OF * tagged_division_unions], goal_cases) + case 1 + then show ?case + using qq by auto + next + case 2 + then show ?case + apply rule + apply rule + apply rule + apply(rule q'(5)) + unfolding r_def + apply auto + done + next + case 3 + then show ?case + apply (rule inter_interior_unions_intervals) + apply fact + apply rule + apply rule + apply (rule q') + defer + apply rule + apply (subst Int_commute) + apply (rule inter_interior_unions_intervals) + apply (rule finite_imageI) + apply (rule p') + apply rule + defer + apply rule + apply (rule q') + using q(1) p' + unfolding r_def + apply auto + done + qed + moreover have "\(snd ` p) \ \r = cbox a b" and "{qq i |i. i \ r} = qq ` r" + unfolding Union_Un_distrib[symmetric] r_def + using q + by auto + ultimately show "?p tagged_division_of (cbox a b)" + by fastforce + qed + + then have "norm ((\(x, k)\p. content k *\<^sub>R f x) + (\(x, k)\\(qq ` r). content k *\<^sub>R f x) - + integral (cbox a b) f) < e" + apply (subst setsum.union_inter_neutral[symmetric]) + apply (rule p') + prefer 3 + apply assumption + apply rule + apply (rule r) + apply safe + apply (drule qq) + proof - + fix x l k + assume as: "(x, l) \ p" "(x, l) \ qq k" "k \ r" + note qq[OF this(3)] + note tagged_division_ofD(3,4)[OF conjunct1[OF this] as(2)] + from this(2) guess u v by (elim exE) note uv=this + have "l\snd ` p" unfolding image_iff apply(rule_tac x="(x,l)" in bexI) using as by auto + then have "l \ q" "k \ q" "l \ k" + using as(1,3) q(1) unfolding r_def by auto + note q'(5)[OF this] + then have "interior l = {}" + using interior_mono[OF \l \ k\] by blast + then show "content l *\<^sub>R f x = 0" + unfolding uv content_eq_0_interior[symmetric] by auto + qed auto + + then have "norm ((\(x, k)\p. content k *\<^sub>R f x) + setsum (setsum (\(x, k). content k *\<^sub>R f x)) + (qq ` r) - integral (cbox a b) f) < e" + apply (subst (asm) setsum.Union_comp) + prefer 2 + unfolding split_paired_all split_conv image_iff + apply (erule bexE)+ + proof - + fix x m k l T1 T2 + assume "(x, m) \ T1" "(x, m) \ T2" "T1 \ T2" "k \ r" "l \ r" "T1 = qq k" "T2 = qq l" + note as = this(1-5)[unfolded this(6-)] + note kl = tagged_division_ofD(3,4)[OF qq[THEN conjunct1]] + from this(2)[OF as(4,1)] guess u v by (elim exE) note uv=this + have *: "interior (k \ l) = {}" + by (metis DiffE \T1 = qq k\ \T1 \ T2\ \T2 = qq l\ as(4) as(5) interior_Int q'(5) r_def) + have "interior m = {}" + unfolding subset_empty[symmetric] + unfolding *[symmetric] + apply (rule interior_mono) + using kl(1)[OF as(4,1)] kl(1)[OF as(5,2)] + apply auto + done + then show "content m *\<^sub>R f x = 0" + unfolding uv content_eq_0_interior[symmetric] + by auto + qed (insert qq, auto) + + then have **: "norm ((\(x, k)\p. content k *\<^sub>R f x) + setsum (setsum (\(x, k). content k *\<^sub>R f x) \ qq) r - + integral (cbox a b) f) < e" + apply (subst (asm) setsum.reindex_nontrivial) + apply fact + apply (rule setsum.neutral) + apply rule + unfolding split_paired_all split_conv + defer + apply assumption + proof - + fix k l x m + assume as: "k \ r" "l \ r" "k \ l" "qq k = qq l" "(x, m) \ qq k" + note tagged_division_ofD(6)[OF qq[THEN conjunct1]] + from this[OF as(1)] this[OF as(2)] show "content m *\<^sub>R f x = 0" + using as(3) unfolding as by auto + qed + + have *: "norm (cp - ip) \ e + k" + if "norm ((cp + cr) - i) < e" + and "norm (cr - ir) < k" + and "ip + ir = i" + for ir ip i cr cp + proof - + from that show ?thesis + using norm_triangle_le[of "cp + cr - i" "- (cr - ir)"] + unfolding that(3)[symmetric] norm_minus_cancel + by (auto simp add: algebra_simps) + qed + + have "?x = norm ((\(x, k)\p. content k *\<^sub>R f x) - (\(x, k)\p. integral k f))" + unfolding split_def setsum_subtractf .. + also have "\ \ e + k" + apply (rule *[OF **, where ir1="setsum (\k. integral k f) r"]) + proof goal_cases + case 1 + have *: "k * real (card r) / (1 + real (card r)) < k" + using k by (auto simp add: field_simps) + show ?case + apply (rule le_less_trans[of _ "setsum (\x. k / (real (card r) + 1)) r"]) + unfolding setsum_subtractf[symmetric] + apply (rule setsum_norm_le) + apply rule + apply (drule qq) + defer + unfolding divide_inverse setsum_left_distrib[symmetric] + unfolding divide_inverse[symmetric] + using * apply (auto simp add: field_simps) + done + next + case 2 + have *: "(\(x, k)\p. integral k f) = (\k\snd ` p. integral k f)" + apply (subst setsum.reindex_nontrivial) + apply fact + unfolding split_paired_all snd_conv split_def o_def + proof - + fix x l y m + assume as: "(x, l) \ p" "(y, m) \ p" "(x, l) \ (y, m)" "l = m" + from p'(4)[OF as(1)] guess u v by (elim exE) note uv=this + show "integral l f = 0" + unfolding uv + apply (rule integral_unique) + apply (rule has_integral_null) + unfolding content_eq_0_interior + using p'(5)[OF as(1-3)] + unfolding uv as(4)[symmetric] + apply auto + done + qed auto + from q(1) have **: "snd ` p \ q = q" by auto + show ?case + unfolding integral_combine_division_topdown[OF assms(1) q(2)] * r_def + using ** q'(1) p'(1) setsum.union_disjoint [of "snd ` p" "q - snd ` p" "\k. integral k f", symmetric] + by simp + qed + finally show "?x \ e + k" . +qed + +lemma henstock_lemma_part2: + fixes f :: "'m::euclidean_space \ 'n::euclidean_space" + assumes "f integrable_on cbox a b" + and "e > 0" + and "gauge d" + and "\p. p tagged_division_of (cbox a b) \ d fine p \ + norm (setsum (\(x,k). content k *\<^sub>R f x) p - integral (cbox a b) f) < e" + and "p tagged_partial_division_of (cbox a b)" + and "d fine p" + shows "setsum (\(x,k). norm (content k *\<^sub>R f x - integral k f)) p \ 2 * real (DIM('n)) * e" + unfolding split_def + apply (rule setsum_norm_allsubsets_bound) + defer + apply (rule henstock_lemma_part1[unfolded split_def,OF assms(1-3)]) + apply safe + apply (rule assms[rule_format,unfolded split_def]) + defer + apply (rule tagged_partial_division_subset) + apply (rule assms) + apply assumption + apply (rule fine_subset) + apply assumption + apply (rule assms) + using assms(5) + apply auto + done + +lemma henstock_lemma: + fixes f :: "'m::euclidean_space \ 'n::euclidean_space" + assumes "f integrable_on cbox a b" + and "e > 0" + obtains d where "gauge d" + and "\p. p tagged_partial_division_of (cbox a b) \ d fine p \ + setsum (\(x,k). norm(content k *\<^sub>R f x - integral k f)) p < e" +proof - + have *: "e / (2 * (real DIM('n) + 1)) > 0" using assms(2) by simp + from integrable_integral[OF assms(1),unfolded has_integral[of f],rule_format,OF this] + guess d .. note d = conjunctD2[OF this] + show thesis + apply (rule that) + apply (rule d) + proof (safe, goal_cases) + case (1 p) + note * = henstock_lemma_part2[OF assms(1) * d this] + show ?case + apply (rule le_less_trans[OF *]) + using \e > 0\ + apply (auto simp add: field_simps) + done + qed +qed + + +subsection \Geometric progression\ + +text \FIXME: Should one or more of these theorems be moved to @{file +"~~/src/HOL/Set_Interval.thy"}, alongside \geometric_sum\?\ + +lemma sum_gp_basic: + fixes x :: "'a::ring_1" + shows "(1 - x) * setsum (\i. x^i) {0 .. n} = (1 - x^(Suc n))" +proof - + define y where "y = 1 - x" + have "y * (\i=0..n. (1 - y) ^ i) = 1 - (1 - y) ^ Suc n" + by (induct n) (simp_all add: algebra_simps) + then show ?thesis + unfolding y_def by simp +qed + +lemma sum_gp_multiplied: + assumes mn: "m \ n" + shows "((1::'a::{field}) - x) * setsum (op ^ x) {m..n} = x^m - x^ Suc n" + (is "?lhs = ?rhs") +proof - + let ?S = "{0..(n - m)}" + from mn have mn': "n - m \ 0" + by arith + let ?f = "op + m" + have i: "inj_on ?f ?S" + unfolding inj_on_def by auto + have f: "?f ` ?S = {m..n}" + using mn + apply (auto simp add: image_iff Bex_def) + apply presburger + done + have th: "op ^ x \ op + m = (\i. x^m * x^i)" + by (rule ext) (simp add: power_add power_mult) + from setsum.reindex[OF i, of "op ^ x", unfolded f th setsum_right_distrib[symmetric]] + have "?lhs = x^m * ((1 - x) * setsum (op ^ x) {0..n - m})" + by simp + then show ?thesis + unfolding sum_gp_basic + using mn + by (simp add: field_simps power_add[symmetric]) +qed + +lemma sum_gp: + "setsum (op ^ (x::'a::{field})) {m .. n} = + (if n < m then 0 + else if x = 1 then of_nat ((n + 1) - m) + else (x^ m - x^ (Suc n)) / (1 - x))" +proof - + { + assume nm: "n < m" + then have ?thesis by simp + } + moreover + { + assume "\ n < m" + then have nm: "m \ n" + by arith + { + assume x: "x = 1" + then have ?thesis + by simp + } + moreover + { + assume x: "x \ 1" + then have nz: "1 - x \ 0" + by simp + from sum_gp_multiplied[OF nm, of x] nz have ?thesis + by (simp add: field_simps) + } + ultimately have ?thesis by blast + } + ultimately show ?thesis by blast +qed + +lemma sum_gp_offset: + "setsum (op ^ (x::'a::{field})) {m .. m+n} = + (if x = 1 then of_nat n + 1 else x^m * (1 - x^Suc n) / (1 - x))" + unfolding sum_gp[of x m "m + n"] power_Suc + by (simp add: field_simps power_add) + + +subsection \Monotone convergence (bounded interval first)\ + +lemma monotone_convergence_interval: + fixes f :: "nat \ 'n::euclidean_space \ real" + assumes "\k. (f k) integrable_on cbox a b" + and "\k. \x\cbox a b.(f k x) \ f (Suc k) x" + and "\x\cbox a b. ((\k. f k x) \ g x) sequentially" + and "bounded {integral (cbox a b) (f k) | k . k \ UNIV}" + shows "g integrable_on cbox a b \ ((\k. integral (cbox a b) (f k)) \ integral (cbox a b) g) sequentially" +proof (cases "content (cbox a b) = 0") + case True + show ?thesis + using integrable_on_null[OF True] + unfolding integral_null[OF True] + using tendsto_const + by auto +next + case False + have fg: "\x\cbox a b. \k. (f k x) \ 1 \ (g x) \ 1" + proof safe + fix x k + assume x: "x \ cbox a b" + note * = Lim_component_ge[OF assms(3)[rule_format, OF x] trivial_limit_sequentially] + show "f k x \ 1 \ g x \ 1" + apply (rule *) + unfolding eventually_sequentially + apply (rule_tac x=k in exI) + apply - + apply (rule transitive_stepwise_le) + using assms(2)[rule_format, OF x] + apply auto + done + qed + have "\i. ((\k. integral (cbox a b) (f k)) \ i) sequentially" + apply (rule bounded_increasing_convergent) + defer + apply rule + apply (rule integral_le) + apply safe + apply (rule assms(1-2)[rule_format])+ + using assms(4) + apply auto + done + then guess i .. note i=this + have i': "\k. (integral(cbox a b) (f k)) \ i\1" + apply (rule Lim_component_ge) + apply (rule i) + apply (rule trivial_limit_sequentially) + unfolding eventually_sequentially + apply (rule_tac x=k in exI) + apply (rule transitive_stepwise_le) + prefer 3 + unfolding inner_simps real_inner_1_right + apply (rule integral_le) + apply (rule assms(1-2)[rule_format])+ + using assms(2) + apply auto + done + + have "(g has_integral i) (cbox a b)" + unfolding has_integral + proof (safe, goal_cases) + case e: (1 e) + then have "\k. (\d. gauge d \ (\p. p tagged_division_of (cbox a b) \ d fine p \ + norm ((\(x, ka)\p. content ka *\<^sub>R f k x) - integral (cbox a b) (f k)) < e / 2 ^ (k + 2)))" + apply - + apply rule + apply (rule assms(1)[unfolded has_integral_integral has_integral,rule_format]) + apply auto + done + from choice[OF this] guess c .. note c=conjunctD2[OF this[rule_format],rule_format] + + have "\r. \k\r. 0 \ i\1 - (integral (cbox a b) (f k)) \ i\1 - (integral (cbox a b) (f k)) < e / 4" + proof - + have "e/4 > 0" + using e by auto + from LIMSEQ_D [OF i this] guess r .. + then show ?thesis + apply (rule_tac x=r in exI) + apply rule + apply (erule_tac x=k in allE) + subgoal for k using i'[of k] by auto + done + qed + then guess r .. note r=conjunctD2[OF this[rule_format]] + + have "\x\cbox a b. \n\r. \k\n. 0 \ (g x)\1 - (f k x)\1 \ + (g x)\1 - (f k x)\1 < e / (4 * content(cbox a b))" + proof (rule, goal_cases) + case prems: (1 x) + have "e / (4 * content (cbox a b)) > 0" + using \e>0\ False content_pos_le[of a b] by auto + from assms(3)[rule_format, OF prems, THEN LIMSEQ_D, OF this] + guess n .. note n=this + then show ?case + apply (rule_tac x="n + r" in exI) + apply safe + apply (erule_tac[2-3] x=k in allE) + unfolding dist_real_def + using fg[rule_format, OF prems] + apply (auto simp add: field_simps) + done + qed + from bchoice[OF this] guess m .. note m=conjunctD2[OF this[rule_format],rule_format] + define d where "d x = c (m x) x" for x + show ?case + apply (rule_tac x=d in exI) + proof safe + show "gauge d" + using c(1) unfolding gauge_def d_def by auto + next + fix p + assume p: "p tagged_division_of (cbox a b)" "d fine p" + note p'=tagged_division_ofD[OF p(1)] + have "\a. \x\p. m (fst x) \ a" + by (metis finite_imageI finite_nat_set_iff_bounded_le p'(1) rev_image_eqI) + then guess s .. note s=this + have *: "\a b c d. norm(a - b) \ e / 4 \ norm(b - c) < e / 2 \ + norm (c - d) < e / 4 \ norm (a - d) < e" + proof (safe, goal_cases) + case (1 a b c d) + then show ?case + using norm_triangle_lt[of "a - b" "b - c" "3* e/4"] + norm_triangle_lt[of "a - b + (b - c)" "c - d" e] + unfolding norm_minus_cancel + by (auto simp add: algebra_simps) + qed + show "norm ((\(x, k)\p. content k *\<^sub>R g x) - i) < e" + apply (rule *[rule_format,where + b="\(x, k)\p. content k *\<^sub>R f (m x) x" and c="\(x, k)\p. integral k (f (m x))"]) + proof (safe, goal_cases) + case 1 + show ?case + apply (rule order_trans[of _ "\(x, k)\p. content k * (e / (4 * content (cbox a b)))"]) + unfolding setsum_subtractf[symmetric] + apply (rule order_trans) + apply (rule norm_setsum) + apply (rule setsum_mono) + unfolding split_paired_all split_conv + unfolding split_def setsum_left_distrib[symmetric] scaleR_diff_right[symmetric] + unfolding additive_content_tagged_division[OF p(1), unfolded split_def] + proof - + fix x k + assume xk: "(x, k) \ p" + then have x: "x \ cbox a b" + using p'(2-3)[OF xk] by auto + from p'(4)[OF xk] guess u v by (elim exE) note uv=this + show "norm (content k *\<^sub>R (g x - f (m x) x)) \ content k * (e / (4 * content (cbox a b)))" + unfolding norm_scaleR uv + unfolding abs_of_nonneg[OF content_pos_le] + apply (rule mult_left_mono) + using m(2)[OF x,of "m x"] + apply auto + done + qed (insert False, auto) + + next + case 2 + show ?case + apply (rule le_less_trans[of _ "norm (\j = 0..s. + \(x, k)\{xk\p. m (fst xk) = j}. content k *\<^sub>R f (m x) x - integral k (f (m x)))"]) + apply (subst setsum_group) + apply fact + apply (rule finite_atLeastAtMost) + defer + apply (subst split_def)+ + unfolding setsum_subtractf + apply rule + proof - + show "norm (\j = 0..s. \(x, k)\{xk \ p. + m (fst xk) = j}. content k *\<^sub>R f (m x) x - integral k (f (m x))) < e / 2" + apply (rule le_less_trans[of _ "setsum (\i. e / 2^(i+2)) {0..s}"]) + apply (rule setsum_norm_le) + proof + show "(\i = 0..s. e / 2 ^ (i + 2)) < e / 2" + unfolding power_add divide_inverse inverse_mult_distrib + unfolding setsum_right_distrib[symmetric] setsum_left_distrib[symmetric] + unfolding power_inverse [symmetric] sum_gp + apply(rule mult_strict_left_mono[OF _ e]) + unfolding power2_eq_square + apply auto + done + fix t + assume "t \ {0..s}" + show "norm (\(x, k)\{xk \ p. m (fst xk) = t}. content k *\<^sub>R f (m x) x - + integral k (f (m x))) \ e / 2 ^ (t + 2)" + apply (rule order_trans + [of _ "norm (setsum (\(x,k). content k *\<^sub>R f t x - integral k (f t)) {xk \ p. m (fst xk) = t})"]) + apply (rule eq_refl) + apply (rule arg_cong[where f=norm]) + apply (rule setsum.cong) + apply (rule refl) + defer + apply (rule henstock_lemma_part1) + apply (rule assms(1)[rule_format]) + apply (simp add: e) + apply safe + apply (rule c)+ + apply rule + apply assumption+ + apply (rule tagged_partial_division_subset[of p]) + apply (rule p(1)[unfolded tagged_division_of_def,THEN conjunct1]) + defer + unfolding fine_def + apply safe + apply (drule p(2)[unfolded fine_def,rule_format]) + unfolding d_def + apply auto + done + qed + qed (insert s, auto) + next + case 3 + note comb = integral_combine_tagged_division_topdown[OF assms(1)[rule_format] p(1)] + have *: "\sr sx ss ks kr::real. kr = sr \ ks = ss \ + ks \ i \ sr \ sx \ sx \ ss \ 0 \ i\1 - kr\1 \ i\1 - kr\1 < e/4 \ \sx - i\ < e/4" + by auto + show ?case + unfolding real_norm_def + apply (rule *[rule_format]) + apply safe + apply (rule comb[of r]) + apply (rule comb[of s]) + apply (rule i'[unfolded real_inner_1_right]) + apply (rule_tac[1-2] setsum_mono) + unfolding split_paired_all split_conv + apply (rule_tac[1-2] integral_le[OF ]) + proof safe + show "0 \ i\1 - (integral (cbox a b) (f r))\1" + using r(1) by auto + show "i\1 - (integral (cbox a b) (f r))\1 < e / 4" + using r(2) by auto + fix x k + assume xk: "(x, k) \ p" + from p'(4)[OF this] guess u v by (elim exE) note uv=this + show "f r integrable_on k" + and "f s integrable_on k" + and "f (m x) integrable_on k" + and "f (m x) integrable_on k" + unfolding uv + apply (rule_tac[!] integrable_on_subcbox[OF assms(1)[rule_format]]) + using p'(3)[OF xk] + unfolding uv + apply auto + done + fix y + assume "y \ k" + then have "y \ cbox a b" + using p'(3)[OF xk] by auto + then have *: "\m. \n\m. f m y \ f n y" + apply - + apply (rule transitive_stepwise_le) + using assms(2) + apply auto + done + show "f r y \ f (m x) y" and "f (m x) y \ f s y" + apply (rule_tac[!] *[rule_format]) + using s[rule_format,OF xk] m(1)[of x] p'(2-3)[OF xk] + apply auto + done + qed + qed + qed + qed note * = this + + have "integral (cbox a b) g = i" + by (rule integral_unique) (rule *) + then show ?thesis + using i * by auto +qed + +lemma monotone_convergence_increasing: + fixes f :: "nat \ 'n::euclidean_space \ real" + assumes "\k. (f k) integrable_on s" + and "\k. \x\s. (f k x) \ (f (Suc k) x)" + and "\x\s. ((\k. f k x) \ g x) sequentially" + and "bounded {integral s (f k)| k. True}" + shows "g integrable_on s \ ((\k. integral s (f k)) \ integral s g) sequentially" +proof - + have lem: "g integrable_on s \ ((\k. integral s (f k)) \ integral s g) sequentially" + if "\k. \x\s. 0 \ f k x" + and "\k. (f k) integrable_on s" + and "\k. \x\s. f k x \ f (Suc k) x" + and "\x\s. ((\k. f k x) \ g x) sequentially" + and "bounded {integral s (f k)| k. True}" + for f :: "nat \ 'n::euclidean_space \ real" and g s + proof - + note assms=that[rule_format] + have "\x\s. \k. (f k x)\1 \ (g x)\1" + apply safe + apply (rule Lim_component_ge) + apply (rule that(4)[rule_format]) + apply assumption + apply (rule trivial_limit_sequentially) + unfolding eventually_sequentially + apply (rule_tac x=k in exI) + apply (rule transitive_stepwise_le) + using that(3) + apply auto + done + note fg=this[rule_format] + + have "\i. ((\k. integral s (f k)) \ i) sequentially" + apply (rule bounded_increasing_convergent) + apply (rule that(5)) + apply rule + apply (rule integral_le) + apply (rule that(2)[rule_format])+ + using that(3) + apply auto + done + then guess i .. note i=this + have "\k. \x\s. \n\k. f k x \ f n x" + apply rule + apply (rule transitive_stepwise_le) + using that(3) + apply auto + done + then have i': "\k. (integral s (f k))\1 \ i\1" + apply - + apply rule + apply (rule Lim_component_ge) + apply (rule i) + apply (rule trivial_limit_sequentially) + unfolding eventually_sequentially + apply (rule_tac x=k in exI) + apply safe + apply (rule integral_component_le) + apply simp + apply (rule that(2)[rule_format])+ + apply auto + done + + note int = assms(2)[unfolded integrable_alt[of _ s],THEN conjunct1,rule_format] + have ifif: "\k t. (\x. if x \ t then if x \ s then f k x else 0 else 0) = + (\x. if x \ t \ s then f k x else 0)" + by (rule ext) auto + have int': "\k a b. f k integrable_on cbox a b \ s" + apply (subst integrable_restrict_univ[symmetric]) + apply (subst ifif[symmetric]) + apply (subst integrable_restrict_univ) + apply (rule int) + done + have "\a b. (\x. if x \ s then g x else 0) integrable_on cbox a b \ + ((\k. integral (cbox a b) (\x. if x \ s then f k x else 0)) \ + integral (cbox a b) (\x. if x \ s then g x else 0)) sequentially" + proof (rule monotone_convergence_interval, safe, goal_cases) + case 1 + show ?case by (rule int) + next + case (2 _ _ _ x) + then show ?case + apply (cases "x \ s") + using assms(3) + apply auto + done + next + case (3 _ _ x) + then show ?case + apply (cases "x \ s") + using assms(4) + apply auto + done + next + case (4 a b) + note * = integral_nonneg + have "\k. norm (integral (cbox a b) (\x. if x \ s then f k x else 0)) \ norm (integral s (f k))" + unfolding real_norm_def + apply (subst abs_of_nonneg) + apply (rule *[OF int]) + apply safe + apply (case_tac "x \ s") + apply (drule assms(1)) + prefer 3 + apply (subst abs_of_nonneg) + apply (rule *[OF assms(2) that(1)[THEN spec]]) + apply (subst integral_restrict_univ[symmetric,OF int]) + unfolding ifif + unfolding integral_restrict_univ[OF int'] + apply (rule integral_subset_le[OF _ int' assms(2)]) + using assms(1) + apply auto + done + then show ?case + using assms(5) + unfolding bounded_iff + apply safe + apply (rule_tac x=aa in exI) + apply safe + apply (erule_tac x="integral s (f k)" in ballE) + apply (rule order_trans) + apply assumption + apply auto + done + qed + note g = conjunctD2[OF this] + + have "(g has_integral i) s" + unfolding has_integral_alt' + apply safe + apply (rule g(1)) + proof goal_cases + case (1 e) + then have "e/4>0" + by auto + from LIMSEQ_D [OF i this] guess N .. note N=this + note assms(2)[of N,unfolded has_integral_integral has_integral_alt'[of "f N"]] + from this[THEN conjunct2,rule_format,OF \e/4>0\] guess B .. note B=conjunctD2[OF this] + show ?case + apply rule + apply rule + apply (rule B) + apply safe + proof - + fix a b :: 'n + assume ab: "ball 0 B \ cbox a b" + from \e > 0\ have "e/2 > 0" + by auto + from LIMSEQ_D [OF g(2)[of a b] this] guess M .. note M=this + have **: "norm (integral (cbox a b) (\x. if x \ s then f N x else 0) - i) < e/2" + apply (rule norm_triangle_half_l) + using B(2)[rule_format,OF ab] N[rule_format,of N] + apply - + defer + apply (subst norm_minus_commute) + apply auto + done + have *: "\f1 f2 g. \f1 - i\ < e / 2 \ \f2 - g\ < e / 2 \ + f1 \ f2 \ f2 \ i \ \g - i\ < e" + unfolding real_inner_1_right by arith + show "norm (integral (cbox a b) (\x. if x \ s then g x else 0) - i) < e" + unfolding real_norm_def + apply (rule *[rule_format]) + apply (rule **[unfolded real_norm_def]) + apply (rule M[rule_format,of "M + N",unfolded real_norm_def]) + apply (rule le_add1) + apply (rule integral_le[OF int int]) + defer + apply (rule order_trans[OF _ i'[rule_format,of "M + N",unfolded real_inner_1_right]]) + proof (safe, goal_cases) + case (2 x) + have "\m. x \ s \ \n\m. (f m x)\1 \ (f n x)\1" + apply (rule transitive_stepwise_le) + using assms(3) + apply auto + done + then show ?case + by auto + next + case 1 + show ?case + apply (subst integral_restrict_univ[symmetric,OF int]) + unfolding ifif integral_restrict_univ[OF int'] + apply (rule integral_subset_le[OF _ int']) + using assms + apply auto + done + qed + qed + qed + then show ?thesis + apply safe + defer + apply (drule integral_unique) + using i + apply auto + done + qed + + have sub: "\k. integral s (\x. f k x - f 0 x) = integral s (f k) - integral s (f 0)" + apply (subst integral_diff) + apply (rule assms(1)[rule_format])+ + apply rule + done + have "\x m. x \ s \ \n\m. f m x \ f n x" + apply (rule transitive_stepwise_le) + using assms(2) + apply auto + done + note * = this[rule_format] + have "(\x. g x - f 0 x) integrable_on s \ ((\k. integral s (\x. f (Suc k) x - f 0 x)) \ + integral s (\x. g x - f 0 x)) sequentially" + apply (rule lem) + apply safe + proof goal_cases + case (1 k x) + then show ?case + using *[of x 0 "Suc k"] by auto + next + case (2 k) + then show ?case + apply (rule integrable_diff) + using assms(1) + apply auto + done + next + case (3 k x) + then show ?case + using *[of x "Suc k" "Suc (Suc k)"] by auto + next + case (4 x) + then show ?case + apply - + apply (rule tendsto_diff) + using LIMSEQ_ignore_initial_segment[OF assms(3)[rule_format],of x 1] + apply auto + done + next + case 5 + then show ?case + using assms(4) + unfolding bounded_iff + apply safe + apply (rule_tac x="a + norm (integral s (\x. f 0 x))" in exI) + apply safe + apply (erule_tac x="integral s (\x. f (Suc k) x)" in ballE) + unfolding sub + apply (rule order_trans[OF norm_triangle_ineq4]) + apply auto + done + qed + note conjunctD2[OF this] + note tendsto_add[OF this(2) tendsto_const[of "integral s (f 0)"]] + integrable_add[OF this(1) assms(1)[rule_format,of 0]] + then show ?thesis + unfolding sub + apply - + apply rule + defer + apply (subst(asm) integral_diff) + using assms(1) + apply auto + apply (rule LIMSEQ_imp_Suc) + apply assumption + done +qed + +lemma has_integral_monotone_convergence_increasing: + fixes f :: "nat \ 'a::euclidean_space \ real" + assumes f: "\k. (f k has_integral x k) s" + assumes "\k x. x \ s \ f k x \ f (Suc k) x" + assumes "\x. x \ s \ (\k. f k x) \ g x" + assumes "x \ x'" + shows "(g has_integral x') s" +proof - + have x_eq: "x = (\i. integral s (f i))" + by (simp add: integral_unique[OF f]) + then have x: "{integral s (f k) |k. True} = range x" + by auto + + have *: "g integrable_on s \ (\k. integral s (f k)) \ integral s g" + proof (intro monotone_convergence_increasing allI ballI assms) + show "bounded {integral s (f k) |k. True}" + unfolding x by (rule convergent_imp_bounded) fact + qed (auto intro: f) + then have "integral s g = x'" + by (intro LIMSEQ_unique[OF _ \x \ x'\]) (simp add: x_eq) + with * show ?thesis + by (simp add: has_integral_integral) +qed + +lemma monotone_convergence_decreasing: + fixes f :: "nat \ 'n::euclidean_space \ real" + assumes "\k. (f k) integrable_on s" + and "\k. \x\s. f (Suc k) x \ f k x" + and "\x\s. ((\k. f k x) \ g x) sequentially" + and "bounded {integral s (f k)| k. True}" + shows "g integrable_on s \ ((\k. integral s (f k)) \ integral s g) sequentially" +proof - + note assm = assms[rule_format] + have *: "{integral s (\x. - f k x) |k. True} = op *\<^sub>R (- 1) ` {integral s (f k)| k. True}" + apply safe + unfolding image_iff + apply (rule_tac x="integral s (f k)" in bexI) + prefer 3 + apply (rule_tac x=k in exI) + apply auto + done + have "(\x. - g x) integrable_on s \ + ((\k. integral s (\x. - f k x)) \ integral s (\x. - g x)) sequentially" + apply (rule monotone_convergence_increasing) + apply safe + apply (rule integrable_neg) + apply (rule assm) + defer + apply (rule tendsto_minus) + apply (rule assm) + apply assumption + unfolding * + apply (rule bounded_scaling) + using assm + apply auto + done + note * = conjunctD2[OF this] + show ?thesis + using integrable_neg[OF *(1)] tendsto_minus[OF *(2)] + by auto +qed + + +subsection \Absolute integrability (this is the same as Lebesgue integrability)\ + +definition absolutely_integrable_on (infixr "absolutely'_integrable'_on" 46) + where "f absolutely_integrable_on s \ f integrable_on s \ (\x. (norm(f x))) integrable_on s" + +lemma absolutely_integrable_onI[intro?]: + "f integrable_on s \ + (\x. (norm(f x))) integrable_on s \ f absolutely_integrable_on s" + unfolding absolutely_integrable_on_def + by auto + +lemma absolutely_integrable_onD[dest]: + assumes "f absolutely_integrable_on s" + shows "f integrable_on s" + and "(\x. norm (f x)) integrable_on s" + using assms + unfolding absolutely_integrable_on_def + by auto + +(*lemma absolutely_integrable_on_trans[simp]: + fixes f::"'n::euclidean_space \ real" + shows "(vec1 \ f) absolutely_integrable_on s \ f absolutely_integrable_on s" + unfolding absolutely_integrable_on_def o_def by auto*) + +lemma integral_norm_bound_integral: + fixes f :: "'n::euclidean_space \ 'a::banach" + assumes "f integrable_on s" + and "g integrable_on s" + and "\x\s. norm (f x) \ g x" + shows "norm (integral s f) \ integral s g" +proof - + have *: "\x y. (\e::real. 0 < e \ x < y + e) \ x \ y" + apply (rule ccontr) + apply (erule_tac x="x - y" in allE) + apply auto + done + have norm: "norm ig < dia + e" + if "norm sg \ dsa" + and "\dsa - dia\ < e / 2" + and "norm (sg - ig) < e / 2" + for e dsa dia and sg ig :: 'a + apply (rule le_less_trans[OF norm_triangle_sub[of ig sg]]) + apply (subst real_sum_of_halves[of e,symmetric]) + unfolding add.assoc[symmetric] + apply (rule add_le_less_mono) + defer + apply (subst norm_minus_commute) + apply (rule that(3)) + apply (rule order_trans[OF that(1)]) + using that(2) + apply arith + done + have lem: "norm (integral(cbox a b) f) \ integral (cbox a b) g" + if "f integrable_on cbox a b" + and "g integrable_on cbox a b" + and "\x\cbox a b. norm (f x) \ g x" + for f :: "'n \ 'a" and g a b + proof (rule *[rule_format]) + fix e :: real + assume "e > 0" + then have *: "e/2 > 0" + by auto + from integrable_integral[OF that(1),unfolded has_integral[of f],rule_format,OF *] + guess d1 .. note d1 = conjunctD2[OF this,rule_format] + from integrable_integral[OF that(2),unfolded has_integral[of g],rule_format,OF *] + guess d2 .. note d2 = conjunctD2[OF this,rule_format] + note gauge_inter[OF d1(1) d2(1)] + from fine_division_exists[OF this, of a b] guess p . note p=this + show "norm (integral (cbox a b) f) < integral (cbox a b) g + e" + apply (rule norm) + defer + apply (rule d2(2)[OF conjI[OF p(1)],unfolded real_norm_def]) + defer + apply (rule d1(2)[OF conjI[OF p(1)]]) + defer + apply (rule setsum_norm_le) + proof safe + fix x k + assume "(x, k) \ p" + note as = tagged_division_ofD(2-4)[OF p(1) this] + from this(3) guess u v by (elim exE) note uv=this + show "norm (content k *\<^sub>R f x) \ content k *\<^sub>R g x" + unfolding uv norm_scaleR + unfolding abs_of_nonneg[OF content_pos_le] real_scaleR_def + apply (rule mult_left_mono) + using that(3) as + apply auto + done + qed (insert p[unfolded fine_inter], auto) + qed + + { presume "\e. 0 < e \ norm (integral s f) < integral s g + e" + then show ?thesis by (rule *[rule_format]) auto } + fix e :: real + assume "e > 0" + then have e: "e/2 > 0" + by auto + note assms(1)[unfolded integrable_alt[of f]] note f=this[THEN conjunct1,rule_format] + note assms(2)[unfolded integrable_alt[of g]] note g=this[THEN conjunct1,rule_format] + from integrable_integral[OF assms(1),unfolded has_integral'[of f],rule_format,OF e] + guess B1 .. note B1=conjunctD2[OF this[rule_format],rule_format] + from integrable_integral[OF assms(2),unfolded has_integral'[of g],rule_format,OF e] + guess B2 .. note B2=conjunctD2[OF this[rule_format],rule_format] + from bounded_subset_cbox[OF bounded_ball, of "0::'n" "max B1 B2"] + guess a b by (elim exE) note ab=this[unfolded ball_max_Un] + + have "ball 0 B1 \ cbox a b" + using ab by auto + from B1(2)[OF this] guess z .. note z=conjunctD2[OF this] + have "ball 0 B2 \ cbox a b" + using ab by auto + from B2(2)[OF this] guess w .. note w=conjunctD2[OF this] + + show "norm (integral s f) < integral s g + e" + apply (rule norm) + apply (rule lem[OF f g, of a b]) + unfolding integral_unique[OF z(1)] integral_unique[OF w(1)] + defer + apply (rule w(2)[unfolded real_norm_def]) + apply (rule z(2)) + apply safe + apply (case_tac "x \ s") + unfolding if_P + apply (rule assms(3)[rule_format]) + apply auto + done +qed + +lemma integral_norm_bound_integral_component: + fixes f :: "'n::euclidean_space \ 'a::banach" + fixes g :: "'n \ 'b::euclidean_space" + assumes "f integrable_on s" + and "g integrable_on s" + and "\x\s. norm(f x) \ (g x)\k" + shows "norm (integral s f) \ (integral s g)\k" +proof - + have "norm (integral s f) \ integral s ((\x. x \ k) \ g)" + apply (rule integral_norm_bound_integral[OF assms(1)]) + apply (rule integrable_linear[OF assms(2)]) + apply rule + unfolding o_def + apply (rule assms) + done + then show ?thesis + unfolding o_def integral_component_eq[OF assms(2)] . +qed + +lemma has_integral_norm_bound_integral_component: + fixes f :: "'n::euclidean_space \ 'a::banach" + fixes g :: "'n \ 'b::euclidean_space" + assumes "(f has_integral i) s" + and "(g has_integral j) s" + and "\x\s. norm (f x) \ (g x)\k" + shows "norm i \ j\k" + using integral_norm_bound_integral_component[of f s g k] + unfolding integral_unique[OF assms(1)] integral_unique[OF assms(2)] + using assms + by auto + +lemma absolutely_integrable_le: + fixes f :: "'n::euclidean_space \ 'a::banach" + assumes "f absolutely_integrable_on s" + shows "norm (integral s f) \ integral s (\x. norm (f x))" + apply (rule integral_norm_bound_integral) + using assms + apply auto + done + +lemma absolutely_integrable_0[intro]: + "(\x. 0) absolutely_integrable_on s" + unfolding absolutely_integrable_on_def + by auto + +lemma absolutely_integrable_cmul[intro]: + "f absolutely_integrable_on s \ + (\x. c *\<^sub>R f x) absolutely_integrable_on s" + unfolding absolutely_integrable_on_def + using integrable_cmul[of f s c] + using integrable_cmul[of "\x. norm (f x)" s "\c\"] + by auto + +lemma absolutely_integrable_neg[intro]: + "f absolutely_integrable_on s \ + (\x. -f(x)) absolutely_integrable_on s" + apply (drule absolutely_integrable_cmul[where c="-1"]) + apply auto + done + +lemma absolutely_integrable_norm[intro]: + "f absolutely_integrable_on s \ + (\x. norm (f x)) absolutely_integrable_on s" + unfolding absolutely_integrable_on_def + by auto + +lemma absolutely_integrable_abs[intro]: + "f absolutely_integrable_on s \ + (\x. \f x::real\) absolutely_integrable_on s" + apply (drule absolutely_integrable_norm) + unfolding real_norm_def + apply assumption + done + +lemma absolutely_integrable_on_subinterval: + fixes f :: "'n::euclidean_space \ 'a::banach" + shows "f absolutely_integrable_on s \ + cbox a b \ s \ f absolutely_integrable_on cbox a b" + unfolding absolutely_integrable_on_def + by (metis integrable_on_subcbox) + +lemma absolutely_integrable_bounded_variation: + fixes f :: "'n::euclidean_space \ 'a::banach" + assumes "f absolutely_integrable_on UNIV" + obtains B where "\d. d division_of (\d) \ setsum (\k. norm(integral k f)) d \ B" + apply (rule that[of "integral UNIV (\x. norm (f x))"]) + apply safe +proof goal_cases + case prems: (1 d) + note d = division_ofD[OF prems(2)] + have "(\k\d. norm (integral k f)) \ (\i\d. integral i (\x. norm (f x)))" + apply (rule setsum_mono,rule absolutely_integrable_le) + apply (drule d(4)) + apply safe + apply (rule absolutely_integrable_on_subinterval[OF assms]) + apply auto + done + also have "\ \ integral (\d) (\x. norm (f x))" + apply (subst integral_combine_division_topdown[OF _ prems(2)]) + using integrable_on_subdivision[OF prems(2)] + using assms + apply auto + done + also have "\ \ integral UNIV (\x. norm (f x))" + apply (rule integral_subset_le) + using integrable_on_subdivision[OF prems(2)] + using assms + apply auto + done + finally show ?case . +qed + +lemma helplemma: + assumes "setsum (\x. norm (f x - g x)) s < e" + and "finite s" + shows "\setsum (\x. norm(f x)) s - setsum (\x. norm(g x)) s\ < e" + unfolding setsum_subtractf[symmetric] + apply (rule le_less_trans[OF setsum_abs]) + apply (rule le_less_trans[OF _ assms(1)]) + apply (rule setsum_mono) + apply (rule norm_triangle_ineq3) + done + +lemma bounded_variation_absolutely_integrable_interval: + fixes f :: "'n::euclidean_space \ 'm::euclidean_space" + assumes f: "f integrable_on cbox a b" + and *: "\d. d division_of (cbox a b) \ setsum (\k. norm(integral k f)) d \ B" + shows "f absolutely_integrable_on cbox a b" +proof - + let ?f = "\d. \k\d. norm (integral k f)" and ?D = "{d. d division_of (cbox a b)}" + have D_1: "?D \ {}" + by (rule elementary_interval[of a b]) auto + have D_2: "bdd_above (?f`?D)" + by (metis * mem_Collect_eq bdd_aboveI2) + note D = D_1 D_2 + let ?S = "SUP x:?D. ?f x" + show ?thesis + apply (rule absolutely_integrable_onI [OF f has_integral_integrable]) + apply (subst has_integral[of _ ?S]) + apply safe + proof goal_cases + case e: (1 e) + then have "?S - e / 2 < ?S" by simp + then obtain d where d: "d division_of (cbox a b)" "?S - e / 2 < (\k\d. norm (integral k f))" + unfolding less_cSUP_iff[OF D] by auto + note d' = division_ofD[OF this(1)] + + have "\x. \e>0. \i\d. x \ i \ ball x e \ i = {}" + proof + fix x + have "\da>0. \xa\\{i \ d. x \ i}. da \ dist x xa" + apply (rule separate_point_closed) + apply (rule closed_Union) + apply (rule finite_subset[OF _ d'(1)]) + using d'(4) + apply auto + done + then show "\e>0. \i\d. x \ i \ ball x e \ i = {}" + by force + qed + from choice[OF this] guess k .. note k=conjunctD2[OF this[rule_format],rule_format] + + have "e/2 > 0" + using e by auto + from henstock_lemma[OF assms(1) this] guess g . note g=this[rule_format] + let ?g = "\x. g x \ ball x (k x)" + show ?case + apply (rule_tac x="?g" in exI) + apply safe + proof - + show "gauge ?g" + using g(1) k(1) + unfolding gauge_def + by auto + fix p + assume "p tagged_division_of (cbox a b)" and "?g fine p" + note p = this(1) conjunctD2[OF this(2)[unfolded fine_inter]] + note p' = tagged_division_ofD[OF p(1)] + define p' where "p' = {(x,k) | x k. \i l. x \ i \ i \ d \ (x,l) \ p \ k = i \ l}" + have gp': "g fine p'" + using p(2) + unfolding p'_def fine_def + by auto + have p'': "p' tagged_division_of (cbox a b)" + apply (rule tagged_division_ofI) + proof - + show "finite p'" + apply (rule finite_subset[of _ "(\(k,(x,l)). (x,k \ l)) ` + {(k,xl) | k xl. k \ d \ xl \ p}"]) + unfolding p'_def + defer + apply (rule finite_imageI,rule finite_product_dependent[OF d'(1) p'(1)]) + apply safe + unfolding image_iff + apply (rule_tac x="(i,x,l)" in bexI) + apply auto + done + fix x k + assume "(x, k) \ p'" + then have "\i l. x \ i \ i \ d \ (x, l) \ p \ k = i \ l" + unfolding p'_def by auto + then guess i l by (elim exE) note il=conjunctD4[OF this] + show "x \ k" and "k \ cbox a b" + using p'(2-3)[OF il(3)] il by auto + show "\a b. k = cbox a b" + unfolding il using p'(4)[OF il(3)] d'(4)[OF il(2)] + apply safe + unfolding inter_interval + apply auto + done + next + fix x1 k1 + assume "(x1, k1) \ p'" + then have "\i l. x1 \ i \ i \ d \ (x1, l) \ p \ k1 = i \ l" + unfolding p'_def by auto + then guess i1 l1 by (elim exE) note il1=conjunctD4[OF this] + fix x2 k2 + assume "(x2,k2)\p'" + then have "\i l. x2 \ i \ i \ d \ (x2, l) \ p \ k2 = i \ l" + unfolding p'_def by auto + then guess i2 l2 by (elim exE) note il2=conjunctD4[OF this] + assume "(x1, k1) \ (x2, k2)" + then have "interior i1 \ interior i2 = {} \ interior l1 \ interior l2 = {}" + using d'(5)[OF il1(2) il2(2)] p'(5)[OF il1(3) il2(3)] + unfolding il1 il2 + by auto + then show "interior k1 \ interior k2 = {}" + unfolding il1 il2 by auto + next + have *: "\(x, X) \ p'. X \ cbox a b" + unfolding p'_def using d' by auto + show "\{k. \x. (x, k) \ p'} = cbox a b" + apply rule + apply (rule Union_least) + unfolding mem_Collect_eq + apply (erule exE) + apply (drule *[rule_format]) + apply safe + proof - + fix y + assume y: "y \ cbox a b" + then have "\x l. (x, l) \ p \ y\l" + unfolding p'(6)[symmetric] by auto + then guess x l by (elim exE) note xl=conjunctD2[OF this] + then have "\k. k \ d \ y \ k" + using y unfolding d'(6)[symmetric] by auto + then guess i .. note i = conjunctD2[OF this] + have "x \ i" + using fineD[OF p(3) xl(1)] + using k(2)[OF i(1), of x] + using i(2) xl(2) + by auto + then show "y \ \{k. \x. (x, k) \ p'}" + unfolding p'_def Union_iff + apply (rule_tac x="i \ l" in bexI) + using i xl + apply auto + done + qed + qed + + then have "(\(x, k)\p'. norm (content k *\<^sub>R f x - integral k f)) < e / 2" + apply - + apply (rule g(2)[rule_format]) + unfolding tagged_division_of_def + apply safe + apply (rule gp') + done + then have **: "\(\(x,k)\p'. norm (content k *\<^sub>R f x)) - (\(x,k)\p'. norm (integral k f))\ < e / 2" + unfolding split_def + using p'' + by (force intro!: helplemma) + + have p'alt: "p' = {(x,(i \ l)) | x i l. (x,l) \ p \ i \ d \ i \ l \ {}}" + proof (safe, goal_cases) + case prems: (2 _ _ x i l) + have "x \ i" + using fineD[OF p(3) prems(1)] k(2)[OF prems(2), of x] prems(4-) + by auto + then have "(x, i \ l) \ p'" + unfolding p'_def + using prems + apply safe + apply (rule_tac x=x in exI) + apply (rule_tac x="i \ l" in exI) + apply safe + using prems + apply auto + done + then show ?case + using prems(3) by auto + next + fix x k + assume "(x, k) \ p'" + then have "\i l. x \ i \ i \ d \ (x, l) \ p \ k = i \ l" + unfolding p'_def by auto + then guess i l by (elim exE) note il=conjunctD4[OF this] + then show "\y i l. (x, k) = (y, i \ l) \ (y, l) \ p \ i \ d \ i \ l \ {}" + apply (rule_tac x=x in exI) + apply (rule_tac x=i in exI) + apply (rule_tac x=l in exI) + using p'(2)[OF il(3)] + apply auto + done + qed + have sum_p': "(\(x, k)\p'. norm (integral k f)) = (\k\snd ` p'. norm (integral k f))" + apply (subst setsum.over_tagged_division_lemma[OF p'',of "\k. norm (integral k f)"]) + unfolding norm_eq_zero + apply (rule integral_null) + apply assumption + apply rule + done + note snd_p = division_ofD[OF division_of_tagged_division[OF p(1)]] + + have *: "\sni sni' sf sf'. \sf' - sni'\ < e / 2 \ ?S - e / 2 < sni \ sni' \ ?S \ + sni \ sni' \ sf' = sf \ \sf - ?S\ < e" + by arith + show "norm ((\(x, k)\p. content k *\<^sub>R norm (f x)) - ?S) < e" + unfolding real_norm_def + apply (rule *[rule_format,OF **]) + apply safe + apply(rule d(2)) + proof goal_cases + case 1 + show ?case + by (auto simp: sum_p' division_of_tagged_division[OF p''] D intro!: cSUP_upper) + next + case 2 + have *: "{k \ l | k l. k \ d \ l \ snd ` p} = + (\(k,l). k \ l) ` {(k,l)|k l. k \ d \ l \ snd ` p}" + by auto + have "(\k\d. norm (integral k f)) \ (\i\d. \l\snd ` p. norm (integral (i \ l) f))" + proof (rule setsum_mono, goal_cases) + case k: (1 k) + from d'(4)[OF this] guess u v by (elim exE) note uv=this + define d' where "d' = {cbox u v \ l |l. l \ snd ` p \ cbox u v \ l \ {}}" + note uvab = d'(2)[OF k[unfolded uv]] + have "d' division_of cbox u v" + apply (subst d'_def) + apply (rule division_inter_1) + apply (rule division_of_tagged_division[OF p(1)]) + apply (rule uvab) + done + then have "norm (integral k f) \ setsum (\k. norm (integral k f)) d'" + unfolding uv + apply (subst integral_combine_division_topdown[of _ _ d']) + apply (rule integrable_on_subcbox[OF assms(1) uvab]) + apply assumption + apply (rule setsum_norm_le) + apply auto + done + also have "\ = (\k\{k \ l |l. l \ snd ` p}. norm (integral k f))" + apply (rule setsum.mono_neutral_left) + apply (subst simple_image) + apply (rule finite_imageI)+ + apply fact + unfolding d'_def uv + apply blast + proof (rule, goal_cases) + case prems: (1 i) + then have "i \ {cbox u v \ l |l. l \ snd ` p}" + by auto + from this[unfolded mem_Collect_eq] guess l .. note l=this + then have "cbox u v \ l = {}" + using prems by auto + then show ?case + using l by auto + qed + also have "\ = (\l\snd ` p. norm (integral (k \ l) f))" + unfolding simple_image + apply (rule setsum.reindex_nontrivial [unfolded o_def]) + apply (rule finite_imageI) + apply (rule p') + proof goal_cases + case prems: (1 l y) + have "interior (k \ l) \ interior (l \ y)" + apply (subst(2) interior_Int) + apply (rule Int_greatest) + defer + apply (subst prems(4)) + apply auto + done + then have *: "interior (k \ l) = {}" + using snd_p(5)[OF prems(1-3)] by auto + from d'(4)[OF k] snd_p(4)[OF prems(1)] guess u1 v1 u2 v2 by (elim exE) note uv=this + show ?case + using * + unfolding uv inter_interval content_eq_0_interior[symmetric] + by auto + qed + finally show ?case . + qed + also have "\ = (\(i,l)\{(i, l) |i l. i \ d \ l \ snd ` p}. norm (integral (i\l) f))" + apply (subst sum_sum_product[symmetric]) + apply fact + using p'(1) + apply auto + done + also have "\ = (\x\{(i, l) |i l. i \ d \ l \ snd ` p}. norm (integral (case_prod op \ x) f))" + unfolding split_def .. + also have "\ = (\k\{i \ l |i l. i \ d \ l \ snd ` p}. norm (integral k f))" + unfolding * + apply (rule setsum.reindex_nontrivial [symmetric, unfolded o_def]) + apply (rule finite_product_dependent) + apply fact + apply (rule finite_imageI) + apply (rule p') + unfolding split_paired_all mem_Collect_eq split_conv o_def + proof - + note * = division_ofD(4,5)[OF division_of_tagged_division,OF p(1)] + fix l1 l2 k1 k2 + assume as: + "(l1, k1) \ (l2, k2)" + "l1 \ k1 = l2 \ k2" + "\i l. (l1, k1) = (i, l) \ i \ d \ l \ snd ` p" + "\i l. (l2, k2) = (i, l) \ i \ d \ l \ snd ` p" + then have "l1 \ d" and "k1 \ snd ` p" + by auto from d'(4)[OF this(1)] *(1)[OF this(2)] + guess u1 v1 u2 v2 by (elim exE) note uv=this + have "l1 \ l2 \ k1 \ k2" + using as by auto + then have "interior k1 \ interior k2 = {} \ interior l1 \ interior l2 = {}" + apply - + apply (erule disjE) + apply (rule disjI2) + apply (rule d'(5)) + prefer 4 + apply (rule disjI1) + apply (rule *) + using as + apply auto + done + moreover have "interior (l1 \ k1) = interior (l2 \ k2)" + using as(2) by auto + ultimately have "interior(l1 \ k1) = {}" + by auto + then show "norm (integral (l1 \ k1) f) = 0" + unfolding uv inter_interval + unfolding content_eq_0_interior[symmetric] + by auto + qed + also have "\ = (\(x, k)\p'. norm (integral k f))" + unfolding sum_p' + apply (rule setsum.mono_neutral_right) + apply (subst *) + apply (rule finite_imageI[OF finite_product_dependent]) + apply fact + apply (rule finite_imageI[OF p'(1)]) + apply safe + proof goal_cases + case (2 i ia l a b) + then have "ia \ b = {}" + unfolding p'alt image_iff Bex_def not_ex + apply (erule_tac x="(a, ia \ b)" in allE) + apply auto + done + then show ?case + by auto + next + case (1 x a b) + then show ?case + unfolding p'_def + apply safe + apply (rule_tac x=i in exI) + apply (rule_tac x=l in exI) + unfolding snd_conv image_iff + apply safe + apply (rule_tac x="(a,l)" in bexI) + apply auto + done + qed + finally show ?case . + next + case 3 + let ?S = "{(x, i \ l) |x i l. (x, l) \ p \ i \ d}" + have Sigma_alt: "\s t. s \ t = {(i, j) |i j. i \ s \ j \ t}" + by auto + have *: "?S = (\(xl,i). (fst xl, snd xl \ i)) ` (p \ d)" + apply safe + unfolding image_iff + apply (rule_tac x="((x,l),i)" in bexI) + apply auto + done + note pdfin = finite_cartesian_product[OF p'(1) d'(1)] + have "(\(x, k)\p'. norm (content k *\<^sub>R f x)) = (\(x, k)\?S. \content k\ * norm (f x))" + unfolding norm_scaleR + apply (rule setsum.mono_neutral_left) + apply (subst *) + apply (rule finite_imageI) + apply fact + unfolding p'alt + apply blast + apply safe + apply (rule_tac x=x in exI) + apply (rule_tac x=i in exI) + apply (rule_tac x=l in exI) + apply auto + done + also have "\ = (\((x,l),i)\p \ d. \content (l \ i)\ * norm (f x))" + unfolding * + apply (subst setsum.reindex_nontrivial) + apply fact + unfolding split_paired_all + unfolding o_def split_def snd_conv fst_conv mem_Sigma_iff prod.inject + apply (elim conjE) + proof - + fix x1 l1 k1 x2 l2 k2 + assume as: "(x1, l1) \ p" "(x2, l2) \ p" "k1 \ d" "k2 \ d" + "x1 = x2" "l1 \ k1 = l2 \ k2" "\ ((x1 = x2 \ l1 = l2) \ k1 = k2)" + from d'(4)[OF as(3)] p'(4)[OF as(1)] guess u1 v1 u2 v2 by (elim exE) note uv=this + from as have "l1 \ l2 \ k1 \ k2" + by auto + then have "interior k1 \ interior k2 = {} \ interior l1 \ interior l2 = {}" + apply - + apply (erule disjE) + apply (rule disjI2) + defer + apply (rule disjI1) + apply (rule d'(5)[OF as(3-4)]) + apply assumption + apply (rule p'(5)[OF as(1-2)]) + apply auto + done + moreover have "interior (l1 \ k1) = interior (l2 \ k2)" + unfolding as .. + ultimately have "interior (l1 \ k1) = {}" + by auto + then show "\content (l1 \ k1)\ * norm (f x1) = 0" + unfolding uv inter_interval + unfolding content_eq_0_interior[symmetric] + by auto + qed safe + also have "\ = (\(x, k)\p. content k *\<^sub>R norm (f x))" + unfolding Sigma_alt + apply (subst sum_sum_product[symmetric]) + apply (rule p') + apply rule + apply (rule d') + apply (rule setsum.cong) + apply (rule refl) + unfolding split_paired_all split_conv + proof - + fix x l + assume as: "(x, l) \ p" + note xl = p'(2-4)[OF this] + from this(3) guess u v by (elim exE) note uv=this + have "(\i\d. \content (l \ i)\) = (\k\d. content (k \ cbox u v))" + apply (rule setsum.cong) + apply (rule refl) + apply (drule d'(4)) + apply safe + apply (subst Int_commute) + unfolding inter_interval uv + apply (subst abs_of_nonneg) + apply auto + done + also have "\ = setsum content {k \ cbox u v| k. k \ d}" + unfolding simple_image + apply (rule setsum.reindex_nontrivial [unfolded o_def, symmetric]) + apply (rule d') + proof goal_cases + case prems: (1 k y) + from d'(4)[OF this(1)] d'(4)[OF this(2)] + guess u1 v1 u2 v2 by (elim exE) note uv=this + have "{} = interior ((k \ y) \ cbox u v)" + apply (subst interior_Int) + using d'(5)[OF prems(1-3)] + apply auto + done + also have "\ = interior (y \ (k \ cbox u v))" + by auto + also have "\ = interior (k \ cbox u v)" + unfolding prems(4) by auto + finally show ?case + unfolding uv inter_interval content_eq_0_interior .. + qed + also have "\ = setsum content {cbox u v \ k |k. k \ d \ cbox u v \ k \ {}}" + apply (rule setsum.mono_neutral_right) + unfolding simple_image + apply (rule finite_imageI) + apply (rule d') + apply blast + apply safe + apply (rule_tac x=k in exI) + proof goal_cases + case prems: (1 i k) + from d'(4)[OF this(1)] guess a b by (elim exE) note ab=this + have "interior (k \ cbox u v) \ {}" + using prems(2) + unfolding ab inter_interval content_eq_0_interior + by auto + then show ?case + using prems(1) + using interior_subset[of "k \ cbox u v"] + by auto + qed + finally show "(\i\d. \content (l \ i)\ * norm (f x)) = content l *\<^sub>R norm (f x)" + unfolding setsum_left_distrib[symmetric] real_scaleR_def + apply (subst(asm) additive_content_division[OF division_inter_1[OF d(1)]]) + using xl(2)[unfolded uv] + unfolding uv + apply auto + done + qed + finally show ?case . + qed + qed + qed +qed + +lemma bounded_variation_absolutely_integrable: + fixes f :: "'n::euclidean_space \ 'm::euclidean_space" + assumes "f integrable_on UNIV" + and "\d. d division_of (\d) \ setsum (\k. norm (integral k f)) d \ B" + shows "f absolutely_integrable_on UNIV" +proof (rule absolutely_integrable_onI, fact, rule) + let ?f = "\d. \k\d. norm (integral k f)" and ?D = "{d. d division_of (\d)}" + have D_1: "?D \ {}" + by (rule elementary_interval) auto + have D_2: "bdd_above (?f`?D)" + by (intro bdd_aboveI2[where M=B] assms(2)[rule_format]) simp + note D = D_1 D_2 + let ?S = "SUP d:?D. ?f d" + have f_int: "\a b. f absolutely_integrable_on cbox a b" + apply (rule bounded_variation_absolutely_integrable_interval[where B=B]) + apply (rule integrable_on_subcbox[OF assms(1)]) + defer + apply safe + apply (rule assms(2)[rule_format]) + apply auto + done + show "((\x. norm (f x)) has_integral ?S) UNIV" + apply (subst has_integral_alt') + apply safe + proof goal_cases + case (1 a b) + show ?case + using f_int[of a b] by auto + next + case prems: (2 e) + have "\y\setsum (\k. norm (integral k f)) ` {d. d division_of \d}. \ y \ ?S - e" + proof (rule ccontr) + assume "\ ?thesis" + then have "?S \ ?S - e" + by (intro cSUP_least[OF D(1)]) auto + then show False + using prems by auto + qed + then obtain K where *: "\x\{d. d division_of \d}. K = (\k\x. norm (integral k f))" + "SUPREMUM {d. d division_of \d} (setsum (\k. norm (integral k f))) - e < K" + by (auto simp add: image_iff not_le) + from this(1) obtain d where "d division_of \d" + and "K = (\k\d. norm (integral k f))" + by auto + note d = this(1) *(2)[unfolded this(2)] + note d'=division_ofD[OF this(1)] + have "bounded (\d)" + by (rule elementary_bounded,fact) + from this[unfolded bounded_pos] obtain K where + K: "0 < K" "\x\\d. norm x \ K" by auto + show ?case + apply (rule_tac x="K + 1" in exI) + apply safe + proof - + fix a b :: 'n + assume ab: "ball 0 (K + 1) \ cbox a b" + have *: "\s s1. ?S - e < s1 \ s1 \ s \ s < ?S + e \ \s - ?S\ < e" + by arith + show "norm (integral (cbox a b) (\x. if x \ UNIV then norm (f x) else 0) - ?S) < e" + unfolding real_norm_def + apply (rule *[rule_format]) + apply safe + apply (rule d(2)) + proof goal_cases + case 1 + have "(\k\d. norm (integral k f)) \ setsum (\k. integral k (\x. norm (f x))) d" + apply (rule setsum_mono) + apply (rule absolutely_integrable_le) + apply (drule d'(4)) + apply safe + apply (rule f_int) + done + also have "\ = integral (\d) (\x. norm (f x))" + apply (rule integral_combine_division_bottomup[symmetric]) + apply (rule d) + unfolding forall_in_division[OF d(1)] + using f_int + apply auto + done + also have "\ \ integral (cbox a b) (\x. if x \ UNIV then norm (f x) else 0)" + proof - + have "\d \ cbox a b" + apply rule + apply (drule K(2)[rule_format]) + apply (rule ab[unfolded subset_eq,rule_format]) + apply (auto simp add: dist_norm) + done + then show ?thesis + apply - + apply (subst if_P) + apply rule + apply (rule integral_subset_le) + defer + apply (rule integrable_on_subdivision[of _ _ _ "cbox a b"]) + apply (rule d) + using f_int[of a b] + apply auto + done + qed + finally show ?case . + next + note f = absolutely_integrable_onD[OF f_int[of a b]] + note * = this(2)[unfolded has_integral_integral has_integral[of "\x. norm (f x)"],rule_format] + have "e/2>0" + using \e > 0\ by auto + from * [OF this] obtain d1 where + d1: "gauge d1" "\p. p tagged_division_of (cbox a b) \ d1 fine p \ + norm ((\(x, k)\p. content k *\<^sub>R norm (f x)) - integral (cbox a b) (\x. norm (f x))) < e / 2" + by auto + from henstock_lemma [OF f(1) \e/2>0\] obtain d2 where + d2: "gauge d2" "\p. p tagged_partial_division_of (cbox a b) \ d2 fine p \ + (\(x, k)\p. norm (content k *\<^sub>R f x - integral k f)) < e / 2" . + obtain p where + p: "p tagged_division_of (cbox a b)" "d1 fine p" "d2 fine p" + by (rule fine_division_exists [OF gauge_inter [OF d1(1) d2(1)], of a b]) + (auto simp add: fine_inter) + have *: "\sf sf' si di. sf' = sf \ si \ ?S \ \sf - si\ < e / 2 \ + \sf' - di\ < e / 2 \ di < ?S + e" + by arith + show "integral (cbox a b) (\x. if x \ UNIV then norm (f x) else 0) < ?S + e" + apply (subst if_P) + apply rule + proof (rule *[rule_format]) + show "\(\(x,k)\p. norm (content k *\<^sub>R f x)) - (\(x,k)\p. norm (integral k f))\ < e / 2" + unfolding split_def + apply (rule helplemma) + using d2(2)[rule_format,of p] + using p(1,3) + unfolding tagged_division_of_def split_def + apply auto + done + show "\(\(x, k)\p. content k *\<^sub>R norm (f x)) - integral (cbox a b) (\x. norm(f x))\ < e / 2" + using d1(2)[rule_format,OF conjI[OF p(1,2)]] + by (simp only: real_norm_def) + show "(\(x, k)\p. content k *\<^sub>R norm (f x)) = (\(x, k)\p. norm (content k *\<^sub>R f x))" + apply (rule setsum.cong) + apply (rule refl) + unfolding split_paired_all split_conv + apply (drule tagged_division_ofD(4)[OF p(1)]) + unfolding norm_scaleR + apply (subst abs_of_nonneg) + apply auto + done + show "(\(x, k)\p. norm (integral k f)) \ ?S" + using partial_division_of_tagged_division[of p "cbox a b"] p(1) + apply (subst setsum.over_tagged_division_lemma[OF p(1)]) + apply (simp add: integral_null) + apply (intro cSUP_upper2[OF D(2), of "snd ` p"]) + apply (auto simp: tagged_partial_division_of_def) + done + qed + qed + qed (insert K, auto) + qed +qed + +lemma absolutely_integrable_restrict_univ: + "(\x. if x \ s then f x else (0::'a::banach)) absolutely_integrable_on UNIV \ + f absolutely_integrable_on s" + unfolding absolutely_integrable_on_def if_distrib norm_zero integrable_restrict_univ .. + +lemma absolutely_integrable_add[intro]: + fixes f g :: "'n::euclidean_space \ 'm::euclidean_space" + assumes "f absolutely_integrable_on s" + and "g absolutely_integrable_on s" + shows "(\x. f x + g x) absolutely_integrable_on s" +proof - + let ?P = "\f g::'n \ 'm. f absolutely_integrable_on UNIV \ + g absolutely_integrable_on UNIV \ (\x. f x + g x) absolutely_integrable_on UNIV" + { + presume as: "PROP ?P" + note a = absolutely_integrable_restrict_univ[symmetric] + have *: "\x. (if x \ s then f x else 0) + (if x \ s then g x else 0) = + (if x \ s then f x + g x else 0)" by auto + show ?thesis + apply (subst a) + using as[OF assms[unfolded a[of f] a[of g]]] + apply (simp only: *) + done + } + fix f g :: "'n \ 'm" + assume assms: "f absolutely_integrable_on UNIV" "g absolutely_integrable_on UNIV" + note absolutely_integrable_bounded_variation + from this[OF assms(1)] this[OF assms(2)] guess B1 B2 . note B=this[rule_format] + show "(\x. f x + g x) absolutely_integrable_on UNIV" + apply (rule bounded_variation_absolutely_integrable[of _ "B1+B2"]) + apply (rule integrable_add) + prefer 3 + apply safe + proof goal_cases + case prems: (1 d) + have "\k. k \ d \ f integrable_on k \ g integrable_on k" + apply (drule division_ofD(4)[OF prems]) + apply safe + apply (rule_tac[!] integrable_on_subcbox[of _ UNIV]) + using assms + apply auto + done + then have "(\k\d. norm (integral k (\x. f x + g x))) \ + (\k\d. norm (integral k f)) + (\k\d. norm (integral k g))" + apply - + unfolding setsum.distrib [symmetric] + apply (rule setsum_mono) + apply (subst integral_add) + prefer 3 + apply (rule norm_triangle_ineq) + apply auto + done + also have "\ \ B1 + B2" + using B(1)[OF prems] B(2)[OF prems] by auto + finally show ?case . + qed (insert assms, auto) +qed + +lemma absolutely_integrable_sub[intro]: + fixes f g :: "'n::euclidean_space \ 'm::euclidean_space" + assumes "f absolutely_integrable_on s" + and "g absolutely_integrable_on s" + shows "(\x. f x - g x) absolutely_integrable_on s" + using absolutely_integrable_add[OF assms(1) absolutely_integrable_neg[OF assms(2)]] + by (simp add: algebra_simps) + +lemma absolutely_integrable_linear: + fixes f :: "'m::euclidean_space \ 'n::euclidean_space" + and h :: "'n::euclidean_space \ 'p::euclidean_space" + assumes "f absolutely_integrable_on s" + and "bounded_linear h" + shows "(h \ f) absolutely_integrable_on s" +proof - + { + presume as: "\f::'m \ 'n. \h::'n \ 'p. f absolutely_integrable_on UNIV \ + bounded_linear h \ (h \ f) absolutely_integrable_on UNIV" + note a = absolutely_integrable_restrict_univ[symmetric] + show ?thesis + apply (subst a) + using as[OF assms[unfolded a[of f] a[of g]]] + apply (simp only: o_def if_distrib linear_simps[OF assms(2)]) + done + } + fix f :: "'m \ 'n" + fix h :: "'n \ 'p" + assume assms: "f absolutely_integrable_on UNIV" "bounded_linear h" + from absolutely_integrable_bounded_variation[OF assms(1)] guess B . note B=this + from bounded_linear.pos_bounded[OF assms(2)] guess b .. note b=conjunctD2[OF this] + show "(h \ f) absolutely_integrable_on UNIV" + apply (rule bounded_variation_absolutely_integrable[of _ "B * b"]) + apply (rule integrable_linear[OF _ assms(2)]) + apply safe + proof goal_cases + case prems: (2 d) + have "(\k\d. norm (integral k (h \ f))) \ setsum (\k. norm(integral k f)) d * b" + unfolding setsum_left_distrib + apply (rule setsum_mono) + proof goal_cases + case (1 k) + from division_ofD(4)[OF prems this] + guess u v by (elim exE) note uv=this + have *: "f integrable_on k" + unfolding uv + apply (rule integrable_on_subcbox[of _ UNIV]) + using assms + apply auto + done + note this[unfolded has_integral_integral] + note has_integral_linear[OF this assms(2)] integrable_linear[OF * assms(2)] + note * = has_integral_unique[OF this(2)[unfolded has_integral_integral] this(1)] + show ?case + unfolding * using b by auto + qed + also have "\ \ B * b" + apply (rule mult_right_mono) + using B prems b + apply auto + done + finally show ?case . + qed (insert assms, auto) +qed + +lemma absolutely_integrable_setsum: + fixes f :: "'a \ 'n::euclidean_space \ 'm::euclidean_space" + assumes "finite t" + and "\a. a \ t \ (f a) absolutely_integrable_on s" + shows "(\x. setsum (\a. f a x) t) absolutely_integrable_on s" + using assms(1,2) + by induct auto + +lemma absolutely_integrable_vector_abs: + fixes f :: "'a::euclidean_space => 'b::euclidean_space" + and T :: "'c::euclidean_space \ 'b" + assumes f: "f absolutely_integrable_on s" + shows "(\x. (\i\Basis. \f x\T i\ *\<^sub>R i)) absolutely_integrable_on s" + (is "?Tf absolutely_integrable_on s") +proof - + have if_distrib: "\P A B x. (if P then A else B) *\<^sub>R x = (if P then A *\<^sub>R x else B *\<^sub>R x)" + by simp + have *: "\x. ?Tf x = (\i\Basis. + ((\y. (\j\Basis. (if j = i then y else 0) *\<^sub>R j)) o + (\x. (norm (\j\Basis. (if j = i then f x\T i else 0) *\<^sub>R j)))) x)" + by (simp add: comp_def if_distrib setsum.If_cases) + show ?thesis + unfolding * + apply (rule absolutely_integrable_setsum[OF finite_Basis]) + apply (rule absolutely_integrable_linear) + apply (rule absolutely_integrable_norm) + apply (rule absolutely_integrable_linear[OF f, unfolded o_def]) + apply (auto simp: linear_linear euclidean_eq_iff[where 'a='c] inner_simps intro!: linearI) + done +qed + +lemma absolutely_integrable_max: + fixes f g :: "'m::euclidean_space \ 'n::euclidean_space" + assumes "f absolutely_integrable_on s" + and "g absolutely_integrable_on s" + shows "(\x. (\i\Basis. max (f(x)\i) (g(x)\i) *\<^sub>R i)::'n) absolutely_integrable_on s" +proof - + have *:"\x. (1 / 2) *\<^sub>R (((\i\Basis. \(f x - g x) \ i\ *\<^sub>R i)::'n) + (f x + g x)) = + (\i\Basis. max (f(x)\i) (g(x)\i) *\<^sub>R i)" + unfolding euclidean_eq_iff[where 'a='n] by (auto simp: inner_simps) + note absolutely_integrable_sub[OF assms] absolutely_integrable_add[OF assms] + note absolutely_integrable_vector_abs[OF this(1), where T="\x. x"] this(2) + note absolutely_integrable_add[OF this] + note absolutely_integrable_cmul[OF this, of "1/2"] + then show ?thesis unfolding * . +qed + +lemma absolutely_integrable_min: + fixes f g::"'m::euclidean_space \ 'n::euclidean_space" + assumes "f absolutely_integrable_on s" + and "g absolutely_integrable_on s" + shows "(\x. (\i\Basis. min (f(x)\i) (g(x)\i) *\<^sub>R i)::'n) absolutely_integrable_on s" +proof - + have *:"\x. (1 / 2) *\<^sub>R ((f x + g x) - (\i\Basis. \(f x - g x) \ i\ *\<^sub>R i::'n)) = + (\i\Basis. min (f(x)\i) (g(x)\i) *\<^sub>R i)" + unfolding euclidean_eq_iff[where 'a='n] by (auto simp: inner_simps) + + note absolutely_integrable_add[OF assms] absolutely_integrable_sub[OF assms] + note this(1) absolutely_integrable_vector_abs[OF this(2), where T="\x. x"] + note absolutely_integrable_sub[OF this] + note absolutely_integrable_cmul[OF this,of "1/2"] + then show ?thesis unfolding * . +qed + +lemma absolutely_integrable_abs_eq: + fixes f::"'n::euclidean_space \ 'm::euclidean_space" + shows "f absolutely_integrable_on s \ f integrable_on s \ + (\x. (\i\Basis. \f x\i\ *\<^sub>R i)::'m) integrable_on s" + (is "?l = ?r") +proof + assume ?l + then show ?r + apply - + apply rule + defer + apply (drule absolutely_integrable_vector_abs) + apply auto + done +next + assume ?r + { + presume lem: "\f::'n \ 'm. f integrable_on UNIV \ + (\x. (\i\Basis. \f x\i\ *\<^sub>R i)::'m) integrable_on UNIV \ + f absolutely_integrable_on UNIV" + have *: "\x. (\i\Basis. \(if x \ s then f x else 0) \ i\ *\<^sub>R i) = + (if x \ s then (\i\Basis. \f x \ i\ *\<^sub>R i) else (0::'m))" + unfolding euclidean_eq_iff[where 'a='m] + by auto + show ?l + apply (subst absolutely_integrable_restrict_univ[symmetric]) + apply (rule lem) + unfolding integrable_restrict_univ * + using \?r\ + apply auto + done + } + fix f :: "'n \ 'm" + assume assms: "f integrable_on UNIV" "(\x. (\i\Basis. \f x\i\ *\<^sub>R i)::'m) integrable_on UNIV" + let ?B = "\i\Basis. integral UNIV (\x. (\i\Basis. \f x\i\ *\<^sub>R i)::'m) \ i" + show "f absolutely_integrable_on UNIV" + apply (rule bounded_variation_absolutely_integrable[OF assms(1), where B="?B"]) + apply safe + proof goal_cases + case d: (1 d) + note d'=division_ofD[OF d] + have "(\k\d. norm (integral k f)) \ + (\k\d. setsum (op \ (integral k (\x. (\i\Basis. \f x\i\ *\<^sub>R i)::'m))) Basis)" + apply (rule setsum_mono) + apply (rule order_trans[OF norm_le_l1]) + apply (rule setsum_mono) + unfolding lessThan_iff + proof - + fix k + fix i :: 'm + assume "k \ d" and i: "i \ Basis" + from d'(4)[OF this(1)] guess a b by (elim exE) note ab=this + show "\integral k f \ i\ \ integral k (\x. (\i\Basis. \f x\i\ *\<^sub>R i)::'m) \ i" + apply (rule abs_leI) + unfolding inner_minus_left[symmetric] + defer + apply (subst integral_neg[symmetric]) + apply (rule_tac[1-2] integral_component_le[OF i]) + using integrable_on_subcbox[OF assms(1),of a b] + integrable_on_subcbox[OF assms(2),of a b] i integrable_neg + unfolding ab + apply auto + done + qed + also have "\ \ setsum (op \ (integral UNIV (\x. (\i\Basis. \f x\i\ *\<^sub>R i)::'m))) Basis" + apply (subst setsum.commute) + apply (rule setsum_mono) + proof goal_cases + case (1 j) + have *: "(\x. \i\Basis. \f x\i\ *\<^sub>R i::'m) integrable_on \d" + using integrable_on_subdivision[OF d assms(2)] by auto + have "(\i\d. integral i (\x. \i\Basis. \f x\i\ *\<^sub>R i::'m) \ j) = + integral (\d) (\x. \i\Basis. \f x\i\ *\<^sub>R i::'m) \ j" + unfolding inner_setsum_left[symmetric] integral_combine_division_topdown[OF * d] .. + also have "\ \ integral UNIV (\x. \i\Basis. \f x\i\ *\<^sub>R i::'m) \ j" + apply (rule integral_subset_component_le) + using assms * \j \ Basis\ + apply auto + done + finally show ?case . + qed + finally show ?case . + qed +qed + +lemma nonnegative_absolutely_integrable: + fixes f :: "'n::euclidean_space \ 'm::euclidean_space" + assumes "\x\s. \i\Basis. 0 \ f x \ i" + and "f integrable_on s" + shows "f absolutely_integrable_on s" + unfolding absolutely_integrable_abs_eq + apply rule + apply (rule assms)thm integrable_eq + apply (rule integrable_eq[of _ f]) + using assms + apply (auto simp: euclidean_eq_iff[where 'a='m]) + done + +lemma absolutely_integrable_integrable_bound: + fixes f :: "'n::euclidean_space \ 'm::euclidean_space" + assumes "\x\s. norm (f x) \ g x" + and "f integrable_on s" + and "g integrable_on s" + shows "f absolutely_integrable_on s" +proof - + { + presume *: "\f::'n \ 'm. \g. \x. norm (f x) \ g x \ f integrable_on UNIV \ + g integrable_on UNIV \ f absolutely_integrable_on UNIV" + show ?thesis + apply (subst absolutely_integrable_restrict_univ[symmetric]) + apply (rule *[of _ "\x. if x\s then g x else 0"]) + using assms + unfolding integrable_restrict_univ + apply auto + done + } + fix g + fix f :: "'n \ 'm" + assume assms: "\x. norm (f x) \ g x" "f integrable_on UNIV" "g integrable_on UNIV" + show "f absolutely_integrable_on UNIV" + apply (rule bounded_variation_absolutely_integrable[OF assms(2),where B="integral UNIV g"]) + apply safe + proof goal_cases + case d: (1 d) + note d'=division_ofD[OF d] + have "(\k\d. norm (integral k f)) \ (\k\d. integral k g)" + apply (rule setsum_mono) + apply (rule integral_norm_bound_integral) + apply (drule_tac[!] d'(4)) + apply safe + apply (rule_tac[1-2] integrable_on_subcbox) + using assms + apply auto + done + also have "\ = integral (\d) g" + apply (rule integral_combine_division_bottomup[symmetric]) + apply (rule d) + apply safe + apply (drule d'(4)) + apply safe + apply (rule integrable_on_subcbox[OF assms(3)]) + apply auto + done + also have "\ \ integral UNIV g" + apply (rule integral_subset_le) + defer + apply (rule integrable_on_subdivision[OF d,of _ UNIV]) + prefer 4 + apply rule + apply (rule_tac y="norm (f x)" in order_trans) + using assms + apply auto + done + finally show ?case . + qed +qed + +lemma absolutely_integrable_absolutely_integrable_bound: + fixes f :: "'n::euclidean_space \ 'm::euclidean_space" + and g :: "'n::euclidean_space \ 'p::euclidean_space" + assumes "\x\s. norm (f x) \ norm (g x)" + and "f integrable_on s" + and "g absolutely_integrable_on s" + shows "f absolutely_integrable_on s" + apply (rule absolutely_integrable_integrable_bound[of s f "\x. norm (g x)"]) + using assms + apply auto + done + +lemma absolutely_integrable_inf_real: + assumes "finite k" + and "k \ {}" + and "\i\k. (\x. (fs x i)::real) absolutely_integrable_on s" + shows "(\x. (Inf ((fs x) ` k))) absolutely_integrable_on s" + using assms +proof induct + case (insert a k) + let ?P = "(\x. + if fs x ` k = {} then fs x a + else min (fs x a) (Inf (fs x ` k))) absolutely_integrable_on s" + show ?case + unfolding image_insert + apply (subst Inf_insert_finite) + apply (rule finite_imageI[OF insert(1)]) + proof (cases "k = {}") + case True + then show ?P + apply (subst if_P) + defer + apply (rule insert(5)[rule_format]) + apply auto + done + next + case False + then show ?P + apply (subst if_not_P) + defer + apply (rule absolutely_integrable_min[where 'n=real, unfolded Basis_real_def, simplified]) + defer + apply(rule insert(3)[OF False]) + using insert(5) + apply auto + done + qed +next + case empty + then show ?case by auto +qed + +lemma absolutely_integrable_sup_real: + assumes "finite k" + and "k \ {}" + and "\i\k. (\x. (fs x i)::real) absolutely_integrable_on s" + shows "(\x. (Sup ((fs x) ` k))) absolutely_integrable_on s" + using assms +proof induct + case (insert a k) + let ?P = "(\x. + if fs x ` k = {} then fs x a + else max (fs x a) (Sup (fs x ` k))) absolutely_integrable_on s" + show ?case + unfolding image_insert + apply (subst Sup_insert_finite) + apply (rule finite_imageI[OF insert(1)]) + proof (cases "k = {}") + case True + then show ?P + apply (subst if_P) + defer + apply (rule insert(5)[rule_format]) + apply auto + done + next + case False + then show ?P + apply (subst if_not_P) + defer + apply (rule absolutely_integrable_max[where 'n=real, unfolded Basis_real_def, simplified]) + defer + apply (rule insert(3)[OF False]) + using insert(5) + apply auto + done + qed +qed auto + + +subsection \differentiation under the integral sign\ + +lemma tube_lemma: + assumes "compact K" + assumes "open W" + assumes "{x0} \ K \ W" + shows "\X0. x0 \ X0 \ open X0 \ X0 \ K \ W" +proof - + { + fix y assume "y \ K" + then have "(x0, y) \ W" using assms by auto + with \open W\ + have "\X0 Y. open X0 \ open Y \ x0 \ X0 \ y \ Y \ X0 \ Y \ W" + by (rule open_prod_elim) blast + } + then obtain X0 Y where + *: "\y \ K. open (X0 y) \ open (Y y) \ x0 \ X0 y \ y \ Y y \ X0 y \ Y y \ W" + by metis + from * have "\t\Y ` K. open t" "K \ \(Y ` K)" by auto + with \compact K\ obtain CC where CC: "CC \ Y ` K" "finite CC" "K \ \CC" + by (rule compactE) + then obtain c where c: "\C. C \ CC \ c C \ K \ C = Y (c C)" + by (force intro!: choice) + with * CC show ?thesis + by (force intro!: exI[where x="\C\CC. X0 (c C)"]) (* SLOW *) +qed + +lemma continuous_on_prod_compactE: + fixes fx::"'a::topological_space \ 'b::topological_space \ 'c::metric_space" + and e::real + assumes cont_fx: "continuous_on (U \ C) fx" + assumes "compact C" + assumes [intro]: "x0 \ U" + notes [continuous_intros] = continuous_on_compose2[OF cont_fx] + assumes "e > 0" + obtains X0 where "x0 \ X0" "open X0" + "\x\X0 \ U. \t \ C. dist (fx (x, t)) (fx (x0, t)) \ e" +proof - + define psi where "psi = (\(x, t). dist (fx (x, t)) (fx (x0, t)))" + define W0 where "W0 = {(x, t) \ U \ C. psi (x, t) < e}" + have W0_eq: "W0 = psi -` {.. U \ C" + by (auto simp: vimage_def W0_def) + have "open {.. C) psi" + by (auto intro!: continuous_intros simp: psi_def split_beta') + from this[unfolded continuous_on_open_invariant, rule_format, OF \open {..] + obtain W where W: "open W" "W \ U \ C = W0 \ U \ C" + unfolding W0_eq by blast + have "{x0} \ C \ W \ U \ C" + unfolding W + by (auto simp: W0_def psi_def \0 < e\) + then have "{x0} \ C \ W" by blast + from tube_lemma[OF \compact C\ \open W\ this] + obtain X0 where X0: "x0 \ X0" "open X0" "X0 \ C \ W" + by blast + + have "\x\X0 \ U. \t \ C. dist (fx (x, t)) (fx (x0, t)) \ e" + proof safe + fix x assume x: "x \ X0" "x \ U" + fix t assume t: "t \ C" + have "dist (fx (x, t)) (fx (x0, t)) = psi (x, t)" + by (auto simp: psi_def) + also + { + have "(x, t) \ X0 \ C" + using t x + by auto + also note \\ \ W\ + finally have "(x, t) \ W" . + with t x have "(x, t) \ W \ U \ C" + by blast + also note \W \ U \ C = W0 \ U \ C\ + finally have "psi (x, t) < e" + by (auto simp: W0_def) + } + finally show "dist (fx (x, t)) (fx (x0, t)) \ e" by simp + qed + from X0(1,2) this show ?thesis .. +qed + +lemma integral_continuous_on_param: + fixes f::"'a::topological_space \ 'b::euclidean_space \ 'c::banach" + assumes cont_fx: "continuous_on (U \ cbox a b) (\(x, t). f x t)" + shows "continuous_on U (\x. integral (cbox a b) (f x))" +proof cases + assume "content (cbox a b) \ 0" + then have ne: "cbox a b \ {}" by auto + + note [continuous_intros] = + continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, + unfolded split_beta fst_conv snd_conv] + + show ?thesis + unfolding continuous_on_def + proof (safe intro!: tendstoI) + fix e'::real and x + assume "e' > 0" + define e where "e = e' / (content (cbox a b) + 1)" + have "e > 0" using \e' > 0\ by (auto simp: e_def intro!: divide_pos_pos add_nonneg_pos) + assume "x \ U" + from continuous_on_prod_compactE[OF cont_fx compact_cbox \x \ U\ \0 < e\] + obtain X0 where X0: "x \ X0" "open X0" + and fx_bound: "\y t. y \ X0 \ U \ t \ cbox a b \ norm (f y t - f x t) \ e" + unfolding split_beta fst_conv snd_conv dist_norm + by metis + have "\\<^sub>F y in at x within U. y \ X0 \ U" + using X0(1) X0(2) eventually_at_topological by auto + then show "\\<^sub>F y in at x within U. dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) < e'" + proof eventually_elim + case (elim y) + have "dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) = + norm (integral (cbox a b) (\t. f y t - f x t))" + using elim \x \ U\ + unfolding dist_norm + by (subst integral_diff) + (auto intro!: integrable_continuous continuous_intros) + also have "\ \ e * content (cbox a b)" + using elim \x \ U\ + by (intro integrable_bound) + (auto intro!: fx_bound \x \ U \ less_imp_le[OF \0 < e\] + integrable_continuous continuous_intros) + also have "\ < e'" + using \0 < e'\ \e > 0\ + by (auto simp: e_def divide_simps) + finally show "dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) < e'" . + qed + qed +qed (auto intro!: continuous_on_const) + +lemma eventually_closed_segment: + fixes x0::"'a::real_normed_vector" + assumes "open X0" "x0 \ X0" + shows "\\<^sub>F x in at x0 within U. closed_segment x0 x \ X0" +proof - + from openE[OF assms] + obtain e where e: "0 < e" "ball x0 e \ X0" . + then have "\\<^sub>F x in at x0 within U. x \ ball x0 e" + by (auto simp: dist_commute eventually_at) + then show ?thesis + proof eventually_elim + case (elim x) + have "x0 \ ball x0 e" using \e > 0\ by simp + from convex_ball[unfolded convex_contains_segment, rule_format, OF this elim] + have "closed_segment x0 x \ ball x0 e" . + also note \\ \ X0\ + finally show ?case . + qed +qed + +lemma leibniz_rule: + fixes f::"'a::banach \ 'b::euclidean_space \ 'c::banach" + assumes fx: "\x t. x \ U \ t \ cbox a b \ + ((\x. f x t) has_derivative blinfun_apply (fx x t)) (at x within U)" + assumes integrable_f2: "\x. x \ U \ f x integrable_on cbox a b" + assumes cont_fx: "continuous_on (U \ (cbox a b)) (\(x, t). fx x t)" + assumes [intro]: "x0 \ U" + assumes "convex U" + shows + "((\x. integral (cbox a b) (f x)) has_derivative integral (cbox a b) (fx x0)) (at x0 within U)" + (is "(?F has_derivative ?dF) _") +proof cases + assume "content (cbox a b) \ 0" + then have ne: "cbox a b \ {}" by auto + note [continuous_intros] = + continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, + unfolded split_beta fst_conv snd_conv] + show ?thesis + proof (intro has_derivativeI bounded_linear_scaleR_left tendstoI, fold norm_conv_dist) + have cont_f1: "\t. t \ cbox a b \ continuous_on U (\x. f x t)" + by (auto simp: continuous_on_eq_continuous_within intro!: has_derivative_continuous fx) + note [continuous_intros] = continuous_on_compose2[OF cont_f1] + fix e'::real + assume "e' > 0" + define e where "e = e' / (content (cbox a b) + 1)" + have "e > 0" using \e' > 0\ by (auto simp: e_def intro!: divide_pos_pos add_nonneg_pos) + from continuous_on_prod_compactE[OF cont_fx compact_cbox \x0 \ U\ \e > 0\] + obtain X0 where X0: "x0 \ X0" "open X0" + and fx_bound: "\x t. x \ X0 \ U \ t \ cbox a b \ norm (fx x t - fx x0 t) \ e" + unfolding split_beta fst_conv snd_conv + by (metis dist_norm) + + note eventually_closed_segment[OF \open X0\ \x0 \ X0\, of U] + moreover + have "\\<^sub>F x in at x0 within U. x \ X0" + using \open X0\ \x0 \ X0\ eventually_at_topological by blast + moreover have "\\<^sub>F x in at x0 within U. x \ x0" + by (auto simp: eventually_at_filter) + moreover have "\\<^sub>F x in at x0 within U. x \ U" + by (auto simp: eventually_at_filter) + ultimately + show "\\<^sub>F x in at x0 within U. norm ((?F x - ?F x0 - ?dF (x - x0)) /\<^sub>R norm (x - x0)) < e'" + proof eventually_elim + case (elim x) + from elim have "0 < norm (x - x0)" by simp + have "closed_segment x0 x \ U" + by (rule \convex U\[unfolded convex_contains_segment, rule_format, OF \x0 \ U\ \x \ U\]) + from elim have [intro]: "x \ U" by auto + + have "?F x - ?F x0 - ?dF (x - x0) = + integral (cbox a b) (\y. f x y - f x0 y - fx x0 y (x - x0))" + (is "_ = ?id") + using \x \ x0\ + by (subst blinfun_apply_integral integral_diff, + auto intro!: integrable_diff integrable_f2 continuous_intros + intro: integrable_continuous)+ + also + { + fix t assume t: "t \ (cbox a b)" + have seg: "\t. t \ {0..1} \ x0 + t *\<^sub>R (x - x0) \ X0 \ U" + using \closed_segment x0 x \ U\ + \closed_segment x0 x \ X0\ + by (force simp: closed_segment_def algebra_simps) + from t have deriv: + "((\x. f x t) has_derivative (fx y t)) (at y within X0 \ U)" + if "y \ X0 \ U" for y + unfolding has_vector_derivative_def[symmetric] + using that \x \ X0\ + by (intro has_derivative_within_subset[OF fx]) auto + have "\x \ X0 \ U. onorm (blinfun_apply (fx x t) - (fx x0 t)) \ e" + using fx_bound t + by (auto simp add: norm_blinfun_def fun_diff_def blinfun.bilinear_simps[symmetric]) + from differentiable_bound_linearization[OF seg deriv this] X0 + have "norm (f x t - f x0 t - fx x0 t (x - x0)) \ e * norm (x - x0)" + by (auto simp add: ac_simps) + } + then have "norm ?id \ integral (cbox a b) (\_. e * norm (x - x0))" + by (intro integral_norm_bound_integral) + (auto intro!: continuous_intros integrable_diff integrable_f2 + intro: integrable_continuous) + also have "\ = content (cbox a b) * e * norm (x - x0)" + by simp + also have "\ < e' * norm (x - x0)" + using \e' > 0\ content_pos_le[of a b] + by (intro mult_strict_right_mono[OF _ \0 < norm (x - x0)\]) + (auto simp: divide_simps e_def) + finally have "norm (?F x - ?F x0 - ?dF (x - x0)) < e' * norm (x - x0)" . + then show ?case + by (auto simp: divide_simps) + qed + qed (rule blinfun.bounded_linear_right) +qed (auto intro!: derivative_eq_intros simp: blinfun.bilinear_simps) + +lemma + has_vector_derivative_eq_has_derivative_blinfun: + "(f has_vector_derivative f') (at x within U) \ + (f has_derivative blinfun_scaleR_left f') (at x within U)" + by (simp add: has_vector_derivative_def) + +lemma leibniz_rule_vector_derivative: + fixes f::"real \ 'b::euclidean_space \ 'c::banach" + assumes fx: "\x t. x \ U \ t \ cbox a b \ + ((\x. f x t) has_vector_derivative (fx x t)) (at x within U)" + assumes integrable_f2: "\x. x \ U \ (f x) integrable_on cbox a b" + assumes cont_fx: "continuous_on (U \ cbox a b) (\(x, t). fx x t)" + assumes U: "x0 \ U" "convex U" + shows "((\x. integral (cbox a b) (f x)) has_vector_derivative integral (cbox a b) (fx x0)) + (at x0 within U)" +proof - + note [continuous_intros] = + continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, + unfolded split_beta fst_conv snd_conv] + have *: "blinfun_scaleR_left (integral (cbox a b) (fx x0)) = + integral (cbox a b) (\t. blinfun_scaleR_left (fx x0 t))" + by (subst integral_linear[symmetric]) + (auto simp: has_vector_derivative_def o_def + intro!: integrable_continuous U continuous_intros bounded_linear_intros) + show ?thesis + unfolding has_vector_derivative_eq_has_derivative_blinfun + apply (rule has_derivative_eq_rhs) + apply (rule leibniz_rule[OF _ integrable_f2 _ U, where fx="\x t. blinfun_scaleR_left (fx x t)"]) + using fx cont_fx + apply (auto simp: has_vector_derivative_def * split_beta intro!: continuous_intros) + done +qed + +lemma + has_field_derivative_eq_has_derivative_blinfun: + "(f has_field_derivative f') (at x within U) \ (f has_derivative blinfun_mult_right f') (at x within U)" + by (simp add: has_field_derivative_def) + +lemma leibniz_rule_field_derivative: + fixes f::"'a::{real_normed_field, banach} \ 'b::euclidean_space \ 'a" + assumes fx: "\x t. x \ U \ t \ cbox a b \ ((\x. f x t) has_field_derivative fx x t) (at x within U)" + assumes integrable_f2: "\x. x \ U \ (f x) integrable_on cbox a b" + assumes cont_fx: "continuous_on (U \ (cbox a b)) (\(x, t). fx x t)" + assumes U: "x0 \ U" "convex U" + shows "((\x. integral (cbox a b) (f x)) has_field_derivative integral (cbox a b) (fx x0)) (at x0 within U)" +proof - + note [continuous_intros] = + continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, + unfolded split_beta fst_conv snd_conv] + have *: "blinfun_mult_right (integral (cbox a b) (fx x0)) = + integral (cbox a b) (\t. blinfun_mult_right (fx x0 t))" + by (subst integral_linear[symmetric]) + (auto simp: has_vector_derivative_def o_def + intro!: integrable_continuous U continuous_intros bounded_linear_intros) + show ?thesis + unfolding has_field_derivative_eq_has_derivative_blinfun + apply (rule has_derivative_eq_rhs) + apply (rule leibniz_rule[OF _ integrable_f2 _ U, where fx="\x t. blinfun_mult_right (fx x t)"]) + using fx cont_fx + apply (auto simp: has_field_derivative_def * split_beta intro!: continuous_intros) + done +qed + + +subsection \Exchange uniform limit and integral\ + +lemma + uniform_limit_integral: + fixes f::"'a \ 'b::euclidean_space \ 'c::banach" + assumes u: "uniform_limit (cbox a b) f g F" + assumes c: "\n. continuous_on (cbox a b) (f n)" + assumes [simp]: "F \ bot" + obtains I J where + "\n. (f n has_integral I n) (cbox a b)" + "(g has_integral J) (cbox a b)" + "(I \ J) F" +proof - + have fi[simp]: "f n integrable_on (cbox a b)" for n + by (auto intro!: integrable_continuous assms) + then obtain I where I: "\n. (f n has_integral I n) (cbox a b)" + by atomize_elim (auto simp: integrable_on_def intro!: choice) + + moreover + + have gi[simp]: "g integrable_on (cbox a b)" + by (auto intro!: integrable_continuous uniform_limit_theorem[OF _ u] eventuallyI c) + then obtain J where J: "(g has_integral J) (cbox a b)" + by blast + + moreover + + have "(I \ J) F" + proof cases + assume "content (cbox a b) = 0" + hence "I = (\_. 0)" "J = 0" + by (auto intro!: has_integral_unique I J) + thus ?thesis by simp + next + assume content_nonzero: "content (cbox a b) \ 0" + show ?thesis + proof (rule tendstoI) + fix e::real + assume "e > 0" + define e' where "e' = e / 2" + with \e > 0\ have "e' > 0" by simp + then have "\\<^sub>F n in F. \x\cbox a b. norm (f n x - g x) < e' / content (cbox a b)" + using u content_nonzero content_pos_le[of a b] + by (auto simp: uniform_limit_iff dist_norm) + then show "\\<^sub>F n in F. dist (I n) J < e" + proof eventually_elim + case (elim n) + have "I n = integral (cbox a b) (f n)" + "J = integral (cbox a b) g" + using I[of n] J by (simp_all add: integral_unique) + then have "dist (I n) J = norm (integral (cbox a b) (\x. f n x - g x))" + by (simp add: integral_diff dist_norm) + also have "\ \ integral (cbox a b) (\x. (e' / content (cbox a b)))" + using elim + by (intro integral_norm_bound_integral) + (auto intro!: integrable_diff absolutely_integrable_onI) + also have "\ < e" + using \0 < e\ + by (simp add: e'_def) + finally show ?case . + qed + qed + qed + ultimately show ?thesis .. +qed + + +subsection \Dominated convergence\ + +context +begin + +private lemma dominated_convergence_real: + fixes f :: "nat \ 'n::euclidean_space \ real" + assumes "\k. (f k) integrable_on s" "h integrable_on s" + and "\k. \x \ s. norm (f k x) \ h x" + and "\x \ s. ((\k. f k x) \ g x) sequentially" + shows "g integrable_on s \ ((\k. integral s (f k)) \ integral s g) sequentially" +proof + have bdd_below[simp]: "\x P. x \ s \ bdd_below {f n x |n. P n}" + proof (safe intro!: bdd_belowI) + fix n x show "x \ s \ - h x \ f n x" + using assms(3)[rule_format, of x n] by simp + qed + have bdd_above[simp]: "\x P. x \ s \ bdd_above {f n x |n. P n}" + proof (safe intro!: bdd_aboveI) + fix n x show "x \ s \ f n x \ h x" + using assms(3)[rule_format, of x n] by simp + qed + + have "\m. (\x. Inf {f j x |j. m \ j}) integrable_on s \ + ((\k. integral s (\x. Inf {f j x |j. j \ {m..m + k}})) \ + integral s (\x. Inf {f j x |j. m \ j}))sequentially" + proof (rule monotone_convergence_decreasing, safe) + fix m :: nat + show "bounded {integral s (\x. Inf {f j x |j. j \ {m..m + k}}) |k. True}" + unfolding bounded_iff + apply (rule_tac x="integral s h" in exI) + proof safe + fix k :: nat + show "norm (integral s (\x. Inf {f j x |j. j \ {m..m + k}})) \ integral s h" + apply (rule integral_norm_bound_integral) + unfolding simple_image + apply (rule absolutely_integrable_onD) + apply (rule absolutely_integrable_inf_real) + prefer 5 + unfolding real_norm_def + apply rule + apply (rule cInf_abs_ge) + prefer 5 + apply rule + apply (rule_tac g = h in absolutely_integrable_integrable_bound) + using assms + unfolding real_norm_def + apply auto + done + qed + fix k :: nat + show "(\x. Inf {f j x |j. j \ {m..m + k}}) integrable_on s" + unfolding simple_image + apply (rule absolutely_integrable_onD) + apply (rule absolutely_integrable_inf_real) + prefer 3 + using absolutely_integrable_integrable_bound[OF assms(3,1,2)] + apply auto + done + fix x + assume x: "x \ s" + show "Inf {f j x |j. j \ {m..m + Suc k}} \ Inf {f j x |j. j \ {m..m + k}}" + by (rule cInf_superset_mono) auto + let ?S = "{f j x| j. m \ j}" + show "((\k. Inf {f j x |j. j \ {m..m + k}}) \ Inf ?S) sequentially" + proof (rule LIMSEQ_I, goal_cases) + case r: (1 r) + + have "\y\?S. y < Inf ?S + r" + by (subst cInf_less_iff[symmetric]) (auto simp: \x\s\ r) + then obtain N where N: "f N x < Inf ?S + r" "m \ N" + by blast + + show ?case + apply (rule_tac x=N in exI) + apply safe + proof goal_cases + case prems: (1 n) + have *: "\y ix. y < Inf ?S + r \ Inf ?S \ ix \ ix \ y \ \ix - Inf ?S\ < r" + by arith + show ?case + unfolding real_norm_def + apply (rule *[rule_format, OF N(1)]) + apply (rule cInf_superset_mono, auto simp: \x\s\) [] + apply (rule cInf_lower) + using N prems + apply auto [] + apply simp + done + qed + qed + qed + note dec1 = conjunctD2[OF this] + + have "\m. (\x. Sup {f j x |j. m \ j}) integrable_on s \ + ((\k. integral s (\x. Sup {f j x |j. j \ {m..m + k}})) \ + integral s (\x. Sup {f j x |j. m \ j})) sequentially" + proof (rule monotone_convergence_increasing,safe) + fix m :: nat + show "bounded {integral s (\x. Sup {f j x |j. j \ {m..m + k}}) |k. True}" + unfolding bounded_iff + apply (rule_tac x="integral s h" in exI) + proof safe + fix k :: nat + show "norm (integral s (\x. Sup {f j x |j. j \ {m..m + k}})) \ integral s h" + apply (rule integral_norm_bound_integral) unfolding simple_image + apply (rule absolutely_integrable_onD) + apply(rule absolutely_integrable_sup_real) + prefer 5 unfolding real_norm_def + apply rule + apply (rule cSup_abs_le) + using assms + apply (force simp add: ) + prefer 4 + apply rule + apply (rule_tac g=h in absolutely_integrable_integrable_bound) + using assms + unfolding real_norm_def + apply auto + done + qed + fix k :: nat + show "(\x. Sup {f j x |j. j \ {m..m + k}}) integrable_on s" + unfolding simple_image + apply (rule absolutely_integrable_onD) + apply (rule absolutely_integrable_sup_real) + prefer 3 + using absolutely_integrable_integrable_bound[OF assms(3,1,2)] + apply auto + done + fix x + assume x: "x\s" + show "Sup {f j x |j. j \ {m..m + Suc k}} \ Sup {f j x |j. j \ {m..m + k}}" + by (rule cSup_subset_mono) auto + let ?S = "{f j x| j. m \ j}" + show "((\k. Sup {f j x |j. j \ {m..m + k}}) \ Sup ?S) sequentially" + proof (rule LIMSEQ_I, goal_cases) + case r: (1 r) + have "\y\?S. Sup ?S - r < y" + by (subst less_cSup_iff[symmetric]) (auto simp: r \x\s\) + then obtain N where N: "Sup ?S - r < f N x" "m \ N" + by blast + + show ?case + apply (rule_tac x=N in exI) + apply safe + proof goal_cases + case prems: (1 n) + have *: "\y ix. Sup ?S - r < y \ ix \ Sup ?S \ y \ ix \ \ix - Sup ?S\ < r" + by arith + show ?case + apply simp + apply (rule *[rule_format, OF N(1)]) + apply (rule cSup_subset_mono, auto simp: \x\s\) [] + apply (subst cSup_upper) + using N prems + apply auto + done + qed + qed + qed + note inc1 = conjunctD2[OF this] + + have "g integrable_on s \ + ((\k. integral s (\x. Inf {f j x |j. k \ j})) \ integral s g) sequentially" + apply (rule monotone_convergence_increasing,safe) + apply fact + proof - + show "bounded {integral s (\x. Inf {f j x |j. k \ j}) |k. True}" + unfolding bounded_iff apply(rule_tac x="integral s h" in exI) + proof safe + fix k :: nat + show "norm (integral s (\x. Inf {f j x |j. k \ j})) \ integral s h" + apply (rule integral_norm_bound_integral) + apply fact+ + unfolding real_norm_def + apply rule + apply (rule cInf_abs_ge) + using assms(3) + apply auto + done + qed + fix k :: nat and x + assume x: "x \ s" + + have *: "\x y::real. x \ - y \ - x \ y" by auto + show "Inf {f j x |j. k \ j} \ Inf {f j x |j. Suc k \ j}" + by (intro cInf_superset_mono) (auto simp: \x\s\) + + show "(\k::nat. Inf {f j x |j. k \ j}) \ g x" + proof (rule LIMSEQ_I, goal_cases) + case r: (1 r) + then have "0 + ((\k. integral s (\x. Sup {f j x |j. k \ j})) \ integral s g) sequentially" + apply (rule monotone_convergence_decreasing,safe) + apply fact + proof - + show "bounded {integral s (\x. Sup {f j x |j. k \ j}) |k. True}" + unfolding bounded_iff + apply (rule_tac x="integral s h" in exI) + proof safe + fix k :: nat + show "norm (integral s (\x. Sup {f j x |j. k \ j})) \ integral s h" + apply (rule integral_norm_bound_integral) + apply fact+ + unfolding real_norm_def + apply rule + apply (rule cSup_abs_le) + using assms(3) + apply auto + done + qed + fix k :: nat + fix x + assume x: "x \ s" + + show "Sup {f j x |j. k \ j} \ Sup {f j x |j. Suc k \ j}" + by (rule cSup_subset_mono) (auto simp: \x\s\) + show "((\k. Sup {f j x |j. k \ j}) \ g x) sequentially" + proof (rule LIMSEQ_I, goal_cases) + case r: (1 r) + then have "0k. integral s (f k)) \ integral s g) sequentially" + proof (rule LIMSEQ_I, goal_cases) + case r: (1 r) + from LIMSEQ_D [OF inc2(2) r] guess N1 .. note N1=this[unfolded real_norm_def] + from LIMSEQ_D [OF dec2(2) r] guess N2 .. note N2=this[unfolded real_norm_def] + show ?case + proof (rule_tac x="N1+N2" in exI, safe) + fix n + assume n: "n \ N1 + N2" + have *: "\i0 i i1 g. \i0 - g\ < r \ \i1 - g\ < r \ i0 \ i \ i \ i1 \ \i - g\ < r" + by arith + show "norm (integral s (f n) - integral s g) < r" + unfolding real_norm_def + proof (rule *[rule_format,OF N1[rule_format] N2[rule_format], of n n]) + show "integral s (\x. Inf {f j x |j. n \ j}) \ integral s (f n)" + by (rule integral_le[OF dec1(1) assms(1)]) (auto intro!: cInf_lower) + show "integral s (f n) \ integral s (\x. Sup {f j x |j. n \ j})" + by (rule integral_le[OF assms(1) inc1(1)]) (auto intro!: cSup_upper) + qed (insert n, auto) + qed + qed +qed + +lemma dominated_convergence: + fixes f :: "nat \ 'n::euclidean_space \ 'm::euclidean_space" + assumes f: "\k. (f k) integrable_on s" and h: "h integrable_on s" + and le: "\k. \x \ s. norm (f k x) \ h x" + and conv: "\x \ s. ((\k. f k x) \ g x) sequentially" + shows "g integrable_on s" + and "((\k. integral s (f k)) \ integral s g) sequentially" +proof - + { + fix b :: 'm assume b: "b \ Basis" + have A: "(\x. g x \ b) integrable_on s \ + (\k. integral s (\x. f k x \ b)) \ integral s (\x. g x \ b)" + proof (rule dominated_convergence_real) + fix k :: nat + from f show "(\x. f k x \ b) integrable_on s" by (rule integrable_component) + next + from h show "h integrable_on s" . + next + fix k :: nat + show "\x\s. norm (f k x \ b) \ h x" + proof + fix x assume x: "x \ s" + have "norm (f k x \ b) \ norm (f k x)" by (simp add: Basis_le_norm b) + also have "\ \ h x" using le[of k] x by simp + finally show "norm (f k x \ b) \ h x" . + qed + next + from conv show "\x\s. (\k. f k x \ b) \ g x \ b" + by (auto intro!: tendsto_inner) + qed + have B: "integral s ((\x. x *\<^sub>R b) \ (\x. f k x \ b)) = integral s (\x. f k x \ b) *\<^sub>R b" + for k by (rule integral_linear) + (simp_all add: f integrable_component bounded_linear_scaleR_left) + have C: "integral s ((\x. x *\<^sub>R b) \ (\x. g x \ b)) = integral s (\x. g x \ b) *\<^sub>R b" + using A by (intro integral_linear) + (simp_all add: integrable_component bounded_linear_scaleR_left) + note A B C + } note A = this + + show "g integrable_on s" by (rule integrable_componentwise) (insert A, blast) + with A f show "((\k. integral s (f k)) \ integral s g) sequentially" + by (subst (1 2) integral_componentwise) + (auto intro!: tendsto_setsum tendsto_scaleR simp: o_def) +qed + +lemma has_integral_dominated_convergence: + fixes f :: "nat \ 'n::euclidean_space \ 'm::euclidean_space" + assumes "\k. (f k has_integral y k) s" "h integrable_on s" + "\k. \x\s. norm (f k x) \ h x" "\x\s. (\k. f k x) \ g x" + and x: "y \ x" + shows "(g has_integral x) s" +proof - + have int_f: "\k. (f k) integrable_on s" + using assms by (auto simp: integrable_on_def) + have "(g has_integral (integral s g)) s" + by (intro integrable_integral dominated_convergence[OF int_f assms(2)]) fact+ + moreover have "integral s g = x" + proof (rule LIMSEQ_unique) + show "(\i. integral s (f i)) \ x" + using integral_unique[OF assms(1)] x by simp + show "(\i. integral s (f i)) \ integral s g" + by (intro dominated_convergence[OF int_f assms(2)]) fact+ + qed + ultimately show ?thesis + by simp +qed + +end + + +subsection \Integration by parts\ + +lemma integration_by_parts_interior_strong: + assumes bilinear: "bounded_bilinear (prod :: _ \ _ \ 'b :: banach)" + assumes s: "finite s" and le: "a \ b" + assumes cont [continuous_intros]: "continuous_on {a..b} f" "continuous_on {a..b} g" + assumes deriv: "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" + "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" + assumes int: "((\x. prod (f x) (g' x)) has_integral + (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}" + shows "((\x. prod (f' x) (g x)) has_integral y) {a..b}" +proof - + interpret bounded_bilinear prod by fact + have "((\x. prod (f x) (g' x) + prod (f' x) (g x)) has_integral + (prod (f b) (g b) - prod (f a) (g a))) {a..b}" + using deriv by (intro fundamental_theorem_of_calculus_interior_strong[OF s le]) + (auto intro!: continuous_intros continuous_on has_vector_derivative) + from has_integral_sub[OF this int] show ?thesis by (simp add: algebra_simps) +qed + +lemma integration_by_parts_interior: + assumes "bounded_bilinear (prod :: _ \ _ \ 'b :: banach)" "a \ b" + "continuous_on {a..b} f" "continuous_on {a..b} g" + assumes "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" + "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" + assumes "((\x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}" + shows "((\x. prod (f' x) (g x)) has_integral y) {a..b}" + by (rule integration_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all) + +lemma integration_by_parts: + assumes "bounded_bilinear (prod :: _ \ _ \ 'b :: banach)" "a \ b" + "continuous_on {a..b} f" "continuous_on {a..b} g" + assumes "\x. x\{a..b} \ (f has_vector_derivative f' x) (at x)" + "\x. x\{a..b} \ (g has_vector_derivative g' x) (at x)" + assumes "((\x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}" + shows "((\x. prod (f' x) (g x)) has_integral y) {a..b}" + by (rule integration_by_parts_interior[of _ _ _ f g f' g']) (insert assms, simp_all) + +lemma integrable_by_parts_interior_strong: + assumes bilinear: "bounded_bilinear (prod :: _ \ _ \ 'b :: banach)" + assumes s: "finite s" and le: "a \ b" + assumes cont [continuous_intros]: "continuous_on {a..b} f" "continuous_on {a..b} g" + assumes deriv: "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" + "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" + assumes int: "(\x. prod (f x) (g' x)) integrable_on {a..b}" + shows "(\x. prod (f' x) (g x)) integrable_on {a..b}" +proof - + from int obtain I where "((\x. prod (f x) (g' x)) has_integral I) {a..b}" + unfolding integrable_on_def by blast + hence "((\x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - + (prod (f b) (g b) - prod (f a) (g a) - I))) {a..b}" by simp + from integration_by_parts_interior_strong[OF assms(1-7) this] + show ?thesis unfolding integrable_on_def by blast +qed + +lemma integrable_by_parts_interior: + assumes "bounded_bilinear (prod :: _ \ _ \ 'b :: banach)" "a \ b" + "continuous_on {a..b} f" "continuous_on {a..b} g" + assumes "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" + "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" + assumes "(\x. prod (f x) (g' x)) integrable_on {a..b}" + shows "(\x. prod (f' x) (g x)) integrable_on {a..b}" + by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all) + +lemma integrable_by_parts: + assumes "bounded_bilinear (prod :: _ \ _ \ 'b :: banach)" "a \ b" + "continuous_on {a..b} f" "continuous_on {a..b} g" + assumes "\x. x\{a..b} \ (f has_vector_derivative f' x) (at x)" + "\x. x\{a..b} \ (g has_vector_derivative g' x) (at x)" + assumes "(\x. prod (f x) (g' x)) integrable_on {a..b}" + shows "(\x. prod (f' x) (g x)) integrable_on {a..b}" + by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all) + + +subsection \Integration by substitution\ + + +lemma has_integral_substitution_strong: + fixes f :: "real \ 'a::euclidean_space" and g :: "real \ real" + assumes s: "finite s" and le: "a \ b" "c \ d" "g a \ g b" + and subset: "g ` {a..b} \ {c..d}" + and f [continuous_intros]: "continuous_on {c..d} f" + and g [continuous_intros]: "continuous_on {a..b} g" + and deriv [derivative_intros]: + "\x. x \ {a..b} - s \ (g has_field_derivative g' x) (at x within {a..b})" + shows "((\x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f)) {a..b}" +proof - + let ?F = "\x. integral {c..g x} f" + have cont_int: "continuous_on {a..b} ?F" + by (rule continuous_on_compose2[OF _ g subset] indefinite_integral_continuous + f integrable_continuous_real)+ + have deriv: "(((\x. integral {c..x} f) \ g) has_vector_derivative g' x *\<^sub>R f (g x)) + (at x within {a..b})" if "x \ {a..b} - s" for x + apply (rule has_vector_derivative_eq_rhs) + apply (rule vector_diff_chain_within) + apply (subst has_field_derivative_iff_has_vector_derivative [symmetric]) + apply (rule deriv that)+ + apply (rule has_vector_derivative_within_subset) + apply (rule integral_has_vector_derivative f)+ + using that le subset + apply blast+ + done + have deriv: "(?F has_vector_derivative g' x *\<^sub>R f (g x)) + (at x)" if "x \ {a..b} - (s \ {a,b})" for x + using deriv[of x] that by (simp add: at_within_closed_interval o_def) + + + have "((\x. g' x *\<^sub>R f (g x)) has_integral (?F b - ?F a)) {a..b}" + using le cont_int s deriv cont_int + by (intro fundamental_theorem_of_calculus_interior_strong[of "s \ {a,b}"]) simp_all + also from subset have "g x \ {c..d}" if "x \ {a..b}" for x using that by blast + from this[of a] this[of b] le have "c \ g a" "g b \ d" by auto + hence "integral {c..g a} f + integral {g a..g b} f = integral {c..g b} f" + by (intro integral_combine integrable_continuous_real continuous_on_subset[OF f] le) simp_all + hence "integral {c..g b} f - integral {c..g a} f = integral {g a..g b} f" + by (simp add: algebra_simps) + finally show ?thesis . +qed + +lemma has_integral_substitution: + fixes f :: "real \ 'a::euclidean_space" and g :: "real \ real" + assumes "a \ b" "c \ d" "g a \ g b" "g ` {a..b} \ {c..d}" + and "continuous_on {c..d} f" + and "\x. x \ {a..b} \ (g has_field_derivative g' x) (at x within {a..b})" + shows "((\x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f)) {a..b}" + by (intro has_integral_substitution_strong[of "{}" a b c d] assms) + (auto intro: DERIV_continuous_on assms) + + +subsection \Compute a double integral using iterated integrals and switching the order of integration\ + +lemma setcomp_dot1: "{z. P (z \ (i,0))} = {(x,y). P(x \ i)}" + by auto + +lemma setcomp_dot2: "{z. P (z \ (0,i))} = {(x,y). P(y \ i)}" + by auto + +lemma Sigma_Int_Paircomp1: "(Sigma A B) \ {(x, y). P x} = Sigma (A \ {x. P x}) B" + by blast + +lemma Sigma_Int_Paircomp2: "(Sigma A B) \ {(x, y). P y} = Sigma A (\z. B z \ {y. P y})" + by blast + +lemma continuous_on_imp_integrable_on_Pair1: + fixes f :: "_ \ 'b::banach" + assumes con: "continuous_on (cbox (a,c) (b,d)) f" and x: "x \ cbox a b" + shows "(\y. f (x, y)) integrable_on (cbox c d)" +proof - + have "f \ (\y. (x, y)) integrable_on (cbox c d)" + apply (rule integrable_continuous) + apply (rule continuous_on_compose [OF _ continuous_on_subset [OF con]]) + using x + apply (auto intro: continuous_on_Pair continuous_on_const continuous_on_id continuous_on_subset con) + done + then show ?thesis + by (simp add: o_def) +qed + +lemma integral_integrable_2dim: + fixes f :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::banach" + assumes "continuous_on (cbox (a,c) (b,d)) f" + shows "(\x. integral (cbox c d) (\y. f (x,y))) integrable_on cbox a b" +proof (cases "content(cbox c d) = 0") +case True + then show ?thesis + by (simp add: True integrable_const) +next + case False + have uc: "uniformly_continuous_on (cbox (a,c) (b,d)) f" + by (simp add: assms compact_cbox compact_uniformly_continuous) + { fix x::'a and e::real + assume x: "x \ cbox a b" and e: "0 < e" + then have e2_gt: "0 < e / 2 / content (cbox c d)" and e2_less: "e / 2 / content (cbox c d) * content (cbox c d) < e" + by (auto simp: False content_lt_nz e) + then obtain dd + where dd: "\x x'. \x\cbox (a, c) (b, d); x'\cbox (a, c) (b, d); norm (x' - x) < dd\ + \ norm (f x' - f x) \ e / (2 * content (cbox c d))" "dd>0" + using uc [unfolded uniformly_continuous_on_def, THEN spec, of "e / (2 * content (cbox c d))"] + by (auto simp: dist_norm intro: less_imp_le) + have "\delta>0. \x'\cbox a b. norm (x' - x) < delta \ norm (integral (cbox c d) (\u. f (x', u) - f (x, u))) < e" + apply (rule_tac x=dd in exI) + using dd e2_gt assms x + apply clarify + apply (rule le_less_trans [OF _ e2_less]) + apply (rule integrable_bound) + apply (auto intro: integrable_diff continuous_on_imp_integrable_on_Pair1) + done + } note * = this + show ?thesis + apply (rule integrable_continuous) + apply (simp add: * continuous_on_iff dist_norm integral_diff [symmetric] continuous_on_imp_integrable_on_Pair1 [OF assms]) + done +qed + +lemma norm_diff2: "\y = y1 + y2; x = x1 + x2; e = e1 + e2; norm(y1 - x1) \ e1; norm(y2 - x2) \ e2\ + \ norm(y - x) \ e" +using norm_triangle_mono [of "y1 - x1" "e1" "y2 - x2" "e2"] + by (simp add: add_diff_add) + +lemma integral_split: + fixes f :: "'a::euclidean_space \ 'b::{real_normed_vector,complete_space}" + assumes f: "f integrable_on (cbox a b)" + and k: "k \ Basis" + shows "integral (cbox a b) f = + integral (cbox a b \ {x. x\k \ c}) f + + integral (cbox a b \ {x. x\k \ c}) f" + apply (rule integral_unique [OF has_integral_split [where c=c]]) + using k f + apply (auto simp: has_integral_integral [symmetric]) + done + +lemma integral_swap_operative: + fixes f :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::banach" + assumes f: "continuous_on s f" and e: "0 < e" + shows "comm_monoid.operative (op \) True + (\k. \a b c d. + cbox (a,c) (b,d) \ k \ cbox (a,c) (b,d) \ s + \ norm(integral (cbox (a,c) (b,d)) f - + integral (cbox a b) (\x. integral (cbox c d) (\y. f((x,y))))) + \ e * content (cbox (a,c) (b,d)))" +proof (auto simp: comm_monoid.operative_def[OF comm_monoid_and]) + fix a::'a and c::'b and b::'a and d::'b and u::'a and v::'a and w::'b and z::'b + assume c0: "content (cbox (a, c) (b, d)) = 0" + and cb1: "cbox (u, w) (v, z) \ cbox (a, c) (b, d)" + and cb2: "cbox (u, w) (v, z) \ s" + have c0': "content (cbox (u, w) (v, z)) = 0" + by (fact content_0_subset [OF c0 cb1]) + show "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y)))) + \ e * content (cbox (u,w) (v,z))" + using content_cbox_pair_eq0_D [OF c0'] + by (force simp add: c0') +next + fix a::'a and c::'b and b::'a and d::'b + and M::real and i::'a and j::'b + and u::'a and v::'a and w::'b and z::'b + assume ij: "(i,j) \ Basis" + and n1: "\a' b' c' d'. + cbox (a',c') (b',d') \ cbox (a,c) (b,d) \ + cbox (a',c') (b',d') \ {x. x \ (i,j) \ M} \ cbox (a',c') (b',d') \ s \ + norm (integral (cbox (a',c') (b',d')) f - integral (cbox a' b') (\x. integral (cbox c' d') (\y. f (x,y)))) + \ e * content (cbox (a',c') (b',d'))" + and n2: "\a' b' c' d'. + cbox (a',c') (b',d') \ cbox (a,c) (b,d) \ + cbox (a',c') (b',d') \ {x. M \ x \ (i,j)} \ cbox (a',c') (b',d') \ s \ + norm (integral (cbox (a',c') (b',d')) f - integral (cbox a' b') (\x. integral (cbox c' d') (\y. f (x,y)))) + \ e * content (cbox (a',c') (b',d'))" + and subs: "cbox (u,w) (v,z) \ cbox (a,c) (b,d)" "cbox (u,w) (v,z) \ s" + have fcont: "continuous_on (cbox (u, w) (v, z)) f" + using assms(1) continuous_on_subset subs(2) by blast + then have fint: "f integrable_on cbox (u, w) (v, z)" + by (metis integrable_continuous) + consider "i \ Basis" "j=0" | "j \ Basis" "i=0" using ij + by (auto simp: Euclidean_Space.Basis_prod_def) + then show "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x,y)))) + \ e * content (cbox (u,w) (v,z))" (is ?normle) + proof cases + case 1 + then have e: "e * content (cbox (u, w) (v, z)) = + e * (content (cbox u v \ {x. x \ i \ M}) * content (cbox w z)) + + e * (content (cbox u v \ {x. M \ x \ i}) * content (cbox w z))" + by (simp add: content_split [where c=M] content_Pair algebra_simps) + have *: "integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y))) = + integral (cbox u v \ {x. x \ i \ M}) (\x. integral (cbox w z) (\y. f (x, y))) + + integral (cbox u v \ {x. M \ x \ i}) (\x. integral (cbox w z) (\y. f (x, y)))" + using 1 f subs integral_integrable_2dim continuous_on_subset + by (blast intro: integral_split) + show ?normle + apply (rule norm_diff2 [OF integral_split [where c=M, OF fint ij] * e]) + using 1 subs + apply (simp_all add: cbox_Pair_eq setcomp_dot1 [of "\u. M\u"] setcomp_dot1 [of "\u. u\M"] Sigma_Int_Paircomp1) + apply (simp_all add: interval_split ij) + apply (simp_all add: cbox_Pair_eq [symmetric] content_Pair [symmetric]) + apply (force simp add: interval_split [symmetric] intro!: n1 [rule_format]) + apply (force simp add: interval_split [symmetric] intro!: n2 [rule_format]) + done + next + case 2 + then have e: "e * content (cbox (u, w) (v, z)) = + e * (content (cbox u v) * content (cbox w z \ {x. x \ j \ M})) + + e * (content (cbox u v) * content (cbox w z \ {x. M \ x \ j}))" + by (simp add: content_split [where c=M] content_Pair algebra_simps) + have "(\x. integral (cbox w z \ {x. x \ j \ M}) (\y. f (x, y))) integrable_on cbox u v" + "(\x. integral (cbox w z \ {x. M \ x \ j}) (\y. f (x, y))) integrable_on cbox u v" + using 2 subs + apply (simp_all add: interval_split) + apply (rule_tac [!] integral_integrable_2dim [OF continuous_on_subset [OF f]]) + apply (auto simp: cbox_Pair_eq interval_split [symmetric]) + done + with 2 have *: "integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y))) = + integral (cbox u v) (\x. integral (cbox w z \ {x. x \ j \ M}) (\y. f (x, y))) + + integral (cbox u v) (\x. integral (cbox w z \ {x. M \ x \ j}) (\y. f (x, y)))" + by (simp add: integral_add [symmetric] integral_split [symmetric] + continuous_on_imp_integrable_on_Pair1 [OF fcont] cong: integral_cong) + show ?normle + apply (rule norm_diff2 [OF integral_split [where c=M, OF fint ij] * e]) + using 2 subs + apply (simp_all add: cbox_Pair_eq setcomp_dot2 [of "\u. M\u"] setcomp_dot2 [of "\u. u\M"] Sigma_Int_Paircomp2) + apply (simp_all add: interval_split ij) + apply (simp_all add: cbox_Pair_eq [symmetric] content_Pair [symmetric]) + apply (force simp add: interval_split [symmetric] intro!: n1 [rule_format]) + apply (force simp add: interval_split [symmetric] intro!: n2 [rule_format]) + done + qed +qed + +lemma integral_Pair_const: + "integral (cbox (a,c) (b,d)) (\x. k) = + integral (cbox a b) (\x. integral (cbox c d) (\y. k))" + by (simp add: content_Pair) + +lemma norm_minus2: "norm (x1-x2, y1-y2) = norm (x2-x1, y2-y1)" + by (simp add: norm_minus_eqI) + +lemma integral_prod_continuous: + fixes f :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::banach" + assumes "continuous_on (cbox (a,c) (b,d)) f" (is "continuous_on ?CBOX f") + shows "integral (cbox (a,c) (b,d)) f = integral (cbox a b) (\x. integral (cbox c d) (\y. f(x,y)))" +proof (cases "content ?CBOX = 0") + case True + then show ?thesis + by (auto simp: content_Pair) +next + case False + then have cbp: "content ?CBOX > 0" + using content_lt_nz by blast + have "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x,y)))) = 0" + proof (rule dense_eq0_I, simp) + fix e::real assume "0 < e" + with cbp have e': "0 < e / content ?CBOX" + by simp + have f_int_acbd: "f integrable_on cbox (a,c) (b,d)" + by (rule integrable_continuous [OF assms]) + { fix p + assume p: "p division_of cbox (a,c) (b,d)" + note opd1 = comm_monoid_set.operative_division [OF comm_monoid_set_and integral_swap_operative [OF assms e'], THEN iffD1, + THEN spec, THEN spec, THEN spec, THEN spec, of p "(a,c)" "(b,d)" a c b d] + have "(\t u v w z. + \t \ p; cbox (u,w) (v,z) \ t; cbox (u,w) (v,z) \ cbox (a,c) (b,d)\ \ + norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x,y)))) + \ e * content (cbox (u,w) (v,z)) / content?CBOX) + \ + norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x,y)))) \ e" + using opd1 [OF p] False by (simp add: comm_monoid_set_F_and) + } note op_acbd = this + { fix k::real and p and u::'a and v w and z::'b and t1 t2 l + assume k: "0 < k" + and nf: "\x y u v. + \x \ cbox a b; y \ cbox c d; u \ cbox a b; v\cbox c d; norm (u-x, v-y) < k\ + \ norm (f(u,v) - f(x,y)) < e / (2 * (content ?CBOX))" + and p_acbd: "p tagged_division_of cbox (a,c) (b,d)" + and fine: "(\x. ball x k) fine p" "((t1,t2), l) \ p" + and uwvz_sub: "cbox (u,w) (v,z) \ l" "cbox (u,w) (v,z) \ cbox (a,c) (b,d)" + have t: "t1 \ cbox a b" "t2 \ cbox c d" + by (meson fine p_acbd cbox_Pair_iff tag_in_interval)+ + have ls: "l \ ball (t1,t2) k" + using fine by (simp add: fine_def Ball_def) + { fix x1 x2 + assume xuvwz: "x1 \ cbox u v" "x2 \ cbox w z" + then have x: "x1 \ cbox a b" "x2 \ cbox c d" + using uwvz_sub by auto + have "norm (x1 - t1, x2 - t2) < k" + using xuvwz ls uwvz_sub unfolding ball_def + by (force simp add: cbox_Pair_eq dist_norm norm_minus2) + then have "norm (f (x1,x2) - f (t1,t2)) \ e / content ?CBOX / 2" + using nf [OF t x] by simp + } note nf' = this + have f_int_uwvz: "f integrable_on cbox (u,w) (v,z)" + using f_int_acbd uwvz_sub integrable_on_subcbox by blast + have f_int_uv: "\x. x \ cbox u v \ (\y. f (x,y)) integrable_on cbox w z" + using assms continuous_on_subset uwvz_sub + by (blast intro: continuous_on_imp_integrable_on_Pair1) + have 1: "norm (integral (cbox (u,w) (v,z)) f - integral (cbox (u,w) (v,z)) (\x. f (t1,t2))) + \ e * content (cbox (u,w) (v,z)) / content ?CBOX / 2" + apply (simp only: integral_diff [symmetric] f_int_uwvz integrable_const) + apply (rule order_trans [OF integrable_bound [of "e / content ?CBOX / 2"]]) + using cbp e' nf' + apply (auto simp: integrable_diff f_int_uwvz integrable_const) + done + have int_integrable: "(\x. integral (cbox w z) (\y. f (x, y))) integrable_on cbox u v" + using assms integral_integrable_2dim continuous_on_subset uwvz_sub(2) by blast + have normint_wz: + "\x. x \ cbox u v \ + norm (integral (cbox w z) (\y. f (x, y)) - integral (cbox w z) (\y. f (t1, t2))) + \ e * content (cbox w z) / content (cbox (a, c) (b, d)) / 2" + apply (simp only: integral_diff [symmetric] f_int_uv integrable_const) + apply (rule order_trans [OF integrable_bound [of "e / content ?CBOX / 2"]]) + using cbp e' nf' + apply (auto simp: integrable_diff f_int_uv integrable_const) + done + have "norm (integral (cbox u v) + (\x. integral (cbox w z) (\y. f (x,y)) - integral (cbox w z) (\y. f (t1,t2)))) + \ e * content (cbox w z) / content ?CBOX / 2 * content (cbox u v)" + apply (rule integrable_bound [OF _ _ normint_wz]) + using cbp e' + apply (auto simp: divide_simps content_pos_le integrable_diff int_integrable integrable_const) + done + also have "... \ e * content (cbox (u,w) (v,z)) / content ?CBOX / 2" + by (simp add: content_Pair divide_simps) + finally have 2: "norm (integral (cbox u v) (\x. integral (cbox w z) (\y. f (x,y))) - + integral (cbox u v) (\x. integral (cbox w z) (\y. f (t1,t2)))) + \ e * content (cbox (u,w) (v,z)) / content ?CBOX / 2" + by (simp only: integral_diff [symmetric] int_integrable integrable_const) + have norm_xx: "\x' = y'; norm(x - x') \ e/2; norm(y - y') \ e/2\ \ norm(x - y) \ e" for x::'c and y x' y' e + using norm_triangle_mono [of "x-y'" "e/2" "y'-y" "e/2"] real_sum_of_halves + by (simp add: norm_minus_commute) + have "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x,y)))) + \ e * content (cbox (u,w) (v,z)) / content ?CBOX" + by (rule norm_xx [OF integral_Pair_const 1 2]) + } note * = this + show "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x,y)))) \ e" + using compact_uniformly_continuous [OF assms compact_cbox] + apply (simp add: uniformly_continuous_on_def dist_norm) + apply (drule_tac x="e / 2 / content?CBOX" in spec) + using cbp \0 < e\ + apply (auto simp: zero_less_mult_iff) + apply (rename_tac k) + apply (rule_tac e1=k in fine_division_exists [OF gauge_ball, where a = "(a,c)" and b = "(b,d)"]) + apply assumption + apply (rule op_acbd) + apply (erule division_of_tagged_division) + using * + apply auto + done + qed + then show ?thesis + by simp +qed + +lemma swap_continuous: + assumes "continuous_on (cbox (a,c) (b,d)) (\(x,y). f x y)" + shows "continuous_on (cbox (c,a) (d,b)) (\(x, y). f y x)" +proof - + have "(\(x, y). f y x) = (\(x, y). f x y) \ prod.swap" + by auto + then show ?thesis + apply (rule ssubst) + apply (rule continuous_on_compose) + apply (simp add: split_def) + apply (rule continuous_intros | simp add: assms)+ + done +qed + +lemma integral_swap_2dim: + fixes f :: "['a::euclidean_space, 'b::euclidean_space] \ 'c::banach" + assumes "continuous_on (cbox (a,c) (b,d)) (\(x,y). f x y)" + shows "integral (cbox (a, c) (b, d)) (\(x, y). f x y) = integral (cbox (c, a) (d, b)) (\(x, y). f y x)" +proof - + have "((\(x, y). f x y) has_integral integral (cbox (c, a) (d, b)) (\(x, y). f y x)) (prod.swap ` (cbox (c, a) (d, b)))" + apply (rule has_integral_twiddle [of 1 prod.swap prod.swap "\(x,y). f y x" "integral (cbox (c, a) (d, b)) (\(x, y). f y x)", simplified]) + apply (auto simp: isCont_swap content_Pair has_integral_integral [symmetric] integrable_continuous swap_continuous assms) + done + then show ?thesis + by force +qed + +theorem integral_swap_continuous: + fixes f :: "['a::euclidean_space, 'b::euclidean_space] \ 'c::banach" + assumes "continuous_on (cbox (a,c) (b,d)) (\(x,y). f x y)" + shows "integral (cbox a b) (\x. integral (cbox c d) (f x)) = + integral (cbox c d) (\y. integral (cbox a b) (\x. f x y))" +proof - + have "integral (cbox a b) (\x. integral (cbox c d) (f x)) = integral (cbox (a, c) (b, d)) (\(x, y). f x y)" + using integral_prod_continuous [OF assms] by auto + also have "... = integral (cbox (c, a) (d, b)) (\(x, y). f y x)" + by (rule integral_swap_2dim [OF assms]) + also have "... = integral (cbox c d) (\y. integral (cbox a b) (\x. f x y))" + using integral_prod_continuous [OF swap_continuous] assms + by auto + finally show ?thesis . +qed + + +subsection \Definite integrals for exponential and power function\ + +lemma has_integral_exp_minus_to_infinity: + assumes a: "a > 0" + shows "((\x::real. exp (-a*x)) has_integral exp (-a*c)/a) {c..}" +proof - + define f where "f = (\k x. if x \ {c..real k} then exp (-a*x) else 0)" + + { + fix k :: nat assume k: "of_nat k \ c" + from k a + have "((\x. exp (-a*x)) has_integral (-exp (-a*real k)/a - (-exp (-a*c)/a))) {c..real k}" + by (intro fundamental_theorem_of_calculus) + (auto intro!: derivative_eq_intros + simp: has_field_derivative_iff_has_vector_derivative [symmetric]) + hence "(f k has_integral (exp (-a*c)/a - exp (-a*real k)/a)) {c..}" unfolding f_def + by (subst has_integral_restrict) simp_all + } note has_integral_f = this + + have [simp]: "f k = (\_. 0)" if "of_nat k < c" for k using that by (auto simp: fun_eq_iff f_def) + have integral_f: "integral {c..} (f k) = + (if real k \ c then exp (-a*c)/a - exp (-a*real k)/a else 0)" + for k using integral_unique[OF has_integral_f[of k]] by simp + + have A: "(\x. exp (-a*x)) integrable_on {c..} \ + (\k. integral {c..} (f k)) \ integral {c..} (\x. exp (-a*x))" + proof (intro monotone_convergence_increasing allI ballI) + fix k ::nat + have "(\x. exp (-a*x)) integrable_on {c..of_real k}" (is ?P) + unfolding f_def by (auto intro!: continuous_intros integrable_continuous_real) + hence int: "(f k) integrable_on {c..of_real k}" + by (rule integrable_eq[rotated]) (simp add: f_def) + show "f k integrable_on {c..}" + by (rule integrable_on_superset[OF _ _ int]) (auto simp: f_def) + next + fix x assume x: "x \ {c..}" + have "sequentially \ principal {nat \x\..}" unfolding at_top_def by (simp add: Inf_lower) + also have "{nat \x\..} \ {k. x \ real k}" by auto + also have "inf (principal \) (principal {k. \x \ real k}) = + principal ({k. x \ real k} \ {k. \x \ real k})" by simp + also have "{k. x \ real k} \ {k. \x \ real k} = {}" by blast + finally have "inf sequentially (principal {k. \x \ real k}) = bot" + by (simp add: inf.coboundedI1 bot_unique) + with x show "(\k. f k x) \ exp (-a*x)" unfolding f_def + by (intro filterlim_If) auto + next + have "\integral {c..} (f k)\ \ exp (-a*c)/a" for k + proof (cases "c > of_nat k") + case False + hence "abs (integral {c..} (f k)) = abs (exp (- (a * c)) / a - exp (- (a * real k)) / a)" + by (simp add: integral_f) + also have "abs (exp (- (a * c)) / a - exp (- (a * real k)) / a) = + exp (- (a * c)) / a - exp (- (a * real k)) / a" + using False a by (intro abs_of_nonneg) (simp_all add: field_simps) + also have "\ \ exp (- a * c) / a" using a by simp + finally show ?thesis . + qed (insert a, simp_all add: integral_f) + thus "bounded {integral {c..} (f k) |k. True}" + by (intro bounded_realI[of _ "exp (-a*c)/a"]) auto + qed (auto simp: f_def) + + from eventually_gt_at_top[of "nat \c\"] have "eventually (\k. of_nat k > c) sequentially" + by eventually_elim linarith + hence "eventually (\k. exp (-a*c)/a - exp (-a * of_nat k)/a = integral {c..} (f k)) sequentially" + by eventually_elim (simp add: integral_f) + moreover have "(\k. exp (-a*c)/a - exp (-a * of_nat k)/a) \ exp (-a*c)/a - 0/a" + by (intro tendsto_intros filterlim_compose[OF exp_at_bot] + filterlim_tendsto_neg_mult_at_bot[OF tendsto_const] filterlim_real_sequentially)+ + (insert a, simp_all) + ultimately have "(\k. integral {c..} (f k)) \ exp (-a*c)/a - 0/a" + by (rule Lim_transform_eventually) + from LIMSEQ_unique[OF conjunct2[OF A] this] + have "integral {c..} (\x. exp (-a*x)) = exp (-a*c)/a" by simp + with conjunct1[OF A] show ?thesis + by (simp add: has_integral_integral) +qed + +lemma integrable_on_exp_minus_to_infinity: "a > 0 \ (\x. exp (-a*x) :: real) integrable_on {c..}" + using has_integral_exp_minus_to_infinity[of a c] unfolding integrable_on_def by blast + +lemma has_integral_powr_from_0: + assumes a: "a > (-1::real)" and c: "c \ 0" + shows "((\x. x powr a) has_integral (c powr (a+1) / (a+1))) {0..c}" +proof (cases "c = 0") + case False + 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}" + proof (cases "inverse (of_nat (Suc k)) \ c") + case True + { + fix x assume x: "x \ inverse (1 + real k)" + have "0 < inverse (1 + real k)" by simp + also note x + finally have "x > 0" . + } note x = this + hence "((\x. x powr a) has_integral c powr (a + 1) / (a + 1) - + inverse (real (Suc k)) powr (a + 1) / (a + 1)) {inverse (real (Suc k))..c}" + using True a by (intro fundamental_theorem_of_calculus) + (auto intro!: derivative_eq_intros continuous_on_powr' continuous_on_const + continuous_on_id simp: has_field_derivative_iff_has_vector_derivative [symmetric]) + with True show ?thesis unfolding f_def F_def by (subst has_integral_restrict) simp_all + next + case False + thus ?thesis unfolding f_def F_def by (subst has_integral_restrict) auto + qed + } note has_integral_f = this + have integral_f: "integral {0..c} (f k) = F k" for k + using has_integral_f[of k] by (rule integral_unique) + + have A: "(\x. x powr a) integrable_on {0..c} \ + (\k. integral {0..c} (f k)) \ integral {0..c} (\x. x powr a)" + proof (intro monotone_convergence_increasing ballI allI) + fix k from has_integral_f[of k] show "f k integrable_on {0..c}" + by (auto simp: integrable_on_def) + next + fix k :: nat and x :: real + { + assume x: "inverse (real (Suc k)) \ x" + have "inverse (real (Suc (Suc k))) \ inverse (real (Suc k))" by (simp add: field_simps) + also note x + finally have "inverse (real (Suc (Suc k))) \ x" . + } + thus "f k x \ f (Suc k) x" by (auto simp: f_def simp del: of_nat_Suc) + next + fix x assume x: "x \ {0..c}" + show "(\k. f k x) \ x powr a" + proof (cases "x = 0") + case False + with x have "x > 0" by simp + from order_tendstoD(2)[OF LIMSEQ_inverse_real_of_nat this] + have "eventually (\k. x powr a = f k x) sequentially" + by eventually_elim (insert x, simp add: f_def) + moreover have "(\_. x powr a) \ x powr a" by simp + ultimately show ?thesis by (rule Lim_transform_eventually) + qed (simp_all add: f_def) + next + { + fix k + from a have "F k \ c powr (a + 1) / (a + 1)" + by (auto simp add: F_def divide_simps) + also from a have "F k \ 0" + by (auto simp: F_def divide_simps simp del: of_nat_Suc intro!: powr_mono2) + hence "F k = abs (F k)" by simp + finally have "abs (F k) \ c powr (a + 1) / (a + 1)" . + } + thus "bounded {integral {0..c} (f k) |k. True}" + by (intro bounded_realI[of _ "c powr (a+1) / (a+1)"]) (auto simp: integral_f) + qed + + from False c have "c > 0" by simp + from order_tendstoD(2)[OF LIMSEQ_inverse_real_of_nat this] + have "eventually (\k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a+1) / (a+1) = + integral {0..c} (f k)) sequentially" + by eventually_elim (simp add: integral_f F_def) + moreover have "(\k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1)) + \ c powr (a + 1) / (a + 1) - 0 powr (a + 1) / (a + 1)" + using a by (intro tendsto_intros LIMSEQ_inverse_real_of_nat) auto + hence "(\k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1)) + \ c powr (a + 1) / (a + 1)" by simp + ultimately have "(\k. integral {0..c} (f k)) \ c powr (a+1) / (a+1)" + by (rule Lim_transform_eventually) + with A have "integral {0..c} (\x. x powr a) = c powr (a+1) / (a+1)" + by (blast intro: LIMSEQ_unique) + with A show ?thesis by (simp add: has_integral_integral) +qed (simp_all add: has_integral_refl) + +lemma integrable_on_powr_from_0: + assumes a: "a > (-1::real)" and c: "c \ 0" + shows "(\x. x powr a) integrable_on {0..c}" + using has_integral_powr_from_0[OF assms] unfolding integrable_on_def by blast + +end diff -r bbcb05504fdc -r bd218a9320b5 src/HOL/Multivariate_Analysis/Integral_Test.thy --- a/src/HOL/Multivariate_Analysis/Integral_Test.thy Thu Aug 04 18:45:28 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Integral_Test.thy Thu Aug 04 19:36:31 2016 +0200 @@ -1,18 +1,18 @@ (* Title: HOL/Multivariate_Analysis/Integral_Test.thy Author: Manuel Eberl, TU München *) - + section \Integral Test for Summability\ theory Integral_Test -imports Integration +imports Henstock_Kurzweil_Integration begin text \ - The integral test for summability. We show here that for a decreasing non-negative - function, the infinite sum over that function evaluated at the natural numbers + The integral test for summability. We show here that for a decreasing non-negative + function, the infinite sum over that function evaluated at the natural numbers converges iff the corresponding integral converges. - + As a useful side result, we also provide some results on the difference between the integral and the partial sum. (This is useful e.g. for the definition of the Euler-Mascheroni constant) @@ -33,7 +33,7 @@ proof - note int = integrable_continuous_real[OF continuous_on_subset[OF cont]] let ?int = "\a b. integral {of_nat a..of_nat b} f" - have "-sum_integral_diff_series n = ?int 0 n - (\k\n. f (of_nat k))" + have "-sum_integral_diff_series n = ?int 0 n - (\k\n. f (of_nat k))" by (simp add: sum_integral_diff_series_def) also have "?int 0 n = (\k sum_integral_diff_series n" for n proof - fix n :: nat - have "sum_integral_diff_series (Suc n) - sum_integral_diff_series n = + have "sum_integral_diff_series (Suc n) - sum_integral_diff_series n = f (of_nat (Suc n)) + (?int 0 n - ?int 0 (Suc n))" unfolding sum_integral_diff_series_def by (simp add: algebra_simps) also have "?int 0 n - ?int 0 (Suc n) = -?int n (Suc n)" @@ -77,7 +77,7 @@ lemma sum_integral_diff_series_Bseq: "Bseq sum_integral_diff_series" proof - - from sum_integral_diff_series_nonneg and sum_integral_diff_series_antimono + from sum_integral_diff_series_nonneg and sum_integral_diff_series_antimono have "norm (sum_integral_diff_series n) \ sum_integral_diff_series 0" for n by simp thus "Bseq sum_integral_diff_series" by (rule BseqI') qed @@ -97,12 +97,12 @@ also have "... \ convergent (\n. integral {0..of_nat n} f)" proof assume "convergent (\n. \k\n. f (of_nat k))" - from convergent_diff[OF this sum_integral_diff_series_convergent] - show "convergent (\n. integral {0..of_nat n} f)" + from convergent_diff[OF this sum_integral_diff_series_convergent] + show "convergent (\n. integral {0..of_nat n} f)" unfolding sum_integral_diff_series_def by simp next assume "convergent (\n. integral {0..of_nat n} f)" - from convergent_add[OF this sum_integral_diff_series_convergent] + from convergent_add[OF this sum_integral_diff_series_convergent] show "convergent (\n. \k\n. f (of_nat k))" unfolding sum_integral_diff_series_def by simp qed finally show ?thesis by simp diff -r bbcb05504fdc -r bd218a9320b5 src/HOL/Multivariate_Analysis/Integration.thy --- a/src/HOL/Multivariate_Analysis/Integration.thy Thu Aug 04 18:45:28 2016 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,12645 +0,0 @@ -(* Author: John Harrison - Author: Robert Himmelmann, TU Muenchen (Translation from HOL light); proofs reworked by LCP -*) - -section \Kurzweil-Henstock Gauge Integration in many dimensions.\ - -theory Integration -imports - Derivative - Uniform_Limit - "~~/src/HOL/Library/Indicator_Function" -begin - -lemmas scaleR_simps = scaleR_zero_left scaleR_minus_left scaleR_left_diff_distrib - scaleR_zero_right scaleR_minus_right scaleR_right_diff_distrib scaleR_eq_0_iff - scaleR_cancel_left scaleR_cancel_right scaleR_add_right scaleR_add_left real_vector_class.scaleR_one - - -subsection \Sundries\ - -lemma conjunctD2: assumes "a \ b" shows a b using assms by auto -lemma conjunctD3: assumes "a \ b \ c" shows a b c using assms by auto -lemma conjunctD4: assumes "a \ b \ c \ d" shows a b c d using assms by auto - -declare norm_triangle_ineq4[intro] - -lemma simple_image: "{f x |x . x \ s} = f ` s" - by blast - -lemma linear_simps: - assumes "bounded_linear f" - shows - "f (a + b) = f a + f b" - "f (a - b) = f a - f b" - "f 0 = 0" - "f (- a) = - f a" - "f (s *\<^sub>R v) = s *\<^sub>R (f v)" -proof - - interpret f: bounded_linear f by fact - show "f (a + b) = f a + f b" by (rule f.add) - show "f (a - b) = f a - f b" by (rule f.diff) - show "f 0 = 0" by (rule f.zero) - show "f (- a) = - f a" by (rule f.minus) - show "f (s *\<^sub>R v) = s *\<^sub>R (f v)" by (rule f.scaleR) -qed - -lemma bounded_linearI: - assumes "\x y. f (x + y) = f x + f y" - and "\r x. f (r *\<^sub>R x) = r *\<^sub>R f x" - and "\x. norm (f x) \ norm x * K" - shows "bounded_linear f" - using assms by (rule bounded_linear_intro) (* FIXME: duplicate *) - -lemma bounded_linear_component [intro]: "bounded_linear (\x::'a::euclidean_space. x \ k)" - by (rule bounded_linear_inner_left) - -lemma transitive_stepwise_lt_eq: - assumes "(\x y z::nat. R x y \ R y z \ R x z)" - shows "((\m. \n>m. R m n) \ (\n. R n (Suc n)))" - (is "?l = ?r") -proof safe - assume ?r - fix n m :: nat - assume "m < n" - then show "R m n" - proof (induct n arbitrary: m) - case 0 - then show ?case by auto - next - case (Suc n) - show ?case - proof (cases "m < n") - case True - show ?thesis - apply (rule assms[OF Suc(1)[OF True]]) - using \?r\ - apply auto - done - next - case False - then have "m = n" - using Suc(2) by auto - then show ?thesis - using \?r\ by auto - qed - qed -qed auto - -lemma transitive_stepwise_gt: - assumes "\x y z. R x y \ R y z \ R x z" "\n. R n (Suc n)" - shows "\n>m. R m n" -proof - - have "\m. \n>m. R m n" - apply (subst transitive_stepwise_lt_eq) - apply (blast intro: assms)+ - done - then show ?thesis by auto -qed - -lemma transitive_stepwise_le_eq: - assumes "\x. R x x" "\x y z. R x y \ R y z \ R x z" - shows "(\m. \n\m. R m n) \ (\n. R n (Suc n))" - (is "?l = ?r") -proof safe - assume ?r - fix m n :: nat - assume "m \ n" - then show "R m n" - proof (induct n arbitrary: m) - case 0 - with assms show ?case by auto - next - case (Suc n) - show ?case - proof (cases "m \ n") - case True - with Suc.hyps \\n. R n (Suc n)\ assms show ?thesis - by blast - next - case False - then have "m = Suc n" - using Suc(2) by auto - then show ?thesis - using assms(1) by auto - qed - qed -qed auto - -lemma transitive_stepwise_le: - assumes "\x. R x x" "\x y z. R x y \ R y z \ R x z" - and "\n. R n (Suc n)" - shows "\n\m. R m n" -proof - - have "\m. \n\m. R m n" - apply (subst transitive_stepwise_le_eq) - apply (blast intro: assms)+ - done - then show ?thesis by auto -qed - - -subsection \Some useful lemmas about intervals.\ - -lemma empty_as_interval: "{} = cbox One (0::'a::euclidean_space)" - using nonempty_Basis - by (fastforce simp add: set_eq_iff mem_box) - -lemma interior_subset_union_intervals: - assumes "i = cbox a b" - and "j = cbox c d" - and "interior j \ {}" - and "i \ j \ s" - and "interior i \ interior j = {}" - shows "interior i \ interior s" -proof - - have "box a b \ cbox c d = {}" - using inter_interval_mixed_eq_empty[of c d a b] and assms(3,5) - unfolding assms(1,2) interior_cbox by auto - moreover - have "box a b \ cbox c d \ s" - apply (rule order_trans,rule box_subset_cbox) - using assms(4) unfolding assms(1,2) - apply auto - done - ultimately - show ?thesis - unfolding assms interior_cbox - by auto (metis IntI UnE empty_iff interior_maximal open_box subsetCE subsetI) -qed - -lemma inter_interior_unions_intervals: - fixes f::"('a::euclidean_space) set set" - assumes "finite f" - and "open s" - and "\t\f. \a b. t = cbox a b" - and "\t\f. s \ (interior t) = {}" - shows "s \ interior (\f) = {}" -proof (clarsimp simp only: all_not_in_conv [symmetric]) - fix x - assume x: "x \ s" "x \ interior (\f)" - have lem1: "\x e s U. ball x e \ s \ interior U \ ball x e \ s \ U" - using interior_subset - by auto (meson Topology_Euclidean_Space.open_ball contra_subsetD interior_maximal mem_ball) - have "\t\f. \x. \e>0. ball x e \ s \ t" - if "finite f" and "\t\f. \a b. t = cbox a b" and "\x. x \ s \ interior (\f)" for f - using that - proof (induct rule: finite_induct) - case empty - obtain x where "x \ s \ interior (\{})" - using empty(2) .. - then have False - unfolding Union_empty interior_empty by auto - then show ?case by auto - next - case (insert i f) - obtain x where x: "x \ s \ interior (\insert i f)" - using insert(5) .. - then obtain e where e: "0 < e \ ball x e \ s \ interior (\insert i f)" - unfolding open_contains_ball_eq[OF open_Int[OF assms(2) open_interior], rule_format] .. - obtain a where "\b. i = cbox a b" - using insert(4)[rule_format,OF insertI1] .. - then obtain b where ab: "i = cbox a b" .. - show ?case - proof (cases "x \ i") - case False - then have "x \ UNIV - cbox a b" - unfolding ab by auto - then obtain d where "0 < d \ ball x d \ UNIV - cbox a b" - unfolding open_contains_ball_eq[OF open_Diff[OF open_UNIV closed_cbox],rule_format] .. - then have "0 < d" "ball x (min d e) \ UNIV - i" - unfolding ab ball_min_Int by auto - then have "ball x (min d e) \ s \ interior (\f)" - using e unfolding lem1 unfolding ball_min_Int by auto - then have "x \ s \ interior (\f)" using \d>0\ e by auto - then have "\t\f. \x e. 0 < e \ ball x e \ s \ t" - using insert.hyps(3) insert.prems(1) by blast - then show ?thesis by auto - next - case True show ?thesis - proof (cases "x\box a b") - case True - then obtain d where "0 < d \ ball x d \ box a b" - unfolding open_contains_ball_eq[OF open_box,rule_format] .. - then show ?thesis - apply (rule_tac x=i in bexI, rule_tac x=x in exI, rule_tac x="min d e" in exI) - unfolding ab - using box_subset_cbox[of a b] and e - apply fastforce+ - done - next - case False - then obtain k where "x\k \ a\k \ x\k \ b\k" and k: "k \ Basis" - unfolding mem_box by (auto simp add: not_less) - then have "x\k = a\k \ x\k = b\k" - using True unfolding ab and mem_box - apply (erule_tac x = k in ballE) - apply auto - done - then have "\x. ball x (e/2) \ s \ (\f)" - proof (rule disjE) - let ?z = "x - (e/2) *\<^sub>R k" - assume as: "x\k = a\k" - have "ball ?z (e / 2) \ i = {}" - proof (clarsimp simp only: all_not_in_conv [symmetric]) - fix y - assume "y \ ball ?z (e / 2)" and yi: "y \ i" - then have "dist ?z y < e/2" by auto - then have "\(?z - y) \ k\ < e/2" - using Basis_le_norm[OF k, of "?z - y"] unfolding dist_norm by auto - then have "y\k < a\k" - using e k - by (auto simp add: field_simps abs_less_iff as inner_simps) - then have "y \ i" - unfolding ab mem_box by (auto intro!: bexI[OF _ k]) - then show False using yi by auto - qed - moreover - have "ball ?z (e/2) \ s \ (\insert i f)" - apply (rule order_trans[OF _ e[THEN conjunct2, unfolded lem1]]) - proof - fix y - assume as: "y \ ball ?z (e/2)" - have "norm (x - y) \ \e\ / 2 + norm (x - y - (e / 2) *\<^sub>R k)" - apply (rule order_trans,rule norm_triangle_sub[of "x - y" "(e/2) *\<^sub>R k"]) - unfolding norm_scaleR norm_Basis[OF k] - apply auto - done - also have "\ < \e\ / 2 + \e\ / 2" - apply (rule add_strict_left_mono) - using as e - apply (auto simp add: field_simps dist_norm) - done - finally show "y \ ball x e" - unfolding mem_ball dist_norm using e by (auto simp add:field_simps) - qed - ultimately show ?thesis - apply (rule_tac x="?z" in exI) - unfolding Union_insert - apply auto - done - next - let ?z = "x + (e/2) *\<^sub>R k" - assume as: "x\k = b\k" - have "ball ?z (e / 2) \ i = {}" - proof (clarsimp simp only: all_not_in_conv [symmetric]) - fix y - assume "y \ ball ?z (e / 2)" and yi: "y \ i" - then have "dist ?z y < e/2" - by auto - then have "\(?z - y) \ k\ < e/2" - using Basis_le_norm[OF k, of "?z - y"] - unfolding dist_norm by auto - then have "y\k > b\k" - using e k - by (auto simp add:field_simps inner_simps inner_Basis as) - then have "y \ i" - unfolding ab mem_box by (auto intro!: bexI[OF _ k]) - then show False using yi by auto - qed - moreover - have "ball ?z (e/2) \ s \ (\insert i f)" - apply (rule order_trans[OF _ e[THEN conjunct2, unfolded lem1]]) - proof - fix y - assume as: "y\ ball ?z (e/2)" - have "norm (x - y) \ \e\ / 2 + norm (x - y + (e / 2) *\<^sub>R k)" - apply (rule order_trans,rule norm_triangle_sub[of "x - y" "- (e/2) *\<^sub>R k"]) - unfolding norm_scaleR - apply (auto simp: k) - done - also have "\ < \e\ / 2 + \e\ / 2" - apply (rule add_strict_left_mono) - using as unfolding mem_ball dist_norm - using e apply (auto simp add: field_simps) - done - finally show "y \ ball x e" - unfolding mem_ball dist_norm using e by (auto simp add:field_simps) - qed - ultimately show ?thesis - apply (rule_tac x="?z" in exI) - unfolding Union_insert - apply auto - done - qed - then obtain x where "ball x (e / 2) \ s \ \f" .. - then have "x \ s \ interior (\f)" - unfolding lem1[where U="\f", symmetric] - using centre_in_ball e by auto - then show ?thesis - using insert.hyps(3) insert.prems(1) by blast - qed - qed - qed - from this[OF assms(1,3)] x - obtain t x e where "t \ f" "0 < e" "ball x e \ s \ t" - by blast - then have "x \ s" "x \ interior t" - using open_subset_interior[OF open_ball, of x e t] - by auto - then show False - using \t \ f\ assms(4) by auto -qed - -subsection \Bounds on intervals where they exist.\ - -definition interval_upperbound :: "('a::euclidean_space) set \ 'a" - where "interval_upperbound s = (\i\Basis. (SUP x:s. x\i) *\<^sub>R i)" - -definition interval_lowerbound :: "('a::euclidean_space) set \ 'a" - where "interval_lowerbound s = (\i\Basis. (INF x:s. x\i) *\<^sub>R i)" - -lemma interval_upperbound[simp]: - "\i\Basis. a\i \ b\i \ - interval_upperbound (cbox a b) = (b::'a::euclidean_space)" - unfolding interval_upperbound_def euclidean_representation_setsum cbox_def - by (safe intro!: cSup_eq) auto - -lemma interval_lowerbound[simp]: - "\i\Basis. a\i \ b\i \ - interval_lowerbound (cbox a b) = (a::'a::euclidean_space)" - unfolding interval_lowerbound_def euclidean_representation_setsum cbox_def - by (safe intro!: cInf_eq) auto - -lemmas interval_bounds = interval_upperbound interval_lowerbound - -lemma - fixes X::"real set" - shows interval_upperbound_real[simp]: "interval_upperbound X = Sup X" - and interval_lowerbound_real[simp]: "interval_lowerbound X = Inf X" - by (auto simp: interval_upperbound_def interval_lowerbound_def) - -lemma interval_bounds'[simp]: - assumes "cbox a b \ {}" - shows "interval_upperbound (cbox a b) = b" - and "interval_lowerbound (cbox a b) = a" - using assms unfolding box_ne_empty by auto - - -lemma interval_upperbound_Times: - assumes "A \ {}" and "B \ {}" - shows "interval_upperbound (A \ B) = (interval_upperbound A, interval_upperbound B)" -proof- - from assms have fst_image_times': "A = fst ` (A \ B)" by simp - have "(\i\Basis. (SUP x:A \ B. x \ (i, 0)) *\<^sub>R i) = (\i\Basis. (SUP x:A. x \ i) *\<^sub>R i)" - by (subst (2) fst_image_times') (simp del: fst_image_times add: o_def inner_Pair_0) - moreover from assms have snd_image_times': "B = snd ` (A \ B)" by simp - have "(\i\Basis. (SUP x:A \ B. x \ (0, i)) *\<^sub>R i) = (\i\Basis. (SUP x:B. x \ i) *\<^sub>R i)" - by (subst (2) snd_image_times') (simp del: snd_image_times add: o_def inner_Pair_0) - ultimately show ?thesis unfolding interval_upperbound_def - by (subst setsum_Basis_prod_eq) (auto simp add: setsum_prod) -qed - -lemma interval_lowerbound_Times: - assumes "A \ {}" and "B \ {}" - shows "interval_lowerbound (A \ B) = (interval_lowerbound A, interval_lowerbound B)" -proof- - from assms have fst_image_times': "A = fst ` (A \ B)" by simp - have "(\i\Basis. (INF x:A \ B. x \ (i, 0)) *\<^sub>R i) = (\i\Basis. (INF x:A. x \ i) *\<^sub>R i)" - by (subst (2) fst_image_times') (simp del: fst_image_times add: o_def inner_Pair_0) - moreover from assms have snd_image_times': "B = snd ` (A \ B)" by simp - have "(\i\Basis. (INF x:A \ B. x \ (0, i)) *\<^sub>R i) = (\i\Basis. (INF x:B. x \ i) *\<^sub>R i)" - by (subst (2) snd_image_times') (simp del: snd_image_times add: o_def inner_Pair_0) - ultimately show ?thesis unfolding interval_lowerbound_def - by (subst setsum_Basis_prod_eq) (auto simp add: setsum_prod) -qed - -subsection \Content (length, area, volume...) of an interval.\ - -definition "content (s::('a::euclidean_space) set) = - (if s = {} then 0 else (\i\Basis. (interval_upperbound s)\i - (interval_lowerbound s)\i))" - -lemma interval_not_empty: "\i\Basis. a\i \ b\i \ cbox a b \ {}" - unfolding box_eq_empty unfolding not_ex not_less by auto - -lemma content_cbox: - fixes a :: "'a::euclidean_space" - assumes "\i\Basis. a\i \ b\i" - shows "content (cbox a b) = (\i\Basis. b\i - a\i)" - using interval_not_empty[OF assms] - unfolding content_def - by auto - -lemma content_cbox': - fixes a :: "'a::euclidean_space" - assumes "cbox a b \ {}" - shows "content (cbox a b) = (\i\Basis. b\i - a\i)" - using assms box_ne_empty(1) content_cbox by blast - -lemma content_real: "a \ b \ content {a..b} = b - a" - by (auto simp: interval_upperbound_def interval_lowerbound_def content_def) - -lemma abs_eq_content: "\y - x\ = (if x\y then content {x .. y} else content {y..x})" - by (auto simp: content_real) - -lemma content_singleton[simp]: "content {a} = 0" -proof - - have "content (cbox a a) = 0" - by (subst content_cbox) (auto simp: ex_in_conv) - then show ?thesis by (simp add: cbox_sing) -qed - -lemma content_unit[iff]: "content(cbox 0 (One::'a::euclidean_space)) = 1" - proof - - have *: "\i\Basis. (0::'a)\i \ (One::'a)\i" - by auto - have "0 \ cbox 0 (One::'a)" - unfolding mem_box by auto - then show ?thesis - unfolding content_def interval_bounds[OF *] using setprod.neutral_const by auto - qed - -lemma content_pos_le[intro]: - fixes a::"'a::euclidean_space" - shows "0 \ content (cbox a b)" -proof (cases "cbox a b = {}") - case False - then have *: "\i\Basis. a \ i \ b \ i" - unfolding box_ne_empty . - have "0 \ (\i\Basis. interval_upperbound (cbox a b) \ i - interval_lowerbound (cbox a b) \ i)" - apply (rule setprod_nonneg) - unfolding interval_bounds[OF *] - using * - apply auto - done - also have "\ = content (cbox a b)" using False by (simp add: content_def) - finally show ?thesis . -qed (simp add: content_def) - -corollary content_nonneg [simp]: - fixes a::"'a::euclidean_space" - shows "~ content (cbox a b) < 0" -using not_le by blast - -lemma content_pos_lt: - fixes a :: "'a::euclidean_space" - assumes "\i\Basis. a\i < b\i" - shows "0 < content (cbox a b)" - using assms - by (auto simp: content_def box_eq_empty intro!: setprod_pos) - -lemma content_eq_0: - "content (cbox a b) = 0 \ (\i\Basis. b\i \ a\i)" - by (auto simp: content_def box_eq_empty intro!: setprod_pos bexI) - -lemma cond_cases: "(P \ Q x) \ (\ P \ Q y) \ Q (if P then x else y)" - by auto - -lemma content_cbox_cases: - "content (cbox a (b::'a::euclidean_space)) = - (if \i\Basis. a\i \ b\i then setprod (\i. b\i - a\i) Basis else 0)" - by (auto simp: not_le content_eq_0 intro: less_imp_le content_cbox) - -lemma content_eq_0_interior: "content (cbox a b) = 0 \ interior(cbox a b) = {}" - unfolding content_eq_0 interior_cbox box_eq_empty - by auto - -lemma content_pos_lt_eq: - "0 < content (cbox a (b::'a::euclidean_space)) \ (\i\Basis. a\i < b\i)" -proof (rule iffI) - assume "0 < content (cbox a b)" - then have "content (cbox a b) \ 0" by auto - then show "\i\Basis. a\i < b\i" - unfolding content_eq_0 not_ex not_le by fastforce -next - assume "\i\Basis. a \ i < b \ i" - then show "0 < content (cbox a b)" - by (metis content_pos_lt) -qed - -lemma content_empty [simp]: "content {} = 0" - unfolding content_def by auto - -lemma content_real_if [simp]: "content {a..b} = (if a \ b then b - a else 0)" - by (simp add: content_real) - -lemma content_subset: - assumes "cbox a b \ cbox c d" - shows "content (cbox a b) \ content (cbox c d)" -proof (cases "cbox a b = {}") - case True - then show ?thesis - using content_pos_le[of c d] by auto -next - case False - then have ab_ne: "\i\Basis. a \ i \ b \ i" - unfolding box_ne_empty by auto - then have ab_ab: "a\cbox a b" "b\cbox a b" - unfolding mem_box by auto - have "cbox c d \ {}" using assms False by auto - then have cd_ne: "\i\Basis. c \ i \ d \ i" - using assms unfolding box_ne_empty by auto - have "\i. i \ Basis \ 0 \ b \ i - a \ i" - using ab_ne by auto - moreover - have "\i. i \ Basis \ b \ i - a \ i \ d \ i - c \ i" - using assms[unfolded subset_eq mem_box,rule_format,OF ab_ab(2)] - assms[unfolded subset_eq mem_box,rule_format,OF ab_ab(1)] - by (metis diff_mono) - ultimately show ?thesis - unfolding content_def interval_bounds[OF ab_ne] interval_bounds[OF cd_ne] - by (simp add: setprod_mono if_not_P[OF False] if_not_P[OF \cbox c d \ {}\]) -qed - -lemma content_lt_nz: "0 < content (cbox a b) \ content (cbox a b) \ 0" - unfolding content_pos_lt_eq content_eq_0 unfolding not_ex not_le by fastforce - -lemma content_times[simp]: "content (A \ B) = content A * content B" -proof (cases "A \ B = {}") - let ?ub1 = "interval_upperbound" and ?lb1 = "interval_lowerbound" - let ?ub2 = "interval_upperbound" and ?lb2 = "interval_lowerbound" - assume nonempty: "A \ B \ {}" - hence "content (A \ B) = (\i\Basis. (?ub1 A, ?ub2 B) \ i - (?lb1 A, ?lb2 B) \ i)" - unfolding content_def by (simp add: interval_upperbound_Times interval_lowerbound_Times) - also have "... = content A * content B" unfolding content_def using nonempty - apply (subst Basis_prod_def, subst setprod.union_disjoint, force, force, force, simp) - apply (subst (1 2) setprod.reindex, auto intro: inj_onI) - done - finally show ?thesis . -qed (auto simp: content_def) - -lemma content_Pair: "content (cbox (a,c) (b,d)) = content (cbox a b) * content (cbox c d)" - by (simp add: cbox_Pair_eq) - -lemma content_cbox_pair_eq0_D: - "content (cbox (a,c) (b,d)) = 0 \ content (cbox a b) = 0 \ content (cbox c d) = 0" - by (simp add: content_Pair) - -lemma content_eq_0_gen: - fixes s :: "'a::euclidean_space set" - assumes "bounded s" - shows "content s = 0 \ (\i\Basis. \v. \x \ s. x \ i = v)" (is "_ = ?rhs") -proof safe - assume "content s = 0" then show ?rhs - apply (clarsimp simp: ex_in_conv content_def split: if_split_asm) - apply (rule_tac x=a in bexI) - apply (rule_tac x="interval_lowerbound s \ a" in exI) - apply (clarsimp simp: interval_upperbound_def interval_lowerbound_def) - apply (drule cSUP_eq_cINF_D) - apply (auto simp: bounded_inner_imp_bdd_above [OF assms] bounded_inner_imp_bdd_below [OF assms]) - done -next - fix i a - assume "i \ Basis" "\x\s. x \ i = a" - then show "content s = 0" - apply (clarsimp simp: content_def) - apply (rule_tac x=i in bexI) - apply (auto simp: interval_upperbound_def interval_lowerbound_def) - done -qed - -lemma content_0_subset_gen: - fixes a :: "'a::euclidean_space" - assumes "content t = 0" "s \ t" "bounded t" shows "content s = 0" -proof - - have "bounded s" - using assms by (metis bounded_subset) - then show ?thesis - using assms - by (auto simp: content_eq_0_gen) -qed - -lemma content_0_subset: "\content(cbox a b) = 0; s \ cbox a b\ \ content s = 0" - by (simp add: content_0_subset_gen bounded_cbox) - - -lemma interval_split: - fixes a :: "'a::euclidean_space" - assumes "k \ Basis" - shows - "cbox a b \ {x. x\k \ c} = cbox a (\i\Basis. (if i = k then min (b\k) c else b\i) *\<^sub>R i)" - "cbox a b \ {x. x\k \ c} = cbox (\i\Basis. (if i = k then max (a\k) c else a\i) *\<^sub>R i) b" - apply (rule_tac[!] set_eqI) - unfolding Int_iff mem_box mem_Collect_eq - using assms - apply auto - done - -lemma content_split: - fixes a :: "'a::euclidean_space" - assumes "k \ Basis" - shows "content (cbox a b) = content(cbox a b \ {x. x\k \ c}) + content(cbox a b \ {x. x\k \ c})" -proof cases - note simps = interval_split[OF assms] content_cbox_cases - have *: "Basis = insert k (Basis - {k})" "\x. finite (Basis-{x})" "\x. x\Basis-{x}" - using assms by auto - have *: "\X Y Z. (\i\Basis. Z i (if i = k then X else Y i)) = Z k X * (\i\Basis-{k}. Z i (Y i))" - "(\i\Basis. b\i - a\i) = (\i\Basis-{k}. b\i - a\i) * (b\k - a\k)" - apply (subst *(1)) - defer - apply (subst *(1)) - unfolding setprod.insert[OF *(2-)] - apply auto - done - assume as: "\i\Basis. a \ i \ b \ i" - moreover - have "\x. min (b \ k) c = max (a \ k) c \ - x * (b\k - a\k) = x * (max (a \ k) c - a \ k) + x * (b \ k - max (a \ k) c)" - by (auto simp add: field_simps) - moreover - have **: "(\i\Basis. ((\i\Basis. (if i = k then min (b \ k) c else b \ i) *\<^sub>R i) \ i - a \ i)) = - (\i\Basis. (if i = k then min (b \ k) c else b \ i) - a \ i)" - "(\i\Basis. b \ i - ((\i\Basis. (if i = k then max (a \ k) c else a \ i) *\<^sub>R i) \ i)) = - (\i\Basis. b \ i - (if i = k then max (a \ k) c else a \ i))" - by (auto intro!: setprod.cong) - have "\ a \ k \ c \ \ c \ b \ k \ False" - unfolding not_le - using as[unfolded ,rule_format,of k] assms - by auto - ultimately show ?thesis - using assms - unfolding simps ** - unfolding *(1)[of "\i x. b\i - x"] *(1)[of "\i x. x - a\i"] - unfolding *(2) - by auto -next - assume "\ (\i\Basis. a \ i \ b \ i)" - then have "cbox a b = {}" - unfolding box_eq_empty by (auto simp: not_le) - then show ?thesis - by (auto simp: not_le) -qed - -subsection \The notion of a gauge --- simply an open set containing the point.\ - -definition "gauge d \ (\x. x \ d x \ open (d x))" - -lemma gaugeI: - assumes "\x. x \ g x" - and "\x. open (g x)" - shows "gauge g" - using assms unfolding gauge_def by auto - -lemma gaugeD[dest]: - assumes "gauge d" - shows "x \ d x" - and "open (d x)" - using assms unfolding gauge_def by auto - -lemma gauge_ball_dependent: "\x. 0 < e x \ gauge (\x. ball x (e x))" - unfolding gauge_def by auto - -lemma gauge_ball[intro]: "0 < e \ gauge (\x. ball x e)" - unfolding gauge_def by auto - -lemma gauge_trivial[intro!]: "gauge (\x. ball x 1)" - by (rule gauge_ball) auto - -lemma gauge_inter[intro]: "gauge d1 \ gauge d2 \ gauge (\x. d1 x \ d2 x)" - unfolding gauge_def by auto - -lemma gauge_inters: - assumes "finite s" - and "\d\s. gauge (f d)" - shows "gauge (\x. \{f d x | d. d \ s})" -proof - - have *: "\x. {f d x |d. d \ s} = (\d. f d x) ` s" - by auto - show ?thesis - unfolding gauge_def unfolding * - using assms unfolding Ball_def Inter_iff mem_Collect_eq gauge_def by auto -qed - -lemma gauge_existence_lemma: - "(\x. \d :: real. p x \ 0 < d \ q d x) \ (\x. \d>0. p x \ q d x)" - by (metis zero_less_one) - - -subsection \Divisions.\ - -definition division_of (infixl "division'_of" 40) -where - "s division_of i \ - finite s \ - (\k\s. k \ i \ k \ {} \ (\a b. k = cbox a b)) \ - (\k1\s. \k2\s. k1 \ k2 \ interior(k1) \ interior(k2) = {}) \ - (\s = i)" - -lemma division_ofD[dest]: - assumes "s division_of i" - shows "finite s" - and "\k. k \ s \ k \ i" - and "\k. k \ s \ k \ {}" - and "\k. k \ s \ \a b. k = cbox a b" - and "\k1 k2. k1 \ s \ k2 \ s \ k1 \ k2 \ interior(k1) \ interior(k2) = {}" - and "\s = i" - using assms unfolding division_of_def by auto - -lemma division_ofI: - assumes "finite s" - and "\k. k \ s \ k \ i" - and "\k. k \ s \ k \ {}" - and "\k. k \ s \ \a b. k = cbox a b" - and "\k1 k2. k1 \ s \ k2 \ s \ k1 \ k2 \ interior k1 \ interior k2 = {}" - and "\s = i" - shows "s division_of i" - using assms unfolding division_of_def by auto - -lemma division_of_finite: "s division_of i \ finite s" - unfolding division_of_def by auto - -lemma division_of_self[intro]: "cbox a b \ {} \ {cbox a b} division_of (cbox a b)" - unfolding division_of_def by auto - -lemma division_of_trivial[simp]: "s division_of {} \ s = {}" - unfolding division_of_def by auto - -lemma division_of_sing[simp]: - "s division_of cbox a (a::'a::euclidean_space) \ s = {cbox a a}" - (is "?l = ?r") -proof - assume ?r - moreover - { fix k - assume "s = {{a}}" "k\s" - then have "\x y. k = cbox x y" - apply (rule_tac x=a in exI)+ - apply (force simp: cbox_sing) - done - } - ultimately show ?l - unfolding division_of_def cbox_sing by auto -next - assume ?l - note * = conjunctD4[OF this[unfolded division_of_def cbox_sing]] - { - fix x - assume x: "x \ s" have "x = {a}" - using *(2)[rule_format,OF x] by auto - } - moreover have "s \ {}" - using *(4) by auto - ultimately show ?r - unfolding cbox_sing by auto -qed - -lemma elementary_empty: obtains p where "p division_of {}" - unfolding division_of_trivial by auto - -lemma elementary_interval: obtains p where "p division_of (cbox a b)" - by (metis division_of_trivial division_of_self) - -lemma division_contains: "s division_of i \ \x\i. \k\s. x \ k" - unfolding division_of_def by auto - -lemma forall_in_division: - "d division_of i \ (\x\d. P x) \ (\a b. cbox a b \ d \ P (cbox a b))" - unfolding division_of_def by fastforce - -lemma division_of_subset: - assumes "p division_of (\p)" - and "q \ p" - shows "q division_of (\q)" -proof (rule division_ofI) - note * = division_ofD[OF assms(1)] - show "finite q" - using "*"(1) assms(2) infinite_super by auto - { - fix k - assume "k \ q" - then have kp: "k \ p" - using assms(2) by auto - show "k \ \q" - using \k \ q\ by auto - show "\a b. k = cbox a b" - using *(4)[OF kp] by auto - show "k \ {}" - using *(3)[OF kp] by auto - } - fix k1 k2 - assume "k1 \ q" "k2 \ q" "k1 \ k2" - then have **: "k1 \ p" "k2 \ p" "k1 \ k2" - using assms(2) by auto - show "interior k1 \ interior k2 = {}" - using *(5)[OF **] by auto -qed auto - -lemma division_of_union_self[intro]: "p division_of s \ p division_of (\p)" - unfolding division_of_def by auto - -lemma division_of_content_0: - assumes "content (cbox a b) = 0" "d division_of (cbox a b)" - shows "\k\d. content k = 0" - unfolding forall_in_division[OF assms(2)] - by (metis antisym_conv assms content_pos_le content_subset division_ofD(2)) - -lemma division_inter: - fixes s1 s2 :: "'a::euclidean_space set" - assumes "p1 division_of s1" - and "p2 division_of s2" - shows "{k1 \ k2 | k1 k2 .k1 \ p1 \ k2 \ p2 \ k1 \ k2 \ {}} division_of (s1 \ s2)" - (is "?A' division_of _") -proof - - let ?A = "{s. s \ (\(k1,k2). k1 \ k2) ` (p1 \ p2) \ s \ {}}" - have *: "?A' = ?A" by auto - show ?thesis - unfolding * - proof (rule division_ofI) - have "?A \ (\(x, y). x \ y) ` (p1 \ p2)" - by auto - moreover have "finite (p1 \ p2)" - using assms unfolding division_of_def by auto - ultimately show "finite ?A" by auto - have *: "\s. \{x\s. x \ {}} = \s" - by auto - show "\?A = s1 \ s2" - apply (rule set_eqI) - unfolding * and UN_iff - using division_ofD(6)[OF assms(1)] and division_ofD(6)[OF assms(2)] - apply auto - done - { - fix k - assume "k \ ?A" - then obtain k1 k2 where k: "k = k1 \ k2" "k1 \ p1" "k2 \ p2" "k \ {}" - by auto - then show "k \ {}" - by auto - show "k \ s1 \ s2" - using division_ofD(2)[OF assms(1) k(2)] and division_ofD(2)[OF assms(2) k(3)] - unfolding k by auto - obtain a1 b1 where k1: "k1 = cbox a1 b1" - using division_ofD(4)[OF assms(1) k(2)] by blast - obtain a2 b2 where k2: "k2 = cbox a2 b2" - using division_ofD(4)[OF assms(2) k(3)] by blast - show "\a b. k = cbox a b" - unfolding k k1 k2 unfolding inter_interval by auto - } - fix k1 k2 - assume "k1 \ ?A" - then obtain x1 y1 where k1: "k1 = x1 \ y1" "x1 \ p1" "y1 \ p2" "k1 \ {}" - by auto - assume "k2 \ ?A" - then obtain x2 y2 where k2: "k2 = x2 \ y2" "x2 \ p1" "y2 \ p2" "k2 \ {}" - by auto - assume "k1 \ k2" - then have th: "x1 \ x2 \ y1 \ y2" - unfolding k1 k2 by auto - have *: "interior x1 \ interior x2 = {} \ interior y1 \ interior y2 = {} \ - interior (x1 \ y1) \ interior x1 \ interior (x1 \ y1) \ interior y1 \ - interior (x2 \ y2) \ interior x2 \ interior (x2 \ y2) \ interior y2 \ - interior (x1 \ y1) \ interior (x2 \ y2) = {}" by auto - show "interior k1 \ interior k2 = {}" - unfolding k1 k2 - apply (rule *) - using assms division_ofD(5) k1 k2(2) k2(3) th apply auto - done - qed -qed - -lemma division_inter_1: - assumes "d division_of i" - and "cbox a (b::'a::euclidean_space) \ i" - shows "{cbox a b \ k | k. k \ d \ cbox a b \ k \ {}} division_of (cbox a b)" -proof (cases "cbox a b = {}") - case True - show ?thesis - unfolding True and division_of_trivial by auto -next - case False - have *: "cbox a b \ i = cbox a b" using assms(2) by auto - show ?thesis - using division_inter[OF division_of_self[OF False] assms(1)] - unfolding * by auto -qed - -lemma elementary_inter: - fixes s t :: "'a::euclidean_space set" - assumes "p1 division_of s" - and "p2 division_of t" - shows "\p. p division_of (s \ t)" -using assms division_inter by blast - -lemma elementary_inters: - assumes "finite f" - and "f \ {}" - and "\s\f. \p. p division_of (s::('a::euclidean_space) set)" - shows "\p. p division_of (\f)" - using assms -proof (induct f rule: finite_induct) - case (insert x f) - show ?case - proof (cases "f = {}") - case True - then show ?thesis - unfolding True using insert by auto - next - case False - obtain p where "p division_of \f" - using insert(3)[OF False insert(5)[unfolded ball_simps,THEN conjunct2]] .. - moreover obtain px where "px division_of x" - using insert(5)[rule_format,OF insertI1] .. - ultimately show ?thesis - by (simp add: elementary_inter Inter_insert) - qed -qed auto - -lemma division_disjoint_union: - assumes "p1 division_of s1" - and "p2 division_of s2" - and "interior s1 \ interior s2 = {}" - shows "(p1 \ p2) division_of (s1 \ s2)" -proof (rule division_ofI) - note d1 = division_ofD[OF assms(1)] - note d2 = division_ofD[OF assms(2)] - show "finite (p1 \ p2)" - using d1(1) d2(1) by auto - show "\(p1 \ p2) = s1 \ s2" - using d1(6) d2(6) by auto - { - fix k1 k2 - assume as: "k1 \ p1 \ p2" "k2 \ p1 \ p2" "k1 \ k2" - moreover - let ?g="interior k1 \ interior k2 = {}" - { - assume as: "k1\p1" "k2\p2" - have ?g - using interior_mono[OF d1(2)[OF as(1)]] interior_mono[OF d2(2)[OF as(2)]] - using assms(3) by blast - } - moreover - { - assume as: "k1\p2" "k2\p1" - have ?g - using interior_mono[OF d1(2)[OF as(2)]] interior_mono[OF d2(2)[OF as(1)]] - using assms(3) by blast - } - ultimately show ?g - using d1(5)[OF _ _ as(3)] and d2(5)[OF _ _ as(3)] by auto - } - fix k - assume k: "k \ p1 \ p2" - show "k \ s1 \ s2" - using k d1(2) d2(2) by auto - show "k \ {}" - using k d1(3) d2(3) by auto - show "\a b. k = cbox a b" - using k d1(4) d2(4) by auto -qed - -lemma partial_division_extend_1: - fixes a b c d :: "'a::euclidean_space" - assumes incl: "cbox c d \ cbox a b" - and nonempty: "cbox c d \ {}" - obtains p where "p division_of (cbox a b)" "cbox c d \ p" -proof - let ?B = "\f::'a\'a \ 'a. - cbox (\i\Basis. (fst (f i) \ i) *\<^sub>R i) (\i\Basis. (snd (f i) \ i) *\<^sub>R i)" - define p where "p = ?B ` (Basis \\<^sub>E {(a, c), (c, d), (d, b)})" - - show "cbox c d \ p" - unfolding p_def - by (auto simp add: box_eq_empty cbox_def intro!: image_eqI[where x="\(i::'a)\Basis. (c, d)"]) - { - fix i :: 'a - assume "i \ Basis" - with incl nonempty have "a \ i \ c \ i" "c \ i \ d \ i" "d \ i \ b \ i" - unfolding box_eq_empty subset_box by (auto simp: not_le) - } - note ord = this - - show "p division_of (cbox a b)" - proof (rule division_ofI) - show "finite p" - unfolding p_def by (auto intro!: finite_PiE) - { - fix k - assume "k \ p" - then obtain f where f: "f \ Basis \\<^sub>E {(a, c), (c, d), (d, b)}" and k: "k = ?B f" - by (auto simp: p_def) - then show "\a b. k = cbox a b" - by auto - have "k \ cbox a b \ k \ {}" - proof (simp add: k box_eq_empty subset_box not_less, safe) - fix i :: 'a - assume i: "i \ Basis" - with f have "f i = (a, c) \ f i = (c, d) \ f i = (d, b)" - by (auto simp: PiE_iff) - with i ord[of i] - show "a \ i \ fst (f i) \ i" "snd (f i) \ i \ b \ i" "fst (f i) \ i \ snd (f i) \ i" - by auto - qed - then show "k \ {}" "k \ cbox a b" - by auto - { - fix l - assume "l \ p" - then obtain g where g: "g \ Basis \\<^sub>E {(a, c), (c, d), (d, b)}" and l: "l = ?B g" - by (auto simp: p_def) - assume "l \ k" - have "\i\Basis. f i \ g i" - proof (rule ccontr) - assume "\ ?thesis" - with f g have "f = g" - by (auto simp: PiE_iff extensional_def intro!: ext) - with \l \ k\ show False - by (simp add: l k) - qed - then obtain i where *: "i \ Basis" "f i \ g i" .. - then have "f i = (a, c) \ f i = (c, d) \ f i = (d, b)" - "g i = (a, c) \ g i = (c, d) \ g i = (d, b)" - using f g by (auto simp: PiE_iff) - with * ord[of i] show "interior l \ interior k = {}" - by (auto simp add: l k interior_cbox disjoint_interval intro!: bexI[of _ i]) - } - note \k \ cbox a b\ - } - moreover - { - fix x assume x: "x \ cbox a b" - have "\i\Basis. \l. x \ i \ {fst l \ i .. snd l \ i} \ l \ {(a, c), (c, d), (d, b)}" - proof - fix i :: 'a - assume "i \ Basis" - with x ord[of i] - have "(a \ i \ x \ i \ x \ i \ c \ i) \ (c \ i \ x \ i \ x \ i \ d \ i) \ - (d \ i \ x \ i \ x \ i \ b \ i)" - by (auto simp: cbox_def) - then show "\l. x \ i \ {fst l \ i .. snd l \ i} \ l \ {(a, c), (c, d), (d, b)}" - by auto - qed - then obtain f where - f: "\i\Basis. x \ i \ {fst (f i) \ i..snd (f i) \ i} \ f i \ {(a, c), (c, d), (d, b)}" - unfolding bchoice_iff .. - moreover from f have "restrict f Basis \ Basis \\<^sub>E {(a, c), (c, d), (d, b)}" - by auto - moreover from f have "x \ ?B (restrict f Basis)" - by (auto simp: mem_box) - ultimately have "\k\p. x \ k" - unfolding p_def by blast - } - ultimately show "\p = cbox a b" - by auto - qed -qed - -lemma partial_division_extend_interval: - assumes "p division_of (\p)" "(\p) \ cbox a b" - obtains q where "p \ q" "q division_of cbox a (b::'a::euclidean_space)" -proof (cases "p = {}") - case True - obtain q where "q division_of (cbox a b)" - by (rule elementary_interval) - then show ?thesis - using True that by blast -next - case False - note p = division_ofD[OF assms(1)] - have div_cbox: "\k\p. \q. q division_of cbox a b \ k \ q" - proof - fix k - assume kp: "k \ p" - obtain c d where k: "k = cbox c d" - using p(4)[OF kp] by blast - have *: "cbox c d \ cbox a b" "cbox c d \ {}" - using p(2,3)[OF kp, unfolded k] using assms(2) - by (blast intro: order.trans)+ - obtain q where "q division_of cbox a b" "cbox c d \ q" - by (rule partial_division_extend_1[OF *]) - then show "\q. q division_of cbox a b \ k \ q" - unfolding k by auto - qed - obtain q where q: "\x. x \ p \ q x division_of cbox a b" "\x. x \ p \ x \ q x" - using bchoice[OF div_cbox] by blast - { fix x - assume x: "x \ p" - have "q x division_of \q x" - apply (rule division_ofI) - using division_ofD[OF q(1)[OF x]] - apply auto - done } - then have "\x. x \ p \ \d. d division_of \(q x - {x})" - by (meson Diff_subset division_of_subset) - then have "\d. d division_of \((\i. \(q i - {i})) ` p)" - apply - - apply (rule elementary_inters [OF finite_imageI[OF p(1)]]) - apply (auto simp: False elementary_inters [OF finite_imageI[OF p(1)]]) - done - then obtain d where d: "d division_of \((\i. \(q i - {i})) ` p)" .. - have "d \ p division_of cbox a b" - proof - - have te: "\s f t. s \ {} \ \i\s. f i \ i = t \ t = \(f ` s) \ \s" by auto - have cbox_eq: "cbox a b = \((\i. \(q i - {i})) ` p) \ \p" - proof (rule te[OF False], clarify) - fix i - assume i: "i \ p" - show "\(q i - {i}) \ i = cbox a b" - using division_ofD(6)[OF q(1)[OF i]] using q(2)[OF i] by auto - qed - { fix k - assume k: "k \ p" - have *: "\u t s. t \ s = {} \ u \ s \ u \ t = {}" - by auto - have "interior (\i\p. \(q i - {i})) \ interior k = {}" - proof (rule *[OF inter_interior_unions_intervals]) - note qk=division_ofD[OF q(1)[OF k]] - show "finite (q k - {k})" "open (interior k)" "\t\q k - {k}. \a b. t = cbox a b" - using qk by auto - show "\t\q k - {k}. interior k \ interior t = {}" - using qk(5) using q(2)[OF k] by auto - show "interior (\i\p. \(q i - {i})) \ interior (\(q k - {k}))" - apply (rule interior_mono)+ - using k - apply auto - done - qed } note [simp] = this - show "d \ p division_of (cbox a b)" - unfolding cbox_eq - apply (rule division_disjoint_union[OF d assms(1)]) - apply (rule inter_interior_unions_intervals) - apply (rule p open_interior ballI)+ - apply simp_all - done - qed - then show ?thesis - by (meson Un_upper2 that) -qed - -lemma elementary_bounded[dest]: - fixes s :: "'a::euclidean_space set" - shows "p division_of s \ bounded s" - unfolding division_of_def by (metis bounded_Union bounded_cbox) - -lemma elementary_subset_cbox: - "p division_of s \ \a b. s \ cbox a (b::'a::euclidean_space)" - by (meson elementary_bounded bounded_subset_cbox) - -lemma division_union_intervals_exists: - fixes a b :: "'a::euclidean_space" - assumes "cbox a b \ {}" - obtains p where "(insert (cbox a b) p) division_of (cbox a b \ cbox c d)" -proof (cases "cbox c d = {}") - case True - show ?thesis - apply (rule that[of "{}"]) - unfolding True - using assms - apply auto - done -next - case False - show ?thesis - proof (cases "cbox a b \ cbox c d = {}") - case True - then show ?thesis - by (metis that False assms division_disjoint_union division_of_self insert_is_Un interior_Int interior_empty) - next - case False - obtain u v where uv: "cbox a b \ cbox c d = cbox u v" - unfolding inter_interval by auto - have uv_sub: "cbox u v \ cbox c d" using uv by auto - obtain p where "p division_of cbox c d" "cbox u v \ p" - by (rule partial_division_extend_1[OF uv_sub False[unfolded uv]]) - note p = this division_ofD[OF this(1)] - have "interior (cbox a b \ \(p - {cbox u v})) = interior(cbox u v \ \(p - {cbox u v}))" - apply (rule arg_cong[of _ _ interior]) - using p(8) uv by auto - also have "\ = {}" - unfolding interior_Int - apply (rule inter_interior_unions_intervals) - using p(6) p(7)[OF p(2)] p(3) - apply auto - done - finally have [simp]: "interior (cbox a b) \ interior (\(p - {cbox u v})) = {}" by simp - have cbe: "cbox a b \ cbox c d = cbox a b \ \(p - {cbox u v})" - using p(8) unfolding uv[symmetric] by auto - have "insert (cbox a b) (p - {cbox u v}) division_of cbox a b \ \(p - {cbox u v})" - proof - - have "{cbox a b} division_of cbox a b" - by (simp add: assms division_of_self) - then show "insert (cbox a b) (p - {cbox u v}) division_of cbox a b \ \(p - {cbox u v})" - by (metis (no_types) Diff_subset \interior (cbox a b) \ interior (\(p - {cbox u v})) = {}\ division_disjoint_union division_of_subset insert_is_Un p(1) p(8)) - qed - with that[of "p - {cbox u v}"] show ?thesis by (simp add: cbe) - qed -qed - -lemma division_of_unions: - assumes "finite f" - and "\p. p \ f \ p division_of (\p)" - and "\k1 k2. k1 \ \f \ k2 \ \f \ k1 \ k2 \ interior k1 \ interior k2 = {}" - shows "\f division_of \\f" - using assms - by (auto intro!: division_ofI) - -lemma elementary_union_interval: - fixes a b :: "'a::euclidean_space" - assumes "p division_of \p" - obtains q where "q division_of (cbox a b \ \p)" -proof - - note assm = division_ofD[OF assms] - have lem1: "\f s. \\(f ` s) = \((\x. \(f x)) ` s)" - by auto - have lem2: "\f s. f \ {} \ \{s \ t |t. t \ f} = s \ \f" - by auto - { - presume "p = {} \ thesis" - "cbox a b = {} \ thesis" - "cbox a b \ {} \ interior (cbox a b) = {} \ thesis" - "p \ {} \ interior (cbox a b)\{} \ cbox a b \ {} \ thesis" - then show thesis by auto - next - assume as: "p = {}" - obtain p where "p division_of (cbox a b)" - by (rule elementary_interval) - then show thesis - using as that by auto - next - assume as: "cbox a b = {}" - show thesis - using as assms that by auto - next - assume as: "interior (cbox a b) = {}" "cbox a b \ {}" - show thesis - apply (rule that[of "insert (cbox a b) p"],rule division_ofI) - unfolding finite_insert - apply (rule assm(1)) unfolding Union_insert - using assm(2-4) as - apply - - apply (fast dest: assm(5))+ - done - next - assume as: "p \ {}" "interior (cbox a b) \ {}" "cbox a b \ {}" - have "\k\p. \q. (insert (cbox a b) q) division_of (cbox a b \ k)" - proof - fix k - assume kp: "k \ p" - from assm(4)[OF kp] obtain c d where "k = cbox c d" by blast - then show "\q. (insert (cbox a b) q) division_of (cbox a b \ k)" - by (meson as(3) division_union_intervals_exists) - qed - from bchoice[OF this] obtain q where "\x\p. insert (cbox a b) (q x) division_of (cbox a b) \ x" .. - note q = division_ofD[OF this[rule_format]] - let ?D = "\{insert (cbox a b) (q k) | k. k \ p}" - show thesis - proof (rule that[OF division_ofI]) - have *: "{insert (cbox a b) (q k) |k. k \ p} = (\k. insert (cbox a b) (q k)) ` p" - by auto - show "finite ?D" - using "*" assm(1) q(1) by auto - show "\?D = cbox a b \ \p" - unfolding * lem1 - unfolding lem2[OF as(1), of "cbox a b", symmetric] - using q(6) - by auto - fix k - assume k: "k \ ?D" - then show "k \ cbox a b \ \p" - using q(2) by auto - show "k \ {}" - using q(3) k by auto - show "\a b. k = cbox a b" - using q(4) k by auto - fix k' - assume k': "k' \ ?D" "k \ k'" - obtain x where x: "k \ insert (cbox a b) (q x)" "x\p" - using k by auto - obtain x' where x': "k'\insert (cbox a b) (q x')" "x'\p" - using k' by auto - show "interior k \ interior k' = {}" - proof (cases "x = x'") - case True - show ?thesis - using True k' q(5) x' x by auto - next - case False - { - presume "k = cbox a b \ ?thesis" - and "k' = cbox a b \ ?thesis" - and "k \ cbox a b \ k' \ cbox a b \ ?thesis" - then show ?thesis by linarith - next - assume as': "k = cbox a b" - show ?thesis - using as' k' q(5) x' by blast - next - assume as': "k' = cbox a b" - show ?thesis - using as' k'(2) q(5) x by blast - } - assume as': "k \ cbox a b" "k' \ cbox a b" - obtain c d where k: "k = cbox c d" - using q(4)[OF x(2,1)] by blast - have "interior k \ interior (cbox a b) = {}" - using as' k'(2) q(5) x by blast - then have "interior k \ interior x" - using interior_subset_union_intervals - by (metis as(2) k q(2) x interior_subset_union_intervals) - moreover - obtain c d where c_d: "k' = cbox c d" - using q(4)[OF x'(2,1)] by blast - have "interior k' \ interior (cbox a b) = {}" - using as'(2) q(5) x' by blast - then have "interior k' \ interior x'" - by (metis as(2) c_d interior_subset_union_intervals q(2) x'(1) x'(2)) - ultimately show ?thesis - using assm(5)[OF x(2) x'(2) False] by auto - qed - qed - } -qed - -lemma elementary_unions_intervals: - assumes fin: "finite f" - and "\s. s \ f \ \a b. s = cbox a (b::'a::euclidean_space)" - obtains p where "p division_of (\f)" -proof - - have "\p. p division_of (\f)" - proof (induct_tac f rule:finite_subset_induct) - show "\p. p division_of \{}" using elementary_empty by auto - next - fix x F - assume as: "finite F" "x \ F" "\p. p division_of \F" "x\f" - from this(3) obtain p where p: "p division_of \F" .. - from assms(2)[OF as(4)] obtain a b where x: "x = cbox a b" by blast - have *: "\F = \p" - using division_ofD[OF p] by auto - show "\p. p division_of \insert x F" - using elementary_union_interval[OF p[unfolded *], of a b] - unfolding Union_insert x * by metis - qed (insert assms, auto) - then show ?thesis - using that by auto -qed - -lemma elementary_union: - fixes s t :: "'a::euclidean_space set" - assumes "ps division_of s" "pt division_of t" - obtains p where "p division_of (s \ t)" -proof - - have *: "s \ t = \ps \ \pt" - using assms unfolding division_of_def by auto - show ?thesis - apply (rule elementary_unions_intervals[of "ps \ pt"]) - using assms apply auto - by (simp add: * that) -qed - -lemma partial_division_extend: - fixes t :: "'a::euclidean_space set" - assumes "p division_of s" - and "q division_of t" - and "s \ t" - obtains r where "p \ r" and "r division_of t" -proof - - note divp = division_ofD[OF assms(1)] and divq = division_ofD[OF assms(2)] - obtain a b where ab: "t \ cbox a b" - using elementary_subset_cbox[OF assms(2)] by auto - obtain r1 where "p \ r1" "r1 division_of (cbox a b)" - using assms - by (metis ab dual_order.trans partial_division_extend_interval divp(6)) - note r1 = this division_ofD[OF this(2)] - obtain p' where "p' division_of \(r1 - p)" - apply (rule elementary_unions_intervals[of "r1 - p"]) - using r1(3,6) - apply auto - done - then obtain r2 where r2: "r2 division_of (\(r1 - p)) \ (\q)" - by (metis assms(2) divq(6) elementary_inter) - { - fix x - assume x: "x \ t" "x \ s" - then have "x\\r1" - unfolding r1 using ab by auto - then obtain r where r: "r \ r1" "x \ r" - unfolding Union_iff .. - moreover - have "r \ p" - proof - assume "r \ p" - then have "x \ s" using divp(2) r by auto - then show False using x by auto - qed - ultimately have "x\\(r1 - p)" by auto - } - then have *: "t = \p \ (\(r1 - p) \ \q)" - unfolding divp divq using assms(3) by auto - show ?thesis - apply (rule that[of "p \ r2"]) - unfolding * - defer - apply (rule division_disjoint_union) - unfolding divp(6) - apply(rule assms r2)+ - proof - - have "interior s \ interior (\(r1-p)) = {}" - proof (rule inter_interior_unions_intervals) - show "finite (r1 - p)" and "open (interior s)" and "\t\r1-p. \a b. t = cbox a b" - using r1 by auto - have *: "\s. (\x. x \ s \ False) \ s = {}" - by auto - show "\t\r1-p. interior s \ interior t = {}" - proof - fix m x - assume as: "m \ r1 - p" - have "interior m \ interior (\p) = {}" - proof (rule inter_interior_unions_intervals) - show "finite p" and "open (interior m)" and "\t\p. \a b. t = cbox a b" - using divp by auto - show "\t\p. interior m \ interior t = {}" - by (metis DiffD1 DiffD2 as r1(1) r1(7) set_rev_mp) - qed - then show "interior s \ interior m = {}" - unfolding divp by auto - qed - qed - then show "interior s \ interior (\(r1-p) \ (\q)) = {}" - using interior_subset by auto - qed auto -qed - -lemma division_split_left_inj: - fixes type :: "'a::euclidean_space" - assumes "d division_of i" - and "k1 \ d" - and "k2 \ d" - and "k1 \ k2" - and "k1 \ {x::'a. x\k \ c} = k2 \ {x. x\k \ c}" - and k: "k\Basis" - shows "content(k1 \ {x. x\k \ c}) = 0" -proof - - note d=division_ofD[OF assms(1)] - have *: "\(a::'a) b c. content (cbox a b \ {x. x\k \ c}) = 0 \ - interior(cbox a b \ {x. x\k \ c}) = {}" - unfolding interval_split[OF k] content_eq_0_interior by auto - guess u1 v1 using d(4)[OF assms(2)] by (elim exE) note uv1=this - guess u2 v2 using d(4)[OF assms(3)] by (elim exE) note uv2=this - have **: "\s t u. s \ t = {} \ u \ s \ u \ t \ u = {}" - by auto - show ?thesis - unfolding uv1 uv2 * - apply (rule **[OF d(5)[OF assms(2-4)]]) - apply (simp add: uv1) - using assms(5) uv1 by auto -qed - -lemma division_split_right_inj: - fixes type :: "'a::euclidean_space" - assumes "d division_of i" - and "k1 \ d" - and "k2 \ d" - and "k1 \ k2" - and "k1 \ {x::'a. x\k \ c} = k2 \ {x. x\k \ c}" - and k: "k \ Basis" - shows "content (k1 \ {x. x\k \ c}) = 0" -proof - - note d=division_ofD[OF assms(1)] - have *: "\a b::'a. \c. content(cbox a b \ {x. x\k \ c}) = 0 \ - interior(cbox a b \ {x. x\k \ c}) = {}" - unfolding interval_split[OF k] content_eq_0_interior by auto - guess u1 v1 using d(4)[OF assms(2)] by (elim exE) note uv1=this - guess u2 v2 using d(4)[OF assms(3)] by (elim exE) note uv2=this - have **: "\s t u. s \ t = {} \ u \ s \ u \ t \ u = {}" - by auto - show ?thesis - unfolding uv1 uv2 * - apply (rule **[OF d(5)[OF assms(2-4)]]) - apply (simp add: uv1) - using assms(5) uv1 by auto -qed - - -lemma division_split: - fixes a :: "'a::euclidean_space" - assumes "p division_of (cbox a b)" - and k: "k\Basis" - shows "{l \ {x. x\k \ c} | l. l \ p \ l \ {x. x\k \ c} \ {}} division_of(cbox a b \ {x. x\k \ c})" - (is "?p1 division_of ?I1") - and "{l \ {x. x\k \ c} | l. l \ p \ l \ {x. x\k \ c} \ {}} division_of (cbox a b \ {x. x\k \ c})" - (is "?p2 division_of ?I2") -proof (rule_tac[!] division_ofI) - note p = division_ofD[OF assms(1)] - show "finite ?p1" "finite ?p2" - using p(1) by auto - show "\?p1 = ?I1" "\?p2 = ?I2" - unfolding p(6)[symmetric] by auto - { - fix k - assume "k \ ?p1" - then guess l unfolding mem_Collect_eq by (elim exE conjE) note l=this - guess u v using p(4)[OF l(2)] by (elim exE) note uv=this - show "k \ ?I1" - using l p(2) uv by force - show "k \ {}" - by (simp add: l) - show "\a b. k = cbox a b" - apply (simp add: l uv p(2-3)[OF l(2)]) - apply (subst interval_split[OF k]) - apply (auto intro: order.trans) - done - fix k' - assume "k' \ ?p1" - then guess l' unfolding mem_Collect_eq by (elim exE conjE) note l'=this - assume "k \ k'" - then show "interior k \ interior k' = {}" - unfolding l l' using p(5)[OF l(2) l'(2)] by auto - } - { - fix k - assume "k \ ?p2" - then guess l unfolding mem_Collect_eq by (elim exE conjE) note l=this - guess u v using p(4)[OF l(2)] by (elim exE) note uv=this - show "k \ ?I2" - using l p(2) uv by force - show "k \ {}" - by (simp add: l) - show "\a b. k = cbox a b" - apply (simp add: l uv p(2-3)[OF l(2)]) - apply (subst interval_split[OF k]) - apply (auto intro: order.trans) - done - fix k' - assume "k' \ ?p2" - then guess l' unfolding mem_Collect_eq by (elim exE conjE) note l'=this - assume "k \ k'" - then show "interior k \ interior k' = {}" - unfolding l l' using p(5)[OF l(2) l'(2)] by auto - } -qed - -subsection \Tagged (partial) divisions.\ - -definition tagged_partial_division_of (infixr "tagged'_partial'_division'_of" 40) - where "s tagged_partial_division_of i \ - finite s \ - (\x k. (x, k) \ s \ x \ k \ k \ i \ (\a b. k = cbox a b)) \ - (\x1 k1 x2 k2. (x1, k1) \ s \ (x2, k2) \ s \ (x1, k1) \ (x2, k2) \ - interior k1 \ interior k2 = {})" - -lemma tagged_partial_division_ofD[dest]: - assumes "s tagged_partial_division_of i" - shows "finite s" - and "\x k. (x,k) \ s \ x \ k" - and "\x k. (x,k) \ s \ k \ i" - and "\x k. (x,k) \ s \ \a b. k = cbox a b" - and "\x1 k1 x2 k2. (x1,k1) \ s \ - (x2, k2) \ s \ (x1, k1) \ (x2, k2) \ interior k1 \ interior k2 = {}" - using assms unfolding tagged_partial_division_of_def by blast+ - -definition tagged_division_of (infixr "tagged'_division'_of" 40) - where "s tagged_division_of i \ s tagged_partial_division_of i \ (\{k. \x. (x,k) \ s} = i)" - -lemma tagged_division_of_finite: "s tagged_division_of i \ finite s" - unfolding tagged_division_of_def tagged_partial_division_of_def by auto - -lemma tagged_division_of: - "s tagged_division_of i \ - finite s \ - (\x k. (x, k) \ s \ x \ k \ k \ i \ (\a b. k = cbox a b)) \ - (\x1 k1 x2 k2. (x1, k1) \ s \ (x2, k2) \ s \ (x1, k1) \ (x2, k2) \ - interior k1 \ interior k2 = {}) \ - (\{k. \x. (x,k) \ s} = i)" - unfolding tagged_division_of_def tagged_partial_division_of_def by auto - -lemma tagged_division_ofI: - assumes "finite s" - and "\x k. (x,k) \ s \ x \ k" - and "\x k. (x,k) \ s \ k \ i" - and "\x k. (x,k) \ s \ \a b. k = cbox a b" - and "\x1 k1 x2 k2. (x1,k1) \ s \ (x2, k2) \ s \ (x1, k1) \ (x2, k2) \ - interior k1 \ interior k2 = {}" - and "(\{k. \x. (x,k) \ s} = i)" - shows "s tagged_division_of i" - unfolding tagged_division_of - using assms - apply auto - apply fastforce+ - done - -lemma tagged_division_ofD[dest]: (*FIXME USE A LOCALE*) - assumes "s tagged_division_of i" - shows "finite s" - and "\x k. (x,k) \ s \ x \ k" - and "\x k. (x,k) \ s \ k \ i" - and "\x k. (x,k) \ s \ \a b. k = cbox a b" - and "\x1 k1 x2 k2. (x1, k1) \ s \ (x2, k2) \ s \ (x1, k1) \ (x2, k2) \ - interior k1 \ interior k2 = {}" - and "(\{k. \x. (x,k) \ s} = i)" - using assms unfolding tagged_division_of by blast+ - -lemma division_of_tagged_division: - assumes "s tagged_division_of i" - shows "(snd ` s) division_of i" -proof (rule division_ofI) - note assm = tagged_division_ofD[OF assms] - show "\(snd ` s) = i" "finite (snd ` s)" - using assm by auto - fix k - assume k: "k \ snd ` s" - then obtain xk where xk: "(xk, k) \ s" - by auto - then show "k \ i" "k \ {}" "\a b. k = cbox a b" - using assm by fastforce+ - fix k' - assume k': "k' \ snd ` s" "k \ k'" - from this(1) obtain xk' where xk': "(xk', k') \ s" - by auto - then show "interior k \ interior k' = {}" - using assm(5) k'(2) xk by blast -qed - -lemma partial_division_of_tagged_division: - assumes "s tagged_partial_division_of i" - shows "(snd ` s) division_of \(snd ` s)" -proof (rule division_ofI) - note assm = tagged_partial_division_ofD[OF assms] - show "finite (snd ` s)" "\(snd ` s) = \(snd ` s)" - using assm by auto - fix k - assume k: "k \ snd ` s" - then obtain xk where xk: "(xk, k) \ s" - by auto - then show "k \ {}" "\a b. k = cbox a b" "k \ \(snd ` s)" - using assm by auto - fix k' - assume k': "k' \ snd ` s" "k \ k'" - from this(1) obtain xk' where xk': "(xk', k') \ s" - by auto - then show "interior k \ interior k' = {}" - using assm(5) k'(2) xk by auto -qed - -lemma tagged_partial_division_subset: - assumes "s tagged_partial_division_of i" - and "t \ s" - shows "t tagged_partial_division_of i" - using assms - unfolding tagged_partial_division_of_def - using finite_subset[OF assms(2)] - by blast - -lemma (in comm_monoid_set) over_tagged_division_lemma: - assumes "p tagged_division_of i" - and "\u v. cbox u v \ {} \ content (cbox u v) = 0 \ d (cbox u v) = \<^bold>1" - shows "F (\(x,k). d k) p = F d (snd ` p)" -proof - - have *: "(\(x,k). d k) = d \ snd" - unfolding o_def by (rule ext) auto - note assm = tagged_division_ofD[OF assms(1)] - show ?thesis - unfolding * - proof (rule reindex_nontrivial[symmetric]) - show "finite p" - using assm by auto - fix x y - assume "x\p" "y\p" "x\y" "snd x = snd y" - obtain a b where ab: "snd x = cbox a b" - using assm(4)[of "fst x" "snd x"] \x\p\ by auto - have "(fst x, snd y) \ p" "(fst x, snd y) \ y" - by (metis prod.collapse \x\p\ \snd x = snd y\ \x \ y\)+ - with \x\p\ \y\p\ have "interior (snd x) \ interior (snd y) = {}" - by (intro assm(5)[of "fst x" _ "fst y"]) auto - then have "content (cbox a b) = 0" - unfolding \snd x = snd y\[symmetric] ab content_eq_0_interior by auto - then have "d (cbox a b) = \<^bold>1" - using assm(2)[of "fst x" "snd x"] \x\p\ ab[symmetric] by (intro assms(2)) auto - then show "d (snd x) = \<^bold>1" - unfolding ab by auto - qed -qed - -lemma tag_in_interval: "p tagged_division_of i \ (x, k) \ p \ x \ i" - by auto - -lemma tagged_division_of_empty: "{} tagged_division_of {}" - unfolding tagged_division_of by auto - -lemma tagged_partial_division_of_trivial[simp]: "p tagged_partial_division_of {} \ p = {}" - unfolding tagged_partial_division_of_def by auto - -lemma tagged_division_of_trivial[simp]: "p tagged_division_of {} \ p = {}" - unfolding tagged_division_of by auto - -lemma tagged_division_of_self: "x \ cbox a b \ {(x,cbox a b)} tagged_division_of (cbox a b)" - by (rule tagged_division_ofI) auto - -lemma tagged_division_of_self_real: "x \ {a .. b::real} \ {(x,{a .. b})} tagged_division_of {a .. b}" - unfolding box_real[symmetric] - by (rule tagged_division_of_self) - -lemma tagged_division_union: - assumes "p1 tagged_division_of s1" - and "p2 tagged_division_of s2" - and "interior s1 \ interior s2 = {}" - shows "(p1 \ p2) tagged_division_of (s1 \ s2)" -proof (rule tagged_division_ofI) - note p1 = tagged_division_ofD[OF assms(1)] - note p2 = tagged_division_ofD[OF assms(2)] - show "finite (p1 \ p2)" - using p1(1) p2(1) by auto - show "\{k. \x. (x, k) \ p1 \ p2} = s1 \ s2" - using p1(6) p2(6) by blast - fix x k - assume xk: "(x, k) \ p1 \ p2" - show "x \ k" "\a b. k = cbox a b" - using xk p1(2,4) p2(2,4) by auto - show "k \ s1 \ s2" - using xk p1(3) p2(3) by blast - fix x' k' - assume xk': "(x', k') \ p1 \ p2" "(x, k) \ (x', k')" - have *: "\a b. a \ s1 \ b \ s2 \ interior a \ interior b = {}" - using assms(3) interior_mono by blast - show "interior k \ interior k' = {}" - apply (cases "(x, k) \ p1") - apply (meson "*" UnE assms(1) assms(2) p1(5) tagged_division_ofD(3) xk'(1) xk'(2)) - by (metis "*" UnE assms(1) assms(2) inf_sup_aci(1) p2(5) tagged_division_ofD(3) xk xk'(1) xk'(2)) -qed - -lemma tagged_division_unions: - assumes "finite iset" - and "\i\iset. pfn i tagged_division_of i" - and "\i1\iset. \i2\iset. i1 \ i2 \ interior(i1) \ interior(i2) = {}" - shows "\(pfn ` iset) tagged_division_of (\iset)" -proof (rule tagged_division_ofI) - note assm = tagged_division_ofD[OF assms(2)[rule_format]] - show "finite (\(pfn ` iset))" - apply (rule finite_Union) - using assms - apply auto - done - have "\{k. \x. (x, k) \ \(pfn ` iset)} = \((\i. \{k. \x. (x, k) \ pfn i}) ` iset)" - by blast - also have "\ = \iset" - using assm(6) by auto - finally show "\{k. \x. (x, k) \ \(pfn ` iset)} = \iset" . - fix x k - assume xk: "(x, k) \ \(pfn ` iset)" - then obtain i where i: "i \ iset" "(x, k) \ pfn i" - by auto - show "x \ k" "\a b. k = cbox a b" "k \ \iset" - using assm(2-4)[OF i] using i(1) by auto - fix x' k' - assume xk': "(x', k') \ \(pfn ` iset)" "(x, k) \ (x', k')" - then obtain i' where i': "i' \ iset" "(x', k') \ pfn i'" - by auto - have *: "\a b. i \ i' \ a \ i \ b \ i' \ interior a \ interior b = {}" - using i(1) i'(1) - using assms(3)[rule_format] interior_mono - by blast - show "interior k \ interior k' = {}" - apply (cases "i = i'") - using assm(5) i' i(2) xk'(2) apply blast - using "*" assm(3) i' i by auto -qed - -lemma tagged_partial_division_of_union_self: - assumes "p tagged_partial_division_of s" - shows "p tagged_division_of (\(snd ` p))" - apply (rule tagged_division_ofI) - using tagged_partial_division_ofD[OF assms] - apply auto - done - -lemma tagged_division_of_union_self: - assumes "p tagged_division_of s" - shows "p tagged_division_of (\(snd ` p))" - apply (rule tagged_division_ofI) - using tagged_division_ofD[OF assms] - apply auto - done - -subsection \Functions closed on boxes: morphisms from boxes to monoids\ - -text \This auxiliary structure is used to sum up over the elements of a division. Main theorem is - @{text operative_division}. Instances for the monoid are @{typ "'a option"}, @{typ real}, and - @{typ bool}.\ - -lemma property_empty_interval: "\a b. content (cbox a b) = 0 \ P (cbox a b) \ P {}" - using content_empty unfolding empty_as_interval by auto - -paragraph \Using additivity of lifted function to encode definedness.\ - -definition lift_option :: "('a \ 'b \ 'c) \ 'a option \ 'b option \ 'c option" -where - "lift_option f a' b' = Option.bind a' (\a. Option.bind b' (\b. Some (f a b)))" - -lemma lift_option_simps[simp]: - "lift_option f (Some a) (Some b) = Some (f a b)" - "lift_option f None b' = None" - "lift_option f a' None = None" - by (auto simp: lift_option_def) - -lemma (in comm_monoid) comm_monoid_lift_option: "comm_monoid (lift_option op \<^bold>*) (Some \<^bold>1)" - proof qed (auto simp: lift_option_def ac_simps split: bind_split) - -lemma (in comm_monoid) comm_monoid_set_lift_option: "comm_monoid_set (lift_option op \<^bold>*) (Some \<^bold>1)" - proof qed (auto simp: lift_option_def ac_simps split: bind_split) - -lemma comm_monoid_and: "comm_monoid op \ True" - proof qed auto - -lemma comm_monoid_set_and: "comm_monoid_set op \ True" - proof qed auto - -paragraph \Operative\ - -definition (in comm_monoid) operative :: "('b::euclidean_space set \ 'a) \ bool" - where "operative g \ - (\a b. content (cbox a b) = 0 \ g (cbox a b) = \<^bold>1) \ - (\a b c. \k\Basis. g (cbox a b) = g (cbox a b \ {x. x\k \ c}) \<^bold>* g (cbox a b \ {x. x\k \ c}))" - -lemma (in comm_monoid) operativeD[dest]: - assumes "operative g" - shows "\a b. content (cbox a b) = 0 \ g (cbox a b) = \<^bold>1" - and "\a b c k. k \ Basis \ g (cbox a b) = g (cbox a b \ {x. x\k \ c}) \<^bold>* g (cbox a b \ {x. x\k \ c})" - using assms unfolding operative_def by auto - -lemma (in comm_monoid) operative_empty: "operative g \ g {} = \<^bold>1" - unfolding operative_def by (rule property_empty_interval) auto - -lemma operative_content[intro]: "add.operative content" - by (force simp add: add.operative_def content_split[symmetric]) - -definition "division_points (k::('a::euclidean_space) set) d = - {(j,x). j \ Basis \ (interval_lowerbound k)\j < x \ x < (interval_upperbound k)\j \ - (\i\d. (interval_lowerbound i)\j = x \ (interval_upperbound i)\j = x)}" - -lemma division_points_finite: - fixes i :: "'a::euclidean_space set" - assumes "d division_of i" - shows "finite (division_points i d)" -proof - - note assm = division_ofD[OF assms] - let ?M = "\j. {(j,x)|x. (interval_lowerbound i)\j < x \ x < (interval_upperbound i)\j \ - (\i\d. (interval_lowerbound i)\j = x \ (interval_upperbound i)\j = x)}" - have *: "division_points i d = \(?M ` Basis)" - unfolding division_points_def by auto - show ?thesis - unfolding * using assm by auto -qed - -lemma division_points_subset: - fixes a :: "'a::euclidean_space" - assumes "d division_of (cbox a b)" - and "\i\Basis. a\i < b\i" "a\k < c" "c < b\k" - and k: "k \ Basis" - shows "division_points (cbox a b \ {x. x\k \ c}) {l \ {x. x\k \ c} | l . l \ d \ l \ {x. x\k \ c} \ {}} \ - division_points (cbox a b) d" (is ?t1) - and "division_points (cbox a b \ {x. x\k \ c}) {l \ {x. x\k \ c} | l . l \ d \ ~(l \ {x. x\k \ c} = {})} \ - division_points (cbox a b) d" (is ?t2) -proof - - note assm = division_ofD[OF assms(1)] - have *: "\i\Basis. a\i \ b\i" - "\i\Basis. a\i \ (\i\Basis. (if i = k then min (b \ k) c else b \ i) *\<^sub>R i) \ i" - "\i\Basis. (\i\Basis. (if i = k then max (a \ k) c else a \ i) *\<^sub>R i) \ i \ b\i" - "min (b \ k) c = c" "max (a \ k) c = c" - using assms using less_imp_le by auto - show ?t1 (*FIXME a horrible mess*) - unfolding division_points_def interval_split[OF k, of a b] - unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)] - unfolding * - apply (rule subsetI) - unfolding mem_Collect_eq split_beta - apply (erule bexE conjE)+ - apply (simp add: ) - apply (erule exE conjE)+ - proof - fix i l x - assume as: - "a \ fst x < snd x" "snd x < (if fst x = k then c else b \ fst x)" - "interval_lowerbound i \ fst x = snd x \ interval_upperbound i \ fst x = snd x" - "i = l \ {x. x \ k \ c}" "l \ d" "l \ {x. x \ k \ c} \ {}" - and fstx: "fst x \ Basis" - from assm(4)[OF this(5)] guess u v apply-by(erule exE)+ note l=this - have *: "\i\Basis. u \ i \ (\i\Basis. (if i = k then min (v \ k) c else v \ i) *\<^sub>R i) \ i" - using as(6) unfolding l interval_split[OF k] box_ne_empty as . - have **: "\i\Basis. u\i \ v\i" - using l using as(6) unfolding box_ne_empty[symmetric] by auto - show "\i\d. interval_lowerbound i \ fst x = snd x \ interval_upperbound i \ fst x = snd x" - apply (rule bexI[OF _ \l \ d\]) - using as(1-3,5) fstx - unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] as - apply (auto split: if_split_asm) - done - show "snd x < b \ fst x" - using as(2) \c < b\k\ by (auto split: if_split_asm) - qed - show ?t2 - unfolding division_points_def interval_split[OF k, of a b] - unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)] - unfolding * - unfolding subset_eq - apply rule - unfolding mem_Collect_eq split_beta - apply (erule bexE conjE)+ - apply (simp only: mem_Collect_eq inner_setsum_left_Basis simp_thms) - apply (erule exE conjE)+ - proof - fix i l x - assume as: - "(if fst x = k then c else a \ fst x) < snd x" "snd x < b \ fst x" - "interval_lowerbound i \ fst x = snd x \ interval_upperbound i \ fst x = snd x" - "i = l \ {x. c \ x \ k}" "l \ d" "l \ {x. c \ x \ k} \ {}" - and fstx: "fst x \ Basis" - from assm(4)[OF this(5)] guess u v by (elim exE) note l=this - have *: "\i\Basis. (\i\Basis. (if i = k then max (u \ k) c else u \ i) *\<^sub>R i) \ i \ v \ i" - using as(6) unfolding l interval_split[OF k] box_ne_empty as . - have **: "\i\Basis. u\i \ v\i" - using l using as(6) unfolding box_ne_empty[symmetric] by auto - show "\i\d. interval_lowerbound i \ fst x = snd x \ interval_upperbound i \ fst x = snd x" - apply (rule bexI[OF _ \l \ d\]) - using as(1-3,5) fstx - unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] as - apply (auto split: if_split_asm) - done - show "a \ fst x < snd x" - using as(1) \a\k < c\ by (auto split: if_split_asm) - qed -qed - -lemma division_points_psubset: - fixes a :: "'a::euclidean_space" - assumes "d division_of (cbox a b)" - and "\i\Basis. a\i < b\i" "a\k < c" "c < b\k" - and "l \ d" - and "interval_lowerbound l\k = c \ interval_upperbound l\k = c" - and k: "k \ Basis" - shows "division_points (cbox a b \ {x. x\k \ c}) {l \ {x. x\k \ c} | l. l\d \ l \ {x. x\k \ c} \ {}} \ - division_points (cbox a b) d" (is "?D1 \ ?D") - and "division_points (cbox a b \ {x. x\k \ c}) {l \ {x. x\k \ c} | l. l\d \ l \ {x. x\k \ c} \ {}} \ - division_points (cbox a b) d" (is "?D2 \ ?D") -proof - - have ab: "\i\Basis. a\i \ b\i" - using assms(2) by (auto intro!:less_imp_le) - guess u v using division_ofD(4)[OF assms(1,5)] by (elim exE) note l=this - have uv: "\i\Basis. u\i \ v\i" "\i\Basis. a\i \ u\i \ v\i \ b\i" - using division_ofD(2,2,3)[OF assms(1,5)] unfolding l box_ne_empty - using subset_box(1) - apply auto - apply blast+ - done - have *: "interval_upperbound (cbox a b \ {x. x \ k \ interval_upperbound l \ k}) \ k = interval_upperbound l \ k" - "interval_upperbound (cbox a b \ {x. x \ k \ interval_lowerbound l \ k}) \ k = interval_lowerbound l \ k" - unfolding l interval_split[OF k] interval_bounds[OF uv(1)] - using uv[rule_format, of k] ab k - by auto - have "\x. x \ ?D - ?D1" - using assms(3-) - unfolding division_points_def interval_bounds[OF ab] - apply - - apply (erule disjE) - apply (rule_tac x="(k,(interval_lowerbound l)\k)" in exI, force simp add: *) - apply (rule_tac x="(k,(interval_upperbound l)\k)" in exI, force simp add: *) - done - moreover have "?D1 \ ?D" - by (auto simp add: assms division_points_subset) - ultimately show "?D1 \ ?D" - by blast - have *: "interval_lowerbound (cbox a b \ {x. x \ k \ interval_lowerbound l \ k}) \ k = interval_lowerbound l \ k" - "interval_lowerbound (cbox a b \ {x. x \ k \ interval_upperbound l \ k}) \ k = interval_upperbound l \ k" - unfolding l interval_split[OF k] interval_bounds[OF uv(1)] - using uv[rule_format, of k] ab k - by auto - have "\x. x \ ?D - ?D2" - using assms(3-) - unfolding division_points_def interval_bounds[OF ab] - apply - - apply (erule disjE) - apply (rule_tac x="(k,(interval_lowerbound l)\k)" in exI, force simp add: *) - apply (rule_tac x="(k,(interval_upperbound l)\k)" in exI, force simp add: *) - done - moreover have "?D2 \ ?D" - by (auto simp add: assms division_points_subset) - ultimately show "?D2 \ ?D" - by blast -qed - -lemma (in comm_monoid_set) operative_division: - fixes g :: "'b::euclidean_space set \ 'a" - assumes g: "operative g" and d: "d division_of (cbox a b)" shows "F g d = g (cbox a b)" -proof - - define C where [abs_def]: "C = card (division_points (cbox a b) d)" - then show ?thesis - using d - proof (induction C arbitrary: a b d rule: less_induct) - case (less a b d) - show ?case - proof cases - show "content (cbox a b) = 0 \ F g d = g (cbox a b)" - using division_of_content_0[OF _ less.prems] operativeD(1)[OF g] division_ofD(4)[OF less.prems] - by (fastforce intro!: neutral) - next - assume "content (cbox a b) \ 0" - note ab = this[unfolded content_lt_nz[symmetric] content_pos_lt_eq] - then have ab': "\i\Basis. a\i \ b\i" - by (auto intro!: less_imp_le) - show "F g d = g (cbox a b)" - proof (cases "division_points (cbox a b) d = {}") - case True - { fix u v and j :: 'b - assume j: "j \ Basis" and as: "cbox u v \ d" - then have "cbox u v \ {}" - using less.prems by blast - then have uv: "\i\Basis. u\i \ v\i" "u\j \ v\j" - using j unfolding box_ne_empty by auto - have *: "\p r Q. \ j\Basis \ p \ r \ (\x\d. Q x) \ p \ r \ Q (cbox u v)" - using as j by auto - have "(j, u\j) \ division_points (cbox a b) d" - "(j, v\j) \ division_points (cbox a b) d" using True by auto - note this[unfolded de_Morgan_conj division_points_def mem_Collect_eq split_conv interval_bounds[OF ab'] bex_simps] - note *[OF this(1)] *[OF this(2)] note this[unfolded interval_bounds[OF uv(1)]] - moreover - have "a\j \ u\j" "v\j \ b\j" - using division_ofD(2,2,3)[OF \d division_of cbox a b\ as] - apply (metis j subset_box(1) uv(1)) - by (metis \cbox u v \ cbox a b\ j subset_box(1) uv(1)) - ultimately have "u\j = a\j \ v\j = a\j \ u\j = b\j \ v\j = b\j \ u\j = a\j \ v\j = b\j" - unfolding not_less de_Morgan_disj using ab[rule_format,of j] uv(2) j by force } - then have d': "\i\d. \u v. i = cbox u v \ - (\j\Basis. u\j = a\j \ v\j = a\j \ u\j = b\j \ v\j = b\j \ u\j = a\j \ v\j = b\j)" - unfolding forall_in_division[OF less.prems] by blast - have "(1/2) *\<^sub>R (a+b) \ cbox a b" - unfolding mem_box using ab by(auto intro!: less_imp_le simp: inner_simps) - note this[unfolded division_ofD(6)[OF \d division_of cbox a b\,symmetric] Union_iff] - then guess i .. note i=this - guess u v using d'[rule_format,OF i(1)] by (elim exE conjE) note uv=this - have "cbox a b \ d" - proof - - have "u = a" "v = b" - unfolding euclidean_eq_iff[where 'a='b] - proof safe - fix j :: 'b - assume j: "j \ Basis" - note i(2)[unfolded uv mem_box,rule_format,of j] - then show "u \ j = a \ j" and "v \ j = b \ j" - using uv(2)[rule_format,of j] j by (auto simp: inner_simps) - qed - then have "i = cbox a b" using uv by auto - then show ?thesis using i by auto - qed - then have deq: "d = insert (cbox a b) (d - {cbox a b})" - by auto - have "F g (d - {cbox a b}) = \<^bold>1" - proof (intro neutral ballI) - fix x - assume x: "x \ d - {cbox a b}" - then have "x\d" - by auto note d'[rule_format,OF this] - then guess u v by (elim exE conjE) note uv=this - have "u \ a \ v \ b" - using x[unfolded uv] by auto - then obtain j where "u\j \ a\j \ v\j \ b\j" and j: "j \ Basis" - unfolding euclidean_eq_iff[where 'a='b] by auto - then have "u\j = v\j" - using uv(2)[rule_format,OF j] by auto - then have "content (cbox u v) = 0" - unfolding content_eq_0 using j - by force - then show "g x = \<^bold>1" - unfolding uv(1) by (rule operativeD(1)[OF g]) - qed - then show "F g d = g (cbox a b)" - using division_ofD[OF less.prems] - apply (subst deq) - apply (subst insert) - apply auto - done - next - case False - then have "\x. x \ division_points (cbox a b) d" - by auto - then guess k c - unfolding split_paired_Ex division_points_def mem_Collect_eq split_conv - apply (elim exE conjE) - done - note this(2-4,1) note kc=this[unfolded interval_bounds[OF ab']] - from this(3) guess j .. note j=this - define d1 where "d1 = {l \ {x. x\k \ c} | l. l \ d \ l \ {x. x\k \ c} \ {}}" - define d2 where "d2 = {l \ {x. x\k \ c} | l. l \ d \ l \ {x. x\k \ c} \ {}}" - define cb where "cb = (\i\Basis. (if i = k then c else b\i) *\<^sub>R i)" - define ca where "ca = (\i\Basis. (if i = k then c else a\i) *\<^sub>R i)" - note division_points_psubset[OF \d division_of cbox a b\ ab kc(1-2) j] - note psubset_card_mono[OF _ this(1)] psubset_card_mono[OF _ this(2)] - then have *: "F g d1 = g (cbox a b \ {x. x\k \ c})" "F g d2 = g (cbox a b \ {x. x\k \ c})" - unfolding interval_split[OF kc(4)] - apply (rule_tac[!] "less.hyps"[rule_format]) - using division_split[OF \d division_of cbox a b\, where k=k and c=c] - apply (simp_all add: interval_split kc d1_def d2_def division_points_finite[OF \d division_of cbox a b\]) - done - { fix l y - assume as: "l \ d" "y \ d" "l \ {x. x \ k \ c} = y \ {x. x \ k \ c}" "l \ y" - from division_ofD(4)[OF \d division_of cbox a b\ this(1)] guess u v by (elim exE) note leq=this - have "g (l \ {x. x \ k \ c}) = \<^bold>1" - unfolding leq interval_split[OF kc(4)] - apply (rule operativeD[OF g]) - unfolding interval_split[symmetric, OF kc(4)] - using division_split_left_inj less as kc leq by blast - } note fxk_le = this - { fix l y - assume as: "l \ d" "y \ d" "l \ {x. c \ x \ k} = y \ {x. c \ x \ k}" "l \ y" - from division_ofD(4)[OF \d division_of cbox a b\ this(1)] guess u v by (elim exE) note leq=this - have "g (l \ {x. x \ k \ c}) = \<^bold>1" - unfolding leq interval_split[OF kc(4)] - apply (rule operativeD(1)[OF g]) - unfolding interval_split[symmetric,OF kc(4)] - using division_split_right_inj less leq as kc by blast - } note fxk_ge = this - have d1_alt: "d1 = (\l. l \ {x. x\k \ c}) ` {l \ d. l \ {x. x\k \ c} \ {}}" - using d1_def by auto - have d2_alt: "d2 = (\l. l \ {x. x\k \ c}) ` {l \ d. l \ {x. x\k \ c} \ {}}" - using d2_def by auto - have "g (cbox a b) = F g d1 \<^bold>* F g d2" (is "_ = ?prev") - unfolding * using g kc(4) by blast - also have "F g d1 = F (\l. g (l \ {x. x\k \ c})) d" - unfolding d1_alt using division_of_finite[OF less.prems] fxk_le - by (subst reindex_nontrivial) (auto intro!: mono_neutral_cong_left simp: operative_empty[OF g]) - also have "F g d2 = F (\l. g (l \ {x. x\k \ c})) d" - unfolding d2_alt using division_of_finite[OF less.prems] fxk_ge - by (subst reindex_nontrivial) (auto intro!: mono_neutral_cong_left simp: operative_empty[OF g]) - also have *: "\x\d. g x = g (x \ {x. x \ k \ c}) \<^bold>* g (x \ {x. c \ x \ k})" - unfolding forall_in_division[OF \d division_of cbox a b\] - using g kc(4) by blast - have "F (\l. g (l \ {x. x\k \ c})) d \<^bold>* F (\l. g (l \ {x. x\k \ c})) d = F g d" - using * by (simp add: distrib) - finally show ?thesis by auto - qed - qed - qed -qed - -lemma (in comm_monoid_set) operative_tagged_division: - assumes f: "operative g" and d: "d tagged_division_of (cbox a b)" - shows "F (\(x, l). g l) d = g (cbox a b)" - unfolding d[THEN division_of_tagged_division, THEN operative_division[OF f], symmetric] - by (simp add: f[THEN operativeD(1)] over_tagged_division_lemma[OF d]) - -lemma additive_content_division: "d division_of (cbox a b) \ setsum content d = content (cbox a b)" - by (metis operative_content setsum.operative_division) - -lemma additive_content_tagged_division: - "d tagged_division_of (cbox a b) \ setsum (\(x,l). content l) d = content (cbox a b)" - unfolding setsum.operative_tagged_division[OF operative_content, symmetric] by blast - -lemma - shows real_inner_1_left: "inner 1 x = x" - and real_inner_1_right: "inner x 1 = x" - by simp_all - -lemma content_real_eq_0: "content {a .. b::real} = 0 \ a \ b" - by (metis atLeastatMost_empty_iff2 content_empty content_real diff_self eq_iff le_cases le_iff_diff_le_0) - -lemma interval_real_split: - "{a .. b::real} \ {x. x \ c} = {a .. min b c}" - "{a .. b} \ {x. c \ x} = {max a c .. b}" - apply (metis Int_atLeastAtMostL1 atMost_def) - apply (metis Int_atLeastAtMostL2 atLeast_def) - done - -lemma (in comm_monoid) operative_1_lt: - "operative (g :: real set \ 'a) \ - ((\a b. b \ a \ g {a .. b} = \<^bold>1) \ (\a b c. a < c \ c < b \ g {a .. c} \<^bold>* g {c .. b} = g {a .. b}))" - apply (simp add: operative_def content_real_eq_0 atMost_def[symmetric] atLeast_def[symmetric] - del: content_real_if) -proof safe - fix a b c :: real - assume *: "\a b c. g {a..b} = g {a..min b c} \<^bold>* g {max a c..b}" - assume "a < c" "c < b" - with *[rule_format, of a b c] show "g {a..c} \<^bold>* g {c..b} = g {a..b}" - by (simp add: less_imp_le min.absorb2 max.absorb2) -next - fix a b c :: real - assume as: "\a b. b \ a \ g {a..b} = \<^bold>1" - "\a b c. a < c \ c < b \ g {a..c} \<^bold>* g {c..b} = g {a..b}" - from as(1)[rule_format, of 0 1] as(1)[rule_format, of a a for a] as(2) - have [simp]: "g {} = \<^bold>1" "\a. g {a} = \<^bold>1" - "\a b c. a < c \ c < b \ g {a..c} \<^bold>* g {c..b} = g {a..b}" - by auto - show "g {a..b} = g {a..min b c} \<^bold>* g {max a c..b}" - by (auto simp: min_def max_def le_less) -qed - -lemma (in comm_monoid) operative_1_le: - "operative (g :: real set \ 'a) \ - ((\a b. b \ a \ g {a..b} = \<^bold>1) \ (\a b c. a \ c \ c \ b \ g {a .. c} \<^bold>* g {c .. b} = g {a .. b}))" - unfolding operative_1_lt -proof safe - fix a b c :: real - assume as: "\a b c. a \ c \ c \ b \ g {a..c} \<^bold>* g {c..b} = g {a..b}" "a < c" "c < b" - show "g {a..c} \<^bold>* g {c..b} = g {a..b}" - apply (rule as(1)[rule_format]) - using as(2-) - apply auto - done -next - fix a b c :: real - assume "\a b. b \ a \ g {a .. b} = \<^bold>1" - and "\a b c. a < c \ c < b \ g {a..c} \<^bold>* g {c..b} = g {a..b}" - and "a \ c" - and "c \ b" - note as = this[rule_format] - show "g {a..c} \<^bold>* g {c..b} = g {a..b}" - proof (cases "c = a \ c = b") - case False - then show ?thesis - apply - - apply (subst as(2)) - using as(3-) - apply auto - done - next - case True - then show ?thesis - proof - assume *: "c = a" - then have "g {a .. c} = \<^bold>1" - apply - - apply (rule as(1)[rule_format]) - apply auto - done - then show ?thesis - unfolding * by auto - next - assume *: "c = b" - then have "g {c .. b} = \<^bold>1" - apply - - apply (rule as(1)[rule_format]) - apply auto - done - then show ?thesis - unfolding * by auto - qed - qed -qed - -subsection \Fine-ness of a partition w.r.t. a gauge.\ - -definition fine (infixr "fine" 46) - where "d fine s \ (\(x,k) \ s. k \ d x)" - -lemma fineI: - assumes "\x k. (x, k) \ s \ k \ d x" - shows "d fine s" - using assms unfolding fine_def by auto - -lemma fineD[dest]: - assumes "d fine s" - shows "\x k. (x,k) \ s \ k \ d x" - using assms unfolding fine_def by auto - -lemma fine_inter: "(\x. d1 x \ d2 x) fine p \ d1 fine p \ d2 fine p" - unfolding fine_def by auto - -lemma fine_inters: - "(\x. \{f d x | d. d \ s}) fine p \ (\d\s. (f d) fine p)" - unfolding fine_def by blast - -lemma fine_union: "d fine p1 \ d fine p2 \ d fine (p1 \ p2)" - unfolding fine_def by blast - -lemma fine_unions: "(\p. p \ ps \ d fine p) \ d fine (\ps)" - unfolding fine_def by auto - -lemma fine_subset: "p \ q \ d fine q \ d fine p" - unfolding fine_def by blast - - -subsection \Gauge integral. Define on compact intervals first, then use a limit.\ - -definition has_integral_compact_interval (infixr "has'_integral'_compact'_interval" 46) - where "(f has_integral_compact_interval y) i \ - (\e>0. \d. gauge d \ - (\p. p tagged_division_of i \ d fine p \ - norm (setsum (\(x,k). content k *\<^sub>R f x) p - y) < e))" - -definition has_integral :: - "('n::euclidean_space \ 'b::real_normed_vector) \ 'b \ 'n set \ bool" - (infixr "has'_integral" 46) - where "(f has_integral y) i \ - (if \a b. i = cbox a b - then (f has_integral_compact_interval y) i - else (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ - (\z. ((\x. if x \ i then f x else 0) has_integral_compact_interval z) (cbox a b) \ - norm (z - y) < e)))" - -lemma has_integral: - "(f has_integral y) (cbox a b) \ - (\e>0. \d. gauge d \ - (\p. p tagged_division_of (cbox a b) \ d fine p \ - norm (setsum (\(x,k). content(k) *\<^sub>R f x) p - y) < e))" - unfolding has_integral_def has_integral_compact_interval_def - by auto - -lemma has_integral_real: - "(f has_integral y) {a .. b::real} \ - (\e>0. \d. gauge d \ - (\p. p tagged_division_of {a .. b} \ d fine p \ - norm (setsum (\(x,k). content(k) *\<^sub>R f x) p - y) < e))" - unfolding box_real[symmetric] - by (rule has_integral) - -lemma has_integralD[dest]: - assumes "(f has_integral y) (cbox a b)" - and "e > 0" - obtains d where "gauge d" - and "\p. p tagged_division_of (cbox a b) \ d fine p \ - norm (setsum (\(x,k). content(k) *\<^sub>R f(x)) p - y) < e" - using assms unfolding has_integral by auto - -lemma has_integral_alt: - "(f has_integral y) i \ - (if \a b. i = cbox a b - then (f has_integral y) i - else (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ - (\z. ((\x. if x \ i then f(x) else 0) has_integral z) (cbox a b) \ norm (z - y) < e)))" - unfolding has_integral - unfolding has_integral_compact_interval_def has_integral_def - by auto - -lemma has_integral_altD: - assumes "(f has_integral y) i" - and "\ (\a b. i = cbox a b)" - and "e>0" - obtains B where "B > 0" - and "\a b. ball 0 B \ cbox a b \ - (\z. ((\x. if x \ i then f(x) else 0) has_integral z) (cbox a b) \ norm(z - y) < e)" - using assms - unfolding has_integral - unfolding has_integral_compact_interval_def has_integral_def - by auto - -definition integrable_on (infixr "integrable'_on" 46) - where "f integrable_on i \ (\y. (f has_integral y) i)" - -definition "integral i f = (SOME y. (f has_integral y) i \ ~ f integrable_on i \ y=0)" - -lemma integrable_integral[dest]: "f integrable_on i \ (f has_integral (integral i f)) i" - unfolding integrable_on_def integral_def by (metis (mono_tags, lifting) someI_ex) - -lemma not_integrable_integral: "~ f integrable_on i \ integral i f = 0" - unfolding integrable_on_def integral_def by blast - -lemma has_integral_integrable[intro]: "(f has_integral i) s \ f integrable_on s" - unfolding integrable_on_def by auto - -lemma has_integral_integral: "f integrable_on s \ (f has_integral (integral s f)) s" - by auto - -lemma setsum_content_null: - assumes "content (cbox a b) = 0" - and "p tagged_division_of (cbox a b)" - shows "setsum (\(x,k). content k *\<^sub>R f x) p = (0::'a::real_normed_vector)" -proof (rule setsum.neutral, rule) - fix y - assume y: "y \ p" - obtain x k where xk: "y = (x, k)" - using surj_pair[of y] by blast - note assm = tagged_division_ofD(3-4)[OF assms(2) y[unfolded xk]] - from this(2) obtain c d where k: "k = cbox c d" by blast - have "(\(x, k). content k *\<^sub>R f x) y = content k *\<^sub>R f x" - unfolding xk by auto - also have "\ = 0" - using content_subset[OF assm(1)[unfolded k]] content_pos_le[of c d] - unfolding assms(1) k - by auto - finally show "(\(x, k). content k *\<^sub>R f x) y = 0" . -qed - - -subsection \Some basic combining lemmas.\ - -lemma tagged_division_unions_exists: - assumes "finite iset" - and "\i\iset. \p. p tagged_division_of i \ d fine p" - and "\i1\iset. \i2\iset. i1 \ i2 \ interior i1 \ interior i2 = {}" - and "\iset = i" - obtains p where "p tagged_division_of i" and "d fine p" -proof - - obtain pfn where pfn: - "\x. x \ iset \ pfn x tagged_division_of x" - "\x. x \ iset \ d fine pfn x" - using bchoice[OF assms(2)] by auto - show thesis - apply (rule_tac p="\(pfn ` iset)" in that) - using assms(1) assms(3) assms(4) pfn(1) tagged_division_unions apply force - by (metis (mono_tags, lifting) fine_unions imageE pfn(2)) -qed - - -subsection \The set we're concerned with must be closed.\ - -lemma division_of_closed: - fixes i :: "'n::euclidean_space set" - shows "s division_of i \ closed i" - unfolding division_of_def by fastforce - -subsection \General bisection principle for intervals; might be useful elsewhere.\ - -lemma interval_bisection_step: - fixes type :: "'a::euclidean_space" - assumes "P {}" - and "\s t. P s \ P t \ interior(s) \ interior(t) = {} \ P (s \ t)" - and "\ P (cbox a (b::'a))" - obtains c d where "\ P (cbox c d)" - and "\i\Basis. a\i \ c\i \ c\i \ d\i \ d\i \ b\i \ 2 * (d\i - c\i) \ b\i - a\i" -proof - - have "cbox a b \ {}" - using assms(1,3) by metis - then have ab: "\i. i\Basis \ a \ i \ b \ i" - by (force simp: mem_box) - { fix f - have "\finite f; - \s. s\f \ P s; - \s. s\f \ \a b. s = cbox a b; - \s t. s\f \ t\f \ s \ t \ interior s \ interior t = {}\ \ P (\f)" - proof (induct f rule: finite_induct) - case empty - show ?case - using assms(1) by auto - next - case (insert x f) - show ?case - unfolding Union_insert - apply (rule assms(2)[rule_format]) - using inter_interior_unions_intervals [of f "interior x"] - apply (auto simp: insert) - by (metis IntI empty_iff insert.hyps(2) insert.prems(3) insert_iff) - qed - } note UN_cases = this - let ?A = "{cbox c d | c d::'a. \i\Basis. (c\i = a\i) \ (d\i = (a\i + b\i) / 2) \ - (c\i = (a\i + b\i) / 2) \ (d\i = b\i)}" - let ?PP = "\c d. \i\Basis. a\i \ c\i \ c\i \ d\i \ d\i \ b\i \ 2 * (d\i - c\i) \ b\i - a\i" - { - presume "\c d. ?PP c d \ P (cbox c d) \ False" - then show thesis - unfolding atomize_not not_all - by (blast intro: that) - } - assume as: "\c d. ?PP c d \ P (cbox c d)" - have "P (\?A)" - proof (rule UN_cases) - let ?B = "(\s. cbox (\i\Basis. (if i \ s then a\i else (a\i + b\i) / 2) *\<^sub>R i::'a) - (\i\Basis. (if i \ s then (a\i + b\i) / 2 else b\i) *\<^sub>R i)) ` {s. s \ Basis}" - have "?A \ ?B" - proof - fix x - assume "x \ ?A" - then obtain c d - where x: "x = cbox c d" - "\i. i \ Basis \ - c \ i = a \ i \ d \ i = (a \ i + b \ i) / 2 \ - c \ i = (a \ i + b \ i) / 2 \ d \ i = b \ i" by blast - show "x \ ?B" - unfolding image_iff x - apply (rule_tac x="{i. i\Basis \ c\i = a\i}" in bexI) - apply (rule arg_cong2 [where f = cbox]) - using x(2) ab - apply (auto simp add: euclidean_eq_iff[where 'a='a]) - by fastforce - qed - then show "finite ?A" - by (rule finite_subset) auto - next - fix s - assume "s \ ?A" - then obtain c d - where s: "s = cbox c d" - "\i. i \ Basis \ - c \ i = a \ i \ d \ i = (a \ i + b \ i) / 2 \ - c \ i = (a \ i + b \ i) / 2 \ d \ i = b \ i" - by blast - show "P s" - unfolding s - apply (rule as[rule_format]) - using ab s(2) by force - show "\a b. s = cbox a b" - unfolding s by auto - fix t - assume "t \ ?A" - then obtain e f where t: - "t = cbox e f" - "\i. i \ Basis \ - e \ i = a \ i \ f \ i = (a \ i + b \ i) / 2 \ - e \ i = (a \ i + b \ i) / 2 \ f \ i = b \ i" - by blast - assume "s \ t" - then have "\ (c = e \ d = f)" - unfolding s t by auto - then obtain i where "c\i \ e\i \ d\i \ f\i" and i': "i \ Basis" - unfolding euclidean_eq_iff[where 'a='a] by auto - then have i: "c\i \ e\i" "d\i \ f\i" - using s(2) t(2) apply fastforce - using t(2)[OF i'] \c \ i \ e \ i \ d \ i \ f \ i\ i' s(2) t(2) by fastforce - have *: "\s t. (\a. a \ s \ a \ t \ False) \ s \ t = {}" - by auto - show "interior s \ interior t = {}" - unfolding s t interior_cbox - proof (rule *) - fix x - assume "x \ box c d" "x \ box e f" - then have x: "c\i < d\i" "e\i < f\i" "c\i < f\i" "e\i < d\i" - unfolding mem_box using i' - by force+ - show False using s(2)[OF i'] - proof safe - assume as: "c \ i = a \ i" "d \ i = (a \ i + b \ i) / 2" - show False - using t(2)[OF i'] and i x unfolding as by (fastforce simp add:field_simps) - next - assume as: "c \ i = (a \ i + b \ i) / 2" "d \ i = b \ i" - show False - using t(2)[OF i'] and i x unfolding as by(fastforce simp add:field_simps) - qed - qed - qed - also have "\?A = cbox a b" - proof (rule set_eqI,rule) - fix x - assume "x \ \?A" - then obtain c d where x: - "x \ cbox c d" - "\i. i \ Basis \ - c \ i = a \ i \ d \ i = (a \ i + b \ i) / 2 \ - c \ i = (a \ i + b \ i) / 2 \ d \ i = b \ i" - by blast - show "x\cbox a b" - unfolding mem_box - proof safe - fix i :: 'a - assume i: "i \ Basis" - then show "a \ i \ x \ i" "x \ i \ b \ i" - using x(2)[OF i] x(1)[unfolded mem_box,THEN bspec, OF i] by auto - qed - next - fix x - assume x: "x \ cbox a b" - have "\i\Basis. - \c d. (c = a\i \ d = (a\i + b\i) / 2 \ c = (a\i + b\i) / 2 \ d = b\i) \ c\x\i \ x\i \ d" - (is "\i\Basis. \c d. ?P i c d") - unfolding mem_box - proof - fix i :: 'a - assume i: "i \ Basis" - have "?P i (a\i) ((a \ i + b \ i) / 2) \ ?P i ((a \ i + b \ i) / 2) (b\i)" - using x[unfolded mem_box,THEN bspec, OF i] by auto - then show "\c d. ?P i c d" - by blast - qed - then show "x\\?A" - unfolding Union_iff Bex_def mem_Collect_eq choice_Basis_iff - apply auto - apply (rule_tac x="cbox xa xaa" in exI) - unfolding mem_box - apply auto - done - qed - finally show False - using assms by auto -qed - -lemma interval_bisection: - fixes type :: "'a::euclidean_space" - assumes "P {}" - and "(\s t. P s \ P t \ interior(s) \ interior(t) = {} \ P(s \ t))" - and "\ P (cbox a (b::'a))" - obtains x where "x \ cbox a b" - and "\e>0. \c d. x \ cbox c d \ cbox c d \ ball x e \ cbox c d \ cbox a b \ \ P (cbox c d)" -proof - - have "\x. \y. \ P (cbox (fst x) (snd x)) \ (\ P (cbox (fst y) (snd y)) \ - (\i\Basis. fst x\i \ fst y\i \ fst y\i \ snd y\i \ snd y\i \ snd x\i \ - 2 * (snd y\i - fst y\i) \ snd x\i - fst x\i))" (is "\x. ?P x") - proof - show "?P x" for x - proof (cases "P (cbox (fst x) (snd x))") - case True - then show ?thesis by auto - next - case as: False - obtain c d where "\ P (cbox c d)" - "\i\Basis. - fst x \ i \ c \ i \ - c \ i \ d \ i \ - d \ i \ snd x \ i \ - 2 * (d \ i - c \ i) \ snd x \ i - fst x \ i" - by (rule interval_bisection_step[of P, OF assms(1-2) as]) - then show ?thesis - apply - - apply (rule_tac x="(c,d)" in exI) - apply auto - done - qed - qed - then obtain f where f: - "\x. - \ P (cbox (fst x) (snd x)) \ - \ P (cbox (fst (f x)) (snd (f x))) \ - (\i\Basis. - fst x \ i \ fst (f x) \ i \ - fst (f x) \ i \ snd (f x) \ i \ - snd (f x) \ i \ snd x \ i \ - 2 * (snd (f x) \ i - fst (f x) \ i) \ snd x \ i - fst x \ i)" - apply - - apply (drule choice) - apply blast - done - define AB A B where ab_def: "AB n = (f ^^ n) (a,b)" "A n = fst(AB n)" "B n = snd(AB n)" for n - have "A 0 = a" "B 0 = b" "\n. \ P (cbox (A(Suc n)) (B(Suc n))) \ - (\i\Basis. A(n)\i \ A(Suc n)\i \ A(Suc n)\i \ B(Suc n)\i \ B(Suc n)\i \ B(n)\i \ - 2 * (B(Suc n)\i - A(Suc n)\i) \ B(n)\i - A(n)\i)" (is "\n. ?P n") - proof - - show "A 0 = a" "B 0 = b" - unfolding ab_def by auto - note S = ab_def funpow.simps o_def id_apply - show "?P n" for n - proof (induct n) - case 0 - then show ?case - unfolding S - apply (rule f[rule_format]) using assms(3) - apply auto - done - next - case (Suc n) - show ?case - unfolding S - apply (rule f[rule_format]) - using Suc - unfolding S - apply auto - done - qed - qed - note AB = this(1-2) conjunctD2[OF this(3),rule_format] - - have interv: "\n. \x\cbox (A n) (B n). \y\cbox (A n) (B n). dist x y < e" - if e: "0 < e" for e - proof - - obtain n where n: "(\i\Basis. b \ i - a \ i) / e < 2 ^ n" - using real_arch_pow[of 2 "(setsum (\i. b\i - a\i) Basis) / e"] by auto - show ?thesis - proof (rule exI [where x=n], clarify) - fix x y - assume xy: "x\cbox (A n) (B n)" "y\cbox (A n) (B n)" - have "dist x y \ setsum (\i. \(x - y)\i\) Basis" - unfolding dist_norm by(rule norm_le_l1) - also have "\ \ setsum (\i. B n\i - A n\i) Basis" - proof (rule setsum_mono) - fix i :: 'a - assume i: "i \ Basis" - show "\(x - y) \ i\ \ B n \ i - A n \ i" - using xy[unfolded mem_box,THEN bspec, OF i] - by (auto simp: inner_diff_left) - qed - also have "\ \ setsum (\i. b\i - a\i) Basis / 2^n" - unfolding setsum_divide_distrib - proof (rule setsum_mono) - show "B n \ i - A n \ i \ (b \ i - a \ i) / 2 ^ n" if i: "i \ Basis" for i - proof (induct n) - case 0 - then show ?case - unfolding AB by auto - next - case (Suc n) - have "B (Suc n) \ i - A (Suc n) \ i \ (B n \ i - A n \ i) / 2" - using AB(4)[of i n] using i by auto - also have "\ \ (b \ i - a \ i) / 2 ^ Suc n" - using Suc by (auto simp add: field_simps) - finally show ?case . - qed - qed - also have "\ < e" - using n using e by (auto simp add: field_simps) - finally show "dist x y < e" . - qed - qed - { - fix n m :: nat - assume "m \ n" then have "cbox (A n) (B n) \ cbox (A m) (B m)" - proof (induction rule: inc_induct) - case (step i) - show ?case - using AB(4) by (intro order_trans[OF step.IH] subset_box_imp) auto - qed simp - } note ABsubset = this - have "\a. \n. a\ cbox (A n) (B n)" - by (rule decreasing_closed_nest[rule_format,OF closed_cbox _ ABsubset interv]) - (metis nat.exhaust AB(1-3) assms(1,3)) - then obtain x0 where x0: "\n. x0 \ cbox (A n) (B n)" - by blast - show thesis - proof (rule that[rule_format, of x0]) - show "x0\cbox a b" - using x0[of 0] unfolding AB . - fix e :: real - assume "e > 0" - from interv[OF this] obtain n - where n: "\x\cbox (A n) (B n). \y\cbox (A n) (B n). dist x y < e" .. - have "\ P (cbox (A n) (B n))" - apply (cases "0 < n") - using AB(3)[of "n - 1"] assms(3) AB(1-2) - apply auto - done - moreover have "cbox (A n) (B n) \ ball x0 e" - using n using x0[of n] by auto - moreover have "cbox (A n) (B n) \ cbox a b" - unfolding AB(1-2)[symmetric] by (rule ABsubset) auto - ultimately show "\c d. x0 \ cbox c d \ cbox c d \ ball x0 e \ cbox c d \ cbox a b \ \ P (cbox c d)" - apply (rule_tac x="A n" in exI) - apply (rule_tac x="B n" in exI) - apply (auto simp: x0) - done - qed -qed - - -subsection \Cousin's lemma.\ - -lemma fine_division_exists: - fixes a b :: "'a::euclidean_space" - assumes "gauge g" - obtains p where "p tagged_division_of (cbox a b)" "g fine p" -proof - - presume "\ (\p. p tagged_division_of (cbox a b) \ g fine p) \ False" - then obtain p where "p tagged_division_of (cbox a b)" "g fine p" - by blast - then show thesis .. -next - assume as: "\ (\p. p tagged_division_of (cbox a b) \ g fine p)" - obtain x where x: - "x \ (cbox a b)" - "\e. 0 < e \ - \c d. - x \ cbox c d \ - cbox c d \ ball x e \ - cbox c d \ (cbox a b) \ - \ (\p. p tagged_division_of cbox c d \ g fine p)" - apply (rule interval_bisection[of "\s. \p. p tagged_division_of s \ g fine p", OF _ _ as]) - apply (simp add: fine_def) - apply (metis tagged_division_union fine_union) - apply (auto simp: ) - done - obtain e where e: "e > 0" "ball x e \ g x" - using gaugeD[OF assms, of x] unfolding open_contains_ball by auto - from x(2)[OF e(1)] - obtain c d where c_d: "x \ cbox c d" - "cbox c d \ ball x e" - "cbox c d \ cbox a b" - "\ (\p. p tagged_division_of cbox c d \ g fine p)" - by blast - have "g fine {(x, cbox c d)}" - unfolding fine_def using e using c_d(2) by auto - then show False - using tagged_division_of_self[OF c_d(1)] using c_d by auto -qed - -lemma fine_division_exists_real: - fixes a b :: real - assumes "gauge g" - obtains p where "p tagged_division_of {a .. b}" "g fine p" - by (metis assms box_real(2) fine_division_exists) - -subsection \Basic theorems about integrals.\ - -lemma has_integral_unique: - fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" - assumes "(f has_integral k1) i" - and "(f has_integral k2) i" - shows "k1 = k2" -proof (rule ccontr) - let ?e = "norm (k1 - k2) / 2" - assume as: "k1 \ k2" - then have e: "?e > 0" - by auto - have lem: False - if f_k1: "(f has_integral k1) (cbox a b)" - and f_k2: "(f has_integral k2) (cbox a b)" - and "k1 \ k2" - for f :: "'n \ 'a" and a b k1 k2 - proof - - let ?e = "norm (k1 - k2) / 2" - from \k1 \ k2\ have e: "?e > 0" by auto - obtain d1 where d1: - "gauge d1" - "\p. p tagged_division_of cbox a b \ - d1 fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x) - k1) < norm (k1 - k2) / 2" - by (rule has_integralD[OF f_k1 e]) blast - obtain d2 where d2: - "gauge d2" - "\p. p tagged_division_of cbox a b \ - d2 fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x) - k2) < norm (k1 - k2) / 2" - by (rule has_integralD[OF f_k2 e]) blast - obtain p where p: - "p tagged_division_of cbox a b" - "(\x. d1 x \ d2 x) fine p" - by (rule fine_division_exists[OF gauge_inter[OF d1(1) d2(1)]]) - let ?c = "(\(x, k)\p. content k *\<^sub>R f x)" - have "norm (k1 - k2) \ norm (?c - k2) + norm (?c - k1)" - using norm_triangle_ineq4[of "k1 - ?c" "k2 - ?c"] - by (auto simp add:algebra_simps norm_minus_commute) - also have "\ < norm (k1 - k2) / 2 + norm (k1 - k2) / 2" - apply (rule add_strict_mono) - apply (rule_tac[!] d2(2) d1(2)) - using p unfolding fine_def - apply auto - done - finally show False by auto - qed - { - presume "\ (\a b. i = cbox a b) \ False" - then show False - using as assms lem by blast - } - assume as: "\ (\a b. i = cbox a b)" - obtain B1 where B1: - "0 < B1" - "\a b. ball 0 B1 \ cbox a b \ - \z. ((\x. if x \ i then f x else 0) has_integral z) (cbox a b) \ - norm (z - k1) < norm (k1 - k2) / 2" - by (rule has_integral_altD[OF assms(1) as,OF e]) blast - obtain B2 where B2: - "0 < B2" - "\a b. ball 0 B2 \ cbox a b \ - \z. ((\x. if x \ i then f x else 0) has_integral z) (cbox a b) \ - norm (z - k2) < norm (k1 - k2) / 2" - by (rule has_integral_altD[OF assms(2) as,OF e]) blast - have "\a b::'n. ball 0 B1 \ ball 0 B2 \ cbox a b" - apply (rule bounded_subset_cbox) - using bounded_Un bounded_ball - apply auto - done - then obtain a b :: 'n where ab: "ball 0 B1 \ cbox a b" "ball 0 B2 \ cbox a b" - by blast - obtain w where w: - "((\x. if x \ i then f x else 0) has_integral w) (cbox a b)" - "norm (w - k1) < norm (k1 - k2) / 2" - using B1(2)[OF ab(1)] by blast - obtain z where z: - "((\x. if x \ i then f x else 0) has_integral z) (cbox a b)" - "norm (z - k2) < norm (k1 - k2) / 2" - using B2(2)[OF ab(2)] by blast - have "z = w" - using lem[OF w(1) z(1)] by auto - then have "norm (k1 - k2) \ norm (z - k2) + norm (w - k1)" - using norm_triangle_ineq4 [of "k1 - w" "k2 - z"] - by (auto simp add: norm_minus_commute) - also have "\ < norm (k1 - k2) / 2 + norm (k1 - k2) / 2" - apply (rule add_strict_mono) - apply (rule_tac[!] z(2) w(2)) - done - finally show False by auto -qed - -lemma integral_unique [intro]: "(f has_integral y) k \ integral k f = y" - unfolding integral_def - by (rule some_equality) (auto intro: has_integral_unique) - -lemma eq_integralD: "integral k f = y \ (f has_integral y) k \ ~ f integrable_on k \ y=0" - unfolding integral_def integrable_on_def - apply (erule subst) - apply (rule someI_ex) - by blast - -lemma has_integral_is_0: - fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" - assumes "\x\s. f x = 0" - shows "(f has_integral 0) s" -proof - - have lem: "\a b. \f::'n \ 'a. - (\x\cbox a b. f(x) = 0) \ (f has_integral 0) (cbox a b)" - unfolding has_integral - proof clarify - fix a b e - fix f :: "'n \ 'a" - assume as: "\x\cbox a b. f x = 0" "0 < (e::real)" - have "norm ((\(x, k)\p. content k *\<^sub>R f x) - 0) < e" - if p: "p tagged_division_of cbox a b" for p - proof - - have "(\(x, k)\p. content k *\<^sub>R f x) = 0" - proof (rule setsum.neutral, rule) - fix x - assume x: "x \ p" - have "f (fst x) = 0" - using tagged_division_ofD(2-3)[OF p, of "fst x" "snd x"] using as x by auto - then show "(\(x, k). content k *\<^sub>R f x) x = 0" - apply (subst surjective_pairing[of x]) - unfolding split_conv - apply auto - done - qed - then show ?thesis - using as by auto - qed - then show "\d. gauge d \ - (\p. p tagged_division_of (cbox a b) \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x) - 0) < e)" - by auto - qed - { - presume "\ (\a b. s = cbox a b) \ ?thesis" - with assms lem show ?thesis - by blast - } - have *: "(\x. if x \ s then f x else 0) = (\x. 0)" - apply (rule ext) - using assms - apply auto - done - assume "\ (\a b. s = cbox a b)" - then show ?thesis - using lem - by (subst has_integral_alt) (force simp add: *) -qed - -lemma has_integral_0[simp]: "((\x::'n::euclidean_space. 0) has_integral 0) s" - by (rule has_integral_is_0) auto - -lemma has_integral_0_eq[simp]: "((\x. 0) has_integral i) s \ i = 0" - using has_integral_unique[OF has_integral_0] by auto - -lemma has_integral_linear: - fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" - assumes "(f has_integral y) s" - and "bounded_linear h" - shows "((h \ f) has_integral ((h y))) s" -proof - - interpret bounded_linear h - using assms(2) . - from pos_bounded obtain B where B: "0 < B" "\x. norm (h x) \ norm x * B" - by blast - have lem: "\(f :: 'n \ 'a) y a b. - (f has_integral y) (cbox a b) \ ((h \ f) has_integral h y) (cbox a b)" - unfolding has_integral - proof (clarify, goal_cases) - case prems: (1 f y a b e) - from pos_bounded - obtain B where B: "0 < B" "\x. norm (h x) \ norm x * B" - by blast - have "e / B > 0" using prems(2) B by simp - then obtain g - where g: "gauge g" - "\p. p tagged_division_of (cbox a b) \ g fine p \ - norm ((\(x, k)\p. content k *\<^sub>R f x) - y) < e / B" - using prems(1) by auto - { - fix p - assume as: "p tagged_division_of (cbox a b)" "g fine p" - have hc: "\x k. h ((\(x, k). content k *\<^sub>R f x) x) = (\(x, k). h (content k *\<^sub>R f x)) x" - by auto - then have "(\(x, k)\p. content k *\<^sub>R (h \ f) x) = setsum (h \ (\(x, k). content k *\<^sub>R f x)) p" - unfolding o_def unfolding scaleR[symmetric] hc by simp - also have "\ = h (\(x, k)\p. content k *\<^sub>R f x)" - using setsum[of "\(x,k). content k *\<^sub>R f x" p] using as by auto - finally have "(\(x, k)\p. content k *\<^sub>R (h \ f) x) = h (\(x, k)\p. content k *\<^sub>R f x)" . - then have "norm ((\(x, k)\p. content k *\<^sub>R (h \ f) x) - h y) < e" - apply (simp add: diff[symmetric]) - apply (rule le_less_trans[OF B(2)]) - using g(2)[OF as] B(1) - apply (auto simp add: field_simps) - done - } - with g show ?case - by (rule_tac x=g in exI) auto - qed - { - presume "\ (\a b. s = cbox a b) \ ?thesis" - then show ?thesis - using assms(1) lem by blast - } - assume as: "\ (\a b. s = cbox a b)" - then show ?thesis - proof (subst has_integral_alt, clarsimp) - fix e :: real - assume e: "e > 0" - have *: "0 < e/B" using e B(1) by simp - obtain M where M: - "M > 0" - "\a b. ball 0 M \ cbox a b \ - \z. ((\x. if x \ s then f x else 0) has_integral z) (cbox a b) \ norm (z - y) < e / B" - using has_integral_altD[OF assms(1) as *] by blast - show "\B>0. \a b. ball 0 B \ cbox a b \ - (\z. ((\x. if x \ s then (h \ f) x else 0) has_integral z) (cbox a b) \ norm (z - h y) < e)" - proof (rule_tac x=M in exI, clarsimp simp add: M, goal_cases) - case prems: (1 a b) - obtain z where z: - "((\x. if x \ s then f x else 0) has_integral z) (cbox a b)" - "norm (z - y) < e / B" - using M(2)[OF prems(1)] by blast - have *: "(\x. if x \ s then (h \ f) x else 0) = h \ (\x. if x \ s then f x else 0)" - using zero by auto - show ?case - apply (rule_tac x="h z" in exI) - apply (simp add: * lem z(1)) - apply (metis B diff le_less_trans pos_less_divide_eq z(2)) - done - qed - qed -qed - -lemma has_integral_scaleR_left: - "(f has_integral y) s \ ((\x. f x *\<^sub>R c) has_integral (y *\<^sub>R c)) s" - using has_integral_linear[OF _ bounded_linear_scaleR_left] by (simp add: comp_def) - -lemma has_integral_mult_left: - fixes c :: "_ :: real_normed_algebra" - shows "(f has_integral y) s \ ((\x. f x * c) has_integral (y * c)) s" - using has_integral_linear[OF _ bounded_linear_mult_left] by (simp add: comp_def) - -text\The case analysis eliminates the condition @{term "f integrable_on s"} at the cost - of the type class constraint \division_ring\\ -corollary integral_mult_left [simp]: - fixes c:: "'a::{real_normed_algebra,division_ring}" - shows "integral s (\x. f x * c) = integral s f * c" -proof (cases "f integrable_on s \ c = 0") - case True then show ?thesis - by (force intro: has_integral_mult_left) -next - case False then have "~ (\x. f x * c) integrable_on s" - using has_integral_mult_left [of "(\x. f x * c)" _ s "inverse c"] - by (force simp add: mult.assoc) - with False show ?thesis by (simp add: not_integrable_integral) -qed - -corollary integral_mult_right [simp]: - fixes c:: "'a::{real_normed_field}" - shows "integral s (\x. c * f x) = c * integral s f" -by (simp add: mult.commute [of c]) - -corollary integral_divide [simp]: - fixes z :: "'a::real_normed_field" - shows "integral S (\x. f x / z) = integral S (\x. f x) / z" -using integral_mult_left [of S f "inverse z"] - by (simp add: divide_inverse_commute) - -lemma has_integral_mult_right: - fixes c :: "'a :: real_normed_algebra" - shows "(f has_integral y) i \ ((\x. c * f x) has_integral (c * y)) i" - using has_integral_linear[OF _ bounded_linear_mult_right] by (simp add: comp_def) - -lemma has_integral_cmul: "(f has_integral k) s \ ((\x. c *\<^sub>R f x) has_integral (c *\<^sub>R k)) s" - unfolding o_def[symmetric] - by (metis has_integral_linear bounded_linear_scaleR_right) - -lemma has_integral_cmult_real: - fixes c :: real - assumes "c \ 0 \ (f has_integral x) A" - shows "((\x. c * f x) has_integral c * x) A" -proof (cases "c = 0") - case True - then show ?thesis by simp -next - case False - from has_integral_cmul[OF assms[OF this], of c] show ?thesis - unfolding real_scaleR_def . -qed - -lemma has_integral_neg: "(f has_integral k) s \ ((\x. -(f x)) has_integral -k) s" - by (drule_tac c="-1" in has_integral_cmul) auto - -lemma has_integral_add: - fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" - assumes "(f has_integral k) s" - and "(g has_integral l) s" - shows "((\x. f x + g x) has_integral (k + l)) s" -proof - - have lem: "((\x. f x + g x) has_integral (k + l)) (cbox a b)" - if f_k: "(f has_integral k) (cbox a b)" - and g_l: "(g has_integral l) (cbox a b)" - for f :: "'n \ 'a" and g a b k l - unfolding has_integral - proof clarify - fix e :: real - assume e: "e > 0" - then have *: "e / 2 > 0" - by auto - obtain d1 where d1: - "gauge d1" - "\p. p tagged_division_of (cbox a b) \ d1 fine p \ - norm ((\(x, k)\p. content k *\<^sub>R f x) - k) < e / 2" - using has_integralD[OF f_k *] by blast - obtain d2 where d2: - "gauge d2" - "\p. p tagged_division_of (cbox a b) \ d2 fine p \ - norm ((\(x, k)\p. content k *\<^sub>R g x) - l) < e / 2" - using has_integralD[OF g_l *] by blast - show "\d. gauge d \ (\p. p tagged_division_of (cbox a b) \ d fine p \ - norm ((\(x, k)\p. content k *\<^sub>R (f x + g x)) - (k + l)) < e)" - proof (rule exI [where x="\x. (d1 x) \ (d2 x)"], clarsimp simp add: gauge_inter[OF d1(1) d2(1)]) - fix p - assume as: "p tagged_division_of (cbox a b)" "(\x. d1 x \ d2 x) fine p" - have *: "(\(x, k)\p. content k *\<^sub>R (f x + g x)) = - (\(x, k)\p. content k *\<^sub>R f x) + (\(x, k)\p. content k *\<^sub>R g x)" - unfolding scaleR_right_distrib setsum.distrib[of "\(x,k). content k *\<^sub>R f x" "\(x,k). content k *\<^sub>R g x" p,symmetric] - by (rule setsum.cong) auto - from as have fine: "d1 fine p" "d2 fine p" - unfolding fine_inter by auto - have "norm ((\(x, k)\p. content k *\<^sub>R (f x + g x)) - (k + l)) = - norm (((\(x, k)\p. content k *\<^sub>R f x) - k) + ((\(x, k)\p. content k *\<^sub>R g x) - l))" - unfolding * by (auto simp add: algebra_simps) - also have "\ < e/2 + e/2" - apply (rule le_less_trans[OF norm_triangle_ineq]) - using as d1 d2 fine - apply (blast intro: add_strict_mono) - done - finally show "norm ((\(x, k)\p. content k *\<^sub>R (f x + g x)) - (k + l)) < e" - by auto - qed - qed - { - presume "\ (\a b. s = cbox a b) \ ?thesis" - then show ?thesis - using assms lem by force - } - assume as: "\ (\a b. s = cbox a b)" - then show ?thesis - proof (subst has_integral_alt, clarsimp, goal_cases) - case (1 e) - then have *: "e / 2 > 0" - by auto - from has_integral_altD[OF assms(1) as *] - obtain B1 where B1: - "0 < B1" - "\a b. ball 0 B1 \ cbox a b \ - \z. ((\x. if x \ s then f x else 0) has_integral z) (cbox a b) \ norm (z - k) < e / 2" - by blast - from has_integral_altD[OF assms(2) as *] - obtain B2 where B2: - "0 < B2" - "\a b. ball 0 B2 \ (cbox a b) \ - \z. ((\x. if x \ s then g x else 0) has_integral z) (cbox a b) \ norm (z - l) < e / 2" - by blast - show ?case - proof (rule_tac x="max B1 B2" in exI, clarsimp simp add: max.strict_coboundedI1 B1) - fix a b - assume "ball 0 (max B1 B2) \ cbox a (b::'n)" - then have *: "ball 0 B1 \ cbox a (b::'n)" "ball 0 B2 \ cbox a (b::'n)" - by auto - obtain w where w: - "((\x. if x \ s then f x else 0) has_integral w) (cbox a b)" - "norm (w - k) < e / 2" - using B1(2)[OF *(1)] by blast - obtain z where z: - "((\x. if x \ s then g x else 0) has_integral z) (cbox a b)" - "norm (z - l) < e / 2" - using B2(2)[OF *(2)] by blast - have *: "\x. (if x \ s then f x + g x else 0) = - (if x \ s then f x else 0) + (if x \ s then g x else 0)" - by auto - show "\z. ((\x. if x \ s then f x + g x else 0) has_integral z) (cbox a b) \ norm (z - (k + l)) < e" - apply (rule_tac x="w + z" in exI) - apply (simp add: lem[OF w(1) z(1), unfolded *[symmetric]]) - using norm_triangle_ineq[of "w - k" "z - l"] w(2) z(2) - apply (auto simp add: field_simps) - done - qed - qed -qed - -lemma has_integral_sub: - "(f has_integral k) s \ (g has_integral l) s \ - ((\x. f x - g x) has_integral (k - l)) s" - using has_integral_add[OF _ has_integral_neg, of f k s g l] - by (auto simp: algebra_simps) - -lemma integral_0 [simp]: - "integral s (\x::'n::euclidean_space. 0::'m::real_normed_vector) = 0" - by (rule integral_unique has_integral_0)+ - -lemma integral_add: "f integrable_on s \ g integrable_on s \ - integral s (\x. f x + g x) = integral s f + integral s g" - by (rule integral_unique) (metis integrable_integral has_integral_add) - -lemma integral_cmul [simp]: "integral s (\x. c *\<^sub>R f x) = c *\<^sub>R integral s f" -proof (cases "f integrable_on s \ c = 0") - case True with has_integral_cmul show ?thesis by force -next - case False then have "~ (\x. c *\<^sub>R f x) integrable_on s" - using has_integral_cmul [of "(\x. c *\<^sub>R f x)" _ s "inverse c"] - by force - with False show ?thesis by (simp add: not_integrable_integral) -qed - -lemma integral_neg [simp]: "integral s (\x. - f x) = - integral s f" -proof (cases "f integrable_on s") - case True then show ?thesis - by (simp add: has_integral_neg integrable_integral integral_unique) -next - case False then have "~ (\x. - f x) integrable_on s" - using has_integral_neg [of "(\x. - f x)" _ s ] - by force - with False show ?thesis by (simp add: not_integrable_integral) -qed - -lemma integral_diff: "f integrable_on s \ g integrable_on s \ - integral s (\x. f x - g x) = integral s f - integral s g" - by (rule integral_unique) (metis integrable_integral has_integral_sub) - -lemma integrable_0: "(\x. 0) integrable_on s" - unfolding integrable_on_def using has_integral_0 by auto - -lemma integrable_add: "f integrable_on s \ g integrable_on s \ (\x. f x + g x) integrable_on s" - unfolding integrable_on_def by(auto intro: has_integral_add) - -lemma integrable_cmul: "f integrable_on s \ (\x. c *\<^sub>R f(x)) integrable_on s" - unfolding integrable_on_def by(auto intro: has_integral_cmul) - -lemma integrable_on_cmult_iff: - fixes c :: real - assumes "c \ 0" - shows "(\x. c * f x) integrable_on s \ f integrable_on s" - using integrable_cmul[of "\x. c * f x" s "1 / c"] integrable_cmul[of f s c] \c \ 0\ - by auto - -lemma integrable_on_cmult_left: - assumes "f integrable_on s" - shows "(\x. of_real c * f x) integrable_on s" - using integrable_cmul[of f s "of_real c"] assms - by (simp add: scaleR_conv_of_real) - -lemma integrable_neg: "f integrable_on s \ (\x. -f(x)) integrable_on s" - unfolding integrable_on_def by(auto intro: has_integral_neg) - -lemma integrable_diff: - "f integrable_on s \ g integrable_on s \ (\x. f x - g x) integrable_on s" - unfolding integrable_on_def by(auto intro: has_integral_sub) - -lemma integrable_linear: - "f integrable_on s \ bounded_linear h \ (h \ f) integrable_on s" - unfolding integrable_on_def by(auto intro: has_integral_linear) - -lemma integral_linear: - "f integrable_on s \ bounded_linear h \ integral s (h \ f) = h (integral s f)" - apply (rule has_integral_unique [where i=s and f = "h \ f"]) - apply (simp_all add: integrable_integral integrable_linear has_integral_linear ) - done - -lemma integral_component_eq[simp]: - fixes f :: "'n::euclidean_space \ 'm::euclidean_space" - assumes "f integrable_on s" - shows "integral s (\x. f x \ k) = integral s f \ k" - unfolding integral_linear[OF assms(1) bounded_linear_component,unfolded o_def] .. - -lemma has_integral_setsum: - assumes "finite t" - and "\a\t. ((f a) has_integral (i a)) s" - shows "((\x. setsum (\a. f a x) t) has_integral (setsum i t)) s" - using assms(1) subset_refl[of t] -proof (induct rule: finite_subset_induct) - case empty - then show ?case by auto -next - case (insert x F) - with assms show ?case - by (simp add: has_integral_add) -qed - -lemma integral_setsum: - "\finite t; \a\t. (f a) integrable_on s\ \ - integral s (\x. setsum (\a. f a x) t) = setsum (\a. integral s (f a)) t" - by (auto intro: has_integral_setsum integrable_integral) - -lemma integrable_setsum: - "\finite t; \a\t. (f a) integrable_on s\ \ (\x. setsum (\a. f a x) t) integrable_on s" - unfolding integrable_on_def - apply (drule bchoice) - using has_integral_setsum[of t] - apply auto - done - -lemma has_integral_eq: - assumes "\x. x \ s \ f x = g x" - and "(f has_integral k) s" - shows "(g has_integral k) s" - using has_integral_sub[OF assms(2), of "\x. f x - g x" 0] - using has_integral_is_0[of s "\x. f x - g x"] - using assms(1) - by auto - -lemma integrable_eq: "(\x. x \ s \ f x = g x) \ f integrable_on s \ g integrable_on s" - unfolding integrable_on_def - using has_integral_eq[of s f g] has_integral_eq by blast - -lemma has_integral_cong: - assumes "\x. x \ s \ f x = g x" - shows "(f has_integral i) s = (g has_integral i) s" - using has_integral_eq[of s f g] has_integral_eq[of s g f] assms - by auto - -lemma integral_cong: - assumes "\x. x \ s \ f x = g x" - shows "integral s f = integral s g" - unfolding integral_def -by (metis (full_types, hide_lams) assms has_integral_cong integrable_eq) - -lemma integrable_on_cmult_left_iff [simp]: - assumes "c \ 0" - shows "(\x. of_real c * f x) integrable_on s \ f integrable_on s" - (is "?lhs = ?rhs") -proof - assume ?lhs - then have "(\x. of_real (1 / c) * (of_real c * f x)) integrable_on s" - using integrable_cmul[of "\x. of_real c * f x" s "1 / of_real c"] - by (simp add: scaleR_conv_of_real) - then have "(\x. (of_real (1 / c) * of_real c * f x)) integrable_on s" - by (simp add: algebra_simps) - with \c \ 0\ show ?rhs - by (metis (no_types, lifting) integrable_eq mult.left_neutral nonzero_divide_eq_eq of_real_1 of_real_mult) -qed (blast intro: integrable_on_cmult_left) - -lemma integrable_on_cmult_right: - fixes f :: "_ \ 'b :: {comm_ring,real_algebra_1,real_normed_vector}" - assumes "f integrable_on s" - shows "(\x. f x * of_real c) integrable_on s" -using integrable_on_cmult_left [OF assms] by (simp add: mult.commute) - -lemma integrable_on_cmult_right_iff [simp]: - fixes f :: "_ \ 'b :: {comm_ring,real_algebra_1,real_normed_vector}" - assumes "c \ 0" - shows "(\x. f x * of_real c) integrable_on s \ f integrable_on s" -using integrable_on_cmult_left_iff [OF assms] by (simp add: mult.commute) - -lemma integrable_on_cdivide: - fixes f :: "_ \ 'b :: real_normed_field" - assumes "f integrable_on s" - shows "(\x. f x / of_real c) integrable_on s" -by (simp add: integrable_on_cmult_right divide_inverse assms of_real_inverse [symmetric] del: of_real_inverse) - -lemma integrable_on_cdivide_iff [simp]: - fixes f :: "_ \ 'b :: real_normed_field" - assumes "c \ 0" - shows "(\x. f x / of_real c) integrable_on s \ f integrable_on s" -by (simp add: divide_inverse assms of_real_inverse [symmetric] del: of_real_inverse) - -lemma has_integral_null [intro]: - assumes "content(cbox a b) = 0" - shows "(f has_integral 0) (cbox a b)" -proof - - have "gauge (\x. ball x 1)" - by auto - moreover - { - fix e :: real - fix p - assume e: "e > 0" - assume p: "p tagged_division_of (cbox a b)" - have "norm ((\(x, k)\p. content k *\<^sub>R f x) - 0) = 0" - unfolding norm_eq_zero diff_0_right - using setsum_content_null[OF assms(1) p, of f] . - then have "norm ((\(x, k)\p. content k *\<^sub>R f x) - 0) < e" - using e by auto - } - ultimately show ?thesis - by (auto simp: has_integral) -qed - -lemma has_integral_null_real [intro]: - assumes "content {a .. b::real} = 0" - shows "(f has_integral 0) {a .. b}" - by (metis assms box_real(2) has_integral_null) - -lemma has_integral_null_eq[simp]: "content (cbox a b) = 0 \ (f has_integral i) (cbox a b) \ i = 0" - by (auto simp add: has_integral_null dest!: integral_unique) - -lemma integral_null [simp]: "content (cbox a b) = 0 \ integral (cbox a b) f = 0" - by (metis has_integral_null integral_unique) - -lemma integrable_on_null [intro]: "content (cbox a b) = 0 \ f integrable_on (cbox a b)" - by (simp add: has_integral_integrable) - -lemma has_integral_empty[intro]: "(f has_integral 0) {}" - by (simp add: has_integral_is_0) - -lemma has_integral_empty_eq[simp]: "(f has_integral i) {} \ i = 0" - by (auto simp add: has_integral_empty has_integral_unique) - -lemma integrable_on_empty[intro]: "f integrable_on {}" - unfolding integrable_on_def by auto - -lemma integral_empty[simp]: "integral {} f = 0" - by (rule integral_unique) (rule has_integral_empty) - -lemma has_integral_refl[intro]: - fixes a :: "'a::euclidean_space" - shows "(f has_integral 0) (cbox a a)" - and "(f has_integral 0) {a}" -proof - - have *: "{a} = cbox a a" - apply (rule set_eqI) - unfolding mem_box singleton_iff euclidean_eq_iff[where 'a='a] - apply safe - prefer 3 - apply (erule_tac x=b in ballE) - apply (auto simp add: field_simps) - done - show "(f has_integral 0) (cbox a a)" "(f has_integral 0) {a}" - unfolding * - apply (rule_tac[!] has_integral_null) - unfolding content_eq_0_interior - unfolding interior_cbox - using box_sing - apply auto - done -qed - -lemma integrable_on_refl[intro]: "f integrable_on cbox a a" - unfolding integrable_on_def by auto - -lemma integral_refl [simp]: "integral (cbox a a) f = 0" - by (rule integral_unique) auto - -lemma integral_singleton [simp]: "integral {a} f = 0" - by auto - -lemma integral_blinfun_apply: - assumes "f integrable_on s" - shows "integral s (\x. blinfun_apply h (f x)) = blinfun_apply h (integral s f)" - by (subst integral_linear[symmetric, OF assms blinfun.bounded_linear_right]) (simp add: o_def) - -lemma blinfun_apply_integral: - assumes "f integrable_on s" - shows "blinfun_apply (integral s f) x = integral s (\y. blinfun_apply (f y) x)" - by (metis (no_types, lifting) assms blinfun.prod_left.rep_eq integral_blinfun_apply integral_cong) - -lemma has_integral_componentwise_iff: - fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" - shows "(f has_integral y) A \ (\b\Basis. ((\x. f x \ b) has_integral (y \ b)) A)" -proof safe - fix b :: 'b assume "(f has_integral y) A" - from has_integral_linear[OF this(1) bounded_linear_component, of b] - show "((\x. f x \ b) has_integral (y \ b)) A" by (simp add: o_def) -next - assume "(\b\Basis. ((\x. f x \ b) has_integral (y \ b)) A)" - hence "\b\Basis. (((\x. x *\<^sub>R b) \ (\x. f x \ b)) has_integral ((y \ b) *\<^sub>R b)) A" - by (intro ballI has_integral_linear) (simp_all add: bounded_linear_scaleR_left) - hence "((\x. \b\Basis. (f x \ b) *\<^sub>R b) has_integral (\b\Basis. (y \ b) *\<^sub>R b)) A" - by (intro has_integral_setsum) (simp_all add: o_def) - thus "(f has_integral y) A" by (simp add: euclidean_representation) -qed - -lemma has_integral_componentwise: - fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" - shows "(\b. b \ Basis \ ((\x. f x \ b) has_integral (y \ b)) A) \ (f has_integral y) A" - by (subst has_integral_componentwise_iff) blast - -lemma integrable_componentwise_iff: - fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" - shows "f integrable_on A \ (\b\Basis. (\x. f x \ b) integrable_on A)" -proof - assume "f integrable_on A" - then obtain y where "(f has_integral y) A" by (auto simp: integrable_on_def) - hence "(\b\Basis. ((\x. f x \ b) has_integral (y \ b)) A)" - by (subst (asm) has_integral_componentwise_iff) - thus "(\b\Basis. (\x. f x \ b) integrable_on A)" by (auto simp: integrable_on_def) -next - assume "(\b\Basis. (\x. f x \ b) integrable_on A)" - then obtain y where "\b\Basis. ((\x. f x \ b) has_integral y b) A" - unfolding integrable_on_def by (subst (asm) bchoice_iff) blast - hence "\b\Basis. (((\x. x *\<^sub>R b) \ (\x. f x \ b)) has_integral (y b *\<^sub>R b)) A" - by (intro ballI has_integral_linear) (simp_all add: bounded_linear_scaleR_left) - hence "((\x. \b\Basis. (f x \ b) *\<^sub>R b) has_integral (\b\Basis. y b *\<^sub>R b)) A" - by (intro has_integral_setsum) (simp_all add: o_def) - thus "f integrable_on A" by (auto simp: integrable_on_def o_def euclidean_representation) -qed - -lemma integrable_componentwise: - fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" - shows "(\b. b \ Basis \ (\x. f x \ b) integrable_on A) \ f integrable_on A" - by (subst integrable_componentwise_iff) blast - -lemma integral_componentwise: - fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" - assumes "f integrable_on A" - shows "integral A f = (\b\Basis. integral A (\x. (f x \ b) *\<^sub>R b))" -proof - - from assms have integrable: "\b\Basis. (\x. x *\<^sub>R b) \ (\x. (f x \ b)) integrable_on A" - by (subst (asm) integrable_componentwise_iff, intro integrable_linear ballI) - (simp_all add: bounded_linear_scaleR_left) - have "integral A f = integral A (\x. \b\Basis. (f x \ b) *\<^sub>R b)" - by (simp add: euclidean_representation) - also from integrable have "\ = (\a\Basis. integral A (\x. (f x \ a) *\<^sub>R a))" - by (subst integral_setsum) (simp_all add: o_def) - finally show ?thesis . -qed - -lemma integrable_component: - "f integrable_on A \ (\x. f x \ (y :: 'b :: euclidean_space)) integrable_on A" - by (drule integrable_linear[OF _ bounded_linear_component[of y]]) (simp add: o_def) - - - -subsection \Cauchy-type criterion for integrability.\ - -(* XXXXXXX *) -lemma integrable_cauchy: - fixes f :: "'n::euclidean_space \ 'a::{real_normed_vector,complete_space}" - shows "f integrable_on cbox a b \ - (\e>0.\d. gauge d \ - (\p1 p2. p1 tagged_division_of (cbox a b) \ d fine p1 \ - p2 tagged_division_of (cbox a b) \ d fine p2 \ - norm (setsum (\(x,k). content k *\<^sub>R f x) p1 - - setsum (\(x,k). content k *\<^sub>R f x) p2) < e))" - (is "?l = (\e>0. \d. ?P e d)") -proof - assume ?l - then guess y unfolding integrable_on_def has_integral .. note y=this - show "\e>0. \d. ?P e d" - proof (clarify, goal_cases) - case (1 e) - then have "e/2 > 0" by auto - then guess d - apply - - apply (drule y[rule_format]) - apply (elim exE conjE) - done - note d=this[rule_format] - show ?case - proof (rule_tac x=d in exI, clarsimp simp: d) - fix p1 p2 - assume as: "p1 tagged_division_of (cbox a b)" "d fine p1" - "p2 tagged_division_of (cbox a b)" "d fine p2" - show "norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e" - apply (rule dist_triangle_half_l[where y=y,unfolded dist_norm]) - using d(2)[OF conjI[OF as(1-2)]] d(2)[OF conjI[OF as(3-4)]] . - qed - qed -next - assume "\e>0. \d. ?P e d" - then have "\n::nat. \d. ?P (inverse(of_nat (n + 1))) d" - by auto - from choice[OF this] guess d .. note d=conjunctD2[OF this[rule_format],rule_format] - have "\n. gauge (\x. \{d i x |i. i \ {0..n}})" - apply (rule gauge_inters) - using d(1) - apply auto - done - then have "\n. \p. p tagged_division_of (cbox a b) \ (\x. \{d i x |i. i \ {0..n}}) fine p" - by (meson fine_division_exists) - from choice[OF this] guess p .. note p = conjunctD2[OF this[rule_format]] - have dp: "\i n. i\n \ d i fine p n" - using p(2) unfolding fine_inters by auto - have "Cauchy (\n. setsum (\(x,k). content k *\<^sub>R (f x)) (p n))" - proof (rule CauchyI, goal_cases) - case (1 e) - then guess N unfolding real_arch_inverse[of e] .. note N=this - show ?case - apply (rule_tac x=N in exI) - proof clarify - fix m n - assume mn: "N \ m" "N \ n" - have *: "N = (N - 1) + 1" using N by auto - show "norm ((\(x, k)\p m. content k *\<^sub>R f x) - (\(x, k)\p n. content k *\<^sub>R f x)) < e" - apply (rule less_trans[OF _ N[THEN conjunct2,THEN conjunct2]]) - apply(subst *) - using dp p(1) mn d(2) by auto - qed - qed - then guess y unfolding convergent_eq_cauchy[symmetric] .. note y=this[THEN LIMSEQ_D] - show ?l - unfolding integrable_on_def has_integral - proof (rule_tac x=y in exI, clarify) - fix e :: real - assume "e>0" - then have *:"e/2 > 0" by auto - then guess N1 unfolding real_arch_inverse[of "e/2"] .. note N1=this - then have N1': "N1 = N1 - 1 + 1" - by auto - guess N2 using y[OF *] .. note N2=this - have "gauge (d (N1 + N2))" - using d by auto - moreover - { - fix q - assume as: "q tagged_division_of (cbox a b)" "d (N1 + N2) fine q" - have *: "inverse (of_nat (N1 + N2 + 1)) < e / 2" - apply (rule less_trans) - using N1 - apply auto - done - have "norm ((\(x, k)\q. content k *\<^sub>R f x) - y) < e" - apply (rule norm_triangle_half_r) - apply (rule less_trans[OF _ *]) - apply (subst N1', rule d(2)[of "p (N1+N2)"]) - using N1' as(1) as(2) dp - apply (simp add: \\x. p x tagged_division_of cbox a b \ (\xa. \{d i xa |i. i \ {0..x}}) fine p x\) - using N2 le_add2 by blast - } - ultimately show "\d. gauge d \ - (\p. p tagged_division_of (cbox a b) \ d fine p \ - norm ((\(x, k)\p. content k *\<^sub>R f x) - y) < e)" - by (rule_tac x="d (N1 + N2)" in exI) auto - qed -qed - - -subsection \Additivity of integral on abutting intervals.\ - -lemma tagged_division_split_left_inj: - fixes x1 :: "'a::euclidean_space" - assumes d: "d tagged_division_of i" - and k12: "(x1, k1) \ d" - "(x2, k2) \ d" - "k1 \ k2" - "k1 \ {x. x\k \ c} = k2 \ {x. x\k \ c}" - "k \ Basis" - shows "content (k1 \ {x. x\k \ c}) = 0" -proof - - have *: "\a b c. (a,b) \ c \ b \ snd ` c" - by force - show ?thesis - using k12 - by (fastforce intro!: division_split_left_inj[OF division_of_tagged_division[OF d]] *) -qed - -lemma tagged_division_split_right_inj: - fixes x1 :: "'a::euclidean_space" - assumes d: "d tagged_division_of i" - and k12: "(x1, k1) \ d" - "(x2, k2) \ d" - "k1 \ k2" - "k1 \ {x. x\k \ c} = k2 \ {x. x\k \ c}" - "k \ Basis" - shows "content (k1 \ {x. x\k \ c}) = 0" -proof - - have *: "\a b c. (a,b) \ c \ b \ snd ` c" - by force - show ?thesis - using k12 - by (fastforce intro!: division_split_right_inj[OF division_of_tagged_division[OF d]] *) -qed - -lemma has_integral_split: - fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" - assumes fi: "(f has_integral i) (cbox a b \ {x. x\k \ c})" - and fj: "(f has_integral j) (cbox a b \ {x. x\k \ c})" - and k: "k \ Basis" - shows "(f has_integral (i + j)) (cbox a b)" -proof (unfold has_integral, rule, rule, goal_cases) - case (1 e) - then have e: "e/2 > 0" - by auto - obtain d1 - where d1: "gauge d1" - and d1norm: - "\p. \p tagged_division_of cbox a b \ {x. x \ k \ c}; - d1 fine p\ \ norm ((\(x, k) \ p. content k *\<^sub>R f x) - i) < e / 2" - apply (rule has_integralD[OF fi[unfolded interval_split[OF k]] e]) - apply (simp add: interval_split[symmetric] k) - done - obtain d2 - where d2: "gauge d2" - and d2norm: - "\p. \p tagged_division_of cbox a b \ {x. c \ x \ k}; - d2 fine p\ \ norm ((\(x, k) \ p. content k *\<^sub>R f x) - j) < e / 2" - apply (rule has_integralD[OF fj[unfolded interval_split[OF k]] e]) - apply (simp add: interval_split[symmetric] k) - done - let ?d = "\x. if x\k = c then (d1 x \ d2 x) else ball x \x\k - c\ \ d1 x \ d2 x" - have "gauge ?d" - using d1 d2 unfolding gauge_def by auto - then show ?case - proof (rule_tac x="?d" in exI, safe) - fix p - assume "p tagged_division_of (cbox a b)" "?d fine p" - note p = this tagged_division_ofD[OF this(1)] - have xk_le_c: "\x kk. (x, kk) \ p \ kk \ {x. x\k \ c} \ {} \ x\k \ c" - proof - - fix x kk - assume as: "(x, kk) \ p" and kk: "kk \ {x. x\k \ c} \ {}" - show "x\k \ c" - proof (rule ccontr) - assume **: "\ ?thesis" - from this[unfolded not_le] - have "kk \ ball x \x \ k - c\" - using p(2)[unfolded fine_def, rule_format,OF as] by auto - with kk obtain y where y: "y \ ball x \x \ k - c\" "y\k \ c" - by blast - then have "\x \ k - y \ k\ < \x \ k - c\" - using Basis_le_norm[OF k, of "x - y"] - by (auto simp add: dist_norm inner_diff_left intro: le_less_trans) - with y show False - using ** by (auto simp add: field_simps) - qed - qed - have xk_ge_c: "\x kk. (x, kk) \ p \ kk \ {x. x\k \ c} \ {} \ x\k \ c" - proof - - fix x kk - assume as: "(x, kk) \ p" and kk: "kk \ {x. x\k \ c} \ {}" - show "x\k \ c" - proof (rule ccontr) - assume **: "\ ?thesis" - from this[unfolded not_le] have "kk \ ball x \x \ k - c\" - using p(2)[unfolded fine_def,rule_format,OF as,unfolded split_conv] by auto - with kk obtain y where y: "y \ ball x \x \ k - c\" "y\k \ c" - by blast - then have "\x \ k - y \ k\ < \x \ k - c\" - using Basis_le_norm[OF k, of "x - y"] - by (auto simp add: dist_norm inner_diff_left intro: le_less_trans) - with y show False - using ** by (auto simp add: field_simps) - qed - qed - - have lem1: "\f P Q. (\x k. (x, k) \ {(x, f k) | x k. P x k} \ Q x k) \ - (\x k. P x k \ Q x (f k))" - by auto - have fin_finite: "finite {(x,f k) | x k. (x,k) \ s \ P x k}" if "finite s" for f s P - proof - - from that have "finite ((\(x, k). (x, f k)) ` s)" - by auto - then show ?thesis - by (rule rev_finite_subset) auto - qed - { fix g :: "'a set \ 'a set" - fix i :: "'a \ 'a set" - assume "i \ (\(x, k). (x, g k)) ` p - {(x, g k) |x k. (x, k) \ p \ g k \ {}}" - then obtain x k where xk: - "i = (x, g k)" "(x, k) \ p" - "(x, g k) \ {(x, g k) |x k. (x, k) \ p \ g k \ {}}" - by auto - have "content (g k) = 0" - using xk using content_empty by auto - then have "(\(x, k). content k *\<^sub>R f x) i = 0" - unfolding xk split_conv by auto - } note [simp] = this - have lem3: "\g :: 'a set \ 'a set. finite p \ - setsum (\(x, k). content k *\<^sub>R f x) {(x,g k) |x k. (x,k) \ p \ g k \ {}} = - setsum (\(x, k). content k *\<^sub>R f x) ((\(x, k). (x, g k)) ` p)" - by (rule setsum.mono_neutral_left) auto - let ?M1 = "{(x, kk \ {x. x\k \ c}) |x kk. (x, kk) \ p \ kk \ {x. x\k \ c} \ {}}" - have d1_fine: "d1 fine ?M1" - by (force intro: fineI dest: fineD[OF p(2)] simp add: split: if_split_asm) - have "norm ((\(x, k)\?M1. content k *\<^sub>R f x) - i) < e/2" - proof (rule d1norm [OF tagged_division_ofI d1_fine]) - show "finite ?M1" - by (rule fin_finite p(3))+ - show "\{k. \x. (x, k) \ ?M1} = cbox a b \ {x. x\k \ c}" - unfolding p(8)[symmetric] by auto - fix x l - assume xl: "(x, l) \ ?M1" - then guess x' l' unfolding mem_Collect_eq unfolding prod.inject by (elim exE conjE) note xl'=this - show "x \ l" "l \ cbox a b \ {x. x \ k \ c}" - unfolding xl' - using p(4-6)[OF xl'(3)] using xl'(4) - using xk_le_c[OF xl'(3-4)] by auto - show "\a b. l = cbox a b" - unfolding xl' - using p(6)[OF xl'(3)] - by (fastforce simp add: interval_split[OF k,where c=c]) - fix y r - let ?goal = "interior l \ interior r = {}" - assume yr: "(y, r) \ ?M1" - then guess y' r' unfolding mem_Collect_eq unfolding prod.inject by (elim exE conjE) note yr'=this - assume as: "(x, l) \ (y, r)" - show "interior l \ interior r = {}" - proof (cases "l' = r' \ x' = y'") - case False - then show ?thesis - using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto - next - case True - then have "l' \ r'" - using as unfolding xl' yr' by auto - then show ?thesis - using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto - qed - qed - moreover - let ?M2 = "{(x,kk \ {x. x\k \ c}) |x kk. (x,kk) \ p \ kk \ {x. x\k \ c} \ {}}" - have d2_fine: "d2 fine ?M2" - by (force intro: fineI dest: fineD[OF p(2)] simp add: split: if_split_asm) - have "norm ((\(x, k)\?M2. content k *\<^sub>R f x) - j) < e/2" - proof (rule d2norm [OF tagged_division_ofI d2_fine]) - show "finite ?M2" - by (rule fin_finite p(3))+ - show "\{k. \x. (x, k) \ ?M2} = cbox a b \ {x. x\k \ c}" - unfolding p(8)[symmetric] by auto - fix x l - assume xl: "(x, l) \ ?M2" - then guess x' l' unfolding mem_Collect_eq unfolding prod.inject by (elim exE conjE) note xl'=this - show "x \ l" "l \ cbox a b \ {x. x \ k \ c}" - unfolding xl' - using p(4-6)[OF xl'(3)] xl'(4) xk_ge_c[OF xl'(3-4)] - by auto - show "\a b. l = cbox a b" - unfolding xl' - using p(6)[OF xl'(3)] - by (fastforce simp add: interval_split[OF k, where c=c]) - fix y r - let ?goal = "interior l \ interior r = {}" - assume yr: "(y, r) \ ?M2" - then guess y' r' unfolding mem_Collect_eq unfolding prod.inject by (elim exE conjE) note yr'=this - assume as: "(x, l) \ (y, r)" - show "interior l \ interior r = {}" - proof (cases "l' = r' \ x' = y'") - case False - then show ?thesis - using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto - next - case True - then have "l' \ r'" - using as unfolding xl' yr' by auto - then show ?thesis - using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto - qed - qed - ultimately - have "norm (((\(x, k)\?M1. content k *\<^sub>R f x) - i) + ((\(x, k)\?M2. content k *\<^sub>R f x) - j)) < e/2 + e/2" - using norm_add_less by blast - also { - have eq0: "\x y. x = (0::real) \ x *\<^sub>R (y::'b) = 0" - using scaleR_zero_left by auto - have cont_eq: "\g. (\(x,l). content l *\<^sub>R f x) \ (\(x,l). (x,g l)) = (\(x,l). content (g l) *\<^sub>R f x)" - by auto - have "((\(x, k)\?M1. content k *\<^sub>R f x) - i) + ((\(x, k)\?M2. content k *\<^sub>R f x) - j) = - (\(x, k)\?M1. content k *\<^sub>R f x) + (\(x, k)\?M2. content k *\<^sub>R f x) - (i + j)" - by auto - also have "\ = (\(x, ka)\p. content (ka \ {x. x \ k \ c}) *\<^sub>R f x) + - (\(x, ka)\p. content (ka \ {x. c \ x \ k}) *\<^sub>R f x) - (i + j)" - unfolding lem3[OF p(3)] - by (subst setsum.reindex_nontrivial[OF p(3)], auto intro!: k eq0 tagged_division_split_left_inj[OF p(1)] tagged_division_split_right_inj[OF p(1)] - simp: cont_eq)+ - also note setsum.distrib[symmetric] - also have "\x. x \ p \ - (\(x,ka). content (ka \ {x. x \ k \ c}) *\<^sub>R f x) x + - (\(x,ka). content (ka \ {x. c \ x \ k}) *\<^sub>R f x) x = - (\(x,ka). content ka *\<^sub>R f x) x" - proof clarify - fix a b - assume "(a, b) \ p" - from p(6)[OF this] guess u v by (elim exE) note uv=this - then show "content (b \ {x. x \ k \ c}) *\<^sub>R f a + content (b \ {x. c \ x \ k}) *\<^sub>R f a = - content b *\<^sub>R f a" - unfolding scaleR_left_distrib[symmetric] - unfolding uv content_split[OF k,of u v c] - by auto - qed - note setsum.cong [OF _ this] - finally have "(\(x, k)\{(x, kk \ {x. x \ k \ c}) |x kk. (x, kk) \ p \ kk \ {x. x \ k \ c} \ {}}. content k *\<^sub>R f x) - i + - ((\(x, k)\{(x, kk \ {x. c \ x \ k}) |x kk. (x, kk) \ p \ kk \ {x. c \ x \ k} \ {}}. content k *\<^sub>R f x) - j) = - (\(x, ka)\p. content ka *\<^sub>R f x) - (i + j)" - by auto - } - finally show "norm ((\(x, k)\p. content k *\<^sub>R f x) - (i + j)) < e" - by auto - qed -qed - - -subsection \A sort of converse, integrability on subintervals.\ - -lemma tagged_division_union_interval: - fixes a :: "'a::euclidean_space" - assumes "p1 tagged_division_of (cbox a b \ {x. x\k \ (c::real)})" - and "p2 tagged_division_of (cbox a b \ {x. x\k \ c})" - and k: "k \ Basis" - shows "(p1 \ p2) tagged_division_of (cbox a b)" -proof - - have *: "cbox a b = (cbox a b \ {x. x\k \ c}) \ (cbox a b \ {x. x\k \ c})" - by auto - show ?thesis - apply (subst *) - apply (rule tagged_division_union[OF assms(1-2)]) - unfolding interval_split[OF k] interior_cbox - using k - apply (auto simp add: box_def elim!: ballE[where x=k]) - done -qed - -lemma tagged_division_union_interval_real: - fixes a :: real - assumes "p1 tagged_division_of ({a .. b} \ {x. x\k \ (c::real)})" - and "p2 tagged_division_of ({a .. b} \ {x. x\k \ c})" - and k: "k \ Basis" - shows "(p1 \ p2) tagged_division_of {a .. b}" - using assms - unfolding box_real[symmetric] - by (rule tagged_division_union_interval) - -lemma has_integral_separate_sides: - fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" - assumes "(f has_integral i) (cbox a b)" - and "e > 0" - and k: "k \ Basis" - obtains d where "gauge d" - "\p1 p2. p1 tagged_division_of (cbox a b \ {x. x\k \ c}) \ d fine p1 \ - p2 tagged_division_of (cbox a b \ {x. x\k \ c}) \ d fine p2 \ - norm ((setsum (\(x,k). content k *\<^sub>R f x) p1 + setsum (\(x,k). content k *\<^sub>R f x) p2) - i) < e" -proof - - guess d using has_integralD[OF assms(1-2)] . note d=this - { fix p1 p2 - assume "p1 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p1" - note p1=tagged_division_ofD[OF this(1)] this - assume "p2 tagged_division_of (cbox a b) \ {x. c \ x \ k}" "d fine p2" - note p2=tagged_division_ofD[OF this(1)] this - note tagged_division_union_interval[OF p1(7) p2(7)] note p12 = tagged_division_ofD[OF this] this - { fix a b - assume ab: "(a, b) \ p1 \ p2" - have "(a, b) \ p1" - using ab by auto - with p1 obtain u v where uv: "b = cbox u v" by auto - have "b \ {x. x\k = c}" - using ab p1(3)[of a b] p2(3)[of a b] by fastforce - moreover - have "interior {x::'a. x \ k = c} = {}" - proof (rule ccontr) - assume "\ ?thesis" - then obtain x where x: "x \ interior {x::'a. x\k = c}" - by auto - then guess e unfolding mem_interior .. note e=this - have x: "x\k = c" - using x interior_subset by fastforce - have *: "\i. i \ Basis \ \(x - (x + (e / 2) *\<^sub>R k)) \ i\ = (if i = k then e/2 else 0)" - using e k by (auto simp: inner_simps inner_not_same_Basis) - have "(\i\Basis. \(x - (x + (e / 2 ) *\<^sub>R k)) \ i\) = - (\i\Basis. (if i = k then e / 2 else 0))" - using "*" by (blast intro: setsum.cong) - also have "\ < e" - apply (subst setsum.delta) - using e - apply auto - done - finally have "x + (e/2) *\<^sub>R k \ ball x e" - unfolding mem_ball dist_norm by(rule le_less_trans[OF norm_le_l1]) - then have "x + (e/2) *\<^sub>R k \ {x. x\k = c}" - using e by auto - then show False - unfolding mem_Collect_eq using e x k by (auto simp: inner_simps) - qed - ultimately have "content b = 0" - unfolding uv content_eq_0_interior - using interior_mono by blast - then have "content b *\<^sub>R f a = 0" - by auto - } - then have "norm ((\(x, k)\p1. content k *\<^sub>R f x) + (\(x, k)\p2. content k *\<^sub>R f x) - i) = - norm ((\(x, k)\p1 \ p2. content k *\<^sub>R f x) - i)" - by (subst setsum.union_inter_neutral) (auto simp: p1 p2) - also have "\ < e" - by (rule k d(2) p12 fine_union p1 p2)+ - finally have "norm ((\(x, k)\p1. content k *\<^sub>R f x) + (\(x, k)\p2. content k *\<^sub>R f x) - i) < e" . - } - then show ?thesis - by (auto intro: that[of d] d elim: ) -qed - -lemma integrable_split[intro]: - fixes f :: "'a::euclidean_space \ 'b::{real_normed_vector,complete_space}" - assumes "f integrable_on cbox a b" - and k: "k \ Basis" - shows "f integrable_on (cbox a b \ {x. x\k \ c})" (is ?t1) - and "f integrable_on (cbox a b \ {x. x\k \ c})" (is ?t2) -proof - - guess y using assms(1) unfolding integrable_on_def .. note y=this - define b' where "b' = (\i\Basis. (if i = k then min (b\k) c else b\i)*\<^sub>R i)" - define a' where "a' = (\i\Basis. (if i = k then max (a\k) c else a\i)*\<^sub>R i)" - show ?t1 ?t2 - unfolding interval_split[OF k] integrable_cauchy - unfolding interval_split[symmetric,OF k] - proof (rule_tac[!] allI impI)+ - fix e :: real - assume "e > 0" - then have "e/2>0" - by auto - from has_integral_separate_sides[OF y this k,of c] guess d . note d=this[rule_format] - let ?P = "\A. \d. gauge d \ (\p1 p2. p1 tagged_division_of (cbox a b) \ A \ d fine p1 \ - p2 tagged_division_of (cbox a b) \ A \ d fine p2 \ - norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e)" - show "?P {x. x \ k \ c}" - proof (rule_tac x=d in exI, clarsimp simp add: d) - fix p1 p2 - assume as: "p1 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p1" - "p2 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p2" - show "norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e" - proof (rule fine_division_exists[OF d(1), of a' b] ) - fix p - assume "p tagged_division_of cbox a' b" "d fine p" - then show ?thesis - using as norm_triangle_half_l[OF d(2)[of p1 p] d(2)[of p2 p]] - unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric] - by (auto simp add: algebra_simps) - qed - qed - show "?P {x. x \ k \ c}" - proof (rule_tac x=d in exI, clarsimp simp add: d) - fix p1 p2 - assume as: "p1 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p1" - "p2 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p2" - show "norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e" - proof (rule fine_division_exists[OF d(1), of a b'] ) - fix p - assume "p tagged_division_of cbox a b'" "d fine p" - then show ?thesis - using as norm_triangle_half_l[OF d(2)[of p p1] d(2)[of p p2]] - unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric] - by (auto simp add: algebra_simps) - qed - qed - qed -qed - -lemma operative_integral: - fixes f :: "'a::euclidean_space \ 'b::banach" - shows "comm_monoid.operative (lift_option op +) (Some 0) (\i. if f integrable_on i then Some (integral i f) else None)" - unfolding comm_monoid.operative_def[OF add.comm_monoid_lift_option] -proof safe - fix a b c - fix k :: 'a - assume k: "k \ Basis" - show "(if f integrable_on cbox a b then Some (integral (cbox a b) f) else None) = - lift_option op + (if f integrable_on cbox a b \ {x. x \ k \ c} then Some (integral (cbox a b \ {x. x \ k \ c}) f) else None) - (if f integrable_on cbox a b \ {x. c \ x \ k} then Some (integral (cbox a b \ {x. c \ x \ k}) f) else None)" - proof (cases "f integrable_on cbox a b") - case True - with k show ?thesis - apply (simp add: integrable_split) - apply (rule integral_unique [OF has_integral_split[OF _ _ k]]) - apply (auto intro: integrable_integral) - done - next - case False - have "\ (f integrable_on cbox a b \ {x. x \ k \ c}) \ \ ( f integrable_on cbox a b \ {x. c \ x \ k})" - proof (rule ccontr) - assume "\ ?thesis" - then have "f integrable_on cbox a b" - unfolding integrable_on_def - apply (rule_tac x="integral (cbox a b \ {x. x \ k \ c}) f + integral (cbox a b \ {x. x \ k \ c}) f" in exI) - apply (rule has_integral_split[OF _ _ k]) - apply (auto intro: integrable_integral) - done - then show False - using False by auto - qed - then show ?thesis - using False by auto - qed -next - fix a b :: 'a - assume "content (cbox a b) = 0" - then show "(if f integrable_on cbox a b then Some (integral (cbox a b) f) else None) = Some 0" - using has_integral_null_eq - by (auto simp: integrable_on_null) -qed - -subsection \Finally, the integral of a constant\ - -lemma has_integral_const [intro]: - fixes a b :: "'a::euclidean_space" - shows "((\x. c) has_integral (content (cbox a b) *\<^sub>R c)) (cbox a b)" - apply (auto intro!: exI [where x="\x. ball x 1"] simp: split_def has_integral) - apply (subst scaleR_left.setsum[symmetric, unfolded o_def]) - apply (subst additive_content_tagged_division[unfolded split_def]) - apply auto - done - -lemma has_integral_const_real [intro]: - fixes a b :: real - shows "((\x. c) has_integral (content {a .. b} *\<^sub>R c)) {a .. b}" - by (metis box_real(2) has_integral_const) - -lemma integral_const [simp]: - fixes a b :: "'a::euclidean_space" - shows "integral (cbox a b) (\x. c) = content (cbox a b) *\<^sub>R c" - by (rule integral_unique) (rule has_integral_const) - -lemma integral_const_real [simp]: - fixes a b :: real - shows "integral {a .. b} (\x. c) = content {a .. b} *\<^sub>R c" - by (metis box_real(2) integral_const) - - -subsection \Bounds on the norm of Riemann sums and the integral itself.\ - -lemma dsum_bound: - assumes "p division_of (cbox a b)" - and "norm c \ e" - shows "norm (setsum (\l. content l *\<^sub>R c) p) \ e * content(cbox a b)" -proof - - have sumeq: "(\i\p. \content i\) = setsum content p" - apply (rule setsum.cong) - using assms - apply simp - apply (metis abs_of_nonneg assms(1) content_pos_le division_ofD(4)) - done - have e: "0 \ e" - using assms(2) norm_ge_zero order_trans by blast - have "norm (setsum (\l. content l *\<^sub>R c) p) \ (\i\p. norm (content i *\<^sub>R c))" - using norm_setsum by blast - also have "... \ e * (\i\p. \content i\)" - apply (simp add: setsum_right_distrib[symmetric] mult.commute) - using assms(2) mult_right_mono by blast - also have "... \ e * content (cbox a b)" - apply (rule mult_left_mono [OF _ e]) - apply (simp add: sumeq) - using additive_content_division assms(1) eq_iff apply blast - done - finally show ?thesis . -qed - -lemma rsum_bound: - assumes p: "p tagged_division_of (cbox a b)" - and "\x\cbox a b. norm (f x) \ e" - shows "norm (setsum (\(x,k). content k *\<^sub>R f x) p) \ e * content (cbox a b)" -proof (cases "cbox a b = {}") - case True show ?thesis - using p unfolding True tagged_division_of_trivial by auto -next - case False - then have e: "e \ 0" - by (meson ex_in_conv assms(2) norm_ge_zero order_trans) - have setsum_le: "setsum (content \ snd) p \ content (cbox a b)" - unfolding additive_content_tagged_division[OF p, symmetric] split_def - by (auto intro: eq_refl) - have con: "\xk. xk \ p \ 0 \ content (snd xk)" - using tagged_division_ofD(4) [OF p] content_pos_le - by force - have norm: "\xk. xk \ p \ norm (f (fst xk)) \ e" - unfolding fst_conv using tagged_division_ofD(2,3)[OF p] assms - by (metis prod.collapse subset_eq) - have "norm (setsum (\(x,k). content k *\<^sub>R f x) p) \ (\i\p. norm (case i of (x, k) \ content k *\<^sub>R f x))" - by (rule norm_setsum) - also have "... \ e * content (cbox a b)" - unfolding split_def norm_scaleR - apply (rule order_trans[OF setsum_mono]) - apply (rule mult_left_mono[OF _ abs_ge_zero, of _ e]) - apply (metis norm) - unfolding setsum_left_distrib[symmetric] - using con setsum_le - apply (auto simp: mult.commute intro: mult_left_mono [OF _ e]) - done - finally show ?thesis . -qed - -lemma rsum_diff_bound: - assumes "p tagged_division_of (cbox a b)" - and "\x\cbox a b. norm (f x - g x) \ e" - shows "norm (setsum (\(x,k). content k *\<^sub>R f x) p - setsum (\(x,k). content k *\<^sub>R g x) p) \ - e * content (cbox a b)" - apply (rule order_trans[OF _ rsum_bound[OF assms]]) - apply (simp add: split_def scaleR_diff_right setsum_subtractf eq_refl) - done - -lemma has_integral_bound: - fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" - assumes "0 \ B" - and "(f has_integral i) (cbox a b)" - and "\x\cbox a b. norm (f x) \ B" - shows "norm i \ B * content (cbox a b)" -proof (rule ccontr) - assume "\ ?thesis" - then have *: "norm i - B * content (cbox a b) > 0" - by auto - from assms(2)[unfolded has_integral,rule_format,OF *] - guess d by (elim exE conjE) note d=this[rule_format] - from fine_division_exists[OF this(1), of a b] guess p . note p=this - have *: "\s B. norm s \ B \ \ norm (s - i) < norm i - B" - unfolding not_less - by (metis norm_triangle_sub[of i] add.commute le_less_trans less_diff_eq linorder_not_le norm_minus_commute) - show False - using d(2)[OF conjI[OF p]] *[OF rsum_bound[OF p(1) assms(3)]] by auto -qed - -corollary has_integral_bound_real: - fixes f :: "real \ 'b::real_normed_vector" - assumes "0 \ B" - and "(f has_integral i) {a .. b}" - and "\x\{a .. b}. norm (f x) \ B" - shows "norm i \ B * content {a .. b}" - by (metis assms box_real(2) has_integral_bound) - -corollary integrable_bound: - fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" - assumes "0 \ B" - and "f integrable_on (cbox a b)" - and "\x. x\cbox a b \ norm (f x) \ B" - shows "norm (integral (cbox a b) f) \ B * content (cbox a b)" -by (metis integrable_integral has_integral_bound assms) - - -subsection \Similar theorems about relationship among components.\ - -lemma rsum_component_le: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "p tagged_division_of (cbox a b)" - and "\x\cbox a b. (f x)\i \ (g x)\i" - shows "(setsum (\(x,k). content k *\<^sub>R f x) p)\i \ (setsum (\(x,k). content k *\<^sub>R g x) p)\i" -unfolding inner_setsum_left -proof (rule setsum_mono, clarify) - fix a b - assume ab: "(a, b) \ p" - note tagged = tagged_division_ofD(2-4)[OF assms(1) ab] - from this(3) guess u v by (elim exE) note b=this - show "(content b *\<^sub>R f a) \ i \ (content b *\<^sub>R g a) \ i" - unfolding b inner_simps real_scaleR_def - apply (rule mult_left_mono) - using assms(2) tagged - by (auto simp add: content_pos_le) -qed - -lemma has_integral_component_le: - fixes f g :: "'a::euclidean_space \ 'b::euclidean_space" - assumes k: "k \ Basis" - assumes "(f has_integral i) s" "(g has_integral j) s" - and "\x\s. (f x)\k \ (g x)\k" - shows "i\k \ j\k" -proof - - have lem: "i\k \ j\k" - if f_i: "(f has_integral i) (cbox a b)" - and g_j: "(g has_integral j) (cbox a b)" - and le: "\x\cbox a b. (f x)\k \ (g x)\k" - for a b i and j :: 'b and f g :: "'a \ 'b" - proof (rule ccontr) - assume "\ ?thesis" - then have *: "0 < (i\k - j\k) / 3" - by auto - guess d1 using f_i[unfolded has_integral,rule_format,OF *] by (elim exE conjE) note d1=this[rule_format] - guess d2 using g_j[unfolded has_integral,rule_format,OF *] by (elim exE conjE) note d2=this[rule_format] - obtain p where p: "p tagged_division_of cbox a b" "d1 fine p" "d2 fine p" - using fine_division_exists[OF gauge_inter[OF d1(1) d2(1)], of a b] unfolding fine_inter - by metis - note le_less_trans[OF Basis_le_norm[OF k]] - then have "\((\(x, k)\p. content k *\<^sub>R f x) - i) \ k\ < (i \ k - j \ k) / 3" - "\((\(x, k)\p. content k *\<^sub>R g x) - j) \ k\ < (i \ k - j \ k) / 3" - using k norm_bound_Basis_lt d1 d2 p - by blast+ - then show False - unfolding inner_simps - using rsum_component_le[OF p(1) le] - by (simp add: abs_real_def split: if_split_asm) - qed - show ?thesis - proof (cases "\a b. s = cbox a b") - case True - with lem assms show ?thesis - by auto - next - case False - show ?thesis - proof (rule ccontr) - assume "\ i\k \ j\k" - then have ij: "(i\k - j\k) / 3 > 0" - by auto - note has_integral_altD[OF _ False this] - from this[OF assms(2)] this[OF assms(3)] guess B1 B2 . note B=this[rule_format] - have "bounded (ball 0 B1 \ ball (0::'a) B2)" - unfolding bounded_Un by(rule conjI bounded_ball)+ - from bounded_subset_cbox[OF this] guess a b by (elim exE) - note ab = conjunctD2[OF this[unfolded Un_subset_iff]] - guess w1 using B(2)[OF ab(1)] .. note w1=conjunctD2[OF this] - guess w2 using B(4)[OF ab(2)] .. note w2=conjunctD2[OF this] - have *: "\w1 w2 j i::real .\w1 - i\ < (i - j) / 3 \ \w2 - j\ < (i - j) / 3 \ w1 \ w2 \ False" - by (simp add: abs_real_def split: if_split_asm) - note le_less_trans[OF Basis_le_norm[OF k]] - note this[OF w1(2)] this[OF w2(2)] - moreover - have "w1\k \ w2\k" - by (rule lem[OF w1(1) w2(1)]) (simp add: assms(4)) - ultimately show False - unfolding inner_simps by(rule *) - qed - qed -qed - -lemma integral_component_le: - fixes g f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "k \ Basis" - and "f integrable_on s" "g integrable_on s" - and "\x\s. (f x)\k \ (g x)\k" - shows "(integral s f)\k \ (integral s g)\k" - apply (rule has_integral_component_le) - using integrable_integral assms - apply auto - done - -lemma has_integral_component_nonneg: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "k \ Basis" - and "(f has_integral i) s" - and "\x\s. 0 \ (f x)\k" - shows "0 \ i\k" - using has_integral_component_le[OF assms(1) has_integral_0 assms(2)] - using assms(3-) - by auto - -lemma integral_component_nonneg: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "k \ Basis" - and "\x\s. 0 \ (f x)\k" - shows "0 \ (integral s f)\k" -proof (cases "f integrable_on s") - case True show ?thesis - apply (rule has_integral_component_nonneg) - using assms True - apply auto - done -next - case False then show ?thesis by (simp add: not_integrable_integral) -qed - -lemma has_integral_component_neg: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "k \ Basis" - and "(f has_integral i) s" - and "\x\s. (f x)\k \ 0" - shows "i\k \ 0" - using has_integral_component_le[OF assms(1,2) has_integral_0] assms(2-) - by auto - -lemma has_integral_component_lbound: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "(f has_integral i) (cbox a b)" - and "\x\cbox a b. B \ f(x)\k" - and "k \ Basis" - shows "B * content (cbox a b) \ i\k" - using has_integral_component_le[OF assms(3) has_integral_const assms(1),of "(\i\Basis. B *\<^sub>R i)::'b"] assms(2-) - by (auto simp add: field_simps) - -lemma has_integral_component_ubound: - fixes f::"'a::euclidean_space => 'b::euclidean_space" - assumes "(f has_integral i) (cbox a b)" - and "\x\cbox a b. f x\k \ B" - and "k \ Basis" - shows "i\k \ B * content (cbox a b)" - using has_integral_component_le[OF assms(3,1) has_integral_const, of "\i\Basis. B *\<^sub>R i"] assms(2-) - by (auto simp add: field_simps) - -lemma integral_component_lbound: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "f integrable_on cbox a b" - and "\x\cbox a b. B \ f(x)\k" - and "k \ Basis" - shows "B * content (cbox a b) \ (integral(cbox a b) f)\k" - apply (rule has_integral_component_lbound) - using assms - unfolding has_integral_integral - apply auto - done - -lemma integral_component_lbound_real: - assumes "f integrable_on {a ::real .. b}" - and "\x\{a .. b}. B \ f(x)\k" - and "k \ Basis" - shows "B * content {a .. b} \ (integral {a .. b} f)\k" - using assms - by (metis box_real(2) integral_component_lbound) - -lemma integral_component_ubound: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "f integrable_on cbox a b" - and "\x\cbox a b. f x\k \ B" - and "k \ Basis" - shows "(integral (cbox a b) f)\k \ B * content (cbox a b)" - apply (rule has_integral_component_ubound) - using assms - unfolding has_integral_integral - apply auto - done - -lemma integral_component_ubound_real: - fixes f :: "real \ 'a::euclidean_space" - assumes "f integrable_on {a .. b}" - and "\x\{a .. b}. f x\k \ B" - and "k \ Basis" - shows "(integral {a .. b} f)\k \ B * content {a .. b}" - using assms - by (metis box_real(2) integral_component_ubound) - -subsection \Uniform limit of integrable functions is integrable.\ - -lemma real_arch_invD: - "0 < (e::real) \ (\n::nat. n \ 0 \ 0 < inverse (real n) \ inverse (real n) < e)" - by (subst(asm) real_arch_inverse) - -lemma integrable_uniform_limit: - fixes f :: "'a::euclidean_space \ 'b::banach" - assumes "\e>0. \g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" - shows "f integrable_on cbox a b" -proof (cases "content (cbox a b) > 0") - case False then show ?thesis - using has_integral_null - by (simp add: content_lt_nz integrable_on_def) -next - case True - have *: "\P. \e>(0::real). P e \ \n::nat. P (inverse (real n + 1))" - by auto - from choice[OF *[OF assms]] guess g .. note g=conjunctD2[OF this[rule_format],rule_format] - from choice[OF allI[OF g(2)[unfolded integrable_on_def], of "\x. x"]] - obtain i where i: "\x. (g x has_integral i x) (cbox a b)" - by auto - have "Cauchy i" - unfolding Cauchy_def - proof clarify - fix e :: real - assume "e>0" - then have "e / 4 / content (cbox a b) > 0" - using True by (auto simp add: field_simps) - then obtain M :: nat - where M: "M \ 0" "0 < inverse (real_of_nat M)" "inverse (of_nat M) < e / 4 / content (cbox a b)" - by (subst (asm) real_arch_inverse) auto - show "\M. \m\M. \n\M. dist (i m) (i n) < e" - proof (rule exI [where x=M], clarify) - fix m n - assume m: "M \ m" and n: "M \ n" - have "e/4>0" using \e>0\ by auto - note * = i[unfolded has_integral,rule_format,OF this] - from *[of m] guess gm by (elim conjE exE) note gm=this[rule_format] - from *[of n] guess gn by (elim conjE exE) note gn=this[rule_format] - from fine_division_exists[OF gauge_inter[OF gm(1) gn(1)], of a b] - obtain p where p: "p tagged_division_of cbox a b" "(\x. gm x \ gn x) fine p" - by auto - { fix s1 s2 i1 and i2::'b - assume no: "norm(s2 - s1) \ e/2" "norm (s1 - i1) < e/4" "norm (s2 - i2) < e/4" - have "norm (i1 - i2) \ norm (i1 - s1) + norm (s1 - s2) + norm (s2 - i2)" - using norm_triangle_ineq[of "i1 - s1" "s1 - i2"] - using norm_triangle_ineq[of "s1 - s2" "s2 - i2"] - by (auto simp add: algebra_simps) - also have "\ < e" - using no - unfolding norm_minus_commute - by (auto simp add: algebra_simps) - finally have "norm (i1 - i2) < e" . - } note triangle3 = this - have finep: "gm fine p" "gn fine p" - using fine_inter p by auto - { fix x - assume x: "x \ cbox a b" - have "norm (f x - g n x) + norm (f x - g m x) \ inverse (real n + 1) + inverse (real m + 1)" - using g(1)[OF x, of n] g(1)[OF x, of m] by auto - also have "\ \ inverse (real M) + inverse (real M)" - apply (rule add_mono) - using M(2) m n by auto - also have "\ = 2 / real M" - unfolding divide_inverse by auto - finally have "norm (g n x - g m x) \ 2 / real M" - using norm_triangle_le[of "g n x - f x" "f x - g m x" "2 / real M"] - by (auto simp add: algebra_simps simp add: norm_minus_commute) - } note norm_le = this - have le_e2: "norm ((\(x, k)\p. content k *\<^sub>R g n x) - (\(x, k)\p. content k *\<^sub>R g m x)) \ e / 2" - apply (rule order_trans [OF rsum_diff_bound[OF p(1), where e="2 / real M"]]) - apply (blast intro: norm_le) - using M True - by (auto simp add: field_simps) - then show "dist (i m) (i n) < e" - unfolding dist_norm - using gm gn p finep - by (auto intro!: triangle3) - qed - qed - then obtain s where s: "i \ s" - using convergent_eq_cauchy[symmetric] by blast - show ?thesis - unfolding integrable_on_def has_integral - proof (rule_tac x=s in exI, clarify) - fix e::real - assume e: "0 < e" - then have *: "e/3 > 0" by auto - then obtain N1 where N1: "\n\N1. norm (i n - s) < e / 3" - using LIMSEQ_D [OF s] by metis - from e True have "e / 3 / content (cbox a b) > 0" - by (auto simp add: field_simps) - from real_arch_invD[OF this] guess N2 by (elim exE conjE) note N2=this - from i[of "N1 + N2",unfolded has_integral,rule_format,OF *] guess g' .. note g'=conjunctD2[OF this,rule_format] - { fix sf sg i - assume no: "norm (sf - sg) \ e / 3" - "norm(i - s) < e / 3" - "norm (sg - i) < e / 3" - have "norm (sf - s) \ norm (sf - sg) + norm (sg - i) + norm (i - s)" - using norm_triangle_ineq[of "sf - sg" "sg - s"] - using norm_triangle_ineq[of "sg - i" " i - s"] - by (auto simp add: algebra_simps) - also have "\ < e" - using no - unfolding norm_minus_commute - by (auto simp add: algebra_simps) - finally have "norm (sf - s) < e" . - } note lem = this - { fix p - assume p: "p tagged_division_of (cbox a b) \ g' fine p" - then have norm_less: "norm ((\(x, k)\p. content k *\<^sub>R g (N1 + N2) x) - i (N1 + N2)) < e / 3" - using g' by blast - have "content (cbox a b) < e / 3 * (of_nat N2)" - using N2 unfolding inverse_eq_divide using True by (auto simp add: field_simps) - moreover have "e / 3 * of_nat N2 \ e / 3 * (of_nat (N1 + N2) + 1)" - using \e>0\ by auto - ultimately have "content (cbox a b) < e / 3 * (of_nat (N1 + N2) + 1)" - by linarith - then have le_e3: "inverse (real (N1 + N2) + 1) * content (cbox a b) \ e / 3" - unfolding inverse_eq_divide - by (auto simp add: field_simps) - have ne3: "norm (i (N1 + N2) - s) < e / 3" - using N1 by auto - have "norm ((\(x, k)\p. content k *\<^sub>R f x) - s) < e" - apply (rule lem[OF order_trans [OF _ le_e3] ne3 norm_less]) - apply (rule rsum_diff_bound[OF p[THEN conjunct1]]) - apply (blast intro: g) - done } - then show "\d. gauge d \ - (\p. p tagged_division_of cbox a b \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x) - s) < e)" - by (blast intro: g') - qed -qed - -lemmas integrable_uniform_limit_real = integrable_uniform_limit [where 'a=real, simplified] - - -subsection \Negligible sets.\ - -definition "negligible (s:: 'a::euclidean_space set) \ - (\a b. ((indicator s :: 'a\real) has_integral 0) (cbox a b))" - - -subsection \Negligibility of hyperplane.\ - -lemma interval_doublesplit: - fixes a :: "'a::euclidean_space" - assumes "k \ Basis" - shows "cbox a b \ {x . \x\k - c\ \ (e::real)} = - cbox (\i\Basis. (if i = k then max (a\k) (c - e) else a\i) *\<^sub>R i) - (\i\Basis. (if i = k then min (b\k) (c + e) else b\i) *\<^sub>R i)" -proof - - have *: "\x c e::real. \x - c\ \ e \ x \ c - e \ x \ c + e" - by auto - have **: "\s P Q. s \ {x. P x \ Q x} = (s \ {x. Q x}) \ {x. P x}" - by blast - show ?thesis - unfolding * ** interval_split[OF assms] by (rule refl) -qed - -lemma division_doublesplit: - fixes a :: "'a::euclidean_space" - assumes "p division_of (cbox a b)" - and k: "k \ Basis" - shows "(\l. l \ {x. \x\k - c\ \ e}) ` {l\p. l \ {x. \x\k - c\ \ e} \ {}} - division_of (cbox a b \ {x. \x\k - c\ \ e})" -proof - - have *: "\x c. \x - c\ \ e \ x \ c - e \ x \ c + e" - by auto - have **: "\p q p' q'. p division_of q \ p = p' \ q = q' \ p' division_of q'" - by auto - note division_split(1)[OF assms, where c="c+e",unfolded interval_split[OF k]] - note division_split(2)[OF this, where c="c-e" and k=k,OF k] - then show ?thesis - apply (rule **) - subgoal - apply (simp add: abs_diff_le_iff field_simps Collect_conj_eq setcompr_eq_image[symmetric]) - apply (rule equalityI) - apply blast - apply clarsimp - apply (rule_tac x="l \ {x. c + e \ x \ k}" in exI) - apply auto - done - by (simp add: interval_split k interval_doublesplit) -qed - -lemma content_doublesplit: - fixes a :: "'a::euclidean_space" - assumes "0 < e" - and k: "k \ Basis" - obtains d where "0 < d" and "content (cbox a b \ {x. \x\k - c\ \ d}) < e" -proof (cases "content (cbox a b) = 0") - case True - then have ce: "content (cbox a b) < e" - by (metis \0 < e\) - show ?thesis - apply (rule that[of 1]) - apply simp - unfolding interval_doublesplit[OF k] - apply (rule le_less_trans[OF content_subset ce]) - apply (auto simp: interval_doublesplit[symmetric] k) - done -next - case False - define d where "d = e / 3 / setprod (\i. b\i - a\i) (Basis - {k})" - note False[unfolded content_eq_0 not_ex not_le, rule_format] - then have "\x. x \ Basis \ b\x > a\x" - by (auto simp add:not_le) - then have prod0: "0 < setprod (\i. b\i - a\i) (Basis - {k})" - by (force simp add: setprod_pos field_simps) - then have "d > 0" - using assms - by (auto simp add: d_def field_simps) - then show ?thesis - proof (rule that[of d]) - have *: "Basis = insert k (Basis - {k})" - using k by auto - have less_e: "(min (b \ k) (c + d) - max (a \ k) (c - d)) * (\i\Basis - {k}. b \ i - a \ i) < e" - proof - - have "(min (b \ k) (c + d) - max (a \ k) (c - d)) \ 2 * d" - by auto - also have "\ < e / (\i\Basis - {k}. b \ i - a \ i)" - unfolding d_def - using assms prod0 - by (auto simp add: field_simps) - finally show "(min (b \ k) (c + d) - max (a \ k) (c - d)) * (\i\Basis - {k}. b \ i - a \ i) < e" - unfolding pos_less_divide_eq[OF prod0] . - qed - show "content (cbox a b \ {x. \x \ k - c\ \ d}) < e" - proof (cases "cbox a b \ {x. \x \ k - c\ \ d} = {}") - case True - then show ?thesis - using assms by simp - next - case False - then have - "(\i\Basis - {k}. interval_upperbound (cbox a b \ {x. \x \ k - c\ \ d}) \ i - - interval_lowerbound (cbox a b \ {x. \x \ k - c\ \ d}) \ i) - = (\i\Basis - {k}. b\i - a\i)" - by (simp add: box_eq_empty interval_doublesplit[OF k]) - then show "content (cbox a b \ {x. \x \ k - c\ \ d}) < e" - unfolding content_def - using assms False - apply (subst *) - apply (subst setprod.insert) - apply (simp_all add: interval_doublesplit[OF k] box_eq_empty not_less less_e) - done - qed - qed -qed - -lemma negligible_standard_hyperplane[intro]: - fixes k :: "'a::euclidean_space" - assumes k: "k \ Basis" - shows "negligible {x. x\k = c}" - unfolding negligible_def has_integral -proof (clarify, goal_cases) - case (1 a b e) - from this and k obtain d where d: "0 < d" "content (cbox a b \ {x. \x \ k - c\ \ d}) < e" - by (rule content_doublesplit) - let ?i = "indicator {x::'a. x\k = c} :: 'a\real" - show ?case - apply (rule_tac x="\x. ball x d" in exI) - apply rule - apply (rule gauge_ball) - apply (rule d) - proof (rule, rule) - fix p - assume p: "p tagged_division_of (cbox a b) \ (\x. ball x d) fine p" - have *: "(\(x, ka)\p. content ka *\<^sub>R ?i x) = - (\(x, ka)\p. content (ka \ {x. \x\k - c\ \ d}) *\<^sub>R ?i x)" - apply (rule setsum.cong) - apply (rule refl) - unfolding split_paired_all real_scaleR_def mult_cancel_right split_conv - apply cases - apply (rule disjI1) - apply assumption - apply (rule disjI2) - proof - - fix x l - assume as: "(x, l) \ p" "?i x \ 0" - then have xk: "x\k = c" - unfolding indicator_def - apply - - apply (rule ccontr) - apply auto - done - show "content l = content (l \ {x. \x \ k - c\ \ d})" - apply (rule arg_cong[where f=content]) - apply (rule set_eqI) - apply rule - apply rule - unfolding mem_Collect_eq - proof - - fix y - assume y: "y \ l" - note p[THEN conjunct2,unfolded fine_def,rule_format,OF as(1),unfolded split_conv] - note this[unfolded subset_eq mem_ball dist_norm,rule_format,OF y] - note le_less_trans[OF Basis_le_norm[OF k] this] - then show "\y \ k - c\ \ d" - unfolding inner_simps xk by auto - qed auto - qed - note p'= tagged_division_ofD[OF p[THEN conjunct1]] and p''=division_of_tagged_division[OF p[THEN conjunct1]] - show "norm ((\(x, ka)\p. content ka *\<^sub>R ?i x) - 0) < e" - unfolding diff_0_right * - unfolding real_scaleR_def real_norm_def - apply (subst abs_of_nonneg) - apply (rule setsum_nonneg) - apply rule - unfolding split_paired_all split_conv - apply (rule mult_nonneg_nonneg) - apply (drule p'(4)) - apply (erule exE)+ - apply(rule_tac b=b in back_subst) - prefer 2 - apply (subst(asm) eq_commute) - apply assumption - apply (subst interval_doublesplit[OF k]) - apply (rule content_pos_le) - apply (rule indicator_pos_le) - proof - - have "(\(x, ka)\p. content (ka \ {x. \x \ k - c\ \ d}) * ?i x) \ - (\(x, ka)\p. content (ka \ {x. \x \ k - c\ \ d}))" - apply (rule setsum_mono) - unfolding split_paired_all split_conv - apply (rule mult_right_le_one_le) - apply (drule p'(4)) - apply (auto simp add:interval_doublesplit[OF k]) - done - also have "\ < e" - proof (subst setsum.over_tagged_division_lemma[OF p[THEN conjunct1]], goal_cases) - case prems: (1 u v) - have "content (cbox u v \ {x. \x \ k - c\ \ d}) \ content (cbox u v)" - unfolding interval_doublesplit[OF k] - apply (rule content_subset) - unfolding interval_doublesplit[symmetric,OF k] - apply auto - done - then show ?case - unfolding prems interval_doublesplit[OF k] - by (blast intro: antisym) - next - have "(\l\snd ` p. content (l \ {x. \x \ k - c\ \ d})) = - setsum content ((\l. l \ {x. \x \ k - c\ \ d})`{l\snd ` p. l \ {x. \x \ k - c\ \ d} \ {}})" - proof (subst (2) setsum.reindex_nontrivial) - fix x y assume "x \ {l \ snd ` p. l \ {x. \x \ k - c\ \ d} \ {}}" "y \ {l \ snd ` p. l \ {x. \x \ k - c\ \ d} \ {}}" - "x \ y" and eq: "x \ {x. \x \ k - c\ \ d} = y \ {x. \x \ k - c\ \ d}" - then obtain x' y' where "(x', x) \ p" "x \ {x. \x \ k - c\ \ d} \ {}" "(y', y) \ p" "y \ {x. \x \ k - c\ \ d} \ {}" - by (auto) - from p'(5)[OF \(x', x) \ p\ \(y', y) \ p\] \x \ y\ have "interior (x \ y) = {}" - by auto - moreover have "interior ((x \ {x. \x \ k - c\ \ d}) \ (y \ {x. \x \ k - c\ \ d})) \ interior (x \ y)" - by (auto intro: interior_mono) - ultimately have "interior (x \ {x. \x \ k - c\ \ d}) = {}" - by (auto simp: eq) - then show "content (x \ {x. \x \ k - c\ \ d}) = 0" - using p'(4)[OF \(x', x) \ p\] by (auto simp: interval_doublesplit[OF k] content_eq_0_interior simp del: interior_Int) - qed (insert p'(1), auto intro!: setsum.mono_neutral_right) - also have "\ \ norm (\l\(\l. l \ {x. \x \ k - c\ \ d})`{l\snd ` p. l \ {x. \x \ k - c\ \ d} \ {}}. content l *\<^sub>R 1::real)" - by simp - also have "\ \ 1 * content (cbox a b \ {x. \x \ k - c\ \ d})" - using division_doublesplit[OF p'' k, unfolded interval_doublesplit[OF k]] - unfolding interval_doublesplit[OF k] by (intro dsum_bound) auto - also have "\ < e" - using d(2) by simp - finally show "(\ka\snd ` p. content (ka \ {x. \x \ k - c\ \ d})) < e" . - qed - finally show "(\(x, ka)\p. content (ka \ {x. \x \ k - c\ \ d}) * ?i x) < e" . - qed - qed -qed - - -subsection \A technical lemma about "refinement" of division.\ - -lemma tagged_division_finer: - fixes p :: "('a::euclidean_space \ ('a::euclidean_space set)) set" - assumes "p tagged_division_of (cbox a b)" - and "gauge d" - obtains q where "q tagged_division_of (cbox a b)" - and "d fine q" - and "\(x,k) \ p. k \ d(x) \ (x,k) \ q" -proof - - let ?P = "\p. p tagged_partial_division_of (cbox a b) \ gauge d \ - (\q. q tagged_division_of (\{k. \x. (x,k) \ p}) \ d fine q \ - (\(x,k) \ p. k \ d(x) \ (x,k) \ q))" - { - have *: "finite p" "p tagged_partial_division_of (cbox a b)" - using assms(1) - unfolding tagged_division_of_def - by auto - presume "\p. finite p \ ?P p" - from this[rule_format,OF * assms(2)] guess q .. note q=this - then show ?thesis - apply - - apply (rule that[of q]) - unfolding tagged_division_ofD[OF assms(1)] - apply auto - done - } - fix p :: "('a::euclidean_space \ ('a::euclidean_space set)) set" - assume as: "finite p" - show "?P p" - apply rule - apply rule - using as - proof (induct p) - case empty - show ?case - apply (rule_tac x="{}" in exI) - unfolding fine_def - apply auto - done - next - case (insert xk p) - guess x k using surj_pair[of xk] by (elim exE) note xk=this - note tagged_partial_division_subset[OF insert(4) subset_insertI] - from insert(3)[OF this insert(5)] guess q1 .. note q1 = conjunctD3[OF this] - have *: "\{l. \y. (y,l) \ insert xk p} = k \ \{l. \y. (y,l) \ p}" - unfolding xk by auto - note p = tagged_partial_division_ofD[OF insert(4)] - from p(4)[unfolded xk, OF insertI1] guess u v by (elim exE) note uv=this - - have "finite {k. \x. (x, k) \ p}" - apply (rule finite_subset[of _ "snd ` p"]) - using p - apply safe - apply (metis image_iff snd_conv) - apply auto - done - then have int: "interior (cbox u v) \ interior (\{k. \x. (x, k) \ p}) = {}" - apply (rule inter_interior_unions_intervals) - apply (rule open_interior) - apply (rule_tac[!] ballI) - unfolding mem_Collect_eq - apply (erule_tac[!] exE) - apply (drule p(4)[OF insertI2]) - apply assumption - apply (rule p(5)) - unfolding uv xk - apply (rule insertI1) - apply (rule insertI2) - apply assumption - using insert(2) - unfolding uv xk - apply auto - done - show ?case - proof (cases "cbox u v \ d x") - case True - then show ?thesis - apply (rule_tac x="{(x,cbox u v)} \ q1" in exI) - apply rule - unfolding * uv - apply (rule tagged_division_union) - apply (rule tagged_division_of_self) - apply (rule p[unfolded xk uv] insertI1)+ - apply (rule q1) - apply (rule int) - apply rule - apply (rule fine_union) - apply (subst fine_def) - defer - apply (rule q1) - unfolding Ball_def split_paired_All split_conv - apply rule - apply rule - apply rule - apply rule - apply (erule insertE) - apply (simp add: uv xk) - apply (rule UnI2) - apply (drule q1(3)[rule_format]) - unfolding xk uv - apply auto - done - next - case False - from fine_division_exists[OF assms(2), of u v] guess q2 . note q2=this - show ?thesis - apply (rule_tac x="q2 \ q1" in exI) - apply rule - unfolding * uv - apply (rule tagged_division_union q2 q1 int fine_union)+ - unfolding Ball_def split_paired_All split_conv - apply rule - apply (rule fine_union) - apply (rule q1 q2)+ - apply rule - apply rule - apply rule - apply rule - apply (erule insertE) - apply (rule UnI2) - apply (simp add: False uv xk) - apply (drule q1(3)[rule_format]) - using False - unfolding xk uv - apply auto - done - qed - qed -qed - - -subsection \Hence the main theorem about negligible sets.\ - -lemma finite_product_dependent: - assumes "finite s" - and "\x. x \ s \ finite (t x)" - shows "finite {(i, j) |i j. i \ s \ j \ t i}" - using assms -proof induct - case (insert x s) - have *: "{(i, j) |i j. i \ insert x s \ j \ t i} = - (\y. (x,y)) ` (t x) \ {(i, j) |i j. i \ s \ j \ t i}" by auto - show ?case - unfolding * - apply (rule finite_UnI) - using insert - apply auto - done -qed auto - -lemma sum_sum_product: - assumes "finite s" - and "\i\s. finite (t i)" - shows "setsum (\i. setsum (x i) (t i)::real) s = - setsum (\(i,j). x i j) {(i,j) | i j. i \ s \ j \ t i}" - using assms -proof induct - case (insert a s) - have *: "{(i, j) |i j. i \ insert a s \ j \ t i} = - (\y. (a,y)) ` (t a) \ {(i, j) |i j. i \ s \ j \ t i}" by auto - show ?case - unfolding * - apply (subst setsum.union_disjoint) - unfolding setsum.insert[OF insert(1-2)] - prefer 4 - apply (subst insert(3)) - unfolding add_right_cancel - proof - - show "setsum (x a) (t a) = (\(xa, y)\ Pair a ` t a. x xa y)" - apply (subst setsum.reindex) - unfolding inj_on_def - apply auto - done - show "finite {(i, j) |i j. i \ s \ j \ t i}" - apply (rule finite_product_dependent) - using insert - apply auto - done - qed (insert insert, auto) -qed auto - -lemma has_integral_negligible: - fixes f :: "'b::euclidean_space \ 'a::real_normed_vector" - assumes "negligible s" - and "\x\(t - s). f x = 0" - shows "(f has_integral 0) t" -proof - - presume P: "\f::'b::euclidean_space \ 'a. - \a b. \x. x \ s \ f x = 0 \ (f has_integral 0) (cbox a b)" - let ?f = "(\x. if x \ t then f x else 0)" - show ?thesis - apply (rule_tac f="?f" in has_integral_eq) - unfolding if_P - apply (rule refl) - apply (subst has_integral_alt) - apply cases - apply (subst if_P, assumption) - unfolding if_not_P - proof - - assume "\a b. t = cbox a b" - then guess a b apply - by (erule exE)+ note t = this - show "(?f has_integral 0) t" - unfolding t - apply (rule P) - using assms(2) - unfolding t - apply auto - done - next - show "\e>0. \B>0. \a b. ball 0 B \ cbox a b \ - (\z. ((\x. if x \ t then ?f x else 0) has_integral z) (cbox a b) \ norm (z - 0) < e)" - apply safe - apply (rule_tac x=1 in exI) - apply rule - apply (rule zero_less_one) - apply safe - apply (rule_tac x=0 in exI) - apply rule - apply (rule P) - using assms(2) - apply auto - done - qed -next - fix f :: "'b \ 'a" - fix a b :: 'b - assume assm: "\x. x \ s \ f x = 0" - show "(f has_integral 0) (cbox a b)" - unfolding has_integral - proof (safe, goal_cases) - case prems: (1 e) - then have "\n. e / 2 / ((real n+1) * (2 ^ n)) > 0" - apply - - apply (rule divide_pos_pos) - defer - apply (rule mult_pos_pos) - apply (auto simp add:field_simps) - done - note assms(1)[unfolded negligible_def has_integral,rule_format,OF this,of a b] - note allI[OF this,of "\x. x"] - from choice[OF this] guess d .. note d=conjunctD2[OF this[rule_format]] - show ?case - apply (rule_tac x="\x. d (nat \norm (f x)\) x" in exI) - proof safe - show "gauge (\x. d (nat \norm (f x)\) x)" - using d(1) unfolding gauge_def by auto - fix p - assume as: "p tagged_division_of (cbox a b)" "(\x. d (nat \norm (f x)\) x) fine p" - let ?goal = "norm ((\(x, k)\p. content k *\<^sub>R f x) - 0) < e" - { - presume "p \ {} \ ?goal" - then show ?goal - apply (cases "p = {}") - using prems - apply auto - done - } - assume as': "p \ {}" - from real_arch_simple[of "Max((\(x,k). norm(f x)) ` p)"] guess N .. - then have N: "\x\(\(x, k). norm (f x)) ` p. x \ real N" - by (meson Max_ge as(1) dual_order.trans finite_imageI tagged_division_of_finite) - have "\i. \q. q tagged_division_of (cbox a b) \ (d i) fine q \ (\(x, k)\p. k \ (d i) x \ (x, k) \ q)" - by (auto intro: tagged_division_finer[OF as(1) d(1)]) - from choice[OF this] guess q .. note q=conjunctD3[OF this[rule_format]] - have *: "\i. (\(x, k)\q i. content k *\<^sub>R indicator s x) \ (0::real)" - apply (rule setsum_nonneg) - apply safe - unfolding real_scaleR_def - apply (drule tagged_division_ofD(4)[OF q(1)]) - apply (auto intro: mult_nonneg_nonneg) - done - have **: "finite s \ finite t \ (\(x,y) \ t. (0::real) \ g(x,y)) \ - (\y\s. \x. (x,y) \ t \ f(y) \ g(x,y)) \ setsum f s \ setsum g t" for f g s t - apply (rule setsum_le_included[of s t g snd f]) - prefer 4 - apply safe - apply (erule_tac x=x in ballE) - apply (erule exE) - apply (rule_tac x="(xa,x)" in bexI) - apply auto - done - have "norm ((\(x, k)\p. content k *\<^sub>R f x) - 0) \ setsum (\i. (real i + 1) * - norm (setsum (\(x,k). content k *\<^sub>R indicator s x :: real) (q i))) {..N+1}" - unfolding real_norm_def setsum_right_distrib abs_of_nonneg[OF *] diff_0_right - apply (rule order_trans) - apply (rule norm_setsum) - apply (subst sum_sum_product) - prefer 3 - proof (rule **, safe) - show "finite {(i, j) |i j. i \ {..N + 1} \ j \ q i}" - apply (rule finite_product_dependent) - using q - apply auto - done - fix i a b - assume as'': "(a, b) \ q i" - show "0 \ (real i + 1) * (content b *\<^sub>R indicator s a)" - unfolding real_scaleR_def - using tagged_division_ofD(4)[OF q(1) as''] - by (auto intro!: mult_nonneg_nonneg) - next - fix i :: nat - show "finite (q i)" - using q by auto - next - fix x k - assume xk: "(x, k) \ p" - define n where "n = nat \norm (f x)\" - have *: "norm (f x) \ (\(x, k). norm (f x)) ` p" - using xk by auto - have nfx: "real n \ norm (f x)" "norm (f x) \ real n + 1" - unfolding n_def by auto - then have "n \ {0..N + 1}" - using N[rule_format,OF *] by auto - moreover - note as(2)[unfolded fine_def,rule_format,OF xk,unfolded split_conv] - note q(3)[rule_format,OF xk,unfolded split_conv,rule_format,OF this] - note this[unfolded n_def[symmetric]] - moreover - have "norm (content k *\<^sub>R f x) \ (real n + 1) * (content k * indicator s x)" - proof (cases "x \ s") - case False - then show ?thesis - using assm by auto - next - case True - have *: "content k \ 0" - using tagged_division_ofD(4)[OF as(1) xk] by auto - moreover - have "content k * norm (f x) \ content k * (real n + 1)" - apply (rule mult_mono) - using nfx * - apply auto - done - ultimately - show ?thesis - unfolding abs_mult - using nfx True - by (auto simp add: field_simps) - qed - ultimately show "\y. (y, x, k) \ {(i, j) |i j. i \ {..N + 1} \ j \ q i} \ norm (content k *\<^sub>R f x) \ - (real y + 1) * (content k *\<^sub>R indicator s x)" - apply (rule_tac x=n in exI) - apply safe - apply (rule_tac x=n in exI) - apply (rule_tac x="(x,k)" in exI) - apply safe - apply auto - done - qed (insert as, auto) - also have "\ \ setsum (\i. e / 2 / 2 ^ i) {..N+1}" - proof (rule setsum_mono, goal_cases) - case (1 i) - then show ?case - apply (subst mult.commute, subst pos_le_divide_eq[symmetric]) - using d(2)[rule_format, of "q i" i] - using q[rule_format] - apply (auto simp add: field_simps) - done - qed - also have "\ < e * inverse 2 * 2" - unfolding divide_inverse setsum_right_distrib[symmetric] - apply (rule mult_strict_left_mono) - unfolding power_inverse [symmetric] lessThan_Suc_atMost[symmetric] - apply (subst geometric_sum) - using prems - apply auto - done - finally show "?goal" by auto - qed - qed -qed - -lemma has_integral_spike: - fixes f :: "'b::euclidean_space \ 'a::real_normed_vector" - assumes "negligible s" - and "(\x\(t - s). g x = f x)" - and "(f has_integral y) t" - shows "(g has_integral y) t" -proof - - { - fix a b :: 'b - fix f g :: "'b \ 'a" - fix y :: 'a - assume as: "\x \ cbox a b - s. g x = f x" "(f has_integral y) (cbox a b)" - have "((\x. f x + (g x - f x)) has_integral (y + 0)) (cbox a b)" - apply (rule has_integral_add[OF as(2)]) - apply (rule has_integral_negligible[OF assms(1)]) - using as - apply auto - done - then have "(g has_integral y) (cbox a b)" - by auto - } note * = this - show ?thesis - apply (subst has_integral_alt) - using assms(2-) - apply - - apply (rule cond_cases) - apply safe - apply (rule *) - apply assumption+ - apply (subst(asm) has_integral_alt) - unfolding if_not_P - apply (erule_tac x=e in allE) - apply safe - apply (rule_tac x=B in exI) - apply safe - apply (erule_tac x=a in allE) - apply (erule_tac x=b in allE) - apply safe - apply (rule_tac x=z in exI) - apply safe - apply (rule *[where fa2="\x. if x\t then f x else 0"]) - apply auto - done -qed - -lemma has_integral_spike_eq: - assumes "negligible s" - and "\x\(t - s). g x = f x" - shows "((f has_integral y) t \ (g has_integral y) t)" - apply rule - apply (rule_tac[!] has_integral_spike[OF assms(1)]) - using assms(2) - apply auto - done - -lemma integrable_spike: - assumes "negligible s" - and "\x\(t - s). g x = f x" - and "f integrable_on t" - shows "g integrable_on t" - using assms - unfolding integrable_on_def - apply - - apply (erule exE) - apply rule - apply (rule has_integral_spike) - apply fastforce+ - done - -lemma integral_spike: - assumes "negligible s" - and "\x\(t - s). g x = f x" - shows "integral t f = integral t g" - using has_integral_spike_eq[OF assms] by (simp add: integral_def integrable_on_def) - - -subsection \Some other trivialities about negligible sets.\ - -lemma negligible_subset[intro]: - assumes "negligible s" - and "t \ s" - shows "negligible t" - unfolding negligible_def -proof (safe, goal_cases) - case (1 a b) - show ?case - using assms(1)[unfolded negligible_def,rule_format,of a b] - apply - - apply (rule has_integral_spike[OF assms(1)]) - defer - apply assumption - using assms(2) - unfolding indicator_def - apply auto - done -qed - -lemma negligible_diff[intro?]: - assumes "negligible s" - shows "negligible (s - t)" - using assms by auto - -lemma negligible_Int: - assumes "negligible s \ negligible t" - shows "negligible (s \ t)" - using assms by auto - -lemma negligible_Un: - assumes "negligible s" - and "negligible t" - shows "negligible (s \ t)" - unfolding negligible_def -proof (safe, goal_cases) - case (1 a b) - note assm = assms[unfolded negligible_def,rule_format,of a b] - then show ?case - apply (subst has_integral_spike_eq[OF assms(2)]) - defer - apply assumption - unfolding indicator_def - apply auto - done -qed - -lemma negligible_Un_eq[simp]: "negligible (s \ t) \ negligible s \ negligible t" - using negligible_Un by auto - -lemma negligible_sing[intro]: "negligible {a::'a::euclidean_space}" - using negligible_standard_hyperplane[OF SOME_Basis, of "a \ (SOME i. i \ Basis)"] by auto - -lemma negligible_insert[simp]: "negligible (insert a s) \ negligible s" - apply (subst insert_is_Un) - unfolding negligible_Un_eq - apply auto - done - -lemma negligible_empty[iff]: "negligible {}" - by auto - -lemma negligible_finite[intro]: - assumes "finite s" - shows "negligible s" - using assms by (induct s) auto - -lemma negligible_Union[intro]: - assumes "finite s" - and "\t\s. negligible t" - shows "negligible(\s)" - using assms by induct auto - -lemma negligible: - "negligible s \ (\t::('a::euclidean_space) set. ((indicator s::'a\real) has_integral 0) t)" - apply safe - defer - apply (subst negligible_def) -proof - - fix t :: "'a set" - assume as: "negligible s" - have *: "(\x. if x \ s \ t then 1 else 0) = (\x. if x\t then if x\s then 1 else 0 else 0)" - by auto - show "((indicator s::'a\real) has_integral 0) t" - apply (subst has_integral_alt) - apply cases - apply (subst if_P,assumption) - unfolding if_not_P - apply safe - apply (rule as[unfolded negligible_def,rule_format]) - apply (rule_tac x=1 in exI) - apply safe - apply (rule zero_less_one) - apply (rule_tac x=0 in exI) - using negligible_subset[OF as,of "s \ t"] - unfolding negligible_def indicator_def [abs_def] - unfolding * - apply auto - done -qed auto - - -subsection \Finite case of the spike theorem is quite commonly needed.\ - -lemma has_integral_spike_finite: - assumes "finite s" - and "\x\t-s. g x = f x" - and "(f has_integral y) t" - shows "(g has_integral y) t" - apply (rule has_integral_spike) - using assms - apply auto - done - -lemma has_integral_spike_finite_eq: - assumes "finite s" - and "\x\t-s. g x = f x" - shows "((f has_integral y) t \ (g has_integral y) t)" - apply rule - apply (rule_tac[!] has_integral_spike_finite) - using assms - apply auto - done - -lemma integrable_spike_finite: - assumes "finite s" - and "\x\t-s. g x = f x" - and "f integrable_on t" - shows "g integrable_on t" - using assms - unfolding integrable_on_def - apply safe - apply (rule_tac x=y in exI) - apply (rule has_integral_spike_finite) - apply auto - done - - -subsection \In particular, the boundary of an interval is negligible.\ - -lemma negligible_frontier_interval: "negligible(cbox (a::'a::euclidean_space) b - box a b)" -proof - - let ?A = "\((\k. {x. x\k = a\k} \ {x::'a. x\k = b\k}) ` Basis)" - have "cbox a b - box a b \ ?A" - apply rule unfolding Diff_iff mem_box - apply simp - apply(erule conjE bexE)+ - apply(rule_tac x=i in bexI) - apply auto - done - then show ?thesis - apply - - apply (rule negligible_subset[of ?A]) - apply (rule negligible_Union[OF finite_imageI]) - apply auto - done -qed - -lemma has_integral_spike_interior: - assumes "\x\box a b. g x = f x" - and "(f has_integral y) (cbox a b)" - shows "(g has_integral y) (cbox a b)" - apply (rule has_integral_spike[OF negligible_frontier_interval _ assms(2)]) - using assms(1) - apply auto - done - -lemma has_integral_spike_interior_eq: - assumes "\x\box a b. g x = f x" - shows "(f has_integral y) (cbox a b) \ (g has_integral y) (cbox a b)" - apply rule - apply (rule_tac[!] has_integral_spike_interior) - using assms - apply auto - done - -lemma integrable_spike_interior: - assumes "\x\box a b. g x = f x" - and "f integrable_on cbox a b" - shows "g integrable_on cbox a b" - using assms - unfolding integrable_on_def - using has_integral_spike_interior[OF assms(1)] - by auto - - -subsection \Integrability of continuous functions.\ - -lemma operative_approximable: - fixes f :: "'b::euclidean_space \ 'a::banach" - assumes "0 \ e" - shows "comm_monoid.operative op \ True (\i. \g. (\x\i. norm (f x - g (x::'b)) \ e) \ g integrable_on i)" - unfolding comm_monoid.operative_def[OF comm_monoid_and] -proof safe - fix a b :: 'b - show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" - if "content (cbox a b) = 0" - apply (rule_tac x=f in exI) - using assms that - apply (auto intro!: integrable_on_null) - done - { - fix c g - fix k :: 'b - assume as: "\x\cbox a b. norm (f x - g x) \ e" "g integrable_on cbox a b" - assume k: "k \ Basis" - show "\g. (\x\cbox a b \ {x. x \ k \ c}. norm (f x - g x) \ e) \ g integrable_on cbox a b \ {x. x \ k \ c}" - "\g. (\x\cbox a b \ {x. c \ x \ k}. norm (f x - g x) \ e) \ g integrable_on cbox a b \ {x. c \ x \ k}" - apply (rule_tac[!] x=g in exI) - using as(1) integrable_split[OF as(2) k] - apply auto - done - } - fix c k g1 g2 - assume as: "\x\cbox a b \ {x. x \ k \ c}. norm (f x - g1 x) \ e" "g1 integrable_on cbox a b \ {x. x \ k \ c}" - "\x\cbox a b \ {x. c \ x \ k}. norm (f x - g2 x) \ e" "g2 integrable_on cbox a b \ {x. c \ x \ k}" - assume k: "k \ Basis" - let ?g = "\x. if x\k = c then f x else if x\k \ c then g1 x else g2 x" - show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" - apply (rule_tac x="?g" in exI) - apply safe - proof goal_cases - case (1 x) - then show ?case - apply - - apply (cases "x\k=c") - apply (case_tac "x\k < c") - using as assms - apply auto - done - next - case 2 - presume "?g integrable_on cbox a b \ {x. x \ k \ c}" - and "?g integrable_on cbox a b \ {x. x \ k \ c}" - then guess h1 h2 unfolding integrable_on_def by auto - from has_integral_split[OF this k] show ?case - unfolding integrable_on_def by auto - next - show "?g integrable_on cbox a b \ {x. x \ k \ c}" "?g integrable_on cbox a b \ {x. x \ k \ c}" - apply(rule_tac[!] integrable_spike[OF negligible_standard_hyperplane[of k c]]) - using k as(2,4) - apply auto - done - qed -qed - -lemma comm_monoid_set_F_and: "comm_monoid_set.F op \ True f s \ (finite s \ (\x\s. f x))" -proof - - interpret bool: comm_monoid_set "op \" True - proof qed auto - show ?thesis - by (induction s rule: infinite_finite_induct) auto -qed - -lemma approximable_on_division: - fixes f :: "'b::euclidean_space \ 'a::banach" - assumes "0 \ e" - and "d division_of (cbox a b)" - and "\i\d. \g. (\x\i. norm (f x - g x) \ e) \ g integrable_on i" - obtains g where "\x\cbox a b. norm (f x - g x) \ e" "g integrable_on cbox a b" -proof - - note * = comm_monoid_set.operative_division[OF comm_monoid_set_and operative_approximable[OF assms(1)] assms(2)] - from assms(3) this[unfolded comm_monoid_set_F_and, of f] division_of_finite[OF assms(2)] - guess g by auto - then show thesis - apply - - apply (rule that[of g]) - apply auto - done -qed - -lemma integrable_continuous: - fixes f :: "'b::euclidean_space \ 'a::banach" - assumes "continuous_on (cbox a b) f" - shows "f integrable_on cbox a b" -proof (rule integrable_uniform_limit, safe) - fix e :: real - assume e: "e > 0" - from compact_uniformly_continuous[OF assms compact_cbox,unfolded uniformly_continuous_on_def,rule_format,OF e] guess d .. - note d=conjunctD2[OF this,rule_format] - from fine_division_exists[OF gauge_ball[OF d(1)], of a b] guess p . note p=this - note p' = tagged_division_ofD[OF p(1)] - have *: "\i\snd ` p. \g. (\x\i. norm (f x - g x) \ e) \ g integrable_on i" - proof (safe, unfold snd_conv) - fix x l - assume as: "(x, l) \ p" - from p'(4)[OF this] guess a b by (elim exE) note l=this - show "\g. (\x\l. norm (f x - g x) \ e) \ g integrable_on l" - apply (rule_tac x="\y. f x" in exI) - proof safe - show "(\y. f x) integrable_on l" - unfolding integrable_on_def l - apply rule - apply (rule has_integral_const) - done - fix y - assume y: "y \ l" - note fineD[OF p(2) as,unfolded subset_eq,rule_format,OF this] - note d(2)[OF _ _ this[unfolded mem_ball]] - then show "norm (f y - f x) \ e" - using y p'(2-3)[OF as] unfolding dist_norm l norm_minus_commute by fastforce - qed - qed - from e have "e \ 0" - by auto - from approximable_on_division[OF this division_of_tagged_division[OF p(1)] *] guess g . - then show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" - by auto -qed - -lemma integrable_continuous_real: - fixes f :: "real \ 'a::banach" - assumes "continuous_on {a .. b} f" - shows "f integrable_on {a .. b}" - by (metis assms box_real(2) integrable_continuous) - -subsection \Specialization of additivity to one dimension.\ - -subsection \Special case of additivity we need for the FTC.\ - -lemma additive_tagged_division_1: - fixes f :: "real \ 'a::real_normed_vector" - assumes "a \ b" - and "p tagged_division_of {a..b}" - shows "setsum (\(x,k). f(Sup k) - f(Inf k)) p = f b - f a" -proof - - let ?f = "(\k::(real) set. if k = {} then 0 else f(interval_upperbound k) - f(interval_lowerbound k))" - have ***: "\i\Basis. a \ i \ b \ i" - using assms by auto - have *: "add.operative ?f" - unfolding add.operative_1_lt box_eq_empty - by auto - have **: "cbox a b \ {}" - using assms(1) by auto - note setsum.operative_tagged_division[OF * assms(2)[simplified box_real[symmetric]]] - note * = this[unfolded if_not_P[OF **] interval_bounds[OF ***],symmetric] - show ?thesis - unfolding * - apply (rule setsum.cong) - unfolding split_paired_all split_conv - using assms(2) - apply auto - done -qed - - -subsection \A useful lemma allowing us to factor out the content size.\ - -lemma has_integral_factor_content: - "(f has_integral i) (cbox a b) \ - (\e>0. \d. gauge d \ (\p. p tagged_division_of (cbox a b) \ d fine p \ - norm (setsum (\(x,k). content k *\<^sub>R f x) p - i) \ e * content (cbox a b)))" -proof (cases "content (cbox a b) = 0") - case True - show ?thesis - unfolding has_integral_null_eq[OF True] - apply safe - apply (rule, rule, rule gauge_trivial, safe) - unfolding setsum_content_null[OF True] True - defer - apply (erule_tac x=1 in allE) - apply safe - defer - apply (rule fine_division_exists[of _ a b]) - apply assumption - apply (erule_tac x=p in allE) - unfolding setsum_content_null[OF True] - apply auto - done -next - case False - note F = this[unfolded content_lt_nz[symmetric]] - let ?P = "\e opp. \d. gauge d \ - (\p. p tagged_division_of (cbox a b) \ d fine p \ opp (norm ((\(x, k)\p. content k *\<^sub>R f x) - i)) e)" - show ?thesis - apply (subst has_integral) - proof safe - fix e :: real - assume e: "e > 0" - { - assume "\e>0. ?P e op <" - then show "?P (e * content (cbox a b)) op \" - apply (erule_tac x="e * content (cbox a b)" in allE) - apply (erule impE) - defer - apply (erule exE,rule_tac x=d in exI) - using F e - apply (auto simp add:field_simps) - done - } - { - assume "\e>0. ?P (e * content (cbox a b)) op \" - then show "?P e op <" - apply (erule_tac x="e / 2 / content (cbox a b)" in allE) - apply (erule impE) - defer - apply (erule exE,rule_tac x=d in exI) - using F e - apply (auto simp add: field_simps) - done - } - qed -qed - -lemma has_integral_factor_content_real: - "(f has_integral i) {a .. b::real} \ - (\e>0. \d. gauge d \ (\p. p tagged_division_of {a .. b} \ d fine p \ - norm (setsum (\(x,k). content k *\<^sub>R f x) p - i) \ e * content {a .. b} ))" - unfolding box_real[symmetric] - by (rule has_integral_factor_content) - - -subsection \Fundamental theorem of calculus.\ - -lemma interval_bounds_real: - fixes q b :: real - assumes "a \ b" - shows "Sup {a..b} = b" - and "Inf {a..b} = a" - using assms by auto - -lemma fundamental_theorem_of_calculus: - fixes f :: "real \ 'a::banach" - assumes "a \ b" - and "\x\{a .. b}. (f has_vector_derivative f' x) (at x within {a .. b})" - shows "(f' has_integral (f b - f a)) {a .. b}" - unfolding has_integral_factor_content box_real[symmetric] -proof safe - fix e :: real - assume e: "e > 0" - note assm = assms(2)[unfolded has_vector_derivative_def has_derivative_within_alt] - have *: "\P Q. \x\{a .. b}. P x \ (\e>0. \d>0. Q x e d) \ \x. \(d::real)>0. x\{a .. b} \ Q x e d" - using e by blast - note this[OF assm,unfolded gauge_existence_lemma] - from choice[OF this,unfolded Ball_def[symmetric]] guess d .. - note d=conjunctD2[OF this[rule_format],rule_format] - show "\d. gauge d \ (\p. p tagged_division_of (cbox a b) \ d fine p \ - norm ((\(x, k)\p. content k *\<^sub>R f' x) - (f b - f a)) \ e * content (cbox a b))" - apply (rule_tac x="\x. ball x (d x)" in exI) - apply safe - apply (rule gauge_ball_dependent) - apply rule - apply (rule d(1)) - proof - - fix p - assume as: "p tagged_division_of cbox a b" "(\x. ball x (d x)) fine p" - show "norm ((\(x, k)\p. content k *\<^sub>R f' x) - (f b - f a)) \ e * content (cbox a b)" - unfolding content_real[OF assms(1), simplified box_real[symmetric]] additive_tagged_division_1[OF assms(1) as(1)[simplified box_real],of f,symmetric] - unfolding additive_tagged_division_1[OF assms(1) as(1)[simplified box_real],of "\x. x",symmetric] - unfolding setsum_right_distrib - defer - unfolding setsum_subtractf[symmetric] - proof (rule setsum_norm_le,safe) - fix x k - assume "(x, k) \ p" - note xk = tagged_division_ofD(2-4)[OF as(1) this] - from this(3) guess u v by (elim exE) note k=this - have *: "u \ v" - using xk unfolding k by auto - have ball: "\xa\k. xa \ ball x (d x)" - using as(2)[unfolded fine_def,rule_format,OF \(x,k)\p\,unfolded split_conv subset_eq] . - have "norm ((v - u) *\<^sub>R f' x - (f v - f u)) \ - norm (f u - f x - (u - x) *\<^sub>R f' x) + norm (f v - f x - (v - x) *\<^sub>R f' x)" - apply (rule order_trans[OF _ norm_triangle_ineq4]) - apply (rule eq_refl) - apply (rule arg_cong[where f=norm]) - unfolding scaleR_diff_left - apply (auto simp add:algebra_simps) - done - also have "\ \ e * norm (u - x) + e * norm (v - x)" - apply (rule add_mono) - apply (rule d(2)[of "x" "u",unfolded o_def]) - prefer 4 - apply (rule d(2)[of "x" "v",unfolded o_def]) - using ball[rule_format,of u] ball[rule_format,of v] - using xk(1-2) - unfolding k subset_eq - apply (auto simp add:dist_real_def) - done - also have "\ \ e * (Sup k - Inf k)" - unfolding k interval_bounds_real[OF *] - using xk(1) - unfolding k - by (auto simp add: dist_real_def field_simps) - finally show "norm (content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) \ - e * (Sup k - Inf k)" - unfolding box_real k interval_bounds_real[OF *] content_real[OF *] - interval_upperbound_real interval_lowerbound_real - . - qed - qed -qed - -lemma ident_has_integral: - fixes a::real - assumes "a \ b" - shows "((\x. x) has_integral (b\<^sup>2 - a\<^sup>2) / 2) {a..b}" -proof - - have "((\x. x) has_integral inverse 2 * b\<^sup>2 - inverse 2 * a\<^sup>2) {a..b}" - apply (rule fundamental_theorem_of_calculus [OF assms], clarify) - unfolding power2_eq_square - by (rule derivative_eq_intros | simp)+ - then show ?thesis - by (simp add: field_simps) -qed - -lemma integral_ident [simp]: - fixes a::real - assumes "a \ b" - shows "integral {a..b} (\x. x) = (if a \ b then (b\<^sup>2 - a\<^sup>2) / 2 else 0)" -using ident_has_integral integral_unique by fastforce - -lemma ident_integrable_on: - fixes a::real - shows "(\x. x) integrable_on {a..b}" -by (metis atLeastatMost_empty_iff integrable_on_def has_integral_empty ident_has_integral) - - -subsection \Taylor series expansion\ - -lemma (in bounded_bilinear) setsum_prod_derivatives_has_vector_derivative: - assumes "p>0" - and f0: "Df 0 = f" - and Df: "\m t. m < p \ a \ t \ t \ b \ - (Df m has_vector_derivative Df (Suc m) t) (at t within {a .. b})" - and g0: "Dg 0 = g" - and Dg: "\m t. m < p \ a \ t \ t \ b \ - (Dg m has_vector_derivative Dg (Suc m) t) (at t within {a .. b})" - and ivl: "a \ t" "t \ b" - shows "((\t. \iR prod (Df i t) (Dg (p - Suc i) t)) - has_vector_derivative - prod (f t) (Dg p t) - (-1)^p *\<^sub>R prod (Df p t) (g t)) - (at t within {a .. b})" - using assms -proof cases - assume p: "p \ 1" - define p' where "p' = p - 2" - from assms p have p': "{..i. i \ p' \ Suc (Suc p' - i) = (Suc (Suc p') - i)" - by auto - let ?f = "\i. (-1) ^ i *\<^sub>R (prod (Df i t) (Dg ((p - i)) t))" - have "(\iR (prod (Df i t) (Dg (Suc (p - Suc i)) t) + - prod (Df (Suc i) t) (Dg (p - Suc i) t))) = - (\i\(Suc p'). ?f i - ?f (Suc i))" - by (auto simp: algebra_simps p'(2) numeral_2_eq_2 * lessThan_Suc_atMost) - also note setsum_telescope - finally - have "(\iR (prod (Df i t) (Dg (Suc (p - Suc i)) t) + - prod (Df (Suc i) t) (Dg (p - Suc i) t))) - = prod (f t) (Dg p t) - (- 1) ^ p *\<^sub>R prod (Df p t) (g t)" - unfolding p'[symmetric] - by (simp add: assms) - thus ?thesis - using assms - by (auto intro!: derivative_eq_intros has_vector_derivative) -qed (auto intro!: derivative_eq_intros has_vector_derivative) - -lemma - fixes f::"real\'a::banach" - assumes "p>0" - and f0: "Df 0 = f" - and Df: "\m t. m < p \ a \ t \ t \ b \ - (Df m has_vector_derivative Df (Suc m) t) (at t within {a .. b})" - and ivl: "a \ b" - defines "i \ \x. ((b - x) ^ (p - 1) / fact (p - 1)) *\<^sub>R Df p x" - shows taylor_has_integral: - "(i has_integral f b - (\iR Df i a)) {a..b}" - and taylor_integral: - "f b = (\iR Df i a) + integral {a..b} i" - and taylor_integrable: - "i integrable_on {a .. b}" -proof goal_cases - case 1 - interpret bounded_bilinear "scaleR::real\'a\'a" - by (rule bounded_bilinear_scaleR) - define g where "g s = (b - s)^(p - 1)/fact (p - 1)" for s - define Dg where [abs_def]: - "Dg n s = (if n < p then (-1)^n * (b - s)^(p - 1 - n) / fact (p - 1 - n) else 0)" for n s - have g0: "Dg 0 = g" - using \p > 0\ - by (auto simp add: Dg_def divide_simps g_def split: if_split_asm) - { - fix m - assume "p > Suc m" - hence "p - Suc m = Suc (p - Suc (Suc m))" - by auto - hence "real (p - Suc m) * fact (p - Suc (Suc m)) = fact (p - Suc m)" - by auto - } note fact_eq = this - have Dg: "\m t. m < p \ a \ t \ t \ b \ - (Dg m has_vector_derivative Dg (Suc m) t) (at t within {a .. b})" - unfolding Dg_def - by (auto intro!: derivative_eq_intros simp: has_vector_derivative_def fact_eq divide_simps) - let ?sum = "\t. \iR Dg i t *\<^sub>R Df (p - Suc i) t" - from setsum_prod_derivatives_has_vector_derivative[of _ Dg _ _ _ Df, - OF \p > 0\ g0 Dg f0 Df] - have deriv: "\t. a \ t \ t \ b \ - (?sum has_vector_derivative - g t *\<^sub>R Df p t - (- 1) ^ p *\<^sub>R Dg p t *\<^sub>R f t) (at t within {a..b})" - by auto - from fundamental_theorem_of_calculus[rule_format, OF \a \ b\ deriv] - have "(i has_integral ?sum b - ?sum a) {a .. b}" - by (simp add: i_def g_def Dg_def) - also - have one: "(- 1) ^ p' * (- 1) ^ p' = (1::real)" - and "{.. {i. p = Suc i} = {p - 1}" - for p' - using \p > 0\ - by (auto simp: power_mult_distrib[symmetric]) - then have "?sum b = f b" - using Suc_pred'[OF \p > 0\] - by (simp add: diff_eq_eq Dg_def power_0_left le_Suc_eq if_distrib - cond_application_beta setsum.If_cases f0) - also - have "{..x. p - x - 1) ` {.. (\x. p - x - 1) ` {..iR Df i a)" - by (rule setsum.reindex_cong) (auto simp add: inj_on_def Dg_def one) - finally show c: ?case . - case 2 show ?case using c integral_unique by force - case 3 show ?case using c by force -qed - - -subsection \Attempt a systematic general set of "offset" results for components.\ - -lemma gauge_modify: - assumes "(\s. open s \ open {x. f(x) \ s})" "gauge d" - shows "gauge (\x. {y. f y \ d (f x)})" - using assms - unfolding gauge_def - apply safe - defer - apply (erule_tac x="f x" in allE) - apply (erule_tac x="d (f x)" in allE) - apply auto - done - - -subsection \Only need trivial subintervals if the interval itself is trivial.\ - -lemma division_of_nontrivial: - fixes s :: "'a::euclidean_space set set" - assumes "s division_of (cbox a b)" - and "content (cbox a b) \ 0" - shows "{k. k \ s \ content k \ 0} division_of (cbox a b)" - using assms(1) - apply - -proof (induct "card s" arbitrary: s rule: nat_less_induct) - fix s::"'a set set" - assume assm: "s division_of (cbox a b)" - "\mx. m = card x \ - x division_of (cbox a b) \ {k \ x. content k \ 0} division_of (cbox a b)" - note s = division_ofD[OF assm(1)] - let ?thesis = "{k \ s. content k \ 0} division_of (cbox a b)" - { - presume *: "{k \ s. content k \ 0} \ s \ ?thesis" - show ?thesis - apply cases - defer - apply (rule *) - apply assumption - using assm(1) - apply auto - done - } - assume noteq: "{k \ s. content k \ 0} \ s" - then obtain k where k: "k \ s" "content k = 0" - by auto - from s(4)[OF k(1)] guess c d by (elim exE) note k=k this - from k have "card s > 0" - unfolding card_gt_0_iff using assm(1) by auto - then have card: "card (s - {k}) < card s" - using assm(1) k(1) - apply (subst card_Diff_singleton_if) - apply auto - done - have *: "closed (\(s - {k}))" - apply (rule closed_Union) - defer - apply rule - apply (drule DiffD1,drule s(4)) - using assm(1) - apply auto - done - have "k \ \(s - {k})" - apply safe - apply (rule *[unfolded closed_limpt,rule_format]) - unfolding islimpt_approachable - proof safe - fix x - fix e :: real - assume as: "x \ k" "e > 0" - from k(2)[unfolded k content_eq_0] guess i .. - then have i:"c\i = d\i" "i\Basis" - using s(3)[OF k(1),unfolded k] unfolding box_ne_empty by auto - then have xi: "x\i = d\i" - using as unfolding k mem_box by (metis antisym) - define y where "y = (\j\Basis. (if j = i then if c\i \ (a\i + b\i) / 2 then c\i + - min e (b\i - c\i) / 2 else c\i - min e (c\i - a\i) / 2 else x\j) *\<^sub>R j)" - show "\x'\\(s - {k}). x' \ x \ dist x' x < e" - apply (rule_tac x=y in bexI) - proof - have "d \ cbox c d" - using s(3)[OF k(1)] - unfolding k box_eq_empty mem_box - by (fastforce simp add: not_less) - then have "d \ cbox a b" - using s(2)[OF k(1)] - unfolding k - by auto - note di = this[unfolded mem_box,THEN bspec[where x=i]] - then have xyi: "y\i \ x\i" - unfolding y_def i xi - using as(2) assms(2)[unfolded content_eq_0] i(2) - by (auto elim!: ballE[of _ _ i]) - then show "y \ x" - unfolding euclidean_eq_iff[where 'a='a] using i by auto - have *: "Basis = insert i (Basis - {i})" - using i by auto - have "norm (y - x) < e + setsum (\i. 0) Basis" - apply (rule le_less_trans[OF norm_le_l1]) - apply (subst *) - apply (subst setsum.insert) - prefer 3 - apply (rule add_less_le_mono) - proof - - show "\(y - x) \ i\ < e" - using di as(2) y_def i xi by (auto simp: inner_simps) - show "(\i\Basis - {i}. \(y - x) \ i\) \ (\i\Basis. 0)" - unfolding y_def by (auto simp: inner_simps) - qed auto - then show "dist y x < e" - unfolding dist_norm by auto - have "y \ k" - unfolding k mem_box - apply rule - apply (erule_tac x=i in ballE) - using xyi k i xi - apply auto - done - moreover - have "y \ \s" - using set_rev_mp[OF as(1) s(2)[OF k(1)]] as(2) di i - unfolding s mem_box y_def - by (auto simp: field_simps elim!: ballE[of _ _ i]) - ultimately - show "y \ \(s - {k})" by auto - qed - qed - then have "\(s - {k}) = cbox a b" - unfolding s(6)[symmetric] by auto - then have "{ka \ s - {k}. content ka \ 0} division_of (cbox a b)" - apply - - apply (rule assm(2)[rule_format,OF card refl]) - apply (rule division_ofI) - defer - apply (rule_tac[1-4] s) - using assm(1) - apply auto - done - moreover - have "{ka \ s - {k}. content ka \ 0} = {k \ s. content k \ 0}" - using k by auto - ultimately show ?thesis by auto -qed - - -subsection \Integrability on subintervals.\ - -lemma operative_integrable: - fixes f :: "'b::euclidean_space \ 'a::banach" - shows "comm_monoid.operative op \ True (\i. f integrable_on i)" - unfolding comm_monoid.operative_def[OF comm_monoid_and] - apply safe - apply (subst integrable_on_def) - unfolding has_integral_null_eq - apply (rule, rule refl) - apply (rule, assumption, assumption)+ - unfolding integrable_on_def - by (auto intro!: has_integral_split) - -lemma integrable_subinterval: - fixes f :: "'b::euclidean_space \ 'a::banach" - assumes "f integrable_on cbox a b" - and "cbox c d \ cbox a b" - shows "f integrable_on cbox c d" - apply (cases "cbox c d = {}") - defer - apply (rule partial_division_extend_1[OF assms(2)],assumption) - using comm_monoid_set.operative_division[OF comm_monoid_set_and operative_integrable,symmetric,of _ _ _ f] assms(1) - apply (auto simp: comm_monoid_set_F_and) - done - -lemma integrable_subinterval_real: - fixes f :: "real \ 'a::banach" - assumes "f integrable_on {a .. b}" - and "{c .. d} \ {a .. b}" - shows "f integrable_on {c .. d}" - by (metis assms(1) assms(2) box_real(2) integrable_subinterval) - - -subsection \Combining adjacent intervals in 1 dimension.\ - -lemma has_integral_combine: - fixes a b c :: real - assumes "a \ c" - and "c \ b" - and "(f has_integral i) {a .. c}" - and "(f has_integral (j::'a::banach)) {c .. b}" - shows "(f has_integral (i + j)) {a .. b}" -proof - - note operative_integral[of f, unfolded comm_monoid.operative_1_le[OF add.comm_monoid_lift_option]] - note conjunctD2[OF this,rule_format] - note * = this(2)[OF conjI[OF assms(1-2)],unfolded if_P[OF assms(3)]] - then have "f integrable_on cbox a b" - apply - - apply (rule ccontr) - apply (subst(asm) if_P) - defer - apply (subst(asm) if_P) - using assms(3-) - apply auto - done - with * - show ?thesis - apply - - apply (subst(asm) if_P) - defer - apply (subst(asm) if_P) - defer - apply (subst(asm) if_P) - using assms(3-) - apply (auto simp add: integrable_on_def integral_unique) - done -qed - -lemma integral_combine: - fixes f :: "real \ 'a::banach" - assumes "a \ c" - and "c \ b" - and "f integrable_on {a .. b}" - shows "integral {a .. c} f + integral {c .. b} f = integral {a .. b} f" - apply (rule integral_unique[symmetric]) - apply (rule has_integral_combine[OF assms(1-2)]) - apply (metis assms(2) assms(3) atLeastatMost_subset_iff box_real(2) content_pos_le content_real_eq_0 integrable_integral integrable_subinterval le_add_same_cancel2 monoid_add_class.add.left_neutral) - by (metis assms(1) assms(3) atLeastatMost_subset_iff box_real(2) content_pos_le content_real_eq_0 integrable_integral integrable_subinterval le_add_same_cancel1 monoid_add_class.add.right_neutral) - -lemma integrable_combine: - fixes f :: "real \ 'a::banach" - assumes "a \ c" - and "c \ b" - and "f integrable_on {a .. c}" - and "f integrable_on {c .. b}" - shows "f integrable_on {a .. b}" - using assms - unfolding integrable_on_def - by (fastforce intro!:has_integral_combine) - - -subsection \Reduce integrability to "local" integrability.\ - -lemma integrable_on_little_subintervals: - fixes f :: "'b::euclidean_space \ 'a::banach" - assumes "\x\cbox a b. \d>0. \u v. x \ cbox u v \ cbox u v \ ball x d \ cbox u v \ cbox a b \ - f integrable_on cbox u v" - shows "f integrable_on cbox a b" -proof - - have "\x. \d. x\cbox a b \ d>0 \ (\u v. x \ cbox u v \ cbox u v \ ball x d \ cbox u v \ cbox a b \ - f integrable_on cbox u v)" - using assms by auto - note this[unfolded gauge_existence_lemma] - from choice[OF this] guess d .. note d=this[rule_format] - guess p - apply (rule fine_division_exists[OF gauge_ball_dependent,of d a b]) - using d - by auto - note p=this(1-2) - note division_of_tagged_division[OF this(1)] - note * = comm_monoid_set.operative_division[OF comm_monoid_set_and operative_integrable, OF this, symmetric, of f] - show ?thesis - unfolding * comm_monoid_set_F_and - apply safe - unfolding snd_conv - proof - - fix x k - assume "(x, k) \ p" - note tagged_division_ofD(2-4)[OF p(1) this] fineD[OF p(2) this] - then show "f integrable_on k" - apply safe - apply (rule d[THEN conjunct2,rule_format,of x]) - apply (auto intro: order.trans) - done - qed -qed - - -subsection \Second FTC or existence of antiderivative.\ - -lemma integrable_const[intro]: "(\x. c) integrable_on cbox a b" - unfolding integrable_on_def - apply rule - apply (rule has_integral_const) - done - -lemma integral_has_vector_derivative_continuous_at: - fixes f :: "real \ 'a::banach" - assumes f: "f integrable_on {a..b}" - and x: "x \ {a..b}" - and fx: "continuous (at x within {a..b}) f" - shows "((\u. integral {a..u} f) has_vector_derivative f x) (at x within {a..b})" -proof - - let ?I = "\a b. integral {a..b} f" - { fix e::real - assume "e > 0" - obtain d where "d>0" and d: "\x'. \x' \ {a..b}; \x' - x\ < d\ \ norm(f x' - f x) \ e" - using \e>0\ fx by (auto simp: continuous_within_eps_delta dist_norm less_imp_le) - have "norm (integral {a..y} f - integral {a..x} f - (y - x) *\<^sub>R f x) \ e * \y - x\" - if y: "y \ {a..b}" and yx: "\y - x\ < d" for y - proof (cases "y < x") - case False - have "f integrable_on {a..y}" - using f y by (simp add: integrable_subinterval_real) - then have Idiff: "?I a y - ?I a x = ?I x y" - using False x by (simp add: algebra_simps integral_combine) - have fux_int: "((\u. f u - f x) has_integral integral {x..y} f - (y - x) *\<^sub>R f x) {x..y}" - apply (rule has_integral_sub) - using x y apply (force intro: integrable_integral [OF integrable_subinterval_real [OF f]]) - using has_integral_const_real [of "f x" x y] False - apply (simp add: ) - done - show ?thesis - using False - apply (simp add: abs_eq_content del: content_real_if) - apply (rule has_integral_bound_real[where f="(\u. f u - f x)"]) - using yx False d x y \e>0\ apply (auto simp add: Idiff fux_int) - done - next - case True - have "f integrable_on {a..x}" - using f x by (simp add: integrable_subinterval_real) - then have Idiff: "?I a x - ?I a y = ?I y x" - using True x y by (simp add: algebra_simps integral_combine) - have fux_int: "((\u. f u - f x) has_integral integral {y..x} f - (x - y) *\<^sub>R f x) {y..x}" - apply (rule has_integral_sub) - using x y apply (force intro: integrable_integral [OF integrable_subinterval_real [OF f]]) - using has_integral_const_real [of "f x" y x] True - apply (simp add: ) - done - have "norm (integral {a..x} f - integral {a..y} f - (x - y) *\<^sub>R f x) \ e * \y - x\" - using True - apply (simp add: abs_eq_content del: content_real_if) - apply (rule has_integral_bound_real[where f="(\u. f u - f x)"]) - using yx True d x y \e>0\ apply (auto simp add: Idiff fux_int) - done - then show ?thesis - by (simp add: algebra_simps norm_minus_commute) - qed - then have "\d>0. \y\{a..b}. \y - x\ < d \ norm (integral {a..y} f - integral {a..x} f - (y - x) *\<^sub>R f x) \ e * \y - x\" - using \d>0\ by blast - } - then show ?thesis - by (simp add: has_vector_derivative_def has_derivative_within_alt bounded_linear_scaleR_left) -qed - -lemma integral_has_vector_derivative: - fixes f :: "real \ 'a::banach" - assumes "continuous_on {a .. b} f" - and "x \ {a .. b}" - shows "((\u. integral {a .. u} f) has_vector_derivative f(x)) (at x within {a .. b})" -apply (rule integral_has_vector_derivative_continuous_at [OF integrable_continuous_real]) -using assms -apply (auto simp: continuous_on_eq_continuous_within) -done - -lemma antiderivative_continuous: - fixes q b :: real - assumes "continuous_on {a .. b} f" - obtains g where "\x\{a .. b}. (g has_vector_derivative (f x::_::banach)) (at x within {a .. b})" - apply (rule that) - apply rule - using integral_has_vector_derivative[OF assms] - apply auto - done - - -subsection \Combined fundamental theorem of calculus.\ - -lemma antiderivative_integral_continuous: - fixes f :: "real \ 'a::banach" - assumes "continuous_on {a .. b} f" - obtains g where "\u\{a .. b}. \v \ {a .. b}. u \ v \ (f has_integral (g v - g u)) {u .. v}" -proof - - from antiderivative_continuous[OF assms] guess g . note g=this - show ?thesis - apply (rule that[of g]) - apply safe - proof goal_cases - case prems: (1 u v) - have "\x\cbox u v. (g has_vector_derivative f x) (at x within cbox u v)" - apply rule - apply (rule has_vector_derivative_within_subset) - apply (rule g[rule_format]) - using prems(1,2) - apply auto - done - then show ?case - using fundamental_theorem_of_calculus[OF prems(3), of g f] by auto - qed -qed - - -subsection \General "twiddling" for interval-to-interval function image.\ - -lemma has_integral_twiddle: - assumes "0 < r" - and "\x. h(g x) = x" - and "\x. g(h x) = x" - and "\x. continuous (at x) g" - and "\u v. \w z. g ` cbox u v = cbox w z" - and "\u v. \w z. h ` cbox u v = cbox w z" - and "\u v. content(g ` cbox u v) = r * content (cbox u v)" - and "(f has_integral i) (cbox a b)" - shows "((\x. f(g x)) has_integral (1 / r) *\<^sub>R i) (h ` cbox a b)" -proof - - show ?thesis when *: "cbox a b \ {} \ ?thesis" - apply cases - defer - apply (rule *) - apply assumption - proof goal_cases - case prems: 1 - then show ?thesis - unfolding prems assms(8)[unfolded prems has_integral_empty_eq] by auto - qed - assume "cbox a b \ {}" - from assms(6)[rule_format,of a b] guess w z by (elim exE) note wz=this - have inj: "inj g" "inj h" - unfolding inj_on_def - apply safe - apply(rule_tac[!] ccontr) - using assms(2) - apply(erule_tac x=x in allE) - using assms(2) - apply(erule_tac x=y in allE) - defer - using assms(3) - apply (erule_tac x=x in allE) - using assms(3) - apply(erule_tac x=y in allE) - apply auto - done - show ?thesis - unfolding has_integral_def has_integral_compact_interval_def - apply (subst if_P) - apply rule - apply rule - apply (rule wz) - proof safe - fix e :: real - assume e: "e > 0" - with assms(1) have "e * r > 0" by simp - from assms(8)[unfolded has_integral,rule_format,OF this] guess d by (elim exE conjE) note d=this[rule_format] - define d' where "d' x = {y. g y \ d (g x)}" for x - have d': "\x. d' x = {y. g y \ (d (g x))}" - unfolding d'_def .. - show "\d. gauge d \ (\p. p tagged_division_of h ` cbox a b \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e)" - proof (rule_tac x=d' in exI, safe) - show "gauge d'" - using d(1) - unfolding gauge_def d' - using continuous_open_preimage_univ[OF assms(4)] - by auto - fix p - assume as: "p tagged_division_of h ` cbox a b" "d' fine p" - note p = tagged_division_ofD[OF as(1)] - have "(\(x, k). (g x, g ` k)) ` p tagged_division_of (cbox a b) \ d fine (\(x, k). (g x, g ` k)) ` p" - unfolding tagged_division_of - proof safe - show "finite ((\(x, k). (g x, g ` k)) ` p)" - using as by auto - show "d fine (\(x, k). (g x, g ` k)) ` p" - using as(2) unfolding fine_def d' by auto - fix x k - assume xk[intro]: "(x, k) \ p" - show "g x \ g ` k" - using p(2)[OF xk] by auto - show "\u v. g ` k = cbox u v" - using p(4)[OF xk] using assms(5-6) by auto - { - fix y - assume "y \ k" - then show "g y \ cbox a b" "g y \ cbox a b" - using p(3)[OF xk,unfolded subset_eq,rule_format,of "h (g y)"] - using assms(2)[rule_format,of y] - unfolding inj_image_mem_iff[OF inj(2)] - by auto - } - fix x' k' - assume xk': "(x', k') \ p" - fix z - assume z: "z \ interior (g ` k)" "z \ interior (g ` k')" - have same: "(x, k) = (x', k')" - apply - - apply (rule ccontr) - apply (drule p(5)[OF xk xk']) - proof - - assume as: "interior k \ interior k' = {}" - have "z \ g ` (interior k \ interior k')" - using interior_image_subset[OF assms(4) inj(1)] z - unfolding image_Int[OF inj(1)] by blast - then show False - using as by blast - qed - then show "g x = g x'" - by auto - { - fix z - assume "z \ k" - then show "g z \ g ` k'" - using same by auto - } - { - fix z - assume "z \ k'" - then show "g z \ g ` k" - using same by auto - } - next - fix x - assume "x \ cbox a b" - then have "h x \ \{k. \x. (x, k) \ p}" - using p(6) by auto - then guess X unfolding Union_iff .. note X=this - from this(1) guess y unfolding mem_Collect_eq .. - then show "x \ \{k. \x. (x, k) \ (\(x, k). (g x, g ` k)) ` p}" - apply - - apply (rule_tac X="g ` X" in UnionI) - defer - apply (rule_tac x="h x" in image_eqI) - using X(2) assms(3)[rule_format,of x] - apply auto - done - qed - note ** = d(2)[OF this] - have *: "inj_on (\(x, k). (g x, g ` k)) p" - using inj(1) unfolding inj_on_def by fastforce - have "(\(x, k)\(\(x, k). (g x, g ` k)) ` p. content k *\<^sub>R f x) - i = r *\<^sub>R (\(x, k)\p. content k *\<^sub>R f (g x)) - i" (is "?l = _") - using assms(7) - apply (simp only: algebra_simps add_left_cancel scaleR_right.setsum) - apply (subst setsum.reindex_bij_betw[symmetric, where h="\(x, k). (g x, g ` k)" and S=p]) - apply (auto intro!: * setsum.cong simp: bij_betw_def dest!: p(4)) - done - also have "\ = r *\<^sub>R ((\(x, k)\p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i)" (is "_ = ?r") - unfolding scaleR_diff_right scaleR_scaleR - using assms(1) - by auto - finally have *: "?l = ?r" . - show "norm ((\(x, k)\p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e" - using ** - unfolding * - unfolding norm_scaleR - using assms(1) - by (auto simp add:field_simps) - qed - qed -qed - - -subsection \Special case of a basic affine transformation.\ - -lemma interval_image_affinity_interval: - "\u v. (\x. m *\<^sub>R (x::'a::euclidean_space) + c) ` cbox a b = cbox u v" - unfolding image_affinity_cbox - by auto - -lemma content_image_affinity_cbox: - "content((\x::'a::euclidean_space. m *\<^sub>R x + c) ` cbox a b) = - \m\ ^ DIM('a) * content (cbox a b)" (is "?l = ?r") -proof (cases "cbox a b = {}") - case True then show ?thesis by simp -next - case False - show ?thesis - proof (cases "m \ 0") - case True - with \cbox a b \ {}\ have "cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c) \ {}" - unfolding box_ne_empty - apply (intro ballI) - apply (erule_tac x=i in ballE) - apply (auto simp: inner_simps mult_left_mono) - done - moreover from True have *: "\i. (m *\<^sub>R b + c) \ i - (m *\<^sub>R a + c) \ i = m *\<^sub>R (b - a) \ i" - by (simp add: inner_simps field_simps) - ultimately show ?thesis - by (simp add: image_affinity_cbox True content_cbox' - setprod.distrib setprod_constant inner_diff_left) - next - case False - with \cbox a b \ {}\ have "cbox (m *\<^sub>R b + c) (m *\<^sub>R a + c) \ {}" - unfolding box_ne_empty - apply (intro ballI) - apply (erule_tac x=i in ballE) - apply (auto simp: inner_simps mult_left_mono) - done - moreover from False have *: "\i. (m *\<^sub>R a + c) \ i - (m *\<^sub>R b + c) \ i = (-m) *\<^sub>R (b - a) \ i" - by (simp add: inner_simps field_simps) - ultimately show ?thesis using False - by (simp add: image_affinity_cbox content_cbox' - setprod.distrib[symmetric] setprod_constant[symmetric] inner_diff_left) - qed -qed - -lemma has_integral_affinity: - fixes a :: "'a::euclidean_space" - assumes "(f has_integral i) (cbox a b)" - and "m \ 0" - shows "((\x. f(m *\<^sub>R x + c)) has_integral ((1 / (\m\ ^ DIM('a))) *\<^sub>R i)) ((\x. (1 / m) *\<^sub>R x + -((1 / m) *\<^sub>R c)) ` cbox a b)" - apply (rule has_integral_twiddle) - using assms - apply (safe intro!: interval_image_affinity_interval content_image_affinity_cbox) - apply (rule zero_less_power) - unfolding scaleR_right_distrib - apply auto - done - -lemma integrable_affinity: - assumes "f integrable_on cbox a b" - and "m \ 0" - shows "(\x. f(m *\<^sub>R x + c)) integrable_on ((\x. (1 / m) *\<^sub>R x + -((1/m) *\<^sub>R c)) ` cbox a b)" - using assms - unfolding integrable_on_def - apply safe - apply (drule has_integral_affinity) - apply auto - done - -lemmas has_integral_affinity01 = has_integral_affinity [of _ _ 0 "1::real", simplified] - -subsection \Special case of stretching coordinate axes separately.\ - -lemma content_image_stretch_interval: - "content ((\x::'a::euclidean_space. (\k\Basis. (m k * (x\k))*\<^sub>R k)::'a) ` cbox a b) = - \setprod m Basis\ * content (cbox a b)" -proof (cases "cbox a b = {}") - case True - then show ?thesis - unfolding content_def image_is_empty image_stretch_interval if_P[OF True] by auto -next - case False - then have "(\x. (\k\Basis. (m k * (x\k))*\<^sub>R k)) ` cbox a b \ {}" - by auto - then show ?thesis - using False - unfolding content_def image_stretch_interval - apply - - unfolding interval_bounds' if_not_P - unfolding abs_setprod setprod.distrib[symmetric] - apply (rule setprod.cong) - apply (rule refl) - unfolding lessThan_iff - apply (simp only: inner_setsum_left_Basis) - proof - - fix i :: 'a - assume i: "i \ Basis" - have "(m i < 0 \ m i > 0) \ m i = 0" - by auto - then show "max (m i * (a \ i)) (m i * (b \ i)) - min (m i * (a \ i)) (m i * (b \ i)) = - \m i\ * (b \ i - a \ i)" - apply - - apply (erule disjE)+ - unfolding min_def max_def - using False[unfolded box_ne_empty,rule_format,of i] i - apply (auto simp add:field_simps not_le mult_le_cancel_left_neg mult_le_cancel_left_pos) - done - qed -qed - -lemma has_integral_stretch: - fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" - assumes "(f has_integral i) (cbox a b)" - and "\k\Basis. m k \ 0" - shows "((\x. f (\k\Basis. (m k * (x\k))*\<^sub>R k)) has_integral - ((1/ \setprod m Basis\) *\<^sub>R i)) ((\x. (\k\Basis. (1 / m k * (x\k))*\<^sub>R k)) ` cbox a b)" - apply (rule has_integral_twiddle[where f=f]) - unfolding zero_less_abs_iff content_image_stretch_interval - unfolding image_stretch_interval empty_as_interval euclidean_eq_iff[where 'a='a] - using assms -proof - - show "\y::'a. continuous (at y) (\x. (\k\Basis. (m k * (x\k))*\<^sub>R k))" - apply rule - apply (rule linear_continuous_at) - unfolding linear_linear - unfolding linear_iff inner_simps euclidean_eq_iff[where 'a='a] - apply (auto simp add: field_simps) - done -qed auto - -lemma integrable_stretch: - fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" - assumes "f integrable_on cbox a b" - and "\k\Basis. m k \ 0" - shows "(\x::'a. f (\k\Basis. (m k * (x\k))*\<^sub>R k)) integrable_on - ((\x. \k\Basis. (1 / m k * (x\k))*\<^sub>R k) ` cbox a b)" - using assms - unfolding integrable_on_def - apply - - apply (erule exE) - apply (drule has_integral_stretch) - apply assumption - apply auto - done - -subsection \even more special cases.\ - -lemma uminus_interval_vector[simp]: - fixes a b :: "'a::euclidean_space" - shows "uminus ` cbox a b = cbox (-b) (-a)" - apply (rule set_eqI) - apply rule - defer - unfolding image_iff - apply (rule_tac x="-x" in bexI) - apply (auto simp add:minus_le_iff le_minus_iff mem_box) - done - -lemma has_integral_reflect_lemma[intro]: - assumes "(f has_integral i) (cbox a b)" - shows "((\x. f(-x)) has_integral i) (cbox (-b) (-a))" - using has_integral_affinity[OF assms, of "-1" 0] - by auto - -lemma has_integral_reflect_lemma_real[intro]: - assumes "(f has_integral i) {a .. b::real}" - shows "((\x. f(-x)) has_integral i) {-b .. -a}" - using assms - unfolding box_real[symmetric] - by (rule has_integral_reflect_lemma) - -lemma has_integral_reflect[simp]: - "((\x. f (-x)) has_integral i) (cbox (-b) (-a)) \ (f has_integral i) (cbox a b)" - apply rule - apply (drule_tac[!] has_integral_reflect_lemma) - apply auto - done - -lemma integrable_reflect[simp]: "(\x. f(-x)) integrable_on cbox (-b) (-a) \ f integrable_on cbox a b" - unfolding integrable_on_def by auto - -lemma integrable_reflect_real[simp]: "(\x. f(-x)) integrable_on {-b .. -a} \ f integrable_on {a .. b::real}" - unfolding box_real[symmetric] - by (rule integrable_reflect) - -lemma integral_reflect[simp]: "integral (cbox (-b) (-a)) (\x. f (-x)) = integral (cbox a b) f" - unfolding integral_def by auto - -lemma integral_reflect_real[simp]: "integral {-b .. -a} (\x. f (-x)) = integral {a .. b::real} f" - unfolding box_real[symmetric] - by (rule integral_reflect) - - -subsection \Stronger form of FCT; quite a tedious proof.\ - -lemma bgauge_existence_lemma: "(\x\s. \d::real. 0 < d \ q d x) \ (\x. \d>0. x\s \ q d x)" - by (meson zero_less_one) - -lemma additive_tagged_division_1': - fixes f :: "real \ 'a::real_normed_vector" - assumes "a \ b" - and "p tagged_division_of {a..b}" - shows "setsum (\(x,k). f (Sup k) - f(Inf k)) p = f b - f a" - using additive_tagged_division_1[OF _ assms(2), of f] - using assms(1) - by auto - -lemma split_minus[simp]: "(\(x, k). f x k) x - (\(x, k). g x k) x = (\(x, k). f x k - g x k) x" - by (simp add: split_def) - -lemma norm_triangle_le_sub: "norm x + norm y \ e \ norm (x - y) \ e" - apply (subst(asm)(2) norm_minus_cancel[symmetric]) - apply (drule norm_triangle_le) - apply (auto simp add: algebra_simps) - done - -lemma fundamental_theorem_of_calculus_interior: - fixes f :: "real \ 'a::real_normed_vector" - assumes "a \ b" - and "continuous_on {a .. b} f" - and "\x\{a <..< b}. (f has_vector_derivative f'(x)) (at x)" - shows "(f' has_integral (f b - f a)) {a .. b}" -proof - - { - presume *: "a < b \ ?thesis" - show ?thesis - proof (cases "a < b") - case True - then show ?thesis by (rule *) - next - case False - then have "a = b" - using assms(1) by auto - then have *: "cbox a b = {b}" "f b - f a = 0" - by (auto simp add: order_antisym) - show ?thesis - unfolding *(2) - unfolding content_eq_0 - using * \a = b\ - by (auto simp: ex_in_conv) - qed - } - assume ab: "a < b" - let ?P = "\e. \d. gauge d \ (\p. p tagged_division_of {a .. b} \ d fine p \ - norm ((\(x, k)\p. content k *\<^sub>R f' x) - (f b - f a)) \ e * content {a .. b})" - { presume "\e. e > 0 \ ?P e" then show ?thesis unfolding has_integral_factor_content_real by auto } - fix e :: real - assume e: "e > 0" - note assms(3)[unfolded has_vector_derivative_def has_derivative_at_alt ball_conj_distrib] - note conjunctD2[OF this] - note bounded=this(1) and this(2) - from this(2) have "\x\box a b. \d>0. \y. norm (y - x) < d \ - norm (f y - f x - (y - x) *\<^sub>R f' x) \ e/2 * norm (y - x)" - apply - - apply safe - apply (erule_tac x=x in ballE) - apply (erule_tac x="e/2" in allE) - using e - apply auto - done - note this[unfolded bgauge_existence_lemma] - from choice[OF this] guess d .. - note conjunctD2[OF this[rule_format]] - note d = this[rule_format] - have "bounded (f ` cbox a b)" - apply (rule compact_imp_bounded compact_continuous_image)+ - using compact_cbox assms - apply auto - done - from this[unfolded bounded_pos] guess B .. note B = this[rule_format] - - have "\da. 0 < da \ (\c. a \ c \ {a .. c} \ {a .. b} \ {a .. c} \ ball a da \ - norm (content {a .. c} *\<^sub>R f' a - (f c - f a)) \ (e * (b - a)) / 4)" - proof - - have "a \ {a .. b}" - using ab by auto - note assms(2)[unfolded continuous_on_eq_continuous_within,rule_format,OF this] - note * = this[unfolded continuous_within Lim_within,rule_format] - have "(e * (b - a)) / 8 > 0" - using e ab by (auto simp add: field_simps) - from *[OF this] guess k .. note k = conjunctD2[OF this,rule_format] - have "\l. 0 < l \ norm(l *\<^sub>R f' a) \ (e * (b - a)) / 8" - proof (cases "f' a = 0") - case True - thus ?thesis using ab e by auto - next - case False - then show ?thesis - apply (rule_tac x="(e * (b - a)) / 8 / norm (f' a)" in exI) - using ab e - apply (auto simp add: field_simps) - done - qed - then guess l .. note l = conjunctD2[OF this] - show ?thesis - apply (rule_tac x="min k l" in exI) - apply safe - unfolding min_less_iff_conj - apply rule - apply (rule l k)+ - proof - - fix c - assume as: "a \ c" "{a .. c} \ {a .. b}" "{a .. c} \ ball a (min k l)" - note as' = this[unfolded subset_eq Ball_def mem_ball dist_real_def mem_box] - have "norm ((c - a) *\<^sub>R f' a - (f c - f a)) \ norm ((c - a) *\<^sub>R f' a) + norm (f c - f a)" - by (rule norm_triangle_ineq4) - also have "\ \ e * (b - a) / 8 + e * (b - a) / 8" - proof (rule add_mono) - have "\c - a\ \ \l\" - using as' by auto - then show "norm ((c - a) *\<^sub>R f' a) \ e * (b - a) / 8" - apply - - apply (rule order_trans[OF _ l(2)]) - unfolding norm_scaleR - apply (rule mult_right_mono) - apply auto - done - next - show "norm (f c - f a) \ e * (b - a) / 8" - apply (rule less_imp_le) - apply (cases "a = c") - defer - apply (rule k(2)[unfolded dist_norm]) - using as' e ab - apply (auto simp add: field_simps) - done - qed - finally show "norm (content {a .. c} *\<^sub>R f' a - (f c - f a)) \ e * (b - a) / 4" - unfolding content_real[OF as(1)] by auto - qed - qed - then guess da .. note da=conjunctD2[OF this,rule_format] - - have "\db>0. \c\b. {c .. b} \ {a .. b} \ {c .. b} \ ball b db \ - norm (content {c .. b} *\<^sub>R f' b - (f b - f c)) \ (e * (b - a)) / 4" - proof - - have "b \ {a .. b}" - using ab by auto - note assms(2)[unfolded continuous_on_eq_continuous_within,rule_format,OF this] - note * = this[unfolded continuous_within Lim_within,rule_format] have "(e * (b - a)) / 8 > 0" - using e ab by (auto simp add: field_simps) - from *[OF this] guess k .. note k = conjunctD2[OF this,rule_format] - have "\l. 0 < l \ norm (l *\<^sub>R f' b) \ (e * (b - a)) / 8" - proof (cases "f' b = 0") - case True - thus ?thesis using ab e by auto - next - case False - then show ?thesis - apply (rule_tac x="(e * (b - a)) / 8 / norm (f' b)" in exI) - using ab e - apply (auto simp add: field_simps) - done - qed - then guess l .. note l = conjunctD2[OF this] - show ?thesis - apply (rule_tac x="min k l" in exI) - apply safe - unfolding min_less_iff_conj - apply rule - apply (rule l k)+ - proof - - fix c - assume as: "c \ b" "{c..b} \ {a..b}" "{c..b} \ ball b (min k l)" - note as' = this[unfolded subset_eq Ball_def mem_ball dist_real_def mem_box] - have "norm ((b - c) *\<^sub>R f' b - (f b - f c)) \ norm ((b - c) *\<^sub>R f' b) + norm (f b - f c)" - by (rule norm_triangle_ineq4) - also have "\ \ e * (b - a) / 8 + e * (b - a) / 8" - proof (rule add_mono) - have "\c - b\ \ \l\" - using as' by auto - then show "norm ((b - c) *\<^sub>R f' b) \ e * (b - a) / 8" - apply - - apply (rule order_trans[OF _ l(2)]) - unfolding norm_scaleR - apply (rule mult_right_mono) - apply auto - done - next - show "norm (f b - f c) \ e * (b - a) / 8" - apply (rule less_imp_le) - apply (cases "b = c") - defer - apply (subst norm_minus_commute) - apply (rule k(2)[unfolded dist_norm]) - using as' e ab - apply (auto simp add: field_simps) - done - qed - finally show "norm (content {c .. b} *\<^sub>R f' b - (f b - f c)) \ e * (b - a) / 4" - unfolding content_real[OF as(1)] by auto - qed - qed - then guess db .. note db=conjunctD2[OF this,rule_format] - - let ?d = "(\x. ball x (if x=a then da else if x=b then db else d x))" - show "?P e" - apply (rule_tac x="?d" in exI) - proof (safe, goal_cases) - case 1 - show ?case - apply (rule gauge_ball_dependent) - using ab db(1) da(1) d(1) - apply auto - done - next - case as: (2 p) - let ?A = "{t. fst t \ {a, b}}" - note p = tagged_division_ofD[OF as(1)] - have pA: "p = (p \ ?A) \ (p - ?A)" "finite (p \ ?A)" "finite (p - ?A)" "(p \ ?A) \ (p - ?A) = {}" - using as by auto - note * = additive_tagged_division_1'[OF assms(1) as(1), symmetric] - have **: "\n1 s1 n2 s2::real. n2 \ s2 / 2 \ n1 - s1 \ s2 / 2 \ n1 + n2 \ s1 + s2" - by arith - show ?case - unfolding content_real[OF assms(1)] and *[of "\x. x"] *[of f] setsum_subtractf[symmetric] split_minus - unfolding setsum_right_distrib - apply (subst(2) pA) - apply (subst pA) - unfolding setsum.union_disjoint[OF pA(2-)] - proof (rule norm_triangle_le, rule **, goal_cases) - case 1 - show ?case - apply (rule order_trans) - apply (rule setsum_norm_le) - defer - apply (subst setsum_divide_distrib) - apply (rule order_refl) - apply safe - apply (unfold not_le o_def split_conv fst_conv) - proof (rule ccontr) - fix x k - assume xk: "(x, k) \ p" - "e * (Sup k - Inf k) / 2 < - norm (content k *\<^sub>R f' x - (f (Sup k) - f (Inf k)))" - from p(4)[OF this(1)] guess u v by (elim exE) note k=this - then have "u \ v" and uv: "{u, v} \ cbox u v" - using p(2)[OF xk(1)] by auto - note result = xk(2)[unfolded k box_real interval_bounds_real[OF this(1)] content_real[OF this(1)]] - - assume as': "x \ a" "x \ b" - then have "x \ box a b" - using p(2-3)[OF xk(1)] by (auto simp: mem_box) - note * = d(2)[OF this] - have "norm ((v - u) *\<^sub>R f' (x) - (f (v) - f (u))) = - norm ((f (u) - f (x) - (u - x) *\<^sub>R f' (x)) - (f (v) - f (x) - (v - x) *\<^sub>R f' (x)))" - apply (rule arg_cong[of _ _ norm]) - unfolding scaleR_left.diff - apply auto - done - also have "\ \ e / 2 * norm (u - x) + e / 2 * norm (v - x)" - apply (rule norm_triangle_le_sub) - apply (rule add_mono) - apply (rule_tac[!] *) - using fineD[OF as(2) xk(1)] as' - unfolding k subset_eq - apply - - apply (erule_tac x=u in ballE) - apply (erule_tac[3] x=v in ballE) - using uv - apply (auto simp:dist_real_def) - done - also have "\ \ e / 2 * norm (v - u)" - using p(2)[OF xk(1)] - unfolding k - by (auto simp add: field_simps) - finally have "e * (v - u) / 2 < e * (v - u) / 2" - apply - - apply (rule less_le_trans[OF result]) - using uv - apply auto - done - then show False by auto - qed - next - have *: "\x s1 s2::real. 0 \ s1 \ x \ (s1 + s2) / 2 \ x - s1 \ s2 / 2" - by auto - case 2 - show ?case - apply (rule *) - apply (rule setsum_nonneg) - apply rule - apply (unfold split_paired_all split_conv) - defer - unfolding setsum.union_disjoint[OF pA(2-),symmetric] pA(1)[symmetric] - unfolding setsum_right_distrib[symmetric] - apply (subst additive_tagged_division_1[OF _ as(1)]) - apply (rule assms) - proof - - fix x k - assume "(x, k) \ p \ {t. fst t \ {a, b}}" - note xk=IntD1[OF this] - from p(4)[OF this] guess u v by (elim exE) note uv=this - with p(2)[OF xk] have "cbox u v \ {}" - by auto - then show "0 \ e * ((Sup k) - (Inf k))" - unfolding uv using e by (auto simp add: field_simps) - next - have *: "\s f t e. setsum f s = setsum f t \ norm (setsum f t) \ e \ norm (setsum f s) \ e" - by auto - show "norm (\(x, k)\p \ ?A. content k *\<^sub>R f' x - - (f ((Sup k)) - f ((Inf k)))) \ e * (b - a) / 2" - apply (rule *[where t1="p \ {t. fst t \ {a, b} \ content(snd t) \ 0}"]) - apply (rule setsum.mono_neutral_right[OF pA(2)]) - defer - apply rule - unfolding split_paired_all split_conv o_def - proof goal_cases - fix x k - assume "(x, k) \ p \ {t. fst t \ {a, b}} - p \ {t. fst t \ {a, b} \ content (snd t) \ 0}" - then have xk: "(x, k) \ p" "content k = 0" - by auto - from p(4)[OF xk(1)] guess u v by (elim exE) note uv=this - have "k \ {}" - using p(2)[OF xk(1)] by auto - then have *: "u = v" - using xk - unfolding uv content_eq_0 box_eq_empty - by auto - then show "content k *\<^sub>R (f' (x)) - (f ((Sup k)) - f ((Inf k))) = 0" - using xk unfolding uv by auto - next - have *: "p \ {t. fst t \ {a, b} \ content(snd t) \ 0} = - {t. t\p \ fst t = a \ content(snd t) \ 0} \ {t. t\p \ fst t = b \ content(snd t) \ 0}" - by blast - have **: "norm (setsum f s) \ e" - if "\x y. x \ s \ y \ s \ x = y" - and "\x. x \ s \ norm (f x) \ e" - and "e > 0" - for s f and e :: real - proof (cases "s = {}") - case True - with that show ?thesis by auto - next - case False - then obtain x where "x \ s" - by auto - then have *: "s = {x}" - using that(1) by auto - then show ?thesis - using \x \ s\ that(2) by auto - qed - case 2 - show ?case - apply (subst *) - apply (subst setsum.union_disjoint) - prefer 4 - apply (rule order_trans[of _ "e * (b - a)/4 + e * (b - a)/4"]) - apply (rule norm_triangle_le,rule add_mono) - apply (rule_tac[1-2] **) - proof - - let ?B = "\x. {t \ p. fst t = x \ content (snd t) \ 0}" - have pa: "\v. k = cbox a v \ a \ v" if "(a, k) \ p" for k - proof - - guess u v using p(4)[OF that] by (elim exE) note uv=this - have *: "u \ v" - using p(2)[OF that] unfolding uv by auto - have u: "u = a" - proof (rule ccontr) - have "u \ cbox u v" - using p(2-3)[OF that(1)] unfolding uv by auto - have "u \ a" - using p(2-3)[OF that(1)] unfolding uv subset_eq by auto - moreover assume "\ ?thesis" - ultimately have "u > a" by auto - then show False - using p(2)[OF that(1)] unfolding uv by (auto simp add:) - qed - then show ?thesis - apply (rule_tac x=v in exI) - unfolding uv - using * - apply auto - done - qed - have pb: "\v. k = cbox v b \ b \ v" if "(b, k) \ p" for k - proof - - guess u v using p(4)[OF that] by (elim exE) note uv=this - have *: "u \ v" - using p(2)[OF that] unfolding uv by auto - have u: "v = b" - proof (rule ccontr) - have "u \ cbox u v" - using p(2-3)[OF that(1)] unfolding uv by auto - have "v \ b" - using p(2-3)[OF that(1)] unfolding uv subset_eq by auto - moreover assume "\ ?thesis" - ultimately have "v < b" by auto - then show False - using p(2)[OF that(1)] unfolding uv by (auto simp add:) - qed - then show ?thesis - apply (rule_tac x=u in exI) - unfolding uv - using * - apply auto - done - qed - show "\x y. x \ ?B a \ y \ ?B a \ x = y" - apply (rule,rule,rule,unfold split_paired_all) - unfolding mem_Collect_eq fst_conv snd_conv - apply safe - proof - - fix x k k' - assume k: "(a, k) \ p" "(a, k') \ p" "content k \ 0" "content k' \ 0" - guess v using pa[OF k(1)] .. note v = conjunctD2[OF this] - guess v' using pa[OF k(2)] .. note v' = conjunctD2[OF this] let ?v = "min v v'" - have "box a ?v \ k \ k'" - unfolding v v' by (auto simp add: mem_box) - note interior_mono[OF this,unfolded interior_Int] - moreover have "(a + ?v)/2 \ box a ?v" - using k(3-) - unfolding v v' content_eq_0 not_le - by (auto simp add: mem_box) - ultimately have "(a + ?v)/2 \ interior k \ interior k'" - unfolding interior_open[OF open_box] by auto - then have *: "k = k'" - apply - - apply (rule ccontr) - using p(5)[OF k(1-2)] - apply auto - done - { assume "x \ k" then show "x \ k'" unfolding * . } - { assume "x \ k'" then show "x \ k" unfolding * . } - qed - show "\x y. x \ ?B b \ y \ ?B b \ x = y" - apply rule - apply rule - apply rule - apply (unfold split_paired_all) - unfolding mem_Collect_eq fst_conv snd_conv - apply safe - proof - - fix x k k' - assume k: "(b, k) \ p" "(b, k') \ p" "content k \ 0" "content k' \ 0" - guess v using pb[OF k(1)] .. note v = conjunctD2[OF this] - guess v' using pb[OF k(2)] .. note v' = conjunctD2[OF this] - let ?v = "max v v'" - have "box ?v b \ k \ k'" - unfolding v v' by (auto simp: mem_box) - note interior_mono[OF this,unfolded interior_Int] - moreover have " ((b + ?v)/2) \ box ?v b" - using k(3-) unfolding v v' content_eq_0 not_le by (auto simp: mem_box) - ultimately have " ((b + ?v)/2) \ interior k \ interior k'" - unfolding interior_open[OF open_box] by auto - then have *: "k = k'" - apply - - apply (rule ccontr) - using p(5)[OF k(1-2)] - apply auto - done - { assume "x \ k" then show "x \ k'" unfolding * . } - { assume "x \ k'" then show "x\k" unfolding * . } - qed - - let ?a = a and ?b = b (* a is something else while proofing the next theorem. *) - show "\x. x \ ?B a \ norm ((\(x, k). content k *\<^sub>R f' x - (f (Sup k) - - f (Inf k))) x) \ e * (b - a) / 4" - apply rule - apply rule - unfolding mem_Collect_eq - unfolding split_paired_all fst_conv snd_conv - proof (safe, goal_cases) - case prems: 1 - guess v using pa[OF prems(1)] .. note v = conjunctD2[OF this] - have "?a \ {?a..v}" - using v(2) by auto - then have "v \ ?b" - using p(3)[OF prems(1)] unfolding subset_eq v by auto - moreover have "{?a..v} \ ball ?a da" - using fineD[OF as(2) prems(1)] - apply - - apply (subst(asm) if_P) - apply (rule refl) - unfolding subset_eq - apply safe - apply (erule_tac x=" x" in ballE) - apply (auto simp add:subset_eq dist_real_def v) - done - ultimately show ?case - unfolding v interval_bounds_real[OF v(2)] box_real - apply - - apply(rule da(2)[of "v"]) - using prems fineD[OF as(2) prems(1)] - unfolding v content_eq_0 - apply auto - done - qed - show "\x. x \ ?B b \ norm ((\(x, k). content k *\<^sub>R f' x - - (f (Sup k) - f (Inf k))) x) \ e * (b - a) / 4" - apply rule - apply rule - unfolding mem_Collect_eq - unfolding split_paired_all fst_conv snd_conv - proof (safe, goal_cases) - case prems: 1 - guess v using pb[OF prems(1)] .. note v = conjunctD2[OF this] - have "?b \ {v.. ?b}" - using v(2) by auto - then have "v \ ?a" using p(3)[OF prems(1)] - unfolding subset_eq v by auto - moreover have "{v..?b} \ ball ?b db" - using fineD[OF as(2) prems(1)] - apply - - apply (subst(asm) if_P, rule refl) - unfolding subset_eq - apply safe - apply (erule_tac x=" x" in ballE) - using ab - apply (auto simp add:subset_eq v dist_real_def) - done - ultimately show ?case - unfolding v - unfolding interval_bounds_real[OF v(2)] box_real - apply - - apply(rule db(2)[of "v"]) - using prems fineD[OF as(2) prems(1)] - unfolding v content_eq_0 - apply auto - done - qed - qed (insert p(1) ab e, auto simp add: field_simps) - qed auto - qed - qed - qed -qed - - -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}" - using assms -proof (induct "card s" arbitrary: s a b) - case 0 - show ?case - apply (rule fundamental_theorem_of_calculus_interior) - using 0 - apply auto - done -next - case (Suc n) - from this(2) guess c s' - apply - - apply (subst(asm) eq_commute) - unfolding card_Suc_eq - apply (subst(asm)(2) eq_commute) - apply (elim exE conjE) - done - note cs = this[rule_format] - show ?case - proof (cases "c \ box a b") - case False - then show ?thesis - apply - - apply (rule Suc(1)[OF cs(3) _ Suc(4,5)]) - apply safe - defer - apply (rule Suc(6)[rule_format]) - using Suc(3) - unfolding cs - apply auto - done - next - have *: "f b - f a = (f c - f a) + (f b - f c)" - by auto - case True - then have "a \ c" "c \ b" - by (auto simp: mem_box) - then show ?thesis - apply (subst *) - apply (rule has_integral_combine) - apply assumption+ - apply (rule_tac[!] Suc(1)[OF cs(3)]) - using Suc(3) - unfolding cs - proof - - show "continuous_on {a .. c} f" "continuous_on {c .. b} f" - apply (rule_tac[!] continuous_on_subset[OF Suc(5)]) - using True - apply (auto simp: mem_box) - done - let ?P = "\i j. \x\{i <..< j} - s'. (f has_vector_derivative f' x) (at x)" - show "?P a c" "?P c b" - apply safe - apply (rule_tac[!] Suc(6)[rule_format]) - using True - unfolding cs - apply (auto simp: mem_box) - done - qed auto - qed -qed - -lemma fundamental_theorem_of_calculus_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}" - apply (rule fundamental_theorem_of_calculus_interior_strong[OF assms(1-3), of f']) - using assms(4) - apply (auto simp: mem_box) - done - -lemma indefinite_integral_continuous_left: - fixes f:: "real \ 'a::banach" - assumes "f integrable_on {a .. b}" - and "a < c" - and "c \ b" - and "e > 0" - obtains d where "d > 0" - and "\t. c - d < t \ t \ c \ norm (integral {a .. c} f - integral {a .. t} f) < e" -proof - - have "\w>0. \t. c - w < t \ t < c \ norm (f c) * norm(c - t) < e / 3" - proof (cases "f c = 0") - case False - hence "0 < e / 3 / norm (f c)" using \e>0\ by simp - then show ?thesis - apply - - apply rule - apply rule - apply assumption - apply safe - proof - - fix t - assume as: "t < c" and "c - e / 3 / norm (f c) < t" - then have "c - t < e / 3 / norm (f c)" - by auto - then have "norm (c - t) < e / 3 / norm (f c)" - using as by auto - then show "norm (f c) * norm (c - t) < e / 3" - using False - apply - - apply (subst mult.commute) - apply (subst pos_less_divide_eq[symmetric]) - apply auto - done - qed - next - case True - show ?thesis - apply (rule_tac x=1 in exI) - unfolding True - using \e > 0\ - apply auto - done - qed - then guess w .. note w = conjunctD2[OF this,rule_format] - - have *: "e / 3 > 0" - using assms by auto - have "f integrable_on {a .. c}" - apply (rule integrable_subinterval_real[OF assms(1)]) - using assms(2-3) - apply auto - done - from integrable_integral[OF this,unfolded has_integral_real,rule_format,OF *] guess d1 .. - note d1 = conjunctD2[OF this,rule_format] - define d where [abs_def]: "d x = ball x w \ d1 x" for x - have "gauge d" - unfolding d_def using w(1) d1 by auto - note this[unfolded gauge_def,rule_format,of c] - note conjunctD2[OF this] - from this(2)[unfolded open_contains_ball,rule_format,OF this(1)] guess k .. - note k=conjunctD2[OF this] - - let ?d = "min k (c - a) / 2" - show ?thesis - apply (rule that[of ?d]) - apply safe - proof - - show "?d > 0" - using k(1) using assms(2) by auto - fix t - assume as: "c - ?d < t" "t \ c" - let ?thesis = "norm (integral ({a .. c}) f - integral ({a .. t}) f) < e" - { - presume *: "t < c \ ?thesis" - show ?thesis - apply (cases "t = c") - defer - apply (rule *) - apply (subst less_le) - using \e > 0\ as(2) - apply auto - done - } - assume "t < c" - - have "f integrable_on {a .. t}" - apply (rule integrable_subinterval_real[OF assms(1)]) - using assms(2-3) as(2) - apply auto - done - from integrable_integral[OF this,unfolded has_integral_real,rule_format,OF *] guess d2 .. - note d2 = conjunctD2[OF this,rule_format] - define d3 where "d3 x = (if x \ t then d1 x \ d2 x else d1 x)" for x - have "gauge d3" - using d2(1) d1(1) unfolding d3_def gauge_def by auto - from fine_division_exists_real[OF this, of a t] guess p . note p=this - note p'=tagged_division_ofD[OF this(1)] - have pt: "\(x,k)\p. x \ t" - proof (safe, goal_cases) - case prems: 1 - from p'(2,3)[OF prems] show ?case - by auto - qed - with p(2) have "d2 fine p" - unfolding fine_def d3_def - apply safe - apply (erule_tac x="(a,b)" in ballE)+ - apply auto - done - note d2_fin = d2(2)[OF conjI[OF p(1) this]] - - have *: "{a .. c} \ {x. x \ 1 \ t} = {a .. t}" "{a .. c} \ {x. x \ 1 \ t} = {t .. c}" - using assms(2-3) as by (auto simp add: field_simps) - have "p \ {(c, {t .. c})} tagged_division_of {a .. c} \ d1 fine p \ {(c, {t .. c})}" - apply rule - apply (rule tagged_division_union_interval_real[of _ _ _ 1 "t"]) - unfolding * - apply (rule p) - apply (rule tagged_division_of_self_real) - unfolding fine_def - apply safe - proof - - fix x k y - assume "(x,k) \ p" and "y \ k" - then show "y \ d1 x" - using p(2) pt - unfolding fine_def d3_def - apply - - apply (erule_tac x="(x,k)" in ballE)+ - apply auto - done - next - fix x assume "x \ {t..c}" - then have "dist c x < k" - unfolding dist_real_def - using as(1) - by (auto simp add: field_simps) - then show "x \ d1 c" - using k(2) - unfolding d_def - by auto - qed (insert as(2), auto) note d1_fin = d1(2)[OF this] - - have *: "integral {a .. c} f - integral {a .. t} f = -(((c - t) *\<^sub>R f c + (\(x, k)\p. content k *\<^sub>R f x)) - - integral {a .. c} f) + ((\(x, k)\p. content k *\<^sub>R f x) - integral {a .. t} f) + (c - t) *\<^sub>R f c" - "e = (e/3 + e/3) + e/3" - by auto - have **: "(\(x, k)\p \ {(c, {t .. c})}. content k *\<^sub>R f x) = - (c - t) *\<^sub>R f c + (\(x, k)\p. content k *\<^sub>R f x)" - proof - - have **: "\x F. F \ {x} = insert x F" - by auto - have "(c, cbox t c) \ p" - proof (safe, goal_cases) - case prems: 1 - from p'(2-3)[OF prems] have "c \ cbox a t" - by auto - then show False using \t < c\ - by auto - qed - then show ?thesis - unfolding ** box_real - apply - - apply (subst setsum.insert) - apply (rule p') - unfolding split_conv - defer - apply (subst content_real) - using as(2) - apply auto - done - qed - have ***: "c - w < t \ t < c" - proof - - have "c - k < t" - using \k>0\ as(1) by (auto simp add: field_simps) - moreover have "k \ w" - apply (rule ccontr) - using k(2) - unfolding subset_eq - apply (erule_tac x="c + ((k + w)/2)" in ballE) - unfolding d_def - using \k > 0\ \w > 0\ - apply (auto simp add: field_simps not_le not_less dist_real_def) - done - ultimately show ?thesis using \t < c\ - by (auto simp add: field_simps) - qed - show ?thesis - unfolding *(1) - apply (subst *(2)) - apply (rule norm_triangle_lt add_strict_mono)+ - unfolding norm_minus_cancel - apply (rule d1_fin[unfolded **]) - apply (rule d2_fin) - using w(2)[OF ***] - unfolding norm_scaleR - apply (auto simp add: field_simps) - done - qed -qed - -lemma indefinite_integral_continuous_right: - fixes f :: "real \ 'a::banach" - assumes "f integrable_on {a .. b}" - and "a \ c" - and "c < b" - and "e > 0" - obtains d where "0 < d" - and "\t. c \ t \ t < c + d \ norm (integral {a .. c} f - integral {a .. t} f) < e" -proof - - have *: "(\x. f (- x)) integrable_on {-b .. -a}" "- b < - c" "- c \ - a" - using assms by auto - from indefinite_integral_continuous_left[OF * \e>0\] guess d . note d = this - let ?d = "min d (b - c)" - show ?thesis - apply (rule that[of "?d"]) - apply safe - proof - - show "0 < ?d" - using d(1) assms(3) by auto - fix t :: real - assume as: "c \ t" "t < c + ?d" - have *: "integral {a .. c} f = integral {a .. b} f - integral {c .. b} f" - "integral {a .. t} f = integral {a .. b} f - integral {t .. b} f" - apply (simp_all only: algebra_simps) - apply (rule_tac[!] integral_combine) - using assms as - apply auto - done - have "(- c) - d < (- t) \ - t \ - c" - using as by auto note d(2)[rule_format,OF this] - then show "norm (integral {a .. c} f - integral {a .. t} f) < e" - unfolding * - unfolding integral_reflect - apply (subst norm_minus_commute) - apply (auto simp add: algebra_simps) - done - qed -qed - -lemma indefinite_integral_continuous: - fixes f :: "real \ 'a::banach" - assumes "f integrable_on {a .. b}" - shows "continuous_on {a .. b} (\x. integral {a .. x} f)" -proof (unfold continuous_on_iff, safe) - fix x e :: real - assume as: "x \ {a .. b}" "e > 0" - let ?thesis = "\d>0. \x'\{a .. b}. dist x' x < d \ dist (integral {a .. x'} f) (integral {a .. x} f) < e" - { - presume *: "a < b \ ?thesis" - show ?thesis - apply cases - apply (rule *) - apply assumption - proof goal_cases - case 1 - then have "cbox a b = {x}" - using as(1) - apply - - apply (rule set_eqI) - apply auto - done - then show ?case using \e > 0\ by auto - qed - } - assume "a < b" - have "(x = a \ x = b) \ (a < x \ x < b)" - using as(1) by auto - then show ?thesis - apply (elim disjE) - proof - - assume "x = a" - have "a \ a" .. - from indefinite_integral_continuous_right[OF assms(1) this \a \e>0\] guess d . note d=this - show ?thesis - apply rule - apply rule - apply (rule d) - apply safe - apply (subst dist_commute) - unfolding \x = a\ dist_norm - apply (rule d(2)[rule_format]) - apply auto - done - next - assume "x = b" - have "b \ b" .. - from indefinite_integral_continuous_left[OF assms(1) \a this \e>0\] guess d . note d=this - show ?thesis - apply rule - apply rule - apply (rule d) - apply safe - apply (subst dist_commute) - unfolding \x = b\ dist_norm - apply (rule d(2)[rule_format]) - apply auto - done - next - assume "a < x \ x < b" - then have xl: "a < x" "x \ b" and xr: "a \ x" "x < b" - by auto - from indefinite_integral_continuous_left [OF assms(1) xl \e>0\] guess d1 . note d1=this - from indefinite_integral_continuous_right[OF assms(1) xr \e>0\] guess d2 . note d2=this - show ?thesis - apply (rule_tac x="min d1 d2" in exI) - proof safe - show "0 < min d1 d2" - using d1 d2 by auto - fix y - assume "y \ {a .. b}" and "dist y x < min d1 d2" - then show "dist (integral {a .. y} f) (integral {a .. x} f) < e" - apply (subst dist_commute) - apply (cases "y < x") - unfolding dist_norm - apply (rule d1(2)[rule_format]) - defer - apply (rule d2(2)[rule_format]) - unfolding not_less - apply (auto simp add: field_simps) - done - qed - qed -qed - - -subsection \This doesn't directly involve integration, but that gives an easy proof.\ - -lemma has_derivative_zero_unique_strong_interval: - fixes f :: "real \ 'a::banach" - assumes "finite k" - and "continuous_on {a .. b} f" - and "f a = y" - and "\x\({a .. b} - k). (f has_derivative (\h. 0)) (at x within {a .. b})" "x \ {a .. b}" - shows "f x = y" -proof - - have ab: "a \ b" - using assms by auto - have *: "a \ x" - using assms(5) by auto - have "((\x. 0::'a) has_integral f x - f a) {a .. x}" - apply (rule fundamental_theorem_of_calculus_interior_strong[OF assms(1) *]) - apply (rule continuous_on_subset[OF assms(2)]) - defer - apply safe - unfolding has_vector_derivative_def - apply (subst has_derivative_within_open[symmetric]) - apply assumption - apply (rule open_greaterThanLessThan) - apply (rule has_derivative_within_subset[where s="{a .. b}"]) - using assms(4) assms(5) - apply (auto simp: mem_box) - done - note this[unfolded *] - note has_integral_unique[OF has_integral_0 this] - then show ?thesis - unfolding assms by auto -qed - - -subsection \Generalize a bit to any convex set.\ - -lemma has_derivative_zero_unique_strong_convex: - fixes f :: "'a::euclidean_space \ 'b::banach" - assumes "convex s" - and "finite k" - and "continuous_on s f" - and "c \ s" - and "f c = y" - and "\x\(s - k). (f has_derivative (\h. 0)) (at x within s)" - and "x \ s" - shows "f x = y" -proof - - { - presume *: "x \ c \ ?thesis" - show ?thesis - apply cases - apply (rule *) - apply assumption - unfolding assms(5)[symmetric] - apply auto - done - } - assume "x \ c" - note conv = assms(1)[unfolded convex_alt,rule_format] - have as1: "continuous_on {0 ..1} (f \ (\t. (1 - t) *\<^sub>R c + t *\<^sub>R x))" - apply (rule continuous_intros)+ - apply (rule continuous_on_subset[OF assms(3)]) - apply safe - apply (rule conv) - using assms(4,7) - apply auto - done - have *: "t = xa" if "(1 - t) *\<^sub>R c + t *\<^sub>R x = (1 - xa) *\<^sub>R c + xa *\<^sub>R x" for t xa - proof - - from that have "(t - xa) *\<^sub>R x = (t - xa) *\<^sub>R c" - unfolding scaleR_simps by (auto simp add: algebra_simps) - then show ?thesis - using \x \ c\ by auto - qed - have as2: "finite {t. ((1 - t) *\<^sub>R c + t *\<^sub>R x) \ k}" - using assms(2) - apply (rule finite_surj[where f="\z. SOME t. (1-t) *\<^sub>R c + t *\<^sub>R x = z"]) - apply safe - unfolding image_iff - apply rule - defer - apply assumption - apply (rule sym) - apply (rule some_equality) - defer - apply (drule *) - apply auto - done - have "(f \ (\t. (1 - t) *\<^sub>R c + t *\<^sub>R x)) 1 = y" - apply (rule has_derivative_zero_unique_strong_interval[OF as2 as1, of ]) - unfolding o_def - using assms(5) - defer - apply - - apply rule - proof - - fix t - assume as: "t \ {0 .. 1} - {t. (1 - t) *\<^sub>R c + t *\<^sub>R x \ k}" - have *: "c - t *\<^sub>R c + t *\<^sub>R x \ s - k" - apply safe - apply (rule conv[unfolded scaleR_simps]) - using \x \ s\ \c \ s\ as - by (auto simp add: algebra_simps) - have "(f \ (\t. (1 - t) *\<^sub>R c + t *\<^sub>R x) has_derivative (\x. 0) \ (\z. (0 - z *\<^sub>R c) + z *\<^sub>R x)) - (at t within {0 .. 1})" - apply (intro derivative_eq_intros) - apply simp_all - apply (simp add: field_simps) - unfolding scaleR_simps - apply (rule has_derivative_within_subset,rule assms(6)[rule_format]) - apply (rule *) - apply safe - apply (rule conv[unfolded scaleR_simps]) - using \x \ s\ \c \ s\ - apply auto - done - then show "((\xa. f ((1 - xa) *\<^sub>R c + xa *\<^sub>R x)) has_derivative (\h. 0)) (at t within {0 .. 1})" - unfolding o_def . - qed auto - then show ?thesis - by auto -qed - - -text \Also to any open connected set with finite set of exceptions. Could - generalize to locally convex set with limpt-free set of exceptions.\ - -lemma has_derivative_zero_unique_strong_connected: - fixes f :: "'a::euclidean_space \ 'b::banach" - assumes "connected s" - and "open s" - and "finite k" - and "continuous_on s f" - and "c \ s" - and "f c = y" - and "\x\(s - k). (f has_derivative (\h. 0)) (at x within s)" - and "x\s" - shows "f x = y" -proof - - have "{x \ s. f x \ {y}} = {} \ {x \ s. f x \ {y}} = s" - apply (rule assms(1)[unfolded connected_clopen,rule_format]) - apply rule - defer - apply (rule continuous_closedin_preimage[OF assms(4) closed_singleton]) - apply (rule open_openin_trans[OF assms(2)]) - unfolding open_contains_ball - proof safe - fix x - assume "x \ s" - from assms(2)[unfolded open_contains_ball,rule_format,OF this] guess e .. note e=conjunctD2[OF this] - show "\e>0. ball x e \ {xa \ s. f xa \ {f x}}" - apply rule - apply rule - apply (rule e) - proof safe - fix y - assume y: "y \ ball x e" - then show "y \ s" - using e by auto - show "f y = f x" - apply (rule has_derivative_zero_unique_strong_convex[OF convex_ball]) - apply (rule assms) - apply (rule continuous_on_subset) - apply (rule assms) - apply (rule e)+ - apply (subst centre_in_ball) - apply (rule e) - apply rule - apply safe - apply (rule has_derivative_within_subset) - apply (rule assms(7)[rule_format]) - using y e - apply auto - done - qed - qed - then show ?thesis - using \x \ s\ \f c = y\ \c \ s\ by auto -qed - -lemma has_derivative_zero_connected_constant: - fixes f :: "'a::euclidean_space \ 'b::banach" - assumes "connected s" - and "open s" - and "finite k" - and "continuous_on s f" - and "\x\(s - k). (f has_derivative (\h. 0)) (at x within s)" - obtains c where "\x. x \ s \ f(x) = c" -proof (cases "s = {}") - case True - then show ?thesis -by (metis empty_iff that) -next - case False - then obtain c where "c \ s" - by (metis equals0I) - then show ?thesis - by (metis has_derivative_zero_unique_strong_connected assms that) -qed - - -subsection \Integrating characteristic function of an interval\ - -lemma has_integral_restrict_open_subinterval: - fixes f :: "'a::euclidean_space \ 'b::banach" - assumes "(f has_integral i) (cbox c d)" - and "cbox c d \ cbox a b" - shows "((\x. if x \ box c d then f x else 0) has_integral i) (cbox a b)" -proof - - define g where [abs_def]: "g x = (if x \box c d then f x else 0)" for x - { - presume *: "cbox c d \ {} \ ?thesis" - show ?thesis - apply cases - apply (rule *) - apply assumption - proof goal_cases - case prems: 1 - then have *: "box c d = {}" - by (metis bot.extremum_uniqueI box_subset_cbox) - show ?thesis - using assms(1) - unfolding * - using prems - by auto - qed - } - assume "cbox c d \ {}" - from partial_division_extend_1[OF assms(2) this] guess p . note p=this - note operat = comm_monoid_set.operative_division[OF add.comm_monoid_set_lift_option operative_integral p(1), symmetric] - let ?P = "(if g integrable_on cbox a b then Some (integral (cbox a b) g) else None) = Some i" - { - presume "?P" - then have "g integrable_on cbox a b \ integral (cbox a b) g = i" - apply - - apply cases - apply (subst(asm) if_P) - apply assumption - apply auto - done - then show ?thesis - using integrable_integral - unfolding g_def - by auto - } - let ?F = "comm_monoid_set.F (lift_option op +) (Some 0)" - have iterate:"?F (\i. if g integrable_on i then Some (integral i g) else None) (p - {cbox c d}) = Some 0" - proof (intro comm_monoid_set.neutral[OF add.comm_monoid_set_lift_option] ballI) - fix x - assume x: "x \ p - {cbox c d}" - then have "x \ p" - by auto - note div = division_ofD(2-5)[OF p(1) this] - from div(3) guess u v by (elim exE) note uv=this - have "interior x \ interior (cbox c d) = {}" - using div(4)[OF p(2)] x by auto - then have "(g has_integral 0) x" - unfolding uv - apply - - apply (rule has_integral_spike_interior[where f="\x. 0"]) - unfolding g_def interior_cbox - apply auto - done - then show "(if g integrable_on x then Some (integral x g) else None) = Some 0" - by auto - qed - - have *: "p = insert (cbox c d) (p - {cbox c d})" - using p by auto - have **: "g integrable_on cbox c d" - apply (rule integrable_spike_interior[where f=f]) - unfolding g_def using assms(1) - apply auto - done - moreover - have "integral (cbox c d) g = i" - apply (rule has_integral_unique[OF _ assms(1)]) - apply (rule has_integral_spike_interior[where f=g]) - defer - apply (rule integrable_integral[OF **]) - unfolding g_def - apply auto - done - ultimately show ?P - unfolding operat - using p - apply (subst *) - apply (subst comm_monoid_set.insert[OF add.comm_monoid_set_lift_option]) - apply (simp_all add: division_of_finite iterate) - done -qed - -lemma has_integral_restrict_closed_subinterval: - fixes f :: "'a::euclidean_space \ 'b::banach" - assumes "(f has_integral i) (cbox c d)" - and "cbox c d \ cbox a b" - shows "((\x. if x \ cbox c d then f x else 0) has_integral i) (cbox a b)" -proof - - note has_integral_restrict_open_subinterval[OF assms] - note * = has_integral_spike[OF negligible_frontier_interval _ this] - show ?thesis - apply (rule *[of c d]) - using box_subset_cbox[of c d] - apply auto - done -qed - -lemma has_integral_restrict_closed_subintervals_eq: - fixes f :: "'a::euclidean_space \ 'b::banach" - assumes "cbox c d \ cbox a b" - shows "((\x. if x \ cbox c d then f x else 0) has_integral i) (cbox a b) \ (f has_integral i) (cbox c d)" - (is "?l = ?r") -proof (cases "cbox c d = {}") - case False - let ?g = "\x. if x \ cbox c d then f x else 0" - show ?thesis - apply rule - defer - apply (rule has_integral_restrict_closed_subinterval[OF _ assms]) - apply assumption - proof - - assume ?l - then have "?g integrable_on cbox c d" - using assms has_integral_integrable integrable_subinterval by blast - then have *: "f integrable_on cbox c d" - apply - - apply (rule integrable_eq) - apply auto - done - then have "i = integral (cbox c d) f" - apply - - apply (rule has_integral_unique) - apply (rule \?l\) - apply (rule has_integral_restrict_closed_subinterval[OF _ assms]) - apply auto - done - then show ?r - using * by auto - qed -qed auto - - -text \Hence we can apply the limit process uniformly to all integrals.\ - -lemma has_integral': - fixes f :: "'n::euclidean_space \ 'a::banach" - shows "(f has_integral i) s \ - (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ - (\z. ((\x. if x \ s then f(x) else 0) has_integral z) (cbox a b) \ norm(z - i) < e))" - (is "?l \ (\e>0. ?r e)") -proof - - { - presume *: "\a b. s = cbox a b \ ?thesis" - show ?thesis - apply cases - apply (rule *) - apply assumption - apply (subst has_integral_alt) - apply auto - done - } - assume "\a b. s = cbox a b" - then guess a b by (elim exE) note s=this - from bounded_cbox[of a b, unfolded bounded_pos] guess B .. - note B = conjunctD2[OF this,rule_format] show ?thesis - apply safe - proof - - fix e :: real - assume ?l and "e > 0" - show "?r e" - apply (rule_tac x="B+1" in exI) - apply safe - defer - apply (rule_tac x=i in exI) - proof - fix c d :: 'n - assume as: "ball 0 (B+1) \ cbox c d" - then show "((\x. if x \ s then f x else 0) has_integral i) (cbox c d)" - unfolding s - apply - - apply (rule has_integral_restrict_closed_subinterval) - apply (rule \?l\[unfolded s]) - apply safe - apply (drule B(2)[rule_format]) - unfolding subset_eq - apply (erule_tac x=x in ballE) - apply (auto simp add: dist_norm) - done - qed (insert B \e>0\, auto) - next - assume as: "\e>0. ?r e" - from this[rule_format,OF zero_less_one] guess C .. note C=conjunctD2[OF this,rule_format] - define c :: 'n where "c = (\i\Basis. (- max B C) *\<^sub>R i)" - define d :: 'n where "d = (\i\Basis. max B C *\<^sub>R i)" - have c_d: "cbox a b \ cbox c d" - apply safe - apply (drule B(2)) - unfolding mem_box - proof - fix x i - show "c \ i \ x \ i \ x \ i \ d \ i" if "norm x \ B" and "i \ Basis" - using that and Basis_le_norm[OF \i\Basis\, of x] - unfolding c_def d_def - by (auto simp add: field_simps setsum_negf) - qed - have "ball 0 C \ cbox c d" - apply (rule subsetI) - unfolding mem_box mem_ball dist_norm - proof - fix x i :: 'n - assume x: "norm (0 - x) < C" and i: "i \ Basis" - show "c \ i \ x \ i \ x \ i \ d \ i" - using Basis_le_norm[OF i, of x] and x i - unfolding c_def d_def - by (auto simp: setsum_negf) - qed - from C(2)[OF this] have "\y. (f has_integral y) (cbox a b)" - unfolding has_integral_restrict_closed_subintervals_eq[OF c_d,symmetric] - unfolding s - by auto - then guess y .. note y=this - - have "y = i" - proof (rule ccontr) - assume "\ ?thesis" - then have "0 < norm (y - i)" - by auto - from as[rule_format,OF this] guess C .. note C=conjunctD2[OF this,rule_format] - define c :: 'n where "c = (\i\Basis. (- max B C) *\<^sub>R i)" - define d :: 'n where "d = (\i\Basis. max B C *\<^sub>R i)" - have c_d: "cbox a b \ cbox c d" - apply safe - apply (drule B(2)) - unfolding mem_box - proof - fix x i :: 'n - assume "norm x \ B" and "i \ Basis" - then show "c \ i \ x \ i \ x \ i \ d \ i" - using Basis_le_norm[of i x] - unfolding c_def d_def - by (auto simp add: field_simps setsum_negf) - qed - have "ball 0 C \ cbox c d" - apply (rule subsetI) - unfolding mem_box mem_ball dist_norm - proof - fix x i :: 'n - assume "norm (0 - x) < C" and "i \ Basis" - then show "c \ i \ x \ i \ x \ i \ d \ i" - using Basis_le_norm[of i x] - unfolding c_def d_def - by (auto simp: setsum_negf) - qed - note C(2)[OF this] then guess z .. note z = conjunctD2[OF this, unfolded s] - note this[unfolded has_integral_restrict_closed_subintervals_eq[OF c_d]] - then have "z = y" and "norm (z - i) < norm (y - i)" - apply - - apply (rule has_integral_unique[OF _ y(1)]) - apply assumption - apply assumption - done - then show False - by auto - qed - then show ?l - using y - unfolding s - by auto - qed -qed - -lemma has_integral_le: - fixes f :: "'n::euclidean_space \ real" - assumes "(f has_integral i) s" - and "(g has_integral j) s" - and "\x\s. f x \ g x" - shows "i \ j" - using has_integral_component_le[OF _ assms(1-2), of 1] - using assms(3) - by auto - -lemma integral_le: - fixes f :: "'n::euclidean_space \ real" - assumes "f integrable_on s" - and "g integrable_on s" - and "\x\s. f x \ g x" - shows "integral s f \ integral s g" - by (rule has_integral_le[OF assms(1,2)[unfolded has_integral_integral] assms(3)]) - -lemma has_integral_nonneg: - fixes f :: "'n::euclidean_space \ real" - assumes "(f has_integral i) s" - and "\x\s. 0 \ f x" - shows "0 \ i" - using has_integral_component_nonneg[of 1 f i s] - unfolding o_def - using assms - by auto - -lemma integral_nonneg: - fixes f :: "'n::euclidean_space \ real" - assumes "f integrable_on s" - and "\x\s. 0 \ f x" - shows "0 \ integral s f" - by (rule has_integral_nonneg[OF assms(1)[unfolded has_integral_integral] assms(2)]) - - -text \Hence a general restriction property.\ - -lemma has_integral_restrict[simp]: - assumes "s \ t" - shows "((\x. if x \ s then f x else (0::'a::banach)) has_integral i) t \ (f has_integral i) s" -proof - - have *: "\x. (if x \ t then if x \ s then f x else 0 else 0) = (if x\s then f x else 0)" - using assms by auto - show ?thesis - apply (subst(2) has_integral') - apply (subst has_integral') - unfolding * - apply rule - done -qed - -lemma has_integral_restrict_univ: - fixes f :: "'n::euclidean_space \ 'a::banach" - shows "((\x. if x \ s then f x else 0) has_integral i) UNIV \ (f has_integral i) s" - by auto - -lemma has_integral_on_superset: - fixes f :: "'n::euclidean_space \ 'a::banach" - assumes "\x. x \ s \ f x = 0" - and "s \ t" - and "(f has_integral i) s" - shows "(f has_integral i) t" -proof - - have "(\x. if x \ s then f x else 0) = (\x. if x \ t then f x else 0)" - apply rule - using assms(1-2) - apply auto - done - then show ?thesis - using assms(3) - apply (subst has_integral_restrict_univ[symmetric]) - apply (subst(asm) has_integral_restrict_univ[symmetric]) - apply auto - done -qed - -lemma integrable_on_superset: - fixes f :: "'n::euclidean_space \ 'a::banach" - assumes "\x. x \ s \ f x = 0" - and "s \ t" - and "f integrable_on s" - shows "f integrable_on t" - using assms - unfolding integrable_on_def - by (auto intro:has_integral_on_superset) - -lemma integral_restrict_univ[intro]: - fixes f :: "'n::euclidean_space \ 'a::banach" - shows "f integrable_on s \ integral UNIV (\x. if x \ s then f x else 0) = integral s f" - apply (rule integral_unique) - unfolding has_integral_restrict_univ - apply auto - done - -lemma integrable_restrict_univ: - fixes f :: "'n::euclidean_space \ 'a::banach" - shows "(\x. if x \ s then f x else 0) integrable_on UNIV \ f integrable_on s" - unfolding integrable_on_def - by auto - -lemma negligible_on_intervals: "negligible s \ (\a b. negligible(s \ cbox a b))" (is "?l \ ?r") -proof - assume ?r - show ?l - unfolding negligible_def - proof safe - fix a b - show "(indicator s has_integral 0) (cbox a b)" - apply (rule has_integral_negligible[OF \?r\[rule_format,of a b]]) - unfolding indicator_def - apply auto - done - qed -qed auto - -lemma negligible_translation: - assumes "negligible S" - shows "negligible (op + c ` S)" -proof - - have inj: "inj (op + c)" - by simp - show ?thesis - using assms - proof (clarsimp simp: negligible_def) - fix a b - assume "\x y. (indicator S has_integral 0) (cbox x y)" - then have *: "(indicator S has_integral 0) (cbox (a-c) (b-c))" - by (meson Diff_iff assms has_integral_negligible indicator_simps(2)) - have eq: "indicator (op + c ` S) = (\x. indicator S (x - c))" - by (force simp add: indicator_def) - show "(indicator (op + c ` S) has_integral 0) (cbox a b)" - using has_integral_affinity [OF *, of 1 "-c"] - cbox_translation [of "c" "-c+a" "-c+b"] - by (simp add: eq add.commute) - qed -qed - -lemma negligible_translation_rev: - assumes "negligible (op + c ` S)" - shows "negligible S" -by (metis negligible_translation [OF assms, of "-c"] translation_galois) - -lemma has_integral_spike_set_eq: - fixes f :: "'n::euclidean_space \ 'a::banach" - assumes "negligible ((s - t) \ (t - s))" - shows "(f has_integral y) s \ (f has_integral y) t" - unfolding has_integral_restrict_univ[symmetric,of f] - apply (rule has_integral_spike_eq[OF assms]) - by (auto split: if_split_asm) - -lemma has_integral_spike_set[dest]: - fixes f :: "'n::euclidean_space \ 'a::banach" - assumes "negligible ((s - t) \ (t - s))" - and "(f has_integral y) s" - shows "(f has_integral y) t" - using assms has_integral_spike_set_eq - by auto - -lemma integrable_spike_set[dest]: - fixes f :: "'n::euclidean_space \ 'a::banach" - assumes "negligible ((s - t) \ (t - s))" - and "f integrable_on s" - shows "f integrable_on t" - using assms(2) - unfolding integrable_on_def - unfolding has_integral_spike_set_eq[OF assms(1)] . - -lemma integrable_spike_set_eq: - fixes f :: "'n::euclidean_space \ 'a::banach" - assumes "negligible ((s - t) \ (t - s))" - shows "f integrable_on s \ f integrable_on t" - apply rule - apply (rule_tac[!] integrable_spike_set) - using assms - apply auto - done - -(*lemma integral_spike_set: - "\f:real^M->real^N g s t. - negligible(s DIFF t \ t DIFF s) - \ integral s f = integral t f" -qed REPEAT STRIP_TAC THEN REWRITE_TAC[integral] THEN - AP_TERM_TAC THEN ABS_TAC THEN MATCH_MP_TAC HAS_INTEGRAL_SPIKE_SET_EQ THEN - ASM_MESON_TAC[]);; - -lemma has_integral_interior: - "\f:real^M->real^N y s. - negligible(frontier s) - \ ((f has_integral y) (interior s) \ (f has_integral y) s)" -qed REPEAT STRIP_TAC THEN MATCH_MP_TAC HAS_INTEGRAL_SPIKE_SET_EQ THEN - FIRST_X_ASSUM(MATCH_MP_TAC o MATCH_MP (REWRITE_RULE[IMP_CONJ] - NEGLIGIBLE_SUBSET)) THEN - REWRITE_TAC[frontier] THEN - MP_TAC(ISPEC `s:real^M->bool` INTERIOR_SUBSET) THEN - MP_TAC(ISPEC `s:real^M->bool` CLOSURE_SUBSET) THEN - SET_TAC[]);; - -lemma has_integral_closure: - "\f:real^M->real^N y s. - negligible(frontier s) - \ ((f has_integral y) (closure s) \ (f has_integral y) s)" -qed REPEAT STRIP_TAC THEN MATCH_MP_TAC HAS_INTEGRAL_SPIKE_SET_EQ THEN - FIRST_X_ASSUM(MATCH_MP_TAC o MATCH_MP (REWRITE_RULE[IMP_CONJ] - NEGLIGIBLE_SUBSET)) THEN - REWRITE_TAC[frontier] THEN - MP_TAC(ISPEC `s:real^M->bool` INTERIOR_SUBSET) THEN - MP_TAC(ISPEC `s:real^M->bool` CLOSURE_SUBSET) THEN - SET_TAC[]);;*) - - -subsection \More lemmas that are useful later\ - -lemma has_integral_subset_component_le: - fixes f :: "'n::euclidean_space \ 'm::euclidean_space" - assumes k: "k \ Basis" - and as: "s \ t" "(f has_integral i) s" "(f has_integral j) t" "\x\t. 0 \ f(x)\k" - shows "i\k \ j\k" -proof - - note has_integral_restrict_univ[symmetric, of f] - note as(2-3)[unfolded this] note * = has_integral_component_le[OF k this] - show ?thesis - apply (rule *) - using as(1,4) - apply auto - done -qed - -lemma has_integral_subset_le: - fixes f :: "'n::euclidean_space \ real" - assumes "s \ t" - and "(f has_integral i) s" - and "(f has_integral j) t" - and "\x\t. 0 \ f x" - shows "i \ j" - using has_integral_subset_component_le[OF _ assms(1), of 1 f i j] - using assms - by auto - -lemma integral_subset_component_le: - fixes f :: "'n::euclidean_space \ 'm::euclidean_space" - assumes "k \ Basis" - and "s \ t" - and "f integrable_on s" - and "f integrable_on t" - and "\x \ t. 0 \ f x \ k" - shows "(integral s f)\k \ (integral t f)\k" - apply (rule has_integral_subset_component_le) - using assms - apply auto - done - -lemma integral_subset_le: - fixes f :: "'n::euclidean_space \ real" - assumes "s \ t" - and "f integrable_on s" - and "f integrable_on t" - and "\x \ t. 0 \ f x" - shows "integral s f \ integral t f" - apply (rule has_integral_subset_le) - using assms - apply auto - done - -lemma has_integral_alt': - fixes f :: "'n::euclidean_space \ 'a::banach" - shows "(f has_integral i) s \ (\a b. (\x. if x \ s then f x else 0) integrable_on cbox a b) \ - (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ - norm (integral (cbox a b) (\x. if x \ s then f x else 0) - i) < e)" - (is "?l = ?r") -proof - assume ?r - show ?l - apply (subst has_integral') - apply safe - proof goal_cases - case (1 e) - from \?r\[THEN conjunct2,rule_format,OF this] guess B .. note B=conjunctD2[OF this] - show ?case - apply rule - apply rule - apply (rule B) - apply safe - apply (rule_tac x="integral (cbox a b) (\x. if x \ s then f x else 0)" in exI) - apply (drule B(2)[rule_format]) - using integrable_integral[OF \?r\[THEN conjunct1,rule_format]] - apply auto - done - qed -next - assume ?l note as = this[unfolded has_integral'[of f],rule_format] - let ?f = "\x. if x \ s then f x else 0" - show ?r - proof safe - fix a b :: 'n - from as[OF zero_less_one] guess B .. note B=conjunctD2[OF this,rule_format] - let ?a = "\i\Basis. min (a\i) (-B) *\<^sub>R i::'n" - let ?b = "\i\Basis. max (b\i) B *\<^sub>R i::'n" - show "?f integrable_on cbox a b" - proof (rule integrable_subinterval[of _ ?a ?b]) - have "ball 0 B \ cbox ?a ?b" - apply (rule subsetI) - unfolding mem_ball mem_box dist_norm - proof (rule, goal_cases) - case (1 x i) - then show ?case using Basis_le_norm[of i x] - by (auto simp add:field_simps) - qed - from B(2)[OF this] guess z .. note conjunct1[OF this] - then show "?f integrable_on cbox ?a ?b" - unfolding integrable_on_def by auto - show "cbox a b \ cbox ?a ?b" - apply safe - unfolding mem_box - apply rule - apply (erule_tac x=i in ballE) - apply auto - done - qed - - fix e :: real - assume "e > 0" - from as[OF this] guess B .. note B=conjunctD2[OF this,rule_format] - show "\B>0. \a b. ball 0 B \ cbox a b \ - norm (integral (cbox a b) (\x. if x \ s then f x else 0) - i) < e" - apply rule - apply rule - apply (rule B) - apply safe - proof goal_cases - case 1 - from B(2)[OF this] guess z .. note z=conjunctD2[OF this] - from integral_unique[OF this(1)] show ?case - using z(2) by auto - qed - qed -qed - - -subsection \Continuity of the integral (for a 1-dimensional interval).\ - -lemma integrable_alt: - fixes f :: "'n::euclidean_space \ 'a::banach" - shows "f integrable_on s \ - (\a b. (\x. if x \ s then f x else 0) integrable_on cbox a b) \ - (\e>0. \B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ - norm (integral (cbox a b) (\x. if x \ s then f x else 0) - - integral (cbox c d) (\x. if x \ s then f x else 0)) < e)" - (is "?l = ?r") -proof - assume ?l - then guess y unfolding integrable_on_def .. note this[unfolded has_integral_alt'[of f]] - note y=conjunctD2[OF this,rule_format] - show ?r - apply safe - apply (rule y) - proof goal_cases - case (1 e) - then have "e/2 > 0" - by auto - from y(2)[OF this] guess B .. note B=conjunctD2[OF this,rule_format] - show ?case - apply rule - apply rule - apply (rule B) - apply safe - proof goal_cases - case prems: (1 a b c d) - show ?case - apply (rule norm_triangle_half_l) - using B(2)[OF prems(1)] B(2)[OF prems(2)] - apply auto - done - qed - qed -next - assume ?r - note as = conjunctD2[OF this,rule_format] - let ?cube = "\n. cbox (\i\Basis. - real n *\<^sub>R i::'n) (\i\Basis. real n *\<^sub>R i)" - have "Cauchy (\n. integral (?cube n) (\x. if x \ s then f x else 0))" - proof (unfold Cauchy_def, safe, goal_cases) - case (1 e) - from as(2)[OF this] guess B .. note B = conjunctD2[OF this,rule_format] - from real_arch_simple[of B] guess N .. note N = this - { - fix n - assume n: "n \ N" - have "ball 0 B \ ?cube n" - apply (rule subsetI) - unfolding mem_ball mem_box dist_norm - proof (rule, goal_cases) - case (1 x i) - then show ?case - using Basis_le_norm[of i x] \i\Basis\ - using n N - by (auto simp add: field_simps setsum_negf) - qed - } - then show ?case - apply - - apply (rule_tac x=N in exI) - apply safe - unfolding dist_norm - apply (rule B(2)) - apply auto - done - qed - from this[unfolded convergent_eq_cauchy[symmetric]] guess i .. - note i = this[THEN LIMSEQ_D] - - show ?l unfolding integrable_on_def has_integral_alt'[of f] - apply (rule_tac x=i in exI) - apply safe - apply (rule as(1)[unfolded integrable_on_def]) - proof goal_cases - case (1 e) - then have *: "e/2 > 0" by auto - from i[OF this] guess N .. note N =this[rule_format] - from as(2)[OF *] guess B .. note B=conjunctD2[OF this,rule_format] - let ?B = "max (real N) B" - show ?case - apply (rule_tac x="?B" in exI) - proof safe - show "0 < ?B" - using B(1) by auto - fix a b :: 'n - assume ab: "ball 0 ?B \ cbox a b" - from real_arch_simple[of ?B] guess n .. note n=this - show "norm (integral (cbox a b) (\x. if x \ s then f x else 0) - i) < e" - apply (rule norm_triangle_half_l) - apply (rule B(2)) - defer - apply (subst norm_minus_commute) - apply (rule N[of n]) - proof safe - show "N \ n" - using n by auto - fix x :: 'n - assume x: "x \ ball 0 B" - then have "x \ ball 0 ?B" - by auto - then show "x \ cbox a b" - using ab by blast - show "x \ ?cube n" - using x - unfolding mem_box mem_ball dist_norm - apply - - proof (rule, goal_cases) - case (1 i) - then show ?case - using Basis_le_norm[of i x] \i \ Basis\ - using n - by (auto simp add: field_simps setsum_negf) - qed - qed - qed - qed -qed - -lemma integrable_altD: - fixes f :: "'n::euclidean_space \ 'a::banach" - assumes "f integrable_on s" - shows "\a b. (\x. if x \ s then f x else 0) integrable_on cbox a b" - and "\e. e > 0 \ \B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ - norm (integral (cbox a b) (\x. if x \ s then f x else 0) - integral (cbox c d) (\x. if x \ s then f x else 0)) < e" - using assms[unfolded integrable_alt[of f]] by auto - -lemma integrable_on_subcbox: - fixes f :: "'n::euclidean_space \ 'a::banach" - assumes "f integrable_on s" - and "cbox a b \ s" - shows "f integrable_on cbox a b" - apply (rule integrable_eq) - defer - apply (rule integrable_altD(1)[OF assms(1)]) - using assms(2) - apply auto - done - - -subsection \A straddling criterion for integrability\ - -lemma integrable_straddle_interval: - fixes f :: "'n::euclidean_space \ real" - assumes "\e>0. \g h i j. (g has_integral i) (cbox a b) \ (h has_integral j) (cbox a b) \ - norm (i - j) < e \ (\x\cbox a b. (g x) \ f x \ f x \ h x)" - shows "f integrable_on cbox a b" -proof (subst integrable_cauchy, safe, goal_cases) - case (1 e) - then have e: "e/3 > 0" - by auto - note assms[rule_format,OF this] - then guess g h i j by (elim exE conjE) note obt = this - from obt(1)[unfolded has_integral[of g], rule_format, OF e] guess d1 .. note d1=conjunctD2[OF this,rule_format] - from obt(2)[unfolded has_integral[of h], rule_format, OF e] guess d2 .. note d2=conjunctD2[OF this,rule_format] - show ?case - apply (rule_tac x="\x. d1 x \ d2 x" in exI) - apply (rule conjI gauge_inter d1 d2)+ - unfolding fine_inter - proof (safe, goal_cases) - have **: "\i j g1 g2 h1 h2 f1 f2. g1 - h2 \ f1 - f2 \ f1 - f2 \ h1 - g2 \ - \i - j\ < e / 3 \ \g2 - i\ < e / 3 \ \g1 - i\ < e / 3 \ - \h2 - j\ < e / 3 \ \h1 - j\ < e / 3 \ \f1 - f2\ < e" - using \e > 0\ by arith - case prems: (1 p1 p2) - note tagged_division_ofD(2-4) note * = this[OF prems(1)] this[OF prems(4)] - - have "(\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p1. content k *\<^sub>R g x) \ 0" - and "0 \ (\(x, k)\p2. content k *\<^sub>R h x) - (\(x, k)\p2. content k *\<^sub>R f x)" - and "(\(x, k)\p2. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R g x) \ 0" - and "0 \ (\(x, k)\p1. content k *\<^sub>R h x) - (\(x, k)\p1. content k *\<^sub>R f x)" - unfolding setsum_subtractf[symmetric] - apply - - apply (rule_tac[!] setsum_nonneg) - apply safe - unfolding real_scaleR_def right_diff_distrib[symmetric] - apply (rule_tac[!] mult_nonneg_nonneg) - proof - - fix a b - assume ab: "(a, b) \ p1" - show "0 \ content b" - using *(3)[OF ab] - apply safe - apply (rule content_pos_le) - done - then show "0 \ content b" . - show "0 \ f a - g a" "0 \ h a - f a" - using *(1-2)[OF ab] - using obt(4)[rule_format,of a] - by auto - next - fix a b - assume ab: "(a, b) \ p2" - show "0 \ content b" - using *(6)[OF ab] - apply safe - apply (rule content_pos_le) - done - then show "0 \ content b" . - show "0 \ f a - g a" and "0 \ h a - f a" - using *(4-5)[OF ab] using obt(4)[rule_format,of a] by auto - qed - then show ?case - apply - - unfolding real_norm_def - apply (rule **) - defer - defer - unfolding real_norm_def[symmetric] - apply (rule obt(3)) - apply (rule d1(2)[OF conjI[OF prems(4,5)]]) - apply (rule d1(2)[OF conjI[OF prems(1,2)]]) - apply (rule d2(2)[OF conjI[OF prems(4,6)]]) - apply (rule d2(2)[OF conjI[OF prems(1,3)]]) - apply auto - done - qed -qed - -lemma integrable_straddle: - fixes f :: "'n::euclidean_space \ real" - assumes "\e>0. \g h i j. (g has_integral i) s \ (h has_integral j) s \ - norm (i - j) < e \ (\x\s. g x \ f x \ f x \ h x)" - shows "f integrable_on s" -proof - - have "\a b. (\x. if x \ s then f x else 0) integrable_on cbox a b" - proof (rule integrable_straddle_interval, safe, goal_cases) - case (1 a b e) - then have *: "e/4 > 0" - by auto - from assms[rule_format,OF this] guess g h i j by (elim exE conjE) note obt=this - note obt(1)[unfolded has_integral_alt'[of g]] - note conjunctD2[OF this, rule_format] - note g = this(1) and this(2)[OF *] - from this(2) guess B1 .. note B1 = conjunctD2[OF this,rule_format] - note obt(2)[unfolded has_integral_alt'[of h]] - note conjunctD2[OF this, rule_format] - note h = this(1) and this(2)[OF *] - from this(2) guess B2 .. note B2 = conjunctD2[OF this,rule_format] - define c :: 'n where "c = (\i\Basis. min (a\i) (- (max B1 B2)) *\<^sub>R i)" - define d :: 'n where "d = (\i\Basis. max (b\i) (max B1 B2) *\<^sub>R i)" - have *: "ball 0 B1 \ cbox c d" "ball 0 B2 \ cbox c d" - apply safe - unfolding mem_ball mem_box dist_norm - apply (rule_tac[!] ballI) - proof goal_cases - case (1 x i) - then show ?case using Basis_le_norm[of i x] - unfolding c_def d_def by auto - next - case (2 x i) - then show ?case using Basis_le_norm[of i x] - unfolding c_def d_def by auto - qed - have **: "\ch cg ag ah::real. norm (ah - ag) \ norm (ch - cg) \ norm (cg - i) < e / 4 \ - norm (ch - j) < e / 4 \ norm (ag - ah) < e" - using obt(3) - unfolding real_norm_def - by arith - show ?case - apply (rule_tac x="\x. if x \ s then g x else 0" in exI) - apply (rule_tac x="\x. if x \ s then h x else 0" in exI) - apply (rule_tac x="integral (cbox a b) (\x. if x \ s then g x else 0)" in exI) - apply (rule_tac x="integral (cbox a b) (\x. if x \ s then h x else 0)" in exI) - apply safe - apply (rule_tac[1-2] integrable_integral,rule g) - apply (rule h) - apply (rule **[OF _ B1(2)[OF *(1)] B2(2)[OF *(2)]]) - proof - - have *: "\x f g. (if x \ s then f x else 0) - (if x \ s then g x else 0) = - (if x \ s then f x - g x else (0::real))" - by auto - note ** = abs_of_nonneg[OF integral_nonneg[OF integrable_diff, OF h g]] - show "norm (integral (cbox a b) (\x. if x \ s then h x else 0) - - integral (cbox a b) (\x. if x \ s then g x else 0)) \ - norm (integral (cbox c d) (\x. if x \ s then h x else 0) - - integral (cbox c d) (\x. if x \ s then g x else 0))" - unfolding integral_diff[OF h g,symmetric] real_norm_def - apply (subst **) - defer - apply (subst **) - defer - apply (rule has_integral_subset_le) - defer - apply (rule integrable_integral integrable_diff h g)+ - proof safe - fix x - assume "x \ cbox a b" - then show "x \ cbox c d" - unfolding mem_box c_def d_def - apply - - apply rule - apply (erule_tac x=i in ballE) - apply auto - done - qed (insert obt(4), auto) - qed (insert obt(4), auto) - qed - note interv = this - - show ?thesis - unfolding integrable_alt[of f] - apply safe - apply (rule interv) - proof goal_cases - case (1 e) - then have *: "e/3 > 0" - by auto - from assms[rule_format,OF this] guess g h i j by (elim exE conjE) note obt=this - note obt(1)[unfolded has_integral_alt'[of g]] - note conjunctD2[OF this, rule_format] - note g = this(1) and this(2)[OF *] - from this(2) guess B1 .. note B1 = conjunctD2[OF this,rule_format] - note obt(2)[unfolded has_integral_alt'[of h]] - note conjunctD2[OF this, rule_format] - note h = this(1) and this(2)[OF *] - from this(2) guess B2 .. note B2 = conjunctD2[OF this,rule_format] - show ?case - apply (rule_tac x="max B1 B2" in exI) - apply safe - apply (rule max.strict_coboundedI1) - apply (rule B1) - proof - - fix a b c d :: 'n - assume as: "ball 0 (max B1 B2) \ cbox a b" "ball 0 (max B1 B2) \ cbox c d" - have **: "ball 0 B1 \ ball (0::'n) (max B1 B2)" "ball 0 B2 \ ball (0::'n) (max B1 B2)" - by auto - have *: "\ga gc ha hc fa fc::real. - \ga - i\ < e / 3 \ \gc - i\ < e / 3 \ \ha - j\ < e / 3 \ - \hc - j\ < e / 3 \ \i - j\ < e / 3 \ ga \ fa \ fa \ ha \ gc \ fc \ fc \ hc \ - \fa - fc\ < e" - by (simp add: abs_real_def split: if_split_asm) - show "norm (integral (cbox a b) (\x. if x \ s then f x else 0) - integral (cbox c d) - (\x. if x \ s then f x else 0)) < e" - unfolding real_norm_def - apply (rule *) - apply safe - unfolding real_norm_def[symmetric] - apply (rule B1(2)) - apply (rule order_trans) - apply (rule **) - apply (rule as(1)) - apply (rule B1(2)) - apply (rule order_trans) - apply (rule **) - apply (rule as(2)) - apply (rule B2(2)) - apply (rule order_trans) - apply (rule **) - apply (rule as(1)) - apply (rule B2(2)) - apply (rule order_trans) - apply (rule **) - apply (rule as(2)) - apply (rule obt) - apply (rule_tac[!] integral_le) - using obt - apply (auto intro!: h g interv) - done - qed - qed -qed - - -subsection \Adding integrals over several sets\ - -lemma has_integral_union: - fixes f :: "'n::euclidean_space \ 'a::banach" - assumes "(f has_integral i) s" - and "(f has_integral j) t" - and "negligible (s \ t)" - shows "(f has_integral (i + j)) (s \ t)" -proof - - note * = has_integral_restrict_univ[symmetric, of f] - show ?thesis - unfolding * - apply (rule has_integral_spike[OF assms(3)]) - defer - apply (rule has_integral_add[OF assms(1-2)[unfolded *]]) - apply auto - done -qed - -lemma integrable_union: - fixes f :: "'a::euclidean_space \ 'b :: banach" - assumes "negligible (A \ B)" "f integrable_on A" "f integrable_on B" - shows "f integrable_on (A \ B)" -proof - - from assms obtain y z where "(f has_integral y) A" "(f has_integral z) B" - by (auto simp: integrable_on_def) - from has_integral_union[OF this assms(1)] show ?thesis by (auto simp: integrable_on_def) -qed - -lemma integrable_union': - fixes f :: "'a::euclidean_space \ 'b :: banach" - assumes "f integrable_on A" "f integrable_on B" "negligible (A \ B)" "C = A \ B" - shows "f integrable_on C" - using integrable_union[of A B f] assms by simp - -lemma has_integral_unions: - fixes f :: "'n::euclidean_space \ 'a::banach" - assumes "finite t" - and "\s\t. (f has_integral (i s)) s" - and "\s\t. \s'\t. s \ s' \ negligible (s \ s')" - shows "(f has_integral (setsum i t)) (\t)" -proof - - note * = has_integral_restrict_univ[symmetric, of f] - have **: "negligible (\((\(a,b). a \ b) ` {(a,b). a \ t \ b \ {y. y \ t \ a \ y}}))" - apply (rule negligible_Union) - apply (rule finite_imageI) - apply (rule finite_subset[of _ "t \ t"]) - defer - apply (rule finite_cartesian_product[OF assms(1,1)]) - using assms(3) - apply auto - done - note assms(2)[unfolded *] - note has_integral_setsum[OF assms(1) this] - then show ?thesis - unfolding * - apply - - apply (rule has_integral_spike[OF **]) - defer - apply assumption - apply safe - proof goal_cases - case prems: (1 x) - then show ?case - proof (cases "x \ \t") - case True - then guess s unfolding Union_iff .. note s=this - then have *: "\b\t. x \ b \ b = s" - using prems(3) by blast - show ?thesis - unfolding if_P[OF True] - apply (rule trans) - defer - apply (rule setsum.cong) - apply (rule refl) - apply (subst *) - apply assumption - apply (rule refl) - unfolding setsum.delta[OF assms(1)] - using s - apply auto - done - qed auto - qed -qed - - -text \In particular adding integrals over a division, maybe not of an interval.\ - -lemma has_integral_combine_division: - fixes f :: "'n::euclidean_space \ 'a::banach" - assumes "d division_of s" - and "\k\d. (f has_integral (i k)) k" - shows "(f has_integral (setsum i d)) s" -proof - - note d = division_ofD[OF assms(1)] - show ?thesis - unfolding d(6)[symmetric] - apply (rule has_integral_unions) - apply (rule d assms)+ - apply rule - apply rule - apply rule - proof goal_cases - case prems: (1 s s') - from d(4)[OF this(1)] d(4)[OF this(2)] guess a c b d by (elim exE) note obt=this - from d(5)[OF prems] show ?case - unfolding obt interior_cbox - apply - - apply (rule negligible_subset[of "(cbox a b-box a b) \ (cbox c d-box c d)"]) - apply (rule negligible_Un negligible_frontier_interval)+ - apply auto - done - qed -qed - -lemma integral_combine_division_bottomup: - fixes f :: "'n::euclidean_space \ 'a::banach" - assumes "d division_of s" - and "\k\d. f integrable_on k" - shows "integral s f = setsum (\i. integral i f) d" - apply (rule integral_unique) - apply (rule has_integral_combine_division[OF assms(1)]) - using assms(2) - unfolding has_integral_integral - apply assumption - done - -lemma has_integral_combine_division_topdown: - fixes f :: "'n::euclidean_space \ 'a::banach" - assumes "f integrable_on s" - and "d division_of k" - and "k \ s" - shows "(f has_integral (setsum (\i. integral i f) d)) k" - apply (rule has_integral_combine_division[OF assms(2)]) - apply safe - unfolding has_integral_integral[symmetric] -proof goal_cases - case (1 k) - from division_ofD(2,4)[OF assms(2) this] - show ?case - apply safe - apply (rule integrable_on_subcbox) - apply (rule assms) - using assms(3) - apply auto - done -qed - -lemma integral_combine_division_topdown: - fixes f :: "'n::euclidean_space \ 'a::banach" - assumes "f integrable_on s" - and "d division_of s" - shows "integral s f = setsum (\i. integral i f) d" - apply (rule integral_unique) - apply (rule has_integral_combine_division_topdown) - using assms - apply auto - done - -lemma integrable_combine_division: - fixes f :: "'n::euclidean_space \ 'a::banach" - assumes "d division_of s" - and "\i\d. f integrable_on i" - shows "f integrable_on s" - using assms(2) - unfolding integrable_on_def - by (metis has_integral_combine_division[OF assms(1)]) - -lemma integrable_on_subdivision: - fixes f :: "'n::euclidean_space \ 'a::banach" - assumes "d division_of i" - and "f integrable_on s" - and "i \ s" - shows "f integrable_on i" - apply (rule integrable_combine_division assms)+ - apply safe -proof goal_cases - case 1 - note division_ofD(2,4)[OF assms(1) this] - then show ?case - apply safe - apply (rule integrable_on_subcbox[OF assms(2)]) - using assms(3) - apply auto - done -qed - - -subsection \Also tagged divisions\ - -lemma has_integral_combine_tagged_division: - fixes f :: "'n::euclidean_space \ 'a::banach" - assumes "p tagged_division_of s" - and "\(x,k) \ p. (f has_integral (i k)) k" - shows "(f has_integral (setsum (\(x,k). i k) p)) s" -proof - - have *: "(f has_integral (setsum (\k. integral k f) (snd ` p))) s" - apply (rule has_integral_combine_division) - apply (rule division_of_tagged_division[OF assms(1)]) - using assms(2) - unfolding has_integral_integral[symmetric] - apply safe - apply auto - done - then show ?thesis - apply - - apply (rule subst[where P="\i. (f has_integral i) s"]) - defer - apply assumption - apply (rule trans[of _ "setsum (\(x,k). integral k f) p"]) - apply (subst eq_commute) - apply (rule setsum.over_tagged_division_lemma[OF assms(1)]) - apply (rule integral_null) - apply assumption - apply (rule setsum.cong) - using assms(2) - apply auto - done -qed - -lemma integral_combine_tagged_division_bottomup: - fixes f :: "'n::euclidean_space \ 'a::banach" - assumes "p tagged_division_of (cbox a b)" - and "\(x,k)\p. f integrable_on k" - shows "integral (cbox a b) f = setsum (\(x,k). integral k f) p" - apply (rule integral_unique) - apply (rule has_integral_combine_tagged_division[OF assms(1)]) - using assms(2) - apply auto - done - -lemma has_integral_combine_tagged_division_topdown: - fixes f :: "'n::euclidean_space \ 'a::banach" - assumes "f integrable_on cbox a b" - and "p tagged_division_of (cbox a b)" - shows "(f has_integral (setsum (\(x,k). integral k f) p)) (cbox a b)" - apply (rule has_integral_combine_tagged_division[OF assms(2)]) - apply safe -proof goal_cases - case 1 - note tagged_division_ofD(3-4)[OF assms(2) this] - then show ?case - using integrable_subinterval[OF assms(1)] by blast -qed - -lemma integral_combine_tagged_division_topdown: - fixes f :: "'n::euclidean_space \ 'a::banach" - assumes "f integrable_on cbox a b" - and "p tagged_division_of (cbox a b)" - shows "integral (cbox a b) f = setsum (\(x,k). integral k f) p" - apply (rule integral_unique) - apply (rule has_integral_combine_tagged_division_topdown) - using assms - apply auto - done - - -subsection \Henstock's lemma\ - -lemma henstock_lemma_part1: - fixes f :: "'n::euclidean_space \ 'a::banach" - assumes "f integrable_on cbox a b" - and "e > 0" - and "gauge d" - and "(\p. p tagged_division_of (cbox a b) \ d fine p \ - norm (setsum (\(x,k). content k *\<^sub>R f x) p - integral(cbox a b) f) < e)" - and p: "p tagged_partial_division_of (cbox a b)" "d fine p" - shows "norm (setsum (\(x,k). content k *\<^sub>R f x - integral k f) p) \ e" - (is "?x \ e") -proof - - { presume "\k. 0 ?x \ e + k" then show ?thesis by (blast intro: field_le_epsilon) } - fix k :: real - assume k: "k > 0" - note p' = tagged_partial_division_ofD[OF p(1)] - have "\(snd ` p) \ cbox a b" - using p'(3) by fastforce - note partial_division_of_tagged_division[OF p(1)] this - from partial_division_extend_interval[OF this] guess q . note q=this and q' = division_ofD[OF this(2)] - define r where "r = q - snd ` p" - have "snd ` p \ r = {}" - unfolding r_def by auto - have r: "finite r" - using q' unfolding r_def by auto - - have "\i\r. \p. p tagged_division_of i \ d fine p \ - norm (setsum (\(x,j). content j *\<^sub>R f x) p - integral i f) < k / (real (card r) + 1)" - apply safe - proof goal_cases - case (1 i) - then have i: "i \ q" - unfolding r_def by auto - from q'(4)[OF this] guess u v by (elim exE) note uv=this - have *: "k / (real (card r) + 1) > 0" using k by simp - have "f integrable_on cbox u v" - apply (rule integrable_subinterval[OF assms(1)]) - using q'(2)[OF i] - unfolding uv - apply auto - done - note integrable_integral[OF this, unfolded has_integral[of f]] - from this[rule_format,OF *] guess dd .. note dd=conjunctD2[OF this,rule_format] - note gauge_inter[OF \gauge d\ dd(1)] - from fine_division_exists[OF this,of u v] guess qq . - then show ?case - apply (rule_tac x=qq in exI) - using dd(2)[of qq] - unfolding fine_inter uv - apply auto - done - qed - from bchoice[OF this] guess qq .. note qq=this[rule_format] - - let ?p = "p \ \(qq ` r)" - have "norm ((\(x, k)\?p. content k *\<^sub>R f x) - integral (cbox a b) f) < e" - apply (rule assms(4)[rule_format]) - proof - show "d fine ?p" - apply (rule fine_union) - apply (rule p) - apply (rule fine_unions) - using qq - apply auto - done - note * = tagged_partial_division_of_union_self[OF p(1)] - have "p \ \(qq ` r) tagged_division_of \(snd ` p) \ \r" - using r - proof (rule tagged_division_union[OF * tagged_division_unions], goal_cases) - case 1 - then show ?case - using qq by auto - next - case 2 - then show ?case - apply rule - apply rule - apply rule - apply(rule q'(5)) - unfolding r_def - apply auto - done - next - case 3 - then show ?case - apply (rule inter_interior_unions_intervals) - apply fact - apply rule - apply rule - apply (rule q') - defer - apply rule - apply (subst Int_commute) - apply (rule inter_interior_unions_intervals) - apply (rule finite_imageI) - apply (rule p') - apply rule - defer - apply rule - apply (rule q') - using q(1) p' - unfolding r_def - apply auto - done - qed - moreover have "\(snd ` p) \ \r = cbox a b" and "{qq i |i. i \ r} = qq ` r" - unfolding Union_Un_distrib[symmetric] r_def - using q - by auto - ultimately show "?p tagged_division_of (cbox a b)" - by fastforce - qed - - then have "norm ((\(x, k)\p. content k *\<^sub>R f x) + (\(x, k)\\(qq ` r). content k *\<^sub>R f x) - - integral (cbox a b) f) < e" - apply (subst setsum.union_inter_neutral[symmetric]) - apply (rule p') - prefer 3 - apply assumption - apply rule - apply (rule r) - apply safe - apply (drule qq) - proof - - fix x l k - assume as: "(x, l) \ p" "(x, l) \ qq k" "k \ r" - note qq[OF this(3)] - note tagged_division_ofD(3,4)[OF conjunct1[OF this] as(2)] - from this(2) guess u v by (elim exE) note uv=this - have "l\snd ` p" unfolding image_iff apply(rule_tac x="(x,l)" in bexI) using as by auto - then have "l \ q" "k \ q" "l \ k" - using as(1,3) q(1) unfolding r_def by auto - note q'(5)[OF this] - then have "interior l = {}" - using interior_mono[OF \l \ k\] by blast - then show "content l *\<^sub>R f x = 0" - unfolding uv content_eq_0_interior[symmetric] by auto - qed auto - - then have "norm ((\(x, k)\p. content k *\<^sub>R f x) + setsum (setsum (\(x, k). content k *\<^sub>R f x)) - (qq ` r) - integral (cbox a b) f) < e" - apply (subst (asm) setsum.Union_comp) - prefer 2 - unfolding split_paired_all split_conv image_iff - apply (erule bexE)+ - proof - - fix x m k l T1 T2 - assume "(x, m) \ T1" "(x, m) \ T2" "T1 \ T2" "k \ r" "l \ r" "T1 = qq k" "T2 = qq l" - note as = this(1-5)[unfolded this(6-)] - note kl = tagged_division_ofD(3,4)[OF qq[THEN conjunct1]] - from this(2)[OF as(4,1)] guess u v by (elim exE) note uv=this - have *: "interior (k \ l) = {}" - by (metis DiffE \T1 = qq k\ \T1 \ T2\ \T2 = qq l\ as(4) as(5) interior_Int q'(5) r_def) - have "interior m = {}" - unfolding subset_empty[symmetric] - unfolding *[symmetric] - apply (rule interior_mono) - using kl(1)[OF as(4,1)] kl(1)[OF as(5,2)] - apply auto - done - then show "content m *\<^sub>R f x = 0" - unfolding uv content_eq_0_interior[symmetric] - by auto - qed (insert qq, auto) - - then have **: "norm ((\(x, k)\p. content k *\<^sub>R f x) + setsum (setsum (\(x, k). content k *\<^sub>R f x) \ qq) r - - integral (cbox a b) f) < e" - apply (subst (asm) setsum.reindex_nontrivial) - apply fact - apply (rule setsum.neutral) - apply rule - unfolding split_paired_all split_conv - defer - apply assumption - proof - - fix k l x m - assume as: "k \ r" "l \ r" "k \ l" "qq k = qq l" "(x, m) \ qq k" - note tagged_division_ofD(6)[OF qq[THEN conjunct1]] - from this[OF as(1)] this[OF as(2)] show "content m *\<^sub>R f x = 0" - using as(3) unfolding as by auto - qed - - have *: "norm (cp - ip) \ e + k" - if "norm ((cp + cr) - i) < e" - and "norm (cr - ir) < k" - and "ip + ir = i" - for ir ip i cr cp - proof - - from that show ?thesis - using norm_triangle_le[of "cp + cr - i" "- (cr - ir)"] - unfolding that(3)[symmetric] norm_minus_cancel - by (auto simp add: algebra_simps) - qed - - have "?x = norm ((\(x, k)\p. content k *\<^sub>R f x) - (\(x, k)\p. integral k f))" - unfolding split_def setsum_subtractf .. - also have "\ \ e + k" - apply (rule *[OF **, where ir1="setsum (\k. integral k f) r"]) - proof goal_cases - case 1 - have *: "k * real (card r) / (1 + real (card r)) < k" - using k by (auto simp add: field_simps) - show ?case - apply (rule le_less_trans[of _ "setsum (\x. k / (real (card r) + 1)) r"]) - unfolding setsum_subtractf[symmetric] - apply (rule setsum_norm_le) - apply rule - apply (drule qq) - defer - unfolding divide_inverse setsum_left_distrib[symmetric] - unfolding divide_inverse[symmetric] - using * apply (auto simp add: field_simps) - done - next - case 2 - have *: "(\(x, k)\p. integral k f) = (\k\snd ` p. integral k f)" - apply (subst setsum.reindex_nontrivial) - apply fact - unfolding split_paired_all snd_conv split_def o_def - proof - - fix x l y m - assume as: "(x, l) \ p" "(y, m) \ p" "(x, l) \ (y, m)" "l = m" - from p'(4)[OF as(1)] guess u v by (elim exE) note uv=this - show "integral l f = 0" - unfolding uv - apply (rule integral_unique) - apply (rule has_integral_null) - unfolding content_eq_0_interior - using p'(5)[OF as(1-3)] - unfolding uv as(4)[symmetric] - apply auto - done - qed auto - from q(1) have **: "snd ` p \ q = q" by auto - show ?case - unfolding integral_combine_division_topdown[OF assms(1) q(2)] * r_def - using ** q'(1) p'(1) setsum.union_disjoint [of "snd ` p" "q - snd ` p" "\k. integral k f", symmetric] - by simp - qed - finally show "?x \ e + k" . -qed - -lemma henstock_lemma_part2: - fixes f :: "'m::euclidean_space \ 'n::euclidean_space" - assumes "f integrable_on cbox a b" - and "e > 0" - and "gauge d" - and "\p. p tagged_division_of (cbox a b) \ d fine p \ - norm (setsum (\(x,k). content k *\<^sub>R f x) p - integral (cbox a b) f) < e" - and "p tagged_partial_division_of (cbox a b)" - and "d fine p" - shows "setsum (\(x,k). norm (content k *\<^sub>R f x - integral k f)) p \ 2 * real (DIM('n)) * e" - unfolding split_def - apply (rule setsum_norm_allsubsets_bound) - defer - apply (rule henstock_lemma_part1[unfolded split_def,OF assms(1-3)]) - apply safe - apply (rule assms[rule_format,unfolded split_def]) - defer - apply (rule tagged_partial_division_subset) - apply (rule assms) - apply assumption - apply (rule fine_subset) - apply assumption - apply (rule assms) - using assms(5) - apply auto - done - -lemma henstock_lemma: - fixes f :: "'m::euclidean_space \ 'n::euclidean_space" - assumes "f integrable_on cbox a b" - and "e > 0" - obtains d where "gauge d" - and "\p. p tagged_partial_division_of (cbox a b) \ d fine p \ - setsum (\(x,k). norm(content k *\<^sub>R f x - integral k f)) p < e" -proof - - have *: "e / (2 * (real DIM('n) + 1)) > 0" using assms(2) by simp - from integrable_integral[OF assms(1),unfolded has_integral[of f],rule_format,OF this] - guess d .. note d = conjunctD2[OF this] - show thesis - apply (rule that) - apply (rule d) - proof (safe, goal_cases) - case (1 p) - note * = henstock_lemma_part2[OF assms(1) * d this] - show ?case - apply (rule le_less_trans[OF *]) - using \e > 0\ - apply (auto simp add: field_simps) - done - qed -qed - - -subsection \Geometric progression\ - -text \FIXME: Should one or more of these theorems be moved to @{file -"~~/src/HOL/Set_Interval.thy"}, alongside \geometric_sum\?\ - -lemma sum_gp_basic: - fixes x :: "'a::ring_1" - shows "(1 - x) * setsum (\i. x^i) {0 .. n} = (1 - x^(Suc n))" -proof - - define y where "y = 1 - x" - have "y * (\i=0..n. (1 - y) ^ i) = 1 - (1 - y) ^ Suc n" - by (induct n) (simp_all add: algebra_simps) - then show ?thesis - unfolding y_def by simp -qed - -lemma sum_gp_multiplied: - assumes mn: "m \ n" - shows "((1::'a::{field}) - x) * setsum (op ^ x) {m..n} = x^m - x^ Suc n" - (is "?lhs = ?rhs") -proof - - let ?S = "{0..(n - m)}" - from mn have mn': "n - m \ 0" - by arith - let ?f = "op + m" - have i: "inj_on ?f ?S" - unfolding inj_on_def by auto - have f: "?f ` ?S = {m..n}" - using mn - apply (auto simp add: image_iff Bex_def) - apply presburger - done - have th: "op ^ x \ op + m = (\i. x^m * x^i)" - by (rule ext) (simp add: power_add power_mult) - from setsum.reindex[OF i, of "op ^ x", unfolded f th setsum_right_distrib[symmetric]] - have "?lhs = x^m * ((1 - x) * setsum (op ^ x) {0..n - m})" - by simp - then show ?thesis - unfolding sum_gp_basic - using mn - by (simp add: field_simps power_add[symmetric]) -qed - -lemma sum_gp: - "setsum (op ^ (x::'a::{field})) {m .. n} = - (if n < m then 0 - else if x = 1 then of_nat ((n + 1) - m) - else (x^ m - x^ (Suc n)) / (1 - x))" -proof - - { - assume nm: "n < m" - then have ?thesis by simp - } - moreover - { - assume "\ n < m" - then have nm: "m \ n" - by arith - { - assume x: "x = 1" - then have ?thesis - by simp - } - moreover - { - assume x: "x \ 1" - then have nz: "1 - x \ 0" - by simp - from sum_gp_multiplied[OF nm, of x] nz have ?thesis - by (simp add: field_simps) - } - ultimately have ?thesis by blast - } - ultimately show ?thesis by blast -qed - -lemma sum_gp_offset: - "setsum (op ^ (x::'a::{field})) {m .. m+n} = - (if x = 1 then of_nat n + 1 else x^m * (1 - x^Suc n) / (1 - x))" - unfolding sum_gp[of x m "m + n"] power_Suc - by (simp add: field_simps power_add) - - -subsection \Monotone convergence (bounded interval first)\ - -lemma monotone_convergence_interval: - fixes f :: "nat \ 'n::euclidean_space \ real" - assumes "\k. (f k) integrable_on cbox a b" - and "\k. \x\cbox a b.(f k x) \ f (Suc k) x" - and "\x\cbox a b. ((\k. f k x) \ g x) sequentially" - and "bounded {integral (cbox a b) (f k) | k . k \ UNIV}" - shows "g integrable_on cbox a b \ ((\k. integral (cbox a b) (f k)) \ integral (cbox a b) g) sequentially" -proof (cases "content (cbox a b) = 0") - case True - show ?thesis - using integrable_on_null[OF True] - unfolding integral_null[OF True] - using tendsto_const - by auto -next - case False - have fg: "\x\cbox a b. \k. (f k x) \ 1 \ (g x) \ 1" - proof safe - fix x k - assume x: "x \ cbox a b" - note * = Lim_component_ge[OF assms(3)[rule_format, OF x] trivial_limit_sequentially] - show "f k x \ 1 \ g x \ 1" - apply (rule *) - unfolding eventually_sequentially - apply (rule_tac x=k in exI) - apply - - apply (rule transitive_stepwise_le) - using assms(2)[rule_format, OF x] - apply auto - done - qed - have "\i. ((\k. integral (cbox a b) (f k)) \ i) sequentially" - apply (rule bounded_increasing_convergent) - defer - apply rule - apply (rule integral_le) - apply safe - apply (rule assms(1-2)[rule_format])+ - using assms(4) - apply auto - done - then guess i .. note i=this - have i': "\k. (integral(cbox a b) (f k)) \ i\1" - apply (rule Lim_component_ge) - apply (rule i) - apply (rule trivial_limit_sequentially) - unfolding eventually_sequentially - apply (rule_tac x=k in exI) - apply (rule transitive_stepwise_le) - prefer 3 - unfolding inner_simps real_inner_1_right - apply (rule integral_le) - apply (rule assms(1-2)[rule_format])+ - using assms(2) - apply auto - done - - have "(g has_integral i) (cbox a b)" - unfolding has_integral - proof (safe, goal_cases) - case e: (1 e) - then have "\k. (\d. gauge d \ (\p. p tagged_division_of (cbox a b) \ d fine p \ - norm ((\(x, ka)\p. content ka *\<^sub>R f k x) - integral (cbox a b) (f k)) < e / 2 ^ (k + 2)))" - apply - - apply rule - apply (rule assms(1)[unfolded has_integral_integral has_integral,rule_format]) - apply auto - done - from choice[OF this] guess c .. note c=conjunctD2[OF this[rule_format],rule_format] - - have "\r. \k\r. 0 \ i\1 - (integral (cbox a b) (f k)) \ i\1 - (integral (cbox a b) (f k)) < e / 4" - proof - - have "e/4 > 0" - using e by auto - from LIMSEQ_D [OF i this] guess r .. - then show ?thesis - apply (rule_tac x=r in exI) - apply rule - apply (erule_tac x=k in allE) - subgoal for k using i'[of k] by auto - done - qed - then guess r .. note r=conjunctD2[OF this[rule_format]] - - have "\x\cbox a b. \n\r. \k\n. 0 \ (g x)\1 - (f k x)\1 \ - (g x)\1 - (f k x)\1 < e / (4 * content(cbox a b))" - proof (rule, goal_cases) - case prems: (1 x) - have "e / (4 * content (cbox a b)) > 0" - using \e>0\ False content_pos_le[of a b] by auto - from assms(3)[rule_format, OF prems, THEN LIMSEQ_D, OF this] - guess n .. note n=this - then show ?case - apply (rule_tac x="n + r" in exI) - apply safe - apply (erule_tac[2-3] x=k in allE) - unfolding dist_real_def - using fg[rule_format, OF prems] - apply (auto simp add: field_simps) - done - qed - from bchoice[OF this] guess m .. note m=conjunctD2[OF this[rule_format],rule_format] - define d where "d x = c (m x) x" for x - show ?case - apply (rule_tac x=d in exI) - proof safe - show "gauge d" - using c(1) unfolding gauge_def d_def by auto - next - fix p - assume p: "p tagged_division_of (cbox a b)" "d fine p" - note p'=tagged_division_ofD[OF p(1)] - have "\a. \x\p. m (fst x) \ a" - by (metis finite_imageI finite_nat_set_iff_bounded_le p'(1) rev_image_eqI) - then guess s .. note s=this - have *: "\a b c d. norm(a - b) \ e / 4 \ norm(b - c) < e / 2 \ - norm (c - d) < e / 4 \ norm (a - d) < e" - proof (safe, goal_cases) - case (1 a b c d) - then show ?case - using norm_triangle_lt[of "a - b" "b - c" "3* e/4"] - norm_triangle_lt[of "a - b + (b - c)" "c - d" e] - unfolding norm_minus_cancel - by (auto simp add: algebra_simps) - qed - show "norm ((\(x, k)\p. content k *\<^sub>R g x) - i) < e" - apply (rule *[rule_format,where - b="\(x, k)\p. content k *\<^sub>R f (m x) x" and c="\(x, k)\p. integral k (f (m x))"]) - proof (safe, goal_cases) - case 1 - show ?case - apply (rule order_trans[of _ "\(x, k)\p. content k * (e / (4 * content (cbox a b)))"]) - unfolding setsum_subtractf[symmetric] - apply (rule order_trans) - apply (rule norm_setsum) - apply (rule setsum_mono) - unfolding split_paired_all split_conv - unfolding split_def setsum_left_distrib[symmetric] scaleR_diff_right[symmetric] - unfolding additive_content_tagged_division[OF p(1), unfolded split_def] - proof - - fix x k - assume xk: "(x, k) \ p" - then have x: "x \ cbox a b" - using p'(2-3)[OF xk] by auto - from p'(4)[OF xk] guess u v by (elim exE) note uv=this - show "norm (content k *\<^sub>R (g x - f (m x) x)) \ content k * (e / (4 * content (cbox a b)))" - unfolding norm_scaleR uv - unfolding abs_of_nonneg[OF content_pos_le] - apply (rule mult_left_mono) - using m(2)[OF x,of "m x"] - apply auto - done - qed (insert False, auto) - - next - case 2 - show ?case - apply (rule le_less_trans[of _ "norm (\j = 0..s. - \(x, k)\{xk\p. m (fst xk) = j}. content k *\<^sub>R f (m x) x - integral k (f (m x)))"]) - apply (subst setsum_group) - apply fact - apply (rule finite_atLeastAtMost) - defer - apply (subst split_def)+ - unfolding setsum_subtractf - apply rule - proof - - show "norm (\j = 0..s. \(x, k)\{xk \ p. - m (fst xk) = j}. content k *\<^sub>R f (m x) x - integral k (f (m x))) < e / 2" - apply (rule le_less_trans[of _ "setsum (\i. e / 2^(i+2)) {0..s}"]) - apply (rule setsum_norm_le) - proof - show "(\i = 0..s. e / 2 ^ (i + 2)) < e / 2" - unfolding power_add divide_inverse inverse_mult_distrib - unfolding setsum_right_distrib[symmetric] setsum_left_distrib[symmetric] - unfolding power_inverse [symmetric] sum_gp - apply(rule mult_strict_left_mono[OF _ e]) - unfolding power2_eq_square - apply auto - done - fix t - assume "t \ {0..s}" - show "norm (\(x, k)\{xk \ p. m (fst xk) = t}. content k *\<^sub>R f (m x) x - - integral k (f (m x))) \ e / 2 ^ (t + 2)" - apply (rule order_trans - [of _ "norm (setsum (\(x,k). content k *\<^sub>R f t x - integral k (f t)) {xk \ p. m (fst xk) = t})"]) - apply (rule eq_refl) - apply (rule arg_cong[where f=norm]) - apply (rule setsum.cong) - apply (rule refl) - defer - apply (rule henstock_lemma_part1) - apply (rule assms(1)[rule_format]) - apply (simp add: e) - apply safe - apply (rule c)+ - apply rule - apply assumption+ - apply (rule tagged_partial_division_subset[of p]) - apply (rule p(1)[unfolded tagged_division_of_def,THEN conjunct1]) - defer - unfolding fine_def - apply safe - apply (drule p(2)[unfolded fine_def,rule_format]) - unfolding d_def - apply auto - done - qed - qed (insert s, auto) - next - case 3 - note comb = integral_combine_tagged_division_topdown[OF assms(1)[rule_format] p(1)] - have *: "\sr sx ss ks kr::real. kr = sr \ ks = ss \ - ks \ i \ sr \ sx \ sx \ ss \ 0 \ i\1 - kr\1 \ i\1 - kr\1 < e/4 \ \sx - i\ < e/4" - by auto - show ?case - unfolding real_norm_def - apply (rule *[rule_format]) - apply safe - apply (rule comb[of r]) - apply (rule comb[of s]) - apply (rule i'[unfolded real_inner_1_right]) - apply (rule_tac[1-2] setsum_mono) - unfolding split_paired_all split_conv - apply (rule_tac[1-2] integral_le[OF ]) - proof safe - show "0 \ i\1 - (integral (cbox a b) (f r))\1" - using r(1) by auto - show "i\1 - (integral (cbox a b) (f r))\1 < e / 4" - using r(2) by auto - fix x k - assume xk: "(x, k) \ p" - from p'(4)[OF this] guess u v by (elim exE) note uv=this - show "f r integrable_on k" - and "f s integrable_on k" - and "f (m x) integrable_on k" - and "f (m x) integrable_on k" - unfolding uv - apply (rule_tac[!] integrable_on_subcbox[OF assms(1)[rule_format]]) - using p'(3)[OF xk] - unfolding uv - apply auto - done - fix y - assume "y \ k" - then have "y \ cbox a b" - using p'(3)[OF xk] by auto - then have *: "\m. \n\m. f m y \ f n y" - apply - - apply (rule transitive_stepwise_le) - using assms(2) - apply auto - done - show "f r y \ f (m x) y" and "f (m x) y \ f s y" - apply (rule_tac[!] *[rule_format]) - using s[rule_format,OF xk] m(1)[of x] p'(2-3)[OF xk] - apply auto - done - qed - qed - qed - qed note * = this - - have "integral (cbox a b) g = i" - by (rule integral_unique) (rule *) - then show ?thesis - using i * by auto -qed - -lemma monotone_convergence_increasing: - fixes f :: "nat \ 'n::euclidean_space \ real" - assumes "\k. (f k) integrable_on s" - and "\k. \x\s. (f k x) \ (f (Suc k) x)" - and "\x\s. ((\k. f k x) \ g x) sequentially" - and "bounded {integral s (f k)| k. True}" - shows "g integrable_on s \ ((\k. integral s (f k)) \ integral s g) sequentially" -proof - - have lem: "g integrable_on s \ ((\k. integral s (f k)) \ integral s g) sequentially" - if "\k. \x\s. 0 \ f k x" - and "\k. (f k) integrable_on s" - and "\k. \x\s. f k x \ f (Suc k) x" - and "\x\s. ((\k. f k x) \ g x) sequentially" - and "bounded {integral s (f k)| k. True}" - for f :: "nat \ 'n::euclidean_space \ real" and g s - proof - - note assms=that[rule_format] - have "\x\s. \k. (f k x)\1 \ (g x)\1" - apply safe - apply (rule Lim_component_ge) - apply (rule that(4)[rule_format]) - apply assumption - apply (rule trivial_limit_sequentially) - unfolding eventually_sequentially - apply (rule_tac x=k in exI) - apply (rule transitive_stepwise_le) - using that(3) - apply auto - done - note fg=this[rule_format] - - have "\i. ((\k. integral s (f k)) \ i) sequentially" - apply (rule bounded_increasing_convergent) - apply (rule that(5)) - apply rule - apply (rule integral_le) - apply (rule that(2)[rule_format])+ - using that(3) - apply auto - done - then guess i .. note i=this - have "\k. \x\s. \n\k. f k x \ f n x" - apply rule - apply (rule transitive_stepwise_le) - using that(3) - apply auto - done - then have i': "\k. (integral s (f k))\1 \ i\1" - apply - - apply rule - apply (rule Lim_component_ge) - apply (rule i) - apply (rule trivial_limit_sequentially) - unfolding eventually_sequentially - apply (rule_tac x=k in exI) - apply safe - apply (rule integral_component_le) - apply simp - apply (rule that(2)[rule_format])+ - apply auto - done - - note int = assms(2)[unfolded integrable_alt[of _ s],THEN conjunct1,rule_format] - have ifif: "\k t. (\x. if x \ t then if x \ s then f k x else 0 else 0) = - (\x. if x \ t \ s then f k x else 0)" - by (rule ext) auto - have int': "\k a b. f k integrable_on cbox a b \ s" - apply (subst integrable_restrict_univ[symmetric]) - apply (subst ifif[symmetric]) - apply (subst integrable_restrict_univ) - apply (rule int) - done - have "\a b. (\x. if x \ s then g x else 0) integrable_on cbox a b \ - ((\k. integral (cbox a b) (\x. if x \ s then f k x else 0)) \ - integral (cbox a b) (\x. if x \ s then g x else 0)) sequentially" - proof (rule monotone_convergence_interval, safe, goal_cases) - case 1 - show ?case by (rule int) - next - case (2 _ _ _ x) - then show ?case - apply (cases "x \ s") - using assms(3) - apply auto - done - next - case (3 _ _ x) - then show ?case - apply (cases "x \ s") - using assms(4) - apply auto - done - next - case (4 a b) - note * = integral_nonneg - have "\k. norm (integral (cbox a b) (\x. if x \ s then f k x else 0)) \ norm (integral s (f k))" - unfolding real_norm_def - apply (subst abs_of_nonneg) - apply (rule *[OF int]) - apply safe - apply (case_tac "x \ s") - apply (drule assms(1)) - prefer 3 - apply (subst abs_of_nonneg) - apply (rule *[OF assms(2) that(1)[THEN spec]]) - apply (subst integral_restrict_univ[symmetric,OF int]) - unfolding ifif - unfolding integral_restrict_univ[OF int'] - apply (rule integral_subset_le[OF _ int' assms(2)]) - using assms(1) - apply auto - done - then show ?case - using assms(5) - unfolding bounded_iff - apply safe - apply (rule_tac x=aa in exI) - apply safe - apply (erule_tac x="integral s (f k)" in ballE) - apply (rule order_trans) - apply assumption - apply auto - done - qed - note g = conjunctD2[OF this] - - have "(g has_integral i) s" - unfolding has_integral_alt' - apply safe - apply (rule g(1)) - proof goal_cases - case (1 e) - then have "e/4>0" - by auto - from LIMSEQ_D [OF i this] guess N .. note N=this - note assms(2)[of N,unfolded has_integral_integral has_integral_alt'[of "f N"]] - from this[THEN conjunct2,rule_format,OF \e/4>0\] guess B .. note B=conjunctD2[OF this] - show ?case - apply rule - apply rule - apply (rule B) - apply safe - proof - - fix a b :: 'n - assume ab: "ball 0 B \ cbox a b" - from \e > 0\ have "e/2 > 0" - by auto - from LIMSEQ_D [OF g(2)[of a b] this] guess M .. note M=this - have **: "norm (integral (cbox a b) (\x. if x \ s then f N x else 0) - i) < e/2" - apply (rule norm_triangle_half_l) - using B(2)[rule_format,OF ab] N[rule_format,of N] - apply - - defer - apply (subst norm_minus_commute) - apply auto - done - have *: "\f1 f2 g. \f1 - i\ < e / 2 \ \f2 - g\ < e / 2 \ - f1 \ f2 \ f2 \ i \ \g - i\ < e" - unfolding real_inner_1_right by arith - show "norm (integral (cbox a b) (\x. if x \ s then g x else 0) - i) < e" - unfolding real_norm_def - apply (rule *[rule_format]) - apply (rule **[unfolded real_norm_def]) - apply (rule M[rule_format,of "M + N",unfolded real_norm_def]) - apply (rule le_add1) - apply (rule integral_le[OF int int]) - defer - apply (rule order_trans[OF _ i'[rule_format,of "M + N",unfolded real_inner_1_right]]) - proof (safe, goal_cases) - case (2 x) - have "\m. x \ s \ \n\m. (f m x)\1 \ (f n x)\1" - apply (rule transitive_stepwise_le) - using assms(3) - apply auto - done - then show ?case - by auto - next - case 1 - show ?case - apply (subst integral_restrict_univ[symmetric,OF int]) - unfolding ifif integral_restrict_univ[OF int'] - apply (rule integral_subset_le[OF _ int']) - using assms - apply auto - done - qed - qed - qed - then show ?thesis - apply safe - defer - apply (drule integral_unique) - using i - apply auto - done - qed - - have sub: "\k. integral s (\x. f k x - f 0 x) = integral s (f k) - integral s (f 0)" - apply (subst integral_diff) - apply (rule assms(1)[rule_format])+ - apply rule - done - have "\x m. x \ s \ \n\m. f m x \ f n x" - apply (rule transitive_stepwise_le) - using assms(2) - apply auto - done - note * = this[rule_format] - have "(\x. g x - f 0 x) integrable_on s \ ((\k. integral s (\x. f (Suc k) x - f 0 x)) \ - integral s (\x. g x - f 0 x)) sequentially" - apply (rule lem) - apply safe - proof goal_cases - case (1 k x) - then show ?case - using *[of x 0 "Suc k"] by auto - next - case (2 k) - then show ?case - apply (rule integrable_diff) - using assms(1) - apply auto - done - next - case (3 k x) - then show ?case - using *[of x "Suc k" "Suc (Suc k)"] by auto - next - case (4 x) - then show ?case - apply - - apply (rule tendsto_diff) - using LIMSEQ_ignore_initial_segment[OF assms(3)[rule_format],of x 1] - apply auto - done - next - case 5 - then show ?case - using assms(4) - unfolding bounded_iff - apply safe - apply (rule_tac x="a + norm (integral s (\x. f 0 x))" in exI) - apply safe - apply (erule_tac x="integral s (\x. f (Suc k) x)" in ballE) - unfolding sub - apply (rule order_trans[OF norm_triangle_ineq4]) - apply auto - done - qed - note conjunctD2[OF this] - note tendsto_add[OF this(2) tendsto_const[of "integral s (f 0)"]] - integrable_add[OF this(1) assms(1)[rule_format,of 0]] - then show ?thesis - unfolding sub - apply - - apply rule - defer - apply (subst(asm) integral_diff) - using assms(1) - apply auto - apply (rule LIMSEQ_imp_Suc) - apply assumption - done -qed - -lemma has_integral_monotone_convergence_increasing: - fixes f :: "nat \ 'a::euclidean_space \ real" - assumes f: "\k. (f k has_integral x k) s" - assumes "\k x. x \ s \ f k x \ f (Suc k) x" - assumes "\x. x \ s \ (\k. f k x) \ g x" - assumes "x \ x'" - shows "(g has_integral x') s" -proof - - have x_eq: "x = (\i. integral s (f i))" - by (simp add: integral_unique[OF f]) - then have x: "{integral s (f k) |k. True} = range x" - by auto - - have *: "g integrable_on s \ (\k. integral s (f k)) \ integral s g" - proof (intro monotone_convergence_increasing allI ballI assms) - show "bounded {integral s (f k) |k. True}" - unfolding x by (rule convergent_imp_bounded) fact - qed (auto intro: f) - then have "integral s g = x'" - by (intro LIMSEQ_unique[OF _ \x \ x'\]) (simp add: x_eq) - with * show ?thesis - by (simp add: has_integral_integral) -qed - -lemma monotone_convergence_decreasing: - fixes f :: "nat \ 'n::euclidean_space \ real" - assumes "\k. (f k) integrable_on s" - and "\k. \x\s. f (Suc k) x \ f k x" - and "\x\s. ((\k. f k x) \ g x) sequentially" - and "bounded {integral s (f k)| k. True}" - shows "g integrable_on s \ ((\k. integral s (f k)) \ integral s g) sequentially" -proof - - note assm = assms[rule_format] - have *: "{integral s (\x. - f k x) |k. True} = op *\<^sub>R (- 1) ` {integral s (f k)| k. True}" - apply safe - unfolding image_iff - apply (rule_tac x="integral s (f k)" in bexI) - prefer 3 - apply (rule_tac x=k in exI) - apply auto - done - have "(\x. - g x) integrable_on s \ - ((\k. integral s (\x. - f k x)) \ integral s (\x. - g x)) sequentially" - apply (rule monotone_convergence_increasing) - apply safe - apply (rule integrable_neg) - apply (rule assm) - defer - apply (rule tendsto_minus) - apply (rule assm) - apply assumption - unfolding * - apply (rule bounded_scaling) - using assm - apply auto - done - note * = conjunctD2[OF this] - show ?thesis - using integrable_neg[OF *(1)] tendsto_minus[OF *(2)] - by auto -qed - - -subsection \Absolute integrability (this is the same as Lebesgue integrability)\ - -definition absolutely_integrable_on (infixr "absolutely'_integrable'_on" 46) - where "f absolutely_integrable_on s \ f integrable_on s \ (\x. (norm(f x))) integrable_on s" - -lemma absolutely_integrable_onI[intro?]: - "f integrable_on s \ - (\x. (norm(f x))) integrable_on s \ f absolutely_integrable_on s" - unfolding absolutely_integrable_on_def - by auto - -lemma absolutely_integrable_onD[dest]: - assumes "f absolutely_integrable_on s" - shows "f integrable_on s" - and "(\x. norm (f x)) integrable_on s" - using assms - unfolding absolutely_integrable_on_def - by auto - -(*lemma absolutely_integrable_on_trans[simp]: - fixes f::"'n::euclidean_space \ real" - shows "(vec1 \ f) absolutely_integrable_on s \ f absolutely_integrable_on s" - unfolding absolutely_integrable_on_def o_def by auto*) - -lemma integral_norm_bound_integral: - fixes f :: "'n::euclidean_space \ 'a::banach" - assumes "f integrable_on s" - and "g integrable_on s" - and "\x\s. norm (f x) \ g x" - shows "norm (integral s f) \ integral s g" -proof - - have *: "\x y. (\e::real. 0 < e \ x < y + e) \ x \ y" - apply (rule ccontr) - apply (erule_tac x="x - y" in allE) - apply auto - done - have norm: "norm ig < dia + e" - if "norm sg \ dsa" - and "\dsa - dia\ < e / 2" - and "norm (sg - ig) < e / 2" - for e dsa dia and sg ig :: 'a - apply (rule le_less_trans[OF norm_triangle_sub[of ig sg]]) - apply (subst real_sum_of_halves[of e,symmetric]) - unfolding add.assoc[symmetric] - apply (rule add_le_less_mono) - defer - apply (subst norm_minus_commute) - apply (rule that(3)) - apply (rule order_trans[OF that(1)]) - using that(2) - apply arith - done - have lem: "norm (integral(cbox a b) f) \ integral (cbox a b) g" - if "f integrable_on cbox a b" - and "g integrable_on cbox a b" - and "\x\cbox a b. norm (f x) \ g x" - for f :: "'n \ 'a" and g a b - proof (rule *[rule_format]) - fix e :: real - assume "e > 0" - then have *: "e/2 > 0" - by auto - from integrable_integral[OF that(1),unfolded has_integral[of f],rule_format,OF *] - guess d1 .. note d1 = conjunctD2[OF this,rule_format] - from integrable_integral[OF that(2),unfolded has_integral[of g],rule_format,OF *] - guess d2 .. note d2 = conjunctD2[OF this,rule_format] - note gauge_inter[OF d1(1) d2(1)] - from fine_division_exists[OF this, of a b] guess p . note p=this - show "norm (integral (cbox a b) f) < integral (cbox a b) g + e" - apply (rule norm) - defer - apply (rule d2(2)[OF conjI[OF p(1)],unfolded real_norm_def]) - defer - apply (rule d1(2)[OF conjI[OF p(1)]]) - defer - apply (rule setsum_norm_le) - proof safe - fix x k - assume "(x, k) \ p" - note as = tagged_division_ofD(2-4)[OF p(1) this] - from this(3) guess u v by (elim exE) note uv=this - show "norm (content k *\<^sub>R f x) \ content k *\<^sub>R g x" - unfolding uv norm_scaleR - unfolding abs_of_nonneg[OF content_pos_le] real_scaleR_def - apply (rule mult_left_mono) - using that(3) as - apply auto - done - qed (insert p[unfolded fine_inter], auto) - qed - - { presume "\e. 0 < e \ norm (integral s f) < integral s g + e" - then show ?thesis by (rule *[rule_format]) auto } - fix e :: real - assume "e > 0" - then have e: "e/2 > 0" - by auto - note assms(1)[unfolded integrable_alt[of f]] note f=this[THEN conjunct1,rule_format] - note assms(2)[unfolded integrable_alt[of g]] note g=this[THEN conjunct1,rule_format] - from integrable_integral[OF assms(1),unfolded has_integral'[of f],rule_format,OF e] - guess B1 .. note B1=conjunctD2[OF this[rule_format],rule_format] - from integrable_integral[OF assms(2),unfolded has_integral'[of g],rule_format,OF e] - guess B2 .. note B2=conjunctD2[OF this[rule_format],rule_format] - from bounded_subset_cbox[OF bounded_ball, of "0::'n" "max B1 B2"] - guess a b by (elim exE) note ab=this[unfolded ball_max_Un] - - have "ball 0 B1 \ cbox a b" - using ab by auto - from B1(2)[OF this] guess z .. note z=conjunctD2[OF this] - have "ball 0 B2 \ cbox a b" - using ab by auto - from B2(2)[OF this] guess w .. note w=conjunctD2[OF this] - - show "norm (integral s f) < integral s g + e" - apply (rule norm) - apply (rule lem[OF f g, of a b]) - unfolding integral_unique[OF z(1)] integral_unique[OF w(1)] - defer - apply (rule w(2)[unfolded real_norm_def]) - apply (rule z(2)) - apply safe - apply (case_tac "x \ s") - unfolding if_P - apply (rule assms(3)[rule_format]) - apply auto - done -qed - -lemma integral_norm_bound_integral_component: - fixes f :: "'n::euclidean_space \ 'a::banach" - fixes g :: "'n \ 'b::euclidean_space" - assumes "f integrable_on s" - and "g integrable_on s" - and "\x\s. norm(f x) \ (g x)\k" - shows "norm (integral s f) \ (integral s g)\k" -proof - - have "norm (integral s f) \ integral s ((\x. x \ k) \ g)" - apply (rule integral_norm_bound_integral[OF assms(1)]) - apply (rule integrable_linear[OF assms(2)]) - apply rule - unfolding o_def - apply (rule assms) - done - then show ?thesis - unfolding o_def integral_component_eq[OF assms(2)] . -qed - -lemma has_integral_norm_bound_integral_component: - fixes f :: "'n::euclidean_space \ 'a::banach" - fixes g :: "'n \ 'b::euclidean_space" - assumes "(f has_integral i) s" - and "(g has_integral j) s" - and "\x\s. norm (f x) \ (g x)\k" - shows "norm i \ j\k" - using integral_norm_bound_integral_component[of f s g k] - unfolding integral_unique[OF assms(1)] integral_unique[OF assms(2)] - using assms - by auto - -lemma absolutely_integrable_le: - fixes f :: "'n::euclidean_space \ 'a::banach" - assumes "f absolutely_integrable_on s" - shows "norm (integral s f) \ integral s (\x. norm (f x))" - apply (rule integral_norm_bound_integral) - using assms - apply auto - done - -lemma absolutely_integrable_0[intro]: - "(\x. 0) absolutely_integrable_on s" - unfolding absolutely_integrable_on_def - by auto - -lemma absolutely_integrable_cmul[intro]: - "f absolutely_integrable_on s \ - (\x. c *\<^sub>R f x) absolutely_integrable_on s" - unfolding absolutely_integrable_on_def - using integrable_cmul[of f s c] - using integrable_cmul[of "\x. norm (f x)" s "\c\"] - by auto - -lemma absolutely_integrable_neg[intro]: - "f absolutely_integrable_on s \ - (\x. -f(x)) absolutely_integrable_on s" - apply (drule absolutely_integrable_cmul[where c="-1"]) - apply auto - done - -lemma absolutely_integrable_norm[intro]: - "f absolutely_integrable_on s \ - (\x. norm (f x)) absolutely_integrable_on s" - unfolding absolutely_integrable_on_def - by auto - -lemma absolutely_integrable_abs[intro]: - "f absolutely_integrable_on s \ - (\x. \f x::real\) absolutely_integrable_on s" - apply (drule absolutely_integrable_norm) - unfolding real_norm_def - apply assumption - done - -lemma absolutely_integrable_on_subinterval: - fixes f :: "'n::euclidean_space \ 'a::banach" - shows "f absolutely_integrable_on s \ - cbox a b \ s \ f absolutely_integrable_on cbox a b" - unfolding absolutely_integrable_on_def - by (metis integrable_on_subcbox) - -lemma absolutely_integrable_bounded_variation: - fixes f :: "'n::euclidean_space \ 'a::banach" - assumes "f absolutely_integrable_on UNIV" - obtains B where "\d. d division_of (\d) \ setsum (\k. norm(integral k f)) d \ B" - apply (rule that[of "integral UNIV (\x. norm (f x))"]) - apply safe -proof goal_cases - case prems: (1 d) - note d = division_ofD[OF prems(2)] - have "(\k\d. norm (integral k f)) \ (\i\d. integral i (\x. norm (f x)))" - apply (rule setsum_mono,rule absolutely_integrable_le) - apply (drule d(4)) - apply safe - apply (rule absolutely_integrable_on_subinterval[OF assms]) - apply auto - done - also have "\ \ integral (\d) (\x. norm (f x))" - apply (subst integral_combine_division_topdown[OF _ prems(2)]) - using integrable_on_subdivision[OF prems(2)] - using assms - apply auto - done - also have "\ \ integral UNIV (\x. norm (f x))" - apply (rule integral_subset_le) - using integrable_on_subdivision[OF prems(2)] - using assms - apply auto - done - finally show ?case . -qed - -lemma helplemma: - assumes "setsum (\x. norm (f x - g x)) s < e" - and "finite s" - shows "\setsum (\x. norm(f x)) s - setsum (\x. norm(g x)) s\ < e" - unfolding setsum_subtractf[symmetric] - apply (rule le_less_trans[OF setsum_abs]) - apply (rule le_less_trans[OF _ assms(1)]) - apply (rule setsum_mono) - apply (rule norm_triangle_ineq3) - done - -lemma bounded_variation_absolutely_integrable_interval: - fixes f :: "'n::euclidean_space \ 'm::euclidean_space" - assumes f: "f integrable_on cbox a b" - and *: "\d. d division_of (cbox a b) \ setsum (\k. norm(integral k f)) d \ B" - shows "f absolutely_integrable_on cbox a b" -proof - - let ?f = "\d. \k\d. norm (integral k f)" and ?D = "{d. d division_of (cbox a b)}" - have D_1: "?D \ {}" - by (rule elementary_interval[of a b]) auto - have D_2: "bdd_above (?f`?D)" - by (metis * mem_Collect_eq bdd_aboveI2) - note D = D_1 D_2 - let ?S = "SUP x:?D. ?f x" - show ?thesis - apply (rule absolutely_integrable_onI [OF f has_integral_integrable]) - apply (subst has_integral[of _ ?S]) - apply safe - proof goal_cases - case e: (1 e) - then have "?S - e / 2 < ?S" by simp - then obtain d where d: "d division_of (cbox a b)" "?S - e / 2 < (\k\d. norm (integral k f))" - unfolding less_cSUP_iff[OF D] by auto - note d' = division_ofD[OF this(1)] - - have "\x. \e>0. \i\d. x \ i \ ball x e \ i = {}" - proof - fix x - have "\da>0. \xa\\{i \ d. x \ i}. da \ dist x xa" - apply (rule separate_point_closed) - apply (rule closed_Union) - apply (rule finite_subset[OF _ d'(1)]) - using d'(4) - apply auto - done - then show "\e>0. \i\d. x \ i \ ball x e \ i = {}" - by force - qed - from choice[OF this] guess k .. note k=conjunctD2[OF this[rule_format],rule_format] - - have "e/2 > 0" - using e by auto - from henstock_lemma[OF assms(1) this] guess g . note g=this[rule_format] - let ?g = "\x. g x \ ball x (k x)" - show ?case - apply (rule_tac x="?g" in exI) - apply safe - proof - - show "gauge ?g" - using g(1) k(1) - unfolding gauge_def - by auto - fix p - assume "p tagged_division_of (cbox a b)" and "?g fine p" - note p = this(1) conjunctD2[OF this(2)[unfolded fine_inter]] - note p' = tagged_division_ofD[OF p(1)] - define p' where "p' = {(x,k) | x k. \i l. x \ i \ i \ d \ (x,l) \ p \ k = i \ l}" - have gp': "g fine p'" - using p(2) - unfolding p'_def fine_def - by auto - have p'': "p' tagged_division_of (cbox a b)" - apply (rule tagged_division_ofI) - proof - - show "finite p'" - apply (rule finite_subset[of _ "(\(k,(x,l)). (x,k \ l)) ` - {(k,xl) | k xl. k \ d \ xl \ p}"]) - unfolding p'_def - defer - apply (rule finite_imageI,rule finite_product_dependent[OF d'(1) p'(1)]) - apply safe - unfolding image_iff - apply (rule_tac x="(i,x,l)" in bexI) - apply auto - done - fix x k - assume "(x, k) \ p'" - then have "\i l. x \ i \ i \ d \ (x, l) \ p \ k = i \ l" - unfolding p'_def by auto - then guess i l by (elim exE) note il=conjunctD4[OF this] - show "x \ k" and "k \ cbox a b" - using p'(2-3)[OF il(3)] il by auto - show "\a b. k = cbox a b" - unfolding il using p'(4)[OF il(3)] d'(4)[OF il(2)] - apply safe - unfolding inter_interval - apply auto - done - next - fix x1 k1 - assume "(x1, k1) \ p'" - then have "\i l. x1 \ i \ i \ d \ (x1, l) \ p \ k1 = i \ l" - unfolding p'_def by auto - then guess i1 l1 by (elim exE) note il1=conjunctD4[OF this] - fix x2 k2 - assume "(x2,k2)\p'" - then have "\i l. x2 \ i \ i \ d \ (x2, l) \ p \ k2 = i \ l" - unfolding p'_def by auto - then guess i2 l2 by (elim exE) note il2=conjunctD4[OF this] - assume "(x1, k1) \ (x2, k2)" - then have "interior i1 \ interior i2 = {} \ interior l1 \ interior l2 = {}" - using d'(5)[OF il1(2) il2(2)] p'(5)[OF il1(3) il2(3)] - unfolding il1 il2 - by auto - then show "interior k1 \ interior k2 = {}" - unfolding il1 il2 by auto - next - have *: "\(x, X) \ p'. X \ cbox a b" - unfolding p'_def using d' by auto - show "\{k. \x. (x, k) \ p'} = cbox a b" - apply rule - apply (rule Union_least) - unfolding mem_Collect_eq - apply (erule exE) - apply (drule *[rule_format]) - apply safe - proof - - fix y - assume y: "y \ cbox a b" - then have "\x l. (x, l) \ p \ y\l" - unfolding p'(6)[symmetric] by auto - then guess x l by (elim exE) note xl=conjunctD2[OF this] - then have "\k. k \ d \ y \ k" - using y unfolding d'(6)[symmetric] by auto - then guess i .. note i = conjunctD2[OF this] - have "x \ i" - using fineD[OF p(3) xl(1)] - using k(2)[OF i(1), of x] - using i(2) xl(2) - by auto - then show "y \ \{k. \x. (x, k) \ p'}" - unfolding p'_def Union_iff - apply (rule_tac x="i \ l" in bexI) - using i xl - apply auto - done - qed - qed - - then have "(\(x, k)\p'. norm (content k *\<^sub>R f x - integral k f)) < e / 2" - apply - - apply (rule g(2)[rule_format]) - unfolding tagged_division_of_def - apply safe - apply (rule gp') - done - then have **: "\(\(x,k)\p'. norm (content k *\<^sub>R f x)) - (\(x,k)\p'. norm (integral k f))\ < e / 2" - unfolding split_def - using p'' - by (force intro!: helplemma) - - have p'alt: "p' = {(x,(i \ l)) | x i l. (x,l) \ p \ i \ d \ i \ l \ {}}" - proof (safe, goal_cases) - case prems: (2 _ _ x i l) - have "x \ i" - using fineD[OF p(3) prems(1)] k(2)[OF prems(2), of x] prems(4-) - by auto - then have "(x, i \ l) \ p'" - unfolding p'_def - using prems - apply safe - apply (rule_tac x=x in exI) - apply (rule_tac x="i \ l" in exI) - apply safe - using prems - apply auto - done - then show ?case - using prems(3) by auto - next - fix x k - assume "(x, k) \ p'" - then have "\i l. x \ i \ i \ d \ (x, l) \ p \ k = i \ l" - unfolding p'_def by auto - then guess i l by (elim exE) note il=conjunctD4[OF this] - then show "\y i l. (x, k) = (y, i \ l) \ (y, l) \ p \ i \ d \ i \ l \ {}" - apply (rule_tac x=x in exI) - apply (rule_tac x=i in exI) - apply (rule_tac x=l in exI) - using p'(2)[OF il(3)] - apply auto - done - qed - have sum_p': "(\(x, k)\p'. norm (integral k f)) = (\k\snd ` p'. norm (integral k f))" - apply (subst setsum.over_tagged_division_lemma[OF p'',of "\k. norm (integral k f)"]) - unfolding norm_eq_zero - apply (rule integral_null) - apply assumption - apply rule - done - note snd_p = division_ofD[OF division_of_tagged_division[OF p(1)]] - - have *: "\sni sni' sf sf'. \sf' - sni'\ < e / 2 \ ?S - e / 2 < sni \ sni' \ ?S \ - sni \ sni' \ sf' = sf \ \sf - ?S\ < e" - by arith - show "norm ((\(x, k)\p. content k *\<^sub>R norm (f x)) - ?S) < e" - unfolding real_norm_def - apply (rule *[rule_format,OF **]) - apply safe - apply(rule d(2)) - proof goal_cases - case 1 - show ?case - by (auto simp: sum_p' division_of_tagged_division[OF p''] D intro!: cSUP_upper) - next - case 2 - have *: "{k \ l | k l. k \ d \ l \ snd ` p} = - (\(k,l). k \ l) ` {(k,l)|k l. k \ d \ l \ snd ` p}" - by auto - have "(\k\d. norm (integral k f)) \ (\i\d. \l\snd ` p. norm (integral (i \ l) f))" - proof (rule setsum_mono, goal_cases) - case k: (1 k) - from d'(4)[OF this] guess u v by (elim exE) note uv=this - define d' where "d' = {cbox u v \ l |l. l \ snd ` p \ cbox u v \ l \ {}}" - note uvab = d'(2)[OF k[unfolded uv]] - have "d' division_of cbox u v" - apply (subst d'_def) - apply (rule division_inter_1) - apply (rule division_of_tagged_division[OF p(1)]) - apply (rule uvab) - done - then have "norm (integral k f) \ setsum (\k. norm (integral k f)) d'" - unfolding uv - apply (subst integral_combine_division_topdown[of _ _ d']) - apply (rule integrable_on_subcbox[OF assms(1) uvab]) - apply assumption - apply (rule setsum_norm_le) - apply auto - done - also have "\ = (\k\{k \ l |l. l \ snd ` p}. norm (integral k f))" - apply (rule setsum.mono_neutral_left) - apply (subst simple_image) - apply (rule finite_imageI)+ - apply fact - unfolding d'_def uv - apply blast - proof (rule, goal_cases) - case prems: (1 i) - then have "i \ {cbox u v \ l |l. l \ snd ` p}" - by auto - from this[unfolded mem_Collect_eq] guess l .. note l=this - then have "cbox u v \ l = {}" - using prems by auto - then show ?case - using l by auto - qed - also have "\ = (\l\snd ` p. norm (integral (k \ l) f))" - unfolding simple_image - apply (rule setsum.reindex_nontrivial [unfolded o_def]) - apply (rule finite_imageI) - apply (rule p') - proof goal_cases - case prems: (1 l y) - have "interior (k \ l) \ interior (l \ y)" - apply (subst(2) interior_Int) - apply (rule Int_greatest) - defer - apply (subst prems(4)) - apply auto - done - then have *: "interior (k \ l) = {}" - using snd_p(5)[OF prems(1-3)] by auto - from d'(4)[OF k] snd_p(4)[OF prems(1)] guess u1 v1 u2 v2 by (elim exE) note uv=this - show ?case - using * - unfolding uv inter_interval content_eq_0_interior[symmetric] - by auto - qed - finally show ?case . - qed - also have "\ = (\(i,l)\{(i, l) |i l. i \ d \ l \ snd ` p}. norm (integral (i\l) f))" - apply (subst sum_sum_product[symmetric]) - apply fact - using p'(1) - apply auto - done - also have "\ = (\x\{(i, l) |i l. i \ d \ l \ snd ` p}. norm (integral (case_prod op \ x) f))" - unfolding split_def .. - also have "\ = (\k\{i \ l |i l. i \ d \ l \ snd ` p}. norm (integral k f))" - unfolding * - apply (rule setsum.reindex_nontrivial [symmetric, unfolded o_def]) - apply (rule finite_product_dependent) - apply fact - apply (rule finite_imageI) - apply (rule p') - unfolding split_paired_all mem_Collect_eq split_conv o_def - proof - - note * = division_ofD(4,5)[OF division_of_tagged_division,OF p(1)] - fix l1 l2 k1 k2 - assume as: - "(l1, k1) \ (l2, k2)" - "l1 \ k1 = l2 \ k2" - "\i l. (l1, k1) = (i, l) \ i \ d \ l \ snd ` p" - "\i l. (l2, k2) = (i, l) \ i \ d \ l \ snd ` p" - then have "l1 \ d" and "k1 \ snd ` p" - by auto from d'(4)[OF this(1)] *(1)[OF this(2)] - guess u1 v1 u2 v2 by (elim exE) note uv=this - have "l1 \ l2 \ k1 \ k2" - using as by auto - then have "interior k1 \ interior k2 = {} \ interior l1 \ interior l2 = {}" - apply - - apply (erule disjE) - apply (rule disjI2) - apply (rule d'(5)) - prefer 4 - apply (rule disjI1) - apply (rule *) - using as - apply auto - done - moreover have "interior (l1 \ k1) = interior (l2 \ k2)" - using as(2) by auto - ultimately have "interior(l1 \ k1) = {}" - by auto - then show "norm (integral (l1 \ k1) f) = 0" - unfolding uv inter_interval - unfolding content_eq_0_interior[symmetric] - by auto - qed - also have "\ = (\(x, k)\p'. norm (integral k f))" - unfolding sum_p' - apply (rule setsum.mono_neutral_right) - apply (subst *) - apply (rule finite_imageI[OF finite_product_dependent]) - apply fact - apply (rule finite_imageI[OF p'(1)]) - apply safe - proof goal_cases - case (2 i ia l a b) - then have "ia \ b = {}" - unfolding p'alt image_iff Bex_def not_ex - apply (erule_tac x="(a, ia \ b)" in allE) - apply auto - done - then show ?case - by auto - next - case (1 x a b) - then show ?case - unfolding p'_def - apply safe - apply (rule_tac x=i in exI) - apply (rule_tac x=l in exI) - unfolding snd_conv image_iff - apply safe - apply (rule_tac x="(a,l)" in bexI) - apply auto - done - qed - finally show ?case . - next - case 3 - let ?S = "{(x, i \ l) |x i l. (x, l) \ p \ i \ d}" - have Sigma_alt: "\s t. s \ t = {(i, j) |i j. i \ s \ j \ t}" - by auto - have *: "?S = (\(xl,i). (fst xl, snd xl \ i)) ` (p \ d)" - apply safe - unfolding image_iff - apply (rule_tac x="((x,l),i)" in bexI) - apply auto - done - note pdfin = finite_cartesian_product[OF p'(1) d'(1)] - have "(\(x, k)\p'. norm (content k *\<^sub>R f x)) = (\(x, k)\?S. \content k\ * norm (f x))" - unfolding norm_scaleR - apply (rule setsum.mono_neutral_left) - apply (subst *) - apply (rule finite_imageI) - apply fact - unfolding p'alt - apply blast - apply safe - apply (rule_tac x=x in exI) - apply (rule_tac x=i in exI) - apply (rule_tac x=l in exI) - apply auto - done - also have "\ = (\((x,l),i)\p \ d. \content (l \ i)\ * norm (f x))" - unfolding * - apply (subst setsum.reindex_nontrivial) - apply fact - unfolding split_paired_all - unfolding o_def split_def snd_conv fst_conv mem_Sigma_iff prod.inject - apply (elim conjE) - proof - - fix x1 l1 k1 x2 l2 k2 - assume as: "(x1, l1) \ p" "(x2, l2) \ p" "k1 \ d" "k2 \ d" - "x1 = x2" "l1 \ k1 = l2 \ k2" "\ ((x1 = x2 \ l1 = l2) \ k1 = k2)" - from d'(4)[OF as(3)] p'(4)[OF as(1)] guess u1 v1 u2 v2 by (elim exE) note uv=this - from as have "l1 \ l2 \ k1 \ k2" - by auto - then have "interior k1 \ interior k2 = {} \ interior l1 \ interior l2 = {}" - apply - - apply (erule disjE) - apply (rule disjI2) - defer - apply (rule disjI1) - apply (rule d'(5)[OF as(3-4)]) - apply assumption - apply (rule p'(5)[OF as(1-2)]) - apply auto - done - moreover have "interior (l1 \ k1) = interior (l2 \ k2)" - unfolding as .. - ultimately have "interior (l1 \ k1) = {}" - by auto - then show "\content (l1 \ k1)\ * norm (f x1) = 0" - unfolding uv inter_interval - unfolding content_eq_0_interior[symmetric] - by auto - qed safe - also have "\ = (\(x, k)\p. content k *\<^sub>R norm (f x))" - unfolding Sigma_alt - apply (subst sum_sum_product[symmetric]) - apply (rule p') - apply rule - apply (rule d') - apply (rule setsum.cong) - apply (rule refl) - unfolding split_paired_all split_conv - proof - - fix x l - assume as: "(x, l) \ p" - note xl = p'(2-4)[OF this] - from this(3) guess u v by (elim exE) note uv=this - have "(\i\d. \content (l \ i)\) = (\k\d. content (k \ cbox u v))" - apply (rule setsum.cong) - apply (rule refl) - apply (drule d'(4)) - apply safe - apply (subst Int_commute) - unfolding inter_interval uv - apply (subst abs_of_nonneg) - apply auto - done - also have "\ = setsum content {k \ cbox u v| k. k \ d}" - unfolding simple_image - apply (rule setsum.reindex_nontrivial [unfolded o_def, symmetric]) - apply (rule d') - proof goal_cases - case prems: (1 k y) - from d'(4)[OF this(1)] d'(4)[OF this(2)] - guess u1 v1 u2 v2 by (elim exE) note uv=this - have "{} = interior ((k \ y) \ cbox u v)" - apply (subst interior_Int) - using d'(5)[OF prems(1-3)] - apply auto - done - also have "\ = interior (y \ (k \ cbox u v))" - by auto - also have "\ = interior (k \ cbox u v)" - unfolding prems(4) by auto - finally show ?case - unfolding uv inter_interval content_eq_0_interior .. - qed - also have "\ = setsum content {cbox u v \ k |k. k \ d \ cbox u v \ k \ {}}" - apply (rule setsum.mono_neutral_right) - unfolding simple_image - apply (rule finite_imageI) - apply (rule d') - apply blast - apply safe - apply (rule_tac x=k in exI) - proof goal_cases - case prems: (1 i k) - from d'(4)[OF this(1)] guess a b by (elim exE) note ab=this - have "interior (k \ cbox u v) \ {}" - using prems(2) - unfolding ab inter_interval content_eq_0_interior - by auto - then show ?case - using prems(1) - using interior_subset[of "k \ cbox u v"] - by auto - qed - finally show "(\i\d. \content (l \ i)\ * norm (f x)) = content l *\<^sub>R norm (f x)" - unfolding setsum_left_distrib[symmetric] real_scaleR_def - apply (subst(asm) additive_content_division[OF division_inter_1[OF d(1)]]) - using xl(2)[unfolded uv] - unfolding uv - apply auto - done - qed - finally show ?case . - qed - qed - qed -qed - -lemma bounded_variation_absolutely_integrable: - fixes f :: "'n::euclidean_space \ 'm::euclidean_space" - assumes "f integrable_on UNIV" - and "\d. d division_of (\d) \ setsum (\k. norm (integral k f)) d \ B" - shows "f absolutely_integrable_on UNIV" -proof (rule absolutely_integrable_onI, fact, rule) - let ?f = "\d. \k\d. norm (integral k f)" and ?D = "{d. d division_of (\d)}" - have D_1: "?D \ {}" - by (rule elementary_interval) auto - have D_2: "bdd_above (?f`?D)" - by (intro bdd_aboveI2[where M=B] assms(2)[rule_format]) simp - note D = D_1 D_2 - let ?S = "SUP d:?D. ?f d" - have f_int: "\a b. f absolutely_integrable_on cbox a b" - apply (rule bounded_variation_absolutely_integrable_interval[where B=B]) - apply (rule integrable_on_subcbox[OF assms(1)]) - defer - apply safe - apply (rule assms(2)[rule_format]) - apply auto - done - show "((\x. norm (f x)) has_integral ?S) UNIV" - apply (subst has_integral_alt') - apply safe - proof goal_cases - case (1 a b) - show ?case - using f_int[of a b] by auto - next - case prems: (2 e) - have "\y\setsum (\k. norm (integral k f)) ` {d. d division_of \d}. \ y \ ?S - e" - proof (rule ccontr) - assume "\ ?thesis" - then have "?S \ ?S - e" - by (intro cSUP_least[OF D(1)]) auto - then show False - using prems by auto - qed - then obtain K where *: "\x\{d. d division_of \d}. K = (\k\x. norm (integral k f))" - "SUPREMUM {d. d division_of \d} (setsum (\k. norm (integral k f))) - e < K" - by (auto simp add: image_iff not_le) - from this(1) obtain d where "d division_of \d" - and "K = (\k\d. norm (integral k f))" - by auto - note d = this(1) *(2)[unfolded this(2)] - note d'=division_ofD[OF this(1)] - have "bounded (\d)" - by (rule elementary_bounded,fact) - from this[unfolded bounded_pos] obtain K where - K: "0 < K" "\x\\d. norm x \ K" by auto - show ?case - apply (rule_tac x="K + 1" in exI) - apply safe - proof - - fix a b :: 'n - assume ab: "ball 0 (K + 1) \ cbox a b" - have *: "\s s1. ?S - e < s1 \ s1 \ s \ s < ?S + e \ \s - ?S\ < e" - by arith - show "norm (integral (cbox a b) (\x. if x \ UNIV then norm (f x) else 0) - ?S) < e" - unfolding real_norm_def - apply (rule *[rule_format]) - apply safe - apply (rule d(2)) - proof goal_cases - case 1 - have "(\k\d. norm (integral k f)) \ setsum (\k. integral k (\x. norm (f x))) d" - apply (rule setsum_mono) - apply (rule absolutely_integrable_le) - apply (drule d'(4)) - apply safe - apply (rule f_int) - done - also have "\ = integral (\d) (\x. norm (f x))" - apply (rule integral_combine_division_bottomup[symmetric]) - apply (rule d) - unfolding forall_in_division[OF d(1)] - using f_int - apply auto - done - also have "\ \ integral (cbox a b) (\x. if x \ UNIV then norm (f x) else 0)" - proof - - have "\d \ cbox a b" - apply rule - apply (drule K(2)[rule_format]) - apply (rule ab[unfolded subset_eq,rule_format]) - apply (auto simp add: dist_norm) - done - then show ?thesis - apply - - apply (subst if_P) - apply rule - apply (rule integral_subset_le) - defer - apply (rule integrable_on_subdivision[of _ _ _ "cbox a b"]) - apply (rule d) - using f_int[of a b] - apply auto - done - qed - finally show ?case . - next - note f = absolutely_integrable_onD[OF f_int[of a b]] - note * = this(2)[unfolded has_integral_integral has_integral[of "\x. norm (f x)"],rule_format] - have "e/2>0" - using \e > 0\ by auto - from * [OF this] obtain d1 where - d1: "gauge d1" "\p. p tagged_division_of (cbox a b) \ d1 fine p \ - norm ((\(x, k)\p. content k *\<^sub>R norm (f x)) - integral (cbox a b) (\x. norm (f x))) < e / 2" - by auto - from henstock_lemma [OF f(1) \e/2>0\] obtain d2 where - d2: "gauge d2" "\p. p tagged_partial_division_of (cbox a b) \ d2 fine p \ - (\(x, k)\p. norm (content k *\<^sub>R f x - integral k f)) < e / 2" . - obtain p where - p: "p tagged_division_of (cbox a b)" "d1 fine p" "d2 fine p" - by (rule fine_division_exists [OF gauge_inter [OF d1(1) d2(1)], of a b]) - (auto simp add: fine_inter) - have *: "\sf sf' si di. sf' = sf \ si \ ?S \ \sf - si\ < e / 2 \ - \sf' - di\ < e / 2 \ di < ?S + e" - by arith - show "integral (cbox a b) (\x. if x \ UNIV then norm (f x) else 0) < ?S + e" - apply (subst if_P) - apply rule - proof (rule *[rule_format]) - show "\(\(x,k)\p. norm (content k *\<^sub>R f x)) - (\(x,k)\p. norm (integral k f))\ < e / 2" - unfolding split_def - apply (rule helplemma) - using d2(2)[rule_format,of p] - using p(1,3) - unfolding tagged_division_of_def split_def - apply auto - done - show "\(\(x, k)\p. content k *\<^sub>R norm (f x)) - integral (cbox a b) (\x. norm(f x))\ < e / 2" - using d1(2)[rule_format,OF conjI[OF p(1,2)]] - by (simp only: real_norm_def) - show "(\(x, k)\p. content k *\<^sub>R norm (f x)) = (\(x, k)\p. norm (content k *\<^sub>R f x))" - apply (rule setsum.cong) - apply (rule refl) - unfolding split_paired_all split_conv - apply (drule tagged_division_ofD(4)[OF p(1)]) - unfolding norm_scaleR - apply (subst abs_of_nonneg) - apply auto - done - show "(\(x, k)\p. norm (integral k f)) \ ?S" - using partial_division_of_tagged_division[of p "cbox a b"] p(1) - apply (subst setsum.over_tagged_division_lemma[OF p(1)]) - apply (simp add: integral_null) - apply (intro cSUP_upper2[OF D(2), of "snd ` p"]) - apply (auto simp: tagged_partial_division_of_def) - done - qed - qed - qed (insert K, auto) - qed -qed - -lemma absolutely_integrable_restrict_univ: - "(\x. if x \ s then f x else (0::'a::banach)) absolutely_integrable_on UNIV \ - f absolutely_integrable_on s" - unfolding absolutely_integrable_on_def if_distrib norm_zero integrable_restrict_univ .. - -lemma absolutely_integrable_add[intro]: - fixes f g :: "'n::euclidean_space \ 'm::euclidean_space" - assumes "f absolutely_integrable_on s" - and "g absolutely_integrable_on s" - shows "(\x. f x + g x) absolutely_integrable_on s" -proof - - let ?P = "\f g::'n \ 'm. f absolutely_integrable_on UNIV \ - g absolutely_integrable_on UNIV \ (\x. f x + g x) absolutely_integrable_on UNIV" - { - presume as: "PROP ?P" - note a = absolutely_integrable_restrict_univ[symmetric] - have *: "\x. (if x \ s then f x else 0) + (if x \ s then g x else 0) = - (if x \ s then f x + g x else 0)" by auto - show ?thesis - apply (subst a) - using as[OF assms[unfolded a[of f] a[of g]]] - apply (simp only: *) - done - } - fix f g :: "'n \ 'm" - assume assms: "f absolutely_integrable_on UNIV" "g absolutely_integrable_on UNIV" - note absolutely_integrable_bounded_variation - from this[OF assms(1)] this[OF assms(2)] guess B1 B2 . note B=this[rule_format] - show "(\x. f x + g x) absolutely_integrable_on UNIV" - apply (rule bounded_variation_absolutely_integrable[of _ "B1+B2"]) - apply (rule integrable_add) - prefer 3 - apply safe - proof goal_cases - case prems: (1 d) - have "\k. k \ d \ f integrable_on k \ g integrable_on k" - apply (drule division_ofD(4)[OF prems]) - apply safe - apply (rule_tac[!] integrable_on_subcbox[of _ UNIV]) - using assms - apply auto - done - then have "(\k\d. norm (integral k (\x. f x + g x))) \ - (\k\d. norm (integral k f)) + (\k\d. norm (integral k g))" - apply - - unfolding setsum.distrib [symmetric] - apply (rule setsum_mono) - apply (subst integral_add) - prefer 3 - apply (rule norm_triangle_ineq) - apply auto - done - also have "\ \ B1 + B2" - using B(1)[OF prems] B(2)[OF prems] by auto - finally show ?case . - qed (insert assms, auto) -qed - -lemma absolutely_integrable_sub[intro]: - fixes f g :: "'n::euclidean_space \ 'm::euclidean_space" - assumes "f absolutely_integrable_on s" - and "g absolutely_integrable_on s" - shows "(\x. f x - g x) absolutely_integrable_on s" - using absolutely_integrable_add[OF assms(1) absolutely_integrable_neg[OF assms(2)]] - by (simp add: algebra_simps) - -lemma absolutely_integrable_linear: - fixes f :: "'m::euclidean_space \ 'n::euclidean_space" - and h :: "'n::euclidean_space \ 'p::euclidean_space" - assumes "f absolutely_integrable_on s" - and "bounded_linear h" - shows "(h \ f) absolutely_integrable_on s" -proof - - { - presume as: "\f::'m \ 'n. \h::'n \ 'p. f absolutely_integrable_on UNIV \ - bounded_linear h \ (h \ f) absolutely_integrable_on UNIV" - note a = absolutely_integrable_restrict_univ[symmetric] - show ?thesis - apply (subst a) - using as[OF assms[unfolded a[of f] a[of g]]] - apply (simp only: o_def if_distrib linear_simps[OF assms(2)]) - done - } - fix f :: "'m \ 'n" - fix h :: "'n \ 'p" - assume assms: "f absolutely_integrable_on UNIV" "bounded_linear h" - from absolutely_integrable_bounded_variation[OF assms(1)] guess B . note B=this - from bounded_linear.pos_bounded[OF assms(2)] guess b .. note b=conjunctD2[OF this] - show "(h \ f) absolutely_integrable_on UNIV" - apply (rule bounded_variation_absolutely_integrable[of _ "B * b"]) - apply (rule integrable_linear[OF _ assms(2)]) - apply safe - proof goal_cases - case prems: (2 d) - have "(\k\d. norm (integral k (h \ f))) \ setsum (\k. norm(integral k f)) d * b" - unfolding setsum_left_distrib - apply (rule setsum_mono) - proof goal_cases - case (1 k) - from division_ofD(4)[OF prems this] - guess u v by (elim exE) note uv=this - have *: "f integrable_on k" - unfolding uv - apply (rule integrable_on_subcbox[of _ UNIV]) - using assms - apply auto - done - note this[unfolded has_integral_integral] - note has_integral_linear[OF this assms(2)] integrable_linear[OF * assms(2)] - note * = has_integral_unique[OF this(2)[unfolded has_integral_integral] this(1)] - show ?case - unfolding * using b by auto - qed - also have "\ \ B * b" - apply (rule mult_right_mono) - using B prems b - apply auto - done - finally show ?case . - qed (insert assms, auto) -qed - -lemma absolutely_integrable_setsum: - fixes f :: "'a \ 'n::euclidean_space \ 'm::euclidean_space" - assumes "finite t" - and "\a. a \ t \ (f a) absolutely_integrable_on s" - shows "(\x. setsum (\a. f a x) t) absolutely_integrable_on s" - using assms(1,2) - by induct auto - -lemma absolutely_integrable_vector_abs: - fixes f :: "'a::euclidean_space => 'b::euclidean_space" - and T :: "'c::euclidean_space \ 'b" - assumes f: "f absolutely_integrable_on s" - shows "(\x. (\i\Basis. \f x\T i\ *\<^sub>R i)) absolutely_integrable_on s" - (is "?Tf absolutely_integrable_on s") -proof - - have if_distrib: "\P A B x. (if P then A else B) *\<^sub>R x = (if P then A *\<^sub>R x else B *\<^sub>R x)" - by simp - have *: "\x. ?Tf x = (\i\Basis. - ((\y. (\j\Basis. (if j = i then y else 0) *\<^sub>R j)) o - (\x. (norm (\j\Basis. (if j = i then f x\T i else 0) *\<^sub>R j)))) x)" - by (simp add: comp_def if_distrib setsum.If_cases) - show ?thesis - unfolding * - apply (rule absolutely_integrable_setsum[OF finite_Basis]) - apply (rule absolutely_integrable_linear) - apply (rule absolutely_integrable_norm) - apply (rule absolutely_integrable_linear[OF f, unfolded o_def]) - apply (auto simp: linear_linear euclidean_eq_iff[where 'a='c] inner_simps intro!: linearI) - done -qed - -lemma absolutely_integrable_max: - fixes f g :: "'m::euclidean_space \ 'n::euclidean_space" - assumes "f absolutely_integrable_on s" - and "g absolutely_integrable_on s" - shows "(\x. (\i\Basis. max (f(x)\i) (g(x)\i) *\<^sub>R i)::'n) absolutely_integrable_on s" -proof - - have *:"\x. (1 / 2) *\<^sub>R (((\i\Basis. \(f x - g x) \ i\ *\<^sub>R i)::'n) + (f x + g x)) = - (\i\Basis. max (f(x)\i) (g(x)\i) *\<^sub>R i)" - unfolding euclidean_eq_iff[where 'a='n] by (auto simp: inner_simps) - note absolutely_integrable_sub[OF assms] absolutely_integrable_add[OF assms] - note absolutely_integrable_vector_abs[OF this(1), where T="\x. x"] this(2) - note absolutely_integrable_add[OF this] - note absolutely_integrable_cmul[OF this, of "1/2"] - then show ?thesis unfolding * . -qed - -lemma absolutely_integrable_min: - fixes f g::"'m::euclidean_space \ 'n::euclidean_space" - assumes "f absolutely_integrable_on s" - and "g absolutely_integrable_on s" - shows "(\x. (\i\Basis. min (f(x)\i) (g(x)\i) *\<^sub>R i)::'n) absolutely_integrable_on s" -proof - - have *:"\x. (1 / 2) *\<^sub>R ((f x + g x) - (\i\Basis. \(f x - g x) \ i\ *\<^sub>R i::'n)) = - (\i\Basis. min (f(x)\i) (g(x)\i) *\<^sub>R i)" - unfolding euclidean_eq_iff[where 'a='n] by (auto simp: inner_simps) - - note absolutely_integrable_add[OF assms] absolutely_integrable_sub[OF assms] - note this(1) absolutely_integrable_vector_abs[OF this(2), where T="\x. x"] - note absolutely_integrable_sub[OF this] - note absolutely_integrable_cmul[OF this,of "1/2"] - then show ?thesis unfolding * . -qed - -lemma absolutely_integrable_abs_eq: - fixes f::"'n::euclidean_space \ 'm::euclidean_space" - shows "f absolutely_integrable_on s \ f integrable_on s \ - (\x. (\i\Basis. \f x\i\ *\<^sub>R i)::'m) integrable_on s" - (is "?l = ?r") -proof - assume ?l - then show ?r - apply - - apply rule - defer - apply (drule absolutely_integrable_vector_abs) - apply auto - done -next - assume ?r - { - presume lem: "\f::'n \ 'm. f integrable_on UNIV \ - (\x. (\i\Basis. \f x\i\ *\<^sub>R i)::'m) integrable_on UNIV \ - f absolutely_integrable_on UNIV" - have *: "\x. (\i\Basis. \(if x \ s then f x else 0) \ i\ *\<^sub>R i) = - (if x \ s then (\i\Basis. \f x \ i\ *\<^sub>R i) else (0::'m))" - unfolding euclidean_eq_iff[where 'a='m] - by auto - show ?l - apply (subst absolutely_integrable_restrict_univ[symmetric]) - apply (rule lem) - unfolding integrable_restrict_univ * - using \?r\ - apply auto - done - } - fix f :: "'n \ 'm" - assume assms: "f integrable_on UNIV" "(\x. (\i\Basis. \f x\i\ *\<^sub>R i)::'m) integrable_on UNIV" - let ?B = "\i\Basis. integral UNIV (\x. (\i\Basis. \f x\i\ *\<^sub>R i)::'m) \ i" - show "f absolutely_integrable_on UNIV" - apply (rule bounded_variation_absolutely_integrable[OF assms(1), where B="?B"]) - apply safe - proof goal_cases - case d: (1 d) - note d'=division_ofD[OF d] - have "(\k\d. norm (integral k f)) \ - (\k\d. setsum (op \ (integral k (\x. (\i\Basis. \f x\i\ *\<^sub>R i)::'m))) Basis)" - apply (rule setsum_mono) - apply (rule order_trans[OF norm_le_l1]) - apply (rule setsum_mono) - unfolding lessThan_iff - proof - - fix k - fix i :: 'm - assume "k \ d" and i: "i \ Basis" - from d'(4)[OF this(1)] guess a b by (elim exE) note ab=this - show "\integral k f \ i\ \ integral k (\x. (\i\Basis. \f x\i\ *\<^sub>R i)::'m) \ i" - apply (rule abs_leI) - unfolding inner_minus_left[symmetric] - defer - apply (subst integral_neg[symmetric]) - apply (rule_tac[1-2] integral_component_le[OF i]) - using integrable_on_subcbox[OF assms(1),of a b] - integrable_on_subcbox[OF assms(2),of a b] i integrable_neg - unfolding ab - apply auto - done - qed - also have "\ \ setsum (op \ (integral UNIV (\x. (\i\Basis. \f x\i\ *\<^sub>R i)::'m))) Basis" - apply (subst setsum.commute) - apply (rule setsum_mono) - proof goal_cases - case (1 j) - have *: "(\x. \i\Basis. \f x\i\ *\<^sub>R i::'m) integrable_on \d" - using integrable_on_subdivision[OF d assms(2)] by auto - have "(\i\d. integral i (\x. \i\Basis. \f x\i\ *\<^sub>R i::'m) \ j) = - integral (\d) (\x. \i\Basis. \f x\i\ *\<^sub>R i::'m) \ j" - unfolding inner_setsum_left[symmetric] integral_combine_division_topdown[OF * d] .. - also have "\ \ integral UNIV (\x. \i\Basis. \f x\i\ *\<^sub>R i::'m) \ j" - apply (rule integral_subset_component_le) - using assms * \j \ Basis\ - apply auto - done - finally show ?case . - qed - finally show ?case . - qed -qed - -lemma nonnegative_absolutely_integrable: - fixes f :: "'n::euclidean_space \ 'm::euclidean_space" - assumes "\x\s. \i\Basis. 0 \ f x \ i" - and "f integrable_on s" - shows "f absolutely_integrable_on s" - unfolding absolutely_integrable_abs_eq - apply rule - apply (rule assms)thm integrable_eq - apply (rule integrable_eq[of _ f]) - using assms - apply (auto simp: euclidean_eq_iff[where 'a='m]) - done - -lemma absolutely_integrable_integrable_bound: - fixes f :: "'n::euclidean_space \ 'm::euclidean_space" - assumes "\x\s. norm (f x) \ g x" - and "f integrable_on s" - and "g integrable_on s" - shows "f absolutely_integrable_on s" -proof - - { - presume *: "\f::'n \ 'm. \g. \x. norm (f x) \ g x \ f integrable_on UNIV \ - g integrable_on UNIV \ f absolutely_integrable_on UNIV" - show ?thesis - apply (subst absolutely_integrable_restrict_univ[symmetric]) - apply (rule *[of _ "\x. if x\s then g x else 0"]) - using assms - unfolding integrable_restrict_univ - apply auto - done - } - fix g - fix f :: "'n \ 'm" - assume assms: "\x. norm (f x) \ g x" "f integrable_on UNIV" "g integrable_on UNIV" - show "f absolutely_integrable_on UNIV" - apply (rule bounded_variation_absolutely_integrable[OF assms(2),where B="integral UNIV g"]) - apply safe - proof goal_cases - case d: (1 d) - note d'=division_ofD[OF d] - have "(\k\d. norm (integral k f)) \ (\k\d. integral k g)" - apply (rule setsum_mono) - apply (rule integral_norm_bound_integral) - apply (drule_tac[!] d'(4)) - apply safe - apply (rule_tac[1-2] integrable_on_subcbox) - using assms - apply auto - done - also have "\ = integral (\d) g" - apply (rule integral_combine_division_bottomup[symmetric]) - apply (rule d) - apply safe - apply (drule d'(4)) - apply safe - apply (rule integrable_on_subcbox[OF assms(3)]) - apply auto - done - also have "\ \ integral UNIV g" - apply (rule integral_subset_le) - defer - apply (rule integrable_on_subdivision[OF d,of _ UNIV]) - prefer 4 - apply rule - apply (rule_tac y="norm (f x)" in order_trans) - using assms - apply auto - done - finally show ?case . - qed -qed - -lemma absolutely_integrable_absolutely_integrable_bound: - fixes f :: "'n::euclidean_space \ 'm::euclidean_space" - and g :: "'n::euclidean_space \ 'p::euclidean_space" - assumes "\x\s. norm (f x) \ norm (g x)" - and "f integrable_on s" - and "g absolutely_integrable_on s" - shows "f absolutely_integrable_on s" - apply (rule absolutely_integrable_integrable_bound[of s f "\x. norm (g x)"]) - using assms - apply auto - done - -lemma absolutely_integrable_inf_real: - assumes "finite k" - and "k \ {}" - and "\i\k. (\x. (fs x i)::real) absolutely_integrable_on s" - shows "(\x. (Inf ((fs x) ` k))) absolutely_integrable_on s" - using assms -proof induct - case (insert a k) - let ?P = "(\x. - if fs x ` k = {} then fs x a - else min (fs x a) (Inf (fs x ` k))) absolutely_integrable_on s" - show ?case - unfolding image_insert - apply (subst Inf_insert_finite) - apply (rule finite_imageI[OF insert(1)]) - proof (cases "k = {}") - case True - then show ?P - apply (subst if_P) - defer - apply (rule insert(5)[rule_format]) - apply auto - done - next - case False - then show ?P - apply (subst if_not_P) - defer - apply (rule absolutely_integrable_min[where 'n=real, unfolded Basis_real_def, simplified]) - defer - apply(rule insert(3)[OF False]) - using insert(5) - apply auto - done - qed -next - case empty - then show ?case by auto -qed - -lemma absolutely_integrable_sup_real: - assumes "finite k" - and "k \ {}" - and "\i\k. (\x. (fs x i)::real) absolutely_integrable_on s" - shows "(\x. (Sup ((fs x) ` k))) absolutely_integrable_on s" - using assms -proof induct - case (insert a k) - let ?P = "(\x. - if fs x ` k = {} then fs x a - else max (fs x a) (Sup (fs x ` k))) absolutely_integrable_on s" - show ?case - unfolding image_insert - apply (subst Sup_insert_finite) - apply (rule finite_imageI[OF insert(1)]) - proof (cases "k = {}") - case True - then show ?P - apply (subst if_P) - defer - apply (rule insert(5)[rule_format]) - apply auto - done - next - case False - then show ?P - apply (subst if_not_P) - defer - apply (rule absolutely_integrable_max[where 'n=real, unfolded Basis_real_def, simplified]) - defer - apply (rule insert(3)[OF False]) - using insert(5) - apply auto - done - qed -qed auto - - -subsection \differentiation under the integral sign\ - -lemma tube_lemma: - assumes "compact K" - assumes "open W" - assumes "{x0} \ K \ W" - shows "\X0. x0 \ X0 \ open X0 \ X0 \ K \ W" -proof - - { - fix y assume "y \ K" - then have "(x0, y) \ W" using assms by auto - with \open W\ - have "\X0 Y. open X0 \ open Y \ x0 \ X0 \ y \ Y \ X0 \ Y \ W" - by (rule open_prod_elim) blast - } - then obtain X0 Y where - *: "\y \ K. open (X0 y) \ open (Y y) \ x0 \ X0 y \ y \ Y y \ X0 y \ Y y \ W" - by metis - from * have "\t\Y ` K. open t" "K \ \(Y ` K)" by auto - with \compact K\ obtain CC where CC: "CC \ Y ` K" "finite CC" "K \ \CC" - by (rule compactE) - then obtain c where c: "\C. C \ CC \ c C \ K \ C = Y (c C)" - by (force intro!: choice) - with * CC show ?thesis - by (force intro!: exI[where x="\C\CC. X0 (c C)"]) (* SLOW *) -qed - -lemma continuous_on_prod_compactE: - fixes fx::"'a::topological_space \ 'b::topological_space \ 'c::metric_space" - and e::real - assumes cont_fx: "continuous_on (U \ C) fx" - assumes "compact C" - assumes [intro]: "x0 \ U" - notes [continuous_intros] = continuous_on_compose2[OF cont_fx] - assumes "e > 0" - obtains X0 where "x0 \ X0" "open X0" - "\x\X0 \ U. \t \ C. dist (fx (x, t)) (fx (x0, t)) \ e" -proof - - define psi where "psi = (\(x, t). dist (fx (x, t)) (fx (x0, t)))" - define W0 where "W0 = {(x, t) \ U \ C. psi (x, t) < e}" - have W0_eq: "W0 = psi -` {.. U \ C" - by (auto simp: vimage_def W0_def) - have "open {.. C) psi" - by (auto intro!: continuous_intros simp: psi_def split_beta') - from this[unfolded continuous_on_open_invariant, rule_format, OF \open {..] - obtain W where W: "open W" "W \ U \ C = W0 \ U \ C" - unfolding W0_eq by blast - have "{x0} \ C \ W \ U \ C" - unfolding W - by (auto simp: W0_def psi_def \0 < e\) - then have "{x0} \ C \ W" by blast - from tube_lemma[OF \compact C\ \open W\ this] - obtain X0 where X0: "x0 \ X0" "open X0" "X0 \ C \ W" - by blast - - have "\x\X0 \ U. \t \ C. dist (fx (x, t)) (fx (x0, t)) \ e" - proof safe - fix x assume x: "x \ X0" "x \ U" - fix t assume t: "t \ C" - have "dist (fx (x, t)) (fx (x0, t)) = psi (x, t)" - by (auto simp: psi_def) - also - { - have "(x, t) \ X0 \ C" - using t x - by auto - also note \\ \ W\ - finally have "(x, t) \ W" . - with t x have "(x, t) \ W \ U \ C" - by blast - also note \W \ U \ C = W0 \ U \ C\ - finally have "psi (x, t) < e" - by (auto simp: W0_def) - } - finally show "dist (fx (x, t)) (fx (x0, t)) \ e" by simp - qed - from X0(1,2) this show ?thesis .. -qed - -lemma integral_continuous_on_param: - fixes f::"'a::topological_space \ 'b::euclidean_space \ 'c::banach" - assumes cont_fx: "continuous_on (U \ cbox a b) (\(x, t). f x t)" - shows "continuous_on U (\x. integral (cbox a b) (f x))" -proof cases - assume "content (cbox a b) \ 0" - then have ne: "cbox a b \ {}" by auto - - note [continuous_intros] = - continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, - unfolded split_beta fst_conv snd_conv] - - show ?thesis - unfolding continuous_on_def - proof (safe intro!: tendstoI) - fix e'::real and x - assume "e' > 0" - define e where "e = e' / (content (cbox a b) + 1)" - have "e > 0" using \e' > 0\ by (auto simp: e_def intro!: divide_pos_pos add_nonneg_pos) - assume "x \ U" - from continuous_on_prod_compactE[OF cont_fx compact_cbox \x \ U\ \0 < e\] - obtain X0 where X0: "x \ X0" "open X0" - and fx_bound: "\y t. y \ X0 \ U \ t \ cbox a b \ norm (f y t - f x t) \ e" - unfolding split_beta fst_conv snd_conv dist_norm - by metis - have "\\<^sub>F y in at x within U. y \ X0 \ U" - using X0(1) X0(2) eventually_at_topological by auto - then show "\\<^sub>F y in at x within U. dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) < e'" - proof eventually_elim - case (elim y) - have "dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) = - norm (integral (cbox a b) (\t. f y t - f x t))" - using elim \x \ U\ - unfolding dist_norm - by (subst integral_diff) - (auto intro!: integrable_continuous continuous_intros) - also have "\ \ e * content (cbox a b)" - using elim \x \ U\ - by (intro integrable_bound) - (auto intro!: fx_bound \x \ U \ less_imp_le[OF \0 < e\] - integrable_continuous continuous_intros) - also have "\ < e'" - using \0 < e'\ \e > 0\ - by (auto simp: e_def divide_simps) - finally show "dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) < e'" . - qed - qed -qed (auto intro!: continuous_on_const) - -lemma eventually_closed_segment: - fixes x0::"'a::real_normed_vector" - assumes "open X0" "x0 \ X0" - shows "\\<^sub>F x in at x0 within U. closed_segment x0 x \ X0" -proof - - from openE[OF assms] - obtain e where e: "0 < e" "ball x0 e \ X0" . - then have "\\<^sub>F x in at x0 within U. x \ ball x0 e" - by (auto simp: dist_commute eventually_at) - then show ?thesis - proof eventually_elim - case (elim x) - have "x0 \ ball x0 e" using \e > 0\ by simp - from convex_ball[unfolded convex_contains_segment, rule_format, OF this elim] - have "closed_segment x0 x \ ball x0 e" . - also note \\ \ X0\ - finally show ?case . - qed -qed - -lemma leibniz_rule: - fixes f::"'a::banach \ 'b::euclidean_space \ 'c::banach" - assumes fx: "\x t. x \ U \ t \ cbox a b \ - ((\x. f x t) has_derivative blinfun_apply (fx x t)) (at x within U)" - assumes integrable_f2: "\x. x \ U \ f x integrable_on cbox a b" - assumes cont_fx: "continuous_on (U \ (cbox a b)) (\(x, t). fx x t)" - assumes [intro]: "x0 \ U" - assumes "convex U" - shows - "((\x. integral (cbox a b) (f x)) has_derivative integral (cbox a b) (fx x0)) (at x0 within U)" - (is "(?F has_derivative ?dF) _") -proof cases - assume "content (cbox a b) \ 0" - then have ne: "cbox a b \ {}" by auto - note [continuous_intros] = - continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, - unfolded split_beta fst_conv snd_conv] - show ?thesis - proof (intro has_derivativeI bounded_linear_scaleR_left tendstoI, fold norm_conv_dist) - have cont_f1: "\t. t \ cbox a b \ continuous_on U (\x. f x t)" - by (auto simp: continuous_on_eq_continuous_within intro!: has_derivative_continuous fx) - note [continuous_intros] = continuous_on_compose2[OF cont_f1] - fix e'::real - assume "e' > 0" - define e where "e = e' / (content (cbox a b) + 1)" - have "e > 0" using \e' > 0\ by (auto simp: e_def intro!: divide_pos_pos add_nonneg_pos) - from continuous_on_prod_compactE[OF cont_fx compact_cbox \x0 \ U\ \e > 0\] - obtain X0 where X0: "x0 \ X0" "open X0" - and fx_bound: "\x t. x \ X0 \ U \ t \ cbox a b \ norm (fx x t - fx x0 t) \ e" - unfolding split_beta fst_conv snd_conv - by (metis dist_norm) - - note eventually_closed_segment[OF \open X0\ \x0 \ X0\, of U] - moreover - have "\\<^sub>F x in at x0 within U. x \ X0" - using \open X0\ \x0 \ X0\ eventually_at_topological by blast - moreover have "\\<^sub>F x in at x0 within U. x \ x0" - by (auto simp: eventually_at_filter) - moreover have "\\<^sub>F x in at x0 within U. x \ U" - by (auto simp: eventually_at_filter) - ultimately - show "\\<^sub>F x in at x0 within U. norm ((?F x - ?F x0 - ?dF (x - x0)) /\<^sub>R norm (x - x0)) < e'" - proof eventually_elim - case (elim x) - from elim have "0 < norm (x - x0)" by simp - have "closed_segment x0 x \ U" - by (rule \convex U\[unfolded convex_contains_segment, rule_format, OF \x0 \ U\ \x \ U\]) - from elim have [intro]: "x \ U" by auto - - have "?F x - ?F x0 - ?dF (x - x0) = - integral (cbox a b) (\y. f x y - f x0 y - fx x0 y (x - x0))" - (is "_ = ?id") - using \x \ x0\ - by (subst blinfun_apply_integral integral_diff, - auto intro!: integrable_diff integrable_f2 continuous_intros - intro: integrable_continuous)+ - also - { - fix t assume t: "t \ (cbox a b)" - have seg: "\t. t \ {0..1} \ x0 + t *\<^sub>R (x - x0) \ X0 \ U" - using \closed_segment x0 x \ U\ - \closed_segment x0 x \ X0\ - by (force simp: closed_segment_def algebra_simps) - from t have deriv: - "((\x. f x t) has_derivative (fx y t)) (at y within X0 \ U)" - if "y \ X0 \ U" for y - unfolding has_vector_derivative_def[symmetric] - using that \x \ X0\ - by (intro has_derivative_within_subset[OF fx]) auto - have "\x \ X0 \ U. onorm (blinfun_apply (fx x t) - (fx x0 t)) \ e" - using fx_bound t - by (auto simp add: norm_blinfun_def fun_diff_def blinfun.bilinear_simps[symmetric]) - from differentiable_bound_linearization[OF seg deriv this] X0 - have "norm (f x t - f x0 t - fx x0 t (x - x0)) \ e * norm (x - x0)" - by (auto simp add: ac_simps) - } - then have "norm ?id \ integral (cbox a b) (\_. e * norm (x - x0))" - by (intro integral_norm_bound_integral) - (auto intro!: continuous_intros integrable_diff integrable_f2 - intro: integrable_continuous) - also have "\ = content (cbox a b) * e * norm (x - x0)" - by simp - also have "\ < e' * norm (x - x0)" - using \e' > 0\ content_pos_le[of a b] - by (intro mult_strict_right_mono[OF _ \0 < norm (x - x0)\]) - (auto simp: divide_simps e_def) - finally have "norm (?F x - ?F x0 - ?dF (x - x0)) < e' * norm (x - x0)" . - then show ?case - by (auto simp: divide_simps) - qed - qed (rule blinfun.bounded_linear_right) -qed (auto intro!: derivative_eq_intros simp: blinfun.bilinear_simps) - -lemma - has_vector_derivative_eq_has_derivative_blinfun: - "(f has_vector_derivative f') (at x within U) \ - (f has_derivative blinfun_scaleR_left f') (at x within U)" - by (simp add: has_vector_derivative_def) - -lemma leibniz_rule_vector_derivative: - fixes f::"real \ 'b::euclidean_space \ 'c::banach" - assumes fx: "\x t. x \ U \ t \ cbox a b \ - ((\x. f x t) has_vector_derivative (fx x t)) (at x within U)" - assumes integrable_f2: "\x. x \ U \ (f x) integrable_on cbox a b" - assumes cont_fx: "continuous_on (U \ cbox a b) (\(x, t). fx x t)" - assumes U: "x0 \ U" "convex U" - shows "((\x. integral (cbox a b) (f x)) has_vector_derivative integral (cbox a b) (fx x0)) - (at x0 within U)" -proof - - note [continuous_intros] = - continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, - unfolded split_beta fst_conv snd_conv] - have *: "blinfun_scaleR_left (integral (cbox a b) (fx x0)) = - integral (cbox a b) (\t. blinfun_scaleR_left (fx x0 t))" - by (subst integral_linear[symmetric]) - (auto simp: has_vector_derivative_def o_def - intro!: integrable_continuous U continuous_intros bounded_linear_intros) - show ?thesis - unfolding has_vector_derivative_eq_has_derivative_blinfun - apply (rule has_derivative_eq_rhs) - apply (rule leibniz_rule[OF _ integrable_f2 _ U, where fx="\x t. blinfun_scaleR_left (fx x t)"]) - using fx cont_fx - apply (auto simp: has_vector_derivative_def * split_beta intro!: continuous_intros) - done -qed - -lemma - has_field_derivative_eq_has_derivative_blinfun: - "(f has_field_derivative f') (at x within U) \ (f has_derivative blinfun_mult_right f') (at x within U)" - by (simp add: has_field_derivative_def) - -lemma leibniz_rule_field_derivative: - fixes f::"'a::{real_normed_field, banach} \ 'b::euclidean_space \ 'a" - assumes fx: "\x t. x \ U \ t \ cbox a b \ ((\x. f x t) has_field_derivative fx x t) (at x within U)" - assumes integrable_f2: "\x. x \ U \ (f x) integrable_on cbox a b" - assumes cont_fx: "continuous_on (U \ (cbox a b)) (\(x, t). fx x t)" - assumes U: "x0 \ U" "convex U" - shows "((\x. integral (cbox a b) (f x)) has_field_derivative integral (cbox a b) (fx x0)) (at x0 within U)" -proof - - note [continuous_intros] = - continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, - unfolded split_beta fst_conv snd_conv] - have *: "blinfun_mult_right (integral (cbox a b) (fx x0)) = - integral (cbox a b) (\t. blinfun_mult_right (fx x0 t))" - by (subst integral_linear[symmetric]) - (auto simp: has_vector_derivative_def o_def - intro!: integrable_continuous U continuous_intros bounded_linear_intros) - show ?thesis - unfolding has_field_derivative_eq_has_derivative_blinfun - apply (rule has_derivative_eq_rhs) - apply (rule leibniz_rule[OF _ integrable_f2 _ U, where fx="\x t. blinfun_mult_right (fx x t)"]) - using fx cont_fx - apply (auto simp: has_field_derivative_def * split_beta intro!: continuous_intros) - done -qed - - -subsection \Exchange uniform limit and integral\ - -lemma - uniform_limit_integral: - fixes f::"'a \ 'b::euclidean_space \ 'c::banach" - assumes u: "uniform_limit (cbox a b) f g F" - assumes c: "\n. continuous_on (cbox a b) (f n)" - assumes [simp]: "F \ bot" - obtains I J where - "\n. (f n has_integral I n) (cbox a b)" - "(g has_integral J) (cbox a b)" - "(I \ J) F" -proof - - have fi[simp]: "f n integrable_on (cbox a b)" for n - by (auto intro!: integrable_continuous assms) - then obtain I where I: "\n. (f n has_integral I n) (cbox a b)" - by atomize_elim (auto simp: integrable_on_def intro!: choice) - - moreover - - have gi[simp]: "g integrable_on (cbox a b)" - by (auto intro!: integrable_continuous uniform_limit_theorem[OF _ u] eventuallyI c) - then obtain J where J: "(g has_integral J) (cbox a b)" - by blast - - moreover - - have "(I \ J) F" - proof cases - assume "content (cbox a b) = 0" - hence "I = (\_. 0)" "J = 0" - by (auto intro!: has_integral_unique I J) - thus ?thesis by simp - next - assume content_nonzero: "content (cbox a b) \ 0" - show ?thesis - proof (rule tendstoI) - fix e::real - assume "e > 0" - define e' where "e' = e / 2" - with \e > 0\ have "e' > 0" by simp - then have "\\<^sub>F n in F. \x\cbox a b. norm (f n x - g x) < e' / content (cbox a b)" - using u content_nonzero content_pos_le[of a b] - by (auto simp: uniform_limit_iff dist_norm) - then show "\\<^sub>F n in F. dist (I n) J < e" - proof eventually_elim - case (elim n) - have "I n = integral (cbox a b) (f n)" - "J = integral (cbox a b) g" - using I[of n] J by (simp_all add: integral_unique) - then have "dist (I n) J = norm (integral (cbox a b) (\x. f n x - g x))" - by (simp add: integral_diff dist_norm) - also have "\ \ integral (cbox a b) (\x. (e' / content (cbox a b)))" - using elim - by (intro integral_norm_bound_integral) - (auto intro!: integrable_diff absolutely_integrable_onI) - also have "\ < e" - using \0 < e\ - by (simp add: e'_def) - finally show ?case . - qed - qed - qed - ultimately show ?thesis .. -qed - - -subsection \Dominated convergence\ - -context -begin - -private lemma dominated_convergence_real: - fixes f :: "nat \ 'n::euclidean_space \ real" - assumes "\k. (f k) integrable_on s" "h integrable_on s" - and "\k. \x \ s. norm (f k x) \ h x" - and "\x \ s. ((\k. f k x) \ g x) sequentially" - shows "g integrable_on s \ ((\k. integral s (f k)) \ integral s g) sequentially" -proof - have bdd_below[simp]: "\x P. x \ s \ bdd_below {f n x |n. P n}" - proof (safe intro!: bdd_belowI) - fix n x show "x \ s \ - h x \ f n x" - using assms(3)[rule_format, of x n] by simp - qed - have bdd_above[simp]: "\x P. x \ s \ bdd_above {f n x |n. P n}" - proof (safe intro!: bdd_aboveI) - fix n x show "x \ s \ f n x \ h x" - using assms(3)[rule_format, of x n] by simp - qed - - have "\m. (\x. Inf {f j x |j. m \ j}) integrable_on s \ - ((\k. integral s (\x. Inf {f j x |j. j \ {m..m + k}})) \ - integral s (\x. Inf {f j x |j. m \ j}))sequentially" - proof (rule monotone_convergence_decreasing, safe) - fix m :: nat - show "bounded {integral s (\x. Inf {f j x |j. j \ {m..m + k}}) |k. True}" - unfolding bounded_iff - apply (rule_tac x="integral s h" in exI) - proof safe - fix k :: nat - show "norm (integral s (\x. Inf {f j x |j. j \ {m..m + k}})) \ integral s h" - apply (rule integral_norm_bound_integral) - unfolding simple_image - apply (rule absolutely_integrable_onD) - apply (rule absolutely_integrable_inf_real) - prefer 5 - unfolding real_norm_def - apply rule - apply (rule cInf_abs_ge) - prefer 5 - apply rule - apply (rule_tac g = h in absolutely_integrable_integrable_bound) - using assms - unfolding real_norm_def - apply auto - done - qed - fix k :: nat - show "(\x. Inf {f j x |j. j \ {m..m + k}}) integrable_on s" - unfolding simple_image - apply (rule absolutely_integrable_onD) - apply (rule absolutely_integrable_inf_real) - prefer 3 - using absolutely_integrable_integrable_bound[OF assms(3,1,2)] - apply auto - done - fix x - assume x: "x \ s" - show "Inf {f j x |j. j \ {m..m + Suc k}} \ Inf {f j x |j. j \ {m..m + k}}" - by (rule cInf_superset_mono) auto - let ?S = "{f j x| j. m \ j}" - show "((\k. Inf {f j x |j. j \ {m..m + k}}) \ Inf ?S) sequentially" - proof (rule LIMSEQ_I, goal_cases) - case r: (1 r) - - have "\y\?S. y < Inf ?S + r" - by (subst cInf_less_iff[symmetric]) (auto simp: \x\s\ r) - then obtain N where N: "f N x < Inf ?S + r" "m \ N" - by blast - - show ?case - apply (rule_tac x=N in exI) - apply safe - proof goal_cases - case prems: (1 n) - have *: "\y ix. y < Inf ?S + r \ Inf ?S \ ix \ ix \ y \ \ix - Inf ?S\ < r" - by arith - show ?case - unfolding real_norm_def - apply (rule *[rule_format, OF N(1)]) - apply (rule cInf_superset_mono, auto simp: \x\s\) [] - apply (rule cInf_lower) - using N prems - apply auto [] - apply simp - done - qed - qed - qed - note dec1 = conjunctD2[OF this] - - have "\m. (\x. Sup {f j x |j. m \ j}) integrable_on s \ - ((\k. integral s (\x. Sup {f j x |j. j \ {m..m + k}})) \ - integral s (\x. Sup {f j x |j. m \ j})) sequentially" - proof (rule monotone_convergence_increasing,safe) - fix m :: nat - show "bounded {integral s (\x. Sup {f j x |j. j \ {m..m + k}}) |k. True}" - unfolding bounded_iff - apply (rule_tac x="integral s h" in exI) - proof safe - fix k :: nat - show "norm (integral s (\x. Sup {f j x |j. j \ {m..m + k}})) \ integral s h" - apply (rule integral_norm_bound_integral) unfolding simple_image - apply (rule absolutely_integrable_onD) - apply(rule absolutely_integrable_sup_real) - prefer 5 unfolding real_norm_def - apply rule - apply (rule cSup_abs_le) - using assms - apply (force simp add: ) - prefer 4 - apply rule - apply (rule_tac g=h in absolutely_integrable_integrable_bound) - using assms - unfolding real_norm_def - apply auto - done - qed - fix k :: nat - show "(\x. Sup {f j x |j. j \ {m..m + k}}) integrable_on s" - unfolding simple_image - apply (rule absolutely_integrable_onD) - apply (rule absolutely_integrable_sup_real) - prefer 3 - using absolutely_integrable_integrable_bound[OF assms(3,1,2)] - apply auto - done - fix x - assume x: "x\s" - show "Sup {f j x |j. j \ {m..m + Suc k}} \ Sup {f j x |j. j \ {m..m + k}}" - by (rule cSup_subset_mono) auto - let ?S = "{f j x| j. m \ j}" - show "((\k. Sup {f j x |j. j \ {m..m + k}}) \ Sup ?S) sequentially" - proof (rule LIMSEQ_I, goal_cases) - case r: (1 r) - have "\y\?S. Sup ?S - r < y" - by (subst less_cSup_iff[symmetric]) (auto simp: r \x\s\) - then obtain N where N: "Sup ?S - r < f N x" "m \ N" - by blast - - show ?case - apply (rule_tac x=N in exI) - apply safe - proof goal_cases - case prems: (1 n) - have *: "\y ix. Sup ?S - r < y \ ix \ Sup ?S \ y \ ix \ \ix - Sup ?S\ < r" - by arith - show ?case - apply simp - apply (rule *[rule_format, OF N(1)]) - apply (rule cSup_subset_mono, auto simp: \x\s\) [] - apply (subst cSup_upper) - using N prems - apply auto - done - qed - qed - qed - note inc1 = conjunctD2[OF this] - - have "g integrable_on s \ - ((\k. integral s (\x. Inf {f j x |j. k \ j})) \ integral s g) sequentially" - apply (rule monotone_convergence_increasing,safe) - apply fact - proof - - show "bounded {integral s (\x. Inf {f j x |j. k \ j}) |k. True}" - unfolding bounded_iff apply(rule_tac x="integral s h" in exI) - proof safe - fix k :: nat - show "norm (integral s (\x. Inf {f j x |j. k \ j})) \ integral s h" - apply (rule integral_norm_bound_integral) - apply fact+ - unfolding real_norm_def - apply rule - apply (rule cInf_abs_ge) - using assms(3) - apply auto - done - qed - fix k :: nat and x - assume x: "x \ s" - - have *: "\x y::real. x \ - y \ - x \ y" by auto - show "Inf {f j x |j. k \ j} \ Inf {f j x |j. Suc k \ j}" - by (intro cInf_superset_mono) (auto simp: \x\s\) - - show "(\k::nat. Inf {f j x |j. k \ j}) \ g x" - proof (rule LIMSEQ_I, goal_cases) - case r: (1 r) - then have "0 - ((\k. integral s (\x. Sup {f j x |j. k \ j})) \ integral s g) sequentially" - apply (rule monotone_convergence_decreasing,safe) - apply fact - proof - - show "bounded {integral s (\x. Sup {f j x |j. k \ j}) |k. True}" - unfolding bounded_iff - apply (rule_tac x="integral s h" in exI) - proof safe - fix k :: nat - show "norm (integral s (\x. Sup {f j x |j. k \ j})) \ integral s h" - apply (rule integral_norm_bound_integral) - apply fact+ - unfolding real_norm_def - apply rule - apply (rule cSup_abs_le) - using assms(3) - apply auto - done - qed - fix k :: nat - fix x - assume x: "x \ s" - - show "Sup {f j x |j. k \ j} \ Sup {f j x |j. Suc k \ j}" - by (rule cSup_subset_mono) (auto simp: \x\s\) - show "((\k. Sup {f j x |j. k \ j}) \ g x) sequentially" - proof (rule LIMSEQ_I, goal_cases) - case r: (1 r) - then have "0k. integral s (f k)) \ integral s g) sequentially" - proof (rule LIMSEQ_I, goal_cases) - case r: (1 r) - from LIMSEQ_D [OF inc2(2) r] guess N1 .. note N1=this[unfolded real_norm_def] - from LIMSEQ_D [OF dec2(2) r] guess N2 .. note N2=this[unfolded real_norm_def] - show ?case - proof (rule_tac x="N1+N2" in exI, safe) - fix n - assume n: "n \ N1 + N2" - have *: "\i0 i i1 g. \i0 - g\ < r \ \i1 - g\ < r \ i0 \ i \ i \ i1 \ \i - g\ < r" - by arith - show "norm (integral s (f n) - integral s g) < r" - unfolding real_norm_def - proof (rule *[rule_format,OF N1[rule_format] N2[rule_format], of n n]) - show "integral s (\x. Inf {f j x |j. n \ j}) \ integral s (f n)" - by (rule integral_le[OF dec1(1) assms(1)]) (auto intro!: cInf_lower) - show "integral s (f n) \ integral s (\x. Sup {f j x |j. n \ j})" - by (rule integral_le[OF assms(1) inc1(1)]) (auto intro!: cSup_upper) - qed (insert n, auto) - qed - qed -qed - -lemma dominated_convergence: - fixes f :: "nat \ 'n::euclidean_space \ 'm::euclidean_space" - assumes f: "\k. (f k) integrable_on s" and h: "h integrable_on s" - and le: "\k. \x \ s. norm (f k x) \ h x" - and conv: "\x \ s. ((\k. f k x) \ g x) sequentially" - shows "g integrable_on s" - and "((\k. integral s (f k)) \ integral s g) sequentially" -proof - - { - fix b :: 'm assume b: "b \ Basis" - have A: "(\x. g x \ b) integrable_on s \ - (\k. integral s (\x. f k x \ b)) \ integral s (\x. g x \ b)" - proof (rule dominated_convergence_real) - fix k :: nat - from f show "(\x. f k x \ b) integrable_on s" by (rule integrable_component) - next - from h show "h integrable_on s" . - next - fix k :: nat - show "\x\s. norm (f k x \ b) \ h x" - proof - fix x assume x: "x \ s" - have "norm (f k x \ b) \ norm (f k x)" by (simp add: Basis_le_norm b) - also have "\ \ h x" using le[of k] x by simp - finally show "norm (f k x \ b) \ h x" . - qed - next - from conv show "\x\s. (\k. f k x \ b) \ g x \ b" - by (auto intro!: tendsto_inner) - qed - have B: "integral s ((\x. x *\<^sub>R b) \ (\x. f k x \ b)) = integral s (\x. f k x \ b) *\<^sub>R b" - for k by (rule integral_linear) - (simp_all add: f integrable_component bounded_linear_scaleR_left) - have C: "integral s ((\x. x *\<^sub>R b) \ (\x. g x \ b)) = integral s (\x. g x \ b) *\<^sub>R b" - using A by (intro integral_linear) - (simp_all add: integrable_component bounded_linear_scaleR_left) - note A B C - } note A = this - - show "g integrable_on s" by (rule integrable_componentwise) (insert A, blast) - with A f show "((\k. integral s (f k)) \ integral s g) sequentially" - by (subst (1 2) integral_componentwise) - (auto intro!: tendsto_setsum tendsto_scaleR simp: o_def) -qed - -lemma has_integral_dominated_convergence: - fixes f :: "nat \ 'n::euclidean_space \ 'm::euclidean_space" - assumes "\k. (f k has_integral y k) s" "h integrable_on s" - "\k. \x\s. norm (f k x) \ h x" "\x\s. (\k. f k x) \ g x" - and x: "y \ x" - shows "(g has_integral x) s" -proof - - have int_f: "\k. (f k) integrable_on s" - using assms by (auto simp: integrable_on_def) - have "(g has_integral (integral s g)) s" - by (intro integrable_integral dominated_convergence[OF int_f assms(2)]) fact+ - moreover have "integral s g = x" - proof (rule LIMSEQ_unique) - show "(\i. integral s (f i)) \ x" - using integral_unique[OF assms(1)] x by simp - show "(\i. integral s (f i)) \ integral s g" - by (intro dominated_convergence[OF int_f assms(2)]) fact+ - qed - ultimately show ?thesis - by simp -qed - -end - - -subsection \Integration by parts\ - -lemma integration_by_parts_interior_strong: - assumes bilinear: "bounded_bilinear (prod :: _ \ _ \ 'b :: banach)" - assumes s: "finite s" and le: "a \ b" - assumes cont [continuous_intros]: "continuous_on {a..b} f" "continuous_on {a..b} g" - assumes deriv: "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" - "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" - assumes int: "((\x. prod (f x) (g' x)) has_integral - (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}" - shows "((\x. prod (f' x) (g x)) has_integral y) {a..b}" -proof - - interpret bounded_bilinear prod by fact - have "((\x. prod (f x) (g' x) + prod (f' x) (g x)) has_integral - (prod (f b) (g b) - prod (f a) (g a))) {a..b}" - using deriv by (intro fundamental_theorem_of_calculus_interior_strong[OF s le]) - (auto intro!: continuous_intros continuous_on has_vector_derivative) - from has_integral_sub[OF this int] show ?thesis by (simp add: algebra_simps) -qed - -lemma integration_by_parts_interior: - assumes "bounded_bilinear (prod :: _ \ _ \ 'b :: banach)" "a \ b" - "continuous_on {a..b} f" "continuous_on {a..b} g" - assumes "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" - "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" - assumes "((\x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}" - shows "((\x. prod (f' x) (g x)) has_integral y) {a..b}" - by (rule integration_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all) - -lemma integration_by_parts: - assumes "bounded_bilinear (prod :: _ \ _ \ 'b :: banach)" "a \ b" - "continuous_on {a..b} f" "continuous_on {a..b} g" - assumes "\x. x\{a..b} \ (f has_vector_derivative f' x) (at x)" - "\x. x\{a..b} \ (g has_vector_derivative g' x) (at x)" - assumes "((\x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}" - shows "((\x. prod (f' x) (g x)) has_integral y) {a..b}" - by (rule integration_by_parts_interior[of _ _ _ f g f' g']) (insert assms, simp_all) - -lemma integrable_by_parts_interior_strong: - assumes bilinear: "bounded_bilinear (prod :: _ \ _ \ 'b :: banach)" - assumes s: "finite s" and le: "a \ b" - assumes cont [continuous_intros]: "continuous_on {a..b} f" "continuous_on {a..b} g" - assumes deriv: "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" - "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" - assumes int: "(\x. prod (f x) (g' x)) integrable_on {a..b}" - shows "(\x. prod (f' x) (g x)) integrable_on {a..b}" -proof - - from int obtain I where "((\x. prod (f x) (g' x)) has_integral I) {a..b}" - unfolding integrable_on_def by blast - hence "((\x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - - (prod (f b) (g b) - prod (f a) (g a) - I))) {a..b}" by simp - from integration_by_parts_interior_strong[OF assms(1-7) this] - show ?thesis unfolding integrable_on_def by blast -qed - -lemma integrable_by_parts_interior: - assumes "bounded_bilinear (prod :: _ \ _ \ 'b :: banach)" "a \ b" - "continuous_on {a..b} f" "continuous_on {a..b} g" - assumes "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" - "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" - assumes "(\x. prod (f x) (g' x)) integrable_on {a..b}" - shows "(\x. prod (f' x) (g x)) integrable_on {a..b}" - by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all) - -lemma integrable_by_parts: - assumes "bounded_bilinear (prod :: _ \ _ \ 'b :: banach)" "a \ b" - "continuous_on {a..b} f" "continuous_on {a..b} g" - assumes "\x. x\{a..b} \ (f has_vector_derivative f' x) (at x)" - "\x. x\{a..b} \ (g has_vector_derivative g' x) (at x)" - assumes "(\x. prod (f x) (g' x)) integrable_on {a..b}" - shows "(\x. prod (f' x) (g x)) integrable_on {a..b}" - by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all) - - -subsection \Integration by substitution\ - - -lemma has_integral_substitution_strong: - fixes f :: "real \ 'a::euclidean_space" and g :: "real \ real" - assumes s: "finite s" and le: "a \ b" "c \ d" "g a \ g b" - and subset: "g ` {a..b} \ {c..d}" - and f [continuous_intros]: "continuous_on {c..d} f" - and g [continuous_intros]: "continuous_on {a..b} g" - and deriv [derivative_intros]: - "\x. x \ {a..b} - s \ (g has_field_derivative g' x) (at x within {a..b})" - shows "((\x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f)) {a..b}" -proof - - let ?F = "\x. integral {c..g x} f" - have cont_int: "continuous_on {a..b} ?F" - by (rule continuous_on_compose2[OF _ g subset] indefinite_integral_continuous - f integrable_continuous_real)+ - have deriv: "(((\x. integral {c..x} f) \ g) has_vector_derivative g' x *\<^sub>R f (g x)) - (at x within {a..b})" if "x \ {a..b} - s" for x - apply (rule has_vector_derivative_eq_rhs) - apply (rule vector_diff_chain_within) - apply (subst has_field_derivative_iff_has_vector_derivative [symmetric]) - apply (rule deriv that)+ - apply (rule has_vector_derivative_within_subset) - apply (rule integral_has_vector_derivative f)+ - using that le subset - apply blast+ - done - have deriv: "(?F has_vector_derivative g' x *\<^sub>R f (g x)) - (at x)" if "x \ {a..b} - (s \ {a,b})" for x - using deriv[of x] that by (simp add: at_within_closed_interval o_def) - - - have "((\x. g' x *\<^sub>R f (g x)) has_integral (?F b - ?F a)) {a..b}" - using le cont_int s deriv cont_int - by (intro fundamental_theorem_of_calculus_interior_strong[of "s \ {a,b}"]) simp_all - also from subset have "g x \ {c..d}" if "x \ {a..b}" for x using that by blast - from this[of a] this[of b] le have "c \ g a" "g b \ d" by auto - hence "integral {c..g a} f + integral {g a..g b} f = integral {c..g b} f" - by (intro integral_combine integrable_continuous_real continuous_on_subset[OF f] le) simp_all - hence "integral {c..g b} f - integral {c..g a} f = integral {g a..g b} f" - by (simp add: algebra_simps) - finally show ?thesis . -qed - -lemma has_integral_substitution: - fixes f :: "real \ 'a::euclidean_space" and g :: "real \ real" - assumes "a \ b" "c \ d" "g a \ g b" "g ` {a..b} \ {c..d}" - and "continuous_on {c..d} f" - and "\x. x \ {a..b} \ (g has_field_derivative g' x) (at x within {a..b})" - shows "((\x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f)) {a..b}" - by (intro has_integral_substitution_strong[of "{}" a b c d] assms) - (auto intro: DERIV_continuous_on assms) - - -subsection \Compute a double integral using iterated integrals and switching the order of integration\ - -lemma setcomp_dot1: "{z. P (z \ (i,0))} = {(x,y). P(x \ i)}" - by auto - -lemma setcomp_dot2: "{z. P (z \ (0,i))} = {(x,y). P(y \ i)}" - by auto - -lemma Sigma_Int_Paircomp1: "(Sigma A B) \ {(x, y). P x} = Sigma (A \ {x. P x}) B" - by blast - -lemma Sigma_Int_Paircomp2: "(Sigma A B) \ {(x, y). P y} = Sigma A (\z. B z \ {y. P y})" - by blast - -lemma continuous_on_imp_integrable_on_Pair1: - fixes f :: "_ \ 'b::banach" - assumes con: "continuous_on (cbox (a,c) (b,d)) f" and x: "x \ cbox a b" - shows "(\y. f (x, y)) integrable_on (cbox c d)" -proof - - have "f \ (\y. (x, y)) integrable_on (cbox c d)" - apply (rule integrable_continuous) - apply (rule continuous_on_compose [OF _ continuous_on_subset [OF con]]) - using x - apply (auto intro: continuous_on_Pair continuous_on_const continuous_on_id continuous_on_subset con) - done - then show ?thesis - by (simp add: o_def) -qed - -lemma integral_integrable_2dim: - fixes f :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::banach" - assumes "continuous_on (cbox (a,c) (b,d)) f" - shows "(\x. integral (cbox c d) (\y. f (x,y))) integrable_on cbox a b" -proof (cases "content(cbox c d) = 0") -case True - then show ?thesis - by (simp add: True integrable_const) -next - case False - have uc: "uniformly_continuous_on (cbox (a,c) (b,d)) f" - by (simp add: assms compact_cbox compact_uniformly_continuous) - { fix x::'a and e::real - assume x: "x \ cbox a b" and e: "0 < e" - then have e2_gt: "0 < e / 2 / content (cbox c d)" and e2_less: "e / 2 / content (cbox c d) * content (cbox c d) < e" - by (auto simp: False content_lt_nz e) - then obtain dd - where dd: "\x x'. \x\cbox (a, c) (b, d); x'\cbox (a, c) (b, d); norm (x' - x) < dd\ - \ norm (f x' - f x) \ e / (2 * content (cbox c d))" "dd>0" - using uc [unfolded uniformly_continuous_on_def, THEN spec, of "e / (2 * content (cbox c d))"] - by (auto simp: dist_norm intro: less_imp_le) - have "\delta>0. \x'\cbox a b. norm (x' - x) < delta \ norm (integral (cbox c d) (\u. f (x', u) - f (x, u))) < e" - apply (rule_tac x=dd in exI) - using dd e2_gt assms x - apply clarify - apply (rule le_less_trans [OF _ e2_less]) - apply (rule integrable_bound) - apply (auto intro: integrable_diff continuous_on_imp_integrable_on_Pair1) - done - } note * = this - show ?thesis - apply (rule integrable_continuous) - apply (simp add: * continuous_on_iff dist_norm integral_diff [symmetric] continuous_on_imp_integrable_on_Pair1 [OF assms]) - done -qed - -lemma norm_diff2: "\y = y1 + y2; x = x1 + x2; e = e1 + e2; norm(y1 - x1) \ e1; norm(y2 - x2) \ e2\ - \ norm(y - x) \ e" -using norm_triangle_mono [of "y1 - x1" "e1" "y2 - x2" "e2"] - by (simp add: add_diff_add) - -lemma integral_split: - fixes f :: "'a::euclidean_space \ 'b::{real_normed_vector,complete_space}" - assumes f: "f integrable_on (cbox a b)" - and k: "k \ Basis" - shows "integral (cbox a b) f = - integral (cbox a b \ {x. x\k \ c}) f + - integral (cbox a b \ {x. x\k \ c}) f" - apply (rule integral_unique [OF has_integral_split [where c=c]]) - using k f - apply (auto simp: has_integral_integral [symmetric]) - done - -lemma integral_swap_operative: - fixes f :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::banach" - assumes f: "continuous_on s f" and e: "0 < e" - shows "comm_monoid.operative (op \) True - (\k. \a b c d. - cbox (a,c) (b,d) \ k \ cbox (a,c) (b,d) \ s - \ norm(integral (cbox (a,c) (b,d)) f - - integral (cbox a b) (\x. integral (cbox c d) (\y. f((x,y))))) - \ e * content (cbox (a,c) (b,d)))" -proof (auto simp: comm_monoid.operative_def[OF comm_monoid_and]) - fix a::'a and c::'b and b::'a and d::'b and u::'a and v::'a and w::'b and z::'b - assume c0: "content (cbox (a, c) (b, d)) = 0" - and cb1: "cbox (u, w) (v, z) \ cbox (a, c) (b, d)" - and cb2: "cbox (u, w) (v, z) \ s" - have c0': "content (cbox (u, w) (v, z)) = 0" - by (fact content_0_subset [OF c0 cb1]) - show "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y)))) - \ e * content (cbox (u,w) (v,z))" - using content_cbox_pair_eq0_D [OF c0'] - by (force simp add: c0') -next - fix a::'a and c::'b and b::'a and d::'b - and M::real and i::'a and j::'b - and u::'a and v::'a and w::'b and z::'b - assume ij: "(i,j) \ Basis" - and n1: "\a' b' c' d'. - cbox (a',c') (b',d') \ cbox (a,c) (b,d) \ - cbox (a',c') (b',d') \ {x. x \ (i,j) \ M} \ cbox (a',c') (b',d') \ s \ - norm (integral (cbox (a',c') (b',d')) f - integral (cbox a' b') (\x. integral (cbox c' d') (\y. f (x,y)))) - \ e * content (cbox (a',c') (b',d'))" - and n2: "\a' b' c' d'. - cbox (a',c') (b',d') \ cbox (a,c) (b,d) \ - cbox (a',c') (b',d') \ {x. M \ x \ (i,j)} \ cbox (a',c') (b',d') \ s \ - norm (integral (cbox (a',c') (b',d')) f - integral (cbox a' b') (\x. integral (cbox c' d') (\y. f (x,y)))) - \ e * content (cbox (a',c') (b',d'))" - and subs: "cbox (u,w) (v,z) \ cbox (a,c) (b,d)" "cbox (u,w) (v,z) \ s" - have fcont: "continuous_on (cbox (u, w) (v, z)) f" - using assms(1) continuous_on_subset subs(2) by blast - then have fint: "f integrable_on cbox (u, w) (v, z)" - by (metis integrable_continuous) - consider "i \ Basis" "j=0" | "j \ Basis" "i=0" using ij - by (auto simp: Euclidean_Space.Basis_prod_def) - then show "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x,y)))) - \ e * content (cbox (u,w) (v,z))" (is ?normle) - proof cases - case 1 - then have e: "e * content (cbox (u, w) (v, z)) = - e * (content (cbox u v \ {x. x \ i \ M}) * content (cbox w z)) + - e * (content (cbox u v \ {x. M \ x \ i}) * content (cbox w z))" - by (simp add: content_split [where c=M] content_Pair algebra_simps) - have *: "integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y))) = - integral (cbox u v \ {x. x \ i \ M}) (\x. integral (cbox w z) (\y. f (x, y))) + - integral (cbox u v \ {x. M \ x \ i}) (\x. integral (cbox w z) (\y. f (x, y)))" - using 1 f subs integral_integrable_2dim continuous_on_subset - by (blast intro: integral_split) - show ?normle - apply (rule norm_diff2 [OF integral_split [where c=M, OF fint ij] * e]) - using 1 subs - apply (simp_all add: cbox_Pair_eq setcomp_dot1 [of "\u. M\u"] setcomp_dot1 [of "\u. u\M"] Sigma_Int_Paircomp1) - apply (simp_all add: interval_split ij) - apply (simp_all add: cbox_Pair_eq [symmetric] content_Pair [symmetric]) - apply (force simp add: interval_split [symmetric] intro!: n1 [rule_format]) - apply (force simp add: interval_split [symmetric] intro!: n2 [rule_format]) - done - next - case 2 - then have e: "e * content (cbox (u, w) (v, z)) = - e * (content (cbox u v) * content (cbox w z \ {x. x \ j \ M})) + - e * (content (cbox u v) * content (cbox w z \ {x. M \ x \ j}))" - by (simp add: content_split [where c=M] content_Pair algebra_simps) - have "(\x. integral (cbox w z \ {x. x \ j \ M}) (\y. f (x, y))) integrable_on cbox u v" - "(\x. integral (cbox w z \ {x. M \ x \ j}) (\y. f (x, y))) integrable_on cbox u v" - using 2 subs - apply (simp_all add: interval_split) - apply (rule_tac [!] integral_integrable_2dim [OF continuous_on_subset [OF f]]) - apply (auto simp: cbox_Pair_eq interval_split [symmetric]) - done - with 2 have *: "integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y))) = - integral (cbox u v) (\x. integral (cbox w z \ {x. x \ j \ M}) (\y. f (x, y))) + - integral (cbox u v) (\x. integral (cbox w z \ {x. M \ x \ j}) (\y. f (x, y)))" - by (simp add: integral_add [symmetric] integral_split [symmetric] - continuous_on_imp_integrable_on_Pair1 [OF fcont] cong: integral_cong) - show ?normle - apply (rule norm_diff2 [OF integral_split [where c=M, OF fint ij] * e]) - using 2 subs - apply (simp_all add: cbox_Pair_eq setcomp_dot2 [of "\u. M\u"] setcomp_dot2 [of "\u. u\M"] Sigma_Int_Paircomp2) - apply (simp_all add: interval_split ij) - apply (simp_all add: cbox_Pair_eq [symmetric] content_Pair [symmetric]) - apply (force simp add: interval_split [symmetric] intro!: n1 [rule_format]) - apply (force simp add: interval_split [symmetric] intro!: n2 [rule_format]) - done - qed -qed - -lemma integral_Pair_const: - "integral (cbox (a,c) (b,d)) (\x. k) = - integral (cbox a b) (\x. integral (cbox c d) (\y. k))" - by (simp add: content_Pair) - -lemma norm_minus2: "norm (x1-x2, y1-y2) = norm (x2-x1, y2-y1)" - by (simp add: norm_minus_eqI) - -lemma integral_prod_continuous: - fixes f :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::banach" - assumes "continuous_on (cbox (a,c) (b,d)) f" (is "continuous_on ?CBOX f") - shows "integral (cbox (a,c) (b,d)) f = integral (cbox a b) (\x. integral (cbox c d) (\y. f(x,y)))" -proof (cases "content ?CBOX = 0") - case True - then show ?thesis - by (auto simp: content_Pair) -next - case False - then have cbp: "content ?CBOX > 0" - using content_lt_nz by blast - have "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x,y)))) = 0" - proof (rule dense_eq0_I, simp) - fix e::real assume "0 < e" - with cbp have e': "0 < e / content ?CBOX" - by simp - have f_int_acbd: "f integrable_on cbox (a,c) (b,d)" - by (rule integrable_continuous [OF assms]) - { fix p - assume p: "p division_of cbox (a,c) (b,d)" - note opd1 = comm_monoid_set.operative_division [OF comm_monoid_set_and integral_swap_operative [OF assms e'], THEN iffD1, - THEN spec, THEN spec, THEN spec, THEN spec, of p "(a,c)" "(b,d)" a c b d] - have "(\t u v w z. - \t \ p; cbox (u,w) (v,z) \ t; cbox (u,w) (v,z) \ cbox (a,c) (b,d)\ \ - norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x,y)))) - \ e * content (cbox (u,w) (v,z)) / content?CBOX) - \ - norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x,y)))) \ e" - using opd1 [OF p] False by (simp add: comm_monoid_set_F_and) - } note op_acbd = this - { fix k::real and p and u::'a and v w and z::'b and t1 t2 l - assume k: "0 < k" - and nf: "\x y u v. - \x \ cbox a b; y \ cbox c d; u \ cbox a b; v\cbox c d; norm (u-x, v-y) < k\ - \ norm (f(u,v) - f(x,y)) < e / (2 * (content ?CBOX))" - and p_acbd: "p tagged_division_of cbox (a,c) (b,d)" - and fine: "(\x. ball x k) fine p" "((t1,t2), l) \ p" - and uwvz_sub: "cbox (u,w) (v,z) \ l" "cbox (u,w) (v,z) \ cbox (a,c) (b,d)" - have t: "t1 \ cbox a b" "t2 \ cbox c d" - by (meson fine p_acbd cbox_Pair_iff tag_in_interval)+ - have ls: "l \ ball (t1,t2) k" - using fine by (simp add: fine_def Ball_def) - { fix x1 x2 - assume xuvwz: "x1 \ cbox u v" "x2 \ cbox w z" - then have x: "x1 \ cbox a b" "x2 \ cbox c d" - using uwvz_sub by auto - have "norm (x1 - t1, x2 - t2) < k" - using xuvwz ls uwvz_sub unfolding ball_def - by (force simp add: cbox_Pair_eq dist_norm norm_minus2) - then have "norm (f (x1,x2) - f (t1,t2)) \ e / content ?CBOX / 2" - using nf [OF t x] by simp - } note nf' = this - have f_int_uwvz: "f integrable_on cbox (u,w) (v,z)" - using f_int_acbd uwvz_sub integrable_on_subcbox by blast - have f_int_uv: "\x. x \ cbox u v \ (\y. f (x,y)) integrable_on cbox w z" - using assms continuous_on_subset uwvz_sub - by (blast intro: continuous_on_imp_integrable_on_Pair1) - have 1: "norm (integral (cbox (u,w) (v,z)) f - integral (cbox (u,w) (v,z)) (\x. f (t1,t2))) - \ e * content (cbox (u,w) (v,z)) / content ?CBOX / 2" - apply (simp only: integral_diff [symmetric] f_int_uwvz integrable_const) - apply (rule order_trans [OF integrable_bound [of "e / content ?CBOX / 2"]]) - using cbp e' nf' - apply (auto simp: integrable_diff f_int_uwvz integrable_const) - done - have int_integrable: "(\x. integral (cbox w z) (\y. f (x, y))) integrable_on cbox u v" - using assms integral_integrable_2dim continuous_on_subset uwvz_sub(2) by blast - have normint_wz: - "\x. x \ cbox u v \ - norm (integral (cbox w z) (\y. f (x, y)) - integral (cbox w z) (\y. f (t1, t2))) - \ e * content (cbox w z) / content (cbox (a, c) (b, d)) / 2" - apply (simp only: integral_diff [symmetric] f_int_uv integrable_const) - apply (rule order_trans [OF integrable_bound [of "e / content ?CBOX / 2"]]) - using cbp e' nf' - apply (auto simp: integrable_diff f_int_uv integrable_const) - done - have "norm (integral (cbox u v) - (\x. integral (cbox w z) (\y. f (x,y)) - integral (cbox w z) (\y. f (t1,t2)))) - \ e * content (cbox w z) / content ?CBOX / 2 * content (cbox u v)" - apply (rule integrable_bound [OF _ _ normint_wz]) - using cbp e' - apply (auto simp: divide_simps content_pos_le integrable_diff int_integrable integrable_const) - done - also have "... \ e * content (cbox (u,w) (v,z)) / content ?CBOX / 2" - by (simp add: content_Pair divide_simps) - finally have 2: "norm (integral (cbox u v) (\x. integral (cbox w z) (\y. f (x,y))) - - integral (cbox u v) (\x. integral (cbox w z) (\y. f (t1,t2)))) - \ e * content (cbox (u,w) (v,z)) / content ?CBOX / 2" - by (simp only: integral_diff [symmetric] int_integrable integrable_const) - have norm_xx: "\x' = y'; norm(x - x') \ e/2; norm(y - y') \ e/2\ \ norm(x - y) \ e" for x::'c and y x' y' e - using norm_triangle_mono [of "x-y'" "e/2" "y'-y" "e/2"] real_sum_of_halves - by (simp add: norm_minus_commute) - have "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x,y)))) - \ e * content (cbox (u,w) (v,z)) / content ?CBOX" - by (rule norm_xx [OF integral_Pair_const 1 2]) - } note * = this - show "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x,y)))) \ e" - using compact_uniformly_continuous [OF assms compact_cbox] - apply (simp add: uniformly_continuous_on_def dist_norm) - apply (drule_tac x="e / 2 / content?CBOX" in spec) - using cbp \0 < e\ - apply (auto simp: zero_less_mult_iff) - apply (rename_tac k) - apply (rule_tac e1=k in fine_division_exists [OF gauge_ball, where a = "(a,c)" and b = "(b,d)"]) - apply assumption - apply (rule op_acbd) - apply (erule division_of_tagged_division) - using * - apply auto - done - qed - then show ?thesis - by simp -qed - -lemma swap_continuous: - assumes "continuous_on (cbox (a,c) (b,d)) (\(x,y). f x y)" - shows "continuous_on (cbox (c,a) (d,b)) (\(x, y). f y x)" -proof - - have "(\(x, y). f y x) = (\(x, y). f x y) \ prod.swap" - by auto - then show ?thesis - apply (rule ssubst) - apply (rule continuous_on_compose) - apply (simp add: split_def) - apply (rule continuous_intros | simp add: assms)+ - done -qed - -lemma integral_swap_2dim: - fixes f :: "['a::euclidean_space, 'b::euclidean_space] \ 'c::banach" - assumes "continuous_on (cbox (a,c) (b,d)) (\(x,y). f x y)" - shows "integral (cbox (a, c) (b, d)) (\(x, y). f x y) = integral (cbox (c, a) (d, b)) (\(x, y). f y x)" -proof - - have "((\(x, y). f x y) has_integral integral (cbox (c, a) (d, b)) (\(x, y). f y x)) (prod.swap ` (cbox (c, a) (d, b)))" - apply (rule has_integral_twiddle [of 1 prod.swap prod.swap "\(x,y). f y x" "integral (cbox (c, a) (d, b)) (\(x, y). f y x)", simplified]) - apply (auto simp: isCont_swap content_Pair has_integral_integral [symmetric] integrable_continuous swap_continuous assms) - done - then show ?thesis - by force -qed - -theorem integral_swap_continuous: - fixes f :: "['a::euclidean_space, 'b::euclidean_space] \ 'c::banach" - assumes "continuous_on (cbox (a,c) (b,d)) (\(x,y). f x y)" - shows "integral (cbox a b) (\x. integral (cbox c d) (f x)) = - integral (cbox c d) (\y. integral (cbox a b) (\x. f x y))" -proof - - have "integral (cbox a b) (\x. integral (cbox c d) (f x)) = integral (cbox (a, c) (b, d)) (\(x, y). f x y)" - using integral_prod_continuous [OF assms] by auto - also have "... = integral (cbox (c, a) (d, b)) (\(x, y). f y x)" - by (rule integral_swap_2dim [OF assms]) - also have "... = integral (cbox c d) (\y. integral (cbox a b) (\x. f x y))" - using integral_prod_continuous [OF swap_continuous] assms - by auto - finally show ?thesis . -qed - - -subsection \Definite integrals for exponential and power function\ - -lemma has_integral_exp_minus_to_infinity: - assumes a: "a > 0" - shows "((\x::real. exp (-a*x)) has_integral exp (-a*c)/a) {c..}" -proof - - define f where "f = (\k x. if x \ {c..real k} then exp (-a*x) else 0)" - - { - fix k :: nat assume k: "of_nat k \ c" - from k a - have "((\x. exp (-a*x)) has_integral (-exp (-a*real k)/a - (-exp (-a*c)/a))) {c..real k}" - by (intro fundamental_theorem_of_calculus) - (auto intro!: derivative_eq_intros - simp: has_field_derivative_iff_has_vector_derivative [symmetric]) - hence "(f k has_integral (exp (-a*c)/a - exp (-a*real k)/a)) {c..}" unfolding f_def - by (subst has_integral_restrict) simp_all - } note has_integral_f = this - - have [simp]: "f k = (\_. 0)" if "of_nat k < c" for k using that by (auto simp: fun_eq_iff f_def) - have integral_f: "integral {c..} (f k) = - (if real k \ c then exp (-a*c)/a - exp (-a*real k)/a else 0)" - for k using integral_unique[OF has_integral_f[of k]] by simp - - have A: "(\x. exp (-a*x)) integrable_on {c..} \ - (\k. integral {c..} (f k)) \ integral {c..} (\x. exp (-a*x))" - proof (intro monotone_convergence_increasing allI ballI) - fix k ::nat - have "(\x. exp (-a*x)) integrable_on {c..of_real k}" (is ?P) - unfolding f_def by (auto intro!: continuous_intros integrable_continuous_real) - hence int: "(f k) integrable_on {c..of_real k}" - by (rule integrable_eq[rotated]) (simp add: f_def) - show "f k integrable_on {c..}" - by (rule integrable_on_superset[OF _ _ int]) (auto simp: f_def) - next - fix x assume x: "x \ {c..}" - have "sequentially \ principal {nat \x\..}" unfolding at_top_def by (simp add: Inf_lower) - also have "{nat \x\..} \ {k. x \ real k}" by auto - also have "inf (principal \) (principal {k. \x \ real k}) = - principal ({k. x \ real k} \ {k. \x \ real k})" by simp - also have "{k. x \ real k} \ {k. \x \ real k} = {}" by blast - finally have "inf sequentially (principal {k. \x \ real k}) = bot" - by (simp add: inf.coboundedI1 bot_unique) - with x show "(\k. f k x) \ exp (-a*x)" unfolding f_def - by (intro filterlim_If) auto - next - have "\integral {c..} (f k)\ \ exp (-a*c)/a" for k - proof (cases "c > of_nat k") - case False - hence "abs (integral {c..} (f k)) = abs (exp (- (a * c)) / a - exp (- (a * real k)) / a)" - by (simp add: integral_f) - also have "abs (exp (- (a * c)) / a - exp (- (a * real k)) / a) = - exp (- (a * c)) / a - exp (- (a * real k)) / a" - using False a by (intro abs_of_nonneg) (simp_all add: field_simps) - also have "\ \ exp (- a * c) / a" using a by simp - finally show ?thesis . - qed (insert a, simp_all add: integral_f) - thus "bounded {integral {c..} (f k) |k. True}" - by (intro bounded_realI[of _ "exp (-a*c)/a"]) auto - qed (auto simp: f_def) - - from eventually_gt_at_top[of "nat \c\"] have "eventually (\k. of_nat k > c) sequentially" - by eventually_elim linarith - hence "eventually (\k. exp (-a*c)/a - exp (-a * of_nat k)/a = integral {c..} (f k)) sequentially" - by eventually_elim (simp add: integral_f) - moreover have "(\k. exp (-a*c)/a - exp (-a * of_nat k)/a) \ exp (-a*c)/a - 0/a" - by (intro tendsto_intros filterlim_compose[OF exp_at_bot] - filterlim_tendsto_neg_mult_at_bot[OF tendsto_const] filterlim_real_sequentially)+ - (insert a, simp_all) - ultimately have "(\k. integral {c..} (f k)) \ exp (-a*c)/a - 0/a" - by (rule Lim_transform_eventually) - from LIMSEQ_unique[OF conjunct2[OF A] this] - have "integral {c..} (\x. exp (-a*x)) = exp (-a*c)/a" by simp - with conjunct1[OF A] show ?thesis - by (simp add: has_integral_integral) -qed - -lemma integrable_on_exp_minus_to_infinity: "a > 0 \ (\x. exp (-a*x) :: real) integrable_on {c..}" - using has_integral_exp_minus_to_infinity[of a c] unfolding integrable_on_def by blast - -lemma has_integral_powr_from_0: - assumes a: "a > (-1::real)" and c: "c \ 0" - shows "((\x. x powr a) has_integral (c powr (a+1) / (a+1))) {0..c}" -proof (cases "c = 0") - case False - 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}" - proof (cases "inverse (of_nat (Suc k)) \ c") - case True - { - fix x assume x: "x \ inverse (1 + real k)" - have "0 < inverse (1 + real k)" by simp - also note x - finally have "x > 0" . - } note x = this - hence "((\x. x powr a) has_integral c powr (a + 1) / (a + 1) - - inverse (real (Suc k)) powr (a + 1) / (a + 1)) {inverse (real (Suc k))..c}" - using True a by (intro fundamental_theorem_of_calculus) - (auto intro!: derivative_eq_intros continuous_on_powr' continuous_on_const - continuous_on_id simp: has_field_derivative_iff_has_vector_derivative [symmetric]) - with True show ?thesis unfolding f_def F_def by (subst has_integral_restrict) simp_all - next - case False - thus ?thesis unfolding f_def F_def by (subst has_integral_restrict) auto - qed - } note has_integral_f = this - have integral_f: "integral {0..c} (f k) = F k" for k - using has_integral_f[of k] by (rule integral_unique) - - have A: "(\x. x powr a) integrable_on {0..c} \ - (\k. integral {0..c} (f k)) \ integral {0..c} (\x. x powr a)" - proof (intro monotone_convergence_increasing ballI allI) - fix k from has_integral_f[of k] show "f k integrable_on {0..c}" - by (auto simp: integrable_on_def) - next - fix k :: nat and x :: real - { - assume x: "inverse (real (Suc k)) \ x" - have "inverse (real (Suc (Suc k))) \ inverse (real (Suc k))" by (simp add: field_simps) - also note x - finally have "inverse (real (Suc (Suc k))) \ x" . - } - thus "f k x \ f (Suc k) x" by (auto simp: f_def simp del: of_nat_Suc) - next - fix x assume x: "x \ {0..c}" - show "(\k. f k x) \ x powr a" - proof (cases "x = 0") - case False - with x have "x > 0" by simp - from order_tendstoD(2)[OF LIMSEQ_inverse_real_of_nat this] - have "eventually (\k. x powr a = f k x) sequentially" - by eventually_elim (insert x, simp add: f_def) - moreover have "(\_. x powr a) \ x powr a" by simp - ultimately show ?thesis by (rule Lim_transform_eventually) - qed (simp_all add: f_def) - next - { - fix k - from a have "F k \ c powr (a + 1) / (a + 1)" - by (auto simp add: F_def divide_simps) - also from a have "F k \ 0" - by (auto simp: F_def divide_simps simp del: of_nat_Suc intro!: powr_mono2) - hence "F k = abs (F k)" by simp - finally have "abs (F k) \ c powr (a + 1) / (a + 1)" . - } - thus "bounded {integral {0..c} (f k) |k. True}" - by (intro bounded_realI[of _ "c powr (a+1) / (a+1)"]) (auto simp: integral_f) - qed - - from False c have "c > 0" by simp - from order_tendstoD(2)[OF LIMSEQ_inverse_real_of_nat this] - have "eventually (\k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a+1) / (a+1) = - integral {0..c} (f k)) sequentially" - by eventually_elim (simp add: integral_f F_def) - moreover have "(\k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1)) - \ c powr (a + 1) / (a + 1) - 0 powr (a + 1) / (a + 1)" - using a by (intro tendsto_intros LIMSEQ_inverse_real_of_nat) auto - hence "(\k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1)) - \ c powr (a + 1) / (a + 1)" by simp - ultimately have "(\k. integral {0..c} (f k)) \ c powr (a+1) / (a+1)" - by (rule Lim_transform_eventually) - with A have "integral {0..c} (\x. x powr a) = c powr (a+1) / (a+1)" - by (blast intro: LIMSEQ_unique) - with A show ?thesis by (simp add: has_integral_integral) -qed (simp_all add: has_integral_refl) - -lemma integrable_on_powr_from_0: - assumes a: "a > (-1::real)" and c: "c \ 0" - shows "(\x. x powr a) integrable_on {0..c}" - using has_integral_powr_from_0[OF assms] unfolding integrable_on_def by blast - -end diff -r bbcb05504fdc -r bd218a9320b5 src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy --- a/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy Thu Aug 04 18:45:28 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy Thu Aug 04 19:36:31 2016 +0200 @@ -1,16 +1,17 @@ theory Multivariate_Analysis imports - Fashoda + Fashoda_Theorem Extended_Real_Limits Determinants Homeomorphism Ordered_Euclidean_Space Bounded_Continuous_Function - Weierstrass + Weierstrass_Theorems Polytope + Poly_Roots Conformal_Mappings Generalised_Binomial_Theorem - Gamma + Gamma_Function begin end diff -r bbcb05504fdc -r bd218a9320b5 src/HOL/Multivariate_Analysis/Path_Connected.thy --- a/src/HOL/Multivariate_Analysis/Path_Connected.thy Thu Aug 04 18:45:28 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Path_Connected.thy Thu Aug 04 19:36:31 2016 +0200 @@ -5,7 +5,7 @@ section \Continuous paths and path-connected sets\ theory Path_Connected -imports Extension +imports Continuous_Extension begin subsection \Paths and Arcs\ @@ -558,9 +558,9 @@ subsection\Some reversed and "if and only if" versions of joining theorems\ -lemma path_join_path_ends: +lemma path_join_path_ends: fixes g1 :: "real \ 'a::metric_space" - assumes "path(g1 +++ g2)" "path g2" + assumes "path(g1 +++ g2)" "path g2" shows "pathfinish g1 = pathstart g2" proof (rule ccontr) define e where "e = dist (g1 1) (g2 0)" @@ -568,14 +568,14 @@ then have "0 < dist (pathfinish g1) (pathstart g2)" by auto then have "e > 0" - by (metis e_def pathfinish_def pathstart_def) - then obtain d1 where "d1 > 0" + by (metis e_def pathfinish_def pathstart_def) + then obtain d1 where "d1 > 0" and d1: "\x'. \x'\{0..1}; norm x' < d1\ \ dist (g2 x') (g2 0) < e/2" using assms(2) unfolding path_def continuous_on_iff apply (drule_tac x=0 in bspec, simp) by (metis half_gt_zero_iff norm_conv_dist) - obtain d2 where "d2 > 0" - and d2: "\x'. \x'\{0..1}; dist x' (1/2) < d2\ + obtain d2 where "d2 > 0" + and d2: "\x'. \x'\{0..1}; dist x' (1/2) < d2\ \ dist ((g1 +++ g2) x') (g1 1) < e/2" using assms(1) \e > 0\ unfolding path_def continuous_on_iff apply (drule_tac x="1/2" in bspec, simp) @@ -597,25 +597,25 @@ using d1 [OF int01_1 dist1] d2 [OF int01_2 dist2] by (simp_all add: joinpaths_def) then have "dist (g1 1) (g2 0) < e/2 + e/2" using dist_triangle_half_r e_def by blast - then show False + then show False by (simp add: e_def [symmetric]) qed -lemma path_join_eq [simp]: +lemma path_join_eq [simp]: fixes g1 :: "real \ 'a::metric_space" assumes "path g1" "path g2" shows "path(g1 +++ g2) \ pathfinish g1 = pathstart g2" using assms by (metis path_join_path_ends path_join_imp) -lemma simple_path_joinE: +lemma simple_path_joinE: assumes "simple_path(g1 +++ g2)" and "pathfinish g1 = pathstart g2" - obtains "arc g1" "arc g2" + obtains "arc g1" "arc g2" "path_image g1 \ path_image g2 \ {pathstart g1, pathstart g2}" proof - - have *: "\x y. \0 \ x; x \ 1; 0 \ y; y \ 1; (g1 +++ g2) x = (g1 +++ g2) y\ + have *: "\x y. \0 \ x; x \ 1; 0 \ y; y \ 1; (g1 +++ g2) x = (g1 +++ g2) y\ \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" using assms by (simp add: simple_path_def) - have "path g1" + have "path g1" using assms path_join simple_path_imp_path by blast moreover have "inj_on g1 {0..1}" proof (clarsimp simp: inj_on_def) @@ -627,7 +627,7 @@ ultimately have "arc g1" using assms by (simp add: arc_def) have [simp]: "g2 0 = g1 1" - using assms by (metis pathfinish_def pathstart_def) + using assms by (metis pathfinish_def pathstart_def) have "path g2" using assms path_join simple_path_imp_path by blast moreover have "inj_on g2 {0..1}" @@ -640,7 +640,7 @@ qed ultimately have "arc g2" using assms by (simp add: arc_def) - have "g2 y = g1 0 \ g2 y = g1 1" + have "g2 y = g1 0 \ g2 y = g1 1" if "g1 x = g2 y" "0 \ x" "x \ 1" "0 \ y" "y \ 1" for x y using * [of "x / 2" "(y + 1) / 2"] that by (auto simp: joinpaths_def split_ifs divide_simps) @@ -650,20 +650,20 @@ qed lemma simple_path_join_loop_eq: - assumes "pathfinish g2 = pathstart g1" "pathfinish g1 = pathstart g2" + assumes "pathfinish g2 = pathstart g1" "pathfinish g1 = pathstart g2" shows "simple_path(g1 +++ g2) \ arc g1 \ arc g2 \ path_image g1 \ path_image g2 \ {pathstart g1, pathstart g2}" by (metis assms simple_path_joinE simple_path_join_loop) lemma arc_join_eq: - assumes "pathfinish g1 = pathstart g2" + assumes "pathfinish g1 = pathstart g2" shows "arc(g1 +++ g2) \ arc g1 \ arc g2 \ path_image g1 \ path_image g2 \ {pathstart g2}" (is "?lhs = ?rhs") -proof +proof assume ?lhs then have "simple_path(g1 +++ g2)" by (rule arc_imp_simple_path) - then have *: "\x y. \0 \ x; x \ 1; 0 \ y; y \ 1; (g1 +++ g2) x = (g1 +++ g2) y\ + then have *: "\x y. \0 \ x; x \ 1; 0 \ y; y \ 1; (g1 +++ g2) x = (g1 +++ g2) y\ \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" using assms by (simp add: simple_path_def) have False if "g1 0 = g2 u" "0 \ u" "u \ 1" for u @@ -681,7 +681,7 @@ by (fastforce simp: pathfinish_def pathstart_def intro!: arc_join) qed -lemma arc_join_eq_alt: +lemma arc_join_eq_alt: "pathfinish g1 = pathstart g2 \ (arc(g1 +++ g2) \ arc g1 \ arc g2 \ @@ -696,27 +696,27 @@ \ path(p +++ (q +++ r)) \ path((p +++ q) +++ r)" by simp -lemma simple_path_assoc: - assumes "pathfinish p = pathstart q" "pathfinish q = pathstart r" +lemma simple_path_assoc: + assumes "pathfinish p = pathstart q" "pathfinish q = pathstart r" shows "simple_path (p +++ (q +++ r)) \ simple_path ((p +++ q) +++ r)" proof (cases "pathstart p = pathfinish r") case True show ?thesis proof assume "simple_path (p +++ q +++ r)" with assms True show "simple_path ((p +++ q) +++ r)" - by (fastforce simp add: simple_path_join_loop_eq arc_join_eq path_image_join + by (fastforce simp add: simple_path_join_loop_eq arc_join_eq path_image_join dest: arc_distinct_ends [of r]) next assume 0: "simple_path ((p +++ q) +++ r)" with assms True have q: "pathfinish r \ path_image q" - using arc_distinct_ends + using arc_distinct_ends by (fastforce simp add: simple_path_join_loop_eq arc_join_eq path_image_join) have "pathstart r \ path_image p" using assms - by (metis 0 IntI arc_distinct_ends arc_join_eq_alt empty_iff insert_iff + by (metis 0 IntI arc_distinct_ends arc_join_eq_alt empty_iff insert_iff pathfinish_in_path_image pathfinish_join simple_path_joinE) with assms 0 q True show "simple_path (p +++ q +++ r)" - by (auto simp: simple_path_join_loop_eq arc_join_eq path_image_join + by (auto simp: simple_path_join_loop_eq arc_join_eq path_image_join dest!: subsetD [OF _ IntI]) qed next @@ -730,11 +730,11 @@ with a have "x = pathstart q" by blast } - with False assms show ?thesis + with False assms show ?thesis by (auto simp: simple_path_eq_arc simple_path_join_loop_eq arc_join_eq path_image_join) qed -lemma arc_assoc: +lemma arc_assoc: "\pathfinish p = pathstart q; pathfinish q = pathstart r\ \ arc(p +++ (q +++ r)) \ arc((p +++ q) +++ r)" by (simp add: arc_simple_path simple_path_assoc) @@ -1262,7 +1262,7 @@ shows "\e > 0. ball z e \ path_image g = {}" proof - obtain a where "a \ path_image g" "\y \ path_image g. dist z a \ dist z y" - apply (rule distance_attains_inf[OF _ path_image_nonempty, of g z]) + apply (rule distance_attains_inf[OF _ path_image_nonempty, of g z]) using compact_path_image[THEN compact_imp_closed, OF assms(1)] by auto then show ?thesis apply (rule_tac x="dist z a" in exI) @@ -2080,8 +2080,8 @@ corollary connected_open_delete_finite: fixes S T::"'a::euclidean_space set" assumes S: "open S" "connected S" and 2: "2 \ DIM('a)" and "finite T" - shows "connected(S - T)" - using \finite T\ S + shows "connected(S - T)" + using \finite T\ S proof (induct T) case empty show ?case using \connected S\ by simp @@ -3068,12 +3068,12 @@ by (force simp: imageI image_mono interiorI interior_subset frontier_def y) have **: "(~(b \ (- S) = {}) \ ~(b - (- S) = {}) \ (b \ f \ {})) \ (b \ S \ {}) \ b \ f = {} \ - b \ S" for b f and S :: "'b set" + b \ S" for b f and S :: "'b set" by blast show ?thesis apply (rule **) (*such a horrible mess*) apply (rule connected_Int_frontier [where t = "f`S", OF connected_ball]) - using \a \ S\ \0 < r\ + using \a \ S\ \0 < r\ apply (auto simp: disjoint_iff_not_equal dist_norm) by (metis dw_le norm_minus_commute not_less order_trans rle wy) qed @@ -4003,7 +4003,7 @@ lemma homotopic_points_eq_path_component: "homotopic_loops S (linepath a a) (linepath b b) \ path_component S a b" -by (auto simp: path_component_imp_homotopic_points +by (auto simp: path_component_imp_homotopic_points dest: homotopic_loops_imp_path_component_value [where t=1]) lemma path_connected_eq_homotopic_points: @@ -4058,11 +4058,11 @@ apply (fastforce simp: simply_connected_imp_path_connected simply_connected_eq_contractible_loop_any) apply (clarsimp simp add: simply_connected_eq_contractible_loop_any) apply (drule_tac x=p in spec) -using homotopic_loops_trans path_connected_eq_homotopic_points +using homotopic_loops_trans path_connected_eq_homotopic_points apply blast done -lemma simply_connected_eq_contractible_loop_all: +lemma simply_connected_eq_contractible_loop_all: fixes S :: "_::real_normed_vector set" shows "simply_connected S \ S = {} \ @@ -4075,21 +4075,21 @@ case False then obtain a where "a \ S" by blast show ?thesis - proof + proof assume "simply_connected S" then show ?rhs - using \a \ S\ \simply_connected S\ simply_connected_eq_contractible_loop_any + using \a \ S\ \simply_connected S\ simply_connected_eq_contractible_loop_any by blast - next + next assume ?rhs then show "simply_connected S" apply (simp add: simply_connected_eq_contractible_loop_any False) - by (meson homotopic_loops_refl homotopic_loops_sym homotopic_loops_trans + by (meson homotopic_loops_refl homotopic_loops_sym homotopic_loops_trans path_component_imp_homotopic_points path_component_refl) qed qed -lemma simply_connected_eq_contractible_path: +lemma simply_connected_eq_contractible_path: fixes S :: "_::real_normed_vector set" shows "simply_connected S \ path_connected S \ @@ -4098,7 +4098,7 @@ apply (rule iffI) apply (simp add: simply_connected_imp_path_connected) apply (metis simply_connected_eq_contractible_loop_some homotopic_loops_imp_homotopic_paths_null) -by (meson homotopic_paths_imp_homotopic_loops pathfinish_linepath pathstart_in_path_image +by (meson homotopic_paths_imp_homotopic_loops pathfinish_linepath pathstart_in_path_image simply_connected_eq_contractible_loop_some subset_iff) lemma simply_connected_eq_homotopic_paths: @@ -4112,23 +4112,23 @@ (is "?lhs = ?rhs") proof assume ?lhs - then have pc: "path_connected S" + then have pc: "path_connected S" and *: "\p. \path p; path_image p \ S; - pathfinish p = pathstart p\ + pathfinish p = pathstart p\ \ homotopic_paths S p (linepath (pathstart p) (pathstart p))" by (auto simp: simply_connected_eq_contractible_path) - have "homotopic_paths S p q" + have "homotopic_paths S p q" if "path p" "path_image p \ S" "path q" "path_image q \ S" "pathstart q = pathstart p" "pathfinish q = pathfinish p" for p q proof - - have "homotopic_paths S p (p +++ linepath (pathfinish p) (pathfinish p))" + have "homotopic_paths S p (p +++ linepath (pathfinish p) (pathfinish p))" by (simp add: homotopic_paths_rid homotopic_paths_sym that) also have "homotopic_paths S (p +++ linepath (pathfinish p) (pathfinish p)) (p +++ reversepath q +++ q)" using that by (metis homotopic_paths_join homotopic_paths_linv homotopic_paths_refl homotopic_paths_sym_eq pathstart_linepath) - also have "homotopic_paths S (p +++ reversepath q +++ q) + also have "homotopic_paths S (p +++ reversepath q +++ q) ((p +++ reversepath q) +++ q)" by (simp add: that homotopic_paths_assoc) also have "homotopic_paths S ((p +++ reversepath q) +++ q) @@ -4142,7 +4142,7 @@ then show ?rhs by (blast intro: pc *) next - assume ?rhs + assume ?rhs then show ?lhs by (force simp: simply_connected_eq_contractible_path) qed @@ -4408,7 +4408,7 @@ lemma is_interval_contractible_1: fixes S :: "real set" shows "is_interval S \ contractible S" -using contractible_imp_simply_connected convex_imp_contractible is_interval_convex_1 +using contractible_imp_simply_connected convex_imp_contractible is_interval_convex_1 is_interval_simply_connected_1 by auto lemma contractible_Times: @@ -4416,14 +4416,14 @@ assumes S: "contractible S" and T: "contractible T" shows "contractible (S \ T)" proof - - obtain a h where conth: "continuous_on ({0..1} \ S) h" + obtain a h where conth: "continuous_on ({0..1} \ S) h" and hsub: "h ` ({0..1} \ S) \ S" - and [simp]: "\x. x \ S \ h (0, x) = x" + and [simp]: "\x. x \ S \ h (0, x) = x" and [simp]: "\x. x \ S \ h (1::real, x) = a" using S by (auto simp add: contractible_def homotopic_with) - obtain b k where contk: "continuous_on ({0..1} \ T) k" + obtain b k where contk: "continuous_on ({0..1} \ T) k" and ksub: "k ` ({0..1} \ T) \ T" - and [simp]: "\x. x \ T \ k (0, x) = x" + and [simp]: "\x. x \ T \ k (0, x) = x" and [simp]: "\x. x \ T \ k (1::real, x) = b" using T by (auto simp add: contractible_def homotopic_with) show ?thesis @@ -4432,16 +4432,16 @@ apply (rule exI [where x=b]) apply (rule exI [where x = "\z. (h (fst z, fst(snd z)), k (fst z, snd(snd z)))"]) apply (intro conjI ballI continuous_intros continuous_on_compose2 [OF conth] continuous_on_compose2 [OF contk]) - using hsub ksub + using hsub ksub apply (auto simp: ) done qed -lemma homotopy_dominated_contractibility: +lemma homotopy_dominated_contractibility: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" assumes S: "contractible S" - and f: "continuous_on S f" "image f S \ T" - and g: "continuous_on T g" "image g T \ S" + and f: "continuous_on S f" "image f S \ T" + and g: "continuous_on T g" "image g T \ S" and hom: "homotopic_with (\x. True) T T (f o g) id" shows "contractible T" proof - diff -r bbcb05504fdc -r bd218a9320b5 src/HOL/Multivariate_Analysis/PolyRoots.thy --- a/src/HOL/Multivariate_Analysis/PolyRoots.thy Thu Aug 04 18:45:28 2016 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,301 +0,0 @@ -(* Author: John Harrison and Valentina Bruno - Ported from "hol_light/Multivariate/complexes.ml" by L C Paulson -*) - -section \polynomial functions: extremal behaviour and root counts\ - -theory PolyRoots -imports Complex_Main -begin - -subsection\Geometric progressions\ - -lemma setsum_gp_basic: - fixes x :: "'a::{comm_ring,monoid_mult}" - shows "(1 - x) * (\i\n. x^i) = 1 - x^Suc n" - by (simp only: one_diff_power_eq [of "Suc n" x] lessThan_Suc_atMost) - -lemma setsum_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 setsum_gp_basic[of x n] - by (simp add: mult.commute divide_simps) - -lemma setsum_power_add: - fixes x :: "'a::{comm_ring,monoid_mult}" - shows "(\i\I. x^(m+i)) = x^m * (\i\I. x^i)" - by (simp add: setsum_right_distrib power_add) - -lemma setsum_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: setsum_right_distrib power_add [symmetric]) - also have "(\i=m..n. x^(i-m)) = (\i\n-m. x^i)" - using \m \ n\ by (intro setsum.reindex_bij_witness[where j="\i. i - m" and i="\i. i + m"]) auto - finally show ?thesis . -qed - -lemma setsum_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 setsum_power_shift) - also have "... =x^m * (1 - x^Suc(n-m))" - by (metis mult.assoc setsum_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 setsum_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 setsum_gp_multiplied [of m n x] -apply auto -by (metis eq_iff_diff_eq_0 mult.commute nonzero_divide_eq_eq) - -lemma setsum_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 setsum_gp [of x m "m+n"] - by (auto simp: power_add algebra_simps) - -lemma setsum_gp_strict: - fixes x :: "'a::{comm_ring,division_ring}" - shows "(\iBasics about polynomial functions: extremal behaviour and root counts.\ - -lemma sub_polyfun: - fixes x :: "'a::{comm_ring,monoid_mult}" - shows "(\i\n. a i * x^i) - (\i\n. a i * y^i) = - (x - y) * (\jk= Suc j..n. a k * y^(k - Suc j) * x^j)" -proof - - have "(\i\n. a i * x^i) - (\i\n. a i * y^i) = - (\i\n. a i * (x^i - y^i))" - by (simp add: algebra_simps setsum_subtractf [symmetric]) - also have "... = (\i\n. a i * (x - y) * (\ji\n. (\jji=Suc j..n. a i * y^(i - Suc j) * x^j))" - by (simp add: nested_setsum_swap') - finally show ?thesis . -qed - -lemma sub_polyfun_alt: - fixes x :: "'a::{comm_ring,monoid_mult}" - shows "(\i\n. a i * x^i) - (\i\n. a i * y^i) = - (x - y) * (\jkk = Suc j..n. a k * y^(k - Suc j) * x^j) = - (\k b. \z. (\i\n. c i * z^i) = - (z-a) * (\ii\n. c i * a^i)" -proof - - { fix z - have "(\i\n. c i * z^i) - (\i\n. c i * a^i) = - (z - a) * (\jk = Suc j..n. c k * a^(k - Suc j)) * z^j)" - by (simp add: sub_polyfun setsum_left_distrib) - then have "(\i\n. c i * z^i) = - (z - a) * (\jk = Suc j..n. c k * a^(k - Suc j)) * z^j) - + (\i\n. c i * a^i)" - by (simp add: algebra_simps) } - then show ?thesis - by (intro exI allI) -qed - -lemma polyfun_linear_factor_root: - fixes a :: "'a::{comm_ring,monoid_mult}" - assumes "(\i\n. c i * a^i) = 0" - shows "\b. \z. (\i\n. c i * z^i) = (z-a) * (\i b ==> norm(x) \ a ==> norm(x + y) \ b" - by (metis norm_triangle_mono order.trans order_refl) - -lemma polyfun_extremal_lemma: - fixes c :: "nat \ 'a::real_normed_div_algebra" - assumes "e > 0" - shows "\M. \z. M \ norm z \ norm(\i\n. c i * z^i) \ e * norm(z) ^ Suc n" -proof (induction n) - case 0 - show ?case - by (rule exI [where x="norm (c 0) / e"]) (auto simp: mult.commute pos_divide_le_eq assms) -next - case (Suc n) - then obtain M where M: "\z. M \ norm z \ norm (\i\n. c i * z^i) \ e * norm z ^ Suc n" .. - show ?case - proof (rule exI [where x="max 1 (max M ((e + norm(c(Suc n))) / e))"], clarify) - fix z::'a - assume "max 1 (max M ((e + norm (c (Suc n))) / e)) \ norm z" - then have norm1: "0 < norm z" "M \ norm z" "(e + norm (c (Suc n))) / e \ norm z" - by auto - then have norm2: "(e + norm (c (Suc n))) \ e * norm z" "(norm z * norm z ^ n) > 0" - apply (metis assms less_divide_eq mult.commute not_le) - using norm1 apply (metis mult_pos_pos zero_less_power) - done - have "e * (norm z * norm z ^ n) + norm (c (Suc n) * (z * z ^ n)) = - (e + norm (c (Suc n))) * (norm z * norm z ^ n)" - by (simp add: norm_mult norm_power algebra_simps) - also have "... \ (e * norm z) * (norm z * norm z ^ n)" - using norm2 by (metis real_mult_le_cancel_iff1) - also have "... = e * (norm z * (norm z * norm z ^ n))" - by (simp add: algebra_simps) - finally have "e * (norm z * norm z ^ n) + norm (c (Suc n) * (z * z ^ n)) - \ e * (norm z * (norm z * norm z ^ n))" . - then show "norm (\i\Suc n. c i * z^i) \ e * norm z ^ Suc (Suc n)" using M norm1 - by (drule_tac x=z in spec) (auto simp: intro!: adhoc_norm_triangle) - qed -qed - -lemma norm_lemma_xy: assumes "\b\ + 1 \ norm(y) - a" "norm(x) \ a" shows "b \ norm(x + y)" -proof - - have "b \ norm y - norm x" - using assms by linarith - then show ?thesis - by (metis (no_types) add.commute norm_diff_ineq order_trans) -qed - -lemma polyfun_extremal: - fixes c :: "nat \ 'a::real_normed_div_algebra" - assumes "\k. k \ 0 \ k \ n \ c k \ 0" - shows "eventually (\z. norm(\i\n. c i * z^i) \ B) at_infinity" -using assms -proof (induction n) - case 0 then show ?case - by simp -next - case (Suc n) - show ?case - proof (cases "c (Suc n) = 0") - case True - with Suc show ?thesis - by auto (metis diff_is_0_eq diffs0_imp_equal less_Suc_eq_le not_less_eq) - next - case False - with polyfun_extremal_lemma [of "norm(c (Suc n)) / 2" c n] - obtain M where M: "\z. M \ norm z \ - norm (\i\n. c i * z^i) \ norm (c (Suc n)) / 2 * norm z ^ Suc n" - by auto - show ?thesis - unfolding eventually_at_infinity - proof (rule exI [where x="max M (max 1 ((\B\ + 1) / (norm (c (Suc n)) / 2)))"], clarsimp) - fix z::'a - assume les: "M \ norm z" "1 \ norm z" "(\B\ * 2 + 2) / norm (c (Suc n)) \ norm z" - then have "\B\ * 2 + 2 \ norm z * norm (c (Suc n))" - by (metis False pos_divide_le_eq zero_less_norm_iff) - then have "\B\ * 2 + 2 \ norm z ^ (Suc n) * norm (c (Suc n))" - by (metis \1 \ norm z\ order.trans mult_right_mono norm_ge_zero self_le_power zero_less_Suc) - then show "B \ norm ((\i\n. c i * z^i) + c (Suc n) * (z * z ^ n))" using M les - apply auto - apply (rule norm_lemma_xy [where a = "norm (c (Suc n)) * norm z ^ (Suc n) / 2"]) - apply (simp_all add: norm_mult norm_power) - done - qed - qed -qed - -lemma polyfun_rootbound: - fixes c :: "nat \ 'a::{comm_ring,real_normed_div_algebra}" - assumes "\k. k \ n \ c k \ 0" - shows "finite {z. (\i\n. c i * z^i) = 0} \ card {z. (\i\n. c i * z^i) = 0} \ n" -using assms -proof (induction n arbitrary: c) - case (Suc n) show ?case - proof (cases "{z. (\i\Suc n. c i * z^i) = 0} = {}") - case False - then obtain a where a: "(\i\Suc n. c i * a^i) = 0" - by auto - from polyfun_linear_factor_root [OF this] - obtain b where "\z. (\i\Suc n. c i * z^i) = (z - a) * (\i< Suc n. b i * z^i)" - by auto - then have b: "\z. (\i\Suc n. c i * z^i) = (z - a) * (\i\n. b i * z^i)" - by (metis lessThan_Suc_atMost) - then have ins_ab: "{z. (\i\Suc n. c i * z^i) = 0} = insert a {z. (\i\n. b i * z^i) = 0}" - by auto - have c0: "c 0 = - (a * b 0)" using b [of 0] - by simp - then have extr_prem: "~ (\k\n. b k \ 0) \ \k. k \ 0 \ k \ Suc n \ c k \ 0" - by (metis Suc.prems le0 minus_zero mult_zero_right) - have "\k\n. b k \ 0" - apply (rule ccontr) - using polyfun_extremal [OF extr_prem, of 1] - apply (auto simp: eventually_at_infinity b simp del: setsum_atMost_Suc) - apply (drule_tac x="of_real ba" in spec, simp) - done - then show ?thesis using Suc.IH [of b] ins_ab - by (auto simp: card_insert_if) - qed simp -qed simp - -corollary - fixes c :: "nat \ 'a::{comm_ring,real_normed_div_algebra}" - assumes "\k. k \ n \ c k \ 0" - shows polyfun_rootbound_finite: "finite {z. (\i\n. c i * z^i) = 0}" - and polyfun_rootbound_card: "card {z. (\i\n. c i * z^i) = 0} \ n" -using polyfun_rootbound [OF assms] by auto - -lemma polyfun_finite_roots: - fixes c :: "nat \ 'a::{comm_ring,real_normed_div_algebra}" - shows "finite {z. (\i\n. c i * z^i) = 0} \ (\k. k \ n \ c k \ 0)" -proof (cases " \k\n. c k \ 0") - case True then show ?thesis - by (blast intro: polyfun_rootbound_finite) -next - case False then show ?thesis - by (auto simp: infinite_UNIV_char_0) -qed - -lemma polyfun_eq_0: - fixes c :: "nat \ 'a::{comm_ring,real_normed_div_algebra}" - shows "(\z. (\i\n. c i * z^i) = 0) \ (\k. k \ n \ c k = 0)" -proof (cases "(\z. (\i\n. c i * z^i) = 0)") - case True - then have "~ finite {z. (\i\n. c i * z^i) = 0}" - by (simp add: infinite_UNIV_char_0) - with True show ?thesis - by (metis (poly_guards_query) polyfun_rootbound_finite) -next - case False - then show ?thesis - by auto -qed - -lemma polyfun_eq_const: - fixes c :: "nat \ 'a::{comm_ring,real_normed_div_algebra}" - shows "(\z. (\i\n. c i * z^i) = k) \ c 0 = k \ (\k. k \ 0 \ k \ n \ c k = 0)" -proof - - {fix z - have "(\i\n. c i * z^i) = (\i\n. (if i = 0 then c 0 - k else c i) * z^i) + k" - by (induct n) auto - } then - have "(\z. (\i\n. c i * z^i) = k) \ (\z. (\i\n. (if i = 0 then c 0 - k else c i) * z^i) = 0)" - by auto - also have "... \ c 0 = k \ (\k. k \ 0 \ k \ n \ c k = 0)" - by (auto simp: polyfun_eq_0) - finally show ?thesis . -qed - -end - diff -r bbcb05504fdc -r bd218a9320b5 src/HOL/Multivariate_Analysis/Poly_Roots.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Multivariate_Analysis/Poly_Roots.thy Thu Aug 04 19:36:31 2016 +0200 @@ -0,0 +1,301 @@ +(* Author: John Harrison and Valentina Bruno + Ported from "hol_light/Multivariate/complexes.ml" by L C Paulson +*) + +section \polynomial functions: extremal behaviour and root counts\ + +theory Poly_Roots +imports Complex_Main +begin + +subsection\Geometric progressions\ + +lemma setsum_gp_basic: + fixes x :: "'a::{comm_ring,monoid_mult}" + shows "(1 - x) * (\i\n. x^i) = 1 - x^Suc n" + by (simp only: one_diff_power_eq [of "Suc n" x] lessThan_Suc_atMost) + +lemma setsum_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 setsum_gp_basic[of x n] + by (simp add: mult.commute divide_simps) + +lemma setsum_power_add: + fixes x :: "'a::{comm_ring,monoid_mult}" + shows "(\i\I. x^(m+i)) = x^m * (\i\I. x^i)" + by (simp add: setsum_right_distrib power_add) + +lemma setsum_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: setsum_right_distrib power_add [symmetric]) + also have "(\i=m..n. x^(i-m)) = (\i\n-m. x^i)" + using \m \ n\ by (intro setsum.reindex_bij_witness[where j="\i. i - m" and i="\i. i + m"]) auto + finally show ?thesis . +qed + +lemma setsum_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 setsum_power_shift) + also have "... =x^m * (1 - x^Suc(n-m))" + by (metis mult.assoc setsum_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 setsum_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 setsum_gp_multiplied [of m n x] +apply auto +by (metis eq_iff_diff_eq_0 mult.commute nonzero_divide_eq_eq) + +lemma setsum_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 setsum_gp [of x m "m+n"] + by (auto simp: power_add algebra_simps) + +lemma setsum_gp_strict: + fixes x :: "'a::{comm_ring,division_ring}" + shows "(\iBasics about polynomial functions: extremal behaviour and root counts.\ + +lemma sub_polyfun: + fixes x :: "'a::{comm_ring,monoid_mult}" + shows "(\i\n. a i * x^i) - (\i\n. a i * y^i) = + (x - y) * (\jk= Suc j..n. a k * y^(k - Suc j) * x^j)" +proof - + have "(\i\n. a i * x^i) - (\i\n. a i * y^i) = + (\i\n. a i * (x^i - y^i))" + by (simp add: algebra_simps setsum_subtractf [symmetric]) + also have "... = (\i\n. a i * (x - y) * (\ji\n. (\jji=Suc j..n. a i * y^(i - Suc j) * x^j))" + by (simp add: nested_setsum_swap') + finally show ?thesis . +qed + +lemma sub_polyfun_alt: + fixes x :: "'a::{comm_ring,monoid_mult}" + shows "(\i\n. a i * x^i) - (\i\n. a i * y^i) = + (x - y) * (\jkk = Suc j..n. a k * y^(k - Suc j) * x^j) = + (\k b. \z. (\i\n. c i * z^i) = + (z-a) * (\ii\n. c i * a^i)" +proof - + { fix z + have "(\i\n. c i * z^i) - (\i\n. c i * a^i) = + (z - a) * (\jk = Suc j..n. c k * a^(k - Suc j)) * z^j)" + by (simp add: sub_polyfun setsum_left_distrib) + then have "(\i\n. c i * z^i) = + (z - a) * (\jk = Suc j..n. c k * a^(k - Suc j)) * z^j) + + (\i\n. c i * a^i)" + by (simp add: algebra_simps) } + then show ?thesis + by (intro exI allI) +qed + +lemma polyfun_linear_factor_root: + fixes a :: "'a::{comm_ring,monoid_mult}" + assumes "(\i\n. c i * a^i) = 0" + shows "\b. \z. (\i\n. c i * z^i) = (z-a) * (\i b ==> norm(x) \ a ==> norm(x + y) \ b" + by (metis norm_triangle_mono order.trans order_refl) + +lemma polyfun_extremal_lemma: + fixes c :: "nat \ 'a::real_normed_div_algebra" + assumes "e > 0" + shows "\M. \z. M \ norm z \ norm(\i\n. c i * z^i) \ e * norm(z) ^ Suc n" +proof (induction n) + case 0 + show ?case + by (rule exI [where x="norm (c 0) / e"]) (auto simp: mult.commute pos_divide_le_eq assms) +next + case (Suc n) + then obtain M where M: "\z. M \ norm z \ norm (\i\n. c i * z^i) \ e * norm z ^ Suc n" .. + show ?case + proof (rule exI [where x="max 1 (max M ((e + norm(c(Suc n))) / e))"], clarify) + fix z::'a + assume "max 1 (max M ((e + norm (c (Suc n))) / e)) \ norm z" + then have norm1: "0 < norm z" "M \ norm z" "(e + norm (c (Suc n))) / e \ norm z" + by auto + then have norm2: "(e + norm (c (Suc n))) \ e * norm z" "(norm z * norm z ^ n) > 0" + apply (metis assms less_divide_eq mult.commute not_le) + using norm1 apply (metis mult_pos_pos zero_less_power) + done + have "e * (norm z * norm z ^ n) + norm (c (Suc n) * (z * z ^ n)) = + (e + norm (c (Suc n))) * (norm z * norm z ^ n)" + by (simp add: norm_mult norm_power algebra_simps) + also have "... \ (e * norm z) * (norm z * norm z ^ n)" + using norm2 by (metis real_mult_le_cancel_iff1) + also have "... = e * (norm z * (norm z * norm z ^ n))" + by (simp add: algebra_simps) + finally have "e * (norm z * norm z ^ n) + norm (c (Suc n) * (z * z ^ n)) + \ e * (norm z * (norm z * norm z ^ n))" . + then show "norm (\i\Suc n. c i * z^i) \ e * norm z ^ Suc (Suc n)" using M norm1 + by (drule_tac x=z in spec) (auto simp: intro!: adhoc_norm_triangle) + qed +qed + +lemma norm_lemma_xy: assumes "\b\ + 1 \ norm(y) - a" "norm(x) \ a" shows "b \ norm(x + y)" +proof - + have "b \ norm y - norm x" + using assms by linarith + then show ?thesis + by (metis (no_types) add.commute norm_diff_ineq order_trans) +qed + +lemma polyfun_extremal: + fixes c :: "nat \ 'a::real_normed_div_algebra" + assumes "\k. k \ 0 \ k \ n \ c k \ 0" + shows "eventually (\z. norm(\i\n. c i * z^i) \ B) at_infinity" +using assms +proof (induction n) + case 0 then show ?case + by simp +next + case (Suc n) + show ?case + proof (cases "c (Suc n) = 0") + case True + with Suc show ?thesis + by auto (metis diff_is_0_eq diffs0_imp_equal less_Suc_eq_le not_less_eq) + next + case False + with polyfun_extremal_lemma [of "norm(c (Suc n)) / 2" c n] + obtain M where M: "\z. M \ norm z \ + norm (\i\n. c i * z^i) \ norm (c (Suc n)) / 2 * norm z ^ Suc n" + by auto + show ?thesis + unfolding eventually_at_infinity + proof (rule exI [where x="max M (max 1 ((\B\ + 1) / (norm (c (Suc n)) / 2)))"], clarsimp) + fix z::'a + assume les: "M \ norm z" "1 \ norm z" "(\B\ * 2 + 2) / norm (c (Suc n)) \ norm z" + then have "\B\ * 2 + 2 \ norm z * norm (c (Suc n))" + by (metis False pos_divide_le_eq zero_less_norm_iff) + then have "\B\ * 2 + 2 \ norm z ^ (Suc n) * norm (c (Suc n))" + by (metis \1 \ norm z\ order.trans mult_right_mono norm_ge_zero self_le_power zero_less_Suc) + then show "B \ norm ((\i\n. c i * z^i) + c (Suc n) * (z * z ^ n))" using M les + apply auto + apply (rule norm_lemma_xy [where a = "norm (c (Suc n)) * norm z ^ (Suc n) / 2"]) + apply (simp_all add: norm_mult norm_power) + done + qed + qed +qed + +lemma polyfun_rootbound: + fixes c :: "nat \ 'a::{comm_ring,real_normed_div_algebra}" + assumes "\k. k \ n \ c k \ 0" + shows "finite {z. (\i\n. c i * z^i) = 0} \ card {z. (\i\n. c i * z^i) = 0} \ n" +using assms +proof (induction n arbitrary: c) + case (Suc n) show ?case + proof (cases "{z. (\i\Suc n. c i * z^i) = 0} = {}") + case False + then obtain a where a: "(\i\Suc n. c i * a^i) = 0" + by auto + from polyfun_linear_factor_root [OF this] + obtain b where "\z. (\i\Suc n. c i * z^i) = (z - a) * (\i< Suc n. b i * z^i)" + by auto + then have b: "\z. (\i\Suc n. c i * z^i) = (z - a) * (\i\n. b i * z^i)" + by (metis lessThan_Suc_atMost) + then have ins_ab: "{z. (\i\Suc n. c i * z^i) = 0} = insert a {z. (\i\n. b i * z^i) = 0}" + by auto + have c0: "c 0 = - (a * b 0)" using b [of 0] + by simp + then have extr_prem: "~ (\k\n. b k \ 0) \ \k. k \ 0 \ k \ Suc n \ c k \ 0" + by (metis Suc.prems le0 minus_zero mult_zero_right) + have "\k\n. b k \ 0" + apply (rule ccontr) + using polyfun_extremal [OF extr_prem, of 1] + apply (auto simp: eventually_at_infinity b simp del: setsum_atMost_Suc) + apply (drule_tac x="of_real ba" in spec, simp) + done + then show ?thesis using Suc.IH [of b] ins_ab + by (auto simp: card_insert_if) + qed simp +qed simp + +corollary + fixes c :: "nat \ 'a::{comm_ring,real_normed_div_algebra}" + assumes "\k. k \ n \ c k \ 0" + shows polyfun_rootbound_finite: "finite {z. (\i\n. c i * z^i) = 0}" + and polyfun_rootbound_card: "card {z. (\i\n. c i * z^i) = 0} \ n" +using polyfun_rootbound [OF assms] by auto + +lemma polyfun_finite_roots: + fixes c :: "nat \ 'a::{comm_ring,real_normed_div_algebra}" + shows "finite {z. (\i\n. c i * z^i) = 0} \ (\k. k \ n \ c k \ 0)" +proof (cases " \k\n. c k \ 0") + case True then show ?thesis + by (blast intro: polyfun_rootbound_finite) +next + case False then show ?thesis + by (auto simp: infinite_UNIV_char_0) +qed + +lemma polyfun_eq_0: + fixes c :: "nat \ 'a::{comm_ring,real_normed_div_algebra}" + shows "(\z. (\i\n. c i * z^i) = 0) \ (\k. k \ n \ c k = 0)" +proof (cases "(\z. (\i\n. c i * z^i) = 0)") + case True + then have "~ finite {z. (\i\n. c i * z^i) = 0}" + by (simp add: infinite_UNIV_char_0) + with True show ?thesis + by (metis (poly_guards_query) polyfun_rootbound_finite) +next + case False + then show ?thesis + by auto +qed + +lemma polyfun_eq_const: + fixes c :: "nat \ 'a::{comm_ring,real_normed_div_algebra}" + shows "(\z. (\i\n. c i * z^i) = k) \ c 0 = k \ (\k. k \ 0 \ k \ n \ c k = 0)" +proof - + {fix z + have "(\i\n. c i * z^i) = (\i\n. (if i = 0 then c 0 - k else c i) * z^i) + k" + by (induct n) auto + } then + have "(\z. (\i\n. c i * z^i) = k) \ (\z. (\i\n. (if i = 0 then c 0 - k else c i) * z^i) = 0)" + by auto + also have "... \ c 0 = k \ (\k. k \ 0 \ k \ n \ c k = 0)" + by (auto simp: polyfun_eq_0) + finally show ?thesis . +qed + +end + diff -r bbcb05504fdc -r bd218a9320b5 src/HOL/Multivariate_Analysis/Summation.thy --- a/src/HOL/Multivariate_Analysis/Summation.thy Thu Aug 04 18:45:28 2016 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,920 +0,0 @@ -(* Title: HOL/Multivariate_Analysis/Summation.thy - Author: Manuel Eberl, TU München -*) - -section \Radius of Convergence and Summation Tests\ - -theory Summation -imports - Complex_Main - "~~/src/HOL/Library/Extended_Real" - "~~/src/HOL/Library/Liminf_Limsup" -begin - -text \ - The definition of the radius of convergence of a power series, - various summability tests, lemmas to compute the radius of convergence etc. -\ - -subsection \Rounded dual logarithm\ - -(* This is required for the Cauchy condensation criterion *) - -definition "natlog2 n = (if n = 0 then 0 else nat \log 2 (real_of_nat n)\)" - -lemma natlog2_0 [simp]: "natlog2 0 = 0" by (simp add: natlog2_def) -lemma natlog2_1 [simp]: "natlog2 1 = 0" by (simp add: natlog2_def) -lemma natlog2_eq_0_iff: "natlog2 n = 0 \ n < 2" by (simp add: natlog2_def) - -lemma natlog2_power_of_two [simp]: "natlog2 (2 ^ n) = n" - by (simp add: natlog2_def log_nat_power) - -lemma natlog2_mono: "m \ n \ natlog2 m \ natlog2 n" - unfolding natlog2_def by (simp_all add: nat_mono floor_mono) - -lemma pow_natlog2_le: "n > 0 \ 2 ^ natlog2 n \ n" -proof - - assume n: "n > 0" - from n have "of_nat (2 ^ natlog2 n) = 2 powr real_of_nat (nat \log 2 (real_of_nat n)\)" - by (subst powr_realpow) (simp_all add: natlog2_def) - also have "\ = 2 powr of_int \log 2 (real_of_nat n)\" using n by simp - also have "\ \ 2 powr log 2 (real_of_nat n)" by (intro powr_mono) (linarith, simp_all) - also have "\ = of_nat n" using n by simp - finally show ?thesis by simp -qed - -lemma pow_natlog2_gt: "n > 0 \ 2 * 2 ^ natlog2 n > n" - and pow_natlog2_ge: "n > 0 \ 2 * 2 ^ natlog2 n \ n" -proof - - assume n: "n > 0" - from n have "of_nat n = 2 powr log 2 (real_of_nat n)" by simp - also have "\ < 2 powr (1 + of_int \log 2 (real_of_nat n)\)" - by (intro powr_less_mono) (linarith, simp_all) - also from n have "\ = 2 powr (1 + real_of_nat (nat \log 2 (real_of_nat n)\))" by simp - also from n have "\ = of_nat (2 * 2 ^ natlog2 n)" - by (simp_all add: natlog2_def powr_real_of_int powr_add) - finally show "2 * 2 ^ natlog2 n > n" by (rule of_nat_less_imp_less) - thus "2 * 2 ^ natlog2 n \ n" by simp -qed - -lemma natlog2_eqI: - assumes "n > 0" "2^k \ n" "n < 2 * 2^k" - shows "natlog2 n = k" -proof - - from assms have "of_nat (2 ^ k) \ real_of_nat n" by (subst of_nat_le_iff) simp_all - hence "real_of_int (int k) \ log (of_nat 2) (real_of_nat n)" - by (subst le_log_iff) (simp_all add: powr_realpow assms del: of_nat_le_iff) - moreover from assms have "real_of_nat n < of_nat (2 ^ Suc k)" by (subst of_nat_less_iff) simp_all - hence "log 2 (real_of_nat n) < of_nat k + 1" - by (subst log_less_iff) (simp_all add: assms powr_realpow powr_add) - ultimately have "\log 2 (real_of_nat n)\ = of_nat k" by (intro floor_unique) simp_all - with assms show ?thesis by (simp add: natlog2_def) -qed - -lemma natlog2_rec: - assumes "n \ 2" - shows "natlog2 n = 1 + natlog2 (n div 2)" -proof (rule natlog2_eqI) - from assms have "2 ^ (1 + natlog2 (n div 2)) \ 2 * (n div 2)" - by (simp add: pow_natlog2_le) - also have "\ \ n" by simp - finally show "2 ^ (1 + natlog2 (n div 2)) \ n" . -next - from assms have "n < 2 * (n div 2 + 1)" by simp - also from assms have "(n div 2) < 2 ^ (1 + natlog2 (n div 2))" - by (simp add: pow_natlog2_gt) - hence "2 * (n div 2 + 1) \ 2 * (2 ^ (1 + natlog2 (n div 2)))" - by (intro mult_left_mono) simp_all - finally show "n < 2 * 2 ^ (1 + natlog2 (n div 2))" . -qed (insert assms, simp_all) - -fun natlog2_aux where - "natlog2_aux n acc = (if (n::nat) < 2 then acc else natlog2_aux (n div 2) (acc + 1))" - -lemma natlog2_aux_correct: - "natlog2_aux n acc = acc + natlog2 n" - by (induction n acc rule: natlog2_aux.induct) (auto simp: natlog2_rec natlog2_eq_0_iff) - -lemma natlog2_code [code]: "natlog2 n = natlog2_aux n 0" - by (subst natlog2_aux_correct) simp - - -subsection \Convergence tests for infinite sums\ - -subsubsection \Root test\ - -lemma limsup_root_powser: - fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" - shows "limsup (\n. ereal (root n (norm (f n * z ^ n)))) = - limsup (\n. ereal (root n (norm (f n)))) * ereal (norm z)" -proof - - have A: "(\n. ereal (root n (norm (f n * z ^ n)))) = - (\n. ereal (root n (norm (f n))) * ereal (norm z))" (is "?g = ?h") - proof - fix n show "?g n = ?h n" - by (cases "n = 0") (simp_all add: norm_mult real_root_mult real_root_pos2 norm_power) - qed - show ?thesis by (subst A, subst limsup_ereal_mult_right) simp_all -qed - -lemma limsup_root_limit: - assumes "(\n. ereal (root n (norm (f n)))) \ l" (is "?g \ _") - shows "limsup (\n. ereal (root n (norm (f n)))) = l" -proof - - from assms have "convergent ?g" "lim ?g = l" - unfolding convergent_def by (blast intro: limI)+ - with convergent_limsup_cl show ?thesis by force -qed - -lemma limsup_root_limit': - assumes "(\n. root n (norm (f n))) \ l" - shows "limsup (\n. ereal (root n (norm (f n)))) = ereal l" - by (intro limsup_root_limit tendsto_ereal assms) - -lemma root_test_convergence': - fixes f :: "nat \ 'a :: banach" - defines "l \ limsup (\n. ereal (root n (norm (f n))))" - assumes l: "l < 1" - shows "summable f" -proof - - have "0 = limsup (\n. 0)" by (simp add: Limsup_const) - also have "... \ l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero) - finally have "l \ 0" by simp - with l obtain l' where l': "l = ereal l'" by (cases l) simp_all - - define c where "c = (1 - l') / 2" - from l and \l \ 0\ have c: "l + c > l" "l' + c \ 0" "l' + c < 1" unfolding c_def - by (simp_all add: field_simps l') - have "\C>l. eventually (\n. ereal (root n (norm (f n))) < C) sequentially" - by (subst Limsup_le_iff[symmetric]) (simp add: l_def) - with c have "eventually (\n. ereal (root n (norm (f n))) < l + ereal c) sequentially" by simp - with eventually_gt_at_top[of "0::nat"] - have "eventually (\n. norm (f n) \ (l' + c) ^ n) sequentially" - proof eventually_elim - fix n :: nat assume n: "n > 0" - assume "ereal (root n (norm (f n))) < l + ereal c" - hence "root n (norm (f n)) \ l' + c" by (simp add: l') - with c n have "root n (norm (f n)) ^ n \ (l' + c) ^ n" - by (intro power_mono) (simp_all add: real_root_ge_zero) - also from n have "root n (norm (f n)) ^ n = norm (f n)" by simp - finally show "norm (f n) \ (l' + c) ^ n" by simp - qed - thus ?thesis - by (rule summable_comparison_test_ev[OF _ summable_geometric]) (simp add: c) -qed - -lemma root_test_divergence: - fixes f :: "nat \ 'a :: banach" - defines "l \ limsup (\n. ereal (root n (norm (f n))))" - assumes l: "l > 1" - shows "\summable f" -proof - assume "summable f" - hence bounded: "Bseq f" by (simp add: summable_imp_Bseq) - - have "0 = limsup (\n. 0)" by (simp add: Limsup_const) - also have "... \ l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero) - finally have l_nonneg: "l \ 0" by simp - - define c where "c = (if l = \ then 2 else 1 + (real_of_ereal l - 1) / 2)" - from l l_nonneg consider "l = \" | "\l'. l = ereal l'" by (cases l) simp_all - hence c: "c > 1 \ ereal c < l" by cases (insert l, auto simp: c_def field_simps) - - have unbounded: "\bdd_above {n. root n (norm (f n)) > c}" - proof - assume "bdd_above {n. root n (norm (f n)) > c}" - then obtain N where "\n. root n (norm (f n)) > c \ n \ N" unfolding bdd_above_def by blast - hence "\N. \n\N. root n (norm (f n)) \ c" - by (intro exI[of _ "N + 1"]) (force simp: not_less_eq_eq[symmetric]) - hence "eventually (\n. root n (norm (f n)) \ c) sequentially" - by (auto simp: eventually_at_top_linorder) - hence "l \ c" unfolding l_def by (intro Limsup_bounded) simp_all - with c show False by auto - qed - - from bounded obtain K where K: "K > 0" "\n. norm (f n) \ K" using BseqE by blast - define n where "n = nat \log c K\" - from unbounded have "\m>n. c < root m (norm (f m))" unfolding bdd_above_def - by (auto simp: not_le) - then guess m by (elim exE conjE) note m = this - from c K have "K = c powr log c K" by (simp add: powr_def log_def) - also from c have "c powr log c K \ c powr real n" unfolding n_def - by (intro powr_mono, linarith, simp) - finally have "K \ c ^ n" using c by (simp add: powr_realpow) - also from c m have "c ^ n < c ^ m" by simp - also from c m have "c ^ m < root m (norm (f m)) ^ m" by (intro power_strict_mono) simp_all - also from m have "... = norm (f m)" by simp - finally show False using K(2)[of m] by simp -qed - - -subsubsection \Cauchy's condensation test\ - -context -fixes f :: "nat \ real" -begin - -private lemma condensation_inequality: - assumes mono: "\m n. 0 < m \ m \ n \ f n \ f m" - shows "(\k=1.. (\k=1..k=1.. (\k=1..k=1..<2^n. f (2 ^ natlog2 k)) = (\k {2^n..<(2^Suc n :: nat)}" by auto - also have "(\k\\. f (2 ^ natlog2 k)) = - (\kk = 2^n..<2^Suc n. f (2^natlog2 k))" - by (subst setsum.union_disjoint) (insert Suc, auto) - also have "natlog2 k = n" if "k \ {2^n..<2^Suc n}" for k using that by (intro natlog2_eqI) simp_all - hence "(\k = 2^n..<2^Suc n. f (2^natlog2 k)) = (\(_::nat) = 2^n..<2^Suc n. f (2^n))" - by (intro setsum.cong) simp_all - also have "\ = 2^n * f (2^n)" by (simp add: of_nat_power) - finally show ?case by simp -qed simp - -private lemma condensation_condense2: "(\k=1..<2^n. f (2 * 2 ^ natlog2 k)) = (\k {2^n..<(2^Suc n :: nat)}" by auto - also have "(\k\\. f (2 * 2 ^ natlog2 k)) = - (\kk = 2^n..<2^Suc n. f (2 * 2^natlog2 k))" - by (subst setsum.union_disjoint) (insert Suc, auto) - also have "natlog2 k = n" if "k \ {2^n..<2^Suc n}" for k using that by (intro natlog2_eqI) simp_all - hence "(\k = 2^n..<2^Suc n. f (2*2^natlog2 k)) = (\(_::nat) = 2^n..<2^Suc n. f (2^Suc n))" - by (intro setsum.cong) simp_all - also have "\ = 2^n * f (2^Suc n)" by (simp add: of_nat_power) - finally show ?case by simp -qed simp - -lemma condensation_test: - assumes mono: "\m. 0 < m \ f (Suc m) \ f m" - assumes nonneg: "\n. f n \ 0" - shows "summable f \ summable (\n. 2^n * f (2^n))" -proof - - define f' where "f' n = (if n = 0 then 0 else f n)" for n - from mono have mono': "decseq (\n. f (Suc n))" by (intro decseq_SucI) simp - hence mono': "f n \ f m" if "m \ n" "m > 0" for m n - using that decseqD[OF mono', of "m - 1" "n - 1"] by simp - - have "(\n. f (Suc n)) = (\n. f' (Suc n))" by (intro ext) (simp add: f'_def) - hence "summable f \ summable f'" - by (subst (1 2) summable_Suc_iff [symmetric]) (simp only:) - also have "\ \ convergent (\n. \kn. \kn. \k Bseq (\n. \k \ Bseq (\n. \k=1.. \ Bseq (\n. \k=1.. \ Bseq (\n. \k=1..<2^n. f k)" - by (rule nonneg_incseq_Bseq_subseq_iff[symmetric]) - (auto intro!: setsum_nonneg incseq_SucI nonneg simp: subseq_def) - also have "\ \ Bseq (\n. \kn. \k=1..<2^n. f k)" - have "eventually (\n. norm (\k norm (\k=1..<2^n. f k)) sequentially" - proof (intro always_eventually allI) - fix n :: nat - have "norm (\kk \ (\k=1..<2^n. f k)" - by (subst condensation_condense2 [symmetric]) (intro condensation_inequality mono') - also have "\ = norm \" unfolding real_norm_def - by (intro abs_of_nonneg[symmetric] setsum_nonneg ballI mult_nonneg_nonneg nonneg) - finally show "norm (\k norm (\k=1..<2^n. f k)" . - qed - from this and A have "Bseq (\n. \kn. \kn. (\k=Suc 0..n. (\k=0..n. (\kn. (\kn. norm (\k=1..<2^n. f k) \ norm (\kk=1..<2^n. f k) = (\k=1..<2^n. f k)" unfolding real_norm_def - by (intro abs_of_nonneg setsum_nonneg ballI mult_nonneg_nonneg nonneg) - also have "\ \ (\k = norm \" unfolding real_norm_def - by (intro abs_of_nonneg [symmetric] setsum_nonneg ballI mult_nonneg_nonneg nonneg) simp_all - finally show "norm (\k=1..<2^n. f k) \ norm (\kn. \k=1..<2^n. f k)" by (rule Bseq_eventually_mono) - qed - also have "monoseq (\n. (\kn. (\k convergent (\n. (\k \ summable (\k. 2^k * f (2^k))" by (simp only: summable_iff_convergent) - finally show ?thesis . -qed - -end - - -subsubsection \Summability of powers\ - -lemma abs_summable_complex_powr_iff: - "summable (\n. norm (exp (of_real (ln (of_nat n)) * s))) \ Re s < -1" -proof (cases "Re s \ 0") - let ?l = "\n. complex_of_real (ln (of_nat n))" - case False - with eventually_gt_at_top[of "0::nat"] - have "eventually (\n. norm (1 :: real) \ norm (exp (?l n * s))) sequentially" - by (auto intro!: ge_one_powr_ge_zero elim!: eventually_mono) - from summable_comparison_test_ev[OF this] False show ?thesis by (auto simp: summable_const_iff) -next - let ?l = "\n. complex_of_real (ln (of_nat n))" - case True - hence "summable (\n. norm (exp (?l n * s))) \ summable (\n. 2^n * norm (exp (?l (2^n) * s)))" - by (intro condensation_test) (auto intro!: mult_right_mono_neg) - also have "(\n. 2^n * norm (exp (?l (2^n) * s))) = (\n. (2 powr (Re s + 1)) ^ n)" - proof - fix n :: nat - have "2^n * norm (exp (?l (2^n) * s)) = exp (real n * ln 2) * exp (real n * ln 2 * Re s)" - using True by (subst exp_of_nat_mult) (simp add: ln_realpow algebra_simps) - also have "\ = exp (real n * (ln 2 * (Re s + 1)))" - by (simp add: algebra_simps exp_add) - also have "\ = exp (ln 2 * (Re s + 1)) ^ n" by (subst exp_of_nat_mult) simp - also have "exp (ln 2 * (Re s + 1)) = 2 powr (Re s + 1)" by (simp add: powr_def) - finally show "2^n * norm (exp (?l (2^n) * s)) = (2 powr (Re s + 1)) ^ n" . - qed - also have "summable \ \ 2 powr (Re s + 1) < 2 powr 0" - by (subst summable_geometric_iff) simp - also have "\ \ Re s < -1" by (subst powr_less_cancel_iff) (simp, linarith) - finally show ?thesis . -qed - -lemma summable_complex_powr_iff: - assumes "Re s < -1" - shows "summable (\n. exp (of_real (ln (of_nat n)) * s))" - by (rule summable_norm_cancel, subst abs_summable_complex_powr_iff) fact - -lemma summable_real_powr_iff: "summable (\n. of_nat n powr s :: real) \ s < -1" -proof - - from eventually_gt_at_top[of "0::nat"] - have "summable (\n. of_nat n powr s) \ summable (\n. exp (ln (of_nat n) * s))" - by (intro summable_cong) (auto elim!: eventually_mono simp: powr_def) - also have "\ \ s < -1" using abs_summable_complex_powr_iff[of "of_real s"] by simp - finally show ?thesis . -qed - -lemma inverse_power_summable: - assumes s: "s \ 2" - shows "summable (\n. inverse (of_nat n ^ s :: 'a :: {real_normed_div_algebra,banach}))" -proof (rule summable_norm_cancel, subst summable_cong) - from eventually_gt_at_top[of "0::nat"] - show "eventually (\n. norm (inverse (of_nat n ^ s:: 'a)) = real_of_nat n powr (-real s)) at_top" - by eventually_elim (simp add: norm_inverse norm_power powr_minus powr_realpow) -qed (insert s summable_real_powr_iff[of "-s"], simp_all) - -lemma not_summable_harmonic: "\summable (\n. inverse (of_nat n) :: 'a :: real_normed_field)" -proof - assume "summable (\n. inverse (of_nat n) :: 'a)" - hence "convergent (\n. norm (of_real (\kn. abs (\kn. abs (\kn. \k \ summable (\k. inverse (of_nat k) :: real)" - by (simp add: summable_iff_convergent) - finally show False using summable_real_powr_iff[of "-1"] by (simp add: powr_minus) -qed - - -subsubsection \Kummer's test\ - -lemma kummers_test_convergence: - fixes f p :: "nat \ real" - assumes pos_f: "eventually (\n. f n > 0) sequentially" - assumes nonneg_p: "eventually (\n. p n \ 0) sequentially" - defines "l \ liminf (\n. ereal (p n * f n / f (Suc n) - p (Suc n)))" - assumes l: "l > 0" - shows "summable f" - unfolding summable_iff_convergent' -proof - - define r where "r = (if l = \ then 1 else real_of_ereal l / 2)" - from l have "r > 0 \ of_real r < l" by (cases l) (simp_all add: r_def) - hence r: "r > 0" "of_real r < l" by simp_all - hence "eventually (\n. p n * f n / f (Suc n) - p (Suc n) > r) sequentially" - unfolding l_def by (force dest: less_LiminfD) - moreover from pos_f have "eventually (\n. f (Suc n) > 0) sequentially" - by (subst eventually_sequentially_Suc) - ultimately have "eventually (\n. p n * f n - p (Suc n) * f (Suc n) > r * f (Suc n)) sequentially" - by eventually_elim (simp add: field_simps) - from eventually_conj[OF pos_f eventually_conj[OF nonneg_p this]] - obtain m where m: "\n. n \ m \ f n > 0" "\n. n \ m \ p n \ 0" - "\n. n \ m \ p n * f n - p (Suc n) * f (Suc n) > r * f (Suc n)" - unfolding eventually_at_top_linorder by blast - - let ?c = "(norm (\k\m. r * f k) + p m * f m) / r" - have "Bseq (\n. (\k\n + Suc m. f k))" - proof (rule BseqI') - fix k :: nat - define n where "n = k + Suc m" - have n: "n > m" by (simp add: n_def) - - from r have "r * norm (\k\n. f k) = norm (\k\n. r * f k)" - by (simp add: setsum_right_distrib[symmetric] abs_mult) - also from n have "{..n} = {..m} \ {Suc m..n}" by auto - hence "(\k\n. r * f k) = (\k\{..m} \ {Suc m..n}. r * f k)" by (simp only:) - also have "\ = (\k\m. r * f k) + (\k=Suc m..n. r * f k)" - by (subst setsum.union_disjoint) auto - also have "norm \ \ norm (\k\m. r * f k) + norm (\k=Suc m..n. r * f k)" - by (rule norm_triangle_ineq) - also from r less_imp_le[OF m(1)] have "(\k=Suc m..n. r * f k) \ 0" - by (intro setsum_nonneg) auto - hence "norm (\k=Suc m..n. r * f k) = (\k=Suc m..n. r * f k)" by simp - also have "(\k=Suc m..n. r * f k) = (\k=m.. \ (\k=m.. = -(\k=m.. = p m * f m - p n * f n" - by (cases n, simp, simp only: atLeastLessThanSuc_atLeastAtMost, subst setsum_Suc_diff) simp_all - also from less_imp_le[OF m(1)] m(2) n have "\ \ p m * f m" by simp - finally show "norm (\k\n. f k) \ (norm (\k\m. r * f k) + p m * f m) / r" using r - by (subst pos_le_divide_eq[OF r(1)]) (simp only: mult_ac) - qed - moreover have "(\k\n. f k) \ (\k\n'. f k)" if "Suc m \ n" "n \ n'" for n n' - using less_imp_le[OF m(1)] that by (intro setsum_mono2) auto - ultimately show "convergent (\n. \k\n. f k)" by (rule Bseq_monoseq_convergent'_inc) -qed - - -lemma kummers_test_divergence: - fixes f p :: "nat \ real" - assumes pos_f: "eventually (\n. f n > 0) sequentially" - assumes pos_p: "eventually (\n. p n > 0) sequentially" - assumes divergent_p: "\summable (\n. inverse (p n))" - defines "l \ limsup (\n. ereal (p n * f n / f (Suc n) - p (Suc n)))" - assumes l: "l < 0" - shows "\summable f" -proof - assume "summable f" - from eventually_conj[OF pos_f eventually_conj[OF pos_p Limsup_lessD[OF l[unfolded l_def]]]] - obtain N where N: "\n. n \ N \ p n > 0" "\n. n \ N \ f n > 0" - "\n. n \ N \ p n * f n / f (Suc n) - p (Suc n) < 0" - by (auto simp: eventually_at_top_linorder) - hence A: "p n * f n < p (Suc n) * f (Suc n)" if "n \ N" for n using that N[of n] N[of "Suc n"] - by (simp add: field_simps) - have "p n * f n \ p N * f N" if "n \ N" for n using that and A - by (induction n rule: dec_induct) (auto intro!: less_imp_le elim!: order.trans) - from eventually_ge_at_top[of N] N this - have "eventually (\n. norm (p N * f N * inverse (p n)) \ f n) sequentially" - by (auto elim!: eventually_mono simp: field_simps abs_of_pos) - from this and \summable f\ have "summable (\n. p N * f N * inverse (p n))" - by (rule summable_comparison_test_ev) - from summable_mult[OF this, of "inverse (p N * f N)"] N[OF le_refl] - have "summable (\n. inverse (p n))" by (simp add: divide_simps) - with divergent_p show False by contradiction -qed - - -subsubsection \Ratio test\ - -lemma ratio_test_convergence: - fixes f :: "nat \ real" - assumes pos_f: "eventually (\n. f n > 0) sequentially" - defines "l \ liminf (\n. ereal (f n / f (Suc n)))" - assumes l: "l > 1" - shows "summable f" -proof (rule kummers_test_convergence[OF pos_f]) - note l - also have "l = liminf (\n. ereal (f n / f (Suc n) - 1)) + 1" - by (subst Liminf_add_ereal_right[symmetric]) (simp_all add: minus_ereal_def l_def one_ereal_def) - finally show "liminf (\n. ereal (1 * f n / f (Suc n) - 1)) > 0" - by (cases "liminf (\n. ereal (1 * f n / f (Suc n) - 1))") simp_all -qed simp - -lemma ratio_test_divergence: - fixes f :: "nat \ real" - assumes pos_f: "eventually (\n. f n > 0) sequentially" - defines "l \ limsup (\n. ereal (f n / f (Suc n)))" - assumes l: "l < 1" - shows "\summable f" -proof (rule kummers_test_divergence[OF pos_f]) - have "limsup (\n. ereal (f n / f (Suc n) - 1)) + 1 = l" - by (subst Limsup_add_ereal_right[symmetric]) (simp_all add: minus_ereal_def l_def one_ereal_def) - also note l - finally show "limsup (\n. ereal (1 * f n / f (Suc n) - 1)) < 0" - by (cases "limsup (\n. ereal (1 * f n / f (Suc n) - 1))") simp_all -qed (simp_all add: summable_const_iff) - - -subsubsection \Raabe's test\ - -lemma raabes_test_convergence: -fixes f :: "nat \ real" - assumes pos: "eventually (\n. f n > 0) sequentially" - defines "l \ liminf (\n. ereal (of_nat n * (f n / f (Suc n) - 1)))" - assumes l: "l > 1" - shows "summable f" -proof (rule kummers_test_convergence) - let ?l' = "liminf (\n. ereal (of_nat n * f n / f (Suc n) - of_nat (Suc n)))" - have "1 < l" by fact - also have "l = liminf (\n. ereal (of_nat n * f n / f (Suc n) - of_nat (Suc n)) + 1)" - by (simp add: l_def algebra_simps) - also have "\ = ?l' + 1" by (subst Liminf_add_ereal_right) simp_all - finally show "?l' > 0" by (cases ?l') (simp_all add: algebra_simps) -qed (simp_all add: pos) - -lemma raabes_test_divergence: -fixes f :: "nat \ real" - assumes pos: "eventually (\n. f n > 0) sequentially" - defines "l \ limsup (\n. ereal (of_nat n * (f n / f (Suc n) - 1)))" - assumes l: "l < 1" - shows "\summable f" -proof (rule kummers_test_divergence) - let ?l' = "limsup (\n. ereal (of_nat n * f n / f (Suc n) - of_nat (Suc n)))" - note l - also have "l = limsup (\n. ereal (of_nat n * f n / f (Suc n) - of_nat (Suc n)) + 1)" - by (simp add: l_def algebra_simps) - also have "\ = ?l' + 1" by (subst Limsup_add_ereal_right) simp_all - finally show "?l' < 0" by (cases ?l') (simp_all add: algebra_simps) -qed (insert pos eventually_gt_at_top[of "0::nat"] not_summable_harmonic, simp_all) - - - -subsection \Radius of convergence\ - -text \ - The radius of convergence of a power series. This value always exists, ranges from - @{term "0::ereal"} to @{term "\::ereal"}, and the power series is guaranteed to converge for - all inputs with a norm that is smaller than that radius and to diverge for all inputs with a - norm that is greater. -\ -definition conv_radius :: "(nat \ 'a :: banach) \ ereal" where - "conv_radius f = inverse (limsup (\n. ereal (root n (norm (f n)))))" - -lemma conv_radius_nonneg: "conv_radius f \ 0" -proof - - have "0 = limsup (\n. 0)" by (subst Limsup_const) simp_all - also have "\ \ limsup (\n. ereal (root n (norm (f n))))" - by (intro Limsup_mono) (simp_all add: real_root_ge_zero) - finally show ?thesis - unfolding conv_radius_def by (auto simp: ereal_inverse_nonneg_iff) -qed - -lemma conv_radius_zero [simp]: "conv_radius (\_. 0) = \" - by (auto simp: conv_radius_def zero_ereal_def [symmetric] Limsup_const) - -lemma conv_radius_cong: - assumes "eventually (\x. f x = g x) sequentially" - shows "conv_radius f = conv_radius g" -proof - - have "eventually (\n. ereal (root n (norm (f n))) = ereal (root n (norm (g n)))) sequentially" - using assms by eventually_elim simp - from Limsup_eq[OF this] show ?thesis unfolding conv_radius_def by simp -qed - -lemma conv_radius_altdef: - "conv_radius f = liminf (\n. inverse (ereal (root n (norm (f n)))))" - by (subst Liminf_inverse_ereal) (simp_all add: real_root_ge_zero conv_radius_def) - -lemma abs_summable_in_conv_radius: - fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" - assumes "ereal (norm z) < conv_radius f" - shows "summable (\n. norm (f n * z ^ n))" -proof (rule root_test_convergence') - define l where "l = limsup (\n. ereal (root n (norm (f n))))" - have "0 = limsup (\n. 0)" by (simp add: Limsup_const) - also have "... \ l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero) - finally have l_nonneg: "l \ 0" . - - have "limsup (\n. root n (norm (f n * z^n))) = l * ereal (norm z)" unfolding l_def - by (rule limsup_root_powser) - also from l_nonneg consider "l = 0" | "l = \" | "\l'. l = ereal l' \ l' > 0" - by (cases "l") (auto simp: less_le) - hence "l * ereal (norm z) < 1" - proof cases - assume "l = \" - hence "conv_radius f = 0" unfolding conv_radius_def l_def by simp - with assms show ?thesis by simp - next - assume "\l'. l = ereal l' \ l' > 0" - then guess l' by (elim exE conjE) note l' = this - hence "l \ \" by auto - have "l * ereal (norm z) < l * conv_radius f" - by (intro ereal_mult_strict_left_mono) (simp_all add: l' assms) - also have "conv_radius f = inverse l" by (simp add: conv_radius_def l_def) - also from l' have "l * inverse l = 1" by simp - finally show ?thesis . - qed simp_all - finally show "limsup (\n. ereal (root n (norm (norm (f n * z ^ n))))) < 1" by simp -qed - -lemma summable_in_conv_radius: - fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" - assumes "ereal (norm z) < conv_radius f" - shows "summable (\n. f n * z ^ n)" - by (rule summable_norm_cancel, rule abs_summable_in_conv_radius) fact+ - -lemma not_summable_outside_conv_radius: - fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" - assumes "ereal (norm z) > conv_radius f" - shows "\summable (\n. f n * z ^ n)" -proof (rule root_test_divergence) - define l where "l = limsup (\n. ereal (root n (norm (f n))))" - have "0 = limsup (\n. 0)" by (simp add: Limsup_const) - also have "... \ l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero) - finally have l_nonneg: "l \ 0" . - from assms have l_nz: "l \ 0" unfolding conv_radius_def l_def by auto - - have "limsup (\n. ereal (root n (norm (f n * z^n)))) = l * ereal (norm z)" - unfolding l_def by (rule limsup_root_powser) - also have "... > 1" - proof (cases l) - assume "l = \" - with assms conv_radius_nonneg[of f] show ?thesis - by (auto simp: zero_ereal_def[symmetric]) - next - fix l' assume l': "l = ereal l'" - from l_nonneg l_nz have "1 = l * inverse l" by (auto simp: l' field_simps) - also from l_nz have "inverse l = conv_radius f" - unfolding l_def conv_radius_def by auto - also from l' l_nz l_nonneg assms have "l * \ < l * ereal (norm z)" - by (intro ereal_mult_strict_left_mono) (auto simp: l') - finally show ?thesis . - qed (insert l_nonneg, simp_all) - finally show "limsup (\n. ereal (root n (norm (f n * z^n)))) > 1" . -qed - - -lemma conv_radius_geI: - assumes "summable (\n. f n * z ^ n :: 'a :: {banach, real_normed_div_algebra})" - shows "conv_radius f \ norm z" - using not_summable_outside_conv_radius[of f z] assms by (force simp: not_le[symmetric]) - -lemma conv_radius_leI: - assumes "\summable (\n. norm (f n * z ^ n :: 'a :: {banach, real_normed_div_algebra}))" - shows "conv_radius f \ norm z" - using abs_summable_in_conv_radius[of z f] assms by (force simp: not_le[symmetric]) - -lemma conv_radius_leI': - assumes "\summable (\n. f n * z ^ n :: 'a :: {banach, real_normed_div_algebra})" - shows "conv_radius f \ norm z" - using summable_in_conv_radius[of z f] assms by (force simp: not_le[symmetric]) - -lemma conv_radius_geI_ex: - fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" - assumes "\r. 0 < r \ ereal r < R \ \z. norm z = r \ summable (\n. f n * z^n)" - shows "conv_radius f \ R" -proof (rule linorder_cases[of "conv_radius f" R]) - assume R: "conv_radius f < R" - with conv_radius_nonneg[of f] obtain conv_radius' - where [simp]: "conv_radius f = ereal conv_radius'" - by (cases "conv_radius f") simp_all - define r where "r = (if R = \ then conv_radius' + 1 else (real_of_ereal R + conv_radius') / 2)" - from R conv_radius_nonneg[of f] have "0 < r \ ereal r < R \ ereal r > conv_radius f" - unfolding r_def by (cases R) (auto simp: r_def field_simps) - with assms(1)[of r] obtain z where "norm z > conv_radius f" "summable (\n. f n * z^n)" by auto - with not_summable_outside_conv_radius[of f z] show ?thesis by simp -qed simp_all - -lemma conv_radius_geI_ex': - fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" - assumes "\r. 0 < r \ ereal r < R \ summable (\n. f n * of_real r^n)" - shows "conv_radius f \ R" -proof (rule conv_radius_geI_ex) - fix r assume "0 < r" "ereal r < R" - with assms[of r] show "\z. norm z = r \ summable (\n. f n * z ^ n)" - by (intro exI[of _ "of_real r :: 'a"]) auto -qed - -lemma conv_radius_leI_ex: - fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" - assumes "R \ 0" - assumes "\r. 0 < r \ ereal r > R \ \z. norm z = r \ \summable (\n. norm (f n * z^n))" - shows "conv_radius f \ R" -proof (rule linorder_cases[of "conv_radius f" R]) - assume R: "conv_radius f > R" - from R assms(1) obtain R' where R': "R = ereal R'" by (cases R) simp_all - define r where - "r = (if conv_radius f = \ then R' + 1 else (R' + real_of_ereal (conv_radius f)) / 2)" - from R conv_radius_nonneg[of f] have "r > R \ r < conv_radius f" unfolding r_def - by (cases "conv_radius f") (auto simp: r_def field_simps R') - with assms(1) assms(2)[of r] R' - obtain z where "norm z < conv_radius f" "\summable (\n. norm (f n * z^n))" by auto - with abs_summable_in_conv_radius[of z f] show ?thesis by auto -qed simp_all - -lemma conv_radius_leI_ex': - fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" - assumes "R \ 0" - assumes "\r. 0 < r \ ereal r > R \ \summable (\n. f n * of_real r^n)" - shows "conv_radius f \ R" -proof (rule conv_radius_leI_ex) - fix r assume "0 < r" "ereal r > R" - with assms(2)[of r] show "\z. norm z = r \ \summable (\n. norm (f n * z ^ n))" - by (intro exI[of _ "of_real r :: 'a"]) (auto dest: summable_norm_cancel) -qed fact+ - -lemma conv_radius_eqI: - fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" - assumes "R \ 0" - assumes "\r. 0 < r \ ereal r < R \ \z. norm z = r \ summable (\n. f n * z^n)" - assumes "\r. 0 < r \ ereal r > R \ \z. norm z = r \ \summable (\n. norm (f n * z^n))" - shows "conv_radius f = R" - by (intro antisym conv_radius_geI_ex conv_radius_leI_ex assms) - -lemma conv_radius_eqI': - fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" - assumes "R \ 0" - assumes "\r. 0 < r \ ereal r < R \ summable (\n. f n * (of_real r)^n)" - assumes "\r. 0 < r \ ereal r > R \ \summable (\n. norm (f n * (of_real r)^n))" - shows "conv_radius f = R" -proof (intro conv_radius_eqI[OF assms(1)]) - fix r assume "0 < r" "ereal r < R" with assms(2)[OF this] - show "\z. norm z = r \ summable (\n. f n * z ^ n)" by force -next - fix r assume "0 < r" "ereal r > R" with assms(3)[OF this] - show "\z. norm z = r \ \summable (\n. norm (f n * z ^ n))" by force -qed - -lemma conv_radius_zeroI: - fixes f :: "nat \ 'a :: {banach,real_normed_div_algebra}" - assumes "\z. z \ 0 \ \summable (\n. f n * z^n)" - shows "conv_radius f = 0" -proof (rule ccontr) - assume "conv_radius f \ 0" - with conv_radius_nonneg[of f] have pos: "conv_radius f > 0" by simp - define r where "r = (if conv_radius f = \ then 1 else real_of_ereal (conv_radius f) / 2)" - from pos have r: "ereal r > 0 \ ereal r < conv_radius f" - by (cases "conv_radius f") (simp_all add: r_def) - hence "summable (\n. f n * of_real r ^ n)" by (intro summable_in_conv_radius) simp - moreover from r and assms[of "of_real r"] have "\summable (\n. f n * of_real r ^ n)" by simp - ultimately show False by contradiction -qed - -lemma conv_radius_inftyI': - fixes f :: "nat \ 'a :: {banach,real_normed_div_algebra}" - assumes "\r. r > c \ \z. norm z = r \ summable (\n. f n * z^n)" - shows "conv_radius f = \" -proof - - { - fix r :: real - have "max r (c + 1) > c" by (auto simp: max_def) - from assms[OF this] obtain z where "norm z = max r (c + 1)" "summable (\n. f n * z^n)" by blast - from conv_radius_geI[OF this(2)] this(1) have "conv_radius f \ r" by simp - } - from this[of "real_of_ereal (conv_radius f + 1)"] show "conv_radius f = \" - by (cases "conv_radius f") simp_all -qed - -lemma conv_radius_inftyI: - fixes f :: "nat \ 'a :: {banach,real_normed_div_algebra}" - assumes "\r. \z. norm z = r \ summable (\n. f n * z^n)" - shows "conv_radius f = \" - using assms by (rule conv_radius_inftyI') - -lemma conv_radius_inftyI'': - fixes f :: "nat \ 'a :: {banach,real_normed_div_algebra}" - assumes "\z. summable (\n. f n * z^n)" - shows "conv_radius f = \" -proof (rule conv_radius_inftyI') - fix r :: real assume "r > 0" - with assms show "\z. norm z = r \ summable (\n. f n * z^n)" - by (intro exI[of _ "of_real r"]) simp -qed - -lemma conv_radius_ratio_limit_ereal: - fixes f :: "nat \ 'a :: {banach,real_normed_div_algebra}" - assumes nz: "eventually (\n. f n \ 0) sequentially" - assumes lim: "(\n. ereal (norm (f n) / norm (f (Suc n)))) \ c" - shows "conv_radius f = c" -proof (rule conv_radius_eqI') - show "c \ 0" by (intro Lim_bounded2_ereal[OF lim]) simp_all -next - fix r assume r: "0 < r" "ereal r < c" - let ?l = "liminf (\n. ereal (norm (f n * of_real r ^ n) / norm (f (Suc n) * of_real r ^ Suc n)))" - have "?l = liminf (\n. ereal (norm (f n) / (norm (f (Suc n)))) * ereal (inverse r))" - using r by (simp add: norm_mult norm_power divide_simps) - also from r have "\ = liminf (\n. ereal (norm (f n) / (norm (f (Suc n))))) * ereal (inverse r)" - by (intro Liminf_ereal_mult_right) simp_all - also have "liminf (\n. ereal (norm (f n) / (norm (f (Suc n))))) = c" - by (intro lim_imp_Liminf lim) simp - finally have l: "?l = c * ereal (inverse r)" by simp - from r have l': "c * ereal (inverse r) > 1" by (cases c) (simp_all add: field_simps) - show "summable (\n. f n * of_real r^n)" - by (rule summable_norm_cancel, rule ratio_test_convergence) - (insert r nz l l', auto elim!: eventually_mono) -next - fix r assume r: "0 < r" "ereal r > c" - let ?l = "limsup (\n. ereal (norm (f n * of_real r ^ n) / norm (f (Suc n) * of_real r ^ Suc n)))" - have "?l = limsup (\n. ereal (norm (f n) / (norm (f (Suc n)))) * ereal (inverse r))" - using r by (simp add: norm_mult norm_power divide_simps) - also from r have "\ = limsup (\n. ereal (norm (f n) / (norm (f (Suc n))))) * ereal (inverse r)" - by (intro Limsup_ereal_mult_right) simp_all - also have "limsup (\n. ereal (norm (f n) / (norm (f (Suc n))))) = c" - by (intro lim_imp_Limsup lim) simp - finally have l: "?l = c * ereal (inverse r)" by simp - from r have l': "c * ereal (inverse r) < 1" by (cases c) (simp_all add: field_simps) - show "\summable (\n. norm (f n * of_real r^n))" - by (rule ratio_test_divergence) (insert r nz l l', auto elim!: eventually_mono) -qed - -lemma conv_radius_ratio_limit_ereal_nonzero: - fixes f :: "nat \ 'a :: {banach,real_normed_div_algebra}" - assumes nz: "c \ 0" - assumes lim: "(\n. ereal (norm (f n) / norm (f (Suc n)))) \ c" - shows "conv_radius f = c" -proof (rule conv_radius_ratio_limit_ereal[OF _ lim], rule ccontr) - assume "\eventually (\n. f n \ 0) sequentially" - hence "frequently (\n. f n = 0) sequentially" by (simp add: frequently_def) - hence "frequently (\n. ereal (norm (f n) / norm (f (Suc n))) = 0) sequentially" - by (force elim!: frequently_elim1) - hence "c = 0" by (intro limit_frequently_eq[OF _ _ lim]) auto - with nz show False by contradiction -qed - -lemma conv_radius_ratio_limit: - fixes f :: "nat \ 'a :: {banach,real_normed_div_algebra}" - assumes "c' = ereal c" - assumes nz: "eventually (\n. f n \ 0) sequentially" - assumes lim: "(\n. norm (f n) / norm (f (Suc n))) \ c" - shows "conv_radius f = c'" - using assms by (intro conv_radius_ratio_limit_ereal) simp_all - -lemma conv_radius_ratio_limit_nonzero: - fixes f :: "nat \ 'a :: {banach,real_normed_div_algebra}" - assumes "c' = ereal c" - assumes nz: "c \ 0" - assumes lim: "(\n. norm (f n) / norm (f (Suc n))) \ c" - shows "conv_radius f = c'" - using assms by (intro conv_radius_ratio_limit_ereal_nonzero) simp_all - -lemma conv_radius_mult_power: - assumes "c \ (0 :: 'a :: {real_normed_div_algebra,banach})" - shows "conv_radius (\n. c ^ n * f n) = conv_radius f / ereal (norm c)" -proof - - have "limsup (\n. ereal (root n (norm (c ^ n * f n)))) = - limsup (\n. ereal (norm c) * ereal (root n (norm (f n))))" - using eventually_gt_at_top[of "0::nat"] - by (intro Limsup_eq) - (auto elim!: eventually_mono simp: norm_mult norm_power real_root_mult real_root_power) - also have "\ = ereal (norm c) * limsup (\n. ereal (root n (norm (f n))))" - using assms by (subst Limsup_ereal_mult_left[symmetric]) simp_all - finally have A: "limsup (\n. ereal (root n (norm (c ^ n * f n)))) = - ereal (norm c) * limsup (\n. ereal (root n (norm (f n))))" . - show ?thesis using assms - apply (cases "limsup (\n. ereal (root n (norm (f n)))) = 0") - apply (simp add: A conv_radius_def) - apply (unfold conv_radius_def A divide_ereal_def, simp add: mult.commute ereal_inverse_mult) - done -qed - -lemma conv_radius_mult_power_right: - assumes "c \ (0 :: 'a :: {real_normed_div_algebra,banach})" - shows "conv_radius (\n. f n * c ^ n) = conv_radius f / ereal (norm c)" - using conv_radius_mult_power[OF assms, of f] - unfolding conv_radius_def by (simp add: mult.commute norm_mult) - -lemma conv_radius_divide_power: - assumes "c \ (0 :: 'a :: {real_normed_div_algebra,banach})" - shows "conv_radius (\n. f n / c^n) = conv_radius f * ereal (norm c)" -proof - - from assms have "inverse c \ 0" by simp - from conv_radius_mult_power_right[OF this, of f] show ?thesis - by (simp add: divide_inverse divide_ereal_def assms norm_inverse power_inverse) -qed - - -lemma conv_radius_add_ge: - "min (conv_radius f) (conv_radius g) \ - conv_radius (\x. f x + g x :: 'a :: {banach,real_normed_div_algebra})" - by (rule conv_radius_geI_ex') - (auto simp: algebra_simps intro!: summable_add summable_in_conv_radius) - -lemma conv_radius_mult_ge: - fixes f g :: "nat \ ('a :: {banach,real_normed_div_algebra})" - shows "conv_radius (\x. \i\x. f i * g (x - i)) \ min (conv_radius f) (conv_radius g)" -proof (rule conv_radius_geI_ex') - fix r assume r: "r > 0" "ereal r < min (conv_radius f) (conv_radius g)" - from r have "summable (\n. (\i\n. (f i * of_real r^i) * (g (n - i) * of_real r^(n - i))))" - by (intro summable_Cauchy_product abs_summable_in_conv_radius) simp_all - thus "summable (\n. (\i\n. f i * g (n - i)) * of_real r ^ n)" - by (simp add: algebra_simps of_real_def power_add [symmetric] scaleR_setsum_right) -qed - -lemma le_conv_radius_iff: - fixes a :: "nat \ 'a::{real_normed_div_algebra,banach}" - shows "r \ conv_radius a \ (\x. norm (x-\) < r \ summable (\i. a i * (x - \) ^ i))" -apply (intro iffI allI impI summable_in_conv_radius conv_radius_geI_ex) -apply (meson less_ereal.simps(1) not_le order_trans) -apply (rule_tac x="of_real ra" in exI, simp) -apply (metis abs_of_nonneg add_diff_cancel_left' less_eq_real_def norm_of_real) -done - -end diff -r bbcb05504fdc -r bd218a9320b5 src/HOL/Multivariate_Analysis/Summation_Tests.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Multivariate_Analysis/Summation_Tests.thy Thu Aug 04 19:36:31 2016 +0200 @@ -0,0 +1,920 @@ +(* Title: HOL/Multivariate_Analysis/Summation.thy + Author: Manuel Eberl, TU München +*) + +section \Radius of Convergence and Summation Tests\ + +theory Summation_Tests +imports + Complex_Main + "~~/src/HOL/Library/Extended_Real" + "~~/src/HOL/Library/Liminf_Limsup" +begin + +text \ + The definition of the radius of convergence of a power series, + various summability tests, lemmas to compute the radius of convergence etc. +\ + +subsection \Rounded dual logarithm\ + +(* This is required for the Cauchy condensation criterion *) + +definition "natlog2 n = (if n = 0 then 0 else nat \log 2 (real_of_nat n)\)" + +lemma natlog2_0 [simp]: "natlog2 0 = 0" by (simp add: natlog2_def) +lemma natlog2_1 [simp]: "natlog2 1 = 0" by (simp add: natlog2_def) +lemma natlog2_eq_0_iff: "natlog2 n = 0 \ n < 2" by (simp add: natlog2_def) + +lemma natlog2_power_of_two [simp]: "natlog2 (2 ^ n) = n" + by (simp add: natlog2_def log_nat_power) + +lemma natlog2_mono: "m \ n \ natlog2 m \ natlog2 n" + unfolding natlog2_def by (simp_all add: nat_mono floor_mono) + +lemma pow_natlog2_le: "n > 0 \ 2 ^ natlog2 n \ n" +proof - + assume n: "n > 0" + from n have "of_nat (2 ^ natlog2 n) = 2 powr real_of_nat (nat \log 2 (real_of_nat n)\)" + by (subst powr_realpow) (simp_all add: natlog2_def) + also have "\ = 2 powr of_int \log 2 (real_of_nat n)\" using n by simp + also have "\ \ 2 powr log 2 (real_of_nat n)" by (intro powr_mono) (linarith, simp_all) + also have "\ = of_nat n" using n by simp + finally show ?thesis by simp +qed + +lemma pow_natlog2_gt: "n > 0 \ 2 * 2 ^ natlog2 n > n" + and pow_natlog2_ge: "n > 0 \ 2 * 2 ^ natlog2 n \ n" +proof - + assume n: "n > 0" + from n have "of_nat n = 2 powr log 2 (real_of_nat n)" by simp + also have "\ < 2 powr (1 + of_int \log 2 (real_of_nat n)\)" + by (intro powr_less_mono) (linarith, simp_all) + also from n have "\ = 2 powr (1 + real_of_nat (nat \log 2 (real_of_nat n)\))" by simp + also from n have "\ = of_nat (2 * 2 ^ natlog2 n)" + by (simp_all add: natlog2_def powr_real_of_int powr_add) + finally show "2 * 2 ^ natlog2 n > n" by (rule of_nat_less_imp_less) + thus "2 * 2 ^ natlog2 n \ n" by simp +qed + +lemma natlog2_eqI: + assumes "n > 0" "2^k \ n" "n < 2 * 2^k" + shows "natlog2 n = k" +proof - + from assms have "of_nat (2 ^ k) \ real_of_nat n" by (subst of_nat_le_iff) simp_all + hence "real_of_int (int k) \ log (of_nat 2) (real_of_nat n)" + by (subst le_log_iff) (simp_all add: powr_realpow assms del: of_nat_le_iff) + moreover from assms have "real_of_nat n < of_nat (2 ^ Suc k)" by (subst of_nat_less_iff) simp_all + hence "log 2 (real_of_nat n) < of_nat k + 1" + by (subst log_less_iff) (simp_all add: assms powr_realpow powr_add) + ultimately have "\log 2 (real_of_nat n)\ = of_nat k" by (intro floor_unique) simp_all + with assms show ?thesis by (simp add: natlog2_def) +qed + +lemma natlog2_rec: + assumes "n \ 2" + shows "natlog2 n = 1 + natlog2 (n div 2)" +proof (rule natlog2_eqI) + from assms have "2 ^ (1 + natlog2 (n div 2)) \ 2 * (n div 2)" + by (simp add: pow_natlog2_le) + also have "\ \ n" by simp + finally show "2 ^ (1 + natlog2 (n div 2)) \ n" . +next + from assms have "n < 2 * (n div 2 + 1)" by simp + also from assms have "(n div 2) < 2 ^ (1 + natlog2 (n div 2))" + by (simp add: pow_natlog2_gt) + hence "2 * (n div 2 + 1) \ 2 * (2 ^ (1 + natlog2 (n div 2)))" + by (intro mult_left_mono) simp_all + finally show "n < 2 * 2 ^ (1 + natlog2 (n div 2))" . +qed (insert assms, simp_all) + +fun natlog2_aux where + "natlog2_aux n acc = (if (n::nat) < 2 then acc else natlog2_aux (n div 2) (acc + 1))" + +lemma natlog2_aux_correct: + "natlog2_aux n acc = acc + natlog2 n" + by (induction n acc rule: natlog2_aux.induct) (auto simp: natlog2_rec natlog2_eq_0_iff) + +lemma natlog2_code [code]: "natlog2 n = natlog2_aux n 0" + by (subst natlog2_aux_correct) simp + + +subsection \Convergence tests for infinite sums\ + +subsubsection \Root test\ + +lemma limsup_root_powser: + fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" + shows "limsup (\n. ereal (root n (norm (f n * z ^ n)))) = + limsup (\n. ereal (root n (norm (f n)))) * ereal (norm z)" +proof - + have A: "(\n. ereal (root n (norm (f n * z ^ n)))) = + (\n. ereal (root n (norm (f n))) * ereal (norm z))" (is "?g = ?h") + proof + fix n show "?g n = ?h n" + by (cases "n = 0") (simp_all add: norm_mult real_root_mult real_root_pos2 norm_power) + qed + show ?thesis by (subst A, subst limsup_ereal_mult_right) simp_all +qed + +lemma limsup_root_limit: + assumes "(\n. ereal (root n (norm (f n)))) \ l" (is "?g \ _") + shows "limsup (\n. ereal (root n (norm (f n)))) = l" +proof - + from assms have "convergent ?g" "lim ?g = l" + unfolding convergent_def by (blast intro: limI)+ + with convergent_limsup_cl show ?thesis by force +qed + +lemma limsup_root_limit': + assumes "(\n. root n (norm (f n))) \ l" + shows "limsup (\n. ereal (root n (norm (f n)))) = ereal l" + by (intro limsup_root_limit tendsto_ereal assms) + +lemma root_test_convergence': + fixes f :: "nat \ 'a :: banach" + defines "l \ limsup (\n. ereal (root n (norm (f n))))" + assumes l: "l < 1" + shows "summable f" +proof - + have "0 = limsup (\n. 0)" by (simp add: Limsup_const) + also have "... \ l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero) + finally have "l \ 0" by simp + with l obtain l' where l': "l = ereal l'" by (cases l) simp_all + + define c where "c = (1 - l') / 2" + from l and \l \ 0\ have c: "l + c > l" "l' + c \ 0" "l' + c < 1" unfolding c_def + by (simp_all add: field_simps l') + have "\C>l. eventually (\n. ereal (root n (norm (f n))) < C) sequentially" + by (subst Limsup_le_iff[symmetric]) (simp add: l_def) + with c have "eventually (\n. ereal (root n (norm (f n))) < l + ereal c) sequentially" by simp + with eventually_gt_at_top[of "0::nat"] + have "eventually (\n. norm (f n) \ (l' + c) ^ n) sequentially" + proof eventually_elim + fix n :: nat assume n: "n > 0" + assume "ereal (root n (norm (f n))) < l + ereal c" + hence "root n (norm (f n)) \ l' + c" by (simp add: l') + with c n have "root n (norm (f n)) ^ n \ (l' + c) ^ n" + by (intro power_mono) (simp_all add: real_root_ge_zero) + also from n have "root n (norm (f n)) ^ n = norm (f n)" by simp + finally show "norm (f n) \ (l' + c) ^ n" by simp + qed + thus ?thesis + by (rule summable_comparison_test_ev[OF _ summable_geometric]) (simp add: c) +qed + +lemma root_test_divergence: + fixes f :: "nat \ 'a :: banach" + defines "l \ limsup (\n. ereal (root n (norm (f n))))" + assumes l: "l > 1" + shows "\summable f" +proof + assume "summable f" + hence bounded: "Bseq f" by (simp add: summable_imp_Bseq) + + have "0 = limsup (\n. 0)" by (simp add: Limsup_const) + also have "... \ l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero) + finally have l_nonneg: "l \ 0" by simp + + define c where "c = (if l = \ then 2 else 1 + (real_of_ereal l - 1) / 2)" + from l l_nonneg consider "l = \" | "\l'. l = ereal l'" by (cases l) simp_all + hence c: "c > 1 \ ereal c < l" by cases (insert l, auto simp: c_def field_simps) + + have unbounded: "\bdd_above {n. root n (norm (f n)) > c}" + proof + assume "bdd_above {n. root n (norm (f n)) > c}" + then obtain N where "\n. root n (norm (f n)) > c \ n \ N" unfolding bdd_above_def by blast + hence "\N. \n\N. root n (norm (f n)) \ c" + by (intro exI[of _ "N + 1"]) (force simp: not_less_eq_eq[symmetric]) + hence "eventually (\n. root n (norm (f n)) \ c) sequentially" + by (auto simp: eventually_at_top_linorder) + hence "l \ c" unfolding l_def by (intro Limsup_bounded) simp_all + with c show False by auto + qed + + from bounded obtain K where K: "K > 0" "\n. norm (f n) \ K" using BseqE by blast + define n where "n = nat \log c K\" + from unbounded have "\m>n. c < root m (norm (f m))" unfolding bdd_above_def + by (auto simp: not_le) + then guess m by (elim exE conjE) note m = this + from c K have "K = c powr log c K" by (simp add: powr_def log_def) + also from c have "c powr log c K \ c powr real n" unfolding n_def + by (intro powr_mono, linarith, simp) + finally have "K \ c ^ n" using c by (simp add: powr_realpow) + also from c m have "c ^ n < c ^ m" by simp + also from c m have "c ^ m < root m (norm (f m)) ^ m" by (intro power_strict_mono) simp_all + also from m have "... = norm (f m)" by simp + finally show False using K(2)[of m] by simp +qed + + +subsubsection \Cauchy's condensation test\ + +context +fixes f :: "nat \ real" +begin + +private lemma condensation_inequality: + assumes mono: "\m n. 0 < m \ m \ n \ f n \ f m" + shows "(\k=1.. (\k=1..k=1.. (\k=1..k=1..<2^n. f (2 ^ natlog2 k)) = (\k {2^n..<(2^Suc n :: nat)}" by auto + also have "(\k\\. f (2 ^ natlog2 k)) = + (\kk = 2^n..<2^Suc n. f (2^natlog2 k))" + by (subst setsum.union_disjoint) (insert Suc, auto) + also have "natlog2 k = n" if "k \ {2^n..<2^Suc n}" for k using that by (intro natlog2_eqI) simp_all + hence "(\k = 2^n..<2^Suc n. f (2^natlog2 k)) = (\(_::nat) = 2^n..<2^Suc n. f (2^n))" + by (intro setsum.cong) simp_all + also have "\ = 2^n * f (2^n)" by (simp add: of_nat_power) + finally show ?case by simp +qed simp + +private lemma condensation_condense2: "(\k=1..<2^n. f (2 * 2 ^ natlog2 k)) = (\k {2^n..<(2^Suc n :: nat)}" by auto + also have "(\k\\. f (2 * 2 ^ natlog2 k)) = + (\kk = 2^n..<2^Suc n. f (2 * 2^natlog2 k))" + by (subst setsum.union_disjoint) (insert Suc, auto) + also have "natlog2 k = n" if "k \ {2^n..<2^Suc n}" for k using that by (intro natlog2_eqI) simp_all + hence "(\k = 2^n..<2^Suc n. f (2*2^natlog2 k)) = (\(_::nat) = 2^n..<2^Suc n. f (2^Suc n))" + by (intro setsum.cong) simp_all + also have "\ = 2^n * f (2^Suc n)" by (simp add: of_nat_power) + finally show ?case by simp +qed simp + +lemma condensation_test: + assumes mono: "\m. 0 < m \ f (Suc m) \ f m" + assumes nonneg: "\n. f n \ 0" + shows "summable f \ summable (\n. 2^n * f (2^n))" +proof - + define f' where "f' n = (if n = 0 then 0 else f n)" for n + from mono have mono': "decseq (\n. f (Suc n))" by (intro decseq_SucI) simp + hence mono': "f n \ f m" if "m \ n" "m > 0" for m n + using that decseqD[OF mono', of "m - 1" "n - 1"] by simp + + have "(\n. f (Suc n)) = (\n. f' (Suc n))" by (intro ext) (simp add: f'_def) + hence "summable f \ summable f'" + by (subst (1 2) summable_Suc_iff [symmetric]) (simp only:) + also have "\ \ convergent (\n. \kn. \kn. \k Bseq (\n. \k \ Bseq (\n. \k=1.. \ Bseq (\n. \k=1.. \ Bseq (\n. \k=1..<2^n. f k)" + by (rule nonneg_incseq_Bseq_subseq_iff[symmetric]) + (auto intro!: setsum_nonneg incseq_SucI nonneg simp: subseq_def) + also have "\ \ Bseq (\n. \kn. \k=1..<2^n. f k)" + have "eventually (\n. norm (\k norm (\k=1..<2^n. f k)) sequentially" + proof (intro always_eventually allI) + fix n :: nat + have "norm (\kk \ (\k=1..<2^n. f k)" + by (subst condensation_condense2 [symmetric]) (intro condensation_inequality mono') + also have "\ = norm \" unfolding real_norm_def + by (intro abs_of_nonneg[symmetric] setsum_nonneg ballI mult_nonneg_nonneg nonneg) + finally show "norm (\k norm (\k=1..<2^n. f k)" . + qed + from this and A have "Bseq (\n. \kn. \kn. (\k=Suc 0..n. (\k=0..n. (\kn. (\kn. norm (\k=1..<2^n. f k) \ norm (\kk=1..<2^n. f k) = (\k=1..<2^n. f k)" unfolding real_norm_def + by (intro abs_of_nonneg setsum_nonneg ballI mult_nonneg_nonneg nonneg) + also have "\ \ (\k = norm \" unfolding real_norm_def + by (intro abs_of_nonneg [symmetric] setsum_nonneg ballI mult_nonneg_nonneg nonneg) simp_all + finally show "norm (\k=1..<2^n. f k) \ norm (\kn. \k=1..<2^n. f k)" by (rule Bseq_eventually_mono) + qed + also have "monoseq (\n. (\kn. (\k convergent (\n. (\k \ summable (\k. 2^k * f (2^k))" by (simp only: summable_iff_convergent) + finally show ?thesis . +qed + +end + + +subsubsection \Summability of powers\ + +lemma abs_summable_complex_powr_iff: + "summable (\n. norm (exp (of_real (ln (of_nat n)) * s))) \ Re s < -1" +proof (cases "Re s \ 0") + let ?l = "\n. complex_of_real (ln (of_nat n))" + case False + with eventually_gt_at_top[of "0::nat"] + have "eventually (\n. norm (1 :: real) \ norm (exp (?l n * s))) sequentially" + by (auto intro!: ge_one_powr_ge_zero elim!: eventually_mono) + from summable_comparison_test_ev[OF this] False show ?thesis by (auto simp: summable_const_iff) +next + let ?l = "\n. complex_of_real (ln (of_nat n))" + case True + hence "summable (\n. norm (exp (?l n * s))) \ summable (\n. 2^n * norm (exp (?l (2^n) * s)))" + by (intro condensation_test) (auto intro!: mult_right_mono_neg) + also have "(\n. 2^n * norm (exp (?l (2^n) * s))) = (\n. (2 powr (Re s + 1)) ^ n)" + proof + fix n :: nat + have "2^n * norm (exp (?l (2^n) * s)) = exp (real n * ln 2) * exp (real n * ln 2 * Re s)" + using True by (subst exp_of_nat_mult) (simp add: ln_realpow algebra_simps) + also have "\ = exp (real n * (ln 2 * (Re s + 1)))" + by (simp add: algebra_simps exp_add) + also have "\ = exp (ln 2 * (Re s + 1)) ^ n" by (subst exp_of_nat_mult) simp + also have "exp (ln 2 * (Re s + 1)) = 2 powr (Re s + 1)" by (simp add: powr_def) + finally show "2^n * norm (exp (?l (2^n) * s)) = (2 powr (Re s + 1)) ^ n" . + qed + also have "summable \ \ 2 powr (Re s + 1) < 2 powr 0" + by (subst summable_geometric_iff) simp + also have "\ \ Re s < -1" by (subst powr_less_cancel_iff) (simp, linarith) + finally show ?thesis . +qed + +lemma summable_complex_powr_iff: + assumes "Re s < -1" + shows "summable (\n. exp (of_real (ln (of_nat n)) * s))" + by (rule summable_norm_cancel, subst abs_summable_complex_powr_iff) fact + +lemma summable_real_powr_iff: "summable (\n. of_nat n powr s :: real) \ s < -1" +proof - + from eventually_gt_at_top[of "0::nat"] + have "summable (\n. of_nat n powr s) \ summable (\n. exp (ln (of_nat n) * s))" + by (intro summable_cong) (auto elim!: eventually_mono simp: powr_def) + also have "\ \ s < -1" using abs_summable_complex_powr_iff[of "of_real s"] by simp + finally show ?thesis . +qed + +lemma inverse_power_summable: + assumes s: "s \ 2" + shows "summable (\n. inverse (of_nat n ^ s :: 'a :: {real_normed_div_algebra,banach}))" +proof (rule summable_norm_cancel, subst summable_cong) + from eventually_gt_at_top[of "0::nat"] + show "eventually (\n. norm (inverse (of_nat n ^ s:: 'a)) = real_of_nat n powr (-real s)) at_top" + by eventually_elim (simp add: norm_inverse norm_power powr_minus powr_realpow) +qed (insert s summable_real_powr_iff[of "-s"], simp_all) + +lemma not_summable_harmonic: "\summable (\n. inverse (of_nat n) :: 'a :: real_normed_field)" +proof + assume "summable (\n. inverse (of_nat n) :: 'a)" + hence "convergent (\n. norm (of_real (\kn. abs (\kn. abs (\kn. \k \ summable (\k. inverse (of_nat k) :: real)" + by (simp add: summable_iff_convergent) + finally show False using summable_real_powr_iff[of "-1"] by (simp add: powr_minus) +qed + + +subsubsection \Kummer's test\ + +lemma kummers_test_convergence: + fixes f p :: "nat \ real" + assumes pos_f: "eventually (\n. f n > 0) sequentially" + assumes nonneg_p: "eventually (\n. p n \ 0) sequentially" + defines "l \ liminf (\n. ereal (p n * f n / f (Suc n) - p (Suc n)))" + assumes l: "l > 0" + shows "summable f" + unfolding summable_iff_convergent' +proof - + define r where "r = (if l = \ then 1 else real_of_ereal l / 2)" + from l have "r > 0 \ of_real r < l" by (cases l) (simp_all add: r_def) + hence r: "r > 0" "of_real r < l" by simp_all + hence "eventually (\n. p n * f n / f (Suc n) - p (Suc n) > r) sequentially" + unfolding l_def by (force dest: less_LiminfD) + moreover from pos_f have "eventually (\n. f (Suc n) > 0) sequentially" + by (subst eventually_sequentially_Suc) + ultimately have "eventually (\n. p n * f n - p (Suc n) * f (Suc n) > r * f (Suc n)) sequentially" + by eventually_elim (simp add: field_simps) + from eventually_conj[OF pos_f eventually_conj[OF nonneg_p this]] + obtain m where m: "\n. n \ m \ f n > 0" "\n. n \ m \ p n \ 0" + "\n. n \ m \ p n * f n - p (Suc n) * f (Suc n) > r * f (Suc n)" + unfolding eventually_at_top_linorder by blast + + let ?c = "(norm (\k\m. r * f k) + p m * f m) / r" + have "Bseq (\n. (\k\n + Suc m. f k))" + proof (rule BseqI') + fix k :: nat + define n where "n = k + Suc m" + have n: "n > m" by (simp add: n_def) + + from r have "r * norm (\k\n. f k) = norm (\k\n. r * f k)" + by (simp add: setsum_right_distrib[symmetric] abs_mult) + also from n have "{..n} = {..m} \ {Suc m..n}" by auto + hence "(\k\n. r * f k) = (\k\{..m} \ {Suc m..n}. r * f k)" by (simp only:) + also have "\ = (\k\m. r * f k) + (\k=Suc m..n. r * f k)" + by (subst setsum.union_disjoint) auto + also have "norm \ \ norm (\k\m. r * f k) + norm (\k=Suc m..n. r * f k)" + by (rule norm_triangle_ineq) + also from r less_imp_le[OF m(1)] have "(\k=Suc m..n. r * f k) \ 0" + by (intro setsum_nonneg) auto + hence "norm (\k=Suc m..n. r * f k) = (\k=Suc m..n. r * f k)" by simp + also have "(\k=Suc m..n. r * f k) = (\k=m.. \ (\k=m.. = -(\k=m.. = p m * f m - p n * f n" + by (cases n, simp, simp only: atLeastLessThanSuc_atLeastAtMost, subst setsum_Suc_diff) simp_all + also from less_imp_le[OF m(1)] m(2) n have "\ \ p m * f m" by simp + finally show "norm (\k\n. f k) \ (norm (\k\m. r * f k) + p m * f m) / r" using r + by (subst pos_le_divide_eq[OF r(1)]) (simp only: mult_ac) + qed + moreover have "(\k\n. f k) \ (\k\n'. f k)" if "Suc m \ n" "n \ n'" for n n' + using less_imp_le[OF m(1)] that by (intro setsum_mono2) auto + ultimately show "convergent (\n. \k\n. f k)" by (rule Bseq_monoseq_convergent'_inc) +qed + + +lemma kummers_test_divergence: + fixes f p :: "nat \ real" + assumes pos_f: "eventually (\n. f n > 0) sequentially" + assumes pos_p: "eventually (\n. p n > 0) sequentially" + assumes divergent_p: "\summable (\n. inverse (p n))" + defines "l \ limsup (\n. ereal (p n * f n / f (Suc n) - p (Suc n)))" + assumes l: "l < 0" + shows "\summable f" +proof + assume "summable f" + from eventually_conj[OF pos_f eventually_conj[OF pos_p Limsup_lessD[OF l[unfolded l_def]]]] + obtain N where N: "\n. n \ N \ p n > 0" "\n. n \ N \ f n > 0" + "\n. n \ N \ p n * f n / f (Suc n) - p (Suc n) < 0" + by (auto simp: eventually_at_top_linorder) + hence A: "p n * f n < p (Suc n) * f (Suc n)" if "n \ N" for n using that N[of n] N[of "Suc n"] + by (simp add: field_simps) + have "p n * f n \ p N * f N" if "n \ N" for n using that and A + by (induction n rule: dec_induct) (auto intro!: less_imp_le elim!: order.trans) + from eventually_ge_at_top[of N] N this + have "eventually (\n. norm (p N * f N * inverse (p n)) \ f n) sequentially" + by (auto elim!: eventually_mono simp: field_simps abs_of_pos) + from this and \summable f\ have "summable (\n. p N * f N * inverse (p n))" + by (rule summable_comparison_test_ev) + from summable_mult[OF this, of "inverse (p N * f N)"] N[OF le_refl] + have "summable (\n. inverse (p n))" by (simp add: divide_simps) + with divergent_p show False by contradiction +qed + + +subsubsection \Ratio test\ + +lemma ratio_test_convergence: + fixes f :: "nat \ real" + assumes pos_f: "eventually (\n. f n > 0) sequentially" + defines "l \ liminf (\n. ereal (f n / f (Suc n)))" + assumes l: "l > 1" + shows "summable f" +proof (rule kummers_test_convergence[OF pos_f]) + note l + also have "l = liminf (\n. ereal (f n / f (Suc n) - 1)) + 1" + by (subst Liminf_add_ereal_right[symmetric]) (simp_all add: minus_ereal_def l_def one_ereal_def) + finally show "liminf (\n. ereal (1 * f n / f (Suc n) - 1)) > 0" + by (cases "liminf (\n. ereal (1 * f n / f (Suc n) - 1))") simp_all +qed simp + +lemma ratio_test_divergence: + fixes f :: "nat \ real" + assumes pos_f: "eventually (\n. f n > 0) sequentially" + defines "l \ limsup (\n. ereal (f n / f (Suc n)))" + assumes l: "l < 1" + shows "\summable f" +proof (rule kummers_test_divergence[OF pos_f]) + have "limsup (\n. ereal (f n / f (Suc n) - 1)) + 1 = l" + by (subst Limsup_add_ereal_right[symmetric]) (simp_all add: minus_ereal_def l_def one_ereal_def) + also note l + finally show "limsup (\n. ereal (1 * f n / f (Suc n) - 1)) < 0" + by (cases "limsup (\n. ereal (1 * f n / f (Suc n) - 1))") simp_all +qed (simp_all add: summable_const_iff) + + +subsubsection \Raabe's test\ + +lemma raabes_test_convergence: +fixes f :: "nat \ real" + assumes pos: "eventually (\n. f n > 0) sequentially" + defines "l \ liminf (\n. ereal (of_nat n * (f n / f (Suc n) - 1)))" + assumes l: "l > 1" + shows "summable f" +proof (rule kummers_test_convergence) + let ?l' = "liminf (\n. ereal (of_nat n * f n / f (Suc n) - of_nat (Suc n)))" + have "1 < l" by fact + also have "l = liminf (\n. ereal (of_nat n * f n / f (Suc n) - of_nat (Suc n)) + 1)" + by (simp add: l_def algebra_simps) + also have "\ = ?l' + 1" by (subst Liminf_add_ereal_right) simp_all + finally show "?l' > 0" by (cases ?l') (simp_all add: algebra_simps) +qed (simp_all add: pos) + +lemma raabes_test_divergence: +fixes f :: "nat \ real" + assumes pos: "eventually (\n. f n > 0) sequentially" + defines "l \ limsup (\n. ereal (of_nat n * (f n / f (Suc n) - 1)))" + assumes l: "l < 1" + shows "\summable f" +proof (rule kummers_test_divergence) + let ?l' = "limsup (\n. ereal (of_nat n * f n / f (Suc n) - of_nat (Suc n)))" + note l + also have "l = limsup (\n. ereal (of_nat n * f n / f (Suc n) - of_nat (Suc n)) + 1)" + by (simp add: l_def algebra_simps) + also have "\ = ?l' + 1" by (subst Limsup_add_ereal_right) simp_all + finally show "?l' < 0" by (cases ?l') (simp_all add: algebra_simps) +qed (insert pos eventually_gt_at_top[of "0::nat"] not_summable_harmonic, simp_all) + + + +subsection \Radius of convergence\ + +text \ + The radius of convergence of a power series. This value always exists, ranges from + @{term "0::ereal"} to @{term "\::ereal"}, and the power series is guaranteed to converge for + all inputs with a norm that is smaller than that radius and to diverge for all inputs with a + norm that is greater. +\ +definition conv_radius :: "(nat \ 'a :: banach) \ ereal" where + "conv_radius f = inverse (limsup (\n. ereal (root n (norm (f n)))))" + +lemma conv_radius_nonneg: "conv_radius f \ 0" +proof - + have "0 = limsup (\n. 0)" by (subst Limsup_const) simp_all + also have "\ \ limsup (\n. ereal (root n (norm (f n))))" + by (intro Limsup_mono) (simp_all add: real_root_ge_zero) + finally show ?thesis + unfolding conv_radius_def by (auto simp: ereal_inverse_nonneg_iff) +qed + +lemma conv_radius_zero [simp]: "conv_radius (\_. 0) = \" + by (auto simp: conv_radius_def zero_ereal_def [symmetric] Limsup_const) + +lemma conv_radius_cong: + assumes "eventually (\x. f x = g x) sequentially" + shows "conv_radius f = conv_radius g" +proof - + have "eventually (\n. ereal (root n (norm (f n))) = ereal (root n (norm (g n)))) sequentially" + using assms by eventually_elim simp + from Limsup_eq[OF this] show ?thesis unfolding conv_radius_def by simp +qed + +lemma conv_radius_altdef: + "conv_radius f = liminf (\n. inverse (ereal (root n (norm (f n)))))" + by (subst Liminf_inverse_ereal) (simp_all add: real_root_ge_zero conv_radius_def) + +lemma abs_summable_in_conv_radius: + fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" + assumes "ereal (norm z) < conv_radius f" + shows "summable (\n. norm (f n * z ^ n))" +proof (rule root_test_convergence') + define l where "l = limsup (\n. ereal (root n (norm (f n))))" + have "0 = limsup (\n. 0)" by (simp add: Limsup_const) + also have "... \ l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero) + finally have l_nonneg: "l \ 0" . + + have "limsup (\n. root n (norm (f n * z^n))) = l * ereal (norm z)" unfolding l_def + by (rule limsup_root_powser) + also from l_nonneg consider "l = 0" | "l = \" | "\l'. l = ereal l' \ l' > 0" + by (cases "l") (auto simp: less_le) + hence "l * ereal (norm z) < 1" + proof cases + assume "l = \" + hence "conv_radius f = 0" unfolding conv_radius_def l_def by simp + with assms show ?thesis by simp + next + assume "\l'. l = ereal l' \ l' > 0" + then guess l' by (elim exE conjE) note l' = this + hence "l \ \" by auto + have "l * ereal (norm z) < l * conv_radius f" + by (intro ereal_mult_strict_left_mono) (simp_all add: l' assms) + also have "conv_radius f = inverse l" by (simp add: conv_radius_def l_def) + also from l' have "l * inverse l = 1" by simp + finally show ?thesis . + qed simp_all + finally show "limsup (\n. ereal (root n (norm (norm (f n * z ^ n))))) < 1" by simp +qed + +lemma summable_in_conv_radius: + fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" + assumes "ereal (norm z) < conv_radius f" + shows "summable (\n. f n * z ^ n)" + by (rule summable_norm_cancel, rule abs_summable_in_conv_radius) fact+ + +lemma not_summable_outside_conv_radius: + fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" + assumes "ereal (norm z) > conv_radius f" + shows "\summable (\n. f n * z ^ n)" +proof (rule root_test_divergence) + define l where "l = limsup (\n. ereal (root n (norm (f n))))" + have "0 = limsup (\n. 0)" by (simp add: Limsup_const) + also have "... \ l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero) + finally have l_nonneg: "l \ 0" . + from assms have l_nz: "l \ 0" unfolding conv_radius_def l_def by auto + + have "limsup (\n. ereal (root n (norm (f n * z^n)))) = l * ereal (norm z)" + unfolding l_def by (rule limsup_root_powser) + also have "... > 1" + proof (cases l) + assume "l = \" + with assms conv_radius_nonneg[of f] show ?thesis + by (auto simp: zero_ereal_def[symmetric]) + next + fix l' assume l': "l = ereal l'" + from l_nonneg l_nz have "1 = l * inverse l" by (auto simp: l' field_simps) + also from l_nz have "inverse l = conv_radius f" + unfolding l_def conv_radius_def by auto + also from l' l_nz l_nonneg assms have "l * \ < l * ereal (norm z)" + by (intro ereal_mult_strict_left_mono) (auto simp: l') + finally show ?thesis . + qed (insert l_nonneg, simp_all) + finally show "limsup (\n. ereal (root n (norm (f n * z^n)))) > 1" . +qed + + +lemma conv_radius_geI: + assumes "summable (\n. f n * z ^ n :: 'a :: {banach, real_normed_div_algebra})" + shows "conv_radius f \ norm z" + using not_summable_outside_conv_radius[of f z] assms by (force simp: not_le[symmetric]) + +lemma conv_radius_leI: + assumes "\summable (\n. norm (f n * z ^ n :: 'a :: {banach, real_normed_div_algebra}))" + shows "conv_radius f \ norm z" + using abs_summable_in_conv_radius[of z f] assms by (force simp: not_le[symmetric]) + +lemma conv_radius_leI': + assumes "\summable (\n. f n * z ^ n :: 'a :: {banach, real_normed_div_algebra})" + shows "conv_radius f \ norm z" + using summable_in_conv_radius[of z f] assms by (force simp: not_le[symmetric]) + +lemma conv_radius_geI_ex: + fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" + assumes "\r. 0 < r \ ereal r < R \ \z. norm z = r \ summable (\n. f n * z^n)" + shows "conv_radius f \ R" +proof (rule linorder_cases[of "conv_radius f" R]) + assume R: "conv_radius f < R" + with conv_radius_nonneg[of f] obtain conv_radius' + where [simp]: "conv_radius f = ereal conv_radius'" + by (cases "conv_radius f") simp_all + define r where "r = (if R = \ then conv_radius' + 1 else (real_of_ereal R + conv_radius') / 2)" + from R conv_radius_nonneg[of f] have "0 < r \ ereal r < R \ ereal r > conv_radius f" + unfolding r_def by (cases R) (auto simp: r_def field_simps) + with assms(1)[of r] obtain z where "norm z > conv_radius f" "summable (\n. f n * z^n)" by auto + with not_summable_outside_conv_radius[of f z] show ?thesis by simp +qed simp_all + +lemma conv_radius_geI_ex': + fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" + assumes "\r. 0 < r \ ereal r < R \ summable (\n. f n * of_real r^n)" + shows "conv_radius f \ R" +proof (rule conv_radius_geI_ex) + fix r assume "0 < r" "ereal r < R" + with assms[of r] show "\z. norm z = r \ summable (\n. f n * z ^ n)" + by (intro exI[of _ "of_real r :: 'a"]) auto +qed + +lemma conv_radius_leI_ex: + fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" + assumes "R \ 0" + assumes "\r. 0 < r \ ereal r > R \ \z. norm z = r \ \summable (\n. norm (f n * z^n))" + shows "conv_radius f \ R" +proof (rule linorder_cases[of "conv_radius f" R]) + assume R: "conv_radius f > R" + from R assms(1) obtain R' where R': "R = ereal R'" by (cases R) simp_all + define r where + "r = (if conv_radius f = \ then R' + 1 else (R' + real_of_ereal (conv_radius f)) / 2)" + from R conv_radius_nonneg[of f] have "r > R \ r < conv_radius f" unfolding r_def + by (cases "conv_radius f") (auto simp: r_def field_simps R') + with assms(1) assms(2)[of r] R' + obtain z where "norm z < conv_radius f" "\summable (\n. norm (f n * z^n))" by auto + with abs_summable_in_conv_radius[of z f] show ?thesis by auto +qed simp_all + +lemma conv_radius_leI_ex': + fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" + assumes "R \ 0" + assumes "\r. 0 < r \ ereal r > R \ \summable (\n. f n * of_real r^n)" + shows "conv_radius f \ R" +proof (rule conv_radius_leI_ex) + fix r assume "0 < r" "ereal r > R" + with assms(2)[of r] show "\z. norm z = r \ \summable (\n. norm (f n * z ^ n))" + by (intro exI[of _ "of_real r :: 'a"]) (auto dest: summable_norm_cancel) +qed fact+ + +lemma conv_radius_eqI: + fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" + assumes "R \ 0" + assumes "\r. 0 < r \ ereal r < R \ \z. norm z = r \ summable (\n. f n * z^n)" + assumes "\r. 0 < r \ ereal r > R \ \z. norm z = r \ \summable (\n. norm (f n * z^n))" + shows "conv_radius f = R" + by (intro antisym conv_radius_geI_ex conv_radius_leI_ex assms) + +lemma conv_radius_eqI': + fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" + assumes "R \ 0" + assumes "\r. 0 < r \ ereal r < R \ summable (\n. f n * (of_real r)^n)" + assumes "\r. 0 < r \ ereal r > R \ \summable (\n. norm (f n * (of_real r)^n))" + shows "conv_radius f = R" +proof (intro conv_radius_eqI[OF assms(1)]) + fix r assume "0 < r" "ereal r < R" with assms(2)[OF this] + show "\z. norm z = r \ summable (\n. f n * z ^ n)" by force +next + fix r assume "0 < r" "ereal r > R" with assms(3)[OF this] + show "\z. norm z = r \ \summable (\n. norm (f n * z ^ n))" by force +qed + +lemma conv_radius_zeroI: + fixes f :: "nat \ 'a :: {banach,real_normed_div_algebra}" + assumes "\z. z \ 0 \ \summable (\n. f n * z^n)" + shows "conv_radius f = 0" +proof (rule ccontr) + assume "conv_radius f \ 0" + with conv_radius_nonneg[of f] have pos: "conv_radius f > 0" by simp + define r where "r = (if conv_radius f = \ then 1 else real_of_ereal (conv_radius f) / 2)" + from pos have r: "ereal r > 0 \ ereal r < conv_radius f" + by (cases "conv_radius f") (simp_all add: r_def) + hence "summable (\n. f n * of_real r ^ n)" by (intro summable_in_conv_radius) simp + moreover from r and assms[of "of_real r"] have "\summable (\n. f n * of_real r ^ n)" by simp + ultimately show False by contradiction +qed + +lemma conv_radius_inftyI': + fixes f :: "nat \ 'a :: {banach,real_normed_div_algebra}" + assumes "\r. r > c \ \z. norm z = r \ summable (\n. f n * z^n)" + shows "conv_radius f = \" +proof - + { + fix r :: real + have "max r (c + 1) > c" by (auto simp: max_def) + from assms[OF this] obtain z where "norm z = max r (c + 1)" "summable (\n. f n * z^n)" by blast + from conv_radius_geI[OF this(2)] this(1) have "conv_radius f \ r" by simp + } + from this[of "real_of_ereal (conv_radius f + 1)"] show "conv_radius f = \" + by (cases "conv_radius f") simp_all +qed + +lemma conv_radius_inftyI: + fixes f :: "nat \ 'a :: {banach,real_normed_div_algebra}" + assumes "\r. \z. norm z = r \ summable (\n. f n * z^n)" + shows "conv_radius f = \" + using assms by (rule conv_radius_inftyI') + +lemma conv_radius_inftyI'': + fixes f :: "nat \ 'a :: {banach,real_normed_div_algebra}" + assumes "\z. summable (\n. f n * z^n)" + shows "conv_radius f = \" +proof (rule conv_radius_inftyI') + fix r :: real assume "r > 0" + with assms show "\z. norm z = r \ summable (\n. f n * z^n)" + by (intro exI[of _ "of_real r"]) simp +qed + +lemma conv_radius_ratio_limit_ereal: + fixes f :: "nat \ 'a :: {banach,real_normed_div_algebra}" + assumes nz: "eventually (\n. f n \ 0) sequentially" + assumes lim: "(\n. ereal (norm (f n) / norm (f (Suc n)))) \ c" + shows "conv_radius f = c" +proof (rule conv_radius_eqI') + show "c \ 0" by (intro Lim_bounded2_ereal[OF lim]) simp_all +next + fix r assume r: "0 < r" "ereal r < c" + let ?l = "liminf (\n. ereal (norm (f n * of_real r ^ n) / norm (f (Suc n) * of_real r ^ Suc n)))" + have "?l = liminf (\n. ereal (norm (f n) / (norm (f (Suc n)))) * ereal (inverse r))" + using r by (simp add: norm_mult norm_power divide_simps) + also from r have "\ = liminf (\n. ereal (norm (f n) / (norm (f (Suc n))))) * ereal (inverse r)" + by (intro Liminf_ereal_mult_right) simp_all + also have "liminf (\n. ereal (norm (f n) / (norm (f (Suc n))))) = c" + by (intro lim_imp_Liminf lim) simp + finally have l: "?l = c * ereal (inverse r)" by simp + from r have l': "c * ereal (inverse r) > 1" by (cases c) (simp_all add: field_simps) + show "summable (\n. f n * of_real r^n)" + by (rule summable_norm_cancel, rule ratio_test_convergence) + (insert r nz l l', auto elim!: eventually_mono) +next + fix r assume r: "0 < r" "ereal r > c" + let ?l = "limsup (\n. ereal (norm (f n * of_real r ^ n) / norm (f (Suc n) * of_real r ^ Suc n)))" + have "?l = limsup (\n. ereal (norm (f n) / (norm (f (Suc n)))) * ereal (inverse r))" + using r by (simp add: norm_mult norm_power divide_simps) + also from r have "\ = limsup (\n. ereal (norm (f n) / (norm (f (Suc n))))) * ereal (inverse r)" + by (intro Limsup_ereal_mult_right) simp_all + also have "limsup (\n. ereal (norm (f n) / (norm (f (Suc n))))) = c" + by (intro lim_imp_Limsup lim) simp + finally have l: "?l = c * ereal (inverse r)" by simp + from r have l': "c * ereal (inverse r) < 1" by (cases c) (simp_all add: field_simps) + show "\summable (\n. norm (f n * of_real r^n))" + by (rule ratio_test_divergence) (insert r nz l l', auto elim!: eventually_mono) +qed + +lemma conv_radius_ratio_limit_ereal_nonzero: + fixes f :: "nat \ 'a :: {banach,real_normed_div_algebra}" + assumes nz: "c \ 0" + assumes lim: "(\n. ereal (norm (f n) / norm (f (Suc n)))) \ c" + shows "conv_radius f = c" +proof (rule conv_radius_ratio_limit_ereal[OF _ lim], rule ccontr) + assume "\eventually (\n. f n \ 0) sequentially" + hence "frequently (\n. f n = 0) sequentially" by (simp add: frequently_def) + hence "frequently (\n. ereal (norm (f n) / norm (f (Suc n))) = 0) sequentially" + by (force elim!: frequently_elim1) + hence "c = 0" by (intro limit_frequently_eq[OF _ _ lim]) auto + with nz show False by contradiction +qed + +lemma conv_radius_ratio_limit: + fixes f :: "nat \ 'a :: {banach,real_normed_div_algebra}" + assumes "c' = ereal c" + assumes nz: "eventually (\n. f n \ 0) sequentially" + assumes lim: "(\n. norm (f n) / norm (f (Suc n))) \ c" + shows "conv_radius f = c'" + using assms by (intro conv_radius_ratio_limit_ereal) simp_all + +lemma conv_radius_ratio_limit_nonzero: + fixes f :: "nat \ 'a :: {banach,real_normed_div_algebra}" + assumes "c' = ereal c" + assumes nz: "c \ 0" + assumes lim: "(\n. norm (f n) / norm (f (Suc n))) \ c" + shows "conv_radius f = c'" + using assms by (intro conv_radius_ratio_limit_ereal_nonzero) simp_all + +lemma conv_radius_mult_power: + assumes "c \ (0 :: 'a :: {real_normed_div_algebra,banach})" + shows "conv_radius (\n. c ^ n * f n) = conv_radius f / ereal (norm c)" +proof - + have "limsup (\n. ereal (root n (norm (c ^ n * f n)))) = + limsup (\n. ereal (norm c) * ereal (root n (norm (f n))))" + using eventually_gt_at_top[of "0::nat"] + by (intro Limsup_eq) + (auto elim!: eventually_mono simp: norm_mult norm_power real_root_mult real_root_power) + also have "\ = ereal (norm c) * limsup (\n. ereal (root n (norm (f n))))" + using assms by (subst Limsup_ereal_mult_left[symmetric]) simp_all + finally have A: "limsup (\n. ereal (root n (norm (c ^ n * f n)))) = + ereal (norm c) * limsup (\n. ereal (root n (norm (f n))))" . + show ?thesis using assms + apply (cases "limsup (\n. ereal (root n (norm (f n)))) = 0") + apply (simp add: A conv_radius_def) + apply (unfold conv_radius_def A divide_ereal_def, simp add: mult.commute ereal_inverse_mult) + done +qed + +lemma conv_radius_mult_power_right: + assumes "c \ (0 :: 'a :: {real_normed_div_algebra,banach})" + shows "conv_radius (\n. f n * c ^ n) = conv_radius f / ereal (norm c)" + using conv_radius_mult_power[OF assms, of f] + unfolding conv_radius_def by (simp add: mult.commute norm_mult) + +lemma conv_radius_divide_power: + assumes "c \ (0 :: 'a :: {real_normed_div_algebra,banach})" + shows "conv_radius (\n. f n / c^n) = conv_radius f * ereal (norm c)" +proof - + from assms have "inverse c \ 0" by simp + from conv_radius_mult_power_right[OF this, of f] show ?thesis + by (simp add: divide_inverse divide_ereal_def assms norm_inverse power_inverse) +qed + + +lemma conv_radius_add_ge: + "min (conv_radius f) (conv_radius g) \ + conv_radius (\x. f x + g x :: 'a :: {banach,real_normed_div_algebra})" + by (rule conv_radius_geI_ex') + (auto simp: algebra_simps intro!: summable_add summable_in_conv_radius) + +lemma conv_radius_mult_ge: + fixes f g :: "nat \ ('a :: {banach,real_normed_div_algebra})" + shows "conv_radius (\x. \i\x. f i * g (x - i)) \ min (conv_radius f) (conv_radius g)" +proof (rule conv_radius_geI_ex') + fix r assume r: "r > 0" "ereal r < min (conv_radius f) (conv_radius g)" + from r have "summable (\n. (\i\n. (f i * of_real r^i) * (g (n - i) * of_real r^(n - i))))" + by (intro summable_Cauchy_product abs_summable_in_conv_radius) simp_all + thus "summable (\n. (\i\n. f i * g (n - i)) * of_real r ^ n)" + by (simp add: algebra_simps of_real_def power_add [symmetric] scaleR_setsum_right) +qed + +lemma le_conv_radius_iff: + fixes a :: "nat \ 'a::{real_normed_div_algebra,banach}" + shows "r \ conv_radius a \ (\x. norm (x-\) < r \ summable (\i. a i * (x - \) ^ i))" +apply (intro iffI allI impI summable_in_conv_radius conv_radius_geI_ex) +apply (meson less_ereal.simps(1) not_le order_trans) +apply (rule_tac x="of_real ra" in exI, simp) +apply (metis abs_of_nonneg add_diff_cancel_left' less_eq_real_def norm_of_real) +done + +end diff -r bbcb05504fdc -r bd218a9320b5 src/HOL/Multivariate_Analysis/Uniform_Limit.thy --- a/src/HOL/Multivariate_Analysis/Uniform_Limit.thy Thu Aug 04 18:45:28 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Uniform_Limit.thy Thu Aug 04 19:36:31 2016 +0200 @@ -6,7 +6,7 @@ section \Uniform Limit and Uniform Convergence\ theory Uniform_Limit -imports Topology_Euclidean_Space Summation +imports Topology_Euclidean_Space Summation_Tests begin definition uniformly_on :: "'a set \ ('a \ 'b::metric_space) \ ('a \ 'b) filter" @@ -18,9 +18,9 @@ definition uniformly_convergent_on where "uniformly_convergent_on X f \ (\l. uniform_limit X f l sequentially)" -definition uniformly_Cauchy_on where +definition uniformly_Cauchy_on where "uniformly_Cauchy_on X f \ (\e>0. \M. \x\X. \(m::nat)\M. \n\M. dist (f m x) (f n x) < e)" - + lemma uniform_limit_iff: "uniform_limit S f l F \ (\e>0. \\<^sub>F n in F. \x\S. dist (f n x) (l x) < e)" unfolding filterlim_iff uniformly_on_def @@ -141,7 +141,7 @@ shows "uniformly_Cauchy_on X f" proof (rule uniformly_Cauchy_onI) fix e :: real assume e: "e > 0" - from assms[OF this] obtain M + from assms[OF this] obtain M where M: "\x m n. x \ X \ m \ M \ n > m \ dist (f m x) (f n x) < e" by fast { fix x m n assume x: "x \ X" and m: "m \ M" and n: "n \ M" @@ -151,7 +151,7 @@ thus "\M. \x\X. \m\M. \n\M. dist (f m x) (f n x) < e" by fast qed -lemma uniformly_Cauchy_imp_Cauchy: +lemma uniformly_Cauchy_imp_Cauchy: "uniformly_Cauchy_on X f \ x \ X \ Cauchy (\n. f n x)" unfolding Cauchy_def uniformly_Cauchy_on_def by fast @@ -167,7 +167,7 @@ and B: "\x. x \ X \ h x = i x" { fix e ::real assume "e > 0" - with C have "eventually (\y. \x\X. dist (f y x) (h x) < e) F" + with C have "eventually (\y. \x\X. dist (f y x) (h x) < e) F" unfolding uniform_limit_iff by blast with A have "eventually (\y. \x\X. dist (g y x) (i x) < e) F" by eventually_elim (insert B, simp_all) @@ -188,7 +188,7 @@ "uniformly_convergent_on X f \ uniform_limit X f (\x. lim (\n. f n x)) sequentially" proof assume "uniformly_convergent_on X f" - then obtain l where l: "uniform_limit X f l sequentially" + then obtain l where l: "uniform_limit X f l sequentially" unfolding uniformly_convergent_on_def by blast from l have "uniform_limit X f (\x. lim (\n. f n x)) sequentially \ uniform_limit X f l sequentially" @@ -221,11 +221,11 @@ show "\x\X. dist (f n x) (?f x) < e" proof fix x assume x: "x \ X" - with assms have "(\n. f n x) \ ?f x" + with assms have "(\n. f n x) \ ?f x" by (auto dest!: Cauchy_convergent uniformly_Cauchy_imp_Cauchy simp: convergent_LIMSEQ_iff) with \e/2 > 0\ have "eventually (\m. m \ N \ dist (f m x) (?f x) < e/2) sequentially" by (intro tendstoD eventually_conj eventually_ge_at_top) - then obtain m where m: "m \ N" "dist (f m x) (?f x) < e/2" + then obtain m where m: "m \ N" "dist (f m x) (?f x) < e/2" unfolding eventually_at_top_linorder by blast have "dist (f n x) (?f x) \ dist (f n x) (f m x) + dist (f m x) (?f x)" by (rule dist_triangle) @@ -306,16 +306,16 @@ assumes "summable M" shows "uniform_limit A (\n x. \ix. suminf (\i. f i x)) sequentially" using assms by (intro weierstrass_m_test_ev always_eventually) auto - + corollary weierstrass_m_test'_ev: fixes f :: "_ \ _ \ _ :: banach" - assumes "eventually (\n. \x\A. norm (f n x) \ M n) sequentially" "summable M" + assumes "eventually (\n. \x\A. norm (f n x) \ M n) sequentially" "summable M" shows "uniformly_convergent_on A (\n x. \i _ \ _ :: banach" - assumes "\n x. x \ A \ norm (f n x) \ M n" "summable M" + assumes "\n x. x \ A \ norm (f n x) \ M n" "summable M" shows "uniformly_convergent_on A (\n x. \i + "uniformly_convergent_on A f \ uniformly_convergent_on A (\k x. c * f k x :: 'a :: {real_normed_algebra})" unfolding uniformly_convergent_on_def by (blast dest: bounded_linear_uniform_limit_intros(13)) diff -r bbcb05504fdc -r bd218a9320b5 src/HOL/Multivariate_Analysis/Weierstrass.thy --- a/src/HOL/Multivariate_Analysis/Weierstrass.thy Thu Aug 04 18:45:28 2016 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1218 +0,0 @@ -section \The Bernstein-Weierstrass and Stone-Weierstrass Theorems\ - -text\By L C Paulson (2015)\ - -theory Weierstrass -imports Uniform_Limit Path_Connected -begin - -subsection \Bernstein polynomials\ - -definition Bernstein :: "[nat,nat,real] \ real" where - "Bernstein n k x \ of_nat (n choose k) * x ^ k * (1 - x) ^ (n - k)" - -lemma Bernstein_nonneg: "\0 \ x; x \ 1\ \ 0 \ Bernstein n k x" - by (simp add: Bernstein_def) - -lemma Bernstein_pos: "\0 < x; x < 1; k \ n\ \ 0 < Bernstein n k x" - by (simp add: Bernstein_def) - -lemma sum_Bernstein [simp]: "(\ k = 0..n. Bernstein n k x) = 1" - using binomial_ring [of x "1-x" n] - by (simp add: Bernstein_def) - -lemma binomial_deriv1: - "(\k=0..n. (of_nat k * of_nat (n choose k)) * a^(k-1) * b^(n-k)) = real_of_nat n * (a+b) ^ (n-1)" - apply (rule DERIV_unique [where f = "\a. (a+b)^n" and x=a]) - apply (subst binomial_ring) - apply (rule derivative_eq_intros setsum.cong | simp)+ - done - -lemma binomial_deriv2: - "(\k=0..n. (of_nat k * of_nat (k-1) * of_nat (n choose k)) * a^(k-2) * b^(n-k)) = - of_nat n * of_nat (n-1) * (a+b::real) ^ (n-2)" - apply (rule DERIV_unique [where f = "\a. of_nat n * (a+b::real) ^ (n-1)" and x=a]) - apply (subst binomial_deriv1 [symmetric]) - apply (rule derivative_eq_intros setsum.cong | simp add: Num.numeral_2_eq_2)+ - done - -lemma sum_k_Bernstein [simp]: "(\k = 0..n. real k * Bernstein n k x) = of_nat n * x" - apply (subst binomial_deriv1 [of n x "1-x", simplified, symmetric]) - apply (simp add: setsum_left_distrib) - apply (auto simp: Bernstein_def algebra_simps realpow_num_eq_if intro!: setsum.cong) - done - -lemma sum_kk_Bernstein [simp]: "(\ k = 0..n. real k * (real k - 1) * Bernstein n k x) = real n * (real n - 1) * x\<^sup>2" -proof - - have "(\ k = 0..n. real k * (real k - 1) * Bernstein n k x) = real_of_nat n * real_of_nat (n - Suc 0) * x\<^sup>2" - apply (subst binomial_deriv2 [of n x "1-x", simplified, symmetric]) - apply (simp add: setsum_left_distrib) - apply (rule setsum.cong [OF refl]) - apply (simp add: Bernstein_def power2_eq_square algebra_simps) - apply (rename_tac k) - apply (subgoal_tac "k = 0 \ k = 1 \ (\k'. k = Suc (Suc k'))") - apply (force simp add: field_simps of_nat_Suc power2_eq_square) - by presburger - also have "... = n * (n - 1) * x\<^sup>2" - by auto - finally show ?thesis - by auto -qed - -subsection \Explicit Bernstein version of the 1D Weierstrass approximation theorem\ - -lemma Bernstein_Weierstrass: - fixes f :: "real \ real" - assumes contf: "continuous_on {0..1} f" and e: "0 < e" - shows "\N. \n x. N \ n \ x \ {0..1} - \ \f x - (\k = 0..n. f(k/n) * Bernstein n k x)\ < e" -proof - - have "bounded (f ` {0..1})" - using compact_continuous_image compact_imp_bounded contf by blast - then obtain M where M: "\x. 0 \ x \ x \ 1 \ \f x\ \ M" - by (force simp add: bounded_iff) - then have Mge0: "0 \ M" by force - have ucontf: "uniformly_continuous_on {0..1} f" - using compact_uniformly_continuous contf by blast - then obtain d where d: "d>0" "\x x'. \ x \ {0..1}; x' \ {0..1}; \x' - x\ < d\ \ \f x' - f x\ < e/2" - apply (rule uniformly_continuous_onE [where e = "e/2"]) - using e by (auto simp: dist_norm) - { fix n::nat and x::real - assume n: "Suc (nat\4*M/(e*d\<^sup>2)\) \ n" and x: "0 \ x" "x \ 1" - have "0 < n" using n by simp - have ed0: "- (e * d\<^sup>2) < 0" - using e \0 by simp - also have "... \ M * 4" - using \0\M\ by simp - finally have [simp]: "real_of_int (nat \4 * M / (e * d\<^sup>2)\) = real_of_int \4 * M / (e * d\<^sup>2)\" - using \0\M\ e \0 - by (simp add: of_nat_Suc field_simps) - have "4*M/(e*d\<^sup>2) + 1 \ real (Suc (nat\4*M/(e*d\<^sup>2)\))" - by (simp add: of_nat_Suc real_nat_ceiling_ge) - also have "... \ real n" - using n by (simp add: of_nat_Suc field_simps) - finally have nbig: "4*M/(e*d\<^sup>2) + 1 \ real n" . - have sum_bern: "(\k = 0..n. (x - k/n)\<^sup>2 * Bernstein n k x) = x * (1 - x) / n" - proof - - have *: "\a b x::real. (a - b)\<^sup>2 * x = a * (a - 1) * x + (1 - 2 * b) * a * x + b * b * x" - by (simp add: algebra_simps power2_eq_square) - have "(\ k = 0..n. (k - n * x)\<^sup>2 * Bernstein n k x) = n * x * (1 - x)" - apply (simp add: * setsum.distrib) - apply (simp add: setsum_right_distrib [symmetric] mult.assoc) - apply (simp add: algebra_simps power2_eq_square) - done - then have "(\ k = 0..n. (k - n * x)\<^sup>2 * Bernstein n k x)/n^2 = x * (1 - x) / n" - by (simp add: power2_eq_square) - then show ?thesis - using n by (simp add: setsum_divide_distrib divide_simps mult.commute power2_commute) - qed - { fix k - assume k: "k \ n" - then have kn: "0 \ k / n" "k / n \ 1" - by (auto simp: divide_simps) - consider (lessd) "\x - k / n\ < d" | (ged) "d \ \x - k / n\" - by linarith - then have "\(f x - f (k/n))\ \ e/2 + 2 * M / d\<^sup>2 * (x - k/n)\<^sup>2" - proof cases - case lessd - then have "\(f x - f (k/n))\ < e/2" - using d x kn by (simp add: abs_minus_commute) - also have "... \ (e/2 + 2 * M / d\<^sup>2 * (x - k/n)\<^sup>2)" - using Mge0 d by simp - finally show ?thesis by simp - next - case ged - then have dle: "d\<^sup>2 \ (x - k/n)\<^sup>2" - by (metis d(1) less_eq_real_def power2_abs power_mono) - have "\(f x - f (k/n))\ \ \f x\ + \f (k/n)\" - by (rule abs_triangle_ineq4) - also have "... \ M+M" - by (meson M add_mono_thms_linordered_semiring(1) kn x) - also have "... \ 2 * M * ((x - k/n)\<^sup>2 / d\<^sup>2)" - apply simp - apply (rule Rings.ordered_semiring_class.mult_left_mono [of 1 "((x - k/n)\<^sup>2 / d\<^sup>2)", simplified]) - using dle \d>0\ \M\0\ by auto - also have "... \ e/2 + 2 * M / d\<^sup>2 * (x - k/n)\<^sup>2" - using e by simp - finally show ?thesis . - qed - } note * = this - have "\f x - (\ k = 0..n. f(k / n) * Bernstein n k x)\ \ \\ k = 0..n. (f x - f(k / n)) * Bernstein n k x\" - by (simp add: setsum_subtractf setsum_right_distrib [symmetric] algebra_simps) - also have "... \ (\ k = 0..n. (e/2 + (2 * M / d\<^sup>2) * (x - k / n)\<^sup>2) * Bernstein n k x)" - apply (rule order_trans [OF setsum_abs setsum_mono]) - using * - apply (simp add: abs_mult Bernstein_nonneg x mult_right_mono) - done - also have "... \ e/2 + (2 * M) / (d\<^sup>2 * n)" - apply (simp only: setsum.distrib Rings.semiring_class.distrib_right setsum_right_distrib [symmetric] mult.assoc sum_bern) - using \d>0\ x - apply (simp add: divide_simps Mge0 mult_le_one mult_left_le) - done - also have "... < e" - apply (simp add: field_simps) - using \d>0\ nbig e \n>0\ - apply (simp add: divide_simps algebra_simps) - using ed0 by linarith - finally have "\f x - (\k = 0..n. f (real k / real n) * Bernstein n k x)\ < e" . - } - then show ?thesis - by auto -qed - - -subsection \General Stone-Weierstrass theorem\ - -text\Source: -Bruno Brosowski and Frank Deutsch. -An Elementary Proof of the Stone-Weierstrass Theorem. -Proceedings of the American Mathematical Society -Volume 81, Number 1, January 1981. -DOI: 10.2307/2043993 http://www.jstor.org/stable/2043993\ - -locale function_ring_on = - fixes R :: "('a::t2_space \ real) set" and s :: "'a set" - assumes compact: "compact s" - assumes continuous: "f \ R \ continuous_on s f" - assumes add: "f \ R \ g \ R \ (\x. f x + g x) \ R" - assumes mult: "f \ R \ g \ R \ (\x. f x * g x) \ R" - assumes const: "(\_. c) \ R" - assumes separable: "x \ s \ y \ s \ x \ y \ \f\R. f x \ f y" - -begin - lemma minus: "f \ R \ (\x. - f x) \ R" - by (frule mult [OF const [of "-1"]]) simp - - lemma diff: "f \ R \ g \ R \ (\x. f x - g x) \ R" - unfolding diff_conv_add_uminus by (metis add minus) - - lemma power: "f \ R \ (\x. f x ^ n) \ R" - by (induct n) (auto simp: const mult) - - lemma setsum: "\finite I; \i. i \ I \ f i \ R\ \ (\x. \i \ I. f i x) \ R" - by (induct I rule: finite_induct; simp add: const add) - - lemma setprod: "\finite I; \i. i \ I \ f i \ R\ \ (\x. \i \ I. f i x) \ R" - by (induct I rule: finite_induct; simp add: const mult) - - definition normf :: "('a::t2_space \ real) \ real" - where "normf f \ SUP x:s. \f x\" - - lemma normf_upper: "\continuous_on s f; x \ s\ \ \f x\ \ normf f" - apply (simp add: normf_def) - apply (rule cSUP_upper, assumption) - by (simp add: bounded_imp_bdd_above compact compact_continuous_image compact_imp_bounded continuous_on_rabs) - - lemma normf_least: "s \ {} \ (\x. x \ s \ \f x\ \ M) \ normf f \ M" - by (simp add: normf_def cSUP_least) - -end - -lemma (in function_ring_on) one: - assumes U: "open U" and t0: "t0 \ s" "t0 \ U" and t1: "t1 \ s-U" - shows "\V. open V \ t0 \ V \ s \ V \ U \ - (\e>0. \f \ R. f ` s \ {0..1} \ (\t \ s \ V. f t < e) \ (\t \ s - U. f t > 1 - e))" -proof - - have "\pt \ R. pt t0 = 0 \ pt t > 0 \ pt ` s \ {0..1}" if t: "t \ s - U" for t - proof - - have "t \ t0" using t t0 by auto - then obtain g where g: "g \ R" "g t \ g t0" - using separable t0 by (metis Diff_subset subset_eq t) - define h where [abs_def]: "h x = g x - g t0" for x - have "h \ R" - unfolding h_def by (fast intro: g const diff) - then have hsq: "(\w. (h w)\<^sup>2) \ R" - by (simp add: power2_eq_square mult) - have "h t \ h t0" - by (simp add: h_def g) - then have "h t \ 0" - by (simp add: h_def) - then have ht2: "0 < (h t)^2" - by simp - also have "... \ normf (\w. (h w)\<^sup>2)" - using t normf_upper [where x=t] continuous [OF hsq] by force - finally have nfp: "0 < normf (\w. (h w)\<^sup>2)" . - define p where [abs_def]: "p x = (1 / normf (\w. (h w)\<^sup>2)) * (h x)^2" for x - have "p \ R" - unfolding p_def by (fast intro: hsq const mult) - moreover have "p t0 = 0" - by (simp add: p_def h_def) - moreover have "p t > 0" - using nfp ht2 by (simp add: p_def) - moreover have "\x. x \ s \ p x \ {0..1}" - using nfp normf_upper [OF continuous [OF hsq] ] by (auto simp: p_def) - ultimately show "\pt \ R. pt t0 = 0 \ pt t > 0 \ pt ` s \ {0..1}" - by auto - qed - then obtain pf where pf: "\t. t \ s-U \ pf t \ R \ pf t t0 = 0 \ pf t t > 0" - and pf01: "\t. t \ s-U \ pf t ` s \ {0..1}" - by metis - have com_sU: "compact (s-U)" - using compact closed_Int_compact U by (simp add: Diff_eq compact_Int_closed open_closed) - have "\t. t \ s-U \ \A. open A \ A \ s = {x\s. 0 < pf t x}" - apply (rule open_Collect_positive) - by (metis pf continuous) - then obtain Uf where Uf: "\t. t \ s-U \ open (Uf t) \ (Uf t) \ s = {x\s. 0 < pf t x}" - by metis - then have open_Uf: "\t. t \ s-U \ open (Uf t)" - by blast - have tUft: "\t. t \ s-U \ t \ Uf t" - using pf Uf by blast - then have *: "s-U \ (\x \ s-U. Uf x)" - by blast - obtain subU where subU: "subU \ s - U" "finite subU" "s - U \ (\x \ subU. Uf x)" - by (blast intro: that open_Uf compactE_image [OF com_sU _ *]) - then have [simp]: "subU \ {}" - using t1 by auto - then have cardp: "card subU > 0" using subU - by (simp add: card_gt_0_iff) - define p where [abs_def]: "p x = (1 / card subU) * (\t \ subU. pf t x)" for x - have pR: "p \ R" - unfolding p_def using subU pf by (fast intro: pf const mult setsum) - have pt0 [simp]: "p t0 = 0" - using subU pf by (auto simp: p_def intro: setsum.neutral) - have pt_pos: "p t > 0" if t: "t \ s-U" for t - proof - - obtain i where i: "i \ subU" "t \ Uf i" using subU t by blast - show ?thesis - using subU i t - apply (clarsimp simp: p_def divide_simps) - apply (rule setsum_pos2 [OF \finite subU\]) - using Uf t pf01 apply auto - apply (force elim!: subsetCE) - done - qed - have p01: "p x \ {0..1}" if t: "x \ s" for x - proof - - have "0 \ p x" - using subU cardp t - apply (simp add: p_def divide_simps setsum_nonneg) - apply (rule setsum_nonneg) - using pf01 by force - moreover have "p x \ 1" - using subU cardp t - apply (simp add: p_def divide_simps setsum_nonneg) - apply (rule setsum_bounded_above [where 'a=real and K=1, simplified]) - using pf01 by force - ultimately show ?thesis - by auto - qed - have "compact (p ` (s-U))" - by (meson Diff_subset com_sU compact_continuous_image continuous continuous_on_subset pR) - then have "open (- (p ` (s-U)))" - by (simp add: compact_imp_closed open_Compl) - moreover have "0 \ - (p ` (s-U))" - by (metis (no_types) ComplI image_iff not_less_iff_gr_or_eq pt_pos) - ultimately obtain delta0 where delta0: "delta0 > 0" "ball 0 delta0 \ - (p ` (s-U))" - by (auto simp: elim!: openE) - then have pt_delta: "\x. x \ s-U \ p x \ delta0" - by (force simp: ball_def dist_norm dest: p01) - define \ where "\ = delta0/2" - have "delta0 \ 1" using delta0 p01 [of t1] t1 - by (force simp: ball_def dist_norm dest: p01) - with delta0 have \01: "0 < \" "\ < 1" - by (auto simp: \_def) - have pt_\: "\x. x \ s-U \ p x \ \" - using pt_delta delta0 by (force simp: \_def) - have "\A. open A \ A \ s = {x\s. p x < \/2}" - by (rule open_Collect_less_Int [OF continuous [OF pR] continuous_on_const]) - then obtain V where V: "open V" "V \ s = {x\s. p x < \/2}" - by blast - define k where "k = nat\1/\\ + 1" - have "k>0" by (simp add: k_def) - have "k-1 \ 1/\" - using \01 by (simp add: k_def) - with \01 have "k \ (1+\)/\" - by (auto simp: algebra_simps add_divide_distrib) - also have "... < 2/\" - using \01 by (auto simp: divide_simps) - finally have k2\: "k < 2/\" . - have "1/\ < k" - using \01 unfolding k_def by linarith - with \01 k2\ have k\: "1 < k*\" "k*\ < 2" - by (auto simp: divide_simps) - define q where [abs_def]: "q n t = (1 - p t ^ n) ^ (k^n)" for n t - have qR: "q n \ R" for n - by (simp add: q_def const diff power pR) - have q01: "\n t. t \ s \ q n t \ {0..1}" - using p01 by (simp add: q_def power_le_one algebra_simps) - have qt0 [simp]: "\n. n>0 \ q n t0 = 1" - using t0 pf by (simp add: q_def power_0_left) - { fix t and n::nat - assume t: "t \ s \ V" - with \k>0\ V have "k * p t < k * \ / 2" - by force - then have "1 - (k * \ / 2)^n \ 1 - (k * p t)^n" - using \k>0\ p01 t by (simp add: power_mono) - also have "... \ q n t" - using Bernoulli_inequality [of "- ((p t)^n)" "k^n"] - apply (simp add: q_def) - by (metis IntE atLeastAtMost_iff p01 power_le_one power_mult_distrib t) - finally have "1 - (k * \ / 2) ^ n \ q n t" . - } note limitV = this - { fix t and n::nat - assume t: "t \ s - U" - with \k>0\ U have "k * \ \ k * p t" - by (simp add: pt_\) - with k\ have kpt: "1 < k * p t" - by (blast intro: less_le_trans) - have ptn_pos: "0 < p t ^ n" - using pt_pos [OF t] by simp - have ptn_le: "p t ^ n \ 1" - by (meson DiffE atLeastAtMost_iff p01 power_le_one t) - have "q n t = (1/(k^n * (p t)^n)) * (1 - p t ^ n) ^ (k^n) * k^n * (p t)^n" - using pt_pos [OF t] \k>0\ by (simp add: q_def) - also have "... \ (1/(k * (p t))^n) * (1 - p t ^ n) ^ (k^n) * (1 + k^n * (p t)^n)" - using pt_pos [OF t] \k>0\ - apply simp - apply (simp only: times_divide_eq_right [symmetric]) - apply (rule mult_left_mono [of "1::real", simplified]) - apply (simp_all add: power_mult_distrib) - apply (rule zero_le_power) - using ptn_le by linarith - also have "... \ (1/(k * (p t))^n) * (1 - p t ^ n) ^ (k^n) * (1 + (p t)^n) ^ (k^n)" - apply (rule mult_left_mono [OF Bernoulli_inequality [of "p t ^ n" "k^n"]]) - using \k>0\ ptn_pos ptn_le - apply (auto simp: power_mult_distrib) - done - also have "... = (1/(k * (p t))^n) * (1 - p t ^ (2*n)) ^ (k^n)" - using pt_pos [OF t] \k>0\ - by (simp add: algebra_simps power_mult power2_eq_square power_mult_distrib [symmetric]) - also have "... \ (1/(k * (p t))^n) * 1" - apply (rule mult_left_mono [OF power_le_one]) - using pt_pos \k>0\ p01 power_le_one t apply auto - done - also have "... \ (1 / (k*\))^n" - using \k>0\ \01 power_mono pt_\ t - by (fastforce simp: field_simps) - finally have "q n t \ (1 / (real k * \)) ^ n " . - } note limitNonU = this - define NN - where "NN e = 1 + nat \max (ln e / ln (real k * \ / 2)) (- ln e / ln (real k * \))\" for e - have NN: "of_nat (NN e) > ln e / ln (real k * \ / 2)" "of_nat (NN e) > - ln e / ln (real k * \)" - if "0e. e>0 \ (k * \ / 2)^NN e < e" - apply (subst Transcendental.ln_less_cancel_iff [symmetric]) - prefer 3 apply (subst ln_realpow) - using \k>0\ \\>0\ NN k\ - apply (force simp add: field_simps)+ - done - have NN0: "\e. e>0 \ (1/(k*\))^NN e < e" - apply (subst Transcendental.ln_less_cancel_iff [symmetric]) - prefer 3 apply (subst ln_realpow) - using \k>0\ \\>0\ NN k\ - apply (force simp add: field_simps ln_div)+ - done - { fix t and e::real - assume "e>0" - have "t \ s \ V \ 1 - q (NN e) t < e" "t \ s - U \ q (NN e) t < e" - proof - - assume t: "t \ s \ V" - show "1 - q (NN e) t < e" - by (metis add.commute diff_le_eq not_le limitV [OF t] less_le_trans [OF NN1 [OF \e>0\]]) - next - assume t: "t \ s - U" - show "q (NN e) t < e" - using limitNonU [OF t] less_le_trans [OF NN0 [OF \e>0\]] not_le by blast - qed - } then have "\e. e > 0 \ \f\R. f ` s \ {0..1} \ (\t \ s \ V. f t < e) \ (\t \ s - U. 1 - e < f t)" - using q01 - by (rule_tac x="\x. 1 - q (NN e) x" in bexI) (auto simp: algebra_simps intro: diff const qR) - moreover have t0V: "t0 \ V" "s \ V \ U" - using pt_\ t0 U V \01 by fastforce+ - ultimately show ?thesis using V t0V - by blast -qed - -text\Non-trivial case, with @{term A} and @{term B} both non-empty\ -lemma (in function_ring_on) two_special: - assumes A: "closed A" "A \ s" "a \ A" - and B: "closed B" "B \ s" "b \ B" - and disj: "A \ B = {}" - and e: "0 < e" "e < 1" - shows "\f \ R. f ` s \ {0..1} \ (\x \ A. f x < e) \ (\x \ B. f x > 1 - e)" -proof - - { fix w - assume "w \ A" - then have "open ( - B)" "b \ s" "w \ B" "w \ s" - using assms by auto - then have "\V. open V \ w \ V \ s \ V \ -B \ - (\e>0. \f \ R. f ` s \ {0..1} \ (\x \ s \ V. f x < e) \ (\x \ s \ B. f x > 1 - e))" - using one [of "-B" w b] assms \w \ A\ by simp - } - then obtain Vf where Vf: - "\w. w \ A \ open (Vf w) \ w \ Vf w \ s \ Vf w \ -B \ - (\e>0. \f \ R. f ` s \ {0..1} \ (\x \ s \ Vf w. f x < e) \ (\x \ s \ B. f x > 1 - e))" - by metis - then have open_Vf: "\w. w \ A \ open (Vf w)" - by blast - have tVft: "\w. w \ A \ w \ Vf w" - using Vf by blast - then have setsum_max_0: "A \ (\x \ A. Vf x)" - by blast - have com_A: "compact A" using A - by (metis compact compact_Int_closed inf.absorb_iff2) - obtain subA where subA: "subA \ A" "finite subA" "A \ (\x \ subA. Vf x)" - by (blast intro: that open_Vf compactE_image [OF com_A _ setsum_max_0]) - then have [simp]: "subA \ {}" - using \a \ A\ by auto - then have cardp: "card subA > 0" using subA - by (simp add: card_gt_0_iff) - have "\w. w \ A \ \f \ R. f ` s \ {0..1} \ (\x \ s \ Vf w. f x < e / card subA) \ (\x \ s \ B. f x > 1 - e / card subA)" - using Vf e cardp by simp - then obtain ff where ff: - "\w. w \ A \ ff w \ R \ ff w ` s \ {0..1} \ - (\x \ s \ Vf w. ff w x < e / card subA) \ (\x \ s \ B. ff w x > 1 - e / card subA)" - by metis - define pff where [abs_def]: "pff x = (\w \ subA. ff w x)" for x - have pffR: "pff \ R" - unfolding pff_def using subA ff by (auto simp: intro: setprod) - moreover - have pff01: "pff x \ {0..1}" if t: "x \ s" for x - proof - - have "0 \ pff x" - using subA cardp t - apply (simp add: pff_def divide_simps setsum_nonneg) - apply (rule Groups_Big.linordered_semidom_class.setprod_nonneg) - using ff by fastforce - moreover have "pff x \ 1" - using subA cardp t - apply (simp add: pff_def divide_simps setsum_nonneg) - apply (rule setprod_mono [where g = "\x. 1", simplified]) - using ff by fastforce - ultimately show ?thesis - by auto - qed - moreover - { fix v x - assume v: "v \ subA" and x: "x \ Vf v" "x \ s" - from subA v have "pff x = ff v x * (\w \ subA - {v}. ff w x)" - unfolding pff_def by (metis setprod.remove) - also have "... \ ff v x * 1" - apply (rule Rings.ordered_semiring_class.mult_left_mono) - apply (rule setprod_mono [where g = "\x. 1", simplified]) - using ff [THEN conjunct2, THEN conjunct1] v subA x - apply auto - apply (meson atLeastAtMost_iff contra_subsetD imageI) - apply (meson atLeastAtMost_iff contra_subsetD image_eqI) - using atLeastAtMost_iff by blast - also have "... < e / card subA" - using ff [THEN conjunct2, THEN conjunct2, THEN conjunct1] v subA x - by auto - also have "... \ e" - using cardp e by (simp add: divide_simps) - finally have "pff x < e" . - } - then have "\x. x \ A \ pff x < e" - using A Vf subA by (metis UN_E contra_subsetD) - moreover - { fix x - assume x: "x \ B" - then have "x \ s" - using B by auto - have "1 - e \ (1 - e / card subA) ^ card subA" - using Bernoulli_inequality [of "-e / card subA" "card subA"] e cardp - by (auto simp: field_simps) - also have "... = (\w \ subA. 1 - e / card subA)" - by (simp add: setprod_constant subA(2)) - also have "... < pff x" - apply (simp add: pff_def) - apply (rule setprod_mono_strict [where f = "\x. 1 - e / card subA", simplified]) - apply (simp_all add: subA(2)) - apply (intro ballI conjI) - using e apply (force simp: divide_simps) - using ff [THEN conjunct2, THEN conjunct2, THEN conjunct2] subA B x - apply blast - done - finally have "1 - e < pff x" . - } - ultimately - show ?thesis by blast -qed - -lemma (in function_ring_on) two: - assumes A: "closed A" "A \ s" - and B: "closed B" "B \ s" - and disj: "A \ B = {}" - and e: "0 < e" "e < 1" - shows "\f \ R. f ` s \ {0..1} \ (\x \ A. f x < e) \ (\x \ B. f x > 1 - e)" -proof (cases "A \ {} \ B \ {}") - case True then show ?thesis - apply (simp add: ex_in_conv [symmetric]) - using assms - apply safe - apply (force simp add: intro!: two_special) - done -next - case False with e show ?thesis - apply simp - apply (erule disjE) - apply (rule_tac [2] x="\x. 0" in bexI) - apply (rule_tac x="\x. 1" in bexI) - apply (auto simp: const) - done -qed - -text\The special case where @{term f} is non-negative and @{term"e<1/3"}\ -lemma (in function_ring_on) Stone_Weierstrass_special: - assumes f: "continuous_on s f" and fpos: "\x. x \ s \ f x \ 0" - and e: "0 < e" "e < 1/3" - shows "\g \ R. \x\s. \f x - g x\ < 2*e" -proof - - define n where "n = 1 + nat \normf f / e\" - define A where "A j = {x \ s. f x \ (j - 1/3)*e}" for j :: nat - define B where "B j = {x \ s. f x \ (j + 1/3)*e}" for j :: nat - have ngt: "(n-1) * e \ normf f" "n\1" - using e - apply (simp_all add: n_def field_simps of_nat_Suc) - by (metis real_nat_ceiling_ge mult.commute not_less pos_less_divide_eq) - then have ge_fx: "(n-1) * e \ f x" if "x \ s" for x - using f normf_upper that by fastforce - { fix j - have A: "closed (A j)" "A j \ s" - apply (simp_all add: A_def Collect_restrict) - apply (rule continuous_on_closed_Collect_le [OF f continuous_on_const]) - apply (simp add: compact compact_imp_closed) - done - have B: "closed (B j)" "B j \ s" - apply (simp_all add: B_def Collect_restrict) - apply (rule continuous_on_closed_Collect_le [OF continuous_on_const f]) - apply (simp add: compact compact_imp_closed) - done - have disj: "(A j) \ (B j) = {}" - using e by (auto simp: A_def B_def field_simps) - have "\f \ R. f ` s \ {0..1} \ (\x \ A j. f x < e/n) \ (\x \ B j. f x > 1 - e/n)" - apply (rule two) - using e A B disj ngt - apply simp_all - done - } - then obtain xf where xfR: "\j. xf j \ R" and xf01: "\j. xf j ` s \ {0..1}" - and xfA: "\x j. x \ A j \ xf j x < e/n" - and xfB: "\x j. x \ B j \ xf j x > 1 - e/n" - by metis - define g where [abs_def]: "g x = e * (\i\n. xf i x)" for x - have gR: "g \ R" - unfolding g_def by (fast intro: mult const setsum xfR) - have gge0: "\x. x \ s \ g x \ 0" - using e xf01 by (simp add: g_def zero_le_mult_iff image_subset_iff setsum_nonneg) - have A0: "A 0 = {}" - using fpos e by (fastforce simp: A_def) - have An: "A n = s" - using e ngt f normf_upper by (fastforce simp: A_def field_simps of_nat_diff) - have Asub: "A j \ A i" if "i\j" for i j - using e that apply (clarsimp simp: A_def) - apply (erule order_trans, simp) - done - { fix t - assume t: "t \ s" - define j where "j = (LEAST j. t \ A j)" - have jn: "j \ n" - using t An by (simp add: Least_le j_def) - have Aj: "t \ A j" - using t An by (fastforce simp add: j_def intro: LeastI) - then have Ai: "t \ A i" if "i\j" for i - using Asub [OF that] by blast - then have fj1: "f t \ (j - 1/3)*e" - by (simp add: A_def) - then have Anj: "t \ A i" if "ii - apply (simp add: j_def) - using not_less_Least by blast - have j1: "1 \ j" - using A0 Aj j_def not_less_eq_eq by (fastforce simp add: j_def) - then have Anj: "t \ A (j-1)" - using Least_le by (fastforce simp add: j_def) - then have fj2: "(j - 4/3)*e < f t" - using j1 t by (simp add: A_def of_nat_diff) - have ***: "xf i t \ e/n" if "i\j" for i - using xfA [OF Ai] that by (simp add: less_eq_real_def) - { fix i - assume "i+2 \ j" - then obtain d where "i+2+d = j" - using le_Suc_ex that by blast - then have "t \ B i" - using Anj e ge_fx [OF t] \1 \ n\ fpos [OF t] t - apply (simp add: A_def B_def) - apply (clarsimp simp add: field_simps of_nat_diff not_le of_nat_Suc) - apply (rule order_trans [of _ "e * 2 + (e * (real d * 3) + e * (real i * 3))"]) - apply auto - done - then have "xf i t > 1 - e/n" - by (rule xfB) - } note **** = this - have xf_le1: "\i. xf i t \ 1" - using xf01 t by force - have "g t = e * (\ii=j..n. xf i t)" - using j1 jn e - apply (simp add: g_def distrib_left [symmetric]) - apply (subst setsum.union_disjoint [symmetric]) - apply (auto simp: ivl_disj_un) - done - also have "... \ e*j + e * ((Suc n - j)*e/n)" - apply (rule add_mono) - apply (simp_all only: mult_le_cancel_left_pos e) - apply (rule setsum_bounded_above [OF xf_le1, where A = "lessThan j", simplified]) - using setsum_bounded_above [of "{j..n}" "\i. xf i t", OF ***] - apply simp - done - also have "... \ j*e + e*(n - j + 1)*e/n " - using \1 \ n\ e by (simp add: field_simps del: of_nat_Suc) - also have "... \ j*e + e*e" - using \1 \ n\ e j1 by (simp add: field_simps del: of_nat_Suc) - also have "... < (j + 1/3)*e" - using e by (auto simp: field_simps) - finally have gj1: "g t < (j + 1 / 3) * e" . - have gj2: "(j - 4/3)*e < g t" - proof (cases "2 \ j") - case False - then have "j=1" using j1 by simp - with t gge0 e show ?thesis by force - next - case True - then have "(j - 4/3)*e < (j-1)*e - e^2" - using e by (auto simp: of_nat_diff algebra_simps power2_eq_square) - also have "... < (j-1)*e - ((j - 1)/n) * e^2" - using e True jn by (simp add: power2_eq_square field_simps) - also have "... = e * (j-1) * (1 - e/n)" - by (simp add: power2_eq_square field_simps) - also have "... \ e * (\i\j-2. xf i t)" - using e - apply simp - apply (rule order_trans [OF _ setsum_bounded_below [OF less_imp_le [OF ****]]]) - using True - apply (simp_all add: of_nat_Suc of_nat_diff) - done - also have "... \ g t" - using jn e - using e xf01 t - apply (simp add: g_def zero_le_mult_iff image_subset_iff setsum_nonneg) - apply (rule Groups_Big.setsum_mono2, auto) - done - finally show ?thesis . - qed - have "\f t - g t\ < 2 * e" - using fj1 fj2 gj1 gj2 by (simp add: abs_less_iff field_simps) - } - then show ?thesis - by (rule_tac x=g in bexI) (auto intro: gR) -qed - -text\The ``unpretentious'' formulation\ -lemma (in function_ring_on) Stone_Weierstrass_basic: - assumes f: "continuous_on s f" and e: "e > 0" - shows "\g \ R. \x\s. \f x - g x\ < e" -proof - - have "\g \ R. \x\s. \(f x + normf f) - g x\ < 2 * min (e/2) (1/4)" - apply (rule Stone_Weierstrass_special) - apply (rule Limits.continuous_on_add [OF f Topological_Spaces.continuous_on_const]) - using normf_upper [OF f] apply force - apply (simp add: e, linarith) - done - then obtain g where "g \ R" "\x\s. \g x - (f x + normf f)\ < e" - by force - then show ?thesis - apply (rule_tac x="\x. g x - normf f" in bexI) - apply (auto simp: algebra_simps intro: diff const) - done -qed - - -theorem (in function_ring_on) Stone_Weierstrass: - assumes f: "continuous_on s f" - shows "\F\UNIV \ R. LIM n sequentially. F n :> uniformly_on s f" -proof - - { fix e::real - assume e: "0 < e" - then obtain N::nat where N: "0 < N" "0 < inverse N" "inverse N < e" - by (auto simp: real_arch_inverse [of e]) - { fix n :: nat and x :: 'a and g :: "'a \ real" - assume n: "N \ n" "\x\s. \f x - g x\ < 1 / (1 + real n)" - assume x: "x \ s" - have "\ real (Suc n) < inverse e" - using \N \ n\ N using less_imp_inverse_less by force - then have "1 / (1 + real n) \ e" - using e by (simp add: field_simps of_nat_Suc) - then have "\f x - g x\ < e" - using n(2) x by auto - } note * = this - have "\\<^sub>F n in sequentially. \x\s. \f x - (SOME g. g \ R \ (\x\s. \f x - g x\ < 1 / (1 + real n))) x\ < e" - apply (rule eventually_sequentiallyI [of N]) - apply (auto intro: someI2_bex [OF Stone_Weierstrass_basic [OF f]] *) - done - } then - show ?thesis - apply (rule_tac x="\n::nat. SOME g. g \ R \ (\x\s. \f x - g x\ < 1 / (1 + n))" in bexI) - prefer 2 apply (force intro: someI2_bex [OF Stone_Weierstrass_basic [OF f]]) - unfolding uniform_limit_iff - apply (auto simp: dist_norm abs_minus_commute) - done -qed - -text\A HOL Light formulation\ -corollary Stone_Weierstrass_HOL: - fixes R :: "('a::t2_space \ real) set" and s :: "'a set" - assumes "compact s" "\c. P(\x. c::real)" - "\f. P f \ continuous_on s f" - "\f g. P(f) \ P(g) \ P(\x. f x + g x)" "\f g. P(f) \ P(g) \ P(\x. f x * g x)" - "\x y. x \ s \ y \ s \ ~(x = y) \ \f. P(f) \ ~(f x = f y)" - "continuous_on s f" - "0 < e" - shows "\g. P(g) \ (\x \ s. \f x - g x\ < e)" -proof - - interpret PR: function_ring_on "Collect P" - apply unfold_locales - using assms - by auto - show ?thesis - using PR.Stone_Weierstrass_basic [OF \continuous_on s f\ \0 < e\] - by blast -qed - - -subsection \Polynomial functions\ - -inductive real_polynomial_function :: "('a::real_normed_vector \ real) \ bool" where - linear: "bounded_linear f \ real_polynomial_function f" - | const: "real_polynomial_function (\x. c)" - | add: "\real_polynomial_function f; real_polynomial_function g\ \ real_polynomial_function (\x. f x + g x)" - | mult: "\real_polynomial_function f; real_polynomial_function g\ \ real_polynomial_function (\x. f x * g x)" - -declare real_polynomial_function.intros [intro] - -definition polynomial_function :: "('a::real_normed_vector \ 'b::real_normed_vector) \ bool" - where - "polynomial_function p \ (\f. bounded_linear f \ real_polynomial_function (f o p))" - -lemma real_polynomial_function_eq: "real_polynomial_function p = polynomial_function p" -unfolding polynomial_function_def -proof - assume "real_polynomial_function p" - then show " \f. bounded_linear f \ real_polynomial_function (f \ p)" - proof (induction p rule: real_polynomial_function.induct) - case (linear h) then show ?case - by (auto simp: bounded_linear_compose real_polynomial_function.linear) - next - case (const h) then show ?case - by (simp add: real_polynomial_function.const) - next - case (add h) then show ?case - by (force simp add: bounded_linear_def linear_add real_polynomial_function.add) - next - case (mult h) then show ?case - by (force simp add: real_bounded_linear const real_polynomial_function.mult) - qed -next - assume [rule_format, OF bounded_linear_ident]: "\f. bounded_linear f \ real_polynomial_function (f \ p)" - then show "real_polynomial_function p" - by (simp add: o_def) -qed - -lemma polynomial_function_const [iff]: "polynomial_function (\x. c)" - by (simp add: polynomial_function_def o_def const) - -lemma polynomial_function_bounded_linear: - "bounded_linear f \ polynomial_function f" - by (simp add: polynomial_function_def o_def bounded_linear_compose real_polynomial_function.linear) - -lemma polynomial_function_id [iff]: "polynomial_function(\x. x)" - by (simp add: polynomial_function_bounded_linear) - -lemma polynomial_function_add [intro]: - "\polynomial_function f; polynomial_function g\ \ polynomial_function (\x. f x + g x)" - by (auto simp: polynomial_function_def bounded_linear_def linear_add real_polynomial_function.add o_def) - -lemma polynomial_function_mult [intro]: - assumes f: "polynomial_function f" and g: "polynomial_function g" - shows "polynomial_function (\x. f x *\<^sub>R g x)" - using g - apply (auto simp: polynomial_function_def bounded_linear_def Real_Vector_Spaces.linear.scaleR const real_polynomial_function.mult o_def) - apply (rule mult) - using f - apply (auto simp: real_polynomial_function_eq) - done - -lemma polynomial_function_cmul [intro]: - assumes f: "polynomial_function f" - shows "polynomial_function (\x. c *\<^sub>R f x)" - by (rule polynomial_function_mult [OF polynomial_function_const f]) - -lemma polynomial_function_minus [intro]: - assumes f: "polynomial_function f" - shows "polynomial_function (\x. - f x)" - using polynomial_function_cmul [OF f, of "-1"] by simp - -lemma polynomial_function_diff [intro]: - "\polynomial_function f; polynomial_function g\ \ polynomial_function (\x. f x - g x)" - unfolding add_uminus_conv_diff [symmetric] - by (metis polynomial_function_add polynomial_function_minus) - -lemma polynomial_function_setsum [intro]: - "\finite I; \i. i \ I \ polynomial_function (\x. f x i)\ \ polynomial_function (\x. setsum (f x) I)" -by (induct I rule: finite_induct) auto - -lemma real_polynomial_function_minus [intro]: - "real_polynomial_function f \ real_polynomial_function (\x. - f x)" - using polynomial_function_minus [of f] - by (simp add: real_polynomial_function_eq) - -lemma real_polynomial_function_diff [intro]: - "\real_polynomial_function f; real_polynomial_function g\ \ real_polynomial_function (\x. f x - g x)" - using polynomial_function_diff [of f] - by (simp add: real_polynomial_function_eq) - -lemma real_polynomial_function_setsum [intro]: - "\finite I; \i. i \ I \ real_polynomial_function (\x. f x i)\ \ real_polynomial_function (\x. setsum (f x) I)" - using polynomial_function_setsum [of I f] - by (simp add: real_polynomial_function_eq) - -lemma real_polynomial_function_power [intro]: - "real_polynomial_function f \ real_polynomial_function (\x. f x ^ n)" - by (induct n) (simp_all add: const mult) - -lemma real_polynomial_function_compose [intro]: - assumes f: "polynomial_function f" and g: "real_polynomial_function g" - shows "real_polynomial_function (g o f)" - using g - apply (induction g rule: real_polynomial_function.induct) - using f - apply (simp_all add: polynomial_function_def o_def const add mult) - done - -lemma polynomial_function_compose [intro]: - assumes f: "polynomial_function f" and g: "polynomial_function g" - shows "polynomial_function (g o f)" - using g real_polynomial_function_compose [OF f] - by (auto simp: polynomial_function_def o_def) - -lemma setsum_max_0: - fixes x::real (*in fact "'a::comm_ring_1"*) - shows "(\i = 0..max m n. x^i * (if i \ m then a i else 0)) = (\i = 0..m. x^i * a i)" -proof - - have "(\i = 0..max m n. x^i * (if i \ m then a i else 0)) = (\i = 0..max m n. (if i \ m then x^i * a i else 0))" - by (auto simp: algebra_simps intro: setsum.cong) - also have "... = (\i = 0..m. (if i \ m then x^i * a i else 0))" - by (rule setsum.mono_neutral_right) auto - also have "... = (\i = 0..m. x^i * a i)" - by (auto simp: algebra_simps intro: setsum.cong) - finally show ?thesis . -qed - -lemma real_polynomial_function_imp_setsum: - assumes "real_polynomial_function f" - shows "\a n::nat. f = (\x. \i=0..n. a i * x ^ i)" -using assms -proof (induct f) - case (linear f) - then show ?case - apply (clarsimp simp add: real_bounded_linear) - apply (rule_tac x="\i. if i=0 then 0 else c" in exI) - apply (rule_tac x=1 in exI) - apply (simp add: mult_ac) - done -next - case (const c) - show ?case - apply (rule_tac x="\i. c" in exI) - apply (rule_tac x=0 in exI) - apply (auto simp: mult_ac of_nat_Suc) - done - case (add f1 f2) - then obtain a1 n1 a2 n2 where - "f1 = (\x. \i = 0..n1. a1 i * x ^ i)" "f2 = (\x. \i = 0..n2. a2 i * x ^ i)" - by auto - then show ?case - apply (rule_tac x="\i. (if i \ n1 then a1 i else 0) + (if i \ n2 then a2 i else 0)" in exI) - apply (rule_tac x="max n1 n2" in exI) - using setsum_max_0 [where m=n1 and n=n2] setsum_max_0 [where m=n2 and n=n1] - apply (simp add: setsum.distrib algebra_simps max.commute) - done - case (mult f1 f2) - then obtain a1 n1 a2 n2 where - "f1 = (\x. \i = 0..n1. a1 i * x ^ i)" "f2 = (\x. \i = 0..n2. a2 i * x ^ i)" - by auto - then obtain b1 b2 where - "f1 = (\x. \i = 0..n1. b1 i * x ^ i)" "f2 = (\x. \i = 0..n2. b2 i * x ^ i)" - "b1 = (\i. if i\n1 then a1 i else 0)" "b2 = (\i. if i\n2 then a2 i else 0)" - by auto - then show ?case - apply (rule_tac x="\i. \k\i. b1 k * b2 (i - k)" in exI) - apply (rule_tac x="n1+n2" in exI) - using polynomial_product [of n1 b1 n2 b2] - apply (simp add: Set_Interval.atLeast0AtMost) - done -qed - -lemma real_polynomial_function_iff_setsum: - "real_polynomial_function f \ (\a n::nat. f = (\x. \i=0..n. a i * x ^ i))" - apply (rule iffI) - apply (erule real_polynomial_function_imp_setsum) - apply (auto simp: linear mult const real_polynomial_function_power real_polynomial_function_setsum) - done - -lemma polynomial_function_iff_Basis_inner: - fixes f :: "'a::real_normed_vector \ 'b::euclidean_space" - shows "polynomial_function f \ (\b\Basis. real_polynomial_function (\x. inner (f x) b))" - (is "?lhs = ?rhs") -unfolding polynomial_function_def -proof (intro iffI allI impI) - assume "\h. bounded_linear h \ real_polynomial_function (h \ f)" - then show ?rhs - by (force simp add: bounded_linear_inner_left o_def) -next - fix h :: "'b \ real" - assume rp: "\b\Basis. real_polynomial_function (\x. f x \ b)" and h: "bounded_linear h" - have "real_polynomial_function (h \ (\x. \b\Basis. (f x \ b) *\<^sub>R b))" - apply (rule real_polynomial_function_compose [OF _ linear [OF h]]) - using rp - apply (auto simp: real_polynomial_function_eq polynomial_function_mult) - done - then show "real_polynomial_function (h \ f)" - by (simp add: euclidean_representation_setsum_fun) -qed - -subsection \Stone-Weierstrass theorem for polynomial functions\ - -text\First, we need to show that they are continous, differentiable and separable.\ - -lemma continuous_real_polymonial_function: - assumes "real_polynomial_function f" - shows "continuous (at x) f" -using assms -by (induct f) (auto simp: linear_continuous_at) - -lemma continuous_polymonial_function: - fixes f :: "'a::real_normed_vector \ 'b::euclidean_space" - assumes "polynomial_function f" - shows "continuous (at x) f" - apply (rule euclidean_isCont) - using assms apply (simp add: polynomial_function_iff_Basis_inner) - apply (force dest: continuous_real_polymonial_function intro: isCont_scaleR) - done - -lemma continuous_on_polymonial_function: - fixes f :: "'a::real_normed_vector \ 'b::euclidean_space" - assumes "polynomial_function f" - shows "continuous_on s f" - using continuous_polymonial_function [OF assms] continuous_at_imp_continuous_on - by blast - -lemma has_real_derivative_polynomial_function: - assumes "real_polynomial_function p" - shows "\p'. real_polynomial_function p' \ - (\x. (p has_real_derivative (p' x)) (at x))" -using assms -proof (induct p) - case (linear p) - then show ?case - by (force simp: real_bounded_linear const intro!: derivative_eq_intros) -next - case (const c) - show ?case - by (rule_tac x="\x. 0" in exI) auto - case (add f1 f2) - then obtain p1 p2 where - "real_polynomial_function p1" "\x. (f1 has_real_derivative p1 x) (at x)" - "real_polynomial_function p2" "\x. (f2 has_real_derivative p2 x) (at x)" - by auto - then show ?case - apply (rule_tac x="\x. p1 x + p2 x" in exI) - apply (auto intro!: derivative_eq_intros) - done - case (mult f1 f2) - then obtain p1 p2 where - "real_polynomial_function p1" "\x. (f1 has_real_derivative p1 x) (at x)" - "real_polynomial_function p2" "\x. (f2 has_real_derivative p2 x) (at x)" - by auto - then show ?case - using mult - apply (rule_tac x="\x. f1 x * p2 x + f2 x * p1 x" in exI) - apply (auto intro!: derivative_eq_intros) - done -qed - -lemma has_vector_derivative_polynomial_function: - fixes p :: "real \ 'a::euclidean_space" - assumes "polynomial_function p" - shows "\p'. polynomial_function p' \ - (\x. (p has_vector_derivative (p' x)) (at x))" -proof - - { fix b :: 'a - assume "b \ Basis" - then - obtain p' where p': "real_polynomial_function p'" and pd: "\x. ((\x. p x \ b) has_real_derivative p' x) (at x)" - using assms [unfolded polynomial_function_iff_Basis_inner, rule_format] \b \ Basis\ - has_real_derivative_polynomial_function - by blast - have "\q. polynomial_function q \ (\x. ((\u. (p u \ b) *\<^sub>R b) has_vector_derivative q x) (at x))" - apply (rule_tac x="\x. p' x *\<^sub>R b" in exI) - using \b \ Basis\ p' - apply (simp add: polynomial_function_iff_Basis_inner inner_Basis) - apply (auto intro: derivative_eq_intros pd) - done - } - then obtain qf where qf: - "\b. b \ Basis \ polynomial_function (qf b)" - "\b x. b \ Basis \ ((\u. (p u \ b) *\<^sub>R b) has_vector_derivative qf b x) (at x)" - by metis - show ?thesis - apply (subst euclidean_representation_setsum_fun [of p, symmetric]) - apply (rule_tac x="\x. \b\Basis. qf b x" in exI) - apply (auto intro: has_vector_derivative_setsum qf) - done -qed - -lemma real_polynomial_function_separable: - fixes x :: "'a::euclidean_space" - assumes "x \ y" shows "\f. real_polynomial_function f \ f x \ f y" -proof - - have "real_polynomial_function (\u. \b\Basis. (inner (x-u) b)^2)" - apply (rule real_polynomial_function_setsum) - apply (auto simp: algebra_simps real_polynomial_function_power real_polynomial_function_diff - const linear bounded_linear_inner_left) - done - then show ?thesis - apply (intro exI conjI, assumption) - using assms - apply (force simp add: euclidean_eq_iff [of x y] setsum_nonneg_eq_0_iff algebra_simps) - done -qed - -lemma Stone_Weierstrass_real_polynomial_function: - fixes f :: "'a::euclidean_space \ real" - assumes "compact s" "continuous_on s f" "0 < e" - shows "\g. real_polynomial_function g \ (\x \ s. \f x - g x\ < e)" -proof - - interpret PR: function_ring_on "Collect real_polynomial_function" - apply unfold_locales - using assms continuous_on_polymonial_function real_polynomial_function_eq - apply (auto intro: real_polynomial_function_separable) - done - show ?thesis - using PR.Stone_Weierstrass_basic [OF \continuous_on s f\ \0 < e\] - by blast -qed - -lemma Stone_Weierstrass_polynomial_function: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes s: "compact s" - and f: "continuous_on s f" - and e: "0 < e" - shows "\g. polynomial_function g \ (\x \ s. norm(f x - g x) < e)" -proof - - { fix b :: 'b - assume "b \ Basis" - have "\p. real_polynomial_function p \ (\x \ s. \f x \ b - p x\ < e / DIM('b))" - apply (rule exE [OF Stone_Weierstrass_real_polynomial_function [OF s _, of "\x. f x \ b" "e / card Basis"]]) - using e f - apply (auto simp: Euclidean_Space.DIM_positive intro: continuous_intros) - done - } - then obtain pf where pf: - "\b. b \ Basis \ real_polynomial_function (pf b) \ (\x \ s. \f x \ b - pf b x\ < e / DIM('b))" - apply (rule bchoice [rule_format, THEN exE]) - apply assumption - apply (force simp add: intro: that) - done - have "polynomial_function (\x. \b\Basis. pf b x *\<^sub>R b)" - using pf - by (simp add: polynomial_function_setsum polynomial_function_mult real_polynomial_function_eq) - moreover - { fix x - assume "x \ s" - have "norm (\b\Basis. (f x \ b) *\<^sub>R b - pf b x *\<^sub>R b) \ (\b\Basis. norm ((f x \ b) *\<^sub>R b - pf b x *\<^sub>R b))" - by (rule norm_setsum) - also have "... < of_nat DIM('b) * (e / DIM('b))" - apply (rule setsum_bounded_above_strict) - apply (simp add: Real_Vector_Spaces.scaleR_diff_left [symmetric] pf \x \ s\) - apply (rule DIM_positive) - done - also have "... = e" - using DIM_positive by (simp add: field_simps) - finally have "norm (\b\Basis. (f x \ b) *\<^sub>R b - pf b x *\<^sub>R b) < e" . - } - ultimately - show ?thesis - apply (subst euclidean_representation_setsum_fun [of f, symmetric]) - apply (rule_tac x="\x. \b\Basis. pf b x *\<^sub>R b" in exI) - apply (auto simp: setsum_subtractf [symmetric]) - done -qed - - -subsection\Polynomial functions as paths\ - -text\One application is to pick a smooth approximation to a path, -or just pick a smooth path anyway in an open connected set\ - -lemma path_polynomial_function: - fixes g :: "real \ 'b::euclidean_space" - shows "polynomial_function g \ path g" - by (simp add: path_def continuous_on_polymonial_function) - -lemma path_approx_polynomial_function: - fixes g :: "real \ 'b::euclidean_space" - assumes "path g" "0 < e" - shows "\p. polynomial_function p \ - pathstart p = pathstart g \ - pathfinish p = pathfinish g \ - (\t \ {0..1}. norm(p t - g t) < e)" -proof - - obtain q where poq: "polynomial_function q" and noq: "\x. x \ {0..1} \ norm (g x - q x) < e/4" - using Stone_Weierstrass_polynomial_function [of "{0..1}" g "e/4"] assms - by (auto simp: path_def) - have pf: "polynomial_function (\t. q t + (g 0 - q 0) + t *\<^sub>R (g 1 - q 1 - (g 0 - q 0)))" - by (force simp add: poq) - have *: "\t. t \ {0..1} \ norm (((q t - g t) + (g 0 - q 0)) + (t *\<^sub>R (g 1 - q 1) + t *\<^sub>R (q 0 - g 0))) < (e/4 + e/4) + (e/4+e/4)" - apply (intro Real_Vector_Spaces.norm_add_less) - using noq - apply (auto simp: norm_minus_commute intro: le_less_trans [OF mult_left_le_one_le noq] simp del: less_divide_eq_numeral1) - done - show ?thesis - apply (intro exI conjI) - apply (rule pf) - using * - apply (auto simp add: pathstart_def pathfinish_def algebra_simps) - done -qed - -lemma connected_open_polynomial_connected: - fixes s :: "'a::euclidean_space set" - assumes s: "open s" "connected s" - and "x \ s" "y \ s" - shows "\g. polynomial_function g \ path_image g \ s \ - pathstart g = x \ pathfinish g = y" -proof - - have "path_connected s" using assms - by (simp add: connected_open_path_connected) - with \x \ s\ \y \ s\ obtain p where p: "path p" "path_image p \ s" "pathstart p = x" "pathfinish p = y" - by (force simp: path_connected_def) - have "\e. 0 < e \ (\x \ path_image p. ball x e \ s)" - proof (cases "s = UNIV") - case True then show ?thesis - by (simp add: gt_ex) - next - case False - then have "- s \ {}" by blast - then show ?thesis - apply (rule_tac x="setdist (path_image p) (-s)" in exI) - using s p - apply (simp add: setdist_gt_0_compact_closed compact_path_image open_closed) - using setdist_le_dist [of _ "path_image p" _ "-s"] - by fastforce - qed - then obtain e where "0 < e"and eb: "\x. x \ path_image p \ ball x e \ s" - by auto - show ?thesis - using path_approx_polynomial_function [OF \path p\ \0 < e\] - apply clarify - apply (intro exI conjI, assumption) - using p - apply (fastforce simp add: dist_norm path_image_def norm_minus_commute intro: eb [THEN subsetD])+ - done -qed - -hide_fact linear add mult const - -end diff -r bbcb05504fdc -r bd218a9320b5 src/HOL/Multivariate_Analysis/Weierstrass_Theorems.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Multivariate_Analysis/Weierstrass_Theorems.thy Thu Aug 04 19:36:31 2016 +0200 @@ -0,0 +1,1218 @@ +section \The Bernstein-Weierstrass and Stone-Weierstrass Theorems\ + +text\By L C Paulson (2015)\ + +theory Weierstrass_Theorems +imports Uniform_Limit Path_Connected +begin + +subsection \Bernstein polynomials\ + +definition Bernstein :: "[nat,nat,real] \ real" where + "Bernstein n k x \ of_nat (n choose k) * x ^ k * (1 - x) ^ (n - k)" + +lemma Bernstein_nonneg: "\0 \ x; x \ 1\ \ 0 \ Bernstein n k x" + by (simp add: Bernstein_def) + +lemma Bernstein_pos: "\0 < x; x < 1; k \ n\ \ 0 < Bernstein n k x" + by (simp add: Bernstein_def) + +lemma sum_Bernstein [simp]: "(\ k = 0..n. Bernstein n k x) = 1" + using binomial_ring [of x "1-x" n] + by (simp add: Bernstein_def) + +lemma binomial_deriv1: + "(\k=0..n. (of_nat k * of_nat (n choose k)) * a^(k-1) * b^(n-k)) = real_of_nat n * (a+b) ^ (n-1)" + apply (rule DERIV_unique [where f = "\a. (a+b)^n" and x=a]) + apply (subst binomial_ring) + apply (rule derivative_eq_intros setsum.cong | simp)+ + done + +lemma binomial_deriv2: + "(\k=0..n. (of_nat k * of_nat (k-1) * of_nat (n choose k)) * a^(k-2) * b^(n-k)) = + of_nat n * of_nat (n-1) * (a+b::real) ^ (n-2)" + apply (rule DERIV_unique [where f = "\a. of_nat n * (a+b::real) ^ (n-1)" and x=a]) + apply (subst binomial_deriv1 [symmetric]) + apply (rule derivative_eq_intros setsum.cong | simp add: Num.numeral_2_eq_2)+ + done + +lemma sum_k_Bernstein [simp]: "(\k = 0..n. real k * Bernstein n k x) = of_nat n * x" + apply (subst binomial_deriv1 [of n x "1-x", simplified, symmetric]) + apply (simp add: setsum_left_distrib) + apply (auto simp: Bernstein_def algebra_simps realpow_num_eq_if intro!: setsum.cong) + done + +lemma sum_kk_Bernstein [simp]: "(\ k = 0..n. real k * (real k - 1) * Bernstein n k x) = real n * (real n - 1) * x\<^sup>2" +proof - + have "(\ k = 0..n. real k * (real k - 1) * Bernstein n k x) = real_of_nat n * real_of_nat (n - Suc 0) * x\<^sup>2" + apply (subst binomial_deriv2 [of n x "1-x", simplified, symmetric]) + apply (simp add: setsum_left_distrib) + apply (rule setsum.cong [OF refl]) + apply (simp add: Bernstein_def power2_eq_square algebra_simps) + apply (rename_tac k) + apply (subgoal_tac "k = 0 \ k = 1 \ (\k'. k = Suc (Suc k'))") + apply (force simp add: field_simps of_nat_Suc power2_eq_square) + by presburger + also have "... = n * (n - 1) * x\<^sup>2" + by auto + finally show ?thesis + by auto +qed + +subsection \Explicit Bernstein version of the 1D Weierstrass approximation theorem\ + +lemma Bernstein_Weierstrass: + fixes f :: "real \ real" + assumes contf: "continuous_on {0..1} f" and e: "0 < e" + shows "\N. \n x. N \ n \ x \ {0..1} + \ \f x - (\k = 0..n. f(k/n) * Bernstein n k x)\ < e" +proof - + have "bounded (f ` {0..1})" + using compact_continuous_image compact_imp_bounded contf by blast + then obtain M where M: "\x. 0 \ x \ x \ 1 \ \f x\ \ M" + by (force simp add: bounded_iff) + then have Mge0: "0 \ M" by force + have ucontf: "uniformly_continuous_on {0..1} f" + using compact_uniformly_continuous contf by blast + then obtain d where d: "d>0" "\x x'. \ x \ {0..1}; x' \ {0..1}; \x' - x\ < d\ \ \f x' - f x\ < e/2" + apply (rule uniformly_continuous_onE [where e = "e/2"]) + using e by (auto simp: dist_norm) + { fix n::nat and x::real + assume n: "Suc (nat\4*M/(e*d\<^sup>2)\) \ n" and x: "0 \ x" "x \ 1" + have "0 < n" using n by simp + have ed0: "- (e * d\<^sup>2) < 0" + using e \0 by simp + also have "... \ M * 4" + using \0\M\ by simp + finally have [simp]: "real_of_int (nat \4 * M / (e * d\<^sup>2)\) = real_of_int \4 * M / (e * d\<^sup>2)\" + using \0\M\ e \0 + by (simp add: of_nat_Suc field_simps) + have "4*M/(e*d\<^sup>2) + 1 \ real (Suc (nat\4*M/(e*d\<^sup>2)\))" + by (simp add: of_nat_Suc real_nat_ceiling_ge) + also have "... \ real n" + using n by (simp add: of_nat_Suc field_simps) + finally have nbig: "4*M/(e*d\<^sup>2) + 1 \ real n" . + have sum_bern: "(\k = 0..n. (x - k/n)\<^sup>2 * Bernstein n k x) = x * (1 - x) / n" + proof - + have *: "\a b x::real. (a - b)\<^sup>2 * x = a * (a - 1) * x + (1 - 2 * b) * a * x + b * b * x" + by (simp add: algebra_simps power2_eq_square) + have "(\ k = 0..n. (k - n * x)\<^sup>2 * Bernstein n k x) = n * x * (1 - x)" + apply (simp add: * setsum.distrib) + apply (simp add: setsum_right_distrib [symmetric] mult.assoc) + apply (simp add: algebra_simps power2_eq_square) + done + then have "(\ k = 0..n. (k - n * x)\<^sup>2 * Bernstein n k x)/n^2 = x * (1 - x) / n" + by (simp add: power2_eq_square) + then show ?thesis + using n by (simp add: setsum_divide_distrib divide_simps mult.commute power2_commute) + qed + { fix k + assume k: "k \ n" + then have kn: "0 \ k / n" "k / n \ 1" + by (auto simp: divide_simps) + consider (lessd) "\x - k / n\ < d" | (ged) "d \ \x - k / n\" + by linarith + then have "\(f x - f (k/n))\ \ e/2 + 2 * M / d\<^sup>2 * (x - k/n)\<^sup>2" + proof cases + case lessd + then have "\(f x - f (k/n))\ < e/2" + using d x kn by (simp add: abs_minus_commute) + also have "... \ (e/2 + 2 * M / d\<^sup>2 * (x - k/n)\<^sup>2)" + using Mge0 d by simp + finally show ?thesis by simp + next + case ged + then have dle: "d\<^sup>2 \ (x - k/n)\<^sup>2" + by (metis d(1) less_eq_real_def power2_abs power_mono) + have "\(f x - f (k/n))\ \ \f x\ + \f (k/n)\" + by (rule abs_triangle_ineq4) + also have "... \ M+M" + by (meson M add_mono_thms_linordered_semiring(1) kn x) + also have "... \ 2 * M * ((x - k/n)\<^sup>2 / d\<^sup>2)" + apply simp + apply (rule Rings.ordered_semiring_class.mult_left_mono [of 1 "((x - k/n)\<^sup>2 / d\<^sup>2)", simplified]) + using dle \d>0\ \M\0\ by auto + also have "... \ e/2 + 2 * M / d\<^sup>2 * (x - k/n)\<^sup>2" + using e by simp + finally show ?thesis . + qed + } note * = this + have "\f x - (\ k = 0..n. f(k / n) * Bernstein n k x)\ \ \\ k = 0..n. (f x - f(k / n)) * Bernstein n k x\" + by (simp add: setsum_subtractf setsum_right_distrib [symmetric] algebra_simps) + also have "... \ (\ k = 0..n. (e/2 + (2 * M / d\<^sup>2) * (x - k / n)\<^sup>2) * Bernstein n k x)" + apply (rule order_trans [OF setsum_abs setsum_mono]) + using * + apply (simp add: abs_mult Bernstein_nonneg x mult_right_mono) + done + also have "... \ e/2 + (2 * M) / (d\<^sup>2 * n)" + apply (simp only: setsum.distrib Rings.semiring_class.distrib_right setsum_right_distrib [symmetric] mult.assoc sum_bern) + using \d>0\ x + apply (simp add: divide_simps Mge0 mult_le_one mult_left_le) + done + also have "... < e" + apply (simp add: field_simps) + using \d>0\ nbig e \n>0\ + apply (simp add: divide_simps algebra_simps) + using ed0 by linarith + finally have "\f x - (\k = 0..n. f (real k / real n) * Bernstein n k x)\ < e" . + } + then show ?thesis + by auto +qed + + +subsection \General Stone-Weierstrass theorem\ + +text\Source: +Bruno Brosowski and Frank Deutsch. +An Elementary Proof of the Stone-Weierstrass Theorem. +Proceedings of the American Mathematical Society +Volume 81, Number 1, January 1981. +DOI: 10.2307/2043993 http://www.jstor.org/stable/2043993\ + +locale function_ring_on = + fixes R :: "('a::t2_space \ real) set" and s :: "'a set" + assumes compact: "compact s" + assumes continuous: "f \ R \ continuous_on s f" + assumes add: "f \ R \ g \ R \ (\x. f x + g x) \ R" + assumes mult: "f \ R \ g \ R \ (\x. f x * g x) \ R" + assumes const: "(\_. c) \ R" + assumes separable: "x \ s \ y \ s \ x \ y \ \f\R. f x \ f y" + +begin + lemma minus: "f \ R \ (\x. - f x) \ R" + by (frule mult [OF const [of "-1"]]) simp + + lemma diff: "f \ R \ g \ R \ (\x. f x - g x) \ R" + unfolding diff_conv_add_uminus by (metis add minus) + + lemma power: "f \ R \ (\x. f x ^ n) \ R" + by (induct n) (auto simp: const mult) + + lemma setsum: "\finite I; \i. i \ I \ f i \ R\ \ (\x. \i \ I. f i x) \ R" + by (induct I rule: finite_induct; simp add: const add) + + lemma setprod: "\finite I; \i. i \ I \ f i \ R\ \ (\x. \i \ I. f i x) \ R" + by (induct I rule: finite_induct; simp add: const mult) + + definition normf :: "('a::t2_space \ real) \ real" + where "normf f \ SUP x:s. \f x\" + + lemma normf_upper: "\continuous_on s f; x \ s\ \ \f x\ \ normf f" + apply (simp add: normf_def) + apply (rule cSUP_upper, assumption) + by (simp add: bounded_imp_bdd_above compact compact_continuous_image compact_imp_bounded continuous_on_rabs) + + lemma normf_least: "s \ {} \ (\x. x \ s \ \f x\ \ M) \ normf f \ M" + by (simp add: normf_def cSUP_least) + +end + +lemma (in function_ring_on) one: + assumes U: "open U" and t0: "t0 \ s" "t0 \ U" and t1: "t1 \ s-U" + shows "\V. open V \ t0 \ V \ s \ V \ U \ + (\e>0. \f \ R. f ` s \ {0..1} \ (\t \ s \ V. f t < e) \ (\t \ s - U. f t > 1 - e))" +proof - + have "\pt \ R. pt t0 = 0 \ pt t > 0 \ pt ` s \ {0..1}" if t: "t \ s - U" for t + proof - + have "t \ t0" using t t0 by auto + then obtain g where g: "g \ R" "g t \ g t0" + using separable t0 by (metis Diff_subset subset_eq t) + define h where [abs_def]: "h x = g x - g t0" for x + have "h \ R" + unfolding h_def by (fast intro: g const diff) + then have hsq: "(\w. (h w)\<^sup>2) \ R" + by (simp add: power2_eq_square mult) + have "h t \ h t0" + by (simp add: h_def g) + then have "h t \ 0" + by (simp add: h_def) + then have ht2: "0 < (h t)^2" + by simp + also have "... \ normf (\w. (h w)\<^sup>2)" + using t normf_upper [where x=t] continuous [OF hsq] by force + finally have nfp: "0 < normf (\w. (h w)\<^sup>2)" . + define p where [abs_def]: "p x = (1 / normf (\w. (h w)\<^sup>2)) * (h x)^2" for x + have "p \ R" + unfolding p_def by (fast intro: hsq const mult) + moreover have "p t0 = 0" + by (simp add: p_def h_def) + moreover have "p t > 0" + using nfp ht2 by (simp add: p_def) + moreover have "\x. x \ s \ p x \ {0..1}" + using nfp normf_upper [OF continuous [OF hsq] ] by (auto simp: p_def) + ultimately show "\pt \ R. pt t0 = 0 \ pt t > 0 \ pt ` s \ {0..1}" + by auto + qed + then obtain pf where pf: "\t. t \ s-U \ pf t \ R \ pf t t0 = 0 \ pf t t > 0" + and pf01: "\t. t \ s-U \ pf t ` s \ {0..1}" + by metis + have com_sU: "compact (s-U)" + using compact closed_Int_compact U by (simp add: Diff_eq compact_Int_closed open_closed) + have "\t. t \ s-U \ \A. open A \ A \ s = {x\s. 0 < pf t x}" + apply (rule open_Collect_positive) + by (metis pf continuous) + then obtain Uf where Uf: "\t. t \ s-U \ open (Uf t) \ (Uf t) \ s = {x\s. 0 < pf t x}" + by metis + then have open_Uf: "\t. t \ s-U \ open (Uf t)" + by blast + have tUft: "\t. t \ s-U \ t \ Uf t" + using pf Uf by blast + then have *: "s-U \ (\x \ s-U. Uf x)" + by blast + obtain subU where subU: "subU \ s - U" "finite subU" "s - U \ (\x \ subU. Uf x)" + by (blast intro: that open_Uf compactE_image [OF com_sU _ *]) + then have [simp]: "subU \ {}" + using t1 by auto + then have cardp: "card subU > 0" using subU + by (simp add: card_gt_0_iff) + define p where [abs_def]: "p x = (1 / card subU) * (\t \ subU. pf t x)" for x + have pR: "p \ R" + unfolding p_def using subU pf by (fast intro: pf const mult setsum) + have pt0 [simp]: "p t0 = 0" + using subU pf by (auto simp: p_def intro: setsum.neutral) + have pt_pos: "p t > 0" if t: "t \ s-U" for t + proof - + obtain i where i: "i \ subU" "t \ Uf i" using subU t by blast + show ?thesis + using subU i t + apply (clarsimp simp: p_def divide_simps) + apply (rule setsum_pos2 [OF \finite subU\]) + using Uf t pf01 apply auto + apply (force elim!: subsetCE) + done + qed + have p01: "p x \ {0..1}" if t: "x \ s" for x + proof - + have "0 \ p x" + using subU cardp t + apply (simp add: p_def divide_simps setsum_nonneg) + apply (rule setsum_nonneg) + using pf01 by force + moreover have "p x \ 1" + using subU cardp t + apply (simp add: p_def divide_simps setsum_nonneg) + apply (rule setsum_bounded_above [where 'a=real and K=1, simplified]) + using pf01 by force + ultimately show ?thesis + by auto + qed + have "compact (p ` (s-U))" + by (meson Diff_subset com_sU compact_continuous_image continuous continuous_on_subset pR) + then have "open (- (p ` (s-U)))" + by (simp add: compact_imp_closed open_Compl) + moreover have "0 \ - (p ` (s-U))" + by (metis (no_types) ComplI image_iff not_less_iff_gr_or_eq pt_pos) + ultimately obtain delta0 where delta0: "delta0 > 0" "ball 0 delta0 \ - (p ` (s-U))" + by (auto simp: elim!: openE) + then have pt_delta: "\x. x \ s-U \ p x \ delta0" + by (force simp: ball_def dist_norm dest: p01) + define \ where "\ = delta0/2" + have "delta0 \ 1" using delta0 p01 [of t1] t1 + by (force simp: ball_def dist_norm dest: p01) + with delta0 have \01: "0 < \" "\ < 1" + by (auto simp: \_def) + have pt_\: "\x. x \ s-U \ p x \ \" + using pt_delta delta0 by (force simp: \_def) + have "\A. open A \ A \ s = {x\s. p x < \/2}" + by (rule open_Collect_less_Int [OF continuous [OF pR] continuous_on_const]) + then obtain V where V: "open V" "V \ s = {x\s. p x < \/2}" + by blast + define k where "k = nat\1/\\ + 1" + have "k>0" by (simp add: k_def) + have "k-1 \ 1/\" + using \01 by (simp add: k_def) + with \01 have "k \ (1+\)/\" + by (auto simp: algebra_simps add_divide_distrib) + also have "... < 2/\" + using \01 by (auto simp: divide_simps) + finally have k2\: "k < 2/\" . + have "1/\ < k" + using \01 unfolding k_def by linarith + with \01 k2\ have k\: "1 < k*\" "k*\ < 2" + by (auto simp: divide_simps) + define q where [abs_def]: "q n t = (1 - p t ^ n) ^ (k^n)" for n t + have qR: "q n \ R" for n + by (simp add: q_def const diff power pR) + have q01: "\n t. t \ s \ q n t \ {0..1}" + using p01 by (simp add: q_def power_le_one algebra_simps) + have qt0 [simp]: "\n. n>0 \ q n t0 = 1" + using t0 pf by (simp add: q_def power_0_left) + { fix t and n::nat + assume t: "t \ s \ V" + with \k>0\ V have "k * p t < k * \ / 2" + by force + then have "1 - (k * \ / 2)^n \ 1 - (k * p t)^n" + using \k>0\ p01 t by (simp add: power_mono) + also have "... \ q n t" + using Bernoulli_inequality [of "- ((p t)^n)" "k^n"] + apply (simp add: q_def) + by (metis IntE atLeastAtMost_iff p01 power_le_one power_mult_distrib t) + finally have "1 - (k * \ / 2) ^ n \ q n t" . + } note limitV = this + { fix t and n::nat + assume t: "t \ s - U" + with \k>0\ U have "k * \ \ k * p t" + by (simp add: pt_\) + with k\ have kpt: "1 < k * p t" + by (blast intro: less_le_trans) + have ptn_pos: "0 < p t ^ n" + using pt_pos [OF t] by simp + have ptn_le: "p t ^ n \ 1" + by (meson DiffE atLeastAtMost_iff p01 power_le_one t) + have "q n t = (1/(k^n * (p t)^n)) * (1 - p t ^ n) ^ (k^n) * k^n * (p t)^n" + using pt_pos [OF t] \k>0\ by (simp add: q_def) + also have "... \ (1/(k * (p t))^n) * (1 - p t ^ n) ^ (k^n) * (1 + k^n * (p t)^n)" + using pt_pos [OF t] \k>0\ + apply simp + apply (simp only: times_divide_eq_right [symmetric]) + apply (rule mult_left_mono [of "1::real", simplified]) + apply (simp_all add: power_mult_distrib) + apply (rule zero_le_power) + using ptn_le by linarith + also have "... \ (1/(k * (p t))^n) * (1 - p t ^ n) ^ (k^n) * (1 + (p t)^n) ^ (k^n)" + apply (rule mult_left_mono [OF Bernoulli_inequality [of "p t ^ n" "k^n"]]) + using \k>0\ ptn_pos ptn_le + apply (auto simp: power_mult_distrib) + done + also have "... = (1/(k * (p t))^n) * (1 - p t ^ (2*n)) ^ (k^n)" + using pt_pos [OF t] \k>0\ + by (simp add: algebra_simps power_mult power2_eq_square power_mult_distrib [symmetric]) + also have "... \ (1/(k * (p t))^n) * 1" + apply (rule mult_left_mono [OF power_le_one]) + using pt_pos \k>0\ p01 power_le_one t apply auto + done + also have "... \ (1 / (k*\))^n" + using \k>0\ \01 power_mono pt_\ t + by (fastforce simp: field_simps) + finally have "q n t \ (1 / (real k * \)) ^ n " . + } note limitNonU = this + define NN + where "NN e = 1 + nat \max (ln e / ln (real k * \ / 2)) (- ln e / ln (real k * \))\" for e + have NN: "of_nat (NN e) > ln e / ln (real k * \ / 2)" "of_nat (NN e) > - ln e / ln (real k * \)" + if "0e. e>0 \ (k * \ / 2)^NN e < e" + apply (subst Transcendental.ln_less_cancel_iff [symmetric]) + prefer 3 apply (subst ln_realpow) + using \k>0\ \\>0\ NN k\ + apply (force simp add: field_simps)+ + done + have NN0: "\e. e>0 \ (1/(k*\))^NN e < e" + apply (subst Transcendental.ln_less_cancel_iff [symmetric]) + prefer 3 apply (subst ln_realpow) + using \k>0\ \\>0\ NN k\ + apply (force simp add: field_simps ln_div)+ + done + { fix t and e::real + assume "e>0" + have "t \ s \ V \ 1 - q (NN e) t < e" "t \ s - U \ q (NN e) t < e" + proof - + assume t: "t \ s \ V" + show "1 - q (NN e) t < e" + by (metis add.commute diff_le_eq not_le limitV [OF t] less_le_trans [OF NN1 [OF \e>0\]]) + next + assume t: "t \ s - U" + show "q (NN e) t < e" + using limitNonU [OF t] less_le_trans [OF NN0 [OF \e>0\]] not_le by blast + qed + } then have "\e. e > 0 \ \f\R. f ` s \ {0..1} \ (\t \ s \ V. f t < e) \ (\t \ s - U. 1 - e < f t)" + using q01 + by (rule_tac x="\x. 1 - q (NN e) x" in bexI) (auto simp: algebra_simps intro: diff const qR) + moreover have t0V: "t0 \ V" "s \ V \ U" + using pt_\ t0 U V \01 by fastforce+ + ultimately show ?thesis using V t0V + by blast +qed + +text\Non-trivial case, with @{term A} and @{term B} both non-empty\ +lemma (in function_ring_on) two_special: + assumes A: "closed A" "A \ s" "a \ A" + and B: "closed B" "B \ s" "b \ B" + and disj: "A \ B = {}" + and e: "0 < e" "e < 1" + shows "\f \ R. f ` s \ {0..1} \ (\x \ A. f x < e) \ (\x \ B. f x > 1 - e)" +proof - + { fix w + assume "w \ A" + then have "open ( - B)" "b \ s" "w \ B" "w \ s" + using assms by auto + then have "\V. open V \ w \ V \ s \ V \ -B \ + (\e>0. \f \ R. f ` s \ {0..1} \ (\x \ s \ V. f x < e) \ (\x \ s \ B. f x > 1 - e))" + using one [of "-B" w b] assms \w \ A\ by simp + } + then obtain Vf where Vf: + "\w. w \ A \ open (Vf w) \ w \ Vf w \ s \ Vf w \ -B \ + (\e>0. \f \ R. f ` s \ {0..1} \ (\x \ s \ Vf w. f x < e) \ (\x \ s \ B. f x > 1 - e))" + by metis + then have open_Vf: "\w. w \ A \ open (Vf w)" + by blast + have tVft: "\w. w \ A \ w \ Vf w" + using Vf by blast + then have setsum_max_0: "A \ (\x \ A. Vf x)" + by blast + have com_A: "compact A" using A + by (metis compact compact_Int_closed inf.absorb_iff2) + obtain subA where subA: "subA \ A" "finite subA" "A \ (\x \ subA. Vf x)" + by (blast intro: that open_Vf compactE_image [OF com_A _ setsum_max_0]) + then have [simp]: "subA \ {}" + using \a \ A\ by auto + then have cardp: "card subA > 0" using subA + by (simp add: card_gt_0_iff) + have "\w. w \ A \ \f \ R. f ` s \ {0..1} \ (\x \ s \ Vf w. f x < e / card subA) \ (\x \ s \ B. f x > 1 - e / card subA)" + using Vf e cardp by simp + then obtain ff where ff: + "\w. w \ A \ ff w \ R \ ff w ` s \ {0..1} \ + (\x \ s \ Vf w. ff w x < e / card subA) \ (\x \ s \ B. ff w x > 1 - e / card subA)" + by metis + define pff where [abs_def]: "pff x = (\w \ subA. ff w x)" for x + have pffR: "pff \ R" + unfolding pff_def using subA ff by (auto simp: intro: setprod) + moreover + have pff01: "pff x \ {0..1}" if t: "x \ s" for x + proof - + have "0 \ pff x" + using subA cardp t + apply (simp add: pff_def divide_simps setsum_nonneg) + apply (rule Groups_Big.linordered_semidom_class.setprod_nonneg) + using ff by fastforce + moreover have "pff x \ 1" + using subA cardp t + apply (simp add: pff_def divide_simps setsum_nonneg) + apply (rule setprod_mono [where g = "\x. 1", simplified]) + using ff by fastforce + ultimately show ?thesis + by auto + qed + moreover + { fix v x + assume v: "v \ subA" and x: "x \ Vf v" "x \ s" + from subA v have "pff x = ff v x * (\w \ subA - {v}. ff w x)" + unfolding pff_def by (metis setprod.remove) + also have "... \ ff v x * 1" + apply (rule Rings.ordered_semiring_class.mult_left_mono) + apply (rule setprod_mono [where g = "\x. 1", simplified]) + using ff [THEN conjunct2, THEN conjunct1] v subA x + apply auto + apply (meson atLeastAtMost_iff contra_subsetD imageI) + apply (meson atLeastAtMost_iff contra_subsetD image_eqI) + using atLeastAtMost_iff by blast + also have "... < e / card subA" + using ff [THEN conjunct2, THEN conjunct2, THEN conjunct1] v subA x + by auto + also have "... \ e" + using cardp e by (simp add: divide_simps) + finally have "pff x < e" . + } + then have "\x. x \ A \ pff x < e" + using A Vf subA by (metis UN_E contra_subsetD) + moreover + { fix x + assume x: "x \ B" + then have "x \ s" + using B by auto + have "1 - e \ (1 - e / card subA) ^ card subA" + using Bernoulli_inequality [of "-e / card subA" "card subA"] e cardp + by (auto simp: field_simps) + also have "... = (\w \ subA. 1 - e / card subA)" + by (simp add: setprod_constant subA(2)) + also have "... < pff x" + apply (simp add: pff_def) + apply (rule setprod_mono_strict [where f = "\x. 1 - e / card subA", simplified]) + apply (simp_all add: subA(2)) + apply (intro ballI conjI) + using e apply (force simp: divide_simps) + using ff [THEN conjunct2, THEN conjunct2, THEN conjunct2] subA B x + apply blast + done + finally have "1 - e < pff x" . + } + ultimately + show ?thesis by blast +qed + +lemma (in function_ring_on) two: + assumes A: "closed A" "A \ s" + and B: "closed B" "B \ s" + and disj: "A \ B = {}" + and e: "0 < e" "e < 1" + shows "\f \ R. f ` s \ {0..1} \ (\x \ A. f x < e) \ (\x \ B. f x > 1 - e)" +proof (cases "A \ {} \ B \ {}") + case True then show ?thesis + apply (simp add: ex_in_conv [symmetric]) + using assms + apply safe + apply (force simp add: intro!: two_special) + done +next + case False with e show ?thesis + apply simp + apply (erule disjE) + apply (rule_tac [2] x="\x. 0" in bexI) + apply (rule_tac x="\x. 1" in bexI) + apply (auto simp: const) + done +qed + +text\The special case where @{term f} is non-negative and @{term"e<1/3"}\ +lemma (in function_ring_on) Stone_Weierstrass_special: + assumes f: "continuous_on s f" and fpos: "\x. x \ s \ f x \ 0" + and e: "0 < e" "e < 1/3" + shows "\g \ R. \x\s. \f x - g x\ < 2*e" +proof - + define n where "n = 1 + nat \normf f / e\" + define A where "A j = {x \ s. f x \ (j - 1/3)*e}" for j :: nat + define B where "B j = {x \ s. f x \ (j + 1/3)*e}" for j :: nat + have ngt: "(n-1) * e \ normf f" "n\1" + using e + apply (simp_all add: n_def field_simps of_nat_Suc) + by (metis real_nat_ceiling_ge mult.commute not_less pos_less_divide_eq) + then have ge_fx: "(n-1) * e \ f x" if "x \ s" for x + using f normf_upper that by fastforce + { fix j + have A: "closed (A j)" "A j \ s" + apply (simp_all add: A_def Collect_restrict) + apply (rule continuous_on_closed_Collect_le [OF f continuous_on_const]) + apply (simp add: compact compact_imp_closed) + done + have B: "closed (B j)" "B j \ s" + apply (simp_all add: B_def Collect_restrict) + apply (rule continuous_on_closed_Collect_le [OF continuous_on_const f]) + apply (simp add: compact compact_imp_closed) + done + have disj: "(A j) \ (B j) = {}" + using e by (auto simp: A_def B_def field_simps) + have "\f \ R. f ` s \ {0..1} \ (\x \ A j. f x < e/n) \ (\x \ B j. f x > 1 - e/n)" + apply (rule two) + using e A B disj ngt + apply simp_all + done + } + then obtain xf where xfR: "\j. xf j \ R" and xf01: "\j. xf j ` s \ {0..1}" + and xfA: "\x j. x \ A j \ xf j x < e/n" + and xfB: "\x j. x \ B j \ xf j x > 1 - e/n" + by metis + define g where [abs_def]: "g x = e * (\i\n. xf i x)" for x + have gR: "g \ R" + unfolding g_def by (fast intro: mult const setsum xfR) + have gge0: "\x. x \ s \ g x \ 0" + using e xf01 by (simp add: g_def zero_le_mult_iff image_subset_iff setsum_nonneg) + have A0: "A 0 = {}" + using fpos e by (fastforce simp: A_def) + have An: "A n = s" + using e ngt f normf_upper by (fastforce simp: A_def field_simps of_nat_diff) + have Asub: "A j \ A i" if "i\j" for i j + using e that apply (clarsimp simp: A_def) + apply (erule order_trans, simp) + done + { fix t + assume t: "t \ s" + define j where "j = (LEAST j. t \ A j)" + have jn: "j \ n" + using t An by (simp add: Least_le j_def) + have Aj: "t \ A j" + using t An by (fastforce simp add: j_def intro: LeastI) + then have Ai: "t \ A i" if "i\j" for i + using Asub [OF that] by blast + then have fj1: "f t \ (j - 1/3)*e" + by (simp add: A_def) + then have Anj: "t \ A i" if "ii + apply (simp add: j_def) + using not_less_Least by blast + have j1: "1 \ j" + using A0 Aj j_def not_less_eq_eq by (fastforce simp add: j_def) + then have Anj: "t \ A (j-1)" + using Least_le by (fastforce simp add: j_def) + then have fj2: "(j - 4/3)*e < f t" + using j1 t by (simp add: A_def of_nat_diff) + have ***: "xf i t \ e/n" if "i\j" for i + using xfA [OF Ai] that by (simp add: less_eq_real_def) + { fix i + assume "i+2 \ j" + then obtain d where "i+2+d = j" + using le_Suc_ex that by blast + then have "t \ B i" + using Anj e ge_fx [OF t] \1 \ n\ fpos [OF t] t + apply (simp add: A_def B_def) + apply (clarsimp simp add: field_simps of_nat_diff not_le of_nat_Suc) + apply (rule order_trans [of _ "e * 2 + (e * (real d * 3) + e * (real i * 3))"]) + apply auto + done + then have "xf i t > 1 - e/n" + by (rule xfB) + } note **** = this + have xf_le1: "\i. xf i t \ 1" + using xf01 t by force + have "g t = e * (\ii=j..n. xf i t)" + using j1 jn e + apply (simp add: g_def distrib_left [symmetric]) + apply (subst setsum.union_disjoint [symmetric]) + apply (auto simp: ivl_disj_un) + done + also have "... \ e*j + e * ((Suc n - j)*e/n)" + apply (rule add_mono) + apply (simp_all only: mult_le_cancel_left_pos e) + apply (rule setsum_bounded_above [OF xf_le1, where A = "lessThan j", simplified]) + using setsum_bounded_above [of "{j..n}" "\i. xf i t", OF ***] + apply simp + done + also have "... \ j*e + e*(n - j + 1)*e/n " + using \1 \ n\ e by (simp add: field_simps del: of_nat_Suc) + also have "... \ j*e + e*e" + using \1 \ n\ e j1 by (simp add: field_simps del: of_nat_Suc) + also have "... < (j + 1/3)*e" + using e by (auto simp: field_simps) + finally have gj1: "g t < (j + 1 / 3) * e" . + have gj2: "(j - 4/3)*e < g t" + proof (cases "2 \ j") + case False + then have "j=1" using j1 by simp + with t gge0 e show ?thesis by force + next + case True + then have "(j - 4/3)*e < (j-1)*e - e^2" + using e by (auto simp: of_nat_diff algebra_simps power2_eq_square) + also have "... < (j-1)*e - ((j - 1)/n) * e^2" + using e True jn by (simp add: power2_eq_square field_simps) + also have "... = e * (j-1) * (1 - e/n)" + by (simp add: power2_eq_square field_simps) + also have "... \ e * (\i\j-2. xf i t)" + using e + apply simp + apply (rule order_trans [OF _ setsum_bounded_below [OF less_imp_le [OF ****]]]) + using True + apply (simp_all add: of_nat_Suc of_nat_diff) + done + also have "... \ g t" + using jn e + using e xf01 t + apply (simp add: g_def zero_le_mult_iff image_subset_iff setsum_nonneg) + apply (rule Groups_Big.setsum_mono2, auto) + done + finally show ?thesis . + qed + have "\f t - g t\ < 2 * e" + using fj1 fj2 gj1 gj2 by (simp add: abs_less_iff field_simps) + } + then show ?thesis + by (rule_tac x=g in bexI) (auto intro: gR) +qed + +text\The ``unpretentious'' formulation\ +lemma (in function_ring_on) Stone_Weierstrass_basic: + assumes f: "continuous_on s f" and e: "e > 0" + shows "\g \ R. \x\s. \f x - g x\ < e" +proof - + have "\g \ R. \x\s. \(f x + normf f) - g x\ < 2 * min (e/2) (1/4)" + apply (rule Stone_Weierstrass_special) + apply (rule Limits.continuous_on_add [OF f Topological_Spaces.continuous_on_const]) + using normf_upper [OF f] apply force + apply (simp add: e, linarith) + done + then obtain g where "g \ R" "\x\s. \g x - (f x + normf f)\ < e" + by force + then show ?thesis + apply (rule_tac x="\x. g x - normf f" in bexI) + apply (auto simp: algebra_simps intro: diff const) + done +qed + + +theorem (in function_ring_on) Stone_Weierstrass: + assumes f: "continuous_on s f" + shows "\F\UNIV \ R. LIM n sequentially. F n :> uniformly_on s f" +proof - + { fix e::real + assume e: "0 < e" + then obtain N::nat where N: "0 < N" "0 < inverse N" "inverse N < e" + by (auto simp: real_arch_inverse [of e]) + { fix n :: nat and x :: 'a and g :: "'a \ real" + assume n: "N \ n" "\x\s. \f x - g x\ < 1 / (1 + real n)" + assume x: "x \ s" + have "\ real (Suc n) < inverse e" + using \N \ n\ N using less_imp_inverse_less by force + then have "1 / (1 + real n) \ e" + using e by (simp add: field_simps of_nat_Suc) + then have "\f x - g x\ < e" + using n(2) x by auto + } note * = this + have "\\<^sub>F n in sequentially. \x\s. \f x - (SOME g. g \ R \ (\x\s. \f x - g x\ < 1 / (1 + real n))) x\ < e" + apply (rule eventually_sequentiallyI [of N]) + apply (auto intro: someI2_bex [OF Stone_Weierstrass_basic [OF f]] *) + done + } then + show ?thesis + apply (rule_tac x="\n::nat. SOME g. g \ R \ (\x\s. \f x - g x\ < 1 / (1 + n))" in bexI) + prefer 2 apply (force intro: someI2_bex [OF Stone_Weierstrass_basic [OF f]]) + unfolding uniform_limit_iff + apply (auto simp: dist_norm abs_minus_commute) + done +qed + +text\A HOL Light formulation\ +corollary Stone_Weierstrass_HOL: + fixes R :: "('a::t2_space \ real) set" and s :: "'a set" + assumes "compact s" "\c. P(\x. c::real)" + "\f. P f \ continuous_on s f" + "\f g. P(f) \ P(g) \ P(\x. f x + g x)" "\f g. P(f) \ P(g) \ P(\x. f x * g x)" + "\x y. x \ s \ y \ s \ ~(x = y) \ \f. P(f) \ ~(f x = f y)" + "continuous_on s f" + "0 < e" + shows "\g. P(g) \ (\x \ s. \f x - g x\ < e)" +proof - + interpret PR: function_ring_on "Collect P" + apply unfold_locales + using assms + by auto + show ?thesis + using PR.Stone_Weierstrass_basic [OF \continuous_on s f\ \0 < e\] + by blast +qed + + +subsection \Polynomial functions\ + +inductive real_polynomial_function :: "('a::real_normed_vector \ real) \ bool" where + linear: "bounded_linear f \ real_polynomial_function f" + | const: "real_polynomial_function (\x. c)" + | add: "\real_polynomial_function f; real_polynomial_function g\ \ real_polynomial_function (\x. f x + g x)" + | mult: "\real_polynomial_function f; real_polynomial_function g\ \ real_polynomial_function (\x. f x * g x)" + +declare real_polynomial_function.intros [intro] + +definition polynomial_function :: "('a::real_normed_vector \ 'b::real_normed_vector) \ bool" + where + "polynomial_function p \ (\f. bounded_linear f \ real_polynomial_function (f o p))" + +lemma real_polynomial_function_eq: "real_polynomial_function p = polynomial_function p" +unfolding polynomial_function_def +proof + assume "real_polynomial_function p" + then show " \f. bounded_linear f \ real_polynomial_function (f \ p)" + proof (induction p rule: real_polynomial_function.induct) + case (linear h) then show ?case + by (auto simp: bounded_linear_compose real_polynomial_function.linear) + next + case (const h) then show ?case + by (simp add: real_polynomial_function.const) + next + case (add h) then show ?case + by (force simp add: bounded_linear_def linear_add real_polynomial_function.add) + next + case (mult h) then show ?case + by (force simp add: real_bounded_linear const real_polynomial_function.mult) + qed +next + assume [rule_format, OF bounded_linear_ident]: "\f. bounded_linear f \ real_polynomial_function (f \ p)" + then show "real_polynomial_function p" + by (simp add: o_def) +qed + +lemma polynomial_function_const [iff]: "polynomial_function (\x. c)" + by (simp add: polynomial_function_def o_def const) + +lemma polynomial_function_bounded_linear: + "bounded_linear f \ polynomial_function f" + by (simp add: polynomial_function_def o_def bounded_linear_compose real_polynomial_function.linear) + +lemma polynomial_function_id [iff]: "polynomial_function(\x. x)" + by (simp add: polynomial_function_bounded_linear) + +lemma polynomial_function_add [intro]: + "\polynomial_function f; polynomial_function g\ \ polynomial_function (\x. f x + g x)" + by (auto simp: polynomial_function_def bounded_linear_def linear_add real_polynomial_function.add o_def) + +lemma polynomial_function_mult [intro]: + assumes f: "polynomial_function f" and g: "polynomial_function g" + shows "polynomial_function (\x. f x *\<^sub>R g x)" + using g + apply (auto simp: polynomial_function_def bounded_linear_def Real_Vector_Spaces.linear.scaleR const real_polynomial_function.mult o_def) + apply (rule mult) + using f + apply (auto simp: real_polynomial_function_eq) + done + +lemma polynomial_function_cmul [intro]: + assumes f: "polynomial_function f" + shows "polynomial_function (\x. c *\<^sub>R f x)" + by (rule polynomial_function_mult [OF polynomial_function_const f]) + +lemma polynomial_function_minus [intro]: + assumes f: "polynomial_function f" + shows "polynomial_function (\x. - f x)" + using polynomial_function_cmul [OF f, of "-1"] by simp + +lemma polynomial_function_diff [intro]: + "\polynomial_function f; polynomial_function g\ \ polynomial_function (\x. f x - g x)" + unfolding add_uminus_conv_diff [symmetric] + by (metis polynomial_function_add polynomial_function_minus) + +lemma polynomial_function_setsum [intro]: + "\finite I; \i. i \ I \ polynomial_function (\x. f x i)\ \ polynomial_function (\x. setsum (f x) I)" +by (induct I rule: finite_induct) auto + +lemma real_polynomial_function_minus [intro]: + "real_polynomial_function f \ real_polynomial_function (\x. - f x)" + using polynomial_function_minus [of f] + by (simp add: real_polynomial_function_eq) + +lemma real_polynomial_function_diff [intro]: + "\real_polynomial_function f; real_polynomial_function g\ \ real_polynomial_function (\x. f x - g x)" + using polynomial_function_diff [of f] + by (simp add: real_polynomial_function_eq) + +lemma real_polynomial_function_setsum [intro]: + "\finite I; \i. i \ I \ real_polynomial_function (\x. f x i)\ \ real_polynomial_function (\x. setsum (f x) I)" + using polynomial_function_setsum [of I f] + by (simp add: real_polynomial_function_eq) + +lemma real_polynomial_function_power [intro]: + "real_polynomial_function f \ real_polynomial_function (\x. f x ^ n)" + by (induct n) (simp_all add: const mult) + +lemma real_polynomial_function_compose [intro]: + assumes f: "polynomial_function f" and g: "real_polynomial_function g" + shows "real_polynomial_function (g o f)" + using g + apply (induction g rule: real_polynomial_function.induct) + using f + apply (simp_all add: polynomial_function_def o_def const add mult) + done + +lemma polynomial_function_compose [intro]: + assumes f: "polynomial_function f" and g: "polynomial_function g" + shows "polynomial_function (g o f)" + using g real_polynomial_function_compose [OF f] + by (auto simp: polynomial_function_def o_def) + +lemma setsum_max_0: + fixes x::real (*in fact "'a::comm_ring_1"*) + shows "(\i = 0..max m n. x^i * (if i \ m then a i else 0)) = (\i = 0..m. x^i * a i)" +proof - + have "(\i = 0..max m n. x^i * (if i \ m then a i else 0)) = (\i = 0..max m n. (if i \ m then x^i * a i else 0))" + by (auto simp: algebra_simps intro: setsum.cong) + also have "... = (\i = 0..m. (if i \ m then x^i * a i else 0))" + by (rule setsum.mono_neutral_right) auto + also have "... = (\i = 0..m. x^i * a i)" + by (auto simp: algebra_simps intro: setsum.cong) + finally show ?thesis . +qed + +lemma real_polynomial_function_imp_setsum: + assumes "real_polynomial_function f" + shows "\a n::nat. f = (\x. \i=0..n. a i * x ^ i)" +using assms +proof (induct f) + case (linear f) + then show ?case + apply (clarsimp simp add: real_bounded_linear) + apply (rule_tac x="\i. if i=0 then 0 else c" in exI) + apply (rule_tac x=1 in exI) + apply (simp add: mult_ac) + done +next + case (const c) + show ?case + apply (rule_tac x="\i. c" in exI) + apply (rule_tac x=0 in exI) + apply (auto simp: mult_ac of_nat_Suc) + done + case (add f1 f2) + then obtain a1 n1 a2 n2 where + "f1 = (\x. \i = 0..n1. a1 i * x ^ i)" "f2 = (\x. \i = 0..n2. a2 i * x ^ i)" + by auto + then show ?case + apply (rule_tac x="\i. (if i \ n1 then a1 i else 0) + (if i \ n2 then a2 i else 0)" in exI) + apply (rule_tac x="max n1 n2" in exI) + using setsum_max_0 [where m=n1 and n=n2] setsum_max_0 [where m=n2 and n=n1] + apply (simp add: setsum.distrib algebra_simps max.commute) + done + case (mult f1 f2) + then obtain a1 n1 a2 n2 where + "f1 = (\x. \i = 0..n1. a1 i * x ^ i)" "f2 = (\x. \i = 0..n2. a2 i * x ^ i)" + by auto + then obtain b1 b2 where + "f1 = (\x. \i = 0..n1. b1 i * x ^ i)" "f2 = (\x. \i = 0..n2. b2 i * x ^ i)" + "b1 = (\i. if i\n1 then a1 i else 0)" "b2 = (\i. if i\n2 then a2 i else 0)" + by auto + then show ?case + apply (rule_tac x="\i. \k\i. b1 k * b2 (i - k)" in exI) + apply (rule_tac x="n1+n2" in exI) + using polynomial_product [of n1 b1 n2 b2] + apply (simp add: Set_Interval.atLeast0AtMost) + done +qed + +lemma real_polynomial_function_iff_setsum: + "real_polynomial_function f \ (\a n::nat. f = (\x. \i=0..n. a i * x ^ i))" + apply (rule iffI) + apply (erule real_polynomial_function_imp_setsum) + apply (auto simp: linear mult const real_polynomial_function_power real_polynomial_function_setsum) + done + +lemma polynomial_function_iff_Basis_inner: + fixes f :: "'a::real_normed_vector \ 'b::euclidean_space" + shows "polynomial_function f \ (\b\Basis. real_polynomial_function (\x. inner (f x) b))" + (is "?lhs = ?rhs") +unfolding polynomial_function_def +proof (intro iffI allI impI) + assume "\h. bounded_linear h \ real_polynomial_function (h \ f)" + then show ?rhs + by (force simp add: bounded_linear_inner_left o_def) +next + fix h :: "'b \ real" + assume rp: "\b\Basis. real_polynomial_function (\x. f x \ b)" and h: "bounded_linear h" + have "real_polynomial_function (h \ (\x. \b\Basis. (f x \ b) *\<^sub>R b))" + apply (rule real_polynomial_function_compose [OF _ linear [OF h]]) + using rp + apply (auto simp: real_polynomial_function_eq polynomial_function_mult) + done + then show "real_polynomial_function (h \ f)" + by (simp add: euclidean_representation_setsum_fun) +qed + +subsection \Stone-Weierstrass theorem for polynomial functions\ + +text\First, we need to show that they are continous, differentiable and separable.\ + +lemma continuous_real_polymonial_function: + assumes "real_polynomial_function f" + shows "continuous (at x) f" +using assms +by (induct f) (auto simp: linear_continuous_at) + +lemma continuous_polymonial_function: + fixes f :: "'a::real_normed_vector \ 'b::euclidean_space" + assumes "polynomial_function f" + shows "continuous (at x) f" + apply (rule euclidean_isCont) + using assms apply (simp add: polynomial_function_iff_Basis_inner) + apply (force dest: continuous_real_polymonial_function intro: isCont_scaleR) + done + +lemma continuous_on_polymonial_function: + fixes f :: "'a::real_normed_vector \ 'b::euclidean_space" + assumes "polynomial_function f" + shows "continuous_on s f" + using continuous_polymonial_function [OF assms] continuous_at_imp_continuous_on + by blast + +lemma has_real_derivative_polynomial_function: + assumes "real_polynomial_function p" + shows "\p'. real_polynomial_function p' \ + (\x. (p has_real_derivative (p' x)) (at x))" +using assms +proof (induct p) + case (linear p) + then show ?case + by (force simp: real_bounded_linear const intro!: derivative_eq_intros) +next + case (const c) + show ?case + by (rule_tac x="\x. 0" in exI) auto + case (add f1 f2) + then obtain p1 p2 where + "real_polynomial_function p1" "\x. (f1 has_real_derivative p1 x) (at x)" + "real_polynomial_function p2" "\x. (f2 has_real_derivative p2 x) (at x)" + by auto + then show ?case + apply (rule_tac x="\x. p1 x + p2 x" in exI) + apply (auto intro!: derivative_eq_intros) + done + case (mult f1 f2) + then obtain p1 p2 where + "real_polynomial_function p1" "\x. (f1 has_real_derivative p1 x) (at x)" + "real_polynomial_function p2" "\x. (f2 has_real_derivative p2 x) (at x)" + by auto + then show ?case + using mult + apply (rule_tac x="\x. f1 x * p2 x + f2 x * p1 x" in exI) + apply (auto intro!: derivative_eq_intros) + done +qed + +lemma has_vector_derivative_polynomial_function: + fixes p :: "real \ 'a::euclidean_space" + assumes "polynomial_function p" + shows "\p'. polynomial_function p' \ + (\x. (p has_vector_derivative (p' x)) (at x))" +proof - + { fix b :: 'a + assume "b \ Basis" + then + obtain p' where p': "real_polynomial_function p'" and pd: "\x. ((\x. p x \ b) has_real_derivative p' x) (at x)" + using assms [unfolded polynomial_function_iff_Basis_inner, rule_format] \b \ Basis\ + has_real_derivative_polynomial_function + by blast + have "\q. polynomial_function q \ (\x. ((\u. (p u \ b) *\<^sub>R b) has_vector_derivative q x) (at x))" + apply (rule_tac x="\x. p' x *\<^sub>R b" in exI) + using \b \ Basis\ p' + apply (simp add: polynomial_function_iff_Basis_inner inner_Basis) + apply (auto intro: derivative_eq_intros pd) + done + } + then obtain qf where qf: + "\b. b \ Basis \ polynomial_function (qf b)" + "\b x. b \ Basis \ ((\u. (p u \ b) *\<^sub>R b) has_vector_derivative qf b x) (at x)" + by metis + show ?thesis + apply (subst euclidean_representation_setsum_fun [of p, symmetric]) + apply (rule_tac x="\x. \b\Basis. qf b x" in exI) + apply (auto intro: has_vector_derivative_setsum qf) + done +qed + +lemma real_polynomial_function_separable: + fixes x :: "'a::euclidean_space" + assumes "x \ y" shows "\f. real_polynomial_function f \ f x \ f y" +proof - + have "real_polynomial_function (\u. \b\Basis. (inner (x-u) b)^2)" + apply (rule real_polynomial_function_setsum) + apply (auto simp: algebra_simps real_polynomial_function_power real_polynomial_function_diff + const linear bounded_linear_inner_left) + done + then show ?thesis + apply (intro exI conjI, assumption) + using assms + apply (force simp add: euclidean_eq_iff [of x y] setsum_nonneg_eq_0_iff algebra_simps) + done +qed + +lemma Stone_Weierstrass_real_polynomial_function: + fixes f :: "'a::euclidean_space \ real" + assumes "compact s" "continuous_on s f" "0 < e" + shows "\g. real_polynomial_function g \ (\x \ s. \f x - g x\ < e)" +proof - + interpret PR: function_ring_on "Collect real_polynomial_function" + apply unfold_locales + using assms continuous_on_polymonial_function real_polynomial_function_eq + apply (auto intro: real_polynomial_function_separable) + done + show ?thesis + using PR.Stone_Weierstrass_basic [OF \continuous_on s f\ \0 < e\] + by blast +qed + +lemma Stone_Weierstrass_polynomial_function: + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes s: "compact s" + and f: "continuous_on s f" + and e: "0 < e" + shows "\g. polynomial_function g \ (\x \ s. norm(f x - g x) < e)" +proof - + { fix b :: 'b + assume "b \ Basis" + have "\p. real_polynomial_function p \ (\x \ s. \f x \ b - p x\ < e / DIM('b))" + apply (rule exE [OF Stone_Weierstrass_real_polynomial_function [OF s _, of "\x. f x \ b" "e / card Basis"]]) + using e f + apply (auto simp: Euclidean_Space.DIM_positive intro: continuous_intros) + done + } + then obtain pf where pf: + "\b. b \ Basis \ real_polynomial_function (pf b) \ (\x \ s. \f x \ b - pf b x\ < e / DIM('b))" + apply (rule bchoice [rule_format, THEN exE]) + apply assumption + apply (force simp add: intro: that) + done + have "polynomial_function (\x. \b\Basis. pf b x *\<^sub>R b)" + using pf + by (simp add: polynomial_function_setsum polynomial_function_mult real_polynomial_function_eq) + moreover + { fix x + assume "x \ s" + have "norm (\b\Basis. (f x \ b) *\<^sub>R b - pf b x *\<^sub>R b) \ (\b\Basis. norm ((f x \ b) *\<^sub>R b - pf b x *\<^sub>R b))" + by (rule norm_setsum) + also have "... < of_nat DIM('b) * (e / DIM('b))" + apply (rule setsum_bounded_above_strict) + apply (simp add: Real_Vector_Spaces.scaleR_diff_left [symmetric] pf \x \ s\) + apply (rule DIM_positive) + done + also have "... = e" + using DIM_positive by (simp add: field_simps) + finally have "norm (\b\Basis. (f x \ b) *\<^sub>R b - pf b x *\<^sub>R b) < e" . + } + ultimately + show ?thesis + apply (subst euclidean_representation_setsum_fun [of f, symmetric]) + apply (rule_tac x="\x. \b\Basis. pf b x *\<^sub>R b" in exI) + apply (auto simp: setsum_subtractf [symmetric]) + done +qed + + +subsection\Polynomial functions as paths\ + +text\One application is to pick a smooth approximation to a path, +or just pick a smooth path anyway in an open connected set\ + +lemma path_polynomial_function: + fixes g :: "real \ 'b::euclidean_space" + shows "polynomial_function g \ path g" + by (simp add: path_def continuous_on_polymonial_function) + +lemma path_approx_polynomial_function: + fixes g :: "real \ 'b::euclidean_space" + assumes "path g" "0 < e" + shows "\p. polynomial_function p \ + pathstart p = pathstart g \ + pathfinish p = pathfinish g \ + (\t \ {0..1}. norm(p t - g t) < e)" +proof - + obtain q where poq: "polynomial_function q" and noq: "\x. x \ {0..1} \ norm (g x - q x) < e/4" + using Stone_Weierstrass_polynomial_function [of "{0..1}" g "e/4"] assms + by (auto simp: path_def) + have pf: "polynomial_function (\t. q t + (g 0 - q 0) + t *\<^sub>R (g 1 - q 1 - (g 0 - q 0)))" + by (force simp add: poq) + have *: "\t. t \ {0..1} \ norm (((q t - g t) + (g 0 - q 0)) + (t *\<^sub>R (g 1 - q 1) + t *\<^sub>R (q 0 - g 0))) < (e/4 + e/4) + (e/4+e/4)" + apply (intro Real_Vector_Spaces.norm_add_less) + using noq + apply (auto simp: norm_minus_commute intro: le_less_trans [OF mult_left_le_one_le noq] simp del: less_divide_eq_numeral1) + done + show ?thesis + apply (intro exI conjI) + apply (rule pf) + using * + apply (auto simp add: pathstart_def pathfinish_def algebra_simps) + done +qed + +lemma connected_open_polynomial_connected: + fixes s :: "'a::euclidean_space set" + assumes s: "open s" "connected s" + and "x \ s" "y \ s" + shows "\g. polynomial_function g \ path_image g \ s \ + pathstart g = x \ pathfinish g = y" +proof - + have "path_connected s" using assms + by (simp add: connected_open_path_connected) + with \x \ s\ \y \ s\ obtain p where p: "path p" "path_image p \ s" "pathstart p = x" "pathfinish p = y" + by (force simp: path_connected_def) + have "\e. 0 < e \ (\x \ path_image p. ball x e \ s)" + proof (cases "s = UNIV") + case True then show ?thesis + by (simp add: gt_ex) + next + case False + then have "- s \ {}" by blast + then show ?thesis + apply (rule_tac x="setdist (path_image p) (-s)" in exI) + using s p + apply (simp add: setdist_gt_0_compact_closed compact_path_image open_closed) + using setdist_le_dist [of _ "path_image p" _ "-s"] + by fastforce + qed + then obtain e where "0 < e"and eb: "\x. x \ path_image p \ ball x e \ s" + by auto + show ?thesis + using path_approx_polynomial_function [OF \path p\ \0 < e\] + apply clarify + apply (intro exI conjI, assumption) + using p + apply (fastforce simp add: dist_norm path_image_def norm_minus_commute intro: eb [THEN subsetD])+ + done +qed + +hide_fact linear add mult const + +end diff -r bbcb05504fdc -r bd218a9320b5 src/HOL/ROOT --- a/src/HOL/ROOT Thu Aug 04 18:45:28 2016 +0200 +++ b/src/HOL/ROOT Thu Aug 04 19:36:31 2016 +0200 @@ -719,12 +719,6 @@ session "HOL-Multivariate_Analysis" (main) in Multivariate_Analysis = HOL + theories Multivariate_Analysis - Determinants - PolyRoots - Polytope - Complex_Analysis_Basics - Complex_Transcendental - Cauchy_Integral_Thm document_files "root.tex"