diff --git a/src/HOL/Analysis/Convex_Euclidean_Space.thy b/src/HOL/Analysis/Convex_Euclidean_Space.thy --- a/src/HOL/Analysis/Convex_Euclidean_Space.thy +++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy @@ -1,2126 +1,2151 @@ (* Title: HOL/Analysis/Convex_Euclidean_Space.thy Author: L C Paulson, University of Cambridge Author: Robert Himmelmann, TU Muenchen Author: Bogdan Grechuk, University of Edinburgh Author: Armin Heller, TU Muenchen Author: Johannes Hoelzl, TU Muenchen *) section \Convex Sets and Functions on (Normed) Euclidean Spaces\ theory Convex_Euclidean_Space imports Convex Topology_Euclidean_Space begin subsection\<^marker>\tag unimportant\ \Topological Properties of Convex Sets and Functions\ lemma aff_dim_cball: fixes a :: "'n::euclidean_space" assumes "e > 0" shows "aff_dim (cball a e) = int (DIM('n))" proof - have "(\x. a + x) ` (cball 0 e) \ cball a e" unfolding cball_def dist_norm by auto then have "aff_dim (cball (0 :: 'n::euclidean_space) e) \ aff_dim (cball a e)" using aff_dim_translation_eq[of a "cball 0 e"] aff_dim_subset[of "(+) a ` cball 0 e" "cball a e"] by auto moreover have "aff_dim (cball (0 :: 'n::euclidean_space) e) = int (DIM('n))" using hull_inc[of "(0 :: 'n::euclidean_space)" "cball 0 e"] centre_in_cball[of "(0 :: 'n::euclidean_space)"] assms by (simp add: dim_cball[of e] aff_dim_zero[of "cball 0 e"]) ultimately show ?thesis using aff_dim_le_DIM[of "cball a e"] by auto qed lemma aff_dim_open: fixes S :: "'n::euclidean_space set" assumes "open S" and "S \ {}" shows "aff_dim S = int (DIM('n))" proof - obtain x where "x \ S" using assms by auto then obtain e where e: "e > 0" "cball x e \ S" using open_contains_cball[of S] assms by auto then have "aff_dim (cball x e) \ aff_dim S" using aff_dim_subset by auto with e show ?thesis using aff_dim_cball[of e x] aff_dim_le_DIM[of S] by auto qed lemma low_dim_interior: fixes S :: "'n::euclidean_space set" assumes "\ aff_dim S = int (DIM('n))" shows "interior S = {}" proof - have "aff_dim(interior S) \ aff_dim S" using interior_subset aff_dim_subset[of "interior S" S] by auto then show ?thesis using aff_dim_open[of "interior S"] aff_dim_le_DIM[of S] assms by auto qed corollary empty_interior_lowdim: fixes S :: "'n::euclidean_space set" shows "dim S < DIM ('n) \ interior S = {}" by (metis low_dim_interior affine_hull_UNIV dim_affine_hull less_not_refl dim_UNIV) corollary aff_dim_nonempty_interior: fixes S :: "'a::euclidean_space set" shows "interior S \ {} \ aff_dim S = DIM('a)" by (metis low_dim_interior) subsection \Relative interior of a set\ definition\<^marker>\tag important\ "rel_interior S = {x. \T. openin (top_of_set (affine hull S)) T \ x \ T \ T \ S}" lemma rel_interior_mono: "\S \ T; affine hull S = affine hull T\ \ (rel_interior S) \ (rel_interior T)" by (auto simp: rel_interior_def) lemma rel_interior_maximal: "\T \ S; openin(top_of_set (affine hull S)) T\ \ T \ (rel_interior S)" by (auto simp: rel_interior_def) lemma rel_interior: "rel_interior S = {x \ S. \T. open T \ x \ T \ T \ affine hull S \ S}" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (force simp add: rel_interior_def openin_open) { fix x T assume *: "x \ S" "open T" "x \ T" "T \ affine hull S \ S" then have **: "x \ T \ affine hull S" using hull_inc by auto with * have "\Tb. (\Ta. open Ta \ Tb = affine hull S \ Ta) \ x \ Tb \ Tb \ S" by (rule_tac x = "T \ (affine hull S)" in exI) auto } then show "?rhs \ ?lhs" by (force simp add: rel_interior_def openin_open) qed lemma mem_rel_interior: "x \ rel_interior S \ (\T. open T \ x \ T \ S \ T \ affine hull S \ S)" by (auto simp: rel_interior) lemma mem_rel_interior_ball: "x \ rel_interior S \ x \ S \ (\e. e > 0 \ ball x e \ affine hull S \ S)" (is "?lhs = ?rhs") proof assume ?rhs then show ?lhs by (simp add: rel_interior) (meson Elementary_Metric_Spaces.open_ball centre_in_ball) qed (force simp: rel_interior open_contains_ball) lemma rel_interior_ball: "rel_interior S = {x \ S. \e. e > 0 \ ball x e \ affine hull S \ S}" using mem_rel_interior_ball [of _ S] by auto lemma mem_rel_interior_cball: "x \ rel_interior S \ x \ S \ (\e. e > 0 \ cball x e \ affine hull S \ S)" (is "?lhs = ?rhs") proof assume ?rhs then obtain e where "x \ S" "e > 0" "cball x e \ affine hull S \ S" by (auto simp: rel_interior) then have "ball x e \ affine hull S \ S" by auto then show ?lhs using \0 < e\ \x \ S\ rel_interior_ball by auto qed (force simp: rel_interior open_contains_cball) lemma rel_interior_cball: "rel_interior S = {x \ S. \e. e > 0 \ cball x e \ affine hull S \ S}" using mem_rel_interior_cball [of _ S] by auto lemma rel_interior_empty [simp]: "rel_interior {} = {}" by (auto simp: rel_interior_def) lemma affine_hull_sing [simp]: "affine hull {a :: 'n::euclidean_space} = {a}" by (metis affine_hull_eq affine_sing) lemma rel_interior_sing [simp]: fixes a :: "'n::euclidean_space" shows "rel_interior {a} = {a}" proof - have "\x::real. 0 < x" using zero_less_one by blast then show ?thesis by (auto simp: rel_interior_ball) qed lemma subset_rel_interior: fixes S T :: "'n::euclidean_space set" assumes "S \ T" and "affine hull S = affine hull T" shows "rel_interior S \ rel_interior T" using assms by (auto simp: rel_interior_def) lemma rel_interior_subset: "rel_interior S \ S" by (auto simp: rel_interior_def) lemma rel_interior_subset_closure: "rel_interior S \ closure S" using rel_interior_subset by (auto simp: closure_def) lemma interior_subset_rel_interior: "interior S \ rel_interior S" by (auto simp: rel_interior interior_def) lemma interior_rel_interior: fixes S :: "'n::euclidean_space set" assumes "aff_dim S = int(DIM('n))" shows "rel_interior S = interior S" proof - have "affine hull S = UNIV" using assms affine_hull_UNIV[of S] by auto then show ?thesis unfolding rel_interior interior_def by auto qed lemma rel_interior_interior: fixes S :: "'n::euclidean_space set" assumes "affine hull S = UNIV" shows "rel_interior S = interior S" using assms unfolding rel_interior interior_def by auto lemma rel_interior_open: fixes S :: "'n::euclidean_space set" assumes "open S" shows "rel_interior S = S" by (metis assms interior_eq interior_subset_rel_interior rel_interior_subset set_eq_subset) lemma interior_rel_interior_gen: fixes S :: "'n::euclidean_space set" shows "interior S = (if aff_dim S = int(DIM('n)) then rel_interior S else {})" by (metis interior_rel_interior low_dim_interior) lemma rel_interior_nonempty_interior: fixes S :: "'n::euclidean_space set" shows "interior S \ {} \ rel_interior S = interior S" by (metis interior_rel_interior_gen) lemma affine_hull_nonempty_interior: fixes S :: "'n::euclidean_space set" shows "interior S \ {} \ affine hull S = UNIV" by (metis affine_hull_UNIV interior_rel_interior_gen) lemma rel_interior_affine_hull [simp]: fixes S :: "'n::euclidean_space set" shows "rel_interior (affine hull S) = affine hull S" proof - have *: "rel_interior (affine hull S) \ affine hull S" using rel_interior_subset by auto { fix x assume x: "x \ affine hull S" define e :: real where "e = 1" then have "e > 0" "ball x e \ affine hull (affine hull S) \ affine hull S" using hull_hull[of _ S] by auto then have "x \ rel_interior (affine hull S)" using x rel_interior_ball[of "affine hull S"] by auto } then show ?thesis using * by auto qed lemma rel_interior_UNIV [simp]: "rel_interior (UNIV :: ('n::euclidean_space) set) = UNIV" by (metis open_UNIV rel_interior_open) lemma rel_interior_convex_shrink: fixes S :: "'a::euclidean_space set" assumes "convex S" and "c \ rel_interior S" and "x \ S" and "0 < e" and "e \ 1" shows "x - e *\<^sub>R (x - c) \ rel_interior S" proof - obtain d where "d > 0" and d: "ball c d \ affine hull S \ S" using assms(2) unfolding mem_rel_interior_ball by auto { fix y assume as: "dist (x - e *\<^sub>R (x - c)) y < e * d" "y \ affine hull S" have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" using \e > 0\ by (auto simp: scaleR_left_diff_distrib scaleR_right_diff_distrib) have "x \ affine hull S" using assms hull_subset[of S] by auto moreover have "1 / e + - ((1 - e) / e) = 1" using \e > 0\ left_diff_distrib[of "1" "(1-e)" "1/e"] by auto ultimately have **: "(1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x \ affine hull S" using as affine_affine_hull[of S] mem_affine[of "affine hull S" y x "(1 / e)" "-((1 - e) / e)"] by (simp add: algebra_simps) have "c - ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = (1 / e) *\<^sub>R (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)" using \e > 0\ by (auto simp: euclidean_eq_iff[where 'a='a] field_simps inner_simps) then have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = \1/e\ * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)" unfolding dist_norm norm_scaleR[symmetric] by auto also have "\ = \1/e\ * norm (x - e *\<^sub>R (x - c) - y)" by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps) also have "\ < d" using as[unfolded dist_norm] and \e > 0\ by (auto simp:pos_divide_less_eq[OF \e > 0\] mult.commute) finally have "(1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x \ S" using "**" d by auto then have "y \ S" using * convexD [OF \convex S\] assms(3-5) by (metis diff_add_cancel diff_ge_0_iff_ge le_add_same_cancel1 less_eq_real_def) } then have "ball (x - e *\<^sub>R (x - c)) (e*d) \ affine hull S \ S" by auto moreover have "e * d > 0" using \e > 0\ \d > 0\ by simp moreover have c: "c \ S" using assms rel_interior_subset by auto moreover from c have "x - e *\<^sub>R (x - c) \ S" using convexD_alt[of S x c e] assms by (metis diff_add_eq diff_diff_eq2 less_eq_real_def scaleR_diff_left scaleR_one scale_right_diff_distrib) ultimately show ?thesis using mem_rel_interior_ball[of "x - e *\<^sub>R (x - c)" S] \e > 0\ by auto qed lemma interior_real_atLeast [simp]: fixes a :: real shows "interior {a..} = {a<..}" proof - { fix y have "ball y (y - a) \ {a..}" by (auto simp: dist_norm) moreover assume "a < y" ultimately have "y \ interior {a..}" by (force simp add: mem_interior) } moreover { fix y assume "y \ interior {a..}" then obtain e where e: "e > 0" "cball y e \ {a..}" using mem_interior_cball[of y "{a..}"] by auto moreover from e have "y - e \ cball y e" by (auto simp: cball_def dist_norm) ultimately have "a \ y - e" by blast then have "a < y" using e by auto } ultimately show ?thesis by auto qed lemma continuous_ge_on_Ioo: assumes "continuous_on {c..d} g" "\x. x \ {c<.. g x \ a" "c < d" "x \ {c..d}" shows "g (x::real) \ (a::real)" proof- from assms(3) have "{c..d} = closure {c<.. (g -` {a..} \ {c..d})" by auto hence "closure {c<.. closure (g -` {a..} \ {c..d})" by (rule closure_mono) also from assms(1) have "closed (g -` {a..} \ {c..d})" by (auto simp: continuous_on_closed_vimage) hence "closure (g -` {a..} \ {c..d}) = g -` {a..} \ {c..d}" by simp finally show ?thesis using \x \ {c..d}\ by auto qed lemma interior_real_atMost [simp]: fixes a :: real shows "interior {..a} = {.. {..a}" by (auto simp: dist_norm) moreover assume "a > y" ultimately have "y \ interior {..a}" by (force simp add: mem_interior) } moreover { fix y assume "y \ interior {..a}" then obtain e where e: "e > 0" "cball y e \ {..a}" using mem_interior_cball[of y "{..a}"] by auto moreover from e have "y + e \ cball y e" by (auto simp: cball_def dist_norm) ultimately have "a \ y + e" by auto then have "a > y" using e by auto } ultimately show ?thesis by auto qed lemma interior_atLeastAtMost_real [simp]: "interior {a..b} = {a<.. {..b}" by auto also have "interior \ = {a<..} \ {.. = {a<.. {}" using assms unfolding set_eq_iff by (auto intro!: exI[of _ "(a + b) / 2"] simp: box_def) then show ?thesis using interior_rel_interior_gen[of "cbox a b", symmetric] by (simp split: if_split_asm del: box_real add: box_real[symmetric]) qed lemma rel_interior_real_semiline [simp]: fixes a :: real shows "rel_interior {a..} = {a<..}" proof - have *: "{a<..} \ {}" unfolding set_eq_iff by (auto intro!: exI[of _ "a + 1"]) then show ?thesis using interior_real_atLeast interior_rel_interior_gen[of "{a..}"] by (auto split: if_split_asm) qed subsubsection \Relative open sets\ definition\<^marker>\tag important\ "rel_open S \ rel_interior S = S" lemma rel_open: "rel_open S \ openin (top_of_set (affine hull S)) S" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding rel_open_def rel_interior_def using openin_subopen[of "top_of_set (affine hull S)" S] by auto qed (auto simp: rel_open_def rel_interior_def) lemma openin_rel_interior: "openin (top_of_set (affine hull S)) (rel_interior S)" using openin_subopen by (fastforce simp add: rel_interior_def) lemma openin_set_rel_interior: "openin (top_of_set S) (rel_interior S)" by (rule openin_subset_trans [OF openin_rel_interior rel_interior_subset hull_subset]) lemma affine_rel_open: fixes S :: "'n::euclidean_space set" assumes "affine S" shows "rel_open S" unfolding rel_open_def using assms rel_interior_affine_hull[of S] affine_hull_eq[of S] by metis lemma affine_closed: fixes S :: "'n::euclidean_space set" assumes "affine S" shows "closed S" proof - { assume "S \ {}" then obtain L where L: "subspace L" "affine_parallel S L" using assms affine_parallel_subspace[of S] by auto then obtain a where a: "S = ((+) a ` L)" using affine_parallel_def[of L S] affine_parallel_commut by auto from L have "closed L" using closed_subspace by auto then have "closed S" using closed_translation a by auto } then show ?thesis by auto qed lemma closure_affine_hull: fixes S :: "'n::euclidean_space set" shows "closure S \ affine hull S" by (intro closure_minimal hull_subset affine_closed affine_affine_hull) lemma closed_affine_hull [iff]: fixes S :: "'n::euclidean_space set" shows "closed (affine hull S)" by (metis affine_affine_hull affine_closed) lemma closure_same_affine_hull [simp]: fixes S :: "'n::euclidean_space set" shows "affine hull (closure S) = affine hull S" proof - have "affine hull (closure S) \ affine hull S" using hull_mono[of "closure S" "affine hull S" "affine"] closure_affine_hull[of S] hull_hull[of "affine" S] by auto moreover have "affine hull (closure S) \ affine hull S" using hull_mono[of "S" "closure S" "affine"] closure_subset by auto ultimately show ?thesis by auto qed lemma closure_aff_dim [simp]: fixes S :: "'n::euclidean_space set" shows "aff_dim (closure S) = aff_dim S" proof - have "aff_dim S \ aff_dim (closure S)" using aff_dim_subset closure_subset by auto moreover have "aff_dim (closure S) \ aff_dim (affine hull S)" using aff_dim_subset closure_affine_hull by blast moreover have "aff_dim (affine hull S) = aff_dim S" using aff_dim_affine_hull by auto ultimately show ?thesis by auto qed lemma rel_interior_closure_convex_shrink: fixes S :: "_::euclidean_space set" assumes "convex S" and "c \ rel_interior S" and "x \ closure S" and "e > 0" and "e \ 1" shows "x - e *\<^sub>R (x - c) \ rel_interior S" proof - obtain d where "d > 0" and d: "ball c d \ affine hull S \ S" using assms(2) unfolding mem_rel_interior_ball by auto have "\y \ S. norm (y - x) * (1 - e) < e * d" proof (cases "x \ S") case True then show ?thesis using \e > 0\ \d > 0\ by force next case False then have x: "x islimpt S" using assms(3)[unfolded closure_def] by auto show ?thesis proof (cases "e = 1") case True obtain y where "y \ S" "y \ x" "dist y x < 1" using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto then show ?thesis unfolding True using \d > 0\ by (force simp add: ) next case False then have "0 < e * d / (1 - e)" and *: "1 - e > 0" using \e \ 1\ \e > 0\ \d > 0\ by auto then obtain y where "y \ S" "y \ x" "dist y x < e * d / (1 - e)" using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto then show ?thesis unfolding dist_norm using pos_less_divide_eq[OF *] by force qed qed then obtain y where "y \ S" and y: "norm (y - x) * (1 - e) < e * d" by auto define z where "z = c + ((1 - e) / e) *\<^sub>R (x - y)" have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)" unfolding z_def using \e > 0\ by (auto simp: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib) have zball: "z \ ball c d" using mem_ball z_def dist_norm[of c] using y and assms(4,5) by (simp add: norm_minus_commute) (simp add: field_simps) have "x \ affine hull S" using closure_affine_hull assms by auto moreover have "y \ affine hull S" using \y \ S\ hull_subset[of S] by auto moreover have "c \ affine hull S" using assms rel_interior_subset hull_subset[of S] by auto ultimately have "z \ affine hull S" using z_def affine_affine_hull[of S] mem_affine_3_minus [of "affine hull S" c x y "(1 - e) / e"] assms by simp then have "z \ S" using d zball by auto obtain d1 where "d1 > 0" and d1: "ball z d1 \ ball c d" using zball open_ball[of c d] openE[of "ball c d" z] by auto then have "ball z d1 \ affine hull S \ ball c d \ affine hull S" by auto then have "ball z d1 \ affine hull S \ S" using d by auto then have "z \ rel_interior S" using mem_rel_interior_ball using \d1 > 0\ \z \ S\ by auto then have "y - e *\<^sub>R (y - z) \ rel_interior S" using rel_interior_convex_shrink[of S z y e] assms \y \ S\ by auto then show ?thesis using * by auto qed lemma rel_interior_eq: "rel_interior s = s \ openin(top_of_set (affine hull s)) s" using rel_open rel_open_def by blast lemma rel_interior_openin: "openin(top_of_set (affine hull s)) s \ rel_interior s = s" by (simp add: rel_interior_eq) lemma rel_interior_affine: fixes S :: "'n::euclidean_space set" shows "affine S \ rel_interior S = S" using affine_rel_open rel_open_def by auto lemma rel_interior_eq_closure: fixes S :: "'n::euclidean_space set" shows "rel_interior S = closure S \ affine S" proof (cases "S = {}") case True then show ?thesis by auto next case False show ?thesis proof assume eq: "rel_interior S = closure S" have "openin (top_of_set (affine hull S)) S" by (metis eq closure_subset openin_rel_interior rel_interior_subset subset_antisym) moreover have "closedin (top_of_set (affine hull S)) S" by (metis closed_subset closure_subset_eq eq hull_subset rel_interior_subset) ultimately have "S = {} \ S = affine hull S" using convex_connected connected_clopen convex_affine_hull by metis with False have "affine hull S = S" by auto then show "affine S" by (metis affine_hull_eq) next assume "affine S" then show "rel_interior S = closure S" by (simp add: rel_interior_affine affine_closed) qed qed subsubsection\<^marker>\tag unimportant\\Relative interior preserves under linear transformations\ lemma rel_interior_translation_aux: fixes a :: "'n::euclidean_space" shows "((\x. a + x) ` rel_interior S) \ rel_interior ((\x. a + x) ` S)" proof - { fix x assume x: "x \ rel_interior S" then obtain T where "open T" "x \ T \ S" "T \ affine hull S \ S" using mem_rel_interior[of x S] by auto then have "open ((\x. a + x) ` T)" and "a + x \ ((\x. a + x) ` T) \ ((\x. a + x) ` S)" and "((\x. a + x) ` T) \ affine hull ((\x. a + x) ` S) \ (\x. a + x) ` S" using affine_hull_translation[of a S] open_translation[of T a] x by auto then have "a + x \ rel_interior ((\x. a + x) ` S)" using mem_rel_interior[of "a+x" "((\x. a + x) ` S)"] by auto } then show ?thesis by auto qed lemma rel_interior_translation: fixes a :: "'n::euclidean_space" shows "rel_interior ((\x. a + x) ` S) = (\x. a + x) ` rel_interior S" proof - have "(\x. (-a) + x) ` rel_interior ((\x. a + x) ` S) \ rel_interior S" using rel_interior_translation_aux[of "-a" "(\x. a + x) ` S"] translation_assoc[of "-a" "a"] by auto then have "((\x. a + x) ` rel_interior S) \ rel_interior ((\x. a + x) ` S)" using translation_inverse_subset[of a "rel_interior ((+) a ` S)" "rel_interior S"] by auto then show ?thesis using rel_interior_translation_aux[of a S] by auto qed lemma affine_hull_linear_image: assumes "bounded_linear f" shows "f ` (affine hull s) = affine hull f ` s" proof - interpret f: bounded_linear f by fact have "affine {x. f x \ affine hull f ` s}" unfolding affine_def by (auto simp: f.scaleR f.add affine_affine_hull[unfolded affine_def, rule_format]) moreover have "affine {x. x \ f ` (affine hull s)}" using affine_affine_hull[unfolded affine_def, of s] unfolding affine_def by (auto simp: f.scaleR [symmetric] f.add [symmetric]) ultimately show ?thesis by (auto simp: hull_inc elim!: hull_induct) qed lemma rel_interior_injective_on_span_linear_image: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" and S :: "'m::euclidean_space set" assumes "bounded_linear f" and "inj_on f (span S)" shows "rel_interior (f ` S) = f ` (rel_interior S)" proof - { fix z assume z: "z \ rel_interior (f ` S)" then have "z \ f ` S" using rel_interior_subset[of "f ` S"] by auto then obtain x where x: "x \ S" "f x = z" by auto obtain e2 where e2: "e2 > 0" "cball z e2 \ affine hull (f ` S) \ (f ` S)" using z rel_interior_cball[of "f ` S"] by auto obtain K where K: "K > 0" "\x. norm (f x) \ norm x * K" using assms Real_Vector_Spaces.bounded_linear.pos_bounded[of f] by auto define e1 where "e1 = 1 / K" then have e1: "e1 > 0" "\x. e1 * norm (f x) \ norm x" using K pos_le_divide_eq[of e1] by auto define e where "e = e1 * e2" then have "e > 0" using e1 e2 by auto { fix y assume y: "y \ cball x e \ affine hull S" then have h1: "f y \ affine hull (f ` S)" using affine_hull_linear_image[of f S] assms by auto from y have "norm (x-y) \ e1 * e2" using cball_def[of x e] dist_norm[of x y] e_def by auto moreover have "f x - f y = f (x - y)" using assms linear_diff[of f x y] linear_conv_bounded_linear[of f] by auto moreover have "e1 * norm (f (x-y)) \ norm (x - y)" using e1 by auto ultimately have "e1 * norm ((f x)-(f y)) \ e1 * e2" by auto then have "f y \ cball z e2" using cball_def[of "f x" e2] dist_norm[of "f x" "f y"] e1 x by auto then have "f y \ f ` S" using y e2 h1 by auto then have "y \ S" using assms y hull_subset[of S] affine_hull_subset_span inj_on_image_mem_iff [OF \inj_on f (span S)\] by (metis Int_iff span_superset subsetCE) } then have "z \ f ` (rel_interior S)" using mem_rel_interior_cball[of x S] \e > 0\ x by auto } moreover { fix x assume x: "x \ rel_interior S" then obtain e2 where e2: "e2 > 0" "cball x e2 \ affine hull S \ S" using rel_interior_cball[of S] by auto have "x \ S" using x rel_interior_subset by auto then have *: "f x \ f ` S" by auto have "\x\span S. f x = 0 \ x = 0" using assms subspace_span linear_conv_bounded_linear[of f] linear_injective_on_subspace_0[of f "span S"] by auto then obtain e1 where e1: "e1 > 0" "\x \ span S. e1 * norm x \ norm (f x)" using assms injective_imp_isometric[of "span S" f] subspace_span[of S] closed_subspace[of "span S"] by auto define e where "e = e1 * e2" hence "e > 0" using e1 e2 by auto { fix y assume y: "y \ cball (f x) e \ affine hull (f ` S)" then have "y \ f ` (affine hull S)" using affine_hull_linear_image[of f S] assms by auto then obtain xy where xy: "xy \ affine hull S" "f xy = y" by auto with y have "norm (f x - f xy) \ e1 * e2" using cball_def[of "f x" e] dist_norm[of "f x" y] e_def by auto moreover have "f x - f xy = f (x - xy)" using assms linear_diff[of f x xy] linear_conv_bounded_linear[of f] by auto moreover have *: "x - xy \ span S" using subspace_diff[of "span S" x xy] subspace_span \x \ S\ xy affine_hull_subset_span[of S] span_superset by auto moreover from * have "e1 * norm (x - xy) \ norm (f (x - xy))" using e1 by auto ultimately have "e1 * norm (x - xy) \ e1 * e2" by auto then have "xy \ cball x e2" using cball_def[of x e2] dist_norm[of x xy] e1 by auto then have "y \ f ` S" using xy e2 by auto } then have "f x \ rel_interior (f ` S)" using mem_rel_interior_cball[of "(f x)" "(f ` S)"] * \e > 0\ by auto } ultimately show ?thesis by auto qed lemma rel_interior_injective_linear_image: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "bounded_linear f" and "inj f" shows "rel_interior (f ` S) = f ` (rel_interior S)" using assms rel_interior_injective_on_span_linear_image[of f S] subset_inj_on[of f "UNIV" "span S"] by auto subsection\<^marker>\tag unimportant\ \Openness and compactness are preserved by convex hull operation\ lemma open_convex_hull[intro]: fixes S :: "'a::real_normed_vector set" assumes "open S" shows "open (convex hull S)" proof (clarsimp simp: open_contains_cball convex_hull_explicit) fix T and u :: "'a\real" assume obt: "finite T" "T\S" "\x\T. 0 \ u x" "sum u T = 1" from assms[unfolded open_contains_cball] obtain b where b: "\x. x\S \ 0 < b x \ cball x (b x) \ S" by metis have "b ` T \ {}" using obt by auto define i where "i = b ` T" let ?\ = "\y. \F. finite F \ F \ S \ (\u. (\x\F. 0 \ u x) \ sum u F = 1 \ (\v\F. u v *\<^sub>R v) = y)" let ?a = "\v\T. u v *\<^sub>R v" show "\e > 0. cball ?a e \ {y. ?\ y}" proof (intro exI subsetI conjI) show "0 < Min i" unfolding i_def and Min_gr_iff[OF finite_imageI[OF obt(1)] \b ` T\{}\] using b \T\S\ by auto next fix y assume "y \ cball ?a (Min i)" then have y: "norm (?a - y) \ Min i" unfolding dist_norm[symmetric] by auto { fix x assume "x \ T" then have "Min i \ b x" by (simp add: i_def obt(1)) then have "x + (y - ?a) \ cball x (b x)" using y unfolding mem_cball dist_norm by auto moreover have "x \ S" using \x\T\ \T\S\ by auto ultimately have "x + (y - ?a) \ S" using y b by blast } moreover have *: "inj_on (\v. v + (y - ?a)) T" unfolding inj_on_def by auto have "(\v\(\v. v + (y - ?a)) ` T. u (v - (y - ?a)) *\<^sub>R v) = y" unfolding sum.reindex[OF *] o_def using obt(4) by (simp add: sum.distrib sum_subtractf scaleR_left.sum[symmetric] scaleR_right_distrib) ultimately show "y \ {y. ?\ y}" proof (intro CollectI exI conjI) show "finite ((\v. v + (y - ?a)) ` T)" by (simp add: obt(1)) show "sum (\v. u (v - (y - ?a))) ((\v. v + (y - ?a)) ` T) = 1" unfolding sum.reindex[OF *] o_def using obt(4) by auto qed (use obt(1, 3) in auto) qed qed lemma compact_convex_combinations: fixes S T :: "'a::real_normed_vector set" assumes "compact S" "compact T" shows "compact { (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \ u \ u \ 1 \ x \ S \ y \ T}" proof - let ?X = "{0..1} \ S \ T" let ?h = "(\z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))" have *: "{ (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \ u \ u \ 1 \ x \ S \ y \ T} = ?h ` ?X" by force have "continuous_on ?X (\z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))" unfolding continuous_on by (rule ballI) (intro tendsto_intros) with assms show ?thesis by (simp add: * compact_Times compact_continuous_image) qed lemma finite_imp_compact_convex_hull: fixes S :: "'a::real_normed_vector set" assumes "finite S" shows "compact (convex hull S)" proof (cases "S = {}") case True then show ?thesis by simp next case False with assms show ?thesis proof (induct rule: finite_ne_induct) case (singleton x) show ?case by simp next case (insert x A) let ?f = "\(u, y::'a). u *\<^sub>R x + (1 - u) *\<^sub>R y" let ?T = "{0..1::real} \ (convex hull A)" have "continuous_on ?T ?f" unfolding split_def continuous_on by (intro ballI tendsto_intros) moreover have "compact ?T" by (intro compact_Times compact_Icc insert) ultimately have "compact (?f ` ?T)" by (rule compact_continuous_image) also have "?f ` ?T = convex hull (insert x A)" unfolding convex_hull_insert [OF \A \ {}\] apply safe apply (rule_tac x=a in exI, simp) apply (rule_tac x="1 - a" in exI, simp, fast) apply (rule_tac x="(u, b)" in image_eqI, simp_all) done finally show "compact (convex hull (insert x A))" . qed qed lemma compact_convex_hull: fixes S :: "'a::euclidean_space set" assumes "compact S" shows "compact (convex hull S)" proof (cases "S = {}") case True then show ?thesis using compact_empty by simp next case False then obtain w where "w \ S" by auto show ?thesis unfolding caratheodory[of S] proof (induct ("DIM('a) + 1")) case 0 have *: "{x.\sa. finite sa \ sa \ S \ card sa \ 0 \ x \ convex hull sa} = {}" using compact_empty by auto from 0 show ?case unfolding * by simp next case (Suc n) show ?case proof (cases "n = 0") case True have "{x. \T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T} = S" unfolding set_eq_iff and mem_Collect_eq proof (rule, rule) fix x assume "\T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T" then obtain T where T: "finite T" "T \ S" "card T \ Suc n" "x \ convex hull T" by auto show "x \ S" proof (cases "card T = 0") case True then show ?thesis using T(4) unfolding card_0_eq[OF T(1)] by simp next case False then have "card T = Suc 0" using T(3) \n=0\ by auto then obtain a where "T = {a}" unfolding card_Suc_eq by auto then show ?thesis using T(2,4) by simp qed next fix x assume "x\S" then show "\T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T" by (rule_tac x="{x}" in exI) (use convex_hull_singleton in auto) qed then show ?thesis using assms by simp next case False have "{x. \T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T} = {(1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \ u \ u \ 1 \ x \ S \ y \ {x. \T. finite T \ T \ S \ card T \ n \ x \ convex hull T}}" unfolding set_eq_iff and mem_Collect_eq proof (rule, rule) fix x assume "\u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \ 0 \ c \ c \ 1 \ u \ S \ (\T. finite T \ T \ S \ card T \ n \ v \ convex hull T)" then obtain u v c T where obt: "x = (1 - c) *\<^sub>R u + c *\<^sub>R v" "0 \ c \ c \ 1" "u \ S" "finite T" "T \ S" "card T \ n" "v \ convex hull T" by auto moreover have "(1 - c) *\<^sub>R u + c *\<^sub>R v \ convex hull insert u T" by (meson convexD_alt convex_convex_hull hull_inc hull_mono in_mono insertCI obt(2) obt(7) subset_insertI) ultimately show "\T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T" by (rule_tac x="insert u T" in exI) (auto simp: card_insert_if) next fix x assume "\T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T" then obtain T where T: "finite T" "T \ S" "card T \ Suc n" "x \ convex hull T" by auto show "\u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \ 0 \ c \ c \ 1 \ u \ S \ (\T. finite T \ T \ S \ card T \ n \ v \ convex hull T)" proof (cases "card T = Suc n") case False then have "card T \ n" using T(3) by auto then show ?thesis using \w\S\ and T by (rule_tac x=w in exI, rule_tac x=x in exI, rule_tac x=1 in exI) auto next case True then obtain a u where au: "T = insert a u" "a\u" by (metis card_le_Suc_iff order_refl) show ?thesis proof (cases "u = {}") case True then have "x = a" using T(4)[unfolded au] by auto show ?thesis unfolding \x = a\ using T \n \ 0\ unfolding au by (rule_tac x=a in exI, rule_tac x=a in exI, rule_tac x=1 in exI) force next case False obtain ux vx b where obt: "ux\0" "vx\0" "ux + vx = 1" "b \ convex hull u" "x = ux *\<^sub>R a + vx *\<^sub>R b" using T(4)[unfolded au convex_hull_insert[OF False]] by auto have *: "1 - vx = ux" using obt(3) by auto show ?thesis using obt T(1-3) card_insert_disjoint[OF _ au(2)] unfolding au * by (rule_tac x=a in exI, rule_tac x=b in exI, rule_tac x=vx in exI) force qed qed qed then show ?thesis using compact_convex_combinations[OF assms Suc] by simp qed qed qed subsection\<^marker>\tag unimportant\ \Extremal points of a simplex are some vertices\ lemma dist_increases_online: fixes a b d :: "'a::real_inner" assumes "d \ 0" shows "dist a (b + d) > dist a b \ dist a (b - d) > dist a b" proof (cases "inner a d - inner b d > 0") case True then have "0 < inner d d + (inner a d * 2 - inner b d * 2)" using assms by (intro add_pos_pos) auto then show ?thesis unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff by (simp add: algebra_simps inner_commute) next case False then have "0 < inner d d + (inner b d * 2 - inner a d * 2)" using assms by (intro add_pos_nonneg) auto then show ?thesis unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff by (simp add: algebra_simps inner_commute) qed lemma norm_increases_online: fixes d :: "'a::real_inner" shows "d \ 0 \ norm (a + d) > norm a \ norm(a - d) > norm a" using dist_increases_online[of d a 0] unfolding dist_norm by auto lemma simplex_furthest_lt: fixes S :: "'a::real_inner set" assumes "finite S" shows "\x \ convex hull S. x \ S \ (\y \ convex hull S. norm (x - a) < norm(y - a))" using assms proof induct fix x S assume as: "finite S" "x\S" "\x\convex hull S. x \ S \ (\y\convex hull S. norm (x - a) < norm (y - a))" show "\xa\convex hull insert x S. xa \ insert x S \ (\y\convex hull insert x S. norm (xa - a) < norm (y - a))" proof (intro impI ballI, cases "S = {}") case False fix y assume y: "y \ convex hull insert x S" "y \ insert x S" obtain u v b where obt: "u\0" "v\0" "u + v = 1" "b \ convex hull S" "y = u *\<^sub>R x + v *\<^sub>R b" using y(1)[unfolded convex_hull_insert[OF False]] by auto show "\z\convex hull insert x S. norm (y - a) < norm (z - a)" proof (cases "y \ convex hull S") case True then obtain z where "z \ convex hull S" "norm (y - a) < norm (z - a)" using as(3)[THEN bspec[where x=y]] and y(2) by auto then show ?thesis by (meson hull_mono subsetD subset_insertI) next case False show ?thesis proof (cases "u = 0 \ v = 0") case True with False show ?thesis using obt y by auto next case False then obtain w where w: "w>0" "w b" proof assume "x = b" then have "y = b" unfolding obt(5) using obt(3) by (auto simp: scaleR_left_distrib[symmetric]) then show False using obt(4) and False using \x = b\ y(2) by blast qed then have *: "w *\<^sub>R (x - b) \ 0" using w(1) by auto show ?thesis using dist_increases_online[OF *, of a y] proof (elim disjE) assume "dist a y < dist a (y + w *\<^sub>R (x - b))" then have "norm (y - a) < norm ((u + w) *\<^sub>R x + (v - w) *\<^sub>R b - a)" unfolding dist_commute[of a] unfolding dist_norm obt(5) by (simp add: algebra_simps) moreover have "(u + w) *\<^sub>R x + (v - w) *\<^sub>R b \ convex hull insert x S" unfolding convex_hull_insert[OF \S\{}\] proof (intro CollectI conjI exI) show "u + w \ 0" "v - w \ 0" using obt(1) w by auto qed (use obt in auto) ultimately show ?thesis by auto next assume "dist a y < dist a (y - w *\<^sub>R (x - b))" then have "norm (y - a) < norm ((u - w) *\<^sub>R x + (v + w) *\<^sub>R b - a)" unfolding dist_commute[of a] unfolding dist_norm obt(5) by (simp add: algebra_simps) moreover have "(u - w) *\<^sub>R x + (v + w) *\<^sub>R b \ convex hull insert x S" unfolding convex_hull_insert[OF \S\{}\] proof (intro CollectI conjI exI) show "u - w \ 0" "v + w \ 0" using obt(1) w by auto qed (use obt in auto) ultimately show ?thesis by auto qed qed qed qed auto qed (auto simp: assms) lemma simplex_furthest_le: fixes S :: "'a::real_inner set" assumes "finite S" and "S \ {}" shows "\y\S. \x\ convex hull S. norm (x - a) \ norm (y - a)" proof - have "convex hull S \ {}" using hull_subset[of S convex] and assms(2) by auto then obtain x where x: "x \ convex hull S" "\y\convex hull S. norm (y - a) \ norm (x - a)" using distance_attains_sup[OF finite_imp_compact_convex_hull[OF \finite S\], of a] unfolding dist_commute[of a] unfolding dist_norm by auto show ?thesis proof (cases "x \ S") case False then obtain y where "y \ convex hull S" "norm (x - a) < norm (y - a)" using simplex_furthest_lt[OF assms(1), THEN bspec[where x=x]] and x(1) by auto then show ?thesis using x(2)[THEN bspec[where x=y]] by auto next case True with x show ?thesis by auto qed qed lemma simplex_furthest_le_exists: fixes S :: "('a::real_inner) set" shows "finite S \ \x\(convex hull S). \y\S. norm (x - a) \ norm (y - a)" using simplex_furthest_le[of S] by (cases "S = {}") auto lemma simplex_extremal_le: fixes S :: "'a::real_inner set" assumes "finite S" and "S \ {}" shows "\u\S. \v\S. \x\convex hull S. \y \ convex hull S. norm (x - y) \ norm (u - v)" proof - have "convex hull S \ {}" using hull_subset[of S convex] and assms(2) by auto then obtain u v where obt: "u \ convex hull S" "v \ convex hull S" "\x\convex hull S. \y\convex hull S. norm (x - y) \ norm (u - v)" using compact_sup_maxdistance[OF finite_imp_compact_convex_hull[OF assms(1)]] by (auto simp: dist_norm) then show ?thesis proof (cases "u\S \ v\S", elim disjE) assume "u \ S" then obtain y where "y \ convex hull S" "norm (u - v) < norm (y - v)" using simplex_furthest_lt[OF assms(1), THEN bspec[where x=u]] and obt(1) by auto then show ?thesis using obt(3)[THEN bspec[where x=y], THEN bspec[where x=v]] and obt(2) by auto next assume "v \ S" then obtain y where "y \ convex hull S" "norm (v - u) < norm (y - u)" using simplex_furthest_lt[OF assms(1), THEN bspec[where x=v]] and obt(2) by auto then show ?thesis using obt(3)[THEN bspec[where x=u], THEN bspec[where x=y]] and obt(1) by (auto simp: norm_minus_commute) qed auto qed lemma simplex_extremal_le_exists: fixes S :: "'a::real_inner set" shows "finite S \ x \ convex hull S \ y \ convex hull S \ \u\S. \v\S. norm (x - y) \ norm (u - v)" using convex_hull_empty simplex_extremal_le[of S] by(cases "S = {}") auto subsection \Closest point of a convex set is unique, with a continuous projection\ definition\<^marker>\tag important\ closest_point :: "'a::{real_inner,heine_borel} set \ 'a \ 'a" where "closest_point S a = (SOME x. x \ S \ (\y\S. dist a x \ dist a y))" lemma closest_point_exists: assumes "closed S" and "S \ {}" shows closest_point_in_set: "closest_point S a \ S" and "\y\S. dist a (closest_point S a) \ dist a y" unfolding closest_point_def by (rule_tac someI2_ex, auto intro: distance_attains_inf[OF assms(1,2), of a])+ lemma closest_point_le: "closed S \ x \ S \ dist a (closest_point S a) \ dist a x" using closest_point_exists[of S] by auto lemma closest_point_self: assumes "x \ S" shows "closest_point S x = x" unfolding closest_point_def by (rule some1_equality, rule ex1I[of _ x]) (use assms in auto) lemma closest_point_refl: "closed S \ S \ {} \ closest_point S x = x \ x \ S" using closest_point_in_set[of S x] closest_point_self[of x S] by auto lemma closer_points_lemma: assumes "inner y z > 0" shows "\u>0. \v>0. v \ u \ norm(v *\<^sub>R z - y) < norm y" proof - have z: "inner z z > 0" unfolding inner_gt_zero_iff using assms by auto have "norm (v *\<^sub>R z - y) < norm y" if "0 < v" and "v \ inner y z / inner z z" for v unfolding norm_lt using z assms that by (simp add: field_simps inner_diff inner_commute mult_strict_left_mono[OF _ \0]) then show ?thesis using assms z by (rule_tac x = "inner y z / inner z z" in exI) auto qed lemma closer_point_lemma: assumes "inner (y - x) (z - x) > 0" shows "\u>0. u \ 1 \ dist (x + u *\<^sub>R (z - x)) y < dist x y" proof - obtain u where "u > 0" and u: "\v. \0 u\ \ norm (v *\<^sub>R (z - x) - (y - x)) < norm (y - x)" using closer_points_lemma[OF assms] by auto show ?thesis using u[of "min u 1"] and \u > 0\ by (metis diff_diff_add dist_commute dist_norm less_eq_real_def not_less u zero_less_one) qed lemma any_closest_point_dot: assumes "convex S" "closed S" "x \ S" "y \ S" "\z\S. dist a x \ dist a z" shows "inner (a - x) (y - x) \ 0" proof (rule ccontr) assume "\ ?thesis" then obtain u where u: "u>0" "u\1" "dist (x + u *\<^sub>R (y - x)) a < dist x a" using closer_point_lemma[of a x y] by auto let ?z = "(1 - u) *\<^sub>R x + u *\<^sub>R y" have "?z \ S" using convexD_alt[OF assms(1,3,4), of u] using u by auto then show False using assms(5)[THEN bspec[where x="?z"]] and u(3) by (auto simp: dist_commute algebra_simps) qed lemma any_closest_point_unique: fixes x :: "'a::real_inner" assumes "convex S" "closed S" "x \ S" "y \ S" "\z\S. dist a x \ dist a z" "\z\S. dist a y \ dist a z" shows "x = y" using any_closest_point_dot[OF assms(1-4,5)] and any_closest_point_dot[OF assms(1-2,4,3,6)] unfolding norm_pths(1) and norm_le_square by (auto simp: algebra_simps) lemma closest_point_unique: assumes "convex S" "closed S" "x \ S" "\z\S. dist a x \ dist a z" shows "x = closest_point S a" using any_closest_point_unique[OF assms(1-3) _ assms(4), of "closest_point S a"] using closest_point_exists[OF assms(2)] and assms(3) by auto lemma closest_point_dot: assumes "convex S" "closed S" "x \ S" shows "inner (a - closest_point S a) (x - closest_point S a) \ 0" using any_closest_point_dot[OF assms(1,2) _ assms(3)] by (metis assms(2) assms(3) closest_point_in_set closest_point_le empty_iff) lemma closest_point_lt: assumes "convex S" "closed S" "x \ S" "x \ closest_point S a" shows "dist a (closest_point S a) < dist a x" using closest_point_unique[where a=a] closest_point_le[where a=a] assms by fastforce lemma setdist_closest_point: "\closed S; S \ {}\ \ setdist {a} S = dist a (closest_point S a)" by (metis closest_point_exists(2) closest_point_in_set emptyE insert_iff setdist_unique) lemma closest_point_lipschitz: assumes "convex S" and "closed S" "S \ {}" shows "dist (closest_point S x) (closest_point S y) \ dist x y" proof - have "inner (x - closest_point S x) (closest_point S y - closest_point S x) \ 0" and "inner (y - closest_point S y) (closest_point S x - closest_point S y) \ 0" by (simp_all add: assms closest_point_dot closest_point_in_set) then show ?thesis unfolding dist_norm and norm_le using inner_ge_zero[of "(x - closest_point S x) - (y - closest_point S y)"] by (simp add: inner_add inner_diff inner_commute) qed lemma continuous_at_closest_point: assumes "convex S" and "closed S" and "S \ {}" shows "continuous (at x) (closest_point S)" unfolding continuous_at_eps_delta using le_less_trans[OF closest_point_lipschitz[OF assms]] by auto lemma continuous_on_closest_point: assumes "convex S" and "closed S" and "S \ {}" shows "continuous_on t (closest_point S)" by (metis continuous_at_imp_continuous_on continuous_at_closest_point[OF assms]) proposition closest_point_in_rel_interior: assumes "closed S" "S \ {}" and x: "x \ affine hull S" shows "closest_point S x \ rel_interior S \ x \ rel_interior S" proof (cases "x \ S") case True then show ?thesis by (simp add: closest_point_self) next case False then have "False" if asm: "closest_point S x \ rel_interior S" proof - obtain e where "e > 0" and clox: "closest_point S x \ S" and e: "cball (closest_point S x) e \ affine hull S \ S" using asm mem_rel_interior_cball by blast then have clo_notx: "closest_point S x \ x" using \x \ S\ by auto define y where "y \ closest_point S x - (min 1 (e / norm(closest_point S x - x))) *\<^sub>R (closest_point S x - x)" have "x - y = (1 - min 1 (e / norm (closest_point S x - x))) *\<^sub>R (x - closest_point S x)" by (simp add: y_def algebra_simps) then have "norm (x - y) = abs ((1 - min 1 (e / norm (closest_point S x - x)))) * norm(x - closest_point S x)" by simp also have "\ < norm(x - closest_point S x)" using clo_notx \e > 0\ by (auto simp: mult_less_cancel_right2 field_split_simps) finally have no_less: "norm (x - y) < norm (x - closest_point S x)" . have "y \ affine hull S" unfolding y_def by (meson affine_affine_hull clox hull_subset mem_affine_3_minus2 subsetD x) moreover have "dist (closest_point S x) y \ e" using \e > 0\ by (auto simp: y_def min_mult_distrib_right) ultimately have "y \ S" using subsetD [OF e] by simp then have "dist x (closest_point S x) \ dist x y" by (simp add: closest_point_le \closed S\) with no_less show False by (simp add: dist_norm) qed moreover have "x \ rel_interior S" using rel_interior_subset False by blast ultimately show ?thesis by blast qed subsubsection\<^marker>\tag unimportant\ \Various point-to-set separating/supporting hyperplane theorems\ lemma supporting_hyperplane_closed_point: fixes z :: "'a::{real_inner,heine_borel}" assumes "convex S" and "closed S" and "S \ {}" and "z \ S" shows "\a b. \y\S. inner a z < b \ inner a y = b \ (\x\S. inner a x \ b)" proof - obtain y where "y \ S" and y: "\x\S. dist z y \ dist z x" by (metis distance_attains_inf[OF assms(2-3)]) show ?thesis proof (intro exI bexI conjI ballI) show "(y - z) \ z < (y - z) \ y" by (metis \y \ S\ assms(4) diff_gt_0_iff_gt inner_commute inner_diff_left inner_gt_zero_iff right_minus_eq) show "(y - z) \ y \ (y - z) \ x" if "x \ S" for x proof (rule ccontr) have *: "\u. 0 \ u \ u \ 1 \ dist z y \ dist z ((1 - u) *\<^sub>R y + u *\<^sub>R x)" using assms(1)[unfolded convex_alt] and y and \x\S\ and \y\S\ by auto assume "\ (y - z) \ y \ (y - z) \ x" then obtain v where "v > 0" "v \ 1" "dist (y + v *\<^sub>R (x - y)) z < dist y z" using closer_point_lemma[of z y x] by (auto simp: inner_diff) then show False using *[of v] by (auto simp: dist_commute algebra_simps) qed qed (use \y \ S\ in auto) qed lemma separating_hyperplane_closed_point: fixes z :: "'a::{real_inner,heine_borel}" assumes "convex S" and "closed S" and "z \ S" shows "\a b. inner a z < b \ (\x\S. inner a x > b)" proof (cases "S = {}") case True then show ?thesis by (simp add: gt_ex) next case False obtain y where "y \ S" and y: "\x. x \ S \ dist z y \ dist z x" by (metis distance_attains_inf[OF assms(2) False]) show ?thesis proof (intro exI conjI ballI) show "(y - z) \ z < inner (y - z) z + (norm (y - z))\<^sup>2 / 2" using \y\S\ \z\S\ by auto next fix x assume "x \ S" have "False" if *: "0 < inner (z - y) (x - y)" proof - obtain u where "u > 0" "u \ 1" "dist (y + u *\<^sub>R (x - y)) z < dist y z" using * closer_point_lemma by blast then show False using y[of "y + u *\<^sub>R (x - y)"] convexD_alt [OF \convex S\] using \x\S\ \y\S\ by (auto simp: dist_commute algebra_simps) qed moreover have "0 < (norm (y - z))\<^sup>2" using \y\S\ \z\S\ by auto then have "0 < inner (y - z) (y - z)" unfolding power2_norm_eq_inner by simp ultimately show "(y - z) \ z + (norm (y - z))\<^sup>2 / 2 < (y - z) \ x" by (force simp: field_simps power2_norm_eq_inner inner_commute inner_diff) qed qed lemma separating_hyperplane_closed_0: assumes "convex (S::('a::euclidean_space) set)" and "closed S" and "0 \ S" shows "\a b. a \ 0 \ 0 < b \ (\x\S. inner a x > b)" proof (cases "S = {}") case True have "(SOME i. i\Basis) \ (0::'a)" by (metis Basis_zero SOME_Basis) then show ?thesis using True zero_less_one by blast next case False then show ?thesis using False using separating_hyperplane_closed_point[OF assms] by (metis all_not_in_conv inner_zero_left inner_zero_right less_eq_real_def not_le) qed subsubsection\<^marker>\tag unimportant\ \Now set-to-set for closed/compact sets\ lemma separating_hyperplane_closed_compact: fixes S :: "'a::euclidean_space set" assumes "convex S" and "closed S" and "convex T" and "compact T" and "T \ {}" and "S \ T = {}" shows "\a b. (\x\S. inner a x < b) \ (\x\T. inner a x > b)" proof (cases "S = {}") case True obtain b where b: "b > 0" "\x\T. norm x \ b" using compact_imp_bounded[OF assms(4)] unfolding bounded_pos by auto obtain z :: 'a where z: "norm z = b + 1" using vector_choose_size[of "b + 1"] and b(1) by auto then have "z \ T" using b(2)[THEN bspec[where x=z]] by auto then obtain a b where ab: "inner a z < b" "\x\T. b < inner a x" using separating_hyperplane_closed_point[OF assms(3) compact_imp_closed[OF assms(4)], of z] by auto then show ?thesis using True by auto next case False then obtain y where "y \ S" by auto obtain a b where "0 < b" and \
: "\x. x \ (\x\ S. \y \ T. {x - y}) \ b < inner a x" using separating_hyperplane_closed_point[OF convex_differences[OF assms(1,3)], of 0] using closed_compact_differences assms by fastforce have ab: "b + inner a y < inner a x" if "x\S" "y\T" for x y using \
[of "x-y"] that by (auto simp add: inner_diff_right less_diff_eq) define k where "k = (SUP x\T. a \ x)" have "k + b / 2 < a \ x" if "x \ S" for x proof - have "k \ inner a x - b" unfolding k_def using \T \ {}\ ab that by (fastforce intro: cSUP_least) then show ?thesis using \0 < b\ by auto qed moreover have "- (k + b / 2) < - a \ x" if "x \ T" for x proof - have "inner a x - b / 2 < k" unfolding k_def proof (subst less_cSUP_iff) show "T \ {}" by fact show "bdd_above ((\) a ` T)" using ab[rule_format, of y] \y \ S\ by (intro bdd_aboveI2[where M="inner a y - b"]) (auto simp: field_simps intro: less_imp_le) show "\y\T. a \ x - b / 2 < a \ y" using \0 < b\ that by force qed then show ?thesis by auto qed ultimately show ?thesis by (metis inner_minus_left neg_less_iff_less) qed lemma separating_hyperplane_compact_closed: fixes S :: "'a::euclidean_space set" assumes "convex S" and "compact S" and "S \ {}" and "convex T" and "closed T" and "S \ T = {}" shows "\a b. (\x\S. inner a x < b) \ (\x\T. inner a x > b)" proof - obtain a b where "(\x\T. inner a x < b) \ (\x\S. b < inner a x)" by (metis disjoint_iff_not_equal separating_hyperplane_closed_compact assms) then show ?thesis by (metis inner_minus_left neg_less_iff_less) qed subsubsection\<^marker>\tag unimportant\ \General case without assuming closure and getting non-strict separation\ lemma separating_hyperplane_set_0: assumes "convex S" "(0::'a::euclidean_space) \ S" shows "\a. a \ 0 \ (\x\S. 0 \ inner a x)" proof - let ?k = "\c. {x::'a. 0 \ inner c x}" have *: "frontier (cball 0 1) \ \f \ {}" if as: "f \ ?k ` S" "finite f" for f proof - obtain c where c: "f = ?k ` c" "c \ S" "finite c" using finite_subset_image[OF as(2,1)] by auto then obtain a b where ab: "a \ 0" "0 < b" "\x\convex hull c. b < inner a x" using separating_hyperplane_closed_0[OF convex_convex_hull, of c] using finite_imp_compact_convex_hull[OF c(3), THEN compact_imp_closed] and assms(2) using subset_hull[of convex, OF assms(1), symmetric, of c] by force have "norm (a /\<^sub>R norm a) = 1" by (simp add: ab(1)) moreover have "(\y\c. 0 \ y \ (a /\<^sub>R norm a))" using hull_subset[of c convex] ab by (force simp: inner_commute) ultimately have "\x. norm x = 1 \ (\y\c. 0 \ inner y x)" by blast then show "frontier (cball 0 1) \ \f \ {}" unfolding c(1) frontier_cball sphere_def dist_norm by auto qed have "frontier (cball 0 1) \ (\(?k ` S)) \ {}" by (rule compact_imp_fip) (use * closed_halfspace_ge in auto) then obtain x where "norm x = 1" "\y\S. x\?k y" unfolding frontier_cball dist_norm sphere_def by auto then show ?thesis by (metis inner_commute mem_Collect_eq norm_eq_zero zero_neq_one) qed lemma separating_hyperplane_sets: fixes S T :: "'a::euclidean_space set" assumes "convex S" and "convex T" and "S \ {}" and "T \ {}" and "S \ T = {}" shows "\a b. a \ 0 \ (\x\S. inner a x \ b) \ (\x\T. inner a x \ b)" proof - from separating_hyperplane_set_0[OF convex_differences[OF assms(2,1)]] obtain a where "a \ 0" "\x\{x - y |x y. x \ T \ y \ S}. 0 \ inner a x" using assms(3-5) by force then have *: "\x y. x \ T \ y \ S \ inner a y \ inner a x" by (force simp: inner_diff) then have bdd: "bdd_above (((\) a)`S)" using \T \ {}\ by (auto intro: bdd_aboveI2[OF *]) show ?thesis using \a\0\ by (intro exI[of _ a] exI[of _ "SUP x\S. a \ x"]) (auto intro!: cSUP_upper bdd cSUP_least \a \ 0\ \S \ {}\ *) qed subsection\<^marker>\tag unimportant\ \More convexity generalities\ lemma convex_closure [intro,simp]: fixes S :: "'a::real_normed_vector set" assumes "convex S" shows "convex (closure S)" apply (rule convexI) unfolding closure_sequential apply (elim exE) subgoal for x y u v f g by (rule_tac x="\n. u *\<^sub>R f n + v *\<^sub>R g n" in exI) (force intro: tendsto_intros dest: convexD [OF assms]) done lemma convex_interior [intro,simp]: fixes S :: "'a::real_normed_vector set" assumes "convex S" shows "convex (interior S)" unfolding convex_alt Ball_def mem_interior proof clarify fix x y u assume u: "0 \ u" "u \ (1::real)" fix e d assume ed: "ball x e \ S" "ball y d \ S" "0e>0. ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) e \ S" proof (intro exI conjI subsetI) fix z assume z: "z \ ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) (min d e)" have "(1- u) *\<^sub>R (z - u *\<^sub>R (y - x)) + u *\<^sub>R (z + (1 - u) *\<^sub>R (y - x)) \ S" proof (rule_tac assms[unfolded convex_alt, rule_format]) show "z - u *\<^sub>R (y - x) \ S" "z + (1 - u) *\<^sub>R (y - x) \ S" using ed z u by (auto simp add: algebra_simps dist_norm) qed (use u in auto) then show "z \ S" using u by (auto simp: algebra_simps) qed(use u ed in auto) qed lemma convex_hull_eq_empty[simp]: "convex hull S = {} \ S = {}" using hull_subset[of S convex] convex_hull_empty by auto subsection\<^marker>\tag unimportant\ \Convex set as intersection of halfspaces\ lemma convex_halfspace_intersection: fixes S :: "('a::euclidean_space) set" assumes "closed S" "convex S" shows "S = \{h. S \ h \ (\a b. h = {x. inner a x \ b})}" proof - { fix z assume "\T. S \ T \ (\a b. T = {x. inner a x \ b}) \ z \ T" "z \ S" then have \
: "\a b. S \ {x. inner a x \ b} \ z \ {x. inner a x \ b}" by blast obtain a b where "inner a z < b" "(\x\S. inner a x > b)" using \z \ S\ assms separating_hyperplane_closed_point by blast then have False using \
[of "-a" "-b"] by fastforce } then show ?thesis by force qed subsection\<^marker>\tag unimportant\ \Convexity of general and special intervals\ lemma is_interval_convex: fixes S :: "'a::euclidean_space set" assumes "is_interval S" shows "convex S" proof (rule convexI) fix x y and u v :: real assume "x \ S" "y \ S" and uv: "0 \ u" "0 \ v" "u + v = 1" then have *: "u = 1 - v" "1 - v \ 0" and **: "v = 1 - u" "1 - u \ 0" by auto { fix a b assume "\ b \ u * a + v * b" then have "u * a < (1 - v) * b" unfolding not_le using \0 \ v\by (auto simp: field_simps) then have "a < b" using "*"(1) less_eq_real_def uv(1) by auto then have "a \ u * a + v * b" unfolding * using \0 \ v\ by (auto simp: field_simps intro!:mult_right_mono) } moreover { fix a b assume "\ u * a + v * b \ a" then have "v * b > (1 - u) * a" unfolding not_le using \0 \ v\ by (auto simp: field_simps) then have "a < b" unfolding * using \0 \ v\ by (rule_tac mult_left_less_imp_less) (auto simp: field_simps) then have "u * a + v * b \ b" unfolding ** using **(2) \0 \ u\ by (auto simp: algebra_simps mult_right_mono) } ultimately show "u *\<^sub>R x + v *\<^sub>R y \ S" using DIM_positive[where 'a='a] by (intro mem_is_intervalI [OF assms \x \ S\ \y \ S\]) (auto simp: inner_simps) qed lemma is_interval_connected: fixes S :: "'a::euclidean_space set" shows "is_interval S \ connected S" using is_interval_convex convex_connected by auto lemma convex_box [simp]: "convex (cbox a b)" "convex (box a (b::'a::euclidean_space))" by (auto simp add: is_interval_convex) text\A non-singleton connected set is perfect (i.e. has no isolated points). \ lemma connected_imp_perfect: fixes a :: "'a::metric_space" assumes "connected S" "a \ S" and S: "\x. S \ {x}" shows "a islimpt S" proof - have False if "a \ T" "open T" "\y. \y \ S; y \ T\ \ y = a" for T proof - obtain e where "e > 0" and e: "cball a e \ T" using \open T\ \a \ T\ by (auto simp: open_contains_cball) have "openin (top_of_set S) {a}" unfolding openin_open using that \a \ S\ by blast moreover have "closedin (top_of_set S) {a}" by (simp add: assms) ultimately show "False" using \connected S\ connected_clopen S by blast qed then show ?thesis unfolding islimpt_def by blast qed +lemma islimpt_Ioc [simp]: + fixes a :: real + assumes "a x \ {a..b}" (is "?lhs = ?rhs") +proof + show "?lhs \ ?rhs" + by (metis assms closed_atLeastAtMost closed_limpt closure_greaterThanAtMost closure_subset islimpt_subset) +next + assume ?rhs + then have "x \ closure {a<.. x \ {a..b}" + by (metis assms closure_atLeastLessThan closure_greaterThanAtMost islimpt_Ioc limpt_of_closure) + +lemma islimpt_Icc [simp]: + fixes a :: real + assumes "a x \ {a..b}" + by (metis assms closure_atLeastLessThan islimpt_Ico limpt_of_closure) + lemma connected_imp_perfect_aff_dim: "\connected S; aff_dim S \ 0; a \ S\ \ a islimpt S" using aff_dim_sing connected_imp_perfect by blast subsection\<^marker>\tag unimportant\ \On \real\, \is_interval\, \convex\ and \connected\ are all equivalent\ lemma mem_is_interval_1_I: fixes a b c::real assumes "is_interval S" assumes "a \ S" "c \ S" assumes "a \ b" "b \ c" shows "b \ S" using assms is_interval_1 by blast lemma is_interval_connected_1: fixes S :: "real set" shows "is_interval S \ connected S" by (meson connected_iff_interval is_interval_1) lemma is_interval_convex_1: fixes S :: "real set" shows "is_interval S \ convex S" by (metis is_interval_convex convex_connected is_interval_connected_1) lemma connected_compact_interval_1: "connected S \ compact S \ (\a b. S = {a..b::real})" by (auto simp: is_interval_connected_1 [symmetric] is_interval_compact) lemma connected_convex_1: fixes S :: "real set" shows "connected S \ convex S" by (metis is_interval_convex convex_connected is_interval_connected_1) lemma connected_convex_1_gen: fixes S :: "'a :: euclidean_space set" assumes "DIM('a) = 1" shows "connected S \ convex S" proof - obtain f:: "'a \ real" where linf: "linear f" and "inj f" using subspace_isomorphism[OF subspace_UNIV subspace_UNIV, where 'a='a and 'b=real] unfolding Euclidean_Space.dim_UNIV by (auto simp: assms) then have "f -` (f ` S) = S" by (simp add: inj_vimage_image_eq) then show ?thesis by (metis connected_convex_1 convex_linear_vimage linf convex_connected connected_linear_image) qed lemma [simp]: fixes r s::real shows is_interval_io: "is_interval {..\tag unimportant\ \Another intermediate value theorem formulation\ lemma ivt_increasing_component_on_1: fixes f :: "real \ 'a::euclidean_space" assumes "a \ b" and "continuous_on {a..b} f" and "(f a)\k \ y" "y \ (f b)\k" shows "\x\{a..b}. (f x)\k = y" proof - have "f a \ f ` cbox a b" "f b \ f ` cbox a b" using \a \ b\ by auto then show ?thesis using connected_ivt_component[of "f ` cbox a b" "f a" "f b" k y] by (simp add: connected_continuous_image assms) qed lemma ivt_increasing_component_1: fixes f :: "real \ 'a::euclidean_space" shows "a \ b \ \x\{a..b}. continuous (at x) f \ f a\k \ y \ y \ f b\k \ \x\{a..b}. (f x)\k = y" by (rule ivt_increasing_component_on_1) (auto simp: continuous_at_imp_continuous_on) lemma ivt_decreasing_component_on_1: fixes f :: "real \ 'a::euclidean_space" assumes "a \ b" and "continuous_on {a..b} f" and "(f b)\k \ y" and "y \ (f a)\k" shows "\x\{a..b}. (f x)\k = y" using ivt_increasing_component_on_1[of a b "\x. - f x" k "- y"] neg_equal_iff_equal using assms continuous_on_minus by force lemma ivt_decreasing_component_1: fixes f :: "real \ 'a::euclidean_space" shows "a \ b \ \x\{a..b}. continuous (at x) f \ f b\k \ y \ y \ f a\k \ \x\{a..b}. (f x)\k = y" by (rule ivt_decreasing_component_on_1) (auto simp: continuous_at_imp_continuous_on) subsection\<^marker>\tag unimportant\ \A bound within an interval\ lemma convex_hull_eq_real_cbox: fixes x y :: real assumes "x \ y" shows "convex hull {x, y} = cbox x y" proof (rule hull_unique) show "{x, y} \ cbox x y" using \x \ y\ by auto show "convex (cbox x y)" by (rule convex_box) next fix S assume "{x, y} \ S" and "convex S" then show "cbox x y \ S" unfolding is_interval_convex_1 [symmetric] is_interval_def Basis_real_def by - (clarify, simp (no_asm_use), fast) qed lemma unit_interval_convex_hull: "cbox (0::'a::euclidean_space) One = convex hull {x. \i\Basis. (x\i = 0) \ (x\i = 1)}" (is "?int = convex hull ?points") proof - have One[simp]: "\i. i \ Basis \ One \ i = 1" by (simp add: inner_sum_left sum.If_cases inner_Basis) have "?int = {x. \i\Basis. x \ i \ cbox 0 1}" by (auto simp: cbox_def) also have "\ = (\i\Basis. (\x. x *\<^sub>R i) ` cbox 0 1)" by (simp only: box_eq_set_sum_Basis) also have "\ = (\i\Basis. (\x. x *\<^sub>R i) ` (convex hull {0, 1}))" by (simp only: convex_hull_eq_real_cbox zero_le_one) also have "\ = (\i\Basis. convex hull ((\x. x *\<^sub>R i) ` {0, 1}))" by (simp add: convex_hull_linear_image) also have "\ = convex hull (\i\Basis. (\x. x *\<^sub>R i) ` {0, 1})" by (simp only: convex_hull_set_sum) also have "\ = convex hull {x. \i\Basis. x\i \ {0, 1}}" by (simp only: box_eq_set_sum_Basis) also have "convex hull {x. \i\Basis. x\i \ {0, 1}} = convex hull ?points" by simp finally show ?thesis . qed text \And this is a finite set of vertices.\ lemma unit_cube_convex_hull: obtains S :: "'a::euclidean_space set" where "finite S" and "cbox 0 (\Basis) = convex hull S" proof show "finite {x::'a. \i\Basis. x \ i = 0 \ x \ i = 1}" proof (rule finite_subset, clarify) show "finite ((\S. \i\Basis. (if i \ S then 1 else 0) *\<^sub>R i) ` Pow Basis)" using finite_Basis by blast fix x :: 'a assume x: "\i\Basis. x \ i = 0 \ x \ i = 1" show "x \ (\S. \i\Basis. (if i\S then 1 else 0) *\<^sub>R i) ` Pow Basis" apply (rule image_eqI[where x="{i. i \ Basis \ x\i = 1}"]) using x by (subst euclidean_eq_iff, auto) qed show "cbox 0 One = convex hull {x. \i\Basis. x \ i = 0 \ x \ i = 1}" using unit_interval_convex_hull by blast qed text \Hence any cube (could do any nonempty interval).\ lemma cube_convex_hull: assumes "d > 0" obtains S :: "'a::euclidean_space set" where "finite S" and "cbox (x - (\i\Basis. d*\<^sub>Ri)) (x + (\i\Basis. d*\<^sub>Ri)) = convex hull S" proof - let ?d = "(\i\Basis. d *\<^sub>R i)::'a" have *: "cbox (x - ?d) (x + ?d) = (\y. x - ?d + (2 * d) *\<^sub>R y) ` cbox 0 (\Basis)" proof (intro set_eqI iffI) fix y assume "y \ cbox (x - ?d) (x + ?d)" then have "inverse (2 * d) *\<^sub>R (y - (x - ?d)) \ cbox 0 (\Basis)" using assms by (simp add: mem_box inner_simps) (simp add: field_simps) with \0 < d\ show "y \ (\y. x - sum ((*\<^sub>R) d) Basis + (2 * d) *\<^sub>R y) ` cbox 0 One" by (auto intro: image_eqI[where x= "inverse (2 * d) *\<^sub>R (y - (x - ?d))"]) next fix y assume "y \ (\y. x - ?d + (2 * d) *\<^sub>R y) ` cbox 0 One" then obtain z where z: "z \ cbox 0 One" "y = x - ?d + (2*d) *\<^sub>R z" by auto then show "y \ cbox (x - ?d) (x + ?d)" using z assms by (auto simp: mem_box inner_simps) qed obtain S where "finite S" "cbox 0 (\Basis::'a) = convex hull S" using unit_cube_convex_hull by auto then show ?thesis by (rule_tac that[of "(\y. x - ?d + (2 * d) *\<^sub>R y)` S"]) (auto simp: convex_hull_affinity *) qed subsection\<^marker>\tag unimportant\\Representation of any interval as a finite convex hull\ lemma image_stretch_interval: "(\x. \k\Basis. (m k * (x\k)) *\<^sub>R k) ` cbox a (b::'a::euclidean_space) = (if (cbox a b) = {} then {} else cbox (\k\Basis. (min (m k * (a\k)) (m k * (b\k))) *\<^sub>R k::'a) (\k\Basis. (max (m k * (a\k)) (m k * (b\k))) *\<^sub>R k))" proof cases assume *: "cbox a b \ {}" show ?thesis unfolding box_ne_empty if_not_P[OF *] apply (simp add: cbox_def image_Collect set_eq_iff euclidean_eq_iff[where 'a='a] ball_conj_distrib[symmetric]) apply (subst choice_Basis_iff[symmetric]) proof (intro allI ball_cong refl) fix x i :: 'a assume "i \ Basis" with * have a_le_b: "a \ i \ b \ i" unfolding box_ne_empty by auto show "(\xa. x \ i = m i * xa \ a \ i \ xa \ xa \ b \ i) \ min (m i * (a \ i)) (m i * (b \ i)) \ x \ i \ x \ i \ max (m i * (a \ i)) (m i * (b \ i))" proof (cases "m i = 0") case True with a_le_b show ?thesis by auto next case False then have *: "\a b. a = m i * b \ b = a / m i" by (auto simp: field_simps) from False have "min (m i * (a \ i)) (m i * (b \ i)) = (if 0 < m i then m i * (a \ i) else m i * (b \ i))" "max (m i * (a \ i)) (m i * (b \ i)) = (if 0 < m i then m i * (b \ i) else m i * (a \ i))" using a_le_b by (auto simp: min_def max_def mult_le_cancel_left) with False show ?thesis using a_le_b * by (simp add: le_divide_eq divide_le_eq) (simp add: ac_simps) qed qed qed simp lemma interval_image_stretch_interval: "\u v. (\x. \k\Basis. (m k * (x\k))*\<^sub>R k) ` cbox a (b::'a::euclidean_space) = cbox u (v::'a::euclidean_space)" unfolding image_stretch_interval by auto lemma cbox_translation: "cbox (c + a) (c + b) = image (\x. c + x) (cbox a b)" using image_affinity_cbox [of 1 c a b] using box_ne_empty [of "a+c" "b+c"] box_ne_empty [of a b] by (auto simp: inner_left_distrib add.commute) lemma cbox_image_unit_interval: fixes a :: "'a::euclidean_space" assumes "cbox a b \ {}" shows "cbox a b = (+) a ` (\x. \k\Basis. ((b \ k - a \ k) * (x \ k)) *\<^sub>R k) ` cbox 0 One" using assms apply (simp add: box_ne_empty image_stretch_interval cbox_translation [symmetric]) apply (simp add: min_def max_def algebra_simps sum_subtractf euclidean_representation) done lemma closed_interval_as_convex_hull: fixes a :: "'a::euclidean_space" obtains S where "finite S" "cbox a b = convex hull S" proof (cases "cbox a b = {}") case True with convex_hull_empty that show ?thesis by blast next case False obtain S::"'a set" where "finite S" and eq: "cbox 0 One = convex hull S" by (blast intro: unit_cube_convex_hull) let ?S = "((+) a ` (\x. \k\Basis. ((b \ k - a \ k) * (x \ k)) *\<^sub>R k) ` S)" show thesis proof show "finite ?S" by (simp add: \finite S\) have lin: "linear (\x. \k\Basis. ((b \ k - a \ k) * (x \ k)) *\<^sub>R k)" by (rule linear_compose_sum) (auto simp: algebra_simps linearI) show "cbox a b = convex hull ?S" using convex_hull_linear_image [OF lin] by (simp add: convex_hull_translation eq cbox_image_unit_interval [OF False]) qed qed subsection\<^marker>\tag unimportant\ \Bounded convex function on open set is continuous\ lemma convex_on_bounded_continuous: fixes S :: "('a::real_normed_vector) set" assumes "open S" and "convex_on S f" and "\x\S. \f x\ \ b" shows "continuous_on S f" proof - have "\d>0. \x'. norm (x' - x) < d \ \f x' - f x\ < e" if "x \ S" "e > 0" for x and e :: real proof - define B where "B = \b\ + 1" then have B: "0 < B""\x. x\S \ \f x\ \ B" using assms(3) by auto obtain k where "k > 0" and k: "cball x k \ S" using \x \ S\ assms(1) open_contains_cball_eq by blast show "\d>0. \x'. norm (x' - x) < d \ \f x' - f x\ < e" proof (intro exI conjI allI impI) fix y assume as: "norm (y - x) < min (k / 2) (e / (2 * B) * k)" show "\f y - f x\ < e" proof (cases "y = x") case False define t where "t = k / norm (y - x)" have "2 < t" "0k>0\ by (auto simp:field_simps) have "y \ S" apply (rule k[THEN subsetD]) unfolding mem_cball dist_norm apply (rule order_trans[of _ "2 * norm (x - y)"]) using as by (auto simp: field_simps norm_minus_commute) { define w where "w = x + t *\<^sub>R (y - x)" have "w \ S" using \k>0\ by (auto simp: dist_norm t_def w_def k[THEN subsetD]) have "(1 / t) *\<^sub>R x + - x + ((t - 1) / t) *\<^sub>R x = (1 / t - 1 + (t - 1) / t) *\<^sub>R x" by (auto simp: algebra_simps) also have "\ = 0" using \t > 0\ by (auto simp:field_simps) finally have w: "(1 / t) *\<^sub>R w + ((t - 1) / t) *\<^sub>R x = y" unfolding w_def using False and \t > 0\ by (auto simp: algebra_simps) have 2: "2 * B < e * t" unfolding t_def using \0 < e\ \0 < k\ \B > 0\ and as and False by (auto simp:field_simps) have "f y - f x \ (f w - f x) / t" using assms(2)[unfolded convex_on_def,rule_format,of w x "1/t" "(t - 1)/t", unfolded w] using \0 < t\ \2 < t\ and \x \ S\ \w \ S\ by (auto simp:field_simps) also have "... < e" using B(2)[OF \w\S\] and B(2)[OF \x\S\] 2 \t > 0\ by (auto simp: field_simps) finally have th1: "f y - f x < e" . } moreover { define w where "w = x - t *\<^sub>R (y - x)" have "w \ S" using \k > 0\ by (auto simp: dist_norm t_def w_def k[THEN subsetD]) have "(1 / (1 + t)) *\<^sub>R x + (t / (1 + t)) *\<^sub>R x = (1 / (1 + t) + t / (1 + t)) *\<^sub>R x" by (auto simp: algebra_simps) also have "\ = x" using \t > 0\ by (auto simp:field_simps) finally have w: "(1 / (1+t)) *\<^sub>R w + (t / (1 + t)) *\<^sub>R y = x" unfolding w_def using False and \t > 0\ by (auto simp: algebra_simps) have "2 * B < e * t" unfolding t_def using \0 < e\ \0 < k\ \B > 0\ and as and False by (auto simp:field_simps) then have *: "(f w - f y) / t < e" using B(2)[OF \w\S\] and B(2)[OF \y\S\] using \t > 0\ by (auto simp:field_simps) have "f x \ 1 / (1 + t) * f w + (t / (1 + t)) * f y" using assms(2)[unfolded convex_on_def,rule_format,of w y "1/(1+t)" "t / (1+t)",unfolded w] using \0 < t\ \2 < t\ and \y \ S\ \w \ S\ by (auto simp:field_simps) also have "\ = (f w + t * f y) / (1 + t)" using \t > 0\ by (simp add: add_divide_distrib) also have "\ < e + f y" using \t > 0\ * \e > 0\ by (auto simp: field_simps) finally have "f x - f y < e" by auto } ultimately show ?thesis by auto qed (use \0 in auto) qed (use \0 \0 \0 in \auto simp: field_simps\) qed then show ?thesis by (metis continuous_on_iff dist_norm real_norm_def) qed subsection\<^marker>\tag unimportant\ \Upper bound on a ball implies upper and lower bounds\ lemma convex_bounds_lemma: fixes x :: "'a::real_normed_vector" assumes "convex_on (cball x e) f" and "\y \ cball x e. f y \ b" and y: "y \ cball x e" shows "\f y\ \ b + 2 * \f x\" proof (cases "0 \ e") case True define z where "z = 2 *\<^sub>R x - y" have *: "x - (2 *\<^sub>R x - y) = y - x" by (simp add: scaleR_2) have z: "z \ cball x e" using y unfolding z_def mem_cball dist_norm * by (auto simp: norm_minus_commute) have "(1 / 2) *\<^sub>R y + (1 / 2) *\<^sub>R z = x" unfolding z_def by (auto simp: algebra_simps) then show "\f y\ \ b + 2 * \f x\" using assms(1)[unfolded convex_on_def,rule_format, OF y z, of "1/2" "1/2"] using assms(2)[rule_format,OF y] assms(2)[rule_format,OF z] by (auto simp:field_simps) next case False have "dist x y < 0" using False y unfolding mem_cball not_le by (auto simp del: dist_not_less_zero) then show "\f y\ \ b + 2 * \f x\" using zero_le_dist[of x y] by auto qed subsubsection\<^marker>\tag unimportant\ \Hence a convex function on an open set is continuous\ lemma real_of_nat_ge_one_iff: "1 \ real (n::nat) \ 1 \ n" by auto lemma convex_on_continuous: assumes "open (s::('a::euclidean_space) set)" "convex_on s f" shows "continuous_on s f" unfolding continuous_on_eq_continuous_at[OF assms(1)] proof note dimge1 = DIM_positive[where 'a='a] fix x assume "x \ s" then obtain e where e: "cball x e \ s" "e > 0" using assms(1) unfolding open_contains_cball by auto define d where "d = e / real DIM('a)" have "0 < d" unfolding d_def using \e > 0\ dimge1 by auto let ?d = "(\i\Basis. d *\<^sub>R i)::'a" obtain c where c: "finite c" and c1: "convex hull c \ cball x e" and c2: "cball x d \ convex hull c" proof define c where "c = (\i\Basis. (\a. a *\<^sub>R i) ` {x\i - d, x\i + d})" show "finite c" unfolding c_def by (simp add: finite_set_sum) have "\i. i \ Basis \ convex hull {x \ i - d, x \ i + d} = cbox (x \ i - d) (x \ i + d)" using \0 < d\ convex_hull_eq_real_cbox by auto then have 1: "convex hull c = {a. \i\Basis. a \ i \ cbox (x \ i - d) (x \ i + d)}" unfolding box_eq_set_sum_Basis c_def convex_hull_set_sum apply (subst convex_hull_linear_image [symmetric]) by (force simp add: linear_iff scaleR_add_left)+ then have 2: "convex hull c = {a. \i\Basis. a \ i \ cball (x \ i) d}" by (simp add: dist_norm abs_le_iff algebra_simps) show "cball x d \ convex hull c" unfolding 2 by (clarsimp simp: dist_norm) (metis inner_commute inner_diff_right norm_bound_Basis_le) have e': "e = (\(i::'a)\Basis. d)" by (simp add: d_def) show "convex hull c \ cball x e" unfolding 2 proof clarsimp show "dist x y \ e" if "\i\Basis. dist (x \ i) (y \ i) \ d" for y proof - have "\i. i \ Basis \ 0 \ dist (x \ i) (y \ i)" by simp have "(\i\Basis. dist (x \ i) (y \ i)) \ e" using e' sum_mono that by fastforce then show ?thesis by (metis (mono_tags) euclidean_dist_l2 order_trans [OF L2_set_le_sum] zero_le_dist) qed qed qed define k where "k = Max (f ` c)" have "convex_on (convex hull c) f" using assms(2) c1 convex_on_subset e(1) by blast then have k: "\y\convex hull c. f y \ k" using c convex_on_convex_hull_bound k_def by fastforce have "e \ e * real DIM('a)" using e(2) real_of_nat_ge_one_iff by auto then have "d \ e" by (simp add: d_def field_split_simps) then have dsube: "cball x d \ cball x e" by (rule subset_cball) have conv: "convex_on (cball x d) f" using \convex_on (convex hull c) f\ c2 convex_on_subset by blast then have "\y. y\cball x d \ \f y\ \ k + 2 * \f x\" by (rule convex_bounds_lemma) (use c2 k in blast) then have "continuous_on (ball x d) f" by (meson Elementary_Metric_Spaces.open_ball ball_subset_cball conv convex_on_bounded_continuous convex_on_subset mem_ball_imp_mem_cball) then show "continuous (at x) f" unfolding continuous_on_eq_continuous_at[OF open_ball] using \d > 0\ by auto qed end diff --git a/src/HOL/Analysis/Derivative.thy b/src/HOL/Analysis/Derivative.thy --- a/src/HOL/Analysis/Derivative.thy +++ b/src/HOL/Analysis/Derivative.thy @@ -1,3475 +1,3504 @@ (* Title: HOL/Analysis/Derivative.thy Author: John Harrison Author: Robert Himmelmann, TU Muenchen (translation from HOL Light); tidied by LCP *) section \Derivative\ theory Derivative imports Bounded_Linear_Function Line_Segment Convex_Euclidean_Space begin declare bounded_linear_inner_left [intro] declare has_derivative_bounded_linear[dest] subsection \Derivatives\ lemma has_derivative_add_const: "(f has_derivative f') net \ ((\x. f x + c) has_derivative f') net" by (intro derivative_eq_intros) auto subsection\<^marker>\tag unimportant\ \Derivative with composed bilinear function\ text \More explicit epsilon-delta forms.\ proposition has_derivative_within': "(f has_derivative f')(at x within s) \ bounded_linear f' \ (\e>0. \d>0. \x'\s. 0 < norm (x' - x) \ norm (x' - x) < d \ norm (f x' - f x - f'(x' - x)) / norm (x' - x) < e)" unfolding has_derivative_within Lim_within dist_norm by (simp add: diff_diff_eq) lemma has_derivative_at': "(f has_derivative f') (at x) \ bounded_linear f' \ (\e>0. \d>0. \x'. 0 < norm (x' - x) \ norm (x' - x) < d \ norm (f x' - f x - f'(x' - x)) / norm (x' - x) < e)" using has_derivative_within' [of f f' x UNIV] by simp lemma has_derivative_componentwise_within: "(f has_derivative f') (at a within S) \ (\i \ Basis. ((\x. f x \ i) has_derivative (\x. f' x \ i)) (at a within S))" apply (simp add: has_derivative_within) apply (subst tendsto_componentwise_iff) apply (simp add: bounded_linear_componentwise_iff [symmetric] ball_conj_distrib) apply (simp add: algebra_simps) done lemma has_derivative_at_withinI: "(f has_derivative f') (at x) \ (f has_derivative f') (at x within s)" unfolding has_derivative_within' has_derivative_at' by blast lemma has_derivative_right: fixes f :: "real \ real" and y :: "real" shows "(f has_derivative ((*) y)) (at x within ({x <..} \ I)) \ ((\t. (f x - f t) / (x - t)) \ y) (at x within ({x <..} \ I))" proof - have "((\t. (f t - (f x + y * (t - x))) / \t - x\) \ 0) (at x within ({x<..} \ I)) \ ((\t. (f t - f x) / (t - x) - y) \ 0) (at x within ({x<..} \ I))" by (intro Lim_cong_within) (auto simp add: diff_divide_distrib add_divide_distrib) also have "\ \ ((\t. (f t - f x) / (t - x)) \ y) (at x within ({x<..} \ I))" by (simp add: Lim_null[symmetric]) also have "\ \ ((\t. (f x - f t) / (x - t)) \ y) (at x within ({x<..} \ I))" by (intro Lim_cong_within) (simp_all add: field_simps) finally show ?thesis by (simp add: bounded_linear_mult_right has_derivative_within) qed subsubsection \Caratheodory characterization\ lemma DERIV_caratheodory_within: "(f has_field_derivative l) (at x within S) \ (\g. (\z. f z - f x = g z * (z - x)) \ continuous (at x within S) g \ g x = l)" (is "?lhs = ?rhs") proof assume ?lhs show ?rhs proof (intro exI conjI) let ?g = "(%z. if z = x then l else (f z - f x) / (z-x))" show "\z. f z - f x = ?g z * (z-x)" by simp show "continuous (at x within S) ?g" using \?lhs\ by (auto simp add: continuous_within has_field_derivative_iff cong: Lim_cong_within) show "?g x = l" by simp qed next assume ?rhs then obtain g where "(\z. f z - f x = g z * (z-x))" and "continuous (at x within S) g" and "g x = l" by blast thus ?lhs by (auto simp add: continuous_within has_field_derivative_iff cong: Lim_cong_within) qed subsection \Differentiability\ definition\<^marker>\tag important\ differentiable_on :: "('a::real_normed_vector \ 'b::real_normed_vector) \ 'a set \ bool" (infix "differentiable'_on" 50) where "f differentiable_on s \ (\x\s. f differentiable (at x within s))" lemma differentiableI: "(f has_derivative f') net \ f differentiable net" unfolding differentiable_def by auto lemma differentiable_onD: "\f differentiable_on S; x \ S\ \ f differentiable (at x within S)" using differentiable_on_def by blast lemma differentiable_at_withinI: "f differentiable (at x) \ f differentiable (at x within s)" unfolding differentiable_def using has_derivative_at_withinI by blast lemma differentiable_at_imp_differentiable_on: "(\x. x \ s \ f differentiable at x) \ f differentiable_on s" by (metis differentiable_at_withinI differentiable_on_def) corollary\<^marker>\tag unimportant\ differentiable_iff_scaleR: fixes f :: "real \ 'a::real_normed_vector" shows "f differentiable F \ (\d. (f has_derivative (\x. x *\<^sub>R d)) F)" by (auto simp: differentiable_def dest: has_derivative_linear linear_imp_scaleR) lemma differentiable_on_eq_differentiable_at: "open s \ f differentiable_on s \ (\x\s. f differentiable at x)" unfolding differentiable_on_def by (metis at_within_interior interior_open) lemma differentiable_transform_within: assumes "f differentiable (at x within s)" and "0 < d" and "x \ s" and "\x'. \x'\s; dist x' x < d\ \ f x' = g x'" shows "g differentiable (at x within s)" using assms has_derivative_transform_within unfolding differentiable_def by blast lemma differentiable_on_ident [simp, derivative_intros]: "(\x. x) differentiable_on S" by (simp add: differentiable_at_imp_differentiable_on) lemma differentiable_on_id [simp, derivative_intros]: "id differentiable_on S" by (simp add: id_def) lemma differentiable_on_const [simp, derivative_intros]: "(\z. c) differentiable_on S" by (simp add: differentiable_on_def) lemma differentiable_on_mult [simp, derivative_intros]: fixes f :: "'M::real_normed_vector \ 'a::real_normed_algebra" shows "\f differentiable_on S; g differentiable_on S\ \ (\z. f z * g z) differentiable_on S" unfolding differentiable_on_def differentiable_def using differentiable_def differentiable_mult by blast lemma differentiable_on_compose: "\g differentiable_on S; f differentiable_on (g ` S)\ \ (\x. f (g x)) differentiable_on S" by (simp add: differentiable_in_compose differentiable_on_def) lemma bounded_linear_imp_differentiable_on: "bounded_linear f \ f differentiable_on S" by (simp add: differentiable_on_def bounded_linear_imp_differentiable) lemma linear_imp_differentiable_on: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" shows "linear f \ f differentiable_on S" by (simp add: differentiable_on_def linear_imp_differentiable) lemma differentiable_on_minus [simp, derivative_intros]: "f differentiable_on S \ (\z. -(f z)) differentiable_on S" by (simp add: differentiable_on_def) lemma differentiable_on_add [simp, derivative_intros]: "\f differentiable_on S; g differentiable_on S\ \ (\z. f z + g z) differentiable_on S" by (simp add: differentiable_on_def) lemma differentiable_on_diff [simp, derivative_intros]: "\f differentiable_on S; g differentiable_on S\ \ (\z. f z - g z) differentiable_on S" by (simp add: differentiable_on_def) lemma differentiable_on_inverse [simp, derivative_intros]: fixes f :: "'a :: real_normed_vector \ 'b :: real_normed_field" shows "f differentiable_on S \ (\x. x \ S \ f x \ 0) \ (\x. inverse (f x)) differentiable_on S" by (simp add: differentiable_on_def) lemma differentiable_on_scaleR [derivative_intros, simp]: "\f differentiable_on S; g differentiable_on S\ \ (\x. f x *\<^sub>R g x) differentiable_on S" unfolding differentiable_on_def by (blast intro: differentiable_scaleR) lemma has_derivative_sqnorm_at [derivative_intros, simp]: "((\x. (norm x)\<^sup>2) has_derivative (\x. 2 *\<^sub>R (a \ x))) (at a)" using bounded_bilinear.FDERIV [of "(\)" id id a _ id id] by (auto simp: inner_commute dot_square_norm bounded_bilinear_inner) lemma differentiable_sqnorm_at [derivative_intros, simp]: fixes a :: "'a :: {real_normed_vector,real_inner}" shows "(\x. (norm x)\<^sup>2) differentiable (at a)" by (force simp add: differentiable_def intro: has_derivative_sqnorm_at) lemma differentiable_on_sqnorm [derivative_intros, simp]: fixes S :: "'a :: {real_normed_vector,real_inner} set" shows "(\x. (norm x)\<^sup>2) differentiable_on S" by (simp add: differentiable_at_imp_differentiable_on) lemma differentiable_norm_at [derivative_intros, simp]: fixes a :: "'a :: {real_normed_vector,real_inner}" shows "a \ 0 \ norm differentiable (at a)" using differentiableI has_derivative_norm by blast lemma differentiable_on_norm [derivative_intros, simp]: fixes S :: "'a :: {real_normed_vector,real_inner} set" shows "0 \ S \ norm differentiable_on S" by (metis differentiable_at_imp_differentiable_on differentiable_norm_at) subsection \Frechet derivative and Jacobian matrix\ definition "frechet_derivative f net = (SOME f'. (f has_derivative f') net)" proposition frechet_derivative_works: "f differentiable net \ (f has_derivative (frechet_derivative f net)) net" unfolding frechet_derivative_def differentiable_def unfolding some_eq_ex[of "\ f' . (f has_derivative f') net"] .. lemma linear_frechet_derivative: "f differentiable net \ linear (frechet_derivative f net)" unfolding frechet_derivative_works has_derivative_def by (auto intro: bounded_linear.linear) lemma frechet_derivative_const [simp]: "frechet_derivative (\x. c) (at a) = (\x. 0)" using differentiable_const frechet_derivative_works has_derivative_const has_derivative_unique by blast lemma frechet_derivative_id [simp]: "frechet_derivative id (at a) = id" using differentiable_def frechet_derivative_works has_derivative_id has_derivative_unique by blast lemma frechet_derivative_ident [simp]: "frechet_derivative (\x. x) (at a) = (\x. x)" by (metis eq_id_iff frechet_derivative_id) subsection \Differentiability implies continuity\ proposition differentiable_imp_continuous_within: "f differentiable (at x within s) \ continuous (at x within s) f" by (auto simp: differentiable_def intro: has_derivative_continuous) lemma differentiable_imp_continuous_on: "f differentiable_on s \ continuous_on s f" unfolding differentiable_on_def continuous_on_eq_continuous_within using differentiable_imp_continuous_within by blast lemma differentiable_on_subset: "f differentiable_on t \ s \ t \ f differentiable_on s" unfolding differentiable_on_def using differentiable_within_subset by blast lemma differentiable_on_empty: "f differentiable_on {}" unfolding differentiable_on_def by auto lemma has_derivative_continuous_on: "(\x. x \ s \ (f has_derivative f' x) (at x within s)) \ continuous_on s f" by (auto intro!: differentiable_imp_continuous_on differentiableI simp: differentiable_on_def) text \Results about neighborhoods filter.\ lemma eventually_nhds_metric_le: "eventually P (nhds a) = (\d>0. \x. dist x a \ d \ P x)" unfolding eventually_nhds_metric by (safe, rule_tac x="d / 2" in exI, auto) lemma le_nhds: "F \ nhds a \ (\S. open S \ a \ S \ eventually (\x. x \ S) F)" unfolding le_filter_def eventually_nhds by (fast elim: eventually_mono) lemma le_nhds_metric: "F \ nhds a \ (\e>0. eventually (\x. dist x a < e) F)" unfolding le_filter_def eventually_nhds_metric by (fast elim: eventually_mono) lemma le_nhds_metric_le: "F \ nhds a \ (\e>0. eventually (\x. dist x a \ e) F)" unfolding le_filter_def eventually_nhds_metric_le by (fast elim: eventually_mono) text \Several results are easier using a "multiplied-out" variant. (I got this idea from Dieudonne's proof of the chain rule).\ lemma has_derivative_within_alt: "(f has_derivative f') (at x within s) \ bounded_linear f' \ (\e>0. \d>0. \y\s. norm(y - x) < d \ norm (f y - f x - f' (y - x)) \ e * norm (y - x))" unfolding has_derivative_within filterlim_def le_nhds_metric_le eventually_filtermap eventually_at dist_norm diff_diff_eq by (force simp add: linear_0 bounded_linear.linear pos_divide_le_eq) lemma has_derivative_within_alt2: "(f has_derivative f') (at x within s) \ bounded_linear f' \ (\e>0. eventually (\y. norm (f y - f x - f' (y - x)) \ e * norm (y - x)) (at x within s))" unfolding has_derivative_within filterlim_def le_nhds_metric_le eventually_filtermap eventually_at dist_norm diff_diff_eq by (force simp add: linear_0 bounded_linear.linear pos_divide_le_eq) lemma has_derivative_at_alt: "(f has_derivative f') (at x) \ bounded_linear f' \ (\e>0. \d>0. \y. norm(y - x) < d \ norm (f y - f x - f'(y - x)) \ e * norm (y - x))" using has_derivative_within_alt[where s=UNIV] by simp subsection \The chain rule\ proposition diff_chain_within[derivative_intros]: assumes "(f has_derivative f') (at x within s)" and "(g has_derivative g') (at (f x) within (f ` s))" shows "((g \ f) has_derivative (g' \ f'))(at x within s)" using has_derivative_in_compose[OF assms] by (simp add: comp_def) lemma diff_chain_at[derivative_intros]: "(f has_derivative f') (at x) \ (g has_derivative g') (at (f x)) \ ((g \ f) has_derivative (g' \ f')) (at x)" using has_derivative_compose[of f f' x UNIV g g'] by (simp add: comp_def) lemma has_vector_derivative_within_open: "a \ S \ open S \ (f has_vector_derivative f') (at a within S) \ (f has_vector_derivative f') (at a)" by (simp only: at_within_interior interior_open) lemma field_vector_diff_chain_within: assumes Df: "(f has_vector_derivative f') (at x within S)" and Dg: "(g has_field_derivative g') (at (f x) within f ` S)" shows "((g \ f) has_vector_derivative (f' * g')) (at x within S)" using diff_chain_within[OF Df[unfolded has_vector_derivative_def] Dg [unfolded has_field_derivative_def]] by (auto simp: o_def mult.commute has_vector_derivative_def) lemma vector_derivative_diff_chain_within: assumes Df: "(f has_vector_derivative f') (at x within S)" and Dg: "(g has_derivative g') (at (f x) within f`S)" shows "((g \ f) has_vector_derivative (g' f')) (at x within S)" using diff_chain_within[OF Df[unfolded has_vector_derivative_def] Dg] linear.scaleR[OF has_derivative_linear[OF Dg]] unfolding has_vector_derivative_def o_def by (auto simp: o_def mult.commute has_vector_derivative_def) subsection\<^marker>\tag unimportant\ \Composition rules stated just for differentiability\ lemma differentiable_chain_at: "f differentiable (at x) \ g differentiable (at (f x)) \ (g \ f) differentiable (at x)" unfolding differentiable_def by (meson diff_chain_at) lemma differentiable_chain_within: "f differentiable (at x within S) \ g differentiable (at(f x) within (f ` S)) \ (g \ f) differentiable (at x within S)" unfolding differentiable_def by (meson diff_chain_within) subsection \Uniqueness of derivative\ text\<^marker>\tag important\ \ The general result is a bit messy because we need approachability of the limit point from any direction. But OK for nontrivial intervals etc. \ proposition frechet_derivative_unique_within: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes 1: "(f has_derivative f') (at x within S)" and 2: "(f has_derivative f'') (at x within S)" and S: "\i e. \i\Basis; e>0\ \ \d. 0 < \d\ \ \d\ < e \ (x + d *\<^sub>R i) \ S" shows "f' = f''" proof - note as = assms(1,2)[unfolded has_derivative_def] then interpret f': bounded_linear f' by auto from as interpret f'': bounded_linear f'' by auto have "x islimpt S" unfolding islimpt_approachable proof (intro allI impI) fix e :: real assume "e > 0" obtain d where "0 < \d\" and "\d\ < e" and "x + d *\<^sub>R (SOME i. i \ Basis) \ S" using assms(3) SOME_Basis \e>0\ by blast then show "\x'\S. x' \ x \ dist x' x < e" by (rule_tac x="x + d *\<^sub>R (SOME i. i \ Basis)" in bexI) (auto simp: dist_norm SOME_Basis nonzero_Basis) qed then have *: "netlimit (at x within S) = x" by (simp add: Lim_ident_at trivial_limit_within) show ?thesis proof (rule linear_eq_stdbasis) show "linear f'" "linear f''" unfolding linear_conv_bounded_linear using as by auto next fix i :: 'a assume i: "i \ Basis" define e where "e = norm (f' i - f'' i)" show "f' i = f'' i" proof (rule ccontr) assume "f' i \ f'' i" then have "e > 0" unfolding e_def by auto obtain d where d: "0 < d" "(\y. y\S \ 0 < dist y x \ dist y x < d \ dist ((f y - f x - f' (y - x)) /\<^sub>R norm (y - x) - (f y - f x - f'' (y - x)) /\<^sub>R norm (y - x)) (0 - 0) < e)" using tendsto_diff [OF as(1,2)[THEN conjunct2]] unfolding * Lim_within using \e>0\ by blast obtain c where c: "0 < \c\" "\c\ < d \ x + c *\<^sub>R i \ S" using assms(3) i d(1) by blast have *: "norm (- ((1 / \c\) *\<^sub>R f' (c *\<^sub>R i)) + (1 / \c\) *\<^sub>R f'' (c *\<^sub>R i)) = norm ((1 / \c\) *\<^sub>R (- (f' (c *\<^sub>R i)) + f'' (c *\<^sub>R i)))" unfolding scaleR_right_distrib by auto also have "\ = norm ((1 / \c\) *\<^sub>R (c *\<^sub>R (- (f' i) + f'' i)))" unfolding f'.scaleR f''.scaleR unfolding scaleR_right_distrib scaleR_minus_right by auto also have "\ = e" unfolding e_def using c(1) using norm_minus_cancel[of "f' i - f'' i"] by auto finally show False using c using d(2)[of "x + c *\<^sub>R i"] unfolding dist_norm unfolding f'.scaleR f''.scaleR f'.add f''.add f'.diff f''.diff scaleR_scaleR scaleR_right_diff_distrib scaleR_right_distrib using i by (auto simp: inverse_eq_divide) qed qed qed proposition frechet_derivative_unique_within_closed_interval: fixes f::"'a::euclidean_space \ 'b::real_normed_vector" assumes ab: "\i. i\Basis \ a\i < b\i" and x: "x \ cbox a b" and "(f has_derivative f' ) (at x within cbox a b)" and "(f has_derivative f'') (at x within cbox a b)" shows "f' = f''" proof (rule frechet_derivative_unique_within) fix e :: real fix i :: 'a assume "e > 0" and i: "i \ Basis" then show "\d. 0 < \d\ \ \d\ < e \ x + d *\<^sub>R i \ cbox a b" proof (cases "x\i = a\i") case True with ab[of i] \e>0\ x i show ?thesis by (rule_tac x="(min (b\i - a\i) e) / 2" in exI) (auto simp add: mem_box field_simps inner_simps inner_Basis) next case False moreover have "a \ i < x \ i" using False i mem_box(2) x by force moreover { have "a \ i * 2 + min (x \ i - a \ i) e \ a\i *2 + x\i - a\i" by auto also have "\ = a\i + x\i" by auto also have "\ \ 2 * (x\i)" using \a \ i < x \ i\ by auto finally have "a \ i * 2 + min (x \ i - a \ i) e \ x \ i * 2" by auto } moreover have "min (x \ i - a \ i) e \ 0" by (simp add: \0 < e\ \a \ i < x \ i\ less_eq_real_def) then have "x \ i * 2 \ b \ i * 2 + min (x \ i - a \ i) e" using i mem_box(2) x by force ultimately show ?thesis using ab[of i] \e>0\ x i by (rule_tac x="- (min (x\i - a\i) e) / 2" in exI) (auto simp add: mem_box field_simps inner_simps inner_Basis) qed qed (use assms in auto) lemma frechet_derivative_unique_within_open_interval: fixes f::"'a::euclidean_space \ 'b::real_normed_vector" assumes x: "x \ box a b" and f: "(f has_derivative f' ) (at x within box a b)" "(f has_derivative f'') (at x within box a b)" shows "f' = f''" proof - have "at x within box a b = at x" by (metis x at_within_interior interior_open open_box) with f show "f' = f''" by (simp add: has_derivative_unique) qed lemma frechet_derivative_at: "(f has_derivative f') (at x) \ f' = frechet_derivative f (at x)" using differentiable_def frechet_derivative_works has_derivative_unique by blast lemma frechet_derivative_compose: "frechet_derivative (f o g) (at x) = frechet_derivative (f) (at (g x)) o frechet_derivative g (at x)" if "g differentiable at x" "f differentiable at (g x)" by (metis diff_chain_at frechet_derivative_at frechet_derivative_works that) lemma frechet_derivative_within_cbox: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "\i. i\Basis \ a\i < b\i" and "x \ cbox a b" and "(f has_derivative f') (at x within cbox a b)" shows "frechet_derivative f (at x within cbox a b) = f'" using assms by (metis Derivative.differentiableI frechet_derivative_unique_within_closed_interval frechet_derivative_works) lemma frechet_derivative_transform_within_open: "frechet_derivative f (at x) = frechet_derivative g (at x)" if "f differentiable at x" "open X" "x \ X" "\x. x \ X \ f x = g x" by (meson frechet_derivative_at frechet_derivative_works has_derivative_transform_within_open that) subsection \Derivatives of local minima and maxima are zero\ lemma has_derivative_local_min: fixes f :: "'a::real_normed_vector \ real" assumes deriv: "(f has_derivative f') (at x)" assumes min: "eventually (\y. f x \ f y) (at x)" shows "f' = (\h. 0)" proof fix h :: 'a interpret f': bounded_linear f' using deriv by (rule has_derivative_bounded_linear) show "f' h = 0" proof (cases "h = 0") case False from min obtain d where d1: "0 < d" and d2: "\y\ball x d. f x \ f y" unfolding eventually_at by (force simp: dist_commute) have "FDERIV (\r. x + r *\<^sub>R h) 0 :> (\r. r *\<^sub>R h)" by (intro derivative_eq_intros) auto then have "FDERIV (\r. f (x + r *\<^sub>R h)) 0 :> (\k. f' (k *\<^sub>R h))" by (rule has_derivative_compose, simp add: deriv) then have "DERIV (\r. f (x + r *\<^sub>R h)) 0 :> f' h" unfolding has_field_derivative_def by (simp add: f'.scaleR mult_commute_abs) moreover have "0 < d / norm h" using d1 and \h \ 0\ by simp moreover have "\y. \0 - y\ < d / norm h \ f (x + 0 *\<^sub>R h) \ f (x + y *\<^sub>R h)" using \h \ 0\ by (auto simp add: d2 dist_norm pos_less_divide_eq) ultimately show "f' h = 0" by (rule DERIV_local_min) qed simp qed lemma has_derivative_local_max: fixes f :: "'a::real_normed_vector \ real" assumes "(f has_derivative f') (at x)" assumes "eventually (\y. f y \ f x) (at x)" shows "f' = (\h. 0)" using has_derivative_local_min [of "\x. - f x" "\h. - f' h" "x"] using assms unfolding fun_eq_iff by simp lemma differential_zero_maxmin: fixes f::"'a::real_normed_vector \ real" assumes "x \ S" and "open S" and deriv: "(f has_derivative f') (at x)" and mono: "(\y\S. f y \ f x) \ (\y\S. f x \ f y)" shows "f' = (\v. 0)" using mono proof assume "\y\S. f y \ f x" with \x \ S\ and \open S\ have "eventually (\y. f y \ f x) (at x)" unfolding eventually_at_topological by auto with deriv show ?thesis by (rule has_derivative_local_max) next assume "\y\S. f x \ f y" with \x \ S\ and \open S\ have "eventually (\y. f x \ f y) (at x)" unfolding eventually_at_topological by auto with deriv show ?thesis by (rule has_derivative_local_min) qed lemma differential_zero_maxmin_component: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes k: "k \ Basis" and ball: "0 < e" "(\y \ ball x e. (f y)\k \ (f x)\k) \ (\y\ball x e. (f x)\k \ (f y)\k)" and diff: "f differentiable (at x)" shows "(\j\Basis. (frechet_derivative f (at x) j \ k) *\<^sub>R j) = (0::'a)" (is "?D k = 0") proof - let ?f' = "frechet_derivative f (at x)" have "x \ ball x e" using \0 < e\ by simp moreover have "open (ball x e)" by simp moreover have "((\x. f x \ k) has_derivative (\h. ?f' h \ k)) (at x)" using bounded_linear_inner_left diff[unfolded frechet_derivative_works] by (rule bounded_linear.has_derivative) ultimately have "(\h. frechet_derivative f (at x) h \ k) = (\v. 0)" using ball(2) by (rule differential_zero_maxmin) then show ?thesis unfolding fun_eq_iff by simp qed subsection \One-dimensional mean value theorem\ lemma mvt_simple: fixes f :: "real \ real" assumes "a < b" and derf: "\x. \a \ x; x \ b\ \ (f has_derivative f' x) (at x within {a..b})" shows "\x\{a<.. real" assumes "a \ b" and derf: "\x. \a \ x; x \ b\ \ (f has_derivative f' x) (at x within {a..b})" shows "\x\{a..b}. f b - f a = f' x (b - a)" proof (cases "a = b") interpret bounded_linear "f' b" using assms(2) assms(1) by auto case True then show ?thesis by force next case False then show ?thesis using mvt_simple[OF _ derf] by (metis \a \ b\ atLeastAtMost_iff dual_order.order_iff_strict greaterThanLessThan_iff) qed text \A nice generalization (see Havin's proof of 5.19 from Rudin's book).\ lemma mvt_general: fixes f :: "real \ 'a::real_inner" assumes "a < b" and contf: "continuous_on {a..b} f" and derf: "\x. \a < x; x < b\ \ (f has_derivative f' x) (at x)" shows "\x\{a<.. norm (f' x (b - a))" proof - have "\x\{a<.. f b - (f b - f a) \ f a = (f b - f a) \ f' x (b - a)" apply (rule mvt [OF \a < b\, where f = "\x. (f b - f a) \ f x"]) apply (intro continuous_intros contf) using derf apply (auto intro: has_derivative_inner_right) done then obtain x where x: "x \ {a<.. f b - (f b - f a) \ f a = (f b - f a) \ f' x (b - a)" .. show ?thesis proof (cases "f a = f b") case False have "norm (f b - f a) * norm (f b - f a) = (norm (f b - f a))\<^sup>2" by (simp add: power2_eq_square) also have "\ = (f b - f a) \ (f b - f a)" unfolding power2_norm_eq_inner .. also have "\ = (f b - f a) \ f' x (b - a)" using x(2) by (simp only: inner_diff_right) also have "\ \ norm (f b - f a) * norm (f' x (b - a))" by (rule norm_cauchy_schwarz) finally show ?thesis using False x(1) by (auto simp add: mult_left_cancel) next case True then show ?thesis using \a < b\ by (rule_tac x="(a + b) /2" in bexI) auto qed qed subsection \More general bound theorems\ proposition differentiable_bound_general: fixes f :: "real \ 'a::real_normed_vector" assumes "a < b" and f_cont: "continuous_on {a..b} f" and phi_cont: "continuous_on {a..b} \" and f': "\x. a < x \ x < b \ (f has_vector_derivative f' x) (at x)" and phi': "\x. a < x \ x < b \ (\ has_vector_derivative \' x) (at x)" and bnd: "\x. a < x \ x < b \ norm (f' x) \ \' x" shows "norm (f b - f a) \ \ b - \ a" proof - { fix x assume x: "a < x" "x < b" have "0 \ norm (f' x)" by simp also have "\ \ \' x" using x by (auto intro!: bnd) finally have "0 \ \' x" . } note phi'_nonneg = this note f_tendsto = assms(2)[simplified continuous_on_def, rule_format] note phi_tendsto = assms(3)[simplified continuous_on_def, rule_format] { fix e::real assume "e > 0" define e2 where "e2 = e / 2" with \e > 0\ have "e2 > 0" by simp let ?le = "\x1. norm (f x1 - f a) \ \ x1 - \ a + e * (x1 - a) + e" define A where "A = {x2. a \ x2 \ x2 \ b \ (\x1\{a ..< x2}. ?le x1)}" have A_subset: "A \ {a..b}" by (auto simp: A_def) { fix x2 assume a: "a \ x2" "x2 \ b" and le: "\x1\{a..e > 0\ proof cases assume "x2 \ a" with a have "a < x2" by simp have "at x2 within {a <.. bot" using \a < x2\ by (auto simp: trivial_limit_within islimpt_in_closure) moreover have "((\x1. (\ x1 - \ a) + e * (x1 - a) + e) \ (\ x2 - \ a) + e * (x2 - a) + e) (at x2 within {a <..x1. norm (f x1 - f a)) \ norm (f x2 - f a)) (at x2 within {a <..x. x > a) (at x2 within {a <.. A" using assms by (auto simp: A_def) hence [simp]: "A \ {}" by auto have A_ivl: "\x1 x2. x2 \ A \ x1 \ {a ..x2} \ x1 \ A" by (simp add: A_def) have [simp]: "bdd_above A" by (auto simp: A_def) define y where "y = Sup A" have "y \ b" unfolding y_def by (simp add: cSup_le_iff) (simp add: A_def) have leI: "\x x1. a \ x1 \ x \ A \ x1 < x \ ?le x1" by (auto simp: A_def intro!: le_cont) have y_all_le: "\x1\{a.. y" by (metis \a \ A\ \bdd_above A\ cSup_upper y_def) have "y \ A" using y_all_le \a \ y\ \y \ b\ by (auto simp: A_def) hence "A = {a .. y}" using A_subset by (auto simp: subset_iff y_def cSup_upper intro: A_ivl) from le_cont[OF \a \ y\ \y \ b\ y_all_le] have le_y: "?le y" . have "y = b" proof (cases "a = y") case True with \a < b\ have "y < b" by simp with \a = y\ f_cont phi_cont \e2 > 0\ have 1: "\\<^sub>F x in at y within {y..b}. dist (f x) (f y) < e2" and 2: "\\<^sub>F x in at y within {y..b}. dist (\ x) (\ y) < e2" by (auto simp: continuous_on_def tendsto_iff) have 3: "eventually (\x. y < x) (at y within {y..b})" by (auto simp: eventually_at_filter) have 4: "eventually (\x::real. x < b) (at y within {y..b})" using _ \y < b\ by (rule order_tendstoD) (auto intro!: tendsto_eq_intros) from 1 2 3 4 have eventually_le: "eventually (\x. ?le x) (at y within {y .. b})" proof eventually_elim case (elim x1) have "norm (f x1 - f a) = norm (f x1 - f y)" by (simp add: \a = y\) also have "norm (f x1 - f y) \ e2" using elim \a = y\ by (auto simp : dist_norm intro!: less_imp_le) also have "\ \ e2 + (\ x1 - \ a + e2 + e * (x1 - a))" using \0 < e\ elim by (intro add_increasing2[OF add_nonneg_nonneg order.refl]) (auto simp: \a = y\ dist_norm intro!: mult_nonneg_nonneg) also have "\ = \ x1 - \ a + e * (x1 - a) + e" by (simp add: e2_def) finally show "?le x1" . qed from this[unfolded eventually_at_topological] \?le y\ obtain S where S: "open S" "y \ S" "\x. x\S \ x \ {y..b} \ ?le x" by metis from \open S\ obtain d where d: "\x. dist x y < d \ x \ S" "d > 0" by (force simp: dist_commute open_dist ball_def dest!: bspec[OF _ \y \ S\]) define d' where "d' = min b (y + (d/2))" have "d' \ A" unfolding A_def proof safe show "a \ d'" using \a = y\ \0 < d\ \y < b\ by (simp add: d'_def) show "d' \ b" by (simp add: d'_def) fix x1 assume "x1 \ {a.. S" "x1 \ {y..b}" by (auto simp: \a = y\ d'_def dist_real_def intro!: d ) thus "?le x1" by (rule S) qed hence "d' \ y" unfolding y_def by (rule cSup_upper) simp then show "y = b" using \d > 0\ \y < b\ by (simp add: d'_def) next case False with \a \ y\ have "a < y" by simp show "y = b" proof (rule ccontr) assume "y \ b" hence "y < b" using \y \ b\ by simp let ?F = "at y within {y.. has_vector_derivative \' y) ?F" using \a < y\ \y < b\ by (auto simp add: at_within_open[of _ "{a<..\<^sub>F x1 in ?F. norm (f x1 - f y - (x1 - y) *\<^sub>R f' y) \ e2 * \x1 - y\" "\\<^sub>F x1 in ?F. norm (\ x1 - \ y - (x1 - y) *\<^sub>R \' y) \ e2 * \x1 - y\" using \e2 > 0\ by (auto simp: has_derivative_within_alt2 has_vector_derivative_def) moreover have "\\<^sub>F x1 in ?F. y \ x1" "\\<^sub>F x1 in ?F. x1 < b" by (auto simp: eventually_at_filter) ultimately have "\\<^sub>F x1 in ?F. norm (f x1 - f y) \ (\ x1 - \ y) + e * \x1 - y\" (is "\\<^sub>F x1 in ?F. ?le' x1") proof eventually_elim case (elim x1) from norm_triangle_ineq2[THEN order_trans, OF elim(1)] have "norm (f x1 - f y) \ norm (f' y) * \x1 - y\ + e2 * \x1 - y\" by (simp add: ac_simps) also have "norm (f' y) \ \' y" using bnd \a < y\ \y < b\ by simp also have "\' y * \x1 - y\ \ \ x1 - \ y + e2 * \x1 - y\" using elim by (simp add: ac_simps) finally have "norm (f x1 - f y) \ \ x1 - \ y + e2 * \x1 - y\ + e2 * \x1 - y\" by (auto simp: mult_right_mono) thus ?case by (simp add: e2_def) qed moreover have "?le' y" by simp ultimately obtain S where S: "open S" "y \ S" "\x. x\S \ x \ {y.. ?le' x" unfolding eventually_at_topological by metis from \open S\ obtain d where d: "\x. dist x y < d \ x \ S" "d > 0" by (force simp: dist_commute open_dist ball_def dest!: bspec[OF _ \y \ S\]) define d' where "d' = min ((y + b)/2) (y + (d/2))" have "d' \ A" unfolding A_def proof safe show "a \ d'" using \a < y\ \0 < d\ \y < b\ by (simp add: d'_def) show "d' \ b" using \y < b\ by (simp add: d'_def min_def) fix x1 assume x1: "x1 \ {a..y \ A\ local.leI x1 by auto next case False hence x1': "x1 \ S" "x1 \ {y.. norm (f x1 - f y) + norm (f y - f a)" by (rule order_trans[OF _ norm_triangle_ineq]) simp also note S(3)[OF x1'] also note le_y finally show "?le x1" using False by (auto simp: algebra_simps) qed qed hence "d' \ y" unfolding y_def by (rule cSup_upper) simp thus False using \d > 0\ \y < b\ by (simp add: d'_def min_def split: if_split_asm) qed qed with le_y have "norm (f b - f a) \ \ b - \ a + e * (b - a + 1)" by (simp add: algebra_simps) } note * = this show ?thesis proof (rule field_le_epsilon) fix e::real assume "e > 0" then show "norm (f b - f a) \ \ b - \ a + e" using *[of "e / (b - a + 1)"] \a < b\ by simp qed qed lemma differentiable_bound: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "convex S" and derf: "\x. x\S \ (f has_derivative f' x) (at x within S)" and B: "\x. x \ S \ onorm (f' x) \ B" and x: "x \ S" and y: "y \ S" shows "norm (f x - f y) \ B * norm (x - y)" proof - let ?p = "\u. x + u *\<^sub>R (y - x)" let ?\ = "\h. h * B * norm (x - y)" have *: "x + u *\<^sub>R (y - x) \ S" if "u \ {0..1}" for u proof - have "u *\<^sub>R y = u *\<^sub>R (y - x) + u *\<^sub>R x" by (simp add: scale_right_diff_distrib) then show "x + u *\<^sub>R (y - x) \ S" using that \convex S\ x y by (simp add: convex_alt) (metis pth_b(2) pth_c(1) scaleR_collapse) qed have "\z. z \ (\u. x + u *\<^sub>R (y - x)) ` {0..1} \ (f has_derivative f' z) (at z within (\u. x + u *\<^sub>R (y - x)) ` {0..1})" by (auto intro: * has_derivative_subset [OF derf]) then have "continuous_on (?p ` {0..1}) f" unfolding continuous_on_eq_continuous_within by (meson has_derivative_continuous) with * have 1: "continuous_on {0 .. 1} (f \ ?p)" by (intro continuous_intros)+ { fix u::real assume u: "u \{0 <..< 1}" let ?u = "?p u" interpret linear "(f' ?u)" using u by (auto intro!: has_derivative_linear derf *) have "(f \ ?p has_derivative (f' ?u) \ (\u. 0 + u *\<^sub>R (y - x))) (at u within box 0 1)" by (intro derivative_intros has_derivative_subset [OF derf]) (use u * in auto) hence "((f \ ?p) has_vector_derivative f' ?u (y - x)) (at u)" by (simp add: at_within_open[OF u open_greaterThanLessThan] scaleR has_vector_derivative_def o_def) } note 2 = this have 3: "continuous_on {0..1} ?\" by (rule continuous_intros)+ have 4: "(?\ has_vector_derivative B * norm (x - y)) (at u)" for u by (auto simp: has_vector_derivative_def intro!: derivative_eq_intros) { fix u::real assume u: "u \{0 <..< 1}" let ?u = "?p u" interpret bounded_linear "(f' ?u)" using u by (auto intro!: has_derivative_bounded_linear derf *) have "norm (f' ?u (y - x)) \ onorm (f' ?u) * norm (y - x)" by (rule onorm) (rule bounded_linear) also have "onorm (f' ?u) \ B" using u by (auto intro!: assms(3)[rule_format] *) finally have "norm ((f' ?u) (y - x)) \ B * norm (x - y)" by (simp add: mult_right_mono norm_minus_commute) } note 5 = this have "norm (f x - f y) = norm ((f \ (\u. x + u *\<^sub>R (y - x))) 1 - (f \ (\u. x + u *\<^sub>R (y - x))) 0)" by (auto simp add: norm_minus_commute) also from differentiable_bound_general[OF zero_less_one 1, OF 3 2 4 5] have "norm ((f \ ?p) 1 - (f \ ?p) 0) \ B * norm (x - y)" by simp finally show ?thesis . qed lemma field_differentiable_bound: fixes S :: "'a::real_normed_field set" assumes cvs: "convex S" and df: "\z. z \ S \ (f has_field_derivative f' z) (at z within S)" and dn: "\z. z \ S \ norm (f' z) \ B" and "x \ S" "y \ S" shows "norm(f x - f y) \ B * norm(x - y)" apply (rule differentiable_bound [OF cvs]) apply (erule df [unfolded has_field_derivative_def]) apply (rule onorm_le, simp_all add: norm_mult mult_right_mono assms) done lemma differentiable_bound_segment: fixes f::"'a::real_normed_vector \ 'b::real_normed_vector" assumes "\t. t \ {0..1} \ x0 + t *\<^sub>R a \ G" assumes f': "\x. x \ G \ (f has_derivative f' x) (at x within G)" assumes B: "\x. x \ {0..1} \ onorm (f' (x0 + x *\<^sub>R a)) \ B" shows "norm (f (x0 + a) - f x0) \ norm a * B" proof - let ?G = "(\x. x0 + x *\<^sub>R a) ` {0..1}" have "?G = (+) x0 ` (\x. x *\<^sub>R a) ` {0..1}" by auto also have "convex \" by (intro convex_translation convex_scaled convex_real_interval) finally have "convex ?G" . moreover have "?G \ G" "x0 \ ?G" "x0 + a \ ?G" using assms by (auto intro: image_eqI[where x=1]) ultimately show ?thesis using has_derivative_subset[OF f' \?G \ G\] B differentiable_bound[of "(\x. x0 + x *\<^sub>R a) ` {0..1}" f f' B "x0 + a" x0] by (force simp: ac_simps) qed lemma differentiable_bound_linearization: fixes f::"'a::real_normed_vector \ 'b::real_normed_vector" assumes S: "\t. t \ {0..1} \ a + t *\<^sub>R (b - a) \ S" assumes f'[derivative_intros]: "\x. x \ S \ (f has_derivative f' x) (at x within S)" assumes B: "\x. x \ S \ onorm (f' x - f' x0) \ B" assumes "x0 \ S" shows "norm (f b - f a - f' x0 (b - a)) \ norm (b - a) * B" proof - define g where [abs_def]: "g x = f x - f' x0 x" for x have g: "\x. x \ S \ (g has_derivative (\i. f' x i - f' x0 i)) (at x within S)" unfolding g_def using assms by (auto intro!: derivative_eq_intros bounded_linear.has_derivative[OF has_derivative_bounded_linear, OF f']) from B have "\x\{0..1}. onorm (\i. f' (a + x *\<^sub>R (b - a)) i - f' x0 i) \ B" using assms by (auto simp: fun_diff_def) with differentiable_bound_segment[OF S g] \x0 \ S\ show ?thesis by (simp add: g_def field_simps linear_diff[OF has_derivative_linear[OF f']]) qed lemma vector_differentiable_bound_linearization: fixes f::"real \ 'b::real_normed_vector" assumes f': "\x. x \ S \ (f has_vector_derivative f' x) (at x within S)" assumes "closed_segment a b \ S" assumes B: "\x. x \ S \ norm (f' x - f' x0) \ B" assumes "x0 \ S" shows "norm (f b - f a - (b - a) *\<^sub>R f' x0) \ norm (b - a) * B" using assms by (intro differentiable_bound_linearization[of a b S f "\x h. h *\<^sub>R f' x" x0 B]) (force simp: closed_segment_real_eq has_vector_derivative_def scaleR_diff_right[symmetric] mult.commute[of B] intro!: onorm_le mult_left_mono)+ text \In particular.\ lemma has_derivative_zero_constant: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "convex s" and "\x. x \ s \ (f has_derivative (\h. 0)) (at x within s)" shows "\c. \x\s. f x = c" proof - { fix x y assume "x \ s" "y \ s" then have "norm (f x - f y) \ 0 * norm (x - y)" using assms by (intro differentiable_bound[of s]) (auto simp: onorm_zero) then have "f x = f y" by simp } then show ?thesis by metis qed lemma has_field_derivative_zero_constant: assumes "convex s" "\x. x \ s \ (f has_field_derivative 0) (at x within s)" shows "\c. \x\s. f (x) = (c :: 'a :: real_normed_field)" proof (rule has_derivative_zero_constant) have A: "(*) 0 = (\_. 0 :: 'a)" by (intro ext) simp fix x assume "x \ s" thus "(f has_derivative (\h. 0)) (at x within s)" using assms(2)[of x] by (simp add: has_field_derivative_def A) qed fact lemma has_vector_derivative_zero_constant: assumes "convex s" assumes "\x. x \ s \ (f has_vector_derivative 0) (at x within s)" obtains c where "\x. x \ s \ f x = c" using has_derivative_zero_constant[of s f] assms by (auto simp: has_vector_derivative_def) lemma has_derivative_zero_unique: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "convex s" and "\x. x \ s \ (f has_derivative (\h. 0)) (at x within s)" and "x \ s" "y \ s" shows "f x = f y" using has_derivative_zero_constant[OF assms(1,2)] assms(3-) by force lemma has_derivative_zero_unique_connected: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "open s" "connected s" assumes f: "\x. x \ s \ (f has_derivative (\x. 0)) (at x)" assumes "x \ s" "y \ s" shows "f x = f y" proof (rule connected_local_const[where f=f, OF \connected s\ \x\s\ \y\s\]) show "\a\s. eventually (\b. f a = f b) (at a within s)" proof fix a assume "a \ s" with \open s\ obtain e where "0 < e" "ball a e \ s" by (rule openE) then have "\c. \x\ball a e. f x = c" by (intro has_derivative_zero_constant) (auto simp: at_within_open[OF _ open_ball] f) with \0 have "\x\ball a e. f a = f x" by auto then show "eventually (\b. f a = f b) (at a within s)" using \0 unfolding eventually_at_topological by (intro exI[of _ "ball a e"]) auto qed qed subsection \Differentiability of inverse function (most basic form)\ lemma has_derivative_inverse_basic: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes derf: "(f has_derivative f') (at (g y))" and ling': "bounded_linear g'" and "g' \ f' = id" and contg: "continuous (at y) g" and "open T" and "y \ T" and fg: "\z. z \ T \ f (g z) = z" shows "(g has_derivative g') (at y)" proof - interpret f': bounded_linear f' using assms unfolding has_derivative_def by auto interpret g': bounded_linear g' using assms by auto obtain C where C: "0 < C" "\x. norm (g' x) \ norm x * C" using bounded_linear.pos_bounded[OF assms(2)] by blast have lem1: "\e>0. \d>0. \z. norm (z - y) < d \ norm (g z - g y - g'(z - y)) \ e * norm (g z - g y)" proof (intro allI impI) fix e :: real assume "e > 0" with C(1) have *: "e / C > 0" by auto obtain d0 where "0 < d0" and d0: "\u. norm (u - g y) < d0 \ norm (f u - f (g y) - f' (u - g y)) \ e / C * norm (u - g y)" using derf * unfolding has_derivative_at_alt by blast obtain d1 where "0 < d1" and d1: "\x. \0 < dist x y; dist x y < d1\ \ dist (g x) (g y) < d0" using contg \0 < d0\ unfolding continuous_at Lim_at by blast obtain d2 where "0 < d2" and d2: "\u. dist u y < d2 \ u \ T" using \open T\ \y \ T\ unfolding open_dist by blast obtain d where d: "0 < d" "d < d1" "d < d2" using field_lbound_gt_zero[OF \0 < d1\ \0 < d2\] by blast show "\d>0. \z. norm (z - y) < d \ norm (g z - g y - g' (z - y)) \ e * norm (g z - g y)" proof (intro exI allI impI conjI) fix z assume as: "norm (z - y) < d" then have "z \ T" using d2 d unfolding dist_norm by auto have "norm (g z - g y - g' (z - y)) \ norm (g' (f (g z) - y - f' (g z - g y)))" unfolding g'.diff f'.diff unfolding assms(3)[unfolded o_def id_def, THEN fun_cong] fg[OF \z\T\] by (simp add: norm_minus_commute) also have "\ \ norm (f (g z) - y - f' (g z - g y)) * C" by (rule C(2)) also have "\ \ (e / C) * norm (g z - g y) * C" proof - have "norm (g z - g y) < d0" by (metis as cancel_comm_monoid_add_class.diff_cancel d(2) \0 < d0\ d1 diff_gt_0_iff_gt diff_strict_mono dist_norm dist_self zero_less_dist_iff) then show ?thesis by (metis C(1) \y \ T\ d0 fg mult_le_cancel_iff1) qed also have "\ \ e * norm (g z - g y)" using C by (auto simp add: field_simps) finally show "norm (g z - g y - g' (z - y)) \ e * norm (g z - g y)" by simp qed (use d in auto) qed have *: "(0::real) < 1 / 2" by auto obtain d where "0 < d" and d: "\z. norm (z - y) < d \ norm (g z - g y - g' (z - y)) \ 1/2 * norm (g z - g y)" using lem1 * by blast define B where "B = C * 2" have "B > 0" unfolding B_def using C by auto have lem2: "norm (g z - g y) \ B * norm (z - y)" if z: "norm(z - y) < d" for z proof - have "norm (g z - g y) \ norm(g' (z - y)) + norm ((g z - g y) - g'(z - y))" by (rule norm_triangle_sub) also have "\ \ norm (g' (z - y)) + 1 / 2 * norm (g z - g y)" by (rule add_left_mono) (use d z in auto) also have "\ \ norm (z - y) * C + 1 / 2 * norm (g z - g y)" by (rule add_right_mono) (use C in auto) finally show "norm (g z - g y) \ B * norm (z - y)" unfolding B_def by (auto simp add: field_simps) qed show ?thesis unfolding has_derivative_at_alt proof (intro conjI assms allI impI) fix e :: real assume "e > 0" then have *: "e / B > 0" by (metis \B > 0\ divide_pos_pos) obtain d' where "0 < d'" and d': "\z. norm (z - y) < d' \ norm (g z - g y - g' (z - y)) \ e / B * norm (g z - g y)" using lem1 * by blast obtain k where k: "0 < k" "k < d" "k < d'" using field_lbound_gt_zero[OF \0 < d\ \0 < d'\] by blast show "\d>0. \ya. norm (ya - y) < d \ norm (g ya - g y - g' (ya - y)) \ e * norm (ya - y)" proof (intro exI allI impI conjI) fix z assume as: "norm (z - y) < k" then have "norm (g z - g y - g' (z - y)) \ e / B * norm(g z - g y)" using d' k by auto also have "\ \ e * norm (z - y)" unfolding times_divide_eq_left pos_divide_le_eq[OF \B>0\] using lem2[of z] k as \e > 0\ by (auto simp add: field_simps) finally show "norm (g z - g y - g' (z - y)) \ e * norm (z - y)" by simp qed (use k in auto) qed qed text\<^marker>\tag unimportant\\Inverse function theorem for complex derivatives\ lemma has_field_derivative_inverse_basic: shows "DERIV f (g y) :> f' \ f' \ 0 \ continuous (at y) g \ open t \ y \ t \ (\z. z \ t \ f (g z) = z) \ DERIV g y :> inverse (f')" unfolding has_field_derivative_def apply (rule has_derivative_inverse_basic) apply (auto simp: bounded_linear_mult_right) done text \Simply rewrite that based on the domain point x.\ lemma has_derivative_inverse_basic_x: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "(f has_derivative f') (at x)" and "bounded_linear g'" and "g' \ f' = id" and "continuous (at (f x)) g" and "g (f x) = x" and "open T" and "f x \ T" and "\y. y \ T \ f (g y) = y" shows "(g has_derivative g') (at (f x))" by (rule has_derivative_inverse_basic) (use assms in auto) text \This is the version in Dieudonne', assuming continuity of f and g.\ lemma has_derivative_inverse_dieudonne: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "open S" and "open (f ` S)" and "continuous_on S f" and "continuous_on (f ` S) g" and "\x. x \ S \ g (f x) = x" and "x \ S" and "(f has_derivative f') (at x)" and "bounded_linear g'" and "g' \ f' = id" shows "(g has_derivative g') (at (f x))" apply (rule has_derivative_inverse_basic_x[OF assms(7-9) _ _ assms(2)]) using assms(3-6) unfolding continuous_on_eq_continuous_at[OF assms(1)] continuous_on_eq_continuous_at[OF assms(2)] apply auto done text \Here's the simplest way of not assuming much about g.\ proposition has_derivative_inverse: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "compact S" and "x \ S" and fx: "f x \ interior (f ` S)" and "continuous_on S f" and gf: "\y. y \ S \ g (f y) = y" and "(f has_derivative f') (at x)" and "bounded_linear g'" and "g' \ f' = id" shows "(g has_derivative g') (at (f x))" proof - have *: "\y. y \ interior (f ` S) \ f (g y) = y" by (metis gf image_iff interior_subset subsetCE) show ?thesis apply (rule has_derivative_inverse_basic_x[OF assms(6-8), where T = "interior (f ` S)"]) apply (rule continuous_on_interior[OF _ fx]) apply (rule continuous_on_inv) apply (simp_all add: assms *) done qed text \Invertible derivative continuous at a point implies local injectivity. It's only for this we need continuity of the derivative, except of course if we want the fact that the inverse derivative is also continuous. So if we know for some other reason that the inverse function exists, it's OK.\ proposition has_derivative_locally_injective: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "a \ S" and "open S" and bling: "bounded_linear g'" and "g' \ f' a = id" and derf: "\x. x \ S \ (f has_derivative f' x) (at x)" and "\e. e > 0 \ \d>0. \x. dist a x < d \ onorm (\v. f' x v - f' a v) < e" obtains r where "r > 0" "ball a r \ S" "inj_on f (ball a r)" proof - interpret bounded_linear g' using assms by auto note f'g' = assms(4)[unfolded id_def o_def,THEN cong] have "g' (f' a (\Basis)) = (\Basis)" "(\Basis) \ (0::'n)" using f'g' by auto then have *: "0 < onorm g'" unfolding onorm_pos_lt[OF assms(3)] by fastforce define k where "k = 1 / onorm g' / 2" have *: "k > 0" unfolding k_def using * by auto obtain d1 where d1: "0 < d1" "\x. dist a x < d1 \ onorm (\v. f' x v - f' a v) < k" using assms(6) * by blast from \open S\ obtain d2 where "d2 > 0" "ball a d2 \ S" using \a\S\ .. obtain d2 where d2: "0 < d2" "ball a d2 \ S" using \0 < d2\ \ball a d2 \ S\ by blast obtain d where d: "0 < d" "d < d1" "d < d2" using field_lbound_gt_zero[OF d1(1) d2(1)] by blast show ?thesis proof show "0 < d" by (fact d) show "ball a d \ S" using \d < d2\ \ball a d2 \ S\ by auto show "inj_on f (ball a d)" unfolding inj_on_def proof (intro strip) fix x y assume as: "x \ ball a d" "y \ ball a d" "f x = f y" define ph where [abs_def]: "ph w = w - g' (f w - f x)" for w have ph':"ph = g' \ (\w. f' a w - (f w - f x))" unfolding ph_def o_def by (simp add: diff f'g') have "norm (ph x - ph y) \ (1 / 2) * norm (x - y)" proof (rule differentiable_bound[OF convex_ball _ _ as(1-2)]) fix u assume u: "u \ ball a d" then have "u \ S" using d d2 by auto have *: "(\v. v - g' (f' u v)) = g' \ (\w. f' a w - f' u w)" unfolding o_def and diff using f'g' by auto have blin: "bounded_linear (f' a)" using \a \ S\ derf by blast show "(ph has_derivative (\v. v - g' (f' u v))) (at u within ball a d)" unfolding ph' * comp_def by (rule \u \ S\ derivative_eq_intros has_derivative_at_withinI [OF derf] bounded_linear.has_derivative [OF blin] bounded_linear.has_derivative [OF bling] |simp)+ have **: "bounded_linear (\x. f' u x - f' a x)" "bounded_linear (\x. f' a x - f' u x)" using \u \ S\ blin bounded_linear_sub derf by auto then have "onorm (\v. v - g' (f' u v)) \ onorm g' * onorm (\w. f' a w - f' u w)" by (simp add: "*" bounded_linear_axioms onorm_compose) also have "\ \ onorm g' * k" apply (rule mult_left_mono) using d1(2)[of u] using onorm_neg[where f="\x. f' u x - f' a x"] d u onorm_pos_le[OF bling] apply (auto simp: algebra_simps) done also have "\ \ 1 / 2" unfolding k_def by auto finally show "onorm (\v. v - g' (f' u v)) \ 1 / 2" . qed moreover have "norm (ph y - ph x) = norm (y - x)" by (simp add: as(3) ph_def) ultimately show "x = y" unfolding norm_minus_commute by auto qed qed qed subsection \Uniformly convergent sequence of derivatives\ lemma has_derivative_sequence_lipschitz_lemma: fixes f :: "nat \ 'a::real_normed_vector \ 'b::real_normed_vector" assumes "convex S" and derf: "\n x. x \ S \ ((f n) has_derivative (f' n x)) (at x within S)" and nle: "\n x h. \n\N; x \ S\ \ norm (f' n x h - g' x h) \ e * norm h" and "0 \ e" shows "\m\N. \n\N. \x\S. \y\S. norm ((f m x - f n x) - (f m y - f n y)) \ 2 * e * norm (x - y)" proof clarify fix m n x y assume as: "N \ m" "N \ n" "x \ S" "y \ S" show "norm ((f m x - f n x) - (f m y - f n y)) \ 2 * e * norm (x - y)" proof (rule differentiable_bound[where f'="\x h. f' m x h - f' n x h", OF \convex S\ _ _ as(3-4)]) fix x assume "x \ S" show "((\a. f m a - f n a) has_derivative (\h. f' m x h - f' n x h)) (at x within S)" by (rule derivative_intros derf \x\S\)+ show "onorm (\h. f' m x h - f' n x h) \ 2 * e" proof (rule onorm_bound) fix h have "norm (f' m x h - f' n x h) \ norm (f' m x h - g' x h) + norm (f' n x h - g' x h)" using norm_triangle_ineq[of "f' m x h - g' x h" "- f' n x h + g' x h"] by (auto simp add: algebra_simps norm_minus_commute) also have "\ \ e * norm h + e * norm h" using nle[OF \N \ m\ \x \ S\, of h] nle[OF \N \ n\ \x \ S\, of h] by (auto simp add: field_simps) finally show "norm (f' m x h - f' n x h) \ 2 * e * norm h" by auto qed (simp add: \0 \ e\) qed qed lemma has_derivative_sequence_Lipschitz: fixes f :: "nat \ 'a::real_normed_vector \ 'b::real_normed_vector" assumes "convex S" and "\n x. x \ S \ ((f n) has_derivative (f' n x)) (at x within S)" and nle: "\e. e > 0 \ \\<^sub>F n in sequentially. \x\S. \h. norm (f' n x h - g' x h) \ e * norm h" and "e > 0" shows "\N. \m\N. \n\N. \x\S. \y\S. norm ((f m x - f n x) - (f m y - f n y)) \ e * norm (x - y)" proof - have *: "2 * (e/2) = e" using \e > 0\ by auto obtain N where "\n\N. \x\S. \h. norm (f' n x h - g' x h) \ (e/2) * norm h" using nle \e > 0\ unfolding eventually_sequentially by (metis less_divide_eq_numeral1(1) mult_zero_left) then show "\N. \m\N. \n\N. \x\S. \y\S. norm (f m x - f n x - (f m y - f n y)) \ e * norm (x - y)" apply (rule_tac x=N in exI) apply (rule has_derivative_sequence_lipschitz_lemma[where e="e/2", unfolded *]) using assms \e > 0\ apply auto done qed proposition has_derivative_sequence: fixes f :: "nat \ 'a::real_normed_vector \ 'b::banach" assumes "convex S" and derf: "\n x. x \ S \ ((f n) has_derivative (f' n x)) (at x within S)" and nle: "\e. e > 0 \ \\<^sub>F n in sequentially. \x\S. \h. norm (f' n x h - g' x h) \ e * norm h" and "x0 \ S" and lim: "((\n. f n x0) \ l) sequentially" shows "\g. \x\S. (\n. f n x) \ g x \ (g has_derivative g'(x)) (at x within S)" proof - have lem1: "\e. e > 0 \ \N. \m\N. \n\N. \x\S. \y\S. norm ((f m x - f n x) - (f m y - f n y)) \ e * norm (x - y)" using assms(1,2,3) by (rule has_derivative_sequence_Lipschitz) have "\g. \x\S. ((\n. f n x) \ g x) sequentially" proof (intro ballI bchoice) fix x assume "x \ S" show "\y. (\n. f n x) \ y" unfolding convergent_eq_Cauchy proof (cases "x = x0") case True then show "Cauchy (\n. f n x)" using LIMSEQ_imp_Cauchy[OF lim] by auto next case False show "Cauchy (\n. f n x)" unfolding Cauchy_def proof (intro allI impI) fix e :: real assume "e > 0" hence *: "e / 2 > 0" "e / 2 / norm (x - x0) > 0" using False by auto obtain M where M: "\m\M. \n\M. dist (f m x0) (f n x0) < e / 2" using LIMSEQ_imp_Cauchy[OF lim] * unfolding Cauchy_def by blast obtain N where N: "\m\N. \n\N. \u\S. \y\S. norm (f m u - f n u - (f m y - f n y)) \ e / 2 / norm (x - x0) * norm (u - y)" using lem1 *(2) by blast show "\M. \m\M. \n\M. dist (f m x) (f n x) < e" proof (intro exI allI impI) fix m n assume as: "max M N \m" "max M N\n" have "dist (f m x) (f n x) \ norm (f m x0 - f n x0) + norm (f m x - f n x - (f m x0 - f n x0))" unfolding dist_norm by (rule norm_triangle_sub) also have "\ \ norm (f m x0 - f n x0) + e / 2" using N \x\S\ \x0\S\ as False by fastforce also have "\ < e / 2 + e / 2" by (rule add_strict_right_mono) (use as M in \auto simp: dist_norm\) finally show "dist (f m x) (f n x) < e" by auto qed qed qed qed then obtain g where g: "\x\S. (\n. f n x) \ g x" .. have lem2: "\N. \n\N. \x\S. \y\S. norm ((f n x - f n y) - (g x - g y)) \ e * norm (x - y)" if "e > 0" for e proof - obtain N where N: "\m\N. \n\N. \x\S. \y\S. norm (f m x - f n x - (f m y - f n y)) \ e * norm (x - y)" using lem1 \e > 0\ by blast show "\N. \n\N. \x\S. \y\S. norm (f n x - f n y - (g x - g y)) \ e * norm (x - y)" proof (intro exI ballI allI impI) fix n x y assume as: "N \ n" "x \ S" "y \ S" have "((\m. norm (f n x - f n y - (f m x - f m y))) \ norm (f n x - f n y - (g x - g y))) sequentially" by (intro tendsto_intros g[rule_format] as) moreover have "eventually (\m. norm (f n x - f n y - (f m x - f m y)) \ e * norm (x - y)) sequentially" unfolding eventually_sequentially proof (intro exI allI impI) fix m assume "N \ m" then show "norm (f n x - f n y - (f m x - f m y)) \ e * norm (x - y)" using N as by (auto simp add: algebra_simps) qed ultimately show "norm (f n x - f n y - (g x - g y)) \ e * norm (x - y)" by (simp add: tendsto_upperbound) qed qed have "\x\S. ((\n. f n x) \ g x) sequentially \ (g has_derivative g' x) (at x within S)" unfolding has_derivative_within_alt2 proof (intro ballI conjI allI impI) fix x assume "x \ S" then show "(\n. f n x) \ g x" by (simp add: g) have tog': "(\n. f' n x u) \ g' x u" for u unfolding filterlim_def le_nhds_metric_le eventually_filtermap dist_norm proof (intro allI impI) fix e :: real assume "e > 0" show "eventually (\n. norm (f' n x u - g' x u) \ e) sequentially" proof (cases "u = 0") case True have "eventually (\n. norm (f' n x u - g' x u) \ e * norm u) sequentially" using nle \0 < e\ \x \ S\ by (fast elim: eventually_mono) then show ?thesis using \u = 0\ \0 < e\ by (auto elim: eventually_mono) next case False with \0 < e\ have "0 < e / norm u" by simp then have "eventually (\n. norm (f' n x u - g' x u) \ e / norm u * norm u) sequentially" using nle \x \ S\ by (fast elim: eventually_mono) then show ?thesis using \u \ 0\ by simp qed qed show "bounded_linear (g' x)" proof fix x' y z :: 'a fix c :: real note lin = assms(2)[rule_format,OF \x\S\,THEN has_derivative_bounded_linear] show "g' x (c *\<^sub>R x') = c *\<^sub>R g' x x'" apply (rule tendsto_unique[OF trivial_limit_sequentially tog']) unfolding lin[THEN bounded_linear.linear, THEN linear_cmul] apply (intro tendsto_intros tog') done show "g' x (y + z) = g' x y + g' x z" apply (rule tendsto_unique[OF trivial_limit_sequentially tog']) unfolding lin[THEN bounded_linear.linear, THEN linear_add] apply (rule tendsto_add) apply (rule tog')+ done obtain N where N: "\h. norm (f' N x h - g' x h) \ 1 * norm h" using nle \x \ S\ unfolding eventually_sequentially by (fast intro: zero_less_one) have "bounded_linear (f' N x)" using derf \x \ S\ by fast from bounded_linear.bounded [OF this] obtain K where K: "\h. norm (f' N x h) \ norm h * K" .. { fix h have "norm (g' x h) = norm (f' N x h - (f' N x h - g' x h))" by simp also have "\ \ norm (f' N x h) + norm (f' N x h - g' x h)" by (rule norm_triangle_ineq4) also have "\ \ norm h * K + 1 * norm h" using N K by (fast intro: add_mono) finally have "norm (g' x h) \ norm h * (K + 1)" by (simp add: ring_distribs) } then show "\K. \h. norm (g' x h) \ norm h * K" by fast qed show "eventually (\y. norm (g y - g x - g' x (y - x)) \ e * norm (y - x)) (at x within S)" if "e > 0" for e proof - have *: "e / 3 > 0" using that by auto obtain N1 where N1: "\n\N1. \x\S. \h. norm (f' n x h - g' x h) \ e / 3 * norm h" using nle * unfolding eventually_sequentially by blast obtain N2 where N2[rule_format]: "\n\N2. \x\S. \y\S. norm (f n x - f n y - (g x - g y)) \ e / 3 * norm (x - y)" using lem2 * by blast let ?N = "max N1 N2" have "eventually (\y. norm (f ?N y - f ?N x - f' ?N x (y - x)) \ e / 3 * norm (y - x)) (at x within S)" using derf[unfolded has_derivative_within_alt2] and \x \ S\ and * by fast moreover have "eventually (\y. y \ S) (at x within S)" unfolding eventually_at by (fast intro: zero_less_one) ultimately show "\\<^sub>F y in at x within S. norm (g y - g x - g' x (y - x)) \ e * norm (y - x)" proof (rule eventually_elim2) fix y assume "y \ S" assume "norm (f ?N y - f ?N x - f' ?N x (y - x)) \ e / 3 * norm (y - x)" moreover have "norm (g y - g x - (f ?N y - f ?N x)) \ e / 3 * norm (y - x)" using N2[OF _ \y \ S\ \x \ S\] by (simp add: norm_minus_commute) ultimately have "norm (g y - g x - f' ?N x (y - x)) \ 2 * e / 3 * norm (y - x)" using norm_triangle_le[of "g y - g x - (f ?N y - f ?N x)" "f ?N y - f ?N x - f' ?N x (y - x)" "2 * e / 3 * norm (y - x)"] by (auto simp add: algebra_simps) moreover have " norm (f' ?N x (y - x) - g' x (y - x)) \ e / 3 * norm (y - x)" using N1 \x \ S\ by auto ultimately show "norm (g y - g x - g' x (y - x)) \ e * norm (y - x)" using norm_triangle_le[of "g y - g x - f' (max N1 N2) x (y - x)" "f' (max N1 N2) x (y - x) - g' x (y - x)"] by (auto simp add: algebra_simps) qed qed qed then show ?thesis by fast qed text \Can choose to line up antiderivatives if we want.\ lemma has_antiderivative_sequence: fixes f :: "nat \ 'a::real_normed_vector \ 'b::banach" assumes "convex S" and der: "\n x. x \ S \ ((f n) has_derivative (f' n x)) (at x within S)" and no: "\e. e > 0 \ \\<^sub>F n in sequentially. \x\S. \h. norm (f' n x h - g' x h) \ e * norm h" shows "\g. \x\S. (g has_derivative g' x) (at x within S)" proof (cases "S = {}") case False then obtain a where "a \ S" by auto have *: "\P Q. \g. \x\S. P g x \ Q g x \ \g. \x\S. Q g x" by auto show ?thesis apply (rule *) apply (rule has_derivative_sequence [OF \convex S\ _ no, of "\n x. f n x + (f 0 a - f n a)"]) apply (metis assms(2) has_derivative_add_const) using \a \ S\ apply auto done qed auto lemma has_antiderivative_limit: fixes g' :: "'a::real_normed_vector \ 'a \ 'b::banach" assumes "convex S" and "\e. e>0 \ \f f'. \x\S. (f has_derivative (f' x)) (at x within S) \ (\h. norm (f' x h - g' x h) \ e * norm h)" shows "\g. \x\S. (g has_derivative g' x) (at x within S)" proof - have *: "\n. \f f'. \x\S. (f has_derivative (f' x)) (at x within S) \ (\h. norm(f' x h - g' x h) \ inverse (real (Suc n)) * norm h)" by (simp add: assms(2)) obtain f where *: "\x. \f'. \xa\S. (f x has_derivative f' xa) (at xa within S) \ (\h. norm (f' xa h - g' xa h) \ inverse (real (Suc x)) * norm h)" using * by metis obtain f' where f': "\x. \z\S. (f x has_derivative f' x z) (at z within S) \ (\h. norm (f' x z h - g' z h) \ inverse (real (Suc x)) * norm h)" using * by metis show ?thesis proof (rule has_antiderivative_sequence[OF \convex S\, of f f']) fix e :: real assume "e > 0" obtain N where N: "inverse (real (Suc N)) < e" using reals_Archimedean[OF \e>0\] .. show "\\<^sub>F n in sequentially. \x\S. \h. norm (f' n x h - g' x h) \ e * norm h" unfolding eventually_sequentially proof (intro exI allI ballI impI) fix n x h assume n: "N \ n" and x: "x \ S" have *: "inverse (real (Suc n)) \ e" apply (rule order_trans[OF _ N[THEN less_imp_le]]) using n apply (auto simp add: field_simps) done show "norm (f' n x h - g' x h) \ e * norm h" by (meson "*" mult_right_mono norm_ge_zero order.trans x f') qed qed (use f' in auto) qed subsection \Differentiation of a series\ proposition has_derivative_series: fixes f :: "nat \ 'a::real_normed_vector \ 'b::banach" assumes "convex S" and "\n x. x \ S \ ((f n) has_derivative (f' n x)) (at x within S)" and "\e. e>0 \ \\<^sub>F n in sequentially. \x\S. \h. norm (sum (\i. f' i x h) {.. e * norm h" and "x \ S" and "(\n. f n x) sums l" shows "\g. \x\S. (\n. f n x) sums (g x) \ (g has_derivative g' x) (at x within S)" unfolding sums_def apply (rule has_derivative_sequence[OF assms(1) _ assms(3)]) apply (metis assms(2) has_derivative_sum) using assms(4-5) unfolding sums_def apply auto done lemma has_field_derivative_series: fixes f :: "nat \ ('a :: {real_normed_field,banach}) \ 'a" assumes "convex S" assumes "\n x. x \ S \ (f n has_field_derivative f' n x) (at x within S)" assumes "uniform_limit S (\n x. \i S" "summable (\n. f n x0)" shows "\g. \x\S. (\n. f n x) sums g x \ (g has_field_derivative g' x) (at x within S)" unfolding has_field_derivative_def proof (rule has_derivative_series) show "\\<^sub>F n in sequentially. \x\S. \h. norm ((\i e * norm h" if "e > 0" for e unfolding eventually_sequentially proof - from that assms(3) obtain N where N: "\n x. n \ N \ x \ S \ norm ((\i N" "x \ S" have "norm ((\iii e" by simp hence "norm ((\i e * norm h" by (intro mult_right_mono) simp_all finally have "norm ((\i e * norm h" . } thus "\N. \n\N. \x\S. \h. norm ((\i e * norm h" by blast qed qed (use assms in \auto simp: has_field_derivative_def\) lemma has_field_derivative_series': fixes f :: "nat \ ('a :: {real_normed_field,banach}) \ 'a" assumes "convex S" assumes "\n x. x \ S \ (f n has_field_derivative f' n x) (at x within S)" assumes "uniformly_convergent_on S (\n x. \i S" "summable (\n. f n x0)" "x \ interior S" shows "summable (\n. f n x)" "((\x. \n. f n x) has_field_derivative (\n. f' n x)) (at x)" proof - from \x \ interior S\ have "x \ S" using interior_subset by blast define g' where [abs_def]: "g' x = (\i. f' i x)" for x from assms(3) have "uniform_limit S (\n x. \ix. x \ S \ (\n. f n x) sums g x" "\x. x \ S \ (g has_field_derivative g' x) (at x within S)" by blast from g(1)[OF \x \ S\] show "summable (\n. f n x)" by (simp add: sums_iff) from g(2)[OF \x \ S\] \x \ interior S\ have "(g has_field_derivative g' x) (at x)" by (simp add: at_within_interior[of x S]) also have "(g has_field_derivative g' x) (at x) \ ((\x. \n. f n x) has_field_derivative g' x) (at x)" using eventually_nhds_in_nhd[OF \x \ interior S\] interior_subset[of S] g(1) by (intro DERIV_cong_ev) (auto elim!: eventually_mono simp: sums_iff) finally show "((\x. \n. f n x) has_field_derivative g' x) (at x)" . qed lemma differentiable_series: fixes f :: "nat \ ('a :: {real_normed_field,banach}) \ 'a" assumes "convex S" "open S" assumes "\n x. x \ S \ (f n has_field_derivative f' n x) (at x)" assumes "uniformly_convergent_on S (\n x. \i S" "summable (\n. f n x0)" and x: "x \ S" shows "summable (\n. f n x)" and "(\x. \n. f n x) differentiable (at x)" proof - from assms(4) obtain g' where A: "uniform_limit S (\n x. \iopen S\ have S: "at x within S = at x" by (rule at_within_open) have "\g. \x\S. (\n. f n x) sums g x \ (g has_field_derivative g' x) (at x within S)" by (intro has_field_derivative_series[of S f f' g' x0] assms A has_field_derivative_at_within) then obtain g where g: "\x. x \ S \ (\n. f n x) sums g x" "\x. x \ S \ (g has_field_derivative g' x) (at x within S)" by blast from g[OF x] show "summable (\n. f n x)" by (auto simp: summable_def) from g(2)[OF x] have g': "(g has_derivative (*) (g' x)) (at x)" by (simp add: has_field_derivative_def S) have "((\x. \n. f n x) has_derivative (*) (g' x)) (at x)" by (rule has_derivative_transform_within_open[OF g' \open S\ x]) (insert g, auto simp: sums_iff) thus "(\x. \n. f n x) differentiable (at x)" unfolding differentiable_def by (auto simp: summable_def differentiable_def has_field_derivative_def) qed lemma differentiable_series': fixes f :: "nat \ ('a :: {real_normed_field,banach}) \ 'a" assumes "convex S" "open S" assumes "\n x. x \ S \ (f n has_field_derivative f' n x) (at x)" assumes "uniformly_convergent_on S (\n x. \i S" "summable (\n. f n x0)" shows "(\x. \n. f n x) differentiable (at x0)" using differentiable_series[OF assms, of x0] \x0 \ S\ by blast+ subsection \Derivative as a vector\ text \Considering derivative \<^typ>\real \ 'b::real_normed_vector\ as a vector.\ definition "vector_derivative f net = (SOME f'. (f has_vector_derivative f') net)" lemma vector_derivative_unique_within: assumes not_bot: "at x within S \ bot" and f': "(f has_vector_derivative f') (at x within S)" and f'': "(f has_vector_derivative f'') (at x within S)" shows "f' = f''" proof - have "(\x. x *\<^sub>R f') = (\x. x *\<^sub>R f'')" proof (rule frechet_derivative_unique_within, simp_all) show "\d. d \ 0 \ \d\ < e \ x + d \ S" if "0 < e" for e proof - from that obtain x' where "x' \ S" "x' \ x" "\x' - x\ < e" using islimpt_approachable_real[of x S] not_bot by (auto simp add: trivial_limit_within) then show ?thesis using eq_iff_diff_eq_0 by fastforce qed qed (use f' f'' in \auto simp: has_vector_derivative_def\) then show ?thesis unfolding fun_eq_iff by (metis scaleR_one) qed lemma vector_derivative_unique_at: "(f has_vector_derivative f') (at x) \ (f has_vector_derivative f'') (at x) \ f' = f''" by (rule vector_derivative_unique_within) auto lemma differentiableI_vector: "(f has_vector_derivative y) F \ f differentiable F" by (auto simp: differentiable_def has_vector_derivative_def) proposition vector_derivative_works: "f differentiable net \ (f has_vector_derivative (vector_derivative f net)) net" (is "?l = ?r") proof assume ?l obtain f' where f': "(f has_derivative f') net" using \?l\ unfolding differentiable_def .. then interpret bounded_linear f' by auto show ?r unfolding vector_derivative_def has_vector_derivative_def by (rule someI[of _ "f' 1"]) (simp add: scaleR[symmetric] f') qed (auto simp: vector_derivative_def has_vector_derivative_def differentiable_def) lemma vector_derivative_within: assumes not_bot: "at x within S \ bot" and y: "(f has_vector_derivative y) (at x within S)" shows "vector_derivative f (at x within S) = y" using y by (intro vector_derivative_unique_within[OF not_bot vector_derivative_works[THEN iffD1] y]) (auto simp: differentiable_def has_vector_derivative_def) lemma frechet_derivative_eq_vector_derivative: assumes "f differentiable (at x)" shows "(frechet_derivative f (at x)) = (\r. r *\<^sub>R vector_derivative f (at x))" using assms by (auto simp: differentiable_iff_scaleR vector_derivative_def has_vector_derivative_def intro: someI frechet_derivative_at [symmetric]) lemma has_real_derivative: fixes f :: "real \ real" assumes "(f has_derivative f') F" obtains c where "(f has_real_derivative c) F" proof - obtain c where "f' = (\x. x * c)" by (metis assms has_derivative_bounded_linear real_bounded_linear) then show ?thesis by (metis assms that has_field_derivative_def mult_commute_abs) qed lemma has_real_derivative_iff: fixes f :: "real \ real" shows "(\c. (f has_real_derivative c) F) = (\D. (f has_derivative D) F)" by (metis has_field_derivative_def has_real_derivative) lemma has_vector_derivative_cong_ev: assumes *: "eventually (\x. x \ S \ f x = g x) (nhds x)" "f x = g x" shows "(f has_vector_derivative f') (at x within S) = (g has_vector_derivative f') (at x within S)" unfolding has_vector_derivative_def has_derivative_def using * apply (cases "at x within S \ bot") apply (intro refl conj_cong filterlim_cong) apply (auto simp: Lim_ident_at eventually_at_filter elim: eventually_mono) done lemma islimpt_closure_open: fixes s :: "'a::perfect_space set" assumes "open s" and t: "t = closure s" "x \ t" shows "x islimpt t" proof cases assume "x \ s" { fix T assume "x \ T" "open T" then have "open (s \ T)" using \open s\ by auto then have "s \ T \ {x}" using not_open_singleton[of x] by auto with \x \ T\ \x \ s\ have "\y\t. y \ T \ y \ x" using closure_subset[of s] by (auto simp: t) } then show ?thesis by (auto intro!: islimptI) next assume "x \ s" with t show ?thesis unfolding t closure_def by (auto intro: islimpt_subset) qed lemma vector_derivative_unique_within_closed_interval: assumes ab: "a < b" "x \ cbox a b" assumes D: "(f has_vector_derivative f') (at x within cbox a b)" "(f has_vector_derivative f'') (at x within cbox a b)" shows "f' = f''" using ab by (intro vector_derivative_unique_within[OF _ D]) (auto simp: trivial_limit_within intro!: islimpt_closure_open[where s="{a <..< b}"]) lemma vector_derivative_at: "(f has_vector_derivative f') (at x) \ vector_derivative f (at x) = f'" by (intro vector_derivative_within at_neq_bot) lemma has_vector_derivative_id_at [simp]: "vector_derivative (\x. x) (at a) = 1" by (simp add: vector_derivative_at) lemma vector_derivative_minus_at [simp]: "f differentiable at a \ vector_derivative (\x. - f x) (at a) = - vector_derivative f (at a)" by (simp add: vector_derivative_at has_vector_derivative_minus vector_derivative_works [symmetric]) lemma vector_derivative_add_at [simp]: "\f differentiable at a; g differentiable at a\ \ vector_derivative (\x. f x + g x) (at a) = vector_derivative f (at a) + vector_derivative g (at a)" by (simp add: vector_derivative_at has_vector_derivative_add vector_derivative_works [symmetric]) lemma vector_derivative_diff_at [simp]: "\f differentiable at a; g differentiable at a\ \ vector_derivative (\x. f x - g x) (at a) = vector_derivative f (at a) - vector_derivative g (at a)" by (simp add: vector_derivative_at has_vector_derivative_diff vector_derivative_works [symmetric]) lemma vector_derivative_mult_at [simp]: fixes f g :: "real \ 'a :: real_normed_algebra" shows "\f differentiable at a; g differentiable at a\ \ vector_derivative (\x. f x * g x) (at a) = f a * vector_derivative g (at a) + vector_derivative f (at a) * g a" by (simp add: vector_derivative_at has_vector_derivative_mult vector_derivative_works [symmetric]) lemma vector_derivative_scaleR_at [simp]: "\f differentiable at a; g differentiable at a\ \ vector_derivative (\x. f x *\<^sub>R g x) (at a) = f a *\<^sub>R vector_derivative g (at a) + vector_derivative f (at a) *\<^sub>R g a" apply (rule vector_derivative_at) apply (rule has_vector_derivative_scaleR) apply (auto simp: vector_derivative_works has_vector_derivative_def has_field_derivative_def mult_commute_abs) done lemma vector_derivative_within_cbox: assumes ab: "a < b" "x \ cbox a b" assumes f: "(f has_vector_derivative f') (at x within cbox a b)" shows "vector_derivative f (at x within cbox a b) = f'" by (intro vector_derivative_unique_within_closed_interval[OF ab _ f] vector_derivative_works[THEN iffD1] differentiableI_vector) fact lemma vector_derivative_within_closed_interval: fixes f::"real \ 'a::euclidean_space" assumes "a < b" and "x \ {a..b}" assumes "(f has_vector_derivative f') (at x within {a..b})" shows "vector_derivative f (at x within {a..b}) = f'" using assms vector_derivative_within_cbox by fastforce lemma has_vector_derivative_within_subset: "(f has_vector_derivative f') (at x within S) \ T \ S \ (f has_vector_derivative f') (at x within T)" by (auto simp: has_vector_derivative_def intro: has_derivative_subset) lemma has_vector_derivative_at_within: "(f has_vector_derivative f') (at x) \ (f has_vector_derivative f') (at x within S)" unfolding has_vector_derivative_def by (rule has_derivative_at_withinI) lemma has_vector_derivative_weaken: fixes x D and f g S T assumes f: "(f has_vector_derivative D) (at x within T)" and "x \ S" "S \ T" and "\x. x \ S \ f x = g x" shows "(g has_vector_derivative D) (at x within S)" proof - have "(f has_vector_derivative D) (at x within S) \ (g has_vector_derivative D) (at x within S)" unfolding has_vector_derivative_def has_derivative_iff_norm using assms by (intro conj_cong Lim_cong_within refl) auto then show ?thesis using has_vector_derivative_within_subset[OF f \S \ T\] by simp qed lemma has_vector_derivative_transform_within: assumes "(f has_vector_derivative f') (at x within S)" and "0 < d" and "x \ S" and "\x'. \x'\S; dist x' x < d\ \ f x' = g x'" shows "(g has_vector_derivative f') (at x within S)" using assms unfolding has_vector_derivative_def by (rule has_derivative_transform_within) lemma has_vector_derivative_transform_within_open: assumes "(f has_vector_derivative f') (at x)" and "open S" and "x \ S" and "\y. y\S \ f y = g y" shows "(g has_vector_derivative f') (at x)" using assms unfolding has_vector_derivative_def by (rule has_derivative_transform_within_open) lemma has_vector_derivative_transform: assumes "x \ S" "\x. x \ S \ g x = f x" assumes f': "(f has_vector_derivative f') (at x within S)" shows "(g has_vector_derivative f') (at x within S)" using assms unfolding has_vector_derivative_def by (rule has_derivative_transform) lemma vector_diff_chain_at: assumes "(f has_vector_derivative f') (at x)" and "(g has_vector_derivative g') (at (f x))" shows "((g \ f) has_vector_derivative (f' *\<^sub>R g')) (at x)" using assms has_vector_derivative_at_within has_vector_derivative_def vector_derivative_diff_chain_within by blast lemma vector_diff_chain_within: assumes "(f has_vector_derivative f') (at x within s)" and "(g has_vector_derivative g') (at (f x) within f ` s)" shows "((g \ f) has_vector_derivative (f' *\<^sub>R g')) (at x within s)" using assms has_vector_derivative_def vector_derivative_diff_chain_within by blast lemma vector_derivative_const_at [simp]: "vector_derivative (\x. c) (at a) = 0" by (simp add: vector_derivative_at) lemma vector_derivative_at_within_ivl: "(f has_vector_derivative f') (at x) \ a \ x \ x \ b \ a vector_derivative f (at x within {a..b}) = f'" using has_vector_derivative_at_within vector_derivative_within_cbox by fastforce lemma vector_derivative_chain_at: assumes "f differentiable at x" "(g differentiable at (f x))" shows "vector_derivative (g \ f) (at x) = vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x))" by (metis vector_diff_chain_at vector_derivative_at vector_derivative_works assms) lemma field_vector_diff_chain_at: (*thanks to Wenda Li*) assumes Df: "(f has_vector_derivative f') (at x)" and Dg: "(g has_field_derivative g') (at (f x))" shows "((g \ f) has_vector_derivative (f' * g')) (at x)" using diff_chain_at[OF Df[unfolded has_vector_derivative_def] Dg [unfolded has_field_derivative_def]] by (auto simp: o_def mult.commute has_vector_derivative_def) lemma vector_derivative_chain_within: assumes "at x within S \ bot" "f differentiable (at x within S)" "(g has_derivative g') (at (f x) within f ` S)" shows "vector_derivative (g \ f) (at x within S) = g' (vector_derivative f (at x within S)) " apply (rule vector_derivative_within [OF \at x within S \ bot\]) apply (rule vector_derivative_diff_chain_within) using assms(2-3) vector_derivative_works by auto subsection \Field differentiability\ definition\<^marker>\tag important\ field_differentiable :: "['a \ 'a::real_normed_field, 'a filter] \ bool" (infixr "(field'_differentiable)" 50) where "f field_differentiable F \ \f'. (f has_field_derivative f') F" lemma field_differentiable_imp_differentiable: "f field_differentiable F \ f differentiable F" unfolding field_differentiable_def differentiable_def using has_field_derivative_imp_has_derivative by auto lemma field_differentiable_imp_continuous_at: "f field_differentiable (at x within S) \ continuous (at x within S) f" by (metis DERIV_continuous field_differentiable_def) lemma field_differentiable_within_subset: "\f field_differentiable (at x within S); T \ S\ \ f field_differentiable (at x within T)" by (metis DERIV_subset field_differentiable_def) lemma field_differentiable_at_within: "\f field_differentiable (at x)\ \ f field_differentiable (at x within S)" unfolding field_differentiable_def by (metis DERIV_subset top_greatest) lemma field_differentiable_linear [simp,derivative_intros]: "((*) c) field_differentiable F" unfolding field_differentiable_def has_field_derivative_def mult_commute_abs by (force intro: has_derivative_mult_right) lemma field_differentiable_const [simp,derivative_intros]: "(\z. c) field_differentiable F" unfolding field_differentiable_def has_field_derivative_def using DERIV_const has_field_derivative_imp_has_derivative by blast lemma field_differentiable_ident [simp,derivative_intros]: "(\z. z) field_differentiable F" unfolding field_differentiable_def has_field_derivative_def using DERIV_ident has_field_derivative_def by blast lemma field_differentiable_id [simp,derivative_intros]: "id field_differentiable F" unfolding id_def by (rule field_differentiable_ident) lemma field_differentiable_minus [derivative_intros]: "f field_differentiable F \ (\z. - (f z)) field_differentiable F" unfolding field_differentiable_def by (metis field_differentiable_minus) lemma field_differentiable_add [derivative_intros]: assumes "f field_differentiable F" "g field_differentiable F" shows "(\z. f z + g z) field_differentiable F" using assms unfolding field_differentiable_def by (metis field_differentiable_add) lemma field_differentiable_add_const [simp,derivative_intros]: "(+) c field_differentiable F" by (simp add: field_differentiable_add) lemma field_differentiable_sum [derivative_intros]: "(\i. i \ I \ (f i) field_differentiable F) \ (\z. \i\I. f i z) field_differentiable F" by (induct I rule: infinite_finite_induct) (auto intro: field_differentiable_add field_differentiable_const) lemma field_differentiable_diff [derivative_intros]: assumes "f field_differentiable F" "g field_differentiable F" shows "(\z. f z - g z) field_differentiable F" using assms unfolding field_differentiable_def by (metis field_differentiable_diff) lemma field_differentiable_inverse [derivative_intros]: assumes "f field_differentiable (at a within S)" "f a \ 0" shows "(\z. inverse (f z)) field_differentiable (at a within S)" using assms unfolding field_differentiable_def by (metis DERIV_inverse_fun) lemma field_differentiable_mult [derivative_intros]: assumes "f field_differentiable (at a within S)" "g field_differentiable (at a within S)" shows "(\z. f z * g z) field_differentiable (at a within S)" using assms unfolding field_differentiable_def by (metis DERIV_mult [of f _ a S g]) lemma field_differentiable_divide [derivative_intros]: assumes "f field_differentiable (at a within S)" "g field_differentiable (at a within S)" "g a \ 0" shows "(\z. f z / g z) field_differentiable (at a within S)" using assms unfolding field_differentiable_def by (metis DERIV_divide [of f _ a S g]) lemma field_differentiable_power [derivative_intros]: assumes "f field_differentiable (at a within S)" shows "(\z. f z ^ n) field_differentiable (at a within S)" using assms unfolding field_differentiable_def by (metis DERIV_power) lemma field_differentiable_transform_within: "0 < d \ x \ S \ (\x'. x' \ S \ dist x' x < d \ f x' = g x') \ f field_differentiable (at x within S) \ g field_differentiable (at x within S)" unfolding field_differentiable_def has_field_derivative_def by (blast intro: has_derivative_transform_within) lemma field_differentiable_compose_within: assumes "f field_differentiable (at a within S)" "g field_differentiable (at (f a) within f`S)" shows "(g o f) field_differentiable (at a within S)" using assms unfolding field_differentiable_def by (metis DERIV_image_chain) lemma field_differentiable_compose: "f field_differentiable at z \ g field_differentiable at (f z) \ (g o f) field_differentiable at z" by (metis field_differentiable_at_within field_differentiable_compose_within) lemma field_differentiable_within_open: "\a \ S; open S\ \ f field_differentiable at a within S \ f field_differentiable at a" unfolding field_differentiable_def by (metis at_within_open) lemma exp_scaleR_has_vector_derivative_right: "((\t. exp (t *\<^sub>R A)) has_vector_derivative exp (t *\<^sub>R A) * A) (at t within T)" unfolding has_vector_derivative_def proof (rule has_derivativeI) let ?F = "at t within (T \ {t - 1 <..< t + 1})" have *: "at t within T = ?F" by (rule at_within_nhd[where S="{t - 1 <..< t + 1}"]) auto let ?e = "\i x. (inverse (1 + real i) * inverse (fact i) * (x - t) ^ i) *\<^sub>R (A * A ^ i)" have "\\<^sub>F n in sequentially. \x\T \ {t - 1<.. norm (A ^ (n + 1) /\<^sub>R fact (n + 1))" apply (auto simp: algebra_split_simps intro!: eventuallyI) apply (rule mult_left_mono) apply (auto simp add: field_simps power_abs intro!: divide_right_mono power_le_one) done then have "uniform_limit (T \ {t - 1<..n x. \ix. \i. ?e i x) sequentially" by (rule Weierstrass_m_test_ev) (intro summable_ignore_initial_segment summable_norm_exp) moreover have "\\<^sub>F x in sequentially. x > 0" by (metis eventually_gt_at_top) then have "\\<^sub>F n in sequentially. ((\x. \i A) ?F" by eventually_elim (auto intro!: tendsto_eq_intros simp: power_0_left if_distrib if_distribR cong: if_cong) ultimately have [tendsto_intros]: "((\x. \i. ?e i x) \ A) ?F" by (auto intro!: swap_uniform_limit[where f="\n x. \i < n. ?e i x" and F = sequentially]) have [tendsto_intros]: "((\x. if x = t then 0 else 1) \ 1) ?F" by (rule tendsto_eventually) (simp add: eventually_at_filter) have "((\y. ((y - t) / abs (y - t)) *\<^sub>R ((\n. ?e n y) - A)) \ 0) (at t within T)" unfolding * by (rule tendsto_norm_zero_cancel) (auto intro!: tendsto_eq_intros) moreover have "\\<^sub>F x in at t within T. x \ t" by (simp add: eventually_at_filter) then have "\\<^sub>F x in at t within T. ((x - t) / \x - t\) *\<^sub>R ((\n. ?e n x) - A) = (exp ((x - t) *\<^sub>R A) - 1 - (x - t) *\<^sub>R A) /\<^sub>R norm (x - t)" proof eventually_elim case (elim x) have "(exp ((x - t) *\<^sub>R A) - 1 - (x - t) *\<^sub>R A) /\<^sub>R norm (x - t) = ((\n. (x - t) *\<^sub>R ?e n x) - (x - t) *\<^sub>R A) /\<^sub>R norm (x - t)" unfolding exp_first_term by (simp add: ac_simps) also have "summable (\n. ?e n x)" proof - from elim have "?e n x = (((x - t) *\<^sub>R A) ^ (n + 1)) /\<^sub>R fact (n + 1) /\<^sub>R (x - t)" for n by simp then show ?thesis by (auto simp only: intro!: summable_scaleR_right summable_ignore_initial_segment summable_exp_generic) qed then have "(\n. (x - t) *\<^sub>R ?e n x) = (x - t) *\<^sub>R (\n. ?e n x)" by (rule suminf_scaleR_right[symmetric]) also have "(\ - (x - t) *\<^sub>R A) /\<^sub>R norm (x - t) = (x - t) *\<^sub>R ((\n. ?e n x) - A) /\<^sub>R norm (x - t)" by (simp add: algebra_simps) finally show ?case by simp (simp add: field_simps) qed ultimately have "((\y. (exp ((y - t) *\<^sub>R A) - 1 - (y - t) *\<^sub>R A) /\<^sub>R norm (y - t)) \ 0) (at t within T)" by (rule Lim_transform_eventually) from tendsto_mult_right_zero[OF this, where c="exp (t *\<^sub>R A)"] show "((\y. (exp (y *\<^sub>R A) - exp (t *\<^sub>R A) - (y - t) *\<^sub>R (exp (t *\<^sub>R A) * A)) /\<^sub>R norm (y - t)) \ 0) (at t within T)" by (rule Lim_transform_eventually) (auto simp: field_split_simps exp_add_commuting[symmetric]) qed (rule bounded_linear_scaleR_left) lemma exp_times_scaleR_commute: "exp (t *\<^sub>R A) * A = A * exp (t *\<^sub>R A)" using exp_times_arg_commute[symmetric, of "t *\<^sub>R A"] by (auto simp: algebra_simps) lemma exp_scaleR_has_vector_derivative_left: "((\t. exp (t *\<^sub>R A)) has_vector_derivative A * exp (t *\<^sub>R A)) (at t)" using exp_scaleR_has_vector_derivative_right[of A t] by (simp add: exp_times_scaleR_commute) lemma field_differentiable_series: fixes f :: "nat \ 'a::{real_normed_field,banach} \ 'a" assumes "convex S" "open S" assumes "\n x. x \ S \ (f n has_field_derivative f' n x) (at x)" assumes "uniformly_convergent_on S (\n x. \i S" "summable (\n. f n x0)" and x: "x \ S" shows "(\x. \n. f n x) field_differentiable (at x)" proof - from assms(4) obtain g' where A: "uniform_limit S (\n x. \iopen S\ have S: "at x within S = at x" by (rule at_within_open) have "\g. \x\S. (\n. f n x) sums g x \ (g has_field_derivative g' x) (at x within S)" by (intro has_field_derivative_series[of S f f' g' x0] assms A has_field_derivative_at_within) then obtain g where g: "\x. x \ S \ (\n. f n x) sums g x" "\x. x \ S \ (g has_field_derivative g' x) (at x within S)" by blast from g(2)[OF x] have g': "(g has_derivative (*) (g' x)) (at x)" by (simp add: has_field_derivative_def S) have "((\x. \n. f n x) has_derivative (*) (g' x)) (at x)" by (rule has_derivative_transform_within_open[OF g' \open S\ x]) (insert g, auto simp: sums_iff) thus "(\x. \n. f n x) field_differentiable (at x)" unfolding differentiable_def by (auto simp: summable_def field_differentiable_def has_field_derivative_def) qed subsubsection\<^marker>\tag unimportant\\Caratheodory characterization\ lemma field_differentiable_caratheodory_at: "f field_differentiable (at z) \ (\g. (\w. f(w) - f(z) = g(w) * (w - z)) \ continuous (at z) g)" using CARAT_DERIV [of f] by (simp add: field_differentiable_def has_field_derivative_def) lemma field_differentiable_caratheodory_within: "f field_differentiable (at z within s) \ (\g. (\w. f(w) - f(z) = g(w) * (w - z)) \ continuous (at z within s) g)" using DERIV_caratheodory_within [of f] by (simp add: field_differentiable_def has_field_derivative_def) subsection \Field derivative\ definition\<^marker>\tag important\ deriv :: "('a \ 'a::real_normed_field) \ 'a \ 'a" where "deriv f x \ SOME D. DERIV f x :> D" lemma DERIV_imp_deriv: "DERIV f x :> f' \ deriv f x = f'" unfolding deriv_def by (metis some_equality DERIV_unique) lemma DERIV_deriv_iff_has_field_derivative: "DERIV f x :> deriv f x \ (\f'. (f has_field_derivative f') (at x))" by (auto simp: has_field_derivative_def DERIV_imp_deriv) lemma DERIV_deriv_iff_real_differentiable: fixes x :: real shows "DERIV f x :> deriv f x \ f differentiable at x" unfolding differentiable_def by (metis DERIV_imp_deriv has_real_derivative_iff) lemma deriv_cong_ev: assumes "eventually (\x. f x = g x) (nhds x)" "x = y" shows "deriv f x = deriv g y" proof - have "(\D. (f has_field_derivative D) (at x)) = (\D. (g has_field_derivative D) (at y))" by (intro ext DERIV_cong_ev refl assms) thus ?thesis by (simp add: deriv_def assms) qed lemma higher_deriv_cong_ev: assumes "eventually (\x. f x = g x) (nhds x)" "x = y" shows "(deriv ^^ n) f x = (deriv ^^ n) g y" proof - from assms(1) have "eventually (\x. (deriv ^^ n) f x = (deriv ^^ n) g x) (nhds x)" proof (induction n arbitrary: f g) case (Suc n) from Suc.prems have "eventually (\y. eventually (\z. f z = g z) (nhds y)) (nhds x)" by (simp add: eventually_eventually) hence "eventually (\x. deriv f x = deriv g x) (nhds x)" by eventually_elim (rule deriv_cong_ev, simp_all) thus ?case by (auto intro!: deriv_cong_ev Suc simp: funpow_Suc_right simp del: funpow.simps) qed auto from eventually_nhds_x_imp_x[OF this] assms(2) show ?thesis by simp qed lemma real_derivative_chain: fixes x :: real shows "f differentiable at x \ g differentiable at (f x) \ deriv (g o f) x = deriv g (f x) * deriv f x" by (metis DERIV_deriv_iff_real_differentiable DERIV_chain DERIV_imp_deriv) lemma field_derivative_eq_vector_derivative: "(deriv f x) = vector_derivative f (at x)" by (simp add: mult.commute deriv_def vector_derivative_def has_vector_derivative_def has_field_derivative_def) proposition field_differentiable_derivI: "f field_differentiable (at x) \ (f has_field_derivative deriv f x) (at x)" by (simp add: field_differentiable_def DERIV_deriv_iff_has_field_derivative) lemma vector_derivative_chain_at_general: assumes "f differentiable at x" "g field_differentiable at (f x)" shows "vector_derivative (g \ f) (at x) = vector_derivative f (at x) * deriv g (f x)" apply (rule vector_derivative_at [OF field_vector_diff_chain_at]) using assms vector_derivative_works by (auto simp: field_differentiable_derivI) lemma DERIV_deriv_iff_field_differentiable: "DERIV f x :> deriv f x \ f field_differentiable at x" unfolding field_differentiable_def by (metis DERIV_imp_deriv) lemma deriv_chain: "f field_differentiable at x \ g field_differentiable at (f x) \ deriv (g o f) x = deriv g (f x) * deriv f x" by (metis DERIV_deriv_iff_field_differentiable DERIV_chain DERIV_imp_deriv) lemma deriv_linear [simp]: "deriv (\w. c * w) = (\z. c)" by (metis DERIV_imp_deriv DERIV_cmult_Id) lemma deriv_uminus [simp]: "deriv (\w. -w) = (\z. -1)" using deriv_linear[of "-1"] by (simp del: deriv_linear) lemma deriv_ident [simp]: "deriv (\w. w) = (\z. 1)" by (metis DERIV_imp_deriv DERIV_ident) lemma deriv_id [simp]: "deriv id = (\z. 1)" by (simp add: id_def) lemma deriv_const [simp]: "deriv (\w. c) = (\z. 0)" by (metis DERIV_imp_deriv DERIV_const) lemma deriv_add [simp]: "\f field_differentiable at z; g field_differentiable at z\ \ deriv (\w. f w + g w) z = deriv f z + deriv g z" unfolding DERIV_deriv_iff_field_differentiable[symmetric] by (auto intro!: DERIV_imp_deriv derivative_intros) lemma deriv_diff [simp]: "\f field_differentiable at z; g field_differentiable at z\ \ deriv (\w. f w - g w) z = deriv f z - deriv g z" unfolding DERIV_deriv_iff_field_differentiable[symmetric] by (auto intro!: DERIV_imp_deriv derivative_intros) lemma deriv_mult [simp]: "\f field_differentiable at z; g field_differentiable at z\ \ deriv (\w. f w * g w) z = f z * deriv g z + deriv f z * g z" unfolding DERIV_deriv_iff_field_differentiable[symmetric] by (auto intro!: DERIV_imp_deriv derivative_eq_intros) lemma deriv_cmult: "f field_differentiable at z \ deriv (\w. c * f w) z = c * deriv f z" by simp lemma deriv_cmult_right: "f field_differentiable at z \ deriv (\w. f w * c) z = deriv f z * c" by simp lemma deriv_inverse [simp]: "\f field_differentiable at z; f z \ 0\ \ deriv (\w. inverse (f w)) z = - deriv f z / f z ^ 2" unfolding DERIV_deriv_iff_field_differentiable[symmetric] by (safe intro!: DERIV_imp_deriv derivative_eq_intros) (auto simp: field_split_simps power2_eq_square) lemma deriv_divide [simp]: "\f field_differentiable at z; g field_differentiable at z; g z \ 0\ \ deriv (\w. f w / g w) z = (deriv f z * g z - f z * deriv g z) / g z ^ 2" by (simp add: field_class.field_divide_inverse field_differentiable_inverse) (simp add: field_split_simps power2_eq_square) lemma deriv_cdivide_right: "f field_differentiable at z \ deriv (\w. f w / c) z = deriv f z / c" by (simp add: field_class.field_divide_inverse) lemma deriv_compose_linear: "f field_differentiable at (c * z) \ deriv (\w. f (c * w)) z = c * deriv f (c * z)" apply (rule DERIV_imp_deriv) unfolding DERIV_deriv_iff_field_differentiable [symmetric] by (metis (full_types) DERIV_chain2 DERIV_cmult_Id mult.commute) lemma nonzero_deriv_nonconstant: assumes df: "DERIV f \ :> df" and S: "open S" "\ \ S" and "df \ 0" shows "\ f constant_on S" unfolding constant_on_def by (metis \df \ 0\ has_field_derivative_transform_within_open [OF df S] DERIV_const DERIV_unique) subsection \Relation between convexity and derivative\ (* TODO: Generalise to real vector spaces? *) proposition convex_on_imp_above_tangent: assumes convex: "convex_on A f" and connected: "connected A" assumes c: "c \ interior A" and x : "x \ A" assumes deriv: "(f has_field_derivative f') (at c within A)" shows "f x - f c \ f' * (x - c)" proof (cases x c rule: linorder_cases) assume xc: "x > c" let ?A' = "interior A \ {c<..}" from c have "c \ interior A \ closure {c<..}" by auto also have "\ \ closure (interior A \ {c<..})" by (intro open_Int_closure_subset) auto finally have "at c within ?A' \ bot" by (subst at_within_eq_bot_iff) auto moreover from deriv have "((\y. (f y - f c) / (y - c)) \ f') (at c within ?A')" unfolding has_field_derivative_iff using interior_subset[of A] by (blast intro: tendsto_mono at_le) moreover from eventually_at_right_real[OF xc] have "eventually (\y. (f y - f c) / (y - c) \ (f x - f c) / (x - c)) (at_right c)" proof eventually_elim fix y assume y: "y \ {c<.. (f x - f c) / (x - c) * (y - c) + f c" using interior_subset[of A] by (intro convex_onD_Icc' convex_on_subset[OF convex] connected_contains_Icc) auto hence "f y - f c \ (f x - f c) / (x - c) * (y - c)" by simp thus "(f y - f c) / (y - c) \ (f x - f c) / (x - c)" using y xc by (simp add: field_split_simps) qed hence "eventually (\y. (f y - f c) / (y - c) \ (f x - f c) / (x - c)) (at c within ?A')" by (blast intro: filter_leD at_le) ultimately have "f' \ (f x - f c) / (x - c)" by (simp add: tendsto_upperbound) thus ?thesis using xc by (simp add: field_simps) next assume xc: "x < c" let ?A' = "interior A \ {.. interior A \ closure {.. \ closure (interior A \ {.. bot" by (subst at_within_eq_bot_iff) auto moreover from deriv have "((\y. (f y - f c) / (y - c)) \ f') (at c within ?A')" unfolding has_field_derivative_iff using interior_subset[of A] by (blast intro: tendsto_mono at_le) moreover from eventually_at_left_real[OF xc] have "eventually (\y. (f y - f c) / (y - c) \ (f x - f c) / (x - c)) (at_left c)" proof eventually_elim fix y assume y: "y \ {x<.. (f x - f c) / (c - x) * (c - y) + f c" using interior_subset[of A] by (intro convex_onD_Icc'' convex_on_subset[OF convex] connected_contains_Icc) auto hence "f y - f c \ (f x - f c) * ((c - y) / (c - x))" by simp also have "(c - y) / (c - x) = (y - c) / (x - c)" using y xc by (simp add: field_simps) finally show "(f y - f c) / (y - c) \ (f x - f c) / (x - c)" using y xc by (simp add: field_split_simps) qed hence "eventually (\y. (f y - f c) / (y - c) \ (f x - f c) / (x - c)) (at c within ?A')" by (blast intro: filter_leD at_le) ultimately have "f' \ (f x - f c) / (x - c)" by (simp add: tendsto_lowerbound) thus ?thesis using xc by (simp add: field_simps) qed simp_all subsection \Partial derivatives\ lemma eventually_at_Pair_within_TimesI1: fixes x::"'a::metric_space" assumes "\\<^sub>F x' in at x within X. P x'" assumes "P x" shows "\\<^sub>F (x', y') in at (x, y) within X \ Y. P x'" proof - from assms[unfolded eventually_at_topological] obtain S where S: "open S" "x \ S" "\x'. x' \ X \ x' \ S \ P x'" by metis show "\\<^sub>F (x', y') in at (x, y) within X \ Y. P x'" unfolding eventually_at_topological by (auto intro!: exI[where x="S \ UNIV"] S open_Times) qed lemma eventually_at_Pair_within_TimesI2: fixes x::"'a::metric_space" assumes "\\<^sub>F y' in at y within Y. P y'" "P y" shows "\\<^sub>F (x', y') in at (x, y) within X \ Y. P y'" proof - from assms[unfolded eventually_at_topological] obtain S where S: "open S" "y \ S" "\y'. y' \ Y \ y' \ S \ P y'" by metis show "\\<^sub>F (x', y') in at (x, y) within X \ Y. P y'" unfolding eventually_at_topological by (auto intro!: exI[where x="UNIV \ S"] S open_Times) qed proposition has_derivative_partialsI: fixes f::"'a::real_normed_vector \ 'b::real_normed_vector \ 'c::real_normed_vector" assumes fx: "((\x. f x y) has_derivative fx) (at x within X)" assumes fy: "\x y. x \ X \ y \ Y \ ((\y. f x y) has_derivative blinfun_apply (fy x y)) (at y within Y)" assumes fy_cont[unfolded continuous_within]: "continuous (at (x, y) within X \ Y) (\(x, y). fy x y)" assumes "y \ Y" "convex Y" shows "((\(x, y). f x y) has_derivative (\(tx, ty). fx tx + fy x y ty)) (at (x, y) within X \ Y)" proof (safe intro!: has_derivativeI tendstoI, goal_cases) case (2 e') interpret fx: bounded_linear "fx" using fx by (rule has_derivative_bounded_linear) define e where "e = e' / 9" have "e > 0" using \e' > 0\ by (simp add: e_def) from fy_cont[THEN tendstoD, OF \e > 0\] have "\\<^sub>F (x', y') in at (x, y) within X \ Y. dist (fy x' y') (fy x y) < e" by (auto simp: split_beta') from this[unfolded eventually_at] obtain d' where "d' > 0" "\x' y'. x' \ X \ y' \ Y \ (x', y') \ (x, y) \ dist (x', y') (x, y) < d' \ dist (fy x' y') (fy x y) < e" by auto then have d': "x' \ X \ y' \ Y \ dist (x', y') (x, y) < d' \ dist (fy x' y') (fy x y) < e" for x' y' using \0 < e\ by (cases "(x', y') = (x, y)") auto define d where "d = d' / sqrt 2" have "d > 0" using \0 < d'\ by (simp add: d_def) have d: "x' \ X \ y' \ Y \ dist x' x < d \ dist y' y < d \ dist (fy x' y') (fy x y) < e" for x' y' by (auto simp: dist_prod_def d_def intro!: d' real_sqrt_sum_squares_less) let ?S = "ball y d \ Y" have "convex ?S" by (auto intro!: convex_Int \convex Y\) { fix x'::'a and y'::'b assume x': "x' \ X" and y': "y' \ Y" assume dx': "dist x' x < d" and dy': "dist y' y < d" have "norm (fy x' y' - fy x' y) \ dist (fy x' y') (fy x y) + dist (fy x' y) (fy x y)" by norm also have "dist (fy x' y') (fy x y) < e" by (rule d; fact) also have "dist (fy x' y) (fy x y) < e" by (auto intro!: d simp: dist_prod_def x' \d > 0\ \y \ Y\ dx') finally have "norm (fy x' y' - fy x' y) < e + e" by arith then have "onorm (blinfun_apply (fy x' y') - blinfun_apply (fy x' y)) < e + e" by (auto simp: norm_blinfun.rep_eq blinfun.diff_left[abs_def] fun_diff_def) } note onorm = this have ev_mem: "\\<^sub>F (x', y') in at (x, y) within X \ Y. (x', y') \ X \ Y" using \y \ Y\ by (auto simp: eventually_at intro!: zero_less_one) moreover have ev_dist: "\\<^sub>F xy in at (x, y) within X \ Y. dist xy (x, y) < d" if "d > 0" for d using eventually_at_ball[OF that] by (rule eventually_elim2) (auto simp: dist_commute intro!: eventually_True) note ev_dist[OF \0 < d\] ultimately have "\\<^sub>F (x', y') in at (x, y) within X \ Y. norm (f x' y' - f x' y - (fy x' y) (y' - y)) \ norm (y' - y) * (e + e)" proof (eventually_elim, safe) fix x' y' assume "x' \ X" and y': "y' \ Y" assume dist: "dist (x', y') (x, y) < d" then have dx: "dist x' x < d" and dy: "dist y' y < d" unfolding dist_prod_def fst_conv snd_conv atomize_conj by (metis le_less_trans real_sqrt_sum_squares_ge1 real_sqrt_sum_squares_ge2) { fix t::real assume "t \ {0 .. 1}" then have "y + t *\<^sub>R (y' - y) \ closed_segment y y'" by (auto simp: closed_segment_def algebra_simps intro!: exI[where x=t]) also have "\ \ ball y d \ Y" using \y \ Y\ \0 < d\ dy y' by (intro \convex ?S\[unfolded convex_contains_segment, rule_format, of y y']) (auto simp: dist_commute) finally have "y + t *\<^sub>R (y' - y) \ ?S" . } note seg = this have "\x. x \ ball y d \ Y \ onorm (blinfun_apply (fy x' x) - blinfun_apply (fy x' y)) \ e + e" by (safe intro!: onorm less_imp_le \x' \ X\ dx) (auto simp: dist_commute \0 < d\ \y \ Y\) with seg has_derivative_subset[OF assms(2)[OF \x' \ X\]] show "norm (f x' y' - f x' y - (fy x' y) (y' - y)) \ norm (y' - y) * (e + e)" by (rule differentiable_bound_linearization[where S="?S"]) (auto intro!: \0 < d\ \y \ Y\) qed moreover let ?le = "\x'. norm (f x' y - f x y - (fx) (x' - x)) \ norm (x' - x) * e" from fx[unfolded has_derivative_within, THEN conjunct2, THEN tendstoD, OF \0 < e\] have "\\<^sub>F x' in at x within X. ?le x'" by eventually_elim (simp, simp add: dist_norm field_split_simps split: if_split_asm) then have "\\<^sub>F (x', y') in at (x, y) within X \ Y. ?le x'" by (rule eventually_at_Pair_within_TimesI1) (simp add: blinfun.bilinear_simps) moreover have "\\<^sub>F (x', y') in at (x, y) within X \ Y. norm ((x', y') - (x, y)) \ 0" unfolding norm_eq_zero right_minus_eq by (auto simp: eventually_at intro!: zero_less_one) moreover from fy_cont[THEN tendstoD, OF \0 < e\] have "\\<^sub>F x' in at x within X. norm (fy x' y - fy x y) < e" unfolding eventually_at using \y \ Y\ by (auto simp: dist_prod_def dist_norm) then have "\\<^sub>F (x', y') in at (x, y) within X \ Y. norm (fy x' y - fy x y) < e" by (rule eventually_at_Pair_within_TimesI1) (simp add: blinfun.bilinear_simps \0 < e\) ultimately have "\\<^sub>F (x', y') in at (x, y) within X \ Y. norm ((f x' y' - f x y - (fx (x' - x) + fy x y (y' - y))) /\<^sub>R norm ((x', y') - (x, y))) < e'" apply eventually_elim proof safe fix x' y' have "norm (f x' y' - f x y - (fx (x' - x) + fy x y (y' - y))) \ norm (f x' y' - f x' y - fy x' y (y' - y)) + norm (fy x y (y' - y) - fy x' y (y' - y)) + norm (f x' y - f x y - fx (x' - x))" by norm also assume nz: "norm ((x', y') - (x, y)) \ 0" and nfy: "norm (fy x' y - fy x y) < e" assume "norm (f x' y' - f x' y - blinfun_apply (fy x' y) (y' - y)) \ norm (y' - y) * (e + e)" also assume "norm (f x' y - f x y - (fx) (x' - x)) \ norm (x' - x) * e" also have "norm ((fy x y) (y' - y) - (fy x' y) (y' - y)) \ norm ((fy x y) - (fy x' y)) * norm (y' - y)" by (auto simp: blinfun.bilinear_simps[symmetric] intro!: norm_blinfun) also have "\ \ (e + e) * norm (y' - y)" using \e > 0\ nfy by (auto simp: norm_minus_commute intro!: mult_right_mono) also have "norm (x' - x) * e \ norm (x' - x) * (e + e)" using \0 < e\ by simp also have "norm (y' - y) * (e + e) + (e + e) * norm (y' - y) + norm (x' - x) * (e + e) \ (norm (y' - y) + norm (x' - x)) * (4 * e)" using \e > 0\ by (simp add: algebra_simps) also have "\ \ 2 * norm ((x', y') - (x, y)) * (4 * e)" using \0 < e\ real_sqrt_sum_squares_ge1[of "norm (x' - x)" "norm (y' - y)"] real_sqrt_sum_squares_ge2[of "norm (y' - y)" "norm (x' - x)"] by (auto intro!: mult_right_mono simp: norm_prod_def simp del: real_sqrt_sum_squares_ge1 real_sqrt_sum_squares_ge2) also have "\ \ norm ((x', y') - (x, y)) * (8 * e)" by simp also have "\ < norm ((x', y') - (x, y)) * e'" using \0 < e'\ nz by (auto simp: e_def) finally show "norm ((f x' y' - f x y - (fx (x' - x) + fy x y (y' - y))) /\<^sub>R norm ((x', y') - (x, y))) < e'" by (simp add: dist_norm) (auto simp add: field_split_simps) qed then show ?case by eventually_elim (auto simp: dist_norm field_simps) next from has_derivative_bounded_linear[OF fx] obtain fxb where "fx = blinfun_apply fxb" by (metis bounded_linear_Blinfun_apply) then show "bounded_linear (\(tx, ty). fx tx + blinfun_apply (fy x y) ty)" by (auto intro!: bounded_linear_intros simp: split_beta') qed subsection\<^marker>\tag unimportant\ \Differentiable case distinction\ lemma has_derivative_within_If_eq: "((\x. if P x then f x else g x) has_derivative f') (at x within S) = (bounded_linear f' \ ((\y.(if P y then (f y - ((if P x then f x else g x) + f' (y - x)))/\<^sub>R norm (y - x) else (g y - ((if P x then f x else g x) + f' (y - x)))/\<^sub>R norm (y - x))) \ 0) (at x within S))" (is "_ = (_ \ (?if \ 0) _)") proof - have "(\y. (1 / norm (y - x)) *\<^sub>R ((if P y then f y else g y) - ((if P x then f x else g x) + f' (y - x)))) = ?if" by (auto simp: inverse_eq_divide) thus ?thesis by (auto simp: has_derivative_within) qed lemma has_derivative_If_within_closures: assumes f': "x \ S \ (closure S \ closure T) \ (f has_derivative f' x) (at x within S \ (closure S \ closure T))" assumes g': "x \ T \ (closure S \ closure T) \ (g has_derivative g' x) (at x within T \ (closure S \ closure T))" assumes connect: "x \ closure S \ x \ closure T \ f x = g x" assumes connect': "x \ closure S \ x \ closure T \ f' x = g' x" assumes x_in: "x \ S \ T" shows "((\x. if x \ S then f x else g x) has_derivative (if x \ S then f' x else g' x)) (at x within (S \ T))" proof - from f' x_in interpret f': bounded_linear "if x \ S then f' x else (\x. 0)" by (auto simp add: has_derivative_within) from g' interpret g': bounded_linear "if x \ T then g' x else (\x. 0)" by (auto simp add: has_derivative_within) have bl: "bounded_linear (if x \ S then f' x else g' x)" using f'.scaleR f'.bounded f'.add g'.scaleR g'.bounded g'.add x_in by (unfold_locales; force) show ?thesis using f' g' closure_subset[of T] closure_subset[of S] unfolding has_derivative_within_If_eq by (intro conjI bl tendsto_If_within_closures x_in) (auto simp: has_derivative_within inverse_eq_divide connect connect' subsetD) qed lemma has_vector_derivative_If_within_closures: assumes x_in: "x \ S \ T" assumes "u = S \ T" assumes f': "x \ S \ (closure S \ closure T) \ (f has_vector_derivative f' x) (at x within S \ (closure S \ closure T))" assumes g': "x \ T \ (closure S \ closure T) \ (g has_vector_derivative g' x) (at x within T \ (closure S \ closure T))" assumes connect: "x \ closure S \ x \ closure T \ f x = g x" assumes connect': "x \ closure S \ x \ closure T \ f' x = g' x" shows "((\x. if x \ S then f x else g x) has_vector_derivative (if x \ S then f' x else g' x)) (at x within u)" unfolding has_vector_derivative_def assms using x_in apply (intro has_derivative_If_within_closures[where ?f' = "\x a. a *\<^sub>R f' x" and ?g' = "\x a. a *\<^sub>R g' x", THEN has_derivative_eq_rhs]) subgoal by (rule f'[unfolded has_vector_derivative_def]; assumption) subgoal by (rule g'[unfolded has_vector_derivative_def]; assumption) by (auto simp: assms) subsection\<^marker>\tag important\\The Inverse Function Theorem\ lemma linear_injective_contraction: assumes "linear f" "c < 1" and le: "\x. norm (f x - x) \ c * norm x" shows "inj f" unfolding linear_injective_0[OF \linear f\] proof safe fix x assume "f x = 0" with le [of x] have "norm x \ c * norm x" by simp then show "x = 0" using \c < 1\ by (simp add: mult_le_cancel_right1) qed text\From an online proof by J. Michael Boardman, Department of Mathematics, Johns Hopkins University\ lemma inverse_function_theorem_scaled: fixes f::"'a::euclidean_space \ 'a" and f'::"'a \ ('a \\<^sub>L 'a)" assumes "open U" and derf: "\x. x \ U \ (f has_derivative blinfun_apply (f' x)) (at x)" and contf: "continuous_on U f'" and "0 \ U" and [simp]: "f 0 = 0" and id: "f' 0 = id_blinfun" obtains U' V g g' where "open U'" "U' \ U" "0 \ U'" "open V" "0 \ V" "homeomorphism U' V f g" "\y. y \ V \ (g has_derivative (g' y)) (at y)" "\y. y \ V \ g' y = inv (blinfun_apply (f'(g y)))" "\y. y \ V \ bij (blinfun_apply (f'(g y)))" proof - obtain d1 where "cball 0 d1 \ U" "d1 > 0" using \open U\ \0 \ U\ open_contains_cball by blast obtain d2 where d2: "\x. \x \ U; dist x 0 \ d2\ \ dist (f' x) (f' 0) < 1/2" "0 < d2" using continuous_onE [OF contf, of 0 "1/2"] by (metis \0 \ U\ half_gt_zero_iff zero_less_one) obtain \ where le: "\x. norm x \ \ \ dist (f' x) id_blinfun \ 1/2" and "0 < \" and subU: "cball 0 \ \ U" proof show "min d1 d2 > 0" by (simp add: \0 < d1\ \0 < d2\) show "cball 0 (min d1 d2) \ U" using \cball 0 d1 \ U\ by auto show "dist (f' x) id_blinfun \ 1/2" if "norm x \ min d1 d2" for x using \cball 0 d1 \ U\ d2 that id by fastforce qed let ?D = "cball 0 \" define V:: "'a set" where "V \ ball 0 (\/2)" have 4: "norm (f (x + h) - f x - h) \ 1/2 * norm h" if "x \ ?D" "x+h \ ?D" for x h proof - let ?w = "\x. f x - x" have B: "\x. x \ ?D \ onorm (blinfun_apply (f' x - id_blinfun)) \ 1/2" by (metis dist_norm le mem_cball_0 norm_blinfun.rep_eq) have "\x. x \ ?D \ (?w has_derivative (blinfun_apply (f' x - id_blinfun))) (at x)" by (rule derivative_eq_intros derf subsetD [OF subU] | force simp: blinfun.diff_left)+ then have Dw: "\x. x \ ?D \ (?w has_derivative (blinfun_apply (f' x - id_blinfun))) (at x within ?D)" using has_derivative_at_withinI by blast have "norm (?w (x+h) - ?w x) \ (1/2) * norm h" using differentiable_bound [OF convex_cball Dw B] that by fastforce then show ?thesis by (auto simp: algebra_simps) qed have for_g: "\!x. norm x < \ \ f x = y" if y: "norm y < \/2" for y proof - let ?u = "\x. x + (y - f x)" have *: "norm (?u x) < \" if "x \ ?D" for x proof - have fxx: "norm (f x - x) \ \/2" using 4 [of 0 x] \0 < \\ \f 0 = 0\ that by auto have "norm (?u x) \ norm y + norm (f x - x)" by (metis add.commute add_diff_eq norm_minus_commute norm_triangle_ineq) also have "\ < \/2 + \/2" using fxx y by auto finally show ?thesis by simp qed have "\!x \ ?D. ?u x = x" proof (rule banach_fix) show "cball 0 \ \ {}" using \0 < \\ by auto show "(\x. x + (y - f x)) ` cball 0 \ \ cball 0 \" using * by force have "dist (x + (y - f x)) (xh + (y - f xh)) * 2 \ dist x xh" if "norm x \ \" and "norm xh \ \" for x xh using that 4 [of x "xh-x"] by (auto simp: dist_norm norm_minus_commute algebra_simps) then show "\x\cball 0 \. \ya\cball 0 \. dist (x + (y - f x)) (ya + (y - f ya)) \ (1/2) * dist x ya" by auto qed (auto simp: complete_eq_closed) then show ?thesis by (metis "*" add_cancel_right_right eq_iff_diff_eq_0 le_less mem_cball_0) qed define g where "g \ \y. THE x. norm x < \ \ f x = y" have g: "norm (g y) < \ \ f (g y) = y" if "norm y < \/2" for y unfolding g_def using that theI' [OF for_g] by meson then have fg[simp]: "f (g y) = y" if "y \ V" for y using that by (auto simp: V_def) have 5: "norm (g y' - g y) \ 2 * norm (y' - y)" if "y \ V" "y' \ V" for y y' proof - have no: "norm (g y) \ \" "norm (g y') \ \" and [simp]: "f (g y) = y" using that g unfolding V_def by force+ have "norm (g y' - g y) \ norm (g y' - g y - (y' - y)) + norm (y' - y)" by (simp add: add.commute norm_triangle_sub) also have "\ \ (1/2) * norm (g y' - g y) + norm (y' - y)" using 4 [of "g y" "g y' - g y"] that no by (simp add: g norm_minus_commute V_def) finally show ?thesis by auto qed have contg: "continuous_on V g" proof fix y::'a and e::real assume "0 < e" and y: "y \ V" show "\d>0. \x'\V. dist x' y < d \ dist (g x') (g y) \ e" proof (intro exI conjI ballI impI) show "0 < e/2" by (simp add: \0 < e\) qed (use 5 y in \force simp: dist_norm\) qed show thesis proof define U' where "U' \ (f -` V) \ ball 0 \" have contf: "continuous_on U f" using derf has_derivative_at_withinI by (fast intro: has_derivative_continuous_on) then have "continuous_on (ball 0 \) f" by (meson ball_subset_cball continuous_on_subset subU) then show "open U'" by (simp add: U'_def V_def Int_commute continuous_open_preimage) show "0 \ U'" "U' \ U" "open V" "0 \ V" using \0 < \\ subU by (auto simp: U'_def V_def) show hom: "homeomorphism U' V f g" proof show "continuous_on U' f" using \U' \ U\ contf continuous_on_subset by blast show "continuous_on V g" using contg by blast show "f ` U' \ V" using U'_def by blast show "g ` V \ U'" by (simp add: U'_def V_def g image_subset_iff) show "g (f x) = x" if "x \ U'" for x by (metis that fg Int_iff U'_def V_def for_g g mem_ball_0 vimage_eq) show "f (g y) = y" if "y \ V" for y using that by (simp add: g V_def) qed show bij: "bij (blinfun_apply (f'(g y)))" if "y \ V" for y proof - have inj: "inj (blinfun_apply (f' (g y)))" proof (rule linear_injective_contraction) show "linear (blinfun_apply (f' (g y)))" using blinfun.bounded_linear_right bounded_linear_def by blast next fix x have "norm (blinfun_apply (f' (g y)) x - x) = norm (blinfun_apply (f' (g y) - id_blinfun) x)" by (simp add: blinfun.diff_left) also have "\ \ norm (f' (g y) - id_blinfun) * norm x" by (rule norm_blinfun) also have "\ \ (1/2) * norm x" proof (rule mult_right_mono) show "norm (f' (g y) - id_blinfun) \ 1/2" using that g [of y] le by (auto simp: V_def dist_norm) qed auto finally show "norm (blinfun_apply (f' (g y)) x - x) \ (1/2) * norm x" . qed auto moreover have "surj (blinfun_apply (f' (g y)))" using blinfun.bounded_linear_right bounded_linear_def by (blast intro!: linear_inj_imp_surj [OF _ inj]) ultimately show ?thesis using bijI by blast qed define g' where "g' \ \y. inv (blinfun_apply (f'(g y)))" show "(g has_derivative g' y) (at y)" if "y \ V" for y proof - have gy: "g y \ U" using g subU that unfolding V_def by fastforce obtain e where e: "\h. f (g y + h) = y + blinfun_apply (f' (g y)) h + e h" and e0: "(\h. norm (e h) / norm h) \0\ 0" using iffD1 [OF has_derivative_iff_Ex derf [OF gy]] \y \ V\ by auto have [simp]: "e 0 = 0" using e [of 0] that by simp let ?INV = "inv (blinfun_apply (f' (g y)))" have inj: "inj (blinfun_apply (f' (g y)))" using bij bij_betw_def that by blast have "(g has_derivative g' y) (at y within V)" unfolding has_derivative_at_within_iff_Ex [OF \y \ V\ \open V\] proof show blinv: "bounded_linear (g' y)" unfolding g'_def using derf gy inj inj_linear_imp_inv_bounded_linear by blast define eg where "eg \ \k. - ?INV (e (g (y+k) - g y))" have "g (y+k) = g y + g' y k + eg k" if "y + k \ V" for k proof - have "?INV k = ?INV (blinfun_apply (f' (g y)) (g (y+k) - g y) + e (g (y+k) - g y))" using e [of "g(y+k) - g y"] that by simp then have "g (y+k) = g y + ?INV k - ?INV (e (g (y+k) - g y))" using inj blinv by (simp add: linear_simps g'_def) then show ?thesis by (auto simp: eg_def g'_def) qed moreover have "(\k. norm (eg k) / norm k) \0\ 0" proof (rule Lim_null_comparison) let ?g = "\k. 2 * onorm ?INV * norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)" show "\\<^sub>F k in at 0. norm (norm (eg k) / norm k) \ ?g k" unfolding eventually_at_topological proof (intro exI conjI ballI impI) show "open ((+)(-y) ` V)" using \open V\ open_translation by blast show "0 \ (+)(-y) ` V" by (simp add: that) show "norm (norm (eg k) / norm k) \ 2 * onorm (inv (blinfun_apply (f' (g y)))) * norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)" if "k \ (+)(-y) ` V" "k \ 0" for k proof - have "y+k \ V" using that by auto have "norm (norm (eg k) / norm k) \ onorm ?INV * norm (e (g (y+k) - g y)) / norm k" using blinv g'_def onorm by (force simp: eg_def divide_simps) also have "\ = (norm (g (y+k) - g y) / norm k) * (onorm ?INV * (norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)))" by (simp add: divide_simps) also have "\ \ 2 * (onorm ?INV * (norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)))" apply (rule mult_right_mono) using 5 [of y "y+k"] \y \ V\ \y + k \ V\ onorm_pos_le [OF blinv] apply (auto simp: divide_simps zero_le_mult_iff zero_le_divide_iff g'_def) done finally show "norm (norm (eg k) / norm k) \ 2 * onorm ?INV * norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)" by simp qed qed have 1: "(\h. norm (e h) / norm h) \0\ (norm (e 0) / norm 0)" using e0 by auto have 2: "(\k. g (y+k) - g y) \0\ 0" using contg \open V\ \y \ V\ LIM_offset_zero_iff LIM_zero_iff at_within_open continuous_on_def by fastforce from tendsto_compose [OF 1 2, simplified] have "(\k. norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)) \0\ 0" . from tendsto_mult_left [OF this] show "?g \0\ 0" by auto qed ultimately show "\e. (\k. y + k \ V \ g (y+k) = g y + g' y k + e k) \ (\k. norm (e k) / norm k) \0\ 0" by blast qed then show ?thesis by (metis \open V\ at_within_open that) qed show "g' y = inv (blinfun_apply (f' (g y)))" if "y \ V" for y by (simp add: g'_def) qed qed text\We need all this to justify the scaling and translations.\ theorem inverse_function_theorem: fixes f::"'a::euclidean_space \ 'a" and f'::"'a \ ('a \\<^sub>L 'a)" assumes "open U" and derf: "\x. x \ U \ (f has_derivative (blinfun_apply (f' x))) (at x)" and contf: "continuous_on U f'" and "x0 \ U" and invf: "invf o\<^sub>L f' x0 = id_blinfun" obtains U' V g g' where "open U'" "U' \ U" "x0 \ U'" "open V" "f x0 \ V" "homeomorphism U' V f g" "\y. y \ V \ (g has_derivative (g' y)) (at y)" "\y. y \ V \ g' y = inv (blinfun_apply (f'(g y)))" "\y. y \ V \ bij (blinfun_apply (f'(g y)))" proof - have apply1 [simp]: "\i. blinfun_apply invf (blinfun_apply (f' x0) i) = i" by (metis blinfun_apply_blinfun_compose blinfun_apply_id_blinfun invf) have apply2 [simp]: "\i. blinfun_apply (f' x0) (blinfun_apply invf i) = i" by (metis apply1 bij_inv_eq_iff blinfun_bij1 invf) have [simp]: "(range (blinfun_apply invf)) = UNIV" using apply1 surjI by blast let ?f = "invf \ (\x. (f \ (+)x0)x - f x0)" let ?f' = "\x. invf o\<^sub>L (f' (x + x0))" obtain U' V g g' where "open U'" and U': "U' \ (+)(-x0) ` U" "0 \ U'" and "open V" "0 \ V" and hom: "homeomorphism U' V ?f g" and derg: "\y. y \ V \ (g has_derivative (g' y)) (at y)" and g': "\y. y \ V \ g' y = inv (?f'(g y))" and bij: "\y. y \ V \ bij (?f'(g y))" proof (rule inverse_function_theorem_scaled [of "(+)(-x0) ` U" ?f "?f'"]) show ope: "open ((+) (- x0) ` U)" using \open U\ open_translation by blast show "(?f has_derivative blinfun_apply (?f' x)) (at x)" if "x \ (+) (- x0) ` U" for x using that apply clarify apply (rule derf derivative_eq_intros | simp add: blinfun_compose.rep_eq)+ done have YY: "(\x. f' (x + x0)) \u-x0\ f' u" if "f' \u\ f' u" "u \ U" for u using that LIM_offset [where k = x0] by (auto simp: algebra_simps) then have "continuous_on ((+) (- x0) ` U) (\x. f' (x + x0))" using contf \open U\ Lim_at_imp_Lim_at_within by (fastforce simp: continuous_on_def at_within_open_NO_MATCH ope) then show "continuous_on ((+) (- x0) ` U) ?f'" by (intro continuous_intros) simp qed (auto simp: invf \x0 \ U\) show thesis proof let ?U' = "(+)x0 ` U'" let ?V = "((+)(f x0) \ f' x0) ` V" let ?g = "(+)x0 \ g \ invf \ (+)(- f x0)" let ?g' = "\y. inv (blinfun_apply (f' (?g y)))" show oU': "open ?U'" by (simp add: \open U'\ open_translation) show subU: "?U' \ U" using ComplI \U' \ (+) (- x0) ` U\ by auto show "x0 \ ?U'" by (simp add: \0 \ U'\) show "open ?V" using blinfun_bij2 [OF invf] by (metis \open V\ bij_is_surj blinfun.bounded_linear_right bounded_linear_def image_comp open_surjective_linear_image open_translation) show "f x0 \ ?V" using \0 \ V\ image_iff by fastforce show "homeomorphism ?U' ?V f ?g" proof show "continuous_on ?U' f" by (meson subU continuous_on_eq_continuous_at derf has_derivative_continuous oU' subsetD) have "?f ` U' \ V" using hom homeomorphism_image1 by blast then show "f ` ?U' \ ?V" unfolding image_subset_iff by (clarsimp simp: image_def) (metis apply2 add.commute diff_add_cancel) show "?g ` ?V \ ?U'" using hom invf by (auto simp: image_def homeomorphism_def) show "?g (f x) = x" if "x \ ?U'" for x using that hom homeomorphism_apply1 by fastforce have "continuous_on V g" using hom homeomorphism_def by blast then show "continuous_on ?V ?g" by (intro continuous_intros) (auto elim!: continuous_on_subset) have fg: "?f (g x) = x" if "x \ V" for x using hom homeomorphism_apply2 that by blast show "f (?g y) = y" if "y \ ?V" for y using that fg by (simp add: image_iff) (metis apply2 add.commute diff_add_cancel) qed show "(?g has_derivative ?g' y) (at y)" "bij (blinfun_apply (f' (?g y)))" if "y \ ?V" for y proof - have 1: "bij (blinfun_apply invf)" using blinfun_bij1 invf by blast then have 2: "bij (blinfun_apply (f' (x0 + g x)))" if "x \ V" for x by (metis add.commute bij bij_betw_comp_iff2 blinfun_compose.rep_eq that top_greatest) then show "bij (blinfun_apply (f' (?g y)))" using that by auto have "g' x \ blinfun_apply invf = inv (blinfun_apply (f' (x0 + g x)))" if "x \ V" for x using that by (simp add: g' o_inv_distrib blinfun_compose.rep_eq 1 2 add.commute bij_is_inj flip: o_assoc) then show "(?g has_derivative ?g' y) (at y)" using that invf by clarsimp (rule derg derivative_eq_intros | simp flip: id_def)+ qed qed auto qed subsection\<^marker>\tag unimportant\ \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 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 \ 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 \ (\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" and fd: "\x. x \ {a..c} - S \ f differentiable at x within {a..c}" and gd: "\x. 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 proof (rule differentiable_transform_within [where d = "dist x c"]) have "f differentiable at x" using x le fd [of x] at_within_interior [of x "{a..c}"] by simp then show "f differentiable at x within {a..b}" by (simp add: differentiable_at_withinI) qed (use x le st dist_real_def in auto) next case ge show ?diff_fg proof (rule differentiable_transform_within [where d = "dist x c"]) have "g differentiable at x" using x ge gd [of x] at_within_interior [of x "{c..b}"] by simp then show "g differentiable at x within {a..b}" by (simp add: differentiable_at_withinI) qed (use x ge st dist_real_def in auto) qed } then have "\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) subsection\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\ldots{} [namely] continuity plus differentiability except on a finite set\ldots{} [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\<^marker>\tag important\ 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))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding C1_differentiable_on_def by (metis (no_types, lifting) continuous_on_eq differentiableI_vector vector_derivative_at) next assume ?rhs then show ?lhs using C1_differentiable_on_def vector_derivative_works by fastforce qed 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: assumes fg: "f C1_differentiable_on S" "g C1_differentiable_on (f ` S)" and fin: "\x. finite (S \ f-`{x})" shows "(g \ f) C1_differentiable_on S" proof - have "\x. x \ S \ g \ f differentiable at x" by (meson C1_differentiable_on_eq assms differentiable_chain_at imageI) moreover have "continuous_on S (\x. vector_derivative (g \ f) (at x))" proof (rule continuous_on_eq [of _ "\x. vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x))"]) show "continuous_on S (\x. vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x)))" using fg apply (clarsimp simp add: C1_differentiable_on_eq) apply (rule Limits.continuous_on_scaleR, assumption) by (metis (mono_tags, lifting) continuous_at_imp_continuous_on continuous_on_compose continuous_on_cong differentiable_imp_continuous_within o_def) show "\x. x \ S \ vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x)) = vector_derivative (g \ f) (at x)" by (metis (mono_tags, hide_lams) C1_differentiable_on_eq fg imageI vector_derivative_chain_at) qed ultimately show ?thesis by (simp add: C1_differentiable_on_eq) qed 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) lemma C1_differentiable_on_const [simp, derivative_intros]: "(\z. a) C1_differentiable_on S" by (auto simp: C1_differentiable_on_eq) 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)+ +lemma C1_differentiable_on_of_real [derivative_intros]: "of_real C1_differentiable_on S" + unfolding C1_differentiable_on_def + by (smt (verit, del_insts) DERIV_ident UNIV_I continuous_on_const has_vector_derivative_of_real has_vector_derivative_transform) + definition\<^marker>\tag important\ 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_withinI) -lemma piecewise_C1_differentiable_compose: +lemma piecewise_C1_differentiable_compose [derivative_intros]: assumes fg: "f piecewise_C1_differentiable_on S" "g piecewise_C1_differentiable_on (f ` S)" and fin: "\x. finite (S \ f-`{x})" shows "(g \ f) piecewise_C1_differentiable_on S" proof - have "continuous_on S (\x. g (f x))" by (metis continuous_on_compose2 fg order_refl piecewise_C1_differentiable_on_def) moreover have "\T. finite T \ g \ f C1_differentiable_on S - T" proof - obtain F where "finite F" and F: "f C1_differentiable_on S - F" and f: "f piecewise_C1_differentiable_on S" using fg by (auto simp: piecewise_C1_differentiable_on_def) obtain G where "finite G" and G: "g C1_differentiable_on f ` S - G" and g: "g piecewise_C1_differentiable_on f ` S" using fg by (auto simp: piecewise_C1_differentiable_on_def) show ?thesis proof (intro exI conjI) show "finite (F \ (\x\G. S \ f-`{x}))" using fin by (auto simp only: Int_Union \finite F\ \finite G\ finite_UN finite_imageI) show "g \ f C1_differentiable_on S - (F \ (\x\G. S \ f -` {x}))" apply (rule C1_differentiable_compose) apply (blast intro: C1_differentiable_on_subset [OF F]) apply (blast intro: C1_differentiable_on_subset [OF G]) by (simp add: C1_differentiable_on_subset G Diff_Int_distrib2 fin) qed qed ultimately show ?thesis by (simp add: piecewise_C1_differentiable_on_def) qed 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 {}" +lemma C1_differentiable_on_empty [iff,derivative_intros]: "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 \ (\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) next case False have *: "\x. finite (S \ {y. m * y + c = x})" using False not_finite_existsD by fastforce show ?thesis apply (rule piecewise_C1_differentiable_compose [OF C1_differentiable_imp_piecewise]) apply (rule * assms derivative_intros | simp add: False vimage_def)+ done qed -lemma piecewise_C1_differentiable_cases: +lemma piecewise_C1_differentiable_cases [derivative_intros]: 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))" proof - have "((\x. if x \ c then f x else g x) has_vector_derivative vector_derivative f (at x)) (at x)" if "a < x" "x < c" "x \ S" for x proof - have f: "f differentiable at x" by (meson C1_differentiable_on_eq Diff_iff atLeastAtMost_iff less_eq_real_def st(1) that) show ?thesis using that apply (rule_tac f=f and d="dist x c" in has_vector_derivative_transform_within) apply (auto simp: dist_norm vector_derivative_works [symmetric] f) done qed then show ?thesis by (metis (no_types, lifting) continuous_on_eq [OF fcon] DiffE greaterThanLessThan_iff vector_derivative_at) qed moreover have "continuous_on ({c<..x. vector_derivative (\x. if x \ c then f x else g x) (at x))" proof - have "((\x. if x \ c then f x else g x) has_vector_derivative vector_derivative g (at x)) (at x)" if "c < x" "x < b" "x \ T" for x proof - have g: "g differentiable at x" by (metis C1_differentiable_on_eq DiffD1 DiffI atLeastAtMost_diff_ends greaterThanLessThan_iff st(2) that) show ?thesis using that apply (rule_tac f=g and d="dist x c" in has_vector_derivative_transform_within) apply (auto simp: dist_norm vector_derivative_works [symmetric] g) done qed then show ?thesis by (metis (no_types, lifting) continuous_on_eq [OF gcon] DiffE greaterThanLessThan_iff vector_derivative_at) qed ultimately have "continuous_on ({a<.. T)) (\x. vector_derivative (\x. if x \ c then f x else g x) (at x))" by (rule continuous_on_subset [OF continuous_on_open_Un], auto) } 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: +lemma piecewise_C1_differentiable_const [derivative_intros]: + "(\x. c) piecewise_C1_differentiable_on S" + by (simp add: C1_differentiable_imp_piecewise) + +lemma piecewise_C1_differentiable_scaleR [derivative_intros]: + "\f piecewise_C1_differentiable_on S\ + \ (\x. c *\<^sub>R f x) piecewise_C1_differentiable_on S" + by (force simp add: piecewise_C1_differentiable_on_def continuous_on_scaleR) + +lemma piecewise_C1_differentiable_neg [derivative_intros]: "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: +lemma piecewise_C1_differentiable_add [derivative_intros]: 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: +lemma piecewise_C1_differentiable_diff [derivative_intros]: "\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_cmult_right [derivative_intros]: + fixes c::complex + shows "f piecewise_C1_differentiable_on S + \ (\x. f x * c) piecewise_C1_differentiable_on S" + by (force simp: piecewise_C1_differentiable_on_def continuous_on_mult_right) + +lemma piecewise_C1_differentiable_cmult_left [derivative_intros]: + fixes c::complex + shows "f piecewise_C1_differentiable_on S + \ (\x. c * f x) piecewise_C1_differentiable_on S" + using piecewise_C1_differentiable_cmult_right [of f S c] by (simp add: mult.commute) + +lemma piecewise_C1_differentiable_on_of_real [derivative_intros]: + "of_real piecewise_C1_differentiable_on S" + by (simp add: C1_differentiable_imp_piecewise C1_differentiable_on_of_real) + end diff --git a/src/HOL/Analysis/Linear_Algebra.thy b/src/HOL/Analysis/Linear_Algebra.thy --- a/src/HOL/Analysis/Linear_Algebra.thy +++ b/src/HOL/Analysis/Linear_Algebra.thy @@ -1,1878 +1,1881 @@ (* Title: HOL/Analysis/Linear_Algebra.thy Author: Amine Chaieb, University of Cambridge *) section \Elementary Linear Algebra on Euclidean Spaces\ theory Linear_Algebra imports Euclidean_Space "HOL-Library.Infinite_Set" begin 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.neg) show "f (s *\<^sub>R v) = s *\<^sub>R (f v)" by (rule f.scale) qed lemma finite_Atleast_Atmost_nat[simp]: "finite {f x |x. x \ (UNIV::'a::finite set)}" using finite finite_image_set by blast lemma substdbasis_expansion_unique: includes inner_syntax assumes d: "d \ Basis" shows "(\i\d. f i *\<^sub>R i) = (x::'a::euclidean_space) \ (\i\Basis. (i \ d \ f i = x \ i) \ (i \ d \ x \ i = 0))" proof - have *: "\x a b P. x * (if P then a else b) = (if P then x * a else x * b)" by auto have **: "finite d" by (auto intro: finite_subset[OF assms]) have ***: "\i. i \ Basis \ (\i\d. f i *\<^sub>R i) \ i = (\x\d. if x = i then f x else 0)" using d by (auto intro!: sum.cong simp: inner_Basis inner_sum_left) show ?thesis unfolding euclidean_eq_iff[where 'a='a] by (auto simp: sum.delta[OF **] ***) qed lemma independent_substdbasis: "d \ Basis \ independent d" by (rule independent_mono[OF independent_Basis]) lemma subset_translation_eq [simp]: fixes a :: "'a::real_vector" shows "(+) a ` s \ (+) a ` t \ s \ t" by auto lemma translate_inj_on: fixes A :: "'a::ab_group_add set" shows "inj_on (\x. a + x) A" unfolding inj_on_def by auto lemma translation_assoc: fixes a b :: "'a::ab_group_add" shows "(\x. b + x) ` ((\x. a + x) ` S) = (\x. (a + b) + x) ` S" by auto lemma translation_invert: fixes a :: "'a::ab_group_add" assumes "(\x. a + x) ` A = (\x. a + x) ` B" shows "A = B" proof - have "(\x. -a + x) ` ((\x. a + x) ` A) = (\x. - a + x) ` ((\x. a + x) ` B)" using assms by auto then show ?thesis using translation_assoc[of "-a" a A] translation_assoc[of "-a" a B] by auto qed lemma translation_galois: fixes a :: "'a::ab_group_add" shows "T = ((\x. a + x) ` S) \ S = ((\x. (- a) + x) ` T)" using translation_assoc[of "-a" a S] apply auto using translation_assoc[of a "-a" T] apply auto done lemma translation_inverse_subset: assumes "((\x. - a + x) ` V) \ (S :: 'n::ab_group_add set)" shows "V \ ((\x. a + x) ` S)" proof - { fix x assume "x \ V" then have "x-a \ S" using assms by auto then have "x \ {a + v |v. v \ S}" apply auto apply (rule exI[of _ "x-a"], simp) done then have "x \ ((\x. a+x) ` S)" by auto } then show ?thesis by auto qed subsection\<^marker>\tag unimportant\ \More interesting properties of the norm\ unbundle inner_syntax text\Equality of vectors in terms of \<^term>\(\)\ products.\ lemma linear_componentwise: fixes f:: "'a::euclidean_space \ 'b::real_inner" assumes lf: "linear f" shows "(f x) \ j = (\i\Basis. (x\i) * (f i\j))" (is "?lhs = ?rhs") proof - interpret linear f by fact have "?rhs = (\i\Basis. (x\i) *\<^sub>R (f i))\j" by (simp add: inner_sum_left) then show ?thesis by (simp add: euclidean_representation sum[symmetric] scale[symmetric]) qed lemma vector_eq: "x = y \ x \ x = x \ y \ y \ y = x \ x" (is "?lhs \ ?rhs") proof assume ?lhs then show ?rhs by simp next assume ?rhs then have "x \ x - x \ y = 0 \ x \ y - y \ y = 0" by simp then have "x \ (x - y) = 0 \ y \ (x - y) = 0" by (simp add: inner_diff inner_commute) then have "(x - y) \ (x - y) = 0" by (simp add: field_simps inner_diff inner_commute) then show "x = y" by simp qed lemma norm_triangle_half_r: "norm (y - x1) < e / 2 \ norm (y - x2) < e / 2 \ norm (x1 - x2) < e" using dist_triangle_half_r unfolding dist_norm[symmetric] by auto lemma norm_triangle_half_l: assumes "norm (x - y) < e / 2" and "norm (x' - y) < e / 2" shows "norm (x - x') < e" using dist_triangle_half_l[OF assms[unfolded dist_norm[symmetric]]] unfolding dist_norm[symmetric] . lemma abs_triangle_half_r: fixes y :: "'a::linordered_field" shows "abs (y - x1) < e / 2 \ abs (y - x2) < e / 2 \ abs (x1 - x2) < e" by linarith lemma abs_triangle_half_l: fixes y :: "'a::linordered_field" assumes "abs (x - y) < e / 2" and "abs (x' - y) < e / 2" shows "abs (x - x') < e" using assms by linarith lemma sum_clauses: shows "sum f {} = 0" and "finite S \ sum f (insert x S) = (if x \ S then sum f S else f x + sum f S)" by (auto simp add: insert_absorb) lemma vector_eq_ldot: "(\x. x \ y = x \ z) \ y = z" proof assume "\x. x \ y = x \ z" then have "\x. x \ (y - z) = 0" by (simp add: inner_diff) then have "(y - z) \ (y - z) = 0" .. then show "y = z" by simp qed simp lemma vector_eq_rdot: "(\z. x \ z = y \ z) \ x = y" proof assume "\z. x \ z = y \ z" then have "\z. (x - y) \ z = 0" by (simp add: inner_diff) then have "(x - y) \ (x - y) = 0" .. then show "x = y" by simp qed simp subsection \Substandard Basis\ lemma ex_card: assumes "n \ card A" shows "\S\A. card S = n" proof (cases "finite A") case True from ex_bij_betw_nat_finite[OF this] obtain f where f: "bij_betw f {0..n \ card A\ have "{..< n} \ {..< card A}" "inj_on f {..< n}" by (auto simp: bij_betw_def intro: subset_inj_on) ultimately have "f ` {..< n} \ A" "card (f ` {..< n}) = n" by (auto simp: bij_betw_def card_image) then show ?thesis by blast next case False with \n \ card A\ show ?thesis by force qed lemma subspace_substandard: "subspace {x::'a::euclidean_space. (\i\Basis. P i \ x\i = 0)}" by (auto simp: subspace_def inner_add_left) lemma dim_substandard: assumes d: "d \ Basis" shows "dim {x::'a::euclidean_space. \i\Basis. i \ d \ x\i = 0} = card d" (is "dim ?A = _") proof (rule dim_unique) from d show "d \ ?A" by (auto simp: inner_Basis) from d show "independent d" by (rule independent_mono [OF independent_Basis]) have "x \ span d" if "\i\Basis. i \ d \ x \ i = 0" for x proof - have "finite d" by (rule finite_subset [OF d finite_Basis]) then have "(\i\d. (x \ i) *\<^sub>R i) \ span d" by (simp add: span_sum span_clauses) also have "(\i\d. (x \ i) *\<^sub>R i) = (\i\Basis. (x \ i) *\<^sub>R i)" by (rule sum.mono_neutral_cong_left [OF finite_Basis d]) (auto simp: that) finally show "x \ span d" by (simp only: euclidean_representation) qed then show "?A \ span d" by auto qed simp subsection \Orthogonality\ definition\<^marker>\tag important\ (in real_inner) "orthogonal x y \ x \ y = 0" context real_inner begin lemma orthogonal_self: "orthogonal x x \ x = 0" by (simp add: orthogonal_def) lemma orthogonal_clauses: "orthogonal a 0" "orthogonal a x \ orthogonal a (c *\<^sub>R x)" "orthogonal a x \ orthogonal a (- x)" "orthogonal a x \ orthogonal a y \ orthogonal a (x + y)" "orthogonal a x \ orthogonal a y \ orthogonal a (x - y)" "orthogonal 0 a" "orthogonal x a \ orthogonal (c *\<^sub>R x) a" "orthogonal x a \ orthogonal (- x) a" "orthogonal x a \ orthogonal y a \ orthogonal (x + y) a" "orthogonal x a \ orthogonal y a \ orthogonal (x - y) a" unfolding orthogonal_def inner_add inner_diff by auto end lemma orthogonal_commute: "orthogonal x y \ orthogonal y x" by (simp add: orthogonal_def inner_commute) lemma orthogonal_scaleR [simp]: "c \ 0 \ orthogonal (c *\<^sub>R x) = orthogonal x" by (rule ext) (simp add: orthogonal_def) lemma pairwise_ortho_scaleR: "pairwise (\i j. orthogonal (f i) (g j)) B \ pairwise (\i j. orthogonal (a i *\<^sub>R f i) (a j *\<^sub>R g j)) B" by (auto simp: pairwise_def orthogonal_clauses) lemma orthogonal_rvsum: "\finite s; \y. y \ s \ orthogonal x (f y)\ \ orthogonal x (sum f s)" by (induction s rule: finite_induct) (auto simp: orthogonal_clauses) lemma orthogonal_lvsum: "\finite s; \x. x \ s \ orthogonal (f x) y\ \ orthogonal (sum f s) y" by (induction s rule: finite_induct) (auto simp: orthogonal_clauses) lemma norm_add_Pythagorean: assumes "orthogonal a b" shows "norm(a + b) ^ 2 = norm a ^ 2 + norm b ^ 2" proof - from assms have "(a - (0 - b)) \ (a - (0 - b)) = a \ a - (0 - b \ b)" by (simp add: algebra_simps orthogonal_def inner_commute) then show ?thesis by (simp add: power2_norm_eq_inner) qed lemma norm_sum_Pythagorean: assumes "finite I" "pairwise (\i j. orthogonal (f i) (f j)) I" shows "(norm (sum f I))\<^sup>2 = (\i\I. (norm (f i))\<^sup>2)" using assms proof (induction I rule: finite_induct) case empty then show ?case by simp next case (insert x I) then have "orthogonal (f x) (sum f I)" by (metis pairwise_insert orthogonal_rvsum) with insert show ?case by (simp add: pairwise_insert norm_add_Pythagorean) qed subsection \Orthogonality of a transformation\ definition\<^marker>\tag important\ "orthogonal_transformation f \ linear f \ (\v w. f v \ f w = v \ w)" lemma\<^marker>\tag unimportant\ orthogonal_transformation: "orthogonal_transformation f \ linear f \ (\v. norm (f v) = norm v)" unfolding orthogonal_transformation_def apply auto apply (erule_tac x=v in allE)+ apply (simp add: norm_eq_sqrt_inner) apply (simp add: dot_norm linear_add[symmetric]) done lemma\<^marker>\tag unimportant\ orthogonal_transformation_id [simp]: "orthogonal_transformation (\x. x)" by (simp add: linear_iff orthogonal_transformation_def) lemma\<^marker>\tag unimportant\ orthogonal_orthogonal_transformation: "orthogonal_transformation f \ orthogonal (f x) (f y) \ orthogonal x y" by (simp add: orthogonal_def orthogonal_transformation_def) lemma\<^marker>\tag unimportant\ orthogonal_transformation_compose: "\orthogonal_transformation f; orthogonal_transformation g\ \ orthogonal_transformation(f \ g)" by (auto simp: orthogonal_transformation_def linear_compose) lemma\<^marker>\tag unimportant\ orthogonal_transformation_neg: "orthogonal_transformation(\x. -(f x)) \ orthogonal_transformation f" by (auto simp: orthogonal_transformation_def dest: linear_compose_neg) lemma\<^marker>\tag unimportant\ orthogonal_transformation_scaleR: "orthogonal_transformation f \ f (c *\<^sub>R v) = c *\<^sub>R f v" by (simp add: linear_iff orthogonal_transformation_def) lemma\<^marker>\tag unimportant\ orthogonal_transformation_linear: "orthogonal_transformation f \ linear f" by (simp add: orthogonal_transformation_def) lemma\<^marker>\tag unimportant\ orthogonal_transformation_inj: "orthogonal_transformation f \ inj f" unfolding orthogonal_transformation_def inj_on_def by (metis vector_eq) lemma\<^marker>\tag unimportant\ orthogonal_transformation_surj: "orthogonal_transformation f \ surj f" for f :: "'a::euclidean_space \ 'a::euclidean_space" by (simp add: linear_injective_imp_surjective orthogonal_transformation_inj orthogonal_transformation_linear) lemma\<^marker>\tag unimportant\ orthogonal_transformation_bij: "orthogonal_transformation f \ bij f" for f :: "'a::euclidean_space \ 'a::euclidean_space" by (simp add: bij_def orthogonal_transformation_inj orthogonal_transformation_surj) lemma\<^marker>\tag unimportant\ orthogonal_transformation_inv: "orthogonal_transformation f \ orthogonal_transformation (inv f)" for f :: "'a::euclidean_space \ 'a::euclidean_space" by (metis (no_types, hide_lams) bijection.inv_right bijection_def inj_linear_imp_inv_linear orthogonal_transformation orthogonal_transformation_bij orthogonal_transformation_inj) lemma\<^marker>\tag unimportant\ orthogonal_transformation_norm: "orthogonal_transformation f \ norm (f x) = norm x" by (metis orthogonal_transformation) subsection \Bilinear functions\ definition\<^marker>\tag important\ bilinear :: "('a::real_vector \ 'b::real_vector \ 'c::real_vector) \ bool" where "bilinear f \ (\x. linear (\y. f x y)) \ (\y. linear (\x. f x y))" lemma bilinear_ladd: "bilinear h \ h (x + y) z = h x z + h y z" by (simp add: bilinear_def linear_iff) lemma bilinear_radd: "bilinear h \ h x (y + z) = h x y + h x z" by (simp add: bilinear_def linear_iff) lemma bilinear_times: fixes c::"'a::real_algebra" shows "bilinear (\x y::'a. x*y)" by (auto simp: bilinear_def distrib_left distrib_right intro!: linearI) lemma bilinear_lmul: "bilinear h \ h (c *\<^sub>R x) y = c *\<^sub>R h x y" by (simp add: bilinear_def linear_iff) lemma bilinear_rmul: "bilinear h \ h x (c *\<^sub>R y) = c *\<^sub>R h x y" by (simp add: bilinear_def linear_iff) lemma bilinear_lneg: "bilinear h \ h (- x) y = - h x y" by (drule bilinear_lmul [of _ "- 1"]) simp lemma bilinear_rneg: "bilinear h \ h x (- y) = - h x y" by (drule bilinear_rmul [of _ _ "- 1"]) simp lemma (in ab_group_add) eq_add_iff: "x = x + y \ y = 0" using add_left_imp_eq[of x y 0] by auto lemma bilinear_lzero: assumes "bilinear h" shows "h 0 x = 0" using bilinear_ladd [OF assms, of 0 0 x] by (simp add: eq_add_iff field_simps) lemma bilinear_rzero: assumes "bilinear h" shows "h x 0 = 0" using bilinear_radd [OF assms, of x 0 0 ] by (simp add: eq_add_iff field_simps) lemma bilinear_lsub: "bilinear h \ h (x - y) z = h x z - h y z" using bilinear_ladd [of h x "- y"] by (simp add: bilinear_lneg) lemma bilinear_rsub: "bilinear h \ h z (x - y) = h z x - h z y" using bilinear_radd [of h _ x "- y"] by (simp add: bilinear_rneg) lemma bilinear_sum: assumes "bilinear h" shows "h (sum f S) (sum g T) = sum (\(i,j). h (f i) (g j)) (S \ T) " proof - interpret l: linear "\x. h x y" for y using assms by (simp add: bilinear_def) interpret r: linear "\y. h x y" for x using assms by (simp add: bilinear_def) have "h (sum f S) (sum g T) = sum (\x. h (f x) (sum g T)) S" by (simp add: l.sum) also have "\ = sum (\x. sum (\y. h (f x) (g y)) T) S" by (rule sum.cong) (simp_all add: r.sum) finally show ?thesis unfolding sum.cartesian_product . qed subsection \Adjoints\ definition\<^marker>\tag important\ adjoint :: "(('a::real_inner) \ ('b::real_inner)) \ 'b \ 'a" where "adjoint f = (SOME f'. \x y. f x \ y = x \ f' y)" lemma adjoint_unique: assumes "\x y. inner (f x) y = inner x (g y)" shows "adjoint f = g" unfolding adjoint_def proof (rule some_equality) show "\x y. inner (f x) y = inner x (g y)" by (rule assms) next fix h assume "\x y. inner (f x) y = inner x (h y)" then have "\x y. inner x (g y) = inner x (h y)" using assms by simp then have "\x y. inner x (g y - h y) = 0" by (simp add: inner_diff_right) then have "\y. inner (g y - h y) (g y - h y) = 0" by simp then have "\y. h y = g y" by simp then show "h = g" by (simp add: ext) qed text \TODO: The following lemmas about adjoints should hold for any Hilbert space (i.e. complete inner product space). (see \<^url>\https://en.wikipedia.org/wiki/Hermitian_adjoint\) \ lemma adjoint_works: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes lf: "linear f" shows "x \ adjoint f y = f x \ y" proof - interpret linear f by fact have "\y. \w. \x. f x \ y = x \ w" proof (intro allI exI) fix y :: "'m" and x let ?w = "(\i\Basis. (f i \ y) *\<^sub>R i) :: 'n" have "f x \ y = f (\i\Basis. (x \ i) *\<^sub>R i) \ y" by (simp add: euclidean_representation) also have "\ = (\i\Basis. (x \ i) *\<^sub>R f i) \ y" by (simp add: sum scale) finally show "f x \ y = x \ ?w" by (simp add: inner_sum_left inner_sum_right mult.commute) qed then show ?thesis unfolding adjoint_def choice_iff by (intro someI2_ex[where Q="\f'. x \ f' y = f x \ y"]) auto qed lemma adjoint_clauses: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes lf: "linear f" shows "x \ adjoint f y = f x \ y" and "adjoint f y \ x = y \ f x" by (simp_all add: adjoint_works[OF lf] inner_commute) lemma adjoint_linear: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes lf: "linear f" shows "linear (adjoint f)" by (simp add: lf linear_iff euclidean_eq_iff[where 'a='n] euclidean_eq_iff[where 'a='m] adjoint_clauses[OF lf] inner_distrib) lemma adjoint_adjoint: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes lf: "linear f" shows "adjoint (adjoint f) = f" by (rule adjoint_unique, simp add: adjoint_clauses [OF lf]) subsection\<^marker>\tag unimportant\ \Euclidean Spaces as Typeclass\ lemma independent_Basis: "independent Basis" by (rule independent_Basis) lemma span_Basis [simp]: "span Basis = UNIV" by (rule span_Basis) lemma in_span_Basis: "x \ span Basis" unfolding span_Basis .. subsection\<^marker>\tag unimportant\ \Linearity and Bilinearity continued\ lemma linear_bounded: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes lf: "linear f" shows "\B. \x. norm (f x) \ B * norm x" proof interpret linear f by fact let ?B = "\b\Basis. norm (f b)" show "\x. norm (f x) \ ?B * norm x" proof fix x :: 'a let ?g = "\b. (x \ b) *\<^sub>R f b" have "norm (f x) = norm (f (\b\Basis. (x \ b) *\<^sub>R b))" unfolding euclidean_representation .. also have "\ = norm (sum ?g Basis)" by (simp add: sum scale) finally have th0: "norm (f x) = norm (sum ?g Basis)" . have th: "norm (?g i) \ norm (f i) * norm x" if "i \ Basis" for i proof - from Basis_le_norm[OF that, of x] show "norm (?g i) \ norm (f i) * norm x" unfolding norm_scaleR by (metis mult.commute mult_left_mono norm_ge_zero) qed from sum_norm_le[of _ ?g, OF th] show "norm (f x) \ ?B * norm x" unfolding th0 sum_distrib_right by metis qed qed lemma linear_conv_bounded_linear: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" shows "linear f \ bounded_linear f" proof assume "linear f" then interpret f: linear f . show "bounded_linear f" proof have "\B. \x. norm (f x) \ B * norm x" using \linear f\ by (rule linear_bounded) then show "\K. \x. norm (f x) \ norm x * K" by (simp add: mult.commute) qed next assume "bounded_linear f" then interpret f: bounded_linear f . show "linear f" .. qed lemmas linear_linear = linear_conv_bounded_linear[symmetric] lemma inj_linear_imp_inv_bounded_linear: fixes f::"'a::euclidean_space \ 'a" shows "\bounded_linear f; inj f\ \ bounded_linear (inv f)" by (simp add: inj_linear_imp_inv_linear linear_linear) lemma linear_bounded_pos: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes lf: "linear f" obtains B where "B > 0" "\x. norm (f x) \ B * norm x" proof - have "\B > 0. \x. norm (f x) \ norm x * B" using lf unfolding linear_conv_bounded_linear by (rule bounded_linear.pos_bounded) with that show ?thesis by (auto simp: mult.commute) qed lemma linear_invertible_bounded_below_pos: fixes f :: "'a::real_normed_vector \ 'b::euclidean_space" assumes "linear f" "linear g" "g \ f = id" obtains B where "B > 0" "\x. B * norm x \ norm(f x)" proof - obtain B where "B > 0" and B: "\x. norm (g x) \ B * norm x" using linear_bounded_pos [OF \linear g\] by blast show thesis proof show "0 < 1/B" by (simp add: \B > 0\) show "1/B * norm x \ norm (f x)" for x proof - have "1/B * norm x = 1/B * norm (g (f x))" using assms by (simp add: pointfree_idE) also have "\ \ norm (f x)" using B [of "f x"] by (simp add: \B > 0\ mult.commute pos_divide_le_eq) finally show ?thesis . qed qed qed lemma linear_inj_bounded_below_pos: fixes f :: "'a::real_normed_vector \ 'b::euclidean_space" assumes "linear f" "inj f" obtains B where "B > 0" "\x. B * norm x \ norm(f x)" using linear_injective_left_inverse [OF assms] linear_invertible_bounded_below_pos assms by blast lemma bounded_linearI': fixes f ::"'a::euclidean_space \ 'b::real_normed_vector" assumes "\x y. f (x + y) = f x + f y" and "\c x. f (c *\<^sub>R x) = c *\<^sub>R f x" shows "bounded_linear f" using assms linearI linear_conv_bounded_linear by blast lemma bilinear_bounded: fixes h :: "'m::euclidean_space \ 'n::euclidean_space \ 'k::real_normed_vector" assumes bh: "bilinear h" shows "\B. \x y. norm (h x y) \ B * norm x * norm y" proof (clarify intro!: exI[of _ "\i\Basis. \j\Basis. norm (h i j)"]) fix x :: 'm fix y :: 'n have "norm (h x y) = norm (h (sum (\i. (x \ i) *\<^sub>R i) Basis) (sum (\i. (y \ i) *\<^sub>R i) Basis))" by (simp add: euclidean_representation) also have "\ = norm (sum (\ (i,j). h ((x \ i) *\<^sub>R i) ((y \ j) *\<^sub>R j)) (Basis \ Basis))" unfolding bilinear_sum[OF bh] .. finally have th: "norm (h x y) = \" . have "\i j. \i \ Basis; j \ Basis\ \ \x \ i\ * (\y \ j\ * norm (h i j)) \ norm x * (norm y * norm (h i j))" by (auto simp add: zero_le_mult_iff Basis_le_norm mult_mono) then show "norm (h x y) \ (\i\Basis. \j\Basis. norm (h i j)) * norm x * norm y" unfolding sum_distrib_right th sum.cartesian_product by (clarsimp simp add: bilinear_rmul[OF bh] bilinear_lmul[OF bh] field_simps simp del: scaleR_scaleR intro!: sum_norm_le) qed lemma bilinear_conv_bounded_bilinear: fixes h :: "'a::euclidean_space \ 'b::euclidean_space \ 'c::real_normed_vector" shows "bilinear h \ bounded_bilinear h" proof assume "bilinear h" show "bounded_bilinear h" proof fix x y z show "h (x + y) z = h x z + h y z" using \bilinear h\ unfolding bilinear_def linear_iff by simp next fix x y z show "h x (y + z) = h x y + h x z" using \bilinear h\ unfolding bilinear_def linear_iff by simp next show "h (scaleR r x) y = scaleR r (h x y)" "h x (scaleR r y) = scaleR r (h x y)" for r x y using \bilinear h\ unfolding bilinear_def linear_iff by simp_all next have "\B. \x y. norm (h x y) \ B * norm x * norm y" using \bilinear h\ by (rule bilinear_bounded) then show "\K. \x y. norm (h x y) \ norm x * norm y * K" by (simp add: ac_simps) qed next assume "bounded_bilinear h" then interpret h: bounded_bilinear h . show "bilinear h" unfolding bilinear_def linear_conv_bounded_linear using h.bounded_linear_left h.bounded_linear_right by simp qed lemma bilinear_bounded_pos: fixes h :: "'a::euclidean_space \ 'b::euclidean_space \ 'c::real_normed_vector" assumes bh: "bilinear h" shows "\B > 0. \x y. norm (h x y) \ B * norm x * norm y" proof - have "\B > 0. \x y. norm (h x y) \ norm x * norm y * B" using bh [unfolded bilinear_conv_bounded_bilinear] by (rule bounded_bilinear.pos_bounded) then show ?thesis by (simp only: ac_simps) qed lemma bounded_linear_imp_has_derivative: "bounded_linear f \ (f has_derivative f) net" by (auto simp add: has_derivative_def linear_diff linear_linear linear_def dest: bounded_linear.linear) lemma linear_imp_has_derivative: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" shows "linear f \ (f has_derivative f) net" by (simp add: bounded_linear_imp_has_derivative linear_conv_bounded_linear) lemma bounded_linear_imp_differentiable: "bounded_linear f \ f differentiable net" using bounded_linear_imp_has_derivative differentiable_def by blast lemma linear_imp_differentiable: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" shows "linear f \ f differentiable net" by (metis linear_imp_has_derivative differentiable_def) +lemma of_real_differentiable [simp,derivative_intros]: "of_real differentiable F" + by (simp add: bounded_linear_imp_differentiable bounded_linear_of_real) + subsection\<^marker>\tag unimportant\ \We continue\ lemma independent_bound: fixes S :: "'a::euclidean_space set" shows "independent S \ finite S \ card S \ DIM('a)" by (metis dim_subset_UNIV finiteI_independent dim_span_eq_card_independent) lemmas independent_imp_finite = finiteI_independent corollary\<^marker>\tag unimportant\ independent_card_le: fixes S :: "'a::euclidean_space set" assumes "independent S" shows "card S \ DIM('a)" using assms independent_bound by auto lemma dependent_biggerset: fixes S :: "'a::euclidean_space set" shows "(finite S \ card S > DIM('a)) \ dependent S" by (metis independent_bound not_less) text \Picking an orthogonal replacement for a spanning set.\ lemma vector_sub_project_orthogonal: fixes b x :: "'a::euclidean_space" shows "b \ (x - ((b \ x) / (b \ b)) *\<^sub>R b) = 0" unfolding inner_simps by auto lemma pairwise_orthogonal_insert: assumes "pairwise orthogonal S" and "\y. y \ S \ orthogonal x y" shows "pairwise orthogonal (insert x S)" using assms unfolding pairwise_def by (auto simp add: orthogonal_commute) lemma basis_orthogonal: fixes B :: "'a::real_inner set" assumes fB: "finite B" shows "\C. finite C \ card C \ card B \ span C = span B \ pairwise orthogonal C" (is " \C. ?P B C") using fB proof (induct rule: finite_induct) case empty then show ?case apply (rule exI[where x="{}"]) apply (auto simp add: pairwise_def) done next case (insert a B) note fB = \finite B\ and aB = \a \ B\ from \\C. finite C \ card C \ card B \ span C = span B \ pairwise orthogonal C\ obtain C where C: "finite C" "card C \ card B" "span C = span B" "pairwise orthogonal C" by blast let ?a = "a - sum (\x. (x \ a / (x \ x)) *\<^sub>R x) C" let ?C = "insert ?a C" from C(1) have fC: "finite ?C" by simp from fB aB C(1,2) have cC: "card ?C \ card (insert a B)" by (simp add: card_insert_if) { fix x k have th0: "\(a::'a) b c. a - (b - c) = c + (a - b)" by (simp add: field_simps) have "x - k *\<^sub>R (a - (\x\C. (x \ a / (x \ x)) *\<^sub>R x)) \ span C \ x - k *\<^sub>R a \ span C" apply (simp only: scaleR_right_diff_distrib th0) apply (rule span_add_eq) apply (rule span_scale) apply (rule span_sum) apply (rule span_scale) apply (rule span_base) apply assumption done } then have SC: "span ?C = span (insert a B)" unfolding set_eq_iff span_breakdown_eq C(3)[symmetric] by auto { fix y assume yC: "y \ C" then have Cy: "C = insert y (C - {y})" by blast have fth: "finite (C - {y})" using C by simp have "orthogonal ?a y" unfolding orthogonal_def unfolding inner_diff inner_sum_left right_minus_eq unfolding sum.remove [OF \finite C\ \y \ C\] apply (clarsimp simp add: inner_commute[of y a]) apply (rule sum.neutral) apply clarsimp apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format]) using \y \ C\ by auto } with \pairwise orthogonal C\ have CPO: "pairwise orthogonal ?C" by (rule pairwise_orthogonal_insert) from fC cC SC CPO have "?P (insert a B) ?C" by blast then show ?case by blast qed lemma orthogonal_basis_exists: fixes V :: "('a::euclidean_space) set" shows "\B. independent B \ B \ span V \ V \ span B \ (card B = dim V) \ pairwise orthogonal B" proof - from basis_exists[of V] obtain B where B: "B \ V" "independent B" "V \ span B" "card B = dim V" by force from B have fB: "finite B" "card B = dim V" using independent_bound by auto from basis_orthogonal[OF fB(1)] obtain C where C: "finite C" "card C \ card B" "span C = span B" "pairwise orthogonal C" by blast from C B have CSV: "C \ span V" by (metis span_superset span_mono subset_trans) from span_mono[OF B(3)] C have SVC: "span V \ span C" by (simp add: span_span) from card_le_dim_spanning[OF CSV SVC C(1)] C(2,3) fB have iC: "independent C" by (simp) from C fB have "card C \ dim V" by simp moreover have "dim V \ card C" using span_card_ge_dim[OF CSV SVC C(1)] by simp ultimately have CdV: "card C = dim V" using C(1) by simp from C B CSV CdV iC show ?thesis by auto qed text \Low-dimensional subset is in a hyperplane (weak orthogonal complement).\ lemma span_not_univ_orthogonal: fixes S :: "'a::euclidean_space set" assumes sU: "span S \ UNIV" shows "\a::'a. a \ 0 \ (\x \ span S. a \ x = 0)" proof - from sU obtain a where a: "a \ span S" by blast from orthogonal_basis_exists obtain B where B: "independent B" "B \ span S" "S \ span B" "card B = dim S" "pairwise orthogonal B" by blast from B have fB: "finite B" "card B = dim S" using independent_bound by auto from span_mono[OF B(2)] span_mono[OF B(3)] have sSB: "span S = span B" by (simp add: span_span) let ?a = "a - sum (\b. (a \ b / (b \ b)) *\<^sub>R b) B" have "sum (\b. (a \ b / (b \ b)) *\<^sub>R b) B \ span S" unfolding sSB apply (rule span_sum) apply (rule span_scale) apply (rule span_base) apply assumption done with a have a0:"?a \ 0" by auto have "?a \ x = 0" if "x\span B" for x proof (rule span_induct [OF that]) show "subspace {x. ?a \ x = 0}" by (auto simp add: subspace_def inner_add) next { fix x assume x: "x \ B" from x have B': "B = insert x (B - {x})" by blast have fth: "finite (B - {x})" using fB by simp have "?a \ x = 0" apply (subst B') using fB fth unfolding sum_clauses(2)[OF fth] apply simp unfolding inner_simps apply (clarsimp simp add: inner_add inner_sum_left) apply (rule sum.neutral, rule ballI) apply (simp only: inner_commute) apply (auto simp add: x field_simps intro: B(5)[unfolded pairwise_def orthogonal_def, rule_format]) done } then show "?a \ x = 0" if "x \ B" for x using that by blast qed with a0 show ?thesis unfolding sSB by (auto intro: exI[where x="?a"]) qed lemma span_not_univ_subset_hyperplane: fixes S :: "'a::euclidean_space set" assumes SU: "span S \ UNIV" shows "\ a. a \0 \ span S \ {x. a \ x = 0}" using span_not_univ_orthogonal[OF SU] by auto lemma lowdim_subset_hyperplane: fixes S :: "'a::euclidean_space set" assumes d: "dim S < DIM('a)" shows "\a::'a. a \ 0 \ span S \ {x. a \ x = 0}" proof - { assume "span S = UNIV" then have "dim (span S) = dim (UNIV :: ('a) set)" by simp then have "dim S = DIM('a)" by (metis Euclidean_Space.dim_UNIV dim_span) with d have False by arith } then have th: "span S \ UNIV" by blast from span_not_univ_subset_hyperplane[OF th] show ?thesis . qed lemma linear_eq_stdbasis: fixes f :: "'a::euclidean_space \ _" assumes lf: "linear f" and lg: "linear g" and fg: "\b. b \ Basis \ f b = g b" shows "f = g" using linear_eq_on_span[OF lf lg, of Basis] fg by auto text \Similar results for bilinear functions.\ lemma bilinear_eq: assumes bf: "bilinear f" and bg: "bilinear g" and SB: "S \ span B" and TC: "T \ span C" and "x\S" "y\T" and fg: "\x y. \x \ B; y\ C\ \ f x y = g x y" shows "f x y = g x y" proof - let ?P = "{x. \y\ span C. f x y = g x y}" from bf bg have sp: "subspace ?P" unfolding bilinear_def linear_iff subspace_def bf bg by (auto simp add: span_zero bilinear_lzero[OF bf] bilinear_lzero[OF bg] span_add Ball_def intro: bilinear_ladd[OF bf]) have sfg: "\x. x \ B \ subspace {a. f x a = g x a}" apply (auto simp add: subspace_def) using bf bg unfolding bilinear_def linear_iff apply (auto simp add: span_zero bilinear_rzero[OF bf] bilinear_rzero[OF bg] span_add Ball_def intro: bilinear_ladd[OF bf]) done have "\y\ span C. f x y = g x y" if "x \ span B" for x apply (rule span_induct [OF that sp]) using fg sfg span_induct by blast then show ?thesis using SB TC assms by auto qed lemma bilinear_eq_stdbasis: fixes f :: "'a::euclidean_space \ 'b::euclidean_space \ _" assumes bf: "bilinear f" and bg: "bilinear g" and fg: "\i j. i \ Basis \ j \ Basis \ f i j = g i j" shows "f = g" using bilinear_eq[OF bf bg equalityD2[OF span_Basis] equalityD2[OF span_Basis]] fg by blast subsection \Infinity norm\ definition\<^marker>\tag important\ "infnorm (x::'a::euclidean_space) = Sup {\x \ b\ |b. b \ Basis}" lemma infnorm_set_image: fixes x :: "'a::euclidean_space" shows "{\x \ i\ |i. i \ Basis} = (\i. \x \ i\) ` Basis" by blast lemma infnorm_Max: fixes x :: "'a::euclidean_space" shows "infnorm x = Max ((\i. \x \ i\) ` Basis)" by (simp add: infnorm_def infnorm_set_image cSup_eq_Max) lemma infnorm_set_lemma: fixes x :: "'a::euclidean_space" shows "finite {\x \ i\ |i. i \ Basis}" and "{\x \ i\ |i. i \ Basis} \ {}" unfolding infnorm_set_image by auto lemma infnorm_pos_le: fixes x :: "'a::euclidean_space" shows "0 \ infnorm x" by (simp add: infnorm_Max Max_ge_iff ex_in_conv) lemma infnorm_triangle: fixes x :: "'a::euclidean_space" shows "infnorm (x + y) \ infnorm x + infnorm y" proof - have *: "\a b c d :: real. \a\ \ c \ \b\ \ d \ \a + b\ \ c + d" by simp show ?thesis by (auto simp: infnorm_Max inner_add_left intro!: *) qed lemma infnorm_eq_0: fixes x :: "'a::euclidean_space" shows "infnorm x = 0 \ x = 0" proof - have "infnorm x \ 0 \ x = 0" unfolding infnorm_Max by (simp add: euclidean_all_zero_iff) then show ?thesis using infnorm_pos_le[of x] by simp qed lemma infnorm_0: "infnorm 0 = 0" by (simp add: infnorm_eq_0) lemma infnorm_neg: "infnorm (- x) = infnorm x" unfolding infnorm_def by simp lemma infnorm_sub: "infnorm (x - y) = infnorm (y - x)" by (metis infnorm_neg minus_diff_eq) lemma absdiff_infnorm: "\infnorm x - infnorm y\ \ infnorm (x - y)" proof - have *: "\(nx::real) n ny. nx \ n + ny \ ny \ n + nx \ \nx - ny\ \ n" by arith show ?thesis proof (rule *) from infnorm_triangle[of "x - y" " y"] infnorm_triangle[of "x - y" "-x"] show "infnorm x \ infnorm (x - y) + infnorm y" "infnorm y \ infnorm (x - y) + infnorm x" by (simp_all add: field_simps infnorm_neg) qed qed lemma real_abs_infnorm: "\infnorm x\ = infnorm x" using infnorm_pos_le[of x] by arith lemma Basis_le_infnorm: fixes x :: "'a::euclidean_space" shows "b \ Basis \ \x \ b\ \ infnorm x" by (simp add: infnorm_Max) lemma infnorm_mul: "infnorm (a *\<^sub>R x) = \a\ * infnorm x" unfolding infnorm_Max proof (safe intro!: Max_eqI) let ?B = "(\i. \x \ i\) ` Basis" { fix b :: 'a assume "b \ Basis" then show "\a *\<^sub>R x \ b\ \ \a\ * Max ?B" by (simp add: abs_mult mult_left_mono) next from Max_in[of ?B] obtain b where "b \ Basis" "Max ?B = \x \ b\" by (auto simp del: Max_in) then show "\a\ * Max ((\i. \x \ i\) ` Basis) \ (\i. \a *\<^sub>R x \ i\) ` Basis" by (intro image_eqI[where x=b]) (auto simp: abs_mult) } qed simp lemma infnorm_mul_lemma: "infnorm (a *\<^sub>R x) \ \a\ * infnorm x" unfolding infnorm_mul .. lemma infnorm_pos_lt: "infnorm x > 0 \ x \ 0" using infnorm_pos_le[of x] infnorm_eq_0[of x] by arith text \Prove that it differs only up to a bound from Euclidean norm.\ lemma infnorm_le_norm: "infnorm x \ norm x" by (simp add: Basis_le_norm infnorm_Max) lemma norm_le_infnorm: fixes x :: "'a::euclidean_space" shows "norm x \ sqrt DIM('a) * infnorm x" - unfolding norm_eq_sqrt_inner id_def + unfolding norm_eq_sqrt_inner id_def proof (rule real_le_lsqrt[OF inner_ge_zero]) show "sqrt DIM('a) * infnorm x \ 0" by (simp add: zero_le_mult_iff infnorm_pos_le) have "x \ x \ (\b\Basis. x \ b * (x \ b))" by (metis euclidean_inner order_refl) also have "... \ DIM('a) * \infnorm x\\<^sup>2" by (rule sum_bounded_above) (metis Basis_le_infnorm abs_le_square_iff power2_eq_square real_abs_infnorm) also have "... \ (sqrt DIM('a) * infnorm x)\<^sup>2" by (simp add: power_mult_distrib) finally show "x \ x \ (sqrt DIM('a) * infnorm x)\<^sup>2" . qed lemma tendsto_infnorm [tendsto_intros]: assumes "(f \ a) F" shows "((\x. infnorm (f x)) \ infnorm a) F" proof (rule tendsto_compose [OF LIM_I assms]) fix r :: real assume "r > 0" then show "\s>0. \x. x \ a \ norm (x - a) < s \ norm (infnorm x - infnorm a) < r" by (metis real_norm_def le_less_trans absdiff_infnorm infnorm_le_norm) qed text \Equality in Cauchy-Schwarz and triangle inequalities.\ lemma norm_cauchy_schwarz_eq: "x \ y = norm x * norm y \ norm x *\<^sub>R y = norm y *\<^sub>R x" (is "?lhs \ ?rhs") proof (cases "x=0") case True - then show ?thesis + then show ?thesis by auto next case False from inner_eq_zero_iff[of "norm y *\<^sub>R x - norm x *\<^sub>R y"] have "?rhs \ (norm y * (norm y * norm x * norm x - norm x * (x \ y)) - norm x * (norm y * (y \ x) - norm x * norm y * norm y) = 0)" using False unfolding inner_simps by (auto simp add: power2_norm_eq_inner[symmetric] power2_eq_square inner_commute field_simps) - also have "\ \ (2 * norm x * norm y * (norm x * norm y - x \ y) = 0)" + also have "\ \ (2 * norm x * norm y * (norm x * norm y - x \ y) = 0)" using False by (simp add: field_simps inner_commute) - also have "\ \ ?lhs" + also have "\ \ ?lhs" using False by auto finally show ?thesis by metis qed lemma norm_cauchy_schwarz_abs_eq: "\x \ y\ = norm x * norm y \ norm x *\<^sub>R y = norm y *\<^sub>R x \ norm x *\<^sub>R y = - norm y *\<^sub>R x" (is "?lhs \ ?rhs") proof - have th: "\(x::real) a. a \ 0 \ \x\ = a \ x = a \ x = - a" by arith have "?rhs \ norm x *\<^sub>R y = norm y *\<^sub>R x \ norm (- x) *\<^sub>R y = norm y *\<^sub>R (- x)" by simp also have "\ \ (x \ y = norm x * norm y \ (- x) \ y = norm x * norm y)" unfolding norm_cauchy_schwarz_eq[symmetric] unfolding norm_minus_cancel norm_scaleR .. also have "\ \ ?lhs" unfolding th[OF mult_nonneg_nonneg, OF norm_ge_zero[of x] norm_ge_zero[of y]] inner_simps by auto finally show ?thesis .. qed lemma norm_triangle_eq: fixes x y :: "'a::real_inner" shows "norm (x + y) = norm x + norm y \ norm x *\<^sub>R y = norm y *\<^sub>R x" proof (cases "x = 0 \ y = 0") case True - then show ?thesis + then show ?thesis by force next case False then have n: "norm x > 0" "norm y > 0" by auto have "norm (x + y) = norm x + norm y \ (norm (x + y))\<^sup>2 = (norm x + norm y)\<^sup>2" by simp also have "\ \ norm x *\<^sub>R y = norm y *\<^sub>R x" unfolding norm_cauchy_schwarz_eq[symmetric] unfolding power2_norm_eq_inner inner_simps by (simp add: power2_norm_eq_inner[symmetric] power2_eq_square inner_commute field_simps) finally show ?thesis . qed subsection \Collinearity\ definition\<^marker>\tag important\ collinear :: "'a::real_vector set \ bool" where "collinear S \ (\u. \x \ S. \ y \ S. \c. x - y = c *\<^sub>R u)" lemma collinear_alt: "collinear S \ (\u v. \x \ S. \c. x = u + c *\<^sub>R v)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding collinear_def by (metis Groups.add_ac(2) diff_add_cancel) next assume ?rhs then obtain u v where *: "\x. x \ S \ \c. x = u + c *\<^sub>R v" by (auto simp: ) have "\c. x - y = c *\<^sub>R v" if "x \ S" "y \ S" for x y by (metis *[OF \x \ S\] *[OF \y \ S\] scaleR_left.diff add_diff_cancel_left) then show ?lhs using collinear_def by blast qed lemma collinear: fixes S :: "'a::{perfect_space,real_vector} set" shows "collinear S \ (\u. u \ 0 \ (\x \ S. \ y \ S. \c. x - y = c *\<^sub>R u))" proof - have "\v. v \ 0 \ (\x\S. \y\S. \c. x - y = c *\<^sub>R v)" if "\x\S. \y\S. \c. x - y = c *\<^sub>R u" "u=0" for u proof - have "\x\S. \y\S. x = y" using that by auto moreover obtain v::'a where "v \ 0" using UNIV_not_singleton [of 0] by auto ultimately have "\x\S. \y\S. \c. x - y = c *\<^sub>R v" by auto then show ?thesis using \v \ 0\ by blast qed then show ?thesis apply (clarsimp simp: collinear_def) by (metis scaleR_zero_right vector_fraction_eq_iff) qed lemma collinear_subset: "\collinear T; S \ T\ \ collinear S" by (meson collinear_def subsetCE) lemma collinear_empty [iff]: "collinear {}" by (simp add: collinear_def) lemma collinear_sing [iff]: "collinear {x}" by (simp add: collinear_def) lemma collinear_2 [iff]: "collinear {x, y}" apply (simp add: collinear_def) apply (rule exI[where x="x - y"]) by (metis minus_diff_eq scaleR_left.minus scaleR_one) lemma collinear_lemma: "collinear {0, x, y} \ x = 0 \ y = 0 \ (\c. y = c *\<^sub>R x)" (is "?lhs \ ?rhs") proof (cases "x = 0 \ y = 0") case True then show ?thesis by (auto simp: insert_commute) next case False - show ?thesis + show ?thesis proof assume h: "?lhs" then obtain u where u: "\ x\ {0,x,y}. \y\ {0,x,y}. \c. x - y = c *\<^sub>R u" unfolding collinear_def by blast from u[rule_format, of x 0] u[rule_format, of y 0] obtain cx and cy where cx: "x = cx *\<^sub>R u" and cy: "y = cy *\<^sub>R u" by auto from cx cy False have cx0: "cx \ 0" and cy0: "cy \ 0" by auto let ?d = "cy / cx" from cx cy cx0 have "y = ?d *\<^sub>R x" by simp then show ?rhs using False by blast next assume h: "?rhs" then obtain c where c: "y = c *\<^sub>R x" using False by blast show ?lhs unfolding collinear_def c apply (rule exI[where x=x]) apply auto apply (rule exI[where x="- 1"], simp) apply (rule exI[where x= "-c"], simp) apply (rule exI[where x=1], simp) apply (rule exI[where x="1 - c"], simp add: scaleR_left_diff_distrib) apply (rule exI[where x="c - 1"], simp add: scaleR_left_diff_distrib) done qed qed lemma norm_cauchy_schwarz_equal: "\x \ y\ = norm x * norm y \ collinear {0, x, y}" proof (cases "x=0") case True then show ?thesis by (auto simp: insert_commute) next case False then have nnz: "norm x \ 0" by auto show ?thesis proof assume "\x \ y\ = norm x * norm y" then show "collinear {0, x, y}" - unfolding norm_cauchy_schwarz_abs_eq collinear_lemma + unfolding norm_cauchy_schwarz_abs_eq collinear_lemma by (meson eq_vector_fraction_iff nnz) next assume "collinear {0, x, y}" with False show "\x \ y\ = norm x * norm y" unfolding norm_cauchy_schwarz_abs_eq collinear_lemma by (auto simp: abs_if) qed qed subsection\Properties of special hyperplanes\ lemma subspace_hyperplane: "subspace {x. a \ x = 0}" by (simp add: subspace_def inner_right_distrib) lemma subspace_hyperplane2: "subspace {x. x \ a = 0}" by (simp add: inner_commute inner_right_distrib subspace_def) lemma special_hyperplane_span: fixes S :: "'n::euclidean_space set" assumes "k \ Basis" shows "{x. k \ x = 0} = span (Basis - {k})" proof - have *: "x \ span (Basis - {k})" if "k \ x = 0" for x proof - have "x = (\b\Basis. (x \ b) *\<^sub>R b)" by (simp add: euclidean_representation) also have "... = (\b \ Basis - {k}. (x \ b) *\<^sub>R b)" by (auto simp: sum.remove [of _ k] inner_commute assms that) finally have "x = (\b\Basis - {k}. (x \ b) *\<^sub>R b)" . then show ?thesis by (simp add: span_finite) qed show ?thesis apply (rule span_subspace [symmetric]) using assms apply (auto simp: inner_not_same_Basis intro: * subspace_hyperplane) done qed lemma dim_special_hyperplane: fixes k :: "'n::euclidean_space" shows "k \ Basis \ dim {x. k \ x = 0} = DIM('n) - 1" apply (simp add: special_hyperplane_span) apply (rule dim_unique [OF subset_refl]) apply (auto simp: independent_substdbasis) apply (metis member_remove remove_def span_base) done proposition dim_hyperplane: fixes a :: "'a::euclidean_space" assumes "a \ 0" shows "dim {x. a \ x = 0} = DIM('a) - 1" proof - have span0: "span {x. a \ x = 0} = {x. a \ x = 0}" by (rule span_unique) (auto simp: subspace_hyperplane) then obtain B where "independent B" and Bsub: "B \ {x. a \ x = 0}" and subspB: "{x. a \ x = 0} \ span B" and card0: "(card B = dim {x. a \ x = 0})" and ortho: "pairwise orthogonal B" using orthogonal_basis_exists by metis with assms have "a \ span B" by (metis (mono_tags, lifting) span_eq inner_eq_zero_iff mem_Collect_eq span0) then have ind: "independent (insert a B)" by (simp add: \independent B\ independent_insert) have "finite B" using \independent B\ independent_bound by blast have "UNIV \ span (insert a B)" proof fix y::'a obtain r z where z: "y = r *\<^sub>R a + z" "a \ z = 0" apply (rule_tac r="(a \ y) / (a \ a)" and z = "y - ((a \ y) / (a \ a)) *\<^sub>R a" in that) using assms by (auto simp: algebra_simps) show "y \ span (insert a B)" by (metis (mono_tags, lifting) z Bsub span_eq_iff add_diff_cancel_left' mem_Collect_eq span0 span_breakdown_eq span_subspace subspB) qed then have dima: "DIM('a) = dim(insert a B)" by (metis independent_Basis span_Basis dim_eq_card top.extremum_uniqueI) then show ?thesis by (metis (mono_tags, lifting) Bsub Diff_insert_absorb \a \ span B\ ind card0 card_Diff_singleton dim_span indep_card_eq_dim_span insertI1 subsetCE subspB) qed lemma lowdim_eq_hyperplane: fixes S :: "'a::euclidean_space set" assumes "dim S = DIM('a) - 1" obtains a where "a \ 0" and "span S = {x. a \ x = 0}" proof - have dimS: "dim S < DIM('a)" by (simp add: assms) then obtain b where b: "b \ 0" "span S \ {a. b \ a = 0}" using lowdim_subset_hyperplane [of S] by fastforce show ?thesis apply (rule that[OF b(1)]) apply (rule subspace_dim_equal) by (auto simp: assms b dim_hyperplane subspace_hyperplane) qed lemma dim_eq_hyperplane: fixes S :: "'n::euclidean_space set" shows "dim S = DIM('n) - 1 \ (\a. a \ 0 \ span S = {x. a \ x = 0})" by (metis One_nat_def dim_hyperplane dim_span lowdim_eq_hyperplane) subsection\ Orthogonal bases and Gram-Schmidt process\ lemma pairwise_orthogonal_independent: assumes "pairwise orthogonal S" and "0 \ S" shows "independent S" proof - have 0: "\x y. \x \ y; x \ S; y \ S\ \ x \ y = 0" using assms by (simp add: pairwise_def orthogonal_def) have "False" if "a \ S" and a: "a \ span (S - {a})" for a proof - obtain T U where "T \ S - {a}" "a = (\v\T. U v *\<^sub>R v)" using a by (force simp: span_explicit) then have "a \ a = a \ (\v\T. U v *\<^sub>R v)" by simp also have "... = 0" apply (simp add: inner_sum_right) apply (rule comm_monoid_add_class.sum.neutral) by (metis "0" DiffE \T \ S - {a}\ mult_not_zero singletonI subsetCE \a \ S\) finally show ?thesis using \0 \ S\ \a \ S\ by auto qed then show ?thesis by (force simp: dependent_def) qed lemma pairwise_orthogonal_imp_finite: fixes S :: "'a::euclidean_space set" assumes "pairwise orthogonal S" shows "finite S" proof - have "independent (S - {0})" apply (rule pairwise_orthogonal_independent) apply (metis Diff_iff assms pairwise_def) by blast then show ?thesis by (meson independent_imp_finite infinite_remove) qed lemma subspace_orthogonal_to_vector: "subspace {y. orthogonal x y}" by (simp add: subspace_def orthogonal_clauses) lemma subspace_orthogonal_to_vectors: "subspace {y. \x \ S. orthogonal x y}" by (simp add: subspace_def orthogonal_clauses) lemma orthogonal_to_span: assumes a: "a \ span S" and x: "\y. y \ S \ orthogonal x y" shows "orthogonal x a" by (metis a orthogonal_clauses(1,2,4) span_induct_alt x) proposition Gram_Schmidt_step: fixes S :: "'a::euclidean_space set" assumes S: "pairwise orthogonal S" and x: "x \ span S" shows "orthogonal x (a - (\b\S. (b \ a / (b \ b)) *\<^sub>R b))" proof - have "finite S" by (simp add: S pairwise_orthogonal_imp_finite) have "orthogonal (a - (\b\S. (b \ a / (b \ b)) *\<^sub>R b)) x" if "x \ S" for x proof - have "a \ x = (\y\S. if y = x then y \ a else 0)" by (simp add: \finite S\ inner_commute that) also have "... = (\b\S. b \ a * (b \ x) / (b \ b))" apply (rule sum.cong [OF refl], simp) by (meson S orthogonal_def pairwise_def that) finally show ?thesis by (simp add: orthogonal_def algebra_simps inner_sum_left) qed then show ?thesis using orthogonal_to_span orthogonal_commute x by blast qed lemma orthogonal_extension_aux: fixes S :: "'a::euclidean_space set" assumes "finite T" "finite S" "pairwise orthogonal S" shows "\U. pairwise orthogonal (S \ U) \ span (S \ U) = span (S \ T)" using assms proof (induction arbitrary: S) case empty then show ?case by simp (metis sup_bot_right) next case (insert a T) have 0: "\x y. \x \ y; x \ S; y \ S\ \ x \ y = 0" using insert by (simp add: pairwise_def orthogonal_def) define a' where "a' = a - (\b\S. (b \ a / (b \ b)) *\<^sub>R b)" obtain U where orthU: "pairwise orthogonal (S \ insert a' U)" and spanU: "span (insert a' S \ U) = span (insert a' S \ T)" by (rule exE [OF insert.IH [of "insert a' S"]]) (auto simp: Gram_Schmidt_step a'_def insert.prems orthogonal_commute pairwise_orthogonal_insert span_clauses) have orthS: "\x. x \ S \ a' \ x = 0" apply (simp add: a'_def) using Gram_Schmidt_step [OF \pairwise orthogonal S\] apply (force simp: orthogonal_def inner_commute span_superset [THEN subsetD]) done have "span (S \ insert a' U) = span (insert a' (S \ T))" using spanU by simp also have "... = span (insert a (S \ T))" apply (rule eq_span_insert_eq) apply (simp add: a'_def span_neg span_sum span_base span_mul) done also have "... = span (S \ insert a T)" by simp finally show ?case by (rule_tac x="insert a' U" in exI) (use orthU in auto) qed proposition orthogonal_extension: fixes S :: "'a::euclidean_space set" assumes S: "pairwise orthogonal S" obtains U where "pairwise orthogonal (S \ U)" "span (S \ U) = span (S \ T)" proof - obtain B where "finite B" "span B = span T" using basis_subspace_exists [of "span T"] subspace_span by metis with orthogonal_extension_aux [of B S] obtain U where "pairwise orthogonal (S \ U)" "span (S \ U) = span (S \ B)" using assms pairwise_orthogonal_imp_finite by auto with \span B = span T\ show ?thesis by (rule_tac U=U in that) (auto simp: span_Un) qed corollary\<^marker>\tag unimportant\ orthogonal_extension_strong: fixes S :: "'a::euclidean_space set" assumes S: "pairwise orthogonal S" obtains U where "U \ (insert 0 S) = {}" "pairwise orthogonal (S \ U)" "span (S \ U) = span (S \ T)" proof - obtain U where "pairwise orthogonal (S \ U)" "span (S \ U) = span (S \ T)" using orthogonal_extension assms by blast then show ?thesis apply (rule_tac U = "U - (insert 0 S)" in that) apply blast apply (force simp: pairwise_def) apply (metis Un_Diff_cancel Un_insert_left span_redundant span_zero) done qed subsection\Decomposing a vector into parts in orthogonal subspaces\ text\existence of orthonormal basis for a subspace.\ lemma orthogonal_spanningset_subspace: fixes S :: "'a :: euclidean_space set" assumes "subspace S" obtains B where "B \ S" "pairwise orthogonal B" "span B = S" proof - obtain B where "B \ S" "independent B" "S \ span B" "card B = dim S" using basis_exists by blast with orthogonal_extension [of "{}" B] show ?thesis by (metis Un_empty_left assms pairwise_empty span_superset span_subspace that) qed lemma orthogonal_basis_subspace: fixes S :: "'a :: euclidean_space set" assumes "subspace S" obtains B where "0 \ B" "B \ S" "pairwise orthogonal B" "independent B" "card B = dim S" "span B = S" proof - obtain B where "B \ S" "pairwise orthogonal B" "span B = S" using assms orthogonal_spanningset_subspace by blast then show ?thesis apply (rule_tac B = "B - {0}" in that) apply (auto simp: indep_card_eq_dim_span pairwise_subset pairwise_orthogonal_independent elim: pairwise_subset) done qed proposition orthonormal_basis_subspace: fixes S :: "'a :: euclidean_space set" assumes "subspace S" obtains B where "B \ S" "pairwise orthogonal B" and "\x. x \ B \ norm x = 1" and "independent B" "card B = dim S" "span B = S" proof - obtain B where "0 \ B" "B \ S" and orth: "pairwise orthogonal B" and "independent B" "card B = dim S" "span B = S" by (blast intro: orthogonal_basis_subspace [OF assms]) have 1: "(\x. x /\<^sub>R norm x) ` B \ S" using \span B = S\ span_superset span_mul by fastforce have 2: "pairwise orthogonal ((\x. x /\<^sub>R norm x) ` B)" using orth by (force simp: pairwise_def orthogonal_clauses) have 3: "\x. x \ (\x. x /\<^sub>R norm x) ` B \ norm x = 1" by (metis (no_types, lifting) \0 \ B\ image_iff norm_sgn sgn_div_norm) have 4: "independent ((\x. x /\<^sub>R norm x) ` B)" by (metis "2" "3" norm_zero pairwise_orthogonal_independent zero_neq_one) have "inj_on (\x. x /\<^sub>R norm x) B" proof fix x y assume "x \ B" "y \ B" "x /\<^sub>R norm x = y /\<^sub>R norm y" moreover have "\i. i \ B \ norm (i /\<^sub>R norm i) = 1" using 3 by blast ultimately show "x = y" by (metis norm_eq_1 orth orthogonal_clauses(7) orthogonal_commute orthogonal_def pairwise_def zero_neq_one) qed then have 5: "card ((\x. x /\<^sub>R norm x) ` B) = dim S" by (metis \card B = dim S\ card_image) have 6: "span ((\x. x /\<^sub>R norm x) ` B) = S" by (metis "1" "4" "5" assms card_eq_dim independent_imp_finite span_subspace) show ?thesis by (rule that [OF 1 2 3 4 5 6]) qed proposition\<^marker>\tag unimportant\ orthogonal_to_subspace_exists_gen: fixes S :: "'a :: euclidean_space set" assumes "span S \ span T" obtains x where "x \ 0" "x \ span T" "\y. y \ span S \ orthogonal x y" proof - obtain B where "B \ span S" and orthB: "pairwise orthogonal B" and "\x. x \ B \ norm x = 1" and "independent B" "card B = dim S" "span B = span S" by (rule orthonormal_basis_subspace [of "span S", OF subspace_span]) (auto) with assms obtain u where spanBT: "span B \ span T" and "u \ span B" "u \ span T" by auto obtain C where orthBC: "pairwise orthogonal (B \ C)" and spanBC: "span (B \ C) = span (B \ {u})" by (blast intro: orthogonal_extension [OF orthB]) show thesis proof (cases "C \ insert 0 B") case True then have "C \ span B" using span_eq by (metis span_insert_0 subset_trans) moreover have "u \ span (B \ C)" using \span (B \ C) = span (B \ {u})\ span_superset by force ultimately show ?thesis using True \u \ span B\ by (metis Un_insert_left span_insert_0 sup.orderE) next case False then obtain x where "x \ C" "x \ 0" "x \ B" by blast then have "x \ span T" by (metis (no_types, lifting) Un_insert_right Un_upper2 \u \ span T\ spanBT spanBC \u \ span T\ insert_subset span_superset span_mono span_span subsetCE subset_trans sup_bot.comm_neutral) moreover have "orthogonal x y" if "y \ span B" for y using that proof (rule span_induct) show "subspace {a. orthogonal x a}" by (simp add: subspace_orthogonal_to_vector) show "\b. b \ B \ orthogonal x b" by (metis Un_iff \x \ C\ \x \ B\ orthBC pairwise_def) qed ultimately show ?thesis using \x \ 0\ that \span B = span S\ by auto qed qed corollary\<^marker>\tag unimportant\ orthogonal_to_subspace_exists: fixes S :: "'a :: euclidean_space set" assumes "dim S < DIM('a)" obtains x where "x \ 0" "\y. y \ span S \ orthogonal x y" proof - have "span S \ UNIV" by (metis (mono_tags) UNIV_I assms inner_eq_zero_iff less_le lowdim_subset_hyperplane mem_Collect_eq top.extremum_strict top.not_eq_extremum) with orthogonal_to_subspace_exists_gen [of S UNIV] that show ?thesis by (auto) qed corollary\<^marker>\tag unimportant\ orthogonal_to_vector_exists: fixes x :: "'a :: euclidean_space" assumes "2 \ DIM('a)" obtains y where "y \ 0" "orthogonal x y" proof - have "dim {x} < DIM('a)" using assms by auto then show thesis by (rule orthogonal_to_subspace_exists) (simp add: orthogonal_commute span_base that) qed proposition\<^marker>\tag unimportant\ orthogonal_subspace_decomp_exists: fixes S :: "'a :: euclidean_space set" obtains y z where "y \ span S" and "\w. w \ span S \ orthogonal z w" and "x = y + z" proof - obtain T where "0 \ T" "T \ span S" "pairwise orthogonal T" "independent T" "card T = dim (span S)" "span T = span S" using orthogonal_basis_subspace subspace_span by blast let ?a = "\b\T. (b \ x / (b \ b)) *\<^sub>R b" have orth: "orthogonal (x - ?a) w" if "w \ span S" for w by (simp add: Gram_Schmidt_step \pairwise orthogonal T\ \span T = span S\ orthogonal_commute that) show ?thesis apply (rule_tac y = "?a" and z = "x - ?a" in that) apply (meson \T \ span S\ span_scale span_sum subsetCE) apply (fact orth, simp) done qed lemma orthogonal_subspace_decomp_unique: fixes S :: "'a :: euclidean_space set" assumes "x + y = x' + y'" and ST: "x \ span S" "x' \ span S" "y \ span T" "y' \ span T" and orth: "\a b. \a \ S; b \ T\ \ orthogonal a b" shows "x = x' \ y = y'" proof - have "x + y - y' = x'" by (simp add: assms) moreover have "\a b. \a \ span S; b \ span T\ \ orthogonal a b" by (meson orth orthogonal_commute orthogonal_to_span) ultimately have "0 = x' - x" by (metis (full_types) add_diff_cancel_left' ST diff_right_commute orthogonal_clauses(10) orthogonal_clauses(5) orthogonal_self) with assms show ?thesis by auto qed lemma vector_in_orthogonal_spanningset: fixes a :: "'a::euclidean_space" obtains S where "a \ S" "pairwise orthogonal S" "span S = UNIV" by (metis UNIV_I Un_iff empty_iff insert_subset orthogonal_extension pairwise_def pairwise_orthogonal_insert span_UNIV subsetI subset_antisym) lemma vector_in_orthogonal_basis: fixes a :: "'a::euclidean_space" assumes "a \ 0" obtains S where "a \ S" "0 \ S" "pairwise orthogonal S" "independent S" "finite S" "span S = UNIV" "card S = DIM('a)" proof - obtain S where S: "a \ S" "pairwise orthogonal S" "span S = UNIV" using vector_in_orthogonal_spanningset . show thesis proof show "pairwise orthogonal (S - {0})" using pairwise_mono S(2) by blast show "independent (S - {0})" by (simp add: \pairwise orthogonal (S - {0})\ pairwise_orthogonal_independent) show "finite (S - {0})" using \independent (S - {0})\ independent_imp_finite by blast show "card (S - {0}) = DIM('a)" using span_delete_0 [of S] S by (simp add: \independent (S - {0})\ indep_card_eq_dim_span) qed (use S \a \ 0\ in auto) qed lemma vector_in_orthonormal_basis: fixes a :: "'a::euclidean_space" assumes "norm a = 1" obtains S where "a \ S" "pairwise orthogonal S" "\x. x \ S \ norm x = 1" "independent S" "card S = DIM('a)" "span S = UNIV" proof - have "a \ 0" using assms by auto then obtain S where "a \ S" "0 \ S" "finite S" and S: "pairwise orthogonal S" "independent S" "span S = UNIV" "card S = DIM('a)" by (metis vector_in_orthogonal_basis) let ?S = "(\x. x /\<^sub>R norm x) ` S" show thesis proof show "a \ ?S" using \a \ S\ assms image_iff by fastforce next show "pairwise orthogonal ?S" using \pairwise orthogonal S\ by (auto simp: pairwise_def orthogonal_def) show "\x. x \ (\x. x /\<^sub>R norm x) ` S \ norm x = 1" using \0 \ S\ by (auto simp: field_split_simps) then show "independent ?S" by (metis \pairwise orthogonal ((\x. x /\<^sub>R norm x) ` S)\ norm_zero pairwise_orthogonal_independent zero_neq_one) have "inj_on (\x. x /\<^sub>R norm x) S" unfolding inj_on_def by (metis (full_types) S(1) \0 \ S\ inverse_nonzero_iff_nonzero norm_eq_zero orthogonal_scaleR orthogonal_self pairwise_def) then show "card ?S = DIM('a)" by (simp add: card_image S) show "span ?S = UNIV" by (metis (no_types) \0 \ S\ \finite S\ \span S = UNIV\ field_class.field_inverse_zero inverse_inverse_eq less_irrefl span_image_scale zero_less_norm_iff) qed qed proposition dim_orthogonal_sum: fixes A :: "'a::euclidean_space set" assumes "\x y. \x \ A; y \ B\ \ x \ y = 0" shows "dim(A \ B) = dim A + dim B" proof - have 1: "\x y. \x \ span A; y \ B\ \ x \ y = 0" by (erule span_induct [OF _ subspace_hyperplane2]; simp add: assms) have "\x y. \x \ span A; y \ span B\ \ x \ y = 0" using 1 by (simp add: span_induct [OF _ subspace_hyperplane]) then have 0: "\x y. \x \ span A; y \ span B\ \ x \ y = 0" by simp have "dim(A \ B) = dim (span (A \ B))" by (simp) also have "span (A \ B) = ((\(a, b). a + b) ` (span A \ span B))" by (auto simp add: span_Un image_def) also have "dim \ = dim {x + y |x y. x \ span A \ y \ span B}" by (auto intro!: arg_cong [where f=dim]) also have "... = dim {x + y |x y. x \ span A \ y \ span B} + dim(span A \ span B)" by (auto simp: dest: 0) also have "... = dim (span A) + dim (span B)" by (rule dim_sums_Int) (auto) also have "... = dim A + dim B" by (simp) finally show ?thesis . qed lemma dim_subspace_orthogonal_to_vectors: fixes A :: "'a::euclidean_space set" assumes "subspace A" "subspace B" "A \ B" shows "dim {y \ B. \x \ A. orthogonal x y} + dim A = dim B" proof - have "dim (span ({y \ B. \x\A. orthogonal x y} \ A)) = dim (span B)" proof (rule arg_cong [where f=dim, OF subset_antisym]) show "span ({y \ B. \x\A. orthogonal x y} \ A) \ span B" by (simp add: \A \ B\ Collect_restrict span_mono) next have *: "x \ span ({y \ B. \x\A. orthogonal x y} \ A)" if "x \ B" for x proof - obtain y z where "x = y + z" "y \ span A" and orth: "\w. w \ span A \ orthogonal z w" using orthogonal_subspace_decomp_exists [of A x] that by auto have "y \ span B" using \y \ span A\ assms(3) span_mono by blast then have "z \ {a \ B. \x. x \ A \ orthogonal x a}" apply simp using \x = y + z\ assms(1) assms(2) orth orthogonal_commute span_add_eq span_eq_iff that by blast then have z: "z \ span {y \ B. \x\A. orthogonal x y}" by (meson span_superset subset_iff) then show ?thesis apply (auto simp: span_Un image_def \x = y + z\ \y \ span A\) using \y \ span A\ add.commute by blast qed show "span B \ span ({y \ B. \x\A. orthogonal x y} \ A)" by (rule span_minimal) (auto intro: * span_minimal) qed then show ?thesis by (metis (no_types, lifting) dim_orthogonal_sum dim_span mem_Collect_eq orthogonal_commute orthogonal_def) qed subsection\Linear functions are (uniformly) continuous on any set\ subsection\<^marker>\tag unimportant\ \Topological properties of linear functions\ lemma linear_lim_0: assumes "bounded_linear f" shows "(f \ 0) (at (0))" proof - interpret f: bounded_linear f by fact have "(f \ f 0) (at 0)" using tendsto_ident_at by (rule f.tendsto) then show ?thesis unfolding f.zero . qed lemma linear_continuous_at: assumes "bounded_linear f" shows "continuous (at a) f" unfolding continuous_at using assms apply (rule bounded_linear.tendsto) apply (rule tendsto_ident_at) done lemma linear_continuous_within: "bounded_linear f \ continuous (at x within s) f" using continuous_at_imp_continuous_at_within linear_continuous_at by blast lemma linear_continuous_on: "bounded_linear f \ continuous_on s f" using continuous_at_imp_continuous_on[of s f] using linear_continuous_at[of f] by auto lemma Lim_linear: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" and h :: "'b \ 'c::real_normed_vector" assumes "(f \ l) F" "linear h" shows "((\x. h(f x)) \ h l) F" proof - obtain B where B: "B > 0" "\x. norm (h x) \ B * norm x" using linear_bounded_pos [OF \linear h\] by blast show ?thesis unfolding tendsto_iff proof (intro allI impI) show "\\<^sub>F x in F. dist (h (f x)) (h l) < e" if "e > 0" for e proof - have "\\<^sub>F x in F. dist (f x) l < e/B" by (simp add: \0 < B\ assms(1) tendstoD that) then show ?thesis unfolding dist_norm proof (rule eventually_mono) show "norm (h (f x) - h l) < e" if "norm (f x - l) < e / B" for x using that B apply (simp add: field_split_simps) by (metis \linear h\ le_less_trans linear_diff) qed qed qed qed lemma linear_continuous_compose: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" and g :: "'b \ 'c::real_normed_vector" assumes "continuous F f" "linear g" shows "continuous F (\x. g(f x))" using assms unfolding continuous_def by (rule Lim_linear) lemma linear_continuous_on_compose: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" and g :: "'b \ 'c::real_normed_vector" assumes "continuous_on S f" "linear g" shows "continuous_on S (\x. g(f x))" using assms by (simp add: continuous_on_eq_continuous_within linear_continuous_compose) text\Also bilinear functions, in composition form\ lemma bilinear_continuous_compose: fixes h :: "'a::euclidean_space \ 'b::euclidean_space \ 'c::real_normed_vector" assumes "continuous F f" "continuous F g" "bilinear h" shows "continuous F (\x. h (f x) (g x))" using assms bilinear_conv_bounded_bilinear bounded_bilinear.continuous by blast lemma bilinear_continuous_on_compose: fixes h :: "'a::euclidean_space \ 'b::euclidean_space \ 'c::real_normed_vector" and f :: "'d::t2_space \ 'a" assumes "continuous_on S f" "continuous_on S g" "bilinear h" shows "continuous_on S (\x. h (f x) (g x))" using assms by (simp add: continuous_on_eq_continuous_within bilinear_continuous_compose) end diff --git a/src/HOL/Analysis/Path_Connected.thy b/src/HOL/Analysis/Path_Connected.thy --- a/src/HOL/Analysis/Path_Connected.thy +++ b/src/HOL/Analysis/Path_Connected.thy @@ -1,4125 +1,4128 @@ (* Title: HOL/Analysis/Path_Connected.thy Authors: LC Paulson and Robert Himmelmann (TU Muenchen), based on material from HOL Light *) section \Path-Connectedness\ theory Path_Connected imports Starlike T1_Spaces begin subsection \Paths and Arcs\ definition\<^marker>\tag important\ path :: "(real \ 'a::topological_space) \ bool" where "path g \ continuous_on {0..1} g" definition\<^marker>\tag important\ pathstart :: "(real \ 'a::topological_space) \ 'a" where "pathstart g = g 0" definition\<^marker>\tag important\ pathfinish :: "(real \ 'a::topological_space) \ 'a" where "pathfinish g = g 1" definition\<^marker>\tag important\ path_image :: "(real \ 'a::topological_space) \ 'a set" where "path_image g = g ` {0 .. 1}" definition\<^marker>\tag important\ reversepath :: "(real \ 'a::topological_space) \ real \ 'a" where "reversepath g = (\x. g(1 - x))" definition\<^marker>\tag important\ joinpaths :: "(real \ 'a::topological_space) \ (real \ 'a) \ real \ 'a" (infixr "+++" 75) where "g1 +++ g2 = (\x. if x \ 1/2 then g1 (2 * x) else g2 (2 * x - 1))" definition\<^marker>\tag important\ simple_path :: "(real \ 'a::topological_space) \ bool" where "simple_path g \ path g \ (\x\{0..1}. \y\{0..1}. g x = g y \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0)" definition\<^marker>\tag important\ arc :: "(real \ 'a :: topological_space) \ bool" where "arc g \ path g \ inj_on g {0..1}" subsection\<^marker>\tag unimportant\\Invariance theorems\ lemma path_eq: "path p \ (\t. t \ {0..1} \ p t = q t) \ path q" using continuous_on_eq path_def by blast lemma path_continuous_image: "path g \ continuous_on (path_image g) f \ path(f \ g)" unfolding path_def path_image_def using continuous_on_compose by blast lemma continuous_on_translation_eq: fixes g :: "'a :: real_normed_vector \ 'b :: real_normed_vector" shows "continuous_on A ((+) a \ g) = continuous_on A g" proof - have g: "g = (\x. -a + x) \ ((\x. a + x) \ g)" by (rule ext) simp show ?thesis by (metis (no_types, hide_lams) g continuous_on_compose homeomorphism_def homeomorphism_translation) qed lemma path_translation_eq: fixes g :: "real \ 'a :: real_normed_vector" shows "path((\x. a + x) \ g) = path g" using continuous_on_translation_eq path_def by blast lemma path_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "path(f \ g) = path g" proof - from linear_injective_left_inverse [OF assms] obtain h where h: "linear h" "h \ f = id" by blast then have g: "g = h \ (f \ g)" by (metis comp_assoc id_comp) show ?thesis unfolding path_def using h assms by (metis g continuous_on_compose linear_continuous_on linear_conv_bounded_linear) qed lemma pathstart_translation: "pathstart((\x. a + x) \ g) = a + pathstart g" by (simp add: pathstart_def) lemma pathstart_linear_image_eq: "linear f \ pathstart(f \ g) = f(pathstart g)" by (simp add: pathstart_def) lemma pathfinish_translation: "pathfinish((\x. a + x) \ g) = a + pathfinish g" by (simp add: pathfinish_def) lemma pathfinish_linear_image: "linear f \ pathfinish(f \ g) = f(pathfinish g)" by (simp add: pathfinish_def) lemma path_image_translation: "path_image((\x. a + x) \ g) = (\x. a + x) ` (path_image g)" by (simp add: image_comp path_image_def) lemma path_image_linear_image: "linear f \ path_image(f \ g) = f ` (path_image g)" by (simp add: image_comp path_image_def) lemma reversepath_translation: "reversepath((\x. a + x) \ g) = (\x. a + x) \ reversepath g" by (rule ext) (simp add: reversepath_def) lemma reversepath_linear_image: "linear f \ reversepath(f \ g) = f \ reversepath g" by (rule ext) (simp add: reversepath_def) lemma joinpaths_translation: "((\x. a + x) \ g1) +++ ((\x. a + x) \ g2) = (\x. a + x) \ (g1 +++ g2)" by (rule ext) (simp add: joinpaths_def) lemma joinpaths_linear_image: "linear f \ (f \ g1) +++ (f \ g2) = f \ (g1 +++ g2)" by (rule ext) (simp add: joinpaths_def) lemma simple_path_translation_eq: fixes g :: "real \ 'a::euclidean_space" shows "simple_path((\x. a + x) \ g) = simple_path g" by (simp add: simple_path_def path_translation_eq) lemma simple_path_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "simple_path(f \ g) = simple_path g" using assms inj_on_eq_iff [of f] by (auto simp: path_linear_image_eq simple_path_def path_translation_eq) lemma arc_translation_eq: fixes g :: "real \ 'a::euclidean_space" shows "arc((\x. a + x) \ g) = arc g" by (auto simp: arc_def inj_on_def path_translation_eq) lemma arc_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "arc(f \ g) = arc g" using assms inj_on_eq_iff [of f] by (auto simp: arc_def inj_on_def path_linear_image_eq) subsection\<^marker>\tag unimportant\\Basic lemmas about paths\ lemma pathin_iff_path_real [simp]: "pathin euclideanreal g \ path g" by (simp add: pathin_def path_def) lemma continuous_on_path: "path f \ t \ {0..1} \ continuous_on t f" using continuous_on_subset path_def by blast lemma arc_imp_simple_path: "arc g \ simple_path g" by (simp add: arc_def inj_on_def simple_path_def) lemma arc_imp_path: "arc g \ path g" using arc_def by blast lemma arc_imp_inj_on: "arc g \ inj_on g {0..1}" by (auto simp: arc_def) lemma simple_path_imp_path: "simple_path g \ path g" using simple_path_def by blast lemma simple_path_cases: "simple_path g \ arc g \ pathfinish g = pathstart g" unfolding simple_path_def arc_def inj_on_def pathfinish_def pathstart_def by force lemma simple_path_imp_arc: "simple_path g \ pathfinish g \ pathstart g \ arc g" using simple_path_cases by auto lemma arc_distinct_ends: "arc g \ pathfinish g \ pathstart g" unfolding arc_def inj_on_def pathfinish_def pathstart_def by fastforce lemma arc_simple_path: "arc g \ simple_path g \ pathfinish g \ pathstart g" using arc_distinct_ends arc_imp_simple_path simple_path_cases by blast lemma simple_path_eq_arc: "pathfinish g \ pathstart g \ (simple_path g = arc g)" by (simp add: arc_simple_path) lemma path_image_const [simp]: "path_image (\t. a) = {a}" by (force simp: path_image_def) lemma path_image_nonempty [simp]: "path_image g \ {}" unfolding path_image_def image_is_empty box_eq_empty by auto lemma pathstart_in_path_image[intro]: "pathstart g \ path_image g" unfolding pathstart_def path_image_def by auto lemma pathfinish_in_path_image[intro]: "pathfinish g \ path_image g" unfolding pathfinish_def path_image_def by auto lemma connected_path_image[intro]: "path g \ connected (path_image g)" unfolding path_def path_image_def using connected_continuous_image connected_Icc by blast lemma compact_path_image[intro]: "path g \ compact (path_image g)" unfolding path_def path_image_def using compact_continuous_image connected_Icc by blast lemma reversepath_reversepath[simp]: "reversepath (reversepath g) = g" unfolding reversepath_def by auto lemma pathstart_reversepath[simp]: "pathstart (reversepath g) = pathfinish g" unfolding pathstart_def reversepath_def pathfinish_def by auto lemma pathfinish_reversepath[simp]: "pathfinish (reversepath g) = pathstart g" unfolding pathstart_def reversepath_def pathfinish_def by auto +lemma reversepath_o: "reversepath g = g \ (-)1" + by (auto simp: reversepath_def) + lemma pathstart_join[simp]: "pathstart (g1 +++ g2) = pathstart g1" unfolding pathstart_def joinpaths_def pathfinish_def by auto lemma pathfinish_join[simp]: "pathfinish (g1 +++ g2) = pathfinish g2" unfolding pathstart_def joinpaths_def pathfinish_def by auto lemma path_image_reversepath[simp]: "path_image (reversepath g) = path_image g" proof - have *: "\g. path_image (reversepath g) \ path_image g" unfolding path_image_def subset_eq reversepath_def Ball_def image_iff by force show ?thesis using *[of g] *[of "reversepath g"] unfolding reversepath_reversepath by auto qed lemma path_reversepath [simp]: "path (reversepath g) \ path g" proof - have *: "\g. path g \ path (reversepath g)" unfolding path_def reversepath_def apply (rule continuous_on_compose[unfolded o_def, of _ "\x. 1 - x"]) apply (auto intro: continuous_intros continuous_on_subset[of "{0..1}"]) done show ?thesis using "*" by force qed lemma arc_reversepath: assumes "arc g" shows "arc(reversepath g)" proof - have injg: "inj_on g {0..1}" using assms by (simp add: arc_def) have **: "\x y::real. 1-x = 1-y \ x = y" by simp show ?thesis using assms by (clarsimp simp: arc_def intro!: inj_onI) (simp add: inj_onD reversepath_def **) qed lemma simple_path_reversepath: "simple_path g \ simple_path (reversepath g)" apply (simp add: simple_path_def) apply (force simp: reversepath_def) done lemmas reversepath_simps = path_reversepath path_image_reversepath pathstart_reversepath pathfinish_reversepath lemma path_join[simp]: assumes "pathfinish g1 = pathstart g2" shows "path (g1 +++ g2) \ path g1 \ path g2" unfolding path_def pathfinish_def pathstart_def proof safe assume cont: "continuous_on {0..1} (g1 +++ g2)" have g1: "continuous_on {0..1} g1 \ continuous_on {0..1} ((g1 +++ g2) \ (\x. x / 2))" by (intro continuous_on_cong refl) (auto simp: joinpaths_def) have g2: "continuous_on {0..1} g2 \ continuous_on {0..1} ((g1 +++ g2) \ (\x. x / 2 + 1/2))" using assms by (intro continuous_on_cong refl) (auto simp: joinpaths_def pathfinish_def pathstart_def) show "continuous_on {0..1} g1" and "continuous_on {0..1} g2" unfolding g1 g2 by (auto intro!: continuous_intros continuous_on_subset[OF cont] simp del: o_apply) next assume g1g2: "continuous_on {0..1} g1" "continuous_on {0..1} g2" have 01: "{0 .. 1} = {0..1/2} \ {1/2 .. 1::real}" by auto { fix x :: real assume "0 \ x" and "x \ 1" then have "x \ (\x. x * 2) ` {0..1 / 2}" by (intro image_eqI[where x="x/2"]) auto } note 1 = this { fix x :: real assume "0 \ x" and "x \ 1" then have "x \ (\x. x * 2 - 1) ` {1 / 2..1}" by (intro image_eqI[where x="x/2 + 1/2"]) auto } note 2 = this show "continuous_on {0..1} (g1 +++ g2)" using assms unfolding joinpaths_def 01 apply (intro continuous_on_cases closed_atLeastAtMost g1g2[THEN continuous_on_compose2] continuous_intros) apply (auto simp: field_simps pathfinish_def pathstart_def intro!: 1 2) done qed subsection\<^marker>\tag unimportant\ \Path Images\ lemma bounded_path_image: "path g \ bounded(path_image g)" by (simp add: compact_imp_bounded compact_path_image) lemma closed_path_image: fixes g :: "real \ 'a::t2_space" shows "path g \ closed(path_image g)" by (metis compact_path_image compact_imp_closed) lemma connected_simple_path_image: "simple_path g \ connected(path_image g)" by (metis connected_path_image simple_path_imp_path) lemma compact_simple_path_image: "simple_path g \ compact(path_image g)" by (metis compact_path_image simple_path_imp_path) lemma bounded_simple_path_image: "simple_path g \ bounded(path_image g)" by (metis bounded_path_image simple_path_imp_path) lemma closed_simple_path_image: fixes g :: "real \ 'a::t2_space" shows "simple_path g \ closed(path_image g)" by (metis closed_path_image simple_path_imp_path) lemma connected_arc_image: "arc g \ connected(path_image g)" by (metis connected_path_image arc_imp_path) lemma compact_arc_image: "arc g \ compact(path_image g)" by (metis compact_path_image arc_imp_path) lemma bounded_arc_image: "arc g \ bounded(path_image g)" by (metis bounded_path_image arc_imp_path) lemma closed_arc_image: fixes g :: "real \ 'a::t2_space" shows "arc g \ closed(path_image g)" by (metis closed_path_image arc_imp_path) lemma path_image_join_subset: "path_image (g1 +++ g2) \ path_image g1 \ path_image g2" unfolding path_image_def joinpaths_def by auto lemma subset_path_image_join: assumes "path_image g1 \ s" and "path_image g2 \ s" shows "path_image (g1 +++ g2) \ s" using path_image_join_subset[of g1 g2] and assms by auto lemma path_image_join: assumes "pathfinish g1 = pathstart g2" shows "path_image(g1 +++ g2) = path_image g1 \ path_image g2" proof - have "path_image g1 \ path_image (g1 +++ g2)" proof (clarsimp simp: path_image_def joinpaths_def) fix u::real assume "0 \ u" "u \ 1" then show "g1 u \ (\x. g1 (2 * x)) ` ({0..1} \ {x. x * 2 \ 1})" by (rule_tac x="u/2" in image_eqI) auto qed moreover have \
: "g2 u \ (\x. g2 (2 * x - 1)) ` ({0..1} \ {x. \ x * 2 \ 1})" if "0 < u" "u \ 1" for u using that assms by (rule_tac x="(u+1)/2" in image_eqI) (auto simp: field_simps pathfinish_def pathstart_def) have "g2 0 \ (\x. g1 (2 * x)) ` ({0..1} \ {x. x * 2 \ 1})" using assms by (rule_tac x="1/2" in image_eqI) (auto simp: pathfinish_def pathstart_def) then have "path_image g2 \ path_image (g1 +++ g2)" by (auto simp: path_image_def joinpaths_def intro!: \
) ultimately show ?thesis using path_image_join_subset by blast qed lemma not_in_path_image_join: assumes "x \ path_image g1" and "x \ path_image g2" shows "x \ path_image (g1 +++ g2)" using assms and path_image_join_subset[of g1 g2] by auto lemma pathstart_compose: "pathstart(f \ p) = f(pathstart p)" by (simp add: pathstart_def) lemma pathfinish_compose: "pathfinish(f \ p) = f(pathfinish p)" by (simp add: pathfinish_def) lemma path_image_compose: "path_image (f \ p) = f ` (path_image p)" by (simp add: image_comp path_image_def) lemma path_compose_join: "f \ (p +++ q) = (f \ p) +++ (f \ q)" by (rule ext) (simp add: joinpaths_def) lemma path_compose_reversepath: "f \ reversepath p = reversepath(f \ p)" by (rule ext) (simp add: reversepath_def) lemma joinpaths_eq: "(\t. t \ {0..1} \ p t = p' t) \ (\t. t \ {0..1} \ q t = q' t) \ t \ {0..1} \ (p +++ q) t = (p' +++ q') t" by (auto simp: joinpaths_def) lemma simple_path_inj_on: "simple_path g \ inj_on g {0<..<1}" by (auto simp: simple_path_def path_image_def inj_on_def less_eq_real_def Ball_def) subsection\<^marker>\tag unimportant\\Simple paths with the endpoints removed\ lemma simple_path_endless: assumes "simple_path c" shows "path_image c - {pathstart c,pathfinish c} = c ` {0<..<1}" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" using less_eq_real_def by (auto simp: path_image_def pathstart_def pathfinish_def) show "?rhs \ ?lhs" using assms apply (auto simp: simple_path_def path_image_def pathstart_def pathfinish_def Ball_def) using less_eq_real_def zero_le_one by blast+ qed lemma connected_simple_path_endless: assumes "simple_path c" shows "connected(path_image c - {pathstart c,pathfinish c})" proof - have "continuous_on {0<..<1} c" using assms by (simp add: simple_path_def continuous_on_path path_def subset_iff) then have "connected (c ` {0<..<1})" using connected_Ioo connected_continuous_image by blast then show ?thesis using assms by (simp add: simple_path_endless) qed lemma nonempty_simple_path_endless: "simple_path c \ path_image c - {pathstart c,pathfinish c} \ {}" by (simp add: simple_path_endless) subsection\<^marker>\tag unimportant\\The operations on paths\ lemma path_image_subset_reversepath: "path_image(reversepath g) \ path_image g" by simp lemma path_imp_reversepath: "path g \ path(reversepath g)" by simp lemma half_bounded_equal: "1 \ x * 2 \ x * 2 \ 1 \ x = (1/2::real)" by simp lemma continuous_on_joinpaths: assumes "continuous_on {0..1} g1" "continuous_on {0..1} g2" "pathfinish g1 = pathstart g2" shows "continuous_on {0..1} (g1 +++ g2)" proof - have "{0..1::real} = {0..1/2} \ {1/2..1}" by auto then show ?thesis using assms by (metis path_def path_join) qed lemma path_join_imp: "\path g1; path g2; pathfinish g1 = pathstart g2\ \ path(g1 +++ g2)" by simp lemma simple_path_join_loop: assumes "arc g1" "arc g2" "pathfinish g1 = pathstart g2" "pathfinish g2 = pathstart g1" "path_image g1 \ path_image g2 \ {pathstart g1, pathstart g2}" shows "simple_path(g1 +++ g2)" proof - have injg1: "inj_on g1 {0..1}" using assms by (simp add: arc_def) have injg2: "inj_on g2 {0..1}" using assms by (simp add: arc_def) have g12: "g1 1 = g2 0" and g21: "g2 1 = g1 0" and sb: "g1 ` {0..1} \ g2 ` {0..1} \ {g1 0, g2 0}" using assms by (simp_all add: arc_def pathfinish_def pathstart_def path_image_def) { fix x and y::real assume g2_eq: "g2 (2 * x - 1) = g1 (2 * y)" and xyI: "x \ 1 \ y \ 0" and xy: "x \ 1" "0 \ y" " y * 2 \ 1" "\ x * 2 \ 1" then consider "g1 (2 * y) = g1 0" | "g1 (2 * y) = g2 0" using sb by force then have False proof cases case 1 then have "y = 0" using xy g2_eq by (auto dest!: inj_onD [OF injg1]) then show ?thesis using xy g2_eq xyI by (auto dest: inj_onD [OF injg2] simp flip: g21) next case 2 then have "2*x = 1" using g2_eq g12 inj_onD [OF injg2] atLeastAtMost_iff xy(1) xy(4) by fastforce with xy show False by auto qed } note * = this { fix x and y::real assume xy: "g1 (2 * x) = g2 (2 * y - 1)" "y \ 1" "0 \ x" "\ y * 2 \ 1" "x * 2 \ 1" then have "x = 0 \ y = 1" using * xy by force } note ** = this show ?thesis using assms apply (simp add: arc_def simple_path_def) apply (auto simp: joinpaths_def split: if_split_asm dest!: * ** dest: inj_onD [OF injg1] inj_onD [OF injg2]) done qed lemma arc_join: assumes "arc g1" "arc g2" "pathfinish g1 = pathstart g2" "path_image g1 \ path_image g2 \ {pathstart g2}" shows "arc(g1 +++ g2)" proof - have injg1: "inj_on g1 {0..1}" using assms by (simp add: arc_def) have injg2: "inj_on g2 {0..1}" using assms by (simp add: arc_def) have g11: "g1 1 = g2 0" and sb: "g1 ` {0..1} \ g2 ` {0..1} \ {g2 0}" using assms by (simp_all add: arc_def pathfinish_def pathstart_def path_image_def) { fix x and y::real assume xy: "g2 (2 * x - 1) = g1 (2 * y)" "x \ 1" "0 \ y" " y * 2 \ 1" "\ x * 2 \ 1" then have "g1 (2 * y) = g2 0" using sb by force then have False using xy inj_onD injg2 by fastforce } note * = this show ?thesis using assms apply (simp add: arc_def inj_on_def) apply (auto simp: joinpaths_def arc_imp_path split: if_split_asm dest: * *[OF sym] inj_onD [OF injg1] inj_onD [OF injg2]) done qed lemma reversepath_joinpaths: "pathfinish g1 = pathstart g2 \ reversepath(g1 +++ g2) = reversepath g2 +++ reversepath g1" unfolding reversepath_def pathfinish_def pathstart_def joinpaths_def by (rule ext) (auto simp: mult.commute) subsection\<^marker>\tag unimportant\\Some reversed and "if and only if" versions of joining theorems\ lemma path_join_path_ends: fixes g1 :: "real \ 'a::metric_space" assumes "path(g1 +++ g2)" "path g2" shows "pathfinish g1 = pathstart g2" proof (rule ccontr) define e where "e = dist (g1 1) (g2 0)" assume Neg: "pathfinish g1 \ pathstart g2" then have "0 < dist (pathfinish g1) (pathstart g2)" by auto then have "e > 0" by (metis e_def pathfinish_def pathstart_def) then have "\e>0. \d>0. \x'\{0..1}. dist x' 0 < d \ dist (g2 x') (g2 0) < e" using \path g2\ atLeastAtMost_iff zero_le_one unfolding path_def continuous_on_iff by blast then obtain d1 where "d1 > 0" and d1: "\x'. \x'\{0..1}; norm x' < d1\ \ dist (g2 x') (g2 0) < e/2" by (metis \0 < e\ half_gt_zero_iff norm_conv_dist) 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) apply (drule_tac x="e/2" in spec, force simp: joinpaths_def) done have int01_1: "min (1/2) (min d1 d2) / 2 \ {0..1}" using \d1 > 0\ \d2 > 0\ by (simp add: min_def) have dist1: "norm (min (1 / 2) (min d1 d2) / 2) < d1" using \d1 > 0\ \d2 > 0\ by (simp add: min_def dist_norm) have int01_2: "1/2 + min (1/2) (min d1 d2) / 4 \ {0..1}" using \d1 > 0\ \d2 > 0\ by (simp add: min_def) have dist2: "dist (1 / 2 + min (1 / 2) (min d1 d2) / 4) (1 / 2) < d2" using \d1 > 0\ \d2 > 0\ by (simp add: min_def dist_norm) have [simp]: "\ min (1 / 2) (min d1 d2) \ 0" using \d1 > 0\ \d2 > 0\ by (simp add: min_def) have "dist (g2 (min (1 / 2) (min d1 d2) / 2)) (g1 1) < e/2" "dist (g2 (min (1 / 2) (min d1 d2) / 2)) (g2 0) < e/2" 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 by (simp add: e_def [symmetric]) qed 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: assumes "simple_path(g1 +++ g2)" and "pathfinish g1 = pathstart 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\ \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" using assms by (simp add: simple_path_def) 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) fix x y assume "g1 x = g1 y" "0 \ x" "x \ 1" "0 \ y" "y \ 1" then show "x = y" using * [of "x/2" "y/2"] by (simp add: joinpaths_def split_ifs) qed 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) have "path g2" using assms path_join simple_path_imp_path by blast moreover have "inj_on g2 {0..1}" proof (clarsimp simp: inj_on_def) fix x y assume "g2 x = g2 y" "0 \ x" "x \ 1" "0 \ y" "y \ 1" then show "x = y" using * [of "(x + 1) / 2" "(y + 1) / 2"] by (force simp: joinpaths_def split_ifs field_split_simps) qed ultimately have "arc g2" using assms by (simp add: arc_def) 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 field_split_simps) then have "path_image g1 \ path_image g2 \ {pathstart g1, pathstart g2}" by (fastforce simp: pathstart_def pathfinish_def path_image_def) with \arc g1\ \arc g2\ show ?thesis using that by blast qed lemma simple_path_join_loop_eq: 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" shows "arc(g1 +++ g2) \ arc g1 \ arc g2 \ path_image g1 \ path_image g2 \ {pathstart g2}" (is "?lhs = ?rhs") 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\ \ 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 using * [of 0 "(u + 1) / 2"] that assms arc_distinct_ends [OF \?lhs\] by (auto simp: joinpaths_def pathstart_def pathfinish_def split_ifs field_split_simps) then have n1: "pathstart g1 \ path_image g2" unfolding pathstart_def path_image_def using atLeastAtMost_iff by blast show ?rhs using \?lhs\ using \simple_path (g1 +++ g2)\ assms n1 simple_path_joinE by auto next assume ?rhs then show ?lhs using assms by (fastforce simp: pathfinish_def pathstart_def intro!: arc_join) qed lemma arc_join_eq_alt: "pathfinish g1 = pathstart g2 \ (arc(g1 +++ g2) \ arc g1 \ arc g2 \ path_image g1 \ path_image g2 = {pathstart g2})" using pathfinish_in_path_image by (fastforce simp: arc_join_eq) subsection\<^marker>\tag unimportant\\The joining of paths is associative\ lemma path_assoc: "\pathfinish p = pathstart q; pathfinish q = pathstart r\ \ path(p +++ (q +++ r)) \ path((p +++ q) +++ r)" by simp 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 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 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 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 dest!: subsetD [OF _ IntI]) qed next case False { fix x :: 'a assume a: "path_image p \ path_image q \ {pathstart q}" "(path_image p \ path_image q) \ path_image r \ {pathstart r}" "x \ path_image p" "x \ path_image r" have "pathstart r \ path_image q" by (metis assms(2) pathfinish_in_path_image) with a have "x = pathstart q" by blast } 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: "\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) subsubsection\<^marker>\tag unimportant\\Symmetry and loops\ lemma path_sym: "\pathfinish p = pathstart q; pathfinish q = pathstart p\ \ path(p +++ q) \ path(q +++ p)" by auto lemma simple_path_sym: "\pathfinish p = pathstart q; pathfinish q = pathstart p\ \ simple_path(p +++ q) \ simple_path(q +++ p)" by (metis (full_types) inf_commute insert_commute simple_path_joinE simple_path_join_loop) lemma path_image_sym: "\pathfinish p = pathstart q; pathfinish q = pathstart p\ \ path_image(p +++ q) = path_image(q +++ p)" by (simp add: path_image_join sup_commute) subsection\Subpath\ definition\<^marker>\tag important\ subpath :: "real \ real \ (real \ 'a) \ real \ 'a::real_normed_vector" where "subpath a b g \ \x. g((b - a) * x + a)" lemma path_image_subpath_gen: fixes g :: "_ \ 'a::real_normed_vector" shows "path_image(subpath u v g) = g ` (closed_segment u v)" by (auto simp add: closed_segment_real_eq path_image_def subpath_def) lemma path_image_subpath: fixes g :: "real \ 'a::real_normed_vector" shows "path_image(subpath u v g) = (if u \ v then g ` {u..v} else g ` {v..u})" by (simp add: path_image_subpath_gen closed_segment_eq_real_ivl) lemma path_image_subpath_commute: fixes g :: "real \ 'a::real_normed_vector" shows "path_image(subpath u v g) = path_image(subpath v u g)" by (simp add: path_image_subpath_gen closed_segment_eq_real_ivl) lemma path_subpath [simp]: fixes g :: "real \ 'a::real_normed_vector" assumes "path g" "u \ {0..1}" "v \ {0..1}" shows "path(subpath u v g)" proof - have "continuous_on {0..1} (g \ (\x. ((v-u) * x+ u)))" using assms apply (intro continuous_intros; simp add: image_affinity_atLeastAtMost [where c=u]) apply (auto simp: path_def continuous_on_subset) done then show ?thesis by (simp add: path_def subpath_def) qed lemma pathstart_subpath [simp]: "pathstart(subpath u v g) = g(u)" by (simp add: pathstart_def subpath_def) lemma pathfinish_subpath [simp]: "pathfinish(subpath u v g) = g(v)" by (simp add: pathfinish_def subpath_def) lemma subpath_trivial [simp]: "subpath 0 1 g = g" by (simp add: subpath_def) lemma subpath_reversepath: "subpath 1 0 g = reversepath g" by (simp add: reversepath_def subpath_def) lemma reversepath_subpath: "reversepath(subpath u v g) = subpath v u g" by (simp add: reversepath_def subpath_def algebra_simps) lemma subpath_translation: "subpath u v ((\x. a + x) \ g) = (\x. a + x) \ subpath u v g" by (rule ext) (simp add: subpath_def) lemma subpath_image: "subpath u v (f \ g) = f \ subpath u v g" by (rule ext) (simp add: subpath_def) lemma affine_ineq: fixes x :: "'a::linordered_idom" assumes "x \ 1" "v \ u" shows "v + x * u \ u + x * v" proof - have "(1-x)*(u-v) \ 0" using assms by auto then show ?thesis by (simp add: algebra_simps) qed lemma sum_le_prod1: fixes a::real shows "\a \ 1; b \ 1\ \ a + b \ 1 + a * b" by (metis add.commute affine_ineq mult.right_neutral) lemma simple_path_subpath_eq: "simple_path(subpath u v g) \ path(subpath u v g) \ u\v \ (\x y. x \ closed_segment u v \ y \ closed_segment u v \ g x = g y \ x = y \ x = u \ y = v \ x = v \ y = u)" (is "?lhs = ?rhs") proof assume ?lhs then have p: "path (\x. g ((v - u) * x + u))" and sim: "(\x y. \x\{0..1}; y\{0..1}; g ((v - u) * x + u) = g ((v - u) * y + u)\ \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0)" by (auto simp: simple_path_def subpath_def) { fix x y assume "x \ closed_segment u v" "y \ closed_segment u v" "g x = g y" then have "x = y \ x = u \ y = v \ x = v \ y = u" using sim [of "(x-u)/(v-u)" "(y-u)/(v-u)"] p by (auto split: if_split_asm simp add: closed_segment_real_eq image_affinity_atLeastAtMost) (simp_all add: field_split_simps) } moreover have "path(subpath u v g) \ u\v" using sim [of "1/3" "2/3"] p by (auto simp: subpath_def) ultimately show ?rhs by metis next assume ?rhs then have d1: "\x y. \g x = g y; u \ x; x \ v; u \ y; y \ v\ \ x = y \ x = u \ y = v \ x = v \ y = u" and d2: "\x y. \g x = g y; v \ x; x \ u; v \ y; y \ u\ \ x = y \ x = u \ y = v \ x = v \ y = u" and ne: "u < v \ v < u" and psp: "path (subpath u v g)" by (auto simp: closed_segment_real_eq image_affinity_atLeastAtMost) have [simp]: "\x. u + x * v = v + x * u \ u=v \ x=1" by algebra show ?lhs using psp ne unfolding simple_path_def subpath_def by (fastforce simp add: algebra_simps affine_ineq mult_left_mono crossproduct_eq dest: d1 d2) qed lemma arc_subpath_eq: "arc(subpath u v g) \ path(subpath u v g) \ u\v \ inj_on g (closed_segment u v)" (is "?lhs = ?rhs") proof assume ?lhs then have p: "path (\x. g ((v - u) * x + u))" and sim: "(\x y. \x\{0..1}; y\{0..1}; g ((v - u) * x + u) = g ((v - u) * y + u)\ \ x = y)" by (auto simp: arc_def inj_on_def subpath_def) { fix x y assume "x \ closed_segment u v" "y \ closed_segment u v" "g x = g y" then have "x = y" using sim [of "(x-u)/(v-u)" "(y-u)/(v-u)"] p by (cases "v = u") (simp_all split: if_split_asm add: inj_on_def closed_segment_real_eq image_affinity_atLeastAtMost, simp add: field_simps) } moreover have "path(subpath u v g) \ u\v" using sim [of "1/3" "2/3"] p by (auto simp: subpath_def) ultimately show ?rhs unfolding inj_on_def by metis next assume ?rhs then have d1: "\x y. \g x = g y; u \ x; x \ v; u \ y; y \ v\ \ x = y" and d2: "\x y. \g x = g y; v \ x; x \ u; v \ y; y \ u\ \ x = y" and ne: "u < v \ v < u" and psp: "path (subpath u v g)" by (auto simp: inj_on_def closed_segment_real_eq image_affinity_atLeastAtMost) show ?lhs using psp ne unfolding arc_def subpath_def inj_on_def by (auto simp: algebra_simps affine_ineq mult_left_mono crossproduct_eq dest: d1 d2) qed lemma simple_path_subpath: assumes "simple_path g" "u \ {0..1}" "v \ {0..1}" "u \ v" shows "simple_path(subpath u v g)" using assms apply (simp add: simple_path_subpath_eq simple_path_imp_path) apply (simp add: simple_path_def closed_segment_real_eq image_affinity_atLeastAtMost, fastforce) done lemma arc_simple_path_subpath: "\simple_path g; u \ {0..1}; v \ {0..1}; g u \ g v\ \ arc(subpath u v g)" by (force intro: simple_path_subpath simple_path_imp_arc) lemma arc_subpath_arc: "\arc g; u \ {0..1}; v \ {0..1}; u \ v\ \ arc(subpath u v g)" by (meson arc_def arc_imp_simple_path arc_simple_path_subpath inj_onD) lemma arc_simple_path_subpath_interior: "\simple_path g; u \ {0..1}; v \ {0..1}; u \ v; \u-v\ < 1\ \ arc(subpath u v g)" by (force simp: simple_path_def intro: arc_simple_path_subpath) lemma path_image_subpath_subset: "\u \ {0..1}; v \ {0..1}\ \ path_image(subpath u v g) \ path_image g" by (metis atLeastAtMost_iff atLeastatMost_subset_iff path_image_def path_image_subpath subset_image_iff) lemma join_subpaths_middle: "subpath (0) ((1 / 2)) p +++ subpath ((1 / 2)) 1 p = p" by (rule ext) (simp add: joinpaths_def subpath_def field_split_simps) subsection\<^marker>\tag unimportant\\There is a subpath to the frontier\ lemma subpath_to_frontier_explicit: fixes S :: "'a::metric_space set" assumes g: "path g" and "pathfinish g \ S" obtains u where "0 \ u" "u \ 1" "\x. 0 \ x \ x < u \ g x \ interior S" "(g u \ interior S)" "(u = 0 \ g u \ closure S)" proof - have gcon: "continuous_on {0..1} g" using g by (simp add: path_def) moreover have "bounded ({u. g u \ closure (- S)} \ {0..1})" using compact_eq_bounded_closed by fastforce ultimately have com: "compact ({0..1} \ {u. g u \ closure (- S)})" using closed_vimage_Int by (metis (full_types) Int_commute closed_atLeastAtMost closed_closure compact_eq_bounded_closed vimage_def) have "1 \ {u. g u \ closure (- S)}" using assms by (simp add: pathfinish_def closure_def) then have dis: "{0..1} \ {u. g u \ closure (- S)} \ {}" using atLeastAtMost_iff zero_le_one by blast then obtain u where "0 \ u" "u \ 1" and gu: "g u \ closure (- S)" and umin: "\t. \0 \ t; t \ 1; g t \ closure (- S)\ \ u \ t" using compact_attains_inf [OF com dis] by fastforce then have umin': "\t. \0 \ t; t \ 1; t < u\ \ g t \ S" using closure_def by fastforce have \
: "g u \ closure S" if "u \ 0" proof - have "u > 0" using that \0 \ u\ by auto { fix e::real assume "e > 0" obtain d where "d>0" and d: "\x'. \x' \ {0..1}; dist x' u \ d\ \ dist (g x') (g u) < e" using continuous_onE [OF gcon _ \e > 0\] \0 \ _\ \_ \ 1\ atLeastAtMost_iff by auto have *: "dist (max 0 (u - d / 2)) u \ d" using \0 \ u\ \u \ 1\ \d > 0\ by (simp add: dist_real_def) have "\y\S. dist y (g u) < e" using \0 < u\ \u \ 1\ \d > 0\ by (force intro: d [OF _ *] umin') } then show ?thesis by (simp add: frontier_def closure_approachable) qed show ?thesis proof show "\x. 0 \ x \ x < u \ g x \ interior S" using \u \ 1\ interior_closure umin by fastforce show "g u \ interior S" by (simp add: gu interior_closure) qed (use \0 \ u\ \u \ 1\ \
in auto) qed lemma subpath_to_frontier_strong: assumes g: "path g" and "pathfinish g \ S" obtains u where "0 \ u" "u \ 1" "g u \ interior S" "u = 0 \ (\x. 0 \ x \ x < 1 \ subpath 0 u g x \ interior S) \ g u \ closure S" proof - obtain u where "0 \ u" "u \ 1" and gxin: "\x. 0 \ x \ x < u \ g x \ interior S" and gunot: "(g u \ interior S)" and u0: "(u = 0 \ g u \ closure S)" using subpath_to_frontier_explicit [OF assms] by blast show ?thesis proof show "g u \ interior S" using gunot by blast qed (use \0 \ u\ \u \ 1\ u0 in \(force simp: subpath_def gxin)+\) qed lemma subpath_to_frontier: assumes g: "path g" and g0: "pathstart g \ closure S" and g1: "pathfinish g \ S" obtains u where "0 \ u" "u \ 1" "g u \ frontier S" "path_image(subpath 0 u g) - {g u} \ interior S" proof - obtain u where "0 \ u" "u \ 1" and notin: "g u \ interior S" and disj: "u = 0 \ (\x. 0 \ x \ x < 1 \ subpath 0 u g x \ interior S) \ g u \ closure S" (is "_ \ ?P") using subpath_to_frontier_strong [OF g g1] by blast show ?thesis proof show "g u \ frontier S" by (metis DiffI disj frontier_def g0 notin pathstart_def) show "path_image (subpath 0 u g) - {g u} \ interior S" using disj proof assume "u = 0" then show ?thesis by (simp add: path_image_subpath) next assume P: ?P show ?thesis proof (clarsimp simp add: path_image_subpath_gen) fix y assume y: "y \ closed_segment 0 u" "g y \ interior S" with \0 \ u\ have "0 \ y" "y \ u" by (auto simp: closed_segment_eq_real_ivl split: if_split_asm) then have "y=u \ subpath 0 u g (y/u) \ interior S" using P less_eq_real_def by force then show "g y = g u" using y by (auto simp: subpath_def split: if_split_asm) qed qed qed (use \0 \ u\ \u \ 1\ in auto) qed lemma exists_path_subpath_to_frontier: fixes S :: "'a::real_normed_vector set" assumes "path g" "pathstart g \ closure S" "pathfinish g \ S" obtains h where "path h" "pathstart h = pathstart g" "path_image h \ path_image g" "path_image h - {pathfinish h} \ interior S" "pathfinish h \ frontier S" proof - obtain u where u: "0 \ u" "u \ 1" "g u \ frontier S" "(path_image(subpath 0 u g) - {g u}) \ interior S" using subpath_to_frontier [OF assms] by blast show ?thesis proof show "path_image (subpath 0 u g) \ path_image g" by (simp add: path_image_subpath_subset u) show "pathstart (subpath 0 u g) = pathstart g" by (metis pathstart_def pathstart_subpath) qed (use assms u in \auto simp: path_image_subpath\) qed lemma exists_path_subpath_to_frontier_closed: fixes S :: "'a::real_normed_vector set" assumes S: "closed S" and g: "path g" and g0: "pathstart g \ S" and g1: "pathfinish g \ S" obtains h where "path h" "pathstart h = pathstart g" "path_image h \ path_image g \ S" "pathfinish h \ frontier S" proof - obtain h where h: "path h" "pathstart h = pathstart g" "path_image h \ path_image g" "path_image h - {pathfinish h} \ interior S" "pathfinish h \ frontier S" using exists_path_subpath_to_frontier [OF g _ g1] closure_closed [OF S] g0 by auto show ?thesis proof show "path_image h \ path_image g \ S" using assms h interior_subset [of S] by (auto simp: frontier_def) qed (use h in auto) qed subsection \Shift Path to Start at Some Given Point\ definition\<^marker>\tag important\ shiftpath :: "real \ (real \ 'a::topological_space) \ real \ 'a" where "shiftpath a f = (\x. if (a + x) \ 1 then f (a + x) else f (a + x - 1))" 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 pathstart_shiftpath: "a \ 1 \ pathstart (shiftpath a g) = g a" unfolding pathstart_def shiftpath_def by auto lemma pathfinish_shiftpath: assumes "0 \ a" and "pathfinish g = pathstart g" shows "pathfinish (shiftpath a g) = g a" using assms unfolding pathstart_def pathfinish_def shiftpath_def by auto lemma endpoints_shiftpath: assumes "pathfinish g = pathstart g" and "a \ {0 .. 1}" shows "pathfinish (shiftpath a g) = g a" and "pathstart (shiftpath a g) = g a" using assms by (auto intro!: pathfinish_shiftpath pathstart_shiftpath) lemma closed_shiftpath: assumes "pathfinish g = pathstart g" and "a \ {0..1}" shows "pathfinish (shiftpath a g) = pathstart (shiftpath a g)" using endpoints_shiftpath[OF assms] by auto lemma path_shiftpath: assumes "path g" and "pathfinish g = pathstart g" and "a \ {0..1}" shows "path (shiftpath a g)" proof - have *: "{0 .. 1} = {0 .. 1-a} \ {1-a .. 1}" using assms(3) by auto have **: "\x. x + a = 1 \ g (x + a - 1) = g (x + a)" using assms(2)[unfolded pathfinish_def pathstart_def] by auto show ?thesis unfolding path_def shiftpath_def * proof (rule continuous_on_closed_Un) have contg: "continuous_on {0..1} g" using \path g\ path_def by blast show "continuous_on {0..1-a} (\x. if a + x \ 1 then g (a + x) else g (a + x - 1))" proof (rule continuous_on_eq) show "continuous_on {0..1-a} (g \ (+) a)" by (intro continuous_intros continuous_on_subset [OF contg]) (use \a \ {0..1}\ in auto) qed auto show "continuous_on {1-a..1} (\x. if a + x \ 1 then g (a + x) else g (a + x - 1))" proof (rule continuous_on_eq) show "continuous_on {1-a..1} (g \ (+) (a - 1))" by (intro continuous_intros continuous_on_subset [OF contg]) (use \a \ {0..1}\ in auto) qed (auto simp: "**" add.commute add_diff_eq) qed auto qed lemma shiftpath_shiftpath: assumes "pathfinish g = pathstart g" and "a \ {0..1}" and "x \ {0..1}" shows "shiftpath (1 - a) (shiftpath a g) x = g x" using assms unfolding pathfinish_def pathstart_def shiftpath_def by auto lemma path_image_shiftpath: assumes a: "a \ {0..1}" and "pathfinish g = pathstart g" shows "path_image (shiftpath a g) = path_image g" proof - { fix x assume g: "g 1 = g 0" "x \ {0..1::real}" and gne: "\y. y\{0..1} \ {x. \ a + x \ 1} \ g x \ g (a + y - 1)" then have "\y\{0..1} \ {x. a + x \ 1}. g x = g (a + y)" proof (cases "a \ x") case False then show ?thesis apply (rule_tac x="1 + x - a" in bexI) using g gne[of "1 + x - a"] a by (force simp: field_simps)+ next case True then show ?thesis using g a by (rule_tac x="x - a" in bexI) (auto simp: field_simps) qed } then show ?thesis using assms unfolding shiftpath_def path_image_def pathfinish_def pathstart_def by (auto simp: image_iff) qed lemma simple_path_shiftpath: assumes "simple_path g" "pathfinish g = pathstart g" and a: "0 \ a" "a \ 1" shows "simple_path (shiftpath a g)" unfolding simple_path_def proof (intro conjI impI ballI) show "path (shiftpath a g)" by (simp add: assms path_shiftpath simple_path_imp_path) have *: "\x y. \g x = g y; x \ {0..1}; y \ {0..1}\ \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" using assms by (simp add: simple_path_def) show "x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" if "x \ {0..1}" "y \ {0..1}" "shiftpath a g x = shiftpath a g y" for x y using that a unfolding shiftpath_def by (force split: if_split_asm dest!: *) qed subsection \Straight-Line Paths\ definition\<^marker>\tag important\ linepath :: "'a::real_normed_vector \ 'a \ real \ 'a" where "linepath a b = (\x. (1 - x) *\<^sub>R a + x *\<^sub>R b)" lemma pathstart_linepath[simp]: "pathstart (linepath a b) = a" unfolding pathstart_def linepath_def by auto lemma pathfinish_linepath[simp]: "pathfinish (linepath a b) = b" unfolding pathfinish_def linepath_def by auto lemma linepath_inner: "linepath a b x \ v = linepath (a \ v) (b \ v) x" by (simp add: linepath_def algebra_simps) lemma Re_linepath': "Re (linepath a b x) = linepath (Re a) (Re b) x" by (simp add: linepath_def) lemma Im_linepath': "Im (linepath a b x) = linepath (Im a) (Im b) x" by (simp add: linepath_def) lemma linepath_0': "linepath a b 0 = a" by (simp add: linepath_def) lemma linepath_1': "linepath a b 1 = b" by (simp add: linepath_def) lemma continuous_linepath_at[intro]: "continuous (at x) (linepath a b)" unfolding linepath_def by (intro continuous_intros) lemma continuous_on_linepath [intro,continuous_intros]: "continuous_on s (linepath a b)" using continuous_linepath_at by (auto intro!: continuous_at_imp_continuous_on) lemma path_linepath[iff]: "path (linepath a b)" unfolding path_def by (rule continuous_on_linepath) lemma path_image_linepath[simp]: "path_image (linepath a b) = closed_segment a b" unfolding path_image_def segment linepath_def by auto lemma reversepath_linepath[simp]: "reversepath (linepath a b) = linepath b a" unfolding reversepath_def linepath_def by auto lemma linepath_0 [simp]: "linepath 0 b x = x *\<^sub>R b" by (simp add: linepath_def) lemma linepath_cnj: "cnj (linepath a b x) = linepath (cnj a) (cnj b) x" by (simp add: linepath_def) lemma arc_linepath: assumes "a \ b" shows [simp]: "arc (linepath a b)" proof - { fix x y :: "real" assume "x *\<^sub>R b + y *\<^sub>R a = x *\<^sub>R a + y *\<^sub>R b" then have "(x - y) *\<^sub>R a = (x - y) *\<^sub>R b" by (simp add: algebra_simps) with assms have "x = y" by simp } then show ?thesis unfolding arc_def inj_on_def by (fastforce simp: algebra_simps linepath_def) qed lemma simple_path_linepath[intro]: "a \ b \ simple_path (linepath a b)" by (simp add: arc_imp_simple_path) lemma linepath_trivial [simp]: "linepath a a x = a" by (simp add: linepath_def real_vector.scale_left_diff_distrib) lemma linepath_refl: "linepath a a = (\x. a)" by auto lemma subpath_refl: "subpath a a g = linepath (g a) (g a)" by (simp add: subpath_def linepath_def algebra_simps) lemma linepath_of_real: "(linepath (of_real a) (of_real b) x) = of_real ((1 - x)*a + x*b)" by (simp add: scaleR_conv_of_real linepath_def) lemma of_real_linepath: "of_real (linepath a b x) = linepath (of_real a) (of_real b) x" by (metis linepath_of_real mult.right_neutral of_real_def real_scaleR_def) lemma inj_on_linepath: assumes "a \ b" shows "inj_on (linepath a b) {0..1}" proof (clarsimp simp: inj_on_def linepath_def) fix x y assume "(1 - x) *\<^sub>R a + x *\<^sub>R b = (1 - y) *\<^sub>R a + y *\<^sub>R b" "0 \ x" "x \ 1" "0 \ y" "y \ 1" then have "x *\<^sub>R (a - b) = y *\<^sub>R (a - b)" by (auto simp: algebra_simps) then show "x=y" using assms by auto qed lemma linepath_le_1: fixes a::"'a::linordered_idom" shows "\a \ 1; b \ 1; 0 \ u; u \ 1\ \ (1 - u) * a + u * b \ 1" using mult_left_le [of a "1-u"] mult_left_le [of b u] by auto 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" proof - have "linepath a b x \ closed_segment a b" using x by (auto simp flip: linepath_image_01) then show ?thesis using a b convex_contains_segment by blast qed 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 bounded_linear_linepath: assumes "bounded_linear f" shows "f (linepath a b x) = linepath (f a) (f b) x" proof - interpret f: bounded_linear f by fact show ?thesis by (simp add: linepath_def f.add f.scale) qed lemma bounded_linear_linepath': assumes "bounded_linear f" shows "f \ linepath a b = linepath (f a) (f b)" using bounded_linear_linepath[OF assms] by (simp add: fun_eq_iff) lemma linepath_cnj': "cnj \ linepath a b = linepath (cnj a) (cnj b)" by (simp add: linepath_def fun_eq_iff) lemma differentiable_linepath [intro]: "linepath a b differentiable at x within A" by (auto simp: linepath_def) lemma has_vector_derivative_linepath_within: "(linepath a b has_vector_derivative (b - a)) (at x within S)" by (force intro: derivative_eq_intros simp add: linepath_def has_vector_derivative_def algebra_simps) subsection\<^marker>\tag unimportant\\Segments via convex hulls\ lemma segments_subset_convex_hull: "closed_segment a b \ (convex hull {a,b,c})" "closed_segment a c \ (convex hull {a,b,c})" "closed_segment b c \ (convex hull {a,b,c})" "closed_segment b a \ (convex hull {a,b,c})" "closed_segment c a \ (convex hull {a,b,c})" "closed_segment c b \ (convex hull {a,b,c})" by (auto simp: segment_convex_hull linepath_of_real elim!: rev_subsetD [OF _ hull_mono]) lemma midpoints_in_convex_hull: assumes "x \ convex hull s" "y \ convex hull s" shows "midpoint x y \ convex hull s" proof - have "(1 - inverse(2)) *\<^sub>R x + inverse(2) *\<^sub>R y \ convex hull s" by (rule convexD_alt) (use assms in auto) then show ?thesis by (simp add: midpoint_def algebra_simps) qed lemma not_in_interior_convex_hull_3: fixes a :: "complex" shows "a \ interior(convex hull {a,b,c})" "b \ interior(convex hull {a,b,c})" "c \ interior(convex hull {a,b,c})" by (auto simp: card_insert_le_m1 not_in_interior_convex_hull) lemma midpoint_in_closed_segment [simp]: "midpoint a b \ closed_segment a b" using midpoints_in_convex_hull segment_convex_hull by blast lemma midpoint_in_open_segment [simp]: "midpoint a b \ open_segment a b \ a \ b" by (simp add: open_segment_def) lemma continuous_IVT_local_extremum: fixes f :: "'a::euclidean_space \ real" assumes contf: "continuous_on (closed_segment a b) f" and "a \ b" "f a = f b" obtains z where "z \ open_segment a b" "(\w \ closed_segment a b. (f w) \ (f z)) \ (\w \ closed_segment a b. (f z) \ (f w))" proof - obtain c where "c \ closed_segment a b" and c: "\y. y \ closed_segment a b \ f y \ f c" using continuous_attains_sup [of "closed_segment a b" f] contf by auto obtain d where "d \ closed_segment a b" and d: "\y. y \ closed_segment a b \ f d \ f y" using continuous_attains_inf [of "closed_segment a b" f] contf by auto show ?thesis proof (cases "c \ open_segment a b \ d \ open_segment a b") case True then show ?thesis using c d that by blast next case False then have "(c = a \ c = b) \ (d = a \ d = b)" by (simp add: \c \ closed_segment a b\ \d \ closed_segment a b\ open_segment_def) with \a \ b\ \f a = f b\ c d show ?thesis by (rule_tac z = "midpoint a b" in that) (fastforce+) qed qed text\An injective map into R is also an open map w.r.T. the universe, and conversely. \ proposition injective_eq_1d_open_map_UNIV: fixes f :: "real \ real" assumes contf: "continuous_on S f" and S: "is_interval S" shows "inj_on f S \ (\T. open T \ T \ S \ open(f ` T))" (is "?lhs = ?rhs") proof safe fix T assume injf: ?lhs and "open T" and "T \ S" have "\U. open U \ f x \ U \ U \ f ` T" if "x \ T" for x proof - obtain \ where "\ > 0" and \: "cball x \ \ T" using \open T\ \x \ T\ open_contains_cball_eq by blast show ?thesis proof (intro exI conjI) have "closed_segment (x-\) (x+\) = {x-\..x+\}" using \0 < \\ by (auto simp: closed_segment_eq_real_ivl) also have "\ \ S" using \ \T \ S\ by (auto simp: dist_norm subset_eq) finally have "f ` (open_segment (x-\) (x+\)) = open_segment (f (x-\)) (f (x+\))" using continuous_injective_image_open_segment_1 by (metis continuous_on_subset [OF contf] inj_on_subset [OF injf]) then show "open (f ` {x-\<..})" using \0 < \\ by (simp add: open_segment_eq_real_ivl) show "f x \ f ` {x - \<..}" by (auto simp: \\ > 0\) show "f ` {x - \<..} \ f ` T" using \ by (auto simp: dist_norm subset_iff) qed qed with open_subopen show "open (f ` T)" by blast next assume R: ?rhs have False if xy: "x \ S" "y \ S" and "f x = f y" "x \ y" for x y proof - have "open (f ` open_segment x y)" using R by (metis S convex_contains_open_segment is_interval_convex open_greaterThanLessThan open_segment_eq_real_ivl xy) moreover have "continuous_on (closed_segment x y) f" by (meson S closed_segment_subset contf continuous_on_subset is_interval_convex that) then obtain \ where "\ \ open_segment x y" and \: "(\w \ closed_segment x y. (f w) \ (f \)) \ (\w \ closed_segment x y. (f \) \ (f w))" using continuous_IVT_local_extremum [of x y f] \f x = f y\ \x \ y\ by blast ultimately obtain e where "e>0" and e: "\u. dist u (f \) < e \ u \ f ` open_segment x y" using open_dist by (metis image_eqI) have fin: "f \ + (e/2) \ f ` open_segment x y" "f \ - (e/2) \ f ` open_segment x y" using e [of "f \ + (e/2)"] e [of "f \ - (e/2)"] \e > 0\ by (auto simp: dist_norm) show ?thesis using \ \0 < e\ fin open_closed_segment by fastforce qed then show ?lhs by (force simp: inj_on_def) qed subsection\<^marker>\tag unimportant\ \Bounding a point away from a path\ lemma not_on_path_ball: fixes g :: "real \ 'a::heine_borel" assumes "path g" and z: "z \ path_image g" shows "\e > 0. ball z e \ path_image g = {}" proof - have "closed (path_image g)" by (simp add: \path g\ closed_path_image) then obtain a where "a \ path_image g" "\y \ path_image g. dist z a \ dist z y" by (auto intro: distance_attains_inf[OF _ path_image_nonempty, of g z]) then show ?thesis by (rule_tac x="dist z a" in exI) (use dist_commute z in auto) qed lemma not_on_path_cball: fixes g :: "real \ 'a::heine_borel" assumes "path g" and "z \ path_image g" shows "\e>0. cball z e \ (path_image g) = {}" proof - obtain e where "ball z e \ path_image g = {}" "e > 0" using not_on_path_ball[OF assms] by auto moreover have "cball z (e/2) \ ball z e" using \e > 0\ by auto ultimately show ?thesis by (rule_tac x="e/2" in exI) auto qed subsection \Path component\ text \Original formalization by Tom Hales\ definition\<^marker>\tag important\ "path_component S x y \ (\g. path g \ path_image g \ S \ pathstart g = x \ pathfinish g = y)" abbreviation\<^marker>\tag important\ "path_component_set S x \ Collect (path_component S x)" lemmas path_defs = path_def pathstart_def pathfinish_def path_image_def path_component_def lemma path_component_mem: assumes "path_component S x y" shows "x \ S" and "y \ S" using assms unfolding path_defs by auto lemma path_component_refl: assumes "x \ S" shows "path_component S x x" using assms unfolding path_defs by (metis (full_types) assms continuous_on_const image_subset_iff path_image_def) lemma path_component_refl_eq: "path_component S x x \ x \ S" by (auto intro!: path_component_mem path_component_refl) lemma path_component_sym: "path_component S x y \ path_component S y x" unfolding path_component_def by (metis (no_types) path_image_reversepath path_reversepath pathfinish_reversepath pathstart_reversepath) lemma path_component_trans: assumes "path_component S x y" and "path_component S y z" shows "path_component S x z" using assms unfolding path_component_def by (metis path_join pathfinish_join pathstart_join subset_path_image_join) lemma path_component_of_subset: "S \ T \ path_component S x y \ path_component T x y" unfolding path_component_def by auto lemma path_component_linepath: fixes S :: "'a::real_normed_vector set" shows "closed_segment a b \ S \ path_component S a b" unfolding path_component_def by (rule_tac x="linepath a b" in exI, auto) subsubsection\<^marker>\tag unimportant\ \Path components as sets\ lemma path_component_set: "path_component_set S x = {y. (\g. path g \ path_image g \ S \ pathstart g = x \ pathfinish g = y)}" by (auto simp: path_component_def) lemma path_component_subset: "path_component_set S x \ S" by (auto simp: path_component_mem(2)) lemma path_component_eq_empty: "path_component_set S x = {} \ x \ S" using path_component_mem path_component_refl_eq by fastforce lemma path_component_mono: "S \ T \ (path_component_set S x) \ (path_component_set T x)" by (simp add: Collect_mono path_component_of_subset) lemma path_component_eq: "y \ path_component_set S x \ path_component_set S y = path_component_set S x" by (metis (no_types, lifting) Collect_cong mem_Collect_eq path_component_sym path_component_trans) subsection \Path connectedness of a space\ definition\<^marker>\tag important\ "path_connected S \ (\x\S. \y\S. \g. path g \ path_image g \ S \ pathstart g = x \ pathfinish g = y)" lemma path_connectedin_iff_path_connected_real [simp]: "path_connectedin euclideanreal S \ path_connected S" by (simp add: path_connectedin path_connected_def path_defs) lemma path_connected_component: "path_connected S \ (\x\S. \y\S. path_component S x y)" unfolding path_connected_def path_component_def by auto lemma path_connected_component_set: "path_connected S \ (\x\S. path_component_set S x = S)" unfolding path_connected_component path_component_subset using path_component_mem by blast lemma path_component_maximal: "\x \ T; path_connected T; T \ S\ \ T \ (path_component_set S x)" by (metis path_component_mono path_connected_component_set) lemma convex_imp_path_connected: fixes S :: "'a::real_normed_vector set" assumes "convex S" shows "path_connected S" unfolding path_connected_def using assms convex_contains_segment by fastforce lemma path_connected_UNIV [iff]: "path_connected (UNIV :: 'a::real_normed_vector set)" by (simp add: convex_imp_path_connected) lemma path_component_UNIV: "path_component_set UNIV x = (UNIV :: 'a::real_normed_vector set)" using path_connected_component_set by auto lemma path_connected_imp_connected: assumes "path_connected S" shows "connected S" proof (rule connectedI) fix e1 e2 assume as: "open e1" "open e2" "S \ e1 \ e2" "e1 \ e2 \ S = {}" "e1 \ S \ {}" "e2 \ S \ {}" then obtain x1 x2 where obt:"x1 \ e1 \ S" "x2 \ e2 \ S" by auto then obtain g where g: "path g" "path_image g \ S" "pathstart g = x1" "pathfinish g = x2" using assms[unfolded path_connected_def,rule_format,of x1 x2] by auto have *: "connected {0..1::real}" by (auto intro!: convex_connected) have "{0..1} \ {x \ {0..1}. g x \ e1} \ {x \ {0..1}. g x \ e2}" using as(3) g(2)[unfolded path_defs] by blast moreover have "{x \ {0..1}. g x \ e1} \ {x \ {0..1}. g x \ e2} = {}" using as(4) g(2)[unfolded path_defs] unfolding subset_eq by auto moreover have "{x \ {0..1}. g x \ e1} \ {} \ {x \ {0..1}. g x \ e2} \ {}" using g(3,4)[unfolded path_defs] using obt by (simp add: ex_in_conv [symmetric], metis zero_le_one order_refl) ultimately show False using *[unfolded connected_local not_ex, rule_format, of "{0..1} \ g -` e1" "{0..1} \ g -` e2"] using continuous_openin_preimage_gen[OF g(1)[unfolded path_def] as(1)] using continuous_openin_preimage_gen[OF g(1)[unfolded path_def] as(2)] by auto qed lemma open_path_component: fixes S :: "'a::real_normed_vector set" assumes "open S" shows "open (path_component_set S x)" unfolding open_contains_ball proof fix y assume as: "y \ path_component_set S x" then have "y \ S" by (simp add: path_component_mem(2)) then obtain e where e: "e > 0" "ball y e \ S" using assms openE by blast have "\u. dist y u < e \ path_component S x u" by (metis (full_types) as centre_in_ball convex_ball convex_imp_path_connected e mem_Collect_eq mem_ball path_component_eq path_component_of_subset path_connected_component) then show "\e > 0. ball y e \ path_component_set S x" using \e>0\ by auto qed lemma open_non_path_component: fixes S :: "'a::real_normed_vector set" assumes "open S" shows "open (S - path_component_set S x)" unfolding open_contains_ball proof fix y assume y: "y \ S - path_component_set S x" then obtain e where e: "e > 0" "ball y e \ S" using assms openE by auto show "\e>0. ball y e \ S - path_component_set S x" proof (intro exI conjI subsetI DiffI notI) show "\x. x \ ball y e \ x \ S" using e by blast show False if "z \ ball y e" "z \ path_component_set S x" for z proof - have "y \ path_component_set S z" by (meson assms convex_ball convex_imp_path_connected e open_contains_ball_eq open_path_component path_component_maximal that(1)) then have "y \ path_component_set S x" using path_component_eq that(2) by blast then show False using y by blast qed qed (use e in auto) qed lemma connected_open_path_connected: fixes S :: "'a::real_normed_vector set" assumes "open S" and "connected S" shows "path_connected S" unfolding path_connected_component_set proof (rule, rule, rule path_component_subset, rule) fix x y assume "x \ S" and "y \ S" show "y \ path_component_set S x" proof (rule ccontr) assume "\ ?thesis" moreover have "path_component_set S x \ S \ {}" using \x \ S\ path_component_eq_empty path_component_subset[of S x] by auto ultimately show False using \y \ S\ open_non_path_component[OF assms(1)] open_path_component[OF assms(1)] using assms(2)[unfolded connected_def not_ex, rule_format, of "path_component_set S x" "S - path_component_set S x"] by auto qed qed lemma path_connected_continuous_image: assumes contf: "continuous_on S f" and "path_connected S" shows "path_connected (f ` S)" unfolding path_connected_def proof (rule, rule) fix x' y' assume "x' \ f ` S" "y' \ f ` S" then obtain x y where x: "x \ S" and y: "y \ S" and x': "x' = f x" and y': "y' = f y" by auto from x y obtain g where "path g \ path_image g \ S \ pathstart g = x \ pathfinish g = y" using assms(2)[unfolded path_connected_def] by fast then show "\g. path g \ path_image g \ f ` S \ pathstart g = x' \ pathfinish g = y'" unfolding x' y' path_defs by (fastforce intro: continuous_on_compose continuous_on_subset[OF contf]) qed lemma path_connected_translationI: fixes a :: "'a :: topological_group_add" assumes "path_connected S" shows "path_connected ((\x. a + x) ` S)" by (intro path_connected_continuous_image assms continuous_intros) lemma path_connected_translation: fixes a :: "'a :: topological_group_add" shows "path_connected ((\x. a + x) ` S) = path_connected S" proof - have "\x y. (+) (x::'a) ` (+) (0 - x) ` y = y" by (simp add: image_image) then show ?thesis by (metis (no_types) path_connected_translationI) qed lemma path_connected_segment [simp]: fixes a :: "'a::real_normed_vector" shows "path_connected (closed_segment a b)" by (simp add: convex_imp_path_connected) lemma path_connected_open_segment [simp]: fixes a :: "'a::real_normed_vector" shows "path_connected (open_segment a b)" by (simp add: convex_imp_path_connected) lemma homeomorphic_path_connectedness: "S homeomorphic T \ path_connected S \ path_connected T" unfolding homeomorphic_def homeomorphism_def by (metis path_connected_continuous_image) lemma path_connected_empty [simp]: "path_connected {}" unfolding path_connected_def by auto lemma path_connected_singleton [simp]: "path_connected {a}" unfolding path_connected_def pathstart_def pathfinish_def path_image_def using path_def by fastforce lemma path_connected_Un: assumes "path_connected S" and "path_connected T" and "S \ T \ {}" shows "path_connected (S \ T)" unfolding path_connected_component proof (intro ballI) fix x y assume x: "x \ S \ T" and y: "y \ S \ T" from assms obtain z where z: "z \ S" "z \ T" by auto show "path_component (S \ T) x y" using x y proof safe assume "x \ S" "y \ S" then show "path_component (S \ T) x y" by (meson Un_upper1 \path_connected S\ path_component_of_subset path_connected_component) next assume "x \ S" "y \ T" then show "path_component (S \ T) x y" by (metis z assms(1-2) le_sup_iff order_refl path_component_of_subset path_component_trans path_connected_component) next assume "x \ T" "y \ S" then show "path_component (S \ T) x y" by (metis z assms(1-2) le_sup_iff order_refl path_component_of_subset path_component_trans path_connected_component) next assume "x \ T" "y \ T" then show "path_component (S \ T) x y" by (metis Un_upper1 assms(2) path_component_of_subset path_connected_component sup_commute) qed qed lemma path_connected_UNION: assumes "\i. i \ A \ path_connected (S i)" and "\i. i \ A \ z \ S i" shows "path_connected (\i\A. S i)" unfolding path_connected_component proof clarify fix x i y j assume *: "i \ A" "x \ S i" "j \ A" "y \ S j" then have "path_component (S i) x z" and "path_component (S j) z y" using assms by (simp_all add: path_connected_component) then have "path_component (\i\A. S i) x z" and "path_component (\i\A. S i) z y" using *(1,3) by (auto elim!: path_component_of_subset [rotated]) then show "path_component (\i\A. S i) x y" by (rule path_component_trans) qed lemma path_component_path_image_pathstart: assumes p: "path p" and x: "x \ path_image p" shows "path_component (path_image p) (pathstart p) x" proof - obtain y where x: "x = p y" and y: "0 \ y" "y \ 1" using x by (auto simp: path_image_def) show ?thesis unfolding path_component_def proof (intro exI conjI) have "continuous_on ((*) y ` {0..1}) p" by (simp add: continuous_on_path image_mult_atLeastAtMost_if p y) then have "continuous_on {0..1} (p \ ((*) y))" using continuous_on_compose continuous_on_mult_const by blast then show "path (\u. p (y * u))" by (simp add: path_def) show "path_image (\u. p (y * u)) \ path_image p" using y mult_le_one by (fastforce simp: path_image_def image_iff) qed (auto simp: pathstart_def pathfinish_def x) qed lemma path_connected_path_image: "path p \ path_connected(path_image p)" unfolding path_connected_component by (meson path_component_path_image_pathstart path_component_sym path_component_trans) lemma path_connected_path_component [simp]: "path_connected (path_component_set s x)" proof - { fix y z assume pa: "path_component s x y" "path_component s x z" then have pae: "path_component_set s x = path_component_set s y" using path_component_eq by auto have yz: "path_component s y z" using pa path_component_sym path_component_trans by blast then have "\g. path g \ path_image g \ path_component_set s x \ pathstart g = y \ pathfinish g = z" apply (simp add: path_component_def) by (metis pae path_component_maximal path_connected_path_image pathstart_in_path_image) } then show ?thesis by (simp add: path_connected_def) qed lemma path_component: "path_component S x y \ (\t. path_connected t \ t \ S \ x \ t \ y \ t)" apply (intro iffI) apply (metis path_connected_path_image path_defs(5) pathfinish_in_path_image pathstart_in_path_image) using path_component_of_subset path_connected_component by blast lemma path_component_path_component [simp]: "path_component_set (path_component_set S x) x = path_component_set S x" proof (cases "x \ S") case True show ?thesis by (metis True mem_Collect_eq path_component_refl path_connected_component_set path_connected_path_component) next case False then show ?thesis by (metis False empty_iff path_component_eq_empty) qed lemma path_component_subset_connected_component: "(path_component_set S x) \ (connected_component_set S x)" proof (cases "x \ S") case True show ?thesis by (simp add: True connected_component_maximal path_component_refl path_component_subset path_connected_imp_connected) next case False then show ?thesis using path_component_eq_empty by auto qed subsection\<^marker>\tag unimportant\\Lemmas about path-connectedness\ lemma path_connected_linear_image: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "path_connected S" "bounded_linear f" shows "path_connected(f ` S)" by (auto simp: linear_continuous_on assms path_connected_continuous_image) lemma is_interval_path_connected: "is_interval S \ path_connected S" by (simp add: convex_imp_path_connected is_interval_convex) lemma path_connected_Ioi[simp]: "path_connected {a<..}" for a :: real by (simp add: convex_imp_path_connected) lemma path_connected_Ici[simp]: "path_connected {a..}" for a :: real by (simp add: convex_imp_path_connected) lemma path_connected_Iio[simp]: "path_connected {.. (\x \ topspace X. \y \ topspace X. \S. path_connectedin X S \ x \ S \ y \ S)" by (metis path_connectedin path_connectedin_topspace path_connected_space_def) lemma connectedin_path_image: "pathin X g \ connectedin X (g ` ({0..1}))" by (simp add: path_connectedin_imp_connectedin path_connectedin_path_image) lemma compactin_path_image: "pathin X g \ compactin X (g ` ({0..1}))" unfolding pathin_def by (rule image_compactin [of "top_of_set {0..1}"]) auto lemma linear_homeomorphism_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" obtains g where "homeomorphism (f ` S) S g f" proof - obtain g where "linear g" "g \ f = id" using assms linear_injective_left_inverse by blast then have "homeomorphism (f ` S) S g f" using assms unfolding homeomorphism_def by (auto simp: eq_id_iff [symmetric] image_comp linear_conv_bounded_linear linear_continuous_on) then show thesis .. qed lemma linear_homeomorphic_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "S homeomorphic f ` S" by (meson homeomorphic_def homeomorphic_sym linear_homeomorphism_image [OF assms]) lemma path_connected_Times: assumes "path_connected s" "path_connected t" shows "path_connected (s \ t)" proof (simp add: path_connected_def Sigma_def, clarify) fix x1 y1 x2 y2 assume "x1 \ s" "y1 \ t" "x2 \ s" "y2 \ t" obtain g where "path g" and g: "path_image g \ s" and gs: "pathstart g = x1" and gf: "pathfinish g = x2" using \x1 \ s\ \x2 \ s\ assms by (force simp: path_connected_def) obtain h where "path h" and h: "path_image h \ t" and hs: "pathstart h = y1" and hf: "pathfinish h = y2" using \y1 \ t\ \y2 \ t\ assms by (force simp: path_connected_def) have "path (\z. (x1, h z))" using \path h\ unfolding path_def by (intro continuous_intros continuous_on_compose2 [where g = "Pair _"]; force) moreover have "path (\z. (g z, y2))" using \path g\ unfolding path_def by (intro continuous_intros continuous_on_compose2 [where g = "Pair _"]; force) ultimately have 1: "path ((\z. (x1, h z)) +++ (\z. (g z, y2)))" by (metis hf gs path_join_imp pathstart_def pathfinish_def) have "path_image ((\z. (x1, h z)) +++ (\z. (g z, y2))) \ path_image (\z. (x1, h z)) \ path_image (\z. (g z, y2))" by (rule Path_Connected.path_image_join_subset) also have "\ \ (\x\s. \x1\t. {(x, x1)})" using g h \x1 \ s\ \y2 \ t\ by (force simp: path_image_def) finally have 2: "path_image ((\z. (x1, h z)) +++ (\z. (g z, y2))) \ (\x\s. \x1\t. {(x, x1)})" . show "\g. path g \ path_image g \ (\x\s. \x1\t. {(x, x1)}) \ pathstart g = (x1, y1) \ pathfinish g = (x2, y2)" using 1 2 gf hs by (metis (no_types, lifting) pathfinish_def pathfinish_join pathstart_def pathstart_join) qed lemma is_interval_path_connected_1: fixes s :: "real set" shows "is_interval s \ path_connected s" using is_interval_connected_1 is_interval_path_connected path_connected_imp_connected by blast subsection\<^marker>\tag unimportant\\Path components\ lemma Union_path_component [simp]: "Union {path_component_set S x |x. x \ S} = S" apply (rule subset_antisym) using path_component_subset apply force using path_component_refl by auto lemma path_component_disjoint: "disjnt (path_component_set S a) (path_component_set S b) \ (a \ path_component_set S b)" unfolding disjnt_iff using path_component_sym path_component_trans by blast lemma path_component_eq_eq: "path_component S x = path_component S y \ (x \ S) \ (y \ S) \ x \ S \ y \ S \ path_component S x y" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (metis (no_types) path_component_mem(1) path_component_refl) next assume ?rhs then show ?lhs proof assume "x \ S \ y \ S" then show ?lhs by (metis Collect_empty_eq_bot path_component_eq_empty) next assume S: "x \ S \ y \ S \ path_component S x y" show ?lhs by (rule ext) (metis S path_component_trans path_component_sym) qed qed lemma path_component_unique: assumes "x \ c" "c \ S" "path_connected c" "\c'. \x \ c'; c' \ S; path_connected c'\ \ c' \ c" shows "path_component_set S x = c" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" using assms by (metis mem_Collect_eq path_component_refl path_component_subset path_connected_path_component subsetD) qed (simp add: assms path_component_maximal) lemma path_component_intermediate_subset: "path_component_set u a \ t \ t \ u \ path_component_set t a = path_component_set u a" by (metis (no_types) path_component_mono path_component_path_component subset_antisym) lemma complement_path_component_Union: fixes x :: "'a :: topological_space" shows "S - path_component_set S x = \({path_component_set S y| y. y \ S} - {path_component_set S x})" proof - have *: "(\x. x \ S - {a} \ disjnt a x) \ \S - a = \(S - {a})" for a::"'a set" and S by (auto simp: disjnt_def) have "\y. y \ {path_component_set S x |x. x \ S} - {path_component_set S x} \ disjnt (path_component_set S x) y" using path_component_disjoint path_component_eq by fastforce then have "\{path_component_set S x |x. x \ S} - path_component_set S x = \({path_component_set S y |y. y \ S} - {path_component_set S x})" by (meson *) then show ?thesis by simp qed subsection\Path components\ definition path_component_of where "path_component_of X x y \ \g. pathin X g \ g 0 = x \ g 1 = y" abbreviation path_component_of_set where "path_component_of_set X x \ Collect (path_component_of X x)" definition path_components_of :: "'a topology \ 'a set set" where "path_components_of X \ path_component_of_set X ` topspace X" lemma pathin_canon_iff: "pathin (top_of_set T) g \ path g \ g ` {0..1} \ T" by (simp add: path_def pathin_def) lemma path_component_of_canon_iff [simp]: "path_component_of (top_of_set T) a b \ path_component T a b" by (simp add: path_component_of_def pathin_canon_iff path_defs) lemma path_component_in_topspace: "path_component_of X x y \ x \ topspace X \ y \ topspace X" by (auto simp: path_component_of_def pathin_def continuous_map_def) lemma path_component_of_refl: "path_component_of X x x \ x \ topspace X" by (metis path_component_in_topspace path_component_of_def pathin_const) lemma path_component_of_sym: assumes "path_component_of X x y" shows "path_component_of X y x" using assms apply (clarsimp simp: path_component_of_def pathin_def) apply (rule_tac x="g \ (\t. 1 - t)" in exI) apply (auto intro!: continuous_map_compose simp: continuous_map_in_subtopology continuous_on_op_minus) done lemma path_component_of_sym_iff: "path_component_of X x y \ path_component_of X y x" by (metis path_component_of_sym) lemma continuous_map_cases_le: assumes contp: "continuous_map X euclideanreal p" and contq: "continuous_map X euclideanreal q" and contf: "continuous_map (subtopology X {x. x \ topspace X \ p x \ q x}) Y f" and contg: "continuous_map (subtopology X {x. x \ topspace X \ q x \ p x}) Y g" and fg: "\x. \x \ topspace X; p x = q x\ \ f x = g x" shows "continuous_map X Y (\x. if p x \ q x then f x else g x)" proof - have "continuous_map X Y (\x. if q x - p x \ {0..} then f x else g x)" proof (rule continuous_map_cases_function) show "continuous_map X euclideanreal (\x. q x - p x)" by (intro contp contq continuous_intros) show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of {0..}}) Y f" by (simp add: contf) show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of (topspace euclideanreal - {0..})}) Y g" by (simp add: contg flip: Compl_eq_Diff_UNIV) qed (auto simp: fg) then show ?thesis by simp qed lemma continuous_map_cases_lt: assumes contp: "continuous_map X euclideanreal p" and contq: "continuous_map X euclideanreal q" and contf: "continuous_map (subtopology X {x. x \ topspace X \ p x \ q x}) Y f" and contg: "continuous_map (subtopology X {x. x \ topspace X \ q x \ p x}) Y g" and fg: "\x. \x \ topspace X; p x = q x\ \ f x = g x" shows "continuous_map X Y (\x. if p x < q x then f x else g x)" proof - have "continuous_map X Y (\x. if q x - p x \ {0<..} then f x else g x)" proof (rule continuous_map_cases_function) show "continuous_map X euclideanreal (\x. q x - p x)" by (intro contp contq continuous_intros) show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of {0<..}}) Y f" by (simp add: contf) show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of (topspace euclideanreal - {0<..})}) Y g" by (simp add: contg flip: Compl_eq_Diff_UNIV) qed (auto simp: fg) then show ?thesis by simp qed lemma path_component_of_trans: assumes "path_component_of X x y" and "path_component_of X y z" shows "path_component_of X x z" unfolding path_component_of_def pathin_def proof - let ?T01 = "top_of_set {0..1::real}" obtain g1 g2 where g1: "continuous_map ?T01 X g1" "x = g1 0" "y = g1 1" and g2: "continuous_map ?T01 X g2" "g2 0 = g1 1" "z = g2 1" using assms unfolding path_component_of_def pathin_def by blast let ?g = "\x. if x \ 1/2 then (g1 \ (\t. 2 * t)) x else (g2 \ (\t. 2 * t -1)) x" show "\g. continuous_map ?T01 X g \ g 0 = x \ g 1 = z" proof (intro exI conjI) show "continuous_map (subtopology euclideanreal {0..1}) X ?g" proof (intro continuous_map_cases_le continuous_map_compose, force, force) show "continuous_map (subtopology ?T01 {x \ topspace ?T01. x \ 1/2}) ?T01 ((*) 2)" by (auto simp: continuous_map_in_subtopology continuous_map_from_subtopology) have "continuous_map (subtopology (top_of_set {0..1}) {x. 0 \ x \ x \ 1 \ 1 \ x * 2}) euclideanreal (\t. 2 * t - 1)" by (intro continuous_intros) (force intro: continuous_map_from_subtopology) then show "continuous_map (subtopology ?T01 {x \ topspace ?T01. 1/2 \ x}) ?T01 (\t. 2 * t - 1)" by (force simp: continuous_map_in_subtopology) show "(g1 \ (*) 2) x = (g2 \ (\t. 2 * t - 1)) x" if "x \ topspace ?T01" "x = 1/2" for x using that by (simp add: g2(2) mult.commute continuous_map_from_subtopology) qed (auto simp: g1 g2) qed (auto simp: g1 g2) qed lemma path_component_of_mono: "\path_component_of (subtopology X S) x y; S \ T\ \ path_component_of (subtopology X T) x y" unfolding path_component_of_def by (metis subsetD pathin_subtopology) lemma path_component_of: "path_component_of X x y \ (\T. path_connectedin X T \ x \ T \ y \ T)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (metis atLeastAtMost_iff image_eqI order_refl path_component_of_def path_connectedin_path_image zero_le_one) next assume ?rhs then show ?lhs by (metis path_component_of_def path_connectedin) qed lemma path_component_of_set: "path_component_of X x y \ (\g. pathin X g \ g 0 = x \ g 1 = y)" by (auto simp: path_component_of_def) lemma path_component_of_subset_topspace: "Collect(path_component_of X x) \ topspace X" using path_component_in_topspace by fastforce lemma path_component_of_eq_empty: "Collect(path_component_of X x) = {} \ (x \ topspace X)" using path_component_in_topspace path_component_of_refl by fastforce lemma path_connected_space_iff_path_component: "path_connected_space X \ (\x \ topspace X. \y \ topspace X. path_component_of X x y)" by (simp add: path_component_of path_connected_space_subconnected) lemma path_connected_space_imp_path_component_of: "\path_connected_space X; a \ topspace X; b \ topspace X\ \ path_component_of X a b" by (simp add: path_connected_space_iff_path_component) lemma path_connected_space_path_component_set: "path_connected_space X \ (\x \ topspace X. Collect(path_component_of X x) = topspace X)" using path_component_of_subset_topspace path_connected_space_iff_path_component by fastforce lemma path_component_of_maximal: "\path_connectedin X s; x \ s\ \ s \ Collect(path_component_of X x)" using path_component_of by fastforce lemma path_component_of_equiv: "path_component_of X x y \ x \ topspace X \ y \ topspace X \ path_component_of X x = path_component_of X y" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply (simp add: fun_eq_iff path_component_in_topspace) apply (meson path_component_of_sym path_component_of_trans) done qed (simp add: path_component_of_refl) lemma path_component_of_disjoint: "disjnt (Collect (path_component_of X x)) (Collect (path_component_of X y)) \ ~(path_component_of X x y)" by (force simp: disjnt_def path_component_of_eq_empty path_component_of_equiv) lemma path_component_of_eq: "path_component_of X x = path_component_of X y \ (x \ topspace X) \ (y \ topspace X) \ x \ topspace X \ y \ topspace X \ path_component_of X x y" by (metis Collect_empty_eq_bot path_component_of_eq_empty path_component_of_equiv) lemma path_component_of_aux: "path_component_of X x y \ path_component_of (subtopology X (Collect (path_component_of X x))) x y" by (meson path_component_of path_component_of_maximal path_connectedin_subtopology) lemma path_connectedin_path_component_of: "path_connectedin X (Collect (path_component_of X x))" proof - have "topspace (subtopology X (path_component_of_set X x)) = path_component_of_set X x" by (meson path_component_of_subset_topspace topspace_subtopology_subset) then have "path_connected_space (subtopology X (path_component_of_set X x))" by (metis (full_types) path_component_of_aux mem_Collect_eq path_component_of_equiv path_connected_space_iff_path_component) then show ?thesis by (simp add: path_component_of_subset_topspace path_connectedin_def) qed lemma path_connectedin_euclidean [simp]: "path_connectedin euclidean S \ path_connected S" by (auto simp: path_connectedin_def path_connected_space_iff_path_component path_connected_component) lemma path_connected_space_euclidean_subtopology [simp]: "path_connected_space(subtopology euclidean S) \ path_connected S" using path_connectedin_topspace by force lemma Union_path_components_of: "\(path_components_of X) = topspace X" by (auto simp: path_components_of_def path_component_of_equiv) lemma path_components_of_maximal: "\C \ path_components_of X; path_connectedin X S; ~disjnt C S\ \ S \ C" apply (auto simp: path_components_of_def path_component_of_equiv) using path_component_of_maximal path_connectedin_def apply fastforce by (meson disjnt_subset2 path_component_of_disjoint path_component_of_equiv path_component_of_maximal) lemma pairwise_disjoint_path_components_of: "pairwise disjnt (path_components_of X)" by (auto simp: path_components_of_def pairwise_def path_component_of_disjoint path_component_of_equiv) lemma complement_path_components_of_Union: "C \ path_components_of X \ topspace X - C = \(path_components_of X - {C})" by (metis Diff_cancel Diff_subset Union_path_components_of cSup_singleton diff_Union_pairwise_disjoint insert_subset pairwise_disjoint_path_components_of) lemma nonempty_path_components_of: assumes "C \ path_components_of X" shows "C \ {}" proof - have "C \ path_component_of_set X ` topspace X" using assms path_components_of_def by blast then show ?thesis using path_component_of_refl by fastforce qed lemma path_components_of_subset: "C \ path_components_of X \ C \ topspace X" by (auto simp: path_components_of_def path_component_of_equiv) lemma path_connectedin_path_components_of: "C \ path_components_of X \ path_connectedin X C" by (auto simp: path_components_of_def path_connectedin_path_component_of) lemma path_component_in_path_components_of: "Collect (path_component_of X a) \ path_components_of X \ a \ topspace X" by (metis imageI nonempty_path_components_of path_component_of_eq_empty path_components_of_def) lemma path_connectedin_Union: assumes \: "\S. S \ \ \ path_connectedin X S" "\\ \ {}" shows "path_connectedin X (\\)" proof - obtain a where "\S. S \ \ \ a \ S" using assms by blast then have "\x. x \ topspace (subtopology X (\\)) \ path_component_of (subtopology X (\\)) a x" by simp (meson Union_upper \ path_component_of path_connectedin_subtopology) then show ?thesis using \ unfolding path_connectedin_def by (metis Sup_le_iff path_component_of_equiv path_connected_space_iff_path_component) qed lemma path_connectedin_Un: "\path_connectedin X S; path_connectedin X T; S \ T \ {}\ \ path_connectedin X (S \ T)" by (blast intro: path_connectedin_Union [of "{S,T}", simplified]) lemma path_connected_space_iff_components_eq: "path_connected_space X \ (\C \ path_components_of X. \C' \ path_components_of X. C = C')" unfolding path_components_of_def proof (intro iffI ballI) assume "\C \ path_component_of_set X ` topspace X. \C' \ path_component_of_set X ` topspace X. C = C'" then show "path_connected_space X" using path_component_of_refl path_connected_space_iff_path_component by fastforce qed (auto simp: path_connected_space_path_component_set) lemma path_components_of_eq_empty: "path_components_of X = {} \ topspace X = {}" using Union_path_components_of nonempty_path_components_of by fastforce lemma path_components_of_empty_space: "topspace X = {} \ path_components_of X = {}" by (simp add: path_components_of_eq_empty) lemma path_components_of_subset_singleton: "path_components_of X \ {S} \ path_connected_space X \ (topspace X = {} \ topspace X = S)" proof (cases "topspace X = {}") case True then show ?thesis by (auto simp: path_components_of_empty_space path_connected_space_topspace_empty) next case False have "(path_components_of X = {S}) \ (path_connected_space X \ topspace X = S)" proof (intro iffI conjI) assume L: "path_components_of X = {S}" then show "path_connected_space X" by (simp add: path_connected_space_iff_components_eq) show "topspace X = S" by (metis L ccpo_Sup_singleton [of S] Union_path_components_of) next assume R: "path_connected_space X \ topspace X = S" then show "path_components_of X = {S}" using ccpo_Sup_singleton [of S] by (metis False all_not_in_conv insert_iff mk_disjoint_insert path_component_in_path_components_of path_connected_space_iff_components_eq path_connected_space_path_component_set) qed with False show ?thesis by (simp add: path_components_of_eq_empty subset_singleton_iff) qed lemma path_connected_space_iff_components_subset_singleton: "path_connected_space X \ (\a. path_components_of X \ {a})" by (simp add: path_components_of_subset_singleton) lemma path_components_of_eq_singleton: "path_components_of X = {S} \ path_connected_space X \ topspace X \ {} \ S = topspace X" by (metis cSup_singleton insert_not_empty path_components_of_subset_singleton subset_singleton_iff) lemma path_components_of_path_connected_space: "path_connected_space X \ path_components_of X = (if topspace X = {} then {} else {topspace X})" by (simp add: path_components_of_eq_empty path_components_of_eq_singleton) lemma path_component_subset_connected_component_of: "path_component_of_set X x \ connected_component_of_set X x" proof (cases "x \ topspace X") case True then show ?thesis by (simp add: connected_component_of_maximal path_component_of_refl path_connectedin_imp_connectedin path_connectedin_path_component_of) next case False then show ?thesis using path_component_of_eq_empty by fastforce qed lemma exists_path_component_of_superset: assumes S: "path_connectedin X S" and ne: "topspace X \ {}" obtains C where "C \ path_components_of X" "S \ C" proof (cases "S = {}") case True then show ?thesis using ne path_components_of_eq_empty that by fastforce next case False then obtain a where "a \ S" by blast show ?thesis proof show "Collect (path_component_of X a) \ path_components_of X" by (meson \a \ S\ S subsetD path_component_in_path_components_of path_connectedin_subset_topspace) show "S \ Collect (path_component_of X a)" by (simp add: S \a \ S\ path_component_of_maximal) qed qed lemma path_component_of_eq_overlap: "path_component_of X x = path_component_of X y \ (x \ topspace X) \ (y \ topspace X) \ Collect (path_component_of X x) \ Collect (path_component_of X y) \ {}" by (metis disjnt_def empty_iff inf_bot_right mem_Collect_eq path_component_of_disjoint path_component_of_eq path_component_of_eq_empty) lemma path_component_of_nonoverlap: "Collect (path_component_of X x) \ Collect (path_component_of X y) = {} \ (x \ topspace X) \ (y \ topspace X) \ path_component_of X x \ path_component_of X y" by (metis inf.idem path_component_of_eq_empty path_component_of_eq_overlap) lemma path_component_of_overlap: "Collect (path_component_of X x) \ Collect (path_component_of X y) \ {} \ x \ topspace X \ y \ topspace X \ path_component_of X x = path_component_of X y" by (meson path_component_of_nonoverlap) lemma path_components_of_disjoint: "\C \ path_components_of X; C' \ path_components_of X\ \ disjnt C C' \ C \ C'" by (auto simp: path_components_of_def path_component_of_disjoint path_component_of_equiv) lemma path_components_of_overlap: "\C \ path_components_of X; C' \ path_components_of X\ \ C \ C' \ {} \ C = C'" by (auto simp: path_components_of_def path_component_of_equiv) lemma path_component_of_unique: "\x \ C; path_connectedin X C; \C'. \x \ C'; path_connectedin X C'\ \ C' \ C\ \ Collect (path_component_of X x) = C" by (meson subsetD eq_iff path_component_of_maximal path_connectedin_path_component_of) lemma path_component_of_discrete_topology [simp]: "Collect (path_component_of (discrete_topology U) x) = (if x \ U then {x} else {})" proof - have "\C'. \x \ C'; path_connectedin (discrete_topology U) C'\ \ C' \ {x}" by (metis path_connectedin_discrete_topology subsetD singletonD) then have "x \ U \ Collect (path_component_of (discrete_topology U) x) = {x}" by (simp add: path_component_of_unique) then show ?thesis using path_component_in_topspace by fastforce qed lemma path_component_of_discrete_topology_iff [simp]: "path_component_of (discrete_topology U) x y \ x \ U \ y=x" by (metis empty_iff insertI1 mem_Collect_eq path_component_of_discrete_topology singletonD) lemma path_components_of_discrete_topology [simp]: "path_components_of (discrete_topology U) = (\x. {x}) ` U" by (auto simp: path_components_of_def image_def fun_eq_iff) lemma homeomorphic_map_path_component_of: assumes f: "homeomorphic_map X Y f" and x: "x \ topspace X" shows "Collect (path_component_of Y (f x)) = f ` Collect(path_component_of X x)" proof - obtain g where g: "homeomorphic_maps X Y f g" using f homeomorphic_map_maps by blast show ?thesis proof have "Collect (path_component_of Y (f x)) \ topspace Y" by (simp add: path_component_of_subset_topspace) moreover have "g ` Collect(path_component_of Y (f x)) \ Collect (path_component_of X (g (f x)))" using g x unfolding homeomorphic_maps_def by (metis f homeomorphic_imp_surjective_map imageI mem_Collect_eq path_component_of_maximal path_component_of_refl path_connectedin_continuous_map_image path_connectedin_path_component_of) ultimately show "Collect (path_component_of Y (f x)) \ f ` Collect (path_component_of X x)" using g x unfolding homeomorphic_maps_def continuous_map_def image_iff subset_iff by metis show "f ` Collect (path_component_of X x) \ Collect (path_component_of Y (f x))" proof (rule path_component_of_maximal) show "path_connectedin Y (f ` Collect (path_component_of X x))" by (meson f homeomorphic_map_path_connectedness_eq path_connectedin_path_component_of) qed (simp add: path_component_of_refl x) qed qed lemma homeomorphic_map_path_components_of: assumes "homeomorphic_map X Y f" shows "path_components_of Y = (image f) ` (path_components_of X)" (is "?lhs = ?rhs") unfolding path_components_of_def homeomorphic_imp_surjective_map [OF assms, symmetric] using assms homeomorphic_map_path_component_of by fastforce subsection \Sphere is path-connected\ lemma path_connected_punctured_universe: assumes "2 \ DIM('a::euclidean_space)" shows "path_connected (- {a::'a})" proof - let ?A = "{x::'a. \i\Basis. x \ i < a \ i}" let ?B = "{x::'a. \i\Basis. a \ i < x \ i}" have A: "path_connected ?A" unfolding Collect_bex_eq proof (rule path_connected_UNION) fix i :: 'a assume "i \ Basis" then show "(\i\Basis. (a \ i - 1)*\<^sub>R i) \ {x::'a. x \ i < a \ i}" by simp show "path_connected {x. x \ i < a \ i}" using convex_imp_path_connected [OF convex_halfspace_lt, of i "a \ i"] by (simp add: inner_commute) qed have B: "path_connected ?B" unfolding Collect_bex_eq proof (rule path_connected_UNION) fix i :: 'a assume "i \ Basis" then show "(\i\Basis. (a \ i + 1) *\<^sub>R i) \ {x::'a. a \ i < x \ i}" by simp show "path_connected {x. a \ i < x \ i}" using convex_imp_path_connected [OF convex_halfspace_gt, of "a \ i" i] by (simp add: inner_commute) qed obtain S :: "'a set" where "S \ Basis" and "card S = Suc (Suc 0)" using ex_card[OF assms] by auto then obtain b0 b1 :: 'a where "b0 \ Basis" and "b1 \ Basis" and "b0 \ b1" unfolding card_Suc_eq by auto then have "a + b0 - b1 \ ?A \ ?B" by (auto simp: inner_simps inner_Basis) then have "?A \ ?B \ {}" by fast with A B have "path_connected (?A \ ?B)" by (rule path_connected_Un) also have "?A \ ?B = {x. \i\Basis. x \ i \ a \ i}" unfolding neq_iff bex_disj_distrib Collect_disj_eq .. also have "\ = {x. x \ a}" unfolding euclidean_eq_iff [where 'a='a] by (simp add: Bex_def) also have "\ = - {a}" by auto finally show ?thesis . qed corollary connected_punctured_universe: "2 \ DIM('N::euclidean_space) \ connected(- {a::'N})" by (simp add: path_connected_punctured_universe path_connected_imp_connected) proposition path_connected_sphere: fixes a :: "'a :: euclidean_space" assumes "2 \ DIM('a)" shows "path_connected(sphere a r)" proof (cases r "0::real" rule: linorder_cases) case less then show ?thesis by (simp) next case equal then show ?thesis by (simp) next case greater then have eq: "(sphere (0::'a) r) = (\x. (r / norm x) *\<^sub>R x) ` (- {0::'a})" by (force simp: image_iff split: if_split_asm) have "continuous_on (- {0::'a}) (\x. (r / norm x) *\<^sub>R x)" by (intro continuous_intros) auto then have "path_connected ((\x. (r / norm x) *\<^sub>R x) ` (- {0::'a}))" by (intro path_connected_continuous_image path_connected_punctured_universe assms) with eq have "path_connected (sphere (0::'a) r)" by auto then have "path_connected((+) a ` (sphere (0::'a) r))" by (simp add: path_connected_translation) then show ?thesis by (metis add.right_neutral sphere_translation) qed lemma connected_sphere: fixes a :: "'a :: euclidean_space" assumes "2 \ DIM('a)" shows "connected(sphere a r)" using path_connected_sphere [OF assms] by (simp add: path_connected_imp_connected) corollary path_connected_complement_bounded_convex: fixes S :: "'a :: euclidean_space set" assumes "bounded S" "convex S" and 2: "2 \ DIM('a)" shows "path_connected (- S)" proof (cases "S = {}") case True then show ?thesis using convex_imp_path_connected by auto next case False then obtain a where "a \ S" by auto have \
[rule_format]: "\y\S. \u. 0 \ u \ u \ 1 \ (1 - u) *\<^sub>R a + u *\<^sub>R y \ S" using \convex S\ \a \ S\ by (simp add: convex_alt) { fix x y assume "x \ S" "y \ S" then have "x \ a" "y \ a" using \a \ S\ by auto then have bxy: "bounded(insert x (insert y S))" by (simp add: \bounded S\) then obtain B::real where B: "0 < B" and Bx: "norm (a - x) < B" and By: "norm (a - y) < B" and "S \ ball a B" using bounded_subset_ballD [OF bxy, of a] by (auto simp: dist_norm) define C where "C = B / norm(x - a)" let ?Cxa = "a + C *\<^sub>R (x - a)" { fix u assume u: "(1 - u) *\<^sub>R x + u *\<^sub>R ?Cxa \ S" and "0 \ u" "u \ 1" have CC: "1 \ 1 + (C - 1) * u" using \x \ a\ \0 \ u\ Bx by (auto simp add: C_def norm_minus_commute) have *: "\v. (1 - u) *\<^sub>R x + u *\<^sub>R (a + v *\<^sub>R (x - a)) = a + (1 + (v - 1) * u) *\<^sub>R (x - a)" by (simp add: algebra_simps) have "a + ((1 / (1 + C * u - u)) *\<^sub>R x + ((u / (1 + C * u - u)) *\<^sub>R a + (C * u / (1 + C * u - u)) *\<^sub>R x)) = (1 + (u / (1 + C * u - u))) *\<^sub>R a + ((1 / (1 + C * u - u)) + (C * u / (1 + C * u - u))) *\<^sub>R x" by (simp add: algebra_simps) also have "\ = (1 + (u / (1 + C * u - u))) *\<^sub>R a + (1 + (u / (1 + C * u - u))) *\<^sub>R x" using CC by (simp add: field_simps) also have "\ = x + (1 + (u / (1 + C * u - u))) *\<^sub>R a + (u / (1 + C * u - u)) *\<^sub>R x" by (simp add: algebra_simps) also have "\ = x + ((1 / (1 + C * u - u)) *\<^sub>R a + ((u / (1 + C * u - u)) *\<^sub>R x + (C * u / (1 + C * u - u)) *\<^sub>R a))" using CC by (simp add: field_simps) (simp add: add_divide_distrib scaleR_add_left) finally have xeq: "(1 - 1 / (1 + (C - 1) * u)) *\<^sub>R a + (1 / (1 + (C - 1) * u)) *\<^sub>R (a + (1 + (C - 1) * u) *\<^sub>R (x - a)) = x" by (simp add: algebra_simps) have False using \
[of "a + (1 + (C - 1) * u) *\<^sub>R (x - a)" "1 / (1 + (C - 1) * u)"] using u \x \ a\ \x \ S\ \0 \ u\ CC by (auto simp: xeq *) } then have pcx: "path_component (- S) x ?Cxa" by (force simp: closed_segment_def intro!: path_component_linepath) define D where "D = B / norm(y - a)" \ \massive duplication with the proof above\ let ?Dya = "a + D *\<^sub>R (y - a)" { fix u assume u: "(1 - u) *\<^sub>R y + u *\<^sub>R ?Dya \ S" and "0 \ u" "u \ 1" have DD: "1 \ 1 + (D - 1) * u" using \y \ a\ \0 \ u\ By by (auto simp add: D_def norm_minus_commute) have *: "\v. (1 - u) *\<^sub>R y + u *\<^sub>R (a + v *\<^sub>R (y - a)) = a + (1 + (v - 1) * u) *\<^sub>R (y - a)" by (simp add: algebra_simps) have "a + ((1 / (1 + D * u - u)) *\<^sub>R y + ((u / (1 + D * u - u)) *\<^sub>R a + (D * u / (1 + D * u - u)) *\<^sub>R y)) = (1 + (u / (1 + D * u - u))) *\<^sub>R a + ((1 / (1 + D * u - u)) + (D * u / (1 + D * u - u))) *\<^sub>R y" by (simp add: algebra_simps) also have "\ = (1 + (u / (1 + D * u - u))) *\<^sub>R a + (1 + (u / (1 + D * u - u))) *\<^sub>R y" using DD by (simp add: field_simps) also have "\ = y + (1 + (u / (1 + D * u - u))) *\<^sub>R a + (u / (1 + D * u - u)) *\<^sub>R y" by (simp add: algebra_simps) also have "\ = y + ((1 / (1 + D * u - u)) *\<^sub>R a + ((u / (1 + D * u - u)) *\<^sub>R y + (D * u / (1 + D * u - u)) *\<^sub>R a))" using DD by (simp add: field_simps) (simp add: add_divide_distrib scaleR_add_left) finally have xeq: "(1 - 1 / (1 + (D - 1) * u)) *\<^sub>R a + (1 / (1 + (D - 1) * u)) *\<^sub>R (a + (1 + (D - 1) * u) *\<^sub>R (y - a)) = y" by (simp add: algebra_simps) have False using \
[of "a + (1 + (D - 1) * u) *\<^sub>R (y - a)" "1 / (1 + (D - 1) * u)"] using u \y \ a\ \y \ S\ \0 \ u\ DD by (auto simp: xeq *) } then have pdy: "path_component (- S) y ?Dya" by (force simp: closed_segment_def intro!: path_component_linepath) have pyx: "path_component (- S) ?Dya ?Cxa" proof (rule path_component_of_subset) show "sphere a B \ - S" using \S \ ball a B\ by (force simp: ball_def dist_norm norm_minus_commute) have aB: "?Dya \ sphere a B" "?Cxa \ sphere a B" using \x \ a\ using \y \ a\ B by (auto simp: dist_norm C_def D_def) then show "path_component (sphere a B) ?Dya ?Cxa" using path_connected_sphere [OF 2] path_connected_component by blast qed have "path_component (- S) x y" by (metis path_component_trans path_component_sym pcx pdy pyx) } then show ?thesis by (auto simp: path_connected_component) qed lemma connected_complement_bounded_convex: fixes S :: "'a :: euclidean_space set" assumes "bounded S" "convex S" "2 \ DIM('a)" shows "connected (- S)" using path_connected_complement_bounded_convex [OF assms] path_connected_imp_connected by blast lemma connected_diff_ball: fixes S :: "'a :: euclidean_space set" assumes "connected S" "cball a r \ S" "2 \ DIM('a)" shows "connected (S - ball a r)" proof (rule connected_diff_open_from_closed [OF ball_subset_cball]) show "connected (cball a r - ball a r)" using assms connected_sphere by (auto simp: cball_diff_eq_sphere) qed (auto simp: assms dist_norm) proposition connected_open_delete: assumes "open S" "connected S" and 2: "2 \ DIM('N::euclidean_space)" shows "connected(S - {a::'N})" proof (cases "a \ S") case True with \open S\ obtain \ where "\ > 0" and \: "cball a \ \ S" using open_contains_cball_eq by blast define b where "b \ a + \ *\<^sub>R (SOME i. i \ Basis)" have "dist a b = \" by (simp add: b_def dist_norm SOME_Basis \0 < \\ less_imp_le) with \ have "b \ \{S - ball a r |r. 0 < r \ r < \}" by auto then have nonemp: "(\{S - ball a r |r. 0 < r \ r < \}) = {} \ False" by auto have con: "\r. r < \ \ connected (S - ball a r)" using \ by (force intro: connected_diff_ball [OF \connected S\ _ 2]) have "x \ \{S - ball a r |r. 0 < r \ r < \}" if "x \ S - {a}" for x using that \0 < \\ by (intro UnionI [of "S - ball a (min \ (dist a x) / 2)"]) auto then have "S - {a} = \{S - ball a r | r. 0 < r \ r < \}" by auto then show ?thesis by (auto intro: connected_Union con dest!: nonemp) next case False then show ?thesis by (simp add: \connected S\) qed corollary path_connected_open_delete: assumes "open S" "connected S" and 2: "2 \ DIM('N::euclidean_space)" shows "path_connected(S - {a::'N})" by (simp add: assms connected_open_delete connected_open_path_connected open_delete) corollary path_connected_punctured_ball: "2 \ DIM('N::euclidean_space) \ path_connected(ball a r - {a::'N})" by (simp add: path_connected_open_delete) corollary connected_punctured_ball: "2 \ DIM('N::euclidean_space) \ connected(ball a r - {a::'N})" by (simp add: connected_open_delete) 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 proof (induct T) case empty show ?case using \connected S\ by simp next case (insert x F) then have "connected (S-F)" by auto moreover have "open (S - F)" using finite_imp_closed[OF \finite F\] \open S\ by auto ultimately have "connected (S - F - {x})" using connected_open_delete[OF _ _ 2] by auto thus ?case by (metis Diff_insert) qed lemma sphere_1D_doubleton_zero: assumes 1: "DIM('a) = 1" and "r > 0" obtains x y::"'a::euclidean_space" where "sphere 0 r = {x,y} \ dist x y = 2*r" proof - obtain b::'a where b: "Basis = {b}" using 1 card_1_singletonE by blast show ?thesis proof (intro that conjI) have "x = norm x *\<^sub>R b \ x = - norm x *\<^sub>R b" if "r = norm x" for x proof - have xb: "(x \ b) *\<^sub>R b = x" using euclidean_representation [of x, unfolded b] by force then have "norm ((x \ b) *\<^sub>R b) = norm x" by simp with b have "\x \ b\ = norm x" using norm_Basis by (simp add: b) with xb show ?thesis by (metis (mono_tags, hide_lams) abs_eq_iff abs_norm_cancel) qed with \r > 0\ b show "sphere 0 r = {r *\<^sub>R b, - r *\<^sub>R b}" by (force simp: sphere_def dist_norm) have "dist (r *\<^sub>R b) (- r *\<^sub>R b) = norm (r *\<^sub>R b + r *\<^sub>R b)" by (simp add: dist_norm) also have "\ = norm ((2*r) *\<^sub>R b)" by (metis mult_2 scaleR_add_left) also have "\ = 2*r" using \r > 0\ b norm_Basis by fastforce finally show "dist (r *\<^sub>R b) (- r *\<^sub>R b) = 2*r" . qed qed lemma sphere_1D_doubleton: fixes a :: "'a :: euclidean_space" assumes "DIM('a) = 1" and "r > 0" obtains x y where "sphere a r = {x,y} \ dist x y = 2*r" proof - have "sphere a r = (+) a ` sphere 0 r" by (metis add.right_neutral sphere_translation) then show ?thesis using sphere_1D_doubleton_zero [OF assms] by (metis (mono_tags, lifting) dist_add_cancel image_empty image_insert that) qed lemma psubset_sphere_Compl_connected: fixes S :: "'a::euclidean_space set" assumes S: "S \ sphere a r" and "0 < r" and 2: "2 \ DIM('a)" shows "connected(- S)" proof - have "S \ sphere a r" using S by blast obtain b where "dist a b = r" and "b \ S" using S mem_sphere by blast have CS: "- S = {x. dist a x \ r \ (x \ S)} \ {x. r \ dist a x \ (x \ S)}" by auto have "{x. dist a x \ r \ x \ S} \ {x. r \ dist a x \ x \ S} \ {}" using \b \ S\ \dist a b = r\ by blast moreover have "connected {x. dist a x \ r \ x \ S}" using assms by (force intro: connected_intermediate_closure [of "ball a r"]) moreover have "connected {x. r \ dist a x \ x \ S}" proof (rule connected_intermediate_closure [of "- cball a r"]) show "{x. r \ dist a x \ x \ S} \ closure (- cball a r)" using interior_closure by (force intro: connected_complement_bounded_convex) qed (use assms connected_complement_bounded_convex in auto) ultimately show ?thesis by (simp add: CS connected_Un) qed subsection\Every annulus is a connected set\ lemma path_connected_2DIM_I: fixes a :: "'N::euclidean_space" assumes 2: "2 \ DIM('N)" and pc: "path_connected {r. 0 \ r \ P r}" shows "path_connected {x. P(norm(x - a))}" proof - have "{x. P(norm(x - a))} = (+) a ` {x. P(norm x)}" by force moreover have "path_connected {x::'N. P(norm x)}" proof - let ?D = "{x. 0 \ x \ P x} \ sphere (0::'N) 1" have "x \ (\z. fst z *\<^sub>R snd z) ` ?D" if "P (norm x)" for x::'N proof (cases "x=0") case True with that show ?thesis apply (simp add: image_iff) by (metis (no_types) mem_sphere_0 order_refl vector_choose_size zero_le_one) next case False with that show ?thesis by (rule_tac x="(norm x, x /\<^sub>R norm x)" in image_eqI) auto qed then have *: "{x::'N. P(norm x)} = (\z. fst z *\<^sub>R snd z) ` ?D" by auto have "continuous_on ?D (\z:: real\'N. fst z *\<^sub>R snd z)" by (intro continuous_intros) moreover have "path_connected ?D" by (metis path_connected_Times [OF pc] path_connected_sphere 2) ultimately show ?thesis by (simp add: "*" path_connected_continuous_image) qed ultimately show ?thesis using path_connected_translation by metis qed proposition path_connected_annulus: fixes a :: "'N::euclidean_space" assumes "2 \ DIM('N)" shows "path_connected {x. r1 < norm(x - a) \ norm(x - a) < r2}" "path_connected {x. r1 < norm(x - a) \ norm(x - a) \ r2}" "path_connected {x. r1 \ norm(x - a) \ norm(x - a) < r2}" "path_connected {x. r1 \ norm(x - a) \ norm(x - a) \ r2}" by (auto simp: is_interval_def intro!: is_interval_convex convex_imp_path_connected path_connected_2DIM_I [OF assms]) proposition connected_annulus: fixes a :: "'N::euclidean_space" assumes "2 \ DIM('N::euclidean_space)" shows "connected {x. r1 < norm(x - a) \ norm(x - a) < r2}" "connected {x. r1 < norm(x - a) \ norm(x - a) \ r2}" "connected {x. r1 \ norm(x - a) \ norm(x - a) < r2}" "connected {x. r1 \ norm(x - a) \ norm(x - a) \ r2}" by (auto simp: path_connected_annulus [OF assms] path_connected_imp_connected) subsection\<^marker>\tag unimportant\\Relations between components and path components\ lemma open_connected_component: fixes S :: "'a::real_normed_vector set" assumes "open S" shows "open (connected_component_set S x)" proof (clarsimp simp: open_contains_ball) fix y assume xy: "connected_component S x y" then obtain e where "e>0" "ball y e \ S" using assms connected_component_in openE by blast then show "\e>0. ball y e \ connected_component_set S x" by (metis xy centre_in_ball connected_ball connected_component_eq_eq connected_component_in connected_component_maximal) qed corollary open_components: fixes S :: "'a::real_normed_vector set" shows "\open u; S \ components u\ \ open S" by (simp add: components_iff) (metis open_connected_component) lemma in_closure_connected_component: fixes S :: "'a::real_normed_vector set" assumes x: "x \ S" and S: "open S" shows "x \ closure (connected_component_set S y) \ x \ connected_component_set S y" proof - { assume "x \ closure (connected_component_set S y)" moreover have "x \ connected_component_set S x" using x by simp ultimately have "x \ connected_component_set S y" using S by (meson Compl_disjoint closure_iff_nhds_not_empty connected_component_disjoint disjoint_eq_subset_Compl open_connected_component) } then show ?thesis by (auto simp: closure_def) qed lemma connected_disjoint_Union_open_pick: assumes "pairwise disjnt B" "\S. S \ A \ connected S \ S \ {}" "\S. S \ B \ open S" "\A \ \B" "S \ A" obtains T where "T \ B" "S \ T" "S \ \(B - {T}) = {}" proof - have "S \ \B" "connected S" "S \ {}" using assms \S \ A\ by blast+ then obtain T where "T \ B" "S \ T \ {}" by (metis Sup_inf_eq_bot_iff inf.absorb_iff2 inf_commute) have 1: "open T" by (simp add: \T \ B\ assms) have 2: "open (\(B-{T}))" using assms by blast have 3: "S \ T \ \(B - {T})" using \S \ \B\ by blast have "T \ \(B - {T}) = {}" using \T \ B\ \pairwise disjnt B\ by (auto simp: pairwise_def disjnt_def) then have 4: "T \ \(B - {T}) \ S = {}" by auto from connectedD [OF \connected S\ 1 2 4 3] have "S \ \(B-{T}) = {}" by (auto simp: Int_commute \S \ T \ {}\) with \T \ B\ have "S \ T" using "3" by auto show ?thesis using \S \ \(B - {T}) = {}\ \S \ T\ \T \ B\ that by auto qed lemma connected_disjoint_Union_open_subset: assumes A: "pairwise disjnt A" and B: "pairwise disjnt B" and SA: "\S. S \ A \ open S \ connected S \ S \ {}" and SB: "\S. S \ B \ open S \ connected S \ S \ {}" and eq [simp]: "\A = \B" shows "A \ B" proof fix S assume "S \ A" obtain T where "T \ B" "S \ T" "S \ \(B - {T}) = {}" using SA SB \S \ A\ connected_disjoint_Union_open_pick [OF B, of A] eq order_refl by blast moreover obtain S' where "S' \ A" "T \ S'" "T \ \(A - {S'}) = {}" using SA SB \T \ B\ connected_disjoint_Union_open_pick [OF A, of B] eq order_refl by blast ultimately have "S' = S" by (metis A Int_subset_iff SA \S \ A\ disjnt_def inf.orderE pairwise_def) with \T \ S'\ have "T \ S" by simp with \S \ T\ have "S = T" by blast with \T \ B\ show "S \ B" by simp qed lemma connected_disjoint_Union_open_unique: assumes A: "pairwise disjnt A" and B: "pairwise disjnt B" and SA: "\S. S \ A \ open S \ connected S \ S \ {}" and SB: "\S. S \ B \ open S \ connected S \ S \ {}" and eq [simp]: "\A = \B" shows "A = B" by (rule subset_antisym; metis connected_disjoint_Union_open_subset assms) proposition components_open_unique: fixes S :: "'a::real_normed_vector set" assumes "pairwise disjnt A" "\A = S" "\X. X \ A \ open X \ connected X \ X \ {}" shows "components S = A" proof - have "open S" using assms by blast show ?thesis proof (rule connected_disjoint_Union_open_unique) show "disjoint (components S)" by (simp add: components_eq disjnt_def pairwise_def) qed (use \open S\ in \simp_all add: assms open_components in_components_connected in_components_nonempty\) qed subsection\<^marker>\tag unimportant\\Existence of unbounded components\ lemma cobounded_unbounded_component: fixes S :: "'a :: euclidean_space set" assumes "bounded (-S)" shows "\x. x \ S \ \ bounded (connected_component_set S x)" proof - obtain i::'a where i: "i \ Basis" using nonempty_Basis by blast obtain B where B: "B>0" "-S \ ball 0 B" using bounded_subset_ballD [OF assms, of 0] by auto then have *: "\x. B \ norm x \ x \ S" by (force simp: ball_def dist_norm) have unbounded_inner: "\ bounded {x. inner i x \ B}" proof (clarsimp simp: bounded_def dist_norm) fix e x show "\y. B \ i \ y \ \ norm (x - y) \ e" using i by (rule_tac x="x + (max B e + 1 + \i \ x\) *\<^sub>R i" in exI) (auto simp: inner_right_distrib) qed have \
: "\x. B \ i \ x \ x \ S" using * Basis_le_norm [OF i] by (metis abs_ge_self inner_commute order_trans) have "{x. B \ i \ x} \ connected_component_set S (B *\<^sub>R i)" by (intro connected_component_maximal) (auto simp: i intro: convex_connected convex_halfspace_ge [of B] \
) then have "\ bounded (connected_component_set S (B *\<^sub>R i))" using bounded_subset unbounded_inner by blast moreover have "B *\<^sub>R i \ S" by (rule *) (simp add: norm_Basis [OF i]) ultimately show ?thesis by blast qed lemma cobounded_unique_unbounded_component: fixes S :: "'a :: euclidean_space set" assumes bs: "bounded (-S)" and "2 \ DIM('a)" and bo: "\ bounded(connected_component_set S x)" "\ bounded(connected_component_set S y)" shows "connected_component_set S x = connected_component_set S y" proof - obtain i::'a where i: "i \ Basis" using nonempty_Basis by blast obtain B where B: "B>0" "-S \ ball 0 B" using bounded_subset_ballD [OF bs, of 0] by auto then have *: "\x. B \ norm x \ x \ S" by (force simp: ball_def dist_norm) obtain x' where x': "connected_component S x x'" "norm x' > B" using bo [unfolded bounded_def dist_norm, simplified, rule_format] by (metis diff_zero norm_minus_commute not_less) obtain y' where y': "connected_component S y y'" "norm y' > B" using bo [unfolded bounded_def dist_norm, simplified, rule_format] by (metis diff_zero norm_minus_commute not_less) have x'y': "connected_component S x' y'" unfolding connected_component_def proof (intro exI conjI) show "connected (- ball 0 B :: 'a set)" using assms by (auto intro: connected_complement_bounded_convex) qed (use x' y' dist_norm * in auto) show ?thesis proof (rule connected_component_eq) show "x \ connected_component_set S y" using x' y' x'y' by (metis (no_types) connected_component_eq_eq connected_component_in mem_Collect_eq) qed qed lemma cobounded_unbounded_components: fixes S :: "'a :: euclidean_space set" shows "bounded (-S) \ \c. c \ components S \ \bounded c" by (metis cobounded_unbounded_component components_def imageI) lemma cobounded_unique_unbounded_components: fixes S :: "'a :: euclidean_space set" shows "\bounded (- S); c \ components S; \ bounded c; c' \ components S; \ bounded c'; 2 \ DIM('a)\ \ c' = c" unfolding components_iff by (metis cobounded_unique_unbounded_component) lemma cobounded_has_bounded_component: fixes S :: "'a :: euclidean_space set" assumes "bounded (- S)" "\ connected S" "2 \ DIM('a)" obtains C where "C \ components S" "bounded C" by (meson cobounded_unique_unbounded_components connected_eq_connected_components_eq assms) subsection\The \inside\ and \outside\ of a Set\ text\<^marker>\tag important\\The inside comprises the points in a bounded connected component of the set's complement. The outside comprises the points in unbounded connected component of the complement.\ definition\<^marker>\tag important\ inside where "inside S \ {x. (x \ S) \ bounded(connected_component_set ( - S) x)}" definition\<^marker>\tag important\ outside where "outside S \ -S \ {x. \ bounded(connected_component_set (- S) x)}" lemma outside: "outside S = {x. \ bounded(connected_component_set (- S) x)}" by (auto simp: outside_def) (metis Compl_iff bounded_empty connected_component_eq_empty) lemma inside_no_overlap [simp]: "inside S \ S = {}" by (auto simp: inside_def) lemma outside_no_overlap [simp]: "outside S \ S = {}" by (auto simp: outside_def) lemma inside_Int_outside [simp]: "inside S \ outside S = {}" by (auto simp: inside_def outside_def) lemma inside_Un_outside [simp]: "inside S \ outside S = (- S)" by (auto simp: inside_def outside_def) lemma inside_eq_outside: "inside S = outside S \ S = UNIV" by (auto simp: inside_def outside_def) lemma inside_outside: "inside S = (- (S \ outside S))" by (force simp: inside_def outside) lemma outside_inside: "outside S = (- (S \ inside S))" by (auto simp: inside_outside) (metis IntI equals0D outside_no_overlap) lemma union_with_inside: "S \ inside S = - outside S" by (auto simp: inside_outside) (simp add: outside_inside) lemma union_with_outside: "S \ outside S = - inside S" by (simp add: inside_outside) lemma outside_mono: "S \ T \ outside T \ outside S" by (auto simp: outside bounded_subset connected_component_mono) lemma inside_mono: "S \ T \ inside S - T \ inside T" by (auto simp: inside_def bounded_subset connected_component_mono) lemma segment_bound_lemma: fixes u::real assumes "x \ B" "y \ B" "0 \ u" "u \ 1" shows "(1 - u) * x + u * y \ B" proof - obtain dx dy where "dx \ 0" "dy \ 0" "x = B + dx" "y = B + dy" using assms by auto (metis add.commute diff_add_cancel) with \0 \ u\ \u \ 1\ show ?thesis by (simp add: add_increasing2 mult_left_le field_simps) qed lemma cobounded_outside: fixes S :: "'a :: real_normed_vector set" assumes "bounded S" shows "bounded (- outside S)" proof - obtain B where B: "B>0" "S \ ball 0 B" using bounded_subset_ballD [OF assms, of 0] by auto { fix x::'a and C::real assume Bno: "B \ norm x" and C: "0 < C" have "\y. connected_component (- S) x y \ norm y > C" proof (cases "x = 0") case True with B Bno show ?thesis by force next case False have "closed_segment x (((B + C) / norm x) *\<^sub>R x) \ - ball 0 B" proof fix w assume "w \ closed_segment x (((B + C) / norm x) *\<^sub>R x)" then obtain u where w: "w = (1 - u + u * (B + C) / norm x) *\<^sub>R x" "0 \ u" "u \ 1" by (auto simp add: closed_segment_def real_vector_class.scaleR_add_left [symmetric]) with False B C have "B \ (1 - u) * norm x + u * (B + C)" using segment_bound_lemma [of B "norm x" "B + C" u] Bno by simp with False B C show "w \ - ball 0 B" using distrib_right [of _ _ "norm x"] by (simp add: ball_def w not_less) qed also have "... \ -S" by (simp add: B) finally have "\T. connected T \ T \ - S \ x \ T \ ((B + C) / norm x) *\<^sub>R x \ T" by (rule_tac x="closed_segment x (((B+C)/norm x) *\<^sub>R x)" in exI) simp with False B show ?thesis by (rule_tac x="((B+C)/norm x) *\<^sub>R x" in exI) (simp add: connected_component_def) qed } then show ?thesis apply (simp add: outside_def assms) apply (rule bounded_subset [OF bounded_ball [of 0 B]]) apply (force simp: dist_norm not_less bounded_pos) done qed lemma unbounded_outside: fixes S :: "'a::{real_normed_vector, perfect_space} set" shows "bounded S \ \ bounded(outside S)" using cobounded_imp_unbounded cobounded_outside by blast lemma bounded_inside: fixes S :: "'a::{real_normed_vector, perfect_space} set" shows "bounded S \ bounded(inside S)" by (simp add: bounded_Int cobounded_outside inside_outside) lemma connected_outside: fixes S :: "'a::euclidean_space set" assumes "bounded S" "2 \ DIM('a)" shows "connected(outside S)" apply (clarsimp simp add: connected_iff_connected_component outside) apply (rule_tac S="connected_component_set (- S) x" in connected_component_of_subset) apply (metis (no_types) assms cobounded_unbounded_component cobounded_unique_unbounded_component connected_component_eq_eq connected_component_idemp double_complement mem_Collect_eq) by (simp add: Collect_mono connected_component_eq) lemma outside_connected_component_lt: "outside S = {x. \B. \y. B < norm(y) \ connected_component (- S) x y}" apply (auto simp: outside bounded_def dist_norm) apply (metis diff_0 norm_minus_cancel not_less) by (metis less_diff_eq norm_minus_commute norm_triangle_ineq2 order.trans pinf(6)) lemma outside_connected_component_le: "outside S = {x. \B. \y. B \ norm(y) \ connected_component (- S) x y}" apply (simp add: outside_connected_component_lt Set.set_eq_iff) by (meson gt_ex leD le_less_linear less_imp_le order.trans) lemma not_outside_connected_component_lt: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" and "2 \ DIM('a)" shows "- (outside S) = {x. \B. \y. B < norm(y) \ \ connected_component (- S) x y}" proof - obtain B::real where B: "0 < B" and Bno: "\x. x \ S \ norm x \ B" using S [simplified bounded_pos] by auto { fix y::'a and z::'a assume yz: "B < norm z" "B < norm y" have "connected_component (- cball 0 B) y z" using assms yz by (force simp: dist_norm intro: connected_componentI [OF _ subset_refl] connected_complement_bounded_convex) then have "connected_component (- S) y z" by (metis connected_component_of_subset Bno Compl_anti_mono mem_cball_0 subset_iff) } note cyz = this show ?thesis apply (auto simp: outside bounded_pos) apply (metis Compl_iff bounded_iff cobounded_imp_unbounded mem_Collect_eq not_le) by (metis B connected_component_trans cyz not_le) qed lemma not_outside_connected_component_le: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" "2 \ DIM('a)" shows "- (outside S) = {x. \B. \y. B \ norm(y) \ \ connected_component (- S) x y}" apply (auto intro: less_imp_le simp: not_outside_connected_component_lt [OF assms]) by (meson gt_ex less_le_trans) lemma inside_connected_component_lt: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" "2 \ DIM('a)" shows "inside S = {x. (x \ S) \ (\B. \y. B < norm(y) \ \ connected_component (- S) x y)}" by (auto simp: inside_outside not_outside_connected_component_lt [OF assms]) lemma inside_connected_component_le: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" "2 \ DIM('a)" shows "inside S = {x. (x \ S) \ (\B. \y. B \ norm(y) \ \ connected_component (- S) x y)}" by (auto simp: inside_outside not_outside_connected_component_le [OF assms]) lemma inside_subset: assumes "connected U" and "\ bounded U" and "T \ U = - S" shows "inside S \ T" apply (auto simp: inside_def) by (metis bounded_subset [of "connected_component_set (- S) _"] connected_component_maximal Compl_iff Un_iff assms subsetI) lemma frontier_not_empty: fixes S :: "'a :: real_normed_vector set" shows "\S \ {}; S \ UNIV\ \ frontier S \ {}" using connected_Int_frontier [of UNIV S] by auto lemma frontier_eq_empty: fixes S :: "'a :: real_normed_vector set" shows "frontier S = {} \ S = {} \ S = UNIV" using frontier_UNIV frontier_empty frontier_not_empty by blast lemma frontier_of_connected_component_subset: fixes S :: "'a::real_normed_vector set" shows "frontier(connected_component_set S x) \ frontier S" proof - { fix y assume y1: "y \ closure (connected_component_set S x)" and y2: "y \ interior (connected_component_set S x)" have "y \ closure S" using y1 closure_mono connected_component_subset by blast moreover have "z \ interior (connected_component_set S x)" if "0 < e" "ball y e \ interior S" "dist y z < e" for e z proof - have "ball y e \ connected_component_set S y" using connected_component_maximal that interior_subset by (metis centre_in_ball connected_ball subset_trans) then show ?thesis using y1 apply (simp add: closure_approachable open_contains_ball_eq [OF open_interior]) by (metis connected_component_eq dist_commute mem_Collect_eq mem_ball mem_interior subsetD \0 < e\ y2) qed then have "y \ interior S" using y2 by (force simp: open_contains_ball_eq [OF open_interior]) ultimately have "y \ frontier S" by (auto simp: frontier_def) } then show ?thesis by (auto simp: frontier_def) qed lemma frontier_Union_subset_closure: fixes F :: "'a::real_normed_vector set set" shows "frontier(\F) \ closure(\t \ F. frontier t)" proof - have "\y\F. \y\frontier y. dist y x < e" if "T \ F" "y \ T" "dist y x < e" "x \ interior (\F)" "0 < e" for x y e T proof (cases "x \ T") case True with that show ?thesis by (metis Diff_iff Sup_upper closure_subset contra_subsetD dist_self frontier_def interior_mono) next case False have 1: "closed_segment x y \ T \ {}" using \y \ T\ by blast have 2: "closed_segment x y - T \ {}" using False by blast obtain c where "c \ closed_segment x y" "c \ frontier T" using False connected_Int_frontier [OF connected_segment 1 2] by auto then show ?thesis proof - have "norm (y - x) < e" by (metis dist_norm \dist y x < e\) moreover have "norm (c - x) \ norm (y - x)" by (simp add: \c \ closed_segment x y\ segment_bound(1)) ultimately have "norm (c - x) < e" by linarith then show ?thesis by (metis (no_types) \c \ frontier T\ dist_norm that(1)) qed qed then show ?thesis by (fastforce simp add: frontier_def closure_approachable) qed lemma frontier_Union_subset: fixes F :: "'a::real_normed_vector set set" shows "finite F \ frontier(\F) \ (\t \ F. frontier t)" by (rule order_trans [OF frontier_Union_subset_closure]) (auto simp: closure_subset_eq) lemma frontier_of_components_subset: fixes S :: "'a::real_normed_vector set" shows "C \ components S \ frontier C \ frontier S" by (metis Path_Connected.frontier_of_connected_component_subset components_iff) lemma frontier_of_components_closed_complement: fixes S :: "'a::real_normed_vector set" shows "\closed S; C \ components (- S)\ \ frontier C \ S" using frontier_complement frontier_of_components_subset frontier_subset_eq by blast lemma frontier_minimal_separating_closed: fixes S :: "'a::real_normed_vector set" assumes "closed S" and nconn: "\ connected(- S)" and C: "C \ components (- S)" and conn: "\T. \closed T; T \ S\ \ connected(- T)" shows "frontier C = S" proof (rule ccontr) assume "frontier C \ S" then have "frontier C \ S" using frontier_of_components_closed_complement [OF \closed S\ C] by blast then have "connected(- (frontier C))" by (simp add: conn) have "\ connected(- (frontier C))" unfolding connected_def not_not proof (intro exI conjI) show "open C" using C \closed S\ open_components by blast show "open (- closure C)" by blast show "C \ - closure C \ - frontier C = {}" using closure_subset by blast show "C \ - frontier C \ {}" using C \open C\ components_eq frontier_disjoint_eq by fastforce show "- frontier C \ C \ - closure C" by (simp add: \open C\ closed_Compl frontier_closures) then show "- closure C \ - frontier C \ {}" by (metis (no_types, lifting) C Compl_subset_Compl_iff \frontier C \ S\ compl_sup frontier_closures in_components_subset psubsetE sup.absorb_iff2 sup.boundedE sup_bot.right_neutral sup_inf_absorb) qed then show False using \connected (- frontier C)\ by blast qed lemma connected_component_UNIV [simp]: fixes x :: "'a::real_normed_vector" shows "connected_component_set UNIV x = UNIV" using connected_iff_eq_connected_component_set [of "UNIV::'a set"] connected_UNIV by auto lemma connected_component_eq_UNIV: fixes x :: "'a::real_normed_vector" shows "connected_component_set s x = UNIV \ s = UNIV" using connected_component_in connected_component_UNIV by blast lemma components_UNIV [simp]: "components UNIV = {UNIV :: 'a::real_normed_vector set}" by (auto simp: components_eq_sing_iff) lemma interior_inside_frontier: fixes S :: "'a::real_normed_vector set" assumes "bounded S" shows "interior S \ inside (frontier S)" proof - { fix x y assume x: "x \ interior S" and y: "y \ S" and cc: "connected_component (- frontier S) x y" have "connected_component_set (- frontier S) x \ frontier S \ {}" proof (rule connected_Int_frontier; simp add: set_eq_iff) show "\u. connected_component (- frontier S) x u \ u \ S" by (meson cc connected_component_in connected_component_refl_eq interior_subset subsetD x) show "\u. connected_component (- frontier S) x u \ u \ S" using y cc by blast qed then have "bounded (connected_component_set (- frontier S) x)" using connected_component_in by auto } then show ?thesis apply (auto simp: inside_def frontier_def) apply (rule classical) apply (rule bounded_subset [OF assms], blast) done qed lemma inside_empty [simp]: "inside {} = ({} :: 'a :: {real_normed_vector, perfect_space} set)" by (simp add: inside_def) lemma outside_empty [simp]: "outside {} = (UNIV :: 'a :: {real_normed_vector, perfect_space} set)" using inside_empty inside_Un_outside by blast lemma inside_same_component: "\connected_component (- S) x y; x \ inside S\ \ y \ inside S" using connected_component_eq connected_component_in by (fastforce simp add: inside_def) lemma outside_same_component: "\connected_component (- S) x y; x \ outside S\ \ y \ outside S" using connected_component_eq connected_component_in by (fastforce simp add: outside_def) lemma convex_in_outside: fixes S :: "'a :: {real_normed_vector, perfect_space} set" assumes S: "convex S" and z: "z \ S" shows "z \ outside S" proof (cases "S={}") case True then show ?thesis by simp next case False then obtain a where "a \ S" by blast with z have zna: "z \ a" by auto { assume "bounded (connected_component_set (- S) z)" with bounded_pos_less obtain B where "B>0" and B: "\x. connected_component (- S) z x \ norm x < B" by (metis mem_Collect_eq) define C where "C = (B + 1 + norm z) / norm (z-a)" have "C > 0" using \0 < B\ zna by (simp add: C_def field_split_simps add_strict_increasing) have "\norm (z + C *\<^sub>R (z-a)) - norm (C *\<^sub>R (z-a))\ \ norm z" by (metis add_diff_cancel norm_triangle_ineq3) moreover have "norm (C *\<^sub>R (z-a)) > norm z + B" using zna \B>0\ by (simp add: C_def le_max_iff_disj) ultimately have C: "norm (z + C *\<^sub>R (z-a)) > B" by linarith { fix u::real assume u: "0\u" "u\1" and ins: "(1 - u) *\<^sub>R z + u *\<^sub>R (z + C *\<^sub>R (z - a)) \ S" then have Cpos: "1 + u * C > 0" by (meson \0 < C\ add_pos_nonneg less_eq_real_def zero_le_mult_iff zero_less_one) then have *: "(1 / (1 + u * C)) *\<^sub>R z + (u * C / (1 + u * C)) *\<^sub>R z = z" by (simp add: scaleR_add_left [symmetric] field_split_simps) then have False using convexD_alt [OF S \a \ S\ ins, of "1/(u*C + 1)"] \C>0\ \z \ S\ Cpos u by (simp add: * field_split_simps) } note contra = this have "connected_component (- S) z (z + C *\<^sub>R (z-a))" proof (rule connected_componentI [OF connected_segment]) show "closed_segment z (z + C *\<^sub>R (z - a)) \ - S" using contra by (force simp add: closed_segment_def) qed auto then have False using zna B [of "z + C *\<^sub>R (z-a)"] C by (auto simp: field_split_simps max_mult_distrib_right) } then show ?thesis by (auto simp: outside_def z) qed lemma outside_convex: fixes S :: "'a :: {real_normed_vector, perfect_space} set" assumes "convex S" shows "outside S = - S" by (metis ComplD assms convex_in_outside equalityI inside_Un_outside subsetI sup.cobounded2) lemma outside_singleton [simp]: fixes x :: "'a :: {real_normed_vector, perfect_space}" shows "outside {x} = -{x}" by (auto simp: outside_convex) lemma inside_convex: fixes S :: "'a :: {real_normed_vector, perfect_space} set" shows "convex S \ inside S = {}" by (simp add: inside_outside outside_convex) lemma inside_singleton [simp]: fixes x :: "'a :: {real_normed_vector, perfect_space}" shows "inside {x} = {}" by (auto simp: inside_convex) lemma outside_subset_convex: fixes S :: "'a :: {real_normed_vector, perfect_space} set" shows "\convex T; S \ T\ \ - T \ outside S" using outside_convex outside_mono by blast lemma outside_Un_outside_Un: fixes S :: "'a::real_normed_vector set" assumes "S \ outside(T \ U) = {}" shows "outside(T \ U) \ outside(T \ S)" proof fix x assume x: "x \ outside (T \ U)" have "Y \ - S" if "connected Y" "Y \ - T" "Y \ - U" "x \ Y" "u \ Y" for u Y proof - have "Y \ connected_component_set (- (T \ U)) x" by (simp add: connected_component_maximal that) also have "\ \ outside(T \ U)" by (metis (mono_tags, lifting) Collect_mono mem_Collect_eq outside outside_same_component x) finally have "Y \ outside(T \ U)" . with assms show ?thesis by auto qed with x show "x \ outside (T \ S)" by (simp add: outside_connected_component_lt connected_component_def) meson qed lemma outside_frontier_misses_closure: fixes S :: "'a::real_normed_vector set" assumes "bounded S" shows "outside(frontier S) \ - closure S" unfolding outside_inside Lattices.boolean_algebra_class.compl_le_compl_iff proof - { assume "interior S \ inside (frontier S)" hence "interior S \ inside (frontier S) = inside (frontier S)" by (simp add: subset_Un_eq) then have "closure S \ frontier S \ inside (frontier S)" using frontier_def by auto } then show "closure S \ frontier S \ inside (frontier S)" using interior_inside_frontier [OF assms] by blast qed lemma outside_frontier_eq_complement_closure: fixes S :: "'a :: {real_normed_vector, perfect_space} set" assumes "bounded S" "convex S" shows "outside(frontier S) = - closure S" by (metis Diff_subset assms convex_closure frontier_def outside_frontier_misses_closure outside_subset_convex subset_antisym) lemma inside_frontier_eq_interior: fixes S :: "'a :: {real_normed_vector, perfect_space} set" shows "\bounded S; convex S\ \ inside(frontier S) = interior S" apply (simp add: inside_outside outside_frontier_eq_complement_closure) using closure_subset interior_subset apply (auto simp: frontier_def) done lemma open_inside: fixes S :: "'a::real_normed_vector set" assumes "closed S" shows "open (inside S)" proof - { fix x assume x: "x \ inside S" have "open (connected_component_set (- S) x)" using assms open_connected_component by blast then obtain e where e: "e>0" and e: "\y. dist y x < e \ connected_component (- S) x y" using dist_not_less_zero apply (simp add: open_dist) by (metis (no_types, lifting) Compl_iff connected_component_refl_eq inside_def mem_Collect_eq x) then have "\e>0. ball x e \ inside S" by (metis e dist_commute inside_same_component mem_ball subsetI x) } then show ?thesis by (simp add: open_contains_ball) qed lemma open_outside: fixes S :: "'a::real_normed_vector set" assumes "closed S" shows "open (outside S)" proof - { fix x assume x: "x \ outside S" have "open (connected_component_set (- S) x)" using assms open_connected_component by blast then obtain e where e: "e>0" and e: "\y. dist y x < e \ connected_component (- S) x y" using dist_not_less_zero x by (auto simp add: open_dist outside_def intro: connected_component_refl) then have "\e>0. ball x e \ outside S" by (metis e dist_commute outside_same_component mem_ball subsetI x) } then show ?thesis by (simp add: open_contains_ball) qed lemma closure_inside_subset: fixes S :: "'a::real_normed_vector set" assumes "closed S" shows "closure(inside S) \ S \ inside S" by (metis assms closure_minimal open_closed open_outside sup.cobounded2 union_with_inside) lemma frontier_inside_subset: fixes S :: "'a::real_normed_vector set" assumes "closed S" shows "frontier(inside S) \ S" proof - have "closure (inside S) \ - inside S = closure (inside S) - interior (inside S)" by (metis (no_types) Diff_Compl assms closure_closed interior_closure open_closed open_inside) moreover have "- inside S \ - outside S = S" by (metis (no_types) compl_sup double_compl inside_Un_outside) moreover have "closure (inside S) \ - outside S" by (metis (no_types) assms closure_inside_subset union_with_inside) ultimately have "closure (inside S) - interior (inside S) \ S" by blast then show ?thesis by (simp add: frontier_def open_inside interior_open) qed lemma closure_outside_subset: fixes S :: "'a::real_normed_vector set" assumes "closed S" shows "closure(outside S) \ S \ outside S" by (metis assms closed_open closure_minimal inside_outside open_inside sup_ge2) lemma frontier_outside_subset: fixes S :: "'a::real_normed_vector set" assumes "closed S" shows "frontier(outside S) \ S" unfolding frontier_def by (metis Diff_subset_conv assms closure_outside_subset interior_eq open_outside sup_aci(1)) lemma inside_complement_unbounded_connected_empty: "\connected (- S); \ bounded (- S)\ \ inside S = {}" using inside_subset by blast lemma inside_bounded_complement_connected_empty: fixes S :: "'a::{real_normed_vector, perfect_space} set" shows "\connected (- S); bounded S\ \ inside S = {}" by (metis inside_complement_unbounded_connected_empty cobounded_imp_unbounded) lemma inside_inside: assumes "S \ inside T" shows "inside S - T \ inside T" unfolding inside_def proof clarify fix x assume x: "x \ T" "x \ S" and bo: "bounded (connected_component_set (- S) x)" show "bounded (connected_component_set (- T) x)" proof (cases "S \ connected_component_set (- T) x = {}") case True then show ?thesis by (metis bounded_subset [OF bo] compl_le_compl_iff connected_component_idemp connected_component_mono disjoint_eq_subset_Compl double_compl) next case False then obtain y where y: "y \ S" "y \ connected_component_set (- T) x" by (meson disjoint_iff) then have "bounded (connected_component_set (- T) y)" using assms [unfolded inside_def] by blast with y show ?thesis by (metis connected_component_eq) qed qed lemma inside_inside_subset: "inside(inside S) \ S" using inside_inside union_with_outside by fastforce lemma inside_outside_intersect_connected: "\connected T; inside S \ T \ {}; outside S \ T \ {}\ \ S \ T \ {}" apply (simp add: inside_def outside_def ex_in_conv [symmetric] disjoint_eq_subset_Compl, clarify) by (metis (no_types, hide_lams) Compl_anti_mono connected_component_eq connected_component_maximal contra_subsetD double_compl) lemma outside_bounded_nonempty: fixes S :: "'a :: {real_normed_vector, perfect_space} set" assumes "bounded S" shows "outside S \ {}" by (metis (no_types, lifting) Collect_empty_eq Collect_mem_eq Compl_eq_Diff_UNIV Diff_cancel Diff_disjoint UNIV_I assms ball_eq_empty bounded_diff cobounded_outside convex_ball double_complement order_refl outside_convex outside_def) lemma outside_compact_in_open: fixes S :: "'a :: {real_normed_vector,perfect_space} set" assumes S: "compact S" and T: "open T" and "S \ T" "T \ {}" shows "outside S \ T \ {}" proof - have "outside S \ {}" by (simp add: compact_imp_bounded outside_bounded_nonempty S) with assms obtain a b where a: "a \ outside S" and b: "b \ T" by auto show ?thesis proof (cases "a \ T") case True with a show ?thesis by blast next case False have front: "frontier T \ - S" using \S \ T\ frontier_disjoint_eq T by auto { fix \ assume "path \" and pimg_sbs: "path_image \ - {pathfinish \} \ interior (- T)" and pf: "pathfinish \ \ frontier T" and ps: "pathstart \ = a" define c where "c = pathfinish \" have "c \ -S" unfolding c_def using front pf by blast moreover have "open (-S)" using S compact_imp_closed by blast ultimately obtain \::real where "\ > 0" and \: "cball c \ \ -S" using open_contains_cball[of "-S"] S by blast then obtain d where "d \ T" and d: "dist d c < \" using closure_approachable [of c T] pf unfolding c_def by (metis Diff_iff frontier_def) then have "d \ -S" using \ using dist_commute by (metis contra_subsetD mem_cball not_le not_less_iff_gr_or_eq) have pimg_sbs_cos: "path_image \ \ -S" using \c \ - S\ \S \ T\ c_def interior_subset pimg_sbs by fastforce have "closed_segment c d \ cball c \" by (metis \0 < \\ centre_in_cball closed_segment_subset convex_cball d dist_commute less_eq_real_def mem_cball) with \ have "closed_segment c d \ -S" by blast moreover have con_gcd: "connected (path_image \ \ closed_segment c d)" by (rule connected_Un) (auto simp: c_def \path \\ connected_path_image) ultimately have "connected_component (- S) a d" unfolding connected_component_def using pimg_sbs_cos ps by blast then have "outside S \ T \ {}" using outside_same_component [OF _ a] by (metis IntI \d \ T\ empty_iff) } note * = this have pal: "pathstart (linepath a b) \ closure (- T)" by (auto simp: False closure_def) show ?thesis by (rule exists_path_subpath_to_frontier [OF path_linepath pal _ *]) (auto simp: b) qed qed lemma inside_inside_compact_connected: fixes S :: "'a :: euclidean_space set" assumes S: "closed S" and T: "compact T" and "connected T" "S \ inside T" shows "inside S \ inside T" proof (cases "inside T = {}") case True with assms show ?thesis by auto next case False consider "DIM('a) = 1" | "DIM('a) \ 2" using antisym not_less_eq_eq by fastforce then show ?thesis proof cases case 1 then show ?thesis using connected_convex_1_gen assms False inside_convex by blast next case 2 have "bounded S" using assms by (meson bounded_inside bounded_subset compact_imp_bounded) then have coms: "compact S" by (simp add: S compact_eq_bounded_closed) then have bst: "bounded (S \ T)" by (simp add: compact_imp_bounded T) then obtain r where "0 < r" and r: "S \ T \ ball 0 r" using bounded_subset_ballD by blast have outst: "outside S \ outside T \ {}" proof - have "- ball 0 r \ outside S" by (meson convex_ball le_supE outside_subset_convex r) moreover have "- ball 0 r \ outside T" by (meson convex_ball le_supE outside_subset_convex r) ultimately show ?thesis by (metis Compl_subset_Compl_iff Int_subset_iff bounded_ball inf.orderE outside_bounded_nonempty outside_no_overlap) qed have "S \ T = {}" using assms by (metis disjoint_iff_not_equal inside_no_overlap subsetCE) moreover have "outside S \ inside T \ {}" by (meson False assms(4) compact_eq_bounded_closed coms open_inside outside_compact_in_open T) ultimately have "inside S \ T = {}" using inside_outside_intersect_connected [OF \connected T\, of S] by (metis "2" compact_eq_bounded_closed coms connected_outside inf.commute inside_outside_intersect_connected outst) then show ?thesis using inside_inside [OF \S \ inside T\] by blast qed qed lemma connected_with_inside: fixes S :: "'a :: real_normed_vector set" assumes S: "closed S" and cons: "connected S" shows "connected(S \ inside S)" proof (cases "S \ inside S = UNIV") case True with assms show ?thesis by auto next case False then obtain b where b: "b \ S" "b \ inside S" by blast have *: "\y T. y \ S \ connected T \ a \ T \ y \ T \ T \ (S \ inside S)" if "a \ S \ inside S" for a using that proof assume "a \ S" then show ?thesis by (rule_tac x=a in exI, rule_tac x="{a}" in exI, simp) next assume a: "a \ inside S" then have ain: "a \ closure (inside S)" by (simp add: closure_def) show ?thesis apply (rule exists_path_subpath_to_frontier [OF path_linepath [of a b], of "inside S"]) apply (simp_all add: ain b) subgoal for h apply (rule_tac x="pathfinish h" in exI) apply (simp add: subsetD [OF frontier_inside_subset[OF S]]) apply (rule_tac x="path_image h" in exI) apply (simp add: pathfinish_in_path_image connected_path_image, auto) by (metis Diff_single_insert S frontier_inside_subset insert_iff interior_subset subsetD) done qed show ?thesis apply (simp add: connected_iff_connected_component) apply (clarsimp simp add: connected_component_def dest!: *) subgoal for x y u u' T t' by (rule_tac x="(S \ T \ t')" in exI) (auto intro!: connected_Un cons) done qed text\The proof is virtually the same as that above.\ lemma connected_with_outside: fixes S :: "'a :: real_normed_vector set" assumes S: "closed S" and cons: "connected S" shows "connected(S \ outside S)" proof (cases "S \ outside S = UNIV") case True with assms show ?thesis by auto next case False then obtain b where b: "b \ S" "b \ outside S" by blast have *: "\y T. y \ S \ connected T \ a \ T \ y \ T \ T \ (S \ outside S)" if "a \ (S \ outside S)" for a using that proof assume "a \ S" then show ?thesis by (rule_tac x=a in exI, rule_tac x="{a}" in exI, simp) next assume a: "a \ outside S" then have ain: "a \ closure (outside S)" by (simp add: closure_def) show ?thesis apply (rule exists_path_subpath_to_frontier [OF path_linepath [of a b], of "outside S"]) apply (simp_all add: ain b) subgoal for h apply (rule_tac x="pathfinish h" in exI) apply (simp add: subsetD [OF frontier_outside_subset[OF S]]) apply (rule_tac x="path_image h" in exI) apply (simp add: pathfinish_in_path_image connected_path_image, auto) by (metis (no_types, lifting) frontier_outside_subset insertE insert_Diff interior_eq open_outside pathfinish_in_path_image S subsetCE) done qed show ?thesis apply (simp add: connected_iff_connected_component) apply (clarsimp simp add: connected_component_def dest!: *) subgoal for x y u u' T t' by (rule_tac x="(S \ T \ t')" in exI) (auto intro!: connected_Un cons) done qed lemma inside_inside_eq_empty [simp]: fixes S :: "'a :: {real_normed_vector, perfect_space} set" assumes S: "closed S" and cons: "connected S" shows "inside (inside S) = {}" by (metis (no_types) unbounded_outside connected_with_outside [OF assms] bounded_Un inside_complement_unbounded_connected_empty unbounded_outside union_with_outside) lemma inside_in_components: "inside S \ components (- S) \ connected(inside S) \ inside S \ {}" (is "?lhs = ?rhs") proof assume R: ?rhs then have "\x. \x \ S; x \ inside S\ \ \ connected (inside S)" by (simp add: inside_outside) with R show ?lhs unfolding in_components_maximal by (auto intro: inside_same_component connected_componentI) qed (simp add: in_components_maximal) text\The proof is like that above.\ lemma outside_in_components: "outside S \ components (- S) \ connected(outside S) \ outside S \ {}" (is "?lhs = ?rhs") proof assume R: ?rhs then have "\x. \x \ S; x \ outside S\ \ \ connected (outside S)" by (meson disjoint_iff outside_no_overlap) with R show ?lhs unfolding in_components_maximal by (auto intro: outside_same_component connected_componentI) qed (simp add: in_components_maximal) lemma bounded_unique_outside: fixes S :: "'a :: euclidean_space set" assumes "bounded S" "DIM('a) \ 2" shows "(c \ components (- S) \ \ bounded c \ c = outside S)" using assms by (metis cobounded_unique_unbounded_components connected_outside double_compl outside_bounded_nonempty outside_in_components unbounded_outside) subsection\Condition for an open map's image to contain a ball\ proposition ball_subset_open_map_image: fixes f :: "'a::heine_borel \ 'b :: {real_normed_vector,heine_borel}" assumes contf: "continuous_on (closure S) f" and oint: "open (f ` interior S)" and le_no: "\z. z \ frontier S \ r \ norm(f z - f a)" and "bounded S" "a \ S" "0 < r" shows "ball (f a) r \ f ` S" proof (cases "f ` S = UNIV") case True then show ?thesis by simp next case False then have "closed (frontier (f ` S))" "frontier (f ` S) \ {}" using \a \ S\ by (auto simp: frontier_eq_empty) then obtain w where w: "w \ frontier (f ` S)" and dw_le: "\y. y \ frontier (f ` S) \ norm (f a - w) \ norm (f a - y)" by (auto simp add: dist_norm intro: distance_attains_inf [of "frontier(f ` S)" "f a"]) then obtain \ where \: "\n. \ n \ f ` S" and tendsw: "\ \ w" by (metis Diff_iff frontier_def closure_sequential) then have "\n. \x \ S. \ n = f x" by force then obtain z where zs: "\n. z n \ S" and fz: "\n. \ n = f (z n)" by metis then obtain y K where y: "y \ closure S" and "strict_mono (K :: nat \ nat)" and Klim: "(z \ K) \ y" using \bounded S\ unfolding compact_closure [symmetric] compact_def by (meson closure_subset subset_iff) then have ftendsw: "((\n. f (z n)) \ K) \ w" by (metis LIMSEQ_subseq_LIMSEQ fun.map_cong0 fz tendsw) have zKs: "\n. (z \ K) n \ S" by (simp add: zs) have fz: "f \ z = \" "(\n. f (z n)) = \" using fz by auto then have "(\ \ K) \ f y" by (metis (no_types) Klim zKs y contf comp_assoc continuous_on_closure_sequentially) with fz have wy: "w = f y" using fz LIMSEQ_unique ftendsw by auto have rle: "r \ norm (f y - f a)" proof (rule le_no) show "y \ frontier S" using w wy oint by (force simp: imageI image_mono interiorI interior_subset frontier_def y) qed have **: "(b \ (- S) \ {} \ b - (- S) \ {} \ b \ f \ {}) \ (b \ S \ {}) \ b \ f = {} \ b \ S" for b f and S :: "'b set" by blast have \
: "\y. \norm (f a - y) < r; y \ frontier (f ` S)\ \ False" by (metis dw_le norm_minus_commute not_less order_trans rle wy) show ?thesis apply (rule ** [OF connected_Int_frontier [where t = "f`S", OF connected_ball]]) (*such a horrible mess*) using \a \ S\ \0 < r\ by (auto simp: disjoint_iff_not_equal dist_norm dest: \
) qed subsubsection\Special characterizations of classes of functions into and out of R.\ lemma Hausdorff_space_euclidean [simp]: "Hausdorff_space (euclidean :: 'a::metric_space topology)" proof - have "\U V. open U \ open V \ x \ U \ y \ V \ disjnt U V" if "x \ y" for x y :: 'a proof (intro exI conjI) let ?r = "dist x y / 2" have [simp]: "?r > 0" by (simp add: that) show "open (ball x ?r)" "open (ball y ?r)" "x \ (ball x ?r)" "y \ (ball y ?r)" by (auto simp add: that) show "disjnt (ball x ?r) (ball y ?r)" unfolding disjnt_def by (simp add: disjoint_ballI) qed then show ?thesis by (simp add: Hausdorff_space_def) qed proposition embedding_map_into_euclideanreal: assumes "path_connected_space X" shows "embedding_map X euclideanreal f \ continuous_map X euclideanreal f \ inj_on f (topspace X)" proof safe show "continuous_map X euclideanreal f" if "embedding_map X euclideanreal f" using continuous_map_in_subtopology homeomorphic_imp_continuous_map that unfolding embedding_map_def by blast show "inj_on f (topspace X)" if "embedding_map X euclideanreal f" using that homeomorphic_imp_injective_map unfolding embedding_map_def by blast show "embedding_map X euclideanreal f" if cont: "continuous_map X euclideanreal f" and inj: "inj_on f (topspace X)" proof - obtain g where gf: "\x. x \ topspace X \ g (f x) = x" using inv_into_f_f [OF inj] by auto show ?thesis unfolding embedding_map_def homeomorphic_map_maps homeomorphic_maps_def proof (intro exI conjI) show "continuous_map X (top_of_set (f ` topspace X)) f" by (simp add: cont continuous_map_in_subtopology) let ?S = "f ` topspace X" have eq: "{x \ ?S. g x \ U} = f ` U" if "openin X U" for U using openin_subset [OF that] by (auto simp: gf) have 1: "g ` ?S \ topspace X" using eq by blast have "openin (top_of_set ?S) {x \ ?S. g x \ T}" if "openin X T" for T proof - have "T \ topspace X" by (simp add: openin_subset that) have RR: "\x \ ?S \ g -` T. \d>0. \x' \ ?S \ ball x d. g x' \ T" proof (clarsimp simp add: gf) have pcS: "path_connectedin euclidean ?S" using assms cont path_connectedin_continuous_map_image path_connectedin_topspace by blast show "\d>0. \x'\f ` topspace X \ ball (f x) d. g x' \ T" if "x \ T" for x proof - have x: "x \ topspace X" using \T \ topspace X\ \x \ T\ by blast obtain u v d where "0 < d" "u \ topspace X" "v \ topspace X" and sub_fuv: "?S \ {f x - d .. f x + d} \ {f u..f v}" proof (cases "\u \ topspace X. f u < f x") case True then obtain u where u: "u \ topspace X" "f u < f x" .. show ?thesis proof (cases "\v \ topspace X. f x < f v") case True then obtain v where v: "v \ topspace X" "f x < f v" .. show ?thesis proof let ?d = "min (f x - f u) (f v - f x)" show "0 < ?d" by (simp add: \f u < f x\ \f x < f v\) show "f ` topspace X \ {f x - ?d..f x + ?d} \ {f u..f v}" by fastforce qed (auto simp: u v) next case False show ?thesis proof let ?d = "f x - f u" show "0 < ?d" by (simp add: u) show "f ` topspace X \ {f x - ?d..f x + ?d} \ {f u..f x}" using x u False by auto qed (auto simp: x u) qed next case False note no_u = False show ?thesis proof (cases "\v \ topspace X. f x < f v") case True then obtain v where v: "v \ topspace X" "f x < f v" .. show ?thesis proof let ?d = "f v - f x" show "0 < ?d" by (simp add: v) show "f ` topspace X \ {f x - ?d..f x + ?d} \ {f x..f v}" using False by auto qed (auto simp: x v) next case False show ?thesis proof show "f ` topspace X \ {f x - 1..f x + 1} \ {f x..f x}" using False no_u by fastforce qed (auto simp: x) qed qed then obtain h where "pathin X h" "h 0 = u" "h 1 = v" using assms unfolding path_connected_space_def by blast obtain C where "compactin X C" "connectedin X C" "u \ C" "v \ C" proof show "compactin X (h ` {0..1})" using that by (simp add: \pathin X h\ compactin_path_image) show "connectedin X (h ` {0..1})" using \pathin X h\ connectedin_path_image by blast qed (use \h 0 = u\ \h 1 = v\ in auto) have "continuous_map (subtopology euclideanreal (?S \ {f x - d .. f x + d})) (subtopology X C) g" proof (rule continuous_inverse_map) show "compact_space (subtopology X C)" using \compactin X C\ compactin_subspace by blast show "continuous_map (subtopology X C) euclideanreal f" by (simp add: cont continuous_map_from_subtopology) have "{f u .. f v} \ f ` topspace (subtopology X C)" proof (rule connected_contains_Icc) show "connected (f ` topspace (subtopology X C))" using connectedin_continuous_map_image [OF cont] by (simp add: \compactin X C\ \connectedin X C\ compactin_subset_topspace inf_absorb2) show "f u \ f ` topspace (subtopology X C)" by (simp add: \u \ C\ \u \ topspace X\) show "f v \ f ` topspace (subtopology X C)" by (simp add: \v \ C\ \v \ topspace X\) qed then show "f ` topspace X \ {f x - d..f x + d} \ f ` topspace (subtopology X C)" using sub_fuv by blast qed (auto simp: gf) then have contg: "continuous_map (subtopology euclideanreal (?S \ {f x - d .. f x + d})) X g" using continuous_map_in_subtopology by blast have "\e>0. \x \ ?S \ {f x - d .. f x + d} \ ball (f x) e. g x \ T" using openin_continuous_map_preimage [OF contg \openin X T\] x \x \ T\ \0 < d\ unfolding openin_euclidean_subtopology_iff by (force simp: gf dist_commute) then obtain e where "e > 0 \ (\x\f ` topspace X \ {f x - d..f x + d} \ ball (f x) e. g x \ T)" by metis with \0 < d\ have "min d e > 0" "\u. u \ topspace X \ \f x - f u\ < min d e \ u \ T" using dist_real_def gf by force+ then show ?thesis by (metis (full_types) Int_iff dist_real_def image_iff mem_ball gf) qed qed then obtain d where d: "\r. r \ ?S \ g -` T \ d r > 0 \ (\x \ ?S \ ball r (d r). g x \ T)" by metis show ?thesis unfolding openin_subtopology proof (intro exI conjI) show "{x \ ?S. g x \ T} = (\r \ ?S \ g -` T. ball r (d r)) \ f ` topspace X" using d by (auto simp: gf) qed auto qed then show "continuous_map (top_of_set ?S) X g" by (simp add: continuous_map_def gf) qed (auto simp: gf) qed qed subsubsection \An injective function into R is a homeomorphism and so an open map.\ lemma injective_into_1d_eq_homeomorphism: fixes f :: "'a::topological_space \ real" assumes f: "continuous_on S f" and S: "path_connected S" shows "inj_on f S \ (\g. homeomorphism S (f ` S) f g)" proof show "\g. homeomorphism S (f ` S) f g" if "inj_on f S" proof - have "embedding_map (top_of_set S) euclideanreal f" using that embedding_map_into_euclideanreal [of "top_of_set S" f] assms by auto then show ?thesis by (simp add: embedding_map_def) (metis all_closedin_homeomorphic_image f homeomorphism_injective_closed_map that) qed qed (metis homeomorphism_def inj_onI) lemma injective_into_1d_imp_open_map: fixes f :: "'a::topological_space \ real" assumes "continuous_on S f" "path_connected S" "inj_on f S" "openin (subtopology euclidean S) T" shows "openin (subtopology euclidean (f ` S)) (f ` T)" using assms homeomorphism_imp_open_map injective_into_1d_eq_homeomorphism by blast lemma homeomorphism_into_1d: fixes f :: "'a::topological_space \ real" assumes "path_connected S" "continuous_on S f" "f ` S = T" "inj_on f S" shows "\g. homeomorphism S T f g" using assms injective_into_1d_eq_homeomorphism by blast subsection\<^marker>\tag unimportant\ \Rectangular paths\ definition\<^marker>\tag unimportant\ rectpath where "rectpath a1 a3 = (let a2 = Complex (Re a3) (Im a1); a4 = Complex (Re a1) (Im a3) in linepath a1 a2 +++ linepath a2 a3 +++ linepath a3 a4 +++ linepath a4 a1)" lemma path_rectpath [simp, intro]: "path (rectpath a b)" by (simp add: Let_def rectpath_def) lemma pathstart_rectpath [simp]: "pathstart (rectpath a1 a3) = a1" by (simp add: rectpath_def Let_def) lemma pathfinish_rectpath [simp]: "pathfinish (rectpath a1 a3) = a1" by (simp add: rectpath_def Let_def) lemma simple_path_rectpath [simp, intro]: assumes "Re a1 \ Re a3" "Im a1 \ Im a3" shows "simple_path (rectpath a1 a3)" unfolding rectpath_def Let_def using assms by (intro simple_path_join_loop arc_join arc_linepath) (auto simp: complex_eq_iff path_image_join closed_segment_same_Re closed_segment_same_Im) lemma path_image_rectpath: assumes "Re a1 \ Re a3" "Im a1 \ Im a3" shows "path_image (rectpath a1 a3) = {z. Re z \ {Re a1, Re a3} \ Im z \ {Im a1..Im a3}} \ {z. Im z \ {Im a1, Im a3} \ Re z \ {Re a1..Re a3}}" (is "?lhs = ?rhs") proof - define a2 a4 where "a2 = Complex (Re a3) (Im a1)" and "a4 = Complex (Re a1) (Im a3)" have "?lhs = closed_segment a1 a2 \ closed_segment a2 a3 \ closed_segment a4 a3 \ closed_segment a1 a4" by (simp_all add: rectpath_def Let_def path_image_join closed_segment_commute a2_def a4_def Un_assoc) also have "\ = ?rhs" using assms by (auto simp: rectpath_def Let_def path_image_join a2_def a4_def closed_segment_same_Re closed_segment_same_Im closed_segment_eq_real_ivl) finally show ?thesis . qed lemma path_image_rectpath_subset_cbox: assumes "Re a \ Re b" "Im a \ Im b" shows "path_image (rectpath a b) \ cbox a b" using assms by (auto simp: path_image_rectpath in_cbox_complex_iff) lemma path_image_rectpath_inter_box: assumes "Re a \ Re b" "Im a \ Im b" shows "path_image (rectpath a b) \ box a b = {}" using assms by (auto simp: path_image_rectpath in_box_complex_iff) lemma path_image_rectpath_cbox_minus_box: assumes "Re a \ Re b" "Im a \ Im b" shows "path_image (rectpath a b) = cbox a b - box a b" using assms by (auto simp: path_image_rectpath in_cbox_complex_iff in_box_complex_iff) end diff --git a/src/HOL/Complex_Analysis/Complex_Singularities.thy b/src/HOL/Complex_Analysis/Complex_Singularities.thy --- a/src/HOL/Complex_Analysis/Complex_Singularities.thy +++ b/src/HOL/Complex_Analysis/Complex_Singularities.thy @@ -1,1533 +1,1538 @@ theory Complex_Singularities imports Conformal_Mappings begin subsection \Non-essential singular points\ definition\<^marker>\tag important\ is_pole :: "('a::topological_space \ 'b::real_normed_vector) \ 'a \ bool" where "is_pole f a = (LIM x (at a). f x :> at_infinity)" lemma is_pole_cong: assumes "eventually (\x. f x = g x) (at a)" "a=b" shows "is_pole f a \ is_pole g b" unfolding is_pole_def using assms by (intro filterlim_cong,auto) lemma is_pole_transform: assumes "is_pole f a" "eventually (\x. f x = g x) (at a)" "a=b" shows "is_pole g b" using is_pole_cong assms by auto +lemma is_pole_shift_iff: + fixes f :: "('a::real_normed_vector \ 'b::real_normed_vector)" + shows "is_pole (f \ (+) d) a = is_pole f (a + d)" + by (metis add_diff_cancel_right' filterlim_shift_iff is_pole_def) + 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 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 non_z:"\x\s-{z}. f x\0" shows "(\x. if x=z then 0 else inverse (f x)) holomorphic_on s" proof - define g where "g \ \x. if x=z then 0 else inverse (f x)" have "isCont g z" unfolding isCont_def using is_pole_tendsto[OF pole] by (simp add: g_def cong: LIM_cong) 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 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\ by (auto simp add:continuous_on_eq_continuous_at) moreover have "(inverse o f) holomorphic_on (s-{z})" unfolding comp_def using f_holo by (auto elim!:holomorphic_on_inverse simp add:non_z) hence "g holomorphic_on (s-{z})" 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) qed lemma not_is_pole_holomorphic: assumes "open A" "x \ A" "f holomorphic_on A" shows "\is_pole f x" proof - have "continuous_on A f" by (intro holomorphic_on_imp_continuous_on) fact with assms have "isCont f x" by (simp add: continuous_on_eq_continuous_at) hence "f \x\ f x" by (simp add: isCont_def) thus "\is_pole f x" unfolding is_pole_def using not_tendsto_and_filterlim_at_infinity[of "at x" f "f x"] by auto qed lemma is_pole_inverse_power: "n > 0 \ is_pole (\z::complex. 1 / (z - a) ^ n) a" unfolding is_pole_def inverse_eq_divide [symmetric] by (intro filterlim_compose[OF filterlim_inverse_at_infinity] tendsto_intros) (auto simp: filterlim_at eventually_at intro!: exI[of _ 1] tendsto_eq_intros) lemma is_pole_inverse: "is_pole (\z::complex. 1 / (z - a)) a" using is_pole_inverse_power[of 1 a] by simp lemma is_pole_divide: fixes f :: "'a :: t2_space \ 'b :: real_normed_field" assumes "isCont f z" "filterlim g (at 0) (at z)" "f z \ 0" shows "is_pole (\z. f z / g z) z" proof - have "filterlim (\z. f z * inverse (g z)) at_infinity (at z)" by (intro tendsto_mult_filterlim_at_infinity[of _ "f z"] filterlim_compose[OF filterlim_inverse_at_infinity])+ (insert assms, auto simp: isCont_def) thus ?thesis by (simp add: field_split_simps is_pole_def) qed lemma is_pole_basic: assumes "f holomorphic_on A" "open A" "z \ A" "f z \ 0" "n > 0" shows "is_pole (\w. f w / (w - z) ^ n) z" proof (rule is_pole_divide) have "continuous_on A f" by (rule holomorphic_on_imp_continuous_on) fact with assms show "isCont f z" by (auto simp: continuous_on_eq_continuous_at) have "filterlim (\w. (w - z) ^ n) (nhds 0) (at z)" using assms by (auto intro!: tendsto_eq_intros) thus "filterlim (\w. (w - z) ^ n) (at 0) (at z)" by (intro filterlim_atI tendsto_eq_intros) (insert assms, auto simp: eventually_at_filter) qed fact+ lemma is_pole_basic': assumes "f holomorphic_on A" "open A" "0 \ A" "f 0 \ 0" "n > 0" shows "is_pole (\w. f w / w ^ n) 0" using is_pole_basic[of f A 0] assms by simp text \The proposition \<^term>\\x. ((f::complex\complex) \ x) (at z) \ is_pole f z\ can be interpreted as the complex function \<^term>\f\ has a non-essential singularity at \<^term>\z\ (i.e. the singularity is either removable or a pole).\ definition not_essential::"[complex \ complex, complex] \ bool" where "not_essential f z = (\x. f\z\x \ is_pole f z)" definition isolated_singularity_at::"[complex \ complex, complex] \ bool" where "isolated_singularity_at f z = (\r>0. f analytic_on ball z r-{z})" named_theorems singularity_intros "introduction rules for singularities" lemma holomorphic_factor_unique: fixes f::"complex \ complex" and z::complex and r::real and m n::int assumes "r>0" "g z\0" "h z\0" and asm:"\w\ball z r-{z}. f w = g w * (w-z) powr n \ g w\0 \ f w = h w * (w - z) powr m \ h w\0" and g_holo:"g holomorphic_on ball z r" and h_holo:"h holomorphic_on ball z r" shows "n=m" proof - have [simp]:"at z within ball z r \ bot" using \r>0\ by (auto simp add:at_within_ball_bot_iff) have False when "n>m" proof - have "(h \ 0) (at z within ball z r)" proof (rule Lim_transform_within[OF _ \r>0\, where f="\w. (w - z) powr (n - m) * g w"]) have "\w\ball z r-{z}. h w = (w-z)powr(n-m) * g w" using \n>m\ asm \r>0\ apply (auto simp add:field_simps powr_diff) by force then show "\x' \ ball z r; 0 < dist x' z;dist x' z < r\ \ (x' - z) powr (n - m) * g x' = h x'" for x' by auto next define F where "F \ at z within ball z r" define f' where "f' \ \x. (x - z) powr (n-m)" have "f' z=0" using \n>m\ unfolding f'_def by auto moreover have "continuous F f'" unfolding f'_def F_def continuous_def apply (subst Lim_ident_at) using \n>m\ by (auto intro!:tendsto_powr_complex_0 tendsto_eq_intros) ultimately have "(f' \ 0) F" unfolding F_def 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\) ultimately have "h z=0" by (auto intro!: tendsto_unique) thus False using \h z\0\ by auto qed moreover have False when "m>n" proof - have "(g \ 0) (at z within ball z r)" proof (rule Lim_transform_within[OF _ \r>0\, where f="\w. (w - z) powr (m - n) * h w"]) have "\w\ball z r -{z}. g w = (w-z) powr (m-n) * h w" using \m>n\ asm apply (auto simp add:field_simps powr_diff) by force then show "\x' \ ball z r; 0 < dist x' z;dist x' z < r\ \ (x' - z) powr (m - n) * h x' = g x'" for x' by auto next define F where "F \ at z within ball z r" define f' where "f' \\x. (x - z) powr (m-n)" have "f' z=0" using \m>n\ unfolding f'_def by auto moreover have "continuous F f'" unfolding f'_def F_def continuous_def apply (subst Lim_ident_at) using \m>n\ by (auto intro!:tendsto_powr_complex_0 tendsto_eq_intros) ultimately have "(f' \ 0) F" unfolding F_def 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\) ultimately have "g z=0" by (auto intro!: tendsto_unique) thus False using \g z\0\ by auto qed ultimately show "n=m" by fastforce qed lemma holomorphic_factor_puncture: assumes f_iso:"isolated_singularity_at f z" and "not_essential f z" \ \\<^term>\f\ has either a removable singularity or a pole at \<^term>\z\\ and non_zero:"\\<^sub>Fw in (at z). f w\0" \ \\<^term>\f\ will not be constantly zero in a neighbour of \<^term>\z\\ shows "\!n::int. \g r. 0 < r \ g holomorphic_on cball z r \ g z\0 \ (\w\cball z r-{z}. f w = g w * (w-z) powr n \ g w\0)" proof - define P where "P = (\f n g r. 0 < r \ g holomorphic_on cball z r \ g z\0 \ (\w\cball z r - {z}. f w = g w * (w-z) powr (of_int n) \ g w\0))" have imp_unique:"\!n::int. \g r. P f n g r" when "\n g r. P f n g r" proof (rule ex_ex1I[OF that]) fix n1 n2 :: int assume g1_asm:"\g1 r1. P f n1 g1 r1" and g2_asm:"\g2 r2. P f n2 g2 r2" define fac where "fac \ \n g r. \w\cball z r-{z}. f w = g w * (w - z) powr (of_int n) \ g w \ 0" obtain g1 r1 where "0 < r1" and g1_holo: "g1 holomorphic_on cball z r1" and "g1 z\0" and "fac n1 g1 r1" using g1_asm unfolding P_def fac_def by auto obtain g2 r2 where "0 < r2" and g2_holo: "g2 holomorphic_on cball z r2" and "g2 z\0" and "fac n2 g2 r2" using g2_asm unfolding P_def fac_def 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-{z}. f w = g1 w * (w-z) powr n1 \ g1 w\0 \ f w = g2 w * (w - z) powr n2 \ g2 w\0" using \fac n1 g1 r1\ \fac n2 g2 r2\ unfolding fac_def r_def by fastforce ultimately show "n1=n2" using g1_holo g2_holo \g1 z\0\ \g2 z\0\ apply (elim holomorphic_factor_unique) by (auto simp add:r_def) qed have P_exist:"\ n g r. P h n g r" when "\z'. (h \ z') (at z)" "isolated_singularity_at h z" "\\<^sub>Fw in (at z). h w\0" for h proof - from that(2) obtain r where "r>0" "h analytic_on ball z r - {z}" unfolding isolated_singularity_at_def by auto obtain z' where "(h \ z') (at z)" using \\z'. (h \ z') (at z)\ by auto define h' where "h'=(\x. if x=z then z' else h x)" have "h' holomorphic_on ball z r" apply (rule no_isolated_singularity'[of "{z}"]) subgoal by (metis LIM_equal Lim_at_imp_Lim_at_within \h \z\ z'\ empty_iff h'_def insert_iff) subgoal using \h analytic_on ball z r - {z}\ analytic_imp_holomorphic h'_def holomorphic_transform by fastforce by auto have ?thesis when "z'=0" proof - have "h' z=0" using that unfolding h'_def by auto moreover have "\ h' constant_on ball z r" using \\\<^sub>Fw in (at z). h w\0\ unfolding constant_on_def frequently_def eventually_at h'_def apply simp by (metis \0 < r\ centre_in_ball dist_commute mem_ball that) moreover note \h' holomorphic_on ball z r\ ultimately obtain g r1 n where "0 < n" "0 < r1" "ball z r1 \ ball z r" and g:"g holomorphic_on ball z r1" "\w. w \ ball z r1 \ h' w = (w - z) ^ n * g w" "\w. w \ ball z r1 \ g w \ 0" using holomorphic_factor_zero_nonconstant[of _ "ball z r" z thesis,simplified, OF \h' holomorphic_on ball z r\ \r>0\ \h' z=0\ \\ h' constant_on ball z r\] by (auto simp add:dist_commute) define rr where "rr=r1/2" have "P h' n g rr" unfolding P_def rr_def using \n>0\ \r1>0\ g by (auto simp add:powr_nat) then have "P h n g rr" unfolding h'_def P_def by auto then show ?thesis unfolding P_def by blast qed moreover have ?thesis when "z'\0" proof - have "h' z\0" using that unfolding h'_def by auto obtain r1 where "r1>0" "cball z r1 \ ball z r" "\x\cball z r1. h' x\0" proof - have "isCont h' z" "h' z\0" by (auto simp add: Lim_cong_within \h \z\ z'\ \z'\0\ continuous_at h'_def) then obtain r2 where r2:"r2>0" "\x\ball z r2. h' x\0" using continuous_at_avoid[of z h' 0 ] unfolding ball_def by auto define r1 where "r1=min r2 r / 2" have "0 < r1" "cball z r1 \ ball z r" using \r2>0\ \r>0\ unfolding r1_def by auto moreover have "\x\cball z r1. h' x \ 0" using r2 unfolding r1_def by simp ultimately show ?thesis using that by auto qed then have "P h' 0 h' r1" using \h' holomorphic_on ball z r\ unfolding P_def by auto then have "P h 0 h' r1" unfolding P_def h'_def by auto then show ?thesis unfolding P_def by blast qed ultimately show ?thesis by auto qed have ?thesis when "\x. (f \ x) (at z)" apply (rule_tac imp_unique[unfolded P_def]) using P_exist[OF that(1) f_iso non_zero] unfolding P_def . moreover have ?thesis when "is_pole f z" proof (rule imp_unique[unfolded P_def]) obtain e where [simp]:"e>0" and e_holo:"f holomorphic_on ball z e - {z}" and e_nz: "\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 by auto then obtain e1 where e1:"e1>0" "\x\ball z e1-{z}. f x\0" using that eventually_at[of "\x. f x\0" z UNIV,simplified] by (auto simp add:dist_commute) obtain e2 where e2:"e2>0" "f holomorphic_on ball z e2 - {z}" using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by auto define e where "e=min e1 e2" show ?thesis apply (rule that[of e]) using e1 e2 unfolding e_def by auto qed define h where "h \ \x. inverse (f x)" have "\n g r. P h n g r" proof - have "h \z\ 0" using Lim_transform_within_open assms(2) h_def is_pole_tendsto that by fastforce moreover have "\\<^sub>Fw in (at z). h w\0" using non_zero apply (elim frequently_rev_mp) unfolding h_def eventually_at by (auto intro:exI[where x=1]) moreover have "isolated_singularity_at h z" unfolding isolated_singularity_at_def h_def apply (rule exI[where x=e]) using e_holo e_nz \e>0\ by (metis open_ball analytic_on_open holomorphic_on_inverse open_delete) ultimately show ?thesis using P_exist[of h] by auto qed then obtain n g r where "0 < r" and g_holo:"g holomorphic_on cball z r" and "g z\0" and g_fac:"(\w\cball z r-{z}. h w = g w * (w - z) powr of_int n \ g w \ 0)" unfolding P_def by auto have "P f (-n) (inverse o g) r" proof - have "f w = inverse (g w) * (w - z) powr of_int (- n)" when "w\cball z r - {z}" for w using g_fac[rule_format,of w] that unfolding h_def apply (auto simp add:powr_minus ) by (metis inverse_inverse_eq inverse_mult_distrib) then show ?thesis unfolding P_def comp_def using \r>0\ g_holo g_fac \g z\0\ by (auto intro:holomorphic_intros) qed then show "\x g r. 0 < r \ g holomorphic_on cball z r \ g z \ 0 \ (\w\cball z r - {z}. f w = g w * (w - z) powr of_int x \ g w \ 0)" unfolding P_def by blast qed ultimately show ?thesis using \not_essential f z\ unfolding not_essential_def by presburger qed lemma not_essential_transform: assumes "not_essential g z" assumes "\\<^sub>F w in (at z). g w = f w" shows "not_essential f z" using assms unfolding not_essential_def by (simp add: filterlim_cong is_pole_cong) lemma isolated_singularity_at_transform: assumes "isolated_singularity_at g z" assumes "\\<^sub>F w in (at z). g w = f w" shows "isolated_singularity_at f z" proof - obtain r1 where "r1>0" and r1:"g analytic_on ball z r1 - {z}" using assms(1) unfolding isolated_singularity_at_def by auto obtain r2 where "r2>0" and r2:" \x. x \ z \ dist x z < r2 \ g x = f x" using assms(2) unfolding eventually_at by auto define r3 where "r3=min r1 r2" have "r3>0" unfolding r3_def using \r1>0\ \r2>0\ by auto moreover have "f analytic_on ball z r3 - {z}" proof - have "g holomorphic_on ball z r3 - {z}" using r1 unfolding r3_def by (subst (asm) analytic_on_open,auto) then have "f holomorphic_on ball z r3 - {z}" using r2 unfolding r3_def by (auto simp add:dist_commute elim!:holomorphic_transform) then show ?thesis by (subst analytic_on_open,auto) qed ultimately show ?thesis unfolding isolated_singularity_at_def by auto qed lemma not_essential_powr[singularity_intros]: assumes "LIM w (at z). f w :> (at x)" shows "not_essential (\w. (f w) powr (of_int n)) z" proof - define fp where "fp=(\w. (f w) powr (of_int n))" have ?thesis when "n>0" proof - have "(\w. (f w) ^ (nat n)) \z\ x ^ nat n" using that assms unfolding filterlim_at by (auto intro!:tendsto_eq_intros) then have "fp \z\ x ^ nat n" unfolding fp_def apply (elim Lim_transform_within[where d=1],simp) by (metis less_le powr_0 powr_of_int that zero_less_nat_eq zero_power) then show ?thesis unfolding not_essential_def fp_def by auto qed moreover have ?thesis when "n=0" proof - have "fp \z\ 1 " apply (subst tendsto_cong[where g="\_.1"]) using that filterlim_at_within_not_equal[OF assms,of 0] unfolding fp_def by auto then show ?thesis unfolding fp_def not_essential_def by auto qed moreover have ?thesis when "n<0" proof (cases "x=0") case True have "LIM w (at z). inverse ((f w) ^ (nat (-n))) :> at_infinity" apply (subst filterlim_inverse_at_iff[symmetric],simp) apply (rule filterlim_atI) subgoal using assms True that unfolding filterlim_at by (auto intro!:tendsto_eq_intros) subgoal using filterlim_at_within_not_equal[OF assms,of 0] by (eventually_elim,insert that,auto) done then have "LIM w (at z). fp w :> at_infinity" proof (elim filterlim_mono_eventually) show "\\<^sub>F x in at z. inverse (f x ^ nat (- n)) = fp x" using filterlim_at_within_not_equal[OF assms,of 0] unfolding fp_def apply eventually_elim using powr_of_int that by auto qed auto then show ?thesis unfolding fp_def not_essential_def is_pole_def by auto next case False let ?xx= "inverse (x ^ (nat (-n)))" have "(\w. inverse ((f w) ^ (nat (-n)))) \z\?xx" using assms False unfolding filterlim_at by (auto intro!:tendsto_eq_intros) then have "fp \z\?xx" apply (elim Lim_transform_within[where d=1],simp) unfolding fp_def by (metis inverse_zero nat_mono_iff nat_zero_as_int neg_0_less_iff_less not_le power_eq_0_iff powr_0 powr_of_int that) then show ?thesis unfolding fp_def not_essential_def by auto qed ultimately show ?thesis by linarith qed lemma isolated_singularity_at_powr[singularity_intros]: assumes "isolated_singularity_at f z" "\\<^sub>F w in (at z). f w\0" shows "isolated_singularity_at (\w. (f w) powr (of_int n)) z" proof - obtain r1 where "r1>0" "f analytic_on ball z r1 - {z}" using assms(1) unfolding isolated_singularity_at_def by auto then have r1:"f holomorphic_on ball z r1 - {z}" using analytic_on_open[of "ball z r1-{z}" f] by blast obtain r2 where "r2>0" and r2:"\w. w \ z \ dist w z < r2 \ f w \ 0" using assms(2) unfolding eventually_at by auto define r3 where "r3=min r1 r2" have "(\w. (f w) powr of_int n) holomorphic_on ball z r3 - {z}" apply (rule holomorphic_on_powr_of_int) subgoal unfolding r3_def using r1 by auto subgoal unfolding r3_def using r2 by (auto simp add:dist_commute) done moreover have "r3>0" unfolding r3_def using \0 < r1\ \0 < r2\ by linarith ultimately show ?thesis unfolding isolated_singularity_at_def apply (subst (asm) analytic_on_open[symmetric]) by auto qed lemma non_zero_neighbour: assumes f_iso:"isolated_singularity_at f z" and f_ness:"not_essential f z" and f_nconst:"\\<^sub>Fw in (at z). f w\0" shows "\\<^sub>F w in (at z). f w\0" proof - obtain fn fp fr where [simp]:"fp z \ 0" and "fr > 0" and fr: "fp holomorphic_on cball z fr" "\w\cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \ fp w \ 0" using holomorphic_factor_puncture[OF f_iso f_ness f_nconst,THEN ex1_implies_ex] by auto have "f w \ 0" when " w \ z" "dist w z < fr" for w proof - have "f w = fp w * (w - z) powr of_int fn" "fp w \ 0" using fr(2)[rule_format, of w] using that by (auto simp add:dist_commute) moreover have "(w - z) powr of_int fn \0" unfolding powr_eq_0_iff using \w\z\ by auto ultimately show ?thesis by auto qed then show ?thesis using \fr>0\ unfolding eventually_at by auto qed lemma non_zero_neighbour_pole: assumes "is_pole f z" shows "\\<^sub>F w in (at z). f w\0" using assms filterlim_at_infinity_imp_eventually_ne[of f "at z" 0] unfolding is_pole_def by auto lemma non_zero_neighbour_alt: assumes holo: "f holomorphic_on S" and "open S" "connected S" "z \ S" "\ \ S" "f \ \ 0" shows "\\<^sub>F w in (at z). f w\0 \ w\S" proof (cases "f z = 0") case True from isolated_zeros[OF holo \open S\ \connected S\ \z \ S\ True \\ \ S\ \f \ \ 0\] obtain r where "0 < r" "ball z r \ S" "\w \ ball z r - {z}.f w \ 0" by metis then show ?thesis unfolding eventually_at apply (rule_tac x=r in exI) by (auto simp add:dist_commute) next case False obtain r1 where r1:"r1>0" "\y. dist z y < r1 \ f y \ 0" using continuous_at_avoid[of z f, OF _ False] assms(2,4) continuous_on_eq_continuous_at holo holomorphic_on_imp_continuous_on by blast obtain r2 where r2:"r2>0" "ball z r2 \ S" using assms(2) assms(4) openE by blast show ?thesis unfolding eventually_at apply (rule_tac x="min r1 r2" in exI) using r1 r2 by (auto simp add:dist_commute) qed lemma not_essential_times[singularity_intros]: assumes f_ness:"not_essential f z" and g_ness:"not_essential g z" assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" shows "not_essential (\w. f w * g w) z" proof - define fg where "fg = (\w. f w * g w)" have ?thesis when "\ ((\\<^sub>Fw in (at z). f w\0) \ (\\<^sub>Fw in (at z). g w\0))" proof - have "\\<^sub>Fw in (at z). fg w=0" using that[unfolded frequently_def, simplified] unfolding fg_def by (auto elim: eventually_rev_mp) from tendsto_cong[OF this] have "fg \z\0" by auto then show ?thesis unfolding not_essential_def fg_def by auto qed moreover have ?thesis when f_nconst:"\\<^sub>Fw in (at z). f w\0" and g_nconst:"\\<^sub>Fw in (at z). g w\0" proof - obtain fn fp fr where [simp]:"fp z \ 0" and "fr > 0" and fr: "fp holomorphic_on cball z fr" "\w\cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \ fp w \ 0" using holomorphic_factor_puncture[OF f_iso f_ness f_nconst,THEN ex1_implies_ex] by auto obtain gn gp gr where [simp]:"gp z \ 0" and "gr > 0" and gr: "gp holomorphic_on cball z gr" "\w\cball z gr - {z}. g w = gp w * (w - z) powr of_int gn \ gp w \ 0" using holomorphic_factor_puncture[OF g_iso g_ness g_nconst,THEN ex1_implies_ex] by auto define r1 where "r1=(min fr gr)" have "r1>0" unfolding r1_def using \fr>0\ \gr>0\ by auto have fg_times:"fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" and fgp_nz:"fp w*gp w\0" when "w\ball z r1 - {z}" for w proof - have "f w = fp w * (w - z) powr of_int fn" "fp w\0" using fr(2)[rule_format,of w] that unfolding r1_def by auto moreover have "g w = gp w * (w - z) powr of_int gn" "gp w \ 0" using gr(2)[rule_format, of w] that unfolding r1_def by auto ultimately show "fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" "fp w*gp w\0" unfolding fg_def by (auto simp add:powr_add) qed have [intro]: "fp \z\fp z" "gp \z\gp z" using fr(1) \fr>0\ gr(1) \gr>0\ by (meson open_ball ball_subset_cball centre_in_ball continuous_on_eq_continuous_at continuous_within holomorphic_on_imp_continuous_on holomorphic_on_subset)+ have ?thesis when "fn+gn>0" proof - have "(\w. (fp w * gp w) * (w - z) ^ (nat (fn+gn))) \z\0" using that by (auto intro!:tendsto_eq_intros) then have "fg \z\ 0" apply (elim Lim_transform_within[OF _ \r1>0\]) by (metis (no_types, hide_lams) Diff_iff cball_trivial dist_commute dist_self eq_iff_diff_eq_0 fg_times less_le linorder_not_le mem_ball mem_cball powr_of_int that) then show ?thesis unfolding not_essential_def fg_def by auto qed moreover have ?thesis when "fn+gn=0" proof - have "(\w. fp w * gp w) \z\fp z*gp z" using that by (auto intro!:tendsto_eq_intros) then have "fg \z\ fp z*gp z" apply (elim Lim_transform_within[OF _ \r1>0\]) apply (subst fg_times) by (auto simp add:dist_commute that) then show ?thesis unfolding not_essential_def fg_def by auto qed moreover have ?thesis when "fn+gn<0" proof - have "LIM w (at z). fp w * gp w / (w-z)^nat (-(fn+gn)) :> at_infinity" apply (rule filterlim_divide_at_infinity) apply (insert that, auto intro!:tendsto_eq_intros filterlim_atI) using eventually_at_topological by blast then have "is_pole fg z" unfolding is_pole_def apply (elim filterlim_transform_within[OF _ _ \r1>0\],simp) apply (subst fg_times,simp add:dist_commute) apply (subst powr_of_int) using that by (auto simp add:field_split_simps) then show ?thesis unfolding not_essential_def fg_def by auto qed ultimately show ?thesis unfolding not_essential_def fg_def by fastforce qed ultimately show ?thesis by auto qed lemma not_essential_inverse[singularity_intros]: assumes f_ness:"not_essential f z" assumes f_iso:"isolated_singularity_at f z" shows "not_essential (\w. inverse (f w)) z" proof - define vf where "vf = (\w. inverse (f w))" have ?thesis when "\(\\<^sub>Fw in (at z). f w\0)" proof - have "\\<^sub>Fw in (at z). f w=0" using that[unfolded frequently_def, simplified] by (auto elim: eventually_rev_mp) then have "\\<^sub>Fw in (at z). vf w=0" unfolding vf_def by auto from tendsto_cong[OF this] have "vf \z\0" unfolding vf_def by auto then show ?thesis unfolding not_essential_def vf_def by auto qed moreover have ?thesis when "is_pole f z" proof - have "vf \z\0" using that filterlim_at filterlim_inverse_at_iff unfolding is_pole_def vf_def by blast then show ?thesis unfolding not_essential_def vf_def by auto qed moreover have ?thesis when "\x. f\z\x " and f_nconst:"\\<^sub>Fw in (at z). f w\0" proof - from that obtain fz where fz:"f\z\fz" by auto have ?thesis when "fz=0" proof - have "(\w. inverse (vf w)) \z\0" using fz that unfolding vf_def by auto moreover have "\\<^sub>F w in at z. inverse (vf w) \ 0" using non_zero_neighbour[OF f_iso f_ness f_nconst] unfolding vf_def by auto ultimately have "is_pole vf z" using filterlim_inverse_at_iff[of vf "at z"] unfolding filterlim_at is_pole_def by auto then show ?thesis unfolding not_essential_def vf_def by auto qed moreover have ?thesis when "fz\0" proof - have "vf \z\inverse fz" using fz that unfolding vf_def by (auto intro:tendsto_eq_intros) then show ?thesis unfolding not_essential_def vf_def by auto qed ultimately show ?thesis by auto qed ultimately show ?thesis using f_ness unfolding not_essential_def by auto qed lemma isolated_singularity_at_inverse[singularity_intros]: assumes f_iso:"isolated_singularity_at f z" and f_ness:"not_essential f z" shows "isolated_singularity_at (\w. inverse (f w)) z" proof - define vf where "vf = (\w. inverse (f w))" have ?thesis when "\(\\<^sub>Fw in (at z). f w\0)" proof - have "\\<^sub>Fw in (at z). f w=0" using that[unfolded frequently_def, simplified] by (auto elim: eventually_rev_mp) then have "\\<^sub>Fw in (at z). vf w=0" unfolding vf_def by auto then obtain d1 where "d1>0" and d1:"\x. x \ z \ dist x z < d1 \ vf x = 0" unfolding eventually_at by auto then have "vf holomorphic_on ball z d1-{z}" apply (rule_tac holomorphic_transform[of "\_. 0"]) by (auto simp add:dist_commute) then have "vf analytic_on ball z d1 - {z}" by (simp add: analytic_on_open open_delete) then show ?thesis using \d1>0\ unfolding isolated_singularity_at_def vf_def by auto qed moreover have ?thesis when f_nconst:"\\<^sub>Fw in (at z). f w\0" proof - have "\\<^sub>F w in at z. f w \ 0" using non_zero_neighbour[OF f_iso f_ness f_nconst] . then obtain d1 where d1:"d1>0" "\x. x \ z \ dist x z < d1 \ f x \ 0" unfolding eventually_at by auto obtain d2 where "d2>0" and d2:"f analytic_on ball z d2 - {z}" using f_iso unfolding isolated_singularity_at_def by auto define d3 where "d3=min d1 d2" have "d3>0" unfolding d3_def using \d1>0\ \d2>0\ by auto moreover have "vf analytic_on ball z d3 - {z}" unfolding vf_def apply (rule analytic_on_inverse) subgoal using d2 unfolding d3_def by (elim analytic_on_subset) auto subgoal for w using d1 unfolding d3_def by (auto simp add:dist_commute) done ultimately show ?thesis unfolding isolated_singularity_at_def vf_def by auto qed ultimately show ?thesis by auto qed lemma not_essential_divide[singularity_intros]: assumes f_ness:"not_essential f z" and g_ness:"not_essential g z" assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" shows "not_essential (\w. f w / g w) z" proof - have "not_essential (\w. f w * inverse (g w)) z" apply (rule not_essential_times[where g="\w. inverse (g w)"]) using assms by (auto intro: isolated_singularity_at_inverse not_essential_inverse) then show ?thesis by (simp add:field_simps) qed lemma assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" shows isolated_singularity_at_times[singularity_intros]: "isolated_singularity_at (\w. f w * g w) z" and isolated_singularity_at_add[singularity_intros]: "isolated_singularity_at (\w. f w + g w) z" proof - obtain d1 d2 where "d1>0" "d2>0" and d1:"f analytic_on ball z d1 - {z}" and d2:"g analytic_on ball z d2 - {z}" using f_iso g_iso unfolding isolated_singularity_at_def by auto define d3 where "d3=min d1 d2" have "d3>0" unfolding d3_def using \d1>0\ \d2>0\ by auto have "(\w. f w * g w) analytic_on ball z d3 - {z}" apply (rule analytic_on_mult) using d1 d2 unfolding d3_def by (auto elim:analytic_on_subset) then show "isolated_singularity_at (\w. f w * g w) z" using \d3>0\ unfolding isolated_singularity_at_def by auto have "(\w. f w + g w) analytic_on ball z d3 - {z}" apply (rule analytic_on_add) using d1 d2 unfolding d3_def by (auto elim:analytic_on_subset) then show "isolated_singularity_at (\w. f w + g w) z" using \d3>0\ unfolding isolated_singularity_at_def by auto qed lemma isolated_singularity_at_uminus[singularity_intros]: assumes f_iso:"isolated_singularity_at f z" shows "isolated_singularity_at (\w. - f w) z" using assms unfolding isolated_singularity_at_def using analytic_on_neg by blast lemma isolated_singularity_at_id[singularity_intros]: "isolated_singularity_at (\w. w) z" unfolding isolated_singularity_at_def by (simp add: gt_ex) lemma isolated_singularity_at_minus[singularity_intros]: assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" shows "isolated_singularity_at (\w. f w - g w) z" using isolated_singularity_at_uminus[THEN isolated_singularity_at_add[OF f_iso,of "\w. - g w"] ,OF g_iso] by simp lemma isolated_singularity_at_divide[singularity_intros]: assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" and g_ness:"not_essential g z" shows "isolated_singularity_at (\w. f w / g w) z" using isolated_singularity_at_inverse[THEN isolated_singularity_at_times[OF f_iso, of "\w. inverse (g w)"],OF g_iso g_ness] by (simp add:field_simps) lemma isolated_singularity_at_const[singularity_intros]: "isolated_singularity_at (\w. c) z" unfolding isolated_singularity_at_def by (simp add: gt_ex) lemma isolated_singularity_at_holomorphic: assumes "f holomorphic_on s-{z}" "open s" "z\s" shows "isolated_singularity_at f z" using assms unfolding isolated_singularity_at_def by (metis analytic_on_holomorphic centre_in_ball insert_Diff openE open_delete subset_insert_iff) subsubsection \The order of non-essential singularities (i.e. removable singularities or poles)\ definition\<^marker>\tag important\ zorder :: "(complex \ complex) \ complex \ int" where "zorder f z = (THE n. (\h r. r>0 \ h holomorphic_on cball z r \ h z\0 \ (\w\cball z r - {z}. f w = h w * (w-z) powr (of_int n) \ h w \0)))" definition\<^marker>\tag important\ zor_poly ::"[complex \ complex, complex] \ complex \ complex" where "zor_poly f z = (SOME h. \r. r > 0 \ h holomorphic_on cball z r \ h z \ 0 \ (\w\cball z r - {z}. f w = h w * (w - z) powr (zorder f z) \ h w \0))" lemma zorder_exist: fixes f::"complex \ complex" and z::complex defines "n\zorder f z" and "g\zor_poly f z" assumes f_iso:"isolated_singularity_at f z" and f_ness:"not_essential f z" and f_nconst:"\\<^sub>Fw in (at z). f w\0" shows "g z\0 \ (\r. r>0 \ g holomorphic_on cball z r \ (\w\cball z r - {z}. f w = g w * (w-z) powr n \ g w \0))" proof - define P where "P = (\n g r. 0 < r \ g holomorphic_on cball z r \ g z\0 \ (\w\cball z r - {z}. f w = g w * (w-z) powr (of_int n) \ g w\0))" have "\!n. \g r. P n g r" using holomorphic_factor_puncture[OF assms(3-)] unfolding P_def by auto then have "\g r. P n g r" unfolding n_def P_def zorder_def by (drule_tac theI',argo) then have "\r. P n g r" unfolding P_def zor_poly_def g_def n_def by (drule_tac someI_ex,argo) then obtain r1 where "P n g r1" by auto then show ?thesis unfolding P_def by auto qed lemma fixes f::"complex \ complex" and z::complex assumes f_iso:"isolated_singularity_at f z" and f_ness:"not_essential f z" and f_nconst:"\\<^sub>Fw in (at z). f w\0" shows zorder_inverse: "zorder (\w. inverse (f w)) z = - zorder f z" and zor_poly_inverse: "\\<^sub>Fw in (at z). zor_poly (\w. inverse (f w)) z w = inverse (zor_poly f z w)" proof - define vf where "vf = (\w. inverse (f w))" define fn vfn where "fn = zorder f z" and "vfn = zorder vf z" define fp vfp where "fp = zor_poly f z" and "vfp = zor_poly vf z" obtain fr where [simp]:"fp z \ 0" and "fr > 0" and fr: "fp holomorphic_on cball z fr" "\w\cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \ fp w \ 0" using zorder_exist[OF f_iso f_ness f_nconst,folded fn_def fp_def] by auto have fr_inverse: "vf w = (inverse (fp w)) * (w - z) powr (of_int (-fn))" and fr_nz: "inverse (fp w)\0" when "w\ball z fr - {z}" for w proof - have "f w = fp w * (w - z) powr of_int fn" "fp w\0" using fr(2)[rule_format,of w] that by auto then show "vf w = (inverse (fp w)) * (w - z) powr (of_int (-fn))" "inverse (fp w)\0" unfolding vf_def by (auto simp add:powr_minus) qed obtain vfr where [simp]:"vfp z \ 0" and "vfr>0" and vfr:"vfp holomorphic_on cball z vfr" "(\w\cball z vfr - {z}. vf w = vfp w * (w - z) powr of_int vfn \ vfp w \ 0)" proof - have "isolated_singularity_at vf z" using isolated_singularity_at_inverse[OF f_iso f_ness] unfolding vf_def . moreover have "not_essential vf z" using not_essential_inverse[OF f_ness f_iso] unfolding vf_def . moreover have "\\<^sub>F w in at z. vf w \ 0" using f_nconst unfolding vf_def by (auto elim:frequently_elim1) ultimately show ?thesis using zorder_exist[of vf z, folded vfn_def vfp_def] that by auto qed define r1 where "r1 = min fr vfr" have "r1>0" using \fr>0\ \vfr>0\ unfolding r1_def by simp show "vfn = - fn" apply (rule holomorphic_factor_unique[of r1 vfp z "\w. inverse (fp w)" vf]) subgoal using \r1>0\ by simp subgoal by simp subgoal by simp subgoal proof (rule ballI) fix w assume "w \ ball z r1 - {z}" then have "w \ ball z fr - {z}" "w \ cball z vfr - {z}" unfolding r1_def by auto from fr_inverse[OF this(1)] fr_nz[OF this(1)] vfr(2)[rule_format,OF this(2)] show "vf w = vfp w * (w - z) powr of_int vfn \ vfp w \ 0 \ vf w = inverse (fp w) * (w - z) powr of_int (- fn) \ inverse (fp w) \ 0" by auto qed subgoal using vfr(1) unfolding r1_def by (auto intro!:holomorphic_intros) subgoal using fr unfolding r1_def by (auto intro!:holomorphic_intros) done have "vfp w = inverse (fp w)" when "w\ball z r1-{z}" for w proof - have "w \ ball z fr - {z}" "w \ cball z vfr - {z}" "w\z" using that unfolding r1_def by auto from fr_inverse[OF this(1)] fr_nz[OF this(1)] vfr(2)[rule_format,OF this(2)] \vfn = - fn\ \w\z\ show ?thesis by auto qed then show "\\<^sub>Fw in (at z). vfp w = inverse (fp w)" unfolding eventually_at using \r1>0\ apply (rule_tac x=r1 in exI) by (auto simp add:dist_commute) qed lemma fixes f g::"complex \ complex" and z::complex assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" and f_ness:"not_essential f z" and g_ness:"not_essential g z" and fg_nconst: "\\<^sub>Fw in (at z). f w * g w\ 0" shows zorder_times:"zorder (\w. f w * g w) z = zorder f z + zorder g z" and zor_poly_times:"\\<^sub>Fw in (at z). zor_poly (\w. f w * g w) z w = zor_poly f z w *zor_poly g z w" proof - define fg where "fg = (\w. f w * g w)" define fn gn fgn where "fn = zorder f z" and "gn = zorder g z" and "fgn = zorder fg z" define fp gp fgp where "fp = zor_poly f z" and "gp = zor_poly g z" and "fgp = zor_poly fg z" have f_nconst:"\\<^sub>Fw in (at z). f w \ 0" and g_nconst:"\\<^sub>Fw in (at z).g w\ 0" using fg_nconst by (auto elim!:frequently_elim1) obtain fr where [simp]:"fp z \ 0" and "fr > 0" and fr: "fp holomorphic_on cball z fr" "\w\cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \ fp w \ 0" using zorder_exist[OF f_iso f_ness f_nconst,folded fp_def fn_def] by auto obtain gr where [simp]:"gp z \ 0" and "gr > 0" and gr: "gp holomorphic_on cball z gr" "\w\cball z gr - {z}. g w = gp w * (w - z) powr of_int gn \ gp w \ 0" using zorder_exist[OF g_iso g_ness g_nconst,folded gn_def gp_def] by auto define r1 where "r1=min fr gr" have "r1>0" unfolding r1_def using \fr>0\ \gr>0\ by auto have fg_times:"fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" and fgp_nz:"fp w*gp w\0" when "w\ball z r1 - {z}" for w proof - have "f w = fp w * (w - z) powr of_int fn" "fp w\0" using fr(2)[rule_format,of w] that unfolding r1_def by auto moreover have "g w = gp w * (w - z) powr of_int gn" "gp w \ 0" using gr(2)[rule_format, of w] that unfolding r1_def by auto ultimately show "fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" "fp w*gp w\0" unfolding fg_def by (auto simp add:powr_add) qed obtain fgr where [simp]:"fgp z \ 0" and "fgr > 0" and fgr: "fgp holomorphic_on cball z fgr" "\w\cball z fgr - {z}. fg w = fgp w * (w - z) powr of_int fgn \ fgp w \ 0" proof - have "fgp z \ 0 \ (\r>0. fgp holomorphic_on cball z r \ (\w\cball z r - {z}. fg w = fgp w * (w - z) powr of_int fgn \ fgp w \ 0))" apply (rule zorder_exist[of fg z, folded fgn_def fgp_def]) subgoal unfolding fg_def using isolated_singularity_at_times[OF f_iso g_iso] . subgoal unfolding fg_def using not_essential_times[OF f_ness g_ness f_iso g_iso] . subgoal unfolding fg_def using fg_nconst . done then show ?thesis using that by blast qed define r2 where "r2 = min fgr r1" have "r2>0" using \r1>0\ \fgr>0\ unfolding r2_def by simp show "fgn = fn + gn " apply (rule holomorphic_factor_unique[of r2 fgp z "\w. fp w * gp w" fg]) subgoal using \r2>0\ by simp subgoal by simp subgoal by simp subgoal proof (rule ballI) fix w assume "w \ ball z r2 - {z}" then have "w \ ball z r1 - {z}" "w \ cball z fgr - {z}" unfolding r2_def by auto from fg_times[OF this(1)] fgp_nz[OF this(1)] fgr(2)[rule_format,OF this(2)] show "fg w = fgp w * (w - z) powr of_int fgn \ fgp w \ 0 \ fg w = fp w * gp w * (w - z) powr of_int (fn + gn) \ fp w * gp w \ 0" by auto qed subgoal using fgr(1) unfolding r2_def r1_def by (auto intro!:holomorphic_intros) subgoal using fr(1) gr(1) unfolding r2_def r1_def by (auto intro!:holomorphic_intros) done have "fgp w = fp w *gp w" when "w\ball z r2-{z}" for w proof - have "w \ ball z r1 - {z}" "w \ cball z fgr - {z}" "w\z" using that unfolding r2_def by auto from fg_times[OF this(1)] fgp_nz[OF this(1)] fgr(2)[rule_format,OF this(2)] \fgn = fn + gn\ \w\z\ show ?thesis by auto qed then show "\\<^sub>Fw in (at z). fgp w = fp w * gp w" using \r2>0\ unfolding eventually_at by (auto simp add:dist_commute) qed lemma fixes f g::"complex \ complex" and z::complex assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" and f_ness:"not_essential f z" and g_ness:"not_essential g z" and fg_nconst: "\\<^sub>Fw in (at z). f w * g w\ 0" shows zorder_divide:"zorder (\w. f w / g w) z = zorder f z - zorder g z" and zor_poly_divide:"\\<^sub>Fw in (at z). zor_poly (\w. f w / g w) z w = zor_poly f z w / zor_poly g z w" proof - have f_nconst:"\\<^sub>Fw in (at z). f w \ 0" and g_nconst:"\\<^sub>Fw in (at z).g w\ 0" using fg_nconst by (auto elim!:frequently_elim1) define vg where "vg=(\w. inverse (g w))" have "zorder (\w. f w * vg w) z = zorder f z + zorder vg z" apply (rule zorder_times[OF f_iso _ f_ness,of vg]) subgoal unfolding vg_def using isolated_singularity_at_inverse[OF g_iso g_ness] . subgoal unfolding vg_def using not_essential_inverse[OF g_ness g_iso] . subgoal unfolding vg_def using fg_nconst by (auto elim!:frequently_elim1) done then show "zorder (\w. f w / g w) z = zorder f z - zorder g z" using zorder_inverse[OF g_iso g_ness g_nconst,folded vg_def] unfolding vg_def by (auto simp add:field_simps) have "\\<^sub>F w in at z. zor_poly (\w. f w * vg w) z w = zor_poly f z w * zor_poly vg z w" apply (rule zor_poly_times[OF f_iso _ f_ness,of vg]) subgoal unfolding vg_def using isolated_singularity_at_inverse[OF g_iso g_ness] . subgoal unfolding vg_def using not_essential_inverse[OF g_ness g_iso] . subgoal unfolding vg_def using fg_nconst by (auto elim!:frequently_elim1) done then show "\\<^sub>Fw in (at z). zor_poly (\w. f w / g w) z w = zor_poly f z w / zor_poly g z w" using zor_poly_inverse[OF g_iso g_ness g_nconst,folded vg_def] unfolding vg_def apply eventually_elim by (auto simp add:field_simps) qed lemma zorder_exist_zero: fixes f::"complex \ complex" and z::complex defines "n\zorder f z" and "g\zor_poly f z" assumes holo: "f holomorphic_on s" and "open s" "connected s" "z\s" and non_const: "\w\s. f w \ 0" shows "(if f z=0 then n > 0 else n=0) \ (\r. r>0 \ cball z r \ s \ g holomorphic_on cball z r \ (\w\cball z r. f w = g w * (w-z) ^ nat n \ g w \0))" proof - obtain r where "g z \ 0" and r: "r>0" "cball z r \ s" "g holomorphic_on cball z r" "(\w\cball z r - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" proof - have "g z \ 0 \ (\r>0. g holomorphic_on cball z r \ (\w\cball z r - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0))" proof (rule zorder_exist[of f z,folded g_def n_def]) show "isolated_singularity_at f z" unfolding isolated_singularity_at_def using holo assms(4,6) by (meson Diff_subset open_ball analytic_on_holomorphic holomorphic_on_subset openE) show "not_essential f z" unfolding not_essential_def using assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on by fastforce have "\\<^sub>F w in at z. f w \ 0 \ w\s" proof - obtain w where "w\s" "f w\0" using non_const by auto then show ?thesis by (rule non_zero_neighbour_alt[OF holo \open s\ \connected s\ \z\s\]) qed then show "\\<^sub>F w in at z. f w \ 0" apply (elim eventually_frequentlyE) by auto qed then obtain r1 where "g z \ 0" "r1>0" and r1:"g holomorphic_on cball z r1" "(\w\cball z r1 - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" by auto obtain r2 where r2: "r2>0" "cball z r2 \ s" using assms(4,6) open_contains_cball_eq by blast define r3 where "r3=min r1 r2" have "r3>0" "cball z r3 \ s" using \r1>0\ r2 unfolding r3_def by auto moreover have "g holomorphic_on cball z r3" using r1(1) unfolding r3_def by auto moreover have "(\w\cball z r3 - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" using r1(2) unfolding r3_def by auto ultimately show ?thesis using that[of r3] \g z\0\ by auto qed have if_0:"if f z=0 then n > 0 else n=0" proof - have "f\ z \ f z" by (metis assms(4,6,7) at_within_open continuous_on holo holomorphic_on_imp_continuous_on) then have "(\w. g w * (w - z) powr of_int n) \z\ f z" apply (elim Lim_transform_within_open[where s="ball z r"]) using r by auto moreover have "g \z\g z" by (metis (mono_tags, lifting) open_ball at_within_open_subset ball_subset_cball centre_in_ball continuous_on holomorphic_on_imp_continuous_on r(1,3) subsetCE) ultimately have "(\w. (g w * (w - z) powr of_int n) / g w) \z\ f z/g z" apply (rule_tac tendsto_divide) using \g z\0\ by auto then have powr_tendsto:"(\w. (w - z) powr of_int n) \z\ f z/g z" apply (elim Lim_transform_within_open[where s="ball z r"]) using r by auto have ?thesis when "n\0" "f z=0" proof - have "(\w. (w - z) ^ nat n) \z\ f z/g z" using powr_tendsto apply (elim Lim_transform_within[where d=r]) by (auto simp add: powr_of_int \n\0\ \r>0\) then have *:"(\w. (w - z) ^ nat n) \z\ 0" using \f z=0\ by simp moreover have False when "n=0" proof - have "(\w. (w - z) ^ nat n) \z\ 1" using \n=0\ by auto then show False using * using LIM_unique zero_neq_one by blast qed ultimately show ?thesis using that by fastforce qed moreover have ?thesis when "n\0" "f z\0" proof - have False when "n>0" proof - have "(\w. (w - z) ^ nat n) \z\ f z/g z" using powr_tendsto apply (elim Lim_transform_within[where d=r]) by (auto simp add: powr_of_int \n\0\ \r>0\) moreover have "(\w. (w - z) ^ nat n) \z\ 0" using \n>0\ by (auto intro!:tendsto_eq_intros) ultimately show False using \f z\0\ \g z\0\ using LIM_unique divide_eq_0_iff by blast qed then show ?thesis using that by force qed moreover have False when "n<0" proof - have "(\w. inverse ((w - z) ^ nat (- n))) \z\ f z/g z" "(\w.((w - z) ^ nat (- n))) \z\ 0" subgoal using powr_tendsto powr_of_int that by (elim Lim_transform_within_open[where s=UNIV],auto) subgoal using that by (auto intro!:tendsto_eq_intros) done from tendsto_mult[OF this,simplified] have "(\x. inverse ((x - z) ^ nat (- n)) * (x - z) ^ nat (- n)) \z\ 0" . then have "(\x. 1::complex) \z\ 0" by (elim Lim_transform_within_open[where s=UNIV],auto) then show False using LIM_const_eq by fastforce qed ultimately show ?thesis by fastforce qed moreover have "f w = g w * (w-z) ^ nat n \ g w \0" when "w\cball z r" for w proof (cases "w=z") case True then have "f \z\f w" using assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on by fastforce then have "(\w. g w * (w-z) ^ nat n) \z\f w" proof (elim Lim_transform_within[OF _ \r>0\]) fix x assume "0 < dist x z" "dist x z < r" then have "x \ cball z r - {z}" "x\z" unfolding cball_def by (auto simp add: dist_commute) then have "f x = g x * (x - z) powr of_int n" using r(4)[rule_format,of x] by simp also have "... = g x * (x - z) ^ nat n" apply (subst powr_of_int) using if_0 \x\z\ by (auto split:if_splits) finally show "f x = g x * (x - z) ^ nat n" . qed moreover have "(\w. g w * (w-z) ^ nat n) \z\ g w * (w-z) ^ nat n" using True apply (auto intro!:tendsto_eq_intros) by (metis open_ball at_within_open_subset ball_subset_cball centre_in_ball continuous_on holomorphic_on_imp_continuous_on r(1) r(3) that) ultimately have "f w = g w * (w-z) ^ nat n" using LIM_unique by blast then show ?thesis using \g z\0\ True by auto next case False then have "f w = g w * (w - z) powr of_int n \ g w \ 0" using r(4) that by auto then show ?thesis using False if_0 powr_of_int by (auto split:if_splits) qed ultimately show ?thesis using r by auto qed lemma zorder_exist_pole: fixes f::"complex \ complex" and z::complex defines "n\zorder f z" and "g\zor_poly f z" assumes holo: "f holomorphic_on s-{z}" and "open s" "z\s" and "is_pole f z" shows "n < 0 \ g z\0 \ (\r. r>0 \ cball z r \ s \ g holomorphic_on cball z r \ (\w\cball z r - {z}. f w = g w / (w-z) ^ nat (- n) \ g w \0))" proof - obtain r where "g z \ 0" and r: "r>0" "cball z r \ s" "g holomorphic_on cball z r" "(\w\cball z r - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" proof - have "g z \ 0 \ (\r>0. g holomorphic_on cball z r \ (\w\cball z r - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0))" proof (rule zorder_exist[of f z,folded g_def n_def]) show "isolated_singularity_at f z" unfolding isolated_singularity_at_def using holo assms(4,5) by (metis analytic_on_holomorphic centre_in_ball insert_Diff openE open_delete subset_insert_iff) show "not_essential f z" unfolding not_essential_def using assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on by fastforce from non_zero_neighbour_pole[OF \is_pole f z\] show "\\<^sub>F w in at z. f w \ 0" apply (elim eventually_frequentlyE) by auto qed then obtain r1 where "g z \ 0" "r1>0" and r1:"g holomorphic_on cball z r1" "(\w\cball z r1 - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" by auto obtain r2 where r2: "r2>0" "cball z r2 \ s" using assms(4,5) open_contains_cball_eq by metis define r3 where "r3=min r1 r2" have "r3>0" "cball z r3 \ s" using \r1>0\ r2 unfolding r3_def by auto moreover have "g holomorphic_on cball z r3" using r1(1) unfolding r3_def by auto moreover have "(\w\cball z r3 - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" using r1(2) unfolding r3_def by auto ultimately show ?thesis using that[of r3] \g z\0\ by auto qed have "n<0" proof (rule ccontr) assume " \ n < 0" define c where "c=(if n=0 then g z else 0)" have [simp]:"g \z\ g z" by (metis open_ball at_within_open ball_subset_cball centre_in_ball continuous_on holomorphic_on_imp_continuous_on holomorphic_on_subset r(1) r(3) ) have "\\<^sub>F x in at z. f x = g x * (x - z) ^ nat n" unfolding eventually_at_topological apply (rule_tac exI[where x="ball z r"]) using r powr_of_int \\ n < 0\ by auto moreover have "(\x. g x * (x - z) ^ nat n) \z\c" proof (cases "n=0") case True then show ?thesis unfolding c_def by simp next case False then have "(\x. (x - z) ^ nat n) \z\ 0" using \\ n < 0\ by (auto intro!:tendsto_eq_intros) from tendsto_mult[OF _ this,of g "g z",simplified] show ?thesis unfolding c_def using False by simp qed ultimately have "f \z\c" using tendsto_cong by fast then show False using \is_pole f z\ at_neq_bot not_tendsto_and_filterlim_at_infinity unfolding is_pole_def by blast qed moreover have "\w\cball z r - {z}. f w = g w / (w-z) ^ nat (- n) \ g w \0" using r(4) \n<0\ powr_of_int by (metis Diff_iff divide_inverse eq_iff_diff_eq_0 insert_iff linorder_not_le) ultimately show ?thesis using r(1-3) \g z\0\ by auto qed lemma zorder_eqI: assumes "open s" "z \ s" "g holomorphic_on s" "g z \ 0" assumes fg_eq:"\w. \w \ s;w\z\ \ f w = g w * (w - z) powr n" shows "zorder f z = n" proof - have "continuous_on s g" by (rule holomorphic_on_imp_continuous_on) fact moreover have "open (-{0::complex})" by auto ultimately have "open ((g -` (-{0})) \ s)" unfolding continuous_on_open_vimage[OF \open s\] by blast moreover from assms have "z \ (g -` (-{0})) \ s" by auto ultimately obtain r where r: "r > 0" "cball z r \ s \ (g -` (-{0}))" unfolding open_contains_cball by blast let ?gg= "(\w. g w * (w - z) powr n)" define P where "P = (\n g r. 0 < r \ g holomorphic_on cball z r \ g z\0 \ (\w\cball z r - {z}. f w = g w * (w-z) powr (of_int n) \ g w\0))" have "P n g r" unfolding P_def using r assms(3,4,5) by auto then have "\g r. P n g r" by auto moreover have unique: "\!n. \g r. P n g r" unfolding P_def proof (rule holomorphic_factor_puncture) have "ball z r-{z} \ s" using r using ball_subset_cball by blast then have "?gg holomorphic_on ball z r-{z}" using \g holomorphic_on s\ r by (auto intro!: holomorphic_intros) then have "f holomorphic_on ball z r - {z}" apply (elim holomorphic_transform) using fg_eq \ball z r-{z} \ s\ by auto then show "isolated_singularity_at f z" unfolding isolated_singularity_at_def using analytic_on_open open_delete r(1) by blast next have "not_essential ?gg z" proof (intro singularity_intros) show "not_essential g z" by (meson \continuous_on s g\ assms(1) assms(2) continuous_on_eq_continuous_at isCont_def not_essential_def) show " \\<^sub>F w in at z. w - z \ 0" by (simp add: eventually_at_filter) then show "LIM w at z. w - z :> at 0" unfolding filterlim_at by (auto intro:tendsto_eq_intros) show "isolated_singularity_at g z" by (meson Diff_subset open_ball analytic_on_holomorphic assms(1,2,3) holomorphic_on_subset isolated_singularity_at_def openE) qed then show "not_essential f z" apply (elim not_essential_transform) unfolding eventually_at using assms(1,2) assms(5)[symmetric] by (metis dist_commute mem_ball openE subsetCE) show "\\<^sub>F w in at z. f w \ 0" unfolding frequently_at proof (rule,rule) fix d::real assume "0 < d" define z' where "z'=z+min d r / 2" have "z' \ z" " dist z' z < d " unfolding z'_def using \d>0\ \r>0\ by (auto simp add:dist_norm) moreover have "f z' \ 0" proof (subst fg_eq[OF _ \z'\z\]) have "z' \ cball z r" unfolding z'_def using \r>0\ \d>0\ by (auto simp add:dist_norm) then show " z' \ s" using r(2) by blast show "g z' * (z' - z) powr of_int n \ 0" using P_def \P n g r\ \z' \ cball z r\ calculation(1) by auto qed ultimately show "\x\UNIV. x \ z \ dist x z < d \ f x \ 0" by auto qed qed ultimately have "(THE n. \g r. P n g r) = n" by (rule_tac the1_equality) then show ?thesis unfolding zorder_def P_def by blast qed lemma simple_zeroI: assumes "open s" "z \ s" "g holomorphic_on s" "g z \ 0" assumes "\w. w \ s \ f w = g w * (w - z)" shows "zorder f z = 1" using assms(1-4) by (rule zorder_eqI) (use assms(5) in auto) lemma higher_deriv_power: shows "(deriv ^^ j) (\w. (w - z) ^ n) w = pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j)" proof (induction j arbitrary: w) case 0 thus ?case by auto next case (Suc j w) have "(deriv ^^ Suc j) (\w. (w - z) ^ n) w = deriv ((deriv ^^ j) (\w. (w - z) ^ n)) w" by simp also have "(deriv ^^ j) (\w. (w - z) ^ n) = (\w. pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j))" using Suc by (intro Suc.IH ext) also { have "(\ has_field_derivative of_nat (n - j) * pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - Suc j)) (at w)" using Suc.prems by (auto intro!: derivative_eq_intros) also have "of_nat (n - j) * pochhammer (of_nat (Suc n - j)) j = pochhammer (of_nat (Suc n - Suc j)) (Suc j)" by (cases "Suc j \ n", subst pochhammer_rec) (insert Suc.prems, simp_all add: algebra_simps Suc_diff_le pochhammer_0_left) finally have "deriv (\w. pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j)) w = \ * (w - z) ^ (n - Suc j)" by (rule DERIV_imp_deriv) } finally show ?case . qed lemma zorder_zero_eqI: assumes f_holo:"f holomorphic_on s" and "open s" "z \ s" assumes zero: "\i. i < nat n \ (deriv ^^ i) f z = 0" assumes nz: "(deriv ^^ nat n) f z \ 0" and "n\0" shows "zorder f z = n" proof - obtain r where [simp]:"r>0" and "ball z r \ s" using \open s\ \z\s\ openE by blast have nz':"\w\ball z r. f w \ 0" proof (rule ccontr) assume "\ (\w\ball z r. f w \ 0)" then have "eventually (\u. f u = 0) (nhds z)" using \r>0\ unfolding eventually_nhds apply (rule_tac x="ball z r" in exI) by auto then have "(deriv ^^ nat n) f z = (deriv ^^ nat n) (\_. 0) z" by (intro higher_deriv_cong_ev) auto also have "(deriv ^^ nat n) (\_. 0) z = 0" by (induction n) simp_all finally show False using nz by contradiction qed define zn g where "zn = zorder f z" and "g = zor_poly f z" obtain e where e_if:"if f z = 0 then 0 < zn else zn = 0" and [simp]:"e>0" and "cball z e \ ball z r" and g_holo:"g holomorphic_on cball z e" and e_fac:"(\w\cball z e. f w = g w * (w - z) ^ nat zn \ g w \ 0)" proof - have "f holomorphic_on ball z r" using f_holo \ball z r \ s\ by auto from that zorder_exist_zero[of f "ball z r" z,simplified,OF this nz',folded zn_def g_def] show ?thesis by blast qed from this(1,2,5) have "zn\0" "g z\0" subgoal by (auto split:if_splits) subgoal using \0 < e\ ball_subset_cball centre_in_ball e_fac by blast done define A where "A = (\i. of_nat (i choose (nat zn)) * fact (nat zn) * (deriv ^^ (i - nat zn)) g z)" have deriv_A:"(deriv ^^ i) f z = (if zn \ int i then A i else 0)" for i proof - have "eventually (\w. w \ ball z e) (nhds z)" using \cball z e \ ball z r\ \e>0\ by (intro eventually_nhds_in_open) auto hence "eventually (\w. f w = (w - z) ^ (nat zn) * g w) (nhds z)" apply eventually_elim by (use e_fac in auto) hence "(deriv ^^ i) f z = (deriv ^^ i) (\w. (w - z) ^ nat zn * g w) z" by (intro higher_deriv_cong_ev) auto also have "\ = (\j=0..i. of_nat (i choose j) * (deriv ^^ j) (\w. (w - z) ^ nat zn) z * (deriv ^^ (i - j)) g z)" using g_holo \e>0\ by (intro higher_deriv_mult[of _ "ball z e"]) (auto intro!: holomorphic_intros) also have "\ = (\j=0..i. if j = nat zn then of_nat (i choose nat zn) * fact (nat zn) * (deriv ^^ (i - nat zn)) g z else 0)" proof (intro sum.cong refl, goal_cases) case (1 j) have "(deriv ^^ j) (\w. (w - z) ^ nat zn) z = pochhammer (of_nat (Suc (nat zn) - j)) j * 0 ^ (nat zn - j)" by (subst higher_deriv_power) auto also have "\ = (if j = nat zn then fact j else 0)" by (auto simp: not_less pochhammer_0_left pochhammer_fact) also have "of_nat (i choose j) * \ * (deriv ^^ (i - j)) g z = (if j = nat zn then of_nat (i choose (nat zn)) * fact (nat zn) * (deriv ^^ (i - nat zn)) g z else 0)" by simp finally show ?case . qed also have "\ = (if i \ zn then A i else 0)" by (auto simp: A_def) finally show "(deriv ^^ i) f z = \" . qed have False when "nn\0\ by auto with nz show False by auto qed moreover have "n\zn" proof - have "g z \ 0" using e_fac[rule_format,of z] \e>0\ by simp then have "(deriv ^^ nat zn) f z \ 0" using deriv_A[of "nat zn"] by(auto simp add:A_def) then have "nat zn \ nat n" using zero[of "nat zn"] by linarith moreover have "zn\0" using e_if by (auto split:if_splits) ultimately show ?thesis using nat_le_eq_zle by blast qed ultimately show ?thesis unfolding zn_def by fastforce qed lemma assumes "eventually (\z. f z = g z) (at z)" "z = z'" shows zorder_cong:"zorder f z = zorder g z'" and zor_poly_cong:"zor_poly f z = zor_poly g z'" proof - define P where "P = (\ff n h r. 0 < r \ h holomorphic_on cball z r \ h z\0 \ (\w\cball z r - {z}. ff w = h w * (w-z) powr (of_int n) \ h w\0))" have "(\r. P f n h r) = (\r. P g n h r)" for n h proof - have *: "\r. P g n h r" if "\r. P f n h r" and "eventually (\x. f x = g x) (at z)" for f g proof - from that(1) obtain r1 where r1_P:"P f n h r1" by auto from that(2) obtain r2 where "r2>0" and r2_dist:"\x. x \ z \ dist x z \ r2 \ f x = g x" unfolding eventually_at_le by auto define r where "r=min r1 r2" have "r>0" "h z\0" using r1_P \r2>0\ unfolding r_def P_def by auto moreover have "h holomorphic_on cball z r" using r1_P unfolding P_def r_def by auto moreover have "g w = h w * (w - z) powr of_int n \ h w \ 0" when "w\cball z r - {z}" for w proof - have "f w = h w * (w - z) powr of_int n \ h w \ 0" using r1_P that unfolding P_def r_def by auto moreover have "f w=g w" using r2_dist[rule_format,of w] that unfolding r_def by (simp add: dist_commute) ultimately show ?thesis by simp qed ultimately show ?thesis unfolding P_def by auto qed from assms have eq': "eventually (\z. g z = f z) (at z)" by (simp add: eq_commute) show ?thesis by (rule iffI[OF *[OF _ assms(1)] *[OF _ eq']]) qed then show "zorder f z = zorder g z'" "zor_poly f z = zor_poly g z'" using \z=z'\ unfolding P_def zorder_def zor_poly_def by auto qed lemma zorder_nonzero_div_power: assumes "open s" "z \ s" "f holomorphic_on s" "f z \ 0" "n > 0" shows "zorder (\w. f w / (w - z) ^ n) z = - n" apply (rule zorder_eqI[OF \open s\ \z\s\ \f holomorphic_on s\ \f z\0\]) apply (subst powr_of_int) using \n>0\ by (auto simp add:field_simps) lemma zor_poly_eq: assumes "isolated_singularity_at f z" "not_essential f z" "\\<^sub>F w in at z. f w \ 0" shows "eventually (\w. zor_poly f z w = f w * (w - z) powr - zorder f z) (at z)" proof - obtain r where r:"r>0" "(\w\cball z r - {z}. f w = zor_poly f z w * (w - z) powr of_int (zorder f z))" using zorder_exist[OF assms] by blast then have *: "\w\ball z r - {z}. zor_poly f z w = f w * (w - z) powr - zorder f z" by (auto simp: field_simps powr_minus) have "eventually (\w. w \ ball z r - {z}) (at z)" using r eventually_at_ball'[of r z UNIV] by auto thus ?thesis by eventually_elim (insert *, auto) qed lemma zor_poly_zero_eq: assumes "f holomorphic_on s" "open s" "connected s" "z \ s" "\w\s. f w \ 0" shows "eventually (\w. zor_poly f z w = f w / (w - z) ^ nat (zorder f z)) (at z)" proof - obtain r where r:"r>0" "(\w\cball z r - {z}. f w = zor_poly f z w * (w - z) ^ nat (zorder f z))" using zorder_exist_zero[OF assms] by auto then have *: "\w\ball z r - {z}. zor_poly f z w = f w / (w - z) ^ nat (zorder f z)" by (auto simp: field_simps powr_minus) have "eventually (\w. w \ ball z r - {z}) (at z)" using r eventually_at_ball'[of r z UNIV] by auto thus ?thesis by eventually_elim (insert *, auto) qed lemma zor_poly_pole_eq: assumes f_iso:"isolated_singularity_at f z" "is_pole f z" shows "eventually (\w. zor_poly f z w = f w * (w - z) ^ nat (- zorder f z)) (at z)" proof - obtain e where [simp]:"e>0" and f_holo:"f holomorphic_on ball z e - {z}" using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by blast obtain r where r:"r>0" "(\w\cball z r - {z}. f w = zor_poly f z w / (w - z) ^ nat (- zorder f z))" using zorder_exist_pole[OF f_holo,simplified,OF \is_pole f z\] by auto then have *: "\w\ball z r - {z}. zor_poly f z w = f w * (w - z) ^ nat (- zorder f z)" by (auto simp: field_simps) have "eventually (\w. w \ ball z r - {z}) (at z)" using r eventually_at_ball'[of r z UNIV] by auto thus ?thesis by eventually_elim (insert *, auto) qed lemma zor_poly_eqI: fixes f :: "complex \ complex" and z0 :: complex defines "n \ zorder f z0" assumes "isolated_singularity_at f z0" "not_essential f z0" "\\<^sub>F w in at z0. f w \ 0" assumes lim: "((\x. f (g x) * (g x - z0) powr - n) \ c) F" assumes g: "filterlim g (at z0) F" and "F \ bot" shows "zor_poly f z0 z0 = c" proof - from zorder_exist[OF assms(2-4)] obtain r where r: "r > 0" "zor_poly f z0 holomorphic_on cball z0 r" "\w. w \ cball z0 r - {z0} \ f w = zor_poly f z0 w * (w - z0) powr n" unfolding n_def by blast from r(1) have "eventually (\w. w \ ball z0 r \ w \ z0) (at z0)" using eventually_at_ball'[of r z0 UNIV] by auto hence "eventually (\w. zor_poly f z0 w = f w * (w - z0) powr - n) (at z0)" by eventually_elim (insert r, auto simp: field_simps powr_minus) moreover have "continuous_on (ball z0 r) (zor_poly f z0)" using r by (intro holomorphic_on_imp_continuous_on) auto with r(1,2) have "isCont (zor_poly f z0) z0" by (auto simp: continuous_on_eq_continuous_at) hence "(zor_poly f z0 \ zor_poly f z0 z0) (at z0)" unfolding isCont_def . ultimately have "((\w. f w * (w - z0) powr - n) \ zor_poly f z0 z0) (at z0)" by (blast intro: Lim_transform_eventually) hence "((\x. f (g x) * (g x - z0) powr - n) \ zor_poly f z0 z0) F" by (rule filterlim_compose[OF _ g]) from tendsto_unique[OF \F \ bot\ this lim] show ?thesis . qed lemma zor_poly_zero_eqI: fixes f :: "complex \ complex" and z0 :: complex defines "n \ zorder f z0" assumes "f holomorphic_on A" "open A" "connected A" "z0 \ A" "\z\A. f z \ 0" assumes lim: "((\x. f (g x) / (g x - z0) ^ nat n) \ c) F" assumes g: "filterlim g (at z0) F" and "F \ bot" shows "zor_poly f z0 z0 = c" proof - from zorder_exist_zero[OF assms(2-6)] obtain r where r: "r > 0" "cball z0 r \ A" "zor_poly f z0 holomorphic_on cball z0 r" "\w. w \ cball z0 r \ f w = zor_poly f z0 w * (w - z0) ^ nat n" unfolding n_def by blast from r(1) have "eventually (\w. w \ ball z0 r \ w \ z0) (at z0)" using eventually_at_ball'[of r z0 UNIV] by auto hence "eventually (\w. zor_poly f z0 w = f w / (w - z0) ^ nat n) (at z0)" by eventually_elim (insert r, auto simp: field_simps) moreover have "continuous_on (ball z0 r) (zor_poly f z0)" using r by (intro holomorphic_on_imp_continuous_on) auto with r(1,2) have "isCont (zor_poly f z0) z0" by (auto simp: continuous_on_eq_continuous_at) hence "(zor_poly f z0 \ zor_poly f z0 z0) (at z0)" unfolding isCont_def . ultimately have "((\w. f w / (w - z0) ^ nat n) \ zor_poly f z0 z0) (at z0)" by (blast intro: Lim_transform_eventually) hence "((\x. f (g x) / (g x - z0) ^ nat n) \ zor_poly f z0 z0) F" by (rule filterlim_compose[OF _ g]) from tendsto_unique[OF \F \ bot\ this lim] show ?thesis . qed lemma zor_poly_pole_eqI: fixes f :: "complex \ complex" and z0 :: complex defines "n \ zorder f z0" assumes f_iso:"isolated_singularity_at f z0" and "is_pole f z0" assumes lim: "((\x. f (g x) * (g x - z0) ^ nat (-n)) \ c) F" assumes g: "filterlim g (at z0) F" and "F \ bot" shows "zor_poly f z0 z0 = c" proof - obtain r where r: "r > 0" "zor_poly f z0 holomorphic_on cball z0 r" proof - have "\\<^sub>F w in at z0. f w \ 0" using non_zero_neighbour_pole[OF \is_pole f z0\] by (auto elim:eventually_frequentlyE) moreover have "not_essential f z0" unfolding not_essential_def using \is_pole f z0\ by simp ultimately show ?thesis using that zorder_exist[OF f_iso,folded n_def] by auto qed from r(1) have "eventually (\w. w \ ball z0 r \ w \ z0) (at z0)" using eventually_at_ball'[of r z0 UNIV] by auto have "eventually (\w. zor_poly f z0 w = f w * (w - z0) ^ nat (-n)) (at z0)" using zor_poly_pole_eq[OF f_iso \is_pole f z0\] unfolding n_def . moreover have "continuous_on (ball z0 r) (zor_poly f z0)" using r by (intro holomorphic_on_imp_continuous_on) auto with r(1,2) have "isCont (zor_poly f z0) z0" by (auto simp: continuous_on_eq_continuous_at) hence "(zor_poly f z0 \ zor_poly f z0 z0) (at z0)" unfolding isCont_def . ultimately have "((\w. f w * (w - z0) ^ nat (-n)) \ zor_poly f z0 z0) (at z0)" by (blast intro: Lim_transform_eventually) hence "((\x. f (g x) * (g x - z0) ^ nat (-n)) \ zor_poly f z0 z0) F" by (rule filterlim_compose[OF _ g]) from tendsto_unique[OF \F \ bot\ this lim] show ?thesis . qed end \ No newline at end of file diff --git a/src/HOL/Complex_Analysis/Contour_Integration.thy b/src/HOL/Complex_Analysis/Contour_Integration.thy --- a/src/HOL/Complex_Analysis/Contour_Integration.thy +++ b/src/HOL/Complex_Analysis/Contour_Integration.thy @@ -1,1704 +1,1704 @@ section \Contour integration\ theory Contour_Integration imports "HOL-Analysis.Analysis" begin lemma lhopital_complex_simple: assumes "(f has_field_derivative f') (at z)" assumes "(g has_field_derivative g') (at z)" assumes "f z = 0" "g z = 0" "g' \ 0" "f' / g' = c" shows "((\w. f w / g w) \ c) (at z)" proof - have "eventually (\w. w \ z) (at z)" by (auto simp: eventually_at_filter) hence "eventually (\w. ((f w - f z) / (w - z)) / ((g w - g z) / (w - z)) = f w / g w) (at z)" by eventually_elim (simp add: assms field_split_simps) moreover have "((\w. ((f w - f z) / (w - z)) / ((g w - g z) / (w - z))) \ f' / g') (at z)" by (intro tendsto_divide has_field_derivativeD assms) ultimately have "((\w. f w / g w) \ f' / g') (at z)" by (blast intro: Lim_transform_eventually) with assms show ?thesis by simp qed subsection\Definition\ text\ This definition is for complex numbers only, and does not generalise to line integrals in a vector field \ definition\<^marker>\tag important\ 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\<^marker>\tag important\ contour_integrable_on (infixr "contour'_integrable'_on" 50) where "f contour_integrable_on g \ \i. (f has_contour_integral i) g" definition\<^marker>\tag important\ contour_integral where "contour_integral g f \ SOME 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 lemma 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 text\Show that we can forget about the localized derivative.\ 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}" + "((\x. f (g x) * vector_derivative p (at x within {a..b})) has_integral i) {a..b} \ + ((\x. f (g x) * vector_derivative p (at x)) has_integral i) {a..b}" proof - have *: "{a..b} - {a,b} = interior {a..b}" by (simp add: atLeastAtMost_diff_ends) show ?thesis by (rule has_integral_spike_eq [of "{a,b}"]) (auto simp: at_within_interior [of _ "{a..b}"]) 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}" + "(\x. f (g x) * vector_derivative p (at x within {a..b})) integrable_on {a..b} \ + (\x. f (g x) * vector_derivative p (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\<^marker>\tag unimportant\ \Reversing a path\ lemma has_contour_integral_reversepath: assumes "valid_path g" and f: "(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 \ (-) 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 \ (\x. 1 - x) has_vector_derivative -1 *\<^sub>R f') (at x)" by (intro vector_diff_chain_within has_vector_derivative_at_within [OF f'] derivative_eq_intros | simp)+ 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 obtain S where S: "continuous_on {0..1} g" "finite S" "g C1_differentiable_on {0..1} - S" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) have "((\x. - (f (g (1 - x)) * vector_derivative g (at (1 - x) within {0..1}))) has_integral -i) {0..1}" using has_integral_affinity01 [where m= "-1" and c=1, OF f [unfolded has_contour_integral_def]] by (simp add: has_integral_neg) then show ?thesis using S unfolding reversepath_def has_contour_integral_def by (rule_tac S = "(\x. 1 - x) ` S" in has_integral_spike_finite) (auto simp: *) 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\<^marker>\tag unimportant\ \Joining two paths together\ 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: "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))" if "0 \ z" "z*2 < 1" "z*2 \ s1" for z proof (rule vector_derivative_at [OF has_vector_derivative_transform_within]) show "0 < \z - 1/2\" using that by auto have "((*) 2 has_vector_derivative 2) (at z)" by (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) moreover have "(g1 has_vector_derivative vector_derivative g1 (at (z * 2))) (at (2 * z))" using s1 that by (auto simp: algebra_simps vector_derivative_works) ultimately show "((\x. g1 (2 * x)) has_vector_derivative 2 *\<^sub>R vector_derivative g1 (at (z * 2))) (at z)" by (intro vector_diff_chain_at [simplified o_def]) qed (use that in \simp_all add: dist_real_def abs_if split: if_split_asm\) have g2: "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))" if "1 < z*2" "z \ 1" "z*2 - 1 \ s2" for z proof (rule vector_derivative_at [OF has_vector_derivative_transform_within]) show "0 < \z - 1/2\" using that by auto have "((\x. 2 * x - 1) has_vector_derivative 2) (at z)" by (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) moreover have "(g2 has_vector_derivative vector_derivative g2 (at (z * 2 - 1))) (at (2 * z - 1))" using s2 that by (auto simp: algebra_simps vector_derivative_works) ultimately show "((\x. g2 (2 * x - 1)) has_vector_derivative 2 *\<^sub>R vector_derivative g2 (at (z * 2 - 1))) (at z)" by (intro vector_diff_chain_at [simplified o_def]) qed (use that in \simp_all add: dist_real_def abs_if split: if_split_asm\) have "((\x. f ((g1 +++ g2) x) * vector_derivative (g1 +++ g2) (at x)) has_integral i1) {0..1/2}" proof (rule has_integral_spike_finite [OF _ _ i1]) show "finite (insert (1/2) ((*) 2 -` s1))" using s1 by (force intro: finite_vimageI [where h = "(*)2"] inj_onI) qed (auto simp add: joinpaths_def scaleR_conv_of_real mult_ac g1) moreover have "((\x. f ((g1 +++ g2) x) * vector_derivative (g1 +++ g2) (at x)) has_integral i2) {1/2..1}" proof (rule has_integral_spike_finite [OF _ _ i2]) show "finite (insert (1/2) ((\x. 2 * x - 1) -` s2))" using s2 by (force intro: finite_vimageI [where h = "\x. 2*x-1"] inj_onI) qed (auto simp add: joinpaths_def scaleR_conv_of_real mult_ac g2) ultimately show ?thesis by (simp add: has_contour_integral has_integral_combine [where c = "1/2"]) 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 integrable_affinity [of _ 0 "1/2" "1/2" 0] integrable_on_subcbox [where a=0 and b="1/2"] by (fastforce simp: contour_integrable_on) 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: "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)" if "0 < z" "z < 1" "z \ s1" for z proof (rule vector_derivative_at [OF has_vector_derivative_transform_within]) show "0 < \(z - 1)/2\" using that by auto have \
: "((\x. x * 2) has_vector_derivative 2) (at (z/2))" using s1 by (auto simp: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) have "(g1 has_vector_derivative vector_derivative g1 (at z)) (at z)" using s1 that by (auto simp: vector_derivative_works) then show "((\x. g1 (2 * x)) has_vector_derivative 2 *\<^sub>R vector_derivative g1 (at z)) (at (z/2))" using vector_diff_chain_at [OF \
] by (auto simp: field_simps o_def) qed (use that in \simp_all add: field_simps dist_real_def abs_if split: if_split_asm\) have fin01: "finite ({0, 1} \ s1)" by (simp add: s1) show ?thesis unfolding contour_integrable_on by (intro integrable_spike_finite [OF fin01 _ *]) (auto simp: joinpaths_def scaleR_conv_of_real g1) 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 integrable_affinity [of _ "1/2::real" 1 "1/2" "1/2"] integrable_on_subcbox [where a="1/2" and b=1] by (fastforce simp: contour_integrable_on image_affinity_atLeastAtMost_diff) 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: "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)" if "0 < z" "z < 1" "z \ s2" for z proof (rule vector_derivative_at [OF has_vector_derivative_transform_within]) show "0 < \z/2\" using that by auto have \
: "((\x. x * 2 - 1) has_vector_derivative 2) (at ((1 + z)/2))" using s2 by (auto simp: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) have "(g2 has_vector_derivative vector_derivative g2 (at z)) (at z)" using s2 that by (auto simp: vector_derivative_works) then show "((\x. g2 (2*x - 1)) has_vector_derivative 2 *\<^sub>R vector_derivative g2 (at z)) (at (z/2 + 1/2))" using vector_diff_chain_at [OF \
] by (auto simp: field_simps o_def) qed (use that in \simp_all add: field_simps dist_real_def abs_if split: if_split_asm\) have fin01: "finite ({0, 1} \ s2)" by (simp add: s2) show ?thesis unfolding contour_integrable_on by (intro integrable_spike_finite [OF fin01 _ *]) (auto simp: joinpaths_def scaleR_conv_of_real g2) qed lemma contour_integrable_join [simp]: "\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]: "\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\<^marker>\tag unimportant\ \Shifting the starting point of a (closed) path\ 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 Henstock_Kurzweil_Integration.integral_combine) using assms * integral_unique by auto have vd1: "vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a))" if "0 \ x" "x + a < 1" "x \ (\x. x - a) ` S" for x unfolding shiftpath_def proof (rule vector_derivative_at [OF has_vector_derivative_transform_within]) have "((\x. g (x + a)) has_vector_derivative vector_derivative g (at (a + x))) (at x)" proof (rule vector_diff_chain_at [of _ 1, simplified o_def scaleR_one]) show "((\x. x + a) has_vector_derivative 1) (at x)" by (rule derivative_eq_intros | simp)+ have "g differentiable at (x + a)" using g a that by force then show "(g has_vector_derivative vector_derivative g (at (a + x))) (at (x + a))" by (metis add.commute vector_derivative_works) qed then show "((\x. g (a + x)) has_vector_derivative vector_derivative g (at (x + a))) (at x)" by (auto simp: field_simps) show "0 < dist (1 - a) x" using that by auto qed (use that in \auto simp: dist_real_def\) have vd2: "vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a - 1))" if "x \ 1" "1 < x + a" "x \ (\x. x - a + 1) ` S" for x unfolding shiftpath_def proof (rule vector_derivative_at [OF has_vector_derivative_transform_within]) have "((\x. g (x + a - 1)) has_vector_derivative vector_derivative g (at (a+x-1))) (at x)" proof (rule vector_diff_chain_at [of _ 1, simplified o_def scaleR_one]) show "((\x. x + a - 1) has_vector_derivative 1) (at x)" by (rule derivative_eq_intros | simp)+ have "g differentiable at (x+a-1)" using g a that by force then show "(g has_vector_derivative vector_derivative g (at (a+x-1))) (at (x + a - 1))" by (metis add.commute vector_derivative_works) qed then show "((\x. g (a + x - 1)) has_vector_derivative vector_derivative g (at (x + a - 1))) (at x)" by (auto simp: field_simps) show "0 < dist (1 - a) x" using that by auto qed (use that in \auto simp: dist_real_def\) 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})" using * a by (force intro: integrable_subinterval_real) have "finite ({1 - a} \ (\x. x - a) ` S)" using S by blast then 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 f = "\x. f(g(a+x)) * vector_derivative g (at(a+x))"]) subgoal using a by (simp add: vd1) (force simp: shiftpath_def add.commute) subgoal using has_integral_affinity [where m=1 and c=a] integrable_integral [OF va1] by (force simp add: add.commute) done moreover have "finite ({1 - a} \ (\x. x - a + 1) ` S)" using S by blast then 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 f = "\x. f(g(a+x-1)) * vector_derivative g (at(a+x-1))"]) subgoal using a by (simp add: vd2) (force simp: shiftpath_def add.commute) subgoal using has_integral_affinity [where m=1 and c="a-1", simplified, OF integrable_integral [OF v0a]] by (force 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 \
: "shiftpath (1 - a) (shiftpath a g) differentiable at x" using assms x by (intro differentiable_transform_within [OF gx, of "min x (1-x)"]) (auto simp: dist_real_def shiftpath_shiftpath abs_if split: if_split_asm) 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 assms x \
apply (auto simp: finite_imp_closed open_Diff shiftpath_shiftpath at_within_interior [of _ "{0..1}"] vector_derivative_works [symmetric]) 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 unfolding has_contour_integral_def proof (rule has_integral_spike_finite [of "{0,1} \ S", OF _ _ fi [unfolded has_contour_integral_def]]) show "finite ({0, 1} \ S)" by (simp add: S) qed (use S assms vd in \auto simp: shiftpath_shiftpath\) 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\<^marker>\tag unimportant\ \More about straight-line paths\ 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) lemma has_contour_integral_trivial [iff]: "(f has_contour_integral 0) (linepath a a)" by (simp add: has_contour_integral_linepath) lemma has_contour_integral_trivial_iff [simp]: "(f has_contour_integral i) (linepath a a) \ i=0" using has_contour_integral_unique by blast 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 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: 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 \
: "(\t. f (g t) * vector_derivative g (at t)) integrable_on {u..v}" using contour_integrable_on f integrable_on_subinterval uv by fastforce then 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 uv False unfolding has_integral_integral apply simp apply (drule has_integral_affinity [where m="v-u" and c=u, simplified]) apply (simp_all add: image_affinity_atLeastAtMost_div_diff scaleR_conv_of_real) apply (simp add: divide_simps) done have vd: "vector_derivative (\x. g ((v-u) * x + u)) (at x) = (v-u) *\<^sub>R vector_derivative g (at ((v-u) * x + u))" if "x \ {0..1}" "x \ (\t. (v-u) *\<^sub>R t + u) -` S" for x proof (rule vector_derivative_at [OF vector_diff_chain_at [simplified o_def]]) show "((\x. (v - u) * x + u) has_vector_derivative v - u) (at x)" by (intro derivative_eq_intros | simp)+ qed (use S uv mult_left_le [of x "v-u"] that in \auto simp: vector_derivative_works\) have fin: "finite ((\t. (v - u) *\<^sub>R t + u) -` S)" using fs by (auto simp: inj_on_def False finite_vimageI) show ?thesis unfolding subpath_def has_contour_integral apply (rule has_integral_spike_finite [OF fin]) using has_integral_cmul [OF *, where c = "v-u"] fs assms by (auto simp: False vd scaleR_conv_of_real) 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)" proof (cases u v rule: linorder_class.le_cases) case le then show ?thesis by (metis contour_integrable_on_def has_contour_integral_subpath [OF assms]) next case ge with assms show ?thesis by (metis (no_types, lifting) contour_integrable_on_def contour_integrable_reversepath_eq has_contour_integral_subpath reversepath_subpath valid_path_subpath) qed 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 proof - have "(\r. f (g r) * vector_derivative g (at r)) integrable_on {u..v}" by (metis (full_types) assms(1) assms(3) assms(4) atLeastAtMost_iff atLeastatMost_subset_iff contour_integrable_on integrable_on_subinterval) then have "((\r. f (g r) * vector_derivative g (at r)) has_integral integral {u..v} (\r. f (g r) * vector_derivative g (at r))) {u..v}" by blast then show ?thesis by (metis (full_types) assms contour_integral_unique has_contour_integral_subpath) qed 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}" "ux. f (g x) * vector_derivative g (at x)) integrable_on {u..w}" using integrable_on_subcbox [where a=u and b=w and S = "{0..1}"] assms by (auto simp: contour_integrable_on) with assms show ?thesis by (auto simp: contour_integral_subcontour_integral Henstock_Kurzweil_Integration.integral_combine) qed lemma contour_integral_subpath_combine: assumes "f contour_integrable_on g" "valid_path g" "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] by (elim disjE) (auto simp: * contour_integral_reversepath contour_integrable_subpath valid_path_subpath algebra_simps) next case False with assms show ?thesis by (metis add.right_neutral contour_integral_reversepath contour_integral_subpath_refl diff_0 eq_diff_eq add_0 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) lemma contour_integral_cong: assumes "g = g'" "\x. x \ path_image g \ f x = f' x" shows "contour_integral g f = contour_integral g' f'" unfolding contour_integral_integral using assms by (intro integral_cong) (auto simp: path_image_def) text \Contour integral along a segment on the real axis\ lemma has_contour_integral_linepath_Reals_iff: fixes a b :: complex and f :: "complex \ complex" assumes "a \ Reals" "b \ Reals" "Re a < Re b" shows "(f has_contour_integral I) (linepath a b) \ ((\x. f (of_real x)) has_integral I) {Re a..Re b}" proof - from assms have [simp]: "of_real (Re a) = a" "of_real (Re b) = b" by (simp_all add: complex_eq_iff) from assms have "a \ b" by auto have "((\x. f (of_real x)) has_integral I) (cbox (Re a) (Re b)) \ ((\x. f (a + b * of_real x - a * of_real x)) has_integral I /\<^sub>R (Re b - Re a)) {0..1}" by (subst has_integral_affinity_iff [of "Re b - Re a" _ "Re a", symmetric]) (insert assms, simp_all add: field_simps scaleR_conv_of_real) also have "(\x. f (a + b * of_real x - a * of_real x)) = (\x. (f (a + b * of_real x - a * of_real x) * (b - a)) /\<^sub>R (Re b - Re a))" using \a \ b\ by (auto simp: field_simps fun_eq_iff scaleR_conv_of_real) also have "(\ has_integral I /\<^sub>R (Re b - Re a)) {0..1} \ ((\x. f (linepath a b x) * (b - a)) has_integral I) {0..1}" using assms by (subst has_integral_cmul_iff) (auto simp: linepath_def scaleR_conv_of_real algebra_simps) also have "\ \ (f has_contour_integral I) (linepath a b)" unfolding has_contour_integral_def by (intro has_integral_cong) (simp add: vector_derivative_linepath_within) finally show ?thesis by simp qed lemma contour_integrable_linepath_Reals_iff: fixes a b :: complex and f :: "complex \ complex" assumes "a \ Reals" "b \ Reals" "Re a < Re b" shows "(f contour_integrable_on linepath a b) \ (\x. f (of_real x)) integrable_on {Re a..Re b}" using has_contour_integral_linepath_Reals_iff[OF assms, of f] by (auto simp: contour_integrable_on_def integrable_on_def) lemma contour_integral_linepath_Reals_eq: fixes a b :: complex and f :: "complex \ complex" assumes "a \ Reals" "b \ Reals" "Re a < Re b" shows "contour_integral (linepath a b) f = integral {Re a..Re b} (\x. f (of_real x))" proof (cases "f contour_integrable_on linepath a b") case True thus ?thesis using has_contour_integral_linepath_Reals_iff[OF assms, of f] using has_contour_integral_integral has_contour_integral_unique by blast next case False thus ?thesis using contour_integrable_linepath_Reals_iff[OF assms, of f] by (simp add: not_integrable_contour_integral not_integrable_integral) qed subsection \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 "finite K" and 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 "continuous_on (g ` {a..b}) f" using assms by (metis field_differentiable_def field_differentiable_imp_continuous_at continuous_on_eq_continuous_within continuous_on_subset image_subset_iff) then have cfg: "continuous_on {a..b} (\x. f (g x))" by (rule continuous_on_compose [OF cg, unfolded o_def]) { 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 (*) (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 using assms cfg * by (force simp: at_within_Icc_at intro: fundamental_theorem_of_calculus_interior_strong [OF \finite K\]) 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 (closed_segment a b) (\x. f x * (b - a))" by (rule continuous_intros | simp add: assms)+ then have "continuous_on {0..1} (\x. f (linepath a b x) * (b - a))" by (metis (no_types, lifting) continuous_on_compose continuous_on_cong continuous_on_linepath linepath_image_01 o_apply) then have "(\x. f (linepath a b x) * vector_derivative (linepath a b) (at x within {0..1})) integrable_on {0..1}" by (metis (no_types, lifting) continuous_on_cong integrable_continuous_real vector_derivative_linepath_within) then show ?thesis by (simp add: contour_integrable_on_def has_contour_integral_def integrable_on_def [symmetric]) 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" using contour_integral_primitive [of UNIV "\x. x^2/2" "\x. x" "linepath a b"] contour_integral_unique by (simp add: has_field_der_id) lemma contour_integrable_on_const [iff]: "(\x. c) contour_integrable_on (linepath a b)" by (simp add: contour_integrable_continuous_linepath) lemma contour_integrable_on_id [iff]: "(\x. x) contour_integrable_on (linepath a b)" by (simp add: contour_integrable_continuous_linepath) subsection\<^marker>\tag unimportant\ \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_diff 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" by (simp add: has_contour_integral_def algebra_simps has_integral_mult_right) lemma has_contour_integral_rmul: "(f has_contour_integral i) g \ ((\x. (f x) * c) has_contour_integral (i*c)) g" by (simp add: mult.commute has_contour_integral_lmul) 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" by (metis (mono_tags, lifting) has_contour_integral_def has_integral_eq image_eqI path_image_def) lemma has_contour_integral_bound_linepath: assumes "(f has_contour_integral i) (linepath a b)" "0 \ B" and B: "\x. x \ closed_segment a b \ norm(f x) \ B" shows "norm i \ B * norm(b - a)" proof - have "norm i \ (B * norm (b - a)) * content (cbox 0 (1::real))" proof (rule has_integral_bound [of _ "\x. f (linepath a b x) * vector_derivative (linepath a b) (at x within {0..1})"]) show "cmod (f (linepath a b x) * vector_derivative (linepath a b) (at x within {0..1})) \ B * cmod (b - a)" if "x \ cbox 0 1" for x::real using that box_real(2) norm_mult by (metis B linepath_in_path mult_right_mono norm_ge_zero vector_derivative_linepath_within) qed (use assms has_contour_integral_def in auto) then show ?thesis by (auto simp: content_real) qed 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_sum: "\finite s; \a. a \ s \ (f a has_contour_integral i a) p\ \ ((\x. sum (\a. f a x) s) has_contour_integral sum i s) p" by (induction s rule: finite_induct) (auto simp: has_contour_integral_0 has_contour_integral_add) subsection\<^marker>\tag unimportant\ \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" using contour_integral_cong contour_integral_def by fastforce 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)" using has_contour_integral_bound_linepath [of f] by (auto simp: has_contour_integral_integral) lemma contour_integral_0 [simp]: "contour_integral g (\x. 0) = 0" by (simp add: contour_integral_unique has_contour_integral_0) lemma contour_integral_sum: "\finite s; \a. a \ s \ (f a) contour_integrable_on p\ \ contour_integral p (\x. sum (\a. f a x) s) = sum (\a. contour_integral p (f a)) s" by (auto simp: contour_integral_unique has_contour_integral_sum 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\<^marker>\tag unimportant\ \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_sum: "\finite s; \a. a \ s \ (f a) contour_integrable_on p\ \ (\x. sum (\a. f a x) s) contour_integrable_on p" unfolding contour_integrable_on_def by (metis has_contour_integral_sum) subsection\<^marker>\tag unimportant\ \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)" by (metis contour_integrable_continuous_linepath contour_integral_unique has_contour_integral_integral has_contour_integral_reverse_linepath) text \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 by auto next case False then have k: "0 < k" "k < 1" "complex_of_real k \ 1" using assms by auto 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. (x / k) *\<^sub>R a + ((k - x) / k) *\<^sub>R a = a" using False by (simp add: field_split_simps flip: real_vector.scale_left_distrib) then have "\x. ((k - x) / k) *\<^sub>R a + (x / k) *\<^sub>R c = (1 - x) *\<^sub>R a + x *\<^sub>R b" using False by (simp add: c' algebra_simps) then have "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral i) {0..k}" using k has_integral_affinity01 [OF *, of "inverse k" "0"] by (force dest: has_integral_cmul [where c = "inverse k"] simp add: divide_simps mult.commute [of _ "k"] image_affinity_atLeastAtMost c) } 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 unfolding c' scaleR_conv_of_real apply (simp add: divide_simps) apply (simp add: distrib_right distrib_left right_diff_distrib left_diff_distrib) done have "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral j) {k..1}" using k has_integral_affinity01 [OF *, of "inverse(1 - k)" "-(k/(1 - k))"] apply (simp add: divide_simps mult.commute [of _ "1-k"] image_affinity_atLeastAtMost ** bc) apply (auto dest: has_integral_cmul [where k = "(1 - k) *\<^sub>R j" and c = "inverse (1 - k)"]) done } note fj = this show ?thesis using f k unfolding has_contour_integral_linepath by (simp add: linepath_def has_integral_combine [OF _ _ fi fj]) 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) have "closed_segment a c \ closed_segment a b" by (metis c' ends_in_segment(1) in_segment(1) k subset_closed_segment) then show "continuous_on (closed_segment a c) f" by (rule continuous_on_subset [OF f]) 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 "closed_segment a c \ closed_segment a b" by (metis c' ends_in_segment(1) in_segment(1) k subset_closed_segment) moreover have "closed_segment c b \ closed_segment a b" by (metis c' ends_in_segment(2) in_segment(1) k subset_closed_segment) ultimately have *: "continuous_on (closed_segment a c) f" "continuous_on (closed_segment c b) f" by (auto intro: continuous_on_subset [OF f]) show ?thesis by (rule contour_integral_unique) (meson "*" c contour_integrable_continuous_linepath has_contour_integral_integral has_contour_integral_split k) 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]) 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) proposition 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) \ (\t. (g x, h t))" by (rule ext) simp have fgh2: "\x. (\t. f (g t) (h x)) = (\(y1,y2). f y1 y2) \ (\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 "continuous_on (cbox (0, 0) (1, 1::real)) ((\x. vector_derivative g (at x)) \ fst)" "continuous_on (cbox (0, 0) (1::real, 1)) ((\x. vector_derivative h (at x)) \ snd)" by (rule continuous_intros | simp add: gvcon hvcon)+ then have gvcon': "continuous_on (cbox (0, 0) (1, 1::real)) (\z. vector_derivative g (at (fst z)))" and hvcon': "continuous_on (cbox (0, 0) (1::real, 1)) (\x. vector_derivative h (at (snd x)))" by auto have "continuous_on (cbox (0, 0) (1, 1)) ((\(y1, y2). f y1 y2) \ (\w. ((g \ fst) w, (h \ snd) w)))" apply (intro gcon hcon continuous_intros | simp)+ apply (auto simp: path_image_def intro: continuous_on_subset [OF fcon]) done then have fgh: "continuous_on (cbox (0, 0) (1, 1)) (\x. f (g (fst x)) (h (snd x)))" by auto 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)))" proof (rule integral_cong [OF contour_integral_rmul [symmetric]]) have "\x. x \ {0..1} \ continuous_on {0..1} (\xa. f (g x) (h xa))" by (subst fgh1) (rule fcon_im1 hcon continuous_intros | simp)+ then show "\x. x \ {0..1} \ f (g x) contour_integrable_on h" unfolding contour_integrable_on using continuous_on_mult hvcon integrable_continuous_real by blast qed also have "\ = integral {0..1} (\y. contour_integral g (\x. f x (h y) * vector_derivative h (at y)))" unfolding contour_integral_integral apply (subst integral_swap_continuous [where 'a = real and 'b = real, of 0 0 1 1, simplified]) subgoal by (rule fgh gvcon' hvcon' continuous_intros | simp add: split_def)+ subgoal unfolding integral_mult_left [symmetric] by (simp only: mult_ac) done also have "\ = contour_integral h (\z. contour_integral g (\w. f w z))" unfolding contour_integral_integral integral_mult_left [symmetric] by (simp add: algebra_simps) finally show ?thesis by (simp add: contour_integral_integral) qed lemma valid_path_negatepath: "valid_path \ \ valid_path (uminus \ \)" unfolding o_def using piecewise_C1_differentiable_neg valid_path_def by blast lemma has_contour_integral_negatepath: assumes \: "valid_path \" and cint: "((\z. f (- z)) has_contour_integral - i) \" shows "(f has_contour_integral i) (uminus \ \)" proof - obtain S where cont: "continuous_on {0..1} \" and "finite S" and diff: "\ C1_differentiable_on {0..1} - S" using \ by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) have "((\x. - (f (- \ x) * vector_derivative \ (at x within {0..1}))) has_integral i) {0..1}" using cint by (auto simp: has_contour_integral_def dest: has_integral_neg) then have "((\x. f (- \ x) * vector_derivative (uminus \ \) (at x within {0..1})) has_integral i) {0..1}" proof (rule rev_iffD1 [OF _ has_integral_spike_eq]) show "negligible S" by (simp add: \finite S\ negligible_finite) show "f (- \ x) * vector_derivative (uminus \ \) (at x within {0..1}) = - (f (- \ x) * vector_derivative \ (at x within {0..1}))" if "x \ {0..1} - S" for x proof - have "vector_derivative (uminus \ \) (at x within cbox 0 1) = - vector_derivative \ (at x within cbox 0 1)" proof (rule vector_derivative_within_cbox) show "(uminus \ \ has_vector_derivative - vector_derivative \ (at x within cbox 0 1)) (at x within cbox 0 1)" using that unfolding o_def by (metis C1_differentiable_on_eq UNIV_I diff differentiable_subset has_vector_derivative_minus subsetI that vector_derivative_works) qed (use that in auto) then show ?thesis by simp qed qed then show ?thesis by (simp add: has_contour_integral_def) qed lemma contour_integrable_negatepath: assumes \: "valid_path \" and pi: "(\z. f (- z)) contour_integrable_on \" shows "f contour_integrable_on (uminus \ \)" by (metis \ add.inverse_inverse contour_integrable_on_def has_contour_integral_negatepath pi) 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) subsection\Partial circle path\ definition\<^marker>\tag important\ 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) lemma reversepath_part_circlepath[simp]: "reversepath (part_circlepath z r s t) = part_circlepath z r t s" unfolding part_circlepath_def reversepath_def linepath_def by (auto simp:algebra_simps) lemma 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)" unfolding part_circlepath_def linepath_def scaleR_conv_of_real by (rule has_vector_derivative_real_field derivative_eq_intros | simp)+ lemma differentiable_part_circlepath: "part_circlepath c r a b differentiable at x within A" using has_vector_derivative_part_circlepath[of c r a b x A] differentiableI_vector by blast lemma 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 lemma 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)" unfolding valid_path_def by (auto simp: C1_differentiable_on_eq vector_derivative_works vector_derivative_part_circlepath has_vector_derivative_part_circlepath intro!: C1_differentiable_imp_piecewise continuous_intros) 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) by (metis (no_types) affine_ineq mult.commute mult_left_mono) } 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: field_split_simps) done } ultimately show ?thesis by (fastforce simp add: path_image_def part_circlepath_def) qed lemma path_image_part_circlepath': "path_image (part_circlepath z r s t) = (\x. z + r * cis x) ` closed_segment s t" proof - have "path_image (part_circlepath z r s t) = (\x. z + r * exp(\ * of_real x)) ` linepath s t ` {0..1}" by (simp add: image_image path_image_def part_circlepath_def) also have "linepath s t ` {0..1} = closed_segment s t" by (rule linepath_image_01) finally show ?thesis by (simp add: cis_conv_exp) qed lemma 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) lemma 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 lemma path_image_part_circlepath_subset': assumes "r \ 0" shows "path_image (part_circlepath z r s t) \ sphere z r" proof (cases "s \ t") case True thus ?thesis using path_image_part_circlepath_subset[of s t r z] assms by simp next case False thus ?thesis using path_image_part_circlepath_subset[of t s r z] assms by (subst reversepath_part_circlepath [symmetric], subst path_image_reversepath) simp_all qed lemma part_circlepath_cnj: "cnj (part_circlepath c r a b x) = part_circlepath (cnj c) r (-a) (-b) x" by (simp add: part_circlepath_def exp_cnj linepath_def algebra_simps) lemma contour_integral_bound_part_circlepath: assumes "f contour_integrable_on part_circlepath c r a b" assumes "B \ 0" "r \ 0" "\x. x \ path_image (part_circlepath c r a b) \ norm (f x) \ B" shows "norm (contour_integral (part_circlepath c r a b) f) \ B * r * \b - a\" proof - let ?I = "integral {0..1} (\x. f (part_circlepath c r a b x) * \ * of_real (r * (b - a)) * exp (\ * linepath a b x))" have "norm ?I \ integral {0..1} (\x::real. B * 1 * (r * \b - a\) * 1)" proof (rule integral_norm_bound_integral, goal_cases) case 1 with assms(1) show ?case by (simp add: contour_integrable_on vector_derivative_part_circlepath mult_ac) next case (3 x) with assms(2-) show ?case unfolding norm_mult norm_of_real abs_mult by (intro mult_mono) (auto simp: path_image_def) qed auto also have "?I = contour_integral (part_circlepath c r a b) f" by (simp add: contour_integral_integral vector_derivative_part_circlepath mult_ac) finally show ?thesis by simp qed lemma has_contour_integral_part_circlepath_iff: assumes "a < b" shows "(f has_contour_integral I) (part_circlepath c r a b) \ ((\t. f (c + r * cis t) * r * \ * cis t) has_integral I) {a..b}" proof - have "(f has_contour_integral I) (part_circlepath c r a b) \ ((\x. f (part_circlepath c r a b x) * vector_derivative (part_circlepath c r a b) (at x within {0..1})) has_integral I) {0..1}" unfolding has_contour_integral_def .. also have "\ \ ((\x. f (part_circlepath c r a b x) * r * (b - a) * \ * cis (linepath a b x)) has_integral I) {0..1}" by (intro has_integral_cong, subst vector_derivative_part_circlepath01) (simp_all add: cis_conv_exp) also have "\ \ ((\x. f (c + r * exp (\ * linepath (of_real a) (of_real b) x)) * r * \ * exp (\ * linepath (of_real a) (of_real b) x) * vector_derivative (linepath (of_real a) (of_real b)) (at x within {0..1})) has_integral I) {0..1}" by (intro has_integral_cong, subst vector_derivative_linepath_within) (auto simp: part_circlepath_def cis_conv_exp of_real_linepath [symmetric]) also have "\ \ ((\z. f (c + r * exp (\ * z)) * r * \ * exp (\ * z)) has_contour_integral I) (linepath (of_real a) (of_real b))" by (simp add: has_contour_integral_def) also have "\ \ ((\t. f (c + r * cis t) * r * \ * cis t) has_integral I) {a..b}" using assms by (subst has_contour_integral_linepath_Reals_iff) (simp_all add: cis_conv_exp) finally show ?thesis . qed lemma contour_integrable_part_circlepath_iff: assumes "a < b" shows "f contour_integrable_on (part_circlepath c r a b) \ (\t. f (c + r * cis t) * r * \ * cis t) integrable_on {a..b}" using assms by (auto simp: contour_integrable_on_def integrable_on_def has_contour_integral_part_circlepath_iff) lemma contour_integral_part_circlepath_eq: assumes "a < b" shows "contour_integral (part_circlepath c r a b) f = integral {a..b} (\t. f (c + r * cis t) * r * \ * cis t)" proof (cases "f contour_integrable_on part_circlepath c r a b") case True hence "(\t. f (c + r * cis t) * r * \ * cis t) integrable_on {a..b}" using assms by (simp add: contour_integrable_part_circlepath_iff) with True show ?thesis using has_contour_integral_part_circlepath_iff[OF assms] contour_integral_unique has_integral_integrable_integral by blast next case False hence "\(\t. f (c + r * cis t) * r * \ * cis t) integrable_on {a..b}" using assms by (simp add: contour_integrable_part_circlepath_iff) with False show ?thesis by (simp add: not_integrable_contour_integral not_integrable_integral) qed lemma contour_integral_part_circlepath_reverse: "contour_integral (part_circlepath c r a b) f = -contour_integral (part_circlepath c r b a) f" by (subst reversepath_part_circlepath [symmetric], subst contour_integral_reversepath) simp_all lemma contour_integral_part_circlepath_reverse': "b < a \ contour_integral (part_circlepath c r a b) f = -contour_integral (part_circlepath c r b a) f" by (rule contour_integral_part_circlepath_reverse) lemma 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 ((2 * real_of_int x * pi) * \) \ b + cmod (Ln w)}" proof (simp add: norm_mult finite_int_iff_bounded_le) show "\k. abs ` {x. 2 * \of_int x\ * pi \ b + cmod (Ln w)} \ {..k}" apply (rule_tac x="\(b + cmod (Ln w)) / (2*pi)\" in exI) apply (auto simp: field_split_simps le_floor_iff) done qed have [simp]: "\P f. {z. P z \ (\n. z = f n)} = f ` {n. P (f n)}" by blast have "finite {z. cmod z \ b \ exp z = exp (Ln w)}" using norm_add_leD by (fastforce intro: finite_subset [OF _ *] simp: exp_eq) then show ?thesis using False by auto 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 lemma 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 - let ?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})" using \s < t\ by (intro finite_vimageI [OF finite_bounded_log2]) (auto simp: inj_of_real) show ?thesis unfolding part_circlepath_def linepath_def vimage_def using le by (intro finite_subset [OF _ fin]) (auto simp: algebra_simps scaleR_conv_of_real exp_add exp_diff) 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}" by (rule has_integral_spike [OF negligible_finite [OF fin01]]) (use fi has_contour_integral in auto) 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 le * "2" \r > 0\ by (auto simp add: norm_mult vector_derivative_part_circlepath) qed qed lemma 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) lemma 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)" unfolding 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 lemma 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 unfolding part_circlepath_def simple_path_def by (rule disjE) (force intro: bexI [where x = "1/4"] bexI [where x = "1/3"])+ 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 **: "\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: "\s - t\ \ 2 * pi" if "\x. 0 \ x \ x \ 1 \ (\n. x * (t - s) = 2 * (real_of_int n * pi)) \ x = 0 \ x = 1" proof (rule ccontr) assume "\ \s - t\ \ 2 * pi" then have *: "\n. t - s \ of_int n * \s - t\" using False that [of "2*pi / \t - s\"] by (simp add: abs_minus_commute divide_simps) show False using * [of 1] * [of "-1"] by auto qed 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 have abs_away: "\P. (\x\{0..1}. \y\{0..1}. P \x - y\) \ (\x::real. 0 \ x \ x \ 1 \ P x)" by force show ?thesis using False apply (simp add: simple_path_def) 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 field_split_simps abs_mult dest: of_int_leD) done qed lemma 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 (rule ccontr) assume "x \ y" 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_div_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]) with \x \ y\ have st: "s-t = (of_int n * (pi * 2) / (y-x))" by (force simp: field_simps) have "\real_of_int n\ < \y - x\" using assms \x \ y\ by (simp add: st abs_mult field_simps) then show False using assms x y st by (auto dest: of_int_lessD) qed then have "inj_on (part_circlepath z r s t) {0..1}" using assms by (force simp add: part_circlepath_def inj_on_def exp_eq) then show ?thesis by (simp add: arc_def) qed subsection\Special case of one complete circle\ definition\<^marker>\tag important\ 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)" proof - have "\x\{0..1}. circlepath z r (y + 1/2) = circlepath z r x" if "0 \ y" "y \ 1" for y proof (cases "y \ 1/2") case False with that show ?thesis by (force simp: circlepath_add_half) qed (use that in force) then show ?thesis by (auto simp add: path_image_def image_def circlepath_minus) qed lemma path_image_circlepath_minus: "path_image (circlepath z (-r)) = path_image (circlepath z r)" using path_image_circlepath_minus_subset by fastforce lemma has_vector_derivative_circlepath [derivative_intros]: "((circlepath z r) has_vector_derivative (2 * pi * \ * r * exp (2 * of_real pi * \ * x))) (at x within X)" unfolding circlepath_def scaleR_conv_of_real by (rule derivative_eq_intros) (simp add: algebra_simps) lemma 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 lemma 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 obtain t where "0 \ t" "t < 2*pi" "Re(w/norm w) = cos t" "Im(w/norm w) = sin t" apply (rule sincos_total_2pi [of "Re(w/(norm w))" "Im(w/(norm w))"]) by (auto simp add: divide_simps \w \ 0\ cmod_power2 [symmetric]) then show ?thesis using False ** w_def \w \ 0\ by (rule_tac x="t / (2*pi)" in image_eqI) (auto simp add: field_simps) 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 lemma path_image_circlepath [simp]: "path_image (circlepath z r) = sphere z \r\" using path_image_circlepath_minus by (force simp: 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) lemma 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) lemma 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) lemma contour_integral_circlepath: assumes "r > 0" shows "contour_integral (circlepath z r) (\w. 1 / (w - z)) = 2 * complex_of_real pi * \" proof (rule contour_integral_unique) show "((\w. 1 / (w - z)) has_contour_integral 2 * complex_of_real pi * \) (circlepath z r)" unfolding has_contour_integral_def using assms has_integral_const_real [of _ 0 1] apply (subst has_integral_cong) apply (simp add: vector_derivative_circlepath01) apply (force simp: circlepath) done qed 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 ul_f: "uniform_limit (path_image \) f l 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 "0 < e / (\B\ + 1)" by simp then have "\\<^sub>F n in F. \x\path_image \. cmod (f n x - l x) < e / (\B\ + 1)" using ul_f [unfolded uniform_limit_iff dist_norm] by auto with ev_fint 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] by (fastforce simp: contour_integrable_on path_image_def) have Ble: "B * e / (\B\ + 1) \ e" using \0 \ B\ \0 < e\ by (simp add: field_split_simps) have "\h. (\x\{0..1}. cmod (l (\ x) * vector_derivative \ (at x) - h x) \ e) \ h integrable_on {0..1}" proof (intro exI conjI ballI) show "cmod (l (\ x) * vector_derivative \ (at x) - f a (\ x) * vector_derivative \ (at x)) \ e" if "x \ {0..1}" for x apply (rule order_trans [OF _ Ble]) using noleB [OF that] fga [OF that] \0 \ B\ \0 < e\ apply (fastforce simp: mult_ac dest: mult_mono [OF less_imp_le] simp add: norm_mult left_diff_distrib [symmetric] norm_minus_commute divide_simps) done qed (rule inta) } then show lintg: "l contour_integrable_on \" unfolding contour_integrable_on by (metis (mono_tags, lifting)integrable_uniform_limit_real) { 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 ul_f [unfolded uniform_limit_iff dist_norm, rule_format, 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 le_e: "\x. \\xa\{0..1}. 2 * cmod (f x (\ xa) - l (\ xa)) < e / B'; f x contour_integrable_on \\ \ cmod (integral {0..1} (\u. f x (\ u) * vector_derivative \ (at u) - l (\ u) * vector_derivative \ (at u))) < e" 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 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] le_e) done } then show "((\n. contour_integral \ (f n)) \ contour_integral \ l) F" by (rule tendstoI) qed corollary\<^marker>\tag unimportant\ contour_integral_uniform_limit_circlepath: assumes "\\<^sub>F n::'a in F. (f n) contour_integrable_on (circlepath z r)" and "uniform_limit (sphere z r) f l F" and "\ 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" using assms by (auto simp: vector_derivative_circlepath norm_mult intro!: contour_integral_uniform_limit) end \ No newline at end of file diff --git a/src/HOL/Deriv.thy b/src/HOL/Deriv.thy --- a/src/HOL/Deriv.thy +++ b/src/HOL/Deriv.thy @@ -1,2368 +1,2385 @@ (* Title: HOL/Deriv.thy Author: Jacques D. Fleuriot, University of Cambridge, 1998 Author: Brian Huffman Author: Lawrence C Paulson, 2004 Author: Benjamin Porter, 2005 *) section \Differentiation\ theory Deriv imports Limits begin subsection \Frechet derivative\ definition has_derivative :: "('a::real_normed_vector \ 'b::real_normed_vector) \ ('a \ 'b) \ 'a filter \ bool" (infix "(has'_derivative)" 50) where "(f has_derivative f') F \ bounded_linear f' \ ((\y. ((f y - f (Lim F (\x. x))) - f' (y - Lim F (\x. x))) /\<^sub>R norm (y - Lim F (\x. x))) \ 0) F" text \ Usually the filter \<^term>\F\ is \<^term>\at x within s\. \<^term>\(f has_derivative D) (at x within s)\ means: \<^term>\D\ is the derivative of function \<^term>\f\ at point \<^term>\x\ within the set \<^term>\s\. Where \<^term>\s\ is used to express left or right sided derivatives. In most cases \<^term>\s\ is either a variable or \<^term>\UNIV\. \ text \These are the only cases we'll care about, probably.\ lemma has_derivative_within: "(f has_derivative f') (at x within s) \ bounded_linear f' \ ((\y. (1 / norm(y - x)) *\<^sub>R (f y - (f x + f' (y - x)))) \ 0) (at x within s)" unfolding has_derivative_def tendsto_iff by (subst eventually_Lim_ident_at) (auto simp add: field_simps) lemma has_derivative_eq_rhs: "(f has_derivative f') F \ f' = g' \ (f has_derivative g') F" by simp definition has_field_derivative :: "('a::real_normed_field \ 'a) \ 'a \ 'a filter \ bool" (infix "(has'_field'_derivative)" 50) where "(f has_field_derivative D) F \ (f has_derivative (*) D) F" lemma DERIV_cong: "(f has_field_derivative X) F \ X = Y \ (f has_field_derivative Y) F" by simp definition has_vector_derivative :: "(real \ 'b::real_normed_vector) \ 'b \ real filter \ bool" (infix "has'_vector'_derivative" 50) where "(f has_vector_derivative f') net \ (f has_derivative (\x. x *\<^sub>R f')) net" lemma has_vector_derivative_eq_rhs: "(f has_vector_derivative X) F \ X = Y \ (f has_vector_derivative Y) F" by simp named_theorems derivative_intros "structural introduction rules for derivatives" setup \ let val eq_thms = @{thms has_derivative_eq_rhs DERIV_cong has_vector_derivative_eq_rhs} fun eq_rule thm = get_first (try (fn eq_thm => eq_thm OF [thm])) eq_thms in Global_Theory.add_thms_dynamic (\<^binding>\derivative_eq_intros\, fn context => Named_Theorems.get (Context.proof_of context) \<^named_theorems>\derivative_intros\ |> map_filter eq_rule) end \ text \ The following syntax is only used as a legacy syntax. \ abbreviation (input) FDERIV :: "('a::real_normed_vector \ 'b::real_normed_vector) \ 'a \ ('a \ 'b) \ bool" ("(FDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where "FDERIV f x :> f' \ (f has_derivative f') (at x)" lemma has_derivative_bounded_linear: "(f has_derivative f') F \ bounded_linear f'" by (simp add: has_derivative_def) lemma has_derivative_linear: "(f has_derivative f') F \ linear f'" using bounded_linear.linear[OF has_derivative_bounded_linear] . lemma has_derivative_ident[derivative_intros, simp]: "((\x. x) has_derivative (\x. x)) F" by (simp add: has_derivative_def) lemma has_derivative_id [derivative_intros, simp]: "(id has_derivative id) (at a)" by (metis eq_id_iff has_derivative_ident) lemma has_derivative_const[derivative_intros, simp]: "((\x. c) has_derivative (\x. 0)) F" by (simp add: has_derivative_def) lemma (in bounded_linear) bounded_linear: "bounded_linear f" .. lemma (in bounded_linear) has_derivative: "(g has_derivative g') F \ ((\x. f (g x)) has_derivative (\x. f (g' x))) F" unfolding has_derivative_def by (auto simp add: bounded_linear_compose [OF bounded_linear] scaleR diff dest: tendsto) lemmas has_derivative_scaleR_right [derivative_intros] = bounded_linear.has_derivative [OF bounded_linear_scaleR_right] lemmas has_derivative_scaleR_left [derivative_intros] = bounded_linear.has_derivative [OF bounded_linear_scaleR_left] lemmas has_derivative_mult_right [derivative_intros] = bounded_linear.has_derivative [OF bounded_linear_mult_right] lemmas has_derivative_mult_left [derivative_intros] = bounded_linear.has_derivative [OF bounded_linear_mult_left] lemmas has_derivative_of_real[derivative_intros, simp] = bounded_linear.has_derivative[OF bounded_linear_of_real] lemma has_derivative_add[simp, derivative_intros]: assumes f: "(f has_derivative f') F" and g: "(g has_derivative g') F" shows "((\x. f x + g x) has_derivative (\x. f' x + g' x)) F" unfolding has_derivative_def proof safe let ?x = "Lim F (\x. x)" let ?D = "\f f' y. ((f y - f ?x) - f' (y - ?x)) /\<^sub>R norm (y - ?x)" have "((\x. ?D f f' x + ?D g g' x) \ (0 + 0)) F" using f g by (intro tendsto_add) (auto simp: has_derivative_def) then show "(?D (\x. f x + g x) (\x. f' x + g' x) \ 0) F" by (simp add: field_simps scaleR_add_right scaleR_diff_right) qed (blast intro: bounded_linear_add f g has_derivative_bounded_linear) lemma has_derivative_sum[simp, derivative_intros]: "(\i. i \ I \ (f i has_derivative f' i) F) \ ((\x. \i\I. f i x) has_derivative (\x. \i\I. f' i x)) F" by (induct I rule: infinite_finite_induct) simp_all lemma has_derivative_minus[simp, derivative_intros]: "(f has_derivative f') F \ ((\x. - f x) has_derivative (\x. - f' x)) F" using has_derivative_scaleR_right[of f f' F "-1"] by simp lemma has_derivative_diff[simp, derivative_intros]: "(f has_derivative f') F \ (g has_derivative g') F \ ((\x. f x - g x) has_derivative (\x. f' x - g' x)) F" by (simp only: diff_conv_add_uminus has_derivative_add has_derivative_minus) lemma has_derivative_at_within: "(f has_derivative f') (at x within s) \ (bounded_linear f' \ ((\y. ((f y - f x) - f' (y - x)) /\<^sub>R norm (y - x)) \ 0) (at x within s))" proof (cases "at x within s = bot") case True then show ?thesis by (metis (no_types, lifting) has_derivative_within tendsto_bot) next case False then show ?thesis by (simp add: Lim_ident_at has_derivative_def) qed lemma has_derivative_iff_norm: "(f has_derivative f') (at x within s) \ bounded_linear f' \ ((\y. norm ((f y - f x) - f' (y - x)) / norm (y - x)) \ 0) (at x within s)" using tendsto_norm_zero_iff[of _ "at x within s", where 'b="'b", symmetric] by (simp add: has_derivative_at_within divide_inverse ac_simps) lemma has_derivative_at: "(f has_derivative D) (at x) \ (bounded_linear D \ (\h. norm (f (x + h) - f x - D h) / norm h) \0\ 0)" by (simp add: has_derivative_iff_norm LIM_offset_zero_iff) lemma field_has_derivative_at: fixes x :: "'a::real_normed_field" shows "(f has_derivative (*) D) (at x) \ (\h. (f (x + h) - f x) / h) \0\ D" (is "?lhs = ?rhs") proof - have "?lhs = (\h. norm (f (x + h) - f x - D * h) / norm h) \0 \ 0" by (simp add: bounded_linear_mult_right has_derivative_at) also have "... = (\y. norm ((f (x + y) - f x - D * y) / y)) \0\ 0" by (simp cong: LIM_cong flip: nonzero_norm_divide) also have "... = (\y. norm ((f (x + y) - f x) / y - D / y * y)) \0\ 0" by (simp only: diff_divide_distrib times_divide_eq_left [symmetric]) also have "... = ?rhs" by (simp add: tendsto_norm_zero_iff LIM_zero_iff cong: LIM_cong) finally show ?thesis . qed lemma has_derivative_iff_Ex: "(f has_derivative f') (at x) \ bounded_linear f' \ (\e. (\h. f (x+h) = f x + f' h + e h) \ ((\h. norm (e h) / norm h) \ 0) (at 0))" unfolding has_derivative_at by force lemma has_derivative_at_within_iff_Ex: assumes "x \ S" "open S" shows "(f has_derivative f') (at x within S) \ bounded_linear f' \ (\e. (\h. x+h \ S \ f (x+h) = f x + f' h + e h) \ ((\h. norm (e h) / norm h) \ 0) (at 0))" (is "?lhs = ?rhs") proof safe show "bounded_linear f'" if "(f has_derivative f') (at x within S)" using has_derivative_bounded_linear that by blast show "\e. (\h. x + h \ S \ f (x + h) = f x + f' h + e h) \ (\h. norm (e h) / norm h) \0\ 0" if "(f has_derivative f') (at x within S)" by (metis (full_types) assms that has_derivative_iff_Ex at_within_open) show "(f has_derivative f') (at x within S)" if "bounded_linear f'" and eq [rule_format]: "\h. x + h \ S \ f (x + h) = f x + f' h + e h" and 0: "(\h. norm (e (h::'a)::'b) / norm h) \0\ 0" for e proof - have 1: "f y - f x = f' (y-x) + e (y-x)" if "y \ S" for y using eq [of "y-x"] that by simp have 2: "((\y. norm (e (y-x)) / norm (y - x)) \ 0) (at x within S)" by (simp add: "0" assms tendsto_offset_zero_iff) have "((\y. norm (f y - f x - f' (y - x)) / norm (y - x)) \ 0) (at x within S)" by (simp add: Lim_cong_within 1 2) then show ?thesis by (simp add: has_derivative_iff_norm \bounded_linear f'\) qed qed lemma has_derivativeI: "bounded_linear f' \ ((\y. ((f y - f x) - f' (y - x)) /\<^sub>R norm (y - x)) \ 0) (at x within s) \ (f has_derivative f') (at x within s)" by (simp add: has_derivative_at_within) lemma has_derivativeI_sandwich: assumes e: "0 < e" and bounded: "bounded_linear f'" and sandwich: "(\y. y \ s \ y \ x \ dist y x < e \ norm ((f y - f x) - f' (y - x)) / norm (y - x) \ H y)" and "(H \ 0) (at x within s)" shows "(f has_derivative f') (at x within s)" unfolding has_derivative_iff_norm proof safe show "((\y. norm (f y - f x - f' (y - x)) / norm (y - x)) \ 0) (at x within s)" proof (rule tendsto_sandwich[where f="\x. 0"]) show "(H \ 0) (at x within s)" by fact show "eventually (\n. norm (f n - f x - f' (n - x)) / norm (n - x) \ H n) (at x within s)" unfolding eventually_at using e sandwich by auto qed (auto simp: le_divide_eq) qed fact lemma has_derivative_subset: "(f has_derivative f') (at x within s) \ t \ s \ (f has_derivative f') (at x within t)" by (auto simp add: has_derivative_iff_norm intro: tendsto_within_subset) lemma has_derivative_within_singleton_iff: "(f has_derivative g) (at x within {x}) \ bounded_linear g" by (auto intro!: has_derivativeI_sandwich[where e=1] has_derivative_bounded_linear) subsubsection \Limit transformation for derivatives\ lemma has_derivative_transform_within: assumes "(f has_derivative f') (at x within s)" and "0 < d" and "x \ s" and "\x'. \x' \ s; dist x' x < d\ \ f x' = g x'" shows "(g has_derivative f') (at x within s)" using assms unfolding has_derivative_within by (force simp add: intro: Lim_transform_within) lemma has_derivative_transform_within_open: assumes "(f has_derivative f') (at x within t)" and "open s" and "x \ s" and "\x. x\s \ f x = g x" shows "(g has_derivative f') (at x within t)" using assms unfolding has_derivative_within by (force simp add: intro: Lim_transform_within_open) lemma has_derivative_transform: assumes "x \ s" "\x. x \ s \ g x = f x" assumes "(f has_derivative f') (at x within s)" shows "(g has_derivative f') (at x within s)" using assms by (intro has_derivative_transform_within[OF _ zero_less_one, where g=g]) auto lemma has_derivative_transform_eventually: assumes "(f has_derivative f') (at x within s)" "(\\<^sub>F x' in at x within s. f x' = g x')" assumes "f x = g x" "x \ s" shows "(g has_derivative f') (at x within s)" using assms proof - from assms(2,3) obtain d where "d > 0" "\x'. x' \ s \ dist x' x < d \ f x' = g x'" by (force simp: eventually_at) from has_derivative_transform_within[OF assms(1) this(1) assms(4) this(2)] show ?thesis . qed lemma has_field_derivative_transform_within: assumes "(f has_field_derivative f') (at a within S)" and "0 < d" and "a \ S" and "\x. \x \ S; dist x a < d\ \ f x = g x" shows "(g has_field_derivative f') (at a within S)" using assms unfolding has_field_derivative_def by (metis has_derivative_transform_within) lemma has_field_derivative_transform_within_open: assumes "(f has_field_derivative f') (at a)" and "open S" "a \ S" and "\x. x \ S \ f x = g x" shows "(g has_field_derivative f') (at a)" using assms unfolding has_field_derivative_def by (metis has_derivative_transform_within_open) subsection \Continuity\ lemma has_derivative_continuous: assumes f: "(f has_derivative f') (at x within s)" shows "continuous (at x within s) f" proof - from f interpret F: bounded_linear f' by (rule has_derivative_bounded_linear) note F.tendsto[tendsto_intros] let ?L = "\f. (f \ 0) (at x within s)" have "?L (\y. norm ((f y - f x) - f' (y - x)) / norm (y - x))" using f unfolding has_derivative_iff_norm by blast then have "?L (\y. norm ((f y - f x) - f' (y - x)) / norm (y - x) * norm (y - x))" (is ?m) by (rule tendsto_mult_zero) (auto intro!: tendsto_eq_intros) also have "?m \ ?L (\y. norm ((f y - f x) - f' (y - x)))" by (intro filterlim_cong) (simp_all add: eventually_at_filter) finally have "?L (\y. (f y - f x) - f' (y - x))" by (rule tendsto_norm_zero_cancel) then have "?L (\y. ((f y - f x) - f' (y - x)) + f' (y - x))" by (rule tendsto_eq_intros) (auto intro!: tendsto_eq_intros simp: F.zero) then have "?L (\y. f y - f x)" by simp from tendsto_add[OF this tendsto_const, of "f x"] show ?thesis by (simp add: continuous_within) qed subsection \Composition\ lemma tendsto_at_iff_tendsto_nhds_within: "f x = y \ (f \ y) (at x within s) \ (f \ y) (inf (nhds x) (principal s))" unfolding tendsto_def eventually_inf_principal eventually_at_filter by (intro ext all_cong imp_cong) (auto elim!: eventually_mono) lemma has_derivative_in_compose: assumes f: "(f has_derivative f') (at x within s)" and g: "(g has_derivative g') (at (f x) within (f`s))" shows "((\x. g (f x)) has_derivative (\x. g' (f' x))) (at x within s)" proof - from f interpret F: bounded_linear f' by (rule has_derivative_bounded_linear) from g interpret G: bounded_linear g' by (rule has_derivative_bounded_linear) from F.bounded obtain kF where kF: "\x. norm (f' x) \ norm x * kF" by fast from G.bounded obtain kG where kG: "\x. norm (g' x) \ norm x * kG" by fast note G.tendsto[tendsto_intros] let ?L = "\f. (f \ 0) (at x within s)" let ?D = "\f f' x y. (f y - f x) - f' (y - x)" let ?N = "\f f' x y. norm (?D f f' x y) / norm (y - x)" let ?gf = "\x. g (f x)" and ?gf' = "\x. g' (f' x)" define Nf where "Nf = ?N f f' x" define Ng where [abs_def]: "Ng y = ?N g g' (f x) (f y)" for y show ?thesis proof (rule has_derivativeI_sandwich[of 1]) show "bounded_linear (\x. g' (f' x))" using f g by (blast intro: bounded_linear_compose has_derivative_bounded_linear) next fix y :: 'a assume neq: "y \ x" have "?N ?gf ?gf' x y = norm (g' (?D f f' x y) + ?D g g' (f x) (f y)) / norm (y - x)" by (simp add: G.diff G.add field_simps) also have "\ \ norm (g' (?D f f' x y)) / norm (y - x) + Ng y * (norm (f y - f x) / norm (y - x))" by (simp add: add_divide_distrib[symmetric] divide_right_mono norm_triangle_ineq G.zero Ng_def) also have "\ \ Nf y * kG + Ng y * (Nf y + kF)" proof (intro add_mono mult_left_mono) have "norm (f y - f x) = norm (?D f f' x y + f' (y - x))" by simp also have "\ \ norm (?D f f' x y) + norm (f' (y - x))" by (rule norm_triangle_ineq) also have "\ \ norm (?D f f' x y) + norm (y - x) * kF" using kF by (intro add_mono) simp finally show "norm (f y - f x) / norm (y - x) \ Nf y + kF" by (simp add: neq Nf_def field_simps) qed (use kG in \simp_all add: Ng_def Nf_def neq zero_le_divide_iff field_simps\) finally show "?N ?gf ?gf' x y \ Nf y * kG + Ng y * (Nf y + kF)" . next have [tendsto_intros]: "?L Nf" using f unfolding has_derivative_iff_norm Nf_def .. from f have "(f \ f x) (at x within s)" by (blast intro: has_derivative_continuous continuous_within[THEN iffD1]) then have f': "LIM x at x within s. f x :> inf (nhds (f x)) (principal (f`s))" unfolding filterlim_def by (simp add: eventually_filtermap eventually_at_filter le_principal) have "((?N g g' (f x)) \ 0) (at (f x) within f`s)" using g unfolding has_derivative_iff_norm .. then have g': "((?N g g' (f x)) \ 0) (inf (nhds (f x)) (principal (f`s)))" by (rule tendsto_at_iff_tendsto_nhds_within[THEN iffD1, rotated]) simp have [tendsto_intros]: "?L Ng" unfolding Ng_def by (rule filterlim_compose[OF g' f']) show "((\y. Nf y * kG + Ng y * (Nf y + kF)) \ 0) (at x within s)" by (intro tendsto_eq_intros) auto qed simp qed lemma has_derivative_compose: "(f has_derivative f') (at x within s) \ (g has_derivative g') (at (f x)) \ ((\x. g (f x)) has_derivative (\x. g' (f' x))) (at x within s)" by (blast intro: has_derivative_in_compose has_derivative_subset) lemma has_derivative_in_compose2: assumes "\x. x \ t \ (g has_derivative g' x) (at x within t)" assumes "f ` s \ t" "x \ s" assumes "(f has_derivative f') (at x within s)" shows "((\x. g (f x)) has_derivative (\y. g' (f x) (f' y))) (at x within s)" using assms by (auto intro: has_derivative_subset intro!: has_derivative_in_compose[of f f' x s g]) lemma (in bounded_bilinear) FDERIV: assumes f: "(f has_derivative f') (at x within s)" and g: "(g has_derivative g') (at x within s)" shows "((\x. f x ** g x) has_derivative (\h. f x ** g' h + f' h ** g x)) (at x within s)" proof - from bounded_linear.bounded [OF has_derivative_bounded_linear [OF f]] obtain KF where norm_F: "\x. norm (f' x) \ norm x * KF" by fast from pos_bounded obtain K where K: "0 < K" and norm_prod: "\a b. norm (a ** b) \ norm a * norm b * K" by fast let ?D = "\f f' y. f y - f x - f' (y - x)" let ?N = "\f f' y. norm (?D f f' y) / norm (y - x)" define Ng where "Ng = ?N g g'" define Nf where "Nf = ?N f f'" let ?fun1 = "\y. norm (f y ** g y - f x ** g x - (f x ** g' (y - x) + f' (y - x) ** g x)) / norm (y - x)" let ?fun2 = "\y. norm (f x) * Ng y * K + Nf y * norm (g y) * K + KF * norm (g y - g x) * K" let ?F = "at x within s" show ?thesis proof (rule has_derivativeI_sandwich[of 1]) show "bounded_linear (\h. f x ** g' h + f' h ** g x)" by (intro bounded_linear_add bounded_linear_compose [OF bounded_linear_right] bounded_linear_compose [OF bounded_linear_left] has_derivative_bounded_linear [OF g] has_derivative_bounded_linear [OF f]) next from g have "(g \ g x) ?F" by (intro continuous_within[THEN iffD1] has_derivative_continuous) moreover from f g have "(Nf \ 0) ?F" "(Ng \ 0) ?F" by (simp_all add: has_derivative_iff_norm Ng_def Nf_def) ultimately have "(?fun2 \ norm (f x) * 0 * K + 0 * norm (g x) * K + KF * norm (0::'b) * K) ?F" by (intro tendsto_intros) (simp_all add: LIM_zero_iff) then show "(?fun2 \ 0) ?F" by simp next fix y :: 'd assume "y \ x" have "?fun1 y = norm (f x ** ?D g g' y + ?D f f' y ** g y + f' (y - x) ** (g y - g x)) / norm (y - x)" by (simp add: diff_left diff_right add_left add_right field_simps) also have "\ \ (norm (f x) * norm (?D g g' y) * K + norm (?D f f' y) * norm (g y) * K + norm (y - x) * KF * norm (g y - g x) * K) / norm (y - x)" by (intro divide_right_mono mult_mono' order_trans [OF norm_triangle_ineq add_mono] order_trans [OF norm_prod mult_right_mono] mult_nonneg_nonneg order_refl norm_ge_zero norm_F K [THEN order_less_imp_le]) also have "\ = ?fun2 y" by (simp add: add_divide_distrib Ng_def Nf_def) finally show "?fun1 y \ ?fun2 y" . qed simp qed lemmas has_derivative_mult[simp, derivative_intros] = bounded_bilinear.FDERIV[OF bounded_bilinear_mult] lemmas has_derivative_scaleR[simp, derivative_intros] = bounded_bilinear.FDERIV[OF bounded_bilinear_scaleR] lemma has_derivative_prod[simp, derivative_intros]: fixes f :: "'i \ 'a::real_normed_vector \ 'b::real_normed_field" shows "(\i. i \ I \ (f i has_derivative f' i) (at x within S)) \ ((\x. \i\I. f i x) has_derivative (\y. \i\I. f' i y * (\j\I - {i}. f j x))) (at x within S)" proof (induct I rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case (insert i I) let ?P = "\y. f i x * (\i\I. f' i y * (\j\I - {i}. f j x)) + (f' i y) * (\i\I. f i x)" have "((\x. f i x * (\i\I. f i x)) has_derivative ?P) (at x within S)" using insert by (intro has_derivative_mult) auto also have "?P = (\y. \i'\insert i I. f' i' y * (\j\insert i I - {i'}. f j x))" using insert(1,2) by (auto simp add: sum_distrib_left insert_Diff_if intro!: ext sum.cong) finally show ?case using insert by simp qed lemma has_derivative_power[simp, derivative_intros]: fixes f :: "'a :: real_normed_vector \ 'b :: real_normed_field" assumes f: "(f has_derivative f') (at x within S)" shows "((\x. f x^n) has_derivative (\y. of_nat n * f' y * f x^(n - 1))) (at x within S)" using has_derivative_prod[OF f, of "{..< n}"] by (simp add: prod_constant ac_simps) lemma has_derivative_inverse': fixes x :: "'a::real_normed_div_algebra" assumes x: "x \ 0" shows "(inverse has_derivative (\h. - (inverse x * h * inverse x))) (at x within S)" (is "(_ has_derivative ?f) _") proof (rule has_derivativeI_sandwich) show "bounded_linear (\h. - (inverse x * h * inverse x))" by (simp add: bounded_linear_minus bounded_linear_mult_const bounded_linear_mult_right) show "0 < norm x" using x by simp have "(inverse \ inverse x) (at x within S)" using tendsto_inverse tendsto_ident_at x by auto then show "((\y. norm (inverse y - inverse x) * norm (inverse x)) \ 0) (at x within S)" by (simp add: LIM_zero_iff tendsto_mult_left_zero tendsto_norm_zero) next fix y :: 'a assume h: "y \ x" "dist y x < norm x" then have "y \ 0" by auto have "norm (inverse y - inverse x - ?f (y -x)) / norm (y - x) = norm (- (inverse y * (y - x) * inverse x - inverse x * (y - x) * inverse x)) / norm (y - x)" by (simp add: \y \ 0\ inverse_diff_inverse x) also have "... = norm ((inverse y - inverse x) * (y - x) * inverse x) / norm (y - x)" by (simp add: left_diff_distrib norm_minus_commute) also have "\ \ norm (inverse y - inverse x) * norm (y - x) * norm (inverse x) / norm (y - x)" by (simp add: norm_mult) also have "\ = norm (inverse y - inverse x) * norm (inverse x)" by simp finally show "norm (inverse y - inverse x - ?f (y -x)) / norm (y - x) \ norm (inverse y - inverse x) * norm (inverse x)" . qed lemma has_derivative_inverse[simp, derivative_intros]: fixes f :: "_ \ 'a::real_normed_div_algebra" assumes x: "f x \ 0" and f: "(f has_derivative f') (at x within S)" shows "((\x. inverse (f x)) has_derivative (\h. - (inverse (f x) * f' h * inverse (f x)))) (at x within S)" using has_derivative_compose[OF f has_derivative_inverse', OF x] . lemma has_derivative_divide[simp, derivative_intros]: fixes f :: "_ \ 'a::real_normed_div_algebra" assumes f: "(f has_derivative f') (at x within S)" and g: "(g has_derivative g') (at x within S)" assumes x: "g x \ 0" shows "((\x. f x / g x) has_derivative (\h. - f x * (inverse (g x) * g' h * inverse (g x)) + f' h / g x)) (at x within S)" using has_derivative_mult[OF f has_derivative_inverse[OF x g]] by (simp add: field_simps) lemma has_derivative_power_int': fixes x :: "'a::real_normed_field" assumes x: "x \ 0" shows "((\x. power_int x n) has_derivative (\y. y * (of_int n * power_int x (n - 1)))) (at x within S)" proof (cases n rule: int_cases4) case (nonneg n) thus ?thesis using x by (cases "n = 0") (auto intro!: derivative_eq_intros simp: field_simps power_int_diff fun_eq_iff simp flip: power_Suc) next case (neg n) thus ?thesis using x by (auto intro!: derivative_eq_intros simp: field_simps power_int_diff power_int_minus simp flip: power_Suc power_Suc2 power_add) qed lemma has_derivative_power_int[simp, derivative_intros]: fixes f :: "_ \ 'a::real_normed_field" assumes x: "f x \ 0" and f: "(f has_derivative f') (at x within S)" shows "((\x. power_int (f x) n) has_derivative (\h. f' h * (of_int n * power_int (f x) (n - 1)))) (at x within S)" using has_derivative_compose[OF f has_derivative_power_int', OF x] . text \Conventional form requires mult-AC laws. Types real and complex only.\ lemma has_derivative_divide'[derivative_intros]: fixes f :: "_ \ 'a::real_normed_field" assumes f: "(f has_derivative f') (at x within S)" and g: "(g has_derivative g') (at x within S)" and x: "g x \ 0" shows "((\x. f x / g x) has_derivative (\h. (f' h * g x - f x * g' h) / (g x * g x))) (at x within S)" proof - have "f' h / g x - f x * (inverse (g x) * g' h * inverse (g x)) = (f' h * g x - f x * g' h) / (g x * g x)" for h by (simp add: field_simps x) then show ?thesis using has_derivative_divide [OF f g] x by simp qed subsection \Uniqueness\ text \ This can not generally shown for \<^const>\has_derivative\, as we need to approach the point from all directions. There is a proof in \Analysis\ for \euclidean_space\. \ lemma has_derivative_at2: "(f has_derivative f') (at x) \ bounded_linear f' \ ((\y. (1 / (norm(y - x))) *\<^sub>R (f y - (f x + f' (y - x)))) \ 0) (at x)" using has_derivative_within [of f f' x UNIV] by simp lemma has_derivative_zero_unique: assumes "((\x. 0) has_derivative F) (at x)" shows "F = (\h. 0)" proof - interpret F: bounded_linear F using assms by (rule has_derivative_bounded_linear) let ?r = "\h. norm (F h) / norm h" have *: "?r \0\ 0" using assms unfolding has_derivative_at by simp show "F = (\h. 0)" proof show "F h = 0" for h proof (rule ccontr) assume **: "\ ?thesis" then have h: "h \ 0" by (auto simp add: F.zero) with ** have "0 < ?r h" by simp from LIM_D [OF * this] obtain S where S: "0 < S" and r: "\x. x \ 0 \ norm x < S \ ?r x < ?r h" by auto from dense [OF S] obtain t where t: "0 < t \ t < S" .. let ?x = "scaleR (t / norm h) h" have "?x \ 0" and "norm ?x < S" using t h by simp_all then have "?r ?x < ?r h" by (rule r) then show False using t h by (simp add: F.scaleR) qed qed qed lemma has_derivative_unique: assumes "(f has_derivative F) (at x)" and "(f has_derivative F') (at x)" shows "F = F'" proof - have "((\x. 0) has_derivative (\h. F h - F' h)) (at x)" using has_derivative_diff [OF assms] by simp then have "(\h. F h - F' h) = (\h. 0)" by (rule has_derivative_zero_unique) then show "F = F'" unfolding fun_eq_iff right_minus_eq . qed lemma has_derivative_Uniq: "\\<^sub>\\<^sub>1F. (f has_derivative F) (at x)" by (simp add: Uniq_def has_derivative_unique) subsection \Differentiability predicate\ definition differentiable :: "('a::real_normed_vector \ 'b::real_normed_vector) \ 'a filter \ bool" (infix "differentiable" 50) where "f differentiable F \ (\D. (f has_derivative D) F)" lemma differentiable_subset: "f differentiable (at x within s) \ t \ s \ f differentiable (at x within t)" unfolding differentiable_def by (blast intro: has_derivative_subset) lemmas differentiable_within_subset = differentiable_subset lemma differentiable_ident [simp, derivative_intros]: "(\x. x) differentiable F" unfolding differentiable_def by (blast intro: has_derivative_ident) lemma differentiable_const [simp, derivative_intros]: "(\z. a) differentiable F" unfolding differentiable_def by (blast intro: has_derivative_const) lemma differentiable_in_compose: "f differentiable (at (g x) within (g`s)) \ g differentiable (at x within s) \ (\x. f (g x)) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_in_compose) lemma differentiable_compose: "f differentiable (at (g x)) \ g differentiable (at x within s) \ (\x. f (g x)) differentiable (at x within s)" by (blast intro: differentiable_in_compose differentiable_subset) lemma differentiable_add [simp, derivative_intros]: "f differentiable F \ g differentiable F \ (\x. f x + g x) differentiable F" unfolding differentiable_def by (blast intro: has_derivative_add) lemma differentiable_sum[simp, derivative_intros]: assumes "finite s" "\a\s. (f a) differentiable net" shows "(\x. sum (\a. f a x) s) differentiable net" proof - from bchoice[OF assms(2)[unfolded differentiable_def]] show ?thesis by (auto intro!: has_derivative_sum simp: differentiable_def) qed lemma differentiable_minus [simp, derivative_intros]: "f differentiable F \ (\x. - f x) differentiable F" unfolding differentiable_def by (blast intro: has_derivative_minus) lemma differentiable_diff [simp, derivative_intros]: "f differentiable F \ g differentiable F \ (\x. f x - g x) differentiable F" unfolding differentiable_def by (blast intro: has_derivative_diff) lemma differentiable_mult [simp, derivative_intros]: fixes f g :: "'a::real_normed_vector \ 'b::real_normed_algebra" shows "f differentiable (at x within s) \ g differentiable (at x within s) \ (\x. f x * g x) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_mult) +lemma differentiable_cmult_left_iff [simp]: + fixes c::"'a::real_normed_field" + shows "(\t. c * q t) differentiable at t \ c = 0 \ (\t. q t) differentiable at t" (is "?lhs = ?rhs") +proof + assume L: ?lhs + {assume "c \ 0" + then have "q differentiable at t" + using differentiable_mult [OF differentiable_const L, of concl: "1/c"] by auto + } then show ?rhs + by auto +qed auto + +lemma differentiable_cmult_right_iff [simp]: + fixes c::"'a::real_normed_field" + shows "(\t. q t * c) differentiable at t \ c = 0 \ (\t. q t) differentiable at t" (is "?lhs = ?rhs") + by (simp add: mult.commute flip: differentiable_cmult_left_iff) + lemma differentiable_inverse [simp, derivative_intros]: fixes f :: "'a::real_normed_vector \ 'b::real_normed_field" shows "f differentiable (at x within s) \ f x \ 0 \ (\x. inverse (f x)) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_inverse) lemma differentiable_divide [simp, derivative_intros]: fixes f g :: "'a::real_normed_vector \ 'b::real_normed_field" shows "f differentiable (at x within s) \ g differentiable (at x within s) \ g x \ 0 \ (\x. f x / g x) differentiable (at x within s)" unfolding divide_inverse by simp lemma differentiable_power [simp, derivative_intros]: fixes f g :: "'a::real_normed_vector \ 'b::real_normed_field" shows "f differentiable (at x within s) \ (\x. f x ^ n) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_power) lemma differentiable_power_int [simp, derivative_intros]: fixes f :: "'a::real_normed_vector \ 'b::real_normed_field" shows "f differentiable (at x within s) \ f x \ 0 \ (\x. power_int (f x) n) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_power_int) lemma differentiable_scaleR [simp, derivative_intros]: "f differentiable (at x within s) \ g differentiable (at x within s) \ (\x. f x *\<^sub>R g x) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_scaleR) lemma has_derivative_imp_has_field_derivative: "(f has_derivative D) F \ (\x. x * D' = D x) \ (f has_field_derivative D') F" unfolding has_field_derivative_def by (rule has_derivative_eq_rhs[of f D]) (simp_all add: fun_eq_iff mult.commute) lemma has_field_derivative_imp_has_derivative: "(f has_field_derivative D) F \ (f has_derivative (*) D) F" by (simp add: has_field_derivative_def) lemma DERIV_subset: "(f has_field_derivative f') (at x within s) \ t \ s \ (f has_field_derivative f') (at x within t)" by (simp add: has_field_derivative_def has_derivative_subset) lemma has_field_derivative_at_within: "(f has_field_derivative f') (at x) \ (f has_field_derivative f') (at x within s)" using DERIV_subset by blast abbreviation (input) DERIV :: "('a::real_normed_field \ 'a) \ 'a \ 'a \ bool" ("(DERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where "DERIV f x :> D \ (f has_field_derivative D) (at x)" abbreviation has_real_derivative :: "(real \ real) \ real \ real filter \ bool" (infix "(has'_real'_derivative)" 50) where "(f has_real_derivative D) F \ (f has_field_derivative D) F" lemma real_differentiable_def: "f differentiable at x within s \ (\D. (f has_real_derivative D) (at x within s))" proof safe assume "f differentiable at x within s" then obtain f' where *: "(f has_derivative f') (at x within s)" unfolding differentiable_def by auto then obtain c where "f' = ((*) c)" by (metis real_bounded_linear has_derivative_bounded_linear mult.commute fun_eq_iff) with * show "\D. (f has_real_derivative D) (at x within s)" unfolding has_field_derivative_def by auto qed (auto simp: differentiable_def has_field_derivative_def) lemma real_differentiableE [elim?]: assumes f: "f differentiable (at x within s)" obtains df where "(f has_real_derivative df) (at x within s)" using assms by (auto simp: real_differentiable_def) lemma has_field_derivative_iff: "(f has_field_derivative D) (at x within S) \ ((\y. (f y - f x) / (y - x)) \ D) (at x within S)" proof - have "((\y. norm (f y - f x - D * (y - x)) / norm (y - x)) \ 0) (at x within S) = ((\y. (f y - f x) / (y - x) - D) \ 0) (at x within S)" apply (subst tendsto_norm_zero_iff[symmetric], rule filterlim_cong) apply (simp_all add: eventually_at_filter field_simps nonzero_norm_divide) done then show ?thesis by (simp add: has_field_derivative_def has_derivative_iff_norm bounded_linear_mult_right LIM_zero_iff) qed lemma DERIV_def: "DERIV f x :> D \ (\h. (f (x + h) - f x) / h) \0\ D" unfolding field_has_derivative_at has_field_derivative_def has_field_derivative_iff .. lemma mult_commute_abs: "(\x. x * c) = (*) c" for c :: "'a::ab_semigroup_mult" by (simp add: fun_eq_iff mult.commute) lemma DERIV_compose_FDERIV: fixes f::"real\real" assumes "DERIV f (g x) :> f'" assumes "(g has_derivative g') (at x within s)" shows "((\x. f (g x)) has_derivative (\x. g' x * f')) (at x within s)" using assms has_derivative_compose[of g g' x s f "(*) f'"] by (auto simp: has_field_derivative_def ac_simps) subsection \Vector derivative\ lemma has_field_derivative_iff_has_vector_derivative: "(f has_field_derivative y) F \ (f has_vector_derivative y) F" unfolding has_vector_derivative_def has_field_derivative_def real_scaleR_def mult_commute_abs .. lemma has_field_derivative_subset: "(f has_field_derivative y) (at x within s) \ t \ s \ (f has_field_derivative y) (at x within t)" unfolding has_field_derivative_def by (rule has_derivative_subset) lemma has_vector_derivative_const[simp, derivative_intros]: "((\x. c) has_vector_derivative 0) net" by (auto simp: has_vector_derivative_def) lemma has_vector_derivative_id[simp, derivative_intros]: "((\x. x) has_vector_derivative 1) net" by (auto simp: has_vector_derivative_def) lemma has_vector_derivative_minus[derivative_intros]: "(f has_vector_derivative f') net \ ((\x. - f x) has_vector_derivative (- f')) net" by (auto simp: has_vector_derivative_def) lemma has_vector_derivative_add[derivative_intros]: "(f has_vector_derivative f') net \ (g has_vector_derivative g') net \ ((\x. f x + g x) has_vector_derivative (f' + g')) net" by (auto simp: has_vector_derivative_def scaleR_right_distrib) lemma has_vector_derivative_sum[derivative_intros]: "(\i. i \ I \ (f i has_vector_derivative f' i) net) \ ((\x. \i\I. f i x) has_vector_derivative (\i\I. f' i)) net" by (auto simp: has_vector_derivative_def fun_eq_iff scaleR_sum_right intro!: derivative_eq_intros) lemma has_vector_derivative_diff[derivative_intros]: "(f has_vector_derivative f') net \ (g has_vector_derivative g') net \ ((\x. f x - g x) has_vector_derivative (f' - g')) net" by (auto simp: has_vector_derivative_def scaleR_diff_right) lemma has_vector_derivative_add_const: "((\t. g t + z) has_vector_derivative f') net = ((\t. g t) has_vector_derivative f') net" apply (intro iffI) apply (force dest: has_vector_derivative_diff [where g = "\t. z", OF _ has_vector_derivative_const]) apply (force dest: has_vector_derivative_add [OF _ has_vector_derivative_const]) done lemma has_vector_derivative_diff_const: "((\t. g t - z) has_vector_derivative f') net = ((\t. g t) has_vector_derivative f') net" using has_vector_derivative_add_const [where z = "-z"] by simp lemma (in bounded_linear) has_vector_derivative: assumes "(g has_vector_derivative g') F" shows "((\x. f (g x)) has_vector_derivative f g') F" using has_derivative[OF assms[unfolded has_vector_derivative_def]] by (simp add: has_vector_derivative_def scaleR) lemma (in bounded_bilinear) has_vector_derivative: assumes "(f has_vector_derivative f') (at x within s)" and "(g has_vector_derivative g') (at x within s)" shows "((\x. f x ** g x) has_vector_derivative (f x ** g' + f' ** g x)) (at x within s)" using FDERIV[OF assms(1-2)[unfolded has_vector_derivative_def]] by (simp add: has_vector_derivative_def scaleR_right scaleR_left scaleR_right_distrib) lemma has_vector_derivative_scaleR[derivative_intros]: "(f has_field_derivative f') (at x within s) \ (g has_vector_derivative g') (at x within s) \ ((\x. f x *\<^sub>R g x) has_vector_derivative (f x *\<^sub>R g' + f' *\<^sub>R g x)) (at x within s)" unfolding has_field_derivative_iff_has_vector_derivative by (rule bounded_bilinear.has_vector_derivative[OF bounded_bilinear_scaleR]) lemma has_vector_derivative_mult[derivative_intros]: "(f has_vector_derivative f') (at x within s) \ (g has_vector_derivative g') (at x within s) \ ((\x. f x * g x) has_vector_derivative (f x * g' + f' * g x)) (at x within s)" for f g :: "real \ 'a::real_normed_algebra" by (rule bounded_bilinear.has_vector_derivative[OF bounded_bilinear_mult]) lemma has_vector_derivative_of_real[derivative_intros]: "(f has_field_derivative D) F \ ((\x. of_real (f x)) has_vector_derivative (of_real D)) F" by (rule bounded_linear.has_vector_derivative[OF bounded_linear_of_real]) (simp add: has_field_derivative_iff_has_vector_derivative) lemma has_vector_derivative_real_field: "(f has_field_derivative f') (at (of_real a)) \ ((\x. f (of_real x)) has_vector_derivative f') (at a within s)" using has_derivative_compose[of of_real of_real a _ f "(*) f'"] by (simp add: scaleR_conv_of_real ac_simps has_vector_derivative_def has_field_derivative_def) lemma has_vector_derivative_continuous: "(f has_vector_derivative D) (at x within s) \ continuous (at x within s) f" by (auto intro: has_derivative_continuous simp: has_vector_derivative_def) lemma continuous_on_vector_derivative: "(\x. x \ S \ (f has_vector_derivative f' x) (at x within S)) \ continuous_on S f" by (auto simp: continuous_on_eq_continuous_within intro!: has_vector_derivative_continuous) lemma has_vector_derivative_mult_right[derivative_intros]: fixes a :: "'a::real_normed_algebra" shows "(f has_vector_derivative x) F \ ((\x. a * f x) has_vector_derivative (a * x)) F" by (rule bounded_linear.has_vector_derivative[OF bounded_linear_mult_right]) lemma has_vector_derivative_mult_left[derivative_intros]: fixes a :: "'a::real_normed_algebra" shows "(f has_vector_derivative x) F \ ((\x. f x * a) has_vector_derivative (x * a)) F" by (rule bounded_linear.has_vector_derivative[OF bounded_linear_mult_left]) subsection \Derivatives\ lemma DERIV_D: "DERIV f x :> D \ (\h. (f (x + h) - f x) / h) \0\ D" by (simp add: DERIV_def) lemma has_field_derivativeD: "(f has_field_derivative D) (at x within S) \ ((\y. (f y - f x) / (y - x)) \ D) (at x within S)" by (simp add: has_field_derivative_iff) lemma DERIV_const [simp, derivative_intros]: "((\x. k) has_field_derivative 0) F" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_const]) auto lemma DERIV_ident [simp, derivative_intros]: "((\x. x) has_field_derivative 1) F" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_ident]) auto lemma field_differentiable_add[derivative_intros]: "(f has_field_derivative f') F \ (g has_field_derivative g') F \ ((\z. f z + g z) has_field_derivative f' + g') F" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_add]) (auto simp: has_field_derivative_def field_simps mult_commute_abs) corollary DERIV_add: "(f has_field_derivative D) (at x within s) \ (g has_field_derivative E) (at x within s) \ ((\x. f x + g x) has_field_derivative D + E) (at x within s)" by (rule field_differentiable_add) lemma field_differentiable_minus[derivative_intros]: "(f has_field_derivative f') F \ ((\z. - (f z)) has_field_derivative -f') F" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_minus]) (auto simp: has_field_derivative_def field_simps mult_commute_abs) corollary DERIV_minus: "(f has_field_derivative D) (at x within s) \ ((\x. - f x) has_field_derivative -D) (at x within s)" by (rule field_differentiable_minus) lemma field_differentiable_diff[derivative_intros]: "(f has_field_derivative f') F \ (g has_field_derivative g') F \ ((\z. f z - g z) has_field_derivative f' - g') F" by (simp only: diff_conv_add_uminus field_differentiable_add field_differentiable_minus) corollary DERIV_diff: "(f has_field_derivative D) (at x within s) \ (g has_field_derivative E) (at x within s) \ ((\x. f x - g x) has_field_derivative D - E) (at x within s)" by (rule field_differentiable_diff) lemma DERIV_continuous: "(f has_field_derivative D) (at x within s) \ continuous (at x within s) f" by (drule has_derivative_continuous[OF has_field_derivative_imp_has_derivative]) simp corollary DERIV_isCont: "DERIV f x :> D \ isCont f x" by (rule DERIV_continuous) lemma DERIV_atLeastAtMost_imp_continuous_on: assumes "\x. \a \ x; x \ b\ \ \y. DERIV f x :> y" shows "continuous_on {a..b} f" by (meson DERIV_isCont assms atLeastAtMost_iff continuous_at_imp_continuous_at_within continuous_on_eq_continuous_within) lemma DERIV_continuous_on: "(\x. x \ s \ (f has_field_derivative (D x)) (at x within s)) \ continuous_on s f" unfolding continuous_on_eq_continuous_within by (intro continuous_at_imp_continuous_on ballI DERIV_continuous) lemma DERIV_mult': "(f has_field_derivative D) (at x within s) \ (g has_field_derivative E) (at x within s) \ ((\x. f x * g x) has_field_derivative f x * E + D * g x) (at x within s)" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_mult]) (auto simp: field_simps mult_commute_abs dest: has_field_derivative_imp_has_derivative) lemma DERIV_mult[derivative_intros]: "(f has_field_derivative Da) (at x within s) \ (g has_field_derivative Db) (at x within s) \ ((\x. f x * g x) has_field_derivative Da * g x + Db * f x) (at x within s)" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_mult]) (auto simp: field_simps dest: has_field_derivative_imp_has_derivative) text \Derivative of linear multiplication\ lemma DERIV_cmult: "(f has_field_derivative D) (at x within s) \ ((\x. c * f x) has_field_derivative c * D) (at x within s)" by (drule DERIV_mult' [OF DERIV_const]) simp lemma DERIV_cmult_right: "(f has_field_derivative D) (at x within s) \ ((\x. f x * c) has_field_derivative D * c) (at x within s)" using DERIV_cmult by (auto simp add: ac_simps) lemma DERIV_cmult_Id [simp]: "((*) c has_field_derivative c) (at x within s)" using DERIV_ident [THEN DERIV_cmult, where c = c and x = x] by simp lemma DERIV_cdivide: "(f has_field_derivative D) (at x within s) \ ((\x. f x / c) has_field_derivative D / c) (at x within s)" using DERIV_cmult_right[of f D x s "1 / c"] by simp lemma DERIV_unique: "DERIV f x :> D \ DERIV f x :> E \ D = E" unfolding DERIV_def by (rule LIM_unique) lemma DERIV_Uniq: "\\<^sub>\\<^sub>1D. DERIV f x :> D" by (simp add: DERIV_unique Uniq_def) lemma DERIV_sum[derivative_intros]: "(\ n. n \ S \ ((\x. f x n) has_field_derivative (f' x n)) F) \ ((\x. sum (f x) S) has_field_derivative sum (f' x) S) F" by (rule has_derivative_imp_has_field_derivative [OF has_derivative_sum]) (auto simp: sum_distrib_left mult_commute_abs dest: has_field_derivative_imp_has_derivative) lemma DERIV_inverse'[derivative_intros]: assumes "(f has_field_derivative D) (at x within s)" and "f x \ 0" shows "((\x. inverse (f x)) has_field_derivative - (inverse (f x) * D * inverse (f x))) (at x within s)" proof - have "(f has_derivative (\x. x * D)) = (f has_derivative (*) D)" by (rule arg_cong [of "\x. x * D"]) (simp add: fun_eq_iff) with assms have "(f has_derivative (\x. x * D)) (at x within s)" by (auto dest!: has_field_derivative_imp_has_derivative) then show ?thesis using \f x \ 0\ by (auto intro: has_derivative_imp_has_field_derivative has_derivative_inverse) qed text \Power of \-1\\ lemma DERIV_inverse: "x \ 0 \ ((\x. inverse(x)) has_field_derivative - (inverse x ^ Suc (Suc 0))) (at x within s)" by (drule DERIV_inverse' [OF DERIV_ident]) simp text \Derivative of inverse\ lemma DERIV_inverse_fun: "(f has_field_derivative d) (at x within s) \ f x \ 0 \ ((\x. inverse (f x)) has_field_derivative (- (d * inverse(f x ^ Suc (Suc 0))))) (at x within s)" by (drule (1) DERIV_inverse') (simp add: ac_simps nonzero_inverse_mult_distrib) text \Derivative of quotient\ lemma DERIV_divide[derivative_intros]: "(f has_field_derivative D) (at x within s) \ (g has_field_derivative E) (at x within s) \ g x \ 0 \ ((\x. f x / g x) has_field_derivative (D * g x - f x * E) / (g x * g x)) (at x within s)" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_divide]) (auto dest: has_field_derivative_imp_has_derivative simp: field_simps) lemma DERIV_quotient: "(f has_field_derivative d) (at x within s) \ (g has_field_derivative e) (at x within s)\ g x \ 0 \ ((\y. f y / g y) has_field_derivative (d * g x - (e * f x)) / (g x ^ Suc (Suc 0))) (at x within s)" by (drule (2) DERIV_divide) (simp add: mult.commute) lemma DERIV_power_Suc: "(f has_field_derivative D) (at x within s) \ ((\x. f x ^ Suc n) has_field_derivative (1 + of_nat n) * (D * f x ^ n)) (at x within s)" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_power]) (auto simp: has_field_derivative_def) lemma DERIV_power[derivative_intros]: "(f has_field_derivative D) (at x within s) \ ((\x. f x ^ n) has_field_derivative of_nat n * (D * f x ^ (n - Suc 0))) (at x within s)" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_power]) (auto simp: has_field_derivative_def) lemma DERIV_pow: "((\x. x ^ n) has_field_derivative real n * (x ^ (n - Suc 0))) (at x within s)" using DERIV_power [OF DERIV_ident] by simp lemma DERIV_power_int [derivative_intros]: assumes [derivative_intros]: "(f has_field_derivative d) (at x within s)" and [simp]: "f x \ 0" shows "((\x. power_int (f x) n) has_field_derivative (of_int n * power_int (f x) (n - 1) * d)) (at x within s)" proof (cases n rule: int_cases4) case (nonneg n) thus ?thesis by (cases "n = 0") (auto intro!: derivative_eq_intros simp: field_simps power_int_diff simp flip: power_Suc power_Suc2 power_add) next case (neg n) thus ?thesis by (auto intro!: derivative_eq_intros simp: field_simps power_int_diff power_int_minus simp flip: power_Suc power_Suc2 power_add) qed lemma DERIV_chain': "(f has_field_derivative D) (at x within s) \ DERIV g (f x) :> E \ ((\x. g (f x)) has_field_derivative E * D) (at x within s)" using has_derivative_compose[of f "(*) D" x s g "(*) E"] by (simp only: has_field_derivative_def mult_commute_abs ac_simps) corollary DERIV_chain2: "DERIV f (g x) :> Da \ (g has_field_derivative Db) (at x within s) \ ((\x. f (g x)) has_field_derivative Da * Db) (at x within s)" by (rule DERIV_chain') text \Standard version\ lemma DERIV_chain: "DERIV f (g x) :> Da \ (g has_field_derivative Db) (at x within s) \ (f \ g has_field_derivative Da * Db) (at x within s)" by (drule (1) DERIV_chain', simp add: o_def mult.commute) lemma DERIV_image_chain: "(f has_field_derivative Da) (at (g x) within (g ` s)) \ (g has_field_derivative Db) (at x within s) \ (f \ g has_field_derivative Da * Db) (at x within s)" using has_derivative_in_compose [of g "(*) Db" x s f "(*) Da "] by (simp add: has_field_derivative_def o_def mult_commute_abs ac_simps) (*These two are from HOL Light: HAS_COMPLEX_DERIVATIVE_CHAIN*) lemma DERIV_chain_s: assumes "(\x. x \ s \ DERIV g x :> g'(x))" and "DERIV f x :> f'" and "f x \ s" shows "DERIV (\x. g(f x)) x :> f' * g'(f x)" by (metis (full_types) DERIV_chain' mult.commute assms) lemma DERIV_chain3: (*HAS_COMPLEX_DERIVATIVE_CHAIN_UNIV*) assumes "(\x. DERIV g x :> g'(x))" and "DERIV f x :> f'" shows "DERIV (\x. g(f x)) x :> f' * g'(f x)" by (metis UNIV_I DERIV_chain_s [of UNIV] assms) text \Alternative definition for differentiability\ lemma DERIV_LIM_iff: fixes f :: "'a::{real_normed_vector,inverse} \ 'a" shows "((\h. (f (a + h) - f a) / h) \0\ D) = ((\x. (f x - f a) / (x - a)) \a\ D)" (is "?lhs = ?rhs") proof assume ?lhs then have "(\x. (f (a + (x + - a)) - f a) / (x + - a)) \0 - - a\ D" by (rule LIM_offset) then show ?rhs by simp next assume ?rhs then have "(\x. (f (x+a) - f a) / ((x+a) - a)) \a-a\ D" by (rule LIM_offset) then show ?lhs by (simp add: add.commute) qed lemma has_field_derivative_cong_ev: assumes "x = y" and *: "eventually (\x. x \ S \ f x = g x) (nhds x)" and "u = v" "S = t" "x \ S" shows "(f has_field_derivative u) (at x within S) = (g has_field_derivative v) (at y within t)" unfolding has_field_derivative_iff proof (rule filterlim_cong) from assms have "f y = g y" by (auto simp: eventually_nhds) with * show "\\<^sub>F z in at x within S. (f z - f x) / (z - x) = (g z - g y) / (z - y)" unfolding eventually_at_filter by eventually_elim (auto simp: assms \f y = g y\) qed (simp_all add: assms) lemma has_field_derivative_cong_eventually: assumes "eventually (\x. f x = g x) (at x within S)" "f x = g x" shows "(f has_field_derivative u) (at x within S) = (g has_field_derivative u) (at x within S)" unfolding has_field_derivative_iff proof (rule tendsto_cong) show "\\<^sub>F y in at x within S. (f y - f x) / (y - x) = (g y - g x) / (y - x)" using assms by (auto elim: eventually_mono) qed lemma DERIV_cong_ev: "x = y \ eventually (\x. f x = g x) (nhds x) \ u = v \ DERIV f x :> u \ DERIV g y :> v" by (rule has_field_derivative_cong_ev) simp_all lemma DERIV_shift: "(f has_field_derivative y) (at (x + z)) = ((\x. f (x + z)) has_field_derivative y) (at x)" by (simp add: DERIV_def field_simps) lemma DERIV_mirror: "(DERIV f (- x) :> y) \ (DERIV (\x. f (- x)) x :> - y)" for f :: "real \ real" and x y :: real by (simp add: DERIV_def filterlim_at_split filterlim_at_left_to_right tendsto_minus_cancel_left field_simps conj_commute) lemma floor_has_real_derivative: fixes f :: "real \ 'a::{floor_ceiling,order_topology}" assumes "isCont f x" and "f x \ \" shows "((\x. floor (f x)) has_real_derivative 0) (at x)" proof (subst DERIV_cong_ev[OF refl _ refl]) show "((\_. floor (f x)) has_real_derivative 0) (at x)" by simp have "\\<^sub>F y in at x. \f y\ = \f x\" by (rule eventually_floor_eq[OF assms[unfolded continuous_at]]) then show "\\<^sub>F y in nhds x. real_of_int \f y\ = real_of_int \f x\" unfolding eventually_at_filter by eventually_elim auto qed lemmas has_derivative_floor[derivative_intros] = floor_has_real_derivative[THEN DERIV_compose_FDERIV] lemma continuous_floor: fixes x::real shows "x \ \ \ continuous (at x) (real_of_int \ floor)" using floor_has_real_derivative [where f=id] by (auto simp: o_def has_field_derivative_def intro: has_derivative_continuous) lemma continuous_frac: fixes x::real assumes "x \ \" shows "continuous (at x) frac" proof - have "isCont (\x. real_of_int \x\) x" using continuous_floor [OF assms] by (simp add: o_def) then have *: "continuous (at x) (\x. x - real_of_int \x\)" by (intro continuous_intros) moreover have "\\<^sub>F x in nhds x. frac x = x - real_of_int \x\" by (simp add: frac_def) ultimately show ?thesis by (simp add: LIM_imp_LIM frac_def isCont_def) qed text \Caratheodory formulation of derivative at a point\ lemma CARAT_DERIV: "(DERIV f x :> l) \ (\g. (\z. f z - f x = g z * (z - x)) \ isCont g x \ g x = l)" (is "?lhs = ?rhs") proof assume ?lhs show "\g. (\z. f z - f x = g z * (z - x)) \ isCont g x \ g x = l" proof (intro exI conjI) let ?g = "(\z. if z = x then l else (f z - f x) / (z-x))" show "\z. f z - f x = ?g z * (z - x)" by simp show "isCont ?g x" using \?lhs\ by (simp add: isCont_iff DERIV_def cong: LIM_equal [rule_format]) show "?g x = l" by simp qed next assume ?rhs then show ?lhs by (auto simp add: isCont_iff DERIV_def cong: LIM_cong) qed subsection \Local extrema\ text \If \<^term>\0 < f' x\ then \<^term>\x\ is Locally Strictly Increasing At The Right.\ lemma has_real_derivative_pos_inc_right: fixes f :: "real \ real" assumes der: "(f has_real_derivative l) (at x within S)" and l: "0 < l" shows "\d > 0. \h > 0. x + h \ S \ h < d \ f x < f (x + h)" using assms proof - from der [THEN has_field_derivativeD, THEN tendstoD, OF l, unfolded eventually_at] obtain s where s: "0 < s" and all: "\xa. xa\S \ xa \ x \ dist xa x < s \ \(f xa - f x) / (xa - x) - l\ < l" by (auto simp: dist_real_def) then show ?thesis proof (intro exI conjI strip) show "0 < s" by (rule s) next fix h :: real assume "0 < h" "h < s" "x + h \ S" with all [of "x + h"] show "f x < f (x+h)" proof (simp add: abs_if dist_real_def pos_less_divide_eq split: if_split_asm) assume "\ (f (x + h) - f x) / h < l" and h: "0 < h" with l have "0 < (f (x + h) - f x) / h" by arith then show "f x < f (x + h)" by (simp add: pos_less_divide_eq h) qed qed qed lemma DERIV_pos_inc_right: fixes f :: "real \ real" assumes der: "DERIV f x :> l" and l: "0 < l" shows "\d > 0. \h > 0. h < d \ f x < f (x + h)" using has_real_derivative_pos_inc_right[OF assms] by auto lemma has_real_derivative_neg_dec_left: fixes f :: "real \ real" assumes der: "(f has_real_derivative l) (at x within S)" and "l < 0" shows "\d > 0. \h > 0. x - h \ S \ h < d \ f x < f (x - h)" proof - from \l < 0\ have l: "- l > 0" by simp from der [THEN has_field_derivativeD, THEN tendstoD, OF l, unfolded eventually_at] obtain s where s: "0 < s" and all: "\xa. xa\S \ xa \ x \ dist xa x < s \ \(f xa - f x) / (xa - x) - l\ < - l" by (auto simp: dist_real_def) then show ?thesis proof (intro exI conjI strip) show "0 < s" by (rule s) next fix h :: real assume "0 < h" "h < s" "x - h \ S" with all [of "x - h"] show "f x < f (x-h)" proof (simp add: abs_if pos_less_divide_eq dist_real_def split: if_split_asm) assume "- ((f (x-h) - f x) / h) < l" and h: "0 < h" with l have "0 < (f (x-h) - f x) / h" by arith then show "f x < f (x - h)" by (simp add: pos_less_divide_eq h) qed qed qed lemma DERIV_neg_dec_left: fixes f :: "real \ real" assumes der: "DERIV f x :> l" and l: "l < 0" shows "\d > 0. \h > 0. h < d \ f x < f (x - h)" using has_real_derivative_neg_dec_left[OF assms] by auto lemma has_real_derivative_pos_inc_left: fixes f :: "real \ real" shows "(f has_real_derivative l) (at x within S) \ 0 < l \ \d>0. \h>0. x - h \ S \ h < d \ f (x - h) < f x" by (rule has_real_derivative_neg_dec_left [of "\x. - f x" "-l" x S, simplified]) (auto simp add: DERIV_minus) lemma DERIV_pos_inc_left: fixes f :: "real \ real" shows "DERIV f x :> l \ 0 < l \ \d > 0. \h > 0. h < d \ f (x - h) < f x" using has_real_derivative_pos_inc_left by blast lemma has_real_derivative_neg_dec_right: fixes f :: "real \ real" shows "(f has_real_derivative l) (at x within S) \ l < 0 \ \d > 0. \h > 0. x + h \ S \ h < d \ f x > f (x + h)" by (rule has_real_derivative_pos_inc_right [of "\x. - f x" "-l" x S, simplified]) (auto simp add: DERIV_minus) lemma DERIV_neg_dec_right: fixes f :: "real \ real" shows "DERIV f x :> l \ l < 0 \ \d > 0. \h > 0. h < d \ f x > f (x + h)" using has_real_derivative_neg_dec_right by blast lemma DERIV_local_max: fixes f :: "real \ real" assumes der: "DERIV f x :> l" and d: "0 < d" and le: "\y. \x - y\ < d \ f y \ f x" shows "l = 0" proof (cases rule: linorder_cases [of l 0]) case equal then show ?thesis . next case less from DERIV_neg_dec_left [OF der less] obtain d' where d': "0 < d'" and lt: "\h > 0. h < d' \ f x < f (x - h)" by blast obtain e where "0 < e \ e < d \ e < d'" using field_lbound_gt_zero [OF d d'] .. with lt le [THEN spec [where x="x - e"]] show ?thesis by (auto simp add: abs_if) next case greater from DERIV_pos_inc_right [OF der greater] obtain d' where d': "0 < d'" and lt: "\h > 0. h < d' \ f x < f (x + h)" by blast obtain e where "0 < e \ e < d \ e < d'" using field_lbound_gt_zero [OF d d'] .. with lt le [THEN spec [where x="x + e"]] show ?thesis by (auto simp add: abs_if) qed text \Similar theorem for a local minimum\ lemma DERIV_local_min: fixes f :: "real \ real" shows "DERIV f x :> l \ 0 < d \ \y. \x - y\ < d \ f x \ f y \ l = 0" by (drule DERIV_minus [THEN DERIV_local_max]) auto text\In particular, if a function is locally flat\ lemma DERIV_local_const: fixes f :: "real \ real" shows "DERIV f x :> l \ 0 < d \ \y. \x - y\ < d \ f x = f y \ l = 0" by (auto dest!: DERIV_local_max) subsection \Rolle's Theorem\ text \Lemma about introducing open ball in open interval\ lemma lemma_interval_lt: fixes a b x :: real assumes "a < x" "x < b" shows "\d. 0 < d \ (\y. \x - y\ < d \ a < y \ y < b)" using linorder_linear [of "x - a" "b - x"] proof assume "x - a \ b - x" with assms show ?thesis by (rule_tac x = "x - a" in exI) auto next assume "b - x \ x - a" with assms show ?thesis by (rule_tac x = "b - x" in exI) auto qed lemma lemma_interval: "a < x \ x < b \ \d. 0 < d \ (\y. \x - y\ < d \ a \ y \ y \ b)" for a b x :: real by (force dest: lemma_interval_lt) text \Rolle's Theorem. If \<^term>\f\ is defined and continuous on the closed interval \[a,b]\ and differentiable on the open interval \(a,b)\, and \<^term>\f a = f b\, then there exists \x0 \ (a,b)\ such that \<^term>\f' x0 = 0\\ theorem Rolle_deriv: fixes f :: "real \ real" assumes "a < b" and fab: "f a = f b" and contf: "continuous_on {a..b} f" and derf: "\x. \a < x; x < b\ \ (f has_derivative f' x) (at x)" shows "\z. a < z \ z < b \ f' z = (\v. 0)" proof - have le: "a \ b" using \a < b\ by simp have "(a + b) / 2 \ {a..b}" using assms(1) by auto then have *: "{a..b} \ {}" by auto obtain x where x_max: "\z. a \ z \ z \ b \ f z \ f x" and "a \ x" "x \ b" using continuous_attains_sup[OF compact_Icc * contf] by (meson atLeastAtMost_iff) obtain x' where x'_min: "\z. a \ z \ z \ b \ f x' \ f z" and "a \ x'" "x' \ b" using continuous_attains_inf[OF compact_Icc * contf] by (meson atLeastAtMost_iff) consider "a < x" "x < b" | "x = a \ x = b" using \a \ x\ \x \ b\ by arith then show ?thesis proof cases case 1 \ \\<^term>\f\ attains its maximum within the interval\ then obtain l where der: "DERIV f x :> l" using derf differentiable_def real_differentiable_def by blast obtain d where d: "0 < d" and bound: "\y. \x - y\ < d \ a \ y \ y \ b" using lemma_interval [OF 1] by blast then have bound': "\y. \x - y\ < d \ f y \ f x" using x_max by blast \ \the derivative at a local maximum is zero\ have "l = 0" by (rule DERIV_local_max [OF der d bound']) with 1 der derf [of x] show ?thesis by (metis has_derivative_unique has_field_derivative_def mult_zero_left) next case 2 then have fx: "f b = f x" by (auto simp add: fab) consider "a < x'" "x' < b" | "x' = a \ x' = b" using \a \ x'\ \x' \ b\ by arith then show ?thesis proof cases case 1 \ \\<^term>\f\ attains its minimum within the interval\ then obtain l where der: "DERIV f x' :> l" using derf differentiable_def real_differentiable_def by blast from lemma_interval [OF 1] obtain d where d: "0y. \x'-y\ < d \ a \ y \ y \ b" by blast then have bound': "\y. \x' - y\ < d \ f x' \ f y" using x'_min by blast have "l = 0" by (rule DERIV_local_min [OF der d bound']) \ \the derivative at a local minimum is zero\ then show ?thesis using 1 der derf [of x'] by (metis has_derivative_unique has_field_derivative_def mult_zero_left) next case 2 \ \\<^term>\f\ is constant throughout the interval\ then have fx': "f b = f x'" by (auto simp: fab) from dense [OF \a < b\] obtain r where r: "a < r" "r < b" by blast obtain d where d: "0 < d" and bound: "\y. \r - y\ < d \ a \ y \ y \ b" using lemma_interval [OF r] by blast have eq_fb: "f z = f b" if "a \ z" and "z \ b" for z proof (rule order_antisym) show "f z \ f b" by (simp add: fx x_max that) show "f b \ f z" by (simp add: fx' x'_min that) qed have bound': "\y. \r - y\ < d \ f r = f y" proof (intro strip) fix y :: real assume lt: "\r - y\ < d" then have "f y = f b" by (simp add: eq_fb bound) then show "f r = f y" by (simp add: eq_fb r order_less_imp_le) qed obtain l where der: "DERIV f r :> l" using derf differentiable_def r(1) r(2) real_differentiable_def by blast have "l = 0" by (rule DERIV_local_const [OF der d bound']) \ \the derivative of a constant function is zero\ with r der derf [of r] show ?thesis by (metis has_derivative_unique has_field_derivative_def mult_zero_left) qed qed qed corollary Rolle: fixes a b :: real assumes ab: "a < b" "f a = f b" "continuous_on {a..b} f" and dif [rule_format]: "\x. \a < x; x < b\ \ f differentiable (at x)" shows "\z. a < z \ z < b \ DERIV f z :> 0" proof - obtain f' where f': "\x. \a < x; x < b\ \ (f has_derivative f' x) (at x)" using dif unfolding differentiable_def by metis then have "\z. a < z \ z < b \ f' z = (\v. 0)" by (metis Rolle_deriv [OF ab]) then show ?thesis using f' has_derivative_imp_has_field_derivative by fastforce qed subsection \Mean Value Theorem\ theorem mvt: fixes f :: "real \ real" assumes "a < b" and contf: "continuous_on {a..b} f" and derf: "\x. \a < x; x < b\ \ (f has_derivative f' x) (at x)" obtains \ where "a < \" "\ < b" "f b - f a = (f' \) (b - a)" proof - have "\x. a < x \ x < b \ (\y. f' x y - (f b - f a) / (b - a) * y) = (\v. 0)" proof (intro Rolle_deriv[OF \a < b\]) fix x assume x: "a < x" "x < b" show "((\x. f x - (f b - f a) / (b - a) * x) has_derivative (\y. f' x y - (f b - f a) / (b - a) * y)) (at x)" by (intro derivative_intros derf[OF x]) qed (use assms in \auto intro!: continuous_intros simp: field_simps\) then obtain \ where "a < \" "\ < b" "(\y. f' \ y - (f b - f a) / (b - a) * y) = (\v. 0)" by metis then show ?thesis by (metis (no_types, hide_lams) that add.right_neutral add_diff_cancel_left' add_diff_eq \a < b\ less_irrefl nonzero_eq_divide_eq) qed theorem MVT: fixes a b :: real assumes lt: "a < b" and contf: "continuous_on {a..b} f" and dif: "\x. \a < x; x < b\ \ f differentiable (at x)" shows "\l z. a < z \ z < b \ DERIV f z :> l \ f b - f a = (b - a) * l" proof - obtain f' :: "real \ real \ real" where derf: "\x. a < x \ x < b \ (f has_derivative f' x) (at x)" using dif unfolding differentiable_def by metis then obtain z where "a < z" "z < b" "f b - f a = (f' z) (b - a)" using mvt [OF lt contf] by blast then show ?thesis by (simp add: ac_simps) (metis derf dif has_derivative_unique has_field_derivative_imp_has_derivative real_differentiable_def) qed corollary MVT2: assumes "a < b" and der: "\x. \a \ x; x \ b\ \ DERIV f x :> f' x" shows "\z::real. a < z \ z < b \ (f b - f a = (b - a) * f' z)" proof - have "\l z. a < z \ z < b \ (f has_real_derivative l) (at z) \ f b - f a = (b - a) * l" proof (rule MVT [OF \a < b\]) show "continuous_on {a..b} f" by (meson DERIV_continuous atLeastAtMost_iff continuous_at_imp_continuous_on der) show "\x. \a < x; x < b\ \ f differentiable (at x)" using assms by (force dest: order_less_imp_le simp add: real_differentiable_def) qed with assms show ?thesis by (blast dest: DERIV_unique order_less_imp_le) qed lemma pos_deriv_imp_strict_mono: assumes "\x. (f has_real_derivative f' x) (at x)" assumes "\x. f' x > 0" shows "strict_mono f" proof (rule strict_monoI) fix x y :: real assume xy: "x < y" from assms and xy have "\z>x. z < y \ f y - f x = (y - x) * f' z" by (intro MVT2) (auto dest: connectedD_interval) then obtain z where z: "z > x" "z < y" "f y - f x = (y - x) * f' z" by blast note \f y - f x = (y - x) * f' z\ also have "(y - x) * f' z > 0" using xy assms by (intro mult_pos_pos) auto finally show "f x < f y" by simp qed proposition deriv_nonneg_imp_mono: assumes deriv: "\x. x \ {a..b} \ (g has_real_derivative g' x) (at x)" assumes nonneg: "\x. x \ {a..b} \ g' x \ 0" assumes ab: "a \ b" shows "g a \ g b" proof (cases "a < b") assume "a < b" from deriv have "\x. \x \ a; x \ b\ \ (g has_real_derivative g' x) (at x)" by simp with MVT2[OF \a < b\] and deriv obtain \ where \_ab: "\ > a" "\ < b" and g_ab: "g b - g a = (b - a) * g' \" by blast from \_ab ab nonneg have "(b - a) * g' \ \ 0" by simp with g_ab show ?thesis by simp qed (insert ab, simp) subsubsection \A function is constant if its derivative is 0 over an interval.\ lemma DERIV_isconst_end: fixes f :: "real \ real" assumes "a < b" and contf: "continuous_on {a..b} f" and 0: "\x. \a < x; x < b\ \ DERIV f x :> 0" shows "f b = f a" using MVT [OF \a < b\] "0" DERIV_unique contf real_differentiable_def by (fastforce simp: algebra_simps) lemma DERIV_isconst2: fixes f :: "real \ real" assumes "a < b" and contf: "continuous_on {a..b} f" and derf: "\x. \a < x; x < b\ \ DERIV f x :> 0" and "a \ x" "x \ b" shows "f x = f a" proof (cases "a < x") case True have *: "continuous_on {a..x} f" using \x \ b\ contf continuous_on_subset by fastforce show ?thesis by (rule DERIV_isconst_end [OF True *]) (use \x \ b\ derf in auto) qed (use \a \ x\ in auto) lemma DERIV_isconst3: fixes a b x y :: real assumes "a < b" and "x \ {a <..< b}" and "y \ {a <..< b}" and derivable: "\x. x \ {a <..< b} \ DERIV f x :> 0" shows "f x = f y" proof (cases "x = y") case False let ?a = "min x y" let ?b = "max x y" have *: "DERIV f z :> 0" if "?a \ z" "z \ ?b" for z proof - have "a < z" and "z < b" using that \x \ {a <..< b}\ and \y \ {a <..< b}\ by auto then have "z \ {a<.. 0" by (rule derivable) qed have isCont: "continuous_on {?a..?b} f" by (meson * DERIV_continuous_on atLeastAtMost_iff has_field_derivative_at_within) have DERIV: "\z. \?a < z; z < ?b\ \ DERIV f z :> 0" using * by auto have "?a < ?b" using \x \ y\ by auto from DERIV_isconst2[OF this isCont DERIV, of x] and DERIV_isconst2[OF this isCont DERIV, of y] show ?thesis by auto qed auto lemma DERIV_isconst_all: fixes f :: "real \ real" shows "\x. DERIV f x :> 0 \ f x = f y" apply (rule linorder_cases [of x y]) apply (metis DERIV_continuous DERIV_isconst_end continuous_at_imp_continuous_on)+ done lemma DERIV_const_ratio_const: fixes f :: "real \ real" assumes "a \ b" and df: "\x. DERIV f x :> k" shows "f b - f a = (b - a) * k" proof (cases a b rule: linorder_cases) case less show ?thesis using MVT [OF less] df by (metis DERIV_continuous DERIV_unique continuous_at_imp_continuous_on real_differentiable_def) next case greater have "f a - f b = (a - b) * k" using MVT [OF greater] df by (metis DERIV_continuous DERIV_unique continuous_at_imp_continuous_on real_differentiable_def) then show ?thesis by (simp add: algebra_simps) qed auto lemma DERIV_const_ratio_const2: fixes f :: "real \ real" assumes "a \ b" and df: "\x. DERIV f x :> k" shows "(f b - f a) / (b - a) = k" using DERIV_const_ratio_const [OF assms] \a \ b\ by auto lemma real_average_minus_first [simp]: "(a + b) / 2 - a = (b - a) / 2" for a b :: real by simp lemma real_average_minus_second [simp]: "(b + a) / 2 - a = (b - a) / 2" for a b :: real by simp text \Gallileo's "trick": average velocity = av. of end velocities.\ lemma DERIV_const_average: fixes v :: "real \ real" and a b :: real assumes neq: "a \ b" and der: "\x. DERIV v x :> k" shows "v ((a + b) / 2) = (v a + v b) / 2" proof (cases rule: linorder_cases [of a b]) case equal with neq show ?thesis by simp next case less have "(v b - v a) / (b - a) = k" by (rule DERIV_const_ratio_const2 [OF neq der]) then have "(b - a) * ((v b - v a) / (b - a)) = (b - a) * k" by simp moreover have "(v ((a + b) / 2) - v a) / ((a + b) / 2 - a) = k" by (rule DERIV_const_ratio_const2 [OF _ der]) (simp add: neq) ultimately show ?thesis using neq by force next case greater have "(v b - v a) / (b - a) = k" by (rule DERIV_const_ratio_const2 [OF neq der]) then have "(b - a) * ((v b - v a) / (b - a)) = (b - a) * k" by simp moreover have " (v ((b + a) / 2) - v a) / ((b + a) / 2 - a) = k" by (rule DERIV_const_ratio_const2 [OF _ der]) (simp add: neq) ultimately show ?thesis using neq by (force simp add: add.commute) qed subsubsection\A function with positive derivative is increasing\ text \A simple proof using the MVT, by Jeremy Avigad. And variants.\ lemma DERIV_pos_imp_increasing_open: fixes a b :: real and f :: "real \ real" assumes "a < b" and "\x. a < x \ x < b \ (\y. DERIV f x :> y \ y > 0)" and con: "continuous_on {a..b} f" shows "f a < f b" proof (rule ccontr) assume f: "\ ?thesis" have "\l z. a < z \ z < b \ DERIV f z :> l \ f b - f a = (b - a) * l" by (rule MVT) (use assms real_differentiable_def in \force+\) then obtain l z where z: "a < z" "z < b" "DERIV f z :> l" and "f b - f a = (b - a) * l" by auto with assms f have "\ l > 0" by (metis linorder_not_le mult_le_0_iff diff_le_0_iff_le) with assms z show False by (metis DERIV_unique) qed lemma DERIV_pos_imp_increasing: fixes a b :: real and f :: "real \ real" assumes "a < b" and der: "\x. \a \ x; x \ b\ \ \y. DERIV f x :> y \ y > 0" shows "f a < f b" by (metis less_le_not_le DERIV_atLeastAtMost_imp_continuous_on DERIV_pos_imp_increasing_open [OF \a < b\] der) lemma DERIV_nonneg_imp_nondecreasing: fixes a b :: real and f :: "real \ real" assumes "a \ b" and "\x. \a \ x; x \ b\ \ \y. DERIV f x :> y \ y \ 0" shows "f a \ f b" proof (rule ccontr, cases "a = b") assume "\ ?thesis" and "a = b" then show False by auto next assume *: "\ ?thesis" assume "a \ b" with \a \ b\ have "a < b" by linarith moreover have "continuous_on {a..b} f" by (meson DERIV_isCont assms(2) atLeastAtMost_iff continuous_at_imp_continuous_on) ultimately have "\l z. a < z \ z < b \ DERIV f z :> l \ f b - f a = (b - a) * l" using assms MVT [OF \a < b\, of f] real_differentiable_def less_eq_real_def by blast then obtain l z where lz: "a < z" "z < b" "DERIV f z :> l" and **: "f b - f a = (b - a) * l" by auto with * have "a < b" "f b < f a" by auto with ** have "\ l \ 0" by (auto simp add: not_le algebra_simps) (metis * add_le_cancel_right assms(1) less_eq_real_def mult_right_mono add_left_mono linear order_refl) with assms lz show False by (metis DERIV_unique order_less_imp_le) qed lemma DERIV_neg_imp_decreasing_open: fixes a b :: real and f :: "real \ real" assumes "a < b" and "\x. a < x \ x < b \ \y. DERIV f x :> y \ y < 0" and con: "continuous_on {a..b} f" shows "f a > f b" proof - have "(\x. -f x) a < (\x. -f x) b" proof (rule DERIV_pos_imp_increasing_open [of a b]) show "\x. \a < x; x < b\ \ \y. ((\x. - f x) has_real_derivative y) (at x) \ 0 < y" using assms by simp (metis field_differentiable_minus neg_0_less_iff_less) show "continuous_on {a..b} (\x. - f x)" using con continuous_on_minus by blast qed (use assms in auto) then show ?thesis by simp qed lemma DERIV_neg_imp_decreasing: fixes a b :: real and f :: "real \ real" assumes "a < b" and der: "\x. \a \ x; x \ b\ \ \y. DERIV f x :> y \ y < 0" shows "f a > f b" by (metis less_le_not_le DERIV_atLeastAtMost_imp_continuous_on DERIV_neg_imp_decreasing_open [OF \a < b\] der) lemma DERIV_nonpos_imp_nonincreasing: fixes a b :: real and f :: "real \ real" assumes "a \ b" and "\x. \a \ x; x \ b\ \ \y. DERIV f x :> y \ y \ 0" shows "f a \ f b" proof - have "(\x. -f x) a \ (\x. -f x) b" using DERIV_nonneg_imp_nondecreasing [of a b "\x. -f x"] assms DERIV_minus by fastforce then show ?thesis by simp qed lemma DERIV_pos_imp_increasing_at_bot: fixes f :: "real \ real" assumes "\x. x \ b \ (\y. DERIV f x :> y \ y > 0)" and lim: "(f \ flim) at_bot" shows "flim < f b" proof - have "\N. \n\N. f n \ f (b - 1)" by (rule_tac x="b - 2" in exI) (force intro: order.strict_implies_order DERIV_pos_imp_increasing assms) then have "flim \ f (b - 1)" by (auto simp: eventually_at_bot_linorder tendsto_upperbound [OF lim]) also have "\ < f b" by (force intro: DERIV_pos_imp_increasing [where f=f] assms) finally show ?thesis . qed lemma DERIV_neg_imp_decreasing_at_top: fixes f :: "real \ real" assumes der: "\x. x \ b \ \y. DERIV f x :> y \ y < 0" and lim: "(f \ flim) at_top" shows "flim < f b" apply (rule DERIV_pos_imp_increasing_at_bot [where f = "\i. f (-i)" and b = "-b", simplified]) apply (metis DERIV_mirror der le_minus_iff neg_0_less_iff_less) apply (metis filterlim_at_top_mirror lim) done text \Derivative of inverse function\ lemma DERIV_inverse_function: fixes f g :: "real \ real" assumes der: "DERIV f (g x) :> D" and neq: "D \ 0" and x: "a < x" "x < b" and inj: "\y. \a < y; y < b\ \ f (g y) = y" and cont: "isCont g x" shows "DERIV g x :> inverse D" unfolding has_field_derivative_iff proof (rule LIM_equal2) show "0 < min (x - a) (b - x)" using x by arith next fix y assume "norm (y - x) < min (x - a) (b - x)" then have "a < y" and "y < b" by (simp_all add: abs_less_iff) then show "(g y - g x) / (y - x) = inverse ((f (g y) - x) / (g y - g x))" by (simp add: inj) next have "(\z. (f z - f (g x)) / (z - g x)) \g x\ D" by (rule der [unfolded has_field_derivative_iff]) then have 1: "(\z. (f z - x) / (z - g x)) \g x\ D" using inj x by simp have 2: "\d>0. \y. y \ x \ norm (y - x) < d \ g y \ g x" proof (rule exI, safe) show "0 < min (x - a) (b - x)" using x by simp next fix y assume "norm (y - x) < min (x - a) (b - x)" then have y: "a < y" "y < b" by (simp_all add: abs_less_iff) assume "g y = g x" then have "f (g y) = f (g x)" by simp then have "y = x" using inj y x by simp also assume "y \ x" finally show False by simp qed have "(\y. (f (g y) - x) / (g y - g x)) \x\ D" using cont 1 2 by (rule isCont_LIM_compose2) then show "(\y. inverse ((f (g y) - x) / (g y - g x))) \x\ inverse D" using neq by (rule tendsto_inverse) qed subsection \Generalized Mean Value Theorem\ theorem GMVT: fixes a b :: real assumes alb: "a < b" and fc: "\x. a \ x \ x \ b \ isCont f x" and fd: "\x. a < x \ x < b \ f differentiable (at x)" and gc: "\x. a \ x \ x \ b \ isCont g x" and gd: "\x. a < x \ x < b \ g differentiable (at x)" shows "\g'c f'c c. DERIV g c :> g'c \ DERIV f c :> f'c \ a < c \ c < b \ (f b - f a) * g'c = (g b - g a) * f'c" proof - let ?h = "\x. (f b - f a) * g x - (g b - g a) * f x" have "\l z. a < z \ z < b \ DERIV ?h z :> l \ ?h b - ?h a = (b - a) * l" proof (rule MVT) from assms show "a < b" by simp show "continuous_on {a..b} ?h" by (simp add: continuous_at_imp_continuous_on fc gc) show "\x. \a < x; x < b\ \ ?h differentiable (at x)" using fd gd by simp qed then obtain l where l: "\z. a < z \ z < b \ DERIV ?h z :> l \ ?h b - ?h a = (b - a) * l" .. then obtain c where c: "a < c \ c < b \ DERIV ?h c :> l \ ?h b - ?h a = (b - a) * l" .. from c have cint: "a < c \ c < b" by auto then obtain g'c where g'c: "DERIV g c :> g'c" using gd real_differentiable_def by blast from c have "a < c \ c < b" by auto then obtain f'c where f'c: "DERIV f c :> f'c" using fd real_differentiable_def by blast from c have "DERIV ?h c :> l" by auto moreover have "DERIV ?h c :> g'c * (f b - f a) - f'c * (g b - g a)" using g'c f'c by (auto intro!: derivative_eq_intros) ultimately have leq: "l = g'c * (f b - f a) - f'c * (g b - g a)" by (rule DERIV_unique) have "?h b - ?h a = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" proof - from c have "?h b - ?h a = (b - a) * l" by auto also from leq have "\ = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" by simp finally show ?thesis by simp qed moreover have "?h b - ?h a = 0" proof - have "?h b - ?h a = ((f b)*(g b) - (f a)*(g b) - (g b)*(f b) + (g a)*(f b)) - ((f b)*(g a) - (f a)*(g a) - (g b)*(f a) + (g a)*(f a))" by (simp add: algebra_simps) then show ?thesis by auto qed ultimately have "(b - a) * (g'c * (f b - f a) - f'c * (g b - g a)) = 0" by auto with alb have "g'c * (f b - f a) - f'c * (g b - g a) = 0" by simp then have "g'c * (f b - f a) = f'c * (g b - g a)" by simp then have "(f b - f a) * g'c = (g b - g a) * f'c" by (simp add: ac_simps) with g'c f'c cint show ?thesis by auto qed lemma GMVT': fixes f g :: "real \ real" assumes "a < b" and isCont_f: "\z. a \ z \ z \ b \ isCont f z" and isCont_g: "\z. a \ z \ z \ b \ isCont g z" and DERIV_g: "\z. a < z \ z < b \ DERIV g z :> (g' z)" and DERIV_f: "\z. a < z \ z < b \ DERIV f z :> (f' z)" shows "\c. a < c \ c < b \ (f b - f a) * g' c = (g b - g a) * f' c" proof - have "\g'c f'c c. DERIV g c :> g'c \ DERIV f c :> f'c \ a < c \ c < b \ (f b - f a) * g'c = (g b - g a) * f'c" using assms by (intro GMVT) (force simp: real_differentiable_def)+ then obtain c where "a < c" "c < b" "(f b - f a) * g' c = (g b - g a) * f' c" using DERIV_f DERIV_g by (force dest: DERIV_unique) then show ?thesis by auto qed subsection \L'Hopitals rule\ lemma isCont_If_ge: fixes a :: "'a :: linorder_topology" assumes "continuous (at_left a) g" and f: "(f \ g a) (at_right a)" shows "isCont (\x. if x \ a then g x else f x) a" (is "isCont ?gf a") proof - have g: "(g \ g a) (at_left a)" using assms continuous_within by blast show ?thesis unfolding isCont_def continuous_within proof (intro filterlim_split_at; simp) show "(?gf \ g a) (at_left a)" by (subst filterlim_cong[OF refl refl, where g=g]) (simp_all add: eventually_at_filter less_le g) show "(?gf \ g a) (at_right a)" by (subst filterlim_cong[OF refl refl, where g=f]) (simp_all add: eventually_at_filter less_le f) qed qed lemma lhopital_right_0: fixes f0 g0 :: "real \ real" assumes f_0: "(f0 \ 0) (at_right 0)" and g_0: "(g0 \ 0) (at_right 0)" and ev: "eventually (\x. g0 x \ 0) (at_right 0)" "eventually (\x. g' x \ 0) (at_right 0)" "eventually (\x. DERIV f0 x :> f' x) (at_right 0)" "eventually (\x. DERIV g0 x :> g' x) (at_right 0)" and lim: "filterlim (\ x. (f' x / g' x)) F (at_right 0)" shows "filterlim (\ x. f0 x / g0 x) F (at_right 0)" proof - define f where [abs_def]: "f x = (if x \ 0 then 0 else f0 x)" for x then have "f 0 = 0" by simp define g where [abs_def]: "g x = (if x \ 0 then 0 else g0 x)" for x then have "g 0 = 0" by simp have "eventually (\x. g0 x \ 0 \ g' x \ 0 \ DERIV f0 x :> (f' x) \ DERIV g0 x :> (g' x)) (at_right 0)" using ev by eventually_elim auto then obtain a where [arith]: "0 < a" and g0_neq_0: "\x. 0 < x \ x < a \ g0 x \ 0" and g'_neq_0: "\x. 0 < x \ x < a \ g' x \ 0" and f0: "\x. 0 < x \ x < a \ DERIV f0 x :> (f' x)" and g0: "\x. 0 < x \ x < a \ DERIV g0 x :> (g' x)" unfolding eventually_at by (auto simp: dist_real_def) have g_neq_0: "\x. 0 < x \ x < a \ g x \ 0" using g0_neq_0 by (simp add: g_def) have f: "DERIV f x :> (f' x)" if x: "0 < x" "x < a" for x using that by (intro DERIV_cong_ev[THEN iffD1, OF _ _ _ f0[OF x]]) (auto simp: f_def eventually_nhds_metric dist_real_def intro!: exI[of _ x]) have g: "DERIV g x :> (g' x)" if x: "0 < x" "x < a" for x using that by (intro DERIV_cong_ev[THEN iffD1, OF _ _ _ g0[OF x]]) (auto simp: g_def eventually_nhds_metric dist_real_def intro!: exI[of _ x]) have "isCont f 0" unfolding f_def by (intro isCont_If_ge f_0 continuous_const) have "isCont g 0" unfolding g_def by (intro isCont_If_ge g_0 continuous_const) have "\\. \x\{0 <..< a}. 0 < \ x \ \ x < x \ f x / g x = f' (\ x) / g' (\ x)" proof (rule bchoice, rule ballI) fix x assume "x \ {0 <..< a}" then have x[arith]: "0 < x" "x < a" by auto with g'_neq_0 g_neq_0 \g 0 = 0\ have g': "\x. 0 < x \ x < a \ 0 \ g' x" "g 0 \ g x" by auto have "\x. 0 \ x \ x < a \ isCont f x" using \isCont f 0\ f by (auto intro: DERIV_isCont simp: le_less) moreover have "\x. 0 \ x \ x < a \ isCont g x" using \isCont g 0\ g by (auto intro: DERIV_isCont simp: le_less) ultimately have "\c. 0 < c \ c < x \ (f x - f 0) * g' c = (g x - g 0) * f' c" using f g \x < a\ by (intro GMVT') auto then obtain c where *: "0 < c" "c < x" "(f x - f 0) * g' c = (g x - g 0) * f' c" by blast moreover from * g'(1)[of c] g'(2) have "(f x - f 0) / (g x - g 0) = f' c / g' c" by (simp add: field_simps) ultimately show "\y. 0 < y \ y < x \ f x / g x = f' y / g' y" using \f 0 = 0\ \g 0 = 0\ by (auto intro!: exI[of _ c]) qed then obtain \ where "\x\{0 <..< a}. 0 < \ x \ \ x < x \ f x / g x = f' (\ x) / g' (\ x)" .. then have \: "eventually (\x. 0 < \ x \ \ x < x \ f x / g x = f' (\ x) / g' (\ x)) (at_right 0)" unfolding eventually_at by (intro exI[of _ a]) (auto simp: dist_real_def) moreover from \ have "eventually (\x. norm (\ x) \ x) (at_right 0)" by eventually_elim auto then have "((\x. norm (\ x)) \ 0) (at_right 0)" by (rule_tac real_tendsto_sandwich[where f="\x. 0" and h="\x. x"]) auto then have "(\ \ 0) (at_right 0)" by (rule tendsto_norm_zero_cancel) with \ have "filterlim \ (at_right 0) (at_right 0)" by (auto elim!: eventually_mono simp: filterlim_at) from this lim have "filterlim (\t. f' (\ t) / g' (\ t)) F (at_right 0)" by (rule_tac filterlim_compose[of _ _ _ \]) ultimately have "filterlim (\t. f t / g t) F (at_right 0)" (is ?P) by (rule_tac filterlim_cong[THEN iffD1, OF refl refl]) (auto elim: eventually_mono) also have "?P \ ?thesis" by (rule filterlim_cong) (auto simp: f_def g_def eventually_at_filter) finally show ?thesis . qed lemma lhopital_right: "(f \ 0) (at_right x) \ (g \ 0) (at_right x) \ eventually (\x. g x \ 0) (at_right x) \ eventually (\x. g' x \ 0) (at_right x) \ eventually (\x. DERIV f x :> f' x) (at_right x) \ eventually (\x. DERIV g x :> g' x) (at_right x) \ filterlim (\ x. (f' x / g' x)) F (at_right x) \ filterlim (\ x. f x / g x) F (at_right x)" for x :: real unfolding eventually_at_right_to_0[of _ x] filterlim_at_right_to_0[of _ _ x] DERIV_shift by (rule lhopital_right_0) lemma lhopital_left: "(f \ 0) (at_left x) \ (g \ 0) (at_left x) \ eventually (\x. g x \ 0) (at_left x) \ eventually (\x. g' x \ 0) (at_left x) \ eventually (\x. DERIV f x :> f' x) (at_left x) \ eventually (\x. DERIV g x :> g' x) (at_left x) \ filterlim (\ x. (f' x / g' x)) F (at_left x) \ filterlim (\ x. f x / g x) F (at_left x)" for x :: real unfolding eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror by (rule lhopital_right[where f'="\x. - f' (- x)"]) (auto simp: DERIV_mirror) lemma lhopital: "(f \ 0) (at x) \ (g \ 0) (at x) \ eventually (\x. g x \ 0) (at x) \ eventually (\x. g' x \ 0) (at x) \ eventually (\x. DERIV f x :> f' x) (at x) \ eventually (\x. DERIV g x :> g' x) (at x) \ filterlim (\ x. (f' x / g' x)) F (at x) \ filterlim (\ x. f x / g x) F (at x)" for x :: real unfolding eventually_at_split filterlim_at_split by (auto intro!: lhopital_right[of f x g g' f'] lhopital_left[of f x g g' f']) lemma lhopital_right_0_at_top: fixes f g :: "real \ real" assumes g_0: "LIM x at_right 0. g x :> at_top" and ev: "eventually (\x. g' x \ 0) (at_right 0)" "eventually (\x. DERIV f x :> f' x) (at_right 0)" "eventually (\x. DERIV g x :> g' x) (at_right 0)" and lim: "((\ x. (f' x / g' x)) \ x) (at_right 0)" shows "((\ x. f x / g x) \ x) (at_right 0)" unfolding tendsto_iff proof safe fix e :: real assume "0 < e" with lim[unfolded tendsto_iff, rule_format, of "e / 4"] have "eventually (\t. dist (f' t / g' t) x < e / 4) (at_right 0)" by simp from eventually_conj[OF eventually_conj[OF ev(1) ev(2)] eventually_conj[OF ev(3) this]] obtain a where [arith]: "0 < a" and g'_neq_0: "\x. 0 < x \ x < a \ g' x \ 0" and f0: "\x. 0 < x \ x \ a \ DERIV f x :> (f' x)" and g0: "\x. 0 < x \ x \ a \ DERIV g x :> (g' x)" and Df: "\t. 0 < t \ t < a \ dist (f' t / g' t) x < e / 4" unfolding eventually_at_le by (auto simp: dist_real_def) from Df have "eventually (\t. t < a) (at_right 0)" "eventually (\t::real. 0 < t) (at_right 0)" unfolding eventually_at by (auto intro!: exI[of _ a] simp: dist_real_def) moreover have "eventually (\t. 0 < g t) (at_right 0)" "eventually (\t. g a < g t) (at_right 0)" using g_0 by (auto elim: eventually_mono simp: filterlim_at_top_dense) moreover have inv_g: "((\x. inverse (g x)) \ 0) (at_right 0)" using tendsto_inverse_0 filterlim_mono[OF g_0 at_top_le_at_infinity order_refl] by (rule filterlim_compose) then have "((\x. norm (1 - g a * inverse (g x))) \ norm (1 - g a * 0)) (at_right 0)" by (intro tendsto_intros) then have "((\x. norm (1 - g a / g x)) \ 1) (at_right 0)" by (simp add: inverse_eq_divide) from this[unfolded tendsto_iff, rule_format, of 1] have "eventually (\x. norm (1 - g a / g x) < 2) (at_right 0)" by (auto elim!: eventually_mono simp: dist_real_def) moreover from inv_g have "((\t. norm ((f a - x * g a) * inverse (g t))) \ norm ((f a - x * g a) * 0)) (at_right 0)" by (intro tendsto_intros) then have "((\t. norm (f a - x * g a) / norm (g t)) \ 0) (at_right 0)" by (simp add: inverse_eq_divide) from this[unfolded tendsto_iff, rule_format, of "e / 2"] \0 < e\ have "eventually (\t. norm (f a - x * g a) / norm (g t) < e / 2) (at_right 0)" by (auto simp: dist_real_def) ultimately show "eventually (\t. dist (f t / g t) x < e) (at_right 0)" proof eventually_elim fix t assume t[arith]: "0 < t" "t < a" "g a < g t" "0 < g t" assume ineq: "norm (1 - g a / g t) < 2" "norm (f a - x * g a) / norm (g t) < e / 2" have "\y. t < y \ y < a \ (g a - g t) * f' y = (f a - f t) * g' y" using f0 g0 t(1,2) by (intro GMVT') (force intro!: DERIV_isCont)+ then obtain y where [arith]: "t < y" "y < a" and D_eq0: "(g a - g t) * f' y = (f a - f t) * g' y" by blast from D_eq0 have D_eq: "(f t - f a) / (g t - g a) = f' y / g' y" using \g a < g t\ g'_neq_0[of y] by (auto simp add: field_simps) have *: "f t / g t - x = ((f t - f a) / (g t - g a) - x) * (1 - g a / g t) + (f a - x * g a) / g t" by (simp add: field_simps) have "norm (f t / g t - x) \ norm (((f t - f a) / (g t - g a) - x) * (1 - g a / g t)) + norm ((f a - x * g a) / g t)" unfolding * by (rule norm_triangle_ineq) also have "\ = dist (f' y / g' y) x * norm (1 - g a / g t) + norm (f a - x * g a) / norm (g t)" by (simp add: abs_mult D_eq dist_real_def) also have "\ < (e / 4) * 2 + e / 2" using ineq Df[of y] \0 < e\ by (intro add_le_less_mono mult_mono) auto finally show "dist (f t / g t) x < e" by (simp add: dist_real_def) qed qed lemma lhopital_right_at_top: "LIM x at_right x. (g::real \ real) x :> at_top \ eventually (\x. g' x \ 0) (at_right x) \ eventually (\x. DERIV f x :> f' x) (at_right x) \ eventually (\x. DERIV g x :> g' x) (at_right x) \ ((\ x. (f' x / g' x)) \ y) (at_right x) \ ((\ x. f x / g x) \ y) (at_right x)" unfolding eventually_at_right_to_0[of _ x] filterlim_at_right_to_0[of _ _ x] DERIV_shift by (rule lhopital_right_0_at_top) lemma lhopital_left_at_top: "LIM x at_left x. g x :> at_top \ eventually (\x. g' x \ 0) (at_left x) \ eventually (\x. DERIV f x :> f' x) (at_left x) \ eventually (\x. DERIV g x :> g' x) (at_left x) \ ((\ x. (f' x / g' x)) \ y) (at_left x) \ ((\ x. f x / g x) \ y) (at_left x)" for x :: real unfolding eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror by (rule lhopital_right_at_top[where f'="\x. - f' (- x)"]) (auto simp: DERIV_mirror) lemma lhopital_at_top: "LIM x at x. (g::real \ real) x :> at_top \ eventually (\x. g' x \ 0) (at x) \ eventually (\x. DERIV f x :> f' x) (at x) \ eventually (\x. DERIV g x :> g' x) (at x) \ ((\ x. (f' x / g' x)) \ y) (at x) \ ((\ x. f x / g x) \ y) (at x)" unfolding eventually_at_split filterlim_at_split by (auto intro!: lhopital_right_at_top[of g x g' f f'] lhopital_left_at_top[of g x g' f f']) lemma lhospital_at_top_at_top: fixes f g :: "real \ real" assumes g_0: "LIM x at_top. g x :> at_top" and g': "eventually (\x. g' x \ 0) at_top" and Df: "eventually (\x. DERIV f x :> f' x) at_top" and Dg: "eventually (\x. DERIV g x :> g' x) at_top" and lim: "((\ x. (f' x / g' x)) \ x) at_top" shows "((\ x. f x / g x) \ x) at_top" unfolding filterlim_at_top_to_right proof (rule lhopital_right_0_at_top) let ?F = "\x. f (inverse x)" let ?G = "\x. g (inverse x)" let ?R = "at_right (0::real)" let ?D = "\f' x. f' (inverse x) * - (inverse x ^ Suc (Suc 0))" show "LIM x ?R. ?G x :> at_top" using g_0 unfolding filterlim_at_top_to_right . show "eventually (\x. DERIV ?G x :> ?D g' x) ?R" unfolding eventually_at_right_to_top using Dg eventually_ge_at_top[where c=1] by eventually_elim (rule derivative_eq_intros DERIV_chain'[where f=inverse] | simp)+ show "eventually (\x. DERIV ?F x :> ?D f' x) ?R" unfolding eventually_at_right_to_top using Df eventually_ge_at_top[where c=1] by eventually_elim (rule derivative_eq_intros DERIV_chain'[where f=inverse] | simp)+ show "eventually (\x. ?D g' x \ 0) ?R" unfolding eventually_at_right_to_top using g' eventually_ge_at_top[where c=1] by eventually_elim auto show "((\x. ?D f' x / ?D g' x) \ x) ?R" unfolding filterlim_at_right_to_top apply (intro filterlim_cong[THEN iffD2, OF refl refl _ lim]) using eventually_ge_at_top[where c=1] by eventually_elim simp qed lemma lhopital_right_at_top_at_top: fixes f g :: "real \ real" assumes f_0: "LIM x at_right a. f x :> at_top" assumes g_0: "LIM x at_right a. g x :> at_top" and ev: "eventually (\x. DERIV f x :> f' x) (at_right a)" "eventually (\x. DERIV g x :> g' x) (at_right a)" and lim: "filterlim (\ x. (f' x / g' x)) at_top (at_right a)" shows "filterlim (\ x. f x / g x) at_top (at_right a)" proof - from lim have pos: "eventually (\x. f' x / g' x > 0) (at_right a)" unfolding filterlim_at_top_dense by blast have "((\x. g x / f x) \ 0) (at_right a)" proof (rule lhopital_right_at_top) from pos show "eventually (\x. f' x \ 0) (at_right a)" by eventually_elim auto from tendsto_inverse_0_at_top[OF lim] show "((\x. g' x / f' x) \ 0) (at_right a)" by simp qed fact+ moreover from f_0 g_0 have "eventually (\x. f x > 0) (at_right a)" "eventually (\x. g x > 0) (at_right a)" unfolding filterlim_at_top_dense by blast+ hence "eventually (\x. g x / f x > 0) (at_right a)" by eventually_elim simp ultimately have "filterlim (\x. inverse (g x / f x)) at_top (at_right a)" by (rule filterlim_inverse_at_top) thus ?thesis by simp qed lemma lhopital_right_at_top_at_bot: fixes f g :: "real \ real" assumes f_0: "LIM x at_right a. f x :> at_top" assumes g_0: "LIM x at_right a. g x :> at_bot" and ev: "eventually (\x. DERIV f x :> f' x) (at_right a)" "eventually (\x. DERIV g x :> g' x) (at_right a)" and lim: "filterlim (\ x. (f' x / g' x)) at_bot (at_right a)" shows "filterlim (\ x. f x / g x) at_bot (at_right a)" proof - from ev(2) have ev': "eventually (\x. DERIV (\x. -g x) x :> -g' x) (at_right a)" by eventually_elim (auto intro: derivative_intros) have "filterlim (\x. f x / (-g x)) at_top (at_right a)" by (rule lhopital_right_at_top_at_top[where f' = f' and g' = "\x. -g' x"]) (insert assms ev', auto simp: filterlim_uminus_at_bot) hence "filterlim (\x. -(f x / g x)) at_top (at_right a)" by simp thus ?thesis by (simp add: filterlim_uminus_at_bot) qed lemma lhopital_left_at_top_at_top: fixes f g :: "real \ real" assumes f_0: "LIM x at_left a. f x :> at_top" assumes g_0: "LIM x at_left a. g x :> at_top" and ev: "eventually (\x. DERIV f x :> f' x) (at_left a)" "eventually (\x. DERIV g x :> g' x) (at_left a)" and lim: "filterlim (\ x. (f' x / g' x)) at_top (at_left a)" shows "filterlim (\ x. f x / g x) at_top (at_left a)" by (insert assms, unfold eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror, rule lhopital_right_at_top_at_top[where f'="\x. - f' (- x)"]) (insert assms, auto simp: DERIV_mirror) lemma lhopital_left_at_top_at_bot: fixes f g :: "real \ real" assumes f_0: "LIM x at_left a. f x :> at_top" assumes g_0: "LIM x at_left a. g x :> at_bot" and ev: "eventually (\x. DERIV f x :> f' x) (at_left a)" "eventually (\x. DERIV g x :> g' x) (at_left a)" and lim: "filterlim (\ x. (f' x / g' x)) at_bot (at_left a)" shows "filterlim (\ x. f x / g x) at_bot (at_left a)" by (insert assms, unfold eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror, rule lhopital_right_at_top_at_bot[where f'="\x. - f' (- x)"]) (insert assms, auto simp: DERIV_mirror) lemma lhopital_at_top_at_top: fixes f g :: "real \ real" assumes f_0: "LIM x at a. f x :> at_top" assumes g_0: "LIM x at a. g x :> at_top" and ev: "eventually (\x. DERIV f x :> f' x) (at a)" "eventually (\x. DERIV g x :> g' x) (at a)" and lim: "filterlim (\ x. (f' x / g' x)) at_top (at a)" shows "filterlim (\ x. f x / g x) at_top (at a)" using assms unfolding eventually_at_split filterlim_at_split by (auto intro!: lhopital_right_at_top_at_top[of f a g f' g'] lhopital_left_at_top_at_top[of f a g f' g']) lemma lhopital_at_top_at_bot: fixes f g :: "real \ real" assumes f_0: "LIM x at a. f x :> at_top" assumes g_0: "LIM x at a. g x :> at_bot" and ev: "eventually (\x. DERIV f x :> f' x) (at a)" "eventually (\x. DERIV g x :> g' x) (at a)" and lim: "filterlim (\ x. (f' x / g' x)) at_bot (at a)" shows "filterlim (\ x. f x / g x) at_bot (at a)" using assms unfolding eventually_at_split filterlim_at_split by (auto intro!: lhopital_right_at_top_at_bot[of f a g f' g'] lhopital_left_at_top_at_bot[of f a g f' g']) end diff --git a/src/HOL/Limits.thy b/src/HOL/Limits.thy --- a/src/HOL/Limits.thy +++ b/src/HOL/Limits.thy @@ -1,3035 +1,3071 @@ (* Title: HOL/Limits.thy Author: Brian Huffman Author: Jacques D. Fleuriot, University of Cambridge Author: Lawrence C Paulson Author: Jeremy Avigad *) section \Limits on Real Vector Spaces\ theory Limits imports Real_Vector_Spaces begin +text \Lemmas related to shifting/scaling\ +lemma range_add [simp]: + fixes a::"'a::group_add" shows "range ((+) a) = UNIV" + by (metis add_minus_cancel surjI) + +lemma range_diff [simp]: + fixes a::"'a::group_add" shows "range ((-) a) = UNIV" + by (metis (full_types) add_minus_cancel diff_minus_eq_add surj_def) + +lemma range_mult [simp]: + fixes a::"real" shows "range ((*) a) = (if a=0 then {0} else UNIV)" + by (simp add: surj_def) (meson dvdE dvd_field_iff) + + subsection \Filter going to infinity norm\ definition at_infinity :: "'a::real_normed_vector filter" where "at_infinity = (INF r. principal {x. r \ norm x})" lemma eventually_at_infinity: "eventually P at_infinity \ (\b. \x. b \ norm x \ P x)" unfolding at_infinity_def by (subst eventually_INF_base) (auto simp: subset_eq eventually_principal intro!: exI[of _ "max a b" for a b]) corollary eventually_at_infinity_pos: "eventually p at_infinity \ (\b. 0 < b \ (\x. norm x \ b \ p x))" unfolding eventually_at_infinity by (meson le_less_trans norm_ge_zero not_le zero_less_one) lemma at_infinity_eq_at_top_bot: "(at_infinity :: real filter) = sup at_top at_bot" proof - have 1: "\\n\u. A n; \n\v. A n\ \ \b. \x. b \ \x\ \ A x" for A and u v::real by (rule_tac x="max (- v) u" in exI) (auto simp: abs_real_def) have 2: "\x. u \ \x\ \ A x \ \N. \n\N. A n" for A and u::real by (meson abs_less_iff le_cases less_le_not_le) have 3: "\x. u \ \x\ \ A x \ \N. \n\N. A n" for A and u::real by (metis (full_types) abs_ge_self abs_minus_cancel le_minus_iff order_trans) show ?thesis by (auto simp: filter_eq_iff eventually_sup eventually_at_infinity eventually_at_top_linorder eventually_at_bot_linorder intro: 1 2 3) qed lemma at_top_le_at_infinity: "at_top \ (at_infinity :: real filter)" unfolding at_infinity_eq_at_top_bot by simp lemma at_bot_le_at_infinity: "at_bot \ (at_infinity :: real filter)" unfolding at_infinity_eq_at_top_bot by simp lemma filterlim_at_top_imp_at_infinity: "filterlim f at_top F \ filterlim f at_infinity F" for f :: "_ \ real" by (rule filterlim_mono[OF _ at_top_le_at_infinity order_refl]) lemma filterlim_real_at_infinity_sequentially: "filterlim real at_infinity sequentially" by (simp add: filterlim_at_top_imp_at_infinity filterlim_real_sequentially) lemma lim_infinity_imp_sequentially: "(f \ l) at_infinity \ ((\n. f(n)) \ l) sequentially" by (simp add: filterlim_at_top_imp_at_infinity filterlim_compose filterlim_real_sequentially) subsubsection \Boundedness\ definition Bfun :: "('a \ 'b::metric_space) \ 'a filter \ bool" where Bfun_metric_def: "Bfun f F = (\y. \K>0. eventually (\x. dist (f x) y \ K) F)" abbreviation Bseq :: "(nat \ 'a::metric_space) \ bool" where "Bseq X \ Bfun X sequentially" lemma Bseq_conv_Bfun: "Bseq X \ Bfun X sequentially" .. lemma Bseq_ignore_initial_segment: "Bseq X \ Bseq (\n. X (n + k))" unfolding Bfun_metric_def by (subst eventually_sequentially_seg) lemma Bseq_offset: "Bseq (\n. X (n + k)) \ Bseq X" unfolding Bfun_metric_def by (subst (asm) eventually_sequentially_seg) lemma Bfun_def: "Bfun f F \ (\K>0. eventually (\x. norm (f x) \ K) F)" unfolding Bfun_metric_def norm_conv_dist proof safe fix y K assume K: "0 < K" and *: "eventually (\x. dist (f x) y \ K) F" moreover have "eventually (\x. dist (f x) 0 \ dist (f x) y + dist 0 y) F" by (intro always_eventually) (metis dist_commute dist_triangle) with * have "eventually (\x. dist (f x) 0 \ K + dist 0 y) F" by eventually_elim auto with \0 < K\ show "\K>0. eventually (\x. dist (f x) 0 \ K) F" by (intro exI[of _ "K + dist 0 y"] add_pos_nonneg conjI zero_le_dist) auto qed (force simp del: norm_conv_dist [symmetric]) lemma BfunI: assumes K: "eventually (\x. norm (f x) \ K) F" shows "Bfun f F" unfolding Bfun_def proof (intro exI conjI allI) show "0 < max K 1" by simp show "eventually (\x. norm (f x) \ max K 1) F" using K by (rule eventually_mono) simp qed lemma BfunE: assumes "Bfun f F" obtains B where "0 < B" and "eventually (\x. norm (f x) \ B) F" using assms unfolding Bfun_def by blast lemma Cauchy_Bseq: assumes "Cauchy X" shows "Bseq X" proof - have "\y K. 0 < K \ (\N. \n\N. dist (X n) y \ K)" if "\m n. \m \ M; n \ M\ \ dist (X m) (X n) < 1" for M by (meson order.order_iff_strict that zero_less_one) with assms show ?thesis by (force simp: Cauchy_def Bfun_metric_def eventually_sequentially) qed subsubsection \Bounded Sequences\ lemma BseqI': "(\n. norm (X n) \ K) \ Bseq X" by (intro BfunI) (auto simp: eventually_sequentially) lemma Bseq_def: "Bseq X \ (\K>0. \n. norm (X n) \ K)" unfolding Bfun_def eventually_sequentially proof safe fix N K assume "0 < K" "\n\N. norm (X n) \ K" then show "\K>0. \n. norm (X n) \ K" by (intro exI[of _ "max (Max (norm ` X ` {..N})) K"] max.strict_coboundedI2) (auto intro!: imageI not_less[where 'a=nat, THEN iffD1] Max_ge simp: le_max_iff_disj) qed auto lemma BseqE: "Bseq X \ (\K. 0 < K \ \n. norm (X n) \ K \ Q) \ Q" unfolding Bseq_def by auto lemma BseqD: "Bseq X \ \K. 0 < K \ (\n. norm (X n) \ K)" by (simp add: Bseq_def) lemma BseqI: "0 < K \ \n. norm (X n) \ K \ Bseq X" by (auto simp: Bseq_def) lemma Bseq_bdd_above: "Bseq X \ bdd_above (range X)" for X :: "nat \ real" proof (elim BseqE, intro bdd_aboveI2) fix K n assume "0 < K" "\n. norm (X n) \ K" then show "X n \ K" by (auto elim!: allE[of _ n]) qed lemma Bseq_bdd_above': "Bseq X \ bdd_above (range (\n. norm (X n)))" for X :: "nat \ 'a :: real_normed_vector" proof (elim BseqE, intro bdd_aboveI2) fix K n assume "0 < K" "\n. norm (X n) \ K" then show "norm (X n) \ K" by (auto elim!: allE[of _ n]) qed lemma Bseq_bdd_below: "Bseq X \ bdd_below (range X)" for X :: "nat \ real" proof (elim BseqE, intro bdd_belowI2) fix K n assume "0 < K" "\n. norm (X n) \ K" then show "- K \ X n" by (auto elim!: allE[of _ n]) qed lemma Bseq_eventually_mono: assumes "eventually (\n. norm (f n) \ norm (g n)) sequentially" "Bseq g" shows "Bseq f" proof - from assms(2) obtain K where "0 < K" and "eventually (\n. norm (g n) \ K) sequentially" unfolding Bfun_def by fast with assms(1) have "eventually (\n. norm (f n) \ K) sequentially" by (fast elim: eventually_elim2 order_trans) with \0 < K\ show "Bseq f" unfolding Bfun_def by fast qed lemma lemma_NBseq_def: "(\K > 0. \n. norm (X n) \ K) \ (\N. \n. norm (X n) \ real(Suc N))" proof safe fix K :: real from reals_Archimedean2 obtain n :: nat where "K < real n" .. then have "K \ real (Suc n)" by auto moreover assume "\m. norm (X m) \ K" ultimately have "\m. norm (X m) \ real (Suc n)" by (blast intro: order_trans) then show "\N. \n. norm (X n) \ real (Suc N)" .. next show "\N. \n. norm (X n) \ real (Suc N) \ \K>0. \n. norm (X n) \ K" using of_nat_0_less_iff by blast qed text \Alternative definition for \Bseq\.\ lemma Bseq_iff: "Bseq X \ (\N. \n. norm (X n) \ real(Suc N))" by (simp add: Bseq_def) (simp add: lemma_NBseq_def) lemma lemma_NBseq_def2: "(\K > 0. \n. norm (X n) \ K) = (\N. \n. norm (X n) < real(Suc N))" proof - have *: "\N. \n. norm (X n) \ 1 + real N \ \N. \n. norm (X n) < 1 + real N" by (metis add.commute le_less_trans less_add_one of_nat_Suc) then show ?thesis unfolding lemma_NBseq_def by (metis less_le_not_le not_less_iff_gr_or_eq of_nat_Suc) qed text \Yet another definition for Bseq.\ lemma Bseq_iff1a: "Bseq X \ (\N. \n. norm (X n) < real (Suc N))" by (simp add: Bseq_def lemma_NBseq_def2) subsubsection \A Few More Equivalence Theorems for Boundedness\ text \Alternative formulation for boundedness.\ lemma Bseq_iff2: "Bseq X \ (\k > 0. \x. \n. norm (X n + - x) \ k)" by (metis BseqE BseqI' add.commute add_cancel_right_left add_uminus_conv_diff norm_add_leD norm_minus_cancel norm_minus_commute) text \Alternative formulation for boundedness.\ lemma Bseq_iff3: "Bseq X \ (\k>0. \N. \n. norm (X n + - X N) \ k)" (is "?P \ ?Q") proof assume ?P then obtain K where *: "0 < K" and **: "\n. norm (X n) \ K" by (auto simp: Bseq_def) from * have "0 < K + norm (X 0)" by (rule order_less_le_trans) simp from ** have "\n. norm (X n - X 0) \ K + norm (X 0)" by (auto intro: order_trans norm_triangle_ineq4) then have "\n. norm (X n + - X 0) \ K + norm (X 0)" by simp with \0 < K + norm (X 0)\ show ?Q by blast next assume ?Q then show ?P by (auto simp: Bseq_iff2) qed subsubsection \Upper Bounds and Lubs of Bounded Sequences\ lemma Bseq_minus_iff: "Bseq (\n. - (X n) :: 'a::real_normed_vector) \ Bseq X" by (simp add: Bseq_def) lemma Bseq_add: fixes f :: "nat \ 'a::real_normed_vector" assumes "Bseq f" shows "Bseq (\x. f x + c)" proof - from assms obtain K where K: "\x. norm (f x) \ K" unfolding Bseq_def by blast { fix x :: nat have "norm (f x + c) \ norm (f x) + norm c" by (rule norm_triangle_ineq) also have "norm (f x) \ K" by (rule K) finally have "norm (f x + c) \ K + norm c" by simp } then show ?thesis by (rule BseqI') qed lemma Bseq_add_iff: "Bseq (\x. f x + c) \ Bseq f" for f :: "nat \ 'a::real_normed_vector" using Bseq_add[of f c] Bseq_add[of "\x. f x + c" "-c"] by auto lemma Bseq_mult: fixes f g :: "nat \ 'a::real_normed_field" assumes "Bseq f" and "Bseq g" shows "Bseq (\x. f x * g x)" proof - from assms obtain K1 K2 where K: "norm (f x) \ K1" "K1 > 0" "norm (g x) \ K2" "K2 > 0" for x unfolding Bseq_def by blast then have "norm (f x * g x) \ K1 * K2" for x by (auto simp: norm_mult intro!: mult_mono) then show ?thesis by (rule BseqI') qed lemma Bfun_const [simp]: "Bfun (\_. c) F" unfolding Bfun_metric_def by (auto intro!: exI[of _ c] exI[of _ "1::real"]) lemma Bseq_cmult_iff: fixes c :: "'a::real_normed_field" assumes "c \ 0" shows "Bseq (\x. c * f x) \ Bseq f" proof assume "Bseq (\x. c * f x)" with Bfun_const have "Bseq (\x. inverse c * (c * f x))" by (rule Bseq_mult) with \c \ 0\ show "Bseq f" by (simp add: field_split_simps) qed (intro Bseq_mult Bfun_const) lemma Bseq_subseq: "Bseq f \ Bseq (\x. f (g x))" for f :: "nat \ 'a::real_normed_vector" unfolding Bseq_def by auto lemma Bseq_Suc_iff: "Bseq (\n. f (Suc n)) \ Bseq f" for f :: "nat \ 'a::real_normed_vector" using Bseq_offset[of f 1] by (auto intro: Bseq_subseq) lemma increasing_Bseq_subseq_iff: assumes "\x y. x \ y \ norm (f x :: 'a::real_normed_vector) \ norm (f y)" "strict_mono g" shows "Bseq (\x. f (g x)) \ Bseq f" proof assume "Bseq (\x. f (g x))" then obtain K where K: "\x. norm (f (g x)) \ K" unfolding Bseq_def by auto { fix x :: nat from filterlim_subseq[OF assms(2)] obtain y where "g y \ x" by (auto simp: filterlim_at_top eventually_at_top_linorder) then have "norm (f x) \ norm (f (g y))" using assms(1) by blast also have "norm (f (g y)) \ K" by (rule K) finally have "norm (f x) \ K" . } then show "Bseq f" by (rule BseqI') qed (use Bseq_subseq[of f g] in simp_all) lemma nonneg_incseq_Bseq_subseq_iff: fixes f :: "nat \ real" and g :: "nat \ nat" assumes "\x. f x \ 0" "incseq f" "strict_mono g" shows "Bseq (\x. f (g x)) \ Bseq f" using assms by (intro increasing_Bseq_subseq_iff) (auto simp: incseq_def) lemma Bseq_eq_bounded: "range f \ {a..b} \ Bseq f" for a b :: real proof (rule BseqI'[where K="max (norm a) (norm b)"]) fix n assume "range f \ {a..b}" then have "f n \ {a..b}" by blast then show "norm (f n) \ max (norm a) (norm b)" by auto qed lemma incseq_bounded: "incseq X \ \i. X i \ B \ Bseq X" for B :: real by (intro Bseq_eq_bounded[of X "X 0" B]) (auto simp: incseq_def) lemma decseq_bounded: "decseq X \ \i. B \ X i \ Bseq X" for B :: real by (intro Bseq_eq_bounded[of X B "X 0"]) (auto simp: decseq_def) subsubsection\<^marker>\tag unimportant\ \Polynomal function extremal theorem, from HOL Light\ lemma polyfun_extremal_lemma: fixes c :: "nat \ 'a::real_normed_div_algebra" assumes "0 < e" shows "\M. \z. M \ norm(z) \ norm (\i\n. c(i) * z^i) \ e * norm(z) ^ (Suc n)" proof (induct n) case 0 with assms show ?case apply (rule_tac x="norm (c 0) / e" in exI) apply (auto simp: field_simps) done next case (Suc n) obtain M where M: "\z. M \ norm z \ norm (\i\n. c i * z^i) \ e * norm z ^ Suc n" using Suc assms by blast show ?case proof (rule exI [where x= "max M (1 + norm(c(Suc n)) / e)"], clarsimp simp del: power_Suc) fix z::'a assume z1: "M \ norm z" and "1 + norm (c (Suc n)) / e \ norm z" then have z2: "e + norm (c (Suc n)) \ e * norm z" using assms by (simp add: field_simps) have "norm (\i\n. c i * z^i) \ e * norm z ^ Suc n" using M [OF z1] by simp then have "norm (\i\n. c i * z^i) + norm (c (Suc n) * z ^ Suc n) \ e * norm z ^ Suc n + norm (c (Suc n) * z ^ Suc n)" by simp then have "norm ((\i\n. c i * z^i) + c (Suc n) * z ^ Suc n) \ e * norm z ^ Suc n + norm (c (Suc n) * z ^ Suc n)" by (blast intro: norm_triangle_le elim: ) also have "... \ (e + norm (c (Suc n))) * norm z ^ Suc n" by (simp add: norm_power norm_mult algebra_simps) also have "... \ (e * norm z) * norm z ^ Suc n" by (metis z2 mult.commute mult_left_mono norm_ge_zero norm_power) finally show "norm ((\i\n. c i * z^i) + c (Suc n) * z ^ Suc n) \ e * norm z ^ Suc (Suc n)" by simp qed qed lemma polyfun_extremal: (*COMPLEX_POLYFUN_EXTREMAL in HOL Light*) fixes c :: "nat \ 'a::real_normed_div_algebra" assumes k: "c k \ 0" "1\k" and kn: "k\n" shows "eventually (\z. norm (\i\n. c(i) * z^i) \ B) at_infinity" using kn proof (induction n) case 0 then show ?case using k by simp next case (Suc m) show ?case proof (cases "c (Suc m) = 0") case True then show ?thesis using Suc k by auto (metis antisym_conv less_eq_Suc_le not_le) next case False then obtain M where M: "\z. M \ norm z \ norm (\i\m. c i * z^i) \ norm (c (Suc m)) / 2 * norm z ^ Suc m" using polyfun_extremal_lemma [of "norm(c (Suc m)) / 2" c m] Suc by auto have "\b. \z. b \ norm z \ B \ norm (\i\Suc m. c i * z^i)" proof (rule exI [where x="max M (max 1 (\B\ / (norm(c (Suc m)) / 2)))"], clarsimp simp del: power_Suc) fix z::'a assume z1: "M \ norm z" "1 \ norm z" and "\B\ * 2 / norm (c (Suc m)) \ norm z" then have z2: "\B\ \ norm (c (Suc m)) * norm z / 2" using False by (simp add: field_simps) have nz: "norm z \ norm z ^ Suc m" by (metis \1 \ norm z\ One_nat_def less_eq_Suc_le power_increasing power_one_right zero_less_Suc) have *: "\y x. norm (c (Suc m)) * norm z / 2 \ norm y - norm x \ B \ norm (x + y)" by (metis abs_le_iff add.commute norm_diff_ineq order_trans z2) have "norm z * norm (c (Suc m)) + 2 * norm (\i\m. c i * z^i) \ norm (c (Suc m)) * norm z + norm (c (Suc m)) * norm z ^ Suc m" using M [of z] Suc z1 by auto also have "... \ 2 * (norm (c (Suc m)) * norm z ^ Suc m)" using nz by (simp add: mult_mono del: power_Suc) finally show "B \ norm ((\i\m. c i * z^i) + c (Suc m) * z ^ Suc m)" using Suc.IH apply (auto simp: eventually_at_infinity) apply (rule *) apply (simp add: field_simps norm_mult norm_power) done qed then show ?thesis by (simp add: eventually_at_infinity) qed qed subsection \Convergence to Zero\ definition Zfun :: "('a \ 'b::real_normed_vector) \ 'a filter \ bool" where "Zfun f F = (\r>0. eventually (\x. norm (f x) < r) F)" lemma ZfunI: "(\r. 0 < r \ eventually (\x. norm (f x) < r) F) \ Zfun f F" by (simp add: Zfun_def) lemma ZfunD: "Zfun f F \ 0 < r \ eventually (\x. norm (f x) < r) F" by (simp add: Zfun_def) lemma Zfun_ssubst: "eventually (\x. f x = g x) F \ Zfun g F \ Zfun f F" unfolding Zfun_def by (auto elim!: eventually_rev_mp) lemma Zfun_zero: "Zfun (\x. 0) F" unfolding Zfun_def by simp lemma Zfun_norm_iff: "Zfun (\x. norm (f x)) F = Zfun (\x. f x) F" unfolding Zfun_def by simp lemma Zfun_imp_Zfun: assumes f: "Zfun f F" and g: "eventually (\x. norm (g x) \ norm (f x) * K) F" shows "Zfun (\x. g x) F" proof (cases "0 < K") case K: True show ?thesis proof (rule ZfunI) fix r :: real assume "0 < r" then have "0 < r / K" using K by simp then have "eventually (\x. norm (f x) < r / K) F" using ZfunD [OF f] by blast with g show "eventually (\x. norm (g x) < r) F" proof eventually_elim case (elim x) then have "norm (f x) * K < r" by (simp add: pos_less_divide_eq K) then show ?case by (simp add: order_le_less_trans [OF elim(1)]) qed qed next case False then have K: "K \ 0" by (simp only: not_less) show ?thesis proof (rule ZfunI) fix r :: real assume "0 < r" from g show "eventually (\x. norm (g x) < r) F" proof eventually_elim case (elim x) also have "norm (f x) * K \ norm (f x) * 0" using K norm_ge_zero by (rule mult_left_mono) finally show ?case using \0 < r\ by simp qed qed qed lemma Zfun_le: "Zfun g F \ \x. norm (f x) \ norm (g x) \ Zfun f F" by (erule Zfun_imp_Zfun [where K = 1]) simp lemma Zfun_add: assumes f: "Zfun f F" and g: "Zfun g F" shows "Zfun (\x. f x + g x) F" proof (rule ZfunI) fix r :: real assume "0 < r" then have r: "0 < r / 2" by simp have "eventually (\x. norm (f x) < r/2) F" using f r by (rule ZfunD) moreover have "eventually (\x. norm (g x) < r/2) F" using g r by (rule ZfunD) ultimately show "eventually (\x. norm (f x + g x) < r) F" proof eventually_elim case (elim x) have "norm (f x + g x) \ norm (f x) + norm (g x)" by (rule norm_triangle_ineq) also have "\ < r/2 + r/2" using elim by (rule add_strict_mono) finally show ?case by simp qed qed lemma Zfun_minus: "Zfun f F \ Zfun (\x. - f x) F" unfolding Zfun_def by simp lemma Zfun_diff: "Zfun f F \ Zfun g F \ Zfun (\x. f x - g x) F" using Zfun_add [of f F "\x. - g x"] by (simp add: Zfun_minus) lemma (in bounded_linear) Zfun: assumes g: "Zfun g F" shows "Zfun (\x. f (g x)) F" proof - obtain K where "norm (f x) \ norm x * K" for x using bounded by blast then have "eventually (\x. norm (f (g x)) \ norm (g x) * K) F" by simp with g show ?thesis by (rule Zfun_imp_Zfun) qed lemma (in bounded_bilinear) Zfun: assumes f: "Zfun f F" and g: "Zfun g F" shows "Zfun (\x. f x ** g x) F" proof (rule ZfunI) fix r :: real assume r: "0 < r" obtain K where K: "0 < K" and norm_le: "norm (x ** y) \ norm x * norm y * K" for x y using pos_bounded by blast from K have K': "0 < inverse K" by (rule positive_imp_inverse_positive) have "eventually (\x. norm (f x) < r) F" using f r by (rule ZfunD) moreover have "eventually (\x. norm (g x) < inverse K) F" using g K' by (rule ZfunD) ultimately show "eventually (\x. norm (f x ** g x) < r) F" proof eventually_elim case (elim x) have "norm (f x ** g x) \ norm (f x) * norm (g x) * K" by (rule norm_le) also have "norm (f x) * norm (g x) * K < r * inverse K * K" by (intro mult_strict_right_mono mult_strict_mono' norm_ge_zero elim K) also from K have "r * inverse K * K = r" by simp finally show ?case . qed qed lemma (in bounded_bilinear) Zfun_left: "Zfun f F \ Zfun (\x. f x ** a) F" by (rule bounded_linear_left [THEN bounded_linear.Zfun]) lemma (in bounded_bilinear) Zfun_right: "Zfun f F \ Zfun (\x. a ** f x) F" by (rule bounded_linear_right [THEN bounded_linear.Zfun]) lemmas Zfun_mult = bounded_bilinear.Zfun [OF bounded_bilinear_mult] lemmas Zfun_mult_right = bounded_bilinear.Zfun_right [OF bounded_bilinear_mult] lemmas Zfun_mult_left = bounded_bilinear.Zfun_left [OF bounded_bilinear_mult] lemma tendsto_Zfun_iff: "(f \ a) F = Zfun (\x. f x - a) F" by (simp only: tendsto_iff Zfun_def dist_norm) lemma tendsto_0_le: "(f \ 0) F \ eventually (\x. norm (g x) \ norm (f x) * K) F \ (g \ 0) F" by (simp add: Zfun_imp_Zfun tendsto_Zfun_iff) subsubsection \Distance and norms\ lemma tendsto_dist [tendsto_intros]: fixes l m :: "'a::metric_space" assumes f: "(f \ l) F" and g: "(g \ m) F" shows "((\x. dist (f x) (g x)) \ dist l m) F" proof (rule tendstoI) fix e :: real assume "0 < e" then have e2: "0 < e/2" by simp from tendstoD [OF f e2] tendstoD [OF g e2] show "eventually (\x. dist (dist (f x) (g x)) (dist l m) < e) F" proof (eventually_elim) case (elim x) then show "dist (dist (f x) (g x)) (dist l m) < e" unfolding dist_real_def using dist_triangle2 [of "f x" "g x" "l"] and dist_triangle2 [of "g x" "l" "m"] and dist_triangle3 [of "l" "m" "f x"] and dist_triangle [of "f x" "m" "g x"] by arith qed qed lemma continuous_dist[continuous_intros]: fixes f g :: "_ \ 'a :: metric_space" shows "continuous F f \ continuous F g \ continuous F (\x. dist (f x) (g x))" unfolding continuous_def by (rule tendsto_dist) lemma continuous_on_dist[continuous_intros]: fixes f g :: "_ \ 'a :: metric_space" shows "continuous_on s f \ continuous_on s g \ continuous_on s (\x. dist (f x) (g x))" unfolding continuous_on_def by (auto intro: tendsto_dist) lemma continuous_at_dist: "isCont (dist a) b" using continuous_on_dist [OF continuous_on_const continuous_on_id] continuous_on_eq_continuous_within by blast lemma tendsto_norm [tendsto_intros]: "(f \ a) F \ ((\x. norm (f x)) \ norm a) F" unfolding norm_conv_dist by (intro tendsto_intros) lemma continuous_norm [continuous_intros]: "continuous F f \ continuous F (\x. norm (f x))" unfolding continuous_def by (rule tendsto_norm) lemma continuous_on_norm [continuous_intros]: "continuous_on s f \ continuous_on s (\x. norm (f x))" unfolding continuous_on_def by (auto intro: tendsto_norm) lemma continuous_on_norm_id [continuous_intros]: "continuous_on S norm" by (intro continuous_on_id continuous_on_norm) lemma tendsto_norm_zero: "(f \ 0) F \ ((\x. norm (f x)) \ 0) F" by (drule tendsto_norm) simp lemma tendsto_norm_zero_cancel: "((\x. norm (f x)) \ 0) F \ (f \ 0) F" unfolding tendsto_iff dist_norm by simp lemma tendsto_norm_zero_iff: "((\x. norm (f x)) \ 0) F \ (f \ 0) F" unfolding tendsto_iff dist_norm by simp lemma tendsto_rabs [tendsto_intros]: "(f \ l) F \ ((\x. \f x\) \ \l\) F" for l :: real by (fold real_norm_def) (rule tendsto_norm) lemma continuous_rabs [continuous_intros]: "continuous F f \ continuous F (\x. \f x :: real\)" unfolding real_norm_def[symmetric] by (rule continuous_norm) lemma continuous_on_rabs [continuous_intros]: "continuous_on s f \ continuous_on s (\x. \f x :: real\)" unfolding real_norm_def[symmetric] by (rule continuous_on_norm) lemma tendsto_rabs_zero: "(f \ (0::real)) F \ ((\x. \f x\) \ 0) F" by (fold real_norm_def) (rule tendsto_norm_zero) lemma tendsto_rabs_zero_cancel: "((\x. \f x\) \ (0::real)) F \ (f \ 0) F" by (fold real_norm_def) (rule tendsto_norm_zero_cancel) lemma tendsto_rabs_zero_iff: "((\x. \f x\) \ (0::real)) F \ (f \ 0) F" by (fold real_norm_def) (rule tendsto_norm_zero_iff) subsection \Topological Monoid\ class topological_monoid_add = topological_space + monoid_add + assumes tendsto_add_Pair: "LIM x (nhds a \\<^sub>F nhds b). fst x + snd x :> nhds (a + b)" class topological_comm_monoid_add = topological_monoid_add + comm_monoid_add lemma tendsto_add [tendsto_intros]: fixes a b :: "'a::topological_monoid_add" shows "(f \ a) F \ (g \ b) F \ ((\x. f x + g x) \ a + b) F" using filterlim_compose[OF tendsto_add_Pair, of "\x. (f x, g x)" a b F] by (simp add: nhds_prod[symmetric] tendsto_Pair) lemma continuous_add [continuous_intros]: fixes f g :: "_ \ 'b::topological_monoid_add" shows "continuous F f \ continuous F g \ continuous F (\x. f x + g x)" unfolding continuous_def by (rule tendsto_add) lemma continuous_on_add [continuous_intros]: fixes f g :: "_ \ 'b::topological_monoid_add" shows "continuous_on s f \ continuous_on s g \ continuous_on s (\x. f x + g x)" unfolding continuous_on_def by (auto intro: tendsto_add) lemma tendsto_add_zero: fixes f g :: "_ \ 'b::topological_monoid_add" shows "(f \ 0) F \ (g \ 0) F \ ((\x. f x + g x) \ 0) F" by (drule (1) tendsto_add) simp lemma tendsto_sum [tendsto_intros]: fixes f :: "'a \ 'b \ 'c::topological_comm_monoid_add" shows "(\i. i \ I \ (f i \ a i) F) \ ((\x. \i\I. f i x) \ (\i\I. a i)) F" by (induct I rule: infinite_finite_induct) (simp_all add: tendsto_add) lemma tendsto_null_sum: fixes f :: "'a \ 'b \ 'c::topological_comm_monoid_add" assumes "\i. i \ I \ ((\x. f x i) \ 0) F" shows "((\i. sum (f i) I) \ 0) F" using tendsto_sum [of I "\x y. f y x" "\x. 0"] assms by simp lemma continuous_sum [continuous_intros]: fixes f :: "'a \ 'b::t2_space \ 'c::topological_comm_monoid_add" shows "(\i. i \ I \ continuous F (f i)) \ continuous F (\x. \i\I. f i x)" unfolding continuous_def by (rule tendsto_sum) lemma continuous_on_sum [continuous_intros]: fixes f :: "'a \ 'b::topological_space \ 'c::topological_comm_monoid_add" shows "(\i. i \ I \ continuous_on S (f i)) \ continuous_on S (\x. \i\I. f i x)" unfolding continuous_on_def by (auto intro: tendsto_sum) instance nat :: topological_comm_monoid_add by standard (simp add: nhds_discrete principal_prod_principal filterlim_principal eventually_principal) instance int :: topological_comm_monoid_add by standard (simp add: nhds_discrete principal_prod_principal filterlim_principal eventually_principal) subsubsection \Topological group\ class topological_group_add = topological_monoid_add + group_add + assumes tendsto_uminus_nhds: "(uminus \ - a) (nhds a)" begin lemma tendsto_minus [tendsto_intros]: "(f \ a) F \ ((\x. - f x) \ - a) F" by (rule filterlim_compose[OF tendsto_uminus_nhds]) end class topological_ab_group_add = topological_group_add + ab_group_add instance topological_ab_group_add < topological_comm_monoid_add .. lemma continuous_minus [continuous_intros]: "continuous F f \ continuous F (\x. - f x)" for f :: "'a::t2_space \ 'b::topological_group_add" unfolding continuous_def by (rule tendsto_minus) lemma continuous_on_minus [continuous_intros]: "continuous_on s f \ continuous_on s (\x. - f x)" for f :: "_ \ 'b::topological_group_add" unfolding continuous_on_def by (auto intro: tendsto_minus) lemma tendsto_minus_cancel: "((\x. - f x) \ - a) F \ (f \ a) F" for a :: "'a::topological_group_add" by (drule tendsto_minus) simp lemma tendsto_minus_cancel_left: "(f \ - (y::_::topological_group_add)) F \ ((\x. - f x) \ y) F" using tendsto_minus_cancel[of f "- y" F] tendsto_minus[of f "- y" F] by auto lemma tendsto_diff [tendsto_intros]: fixes a b :: "'a::topological_group_add" shows "(f \ a) F \ (g \ b) F \ ((\x. f x - g x) \ a - b) F" using tendsto_add [of f a F "\x. - g x" "- b"] by (simp add: tendsto_minus) lemma continuous_diff [continuous_intros]: fixes f g :: "'a::t2_space \ 'b::topological_group_add" shows "continuous F f \ continuous F g \ continuous F (\x. f x - g x)" unfolding continuous_def by (rule tendsto_diff) lemma continuous_on_diff [continuous_intros]: fixes f g :: "_ \ 'b::topological_group_add" shows "continuous_on s f \ continuous_on s g \ continuous_on s (\x. f x - g x)" unfolding continuous_on_def by (auto intro: tendsto_diff) lemma continuous_on_op_minus: "continuous_on (s::'a::topological_group_add set) ((-) x)" by (rule continuous_intros | simp)+ instance real_normed_vector < topological_ab_group_add proof fix a b :: 'a show "((\x. fst x + snd x) \ a + b) (nhds a \\<^sub>F nhds b)" unfolding tendsto_Zfun_iff add_diff_add using tendsto_fst[OF filterlim_ident, of "(a,b)"] tendsto_snd[OF filterlim_ident, of "(a,b)"] by (intro Zfun_add) (auto simp: tendsto_Zfun_iff[symmetric] nhds_prod[symmetric] intro!: tendsto_fst) show "(uminus \ - a) (nhds a)" unfolding tendsto_Zfun_iff minus_diff_minus using filterlim_ident[of "nhds a"] by (intro Zfun_minus) (simp add: tendsto_Zfun_iff) qed lemmas real_tendsto_sandwich = tendsto_sandwich[where 'a=real] subsubsection \Linear operators and multiplication\ lemma linear_times [simp]: "linear (\x. c * x)" for c :: "'a::real_algebra" by (auto simp: linearI distrib_left) lemma (in bounded_linear) tendsto: "(g \ a) F \ ((\x. f (g x)) \ f a) F" by (simp only: tendsto_Zfun_iff diff [symmetric] Zfun) lemma (in bounded_linear) continuous: "continuous F g \ continuous F (\x. f (g x))" using tendsto[of g _ F] by (auto simp: continuous_def) lemma (in bounded_linear) continuous_on: "continuous_on s g \ continuous_on s (\x. f (g x))" using tendsto[of g] by (auto simp: continuous_on_def) lemma (in bounded_linear) tendsto_zero: "(g \ 0) F \ ((\x. f (g x)) \ 0) F" by (drule tendsto) (simp only: zero) lemma (in bounded_bilinear) tendsto: "(f \ a) F \ (g \ b) F \ ((\x. f x ** g x) \ a ** b) F" by (simp only: tendsto_Zfun_iff prod_diff_prod Zfun_add Zfun Zfun_left Zfun_right) lemma (in bounded_bilinear) continuous: "continuous F f \ continuous F g \ continuous F (\x. f x ** g x)" using tendsto[of f _ F g] by (auto simp: continuous_def) lemma (in bounded_bilinear) continuous_on: "continuous_on s f \ continuous_on s g \ continuous_on s (\x. f x ** g x)" using tendsto[of f _ _ g] by (auto simp: continuous_on_def) lemma (in bounded_bilinear) tendsto_zero: assumes f: "(f \ 0) F" and g: "(g \ 0) F" shows "((\x. f x ** g x) \ 0) F" using tendsto [OF f g] by (simp add: zero_left) lemma (in bounded_bilinear) tendsto_left_zero: "(f \ 0) F \ ((\x. f x ** c) \ 0) F" by (rule bounded_linear.tendsto_zero [OF bounded_linear_left]) lemma (in bounded_bilinear) tendsto_right_zero: "(f \ 0) F \ ((\x. c ** f x) \ 0) F" by (rule bounded_linear.tendsto_zero [OF bounded_linear_right]) lemmas tendsto_of_real [tendsto_intros] = bounded_linear.tendsto [OF bounded_linear_of_real] lemmas tendsto_scaleR [tendsto_intros] = bounded_bilinear.tendsto [OF bounded_bilinear_scaleR] text\Analogous type class for multiplication\ class topological_semigroup_mult = topological_space + semigroup_mult + assumes tendsto_mult_Pair: "LIM x (nhds a \\<^sub>F nhds b). fst x * snd x :> nhds (a * b)" instance real_normed_algebra < topological_semigroup_mult proof fix a b :: 'a show "((\x. fst x * snd x) \ a * b) (nhds a \\<^sub>F nhds b)" unfolding nhds_prod[symmetric] using tendsto_fst[OF filterlim_ident, of "(a,b)"] tendsto_snd[OF filterlim_ident, of "(a,b)"] by (simp add: bounded_bilinear.tendsto [OF bounded_bilinear_mult]) qed lemma tendsto_mult [tendsto_intros]: fixes a b :: "'a::topological_semigroup_mult" shows "(f \ a) F \ (g \ b) F \ ((\x. f x * g x) \ a * b) F" using filterlim_compose[OF tendsto_mult_Pair, of "\x. (f x, g x)" a b F] by (simp add: nhds_prod[symmetric] tendsto_Pair) lemma tendsto_mult_left: "(f \ l) F \ ((\x. c * (f x)) \ c * l) F" for c :: "'a::topological_semigroup_mult" by (rule tendsto_mult [OF tendsto_const]) lemma tendsto_mult_right: "(f \ l) F \ ((\x. (f x) * c) \ l * c) F" for c :: "'a::topological_semigroup_mult" by (rule tendsto_mult [OF _ tendsto_const]) lemma tendsto_mult_left_iff [simp]: "c \ 0 \ tendsto(\x. c * f x) (c * l) F \ tendsto f l F" for c :: "'a::{topological_semigroup_mult,field}" by (auto simp: tendsto_mult_left dest: tendsto_mult_left [where c = "1/c"]) lemma tendsto_mult_right_iff [simp]: "c \ 0 \ tendsto(\x. f x * c) (l * c) F \ tendsto f l F" for c :: "'a::{topological_semigroup_mult,field}" by (auto simp: tendsto_mult_right dest: tendsto_mult_left [where c = "1/c"]) lemma tendsto_zero_mult_left_iff [simp]: fixes c::"'a::{topological_semigroup_mult,field}" assumes "c \ 0" shows "(\n. c * a n)\ 0 \ a \ 0" using assms tendsto_mult_left tendsto_mult_left_iff by fastforce lemma tendsto_zero_mult_right_iff [simp]: fixes c::"'a::{topological_semigroup_mult,field}" assumes "c \ 0" shows "(\n. a n * c)\ 0 \ a \ 0" using assms tendsto_mult_right tendsto_mult_right_iff by fastforce lemma tendsto_zero_divide_iff [simp]: fixes c::"'a::{topological_semigroup_mult,field}" assumes "c \ 0" shows "(\n. a n / c)\ 0 \ a \ 0" using tendsto_zero_mult_right_iff [of "1/c" a] assms by (simp add: field_simps) lemma lim_const_over_n [tendsto_intros]: fixes a :: "'a::real_normed_field" shows "(\n. a / of_nat n) \ 0" using tendsto_mult [OF tendsto_const [of a] lim_1_over_n] by simp lemmas continuous_of_real [continuous_intros] = bounded_linear.continuous [OF bounded_linear_of_real] lemmas continuous_scaleR [continuous_intros] = bounded_bilinear.continuous [OF bounded_bilinear_scaleR] lemmas continuous_mult [continuous_intros] = bounded_bilinear.continuous [OF bounded_bilinear_mult] lemmas continuous_on_of_real [continuous_intros] = bounded_linear.continuous_on [OF bounded_linear_of_real] lemmas continuous_on_scaleR [continuous_intros] = bounded_bilinear.continuous_on [OF bounded_bilinear_scaleR] lemmas continuous_on_mult [continuous_intros] = bounded_bilinear.continuous_on [OF bounded_bilinear_mult] lemmas tendsto_mult_zero = bounded_bilinear.tendsto_zero [OF bounded_bilinear_mult] lemmas tendsto_mult_left_zero = bounded_bilinear.tendsto_left_zero [OF bounded_bilinear_mult] lemmas tendsto_mult_right_zero = bounded_bilinear.tendsto_right_zero [OF bounded_bilinear_mult] lemma continuous_mult_left: fixes c::"'a::real_normed_algebra" shows "continuous F f \ continuous F (\x. c * f x)" by (rule continuous_mult [OF continuous_const]) lemma continuous_mult_right: fixes c::"'a::real_normed_algebra" shows "continuous F f \ continuous F (\x. f x * c)" by (rule continuous_mult [OF _ continuous_const]) lemma continuous_on_mult_left: fixes c::"'a::real_normed_algebra" shows "continuous_on s f \ continuous_on s (\x. c * f x)" by (rule continuous_on_mult [OF continuous_on_const]) lemma continuous_on_mult_right: fixes c::"'a::real_normed_algebra" shows "continuous_on s f \ continuous_on s (\x. f x * c)" by (rule continuous_on_mult [OF _ continuous_on_const]) lemma continuous_on_mult_const [simp]: fixes c::"'a::real_normed_algebra" shows "continuous_on s ((*) c)" by (intro continuous_on_mult_left continuous_on_id) lemma tendsto_divide_zero: fixes c :: "'a::real_normed_field" shows "(f \ 0) F \ ((\x. f x / c) \ 0) F" by (cases "c=0") (simp_all add: divide_inverse tendsto_mult_left_zero) lemma tendsto_power [tendsto_intros]: "(f \ a) F \ ((\x. f x ^ n) \ a ^ n) F" for f :: "'a \ 'b::{power,real_normed_algebra}" by (induct n) (simp_all add: tendsto_mult) lemma tendsto_null_power: "\(f \ 0) F; 0 < n\ \ ((\x. f x ^ n) \ 0) F" for f :: "'a \ 'b::{power,real_normed_algebra_1}" using tendsto_power [of f 0 F n] by (simp add: power_0_left) lemma continuous_power [continuous_intros]: "continuous F f \ continuous F (\x. (f x)^n)" for f :: "'a::t2_space \ 'b::{power,real_normed_algebra}" unfolding continuous_def by (rule tendsto_power) lemma continuous_on_power [continuous_intros]: fixes f :: "_ \ 'b::{power,real_normed_algebra}" shows "continuous_on s f \ continuous_on s (\x. (f x)^n)" unfolding continuous_on_def by (auto intro: tendsto_power) lemma tendsto_prod [tendsto_intros]: fixes f :: "'a \ 'b \ 'c::{real_normed_algebra,comm_ring_1}" shows "(\i. i \ S \ (f i \ L i) F) \ ((\x. \i\S. f i x) \ (\i\S. L i)) F" by (induct S rule: infinite_finite_induct) (simp_all add: tendsto_mult) lemma continuous_prod [continuous_intros]: fixes f :: "'a \ 'b::t2_space \ 'c::{real_normed_algebra,comm_ring_1}" shows "(\i. i \ S \ continuous F (f i)) \ continuous F (\x. \i\S. f i x)" unfolding continuous_def by (rule tendsto_prod) lemma continuous_on_prod [continuous_intros]: fixes f :: "'a \ _ \ 'c::{real_normed_algebra,comm_ring_1}" shows "(\i. i \ S \ continuous_on s (f i)) \ continuous_on s (\x. \i\S. f i x)" unfolding continuous_on_def by (auto intro: tendsto_prod) lemma tendsto_of_real_iff: "((\x. of_real (f x) :: 'a::real_normed_div_algebra) \ of_real c) F \ (f \ c) F" unfolding tendsto_iff by simp lemma tendsto_add_const_iff: "((\x. c + f x :: 'a::real_normed_vector) \ c + d) F \ (f \ d) F" using tendsto_add[OF tendsto_const[of c], of f d] and tendsto_add[OF tendsto_const[of "-c"], of "\x. c + f x" "c + d"] by auto class topological_monoid_mult = topological_semigroup_mult + monoid_mult class topological_comm_monoid_mult = topological_monoid_mult + comm_monoid_mult lemma tendsto_power_strong [tendsto_intros]: fixes f :: "_ \ 'b :: topological_monoid_mult" assumes "(f \ a) F" "(g \ b) F" shows "((\x. f x ^ g x) \ a ^ b) F" proof - have "((\x. f x ^ b) \ a ^ b) F" by (induction b) (auto intro: tendsto_intros assms) also from assms(2) have "eventually (\x. g x = b) F" by (simp add: nhds_discrete filterlim_principal) hence "eventually (\x. f x ^ b = f x ^ g x) F" by eventually_elim simp hence "((\x. f x ^ b) \ a ^ b) F \ ((\x. f x ^ g x) \ a ^ b) F" by (intro filterlim_cong refl) finally show ?thesis . qed lemma continuous_mult' [continuous_intros]: fixes f g :: "_ \ 'b::topological_semigroup_mult" shows "continuous F f \ continuous F g \ continuous F (\x. f x * g x)" unfolding continuous_def by (rule tendsto_mult) lemma continuous_power' [continuous_intros]: fixes f :: "_ \ 'b::topological_monoid_mult" shows "continuous F f \ continuous F g \ continuous F (\x. f x ^ g x)" unfolding continuous_def by (rule tendsto_power_strong) auto lemma continuous_on_mult' [continuous_intros]: fixes f g :: "_ \ 'b::topological_semigroup_mult" shows "continuous_on A f \ continuous_on A g \ continuous_on A (\x. f x * g x)" unfolding continuous_on_def by (auto intro: tendsto_mult) lemma continuous_on_power' [continuous_intros]: fixes f :: "_ \ 'b::topological_monoid_mult" shows "continuous_on A f \ continuous_on A g \ continuous_on A (\x. f x ^ g x)" unfolding continuous_on_def by (auto intro: tendsto_power_strong) lemma tendsto_mult_one: fixes f g :: "_ \ 'b::topological_monoid_mult" shows "(f \ 1) F \ (g \ 1) F \ ((\x. f x * g x) \ 1) F" by (drule (1) tendsto_mult) simp lemma tendsto_prod' [tendsto_intros]: fixes f :: "'a \ 'b \ 'c::topological_comm_monoid_mult" shows "(\i. i \ I \ (f i \ a i) F) \ ((\x. \i\I. f i x) \ (\i\I. a i)) F" by (induct I rule: infinite_finite_induct) (simp_all add: tendsto_mult) lemma tendsto_one_prod': fixes f :: "'a \ 'b \ 'c::topological_comm_monoid_mult" assumes "\i. i \ I \ ((\x. f x i) \ 1) F" shows "((\i. prod (f i) I) \ 1) F" using tendsto_prod' [of I "\x y. f y x" "\x. 1"] assms by simp lemma continuous_prod' [continuous_intros]: fixes f :: "'a \ 'b::t2_space \ 'c::topological_comm_monoid_mult" shows "(\i. i \ I \ continuous F (f i)) \ continuous F (\x. \i\I. f i x)" unfolding continuous_def by (rule tendsto_prod') lemma continuous_on_prod' [continuous_intros]: fixes f :: "'a \ 'b::topological_space \ 'c::topological_comm_monoid_mult" shows "(\i. i \ I \ continuous_on S (f i)) \ continuous_on S (\x. \i\I. f i x)" unfolding continuous_on_def by (auto intro: tendsto_prod') instance nat :: topological_comm_monoid_mult by standard (simp add: nhds_discrete principal_prod_principal filterlim_principal eventually_principal) instance int :: topological_comm_monoid_mult by standard (simp add: nhds_discrete principal_prod_principal filterlim_principal eventually_principal) class comm_real_normed_algebra_1 = real_normed_algebra_1 + comm_monoid_mult context real_normed_field begin subclass comm_real_normed_algebra_1 proof from norm_mult[of "1 :: 'a" 1] show "norm 1 = 1" by simp qed (simp_all add: norm_mult) end subsubsection \Inverse and division\ lemma (in bounded_bilinear) Zfun_prod_Bfun: assumes f: "Zfun f F" and g: "Bfun g F" shows "Zfun (\x. f x ** g x) F" proof - obtain K where K: "0 \ K" and norm_le: "\x y. norm (x ** y) \ norm x * norm y * K" using nonneg_bounded by blast obtain B where B: "0 < B" and norm_g: "eventually (\x. norm (g x) \ B) F" using g by (rule BfunE) have "eventually (\x. norm (f x ** g x) \ norm (f x) * (B * K)) F" using norm_g proof eventually_elim case (elim x) have "norm (f x ** g x) \ norm (f x) * norm (g x) * K" by (rule norm_le) also have "\ \ norm (f x) * B * K" by (intro mult_mono' order_refl norm_g norm_ge_zero mult_nonneg_nonneg K elim) also have "\ = norm (f x) * (B * K)" by (rule mult.assoc) finally show "norm (f x ** g x) \ norm (f x) * (B * K)" . qed with f show ?thesis by (rule Zfun_imp_Zfun) qed lemma (in bounded_bilinear) Bfun_prod_Zfun: assumes f: "Bfun f F" and g: "Zfun g F" shows "Zfun (\x. f x ** g x) F" using flip g f by (rule bounded_bilinear.Zfun_prod_Bfun) lemma Bfun_inverse: fixes a :: "'a::real_normed_div_algebra" assumes f: "(f \ a) F" assumes a: "a \ 0" shows "Bfun (\x. inverse (f x)) F" proof - from a have "0 < norm a" by simp then have "\r>0. r < norm a" by (rule dense) then obtain r where r1: "0 < r" and r2: "r < norm a" by blast have "eventually (\x. dist (f x) a < r) F" using tendstoD [OF f r1] by blast then have "eventually (\x. norm (inverse (f x)) \ inverse (norm a - r)) F" proof eventually_elim case (elim x) then have 1: "norm (f x - a) < r" by (simp add: dist_norm) then have 2: "f x \ 0" using r2 by auto then have "norm (inverse (f x)) = inverse (norm (f x))" by (rule nonzero_norm_inverse) also have "\ \ inverse (norm a - r)" proof (rule le_imp_inverse_le) show "0 < norm a - r" using r2 by simp have "norm a - norm (f x) \ norm (a - f x)" by (rule norm_triangle_ineq2) also have "\ = norm (f x - a)" by (rule norm_minus_commute) also have "\ < r" using 1 . finally show "norm a - r \ norm (f x)" by simp qed finally show "norm (inverse (f x)) \ inverse (norm a - r)" . qed then show ?thesis by (rule BfunI) qed lemma tendsto_inverse [tendsto_intros]: fixes a :: "'a::real_normed_div_algebra" assumes f: "(f \ a) F" and a: "a \ 0" shows "((\x. inverse (f x)) \ inverse a) F" proof - from a have "0 < norm a" by simp with f have "eventually (\x. dist (f x) a < norm a) F" by (rule tendstoD) then have "eventually (\x. f x \ 0) F" unfolding dist_norm by (auto elim!: eventually_mono) with a have "eventually (\x. inverse (f x) - inverse a = - (inverse (f x) * (f x - a) * inverse a)) F" by (auto elim!: eventually_mono simp: inverse_diff_inverse) moreover have "Zfun (\x. - (inverse (f x) * (f x - a) * inverse a)) F" by (intro Zfun_minus Zfun_mult_left bounded_bilinear.Bfun_prod_Zfun [OF bounded_bilinear_mult] Bfun_inverse [OF f a] f [unfolded tendsto_Zfun_iff]) ultimately show ?thesis unfolding tendsto_Zfun_iff by (rule Zfun_ssubst) qed lemma continuous_inverse: fixes f :: "'a::t2_space \ 'b::real_normed_div_algebra" assumes "continuous F f" and "f (Lim F (\x. x)) \ 0" shows "continuous F (\x. inverse (f x))" using assms unfolding continuous_def by (rule tendsto_inverse) lemma continuous_at_within_inverse[continuous_intros]: fixes f :: "'a::t2_space \ 'b::real_normed_div_algebra" assumes "continuous (at a within s) f" and "f a \ 0" shows "continuous (at a within s) (\x. inverse (f x))" using assms unfolding continuous_within by (rule tendsto_inverse) lemma continuous_on_inverse[continuous_intros]: fixes f :: "'a::topological_space \ 'b::real_normed_div_algebra" assumes "continuous_on s f" and "\x\s. f x \ 0" shows "continuous_on s (\x. inverse (f x))" using assms unfolding continuous_on_def by (blast intro: tendsto_inverse) lemma tendsto_divide [tendsto_intros]: fixes a b :: "'a::real_normed_field" shows "(f \ a) F \ (g \ b) F \ b \ 0 \ ((\x. f x / g x) \ a / b) F" by (simp add: tendsto_mult tendsto_inverse divide_inverse) lemma continuous_divide: fixes f g :: "'a::t2_space \ 'b::real_normed_field" assumes "continuous F f" and "continuous F g" and "g (Lim F (\x. x)) \ 0" shows "continuous F (\x. (f x) / (g x))" using assms unfolding continuous_def by (rule tendsto_divide) lemma continuous_at_within_divide[continuous_intros]: fixes f g :: "'a::t2_space \ 'b::real_normed_field" assumes "continuous (at a within s) f" "continuous (at a within s) g" and "g a \ 0" shows "continuous (at a within s) (\x. (f x) / (g x))" using assms unfolding continuous_within by (rule tendsto_divide) lemma isCont_divide[continuous_intros, simp]: fixes f g :: "'a::t2_space \ 'b::real_normed_field" assumes "isCont f a" "isCont g a" "g a \ 0" shows "isCont (\x. (f x) / g x) a" using assms unfolding continuous_at by (rule tendsto_divide) lemma continuous_on_divide[continuous_intros]: fixes f :: "'a::topological_space \ 'b::real_normed_field" assumes "continuous_on s f" "continuous_on s g" and "\x\s. g x \ 0" shows "continuous_on s (\x. (f x) / (g x))" using assms unfolding continuous_on_def by (blast intro: tendsto_divide) lemma tendsto_power_int [tendsto_intros]: fixes a :: "'a::real_normed_div_algebra" assumes f: "(f \ a) F" and a: "a \ 0" shows "((\x. power_int (f x) n) \ power_int a n) F" using assms by (cases n rule: int_cases4) (auto intro!: tendsto_intros simp: power_int_minus) lemma continuous_power_int: fixes f :: "'a::t2_space \ 'b::real_normed_div_algebra" assumes "continuous F f" and "f (Lim F (\x. x)) \ 0" shows "continuous F (\x. power_int (f x) n)" using assms unfolding continuous_def by (rule tendsto_power_int) lemma continuous_at_within_power_int[continuous_intros]: fixes f :: "'a::t2_space \ 'b::real_normed_div_algebra" assumes "continuous (at a within s) f" and "f a \ 0" shows "continuous (at a within s) (\x. power_int (f x) n)" using assms unfolding continuous_within by (rule tendsto_power_int) lemma continuous_on_power_int [continuous_intros]: fixes f :: "'a::topological_space \ 'b::real_normed_div_algebra" assumes "continuous_on s f" and "\x\s. f x \ 0" shows "continuous_on s (\x. power_int (f x) n)" using assms unfolding continuous_on_def by (blast intro: tendsto_power_int) lemma tendsto_sgn [tendsto_intros]: "(f \ l) F \ l \ 0 \ ((\x. sgn (f x)) \ sgn l) F" for l :: "'a::real_normed_vector" unfolding sgn_div_norm by (simp add: tendsto_intros) lemma continuous_sgn: fixes f :: "'a::t2_space \ 'b::real_normed_vector" assumes "continuous F f" and "f (Lim F (\x. x)) \ 0" shows "continuous F (\x. sgn (f x))" using assms unfolding continuous_def by (rule tendsto_sgn) lemma continuous_at_within_sgn[continuous_intros]: fixes f :: "'a::t2_space \ 'b::real_normed_vector" assumes "continuous (at a within s) f" and "f a \ 0" shows "continuous (at a within s) (\x. sgn (f x))" using assms unfolding continuous_within by (rule tendsto_sgn) lemma isCont_sgn[continuous_intros]: fixes f :: "'a::t2_space \ 'b::real_normed_vector" assumes "isCont f a" and "f a \ 0" shows "isCont (\x. sgn (f x)) a" using assms unfolding continuous_at by (rule tendsto_sgn) lemma continuous_on_sgn[continuous_intros]: fixes f :: "'a::topological_space \ 'b::real_normed_vector" assumes "continuous_on s f" and "\x\s. f x \ 0" shows "continuous_on s (\x. sgn (f x))" using assms unfolding continuous_on_def by (blast intro: tendsto_sgn) lemma filterlim_at_infinity: fixes f :: "_ \ 'a::real_normed_vector" assumes "0 \ c" shows "(LIM x F. f x :> at_infinity) \ (\r>c. eventually (\x. r \ norm (f x)) F)" unfolding filterlim_iff eventually_at_infinity proof safe fix P :: "'a \ bool" fix b assume *: "\r>c. eventually (\x. r \ norm (f x)) F" assume P: "\x. b \ norm x \ P x" have "max b (c + 1) > c" by auto with * have "eventually (\x. max b (c + 1) \ norm (f x)) F" by auto then show "eventually (\x. P (f x)) F" proof eventually_elim case (elim x) with P show "P (f x)" by auto qed qed force lemma filterlim_at_infinity_imp_norm_at_top: fixes F assumes "filterlim f at_infinity F" shows "filterlim (\x. norm (f x)) at_top F" proof - { fix r :: real have "\\<^sub>F x in F. r \ norm (f x)" using filterlim_at_infinity[of 0 f F] assms by (cases "r > 0") (auto simp: not_less intro: always_eventually order.trans[OF _ norm_ge_zero]) } thus ?thesis by (auto simp: filterlim_at_top) qed lemma filterlim_norm_at_top_imp_at_infinity: fixes F assumes "filterlim (\x. norm (f x)) at_top F" shows "filterlim f at_infinity F" using filterlim_at_infinity[of 0 f F] assms by (auto simp: filterlim_at_top) lemma filterlim_norm_at_top: "filterlim norm at_top at_infinity" by (rule filterlim_at_infinity_imp_norm_at_top) (rule filterlim_ident) lemma filterlim_at_infinity_conv_norm_at_top: "filterlim f at_infinity G \ filterlim (\x. norm (f x)) at_top G" by (auto simp: filterlim_at_infinity[OF order.refl] filterlim_at_top_gt[of _ _ 0]) lemma eventually_not_equal_at_infinity: "eventually (\x. x \ (a :: 'a :: {real_normed_vector})) at_infinity" proof - from filterlim_norm_at_top[where 'a = 'a] have "\\<^sub>F x in at_infinity. norm a < norm (x::'a)" by (auto simp: filterlim_at_top_dense) thus ?thesis by eventually_elim auto qed lemma filterlim_int_of_nat_at_topD: fixes F assumes "filterlim (\x. f (int x)) F at_top" shows "filterlim f F at_top" proof - have "filterlim (\x. f (int (nat x))) F at_top" by (rule filterlim_compose[OF assms filterlim_nat_sequentially]) also have "?this \ filterlim f F at_top" by (intro filterlim_cong refl eventually_mono [OF eventually_ge_at_top[of "0::int"]]) auto finally show ?thesis . qed lemma filterlim_int_sequentially [tendsto_intros]: "filterlim int at_top sequentially" unfolding filterlim_at_top proof fix C :: int show "eventually (\n. int n \ C) at_top" using eventually_ge_at_top[of "nat \C\"] by eventually_elim linarith qed lemma filterlim_real_of_int_at_top [tendsto_intros]: "filterlim real_of_int at_top at_top" unfolding filterlim_at_top proof fix C :: real show "eventually (\n. real_of_int n \ C) at_top" using eventually_ge_at_top[of "\C\"] by eventually_elim linarith qed lemma filterlim_abs_real: "filterlim (abs::real \ real) at_top at_top" proof (subst filterlim_cong[OF refl refl]) from eventually_ge_at_top[of "0::real"] show "eventually (\x::real. \x\ = x) at_top" by eventually_elim simp qed (simp_all add: filterlim_ident) lemma filterlim_of_real_at_infinity [tendsto_intros]: "filterlim (of_real :: real \ 'a :: real_normed_algebra_1) at_infinity at_top" by (intro filterlim_norm_at_top_imp_at_infinity) (auto simp: filterlim_abs_real) lemma not_tendsto_and_filterlim_at_infinity: fixes c :: "'a::real_normed_vector" assumes "F \ bot" and "(f \ c) F" and "filterlim f at_infinity F" shows False proof - from tendstoD[OF assms(2), of "1/2"] have "eventually (\x. dist (f x) c < 1/2) F" by simp moreover from filterlim_at_infinity[of "norm c" f F] assms(3) have "eventually (\x. norm (f x) \ norm c + 1) F" by simp ultimately have "eventually (\x. False) F" proof eventually_elim fix x assume A: "dist (f x) c < 1/2" assume "norm (f x) \ norm c + 1" also have "norm (f x) = dist (f x) 0" by simp also have "\ \ dist (f x) c + dist c 0" by (rule dist_triangle) finally show False using A by simp qed with assms show False by simp qed lemma filterlim_at_infinity_imp_not_convergent: assumes "filterlim f at_infinity sequentially" shows "\ convergent f" by (rule notI, rule not_tendsto_and_filterlim_at_infinity[OF _ _ assms]) (simp_all add: convergent_LIMSEQ_iff) lemma filterlim_at_infinity_imp_eventually_ne: assumes "filterlim f at_infinity F" shows "eventually (\z. f z \ c) F" proof - have "norm c + 1 > 0" by (intro add_nonneg_pos) simp_all with filterlim_at_infinity[OF order.refl, of f F] assms have "eventually (\z. norm (f z) \ norm c + 1) F" by blast then show ?thesis by eventually_elim auto qed lemma tendsto_of_nat [tendsto_intros]: "filterlim (of_nat :: nat \ 'a::real_normed_algebra_1) at_infinity sequentially" proof (subst filterlim_at_infinity[OF order.refl], intro allI impI) fix r :: real assume r: "r > 0" define n where "n = nat \r\" from r have n: "\m\n. of_nat m \ r" unfolding n_def by linarith from eventually_ge_at_top[of n] show "eventually (\m. norm (of_nat m :: 'a) \ r) sequentially" by eventually_elim (use n in simp_all) qed subsection \Relate \<^const>\at\, \<^const>\at_left\ and \<^const>\at_right\\ text \ This lemmas are useful for conversion between \<^term>\at x\ to \<^term>\at_left x\ and \<^term>\at_right x\ and also \<^term>\at_right 0\. \ lemmas filterlim_split_at_real = filterlim_split_at[where 'a=real] lemma filtermap_nhds_shift: "filtermap (\x. x - d) (nhds a) = nhds (a - d)" for a d :: "'a::real_normed_vector" by (rule filtermap_fun_inverse[where g="\x. x + d"]) (auto intro!: tendsto_eq_intros filterlim_ident) lemma filtermap_nhds_minus: "filtermap (\x. - x) (nhds a) = nhds (- a)" for a :: "'a::real_normed_vector" by (rule filtermap_fun_inverse[where g=uminus]) (auto intro!: tendsto_eq_intros filterlim_ident) lemma filtermap_at_shift: "filtermap (\x. x - d) (at a) = at (a - d)" for a d :: "'a::real_normed_vector" by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_shift[symmetric]) lemma filtermap_at_right_shift: "filtermap (\x. x - d) (at_right a) = at_right (a - d)" for a d :: "real" by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_shift[symmetric]) +lemma filterlim_shift: + fixes d :: "'a::real_normed_vector" + assumes "filterlim f F (at a)" + shows "filterlim (f \ (+) d) F (at (a - d))" + unfolding filterlim_iff +proof (intro strip) + fix P + assume "eventually P F" + then have "\\<^sub>F x in filtermap (\y. y - d) (at a). P (f (d + x))" + using assms by (force simp add: filterlim_iff eventually_filtermap) + then show "(\\<^sub>F x in at (a - d). P ((f \ (+) d) x))" + by (force simp add: filtermap_at_shift) +qed + +lemma filterlim_shift_iff: + fixes d :: "'a::real_normed_vector" + shows "filterlim (f \ (+) d) F (at (a - d)) = filterlim f F (at a)" (is "?lhs = ?rhs") +proof + assume L: ?lhs show ?rhs + using filterlim_shift [OF L, of "-d"] by (simp add: filterlim_iff) +qed (metis filterlim_shift) + lemma at_right_to_0: "at_right a = filtermap (\x. x + a) (at_right 0)" for a :: real using filtermap_at_right_shift[of "-a" 0] by simp lemma filterlim_at_right_to_0: "filterlim f F (at_right a) \ filterlim (\x. f (x + a)) F (at_right 0)" for a :: real unfolding filterlim_def filtermap_filtermap at_right_to_0[of a] .. lemma eventually_at_right_to_0: "eventually P (at_right a) \ eventually (\x. P (x + a)) (at_right 0)" for a :: real unfolding at_right_to_0[of a] by (simp add: eventually_filtermap) lemma at_to_0: "at a = filtermap (\x. x + a) (at 0)" for a :: "'a::real_normed_vector" using filtermap_at_shift[of "-a" 0] by simp lemma filterlim_at_to_0: "filterlim f F (at a) \ filterlim (\x. f (x + a)) F (at 0)" for a :: "'a::real_normed_vector" unfolding filterlim_def filtermap_filtermap at_to_0[of a] .. lemma eventually_at_to_0: "eventually P (at a) \ eventually (\x. P (x + a)) (at 0)" for a :: "'a::real_normed_vector" unfolding at_to_0[of a] by (simp add: eventually_filtermap) lemma filtermap_at_minus: "filtermap (\x. - x) (at a) = at (- a)" for a :: "'a::real_normed_vector" by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_minus[symmetric]) lemma at_left_minus: "at_left a = filtermap (\x. - x) (at_right (- a))" for a :: real by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_minus[symmetric]) lemma at_right_minus: "at_right a = filtermap (\x. - x) (at_left (- a))" for a :: real by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_minus[symmetric]) lemma filterlim_at_left_to_right: "filterlim f F (at_left a) \ filterlim (\x. f (- x)) F (at_right (-a))" for a :: real unfolding filterlim_def filtermap_filtermap at_left_minus[of a] .. lemma eventually_at_left_to_right: "eventually P (at_left a) \ eventually (\x. P (- x)) (at_right (-a))" for a :: real unfolding at_left_minus[of a] by (simp add: eventually_filtermap) lemma filterlim_uminus_at_top_at_bot: "LIM x at_bot. - x :: real :> at_top" unfolding filterlim_at_top eventually_at_bot_dense by (metis leI minus_less_iff order_less_asym) lemma filterlim_uminus_at_bot_at_top: "LIM x at_top. - x :: real :> at_bot" unfolding filterlim_at_bot eventually_at_top_dense by (metis leI less_minus_iff order_less_asym) lemma at_bot_mirror : shows "(at_bot::('a::{ordered_ab_group_add,linorder} filter)) = filtermap uminus at_top" proof (rule filtermap_fun_inverse[symmetric]) show "filterlim uminus at_top (at_bot::'a filter)" using eventually_at_bot_linorder filterlim_at_top le_minus_iff by force show "filterlim uminus (at_bot::'a filter) at_top" by (simp add: filterlim_at_bot minus_le_iff) qed auto lemma at_top_mirror : shows "(at_top::('a::{ordered_ab_group_add,linorder} filter)) = filtermap uminus at_bot" apply (subst at_bot_mirror) by (auto simp: filtermap_filtermap) lemma filterlim_at_top_mirror: "(LIM x at_top. f x :> F) \ (LIM x at_bot. f (-x::real) :> F)" unfolding filterlim_def at_top_mirror filtermap_filtermap .. lemma filterlim_at_bot_mirror: "(LIM x at_bot. f x :> F) \ (LIM x at_top. f (-x::real) :> F)" unfolding filterlim_def at_bot_mirror filtermap_filtermap .. lemma filterlim_uminus_at_top: "(LIM x F. f x :> at_top) \ (LIM x F. - (f x) :: real :> at_bot)" using filterlim_compose[OF filterlim_uminus_at_bot_at_top, of f F] and filterlim_compose[OF filterlim_uminus_at_top_at_bot, of "\x. - f x" F] by auto lemma tendsto_at_botI_sequentially: fixes f :: "real \ 'b::first_countable_topology" assumes *: "\X. filterlim X at_bot sequentially \ (\n. f (X n)) \ y" shows "(f \ y) at_bot" unfolding filterlim_at_bot_mirror proof (rule tendsto_at_topI_sequentially) fix X :: "nat \ real" assume "filterlim X at_top sequentially" thus "(\n. f (-X n)) \ y" by (intro *) (auto simp: filterlim_uminus_at_top) qed lemma filterlim_at_infinity_imp_filterlim_at_top: assumes "filterlim (f :: 'a \ real) at_infinity F" assumes "eventually (\x. f x > 0) F" shows "filterlim f at_top F" proof - from assms(2) have *: "eventually (\x. norm (f x) = f x) F" by eventually_elim simp from assms(1) show ?thesis unfolding filterlim_at_infinity_conv_norm_at_top by (subst (asm) filterlim_cong[OF refl refl *]) qed lemma filterlim_at_infinity_imp_filterlim_at_bot: assumes "filterlim (f :: 'a \ real) at_infinity F" assumes "eventually (\x. f x < 0) F" shows "filterlim f at_bot F" proof - from assms(2) have *: "eventually (\x. norm (f x) = -f x) F" by eventually_elim simp from assms(1) have "filterlim (\x. - f x) at_top F" unfolding filterlim_at_infinity_conv_norm_at_top by (subst (asm) filterlim_cong[OF refl refl *]) thus ?thesis by (simp add: filterlim_uminus_at_top) qed lemma filterlim_uminus_at_bot: "(LIM x F. f x :> at_bot) \ (LIM x F. - (f x) :: real :> at_top)" unfolding filterlim_uminus_at_top by simp lemma filterlim_inverse_at_top_right: "LIM x at_right (0::real). inverse x :> at_top" unfolding filterlim_at_top_gt[where c=0] eventually_at_filter proof safe fix Z :: real assume [arith]: "0 < Z" then have "eventually (\x. x < inverse Z) (nhds 0)" by (auto simp: eventually_nhds_metric dist_real_def intro!: exI[of _ "\inverse Z\"]) then show "eventually (\x. x \ 0 \ x \ {0<..} \ Z \ inverse x) (nhds 0)" by (auto elim!: eventually_mono simp: inverse_eq_divide field_simps) qed lemma tendsto_inverse_0: fixes x :: "_ \ 'a::real_normed_div_algebra" shows "(inverse \ (0::'a)) at_infinity" unfolding tendsto_Zfun_iff diff_0_right Zfun_def eventually_at_infinity proof safe fix r :: real assume "0 < r" show "\b. \x. b \ norm x \ norm (inverse x :: 'a) < r" proof (intro exI[of _ "inverse (r / 2)"] allI impI) fix x :: 'a from \0 < r\ have "0 < inverse (r / 2)" by simp also assume *: "inverse (r / 2) \ norm x" finally show "norm (inverse x) < r" using * \0 < r\ by (subst nonzero_norm_inverse) (simp_all add: inverse_eq_divide field_simps) qed qed lemma tendsto_add_filterlim_at_infinity: fixes c :: "'b::real_normed_vector" and F :: "'a filter" assumes "(f \ c) F" and "filterlim g at_infinity F" shows "filterlim (\x. f x + g x) at_infinity F" proof (subst filterlim_at_infinity[OF order_refl], safe) fix r :: real assume r: "r > 0" from assms(1) have "((\x. norm (f x)) \ norm c) F" by (rule tendsto_norm) then have "eventually (\x. norm (f x) < norm c + 1) F" by (rule order_tendstoD) simp_all moreover from r have "r + norm c + 1 > 0" by (intro add_pos_nonneg) simp_all with assms(2) have "eventually (\x. norm (g x) \ r + norm c + 1) F" unfolding filterlim_at_infinity[OF order_refl] by (elim allE[of _ "r + norm c + 1"]) simp_all ultimately show "eventually (\x. norm (f x + g x) \ r) F" proof eventually_elim fix x :: 'a assume A: "norm (f x) < norm c + 1" and B: "r + norm c + 1 \ norm (g x)" from A B have "r \ norm (g x) - norm (f x)" by simp also have "norm (g x) - norm (f x) \ norm (g x + f x)" by (rule norm_diff_ineq) finally show "r \ norm (f x + g x)" by (simp add: add_ac) qed qed lemma tendsto_add_filterlim_at_infinity': fixes c :: "'b::real_normed_vector" and F :: "'a filter" assumes "filterlim f at_infinity F" and "(g \ c) F" shows "filterlim (\x. f x + g x) at_infinity F" by (subst add.commute) (rule tendsto_add_filterlim_at_infinity assms)+ lemma filterlim_inverse_at_right_top: "LIM x at_top. inverse x :> at_right (0::real)" unfolding filterlim_at by (auto simp: eventually_at_top_dense) (metis tendsto_inverse_0 filterlim_mono at_top_le_at_infinity order_refl) lemma filterlim_inverse_at_top: "(f \ (0 :: real)) F \ eventually (\x. 0 < f x) F \ LIM x F. inverse (f x) :> at_top" by (intro filterlim_compose[OF filterlim_inverse_at_top_right]) (simp add: filterlim_def eventually_filtermap eventually_mono at_within_def le_principal) lemma filterlim_inverse_at_bot_neg: "LIM x (at_left (0::real)). inverse x :> at_bot" by (simp add: filterlim_inverse_at_top_right filterlim_uminus_at_bot filterlim_at_left_to_right) lemma filterlim_inverse_at_bot: "(f \ (0 :: real)) F \ eventually (\x. f x < 0) F \ LIM x F. inverse (f x) :> at_bot" unfolding filterlim_uminus_at_bot inverse_minus_eq[symmetric] by (rule filterlim_inverse_at_top) (simp_all add: tendsto_minus_cancel_left[symmetric]) lemma at_right_to_top: "(at_right (0::real)) = filtermap inverse at_top" by (intro filtermap_fun_inverse[symmetric, where g=inverse]) (auto intro: filterlim_inverse_at_top_right filterlim_inverse_at_right_top) lemma eventually_at_right_to_top: "eventually P (at_right (0::real)) \ eventually (\x. P (inverse x)) at_top" unfolding at_right_to_top eventually_filtermap .. lemma filterlim_at_right_to_top: "filterlim f F (at_right (0::real)) \ (LIM x at_top. f (inverse x) :> F)" unfolding filterlim_def at_right_to_top filtermap_filtermap .. lemma at_top_to_right: "at_top = filtermap inverse (at_right (0::real))" unfolding at_right_to_top filtermap_filtermap inverse_inverse_eq filtermap_ident .. lemma eventually_at_top_to_right: "eventually P at_top \ eventually (\x. P (inverse x)) (at_right (0::real))" unfolding at_top_to_right eventually_filtermap .. lemma filterlim_at_top_to_right: "filterlim f F at_top \ (LIM x (at_right (0::real)). f (inverse x) :> F)" unfolding filterlim_def at_top_to_right filtermap_filtermap .. lemma filterlim_inverse_at_infinity: fixes x :: "_ \ 'a::{real_normed_div_algebra, division_ring}" shows "filterlim inverse at_infinity (at (0::'a))" unfolding filterlim_at_infinity[OF order_refl] proof safe fix r :: real assume "0 < r" then show "eventually (\x::'a. r \ norm (inverse x)) (at 0)" unfolding eventually_at norm_inverse by (intro exI[of _ "inverse r"]) (auto simp: norm_conv_dist[symmetric] field_simps inverse_eq_divide) qed lemma filterlim_inverse_at_iff: fixes g :: "'a \ 'b::{real_normed_div_algebra, division_ring}" shows "(LIM x F. inverse (g x) :> at 0) \ (LIM x F. g x :> at_infinity)" unfolding filterlim_def filtermap_filtermap[symmetric] proof assume "filtermap g F \ at_infinity" then have "filtermap inverse (filtermap g F) \ filtermap inverse at_infinity" by (rule filtermap_mono) also have "\ \ at 0" using tendsto_inverse_0[where 'a='b] by (auto intro!: exI[of _ 1] simp: le_principal eventually_filtermap filterlim_def at_within_def eventually_at_infinity) finally show "filtermap inverse (filtermap g F) \ at 0" . next assume "filtermap inverse (filtermap g F) \ at 0" then have "filtermap inverse (filtermap inverse (filtermap g F)) \ filtermap inverse (at 0)" by (rule filtermap_mono) with filterlim_inverse_at_infinity show "filtermap g F \ at_infinity" by (auto intro: order_trans simp: filterlim_def filtermap_filtermap) qed lemma tendsto_mult_filterlim_at_infinity: fixes c :: "'a::real_normed_field" assumes "(f \ c) F" "c \ 0" assumes "filterlim g at_infinity F" shows "filterlim (\x. f x * g x) at_infinity F" proof - have "((\x. inverse (f x) * inverse (g x)) \ inverse c * 0) F" by (intro tendsto_mult tendsto_inverse assms filterlim_compose[OF tendsto_inverse_0]) then have "filterlim (\x. inverse (f x) * inverse (g x)) (at (inverse c * 0)) F" unfolding filterlim_at using assms by (auto intro: filterlim_at_infinity_imp_eventually_ne tendsto_imp_eventually_ne eventually_conj) then show ?thesis by (subst filterlim_inverse_at_iff[symmetric]) simp_all qed lemma tendsto_inverse_0_at_top: "LIM x F. f x :> at_top \ ((\x. inverse (f x) :: real) \ 0) F" by (metis filterlim_at filterlim_mono[OF _ at_top_le_at_infinity order_refl] filterlim_inverse_at_iff) lemma real_tendsto_divide_at_top: fixes c::"real" assumes "(f \ c) F" assumes "filterlim g at_top F" shows "((\x. f x / g x) \ 0) F" by (auto simp: divide_inverse_commute intro!: tendsto_mult[THEN tendsto_eq_rhs] tendsto_inverse_0_at_top assms) lemma mult_nat_left_at_top: "c > 0 \ filterlim (\x. c * x) at_top sequentially" for c :: nat by (rule filterlim_subseq) (auto simp: strict_mono_def) lemma mult_nat_right_at_top: "c > 0 \ filterlim (\x. x * c) at_top sequentially" for c :: nat by (rule filterlim_subseq) (auto simp: strict_mono_def) lemma filterlim_times_pos: "LIM x F1. c * f x :> at_right l" if "filterlim f (at_right p) F1" "0 < c" "l = c * p" for c::"'a::{linordered_field, linorder_topology}" unfolding filterlim_iff proof safe fix P assume "\\<^sub>F x in at_right l. P x" then obtain d where "c * p < d" "\y. y > c * p \ y < d \ P y" unfolding \l = _ \ eventually_at_right_field by auto then have "\\<^sub>F a in at_right p. P (c * a)" by (auto simp: eventually_at_right_field \0 < c\ field_simps intro!: exI[where x="d/c"]) from that(1)[unfolded filterlim_iff, rule_format, OF this] show "\\<^sub>F x in F1. P (c * f x)" . qed lemma filtermap_nhds_times: "c \ 0 \ filtermap (times c) (nhds a) = nhds (c * a)" for a c :: "'a::real_normed_field" by (rule filtermap_fun_inverse[where g="\x. inverse c * x"]) (auto intro!: tendsto_eq_intros filterlim_ident) lemma filtermap_times_pos_at_right: fixes c::"'a::{linordered_field, linorder_topology}" assumes "c > 0" shows "filtermap (times c) (at_right p) = at_right (c * p)" using assms by (intro filtermap_fun_inverse[where g="\x. inverse c * x"]) (auto intro!: filterlim_ident filterlim_times_pos) lemma at_to_infinity: "(at (0::'a::{real_normed_field,field})) = filtermap inverse at_infinity" proof (rule antisym) have "(inverse \ (0::'a)) at_infinity" by (fact tendsto_inverse_0) then show "filtermap inverse at_infinity \ at (0::'a)" using filterlim_def filterlim_ident filterlim_inverse_at_iff by fastforce next have "filtermap inverse (filtermap inverse (at (0::'a))) \ filtermap inverse at_infinity" using filterlim_inverse_at_infinity unfolding filterlim_def by (rule filtermap_mono) then show "at (0::'a) \ filtermap inverse at_infinity" by (simp add: filtermap_ident filtermap_filtermap) qed lemma lim_at_infinity_0: fixes l :: "'a::{real_normed_field,field}" shows "(f \ l) at_infinity \ ((f \ inverse) \ l) (at (0::'a))" by (simp add: tendsto_compose_filtermap at_to_infinity filtermap_filtermap) lemma lim_zero_infinity: fixes l :: "'a::{real_normed_field,field}" shows "((\x. f(1 / x)) \ l) (at (0::'a)) \ (f \ l) at_infinity" by (simp add: inverse_eq_divide lim_at_infinity_0 comp_def) text \ We only show rules for multiplication and addition when the functions are either against a real value or against infinity. Further rules are easy to derive by using @{thm filterlim_uminus_at_top}. \ lemma filterlim_tendsto_pos_mult_at_top: assumes f: "(f \ c) F" and c: "0 < c" and g: "LIM x F. g x :> at_top" shows "LIM x F. (f x * g x :: real) :> at_top" unfolding filterlim_at_top_gt[where c=0] proof safe fix Z :: real assume "0 < Z" from f \0 < c\ have "eventually (\x. c / 2 < f x) F" by (auto dest!: tendstoD[where e="c / 2"] elim!: eventually_mono simp: dist_real_def abs_real_def split: if_split_asm) moreover from g have "eventually (\x. (Z / c * 2) \ g x) F" unfolding filterlim_at_top by auto ultimately show "eventually (\x. Z \ f x * g x) F" proof eventually_elim case (elim x) with \0 < Z\ \0 < c\ have "c / 2 * (Z / c * 2) \ f x * g x" by (intro mult_mono) (auto simp: zero_le_divide_iff) with \0 < c\ show "Z \ f x * g x" by simp qed qed lemma filterlim_at_top_mult_at_top: assumes f: "LIM x F. f x :> at_top" and g: "LIM x F. g x :> at_top" shows "LIM x F. (f x * g x :: real) :> at_top" unfolding filterlim_at_top_gt[where c=0] proof safe fix Z :: real assume "0 < Z" from f have "eventually (\x. 1 \ f x) F" unfolding filterlim_at_top by auto moreover from g have "eventually (\x. Z \ g x) F" unfolding filterlim_at_top by auto ultimately show "eventually (\x. Z \ f x * g x) F" proof eventually_elim case (elim x) with \0 < Z\ have "1 * Z \ f x * g x" by (intro mult_mono) (auto simp: zero_le_divide_iff) then show "Z \ f x * g x" by simp qed qed lemma filterlim_at_top_mult_tendsto_pos: assumes f: "(f \ c) F" and c: "0 < c" and g: "LIM x F. g x :> at_top" shows "LIM x F. (g x * f x:: real) :> at_top" by (auto simp: mult.commute intro!: filterlim_tendsto_pos_mult_at_top f c g) lemma filterlim_tendsto_pos_mult_at_bot: fixes c :: real assumes "(f \ c) F" "0 < c" "filterlim g at_bot F" shows "LIM x F. f x * g x :> at_bot" using filterlim_tendsto_pos_mult_at_top[OF assms(1,2), of "\x. - g x"] assms(3) unfolding filterlim_uminus_at_bot by simp lemma filterlim_tendsto_neg_mult_at_bot: fixes c :: real assumes c: "(f \ c) F" "c < 0" and g: "filterlim g at_top F" shows "LIM x F. f x * g x :> at_bot" using c filterlim_tendsto_pos_mult_at_top[of "\x. - f x" "- c" F, OF _ _ g] unfolding filterlim_uminus_at_bot tendsto_minus_cancel_left by simp lemma filterlim_pow_at_top: fixes f :: "'a \ real" assumes "0 < n" and f: "LIM x F. f x :> at_top" shows "LIM x F. (f x)^n :: real :> at_top" using \0 < n\ proof (induct n) case 0 then show ?case by simp next case (Suc n) with f show ?case by (cases "n = 0") (auto intro!: filterlim_at_top_mult_at_top) qed lemma filterlim_pow_at_bot_even: fixes f :: "real \ real" shows "0 < n \ LIM x F. f x :> at_bot \ even n \ LIM x F. (f x)^n :> at_top" using filterlim_pow_at_top[of n "\x. - f x" F] by (simp add: filterlim_uminus_at_top) lemma filterlim_pow_at_bot_odd: fixes f :: "real \ real" shows "0 < n \ LIM x F. f x :> at_bot \ odd n \ LIM x F. (f x)^n :> at_bot" using filterlim_pow_at_top[of n "\x. - f x" F] by (simp add: filterlim_uminus_at_bot) lemma filterlim_power_at_infinity [tendsto_intros]: fixes F and f :: "'a \ 'b :: real_normed_div_algebra" assumes "filterlim f at_infinity F" "n > 0" shows "filterlim (\x. f x ^ n) at_infinity F" by (rule filterlim_norm_at_top_imp_at_infinity) (auto simp: norm_power intro!: filterlim_pow_at_top assms intro: filterlim_at_infinity_imp_norm_at_top) lemma filterlim_tendsto_add_at_top: assumes f: "(f \ c) F" and g: "LIM x F. g x :> at_top" shows "LIM x F. (f x + g x :: real) :> at_top" unfolding filterlim_at_top_gt[where c=0] proof safe fix Z :: real assume "0 < Z" from f have "eventually (\x. c - 1 < f x) F" by (auto dest!: tendstoD[where e=1] elim!: eventually_mono simp: dist_real_def) moreover from g have "eventually (\x. Z - (c - 1) \ g x) F" unfolding filterlim_at_top by auto ultimately show "eventually (\x. Z \ f x + g x) F" by eventually_elim simp qed lemma LIM_at_top_divide: fixes f g :: "'a \ real" assumes f: "(f \ a) F" "0 < a" and g: "(g \ 0) F" "eventually (\x. 0 < g x) F" shows "LIM x F. f x / g x :> at_top" unfolding divide_inverse by (rule filterlim_tendsto_pos_mult_at_top[OF f]) (rule filterlim_inverse_at_top[OF g]) lemma filterlim_at_top_add_at_top: assumes f: "LIM x F. f x :> at_top" and g: "LIM x F. g x :> at_top" shows "LIM x F. (f x + g x :: real) :> at_top" unfolding filterlim_at_top_gt[where c=0] proof safe fix Z :: real assume "0 < Z" from f have "eventually (\x. 0 \ f x) F" unfolding filterlim_at_top by auto moreover from g have "eventually (\x. Z \ g x) F" unfolding filterlim_at_top by auto ultimately show "eventually (\x. Z \ f x + g x) F" by eventually_elim simp qed lemma tendsto_divide_0: fixes f :: "_ \ 'a::{real_normed_div_algebra, division_ring}" assumes f: "(f \ c) F" and g: "LIM x F. g x :> at_infinity" shows "((\x. f x / g x) \ 0) F" using tendsto_mult[OF f filterlim_compose[OF tendsto_inverse_0 g]] by (simp add: divide_inverse) lemma linear_plus_1_le_power: fixes x :: real assumes x: "0 \ x" shows "real n * x + 1 \ (x + 1) ^ n" proof (induct n) case 0 then show ?case by simp next case (Suc n) from x have "real (Suc n) * x + 1 \ (x + 1) * (real n * x + 1)" by (simp add: field_simps) also have "\ \ (x + 1)^Suc n" using Suc x by (simp add: mult_left_mono) finally show ?case . qed lemma filterlim_realpow_sequentially_gt1: fixes x :: "'a :: real_normed_div_algebra" assumes x[arith]: "1 < norm x" shows "LIM n sequentially. x ^ n :> at_infinity" proof (intro filterlim_at_infinity[THEN iffD2] allI impI) fix y :: real assume "0 < y" obtain N :: nat where "y < real N * (norm x - 1)" by (meson diff_gt_0_iff_gt reals_Archimedean3 x) also have "\ \ real N * (norm x - 1) + 1" by simp also have "\ \ (norm x - 1 + 1) ^ N" by (rule linear_plus_1_le_power) simp also have "\ = norm x ^ N" by simp finally have "\n\N. y \ norm x ^ n" by (metis order_less_le_trans power_increasing order_less_imp_le x) then show "eventually (\n. y \ norm (x ^ n)) sequentially" unfolding eventually_sequentially by (auto simp: norm_power) qed simp lemma filterlim_divide_at_infinity: fixes f g :: "'a \ 'a :: real_normed_field" assumes "filterlim f (nhds c) F" "filterlim g (at 0) F" "c \ 0" shows "filterlim (\x. f x / g x) at_infinity F" proof - have "filterlim (\x. f x * inverse (g x)) at_infinity F" by (intro tendsto_mult_filterlim_at_infinity[OF assms(1,3)] filterlim_compose [OF filterlim_inverse_at_infinity assms(2)]) thus ?thesis by (simp add: field_simps) qed subsection \Floor and Ceiling\ lemma eventually_floor_less: fixes f :: "'a \ 'b::{order_topology,floor_ceiling}" assumes f: "(f \ l) F" and l: "l \ \" shows "\\<^sub>F x in F. of_int (floor l) < f x" by (intro order_tendstoD[OF f]) (metis Ints_of_int antisym_conv2 floor_correct l) lemma eventually_less_ceiling: fixes f :: "'a \ 'b::{order_topology,floor_ceiling}" assumes f: "(f \ l) F" and l: "l \ \" shows "\\<^sub>F x in F. f x < of_int (ceiling l)" by (intro order_tendstoD[OF f]) (metis Ints_of_int l le_of_int_ceiling less_le) lemma eventually_floor_eq: fixes f::"'a \ 'b::{order_topology,floor_ceiling}" assumes f: "(f \ l) F" and l: "l \ \" shows "\\<^sub>F x in F. floor (f x) = floor l" using eventually_floor_less[OF assms] eventually_less_ceiling[OF assms] by eventually_elim (meson floor_less_iff less_ceiling_iff not_less_iff_gr_or_eq) lemma eventually_ceiling_eq: fixes f::"'a \ 'b::{order_topology,floor_ceiling}" assumes f: "(f \ l) F" and l: "l \ \" shows "\\<^sub>F x in F. ceiling (f x) = ceiling l" using eventually_floor_less[OF assms] eventually_less_ceiling[OF assms] by eventually_elim (meson floor_less_iff less_ceiling_iff not_less_iff_gr_or_eq) lemma tendsto_of_int_floor: fixes f::"'a \ 'b::{order_topology,floor_ceiling}" assumes "(f \ l) F" and "l \ \" shows "((\x. of_int (floor (f x)) :: 'c::{ring_1,topological_space}) \ of_int (floor l)) F" using eventually_floor_eq[OF assms] by (simp add: eventually_mono topological_tendstoI) lemma tendsto_of_int_ceiling: fixes f::"'a \ 'b::{order_topology,floor_ceiling}" assumes "(f \ l) F" and "l \ \" shows "((\x. of_int (ceiling (f x)):: 'c::{ring_1,topological_space}) \ of_int (ceiling l)) F" using eventually_ceiling_eq[OF assms] by (simp add: eventually_mono topological_tendstoI) lemma continuous_on_of_int_floor: "continuous_on (UNIV - \::'a::{order_topology, floor_ceiling} set) (\x. of_int (floor x)::'b::{ring_1, topological_space})" unfolding continuous_on_def by (auto intro!: tendsto_of_int_floor) lemma continuous_on_of_int_ceiling: "continuous_on (UNIV - \::'a::{order_topology, floor_ceiling} set) (\x. of_int (ceiling x)::'b::{ring_1, topological_space})" unfolding continuous_on_def by (auto intro!: tendsto_of_int_ceiling) subsection \Limits of Sequences\ lemma [trans]: "X = Y \ Y \ z \ X \ z" by simp lemma LIMSEQ_iff: fixes L :: "'a::real_normed_vector" shows "(X \ L) = (\r>0. \no. \n \ no. norm (X n - L) < r)" unfolding lim_sequentially dist_norm .. lemma LIMSEQ_I: "(\r. 0 < r \ \no. \n\no. norm (X n - L) < r) \ X \ L" for L :: "'a::real_normed_vector" by (simp add: LIMSEQ_iff) lemma LIMSEQ_D: "X \ L \ 0 < r \ \no. \n\no. norm (X n - L) < r" for L :: "'a::real_normed_vector" by (simp add: LIMSEQ_iff) lemma LIMSEQ_linear: "X \ x \ l > 0 \ (\ n. X (n * l)) \ x" unfolding tendsto_def eventually_sequentially by (metis div_le_dividend div_mult_self1_is_m le_trans mult.commute) text \Transformation of limit.\ lemma Lim_transform: "(g \ a) F \ ((\x. f x - g x) \ 0) F \ (f \ a) F" for a b :: "'a::real_normed_vector" using tendsto_add [of g a F "\x. f x - g x" 0] by simp lemma Lim_transform2: "(f \ a) F \ ((\x. f x - g x) \ 0) F \ (g \ a) F" for a b :: "'a::real_normed_vector" by (erule Lim_transform) (simp add: tendsto_minus_cancel) proposition Lim_transform_eq: "((\x. f x - g x) \ 0) F \ (f \ a) F \ (g \ a) F" for a :: "'a::real_normed_vector" using Lim_transform Lim_transform2 by blast lemma Lim_transform_eventually: "\(f \ l) F; eventually (\x. f x = g x) F\ \ (g \ l) F" using eventually_elim2 by (fastforce simp add: tendsto_def) lemma Lim_transform_within: assumes "(f \ l) (at x within S)" and "0 < d" and "\x'. x'\S \ 0 < dist x' x \ dist x' x < d \ f x' = g x'" shows "(g \ l) (at x within S)" proof (rule Lim_transform_eventually) show "eventually (\x. f x = g x) (at x within S)" using assms by (auto simp: eventually_at) show "(f \ l) (at x within S)" by fact qed lemma filterlim_transform_within: assumes "filterlim g G (at x within S)" assumes "G \ F" "0x'. x' \ S \ 0 < dist x' x \ dist x' x < d \ f x' = g x') " shows "filterlim f F (at x within S)" using assms apply (elim filterlim_mono_eventually) unfolding eventually_at by auto text \Common case assuming being away from some crucial point like 0.\ lemma Lim_transform_away_within: fixes a b :: "'a::t1_space" assumes "a \ b" and "\x\S. x \ a \ x \ b \ f x = g x" and "(f \ l) (at a within S)" shows "(g \ l) (at a within S)" proof (rule Lim_transform_eventually) show "(f \ l) (at a within S)" by fact show "eventually (\x. f x = g x) (at a within S)" unfolding eventually_at_topological by (rule exI [where x="- {b}"]) (simp add: open_Compl assms) qed lemma Lim_transform_away_at: fixes a b :: "'a::t1_space" assumes ab: "a \ b" and fg: "\x. x \ a \ x \ b \ f x = g x" and fl: "(f \ l) (at a)" shows "(g \ l) (at a)" using Lim_transform_away_within[OF ab, of UNIV f g l] fg fl by simp text \Alternatively, within an open set.\ lemma Lim_transform_within_open: assumes "(f \ l) (at a within T)" and "open s" and "a \ s" and "\x. x\s \ x \ a \ f x = g x" shows "(g \ l) (at a within T)" proof (rule Lim_transform_eventually) show "eventually (\x. f x = g x) (at a within T)" unfolding eventually_at_topological using assms by auto show "(f \ l) (at a within T)" by fact qed text \A congruence rule allowing us to transform limits assuming not at point.\ lemma Lim_cong_within: assumes "a = b" and "x = y" and "S = T" and "\x. x \ b \ x \ T \ f x = g x" shows "(f \ x) (at a within S) \ (g \ y) (at b within T)" unfolding tendsto_def eventually_at_topological using assms by simp text \An unbounded sequence's inverse tends to 0.\ lemma LIMSEQ_inverse_zero: assumes "\r::real. \N. \n\N. r < X n" shows "(\n. inverse (X n)) \ 0" apply (rule filterlim_compose[OF tendsto_inverse_0]) by (metis assms eventually_at_top_linorderI filterlim_at_top_dense filterlim_at_top_imp_at_infinity) text \The sequence \<^term>\1/n\ tends to 0 as \<^term>\n\ tends to infinity.\ lemma LIMSEQ_inverse_real_of_nat: "(\n. inverse (real (Suc n))) \ 0" by (metis filterlim_compose tendsto_inverse_0 filterlim_mono order_refl filterlim_Suc filterlim_compose[OF filterlim_real_sequentially] at_top_le_at_infinity) text \ The sequence \<^term>\r + 1/n\ tends to \<^term>\r\ as \<^term>\n\ tends to infinity is now easily proved. \ lemma LIMSEQ_inverse_real_of_nat_add: "(\n. r + inverse (real (Suc n))) \ r" using tendsto_add [OF tendsto_const LIMSEQ_inverse_real_of_nat] by auto lemma LIMSEQ_inverse_real_of_nat_add_minus: "(\n. r + -inverse (real (Suc n))) \ r" using tendsto_add [OF tendsto_const tendsto_minus [OF LIMSEQ_inverse_real_of_nat]] by auto lemma LIMSEQ_inverse_real_of_nat_add_minus_mult: "(\n. r * (1 + - inverse (real (Suc n)))) \ r" using tendsto_mult [OF tendsto_const LIMSEQ_inverse_real_of_nat_add_minus [of 1]] by auto lemma lim_inverse_n: "((\n. inverse(of_nat n)) \ (0::'a::real_normed_field)) sequentially" using lim_1_over_n by (simp add: inverse_eq_divide) lemma lim_inverse_n': "((\n. 1 / n) \ 0) sequentially" using lim_inverse_n by (simp add: inverse_eq_divide) lemma LIMSEQ_Suc_n_over_n: "(\n. of_nat (Suc n) / of_nat n :: 'a :: real_normed_field) \ 1" proof (rule Lim_transform_eventually) show "eventually (\n. 1 + inverse (of_nat n :: 'a) = of_nat (Suc n) / of_nat n) sequentially" using eventually_gt_at_top[of "0::nat"] by eventually_elim (simp add: field_simps) have "(\n. 1 + inverse (of_nat n) :: 'a) \ 1 + 0" by (intro tendsto_add tendsto_const lim_inverse_n) then show "(\n. 1 + inverse (of_nat n) :: 'a) \ 1" by simp qed lemma LIMSEQ_n_over_Suc_n: "(\n. of_nat n / of_nat (Suc n) :: 'a :: real_normed_field) \ 1" proof (rule Lim_transform_eventually) show "eventually (\n. inverse (of_nat (Suc n) / of_nat n :: 'a) = of_nat n / of_nat (Suc n)) sequentially" using eventually_gt_at_top[of "0::nat"] by eventually_elim (simp add: field_simps del: of_nat_Suc) have "(\n. inverse (of_nat (Suc n) / of_nat n :: 'a)) \ inverse 1" by (intro tendsto_inverse LIMSEQ_Suc_n_over_n) simp_all then show "(\n. inverse (of_nat (Suc n) / of_nat n :: 'a)) \ 1" by simp qed subsection \Convergence on sequences\ lemma convergent_cong: assumes "eventually (\x. f x = g x) sequentially" shows "convergent f \ convergent g" unfolding convergent_def by (subst filterlim_cong[OF refl refl assms]) (rule refl) lemma convergent_Suc_iff: "convergent (\n. f (Suc n)) \ convergent f" by (auto simp: convergent_def filterlim_sequentially_Suc) lemma convergent_ignore_initial_segment: "convergent (\n. f (n + m)) = convergent f" proof (induct m arbitrary: f) case 0 then show ?case by simp next case (Suc m) have "convergent (\n. f (n + Suc m)) \ convergent (\n. f (Suc n + m))" by simp also have "\ \ convergent (\n. f (n + m))" by (rule convergent_Suc_iff) also have "\ \ convergent f" by (rule Suc) finally show ?case . qed lemma convergent_add: fixes X Y :: "nat \ 'a::topological_monoid_add" assumes "convergent (\n. X n)" and "convergent (\n. Y n)" shows "convergent (\n. X n + Y n)" using assms unfolding convergent_def by (blast intro: tendsto_add) lemma convergent_sum: fixes X :: "'a \ nat \ 'b::topological_comm_monoid_add" shows "(\i. i \ A \ convergent (\n. X i n)) \ convergent (\n. \i\A. X i n)" by (induct A rule: infinite_finite_induct) (simp_all add: convergent_const convergent_add) lemma (in bounded_linear) convergent: assumes "convergent (\n. X n)" shows "convergent (\n. f (X n))" using assms unfolding convergent_def by (blast intro: tendsto) lemma (in bounded_bilinear) convergent: assumes "convergent (\n. X n)" and "convergent (\n. Y n)" shows "convergent (\n. X n ** Y n)" using assms unfolding convergent_def by (blast intro: tendsto) lemma convergent_minus_iff: fixes X :: "nat \ 'a::topological_group_add" shows "convergent X \ convergent (\n. - X n)" unfolding convergent_def by (force dest: tendsto_minus) lemma convergent_diff: fixes X Y :: "nat \ 'a::topological_group_add" assumes "convergent (\n. X n)" assumes "convergent (\n. Y n)" shows "convergent (\n. X n - Y n)" using assms unfolding convergent_def by (blast intro: tendsto_diff) lemma convergent_norm: assumes "convergent f" shows "convergent (\n. norm (f n))" proof - from assms have "f \ lim f" by (simp add: convergent_LIMSEQ_iff) then have "(\n. norm (f n)) \ norm (lim f)" by (rule tendsto_norm) then show ?thesis by (auto simp: convergent_def) qed lemma convergent_of_real: "convergent f \ convergent (\n. of_real (f n) :: 'a::real_normed_algebra_1)" unfolding convergent_def by (blast intro!: tendsto_of_real) lemma convergent_add_const_iff: "convergent (\n. c + f n :: 'a::topological_ab_group_add) \ convergent f" proof assume "convergent (\n. c + f n)" from convergent_diff[OF this convergent_const[of c]] show "convergent f" by simp next assume "convergent f" from convergent_add[OF convergent_const[of c] this] show "convergent (\n. c + f n)" by simp qed lemma convergent_add_const_right_iff: "convergent (\n. f n + c :: 'a::topological_ab_group_add) \ convergent f" using convergent_add_const_iff[of c f] by (simp add: add_ac) lemma convergent_diff_const_right_iff: "convergent (\n. f n - c :: 'a::topological_ab_group_add) \ convergent f" using convergent_add_const_right_iff[of f "-c"] by (simp add: add_ac) lemma convergent_mult: fixes X Y :: "nat \ 'a::topological_semigroup_mult" assumes "convergent (\n. X n)" and "convergent (\n. Y n)" shows "convergent (\n. X n * Y n)" using assms unfolding convergent_def by (blast intro: tendsto_mult) lemma convergent_mult_const_iff: assumes "c \ 0" shows "convergent (\n. c * f n :: 'a::{field,topological_semigroup_mult}) \ convergent f" proof assume "convergent (\n. c * f n)" from assms convergent_mult[OF this convergent_const[of "inverse c"]] show "convergent f" by (simp add: field_simps) next assume "convergent f" from convergent_mult[OF convergent_const[of c] this] show "convergent (\n. c * f n)" by simp qed lemma convergent_mult_const_right_iff: fixes c :: "'a::{field,topological_semigroup_mult}" assumes "c \ 0" shows "convergent (\n. f n * c) \ convergent f" using convergent_mult_const_iff[OF assms, of f] by (simp add: mult_ac) lemma convergent_imp_Bseq: "convergent f \ Bseq f" by (simp add: Cauchy_Bseq convergent_Cauchy) text \A monotone sequence converges to its least upper bound.\ lemma LIMSEQ_incseq_SUP: fixes X :: "nat \ 'a::{conditionally_complete_linorder,linorder_topology}" assumes u: "bdd_above (range X)" and X: "incseq X" shows "X \ (SUP i. X i)" by (rule order_tendstoI) (auto simp: eventually_sequentially u less_cSUP_iff intro: X[THEN incseqD] less_le_trans cSUP_lessD[OF u]) lemma LIMSEQ_decseq_INF: fixes X :: "nat \ 'a::{conditionally_complete_linorder, linorder_topology}" assumes u: "bdd_below (range X)" and X: "decseq X" shows "X \ (INF i. X i)" by (rule order_tendstoI) (auto simp: eventually_sequentially u cINF_less_iff intro: X[THEN decseqD] le_less_trans less_cINF_D[OF u]) text \Main monotonicity theorem.\ lemma Bseq_monoseq_convergent: "Bseq X \ monoseq X \ convergent X" for X :: "nat \ real" by (auto simp: monoseq_iff convergent_def intro: LIMSEQ_decseq_INF LIMSEQ_incseq_SUP dest: Bseq_bdd_above Bseq_bdd_below) lemma Bseq_mono_convergent: "Bseq X \ (\m n. m \ n \ X m \ X n) \ convergent X" for X :: "nat \ real" by (auto intro!: Bseq_monoseq_convergent incseq_imp_monoseq simp: incseq_def) lemma monoseq_imp_convergent_iff_Bseq: "monoseq f \ convergent f \ Bseq f" for f :: "nat \ real" using Bseq_monoseq_convergent[of f] convergent_imp_Bseq[of f] by blast lemma Bseq_monoseq_convergent'_inc: fixes f :: "nat \ real" shows "Bseq (\n. f (n + M)) \ (\m n. M \ m \ m \ n \ f m \ f n) \ convergent f" by (subst convergent_ignore_initial_segment [symmetric, of _ M]) (auto intro!: Bseq_monoseq_convergent simp: monoseq_def) lemma Bseq_monoseq_convergent'_dec: fixes f :: "nat \ real" shows "Bseq (\n. f (n + M)) \ (\m n. M \ m \ m \ n \ f m \ f n) \ convergent f" by (subst convergent_ignore_initial_segment [symmetric, of _ M]) (auto intro!: Bseq_monoseq_convergent simp: monoseq_def) lemma Cauchy_iff: "Cauchy X \ (\e>0. \M. \m\M. \n\M. norm (X m - X n) < e)" for X :: "nat \ 'a::real_normed_vector" unfolding Cauchy_def dist_norm .. lemma CauchyI: "(\e. 0 < e \ \M. \m\M. \n\M. norm (X m - X n) < e) \ Cauchy X" for X :: "nat \ 'a::real_normed_vector" by (simp add: Cauchy_iff) lemma CauchyD: "Cauchy X \ 0 < e \ \M. \m\M. \n\M. norm (X m - X n) < e" for X :: "nat \ 'a::real_normed_vector" by (simp add: Cauchy_iff) lemma incseq_convergent: fixes X :: "nat \ real" assumes "incseq X" and "\i. X i \ B" obtains L where "X \ L" "\i. X i \ L" proof atomize_elim from incseq_bounded[OF assms] \incseq X\ Bseq_monoseq_convergent[of X] obtain L where "X \ L" by (auto simp: convergent_def monoseq_def incseq_def) with \incseq X\ show "\L. X \ L \ (\i. X i \ L)" by (auto intro!: exI[of _ L] incseq_le) qed lemma decseq_convergent: fixes X :: "nat \ real" assumes "decseq X" and "\i. B \ X i" obtains L where "X \ L" "\i. L \ X i" proof atomize_elim from decseq_bounded[OF assms] \decseq X\ Bseq_monoseq_convergent[of X] obtain L where "X \ L" by (auto simp: convergent_def monoseq_def decseq_def) with \decseq X\ show "\L. X \ L \ (\i. L \ X i)" by (auto intro!: exI[of _ L] decseq_ge) qed lemma monoseq_convergent: fixes X :: "nat \ real" assumes X: "monoseq X" and B: "\i. \X i\ \ B" obtains L where "X \ L" using X unfolding monoseq_iff proof assume "incseq X" show thesis using abs_le_D1 [OF B] incseq_convergent [OF \incseq X\] that by meson next assume "decseq X" show thesis using decseq_convergent [OF \decseq X\] that by (metis B abs_le_iff add.inverse_inverse neg_le_iff_le) qed subsection \Power Sequences\ lemma Bseq_realpow: "0 \ x \ x \ 1 \ Bseq (\n. x ^ n)" for x :: real by (metis decseq_bounded decseq_def power_decreasing zero_le_power) lemma monoseq_realpow: "0 \ x \ x \ 1 \ monoseq (\n. x ^ n)" for x :: real using monoseq_def power_decreasing by blast lemma convergent_realpow: "0 \ x \ x \ 1 \ convergent (\n. x ^ n)" for x :: real by (blast intro!: Bseq_monoseq_convergent Bseq_realpow monoseq_realpow) lemma LIMSEQ_inverse_realpow_zero: "1 < x \ (\n. inverse (x ^ n)) \ 0" for x :: real by (rule filterlim_compose[OF tendsto_inverse_0 filterlim_realpow_sequentially_gt1]) simp lemma LIMSEQ_realpow_zero: fixes x :: real assumes "0 \ x" "x < 1" shows "(\n. x ^ n) \ 0" proof (cases "x = 0") case False with \0 \ x\ have x0: "0 < x" by simp then have "1 < inverse x" using \x < 1\ by (rule one_less_inverse) then have "(\n. inverse (inverse x ^ n)) \ 0" by (rule LIMSEQ_inverse_realpow_zero) then show ?thesis by (simp add: power_inverse) next case True show ?thesis by (rule LIMSEQ_imp_Suc) (simp add: True) qed lemma LIMSEQ_power_zero [tendsto_intros]: "norm x < 1 \ (\n. x ^ n) \ 0" for x :: "'a::real_normed_algebra_1" apply (drule LIMSEQ_realpow_zero [OF norm_ge_zero]) by (simp add: Zfun_le norm_power_ineq tendsto_Zfun_iff) lemma LIMSEQ_divide_realpow_zero: "1 < x \ (\n. a / (x ^ n) :: real) \ 0" by (rule tendsto_divide_0 [OF tendsto_const filterlim_realpow_sequentially_gt1]) simp lemma tendsto_power_zero: fixes x::"'a::real_normed_algebra_1" assumes "filterlim f at_top F" assumes "norm x < 1" shows "((\y. x ^ (f y)) \ 0) F" proof (rule tendstoI) fix e::real assume "0 < e" from tendstoD[OF LIMSEQ_power_zero[OF \norm x < 1\] \0 < e\] have "\\<^sub>F xa in sequentially. norm (x ^ xa) < e" by simp then obtain N where N: "norm (x ^ n) < e" if "n \ N" for n by (auto simp: eventually_sequentially) have "\\<^sub>F i in F. f i \ N" using \filterlim f sequentially F\ by (simp add: filterlim_at_top) then show "\\<^sub>F i in F. dist (x ^ f i) 0 < e" by eventually_elim (auto simp: N) qed text \Limit of \<^term>\c^n\ for \<^term>\\c\ < 1\.\ lemma LIMSEQ_abs_realpow_zero: "\c\ < 1 \ (\n. \c\ ^ n :: real) \ 0" by (rule LIMSEQ_realpow_zero [OF abs_ge_zero]) lemma LIMSEQ_abs_realpow_zero2: "\c\ < 1 \ (\n. c ^ n :: real) \ 0" by (rule LIMSEQ_power_zero) simp subsection \Limits of Functions\ lemma LIM_eq: "f \a\ L = (\r>0. \s>0. \x. x \ a \ norm (x - a) < s \ norm (f x - L) < r)" for a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector" by (simp add: LIM_def dist_norm) lemma LIM_I: "(\r. 0 < r \ \s>0. \x. x \ a \ norm (x - a) < s \ norm (f x - L) < r) \ f \a\ L" for a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector" by (simp add: LIM_eq) lemma LIM_D: "f \a\ L \ 0 < r \ \s>0.\x. x \ a \ norm (x - a) < s \ norm (f x - L) < r" for a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector" by (simp add: LIM_eq) lemma LIM_offset: "f \a\ L \ (\x. f (x + k)) \(a - k)\ L" for a :: "'a::real_normed_vector" by (simp add: filtermap_at_shift[symmetric, of a k] filterlim_def filtermap_filtermap) lemma LIM_offset_zero: "f \a\ L \ (\h. f (a + h)) \0\ L" for a :: "'a::real_normed_vector" by (drule LIM_offset [where k = a]) (simp add: add.commute) lemma LIM_offset_zero_cancel: "(\h. f (a + h)) \0\ L \ f \a\ L" for a :: "'a::real_normed_vector" by (drule LIM_offset [where k = "- a"]) simp lemma LIM_offset_zero_iff: "NO_MATCH 0 a \ f \a\ L \ (\h. f (a + h)) \0\ L" for f :: "'a :: real_normed_vector \ _" using LIM_offset_zero_cancel[of f a L] LIM_offset_zero[of f L a] by auto lemma tendsto_offset_zero_iff: fixes f :: "'a :: real_normed_vector \ _" assumes " NO_MATCH 0 a" "a \ S" "open S" shows "(f \ L) (at a within S) \ ((\h. f (a + h)) \ L) (at 0)" using assms by (simp add: tendsto_within_open_NO_MATCH LIM_offset_zero_iff) lemma LIM_zero: "(f \ l) F \ ((\x. f x - l) \ 0) F" for f :: "'a \ 'b::real_normed_vector" unfolding tendsto_iff dist_norm by simp lemma LIM_zero_cancel: fixes f :: "'a \ 'b::real_normed_vector" shows "((\x. f x - l) \ 0) F \ (f \ l) F" unfolding tendsto_iff dist_norm by simp lemma LIM_zero_iff: "((\x. f x - l) \ 0) F = (f \ l) F" for f :: "'a \ 'b::real_normed_vector" unfolding tendsto_iff dist_norm by simp lemma LIM_imp_LIM: fixes f :: "'a::topological_space \ 'b::real_normed_vector" fixes g :: "'a::topological_space \ 'c::real_normed_vector" assumes f: "f \a\ l" and le: "\x. x \ a \ norm (g x - m) \ norm (f x - l)" shows "g \a\ m" by (rule metric_LIM_imp_LIM [OF f]) (simp add: dist_norm le) lemma LIM_equal2: fixes f g :: "'a::real_normed_vector \ 'b::topological_space" assumes "0 < R" and "\x. x \ a \ norm (x - a) < R \ f x = g x" shows "g \a\ l \ f \a\ l" by (rule metric_LIM_equal2 [OF _ assms]) (simp_all add: dist_norm) lemma LIM_compose2: fixes a :: "'a::real_normed_vector" assumes f: "f \a\ b" and g: "g \b\ c" and inj: "\d>0. \x. x \ a \ norm (x - a) < d \ f x \ b" shows "(\x. g (f x)) \a\ c" by (rule metric_LIM_compose2 [OF f g inj [folded dist_norm]]) lemma real_LIM_sandwich_zero: fixes f g :: "'a::topological_space \ real" assumes f: "f \a\ 0" and 1: "\x. x \ a \ 0 \ g x" and 2: "\x. x \ a \ g x \ f x" shows "g \a\ 0" proof (rule LIM_imp_LIM [OF f]) (* FIXME: use tendsto_sandwich *) fix x assume x: "x \ a" with 1 have "norm (g x - 0) = g x" by simp also have "g x \ f x" by (rule 2 [OF x]) also have "f x \ \f x\" by (rule abs_ge_self) also have "\f x\ = norm (f x - 0)" by simp finally show "norm (g x - 0) \ norm (f x - 0)" . qed subsection \Continuity\ lemma LIM_isCont_iff: "(f \a\ f a) = ((\h. f (a + h)) \0\ f a)" for f :: "'a::real_normed_vector \ 'b::topological_space" by (rule iffI [OF LIM_offset_zero LIM_offset_zero_cancel]) lemma isCont_iff: "isCont f x = (\h. f (x + h)) \0\ f x" for f :: "'a::real_normed_vector \ 'b::topological_space" by (simp add: isCont_def LIM_isCont_iff) lemma isCont_LIM_compose2: fixes a :: "'a::real_normed_vector" assumes f [unfolded isCont_def]: "isCont f a" and g: "g \f a\ l" and inj: "\d>0. \x. x \ a \ norm (x - a) < d \ f x \ f a" shows "(\x. g (f x)) \a\ l" by (rule LIM_compose2 [OF f g inj]) lemma isCont_norm [simp]: "isCont f a \ isCont (\x. norm (f x)) a" for f :: "'a::t2_space \ 'b::real_normed_vector" by (fact continuous_norm) lemma isCont_rabs [simp]: "isCont f a \ isCont (\x. \f x\) a" for f :: "'a::t2_space \ real" by (fact continuous_rabs) lemma isCont_add [simp]: "isCont f a \ isCont g a \ isCont (\x. f x + g x) a" for f :: "'a::t2_space \ 'b::topological_monoid_add" by (fact continuous_add) lemma isCont_minus [simp]: "isCont f a \ isCont (\x. - f x) a" for f :: "'a::t2_space \ 'b::real_normed_vector" by (fact continuous_minus) lemma isCont_diff [simp]: "isCont f a \ isCont g a \ isCont (\x. f x - g x) a" for f :: "'a::t2_space \ 'b::real_normed_vector" by (fact continuous_diff) lemma isCont_mult [simp]: "isCont f a \ isCont g a \ isCont (\x. f x * g x) a" for f g :: "'a::t2_space \ 'b::real_normed_algebra" by (fact continuous_mult) lemma (in bounded_linear) isCont: "isCont g a \ isCont (\x. f (g x)) a" by (fact continuous) lemma (in bounded_bilinear) isCont: "isCont f a \ isCont g a \ isCont (\x. f x ** g x) a" by (fact continuous) lemmas isCont_scaleR [simp] = bounded_bilinear.isCont [OF bounded_bilinear_scaleR] lemmas isCont_of_real [simp] = bounded_linear.isCont [OF bounded_linear_of_real] lemma isCont_power [simp]: "isCont f a \ isCont (\x. f x ^ n) a" for f :: "'a::t2_space \ 'b::{power,real_normed_algebra}" by (fact continuous_power) lemma isCont_sum [simp]: "\i\A. isCont (f i) a \ isCont (\x. \i\A. f i x) a" for f :: "'a \ 'b::t2_space \ 'c::topological_comm_monoid_add" by (auto intro: continuous_sum) subsection \Uniform Continuity\ lemma uniformly_continuous_on_def: fixes f :: "'a::metric_space \ 'b::metric_space" shows "uniformly_continuous_on s f \ (\e>0. \d>0. \x\s. \x'\s. dist x' x < d \ dist (f x') (f x) < e)" unfolding uniformly_continuous_on_uniformity uniformity_dist filterlim_INF filterlim_principal eventually_inf_principal by (force simp: Ball_def uniformity_dist[symmetric] eventually_uniformity_metric) abbreviation isUCont :: "['a::metric_space \ 'b::metric_space] \ bool" where "isUCont f \ uniformly_continuous_on UNIV f" lemma isUCont_def: "isUCont f \ (\r>0. \s>0. \x y. dist x y < s \ dist (f x) (f y) < r)" by (auto simp: uniformly_continuous_on_def dist_commute) lemma isUCont_isCont: "isUCont f \ isCont f x" by (drule uniformly_continuous_imp_continuous) (simp add: continuous_on_eq_continuous_at) lemma uniformly_continuous_on_Cauchy: fixes f :: "'a::metric_space \ 'b::metric_space" assumes "uniformly_continuous_on S f" "Cauchy X" "\n. X n \ S" shows "Cauchy (\n. f (X n))" using assms unfolding uniformly_continuous_on_def by (meson Cauchy_def) lemma isUCont_Cauchy: "isUCont f \ Cauchy X \ Cauchy (\n. f (X n))" by (rule uniformly_continuous_on_Cauchy[where S=UNIV and f=f]) simp_all lemma uniformly_continuous_imp_Cauchy_continuous: fixes f :: "'a::metric_space \ 'b::metric_space" shows "\uniformly_continuous_on S f; Cauchy \; \n. (\ n) \ S\ \ Cauchy(f \ \)" by (simp add: uniformly_continuous_on_def Cauchy_def) meson lemma (in bounded_linear) isUCont: "isUCont f" unfolding isUCont_def dist_norm proof (intro allI impI) fix r :: real assume r: "0 < r" obtain K where K: "0 < K" and norm_le: "norm (f x) \ norm x * K" for x using pos_bounded by blast show "\s>0. \x y. norm (x - y) < s \ norm (f x - f y) < r" proof (rule exI, safe) from r K show "0 < r / K" by simp next fix x y :: 'a assume xy: "norm (x - y) < r / K" have "norm (f x - f y) = norm (f (x - y))" by (simp only: diff) also have "\ \ norm (x - y) * K" by (rule norm_le) also from K xy have "\ < r" by (simp only: pos_less_divide_eq) finally show "norm (f x - f y) < r" . qed qed lemma (in bounded_linear) Cauchy: "Cauchy X \ Cauchy (\n. f (X n))" by (rule isUCont [THEN isUCont_Cauchy]) lemma LIM_less_bound: fixes f :: "real \ real" assumes ev: "b < x" "\ x' \ { b <..< x}. 0 \ f x'" and "isCont f x" shows "0 \ f x" proof (rule tendsto_lowerbound) show "(f \ f x) (at_left x)" using \isCont f x\ by (simp add: filterlim_at_split isCont_def) show "eventually (\x. 0 \ f x) (at_left x)" using ev by (auto simp: eventually_at dist_real_def intro!: exI[of _ "x - b"]) qed simp subsection \Nested Intervals and Bisection -- Needed for Compactness\ lemma nested_sequence_unique: assumes "\n. f n \ f (Suc n)" "\n. g (Suc n) \ g n" "\n. f n \ g n" "(\n. f n - g n) \ 0" shows "\l::real. ((\n. f n \ l) \ f \ l) \ ((\n. l \ g n) \ g \ l)" proof - have "incseq f" unfolding incseq_Suc_iff by fact have "decseq g" unfolding decseq_Suc_iff by fact have "f n \ g 0" for n proof - from \decseq g\ have "g n \ g 0" by (rule decseqD) simp with \\n. f n \ g n\[THEN spec, of n] show ?thesis by auto qed then obtain u where "f \ u" "\i. f i \ u" using incseq_convergent[OF \incseq f\] by auto moreover have "f 0 \ g n" for n proof - from \incseq f\ have "f 0 \ f n" by (rule incseqD) simp with \\n. f n \ g n\[THEN spec, of n] show ?thesis by simp qed then obtain l where "g \ l" "\i. l \ g i" using decseq_convergent[OF \decseq g\] by auto moreover note LIMSEQ_unique[OF assms(4) tendsto_diff[OF \f \ u\ \g \ l\]] ultimately show ?thesis by auto qed lemma Bolzano[consumes 1, case_names trans local]: fixes P :: "real \ real \ bool" assumes [arith]: "a \ b" and trans: "\a b c. P a b \ P b c \ a \ b \ b \ c \ P a c" and local: "\x. a \ x \ x \ b \ \d>0. \a b. a \ x \ x \ b \ b - a < d \ P a b" shows "P a b" proof - define bisect where "bisect = rec_nat (a, b) (\n (x, y). if P x ((x+y) / 2) then ((x+y)/2, y) else (x, (x+y)/2))" define l u where "l n = fst (bisect n)" and "u n = snd (bisect n)" for n have l[simp]: "l 0 = a" "\n. l (Suc n) = (if P (l n) ((l n + u n) / 2) then (l n + u n) / 2 else l n)" and u[simp]: "u 0 = b" "\n. u (Suc n) = (if P (l n) ((l n + u n) / 2) then u n else (l n + u n) / 2)" by (simp_all add: l_def u_def bisect_def split: prod.split) have [simp]: "l n \ u n" for n by (induct n) auto have "\x. ((\n. l n \ x) \ l \ x) \ ((\n. x \ u n) \ u \ x)" proof (safe intro!: nested_sequence_unique) show "l n \ l (Suc n)" "u (Suc n) \ u n" for n by (induct n) auto next have "l n - u n = (a - b) / 2^n" for n by (induct n) (auto simp: field_simps) then show "(\n. l n - u n) \ 0" by (simp add: LIMSEQ_divide_realpow_zero) qed fact then obtain x where x: "\n. l n \ x" "\n. x \ u n" and "l \ x" "u \ x" by auto obtain d where "0 < d" and d: "a \ x \ x \ b \ b - a < d \ P a b" for a b using \l 0 \ x\ \x \ u 0\ local[of x] by auto show "P a b" proof (rule ccontr) assume "\ P a b" have "\ P (l n) (u n)" for n proof (induct n) case 0 then show ?case by (simp add: \\ P a b\) next case (Suc n) with trans[of "l n" "(l n + u n) / 2" "u n"] show ?case by auto qed moreover { have "eventually (\n. x - d / 2 < l n) sequentially" using \0 < d\ \l \ x\ by (intro order_tendstoD[of _ x]) auto moreover have "eventually (\n. u n < x + d / 2) sequentially" using \0 < d\ \u \ x\ by (intro order_tendstoD[of _ x]) auto ultimately have "eventually (\n. P (l n) (u n)) sequentially" proof eventually_elim case (elim n) from add_strict_mono[OF this] have "u n - l n < d" by simp with x show "P (l n) (u n)" by (rule d) qed } ultimately show False by simp qed qed lemma compact_Icc[simp, intro]: "compact {a .. b::real}" proof (cases "a \ b", rule compactI) fix C assume C: "a \ b" "\t\C. open t" "{a..b} \ \C" define T where "T = {a .. b}" from C(1,3) show "\C'\C. finite C' \ {a..b} \ \C'" proof (induct rule: Bolzano) case (trans a b c) then have *: "{a..c} = {a..b} \ {b..c}" by auto with trans obtain C1 C2 where "C1\C" "finite C1" "{a..b} \ \C1" "C2\C" "finite C2" "{b..c} \ \C2" by auto with trans show ?case unfolding * by (intro exI[of _ "C1 \ C2"]) auto next case (local x) with C have "x \ \C" by auto with C(2) obtain c where "x \ c" "open c" "c \ C" by auto then obtain e where "0 < e" "{x - e <..< x + e} \ c" by (auto simp: open_dist dist_real_def subset_eq Ball_def abs_less_iff) with \c \ C\ show ?case by (safe intro!: exI[of _ "e/2"] exI[of _ "{c}"]) auto qed qed simp lemma continuous_image_closed_interval: fixes a b and f :: "real \ real" defines "S \ {a..b}" assumes "a \ b" and f: "continuous_on S f" shows "\c d. f`S = {c..d} \ c \ d" proof - have S: "compact S" "S \ {}" using \a \ b\ by (auto simp: S_def) obtain c where "c \ S" "\d\S. f d \ f c" using continuous_attains_sup[OF S f] by auto moreover obtain d where "d \ S" "\c\S. f d \ f c" using continuous_attains_inf[OF S f] by auto moreover have "connected (f`S)" using connected_continuous_image[OF f] connected_Icc by (auto simp: S_def) ultimately have "f ` S = {f d .. f c} \ f d \ f c" by (auto simp: connected_iff_interval) then show ?thesis by auto qed lemma open_Collect_positive: fixes f :: "'a::topological_space \ real" assumes f: "continuous_on s f" shows "\A. open A \ A \ s = {x\s. 0 < f x}" using continuous_on_open_invariant[THEN iffD1, OF f, rule_format, of "{0 <..}"] by (auto simp: Int_def field_simps) lemma open_Collect_less_Int: fixes f g :: "'a::topological_space \ real" assumes f: "continuous_on s f" and g: "continuous_on s g" shows "\A. open A \ A \ s = {x\s. f x < g x}" using open_Collect_positive[OF continuous_on_diff[OF g f]] by (simp add: field_simps) subsection \Boundedness of continuous functions\ text\By bisection, function continuous on closed interval is bounded above\ lemma isCont_eq_Ub: fixes f :: "real \ 'a::linorder_topology" shows "a \ b \ \x::real. a \ x \ x \ b \ isCont f x \ \M. (\x. a \ x \ x \ b \ f x \ M) \ (\x. a \ x \ x \ b \ f x = M)" using continuous_attains_sup[of "{a..b}" f] by (auto simp: continuous_at_imp_continuous_on Ball_def Bex_def) lemma isCont_eq_Lb: fixes f :: "real \ 'a::linorder_topology" shows "a \ b \ \x. a \ x \ x \ b \ isCont f x \ \M. (\x. a \ x \ x \ b \ M \ f x) \ (\x. a \ x \ x \ b \ f x = M)" using continuous_attains_inf[of "{a..b}" f] by (auto simp: continuous_at_imp_continuous_on Ball_def Bex_def) lemma isCont_bounded: fixes f :: "real \ 'a::linorder_topology" shows "a \ b \ \x. a \ x \ x \ b \ isCont f x \ \M. \x. a \ x \ x \ b \ f x \ M" using isCont_eq_Ub[of a b f] by auto lemma isCont_has_Ub: fixes f :: "real \ 'a::linorder_topology" shows "a \ b \ \x. a \ x \ x \ b \ isCont f x \ \M. (\x. a \ x \ x \ b \ f x \ M) \ (\N. N < M \ (\x. a \ x \ x \ b \ N < f x))" using isCont_eq_Ub[of a b f] by auto lemma isCont_Lb_Ub: fixes f :: "real \ real" assumes "a \ b" "\x. a \ x \ x \ b \ isCont f x" shows "\L M. (\x. a \ x \ x \ b \ L \ f x \ f x \ M) \ (\y. L \ y \ y \ M \ (\x. a \ x \ x \ b \ (f x = y)))" proof - obtain M where M: "a \ M" "M \ b" "\x. a \ x \ x \ b \ f x \ f M" using isCont_eq_Ub[OF assms] by auto obtain L where L: "a \ L" "L \ b" "\x. a \ x \ x \ b \ f L \ f x" using isCont_eq_Lb[OF assms] by auto have "(\x. a \ x \ x \ b \ f L \ f x \ f x \ f M)" using M L by simp moreover have "(\y. f L \ y \ y \ f M \ (\x\a. x \ b \ f x = y))" proof (cases "L \ M") case True then show ?thesis using IVT[of f L _ M] M L assms by (metis order.trans) next case False then show ?thesis using IVT2[of f L _ M] by (metis L(2) M(1) assms(2) le_cases order.trans) qed ultimately show ?thesis by blast qed text \Continuity of inverse function.\ lemma isCont_inverse_function: fixes f g :: "real \ real" assumes d: "0 < d" and inj: "\z. \z-x\ \ d \ g (f z) = z" and cont: "\z. \z-x\ \ d \ isCont f z" shows "isCont g (f x)" proof - let ?A = "f (x - d)" let ?B = "f (x + d)" let ?D = "{x - d..x + d}" have f: "continuous_on ?D f" using cont by (intro continuous_at_imp_continuous_on ballI) auto then have g: "continuous_on (f`?D) g" using inj by (intro continuous_on_inv) auto from d f have "{min ?A ?B <..< max ?A ?B} \ f ` ?D" by (intro connected_contains_Ioo connected_continuous_image) (auto split: split_min split_max) with g have "continuous_on {min ?A ?B <..< max ?A ?B} g" by (rule continuous_on_subset) moreover have "(?A < f x \ f x < ?B) \ (?B < f x \ f x < ?A)" using d inj by (intro continuous_inj_imp_mono[OF _ _ f] inj_on_imageI2[of g, OF inj_onI]) auto then have "f x \ {min ?A ?B <..< max ?A ?B}" by auto ultimately show ?thesis by (simp add: continuous_on_eq_continuous_at) qed lemma isCont_inverse_function2: fixes f g :: "real \ real" shows "\a < x; x < b; \z. \a \ z; z \ b\ \ g (f z) = z; \z. \a \ z; z \ b\ \ isCont f z\ \ isCont g (f x)" apply (rule isCont_inverse_function [where f=f and d="min (x - a) (b - x)"]) apply (simp_all add: abs_le_iff) done text \Bartle/Sherbert: Introduction to Real Analysis, Theorem 4.2.9, p. 110.\ lemma LIM_fun_gt_zero: "f \c\ l \ 0 < l \ \r. 0 < r \ (\x. x \ c \ \c - x\ < r \ 0 < f x)" for f :: "real \ real" by (force simp: dest: LIM_D) lemma LIM_fun_less_zero: "f \c\ l \ l < 0 \ \r. 0 < r \ (\x. x \ c \ \c - x\ < r \ f x < 0)" for f :: "real \ real" by (drule LIM_D [where r="-l"]) force+ lemma LIM_fun_not_zero: "f \c\ l \ l \ 0 \ \r. 0 < r \ (\x. x \ c \ \c - x\ < r \ f x \ 0)" for f :: "real \ real" using LIM_fun_gt_zero[of f l c] LIM_fun_less_zero[of f l c] by (auto simp: neq_iff) end