diff --git a/src/HOL/Analysis/Further_Topology.thy b/src/HOL/Analysis/Further_Topology.thy --- a/src/HOL/Analysis/Further_Topology.thy +++ b/src/HOL/Analysis/Further_Topology.thy @@ -1,5706 +1,5573 @@ section \Extending Continous Maps, Invariance of Domain, etc\ (*FIX rename? *) text\Ported from HOL Light (moretop.ml) by L C Paulson\ theory Further_Topology imports Weierstrass_Theorems Polytope Complex_Transcendental Equivalence_Lebesgue_Henstock_Integration Retracts begin subsection\A map from a sphere to a higher dimensional sphere is nullhomotopic\ lemma spheremap_lemma1: fixes f :: "'a::euclidean_space \ 'a::euclidean_space" assumes "subspace S" "subspace T" and dimST: "dim S < dim T" and "S \ T" and diff_f: "f differentiable_on sphere 0 1 \ S" shows "f ` (sphere 0 1 \ S) \ sphere 0 1 \ T" proof assume fim: "f ` (sphere 0 1 \ S) = sphere 0 1 \ T" have inS: "\x. \x \ S; x \ 0\ \ (x /\<^sub>R norm x) \ S" using subspace_mul \subspace S\ by blast have subS01: "(\x. x /\<^sub>R norm x) ` (S - {0}) \ sphere 0 1 \ S" using \subspace S\ subspace_mul by fastforce then have diff_f': "f differentiable_on (\x. x /\<^sub>R norm x) ` (S - {0})" by (rule differentiable_on_subset [OF diff_f]) define g where "g \ \x. norm x *\<^sub>R f(inverse(norm x) *\<^sub>R x)" have gdiff: "g differentiable_on S - {0}" unfolding g_def by (rule diff_f' derivative_intros differentiable_on_compose [where f=f] | force)+ have geq: "g ` (S - {0}) = T - {0}" proof have "\u. \u \ S; norm u *\<^sub>R f (u /\<^sub>R norm u) \ T\ \ u = 0" by (metis (mono_tags, lifting) DiffI subS01 subspace_mul [OF \subspace T\] fim image_subset_iff inf_le2 singletonD) then have "g ` (S - {0}) \ T" using g_def by blast moreover have "g ` (S - {0}) \ UNIV - {0}" proof (clarsimp simp: g_def) fix y assume "y \ S" and f0: "f (y /\<^sub>R norm y) = 0" then have "y \ 0 \ y /\<^sub>R norm y \ sphere 0 1 \ S" by (auto simp: subspace_mul [OF \subspace S\]) then show "y = 0" by (metis fim f0 Int_iff image_iff mem_sphere_0 norm_eq_zero zero_neq_one) qed ultimately show "g ` (S - {0}) \ T - {0}" by auto next have *: "sphere 0 1 \ T \ f ` (sphere 0 1 \ S)" using fim by (simp add: image_subset_iff) have "x \ (\x. norm x *\<^sub>R f (x /\<^sub>R norm x)) ` (S - {0})" if "x \ T" "x \ 0" for x proof - have "x /\<^sub>R norm x \ T" using \subspace T\ subspace_mul that by blast then obtain u where u: "f u \ T" "x /\<^sub>R norm x = f u" "norm u = 1" "u \ S" using * [THEN subsetD, of "x /\<^sub>R norm x"] \x \ 0\ by auto with that have [simp]: "norm x *\<^sub>R f u = x" by (metis divideR_right norm_eq_zero) moreover have "norm x *\<^sub>R u \ S - {0}" using \subspace S\ subspace_scale that(2) u by auto with u show ?thesis by (simp add: image_eqI [where x="norm x *\<^sub>R u"]) qed then have "T - {0} \ (\x. norm x *\<^sub>R f (x /\<^sub>R norm x)) ` (S - {0})" by force then show "T - {0} \ g ` (S - {0})" by (simp add: g_def) qed define T' where "T' \ {y. \x \ T. orthogonal x y}" have "subspace T'" by (simp add: subspace_orthogonal_to_vectors T'_def) have dim_eq: "dim T' + dim T = DIM('a)" using dim_subspace_orthogonal_to_vectors [of T UNIV] \subspace T\ by (simp add: T'_def) have "\v1 v2. v1 \ span T \ (\w \ span T. orthogonal v2 w) \ x = v1 + v2" for x by (force intro: orthogonal_subspace_decomp_exists [of T x]) then obtain p1 p2 where p1span: "p1 x \ span T" and "\w. w \ span T \ orthogonal (p2 x) w" and eq: "p1 x + p2 x = x" for x by metis then have p1: "\z. p1 z \ T" and ortho: "\w. w \ T \ orthogonal (p2 x) w" for x using span_eq_iff \subspace T\ by blast+ then have p2: "\z. p2 z \ T'" by (simp add: T'_def orthogonal_commute) have p12_eq: "\x y. \x \ T; y \ T'\ \ p1(x + y) = x \ p2(x + y) = y" proof (rule orthogonal_subspace_decomp_unique [OF eq p1span, where T=T']) show "\x y. \x \ T; y \ T'\ \ p2 (x + y) \ span T'" using span_eq_iff p2 \subspace T'\ by blast show "\a b. \a \ T; b \ T'\ \ orthogonal a b" using T'_def by blast qed (auto simp: span_base) then have "\c x. p1 (c *\<^sub>R x) = c *\<^sub>R p1 x \ p2 (c *\<^sub>R x) = c *\<^sub>R p2 x" proof - fix c :: real and x :: 'a have f1: "c *\<^sub>R x = c *\<^sub>R p1 x + c *\<^sub>R p2 x" by (metis eq pth_6) have f2: "c *\<^sub>R p2 x \ T'" by (simp add: \subspace T'\ p2 subspace_scale) have "c *\<^sub>R p1 x \ T" by (metis (full_types) assms(2) p1span span_eq_iff subspace_scale) then show "p1 (c *\<^sub>R x) = c *\<^sub>R p1 x \ p2 (c *\<^sub>R x) = c *\<^sub>R p2 x" using f2 f1 p12_eq by presburger qed moreover have lin_add: "\x y. p1 (x + y) = p1 x + p1 y \ p2 (x + y) = p2 x + p2 y" proof (rule orthogonal_subspace_decomp_unique [OF _ p1span, where T=T']) show "\x y. p1 (x + y) + p2 (x + y) = p1 x + p1 y + (p2 x + p2 y)" by (simp add: add.assoc add.left_commute eq) show "\a b. \a \ T; b \ T'\ \ orthogonal a b" using T'_def by blast qed (auto simp: p1span p2 span_base span_add) ultimately have "linear p1" "linear p2" by unfold_locales auto have "g differentiable_on p1 ` {x + y |x y. x \ S - {0} \ y \ T'}" using p12_eq \S \ T\ by (force intro: differentiable_on_subset [OF gdiff]) then have "(\z. g (p1 z)) differentiable_on {x + y |x y. x \ S - {0} \ y \ T'}" by (rule differentiable_on_compose [OF linear_imp_differentiable_on [OF \linear p1\]]) then have diff: "(\x. g (p1 x) + p2 x) differentiable_on {x + y |x y. x \ S - {0} \ y \ T'}" by (intro derivative_intros linear_imp_differentiable_on [OF \linear p2\]) have "dim {x + y |x y. x \ S - {0} \ y \ T'} \ dim {x + y |x y. x \ S \ y \ T'}" by (blast intro: dim_subset) also have "... = dim S + dim T' - dim (S \ T')" using dim_sums_Int [OF \subspace S\ \subspace T'\] by (simp add: algebra_simps) also have "... < DIM('a)" using dimST dim_eq by auto finally have neg: "negligible {x + y |x y. x \ S - {0} \ y \ T'}" by (rule negligible_lowdim) have "negligible ((\x. g (p1 x) + p2 x) ` {x + y |x y. x \ S - {0} \ y \ T'})" by (rule negligible_differentiable_image_negligible [OF order_refl neg diff]) then have "negligible {x + y |x y. x \ g ` (S - {0}) \ y \ T'}" proof (rule negligible_subset) have "\t' \ T'; s \ S; s \ 0\ \ g s + t' \ (\x. g (p1 x) + p2 x) ` {x + t' |x t'. x \ S \ x \ 0 \ t' \ T'}" for t' s - apply (rule_tac x="s + t'" in image_eqI) - using \S \ T\ p12_eq by auto + using \S \ T\ p12_eq by (rule_tac x="s + t'" in image_eqI) auto then show "{x + y |x y. x \ g ` (S - {0}) \ y \ T'} \ (\x. g (p1 x) + p2 x) ` {x + y |x y. x \ S - {0} \ y \ T'}" by auto qed moreover have "- T' \ {x + y |x y. x \ g ` (S - {0}) \ y \ T'}" proof clarsimp fix z assume "z \ T'" show "\x y. z = x + y \ x \ g ` (S - {0}) \ y \ T'" by (metis Diff_iff \z \ T'\ add.left_neutral eq geq p1 p2 singletonD) qed ultimately have "negligible (-T')" using negligible_subset by blast moreover have "negligible T'" using negligible_lowdim by (metis add.commute assms(3) diff_add_inverse2 diff_self_eq_0 dim_eq le_add1 le_antisym linordered_semidom_class.add_diff_inverse not_less0) ultimately have "negligible (-T' \ T')" by (metis negligible_Un_eq) then show False using negligible_Un_eq non_negligible_UNIV by simp qed lemma spheremap_lemma2: fixes f :: "'a::euclidean_space \ 'a::euclidean_space" assumes ST: "subspace S" "subspace T" "dim S < dim T" and "S \ T" and contf: "continuous_on (sphere 0 1 \ S) f" and fim: "f ` (sphere 0 1 \ S) \ sphere 0 1 \ T" shows "\c. homotopic_with_canon (\x. True) (sphere 0 1 \ S) (sphere 0 1 \ T) f (\x. c)" proof - have [simp]: "\x. \norm x = 1; x \ S\ \ norm (f x) = 1" using fim by (simp add: image_subset_iff) have "compact (sphere 0 1 \ S)" by (simp add: \subspace S\ closed_subspace compact_Int_closed) then obtain g where pfg: "polynomial_function g" and gim: "g ` (sphere 0 1 \ S) \ T" and g12: "\x. x \ sphere 0 1 \ S \ norm(f x - g x) < 1/2" apply (rule Stone_Weierstrass_polynomial_function_subspace [OF _ contf _ \subspace T\, of "1/2"]) using fim by auto have gnz: "g x \ 0" if "x \ sphere 0 1 \ S" for x proof - have "norm (f x) = 1" using fim that by (simp add: image_subset_iff) then show ?thesis using g12 [OF that] by auto qed have diffg: "g differentiable_on sphere 0 1 \ S" by (metis pfg differentiable_on_polynomial_function) define h where "h \ \x. inverse(norm(g x)) *\<^sub>R g x" have h: "x \ sphere 0 1 \ S \ h x \ sphere 0 1 \ T" for x unfolding h_def using gnz [of x] by (auto simp: subspace_mul [OF \subspace T\] subsetD [OF gim]) have diffh: "h differentiable_on sphere 0 1 \ S" unfolding h_def using gnz by (fastforce intro: derivative_intros diffg differentiable_on_compose [OF diffg]) have homfg: "homotopic_with_canon (\z. True) (sphere 0 1 \ S) (T - {0}) f g" proof (rule homotopic_with_linear [OF contf]) show "continuous_on (sphere 0 1 \ S) g" using pfg by (simp add: differentiable_imp_continuous_on diffg) next have non0fg: "0 \ closed_segment (f x) (g x)" if "norm x = 1" "x \ S" for x proof - have "f x \ sphere 0 1" using fim that by (simp add: image_subset_iff) moreover have "norm(f x - g x) < 1/2" using g12 that by auto ultimately show ?thesis by (auto simp: norm_minus_commute dest: segment_bound) qed show "closed_segment (f x) (g x) \ T - {0}" if "x \ sphere 0 1 \ S" for x proof - have "convex T" by (simp add: \subspace T\ subspace_imp_convex) then have "convex hull {f x, g x} \ T" by (metis IntD2 closed_segment_subset fim gim image_subset_iff segment_convex_hull that) then show ?thesis using that non0fg segment_convex_hull by fastforce qed qed obtain d where d: "d \ (sphere 0 1 \ T) - h ` (sphere 0 1 \ S)" using h spheremap_lemma1 [OF ST \S \ T\ diffh] by force then have non0hd: "0 \ closed_segment (h x) (- d)" if "norm x = 1" "x \ S" for x using midpoint_between [of 0 "h x" "-d"] that h [of x] by (auto simp: between_mem_segment midpoint_def) have conth: "continuous_on (sphere 0 1 \ S) h" using differentiable_imp_continuous_on diffh by blast have hom_hd: "homotopic_with_canon (\z. True) (sphere 0 1 \ S) (T - {0}) h (\x. -d)" proof (rule homotopic_with_linear [OF conth continuous_on_const]) fix x assume x: "x \ sphere 0 1 \ S" have "convex hull {h x, - d} \ T" proof (rule hull_minimal) show "{h x, - d} \ T" using h d x by (force simp: subspace_neg [OF \subspace T\]) qed (simp add: subspace_imp_convex [OF \subspace T\]) with x segment_convex_hull show "closed_segment (h x) (- d) \ T - {0}" by (auto simp add: subset_Diff_insert non0hd) qed have conT0: "continuous_on (T - {0}) (\y. inverse(norm y) *\<^sub>R y)" by (intro continuous_intros) auto have sub0T: "(\y. y /\<^sub>R norm y) ` (T - {0}) \ sphere 0 1 \ T" by (fastforce simp: assms(2) subspace_mul) obtain c where homhc: "homotopic_with_canon (\z. True) (sphere 0 1 \ S) (sphere 0 1 \ T) h (\x. c)" - apply (rule_tac c="-d" in that) - apply (rule homotopic_with_eq) - apply (rule homotopic_with_compose_continuous_left [OF hom_hd conT0 sub0T]) - using d apply (auto simp: h_def) - done + proof + show "homotopic_with_canon (\z. True) (sphere 0 1 \ S) (sphere 0 1 \ T) h (\x. - d)" + using d + by (force simp: h_def + intro: homotopic_with_eq homotopic_with_compose_continuous_left [OF hom_hd conT0 sub0T]) + qed have "homotopic_with_canon (\x. True) (sphere 0 1 \ S) (sphere 0 1 \ T) f h" - apply (rule homotopic_with_eq [OF homotopic_with_compose_continuous_left [OF homfg conT0 sub0T]]) - by (auto simp: h_def) + by (force simp: h_def + intro: homotopic_with_eq homotopic_with_compose_continuous_left [OF homfg conT0 sub0T]) then show ?thesis by (metis homotopic_with_trans [OF _ homhc]) qed lemma spheremap_lemma3: assumes "bounded S" "convex S" "subspace U" and affSU: "aff_dim S \ dim U" obtains T where "subspace T" "T \ U" "S \ {} \ aff_dim T = aff_dim S" "(rel_frontier S) homeomorphic (sphere 0 1 \ T)" proof (cases "S = {}") case True with \subspace U\ subspace_0 show ?thesis by (rule_tac T = "{0}" in that) auto next case False then obtain a where "a \ S" by auto then have affS: "aff_dim S = int (dim ((\x. -a+x) ` S))" by (metis hull_inc aff_dim_eq_dim) with affSU have "dim ((\x. -a+x) ` S) \ dim U" by linarith with choose_subspace_of_subspace obtain T where "subspace T" "T \ span U" and dimT: "dim T = dim ((\x. -a+x) ` S)" . show ?thesis proof (rule that [OF \subspace T\]) show "T \ U" using span_eq_iff \subspace U\ \T \ span U\ by blast show "aff_dim T = aff_dim S" using dimT \subspace T\ affS aff_dim_subspace by fastforce show "rel_frontier S homeomorphic sphere 0 1 \ T" proof - have "aff_dim (ball 0 1 \ T) = aff_dim (T)" by (metis IntI interior_ball \subspace T\ aff_dim_convex_Int_nonempty_interior centre_in_ball empty_iff inf_commute subspace_0 subspace_imp_convex zero_less_one) then have affS_eq: "aff_dim S = aff_dim (ball 0 1 \ T)" using \aff_dim T = aff_dim S\ by simp have "rel_frontier S homeomorphic rel_frontier(ball 0 1 \ T)" proof (rule homeomorphic_rel_frontiers_convex_bounded_sets [OF \convex S\ \bounded S\]) show "convex (ball 0 1 \ T)" by (simp add: \subspace T\ convex_Int subspace_imp_convex) show "bounded (ball 0 1 \ T)" by (simp add: bounded_Int) show "aff_dim S = aff_dim (ball 0 1 \ T)" by (rule affS_eq) qed also have "... = frontier (ball 0 1) \ T" proof (rule convex_affine_rel_frontier_Int [OF convex_ball]) show "affine T" by (simp add: \subspace T\ subspace_imp_affine) show "interior (ball 0 1) \ T \ {}" using \subspace T\ subspace_0 by force qed also have "... = sphere 0 1 \ T" by auto finally show ?thesis . qed qed qed proposition inessential_spheremap_lowdim_gen: fixes f :: "'M::euclidean_space \ 'a::euclidean_space" assumes "convex S" "bounded S" "convex T" "bounded T" and affST: "aff_dim S < aff_dim T" and contf: "continuous_on (rel_frontier S) f" and fim: "f ` (rel_frontier S) \ rel_frontier T" obtains c where "homotopic_with_canon (\z. True) (rel_frontier S) (rel_frontier T) f (\x. c)" proof (cases "S = {}") case True then show ?thesis by (simp add: that) next case False then show ?thesis proof (cases "T = {}") case True then show ?thesis using fim that by auto next case False obtain T':: "'a set" where "subspace T'" and affT': "aff_dim T' = aff_dim T" and homT: "rel_frontier T homeomorphic sphere 0 1 \ T'" apply (rule spheremap_lemma3 [OF \bounded T\ \convex T\ subspace_UNIV, where 'b='a]) - apply (simp add: aff_dim_le_DIM) - using \T \ {}\ by blast + using \T \ {}\ by (auto simp add: aff_dim_le_DIM) with homeomorphic_imp_homotopy_eqv have relT: "sphere 0 1 \ T' homotopy_eqv rel_frontier T" using homotopy_equivalent_space_sym by blast have "aff_dim S \ int (dim T')" using affT' \subspace T'\ affST aff_dim_subspace by force with spheremap_lemma3 [OF \bounded S\ \convex S\ \subspace T'\] \S \ {}\ obtain S':: "'a set" where "subspace S'" "S' \ T'" and affS': "aff_dim S' = aff_dim S" and homT: "rel_frontier S homeomorphic sphere 0 1 \ S'" by metis with homeomorphic_imp_homotopy_eqv have relS: "sphere 0 1 \ S' homotopy_eqv rel_frontier S" using homotopy_equivalent_space_sym by blast have dimST': "dim S' < dim T'" by (metis \S' \ T'\ \subspace S'\ \subspace T'\ affS' affST affT' less_irrefl not_le subspace_dim_equal) have "\c. homotopic_with_canon (\z. True) (rel_frontier S) (rel_frontier T) f (\x. c)" apply (rule homotopy_eqv_homotopic_triviality_null_imp [OF relT contf fim]) apply (rule homotopy_eqv_cohomotopic_triviality_null[OF relS, THEN iffD1, rule_format]) apply (metis dimST' \subspace S'\ \subspace T'\ \S' \ T'\ spheremap_lemma2, blast) done with that show ?thesis by blast qed qed lemma inessential_spheremap_lowdim: fixes f :: "'M::euclidean_space \ 'a::euclidean_space" assumes "DIM('M) < DIM('a)" and f: "continuous_on (sphere a r) f" "f ` (sphere a r) \ (sphere b s)" obtains c where "homotopic_with_canon (\z. True) (sphere a r) (sphere b s) f (\x. c)" proof (cases "s \ 0") case True then show ?thesis by (meson nullhomotopic_into_contractible f contractible_sphere that) next case False show ?thesis proof (cases "r \ 0") case True then show ?thesis by (meson f nullhomotopic_from_contractible contractible_sphere that) next case False with \\ s \ 0\ have "r > 0" "s > 0" by auto show thesis apply (rule inessential_spheremap_lowdim_gen [of "cball a r" "cball b s" f]) - using \0 < r\ \0 < s\ assms(1) that - by (simp_all add: f aff_dim_cball inessential_spheremap_lowdim_gen [of "cball a r" "cball b s" f]) + using \0 < r\ \0 < s\ assms(1) that by (simp_all add: f aff_dim_cball) qed qed subsection\ Some technical lemmas about extending maps from cell complexes\ lemma extending_maps_Union_aux: assumes fin: "finite \" and "\S. S \ \ \ closed S" and "\S T. \S \ \; T \ \; S \ T\ \ S \ T \ K" and "\S. S \ \ \ \g. continuous_on S g \ g ` S \ T \ (\x \ S \ K. g x = h x)" shows "\g. continuous_on (\\) g \ g ` (\\) \ T \ (\x \ \\ \ K. g x = h x)" using assms proof (induction \) case empty show ?case by simp next case (insert S \) then obtain f where contf: "continuous_on (S) f" and fim: "f ` S \ T" and feq: "\x \ S \ K. f x = h x" by (meson insertI1) obtain g where contg: "continuous_on (\\) g" and gim: "g ` \\ \ T" and geq: "\x \ \\ \ K. g x = h x" using insert by auto have fg: "f x = g x" if "x \ T" "T \ \" "x \ S" for x T proof - have "T \ S \ K \ S = T" using that by (metis (no_types) insert.prems(2) insertCI) then show ?thesis using UnionI feq geq \S \ \\ subsetD that by fastforce qed show ?case apply (rule_tac x="\x. if x \ S then f x else g x" in exI, simp) apply (intro conjI continuous_on_cases) using fim gim feq geq apply (force simp: insert closed_Union contf contg inf_commute intro: fg)+ done qed lemma extending_maps_Union: assumes fin: "finite \" and "\S. S \ \ \ \g. continuous_on S g \ g ` S \ T \ (\x \ S \ K. g x = h x)" and "\S. S \ \ \ closed S" and K: "\X Y. \X \ \; Y \ \; \ X \ Y; \ Y \ X\ \ X \ Y \ K" shows "\g. continuous_on (\\) g \ g ` (\\) \ T \ (\x \ \\ \ K. g x = h x)" -apply (simp add: Union_maximal_sets [OF fin, symmetric]) +apply (simp flip: Union_maximal_sets [OF fin]) apply (rule extending_maps_Union_aux) apply (simp_all add: Union_maximal_sets [OF fin] assms) by (metis K psubsetI) lemma extend_map_lemma: assumes "finite \" "\ \ \" "convex T" "bounded T" and poly: "\X. X \ \ \ polytope X" and aff: "\X. X \ \ - \ \ aff_dim X < aff_dim T" and face: "\S T. \S \ \; T \ \\ \ (S \ T) face_of S" and contf: "continuous_on (\\) f" and fim: "f ` (\\) \ rel_frontier T" obtains g where "continuous_on (\\) g" "g ` (\\) \ rel_frontier T" "\x. x \ \\ \ g x = f x" proof (cases "\ - \ = {}") case True show ?thesis proof show "continuous_on (\ \) f" using True \\ \ \\ contf by auto show "f ` \ \ \ rel_frontier T" using True fim by auto qed auto next case False then have "0 \ aff_dim T" by (metis aff aff_dim_empty aff_dim_geq aff_dim_negative_iff all_not_in_conv not_less) then obtain i::nat where i: "int i = aff_dim T" by (metis nonneg_eq_int) have Union_empty_eq: "\{D. D = {} \ P D} = {}" for P :: "'a set \ bool" by auto have face': "\S T. \S \ \; T \ \\ \ (S \ T) face_of S \ (S \ T) face_of T" by (metis face inf_commute) have extendf: "\g. continuous_on (\(\ \ {D. \C \ \. D face_of C \ aff_dim D < i})) g \ g ` (\ (\ \ {D. \C \ \. D face_of C \ aff_dim D < i})) \ rel_frontier T \ (\x \ \\. g x = f x)" if "i \ aff_dim T" for i::nat using that proof (induction i) case 0 show ?case using 0 contf fim by (auto simp add: Union_empty_eq) next case (Suc p) with \bounded T\ have "rel_frontier T \ {}" by (auto simp: rel_frontier_eq_empty affine_bounded_eq_lowdim [of T]) then obtain t where t: "t \ rel_frontier T" by auto have ple: "int p \ aff_dim T" using Suc.prems by force obtain h where conth: "continuous_on (\(\ \ {D. \C \ \. D face_of C \ aff_dim D < p})) h" and him: "h ` (\ (\ \ {D. \C \ \. D face_of C \ aff_dim D < p})) \ rel_frontier T" and heq: "\x. x \ \\ \ h x = f x" using Suc.IH [OF ple] by auto let ?Faces = "{D. \C \ \. D face_of C \ aff_dim D \ p}" have extendh: "\g. continuous_on D g \ g ` D \ rel_frontier T \ (\x \ D \ \(\ \ {D. \C \ \. D face_of C \ aff_dim D < p}). g x = h x)" if D: "D \ \ \ ?Faces" for D proof (cases "D \ \(\ \ {D. \C \ \. D face_of C \ aff_dim D < p})") case True have "continuous_on D h" using True conth continuous_on_subset by blast moreover have "h ` D \ rel_frontier T" using True him by blast ultimately show ?thesis by blast next case False note notDsub = False show ?thesis proof (cases "\a. D = {a}") case True then obtain a where "D = {a}" by auto with notDsub t show ?thesis by (rule_tac x="\x. t" in exI) simp next case False have "D \ {}" using notDsub by auto have Dnotin: "D \ \ \ {D. \C \ \. D face_of C \ aff_dim D < p}" using notDsub by auto then have "D \ \" by simp have "D \ ?Faces - {D. \C \ \. D face_of C \ aff_dim D < p}" using Dnotin that by auto then obtain C where "C \ \" "D face_of C" and affD: "aff_dim D = int p" by auto then have "bounded D" using face_of_polytope_polytope poly polytope_imp_bounded by blast then have [simp]: "\ affine D" using affine_bounded_eq_trivial False \D \ {}\ \bounded D\ by blast have "{F. F facet_of D} \ {E. E face_of C \ aff_dim E < int p}" by clarify (metis \D face_of C\ affD eq_iff face_of_trans facet_of_def zle_diff1_eq) moreover have "polyhedron D" using \C \ \\ \D face_of C\ face_of_polytope_polytope poly polytope_imp_polyhedron by auto ultimately have relf_sub: "rel_frontier D \ \ {E. E face_of C \ aff_dim E < p}" by (simp add: rel_frontier_of_polyhedron Union_mono) then have him_relf: "h ` rel_frontier D \ rel_frontier T" using \C \ \\ him by blast have "convex D" by (simp add: \polyhedron D\ polyhedron_imp_convex) have affD_lessT: "aff_dim D < aff_dim T" using Suc.prems affD by linarith have contDh: "continuous_on (rel_frontier D) h" using \C \ \\ relf_sub by (blast intro: continuous_on_subset [OF conth]) then have *: "(\c. homotopic_with_canon (\x. True) (rel_frontier D) (rel_frontier T) h (\x. c)) = (\g. continuous_on UNIV g \ range g \ rel_frontier T \ (\x\rel_frontier D. g x = h x))" - apply (rule nullhomotopic_into_rel_frontier_extension [OF closed_rel_frontier]) - apply (simp_all add: assms rel_frontier_eq_empty him_relf) - done - have "(\c. homotopic_with_canon (\x. True) (rel_frontier D) - (rel_frontier T) h (\x. c))" + by (simp add: assms rel_frontier_eq_empty him_relf nullhomotopic_into_rel_frontier_extension [OF closed_rel_frontier]) + have "(\c. homotopic_with_canon (\x. True) (rel_frontier D) (rel_frontier T) h (\x. c))" by (metis inessential_spheremap_lowdim_gen [OF \convex D\ \bounded D\ \convex T\ \bounded T\ affD_lessT contDh him_relf]) then obtain g where contg: "continuous_on UNIV g" and gim: "range g \ rel_frontier T" and gh: "\x. x \ rel_frontier D \ g x = h x" by (metis *) have "D \ E \ rel_frontier D" if "E \ \ \ {D. Bex \ ((face_of) D) \ aff_dim D < int p}" for E proof (rule face_of_subset_rel_frontier) show "D \ E face_of D" using that proof safe assume "E \ \" then show "D \ E face_of D" by (meson \C \ \\ \D face_of C\ assms(2) face' face_of_Int_subface face_of_refl_eq poly polytope_imp_convex subsetD) next fix x assume "aff_dim E < int p" "x \ \" "E face_of x" then show "D \ E face_of D" by (meson \C \ \\ \D face_of C\ face' face_of_Int_subface that) qed show "D \ E \ D" using that notDsub by auto qed moreover have "continuous_on D g" using contg continuous_on_subset by blast ultimately show ?thesis by (rule_tac x=g in exI) (use gh gim in fastforce) qed qed have intle: "i < 1 + int j \ i \ int j" for i j by auto have "finite \" using \finite \\ \\ \ \\ rev_finite_subset by blast moreover have "finite (?Faces)" proof - have \
: "finite (\ {{D. D face_of C} |C. C \ \})" by (auto simp: \finite \\ finite_polytope_faces poly) show ?thesis by (auto intro: finite_subset [OF _ \
]) qed ultimately have fin: "finite (\ \ ?Faces)" by simp have clo: "closed S" if "S \ \ \ ?Faces" for S using that \\ \ \\ face_of_polytope_polytope poly polytope_imp_closed by blast have K: "X \ Y \ \(\ \ {D. \C\\. D face_of C \ aff_dim D < int p})" if "X \ \ \ ?Faces" "Y \ \ \ ?Faces" "\ Y \ X" for X Y proof - have ff: "X \ Y face_of X \ X \ Y face_of Y" if XY: "X face_of D" "Y face_of E" and DE: "D \ \" "E \ \" for D E by (rule face_of_Int_subface [OF _ _ XY]) (auto simp: face' DE) show ?thesis using that apply auto apply (drule_tac x="X \ Y" in spec, safe) using ff face_of_imp_convex [of X] face_of_imp_convex [of Y] apply (fastforce dest: face_of_aff_dim_lt) by (meson face_of_trans ff) qed obtain g where "continuous_on (\(\ \ ?Faces)) g" "g ` \(\ \ ?Faces) \ rel_frontier T" "(\x \ \(\ \ ?Faces) \ \(\ \ {D. \C\\. D face_of C \ aff_dim D < p}). g x = h x)" by (rule exE [OF extending_maps_Union [OF fin extendh clo K]], blast+) then show ?case by (simp add: intle local.heq [symmetric], blast) qed have eq: "\(\ \ {D. \C \ \. D face_of C \ aff_dim D < i}) = \\" proof show "\(\ \ {D. \C\\. D face_of C \ aff_dim D < int i}) \ \\" using \\ \ \\ face_of_imp_subset by fastforce show "\\ \ \(\ \ {D. \C\\. D face_of C \ aff_dim D < i})" proof (rule Union_mono) show "\ \ \ \ {D. \C\\. D face_of C \ aff_dim D < int i}" using face by (fastforce simp: aff i) qed qed have "int i \ aff_dim T" by (simp add: i) then show ?thesis using extendf [of i] unfolding eq by (metis that) qed lemma extend_map_lemma_cofinite0: assumes "finite \" and "pairwise (\S T. S \ T \ K) \" and "\S. S \ \ \ \a g. a \ U \ continuous_on (S - {a}) g \ g ` (S - {a}) \ T \ (\x \ S \ K. g x = h x)" and "\S. S \ \ \ closed S" shows "\C g. finite C \ disjnt C U \ card C \ card \ \ continuous_on (\\ - C) g \ g ` (\\ - C) \ T \ (\x \ (\\ - C) \ K. g x = h x)" using assms proof induction case empty then show ?case by force next case (insert X \) then have "closed X" and clo: "\X. X \ \ \ closed X" and \: "\S. S \ \ \ \a g. a \ U \ continuous_on (S - {a}) g \ g ` (S - {a}) \ T \ (\x \ S \ K. g x = h x)" and pwX: "\Y. Y \ \ \ Y \ X \ X \ Y \ K \ Y \ X \ K" and pwF: "pairwise (\ S T. S \ T \ K) \" by (simp_all add: pairwise_insert) obtain C g where C: "finite C" "disjnt C U" "card C \ card \" and contg: "continuous_on (\\ - C) g" and gim: "g ` (\\ - C) \ T" and gh: "\x. x \ (\\ - C) \ K \ g x = h x" using insert.IH [OF pwF \ clo] by auto obtain a f where "a \ U" and contf: "continuous_on (X - {a}) f" and fim: "f ` (X - {a}) \ T" and fh: "(\x \ X \ K. f x = h x)" using insert.prems by (meson insertI1) show ?case proof (intro exI conjI) show "finite (insert a C)" by (simp add: C) show "disjnt (insert a C) U" using C \a \ U\ by simp show "card (insert a C) \ card (insert X \)" by (simp add: C card_insert_if insert.hyps le_SucI) have "closed (\\)" using clo insert.hyps by blast have "continuous_on (X - insert a C) f" using contf by (force simp: elim: continuous_on_subset) moreover have "continuous_on (\ \ - insert a C) g" using contg by (force simp: elim: continuous_on_subset) ultimately have "continuous_on (X - insert a C \ (\\ - insert a C)) (\x. if x \ X then f x else g x)" apply (intro continuous_on_cases_local; simp add: closedin_closed) using \closed X\ apply blast using \closed (\\)\ apply blast using fh gh insert.hyps pwX by fastforce then show "continuous_on (\(insert X \) - insert a C) (\a. if a \ X then f a else g a)" by (blast intro: continuous_on_subset) show "\x\(\(insert X \) - insert a C) \ K. (if x \ X then f x else g x) = h x" using gh by (auto simp: fh) show "(\a. if a \ X then f a else g a) ` (\(insert X \) - insert a C) \ T" using fim gim by auto force qed qed lemma extend_map_lemma_cofinite1: assumes "finite \" and \: "\X. X \ \ \ \a g. a \ U \ continuous_on (X - {a}) g \ g ` (X - {a}) \ T \ (\x \ X \ K. g x = h x)" and clo: "\X. X \ \ \ closed X" and K: "\X Y. \X \ \; Y \ \; \ X \ Y; \ Y \ X\ \ X \ Y \ K" obtains C g where "finite C" "disjnt C U" "card C \ card \" "continuous_on (\\ - C) g" "g ` (\\ - C) \ T" "\x. x \ (\\ - C) \ K \ g x = h x" proof - let ?\ = "{X \ \. \Y\\. \ X \ Y}" have [simp]: "\?\ = \\" by (simp add: Union_maximal_sets assms) have fin: "finite ?\" by (force intro: finite_subset [OF _ \finite \\]) have pw: "pairwise (\ S T. S \ T \ K) ?\" by (simp add: pairwise_def) (metis K psubsetI) have "card {X \ \. \Y\\. \ X \ Y} \ card \" by (simp add: \finite \\ card_mono) moreover obtain C g where "finite C \ disjnt C U \ card C \ card ?\ \ continuous_on (\?\ - C) g \ g ` (\?\ - C) \ T \ (\x \ (\?\ - C) \ K. g x = h x)" using extend_map_lemma_cofinite0 [OF fin pw, of U T h] by (fastforce intro!: clo \) ultimately show ?thesis by (rule_tac C=C and g=g in that) auto qed lemma extend_map_lemma_cofinite: assumes "finite \" "\ \ \" and T: "convex T" "bounded T" and poly: "\X. X \ \ \ polytope X" and contf: "continuous_on (\\) f" and fim: "f ` (\\) \ rel_frontier T" and face: "\X Y. \X \ \; Y \ \\ \ (X \ Y) face_of X" and aff: "\X. X \ \ - \ \ aff_dim X \ aff_dim T" obtains C g where "finite C" "disjnt C (\\)" "card C \ card \" "continuous_on (\\ - C) g" "g ` (\ \ - C) \ rel_frontier T" "\x. x \ \\ \ g x = f x" proof - define \ where "\ \ \ \ {D. \C \ \ - \. D face_of C \ aff_dim D < aff_dim T}" have "finite \" using assms finite_subset by blast have *: "finite (\{{D. D face_of C} |C. C \ \})" using finite_polytope_faces poly \finite \\ by force then have "finite \" by (auto simp: \_def \finite \\ intro: finite_subset [OF _ *]) have face': "\S T. \S \ \; T \ \\ \ (S \ T) face_of S \ (S \ T) face_of T" by (metis face inf_commute) have *: "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X" unfolding \_def using subsetD [OF \\ \ \\] apply (auto simp add: face) apply (meson face' face_of_Int_subface face_of_refl_eq poly polytope_imp_convex)+ done obtain h where conth: "continuous_on (\\) h" and him: "h ` (\\) \ rel_frontier T" and hf: "\x. x \ \\ \ h x = f x" proof (rule extend_map_lemma [OF \finite \\ [unfolded \_def] Un_upper1 T]) show "\X. \X \ \ \ {D. \C\\ - \. D face_of C \ aff_dim D < aff_dim T}\ \ polytope X" using \\ \ \\ face_of_polytope_polytope poly by fastforce qed (use * \_def contf fim in auto) have "bounded (\\)" using \finite \\ \\ \ \\ poly polytope_imp_bounded by blast then have "\\ \ UNIV" by auto then obtain a where a: "a \ \\" by blast have \: "\a g. a \ \\ \ continuous_on (D - {a}) g \ g ` (D - {a}) \ rel_frontier T \ (\x \ D \ \\. g x = h x)" if "D \ \" for D proof (cases "D \ \\") case True + then have "h ` (D - {a}) \ rel_frontier T" "continuous_on (D - {a}) h" + using him by (blast intro!: \a \ \\\ continuous_on_subset [OF conth])+ then show ?thesis - apply (rule_tac x=a in exI) - apply (rule_tac x=h in exI) - using him apply (blast intro!: \a \ \\\ continuous_on_subset [OF conth]) + - done + using a by blast next case False note D_not_subset = False show ?thesis proof (cases "D \ \") case True with D_not_subset show ?thesis by (auto simp: \_def) next case False then have affD: "aff_dim D \ aff_dim T" by (simp add: \D \ \\ aff) show ?thesis proof (cases "rel_interior D = {}") case True with \D \ \\ poly a show ?thesis by (force simp: rel_interior_eq_empty polytope_imp_convex) next case False then obtain b where brelD: "b \ rel_interior D" by blast have "polyhedron D" by (simp add: poly polytope_imp_polyhedron that) have "rel_frontier D retract_of affine hull D - {b}" by (simp add: rel_frontier_retract_of_punctured_affine_hull poly polytope_imp_bounded polytope_imp_convex that brelD) then obtain r where relfD: "rel_frontier D \ affine hull D - {b}" and contr: "continuous_on (affine hull D - {b}) r" and rim: "r ` (affine hull D - {b}) \ rel_frontier D" and rid: "\x. x \ rel_frontier D \ r x = x" by (auto simp: retract_of_def retraction_def) show ?thesis proof (intro exI conjI ballI) show "b \ \\" proof clarify fix E assume "b \ E" "E \ \" then have "E \ D face_of E \ E \ D face_of D" using \\ \ \\ face' that by auto with face_of_subset_rel_frontier \E \ \\ \b \ E\ brelD rel_interior_subset [of D] D_not_subset rel_frontier_def \_def show False by blast qed have "r ` (D - {b}) \ r ` (affine hull D - {b})" by (simp add: Diff_mono hull_subset image_mono) also have "... \ rel_frontier D" by (rule rim) also have "... \ \{E. E face_of D \ aff_dim E < aff_dim T}" using affD by (force simp: rel_frontier_of_polyhedron [OF \polyhedron D\] facet_of_def) also have "... \ \(\)" using D_not_subset \_def that by fastforce finally have rsub: "r ` (D - {b}) \ \(\)" . show "continuous_on (D - {b}) (h \ r)" - apply (intro conjI \b \ \\\ continuous_on_compose) - apply (rule continuous_on_subset [OF contr]) - apply (simp add: Diff_mono hull_subset) - apply (rule continuous_on_subset [OF conth rsub]) - done + proof (rule continuous_on_compose) + show "continuous_on (D - {b}) r" + by (meson Diff_mono continuous_on_subset contr hull_subset order_refl) + show "continuous_on (r ` (D - {b})) h" + by (simp add: Diff_mono hull_subset continuous_on_subset [OF conth rsub]) + qed show "(h \ r) ` (D - {b}) \ rel_frontier T" using brelD him rsub by fastforce show "(h \ r) x = h x" if x: "x \ D \ \\" for x proof - consider A where "x \ D" "A \ \" "x \ A" | A B where "x \ D" "A face_of B" "B \ \" "B \ \" "aff_dim A < aff_dim T" "x \ A" using x by (auto simp: \_def) then have xrel: "x \ rel_frontier D" proof cases case 1 show ?thesis proof (rule face_of_subset_rel_frontier [THEN subsetD]) show "D \ A face_of D" using \A \ \\ \\ \ \\ face \D \ \\ by blast show "D \ A \ D" using \A \ \\ D_not_subset \_def by blast qed (auto simp: 1) next case 2 show ?thesis proof (rule face_of_subset_rel_frontier [THEN subsetD]) have "D face_of D" by (simp add: \polyhedron D\ polyhedron_imp_convex face_of_refl) then show "D \ A face_of D" by (meson "2"(2) "2"(3) \D \ \\ face' face_of_Int_Int face_of_face) show "D \ A \ D" using "2" D_not_subset \_def by blast qed (auto simp: 2) qed show ?thesis by (simp add: rid xrel) qed qed qed qed qed have clo: "\S. S \ \ \ closed S" by (simp add: poly polytope_imp_closed) obtain C g where "finite C" "disjnt C (\\)" "card C \ card \" "continuous_on (\\ - C) g" "g ` (\\ - C) \ rel_frontier T" and gh: "\x. x \ (\\ - C) \ \\ \ g x = h x" proof (rule extend_map_lemma_cofinite1 [OF \finite \\ \ clo]) show "X \ Y \ \\" if XY: "X \ \" "Y \ \" and "\ X \ Y" "\ Y \ X" for X Y proof (cases "X \ \") case True then show ?thesis by (auto simp: \_def) next case False have "X \ Y \ X" using \\ X \ Y\ by blast with XY show ?thesis by (clarsimp simp: \_def) (metis Diff_iff Int_iff aff antisym_conv face face_of_aff_dim_lt face_of_refl not_le poly polytope_imp_convex) qed qed (blast)+ with \\ \ \\ show ?thesis - apply (rule_tac C=C and g=g in that) - apply (auto simp: disjnt_def hf [symmetric] \_def intro!: gh) - done + by (rule_tac C=C and g=g in that) (auto simp: disjnt_def hf [symmetric] \_def intro!: gh) qed text\The next two proofs are similar\ theorem extend_map_cell_complex_to_sphere: assumes "finite \" and S: "S \ \\" "closed S" and T: "convex T" "bounded T" and poly: "\X. X \ \ \ polytope X" and aff: "\X. X \ \ \ aff_dim X < aff_dim T" and face: "\X Y. \X \ \; Y \ \\ \ (X \ Y) face_of X" and contf: "continuous_on S f" and fim: "f ` S \ rel_frontier T" obtains g where "continuous_on (\\) g" "g ` (\\) \ rel_frontier T" "\x. x \ S \ g x = f x" proof - obtain V g where "S \ V" "open V" "continuous_on V g" and gim: "g ` V \ rel_frontier T" and gf: "\x. x \ S \ g x = f x" using neighbourhood_extension_into_ANR [OF contf fim _ \closed S\] ANR_rel_frontier_convex T by blast have "compact S" by (meson assms compact_Union poly polytope_imp_compact seq_compact_closed_subset seq_compact_eq_compact) then obtain d where "d > 0" and d: "\x y. \x \ S; y \ - V\ \ d \ dist x y" using separate_compact_closed [of S "-V"] \open V\ \S \ V\ by force obtain \ where "finite \" "\\ = \\" and diaG: "\X. X \ \ \ diameter X < d" and polyG: "\X. X \ \ \ polytope X" and affG: "\X. X \ \ \ aff_dim X \ aff_dim T - 1" and faceG: "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X" proof (rule cell_complex_subdivision_exists [OF \d>0\ \finite \\ poly _ face]) show "\X. X \ \ \ aff_dim X \ aff_dim T - 1" by (simp add: aff) qed auto obtain h where conth: "continuous_on (\\) h" and him: "h ` \\ \ rel_frontier T" and hg: "\x. x \ \(\ \ Pow V) \ h x = g x" proof (rule extend_map_lemma [of \ "\ \ Pow V" T g]) show "continuous_on (\(\ \ Pow V)) g" by (metis Union_Int_subset Union_Pow_eq \continuous_on V g\ continuous_on_subset le_inf_iff) qed (use \finite \\ T polyG affG faceG gim in fastforce)+ show ?thesis proof show "continuous_on (\\) h" using \\\ = \\\ conth by auto show "h ` \\ \ rel_frontier T" using \\\ = \\\ him by auto show "h x = f x" if "x \ S" for x proof - have "x \ \\" using \\\ = \\\ \S \ \\\ that by auto then obtain X where "x \ X" "X \ \" by blast then have "diameter X < d" "bounded X" by (auto simp: diaG \X \ \\ polyG polytope_imp_bounded) then have "X \ V" using d [OF \x \ S\] diameter_bounded_bound [OF \bounded X\ \x \ X\] by fastforce have "h x = g x" - apply (rule hg) - using \X \ \\ \X \ V\ \x \ X\ by blast + using \X \ \\ \X \ V\ \x \ X\ hg by auto also have "... = f x" by (simp add: gf that) finally show "h x = f x" . qed qed qed theorem extend_map_cell_complex_to_sphere_cofinite: assumes "finite \" and S: "S \ \\" "closed S" and T: "convex T" "bounded T" and poly: "\X. X \ \ \ polytope X" and aff: "\X. X \ \ \ aff_dim X \ aff_dim T" and face: "\X Y. \X \ \; Y \ \\ \ (X \ Y) face_of X" and contf: "continuous_on S f" and fim: "f ` S \ rel_frontier T" obtains C g where "finite C" "disjnt C S" "continuous_on (\\ - C) g" "g ` (\\ - C) \ rel_frontier T" "\x. x \ S \ g x = f x" proof - obtain V g where "S \ V" "open V" "continuous_on V g" and gim: "g ` V \ rel_frontier T" and gf: "\x. x \ S \ g x = f x" using neighbourhood_extension_into_ANR [OF contf fim _ \closed S\] ANR_rel_frontier_convex T by blast have "compact S" by (meson assms compact_Union poly polytope_imp_compact seq_compact_closed_subset seq_compact_eq_compact) then obtain d where "d > 0" and d: "\x y. \x \ S; y \ - V\ \ d \ dist x y" using separate_compact_closed [of S "-V"] \open V\ \S \ V\ by force obtain \ where "finite \" "\\ = \\" and diaG: "\X. X \ \ \ diameter X < d" and polyG: "\X. X \ \ \ polytope X" and affG: "\X. X \ \ \ aff_dim X \ aff_dim T" and faceG: "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X" by (rule cell_complex_subdivision_exists [OF \d>0\ \finite \\ poly aff face]) auto obtain C h where "finite C" and dis: "disjnt C (\(\ \ Pow V))" and card: "card C \ card \" and conth: "continuous_on (\\ - C) h" and him: "h ` (\\ - C) \ rel_frontier T" and hg: "\x. x \ \(\ \ Pow V) \ h x = g x" proof (rule extend_map_lemma_cofinite [of \ "\ \ Pow V" T g]) show "continuous_on (\(\ \ Pow V)) g" by (metis Union_Int_subset Union_Pow_eq \continuous_on V g\ continuous_on_subset le_inf_iff) show "g ` \(\ \ Pow V) \ rel_frontier T" using gim by force qed (auto intro: \finite \\ T polyG affG dest: faceG) have Ssub: "S \ \(\ \ Pow V)" proof fix x assume "x \ S" then have "x \ \\" using \\\ = \\\ \S \ \\\ by auto then obtain X where "x \ X" "X \ \" by blast then have "diameter X < d" "bounded X" by (auto simp: diaG \X \ \\ polyG polytope_imp_bounded) then have "X \ V" using d [OF \x \ S\] diameter_bounded_bound [OF \bounded X\ \x \ X\] by fastforce then show "x \ \(\ \ Pow V)" using \X \ \\ \x \ X\ by blast qed show ?thesis proof show "continuous_on (\\-C) h" using \\\ = \\\ conth by auto show "h ` (\\ - C) \ rel_frontier T" using \\\ = \\\ him by auto show "h x = f x" if "x \ S" for x proof - have "h x = g x" using Ssub hg that by blast also have "... = f x" by (simp add: gf that) finally show "h x = f x" . qed show "disjnt C S" using dis Ssub by (meson disjnt_iff subset_eq) qed (intro \finite C\) qed subsection\ Special cases and corollaries involving spheres\ lemma disjnt_Diff1: "X \ Y' \ disjnt (X - Y) (X' - Y')" by (auto simp: disjnt_def) proposition extend_map_affine_to_sphere_cofinite_simple: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "compact S" "convex U" "bounded U" and aff: "aff_dim T \ aff_dim U" and "S \ T" and contf: "continuous_on S f" and fim: "f ` S \ rel_frontier U" obtains K g where "finite K" "K \ T" "disjnt K S" "continuous_on (T - K) g" "g ` (T - K) \ rel_frontier U" "\x. x \ S \ g x = f x" proof - have "\K g. finite K \ disjnt K S \ continuous_on (T - K) g \ g ` (T - K) \ rel_frontier U \ (\x \ S. g x = f x)" if "affine T" "S \ T" and aff: "aff_dim T \ aff_dim U" for T proof (cases "S = {}") case True show ?thesis proof (cases "rel_frontier U = {}") case True with \bounded U\ have "aff_dim U \ 0" using affine_bounded_eq_lowdim rel_frontier_eq_empty by auto with aff have "aff_dim T \ 0" by auto then obtain a where "T \ {a}" using \affine T\ affine_bounded_eq_lowdim affine_bounded_eq_trivial by auto then show ?thesis using \S = {}\ fim by (metis Diff_cancel contf disjnt_empty2 finite.emptyI finite_insert finite_subset) next case False then obtain a where "a \ rel_frontier U" by auto then show ?thesis using continuous_on_const [of _ a] \S = {}\ by force qed next case False have "bounded S" by (simp add: \compact S\ compact_imp_bounded) then obtain b where b: "S \ cbox (-b) b" using bounded_subset_cbox_symmetric by blast define bbox where "bbox \ cbox (-(b+One)) (b+One)" have "cbox (-b) b \ bbox" by (auto simp: bbox_def algebra_simps intro!: subset_box_imp) with b \S \ T\ have "S \ bbox \ T" by auto then have Ssub: "S \ \{bbox \ T}" by auto then have "aff_dim (bbox \ T) \ aff_dim U" by (metis aff aff_dim_subset inf_commute inf_le1 order_trans) obtain K g where K: "finite K" "disjnt K S" and contg: "continuous_on (\{bbox \ T} - K) g" and gim: "g ` (\{bbox \ T} - K) \ rel_frontier U" and gf: "\x. x \ S \ g x = f x" proof (rule extend_map_cell_complex_to_sphere_cofinite [OF _ Ssub _ \convex U\ \bounded U\ _ _ _ contf fim]) show "closed S" using \compact S\ compact_eq_bounded_closed by auto show poly: "\X. X \ {bbox \ T} \ polytope X" by (simp add: polytope_Int_polyhedron bbox_def polytope_interval affine_imp_polyhedron \affine T\) show "\X Y. \X \ {bbox \ T}; Y \ {bbox \ T}\ \ X \ Y face_of X" by (simp add:poly face_of_refl polytope_imp_convex) show "\X. X \ {bbox \ T} \ aff_dim X \ aff_dim U" by (simp add: \aff_dim (bbox \ T) \ aff_dim U\) qed auto define fro where "fro \ \d. frontier(cbox (-(b + d *\<^sub>R One)) (b + d *\<^sub>R One))" obtain d where d12: "1/2 \ d" "d \ 1" and dd: "disjnt K (fro d)" proof (rule disjoint_family_elem_disjnt [OF _ \finite K\]) show "infinite {1/2..1::real}" by (simp add: infinite_Icc) have dis1: "disjnt (fro x) (fro y)" if "x b + d *\<^sub>R One" - have cbsub: "cbox (-b) b \ box (-c) c" "cbox (-b) b \ cbox (-c) c" "cbox (-c) c \ bbox" + have cbsub: "cbox (-b) b \ box (-c) c" "cbox (-b) b \ cbox (-c) c" "cbox (-c) c \ bbox" using d12 by (auto simp: algebra_simps subset_box_imp c_def bbox_def) have clo_cbT: "closed (cbox (- c) c \ T)" by (simp add: affine_closed closed_Int closed_cbox \affine T\) have cpT_ne: "cbox (- c) c \ T \ {}" using \S \ {}\ b cbsub(2) \S \ T\ by fastforce have "closest_point (cbox (- c) c \ T) x \ K" if "x \ T" "x \ K" for x proof (cases "x \ cbox (-c) c") case True with that show ?thesis by (simp add: closest_point_self) next case False have int_ne: "interior (cbox (-c) c) \ T \ {}" using \S \ {}\ \S \ T\ b \cbox (- b) b \ box (- c) c\ by force have "convex T" by (meson \affine T\ affine_imp_convex) then have "x \ affine hull (cbox (- c) c \ T)" by (metis Int_commute Int_iff \S \ {}\ \S \ T\ cbsub(1) \x \ T\ affine_hull_convex_Int_nonempty_interior all_not_in_conv b hull_inc inf.orderE interior_cbox) then have "x \ affine hull (cbox (- c) c \ T) - rel_interior (cbox (- c) c \ T)" by (meson DiffI False Int_iff rel_interior_subset subsetCE) then have "closest_point (cbox (- c) c \ T) x \ rel_frontier (cbox (- c) c \ T)" by (rule closest_point_in_rel_frontier [OF clo_cbT cpT_ne]) moreover have "(rel_frontier (cbox (- c) c \ T)) \ fro d" - apply (subst convex_affine_rel_frontier_Int [OF _ \affine T\ int_ne]) - apply (auto simp: fro_def c_def) - done + by (subst convex_affine_rel_frontier_Int [OF _ \affine T\ int_ne]) (auto simp: fro_def c_def) ultimately show ?thesis using dd by (force simp: disjnt_def) qed then have cpt_subset: "closest_point (cbox (- c) c \ T) ` (T - K) \ \{bbox \ T} - K" using closest_point_in_set [OF clo_cbT cpT_ne] cbsub(3) by force show ?thesis proof (intro conjI ballI exI) have "continuous_on (T - K) (closest_point (cbox (- c) c \ T))" - apply (rule continuous_on_closest_point) - using \S \ {}\ cbsub(2) b that - by (auto simp: affine_imp_convex convex_Int affine_closed closed_Int closed_cbox \affine T\) + proof (rule continuous_on_closest_point) + show "convex (cbox (- c) c \ T)" + by (simp add: affine_imp_convex convex_Int \affine T\) + show "closed (cbox (- c) c \ T)" + using clo_cbT by blast + show "cbox (- c) c \ T \ {}" + using \S \ {}\ cbsub(2) b that by auto + qed then show "continuous_on (T - K) (g \ closest_point (cbox (- c) c \ T))" by (metis continuous_on_compose continuous_on_subset [OF contg cpt_subset]) have "(g \ closest_point (cbox (- c) c \ T)) ` (T - K) \ g ` (\{bbox \ T} - K)" by (metis image_comp image_mono cpt_subset) also have "... \ rel_frontier U" by (rule gim) finally show "(g \ closest_point (cbox (- c) c \ T)) ` (T - K) \ rel_frontier U" . show "(g \ closest_point (cbox (- c) c \ T)) x = f x" if "x \ S" for x proof - have "(g \ closest_point (cbox (- c) c \ T)) x = g x" unfolding o_def by (metis IntI \S \ T\ b cbsub(2) closest_point_self subset_eq that) also have "... = f x" by (simp add: that gf) finally show ?thesis . qed qed (auto simp: K) qed then obtain K g where "finite K" "disjnt K S" and contg: "continuous_on (affine hull T - K) g" and gim: "g ` (affine hull T - K) \ rel_frontier U" and gf: "\x. x \ S \ g x = f x" by (metis aff affine_affine_hull aff_dim_affine_hull order_trans [OF \S \ T\ hull_subset [of T affine]]) then obtain K g where "finite K" "disjnt K S" and contg: "continuous_on (T - K) g" and gim: "g ` (T - K) \ rel_frontier U" and gf: "\x. x \ S \ g x = f x" by (rule_tac K=K and g=g in that) (auto simp: hull_inc elim: continuous_on_subset) then show ?thesis by (rule_tac K="K \ T" and g=g in that) (auto simp: disjnt_iff Diff_Int contg) qed subsection\Extending maps to spheres\ (*Up to extend_map_affine_to_sphere_cofinite_gen*) lemma extend_map_affine_to_sphere1: fixes f :: "'a::euclidean_space \ 'b::topological_space" assumes "finite K" "affine U" and contf: "continuous_on (U - K) f" and fim: "f ` (U - K) \ T" and comps: "\C. \C \ components(U - S); C \ K \ {}\ \ C \ L \ {}" and clo: "closedin (top_of_set U) S" and K: "disjnt K S" "K \ U" obtains g where "continuous_on (U - L) g" "g ` (U - L) \ T" "\x. x \ S \ g x = f x" proof (cases "K = {}") case True then show ?thesis by (metis Diff_empty Diff_subset contf fim continuous_on_subset image_subsetI rev_image_eqI subset_iff that) next case False have "S \ U" using clo closedin_limpt by blast then have "(U - S) \ K \ {}" by (metis Diff_triv False Int_Diff K disjnt_def inf.absorb_iff2 inf_commute) then have "\(components (U - S)) \ K \ {}" using Union_components by simp then obtain C0 where C0: "C0 \ components (U - S)" "C0 \ K \ {}" by blast have "convex U" by (simp add: affine_imp_convex \affine U\) then have "locally connected U" by (rule convex_imp_locally_connected) have "\a g. a \ C \ a \ L \ continuous_on (S \ (C - {a})) g \ g ` (S \ (C - {a})) \ T \ (\x \ S. g x = f x)" if C: "C \ components (U - S)" and CK: "C \ K \ {}" for C proof - have "C \ U-S" "C \ L \ {}" by (simp_all add: in_components_subset comps that) then obtain a where a: "a \ C" "a \ L" by auto have opeUC: "openin (top_of_set U) C" proof (rule openin_trans) show "openin (top_of_set (U-S)) C" by (simp add: \locally connected U\ clo locally_diff_closed openin_components_locally_connected [OF _ C]) show "openin (top_of_set U) (U - S)" by (simp add: clo openin_diff) qed then obtain d where "C \ U" "0 < d" and d: "cball a d \ U \ C" using openin_contains_cball by (metis \a \ C\) then have "ball a d \ U \ C" by auto obtain h k where homhk: "homeomorphism (S \ C) (S \ C) h k" and subC: "{x. (\ (h x = x \ k x = x))} \ C" and bou: "bounded {x. (\ (h x = x \ k x = x))}" and hin: "\x. x \ C \ K \ h x \ ball a d \ U" proof (rule homeomorphism_grouping_points_exists_gen [of C "ball a d \ U" "C \ K" "S \ C"]) show "openin (top_of_set C) (ball a d \ U)" by (metis open_ball \C \ U\ \ball a d \ U \ C\ inf.absorb_iff2 inf.orderE inf_assoc open_openin openin_subtopology) show "openin (top_of_set (affine hull C)) C" by (metis \a \ C\ \openin (top_of_set U) C\ affine_hull_eq affine_hull_openin all_not_in_conv \affine U\) show "ball a d \ U \ {}" using \0 < d\ \C \ U\ \a \ C\ by force show "finite (C \ K)" by (simp add: \finite K\) show "S \ C \ affine hull C" by (metis \C \ U\ \S \ U\ \a \ C\ opeUC affine_hull_eq affine_hull_openin all_not_in_conv assms(2) sup.bounded_iff) show "connected C" by (metis C in_components_connected) qed auto have a_BU: "a \ ball a d \ U" using \0 < d\ \C \ U\ \a \ C\ by auto have "rel_frontier (cball a d \ U) retract_of (affine hull (cball a d \ U) - {a})" - apply (rule rel_frontier_retract_of_punctured_affine_hull) - apply (auto simp: \convex U\ convex_Int) - by (metis \affine U\ convex_cball empty_iff interior_cball a_BU rel_interior_convex_Int_affine) + proof (rule rel_frontier_retract_of_punctured_affine_hull) + show "bounded (cball a d \ U)" "convex (cball a d \ U)" + by (auto simp: \convex U\ convex_Int) + show "a \ rel_interior (cball a d \ U)" + by (metis \affine U\ convex_cball empty_iff interior_cball a_BU rel_interior_convex_Int_affine) + qed moreover have "rel_frontier (cball a d \ U) = frontier (cball a d) \ U" - apply (rule convex_affine_rel_frontier_Int) - using a_BU by (force simp: \affine U\)+ + by (metis a_BU \affine U\ convex_affine_rel_frontier_Int convex_cball equals0D interior_cball) moreover have "affine hull (cball a d \ U) = U" by (metis \convex U\ a_BU affine_hull_convex_Int_nonempty_interior affine_hull_eq \affine U\ equals0D inf.commute interior_cball) ultimately have "frontier (cball a d) \ U retract_of (U - {a})" by metis then obtain r where contr: "continuous_on (U - {a}) r" and rim: "r ` (U - {a}) \ sphere a d" "r ` (U - {a}) \ U" and req: "\x. x \ sphere a d \ U \ r x = x" using \affine U\ by (auto simp: retract_of_def retraction_def hull_same) define j where "j \ \x. if x \ ball a d then r x else x" have kj: "\x. x \ S \ k (j x) = x" using \C \ U - S\ \S \ U\ \ball a d \ U \ C\ j_def subC by auto have Uaeq: "U - {a} = (cball a d - {a}) \ U \ (U - ball a d)" using \0 < d\ by auto have jim: "j ` (S \ (C - {a})) \ (S \ C) - ball a d" proof clarify fix y assume "y \ S \ (C - {a})" then have "y \ U - {a}" using \C \ U - S\ \S \ U\ \a \ C\ by auto then have "r y \ sphere a d" using rim by auto then show "j y \ S \ C - ball a d" - apply (simp add: j_def) - using \r y \ sphere a d\ \y \ U - {a}\ \y \ S \ (C - {a})\ d rim by fastforce + unfolding j_def + using \r y \ sphere a d\ \y \ U - {a}\ \y \ S \ (C - {a})\ d rim + by (metis Diff_iff Int_iff Un_iff subsetD cball_diff_eq_sphere image_subset_iff) qed have contj: "continuous_on (U - {a}) j" unfolding j_def Uaeq proof (intro continuous_on_cases_local continuous_on_id, simp_all add: req closedin_closed Uaeq [symmetric]) show "\T. closed T \ (cball a d - {a}) \ U = (U - {a}) \ T" - apply (rule_tac x="(cball a d) \ U" in exI) - using affine_closed \affine U\ by blast + using affine_closed \affine U\ by (rule_tac x="(cball a d) \ U" in exI) blast show "\T. closed T \ U - ball a d = (U - {a}) \ T" - apply (rule_tac x="U - ball a d" in exI) - using \0 < d\ by (force simp: affine_closed \affine U\ closed_Diff) + using \0 < d\ \affine U\ + by (rule_tac x="U - ball a d" in exI) (force simp: affine_closed) show "continuous_on ((cball a d - {a}) \ U) r" by (force intro: continuous_on_subset [OF contr]) qed have fT: "x \ U - K \ f x \ T" for x using fim by blast show ?thesis proof (intro conjI exI) show "continuous_on (S \ (C - {a})) (f \ k \ j)" proof (intro continuous_on_compose) - show "continuous_on (S \ (C - {a})) j" - apply (rule continuous_on_subset [OF contj]) + have "S \ (C - {a}) \ U - {a}" using \C \ U - S\ \S \ U\ \a \ C\ by force - show "continuous_on (j ` (S \ (C - {a}))) k" - apply (rule continuous_on_subset [OF homeomorphism_cont2 [OF homhk]]) + then show "continuous_on (S \ (C - {a})) j" + by (rule continuous_on_subset [OF contj]) + have "j ` (S \ (C - {a})) \ S \ C" using jim \C \ U - S\ \S \ U\ \ball a d \ U \ C\ j_def by blast + then show "continuous_on (j ` (S \ (C - {a}))) k" + by (rule continuous_on_subset [OF homeomorphism_cont2 [OF homhk]]) show "continuous_on (k ` j ` (S \ (C - {a}))) f" proof (clarify intro!: continuous_on_subset [OF contf]) fix y assume "y \ S \ (C - {a})" have ky: "k y \ S \ C" using homeomorphism_image2 [OF homhk] \y \ S \ (C - {a})\ by blast have jy: "j y \ S \ C - ball a d" using Un_iff \y \ S \ (C - {a})\ jim by auto - show "k (j y) \ U - K" - apply safe - using \C \ U\ \S \ U\ homeomorphism_image2 [OF homhk] jy apply blast - by (metis DiffD1 DiffD2 Int_iff Un_iff \disjnt K S\ disjnt_def empty_iff hin homeomorphism_apply2 homeomorphism_image2 homhk imageI jy) + have "k (j y) \ U" + using \C \ U\ \S \ U\ homeomorphism_image2 [OF homhk] jy by blast + moreover have "k (j y) \ K" + using K unfolding disjnt_iff + by (metis DiffE Int_iff Un_iff hin homeomorphism_def homhk image_eqI jy) + ultimately show "k (j y) \ U - K" + by blast qed qed have ST: "\x. x \ S \ (f \ k \ j) x \ T" - apply (simp add: kj) - apply (metis DiffI \S \ U\ \disjnt K S\ subsetD disjnt_iff fim image_subset_iff) - done + proof (simp add: kj) + show "\x. x \ S \ f x \ T" + using K unfolding disjnt_iff by (metis DiffI \S \ U\ subsetD fim image_subset_iff) + qed moreover have "(f \ k \ j) x \ T" if "x \ C" "x \ a" "x \ S" for x proof - have rx: "r x \ sphere a d" using \C \ U\ rim that by fastforce have jj: "j x \ S \ C - ball a d" using jim that by blast have "k (j x) = j x \ k (j x) \ C \ j x \ C" by (metis Diff_iff Int_iff Un_iff \S \ U\ subsetD d j_def jj rx sphere_cball that(1)) - then have "k (j x) \ C" + then have kj: "k (j x) \ C" using homeomorphism_apply2 [OF homhk, of "j x"] \C \ U\ \S \ U\ a rx by (metis (mono_tags, lifting) Diff_iff subsetD jj mem_Collect_eq subC) - with jj \C \ U\ show ?thesis - apply safe - using ST j_def apply fastforce - apply (auto simp: not_less intro!: fT) - by (metis DiffD1 DiffD2 Int_iff hin homeomorphism_apply2 [OF homhk] jj) + then show ?thesis + by (metis DiffE DiffI IntD1 IntI \C \ U\ comp_apply fT hin homeomorphism_apply2 homhk jj kj subset_eq) qed ultimately show "(f \ k \ j) ` (S \ (C - {a})) \ T" by force show "\x\S. (f \ k \ j) x = f x" using kj by simp qed (auto simp: a) qed then obtain a h where ah: "\C. \C \ components (U - S); C \ K \ {}\ \ a C \ C \ a C \ L \ continuous_on (S \ (C - {a C})) (h C) \ h C ` (S \ (C - {a C})) \ T \ (\x \ S. h C x = f x)" using that by metis define F where "F \ {C \ components (U - S). C \ K \ {}}" define G where "G \ {C \ components (U - S). C \ K = {}}" define UF where "UF \ (\C\F. C - {a C})" have "C0 \ F" by (auto simp: F_def C0) have "finite F" proof (subst finite_image_iff [of "\C. C \ K" F, symmetric]) show "inj_on (\C. C \ K) F" unfolding F_def inj_on_def using components_nonoverlap by blast show "finite ((\C. C \ K) ` F)" unfolding F_def by (rule finite_subset [of _ "Pow K"]) (auto simp: \finite K\) qed obtain g where contg: "continuous_on (S \ UF) g" and gh: "\x i. \i \ F; x \ (S \ UF) \ (S \ (i - {a i}))\ \ g x = h i x" proof (rule pasting_lemma_exists_closed [OF \finite F\]) let ?X = "top_of_set (S \ UF)" show "topspace ?X \ (\C\F. S \ (C - {a C}))" using \C0 \ F\ by (force simp: UF_def) show "closedin (top_of_set (S \ UF)) (S \ (C - {a C}))" if "C \ F" for C proof (rule closedin_closed_subset [of U "S \ C"]) - show "closedin (top_of_set U) (S \ C)" - apply (rule closedin_Un_complement_component [OF \locally connected U\ clo]) + have "C \ components (U - S)" using F_def that by blast + then show "closedin (top_of_set U) (S \ C)" + by (rule closedin_Un_complement_component [OF \locally connected U\ clo]) next have "x = a C'" if "C' \ F" "x \ C'" "x \ U" for x C' proof - have "\A. x \ \A \ C' \ A" using \x \ C'\ by blast with that show "x = a C'" by (metis (lifting) DiffD1 F_def Union_components mem_Collect_eq) qed then show "S \ UF \ U" using \S \ U\ by (force simp: UF_def) next show "S \ (C - {a C}) = (S \ C) \ (S \ UF)" using F_def UF_def components_nonoverlap that by auto qed show "continuous_map (subtopology ?X (S \ (C' - {a C'}))) euclidean (h C')" if "C' \ F" for C' proof - have C': "C' \ components (U - S)" "C' \ K \ {}" using F_def that by blast+ show ?thesis using ah [OF C'] by (auto simp: F_def subtopology_subtopology intro: continuous_on_subset) qed show "\i j x. \i \ F; j \ F; x \ topspace ?X \ (S \ (i - {a i})) \ (S \ (j - {a j}))\ \ h i x = h j x" using components_eq by (fastforce simp: components_eq F_def ah) qed auto have SU': "S \ \G \ (S \ UF) \ U" using \S \ U\ in_components_subset by (auto simp: F_def G_def UF_def) have clo1: "closedin (top_of_set (S \ \G \ (S \ UF))) (S \ \G)" proof (rule closedin_closed_subset [OF _ SU']) have *: "\C. C \ F \ openin (top_of_set U) C" unfolding F_def by clarify (metis (no_types, lifting) \locally connected U\ clo closedin_def locally_diff_closed openin_components_locally_connected openin_trans topspace_euclidean_subtopology) show "closedin (top_of_set U) (U - UF)" unfolding UF_def by (force intro: openin_delete *) show "S \ \G = (U - UF) \ (S \ \G \ (S \ UF))" using \S \ U\ apply (auto simp: F_def G_def UF_def) apply (metis Diff_iff UnionI Union_components) apply (metis DiffD1 UnionI Union_components) by (metis (no_types, lifting) IntI components_nonoverlap empty_iff) qed have clo2: "closedin (top_of_set (S \ \G \ (S \ UF))) (S \ UF)" proof (rule closedin_closed_subset [OF _ SU']) show "closedin (top_of_set U) (\C\F. S \ C)" - apply (rule closedin_Union) - apply (simp add: \finite F\) - using F_def \locally connected U\ clo closedin_Un_complement_component by blast + proof (rule closedin_Union) + show "\T. T \ (\) S ` F \ closedin (top_of_set U) T" + using F_def \locally connected U\ clo closedin_Un_complement_component by blast + qed (simp add: \finite F\) show "S \ UF = (\C\F. S \ C) \ (S \ \G \ (S \ UF))" using \S \ U\ apply (auto simp: F_def G_def UF_def) using C0 apply blast - by (metis components_nonoverlap disjnt_def disjnt_iff) + by (metis components_nonoverlap disjoint_iff) qed have SUG: "S \ \G \ U - K" using \S \ U\ K apply (auto simp: G_def disjnt_iff) by (meson Diff_iff subsetD in_components_subset) then have contf': "continuous_on (S \ \G) f" by (rule continuous_on_subset [OF contf]) have contg': "continuous_on (S \ UF) g" - apply (rule continuous_on_subset [OF contg]) - using \S \ U\ by (auto simp: F_def G_def) + by (simp add: contg) have "\x. \S \ U; x \ S\ \ f x = g x" by (subst gh) (auto simp: ah C0 intro: \C0 \ F\) then have f_eq_g: "\x. x \ S \ UF \ x \ S \ \G \ f x = g x" using \S \ U\ apply (auto simp: F_def G_def UF_def dest: in_components_subset) using components_eq by blast have cont: "continuous_on (S \ \G \ (S \ UF)) (\x. if x \ S \ \G then f x else g x)" by (blast intro: continuous_on_cases_local [OF clo1 clo2 contf' contg' f_eq_g, of "\x. x \ S \ \G"]) show ?thesis proof have UF: "\F - L \ UF" unfolding F_def UF_def using ah by blast have "U - S - L = \(components (U - S)) - L" by simp also have "... = \F \ \G - L" unfolding F_def G_def by blast also have "... \ UF \ \G" using UF by blast finally have "U - L \ S \ \G \ (S \ UF)" by blast then show "continuous_on (U - L) (\x. if x \ S \ \G then f x else g x)" by (rule continuous_on_subset [OF cont]) have "((U - L) \ {x. x \ S \ (\xa\G. x \ xa)}) \ ((U - L) \ (-S \ UF))" using \U - L \ S \ \G \ (S \ UF)\ by auto moreover have "g ` ((U - L) \ (-S \ UF)) \ T" proof - have "g x \ T" if "x \ U" "x \ L" "x \ S" "C \ F" "x \ C" "x \ a C" for x C proof (subst gh) show "x \ (S \ UF) \ (S \ (C - {a C}))" using that by (auto simp: UF_def) show "h C x \ T" using ah that by (fastforce simp add: F_def) qed (rule that) then show ?thesis by (force simp: UF_def) qed ultimately have "g ` ((U - L) \ {x. x \ S \ (\xa\G. x \ xa)}) \ T" using image_mono order_trans by blast moreover have "f ` ((U - L) \ (S \ \G)) \ T" using fim SUG by blast ultimately show "(\x. if x \ S \ \G then f x else g x) ` (U - L) \ T" by force show "\x. x \ S \ (if x \ S \ \G then f x else g x) = f x" by (simp add: F_def G_def) qed qed lemma extend_map_affine_to_sphere2: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "compact S" "convex U" "bounded U" "affine T" "S \ T" and affTU: "aff_dim T \ aff_dim U" and contf: "continuous_on S f" and fim: "f ` S \ rel_frontier U" and ovlap: "\C. C \ components(T - S) \ C \ L \ {}" obtains K g where "finite K" "K \ L" "K \ T" "disjnt K S" "continuous_on (T - K) g" "g ` (T - K) \ rel_frontier U" "\x. x \ S \ g x = f x" proof - obtain K g where K: "finite K" "K \ T" "disjnt K S" and contg: "continuous_on (T - K) g" and gim: "g ` (T - K) \ rel_frontier U" and gf: "\x. x \ S \ g x = f x" using assms extend_map_affine_to_sphere_cofinite_simple by metis have "(\y C. C \ components (T - S) \ x \ C \ y \ C \ y \ L)" if "x \ K" for x proof - have "x \ T-S" using \K \ T\ \disjnt K S\ disjnt_def that by fastforce then obtain C where "C \ components(T - S)" "x \ C" by (metis UnionE Union_components) with ovlap [of C] show ?thesis by blast qed then obtain \ where \: "\x. x \ K \ \C. C \ components (T - S) \ x \ C \ \ x \ C \ \ x \ L" by metis obtain h where conth: "continuous_on (T - \ ` K) h" and him: "h ` (T - \ ` K) \ rel_frontier U" and hg: "\x. x \ S \ h x = g x" proof (rule extend_map_affine_to_sphere1 [OF \finite K\ \affine T\ contg gim, of S "\ ` K"]) show cloTS: "closedin (top_of_set T) S" by (simp add: \compact S\ \S \ T\ closed_subset compact_imp_closed) show "\C. \C \ components (T - S); C \ K \ {}\ \ C \ \ ` K \ {}" using \ components_eq by blast qed (use K in auto) show ?thesis proof show *: "\ ` K \ L" using \ by blast show "finite (\ ` K)" by (simp add: K) show "\ ` K \ T" by clarify (meson \ Diff_iff contra_subsetD in_components_subset) show "continuous_on (T - \ ` K) h" by (rule conth) show "disjnt (\ ` K) S" - using K - apply (auto simp: disjnt_def) - by (metis \ DiffD2 UnionI Union_components) + using K \ in_components_subset by (fastforce simp: disjnt_def) qed (simp_all add: him hg gf) qed proposition extend_map_affine_to_sphere_cofinite_gen: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes SUT: "compact S" "convex U" "bounded U" "affine T" "S \ T" and aff: "aff_dim T \ aff_dim U" and contf: "continuous_on S f" and fim: "f ` S \ rel_frontier U" and dis: "\C. \C \ components(T - S); bounded C\ \ C \ L \ {}" obtains K g where "finite K" "K \ L" "K \ T" "disjnt K S" "continuous_on (T - K) g" "g ` (T - K) \ rel_frontier U" "\x. x \ S \ g x = f x" proof (cases "S = {}") case True show ?thesis proof (cases "rel_frontier U = {}") case True with aff have "aff_dim T \ 0" - apply (simp add: rel_frontier_eq_empty) - using affine_bounded_eq_lowdim \bounded U\ order_trans by auto + using affine_bounded_eq_lowdim \bounded U\ order_trans + by (auto simp add: rel_frontier_eq_empty) with aff_dim_geq [of T] consider "aff_dim T = -1" | "aff_dim T = 0" by linarith then show ?thesis proof cases assume "aff_dim T = -1" then have "T = {}" by (simp add: aff_dim_empty) then show ?thesis by (rule_tac K="{}" in that) auto next assume "aff_dim T = 0" then obtain a where "T = {a}" using aff_dim_eq_0 by blast then have "a \ L" using dis [of "{a}"] \S = {}\ by (auto simp: in_components_self) with \S = {}\ \T = {a}\ show ?thesis by (rule_tac K="{a}" and g=f in that) auto qed next case False then obtain y where "y \ rel_frontier U" by auto with \S = {}\ show ?thesis by (rule_tac K="{}" and g="\x. y" in that) (auto) qed next case False have "bounded S" by (simp add: assms compact_imp_bounded) then obtain b where b: "S \ cbox (-b) b" using bounded_subset_cbox_symmetric by blast define LU where "LU \ L \ (\ {C \ components (T - S). \bounded C} - cbox (-(b+One)) (b+One))" obtain K g where "finite K" "K \ LU" "K \ T" "disjnt K S" and contg: "continuous_on (T - K) g" and gim: "g ` (T - K) \ rel_frontier U" and gf: "\x. x \ S \ g x = f x" proof (rule extend_map_affine_to_sphere2 [OF SUT aff contf fim]) show "C \ LU \ {}" if "C \ components (T - S)" for C proof (cases "bounded C") case True with dis that show ?thesis unfolding LU_def by fastforce next case False then have "\ bounded (\{C \ components (T - S). \ bounded C})" by (metis (no_types, lifting) Sup_upper bounded_subset mem_Collect_eq that) then show ?thesis apply (clarsimp simp: LU_def Int_Un_distrib Diff_Int_distrib Int_UN_distrib) by (metis (no_types, lifting) False Sup_upper bounded_cbox bounded_subset inf.orderE mem_Collect_eq that) qed qed blast have *: False if "x \ cbox (- b - m *\<^sub>R One) (b + m *\<^sub>R One)" "x \ box (- b - n *\<^sub>R One) (b + n *\<^sub>R One)" "0 \ m" "m < n" "n \ 1" for m n x using that by (auto simp: mem_box algebra_simps) have "disjoint_family_on (\d. frontier (cbox (- b - d *\<^sub>R One) (b + d *\<^sub>R One))) {1 / 2..1}" by (auto simp: disjoint_family_on_def neq_iff frontier_def dest: *) then obtain d where d12: "1/2 \ d" "d \ 1" and ddis: "disjnt K (frontier (cbox (-(b + d *\<^sub>R One)) (b + d *\<^sub>R One)))" using disjoint_family_elem_disjnt [of "{1/2..1::real}" K "\d. frontier (cbox (-(b + d *\<^sub>R One)) (b + d *\<^sub>R One))"] by (auto simp: \finite K\) define c where "c \ b + d *\<^sub>R One" have cbsub: "cbox (-b) b \ box (-c) c" "cbox (-b) b \ cbox (-c) c" "cbox (-c) c \ cbox (-(b+One)) (b+One)" using d12 by (simp_all add: subset_box c_def inner_diff_left inner_left_distrib) have clo_cT: "closed (cbox (- c) c \ T)" using affine_closed \affine T\ by blast have cT_ne: "cbox (- c) c \ T \ {}" using \S \ {}\ \S \ T\ b cbsub by fastforce have S_sub_cc: "S \ cbox (- c) c" using \cbox (- b) b \ cbox (- c) c\ b by auto show ?thesis proof show "finite (K \ cbox (-(b+One)) (b+One))" using \finite K\ by blast show "K \ cbox (- (b + One)) (b + One) \ L" using \K \ LU\ by (auto simp: LU_def) show "K \ cbox (- (b + One)) (b + One) \ T" using \K \ T\ by auto show "disjnt (K \ cbox (- (b + One)) (b + One)) S" using \disjnt K S\ by (simp add: disjnt_def disjoint_eq_subset_Compl inf.coboundedI1) have cloTK: "closest_point (cbox (- c) c \ T) x \ T - K" if "x \ T" and Knot: "x \ K \ x \ cbox (- b - One) (b + One)" for x proof (cases "x \ cbox (- c) c") case True with \x \ T\ show ?thesis using cbsub(3) Knot by (force simp: closest_point_self) next case False have clo_in_rf: "closest_point (cbox (- c) c \ T) x \ rel_frontier (cbox (- c) c \ T)" proof (intro closest_point_in_rel_frontier [OF clo_cT cT_ne] DiffI notI) have "T \ interior (cbox (- c) c) \ {}" using \S \ {}\ \S \ T\ b cbsub(1) by fastforce then show "x \ affine hull (cbox (- c) c \ T)" by (simp add: Int_commute affine_hull_affine_Int_nonempty_interior \affine T\ hull_inc that(1)) next show "False" if "x \ rel_interior (cbox (- c) c \ T)" proof - have "interior (cbox (- c) c) \ T \ {}" using \S \ {}\ \S \ T\ b cbsub(1) by fastforce then have "affine hull (T \ cbox (- c) c) = T" using affine_hull_convex_Int_nonempty_interior [of T "cbox (- c) c"] by (simp add: affine_imp_convex \affine T\ inf_commute) then show ?thesis by (meson subsetD le_inf_iff rel_interior_subset that False) qed qed have "closest_point (cbox (- c) c \ T) x \ K" proof assume inK: "closest_point (cbox (- c) c \ T) x \ K" have "\x. x \ K \ x \ frontier (cbox (- (b + d *\<^sub>R One)) (b + d *\<^sub>R One))" by (metis ddis disjnt_iff) then show False by (metis DiffI Int_iff \affine T\ cT_ne c_def clo_cT clo_in_rf closest_point_in_set convex_affine_rel_frontier_Int convex_box(1) empty_iff frontier_cbox inK interior_cbox) qed then show ?thesis using cT_ne clo_cT closest_point_in_set by blast qed show "continuous_on (T - K \ cbox (- (b + One)) (b + One)) (g \ closest_point (cbox (-c) c \ T))" + using cloTK apply (intro continuous_on_compose continuous_on_closest_point continuous_on_subset [OF contg]) - apply (simp_all add: clo_cT affine_imp_convex \affine T\ convex_Int cT_ne) - using cloTK by blast + by (auto simp add: clo_cT affine_imp_convex \affine T\ convex_Int cT_ne) have "g (closest_point (cbox (- c) c \ T) x) \ rel_frontier U" if "x \ T" "x \ K \ x \ cbox (- b - One) (b + One)" for x - apply (rule gim [THEN subsetD]) - using that cloTK by blast + using gim [THEN subsetD] that cloTK by blast then show "(g \ closest_point (cbox (- c) c \ T)) ` (T - K \ cbox (- (b + One)) (b + One)) \ rel_frontier U" by force show "\x. x \ S \ (g \ closest_point (cbox (- c) c \ T)) x = f x" by simp (metis (mono_tags, lifting) IntI \S \ T\ cT_ne clo_cT closest_point_refl gf subsetD S_sub_cc) qed qed corollary extend_map_affine_to_sphere_cofinite: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes SUT: "compact S" "affine T" "S \ T" and aff: "aff_dim T \ DIM('b)" and "0 \ r" and contf: "continuous_on S f" and fim: "f ` S \ sphere a r" and dis: "\C. \C \ components(T - S); bounded C\ \ C \ L \ {}" obtains K g where "finite K" "K \ L" "K \ T" "disjnt K S" "continuous_on (T - K) g" "g ` (T - K) \ sphere a r" "\x. x \ S \ g x = f x" proof (cases "r = 0") case True with fim show ?thesis by (rule_tac K="{}" and g = "\x. a" in that) (auto) next case False with assms have "0 < r" by auto then have "aff_dim T \ aff_dim (cball a r)" by (simp add: aff aff_dim_cball) then show ?thesis apply (rule extend_map_affine_to_sphere_cofinite_gen [OF \compact S\ convex_cball bounded_cball \affine T\ \S \ T\ _ contf]) using fim apply (auto simp: assms False that dest: dis) done qed corollary extend_map_UNIV_to_sphere_cofinite: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes aff: "DIM('a) \ DIM('b)" and "0 \ r" - and SUT: "compact S" - and contf: "continuous_on S f" - and fim: "f ` S \ sphere a r" - and dis: "\C. \C \ components(- S); bounded C\ \ C \ L \ {}" + assumes "DIM('a) \ DIM('b)" and "0 \ r" + and "compact S" + and "continuous_on S f" + and "f ` S \ sphere a r" + and "\C. \C \ components(- S); bounded C\ \ C \ L \ {}" obtains K g where "finite K" "K \ L" "disjnt K S" "continuous_on (- K) g" "g ` (- K) \ sphere a r" "\x. x \ S \ g x = f x" -apply (rule extend_map_affine_to_sphere_cofinite - [OF \compact S\ affine_UNIV subset_UNIV _ \0 \ r\ contf fim dis]) - apply (auto simp: assms that Compl_eq_Diff_UNIV [symmetric]) -done + using extend_map_affine_to_sphere_cofinite + [OF \compact S\ affine_UNIV subset_UNIV] assms + by (metis Compl_eq_Diff_UNIV aff_dim_UNIV of_nat_le_iff) corollary extend_map_UNIV_to_sphere_no_bounded_component: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes aff: "DIM('a) \ DIM('b)" and "0 \ r" and SUT: "compact S" and contf: "continuous_on S f" and fim: "f ` S \ sphere a r" and dis: "\C. C \ components(- S) \ \ bounded C" obtains g where "continuous_on UNIV g" "g ` UNIV \ sphere a r" "\x. x \ S \ g x = f x" -apply (rule extend_map_UNIV_to_sphere_cofinite [OF aff \0 \ r\ \compact S\ contf fim, of "{}"]) - apply (auto simp: that dest: dis) + apply (rule extend_map_UNIV_to_sphere_cofinite [OF aff \0 \ r\ \compact S\ contf fim, of "{}"]) + apply (auto dest: dis) done theorem Borsuk_separation_theorem_gen: fixes S :: "'a::euclidean_space set" assumes "compact S" shows "(\c \ components(- S). \bounded c) \ (\f. continuous_on S f \ f ` S \ sphere (0::'a) 1 \ (\c. homotopic_with_canon (\x. True) S (sphere 0 1) f (\x. c)))" (is "?lhs = ?rhs") proof assume L [rule_format]: ?lhs show ?rhs proof clarify fix f :: "'a \ 'a" assume contf: "continuous_on S f" and fim: "f ` S \ sphere 0 1" obtain g where contg: "continuous_on UNIV g" and gim: "range g \ sphere 0 1" and gf: "\x. x \ S \ g x = f x" by (rule extend_map_UNIV_to_sphere_no_bounded_component [OF _ _ \compact S\ contf fim L]) auto then obtain c where c: "homotopic_with_canon (\h. True) UNIV (sphere 0 1) g (\x. c)" using contractible_UNIV nullhomotopic_from_contractible by blast then show "\c. homotopic_with_canon (\x. True) S (sphere 0 1) f (\x. c)" by (metis assms compact_imp_closed contf contg contractible_empty fim gf gim nullhomotopic_from_contractible nullhomotopic_into_sphere_extension) qed next assume R [rule_format]: ?rhs show ?lhs unfolding components_def proof clarify fix a assume "a \ S" and a: "bounded (connected_component_set (- S) a)" - have cont: "continuous_on S (\x. inverse(norm(x - a)) *\<^sub>R (x - a))" - apply (intro continuous_intros) + have "\x\S. norm (x - a) \ 0" using \a \ S\ by auto + then have cont: "continuous_on S (\x. inverse(norm(x - a)) *\<^sub>R (x - a))" + by (intro continuous_intros) have im: "(\x. inverse(norm(x - a)) *\<^sub>R (x - a)) ` S \ sphere 0 1" by clarsimp (metis \a \ S\ eq_iff_diff_eq_0 left_inverse norm_eq_zero) show False using R cont im Borsuk_map_essential_bounded_component [OF \compact S\ \a \ S\] a by blast qed qed corollary Borsuk_separation_theorem: fixes S :: "'a::euclidean_space set" assumes "compact S" and 2: "2 \ DIM('a)" shows "connected(- S) \ (\f. continuous_on S f \ f ` S \ sphere (0::'a) 1 \ (\c. homotopic_with_canon (\x. True) S (sphere 0 1) f (\x. c)))" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs proof (cases "S = {}") case True then show ?thesis by auto next case False then have "(\c\components (- S). \ bounded c)" by (metis L assms(1) bounded_empty cobounded_imp_unbounded compact_imp_bounded in_components_maximal order_refl) then show ?thesis by (simp add: Borsuk_separation_theorem_gen [OF \compact S\]) qed next assume R: ?rhs then show ?lhs apply (simp add: Borsuk_separation_theorem_gen [OF \compact S\, symmetric]) apply (auto simp: components_def connected_iff_eq_connected_component_set) using connected_component_in apply fastforce using cobounded_unique_unbounded_component [OF _ 2, of "-S"] \compact S\ compact_eq_bounded_closed by fastforce qed lemma homotopy_eqv_separation: fixes S :: "'a::euclidean_space set" and T :: "'a set" assumes "S homotopy_eqv T" and "compact S" and "compact T" shows "connected(- S) \ connected(- T)" proof - consider "DIM('a) = 1" | "2 \ DIM('a)" by (metis DIM_ge_Suc0 One_nat_def Suc_1 dual_order.antisym not_less_eq_eq) then show ?thesis proof cases case 1 then show ?thesis using bounded_connected_Compl_1 compact_imp_bounded homotopy_eqv_empty1 homotopy_eqv_empty2 assms by metis next case 2 with assms show ?thesis by (simp add: Borsuk_separation_theorem homotopy_eqv_cohomotopic_triviality_null) qed qed proposition Jordan_Brouwer_separation: fixes S :: "'a::euclidean_space set" and a::'a assumes hom: "S homeomorphic sphere a r" and "0 < r" shows "\ connected(- S)" proof - have "- sphere a r \ ball a r \ {}" using \0 < r\ by (simp add: Int_absorb1 subset_eq) moreover have eq: "- sphere a r - ball a r = - cball a r" by auto have "- cball a r \ {}" proof - have "frontier (cball a r) \ {}" using \0 < r\ by auto then show ?thesis by (metis frontier_complement frontier_empty) qed with eq have "- sphere a r - ball a r \ {}" by auto moreover have "connected (- S) = connected (- sphere a r)" proof (rule homotopy_eqv_separation) show "S homotopy_eqv sphere a r" using hom homeomorphic_imp_homotopy_eqv by blast show "compact (sphere a r)" by simp then show " compact S" using hom homeomorphic_compactness by blast qed ultimately show ?thesis using connected_Int_frontier [of "- sphere a r" "ball a r"] by (auto simp: \0 < r\) qed proposition Jordan_Brouwer_frontier: fixes S :: "'a::euclidean_space set" and a::'a assumes S: "S homeomorphic sphere a r" and T: "T \ components(- S)" and 2: "2 \ DIM('a)" shows "frontier T = S" proof (cases r rule: linorder_cases) assume "r < 0" with S T show ?thesis by auto next assume "r = 0" with S T card_eq_SucD obtain b where "S = {b}" by (auto simp: homeomorphic_finite [of "{a}" S]) have "components (- {b}) = { -{b}}" using T \S = {b}\ by (auto simp: components_eq_sing_iff connected_punctured_universe 2) with T show ?thesis by (metis \S = {b}\ cball_trivial frontier_cball frontier_complement singletonD sphere_trivial) next assume "r > 0" have "compact S" using homeomorphic_compactness compact_sphere S by blast show ?thesis proof (rule frontier_minimal_separating_closed) show "closed S" using \compact S\ compact_eq_bounded_closed by blast show "\ connected (- S)" using Jordan_Brouwer_separation S \0 < r\ by blast obtain f g where hom: "homeomorphism S (sphere a r) f g" using S by (auto simp: homeomorphic_def) show "connected (- T)" if "closed T" "T \ S" for T proof - have "f ` T \ sphere a r" using \T \ S\ hom homeomorphism_image1 by blast moreover have "f ` T \ sphere a r" using \T \ S\ hom by (metis homeomorphism_image2 homeomorphism_of_subsets order_refl psubsetE) ultimately have "f ` T \ sphere a r" by blast then have "connected (- f ` T)" by (rule psubset_sphere_Compl_connected [OF _ \0 < r\ 2]) moreover have "compact T" using \compact S\ bounded_subset compact_eq_bounded_closed that by blast moreover then have "compact (f ` T)" by (meson compact_continuous_image continuous_on_subset hom homeomorphism_def psubsetE \T \ S\) moreover have "T homotopy_eqv f ` T" by (meson \f ` T \ sphere a r\ dual_order.strict_implies_order hom homeomorphic_def homeomorphic_imp_homotopy_eqv homeomorphism_of_subsets \T \ S\) ultimately show ?thesis using homotopy_eqv_separation [of T "f`T"] by blast qed qed (rule T) qed proposition Jordan_Brouwer_nonseparation: fixes S :: "'a::euclidean_space set" and a::'a assumes S: "S homeomorphic sphere a r" and "T \ S" and 2: "2 \ DIM('a)" shows "connected(- T)" proof - have *: "connected(C \ (S - T))" if "C \ components(- S)" for C proof (rule connected_intermediate_closure) show "connected C" using in_components_connected that by auto have "S = frontier C" using "2" Jordan_Brouwer_frontier S that by blast with closure_subset show "C \ (S - T) \ closure C" by (auto simp: frontier_def) qed auto have "components(- S) \ {}" by (metis S bounded_empty cobounded_imp_unbounded compact_eq_bounded_closed compact_sphere components_eq_empty homeomorphic_compactness) then have "- T = (\C \ components(- S). C \ (S - T))" using Union_components [of "-S"] \T \ S\ by auto - then show ?thesis - apply (rule ssubst) - apply (rule connected_Union) - using \T \ S\ apply (auto simp: *) - done + moreover have "connected ..." + using \T \ S\ by (intro connected_Union) (auto simp: *) + ultimately show ?thesis + by simp qed subsection\ Invariance of domain and corollaries\ lemma invariance_of_domain_ball: fixes f :: "'a \ 'a::euclidean_space" assumes contf: "continuous_on (cball a r) f" and "0 < r" and inj: "inj_on f (cball a r)" shows "open(f ` ball a r)" proof (cases "DIM('a) = 1") case True obtain h::"'a\real" and k - where "linear h" "linear k" "h ` UNIV = UNIV" "k ` UNIV = UNIV" - "\x. norm(h x) = norm x" "\x. norm(k x) = norm x" - "\x. k(h x) = x" "\x. h(k x) = x" - apply (rule isomorphisms_UNIV_UNIV [where 'M='a and 'N=real]) - using True - apply force - by (metis UNIV_I UNIV_eq_I imageI) - have cont: "continuous_on S h" "continuous_on T k" for S T + where "linear h" "linear k" "h ` UNIV = UNIV" "k ` UNIV = UNIV" + "\x. norm(h x) = norm x" "\x. norm(k x) = norm x" + and kh: "\x. k(h x) = x" and "\x. h(k x) = x" + proof (rule isomorphisms_UNIV_UNIV) + show "DIM('a) = DIM(real)" + using True by force + qed (metis UNIV_I UNIV_eq_I imageI) + have cont: "continuous_on S h" "continuous_on T k" for S T by (simp_all add: \linear h\ \linear k\ linear_continuous_on linear_linear) have "continuous_on (h ` cball a r) (h \ f \ k)" - apply (intro continuous_on_compose cont continuous_on_subset [OF contf]) - apply (auto simp: \\x. k (h x) = x\) - done + by (intro continuous_on_compose cont continuous_on_subset [OF contf]) (auto simp: kh) moreover have "is_interval (h ` cball a r)" by (simp add: is_interval_connected_1 \linear h\ linear_continuous_on linear_linear connected_continuous_image) moreover have "inj_on (h \ f \ k) (h ` cball a r)" using inj by (simp add: inj_on_def) (metis \\x. k (h x) = x\) ultimately have *: "\T. \open T; T \ h ` cball a r\ \ open ((h \ f \ k) ` T)" using injective_eq_1d_open_map_UNIV by blast have "open ((h \ f \ k) ` (h ` ball a r))" by (rule *) (auto simp: \linear h\ \range h = UNIV\ open_surjective_linear_image) then have "open ((h \ f) ` ball a r)" by (simp add: image_comp \\x. k (h x) = x\ cong: image_cong) then show ?thesis - apply (simp only: image_comp [symmetric]) - - apply (metis open_bijective_linear_image_eq \linear h\ \\x. k (h x) = x\ \range h = UNIV\ bijI inj_on_def) - done + unfolding image_comp [symmetric] + by (metis open_bijective_linear_image_eq \linear h\ kh \range h = UNIV\ bijI inj_on_def) next case False then have 2: "DIM('a) \ 2" by (metis DIM_ge_Suc0 One_nat_def Suc_1 antisym not_less_eq_eq) have fimsub: "f ` ball a r \ - f ` sphere a r" using inj by clarsimp (metis inj_onD less_eq_real_def mem_cball order_less_irrefl) have hom: "f ` sphere a r homeomorphic sphere a r" by (meson compact_sphere contf continuous_on_subset homeomorphic_compact homeomorphic_sym inj inj_on_subset sphere_cball) then have nconn: "\ connected (- f ` sphere a r)" by (rule Jordan_Brouwer_separation) (auto simp: \0 < r\) - obtain C where C: "C \ components (- f ` sphere a r)" and "bounded C" - apply (rule cobounded_has_bounded_component [OF _ nconn]) - apply (simp_all add: 2) + have "bounded (f ` sphere a r)" by (meson compact_imp_bounded compact_continuous_image_eq compact_sphere contf inj sphere_cball) + then obtain C where C: "C \ components (- f ` sphere a r)" and "bounded C" + using cobounded_has_bounded_component [OF _ nconn] "2" by auto moreover have "f ` (ball a r) = C" proof have "C \ {}" by (rule in_components_nonempty [OF C]) show "C \ f ` ball a r" proof (rule ccontr) assume nonsub: "\ C \ f ` ball a r" have "- f ` cball a r \ C" proof (rule components_maximal [OF C]) have "f ` cball a r homeomorphic cball a r" using compact_cball contf homeomorphic_compact homeomorphic_sym inj by blast then show "connected (- f ` cball a r)" by (auto intro: connected_complement_homeomorphic_convex_compact 2) show "- f ` cball a r \ - f ` sphere a r" by auto then show "C \ - f ` cball a r \ {}" using \C \ {}\ in_components_subset [OF C] nonsub using image_iff by fastforce qed then have "bounded (- f ` cball a r)" using bounded_subset \bounded C\ by auto then have "\ bounded (f ` cball a r)" using cobounded_imp_unbounded by blast then show "False" using compact_continuous_image [OF contf] compact_cball compact_imp_bounded by blast qed with \C \ {}\ have "C \ f ` ball a r \ {}" by (simp add: inf.absorb_iff1) then show "f ` ball a r \ C" by (metis components_maximal [OF C _ fimsub] connected_continuous_image ball_subset_cball connected_ball contf continuous_on_subset) qed moreover have "open (- f ` sphere a r)" using hom compact_eq_bounded_closed compact_sphere homeomorphic_compactness by blast ultimately show ?thesis using open_components by blast qed text\Proved by L. E. J. Brouwer (1912)\ theorem invariance_of_domain: fixes f :: "'a \ 'a::euclidean_space" assumes "continuous_on S f" "open S" "inj_on f S" shows "open(f ` S)" unfolding open_subopen [of "f`S"] proof clarify fix a assume "a \ S" obtain \ where "\ > 0" and \: "cball a \ \ S" using \open S\ \a \ S\ open_contains_cball_eq by blast show "\T. open T \ f a \ T \ T \ f ` S" proof (intro exI conjI) show "open (f ` (ball a \))" by (meson \ \0 < \\ assms continuous_on_subset inj_on_subset invariance_of_domain_ball) show "f a \ f ` ball a \" by (simp add: \0 < \\) show "f ` ball a \ \ f ` S" using \ ball_subset_cball by blast qed qed lemma inv_of_domain_ss0: fixes f :: "'a \ 'a::euclidean_space" assumes contf: "continuous_on U f" and injf: "inj_on f U" and fim: "f ` U \ S" and "subspace S" and dimS: "dim S = DIM('b::euclidean_space)" and ope: "openin (top_of_set S) U" shows "openin (top_of_set S) (f ` U)" proof - have "U \ S" using ope openin_imp_subset by blast have "(UNIV::'b set) homeomorphic S" by (simp add: \subspace S\ dimS homeomorphic_subspaces) then obtain h k where homhk: "homeomorphism (UNIV::'b set) S h k" using homeomorphic_def by blast have homkh: "homeomorphism S (k ` S) k h" using homhk homeomorphism_image2 homeomorphism_sym by fastforce have "open ((k \ f \ h) ` k ` U)" proof (rule invariance_of_domain) show "continuous_on (k ` U) (k \ f \ h)" proof (intro continuous_intros) show "continuous_on (k ` U) h" by (meson continuous_on_subset [OF homeomorphism_cont1 [OF homhk]] top_greatest) - show "continuous_on (h ` k ` U) f" - apply (rule continuous_on_subset [OF contf], clarify) - apply (metis homhk homeomorphism_def ope openin_imp_subset rev_subsetD) - done - show "continuous_on (f ` h ` k ` U) k" - apply (rule continuous_on_subset [OF homeomorphism_cont2 [OF homhk]]) - using fim homhk homeomorphism_apply2 ope openin_subset by fastforce + have "h ` k ` U \ U" + by (metis \U \ S\ dual_order.eq_iff homeomorphism_image2 homeomorphism_of_subsets homkh) + then show "continuous_on (h ` k ` U) f" + by (rule continuous_on_subset [OF contf]) + have "f ` h ` k ` U \ S" + using \h ` k ` U \ U\ fim by blast + then show "continuous_on (f ` h ` k ` U) k" + by (rule continuous_on_subset [OF homeomorphism_cont2 [OF homhk]]) qed have ope_iff: "\T. open T \ openin (top_of_set (k ` S)) T" using homhk homeomorphism_image2 open_openin by fastforce show "open (k ` U)" by (simp add: ope_iff homeomorphism_imp_open_map [OF homkh ope]) show "inj_on (k \ f \ h) (k ` U)" apply (clarsimp simp: inj_on_def) - by (metis subsetD fim homeomorphism_apply2 [OF homhk] image_subset_iff inj_on_eq_iff injf \U \ S\) + by (metis \U \ S\ fim homeomorphism_apply2 homhk image_subset_iff inj_onD injf subsetD) qed moreover have eq: "f ` U = h ` (k \ f \ h \ k) ` U" unfolding image_comp [symmetric] using \U \ S\ fim by (metis homeomorphism_image2 homeomorphism_of_subsets homkh subset_image_iff) ultimately show ?thesis by (metis (no_types, hide_lams) homeomorphism_imp_open_map homhk image_comp open_openin subtopology_UNIV) qed lemma inv_of_domain_ss1: fixes f :: "'a \ 'a::euclidean_space" assumes contf: "continuous_on U f" and injf: "inj_on f U" and fim: "f ` U \ S" and "subspace S" and ope: "openin (top_of_set S) U" shows "openin (top_of_set S) (f ` U)" proof - define S' where "S' \ {y. \x \ S. orthogonal x y}" have "subspace S'" by (simp add: S'_def subspace_orthogonal_to_vectors) define g where "g \ \z::'a*'a. ((f \ fst)z, snd z)" have "openin (top_of_set (S \ S')) (g ` (U \ S'))" proof (rule inv_of_domain_ss0) show "continuous_on (U \ S') g" - apply (simp add: g_def) - apply (intro continuous_intros continuous_on_compose2 [OF contf continuous_on_fst], auto) - done + unfolding g_def + by (auto intro!: continuous_intros continuous_on_compose2 [OF contf continuous_on_fst]) show "g ` (U \ S') \ S \ S'" using fim by (auto simp: g_def) show "inj_on g (U \ S')" using injf by (auto simp: g_def inj_on_def) show "subspace (S \ S')" by (simp add: \subspace S'\ \subspace S\ subspace_Times) show "openin (top_of_set (S \ S')) (U \ S')" by (simp add: openin_Times [OF ope]) have "dim (S \ S') = dim S + dim S'" by (simp add: \subspace S'\ \subspace S\ dim_Times) also have "... = DIM('a)" using dim_subspace_orthogonal_to_vectors [OF \subspace S\ subspace_UNIV] by (simp add: add.commute S'_def) finally show "dim (S \ S') = DIM('a)" . qed moreover have "g ` (U \ S') = f ` U \ S'" by (auto simp: g_def image_iff) moreover have "0 \ S'" using \subspace S'\ subspace_affine by blast ultimately show ?thesis by (auto simp: openin_Times_eq) qed corollary invariance_of_domain_subspaces: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes ope: "openin (top_of_set U) S" and "subspace U" "subspace V" and VU: "dim V \ dim U" and contf: "continuous_on S f" and fim: "f ` S \ V" and injf: "inj_on f S" shows "openin (top_of_set V) (f ` S)" proof - obtain V' where "subspace V'" "V' \ U" "dim V' = dim V" using choose_subspace_of_subspace [OF VU] by (metis span_eq_iff \subspace U\) then have "V homeomorphic V'" by (simp add: \subspace V\ homeomorphic_subspaces) then obtain h k where homhk: "homeomorphism V V' h k" using homeomorphic_def by blast have eq: "f ` S = k ` (h \ f) ` S" proof - have "k ` h ` f ` S = f ` S" by (meson fim homeomorphism_def homeomorphism_of_subsets homhk subset_refl) then show ?thesis by (simp add: image_comp) qed show ?thesis unfolding eq proof (rule homeomorphism_imp_open_map) show homkh: "homeomorphism V' V k h" by (simp add: homeomorphism_symD homhk) have hfV': "(h \ f) ` S \ V'" using fim homeomorphism_image1 homhk by fastforce moreover have "openin (top_of_set U) ((h \ f) ` S)" proof (rule inv_of_domain_ss1) show "continuous_on S (h \ f)" by (meson contf continuous_on_compose continuous_on_subset fim homeomorphism_cont1 homhk) show "inj_on (h \ f) S" apply (clarsimp simp: inj_on_def) by (metis fim homeomorphism_apply2 [OF homkh] image_subset_iff inj_onD injf) show "(h \ f) ` S \ U" using \V' \ U\ hfV' by auto qed (auto simp: assms) ultimately show "openin (top_of_set V') ((h \ f) ` S)" using openin_subset_trans \V' \ U\ by force qed qed corollary invariance_of_dimension_subspaces: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes ope: "openin (top_of_set U) S" and "subspace U" "subspace V" and contf: "continuous_on S f" and fim: "f ` S \ V" and injf: "inj_on f S" and "S \ {}" shows "dim U \ dim V" proof - have "False" if "dim V < dim U" proof - obtain T where "subspace T" "T \ U" "dim T = dim V" using choose_subspace_of_subspace [of "dim V" U] by (metis \dim V < dim U\ assms(2) order.strict_implies_order span_eq_iff) then have "V homeomorphic T" by (simp add: \subspace V\ homeomorphic_subspaces) then obtain h k where homhk: "homeomorphism V T h k" using homeomorphic_def by blast have "continuous_on S (h \ f)" by (meson contf continuous_on_compose continuous_on_subset fim homeomorphism_cont1 homhk) moreover have "(h \ f) ` S \ U" using \T \ U\ fim homeomorphism_image1 homhk by fastforce moreover have "inj_on (h \ f) S" apply (clarsimp simp: inj_on_def) by (metis fim homeomorphism_apply1 homhk image_subset_iff inj_onD injf) ultimately have ope_hf: "openin (top_of_set U) ((h \ f) ` S)" using invariance_of_domain_subspaces [OF ope \subspace U\ \subspace U\] by blast have "(h \ f) ` S \ T" using fim homeomorphism_image1 homhk by fastforce then have "dim ((h \ f) ` S) \ dim T" by (rule dim_subset) also have "dim ((h \ f) ` S) = dim U" using \S \ {}\ \subspace U\ by (blast intro: dim_openin ope_hf) finally show False using \dim V < dim U\ \dim T = dim V\ by simp qed then show ?thesis using not_less by blast qed corollary invariance_of_domain_affine_sets: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes ope: "openin (top_of_set U) S" and aff: "affine U" "affine V" "aff_dim V \ aff_dim U" and contf: "continuous_on S f" and fim: "f ` S \ V" and injf: "inj_on f S" shows "openin (top_of_set V) (f ` S)" proof (cases "S = {}") case True then show ?thesis by auto next case False obtain a b where "a \ S" "a \ U" "b \ V" using False fim ope openin_contains_cball by fastforce have "openin (top_of_set ((+) (- b) ` V)) (((+) (- b) \ f \ (+) a) ` (+) (- a) ` S)" proof (rule invariance_of_domain_subspaces) show "openin (top_of_set ((+) (- a) ` U)) ((+) (- a) ` S)" by (metis ope homeomorphism_imp_open_map homeomorphism_translation translation_galois) show "subspace ((+) (- a) ` U)" by (simp add: \a \ U\ affine_diffs_subspace_subtract \affine U\ cong: image_cong_simp) show "subspace ((+) (- b) ` V)" by (simp add: \b \ V\ affine_diffs_subspace_subtract \affine V\ cong: image_cong_simp) show "dim ((+) (- b) ` V) \ dim ((+) (- a) ` U)" by (metis \a \ U\ \b \ V\ aff_dim_eq_dim affine_hull_eq aff of_nat_le_iff) show "continuous_on ((+) (- a) ` S) ((+) (- b) \ f \ (+) a)" by (metis contf continuous_on_compose homeomorphism_cont2 homeomorphism_translation translation_galois) show "((+) (- b) \ f \ (+) a) ` (+) (- a) ` S \ (+) (- b) ` V" using fim by auto show "inj_on ((+) (- b) \ f \ (+) a) ((+) (- a) ` S)" by (auto simp: inj_on_def) (meson inj_onD injf) qed then show ?thesis by (metis (no_types, lifting) homeomorphism_imp_open_map homeomorphism_translation image_comp translation_galois) qed corollary invariance_of_dimension_affine_sets: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes ope: "openin (top_of_set U) S" and aff: "affine U" "affine V" and contf: "continuous_on S f" and fim: "f ` S \ V" and injf: "inj_on f S" and "S \ {}" shows "aff_dim U \ aff_dim V" proof - obtain a b where "a \ S" "a \ U" "b \ V" using \S \ {}\ fim ope openin_contains_cball by fastforce have "dim ((+) (- a) ` U) \ dim ((+) (- b) ` V)" proof (rule invariance_of_dimension_subspaces) show "openin (top_of_set ((+) (- a) ` U)) ((+) (- a) ` S)" by (metis ope homeomorphism_imp_open_map homeomorphism_translation translation_galois) show "subspace ((+) (- a) ` U)" by (simp add: \a \ U\ affine_diffs_subspace_subtract \affine U\ cong: image_cong_simp) show "subspace ((+) (- b) ` V)" by (simp add: \b \ V\ affine_diffs_subspace_subtract \affine V\ cong: image_cong_simp) show "continuous_on ((+) (- a) ` S) ((+) (- b) \ f \ (+) a)" by (metis contf continuous_on_compose homeomorphism_cont2 homeomorphism_translation translation_galois) show "((+) (- b) \ f \ (+) a) ` (+) (- a) ` S \ (+) (- b) ` V" using fim by auto show "inj_on ((+) (- b) \ f \ (+) a) ((+) (- a) ` S)" by (auto simp: inj_on_def) (meson inj_onD injf) qed (use \S \ {}\ in auto) then show ?thesis by (metis \a \ U\ \b \ V\ aff_dim_eq_dim affine_hull_eq aff of_nat_le_iff) qed corollary invariance_of_dimension: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and "open S" and injf: "inj_on f S" and "S \ {}" shows "DIM('a) \ DIM('b)" using invariance_of_dimension_subspaces [of UNIV S UNIV f] assms by auto corollary continuous_injective_image_subspace_dim_le: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "subspace S" "subspace T" and contf: "continuous_on S f" and fim: "f ` S \ T" and injf: "inj_on f S" shows "dim S \ dim T" - apply (rule invariance_of_dimension_subspaces [of S S _ f]) - using assms by (auto simp: subspace_affine) + using invariance_of_dimension_subspaces [of S S _ f] assms by (auto simp: subspace_affine) lemma invariance_of_dimension_convex_domain: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "convex S" and contf: "continuous_on S f" and fim: "f ` S \ affine hull T" and injf: "inj_on f S" shows "aff_dim S \ aff_dim T" proof (cases "S = {}") case True then show ?thesis by (simp add: aff_dim_geq) next case False have "aff_dim (affine hull S) \ aff_dim (affine hull T)" proof (rule invariance_of_dimension_affine_sets) show "openin (top_of_set (affine hull S)) (rel_interior S)" by (simp add: openin_rel_interior) show "continuous_on (rel_interior S) f" using contf continuous_on_subset rel_interior_subset by blast show "f ` rel_interior S \ affine hull T" using fim rel_interior_subset by blast show "inj_on f (rel_interior S)" using inj_on_subset injf rel_interior_subset by blast show "rel_interior S \ {}" by (simp add: False \convex S\ rel_interior_eq_empty) qed auto then show ?thesis by simp qed lemma homeomorphic_convex_sets_le: assumes "convex S" "S homeomorphic T" shows "aff_dim S \ aff_dim T" proof - obtain h k where homhk: "homeomorphism S T h k" using homeomorphic_def assms by blast show ?thesis proof (rule invariance_of_dimension_convex_domain [OF \convex S\]) show "continuous_on S h" using homeomorphism_def homhk by blast show "h ` S \ affine hull T" by (metis homeomorphism_def homhk hull_subset) show "inj_on h S" by (meson homeomorphism_apply1 homhk inj_on_inverseI) qed qed lemma homeomorphic_convex_sets: assumes "convex S" "convex T" "S homeomorphic T" shows "aff_dim S = aff_dim T" by (meson assms dual_order.antisym homeomorphic_convex_sets_le homeomorphic_sym) lemma homeomorphic_convex_compact_sets_eq: assumes "convex S" "compact S" "convex T" "compact T" shows "S homeomorphic T \ aff_dim S = aff_dim T" by (meson assms homeomorphic_convex_compact_sets homeomorphic_convex_sets) lemma invariance_of_domain_gen: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "open S" "continuous_on S f" "inj_on f S" "DIM('b) \ DIM('a)" shows "open(f ` S)" using invariance_of_domain_subspaces [of UNIV S UNIV f] assms by auto lemma injective_into_1d_imp_open_map_UNIV: fixes f :: "'a::euclidean_space \ real" assumes "open T" "continuous_on S f" "inj_on f S" "T \ S" shows "open (f ` T)" apply (rule invariance_of_domain_gen [OF \open T\]) - using assms apply (auto simp: elim: continuous_on_subset subset_inj_on) - done + using assms by (auto simp: elim: continuous_on_subset subset_inj_on) lemma continuous_on_inverse_open: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "open S" "continuous_on S f" "DIM('b) \ DIM('a)" and gf: "\x. x \ S \ g(f x) = x" shows "continuous_on (f ` S) g" proof (clarsimp simp add: continuous_openin_preimage_eq) fix T :: "'a set" assume "open T" have eq: "f ` S \ g -` T = f ` (S \ T)" by (auto simp: gf) - show "openin (top_of_set (f ` S)) (f ` S \ g -` T)" - apply (subst eq) - apply (rule open_openin_trans) - apply (rule invariance_of_domain_gen) + have "open (f ` S)" + by (rule invariance_of_domain_gen) (use assms inj_on_inverseI in auto) + moreover have "open (f ` (S \ T))" using assms - apply auto - using inj_on_inverseI apply auto[1] by (metis \open T\ continuous_on_subset inj_onI inj_on_subset invariance_of_domain_gen openin_open openin_open_eq) + ultimately show "openin (top_of_set (f ` S)) (f ` S \ g -` T)" + unfolding eq by (auto intro: open_openin_trans) qed lemma invariance_of_domain_homeomorphism: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "open S" "continuous_on S f" "DIM('b) \ DIM('a)" "inj_on f S" obtains g where "homeomorphism S (f ` S) f g" proof show "homeomorphism S (f ` S) f (inv_into S f)" by (simp add: assms continuous_on_inverse_open homeomorphism_def) qed corollary invariance_of_domain_homeomorphic: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "open S" "continuous_on S f" "DIM('b) \ DIM('a)" "inj_on f S" shows "S homeomorphic (f ` S)" using invariance_of_domain_homeomorphism [OF assms] by (meson homeomorphic_def) lemma continuous_image_subset_interior: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "continuous_on S f" "inj_on f S" "DIM('b) \ DIM('a)" shows "f ` (interior S) \ interior(f ` S)" - apply (rule interior_maximal) - apply (simp add: image_mono interior_subset) - apply (rule invariance_of_domain_gen) - using assms - apply (auto simp: subset_inj_on interior_subset continuous_on_subset) - done +proof - + have "open (f ` interior S)" + using assms + by (intro invariance_of_domain_gen) (auto simp: subset_inj_on interior_subset continuous_on_subset) + then show ?thesis + by (simp add: image_mono interior_maximal interior_subset) +qed lemma homeomorphic_interiors_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" and dimeq: "DIM('a) = DIM('b)" shows "(interior S) homeomorphic (interior T)" using assms [unfolded homeomorphic_minimal] unfolding homeomorphic_def proof (clarify elim!: ex_forward) fix f g assume S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T" by (auto simp: inj_on_def intro: rev_image_eqI) metis+ have fim: "f ` interior S \ interior T" using continuous_image_subset_interior [OF contf \inj_on f S\] dimeq fST by simp have gim: "g ` interior T \ interior S" using continuous_image_subset_interior [OF contg \inj_on g T\] dimeq gTS by simp show "homeomorphism (interior S) (interior T) f g" unfolding homeomorphism_def proof (intro conjI ballI) show "\x. x \ interior S \ g (f x) = x" by (meson \\x\S. f x \ T \ g (f x) = x\ subsetD interior_subset) have "interior T \ f ` interior S" proof fix x assume "x \ interior T" then have "g x \ interior S" using gim by blast then show "x \ f ` interior S" by (metis T \x \ interior T\ image_iff interior_subset subsetCE) qed then show "f ` interior S = interior T" using fim by blast show "continuous_on (interior S) f" by (metis interior_subset continuous_on_subset contf) show "\y. y \ interior T \ f (g y) = y" by (meson T subsetD interior_subset) have "interior S \ g ` interior T" proof fix x assume "x \ interior S" then have "f x \ interior T" using fim by blast then show "x \ g ` interior T" by (metis S \x \ interior S\ image_iff interior_subset subsetCE) qed then show "g ` interior T = interior S" using gim by blast show "continuous_on (interior T) g" by (metis interior_subset continuous_on_subset contg) qed qed lemma homeomorphic_open_imp_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "open S" "S \ {}" "open T" "T \ {}" shows "DIM('a) = DIM('b)" using assms apply (simp add: homeomorphic_minimal) apply (rule order_antisym; metis inj_onI invariance_of_dimension) done proposition homeomorphic_interiors: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "interior S = {} \ interior T = {}" shows "(interior S) homeomorphic (interior T)" proof (cases "interior T = {}") case True with assms show ?thesis by auto next case False then have "DIM('a) = DIM('b)" using assms apply (simp add: homeomorphic_minimal) apply (rule order_antisym; metis continuous_on_subset inj_onI inj_on_subset interior_subset invariance_of_dimension open_interior) done then show ?thesis by (rule homeomorphic_interiors_same_dimension [OF \S homeomorphic T\]) qed lemma homeomorphic_frontiers_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "closed S" "closed T" and dimeq: "DIM('a) = DIM('b)" shows "(frontier S) homeomorphic (frontier T)" using assms [unfolded homeomorphic_minimal] unfolding homeomorphic_def proof (clarify elim!: ex_forward) fix f g assume S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T" by (auto simp: inj_on_def intro: rev_image_eqI) metis+ have "g ` interior T \ interior S" using continuous_image_subset_interior [OF contg \inj_on g T\] dimeq gTS by simp then have fim: "f ` frontier S \ frontier T" - apply (simp add: frontier_def) + unfolding frontier_def using continuous_image_subset_interior assms(2) assms(3) S by auto have "f ` interior S \ interior T" using continuous_image_subset_interior [OF contf \inj_on f S\] dimeq fST by simp then have gim: "g ` frontier T \ frontier S" - apply (simp add: frontier_def) + unfolding frontier_def using continuous_image_subset_interior T assms(2) assms(3) by auto show "homeomorphism (frontier S) (frontier T) f g" unfolding homeomorphism_def proof (intro conjI ballI) show gf: "\x. x \ frontier S \ g (f x) = x" by (simp add: S assms(2) frontier_def) show fg: "\y. y \ frontier T \ f (g y) = y" by (simp add: T assms(3) frontier_def) have "frontier T \ f ` frontier S" proof fix x assume "x \ frontier T" then have "g x \ frontier S" using gim by blast then show "x \ f ` frontier S" by (metis fg \x \ frontier T\ imageI) qed then show "f ` frontier S = frontier T" using fim by blast show "continuous_on (frontier S) f" by (metis Diff_subset assms(2) closure_eq contf continuous_on_subset frontier_def) have "frontier S \ g ` frontier T" proof fix x assume "x \ frontier S" then have "f x \ frontier T" using fim by blast then show "x \ g ` frontier T" by (metis gf \x \ frontier S\ imageI) qed then show "g ` frontier T = frontier S" using gim by blast show "continuous_on (frontier T) g" by (metis Diff_subset assms(3) closure_closed contg continuous_on_subset frontier_def) qed qed lemma homeomorphic_frontiers: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "closed S" "closed T" "interior S = {} \ interior T = {}" shows "(frontier S) homeomorphic (frontier T)" proof (cases "interior T = {}") case True then show ?thesis by (metis Diff_empty assms closure_eq frontier_def) next case False - show ?thesis - apply (rule homeomorphic_frontiers_same_dimension) - apply (simp_all add: assms) - using False assms homeomorphic_interiors homeomorphic_open_imp_same_dimension by blast + then have "DIM('a) = DIM('b)" + using assms homeomorphic_interiors homeomorphic_open_imp_same_dimension by blast + then show ?thesis + using assms homeomorphic_frontiers_same_dimension by blast qed lemma continuous_image_subset_rel_interior: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \ T" and TS: "aff_dim T \ aff_dim S" shows "f ` (rel_interior S) \ rel_interior(f ` S)" proof (rule rel_interior_maximal) show "f ` rel_interior S \ f ` S" by(simp add: image_mono rel_interior_subset) show "openin (top_of_set (affine hull f ` S)) (f ` rel_interior S)" proof (rule invariance_of_domain_affine_sets) show "openin (top_of_set (affine hull S)) (rel_interior S)" by (simp add: openin_rel_interior) show "aff_dim (affine hull f ` S) \ aff_dim (affine hull S)" by (metis aff_dim_affine_hull aff_dim_subset fim TS order_trans) show "f ` rel_interior S \ affine hull f ` S" by (meson \f ` rel_interior S \ f ` S\ hull_subset order_trans) show "continuous_on (rel_interior S) f" using contf continuous_on_subset rel_interior_subset by blast show "inj_on f (rel_interior S)" using inj_on_subset injf rel_interior_subset by blast qed auto qed lemma homeomorphic_rel_interiors_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" and aff: "aff_dim S = aff_dim T" shows "(rel_interior S) homeomorphic (rel_interior T)" using assms [unfolded homeomorphic_minimal] unfolding homeomorphic_def proof (clarify elim!: ex_forward) fix f g assume S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T" by (auto simp: inj_on_def intro: rev_image_eqI) metis+ have fim: "f ` rel_interior S \ rel_interior T" by (metis \inj_on f S\ aff contf continuous_image_subset_rel_interior fST order_refl) have gim: "g ` rel_interior T \ rel_interior S" by (metis \inj_on g T\ aff contg continuous_image_subset_rel_interior gTS order_refl) show "homeomorphism (rel_interior S) (rel_interior T) f g" unfolding homeomorphism_def proof (intro conjI ballI) show gf: "\x. x \ rel_interior S \ g (f x) = x" using S rel_interior_subset by blast show fg: "\y. y \ rel_interior T \ f (g y) = y" using T mem_rel_interior_ball by blast have "rel_interior T \ f ` rel_interior S" proof fix x assume "x \ rel_interior T" then have "g x \ rel_interior S" using gim by blast then show "x \ f ` rel_interior S" by (metis fg \x \ rel_interior T\ imageI) qed moreover have "f ` rel_interior S \ rel_interior T" by (metis \inj_on f S\ aff contf continuous_image_subset_rel_interior fST order_refl) ultimately show "f ` rel_interior S = rel_interior T" by blast show "continuous_on (rel_interior S) f" using contf continuous_on_subset rel_interior_subset by blast have "rel_interior S \ g ` rel_interior T" proof fix x assume "x \ rel_interior S" then have "f x \ rel_interior T" using fim by blast then show "x \ g ` rel_interior T" by (metis gf \x \ rel_interior S\ imageI) qed then show "g ` rel_interior T = rel_interior S" using gim by blast show "continuous_on (rel_interior T) g" using contg continuous_on_subset rel_interior_subset by blast qed qed + +lemma homeomorphic_aff_dim_le: + fixes S :: "'a::euclidean_space set" + assumes "S homeomorphic T" "rel_interior S \ {}" + shows "aff_dim (affine hull S) \ aff_dim (affine hull T)" +proof - + obtain f g + where S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" + and contf: "continuous_on S f" and contg: "continuous_on T g" + using assms [unfolded homeomorphic_minimal] by auto + show ?thesis + proof (rule invariance_of_dimension_affine_sets) + show "continuous_on (rel_interior S) f" + using contf continuous_on_subset rel_interior_subset by blast + show "f ` rel_interior S \ affine hull T" + by (meson S hull_subset image_subsetI rel_interior_subset rev_subsetD) + show "inj_on f (rel_interior S)" + by (metis S inj_on_inverseI inj_on_subset rel_interior_subset) + qed (simp_all add: openin_rel_interior assms) +qed + lemma homeomorphic_rel_interiors: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "rel_interior S = {} \ rel_interior T = {}" shows "(rel_interior S) homeomorphic (rel_interior T)" proof (cases "rel_interior T = {}") case True with assms show ?thesis by auto next case False - obtain f g - where S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" - and contf: "continuous_on S f" and contg: "continuous_on T g" - using assms [unfolded homeomorphic_minimal] by auto have "aff_dim (affine hull S) \ aff_dim (affine hull T)" - apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior S" _ f]) - apply (simp_all add: openin_rel_interior False assms) - using contf continuous_on_subset rel_interior_subset apply blast - apply (meson S hull_subset image_subsetI rel_interior_subset rev_subsetD) - apply (metis S inj_on_inverseI inj_on_subset rel_interior_subset) - done + using False assms homeomorphic_aff_dim_le by blast moreover have "aff_dim (affine hull T) \ aff_dim (affine hull S)" - apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior T" _ g]) - apply (simp_all add: openin_rel_interior False assms) - using contg continuous_on_subset rel_interior_subset apply blast - apply (meson T hull_subset image_subsetI rel_interior_subset rev_subsetD) - apply (metis T inj_on_inverseI inj_on_subset rel_interior_subset) - done + using False assms(1) homeomorphic_aff_dim_le homeomorphic_sym by auto ultimately have "aff_dim S = aff_dim T" by force then show ?thesis by (rule homeomorphic_rel_interiors_same_dimension [OF \S homeomorphic T\]) qed lemma homeomorphic_rel_boundaries_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" and aff: "aff_dim S = aff_dim T" shows "(S - rel_interior S) homeomorphic (T - rel_interior T)" using assms [unfolded homeomorphic_minimal] unfolding homeomorphic_def proof (clarify elim!: ex_forward) fix f g assume S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T" by (auto simp: inj_on_def intro: rev_image_eqI) metis+ have fim: "f ` rel_interior S \ rel_interior T" by (metis \inj_on f S\ aff contf continuous_image_subset_rel_interior fST order_refl) have gim: "g ` rel_interior T \ rel_interior S" by (metis \inj_on g T\ aff contg continuous_image_subset_rel_interior gTS order_refl) show "homeomorphism (S - rel_interior S) (T - rel_interior T) f g" unfolding homeomorphism_def proof (intro conjI ballI) show gf: "\x. x \ S - rel_interior S \ g (f x) = x" using S rel_interior_subset by blast show fg: "\y. y \ T - rel_interior T \ f (g y) = y" using T mem_rel_interior_ball by blast show "f ` (S - rel_interior S) = T - rel_interior T" using S fST fim gim by auto show "continuous_on (S - rel_interior S) f" using contf continuous_on_subset rel_interior_subset by blast show "g ` (T - rel_interior T) = S - rel_interior S" using T gTS gim fim by auto show "continuous_on (T - rel_interior T) g" using contg continuous_on_subset rel_interior_subset by blast qed qed lemma homeomorphic_rel_boundaries: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "rel_interior S = {} \ rel_interior T = {}" shows "(S - rel_interior S) homeomorphic (T - rel_interior T)" proof (cases "rel_interior T = {}") case True with assms show ?thesis by auto next case False obtain f g where S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" using assms [unfolded homeomorphic_minimal] by auto have "aff_dim (affine hull S) \ aff_dim (affine hull T)" - apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior S" _ f]) - apply (simp_all add: openin_rel_interior False assms) - using contf continuous_on_subset rel_interior_subset apply blast - apply (meson S hull_subset image_subsetI rel_interior_subset rev_subsetD) - apply (metis S inj_on_inverseI inj_on_subset rel_interior_subset) - done + using False assms homeomorphic_aff_dim_le by blast moreover have "aff_dim (affine hull T) \ aff_dim (affine hull S)" - apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior T" _ g]) - apply (simp_all add: openin_rel_interior False assms) - using contg continuous_on_subset rel_interior_subset apply blast - apply (meson T hull_subset image_subsetI rel_interior_subset rev_subsetD) - apply (metis T inj_on_inverseI inj_on_subset rel_interior_subset) - done + by (meson False assms(1) homeomorphic_aff_dim_le homeomorphic_sym) ultimately have "aff_dim S = aff_dim T" by force then show ?thesis by (rule homeomorphic_rel_boundaries_same_dimension [OF \S homeomorphic T\]) qed proposition uniformly_continuous_homeomorphism_UNIV_trivial: fixes f :: "'a::euclidean_space \ 'a" assumes contf: "uniformly_continuous_on S f" and hom: "homeomorphism S UNIV f g" shows "S = UNIV" proof (cases "S = {}") case True then show ?thesis by (metis UNIV_I hom empty_iff homeomorphism_def image_eqI) next case False have "inj g" by (metis UNIV_I hom homeomorphism_apply2 injI) then have "open (g ` UNIV)" by (blast intro: invariance_of_domain hom homeomorphism_cont2) then have "open S" using hom homeomorphism_image2 by blast moreover have "complete S" unfolding complete_def proof clarify fix \ assume \: "\n. \ n \ S" and "Cauchy \" have "Cauchy (f o \)" using uniformly_continuous_imp_Cauchy_continuous \Cauchy \\ \ contf by blast then obtain l where "(f \ \) \ l" by (auto simp: convergent_eq_Cauchy [symmetric]) show "\l\S. \ \ l" proof show "g l \ S" using hom homeomorphism_image2 by blast have "(g \ (f \ \)) \ g l" by (meson UNIV_I \(f \ \) \ l\ continuous_on_sequentially hom homeomorphism_cont2) then show "\ \ g l" proof - have "\n. \ n = (g \ (f \ \)) n" by (metis (no_types) \ comp_eq_dest_lhs hom homeomorphism_apply1) then show ?thesis by (metis (no_types) LIMSEQ_iff \(g \ (f \ \)) \ g l\) qed qed qed then have "closed S" by (simp add: complete_eq_closed) ultimately show ?thesis using clopen [of S] False by simp qed subsection\Formulation of loop homotopy in terms of maps out of type complex\ lemma homotopic_circlemaps_imp_homotopic_loops: assumes "homotopic_with_canon (\h. True) (sphere 0 1) S f g" shows "homotopic_loops S (f \ exp \ (\t. 2 * of_real pi * of_real t * \)) (g \ exp \ (\t. 2 * of_real pi * of_real t * \))" proof - have "homotopic_with_canon (\f. True) {z. cmod z = 1} S f g" using assms by (auto simp: sphere_def) moreover have "continuous_on {0..1} (exp \ (\t. 2 * of_real pi * of_real t * \))" by (intro continuous_intros) moreover have "(exp \ (\t. 2 * of_real pi * of_real t * \)) ` {0..1} \ {z. cmod z = 1}" by (auto simp: norm_mult) ultimately show ?thesis apply (simp add: homotopic_loops_def comp_assoc) apply (rule homotopic_with_compose_continuous_right) apply (auto simp: pathstart_def pathfinish_def) done qed lemma homotopic_loops_imp_homotopic_circlemaps: assumes "homotopic_loops S p q" shows "homotopic_with_canon (\h. True) (sphere 0 1) S (p \ (\z. (Arg2pi z / (2 * pi)))) (q \ (\z. (Arg2pi z / (2 * pi))))" proof - obtain h where conth: "continuous_on ({0..1::real} \ {0..1}) h" and him: "h ` ({0..1} \ {0..1}) \ S" and h0: "(\x. h (0, x) = p x)" and h1: "(\x. h (1, x) = q x)" and h01: "(\t\{0..1}. h (t, 1) = h (t, 0)) " using assms by (auto simp: homotopic_loops_def sphere_def homotopic_with_def pathstart_def pathfinish_def) define j where "j \ \z. if 0 \ Im (snd z) then h (fst z, Arg2pi (snd z) / (2 * pi)) else h (fst z, 1 - Arg2pi (cnj (snd z)) / (2 * pi))" have Arg2pi_eq: "1 - Arg2pi (cnj y) / (2 * pi) = Arg2pi y / (2 * pi) \ Arg2pi y = 0 \ Arg2pi (cnj y) = 0" if "cmod y = 1" for y using that Arg2pi_eq_0_pi Arg2pi_eq_pi by (force simp: Arg2pi_cnj field_split_simps) show ?thesis proof (simp add: homotopic_with; intro conjI ballI exI) show "continuous_on ({0..1} \ sphere 0 1) (\w. h (fst w, Arg2pi (snd w) / (2 * pi)))" proof (rule continuous_on_eq) show j: "j x = h (fst x, Arg2pi (snd x) / (2 * pi))" if "x \ {0..1} \ sphere 0 1" for x using Arg2pi_eq that h01 by (force simp: j_def) have eq: "S = S \ (UNIV \ {z. 0 \ Im z}) \ S \ (UNIV \ {z. Im z \ 0})" for S :: "(real*complex)set" by auto have c1: "continuous_on ({0..1} \ sphere 0 1 \ UNIV \ {z. 0 \ Im z}) (\x. h (fst x, Arg2pi (snd x) / (2 * pi)))" apply (intro continuous_intros continuous_on_compose2 [OF conth] continuous_on_compose2 [OF continuous_on_upperhalf_Arg2pi]) apply (auto simp: Arg2pi) apply (meson Arg2pi_lt_2pi linear not_le) done have c2: "continuous_on ({0..1} \ sphere 0 1 \ UNIV \ {z. Im z \ 0}) (\x. h (fst x, 1 - Arg2pi (cnj (snd x)) / (2 * pi)))" apply (intro continuous_intros continuous_on_compose2 [OF conth] continuous_on_compose2 [OF continuous_on_upperhalf_Arg2pi]) apply (auto simp: Arg2pi) apply (meson Arg2pi_lt_2pi linear not_le) done show "continuous_on ({0..1} \ sphere 0 1) j" apply (simp add: j_def) apply (subst eq) apply (rule continuous_on_cases_local) - apply (simp_all add: eq [symmetric] closedin_closed_Int closed_Times closed_halfspace_Im_le closed_halfspace_Im_ge c1 c2) using Arg2pi_eq h01 - by force + by (force simp add: eq [symmetric] closedin_closed_Int closed_Times closed_halfspace_Im_le closed_halfspace_Im_ge c1 c2)+ qed have "(\w. h (fst w, Arg2pi (snd w) / (2 * pi))) ` ({0..1} \ sphere 0 1) \ h ` ({0..1} \ {0..1})" by (auto simp: Arg2pi_ge_0 Arg2pi_lt_2pi less_imp_le) also have "... \ S" using him by blast finally show "(\w. h (fst w, Arg2pi (snd w) / (2 * pi))) ` ({0..1} \ sphere 0 1) \ S" . qed (auto simp: h0 h1) qed lemma simply_connected_homotopic_loops: "simply_connected S \ (\p q. homotopic_loops S p p \ homotopic_loops S q q \ homotopic_loops S p q)" unfolding simply_connected_def using homotopic_loops_refl by metis lemma simply_connected_eq_homotopic_circlemaps1: fixes f :: "complex \ 'a::topological_space" and g :: "complex \ 'a" assumes S: "simply_connected S" and contf: "continuous_on (sphere 0 1) f" and fim: "f ` (sphere 0 1) \ S" and contg: "continuous_on (sphere 0 1) g" and gim: "g ` (sphere 0 1) \ S" shows "homotopic_with_canon (\h. True) (sphere 0 1) S f g" proof - have "homotopic_loops S (f \ exp \ (\t. of_real(2 * pi * t) * \)) (g \ exp \ (\t. of_real(2 * pi * t) * \))" apply (rule S [unfolded simply_connected_homotopic_loops, rule_format]) apply (simp add: homotopic_circlemaps_imp_homotopic_loops contf fim contg gim) done then show ?thesis apply (rule homotopic_with_eq [OF homotopic_loops_imp_homotopic_circlemaps]) apply (auto simp: o_def complex_norm_eq_1_exp mult.commute) done qed lemma simply_connected_eq_homotopic_circlemaps2a: fixes h :: "complex \ 'a::topological_space" assumes conth: "continuous_on (sphere 0 1) h" and him: "h ` (sphere 0 1) \ S" and hom: "\f g::complex \ 'a. \continuous_on (sphere 0 1) f; f ` (sphere 0 1) \ S; continuous_on (sphere 0 1) g; g ` (sphere 0 1) \ S\ \ homotopic_with_canon (\h. True) (sphere 0 1) S f g" shows "\a. homotopic_with_canon (\h. True) (sphere 0 1) S h (\x. a)" apply (rule_tac x="h 1" in exI) apply (rule hom) using assms by (auto) lemma simply_connected_eq_homotopic_circlemaps2b: fixes S :: "'a::real_normed_vector set" assumes "\f g::complex \ 'a. \continuous_on (sphere 0 1) f; f ` (sphere 0 1) \ S; continuous_on (sphere 0 1) g; g ` (sphere 0 1) \ S\ \ homotopic_with_canon (\h. True) (sphere 0 1) S f g" shows "path_connected S" proof (clarsimp simp add: path_connected_eq_homotopic_points) fix a b assume "a \ S" "b \ S" then show "homotopic_loops S (linepath a a) (linepath b b)" using homotopic_circlemaps_imp_homotopic_loops [OF assms [of "\x. a" "\x. b"]] by (auto simp: o_def linepath_def) qed lemma simply_connected_eq_homotopic_circlemaps3: fixes h :: "complex \ 'a::real_normed_vector" assumes "path_connected S" and hom: "\f::complex \ 'a. \continuous_on (sphere 0 1) f; f `(sphere 0 1) \ S\ \ \a. homotopic_with_canon (\h. True) (sphere 0 1) S f (\x. a)" shows "simply_connected S" proof (clarsimp simp add: simply_connected_eq_contractible_loop_some assms) fix p assume p: "path p" "path_image p \ S" "pathfinish p = pathstart p" then have "homotopic_loops S p p" by (simp add: homotopic_loops_refl) then obtain a where homp: "homotopic_with_canon (\h. True) (sphere 0 1) S (p \ (\z. Arg2pi z / (2 * pi))) (\x. a)" by (metis homotopic_with_imp_subset2 homotopic_loops_imp_homotopic_circlemaps homotopic_with_imp_continuous hom) show "\a. a \ S \ homotopic_loops S p (linepath a a)" proof (intro exI conjI) show "a \ S" using homotopic_with_imp_subset2 [OF homp] by (metis dist_0_norm image_subset_iff mem_sphere norm_one) have teq: "\t. \0 \ t; t \ 1\ \ t = Arg2pi (exp (2 * of_real pi * of_real t * \)) / (2 * pi) \ t=1 \ Arg2pi (exp (2 * of_real pi * of_real t * \)) = 0" - apply (rule disjCI) - using Arg2pi_of_real [of 1] apply (auto simp: Arg2pi_exp) - done + using Arg2pi_of_real [of 1] by (force simp: Arg2pi_exp) have "homotopic_loops S p (p \ (\z. Arg2pi z / (2 * pi)) \ exp \ (\t. 2 * complex_of_real pi * complex_of_real t * \))" - apply (rule homotopic_loops_eq [OF p]) - using p teq apply (fastforce simp: pathfinish_def pathstart_def) - done - then - show "homotopic_loops S p (linepath a a)" + using p teq by (fastforce simp: pathfinish_def pathstart_def intro: homotopic_loops_eq [OF p]) + then show "homotopic_loops S p (linepath a a)" by (simp add: linepath_refl homotopic_loops_trans [OF _ homotopic_circlemaps_imp_homotopic_loops [OF homp, simplified K_record_comp]]) qed qed proposition simply_connected_eq_homotopic_circlemaps: fixes S :: "'a::real_normed_vector set" shows "simply_connected S \ (\f g::complex \ 'a. continuous_on (sphere 0 1) f \ f ` (sphere 0 1) \ S \ continuous_on (sphere 0 1) g \ g ` (sphere 0 1) \ S \ homotopic_with_canon (\h. True) (sphere 0 1) S f g)" apply (rule iffI) - apply (blast elim: dest: simply_connected_eq_homotopic_circlemaps1) + apply (blast dest: simply_connected_eq_homotopic_circlemaps1) by (simp add: simply_connected_eq_homotopic_circlemaps2a simply_connected_eq_homotopic_circlemaps2b simply_connected_eq_homotopic_circlemaps3) proposition simply_connected_eq_contractible_circlemap: fixes S :: "'a::real_normed_vector set" shows "simply_connected S \ path_connected S \ (\f::complex \ 'a. continuous_on (sphere 0 1) f \ f `(sphere 0 1) \ S \ (\a. homotopic_with_canon (\h. True) (sphere 0 1) S f (\x. a)))" apply (rule iffI) apply (simp add: simply_connected_eq_homotopic_circlemaps1 simply_connected_eq_homotopic_circlemaps2a simply_connected_eq_homotopic_circlemaps2b) using simply_connected_eq_homotopic_circlemaps3 by blast corollary homotopy_eqv_simple_connectedness: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" shows "S homotopy_eqv T \ simply_connected S \ simply_connected T" by (simp add: simply_connected_eq_homotopic_circlemaps homotopy_eqv_homotopic_triviality) subsection\Homeomorphism of simple closed curves to circles\ proposition homeomorphic_simple_path_image_circle: fixes a :: complex and \ :: "real \ 'a::t2_space" assumes "simple_path \" and loop: "pathfinish \ = pathstart \" and "0 < r" shows "(path_image \) homeomorphic sphere a r" proof - have "homotopic_loops (path_image \) \ \" by (simp add: assms homotopic_loops_refl simple_path_imp_path) then have hom: "homotopic_with_canon (\h. True) (sphere 0 1) (path_image \) (\ \ (\z. Arg2pi z / (2*pi))) (\ \ (\z. Arg2pi z / (2*pi)))" by (rule homotopic_loops_imp_homotopic_circlemaps) have "\g. homeomorphism (sphere 0 1) (path_image \) (\ \ (\z. Arg2pi z / (2*pi))) g" proof (rule homeomorphism_compact) show "continuous_on (sphere 0 1) (\ \ (\z. Arg2pi z / (2*pi)))" using hom homotopic_with_imp_continuous by blast show "inj_on (\ \ (\z. Arg2pi z / (2*pi))) (sphere 0 1)" proof fix x y assume xy: "x \ sphere 0 1" "y \ sphere 0 1" and eq: "(\ \ (\z. Arg2pi z / (2*pi))) x = (\ \ (\z. Arg2pi z / (2*pi))) y" then have "(Arg2pi x / (2*pi)) = (Arg2pi y / (2*pi))" proof - have "(Arg2pi x / (2*pi)) \ {0..1}" "(Arg2pi y / (2*pi)) \ {0..1}" using Arg2pi_ge_0 Arg2pi_lt_2pi dual_order.strict_iff_order by fastforce+ with eq show ?thesis using \simple_path \\ Arg2pi_lt_2pi unfolding simple_path_def o_def by (metis eq_divide_eq_1 not_less_iff_gr_or_eq) qed with xy show "x = y" by (metis is_Arg_def Arg2pi Arg2pi_0 dist_0_norm divide_cancel_right dual_order.strict_iff_order mem_sphere) qed have "\z. cmod z = 1 \ \x\{0..1}. \ (Arg2pi z / (2*pi)) = \ x" by (metis Arg2pi_ge_0 Arg2pi_lt_2pi atLeastAtMost_iff divide_less_eq_1 less_eq_real_def zero_less_mult_iff pi_gt_zero zero_le_divide_iff zero_less_numeral) moreover have "\z\sphere 0 1. \ x = \ (Arg2pi z / (2*pi))" if "0 \ x" "x \ 1" for x proof (cases "x=1") case True with Arg2pi_of_real [of 1] loop show ?thesis by (rule_tac x=1 in bexI) (auto simp: pathfinish_def pathstart_def \0 \ x\) next case False then have *: "(Arg2pi (exp (\*(2* of_real pi* of_real x))) / (2*pi)) = x" using that by (auto simp: Arg2pi_exp field_split_simps) show ?thesis by (rule_tac x="exp(\ * of_real(2*pi*x))" in bexI) (auto simp: *) qed ultimately show "(\ \ (\z. Arg2pi z / (2*pi))) ` sphere 0 1 = path_image \" by (auto simp: path_image_def image_iff) qed auto then have "path_image \ homeomorphic sphere (0::complex) 1" using homeomorphic_def homeomorphic_sym by blast also have "... homeomorphic sphere a r" by (simp add: assms homeomorphic_spheres) finally show ?thesis . qed lemma homeomorphic_simple_path_images: fixes \1 :: "real \ 'a::t2_space" and \2 :: "real \ 'b::t2_space" assumes "simple_path \1" and loop: "pathfinish \1 = pathstart \1" assumes "simple_path \2" and loop: "pathfinish \2 = pathstart \2" shows "(path_image \1) homeomorphic (path_image \2)" by (meson assms homeomorphic_simple_path_image_circle homeomorphic_sym homeomorphic_trans loop pi_gt_zero) subsection\Dimension-based conditions for various homeomorphisms\ lemma homeomorphic_subspaces_eq: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "subspace S" "subspace T" shows "S homeomorphic T \ dim S = dim T" proof assume "S homeomorphic T" then obtain f g where hom: "homeomorphism S T f g" using homeomorphic_def by blast show "dim S = dim T" proof (rule order_antisym) show "dim S \ dim T" by (metis assms dual_order.refl inj_onI homeomorphism_cont1 [OF hom] homeomorphism_apply1 [OF hom] homeomorphism_image1 [OF hom] continuous_injective_image_subspace_dim_le) show "dim T \ dim S" by (metis assms dual_order.refl inj_onI homeomorphism_cont2 [OF hom] homeomorphism_apply2 [OF hom] homeomorphism_image2 [OF hom] continuous_injective_image_subspace_dim_le) qed next assume "dim S = dim T" then show "S homeomorphic T" by (simp add: assms homeomorphic_subspaces) qed lemma homeomorphic_affine_sets_eq: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "affine S" "affine T" shows "S homeomorphic T \ aff_dim S = aff_dim T" proof (cases "S = {} \ T = {}") case True then show ?thesis using assms homeomorphic_affine_sets by force next case False then obtain a b where "a \ S" "b \ T" by blast then have "subspace ((+) (- a) ` S)" "subspace ((+) (- b) ` T)" using affine_diffs_subspace assms by blast+ then show ?thesis by (metis affine_imp_convex assms homeomorphic_affine_sets homeomorphic_convex_sets) qed lemma homeomorphic_hyperplanes_eq: fixes a :: "'a::euclidean_space" and c :: "'b::euclidean_space" assumes "a \ 0" "c \ 0" shows "({x. a \ x = b} homeomorphic {x. c \ x = d} \ DIM('a) = DIM('b))" apply (auto simp: homeomorphic_affine_sets_eq affine_hyperplane assms) by (metis DIM_positive Suc_pred) lemma homeomorphic_UNIV_UNIV: shows "(UNIV::'a set) homeomorphic (UNIV::'b set) \ DIM('a::euclidean_space) = DIM('b::euclidean_space)" by (simp add: homeomorphic_subspaces_eq) lemma simply_connected_sphere_gen: assumes "convex S" "bounded S" and 3: "3 \ aff_dim S" shows "simply_connected(rel_frontier S)" proof - have pa: "path_connected (rel_frontier S)" using assms by (simp add: path_connected_sphere_gen) show ?thesis proof (clarsimp simp add: simply_connected_eq_contractible_circlemap pa) fix f assume f: "continuous_on (sphere (0::complex) 1) f" "f ` sphere 0 1 \ rel_frontier S" have eq: "sphere (0::complex) 1 = rel_frontier(cball 0 1)" by simp have "convex (cball (0::complex) 1)" by (rule convex_cball) then obtain c where "homotopic_with_canon (\z. True) (sphere (0::complex) 1) (rel_frontier S) f (\x. c)" apply (rule inessential_spheremap_lowdim_gen [OF _ bounded_cball \convex S\ \bounded S\, where f=f]) using f 3 apply (auto simp: aff_dim_cball) done then show "\a. homotopic_with_canon (\h. True) (sphere 0 1) (rel_frontier S) f (\x. a)" by blast qed qed subsection\more invariance of domain\(*FIX ME title? *) proposition invariance_of_domain_sphere_affine_set_gen: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \ T" and U: "bounded U" "convex U" and "affine T" and affTU: "aff_dim T < aff_dim U" and ope: "openin (top_of_set (rel_frontier U)) S" shows "openin (top_of_set T) (f ` S)" proof (cases "rel_frontier U = {}") case True then show ?thesis using ope openin_subset by force next case False obtain b c where b: "b \ rel_frontier U" and c: "c \ rel_frontier U" and "b \ c" using \bounded U\ rel_frontier_not_sing [of U] subset_singletonD False by fastforce obtain V :: "'a set" where "affine V" and affV: "aff_dim V = aff_dim U - 1" proof (rule choose_affine_subset [OF affine_UNIV]) show "- 1 \ aff_dim U - 1" by (metis aff_dim_empty aff_dim_geq aff_dim_negative_iff affTU diff_0 diff_right_mono not_le) show "aff_dim U - 1 \ aff_dim (UNIV::'a set)" by (metis aff_dim_UNIV aff_dim_le_DIM le_cases not_le zle_diff1_eq) qed auto have SU: "S \ rel_frontier U" using ope openin_imp_subset by auto have homb: "rel_frontier U - {b} homeomorphic V" and homc: "rel_frontier U - {c} homeomorphic V" using homeomorphic_punctured_sphere_affine_gen [of U _ V] by (simp_all add: \affine V\ affV U b c) then obtain g h j k where gh: "homeomorphism (rel_frontier U - {b}) V g h" and jk: "homeomorphism (rel_frontier U - {c}) V j k" by (auto simp: homeomorphic_def) with SU have hgsub: "(h ` g ` (S - {b})) \ S" and kjsub: "(k ` j ` (S - {c})) \ S" by (simp_all add: homeomorphism_def subset_eq) have [simp]: "aff_dim T \ aff_dim V" by (simp add: affTU affV) have "openin (top_of_set T) ((f \ h) ` g ` (S - {b}))" proof (rule invariance_of_domain_affine_sets [OF _ \affine V\]) - show "openin (top_of_set V) (g ` (S - {b}))" - apply (rule homeomorphism_imp_open_map [OF gh]) + have "openin (top_of_set (rel_frontier U - {b})) (S - {b})" by (meson Diff_mono Diff_subset SU ope openin_delete openin_subset_trans order_refl) + then show "openin (top_of_set V) (g ` (S - {b}))" + by (rule homeomorphism_imp_open_map [OF gh]) show "continuous_on (g ` (S - {b})) (f \ h)" - apply (rule continuous_on_compose) - apply (meson Diff_mono SU homeomorphism_def homeomorphism_of_subsets gh set_eq_subset) - using contf continuous_on_subset hgsub by blast + proof (rule continuous_on_compose) + show "continuous_on (g ` (S - {b})) h" + by (meson Diff_mono SU homeomorphism_def homeomorphism_of_subsets gh set_eq_subset) + qed (use contf continuous_on_subset hgsub in blast) show "inj_on (f \ h) (g ` (S - {b}))" using kjsub apply (clarsimp simp add: inj_on_def) by (metis SU b homeomorphism_def inj_onD injf insert_Diff insert_iff gh rev_subsetD) show "(f \ h) ` g ` (S - {b}) \ T" by (metis fim image_comp image_mono hgsub subset_trans) qed (auto simp: assms) moreover have "openin (top_of_set T) ((f \ k) ` j ` (S - {c}))" proof (rule invariance_of_domain_affine_sets [OF _ \affine V\]) show "openin (top_of_set V) (j ` (S - {c}))" - apply (rule homeomorphism_imp_open_map [OF jk]) - by (meson Diff_mono Diff_subset SU ope openin_delete openin_subset_trans order_refl) + by (meson Diff_mono Diff_subset SU ope openin_delete openin_subset_trans order_refl homeomorphism_imp_open_map [OF jk]) show "continuous_on (j ` (S - {c})) (f \ k)" - apply (rule continuous_on_compose) - apply (meson Diff_mono SU homeomorphism_def homeomorphism_of_subsets jk set_eq_subset) - using contf continuous_on_subset kjsub by blast + proof (rule continuous_on_compose) + show "continuous_on (j ` (S - {c})) k" + by (meson Diff_mono SU homeomorphism_def homeomorphism_of_subsets jk set_eq_subset) + qed (use contf continuous_on_subset kjsub in blast) show "inj_on (f \ k) (j ` (S - {c}))" using kjsub apply (clarsimp simp add: inj_on_def) by (metis SU c homeomorphism_def inj_onD injf insert_Diff insert_iff jk rev_subsetD) show "(f \ k) ` j ` (S - {c}) \ T" by (metis fim image_comp image_mono kjsub subset_trans) qed (auto simp: assms) ultimately have "openin (top_of_set T) ((f \ h) ` g ` (S - {b}) \ ((f \ k) ` j ` (S - {c})))" by (rule openin_Un) moreover have "(f \ h) ` g ` (S - {b}) = f ` (S - {b})" proof - have "h ` g ` (S - {b}) = (S - {b})" proof show "h ` g ` (S - {b}) \ S - {b}" using homeomorphism_apply1 [OF gh] SU by (fastforce simp add: image_iff image_subset_iff) show "S - {b} \ h ` g ` (S - {b})" apply clarify by (metis SU subsetD homeomorphism_apply1 [OF gh] image_iff member_remove remove_def) qed then show ?thesis by (metis image_comp) qed moreover have "(f \ k) ` j ` (S - {c}) = f ` (S - {c})" proof - have "k ` j ` (S - {c}) = (S - {c})" proof show "k ` j ` (S - {c}) \ S - {c}" using homeomorphism_apply1 [OF jk] SU by (fastforce simp add: image_iff image_subset_iff) show "S - {c} \ k ` j ` (S - {c})" apply clarify by (metis SU subsetD homeomorphism_apply1 [OF jk] image_iff member_remove remove_def) qed then show ?thesis by (metis image_comp) qed moreover have "f ` (S - {b}) \ f ` (S - {c}) = f ` (S)" using \b \ c\ by blast ultimately show ?thesis by simp qed lemma invariance_of_domain_sphere_affine_set: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \ T" and "r \ 0" "affine T" and affTU: "aff_dim T < DIM('a)" and ope: "openin (top_of_set (sphere a r)) S" shows "openin (top_of_set T) (f ` S)" proof (cases "sphere a r = {}") case True then show ?thesis using ope openin_subset by force next case False show ?thesis proof (rule invariance_of_domain_sphere_affine_set_gen [OF contf injf fim bounded_cball convex_cball \affine T\]) show "aff_dim T < aff_dim (cball a r)" by (metis False affTU aff_dim_cball assms(4) linorder_cases sphere_empty) show "openin (top_of_set (rel_frontier (cball a r))) S" by (simp add: \r \ 0\ ope) qed qed lemma no_embedding_sphere_lowdim: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on (sphere a r) f" and injf: "inj_on f (sphere a r)" and "r > 0" shows "DIM('a) \ DIM('b)" proof - have "False" if "DIM('a) > DIM('b)" proof - have "compact (f ` sphere a r)" using compact_continuous_image by (simp add: compact_continuous_image contf) then have "\ open (f ` sphere a r)" using compact_open by (metis assms(3) image_is_empty not_less_iff_gr_or_eq sphere_eq_empty) then show False using invariance_of_domain_sphere_affine_set [OF contf injf subset_UNIV] \r > 0\ by (metis aff_dim_UNIV affine_UNIV less_irrefl of_nat_less_iff open_openin openin_subtopology_self subtopology_UNIV that) qed then show ?thesis using not_less by blast qed lemma simply_connected_sphere: fixes a :: "'a::euclidean_space" assumes "3 \ DIM('a)" shows "simply_connected(sphere a r)" proof (cases rule: linorder_cases [of r 0]) case less then show ?thesis by simp next case equal then show ?thesis by (auto simp: convex_imp_simply_connected) next case greater then show ?thesis using simply_connected_sphere_gen [of "cball a r"] assms by (simp add: aff_dim_cball) qed lemma simply_connected_sphere_eq: fixes a :: "'a::euclidean_space" shows "simply_connected(sphere a r) \ 3 \ DIM('a) \ r \ 0" (is "?lhs = ?rhs") proof (cases "r \ 0") case True have "simply_connected (sphere a r)" - apply (rule convex_imp_simply_connected) - using True less_eq_real_def by auto + using True less_eq_real_def by (auto intro: convex_imp_simply_connected) with True show ?thesis by auto next case False show ?thesis proof assume L: ?lhs have "False" if "DIM('a) = 1 \ DIM('a) = 2" using that proof assume "DIM('a) = 1" with L show False using connected_sphere_eq simply_connected_imp_connected by (metis False Suc_1 not_less_eq_eq order_refl) next assume "DIM('a) = 2" then have "sphere a r homeomorphic sphere (0::complex) 1" by (metis DIM_complex False homeomorphic_spheres_gen not_less zero_less_one) then have "simply_connected(sphere (0::complex) 1)" using L homeomorphic_simply_connected_eq by blast then obtain a::complex where "homotopic_with_canon (\h. True) (sphere 0 1) (sphere 0 1) id (\x. a)" - apply (simp add: simply_connected_eq_contractible_circlemap) - by (metis continuous_on_id' id_apply image_id subset_refl) + by (metis continuous_on_id' id_apply image_id subset_refl simply_connected_eq_contractible_circlemap) then show False using contractible_sphere contractible_def not_one_le_zero by blast qed with False show ?rhs apply simp by (metis DIM_ge_Suc0 le_antisym not_less_eq_eq numeral_2_eq_2 numeral_3_eq_3) next assume ?rhs with False show ?lhs by (simp add: simply_connected_sphere) qed qed lemma simply_connected_punctured_universe_eq: fixes a :: "'a::euclidean_space" shows "simply_connected(- {a}) \ 3 \ DIM('a)" proof - have [simp]: "a \ rel_interior (cball a 1)" by (simp add: rel_interior_nonempty_interior) have [simp]: "affine hull cball a 1 - {a} = -{a}" by (metis Compl_eq_Diff_UNIV aff_dim_cball aff_dim_lt_full not_less_iff_gr_or_eq zero_less_one) - have "simply_connected(- {a}) \ simply_connected(sphere a 1)" - apply (rule sym) - apply (rule homotopy_eqv_simple_connectedness) - using homotopy_eqv_rel_frontier_punctured_affine_hull [of "cball a 1" a] apply auto - done + have "sphere a 1 homotopy_eqv - {a}" + using homotopy_eqv_rel_frontier_punctured_affine_hull [of "cball a 1" a] by auto + then have "simply_connected(- {a}) \ simply_connected(sphere a 1)" + using homotopy_eqv_simple_connectedness by blast also have "... \ 3 \ DIM('a)" by (simp add: simply_connected_sphere_eq) finally show ?thesis . qed lemma not_simply_connected_circle: fixes a :: complex shows "0 < r \ \ simply_connected(sphere a r)" by (simp add: simply_connected_sphere_eq) proposition simply_connected_punctured_convex: fixes a :: "'a::euclidean_space" assumes "convex S" and 3: "3 \ aff_dim S" shows "simply_connected(S - {a})" proof (cases "a \ rel_interior S") case True then obtain e where "a \ S" "0 < e" and e: "cball a e \ affine hull S \ S" by (auto simp: rel_interior_cball) have con: "convex (cball a e \ affine hull S)" by (simp add: convex_Int) have bo: "bounded (cball a e \ affine hull S)" by (simp add: bounded_Int) have "affine hull S \ interior (cball a e) \ {}" using \0 < e\ \a \ S\ hull_subset by fastforce then have "3 \ aff_dim (affine hull S \ cball a e)" by (simp add: 3 aff_dim_convex_Int_nonempty_interior [OF convex_affine_hull]) also have "... = aff_dim (cball a e \ affine hull S)" by (simp add: Int_commute) finally have "3 \ aff_dim (cball a e \ affine hull S)" . moreover have "rel_frontier (cball a e \ affine hull S) homotopy_eqv S - {a}" proof (rule homotopy_eqv_rel_frontier_punctured_convex) show "a \ rel_interior (cball a e \ affine hull S)" by (meson IntI Int_mono \a \ S\ \0 < e\ e \cball a e \ affine hull S \ S\ ball_subset_cball centre_in_cball dual_order.strict_implies_order hull_inc hull_mono mem_rel_interior_ball) have "closed (cball a e \ affine hull S)" by blast then show "rel_frontier (cball a e \ affine hull S) \ S" - apply (simp add: rel_frontier_def) - using e by blast + by (metis Diff_subset closure_closed dual_order.trans e rel_frontier_def) show "S \ affine hull (cball a e \ affine hull S)" by (metis (no_types, lifting) IntI \a \ S\ \0 < e\ affine_hull_convex_Int_nonempty_interior centre_in_ball convex_affine_hull empty_iff hull_subset inf_commute interior_cball subsetCE subsetI) qed (auto simp: assms con bo) ultimately show ?thesis using homotopy_eqv_simple_connectedness simply_connected_sphere_gen [OF con bo] by blast next case False - show ?thesis - apply (rule contractible_imp_simply_connected) - apply (rule contractible_convex_tweak_boundary_points [OF \convex S\]) - apply (simp add: False rel_interior_subset subset_Diff_insert) + then have "rel_interior S \ S - {a}" + by (simp add: False rel_interior_subset subset_Diff_insert) + moreover have "S - {a} \ closure S" by (meson Diff_subset closure_subset subset_trans) + ultimately show ?thesis + by (metis contractible_imp_simply_connected contractible_convex_tweak_boundary_points [OF \convex S\]) qed corollary simply_connected_punctured_universe: fixes a :: "'a::euclidean_space" assumes "3 \ DIM('a)" shows "simply_connected(- {a})" proof - have [simp]: "affine hull cball a 1 = UNIV" - apply auto - by (metis UNIV_I aff_dim_cball aff_dim_lt_full zero_less_one not_less_iff_gr_or_eq) + by (simp add: aff_dim_cball affine_hull_UNIV) + have "a \ rel_interior (cball a 1)" + by (simp add: rel_interior_interior) + then have "simply_connected (rel_frontier (cball a 1)) = simply_connected (affine hull cball a 1 - {a})" - apply (rule homotopy_eqv_simple_connectedness) - apply (rule homotopy_eqv_rel_frontier_punctured_affine_hull) - apply (force simp: rel_interior_cball intro: homotopy_eqv_simple_connectedness homotopy_eqv_rel_frontier_punctured_affine_hull)+ - done + using homotopy_eqv_rel_frontier_punctured_affine_hull homotopy_eqv_simple_connectedness by blast then show ?thesis using simply_connected_sphere [of a 1, OF assms] by (auto simp: Compl_eq_Diff_UNIV) qed subsection\The power, squaring and exponential functions as covering maps\ proposition covering_space_power_punctured_plane: assumes "0 < n" shows "covering_space (- {0}) (\z::complex. z^n) (- {0})" proof - consider "n = 1" | "2 \ n" using assms by linarith then obtain e where "0 < e" and e: "\w z. cmod(w - z) < e * cmod z \ (w^n = z^n \ w = z)" proof cases assume "n = 1" then show ?thesis by (rule_tac e=1 in that) auto next assume "2 \ n" have eq_if_pow_eq: "w = z" if lt: "cmod (w - z) < 2 * sin (pi / real n) * cmod z" and eq: "w^n = z^n" for w z proof (cases "z = 0") case True with eq assms show ?thesis by (auto simp: power_0_left) next case False then have "z \ 0" by auto have "(w/z)^n = 1" by (metis False divide_self_if eq power_divide power_one) then obtain j where j: "w / z = exp (2 * of_real pi * \ * j / n)" and "j < n" using Suc_leI assms \2 \ n\ complex_roots_unity [THEN eqset_imp_iff, of n "w/z"] by force have "cmod (w/z - 1) < 2 * sin (pi / real n)" using lt assms \z \ 0\ by (simp add: field_split_simps norm_divide) then have "cmod (exp (\ * of_real (2 * pi * j / n)) - 1) < 2 * sin (pi / real n)" by (simp add: j field_simps) then have "2 * \sin((2 * pi * j / n) / 2)\ < 2 * sin (pi / real n)" by (simp only: dist_exp_i_1) then have sin_less: "sin((pi * j / n)) < sin (pi / real n)" by (simp add: field_simps) then have "w / z = 1" proof (cases "j = 0") case True then show ?thesis by (auto simp: j) next case False then have "sin (pi / real n) \ sin((pi * j / n))" proof (cases "j / n \ 1/2") case True show ?thesis - apply (rule sin_monotone_2pi_le) using \j \ 0 \ \j < n\ True - apply (auto simp: field_simps intro: order_trans [of _ 0]) - done + by (intro sin_monotone_2pi_le) (auto simp: field_simps intro: order_trans [of _ 0]) next case False then have seq: "sin(pi * j / n) = sin(pi * (n - j) / n)" using \j < n\ by (simp add: algebra_simps diff_divide_distrib of_nat_diff) show ?thesis - apply (simp only: seq) - apply (rule sin_monotone_2pi_le) + unfolding seq using \j < n\ False - apply (auto simp: field_simps intro: order_trans [of _ 0]) - done + by (intro sin_monotone_2pi_le) (auto simp: field_simps intro: order_trans [of _ 0]) qed with sin_less show ?thesis by force qed then show ?thesis by simp qed show ?thesis - apply (rule_tac e = "2 * sin(pi / n)" in that) - apply (force simp: \2 \ n\ sin_pi_divide_n_gt_0) - apply (meson eq_if_pow_eq) - done + proof + show "0 < 2 * sin (pi / real n)" + by (force simp: \2 \ n\ sin_pi_divide_n_gt_0) + qed (meson eq_if_pow_eq) qed have zn1: "continuous_on (- {0}) (\z::complex. z^n)" by (rule continuous_intros)+ have zn2: "(\z::complex. z^n) ` (- {0}) = - {0}" using assms by (auto simp: image_def elim: exists_complex_root_nonzero [where n = n]) have zn3: "\T. z^n \ T \ open T \ 0 \ T \ (\v. \v = -{0} \ (\z. z ^ n) -` T \ (\u\v. open u \ 0 \ u) \ pairwise disjnt v \ (\u\v. Ex (homeomorphism u T (\z. z^n))))" if "z \ 0" for z::complex proof - define d where "d \ min (1/2) (e/4) * norm z" have "0 < d" by (simp add: d_def \0 < e\ \z \ 0\) have iff_x_eq_y: "x^n = y^n \ x = y" if eq: "w^n = z^n" and x: "x \ ball w d" and y: "y \ ball w d" for w x y proof - have [simp]: "norm z = norm w" using that by (simp add: assms power_eq_imp_eq_norm) show ?thesis proof (cases "w = 0") case True with \z \ 0\ assms eq show ?thesis by (auto simp: power_0_left) next case False have "cmod (x - y) < 2*d" using x y by (simp add: dist_norm [symmetric]) (metis dist_commute mult_2 dist_triangle_less_add) also have "... \ 2 * e / 4 * norm w" using \e > 0\ by (simp add: d_def min_mult_distrib_right) also have "... = e * (cmod w / 2)" by simp also have "... \ e * cmod y" - apply (rule mult_left_mono) - using \e > 0\ y - apply (simp_all add: dist_norm d_def min_mult_distrib_right del: divide_const_simps) - apply (metis dist_0_norm dist_complex_def dist_triangle_half_l linorder_not_less order_less_irrefl) - done + proof (rule mult_left_mono) + have "cmod (w - y) < cmod w / 2 \ cmod w / 2 \ cmod y" + by (metis (no_types) dist_0_norm dist_norm norm_triangle_half_l not_le order_less_irrefl) + then show "cmod w / 2 \ cmod y" + using y by (simp add: dist_norm d_def min_mult_distrib_right) + qed (use \e > 0\ in auto) finally have "cmod (x - y) < e * cmod y" . then show ?thesis by (rule e) qed qed then have inj: "inj_on (\w. w^n) (ball z d)" by (simp add: inj_on_def) have cont: "continuous_on (ball z d) (\w. w ^ n)" by (intro continuous_intros) have noncon: "\ (\w::complex. w^n) constant_on UNIV" by (metis UNIV_I assms constant_on_def power_one zero_neq_one zero_power) have im_eq: "(\w. w^n) ` ball z' d = (\w. w^n) ` ball z d" if z': "z'^n = z^n" for z' proof - have nz': "norm z' = norm z" using that assms power_eq_imp_eq_norm by blast have "(w \ (\w. w^n) ` ball z' d) = (w \ (\w. w^n) ` ball z d)" for w proof (cases "w=0") case True with assms show ?thesis by (simp add: image_def ball_def nz') next case False have "z' \ 0" using \z \ 0\ nz' by force - have [simp]: "(z*x / z')^n = x^n" if "x \ 0" for x + have 1: "(z*x / z')^n = x^n" if "x \ 0" for x using z' that by (simp add: field_simps \z \ 0\) - have [simp]: "cmod (z - z * x / z') = cmod (z' - x)" if "x \ 0" for x + have 2: "cmod (z - z * x / z') = cmod (z' - x)" if "x \ 0" for x proof - have "cmod (z - z * x / z') = cmod z * cmod (1 - x / z')" by (metis (no_types) ab_semigroup_mult_class.mult_ac(1) divide_complex_def mult.right_neutral norm_mult right_diff_distrib') also have "... = cmod z' * cmod (1 - x / z')" by (simp add: nz') also have "... = cmod (z' - x)" by (simp add: \z' \ 0\ diff_divide_eq_iff norm_divide) finally show ?thesis . qed - have [simp]: "(z'*x / z)^n = x^n" if "x \ 0" for x + have 3: "(z'*x / z)^n = x^n" if "x \ 0" for x using z' that by (simp add: field_simps \z \ 0\) - have [simp]: "cmod (z' - z' * x / z) = cmod (z - x)" if "x \ 0" for x + have 4: "cmod (z' - z' * x / z) = cmod (z - x)" if "x \ 0" for x proof - have "cmod (z * (1 - x * inverse z)) = cmod (z - x)" by (metis \z \ 0\ diff_divide_distrib divide_complex_def divide_self_if nonzero_eq_divide_eq semiring_normalization_rules(7)) then show ?thesis by (metis (no_types) mult.assoc divide_complex_def mult.right_neutral norm_mult nz' right_diff_distrib') qed show ?thesis - unfolding image_def ball_def - apply safe - apply simp_all - apply (rule_tac x="z/z' * x" in exI) - using assms False apply (simp add: dist_norm) - apply (rule_tac x="z'/z * x" in exI) - using assms False apply (simp add: dist_norm) - done + by (simp add: set_eq_iff image_def ball_def) (metis 1 2 3 4 diff_zero dist_norm nz') qed then show ?thesis by blast qed have ex_ball: "\B. (\z'. B = ball z' d \ z'^n = z^n) \ x \ B" if "x \ 0" and eq: "x^n = w^n" and dzw: "dist z w < d" for x w proof - have "w \ 0" by (metis assms power_eq_0_iff that(1) that(2)) have [simp]: "cmod x = cmod w" using assms power_eq_imp_eq_norm eq by blast have [simp]: "cmod (x * z / w - x) = cmod (z - w)" proof - have "cmod (x * z / w - x) = cmod x * cmod (z / w - 1)" by (metis (no_types) mult.right_neutral norm_mult right_diff_distrib' times_divide_eq_right) also have "... = cmod w * cmod (z / w - 1)" by simp also have "... = cmod (z - w)" by (simp add: \w \ 0\ divide_diff_eq_iff nonzero_norm_divide) finally show ?thesis . qed show ?thesis - apply (rule_tac x="ball (z / w * x) d" in exI) - using \d > 0\ that - apply (simp add: ball_eq_ball_iff) - apply (simp add: \z \ 0\ \w \ 0\ field_simps) - apply (simp add: dist_norm) - done + proof (intro exI conjI) + show "(z / w * x) ^ n = z ^ n" + by (metis \w \ 0\ eq nonzero_eq_divide_eq power_mult_distrib) + show "x \ ball (z / w * x) d" + using \d > 0\ that + by (simp add: ball_eq_ball_iff \z \ 0\ \w \ 0\ field_simps) (simp add: dist_norm) + qed auto qed show ?thesis proof (rule exI, intro conjI) show "z ^ n \ (\w. w ^ n) ` ball z d" using \d > 0\ by simp show "open ((\w. w ^ n) ` ball z d)" by (rule invariance_of_domain [OF cont open_ball inj]) show "0 \ (\w. w ^ n) ` ball z d" using \z \ 0\ assms by (force simp: d_def) show "\v. \v = - {0} \ (\z. z ^ n) -` (\w. w ^ n) ` ball z d \ (\u\v. open u \ 0 \ u) \ disjoint v \ (\u\v. Ex (homeomorphism u ((\w. w ^ n) ` ball z d) (\z. z ^ n)))" proof (rule exI, intro ballI conjI) show "\{ball z' d |z'. z'^n = z^n} = - {0} \ (\z. z ^ n) -` (\w. w ^ n) ` ball z d" (is "?l = ?r") proof - show "?l \ ?r" - apply auto - apply (simp add: assms d_def power_eq_imp_eq_norm that) - by (metis im_eq image_eqI mem_ball) + have "\z'. cmod z' < d \ z' ^ n \ z ^ n" + by (auto simp add: assms d_def power_eq_imp_eq_norm that) + then show "?l \ ?r" + by auto (metis im_eq image_eqI mem_ball) show "?r \ ?l" by auto (meson ex_ball) qed show "\u. u \ {ball z' d |z'. z' ^ n = z ^ n} \ 0 \ u" by (force simp add: assms d_def power_eq_imp_eq_norm that) show "disjoint {ball z' d |z'. z' ^ n = z ^ n}" proof (clarsimp simp add: pairwise_def disjnt_iff) fix \ \ x assume "\^n = z^n" "\^n = z^n" "ball \ d \ ball \ d" and "dist \ x < d" "dist \ x < d" then have "dist \ \ < d+d" using dist_triangle_less_add by blast then have "cmod (\ - \) < 2*d" by (simp add: dist_norm) also have "... \ e * cmod z" using mult_right_mono \0 < e\ that by (auto simp: d_def) finally have "cmod (\ - \) < e * cmod z" . with e have "\ = \" by (metis \\^n = z^n\ \\^n = z^n\ assms power_eq_imp_eq_norm) then show "False" using \ball \ d \ ball \ d\ by blast qed show "Ex (homeomorphism u ((\w. w ^ n) ` ball z d) (\z. z ^ n))" if "u \ {ball z' d |z'. z' ^ n = z ^ n}" for u proof (rule invariance_of_domain_homeomorphism [of "u" "\z. z^n"]) show "open u" using that by auto show "continuous_on u (\z. z ^ n)" by (intro continuous_intros) show "inj_on (\z. z ^ n) u" using that by (auto simp: iff_x_eq_y inj_on_def) show "\g. homeomorphism u ((\z. z ^ n) ` u) (\z. z ^ n) g \ Ex (homeomorphism u ((\w. w ^ n) ` ball z d) (\z. z ^ n))" using im_eq that by clarify metis qed auto qed auto qed qed show ?thesis using assms apply (simp add: covering_space_def zn1 zn2) apply (subst zn2 [symmetric]) - apply (simp add: openin_open_eq open_Compl) - apply (blast intro: zn3) + apply (simp add: openin_open_eq open_Compl zn3) done qed corollary covering_space_square_punctured_plane: "covering_space (- {0}) (\z::complex. z^2) (- {0})" by (simp add: covering_space_power_punctured_plane) proposition covering_space_exp_punctured_plane: "covering_space UNIV (\z::complex. exp z) (- {0})" proof (simp add: covering_space_def, intro conjI ballI) show "continuous_on UNIV (\z::complex. exp z)" by (rule continuous_on_exp [OF continuous_on_id]) show "range exp = - {0::complex}" by auto (metis exp_Ln range_eqI) show "\T. z \ T \ openin (top_of_set (- {0})) T \ (\v. \v = exp -` T \ (\u\v. open u) \ disjoint v \ (\u\v. \q. homeomorphism u T exp q))" if "z \ - {0::complex}" for z proof - have "z \ 0" using that by auto - have inj_exp: "inj_on exp (ball (Ln z) 1)" - apply (rule inj_on_subset [OF inj_on_exp_pi [of "Ln z"]]) + have "ball (Ln z) 1 \ ball (Ln z) pi" using pi_ge_two by (simp add: ball_subset_ball_iff) + then have inj_exp: "inj_on exp (ball (Ln z) 1)" + using inj_on_exp_pi inj_on_subset by blast define \ where "\ \ range (\n. (\x. x + of_real (2 * of_int n * pi) * \) ` (ball(Ln z) 1))" show ?thesis proof (intro exI conjI) show "z \ exp ` (ball(Ln z) 1)" by (metis \z \ 0\ centre_in_ball exp_Ln rev_image_eqI zero_less_one) have "open (- {0::complex})" by blast - moreover have "inj_on exp (ball (Ln z) 1)" - apply (rule inj_on_subset [OF inj_on_exp_pi [of "Ln z"]]) - using pi_ge_two by (simp add: ball_subset_ball_iff) - ultimately show "openin (top_of_set (- {0})) (exp ` ball (Ln z) 1)" + with inj_exp show "openin (top_of_set (- {0})) (exp ` ball (Ln z) 1)" by (auto simp: openin_open_eq invariance_of_domain continuous_on_exp [OF continuous_on_id]) show "\\ = exp -` exp ` ball (Ln z) 1" by (force simp: \_def Complex_Transcendental.exp_eq image_iff) show "\V\\. open V" by (auto simp: \_def inj_on_def continuous_intros invariance_of_domain) have xy: "2 \ cmod (2 * of_int x * of_real pi * \ - 2 * of_int y * of_real pi * \)" if "x < y" for x y proof - have "1 \ abs (x - y)" using that by linarith then have "1 \ cmod (of_int x - of_int y) * 1" by (metis mult.right_neutral norm_of_int of_int_1_le_iff of_int_abs of_int_diff) also have "... \ cmod (of_int x - of_int y) * of_real pi" - apply (rule mult_left_mono) - using pi_ge_two by auto + using pi_ge_two + by (intro mult_left_mono) auto also have "... \ cmod ((of_int x - of_int y) * of_real pi * \)" by (simp add: norm_mult) also have "... \ cmod (of_int x * of_real pi * \ - of_int y * of_real pi * \)" by (simp add: algebra_simps) finally have "1 \ cmod (of_int x * of_real pi * \ - of_int y * of_real pi * \)" . then have "2 * 1 \ cmod (2 * (of_int x * of_real pi * \ - of_int y * of_real pi * \))" by (metis mult_le_cancel_left_pos norm_mult_numeral1 zero_less_numeral) then show ?thesis by (simp add: algebra_simps) qed show "disjoint \" apply (clarsimp simp add: \_def pairwise_def disjnt_def add.commute [of _ "x*y" for x y] - ball_eq_ball_iff) - apply (rule disjoint_ballI) + ball_eq_ball_iff intro!: disjoint_ballI) apply (auto simp: dist_norm neq_iff) by (metis norm_minus_commute xy)+ show "\u\\. \q. homeomorphism u (exp ` ball (Ln z) 1) exp q" proof fix u assume "u \ \" then obtain n where n: "u = (\x. x + of_real (2 * of_int n * pi) * \) ` (ball(Ln z) 1)" by (auto simp: \_def) have "compact (cball (Ln z) 1)" by simp moreover have "continuous_on (cball (Ln z) 1) exp" by (rule continuous_on_exp [OF continuous_on_id]) moreover have "inj_on exp (cball (Ln z) 1)" apply (rule inj_on_subset [OF inj_on_exp_pi [of "Ln z"]]) using pi_ge_two by (simp add: cball_subset_ball_iff) ultimately obtain \ where hom: "homeomorphism (cball (Ln z) 1) (exp ` cball (Ln z) 1) exp \" using homeomorphism_compact by blast have eq1: "exp ` u = exp ` ball (Ln z) 1" - unfolding n - apply (auto simp: algebra_simps) - apply (rename_tac w) - apply (rule_tac x = "w + \ * (of_int n * (of_real pi * 2))" in image_eqI) + apply (auto simp: algebra_simps n) + apply (rule_tac x = "_ + \ * (of_int n * (of_real pi * 2))" in image_eqI) apply (auto simp: image_iff) done have \exp: "\ (exp x) + 2 * of_int n * of_real pi * \ = x" if "x \ u" for x proof - have "exp x = exp (x - 2 * of_int n * of_real pi * \)" by (simp add: exp_eq) then have "\ (exp x) = \ (exp (x - 2 * of_int n * of_real pi * \))" by simp also have "... = x - 2 * of_int n * of_real pi * \" - apply (rule homeomorphism_apply1 [OF hom]) - using \x \ u\ by (auto simp: n) + using \x \ u\ by (auto simp: n intro: homeomorphism_apply1 [OF hom]) finally show ?thesis by simp qed have exp2n: "exp (\ (exp x) + 2 * of_int n * complex_of_real pi * \) = exp x" if "dist (Ln z) x < 1" for x using that by (auto simp: exp_eq homeomorphism_apply1 [OF hom]) - have cont: "continuous_on (exp ` ball (Ln z) 1) (\x. \ x + 2 * of_int n * complex_of_real pi * \)" - apply (intro continuous_intros) - apply (rule continuous_on_subset [OF homeomorphism_cont2 [OF hom]]) - apply (force simp:) - done + have "continuous_on (exp ` ball (Ln z) 1) \" + by (meson ball_subset_cball continuous_on_subset hom homeomorphism_cont2 image_mono) + then have cont: "continuous_on (exp ` ball (Ln z) 1) (\x. \ x + 2 * of_int n * complex_of_real pi * \)" + by (intro continuous_intros) show "\q. homeomorphism u (exp ` ball (Ln z) 1) exp q" apply (rule_tac x="(\x. x + of_real(2 * n * pi) * \) \ \" in exI) unfolding homeomorphism_def apply (intro conjI ballI eq1 continuous_on_exp [OF continuous_on_id]) apply (auto simp: \exp exp2n cont n) - apply (simp add: homeomorphism_apply1 [OF hom]) - using hom homeomorphism_apply1 apply (force simp: image_iff) + apply (force simp: image_iff homeomorphism_apply1 [OF hom])+ done qed qed qed qed subsection\Hence the Borsukian results about mappings into circles\(*FIX ME title *) lemma inessential_eq_continuous_logarithm: fixes f :: "'a::real_normed_vector \ complex" shows "(\a. homotopic_with_canon (\h. True) S (-{0}) f (\t. a)) \ (\g. continuous_on S g \ (\x \ S. f x = exp(g x)))" (is "?lhs \ ?rhs") proof assume ?lhs thus ?rhs by (metis covering_space_lift_inessential_function covering_space_exp_punctured_plane) next assume ?rhs then obtain g where contg: "continuous_on S g" and f: "\x. x \ S \ f x = exp(g x)" by metis obtain a where "homotopic_with_canon (\h. True) S (- {of_real 0}) (exp \ g) (\x. a)" proof (rule nullhomotopic_through_contractible [OF contg subset_UNIV _ _ contractible_UNIV]) show "continuous_on (UNIV::complex set) exp" by (intro continuous_intros) show "range exp \ - {0}" by auto qed force then have "homotopic_with_canon (\h. True) S (- {0}) f (\t. a)" using f homotopic_with_eq by fastforce then show ?lhs .. qed corollary inessential_imp_continuous_logarithm_circle: fixes f :: "'a::real_normed_vector \ complex" assumes "homotopic_with_canon (\h. True) S (sphere 0 1) f (\t. a)" obtains g where "continuous_on S g" and "\x. x \ S \ f x = exp(g x)" proof - have "homotopic_with_canon (\h. True) S (- {0}) f (\t. a)" using assms homotopic_with_subset_right by fastforce then show ?thesis by (metis inessential_eq_continuous_logarithm that) qed lemma inessential_eq_continuous_logarithm_circle: fixes f :: "'a::real_normed_vector \ complex" shows "(\a. homotopic_with_canon (\h. True) S (sphere 0 1) f (\t. a)) \ (\g. continuous_on S g \ (\x \ S. f x = exp(\ * of_real(g x))))" (is "?lhs \ ?rhs") proof assume L: ?lhs then obtain g where contg: "continuous_on S g" and g: "\x. x \ S \ f x = exp(g x)" using inessential_imp_continuous_logarithm_circle by blast have "f ` S \ sphere 0 1" by (metis L homotopic_with_imp_subset1) then have "\x. x \ S \ Re (g x) = 0" using g by auto then show ?rhs - apply (rule_tac x="Im \ g" in exI) - apply (intro conjI contg continuous_intros) - apply (auto simp: Euler g) - done + by (rule_tac x="Im \ g" in exI) (auto simp: Euler g intro: contg continuous_intros) next assume ?rhs then obtain g where contg: "continuous_on S g" and g: "\x. x \ S \ f x = exp(\* of_real(g x))" by metis obtain a where "homotopic_with_canon (\h. True) S (sphere 0 1) ((exp \ (\z. \*z)) \ (of_real \ g)) (\x. a)" proof (rule nullhomotopic_through_contractible) show "continuous_on S (complex_of_real \ g)" by (intro conjI contg continuous_intros) show "(complex_of_real \ g) ` S \ \" by auto show "continuous_on \ (exp \ (*)\)" by (intro continuous_intros) show "(exp \ (*)\) ` \ \ sphere 0 1" by (auto simp: complex_is_Real_iff) qed (auto simp: convex_Reals convex_imp_contractible) moreover have "\x. x \ S \ (exp \ (*)\ \ (complex_of_real \ g)) x = f x" by (simp add: g) ultimately have "homotopic_with_canon (\h. True) S (sphere 0 1) f (\t. a)" using homotopic_with_eq by force then show ?lhs .. qed proposition homotopic_with_sphere_times: fixes f :: "'a::real_normed_vector \ complex" assumes hom: "homotopic_with_canon (\x. True) S (sphere 0 1) f g" and conth: "continuous_on S h" and hin: "\x. x \ S \ h x \ sphere 0 1" shows "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x * h x) (\x. g x * h x)" proof - obtain k where contk: "continuous_on ({0..1::real} \ S) k" and kim: "k ` ({0..1} \ S) \ sphere 0 1" and k0: "\x. k(0, x) = f x" and k1: "\x. k(1, x) = g x" using hom by (auto simp: homotopic_with_def) show ?thesis apply (simp add: homotopic_with) apply (rule_tac x="\z. k z*(h \ snd)z" in exI) - apply (intro conjI contk continuous_intros) - apply (simp add: conth) - using kim hin apply (force simp: norm_mult k0 k1)+ - done + using kim hin by (fastforce simp: conth norm_mult k0 k1 intro!: contk continuous_intros)+ qed proposition homotopic_circlemaps_divide: fixes f :: "'a::real_normed_vector \ complex" shows "homotopic_with_canon (\x. True) S (sphere 0 1) f g \ continuous_on S f \ f ` S \ sphere 0 1 \ continuous_on S g \ g ` S \ sphere 0 1 \ (\c. homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. c))" proof - have "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. 1)" if "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. c)" for c proof - have "S = {} \ path_component (sphere 0 1) 1 c" using homotopic_with_imp_subset2 [OF that] path_connected_sphere [of "0::complex" 1] by (auto simp: path_connected_component) then have "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. 1) (\x. c)" by (simp add: homotopic_constant_maps) then show ?thesis using homotopic_with_symD homotopic_with_trans that by blast qed then have *: "(\c. homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. c)) \ homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. 1)" by auto have "homotopic_with_canon (\x. True) S (sphere 0 1) f g \ continuous_on S f \ f ` S \ sphere 0 1 \ continuous_on S g \ g ` S \ sphere 0 1 \ homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. 1)" (is "?lhs \ ?rhs") proof assume L: ?lhs have geq1 [simp]: "\x. x \ S \ cmod (g x) = 1" using homotopic_with_imp_subset2 [OF L] by (simp add: image_subset_iff) have cont: "continuous_on S (inverse \ g)" - apply (rule continuous_intros) - using homotopic_with_imp_continuous [OF L] apply blast - apply (rule continuous_on_subset [of "sphere 0 1", OF continuous_on_inverse]) - apply (auto) - done + proof (rule continuous_intros) + show "continuous_on S g" + using homotopic_with_imp_continuous [OF L] by blast + show "continuous_on (g ` S) inverse" + by (rule continuous_on_subset [of "sphere 0 1", OF continuous_on_inverse]) auto + qed + have [simp]: "\x. x \ S \ g x \ 0" + using geq1 by fastforce have "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. 1)" - using homotopic_with_sphere_times [OF L cont] - apply (rule homotopic_with_eq) - apply (auto simp: division_ring_class.divide_inverse norm_inverse) - by (metis geq1 norm_zero right_inverse zero_neq_one) + apply (rule homotopic_with_eq [OF homotopic_with_sphere_times [OF L cont]]) + by (auto simp: divide_inverse norm_inverse) with L show ?rhs by (auto simp: homotopic_with_imp_continuous dest: homotopic_with_imp_subset1 homotopic_with_imp_subset2) next assume ?rhs then show ?lhs by (elim conjE homotopic_with_eq [OF homotopic_with_sphere_times]; force) qed then show ?thesis by (simp add: *) qed subsection\Upper and lower hemicontinuous functions\ text\And relation in the case of preimage map to open and closed maps, and fact that upper and lower hemicontinuity together imply continuity in the sense of the Hausdorff metric (at points where the function gives a bounded and nonempty set).\ text\Many similar proofs below.\ lemma upper_hemicontinuous: assumes "\x. x \ S \ f x \ T" shows "((\U. openin (top_of_set T) U \ openin (top_of_set S) {x \ S. f x \ U}) \ (\U. closedin (top_of_set T) U \ closedin (top_of_set S) {x \ S. f x \ U \ {}}))" (is "?lhs = ?rhs") proof (intro iffI allI impI) fix U assume * [rule_format]: ?lhs and "closedin (top_of_set T) U" then have "openin (top_of_set T) (T - U)" by (simp add: openin_diff) then have "openin (top_of_set S) {x \ S. f x \ T - U}" using * [of "T-U"] by blast moreover have "S - {x \ S. f x \ T - U} = {x \ S. f x \ U \ {}}" using assms by blast ultimately show "closedin (top_of_set S) {x \ S. f x \ U \ {}}" by (simp add: openin_closedin_eq) next fix U assume * [rule_format]: ?rhs and "openin (top_of_set T) U" then have "closedin (top_of_set T) (T - U)" by (simp add: closedin_diff) then have "closedin (top_of_set S) {x \ S. f x \ (T - U) \ {}}" using * [of "T-U"] by blast moreover have "{x \ S. f x \ (T - U) \ {}} = S - {x \ S. f x \ U}" using assms by auto ultimately show "openin (top_of_set S) {x \ S. f x \ U}" by (simp add: openin_closedin_eq) qed lemma lower_hemicontinuous: assumes "\x. x \ S \ f x \ T" shows "((\U. closedin (top_of_set T) U \ closedin (top_of_set S) {x \ S. f x \ U}) \ (\U. openin (top_of_set T) U \ openin (top_of_set S) {x \ S. f x \ U \ {}}))" (is "?lhs = ?rhs") proof (intro iffI allI impI) fix U assume * [rule_format]: ?lhs and "openin (top_of_set T) U" then have "closedin (top_of_set T) (T - U)" by (simp add: closedin_diff) then have "closedin (top_of_set S) {x \ S. f x \ T-U}" using * [of "T-U"] by blast moreover have "{x \ S. f x \ T-U} = S - {x \ S. f x \ U \ {}}" using assms by auto ultimately show "openin (top_of_set S) {x \ S. f x \ U \ {}}" by (simp add: openin_closedin_eq) next fix U assume * [rule_format]: ?rhs and "closedin (top_of_set T) U" then have "openin (top_of_set T) (T - U)" by (simp add: openin_diff) then have "openin (top_of_set S) {x \ S. f x \ (T - U) \ {}}" using * [of "T-U"] by blast moreover have "S - {x \ S. f x \ (T - U) \ {}} = {x \ S. f x \ U}" using assms by blast ultimately show "closedin (top_of_set S) {x \ S. f x \ U}" by (simp add: openin_closedin_eq) qed lemma open_map_iff_lower_hemicontinuous_preimage: assumes "f ` S \ T" shows "((\U. openin (top_of_set S) U \ openin (top_of_set T) (f ` U)) \ (\U. closedin (top_of_set S) U \ closedin (top_of_set T) {y \ T. {x. x \ S \ f x = y} \ U}))" (is "?lhs = ?rhs") proof (intro iffI allI impI) fix U assume * [rule_format]: ?lhs and "closedin (top_of_set S) U" then have "openin (top_of_set S) (S - U)" by (simp add: openin_diff) then have "openin (top_of_set T) (f ` (S - U))" using * [of "S-U"] by blast moreover have "T - (f ` (S - U)) = {y \ T. {x \ S. f x = y} \ U}" using assms by blast ultimately show "closedin (top_of_set T) {y \ T. {x \ S. f x = y} \ U}" by (simp add: openin_closedin_eq) next fix U assume * [rule_format]: ?rhs and opeSU: "openin (top_of_set S) U" then have "closedin (top_of_set S) (S - U)" by (simp add: closedin_diff) then have "closedin (top_of_set T) {y \ T. {x \ S. f x = y} \ S - U}" using * [of "S-U"] by blast moreover have "{y \ T. {x \ S. f x = y} \ S - U} = T - (f ` U)" using assms openin_imp_subset [OF opeSU] by auto ultimately show "openin (top_of_set T) (f ` U)" using assms openin_imp_subset [OF opeSU] by (force simp: openin_closedin_eq) qed lemma closed_map_iff_upper_hemicontinuous_preimage: assumes "f ` S \ T" shows "((\U. closedin (top_of_set S) U \ closedin (top_of_set T) (f ` U)) \ (\U. openin (top_of_set S) U \ openin (top_of_set T) {y \ T. {x. x \ S \ f x = y} \ U}))" (is "?lhs = ?rhs") proof (intro iffI allI impI) fix U assume * [rule_format]: ?lhs and opeSU: "openin (top_of_set S) U" then have "closedin (top_of_set S) (S - U)" by (simp add: closedin_diff) then have "closedin (top_of_set T) (f ` (S - U))" using * [of "S-U"] by blast moreover have "f ` (S - U) = T - {y \ T. {x. x \ S \ f x = y} \ U}" using assms openin_imp_subset [OF opeSU] by auto ultimately show "openin (top_of_set T) {y \ T. {x. x \ S \ f x = y} \ U}" using assms openin_imp_subset [OF opeSU] by (force simp: openin_closedin_eq) next fix U assume * [rule_format]: ?rhs and cloSU: "closedin (top_of_set S) U" then have "openin (top_of_set S) (S - U)" by (simp add: openin_diff) then have "openin (top_of_set T) {y \ T. {x \ S. f x = y} \ S - U}" using * [of "S-U"] by blast moreover have "(f ` U) = T - {y \ T. {x \ S. f x = y} \ S - U}" using assms closedin_imp_subset [OF cloSU] by auto ultimately show "closedin (top_of_set T) (f ` U)" by (simp add: openin_closedin_eq) qed proposition upper_lower_hemicontinuous_explicit: fixes T :: "('b::{real_normed_vector,heine_borel}) set" assumes fST: "\x. x \ S \ f x \ T" and ope: "\U. openin (top_of_set T) U \ openin (top_of_set S) {x \ S. f x \ U}" and clo: "\U. closedin (top_of_set T) U \ closedin (top_of_set S) {x \ S. f x \ U}" and "x \ S" "0 < e" and bofx: "bounded(f x)" and fx_ne: "f x \ {}" obtains d where "0 < d" "\x'. \x' \ S; dist x x' < d\ \ (\y \ f x. \y'. y' \ f x' \ dist y y' < e) \ (\y' \ f x'. \y. y \ f x \ dist y' y < e)" proof - have "openin (top_of_set T) (T \ (\a\f x. \b\ball 0 e. {a + b}))" by (auto simp: open_sums openin_open_Int) with ope have "openin (top_of_set S) {u \ S. f u \ T \ (\a\f x. \b\ball 0 e. {a + b})}" by blast with \0 < e\ \x \ S\ obtain d1 where "d1 > 0" and d1: "\x'. \x' \ S; dist x' x < d1\ \ f x' \ T \ f x' \ (\a \ f x. \b \ ball 0 e. {a + b})" by (force simp: openin_euclidean_subtopology_iff dest: fST) have oo: "\U. openin (top_of_set T) U \ openin (top_of_set S) {x \ S. f x \ U \ {}}" apply (rule lower_hemicontinuous [THEN iffD1, rule_format]) using fST clo by auto have "compact (closure(f x))" by (simp add: bofx) moreover have "closure(f x) \ (\a \ f x. ball a (e/2))" using \0 < e\ by (force simp: closure_approachable simp del: divide_const_simps) ultimately obtain C where "C \ f x" "finite C" "closure(f x) \ (\a \ C. ball a (e/2))" apply (rule compactE, force) by (metis finite_subset_image) then have fx_cover: "f x \ (\a \ C. ball a (e/2))" by (meson closure_subset order_trans) with fx_ne have "C \ {}" by blast have xin: "x \ (\a \ C. {x \ S. f x \ T \ ball a (e/2) \ {}})" using \x \ S\ \0 < e\ fST \C \ f x\ by force have "openin (top_of_set S) {x \ S. f x \ (T \ ball a (e/2)) \ {}}" for a by (simp add: openin_open_Int oo) then have "openin (top_of_set S) (\a \ C. {x \ S. f x \ T \ ball a (e/2) \ {}})" by (simp add: Int_assoc openin_INT2 [OF \finite C\ \C \ {}\]) with xin obtain d2 where "d2>0" and d2: "\u v. \u \ S; dist u x < d2; v \ C\ \ f u \ T \ ball v (e/2) \ {}" unfolding openin_euclidean_subtopology_iff using xin by fastforce show ?thesis proof (intro that conjI ballI) show "0 < min d1 d2" using \0 < d1\ \0 < d2\ by linarith next fix x' y assume "x' \ S" "dist x x' < min d1 d2" "y \ f x" then have dd2: "dist x' x < d2" by (auto simp: dist_commute) obtain a where "a \ C" "y \ ball a (e/2)" using fx_cover \y \ f x\ by auto then show "\y'. y' \ f x' \ dist y y' < e" using d2 [OF \x' \ S\ dd2] dist_triangle_half_r by fastforce next fix x' y' assume "x' \ S" "dist x x' < min d1 d2" "y' \ f x'" then have "dist x' x < d1" by (auto simp: dist_commute) then have "y' \ (\a\f x. \b\ball 0 e. {a + b})" using d1 [OF \x' \ S\] \y' \ f x'\ by force then show "\y. y \ f x \ dist y' y < e" - apply auto - by (metis add_diff_cancel_left' dist_norm) + by clarsimp (metis add_diff_cancel_left' dist_norm) qed qed subsection\Complex logs exist on various "well-behaved" sets\ lemma continuous_logarithm_on_contractible: fixes f :: "'a::real_normed_vector \ complex" assumes "continuous_on S f" "contractible S" "\z. z \ S \ f z \ 0" obtains g where "continuous_on S g" "\x. x \ S \ f x = exp(g x)" proof - obtain c where hom: "homotopic_with_canon (\h. True) S (-{0}) f (\x. c)" using nullhomotopic_from_contractible assms by (metis imageE subset_Compl_singleton) then show ?thesis by (metis inessential_eq_continuous_logarithm that) qed lemma continuous_logarithm_on_simply_connected: fixes f :: "'a::real_normed_vector \ complex" assumes contf: "continuous_on S f" and S: "simply_connected S" "locally path_connected S" and f: "\z. z \ S \ f z \ 0" obtains g where "continuous_on S g" "\x. x \ S \ f x = exp(g x)" using covering_space_lift [OF covering_space_exp_punctured_plane S contf] by (metis (full_types) f imageE subset_Compl_singleton) lemma continuous_logarithm_on_cball: fixes f :: "'a::real_normed_vector \ complex" assumes "continuous_on (cball a r) f" and "\z. z \ cball a r \ f z \ 0" obtains h where "continuous_on (cball a r) h" "\z. z \ cball a r \ f z = exp(h z)" using assms continuous_logarithm_on_contractible convex_imp_contractible by blast lemma continuous_logarithm_on_ball: fixes f :: "'a::real_normed_vector \ complex" assumes "continuous_on (ball a r) f" and "\z. z \ ball a r \ f z \ 0" obtains h where "continuous_on (ball a r) h" "\z. z \ ball a r \ f z = exp(h z)" using assms continuous_logarithm_on_contractible convex_imp_contractible by blast lemma continuous_sqrt_on_contractible: fixes f :: "'a::real_normed_vector \ complex" assumes "continuous_on S f" "contractible S" and "\z. z \ S \ f z \ 0" obtains g where "continuous_on S g" "\x. x \ S \ f x = (g x) ^ 2" proof - obtain g where contg: "continuous_on S g" and feq: "\x. x \ S \ f x = exp(g x)" using continuous_logarithm_on_contractible [OF assms] by blast show ?thesis proof show "continuous_on S (\z. exp (g z / 2))" by (rule continuous_on_compose2 [of UNIV exp]; intro continuous_intros contg subset_UNIV) auto show "\x. x \ S \ f x = (exp (g x / 2))\<^sup>2" by (metis exp_double feq nonzero_mult_div_cancel_left times_divide_eq_right zero_neq_numeral) qed qed lemma continuous_sqrt_on_simply_connected: fixes f :: "'a::real_normed_vector \ complex" assumes contf: "continuous_on S f" and S: "simply_connected S" "locally path_connected S" and f: "\z. z \ S \ f z \ 0" obtains g where "continuous_on S g" "\x. x \ S \ f x = (g x) ^ 2" proof - obtain g where contg: "continuous_on S g" and feq: "\x. x \ S \ f x = exp(g x)" using continuous_logarithm_on_simply_connected [OF assms] by blast show ?thesis proof show "continuous_on S (\z. exp (g z / 2))" by (rule continuous_on_compose2 [of UNIV exp]; intro continuous_intros contg subset_UNIV) auto show "\x. x \ S \ f x = (exp (g x / 2))\<^sup>2" by (metis exp_double feq nonzero_mult_div_cancel_left times_divide_eq_right zero_neq_numeral) qed qed subsection\Another simple case where sphere maps are nullhomotopic\ lemma inessential_spheremap_2_aux: fixes f :: "'a::euclidean_space \ complex" assumes 2: "2 < DIM('a)" and contf: "continuous_on (sphere a r) f" and fim: "f `(sphere a r) \ (sphere 0 1)" obtains c where "homotopic_with_canon (\z. True) (sphere a r) (sphere 0 1) f (\x. c)" proof - obtain g where contg: "continuous_on (sphere a r) g" and feq: "\x. x \ sphere a r \ f x = exp(g x)" proof (rule continuous_logarithm_on_simply_connected [OF contf]) show "simply_connected (sphere a r)" using 2 by (simp add: simply_connected_sphere_eq) show "locally path_connected (sphere a r)" by (simp add: locally_path_connected_sphere) show "\z. z \ sphere a r \ f z \ 0" using fim by force qed auto have "\g. continuous_on (sphere a r) g \ (\x\sphere a r. f x = exp (\ * complex_of_real (g x)))" proof (intro exI conjI) show "continuous_on (sphere a r) (Im \ g)" by (intro contg continuous_intros continuous_on_compose) show "\x\sphere a r. f x = exp (\ * complex_of_real ((Im \ g) x))" using exp_eq_polar feq fim norm_exp_eq_Re by auto qed with inessential_eq_continuous_logarithm_circle that show ?thesis by metis qed lemma inessential_spheremap_2: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes a2: "2 < DIM('a)" and b2: "DIM('b) = 2" and contf: "continuous_on (sphere a r) f" and fim: "f `(sphere a r) \ (sphere b s)" obtains c where "homotopic_with_canon (\z. True) (sphere a r) (sphere b s) f (\x. c)" proof (cases "s \ 0") case True then show ?thesis using contf contractible_sphere fim nullhomotopic_into_contractible that by blast next case False then have "sphere b s homeomorphic sphere (0::complex) 1" using assms by (simp add: homeomorphic_spheres_gen) then obtain h k where hk: "homeomorphism (sphere b s) (sphere (0::complex) 1) h k" by (auto simp: homeomorphic_def) then have conth: "continuous_on (sphere b s) h" and contk: "continuous_on (sphere 0 1) k" and him: "h ` sphere b s \ sphere 0 1" and kim: "k ` sphere 0 1 \ sphere b s" by (simp_all add: homeomorphism_def) obtain c where "homotopic_with_canon (\z. True) (sphere a r) (sphere 0 1) (h \ f) (\x. c)" proof (rule inessential_spheremap_2_aux [OF a2]) show "continuous_on (sphere a r) (h \ f)" by (meson continuous_on_compose [OF contf] conth continuous_on_subset fim) show "(h \ f) ` sphere a r \ sphere 0 1" using fim him by force qed auto then have "homotopic_with_canon (\f. True) (sphere a r) (sphere b s) (k \ (h \ f)) (k \ (\x. c))" by (rule homotopic_with_compose_continuous_left [OF _ contk kim]) then have "homotopic_with_canon (\z. True) (sphere a r) (sphere b s) f (\x. k c)" apply (rule homotopic_with_eq, auto) by (metis fim hk homeomorphism_def image_subset_iff mem_sphere) then show ?thesis by (metis that) qed subsection\Holomorphic logarithms and square roots\ -lemma contractible_imp_holomorphic_log: +lemma g_imp_holomorphic_log: assumes holf: "f holomorphic_on S" - and S: "contractible S" + and contg: "continuous_on S g" and feq: "\x. x \ S \ f x = exp (g x)" and fnz: "\z. z \ S \ f z \ 0" obtains g where "g holomorphic_on S" "\z. z \ S \ f z = exp(g z)" proof - have contf: "continuous_on S f" by (simp add: holf holomorphic_on_imp_continuous_on) - obtain g where contg: "continuous_on S g" and feq: "\x. x \ S \ f x = exp (g x)" - by (metis continuous_logarithm_on_contractible [OF contf S fnz]) have "g field_differentiable at z within S" if "f field_differentiable at z within S" "z \ S" for z proof - obtain f' where f': "((\y. (f y - f z) / (y - z)) \ f') (at z within S)" using \f field_differentiable at z within S\ by (auto simp: field_differentiable_def has_field_derivative_iff) then have ee: "((\x. (exp(g x) - exp(g z)) / (x - z)) \ f') (at z within S)" by (simp add: feq \z \ S\ Lim_transform_within [OF _ zero_less_one]) have "(((\y. if y = g z then exp (g z) else (exp y - exp (g z)) / (y - g z)) \ g) \ exp (g z)) (at z within S)" proof (rule tendsto_compose_at) show "(g \ g z) (at z within S)" using contg continuous_on \z \ S\ by blast show "(\y. if y = g z then exp (g z) else (exp y - exp (g z)) / (y - g z)) \g z\ exp (g z)" - apply (subst Lim_at_zero) - apply (simp add: DERIV_D cong: if_cong Lim_cong_within) - done + by (simp add: LIM_offset_zero_iff DERIV_D cong: if_cong Lim_cong_within) qed auto then have dd: "((\x. if g x = g z then exp(g z) else (exp(g x) - exp(g z)) / (g x - g z)) \ exp(g z)) (at z within S)" by (simp add: o_def) have "continuous (at z within S) g" using contg continuous_on_eq_continuous_within \z \ S\ by blast then have "(\\<^sub>F x in at z within S. dist (g x) (g z) < 2*pi)" by (simp add: continuous_within tendsto_iff) then have "\\<^sub>F x in at z within S. exp (g x) = exp (g z) \ g x \ g z \ x = z" - apply (rule eventually_mono) - apply (auto simp: exp_eq dist_norm norm_mult) - done + by (rule eventually_mono) (auto simp: exp_eq dist_norm norm_mult) then have "((\y. (g y - g z) / (y - z)) \ f' / exp (g z)) (at z within S)" by (auto intro!: Lim_transform_eventually [OF tendsto_divide [OF ee dd]]) then show ?thesis by (auto simp: field_differentiable_def has_field_derivative_iff) qed then have "g holomorphic_on S" using holf holomorphic_on_def by auto then show ?thesis using feq that by auto qed -(*Identical proofs*) +lemma contractible_imp_holomorphic_log: + assumes holf: "f holomorphic_on S" + and S: "contractible S" + and fnz: "\z. z \ S \ f z \ 0" + obtains g where "g holomorphic_on S" "\z. z \ S \ f z = exp(g z)" +proof - + have contf: "continuous_on S f" + by (simp add: holf holomorphic_on_imp_continuous_on) + obtain g where contg: "continuous_on S g" and feq: "\x. x \ S \ f x = exp (g x)" + by (metis continuous_logarithm_on_contractible [OF contf S fnz]) + then show thesis + using fnz g_imp_holomorphic_log holf that by blast +qed + lemma simply_connected_imp_holomorphic_log: assumes holf: "f holomorphic_on S" and S: "simply_connected S" "locally path_connected S" and fnz: "\z. z \ S \ f z \ 0" obtains g where "g holomorphic_on S" "\z. z \ S \ f z = exp(g z)" proof - have contf: "continuous_on S f" by (simp add: holf holomorphic_on_imp_continuous_on) obtain g where contg: "continuous_on S g" and feq: "\x. x \ S \ f x = exp (g x)" by (metis continuous_logarithm_on_simply_connected [OF contf S fnz]) - have "g field_differentiable at z within S" if "f field_differentiable at z within S" "z \ S" for z - proof - - obtain f' where f': "((\y. (f y - f z) / (y - z)) \ f') (at z within S)" - using \f field_differentiable at z within S\ by (auto simp: field_differentiable_def has_field_derivative_iff) - then have ee: "((\x. (exp(g x) - exp(g z)) / (x - z)) \ f') (at z within S)" - by (simp add: feq \z \ S\ Lim_transform_within [OF _ zero_less_one]) - have "(((\y. if y = g z then exp (g z) else (exp y - exp (g z)) / (y - g z)) \ g) \ exp (g z)) - (at z within S)" - proof (rule tendsto_compose_at) - show "(g \ g z) (at z within S)" - using contg continuous_on \z \ S\ by blast - show "(\y. if y = g z then exp (g z) else (exp y - exp (g z)) / (y - g z)) \g z\ exp (g z)" - apply (subst Lim_at_zero) - apply (simp add: DERIV_D cong: if_cong Lim_cong_within) - done - qed auto - then have dd: "((\x. if g x = g z then exp(g z) else (exp(g x) - exp(g z)) / (g x - g z)) \ exp(g z)) (at z within S)" - by (simp add: o_def) - have "continuous (at z within S) g" - using contg continuous_on_eq_continuous_within \z \ S\ by blast - then have "(\\<^sub>F x in at z within S. dist (g x) (g z) < 2*pi)" - by (simp add: continuous_within tendsto_iff) - then have "\\<^sub>F x in at z within S. exp (g x) = exp (g z) \ g x \ g z \ x = z" - apply (rule eventually_mono) - apply (auto simp: exp_eq dist_norm norm_mult) - done - then have "((\y. (g y - g z) / (y - z)) \ f' / exp (g z)) (at z within S)" - by (auto intro!: Lim_transform_eventually [OF tendsto_divide [OF ee dd]]) - then show ?thesis - by (auto simp: field_differentiable_def has_field_derivative_iff) - qed - then have "g holomorphic_on S" - using holf holomorphic_on_def by auto - then show ?thesis - using feq that by auto + then show thesis + using fnz g_imp_holomorphic_log holf that by blast qed - lemma contractible_imp_holomorphic_sqrt: assumes holf: "f holomorphic_on S" and S: "contractible S" and fnz: "\z. z \ S \ f z \ 0" obtains g where "g holomorphic_on S" "\z. z \ S \ f z = g z ^ 2" proof - obtain g where holg: "g holomorphic_on S" and feq: "\z. z \ S \ f z = exp(g z)" using contractible_imp_holomorphic_log [OF assms] by blast show ?thesis proof show "exp \ (\z. z / 2) \ g holomorphic_on S" by (intro holomorphic_on_compose holg holomorphic_intros) auto show "\z. z \ S \ f z = ((exp \ (\z. z / 2) \ g) z)\<^sup>2" - apply (auto simp: feq) - by (metis eq_divide_eq_numeral1(1) exp_double mult.commute zero_neq_numeral) + by (simp add: feq flip: exp_double) qed qed lemma simply_connected_imp_holomorphic_sqrt: assumes holf: "f holomorphic_on S" and S: "simply_connected S" "locally path_connected S" and fnz: "\z. z \ S \ f z \ 0" obtains g where "g holomorphic_on S" "\z. z \ S \ f z = g z ^ 2" proof - obtain g where holg: "g holomorphic_on S" and feq: "\z. z \ S \ f z = exp(g z)" using simply_connected_imp_holomorphic_log [OF assms] by blast show ?thesis proof show "exp \ (\z. z / 2) \ g holomorphic_on S" by (intro holomorphic_on_compose holg holomorphic_intros) auto show "\z. z \ S \ f z = ((exp \ (\z. z / 2) \ g) z)\<^sup>2" - apply (auto simp: feq) - by (metis eq_divide_eq_numeral1(1) exp_double mult.commute zero_neq_numeral) + by (simp add: feq flip: exp_double) qed qed text\ Related theorems about holomorphic inverse cosines.\ lemma contractible_imp_holomorphic_arccos: assumes holf: "f holomorphic_on S" and S: "contractible S" and non1: "\z. z \ S \ f z \ 1 \ f z \ -1" obtains g where "g holomorphic_on S" "\z. z \ S \ f z = cos(g z)" proof - have hol1f: "(\z. 1 - f z ^ 2) holomorphic_on S" by (intro holomorphic_intros holf) obtain g where holg: "g holomorphic_on S" and eq: "\z. z \ S \ 1 - (f z)\<^sup>2 = (g z)\<^sup>2" using contractible_imp_holomorphic_sqrt [OF hol1f S] by (metis eq_iff_diff_eq_0 non1 power2_eq_1_iff) have holfg: "(\z. f z + \*g z) holomorphic_on S" by (intro holf holg holomorphic_intros) have "\z. z \ S \ f z + \*g z \ 0" by (metis Arccos_body_lemma eq add.commute add.inverse_unique complex_i_mult_minus power2_csqrt power2_eq_iff) then obtain h where holh: "h holomorphic_on S" and fgeq: "\z. z \ S \ f z + \*g z = exp (h z)" using contractible_imp_holomorphic_log [OF holfg S] by metis show ?thesis proof show "(\z. -\*h z) holomorphic_on S" by (intro holh holomorphic_intros) show "f z = cos (- \*h z)" if "z \ S" for z proof - have "(f z + \*g z)*(f z - \*g z) = 1" using that eq by (auto simp: algebra_simps power2_eq_square) then have "f z - \*g z = inverse (f z + \*g z)" using inverse_unique by force also have "... = exp (- h z)" by (simp add: exp_minus fgeq that) finally have "f z = exp (- h z) + \*g z" by (simp add: diff_eq_eq) then show ?thesis apply (simp add: cos_exp_eq) by (metis fgeq add.assoc mult_2_right that) qed qed qed lemma contractible_imp_holomorphic_arccos_bounded: assumes holf: "f holomorphic_on S" and S: "contractible S" and "a \ S" and non1: "\z. z \ S \ f z \ 1 \ f z \ -1" obtains g where "g holomorphic_on S" "norm(g a) \ pi + norm(f a)" "\z. z \ S \ f z = cos(g z)" proof - obtain g where holg: "g holomorphic_on S" and feq: "\z. z \ S \ f z = cos (g z)" using contractible_imp_holomorphic_arccos [OF holf S non1] by blast obtain b where "cos b = f a" "norm b \ pi + norm (f a)" using cos_Arccos norm_Arccos_bounded by blast then have "cos b = cos (g a)" by (simp add: \a \ S\ feq) then consider n where "n \ \" "b = g a + of_real(2*n*pi)" | n where "n \ \" "b = -g a + of_real(2*n*pi)" by (auto simp: complex_cos_eq) then show ?thesis proof cases case 1 show ?thesis proof show "(\z. g z + of_real(2*n*pi)) holomorphic_on S" by (intro holomorphic_intros holg) show "cmod (g a + of_real(2*n*pi)) \ pi + cmod (f a)" using "1" \cmod b \ pi + cmod (f a)\ by blast show "\z. z \ S \ f z = cos (g z + complex_of_real (2*n*pi))" by (metis \n \ \\ complex_cos_eq feq) qed next case 2 show ?thesis proof show "(\z. -g z + of_real(2*n*pi)) holomorphic_on S" by (intro holomorphic_intros holg) show "cmod (-g a + of_real(2*n*pi)) \ pi + cmod (f a)" using "2" \cmod b \ pi + cmod (f a)\ by blast show "\z. z \ S \ f z = cos (-g z + complex_of_real (2*n*pi))" by (metis \n \ \\ complex_cos_eq feq) qed qed qed subsection\The "Borsukian" property of sets\ text\This doesn't have a standard name. Kuratowski uses ``contractible with respect to \[S\<^sup>1]\'' while Whyburn uses ``property b''. It's closely related to unicoherence.\ definition\<^marker>\tag important\ Borsukian where "Borsukian S \ \f. continuous_on S f \ f ` S \ (- {0::complex}) \ (\a. homotopic_with_canon (\h. True) S (- {0}) f (\x. a))" lemma Borsukian_retraction_gen: assumes "Borsukian S" "continuous_on S h" "h ` S = T" "continuous_on T k" "k ` T \ S" "\y. y \ T \ h(k y) = y" shows "Borsukian T" proof - interpret R: Retracts S h T k using assms by (simp add: Retracts.intro) show ?thesis using assms - apply (simp add: Borsukian_def, clarify) + apply (clarsimp simp add: Borsukian_def) apply (rule R.cohomotopically_trivial_retraction_null_gen [OF TrueI TrueI refl, of "-{0}"], auto) done qed lemma retract_of_Borsukian: "\Borsukian T; S retract_of T\ \ Borsukian S" apply (auto simp: retract_of_def retraction_def) apply (erule (1) Borsukian_retraction_gen) apply (meson retraction retraction_def) apply (auto) done lemma homeomorphic_Borsukian: "\Borsukian S; S homeomorphic T\ \ Borsukian T" using Borsukian_retraction_gen order_refl by (fastforce simp add: homeomorphism_def homeomorphic_def) lemma homeomorphic_Borsukian_eq: "S homeomorphic T \ Borsukian S \ Borsukian T" by (meson homeomorphic_Borsukian homeomorphic_sym) lemma Borsukian_translation: fixes S :: "'a::real_normed_vector set" shows "Borsukian (image (\x. a + x) S) \ Borsukian S" - apply (rule homeomorphic_Borsukian_eq) - using homeomorphic_translation homeomorphic_sym by blast + using homeomorphic_Borsukian_eq homeomorphic_translation by blast lemma Borsukian_injective_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "Borsukian(f ` S) \ Borsukian S" - apply (rule homeomorphic_Borsukian_eq) - using assms homeomorphic_sym linear_homeomorphic_image by blast + using assms homeomorphic_Borsukian_eq linear_homeomorphic_image by blast lemma homotopy_eqv_Borsukianness: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" assumes "S homotopy_eqv T" shows "(Borsukian S \ Borsukian T)" by (meson Borsukian_def assms homotopy_eqv_cohomotopic_triviality_null) lemma Borsukian_alt: fixes S :: "'a::real_normed_vector set" shows "Borsukian S \ (\f g. continuous_on S f \ f ` S \ -{0} \ continuous_on S g \ g ` S \ -{0} \ homotopic_with_canon (\h. True) S (- {0::complex}) f g)" unfolding Borsukian_def homotopic_triviality by (simp add: path_connected_punctured_universe) lemma Borsukian_continuous_logarithm: fixes S :: "'a::real_normed_vector set" shows "Borsukian S \ (\f. continuous_on S f \ f ` S \ (- {0::complex}) \ (\g. continuous_on S g \ (\x \ S. f x = exp(g x))))" by (simp add: Borsukian_def inessential_eq_continuous_logarithm) lemma Borsukian_continuous_logarithm_circle: fixes S :: "'a::real_normed_vector set" shows "Borsukian S \ (\f. continuous_on S f \ f ` S \ sphere (0::complex) 1 \ (\g. continuous_on S g \ (\x \ S. f x = exp(g x))))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (force simp: Borsukian_continuous_logarithm) next assume RHS [rule_format]: ?rhs show ?lhs proof (clarsimp simp: Borsukian_continuous_logarithm) fix f :: "'a \ complex" assume contf: "continuous_on S f" and 0: "0 \ f ` S" then have "continuous_on S (\x. f x / complex_of_real (cmod (f x)))" by (intro continuous_intros) auto moreover have "(\x. f x / complex_of_real (cmod (f x))) ` S \ sphere 0 1" using 0 by (auto simp: norm_divide) ultimately obtain g where contg: "continuous_on S g" and fg: "\x \ S. f x / complex_of_real (cmod (f x)) = exp(g x)" using RHS [of "\x. f x / of_real(norm(f x))"] by auto show "\g. continuous_on S g \ (\x\S. f x = exp (g x))" proof (intro exI ballI conjI) show "continuous_on S (\x. (Ln \ of_real \ norm \ f)x + g x)" by (intro continuous_intros contf contg conjI) (use "0" in auto) show "f x = exp ((Ln \ complex_of_real \ cmod \ f) x + g x)" if "x \ S" for x using 0 that - apply (clarsimp simp: exp_add) - apply (subst exp_Ln, force) - by (metis eq_divide_eq exp_not_eq_zero fg mult.commute) + apply (simp add: exp_add) + by (metis div_by_0 exp_Ln exp_not_eq_zero fg mult.commute nonzero_eq_divide_eq) qed qed qed lemma Borsukian_continuous_logarithm_circle_real: fixes S :: "'a::real_normed_vector set" shows "Borsukian S \ (\f. continuous_on S f \ f ` S \ sphere (0::complex) 1 \ (\g. continuous_on S (complex_of_real \ g) \ (\x \ S. f x = exp(\ * of_real(g x)))))" (is "?lhs = ?rhs") proof assume LHS: ?lhs show ?rhs proof (clarify) fix f :: "'a \ complex" assume "continuous_on S f" and f01: "f ` S \ sphere 0 1" then obtain g where contg: "continuous_on S g" and "\x. x \ S \ f x = exp(g x)" using LHS by (auto simp: Borsukian_continuous_logarithm_circle) then have "\x\S. f x = exp (\ * complex_of_real ((Im \ g) x))" - using f01 apply (simp add: image_iff subset_iff) - by (metis cis_conv_exp exp_eq_polar mult.left_neutral norm_exp_eq_Re of_real_1) + using f01 exp_eq_polar norm_exp_eq_Re by auto then show "\g. continuous_on S (complex_of_real \ g) \ (\x\S. f x = exp (\ * complex_of_real (g x)))" by (rule_tac x="Im \ g" in exI) (force intro: continuous_intros contg) qed next assume RHS [rule_format]: ?rhs show ?lhs proof (clarsimp simp: Borsukian_continuous_logarithm_circle) fix f :: "'a \ complex" assume "continuous_on S f" and f01: "f ` S \ sphere 0 1" then obtain g where contg: "continuous_on S (complex_of_real \ g)" and "\x. x \ S \ f x = exp(\ * of_real(g x))" by (metis RHS) then show "\g. continuous_on S g \ (\x\S. f x = exp (g x))" by (rule_tac x="\x. \* of_real(g x)" in exI) (auto simp: continuous_intros contg) qed qed lemma Borsukian_circle: fixes S :: "'a::real_normed_vector set" shows "Borsukian S \ (\f. continuous_on S f \ f ` S \ sphere (0::complex) 1 \ (\a. homotopic_with_canon (\h. True) S (sphere (0::complex) 1) f (\x. a)))" by (simp add: inessential_eq_continuous_logarithm_circle Borsukian_continuous_logarithm_circle_real) lemma contractible_imp_Borsukian: "contractible S \ Borsukian S" by (meson Borsukian_def nullhomotopic_from_contractible) lemma simply_connected_imp_Borsukian: fixes S :: "'a::real_normed_vector set" shows "\simply_connected S; locally path_connected S\ \ Borsukian S" - apply (simp add: Borsukian_continuous_logarithm) - by (metis (no_types, lifting) continuous_logarithm_on_simply_connected image_iff) + by (metis (no_types, lifting) Borsukian_continuous_logarithm continuous_logarithm_on_simply_connected image_eqI subset_Compl_singleton) lemma starlike_imp_Borsukian: fixes S :: "'a::real_normed_vector set" shows "starlike S \ Borsukian S" by (simp add: contractible_imp_Borsukian starlike_imp_contractible) lemma Borsukian_empty: "Borsukian {}" by (auto simp: contractible_imp_Borsukian) lemma Borsukian_UNIV: "Borsukian (UNIV :: 'a::real_normed_vector set)" by (auto simp: contractible_imp_Borsukian) lemma convex_imp_Borsukian: fixes S :: "'a::real_normed_vector set" shows "convex S \ Borsukian S" by (meson Borsukian_def convex_imp_contractible nullhomotopic_from_contractible) proposition Borsukian_sphere: fixes a :: "'a::euclidean_space" shows "3 \ DIM('a) \ Borsukian (sphere a r)" - apply (rule simply_connected_imp_Borsukian) - using simply_connected_sphere apply blast - using ENR_imp_locally_path_connected ENR_sphere by blast - -proposition Borsukian_open_Un: + using ENR_sphere + by (blast intro: simply_connected_imp_Borsukian ENR_imp_locally_path_connected simply_connected_sphere) + +lemma Borsukian_Un_lemma: fixes S :: "'a::real_normed_vector set" - assumes opeS: "openin (top_of_set (S \ T)) S" - and opeT: "openin (top_of_set (S \ T)) T" - and BS: "Borsukian S" and BT: "Borsukian T" and ST: "connected(S \ T)" - shows "Borsukian(S \ T)" + assumes BS: "Borsukian S" and BT: "Borsukian T" and ST: "connected(S \ T)" + and *: "\f g::'a \ complex. + \continuous_on S f; continuous_on T g; \x. x \ S \ x \ T \ f x = g x\ + \ continuous_on (S \ T) (\x. if x \ S then f x else g x)" + shows "Borsukian(S \ T)" proof (clarsimp simp add: Borsukian_continuous_logarithm) fix f :: "'a \ complex" assume contf: "continuous_on (S \ T) f" and 0: "0 \ f ` (S \ T)" then have contfS: "continuous_on S f" and contfT: "continuous_on T f" using continuous_on_subset by auto have "\continuous_on S f; f ` S \ -{0}\ \ \g. continuous_on S g \ (\x \ S. f x = exp(g x))" using BS by (auto simp: Borsukian_continuous_logarithm) then obtain g where contg: "continuous_on S g" and fg: "\x. x \ S \ f x = exp(g x)" using "0" contfS by blast have "\continuous_on T f; f ` T \ -{0}\ \ \g. continuous_on T g \ (\x \ T. f x = exp(g x))" using BT by (auto simp: Borsukian_continuous_logarithm) then obtain h where conth: "continuous_on T h" and fh: "\x. x \ T \ f x = exp(h x)" using "0" contfT by blast show "\g. continuous_on (S \ T) g \ (\x\S \ T. f x = exp (g x))" proof (cases "S \ T = {}") case True show ?thesis proof (intro exI conjI) show "continuous_on (S \ T) (\x. if x \ S then g x else h x)" - apply (rule continuous_on_cases_local_open [OF opeS opeT contg conth]) - using True by blast + using True * [OF contg conth] + by (meson disjoint_iff) show "\x\S \ T. f x = exp (if x \ S then g x else h x)" using fg fh by auto qed next case False have "(\x. g x - h x) constant_on S \ T" proof (rule continuous_discrete_range_constant [OF ST]) show "continuous_on (S \ T) (\x. g x - h x)" - apply (intro continuous_intros) - apply (meson contg continuous_on_subset inf_le1) - by (meson conth continuous_on_subset inf_sup_ord(2)) + proof (intro continuous_intros) + show "continuous_on (S \ T) g" + by (meson contg continuous_on_subset inf_le1) + show "continuous_on (S \ T) h" + by (meson conth continuous_on_subset inf_sup_ord(2)) + qed show "\e>0. \y. y \ S \ T \ g y - h y \ g x - h x \ e \ cmod (g y - h y - (g x - h x))" if "x \ S \ T" for x proof - have "g y - g x = h y - h x" if "y \ S" "y \ T" "cmod (g y - g x - (h y - h x)) < 2 * pi" for y proof (rule exp_complex_eqI) have "\Im (g y - g x) - Im (h y - h x)\ \ cmod (g y - g x - (h y - h x))" by (metis abs_Im_le_cmod minus_complex.simps(2)) then show "\Im (g y - g x) - Im (h y - h x)\ < 2 * pi" using that by linarith have "exp (g x) = exp (h x)" "exp (g y) = exp (h y)" using fg fh that \x \ S \ T\ by fastforce+ then show "exp (g y - g x) = exp (h y - h x)" by (simp add: exp_diff) qed then show ?thesis by (rule_tac x="2*pi" in exI) (fastforce simp add: algebra_simps) qed qed then obtain a where a: "\x. x \ S \ T \ g x - h x = a" by (auto simp: constant_on_def) with False have "exp a = 1" by (metis IntI disjoint_iff_not_equal divide_self_if exp_diff exp_not_eq_zero fg fh) with a show ?thesis apply (rule_tac x="\x. if x \ S then g x else a + h x" in exI) - apply (intro continuous_on_cases_local_open opeS opeT contg conth continuous_intros conjI) + apply (intro * contg conth continuous_intros conjI) apply (auto simp: algebra_simps fg fh exp_add) done qed qed -text\The proof is a duplicate of that of \Borsukian_open_Un\.\ +proposition Borsukian_open_Un: + fixes S :: "'a::real_normed_vector set" + assumes opeS: "openin (top_of_set (S \ T)) S" + and opeT: "openin (top_of_set (S \ T)) T" + and BS: "Borsukian S" and BT: "Borsukian T" and ST: "connected(S \ T)" + shows "Borsukian(S \ T)" + by (force intro: Borsukian_Un_lemma [OF BS BT ST] continuous_on_cases_local_open [OF opeS opeT]) + lemma Borsukian_closed_Un: fixes S :: "'a::real_normed_vector set" assumes cloS: "closedin (top_of_set (S \ T)) S" and cloT: "closedin (top_of_set (S \ T)) T" and BS: "Borsukian S" and BT: "Borsukian T" and ST: "connected(S \ T)" shows "Borsukian(S \ T)" -proof (clarsimp simp add: Borsukian_continuous_logarithm) - fix f :: "'a \ complex" - assume contf: "continuous_on (S \ T) f" and 0: "0 \ f ` (S \ T)" - then have contfS: "continuous_on S f" and contfT: "continuous_on T f" - using continuous_on_subset by auto - have "\continuous_on S f; f ` S \ -{0}\ \ \g. continuous_on S g \ (\x \ S. f x = exp(g x))" - using BS by (auto simp: Borsukian_continuous_logarithm) - then obtain g where contg: "continuous_on S g" and fg: "\x. x \ S \ f x = exp(g x)" - using "0" contfS by blast - have "\continuous_on T f; f ` T \ -{0}\ \ \g. continuous_on T g \ (\x \ T. f x = exp(g x))" - using BT by (auto simp: Borsukian_continuous_logarithm) - then obtain h where conth: "continuous_on T h" and fh: "\x. x \ T \ f x = exp(h x)" - using "0" contfT by blast - show "\g. continuous_on (S \ T) g \ (\x\S \ T. f x = exp (g x))" - proof (cases "S \ T = {}") - case True - show ?thesis - proof (intro exI conjI) - show "continuous_on (S \ T) (\x. if x \ S then g x else h x)" - apply (rule continuous_on_cases_local [OF cloS cloT contg conth]) - using True by blast - show "\x\S \ T. f x = exp (if x \ S then g x else h x)" - using fg fh by auto - qed - next - case False - have "(\x. g x - h x) constant_on S \ T" - proof (rule continuous_discrete_range_constant [OF ST]) - show "continuous_on (S \ T) (\x. g x - h x)" - apply (intro continuous_intros) - apply (meson contg continuous_on_subset inf_le1) - by (meson conth continuous_on_subset inf_sup_ord(2)) - show "\e>0. \y. y \ S \ T \ g y - h y \ g x - h x \ e \ cmod (g y - h y - (g x - h x))" - if "x \ S \ T" for x - proof - - have "g y - g x = h y - h x" - if "y \ S" "y \ T" "cmod (g y - g x - (h y - h x)) < 2 * pi" for y - proof (rule exp_complex_eqI) - have "\Im (g y - g x) - Im (h y - h x)\ \ cmod (g y - g x - (h y - h x))" - by (metis abs_Im_le_cmod minus_complex.simps(2)) - then show "\Im (g y - g x) - Im (h y - h x)\ < 2 * pi" - using that by linarith - have "exp (g x) = exp (h x)" "exp (g y) = exp (h y)" - using fg fh that \x \ S \ T\ by fastforce+ - then show "exp (g y - g x) = exp (h y - h x)" - by (simp add: exp_diff) - qed - then show ?thesis - by (rule_tac x="2*pi" in exI) (fastforce simp add: algebra_simps) - qed - qed - then obtain a where a: "\x. x \ S \ T \ g x - h x = a" - by (auto simp: constant_on_def) - with False have "exp a = 1" - by (metis IntI disjoint_iff_not_equal divide_self_if exp_diff exp_not_eq_zero fg fh) - with a show ?thesis - apply (rule_tac x="\x. if x \ S then g x else a + h x" in exI) - apply (intro continuous_on_cases_local cloS cloT contg conth continuous_intros conjI) - apply (auto simp: algebra_simps fg fh exp_add) - done - qed -qed + by (force intro: Borsukian_Un_lemma [OF BS BT ST] continuous_on_cases_local [OF cloS cloT]) lemma Borsukian_separation_compact: fixes S :: "complex set" assumes "compact S" shows "Borsukian S \ connected(- S)" by (simp add: Borsuk_separation_theorem Borsukian_circle assms) lemma Borsukian_monotone_image_compact: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "Borsukian S" and contf: "continuous_on S f" and fim: "f ` S = T" and "compact S" and conn: "\y. y \ T \ connected {x. x \ S \ f x = y}" shows "Borsukian T" proof (clarsimp simp add: Borsukian_continuous_logarithm) fix g :: "'b \ complex" assume contg: "continuous_on T g" and 0: "0 \ g ` T" have "continuous_on S (g \ f)" using contf contg continuous_on_compose fim by blast moreover have "(g \ f) ` S \ -{0}" using fim 0 by auto ultimately obtain h where conth: "continuous_on S h" and gfh: "\x. x \ S \ (g \ f) x = exp(h x)" using \Borsukian S\ by (auto simp: Borsukian_continuous_logarithm) have "\y. \x. y \ T \ x \ S \ f x = y" using fim by auto then obtain f' where f': "\y. y \ T \ f' y \ S \ f (f' y) = y" by metis have *: "(\x. h x - h(f' y)) constant_on {x. x \ S \ f x = y}" if "y \ T" for y proof (rule continuous_discrete_range_constant [OF conn [OF that], of "\x. h x - h (f' y)"], simp_all add: algebra_simps) show "continuous_on {x \ S. f x = y} (\x. h x - h (f' y))" by (intro continuous_intros continuous_on_subset [OF conth]) auto show "\e>0. \u. u \ S \ f u = y \ h u \ h x \ e \ cmod (h u - h x)" if x: "x \ S \ f x = y" for x proof - have "h u = h x" if "u \ S" "f u = y" "cmod (h u - h x) < 2 * pi" for u proof (rule exp_complex_eqI) have "\Im (h u) - Im (h x)\ \ cmod (h u - h x)" by (metis abs_Im_le_cmod minus_complex.simps(2)) then show "\Im (h u) - Im (h x)\ < 2 * pi" using that by linarith show "exp (h u) = exp (h x)" by (simp add: gfh [symmetric] x that) qed then show ?thesis by (rule_tac x="2*pi" in exI) (fastforce simp add: algebra_simps) qed qed - have "h x = h (f' (f x))" if "x \ S" for x - using * [of "f x"] fim that unfolding constant_on_def by clarsimp (metis f' imageI right_minus_eq) - moreover have "\x. x \ T \ \u. u \ S \ x = f u \ h (f' x) = h u" - using f' by fastforce - ultimately - have eq: "((\x. (x, (h \ f') x)) ` T) = - {p. \x. x \ S \ (x, p) \ (S \ UNIV) \ ((\z. snd z - ((f \ fst) z, (h \ fst) z)) -` {0})}" - using fim by (auto simp: image_iff) show "\h. continuous_on T h \ (\x\T. g x = exp (h x))" proof (intro exI conjI) show "continuous_on T (h \ f')" proof (rule continuous_from_closed_graph [of "h ` S"]) show "compact (h ` S)" by (simp add: \compact S\ compact_continuous_image conth) show "(h \ f') ` T \ h ` S" by (auto simp: f') - show "closed ((\x. (x, (h \ f') x)) ` T)" - apply (subst eq) + have "h x = h (f' (f x))" if "x \ S" for x + using * [of "f x"] fim that unfolding constant_on_def by clarsimp (metis f' imageI right_minus_eq) + moreover have "\x. x \ T \ \u. u \ S \ x = f u \ h (f' x) = h u" + using f' by fastforce + ultimately + have eq: "((\x. (x, (h \ f') x)) ` T) = + {p. \x. x \ S \ (x, p) \ (S \ UNIV) \ ((\z. snd z - ((f \ fst) z, (h \ fst) z)) -` {0})}" + using fim by (auto simp: image_iff) + moreover have "closed \" apply (intro closed_compact_projection [OF \compact S\] continuous_closed_preimage continuous_intros continuous_on_subset [OF contf] continuous_on_subset [OF conth]) - apply (auto simp: \compact S\ closed_Times compact_imp_closed) - done + by (auto simp: \compact S\ closed_Times compact_imp_closed) + ultimately show "closed ((\x. (x, (h \ f') x)) ` T)" + by simp qed qed (use f' gfh in fastforce) qed lemma Borsukian_open_map_image_compact: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "Borsukian S" and contf: "continuous_on S f" and fim: "f ` S = T" and "compact S" and ope: "\U. openin (top_of_set S) U \ openin (top_of_set T) (f ` U)" shows "Borsukian T" proof (clarsimp simp add: Borsukian_continuous_logarithm_circle_real) fix g :: "'b \ complex" assume contg: "continuous_on T g" and gim: "g ` T \ sphere 0 1" have "continuous_on S (g \ f)" using contf contg continuous_on_compose fim by blast moreover have "(g \ f) ` S \ sphere 0 1" using fim gim by auto ultimately obtain h where cont_cxh: "continuous_on S (complex_of_real \ h)" and gfh: "\x. x \ S \ (g \ f) x = exp(\ * of_real(h x))" using \Borsukian S\ Borsukian_continuous_logarithm_circle_real by metis then have conth: "continuous_on S h" by simp have "\x. x \ S \ f x = y \ (\x' \ S. f x' = y \ h x \ h x')" if "y \ T" for y proof - have 1: "compact (h ` {x \ S. f x = y})" proof (rule compact_continuous_image) show "continuous_on {x \ S. f x = y} h" by (rule continuous_on_subset [OF conth]) auto have "compact (S \ f -` {y})" by (rule proper_map_from_compact [OF contf _ \compact S\, of T]) (simp_all add: fim that) then show "compact {x \ S. f x = y}" by (auto simp: vimage_def Int_def) qed have 2: "h ` {x \ S. f x = y} \ {}" using fim that by auto have "\s \ h ` {x \ S. f x = y}. \t \ h ` {x \ S. f x = y}. s \ t" using compact_attains_inf [OF 1 2] by blast then show ?thesis by auto qed then obtain k where kTS: "\y. y \ T \ k y \ S" and fk: "\y. y \ T \ f (k y) = y " and hle: "\x' y. \y \ T; x' \ S; f x' = y\ \ h (k y) \ h x'" by metis have "continuous_on T (h \ k)" proof (clarsimp simp add: continuous_on_iff) fix y and e::real assume "y \ T" "0 < e" moreover have "uniformly_continuous_on S (complex_of_real \ h)" using \compact S\ cont_cxh compact_uniformly_continuous by blast ultimately obtain d where "0 < d" and d: "\x x'. \x\S; x'\S; dist x' x < d\ \ dist (h x') (h x) < e" by (force simp: uniformly_continuous_on_def) obtain \ where "0 < \" and \: "\x'. \x' \ T; dist y x' < \\ \ (\v \ {z \ S. f z = y}. \v'. v' \ {z \ S. f z = x'} \ dist v v' < d) \ (\v' \ {z \ S. f z = x'}. \v. v \ {z \ S. f z = y} \ dist v' v < d)" proof (rule upper_lower_hemicontinuous_explicit [of T "\y. {z \ S. f z = y}" S]) show "\U. openin (top_of_set S) U \ openin (top_of_set T) {x \ T. {z \ S. f z = x} \ U}" using closed_map_iff_upper_hemicontinuous_preimage [OF fim [THEN equalityD1]] by (simp add: Abstract_Topology_2.continuous_imp_closed_map \compact S\ contf fim) show "\U. closedin (top_of_set S) U \ closedin (top_of_set T) {x \ T. {z \ S. f z = x} \ U}" using ope open_map_iff_lower_hemicontinuous_preimage [OF fim [THEN equalityD1]] by meson show "bounded {z \ S. f z = y}" by (metis (no_types, lifting) compact_imp_bounded [OF \compact S\] bounded_subset mem_Collect_eq subsetI) qed (use \y \ T\ \0 < d\ fk kTS in \force+\) have "dist (h (k y')) (h (k y)) < e" if "y' \ T" "dist y y' < \" for y' proof - have k1: "k y \ S" "f (k y) = y" and k2: "k y' \ S" "f (k y') = y'" by (auto simp: \y \ T\ \y' \ T\ kTS fk) have 1: "\v. \v \ S; f v = y\ \ \v'. v' \ {z \ S. f z = y'} \ dist v v' < d" and 2: "\v'. \v' \ S; f v' = y'\ \ \v. v \ {z \ S. f z = y} \ dist v' v < d" using \ [OF that] by auto then obtain w' w where "w' \ S" "f w' = y'" "dist (k y) w' < d" and "w \ S" "f w = y" "dist (k y') w < d" using 1 [OF k1] 2 [OF k2] by auto then show ?thesis using d [of w "k y'"] d [of w' "k y"] k1 k2 \y' \ T\ \y \ T\ hle by (fastforce simp: dist_norm abs_diff_less_iff algebra_simps) qed then show "\d>0. \x'\T. dist x' y < d \ dist (h (k x')) (h (k y)) < e" using \0 < \\ by (auto simp: dist_commute) qed then show "\h. continuous_on T h \ (\x\T. g x = exp (\ * complex_of_real (h x)))" using fk gfh kTS by force qed text\If two points are separated by a closed set, there's a minimal one.\ proposition closed_irreducible_separator: fixes a :: "'a::real_normed_vector" assumes "closed S" and ab: "\ connected_component (- S) a b" obtains T where "T \ S" "closed T" "T \ {}" "\ connected_component (- T) a b" "\U. U \ T \ connected_component (- U) a b" proof (cases "a \ S \ b \ S") case True then show ?thesis proof assume *: "a \ S" show ?thesis proof show "{a} \ S" using * by blast show "\ connected_component (- {a}) a b" using connected_component_in by auto show "\U. U \ {a} \ connected_component (- U) a b" by (metis connected_component_UNIV UNIV_I compl_bot_eq connected_component_eq_eq less_le_not_le subset_singletonD) qed auto next assume *: "b \ S" show ?thesis proof show "{b} \ S" using * by blast show "\ connected_component (- {b}) a b" using connected_component_in by auto show "\U. U \ {b} \ connected_component (- U) a b" by (metis connected_component_UNIV UNIV_I compl_bot_eq connected_component_eq_eq less_le_not_le subset_singletonD) qed auto qed next case False define A where "A \ connected_component_set (- S) a" define B where "B \ connected_component_set (- (closure A)) b" have "a \ A" using False A_def by auto have "b \ B" unfolding A_def B_def closure_Un_frontier using ab False \closed S\ frontier_complement frontier_of_connected_component_subset frontier_subset_closed by force have "frontier B \ frontier (connected_component_set (- closure A) b)" using B_def by blast also have frsub: "... \ frontier A" proof - have "\A. closure (- closure (- A)) \ closure A" by (metis (no_types) closure_mono closure_subset compl_le_compl_iff double_compl) then show ?thesis by (metis (no_types) closure_closure double_compl frontier_closures frontier_of_connected_component_subset le_inf_iff subset_trans) qed finally have frBA: "frontier B \ frontier A" . show ?thesis proof show "frontier B \ S" proof - have "frontier S \ S" by (simp add: \closed S\ frontier_subset_closed) then show ?thesis using frsub frontier_complement frontier_of_connected_component_subset unfolding A_def B_def by blast qed show "closed (frontier B)" by simp show "\ connected_component (- frontier B) a b" unfolding connected_component_def proof clarify fix T assume "connected T" and TB: "T \ - frontier B" and "a \ T" and "b \ T" have "a \ B" by (metis A_def B_def ComplD \a \ A\ assms(1) closed_open connected_component_subset in_closure_connected_component subsetD) have "T \ B \ {}" using \b \ B\ \b \ T\ by blast moreover have "T - B \ {}" using \a \ B\ \a \ T\ by blast ultimately show "False" using connected_Int_frontier [of T B] TB \connected T\ by blast qed moreover have "connected_component (- frontier B) a b" if "frontier B = {}" - apply (simp add: that) - using connected_component_eq_UNIV by blast + using connected_component_eq_UNIV that by auto ultimately show "frontier B \ {}" by blast show "connected_component (- U) a b" if "U \ frontier B" for U proof - obtain p where Usub: "U \ frontier B" and p: "p \ frontier B" "p \ U" using \U \ frontier B\ by blast show ?thesis unfolding connected_component_def proof (intro exI conjI) have "connected ((insert p A) \ (insert p B))" proof (rule connected_Un) show "connected (insert p A)" by (metis A_def IntD1 frBA \p \ frontier B\ closure_insert closure_subset connected_connected_component connected_intermediate_closure frontier_closures insert_absorb subsetCE subset_insertI) show "connected (insert p B)" by (metis B_def IntD1 \p \ frontier B\ closure_insert closure_subset connected_connected_component connected_intermediate_closure frontier_closures insert_absorb subset_insertI) qed blast then show "connected (insert p (B \ A))" by (simp add: sup.commute) have "A \ - U" using A_def Usub \frontier B \ S\ connected_component_subset by fastforce moreover have "B \ - U" using B_def Usub connected_component_subset frBA frontier_closures by fastforce ultimately show "insert p (B \ A) \ - U" using p by auto qed (auto simp: \a \ A\ \b \ B\) qed qed qed lemma frontier_minimal_separating_closed_pointwise: fixes S :: "'a::real_normed_vector set" assumes S: "closed S" "a \ S" and nconn: "\ connected_component (- S) a b" and conn: "\T. \closed T; T \ S\ \ connected_component (- T) a b" shows "frontier(connected_component_set (- S) a) = S" (is "?F = S") proof - have "?F \ S" by (simp add: S componentsI frontier_of_components_closed_complement) moreover have False if "?F \ S" proof - have "connected_component (- ?F) a b" by (simp add: conn that) then obtain T where "connected T" "T \ -?F" "a \ T" "b \ T" by (auto simp: connected_component_def) moreover have "T \ ?F \ {}" proof (rule connected_Int_frontier [OF \connected T\]) show "T \ connected_component_set (- S) a \ {}" using \a \ S\ \a \ T\ by fastforce show "T - connected_component_set (- S) a \ {}" using \b \ T\ nconn by blast qed ultimately show ?thesis by blast qed ultimately show ?thesis by blast qed subsection\Unicoherence (closed)\ definition\<^marker>\tag important\ unicoherent where "unicoherent U \ \S T. connected S \ connected T \ S \ T = U \ closedin (top_of_set U) S \ closedin (top_of_set U) T \ connected (S \ T)" lemma unicoherentI [intro?]: assumes "\S T. \connected S; connected T; U = S \ T; closedin (top_of_set U) S; closedin (top_of_set U) T\ \ connected (S \ T)" shows "unicoherent U" using assms unfolding unicoherent_def by blast lemma unicoherentD: assumes "unicoherent U" "connected S" "connected T" "U = S \ T" "closedin (top_of_set U) S" "closedin (top_of_set U) T" shows "connected (S \ T)" using assms unfolding unicoherent_def by blast proposition homeomorphic_unicoherent: assumes ST: "S homeomorphic T" and S: "unicoherent S" shows "unicoherent T" proof - obtain f g where gf: "\x. x \ S \ g (f x) = x" and fim: "T = f ` S" and gfim: "g ` f ` S = S" and contf: "continuous_on S f" and contg: "continuous_on (f ` S) g" using ST by (auto simp: homeomorphic_def homeomorphism_def) show ?thesis proof fix U V assume "connected U" "connected V" and T: "T = U \ V" and cloU: "closedin (top_of_set T) U" and cloV: "closedin (top_of_set T) V" have "f ` (g ` U \ g ` V) \ U" "f ` (g ` U \ g ` V) \ V" using gf fim T by auto (metis UnCI image_iff)+ moreover have "U \ V \ f ` (g ` U \ g ` V)" using gf fim by (force simp: image_iff T) ultimately have "U \ V = f ` (g ` U \ g ` V)" by blast moreover have "connected (f ` (g ` U \ g ` V))" proof (rule connected_continuous_image) show "continuous_on (g ` U \ g ` V) f" - apply (rule continuous_on_subset [OF contf]) - using T fim gfim by blast + using T fim gfim by (metis Un_upper1 contf continuous_on_subset image_mono inf_le1) show "connected (g ` U \ g ` V)" proof (intro conjI unicoherentD [OF S]) show "connected (g ` U)" "connected (g ` V)" using \connected U\ cloU \connected V\ cloV by (metis Topological_Spaces.connected_continuous_image closedin_imp_subset contg continuous_on_subset fim)+ show "S = g ` U \ g ` V" using T fim gfim by auto have hom: "homeomorphism T S g f" by (simp add: contf contg fim gf gfim homeomorphism_def) have "closedin (top_of_set T) U" "closedin (top_of_set T) V" by (simp_all add: cloU cloV) then show "closedin (top_of_set S) (g ` U)" "closedin (top_of_set S) (g ` V)" by (blast intro: homeomorphism_imp_closed_map [OF hom])+ qed qed ultimately show "connected (U \ V)" by metis qed qed lemma homeomorphic_unicoherent_eq: "S homeomorphic T \ (unicoherent S \ unicoherent T)" by (meson homeomorphic_sym homeomorphic_unicoherent) lemma unicoherent_translation: fixes S :: "'a::real_normed_vector set" shows "unicoherent (image (\x. a + x) S) \ unicoherent S" using homeomorphic_translation homeomorphic_unicoherent_eq by blast lemma unicoherent_injective_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "(unicoherent(f ` S) \ unicoherent S)" using assms homeomorphic_unicoherent_eq linear_homeomorphic_image by blast lemma Borsukian_imp_unicoherent: fixes U :: "'a::euclidean_space set" assumes "Borsukian U" shows "unicoherent U" unfolding unicoherent_def proof clarify fix S T assume "connected S" "connected T" "U = S \ T" and cloS: "closedin (top_of_set (S \ T)) S" and cloT: "closedin (top_of_set (S \ T)) T" show "connected (S \ T)" unfolding connected_closedin_eq proof clarify fix V W assume "closedin (top_of_set (S \ T)) V" and "closedin (top_of_set (S \ T)) W" and VW: "V \ W = S \ T" "V \ W = {}" and "V \ {}" "W \ {}" then have cloV: "closedin (top_of_set U) V" and cloW: "closedin (top_of_set U) W" using \U = S \ T\ cloS cloT closedin_trans by blast+ obtain q where contq: "continuous_on U q" and q01: "\x. x \ U \ q x \ {0..1::real}" and qV: "\x. x \ V \ q x = 0" and qW: "\x. x \ W \ q x = 1" by (rule Urysohn_local [OF cloV cloW \V \ W = {}\, of 0 1]) (fastforce simp: closed_segment_eq_real_ivl) let ?h = "\x. if x \ S then exp(pi * \ * q x) else 1 / exp(pi * \ * q x)" have eqST: "exp(pi * \ * q x) = 1 / exp(pi * \ * q x)" if "x \ S \ T" for x proof - have "x \ V \ W" using that \V \ W = S \ T\ by blast with qV qW show ?thesis by force qed obtain g where contg: "continuous_on U g" and circle: "g ` U \ sphere 0 1" and S: "\x. x \ S \ g x = exp(pi * \ * q x)" and T: "\x. x \ T \ g x = 1 / exp(pi * \ * q x)" proof show "continuous_on U ?h" unfolding \U = S \ T\ proof (rule continuous_on_cases_local [OF cloS cloT]) show "continuous_on S (\x. exp (pi * \ * q x))" - apply (intro continuous_intros) - using \U = S \ T\ continuous_on_subset contq by blast + proof (intro continuous_intros) + show "continuous_on S q" + using \U = S \ T\ continuous_on_subset contq by blast + qed show "continuous_on T (\x. 1 / exp (pi * \ * q x))" - apply (intro continuous_intros) - using \U = S \ T\ continuous_on_subset contq by auto + proof (intro continuous_intros) + show "continuous_on T q" + using \U = S \ T\ continuous_on_subset contq by auto + qed auto qed (use eqST in auto) qed (use eqST in \auto simp: norm_divide\) then obtain h where conth: "continuous_on U h" and heq: "\x. x \ U \ g x = exp (h x)" by (metis Borsukian_continuous_logarithm_circle assms) obtain v w where "v \ V" "w \ W" using \V \ {}\ \W \ {}\ by blast then have vw: "v \ S \ T" "w \ S \ T" using VW by auto have iff: "2 * pi \ cmod (2 * of_int m * of_real pi * \ - 2 * of_int n * of_real pi * \) \ 1 \ abs (m - n)" for m n proof - have "2 * pi \ cmod (2 * of_int m * of_real pi * \ - 2 * of_int n * of_real pi * \) \ 2 * pi \ cmod ((2 * pi * \) * (of_int m - of_int n))" by (simp add: algebra_simps) also have "... \ 2 * pi \ 2 * pi * cmod (of_int m - of_int n)" by (simp add: norm_mult) also have "... \ 1 \ abs (m - n)" by simp (metis norm_of_int of_int_1_le_iff of_int_abs of_int_diff) finally show ?thesis . qed have *: "\n::int. h x - (pi * \ * q x) = (of_int(2*n) * pi) * \" if "x \ S" for x using that S \U = S \ T\ heq exp_eq [symmetric] by (simp add: algebra_simps) moreover have "(\x. h x - (pi * \ * q x)) constant_on S" proof (rule continuous_discrete_range_constant [OF \connected S\]) have "continuous_on S h" "continuous_on S q" using \U = S \ T\ continuous_on_subset conth contq by blast+ then show "continuous_on S (\x. h x - (pi * \ * q x))" by (intro continuous_intros) have "2*pi \ cmod (h y - (pi * \ * q y) - (h x - (pi * \ * q x)))" if "x \ S" "y \ S" and ne: "h y - (pi * \ * q y) \ h x - (pi * \ * q x)" for x y using * [OF \x \ S\] * [OF \y \ S\] ne by (auto simp: iff) then show "\x. x \ S \ \e>0. \y. y \ S \ h y - (pi * \ * q y) \ h x - (pi * \ * q x) \ e \ cmod (h y - (pi * \ * q y) - (h x - (pi * \ * q x)))" by (rule_tac x="2*pi" in exI) auto qed ultimately obtain m where m: "\x. x \ S \ h x - (pi * \ * q x) = (of_int(2*m) * pi) * \" using vw by (force simp: constant_on_def) have *: "\n::int. h x = - (pi * \ * q x) + (of_int(2*n) * pi) * \" if "x \ T" for x unfolding exp_eq [symmetric] using that T \U = S \ T\ by (simp add: exp_minus field_simps heq [symmetric]) moreover have "(\x. h x + (pi * \ * q x)) constant_on T" proof (rule continuous_discrete_range_constant [OF \connected T\]) have "continuous_on T h" "continuous_on T q" using \U = S \ T\ continuous_on_subset conth contq by blast+ then show "continuous_on T (\x. h x + (pi * \ * q x))" by (intro continuous_intros) have "2*pi \ cmod (h y + (pi * \ * q y) - (h x + (pi * \ * q x)))" if "x \ T" "y \ T" and ne: "h y + (pi * \ * q y) \ h x + (pi * \ * q x)" for x y using * [OF \x \ T\] * [OF \y \ T\] ne by (auto simp: iff) then show "\x. x \ T \ \e>0. \y. y \ T \ h y + (pi * \ * q y) \ h x + (pi * \ * q x) \ e \ cmod (h y + (pi * \ * q y) - (h x + (pi * \ * q x)))" by (rule_tac x="2*pi" in exI) auto qed ultimately obtain n where n: "\x. x \ T \ h x + (pi * \ * q x) = (of_int(2*n) * pi) * \" using vw by (force simp: constant_on_def) show "False" using m [of v] m [of w] n [of v] n [of w] vw by (auto simp: algebra_simps \v \ V\ \w \ W\ qV qW) qed qed corollary contractible_imp_unicoherent: fixes U :: "'a::euclidean_space set" assumes "contractible U" shows "unicoherent U" by (simp add: Borsukian_imp_unicoherent assms contractible_imp_Borsukian) corollary convex_imp_unicoherent: fixes U :: "'a::euclidean_space set" assumes "convex U" shows "unicoherent U" by (simp add: Borsukian_imp_unicoherent assms convex_imp_Borsukian) text\If the type class constraint can be relaxed, I don't know how!\ corollary unicoherent_UNIV: "unicoherent (UNIV :: 'a :: euclidean_space set)" by (simp add: convex_imp_unicoherent) lemma unicoherent_monotone_image_compact: fixes T :: "'b :: t2_space set" assumes S: "unicoherent S" "compact S" and contf: "continuous_on S f" and fim: "f ` S = T" and conn: "\y. y \ T \ connected (S \ f -` {y})" shows "unicoherent T" proof fix U V assume UV: "connected U" "connected V" "T = U \ V" and cloU: "closedin (top_of_set T) U" and cloV: "closedin (top_of_set T) V" moreover have "compact T" using \compact S\ compact_continuous_image contf fim by blast ultimately have "closed U" "closed V" by (auto simp: closedin_closed_eq compact_imp_closed) let ?SUV = "(S \ f -` U) \ (S \ f -` V)" have UV_eq: "f ` ?SUV = U \ V" using \T = U \ V\ fim by force+ have "connected (f ` ?SUV)" proof (rule connected_continuous_image) show "continuous_on ?SUV f" by (meson contf continuous_on_subset inf_le1) show "connected ?SUV" proof (rule unicoherentD [OF \unicoherent S\, of "S \ f -` U" "S \ f -` V"]) have "\C. closedin (top_of_set S) C \ closedin (top_of_set T) (f ` C)" by (metis \compact S\ closed_subset closedin_compact closedin_imp_subset compact_continuous_image compact_imp_closed contf continuous_on_subset fim image_mono) then show "connected (S \ f -` U)" "connected (S \ f -` V)" using UV by (auto simp: conn intro: connected_closed_monotone_preimage [OF contf fim]) show "S = (S \ f -` U) \ (S \ f -` V)" using UV fim by blast show "closedin (top_of_set S) (S \ f -` U)" "closedin (top_of_set S) (S \ f -` V)" by (auto simp: continuous_on_imp_closedin cloU cloV contf fim) qed qed with UV_eq show "connected (U \ V)" by simp qed subsection\Several common variants of unicoherence\ lemma connected_frontier_simple: fixes S :: "'a :: euclidean_space set" assumes "connected S" "connected(- S)" shows "connected(frontier S)" unfolding frontier_closures - apply (rule unicoherentD [OF unicoherent_UNIV]) - apply (simp_all add: assms connected_imp_connected_closure) - by (simp add: closure_def) + by (rule unicoherentD [OF unicoherent_UNIV]; simp add: assms connected_imp_connected_closure flip: closure_Un) lemma connected_frontier_component_complement: fixes S :: "'a :: euclidean_space set" - assumes "connected S" and C: "C \ components(- S)" shows "connected(frontier C)" - apply (rule connected_frontier_simple) - using C in_components_connected apply blast - by (metis assms component_complement_connected) + assumes "connected S" "C \ components(- S)" shows "connected(frontier C)" + by (meson assms component_complement_connected connected_frontier_simple in_components_connected) lemma connected_frontier_disjoint: fixes S :: "'a :: euclidean_space set" assumes "connected S" "connected T" "disjnt S T" and ST: "frontier S \ frontier T" shows "connected(frontier S)" proof (cases "S = UNIV") case True then show ?thesis by simp next case False then have "-S \ {}" by blast then obtain C where C: "C \ components(- S)" and "T \ C" by (metis ComplI disjnt_iff subsetI exists_component_superset \disjnt S T\ \connected T\) moreover have "frontier S = frontier C" proof - have "frontier C \ frontier S" using C frontier_complement frontier_of_components_subset by blast moreover have "x \ frontier C" if "x \ frontier S" for x proof - have "x \ closure C" using that unfolding frontier_def by (metis (no_types) Diff_eq ST \T \ C\ closure_mono contra_subsetD frontier_def le_inf_iff that) moreover have "x \ interior C" using that unfolding frontier_def by (metis C Compl_eq_Diff_UNIV Diff_iff subsetD in_components_subset interior_diff interior_mono) ultimately show ?thesis by (auto simp: frontier_def) qed ultimately show ?thesis by blast qed ultimately show ?thesis using \connected S\ connected_frontier_component_complement by auto qed subsection\Some separation results\ lemma separation_by_component_closed_pointwise: fixes S :: "'a :: euclidean_space set" assumes "closed S" "\ connected_component (- S) a b" obtains C where "C \ components S" "\ connected_component(- C) a b" proof (cases "a \ S \ b \ S") case True then show ?thesis using connected_component_in componentsI that by fastforce next case False obtain T where "T \ S" "closed T" "T \ {}" and nab: "\ connected_component (- T) a b" and conn: "\U. U \ T \ connected_component (- U) a b" using closed_irreducible_separator [OF assms] by metis moreover have "connected T" proof - have ab: "frontier(connected_component_set (- T) a) = T" "frontier(connected_component_set (- T) b) = T" using frontier_minimal_separating_closed_pointwise by (metis False \T \ S\ \closed T\ connected_component_sym conn connected_component_eq_empty connected_component_intermediate_subset empty_subsetI nab)+ have "connected (frontier (connected_component_set (- T) a))" proof (rule connected_frontier_disjoint) show "disjnt (connected_component_set (- T) a) (connected_component_set (- T) b)" unfolding disjnt_iff by (metis connected_component_eq connected_component_eq_empty connected_component_idemp mem_Collect_eq nab) show "frontier (connected_component_set (- T) a) \ frontier (connected_component_set (- T) b)" by (simp add: ab) qed auto with ab \closed T\ show ?thesis by simp qed ultimately obtain C where "C \ components S" "T \ C" using exists_component_superset [of T S] by blast then show ?thesis by (meson Compl_anti_mono connected_component_of_subset nab that) qed lemma separation_by_component_closed: fixes S :: "'a :: euclidean_space set" assumes "closed S" "\ connected(- S)" obtains C where "C \ components S" "\ connected(- C)" proof - obtain x y where "closed S" "x \ S" "y \ S" and "\ connected_component (- S) x y" using assms by (auto simp: connected_iff_connected_component) then obtain C where "C \ components S" "\ connected_component(- C) x y" using separation_by_component_closed_pointwise by metis then show "thesis" - apply (clarify elim!: componentsE) - by (metis Compl_iff \C \ components S\ \x \ S\ \y \ S\ connected_component_eq connected_component_eq_eq connected_iff_connected_component that) + by (metis Compl_iff \x \ S\ \y \ S\ connected_component_eq_self in_components_subset mem_Collect_eq subsetD that) qed lemma separation_by_Un_closed_pointwise: fixes S :: "'a :: euclidean_space set" assumes ST: "closed S" "closed T" "S \ T = {}" and conS: "connected_component (- S) a b" and conT: "connected_component (- T) a b" shows "connected_component (- (S \ T)) a b" proof (rule ccontr) have "a \ S" "b \ S" "a \ T" "b \ T" using conS conT connected_component_in by auto assume "\ connected_component (- (S \ T)) a b" then obtain C where "C \ components (S \ T)" and C: "\ connected_component(- C) a b" using separation_by_component_closed_pointwise assms by blast then have "C \ S \ C \ T" proof - have "connected C" "C \ S \ T" using \C \ components (S \ T)\ in_components_subset by (blast elim: componentsE)+ moreover then have "C \ T = {} \ C \ S = {}" by (metis Int_empty_right ST inf.commute connected_closed) ultimately show ?thesis by blast qed then show False by (meson Compl_anti_mono C conS conT connected_component_of_subset) qed lemma separation_by_Un_closed: fixes S :: "'a :: euclidean_space set" assumes ST: "closed S" "closed T" "S \ T = {}" and conS: "connected(- S)" and conT: "connected(- T)" shows "connected(- (S \ T))" using assms separation_by_Un_closed_pointwise by (fastforce simp add: connected_iff_connected_component) lemma open_unicoherent_UNIV: fixes S :: "'a :: euclidean_space set" assumes "open S" "open T" "connected S" "connected T" "S \ T = UNIV" shows "connected(S \ T)" proof - have "connected(- (-S \ -T))" by (metis closed_Compl compl_sup compl_top_eq double_compl separation_by_Un_closed assms) then show ?thesis by simp qed lemma separation_by_component_open_aux: fixes S :: "'a :: euclidean_space set" assumes ST: "closed S" "closed T" "S \ T = {}" and "S \ {}" "T \ {}" obtains C where "C \ components(-(S \ T))" "C \ {}" "frontier C \ S \ {}" "frontier C \ T \ {}" proof (rule ccontr) let ?S = "S \ \{C \ components(- (S \ T)). frontier C \ S}" let ?T = "T \ \{C \ components(- (S \ T)). frontier C \ T}" assume "\ thesis" with that have *: "frontier C \ S = {} \ frontier C \ T = {}" if C: "C \ components (- (S \ T))" "C \ {}" for C using C by blast have "\A B::'a set. closed A \ closed B \ UNIV \ A \ B \ A \ B = {} \ A \ {} \ B \ {}" proof (intro exI conjI) have "frontier (\{C \ components (- S \ - T). frontier C \ S}) \ S" - apply (rule subset_trans [OF frontier_Union_subset_closure]) + using subset_trans [OF frontier_Union_subset_closure] by (metis (no_types, lifting) SUP_least \closed S\ closure_minimal mem_Collect_eq) then have "frontier ?S \ S" by (simp add: frontier_subset_eq assms subset_trans [OF frontier_Un_subset]) then show "closed ?S" using frontier_subset_eq by fastforce have "frontier (\{C \ components (- S \ - T). frontier C \ T}) \ T" - apply (rule subset_trans [OF frontier_Union_subset_closure]) + using subset_trans [OF frontier_Union_subset_closure] by (metis (no_types, lifting) SUP_least \closed T\ closure_minimal mem_Collect_eq) then have "frontier ?T \ T" by (simp add: frontier_subset_eq assms subset_trans [OF frontier_Un_subset]) then show "closed ?T" using frontier_subset_eq by fastforce have "UNIV \ (S \ T) \ \(components(- (S \ T)))" using Union_components by blast also have "... \ ?S \ ?T" proof - have "C \ components (-(S \ T)) \ frontier C \ S \ C \ components (-(S \ T)) \ frontier C \ T" if "C \ components (- (S \ T))" "C \ {}" for C using * [OF that] that by clarify (metis (no_types, lifting) UnE \closed S\ \closed T\ closed_Un disjoint_iff_not_equal frontier_of_components_closed_complement subsetCE) then show ?thesis by blast qed finally show "UNIV \ ?S \ ?T" . have "\{C \ components (- (S \ T)). frontier C \ S} \ \{C \ components (- (S \ T)). frontier C \ T} \ - (S \ T)" using in_components_subset by fastforce moreover have "\{C \ components (- (S \ T)). frontier C \ S} \ \{C \ components (- (S \ T)). frontier C \ T} = {}" proof - have "C \ C' = {}" if "C \ components (- (S \ T))" "frontier C \ S" "C' \ components (- (S \ T))" "frontier C' \ T" for C C' proof - have NUN: "- S \ - T \ UNIV" using \T \ {}\ by blast have "C \ C'" proof assume "C = C'" with that have "frontier C' \ S \ T" by simp also have "... = {}" using \S \ T = {}\ by blast finally have "C' = {} \ C' = UNIV" using frontier_eq_empty by auto then show False using \C = C'\ NUN that by (force simp: dest: in_components_nonempty in_components_subset) qed with that show ?thesis by (simp add: components_nonoverlap [of _ "-(S \ T)"]) qed then show ?thesis by blast qed ultimately show "?S \ ?T = {}" using ST by blast show "?S \ {}" "?T \ {}" using \S \ {}\ \T \ {}\ by blast+ qed then show False by (metis Compl_disjoint connected_UNIV compl_bot_eq compl_unique connected_closedD inf_sup_absorb sup_compl_top_left1 top.extremum_uniqueI) qed proposition separation_by_component_open: fixes S :: "'a :: euclidean_space set" assumes "open S" and non: "\ connected(- S)" obtains C where "C \ components S" "\ connected(- C)" proof - obtain T U where "closed T" "closed U" and TU: "T \ U = - S" "T \ U = {}" "T \ {}" "U \ {}" using assms by (auto simp: connected_closed_set closed_def) then obtain C where C: "C \ components(-(T \ U))" "C \ {}" and "frontier C \ T \ {}" "frontier C \ U \ {}" using separation_by_component_open_aux [OF \closed T\ \closed U\ \T \ U = {}\] by force show "thesis" proof show "C \ components S" using C(1) TU(1) by auto show "\ connected (- C)" proof assume "connected (- C)" then have "connected (frontier C)" using connected_frontier_simple [of C] \C \ components S\ in_components_connected by blast then show False unfolding connected_closed by (metis C(1) TU(2) \closed T\ \closed U\ \frontier C \ T \ {}\ \frontier C \ U \ {}\ closed_Un frontier_of_components_closed_complement inf_bot_right inf_commute) qed qed qed lemma separation_by_Un_open: fixes S :: "'a :: euclidean_space set" assumes "open S" "open T" "S \ T = {}" and cS: "connected(-S)" and cT: "connected(-T)" shows "connected(- (S \ T))" using assms unicoherent_UNIV unfolding unicoherent_def by force lemma nonseparation_by_component_eq: fixes S :: "'a :: euclidean_space set" assumes "open S \ closed S" shows "((\C \ components S. connected(-C)) \ connected(- S))" (is "?lhs = ?rhs") proof assume ?lhs with assms show ?rhs by (meson separation_by_component_closed separation_by_component_open) next assume ?rhs with assms show ?lhs using component_complement_connected by force qed text\Another interesting equivalent of an inessential mapping into C-{0}\ proposition inessential_eq_extensible: fixes f :: "'a::euclidean_space \ complex" assumes "closed S" shows "(\a. homotopic_with_canon (\h. True) S (-{0}) f (\t. a)) \ (\g. continuous_on UNIV g \ (\x \ S. g x = f x) \ (\x. g x \ 0))" (is "?lhs = ?rhs") proof assume ?lhs then obtain a where a: "homotopic_with_canon (\h. True) S (-{0}) f (\t. a)" .. show ?rhs proof (cases "S = {}") case True with a show ?thesis by force next case False have anr: "ANR (-{0::complex})" by (simp add: ANR_delete open_Compl open_imp_ANR) obtain g where contg: "continuous_on UNIV g" and gim: "g ` UNIV \ -{0}" and gf: "\x. x \ S \ g x = f x" proof (rule Borsuk_homotopy_extension_homotopic [OF _ _ continuous_on_const _ homotopic_with_symD [OF a]]) show "closedin (top_of_set UNIV) S" using assms by auto show "range (\t. a) \ - {0}" using a homotopic_with_imp_subset2 False by blast qed (use anr that in \force+\) then show ?thesis by force qed next assume ?rhs then obtain g where contg: "continuous_on UNIV g" and gf: "\x. x \ S \ g x = f x" and non0: "\x. g x \ 0" by metis obtain h k::"'a\'a" where hk: "homeomorphism (ball 0 1) UNIV h k" using homeomorphic_ball01_UNIV homeomorphic_def by blast then have "continuous_on (ball 0 1) (g \ h)" by (meson contg continuous_on_compose continuous_on_subset homeomorphism_cont1 top_greatest) then obtain j where contj: "continuous_on (ball 0 1) j" and j: "\z. z \ ball 0 1 \ exp(j z) = (g \ h) z" by (metis (mono_tags, hide_lams) continuous_logarithm_on_ball comp_apply non0) have [simp]: "\x. x \ S \ h (k x) = x" using hk homeomorphism_apply2 by blast have "\\. continuous_on S \\ (\x\S. f x = exp (\ x))" proof (intro exI conjI ballI) show "continuous_on S (j \ k)" proof (rule continuous_on_compose) show "continuous_on S k" by (meson continuous_on_subset hk homeomorphism_cont2 top_greatest) show "continuous_on (k ` S) j" - apply (rule continuous_on_subset [OF contj]) - using homeomorphism_image2 [OF hk] continuous_on_subset [OF contj] by blast + by (auto intro: continuous_on_subset [OF contj] simp flip: homeomorphism_image2 [OF hk]) qed show "f x = exp ((j \ k) x)" if "x \ S" for x proof - have "f x = (g \ h) (k x)" by (simp add: gf that) also have "... = exp (j (k x))" by (metis rangeI homeomorphism_image2 [OF hk] j) finally show ?thesis by simp qed qed then show ?lhs by (simp add: inessential_eq_continuous_logarithm) qed lemma inessential_on_clopen_Union: fixes \ :: "'a::euclidean_space set set" assumes T: "path_connected T" and "\S. S \ \ \ closedin (top_of_set (\\)) S" and "\S. S \ \ \ openin (top_of_set (\\)) S" and hom: "\S. S \ \ \ \a. homotopic_with_canon (\x. True) S T f (\x. a)" obtains a where "homotopic_with_canon (\x. True) (\\) T f (\x. a)" proof (cases "\\ = {}") case True with that show ?thesis by force next case False then obtain C where "C \ \" "C \ {}" by blast then obtain a where clo: "closedin (top_of_set (\\)) C" and ope: "openin (top_of_set (\\)) C" and "homotopic_with_canon (\x. True) C T f (\x. a)" using assms by blast with \C \ {}\ have "f ` C \ T" "a \ T" using homotopic_with_imp_subset1 homotopic_with_imp_subset2 by blast+ have "homotopic_with_canon (\x. True) (\\) T f (\x. a)" proof (rule homotopic_on_clopen_Union) show "\S. S \ \ \ closedin (top_of_set (\\)) S" "\S. S \ \ \ openin (top_of_set (\\)) S" by (simp_all add: assms) show "homotopic_with_canon (\x. True) S T f (\x. a)" if "S \ \" for S proof (cases "S = {}") case True then show ?thesis by auto next case False then obtain b where "b \ S" by blast obtain c where c: "homotopic_with_canon (\x. True) S T f (\x. c)" using \S \ \\ hom by blast then have "c \ T" using \b \ S\ homotopic_with_imp_subset2 by blast then have "homotopic_with_canon (\x. True) S T (\x. a) (\x. c)" using T \a \ T\ homotopic_constant_maps path_connected_component by (simp add: homotopic_constant_maps path_connected_component) then show ?thesis using c homotopic_with_symD homotopic_with_trans by blast qed qed then show ?thesis .. qed proposition Janiszewski_dual: fixes S :: "complex set" assumes "compact S" "compact T" "connected S" "connected T" "connected(- (S \ T))" shows "connected(S \ T)" proof - have ST: "compact (S \ T)" by (simp add: assms compact_Un) with Borsukian_imp_unicoherent [of "S \ T"] ST assms show ?thesis by (auto simp: closed_subset compact_imp_closed Borsukian_separation_compact unicoherent_def) qed end diff --git a/src/HOL/Deriv.thy b/src/HOL/Deriv.thy --- a/src/HOL/Deriv.thy +++ b/src/HOL/Deriv.thy @@ -1,2370 +1,2370 @@ (* 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)" - unfolding has_derivative_iff_norm LIM_offset_zero_iff[of _ _ x] by simp + 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) lemmas has_derivative_within_subset = has_derivative_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_within_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_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_within_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,3035 @@ (* 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 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 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: "f \a\ L \ (\h. f (a + h)) \0\ L" +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 "a \ S" "open S" + assumes " NO_MATCH 0 a" "a \ S" "open S" shows "(f \ L) (at a within S) \ ((\h. f (a + h)) \ L) (at 0)" - by (metis LIM_offset_zero_iff assms tendsto_within_open) + 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