diff --git a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy --- a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy +++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy @@ -1,7539 +1,7549 @@ (* Author: John Harrison Author: Robert Himmelmann, TU Muenchen (Translation from HOL light) Huge cleanup by LCP *) section \Henstock-Kurzweil Gauge Integration in Many Dimensions\ theory Henstock_Kurzweil_Integration imports Lebesgue_Measure Tagged_Division begin lemma norm_diff2: "\y = y1 + y2; x = x1 + x2; e = e1 + e2; norm(y1 - x1) \ e1; norm(y2 - x2) \ e2\ \ norm(y-x) \ e" by (smt (verit, ccfv_SIG) norm_diff_triangle_ineq) lemma setcomp_dot1: "{z. P (z \ (i,0))} = {(x,y). P(x \ i)}" by auto lemma setcomp_dot2: "{z. P (z \ (0,i))} = {(x,y). P(y \ i)}" by auto lemma Sigma_Int_Paircomp1: "(Sigma A B) \ {(x, y). P x} = Sigma (A \ {x. P x}) B" by blast lemma Sigma_Int_Paircomp2: "(Sigma A B) \ {(x, y). P y} = Sigma A (\z. B z \ {y. P y})" by blast (* END MOVE *) subsection \Content (length, area, volume...) of an interval\ abbreviation content :: "'a::euclidean_space set \ real" where "content s \ measure lborel s" lemma content_cbox_cases: "content (cbox a b) = (if \i\Basis. a\i \ b\i then prod (\i. b\i - a\i) Basis else 0)" by (simp add: measure_lborel_cbox_eq inner_diff) lemma content_cbox: "\i\Basis. a\i \ b\i \ content (cbox a b) = (\i\Basis. b\i - a\i)" unfolding content_cbox_cases by simp lemma content_cbox': "cbox a b \ {} \ content (cbox a b) = (\i\Basis. b\i - a\i)" by (simp add: box_ne_empty inner_diff) lemma content_cbox_if: "content (cbox a b) = (if cbox a b = {} then 0 else \i\Basis. b\i - a\i)" by (simp add: content_cbox') lemma content_cbox_cart: "cbox a b \ {} \ content(cbox a b) = prod (\i. b$i - a$i) UNIV" by (simp add: content_cbox_if Basis_vec_def cart_eq_inner_axis axis_eq_axis prod.UNION_disjoint) lemma content_cbox_if_cart: "content(cbox a b) = (if cbox a b = {} then 0 else prod (\i. b$i - a$i) UNIV)" by (simp add: content_cbox_cart) lemma content_division_of: assumes "K \ \" "\ division_of S" shows "content K = (\i \ Basis. interval_upperbound K \ i - interval_lowerbound K \ i)" proof - obtain a b where "K = cbox a b" using cbox_division_memE assms by metis then show ?thesis using assms by (force simp: division_of_def content_cbox') qed lemma content_real: "a \ b \ content {a..b} = b - a" by simp lemma abs_eq_content: "\y - x\ = (if x\y then content {x..y} else content {y..x})" by (auto simp: content_real) lemma content_singleton: "content {a} = 0" by simp lemma content_unit[iff]: "content (cbox 0 (One::'a::euclidean_space)) = 1" by simp lemma content_pos_le [iff]: "0 \ content X" by simp corollary\<^marker>\tag unimportant\ content_nonneg [simp]: "\ content (cbox a b) < 0" using not_le by blast lemma content_pos_lt: "\i\Basis. a\i < b\i \ 0 < content (cbox a b)" by (auto simp: less_imp_le inner_diff box_eq_empty intro!: prod_pos) lemma content_eq_0: "content (cbox a b) = 0 \ (\i\Basis. b\i \ a\i)" by (auto simp: content_cbox_cases not_le intro: less_imp_le antisym eq_refl) lemma content_eq_0_interior: "content (cbox a b) = 0 \ interior(cbox a b) = {}" unfolding content_eq_0 interior_cbox box_eq_empty by auto lemma content_pos_lt_eq: "0 < content (cbox a (b::'a::euclidean_space)) \ (\i\Basis. a\i < b\i)" by (auto simp add: content_cbox_cases less_le prod_nonneg) lemma content_empty [simp]: "content {} = 0" by simp lemma content_real_if [simp]: "content {a..b} = (if a \ b then b - a else 0)" by (simp add: content_real) lemma content_subset: "cbox a b \ cbox c d \ content (cbox a b) \ content (cbox c d)" unfolding measure_def by (intro enn2real_mono emeasure_mono) (auto simp: emeasure_lborel_cbox_eq) lemma content_lt_nz: "0 < content (cbox a b) \ content (cbox a b) \ 0" by (fact zero_less_measure_iff) lemma content_Pair: "content (cbox (a,c) (b,d)) = content (cbox a b) * content (cbox c d)" unfolding measure_lborel_cbox_eq Basis_prod_def apply (subst prod.union_disjoint) apply (auto simp: bex_Un ball_Un) apply (subst (1 2) prod.reindex_nontrivial) apply auto done lemma content_cbox_pair_eq0_D: "content (cbox (a,c) (b,d)) = 0 \ content (cbox a b) = 0 \ content (cbox c d) = 0" by (simp add: content_Pair) lemma content_cbox_plus: fixes x :: "'a::euclidean_space" shows "content(cbox x (x + h *\<^sub>R One)) = (if h \ 0 then h ^ DIM('a) else 0)" by (simp add: algebra_simps content_cbox_if box_eq_empty) lemma content_0_subset: "content(cbox a b) = 0 \ s \ cbox a b \ content s = 0" using emeasure_mono[of s "cbox a b" lborel] by (auto simp: measure_def enn2real_eq_0_iff emeasure_lborel_cbox_eq) lemma content_ball_pos: assumes "r > 0" shows "content (ball c r) > 0" proof - from rational_boxes[OF assms, of c] obtain a b where ab: "c \ box a b" "box a b \ ball c r" by auto then have "0 < content (box a b)" by (subst measure_lborel_box_eq) (auto intro!: prod_pos simp: algebra_simps box_def) have "emeasure lborel (box a b) \ emeasure lborel (ball c r)" using ab by (intro emeasure_mono) auto then show ?thesis using \content (box a b) > 0\ by (smt (verit, best) Sigma_Algebra.measure_def emeasure_lborel_ball_finite enn2real_mono infinity_ennreal_def) qed lemma content_cball_pos: assumes "r > 0" shows "content (cball c r) > 0" proof - from rational_boxes[OF assms, of c] obtain a b where ab: "c \ box a b" "box a b \ ball c r" by auto then have "0 < content (box a b)" by (subst measure_lborel_box_eq) (auto intro!: prod_pos simp: algebra_simps box_def) have "emeasure lborel (box a b) \ emeasure lborel (cball c r)" using ab by (intro emeasure_mono) auto also have "emeasure lborel (box a b) = ennreal (content (box a b))" using emeasure_lborel_box_finite[of a b] by (intro emeasure_eq_ennreal_measure) auto also have "emeasure lborel (cball c r) = ennreal (content (cball c r))" using emeasure_lborel_cball_finite[of c r] by (intro emeasure_eq_ennreal_measure) auto finally show ?thesis using \content (box a b) > 0\ by simp qed lemma content_split: fixes a :: "'a::euclidean_space" assumes "k \ Basis" shows "content (cbox a b) = content(cbox a b \ {x. x\k \ c}) + content(cbox a b \ {x. x\k \ c})" \ \Prove using measure theory\ proof (cases "\i\Basis. a \ i \ b \ i") case True have 1: "\X Y Z. (\i\Basis. Z i (if i = k then X else Y i)) = Z k X * (\i\Basis-{k}. Z i (Y i))" by (simp add: if_distrib prod.delta_remove assms) note simps = interval_split[OF assms] content_cbox_cases have 2: "(\i\Basis. b\i - a\i) = (\i\Basis-{k}. b\i - a\i) * (b\k - a\k)" by (metis (no_types, lifting) assms finite_Basis mult.commute prod.remove) have "\x. min (b \ k) c = max (a \ k) c \ x * (b\k - a\k) = x * (max (a \ k) c - a \ k) + x * (b \ k - max (a \ k) c)" by (auto simp add: field_simps) moreover have **: "(\i\Basis. ((\i\Basis. (if i = k then min (b \ k) c else b \ i) *\<^sub>R i) \ i - a \ i)) = (\i\Basis. (if i = k then min (b \ k) c else b \ i) - a \ i)" "(\i\Basis. b \ i - ((\i\Basis. (if i = k then max (a \ k) c else a \ i) *\<^sub>R i) \ i)) = (\i\Basis. b \ i - (if i = k then max (a \ k) c else a \ i))" by (auto intro!: prod.cong) have "\ a \ k \ c \ \ c \ b \ k \ False" unfolding not_le using True assms by auto ultimately show ?thesis using assms unfolding simps ** 1[of "\i x. b\i - x"] 1[of "\i x. x - a\i"] 2 by auto next case False then have "cbox a b = {}" unfolding box_eq_empty by (auto simp: not_le) then show ?thesis by (auto simp: not_le) qed lemma division_of_content_0: assumes "content (cbox a b) = 0" "d division_of (cbox a b)" "K \ d" shows "content K = 0" unfolding forall_in_division[OF assms(2)] by (meson assms content_0_subset division_of_def) lemma sum_content_null: assumes "content (cbox a b) = 0" and "p tagged_division_of (cbox a b)" shows "(\(x,K)\p. content K *\<^sub>R f x) = (0::'a::real_normed_vector)" proof (intro sum.neutral strip) fix y assume y: "y \ p" obtain x K where xk: "y = (x, K)" using surj_pair[of y] by blast then obtain c d where k: "K = cbox c d" "K \ cbox a b" by (metis assms(2) tagged_division_ofD(3) tagged_division_ofD(4) y) have "(\(x',K'). content K' *\<^sub>R f x') y = content K *\<^sub>R f x" unfolding xk by auto also have "\ = 0" using assms(1) content_0_subset k(2) by auto finally show "(\(x, k). content k *\<^sub>R f x) y = 0" . qed global_interpretation sum_content: operative plus 0 content rewrites "comm_monoid_set.F plus 0 = sum" proof - interpret operative plus 0 content by standard (auto simp add: content_split [symmetric] content_eq_0_interior) show "operative plus 0 content" by standard show "comm_monoid_set.F plus 0 = sum" by (simp add: sum_def) qed lemma additive_content_division: "d division_of (cbox a b) \ sum content d = content (cbox a b)" by (fact sum_content.division) lemma additive_content_tagged_division: "d tagged_division_of (cbox a b) \ sum (\(x,l). content l) d = content (cbox a b)" by (fact sum_content.tagged_division) lemma subadditive_content_division: assumes "\ division_of S" "S \ cbox a b" shows "sum content \ \ content(cbox a b)" proof - have "\ division_of \\" "\\ \ cbox a b" using assms by auto then obtain \' where "\ \ \'" "\' division_of cbox a b" using partial_division_extend_interval by metis then have "sum content \ \ sum content \'" using sum_mono2 by blast also have "... \ content(cbox a b)" by (simp add: \\' division_of cbox a b\ additive_content_division less_eq_real_def) finally show ?thesis . qed lemma content_real_eq_0: "content {a..b::real} = 0 \ a \ b" by simp lemma property_empty_interval: "\a b. content (cbox a b) = 0 \ P (cbox a b) \ P {}" using content_empty unfolding empty_as_interval by auto lemma interval_bounds_nz_content [simp]: assumes "content (cbox a b) \ 0" shows "interval_upperbound (cbox a b) = b" and "interval_lowerbound (cbox a b) = a" by (metis assms content_empty interval_bounds')+ subsection \Gauge integral\ text \Case distinction to define it first on compact intervals first, then use a limit. This is only much later unified. In Fremlin: Measure Theory, Volume 4I this is generalized using residual sets.\ definition has_integral :: "('n::euclidean_space \ 'b::real_normed_vector) \ 'b \ 'n set \ bool" (infixr "has'_integral" 46) where "(f has_integral I) s \ (if \a b. s = cbox a b then ((\p. \(x,k)\p. content k *\<^sub>R f x) \ I) (division_filter s) else (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ (\z. ((\p. \(x,k)\p. content k *\<^sub>R (if x \ s then f x else 0)) \ z) (division_filter (cbox a b)) \ norm (z - I) < e)))" lemma has_integral_cbox: "(f has_integral I) (cbox a b) \ ((\p. \(x,k)\p. content k *\<^sub>R f x) \ I) (division_filter (cbox a b))" by (auto simp add: has_integral_def) lemma has_integral: "(f has_integral y) (cbox a b) \ (\e>0. \\. gauge \ \ (\\. \ tagged_division_of (cbox a b) \ \ fine \ \ norm (sum (\(x,k). content(k) *\<^sub>R f x) \ - y) < e))" by (auto simp: dist_norm eventually_division_filter has_integral_def tendsto_iff) lemma has_integral_real: "(f has_integral y) {a..b::real} \ (\e>0. \\. gauge \ \ (\\. \ tagged_division_of {a..b} \ \ fine \ \ norm (sum (\(x,k). content(k) *\<^sub>R f x) \ - y) < e))" unfolding box_real[symmetric] by (rule has_integral) lemma has_integralD[dest]: assumes "(f has_integral y) (cbox a b)" and "e > 0" obtains \ where "gauge \" and "\\. \ tagged_division_of (cbox a b) \ \ fine \ \ norm ((\(x,k)\\. content k *\<^sub>R f x) - y) < e" using assms unfolding has_integral by auto lemma has_integral_alt: "(f has_integral y) i \ (if \a b. i = cbox a b then (f has_integral y) i else (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ (\z. ((\x. if x \ i then f x else 0) has_integral z) (cbox a b) \ norm (z - y) < e)))" by (subst has_integral_def) (auto simp add: has_integral_cbox) lemma has_integral_altD: assumes "(f has_integral y) i" and "\ (\a b. i = cbox a b)" and "e>0" obtains B where "B > 0" and "\a b. ball 0 B \ cbox a b \ (\z. ((\x. if x \ i then f(x) else 0) has_integral z) (cbox a b) \ norm(z - y) < e)" using assms has_integral_alt[of f y i] by auto definition integrable_on (infixr "integrable'_on" 46) where "f integrable_on i \ (\y. (f has_integral y) i)" definition "integral i f = (SOME y. (f has_integral y) i \ \ f integrable_on i \ y=0)" lemma integrable_integral[intro]: "f integrable_on i \ (f has_integral (integral i f)) i" unfolding integrable_on_def integral_def by (metis (mono_tags, lifting)) lemma not_integrable_integral: "\ f integrable_on i \ integral i f = 0" unfolding integrable_on_def integral_def by blast lemma has_integral_integrable[dest]: "(f has_integral i) s \ f integrable_on s" unfolding integrable_on_def by auto lemma has_integral_integral: "f integrable_on s \ (f has_integral (integral s f)) s" by auto subsection \Basic theorems about integrals\ lemma has_integral_eq_rhs: "(f has_integral j) S \ i = j \ (f has_integral i) S" by (rule forw_subst) lemma has_integral_unique_cbox: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" shows "(f has_integral k1) (cbox a b) \ (f has_integral k2) (cbox a b) \ k1 = k2" by (meson division_filter_not_empty has_integral_cbox tendsto_unique) lemma has_integral_unique: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes "(f has_integral k1) i" "(f has_integral k2) i" shows "k1 = k2" proof (rule ccontr) let ?e = "norm (k1 - k2)/2" let ?F = "(\x. if x \ i then f x else 0)" assume "k1 \ k2" then have e: "?e > 0" by auto have nonbox: "\ (\a b. i = cbox a b)" using \k1 \ k2\ assms has_integral_unique_cbox by blast obtain B1 where B1: "0 < B1" "\a b. ball 0 B1 \ cbox a b \ \z. (?F has_integral z) (cbox a b) \ norm (z - k1) < norm (k1 - k2)/2" by (rule has_integral_altD[OF assms(1) nonbox,OF e]) blast obtain B2 where B2: "0 < B2" "\a b. ball 0 B2 \ cbox a b \ \z. (?F has_integral z) (cbox a b) \ norm (z - k2) < norm (k1 - k2)/2" by (rule has_integral_altD[OF assms(2) nonbox,OF e]) blast obtain a b :: 'n where ab: "ball 0 B1 \ cbox a b" "ball 0 B2 \ cbox a b" by (metis Un_subset_iff bounded_Un bounded_ball bounded_subset_cbox_symmetric) obtain w where w: "(?F has_integral w) (cbox a b)" "norm (w - k1) < norm (k1 - k2)/2" using B1(2)[OF ab(1)] by blast obtain z where z: "(?F has_integral z) (cbox a b)" "norm (z - k2) < norm (k1 - k2)/2" using B2(2)[OF ab(2)] by blast have "z = w" using has_integral_unique_cbox[OF w(1) z(1)] by auto then have "norm (k1 - k2) \ norm (z - k2) + norm (w - k1)" using norm_triangle_ineq4 [of "k1 - w" "k2 - z"] by (auto simp add: norm_minus_commute) also have "\ < norm (k1 - k2)/2 + norm (k1 - k2)/2" by (metis add_strict_mono z(2) w(2)) finally show False by auto qed lemma integral_unique [intro]: "(f has_integral y) k \ integral k f = y" unfolding integral_def by (rule some_equality) (auto intro: has_integral_unique) lemma has_integral_iff: "(f has_integral i) S \ (f integrable_on S \ integral S f = i)" by blast lemma eq_integralD: "integral k f = y \ (f has_integral y) k \ \ f integrable_on k \ y=0" unfolding integral_def integrable_on_def by (metis (mono_tags, lifting)) lemma has_integral_const [intro]: fixes a b :: "'a::euclidean_space" shows "((\x. c) has_integral (content (cbox a b) *\<^sub>R c)) (cbox a b)" using eventually_division_filter_tagged_division[of "cbox a b"] additive_content_tagged_division[of _ a b] by (auto simp: has_integral_cbox split_beta' scaleR_sum_left[symmetric] elim!: eventually_mono intro!: tendsto_cong[THEN iffD1, OF _ tendsto_const]) lemma has_integral_const_real [intro]: fixes a b :: real shows "((\x. c) has_integral (content {a..b} *\<^sub>R c)) {a..b}" by (metis box_real(2) has_integral_const) lemma has_integral_integrable_integral: "(f has_integral i) s \ f integrable_on s \ integral s f = i" by blast lemma integral_const [simp]: fixes a b :: "'a::euclidean_space" shows "integral (cbox a b) (\x. c) = content (cbox a b) *\<^sub>R c" by blast lemma integral_const_real [simp]: fixes a b :: real shows "integral {a..b} (\x. c) = content {a..b} *\<^sub>R c" by blast lemma has_integral_is_0_cbox: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes "\x. x \ cbox a b \ f x = 0" shows "(f has_integral 0) (cbox a b)" unfolding has_integral_cbox using eventually_division_filter_tagged_division[of "cbox a b"] assms by (subst tendsto_cong[where g="\_. 0"]) (auto elim!: eventually_mono intro!: sum.neutral simp: tag_in_interval) lemma has_integral_is_0: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes "\x. x \ S \ f x = 0" shows "(f has_integral 0) S" proof (cases "(\a b. S = cbox a b)") case True with assms has_integral_is_0_cbox show ?thesis by blast next case False have *: "(\x. if x \ S then f x else 0) = (\x. 0)" by (auto simp add: assms) show ?thesis using has_integral_is_0_cbox False by (subst has_integral_alt) (force simp add: *) qed lemma has_integral_0[simp]: "((\x::'n::euclidean_space. 0) has_integral 0) S" by (rule has_integral_is_0) auto lemma has_integral_0_eq[simp]: "((\x. 0) has_integral i) S \ i = 0" using has_integral_unique[OF has_integral_0] by auto lemma has_integral_linear_cbox: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes f: "(f has_integral y) (cbox a b)" and h: "bounded_linear h" shows "((h \ f) has_integral (h y)) (cbox a b)" proof - interpret bounded_linear h using h . show ?thesis unfolding has_integral_cbox using tendsto [OF f [unfolded has_integral_cbox]] by (simp add: sum scaleR split_beta') qed lemma has_integral_linear: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes f: "(f has_integral y) S" and h: "bounded_linear h" shows "((h \ f) has_integral (h y)) S" proof (cases "(\a b. S = cbox a b)") case True with f h has_integral_linear_cbox show ?thesis by blast next case False interpret bounded_linear h using h . from pos_bounded obtain B where B: "0 < B" "\x. norm (h x) \ norm x * B" by blast let ?S = "\f x. if x \ S then f x else 0" show ?thesis proof (subst has_integral_alt, clarsimp simp: False) fix e :: real assume e: "e > 0" have *: "0 < e/B" using e B(1) by simp obtain M where M: "M > 0" "\a b. ball 0 M \ cbox a b \ \z. (?S f has_integral z) (cbox a b) \ norm (z - y) < e/B" using has_integral_altD[OF f False *] by blast show "\B>0. \a b. ball 0 B \ cbox a b \ (\z. (?S(h \ f) has_integral z) (cbox a b) \ norm (z - h y) < e)" proof (rule exI, intro allI conjI impI) show "M > 0" using M by metis next fix a b::'n assume sb: "ball 0 M \ cbox a b" obtain z where z: "(?S f has_integral z) (cbox a b)" "norm (z - y) < e/B" using M(2)[OF sb] by blast have *: "?S(h \ f) = h \ ?S f" using zero by auto show "\z. (?S(h \ f) has_integral z) (cbox a b) \ norm (z - h y) < e" proof (intro exI conjI) show "(?S(h \ f) has_integral h z) (cbox a b)" by (simp add: * has_integral_linear_cbox[OF z(1) h]) show "norm (h z - h y) < e" by (metis B diff le_less_trans pos_less_divide_eq z(2)) qed qed qed qed lemma has_integral_scaleR_left: "(f has_integral y) S \ ((\x. f x *\<^sub>R c) has_integral (y *\<^sub>R c)) S" using has_integral_linear[OF _ bounded_linear_scaleR_left] by (simp add: comp_def) lemma integrable_on_scaleR_left: assumes "f integrable_on A" shows "(\x. f x *\<^sub>R y) integrable_on A" using assms has_integral_scaleR_left unfolding integrable_on_def by blast lemma has_integral_mult_left: fixes c :: "_ :: real_normed_algebra" shows "(f has_integral y) S \ ((\x. f x * c) has_integral (y * c)) S" using has_integral_linear[OF _ bounded_linear_mult_left] by (simp add: comp_def) lemma integrable_on_mult_left: fixes c :: "'a :: real_normed_algebra" assumes "f integrable_on A" shows "(\x. f x * c) integrable_on A" using assms has_integral_mult_left by blast lemma has_integral_divide: fixes c :: "_ :: real_normed_div_algebra" shows "(f has_integral y) S \ ((\x. f x / c) has_integral (y / c)) S" unfolding divide_inverse by (simp add: has_integral_mult_left) lemma integrable_on_divide: fixes c :: "'a :: real_normed_div_algebra" assumes "f integrable_on A" shows "(\x. f x / c) integrable_on A" using assms has_integral_divide by blast text\The case analysis eliminates the condition \<^term>\f integrable_on S\ at the cost of the type class constraint \division_ring\\ corollary integral_mult_left [simp]: fixes c:: "'a::{real_normed_algebra,division_ring}" shows "integral S (\x. f x * c) = integral S f * c" proof (cases "f integrable_on S \ c = 0") case True then show ?thesis by (force intro: has_integral_mult_left) next case False then have "\ (\x. f x * c) integrable_on S" using has_integral_mult_left [of "(\x. f x * c)" _ S "inverse c"] by (auto simp add: mult.assoc) with False show ?thesis by (simp add: not_integrable_integral) qed corollary integral_mult_right [simp]: fixes c:: "'a::{real_normed_field}" shows "integral S (\x. c * f x) = c * integral S f" by (simp add: mult.commute [of c]) corollary integral_divide [simp]: fixes z :: "'a::real_normed_field" shows "integral S (\x. f x / z) = integral S (\x. f x) / z" using integral_mult_left [of S f "inverse z"] by (simp add: divide_inverse_commute) lemma has_integral_mult_right: fixes c :: "'a :: real_normed_algebra" shows "(f has_integral y) A \ ((\x. c * f x) has_integral (c * y)) A" using has_integral_linear[OF _ bounded_linear_mult_right] by (simp add: comp_def) lemma integrable_on_mult_right: fixes c :: "'a :: real_normed_algebra" assumes "f integrable_on A" shows "(\x. c * f x) integrable_on A" using assms has_integral_mult_right by blast lemma integrable_on_mult_right_iff [simp]: fixes c :: "'a :: real_normed_field" assumes "c \ 0" shows "(\x. c * f x) integrable_on A \ f integrable_on A" using integrable_on_mult_right[of f A c] integrable_on_mult_right[of "\x. c * f x" A "inverse c"] assms by (auto simp: field_simps) lemma integrable_on_mult_left_iff [simp]: fixes c :: "'a :: real_normed_field" assumes "c \ 0" shows "(\x. f x * c) integrable_on A \ f integrable_on A" using integrable_on_mult_right_iff[OF assms, of f A] by (simp add: mult.commute) lemma integrable_on_div_iff [simp]: fixes c :: "'a :: real_normed_field" assumes "c \ 0" shows "(\x. f x / c) integrable_on A \ f integrable_on A" using integrable_on_mult_right_iff[of "inverse c" f A] assms by (simp add: field_simps) lemma has_integral_cmul: "(f has_integral k) S \ ((\x. c *\<^sub>R f x) has_integral (c *\<^sub>R k)) S" unfolding o_def[symmetric] by (metis has_integral_linear bounded_linear_scaleR_right) lemma has_integral_cmult_real: fixes c :: real assumes "c \ 0 \ (f has_integral x) A" shows "((\x. c * f x) has_integral c * x) A" by (metis assms has_integral_is_0 has_integral_mult_right lambda_zero) lemma has_integral_neg: "(f has_integral k) S \ ((\x. -(f x)) has_integral -k) S" by (drule_tac c="-1" in has_integral_cmul) auto lemma has_integral_neg_iff: "((\x. - f x) has_integral k) S \ (f has_integral - k) S" using has_integral_neg[of f "- k"] has_integral_neg[of "\x. - f x" k] by auto lemma has_integral_add_cbox: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes "(f has_integral k) (cbox a b)" "(g has_integral l) (cbox a b)" shows "((\x. f x + g x) has_integral (k + l)) (cbox a b)" using assms unfolding has_integral_cbox by (simp add: split_beta' scaleR_add_right sum.distrib[abs_def] tendsto_add) lemma has_integral_add: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes f: "(f has_integral k) S" and g: "(g has_integral l) S" shows "((\x. f x + g x) has_integral (k + l)) S" proof (cases "\a b. S = cbox a b") case True with has_integral_add_cbox assms show ?thesis by blast next let ?S = "\f x. if x \ S then f x else 0" case False then show ?thesis proof (subst has_integral_alt, clarsimp, goal_cases) case (1 e) then have e2: "e/2 > 0" by auto obtain Bf where "0 < Bf" and Bf: "\a b. ball 0 Bf \ cbox a b \ \z. (?S f has_integral z) (cbox a b) \ norm (z - k) < e/2" using has_integral_altD[OF f False e2] by blast obtain Bg where "0 < Bg" and Bg: "\a b. ball 0 Bg \ (cbox a b) \ \z. (?S g has_integral z) (cbox a b) \ norm (z - l) < e/2" using has_integral_altD[OF g False e2] by blast show ?case proof (rule_tac x="max Bf Bg" in exI, clarsimp simp add: max.strict_coboundedI1 \0 < Bf\) fix a b assume "ball 0 (max Bf Bg) \ cbox a (b::'n)" then have fs: "ball 0 Bf \ cbox a (b::'n)" and gs: "ball 0 Bg \ cbox a (b::'n)" by auto obtain w where w: "(?S f has_integral w) (cbox a b)" "norm (w - k) < e/2" using Bf[OF fs] by blast obtain z where z: "(?S g has_integral z) (cbox a b)" "norm (z - l) < e/2" using Bg[OF gs] by blast have *: "\x. (if x \ S then f x + g x else 0) = (?S f x) + (?S g x)" by auto show "\z. (?S(\x. f x + g x) has_integral z) (cbox a b) \ norm (z - (k + l)) < e" proof (intro exI conjI) show "(?S(\x. f x + g x) has_integral (w + z)) (cbox a b)" by (simp add: has_integral_add_cbox[OF w(1) z(1), unfolded *[symmetric]]) show "norm (w + z - (k + l)) < e" by (metis dist_norm dist_triangle_add_half w(2) z(2)) qed qed qed qed lemma has_integral_diff: "(f has_integral k) S \ (g has_integral l) S \ ((\x. f x - g x) has_integral (k - l)) S" using has_integral_add[OF _ has_integral_neg, of f k S g l] by (auto simp: algebra_simps) lemma integral_0 [simp]: "integral S (\x::'n::euclidean_space. 0::'m::real_normed_vector) = 0" by auto lemma integral_add: "f integrable_on S \ g integrable_on S \ integral S (\x. f x + g x) = integral S f + integral S g" by (rule integral_unique) (metis integrable_integral has_integral_add) lemma integral_cmul [simp]: "integral S (\x. c *\<^sub>R f x) = c *\<^sub>R integral S f" proof (cases "f integrable_on S \ c = 0") case True with has_integral_cmul integrable_integral show ?thesis by fastforce next case False then have "\ (\x. c *\<^sub>R f x) integrable_on S" using has_integral_cmul [of "(\x. c *\<^sub>R f x)" _ S "inverse c"] by auto with False show ?thesis by (simp add: not_integrable_integral) qed lemma integral_mult: fixes K::real shows "f integrable_on X \ K * integral X f = integral X (\x. K * f x)" by simp lemma integral_neg [simp]: "integral S (\x. - f x) = - integral S f" by (metis eq_integralD equation_minus_iff has_integral_iff has_integral_neg_iff neg_equal_0_iff_equal) lemma integral_diff: "f integrable_on S \ g integrable_on S \ integral S (\x. f x - g x) = integral S f - integral S g" by (rule integral_unique) (metis integrable_integral has_integral_diff) lemma integrable_0: "(\x. 0) integrable_on S" unfolding integrable_on_def using has_integral_0 by auto lemma integrable_add: "f integrable_on S \ g integrable_on S \ (\x. f x + g x) integrable_on S" unfolding integrable_on_def by(auto intro: has_integral_add) lemma integrable_cmul: "f integrable_on S \ (\x. c *\<^sub>R f(x)) integrable_on S" unfolding integrable_on_def by(auto intro: has_integral_cmul) lemma integrable_on_scaleR_iff [simp]: fixes c :: real assumes "c \ 0" shows "(\x. c *\<^sub>R f x) integrable_on S \ f integrable_on S" using integrable_cmul[of "\x. c *\<^sub>R f x" S "1 / c"] integrable_cmul[of f S c] \c \ 0\ by auto lemma integrable_on_cmult_iff [simp]: fixes c :: real assumes "c \ 0" shows "(\x. c * f x) integrable_on S \ f integrable_on S" using integrable_on_scaleR_iff [of c f] assms by simp lemma integrable_on_cmult_left: assumes "f integrable_on S" shows "(\x. of_real c * f x) integrable_on S" using integrable_cmul[of f S "of_real c"] assms by (simp add: scaleR_conv_of_real) lemma integrable_neg: "f integrable_on S \ (\x. -f(x)) integrable_on S" unfolding integrable_on_def by(auto intro: has_integral_neg) lemma integrable_neg_iff: "(\x. -f(x)) integrable_on S \ f integrable_on S" using integrable_neg by fastforce lemma integrable_diff: "f integrable_on S \ g integrable_on S \ (\x. f x - g x) integrable_on S" unfolding integrable_on_def by(auto intro: has_integral_diff) lemma integrable_linear: "f integrable_on S \ bounded_linear h \ (h \ f) integrable_on S" unfolding integrable_on_def by(auto intro: has_integral_linear) lemma integral_linear: "f integrable_on S \ bounded_linear h \ integral S (h \ f) = h (integral S f)" by (meson has_integral_iff has_integral_linear) lemma integrable_on_cnj_iff: "(\x. cnj (f x)) integrable_on A \ f integrable_on A" using integrable_linear[OF _ bounded_linear_cnj, of f A] integrable_linear[OF _ bounded_linear_cnj, of "cnj \ f" A] by (auto simp: o_def) lemma integral_cnj: "cnj (integral A f) = integral A (\x. cnj (f x))" by (cases "f integrable_on A") (simp_all add: integral_linear[OF _ bounded_linear_cnj, symmetric] o_def integrable_on_cnj_iff not_integrable_integral) +lemma has_integral_cnj: "(cnj \ f has_integral (cnj I)) A = (f has_integral I) A" + unfolding has_integral_iff comp_def + by (metis integral_cnj complex_cnj_cancel_iff integrable_on_cnj_iff) + lemma integral_component_eq[simp]: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "f integrable_on S" shows "integral S (\x. f x \ k) = integral S f \ k" unfolding integral_linear[OF assms(1) bounded_linear_inner_left,unfolded o_def] .. lemma has_integral_sum: assumes "finite T" and "\a. a \ T \ ((f a) has_integral (i a)) S" shows "((\x. sum (\a. f a x) T) has_integral (sum i T)) S" using \finite T\ subset_refl[of T] by (induct rule: finite_subset_induct) (use assms in \auto simp: has_integral_add\) lemma integral_sum: "\finite I; \a. a \ I \ f a integrable_on S\ \ integral S (\x. \a\I. f a x) = (\a\I. integral S (f a))" by (simp add: has_integral_sum integrable_integral integral_unique) lemma integrable_sum: "\finite I; \a. a \ I \ f a integrable_on S\ \ (\x. \a\I. f a x) integrable_on S" unfolding integrable_on_def using has_integral_sum[of I] by metis lemma has_integral_eq: assumes "\x. x \ s \ f x = g x" and f: "(f has_integral k) s" shows "(g has_integral k) s" using has_integral_diff[OF f, of "\x. f x - g x" 0] using has_integral_is_0[of s "\x. f x - g x"] using assms by auto lemma integrable_eq: "\f integrable_on s; \x. x \ s \ f x = g x\ \ g integrable_on s" unfolding integrable_on_def using has_integral_eq[of s f g] has_integral_eq by blast lemma has_integral_cong: assumes "\x. x \ s \ f x = g x" shows "(f has_integral i) s = (g has_integral i) s" by (metis assms has_integral_eq) lemma integrable_cong: assumes "\x. x \ A \ f x = g x" shows "f integrable_on A \ g integrable_on A" using has_integral_cong [OF assms] by fast lemma integral_cong: assumes "\x. x \ s \ f x = g x" shows "integral s f = integral s g" unfolding integral_def by (metis (full_types, opaque_lifting) assms has_integral_cong integrable_eq) lemma integrable_on_cmult_left_iff [simp]: assumes "c \ 0" shows "(\x. of_real c * f x) integrable_on s \ f integrable_on s" (is "?lhs = ?rhs") proof assume ?lhs then have "(\x. of_real (1 / c) * (of_real c * f x)) integrable_on s" using integrable_cmul[of "\x. of_real c * f x" s "1 / of_real c"] by (simp add: scaleR_conv_of_real) then have "(\x. (of_real (1 / c) * of_real c * f x)) integrable_on s" by (simp add: algebra_simps) with \c \ 0\ show ?rhs by (metis (no_types, lifting) integrable_eq mult.left_neutral nonzero_divide_eq_eq of_real_1 of_real_mult) qed (blast intro: integrable_on_cmult_left) lemma integrable_on_cmult_right: fixes f :: "_ \ 'b :: {comm_ring,real_algebra_1,real_normed_vector}" assumes "f integrable_on s" shows "(\x. f x * of_real c) integrable_on s" using integrable_on_cmult_left [OF assms] by (simp add: mult.commute) lemma integrable_on_cmult_right_iff [simp]: fixes f :: "_ \ 'b :: {comm_ring,real_algebra_1,real_normed_vector}" assumes "c \ 0" shows "(\x. f x * of_real c) integrable_on s \ f integrable_on s" using integrable_on_cmult_left_iff [OF assms] by (simp add: mult.commute) lemma integrable_on_cdivide_iff [simp]: fixes f :: "_ \ 'b :: real_normed_field" assumes "c \ 0" shows "(\x. f x / of_real c) integrable_on s \ f integrable_on s" by (simp add: divide_inverse assms flip: of_real_inverse) lemma has_integral_null [intro]: "content(cbox a b) = 0 \ (f has_integral 0) (cbox a b)" unfolding has_integral_cbox using eventually_division_filter_tagged_division[of "cbox a b"] by (subst tendsto_cong[where g="\_. 0"]) (auto elim: eventually_mono intro: sum_content_null) lemma has_integral_null_real [intro]: "content {a..b::real} = 0 \ (f has_integral 0) {a..b}" by (metis box_real(2) has_integral_null) lemma has_integral_null_eq[simp]: "content (cbox a b) = 0 \ (f has_integral i) (cbox a b) \ i = 0" by (auto simp add: has_integral_null dest!: integral_unique) lemma integral_null [simp]: "content (cbox a b) = 0 \ integral (cbox a b) f = 0" by (metis has_integral_null integral_unique) lemma integrable_on_null [intro]: "content (cbox a b) = 0 \ f integrable_on (cbox a b)" by (simp add: has_integral_integrable) lemma has_integral_empty[intro]: "(f has_integral 0) {}" by (meson ex_in_conv has_integral_is_0) lemma has_integral_empty_eq[simp]: "(f has_integral i) {} \ i = 0" by (auto simp add: has_integral_empty has_integral_unique) lemma integrable_on_empty[intro]: "f integrable_on {}" unfolding integrable_on_def by auto lemma integral_empty[simp]: "integral {} f = 0" by blast lemma has_integral_refl[intro]: fixes a :: "'a::euclidean_space" shows "(f has_integral 0) (cbox a a)" and "(f has_integral 0) {a}" proof - show "(f has_integral 0) (cbox a a)" by (rule has_integral_null) simp then show "(f has_integral 0) {a}" by simp qed lemma integrable_on_refl[intro]: "f integrable_on cbox a a" unfolding integrable_on_def by auto lemma integral_refl [simp]: "integral (cbox a a) f = 0" by auto lemma integral_singleton [simp]: "integral {a} f = 0" by auto lemma integral_blinfun_apply: assumes "f integrable_on s" shows "integral s (\x. blinfun_apply h (f x)) = blinfun_apply h (integral s f)" by (subst integral_linear[symmetric, OF assms blinfun.bounded_linear_right]) (simp add: o_def) lemma blinfun_apply_integral: assumes "f integrable_on s" shows "blinfun_apply (integral s f) x = integral s (\y. blinfun_apply (f y) x)" by (metis (no_types, lifting) assms blinfun.prod_left.rep_eq integral_blinfun_apply integral_cong) lemma has_integral_componentwise_iff: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" shows "(f has_integral y) A \ (\b\Basis. ((\x. f x \ b) has_integral (y \ b)) A)" proof safe fix b :: 'b assume "(f has_integral y) A" from has_integral_linear[OF this(1) bounded_linear_inner_left, of b] show "((\x. f x \ b) has_integral (y \ b)) A" by (simp add: o_def) next assume "(\b\Basis. ((\x. f x \ b) has_integral (y \ b)) A)" hence "\b\Basis. (((\x. x *\<^sub>R b) \ (\x. f x \ b)) has_integral ((y \ b) *\<^sub>R b)) A" by (intro ballI has_integral_linear) (simp_all add: bounded_linear_scaleR_left) hence "((\x. \b\Basis. (f x \ b) *\<^sub>R b) has_integral (\b\Basis. (y \ b) *\<^sub>R b)) A" by (intro has_integral_sum) (simp_all add: o_def) thus "(f has_integral y) A" by (simp add: euclidean_representation) qed lemma has_integral_componentwise: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" shows "(\b. b \ Basis \ ((\x. f x \ b) has_integral (y \ b)) A) \ (f has_integral y) A" by (subst has_integral_componentwise_iff) blast lemma integrable_componentwise_iff: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" shows "f integrable_on A \ (\b\Basis. (\x. f x \ b) integrable_on A)" proof assume "f integrable_on A" then obtain y where "(f has_integral y) A" by (auto simp: integrable_on_def) hence "(\b\Basis. ((\x. f x \ b) has_integral (y \ b)) A)" by (subst (asm) has_integral_componentwise_iff) thus "(\b\Basis. (\x. f x \ b) integrable_on A)" by (auto simp: integrable_on_def) next assume "(\b\Basis. (\x. f x \ b) integrable_on A)" then obtain y where "\b\Basis. ((\x. f x \ b) has_integral y b) A" unfolding integrable_on_def by (subst (asm) bchoice_iff) blast hence "\b\Basis. (((\x. x *\<^sub>R b) \ (\x. f x \ b)) has_integral (y b *\<^sub>R b)) A" by (intro ballI has_integral_linear) (simp_all add: bounded_linear_scaleR_left) hence "((\x. \b\Basis. (f x \ b) *\<^sub>R b) has_integral (\b\Basis. y b *\<^sub>R b)) A" by (intro has_integral_sum) (simp_all add: o_def) thus "f integrable_on A" by (auto simp: integrable_on_def o_def euclidean_representation) qed lemma integrable_componentwise: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" shows "(\b. b \ Basis \ (\x. f x \ b) integrable_on A) \ f integrable_on A" by (subst integrable_componentwise_iff) blast lemma integral_componentwise: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes "f integrable_on A" shows "integral A f = (\b\Basis. integral A (\x. (f x \ b) *\<^sub>R b))" proof - from assms have integrable: "\b\Basis. (\x. x *\<^sub>R b) \ (\x. (f x \ b)) integrable_on A" by (subst (asm) integrable_componentwise_iff, intro integrable_linear ballI) (simp_all add: bounded_linear_scaleR_left) have "integral A f = integral A (\x. \b\Basis. (f x \ b) *\<^sub>R b)" by (simp add: euclidean_representation) also from integrable have "\ = (\a\Basis. integral A (\x. (f x \ a) *\<^sub>R a))" by (subst integral_sum) (simp_all add: o_def) finally show ?thesis . qed lemma integrable_component: "f integrable_on A \ (\x. f x \ (y :: 'b :: euclidean_space)) integrable_on A" by (drule integrable_linear[OF _ bounded_linear_inner_left[of y]]) (simp add: o_def) subsection \Cauchy-type criterion for integrability\ proposition integrable_Cauchy: fixes f :: "'n::euclidean_space \ 'a::{real_normed_vector,complete_space}" shows "f integrable_on cbox a b \ (\e>0. \\. gauge \ \ (\\1 \2. \1 tagged_division_of (cbox a b) \ \ fine \1 \ \2 tagged_division_of (cbox a b) \ \ fine \2 \ norm ((\(x,K)\\1. content K *\<^sub>R f x) - (\(x,K)\\2. content K *\<^sub>R f x)) < e))" (is "?l = (\e>0. \\. ?P e \)") proof (intro iffI allI impI) assume ?l then obtain y where y: "\e. e > 0 \ \\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - y) < e)" by (auto simp: integrable_on_def has_integral) show "\\. ?P e \" if "e > 0" for e proof - have "e/2 > 0" using that by auto with y obtain \ where "gauge \" and \: "\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x,K)\\. content K *\<^sub>R f x) - y) < e/2" by meson show ?thesis apply (rule_tac x=\ in exI, clarsimp simp: \gauge \\) by (blast intro!: \ dist_triangle_half_l[where y=y,unfolded dist_norm]) qed next assume "\e>0. \\. ?P e \" then have "\n::nat. \\. ?P (1 / (n + 1)) \" by auto then obtain \ :: "nat \ 'n \ 'n set" where \: "\m. gauge (\ m)" "\m \1 \2. \\1 tagged_division_of cbox a b; \ m fine \1; \2 tagged_division_of cbox a b; \ m fine \2\ \ norm ((\(x,K) \ \1. content K *\<^sub>R f x) - (\(x,K) \ \2. content K *\<^sub>R f x)) < 1 / (m + 1)" by metis have "gauge (\x. \{\ i x |i. i \ {0..n}})" for n using \ by (intro gauge_Inter) auto then have "\n. \p. p tagged_division_of (cbox a b) \ (\x. \{\ i x |i. i \ {0..n}}) fine p" by (meson fine_division_exists) then obtain p where p: "\z. p z tagged_division_of cbox a b" "\z. (\x. \{\ i x |i. i \ {0..z}}) fine p z" by meson have dp: "\i n. i\n \ \ i fine p n" using p unfolding fine_Inter using atLeastAtMost_iff by blast have "Cauchy (\n. sum (\(x,K). content K *\<^sub>R (f x)) (p n))" proof (rule CauchyI) fix e::real assume "0 < e" then obtain N where "N \ 0" and N: "inverse (real N) < e" using real_arch_inverse[of e] by blast show "\M. \m\M. \n\M. norm ((\(x,K) \ p m. content K *\<^sub>R f x) - (\(x,K) \ p n. content K *\<^sub>R f x)) < e" proof (intro exI allI impI) fix m n assume mn: "N \ m" "N \ n" have "norm ((\(x,K) \ p m. content K *\<^sub>R f x) - (\(x,K) \ p n. content K *\<^sub>R f x)) < 1 / (real N + 1)" by (simp add: p(1) dp mn \) also have "... < e" using N \N \ 0\ \0 < e\ by (auto simp: field_simps) finally show "norm ((\(x,K) \ p m. content K *\<^sub>R f x) - (\(x,K) \ p n. content K *\<^sub>R f x)) < e" . qed qed then obtain y where y: "\no. \n\no. norm ((\(x,K) \ p n. content K *\<^sub>R f x) - y) < r" if "r > 0" for r by (auto simp: convergent_eq_Cauchy[symmetric] dest: LIMSEQ_D) show ?l unfolding integrable_on_def has_integral proof (rule_tac x=y in exI, clarify) fix e :: real assume "e>0" then have e2: "e/2 > 0" by auto then obtain N1::nat where N1: "N1 \ 0" "inverse (real N1) < e/2" using real_arch_inverse by blast obtain N2::nat where N2: "\n. n \ N2 \ norm ((\(x,K) \ p n. content K *\<^sub>R f x) - y) < e/2" using y[OF e2] by metis show "\\. gauge \ \ (\\. \ tagged_division_of (cbox a b) \ \ fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - y) < e)" proof (intro exI conjI allI impI) show "gauge (\ (N1+N2))" using \ by auto show "norm ((\(x,K) \ q. content K *\<^sub>R f x) - y) < e" if "q tagged_division_of cbox a b \ \ (N1+N2) fine q" for q proof (rule norm_triangle_half_r) have "norm ((\(x,K) \ p (N1+N2). content K *\<^sub>R f x) - (\(x,K) \ q. content K *\<^sub>R f x)) < 1 / (real (N1+N2) + 1)" by (rule \; simp add: dp p that) also have "... < e/2" using N1 \0 < e\ by (auto simp: field_simps intro: less_le_trans) finally show "norm ((\(x,K) \ p (N1+N2). content K *\<^sub>R f x) - (\(x,K) \ q. content K *\<^sub>R f x)) < e/2" . show "norm ((\(x,K) \ p (N1+N2). content K *\<^sub>R f x) - y) < e/2" using N2 le_add_same_cancel2 by blast qed qed qed qed subsection \Additivity of integral on abutting intervals\ lemma tagged_division_split_left_inj_content: assumes \: "\ tagged_division_of S" and "(x1, K1) \ \" "(x2, K2) \ \" "K1 \ K2" "K1 \ {x. x\k \ c} = K2 \ {x. x\k \ c}" "k \ Basis" shows "content (K1 \ {x. x\k \ c}) = 0" proof - from tagged_division_ofD(4)[OF \ \(x1, K1) \ \\] obtain a b where K1: "K1 = cbox a b" by auto then have "interior (K1 \ {x. x \ k \ c}) = {}" by (metis tagged_division_split_left_inj assms) then show ?thesis unfolding K1 interval_split[OF \k \ Basis\] by (auto simp: content_eq_0_interior) qed lemma tagged_division_split_right_inj_content: assumes \: "\ tagged_division_of S" and "(x1, K1) \ \" "(x2, K2) \ \" "K1 \ K2" "K1 \ {x. x\k \ c} = K2 \ {x. x\k \ c}" "k \ Basis" shows "content (K1 \ {x. x\k \ c}) = 0" proof - from tagged_division_ofD(4)[OF \ \(x1, K1) \ \\] obtain a b where K1: "K1 = cbox a b" by auto then have "interior (K1 \ {x. c \ x \ k}) = {}" by (metis tagged_division_split_right_inj assms) then show ?thesis unfolding K1 interval_split[OF \k \ Basis\] by (auto simp: content_eq_0_interior) qed proposition has_integral_split: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes fi: "(f has_integral i) (cbox a b \ {x. x\k \ c})" and fj: "(f has_integral j) (cbox a b \ {x. x\k \ c})" and k: "k \ Basis" shows "(f has_integral (i + j)) (cbox a b)" unfolding has_integral proof clarify fix e::real assume "0 < e" then have e: "e/2 > 0" by auto obtain \1 where \1: "gauge \1" and \1norm: "\\. \\ tagged_division_of cbox a b \ {x. x \ k \ c}; \1 fine \\ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - i) < e/2" apply (rule has_integralD[OF fi[unfolded interval_split[OF k]] e]) apply (simp add: interval_split[symmetric] k) done obtain \2 where \2: "gauge \2" and \2norm: "\\. \\ tagged_division_of cbox a b \ {x. c \ x \ k}; \2 fine \\ \ norm ((\(x, k) \ \. content k *\<^sub>R f x) - j) < e/2" apply (rule has_integralD[OF fj[unfolded interval_split[OF k]] e]) apply (simp add: interval_split[symmetric] k) done let ?\ = "\x. if x\k = c then (\1 x \ \2 x) else ball x \x\k - c\ \ \1 x \ \2 x" have "gauge ?\" using \1 \2 unfolding gauge_def by auto then show "\\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x, k)\\. content k *\<^sub>R f x) - (i + j)) < e)" proof (rule_tac x="?\" in exI, safe) fix p assume p: "p tagged_division_of (cbox a b)" and "?\ fine p" have ab_eqp: "cbox a b = \{K. \x. (x, K) \ p}" using p by blast have xk_le_c: "x\k \ c" if as: "(x,K) \ p" and K: "K \ {x. x\k \ c} \ {}" for x K proof (rule ccontr) assume **: "\ x \ k \ c" then have "K \ ball x \x \ k - c\" using \?\ fine p\ as by (fastforce simp: not_le algebra_simps) with K obtain y where y: "y \ ball x \x \ k - c\" "y\k \ c" by blast then have "\x \ k - y \ k\ < \x \ k - c\" using Basis_le_norm[OF k, of "x - y"] by (auto simp add: dist_norm inner_diff_left intro: le_less_trans) with y show False using ** by (auto simp add: field_simps) qed have xk_ge_c: "x\k \ c" if as: "(x,K) \ p" and K: "K \ {x. x\k \ c} \ {}" for x K proof (rule ccontr) assume **: "\ x \ k \ c" then have "K \ ball x \x \ k - c\" using \?\ fine p\ as by (fastforce simp: not_le algebra_simps) with K obtain y where y: "y \ ball x \x \ k - c\" "y\k \ c" by blast then have "\x \ k - y \ k\ < \x \ k - c\" using Basis_le_norm[OF k, of "x - y"] by (auto simp add: dist_norm inner_diff_left intro: le_less_trans) with y show False using ** by (auto simp add: field_simps) qed have fin_finite: "finite {(x,f K) | x K. (x,K) \ s \ P x K}" if "finite s" for s and f :: "'a set \ 'a set" and P :: "'a \ 'a set \ bool" proof - from that have "finite ((\(x,K). (x, f K)) ` s)" by auto then show ?thesis by (rule rev_finite_subset) auto qed { fix \ :: "'a set \ 'a set" fix i :: "'a \ 'a set" assume "i \ (\(x, k). (x, \ k)) ` p - {(x, \ k) |x k. (x, k) \ p \ \ k \ {}}" then obtain x K where xk: "i = (x, \ K)" "(x,K) \ p" "(x, \ K) \ {(x, \ K) |x K. (x,K) \ p \ \ K \ {}}" by auto have "content (\ K) = 0" using xk using content_empty by auto then have "(\(x,K). content K *\<^sub>R f x) i = 0" unfolding xk split_conv by auto } note [simp] = this have "finite p" using p by blast let ?M1 = "{(x, K \ {x. x\k \ c}) |x K. (x,K) \ p \ K \ {x. x\k \ c} \ {}}" have \1_fine: "\1 fine ?M1" using \?\ fine p\ by (fastforce simp: fine_def split: if_split_asm) have "norm ((\(x, k)\?M1. content k *\<^sub>R f x) - i) < e/2" proof (rule \1norm [OF tagged_division_ofI \1_fine]) show "finite ?M1" by (rule fin_finite) (use p in blast) show "\{k. \x. (x, k) \ ?M1} = cbox a b \ {x. x\k \ c}" by (auto simp: ab_eqp) fix x L assume xL: "(x, L) \ ?M1" then obtain x' L' where xL': "x = x'" "L = L' \ {x. x \ k \ c}" "(x', L') \ p" "L' \ {x. x \ k \ c} \ {}" by blast then obtain a' b' where ab': "L' = cbox a' b'" using p by blast show "x \ L" "L \ cbox a b \ {x. x \ k \ c}" using p xk_le_c xL' by auto show "\a b. L = cbox a b" using p xL' ab' by (auto simp add: interval_split[OF k,where c=c]) fix y R assume yR: "(y, R) \ ?M1" then obtain y' R' where yR': "y = y'" "R = R' \ {x. x \ k \ c}" "(y', R') \ p" "R' \ {x. x \ k \ c} \ {}" by blast assume as: "(x, L) \ (y, R)" show "interior L \ interior R = {}" proof (cases "L' = R' \ x' = y'") case False have "interior R' = {}" by (metis (no_types) False Pair_inject inf.idem tagged_division_ofD(5) [OF p] xL'(3) yR'(3)) then show ?thesis using yR' by simp next case True then have "L' \ R'" using as unfolding xL' yR' by auto have "interior L' \ interior R' = {}" by (metis (no_types) Pair_inject \L' \ R'\ p tagged_division_ofD(5) xL'(3) yR'(3)) then show ?thesis using xL'(2) yR'(2) by auto qed qed moreover let ?M2 = "{(x,K \ {x. x\k \ c}) |x K. (x,K) \ p \ K \ {x. x\k \ c} \ {}}" have \2_fine: "\2 fine ?M2" using \?\ fine p\ by (fastforce simp: fine_def split: if_split_asm) have "norm ((\(x, k)\?M2. content k *\<^sub>R f x) - j) < e/2" proof (rule \2norm [OF tagged_division_ofI \2_fine]) show "finite ?M2" by (rule fin_finite) (use p in blast) show "\{k. \x. (x, k) \ ?M2} = cbox a b \ {x. x\k \ c}" by (auto simp: ab_eqp) fix x L assume xL: "(x, L) \ ?M2" then obtain x' L' where xL': "x = x'" "L = L' \ {x. x \ k \ c}" "(x', L') \ p" "L' \ {x. x \ k \ c} \ {}" by blast then obtain a' b' where ab': "L' = cbox a' b'" using p by blast show "x \ L" "L \ cbox a b \ {x. x \ k \ c}" using p xk_ge_c xL' by auto show "\a b. L = cbox a b" using p xL' ab' by (auto simp add: interval_split[OF k,where c=c]) fix y R assume yR: "(y, R) \ ?M2" then obtain y' R' where yR': "y = y'" "R = R' \ {x. x \ k \ c}" "(y', R') \ p" "R' \ {x. x \ k \ c} \ {}" by blast assume as: "(x, L) \ (y, R)" show "interior L \ interior R = {}" proof (cases "L' = R' \ x' = y'") case False have "interior R' = {}" by (metis (no_types) False Pair_inject inf.idem tagged_division_ofD(5) [OF p] xL'(3) yR'(3)) then show ?thesis using yR' by simp next case True then have "L' \ R'" using as unfolding xL' yR' by auto have "interior L' \ interior R' = {}" by (metis (no_types) Pair_inject \L' \ R'\ p tagged_division_ofD(5) xL'(3) yR'(3)) then show ?thesis using xL'(2) yR'(2) by auto qed qed ultimately have "norm (((\(x,K) \ ?M1. content K *\<^sub>R f x) - i) + ((\(x,K) \ ?M2. content K *\<^sub>R f x) - j)) < e/2 + e/2" using norm_add_less by blast moreover have "((\(x,K) \ ?M1. content K *\<^sub>R f x) - i) + ((\(x,K) \ ?M2. content K *\<^sub>R f x) - j) = (\(x, ka)\p. content ka *\<^sub>R f x) - (i + j)" proof - have eq0: "\x y. x = (0::real) \ x *\<^sub>R (y::'b) = 0" by auto have cont_eq: "\g. (\(x,l). content l *\<^sub>R f x) \ (\(x,l). (x,g l)) = (\(x,l). content (g l) *\<^sub>R f x)" by auto have *: "\\ :: 'a set \ 'a set. (\(x,K)\{(x, \ K) |x K. (x,K) \ p \ \ K \ {}}. content K *\<^sub>R f x) = (\(x,K)\(\(x,K). (x, \ K)) ` p. content K *\<^sub>R f x)" by (rule sum.mono_neutral_left) (auto simp: \finite p\) have "((\(x, k)\?M1. content k *\<^sub>R f x) - i) + ((\(x, k)\?M2. content k *\<^sub>R f x) - j) = (\(x, k)\?M1. content k *\<^sub>R f x) + (\(x, k)\?M2. content k *\<^sub>R f x) - (i + j)" by auto moreover have "\ = (\(x,K) \ p. content (K \ {x. x \ k \ c}) *\<^sub>R f x) + (\(x,K) \ p. content (K \ {x. c \ x \ k}) *\<^sub>R f x) - (i + j)" unfolding * apply (subst (1 2) sum.reindex_nontrivial) apply (auto intro!: k p eq0 tagged_division_split_left_inj_content tagged_division_split_right_inj_content simp: cont_eq \finite p\) done moreover have "\x. x \ p \ (\(a,B). content (B \ {a. a \ k \ c}) *\<^sub>R f a) x + (\(a,B). content (B \ {a. c \ a \ k}) *\<^sub>R f a) x = (\(a,B). content B *\<^sub>R f a) x" proof clarify fix a B assume "(a, B) \ p" with p obtain u v where uv: "B = cbox u v" by blast then show "content (B \ {x. x \ k \ c}) *\<^sub>R f a + content (B \ {x. c \ x \ k}) *\<^sub>R f a = content B *\<^sub>R f a" by (auto simp: scaleR_left_distrib uv content_split[OF k,of u v c]) qed ultimately show ?thesis by (auto simp: sum.distrib[symmetric]) qed ultimately show "norm ((\(x, k)\p. content k *\<^sub>R f x) - (i + j)) < e" by auto qed qed subsection \A sort of converse, integrability on subintervals\ lemma has_integral_separate_sides: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes f: "(f has_integral i) (cbox a b)" and "e > 0" and k: "k \ Basis" obtains d where "gauge d" "\p1 p2. p1 tagged_division_of (cbox a b \ {x. x\k \ c}) \ d fine p1 \ p2 tagged_division_of (cbox a b \ {x. x\k \ c}) \ d fine p2 \ norm ((sum (\(x,k). content k *\<^sub>R f x) p1 + sum (\(x,k). content k *\<^sub>R f x) p2) - i) < e" proof - obtain \ where d: "gauge \" "\p. \p tagged_division_of cbox a b; \ fine p\ \ norm ((\(x, k)\p. content k *\<^sub>R f x) - i) < e" using has_integralD[OF f \e > 0\] by metis { fix p1 p2 assume tdiv1: "p1 tagged_division_of (cbox a b) \ {x. x \ k \ c}" and "\ fine p1" note p1=tagged_division_ofD[OF this(1)] assume tdiv2: "p2 tagged_division_of (cbox a b) \ {x. c \ x \ k}" and "\ fine p2" note p2=tagged_division_ofD[OF this(1)] note tagged_division_Un_interval[OF tdiv1 tdiv2] note p12 = tagged_division_ofD[OF this] this { fix a b assume ab: "(a, b) \ p1 \ p2" have "(a, b) \ p1" using ab by auto obtain u v where uv: "b = cbox u v" using \(a, b) \ p1\ p1(4) by moura have "b \ {x. x\k = c}" using ab p1(3)[of a b] p2(3)[of a b] by fastforce moreover have "interior {x::'a. x \ k = c} = {}" proof (rule ccontr) assume "\ ?thesis" then obtain x where x: "x \ interior {x::'a. x\k = c}" by auto then obtain \ where "0 < \" and \: "ball x \ \ {x. x \ k = c}" using mem_interior by metis have x: "x\k = c" using x interior_subset by fastforce have *: "\i. i \ Basis \ \(x - (x + (\/2) *\<^sub>R k)) \ i\ = (if i = k then \/2 else 0)" using \0 < \\ k by (auto simp: inner_simps inner_not_same_Basis) have "(\i\Basis. \(x - (x + (\/2 ) *\<^sub>R k)) \ i\) = (\i\Basis. (if i = k then \/2 else 0))" using "*" by (blast intro: sum.cong) also have "\ < \" by (subst sum.delta) (use \0 < \\ in auto) finally have "x + (\/2) *\<^sub>R k \ ball x \" unfolding mem_ball dist_norm by(rule le_less_trans[OF norm_le_l1]) then have "x + (\/2) *\<^sub>R k \ {x. x\k = c}" using \ by auto then show False using \0 < \\ x k by (auto simp: inner_simps) qed ultimately have "content b = 0" unfolding uv content_eq_0_interior using interior_mono by blast then have "content b *\<^sub>R f a = 0" by auto } then have "norm ((\(x, k)\p1. content k *\<^sub>R f x) + (\(x, k)\p2. content k *\<^sub>R f x) - i) = norm ((\(x, k)\p1 \ p2. content k *\<^sub>R f x) - i)" by (subst sum.union_inter_neutral) (auto simp: p1 p2) also have "\ < e" using d(2) p12 by (simp add: fine_Un k \\ fine p1\ \\ fine p2\) finally have "norm ((\(x, k)\p1. content k *\<^sub>R f x) + (\(x, k)\p2. content k *\<^sub>R f x) - i) < e" . } then show ?thesis using d(1) that by auto qed lemma integrable_split [intro]: fixes f :: "'a::euclidean_space \ 'b::{real_normed_vector,complete_space}" assumes f: "f integrable_on cbox a b" and k: "k \ Basis" shows "f integrable_on (cbox a b \ {x. x\k \ c})" (is ?thesis1) and "f integrable_on (cbox a b \ {x. x\k \ c})" (is ?thesis2) proof - obtain y where y: "(f has_integral y) (cbox a b)" using f by blast define a' where "a' = (\i\Basis. (if i = k then max (a\k) c else a\i)*\<^sub>R i)" define b' where "b' = (\i\Basis. (if i = k then min (b\k) c else b\i)*\<^sub>R i)" have "\d. gauge d \ (\p1 p2. p1 tagged_division_of cbox a b \ {x. x \ k \ c} \ d fine p1 \ p2 tagged_division_of cbox a b \ {x. x \ k \ c} \ d fine p2 \ norm ((\(x,K) \ p1. content K *\<^sub>R f x) - (\(x,K) \ p2. content K *\<^sub>R f x)) < e)" if "e > 0" for e proof - have "e/2 > 0" using that by auto with has_integral_separate_sides[OF y this k, of c] obtain d where "gauge d" and d: "\p1 p2. \p1 tagged_division_of cbox a b \ {x. x \ k \ c}; d fine p1; p2 tagged_division_of cbox a b \ {x. c \ x \ k}; d fine p2\ \ norm ((\(x,K)\p1. content K *\<^sub>R f x) + (\(x,K)\p2. content K *\<^sub>R f x) - y) < e/2" by metis show ?thesis proof (rule_tac x=d in exI, clarsimp simp add: \gauge d\) fix p1 p2 assume as: "p1 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p1" "p2 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p2" show "norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e" proof (rule fine_division_exists[OF \gauge d\, of a' b]) fix p assume "p tagged_division_of cbox a' b" "d fine p" then show ?thesis using as norm_triangle_half_l[OF d[of p1 p] d[of p2 p]] unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric] by (auto simp add: algebra_simps) qed qed qed with f show ?thesis1 by (simp add: interval_split[OF k] integrable_Cauchy) have "\d. gauge d \ (\p1 p2. p1 tagged_division_of cbox a b \ {x. x \ k \ c} \ d fine p1 \ p2 tagged_division_of cbox a b \ {x. x \ k \ c} \ d fine p2 \ norm ((\(x,K) \ p1. content K *\<^sub>R f x) - (\(x,K) \ p2. content K *\<^sub>R f x)) < e)" if "e > 0" for e proof - have "e/2 > 0" using that by auto with has_integral_separate_sides[OF y this k, of c] obtain d where "gauge d" and d: "\p1 p2. \p1 tagged_division_of cbox a b \ {x. x \ k \ c}; d fine p1; p2 tagged_division_of cbox a b \ {x. c \ x \ k}; d fine p2\ \ norm ((\(x,K)\p1. content K *\<^sub>R f x) + (\(x,K)\p2. content K *\<^sub>R f x) - y) < e/2" by metis show ?thesis proof (rule_tac x=d in exI, clarsimp simp add: \gauge d\) fix p1 p2 assume as: "p1 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p1" "p2 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p2" show "norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e" proof (rule fine_division_exists[OF \gauge d\, of a b']) fix p assume "p tagged_division_of cbox a b'" "d fine p" then show ?thesis using as norm_triangle_half_l[OF d[of p p1] d[of p p2]] unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric] by (auto simp add: algebra_simps) qed qed qed with f show ?thesis2 by (simp add: interval_split[OF k] integrable_Cauchy) qed lemma operative_integralI: fixes f :: "'a::euclidean_space \ 'b::banach" shows "operative (lift_option (+)) (Some 0) (\i. if f integrable_on i then Some (integral i f) else None)" proof - interpret comm_monoid "lift_option plus" "Some (0::'b)" by (rule comm_monoid_lift_option) (rule add.comm_monoid_axioms) show ?thesis proof fix a b c fix k :: 'a assume k: "k \ Basis" show "(if f integrable_on cbox a b then Some (integral (cbox a b) f) else None) = lift_option (+) (if f integrable_on cbox a b \ {x. x \ k \ c} then Some (integral (cbox a b \ {x. x \ k \ c}) f) else None) (if f integrable_on cbox a b \ {x. c \ x \ k} then Some (integral (cbox a b \ {x. c \ x \ k}) f) else None)" proof (cases "f integrable_on cbox a b") case True with k show ?thesis by (auto simp: integrable_split intro: integral_unique [OF has_integral_split[OF _ _ k]]) next case False have "\ (f integrable_on cbox a b \ {x. x \ k \ c}) \ \ ( f integrable_on cbox a b \ {x. c \ x \ k})" proof (rule ccontr) assume "\ ?thesis" then have "f integrable_on cbox a b" unfolding integrable_on_def apply (rule_tac x="integral (cbox a b \ {x. x \ k \ c}) f + integral (cbox a b \ {x. x \ k \ c}) f" in exI) apply (auto intro: has_integral_split[OF _ _ k]) done then show False using False by auto qed then show ?thesis using False by auto qed next fix a b :: 'a assume "box a b = {}" then show "(if f integrable_on cbox a b then Some (integral (cbox a b) f) else None) = Some 0" using has_integral_null_eq by (auto simp: integrable_on_null content_eq_0_interior) qed qed subsection \Bounds on the norm of Riemann sums and the integral itself\ lemma dsum_bound: assumes p: "p division_of (cbox a b)" and "norm c \ e" shows "norm (sum (\l. content l *\<^sub>R c) p) \ e * content(cbox a b)" proof - have sumeq: "(\i\p. \content i\) = sum content p" by simp have e: "0 \ e" using assms(2) norm_ge_zero order_trans by blast have "norm (sum (\l. content l *\<^sub>R c) p) \ (\i\p. norm (content i *\<^sub>R c))" using norm_sum by blast also have "... \ e * (\i\p. \content i\)" by (simp add: sum_distrib_left[symmetric] mult.commute assms(2) mult_right_mono sum_nonneg) also have "... \ e * content (cbox a b)" by (metis additive_content_division p eq_iff sumeq) finally show ?thesis . qed lemma rsum_bound: assumes p: "p tagged_division_of (cbox a b)" and "\x\cbox a b. norm (f x) \ e" shows "norm (sum (\(x,k). content k *\<^sub>R f x) p) \ e * content (cbox a b)" proof (cases "cbox a b = {}") case True show ?thesis using p unfolding True tagged_division_of_trivial by auto next case False then have e: "e \ 0" by (meson ex_in_conv assms(2) norm_ge_zero order_trans) have sum_le: "sum (content \ snd) p \ content (cbox a b)" unfolding additive_content_tagged_division[OF p, symmetric] split_def by (auto intro: eq_refl) have con: "\xk. xk \ p \ 0 \ content (snd xk)" using tagged_division_ofD(4) [OF p] content_pos_le by force have "norm (sum (\(x,k). content k *\<^sub>R f x) p) \ (\i\p. norm (case i of (x, k) \ content k *\<^sub>R f x))" by (rule norm_sum) also have "... \ e * content (cbox a b)" proof - have "\xk. xk \ p \ norm (f (fst xk)) \ e" using assms(2) p tag_in_interval by force moreover have "(\i\p. \content (snd i)\ * e) \ e * content (cbox a b)" unfolding sum_distrib_right[symmetric] using con sum_le by (auto simp: mult.commute intro: mult_left_mono [OF _ e]) ultimately show ?thesis unfolding split_def norm_scaleR by (metis (no_types, lifting) mult_left_mono[OF _ abs_ge_zero] order_trans[OF sum_mono]) qed finally show ?thesis . qed lemma rsum_diff_bound: assumes "p tagged_division_of (cbox a b)" and "\x\cbox a b. norm (f x - g x) \ e" shows "norm (sum (\(x,k). content k *\<^sub>R f x) p - sum (\(x,k). content k *\<^sub>R g x) p) \ e * content (cbox a b)" using order_trans[OF _ rsum_bound[OF assms]] by (simp add: split_def scaleR_diff_right sum_subtractf eq_refl) lemma has_integral_bound: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "0 \ B" and f: "(f has_integral i) (cbox a b)" and "\x. x\cbox a b \ norm (f x) \ B" shows "norm i \ B * content (cbox a b)" proof (rule ccontr) assume "\ ?thesis" then have "norm i - B * content (cbox a b) > 0" by auto with f[unfolded has_integral] obtain \ where "gauge \" and \: "\p. \p tagged_division_of cbox a b; \ fine p\ \ norm ((\(x, K)\p. content K *\<^sub>R f x) - i) < norm i - B * content (cbox a b)" by metis then obtain p where p: "p tagged_division_of cbox a b" and "\ fine p" using fine_division_exists by blast have "\s B. norm s \ B \ \ norm (s - i) < norm i - B" unfolding not_less by (metis diff_left_mono dist_commute dist_norm norm_triangle_ineq2 order_trans) then show False using \ [OF p \\ fine p\] rsum_bound[OF p] assms by metis qed corollary integrable_bound: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "0 \ B" and "f integrable_on (cbox a b)" and "\x. x\cbox a b \ norm (f x) \ B" shows "norm (integral (cbox a b) f) \ B * content (cbox a b)" by (metis integrable_integral has_integral_bound assms) subsection \Similar theorems about relationship among components\ lemma rsum_component_le: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes p: "p tagged_division_of (cbox a b)" and "\x. x \ cbox a b \ (f x)\i \ (g x)\i" shows "(\(x, K)\p. content K *\<^sub>R f x) \ i \ (\(x, K)\p. content K *\<^sub>R g x) \ i" unfolding inner_sum_left proof (rule sum_mono, clarify) fix x K assume ab: "(x, K) \ p" with p obtain u v where K: "K = cbox u v" by blast then show "(content K *\<^sub>R f x) \ i \ (content K *\<^sub>R g x) \ i" by (metis ab assms inner_scaleR_left measure_nonneg mult_left_mono tag_in_interval) qed lemma has_integral_component_le: fixes f g :: "'a::euclidean_space \ 'b::euclidean_space" assumes k: "k \ Basis" assumes "(f has_integral i) S" "(g has_integral j) S" and f_le_g: "\x. x \ S \ (f x)\k \ (g x)\k" shows "i\k \ j\k" proof - have ik_le_jk: "i\k \ j\k" if f_i: "(f has_integral i) (cbox a b)" and g_j: "(g has_integral j) (cbox a b)" and le: "\x\cbox a b. (f x)\k \ (g x)\k" for a b i and j :: 'b and f g :: "'a \ 'b" proof (rule ccontr) assume "\ ?thesis" then have *: "0 < (i\k - j\k) / 3" by auto obtain \1 where "gauge \1" and \1: "\p. \p tagged_division_of cbox a b; \1 fine p\ \ norm ((\(x, k)\p. content k *\<^sub>R f x) - i) < (i \ k - j \ k) / 3" using f_i[unfolded has_integral,rule_format,OF *] by fastforce obtain \2 where "gauge \2" and \2: "\p. \p tagged_division_of cbox a b; \2 fine p\ \ norm ((\(x, k)\p. content k *\<^sub>R g x) - j) < (i \ k - j \ k) / 3" using g_j[unfolded has_integral,rule_format,OF *] by fastforce obtain p where p: "p tagged_division_of cbox a b" and "\1 fine p" "\2 fine p" using fine_division_exists[OF gauge_Int[OF \gauge \1\ \gauge \2\], of a b] unfolding fine_Int by metis then have "\((\(x, k)\p. content k *\<^sub>R f x) - i) \ k\ < (i \ k - j \ k) / 3" "\((\(x, k)\p. content k *\<^sub>R g x) - j) \ k\ < (i \ k - j \ k) / 3" using le_less_trans[OF Basis_le_norm[OF k]] k \1 \2 by metis+ then show False unfolding inner_simps using rsum_component_le[OF p] le by (fastforce simp add: abs_real_def split: if_split_asm) qed show ?thesis proof (cases "\a b. S = cbox a b") case True with ik_le_jk assms show ?thesis by auto next case False show ?thesis proof (rule ccontr) assume "\ i\k \ j\k" then have ij: "(i\k - j\k) / 3 > 0" by auto obtain B1 where "0 < B1" and B1: "\a b. ball 0 B1 \ cbox a b \ \z. ((\x. if x \ S then f x else 0) has_integral z) (cbox a b) \ norm (z - i) < (i \ k - j \ k) / 3" using has_integral_altD[OF _ False ij] assms by blast obtain B2 where "0 < B2" and B2: "\a b. ball 0 B2 \ cbox a b \ \z. ((\x. if x \ S then g x else 0) has_integral z) (cbox a b) \ norm (z - j) < (i \ k - j \ k) / 3" using has_integral_altD[OF _ False ij] assms by blast have "bounded (ball 0 B1 \ ball (0::'a) B2)" unfolding bounded_Un by(rule conjI bounded_ball)+ from bounded_subset_cbox_symmetric[OF this] obtain a b::'a where ab: "ball 0 B1 \ cbox a b" "ball 0 B2 \ cbox a b" by (meson Un_subset_iff) then obtain w1 w2 where int_w1: "((\x. if x \ S then f x else 0) has_integral w1) (cbox a b)" and norm_w1: "norm (w1 - i) < (i \ k - j \ k) / 3" and int_w2: "((\x. if x \ S then g x else 0) has_integral w2) (cbox a b)" and norm_w2: "norm (w2 - j) < (i \ k - j \ k) / 3" using B1 B2 by blast have *: "\w1 w2 j i::real .\w1 - i\ < (i - j) / 3 \ \w2 - j\ < (i - j) / 3 \ w1 \ w2 \ False" by (simp add: abs_real_def split: if_split_asm) have "\(w1 - i) \ k\ < (i \ k - j \ k) / 3" "\(w2 - j) \ k\ < (i \ k - j \ k) / 3" using Basis_le_norm k le_less_trans norm_w1 norm_w2 by blast+ moreover have "w1\k \ w2\k" using ik_le_jk int_w1 int_w2 f_le_g by auto ultimately show False unfolding inner_simps by(rule *) qed qed qed lemma integral_component_le: fixes g f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "k \ Basis" and "f integrable_on S" "g integrable_on S" and "\x. x \ S \ (f x)\k \ (g x)\k" shows "(integral S f)\k \ (integral S g)\k" using has_integral_component_le assms by blast lemma has_integral_component_nonneg: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "k \ Basis" and "(f has_integral i) S" and "\x. x \ S \ 0 \ (f x)\k" shows "0 \ i\k" by (metis (no_types, lifting) assms euclidean_all_zero_iff has_integral_0 has_integral_component_le) lemma integral_component_nonneg: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "k \ Basis" and "\x. x \ S \ 0 \ (f x)\k" shows "0 \ (integral S f)\k" by (smt (verit, ccfv_threshold) assms eq_integralD euclidean_all_zero_iff has_integral_component_nonneg) lemma has_integral_component_neg: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "k \ Basis" and "(f has_integral i) S" and "\x. x \ S \ (f x)\k \ 0" shows "i\k \ 0" by (metis (no_types, lifting) assms has_integral_0 has_integral_component_le inner_zero_left) lemma has_integral_component_lbound: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "(f has_integral i) (cbox a b)" and "\x\cbox a b. B \ f(x)\k" and "k \ Basis" shows "B * content (cbox a b) \ i\k" using has_integral_component_le[OF assms(3) has_integral_const assms(1),of "(\i\Basis. B *\<^sub>R i)::'b"] assms(2-) by (auto simp add: field_simps) lemma has_integral_component_ubound: fixes f::"'a::euclidean_space => 'b::euclidean_space" assumes "(f has_integral i) (cbox a b)" and "\x\cbox a b. f x\k \ B" and "k \ Basis" shows "i\k \ B * content (cbox a b)" using has_integral_component_le[OF assms(3,1) has_integral_const, of "\i\Basis. B *\<^sub>R i"] assms(2-) by (auto simp add: field_simps) lemma integral_component_lbound: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f integrable_on cbox a b" and "\x\cbox a b. B \ f(x)\k" and "k \ Basis" shows "B * content (cbox a b) \ (integral(cbox a b) f)\k" using assms has_integral_component_lbound by blast lemma integral_component_lbound_real: assumes "f integrable_on {a ::real..b}" and "\x\{a..b}. B \ f(x)\k" and "k \ Basis" shows "B * content {a..b} \ (integral {a..b} f)\k" using assms by (metis box_real(2) integral_component_lbound) lemma integral_component_ubound: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f integrable_on cbox a b" and "\x\cbox a b. f x\k \ B" and "k \ Basis" shows "(integral (cbox a b) f)\k \ B * content (cbox a b)" using assms has_integral_component_ubound by blast lemma integral_component_ubound_real: fixes f :: "real \ 'a::euclidean_space" assumes "f integrable_on {a..b}" and "\x\{a..b}. f x\k \ B" and "k \ Basis" shows "(integral {a..b} f)\k \ B * content {a..b}" using assms by (metis box_real(2) integral_component_ubound) subsection \Uniform limit of integrable functions is integrable\ lemma real_arch_invD: "0 < (e::real) \ (\n::nat. n \ 0 \ 0 < inverse (real n) \ inverse (real n) < e)" by (subst(asm) real_arch_inverse) lemma integrable_uniform_limit: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "\e. e > 0 \ \g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" shows "f integrable_on cbox a b" proof (cases "content (cbox a b) > 0") case False then show ?thesis using has_integral_null by (simp add: content_lt_nz integrable_on_def) next case True have "1 / (real n + 1) > 0" for n by auto then have "\g. (\x\cbox a b. norm (f x - g x) \ 1 / (real n + 1)) \ g integrable_on cbox a b" for n using assms by blast then obtain g where g_near_f: "\n x. x \ cbox a b \ norm (f x - g n x) \ 1 / (real n + 1)" and int_g: "\n. g n integrable_on cbox a b" by metis then obtain h where h: "\n. (g n has_integral h n) (cbox a b)" unfolding integrable_on_def by metis have "Cauchy h" unfolding Cauchy_def proof clarify fix e :: real assume "e>0" then have "e/4 / content (cbox a b) > 0" using True by (auto simp: field_simps) then obtain M where "M \ 0" and M: "1 / (real M) < e/4 / content (cbox a b)" by (metis inverse_eq_divide real_arch_inverse) show "\M. \m\M. \n\M. dist (h m) (h n) < e" proof (intro exI strip) fix m n assume m: "M \ m" and n: "M \ n" have "e/4>0" using \e>0\ by auto then obtain gm gn where "gauge gm" "gauge gn" and gm: "\\. \ tagged_division_of cbox a b \ gm fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R g m x) - h m) < e/4" and gn: "\\. \ tagged_division_of cbox a b \ gn fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R g n x) - h n) < e/4" using h[unfolded has_integral] by meson then obtain \ where \: "\ tagged_division_of cbox a b" "(\x. gm x \ gn x) fine \" by (metis (full_types) fine_division_exists gauge_Int) have triangle3: "norm (i1 - i2) < e" if no: "norm(s2 - s1) \ e/2" "norm (s1 - i1) < e/4" "norm (s2 - i2) < e/4" for s1 s2 i1 and i2::'b proof - have "norm (i1 - i2) \ norm (i1 - s1) + norm (s1 - s2) + norm (s2 - i2)" using norm_triangle_ineq[of "i1 - s1" "s1 - i2"] using norm_triangle_ineq[of "s1 - s2" "s2 - i2"] by (auto simp: algebra_simps) also have "\ < e" using no by (auto simp: algebra_simps norm_minus_commute) finally show ?thesis . qed have finep: "gm fine \" "gn fine \" using fine_Int \ by auto have norm_le: "norm (g n x - g m x) \ 2 / real M" if x: "x \ cbox a b" for x proof - have "norm (f x - g n x) + norm (f x - g m x) \ 1 / (real n + 1) + 1 / (real m + 1)" using g_near_f[OF x, of n] g_near_f[OF x, of m] by simp also have "\ \ 1 / (real M) + 1 / (real M)" using \M \ 0\ m n by (intro add_mono; force simp: field_split_simps) also have "\ = 2 / real M" by auto finally show "norm (g n x - g m x) \ 2 / real M" using norm_triangle_le[of "g n x - f x" "f x - g m x" "2 / real M"] by (auto simp: algebra_simps simp add: norm_minus_commute) qed have "norm ((\(x,K) \ \. content K *\<^sub>R g n x) - (\(x,K) \ \. content K *\<^sub>R g m x)) \ 2 / real M * content (cbox a b)" by (blast intro: norm_le rsum_diff_bound[OF \(1), where e="2 / real M"]) also have "... \ e/2" using M True by (auto simp: field_simps) finally have le_e2: "norm ((\(x,K) \ \. content K *\<^sub>R g n x) - (\(x,K) \ \. content K *\<^sub>R g m x)) \ e/2" . then show "dist (h m) (h n) < e" unfolding dist_norm using gm gn \ finep by (auto intro!: triangle3) qed qed then obtain s where s: "h \ s" using convergent_eq_Cauchy[symmetric] by blast show ?thesis unfolding integrable_on_def has_integral proof (rule_tac x=s in exI, clarify) fix e::real assume e: "0 < e" then have "e/3 > 0" by auto then obtain N1 where N1: "\n\N1. norm (h n - s) < e/3" using LIMSEQ_D [OF s] by metis from e True have "e/3 / content (cbox a b) > 0" by (auto simp: field_simps) then obtain N2 :: nat where "N2 \ 0" and N2: "1 / (real N2) < e/3 / content (cbox a b)" by (metis inverse_eq_divide real_arch_inverse) obtain g' where "gauge g'" and g': "\\. \ tagged_division_of cbox a b \ g' fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R g (N1 + N2) x) - h (N1 + N2)) < e/3" by (metis h has_integral \e/3 > 0\) have *: "norm (sf - s) < e" if no: "norm (sf - sg) \ e/3" "norm(h - s) < e/3" "norm (sg - h) < e/3" for sf sg h proof - have "norm (sf - s) \ norm (sf - sg) + norm (sg - h) + norm (h - s)" using norm_triangle_ineq[of "sf - sg" "sg - s"] using norm_triangle_ineq[of "sg - h" " h - s"] by (auto simp: algebra_simps) also have "\ < e" using no by (auto simp: algebra_simps norm_minus_commute) finally show ?thesis . qed { fix \ assume ptag: "\ tagged_division_of (cbox a b)" and "g' fine \" then have norm_less: "norm ((\(x,K) \ \. content K *\<^sub>R g (N1 + N2) x) - h (N1 + N2)) < e/3" using g' by blast have "content (cbox a b) < e/3 * (of_nat N2)" using \N2 \ 0\ N2 using True by (auto simp: field_split_simps) moreover have "e/3 * of_nat N2 \ e/3 * (of_nat (N1 + N2) + 1)" using \e>0\ by auto ultimately have "content (cbox a b) < e/3 * (of_nat (N1 + N2) + 1)" by linarith then have le_e3: "1 / (real (N1 + N2) + 1) * content (cbox a b) \ e/3" unfolding inverse_eq_divide by (auto simp: field_simps) have ne3: "norm (h (N1 + N2) - s) < e/3" using N1 by auto have "norm ((\(x,K) \ \. content K *\<^sub>R f x) - (\(x,K) \ \. content K *\<^sub>R g (N1 + N2) x)) \ 1 / (real (N1 + N2) + 1) * content (cbox a b)" by (blast intro: g_near_f rsum_diff_bound[OF ptag]) then have "norm ((\(x,K) \ \. content K *\<^sub>R f x) - s) < e" by (rule *[OF order_trans [OF _ le_e3] ne3 norm_less]) } then show "\d. gauge d \ (\\. \ tagged_division_of cbox a b \ d fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - s) < e)" by (blast intro: g' \gauge g'\) qed qed lemmas integrable_uniform_limit_real = integrable_uniform_limit [where 'a=real, simplified] subsection \Negligible sets\ definition "negligible (s:: 'a::euclidean_space set) \ (\a b. ((indicator s :: 'a\real) has_integral 0) (cbox a b))" subsubsection \Negligibility of hyperplane\ lemma content_doublesplit: fixes a :: "'a::euclidean_space" assumes "0 < e" and k: "k \ Basis" obtains d where "0 < d" and "content (cbox a b \ {x. \x\k - c\ \ d}) < e" proof cases assume *: "a \ k \ c \ c \ b \ k \ (\j\Basis. a \ j \ b \ j)" define a' where "a' d = (\j\Basis. (if j = k then max (a\j) (c - d) else a\j) *\<^sub>R j)" for d define b' where "b' d = (\j\Basis. (if j = k then min (b\j) (c + d) else b\j) *\<^sub>R j)" for d have "((\d. \j\Basis. (b' d - a' d) \ j) \ (\j\Basis. (b' 0 - a' 0) \ j)) (at_right 0)" by (auto simp: b'_def a'_def intro!: tendsto_min tendsto_max tendsto_eq_intros) also have "(\j\Basis. (b' 0 - a' 0) \ j) = 0" using k * by (intro prod_zero bexI[OF _ k]) (auto simp: b'_def a'_def inner_diff inner_sum_left inner_not_same_Basis intro!: sum.cong) also have "((\d. \j\Basis. (b' d - a' d) \ j) \ 0) (at_right 0) = ((\d. content (cbox a b \ {x. \x\k - c\ \ d})) \ 0) (at_right 0)" proof (intro tendsto_cong eventually_at_rightI) fix d :: real assume d: "d \ {0<..<1}" have "cbox a b \ {x. \x\k - c\ \ d} = cbox (a' d) (b' d)" for d using * d k by (auto simp add: cbox_def set_eq_iff Int_def ball_conj_distrib abs_diff_le_iff a'_def b'_def) moreover have "j \ Basis \ a' d \ j \ b' d \ j" for j using * d k by (auto simp: a'_def b'_def) ultimately show "(\j\Basis. (b' d - a' d) \ j) = content (cbox a b \ {x. \x\k - c\ \ d})" by simp qed simp finally have "((\d. content (cbox a b \ {x. \x \ k - c\ \ d})) \ 0) (at_right 0)" . from order_tendstoD(2)[OF this \0] obtain d' where "0 < d'" and d': "\y. y > 0 \ y < d' \ content (cbox a b \ {x. \x \ k - c\ \ y}) < e" by (subst (asm) eventually_at_right[of _ 1]) auto show ?thesis by (rule that[of "d'/2"], insert \0 d'[of "d'/2"], auto) next assume *: "\ (a \ k \ c \ c \ b \ k \ (\j\Basis. a \ j \ b \ j))" then have "(\j\Basis. b \ j < a \ j) \ (c < a \ k \ b \ k < c)" by (auto simp: not_le) show thesis proof cases assume "\j\Basis. b \ j < a \ j" then have [simp]: "cbox a b = {}" using box_ne_empty(1)[of a b] by auto show ?thesis by (rule that[of 1]) (simp_all add: \0) next assume "\ (\j\Basis. b \ j < a \ j)" with * have "c < a \ k \ b \ k < c" by auto then show thesis proof assume c: "c < a \ k" moreover have "x \ cbox a b \ c \ x \ k" for x using k c by (auto simp: cbox_def) ultimately have "cbox a b \ {x. \x \ k - c\ \ (a \ k - c)/2} = {}" using k by (auto simp: cbox_def) with \0 c that[of "(a \ k - c)/2"] show ?thesis by auto next assume c: "b \ k < c" moreover have "x \ cbox a b \ x \ k \ c" for x using k c by (auto simp: cbox_def) ultimately have "cbox a b \ {x. \x \ k - c\ \ (c - b \ k)/2} = {}" using k by (auto simp: cbox_def) with \0 c that[of "(c - b \ k)/2"] show ?thesis by auto qed qed qed proposition negligible_standard_hyperplane[intro]: fixes k :: "'a::euclidean_space" assumes k: "k \ Basis" shows "negligible {x. x\k = c}" unfolding negligible_def has_integral proof clarsimp fix a b and e::real assume "e > 0" with k obtain d where "0 < d" and d: "content (cbox a b \ {x. \x \ k - c\ \ d}) < e" by (metis content_doublesplit) let ?i = "indicator {x::'a. x\k = c} :: 'a\real" show "\\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ \\(x,K) \ \. content K * ?i x\ < e)" proof (intro exI, safe) show "gauge (\x. ball x d)" using \0 < d\ by blast next fix \ assume p: "\ tagged_division_of (cbox a b)" "(\x. ball x d) fine \" have "content L = content (L \ {x. \x \ k - c\ \ d})" if "(x, L) \ \" "?i x \ 0" for x L proof - have xk: "x\k = c" using that by (simp add: indicator_def split: if_split_asm) have "L \ {x. \x \ k - c\ \ d}" proof fix y assume y: "y \ L" have "L \ ball x d" using p(2) that(1) by auto then have "norm (x - y) < d" by (simp add: dist_norm subset_iff y) then have "\(x - y) \ k\ < d" using k norm_bound_Basis_lt by blast then show "y \ {x. \x \ k - c\ \ d}" unfolding inner_simps xk by auto qed then show "content L = content (L \ {x. \x \ k - c\ \ d})" by (metis inf.orderE) qed then have *: "(\(x,K)\\. content K * ?i x) = (\(x,K)\\. content (K \ {x. \x\k - c\ \ d}) *\<^sub>R ?i x)" by (force simp add: split_paired_all intro!: sum.cong [OF refl]) note p'= tagged_division_ofD[OF p(1)] and p''=division_of_tagged_division[OF p(1)] have "(\(x,K)\\. content (K \ {x. \x \ k - c\ \ d}) * indicator {x. x \ k = c} x) < e" proof - have "(\(x,K)\\. content (K \ {x. \x \ k - c\ \ d}) * ?i x) \ (\(x,K)\\. content (K \ {x. \x \ k - c\ \ d}))" by (force simp add: indicator_def intro!: sum_mono) also have "\ < e" proof (subst sum.over_tagged_division_lemma[OF p(1)]) fix u v::'a assume "box u v = {}" then have *: "content (cbox u v) = 0" unfolding content_eq_0_interior by simp have "cbox u v \ {x. \x \ k - c\ \ d} \ cbox u v" by auto then have "content (cbox u v \ {x. \x \ k - c\ \ d}) \ content (cbox u v)" unfolding interval_doublesplit[OF k] by (rule content_subset) then show "content (cbox u v \ {x. \x \ k - c\ \ d}) = 0" unfolding * interval_doublesplit[OF k] by (blast intro: antisym) next have "(\l\snd ` \. content (l \ {x. \x \ k - c\ \ d})) = sum content ((\l. l \ {x. \x \ k - c\ \ d})`{l\snd ` \. l \ {x. \x \ k - c\ \ d} \ {}})" proof (subst (2) sum.reindex_nontrivial) fix x y assume "x \ {l \ snd ` \. l \ {x. \x \ k - c\ \ d} \ {}}" "y \ {l \ snd ` \. l \ {x. \x \ k - c\ \ d} \ {}}" "x \ y" and eq: "x \ {x. \x \ k - c\ \ d} = y \ {x. \x \ k - c\ \ d}" then obtain x' y' where "(x', x) \ \" "x \ {x. \x \ k - c\ \ d} \ {}" "(y', y) \ \" "y \ {x. \x \ k - c\ \ d} \ {}" by (auto) from p'(5)[OF \(x', x) \ \\ \(y', y) \ \\] \x \ y\ have "interior (x \ y) = {}" by auto moreover have "interior ((x \ {x. \x \ k - c\ \ d}) \ (y \ {x. \x \ k - c\ \ d})) \ interior (x \ y)" by (auto intro: interior_mono) ultimately have "interior (x \ {x. \x \ k - c\ \ d}) = {}" by (auto simp: eq) then show "content (x \ {x. \x \ k - c\ \ d}) = 0" using p'(4)[OF \(x', x) \ \\] by (auto simp: interval_doublesplit[OF k] content_eq_0_interior simp del: interior_Int) qed (insert p'(1), auto intro!: sum.mono_neutral_right) also have "\ \ norm (\l\(\l. l \ {x. \x \ k - c\ \ d})`{l\snd ` \. l \ {x. \x \ k - c\ \ d} \ {}}. content l *\<^sub>R 1::real)" by simp also have "\ \ 1 * content (cbox a b \ {x. \x \ k - c\ \ d})" using division_doublesplit[OF p'' k, unfolded interval_doublesplit[OF k]] unfolding interval_doublesplit[OF k] by (intro dsum_bound) auto also have "\ < e" using d by simp finally show "(\K\snd ` \. content (K \ {x. \x \ k - c\ \ d})) < e" . qed finally show "(\(x, K)\\. content (K \ {x. \x \ k - c\ \ d}) * ?i x) < e" . qed then show "\\(x, K)\\. content K * ?i x\ < e" unfolding * by (simp add: sum_nonneg split: prod.split) qed qed corollary negligible_standard_hyperplane_cart: fixes k :: "'a::finite" shows "negligible {x. x$k = (0::real)}" by (simp add: cart_eq_inner_axis negligible_standard_hyperplane) subsubsection \Hence the main theorem about negligible sets\ lemma has_integral_negligible_cbox: fixes f :: "'b::euclidean_space \ 'a::real_normed_vector" assumes negs: "negligible S" and 0: "\x. x \ S \ f x = 0" shows "(f has_integral 0) (cbox a b)" unfolding has_integral proof clarify fix e::real assume "e > 0" then have nn_gt0: "e/2 / ((real n+1) * (2 ^ n)) > 0" for n by simp then have "\\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ \\(x,K) \ \. content K *\<^sub>R indicator S x\ < e/2 / ((real n + 1) * 2 ^ n))" for n using negs [unfolded negligible_def has_integral] by auto then obtain \ where gd: "\n. gauge (\ n)" and \: "\n \. \\ tagged_division_of cbox a b; \ n fine \\ \ \\(x,K) \ \. content K *\<^sub>R indicator S x\ < e/2 / ((real n + 1) * 2 ^ n)" by metis show "\\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - 0) < e)" proof (intro exI, safe) show "gauge (\x. \ (nat \norm (f x)\) x)" using gd by (auto simp: gauge_def) show "norm ((\(x,K) \ \. content K *\<^sub>R f x) - 0) < e" if "\ tagged_division_of (cbox a b)" "(\x. \ (nat \norm (f x)\) x) fine \" for \ proof (cases "\ = {}") case True with \0 < e\ show ?thesis by simp next case False obtain N where "Max ((\(x, K). norm (f x)) ` \) \ real N" using real_arch_simple by blast then have N: "\x. x \ (\(x, K). norm (f x)) ` \ \ x \ real N" by (meson Max_ge that(1) dual_order.trans finite_imageI tagged_division_of_finite) have "\i. \q. q tagged_division_of (cbox a b) \ (\ i) fine q \ (\(x,K) \ \. K \ (\ i) x \ (x, K) \ q)" by (auto intro: tagged_division_finer[OF that(1) gd]) from choice[OF this] obtain q where q: "\n. q n tagged_division_of cbox a b" "\n. \ n fine q n" "\n x K. \(x, K) \ \; K \ \ n x\ \ (x, K) \ q n" by fastforce have "finite \" using that(1) by blast then have sum_le_inc: "\finite T; \x y. (x,y) \ T \ (0::real) \ g(x,y); \y. y\\ \ \x. (x,y) \ T \ f(y) \ g(x,y)\ \ sum f \ \ sum g T" for f g T by (rule sum_le_included[of \ T g snd f]; force) have "norm (\(x,K) \ \. content K *\<^sub>R f x) \ (\(x,K) \ \. norm (content K *\<^sub>R f x))" unfolding split_def by (rule norm_sum) also have "... \ (\(i, j) \ Sigma {..N + 1} q. (real i + 1) * (case j of (x, K) \ content K *\<^sub>R indicator S x))" proof (rule sum_le_inc, safe) show "finite (Sigma {..N+1} q)" by (meson finite_SigmaI finite_atMost tagged_division_of_finite q(1)) next fix x K assume xk: "(x, K) \ \" define n where "n = nat \norm (f x)\" have *: "norm (f x) \ (\(x, K). norm (f x)) ` \" using xk by auto have nfx: "real n \ norm (f x)" "norm (f x) \ real n + 1" unfolding n_def by auto then have "n \ {0..N + 1}" using N[OF *] by auto moreover have "K \ \ (nat \norm (f x)\) x" using that(2) xk by auto moreover then have "(x, K) \ q (nat \norm (f x)\)" by (simp add: q(3) xk) moreover then have "(x, K) \ q n" using n_def by blast moreover have "norm (content K *\<^sub>R f x) \ (real n + 1) * (content K * indicator S x)" proof (cases "x \ S") case False then show ?thesis by (simp add: 0) next case True have *: "content K \ 0" using tagged_division_ofD(4)[OF that(1) xk] by auto moreover have "content K * norm (f x) \ content K * (real n + 1)" by (simp add: mult_left_mono nfx(2)) ultimately show ?thesis using nfx True by (auto simp: field_simps) qed ultimately show "\y. (y, x, K) \ (Sigma {..N + 1} q) \ norm (content K *\<^sub>R f x) \ (real y + 1) * (content K *\<^sub>R indicator S x)" by force qed auto also have "... = (\i\N + 1. \j\q i. (real i + 1) * (case j of (x, K) \ content K *\<^sub>R indicator S x))" using q(1) by (intro sum_Sigma_product [symmetric]) auto also have "... \ (\i\N + 1. (real i + 1) * \\(x,K) \ q i. content K *\<^sub>R indicator S x\)" by (rule sum_mono) (simp add: sum_distrib_left [symmetric]) also have "... \ (\i\N + 1. e/2/2 ^ i)" proof (rule sum_mono) show "(real i + 1) * \\(x,K) \ q i. content K *\<^sub>R indicator S x\ \ e/2/2 ^ i" if "i \ {..N + 1}" for i using \[of "q i" i] q by (simp add: divide_simps mult.left_commute) qed also have "... = e/2 * (\i\N + 1. (1/2) ^ i)" unfolding sum_distrib_left by (metis divide_inverse inverse_eq_divide power_one_over) also have "\ < e/2 * 2" proof (rule mult_strict_left_mono) have "sum (power (1/2)) {..N + 1} = sum (power (1/2::real)) {..0 < e\ in auto) finally show ?thesis by auto qed qed qed proposition has_integral_negligible: fixes f :: "'b::euclidean_space \ 'a::real_normed_vector" assumes negs: "negligible S" and "\x. x \ (T - S) \ f x = 0" shows "(f has_integral 0) T" proof (cases "\a b. T = cbox a b") case True then have "((\x. if x \ T then f x else 0) has_integral 0) T" using assms by (auto intro!: has_integral_negligible_cbox) then show ?thesis by (rule has_integral_eq [rotated]) auto next case False let ?f = "(\x. if x \ T then f x else 0)" have "((\x. if x \ T then f x else 0) has_integral 0) T" apply (auto simp: False has_integral_alt [of ?f]) apply (rule_tac x=1 in exI, auto) apply (rule_tac x=0 in exI, simp add: has_integral_negligible_cbox [OF negs] assms) done then show ?thesis by (rule_tac f="?f" in has_integral_eq) auto qed lemma assumes "negligible S" shows integrable_negligible: "f integrable_on S" and integral_negligible: "integral S f = 0" using has_integral_negligible [OF assms] by (auto simp: has_integral_iff) lemma has_integral_spike: fixes f :: "'b::euclidean_space \ 'a::real_normed_vector" assumes "negligible S" and gf: "\x. x \ T - S \ g x = f x" and fint: "(f has_integral y) T" shows "(g has_integral y) T" proof - have *: "(g has_integral y) (cbox a b)" if "(f has_integral y) (cbox a b)" "\x \ cbox a b - S. g x = f x" for a b f and g:: "'b \ 'a" and y proof - have "((\x. f x + (g x - f x)) has_integral (y + 0)) (cbox a b)" using that by (intro has_integral_add has_integral_negligible) (auto intro!: \negligible S\) then show ?thesis by auto qed have \
: "\a b z. \\x. x \ T \ x \ S \ g x = f x; ((\x. if x \ T then f x else 0) has_integral z) (cbox a b)\ \ ((\x. if x \ T then g x else 0) has_integral z) (cbox a b)" by (auto dest!: *[where f="\x. if x\T then f x else 0" and g="\x. if x \ T then g x else 0"]) show ?thesis using fint gf apply (subst has_integral_alt) apply (subst (asm) has_integral_alt) apply (auto split: if_split_asm) apply (blast dest: *) using \
by meson qed lemma has_integral_spike_eq: assumes "negligible S" and "\x. x \ T - S \ g x = f x" shows "(f has_integral y) T \ (g has_integral y) T" by (metis assms has_integral_spike) lemma integrable_spike: assumes "f integrable_on T" "negligible S" "\x. x \ T - S \ g x = f x" shows "g integrable_on T" using assms unfolding integrable_on_def by (blast intro: has_integral_spike) lemma integral_spike: assumes "negligible S" and "\x. x \ T - S \ g x = f x" shows "integral T f = integral T g" using has_integral_spike_eq[OF assms] by (auto simp: integral_def integrable_on_def) subsection \Some other trivialities about negligible sets\ lemma negligible_subset: assumes "negligible s" "t \ s" shows "negligible t" unfolding negligible_def by (metis (no_types) Diff_iff assms contra_subsetD has_integral_negligible indicator_simps(2)) lemma negligible_diff[intro?]: assumes "negligible s" shows "negligible (s - t)" using assms by (meson Diff_subset negligible_subset) lemma negligible_Int: assumes "negligible s \ negligible t" shows "negligible (s \ t)" using assms negligible_subset by force lemma negligible_Un: assumes "negligible S" and T: "negligible T" shows "negligible (S \ T)" proof - have "(indicat_real (S \ T) has_integral 0) (cbox a b)" if S0: "(indicat_real S has_integral 0) (cbox a b)" and "(indicat_real T has_integral 0) (cbox a b)" for a b proof (subst has_integral_spike_eq[OF T]) show "indicat_real S x = indicat_real (S \ T) x" if "x \ cbox a b - T" for x using that by (simp add: indicator_def) show "(indicat_real S has_integral 0) (cbox a b)" by (simp add: S0) qed with assms show ?thesis unfolding negligible_def by blast qed lemma negligible_Un_eq[simp]: "negligible (s \ t) \ negligible s \ negligible t" using negligible_Un negligible_subset by blast lemma negligible_sing[intro]: "negligible {a::'a::euclidean_space}" using negligible_standard_hyperplane[OF SOME_Basis, of "a \ (SOME i. i \ Basis)"] negligible_subset by blast lemma negligible_insert[simp]: "negligible (insert a s) \ negligible s" by (metis insert_is_Un negligible_Un_eq negligible_sing) lemma negligible_empty[iff]: "negligible {}" using negligible_insert by blast text\Useful in this form for backchaining\ lemma empty_imp_negligible: "S = {} \ negligible S" by simp lemma negligible_finite[intro]: assumes "finite s" shows "negligible s" using assms by (induct s) auto lemma negligible_Union[intro]: assumes "finite \" and "\t. t \ \ \ negligible t" shows "negligible(\\)" using assms by induct auto lemma negligible: "negligible S \ (\T. (indicat_real S has_integral 0) T)" proof (intro iffI allI) fix T assume "negligible S" then show "(indicator S has_integral 0) T" by (meson Diff_iff has_integral_negligible indicator_simps(2)) qed (simp add: negligible_def) subsection \Finite case of the spike theorem is quite commonly needed\ lemma has_integral_spike_finite: assumes "finite S" and "\x. x \ T - S \ g x = f x" and "(f has_integral y) T" shows "(g has_integral y) T" using assms has_integral_spike negligible_finite by blast lemma has_integral_spike_finite_eq: assumes "finite S" and "\x. x \ T - S \ g x = f x" shows "((f has_integral y) T \ (g has_integral y) T)" by (metis assms has_integral_spike_finite) lemma integrable_spike_finite: assumes "finite S" and "\x. x \ T - S \ g x = f x" and "f integrable_on T" shows "g integrable_on T" using assms has_integral_spike_finite by blast +lemma integrable_spike_finite_eq: + assumes "finite S" + and "\x. x \ T - S \ f x = g x" + shows "f integrable_on T \ g integrable_on T" + by (metis assms integrable_spike_finite) + lemma has_integral_bound_spike_finite: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "0 \ B" "finite S" and f: "(f has_integral i) (cbox a b)" and leB: "\x. x \ cbox a b - S \ norm (f x) \ B" shows "norm i \ B * content (cbox a b)" proof - define g where "g \ (\x. if x \ S then 0 else f x)" then have "\x. x \ cbox a b - S \ norm (g x) \ B" using leB by simp moreover have "(g has_integral i) (cbox a b)" using has_integral_spike_finite [OF \finite S\ _ f] by (simp add: g_def) ultimately show ?thesis by (simp add: \0 \ B\ g_def has_integral_bound) qed corollary has_integral_bound_real: fixes f :: "real \ 'b::real_normed_vector" assumes "0 \ B" "finite S" and "(f has_integral i) {a..b}" and "\x. x \ {a..b} - S \ norm (f x) \ B" shows "norm i \ B * content {a..b}" by (metis assms box_real(2) has_integral_bound_spike_finite) subsection \In particular, the boundary of an interval is negligible\ lemma negligible_frontier_interval: "negligible(cbox (a::'a::euclidean_space) b - box a b)" proof - let ?A = "\((\k. {x. x\k = a\k} \ {x::'a. x\k = b\k}) ` Basis)" have "negligible ?A" by (force simp add: negligible_Union[OF finite_imageI]) moreover have "cbox a b - box a b \ ?A" by (force simp add: mem_box) ultimately show ?thesis by (rule negligible_subset) qed lemma has_integral_spike_interior: assumes f: "(f has_integral y) (cbox a b)" and gf: "\x. x \ box a b \ g x = f x" shows "(g has_integral y) (cbox a b)" by (meson Diff_iff gf has_integral_spike[OF negligible_frontier_interval _ f]) lemma has_integral_spike_interior_eq: assumes "\x. x \ box a b \ g x = f x" shows "(f has_integral y) (cbox a b) \ (g has_integral y) (cbox a b)" by (metis assms has_integral_spike_interior) lemma integrable_spike_interior: assumes "\x. x \ box a b \ g x = f x" and "f integrable_on cbox a b" shows "g integrable_on cbox a b" using assms has_integral_spike_interior_eq by blast subsection \Integrability of continuous functions\ lemma operative_approximableI: fixes f :: "'b::euclidean_space \ 'a::banach" assumes "0 \ e" shows "operative conj True (\i. \g. (\x\i. norm (f x - g (x::'b)) \ e) \ g integrable_on i)" proof - interpret comm_monoid conj True by standard auto show ?thesis proof (standard, safe) fix a b :: 'b show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" if "box a b = {}" for a b using assms that by (metis content_eq_0_interior integrable_on_null interior_cbox norm_zero right_minus_eq) { fix c g and k :: 'b assume fg: "\x\cbox a b. norm (f x - g x) \ e" and g: "g integrable_on cbox a b" assume k: "k \ Basis" show "\g. (\x\cbox a b \ {x. x \ k \ c}. norm (f x - g x) \ e) \ g integrable_on cbox a b \ {x. x \ k \ c}" "\g. (\x\cbox a b \ {x. c \ x \ k}. norm (f x - g x) \ e) \ g integrable_on cbox a b \ {x. c \ x \ k}" using fg g k by auto } show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" if fg1: "\x\cbox a b \ {x. x \ k \ c}. norm (f x - g1 x) \ e" and g1: "g1 integrable_on cbox a b \ {x. x \ k \ c}" and fg2: "\x\cbox a b \ {x. c \ x \ k}. norm (f x - g2 x) \ e" and g2: "g2 integrable_on cbox a b \ {x. c \ x \ k}" and k: "k \ Basis" for c k g1 g2 proof - let ?g = "\x. if x\k = c then f x else if x\k \ c then g1 x else g2 x" show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" proof (intro exI conjI ballI) show "norm (f x - ?g x) \ e" if "x \ cbox a b" for x by (auto simp: that assms fg1 fg2) show "?g integrable_on cbox a b" proof - have "?g integrable_on cbox a b \ {x. x \ k \ c}" "?g integrable_on cbox a b \ {x. x \ k \ c}" by(rule integrable_spike[OF _ negligible_standard_hyperplane[of k c]], use k g1 g2 in auto)+ with has_integral_split[OF _ _ k] show ?thesis unfolding integrable_on_def by blast qed qed qed qed qed lemma comm_monoid_set_F_and: "comm_monoid_set.F (\) True f s \ (finite s \ (\x\s. f x))" proof - interpret bool: comm_monoid_set \(\)\ True .. show ?thesis by (induction s rule: infinite_finite_induct) auto qed lemma approximable_on_division: fixes f :: "'b::euclidean_space \ 'a::banach" assumes "0 \ e" and d: "d division_of (cbox a b)" and f: "\i\d. \g. (\x\i. norm (f x - g x) \ e) \ g integrable_on i" obtains g where "\x\cbox a b. norm (f x - g x) \ e" "g integrable_on cbox a b" proof - interpret operative conj True "\i. \g. (\x\i. norm (f x - g (x::'b)) \ e) \ g integrable_on i" using \0 \ e\ by (rule operative_approximableI) from f local.division [OF d] that show thesis by auto qed lemma integrable_continuous: fixes f :: "'b::euclidean_space \ 'a::banach" assumes "continuous_on (cbox a b) f" shows "f integrable_on cbox a b" proof (rule integrable_uniform_limit) fix e :: real assume e: "e > 0" then obtain d where "0 < d" and d: "\x x'. \x \ cbox a b; x' \ cbox a b; dist x' x < d\ \ dist (f x') (f x) < e" using compact_uniformly_continuous[OF assms compact_cbox] unfolding uniformly_continuous_on_def by metis obtain p where ptag: "p tagged_division_of cbox a b" and finep: "(\x. ball x d) fine p" using fine_division_exists[OF gauge_ball[OF \0 < d\], of a b] . have *: "\i\snd ` p. \g. (\x\i. norm (f x - g x) \ e) \ g integrable_on i" proof (safe, unfold snd_conv) fix x l assume as: "(x, l) \ p" obtain a b where l: "l = cbox a b" using as ptag by blast then have x: "x \ cbox a b" using as ptag by auto show "\g. (\x\l. norm (f x - g x) \ e) \ g integrable_on l" proof (intro exI conjI strip) show "(\y. f x) integrable_on l" unfolding integrable_on_def l by blast next fix y assume y: "y \ l" then have "y \ ball x d" using as finep by fastforce then show "norm (f y - f x) \ e" using d x y as l by (metis dist_commute dist_norm less_imp_le mem_ball ptag subsetCE tagged_division_ofD(3)) qed qed from e have "e \ 0" by auto from approximable_on_division[OF this division_of_tagged_division[OF ptag] *] show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" by metis qed lemma integrable_continuous_interval: fixes f :: "'b::ordered_euclidean_space \ 'a::banach" assumes "continuous_on {a..b} f" shows "f integrable_on {a..b}" by (metis assms integrable_continuous interval_cbox) lemmas integrable_continuous_real = integrable_continuous_interval[where 'b=real] lemma integrable_continuous_closed_segment: fixes f :: "real \ 'a::banach" assumes "continuous_on (closed_segment a b) f" shows "f integrable_on (closed_segment a b)" by (metis assms closed_segment_eq_real_ivl integrable_continuous_interval) subsection \Specialization of additivity to one dimension\ subsection \A useful lemma allowing us to factor out the content size\ lemma has_integral_factor_content: "(f has_integral i) (cbox a b) \ (\e>0. \d. gauge d \ (\p. p tagged_division_of (cbox a b) \ d fine p \ norm (sum (\(x,k). content k *\<^sub>R f x) p - i) \ e * content (cbox a b)))" proof (cases "content (cbox a b) = 0") case True have "\e p. p tagged_division_of cbox a b \ norm ((\(x, k)\p. content k *\<^sub>R f x)) \ e * content (cbox a b)" unfolding sum_content_null[OF True] True by force moreover have "i = 0" if "\e. e > 0 \ \d. gauge d \ (\p. p tagged_division_of cbox a b \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x) - i) \ e * content (cbox a b))" using that [of 1] by (force simp add: True sum_content_null[OF True] intro: fine_division_exists[of _ a b]) ultimately show ?thesis unfolding has_integral_null_eq[OF True] by (force simp add: ) next case False then have F: "0 < content (cbox a b)" using zero_less_measure_iff by blast let ?P = "\e opp. \d. gauge d \ (\p. p tagged_division_of (cbox a b) \ d fine p \ opp (norm ((\(x, k)\p. content k *\<^sub>R f x) - i)) e)" show ?thesis proof (subst has_integral, safe) fix e :: real assume e: "e > 0" show "?P (e * content (cbox a b)) (\)" if \
[rule_format]: "\\>0. ?P \ (<)" using \
[of "e * content (cbox a b)"] by (meson F e less_imp_le mult_pos_pos) show "?P e (<)" if \
[rule_format]: "\\>0. ?P (\ * content (cbox a b)) (\)" using \
[of "e/2 / content (cbox a b)"] using F e by (force simp add: algebra_simps) qed qed lemma has_integral_factor_content_real: "(f has_integral i) {a..b::real} \ (\e>0. \d. gauge d \ (\p. p tagged_division_of {a..b} \ d fine p \ norm (sum (\(x,k). content k *\<^sub>R f x) p - i) \ e * content {a..b} ))" unfolding box_real[symmetric] by (rule has_integral_factor_content) subsection \Fundamental theorem of calculus\ lemma interval_bounds_real: fixes q b :: real assumes "a \ b" shows "Sup {a..b} = b" and "Inf {a..b} = a" using assms by auto theorem fundamental_theorem_of_calculus: fixes f :: "real \ 'a::banach" assumes "a \ b" and vecd: "\x. x \ {a..b} \ (f has_vector_derivative f' x) (at x within {a..b})" shows "(f' has_integral (f b - f a)) {a..b}" unfolding has_integral_factor_content box_real[symmetric] proof safe fix e :: real assume "e > 0" then have "\x. \d>0. x \ {a..b} \ (\y\{a..b}. norm (y-x) < d \ norm (f y - f x - (y-x) *\<^sub>R f' x) \ e * norm (y-x))" using vecd[unfolded has_vector_derivative_def has_derivative_within_alt] by blast then obtain d where d: "\x. 0 < d x" "\x y. \x \ {a..b}; y \ {a..b}; norm (y-x) < d x\ \ norm (f y - f x - (y-x) *\<^sub>R f' x) \ e * norm (y-x)" by metis show "\d. gauge d \ (\p. p tagged_division_of (cbox a b) \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f' x) - (f b - f a)) \ e * content (cbox a b))" proof (rule exI, safe) show "gauge (\x. ball x (d x))" using d(1) gauge_ball_dependent by blast next fix p assume ptag: "p tagged_division_of cbox a b" and finep: "(\x. ball x (d x)) fine p" have ba: "b - a = (\(x,K)\p. Sup K - Inf K)" "f b - f a = (\(x,K)\p. f(Sup K) - f(Inf K))" using additive_tagged_division_1[where f= "\x. x"] additive_tagged_division_1[where f= f] \a \ b\ ptag by auto have "norm (\(x, K) \ p. (content K *\<^sub>R f' x) - (f (Sup K) - f (Inf K))) \ (\n\p. e * (case n of (x, k) \ Sup k - Inf k))" proof (rule sum_norm_le,safe) fix x K assume "(x, K) \ p" then have "x \ K" and kab: "K \ cbox a b" using ptag by blast+ then obtain u v where k: "K = cbox u v" and "x \ K" and kab: "K \ cbox a b" using ptag \(x, K) \ p\ by auto have "u \ v" using \x \ K\ unfolding k by auto have ball: "\y\K. y \ ball x (d x)" using finep \(x, K) \ p\ by blast have "u \ K" "v \ K" by (simp_all add: \u \ v\ k) have "norm ((v - u) *\<^sub>R f' x - (f v - f u)) = norm (f u - f x - (u - x) *\<^sub>R f' x - (f v - f x - (v - x) *\<^sub>R f' x))" by (auto simp add: algebra_simps) also have "... \ norm (f u - f x - (u - x) *\<^sub>R f' x) + norm (f v - f x - (v - x) *\<^sub>R f' x)" by (rule norm_triangle_ineq4) finally have "norm ((v - u) *\<^sub>R f' x - (f v - f u)) \ norm (f u - f x - (u - x) *\<^sub>R f' x) + norm (f v - f x - (v - x) *\<^sub>R f' x)" . also have "\ \ e * norm (u - x) + e * norm (v - x)" proof (rule add_mono) show "norm (f u - f x - (u - x) *\<^sub>R f' x) \ e * norm (u - x)" proof (rule d) show "norm (u - x) < d x" using \u \ K\ ball by (auto simp add: dist_real_def) qed (use \x \ K\ \u \ K\ kab in auto) show "norm (f v - f x - (v - x) *\<^sub>R f' x) \ e * norm (v - x)" proof (rule d) show "norm (v - x) < d x" using \v \ K\ ball by (auto simp add: dist_real_def) qed (use \x \ K\ \v \ K\ kab in auto) qed also have "\ \ e * (Sup K - Inf K)" using \x \ K\ by (auto simp: k interval_bounds_real[OF \u \ v\] field_simps) finally show "norm (content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ e * (Sup K - Inf K)" using \u \ v\ by (simp add: k) qed with \a \ b\ show "norm ((\(x, K)\p. content K *\<^sub>R f' x) - (f b - f a)) \ e * content (cbox a b)" by (auto simp: ba split_def sum_subtractf [symmetric] sum_distrib_left) qed qed lemma has_complex_derivative_imp_has_vector_derivative: fixes f :: "complex \ complex" assumes "(f has_field_derivative f') (at (of_real a) within (cbox (of_real x) (of_real y)))" shows "((f o of_real) has_vector_derivative f') (at a within {x..y})" using has_derivative_in_compose[of of_real of_real a "{x..y}" f "(*) f'"] assms by (simp add: scaleR_conv_of_real ac_simps has_vector_derivative_def has_field_derivative_def o_def cbox_complex_of_real) lemma ident_has_integral: fixes a::real assumes "a \ b" shows "((\x. x) has_integral (b\<^sup>2 - a\<^sup>2)/2) {a..b}" proof - have "((\x. x) has_integral inverse 2 * b\<^sup>2 - inverse 2 * a\<^sup>2) {a..b}" unfolding power2_eq_square by (rule fundamental_theorem_of_calculus [OF assms] derivative_eq_intros | simp)+ then show ?thesis by (simp add: field_simps) qed lemma integral_ident [simp]: fixes a::real assumes "a \ b" shows "integral {a..b} (\x. x) = (if a \ b then (b\<^sup>2 - a\<^sup>2)/2 else 0)" by (metis assms ident_has_integral integral_unique) lemma ident_integrable_on: fixes a::real shows "(\x. x) integrable_on {a..b}" using continuous_on_id integrable_continuous_real by blast lemma integral_sin [simp]: fixes a::real assumes "a \ b" shows "integral {a..b} sin = cos a - cos b" proof - have "(sin has_integral (- cos b - - cos a)) {a..b}" proof (rule fundamental_theorem_of_calculus) show "((\a. - cos a) has_vector_derivative sin x) (at x within {a..b})" for x unfolding has_real_derivative_iff_has_vector_derivative [symmetric] by (rule derivative_eq_intros | force)+ qed (use assms in auto) then show ?thesis by (simp add: integral_unique) qed lemma integral_cos [simp]: fixes a::real assumes "a \ b" shows "integral {a..b} cos = sin b - sin a" proof - have "(cos has_integral (sin b - sin a)) {a..b}" proof (rule fundamental_theorem_of_calculus) show "(sin has_vector_derivative cos x) (at x within {a..b})" for x unfolding has_real_derivative_iff_has_vector_derivative [symmetric] by (rule derivative_eq_intros | force)+ qed (use assms in auto) then show ?thesis by (simp add: integral_unique) qed lemma integral_exp [simp]: fixes a::real assumes "a \ b" shows "integral {a..b} exp = exp b - exp a" by (meson DERIV_exp assms fundamental_theorem_of_calculus has_real_derivative_iff_has_vector_derivative has_vector_derivative_at_within integral_unique) lemma has_integral_sin_nx: "((\x. sin(real_of_int n * x)) has_integral 0) {-pi..pi}" proof (cases "n = 0") case False have "((\x. sin (n * x)) has_integral (- cos (n * pi)/n - - cos (n * - pi)/n)) {-pi..pi}" proof (rule fundamental_theorem_of_calculus) show "((\x. - cos (n * x) / n) has_vector_derivative sin (n * a)) (at a within {-pi..pi})" if "a \ {-pi..pi}" for a :: real using that False unfolding has_vector_derivative_def by (intro derivative_eq_intros | force)+ qed auto then show ?thesis by simp qed auto lemma integral_sin_nx: "integral {-pi..pi} (\x. sin(x * real_of_int n)) = 0" using has_integral_sin_nx [of n] by (force simp: mult.commute) lemma has_integral_cos_nx: "((\x. cos(real_of_int n * x)) has_integral (if n = 0 then 2 * pi else 0)) {-pi..pi}" proof (cases "n = 0") case True then show ?thesis using has_integral_const_real [of "1::real" "-pi" pi] by auto next case False have "((\x. cos (n * x)) has_integral (sin (n * pi)/n - sin (n * - pi)/n)) {-pi..pi}" proof (rule fundamental_theorem_of_calculus) show "((\x. sin (n * x) / n) has_vector_derivative cos (n * x)) (at x within {-pi..pi})" if "x \ {-pi..pi}" for x :: real using that False unfolding has_vector_derivative_def by (intro derivative_eq_intros | force)+ qed auto with False show ?thesis by (simp add: mult.commute) qed lemma integral_cos_nx: "integral {-pi..pi} (\x. cos(x * real_of_int n)) = (if n = 0 then 2 * pi else 0)" using has_integral_cos_nx [of n] by (force simp: mult.commute) subsection \Taylor series expansion\ lemma mvt_integral: fixes f::"'a::real_normed_vector\'b::banach" assumes f'[derivative_intros]: "\x. x \ S \ (f has_derivative f' x) (at x within S)" assumes line_in: "\t. t \ {0..1} \ x + t *\<^sub>R y \ S" shows "f (x + y) - f x = integral {0..1} (\t. f' (x + t *\<^sub>R y) y)" proof - from assms have subset: "(\xa. x + xa *\<^sub>R y) ` {0..1} \ S" by auto note [derivative_intros] = has_derivative_subset[OF _ subset] has_derivative_in_compose[where f="(\xa. x + xa *\<^sub>R y)" and g = f] note [continuous_intros] = continuous_on_compose2[where f="(\xa. x + xa *\<^sub>R y)"] continuous_on_subset[OF _ subset] have "\t. t \ {0..1} \ ((\t. f (x + t *\<^sub>R y)) has_vector_derivative f' (x + t *\<^sub>R y) y) (at t within {0..1})" using assms by (auto simp: has_vector_derivative_def linear_cmul[OF has_derivative_linear[OF f'], symmetric] intro!: derivative_eq_intros) from fundamental_theorem_of_calculus[rule_format, OF _ this] show ?thesis by (auto intro!: integral_unique[symmetric]) qed lemma (in bounded_bilinear) sum_prod_derivatives_has_vector_derivative: assumes "p>0" and f0: "Df 0 = f" and Df: "\m t. m < p \ a \ t \ t \ b \ (Df m has_vector_derivative Df (Suc m) t) (at t within {a..b})" and g0: "Dg 0 = g" and Dg: "\m t. m < p \ a \ t \ t \ b \ (Dg m has_vector_derivative Dg (Suc m) t) (at t within {a..b})" and ivl: "a \ t" "t \ b" shows "((\t. \iR prod (Df i t) (Dg (p - Suc i) t)) has_vector_derivative prod (f t) (Dg p t) - (-1)^p *\<^sub>R prod (Df p t) (g t)) (at t within {a..b})" using assms proof cases assume p: "p \ 1" define p' where "p' = p - 2" from assms p have p': "{..i. i \ p' \ Suc (Suc p' - i) = (Suc (Suc p') - i)" by auto let ?f = "\i. (-1) ^ i *\<^sub>R (prod (Df i t) (Dg ((p - i)) t))" have "(\iR (prod (Df i t) (Dg (Suc (p - Suc i)) t) + prod (Df (Suc i) t) (Dg (p - Suc i) t))) = (\i\(Suc p'). ?f i - ?f (Suc i))" by (auto simp: algebra_simps p'(2) numeral_2_eq_2 * lessThan_Suc_atMost) also note sum_telescope finally have "(\iR (prod (Df i t) (Dg (Suc (p - Suc i)) t) + prod (Df (Suc i) t) (Dg (p - Suc i) t))) = prod (f t) (Dg p t) - (- 1) ^ p *\<^sub>R prod (Df p t) (g t)" unfolding p'[symmetric] by (simp add: assms) thus ?thesis using assms by (auto intro!: derivative_eq_intros has_vector_derivative) qed (auto intro!: derivative_eq_intros has_vector_derivative) lemma fixes f::"real\'a::banach" assumes "p>0" and f0: "Df 0 = f" and Df: "\m t. m < p \ a \ t \ t \ b \ (Df m has_vector_derivative Df (Suc m) t) (at t within {a..b})" and ivl: "a \ b" defines "i \ \x. ((b - x) ^ (p - 1) / fact (p - 1)) *\<^sub>R Df p x" shows Taylor_has_integral: "(i has_integral f b - (\iR Df i a)) {a..b}" and Taylor_integral: "f b = (\iR Df i a) + integral {a..b} i" and Taylor_integrable: "i integrable_on {a..b}" proof goal_cases case 1 interpret bounded_bilinear "scaleR::real\'a\'a" by (rule bounded_bilinear_scaleR) define g where "g s = (b - s)^(p - 1)/fact (p - 1)" for s define Dg where [abs_def]: "Dg n s = (if n < p then (-1)^n * (b - s)^(p - 1 - n) / fact (p - 1 - n) else 0)" for n s have g0: "Dg 0 = g" using \p > 0\ by (auto simp add: Dg_def field_split_simps g_def split: if_split_asm) { fix m assume "p > Suc m" hence "p - Suc m = Suc (p - Suc (Suc m))" by auto hence "real (p - Suc m) * fact (p - Suc (Suc m)) = fact (p - Suc m)" by auto } note fact_eq = this have Dg: "\m t. m < p \ a \ t \ t \ b \ (Dg m has_vector_derivative Dg (Suc m) t) (at t within {a..b})" unfolding Dg_def by (auto intro!: derivative_eq_intros simp: has_vector_derivative_def fact_eq field_split_simps) let ?sum = "\t. \iR Dg i t *\<^sub>R Df (p - Suc i) t" from sum_prod_derivatives_has_vector_derivative[of _ Dg _ _ _ Df, OF \p > 0\ g0 Dg f0 Df] have deriv: "\t. a \ t \ t \ b \ (?sum has_vector_derivative g t *\<^sub>R Df p t - (- 1) ^ p *\<^sub>R Dg p t *\<^sub>R f t) (at t within {a..b})" by auto from fundamental_theorem_of_calculus[rule_format, OF \a \ b\ deriv] have "(i has_integral ?sum b - ?sum a) {a..b}" using atLeastatMost_empty'[simp del] by (simp add: i_def g_def Dg_def) also have one: "(- 1) ^ p' * (- 1) ^ p' = (1::real)" "{.. {i. p = Suc i} = {p - 1}" for p' using \p > 0\ by (auto simp: power_mult_distrib) then have "?sum b = f b" using Suc_pred'[OF \p > 0\] by (simp add: diff_eq_eq Dg_def power_0_left le_Suc_eq if_distrib if_distribR sum.If_cases f0) also have "?sum a = (\iR Df i a)" proof (rule sum.reindex_cong) have "\i. i < p \ \jp>0\ diff_Suc_less diff_diff_cancel less_or_eq_imp_le) then show "{..x. p - x - 1) ` {..Only need trivial subintervals if the interval itself is trivial\ proposition division_of_nontrivial: fixes \ :: "'a::euclidean_space set set" assumes sdiv: "\ division_of (cbox a b)" and cont0: "content (cbox a b) \ 0" shows "{k. k \ \ \ content k \ 0} division_of (cbox a b)" using sdiv proof (induction "card \" arbitrary: \ rule: less_induct) case less note \ = division_ofD[OF less.prems] { presume *: "{k \ \. content k \ 0} \ \ \ ?case" then show ?case using less.prems by fastforce } assume noteq: "{k \ \. content k \ 0} \ \" then obtain K c d where "K \ \" and contk: "content K = 0" and keq: "K = cbox c d" using \(4) by blast then have "card \ > 0" unfolding card_gt_0_iff using less by auto then have card: "card (\ - {K}) < card \" using less \K \ \\ by (simp add: \(1)) have closed: "closed (\(\ - {K}))" using less.prems by auto have "x islimpt \(\ - {K})" if "x \ K" for x unfolding islimpt_approachable proof (intro allI impI) fix e::real assume "e > 0" obtain i where i: "c\i = d\i" "i\Basis" using contk \(3) [OF \K \ \\] unfolding box_ne_empty keq by (meson content_eq_0 dual_order.antisym) then have xi: "x\i = d\i" using \x \ K\ unfolding keq mem_box by (metis antisym) define y where "y = (\j\Basis. (if j = i then if c\i \ (a\i + b\i)/2 then c\i + min e (b\i - c\i)/2 else c\i - min e (c\i - a\i)/2 else x\j) *\<^sub>R j)" show "\x'\\(\ - {K}). x' \ x \ dist x' x < e" proof (intro bexI conjI) have "d \ cbox c d" using \(3)[OF \K \ \\] by (simp add: box_ne_empty(1) keq mem_box(2)) then have "d \ cbox a b" using \(2)[OF \K \ \\] by (auto simp: keq) then have di: "a \ i \ d \ i \ d \ i \ b \ i" using \i \ Basis\ mem_box(2) by blast then have xyi: "y\i \ x\i" unfolding y_def i xi using \e > 0\ cont0 \i \ Basis\ by (auto simp: content_eq_0 elim!: ballE[of _ _ i]) then show "y \ x" unfolding euclidean_eq_iff[where 'a='a] using i by auto have "norm (y-x) \ (\b\Basis. \(y - x) \ b\)" by (rule norm_le_l1) also have "... = \(y - x) \ i\ + (\b \ Basis - {i}. \(y - x) \ b\)" by (meson finite_Basis i(2) sum.remove) also have "... < e + sum (\i. 0) Basis" proof (rule add_less_le_mono) show "\(y-x) \ i\ < e" using di \e > 0\ y_def i xi by (auto simp: inner_simps) show "(\i\Basis - {i}. \(y-x) \ i\) \ (\i\Basis. 0)" unfolding y_def by (auto simp: inner_simps) qed finally have "norm (y-x) < e + sum (\i. 0) Basis" . then show "dist y x < e" unfolding dist_norm by auto have "y \ K" unfolding keq mem_box using i(1) i(2) xi xyi by fastforce moreover have "y \ \\" using subsetD[OF \(2)[OF \K \ \\] \x \ K\] \e > 0\ di i by (auto simp: \ mem_box y_def field_simps elim!: ballE[of _ _ i]) ultimately show "y \ \(\ - {K})" by auto qed qed then have "K \ \(\ - {K})" using closed closed_limpt by blast then have "\(\ - {K}) = cbox a b" unfolding \(6)[symmetric] by auto then have "\ - {K} division_of cbox a b" by (metis Diff_subset less.prems division_of_subset \(6)) then have "{ka \ \ - {K}. content ka \ 0} division_of (cbox a b)" using card less.hyps by blast moreover have "{ka \ \ - {K}. content ka \ 0} = {K \ \. content K \ 0}" using contk by auto ultimately show ?case by auto qed subsection \Integrability on subintervals\ lemma operative_integrableI: fixes f :: "'b::euclidean_space \ 'a::banach" assumes "0 \ e" shows "operative conj True (\i. f integrable_on i)" proof - interpret comm_monoid conj True proof qed show ?thesis proof show "\a b. box a b = {} \ (f integrable_on cbox a b) = True" by (simp add: content_eq_0_interior integrable_on_null) show "\a b c k. k \ Basis \ (f integrable_on cbox a b) \ (f integrable_on cbox a b \ {x. x \ k \ c} \ f integrable_on cbox a b \ {x. c \ x \ k})" unfolding integrable_on_def by (auto intro!: has_integral_split) qed qed lemma integrable_subinterval: fixes f :: "'b::euclidean_space \ 'a::banach" assumes f: "f integrable_on cbox a b" and cd: "cbox c d \ cbox a b" shows "f integrable_on cbox c d" proof - interpret operative conj True "\i. f integrable_on i" using order_refl by (rule operative_integrableI) show ?thesis by (metis cd division division_of_finite empty f partial_division_extend_1 remove) qed lemma integrable_subinterval_real: fixes f :: "real \ 'a::banach" assumes "f integrable_on {a..b}" and "{c..d} \ {a..b}" shows "f integrable_on {c..d}" by (metis assms box_real(2) integrable_subinterval) subsection \Combining adjacent intervals in 1 dimension\ lemma has_integral_combine: fixes a b c :: real and j :: "'a::banach" assumes "a \ c" and "c \ b" and ac: "(f has_integral i) {a..c}" and cb: "(f has_integral j) {c..b}" shows "(f has_integral (i + j)) {a..b}" proof - interpret operative_real "lift_option plus" "Some 0" "\i. if f integrable_on i then Some (integral i f) else None" using operative_integralI by (rule operative_realI) from \a \ c\ \c \ b\ ac cb coalesce_less_eq have *: "lift_option (+) (if f integrable_on {a..c} then Some (integral {a..c} f) else None) (if f integrable_on {c..b} then Some (integral {c..b} f) else None) = (if f integrable_on {a..b} then Some (integral {a..b} f) else None)" by (auto simp: split: if_split_asm) then have "f integrable_on cbox a b" using ac cb by (auto split: if_split_asm) with * show ?thesis using ac cb by (auto simp add: integrable_on_def integral_unique split: if_split_asm) qed lemma integral_combine: fixes f :: "real \ 'a::banach" assumes "a \ c" and "c \ b" and ab: "f integrable_on {a..b}" shows "integral {a..c} f + integral {c..b} f = integral {a..b} f" proof - have "(f has_integral integral {a..c} f) {a..c}" using ab \c \ b\ integrable_subinterval_real by fastforce moreover have "(f has_integral integral {c..b} f) {c..b}" using ab \a \ c\ integrable_subinterval_real by fastforce ultimately show ?thesis by (smt (verit, best) assms has_integral_combine integral_unique) qed lemma integrable_combine: fixes f :: "real \ 'a::banach" assumes "a \ c" and "c \ b" and "f integrable_on {a..c}" and "f integrable_on {c..b}" shows "f integrable_on {a..b}" using assms has_integral_combine by blast lemma integral_minus_sets: fixes f::"real \ 'a::banach" shows "c \ a \ c \ b \ f integrable_on {c .. max a b} \ integral {c .. a} f - integral {c .. b} f = (if a \ b then - integral {a .. b} f else integral {b .. a} f)" using integral_combine[of c a b f] integral_combine[of c b a f] by (auto simp: algebra_simps max_def) lemma integral_minus_sets': fixes f::"real \ 'a::banach" shows "c \ a \ c \ b \ f integrable_on {min a b .. c} \ integral {a .. c} f - integral {b .. c} f = (if a \ b then integral {a .. b} f else - integral {b .. a} f)" using integral_combine[of b a c f] integral_combine[of a b c f] by (auto simp: algebra_simps min_def) subsection \Reduce integrability to "local" integrability\ lemma integrable_on_little_subintervals: fixes f :: "'b::euclidean_space \ 'a::banach" assumes "\x\cbox a b. \d>0. \u v. x \ cbox u v \ cbox u v \ ball x d \ cbox u v \ cbox a b \ f integrable_on cbox u v" shows "f integrable_on cbox a b" proof - interpret operative conj True "\i. f integrable_on i" using order_refl by (rule operative_integrableI) have "\x. \d>0. x\cbox a b \ (\u v. x \ cbox u v \ cbox u v \ ball x d \ cbox u v \ cbox a b \ f integrable_on cbox u v)" using assms by (metis zero_less_one) then obtain d where d: "\x. 0 < d x" "\x u v. \x \ cbox a b; x \ cbox u v; cbox u v \ ball x (d x); cbox u v \ cbox a b\ \ f integrable_on cbox u v" by metis obtain p where p: "p tagged_division_of cbox a b" "(\x. ball x (d x)) fine p" using fine_division_exists[OF gauge_ball_dependent,of d a b] d(1) by blast then have sndp: "snd ` p division_of cbox a b" by (metis division_of_tagged_division) have "f integrable_on k" if "(x, k) \ p" for x k using tagged_division_ofD(2-4)[OF p(1) that] fineD[OF p(2) that] d[of x] by auto then show ?thesis unfolding division [symmetric, OF sndp] comm_monoid_set_F_and by auto qed subsection \Second FTC or existence of antiderivative\ lemma integrable_const[intro]: "(\x. c) integrable_on cbox a b" unfolding integrable_on_def by blast lemma integral_has_vector_derivative_continuous_at: fixes f :: "real \ 'a::banach" assumes f: "f integrable_on {a..b}" and x: "x \ {a..b} - S" and "finite S" and fx: "continuous (at x within ({a..b} - S)) f" shows "((\u. integral {a..u} f) has_vector_derivative f x) (at x within ({a..b} - S))" proof - let ?I = "\a b. integral {a..b} f" { fix e::real assume "e > 0" obtain d where "d>0" and d: "\x'. \x' \ {a..b} - S; \x' - x\ < d\ \ norm(f x' - f x) \ e" using \e>0\ fx by (auto simp: continuous_within_eps_delta dist_norm less_imp_le) have "norm (integral {a..y} f - integral {a..x} f - (y-x) *\<^sub>R f x) \ e * \y - x\" (is "?lhs \ ?rhs") if y: "y \ {a..b} - S" and yx: "\y - x\ < d" for y proof (cases "y < x") case False have "f integrable_on {a..y}" using f y by (simp add: integrable_subinterval_real) then have Idiff: "?I a y - ?I a x = ?I x y" using False x by (simp add: algebra_simps integral_combine) have fux_int: "((\u. f u - f x) has_integral integral {x..y} f - (y-x) *\<^sub>R f x) {x..y}" proof (rule has_integral_diff) show "(f has_integral integral {x..y} f) {x..y}" using x y by (auto intro: integrable_integral [OF integrable_subinterval_real [OF f]]) show "((\u. f x) has_integral (y - x) *\<^sub>R f x) {x..y}" using has_integral_const_real [of "f x" x y] False by simp qed have "?lhs \ e * content {x..y}" using yx False d x y \e>0\ assms by (intro has_integral_bound_real[where f="(\u. f u - f x)"]) (auto simp: Idiff fux_int) also have "... \ ?rhs" using False by auto finally show ?thesis . next case True have "f integrable_on {a..x}" using f x by (simp add: integrable_subinterval_real) then have Idiff: "?I a x - ?I a y = ?I y x" using True x y by (simp add: algebra_simps integral_combine) have fux_int: "((\u. f u - f x) has_integral integral {y..x} f - (x - y) *\<^sub>R f x) {y..x}" proof (rule has_integral_diff) show "(f has_integral integral {y..x} f) {y..x}" using x y by (auto intro: integrable_integral [OF integrable_subinterval_real [OF f]]) show "((\u. f x) has_integral (x - y) *\<^sub>R f x) {y..x}" using has_integral_const_real [of "f x" y x] True by simp qed have "norm (integral {a..x} f - integral {a..y} f - (x - y) *\<^sub>R f x) \ e * content {y..x}" using yx True d x y \e>0\ assms by (intro has_integral_bound_real[where f="(\u. f u - f x)"]) (auto simp: Idiff fux_int) also have "... \ e * \y - x\" using True by auto finally have "norm (integral {a..x} f - integral {a..y} f - (x - y) *\<^sub>R f x) \ e * \y - x\" . then show ?thesis by (simp add: algebra_simps norm_minus_commute) qed then have "\d>0. \y\{a..b} - S. \y - x\ < d \ norm (integral {a..y} f - integral {a..x} f - (y-x) *\<^sub>R f x) \ e * \y - x\" using \d>0\ by blast } then show ?thesis by (simp add: has_vector_derivative_def has_derivative_within_alt bounded_linear_scaleR_left) qed lemma integral_has_vector_derivative: fixes f :: "real \ 'a::banach" assumes "continuous_on {a..b} f" and "x \ {a..b}" shows "((\u. integral {a..u} f) has_vector_derivative f(x)) (at x within {a..b})" using assms integral_has_vector_derivative_continuous_at [OF integrable_continuous_real] by (fastforce simp: continuous_on_eq_continuous_within) lemma integral_has_real_derivative: assumes "continuous_on {a..b} g" assumes "t \ {a..b}" shows "((\x. integral {a..x} g) has_real_derivative g t) (at t within {a..b})" using integral_has_vector_derivative[of a b g t] assms by (auto simp: has_real_derivative_iff_has_vector_derivative) lemma antiderivative_continuous: fixes q b :: real assumes "continuous_on {a..b} f" obtains g where "\x. x \ {a..b} \ (g has_vector_derivative (f x::_::banach)) (at x within {a..b})" using integral_has_vector_derivative[OF assms] by auto subsection \Combined fundamental theorem of calculus\ lemma antiderivative_integral_continuous: fixes f :: "real \ 'a::banach" assumes "continuous_on {a..b} f" obtains g where "\u\{a..b}. \v \ {a..b}. u \ v \ (f has_integral (g v - g u)) {u..v}" proof - obtain g where g: "\x. x \ {a..b} \ (g has_vector_derivative f x) (at x within {a..b})" using antiderivative_continuous[OF assms] by metis have "(f has_integral g v - g u) {u..v}" if "u \ {a..b}" "v \ {a..b}" "u \ v" for u v proof - have "\x. x \ cbox u v \ (g has_vector_derivative f x) (at x within cbox u v)" by (metis atLeastAtMost_iff atLeastatMost_subset_iff box_real(2) g has_vector_derivative_within_subset subsetCE that(1,2)) then show ?thesis by (metis box_real(2) that(3) fundamental_theorem_of_calculus) qed then show ?thesis using that by blast qed subsection \General "twiddling" for interval-to-interval function image\ lemma has_integral_twiddle: assumes "0 < r" and hg: "\x. h(g x) = x" and gh: "\x. g(h x) = x" and contg: "\x. continuous (at x) g" and g: "\u v. \w z. g ` cbox u v = cbox w z" and h: "\u v. \w z. h ` cbox u v = cbox w z" and r: "\u v. content(g ` cbox u v) = r * content (cbox u v)" and intfi: "(f has_integral i) (cbox a b)" shows "((\x. f(g x)) has_integral (1 / r) *\<^sub>R i) (h ` cbox a b)" proof (cases "cbox a b = {}") case True then show ?thesis using intfi by auto next case False obtain w z where wz: "h ` cbox a b = cbox w z" using h by blast have inj: "inj g" "inj h" using hg gh injI by metis+ from h obtain ha hb where h_eq: "h ` cbox a b = cbox ha hb" by blast have "\d. gauge d \ (\p. p tagged_division_of h ` cbox a b \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e)" if "e > 0" for e proof - have "e * r > 0" using that \0 < r\ by simp with intfi[unfolded has_integral] obtain d where "gauge d" and d: "\p. p tagged_division_of cbox a b \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x) - i) < e * r" by metis define d' where "d' x = g -` d (g x)" for x show ?thesis proof (rule_tac x=d' in exI, safe) show "gauge d'" using \gauge d\ continuous_open_vimage[OF _ contg] by (auto simp: gauge_def d'_def) next fix p assume ptag: "p tagged_division_of h ` cbox a b" and finep: "d' fine p" note p = tagged_division_ofD[OF ptag] have gab: "g y \ cbox a b" if "y \ K" "(x, K) \ p" for x y K by (metis hg inj(2) inj_image_mem_iff p(3) subsetCE that that) have gimp: "(\(x,K). (g x, g ` K)) ` p tagged_division_of (cbox a b) \ d fine (\(x, k). (g x, g ` k)) ` p" unfolding tagged_division_of proof safe show "finite ((\(x, k). (g x, g ` k)) ` p)" using ptag by auto show "d fine (\(x, k). (g x, g ` k)) ` p" using finep unfolding fine_def d'_def by auto next fix x K assume xk: "(x, K) \ p" show "g x \ g ` K" using p(2)[OF xk] by auto show "\u v. g ` K = cbox u v" using p(4)[OF xk] using assms(5-6) by auto fix x' K' u assume xk': "(x', K') \ p" and u: "u \ interior (g ` K)" "u \ interior (g ` K')" have "interior K \ interior K' \ {}" proof assume "interior K \ interior K' = {}" moreover have "u \ g ` (interior K \ interior K')" using interior_image_subset[OF \inj g\ contg] u unfolding image_Int[OF inj(1)] by blast ultimately show False by blast qed then have same: "(x, K) = (x', K')" using ptag xk' xk by blast then show "g x = g x'" by auto show "g u \ g ` K'"if "u \ K" for u using that same by auto show "g u \ g ` K"if "u \ K'" for u using that same by auto next fix x assume "x \ cbox a b" then have "h x \ \{k. \x. (x, k) \ p}" using p(6) by auto then obtain X y where "h x \ X" "(y, X) \ p" by blast then show "x \ \{k. \x. (x, k) \ (\(x, k). (g x, g ` k)) ` p}" by clarsimp (metis (no_types, lifting) gh image_eqI pair_imageI) qed (use gab in auto) have *: "inj_on (\(x, k). (g x, g ` k)) p" using inj(1) unfolding inj_on_def by fastforce have "(\(x,K)\(\(y,L). (g y, g ` L)) ` p. content K *\<^sub>R f x) = (\u\p. case case u of (x,K) \ (g x, g ` K) of (y,L) \ content L *\<^sub>R f y)" by (metis (mono_tags, lifting) "*" sum.reindex_cong) also have "... = (\(x,K)\p. r *\<^sub>R content K *\<^sub>R f (g x))" using r by (auto intro!: * sum.cong simp: bij_betw_def dest!: p(4)) finally have "(\(x, K)\(\(x,K). (g x, g ` K)) ` p. content K *\<^sub>R f x) - i = r *\<^sub>R (\(x,K)\p. content K *\<^sub>R f (g x)) - i" by (simp add: scaleR_right.sum split_def) also have "\ = r *\<^sub>R ((\(x,K)\p. content K *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i)" using \0 < r\ by (auto simp: scaleR_diff_right) finally show "norm ((\(x,K)\p. content K *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e" using d[OF gimp] \0 < r\ by auto qed qed then show ?thesis by (auto simp: h_eq has_integral) qed subsection \Special case of a basic affine transformation\ lemma AE_lborel_inner_neq: assumes k: "k \ Basis" shows "AE x in lborel. x \ k \ c" proof - interpret finite_product_sigma_finite "\_. lborel" Basis proof qed simp have "emeasure lborel {x\space lborel. x \ k = c} = emeasure (\\<^sub>M j::'a\Basis. lborel) (\\<^sub>E j\Basis. if j = k then {c} else UNIV)" using k by (auto simp add: lborel_eq[where 'a='a] emeasure_distr intro!: arg_cong2[where f=emeasure]) (auto simp: space_PiM PiE_iff extensional_def split: if_split_asm) also have "\ = (\j\Basis. emeasure lborel (if j = k then {c} else UNIV))" by (intro measure_times) auto also have "\ = 0" by (intro prod_zero bexI[OF _ k]) auto finally show ?thesis by (subst AE_iff_measurable[OF _ refl]) auto qed lemma content_image_stretch_interval: fixes m :: "'a::euclidean_space \ real" defines "s f x \ (\k::'a\Basis. (f k * (x\k)) *\<^sub>R k)" shows "content (s m ` cbox a b) = \\k\Basis. m k\ * content (cbox a b)" proof cases have s[measurable]: "s f \ borel \\<^sub>M borel" for f by (auto simp: s_def[abs_def]) assume m: "\k\Basis. m k \ 0" then have s_comp_s: "s (\k. 1 / m k) \ s m = id" "s m \ s (\k. 1 / m k) = id" by (auto simp: s_def[abs_def] fun_eq_iff euclidean_representation) then have "inv (s (\k. 1 / m k)) = s m" "bij (s (\k. 1 / m k))" by (auto intro: inv_unique_comp o_bij) then have eq: "s m ` cbox a b = s (\k. 1 / m k) -` cbox a b" using bij_vimage_eq_inv_image[OF \bij (s (\k. 1 / m k))\, of "cbox a b"] by auto show ?thesis using m unfolding eq measure_def by (subst lborel_affine_euclidean[where c=m and t=0]) (simp_all add: emeasure_density measurable_sets_borel[OF s] abs_prod nn_integral_cmult s_def[symmetric] emeasure_distr vimage_comp s_comp_s enn2real_mult prod_nonneg) next assume "\ (\k\Basis. m k \ 0)" then obtain k where k: "k \ Basis" "m k = 0" by auto then have [simp]: "(\k\Basis. m k) = 0" by (intro prod_zero) auto have "emeasure lborel {x\space lborel. x \ s m ` cbox a b} = 0" proof (rule emeasure_eq_0_AE) show "AE x in lborel. x \ s m ` cbox a b" using AE_lborel_inner_neq[OF \k\Basis\] proof eventually_elim show "x \ k \ 0 \ x \ s m ` cbox a b " for x using k by (auto simp: s_def[abs_def] cbox_def) qed qed then show ?thesis by (simp add: measure_def) qed lemma interval_image_affinity_interval: "\u v. (\x. m *\<^sub>R (x::'a::euclidean_space) + c) ` cbox a b = cbox u v" unfolding image_affinity_cbox by auto lemma content_image_affinity_cbox: "content((\x::'a::euclidean_space. m *\<^sub>R x + c) ` cbox a b) = \m\ ^ DIM('a) * content (cbox a b)" (is "?l = ?r") proof (cases "cbox a b = {}") case True then show ?thesis by simp next case False show ?thesis proof (cases "m \ 0") case True with \cbox a b \ {}\ have "cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c) \ {}" by (simp add: box_ne_empty inner_left_distrib mult_left_mono) moreover from True have *: "\i. (m *\<^sub>R b + c) \ i - (m *\<^sub>R a + c) \ i = m *\<^sub>R (b-a) \ i" by (simp add: inner_simps field_simps) ultimately show ?thesis by (simp add: image_affinity_cbox True content_cbox' prod.distrib inner_diff_left) next case False with \cbox a b \ {}\ have "cbox (m *\<^sub>R b + c) (m *\<^sub>R a + c) \ {}" by (simp add: box_ne_empty inner_left_distrib mult_left_mono) moreover from False have *: "\i. (m *\<^sub>R a + c) \ i - (m *\<^sub>R b + c) \ i = (-m) *\<^sub>R (b-a) \ i" by (simp add: inner_simps field_simps) ultimately show ?thesis using False by (simp add: image_affinity_cbox content_cbox' prod.distrib[symmetric] inner_diff_left flip: prod_constant) qed qed lemma has_integral_affinity: fixes a :: "'a::euclidean_space" assumes "(f has_integral i) (cbox a b)" and "m \ 0" shows "((\x. f(m *\<^sub>R x + c)) has_integral (1 / (\m\ ^ DIM('a))) *\<^sub>R i) ((\x. (1 / m) *\<^sub>R x + -((1 / m) *\<^sub>R c)) ` cbox a b)" proof (rule has_integral_twiddle) show "\w z. (\x. (1 / m) *\<^sub>R x + - ((1 / m) *\<^sub>R c)) ` cbox u v = cbox w z" "\w z. (\x. m *\<^sub>R x + c) ` cbox u v = cbox w z" for u v using interval_image_affinity_interval by blast+ show "content ((\x. m *\<^sub>R x + c) ` cbox u v) = \m\ ^ DIM('a) * content (cbox u v)" for u v using content_image_affinity_cbox by blast qed (use assms zero_less_power in \auto simp: field_simps\) lemma integrable_affinity: assumes "f integrable_on cbox a b" and "m \ 0" shows "(\x. f(m *\<^sub>R x + c)) integrable_on ((\x. (1 / m) *\<^sub>R x + -((1/m) *\<^sub>R c)) ` cbox a b)" using has_integral_affinity assms unfolding integrable_on_def by blast lemmas has_integral_affinity01 = has_integral_affinity [of _ _ 0 "1::real", simplified] lemma integrable_on_affinity: assumes "m \ 0" "f integrable_on (cbox a b)" shows "(\x. f (m *\<^sub>R x + c)) integrable_on ((\x. (1 / m) *\<^sub>R x - ((1 / m) *\<^sub>R c)) ` cbox a b)" proof - from assms obtain I where "(f has_integral I) (cbox a b)" by (auto simp: integrable_on_def) from has_integral_affinity[OF this assms(1), of c] show ?thesis by (auto simp: integrable_on_def) qed lemma has_integral_cmul_iff: assumes "c \ 0" shows "((\x. c *\<^sub>R f x) has_integral (c *\<^sub>R I)) A \ (f has_integral I) A" using assms has_integral_cmul[of f I A c] has_integral_cmul[of "\x. c *\<^sub>R f x" "c *\<^sub>R I" A "inverse c"] by (auto simp: field_simps) lemma has_integral_cmul_iff': assumes "c \ 0" shows "((\x. c *\<^sub>R f x) has_integral I) A \ (f has_integral I /\<^sub>R c) A" using assms by (metis divideR_right has_integral_cmul_iff) lemma has_integral_affinity': fixes a :: "'a::euclidean_space" assumes "(f has_integral i) (cbox a b)" and "m > 0" shows "((\x. f(m *\<^sub>R x + c)) has_integral (i /\<^sub>R m ^ DIM('a))) (cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m))" proof (cases "cbox a b = {}") case True hence "(cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)) = {}" using \m > 0\ unfolding box_eq_empty by (auto simp: algebra_simps) with True and assms show ?thesis by simp next case False have "((\x. f (m *\<^sub>R x + c)) has_integral (1 / \m\ ^ DIM('a)) *\<^sub>R i) ((\x. (1 / m) *\<^sub>R x + - ((1 / m) *\<^sub>R c)) ` cbox a b)" using assms by (intro has_integral_affinity) auto also have "((\x. (1 / m) *\<^sub>R x + - ((1 / m) *\<^sub>R c)) ` cbox a b) = ((\x. - ((1 / m) *\<^sub>R c) + x) ` (\x. (1 / m) *\<^sub>R x) ` cbox a b)" by (simp add: image_image algebra_simps) also have "(\x. (1 / m) *\<^sub>R x) ` cbox a b = cbox ((1 / m) *\<^sub>R a) ((1 / m) *\<^sub>R b)" using \m > 0\ False by (subst image_smult_cbox) simp_all also have "(\x. - ((1 / m) *\<^sub>R c) + x) ` \ = cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)" by (subst cbox_translation [symmetric]) (simp add: field_simps vector_add_divide_simps) finally show ?thesis using \m > 0\ by (simp add: field_simps) qed lemma has_integral_affinity_iff: fixes f :: "'a :: euclidean_space \ 'b :: real_normed_vector" assumes "m > 0" shows "((\x. f (m *\<^sub>R x + c)) has_integral (I /\<^sub>R m ^ DIM('a))) (cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)) \ (f has_integral I) (cbox a b)" (is "?lhs = ?rhs") proof assume ?lhs from has_integral_affinity'[OF this, of "1 / m" "-c /\<^sub>R m"] and \m > 0\ show ?rhs by (simp add: vector_add_divide_simps) (simp add: field_simps) next assume ?rhs from has_integral_affinity'[OF this, of m c] and \m > 0\ show ?lhs by simp qed subsection \Special case of stretching coordinate axes separately\ lemma has_integral_stretch: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "(f has_integral i) (cbox a b)" and "\k\Basis. m k \ 0" shows "((\x. f (\k\Basis. (m k * (x\k))*\<^sub>R k)) has_integral ((1/ \prod m Basis\) *\<^sub>R i)) ((\x. (\k\Basis. (1 / m k * (x\k))*\<^sub>R k)) ` cbox a b)" apply (rule has_integral_twiddle[where f=f]) unfolding zero_less_abs_iff content_image_stretch_interval unfolding image_stretch_interval empty_as_interval euclidean_eq_iff[where 'a='a] using assms by auto lemma has_integral_stretch_real: fixes f :: "real \ 'b::real_normed_vector" assumes "(f has_integral i) {a..b}" and "m \ 0" shows "((\x. f (m * x)) has_integral (1 / \m\) *\<^sub>R i) ((\x. x / m) ` {a..b})" using has_integral_stretch [of f i a b "\b. m"] assms by simp lemma integrable_stretch: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "f integrable_on cbox a b" and "\k\Basis. m k \ 0" shows "(\x::'a. f (\k\Basis. (m k * (x\k))*\<^sub>R k)) integrable_on ((\x. \k\Basis. (1 / m k * (x\k))*\<^sub>R k) ` cbox a b)" using assms unfolding integrable_on_def by (force dest: has_integral_stretch) lemma vec_lambda_eq_sum: "(\ k. f k (x $ k)) = (\k\Basis. (f (axis_index k) (x \ k)) *\<^sub>R k)" (is "?lhs = ?rhs") proof - have "?lhs = (\ k. f k (x \ axis k 1))" by (simp add: cart_eq_inner_axis) also have "... = (\u\UNIV. f u (x \ axis u 1) *\<^sub>R axis u 1)" by (simp add: vec_eq_iff axis_def if_distrib cong: if_cong) also have "... = ?rhs" by (simp add: Basis_vec_def UNION_singleton_eq_range sum.reindex axis_eq_axis inj_on_def) finally show ?thesis . qed lemma has_integral_stretch_cart: fixes m :: "'n::finite \ real" assumes f: "(f has_integral i) (cbox a b)" and m: "\k. m k \ 0" shows "((\x. f(\ k. m k * x$k)) has_integral i /\<^sub>R \prod m UNIV\) ((\x. \ k. x$k / m k) ` (cbox a b))" proof - have *: "\k:: real^'n \ Basis. m (axis_index k) \ 0" using axis_index by (simp add: m) have eqp: "(\k:: real^'n \ Basis. m (axis_index k)) = prod m UNIV" by (simp add: Basis_vec_def UNION_singleton_eq_range prod.reindex axis_eq_axis inj_on_def) show ?thesis using has_integral_stretch [OF f *] vec_lambda_eq_sum [where f="\i x. m i * x"] vec_lambda_eq_sum [where f="\i x. x / m i"] by (simp add: field_simps eqp) qed lemma image_stretch_interval_cart: fixes m :: "'n::finite \ real" shows "(\x. \ k. m k * x$k) ` cbox a b = (if cbox a b = {} then {} else cbox (\ k. min (m k * a$k) (m k * b$k)) (\ k. max (m k * a$k) (m k * b$k)))" proof - have *: "(\k\Basis. min (m (axis_index k) * (a \ k)) (m (axis_index k) * (b \ k)) *\<^sub>R k) = (\ k. min (m k * a $ k) (m k * b $ k))" "(\k\Basis. max (m (axis_index k) * (a \ k)) (m (axis_index k) * (b \ k)) *\<^sub>R k) = (\ k. max (m k * a $ k) (m k * b $ k))" apply (simp_all add: Basis_vec_def cart_eq_inner_axis UNION_singleton_eq_range sum.reindex axis_eq_axis inj_on_def) apply (simp_all add: vec_eq_iff axis_def if_distrib cong: if_cong) done show ?thesis by (simp add: vec_lambda_eq_sum [where f="\i x. m i * x"] image_stretch_interval eq_cbox *) qed subsection \even more special cases\ lemma uminus_interval_vector[simp]: fixes a b :: "'a::euclidean_space" shows "uminus ` cbox a b = cbox (-b) (-a)" proof - have "x \ uminus ` cbox a b" if "x \ cbox (- b) (- a)" for x by (smt (verit) add.inverse_inverse image_iff inner_minus_left mem_box(2) that) then show ?thesis by (auto simp: mem_box) qed lemma has_integral_reflect_lemma[intro]: assumes "(f has_integral i) (cbox a b)" shows "((\x. f(-x)) has_integral i) (cbox (-b) (-a))" using has_integral_affinity[OF assms, of "-1" 0] by auto lemma has_integral_reflect_lemma_real[intro]: assumes "(f has_integral i) {a..b::real}" shows "((\x. f(-x)) has_integral i) {-b .. -a}" by (metis has_integral_reflect_lemma interval_cbox assms) lemma has_integral_reflect[simp]: "((\x. f (-x)) has_integral i) (cbox (-b) (-a)) \ (f has_integral i) (cbox a b)" by (auto dest: has_integral_reflect_lemma) lemma has_integral_reflect_real[simp]: fixes a b::real shows "((\x. f (-x)) has_integral i) {-b..-a} \ (f has_integral i) {a..b}" by (metis has_integral_reflect interval_cbox) lemma integrable_reflect[simp]: "(\x. f(-x)) integrable_on cbox (-b) (-a) \ f integrable_on cbox a b" unfolding integrable_on_def by auto lemma integrable_reflect_real[simp]: "(\x. f(-x)) integrable_on {-b .. -a} \ f integrable_on {a..b::real}" unfolding box_real[symmetric] by (rule integrable_reflect) lemma integral_reflect[simp]: "integral (cbox (-b) (-a)) (\x. f (-x)) = integral (cbox a b) f" unfolding integral_def by auto lemma integral_reflect_real[simp]: "integral {-b .. -a} (\x. f (-x)) = integral {a..b::real} f" unfolding box_real[symmetric] by (rule integral_reflect) subsection \Stronger form of FCT; quite a tedious proof\ lemma split_minus[simp]: "(\(x, k). f x k) x - (\(x, k). g x k) x = (\(x, k). f x k - g x k) x" by (simp add: split_def) theorem fundamental_theorem_of_calculus_interior: fixes f :: "real \ 'a::real_normed_vector" assumes "a \ b" and contf: "continuous_on {a..b} f" and derf: "\x. x \ {a <..< b} \ (f has_vector_derivative f' x) (at x)" shows "(f' has_integral (f b - f a)) {a..b}" proof (cases "a = b") case True then have *: "cbox a b = {b}" "f b - f a = 0" by (auto simp add: order_antisym) with True show ?thesis by auto next case False with \a \ b\ have ab: "a < b" by arith show ?thesis unfolding has_integral_factor_content_real proof (intro allI impI) fix e :: real assume e: "e > 0" then have eba8: "(e * (b-a)) / 8 > 0" using ab by (auto simp add: field_simps) note derf_exp = derf[unfolded has_vector_derivative_def has_derivative_at_alt, THEN conjunct2, rule_format] have bounded: "\x. x \ {a<.. bounded_linear (\u. u *\<^sub>R f' x)" by (simp add: bounded_linear_scaleR_left) have "\x \ box a b. \d>0. \y. norm (y-x) < d \ norm (f y - f x - (y-x) *\<^sub>R f' x) \ e/2 * norm (y-x)" (is "\x \ box a b. ?Q x") \\The explicit quantifier is required by the following step\ proof fix x assume "x \ box a b" with e show "?Q x" using derf_exp [of x "e/2"] by auto qed then obtain d where d: "\x. 0 < d x" "\x y. \x \ box a b; norm (y-x) < d x\ \ norm (f y - f x - (y-x) *\<^sub>R f' x) \ e/2 * norm (y-x)" unfolding bgauge_existence_lemma by metis have "bounded (f ` cbox a b)" using compact_cbox assms by (auto simp: compact_imp_bounded compact_continuous_image) then obtain B where "0 < B" and B: "\x. x \ f ` cbox a b \ norm x \ B" unfolding bounded_pos by metis obtain da where "0 < da" and da: "\c. \a \ c; {a..c} \ {a..b}; {a..c} \ ball a da\ \ norm (content {a..c} *\<^sub>R f' a - (f c - f a)) \ (e * (b-a)) / 4" proof - have "continuous (at a within {a..b}) f" using contf continuous_on_eq_continuous_within by force with eba8 obtain k where "0 < k" and k: "\x. \x \ {a..b}; 0 < norm (x-a); norm (x-a) < k\ \ norm (f x - f a) < e * (b-a) / 8" unfolding continuous_within Lim_within dist_norm by metis obtain l where l: "0 < l" "norm (l *\<^sub>R f' a) \ e * (b-a) / 8" proof (cases "f' a = 0") case True with ab e that show ?thesis by auto next case False show ?thesis proof show "norm ((e * (b - a) / 8 / norm (f' a)) *\<^sub>R f' a) \ e * (b - a) / 8" "0 < e * (b - a) / 8 / norm (f' a)" using False ab e by (auto simp add: field_simps) qed qed have "norm (content {a..c} *\<^sub>R f' a - (f c - f a)) \ e * (b-a) / 4" if "a \ c" "{a..c} \ {a..b}" and bmin: "{a..c} \ ball a (min k l)" for c proof - have minkl: "\a - x\ < min k l" if "x \ {a..c}" for x using bmin dist_real_def that by auto then have lel: "\c - a\ \ \l\" using that by force have "norm ((c - a) *\<^sub>R f' a - (f c - f a)) \ norm ((c - a) *\<^sub>R f' a) + norm (f c - f a)" by (rule norm_triangle_ineq4) also have "\ \ e * (b-a) / 8 + e * (b-a) / 8" proof (rule add_mono) have "norm ((c - a) *\<^sub>R f' a) \ norm (l *\<^sub>R f' a)" by (auto intro: mult_right_mono [OF lel]) with l show "norm ((c - a) *\<^sub>R f' a) \ e * (b-a) / 8" by linarith next have "norm (f c - f a) < e * (b-a) / 8" proof (cases "a = c") case True then show ?thesis using eba8 by auto next case False show ?thesis by (rule k) (use minkl \a \ c\ that False in auto) qed then show "norm (f c - f a) \ e * (b-a) / 8" by simp qed finally show "norm (content {a..c} *\<^sub>R f' a - (f c - f a)) \ e * (b-a) / 4" unfolding content_real[OF \a \ c\] by auto qed then show ?thesis by (rule_tac da="min k l" in that) (auto simp: l \0 < k\) qed obtain db where "0 < db" and db: "\c. \c \ b; {c..b} \ {a..b}; {c..b} \ ball b db\ \ norm (content {c..b} *\<^sub>R f' b - (f b - f c)) \ (e * (b-a)) / 4" proof - have "continuous (at b within {a..b}) f" using contf continuous_on_eq_continuous_within by force with eba8 obtain k where "0 < k" and k: "\x. \x \ {a..b}; 0 < norm(x-b); norm(x-b) < k\ \ norm (f b - f x) < e * (b-a) / 8" unfolding continuous_within Lim_within dist_norm norm_minus_commute by metis obtain l where l: "0 < l" "norm (l *\<^sub>R f' b) \ (e * (b-a)) / 8" proof (cases "f' b = 0") case True thus ?thesis using ab e that by auto next case False show ?thesis proof show "norm ((e * (b - a) / 8 / norm (f' b)) *\<^sub>R f' b) \ e * (b - a) / 8" "0 < e * (b - a) / 8 / norm (f' b)" using False ab e by (auto simp add: field_simps) qed qed have "norm (content {c..b} *\<^sub>R f' b - (f b - f c)) \ e * (b-a) / 4" if "c \ b" "{c..b} \ {a..b}" and bmin: "{c..b} \ ball b (min k l)" for c proof - have minkl: "\b - x\ < min k l" if "x \ {c..b}" for x using bmin dist_real_def that by auto then have lel: "\b - c\ \ \l\" using that by force have "norm ((b - c) *\<^sub>R f' b - (f b - f c)) \ norm ((b - c) *\<^sub>R f' b) + norm (f b - f c)" by (rule norm_triangle_ineq4) also have "\ \ e * (b-a) / 8 + e * (b-a) / 8" proof (rule add_mono) have "norm ((b - c) *\<^sub>R f' b) \ norm (l *\<^sub>R f' b)" by (auto intro: mult_right_mono [OF lel]) also have "... \ e * (b-a) / 8" by (rule l) finally show "norm ((b - c) *\<^sub>R f' b) \ e * (b-a) / 8" . next have "norm (f b - f c) < e * (b-a) / 8" proof (cases "b = c") case True with eba8 show ?thesis by auto next case False show ?thesis by (rule k) (use minkl \c \ b\ that False in auto) qed then show "norm (f b - f c) \ e * (b-a) / 8" by simp qed finally show "norm (content {c..b} *\<^sub>R f' b - (f b - f c)) \ e * (b-a) / 4" unfolding content_real[OF \c \ b\] by auto qed then show ?thesis by (rule_tac db="min k l" in that) (auto simp: l \0 < k\) qed let ?d = "(\x. ball x (if x=a then da else if x=b then db else d x))" show "\d. gauge d \ (\p. p tagged_division_of {a..b} \ d fine p \ norm ((\(x,K)\p. content K *\<^sub>R f' x) - (f b - f a)) \ e * content {a..b})" proof (rule exI, safe) show "gauge ?d" using ab \db > 0\ \da > 0\ d(1) by (auto intro: gauge_ball_dependent) next fix p assume ptag: "p tagged_division_of {a..b}" and fine: "?d fine p" let ?A = "{t. fst t \ {a, b}}" note p = tagged_division_ofD[OF ptag] have pA: "p = (p \ ?A) \ (p - ?A)" "finite (p \ ?A)" "finite (p - ?A)" "(p \ ?A) \ (p - ?A) = {}" using ptag fine by auto have le_xz: "\w x y z::real. y \ z/2 \ w - x \ z/2 \ w + y \ x + z" by arith have non: False if xk: "(x,K) \ p" and "x \ a" "x \ b" and less: "e * (Sup K - Inf K)/2 < norm (content K *\<^sub>R f' x - (f (Sup K) - f (Inf K)))" for x K proof - obtain u v where k: "K = cbox u v" using p(4) xk by blast then have "u \ v" and uv: "{u, v} \ cbox u v" using p(2)[OF xk] by auto then have result: "e * (v - u)/2 < norm ((v - u) *\<^sub>R f' x - (f v - f u))" using less[unfolded k box_real interval_bounds_real content_real] by auto then have "x \ box a b" using p(2) p(3) \x \ a\ \x \ b\ xk by fastforce with d have *: "\y. norm (y-x) < d x \ norm (f y - f x - (y-x) *\<^sub>R f' x) \ e/2 * norm (y-x)" by metis have xd: "norm (u - x) < d x" "norm (v - x) < d x" using fineD[OF fine xk] \x \ a\ \x \ b\ uv by (auto simp add: k subset_eq dist_commute dist_real_def) have "norm ((v - u) *\<^sub>R f' x - (f v - f u)) = norm ((f u - f x - (u - x) *\<^sub>R f' x) - (f v - f x - (v - x) *\<^sub>R f' x))" by (rule arg_cong[where f=norm]) (auto simp: scaleR_left.diff) also have "\ \ e/2 * norm (u - x) + e/2 * norm (v - x)" by (metis norm_triangle_le_diff add_mono * xd) also have "\ \ e/2 * norm (v - u)" using p(2)[OF xk] by (auto simp add: field_simps k) also have "\ < norm ((v - u) *\<^sub>R f' x - (f v - f u))" using result by (simp add: \u \ v\) finally have "e * (v - u)/2 < e * (v - u)/2" using uv by auto then show False by auto qed have "norm (\(x, K)\p - ?A. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ (\(x, K)\p - ?A. norm (content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))))" by (auto intro: sum_norm_le) also have "... \ (\n\p - ?A. e * (case n of (x, k) \ Sup k - Inf k)/2)" using non by (fastforce intro: sum_mono) finally have I: "norm (\(x, k)\p - ?A. content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) \ (\n\p - ?A. e * (case n of (x, k) \ Sup k - Inf k))/2" by (simp add: sum_divide_distrib) have II: "norm (\(x, k)\p \ ?A. content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) - (\n\p \ ?A. e * (case n of (x, k) \ Sup k - Inf k)) \ (\n\p - ?A. e * (case n of (x, k) \ Sup k - Inf k))/2" proof - have ge0: "0 \ e * (Sup k - Inf k)" if xkp: "(x, k) \ p \ ?A" for x k proof - obtain u v where uv: "k = cbox u v" by (meson Int_iff xkp p(4)) with p that have "cbox u v \ {}" by blast then show "0 \ e * ((Sup k) - (Inf k))" unfolding uv using e by (auto simp add: field_simps) qed let ?B = "\x. {t \ p. fst t = x \ content (snd t) \ 0}" let ?C = "{t \ p. fst t \ {a, b} \ content (snd t) \ 0}" have "norm (\(x, k)\p \ {t. fst t \ {a, b}}. content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) \ e * (b-a)/2" proof - have *: "\S f e. sum f S = sum f (p \ ?C) \ norm (sum f (p \ ?C)) \ e \ norm (sum f S) \ e" by auto have 1: "content K *\<^sub>R (f' x) - (f ((Sup K)) - f ((Inf K))) = 0" if "(x,K) \ p \ {t. fst t \ {a, b}} - p \ ?C" for x K proof - have xk: "(x,K) \ p" and k0: "content K = 0" using that by auto then obtain u v where uv: "K = cbox u v" "u = v" using xk k0 p by fastforce then show "content K *\<^sub>R (f' x) - (f ((Sup K)) - f ((Inf K))) = 0" using xk unfolding uv by auto qed have 2: "norm(\(x,K)\p \ ?C. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ e * (b-a)/2" proof - have norm_le: "norm (sum f S) \ e" if "\x y. \x \ S; y \ S\ \ x = y" "\x. x \ S \ norm (f x) \ e" "e > 0" for S f and e :: real proof (cases "S = {}") case True with that show ?thesis by auto next case False then obtain x where "x \ S" by auto then have "S = {x}" using that(1) by auto then show ?thesis using \x \ S\ that(2) by auto qed have *: "p \ ?C = ?B a \ ?B b" by blast then have "norm (\(x,K)\p \ ?C. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) = norm (\(x,K) \ ?B a \ ?B b. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K)))" by simp also have "... = norm ((\(x,K) \ ?B a. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) + (\(x,K) \ ?B b. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))))" using p(1) ab e by (subst sum.union_disjoint) auto also have "... \ e * (b - a) / 4 + e * (b - a) / 4" proof (rule norm_triangle_le [OF add_mono]) have pa: "\v. k = cbox a v \ a \ v" if "(a, k) \ p" for k using p that by fastforce show "norm (\(x,K) \ ?B a. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ e * (b - a) / 4" proof (intro norm_le; clarsimp) fix K K' assume K: "(a, K) \ p" "(a, K') \ p" and ne0: "content K \ 0" "content K' \ 0" with pa obtain v v' where v: "K = cbox a v" "a \ v" and v': "K' = cbox a v'" "a \ v'" by blast let ?v = "min v v'" have "box a ?v \ K \ K'" unfolding v v' by (auto simp add: mem_box) then have "interior (box a (min v v')) \ interior K \ interior K'" using interior_Int interior_mono by blast moreover have "(a + ?v)/2 \ box a ?v" using ne0 unfolding v v' content_eq_0 not_le by (auto simp add: mem_box) ultimately have "(a + ?v)/2 \ interior K \ interior K'" unfolding interior_open[OF open_box] by auto then show "K = K'" using p(5)[OF K] by auto next fix K assume K: "(a, K) \ p" and ne0: "content K \ 0" show "norm (content c *\<^sub>R f' a - (f (Sup c) - f (Inf c))) * 4 \ e * (b-a)" if "(a, c) \ p" and ne0: "content c \ 0" for c proof - obtain v where v: "c = cbox a v" and "a \ v" using pa[OF \(a, c) \ p\] by metis then have "a \ {a..v}" "v \ b" using p(3)[OF \(a, c) \ p\] by auto moreover have "{a..v} \ ball a da" using fineD[OF \?d fine p\ \(a, c) \ p\] by (simp add: v split: if_split_asm) ultimately show ?thesis unfolding v interval_bounds_real[OF \a \ v\] box_real using da \a \ v\ by auto qed qed (use ab e in auto) next have pb: "\v. k = cbox v b \ b \ v" if "(b, k) \ p" for k using p that by fastforce show "norm (\(x,K) \ ?B b. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ e * (b - a) / 4" proof (intro norm_le; clarsimp) fix K K' assume K: "(b, K) \ p" "(b, K') \ p" and ne0: "content K \ 0" "content K' \ 0" with pb obtain v v' where v: "K = cbox v b" "v \ b" and v': "K' = cbox v' b" "v' \ b" by blast let ?v = "max v v'" have "box ?v b \ K \ K'" unfolding v v' by (auto simp: mem_box) then have "interior (box (max v v') b) \ interior K \ interior K'" using interior_Int interior_mono by blast moreover have "((b + ?v)/2) \ box ?v b" using ne0 unfolding v v' content_eq_0 not_le by (auto simp: mem_box) ultimately have "((b + ?v)/2) \ interior K \ interior K'" unfolding interior_open[OF open_box] by auto then show "K = K'" using p(5)[OF K] by auto next fix K assume K: "(b, K) \ p" and ne0: "content K \ 0" show "norm (content c *\<^sub>R f' b - (f (Sup c) - f (Inf c))) * 4 \ e * (b-a)" if "(b, c) \ p" and ne0: "content c \ 0" for c proof - obtain v where v: "c = cbox v b" and "v \ b" using \(b, c) \ p\ pb by blast then have "v \ a""b \ {v.. b}" using p(3)[OF \(b, c) \ p\] by auto moreover have "{v..b} \ ball b db" using fineD[OF \?d fine p\ \(b, c) \ p\] box_real(2) v False by force ultimately show ?thesis using db v by auto qed qed (use ab e in auto) qed also have "... = e * (b-a)/2" by simp finally show "norm (\(x,k)\p \ ?C. content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) \ e * (b-a)/2" . qed show "norm (\(x, k)\p \ ?A. content k *\<^sub>R f' x - (f ((Sup k)) - f ((Inf k)))) \ e * (b-a)/2" apply (rule * [OF sum.mono_neutral_right[OF pA(2)]]) using 1 2 by (auto simp: split_paired_all) qed also have "... = (\n\p. e * (case n of (x, k) \ Sup k - Inf k))/2" unfolding sum_distrib_left[symmetric] by (subst additive_tagged_division_1[OF \a \ b\ ptag]) auto finally have norm_le: "norm (\(x,K)\p \ {t. fst t \ {a, b}}. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ (\n\p. e * (case n of (x, K) \ Sup K - Inf K))/2" . have le2: "\x s1 s2::real. 0 \ s1 \ x \ (s1 + s2)/2 \ x - s1 \ s2/2" by auto show ?thesis apply (rule le2 [OF sum_nonneg]) using ge0 apply force by (metis (no_types, lifting) Diff_Diff_Int Diff_subset norm_le p(1) sum.subset_diff) qed note * = additive_tagged_division_1[OF assms(1) ptag, symmetric] have "norm (\(x,K)\p \ ?A \ (p - ?A). content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ e * (\(x,K)\p \ ?A \ (p - ?A). Sup K - Inf K)" unfolding sum_distrib_left unfolding sum.union_disjoint[OF pA(2-)] using le_xz norm_triangle_le I II by blast then show "norm ((\(x,K)\p. content K *\<^sub>R f' x) - (f b - f a)) \ e * content {a..b}" by (simp only: content_real[OF \a \ b\] *[of "\x. x"] *[of f] sum_subtractf[symmetric] split_minus pA(1) [symmetric]) qed qed qed subsection \Stronger form with finite number of exceptional points\ lemma fundamental_theorem_of_calculus_interior_strong: fixes f :: "real \ 'a::banach" assumes "finite S" and "a \ b" "\x. x \ {a <..< b} - S \ (f has_vector_derivative f'(x)) (at x)" and "continuous_on {a .. b} f" shows "(f' has_integral (f b - f a)) {a .. b}" using assms proof (induction arbitrary: a b) case empty then show ?case using fundamental_theorem_of_calculus_interior by force next case (insert x S) show ?case proof (cases "x \ {a<..continuous_on {a..b} f\ \a < x\ \x < b\ continuous_on_subset by (force simp: intro!: insert)+ then have "(f' has_integral f x - f a + (f b - f x)) {a..b}" using \a < x\ \x < b\ has_integral_combine less_imp_le by blast then show ?thesis by simp qed qed corollary fundamental_theorem_of_calculus_strong: fixes f :: "real \ 'a::banach" assumes "finite S" and "a \ b" and vec: "\x. x \ {a..b} - S \ (f has_vector_derivative f'(x)) (at x)" and "continuous_on {a..b} f" shows "(f' has_integral (f b - f a)) {a..b}" by (rule fundamental_theorem_of_calculus_interior_strong [OF \finite S\]) (force simp: assms)+ proposition indefinite_integral_continuous_left: fixes f:: "real \ 'a::banach" assumes intf: "f integrable_on {a..b}" and "a < c" "c \ b" "e > 0" obtains d where "d > 0" and "\t. c - d < t \ t \ c \ norm (integral {a..c} f - integral {a..t} f) < e" proof - obtain w where "w > 0" and w: "\t. \c - w < t; t < c\ \ norm (f c) * norm(c - t) < e/3" proof (cases "f c = 0") case False hence e3: "0 < e/3 / norm (f c)" using \e>0\ by simp moreover have "norm (f c) * norm (c - t) < e/3" if "t < c" and "c - e/3 / norm (f c) < t" for t unfolding real_norm_def by (smt (verit) False divide_right_mono nonzero_mult_div_cancel_left norm_eq_zero norm_ge_zero that) ultimately show ?thesis using that by auto qed (use \e > 0\ in auto) let ?SUM = "\p. (\(x,K) \ p. content K *\<^sub>R f x)" have e3: "e/3 > 0" using \e > 0\ by auto have "f integrable_on {a..c}" using \a < c\ \c \ b\ by (auto intro: integrable_subinterval_real[OF intf]) then obtain d1 where "gauge d1" and d1: "\p. \p tagged_division_of {a..c}; d1 fine p\ \ norm (?SUM p - integral {a..c} f) < e/3" using integrable_integral has_integral_real e3 by metis define d where [abs_def]: "d x = ball x w \ d1 x" for x have "gauge d" unfolding d_def using \w > 0\ \gauge d1\ by auto then obtain k where "0 < k" and k: "ball c k \ d c" by (meson gauge_def open_contains_ball) let ?d = "min k (c - a)/2" show thesis proof (intro that[of ?d] allI impI, safe) show "?d > 0" using \0 < k\ \a < c\ by auto next fix t assume t: "c - ?d < t" "t \ c" show "norm (integral ({a..c}) f - integral ({a..t}) f) < e" proof (cases "t < c") case False with \t \ c\ show ?thesis by (simp add: \e > 0\) next case True have "f integrable_on {a..t}" using \t < c\ \c \ b\ by (auto intro: integrable_subinterval_real[OF intf]) then obtain d2 where d2: "gauge d2" "\p. p tagged_division_of {a..t} \ d2 fine p \ norm (?SUM p - integral {a..t} f) < e/3" using integrable_integral has_integral_real e3 by metis define d3 where "d3 x = (if x \ t then d1 x \ d2 x else d1 x)" for x have "gauge d3" using \gauge d1\ \gauge d2\ unfolding d3_def gauge_def by auto then obtain p where ptag: "p tagged_division_of {a..t}" and pfine: "d3 fine p" by (metis box_real(2) fine_division_exists) note p' = tagged_division_ofD[OF ptag] have pt: "(x,K)\p \ x \ t" for x K by (meson atLeastAtMost_iff p'(2) p'(3) subsetCE) with pfine have "d2 fine p" unfolding fine_def d3_def by fastforce then have d2_fin: "norm (?SUM p - integral {a..t} f) < e/3" using d2(2) ptag by auto have eqs: "{a..c} \ {x. x \ t} = {a..t}" "{a..c} \ {x. x \ t} = {t..c}" using t by (auto simp add: field_simps) have "p \ {(c, {t..c})} tagged_division_of {a..c}" proof (intro tagged_division_Un_interval_real) show "{(c, {t..c})} tagged_division_of {a..c} \ {x. t \ x \ 1}" using \t \ c\ by (auto simp: eqs tagged_division_of_self_real) qed (auto simp: eqs ptag) moreover have "d1 fine p \ {(c, {t..c})}" unfolding fine_def proof safe fix x K y assume "(x,K) \ p" and "y \ K" then show "y \ d1 x" by (metis Int_iff d3_def subsetD fineD pfine) next fix x assume "x \ {t..c}" then have "dist c x < k" using t(1) by (auto simp add: field_simps dist_real_def) with k show "x \ d1 c" unfolding d_def by auto qed ultimately have d1_fin: "norm (?SUM(p \ {(c, {t..c})}) - integral {a..c} f) < e/3" using d1 by metis have SUMEQ: "?SUM(p \ {(c, {t..c})}) = (c - t) *\<^sub>R f c + ?SUM p" proof - have "?SUM(p \ {(c, {t..c})}) = (content{t..c} *\<^sub>R f c) + ?SUM p" proof (subst sum.union_disjoint) show "p \ {(c, {t..c})} = {}" using \t < c\ pt by force qed (use p'(1) in auto) also have "... = (c - t) *\<^sub>R f c + ?SUM p" using \t \ c\ by auto finally show ?thesis . qed have "c - k < t" using \k>0\ t(1) by (auto simp add: field_simps) moreover have "k \ w" proof (rule ccontr) assume "\ k \ w" then have "c + (k + w) / 2 \ d c" by (auto simp add: field_simps not_le not_less dist_real_def d_def) then have "c + (k + w) / 2 \ ball c k" using k by blast then show False using \0 < w\ \\ k \ w\ dist_real_def by auto qed ultimately have cwt: "c - w < t" by (auto simp add: field_simps) have eq: "integral {a..c} f - integral {a..t} f = -(((c - t) *\<^sub>R f c + ?SUM p) - integral {a..c} f) + (?SUM p - integral {a..t} f) + (c - t) *\<^sub>R f c" by auto have "norm (integral {a..c} f - integral {a..t} f) < e/3 + e/3 + e/3" unfolding eq proof (intro norm_triangle_lt add_strict_mono) show "norm (- ((c - t) *\<^sub>R f c + ?SUM p - integral {a..c} f)) < e/3" by (metis SUMEQ d1_fin norm_minus_cancel) show "norm (?SUM p - integral {a..t} f) < e/3" using d2_fin by blast show "norm ((c - t) *\<^sub>R f c) < e/3" using w cwt \t < c\ by simp (simp add: field_simps) qed then show ?thesis by simp qed qed qed lemma indefinite_integral_continuous_right: fixes f :: "real \ 'a::banach" assumes "f integrable_on {a..b}" and "a \ c" and "c < b" and "e > 0" obtains d where "0 < d" and "\t. c \ t \ t < c + d \ norm (integral {a..c} f - integral {a..t} f) < e" proof - have intm: "(\x. f (- x)) integrable_on {-b .. -a}" "- b < - c" "- c \ - a" using assms by auto from indefinite_integral_continuous_left[OF intm \e>0\] obtain d where "0 < d" and d: "\t. \- c - d < t; t \ -c\ \ norm (integral {- b..- c} (\x. f (-x)) - integral {- b..t} (\x. f (-x))) < e" by metis let ?d = "min d (b - c)" show ?thesis proof (intro that[of "?d"] allI impI) show "0 < ?d" using \0 < d\ \c < b\ by auto fix t :: real assume t: "c \ t \ t < c + ?d" have *: "integral {a..c} f = integral {a..b} f - integral {c..b} f" "integral {a..t} f = integral {a..b} f - integral {t..b} f" using assms t by (auto simp: algebra_simps integral_combine) have "(- c) - d < (- t)" "- t \ - c" using t by auto from d[OF this] show "norm (integral {a..c} f - integral {a..t} f) < e" by (auto simp add: algebra_simps norm_minus_commute *) qed qed lemma indefinite_integral_continuous_1: fixes f :: "real \ 'a::banach" assumes "f integrable_on {a..b}" shows "continuous_on {a..b} (\x. integral {a..x} f)" proof - have "\d>0. \x'\{a..b}. dist x' x < d \ dist (integral {a..x'} f) (integral {a..x} f) < e" if x: "x \ {a..b}" and "e > 0" for x e :: real proof (cases "a = b") case True with that show ?thesis by force next case False with x have "a < b" by force with x consider "x = a" | "x = b" | "a < x" "x < b" by force then show ?thesis proof cases case 1 then show ?thesis by (force simp: dist_norm algebra_simps intro: indefinite_integral_continuous_right [OF assms _ \a < b\ \e > 0\]) next case 2 then show ?thesis by (force simp: dist_norm norm_minus_commute algebra_simps intro: indefinite_integral_continuous_left [OF assms \a < b\ _ \e > 0\]) next case 3 obtain d1 where "0 < d1" and d1: "\t. \x - d1 < t; t \ x\ \ norm (integral {a..x} f - integral {a..t} f) < e" using 3 by (auto intro: indefinite_integral_continuous_left [OF assms \a < x\ _ \e > 0\]) obtain d2 where "0 < d2" and d2: "\t. \x \ t; t < x + d2\ \ norm (integral {a..x} f - integral {a..t} f) < e" using 3 by (auto intro: indefinite_integral_continuous_right [OF assms _ \x < b\ \e > 0\]) show ?thesis proof (intro exI ballI conjI impI) show "0 < min d1 d2" using \0 < d1\ \0 < d2\ by simp show "dist (integral {a..y} f) (integral {a..x} f) < e" if "dist y x < min d1 d2" for y by (smt (verit) d1 d2 dist_norm dist_real_def norm_minus_commute that) qed qed qed then show ?thesis by (auto simp: continuous_on_iff) qed lemma indefinite_integral_continuous_1': fixes f::"real \ 'a::banach" assumes "f integrable_on {a..b}" shows "continuous_on {a..b} (\x. integral {x..b} f)" proof - have "integral {a..b} f - integral {a..x} f = integral {x..b} f" if "x \ {a..b}" for x using integral_combine[OF _ _ assms, of x] that by (auto simp: algebra_simps) with _ show ?thesis by (rule continuous_on_eq) (auto intro!: continuous_intros indefinite_integral_continuous_1 assms) qed theorem integral_has_vector_derivative': fixes f :: "real \ 'b::banach" assumes "continuous_on {a..b} f" and "x \ {a..b}" shows "((\u. integral {u..b} f) has_vector_derivative - f x) (at x within {a..b})" proof - have *: "integral {x..b} f = integral {a .. b} f - integral {a .. x} f" if "a \ x" "x \ b" for x using integral_combine[of a x b for x, OF that integrable_continuous_real[OF assms(1)]] by (simp add: algebra_simps) show ?thesis using \x \ _\ * by (rule has_vector_derivative_transform) (auto intro!: derivative_eq_intros assms integral_has_vector_derivative) qed lemma integral_has_real_derivative': assumes "continuous_on {a..b} g" assumes "t \ {a..b}" shows "((\x. integral {x..b} g) has_real_derivative -g t) (at t within {a..b})" using integral_has_vector_derivative'[OF assms] by (auto simp: has_real_derivative_iff_has_vector_derivative) subsection \This doesn't directly involve integration, but that gives an easy proof\ lemma has_derivative_zero_unique_strong_interval: fixes f :: "real \ 'a::banach" assumes "finite k" and contf: "continuous_on {a..b} f" and "f a = y" and fder: "\x. x \ {a..b} - k \ (f has_derivative (\h. 0)) (at x within {a..b})" and x: "x \ {a..b}" shows "f x = y" proof - have "a \ b" "a \ x" using assms by auto have "((\x. 0::'a) has_integral f x - f a) {a..x}" proof (rule fundamental_theorem_of_calculus_interior_strong[OF \finite k\ \a \ x\]; clarify?) have "{a..x} \ {a..b}" using x by auto then show "continuous_on {a..x} f" by (rule continuous_on_subset[OF contf]) show "(f has_vector_derivative 0) (at z)" if z: "z \ {a<.. k" for z unfolding has_vector_derivative_def proof (simp add: at_within_open[OF z, symmetric]) show "(f has_derivative (\x. 0)) (at z within {a<..Generalize a bit to any convex set\ lemma has_derivative_zero_unique_strong_convex: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "convex S" "finite K" and contf: "continuous_on S f" and "c \ S" "f c = y" and derf: "\x. x \ S - K \ (f has_derivative (\h. 0)) (at x within S)" and "x \ S" shows "f x = y" proof (cases "x = c") case True with \f c = y\ show ?thesis by blast next case False let ?\ = "\u. (1 - u) *\<^sub>R c + u *\<^sub>R x" have contf': "continuous_on {0 ..1} (f \ ?\)" proof (rule continuous_intros continuous_on_subset[OF contf])+ show "(\u. (1 - u) *\<^sub>R c + u *\<^sub>R x) ` {0..1} \ S" using \convex S\ \x \ S\ \c \ S\ by (auto simp add: convex_alt algebra_simps) qed have "t = u" if "?\ t = ?\ u" for t u proof - from that have "(t - u) *\<^sub>R x = (t - u) *\<^sub>R c" by (auto simp add: algebra_simps) then show ?thesis using \x \ c\ by auto qed then have eq: "(SOME t. ?\ t = ?\ u) = u" for u by blast then have "(?\ -` K) \ (\z. SOME t. ?\ t = z) ` K" by (clarsimp simp: image_iff) (metis (no_types) eq) then have fin: "finite (?\ -` K)" by (rule finite_surj[OF \finite K\]) have derf': "((\u. f (?\ u)) has_derivative (\h. 0)) (at t within {0..1})" if "t \ {0..1} - {t. ?\ t \ K}" for t proof - have df: "(f has_derivative (\h. 0)) (at (?\ t) within ?\ ` {0..1})" using \convex S\ \x \ S\ \c \ S\ that by (auto simp add: convex_alt algebra_simps intro: has_derivative_subset [OF derf]) have "(f \ ?\ has_derivative (\x. 0) \ (\z. (0 - z *\<^sub>R c) + z *\<^sub>R x)) (at t within {0..1})" by (rule derivative_eq_intros df | simp)+ then show ?thesis unfolding o_def . qed have "(f \ ?\) 1 = y" apply (rule has_derivative_zero_unique_strong_interval[OF fin contf']) unfolding o_def using \f c = y\ derf' by auto then show ?thesis by auto qed text \Also to any open connected set with finite set of exceptions. Could generalize to locally convex set with limpt-free set of exceptions.\ lemma has_derivative_zero_unique_strong_connected: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "connected S" and "open S" and "finite K" and contf: "continuous_on S f" and "c \ S" and "f c = y" and derf: "\x. x \ S - K \ (f has_derivative (\h. 0)) (at x within S)" and "x \ S" shows "f x = y" proof - have "\e>0. ball x e \ (S \ f -` {f x})" if "x \ S" for x proof - obtain e where "0 < e" and e: "ball x e \ S" using \x \ S\ \open S\ open_contains_ball by blast have "ball x e \ {u \ S. f u \ {f x}}" proof safe fix y assume y: "y \ ball x e" then show "y \ S" using e by auto show "f y = f x" proof (rule has_derivative_zero_unique_strong_convex[OF convex_ball \finite K\]) show "continuous_on (ball x e) f" using contf continuous_on_subset e by blast show "(f has_derivative (\h. 0)) (at u within ball x e)" if "u \ ball x e - K" for u by (metis Diff_iff contra_subsetD derf e has_derivative_subset that) qed (use y e \0 < e\ in auto) qed then show "\e>0. ball x e \ (S \ f -` {f x})" using \0 < e\ by blast qed then have "openin (top_of_set S) (S \ f -` {y})" by (auto intro!: open_openin_trans[OF \open S\] simp: open_contains_ball) moreover have "closedin (top_of_set S) (S \ f -` {y})" by (force intro!: continuous_closedin_preimage [OF contf]) ultimately have "(S \ f -` {y}) = {} \ (S \ f -` {y}) = S" using \connected S\ by (simp add: connected_clopen) then show ?thesis using \x \ S\ \f c = y\ \c \ S\ by auto qed lemma has_derivative_zero_connected_constant_on: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "connected S" "open S" "finite K" "continuous_on S f" and "\x\S-K. (f has_derivative (\h. 0)) (at x within S)" shows "f constant_on S" by (smt (verit, best) assms constant_on_def has_derivative_zero_unique_strong_connected) lemma DERIV_zero_connected_constant_on: fixes f :: "'a::{real_normed_field,euclidean_space} \ 'a" assumes *: "connected S" "open S" "finite K" "continuous_on S f" and 0: "\x\S-K. DERIV f x :> 0" shows "f constant_on S" using has_derivative_zero_connected_constant_on [OF *] 0 by (metis has_derivative_at_withinI has_field_derivative_def lambda_zero) lemma DERIV_zero_connected_constant: fixes f :: "'a::{real_normed_field,euclidean_space} \ 'a" assumes "connected S" and "open S" and "finite K" and "continuous_on S f" and "\x\S-K. DERIV f x :> 0" obtains c where "\x. x \ S \ f(x) = c" by (metis DERIV_zero_connected_constant_on [OF assms] constant_on_def) lemma has_field_derivative_0_imp_constant_on: fixes f :: "'a::{real_normed_field,euclidean_space} \ 'a" assumes "\z. z \ S \ (f has_field_derivative 0) (at z)" and S: "connected S" "open S" shows "f constant_on S" using DERIV_zero_connected_constant_on [where K="Basis"] by (metis DERIV_isCont Diff_iff assms continuous_at_imp_continuous_on eucl.finite_Basis) subsection \Integrating characteristic function of an interval\ lemma has_integral_restrict_open_subinterval: fixes f :: "'a::euclidean_space \ 'b::banach" assumes intf: "(f has_integral i) (cbox c d)" and cb: "cbox c d \ cbox a b" shows "((\x. if x \ box c d then f x else 0) has_integral i) (cbox a b)" proof (cases "cbox c d = {}") case True then have "box c d = {}" by (metis bot.extremum_uniqueI box_subset_cbox) then show ?thesis using True intf by auto next case False then obtain p where pdiv: "p division_of cbox a b" and inp: "cbox c d \ p" using cb partial_division_extend_1 by blast define g where [abs_def]: "g x = (if x \box c d then f x else 0)" for x interpret operative "lift_option plus" "Some (0 :: 'b)" "\i. if g integrable_on i then Some (integral i g) else None" by (fact operative_integralI) note operat = division [OF pdiv, symmetric] show ?thesis proof (cases "(g has_integral i) (cbox a b)") case True then show ?thesis by (simp add: g_def) next case False have iterate:"F (\i. if g integrable_on i then Some (integral i g) else None) (p - {cbox c d}) = Some 0" proof (intro neutral ballI) fix x assume x: "x \ p - {cbox c d}" then have "x \ p" by auto then obtain u v where uv: "x = cbox u v" using pdiv by blast have "interior x \ interior (cbox c d) = {}" using pdiv inp x by blast then have "(g has_integral 0) x" unfolding uv using has_integral_spike_interior[where f="\x. 0"] by (metis (no_types, lifting) disjoint_iff_not_equal g_def has_integral_0_eq interior_cbox) then show "(if g integrable_on x then Some (integral x g) else None) = Some 0" by auto qed interpret comm_monoid_set "lift_option plus" "Some (0 :: 'b)" by (intro comm_monoid_set.intro comm_monoid_lift_option add.comm_monoid_axioms) have intg: "g integrable_on cbox c d" using integrable_spike_interior[where f=f] by (meson g_def has_integral_integrable intf) moreover have "integral (cbox c d) g = i" by (meson g_def has_integral_iff has_integral_spike_interior intf) ultimately have "F (\A. if g integrable_on A then Some (integral A g) else None) p = Some i" by (metis (full_types, lifting) division_of_finite inp iterate pdiv remove right_neutral) with False show ?thesis by (metis integrable_integral not_None_eq operat option.inject) qed qed lemma has_integral_restrict_closed_subinterval: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "(f has_integral i) (cbox c d)" and "cbox c d \ cbox a b" shows "((\x. if x \ cbox c d then f x else 0) has_integral i) (cbox a b)" proof - note has_integral_restrict_open_subinterval[OF assms] note * = has_integral_spike[OF negligible_frontier_interval _ this] show ?thesis by (rule *[of c d]) (use box_subset_cbox[of c d] in auto) qed lemma has_integral_restrict_closed_subintervals_eq: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "cbox c d \ cbox a b" shows "((\x. if x \ cbox c d then f x else 0) has_integral i) (cbox a b) \ (f has_integral i) (cbox c d)" (is "?l = ?r") proof (cases "cbox c d = {}") case False let ?g = "\x. if x \ cbox c d then f x else 0" show ?thesis proof assume ?l then have "?g integrable_on cbox c d" using assms has_integral_integrable integrable_subinterval by blast then have "f integrable_on cbox c d" by (rule integrable_eq) auto moreover then have "i = integral (cbox c d) f" by (meson \((\x. if x \ cbox c d then f x else 0) has_integral i) (cbox a b)\ assms has_integral_restrict_closed_subinterval has_integral_unique integrable_integral) ultimately show ?r by auto next assume ?r then show ?l by (rule has_integral_restrict_closed_subinterval[OF _ assms]) qed qed auto text \Hence we can apply the limit process uniformly to all integrals.\ lemma has_integral': fixes f :: "'n::euclidean_space \ 'a::banach" shows "(f has_integral i) S \ (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ (\z. ((\x. if x \ S then f(x) else 0) has_integral z) (cbox a b) \ norm(z - i) < e))" (is "?l \ (\e>0. ?r e)") proof (cases "\a b. S = cbox a b") case False then show ?thesis by (simp add: has_integral_alt) next case True then obtain a b where S: "S = cbox a b" by blast obtain B where " 0 < B" and B: "\x. x \ cbox a b \ norm x \ B" using bounded_cbox[unfolded bounded_pos] by blast show ?thesis proof safe fix e :: real assume ?l and "e > 0" have "((\x. if x \ S then f x else 0) has_integral i) (cbox c d)" if "ball 0 (B+1) \ cbox c d" for c d unfolding S using B that by (force intro: \?l\[unfolded S] has_integral_restrict_closed_subinterval) then show "?r e" by (meson \0 < B\ \0 < e\ add_pos_pos le_less_trans zero_less_one norm_pths(2)) next assume as: "\e>0. ?r e" then obtain C where C: "\a b. ball 0 C \ cbox a b \ \z. ((\x. if x \ S then f x else 0) has_integral z) (cbox a b)" by (meson zero_less_one) define c :: 'n where "c = (\i\Basis. (- max B C) *\<^sub>R i)" define d :: 'n where "d = (\i\Basis. max B C *\<^sub>R i)" have "c \ i \ x \ i \ x \ i \ d \ i" if "norm x \ B" "i \ Basis" for x i using that and Basis_le_norm[OF \i\Basis\, of x] by (auto simp add: field_simps sum_negf c_def d_def) then have c_d: "cbox a b \ cbox c d" by (meson B mem_box(2) subsetI) have "c \ i \ x \ i \ x \ i \ d \ i" if x: "norm (0 - x) < C" and i: "i \ Basis" for x i using Basis_le_norm[OF i, of x] x i by (auto simp: sum_negf c_def d_def) then have "ball 0 C \ cbox c d" by (auto simp: mem_box dist_norm) with C obtain y where y: "(f has_integral y) (cbox a b)" using c_d has_integral_restrict_closed_subintervals_eq S by blast have "y = i" proof (rule ccontr) assume "y \ i" then have "0 < norm (y - i)" by auto from as[rule_format,OF this] obtain C where C: "\a b. ball 0 C \ cbox a b \ \z. ((\x. if x \ S then f x else 0) has_integral z) (cbox a b) \ norm (z-i) < norm (y-i)" by auto define c :: 'n where "c = (\i\Basis. (- max B C) *\<^sub>R i)" define d :: 'n where "d = (\i\Basis. max B C *\<^sub>R i)" have "c \ i \ x \ i \ x \ i \ d \ i" if "norm x \ B" and "i \ Basis" for x i using that Basis_le_norm[of i x] by (auto simp add: field_simps sum_negf c_def d_def) then have c_d: "cbox a b \ cbox c d" by (simp add: B mem_box(2) subset_eq) have "c \ i \ x \ i \ x \ i \ d \ i" if "norm (0 - x) < C" and "i \ Basis" for x i using Basis_le_norm[of i x] that by (auto simp: sum_negf c_def d_def) then have "ball 0 C \ cbox c d" by (auto simp: mem_box dist_norm) with C obtain z where z: "(f has_integral z) (cbox a b)" "norm (z-i) < norm (y-i)" using has_integral_restrict_closed_subintervals_eq[OF c_d] S by blast moreover then have "z = y" by (blast intro: has_integral_unique[OF _ y]) ultimately show False by auto qed then show ?l using y by (auto simp: S) qed qed lemma has_integral_le: fixes f :: "'n::euclidean_space \ real" assumes fg: "(f has_integral i) S" "(g has_integral j) S" and le: "\x. x \ S \ f x \ g x" shows "i \ j" using has_integral_component_le[OF _ fg, of 1] le by auto lemma integral_le: fixes f :: "'n::euclidean_space \ real" assumes "f integrable_on S" and "g integrable_on S" and "\x. x \ S \ f x \ g x" shows "integral S f \ integral S g" by (meson assms has_integral_le integrable_integral) lemma has_integral_nonneg: fixes f :: "'n::euclidean_space \ real" assumes "(f has_integral i) S" and "\x. x \ S \ 0 \ f x" shows "0 \ i" using assms has_integral_0 has_integral_le by blast lemma integral_nonneg: fixes f :: "'n::euclidean_space \ real" assumes f: "f integrable_on S" and 0: "\x. x \ S \ 0 \ f x" shows "0 \ integral S f" by (rule has_integral_nonneg[OF f[unfolded has_integral_integral] 0]) text \Hence a general restriction property.\ lemma has_integral_restrict [simp]: fixes f :: "'a :: euclidean_space \ 'b :: banach" assumes "S \ T" shows "((\x. if x \ S then f x else 0) has_integral i) T \ (f has_integral i) S" proof - have *: "\x. (if x \ T then if x \ S then f x else 0 else 0) = (if x\S then f x else 0)" using assms by auto show ?thesis apply (subst(2) has_integral') apply (subst has_integral') apply (simp add: *) done qed corollary has_integral_restrict_UNIV: fixes f :: "'n::euclidean_space \ 'a::banach" shows "((\x. if x \ s then f x else 0) has_integral i) UNIV \ (f has_integral i) s" by auto lemma has_integral_restrict_Int: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "((\x. if x \ S then f x else 0) has_integral i) T \ (f has_integral i) (S \ T)" proof - have "((\x. if x \ T then if x \ S then f x else 0 else 0) has_integral i) UNIV = ((\x. if x \ S \ T then f x else 0) has_integral i) UNIV" by (rule has_integral_cong) auto then show ?thesis using has_integral_restrict_UNIV by fastforce qed lemma integral_restrict_Int: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "integral T (\x. if x \ S then f x else 0) = integral (S \ T) f" by (metis (no_types, lifting) has_integral_cong has_integral_restrict_Int integrable_integral integral_unique not_integrable_integral) lemma integrable_restrict_Int: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "(\x. if x \ S then f x else 0) integrable_on T \ f integrable_on (S \ T)" using has_integral_restrict_Int by fastforce lemma has_integral_on_superset: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "(f has_integral i) S" and "\x. x \ S \ f x = 0" and "S \ T" shows "(f has_integral i) T" by (smt (verit, ccfv_SIG) assms has_integral_cong has_integral_restrict) lemma integrable_on_superset: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on S" and "\x. x \ S \ f x = 0" and "S \ t" shows "f integrable_on t" by (meson assms has_integral_on_superset integrable_integral integrable_on_def) lemma integral_subset_negligible: fixes f :: "'a :: euclidean_space \ 'b :: banach" assumes "S \ T" "negligible (T - S)" shows "integral S f = integral T f" proof - have "integral T f = integral T (\x. if x \ S then f x else 0)" by (rule integral_spike[of "T - S"]) (use assms in auto) also have "\ = integral (S \ T) f" by (subst integral_restrict_Int) auto also have "S \ T = S" using assms by auto finally show ?thesis .. qed lemma integral_restrict_UNIV: fixes f :: "'n::euclidean_space \ 'a::banach" shows "integral UNIV (\x. if x \ S then f x else 0) = integral S f" by (simp add: integral_restrict_Int) lemma integrable_restrict_UNIV: fixes f :: "'n::euclidean_space \ 'a::banach" shows "(\x. if x \ s then f x else 0) integrable_on UNIV \ f integrable_on s" unfolding integrable_on_def by auto lemma has_integral_subset_component_le: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes k: "k \ Basis" and "S \ T" "(f has_integral i) S" "(f has_integral j) T" "\x. x\T \ 0 \ f(x)\k" shows "i\k \ j\k" proof - have \
: "((\x. if x \ S then f x else 0) has_integral i) UNIV" "((\x. if x \ T then f x else 0) has_integral j) UNIV" by (simp_all add: assms) show ?thesis using assms by (force intro!: has_integral_component_le[OF k \
]) qed subsection\Integrals on set differences\ lemma has_integral_setdiff: fixes f :: "'a::euclidean_space \ 'b::banach" assumes S: "(f has_integral i) S" and T: "(f has_integral j) T" and neg: "negligible (T - S)" shows "(f has_integral (i - j)) (S - T)" proof - show ?thesis unfolding has_integral_restrict_UNIV [symmetric, of f] proof (rule has_integral_spike [OF neg]) have eq: "(\x. (if x \ S then f x else 0) - (if x \ T then f x else 0)) = (\x. if x \ T - S then - f x else if x \ S - T then f x else 0)" by (force simp add: ) have "((\x. if x \ S then f x else 0) has_integral i) UNIV" "((\x. if x \ T then f x else 0) has_integral j) UNIV" using S T has_integral_restrict_UNIV by auto from has_integral_diff [OF this] show "((\x. if x \ T - S then - f x else if x \ S - T then f x else 0) has_integral i-j) UNIV" by (simp add: eq) qed force qed lemma integral_setdiff: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "f integrable_on S" "f integrable_on T" "negligible(T - S)" shows "integral (S - T) f = integral S f - integral T f" by (rule integral_unique) (simp add: assms has_integral_setdiff integrable_integral) lemma integrable_setdiff: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "(f has_integral i) S" "(f has_integral j) T" "negligible (T - S)" shows "f integrable_on (S - T)" using has_integral_setdiff [OF assms] by (simp add: has_integral_iff) lemma negligible_setdiff [simp]: "T \ S \ negligible (T - S)" by (metis Diff_eq_empty_iff negligible_empty) lemma negligible_on_intervals: "negligible s \ (\a b. negligible(s \ cbox a b))" (is "?l \ ?r") proof assume R: ?r show ?l unfolding negligible_def by (metis Diff_iff Int_iff R has_integral_negligible indicator_simps(2)) qed (simp add: negligible_Int) lemma negligible_translation: assumes "negligible S" shows "negligible ((+) c ` S)" proof - have inj: "inj ((+) c)" by simp show ?thesis using assms proof (clarsimp simp: negligible_def) fix a b assume "\x y. (indicator S has_integral 0) (cbox x y)" then have *: "(indicator S has_integral 0) (cbox (a-c) (b-c))" by (meson Diff_iff assms has_integral_negligible indicator_simps(2)) have eq: "indicator ((+) c ` S) = (\x. indicator S (x - c))" by (force simp add: indicator_def) show "(indicator ((+) c ` S) has_integral 0) (cbox a b)" using has_integral_affinity [OF *, of 1 "-c"] cbox_translation [of "c" "-c+a" "-c+b"] by (simp add: eq) (simp add: ac_simps) qed qed lemma negligible_translation_rev: assumes "negligible ((+) c ` S)" shows "negligible S" by (metis negligible_translation [OF assms, of "-c"] translation_galois) lemma negligible_atLeastAtMostI: "b \ a \ negligible {a..(b::real)}" using negligible_insert by fastforce lemma has_integral_spike_set_eq: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "negligible {x \ S - T. f x \ 0}" "negligible {x \ T - S. f x \ 0}" shows "(f has_integral y) S \ (f has_integral y) T" proof - have "((\x. if x \ S then f x else 0) has_integral y) UNIV = ((\x. if x \ T then f x else 0) has_integral y) UNIV" proof (rule has_integral_spike_eq) show "negligible ({x \ S - T. f x \ 0} \ {x \ T - S. f x \ 0})" by (rule negligible_Un [OF assms]) qed auto then show ?thesis by (simp add: has_integral_restrict_UNIV) qed corollary integral_spike_set: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "negligible {x \ S - T. f x \ 0}" "negligible {x \ T - S. f x \ 0}" shows "integral S f = integral T f" using has_integral_spike_set_eq [OF assms] by (metis eq_integralD integral_unique) lemma integrable_spike_set: fixes f :: "'n::euclidean_space \ 'a::banach" assumes f: "f integrable_on S" and neg: "negligible {x \ S - T. f x \ 0}" "negligible {x \ T - S. f x \ 0}" shows "f integrable_on T" using has_integral_spike_set_eq [OF neg] f by blast lemma integrable_spike_set_eq: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "negligible ((S - T) \ (T - S))" shows "f integrable_on S \ f integrable_on T" by (blast intro: integrable_spike_set assms negligible_subset) lemma integrable_on_insert_iff: "f integrable_on (insert x X) \ f integrable_on X" for f::"_ \ 'a::banach" by (rule integrable_spike_set_eq) (auto simp: insert_Diff_if) lemma has_integral_interior: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "negligible(frontier S) \ (f has_integral y) (interior S) \ (f has_integral y) S" by (rule has_integral_spike_set_eq [OF empty_imp_negligible negligible_subset]) (use interior_subset in \auto simp: frontier_def closure_def\) lemma has_integral_closure: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "negligible(frontier S) \ (f has_integral y) (closure S) \ (f has_integral y) S" by (rule has_integral_spike_set_eq [OF negligible_subset empty_imp_negligible]) (auto simp: closure_Un_frontier ) lemma has_integral_open_interval: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "(f has_integral y) (box a b) \ (f has_integral y) (cbox a b)" unfolding interior_cbox [symmetric] by (metis frontier_cbox has_integral_interior negligible_frontier_interval) lemma integrable_on_open_interval: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "f integrable_on box a b \ f integrable_on cbox a b" by (simp add: has_integral_open_interval integrable_on_def) lemma integral_open_interval: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "integral(box a b) f = integral(cbox a b) f" by (metis has_integral_integrable_integral has_integral_open_interval not_integrable_integral) lemma integrable_on_open_interval_real: fixes f :: "real \ 'b :: banach" shows "f integrable_on {a<.. f integrable_on {a..b}" using integrable_on_open_interval[of f a b] by simp lemma integral_open_interval_real: "integral {a..b} (f :: real \ 'a :: banach) = integral {a<..<(b::real)} f" using integral_open_interval[of a b f] by simp lemma has_integral_Icc_iff_Ioo: fixes f :: "real \ 'a :: banach" shows "(f has_integral I) {a..b} \ (f has_integral I) {a<.. 'a :: banach" shows "f integrable_on {a..b} \ f integrable_on {a<..More lemmas that are useful later\ lemma has_integral_subset_le: fixes f :: "'n::euclidean_space \ real" assumes "s \ t" and "(f has_integral i) s" and "(f has_integral j) t" and "\x\t. 0 \ f x" shows "i \ j" using assms has_integral_subset_component_le[OF _ assms(1), of 1 f i j] by auto lemma integral_subset_component_le: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "k \ Basis" and "s \ t" and "f integrable_on s" and "f integrable_on t" and "\x \ t. 0 \ f x \ k" shows "(integral s f)\k \ (integral t f)\k" by (meson assms has_integral_subset_component_le integrable_integral) lemma integral_subset_le: fixes f :: "'n::euclidean_space \ real" assumes "s \ t" and "f integrable_on s" and "f integrable_on t" and "\x \ t. 0 \ f x" shows "integral s f \ integral t f" using assms has_integral_subset_le by blast lemma has_integral_alt': fixes f :: "'n::euclidean_space \ 'a::banach" shows "(f has_integral i) s \ (\a b. (\x. if x \ s then f x else 0) integrable_on cbox a b) \ (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) (\x. if x \ s then f x else 0) - i) < e)" (is "?l = ?r") proof assume rhs: ?r show ?l proof (subst has_integral', intro allI impI) fix e::real assume "e > 0" from rhs[THEN conjunct2,rule_format,OF this] show "\B>0. \a b. ball 0 B \ cbox a b \ (\z. ((\x. if x \ s then f x else 0) has_integral z) (cbox a b) \ norm (z - i) < e)" by (simp add: has_integral_iff rhs) qed next let ?\ = "\e a b. \z. ((\x. if x \ s then f x else 0) has_integral z) (cbox a b) \ norm (z - i) < e" assume ?l then have lhs: "\B>0. \a b. ball 0 B \ cbox a b \ ?\ e a b" if "e > 0" for e using that has_integral'[of f] by auto let ?f = "\x. if x \ s then f x else 0" show ?r proof (intro conjI allI impI) fix a b :: 'n from lhs[OF zero_less_one] obtain B where "0 < B" and B: "\a b. ball 0 B \ cbox a b \ ?\ 1 a b" by blast let ?a = "\i\Basis. min (a\i) (-B) *\<^sub>R i::'n" let ?b = "\i\Basis. max (b\i) B *\<^sub>R i::'n" show "?f integrable_on cbox a b" proof (rule integrable_subinterval[of _ ?a ?b]) have "?a \ i \ x \ i \ x \ i \ ?b \ i" if "norm (0 - x) < B" "i \ Basis" for x i using Basis_le_norm[of i x] that by (auto simp add:field_simps) then have "ball 0 B \ cbox ?a ?b" by (auto simp: mem_box dist_norm) then show "?f integrable_on cbox ?a ?b" unfolding integrable_on_def using B by blast show "cbox a b \ cbox ?a ?b" by (force simp: mem_box) qed fix e :: real assume "e > 0" with lhs show "\B>0. \a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) (\x. if x \ s then f x else 0) - i) < e" by (metis (no_types, lifting) has_integral_integrable_integral) qed qed subsection \Continuity of the integral (for a 1-dimensional interval)\ lemma integrable_alt: fixes f :: "'n::euclidean_space \ 'a::banach" shows "f integrable_on s \ (\a b. (\x. if x \ s then f x else 0) integrable_on cbox a b) \ (\e>0. \B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) (\x. if x \ s then f x else 0) - integral (cbox c d) (\x. if x \ s then f x else 0)) < e)" (is "?l = ?r") proof let ?F = "\x. if x \ s then f x else 0" assume ?l then obtain y where intF: "\a b. ?F integrable_on cbox a b" and y: "\e. 0 < e \ \B>0. \a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) ?F - y) < e" unfolding integrable_on_def has_integral_alt'[of f] by auto show ?r proof (intro conjI allI impI intF) fix e::real assume "e > 0" then have "e/2 > 0" by auto obtain B where "0 < B" and B: "\a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) ?F - y) < e/2" using \0 < e/2\ y by blast show "\B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) ?F - integral (cbox c d) ?F) < e" proof (intro conjI exI impI allI, rule \0 < B\) fix a b c d::'n assume sub: "ball 0 B \ cbox a b \ ball 0 B \ cbox c d" show "norm (integral (cbox a b) ?F - integral (cbox c d) ?F) < e" using sub by (auto intro: norm_triangle_half_l dest: B) qed qed next let ?F = "\x. if x \ s then f x else 0" assume rhs: ?r let ?cube = "\n. cbox (\i\Basis. - real n *\<^sub>R i::'n) (\i\Basis. real n *\<^sub>R i)" have "Cauchy (\n. integral (?cube n) ?F)" unfolding Cauchy_def proof (intro allI impI) fix e::real assume "e > 0" with rhs obtain B where "0 < B" and B: "\a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) ?F - integral (cbox c d) ?F) < e" by blast obtain N where N: "B \ real N" using real_arch_simple by blast have "ball 0 B \ ?cube n" if n: "n \ N" for n proof - have "sum ((*\<^sub>R) (- real n)) Basis \ i \ x \ i \ x \ i \ sum ((*\<^sub>R) (real n)) Basis \ i" if "norm x < B" "i \ Basis" for x i::'n using Basis_le_norm[of i x] n N that by (auto simp add: field_simps sum_negf) then show ?thesis by (auto simp: mem_box dist_norm) qed then show "\M. \m\M. \n\M. dist (integral (?cube m) ?F) (integral (?cube n) ?F) < e" by (fastforce simp add: dist_norm intro!: B) qed then obtain i where i: "(\n. integral (?cube n) ?F) \ i" using convergent_eq_Cauchy by blast have "\B>0. \a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) ?F - i) < e" if "e > 0" for e proof - have *: "e/2 > 0" using that by auto then obtain N where N: "\n. N \ n \ norm (i - integral (?cube n) ?F) < e/2" using i[THEN LIMSEQ_D, simplified norm_minus_commute] by meson obtain B where "0 < B" and B: "\a b c d. \ball 0 B \ cbox a b; ball 0 B \ cbox c d\ \ norm (integral (cbox a b) ?F - integral (cbox c d) ?F) < e/2" using rhs * by meson let ?B = "max (real N) B" show ?thesis proof (intro exI conjI allI impI) show "0 < ?B" using \B > 0\ by auto fix a b :: 'n assume "ball 0 ?B \ cbox a b" moreover obtain n where n: "max (real N) B \ real n" using real_arch_simple by blast moreover have "ball 0 B \ ?cube n" proof fix x :: 'n assume x: "x \ ball 0 B" have "\norm (0 - x) < B; i \ Basis\ \ sum ((*\<^sub>R) (-n)) Basis \ i\ x \ i \ x \ i \ sum ((*\<^sub>R) n) Basis \ i" for i using Basis_le_norm[of i x] n by (auto simp add: field_simps sum_negf) then show "x \ ?cube n" using x by (auto simp: mem_box dist_norm) qed ultimately show "norm (integral (cbox a b) ?F - i) < e" using norm_triangle_half_l [OF B N] by force qed qed then show ?l unfolding integrable_on_def has_integral_alt'[of f] using rhs by blast qed lemma integrable_altD: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on s" shows "\a b. (\x. if x \ s then f x else 0) integrable_on cbox a b" and "\e. e > 0 \ \B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) (\x. if x \ s then f x else 0) - integral (cbox c d) (\x. if x \ s then f x else 0)) < e" using assms[unfolded integrable_alt[of f]] by auto lemma integrable_alt_subset: fixes f :: "'a::euclidean_space \ 'b::banach" shows "f integrable_on S \ (\a b. (\x. if x \ S then f x else 0) integrable_on cbox a b) \ (\e>0. \B>0. \a b c d. ball 0 B \ cbox a b \ cbox a b \ cbox c d \ norm(integral (cbox a b) (\x. if x \ S then f x else 0) - integral (cbox c d) (\x. if x \ S then f x else 0)) < e)" (is "_ = ?rhs") proof - let ?g = "\x. if x \ S then f x else 0" have "f integrable_on S \ (\a b. ?g integrable_on cbox a b) \ (\e>0. \B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e)" by (rule integrable_alt) also have "\ = ?rhs" proof - { fix e :: "real" assume e: "\e. e>0 \ \B>0. \a b c d. ball 0 B \ cbox a b \ cbox a b \ cbox c d \ norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e" and "e > 0" obtain B where "B > 0" and B: "\a b c d. \ball 0 B \ cbox a b; cbox a b \ cbox c d\ \ norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e/2" using \e > 0\ e [of "e/2"] by force have "\B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e" proof (intro exI allI conjI impI) fix a b c d :: "'a" let ?\ = "\i\Basis. max (a \ i) (c \ i) *\<^sub>R i" let ?\ = "\i\Basis. min (b \ i) (d \ i) *\<^sub>R i" show "norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e" if ball: "ball 0 B \ cbox a b \ ball 0 B \ cbox c d" proof - have B': "norm (integral (cbox a b \ cbox c d) ?g - integral (cbox x y) ?g) < e/2" if "cbox a b \ cbox c d \ cbox x y" for x y using B [of ?\ ?\ x y] ball that by (simp add: Int_interval [symmetric]) show ?thesis using B' [of a b] B' [of c d] norm_triangle_half_r by blast qed qed (use \B > 0\ in auto)} then show ?thesis by force qed finally show ?thesis . qed lemma integrable_on_subcbox: fixes f :: "'n::euclidean_space \ 'a::banach" assumes intf: "f integrable_on S" and sub: "cbox a b \ S" shows "f integrable_on cbox a b" proof - have "(\x. if x \ S then f x else 0) integrable_on cbox a b" by (simp add: intf integrable_altD(1)) then show ?thesis by (metis (mono_tags) sub integrable_restrict_Int le_inf_iff order_refl subset_antisym) qed subsection \A straddling criterion for integrability\ lemma integrable_straddle_interval: fixes f :: "'n::euclidean_space \ real" assumes "\e. e>0 \ \g h i j. (g has_integral i) (cbox a b) \ (h has_integral j) (cbox a b) \ \i - j\ < e \ (\x\cbox a b. (g x) \ f x \ f x \ h x)" shows "f integrable_on cbox a b" proof - have "\d. gauge d \ (\p1 p2. p1 tagged_division_of cbox a b \ d fine p1 \ p2 tagged_division_of cbox a b \ d fine p2 \ \(\(x,K)\p1. content K *\<^sub>R f x) - (\(x,K)\p2. content K *\<^sub>R f x)\ < e)" if "e > 0" for e proof - have e: "e/3 > 0" using that by auto then obtain g h i j where ij: "\i - j\ < e/3" and "(g has_integral i) (cbox a b)" and "(h has_integral j) (cbox a b)" and fgh: "\x. x \ cbox a b \ g x \ f x \ f x \ h x" using assms real_norm_def by metis then obtain d1 d2 where "gauge d1" "gauge d2" and d1: "\p. \p tagged_division_of cbox a b; d1 fine p\ \ \(\(x,K)\p. content K *\<^sub>R g x) - i\ < e/3" and d2: "\p. \p tagged_division_of cbox a b; d2 fine p\ \ \(\(x,K) \ p. content K *\<^sub>R h x) - j\ < e/3" by (metis e has_integral real_norm_def) have "\(\(x,K) \ p1. content K *\<^sub>R f x) - (\(x,K) \ p2. content K *\<^sub>R f x)\ < e" if p1: "p1 tagged_division_of cbox a b" and 11: "d1 fine p1" and 21: "d2 fine p1" and p2: "p2 tagged_division_of cbox a b" and 12: "d1 fine p2" and 22: "d2 fine p2" for p1 p2 proof - have *: "\g1 g2 h1 h2 f1 f2. \\g2 - i\ < e/3; \g1 - i\ < e/3; \h2 - j\ < e/3; \h1 - j\ < e/3; g1 - h2 \ f1 - f2; f1 - f2 \ h1 - g2\ \ \f1 - f2\ < e" using \e > 0\ ij by arith have 0: "(\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p1. content k *\<^sub>R g x) \ 0" "0 \ (\(x, k)\p2. content k *\<^sub>R h x) - (\(x, k)\p2. content k *\<^sub>R f x)" "(\(x, k)\p2. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R g x) \ 0" "0 \ (\(x, k)\p1. content k *\<^sub>R h x) - (\(x, k)\p1. content k *\<^sub>R f x)" unfolding sum_subtractf[symmetric] apply (auto intro!: sum_nonneg) apply (meson fgh measure_nonneg mult_left_mono tag_in_interval that sum_nonneg)+ done show ?thesis proof (rule *) show "\(\(x,K) \ p2. content K *\<^sub>R g x) - i\ < e/3" by (rule d1[OF p2 12]) show "\(\(x,K) \ p1. content K *\<^sub>R g x) - i\ < e/3" by (rule d1[OF p1 11]) show "\(\(x,K) \ p2. content K *\<^sub>R h x) - j\ < e/3" by (rule d2[OF p2 22]) show "\(\(x,K) \ p1. content K *\<^sub>R h x) - j\ < e/3" by (rule d2[OF p1 21]) qed (use 0 in auto) qed then show ?thesis by (rule_tac x="\x. d1 x \ d2 x" in exI) (auto simp: fine_Int intro: \gauge d1\ \gauge d2\ d1 d2) qed then show ?thesis by (simp add: integrable_Cauchy) qed lemma integrable_straddle: fixes f :: "'n::euclidean_space \ real" assumes "\e. e>0 \ \g h i j. (g has_integral i) s \ (h has_integral j) s \ \i - j\ < e \ (\x\s. g x \ f x \ f x \ h x)" shows "f integrable_on s" proof - let ?fs = "(\x. if x \ s then f x else 0)" have "?fs integrable_on cbox a b" for a b proof (rule integrable_straddle_interval) fix e::real assume "e > 0" then have *: "e/4 > 0" by auto with assms obtain g h i j where g: "(g has_integral i) s" and h: "(h has_integral j) s" and ij: "\i - j\ < e/4" and fgh: "\x. x \ s \ g x \ f x \ f x \ h x" by metis let ?gs = "(\x. if x \ s then g x else 0)" let ?hs = "(\x. if x \ s then h x else 0)" obtain Bg where Bg: "\a b. ball 0 Bg \ cbox a b \ \integral (cbox a b) ?gs - i\ < e/4" and int_g: "\a b. ?gs integrable_on cbox a b" using g * unfolding has_integral_alt' real_norm_def by meson obtain Bh where Bh: "\a b. ball 0 Bh \ cbox a b \ \integral (cbox a b) ?hs - j\ < e/4" and int_h: "\a b. ?hs integrable_on cbox a b" using h * unfolding has_integral_alt' real_norm_def by meson define c where "c = (\i\Basis. min (a\i) (- (max Bg Bh)) *\<^sub>R i)" define d where "d = (\i\Basis. max (b\i) (max Bg Bh) *\<^sub>R i)" have "\norm (0 - x) < Bg; i \ Basis\ \ c \ i \ x \ i \ x \ i \ d \ i" for x i using Basis_le_norm[of i x] unfolding c_def d_def by auto then have ballBg: "ball 0 Bg \ cbox c d" by (auto simp: mem_box dist_norm) have "\norm (0 - x) < Bh; i \ Basis\ \ c \ i \ x \ i \ x \ i \ d \ i" for x i using Basis_le_norm[of i x] unfolding c_def d_def by auto then have ballBh: "ball 0 Bh \ cbox c d" by (auto simp: mem_box dist_norm) have ab_cd: "cbox a b \ cbox c d" by (auto simp: c_def d_def subset_box_imp) have **: "\ch cg ag ah::real. \\ah - ag\ \ \ch - cg\; \cg - i\ < e/4; \ch - j\ < e/4\ \ \ag - ah\ < e" using ij by arith show "\g h i j. (g has_integral i) (cbox a b) \ (h has_integral j) (cbox a b) \ \i - j\ < e \ (\x\cbox a b. g x \ (if x \ s then f x else 0) \ (if x \ s then f x else 0) \ h x)" proof (intro exI ballI conjI) have eq: "\x f g. (if x \ s then f x else 0) - (if x \ s then g x else 0) = (if x \ s then f x - g x else (0::real))" by auto have int_hg: "(\x. if x \ s then h x - g x else 0) integrable_on cbox a b" "(\x. if x \ s then h x - g x else 0) integrable_on cbox c d" by (metis (no_types) integrable_diff g h has_integral_integrable integrable_altD(1))+ show "(?gs has_integral integral (cbox a b) ?gs) (cbox a b)" "(?hs has_integral integral (cbox a b) ?hs) (cbox a b)" by (intro integrable_integral int_g int_h)+ then have "integral (cbox a b) ?gs \ integral (cbox a b) ?hs" using fgh by (force intro: has_integral_le) then have "0 \ integral (cbox a b) ?hs - integral (cbox a b) ?gs" by simp then have "\integral (cbox a b) ?hs - integral (cbox a b) ?gs\ \ \integral (cbox c d) ?hs - integral (cbox c d) ?gs\" apply (simp add: integral_diff [symmetric] int_g int_h) apply (subst abs_of_nonneg[OF integral_nonneg[OF integrable_diff, OF int_h int_g]]) using fgh apply (force simp: eq intro!: integral_subset_le [OF ab_cd int_hg])+ done then show "\integral (cbox a b) ?gs - integral (cbox a b) ?hs\ < e" using ** Bg ballBg Bh ballBh by blast show "\x. x \ cbox a b \ ?gs x \ ?fs x" "\x. x \ cbox a b \ ?fs x \ ?hs x" using fgh by auto qed qed then have int_f: "?fs integrable_on cbox a b" for a b by simp have "\B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ abs (integral (cbox a b) ?fs - integral (cbox c d) ?fs) < e" if "0 < e" for e proof - have *: "e/3 > 0" using that by auto with assms obtain g h i j where g: "(g has_integral i) s" and h: "(h has_integral j) s" and ij: "\i - j\ < e/3" and fgh: "\x. x \ s \ g x \ f x \ f x \ h x" by metis let ?gs = "(\x. if x \ s then g x else 0)" let ?hs = "(\x. if x \ s then h x else 0)" obtain Bg where "Bg > 0" and Bg: "\a b. ball 0 Bg \ cbox a b \ \integral (cbox a b) ?gs - i\ < e/3" and int_g: "\a b. ?gs integrable_on cbox a b" using g * unfolding has_integral_alt' real_norm_def by meson obtain Bh where "Bh > 0" and Bh: "\a b. ball 0 Bh \ cbox a b \ \integral (cbox a b) ?hs - j\ < e/3" and int_h: "\a b. ?hs integrable_on cbox a b" using h * unfolding has_integral_alt' real_norm_def by meson { fix a b c d :: 'n assume as: "ball 0 (max Bg Bh) \ cbox a b" "ball 0 (max Bg Bh) \ cbox c d" have **: "ball 0 Bg \ ball (0::'n) (max Bg Bh)" "ball 0 Bh \ ball (0::'n) (max Bg Bh)" by auto have *: "\ga gc ha hc fa fc. \\ga - i\ < e/3; \gc - i\ < e/3; \ha - j\ < e/3; \hc - j\ < e/3; ga \ fa; fa \ ha; gc \ fc; fc \ hc\ \ \fa - fc\ < e" using ij by arith have "abs (integral (cbox a b) (\x. if x \ s then f x else 0) - integral (cbox c d) (\x. if x \ s then f x else 0)) < e" proof (rule *) show "\integral (cbox a b) ?gs - i\ < e/3" using "**" Bg as by blast show "\integral (cbox c d) ?gs - i\ < e/3" using "**" Bg as by blast show "\integral (cbox a b) ?hs - j\ < e/3" using "**" Bh as by blast show "\integral (cbox c d) ?hs - j\ < e/3" using "**" Bh as by blast qed (use int_f int_g int_h fgh in \simp_all add: integral_le\) } then show ?thesis apply (rule_tac x="max Bg Bh" in exI) using \Bg > 0\ by auto qed then show ?thesis unfolding integrable_alt[of f] real_norm_def by (blast intro: int_f) qed subsection \Adding integrals over several sets\ lemma has_integral_Un: fixes f :: "'n::euclidean_space \ 'a::banach" assumes f: "(f has_integral i) S" "(f has_integral j) T" and neg: "negligible (S \ T)" shows "(f has_integral (i + j)) (S \ T)" unfolding has_integral_restrict_UNIV[symmetric, of f] proof (rule has_integral_spike[OF neg]) let ?f = "\x. (if x \ S then f x else 0) + (if x \ T then f x else 0)" show "(?f has_integral i + j) UNIV" by (simp add: f has_integral_add) qed auto lemma integral_Un [simp]: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on S" "f integrable_on T" "negligible (S \ T)" shows "integral (S \ T) f = integral S f + integral T f" by (simp add: has_integral_Un assms integrable_integral integral_unique) lemma integrable_Un: fixes f :: "'a::euclidean_space \ 'b :: banach" assumes "negligible (A \ B)" "f integrable_on A" "f integrable_on B" shows "f integrable_on (A \ B)" proof - from assms obtain y z where "(f has_integral y) A" "(f has_integral z) B" by (auto simp: integrable_on_def) from has_integral_Un[OF this assms(1)] show ?thesis by (auto simp: integrable_on_def) qed lemma integrable_Un': fixes f :: "'a::euclidean_space \ 'b :: banach" assumes "f integrable_on A" "f integrable_on B" "negligible (A \ B)" "C = A \ B" shows "f integrable_on C" using integrable_Un[of A B f] assms by simp lemma has_integral_UN: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "finite I" and int: "\i. i \ I \ (f has_integral (g i)) (\ i)" and neg: "pairwise (\i i'. negligible (\ i \ \ i')) I" shows "(f has_integral (sum g I)) (\i\I. \ i)" proof - let ?\ = "((\(a,b). \ a \ \ b) ` {(a,b). a \ I \ b \ I-{a}})" have "((\x. if x \ (\i\I. \ i) then f x else 0) has_integral sum g I) UNIV" proof (rule has_integral_spike) show "negligible (\?\)" proof (rule negligible_Union) have "finite (I \ I)" by (simp add: \finite I\) moreover have "{(a,b). a \ I \ b \ I-{a}} \ I \ I" by auto ultimately show "finite ?\" by (simp add: finite_subset) show "\t. t \ ?\ \ negligible t" using neg unfolding pairwise_def by auto qed next show "(if x \ (\i\I. \ i) then f x else 0) = (\i\I. if x \ \ i then f x else 0)" if "x \ UNIV - (\?\)" for x proof clarsimp fix i assume i: "i \ I" "x \ \ i" then have "\j\I. x \ \ j \ j = i" using that by blast with i show "f x = (\i\I. if x \ \ i then f x else 0)" by (simp add: sum.delta[OF \finite I\]) qed next show "((\x. (\i\I. if x \ \ i then f x else 0)) has_integral sum g I) UNIV" using int by (simp add: has_integral_restrict_UNIV has_integral_sum [OF \finite I\]) qed then show ?thesis using has_integral_restrict_UNIV by blast qed lemma has_integral_Union: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "finite \" and "\S. S \ \ \ (f has_integral (i S)) S" and "pairwise (\S S'. negligible (S \ S')) \" shows "(f has_integral (sum i \)) (\\)" proof - have "(f has_integral (sum i \)) (\S\\. S)" by (intro has_integral_UN assms) then show ?thesis by force qed text \In particular adding integrals over a division, maybe not of an interval.\ lemma has_integral_combine_division: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "\ division_of S" and "\k. k \ \ \ (f has_integral (i k)) k" shows "(f has_integral (sum i \)) S" proof - note \ = division_ofD[OF assms(1)] have neg: "negligible (S \ s')" if "S \ \" "s' \ \" "S \ s'" for S s' proof - obtain a c b \ where obt: "S = cbox a b" "s' = cbox c \" by (meson \S \ \\ \s' \ \\ \(4)) from \(5)[OF that] show ?thesis unfolding obt interior_cbox by (metis (no_types, lifting) Diff_empty Int_interval box_Int_box negligible_frontier_interval) qed show ?thesis unfolding \(6)[symmetric] by (auto intro: \ neg assms has_integral_Union pairwiseI) qed lemma integral_combine_division_bottomup: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "\ division_of S" "\k. k \ \ \ f integrable_on k" shows "integral S f = sum (\i. integral i f) \" by (meson assms integral_unique has_integral_combine_division has_integral_integrable_integral) lemma has_integral_combine_division_topdown: fixes f :: "'n::euclidean_space \ 'a::banach" assumes f: "f integrable_on S" and \: "\ division_of K" and "K \ S" shows "(f has_integral (sum (\i. integral i f) \)) K" proof - have "f integrable_on L" if "L \ \" for L by (smt (verit, best) assms cbox_division_memE f integrable_on_subcbox subset_trans that) then show ?thesis by (meson \ has_integral_combine_division has_integral_integrable_integral) qed lemma integral_combine_division_topdown: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on S" and "\ division_of S" shows "integral S f = sum (\i. integral i f) \" using assms has_integral_combine_division_topdown by blast lemma integrable_combine_division: fixes f :: "'n::euclidean_space \ 'a::banach" assumes \: "\ division_of S" and f: "\i. i \ \ \ f integrable_on i" shows "f integrable_on S" using f unfolding integrable_on_def by (metis has_integral_combine_division[OF \]) lemma integrable_on_subdivision: fixes f :: "'n::euclidean_space \ 'a::banach" assumes \: "\ division_of i" and f: "f integrable_on S" and "i \ S" shows "f integrable_on i" proof - have "f integrable_on i" if "i \ \" for i by (smt (verit, best) assms cbox_division_memE f integrable_on_subcbox order_trans that) then show ?thesis using \ integrable_combine_division by blast qed subsection \Also tagged divisions\ lemma has_integral_combine_tagged_division: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "p tagged_division_of S" and "\x k. (x,k) \ p \ (f has_integral (i k)) k" shows "(f has_integral (\(x,k)\p. i k)) S" proof - have *: "(f has_integral (\k\snd`p. integral k f)) S" by (smt (verit, del_insts) assms division_of_tagged_division has_integral_combine_division has_integral_iff imageE prod.collapse) also have "(\k\snd`p. integral k f) = (\(x, k)\p. integral k f)" by (intro sum.over_tagged_division_lemma[OF assms(1), symmetric] integral_null) (simp add: content_eq_0_interior) finally show ?thesis using assms by (auto simp add: has_integral_iff intro!: sum.cong) qed lemma integral_combine_tagged_division_bottomup: fixes f :: "'n::euclidean_space \ 'a::banach" assumes p: "p tagged_division_of (cbox a b)" and f: "\x k. (x,k)\p \ f integrable_on k" shows "integral (cbox a b) f = sum (\(x,k). integral k f) p" by (simp add: has_integral_combine_tagged_division[OF p] integral_unique f integrable_integral) lemma has_integral_combine_tagged_division_topdown: fixes f :: "'n::euclidean_space \ 'a::banach" assumes f: "f integrable_on cbox a b" and p: "p tagged_division_of (cbox a b)" shows "(f has_integral (sum (\(x,K). integral K f) p)) (cbox a b)" proof - have "(f has_integral integral K f) K" if "(x,K) \ p" for x K by (metis assms integrable_integral integrable_on_subcbox tagged_division_ofD(3,4) that) then show ?thesis by (simp add: has_integral_combine_tagged_division p) qed lemma integral_combine_tagged_division_topdown: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on cbox a b" and "p tagged_division_of (cbox a b)" shows "integral (cbox a b) f = sum (\(x,k). integral k f) p" using assms by (auto intro: integral_unique [OF has_integral_combine_tagged_division_topdown]) subsection \Henstock's lemma\ lemma Henstock_lemma_part1: fixes f :: "'n::euclidean_space \ 'a::banach" assumes intf: "f integrable_on cbox a b" and "e > 0" and "gauge d" and less_e: "\p. \p tagged_division_of (cbox a b); d fine p\ \ norm (sum (\(x,K). content K *\<^sub>R f x) p - integral(cbox a b) f) < e" and p: "p tagged_partial_division_of (cbox a b)" "d fine p" shows "norm (sum (\(x,K). content K *\<^sub>R f x - integral K f) p) \ e" (is "?lhs \ e") proof (rule field_le_epsilon) fix k :: real assume "k > 0" let ?SUM = "\p. (\(x,K) \ p. content K *\<^sub>R f x)" note p' = tagged_partial_division_ofD[OF p(1)] have "\(snd ` p) \ cbox a b" using p'(3) by fastforce then obtain q where q: "snd ` p \ q" and qdiv: "q division_of cbox a b" by (meson p(1) partial_division_extend_interval partial_division_of_tagged_division) note q' = division_ofD[OF qdiv] define r where "r = q - snd ` p" have "snd ` p \ r = {}" unfolding r_def by auto have "finite r" using q' unfolding r_def by auto have "\p. p tagged_division_of i \ d fine p \ norm (?SUM p - integral i f) < k / (real (card r) + 1)" if "i\r" for i proof - have gt0: "k / (real (card r) + 1) > 0" using \k > 0\ by simp have i: "i \ q" using that unfolding r_def by auto then obtain u v where uv: "i = cbox u v" using q'(4) by blast then have "cbox u v \ cbox a b" using i q'(2) by auto then have "f integrable_on cbox u v" by (rule integrable_subinterval[OF intf]) with integrable_integral[OF this, unfolded has_integral[of f]] obtain dd where "gauge dd" and dd: "\\. \\ tagged_division_of cbox u v; dd fine \\ \ norm (?SUM \ - integral (cbox u v) f) < k / (real (card r) + 1)" using gt0 by auto with gauge_Int[OF \gauge d\ \gauge dd\] obtain qq where qq: "qq tagged_division_of cbox u v" "(\x. d x \ dd x) fine qq" using fine_division_exists by blast with dd[of qq] show ?thesis by (auto simp: fine_Int uv) qed then obtain qq where qq: "\i. i \ r \ qq i tagged_division_of i \ d fine qq i \ norm (?SUM (qq i) - integral i f) < k / (real (card r) + 1)" by metis let ?p = "p \ \(qq ` r)" have "norm (?SUM ?p - integral (cbox a b) f) < e" proof (rule less_e) show "d fine ?p" by (metis (mono_tags, opaque_lifting) qq fine_Un fine_Union imageE p(2)) note ptag = tagged_partial_division_of_Union_self[OF p(1)] have "p \ \(qq ` r) tagged_division_of \(snd ` p) \ \r" proof (rule tagged_division_Un[OF ptag tagged_division_Union [OF \finite r\]]) show "\i. i \ r \ qq i tagged_division_of i" using qq by auto show "\i1 i2. \i1 \ r; i2 \ r; i1 \ i2\ \ interior i1 \ interior i2 = {}" by (simp add: q'(5) r_def) show "interior (\(snd ` p)) \ interior (\r) = {}" proof (rule Int_interior_Union_intervals [OF \finite r\]) show "open (interior (\(snd ` p)))" by blast show "\T. T \ r \ \a b. T = cbox a b" by (simp add: q'(4) r_def) have "interior T \ interior (\(snd ` p)) = {}" if "T \ r" for T proof (rule Int_interior_Union_intervals) show "\U. U \ snd ` p \ \a b. U = cbox a b" using q q'(4) by blast show "\U. U \ snd ` p \ interior T \ interior U = {}" by (metis DiffE q q'(5) r_def subsetD that) qed (use p' in auto) then show "\T. T \ r \ interior (\(snd ` p)) \ interior T = {}" by (metis Int_commute) qed qed moreover have "\(snd ` p) \ \r = cbox a b" and "{qq i |i. i \ r} = qq ` r" using qdiv q unfolding Union_Un_distrib[symmetric] r_def by auto ultimately show "?p tagged_division_of (cbox a b)" by fastforce qed then have "norm (?SUM p + (?SUM (\(qq ` r))) - integral (cbox a b) f) < e" proof (subst sum.union_inter_neutral[symmetric, OF \finite p\], safe) show "content L *\<^sub>R f x = 0" if "(x, L) \ p" "(x, L) \ qq K" "K \ r" for x K L proof - obtain u v where uv: "L = cbox u v" using \(x,L) \ p\ p'(4) by blast have "L \ K" using qq[OF that(3)] tagged_division_ofD(3) \(x,L) \ qq K\ by metis have "L \ snd ` p" using \(x,L) \ p\ image_iff by fastforce then have "L \ q" "K \ q" "L \ K" using that q(1) unfolding r_def by auto with q'(5) show "content L *\<^sub>R f x = 0" by (metis \L \ K\ content_eq_0_interior inf.orderE interior_Int scaleR_eq_0_iff uv) qed show "finite (\(qq ` r))" by (meson finite_UN qq \finite r\ tagged_division_of_finite) qed moreover have "content M *\<^sub>R f x = 0" if x: "(x,M) \ qq K" "(x,M) \ qq L" and KL: "qq K \ qq L" and r: "K \ r" "L \ r" for x M K L proof - note kl = tagged_division_ofD(3,4)[OF qq[THEN conjunct1]] obtain u v where uv: "M = cbox u v" using \(x, M) \ qq L\ \L \ r\ kl(2) by blast have empty: "interior (K \ L) = {}" by (metis DiffD1 interior_Int q'(5) r_def KL r) with that kl show "content M *\<^sub>R f x = 0" by (metis content_eq_0_interior dual_order.refl inf.orderE inf_mono interior_mono scaleR_eq_0_iff subset_empty uv x) qed ultimately have "norm (?SUM p + sum ?SUM (qq ` r) - integral (cbox a b) f) < e" apply (subst (asm) sum.Union_comp) using qq by (force simp: split_paired_all)+ moreover have "content M *\<^sub>R f x = 0" if "K \ r" "L \ r" "K \ L" "qq K = qq L" "(x, M) \ qq K" for K L x M using tagged_division_ofD(6) qq that by (metis (no_types, lifting)) ultimately have less_e: "norm (?SUM p + sum (?SUM \ qq) r - integral (cbox a b) f) < e" proof (subst (asm) sum.reindex_nontrivial [OF \finite r\]) qed (auto simp: split_paired_all sum.neutral) have norm_le: "norm (cp - ip) \ e + k" if "norm ((cp + cr) - i) < e" "norm (cr - ir) < k" "ip + ir = i" for ir ip i cr cp::'a using norm_triangle_le[of "cp + cr - i" "- (cr - ir)"] that unfolding that(3)[symmetric] norm_minus_cancel by (auto simp add: algebra_simps) have "?lhs = norm (?SUM p - (\(x, k)\p. integral k f))" unfolding split_def sum_subtractf .. also have "\ \ e + k" proof (rule norm_le[OF less_e]) have lessk: "k * real (card r) / (1 + real (card r)) < k" using \k>0\ by (auto simp add: field_simps) have "norm (sum (?SUM \ qq) r - (\k\r. integral k f)) \ (\x\r. k / (real (card r) + 1))" unfolding sum_subtractf[symmetric] by (force dest: qq intro!: sum_norm_le) also have "... < k" by (simp add: lessk add.commute mult.commute) finally show "norm (sum (?SUM \ qq) r - (\k\r. integral k f)) < k" . next from q(1) have [simp]: "snd ` p \ q = q" by auto have "integral l f = 0" if inp: "(x, l) \ p" "(y, m) \ p" and ne: "(x, l) \ (y, m)" and "l = m" for x l y m proof - obtain u v where uv: "l = cbox u v" using inp p'(4) by blast then show ?thesis using uv that p by (metis content_eq_0_interior dual_order.refl inf.orderE integral_null ne tagged_partial_division_ofD(5)) qed then have "(\(x, K)\p. integral K f) = (\K\snd ` p. integral K f)" apply (subst sum.reindex_nontrivial [OF \finite p\]) unfolding split_paired_all split_def by auto then show "(\(x, k)\p. integral k f) + (\k\r. integral k f) = integral (cbox a b) f" unfolding integral_combine_division_topdown[OF intf qdiv] r_def using q'(1) p'(1) sum.union_disjoint [of "snd ` p" "q - snd ` p", symmetric] by simp qed finally show "?lhs \ e + k" . qed lemma Henstock_lemma_part2: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes fed: "f integrable_on cbox a b" "e > 0" "gauge d" and less_e: "\\. \\ tagged_division_of (cbox a b); d fine \\ \ norm (sum (\(x,k). content k *\<^sub>R f x) \ - integral (cbox a b) f) < e" and tag: "p tagged_partial_division_of (cbox a b)" and "d fine p" shows "sum (\(x,k). norm (content k *\<^sub>R f x - integral k f)) p \ 2 * real (DIM('n)) * e" proof - have "finite p" using tag tagged_partial_division_ofD by blast then show ?thesis unfolding split_def proof (rule sum_norm_allsubsets_bound) fix Q assume Q: "Q \ p" then have fine: "d fine Q" by (simp add: \d fine p\ fine_subset) show "norm (\x\Q. content (snd x) *\<^sub>R f (fst x) - integral (snd x) f) \ e" apply (rule Henstock_lemma_part1[OF fed less_e, unfolded split_def]) using Q tag tagged_partial_division_subset by (force simp add: fine)+ qed qed lemma Henstock_lemma: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes intf: "f integrable_on cbox a b" and "e > 0" obtains \ where "gauge \" and "\p. \p tagged_partial_division_of (cbox a b); \ fine p\ \ sum (\(x,k). norm(content k *\<^sub>R f x - integral k f)) p < e" proof - have *: "e/(2 * (real DIM('n) + 1)) > 0" using \e > 0\ by simp with integrable_integral[OF intf, unfolded has_integral] obtain \ where "gauge \" and \: "\\. \\ tagged_division_of cbox a b; \ fine \\ \ norm ((\(x,K)\\. content K *\<^sub>R f x) - integral (cbox a b) f) < e/(2 * (real DIM('n) + 1))" by metis show thesis proof (rule that [OF \gauge \\]) fix p assume p: "p tagged_partial_division_of cbox a b" "\ fine p" have "(\(x,K)\p. norm (content K *\<^sub>R f x - integral K f)) \ 2 * real DIM('n) * (e/(2 * (real DIM('n) + 1)))" using Henstock_lemma_part2[OF intf * \gauge \\ \ p] by metis also have "... < e" using \e > 0\ by (auto simp add: field_simps) finally show "(\(x,K)\p. norm (content K *\<^sub>R f x - integral K f)) < e" . qed qed subsection \Monotone convergence (bounded interval first)\ lemma bounded_increasing_convergent: fixes f :: "nat \ real" shows "\bounded (range f); \n. f n \ f (Suc n)\ \ \l. f \ l" using Bseq_mono_convergent[of f] incseq_Suc_iff[of f] by (auto simp: image_def Bseq_eq_bounded convergent_def incseq_def) lemma monotone_convergence_interval: fixes f :: "nat \ 'n::euclidean_space \ real" assumes intf: "\k. (f k) integrable_on cbox a b" and le: "\k x. x \ cbox a b \ (f k x) \ f (Suc k) x" and fg: "\x. x \ cbox a b \ ((\k. f k x) \ g x) sequentially" and bou: "bounded (range (\k. integral (cbox a b) (f k)))" shows "g integrable_on cbox a b \ ((\k. integral (cbox a b) (f k)) \ integral (cbox a b) g) sequentially" proof (cases "content (cbox a b) = 0") case True then show ?thesis by auto next case False have fg1: "(f k x) \ (g x)" if x: "x \ cbox a b" for x k proof - have "\\<^sub>F j in sequentially. f k x \ f j x" proof (rule eventually_sequentiallyI [of k]) show "\j. k \ j \ f k x \ f j x" using le x by (force intro: transitive_stepwise_le) qed then show "f k x \ g x" using tendsto_lowerbound [OF fg] x trivial_limit_sequentially by blast qed have int_inc: "\n. integral (cbox a b) (f n) \ integral (cbox a b) (f (Suc n))" by (metis integral_le intf le) then obtain i where i: "(\k. integral (cbox a b) (f k)) \ i" using bounded_increasing_convergent bou by blast have "\k. \\<^sub>F x in sequentially. integral (cbox a b) (f k) \ integral (cbox a b) (f x)" unfolding eventually_sequentially by (force intro: transitive_stepwise_le int_inc) then have i': "\k. (integral(cbox a b) (f k)) \ i" using tendsto_le [OF trivial_limit_sequentially i] by blast have "(g has_integral i) (cbox a b)" unfolding has_integral real_norm_def proof clarify fix e::real assume e: "e > 0" have "\k. (\\. gauge \ \ (\\. \ tagged_division_of (cbox a b) \ \ fine \ \ abs ((\(x,K)\\. content K *\<^sub>R f k x) - integral (cbox a b) (f k)) < e/2 ^ (k + 2)))" using intf e by (auto simp: has_integral_integral has_integral) then obtain c where c: "\x. gauge (c x)" "\x \. \\ tagged_division_of cbox a b; c x fine \\ \ abs ((\(u,K)\\. content K *\<^sub>R f x u) - integral (cbox a b) (f x)) < e/2 ^ (x + 2)" by metis have "\r. \k\r. 0 \ i - (integral (cbox a b) (f k)) \ i - (integral (cbox a b) (f k)) < e/4" proof - have "e/4 > 0" using e by auto show ?thesis using LIMSEQ_D [OF i \e/4 > 0\] i' by auto qed then obtain r where r: "\k. r \ k \ 0 \ i - integral (cbox a b) (f k)" "\k. r \ k \ i - integral (cbox a b) (f k) < e/4" by metis have "\n\r. \k\n. 0 \ (g x) - (f k x) \ (g x) - (f k x) < e/(4 * content(cbox a b))" if "x \ cbox a b" for x proof - have "e/(4 * content (cbox a b)) > 0" by (simp add: False content_lt_nz e) with fg that LIMSEQ_D obtain N where "\n\N. norm (f n x - g x) < e/(4 * content (cbox a b))" by metis then show "\n\r. \k\n. 0 \ g x - f k x \ g x - f k x < e/(4 * content (cbox a b))" apply (rule_tac x="N + r" in exI) using fg1[OF that] by (auto simp add: field_simps) qed then obtain m where r_le_m: "\x. x \ cbox a b \ r \ m x" and m: "\x k. \x \ cbox a b; m x \ k\ \ 0 \ g x - f k x \ g x - f k x < e/(4 * content (cbox a b))" by metis define d where "d x = c (m x) x" for x show "\\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ abs ((\(x,K)\\. content K *\<^sub>R g x) - i) < e)" proof (rule exI, safe) show "gauge d" using c(1) unfolding gauge_def d_def by auto next fix \ assume ptag: "\ tagged_division_of (cbox a b)" and "d fine \" note p'=tagged_division_ofD[OF ptag] obtain s where s: "\x. x \ \ \ m (fst x) \ s" by (metis finite_imageI finite_nat_set_iff_bounded_le p'(1) rev_image_eqI) have *: "\a - d\ < e" if "\a - b\ \ e/4" "\b - c\ < e/2" "\c - d\ < e/4" for a b c d using that norm_triangle_lt[of "a - b" "b - c" "3* e/4"] norm_triangle_lt[of "a - b + (b - c)" "c - d" e] by (auto simp add: algebra_simps) show "\(\(x, k)\\. content k *\<^sub>R g x) - i\ < e" proof (rule *) have "\(\(x,K)\\. content K *\<^sub>R g x) - (\(x,K)\\. content K *\<^sub>R f (m x) x)\ \ (\i\\. \(case i of (x, K) \ content K *\<^sub>R g x) - (case i of (x, K) \ content K *\<^sub>R f (m x) x)\)" by (metis (mono_tags) sum_subtractf sum_abs) also have "... \ (\(x, k)\\. content k * (e/(4 * content (cbox a b))))" proof (rule sum_mono, simp add: split_paired_all) fix x K assume xk: "(x,K) \ \" with ptag have x: "x \ cbox a b" by blast then have "abs (content K * (g x - f (m x) x)) \ content K * (e/(4 * content (cbox a b)))" by (metis m[OF x] mult_nonneg_nonneg abs_of_nonneg less_eq_real_def measure_nonneg mult_left_mono order_refl) then show "\content K * g x - content K * f (m x) x\ \ content K * e/(4 * content (cbox a b))" by (simp add: algebra_simps) qed also have "... = (e/(4 * content (cbox a b))) * (\(x, k)\\. content k)" by (simp add: sum_distrib_left sum_divide_distrib split_def mult.commute) also have "... \ e/4" by (metis False additive_content_tagged_division [OF ptag] nonzero_mult_divide_mult_cancel_right order_refl times_divide_eq_left) finally show "\(\(x,K)\\. content K *\<^sub>R g x) - (\(x,K)\\. content K *\<^sub>R f (m x) x)\ \ e/4" . next have "norm ((\(x,K)\\. content K *\<^sub>R f (m x) x) - (\(x,K)\\. integral K (f (m x)))) \ norm (\j = 0..s. \(x,K)\{xk \ \. m (fst xk) = j}. content K *\<^sub>R f (m x) x - integral K (f (m x)))" apply (subst sum.group) using s by (auto simp: sum_subtractf split_def p'(1)) also have "\ < e/2" proof - have "norm (\j = 0..s. \(x, k)\{xk \ \. m (fst xk) = j}. content k *\<^sub>R f (m x) x - integral k (f (m x))) \ (\i = 0..s. e/2 ^ (i + 2))" proof (rule sum_norm_le) fix t assume "t \ {0..s}" have "norm (\(x,k)\{xk \ \. m (fst xk) = t}. content k *\<^sub>R f (m x) x - integral k (f (m x))) = norm (\(x,k)\{xk \ \. m (fst xk) = t}. content k *\<^sub>R f t x - integral k (f t))" by (force intro!: sum.cong arg_cong[where f=norm]) also have "... \ e/2 ^ (t + 2)" proof (rule Henstock_lemma_part1 [OF intf]) show "{xk \ \. m (fst xk) = t} tagged_partial_division_of cbox a b" proof (rule tagged_partial_division_subset[of \]) show "\ tagged_partial_division_of cbox a b" using ptag tagged_division_of_def by blast qed auto show "c t fine {xk \ \. m (fst xk) = t}" using \d fine \\ by (auto simp: fine_def d_def) qed (use c e in auto) finally show "norm (\(x,K)\{xk \ \. m (fst xk) = t}. content K *\<^sub>R f (m x) x - integral K (f (m x))) \ e/2 ^ (t + 2)" . qed also have "... = (e/2/2) * (\i = 0..s. (1/2) ^ i)" by (simp add: sum_distrib_left field_simps) also have "\ < e/2" by (simp add: sum_gp mult_strict_left_mono[OF _ e]) finally show "norm (\j = 0..s. \(x, k)\{xk \ \. m (fst xk) = j}. content k *\<^sub>R f (m x) x - integral k (f (m x))) < e/2" . qed finally show "\(\(x,K)\\. content K *\<^sub>R f (m x) x) - (\(x,K)\\. integral K (f (m x)))\ < e/2" by simp next have comb: "integral (cbox a b) (f y) = (\(x, k)\\. integral k (f y))" for y using integral_combine_tagged_division_topdown[OF intf ptag] by metis have f_le: "\y m n. \y \ cbox a b; n\m\ \ f m y \ f n y" using le by (auto intro: transitive_stepwise_le) have "(\(x, k)\\. integral k (f r)) \ (\(x, K)\\. integral K (f (m x)))" proof (rule sum_mono, simp add: split_paired_all) fix x K assume xK: "(x, K) \ \" show "integral K (f r) \ integral K (f (m x))" proof (rule integral_le) show "f r integrable_on K" by (metis integrable_on_subcbox intf p'(3) p'(4) xK) show "f (m x) integrable_on K" by (metis elementary_interval integrable_on_subdivision intf p'(3) p'(4) xK) show "f r y \ f (m x) y" if "y \ K" for y using that r_le_m[of x] p'(2-3)[OF xK] f_le by auto qed qed moreover have "(\(x, K)\\. integral K (f (m x))) \ (\(x, k)\\. integral k (f s))" proof (rule sum_mono, simp add: split_paired_all) fix x K assume xK: "(x, K) \ \" show "integral K (f (m x)) \ integral K (f s)" proof (rule integral_le) show "f (m x) integrable_on K" by (metis elementary_interval integrable_on_subdivision intf p'(3) p'(4) xK) show "f s integrable_on K" by (metis integrable_on_subcbox intf p'(3) p'(4) xK) show "f (m x) y \ f s y" if "y \ K" for y using that s xK f_le p'(3) by fastforce qed qed moreover have "0 \ i - integral (cbox a b) (f r)" "i - integral (cbox a b) (f r) < e/4" using r by auto ultimately show "\(\(x,K)\\. integral K (f (m x))) - i\ < e/4" using comb i'[of s] by auto qed qed qed with i integral_unique show ?thesis by blast qed lemma monotone_convergence_increasing: fixes f :: "nat \ 'n::euclidean_space \ real" assumes int_f: "\k. (f k) integrable_on S" and "\k x. x \ S \ (f k x) \ (f (Suc k) x)" and fg: "\x. x \ S \ ((\k. f k x) \ g x) sequentially" and bou: "bounded (range (\k. integral S (f k)))" shows "g integrable_on S \ ((\k. integral S (f k)) \ integral S g) sequentially" proof - have lem: "g integrable_on S \ ((\k. integral S (f k)) \ integral S g) sequentially" if f0: "\k x. x \ S \ 0 \ f k x" and int_f: "\k. (f k) integrable_on S" and le: "\k x. x \ S \ f k x \ f (Suc k) x" and lim: "\x. x \ S \ ((\k. f k x) \ g x) sequentially" and bou: "bounded (range(\k. integral S (f k)))" for f :: "nat \ 'n::euclidean_space \ real" and g S proof - have fg: "(f k x) \ (g x)" if "x \ S" for x k proof - have "\xa. k \ xa \ f k x \ f xa x" using le by (force intro: transitive_stepwise_le that) then show ?thesis using tendsto_lowerbound [OF lim [OF that]] eventually_sequentiallyI by force qed obtain i where i: "(\k. integral S (f k)) \ i" using bounded_increasing_convergent [OF bou] le int_f integral_le by blast have i': "(integral S (f k)) \ i" for k proof - have "\k. \x. x \ S \ \n\k. f k x \ f n x" using le by (force intro: transitive_stepwise_le) then show ?thesis using tendsto_lowerbound [OF i eventually_sequentiallyI trivial_limit_sequentially] by (meson int_f integral_le) qed let ?f = "(\k x. if x \ S then f k x else 0)" let ?g = "(\x. if x \ S then g x else 0)" have int: "?f k integrable_on cbox a b" for a b k by (simp add: int_f integrable_altD(1)) have int': "\k a b. f k integrable_on cbox a b \ S" using int by (simp add: Int_commute integrable_restrict_Int) have g: "?g integrable_on cbox a b \ (\k. integral (cbox a b) (?f k)) \ integral (cbox a b) ?g" for a b proof (rule monotone_convergence_interval) have "norm (integral (cbox a b) (?f k)) \ norm (integral S (f k))" for k proof - have "0 \ integral (cbox a b) (?f k)" by (metis (no_types) integral_nonneg Int_iff f0 inf_commute integral_restrict_Int int') moreover have "0 \ integral S (f k)" by (simp add: integral_nonneg f0 int_f) moreover have "integral (S \ cbox a b) (f k) \ integral S (f k)" by (metis f0 inf_commute int' int_f integral_subset_le le_inf_iff order_refl) ultimately show ?thesis by (simp add: integral_restrict_Int) qed moreover obtain B where "\x. x \ range (\k. integral S (f k)) \ norm x \ B" using bou unfolding bounded_iff by blast ultimately show "bounded (range (\k. integral (cbox a b) (?f k)))" unfolding bounded_iff by (blast intro: order_trans) qed (use int le lim in auto) moreover have "\B>0. \a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) ?g - i) < e" if "0 < e" for e proof - have "e/4>0" using that by auto with LIMSEQ_D [OF i] obtain N where N: "\n. n \ N \ norm (integral S (f n) - i) < e/4" by metis with int_f[of N, unfolded has_integral_integral has_integral_alt'[of "f N"]] obtain B where "0 < B" and B: "\a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) (?f N) - integral S (f N)) < e/4" by (meson \0 < e/4\) have "norm (integral (cbox a b) ?g - i) < e" if ab: "ball 0 B \ cbox a b" for a b proof - obtain M where M: "\n. n \ M \ abs (integral (cbox a b) (?f n) - integral (cbox a b) ?g) < e/2" using \e > 0\ g by (fastforce simp add: dest!: LIMSEQ_D [where r = "e/2"]) have *: "\\ \ g. \\\ - i\ < e/2; \\ - g\ < e/2; \ \ \; \ \ i\ \ \g - i\ < e" unfolding real_inner_1_right by arith show "norm (integral (cbox a b) ?g - i) < e" unfolding real_norm_def proof (rule *) show "\integral (cbox a b) (?f N) - i\ < e/2" proof (rule abs_triangle_half_l) show "\integral (cbox a b) (?f N) - integral S (f N)\ < e/2/2" using B[OF ab] by simp show "abs (i - integral S (f N)) < e/2/2" using N by (simp add: abs_minus_commute) qed show "\integral (cbox a b) (?f (M + N)) - integral (cbox a b) ?g\ < e/2" by (metis le_add1 M[of "M + N"]) show "integral (cbox a b) (?f N) \ integral (cbox a b) (?f (M + N))" proof (intro ballI integral_le[OF int int]) fix x assume "x \ cbox a b" have "(f m x) \ (f n x)" if "x \ S" "n \ m" for m n proof (rule transitive_stepwise_le [OF \n \ m\ order_refl]) show "\u y z. \f u x \ f y x; f y x \ f z x\ \ f u x \ f z x" using dual_order.trans by blast qed (simp add: le \x \ S\) then show "(?f N)x \ (?f (M+N))x" by auto qed have "integral (cbox a b \ S) (f (M + N)) \ integral S (f (M + N))" by (metis Int_lower1 f0 inf_commute int' int_f integral_subset_le) then have "integral (cbox a b) (?f (M + N)) \ integral S (f (M + N))" by (metis (no_types) inf_commute integral_restrict_Int) also have "... \ i" using i'[of "M + N"] by auto finally show "integral (cbox a b) (?f (M + N)) \ i" . qed qed then show ?thesis using \0 < B\ by blast qed ultimately have "(g has_integral i) S" unfolding has_integral_alt' by auto then show ?thesis using has_integral_integrable_integral i integral_unique by metis qed have sub: "\k. integral S (\x. f k x - f 0 x) = integral S (f k) - integral S (f 0)" by (simp add: integral_diff int_f) have *: "\x m n. x \ S \ n\m \ f m x \ f n x" using assms(2) by (force intro: transitive_stepwise_le) have gf: "(\x. g x - f 0 x) integrable_on S \ ((\k. integral S (\x. f (Suc k) x - f 0 x)) \ integral S (\x. g x - f 0 x)) sequentially" proof (rule lem) show "\k. (\x. f (Suc k) x - f 0 x) integrable_on S" by (simp add: integrable_diff int_f) show "(\k. f (Suc k) x - f 0 x) \ g x - f 0 x" if "x \ S" for x proof - have "(\n. f (Suc n) x) \ g x" using LIMSEQ_ignore_initial_segment[OF fg[OF \x \ S\], of 1] by simp then show ?thesis by (simp add: tendsto_diff) qed show "bounded (range (\k. integral S (\x. f (Suc k) x - f 0 x)))" proof - obtain B where B: "\k. norm (integral S (f k)) \ B" using bou by (auto simp: bounded_iff) then have "norm (integral S (\x. f (Suc k) x - f 0 x)) \ B + norm (integral S (f 0))" for k unfolding sub by (meson add_le_cancel_right norm_triangle_le_diff) then show ?thesis unfolding bounded_iff by blast qed qed (use * in auto) then have "(\x. integral S (\xa. f (Suc x) xa - f 0 xa) + integral S (f 0)) \ integral S (\x. g x - f 0 x) + integral S (f 0)" by (auto simp add: tendsto_add) moreover have "(\x. g x - f 0 x + f 0 x) integrable_on S" using gf integrable_add int_f [of 0] by metis ultimately show ?thesis by (simp add: integral_diff int_f LIMSEQ_imp_Suc sub) qed lemma has_integral_monotone_convergence_increasing: fixes f :: "nat \ 'a::euclidean_space \ real" assumes f: "\k. (f k has_integral x k) s" assumes "\k x. x \ s \ f k x \ f (Suc k) x" assumes "\x. x \ s \ (\k. f k x) \ g x" assumes "x \ x'" shows "(g has_integral x') s" proof - have x_eq: "x = (\i. integral s (f i))" by (simp add: integral_unique[OF f]) then have x: "range(\k. integral s (f k)) = range x" by auto have *: "g integrable_on s \ (\k. integral s (f k)) \ integral s g" proof (intro monotone_convergence_increasing allI ballI assms) show "bounded (range(\k. integral s (f k)))" using x convergent_imp_bounded assms by metis qed (use f in auto) then have "integral s g = x'" by (intro LIMSEQ_unique[OF _ \x \ x'\]) (simp add: x_eq) with * show ?thesis by (simp add: has_integral_integral) qed lemma monotone_convergence_decreasing: fixes f :: "nat \ 'n::euclidean_space \ real" assumes intf: "\k. (f k) integrable_on S" and le: "\k x. x \ S \ f (Suc k) x \ f k x" and fg: "\x. x \ S \ ((\k. f k x) \ g x) sequentially" and bou: "bounded (range(\k. integral S (f k)))" shows "g integrable_on S \ (\k. integral S (f k)) \ integral S g" proof - have *: "range(\k. integral S (\x. - f k x)) = (*\<^sub>R) (- 1) ` (range(\k. integral S (f k)))" by force have "(\x. - g x) integrable_on S \ (\k. integral S (\x. - f k x)) \ integral S (\x. - g x)" proof (rule monotone_convergence_increasing) show "\k. (\x. - f k x) integrable_on S" by (blast intro: integrable_neg intf) show "\k x. x \ S \ - f k x \ - f (Suc k) x" by (simp add: le) show "\x. x \ S \ (\k. - f k x) \ - g x" by (simp add: fg tendsto_minus) show "bounded (range(\k. integral S (\x. - f k x)))" using "*" bou bounded_scaling by auto qed then show ?thesis by (force dest: integrable_neg tendsto_minus) qed lemma integral_norm_bound_integral: fixes f :: "'n::euclidean_space \ 'a::banach" assumes int_f: "f integrable_on S" and int_g: "g integrable_on S" and le_g: "\x. x \ S \ norm (f x) \ g x" shows "norm (integral S f) \ integral S g" proof - have norm: "norm \ \ y + e" if "norm \ \ x" and "\x - y\ < e/2" and "norm (\ - \) < e/2" for e x y and \ \ :: 'a proof - have "norm (\ - \) < e/2" by (metis norm_minus_commute that(3)) moreover have "x \ y + e/2" using that(2) by linarith ultimately show ?thesis using that(1) le_less_trans[OF norm_triangle_sub[of \ \]] by (auto simp: less_imp_le) qed have lem: "norm (integral(cbox a b) f) \ integral (cbox a b) g" if f: "f integrable_on cbox a b" and g: "g integrable_on cbox a b" and nle: "\x. x \ cbox a b \ norm (f x) \ g x" for f :: "'n \ 'a" and g a b proof (rule field_le_epsilon) fix e :: real assume "e > 0" then have e: "e/2 > 0" by auto with integrable_integral[OF f,unfolded has_integral[of f]] obtain \ where \: "gauge \" "\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x, k)\\. content k *\<^sub>R f x) - integral (cbox a b) f) < e/2" by meson moreover from integrable_integral[OF g,unfolded has_integral[of g]] e obtain \ where \: "gauge \" "\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x, k)\\. content k *\<^sub>R g x) - integral (cbox a b) g) < e/2" by meson ultimately have "gauge (\x. \ x \ \ x)" using gauge_Int by blast with fine_division_exists obtain \ where p: "\ tagged_division_of cbox a b" "(\x. \ x \ \ x) fine \" by metis have "\ fine \" "\ fine \" using fine_Int p(2) by blast+ show "norm (integral (cbox a b) f) \ integral (cbox a b) g + e" proof (rule norm) have "norm (content K *\<^sub>R f x) \ content K *\<^sub>R g x" if "(x, K) \ \" for x K proof- have K: "x \ K" "K \ cbox a b" using \(x, K) \ \\ p(1) by blast+ obtain u v where "K = cbox u v" using \(x, K) \ \\ p(1) by blast moreover have "content K * norm (f x) \ content K * g x" by (meson K(1) K(2) content_pos_le mult_left_mono nle subsetD) then show ?thesis by simp qed then show "norm (\(x, k)\\. content k *\<^sub>R f x) \ (\(x, k)\\. content k *\<^sub>R g x)" by (simp add: sum_norm_le split_def) show "norm ((\(x, k)\\. content k *\<^sub>R f x) - integral (cbox a b) f) < e/2" using \\ fine \\ \ p(1) by simp show "\(\(x, k)\\. content k *\<^sub>R g x) - integral (cbox a b) g\ < e/2" using \\ fine \\ \ p(1) by simp qed qed show ?thesis proof (rule field_le_epsilon) fix e :: real assume "e > 0" then have e: "e/2 > 0" by auto let ?f = "(\x. if x \ S then f x else 0)" let ?g = "(\x. if x \ S then g x else 0)" have f: "?f integrable_on cbox a b" and g: "?g integrable_on cbox a b" for a b using int_f int_g integrable_altD by auto obtain Bf where "0 < Bf" and Bf: "\a b. ball 0 Bf \ cbox a b \ \z. (?f has_integral z) (cbox a b) \ norm (z - integral S f) < e/2" using integrable_integral [OF int_f,unfolded has_integral'[of f]] e that by blast obtain Bg where "0 < Bg" and Bg: "\a b. ball 0 Bg \ cbox a b \ \z. (?g has_integral z) (cbox a b) \ norm (z - integral S g) < e/2" using integrable_integral [OF int_g,unfolded has_integral'[of g]] e that by blast obtain a b::'n where ab: "ball 0 Bf \ ball 0 Bg \ cbox a b" using ball_max_Un by (metis bounded_ball bounded_subset_cbox_symmetric) have "ball 0 Bf \ cbox a b" using ab by auto with Bf obtain z where int_fz: "(?f has_integral z) (cbox a b)" and z: "norm (z - integral S f) < e/2" by meson have "ball 0 Bg \ cbox a b" using ab by auto with Bg obtain w where int_gw: "(?g has_integral w) (cbox a b)" and w: "norm (w - integral S g) < e/2" by meson show "norm (integral S f) \ integral S g + e" proof (rule norm) show "norm (integral (cbox a b) ?f) \ integral (cbox a b) ?g" by (simp add: le_g lem[OF f g, of a b]) show "\integral (cbox a b) ?g - integral S g\ < e/2" using int_gw integral_unique w by auto show "norm (integral (cbox a b) ?f - integral S f) < e/2" using int_fz integral_unique z by blast qed qed qed lemma continuous_on_imp_absolutely_integrable_on: fixes f::"real \ 'a::banach" shows "continuous_on {a..b} f \ norm (integral {a..b} f) \ integral {a..b} (\x. norm (f x))" by (intro integral_norm_bound_integral integrable_continuous_real continuous_on_norm) auto lemma integral_bound: fixes f::"real \ 'a::banach" assumes "a \ b" assumes "continuous_on {a .. b} f" assumes "\t. t \ {a .. b} \ norm (f t) \ B" shows "norm (integral {a .. b} f) \ B * (b - a)" proof - note continuous_on_imp_absolutely_integrable_on[OF assms(2)] also have "integral {a..b} (\x. norm (f x)) \ integral {a..b} (\_. B)" by (rule integral_le) (auto intro!: integrable_continuous_real continuous_intros assms) also have "\ = B * (b - a)" using assms by simp finally show ?thesis . qed lemma integral_norm_bound_integral_component: fixes f :: "'n::euclidean_space \ 'a::banach" fixes g :: "'n \ 'b::euclidean_space" assumes f: "f integrable_on S" and g: "g integrable_on S" and fg: "\x. x \ S \ norm(f x) \ (g x)\k" shows "norm (integral S f) \ (integral S g)\k" proof - have "norm (integral S f) \ integral S ((\x. x \ k) \ g)" using integral_norm_bound_integral[OF f integrable_linear[OF g]] by (simp add: bounded_linear_inner_left fg) then show ?thesis unfolding o_def integral_component_eq[OF g] . qed lemma has_integral_norm_bound_integral_component: fixes f :: "'n::euclidean_space \ 'a::banach" fixes g :: "'n \ 'b::euclidean_space" assumes "(f has_integral i) S" and "(g has_integral j) S" and "\x. x \ S \ norm (f x) \ (g x)\k" shows "norm i \ j\k" by (metis assms has_integral_integrable integral_norm_bound_integral_component integral_unique) lemma uniformly_convergent_improper_integral: fixes f :: "'b \ real \ 'a :: {banach}" assumes deriv: "\x. x \ a \ (G has_field_derivative g x) (at x within {a..})" assumes integrable: "\a' b x. x \ A \ a' \ a \ b \ a' \ f x integrable_on {a'..b}" assumes G: "convergent G" assumes le: "\y x. y \ A \ x \ a \ norm (f y x) \ g x" shows "uniformly_convergent_on A (\b x. integral {a..b} (f x))" proof (intro Cauchy_uniformly_convergent uniformly_Cauchy_onI', goal_cases) case (1 \) from G have "Cauchy G" by (auto intro!: convergent_Cauchy) with 1 obtain M where M: "dist (G (real m)) (G (real n)) < \" if "m \ M" "n \ M" for m n by (force simp: Cauchy_def) define M' where "M' = max (nat \a\) M" show ?case proof (rule exI[of _ M'], safe, goal_cases) case (1 x m n) have M': "M' \ a" "M' \ M" unfolding M'_def by linarith+ have int_g: "(g has_integral (G (real n) - G (real m))) {real m..real n}" using 1 M' by (intro fundamental_theorem_of_calculus) (auto simp: has_real_derivative_iff_has_vector_derivative [symmetric] intro!: DERIV_subset[OF deriv]) have int_f: "f x integrable_on {a'..real n}" if "a' \ a" for a' using that 1 by (cases "a' \ real n") (auto intro: integrable) have "dist (integral {a..real m} (f x)) (integral {a..real n} (f x)) = norm (integral {a..real n} (f x) - integral {a..real m} (f x))" by (simp add: dist_norm norm_minus_commute) also have "integral {a..real m} (f x) + integral {real m..real n} (f x) = integral {a..real n} (f x)" using M' and 1 by (intro integral_combine int_f) auto hence "integral {a..real n} (f x) - integral {a..real m} (f x) = integral {real m..real n} (f x)" by (simp add: algebra_simps) also have "norm \ \ integral {real m..real n} g" using le 1 M' int_f int_g by (intro integral_norm_bound_integral) auto also from int_g have "integral {real m..real n} g = G (real n) - G (real m)" by (simp add: has_integral_iff) also have "\ \ dist (G m) (G n)" by (simp add: dist_norm) also from 1 and M' have "\ < \" by (intro M) auto finally show ?case . qed qed lemma uniformly_convergent_improper_integral': fixes f :: "'b \ real \ 'a :: {banach, real_normed_algebra}" assumes deriv: "\x. x \ a \ (G has_field_derivative g x) (at x within {a..})" assumes integrable: "\a' b x. x \ A \ a' \ a \ b \ a' \ f x integrable_on {a'..b}" assumes G: "convergent G" assumes le: "eventually (\x. \y\A. norm (f y x) \ g x) at_top" shows "uniformly_convergent_on A (\b x. integral {a..b} (f x))" proof - from le obtain a'' where le: "\y x. y \ A \ x \ a'' \ norm (f y x) \ g x" by (auto simp: eventually_at_top_linorder) define a' where "a' = max a a''" have "uniformly_convergent_on A (\b x. integral {a'..real b} (f x))" proof (rule uniformly_convergent_improper_integral) fix t assume t: "t \ a'" hence "(G has_field_derivative g t) (at t within {a..})" by (intro deriv) (auto simp: a'_def) moreover have "{a'..} \ {a..}" unfolding a'_def by auto ultimately show "(G has_field_derivative g t) (at t within {a'..})" by (rule DERIV_subset) qed (insert le, auto simp: a'_def intro: integrable G) hence "uniformly_convergent_on A (\b x. integral {a..a'} (f x) + integral {a'..real b} (f x))" (is ?P) by (intro uniformly_convergent_add) auto also have "eventually (\x. \y\A. integral {a..a'} (f y) + integral {a'..x} (f y) = integral {a..x} (f y)) sequentially" by (intro eventually_mono [OF eventually_ge_at_top[of "nat \a'\"]] ballI integral_combine) (auto simp: a'_def intro: integrable) hence "?P \ ?thesis" by (intro uniformly_convergent_cong) simp_all finally show ?thesis . qed subsection \differentiation under the integral sign\ lemma integral_continuous_on_param: fixes f::"'a::topological_space \ 'b::euclidean_space \ 'c::banach" assumes cont_fx: "continuous_on (U \ cbox a b) (\(x, t). f x t)" shows "continuous_on U (\x. integral (cbox a b) (f x))" proof cases assume "content (cbox a b) \ 0" then have ne: "cbox a b \ {}" by auto note [continuous_intros] = continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, unfolded split_beta fst_conv snd_conv] show ?thesis unfolding continuous_on_def proof (intro strip tendstoI) fix e'::real and x assume "e' > 0" define e where "e = e' / (content (cbox a b) + 1)" have "e > 0" using \e' > 0\ by (auto simp: e_def intro!: divide_pos_pos add_nonneg_pos) assume "x \ U" from continuous_on_prod_compactE[OF cont_fx compact_cbox \x \ U\ \0 < e\] obtain X0 where X0: "x \ X0" "open X0" and fx_bound: "\y t. y \ X0 \ U \ t \ cbox a b \ norm (f y t - f x t) \ e" unfolding split_beta fst_conv snd_conv dist_norm by metis have "\\<^sub>F y in at x within U. y \ X0 \ U" using X0(1) X0(2) eventually_at_topological by auto then show "\\<^sub>F y in at x within U. dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) < e'" proof eventually_elim case (elim y) have "dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) = norm (integral (cbox a b) (\t. f y t - f x t))" using elim \x \ U\ unfolding dist_norm by (subst integral_diff) (auto intro!: integrable_continuous continuous_intros) also have "\ \ e * content (cbox a b)" using elim \x \ U\ by (intro integrable_bound) (auto intro!: fx_bound \x \ U \ less_imp_le[OF \0 < e\] integrable_continuous continuous_intros) also have "\ < e'" using \0 < e'\ \e > 0\ by (auto simp: e_def field_split_simps) finally show "dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) < e'" . qed qed qed (auto intro!: continuous_on_const) lemma leibniz_rule: fixes f::"'a::banach \ 'b::euclidean_space \ 'c::banach" assumes fx: "\x t. x \ U \ t \ cbox a b \ ((\x. f x t) has_derivative blinfun_apply (fx x t)) (at x within U)" assumes integrable_f2: "\x. x \ U \ f x integrable_on cbox a b" assumes cont_fx: "continuous_on (U \ (cbox a b)) (\(x, t). fx x t)" assumes [intro]: "x0 \ U" assumes "convex U" shows "((\x. integral (cbox a b) (f x)) has_derivative integral (cbox a b) (fx x0)) (at x0 within U)" (is "(?F has_derivative ?dF) _") proof cases assume "content (cbox a b) \ 0" then have ne: "cbox a b \ {}" by auto note [continuous_intros] = continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, unfolded split_beta fst_conv snd_conv] show ?thesis proof (intro has_derivativeI bounded_linear_scaleR_left tendstoI, fold norm_conv_dist) have cont_f1: "\t. t \ cbox a b \ continuous_on U (\x. f x t)" by (auto simp: continuous_on_eq_continuous_within intro!: has_derivative_continuous fx) note [continuous_intros] = continuous_on_compose2[OF cont_f1] fix e'::real assume "e' > 0" define e where "e = e' / (content (cbox a b) + 1)" have "e > 0" using \e' > 0\ by (auto simp: e_def intro!: divide_pos_pos add_nonneg_pos) from continuous_on_prod_compactE[OF cont_fx compact_cbox \x0 \ U\ \e > 0\] obtain X0 where X0: "x0 \ X0" "open X0" and fx_bound: "\x t. x \ X0 \ U \ t \ cbox a b \ norm (fx x t - fx x0 t) \ e" unfolding split_beta fst_conv snd_conv by (metis dist_norm) note eventually_closed_segment[OF \open X0\ \x0 \ X0\, of U] moreover have "\\<^sub>F x in at x0 within U. x \ X0" using \open X0\ \x0 \ X0\ eventually_at_topological by blast moreover have "\\<^sub>F x in at x0 within U. x \ x0" by (auto simp: eventually_at_filter) moreover have "\\<^sub>F x in at x0 within U. x \ U" by (auto simp: eventually_at_filter) ultimately show "\\<^sub>F x in at x0 within U. norm ((?F x - ?F x0 - ?dF (x - x0)) /\<^sub>R norm (x - x0)) < e'" proof eventually_elim case (elim x) from elim have "0 < norm (x - x0)" by simp have "closed_segment x0 x \ U" by (simp add: assms closed_segment_subset elim(4)) from elim have [intro]: "x \ U" by auto have "?F x - ?F x0 - ?dF (x - x0) = integral (cbox a b) (\y. f x y - f x0 y - fx x0 y (x - x0))" (is "_ = ?id") using \x \ x0\ by (subst blinfun_apply_integral integral_diff, auto intro!: integrable_diff integrable_f2 continuous_intros intro: integrable_continuous)+ also { fix t assume t: "t \ (cbox a b)" then have deriv: "((\x. f x t) has_derivative (fx y t)) (at y within X0 \ U)" if "y \ X0 \ U" for y using fx has_derivative_subset that by fastforce have seg: "\t. t \ {0..1} \ x0 + t *\<^sub>R (x - x0) \ X0 \ U" using \closed_segment x0 x \ U\ \closed_segment x0 x \ X0\ by (force simp: closed_segment_def algebra_simps) have "\x. x \ X0 \ U \ onorm (blinfun_apply (fx x t) - (fx x0 t)) \ e" using fx_bound t by (auto simp add: norm_blinfun_def fun_diff_def blinfun.bilinear_simps[symmetric]) from differentiable_bound_linearization[OF seg deriv this] X0 have "norm (f x t - f x0 t - fx x0 t (x - x0)) \ e * norm (x - x0)" by (auto simp add: ac_simps) } then have "norm ?id \ integral (cbox a b) (\_. e * norm (x - x0))" by (intro integral_norm_bound_integral) (auto intro!: continuous_intros integrable_diff integrable_f2 intro: integrable_continuous) also have "\ = content (cbox a b) * e * norm (x - x0)" by simp also have "\ < e' * norm (x - x0)" proof (intro mult_strict_right_mono[OF _ \0 < norm (x - x0)\]) show "content (cbox a b) * e < e'" using \e' > 0\ by (simp add: divide_simps e_def not_less) qed finally have "norm (?F x - ?F x0 - ?dF (x - x0)) < e' * norm (x - x0)" . then show ?case by (auto simp: divide_simps) qed qed (rule blinfun.bounded_linear_right) qed (auto intro!: derivative_eq_intros simp: blinfun.bilinear_simps) lemma has_vector_derivative_eq_has_derivative_blinfun: "(f has_vector_derivative f') (at x within U) \ (f has_derivative blinfun_scaleR_left f') (at x within U)" by (simp add: has_vector_derivative_def) lemma leibniz_rule_vector_derivative: fixes f::"real \ 'b::euclidean_space \ 'c::banach" assumes fx: "\x t. x \ U \ t \ cbox a b \ ((\x. f x t) has_vector_derivative (fx x t)) (at x within U)" assumes integrable_f2: "\x. x \ U \ (f x) integrable_on cbox a b" assumes cont_fx: "continuous_on (U \ cbox a b) (\(x, t). fx x t)" assumes U: "x0 \ U" "convex U" shows "((\x. integral (cbox a b) (f x)) has_vector_derivative integral (cbox a b) (fx x0)) (at x0 within U)" proof - note [continuous_intros] = continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, unfolded split_beta fst_conv snd_conv] show ?thesis unfolding has_vector_derivative_eq_has_derivative_blinfun proof (rule has_derivative_eq_rhs [OF leibniz_rule[OF _ integrable_f2 _ U]]) show "continuous_on (U \ cbox a b) (\(x, t). blinfun_scaleR_left (fx x t))" using cont_fx by (auto simp: split_beta intro!: continuous_intros) show "blinfun_apply (integral (cbox a b) (\t. blinfun_scaleR_left (fx x0 t))) = blinfun_apply (blinfun_scaleR_left (integral (cbox a b) (fx x0)))" by (subst integral_linear[symmetric]) (auto simp: has_vector_derivative_def o_def intro!: integrable_continuous U continuous_intros bounded_linear_intros) qed (use fx in \auto simp: has_vector_derivative_def\) qed lemma has_field_derivative_eq_has_derivative_blinfun: "(f has_field_derivative f') (at x within U) \ (f has_derivative blinfun_mult_right f') (at x within U)" by (simp add: has_field_derivative_def) lemma leibniz_rule_field_derivative: fixes f::"'a::{real_normed_field, banach} \ 'b::euclidean_space \ 'a" assumes fx: "\x t. x \ U \ t \ cbox a b \ ((\x. f x t) has_field_derivative fx x t) (at x within U)" assumes integrable_f2: "\x. x \ U \ (f x) integrable_on cbox a b" assumes cont_fx: "continuous_on (U \ (cbox a b)) (\(x, t). fx x t)" assumes U: "x0 \ U" "convex U" shows "((\x. integral (cbox a b) (f x)) has_field_derivative integral (cbox a b) (fx x0)) (at x0 within U)" proof - note [continuous_intros] = continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, unfolded split_beta fst_conv snd_conv] have *: "blinfun_mult_right (integral (cbox a b) (fx x0)) = integral (cbox a b) (\t. blinfun_mult_right (fx x0 t))" by (subst integral_linear[symmetric]) (auto simp: has_vector_derivative_def o_def intro!: integrable_continuous U continuous_intros bounded_linear_intros) show ?thesis unfolding has_field_derivative_eq_has_derivative_blinfun proof (rule has_derivative_eq_rhs [OF leibniz_rule[OF _ integrable_f2 _ U, where fx="\x t. blinfun_mult_right (fx x t)"]]) show "continuous_on (U \ cbox a b) (\(x, t). blinfun_mult_right (fx x t))" using cont_fx by (auto simp: split_beta intro!: continuous_intros) show "blinfun_apply (integral (cbox a b) (\t. blinfun_mult_right (fx x0 t))) = blinfun_apply (blinfun_mult_right (integral (cbox a b) (fx x0)))" by (subst integral_linear[symmetric]) (auto simp: has_vector_derivative_def o_def intro!: integrable_continuous U continuous_intros bounded_linear_intros) qed (use fx in \auto simp: has_field_derivative_def\) qed lemma leibniz_rule_field_differentiable: fixes f::"'a::{real_normed_field, banach} \ 'b::euclidean_space \ 'a" assumes "\x t. x \ U \ t \ cbox a b \ ((\x. f x t) has_field_derivative fx x t) (at x within U)" assumes "\x. x \ U \ (f x) integrable_on cbox a b" assumes "continuous_on (U \ (cbox a b)) (\(x, t). fx x t)" assumes "x0 \ U" "convex U" shows "(\x. integral (cbox a b) (f x)) field_differentiable at x0 within U" using leibniz_rule_field_derivative[OF assms] by (auto simp: field_differentiable_def) subsection \Exchange uniform limit and integral\ lemma uniform_limit_integral_cbox: fixes f::"'a \ 'b::euclidean_space \ 'c::banach" assumes u: "uniform_limit (cbox a b) f g F" assumes c: "\n. continuous_on (cbox a b) (f n)" assumes [simp]: "F \ bot" obtains I J where "\n. (f n has_integral I n) (cbox a b)" "(g has_integral J) (cbox a b)" "(I \ J) F" proof - have fi[simp]: "f n integrable_on (cbox a b)" for n by (auto intro!: integrable_continuous assms) then obtain I where I: "\n. (f n has_integral I n) (cbox a b)" unfolding integrable_on_def by metis moreover have gi[simp]: "g integrable_on (cbox a b)" by (auto intro!: integrable_continuous uniform_limit_theorem[OF _ u] eventuallyI c) then obtain J where J: "(g has_integral J) (cbox a b)" by blast moreover have "(I \ J) F" proof cases assume "content (cbox a b) = 0" hence "I = (\_. 0)" "J = 0" by (auto intro!: has_integral_unique I J) thus ?thesis by simp next assume content_nonzero: "content (cbox a b) \ 0" show ?thesis proof (rule tendstoI) fix e::real assume "e > 0" define e' where "e' = e/2" with \e > 0\ have "e' > 0" by simp then have "\\<^sub>F n in F. \x\cbox a b. norm (f n x - g x) < e' / content (cbox a b)" using u content_nonzero by (auto simp: uniform_limit_iff dist_norm zero_less_measure_iff) then show "\\<^sub>F n in F. dist (I n) J < e" proof eventually_elim case (elim n) have "I n = integral (cbox a b) (f n)" "J = integral (cbox a b) g" using I[of n] J by (simp_all add: integral_unique) then have "dist (I n) J = norm (integral (cbox a b) (\x. f n x - g x))" by (simp add: integral_diff dist_norm) also have "\ \ integral (cbox a b) (\x. (e' / content (cbox a b)))" using elim by (intro integral_norm_bound_integral) (auto intro!: integrable_diff) also have "\ < e" using \0 < e\ by (simp add: e'_def) finally show ?case . qed qed qed ultimately show ?thesis .. qed lemma uniform_limit_integral: fixes f::"'a \ 'b::ordered_euclidean_space \ 'c::banach" assumes u: "uniform_limit {a..b} f g F" assumes c: "\n. continuous_on {a..b} (f n)" assumes [simp]: "F \ bot" obtains I J where "\n. (f n has_integral I n) {a..b}" "(g has_integral J) {a..b}" "(I \ J) F" by (metis interval_cbox assms uniform_limit_integral_cbox) subsection \Integration by parts\ lemma integration_by_parts_interior_strong: fixes prod :: "_ \ _ \ 'b :: banach" assumes bilinear: "bounded_bilinear (prod)" assumes s: "finite s" and le: "a \ b" assumes cont [continuous_intros]: "continuous_on {a..b} f" "continuous_on {a..b} g" assumes deriv: "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" assumes int: "((\x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}" shows "((\x. prod (f' x) (g x)) has_integral y) {a..b}" proof - interpret bounded_bilinear prod by fact have "((\x. prod (f x) (g' x) + prod (f' x) (g x)) has_integral (prod (f b) (g b) - prod (f a) (g a))) {a..b}" using deriv by (intro fundamental_theorem_of_calculus_interior_strong[OF s le]) (auto intro!: continuous_intros continuous_on has_vector_derivative) from has_integral_diff[OF this int] show ?thesis by (simp add: algebra_simps) qed lemma integration_by_parts_interior: fixes prod :: "_ \ _ \ 'b :: banach" assumes "bounded_bilinear (prod)" "a \ b" "continuous_on {a..b} f" "continuous_on {a..b} g" assumes "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" assumes "((\x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}" shows "((\x. prod (f' x) (g x)) has_integral y) {a..b}" by (rule integration_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (use assms in simp_all) lemma integration_by_parts: fixes prod :: "_ \ _ \ 'b :: banach" assumes "bounded_bilinear (prod)" "a \ b" "continuous_on {a..b} f" "continuous_on {a..b} g" assumes "\x. x\{a..b} \ (f has_vector_derivative f' x) (at x)" "\x. x\{a..b} \ (g has_vector_derivative g' x) (at x)" assumes "((\x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}" shows "((\x. prod (f' x) (g x)) has_integral y) {a..b}" by (rule integration_by_parts_interior[of _ _ _ f g f' g']) (use assms in simp_all) lemma integrable_by_parts_interior_strong: fixes prod :: "_ \ _ \ 'b :: banach" assumes bilinear: "bounded_bilinear (prod)" assumes s: "finite s" and le: "a \ b" assumes cont [continuous_intros]: "continuous_on {a..b} f" "continuous_on {a..b} g" assumes deriv: "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" assumes int: "(\x. prod (f x) (g' x)) integrable_on {a..b}" shows "(\x. prod (f' x) (g x)) integrable_on {a..b}" proof - from int obtain I where "((\x. prod (f x) (g' x)) has_integral I) {a..b}" unfolding integrable_on_def by blast hence "((\x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - (prod (f b) (g b) - prod (f a) (g a) - I))) {a..b}" by simp from integration_by_parts_interior_strong[OF assms(1-7) this] show ?thesis unfolding integrable_on_def by blast qed lemma integrable_by_parts_interior: fixes prod :: "_ \ _ \ 'b :: banach" assumes "bounded_bilinear (prod)" "a \ b" "continuous_on {a..b} f" "continuous_on {a..b} g" assumes "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" assumes "(\x. prod (f x) (g' x)) integrable_on {a..b}" shows "(\x. prod (f' x) (g x)) integrable_on {a..b}" by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (use assms in simp_all) lemma integrable_by_parts: fixes prod :: "_ \ _ \ 'b :: banach" assumes "bounded_bilinear (prod)" "a \ b" "continuous_on {a..b} f" "continuous_on {a..b} g" assumes "\x. x\{a..b} \ (f has_vector_derivative f' x) (at x)" "\x. x\{a..b} \ (g has_vector_derivative g' x) (at x)" assumes "(\x. prod (f x) (g' x)) integrable_on {a..b}" shows "(\x. prod (f' x) (g x)) integrable_on {a..b}" by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (use assms in simp_all) subsection \Integration by substitution\ lemma has_integral_substitution_general: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ real" assumes s: "finite s" and le: "a \ b" and subset: "g ` {a..b} \ {c..d}" and f [continuous_intros]: "continuous_on {c..d} f" and g [continuous_intros]: "continuous_on {a..b} g" and deriv [derivative_intros]: "\x. x \ {a..b} - s \ (g has_field_derivative g' x) (at x within {a..b})" shows "((\x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f - integral {g b..g a} f)) {a..b}" proof - let ?F = "\x. integral {c..g x} f" have cont_int: "continuous_on {a..b} ?F" by (rule continuous_on_compose2[OF _ g subset] indefinite_integral_continuous_1 f integrable_continuous_real)+ have deriv: "(((\x. integral {c..x} f) \ g) has_vector_derivative g' x *\<^sub>R f (g x)) (at x within {a..b})" if "x \ {a..b} - s" for x proof (rule has_vector_derivative_eq_rhs [OF vector_diff_chain_within refl]) show "(g has_vector_derivative g' x) (at x within {a..b})" using deriv has_real_derivative_iff_has_vector_derivative that by blast show "((\x. integral {c..x} f) has_vector_derivative f (g x)) (at (g x) within g ` {a..b})" using that le subset by (blast intro: has_vector_derivative_within_subset integral_has_vector_derivative f) qed have deriv: "(?F has_vector_derivative g' x *\<^sub>R f (g x)) (at x)" if "x \ {a..b} - (s \ {a,b})" for x using deriv[of x] that by (simp add: at_within_Icc_at o_def) have "((\x. g' x *\<^sub>R f (g x)) has_integral (?F b - ?F a)) {a..b}" using le cont_int s deriv cont_int by (intro fundamental_theorem_of_calculus_interior_strong[of "s \ {a,b}"]) simp_all also from subset have "g x \ {c..d}" if "x \ {a..b}" for x using that by blast from this[of a] this[of b] le have cd: "c \ g a" "g b \ d" "c \ g b" "g a \ d" by auto have "integral {c..g b} f - integral {c..g a} f = integral {g a..g b} f - integral {g b..g a} f" proof cases assume "g a \ g b" note le = le this from cd have "integral {c..g a} f + integral {g a..g b} f = integral {c..g b} f" by (intro integral_combine integrable_continuous_real continuous_on_subset[OF f] le) simp_all with le show ?thesis by (cases "g a = g b") (simp_all add: algebra_simps) next assume less: "\g a \ g b" then have "g a \ g b" by simp note le = le this from cd have "integral {c..g b} f + integral {g b..g a} f = integral {c..g a} f" by (intro integral_combine integrable_continuous_real continuous_on_subset[OF f] le) simp_all with less show ?thesis by (simp_all add: algebra_simps) qed finally show ?thesis . qed lemma has_integral_substitution_strong: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ real" assumes s: "finite s" and le: "a \ b" "g a \ g b" and subset: "g ` {a..b} \ {c..d}" and f [continuous_intros]: "continuous_on {c..d} f" and g [continuous_intros]: "continuous_on {a..b} g" and deriv [derivative_intros]: "\x. x \ {a..b} - s \ (g has_field_derivative g' x) (at x within {a..b})" shows "((\x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f)) {a..b}" using has_integral_substitution_general[OF s le(1) subset f g deriv] le(2) by (cases "g a = g b") auto lemma has_integral_substitution: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ real" assumes "a \ b" "g a \ g b" "g ` {a..b} \ {c..d}" and "continuous_on {c..d} f" and "\x. x \ {a..b} \ (g has_field_derivative g' x) (at x within {a..b})" shows "((\x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f)) {a..b}" by (intro has_integral_substitution_strong[of "{}" a b g c d] assms) (auto intro: DERIV_continuous_on assms) lemma integral_shift: fixes f :: "real \ 'a::euclidean_space" assumes cont: "continuous_on {a + c..b + c} f" shows "integral {a..b} (f \ (\x. x + c)) = integral {a + c..b + c} f" proof (cases "a \ b") case True have "((\x. 1 *\<^sub>R f (x + c)) has_integral integral {a+c..b+c} f) {a..b}" using True cont by (intro has_integral_substitution[where c = "a + c" and d = "b + c"]) (auto intro!: derivative_eq_intros) thus ?thesis by (simp add: has_integral_iff o_def) qed auto subsection \Compute a double integral using iterated integrals and switching the order of integration\ lemma continuous_on_imp_integrable_on_Pair1: fixes f :: "_ \ 'b::banach" assumes con: "continuous_on (cbox (a,c) (b,d)) f" and x: "x \ cbox a b" shows "(\y. f (x, y)) integrable_on (cbox c d)" proof - have "f \ (\y. (x, y)) integrable_on (cbox c d)" proof (intro integrable_continuous continuous_on_compose [OF _ continuous_on_subset [OF con]]) show "continuous_on (cbox c d) (Pair x)" by (simp add: continuous_on_Pair) show "Pair x ` cbox c d \ cbox (a,c) (b,d)" using x by blast qed then show ?thesis by (simp add: o_def) qed lemma integral_integrable_2dim: fixes f :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::banach" assumes "continuous_on (cbox (a,c) (b,d)) f" shows "(\x. integral (cbox c d) (\y. f (x,y))) integrable_on cbox a b" proof (cases "content(cbox c d) = 0") case True then show ?thesis by (simp add: True integrable_const) next case False have uc: "uniformly_continuous_on (cbox (a,c) (b,d)) f" by (simp add: assms compact_cbox compact_uniformly_continuous) { fix x::'a and e::real assume x: "x \ cbox a b" and e: "0 < e" then have e2_gt: "0 < e/2 / content (cbox c d)" and e2_less: "e/2 / content (cbox c d) * content (cbox c d) < e" by (auto simp: False content_lt_nz e) then obtain dd where dd: "\x x'. \x\cbox (a, c) (b, d); x'\cbox (a, c) (b, d); norm (x' - x) < dd\ \ norm (f x' - f x) \ e/(2 * content (cbox c d))" "dd>0" using uc [unfolded uniformly_continuous_on_def, THEN spec, of "e/(2 * content (cbox c d))"] by (auto simp: dist_norm intro: less_imp_le) have "\delta>0. \x'\cbox a b. norm (x' - x) < delta \ norm (integral (cbox c d) (\u. f (x', u) - f (x, u))) < e" using dd e2_gt assms x apply (rule_tac x=dd in exI) apply clarify apply (rule le_less_trans [OF integrable_bound e2_less]) apply (auto intro: integrable_diff continuous_on_imp_integrable_on_Pair1) done } note * = this show ?thesis proof (rule integrable_continuous) show "continuous_on (cbox a b) (\x. integral (cbox c d) (\y. f (x, y)))" by (simp add: * continuous_on_iff dist_norm integral_diff [symmetric] continuous_on_imp_integrable_on_Pair1 [OF assms]) qed qed lemma integral_split: fixes f :: "'a::euclidean_space \ 'b::{real_normed_vector,complete_space}" assumes f: "f integrable_on (cbox a b)" and k: "k \ Basis" shows "integral (cbox a b) f = integral (cbox a b \ {x. x\k \ c}) f + integral (cbox a b \ {x. x\k \ c}) f" using k f by (auto simp: has_integral_integral intro: integral_unique [OF has_integral_split]) lemma integral_swap_operativeI: fixes f :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::banach" assumes f: "continuous_on s f" and e: "0 < e" shows "operative conj True (\k. \a b c d. cbox (a,c) (b,d) \ k \ cbox (a,c) (b,d) \ s \ norm(integral (cbox (a,c) (b,d)) f - integral (cbox a b) (\x. integral (cbox c d) (\y. f((x,y))))) \ e * content (cbox (a,c) (b,d)))" proof (standard, auto) fix a::'a and c::'b and b::'a and d::'b and u::'a and v::'a and w::'b and z::'b assume *: "box (a, c) (b, d) = {}" and cb1: "cbox (u, w) (v, z) \ cbox (a, c) (b, d)" and cb2: "cbox (u, w) (v, z) \ s" then have c0: "content (cbox (a, c) (b, d)) = 0" using * unfolding content_eq_0_interior by simp have c0': "content (cbox (u, w) (v, z)) = 0" by (fact content_0_subset [OF c0 cb1]) show "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y)))) \ e * content (cbox (u,w) (v,z))" using content_cbox_pair_eq0_D [OF c0'] by (force simp add: c0') next fix a::'a and c::'b and b::'a and d::'b and M::real and i::'a and j::'b and u::'a and v::'a and w::'b and z::'b assume ij: "(i,j) \ Basis" and n1: "\a' b' c' d'. cbox (a',c') (b',d') \ cbox (a,c) (b,d) \ cbox (a',c') (b',d') \ {x. x \ (i,j) \ M} \ cbox (a',c') (b',d') \ s \ norm (integral (cbox (a',c') (b',d')) f - integral (cbox a' b') (\x. integral (cbox c' d') (\y. f (x,y)))) \ e * content (cbox (a',c') (b',d'))" and n2: "\a' b' c' d'. cbox (a',c') (b',d') \ cbox (a,c) (b,d) \ cbox (a',c') (b',d') \ {x. M \ x \ (i,j)} \ cbox (a',c') (b',d') \ s \ norm (integral (cbox (a',c') (b',d')) f - integral (cbox a' b') (\x. integral (cbox c' d') (\y. f (x,y)))) \ e * content (cbox (a',c') (b',d'))" and subs: "cbox (u,w) (v,z) \ cbox (a,c) (b,d)" "cbox (u,w) (v,z) \ s" have fcont: "continuous_on (cbox (u, w) (v, z)) f" using assms(1) continuous_on_subset subs(2) by blast then have fint: "f integrable_on cbox (u, w) (v, z)" by (metis integrable_continuous) consider "i \ Basis" "j=0" | "j \ Basis" "i=0" using ij by (auto simp: Euclidean_Space.Basis_prod_def) then show "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x,y)))) \ e * content (cbox (u,w) (v,z))" (is ?normle) proof cases case 1 then have e: "e * content (cbox (u, w) (v, z)) = e * (content (cbox u v \ {x. x \ i \ M}) * content (cbox w z)) + e * (content (cbox u v \ {x. M \ x \ i}) * content (cbox w z))" by (simp add: content_split [where c=M] content_Pair algebra_simps) have *: "integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y))) = integral (cbox u v \ {x. x \ i \ M}) (\x. integral (cbox w z) (\y. f (x, y))) + integral (cbox u v \ {x. M \ x \ i}) (\x. integral (cbox w z) (\y. f (x, y)))" using 1 f subs integral_integrable_2dim continuous_on_subset by (blast intro: integral_split) show ?normle apply (rule norm_diff2 [OF integral_split [where c=M, OF fint ij] * e]) using 1 subs apply (simp_all add: cbox_Pair_eq setcomp_dot1 [of "\u. M\u"] setcomp_dot1 [of "\u. u\M"] Sigma_Int_Paircomp1) apply (simp_all add: interval_split ij flip: cbox_Pair_eq content_Pair) apply (force simp flip: interval_split intro!: n1 [rule_format]) apply (force simp flip: interval_split intro!: n2 [rule_format]) done next case 2 then have e: "e * content (cbox (u, w) (v, z)) = e * (content (cbox u v) * content (cbox w z \ {x. x \ j \ M})) + e * (content (cbox u v) * content (cbox w z \ {x. M \ x \ j}))" by (simp add: content_split [where c=M] content_Pair algebra_simps) have "(\x. integral (cbox w z \ {x. x \ j \ M}) (\y. f (x, y))) integrable_on cbox u v" "(\x. integral (cbox w z \ {x. M \ x \ j}) (\y. f (x, y))) integrable_on cbox u v" using 2 subs apply (simp_all add: interval_split) apply (rule integral_integrable_2dim [OF continuous_on_subset [OF f]]; auto simp: cbox_Pair_eq interval_split [symmetric])+ done with 2 have *: "integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y))) = integral (cbox u v) (\x. integral (cbox w z \ {x. x \ j \ M}) (\y. f (x, y))) + integral (cbox u v) (\x. integral (cbox w z \ {x. M \ x \ j}) (\y. f (x, y)))" by (simp add: integral_add [symmetric] integral_split [symmetric] continuous_on_imp_integrable_on_Pair1 [OF fcont] cong: integral_cong) show ?normle apply (rule norm_diff2 [OF integral_split [where c=M, OF fint ij] * e]) using 2 subs apply (simp_all add: cbox_Pair_eq setcomp_dot2 [of "\u. M\u"] setcomp_dot2 [of "\u. u\M"] Sigma_Int_Paircomp2) apply (simp_all add: interval_split ij flip: cbox_Pair_eq content_Pair) apply (force simp flip: interval_split intro!: n1 [rule_format]) apply (force simp flip: interval_split intro!: n2 [rule_format]) done qed qed lemma integral_Pair_const: "integral (cbox (a,c) (b,d)) (\x. k) = integral (cbox a b) (\x. integral (cbox c d) (\y. k))" by (simp add: content_Pair) lemma integral_prod_continuous: fixes f :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::banach" assumes "continuous_on (cbox (a, c) (b, d)) f" (is "continuous_on ?CBOX f") shows "integral (cbox (a, c) (b, d)) f = integral (cbox a b) (\x. integral (cbox c d) (\y. f (x, y)))" proof (cases "content ?CBOX = 0") case True then show ?thesis by (auto simp: content_Pair) next case False then have cbp: "content ?CBOX > 0" using content_lt_nz by blast have "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x,y)))) = 0" proof (rule dense_eq0_I, simp) fix e :: real assume "0 < e" with \content ?CBOX > 0\ have "0 < e/content ?CBOX" by simp have f_int_acbd: "f integrable_on ?CBOX" by (rule integrable_continuous [OF assms]) { fix p assume p: "p division_of ?CBOX" then have "finite p" by blast define e' where "e' = e/content ?CBOX" with \0 < e\ \0 < e/content ?CBOX\ have "0 < e'" by simp interpret operative conj True "\k. \a' b' c' d'. cbox (a', c') (b', d') \ k \ cbox (a', c') (b', d') \ ?CBOX \ norm (integral (cbox (a', c') (b', d')) f - integral (cbox a' b') (\x. integral (cbox c' d') (\y. f ((x, y))))) \ e' * content (cbox (a', c') (b', d'))" using assms \0 < e'\ by (rule integral_swap_operativeI) have "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x, y)))) \ e' * content ?CBOX" if "\t u v w z. t \ p \ cbox (u, w) (v, z) \ t \ cbox (u, w) (v, z) \ ?CBOX \ norm (integral (cbox (u, w) (v, z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y)))) \ e' * content (cbox (u, w) (v, z))" using that division [of p "(a, c)" "(b, d)"] p \finite p\ by (auto simp add: comm_monoid_set_F_and) with False have "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x, y)))) \ e" if "\t u v w z. t \ p \ cbox (u, w) (v, z) \ t \ cbox (u, w) (v, z) \ ?CBOX \ norm (integral (cbox (u, w) (v, z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y)))) \ e * content (cbox (u, w) (v, z)) / content ?CBOX" using that by (simp add: e'_def) } note op_acbd = this { fix k::real and \ and u::'a and v w and z::'b and t1 t2 l assume k: "0 < k" and nf: "\x y u v. \x \ cbox a b; y \ cbox c d; u \ cbox a b; v\cbox c d; norm (u-x, v-y) < k\ \ norm (f(u,v) - f(x,y)) < e/(2 * (content ?CBOX))" and p_acbd: "\ tagged_division_of cbox (a,c) (b,d)" and fine: "(\x. ball x k) fine \" "((t1,t2), l) \ \" and uwvz_sub: "cbox (u,w) (v,z) \ l" "cbox (u,w) (v,z) \ cbox (a,c) (b,d)" have t: "t1 \ cbox a b" "t2 \ cbox c d" by (meson fine p_acbd cbox_Pair_iff tag_in_interval)+ have ls: "l \ ball (t1,t2) k" using fine by (simp add: fine_def Ball_def) { fix x1 x2 assume xuvwz: "x1 \ cbox u v" "x2 \ cbox w z" then have x: "x1 \ cbox a b" "x2 \ cbox c d" using uwvz_sub by auto have "norm (x1 - t1, x2 - t2) = norm (t1 - x1, t2 - x2)" by (simp add: norm_Pair norm_minus_commute) also have "norm (t1 - x1, t2 - x2) < k" using xuvwz ls uwvz_sub unfolding ball_def by (force simp add: cbox_Pair_eq dist_norm ) finally have "norm (f (x1,x2) - f (t1,t2)) \ e/content ?CBOX/2" using nf [OF t x] by simp } note nf' = this have f_int_uwvz: "f integrable_on cbox (u,w) (v,z)" using f_int_acbd uwvz_sub integrable_on_subcbox by blast have f_int_uv: "\x. x \ cbox u v \ (\y. f (x,y)) integrable_on cbox w z" using assms continuous_on_subset uwvz_sub by (blast intro: continuous_on_imp_integrable_on_Pair1) have 1: "norm (integral (cbox (u,w) (v,z)) f - integral (cbox (u,w) (v,z)) (\x. f (t1,t2))) \ e * content (cbox (u,w) (v,z)) / content ?CBOX/2" using cbp \0 < e/content ?CBOX\ nf' apply (simp only: integral_diff [symmetric] f_int_uwvz integrable_const) apply (auto simp: integrable_diff f_int_uwvz integrable_const intro: order_trans [OF integrable_bound [of "e/content ?CBOX/2"]]) done have int_integrable: "(\x. integral (cbox w z) (\y. f (x, y))) integrable_on cbox u v" using assms integral_integrable_2dim continuous_on_subset uwvz_sub(2) by blast have normint_wz: "\x. x \ cbox u v \ norm (integral (cbox w z) (\y. f (x, y)) - integral (cbox w z) (\y. f (t1, t2))) \ e * content (cbox w z) / content (cbox (a, c) (b, d))/2" using cbp \0 < e/content ?CBOX\ nf' apply (simp only: integral_diff [symmetric] f_int_uv integrable_const) apply (auto simp: integrable_diff f_int_uv integrable_const intro: order_trans [OF integrable_bound [of "e/content ?CBOX/2"]]) done have "norm (integral (cbox u v) (\x. integral (cbox w z) (\y. f (x,y)) - integral (cbox w z) (\y. f (t1,t2)))) \ e * content (cbox w z) / content ?CBOX/2 * content (cbox u v)" using cbp \0 < e/content ?CBOX\ apply (intro integrable_bound [OF _ _ normint_wz]) apply (auto simp: field_split_simps integrable_diff int_integrable integrable_const) done also have "... \ e * content (cbox (u,w) (v,z)) / content ?CBOX/2" by (simp add: content_Pair field_split_simps) finally have 2: "norm (integral (cbox u v) (\x. integral (cbox w z) (\y. f (x,y))) - integral (cbox u v) (\x. integral (cbox w z) (\y. f (t1,t2)))) \ e * content (cbox (u,w) (v,z)) / content ?CBOX/2" by (simp only: integral_diff [symmetric] int_integrable integrable_const) have norm_xx: "\x' = y'; norm(x - x') \ e/2; norm(y - y') \ e/2\ \ norm(x - y) \ e" for x::'c and y x' y' e using norm_triangle_mono [of "x-y'" "e/2" "y'-y" "e/2"] field_sum_of_halves by (simp add: norm_minus_commute) have "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x,y)))) \ e * content (cbox (u,w) (v,z)) / content ?CBOX" by (rule norm_xx [OF integral_Pair_const 1 2]) } note * = this have "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x,y)))) \ e" if "\x\?CBOX. \x'\?CBOX. norm (x' - x) < k \ norm (f x' - f x) < e /(2 * content (?CBOX))" "0 < k" for k proof - obtain p where ptag: "p tagged_division_of cbox (a, c) (b, d)" and fine: "(\x. ball x k) fine p" using fine_division_exists \0 < k\ by blast show ?thesis using that fine ptag \0 < k\ by (auto simp: * intro: op_acbd [OF division_of_tagged_division [OF ptag]]) qed then show "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x,y)))) \ e" using compact_uniformly_continuous [OF assms compact_cbox] apply (simp add: uniformly_continuous_on_def dist_norm) apply (drule_tac x="e/2 / content?CBOX" in spec) using cbp \0 < e\ by (auto simp: zero_less_mult_iff) qed then show ?thesis by simp qed lemma integral_swap_2dim: fixes f :: "['a::euclidean_space, 'b::euclidean_space] \ 'c::banach" assumes "continuous_on (cbox (a,c) (b,d)) (\(x,y). f x y)" shows "integral (cbox (a, c) (b, d)) (\(x, y). f x y) = integral (cbox (c, a) (d, b)) (\(x, y). f y x)" proof - have "((\(x, y). f x y) has_integral integral (cbox (c, a) (d, b)) (\(x, y). f y x)) (prod.swap ` (cbox (c, a) (d, b)))" proof (rule has_integral_twiddle [of 1 prod.swap prod.swap "\(x,y). f y x" "integral (cbox (c, a) (d, b)) (\(x, y). f y x)", simplified]) show "\u v. content (prod.swap ` cbox u v) = content (cbox u v)" by (metis content_Pair mult.commute old.prod.exhaust swap_cbox_Pair) show "((\(x, y). f y x) has_integral integral (cbox (c, a) (d, b)) (\(x, y). f y x)) (cbox (c, a) (d, b))" by (simp add: assms integrable_continuous integrable_integral swap_continuous) qed (use isCont_swap in \fastforce+\) then show ?thesis by force qed theorem integral_swap_continuous: fixes f :: "['a::euclidean_space, 'b::euclidean_space] \ 'c::banach" assumes "continuous_on (cbox (a,c) (b,d)) (\(x,y). f x y)" shows "integral (cbox a b) (\x. integral (cbox c d) (f x)) = integral (cbox c d) (\y. integral (cbox a b) (\x. f x y))" proof - have "integral (cbox a b) (\x. integral (cbox c d) (f x)) = integral (cbox (a, c) (b, d)) (\(x, y). f x y)" using integral_prod_continuous [OF assms] by auto also have "... = integral (cbox (c, a) (d, b)) (\(x, y). f y x)" by (rule integral_swap_2dim [OF assms]) also have "... = integral (cbox c d) (\y. integral (cbox a b) (\x. f x y))" using integral_prod_continuous [OF swap_continuous] assms by auto finally show ?thesis . qed subsection \Definite integrals for exponential and power function\ lemma has_integral_exp_minus_to_infinity: assumes a: "a > 0" shows "((\x::real. exp (-a*x)) has_integral exp (-a*c)/a) {c..}" proof - define f where "f = (\k x. if x \ {c..real k} then exp (-a*x) else 0)" { fix k :: nat assume k: "of_nat k \ c" from k a have "((\x. exp (-a*x)) has_integral (-exp (-a*real k)/a - (-exp (-a*c)/a))) {c..real k}" by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros simp: has_real_derivative_iff_has_vector_derivative [symmetric]) hence "(f k has_integral (exp (-a*c)/a - exp (-a*real k)/a)) {c..}" unfolding f_def by (subst has_integral_restrict) simp_all } note has_integral_f = this have [simp]: "f k = (\_. 0)" if "of_nat k < c" for k using that by (auto simp: fun_eq_iff f_def) have integral_f: "integral {c..} (f k) = (if real k \ c then exp (-a*c)/a - exp (-a*real k)/a else 0)" for k using integral_unique[OF has_integral_f[of k]] by simp have A: "(\x. exp (-a*x)) integrable_on {c..} \ (\k. integral {c..} (f k)) \ integral {c..} (\x. exp (-a*x))" proof (intro monotone_convergence_increasing allI ballI) fix k ::nat have "(\x. exp (-a*x)) integrable_on {c..of_real k}" unfolding f_def by (auto intro!: continuous_intros integrable_continuous_real) hence "(f k) integrable_on {c..of_real k}" by (rule integrable_eq) (simp add: f_def) then show "f k integrable_on {c..}" by (rule integrable_on_superset) (auto simp: f_def) next fix x assume x: "x \ {c..}" have "sequentially \ principal {nat \x\..}" unfolding at_top_def by (simp add: Inf_lower) also have "{nat \x\..} \ {k. x \ real k}" by auto also have "inf (principal \) (principal {k. \x \ real k}) = principal ({k. x \ real k} \ {k. \x \ real k})" by simp also have "{k. x \ real k} \ {k. \x \ real k} = {}" by blast finally have "inf sequentially (principal {k. \x \ real k}) = bot" by (simp add: inf.coboundedI1 bot_unique) with x show "(\k. f k x) \ exp (-a*x)" unfolding f_def by (intro filterlim_If) auto next have "\integral {c..} (f k)\ \ exp (-a*c)/a" for k proof (cases "c > of_nat k") case False hence "abs (integral {c..} (f k)) = abs (exp (- (a * c)) / a - exp (- (a * real k)) / a)" by (simp add: integral_f) also have "abs (exp (- (a * c)) / a - exp (- (a * real k)) / a) = exp (- (a * c)) / a - exp (- (a * real k)) / a" using False a by (intro abs_of_nonneg) (simp_all add: field_simps) also have "\ \ exp (- a * c) / a" using a by simp finally show ?thesis . qed (insert a, simp_all add: integral_f) thus "bounded (range(\k. integral {c..} (f k)))" by (intro boundedI[of _ "exp (-a*c)/a"]) auto qed (auto simp: f_def) have "(\k. exp (-a*c)/a - exp (-a * of_nat k)/a) \ exp (-a*c)/a - 0/a" by (intro tendsto_intros filterlim_compose[OF exp_at_bot] filterlim_tendsto_neg_mult_at_bot[OF tendsto_const] filterlim_real_sequentially)+ (use a in simp_all) moreover from eventually_gt_at_top[of "nat \c\"] have "eventually (\k. of_nat k > c) sequentially" by eventually_elim linarith hence "eventually (\k. exp (-a*c)/a - exp (-a * of_nat k)/a = integral {c..} (f k)) sequentially" by eventually_elim (simp add: integral_f) ultimately have "(\k. integral {c..} (f k)) \ exp (-a*c)/a - 0/a" by (rule Lim_transform_eventually) from LIMSEQ_unique[OF conjunct2[OF A] this] have "integral {c..} (\x. exp (-a*x)) = exp (-a*c)/a" by simp with conjunct1[OF A] show ?thesis by (simp add: has_integral_integral) qed lemma integrable_on_exp_minus_to_infinity: "a > 0 \ (\x. exp (-a*x) :: real) integrable_on {c..}" using has_integral_exp_minus_to_infinity[of a c] unfolding integrable_on_def by blast lemma has_integral_powr_from_0: assumes a: "a > (-1::real)" and c: "c \ 0" shows "((\x. x powr a) has_integral (c powr (a+1) / (a+1))) {0..c}" proof (cases "c = 0") case False define f where "f = (\k x. if x \ {inverse (of_nat (Suc k))..c} then x powr a else 0)" define F where "F = (\k. if inverse (of_nat (Suc k)) \ c then c powr (a+1)/(a+1) - inverse (real (Suc k)) powr (a+1)/(a+1) else 0)" have has_integral_f: "(f k has_integral F k) {0..c}" for k::nat proof (cases "inverse (of_nat (Suc k)) \ c") case True have x: "x > 0" if "x \ inverse (1 + real k)" for x by (smt (verit) that inverse_Suc of_nat_Suc) hence "((\x. x powr a) has_integral c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1)) {inverse (real (Suc k))..c}" using True a by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros continuous_on_powr' continuous_on_const simp: has_real_derivative_iff_has_vector_derivative [symmetric]) with True show ?thesis unfolding f_def F_def by (subst has_integral_restrict) simp_all next case False thus ?thesis unfolding f_def F_def by (subst has_integral_restrict) auto qed then have integral_f: "integral {0..c} (f k) = F k" for k by blast have A: "(\x. x powr a) integrable_on {0..c} \ (\k. integral {0..c} (f k)) \ integral {0..c} (\x. x powr a)" proof (intro monotone_convergence_increasing ballI allI) fix k from has_integral_f[of k] show "f k integrable_on {0..c}" by (auto simp: integrable_on_def) next fix k :: nat and x :: real { assume x: "inverse (real (Suc k)) \ x" then have "inverse (real (Suc (Suc k))) \ x" using dual_order.trans by fastforce } thus "f k x \ f (Suc k) x" by (auto simp: f_def simp del: of_nat_Suc) next fix x assume x: "x \ {0..c}" show "(\k. f k x) \ x powr a" proof (cases "x = 0") case False with x have "x > 0" by simp from order_tendstoD(2)[OF LIMSEQ_inverse_real_of_nat this] have "eventually (\k. x powr a = f k x) sequentially" by eventually_elim (insert x, simp add: f_def) moreover have "(\_. x powr a) \ x powr a" by simp ultimately show ?thesis by (blast intro: Lim_transform_eventually) qed (simp_all add: f_def) next { fix k from a have "F k \ c powr (a + 1) / (a + 1)" by (auto simp add: F_def divide_simps) also from a have "F k \ 0" by (auto simp: F_def divide_simps simp del: of_nat_Suc intro!: powr_mono2) hence "F k = abs (F k)" by simp finally have "abs (F k) \ c powr (a + 1) / (a + 1)" . } thus "bounded (range(\k. integral {0..c} (f k)))" by (intro boundedI[of _ "c powr (a+1) / (a+1)"]) (auto simp: integral_f) qed from False c have "c > 0" by simp from order_tendstoD(2)[OF LIMSEQ_inverse_real_of_nat this] have "eventually (\k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a+1) / (a+1) = integral {0..c} (f k)) sequentially" by eventually_elim (simp add: integral_f F_def) moreover have "(\k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1)) \ c powr (a + 1) / (a + 1) - 0 powr (a + 1) / (a + 1)" using a by (intro tendsto_intros LIMSEQ_inverse_real_of_nat) auto hence "(\k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1)) \ c powr (a + 1) / (a + 1)" by simp ultimately have "(\k. integral {0..c} (f k)) \ c powr (a+1) / (a+1)" by (blast intro: Lim_transform_eventually) with A have "integral {0..c} (\x. x powr a) = c powr (a+1) / (a+1)" by (blast intro: LIMSEQ_unique) with A show ?thesis by (simp add: has_integral_integral) qed (simp_all add: has_integral_refl) lemma integrable_on_powr_from_0: assumes a: "a > (-1::real)" and c: "c \ 0" shows "(\x. x powr a) integrable_on {0..c}" using has_integral_powr_from_0[OF assms] unfolding integrable_on_def by blast lemma has_integral_powr_to_inf: fixes a e :: real assumes "e < -1" "a > 0" shows "((\x. x powr e) has_integral -(a powr (e + 1)) / (e + 1)) {a..}" proof - define f :: "nat \ real \ real" where "f = (\n x. if x \ {a..n} then x powr e else 0)" define F where "F = (\x. x powr (e + 1) / (e + 1))" have has_integral_f: "(f n has_integral (F n - F a)) {a..}" if n: "n \ a" for n :: nat proof - from n assms have "((\x. x powr e) has_integral (F n - F a)) {a..n}" by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros simp: has_real_derivative_iff_has_vector_derivative [symmetric] F_def) hence "(f n has_integral (F n - F a)) {a..n}" by (rule has_integral_eq [rotated]) (simp add: f_def) thus "(f n has_integral (F n - F a)) {a..}" by (rule has_integral_on_superset) (auto simp: f_def) qed have integral_f: "integral {a..} (f n) = (if n \ a then F n - F a else 0)" for n :: nat proof (cases "n \ a") case True with has_integral_f[OF this] show ?thesis by (simp add: integral_unique) next case False have "(f n has_integral 0) {a}" by (rule has_integral_refl) hence "(f n has_integral 0) {a..}" using False f_def by force with False show ?thesis by (simp add: integral_unique) qed have *: "(\x. x powr e) integrable_on {a..} \ (\n. integral {a..} (f n)) \ integral {a..} (\x. x powr e)" proof (intro monotone_convergence_increasing allI ballI) fix n :: nat from assms have "(\x. x powr e) integrable_on {a..n}" by (auto intro!: integrable_continuous_real continuous_intros) hence "f n integrable_on {a..n}" by (rule integrable_eq) (auto simp: f_def) thus "f n integrable_on {a..}" by (rule integrable_on_superset) (auto simp: f_def) next fix n :: nat and x :: real show "f n x \ f (Suc n) x" by (auto simp: f_def) next fix x :: real assume x: "x \ {a..}" from filterlim_real_sequentially have "eventually (\n. real n \ x) at_top" by (simp add: filterlim_at_top) with x have "eventually (\n. f n x = x powr e) at_top" by (auto elim!: eventually_mono simp: f_def) thus "(\n. f n x) \ x powr e" by (simp add: tendsto_eventually) next have "norm (integral {a..} (f n)) \ -F a" for n :: nat proof (cases "n \ a") case True with assms have "a powr (e + 1) \ n powr (e + 1)" by (intro powr_mono2') simp_all with assms show ?thesis by (auto simp: divide_simps F_def integral_f) qed (use assms in \simp add: integral_f F_def field_split_simps\) thus "bounded (range(\k. integral {a..} (f k)))" unfolding bounded_iff by (intro exI[of _ "-F a"]) auto qed from filterlim_real_sequentially have "eventually (\n. real n \ a) at_top" by (simp add: filterlim_at_top) hence "eventually (\n. F n - F a = integral {a..} (f n)) at_top" by eventually_elim (simp add: integral_f) moreover have "(\n. F n - F a) \ 0 / (e + 1) - F a" unfolding F_def by (insert assms, (rule tendsto_intros filterlim_compose[OF tendsto_neg_powr] filterlim_ident filterlim_real_sequentially | simp)+) hence "(\n. F n - F a) \ -F a" by simp ultimately have "(\n. integral {a..} (f n)) \ -F a" by (blast intro: Lim_transform_eventually) then have "integral {a..} (\x. x powr e) = -F a" using "*" LIMSEQ_unique by blast with * show ?thesis by (simp add: has_integral_integral F_def) qed lemma has_integral_inverse_power_to_inf: fixes a :: real and n :: nat assumes "n > 1" "a > 0" shows "((\x. 1 / x ^ n) has_integral 1 / (real (n - 1) * a ^ (n - 1))) {a..}" proof - from assms have "real_of_int (-int n) < -1" by simp note has_integral_powr_to_inf[OF this \a > 0\] also have "- (a powr (real_of_int (- int n) + 1)) / (real_of_int (- int n) + 1) = 1 / (real (n - 1) * a powr (real (n - 1)))" using assms by (simp add: field_split_simps powr_add [symmetric] of_nat_diff) also from assms have "a powr (real (n - 1)) = a ^ (n - 1)" by (intro powr_realpow) finally show ?thesis by (rule has_integral_eq [rotated]) (insert assms, simp_all add: powr_minus powr_realpow field_split_simps) qed subsubsection \Adaption to ordered Euclidean spaces and the Cartesian Euclidean space\ lemma integral_component_eq_cart[simp]: fixes f :: "'n::euclidean_space \ real^'m" assumes "f integrable_on s" shows "integral s (\x. f x $ k) = integral s f $ k" using integral_linear[OF assms(1) bounded_linear_vec_nth,unfolded o_def] . lemma content_closed_interval: fixes a :: "'a::ordered_euclidean_space" assumes "a \ b" shows "content {a..b} = (\i\Basis. b\i - a\i)" using content_cbox[of a b] assms by (simp add: cbox_interval eucl_le[where 'a='a]) lemma integrable_const_ivl[intro]: fixes a::"'a::ordered_euclidean_space" shows "(\x. c) integrable_on {a..b}" unfolding cbox_interval[symmetric] by (rule integrable_const) lemma integrable_on_subinterval: fixes f :: "'n::ordered_euclidean_space \ 'a::banach" assumes "f integrable_on S" "{a..b} \ S" shows "f integrable_on {a..b}" using integrable_on_subcbox[of f S a b] assms by (simp add: cbox_interval) end diff --git a/src/HOL/Analysis/Homotopy.thy b/src/HOL/Analysis/Homotopy.thy --- a/src/HOL/Analysis/Homotopy.thy +++ b/src/HOL/Analysis/Homotopy.thy @@ -1,5381 +1,5390 @@ (* Title: HOL/Analysis/Path_Connected.thy Authors: LC Paulson and Robert Himmelmann (TU Muenchen), based on material from HOL Light *) section \Homotopy of Maps\ theory Homotopy imports Path_Connected Product_Topology Uncountable_Sets begin definition\<^marker>\tag important\ homotopic_with where "homotopic_with P X Y f g \ (\h. continuous_map (prod_topology (top_of_set {0..1::real}) X) Y h \ (\x. h(0, x) = f x) \ (\x. h(1, x) = g x) \ (\t \ {0..1}. P(\x. h(t,x))))" text\\p\, \q\ are functions \X \ Y\, and the property \P\ restricts all intermediate maps. We often just want to require that \P\ fixes some subset, but to include the case of a loop homotopy, it is convenient to have a general property \P\.\ abbreviation homotopic_with_canon :: "[('a::topological_space \ 'b::topological_space) \ bool, 'a set, 'b set, 'a \ 'b, 'a \ 'b] \ bool" where "homotopic_with_canon P S T p q \ homotopic_with P (top_of_set S) (top_of_set T) p q" lemma split_01: "{0..1::real} = {0..1/2} \ {1/2..1}" by force lemma split_01_prod: "{0..1::real} \ X = ({0..1/2} \ X) \ ({1/2..1} \ X)" by force lemma image_Pair_const: "(\x. (x, c)) ` A = A \ {c}" by auto lemma fst_o_paired [simp]: "fst \ (\(x,y). (f x y, g x y)) = (\(x,y). f x y)" by auto lemma snd_o_paired [simp]: "snd \ (\(x,y). (f x y, g x y)) = (\(x,y). g x y)" by auto lemma continuous_on_o_Pair: "\continuous_on (T \ X) h; t \ T\ \ continuous_on X (h \ Pair t)" by (fast intro: continuous_intros elim!: continuous_on_subset) lemma continuous_map_o_Pair: assumes h: "continuous_map (prod_topology X Y) Z h" and t: "t \ topspace X" shows "continuous_map Y Z (h \ Pair t)" by (intro continuous_map_compose [OF _ h] continuous_intros; simp add: t) subsection\<^marker>\tag unimportant\\Trivial properties\ text \We often want to just localize the ending function equality or whatever.\ text\<^marker>\tag important\ \%whitespace\ proposition homotopic_with: assumes "\h k. (\x. x \ topspace X \ h x = k x) \ (P h \ P k)" shows "homotopic_with P X Y p q \ (\h. continuous_map (prod_topology (subtopology euclideanreal {0..1}) X) Y h \ (\x \ topspace X. h(0,x) = p x) \ (\x \ topspace X. h(1,x) = q x) \ (\t \ {0..1}. P(\x. h(t, x))))" unfolding homotopic_with_def apply (rule iffI, blast, clarify) apply (rule_tac x="\(u,v). if v \ topspace X then h(u,v) else if u = 0 then p v else q v" in exI) apply simp by (smt (verit, best) SigmaE assms case_prod_conv continuous_map_eq topspace_prod_topology) lemma homotopic_with_mono: assumes hom: "homotopic_with P X Y f g" and Q: "\h. \continuous_map X Y h; P h\ \ Q h" shows "homotopic_with Q X Y f g" using hom unfolding homotopic_with_def by (force simp: o_def dest: continuous_map_o_Pair intro: Q) lemma homotopic_with_imp_continuous_maps: assumes "homotopic_with P X Y f g" shows "continuous_map X Y f \ continuous_map X Y g" proof - obtain h :: "real \ 'a \ 'b" where conth: "continuous_map (prod_topology (top_of_set {0..1}) X) Y h" and h: "\x. h (0, x) = f x" "\x. h (1, x) = g x" using assms by (auto simp: homotopic_with_def) have *: "t \ {0..1} \ continuous_map X Y (h \ (\x. (t,x)))" for t by (rule continuous_map_compose [OF _ conth]) (simp add: o_def continuous_map_pairwise) show ?thesis using h *[of 0] *[of 1] by (simp add: continuous_map_eq) qed lemma homotopic_with_imp_continuous: assumes "homotopic_with_canon P X Y f g" shows "continuous_on X f \ continuous_on X g" by (meson assms continuous_map_subtopology_eu homotopic_with_imp_continuous_maps) lemma homotopic_with_imp_property: assumes "homotopic_with P X Y f g" shows "P f \ P g" proof obtain h where h: "\x. h(0, x) = f x" "\x. h(1, x) = g x" and P: "\t. t \ {0..1::real} \ P(\x. h(t,x))" using assms by (force simp: homotopic_with_def) show "P f" "P g" using P [of 0] P [of 1] by (force simp: h)+ qed lemma homotopic_with_equal: assumes "P f" "P g" and contf: "continuous_map X Y f" and fg: "\x. x \ topspace X \ f x = g x" shows "homotopic_with P X Y f g" unfolding homotopic_with_def proof (intro exI conjI allI ballI) let ?h = "\(t::real,x). if t = 1 then g x else f x" show "continuous_map (prod_topology (top_of_set {0..1}) X) Y ?h" proof (rule continuous_map_eq) show "continuous_map (prod_topology (top_of_set {0..1}) X) Y (f \ snd)" by (simp add: contf continuous_map_of_snd) qed (auto simp: fg) show "P (\x. ?h (t, x))" if "t \ {0..1}" for t by (cases "t = 1") (simp_all add: assms) qed auto lemma homotopic_with_imp_subset1: "homotopic_with_canon P X Y f g \ f ` X \ Y" by (meson continuous_map_subtopology_eu homotopic_with_imp_continuous_maps) lemma homotopic_with_imp_subset2: "homotopic_with_canon P X Y f g \ g ` X \ Y" by (meson continuous_map_subtopology_eu homotopic_with_imp_continuous_maps) lemma homotopic_with_imp_funspace1: "homotopic_with_canon P X Y f g \ f \ X \ Y" using homotopic_with_imp_subset1 by blast lemma homotopic_with_imp_funspace2: "homotopic_with_canon P X Y f g \ g \ X \ Y" using homotopic_with_imp_subset2 by blast lemma homotopic_with_subset_left: "\homotopic_with_canon P X Y f g; Z \ X\ \ homotopic_with_canon P Z Y f g" unfolding homotopic_with_def by (auto elim!: continuous_on_subset ex_forward) lemma homotopic_with_subset_right: "\homotopic_with_canon P X Y f g; Y \ Z\ \ homotopic_with_canon P X Z f g" unfolding homotopic_with_def by (auto elim!: continuous_on_subset ex_forward) subsection\Homotopy with P is an equivalence relation\ text \(on continuous functions mapping X into Y that satisfy P, though this only affects reflexivity)\ lemma homotopic_with_refl [simp]: "homotopic_with P X Y f f \ continuous_map X Y f \ P f" by (metis homotopic_with_equal homotopic_with_imp_continuous_maps homotopic_with_imp_property) lemma homotopic_with_symD: assumes "homotopic_with P X Y f g" shows "homotopic_with P X Y g f" proof - let ?I01 = "subtopology euclideanreal {0..1}" let ?j = "\y. (1 - fst y, snd y)" have 1: "continuous_map (prod_topology ?I01 X) (prod_topology euclideanreal X) ?j" by (intro continuous_intros; simp add: continuous_map_subtopology_fst prod_topology_subtopology) have *: "continuous_map (prod_topology ?I01 X) (prod_topology ?I01 X) ?j" proof - have "continuous_map (prod_topology ?I01 X) (subtopology (prod_topology euclideanreal X) ({0..1} \ topspace X)) ?j" by (simp add: continuous_map_into_subtopology [OF 1] image_subset_iff flip: image_subset_iff_funcset) then show ?thesis by (simp add: prod_topology_subtopology(1)) qed show ?thesis using assms apply (clarsimp simp: homotopic_with_def) subgoal for h by (rule_tac x="h \ (\y. (1 - fst y, snd y))" in exI) (simp add: continuous_map_compose [OF *]) done qed lemma homotopic_with_sym: "homotopic_with P X Y f g \ homotopic_with P X Y g f" by (metis homotopic_with_symD) proposition homotopic_with_trans: assumes "homotopic_with P X Y f g" "homotopic_with P X Y g h" shows "homotopic_with P X Y f h" proof - let ?X01 = "prod_topology (subtopology euclideanreal {0..1}) X" obtain k1 k2 where contk1: "continuous_map ?X01 Y k1" and contk2: "continuous_map ?X01 Y k2" and k12: "\x. k1 (1, x) = g x" "\x. k2 (0, x) = g x" "\x. k1 (0, x) = f x" "\x. k2 (1, x) = h x" and P: "\t\{0..1}. P (\x. k1 (t, x))" "\t\{0..1}. P (\x. k2 (t, x))" using assms by (auto simp: homotopic_with_def) define k where "k \ \y. if fst y \ 1/2 then (k1 \ (\x. (2 *\<^sub>R fst x, snd x))) y else (k2 \ (\x. (2 *\<^sub>R fst x -1, snd x))) y" have keq: "k1 (2 * u, v) = k2 (2 * u -1, v)" if "u = 1/2" for u v by (simp add: k12 that) show ?thesis unfolding homotopic_with_def proof (intro exI conjI) show "continuous_map ?X01 Y k" unfolding k_def proof (rule continuous_map_cases_le) show fst: "continuous_map ?X01 euclideanreal fst" using continuous_map_fst continuous_map_in_subtopology by blast show "continuous_map ?X01 euclideanreal (\x. 1/2)" by simp show "continuous_map (subtopology ?X01 {y \ topspace ?X01. fst y \ 1/2}) Y (k1 \ (\x. (2 *\<^sub>R fst x, snd x)))" apply (intro fst continuous_map_compose [OF _ contk1] continuous_intros continuous_map_into_subtopology continuous_map_from_subtopology | simp)+ by (force simp: prod_topology_subtopology) show "continuous_map (subtopology ?X01 {y \ topspace ?X01. 1/2 \ fst y}) Y (k2 \ (\x. (2 *\<^sub>R fst x -1, snd x)))" apply (intro fst continuous_map_compose [OF _ contk2] continuous_intros continuous_map_into_subtopology continuous_map_from_subtopology | simp)+ by (force simp: prod_topology_subtopology) show "(k1 \ (\x. (2 *\<^sub>R fst x, snd x))) y = (k2 \ (\x. (2 *\<^sub>R fst x -1, snd x))) y" if "y \ topspace ?X01" and "fst y = 1/2" for y using that by (simp add: keq) qed show "\x. k (0, x) = f x" by (simp add: k12 k_def) show "\x. k (1, x) = h x" by (simp add: k12 k_def) show "\t\{0..1}. P (\x. k (t, x))" proof fix t show "t\{0..1} \ P (\x. k (t, x))" by (cases "t \ 1/2") (auto simp: k_def P) qed qed qed lemma homotopic_with_id2: "(\x. x \ topspace X \ g (f x) = x) \ homotopic_with (\x. True) X X (g \ f) id" by (metis comp_apply continuous_map_id eq_id_iff homotopic_with_equal homotopic_with_symD) subsection\Continuity lemmas\ lemma homotopic_with_compose_continuous_map_left: "\homotopic_with p X1 X2 f g; continuous_map X2 X3 h; \j. p j \ q(h \ j)\ \ homotopic_with q X1 X3 (h \ f) (h \ g)" unfolding homotopic_with_def apply clarify subgoal for k by (rule_tac x="h \ k" in exI) (rule conjI continuous_map_compose | simp add: o_def)+ done lemma homotopic_with_compose_continuous_map_right: assumes hom: "homotopic_with p X2 X3 f g" and conth: "continuous_map X1 X2 h" and q: "\j. p j \ q(j \ h)" shows "homotopic_with q X1 X3 (f \ h) (g \ h)" proof - obtain k where contk: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X2) X3 k" and k: "\x. k (0, x) = f x" "\x. k (1, x) = g x" and p: "\t. t\{0..1} \ p (\x. k (t, x))" using hom unfolding homotopic_with_def by blast have hsnd: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X1) X2 (h \ snd)" by (rule continuous_map_compose [OF continuous_map_snd conth]) let ?h = "k \ (\(t,x). (t,h x))" show ?thesis unfolding homotopic_with_def proof (intro exI conjI allI ballI) have "continuous_map (prod_topology (top_of_set {0..1}) X1) (prod_topology (top_of_set {0..1::real}) X2) (\(t, x). (t, h x))" by (metis (mono_tags, lifting) case_prod_beta' comp_def continuous_map_eq continuous_map_fst continuous_map_pairedI hsnd) then show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X1) X3 ?h" by (intro conjI continuous_map_compose [OF _ contk]) show "q (\x. ?h (t, x))" if "t \ {0..1}" for t using q [OF p [OF that]] by (simp add: o_def) qed (auto simp: k) qed corollary homotopic_compose: assumes "homotopic_with (\x. True) X Y f f'" "homotopic_with (\x. True) Y Z g g'" shows "homotopic_with (\x. True) X Z (g \ f) (g' \ f')" by (metis assms homotopic_with_compose_continuous_map_left homotopic_with_compose_continuous_map_right homotopic_with_imp_continuous_maps homotopic_with_trans) proposition homotopic_with_compose_continuous_right: "\homotopic_with_canon (\f. p (f \ h)) X Y f g; continuous_on W h; h \ W \ X\ \ homotopic_with_canon p W Y (f \ h) (g \ h)" by (simp add: homotopic_with_compose_continuous_map_right image_subset_iff_funcset) proposition homotopic_with_compose_continuous_left: "\homotopic_with_canon (\f. p (h \ f)) X Y f g; continuous_on Y h; h \ Y \ Z\ \ homotopic_with_canon p X Z (h \ f) (h \ g)" by (simp add: homotopic_with_compose_continuous_map_left image_subset_iff_funcset) lemma homotopic_from_subtopology: "homotopic_with P X X' f g \ homotopic_with P (subtopology X S) X' f g" by (metis continuous_map_id_subt homotopic_with_compose_continuous_map_right o_id) lemma homotopic_on_emptyI: assumes "P f" "P g" shows "homotopic_with P trivial_topology X f g" by (metis assms continuous_map_on_empty empty_iff homotopic_with_equal topspace_discrete_topology) lemma homotopic_on_empty: "(homotopic_with P trivial_topology X f g \ P f \ P g)" using homotopic_on_emptyI homotopic_with_imp_property by metis lemma homotopic_with_canon_on_empty: "homotopic_with_canon (\x. True) {} t f g" by (auto intro: homotopic_with_equal) lemma homotopic_constant_maps: "homotopic_with (\x. True) X X' (\x. a) (\x. b) \ X = trivial_topology \ path_component_of X' a b" (is "?lhs = ?rhs") proof (cases "X = trivial_topology") case False then obtain c where c: "c \ topspace X" by fastforce have "\g. continuous_map (top_of_set {0..1::real}) X' g \ g 0 = a \ g 1 = b" if "x \ topspace X" and hom: "homotopic_with (\x. True) X X' (\x. a) (\x. b)" for x proof - obtain h :: "real \ 'a \ 'b" where conth: "continuous_map (prod_topology (top_of_set {0..1}) X) X' h" and h: "\x. h (0, x) = a" "\x. h (1, x) = b" using hom by (auto simp: homotopic_with_def) have cont: "continuous_map (top_of_set {0..1}) X' (h \ (\t. (t, c)))" by (rule continuous_map_compose [OF _ conth] continuous_intros | simp add: c)+ then show ?thesis by (force simp: h) qed moreover have "homotopic_with (\x. True) X X' (\x. g 0) (\x. g 1)" if "x \ topspace X" "a = g 0" "b = g 1" "continuous_map (top_of_set {0..1}) X' g" for x and g :: "real \ 'b" unfolding homotopic_with_def by (force intro!: continuous_map_compose continuous_intros c that) ultimately show ?thesis using False by (metis c path_component_of_set pathin_def) qed (simp add: homotopic_on_empty) proposition homotopic_with_eq: assumes h: "homotopic_with P X Y f g" and f': "\x. x \ topspace X \ f' x = f x" and g': "\x. x \ topspace X \ g' x = g x" and P: "(\h k. (\x. x \ topspace X \ h x = k x) \ P h \ P k)" shows "homotopic_with P X Y f' g'" by (smt (verit, ccfv_SIG) assms homotopic_with) lemma homotopic_with_prod_topology: assumes "homotopic_with p X1 Y1 f f'" and "homotopic_with q X2 Y2 g g'" and r: "\i j. \p i; q j\ \ r(\(x,y). (i x, j y))" shows "homotopic_with r (prod_topology X1 X2) (prod_topology Y1 Y2) (\z. (f(fst z),g(snd z))) (\z. (f'(fst z), g'(snd z)))" proof - obtain h where h: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X1) Y1 h" and h0: "\x. h (0, x) = f x" and h1: "\x. h (1, x) = f' x" and p: "\t. \0 \ t; t \ 1\ \ p (\x. h (t,x))" using assms unfolding homotopic_with_def by auto obtain k where k: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X2) Y2 k" and k0: "\x. k (0, x) = g x" and k1: "\x. k (1, x) = g' x" and q: "\t. \0 \ t; t \ 1\ \ q (\x. k (t,x))" using assms unfolding homotopic_with_def by auto let ?hk = "\(t,x,y). (h(t,x), k(t,y))" show ?thesis unfolding homotopic_with_def proof (intro conjI allI exI) show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) (prod_topology X1 X2)) (prod_topology Y1 Y2) ?hk" unfolding continuous_map_pairwise case_prod_unfold by (rule conjI continuous_map_pairedI continuous_intros continuous_map_id [unfolded id_def] continuous_map_fst_of [unfolded o_def] continuous_map_snd_of [unfolded o_def] continuous_map_compose [OF _ h, unfolded o_def] continuous_map_compose [OF _ k, unfolded o_def])+ next fix x show "?hk (0, x) = (f (fst x), g (snd x))" "?hk (1, x) = (f' (fst x), g' (snd x))" by (auto simp: case_prod_beta h0 k0 h1 k1) qed (auto simp: p q r) qed lemma homotopic_with_product_topology: assumes ht: "\i. i \ I \ homotopic_with (p i) (X i) (Y i) (f i) (g i)" and pq: "\h. (\i. i \ I \ p i (h i)) \ q(\x. (\i\I. h i (x i)))" shows "homotopic_with q (product_topology X I) (product_topology Y I) (\z. (\i\I. (f i) (z i))) (\z. (\i\I. (g i) (z i)))" proof - obtain h where h: "\i. i \ I \ continuous_map (prod_topology (subtopology euclideanreal {0..1}) (X i)) (Y i) (h i)" and h0: "\i x. i \ I \ h i (0, x) = f i x" and h1: "\i x. i \ I \ h i (1, x) = g i x" and p: "\i t. \i \ I; t \ {0..1}\ \ p i (\x. h i (t,x))" using ht unfolding homotopic_with_def by metis show ?thesis unfolding homotopic_with_def proof (intro conjI allI exI) let ?h = "\(t,z). \i\I. h i (t,z i)" have "continuous_map (prod_topology (subtopology euclideanreal {0..1}) (product_topology X I)) (Y i) (\x. h i (fst x, snd x i))" if "i \ I" for i proof - have \
: "continuous_map (prod_topology (top_of_set {0..1}) (product_topology X I)) (X i) (\x. snd x i)" using continuous_map_componentwise continuous_map_snd that by fastforce show ?thesis unfolding continuous_map_pairwise case_prod_unfold by (intro conjI that \
continuous_intros continuous_map_compose [OF _ h, unfolded o_def]) qed then show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) (product_topology X I)) (product_topology Y I) ?h" by (auto simp: continuous_map_componentwise case_prod_beta) show "?h (0, x) = (\i\I. f i (x i))" "?h (1, x) = (\i\I. g i (x i))" for x by (auto simp: case_prod_beta h0 h1) show "\t\{0..1}. q (\x. ?h (t, x))" by (force intro: p pq) qed qed text\Homotopic triviality implicitly incorporates path-connectedness.\ lemma homotopic_triviality: shows "(\f g. continuous_on S f \ f \ S \ T \ continuous_on S g \ g \ S \ T \ homotopic_with_canon (\x. True) S T f g) \ (S = {} \ path_connected T) \ (\f. continuous_on S f \ f \ S \ T \ (\c. homotopic_with_canon (\x. True) S T f (\x. c)))" (is "?lhs = ?rhs") proof (cases "S = {} \ T = {}") case True then show ?thesis by (auto simp: homotopic_on_emptyI simp flip: image_subset_iff_funcset) next case False show ?thesis proof assume LHS [rule_format]: ?lhs have pab: "path_component T a b" if "a \ T" "b \ T" for a b proof - have "homotopic_with_canon (\x. True) S T (\x. a) (\x. b)" by (simp add: LHS image_subset_iff that) then show ?thesis using False homotopic_constant_maps [of "top_of_set S" "top_of_set T" a b] by (metis path_component_of_canon_iff topspace_discrete_topology topspace_euclidean_subtopology) qed moreover have "\c. homotopic_with_canon (\x. True) S T f (\x. c)" if "continuous_on S f" "f \ S \ T" for f using False LHS continuous_on_const that by blast ultimately show ?rhs by (simp add: path_connected_component) next assume RHS: ?rhs with False have T: "path_connected T" by blast show ?lhs proof clarify fix f g assume "continuous_on S f" "f \ S \ T" "continuous_on S g" "g \ S \ T" obtain c d where c: "homotopic_with_canon (\x. True) S T f (\x. c)" and d: "homotopic_with_canon (\x. True) S T g (\x. d)" using RHS \continuous_on S f\ \continuous_on S g\ \f \ S \ T\ \g \ S \ T\ by presburger with T have "path_component T c d" by (metis False ex_in_conv homotopic_with_imp_subset2 image_subset_iff path_connected_component) then have "homotopic_with_canon (\x. True) S T (\x. c) (\x. d)" by (simp add: homotopic_constant_maps) with c d show "homotopic_with_canon (\x. True) S T f g" by (meson homotopic_with_symD homotopic_with_trans) qed qed qed subsection\Homotopy of paths, maintaining the same endpoints\ definition\<^marker>\tag important\ homotopic_paths :: "['a set, real \ 'a, real \ 'a::topological_space] \ bool" where "homotopic_paths S p q \ homotopic_with_canon (\r. pathstart r = pathstart p \ pathfinish r = pathfinish p) {0..1} S p q" lemma homotopic_paths: "homotopic_paths S p q \ (\h. continuous_on ({0..1} \ {0..1}) h \ h \ ({0..1} \ {0..1}) \ S \ (\x \ {0..1}. h(0,x) = p x) \ (\x \ {0..1}. h(1,x) = q x) \ (\t \ {0..1::real}. pathstart(h \ Pair t) = pathstart p \ pathfinish(h \ Pair t) = pathfinish p))" by (auto simp: homotopic_paths_def homotopic_with pathstart_def pathfinish_def) proposition homotopic_paths_imp_pathstart: "homotopic_paths S p q \ pathstart p = pathstart q" by (metis (mono_tags, lifting) homotopic_paths_def homotopic_with_imp_property) proposition homotopic_paths_imp_pathfinish: "homotopic_paths S p q \ pathfinish p = pathfinish q" by (metis (mono_tags, lifting) homotopic_paths_def homotopic_with_imp_property) lemma homotopic_paths_imp_path: "homotopic_paths S p q \ path p \ path q" using homotopic_paths_def homotopic_with_imp_continuous_maps path_def continuous_map_subtopology_eu by blast lemma homotopic_paths_imp_subset: "homotopic_paths S p q \ path_image p \ S \ path_image q \ S" by (metis (mono_tags) continuous_map_subtopology_eu homotopic_paths_def homotopic_with_imp_continuous_maps path_image_def) proposition homotopic_paths_refl [simp]: "homotopic_paths S p p \ path p \ path_image p \ S" by (simp add: homotopic_paths_def path_def path_image_def) proposition homotopic_paths_sym: "homotopic_paths S p q \ homotopic_paths S q p" by (metis (mono_tags) homotopic_paths_def homotopic_paths_imp_pathfinish homotopic_paths_imp_pathstart homotopic_with_symD) proposition homotopic_paths_sym_eq: "homotopic_paths S p q \ homotopic_paths S q p" by (metis homotopic_paths_sym) proposition homotopic_paths_trans [trans]: assumes "homotopic_paths S p q" "homotopic_paths S q r" shows "homotopic_paths S p r" using assms homotopic_paths_imp_pathfinish homotopic_paths_imp_pathstart unfolding homotopic_paths_def by (smt (verit, ccfv_SIG) homotopic_with_mono homotopic_with_trans) proposition homotopic_paths_eq: "\path p; path_image p \ S; \t. t \ {0..1} \ p t = q t\ \ homotopic_paths S p q" by (smt (verit, best) homotopic_paths homotopic_paths_refl) proposition homotopic_paths_reparametrize: assumes "path p" and pips: "path_image p \ S" and contf: "continuous_on {0..1} f" and f01 :"f \ {0..1} \ {0..1}" and [simp]: "f(0) = 0" "f(1) = 1" and q: "\t. t \ {0..1} \ q(t) = p(f t)" shows "homotopic_paths S p q" proof - have contp: "continuous_on {0..1} p" by (metis \path p\ path_def) then have "continuous_on {0..1} (p \ f)" by (meson assms(4) contf continuous_on_compose continuous_on_subset image_subset_iff_funcset) then have "path q" by (simp add: path_def) (metis q continuous_on_cong) have piqs: "path_image q \ S" by (smt (verit, ccfv_threshold) Pi_iff assms(2) assms(4) assms(7) image_subset_iff path_defs(4)) have fb0: "\a b. \0 \ a; a \ 1; 0 \ b; b \ 1\ \ 0 \ (1 - a) * f b + a * b" using f01 by force have fb1: "\0 \ a; a \ 1; 0 \ b; b \ 1\ \ (1 - a) * f b + a * b \ 1" for a b by (intro convex_bound_le) (use f01 in auto) have "homotopic_paths S q p" proof (rule homotopic_paths_trans) show "homotopic_paths S q (p \ f)" using q by (force intro: homotopic_paths_eq [OF \path q\ piqs]) next show "homotopic_paths S (p \ f) p" using pips [unfolded path_image_def] apply (simp add: homotopic_paths_def homotopic_with_def) apply (rule_tac x="p \ (\y. (1 - (fst y)) *\<^sub>R ((f \ snd) y) + (fst y) *\<^sub>R snd y)" in exI) apply (rule conjI contf continuous_intros continuous_on_subset [OF contp] | simp)+ by (auto simp: fb0 fb1 pathstart_def pathfinish_def) qed then show ?thesis by (simp add: homotopic_paths_sym) qed lemma homotopic_paths_subset: "\homotopic_paths S p q; S \ t\ \ homotopic_paths t p q" unfolding homotopic_paths by fast text\ A slightly ad-hoc but useful lemma in constructing homotopies.\ lemma continuous_on_homotopic_join_lemma: fixes q :: "[real,real] \ 'a::topological_space" assumes p: "continuous_on ({0..1} \ {0..1}) (\y. p (fst y) (snd y))" (is "continuous_on ?A ?p") and q: "continuous_on ({0..1} \ {0..1}) (\y. q (fst y) (snd y))" (is "continuous_on ?A ?q") and pf: "\t. t \ {0..1} \ pathfinish(p t) = pathstart(q t)" shows "continuous_on ({0..1} \ {0..1}) (\y. (p(fst y) +++ q(fst y)) (snd y))" proof - have \
: "(\t. p (fst t) (2 * snd t)) = ?p \ (\y. (fst y, 2 * snd y))" "(\t. q (fst t) (2 * snd t - 1)) = ?q \ (\y. (fst y, 2 * snd y - 1))" by force+ show ?thesis unfolding joinpaths_def proof (rule continuous_on_cases_le) show "continuous_on {y \ ?A. snd y \ 1/2} (\t. p (fst t) (2 * snd t))" "continuous_on {y \ ?A. 1/2 \ snd y} (\t. q (fst t) (2 * snd t - 1))" "continuous_on ?A snd" unfolding \
by (rule continuous_intros continuous_on_subset [OF p] continuous_on_subset [OF q] | force)+ qed (use pf in \auto simp: mult.commute pathstart_def pathfinish_def\) qed text\ Congruence properties of homotopy w.r.t. path-combining operations.\ lemma homotopic_paths_reversepath_D: assumes "homotopic_paths S p q" shows "homotopic_paths S (reversepath p) (reversepath q)" using assms apply (simp add: homotopic_paths_def homotopic_with_def, clarify) apply (rule_tac x="h \ (\x. (fst x, 1 - snd x))" in exI) apply (rule conjI continuous_intros)+ apply (auto simp: reversepath_def pathstart_def pathfinish_def elim!: continuous_on_subset) done proposition homotopic_paths_reversepath: "homotopic_paths S (reversepath p) (reversepath q) \ homotopic_paths S p q" using homotopic_paths_reversepath_D by force proposition homotopic_paths_join: "\homotopic_paths S p p'; homotopic_paths S q q'; pathfinish p = pathstart q\ \ homotopic_paths S (p +++ q) (p' +++ q')" apply (clarsimp simp: homotopic_paths_def homotopic_with_def) apply (rename_tac k1 k2) apply (rule_tac x="(\y. ((k1 \ Pair (fst y)) +++ (k2 \ Pair (fst y))) (snd y))" in exI) apply (intro conjI continuous_intros continuous_on_homotopic_join_lemma; force simp: joinpaths_def pathstart_def pathfinish_def path_image_def) done proposition homotopic_paths_continuous_image: "\homotopic_paths S f g; continuous_on S h; h \ S \ t\ \ homotopic_paths t (h \ f) (h \ g)" unfolding homotopic_paths_def by (simp add: homotopic_with_compose_continuous_map_left pathfinish_compose pathstart_compose image_subset_iff_funcset) subsection\Group properties for homotopy of paths\ text\<^marker>\tag important\\So taking equivalence classes under homotopy would give the fundamental group\ proposition homotopic_paths_rid: assumes "path p" "path_image p \ S" shows "homotopic_paths S (p +++ linepath (pathfinish p) (pathfinish p)) p" proof - have \
: "continuous_on {0..1} (\t::real. if t \ 1/2 then 2 *\<^sub>R t else 1)" unfolding split_01 by (rule continuous_on_cases continuous_intros | force simp: pathfinish_def joinpaths_def)+ show ?thesis apply (rule homotopic_paths_sym) using assms unfolding pathfinish_def joinpaths_def by (intro \
continuous_on_cases continuous_intros homotopic_paths_reparametrize [where f = "\t. if t \ 1/2 then 2 *\<^sub>R t else 1"]; force) qed proposition homotopic_paths_lid: "\path p; path_image p \ S\ \ homotopic_paths S (linepath (pathstart p) (pathstart p) +++ p) p" using homotopic_paths_rid [of "reversepath p" S] by (metis homotopic_paths_reversepath path_image_reversepath path_reversepath pathfinish_linepath pathfinish_reversepath reversepath_joinpaths reversepath_linepath) +lemma homotopic_paths_rid': + assumes "path p" "path_image p \ s" "x = pathfinish p" + shows "homotopic_paths s (p +++ linepath x x) p" + using homotopic_paths_rid[of p s] assms by simp + +lemma homotopic_paths_lid': + "\path p; path_image p \ s; x = pathstart p\ \ homotopic_paths s (linepath x x +++ p) p" + using homotopic_paths_lid[of p s] by simp + proposition homotopic_paths_assoc: "\path p; path_image p \ S; path q; path_image q \ S; path r; path_image r \ S; pathfinish p = pathstart q; pathfinish q = pathstart r\ \ homotopic_paths S (p +++ (q +++ r)) ((p +++ q) +++ r)" apply (subst homotopic_paths_sym) apply (rule homotopic_paths_reparametrize [where f = "\t. if t \ 1/2 then inverse 2 *\<^sub>R t else if t \ 3 / 4 then t - (1 / 4) else 2 *\<^sub>R t - 1"]) apply (simp_all del: le_divide_eq_numeral1 add: subset_path_image_join) apply (rule continuous_on_cases_1 continuous_intros | auto simp: joinpaths_def)+ done proposition homotopic_paths_rinv: assumes "path p" "path_image p \ S" shows "homotopic_paths S (p +++ reversepath p) (linepath (pathstart p) (pathstart p))" proof - have p: "continuous_on {0..1} p" using assms by (auto simp: path_def) let ?A = "{0..1} \ {0..1}" have "continuous_on ?A (\x. (subpath 0 (fst x) p +++ reversepath (subpath 0 (fst x) p)) (snd x))" unfolding joinpaths_def subpath_def reversepath_def path_def add_0_right diff_0_right proof (rule continuous_on_cases_le) show "continuous_on {x \ ?A. snd x \ 1/2} (\t. p (fst t * (2 * snd t)))" "continuous_on {x \ ?A. 1/2 \ snd x} (\t. p (fst t * (1 - (2 * snd t - 1))))" "continuous_on ?A snd" by (intro continuous_on_compose2 [OF p] continuous_intros; auto simp: mult_le_one)+ qed (auto simp: algebra_simps) then show ?thesis using assms apply (subst homotopic_paths_sym_eq) unfolding homotopic_paths_def homotopic_with_def apply (rule_tac x="(\y. (subpath 0 (fst y) p +++ reversepath(subpath 0 (fst y) p)) (snd y))" in exI) apply (force simp: mult_le_one path_defs joinpaths_def subpath_def reversepath_def) done qed proposition homotopic_paths_linv: assumes "path p" "path_image p \ S" shows "homotopic_paths S (reversepath p +++ p) (linepath (pathfinish p) (pathfinish p))" using homotopic_paths_rinv [of "reversepath p" S] assms by simp subsection\Homotopy of loops without requiring preservation of endpoints\ definition\<^marker>\tag important\ homotopic_loops :: "'a::topological_space set \ (real \ 'a) \ (real \ 'a) \ bool" where "homotopic_loops S p q \ homotopic_with_canon (\r. pathfinish r = pathstart r) {0..1} S p q" lemma homotopic_loops: "homotopic_loops S p q \ (\h. continuous_on ({0..1::real} \ {0..1}) h \ image h ({0..1} \ {0..1}) \ S \ (\x \ {0..1}. h(0,x) = p x) \ (\x \ {0..1}. h(1,x) = q x) \ (\t \ {0..1}. pathfinish(h \ Pair t) = pathstart(h \ Pair t)))" by (simp add: homotopic_loops_def pathstart_def pathfinish_def homotopic_with) proposition homotopic_loops_imp_loop: "homotopic_loops S p q \ pathfinish p = pathstart p \ pathfinish q = pathstart q" using homotopic_with_imp_property homotopic_loops_def by blast proposition homotopic_loops_imp_path: "homotopic_loops S p q \ path p \ path q" unfolding homotopic_loops_def path_def using homotopic_with_imp_continuous_maps continuous_map_subtopology_eu by blast proposition homotopic_loops_imp_subset: "homotopic_loops S p q \ path_image p \ S \ path_image q \ S" unfolding homotopic_loops_def path_image_def by (meson continuous_map_subtopology_eu homotopic_with_imp_continuous_maps) proposition homotopic_loops_refl: "homotopic_loops S p p \ path p \ path_image p \ S \ pathfinish p = pathstart p" by (simp add: homotopic_loops_def path_image_def path_def) proposition homotopic_loops_sym: "homotopic_loops S p q \ homotopic_loops S q p" by (simp add: homotopic_loops_def homotopic_with_sym) proposition homotopic_loops_sym_eq: "homotopic_loops S p q \ homotopic_loops S q p" by (metis homotopic_loops_sym) proposition homotopic_loops_trans: "\homotopic_loops S p q; homotopic_loops S q r\ \ homotopic_loops S p r" unfolding homotopic_loops_def by (blast intro: homotopic_with_trans) proposition homotopic_loops_subset: "\homotopic_loops S p q; S \ t\ \ homotopic_loops t p q" by (fastforce simp: homotopic_loops) proposition homotopic_loops_eq: "\path p; path_image p \ S; pathfinish p = pathstart p; \t. t \ {0..1} \ p(t) = q(t)\ \ homotopic_loops S p q" unfolding homotopic_loops_def path_image_def path_def pathstart_def pathfinish_def by (auto intro: homotopic_with_eq [OF homotopic_with_refl [where f = p, THEN iffD2]]) proposition homotopic_loops_continuous_image: "\homotopic_loops S f g; continuous_on S h; h \ S \ t\ \ homotopic_loops t (h \ f) (h \ g)" unfolding homotopic_loops_def by (simp add: homotopic_with_compose_continuous_map_left pathfinish_def pathstart_def image_subset_iff_funcset) subsection\Relations between the two variants of homotopy\ proposition homotopic_paths_imp_homotopic_loops: "\homotopic_paths S p q; pathfinish p = pathstart p; pathfinish q = pathstart p\ \ homotopic_loops S p q" by (auto simp: homotopic_with_def homotopic_paths_def homotopic_loops_def) proposition homotopic_loops_imp_homotopic_paths_null: assumes "homotopic_loops S p (linepath a a)" shows "homotopic_paths S p (linepath (pathstart p) (pathstart p))" proof - have "path p" by (metis assms homotopic_loops_imp_path) have ploop: "pathfinish p = pathstart p" by (metis assms homotopic_loops_imp_loop) have pip: "path_image p \ S" by (metis assms homotopic_loops_imp_subset) let ?A = "{0..1::real} \ {0..1::real}" obtain h where conth: "continuous_on ?A h" and hs: "h \ ?A \ S" and h0[simp]: "\x. x \ {0..1} \ h(0,x) = p x" and h1[simp]: "\x. x \ {0..1} \ h(1,x) = a" and ends: "\t. t \ {0..1} \ pathfinish (h \ Pair t) = pathstart (h \ Pair t)" using assms by (auto simp: homotopic_loops homotopic_with image_subset_iff_funcset) have conth0: "path (\u. h (u, 0))" unfolding path_def proof (rule continuous_on_compose [of _ _ h, unfolded o_def]) show "continuous_on ((\x. (x, 0)) ` {0..1}) h" by (force intro: continuous_on_subset [OF conth]) qed (force intro: continuous_intros) have pih0: "path_image (\u. h (u, 0)) \ S" using hs by (force simp: path_image_def) have c1: "continuous_on ?A (\x. h (fst x * snd x, 0))" proof (rule continuous_on_compose [of _ _ h, unfolded o_def]) show "continuous_on ((\x. (fst x * snd x, 0)) ` ?A) h" by (force simp: mult_le_one intro: continuous_on_subset [OF conth]) qed (force intro: continuous_intros)+ have c2: "continuous_on ?A (\x. h (fst x - fst x * snd x, 0))" proof (rule continuous_on_compose [of _ _ h, unfolded o_def]) show "continuous_on ((\x. (fst x - fst x * snd x, 0)) ` ?A) h" by (auto simp: algebra_simps add_increasing2 mult_left_le intro: continuous_on_subset [OF conth]) qed (force intro: continuous_intros) have [simp]: "\t. \0 \ t \ t \ 1\ \ h (t, 1) = h (t, 0)" using ends by (simp add: pathfinish_def pathstart_def) have adhoc_le: "c * 4 \ 1 + c * (d * 4)" if "\ d * 4 \ 3" "0 \ c" "c \ 1" for c d::real proof - have "c * 3 \ c * (d * 4)" using that less_eq_real_def by auto with \c \ 1\ show ?thesis by fastforce qed have *: "\p x. \path p \ path(reversepath p); path_image p \ S \ path_image(reversepath p) \ S; pathfinish p = pathstart(linepath a a +++ reversepath p) \ pathstart(reversepath p) = a \ pathstart p = x\ \ homotopic_paths S (p +++ linepath a a +++ reversepath p) (linepath x x)" by (metis homotopic_paths_lid homotopic_paths_join homotopic_paths_trans homotopic_paths_sym homotopic_paths_rinv) have 1: "homotopic_paths S p (p +++ linepath (pathfinish p) (pathfinish p))" using \path p\ homotopic_paths_rid homotopic_paths_sym pip by blast moreover have "homotopic_paths S (p +++ linepath (pathfinish p) (pathfinish p)) (linepath (pathstart p) (pathstart p) +++ p +++ linepath (pathfinish p) (pathfinish p))" using homotopic_paths_lid [of "p +++ linepath (pathfinish p) (pathfinish p)" S] by (metis 1 homotopic_paths_imp_path homotopic_paths_imp_subset homotopic_paths_sym pathstart_join) moreover have "homotopic_paths S (linepath (pathstart p) (pathstart p) +++ p +++ linepath (pathfinish p) (pathfinish p)) ((\u. h (u, 0)) +++ linepath a a +++ reversepath (\u. h (u, 0)))" unfolding homotopic_paths_def homotopic_with_def proof (intro exI strip conjI) let ?h = "\y. (subpath 0 (fst y) (\u. h (u, 0)) +++ (\u. h (Pair (fst y) u)) +++ subpath (fst y) 0 (\u. h (u, 0))) (snd y)" have "continuous_on ?A ?h" by (intro continuous_on_homotopic_join_lemma; simp add: path_defs joinpaths_def subpath_def conth c1 c2) moreover have "?h \ ?A \ S" using hs unfolding joinpaths_def subpath_def by (force simp: algebra_simps mult_le_one mult_left_le intro: adhoc_le) ultimately show "continuous_map (prod_topology (top_of_set {0..1}) (top_of_set {0..1})) (top_of_set S) ?h" by (simp add: subpath_reversepath image_subset_iff_funcset) qed (use ploop in \simp_all add: reversepath_def path_defs joinpaths_def o_def subpath_def conth c1 c2\) moreover have "homotopic_paths S ((\u. h (u, 0)) +++ linepath a a +++ reversepath (\u. h (u, 0))) (linepath (pathstart p) (pathstart p))" by (rule *; simp add: pih0 pathstart_def pathfinish_def conth0; simp add: reversepath_def joinpaths_def) ultimately show ?thesis by (blast intro: homotopic_paths_trans) qed proposition homotopic_loops_conjugate: fixes S :: "'a::real_normed_vector set" assumes "path p" "path q" and pip: "path_image p \ S" and piq: "path_image q \ S" and pq: "pathfinish p = pathstart q" and qloop: "pathfinish q = pathstart q" shows "homotopic_loops S (p +++ q +++ reversepath p) q" proof - have contp: "continuous_on {0..1} p" using \path p\ [unfolded path_def] by blast have contq: "continuous_on {0..1} q" using \path q\ [unfolded path_def] by blast let ?A = "{0..1::real} \ {0..1::real}" have c1: "continuous_on ?A (\x. p ((1 - fst x) * snd x + fst x))" proof (rule continuous_on_compose [of _ _ p, unfolded o_def]) show "continuous_on ((\x. (1 - fst x) * snd x + fst x) ` ?A) p" by (auto intro: continuous_on_subset [OF contp] simp: algebra_simps add_increasing2 mult_right_le_one_le sum_le_prod1) qed (force intro: continuous_intros) have c2: "continuous_on ?A (\x. p ((fst x - 1) * snd x + 1))" proof (rule continuous_on_compose [of _ _ p, unfolded o_def]) show "continuous_on ((\x. (fst x - 1) * snd x + 1) ` ?A) p" by (auto intro: continuous_on_subset [OF contp] simp: algebra_simps add_increasing2 mult_left_le_one_le) qed (force intro: continuous_intros) have ps1: "\a b. \b * 2 \ 1; 0 \ b; 0 \ a; a \ 1\ \ p ((1 - a) * (2 * b) + a) \ S" using sum_le_prod1 by (force simp: algebra_simps add_increasing2 mult_left_le intro: pip [unfolded path_image_def, THEN subsetD]) have ps2: "\a b. \\ 4 * b \ 3; b \ 1; 0 \ a; a \ 1\ \ p ((a - 1) * (4 * b - 3) + 1) \ S" apply (rule pip [unfolded path_image_def, THEN subsetD]) apply (rule image_eqI, blast) apply (simp add: algebra_simps) by (metis add_mono affine_ineq linear mult.commute mult.left_neutral mult_right_mono add.commute zero_le_numeral) have qs: "\a b. \4 * b \ 3; \ b * 2 \ 1\ \ q (4 * b - 2) \ S" using path_image_def piq by fastforce have "homotopic_loops S (p +++ q +++ reversepath p) (linepath (pathstart q) (pathstart q) +++ q +++ linepath (pathstart q) (pathstart q))" unfolding homotopic_loops_def homotopic_with_def proof (intro exI strip conjI) let ?h = "(\y. (subpath (fst y) 1 p +++ q +++ subpath 1 (fst y) p) (snd y))" have "continuous_on ?A (\y. q (snd y))" by (force simp: contq intro: continuous_on_compose [of _ _ q, unfolded o_def] continuous_on_id continuous_on_snd) then have "continuous_on ?A ?h" using pq qloop by (intro continuous_on_homotopic_join_lemma) (auto simp: path_defs joinpaths_def subpath_def c1 c2) then show "continuous_map (prod_topology (top_of_set {0..1}) (top_of_set {0..1})) (top_of_set S) ?h" by (auto simp: joinpaths_def subpath_def ps1 ps2 qs) show "?h (1,x) = (linepath (pathstart q) (pathstart q) +++ q +++ linepath (pathstart q) (pathstart q)) x" for x using pq by (simp add: pathfinish_def subpath_refl) qed (auto simp: subpath_reversepath) moreover have "homotopic_loops S (linepath (pathstart q) (pathstart q) +++ q +++ linepath (pathstart q) (pathstart q)) q" proof - have "homotopic_paths S (linepath (pathfinish q) (pathfinish q) +++ q) q" using \path q\ homotopic_paths_lid qloop piq by auto hence 1: "\f. homotopic_paths S f q \ \ homotopic_paths S f (linepath (pathfinish q) (pathfinish q) +++ q)" using homotopic_paths_trans by blast hence "homotopic_paths S (linepath (pathfinish q) (pathfinish q) +++ q +++ linepath (pathfinish q) (pathfinish q)) q" by (smt (verit, best) \path q\ homotopic_paths_imp_path homotopic_paths_imp_subset homotopic_paths_lid homotopic_paths_rid homotopic_paths_trans pathstart_join piq qloop) thus ?thesis by (metis (no_types) qloop homotopic_loops_sym homotopic_paths_imp_homotopic_loops homotopic_paths_imp_pathfinish homotopic_paths_sym) qed ultimately show ?thesis by (blast intro: homotopic_loops_trans) qed lemma homotopic_paths_loop_parts: assumes loops: "homotopic_loops S (p +++ reversepath q) (linepath a a)" and "path q" shows "homotopic_paths S p q" proof - have paths: "homotopic_paths S (p +++ reversepath q) (linepath (pathstart p) (pathstart p))" using homotopic_loops_imp_homotopic_paths_null [OF loops] by simp then have "path p" using \path q\ homotopic_loops_imp_path loops path_join path_join_path_ends path_reversepath by blast show ?thesis proof (cases "pathfinish p = pathfinish q") case True obtain pipq: "path_image p \ S" "path_image q \ S" by (metis Un_subset_iff paths \path p\ \path q\ homotopic_loops_imp_subset homotopic_paths_imp_path loops path_image_join path_image_reversepath path_imp_reversepath path_join_eq) have "homotopic_paths S p (p +++ (linepath (pathfinish p) (pathfinish p)))" using \path p\ \path_image p \ S\ homotopic_paths_rid homotopic_paths_sym by blast moreover have "homotopic_paths S (p +++ (linepath (pathfinish p) (pathfinish p))) (p +++ (reversepath q +++ q))" by (simp add: True \path p\ \path q\ pipq homotopic_paths_join homotopic_paths_linv homotopic_paths_sym) moreover have "homotopic_paths S (p +++ (reversepath q +++ q)) ((p +++ reversepath q) +++ q)" by (simp add: True \path p\ \path q\ homotopic_paths_assoc pipq) moreover have "homotopic_paths S ((p +++ reversepath q) +++ q) (linepath (pathstart p) (pathstart p) +++ q)" by (simp add: \path q\ homotopic_paths_join paths pipq) ultimately show ?thesis by (metis \path q\ homotopic_paths_imp_path homotopic_paths_lid homotopic_paths_trans path_join_path_ends pathfinish_linepath pipq(2)) next case False then show ?thesis using \path q\ homotopic_loops_imp_path loops path_join_path_ends by fastforce qed qed subsection\<^marker>\tag unimportant\\Homotopy of "nearby" function, paths and loops\ lemma homotopic_with_linear: fixes f g :: "_ \ 'b::real_normed_vector" assumes contf: "continuous_on S f" and contg:"continuous_on S g" and sub: "\x. x \ S \ closed_segment (f x) (g x) \ t" shows "homotopic_with_canon (\z. True) S t f g" unfolding homotopic_with_def apply (rule_tac x="\y. ((1 - (fst y)) *\<^sub>R f(snd y) + (fst y) *\<^sub>R g(snd y))" in exI) using sub closed_segment_def by (fastforce intro: continuous_intros continuous_on_subset [OF contf] continuous_on_compose2 [where g=f] continuous_on_subset [OF contg] continuous_on_compose2 [where g=g]) lemma homotopic_paths_linear: fixes g h :: "real \ 'a::real_normed_vector" assumes "path g" "path h" "pathstart h = pathstart g" "pathfinish h = pathfinish g" "\t. t \ {0..1} \ closed_segment (g t) (h t) \ S" shows "homotopic_paths S g h" using assms unfolding path_def apply (simp add: closed_segment_def pathstart_def pathfinish_def homotopic_paths_def homotopic_with_def) apply (rule_tac x="\y. ((1 - (fst y)) *\<^sub>R (g \ snd) y + (fst y) *\<^sub>R (h \ snd) y)" in exI) apply (intro conjI subsetI continuous_intros; force) done lemma homotopic_loops_linear: fixes g h :: "real \ 'a::real_normed_vector" assumes "path g" "path h" "pathfinish g = pathstart g" "pathfinish h = pathstart h" "\t x. t \ {0..1} \ closed_segment (g t) (h t) \ S" shows "homotopic_loops S g h" using assms unfolding path_defs homotopic_loops_def homotopic_with_def apply (rule_tac x="\y. ((1 - (fst y)) *\<^sub>R g(snd y) + (fst y) *\<^sub>R h(snd y))" in exI) by (force simp: closed_segment_def intro!: continuous_intros intro: continuous_on_compose2 [where g=g] continuous_on_compose2 [where g=h]) lemma homotopic_paths_nearby_explicit: assumes \
: "path g" "path h" "pathstart h = pathstart g" "pathfinish h = pathfinish g" and no: "\t x. \t \ {0..1}; x \ S\ \ norm(h t - g t) < norm(g t - x)" shows "homotopic_paths S g h" using homotopic_paths_linear [OF \
] by (metis linorder_not_le no norm_minus_commute segment_bound1 subsetI) lemma homotopic_loops_nearby_explicit: assumes \
: "path g" "path h" "pathfinish g = pathstart g" "pathfinish h = pathstart h" and no: "\t x. \t \ {0..1}; x \ S\ \ norm(h t - g t) < norm(g t - x)" shows "homotopic_loops S g h" using homotopic_loops_linear [OF \
] by (metis linorder_not_le no norm_minus_commute segment_bound1 subsetI) lemma homotopic_nearby_paths: fixes g h :: "real \ 'a::euclidean_space" assumes "path g" "open S" "path_image g \ S" shows "\e. 0 < e \ (\h. path h \ pathstart h = pathstart g \ pathfinish h = pathfinish g \ (\t \ {0..1}. norm(h t - g t) < e) \ homotopic_paths S g h)" proof - obtain e where "e > 0" and e: "\x y. x \ path_image g \ y \ - S \ e \ dist x y" using separate_compact_closed [of "path_image g" "-S"] assms by force show ?thesis using e [unfolded dist_norm] \e > 0\ by (fastforce simp: path_image_def intro!: homotopic_paths_nearby_explicit assms exI) qed lemma homotopic_nearby_loops: fixes g h :: "real \ 'a::euclidean_space" assumes "path g" "open S" "path_image g \ S" "pathfinish g = pathstart g" shows "\e. 0 < e \ (\h. path h \ pathfinish h = pathstart h \ (\t \ {0..1}. norm(h t - g t) < e) \ homotopic_loops S g h)" proof - obtain e where "e > 0" and e: "\x y. x \ path_image g \ y \ - S \ e \ dist x y" using separate_compact_closed [of "path_image g" "-S"] assms by force show ?thesis using e [unfolded dist_norm] \e > 0\ by (fastforce simp: path_image_def intro!: homotopic_loops_nearby_explicit assms exI) qed subsection\ Homotopy and subpaths\ lemma homotopic_join_subpaths1: assumes "path g" and pag: "path_image g \ S" and u: "u \ {0..1}" and v: "v \ {0..1}" and w: "w \ {0..1}" "u \ v" "v \ w" shows "homotopic_paths S (subpath u v g +++ subpath v w g) (subpath u w g)" proof - have 1: "t * 2 \ 1 \ u + t * (v * 2) \ v + t * (u * 2)" for t using affine_ineq \u \ v\ by fastforce have 2: "t * 2 > 1 \ u + (2*t - 1) * v \ v + (2*t - 1) * w" for t by (metis add_mono_thms_linordered_semiring(1) diff_gt_0_iff_gt less_eq_real_def mult.commute mult_right_mono \u \ v\ \v \ w\) have t2: "\t::real. t*2 = 1 \ t = 1/2" by auto have "homotopic_paths (path_image g) (subpath u v g +++ subpath v w g) (subpath u w g)" proof (cases "w = u") case True then show ?thesis by (metis \path g\ homotopic_paths_rinv path_image_subpath_subset path_subpath pathstart_subpath reversepath_subpath subpath_refl u v) next case False let ?f = "\t. if t \ 1/2 then inverse((w - u)) *\<^sub>R (2 * (v - u)) *\<^sub>R t else inverse((w - u)) *\<^sub>R ((v - u) + (w - v) *\<^sub>R (2 *\<^sub>R t - 1))" show ?thesis proof (rule homotopic_paths_sym [OF homotopic_paths_reparametrize [where f = ?f]]) show "path (subpath u w g)" using assms(1) path_subpath u w(1) by blast show "path_image (subpath u w g) \ path_image g" by (meson path_image_subpath_subset u w(1)) show "continuous_on {0..1} ?f" unfolding split_01 by (rule continuous_on_cases continuous_intros | force simp: pathfinish_def joinpaths_def dest!: t2)+ show "?f \ {0..1} \ {0..1}" using False assms by (force simp: field_simps not_le mult_left_mono affine_ineq dest!: 1 2) show "(subpath u v g +++ subpath v w g) t = subpath u w g (?f t)" if "t \ {0..1}" for t using assms unfolding joinpaths_def subpath_def by (auto simp: divide_simps add.commute mult.commute mult.left_commute) qed (use False in auto) qed then show ?thesis by (rule homotopic_paths_subset [OF _ pag]) qed lemma homotopic_join_subpaths2: assumes "homotopic_paths S (subpath u v g +++ subpath v w g) (subpath u w g)" shows "homotopic_paths S (subpath w v g +++ subpath v u g) (subpath w u g)" by (metis assms homotopic_paths_reversepath_D pathfinish_subpath pathstart_subpath reversepath_joinpaths reversepath_subpath) lemma homotopic_join_subpaths3: assumes hom: "homotopic_paths S (subpath u v g +++ subpath v w g) (subpath u w g)" and "path g" and pag: "path_image g \ S" and u: "u \ {0..1}" and v: "v \ {0..1}" and w: "w \ {0..1}" shows "homotopic_paths S (subpath v w g +++ subpath w u g) (subpath v u g)" proof - obtain wvg: "path (subpath w v g)" "path_image (subpath w v g) \ S" and wug: "path (subpath w u g)" "path_image (subpath w u g) \ S" and vug: "path (subpath v u g)" "path_image (subpath v u g) \ S" by (meson \path g\ pag path_image_subpath_subset path_subpath subset_trans u v w) have "homotopic_paths S (subpath u w g +++ subpath w v g) ((subpath u v g +++ subpath v w g) +++ subpath w v g)" by (simp add: hom homotopic_paths_join homotopic_paths_sym wvg) also have "homotopic_paths S \ (subpath u v g +++ subpath v w g +++ subpath w v g)" using wvg vug \path g\ by (metis homotopic_paths_assoc homotopic_paths_sym path_image_subpath_commute path_subpath pathfinish_subpath pathstart_subpath u v w) also have "homotopic_paths S \ (subpath u v g +++ linepath (pathfinish (subpath u v g)) (pathfinish (subpath u v g)))" using wvg vug \path g\ by (metis homotopic_paths_join homotopic_paths_linv homotopic_paths_refl path_image_subpath_commute path_subpath pathfinish_subpath pathstart_join pathstart_subpath reversepath_subpath u v) also have "homotopic_paths S \ (subpath u v g)" using vug \path g\ by (metis homotopic_paths_rid path_image_subpath_commute path_subpath u v) finally have "homotopic_paths S (subpath u w g +++ subpath w v g) (subpath u v g)" . then show ?thesis using homotopic_join_subpaths2 by blast qed proposition homotopic_join_subpaths: "\path g; path_image g \ S; u \ {0..1}; v \ {0..1}; w \ {0..1}\ \ homotopic_paths S (subpath u v g +++ subpath v w g) (subpath u w g)" by (smt (verit, del_insts) homotopic_join_subpaths1 homotopic_join_subpaths2 homotopic_join_subpaths3) text\Relating homotopy of trivial loops to path-connectedness.\ lemma path_component_imp_homotopic_points: assumes "path_component S a b" shows "homotopic_loops S (linepath a a) (linepath b b)" proof - obtain g :: "real \ 'a" where g: "continuous_on {0..1} g" "g \ {0..1} \ S" "g 0 = a" "g 1 = b" using assms by (auto simp: path_defs) then have "continuous_on ({0..1} \ {0..1}) (g \ fst)" by (fastforce intro!: continuous_intros)+ with g show ?thesis by (auto simp: homotopic_loops_def homotopic_with_def path_defs Pi_iff) qed lemma homotopic_loops_imp_path_component_value: "\homotopic_loops S p q; 0 \ t; t \ 1\ \ path_component S (p t) (q t)" apply (clarsimp simp: homotopic_loops_def homotopic_with_def path_defs) apply (rule_tac x="h \ (\u. (u, t))" in exI) apply (fastforce elim!: continuous_on_subset intro!: continuous_intros) done lemma homotopic_points_eq_path_component: "homotopic_loops S (linepath a a) (linepath b b) \ path_component S a b" using homotopic_loops_imp_path_component_value path_component_imp_homotopic_points by fastforce lemma path_connected_eq_homotopic_points: "path_connected S \ (\a b. a \ S \ b \ S \ homotopic_loops S (linepath a a) (linepath b b))" by (auto simp: path_connected_def path_component_def homotopic_points_eq_path_component) subsection\Simply connected sets\ text\<^marker>\tag important\\defined as "all loops are homotopic (as loops)\ definition\<^marker>\tag important\ simply_connected where "simply_connected S \ \p q. path p \ pathfinish p = pathstart p \ path_image p \ S \ path q \ pathfinish q = pathstart q \ path_image q \ S \ homotopic_loops S p q" lemma simply_connected_empty [iff]: "simply_connected {}" by (simp add: simply_connected_def) lemma simply_connected_imp_path_connected: fixes S :: "_::real_normed_vector set" shows "simply_connected S \ path_connected S" by (simp add: simply_connected_def path_connected_eq_homotopic_points) lemma simply_connected_imp_connected: fixes S :: "_::real_normed_vector set" shows "simply_connected S \ connected S" by (simp add: path_connected_imp_connected simply_connected_imp_path_connected) lemma simply_connected_eq_contractible_loop_any: fixes S :: "_::real_normed_vector set" shows "simply_connected S \ (\p a. path p \ path_image p \ S \ pathfinish p = pathstart p \ a \ S \ homotopic_loops S p (linepath a a))" (is "?lhs = ?rhs") proof assume ?rhs then show ?lhs unfolding simply_connected_def by (metis pathfinish_in_path_image subsetD homotopic_loops_trans homotopic_loops_sym) qed (force simp: simply_connected_def) lemma simply_connected_eq_contractible_loop_some: fixes S :: "_::real_normed_vector set" shows "simply_connected S \ path_connected S \ (\p. path p \ path_image p \ S \ pathfinish p = pathstart p \ (\a. a \ S \ homotopic_loops S p (linepath a a)))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using simply_connected_eq_contractible_loop_any by (blast intro: simply_connected_imp_path_connected) next assume ?rhs then show ?lhs by (meson homotopic_loops_trans path_connected_eq_homotopic_points simply_connected_eq_contractible_loop_any) qed lemma simply_connected_eq_contractible_loop_all: fixes S :: "_::real_normed_vector set" shows "simply_connected S \ S = {} \ (\a \ S. \p. path p \ path_image p \ S \ pathfinish p = pathstart p \ homotopic_loops S p (linepath a a))" by (meson ex_in_conv homotopic_loops_sym homotopic_loops_trans simply_connected_def simply_connected_eq_contractible_loop_any) lemma simply_connected_eq_contractible_path: fixes S :: "_::real_normed_vector set" shows "simply_connected S \ path_connected S \ (\p. path p \ path_image p \ S \ pathfinish p = pathstart p \ homotopic_paths S p (linepath (pathstart p) (pathstart p)))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding simply_connected_imp_path_connected by (metis simply_connected_eq_contractible_loop_some homotopic_loops_imp_homotopic_paths_null) next assume ?rhs then show ?lhs using homotopic_paths_imp_homotopic_loops simply_connected_eq_contractible_loop_some by fastforce qed lemma simply_connected_eq_homotopic_paths: fixes S :: "_::real_normed_vector set" shows "simply_connected S \ path_connected S \ (\p q. path p \ path_image p \ S \ path q \ path_image q \ S \ pathstart q = pathstart p \ pathfinish q = pathfinish p \ homotopic_paths S p q)" (is "?lhs = ?rhs") proof assume ?lhs then have pc: "path_connected S" and *: "\p. \path p; path_image p \ S; pathfinish p = pathstart p\ \ homotopic_paths S p (linepath (pathstart p) (pathstart p))" by (auto simp: simply_connected_eq_contractible_path) have "homotopic_paths S p q" if "path p" "path_image p \ S" "path q" "path_image q \ S" "pathstart q = pathstart p" "pathfinish q = pathfinish p" for p q proof - have "homotopic_paths S p (p +++ reversepath q +++ q)" using that by (smt (verit, best) homotopic_paths_join homotopic_paths_linv homotopic_paths_rid homotopic_paths_sym homotopic_paths_trans pathstart_linepath) also have "homotopic_paths S \ ((p +++ reversepath q) +++ q)" by (simp add: that homotopic_paths_assoc) also have "homotopic_paths S \ (linepath (pathstart q) (pathstart q) +++ q)" using * [of "p +++ reversepath q"] that by (simp add: homotopic_paths_assoc homotopic_paths_join path_image_join) also have "homotopic_paths S \ q" using that homotopic_paths_lid by blast finally show ?thesis . qed then show ?rhs by (blast intro: pc *) next assume ?rhs then show ?lhs by (force simp: simply_connected_eq_contractible_path) qed proposition simply_connected_Times: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" assumes S: "simply_connected S" and T: "simply_connected T" shows "simply_connected(S \ T)" proof - have "homotopic_loops (S \ T) p (linepath (a, b) (a, b))" if "path p" "path_image p \ S \ T" "p 1 = p 0" "a \ S" "b \ T" for p a b proof - have "path (fst \ p)" by (simp add: continuous_on_fst Path_Connected.path_continuous_image [OF \path p\]) moreover have "path_image (fst \ p) \ S" using that by (force simp: path_image_def) ultimately have p1: "homotopic_loops S (fst \ p) (linepath a a)" using S that by (simp add: simply_connected_eq_contractible_loop_any pathfinish_def pathstart_def) have "path (snd \ p)" by (simp add: continuous_on_snd Path_Connected.path_continuous_image [OF \path p\]) moreover have "path_image (snd \ p) \ T" using that by (force simp: path_image_def) ultimately have p2: "homotopic_loops T (snd \ p) (linepath b b)" using T that by (simp add: simply_connected_eq_contractible_loop_any pathfinish_def pathstart_def) show ?thesis using p1 p2 unfolding homotopic_loops apply clarify subgoal for h k by (rule_tac x="\z. (h z, k z)" in exI) (force intro: continuous_intros simp: path_defs) done qed with assms show ?thesis by (simp add: simply_connected_eq_contractible_loop_any pathfinish_def pathstart_def) qed subsection\Contractible sets\ definition\<^marker>\tag important\ contractible where "contractible S \ \a. homotopic_with_canon (\x. True) S S id (\x. a)" proposition contractible_imp_simply_connected: fixes S :: "_::real_normed_vector set" assumes "contractible S" shows "simply_connected S" proof (cases "S = {}") case True then show ?thesis by force next case False obtain a where a: "homotopic_with_canon (\x. True) S S id (\x. a)" using assms by (force simp: contractible_def) then have "a \ S" using False homotopic_with_imp_funspace2 by fastforce have "\p. path p \ path_image p \ S \ pathfinish p = pathstart p \ homotopic_loops S p (linepath a a)" using a apply (clarsimp simp: homotopic_loops_def homotopic_with_def path_defs) apply (rule_tac x="(h \ (\y. (fst y, (p \ snd) y)))" in exI) apply (intro conjI continuous_on_compose continuous_intros; force elim: continuous_on_subset) done with \a \ S\ show ?thesis by (auto simp: simply_connected_eq_contractible_loop_all False) qed corollary contractible_imp_connected: fixes S :: "_::real_normed_vector set" shows "contractible S \ connected S" by (simp add: contractible_imp_simply_connected simply_connected_imp_connected) lemma contractible_imp_path_connected: fixes S :: "_::real_normed_vector set" shows "contractible S \ path_connected S" by (simp add: contractible_imp_simply_connected simply_connected_imp_path_connected) lemma nullhomotopic_through_contractible: fixes S :: "_::topological_space set" assumes f: "continuous_on S f" "f \ S \ T" and g: "continuous_on T g" "g \ T \ U" and T: "contractible T" obtains c where "homotopic_with_canon (\h. True) S U (g \ f) (\x. c)" proof - obtain b where b: "homotopic_with_canon (\x. True) T T id (\x. b)" using assms by (force simp: contractible_def) have "homotopic_with_canon (\f. True) T U (g \ id) (g \ (\x. b))" by (metis b continuous_map_subtopology_eu g homotopic_with_compose_continuous_map_left image_subset_iff_funcset) then have "homotopic_with_canon (\f. True) S U (g \ id \ f) (g \ (\x. b) \ f)" by (simp add: f homotopic_with_compose_continuous_map_right image_subset_iff_funcset) then show ?thesis by (simp add: comp_def that) qed lemma nullhomotopic_into_contractible: assumes f: "continuous_on S f" "f \ S \ T" and T: "contractible T" obtains c where "homotopic_with_canon (\h. True) S T f (\x. c)" by (rule nullhomotopic_through_contractible [OF f, of id T]) (use assms in auto) lemma nullhomotopic_from_contractible: assumes f: "continuous_on S f" "f \ S \ T" and S: "contractible S" obtains c where "homotopic_with_canon (\h. True) S T f (\x. c)" by (auto simp: comp_def intro: nullhomotopic_through_contractible [OF continuous_on_id _ f S]) lemma homotopic_through_contractible: fixes S :: "_::real_normed_vector set" assumes "continuous_on S f1" "f1 \ S \ T" "continuous_on T g1" "g1 \ T \ U" "continuous_on S f2" "f2 \ S \ T" "continuous_on T g2" "g2 \ T \ U" "contractible T" "path_connected U" shows "homotopic_with_canon (\h. True) S U (g1 \ f1) (g2 \ f2)" proof - obtain c1 where c1: "homotopic_with_canon (\h. True) S U (g1 \ f1) (\x. c1)" by (rule nullhomotopic_through_contractible [of S f1 T g1 U]) (use assms in auto) obtain c2 where c2: "homotopic_with_canon (\h. True) S U (g2 \ f2) (\x. c2)" by (rule nullhomotopic_through_contractible [of S f2 T g2 U]) (use assms in auto) have "S = {} \ (\t. path_connected t \ t \ U \ c2 \ t \ c1 \ t)" proof (cases "S = {}") case True then show ?thesis by force next case False with c1 c2 have "c1 \ U" "c2 \ U" using homotopic_with_imp_continuous_maps by (metis PiE equals0I homotopic_with_imp_funspace2)+ with \path_connected U\ show ?thesis by blast qed then have "homotopic_with_canon (\h. True) S U (\x. c2) (\x. c1)" by (auto simp: path_component homotopic_constant_maps) then show ?thesis using c1 c2 homotopic_with_symD homotopic_with_trans by blast qed lemma homotopic_into_contractible: fixes S :: "'a::real_normed_vector set" and T:: "'b::real_normed_vector set" assumes f: "continuous_on S f" "f \ S \ T" and g: "continuous_on S g" "g \ S \ T" and T: "contractible T" shows "homotopic_with_canon (\h. True) S T f g" using homotopic_through_contractible [of S f T id T g id] by (simp add: assms contractible_imp_path_connected) lemma homotopic_from_contractible: fixes S :: "'a::real_normed_vector set" and T:: "'b::real_normed_vector set" assumes f: "continuous_on S f" "f \ S \ T" and g: "continuous_on S g" "g \ S \ T" and "contractible S" "path_connected T" shows "homotopic_with_canon (\h. True) S T f g" using homotopic_through_contractible [of S id S f T id g] by (simp add: assms contractible_imp_path_connected) subsection\Starlike sets\ definition\<^marker>\tag important\ "starlike S \ (\a\S. \x\S. closed_segment a x \ S)" lemma starlike_UNIV [simp]: "starlike UNIV" by (simp add: starlike_def) lemma convex_imp_starlike: "convex S \ S \ {} \ starlike S" unfolding convex_contains_segment starlike_def by auto lemma starlike_convex_tweak_boundary_points: fixes S :: "'a::euclidean_space set" assumes "convex S" "S \ {}" and ST: "rel_interior S \ T" and TS: "T \ closure S" shows "starlike T" proof - have "rel_interior S \ {}" by (simp add: assms rel_interior_eq_empty) with ST obtain a where a: "a \ rel_interior S" and "a \ T" by blast have "\x. x \ T \ open_segment a x \ rel_interior S" by (rule rel_interior_closure_convex_segment [OF \convex S\ a]) (use assms in auto) then have "\x\T. a \ T \ open_segment a x \ T" using ST by (blast intro: a \a \ T\ rel_interior_closure_convex_segment [OF \convex S\ a]) then show ?thesis unfolding starlike_def using bexI [OF _ \a \ T\] by (simp add: closed_segment_eq_open) qed lemma starlike_imp_contractible_gen: fixes S :: "'a::real_normed_vector set" assumes S: "starlike S" and P: "\a T. \a \ S; 0 \ T; T \ 1\ \ P(\x. (1 - T) *\<^sub>R x + T *\<^sub>R a)" obtains a where "homotopic_with_canon P S S (\x. x) (\x. a)" proof - obtain a where "a \ S" and a: "\x. x \ S \ closed_segment a x \ S" using S by (auto simp: starlike_def) have "\t b. 0 \ t \ t \ 1 \ \u. (1 - t) *\<^sub>R b + t *\<^sub>R a = (1 - u) *\<^sub>R a + u *\<^sub>R b \ 0 \ u \ u \ 1" by (metis add_diff_cancel_right' diff_ge_0_iff_ge le_add_diff_inverse pth_c(1)) then have "(\y. (1 - fst y) *\<^sub>R snd y + fst y *\<^sub>R a) ` ({0..1} \ S) \ S" using a [unfolded closed_segment_def] by force then have "homotopic_with_canon P S S (\x. x) (\x. a)" using \a \ S\ unfolding homotopic_with_def apply (rule_tac x="\y. (1 - (fst y)) *\<^sub>R snd y + (fst y) *\<^sub>R a" in exI) apply (force simp: P intro: continuous_intros) done then show ?thesis using that by blast qed lemma starlike_imp_contractible: fixes S :: "'a::real_normed_vector set" shows "starlike S \ contractible S" using starlike_imp_contractible_gen contractible_def by (fastforce simp: id_def) lemma contractible_UNIV [simp]: "contractible (UNIV :: 'a::real_normed_vector set)" by (simp add: starlike_imp_contractible) lemma starlike_imp_simply_connected: fixes S :: "'a::real_normed_vector set" shows "starlike S \ simply_connected S" by (simp add: contractible_imp_simply_connected starlike_imp_contractible) lemma convex_imp_simply_connected: fixes S :: "'a::real_normed_vector set" shows "convex S \ simply_connected S" using convex_imp_starlike starlike_imp_simply_connected by blast lemma starlike_imp_path_connected: fixes S :: "'a::real_normed_vector set" shows "starlike S \ path_connected S" by (simp add: simply_connected_imp_path_connected starlike_imp_simply_connected) lemma starlike_imp_connected: fixes S :: "'a::real_normed_vector set" shows "starlike S \ connected S" by (simp add: path_connected_imp_connected starlike_imp_path_connected) lemma is_interval_simply_connected_1: fixes S :: "real set" shows "is_interval S \ simply_connected S" by (meson convex_imp_simply_connected is_interval_connected_1 is_interval_convex_1 simply_connected_imp_connected) lemma contractible_empty [simp]: "contractible {}" by (simp add: contractible_def homotopic_on_emptyI) lemma contractible_convex_tweak_boundary_points: fixes S :: "'a::euclidean_space set" assumes "convex S" and TS: "rel_interior S \ T" "T \ closure S" shows "contractible T" by (metis assms closure_eq_empty contractible_empty empty_subsetI starlike_convex_tweak_boundary_points starlike_imp_contractible subset_antisym) lemma convex_imp_contractible: fixes S :: "'a::real_normed_vector set" shows "convex S \ contractible S" using contractible_empty convex_imp_starlike starlike_imp_contractible by blast lemma contractible_sing [simp]: fixes a :: "'a::real_normed_vector" shows "contractible {a}" by (rule convex_imp_contractible [OF convex_singleton]) lemma is_interval_contractible_1: fixes S :: "real set" shows "is_interval S \ contractible S" using contractible_imp_simply_connected convex_imp_contractible is_interval_convex_1 is_interval_simply_connected_1 by auto lemma contractible_Times: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes S: "contractible S" and T: "contractible T" shows "contractible (S \ T)" proof - obtain a h where conth: "continuous_on ({0..1} \ S) h" and hsub: "h \ ({0..1} \ S) \ S" and [simp]: "\x. x \ S \ h (0, x) = x" and [simp]: "\x. x \ S \ h (1::real, x) = a" using S by (force simp: contractible_def homotopic_with) obtain b k where contk: "continuous_on ({0..1} \ T) k" and ksub: "k \ ({0..1} \ T) \ T" and [simp]: "\x. x \ T \ k (0, x) = x" and [simp]: "\x. x \ T \ k (1::real, x) = b" using T by (force simp: contractible_def homotopic_with) show ?thesis apply (simp add: contractible_def homotopic_with) apply (rule exI [where x=a]) apply (rule exI [where x=b]) apply (rule exI [where x = "\z. (h (fst z, fst(snd z)), k (fst z, snd(snd z)))"]) using hsub ksub apply (fastforce intro!: continuous_intros continuous_on_compose2 [OF conth] continuous_on_compose2 [OF contk]) done qed subsection\Local versions of topological properties in general\ definition\<^marker>\tag important\ locally :: "('a::topological_space set \ bool) \ 'a set \ bool" where "locally P S \ \w x. openin (top_of_set S) w \ x \ w \ (\U V. openin (top_of_set S) U \ P V \ x \ U \ U \ V \ V \ w)" lemma locallyI: assumes "\w x. \openin (top_of_set S) w; x \ w\ \ \U V. openin (top_of_set S) U \ P V \ x \ U \ U \ V \ V \ w" shows "locally P S" using assms by (force simp: locally_def) lemma locallyE: assumes "locally P S" "openin (top_of_set S) w" "x \ w" obtains U V where "openin (top_of_set S) U" "P V" "x \ U" "U \ V" "V \ w" using assms unfolding locally_def by meson lemma locally_mono: assumes "locally P S" "\T. P T \ Q T" shows "locally Q S" by (metis assms locally_def) lemma locally_open_subset: assumes "locally P S" "openin (top_of_set S) t" shows "locally P t" by (smt (verit, ccfv_SIG) assms order.trans locally_def openin_imp_subset openin_subset_trans openin_trans) lemma locally_diff_closed: "\locally P S; closedin (top_of_set S) t\ \ locally P (S - t)" using locally_open_subset closedin_def by fastforce lemma locally_empty [iff]: "locally P {}" by (simp add: locally_def openin_subtopology) lemma locally_singleton [iff]: fixes a :: "'a::metric_space" shows "locally P {a} \ P {a}" proof - have "\x::real. \ 0 < x \ P {a}" using zero_less_one by blast then show ?thesis unfolding locally_def by (auto simp: openin_euclidean_subtopology_iff subset_singleton_iff conj_disj_distribR) qed lemma locally_iff: "locally P S \ (\T x. open T \ x \ S \ T \ (\U. open U \ (\V. P V \ x \ S \ U \ S \ U \ V \ V \ S \ T)))" by (smt (verit) locally_def openin_open) lemma locally_Int: assumes S: "locally P S" and T: "locally P T" and P: "\S T. P S \ P T \ P(S \ T)" shows "locally P (S \ T)" unfolding locally_iff proof clarify fix A x assume "open A" "x \ A" "x \ S" "x \ T" then obtain U1 V1 U2 V2 where "open U1" "P V1" "x \ S \ U1" "S \ U1 \ V1 \ V1 \ S \ A" "open U2" "P V2" "x \ T \ U2" "T \ U2 \ V2 \ V2 \ T \ A" using S T unfolding locally_iff by (meson IntI) then have "S \ T \ (U1 \ U2) \ V1 \ V2" "V1 \ V2 \ S \ T \ A" "x \ S \ T \ (U1 \ U2)" by blast+ moreover have "P (V1 \ V2)" by (simp add: P \P V1\ \P V2\) ultimately show "\U. open U \ (\V. P V \ x \ S \ T \ U \ S \ T \ U \ V \ V \ S \ T \ A)" using \open U1\ \open U2\ by blast qed lemma locally_Times: fixes S :: "('a::metric_space) set" and T :: "('b::metric_space) set" assumes PS: "locally P S" and QT: "locally Q T" and R: "\S T. P S \ Q T \ R(S \ T)" shows "locally R (S \ T)" unfolding locally_def proof (clarify) fix W x y assume W: "openin (top_of_set (S \ T)) W" and xy: "(x, y) \ W" then obtain U V where "openin (top_of_set S) U" "x \ U" "openin (top_of_set T) V" "y \ V" "U \ V \ W" using Times_in_interior_subtopology by metis then obtain U1 U2 V1 V2 where opeS: "openin (top_of_set S) U1 \ P U2 \ x \ U1 \ U1 \ U2 \ U2 \ U" and opeT: "openin (top_of_set T) V1 \ Q V2 \ y \ V1 \ V1 \ V2 \ V2 \ V" by (meson PS QT locallyE) then have "openin (top_of_set (S \ T)) (U1 \ V1)" by (simp add: openin_Times) moreover have "R (U2 \ V2)" by (simp add: R opeS opeT) moreover have "U1 \ V1 \ U2 \ V2 \ U2 \ V2 \ W" using opeS opeT \U \ V \ W\ by auto ultimately show "\U V. openin (top_of_set (S \ T)) U \ R V \ (x,y) \ U \ U \ V \ V \ W" using opeS opeT by auto qed proposition homeomorphism_locally_imp: fixes S :: "'a::metric_space set" and T :: "'b::t2_space set" assumes S: "locally P S" and hom: "homeomorphism S T f g" and Q: "\S S'. \P S; homeomorphism S S' f g\ \ Q S'" shows "locally Q T" proof (clarsimp simp: locally_def) fix W y assume "y \ W" and "openin (top_of_set T) W" then obtain A where T: "open A" "W = T \ A" by (force simp: openin_open) then have "W \ T" by auto have f: "\x. x \ S \ g(f x) = x" "f ` S = T" "continuous_on S f" and g: "\y. y \ T \ f(g y) = y" "g ` T = S" "continuous_on T g" using hom by (auto simp: homeomorphism_def) have gw: "g ` W = S \ f -` W" using \W \ T\ g by force have "openin (top_of_set S) (g ` W)" using \openin (top_of_set T) W\ continuous_on_open f gw by auto then obtain U V where osu: "openin (top_of_set S) U" and uv: "P V" "g y \ U" "U \ V" "V \ g ` W" by (metis S \y \ W\ image_eqI locallyE) have "V \ S" using uv by (simp add: gw) have fv: "f ` V = T \ {x. g x \ V}" using \f ` S = T\ f \V \ S\ by auto have contvf: "continuous_on V f" using \V \ S\ continuous_on_subset f(3) by blast have "openin (top_of_set (g ` T)) U" using \g ` T = S\ by (simp add: osu) then have "openin (top_of_set T) (T \ g -` U)" using \continuous_on T g\ continuous_on_open [THEN iffD1] by blast moreover have "\V. Q V \ y \ (T \ g -` U) \ (T \ g -` U) \ V \ V \ W" proof (intro exI conjI) show "f ` V \ W" using uv using Int_lower2 gw image_subsetI mem_Collect_eq subset_iff by auto then have contvg: "continuous_on (f ` V) g" using \W \ T\ continuous_on_subset [OF g(3)] by blast have "V \ g ` f ` V" by (metis \V \ S\ hom homeomorphism_def homeomorphism_of_subsets order_refl) then have homv: "homeomorphism V (f ` V) f g" using \V \ S\ f by (auto simp: homeomorphism_def contvf contvg) show "Q (f ` V)" using Q homv \P V\ by blast show "y \ T \ g -` U" using T(2) \y \ W\ \g y \ U\ by blast show "T \ g -` U \ f ` V" using g(1) image_iff uv(3) by fastforce qed ultimately show "\U. openin (top_of_set T) U \ (\v. Q v \ y \ U \ U \ v \ v \ W)" by meson qed lemma homeomorphism_locally: fixes f:: "'a::metric_space \ 'b::metric_space" assumes "homeomorphism S T f g" and "\S T. homeomorphism S T f g \ (P S \ Q T)" shows "locally P S \ locally Q T" by (smt (verit) assms homeomorphism_locally_imp homeomorphism_symD) lemma homeomorphic_locally: fixes S:: "'a::metric_space set" and T:: "'b::metric_space set" assumes hom: "S homeomorphic T" and iff: "\X Y. X homeomorphic Y \ (P X \ Q Y)" shows "locally P S \ locally Q T" by (smt (verit, ccfv_SIG) hom homeomorphic_def homeomorphism_locally homeomorphism_locally_imp iff) lemma homeomorphic_local_compactness: fixes S:: "'a::metric_space set" and T:: "'b::metric_space set" shows "S homeomorphic T \ locally compact S \ locally compact T" by (simp add: homeomorphic_compactness homeomorphic_locally) lemma locally_translation: fixes P :: "'a :: real_normed_vector set \ bool" shows "(\S. P ((+) a ` S) = P S) \ locally P ((+) a ` S) = locally P S" using homeomorphism_locally [OF homeomorphism_translation] by (metis (full_types) homeomorphism_image2) lemma locally_injective_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "linear f" "inj f" and iff: "\S. P (f ` S) \ Q S" shows "locally P (f ` S) \ locally Q S" by (smt (verit) f homeomorphism_image2 homeomorphism_locally iff linear_homeomorphism_image) lemma locally_open_map_image: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes P: "locally P S" and f: "continuous_on S f" and oo: "\T. openin (top_of_set S) T \ openin (top_of_set (f ` S)) (f ` T)" and Q: "\T. \T \ S; P T\ \ Q(f ` T)" shows "locally Q (f ` S)" proof (clarsimp simp: locally_def) fix W y assume oiw: "openin (top_of_set (f ` S)) W" and "y \ W" then have "W \ f ` S" by (simp add: openin_euclidean_subtopology_iff) have oivf: "openin (top_of_set S) (S \ f -` W)" by (rule continuous_on_open [THEN iffD1, rule_format, OF f oiw]) then obtain x where "x \ S" "f x = y" using \W \ f ` S\ \y \ W\ by blast then obtain U V where "openin (top_of_set S) U" "P V" "x \ U" "U \ V" "V \ S \ f -` W" by (metis IntI P \y \ W\ locallyE oivf vimageI) then have "openin (top_of_set (f ` S)) (f ` U)" by (simp add: oo) then show "\X. openin (top_of_set (f ` S)) X \ (\Y. Q Y \ y \ X \ X \ Y \ Y \ W)" using Q \P V\ \U \ V\ \V \ S \ f -` W\ \f x = y\ \x \ U\ by blast qed subsection\An induction principle for connected sets\ proposition connected_induction: assumes "connected S" and opD: "\T a. \openin (top_of_set S) T; a \ T\ \ \z. z \ T \ P z" and opI: "\a. a \ S \ \T. openin (top_of_set S) T \ a \ T \ (\x \ T. \y \ T. P x \ P y \ Q x \ Q y)" and etc: "a \ S" "b \ S" "P a" "P b" "Q a" shows "Q b" proof - let ?A = "{b. \T. openin (top_of_set S) T \ b \ T \ (\x\T. P x \ Q x)}" let ?B = "{b. \T. openin (top_of_set S) T \ b \ T \ (\x\T. P x \ \ Q x)}" have "?A \ ?B = {}" by (clarsimp simp: set_eq_iff) (metis (no_types, opaque_lifting) Int_iff opD openin_Int) moreover have "S \ ?A \ ?B" by clarsimp (meson opI) moreover have "openin (top_of_set S) ?A" by (subst openin_subopen, blast) moreover have "openin (top_of_set S) ?B" by (subst openin_subopen, blast) ultimately have "?A = {} \ ?B = {}" by (metis (no_types, lifting) \connected S\ connected_openin) then show ?thesis by clarsimp (meson opI etc) qed lemma connected_equivalence_relation_gen: assumes "connected S" and etc: "a \ S" "b \ S" "P a" "P b" and trans: "\x y z. \R x y; R y z\ \ R x z" and opD: "\T a. \openin (top_of_set S) T; a \ T\ \ \z. z \ T \ P z" and opI: "\a. a \ S \ \T. openin (top_of_set S) T \ a \ T \ (\x \ T. \y \ T. P x \ P y \ R x y)" shows "R a b" proof - have "\a b c. \a \ S; P a; b \ S; c \ S; P b; P c; R a b\ \ R a c" apply (rule connected_induction [OF \connected S\ opD], simp_all) by (meson trans opI) then show ?thesis by (metis etc opI) qed lemma connected_induction_simple: assumes "connected S" and etc: "a \ S" "b \ S" "P a" and opI: "\a. a \ S \ \T. openin (top_of_set S) T \ a \ T \ (\x \ T. \y \ T. P x \ P y)" shows "P b" by (rule connected_induction [OF \connected S\ _, where P = "\x. True"]) (use opI etc in auto) lemma connected_equivalence_relation: assumes "connected S" and etc: "a \ S" "b \ S" and sym: "\x y. \R x y; x \ S; y \ S\ \ R y x" and trans: "\x y z. \R x y; R y z; x \ S; y \ S; z \ S\ \ R x z" and opI: "\a. a \ S \ \T. openin (top_of_set S) T \ a \ T \ (\x \ T. R a x)" shows "R a b" proof - have "\a b c. \a \ S; b \ S; c \ S; R a b\ \ R a c" by (smt (verit, ccfv_threshold) connected_induction_simple [OF \connected S\] assms openin_imp_subset subset_eq) then show ?thesis by (metis etc opI) qed lemma locally_constant_imp_constant: assumes "connected S" and opI: "\a. a \ S \ \T. openin (top_of_set S) T \ a \ T \ (\x \ T. f x = f a)" shows "f constant_on S" proof - have "\x y. x \ S \ y \ S \ f x = f y" apply (rule connected_equivalence_relation [OF \connected S\], simp_all) by (metis opI) then show ?thesis by (metis constant_on_def) qed lemma locally_constant: assumes "connected S" shows "locally (\U. f constant_on U) S \ f constant_on S" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (smt (verit, del_insts) assms constant_on_def locally_constant_imp_constant locally_def openin_subtopology_self subset_iff) next assume ?rhs then show ?lhs by (metis constant_on_subset locallyI openin_imp_subset order_refl) qed subsection\Basic properties of local compactness\ proposition locally_compact: fixes S :: "'a :: metric_space set" shows "locally compact S \ (\x \ S. \u v. x \ u \ u \ v \ v \ S \ openin (top_of_set S) u \ compact v)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (meson locallyE openin_subtopology_self) next assume r [rule_format]: ?rhs have *: "\u v. openin (top_of_set S) u \ compact v \ x \ u \ u \ v \ v \ S \ T" if "open T" "x \ S" "x \ T" for x T proof - obtain U V where uv: "x \ U" "U \ V" "V \ S" "compact V" "openin (top_of_set S) U" using r [OF \x \ S\] by auto obtain e where "e>0" and e: "cball x e \ T" using open_contains_cball \open T\ \x \ T\ by blast show ?thesis apply (rule_tac x="(S \ ball x e) \ U" in exI) apply (rule_tac x="cball x e \ V" in exI) using that \e > 0\ e uv apply auto done qed show ?lhs by (rule locallyI) (metis "*" Int_iff openin_open) qed lemma locally_compactE: fixes S :: "'a :: metric_space set" assumes "locally compact S" obtains u v where "\x. x \ S \ x \ u x \ u x \ v x \ v x \ S \ openin (top_of_set S) (u x) \ compact (v x)" using assms unfolding locally_compact by metis lemma locally_compact_alt: fixes S :: "'a :: heine_borel set" shows "locally compact S \ (\x \ S. \U. x \ U \ openin (top_of_set S) U \ compact(closure U) \ closure U \ S)" by (smt (verit, ccfv_threshold) bounded_subset closure_closed closure_mono closure_subset compact_closure compact_imp_closed order.trans locally_compact) lemma locally_compact_Int_cball: fixes S :: "'a :: heine_borel set" shows "locally compact S \ (\x \ S. \e. 0 < e \ closed(cball x e \ S))" (is "?lhs = ?rhs") proof assume L: ?lhs then have "\x U V e. \U \ V; V \ S; compact V; 0 < e; cball x e \ S \ U\ \ closed (cball x e \ S)" by (metis compact_Int compact_cball compact_imp_closed inf.absorb_iff2 inf.assoc inf.orderE) with L show ?rhs by (meson locally_compactE openin_contains_cball) next assume R: ?rhs show ?lhs unfolding locally_compact proof fix x assume "x \ S" then obtain e where "e>0" and "compact (cball x e \ S)" by (metis Int_commute compact_Int_closed compact_cball inf.right_idem R) moreover have "\y\ball x e \ S. \\>0. cball y \ \ S \ ball x e" by (meson Elementary_Metric_Spaces.open_ball IntD1 le_infI1 open_contains_cball_eq) moreover have "openin (top_of_set S) (ball x e \ S)" by (simp add: inf_commute openin_open_Int) ultimately show "\U V. x \ U \ U \ V \ V \ S \ openin (top_of_set S) U \ compact V" by (metis Int_iff \0 < e\ \x \ S\ ball_subset_cball centre_in_ball inf_commute inf_le1 inf_mono order_refl) qed qed lemma locally_compact_compact: fixes S :: "'a :: heine_borel set" shows "locally compact S \ (\K. K \ S \ compact K \ (\U V. K \ U \ U \ V \ V \ S \ openin (top_of_set S) U \ compact V))" (is "?lhs = ?rhs") proof assume ?lhs then obtain u v where uv: "\x. x \ S \ x \ u x \ u x \ v x \ v x \ S \ openin (top_of_set S) (u x) \ compact (v x)" by (metis locally_compactE) have *: "\U V. K \ U \ U \ V \ V \ S \ openin (top_of_set S) U \ compact V" if "K \ S" "compact K" for K proof - have "\C. (\c\C. openin (top_of_set K) c) \ K \ \C \ \D\C. finite D \ K \ \D" using that by (simp add: compact_eq_openin_cover) moreover have "\c \ (\x. K \ u x) ` K. openin (top_of_set K) c" using that by clarify (metis subsetD inf.absorb_iff2 openin_subset openin_subtopology_Int_subset topspace_euclidean_subtopology uv) moreover have "K \ \((\x. K \ u x) ` K)" using that by clarsimp (meson subsetCE uv) ultimately obtain D where "D \ (\x. K \ u x) ` K" "finite D" "K \ \D" by metis then obtain T where T: "T \ K" "finite T" "K \ \((\x. K \ u x) ` T)" by (metis finite_subset_image) have Tuv: "\(u ` T) \ \(v ` T)" using T that by (force dest!: uv) moreover have "openin (top_of_set S) (\ (u ` T))" using T that uv by fastforce moreover obtain "compact (\ (v ` T))" "\ (v ` T) \ S" by (metis T UN_subset_iff \K \ S\ compact_UN subset_iff uv) ultimately show ?thesis using T by auto qed show ?rhs by (blast intro: *) next assume ?rhs then show ?lhs apply (clarsimp simp: locally_compact) apply (drule_tac x="{x}" in spec, simp) done qed lemma open_imp_locally_compact: fixes S :: "'a :: heine_borel set" assumes "open S" shows "locally compact S" proof - have *: "\U V. x \ U \ U \ V \ V \ S \ openin (top_of_set S) U \ compact V" if "x \ S" for x proof - obtain e where "e>0" and e: "cball x e \ S" using open_contains_cball assms \x \ S\ by blast have ope: "openin (top_of_set S) (ball x e)" by (meson e open_ball ball_subset_cball dual_order.trans open_subset) show ?thesis by (meson \0 < e\ ball_subset_cball centre_in_ball compact_cball e ope) qed show ?thesis unfolding locally_compact by (blast intro: *) qed lemma closed_imp_locally_compact: fixes S :: "'a :: heine_borel set" assumes "closed S" shows "locally compact S" proof - have *: "\U V. x \ U \ U \ V \ V \ S \ openin (top_of_set S) U \ compact V" if "x \ S" for x apply (rule_tac x = "S \ ball x 1" in exI, rule_tac x = "S \ cball x 1" in exI) using \x \ S\ assms by auto show ?thesis unfolding locally_compact by (blast intro: *) qed lemma locally_compact_UNIV: "locally compact (UNIV :: 'a :: heine_borel set)" by (simp add: closed_imp_locally_compact) lemma locally_compact_Int: fixes S :: "'a :: t2_space set" shows "\locally compact S; locally compact T\ \ locally compact (S \ T)" by (simp add: compact_Int locally_Int) lemma locally_compact_closedin: fixes S :: "'a :: heine_borel set" shows "\closedin (top_of_set S) T; locally compact S\ \ locally compact T" unfolding closedin_closed using closed_imp_locally_compact locally_compact_Int by blast lemma locally_compact_delete: fixes S :: "'a :: t1_space set" shows "locally compact S \ locally compact (S - {a})" by (auto simp: openin_delete locally_open_subset) lemma locally_closed: fixes S :: "'a :: heine_borel set" shows "locally closed S \ locally compact S" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding locally_def apply (elim all_forward imp_forward asm_rl exE) apply (rename_tac U V) apply (rule_tac x = "U \ ball x 1" in exI) apply (rule_tac x = "V \ cball x 1" in exI) apply (force intro: openin_trans) done next assume ?rhs then show ?lhs using compact_eq_bounded_closed locally_mono by blast qed lemma locally_compact_openin_Un: fixes S :: "'a::euclidean_space set" assumes LCS: "locally compact S" and LCT: "locally compact T" and opS: "openin (top_of_set (S \ T)) S" and opT: "openin (top_of_set (S \ T)) T" shows "locally compact (S \ T)" proof - have "\e>0. closed (cball x e \ (S \ T))" if "x \ S" for x proof - obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ S)" using LCS \x \ S\ unfolding locally_compact_Int_cball by blast moreover obtain e2 where "e2 > 0" and e2: "cball x e2 \ (S \ T) \ S" by (meson \x \ S\ opS openin_contains_cball) then have "cball x e2 \ (S \ T) = cball x e2 \ S" by force ultimately have "closed (cball x (min e1 e2) \ (S \ T))" by (metis (no_types, lifting) cball_min_Int closed_Int closed_cball inf_assoc inf_commute) then show ?thesis by (metis \0 < e1\ \0 < e2\ min_def) qed moreover have "\e>0. closed (cball x e \ (S \ T))" if "x \ T" for x proof - obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ T)" using LCT \x \ T\ unfolding locally_compact_Int_cball by blast moreover obtain e2 where "e2 > 0" and e2: "cball x e2 \ (S \ T) \ T" by (meson \x \ T\ opT openin_contains_cball) then have "cball x e2 \ (S \ T) = cball x e2 \ T" by force moreover have "closed (cball x e1 \ (cball x e2 \ T))" by (metis closed_Int closed_cball e1 inf_left_commute) ultimately show ?thesis by (rule_tac x="min e1 e2" in exI) (simp add: \0 < e2\ cball_min_Int inf_assoc) qed ultimately show ?thesis by (force simp: locally_compact_Int_cball) qed lemma locally_compact_closedin_Un: fixes S :: "'a::euclidean_space set" assumes LCS: "locally compact S" and LCT:"locally compact T" and clS: "closedin (top_of_set (S \ T)) S" and clT: "closedin (top_of_set (S \ T)) T" shows "locally compact (S \ T)" proof - have "\e>0. closed (cball x e \ (S \ T))" if "x \ S" "x \ T" for x proof - obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ S)" using LCS \x \ S\ unfolding locally_compact_Int_cball by blast moreover obtain e2 where "e2 > 0" and e2: "closed (cball x e2 \ T)" using LCT \x \ T\ unfolding locally_compact_Int_cball by blast moreover have "closed (cball x (min e1 e2) \ (S \ T))" by (smt (verit) Int_Un_distrib2 Int_commute cball_min_Int closed_Int closed_Un closed_cball e1 e2 inf_left_commute) ultimately show ?thesis by (rule_tac x="min e1 e2" in exI) linarith qed moreover have "\e>0. closed (cball x e \ (S \ T))" if x: "x \ S" "x \ T" for x proof - obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ S)" using LCS \x \ S\ unfolding locally_compact_Int_cball by blast moreover obtain e2 where "e2>0" and "cball x e2 \ (S \ T) \ S - T" using clT x by (fastforce simp: openin_contains_cball closedin_def) then have "closed (cball x e2 \ T)" proof - have "{} = T - (T - cball x e2)" using Diff_subset Int_Diff \cball x e2 \ (S \ T) \ S - T\ by auto then show ?thesis by (simp add: Diff_Diff_Int inf_commute) qed with e1 have "closed ((cball x e1 \ cball x e2) \ (S \ T))" apply (simp add: inf_commute inf_sup_distrib2) by (metis closed_Int closed_Un closed_cball inf_assoc inf_left_commute) then have "closed (cball x (min e1 e2) \ (S \ T))" by (simp add: cball_min_Int inf_commute) ultimately show ?thesis using \0 < e2\ by (rule_tac x="min e1 e2" in exI) linarith qed moreover have "\e>0. closed (cball x e \ (S \ T))" if x: "x \ S" "x \ T" for x proof - obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ T)" using LCT \x \ T\ unfolding locally_compact_Int_cball by blast moreover obtain e2 where "e2>0" and "cball x e2 \ (S \ T) \ S \ T - S" using clS x by (fastforce simp: openin_contains_cball closedin_def) then have "closed (cball x e2 \ S)" by (metis Diff_disjoint Int_empty_right closed_empty inf.left_commute inf.orderE inf_sup_absorb) with e1 have "closed ((cball x e1 \ cball x e2) \ (S \ T))" apply (simp add: inf_commute inf_sup_distrib2) by (metis closed_Int closed_Un closed_cball inf_assoc inf_left_commute) then have "closed (cball x (min e1 e2) \ (S \ T))" by (auto simp: cball_min_Int) ultimately show ?thesis using \0 < e2\ by (rule_tac x="min e1 e2" in exI) linarith qed ultimately show ?thesis by (auto simp: locally_compact_Int_cball) qed lemma locally_compact_Times: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" shows "\locally compact S; locally compact T\ \ locally compact (S \ T)" by (auto simp: compact_Times locally_Times) lemma locally_compact_compact_subopen: fixes S :: "'a :: heine_borel set" shows "locally compact S \ (\K T. K \ S \ compact K \ open T \ K \ T \ (\U V. K \ U \ U \ V \ U \ T \ V \ S \ openin (top_of_set S) U \ compact V))" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs proof clarify fix K :: "'a set" and T :: "'a set" assume "K \ S" and "compact K" and "open T" and "K \ T" obtain U V where "K \ U" "U \ V" "V \ S" "compact V" and ope: "openin (top_of_set S) U" using L unfolding locally_compact_compact by (meson \K \ S\ \compact K\) show "\U V. K \ U \ U \ V \ U \ T \ V \ S \ openin (top_of_set S) U \ compact V" proof (intro exI conjI) show "K \ U \ T" by (simp add: \K \ T\ \K \ U\) show "U \ T \ closure(U \ T)" by (rule closure_subset) show "closure (U \ T) \ S" by (metis \U \ V\ \V \ S\ \compact V\ closure_closed closure_mono compact_imp_closed inf.cobounded1 subset_trans) show "openin (top_of_set S) (U \ T)" by (simp add: \open T\ ope openin_Int_open) show "compact (closure (U \ T))" by (meson Int_lower1 \U \ V\ \compact V\ bounded_subset compact_closure compact_eq_bounded_closed) qed auto qed next assume ?rhs then show ?lhs unfolding locally_compact_compact by (metis open_openin openin_topspace subtopology_superset top.extremum topspace_euclidean_subtopology) qed subsection\Sura-Bura's results about compact components of sets\ proposition Sura_Bura_compact: fixes S :: "'a::euclidean_space set" assumes "compact S" and C: "C \ components S" shows "C = \{T. C \ T \ openin (top_of_set S) T \ closedin (top_of_set S) T}" (is "C = \?\") proof obtain x where x: "C = connected_component_set S x" and "x \ S" using C by (auto simp: components_def) have "C \ S" by (simp add: C in_components_subset) have "\?\ \ connected_component_set S x" proof (rule connected_component_maximal) have "x \ C" by (simp add: \x \ S\ x) then show "x \ \?\" by blast have clo: "closed (\?\)" by (simp add: \compact S\ closed_Inter closedin_compact_eq compact_imp_closed) have False if K1: "closedin (top_of_set (\?\)) K1" and K2: "closedin (top_of_set (\?\)) K2" and K12_Int: "K1 \ K2 = {}" and K12_Un: "K1 \ K2 = \?\" and "K1 \ {}" "K2 \ {}" for K1 K2 proof - have "closed K1" "closed K2" using closedin_closed_trans clo K1 K2 by blast+ then obtain V1 V2 where "open V1" "open V2" "K1 \ V1" "K2 \ V2" and V12: "V1 \ V2 = {}" using separation_normal \K1 \ K2 = {}\ by metis have SV12_ne: "(S - (V1 \ V2)) \ (\?\) \ {}" proof (rule compact_imp_fip) show "compact (S - (V1 \ V2))" by (simp add: \open V1\ \open V2\ \compact S\ compact_diff open_Un) show clo\: "closed T" if "T \ ?\" for T using that \compact S\ by (force intro: closedin_closed_trans simp add: compact_imp_closed) show "(S - (V1 \ V2)) \ \\ \ {}" if "finite \" and \: "\ \ ?\" for \ proof assume djo: "(S - (V1 \ V2)) \ \\ = {}" obtain D where opeD: "openin (top_of_set S) D" and cloD: "closedin (top_of_set S) D" and "C \ D" and DV12: "D \ V1 \ V2" proof (cases "\ = {}") case True with \C \ S\ djo that show ?thesis by force next case False show ?thesis proof show ope: "openin (top_of_set S) (\\)" using openin_Inter \finite \\ False \ by blast then show "closedin (top_of_set S) (\\)" by (meson clo\ \ closed_Inter closed_subset openin_imp_subset subset_eq) show "C \ \\" using \ by auto show "\\ \ V1 \ V2" using ope djo openin_imp_subset by fastforce qed qed have "connected C" by (simp add: x) have "closed D" using \compact S\ cloD closedin_closed_trans compact_imp_closed by blast have cloV1: "closedin (top_of_set D) (D \ closure V1)" and cloV2: "closedin (top_of_set D) (D \ closure V2)" by (simp_all add: closedin_closed_Int) moreover have "D \ closure V1 = D \ V1" "D \ closure V2 = D \ V2" using \D \ V1 \ V2\ \open V1\ \open V2\ V12 by (auto simp: closure_subset [THEN subsetD] closure_iff_nhds_not_empty, blast+) ultimately have cloDV1: "closedin (top_of_set D) (D \ V1)" and cloDV2: "closedin (top_of_set D) (D \ V2)" by metis+ then obtain U1 U2 where "closed U1" "closed U2" and D1: "D \ V1 = D \ U1" and D2: "D \ V2 = D \ U2" by (auto simp: closedin_closed) have "D \ U1 \ C \ {}" proof assume "D \ U1 \ C = {}" then have *: "C \ D \ V2" using D1 DV12 \C \ D\ by auto have 1: "openin (top_of_set S) (D \ V2)" by (simp add: \open V2\ opeD openin_Int_open) have 2: "closedin (top_of_set S) (D \ V2)" using cloD cloDV2 closedin_trans by blast have "\ ?\ \ D \ V2" by (rule Inter_lower) (use * 1 2 in simp) then show False using K1 V12 \K1 \ {}\ \K1 \ V1\ closedin_imp_subset by blast qed moreover have "D \ U2 \ C \ {}" proof assume "D \ U2 \ C = {}" then have *: "C \ D \ V1" using D2 DV12 \C \ D\ by auto have 1: "openin (top_of_set S) (D \ V1)" by (simp add: \open V1\ opeD openin_Int_open) have 2: "closedin (top_of_set S) (D \ V1)" using cloD cloDV1 closedin_trans by blast have "\?\ \ D \ V1" by (rule Inter_lower) (use * 1 2 in simp) then show False using K2 V12 \K2 \ {}\ \K2 \ V2\ closedin_imp_subset by blast qed ultimately show False using \connected C\ [unfolded connected_closed, simplified, rule_format, of concl: "D \ U1" "D \ U2"] using \C \ D\ D1 D2 V12 DV12 \closed U1\ \closed U2\ \closed D\ by blast qed qed show False by (metis (full_types) DiffE UnE Un_upper2 SV12_ne \K1 \ V1\ \K2 \ V2\ disjoint_iff_not_equal subsetCE sup_ge1 K12_Un) qed then show "connected (\?\)" by (auto simp: connected_closedin_eq) show "\?\ \ S" by (fastforce simp: C in_components_subset) qed with x show "\?\ \ C" by simp qed auto corollary Sura_Bura_clopen_subset: fixes S :: "'a::euclidean_space set" assumes S: "locally compact S" and C: "C \ components S" and "compact C" and U: "open U" "C \ U" obtains K where "openin (top_of_set S) K" "compact K" "C \ K" "K \ U" proof (rule ccontr) assume "\ thesis" with that have neg: "\K. openin (top_of_set S) K \ compact K \ C \ K \ K \ U" by metis obtain V K where "C \ V" "V \ U" "V \ K" "K \ S" "compact K" and opeSV: "openin (top_of_set S) V" using S U \compact C\ by (meson C in_components_subset locally_compact_compact_subopen) let ?\ = "{T. C \ T \ openin (top_of_set K) T \ compact T \ T \ K}" have CK: "C \ components K" by (meson C \C \ V\ \K \ S\ \V \ K\ components_intermediate_subset subset_trans) with \compact K\ have "C = \{T. C \ T \ openin (top_of_set K) T \ closedin (top_of_set K) T}" by (simp add: Sura_Bura_compact) then have Ceq: "C = \?\" by (simp add: closedin_compact_eq \compact K\) obtain W where "open W" and W: "V = S \ W" using opeSV by (auto simp: openin_open) have "-(U \ W) \ \?\ \ {}" proof (rule closed_imp_fip_compact) show "- (U \ W) \ \\ \ {}" if "finite \" and \: "\ \ ?\" for \ proof (cases "\ = {}") case True have False if "U = UNIV" "W = UNIV" proof - have "V = S" by (simp add: W \W = UNIV\) with neg show False using \C \ V\ \K \ S\ \V \ K\ \V \ U\ \compact K\ by auto qed with True show ?thesis by auto next case False show ?thesis proof assume "- (U \ W) \ \\ = {}" then have FUW: "\\ \ U \ W" by blast have "C \ \\" using \ by auto moreover have "compact (\\)" by (metis (no_types, lifting) compact_Inter False mem_Collect_eq subsetCE \) moreover have "\\ \ K" using False that(2) by fastforce moreover have opeKF: "openin (top_of_set K) (\\)" using False \ \finite \\ by blast then have opeVF: "openin (top_of_set V) (\\)" using W \K \ S\ \V \ K\ opeKF \\\ \ K\ FUW openin_subset_trans by fastforce then have "openin (top_of_set S) (\\)" by (metis opeSV openin_trans) moreover have "\\ \ U" by (meson \V \ U\ opeVF dual_order.trans openin_imp_subset) ultimately show False using neg by blast qed qed qed (use \open W\ \open U\ in auto) with W Ceq \C \ V\ \C \ U\ show False by auto qed corollary Sura_Bura_clopen_subset_alt: fixes S :: "'a::euclidean_space set" assumes S: "locally compact S" and C: "C \ components S" and "compact C" and opeSU: "openin (top_of_set S) U" and "C \ U" obtains K where "openin (top_of_set S) K" "compact K" "C \ K" "K \ U" proof - obtain V where "open V" "U = S \ V" using opeSU by (auto simp: openin_open) with \C \ U\ have "C \ V" by auto then show ?thesis using Sura_Bura_clopen_subset [OF S C \compact C\ \open V\] by (metis \U = S \ V\ inf.bounded_iff openin_imp_subset that) qed corollary Sura_Bura: fixes S :: "'a::euclidean_space set" assumes "locally compact S" "C \ components S" "compact C" shows "C = \ {K. C \ K \ compact K \ openin (top_of_set S) K}" (is "C = ?rhs") proof show "?rhs \ C" proof (clarsimp, rule ccontr) fix x assume *: "\X. C \ X \ compact X \ openin (top_of_set S) X \ x \ X" and "x \ C" obtain U V where "open U" "open V" "{x} \ U" "C \ V" "U \ V = {}" using separation_normal [of "{x}" C] by (metis Int_empty_left \x \ C\ \compact C\ closed_empty closed_insert compact_imp_closed insert_disjoint(1)) have "x \ V" using \U \ V = {}\ \{x} \ U\ by blast then show False by (meson "*" Sura_Bura_clopen_subset \C \ V\ \open V\ assms(1) assms(2) assms(3) subsetCE) qed qed blast subsection\Special cases of local connectedness and path connectedness\ lemma locally_connected_1: assumes "\V x. \openin (top_of_set S) V; x \ V\ \ \U. openin (top_of_set S) U \ connected U \ x \ U \ U \ V" shows "locally connected S" by (metis assms locally_def) lemma locally_connected_2: assumes "locally connected S" "openin (top_of_set S) t" "x \ t" shows "openin (top_of_set S) (connected_component_set t x)" proof - { fix y :: 'a let ?SS = "top_of_set S" assume 1: "openin ?SS t" "\w x. openin ?SS w \ x \ w \ (\u. openin ?SS u \ (\v. connected v \ x \ u \ u \ v \ v \ w))" and "connected_component t x y" then have "y \ t" and y: "y \ connected_component_set t x" using connected_component_subset by blast+ obtain F where "\x y. (\w. openin ?SS w \ (\u. connected u \ x \ w \ w \ u \ u \ y)) = (openin ?SS (F x y) \ (\u. connected u \ x \ F x y \ F x y \ u \ u \ y))" by moura then obtain G where "\a A. (\U. openin ?SS U \ (\V. connected V \ a \ U \ U \ V \ V \ A)) = (openin ?SS (F a A) \ connected (G a A) \ a \ F a A \ F a A \ G a A \ G a A \ A)" by moura then have *: "openin ?SS (F y t) \ connected (G y t) \ y \ F y t \ F y t \ G y t \ G y t \ t" using 1 \y \ t\ by presburger have "G y t \ connected_component_set t y" by (metis (no_types) * connected_component_eq_self connected_component_mono contra_subsetD) then have "\A. openin ?SS A \ y \ A \ A \ connected_component_set t x" by (metis (no_types) * connected_component_eq dual_order.trans y) } then show ?thesis using assms openin_subopen by (force simp: locally_def) qed lemma locally_connected_3: assumes "\t x. \openin (top_of_set S) t; x \ t\ \ openin (top_of_set S) (connected_component_set t x)" "openin (top_of_set S) v" "x \ v" shows "\u. openin (top_of_set S) u \ connected u \ x \ u \ u \ v" using assms connected_component_subset by fastforce lemma locally_connected: "locally connected S \ (\v x. openin (top_of_set S) v \ x \ v \ (\u. openin (top_of_set S) u \ connected u \ x \ u \ u \ v))" by (metis locally_connected_1 locally_connected_2 locally_connected_3) lemma locally_connected_open_connected_component: "locally connected S \ (\t x. openin (top_of_set S) t \ x \ t \ openin (top_of_set S) (connected_component_set t x))" by (metis locally_connected_1 locally_connected_2 locally_connected_3) lemma locally_path_connected_1: assumes "\v x. \openin (top_of_set S) v; x \ v\ \ \u. openin (top_of_set S) u \ path_connected u \ x \ u \ u \ v" shows "locally path_connected S" by (force simp: locally_def dest: assms) lemma locally_path_connected_2: assumes "locally path_connected S" "openin (top_of_set S) t" "x \ t" shows "openin (top_of_set S) (path_component_set t x)" proof - { fix y :: 'a let ?SS = "top_of_set S" assume 1: "openin ?SS t" "\w x. openin ?SS w \ x \ w \ (\u. openin ?SS u \ (\v. path_connected v \ x \ u \ u \ v \ v \ w))" and "path_component t x y" then have "y \ t" and y: "y \ path_component_set t x" using path_component_mem(2) by blast+ obtain F where "\x y. (\w. openin ?SS w \ (\u. path_connected u \ x \ w \ w \ u \ u \ y)) = (openin ?SS (F x y) \ (\u. path_connected u \ x \ F x y \ F x y \ u \ u \ y))" by moura then obtain G where "\a A. (\U. openin ?SS U \ (\V. path_connected V \ a \ U \ U \ V \ V \ A)) = (openin ?SS (F a A) \ path_connected (G a A) \ a \ F a A \ F a A \ G a A \ G a A \ A)" by moura then have *: "openin ?SS (F y t) \ path_connected (G y t) \ y \ F y t \ F y t \ G y t \ G y t \ t" using 1 \y \ t\ by presburger have "G y t \ path_component_set t y" using * path_component_maximal rev_subsetD by blast then have "\A. openin ?SS A \ y \ A \ A \ path_component_set t x" by (metis "*" \G y t \ path_component_set t y\ dual_order.trans path_component_eq y) } then show ?thesis using assms openin_subopen by (force simp: locally_def) qed lemma locally_path_connected_3: assumes "\t x. \openin (top_of_set S) t; x \ t\ \ openin (top_of_set S) (path_component_set t x)" "openin (top_of_set S) v" "x \ v" shows "\u. openin (top_of_set S) u \ path_connected u \ x \ u \ u \ v" proof - have "path_component v x x" by (meson assms(3) path_component_refl) then show ?thesis by (metis assms mem_Collect_eq path_component_subset path_connected_path_component) qed proposition locally_path_connected: "locally path_connected S \ (\V x. openin (top_of_set S) V \ x \ V \ (\U. openin (top_of_set S) U \ path_connected U \ x \ U \ U \ V))" by (metis locally_path_connected_1 locally_path_connected_2 locally_path_connected_3) proposition locally_path_connected_open_path_component: "locally path_connected S \ (\t x. openin (top_of_set S) t \ x \ t \ openin (top_of_set S) (path_component_set t x))" by (metis locally_path_connected_1 locally_path_connected_2 locally_path_connected_3) lemma locally_connected_open_component: "locally connected S \ (\t c. openin (top_of_set S) t \ c \ components t \ openin (top_of_set S) c)" by (metis components_iff locally_connected_open_connected_component) proposition locally_connected_im_kleinen: "locally connected S \ (\v x. openin (top_of_set S) v \ x \ v \ (\u. openin (top_of_set S) u \ x \ u \ u \ v \ (\y. y \ u \ (\c. connected c \ c \ v \ x \ c \ y \ c))))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (fastforce simp: locally_connected) next assume ?rhs have *: "\T. openin (top_of_set S) T \ x \ T \ T \ c" if "openin (top_of_set S) t" and c: "c \ components t" and "x \ c" for t c x proof - from that \?rhs\ [rule_format, of t x] obtain u where u: "openin (top_of_set S) u \ x \ u \ u \ t \ (\y. y \ u \ (\c. connected c \ c \ t \ x \ c \ y \ c))" using in_components_subset by auto obtain F :: "'a set \ 'a set \ 'a" where "\x y. (\z. z \ x \ y = connected_component_set x z) = (F x y \ x \ y = connected_component_set x (F x y))" by moura then have F: "F t c \ t \ c = connected_component_set t (F t c)" by (meson components_iff c) obtain G :: "'a set \ 'a set \ 'a" where G: "\x y. (\z. z \ y \ z \ x) = (G x y \ y \ G x y \ x)" by moura have "G c u \ u \ G c u \ c" using F by (metis (full_types) u connected_componentI connected_component_eq mem_Collect_eq that(3)) then show ?thesis using G u by auto qed show ?lhs unfolding locally_connected_open_component by (meson "*" openin_subopen) qed proposition locally_path_connected_im_kleinen: "locally path_connected S \ (\v x. openin (top_of_set S) v \ x \ v \ (\u. openin (top_of_set S) u \ x \ u \ u \ v \ (\y. y \ u \ (\p. path p \ path_image p \ v \ pathstart p = x \ pathfinish p = y))))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply (simp add: locally_path_connected path_connected_def) apply (erule all_forward ex_forward imp_forward conjE | simp)+ by (meson dual_order.trans) next assume ?rhs have *: "\T. openin (top_of_set S) T \ x \ T \ T \ path_component_set u z" if "openin (top_of_set S) u" and "z \ u" and c: "path_component u z x" for u z x proof - have "x \ u" by (meson c path_component_mem(2)) with that \?rhs\ [rule_format, of u x] obtain U where U: "openin (top_of_set S) U \ x \ U \ U \ u \ (\y. y \ U \ (\p. path p \ path_image p \ u \ pathstart p = x \ pathfinish p = y))" by blast show ?thesis by (metis U c mem_Collect_eq path_component_def path_component_eq subsetI) qed show ?lhs unfolding locally_path_connected_open_path_component using "*" openin_subopen by fastforce qed lemma locally_path_connected_imp_locally_connected: "locally path_connected S \ locally connected S" using locally_mono path_connected_imp_connected by blast lemma locally_connected_components: "\locally connected S; c \ components S\ \ locally connected c" by (meson locally_connected_open_component locally_open_subset openin_subtopology_self) lemma locally_path_connected_components: "\locally path_connected S; c \ components S\ \ locally path_connected c" by (meson locally_connected_open_component locally_open_subset locally_path_connected_imp_locally_connected openin_subtopology_self) lemma locally_path_connected_connected_component: "locally path_connected S \ locally path_connected (connected_component_set S x)" by (metis components_iff connected_component_eq_empty locally_empty locally_path_connected_components) lemma open_imp_locally_path_connected: fixes S :: "'a :: real_normed_vector set" assumes "open S" shows "locally path_connected S" proof (rule locally_mono) show "locally convex S" using assms unfolding locally_def by (meson open_ball centre_in_ball convex_ball openE open_subset openin_imp_subset openin_open_trans subset_trans) show "\T::'a set. convex T \ path_connected T" using convex_imp_path_connected by blast qed lemma open_imp_locally_connected: fixes S :: "'a :: real_normed_vector set" shows "open S \ locally connected S" by (simp add: locally_path_connected_imp_locally_connected open_imp_locally_path_connected) lemma locally_path_connected_UNIV: "locally path_connected (UNIV::'a :: real_normed_vector set)" by (simp add: open_imp_locally_path_connected) lemma locally_connected_UNIV: "locally connected (UNIV::'a :: real_normed_vector set)" by (simp add: open_imp_locally_connected) lemma openin_connected_component_locally_connected: "locally connected S \ openin (top_of_set S) (connected_component_set S x)" by (metis connected_component_eq_empty locally_connected_2 openin_empty openin_subtopology_self) lemma openin_components_locally_connected: "\locally connected S; c \ components S\ \ openin (top_of_set S) c" using locally_connected_open_component openin_subtopology_self by blast lemma openin_path_component_locally_path_connected: "locally path_connected S \ openin (top_of_set S) (path_component_set S x)" by (metis (no_types) empty_iff locally_path_connected_2 openin_subopen openin_subtopology_self path_component_eq_empty) lemma closedin_path_component_locally_path_connected: assumes "locally path_connected S" shows "closedin (top_of_set S) (path_component_set S x)" proof - have "openin (top_of_set S) (\ ({path_component_set S y |y. y \ S} - {path_component_set S x}))" using locally_path_connected_2 assms by fastforce then show ?thesis by (simp add: closedin_def path_component_subset complement_path_component_Union) qed lemma convex_imp_locally_path_connected: fixes S :: "'a:: real_normed_vector set" assumes "convex S" shows "locally path_connected S" proof (clarsimp simp: locally_path_connected) fix V x assume "openin (top_of_set S) V" and "x \ V" then obtain T e where "V = S \ T" "x \ S" "0 < e" "ball x e \ T" by (metis Int_iff openE openin_open) then have "openin (top_of_set S) (S \ ball x e)" "path_connected (S \ ball x e)" by (simp_all add: assms convex_Int convex_imp_path_connected openin_open_Int) then show "\U. openin (top_of_set S) U \ path_connected U \ x \ U \ U \ V" using \0 < e\ \V = S \ T\ \ball x e \ T\ \x \ S\ by auto qed lemma convex_imp_locally_connected: fixes S :: "'a:: real_normed_vector set" shows "convex S \ locally connected S" by (simp add: locally_path_connected_imp_locally_connected convex_imp_locally_path_connected) subsection\Relations between components and path components\ lemma path_component_eq_connected_component: assumes "locally path_connected S" shows "(path_component S x = connected_component S x)" proof (cases "x \ S") case True have "openin (top_of_set (connected_component_set S x)) (path_component_set S x)" proof (rule openin_subset_trans) show "openin (top_of_set S) (path_component_set S x)" by (simp add: True assms locally_path_connected_2) show "connected_component_set S x \ S" by (simp add: connected_component_subset) qed (simp add: path_component_subset_connected_component) moreover have "closedin (top_of_set (connected_component_set S x)) (path_component_set S x)" proof (rule closedin_subset_trans [of S]) show "closedin (top_of_set S) (path_component_set S x)" by (simp add: assms closedin_path_component_locally_path_connected) show "connected_component_set S x \ S" by (simp add: connected_component_subset) qed (simp add: path_component_subset_connected_component) ultimately have *: "path_component_set S x = connected_component_set S x" by (metis connected_connected_component connected_clopen True path_component_eq_empty) then show ?thesis by blast next case False then show ?thesis by (metis Collect_empty_eq_bot connected_component_eq_empty path_component_eq_empty) qed lemma path_component_eq_connected_component_set: "locally path_connected S \ (path_component_set S x = connected_component_set S x)" by (simp add: path_component_eq_connected_component) lemma locally_path_connected_path_component: "locally path_connected S \ locally path_connected (path_component_set S x)" using locally_path_connected_connected_component path_component_eq_connected_component by fastforce lemma open_path_connected_component: fixes S :: "'a :: real_normed_vector set" shows "open S \ path_component S x = connected_component S x" by (simp add: path_component_eq_connected_component open_imp_locally_path_connected) lemma open_path_connected_component_set: fixes S :: "'a :: real_normed_vector set" shows "open S \ path_component_set S x = connected_component_set S x" by (simp add: open_path_connected_component) proposition locally_connected_quotient_image: assumes lcS: "locally connected S" and oo: "\T. T \ f ` S \ openin (top_of_set S) (S \ f -` T) \ openin (top_of_set (f ` S)) T" shows "locally connected (f ` S)" proof (clarsimp simp: locally_connected_open_component) fix U C assume opefSU: "openin (top_of_set (f ` S)) U" and "C \ components U" then have "C \ U" "U \ f ` S" by (meson in_components_subset openin_imp_subset)+ then have "openin (top_of_set (f ` S)) C \ openin (top_of_set S) (S \ f -` C)" by (auto simp: oo) moreover have "openin (top_of_set S) (S \ f -` C)" proof (subst openin_subopen, clarify) fix x assume "x \ S" "f x \ C" show "\T. openin (top_of_set S) T \ x \ T \ T \ (S \ f -` C)" proof (intro conjI exI) show "openin (top_of_set S) (connected_component_set (S \ f -` U) x)" proof (rule ccontr) assume **: "\ openin (top_of_set S) (connected_component_set (S \ f -` U) x)" then have "x \ (S \ f -` U)" using \U \ f ` S\ opefSU lcS locally_connected_2 oo by blast with ** show False by (metis (no_types) connected_component_eq_empty empty_iff openin_subopen) qed next show "x \ connected_component_set (S \ f -` U) x" using \C \ U\ \f x \ C\ \x \ S\ by auto next have contf: "continuous_on S f" by (simp add: continuous_on_open oo openin_imp_subset) then have "continuous_on (connected_component_set (S \ f -` U) x) f" by (meson connected_component_subset continuous_on_subset inf.boundedE) then have "connected (f ` connected_component_set (S \ f -` U) x)" by (rule connected_continuous_image [OF _ connected_connected_component]) moreover have "f ` connected_component_set (S \ f -` U) x \ U" using connected_component_in by blast moreover have "C \ f ` connected_component_set (S \ f -` U) x \ {}" using \C \ U\ \f x \ C\ \x \ S\ by fastforce ultimately have fC: "f ` (connected_component_set (S \ f -` U) x) \ C" by (rule components_maximal [OF \C \ components U\]) have cUC: "connected_component_set (S \ f -` U) x \ (S \ f -` C)" using connected_component_subset fC by blast have "connected_component_set (S \ f -` U) x \ connected_component_set (S \ f -` C) x" proof - { assume "x \ connected_component_set (S \ f -` U) x" then have ?thesis using cUC connected_component_idemp connected_component_mono by blast } then show ?thesis using connected_component_eq_empty by auto qed also have "\ \ (S \ f -` C)" by (rule connected_component_subset) finally show "connected_component_set (S \ f -` U) x \ (S \ f -` C)" . qed qed ultimately show "openin (top_of_set (f ` S)) C" by metis qed text\The proof resembles that above but is not identical!\ proposition locally_path_connected_quotient_image: assumes lcS: "locally path_connected S" and oo: "\T. T \ f ` S \ openin (top_of_set S) (S \ f -` T) \ openin (top_of_set (f ` S)) T" shows "locally path_connected (f ` S)" proof (clarsimp simp: locally_path_connected_open_path_component) fix U y assume opefSU: "openin (top_of_set (f ` S)) U" and "y \ U" then have "path_component_set U y \ U" "U \ f ` S" by (meson path_component_subset openin_imp_subset)+ then have "openin (top_of_set (f ` S)) (path_component_set U y) \ openin (top_of_set S) (S \ f -` path_component_set U y)" proof - have "path_component_set U y \ f ` S" using \U \ f ` S\ \path_component_set U y \ U\ by blast then show ?thesis using oo by blast qed moreover have "openin (top_of_set S) (S \ f -` path_component_set U y)" proof (subst openin_subopen, clarify) fix x assume "x \ S" and Uyfx: "path_component U y (f x)" then have "f x \ U" using path_component_mem by blast show "\T. openin (top_of_set S) T \ x \ T \ T \ (S \ f -` path_component_set U y)" proof (intro conjI exI) show "openin (top_of_set S) (path_component_set (S \ f -` U) x)" proof (rule ccontr) assume **: "\ openin (top_of_set S) (path_component_set (S \ f -` U) x)" then have "x \ (S \ f -` U)" by (metis (no_types, lifting) \U \ f ` S\ opefSU lcS oo locally_path_connected_open_path_component) then show False using ** \path_component_set U y \ U\ \x \ S\ \path_component U y (f x)\ by blast qed next show "x \ path_component_set (S \ f -` U) x" by (simp add: \f x \ U\ \x \ S\ path_component_refl) next have contf: "continuous_on S f" by (simp add: continuous_on_open oo openin_imp_subset) then have "continuous_on (path_component_set (S \ f -` U) x) f" by (meson Int_lower1 continuous_on_subset path_component_subset) then have "path_connected (f ` path_component_set (S \ f -` U) x)" by (simp add: path_connected_continuous_image) moreover have "f ` path_component_set (S \ f -` U) x \ U" using path_component_mem by fastforce moreover have "f x \ f ` path_component_set (S \ f -` U) x" by (force simp: \x \ S\ \f x \ U\ path_component_refl_eq) ultimately have "f ` (path_component_set (S \ f -` U) x) \ path_component_set U (f x)" by (meson path_component_maximal) also have "\ \ path_component_set U y" by (simp add: Uyfx path_component_maximal path_component_subset path_component_sym) finally have fC: "f ` (path_component_set (S \ f -` U) x) \ path_component_set U y" . have cUC: "path_component_set (S \ f -` U) x \ (S \ f -` path_component_set U y)" using path_component_subset fC by blast have "path_component_set (S \ f -` U) x \ path_component_set (S \ f -` path_component_set U y) x" proof - have "\a. path_component_set (path_component_set (S \ f -` U) x) a \ path_component_set (S \ f -` path_component_set U y) a" using cUC path_component_mono by blast then show ?thesis using path_component_path_component by blast qed also have "\ \ (S \ f -` path_component_set U y)" by (rule path_component_subset) finally show "path_component_set (S \ f -` U) x \ (S \ f -` path_component_set U y)" . qed qed ultimately show "openin (top_of_set (f ` S)) (path_component_set U y)" by metis qed subsection\<^marker>\tag unimportant\\Components, continuity, openin, closedin\ lemma continuous_on_components_gen: fixes f :: "'a::topological_space \ 'b::topological_space" assumes "\C. C \ components S \ openin (top_of_set S) C \ continuous_on C f" shows "continuous_on S f" proof (clarsimp simp: continuous_openin_preimage_eq) fix t :: "'b set" assume "open t" have *: "S \ f -` t = (\c \ components S. c \ f -` t)" by auto show "openin (top_of_set S) (S \ f -` t)" unfolding * using \open t\ assms continuous_openin_preimage_gen openin_trans openin_Union by blast qed lemma continuous_on_components: fixes f :: "'a::topological_space \ 'b::topological_space" assumes "locally connected S " "\C. C \ components S \ continuous_on C f" shows "continuous_on S f" proof (rule continuous_on_components_gen) fix C assume "C \ components S" then show "openin (top_of_set S) C \ continuous_on C f" by (simp add: assms openin_components_locally_connected) qed lemma continuous_on_components_eq: "locally connected S \ (continuous_on S f \ (\c \ components S. continuous_on c f))" by (meson continuous_on_components continuous_on_subset in_components_subset) lemma continuous_on_components_open: fixes S :: "'a::real_normed_vector set" assumes "open S " "\c. c \ components S \ continuous_on c f" shows "continuous_on S f" using continuous_on_components open_imp_locally_connected assms by blast lemma continuous_on_components_open_eq: fixes S :: "'a::real_normed_vector set" shows "open S \ (continuous_on S f \ (\c \ components S. continuous_on c f))" using continuous_on_subset in_components_subset by (blast intro: continuous_on_components_open) lemma closedin_union_complement_components: assumes U: "locally connected U" and S: "closedin (top_of_set U) S" and cuS: "c \ components(U - S)" shows "closedin (top_of_set U) (S \ \c)" proof - have di: "(\S T. S \ c \ T \ c' \ disjnt S T) \ disjnt (\ c) (\ c')" for c' by (simp add: disjnt_def) blast have "S \ U" using S closedin_imp_subset by blast moreover have "U - S = \c \ \(components (U - S) - c)" by (metis Diff_partition Union_components Union_Un_distrib assms(3)) moreover have "disjnt (\c) (\(components (U - S) - c))" apply (rule di) by (metis di DiffD1 DiffD2 assms(3) components_nonoverlap disjnt_def subsetCE) ultimately have eq: "S \ \c = U - (\(components(U - S) - c))" by (auto simp: disjnt_def) have *: "openin (top_of_set U) (\(components (U - S) - c))" proof (rule openin_Union [OF openin_trans [of "U - S"]]) show "openin (top_of_set (U - S)) T" if "T \ components (U - S) - c" for T using that by (simp add: U S locally_diff_closed openin_components_locally_connected) show "openin (top_of_set U) (U - S)" if "T \ components (U - S) - c" for T using that by (simp add: openin_diff S) qed have "closedin (top_of_set U) (U - \ (components (U - S) - c))" by (metis closedin_diff closedin_topspace topspace_euclidean_subtopology *) then have "openin (top_of_set U) (U - (U - \(components (U - S) - c)))" by (simp add: openin_diff) then show ?thesis by (force simp: eq closedin_def) qed lemma closed_union_complement_components: fixes S :: "'a::real_normed_vector set" assumes S: "closed S" and c: "c \ components(- S)" shows "closed(S \ \ c)" proof - have "closedin (top_of_set UNIV) (S \ \c)" by (metis Compl_eq_Diff_UNIV S c closed_closedin closedin_union_complement_components locally_connected_UNIV subtopology_UNIV) then show ?thesis by simp qed lemma closedin_Un_complement_component: fixes S :: "'a::real_normed_vector set" assumes u: "locally connected u" and S: "closedin (top_of_set u) S" and c: " c \ components(u - S)" shows "closedin (top_of_set u) (S \ c)" proof - have "closedin (top_of_set u) (S \ \{c})" using c by (blast intro: closedin_union_complement_components [OF u S]) then show ?thesis by simp qed lemma closed_Un_complement_component: fixes S :: "'a::real_normed_vector set" assumes S: "closed S" and c: " c \ components(-S)" shows "closed (S \ c)" by (metis Compl_eq_Diff_UNIV S c closed_closedin closedin_Un_complement_component locally_connected_UNIV subtopology_UNIV) subsection\Existence of isometry between subspaces of same dimension\ lemma isometry_subset_subspace: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes S: "subspace S" and T: "subspace T" and d: "dim S \ dim T" obtains f where "linear f" "f \ S \ T" "\x. x \ S \ norm(f x) = norm x" proof - obtain B where "B \ S" and Borth: "pairwise orthogonal B" and B1: "\x. x \ B \ norm x = 1" and "independent B" "finite B" "card B = dim S" "span B = S" by (metis orthonormal_basis_subspace [OF S] independent_imp_finite) obtain C where "C \ T" and Corth: "pairwise orthogonal C" and C1:"\x. x \ C \ norm x = 1" and "independent C" "finite C" "card C = dim T" "span C = T" by (metis orthonormal_basis_subspace [OF T] independent_imp_finite) obtain fb where "fb ` B \ C" "inj_on fb B" by (metis \card B = dim S\ \card C = dim T\ \finite B\ \finite C\ card_le_inj d) then have pairwise_orth_fb: "pairwise (\v j. orthogonal (fb v) (fb j)) B" using Corth unfolding pairwise_def inj_on_def by (blast intro: orthogonal_clauses) obtain f where "linear f" and ffb: "\x. x \ B \ f x = fb x" using linear_independent_extend \independent B\ by fastforce have "span (f ` B) \ span C" by (metis \fb ` B \ C\ ffb image_cong span_mono) then have "f ` S \ T" unfolding \span B = S\ \span C = T\ span_linear_image[OF \linear f\] . have [simp]: "\x. x \ B \ norm (fb x) = norm x" using B1 C1 \fb ` B \ C\ by auto have "norm (f x) = norm x" if "x \ S" for x proof - interpret linear f by fact obtain a where x: "x = (\v \ B. a v *\<^sub>R v)" using \finite B\ \span B = S\ \x \ S\ span_finite by fastforce have "norm (f x)^2 = norm (\v\B. a v *\<^sub>R fb v)^2" by (simp add: sum scale ffb x) also have "\ = (\v\B. norm ((a v *\<^sub>R fb v))^2)" proof (rule norm_sum_Pythagorean [OF \finite B\]) show "pairwise (\v j. orthogonal (a v *\<^sub>R fb v) (a j *\<^sub>R fb j)) B" by (rule pairwise_ortho_scaleR [OF pairwise_orth_fb]) qed also have "\ = norm x ^2" by (simp add: x pairwise_ortho_scaleR Borth norm_sum_Pythagorean [OF \finite B\]) finally show ?thesis by (simp add: norm_eq_sqrt_inner) qed then show ?thesis by (meson \f ` S \ T\ \linear f\ image_subset_iff_funcset that) qed proposition isometries_subspaces: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes S: "subspace S" and T: "subspace T" and d: "dim S = dim T" obtains f g where "linear f" "linear g" "f ` S = T" "g ` T = S" "\x. x \ S \ norm(f x) = norm x" "\x. x \ T \ norm(g x) = norm x" "\x. x \ S \ g(f x) = x" "\x. x \ T \ f(g x) = x" proof - obtain B where "B \ S" and Borth: "pairwise orthogonal B" and B1: "\x. x \ B \ norm x = 1" and "independent B" "finite B" "card B = dim S" "span B = S" by (metis orthonormal_basis_subspace [OF S] independent_imp_finite) obtain C where "C \ T" and Corth: "pairwise orthogonal C" and C1:"\x. x \ C \ norm x = 1" and "independent C" "finite C" "card C = dim T" "span C = T" by (metis orthonormal_basis_subspace [OF T] independent_imp_finite) obtain fb where "bij_betw fb B C" by (metis \finite B\ \finite C\ bij_betw_iff_card \card B = dim S\ \card C = dim T\ d) then have pairwise_orth_fb: "pairwise (\v j. orthogonal (fb v) (fb j)) B" using Corth unfolding pairwise_def inj_on_def bij_betw_def by (blast intro: orthogonal_clauses) obtain f where "linear f" and ffb: "\x. x \ B \ f x = fb x" using linear_independent_extend \independent B\ by fastforce interpret f: linear f by fact define gb where "gb \ inv_into B fb" then have pairwise_orth_gb: "pairwise (\v j. orthogonal (gb v) (gb j)) C" using Borth \bij_betw fb B C\ unfolding pairwise_def bij_betw_def by force obtain g where "linear g" and ggb: "\x. x \ C \ g x = gb x" using linear_independent_extend \independent C\ by fastforce interpret g: linear g by fact have "span (f ` B) \ span C" by (metis \bij_betw fb B C\ bij_betw_imp_surj_on eq_iff ffb image_cong) then have "f ` S \ T" unfolding \span B = S\ \span C = T\ span_linear_image[OF \linear f\] . have [simp]: "\x. x \ B \ norm (fb x) = norm x" using B1 C1 \bij_betw fb B C\ bij_betw_imp_surj_on by fastforce have f [simp]: "norm (f x) = norm x" "g (f x) = x" if "x \ S" for x proof - obtain a where x: "x = (\v \ B. a v *\<^sub>R v)" using \finite B\ \span B = S\ \x \ S\ span_finite by fastforce have "f x = (\v \ B. f (a v *\<^sub>R v))" using linear_sum [OF \linear f\] x by auto also have "\ = (\v \ B. a v *\<^sub>R f v)" by (simp add: f.sum f.scale) also have "\ = (\v \ B. a v *\<^sub>R fb v)" by (simp add: ffb cong: sum.cong) finally have *: "f x = (\v\B. a v *\<^sub>R fb v)" . then have "(norm (f x))\<^sup>2 = (norm (\v\B. a v *\<^sub>R fb v))\<^sup>2" by simp also have "\ = (\v\B. norm ((a v *\<^sub>R fb v))^2)" proof (rule norm_sum_Pythagorean [OF \finite B\]) show "pairwise (\v j. orthogonal (a v *\<^sub>R fb v) (a j *\<^sub>R fb j)) B" by (rule pairwise_ortho_scaleR [OF pairwise_orth_fb]) qed also have "\ = (norm x)\<^sup>2" by (simp add: x pairwise_ortho_scaleR Borth norm_sum_Pythagorean [OF \finite B\]) finally show "norm (f x) = norm x" by (simp add: norm_eq_sqrt_inner) have "g (f x) = g (\v\B. a v *\<^sub>R fb v)" by (simp add: *) also have "\ = (\v\B. g (a v *\<^sub>R fb v))" by (simp add: g.sum g.scale) also have "\ = (\v\B. a v *\<^sub>R g (fb v))" by (simp add: g.scale) also have "\ = (\v\B. a v *\<^sub>R v)" proof (rule sum.cong [OF refl]) show "a x *\<^sub>R g (fb x) = a x *\<^sub>R x" if "x \ B" for x using that \bij_betw fb B C\ bij_betwE bij_betw_inv_into_left gb_def ggb by fastforce qed also have "\ = x" using x by blast finally show "g (f x) = x" . qed have [simp]: "\x. x \ C \ norm (gb x) = norm x" by (metis B1 C1 \bij_betw fb B C\ bij_betw_imp_surj_on gb_def inv_into_into) have g [simp]: "f (g x) = x" if "x \ T" for x proof - obtain a where x: "x = (\v \ C. a v *\<^sub>R v)" using \finite C\ \span C = T\ \x \ T\ span_finite by fastforce have "g x = (\v \ C. g (a v *\<^sub>R v))" by (simp add: x g.sum) also have "\ = (\v \ C. a v *\<^sub>R g v)" by (simp add: g.scale) also have "\ = (\v \ C. a v *\<^sub>R gb v)" by (simp add: ggb cong: sum.cong) finally have "f (g x) = f (\v\C. a v *\<^sub>R gb v)" by simp also have "\ = (\v\C. f (a v *\<^sub>R gb v))" by (simp add: f.scale f.sum) also have "\ = (\v\C. a v *\<^sub>R f (gb v))" by (simp add: f.scale f.sum) also have "\ = (\v\C. a v *\<^sub>R v)" using \bij_betw fb B C\ by (simp add: bij_betw_def gb_def bij_betw_inv_into_right ffb inv_into_into) also have "\ = x" using x by blast finally show "f (g x) = x" . qed have gim: "g ` T = S" by (metis (full_types) S T \f ` S \ T\ d dim_eq_span dim_image_le f(2) g.linear_axioms image_iff linear_subspace_image span_eq_iff subset_iff) have fim: "f ` S = T" using \g ` T = S\ image_iff by fastforce have [simp]: "norm (g x) = norm x" if "x \ T" for x using fim that by auto show ?thesis by (rule that [OF \linear f\ \linear g\]) (simp_all add: fim gim) qed corollary isometry_subspaces: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes S: "subspace S" and T: "subspace T" and d: "dim S = dim T" obtains f where "linear f" "f ` S = T" "\x. x \ S \ norm(f x) = norm x" using isometries_subspaces [OF assms] by metis corollary isomorphisms_UNIV_UNIV: assumes "DIM('M) = DIM('N)" obtains f::"'M::euclidean_space \'N::euclidean_space" and g where "linear f" "linear g" "\x. norm(f x) = norm x" "\y. norm(g y) = norm y" "\x. g (f x) = x" "\y. f(g y) = y" using assms by (auto intro: isometries_subspaces [of "UNIV::'M set" "UNIV::'N set"]) lemma homeomorphic_subspaces: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes S: "subspace S" and T: "subspace T" and d: "dim S = dim T" shows "S homeomorphic T" proof - obtain f g where "linear f" "linear g" "f ` S = T" "g ` T = S" "\x. x \ S \ g(f x) = x" "\x. x \ T \ f(g x) = x" by (blast intro: isometries_subspaces [OF assms]) then show ?thesis unfolding homeomorphic_def homeomorphism_def apply (rule_tac x=f in exI, rule_tac x=g in exI) apply (auto simp: linear_continuous_on linear_conv_bounded_linear) done qed lemma homeomorphic_affine_sets: assumes "affine S" "affine T" "aff_dim S = aff_dim T" shows "S homeomorphic T" proof (cases "S = {} \ T = {}") case True with assms aff_dim_empty homeomorphic_empty show ?thesis by metis next case False then obtain a b where ab: "a \ S" "b \ T" by auto then have ss: "subspace ((+) (- a) ` S)" "subspace ((+) (- b) ` T)" using affine_diffs_subspace assms by blast+ have dd: "dim ((+) (- a) ` S) = dim ((+) (- b) ` T)" using assms ab by (simp add: aff_dim_eq_dim [OF hull_inc] image_def) have "S homeomorphic ((+) (- a) ` S)" by (fact homeomorphic_translation) also have "\ homeomorphic ((+) (- b) ` T)" by (rule homeomorphic_subspaces [OF ss dd]) also have "\ homeomorphic T" using homeomorphic_translation [of T "- b"] by (simp add: homeomorphic_sym [of T]) finally show ?thesis . qed subsection\Retracts, in a general sense, preserve (co)homotopic triviality)\ locale\<^marker>\tag important\ Retracts = fixes S h t k assumes conth: "continuous_on S h" and imh: "h ` S = t" and contk: "continuous_on t k" and imk: "k \ t \ S" and idhk: "\y. y \ t \ h(k y) = y" begin lemma homotopically_trivial_retraction_gen: assumes P: "\f. \continuous_on U f; f \ U \ t; Q f\ \ P(k \ f)" and Q: "\f. \continuous_on U f; f \ U \ S; P f\ \ Q(h \ f)" and Qeq: "\h k. (\x. x \ U \ h x = k x) \ Q h = Q k" and hom: "\f g. \continuous_on U f; f \ U \ S; P f; continuous_on U g; g \ U \ S; P g\ \ homotopic_with_canon P U S f g" and contf: "continuous_on U f" and imf: "f \ U \ t" and Qf: "Q f" and contg: "continuous_on U g" and img: "g \ U \ t" and Qg: "Q g" shows "homotopic_with_canon Q U t f g" proof - have "continuous_on U (k \ f)" by (meson contf continuous_on_compose continuous_on_subset contk funcset_image imf) moreover have "(k \ f) ` U \ S" using imf imk by fastforce moreover have "P (k \ f)" by (simp add: P Qf contf imf) moreover have "continuous_on U (k \ g)" by (meson contg continuous_on_compose continuous_on_subset contk funcset_image img) moreover have "(k \ g) ` U \ S" using img imk by fastforce moreover have "P (k \ g)" by (simp add: P Qg contg img) ultimately have "homotopic_with_canon P U S (k \ f) (k \ g)" by (simp add: hom image_subset_iff) then have "homotopic_with_canon Q U t (h \ (k \ f)) (h \ (k \ g))" apply (rule homotopic_with_compose_continuous_left [OF homotopic_with_mono]) using Q conth imh by force+ then show ?thesis proof (rule homotopic_with_eq; simp) show "\h k. (\x. x \ U \ h x = k x) \ Q h = Q k" using Qeq topspace_euclidean_subtopology by blast show "\x. x \ U \ f x = h (k (f x))" "\x. x \ U \ g x = h (k (g x))" using idhk imf img by fastforce+ qed qed lemma homotopically_trivial_retraction_null_gen: assumes P: "\f. \continuous_on U f; f \ U \ t; Q f\ \ P(k \ f)" and Q: "\f. \continuous_on U f; f \ U \ S; P f\ \ Q(h \ f)" and Qeq: "\h k. (\x. x \ U \ h x = k x) \ Q h = Q k" and hom: "\f. \continuous_on U f; f \ U \ S; P f\ \ \c. homotopic_with_canon P U S f (\x. c)" and contf: "continuous_on U f" and imf:"f \ U \ t" and Qf: "Q f" obtains c where "homotopic_with_canon Q U t f (\x. c)" proof - have feq: "\x. x \ U \ (h \ (k \ f)) x = f x" using idhk imf by auto have "continuous_on U (k \ f)" by (meson contf continuous_on_compose continuous_on_subset contk funcset_image imf) moreover have "(k \ f) \ U \ S" using imf imk by fastforce moreover have "P (k \ f)" by (simp add: P Qf contf imf) ultimately obtain c where "homotopic_with_canon P U S (k \ f) (\x. c)" by (metis hom) then have "homotopic_with_canon Q U t (h \ (k \ f)) (h \ (\x. c))" apply (rule homotopic_with_compose_continuous_left [OF homotopic_with_mono]) using Q conth imh by force+ then have "homotopic_with_canon Q U t f (\x. h c)" proof (rule homotopic_with_eq) show "\x. x \ topspace (top_of_set U) \ f x = (h \ (k \ f)) x" using feq by auto show "\h k. (\x. x \ topspace (top_of_set U) \ h x = k x) \ Q h = Q k" using Qeq topspace_euclidean_subtopology by blast qed auto then show ?thesis using that by blast qed lemma cohomotopically_trivial_retraction_gen: assumes P: "\f. \continuous_on t f; f \ t \ U; Q f\ \ P(f \ h)" and Q: "\f. \continuous_on S f; f \ S \ U; P f\ \ Q(f \ k)" and Qeq: "\h k. (\x. x \ t \ h x = k x) \ Q h = Q k" and hom: "\f g. \continuous_on S f; f \ S \ U; P f; continuous_on S g; g \ S \ U; P g\ \ homotopic_with_canon P S U f g" and contf: "continuous_on t f" and imf: "f \ t \ U" and Qf: "Q f" and contg: "continuous_on t g" and img: "g \ t \ U" and Qg: "Q g" shows "homotopic_with_canon Q t U f g" proof - have feq: "\x. x \ t \ (f \ h \ k) x = f x" using idhk imf by auto have geq: "\x. x \ t \ (g \ h \ k) x = g x" using idhk img by auto have "continuous_on S (f \ h)" using contf conth continuous_on_compose imh by blast moreover have "(f \ h) \ S \ U" using imf imh by fastforce moreover have "P (f \ h)" by (simp add: P Qf contf imf) moreover have "continuous_on S (g \ h)" using contg continuous_on_compose continuous_on_subset conth imh by blast moreover have "(g \ h) \ S \ U" using img imh by fastforce moreover have "P (g \ h)" by (simp add: P Qg contg img) ultimately have "homotopic_with_canon P S U (f \ h) (g \ h)" by (simp add: hom) then have "homotopic_with_canon Q t U (f \ h \ k) (g \ h \ k)" apply (rule homotopic_with_compose_continuous_right [OF homotopic_with_mono]) using Q contk imk by force+ then show ?thesis proof (rule homotopic_with_eq) show "f x = (f \ h \ k) x" "g x = (g \ h \ k) x" if "x \ topspace (top_of_set t)" for x using feq geq that by force+ qed (use Qeq topspace_euclidean_subtopology in blast) qed lemma cohomotopically_trivial_retraction_null_gen: assumes P: "\f. \continuous_on t f; f \ t \ U; Q f\ \ P(f \ h)" and Q: "\f. \continuous_on S f; f \ S \ U; P f\ \ Q(f \ k)" and Qeq: "\h k. (\x. x \ t \ h x = k x) \ Q h = Q k" and hom: "\f g. \continuous_on S f; f \ S \ U; P f\ \ \c. homotopic_with_canon P S U f (\x. c)" and contf: "continuous_on t f" and imf: "f \ t \ U" and Qf: "Q f" obtains c where "homotopic_with_canon Q t U f (\x. c)" proof - have feq: "\x. x \ t \ (f \ h \ k) x = f x" using idhk imf by auto have "continuous_on S (f \ h)" using contf conth continuous_on_compose imh by blast moreover have "(f \ h) \ S \ U" using imf imh by fastforce moreover have "P (f \ h)" by (simp add: P Qf contf imf) ultimately obtain c where "homotopic_with_canon P S U (f \ h) (\x. c)" by (metis hom) then have \
: "homotopic_with_canon Q t U (f \ h \ k) ((\x. c) \ k)" proof (rule homotopic_with_compose_continuous_right [OF homotopic_with_mono]) show "\h. \continuous_map (top_of_set S) (top_of_set U) h; P h\ \ Q (h \ k)" using Q by auto qed (use contk imk in force)+ moreover have "homotopic_with_canon Q t U f (\x. c)" using homotopic_with_eq [OF \
] feq Qeq by fastforce ultimately show ?thesis using that by blast qed end lemma simply_connected_retraction_gen: shows "\simply_connected S; continuous_on S h; h ` S = T; continuous_on T k; k \ T \ S; \y. y \ T \ h(k y) = y\ \ simply_connected T" apply (simp add: simply_connected_def path_def path_image_def homotopic_loops_def, clarify) apply (rule Retracts.homotopically_trivial_retraction_gen [of S h _ k _ "\p. pathfinish p = pathstart p" "\p. pathfinish p = pathstart p"]) apply (simp_all add: Retracts_def pathfinish_def pathstart_def image_subset_iff_funcset) done lemma homeomorphic_simply_connected: "\S homeomorphic T; simply_connected S\ \ simply_connected T" by (auto simp: homeomorphic_def homeomorphism_def intro: simply_connected_retraction_gen) lemma homeomorphic_simply_connected_eq: "S homeomorphic T \ (simply_connected S \ simply_connected T)" by (metis homeomorphic_simply_connected homeomorphic_sym) subsection\Homotopy equivalence\ subsection\Homotopy equivalence of topological spaces.\ definition\<^marker>\tag important\ homotopy_equivalent_space (infix "homotopy'_equivalent'_space" 50) where "X homotopy_equivalent_space Y \ (\f g. continuous_map X Y f \ continuous_map Y X g \ homotopic_with (\x. True) X X (g \ f) id \ homotopic_with (\x. True) Y Y (f \ g) id)" lemma homeomorphic_imp_homotopy_equivalent_space: "X homeomorphic_space Y \ X homotopy_equivalent_space Y" unfolding homeomorphic_space_def homotopy_equivalent_space_def apply (erule ex_forward)+ by (simp add: homotopic_with_equal homotopic_with_sym homeomorphic_maps_def) lemma homotopy_equivalent_space_refl: "X homotopy_equivalent_space X" by (simp add: homeomorphic_imp_homotopy_equivalent_space homeomorphic_space_refl) lemma homotopy_equivalent_space_sym: "X homotopy_equivalent_space Y \ Y homotopy_equivalent_space X" by (meson homotopy_equivalent_space_def) lemma homotopy_eqv_trans [trans]: assumes 1: "X homotopy_equivalent_space Y" and 2: "Y homotopy_equivalent_space U" shows "X homotopy_equivalent_space U" proof - obtain f1 g1 where f1: "continuous_map X Y f1" and g1: "continuous_map Y X g1" and hom1: "homotopic_with (\x. True) X X (g1 \ f1) id" "homotopic_with (\x. True) Y Y (f1 \ g1) id" using 1 by (auto simp: homotopy_equivalent_space_def) obtain f2 g2 where f2: "continuous_map Y U f2" and g2: "continuous_map U Y g2" and hom2: "homotopic_with (\x. True) Y Y (g2 \ f2) id" "homotopic_with (\x. True) U U (f2 \ g2) id" using 2 by (auto simp: homotopy_equivalent_space_def) have "homotopic_with (\f. True) X Y (g2 \ f2 \ f1) (id \ f1)" using f1 hom2(1) homotopic_with_compose_continuous_map_right by metis then have "homotopic_with (\f. True) X Y (g2 \ (f2 \ f1)) (id \ f1)" by (simp add: o_assoc) then have "homotopic_with (\x. True) X X (g1 \ (g2 \ (f2 \ f1))) (g1 \ (id \ f1))" by (simp add: g1 homotopic_with_compose_continuous_map_left) moreover have "homotopic_with (\x. True) X X (g1 \ id \ f1) id" using hom1 by simp ultimately have SS: "homotopic_with (\x. True) X X (g1 \ g2 \ (f2 \ f1)) id" by (metis comp_assoc homotopic_with_trans id_comp) have "homotopic_with (\f. True) U Y (f1 \ g1 \ g2) (id \ g2)" using g2 hom1(2) homotopic_with_compose_continuous_map_right by fastforce then have "homotopic_with (\f. True) U Y (f1 \ (g1 \ g2)) (id \ g2)" by (simp add: o_assoc) then have "homotopic_with (\x. True) U U (f2 \ (f1 \ (g1 \ g2))) (f2 \ (id \ g2))" by (simp add: f2 homotopic_with_compose_continuous_map_left) moreover have "homotopic_with (\x. True) U U (f2 \ id \ g2) id" using hom2 by simp ultimately have UU: "homotopic_with (\x. True) U U (f2 \ f1 \ (g1 \ g2)) id" by (simp add: fun.map_comp hom2(2) homotopic_with_trans) show ?thesis unfolding homotopy_equivalent_space_def by (blast intro: f1 f2 g1 g2 continuous_map_compose SS UU) qed lemma deformation_retraction_imp_homotopy_equivalent_space: "\homotopic_with (\x. True) X X (S \ r) id; retraction_maps X Y r S\ \ X homotopy_equivalent_space Y" unfolding homotopy_equivalent_space_def retraction_maps_def using homotopic_with_id2 by fastforce lemma deformation_retract_imp_homotopy_equivalent_space: "\homotopic_with (\x. True) X X r id; retraction_maps X Y r id\ \ X homotopy_equivalent_space Y" using deformation_retraction_imp_homotopy_equivalent_space by force lemma deformation_retract_of_space: "S \ topspace X \ (\r. homotopic_with (\x. True) X X id r \ retraction_maps X (subtopology X S) r id) \ S retract_of_space X \ (\f. homotopic_with (\x. True) X X id f \ f ` (topspace X) \ S)" proof (cases "S \ topspace X") case True moreover have "(\r. homotopic_with (\x. True) X X id r \ retraction_maps X (subtopology X S) r id) \ (S retract_of_space X \ (\f. homotopic_with (\x. True) X X id f \ f ` topspace X \ S))" unfolding retract_of_space_def proof safe fix f r assume f: "homotopic_with (\x. True) X X id f" and fS: "f ` topspace X \ S" and r: "continuous_map X (subtopology X S) r" and req: "\x\S. r x = x" show "\r. homotopic_with (\x. True) X X id r \ retraction_maps X (subtopology X S) r id" proof (intro exI conjI) have "homotopic_with (\x. True) X X f r" proof (rule homotopic_with_eq) show "homotopic_with (\x. True) X X (r \ f) (r \ id)" by (metis continuous_map_into_fulltopology f homotopic_with_compose_continuous_map_left homotopic_with_symD r) show "f x = (r \ f) x" if "x \ topspace X" for x using that fS req by auto qed auto then show "homotopic_with (\x. True) X X id r" by (rule homotopic_with_trans [OF f]) next show "retraction_maps X (subtopology X S) r id" by (simp add: r req retraction_maps_def) qed qed (use True in \auto simp: retraction_maps_def topspace_subtopology_subset continuous_map_in_subtopology\) ultimately show ?thesis by simp qed (auto simp: retract_of_space_def retraction_maps_def) subsection\Contractible spaces\ text\The definition (which agrees with "contractible" on subsets of Euclidean space) is a little cryptic because we don't in fact assume that the constant "a" is in the space. This forces the convention that the empty space / set is contractible, avoiding some special cases. \ definition contractible_space where "contractible_space X \ \a. homotopic_with (\x. True) X X id (\x. a)" lemma contractible_space_top_of_set [simp]:"contractible_space (top_of_set S) \ contractible S" by (auto simp: contractible_space_def contractible_def) lemma contractible_space_empty [simp]: "contractible_space trivial_topology" unfolding contractible_space_def homotopic_with_def apply (rule_tac x=undefined in exI) apply (rule_tac x="\(t,x). if t = 0 then x else undefined" in exI) apply (auto simp: continuous_map_on_empty) done lemma contractible_space_singleton [simp]: "contractible_space (discrete_topology{a})" unfolding contractible_space_def homotopic_with_def apply (rule_tac x=a in exI) apply (rule_tac x="\(t,x). if t = 0 then x else a" in exI) apply (auto intro: continuous_map_eq [where f = "\z. a"]) done lemma contractible_space_subset_singleton: "topspace X \ {a} \ contractible_space X" by (metis contractible_space_empty contractible_space_singleton null_topspace_iff_trivial subset_singletonD subtopology_eq_discrete_topology_sing) lemma contractible_space_subtopology_singleton [simp]: "contractible_space (subtopology X {a})" by (meson contractible_space_subset_singleton insert_subset path_connectedin_singleton path_connectedin_subtopology subsetI) lemma contractible_space: "contractible_space X \ X = trivial_topology \ (\a \ topspace X. homotopic_with (\x. True) X X id (\x. a))" proof (cases "X = trivial_topology") case False then show ?thesis using homotopic_with_imp_continuous_maps by (fastforce simp: contractible_space_def) qed (simp add: contractible_space_empty) lemma contractible_imp_path_connected_space: assumes "contractible_space X" shows "path_connected_space X" proof (cases "X = trivial_topology") case False have *: "path_connected_space X" if "a \ topspace X" and conth: "continuous_map (prod_topology (top_of_set {0..1}) X) X h" and h: "\x. h (0, x) = x" "\x. h (1, x) = a" for a and h :: "real \ 'a \ 'a" proof - have "path_component_of X b a" if "b \ topspace X" for b unfolding path_component_of_def proof (intro exI conjI) let ?g = "h \ (\x. (x,b))" show "pathin X ?g" unfolding pathin_def proof (rule continuous_map_compose [OF _ conth]) show "continuous_map (top_of_set {0..1}) (prod_topology (top_of_set {0..1}) X) (\x. (x, b))" using that by (auto intro!: continuous_intros) qed qed (use h in auto) then show ?thesis by (metis path_component_of_equiv path_connected_space_iff_path_component) qed show ?thesis using assms False by (auto simp: contractible_space homotopic_with_def *) qed (simp add: path_connected_space_topspace_empty) lemma contractible_imp_connected_space: "contractible_space X \ connected_space X" by (simp add: contractible_imp_path_connected_space path_connected_imp_connected_space) lemma contractible_space_alt: "contractible_space X \ (\a \ topspace X. homotopic_with (\x. True) X X id (\x. a))" (is "?lhs = ?rhs") proof assume X: ?lhs then obtain a where a: "homotopic_with (\x. True) X X id (\x. a)" by (auto simp: contractible_space_def) show ?rhs proof show "homotopic_with (\x. True) X X id (\x. b)" if "b \ topspace X" for b proof (rule homotopic_with_trans [OF a]) show "homotopic_with (\x. True) X X (\x. a) (\x. b)" using homotopic_constant_maps path_connected_space_imp_path_component_of by (metis X a contractible_imp_path_connected_space homotopic_with_sym homotopic_with_trans path_component_of_equiv that) qed qed next assume R: ?rhs then show ?lhs using contractible_space_def by fastforce qed lemma compose_const [simp]: "f \ (\x. a) = (\x. f a)" "(\x. a) \ g = (\x. a)" by (simp_all add: o_def) lemma nullhomotopic_through_contractible_space: assumes f: "continuous_map X Y f" and g: "continuous_map Y Z g" and Y: "contractible_space Y" obtains c where "homotopic_with (\h. True) X Z (g \ f) (\x. c)" proof - obtain b where b: "homotopic_with (\x. True) Y Y id (\x. b)" using Y by (auto simp: contractible_space_def) show thesis using homotopic_with_compose_continuous_map_right [OF homotopic_with_compose_continuous_map_left [OF b g] f] by (force simp: that) qed lemma nullhomotopic_into_contractible_space: assumes f: "continuous_map X Y f" and Y: "contractible_space Y" obtains c where "homotopic_with (\h. True) X Y f (\x. c)" using nullhomotopic_through_contractible_space [OF f _ Y] by (metis continuous_map_id id_comp) lemma nullhomotopic_from_contractible_space: assumes f: "continuous_map X Y f" and X: "contractible_space X" obtains c where "homotopic_with (\h. True) X Y f (\x. c)" using nullhomotopic_through_contractible_space [OF _ f X] by (metis comp_id continuous_map_id) lemma homotopy_dominated_contractibility: assumes f: "continuous_map X Y f" and g: "continuous_map Y X g" and hom: "homotopic_with (\x. True) Y Y (f \ g) id" and X: "contractible_space X" shows "contractible_space Y" proof - obtain c where c: "homotopic_with (\h. True) X Y f (\x. c)" using nullhomotopic_from_contractible_space [OF f X] . have "homotopic_with (\x. True) Y Y (f \ g) (\x. c)" using homotopic_with_compose_continuous_map_right [OF c g] by fastforce then have "homotopic_with (\x. True) Y Y id (\x. c)" using homotopic_with_trans [OF _ hom] homotopic_with_symD by blast then show ?thesis unfolding contractible_space_def .. qed lemma homotopy_equivalent_space_contractibility: "X homotopy_equivalent_space Y \ (contractible_space X \ contractible_space Y)" unfolding homotopy_equivalent_space_def by (blast intro: homotopy_dominated_contractibility) lemma homeomorphic_space_contractibility: "X homeomorphic_space Y \ (contractible_space X \ contractible_space Y)" by (simp add: homeomorphic_imp_homotopy_equivalent_space homotopy_equivalent_space_contractibility) lemma homotopic_through_contractible_space: "continuous_map X Y f \ continuous_map X Y f' \ continuous_map Y Z g \ continuous_map Y Z g' \ contractible_space Y \ path_connected_space Z \ homotopic_with (\h. True) X Z (g \ f) (g' \ f')" using nullhomotopic_through_contractible_space [of X Y f Z g] using nullhomotopic_through_contractible_space [of X Y f' Z g'] by (smt (verit) continuous_map_const homotopic_constant_maps homotopic_with_imp_continuous_maps homotopic_with_symD homotopic_with_trans path_connected_space_imp_path_component_of) lemma homotopic_from_contractible_space: "continuous_map X Y f \ continuous_map X Y g \ contractible_space X \ path_connected_space Y \ homotopic_with (\x. True) X Y f g" by (metis comp_id continuous_map_id homotopic_through_contractible_space) lemma homotopic_into_contractible_space: "continuous_map X Y f \ continuous_map X Y g \ contractible_space Y \ homotopic_with (\x. True) X Y f g" by (metis continuous_map_id contractible_imp_path_connected_space homotopic_through_contractible_space id_comp) lemma contractible_eq_homotopy_equivalent_singleton_subtopology: "contractible_space X \ X = trivial_topology \ (\a \ topspace X. X homotopy_equivalent_space (subtopology X {a}))"(is "?lhs = ?rhs") proof (cases "X = trivial_topology") case False show ?thesis proof assume ?lhs then obtain a where a: "homotopic_with (\x. True) X X id (\x. a)" by (auto simp: contractible_space_def) then have "a \ topspace X" by (metis False continuous_map_const homotopic_with_imp_continuous_maps) then have "homotopic_with (\x. True) (subtopology X {a}) (subtopology X {a}) id (\x. a)" using connectedin_absolute connectedin_sing contractible_space_alt contractible_space_subtopology_singleton by fastforce then have "X homotopy_equivalent_space subtopology X {a}" unfolding homotopy_equivalent_space_def using \a \ topspace X\ by (metis (full_types) a comp_id continuous_map_const continuous_map_id_subt empty_subsetI homotopic_with_symD id_comp insertI1 insert_subset topspace_subtopology_subset) with \a \ topspace X\ show ?rhs by blast next assume ?rhs then show ?lhs by (meson False contractible_space_subtopology_singleton homotopy_equivalent_space_contractibility) qed qed (simp add: contractible_space_empty) lemma contractible_space_retraction_map_image: assumes "retraction_map X Y f" and X: "contractible_space X" shows "contractible_space Y" proof - obtain g where f: "continuous_map X Y f" and g: "continuous_map Y X g" and fg: "\y \ topspace Y. f(g y) = y" using assms by (auto simp: retraction_map_def retraction_maps_def) obtain a where a: "homotopic_with (\x. True) X X id (\x. a)" using X by (auto simp: contractible_space_def) have "homotopic_with (\x. True) Y Y id (\x. f a)" proof (rule homotopic_with_eq) show "homotopic_with (\x. True) Y Y (f \ id \ g) (f \ (\x. a) \ g)" using f g a homotopic_with_compose_continuous_map_left homotopic_with_compose_continuous_map_right by metis qed (use fg in auto) then show ?thesis unfolding contractible_space_def by blast qed lemma contractible_space_prod_topology: "contractible_space(prod_topology X Y) \ X = trivial_topology \ Y = trivial_topology \ contractible_space X \ contractible_space Y" proof (cases "X = trivial_topology \ Y = trivial_topology") case True then have "(prod_topology X Y) = trivial_topology" by simp then show ?thesis by (auto simp: contractible_space_empty) next case False have "contractible_space(prod_topology X Y) \ contractible_space X \ contractible_space Y" proof safe assume XY: "contractible_space (prod_topology X Y)" with False have "retraction_map (prod_topology X Y) X fst" by (auto simp: contractible_space False retraction_map_fst) then show "contractible_space X" by (rule contractible_space_retraction_map_image [OF _ XY]) have "retraction_map (prod_topology X Y) Y snd" using False XY by (auto simp: contractible_space False retraction_map_snd) then show "contractible_space Y" by (rule contractible_space_retraction_map_image [OF _ XY]) next assume "contractible_space X" and "contractible_space Y" with False obtain a b where "a \ topspace X" and a: "homotopic_with (\x. True) X X id (\x. a)" and "b \ topspace Y" and b: "homotopic_with (\x. True) Y Y id (\x. b)" by (auto simp: contractible_space) with False show "contractible_space (prod_topology X Y)" apply (simp add: contractible_space) apply (rule_tac x=a in bexI) apply (rule_tac x=b in bexI) using homotopic_with_prod_topology [OF a b] apply (metis (no_types, lifting) case_prod_Pair case_prod_beta' eq_id_iff) apply auto done qed with False show ?thesis by auto qed lemma contractible_space_product_topology: "contractible_space(product_topology X I) \ (product_topology X I) = trivial_topology \ (\i \ I. contractible_space(X i))" proof (cases "(product_topology X I) = trivial_topology") case False have 1: "contractible_space (X i)" if XI: "contractible_space (product_topology X I)" and "i \ I" for i proof (rule contractible_space_retraction_map_image [OF _ XI]) show "retraction_map (product_topology X I) (X i) (\x. x i)" using False by (simp add: retraction_map_product_projection \i \ I\) qed have 2: "contractible_space (product_topology X I)" if "x \ topspace (product_topology X I)" and cs: "\i\I. contractible_space (X i)" for x :: "'a \ 'b" proof - obtain f where f: "\i. i\I \ homotopic_with (\x. True) (X i) (X i) id (\x. f i)" using cs unfolding contractible_space_def by metis have "homotopic_with (\x. True) (product_topology X I) (product_topology X I) id (\x. restrict f I)" by (rule homotopic_with_eq [OF homotopic_with_product_topology [OF f]]) (auto) then show ?thesis by (auto simp: contractible_space_def) qed show ?thesis using False 1 2 by (meson equals0I subtopology_eq_discrete_topology_empty) qed auto lemma contractible_space_subtopology_euclideanreal [simp]: "contractible_space(subtopology euclideanreal S) \ is_interval S" (is "?lhs = ?rhs") proof assume ?lhs then have "path_connectedin (subtopology euclideanreal S) S" using contractible_imp_path_connected_space path_connectedin_topspace path_connectedin_absolute by (simp add: contractible_imp_path_connected) then show ?rhs by (simp add: is_interval_path_connected_1) next assume ?rhs then have "convex S" by (simp add: is_interval_convex_1) show ?lhs proof (cases "S = {}") case False then obtain z where "z \ S" by blast show ?thesis unfolding contractible_space_def homotopic_with_def proof (intro exI conjI allI) note \
= convexD [OF \convex S\, simplified] show "continuous_map (prod_topology (top_of_set {0..1}) (top_of_set S)) (top_of_set S) (\(t,x). (1 - t) * x + t * z)" using \z \ S\ by (auto simp: case_prod_unfold intro!: continuous_intros \
) qed auto qed (simp add: contractible_space_empty) qed corollary contractible_space_euclideanreal: "contractible_space euclideanreal" proof - have "contractible_space (subtopology euclideanreal UNIV)" using contractible_space_subtopology_euclideanreal by blast then show ?thesis by simp qed abbreviation\<^marker>\tag important\ homotopy_eqv :: "'a::topological_space set \ 'b::topological_space set \ bool" (infix "homotopy'_eqv" 50) where "S homotopy_eqv T \ top_of_set S homotopy_equivalent_space top_of_set T" lemma homeomorphic_imp_homotopy_eqv: "S homeomorphic T \ S homotopy_eqv T" unfolding homeomorphic_def homeomorphism_def homotopy_equivalent_space_def by (metis continuous_map_subtopology_eu homotopic_with_id2 openin_imp_subset openin_subtopology_self topspace_euclidean_subtopology) lemma homotopy_eqv_inj_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "(f ` S) homotopy_eqv S" by (metis assms homeomorphic_sym homeomorphic_imp_homotopy_eqv linear_homeomorphic_image) lemma homotopy_eqv_translation: fixes S :: "'a::real_normed_vector set" shows "(+) a ` S homotopy_eqv S" using homeomorphic_imp_homotopy_eqv homeomorphic_translation homeomorphic_sym by blast lemma homotopy_eqv_homotopic_triviality_imp: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" and U :: "'c::real_normed_vector set" assumes "S homotopy_eqv T" and f: "continuous_on U f" "f \ U \ T" and g: "continuous_on U g" "g \ U \ T" and homUS: "\f g. \continuous_on U f; f \ U \ S; continuous_on U g; g \ U \ S\ \ homotopic_with_canon (\x. True) U S f g" shows "homotopic_with_canon (\x. True) U T f g" proof - obtain h k where h: "continuous_on S h" "h \ S \ T" and k: "continuous_on T k" "k \ T \ S" and hom: "homotopic_with_canon (\x. True) S S (k \ h) id" "homotopic_with_canon (\x. True) T T (h \ k) id" using assms by (force simp: homotopy_equivalent_space_def image_subset_iff_funcset) have "homotopic_with_canon (\f. True) U S (k \ f) (k \ g)" proof (rule homUS) show "continuous_on U (k \ f)" "continuous_on U (k \ g)" using continuous_on_compose continuous_on_subset f g k by (metis funcset_image)+ qed (use f g k in \(force simp: o_def)+\ ) then have "homotopic_with_canon (\x. True) U T (h \ (k \ f)) (h \ (k \ g))" by (simp add: h homotopic_with_compose_continuous_map_left image_subset_iff_funcset) moreover have "homotopic_with_canon (\x. True) U T (h \ k \ f) (id \ f)" by (rule homotopic_with_compose_continuous_right [where X=T and Y=T]; simp add: hom f) moreover have "homotopic_with_canon (\x. True) U T (h \ k \ g) (id \ g)" by (rule homotopic_with_compose_continuous_right [where X=T and Y=T]; simp add: hom g) ultimately show "homotopic_with_canon (\x. True) U T f g" unfolding o_assoc by (metis homotopic_with_trans homotopic_with_sym id_comp) qed lemma homotopy_eqv_homotopic_triviality: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" and U :: "'c::real_normed_vector set" assumes "S homotopy_eqv T" shows "(\f g. continuous_on U f \ f \ U \ S \ continuous_on U g \ g \ U \ S \ homotopic_with_canon (\x. True) U S f g) \ (\f g. continuous_on U f \ f \ U \ T \ continuous_on U g \ g \ U \ T \ homotopic_with_canon (\x. True) U T f g)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (metis assms homotopy_eqv_homotopic_triviality_imp) next assume ?rhs moreover have "T homotopy_eqv S" using assms homotopy_equivalent_space_sym by blast ultimately show ?lhs by (blast intro: homotopy_eqv_homotopic_triviality_imp) qed lemma homotopy_eqv_cohomotopic_triviality_null_imp: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" and U :: "'c::real_normed_vector set" assumes "S homotopy_eqv T" and f: "continuous_on T f" "f \ T \ U" and homSU: "\f. \continuous_on S f; f \ S \ U\ \ \c. homotopic_with_canon (\x. True) S U f (\x. c)" obtains c where "homotopic_with_canon (\x. True) T U f (\x. c)" proof - obtain h k where h: "continuous_on S h" "h \ S \ T" and k: "continuous_on T k" "k \ T \ S" and hom: "homotopic_with_canon (\x. True) S S (k \ h) id" "homotopic_with_canon (\x. True) T T (h \ k) id" using assms by (force simp: homotopy_equivalent_space_def image_subset_iff_funcset) obtain c where "homotopic_with_canon (\x. True) S U (f \ h) (\x. c)" proof (rule exE [OF homSU]) show "continuous_on S (f \ h)" by (metis continuous_on_compose continuous_on_subset f h funcset_image) qed (use f h in force) then have "homotopic_with_canon (\x. True) T U ((f \ h) \ k) ((\x. c) \ k)" by (rule homotopic_with_compose_continuous_right [where X=S]) (use k in auto) moreover have "homotopic_with_canon (\x. True) T U (f \ id) (f \ (h \ k))" by (rule homotopic_with_compose_continuous_left [where Y=T]) (use f in \auto simp: hom homotopic_with_symD\) ultimately show ?thesis using that homotopic_with_trans by (fastforce simp: o_def) qed lemma homotopy_eqv_cohomotopic_triviality_null: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" and U :: "'c::real_normed_vector set" assumes "S homotopy_eqv T" shows "(\f. continuous_on S f \ f \ S \ U \ (\c. homotopic_with_canon (\x. True) S U f (\x. c))) \ (\f. continuous_on T f \ f \ T \ U \ (\c. homotopic_with_canon (\x. True) T U f (\x. c)))" by (rule iffI; metis assms homotopy_eqv_cohomotopic_triviality_null_imp homotopy_equivalent_space_sym) text \Similar to the proof above\ lemma homotopy_eqv_homotopic_triviality_null_imp: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" and U :: "'c::real_normed_vector set" assumes "S homotopy_eqv T" and f: "continuous_on U f" "f \ U \ T" and homSU: "\f. \continuous_on U f; f \ U \ S\ \ \c. homotopic_with_canon (\x. True) U S f (\x. c)" shows "\c. homotopic_with_canon (\x. True) U T f (\x. c)" proof - obtain h k where h: "continuous_on S h" "h \ S \ T" and k: "continuous_on T k" "k \ T \ S" and hom: "homotopic_with_canon (\x. True) S S (k \ h) id" "homotopic_with_canon (\x. True) T T (h \ k) id" using assms by (force simp: homotopy_equivalent_space_def image_subset_iff_funcset) obtain c::'a where "homotopic_with_canon (\x. True) U S (k \ f) (\x. c)" proof (rule exE [OF homSU [of "k \ f"]]) show "continuous_on U (k \ f)" using continuous_on_compose continuous_on_subset f k by (metis funcset_image) qed (use f k in force) then have "homotopic_with_canon (\x. True) U T (h \ (k \ f)) (h \ (\x. c))" by (rule homotopic_with_compose_continuous_left [where Y=S]) (use h in auto) moreover have "homotopic_with_canon (\x. True) U T (id \ f) ((h \ k) \ f)" by (rule homotopic_with_compose_continuous_right [where X=T]) (use f in \auto simp: hom homotopic_with_symD\) ultimately show ?thesis using homotopic_with_trans by (fastforce simp: o_def) qed lemma homotopy_eqv_homotopic_triviality_null: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" and U :: "'c::real_normed_vector set" assumes "S homotopy_eqv T" shows "(\f. continuous_on U f \ f \ U \ S \ (\c. homotopic_with_canon (\x. True) U S f (\x. c))) \ (\f. continuous_on U f \ f \ U \ T \ (\c. homotopic_with_canon (\x. True) U T f (\x. c)))" by (rule iffI; metis assms homotopy_eqv_homotopic_triviality_null_imp homotopy_equivalent_space_sym) lemma homotopy_eqv_contractible_sets: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" assumes "contractible S" "contractible T" "S = {} \ T = {}" shows "S homotopy_eqv T" proof (cases "S = {}") case True with assms show ?thesis using homeomorphic_imp_homotopy_eqv by fastforce next case False with assms obtain a b where "a \ S" "b \ T" by auto then show ?thesis unfolding homotopy_equivalent_space_def apply (rule_tac x="\x. b" in exI, rule_tac x="\x. a" in exI) apply (intro assms conjI continuous_on_id' homotopic_into_contractible; force) done qed lemma homotopy_eqv_empty1 [simp]: fixes S :: "'a::real_normed_vector set" shows "S homotopy_eqv ({}::'b::real_normed_vector set) \ S = {}" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (metis continuous_map_subtopology_eu empty_iff equalityI homotopy_equivalent_space_def image_subset_iff subsetI) qed (use homeomorphic_imp_homotopy_eqv in force) lemma homotopy_eqv_empty2 [simp]: fixes S :: "'a::real_normed_vector set" shows "({}::'b::real_normed_vector set) homotopy_eqv S \ S = {}" using homotopy_equivalent_space_sym homotopy_eqv_empty1 by blast lemma homotopy_eqv_contractibility: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" shows "S homotopy_eqv T \ (contractible S \ contractible T)" by (meson contractible_space_top_of_set homotopy_equivalent_space_contractibility) lemma homotopy_eqv_sing: fixes S :: "'a::real_normed_vector set" and a :: "'b::real_normed_vector" shows "S homotopy_eqv {a} \ S \ {} \ contractible S" by (metis contractible_sing empty_not_insert homotopy_eqv_contractibility homotopy_eqv_contractible_sets homotopy_eqv_empty2) lemma homeomorphic_contractible_eq: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" shows "S homeomorphic T \ (contractible S \ contractible T)" by (simp add: homeomorphic_imp_homotopy_eqv homotopy_eqv_contractibility) lemma homeomorphic_contractible: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" shows "\contractible S; S homeomorphic T\ \ contractible T" by (metis homeomorphic_contractible_eq) subsection\<^marker>\tag unimportant\\Misc other results\ lemma bounded_connected_Compl_real: fixes S :: "real set" assumes "bounded S" and conn: "connected(- S)" shows "S = {}" proof - obtain a b where "S \ box a b" by (meson assms bounded_subset_box_symmetric) then have "a \ S" "b \ S" by auto then have "\x. a \ x \ x \ b \ x \ - S" by (meson Compl_iff conn connected_iff_interval) then show ?thesis using \S \ box a b\ by auto qed corollary bounded_path_connected_Compl_real: fixes S :: "real set" assumes "bounded S" "path_connected(- S)" shows "S = {}" by (simp add: assms bounded_connected_Compl_real path_connected_imp_connected) lemma bounded_connected_Compl_1: fixes S :: "'a::{euclidean_space} set" assumes "bounded S" and conn: "connected(- S)" and 1: "DIM('a) = 1" shows "S = {}" proof - have "DIM('a) = DIM(real)" by (simp add: "1") then obtain f::"'a \ real" and g where "linear f" "\x. norm(f x) = norm x" and fg: "\x. g(f x) = x" "\y. f(g y) = y" by (rule isomorphisms_UNIV_UNIV) blast with \bounded S\ have "bounded (f ` S)" using bounded_linear_image linear_linear by blast have "bij f" by (metis fg bijI') have "connected (f ` (-S))" using connected_linear_image assms \linear f\ by blast moreover have "f ` (-S) = - (f ` S)" by (simp add: \bij f\ bij_image_Compl_eq) finally have "connected (- (f ` S))" by simp then have "f ` S = {}" using \bounded (f ` S)\ bounded_connected_Compl_real by blast then show ?thesis by blast qed lemma connected_card_eq_iff_nontrivial: fixes S :: "'a::metric_space set" shows "connected S \ uncountable S \ \(\a. S \ {a})" by (metis connected_uncountable finite.emptyI finite.insertI rev_finite_subset singleton_iff subsetI uncountable_infinite) lemma connected_finite_iff_sing: fixes S :: "'a::metric_space set" assumes "connected S" shows "finite S \ S = {} \ (\a. S = {a})" using assms connected_uncountable countable_finite by blast subsection\<^marker>\tag unimportant\\ Some simple positive connection theorems\ proposition path_connected_convex_diff_countable: fixes U :: "'a::euclidean_space set" assumes "convex U" "\ collinear U" "countable S" shows "path_connected(U - S)" proof (clarsimp simp: path_connected_def) fix a b assume "a \ U" "a \ S" "b \ U" "b \ S" let ?m = "midpoint a b" show "\g. path g \ path_image g \ U - S \ pathstart g = a \ pathfinish g = b" proof (cases "a = b") case True then show ?thesis by (metis DiffI \a \ U\ \a \ S\ path_component_def path_component_refl) next case False then have "a \ ?m" "b \ ?m" using midpoint_eq_endpoint by fastforce+ have "?m \ U" using \a \ U\ \b \ U\ \convex U\ convex_contains_segment by force obtain c where "c \ U" and nc_abc: "\ collinear {a,b,c}" by (metis False \a \ U\ \b \ U\ \\ collinear U\ collinear_triples insert_absorb) have ncoll_mca: "\ collinear {?m,c,a}" by (metis (full_types) \a \ ?m\ collinear_3_trans collinear_midpoint insert_commute nc_abc) have ncoll_mcb: "\ collinear {?m,c,b}" by (metis (full_types) \b \ ?m\ collinear_3_trans collinear_midpoint insert_commute nc_abc) have "c \ ?m" by (metis collinear_midpoint insert_commute nc_abc) then have "closed_segment ?m c \ U" by (simp add: \c \ U\ \?m \ U\ \convex U\ closed_segment_subset) then obtain z where z: "z \ closed_segment ?m c" and disjS: "(closed_segment a z \ closed_segment z b) \ S = {}" proof - have False if "closed_segment ?m c \ {z. (closed_segment a z \ closed_segment z b) \ S \ {}}" proof - have closb: "closed_segment ?m c \ {z \ closed_segment ?m c. closed_segment a z \ S \ {}} \ {z \ closed_segment ?m c. closed_segment z b \ S \ {}}" using that by blast have *: "countable {z \ closed_segment ?m c. closed_segment z u \ S \ {}}" if "u \ U" "u \ S" and ncoll: "\ collinear {?m, c, u}" for u proof - have **: False if x1: "x1 \ closed_segment ?m c" and x2: "x2 \ closed_segment ?m c" and "x1 \ x2" "x1 \ u" and w: "w \ closed_segment x1 u" "w \ closed_segment x2 u" and "w \ S" for x1 x2 w proof - have "x1 \ affine hull {?m,c}" "x2 \ affine hull {?m,c}" using segment_as_ball x1 x2 by auto then have coll_x1: "collinear {x1, ?m, c}" and coll_x2: "collinear {?m, c, x2}" by (simp_all add: affine_hull_3_imp_collinear) (metis affine_hull_3_imp_collinear insert_commute) have "\ collinear {x1, u, x2}" proof assume "collinear {x1, u, x2}" then have "collinear {?m, c, u}" by (metis (full_types) \c \ ?m\ coll_x1 coll_x2 collinear_3_trans insert_commute ncoll \x1 \ x2\) with ncoll show False .. qed then have "closed_segment x1 u \ closed_segment u x2 = {u}" by (blast intro!: Int_closed_segment) then have "w = u" using closed_segment_commute w by auto show ?thesis using \u \ S\ \w = u\ that(7) by auto qed then have disj: "disjoint ((\z\closed_segment ?m c. {closed_segment z u \ S}))" by (fastforce simp: pairwise_def disjnt_def) have cou: "countable ((\z \ closed_segment ?m c. {closed_segment z u \ S}) - {{}})" apply (rule pairwise_disjnt_countable_Union [OF _ pairwise_subset [OF disj]]) apply (rule countable_subset [OF _ \countable S\], auto) done define f where "f \ \X. (THE z. z \ closed_segment ?m c \ X = closed_segment z u \ S)" show ?thesis proof (rule countable_subset [OF _ countable_image [OF cou, where f=f]], clarify) fix x assume x: "x \ closed_segment ?m c" "closed_segment x u \ S \ {}" show "x \ f ` ((\z\closed_segment ?m c. {closed_segment z u \ S}) - {{}})" proof (rule_tac x="closed_segment x u \ S" in image_eqI) show "x = f (closed_segment x u \ S)" unfolding f_def by (rule the_equality [symmetric]) (use x in \auto dest: **\) qed (use x in auto) qed qed have "uncountable (closed_segment ?m c)" by (metis \c \ ?m\ uncountable_closed_segment) then show False using closb * [OF \a \ U\ \a \ S\ ncoll_mca] * [OF \b \ U\ \b \ S\ ncoll_mcb] by (simp add: closed_segment_commute countable_subset) qed then show ?thesis by (force intro: that) qed show ?thesis proof (intro exI conjI) have "path_image (linepath a z +++ linepath z b) \ U" by (metis \a \ U\ \b \ U\ \closed_segment ?m c \ U\ z \convex U\ closed_segment_subset contra_subsetD path_image_linepath subset_path_image_join) with disjS show "path_image (linepath a z +++ linepath z b) \ U - S" by (force simp: path_image_join) qed auto qed qed corollary connected_convex_diff_countable: fixes U :: "'a::euclidean_space set" assumes "convex U" "\ collinear U" "countable S" shows "connected(U - S)" by (simp add: assms path_connected_convex_diff_countable path_connected_imp_connected) lemma path_connected_punctured_convex: assumes "convex S" and aff: "aff_dim S \ 1" shows "path_connected(S - {a})" proof - consider "aff_dim S = -1" | "aff_dim S = 0" | "aff_dim S \ 2" using assms aff_dim_geq [of S] by linarith then show ?thesis proof cases assume "aff_dim S = -1" then show ?thesis by (metis aff_dim_empty empty_Diff path_connected_empty) next assume "aff_dim S = 0" then show ?thesis by (metis aff_dim_eq_0 Diff_cancel Diff_empty Diff_insert0 convex_empty convex_imp_path_connected path_connected_singleton singletonD) next assume ge2: "aff_dim S \ 2" then have "\ collinear S" proof (clarsimp simp: collinear_affine_hull) fix u v assume "S \ affine hull {u, v}" then have "aff_dim S \ aff_dim {u, v}" by (metis (no_types) aff_dim_affine_hull aff_dim_subset) with ge2 show False by (metis (no_types) aff_dim_2 antisym aff not_numeral_le_zero one_le_numeral order_trans) qed moreover have "countable {a}" by simp ultimately show ?thesis by (metis path_connected_convex_diff_countable [OF \convex S\]) qed qed lemma connected_punctured_convex: shows "\convex S; aff_dim S \ 1\ \ connected(S - {a})" using path_connected_imp_connected path_connected_punctured_convex by blast lemma path_connected_complement_countable: fixes S :: "'a::euclidean_space set" assumes "2 \ DIM('a)" "countable S" shows "path_connected(- S)" proof - have "\ collinear (UNIV::'a set)" using assms by (auto simp: collinear_aff_dim [of "UNIV :: 'a set"]) then have "path_connected(UNIV - S)" by (simp add: \countable S\ path_connected_convex_diff_countable) then show ?thesis by (simp add: Compl_eq_Diff_UNIV) qed proposition path_connected_openin_diff_countable: fixes S :: "'a::euclidean_space set" assumes "connected S" and ope: "openin (top_of_set (affine hull S)) S" and "\ collinear S" "countable T" shows "path_connected(S - T)" proof (clarsimp simp: path_connected_component) fix x y assume xy: "x \ S" "x \ T" "y \ S" "y \ T" show "path_component (S - T) x y" proof (rule connected_equivalence_relation_gen [OF \connected S\, where P = "\x. x \ T"]) show "\z. z \ U \ z \ T" if opeU: "openin (top_of_set S) U" and "x \ U" for U x proof - have "openin (top_of_set (affine hull S)) U" using opeU ope openin_trans by blast with \x \ U\ obtain r where Usub: "U \ affine hull S" and "r > 0" and subU: "ball x r \ affine hull S \ U" by (auto simp: openin_contains_ball) with \x \ U\ have x: "x \ ball x r \ affine hull S" by auto have "\ S \ {x}" using \\ collinear S\ collinear_subset by blast then obtain x' where "x' \ x" "x' \ S" by blast obtain y where y: "y \ x" "y \ ball x r \ affine hull S" proof show "x + (r / 2 / norm(x' - x)) *\<^sub>R (x' - x) \ x" using \x' \ x\ \r > 0\ by auto show "x + (r / 2 / norm (x' - x)) *\<^sub>R (x' - x) \ ball x r \ affine hull S" using \x' \ x\ \r > 0\ \x' \ S\ x by (simp add: dist_norm mem_affine_3_minus hull_inc) qed have "convex (ball x r \ affine hull S)" by (simp add: affine_imp_convex convex_Int) with x y subU have "uncountable U" by (meson countable_subset uncountable_convex) then have "\ U \ T" using \countable T\ countable_subset by blast then show ?thesis by blast qed show "\U. openin (top_of_set S) U \ x \ U \ (\x\U. \y\U. x \ T \ y \ T \ path_component (S - T) x y)" if "x \ S" for x proof - obtain r where Ssub: "S \ affine hull S" and "r > 0" and subS: "ball x r \ affine hull S \ S" using ope \x \ S\ by (auto simp: openin_contains_ball) then have conv: "convex (ball x r \ affine hull S)" by (simp add: affine_imp_convex convex_Int) have "\ aff_dim (affine hull S) \ 1" using \\ collinear S\ collinear_aff_dim by auto then have "\ aff_dim (ball x r \ affine hull S) \ 1" by (metis (no_types, opaque_lifting) aff_dim_convex_Int_open IntI open_ball \0 < r\ aff_dim_affine_hull affine_affine_hull affine_imp_convex centre_in_ball empty_iff hull_subset inf_commute subsetCE that) then have "\ collinear (ball x r \ affine hull S)" by (simp add: collinear_aff_dim) then have *: "path_connected ((ball x r \ affine hull S) - T)" by (rule path_connected_convex_diff_countable [OF conv _ \countable T\]) have ST: "ball x r \ affine hull S - T \ S - T" using subS by auto show ?thesis proof (intro exI conjI) show "x \ ball x r \ affine hull S" using \x \ S\ \r > 0\ by (simp add: hull_inc) have "openin (top_of_set (affine hull S)) (ball x r \ affine hull S)" by (subst inf.commute) (simp add: openin_Int_open) then show "openin (top_of_set S) (ball x r \ affine hull S)" by (rule openin_subset_trans [OF _ subS Ssub]) qed (use * path_component_trans in \auto simp: path_connected_component path_component_of_subset [OF ST]\) qed qed (use xy path_component_trans in auto) qed corollary connected_openin_diff_countable: fixes S :: "'a::euclidean_space set" assumes "connected S" and ope: "openin (top_of_set (affine hull S)) S" and "\ collinear S" "countable T" shows "connected(S - T)" by (metis path_connected_imp_connected path_connected_openin_diff_countable [OF assms]) corollary path_connected_open_diff_countable: fixes S :: "'a::euclidean_space set" assumes "2 \ DIM('a)" "open S" "connected S" "countable T" shows "path_connected(S - T)" proof (cases "S = {}") case True then show ?thesis by (simp) next case False show ?thesis proof (rule path_connected_openin_diff_countable) show "openin (top_of_set (affine hull S)) S" by (simp add: assms hull_subset open_subset) show "\ collinear S" using assms False by (simp add: collinear_aff_dim aff_dim_open) qed (simp_all add: assms) qed corollary connected_open_diff_countable: fixes S :: "'a::euclidean_space set" assumes "2 \ DIM('a)" "open S" "connected S" "countable T" shows "connected(S - T)" by (simp add: assms path_connected_imp_connected path_connected_open_diff_countable) subsection\<^marker>\tag unimportant\ \Self-homeomorphisms shuffling points about\ subsubsection\<^marker>\tag unimportant\\The theorem \homeomorphism_moving_points_exists\\ lemma homeomorphism_moving_point_1: fixes a :: "'a::euclidean_space" assumes "affine T" "a \ T" and u: "u \ ball a r \ T" obtains f g where "homeomorphism (cball a r \ T) (cball a r \ T) f g" "f a = u" "\x. x \ sphere a r \ f x = x" proof - have nou: "norm (u - a) < r" and "u \ T" using u by (auto simp: dist_norm norm_minus_commute) then have "0 < r" by (metis DiffD1 Diff_Diff_Int ball_eq_empty centre_in_ball not_le u) define f where "f \ \x. (1 - norm(x - a) / r) *\<^sub>R (u - a) + x" have *: "False" if eq: "x + (norm y / r) *\<^sub>R u = y + (norm x / r) *\<^sub>R u" and nou: "norm u < r" and yx: "norm y < norm x" for x y and u::'a proof - have "x = y + (norm x / r - (norm y / r)) *\<^sub>R u" using eq by (simp add: algebra_simps) then have "norm x = norm (y + ((norm x - norm y) / r) *\<^sub>R u)" by (metis diff_divide_distrib) also have "\ \ norm y + norm(((norm x - norm y) / r) *\<^sub>R u)" using norm_triangle_ineq by blast also have "\ = norm y + (norm x - norm y) * (norm u / r)" using yx \r > 0\ by (simp add: field_split_simps) also have "\ < norm y + (norm x - norm y) * 1" proof (subst add_less_cancel_left) show "(norm x - norm y) * (norm u / r) < (norm x - norm y) * 1" proof (rule mult_strict_left_mono) show "norm u / r < 1" using \0 < r\ divide_less_eq_1_pos nou by blast qed (simp add: yx) qed also have "\ = norm x" by simp finally show False by simp qed have "inj f" unfolding f_def proof (clarsimp simp: inj_on_def) fix x y assume "(1 - norm (x - a) / r) *\<^sub>R (u - a) + x = (1 - norm (y - a) / r) *\<^sub>R (u - a) + y" then have eq: "(x - a) + (norm (y - a) / r) *\<^sub>R (u - a) = (y - a) + (norm (x - a) / r) *\<^sub>R (u - a)" by (auto simp: algebra_simps) show "x=y" proof (cases "norm (x - a) = norm (y - a)") case True then show ?thesis using eq by auto next case False then consider "norm (x - a) < norm (y - a)" | "norm (x - a) > norm (y - a)" by linarith then have "False" proof cases case 1 show False using * [OF _ nou 1] eq by simp next case 2 with * [OF eq nou] show False by auto qed then show "x=y" .. qed qed then have inj_onf: "inj_on f (cball a r \ T)" using inj_on_Int by fastforce have contf: "continuous_on (cball a r \ T) f" unfolding f_def using \0 < r\ by (intro continuous_intros) blast have fim: "f ` (cball a r \ T) = cball a r \ T" proof have *: "norm (y + (1 - norm y / r) *\<^sub>R u) \ r" if "norm y \ r" "norm u < r" for y u::'a proof - have "norm (y + (1 - norm y / r) *\<^sub>R u) \ norm y + norm((1 - norm y / r) *\<^sub>R u)" using norm_triangle_ineq by blast also have "\ = norm y + abs(1 - norm y / r) * norm u" by simp also have "\ \ r" proof - have "(r - norm u) * (r - norm y) \ 0" using that by auto then have "r * norm u + r * norm y \ r * r + norm u * norm y" by (simp add: algebra_simps) then show ?thesis using that \0 < r\ by (simp add: abs_if field_simps) qed finally show ?thesis . qed have "f ` (cball a r) \ cball a r" using * nou apply (clarsimp simp: dist_norm norm_minus_commute f_def) by (metis diff_add_eq diff_diff_add diff_diff_eq2 norm_minus_commute) moreover have "f ` T \ T" unfolding f_def using \affine T\ \a \ T\ \u \ T\ by (force simp: add.commute mem_affine_3_minus) ultimately show "f ` (cball a r \ T) \ cball a r \ T" by blast next show "cball a r \ T \ f ` (cball a r \ T)" proof (clarsimp simp: dist_norm norm_minus_commute) fix x assume x: "norm (x - a) \ r" and "x \ T" have "\v \ {0..1}. ((1 - v) * r - norm ((x - a) - v *\<^sub>R (u - a))) \ 1 = 0" by (rule ivt_decreasing_component_on_1) (auto simp: x continuous_intros) then obtain v where "0 \ v" "v \ 1" and v: "(1 - v) * r = norm ((x - a) - v *\<^sub>R (u - a))" by auto then have n: "norm (a - (x - v *\<^sub>R (u - a))) = r - r * v" by (simp add: field_simps norm_minus_commute) show "x \ f ` (cball a r \ T)" proof (rule image_eqI) show "x = f (x - v *\<^sub>R (u - a))" using \r > 0\ v by (simp add: f_def) (simp add: field_simps) have "x - v *\<^sub>R (u - a) \ cball a r" using \r > 0\\0 \ v\ by (simp add: dist_norm n) moreover have "x - v *\<^sub>R (u - a) \ T" by (simp add: f_def \u \ T\ \x \ T\ assms mem_affine_3_minus2) ultimately show "x - v *\<^sub>R (u - a) \ cball a r \ T" by blast qed qed qed have "compact (cball a r \ T)" by (simp add: affine_closed compact_Int_closed \affine T\) then obtain g where "homeomorphism (cball a r \ T) (cball a r \ T) f g" by (metis homeomorphism_compact [OF _ contf fim inj_onf]) then show thesis apply (rule_tac f=f in that) using \r > 0\ by (simp_all add: f_def dist_norm norm_minus_commute) qed corollary\<^marker>\tag unimportant\ homeomorphism_moving_point_2: fixes a :: "'a::euclidean_space" assumes "affine T" "a \ T" and u: "u \ ball a r \ T" and v: "v \ ball a r \ T" obtains f g where "homeomorphism (cball a r \ T) (cball a r \ T) f g" "f u = v" "\x. \x \ sphere a r; x \ T\ \ f x = x" proof - have "0 < r" by (metis DiffD1 Diff_Diff_Int ball_eq_empty centre_in_ball not_le u) obtain f1 g1 where hom1: "homeomorphism (cball a r \ T) (cball a r \ T) f1 g1" and "f1 a = u" and f1: "\x. x \ sphere a r \ f1 x = x" using homeomorphism_moving_point_1 [OF \affine T\ \a \ T\ u] by blast obtain f2 g2 where hom2: "homeomorphism (cball a r \ T) (cball a r \ T) f2 g2" and "f2 a = v" and f2: "\x. x \ sphere a r \ f2 x = x" using homeomorphism_moving_point_1 [OF \affine T\ \a \ T\ v] by blast show ?thesis proof show "homeomorphism (cball a r \ T) (cball a r \ T) (f2 \ g1) (f1 \ g2)" by (metis homeomorphism_compose homeomorphism_symD hom1 hom2) have "g1 u = a" using \0 < r\ \f1 a = u\ assms hom1 homeomorphism_apply1 by fastforce then show "(f2 \ g1) u = v" by (simp add: \f2 a = v\) show "\x. \x \ sphere a r; x \ T\ \ (f2 \ g1) x = x" using f1 f2 hom1 homeomorphism_apply1 by fastforce qed qed corollary\<^marker>\tag unimportant\ homeomorphism_moving_point_3: fixes a :: "'a::euclidean_space" assumes "affine T" "a \ T" and ST: "ball a r \ T \ S" "S \ T" and u: "u \ ball a r \ T" and v: "v \ ball a r \ T" obtains f g where "homeomorphism S S f g" "f u = v" "{x. \ (f x = x \ g x = x)} \ ball a r \ T" proof - obtain f g where hom: "homeomorphism (cball a r \ T) (cball a r \ T) f g" and "f u = v" and fid: "\x. \x \ sphere a r; x \ T\ \ f x = x" using homeomorphism_moving_point_2 [OF \affine T\ \a \ T\ u v] by blast have gid: "\x. \x \ sphere a r; x \ T\ \ g x = x" using fid hom homeomorphism_apply1 by fastforce define ff where "ff \ \x. if x \ ball a r \ T then f x else x" define gg where "gg \ \x. if x \ ball a r \ T then g x else x" show ?thesis proof show "homeomorphism S S ff gg" proof (rule homeomorphismI) have "continuous_on ((cball a r \ T) \ (T - ball a r)) ff" unfolding ff_def using homeomorphism_cont1 [OF hom] by (intro continuous_on_cases) (auto simp: affine_closed \affine T\ fid) then show "continuous_on S ff" by (rule continuous_on_subset) (use ST in auto) have "continuous_on ((cball a r \ T) \ (T - ball a r)) gg" unfolding gg_def using homeomorphism_cont2 [OF hom] by (intro continuous_on_cases) (auto simp: affine_closed \affine T\ gid) then show "continuous_on S gg" by (rule continuous_on_subset) (use ST in auto) show "ff ` S \ S" proof (clarsimp simp: ff_def) fix x assume "x \ S" and x: "dist a x < r" and "x \ T" then have "f x \ cball a r \ T" using homeomorphism_image1 [OF hom] by force then show "f x \ S" using ST(1) \x \ T\ gid hom homeomorphism_def x by fastforce qed show "gg ` S \ S" proof (clarsimp simp: gg_def) fix x assume "x \ S" and x: "dist a x < r" and "x \ T" then have "g x \ cball a r \ T" using homeomorphism_image2 [OF hom] by force then have "g x \ ball a r" using homeomorphism_apply2 [OF hom] by (metis Diff_Diff_Int Diff_iff \x \ T\ cball_def fid le_less mem_Collect_eq mem_ball mem_sphere x) then show "g x \ S" using ST(1) \g x \ cball a r \ T\ by force qed show "\x. x \ S \ gg (ff x) = x" unfolding ff_def gg_def using homeomorphism_apply1 [OF hom] homeomorphism_image1 [OF hom] by simp (metis Int_iff homeomorphism_apply1 [OF hom] fid image_eqI less_eq_real_def mem_cball mem_sphere) show "\x. x \ S \ ff (gg x) = x" unfolding ff_def gg_def using homeomorphism_apply2 [OF hom] homeomorphism_image2 [OF hom] by simp (metis Int_iff fid image_eqI less_eq_real_def mem_cball mem_sphere) qed show "ff u = v" using u by (auto simp: ff_def \f u = v\) show "{x. \ (ff x = x \ gg x = x)} \ ball a r \ T" by (auto simp: ff_def gg_def) qed qed proposition\<^marker>\tag unimportant\ homeomorphism_moving_point: fixes a :: "'a::euclidean_space" assumes ope: "openin (top_of_set (affine hull S)) S" and "S \ T" and TS: "T \ affine hull S" and S: "connected S" "a \ S" "b \ S" obtains f g where "homeomorphism T T f g" "f a = b" "{x. \ (f x = x \ g x = x)} \ S" "bounded {x. \ (f x = x \ g x = x)}" proof - have 1: "\h k. homeomorphism T T h k \ h (f d) = d \ {x. \ (h x = x \ k x = x)} \ S \ bounded {x. \ (h x = x \ k x = x)}" if "d \ S" "f d \ S" and homfg: "homeomorphism T T f g" and S: "{x. \ (f x = x \ g x = x)} \ S" and bo: "bounded {x. \ (f x = x \ g x = x)}" for d f g proof (intro exI conjI) show homgf: "homeomorphism T T g f" by (metis homeomorphism_symD homfg) then show "g (f d) = d" by (meson \S \ T\ homeomorphism_def subsetD \d \ S\) show "{x. \ (g x = x \ f x = x)} \ S" using S by blast show "bounded {x. \ (g x = x \ f x = x)}" using bo by (simp add: conj_commute) qed have 2: "\f g. homeomorphism T T f g \ f x = f2 (f1 x) \ {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" if "x \ S" "f1 x \ S" "f2 (f1 x) \ S" and hom: "homeomorphism T T f1 g1" "homeomorphism T T f2 g2" and sub: "{x. \ (f1 x = x \ g1 x = x)} \ S" "{x. \ (f2 x = x \ g2 x = x)} \ S" and bo: "bounded {x. \ (f1 x = x \ g1 x = x)}" "bounded {x. \ (f2 x = x \ g2 x = x)}" for x f1 f2 g1 g2 proof (intro exI conjI) show homgf: "homeomorphism T T (f2 \ f1) (g1 \ g2)" by (metis homeomorphism_compose hom) then show "(f2 \ f1) x = f2 (f1 x)" by force show "{x. \ ((f2 \ f1) x = x \ (g1 \ g2) x = x)} \ S" using sub by force have "bounded ({x. \(f1 x = x \ g1 x = x)} \ {x. \(f2 x = x \ g2 x = x)})" using bo by simp then show "bounded {x. \ ((f2 \ f1) x = x \ (g1 \ g2) x = x)}" by (rule bounded_subset) auto qed have 3: "\U. openin (top_of_set S) U \ d \ U \ (\x\U. \f g. homeomorphism T T f g \ f d = x \ {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)})" if "d \ S" for d proof - obtain r where "r > 0" and r: "ball d r \ affine hull S \ S" by (metis \d \ S\ ope openin_contains_ball) have *: "\f g. homeomorphism T T f g \ f d = e \ {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" if "e \ S" "e \ ball d r" for e apply (rule homeomorphism_moving_point_3 [of "affine hull S" d r T d e]) using r \S \ T\ TS that apply (auto simp: \d \ S\ \0 < r\ hull_inc) using bounded_subset by blast show ?thesis by (rule_tac x="S \ ball d r" in exI) (fastforce simp: openin_open_Int \0 < r\ that intro: *) qed have "\f g. homeomorphism T T f g \ f a = b \ {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" by (rule connected_equivalence_relation [OF S]; blast intro: 1 2 3) then show ?thesis using that by auto qed lemma homeomorphism_moving_points_exists_gen: assumes K: "finite K" "\i. i \ K \ x i \ S \ y i \ S" "pairwise (\i j. (x i \ x j) \ (y i \ y j)) K" and "2 \ aff_dim S" and ope: "openin (top_of_set (affine hull S)) S" and "S \ T" "T \ affine hull S" "connected S" shows "\f g. homeomorphism T T f g \ (\i \ K. f(x i) = y i) \ {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" using assms proof (induction K) case empty then show ?case by (force simp: homeomorphism_ident) next case (insert i K) then have xney: "\j. \j \ K; j \ i\ \ x i \ x j \ y i \ y j" and pw: "pairwise (\i j. x i \ x j \ y i \ y j) K" and "x i \ S" "y i \ S" and xyS: "\i. i \ K \ x i \ S \ y i \ S" by (simp_all add: pairwise_insert) obtain f g where homfg: "homeomorphism T T f g" and feq: "\i. i \ K \ f(x i) = y i" and fg_sub: "{x. \ (f x = x \ g x = x)} \ S" and bo_fg: "bounded {x. \ (f x = x \ g x = x)}" using insert.IH [OF xyS pw] insert.prems by (blast intro: that) then have "\f g. homeomorphism T T f g \ (\i \ K. f(x i) = y i) \ {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" using insert by blast have aff_eq: "affine hull (S - y ` K) = affine hull S" proof (rule affine_hull_Diff [OF ope]) show "finite (y ` K)" by (simp add: insert.hyps(1)) show "y ` K \ S" using \y i \ S\ insert.hyps(2) xney xyS by fastforce qed have f_in_S: "f x \ S" if "x \ S" for x using homfg fg_sub homeomorphism_apply1 \S \ T\ proof - have "(f (f x) \ f x \ g (f x) \ f x) \ f x \ S" by (metis \S \ T\ homfg subsetD homeomorphism_apply1 that) then show ?thesis using fg_sub by force qed obtain h k where homhk: "homeomorphism T T h k" and heq: "h (f (x i)) = y i" and hk_sub: "{x. \ (h x = x \ k x = x)} \ S - y ` K" and bo_hk: "bounded {x. \ (h x = x \ k x = x)}" proof (rule homeomorphism_moving_point [of "S - y`K" T "f(x i)" "y i"]) show "openin (top_of_set (affine hull (S - y ` K))) (S - y ` K)" by (simp add: aff_eq openin_diff finite_imp_closedin image_subset_iff hull_inc insert xyS) show "S - y ` K \ T" using \S \ T\ by auto show "T \ affine hull (S - y ` K)" using insert by (simp add: aff_eq) show "connected (S - y ` K)" proof (rule connected_openin_diff_countable [OF \connected S\ ope]) show "\ collinear S" using collinear_aff_dim \2 \ aff_dim S\ by force show "countable (y ` K)" using countable_finite insert.hyps(1) by blast qed have "\k. \f (x i) = y k; k \ K\ \ False" by (metis feq homfg \x i \ S\ homeomorphism_def \S \ T\ \i \ K\ subsetCE xney xyS) then show "f (x i) \ S - y ` K" by (auto simp: f_in_S \x i \ S\) show "y i \ S - y ` K" using insert.hyps xney by (auto simp: \y i \ S\) qed blast show ?case proof (intro exI conjI) show "homeomorphism T T (h \ f) (g \ k)" using homfg homhk homeomorphism_compose by blast show "\i \ insert i K. (h \ f) (x i) = y i" using feq hk_sub by (auto simp: heq) show "{x. \ ((h \ f) x = x \ (g \ k) x = x)} \ S" using fg_sub hk_sub by force have "bounded ({x. \(f x = x \ g x = x)} \ {x. \(h x = x \ k x = x)})" using bo_fg bo_hk bounded_Un by blast then show "bounded {x. \ ((h \ f) x = x \ (g \ k) x = x)}" by (rule bounded_subset) auto qed qed proposition\<^marker>\tag unimportant\ homeomorphism_moving_points_exists: fixes S :: "'a::euclidean_space set" assumes 2: "2 \ DIM('a)" "open S" "connected S" "S \ T" "finite K" and KS: "\i. i \ K \ x i \ S \ y i \ S" and pw: "pairwise (\i j. (x i \ x j) \ (y i \ y j)) K" and S: "S \ T" "T \ affine hull S" "connected S" obtains f g where "homeomorphism T T f g" "\i. i \ K \ f(x i) = y i" "{x. \ (f x = x \ g x = x)} \ S" "bounded {x. (\ (f x = x \ g x = x))}" proof (cases "S = {}") case True then show ?thesis using KS homeomorphism_ident that by fastforce next case False then have affS: "affine hull S = UNIV" by (simp add: affine_hull_open \open S\) then have ope: "openin (top_of_set (affine hull S)) S" using \open S\ open_openin by auto have "2 \ DIM('a)" by (rule 2) also have "\ = aff_dim (UNIV :: 'a set)" by simp also have "\ \ aff_dim S" by (metis aff_dim_UNIV aff_dim_affine_hull aff_dim_le_DIM affS) finally have "2 \ aff_dim S" by linarith then show ?thesis using homeomorphism_moving_points_exists_gen [OF \finite K\ KS pw _ ope S] that by fastforce qed subsubsection\<^marker>\tag unimportant\\The theorem \homeomorphism_grouping_points_exists\\ lemma homeomorphism_grouping_point_1: fixes a::real and c::real assumes "a < b" "c < d" obtains f g where "homeomorphism (cbox a b) (cbox c d) f g" "f a = c" "f b = d" proof - define f where "f \ \x. ((d - c) / (b - a)) * x + (c - a * ((d - c) / (b - a)))" have "\g. homeomorphism (cbox a b) (cbox c d) f g" proof (rule homeomorphism_compact) show "continuous_on (cbox a b) f" unfolding f_def by (intro continuous_intros) have "f ` {a..b} = {c..d}" unfolding f_def image_affinity_atLeastAtMost using assms sum_sqs_eq by (auto simp: field_split_simps) then show "f ` cbox a b = cbox c d" by auto show "inj_on f (cbox a b)" unfolding f_def inj_on_def using assms by auto qed auto then obtain g where "homeomorphism (cbox a b) (cbox c d) f g" .. then show ?thesis proof show "f a = c" by (simp add: f_def) show "f b = d" using assms sum_sqs_eq [of a b] by (auto simp: f_def field_split_simps) qed qed lemma homeomorphism_grouping_point_2: fixes a::real and w::real assumes hom_ab: "homeomorphism (cbox a b) (cbox u v) f1 g1" and hom_bc: "homeomorphism (cbox b c) (cbox v w) f2 g2" and "b \ cbox a c" "v \ cbox u w" and eq: "f1 a = u" "f1 b = v" "f2 b = v" "f2 c = w" obtains f g where "homeomorphism (cbox a c) (cbox u w) f g" "f a = u" "f c = w" "\x. x \ cbox a b \ f x = f1 x" "\x. x \ cbox b c \ f x = f2 x" proof - have le: "a \ b" "b \ c" "u \ v" "v \ w" using assms by simp_all then have ac: "cbox a c = cbox a b \ cbox b c" and uw: "cbox u w = cbox u v \ cbox v w" by auto define f where "f \ \x. if x \ b then f1 x else f2 x" have "\g. homeomorphism (cbox a c) (cbox u w) f g" proof (rule homeomorphism_compact) have cf1: "continuous_on (cbox a b) f1" using hom_ab homeomorphism_cont1 by blast have cf2: "continuous_on (cbox b c) f2" using hom_bc homeomorphism_cont1 by blast show "continuous_on (cbox a c) f" unfolding f_def using le eq by (force intro: continuous_on_cases_le [OF continuous_on_subset [OF cf1] continuous_on_subset [OF cf2]]) have "f ` cbox a b = f1 ` cbox a b" "f ` cbox b c = f2 ` cbox b c" unfolding f_def using eq by force+ then show "f ` cbox a c = cbox u w" unfolding ac uw image_Un by (metis hom_ab hom_bc homeomorphism_def) have neq12: "f1 x \ f2 y" if x: "a \ x" "x \ b" and y: "b < y" "y \ c" for x y proof - have "f1 x \ cbox u v" by (metis hom_ab homeomorphism_def image_eqI mem_box_real(2) x) moreover have "f2 y \ cbox v w" by (metis (full_types) hom_bc homeomorphism_def image_subset_iff mem_box_real(2) not_le not_less_iff_gr_or_eq order_refl y) moreover have "f2 y \ f2 b" by (metis cancel_comm_monoid_add_class.diff_cancel diff_gt_0_iff_gt hom_bc homeomorphism_def le(2) less_imp_le less_numeral_extra(3) mem_box_real(2) order_refl y) ultimately show ?thesis using le eq by simp qed have "inj_on f1 (cbox a b)" by (metis (full_types) hom_ab homeomorphism_def inj_onI) moreover have "inj_on f2 (cbox b c)" by (metis (full_types) hom_bc homeomorphism_def inj_onI) ultimately show "inj_on f (cbox a c)" apply (simp (no_asm) add: inj_on_def) apply (simp add: f_def inj_on_eq_iff) using neq12 by force qed auto then obtain g where "homeomorphism (cbox a c) (cbox u w) f g" .. then show ?thesis using eq f_def le that by force qed lemma homeomorphism_grouping_point_3: fixes a::real assumes cbox_sub: "cbox c d \ box a b" "cbox u v \ box a b" and box_ne: "box c d \ {}" "box u v \ {}" obtains f g where "homeomorphism (cbox a b) (cbox a b) f g" "f a = a" "f b = b" "\x. x \ cbox c d \ f x \ cbox u v" proof - have less: "a < c" "a < u" "d < b" "v < b" "c < d" "u < v" "cbox c d \ {}" using assms by (simp_all add: cbox_sub subset_eq) obtain f1 g1 where 1: "homeomorphism (cbox a c) (cbox a u) f1 g1" and f1_eq: "f1 a = a" "f1 c = u" using homeomorphism_grouping_point_1 [OF \a < c\ \a < u\] . obtain f2 g2 where 2: "homeomorphism (cbox c d) (cbox u v) f2 g2" and f2_eq: "f2 c = u" "f2 d = v" using homeomorphism_grouping_point_1 [OF \c < d\ \u < v\] . obtain f3 g3 where 3: "homeomorphism (cbox d b) (cbox v b) f3 g3" and f3_eq: "f3 d = v" "f3 b = b" using homeomorphism_grouping_point_1 [OF \d < b\ \v < b\] . obtain f4 g4 where 4: "homeomorphism (cbox a d) (cbox a v) f4 g4" and "f4 a = a" "f4 d = v" and f4_eq: "\x. x \ cbox a c \ f4 x = f1 x" "\x. x \ cbox c d \ f4 x = f2 x" using homeomorphism_grouping_point_2 [OF 1 2] less by (auto simp: f1_eq f2_eq) obtain f g where fg: "homeomorphism (cbox a b) (cbox a b) f g" "f a = a" "f b = b" and f_eq: "\x. x \ cbox a d \ f x = f4 x" "\x. x \ cbox d b \ f x = f3 x" using homeomorphism_grouping_point_2 [OF 4 3] less by (auto simp: f4_eq f3_eq f2_eq f1_eq) show ?thesis proof (rule that [OF fg]) show "f x \ cbox u v" if "x \ cbox c d" for x using that f4_eq f_eq homeomorphism_image1 [OF 2] by (metis atLeastAtMost_iff box_real(2) image_eqI less(1) less_eq_real_def order_trans) qed qed lemma homeomorphism_grouping_point_4: fixes T :: "real set" assumes "open U" "open S" "connected S" "U \ {}" "finite K" "K \ S" "U \ S" "S \ T" obtains f g where "homeomorphism T T f g" "\x. x \ K \ f x \ U" "{x. (\ (f x = x \ g x = x))} \ S" "bounded {x. (\ (f x = x \ g x = x))}" proof - obtain c d where "box c d \ {}" "cbox c d \ U" proof - obtain u where "u \ U" using \U \ {}\ by blast then obtain e where "e > 0" "cball u e \ U" using \open U\ open_contains_cball by blast then show ?thesis by (rule_tac c=u and d="u+e" in that) (auto simp: dist_norm subset_iff) qed have "compact K" by (simp add: \finite K\ finite_imp_compact) obtain a b where "box a b \ {}" "K \ cbox a b" "cbox a b \ S" proof (cases "K = {}") case True then show ?thesis using \box c d \ {}\ \cbox c d \ U\ \U \ S\ that by blast next case False then obtain a b where "a \ K" "b \ K" and a: "\x. x \ K \ a \ x" and b: "\x. x \ K \ x \ b" using compact_attains_inf compact_attains_sup by (metis \compact K\)+ obtain e where "e > 0" "cball b e \ S" using \open S\ open_contains_cball by (metis \b \ K\ \K \ S\ subsetD) show ?thesis proof show "box a (b + e) \ {}" using \0 < e\ \b \ K\ a by force show "K \ cbox a (b + e)" using \0 < e\ a b by fastforce have "a \ S" using \a \ K\ assms(6) by blast have "b + e \ S" using \0 < e\ \cball b e \ S\ by (force simp: dist_norm) show "cbox a (b + e) \ S" using \a \ S\ \b + e \ S\ \connected S\ connected_contains_Icc by auto qed qed obtain w z where "cbox w z \ S" and sub_wz: "cbox a b \ cbox c d \ box w z" proof - have "a \ S" "b \ S" using \box a b \ {}\ \cbox a b \ S\ by auto moreover have "c \ S" "d \ S" using \box c d \ {}\ \cbox c d \ U\ \U \ S\ by force+ ultimately have "min a c \ S" "max b d \ S" by linarith+ then obtain e1 e2 where "e1 > 0" "cball (min a c) e1 \ S" "e2 > 0" "cball (max b d) e2 \ S" using \open S\ open_contains_cball by metis then have *: "min a c - e1 \ S" "max b d + e2 \ S" by (auto simp: dist_norm) show ?thesis proof show "cbox (min a c - e1) (max b d+ e2) \ S" using * \connected S\ connected_contains_Icc by auto show "cbox a b \ cbox c d \ box (min a c - e1) (max b d + e2)" using \0 < e1\ \0 < e2\ by auto qed qed then obtain f g where hom: "homeomorphism (cbox w z) (cbox w z) f g" and "f w = w" "f z = z" and fin: "\x. x \ cbox a b \ f x \ cbox c d" using homeomorphism_grouping_point_3 [of a b w z c d] using \box a b \ {}\ \box c d \ {}\ by blast have contfg: "continuous_on (cbox w z) f" "continuous_on (cbox w z) g" using hom homeomorphism_def by blast+ define f' where "f' \ \x. if x \ cbox w z then f x else x" define g' where "g' \ \x. if x \ cbox w z then g x else x" show ?thesis proof have T: "cbox w z \ (T - box w z) = T" using \cbox w z \ S\ \S \ T\ by auto show "homeomorphism T T f' g'" proof have clo: "closedin (top_of_set (cbox w z \ (T - box w z))) (T - box w z)" by (metis Diff_Diff_Int Diff_subset T closedin_def open_box openin_open_Int topspace_euclidean_subtopology) have "\x. \w \ x \ x \ z; w < x \ \ x < z\ \ f x = x" using \f w = w\ \f z = z\ by auto moreover have "\x. \w \ x \ x \ z; w < x \ \ x < z\ \ g x = x" using \f w = w\ \f z = z\ hom homeomorphism_apply1 by fastforce ultimately have "continuous_on (cbox w z \ (T - box w z)) f'" "continuous_on (cbox w z \ (T - box w z)) g'" unfolding f'_def g'_def by (intro continuous_on_cases_local contfg continuous_on_id clo; auto simp: closed_subset)+ then show "continuous_on T f'" "continuous_on T g'" by (simp_all only: T) show "f' ` T \ T" unfolding f'_def by clarsimp (metis \cbox w z \ S\ \S \ T\ subsetD hom homeomorphism_def imageI mem_box_real(2)) show "g' ` T \ T" unfolding g'_def by clarsimp (metis \cbox w z \ S\ \S \ T\ subsetD hom homeomorphism_def imageI mem_box_real(2)) show "\x. x \ T \ g' (f' x) = x" unfolding f'_def g'_def using homeomorphism_apply1 [OF hom] homeomorphism_image1 [OF hom] by fastforce show "\y. y \ T \ f' (g' y) = y" unfolding f'_def g'_def using homeomorphism_apply2 [OF hom] homeomorphism_image2 [OF hom] by fastforce qed show "\x. x \ K \ f' x \ U" using fin sub_wz \K \ cbox a b\ \cbox c d \ U\ by (force simp: f'_def) show "{x. \ (f' x = x \ g' x = x)} \ S" using \cbox w z \ S\ by (auto simp: f'_def g'_def) show "bounded {x. \ (f' x = x \ g' x = x)}" proof (rule bounded_subset [of "cbox w z"]) show "bounded (cbox w z)" using bounded_cbox by blast show "{x. \ (f' x = x \ g' x = x)} \ cbox w z" by (auto simp: f'_def g'_def) qed qed qed proposition\<^marker>\tag unimportant\ homeomorphism_grouping_points_exists: fixes S :: "'a::euclidean_space set" assumes "open U" "open S" "connected S" "U \ {}" "finite K" "K \ S" "U \ S" "S \ T" obtains f g where "homeomorphism T T f g" "{x. (\ (f x = x \ g x = x))} \ S" "bounded {x. (\ (f x = x \ g x = x))}" "\x. x \ K \ f x \ U" proof (cases "2 \ DIM('a)") case True have TS: "T \ affine hull S" using affine_hull_open assms by blast have "infinite U" using \open U\ \U \ {}\ finite_imp_not_open by blast then obtain P where "P \ U" "finite P" "card K = card P" using infinite_arbitrarily_large by metis then obtain \ where \: "bij_betw \ K P" using \finite K\ finite_same_card_bij by blast obtain f g where "homeomorphism T T f g" "\i. i \ K \ f (id i) = \ i" "{x. \ (f x = x \ g x = x)} \ S" "bounded {x. \ (f x = x \ g x = x)}" proof (rule homeomorphism_moving_points_exists [OF True \open S\ \connected S\ \S \ T\ \finite K\]) show "\i. i \ K \ id i \ S \ \ i \ S" using \P \ U\ \bij_betw \ K P\ \K \ S\ \U \ S\ bij_betwE by blast show "pairwise (\i j. id i \ id j \ \ i \ \ j) K" using \ by (auto simp: pairwise_def bij_betw_def inj_on_def) qed (use affine_hull_open assms that in auto) then show ?thesis using \ \P \ U\ bij_betwE by (fastforce simp: intro!: that) next case False with DIM_positive have "DIM('a) = 1" by (simp add: dual_order.antisym) then obtain h::"'a \real" and j where "linear h" "linear j" and noh: "\x. norm(h x) = norm x" and noj: "\y. norm(j y) = norm y" and hj: "\x. j(h x) = x" "\y. h(j y) = y" and ranh: "surj h" using isomorphisms_UNIV_UNIV by (metis (mono_tags, opaque_lifting) DIM_real UNIV_eq_I range_eqI) obtain f g where hom: "homeomorphism (h ` T) (h ` T) f g" and f: "\x. x \ h ` K \ f x \ h ` U" and sub: "{x. \ (f x = x \ g x = x)} \ h ` S" and bou: "bounded {x. \ (f x = x \ g x = x)}" apply (rule homeomorphism_grouping_point_4 [of "h ` U" "h ` S" "h ` K" "h ` T"]) by (simp_all add: assms image_mono \linear h\ open_surjective_linear_image connected_linear_image ranh) have jf: "j (f (h x)) = x \ f (h x) = h x" for x by (metis hj) have jg: "j (g (h x)) = x \ g (h x) = h x" for x by (metis hj) have cont_hj: "continuous_on X h" "continuous_on Y j" for X Y by (simp_all add: \linear h\ \linear j\ linear_linear linear_continuous_on) show ?thesis proof show "homeomorphism T T (j \ f \ h) (j \ g \ h)" proof show "continuous_on T (j \ f \ h)" "continuous_on T (j \ g \ h)" using hom homeomorphism_def by (blast intro: continuous_on_compose cont_hj)+ show "(j \ f \ h) ` T \ T" "(j \ g \ h) ` T \ T" by auto (metis (mono_tags, opaque_lifting) hj(1) hom homeomorphism_def imageE imageI)+ show "\x. x \ T \ (j \ g \ h) ((j \ f \ h) x) = x" using hj hom homeomorphism_apply1 by fastforce show "\y. y \ T \ (j \ f \ h) ((j \ g \ h) y) = y" using hj hom homeomorphism_apply2 by fastforce qed show "{x. \ ((j \ f \ h) x = x \ (j \ g \ h) x = x)} \ S" proof (clarsimp simp: jf jg hj) show "f (h x) = h x \ g (h x) \ h x \ x \ S" for x using sub [THEN subsetD, of "h x"] hj by simp (metis imageE) qed have "bounded (j ` {x. (\ (f x = x \ g x = x))})" by (rule bounded_linear_image [OF bou]) (use \linear j\ linear_conv_bounded_linear in auto) moreover have *: "{x. \((j \ f \ h) x = x \ (j \ g \ h) x = x)} = j ` {x. (\ (f x = x \ g x = x))}" using hj by (auto simp: jf jg image_iff, metis+) ultimately show "bounded {x. \ ((j \ f \ h) x = x \ (j \ g \ h) x = x)}" by metis show "\x. x \ K \ (j \ f \ h) x \ U" using f hj by fastforce qed qed proposition\<^marker>\tag unimportant\ homeomorphism_grouping_points_exists_gen: fixes S :: "'a::euclidean_space set" assumes opeU: "openin (top_of_set S) U" and opeS: "openin (top_of_set (affine hull S)) S" and "U \ {}" "finite K" "K \ S" and S: "S \ T" "T \ affine hull S" "connected S" obtains f g where "homeomorphism T T f g" "{x. (\ (f x = x \ g x = x))} \ S" "bounded {x. (\ (f x = x \ g x = x))}" "\x. x \ K \ f x \ U" proof (cases "2 \ aff_dim S") case True have opeU': "openin (top_of_set (affine hull S)) U" using opeS opeU openin_trans by blast obtain u where "u \ U" "u \ S" using \U \ {}\ opeU openin_imp_subset by fastforce+ have "infinite U" proof (rule infinite_openin [OF opeU \u \ U\]) show "u islimpt S" using True \u \ S\ assms(8) connected_imp_perfect_aff_dim by fastforce qed then obtain P where "P \ U" "finite P" "card K = card P" using infinite_arbitrarily_large by metis then obtain \ where \: "bij_betw \ K P" using \finite K\ finite_same_card_bij by blast have "\f g. homeomorphism T T f g \ (\i \ K. f(id i) = \ i) \ {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" proof (rule homeomorphism_moving_points_exists_gen [OF \finite K\ _ _ True opeS S]) show "\i. i \ K \ id i \ S \ \ i \ S" by (metis id_apply opeU openin_contains_cball subsetCE \P \ U\ \bij_betw \ K P\ \K \ S\ bij_betwE) show "pairwise (\i j. id i \ id j \ \ i \ \ j) K" using \ by (auto simp: pairwise_def bij_betw_def inj_on_def) qed then show ?thesis using \ \P \ U\ bij_betwE by (fastforce simp: intro!: that) next case False with aff_dim_geq [of S] consider "aff_dim S = -1" | "aff_dim S = 0" | "aff_dim S = 1" by linarith then show ?thesis proof cases assume "aff_dim S = -1" then have "S = {}" using aff_dim_empty by blast then have "False" using \U \ {}\ \K \ S\ openin_imp_subset [OF opeU] by blast then show ?thesis .. next assume "aff_dim S = 0" then obtain a where "S = {a}" using aff_dim_eq_0 by blast then have "K \ U" using \U \ {}\ \K \ S\ openin_imp_subset [OF opeU] by blast show ?thesis using \K \ U\ by (intro that [of id id]) (auto intro: homeomorphismI) next assume "aff_dim S = 1" then have "affine hull S homeomorphic (UNIV :: real set)" by (auto simp: homeomorphic_affine_sets) then obtain h::"'a\real" and j where homhj: "homeomorphism (affine hull S) UNIV h j" using homeomorphic_def by blast then have h: "\x. x \ affine hull S \ j(h(x)) = x" and j: "\y. j y \ affine hull S \ h(j y) = y" by (auto simp: homeomorphism_def) have connh: "connected (h ` S)" by (meson Topological_Spaces.connected_continuous_image \connected S\ homeomorphism_cont1 homeomorphism_of_subsets homhj hull_subset top_greatest) have hUS: "h ` U \ h ` S" by (meson homeomorphism_imp_open_map homeomorphism_of_subsets homhj hull_subset opeS opeU open_UNIV openin_open_eq) have opn: "openin (top_of_set (affine hull S)) U \ open (h ` U)" for U using homeomorphism_imp_open_map [OF homhj] by simp have "open (h ` U)" "open (h ` S)" by (auto intro: opeS opeU openin_trans opn) then obtain f g where hom: "homeomorphism (h ` T) (h ` T) f g" and f: "\x. x \ h ` K \ f x \ h ` U" and sub: "{x. \ (f x = x \ g x = x)} \ h ` S" and bou: "bounded {x. \ (f x = x \ g x = x)}" apply (rule homeomorphism_grouping_points_exists [of "h ` U" "h ` S" "h ` K" "h ` T"]) using assms by (auto simp: connh hUS) have jf: "\x. x \ affine hull S \ j (f (h x)) = x \ f (h x) = h x" by (metis h j) have jg: "\x. x \ affine hull S \ j (g (h x)) = x \ g (h x) = h x" by (metis h j) have cont_hj: "continuous_on T h" "continuous_on Y j" for Y proof (rule continuous_on_subset [OF _ \T \ affine hull S\]) show "continuous_on (affine hull S) h" using homeomorphism_def homhj by blast qed (meson continuous_on_subset homeomorphism_def homhj top_greatest) define f' where "f' \ \x. if x \ affine hull S then (j \ f \ h) x else x" define g' where "g' \ \x. if x \ affine hull S then (j \ g \ h) x else x" show ?thesis proof show "homeomorphism T T f' g'" proof have "continuous_on T (j \ f \ h)" using hom homeomorphism_def by (intro continuous_on_compose cont_hj) blast then show "continuous_on T f'" apply (rule continuous_on_eq) using \T \ affine hull S\ f'_def by auto have "continuous_on T (j \ g \ h)" using hom homeomorphism_def by (intro continuous_on_compose cont_hj) blast then show "continuous_on T g'" apply (rule continuous_on_eq) using \T \ affine hull S\ g'_def by auto show "f' ` T \ T" proof (clarsimp simp: f'_def) fix x assume "x \ T" then have "f (h x) \ h ` T" by (metis (no_types) hom homeomorphism_def image_subset_iff subset_refl) then show "j (f (h x)) \ T" using \T \ affine hull S\ h by auto qed show "g' ` T \ T" proof (clarsimp simp: g'_def) fix x assume "x \ T" then have "g (h x) \ h ` T" by (metis (no_types) hom homeomorphism_def image_subset_iff subset_refl) then show "j (g (h x)) \ T" using \T \ affine hull S\ h by auto qed show "\x. x \ T \ g' (f' x) = x" using h j hom homeomorphism_apply1 by (fastforce simp: f'_def g'_def) show "\y. y \ T \ f' (g' y) = y" using h j hom homeomorphism_apply2 by (fastforce simp: f'_def g'_def) qed next have \
: "\x y. \x \ affine hull S; h x = h y; y \ S\ \ x \ S" by (metis h hull_inc) show "{x. \ (f' x = x \ g' x = x)} \ S" using sub by (simp add: f'_def g'_def jf jg) (force elim: \
) next have "compact (j ` closure {x. \ (f x = x \ g x = x)})" using bou by (auto simp: compact_continuous_image cont_hj) then have "bounded (j ` {x. \ (f x = x \ g x = x)})" by (rule bounded_closure_image [OF compact_imp_bounded]) moreover have *: "{x \ affine hull S. j (f (h x)) \ x \ j (g (h x)) \ x} = j ` {x. (\ (f x = x \ g x = x))}" using h j by (auto simp: image_iff; metis) ultimately have "bounded {x \ affine hull S. j (f (h x)) \ x \ j (g (h x)) \ x}" by metis then show "bounded {x. \ (f' x = x \ g' x = x)}" by (simp add: f'_def g'_def Collect_mono bounded_subset) next show "f' x \ U" if "x \ K" for x proof - have "U \ S" using opeU openin_imp_subset by blast then have "j (f (h x)) \ U" using f h hull_subset that by fastforce then show "f' x \ U" using \K \ S\ S f'_def that by auto qed qed qed qed subsection\Nullhomotopic mappings\ text\ A mapping out of a sphere is nullhomotopic iff it extends to the ball. This even works out in the degenerate cases when the radius is \\\ 0, and we also don't need to explicitly assume continuity since it's already implicit in both sides of the equivalence.\ lemma nullhomotopic_from_lemma: assumes contg: "continuous_on (cball a r - {a}) g" and fa: "\e. 0 < e \ \d. 0 < d \ (\x. x \ a \ norm(x - a) < d \ norm(g x - f a) < e)" and r: "\x. x \ cball a r \ x \ a \ f x = g x" shows "continuous_on (cball a r) f" proof (clarsimp simp: continuous_on_eq_continuous_within Ball_def) fix x assume x: "dist a x \ r" show "continuous (at x within cball a r) f" proof (cases "x=a") case True then show ?thesis by (metis continuous_within_eps_delta fa dist_norm dist_self r) next case False show ?thesis proof (rule continuous_transform_within [where f=g and d = "norm(x-a)"]) have "\d>0. \x'\cball a r. dist x' x < d \ dist (g x') (g x) < e" if "e>0" for e proof - obtain d where "d > 0" and d: "\x'. \dist x' a \ r; x' \ a; dist x' x < d\ \ dist (g x') (g x) < e" using contg False x \e>0\ unfolding continuous_on_iff by (fastforce simp: dist_commute intro: that) show ?thesis using \d > 0\ \x \ a\ by (rule_tac x="min d (norm(x - a))" in exI) (auto simp: dist_commute dist_norm [symmetric] intro!: d) qed then show "continuous (at x within cball a r) g" using contg False by (auto simp: continuous_within_eps_delta) show "0 < norm (x - a)" using False by force show "x \ cball a r" by (simp add: x) show "\x'. \x' \ cball a r; dist x' x < norm (x - a)\ \ g x' = f x'" by (metis dist_commute dist_norm less_le r) qed qed qed proposition nullhomotopic_from_sphere_extension: fixes f :: "'M::euclidean_space \ 'a::real_normed_vector" shows "(\c. homotopic_with_canon (\x. True) (sphere a r) S f (\x. c)) \ (\g. continuous_on (cball a r) g \ g ` (cball a r) \ S \ (\x \ sphere a r. g x = f x))" (is "?lhs = ?rhs") proof (cases r "0::real" rule: linorder_cases) case less then show ?thesis by (simp add: homotopic_on_emptyI) next case equal show ?thesis proof assume L: ?lhs with equal have [simp]: "f a \ S" using homotopic_with_imp_subset1 by fastforce obtain h:: "real \ 'M \ 'a" where h: "continuous_on ({0..1} \ {a}) h" "h ` ({0..1} \ {a}) \ S" "h (0, a) = f a" using L equal by (auto simp: homotopic_with) then have "continuous_on (cball a r) (\x. h (0, a))" "(\x. h (0, a)) ` cball a r \ S" by (auto simp: equal) then show ?rhs using h(3) local.equal by force next assume ?rhs then show ?lhs using equal continuous_on_const by (force simp: homotopic_with) qed next case greater let ?P = "continuous_on {x. norm(x - a) = r} f \ f ` {x. norm(x - a) = r} \ S" have ?P if ?lhs using that proof fix c assume c: "homotopic_with_canon (\x. True) (sphere a r) S f (\x. c)" then have contf: "continuous_on (sphere a r) f" by (metis homotopic_with_imp_continuous) moreover have fim: "f ` sphere a r \ S" by (meson continuous_map_subtopology_eu c homotopic_with_imp_continuous_maps) show ?P using contf fim by (auto simp: sphere_def dist_norm norm_minus_commute) qed moreover have ?P if ?rhs using that proof fix g assume g: "continuous_on (cball a r) g \ g ` cball a r \ S \ (\xa\sphere a r. g xa = f xa)" then have "f ` {x. norm (x - a) = r} \ S" using sphere_cball [of a r] unfolding image_subset_iff sphere_def by (metis dist_commute dist_norm mem_Collect_eq subset_eq) with g show ?P by (auto simp: dist_norm norm_minus_commute elim!: continuous_on_eq [OF continuous_on_subset]) qed moreover have ?thesis if ?P proof assume ?lhs then obtain c where "homotopic_with_canon (\x. True) (sphere a r) S (\x. c) f" using homotopic_with_sym by blast then obtain h where conth: "continuous_on ({0..1::real} \ sphere a r) h" and him: "h ` ({0..1} \ sphere a r) \ S" and h: "\x. h(0, x) = c" "\x. h(1, x) = f x" by (auto simp: homotopic_with_def) obtain b1::'M where "b1 \ Basis" using SOME_Basis by auto have "c \ h ` ({0..1} \ sphere a r)" proof show "c = h (0, a + r *\<^sub>R b1)" by (simp add: h) show "(0, a + r *\<^sub>R b1) \ {0..1::real} \ sphere a r" using greater \b1 \ Basis\ by (auto simp: dist_norm) qed then have "c \ S" using him by blast have uconth: "uniformly_continuous_on ({0..1::real} \ (sphere a r)) h" by (force intro: compact_Times conth compact_uniformly_continuous) let ?g = "\x. h (norm (x - a)/r, a + (if x = a then r *\<^sub>R b1 else (r / norm(x - a)) *\<^sub>R (x - a)))" let ?g' = "\x. h (norm (x - a)/r, a + (r / norm(x - a)) *\<^sub>R (x - a))" show ?rhs proof (intro exI conjI) have "continuous_on (cball a r - {a}) ?g'" using greater by (force simp: dist_norm norm_minus_commute intro: continuous_on_compose2 [OF conth] continuous_intros) then show "continuous_on (cball a r) ?g" proof (rule nullhomotopic_from_lemma) show "\d>0. \x. x \ a \ norm (x - a) < d \ norm (?g' x - ?g a) < e" if "0 < e" for e proof - obtain d where "0 < d" and d: "\x x'. \x \ {0..1} \ sphere a r; x' \ {0..1} \ sphere a r; norm ( x' - x) < d\ \ norm (h x' - h x) < e" using uniformly_continuous_onE [OF uconth \0 < e\] by (auto simp: dist_norm) have *: "norm (h (norm (x - a) / r, a + (r / norm (x - a)) *\<^sub>R (x - a)) - h (0, a + r *\<^sub>R b1)) < e" (is "norm (?ha - ?hb) < e") if "x \ a" "norm (x - a) < r" "norm (x - a) < d * r" for x proof - have "norm (?ha - ?hb) = norm (?ha - h (0, a + (r / norm (x - a)) *\<^sub>R (x - a)))" by (simp add: h) also have "\ < e" using greater \0 < d\ \b1 \ Basis\ that by (intro d) (simp_all add: dist_norm, simp add: field_simps) finally show ?thesis . qed show ?thesis using greater \0 < d\ by (rule_tac x = "min r (d * r)" in exI) (auto simp: *) qed show "\x. x \ cball a r \ x \ a \ ?g x = ?g' x" by auto qed next show "?g ` cball a r \ S" using greater him \c \ S\ by (force simp: h dist_norm norm_minus_commute) next show "\x\sphere a r. ?g x = f x" using greater by (auto simp: h dist_norm norm_minus_commute) qed next assume ?rhs then obtain g where contg: "continuous_on (cball a r) g" and gim: "g ` cball a r \ S" and gf: "\x \ sphere a r. g x = f x" by auto let ?h = "\y. g (a + (fst y) *\<^sub>R (snd y - a))" have "continuous_on ({0..1} \ sphere a r) ?h" proof (rule continuous_on_compose2 [OF contg]) show "continuous_on ({0..1} \ sphere a r) (\x. a + fst x *\<^sub>R (snd x - a))" by (intro continuous_intros) qed (auto simp: dist_norm norm_minus_commute mult_left_le_one_le) moreover have "?h ` ({0..1} \ sphere a r) \ S" by (auto simp: dist_norm norm_minus_commute mult_left_le_one_le gim [THEN subsetD]) moreover have "\x\sphere a r. ?h (0, x) = g a" "\x\sphere a r. ?h (1, x) = f x" by (auto simp: dist_norm norm_minus_commute mult_left_le_one_le gf) ultimately have "homotopic_with_canon (\x. True) (sphere a r) S (\x. g a) f" by (auto simp: homotopic_with) then show ?lhs using homotopic_with_symD by blast qed ultimately show ?thesis by meson qed end \ No newline at end of file diff --git a/src/HOL/Analysis/Path_Connected.thy b/src/HOL/Analysis/Path_Connected.thy --- a/src/HOL/Analysis/Path_Connected.thy +++ b/src/HOL/Analysis/Path_Connected.thy @@ -1,4066 +1,4096 @@ (* Title: HOL/Analysis/Path_Connected.thy Authors: LC Paulson and Robert Himmelmann (TU Muenchen), based on material from HOL Light *) section \Path-Connectedness\ theory Path_Connected imports Starlike T1_Spaces begin subsection \Paths and Arcs\ definition\<^marker>\tag important\ path :: "(real \ 'a::topological_space) \ bool" where "path g \ continuous_on {0..1} g" definition\<^marker>\tag important\ pathstart :: "(real \ 'a::topological_space) \ 'a" where "pathstart g \ g 0" definition\<^marker>\tag important\ pathfinish :: "(real \ 'a::topological_space) \ 'a" where "pathfinish g \ g 1" definition\<^marker>\tag important\ path_image :: "(real \ 'a::topological_space) \ 'a set" where "path_image g \ g ` {0 .. 1}" definition\<^marker>\tag important\ reversepath :: "(real \ 'a::topological_space) \ real \ 'a" where "reversepath g \ (\x. g(1 - x))" definition\<^marker>\tag important\ joinpaths :: "(real \ 'a::topological_space) \ (real \ 'a) \ real \ 'a" (infixr "+++" 75) where "g1 +++ g2 \ (\x. if x \ 1/2 then g1 (2 * x) else g2 (2 * x - 1))" definition\<^marker>\tag important\ loop_free :: "(real \ 'a::topological_space) \ bool" where "loop_free g \ \x\{0..1}. \y\{0..1}. g x = g y \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" definition\<^marker>\tag important\ simple_path :: "(real \ 'a::topological_space) \ bool" where "simple_path g \ path g \ loop_free g" definition\<^marker>\tag important\ arc :: "(real \ 'a :: topological_space) \ bool" where "arc g \ path g \ inj_on g {0..1}" subsection\<^marker>\tag unimportant\\Invariance theorems\ lemma path_eq: "path p \ (\t. t \ {0..1} \ p t = q t) \ path q" using continuous_on_eq path_def by blast lemma path_continuous_image: "path g \ continuous_on (path_image g) f \ path(f \ g)" unfolding path_def path_image_def using continuous_on_compose by blast lemma path_translation_eq: fixes g :: "real \ 'a :: real_normed_vector" shows "path((\x. a + x) \ g) = path g" using continuous_on_translation_eq path_def by blast lemma path_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "path(f \ g) = path g" proof - from linear_injective_left_inverse [OF assms] obtain h where h: "linear h" "h \ f = id" by blast with assms show ?thesis by (metis comp_assoc id_comp linear_continuous_on linear_linear path_continuous_image) qed lemma pathstart_translation: "pathstart((\x. a + x) \ g) = a + pathstart g" by (simp add: pathstart_def) lemma pathstart_linear_image_eq: "linear f \ pathstart(f \ g) = f(pathstart g)" by (simp add: pathstart_def) lemma pathfinish_translation: "pathfinish((\x. a + x) \ g) = a + pathfinish g" by (simp add: pathfinish_def) lemma pathfinish_linear_image: "linear f \ pathfinish(f \ g) = f(pathfinish g)" by (simp add: pathfinish_def) lemma path_image_translation: "path_image((\x. a + x) \ g) = (\x. a + x) ` (path_image g)" by (simp add: image_comp path_image_def) lemma path_image_linear_image: "linear f \ path_image(f \ g) = f ` (path_image g)" by (simp add: image_comp path_image_def) lemma reversepath_translation: "reversepath((\x. a + x) \ g) = (\x. a + x) \ reversepath g" by (rule ext) (simp add: reversepath_def) lemma reversepath_linear_image: "linear f \ reversepath(f \ g) = f \ reversepath g" by (rule ext) (simp add: reversepath_def) lemma joinpaths_translation: "((\x. a + x) \ g1) +++ ((\x. a + x) \ g2) = (\x. a + x) \ (g1 +++ g2)" by (rule ext) (simp add: joinpaths_def) lemma joinpaths_linear_image: "linear f \ (f \ g1) +++ (f \ g2) = f \ (g1 +++ g2)" by (rule ext) (simp add: joinpaths_def) lemma loop_free_translation_eq: fixes g :: "real \ 'a::euclidean_space" shows "loop_free((\x. a + x) \ g) = loop_free g" by (simp add: loop_free_def) lemma simple_path_translation_eq: fixes g :: "real \ 'a::euclidean_space" shows "simple_path((\x. a + x) \ g) = simple_path g" by (simp add: simple_path_def loop_free_translation_eq path_translation_eq) lemma loop_free_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "loop_free(f \ g) = loop_free g" using assms inj_on_eq_iff [of f] by (auto simp: loop_free_def) lemma simple_path_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "simple_path(f \ g) = simple_path g" using assms by (simp add: loop_free_linear_image_eq path_linear_image_eq simple_path_def) +lemma simple_pathI [intro?]: + assumes "path p" + assumes "\x y. 0 \ x \ x < y \ y \ 1 \ p x = p y \ x = 0 \ y = 1" + shows "simple_path p" + unfolding simple_path_def loop_free_def +proof (intro ballI conjI impI) + fix x y assume "x \ {0..1}" "y \ {0..1}" "p x = p y" + thus "x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" + by (metis assms(2) atLeastAtMost_iff linorder_less_linear) +qed fact+ + +lemma arcD: "arc p \ p x = p y \ x \ {0..1} \ y \ {0..1} \ x = y" + by (auto simp: arc_def inj_on_def) + lemma arc_translation_eq: fixes g :: "real \ 'a::euclidean_space" shows "arc((\x. a + x) \ g) \ arc g" by (auto simp: arc_def inj_on_def path_translation_eq) lemma arc_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "arc(f \ g) = arc g" using assms inj_on_eq_iff [of f] by (auto simp: arc_def inj_on_def path_linear_image_eq) subsection\<^marker>\tag unimportant\\Basic lemmas about paths\ lemma path_of_real: "path complex_of_real" unfolding path_def by (intro continuous_intros) lemma path_const: "path (\t. a)" for a::"'a::real_normed_vector" unfolding path_def by (intro continuous_intros) lemma path_minus: "path g \ path (\t. - g t)" for g::"real\'a::real_normed_vector" unfolding path_def by (intro continuous_intros) lemma path_add: "\path f; path g\ \ path (\t. f t + g t)" for f::"real\'a::real_normed_vector" unfolding path_def by (intro continuous_intros) lemma path_diff: "\path f; path g\ \ path (\t. f t - g t)" for f::"real\'a::real_normed_vector" unfolding path_def by (intro continuous_intros) lemma path_mult: "\path f; path g\ \ path (\t. f t * g t)" for f::"real\'a::real_normed_field" unfolding path_def by (intro continuous_intros) lemma pathin_iff_path_real [simp]: "pathin euclideanreal g \ path g" by (simp add: pathin_def path_def) lemma continuous_on_path: "path f \ t \ {0..1} \ continuous_on t f" using continuous_on_subset path_def by blast lemma inj_on_imp_loop_free: "inj_on g {0..1} \ loop_free g" by (simp add: inj_onD loop_free_def) lemma arc_imp_simple_path: "arc g \ simple_path g" by (simp add: arc_def inj_on_imp_loop_free simple_path_def) lemma arc_imp_path: "arc g \ path g" using arc_def by blast lemma arc_imp_inj_on: "arc g \ inj_on g {0..1}" by (auto simp: arc_def) lemma simple_path_imp_path: "simple_path g \ path g" using simple_path_def by blast lemma loop_free_cases: "loop_free g \ inj_on g {0..1} \ pathfinish g = pathstart g" by (force simp: inj_on_def loop_free_def pathfinish_def pathstart_def) lemma simple_path_cases: "simple_path g \ arc g \ pathfinish g = pathstart g" using arc_def loop_free_cases simple_path_def by blast lemma simple_path_imp_arc: "simple_path g \ pathfinish g \ pathstart g \ arc g" using simple_path_cases by auto lemma arc_distinct_ends: "arc g \ pathfinish g \ pathstart g" unfolding arc_def inj_on_def pathfinish_def pathstart_def by fastforce lemma arc_simple_path: "arc g \ simple_path g \ pathfinish g \ pathstart g" using arc_distinct_ends arc_imp_simple_path simple_path_cases by blast lemma simple_path_eq_arc: "pathfinish g \ pathstart g \ (simple_path g = arc g)" by (simp add: arc_simple_path) lemma path_image_const [simp]: "path_image (\t. a) = {a}" by (force simp: path_image_def) lemma path_image_nonempty [simp]: "path_image g \ {}" unfolding path_image_def image_is_empty box_eq_empty by auto lemma pathstart_in_path_image[intro]: "pathstart g \ path_image g" unfolding pathstart_def path_image_def by auto lemma pathfinish_in_path_image[intro]: "pathfinish g \ path_image g" unfolding pathfinish_def path_image_def by auto lemma connected_path_image[intro]: "path g \ connected (path_image g)" unfolding path_def path_image_def using connected_continuous_image connected_Icc by blast lemma compact_path_image[intro]: "path g \ compact (path_image g)" unfolding path_def path_image_def using compact_continuous_image connected_Icc by blast lemma reversepath_reversepath[simp]: "reversepath (reversepath g) = g" unfolding reversepath_def by auto lemma pathstart_reversepath[simp]: "pathstart (reversepath g) = pathfinish g" unfolding pathstart_def reversepath_def pathfinish_def by auto lemma pathfinish_reversepath[simp]: "pathfinish (reversepath g) = pathstart g" unfolding pathstart_def reversepath_def pathfinish_def by auto lemma reversepath_o: "reversepath g = g \ (-)1" by (auto simp: reversepath_def) lemma pathstart_join[simp]: "pathstart (g1 +++ g2) = pathstart g1" unfolding pathstart_def joinpaths_def pathfinish_def by auto lemma pathfinish_join[simp]: "pathfinish (g1 +++ g2) = pathfinish g2" unfolding pathstart_def joinpaths_def pathfinish_def by auto lemma path_image_reversepath[simp]: "path_image (reversepath g) = path_image g" by (metis cancel_comm_monoid_add_class.diff_cancel diff_zero image_comp image_diff_atLeastAtMost path_image_def reversepath_o) lemma path_reversepath [simp]: "path (reversepath g) \ path g" by (metis continuous_on_compose continuous_on_op_minus image_comp image_ident path_def path_image_def path_image_reversepath reversepath_o reversepath_reversepath) lemma arc_reversepath: assumes "arc g" shows "arc(reversepath g)" proof - have injg: "inj_on g {0..1}" using assms by (simp add: arc_def) have **: "\x y::real. 1-x = 1-y \ x = y" by simp show ?thesis using assms by (clarsimp simp: arc_def intro!: inj_onI) (simp add: inj_onD reversepath_def **) qed lemma loop_free_reversepath: assumes "loop_free g" shows "loop_free(reversepath g)" using assms by (simp add: reversepath_def loop_free_def Ball_def) (smt (verit)) lemma simple_path_reversepath: "simple_path g \ simple_path (reversepath g)" by (simp add: loop_free_reversepath simple_path_def) lemmas reversepath_simps = path_reversepath path_image_reversepath pathstart_reversepath pathfinish_reversepath lemma path_join[simp]: assumes "pathfinish g1 = pathstart g2" shows "path (g1 +++ g2) \ path g1 \ path g2" unfolding path_def pathfinish_def pathstart_def proof safe assume cont: "continuous_on {0..1} (g1 +++ g2)" have g1: "continuous_on {0..1} g1 \ continuous_on {0..1} ((g1 +++ g2) \ (\x. x / 2))" by (intro continuous_on_cong refl) (auto simp: joinpaths_def) have g2: "continuous_on {0..1} g2 \ continuous_on {0..1} ((g1 +++ g2) \ (\x. x / 2 + 1/2))" using assms by (intro continuous_on_cong refl) (auto simp: joinpaths_def pathfinish_def pathstart_def) show "continuous_on {0..1} g1" and "continuous_on {0..1} g2" unfolding g1 g2 by (auto intro!: continuous_intros continuous_on_subset[OF cont] simp del: o_apply) next assume g1g2: "continuous_on {0..1} g1" "continuous_on {0..1} g2" have 01: "{0 .. 1} = {0..1/2} \ {1/2 .. 1::real}" by auto { fix x :: real assume "0 \ x" and "x \ 1" then have "x \ (\x. x * 2) ` {0..1 / 2}" by (intro image_eqI[where x="x/2"]) auto } note 1 = this { fix x :: real assume "0 \ x" and "x \ 1" then have "x \ (\x. x * 2 - 1) ` {1 / 2..1}" by (intro image_eqI[where x="x/2 + 1/2"]) auto } note 2 = this show "continuous_on {0..1} (g1 +++ g2)" using assms unfolding joinpaths_def 01 apply (intro continuous_on_cases closed_atLeastAtMost g1g2[THEN continuous_on_compose2] continuous_intros) apply (auto simp: field_simps pathfinish_def pathstart_def intro!: 1 2) done qed subsection\<^marker>\tag unimportant\ \Path Images\ lemma bounded_path_image: "path g \ bounded(path_image g)" by (simp add: compact_imp_bounded compact_path_image) lemma closed_path_image: fixes g :: "real \ 'a::t2_space" shows "path g \ closed(path_image g)" by (metis compact_path_image compact_imp_closed) lemma connected_simple_path_image: "simple_path g \ connected(path_image g)" by (metis connected_path_image simple_path_imp_path) lemma compact_simple_path_image: "simple_path g \ compact(path_image g)" by (metis compact_path_image simple_path_imp_path) lemma bounded_simple_path_image: "simple_path g \ bounded(path_image g)" by (metis bounded_path_image simple_path_imp_path) lemma closed_simple_path_image: fixes g :: "real \ 'a::t2_space" shows "simple_path g \ closed(path_image g)" by (metis closed_path_image simple_path_imp_path) lemma connected_arc_image: "arc g \ connected(path_image g)" by (metis connected_path_image arc_imp_path) lemma compact_arc_image: "arc g \ compact(path_image g)" by (metis compact_path_image arc_imp_path) lemma bounded_arc_image: "arc g \ bounded(path_image g)" by (metis bounded_path_image arc_imp_path) lemma closed_arc_image: fixes g :: "real \ 'a::t2_space" shows "arc g \ closed(path_image g)" by (metis closed_path_image arc_imp_path) lemma path_image_join_subset: "path_image (g1 +++ g2) \ path_image g1 \ path_image g2" unfolding path_image_def joinpaths_def by auto lemma subset_path_image_join: assumes "path_image g1 \ S" and "path_image g2 \ S" shows "path_image (g1 +++ g2) \ S" using path_image_join_subset[of g1 g2] and assms by auto lemma path_image_join: assumes "pathfinish g1 = pathstart g2" shows "path_image(g1 +++ g2) = path_image g1 \ path_image g2" proof - have "path_image g1 \ path_image (g1 +++ g2)" proof (clarsimp simp: path_image_def joinpaths_def) fix u::real assume "0 \ u" "u \ 1" then show "g1 u \ (\x. g1 (2 * x)) ` ({0..1} \ {x. x * 2 \ 1})" by (rule_tac x="u/2" in image_eqI) auto qed moreover have \
: "g2 u \ (\x. g2 (2 * x - 1)) ` ({0..1} \ {x. \ x * 2 \ 1})" if "0 < u" "u \ 1" for u using that assms by (rule_tac x="(u+1)/2" in image_eqI) (auto simp: field_simps pathfinish_def pathstart_def) have "g2 0 \ (\x. g1 (2 * x)) ` ({0..1} \ {x. x * 2 \ 1})" using assms by (rule_tac x="1/2" in image_eqI) (auto simp: pathfinish_def pathstart_def) then have "path_image g2 \ path_image (g1 +++ g2)" by (auto simp: path_image_def joinpaths_def intro!: \
) ultimately show ?thesis using path_image_join_subset by blast qed lemma not_in_path_image_join: assumes "x \ path_image g1" and "x \ path_image g2" shows "x \ path_image (g1 +++ g2)" using assms and path_image_join_subset[of g1 g2] by auto lemma pathstart_compose: "pathstart(f \ p) = f(pathstart p)" by (simp add: pathstart_def) lemma pathfinish_compose: "pathfinish(f \ p) = f(pathfinish p)" by (simp add: pathfinish_def) lemma path_image_compose: "path_image (f \ p) = f ` (path_image p)" by (simp add: image_comp path_image_def) lemma path_compose_join: "f \ (p +++ q) = (f \ p) +++ (f \ q)" by (rule ext) (simp add: joinpaths_def) lemma path_compose_reversepath: "f \ reversepath p = reversepath(f \ p)" by (rule ext) (simp add: reversepath_def) lemma joinpaths_eq: "(\t. t \ {0..1} \ p t = p' t) \ (\t. t \ {0..1} \ q t = q' t) \ t \ {0..1} \ (p +++ q) t = (p' +++ q') t" by (auto simp: joinpaths_def) lemma loop_free_inj_on: "loop_free g \ inj_on g {0<..<1}" by (force simp: inj_on_def loop_free_def) lemma simple_path_inj_on: "simple_path g \ inj_on g {0<..<1}" using loop_free_inj_on simple_path_def by auto subsection\<^marker>\tag unimportant\\Simple paths with the endpoints removed\ lemma simple_path_endless: assumes "simple_path c" shows "path_image c - {pathstart c,pathfinish c} = c ` {0<..<1}" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" using less_eq_real_def by (auto simp: path_image_def pathstart_def pathfinish_def) show "?rhs \ ?lhs" using assms apply (simp add: image_subset_iff path_image_def pathstart_def pathfinish_def simple_path_def loop_free_def Ball_def) by (smt (verit)) qed lemma connected_simple_path_endless: assumes "simple_path c" shows "connected(path_image c - {pathstart c,pathfinish c})" proof - have "continuous_on {0<..<1} c" using assms by (simp add: simple_path_def continuous_on_path path_def subset_iff) then have "connected (c ` {0<..<1})" using connected_Ioo connected_continuous_image by blast then show ?thesis using assms by (simp add: simple_path_endless) qed lemma nonempty_simple_path_endless: "simple_path c \ path_image c - {pathstart c,pathfinish c} \ {}" by (simp add: simple_path_endless) lemma simple_path_continuous_image: assumes "simple_path f" "continuous_on (path_image f) g" "inj_on g (path_image f)" shows "simple_path (g \ f)" unfolding simple_path_def proof show "path (g \ f)" using assms unfolding simple_path_def by (intro path_continuous_image) auto from assms have [simp]: "g (f x) = g (f y) \ f x = f y" if "x \ {0..1}" "y \ {0..1}" for x y unfolding inj_on_def path_image_def using that by fastforce show "loop_free (g \ f)" using assms(1) by (auto simp: loop_free_def simple_path_def) qed subsection\<^marker>\tag unimportant\\The operations on paths\ lemma path_image_subset_reversepath: "path_image(reversepath g) \ path_image g" by simp lemma path_imp_reversepath: "path g \ path(reversepath g)" by simp lemma half_bounded_equal: "1 \ x * 2 \ x * 2 \ 1 \ x = (1/2::real)" by simp lemma continuous_on_joinpaths: assumes "continuous_on {0..1} g1" "continuous_on {0..1} g2" "pathfinish g1 = pathstart g2" shows "continuous_on {0..1} (g1 +++ g2)" using assms path_def path_join by blast lemma path_join_imp: "\path g1; path g2; pathfinish g1 = pathstart g2\ \ path(g1 +++ g2)" by simp lemma arc_join: assumes "arc g1" "arc g2" "pathfinish g1 = pathstart g2" "path_image g1 \ path_image g2 \ {pathstart g2}" shows "arc(g1 +++ g2)" proof - have injg1: "inj_on g1 {0..1}" and injg2: "inj_on g2 {0..1}" and g11: "g1 1 = g2 0" and sb: "g1 ` {0..1} \ g2 ` {0..1} \ {g2 0}" using assms by (auto simp: arc_def pathfinish_def pathstart_def path_image_def) { fix x and y::real assume xy: "g2 (2 * x - 1) = g1 (2 * y)" "x \ 1" "0 \ y" " y * 2 \ 1" "\ x * 2 \ 1" then have "g1 (2 * y) = g2 0" using sb by force then have False using xy inj_onD injg2 by fastforce } note * = this have "inj_on (g1 +++ g2) {0..1}" using inj_onD [OF injg1] inj_onD [OF injg2] * by (simp add: inj_on_def joinpaths_def Ball_def) (smt (verit)) then show ?thesis using arc_def assms path_join_imp by blast qed lemma simple_path_join_loop: assumes "arc g1" "arc g2" "pathfinish g1 = pathstart g2" "pathfinish g2 = pathstart g1" "path_image g1 \ path_image g2 \ {pathstart g1, pathstart g2}" shows "simple_path(g1 +++ g2)" proof - have injg1: "inj_on g1 {0..1}" and injg2: "inj_on g2 {0..1}" using assms by (auto simp add: arc_def) have g12: "g1 1 = g2 0" and g21: "g2 1 = g1 0" and sb: "g1 ` {0..1} \ g2 ` {0..1} \ {g1 0, g2 0}" using assms by (simp_all add: arc_def pathfinish_def pathstart_def path_image_def) { fix x and y::real assume g2_eq: "g2 (2 * x - 1) = g1 (2 * y)" and xyI: "x \ 1 \ y \ 0" and xy: "x \ 1" "0 \ y" " y * 2 \ 1" "\ x * 2 \ 1" then consider "g1 (2 * y) = g1 0" | "g1 (2 * y) = g2 0" using sb by force then have False proof cases case 1 then have "y = 0" using xy g2_eq by (auto dest!: inj_onD [OF injg1]) then show ?thesis using xy g2_eq xyI by (auto dest: inj_onD [OF injg2] simp flip: g21) next case 2 then have "2*x = 1" using g2_eq g12 inj_onD [OF injg2] atLeastAtMost_iff xy(1) xy(4) by fastforce with xy show False by auto qed } note * = this have "loop_free(g1 +++ g2)" using inj_onD [OF injg1] inj_onD [OF injg2] * by (simp add: loop_free_def joinpaths_def Ball_def) (smt (verit)) then show ?thesis by (simp add: arc_imp_path assms simple_path_def) qed lemma reversepath_joinpaths: "pathfinish g1 = pathstart g2 \ reversepath(g1 +++ g2) = reversepath g2 +++ reversepath g1" unfolding reversepath_def pathfinish_def pathstart_def joinpaths_def by (rule ext) (auto simp: mult.commute) subsection\<^marker>\tag unimportant\\Some reversed and "if and only if" versions of joining theorems\ lemma path_join_path_ends: fixes g1 :: "real \ 'a::metric_space" assumes "path(g1 +++ g2)" "path g2" shows "pathfinish g1 = pathstart g2" proof (rule ccontr) define e where "e = dist (g1 1) (g2 0)" assume Neg: "pathfinish g1 \ pathstart g2" then have "0 < dist (pathfinish g1) (pathstart g2)" by auto then have "e > 0" by (metis e_def pathfinish_def pathstart_def) then have "\e>0. \d>0. \x'\{0..1}. dist x' 0 < d \ dist (g2 x') (g2 0) < e" using \path g2\ atLeastAtMost_iff zero_le_one unfolding path_def continuous_on_iff by blast then obtain d1 where "d1 > 0" and d1: "\x'. \x'\{0..1}; norm x' < d1\ \ dist (g2 x') (g2 0) < e/2" by (metis \0 < e\ half_gt_zero_iff norm_conv_dist) obtain d2 where "d2 > 0" and d2: "\x'. \x'\{0..1}; dist x' (1/2) < d2\ \ dist ((g1 +++ g2) x') (g1 1) < e/2" using assms(1) \e > 0\ unfolding path_def continuous_on_iff apply (drule_tac x="1/2" in bspec, simp) apply (drule_tac x="e/2" in spec, force simp: joinpaths_def) done have int01_1: "min (1/2) (min d1 d2) / 2 \ {0..1}" using \d1 > 0\ \d2 > 0\ by (simp add: min_def) have dist1: "norm (min (1 / 2) (min d1 d2) / 2) < d1" using \d1 > 0\ \d2 > 0\ by (simp add: min_def dist_norm) have int01_2: "1/2 + min (1/2) (min d1 d2) / 4 \ {0..1}" using \d1 > 0\ \d2 > 0\ by (simp add: min_def) have dist2: "dist (1 / 2 + min (1 / 2) (min d1 d2) / 4) (1 / 2) < d2" using \d1 > 0\ \d2 > 0\ by (simp add: min_def dist_norm) have [simp]: "\ min (1 / 2) (min d1 d2) \ 0" using \d1 > 0\ \d2 > 0\ by (simp add: min_def) have "dist (g2 (min (1 / 2) (min d1 d2) / 2)) (g1 1) < e/2" "dist (g2 (min (1 / 2) (min d1 d2) / 2)) (g2 0) < e/2" using d1 [OF int01_1 dist1] d2 [OF int01_2 dist2] by (simp_all add: joinpaths_def) then have "dist (g1 1) (g2 0) < e/2 + e/2" using dist_triangle_half_r e_def by blast then show False by (simp add: e_def [symmetric]) qed lemma path_join_eq [simp]: fixes g1 :: "real \ 'a::metric_space" assumes "path g1" "path g2" shows "path(g1 +++ g2) \ pathfinish g1 = pathstart g2" using assms by (metis path_join_path_ends path_join_imp) lemma simple_path_joinE: assumes "simple_path(g1 +++ g2)" and "pathfinish g1 = pathstart g2" obtains "arc g1" "arc g2" "path_image g1 \ path_image g2 \ {pathstart g1, pathstart g2}" proof - have *: "\x y. \0 \ x; x \ 1; 0 \ y; y \ 1; (g1 +++ g2) x = (g1 +++ g2) y\ \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" using assms by (simp add: simple_path_def loop_free_def) have "path g1" using assms path_join simple_path_imp_path by blast moreover have "inj_on g1 {0..1}" proof (clarsimp simp: inj_on_def) fix x y assume "g1 x = g1 y" "0 \ x" "x \ 1" "0 \ y" "y \ 1" then show "x = y" using * [of "x/2" "y/2"] by (simp add: joinpaths_def split_ifs) qed ultimately have "arc g1" using assms by (simp add: arc_def) have [simp]: "g2 0 = g1 1" using assms by (metis pathfinish_def pathstart_def) have "path g2" using assms path_join simple_path_imp_path by blast moreover have "inj_on g2 {0..1}" proof (clarsimp simp: inj_on_def) fix x y assume "g2 x = g2 y" "0 \ x" "x \ 1" "0 \ y" "y \ 1" then show "x = y" using * [of "(x+1) / 2" "(y+1) / 2"] by (force simp: joinpaths_def split_ifs field_split_simps) qed ultimately have "arc g2" using assms by (simp add: arc_def) have "g2 y = g1 0 \ g2 y = g1 1" if "g1 x = g2 y" "0 \ x" "x \ 1" "0 \ y" "y \ 1" for x y using * [of "x / 2" "(y + 1) / 2"] that by (auto simp: joinpaths_def split_ifs field_split_simps) then have "path_image g1 \ path_image g2 \ {pathstart g1, pathstart g2}" by (fastforce simp: pathstart_def pathfinish_def path_image_def) with \arc g1\ \arc g2\ show ?thesis using that by blast qed lemma simple_path_join_loop_eq: assumes "pathfinish g2 = pathstart g1" "pathfinish g1 = pathstart g2" shows "simple_path(g1 +++ g2) \ arc g1 \ arc g2 \ path_image g1 \ path_image g2 \ {pathstart g1, pathstart g2}" by (metis assms simple_path_joinE simple_path_join_loop) lemma arc_join_eq: assumes "pathfinish g1 = pathstart g2" shows "arc(g1 +++ g2) \ arc g1 \ arc g2 \ path_image g1 \ path_image g2 \ {pathstart g2}" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using reversepath_simps assms by (smt (verit, best) Int_commute arc_reversepath arc_simple_path in_mono insertE pathstart_join reversepath_joinpaths simple_path_joinE subsetI) next assume ?rhs then show ?lhs using assms by (fastforce simp: pathfinish_def pathstart_def intro!: arc_join) qed lemma arc_join_eq_alt: "pathfinish g1 = pathstart g2 \ arc(g1 +++ g2) \ arc g1 \ arc g2 \ path_image g1 \ path_image g2 = {pathstart g2}" using pathfinish_in_path_image by (fastforce simp: arc_join_eq) +subsubsection\<^marker>\tag unimportant\\Symmetry and loops\ + +lemma path_sym: + "\pathfinish p = pathstart q; pathfinish q = pathstart p\ \ path(p +++ q) \ path(q +++ p)" + by auto + +lemma simple_path_sym: + "\pathfinish p = pathstart q; pathfinish q = pathstart p\ + \ simple_path(p +++ q) \ simple_path(q +++ p)" + by (metis (full_types) inf_commute insert_commute simple_path_joinE simple_path_join_loop) + +lemma path_image_sym: + "\pathfinish p = pathstart q; pathfinish q = pathstart p\ + \ path_image(p +++ q) = path_image(q +++ p)" + by (simp add: path_image_join sup_commute) + +lemma simple_path_joinI: + assumes "arc p1" "arc p2" "pathfinish p1 = pathstart p2" + assumes "path_image p1 \ path_image p2 + \ insert (pathstart p2) (if pathstart p1 = pathfinish p2 then {pathstart p1} else {})" + shows "simple_path (p1 +++ p2)" + by (smt (verit, best) Int_commute arc_imp_simple_path arc_join assms insert_commute simple_path_join_loop simple_path_sym) + +lemma simple_path_join3I: + assumes "arc p1" "arc p2" "arc p3" + assumes "path_image p1 \ path_image p2 \ {pathstart p2}" + assumes "path_image p2 \ path_image p3 \ {pathstart p3}" + assumes "path_image p1 \ path_image p3 \ {pathstart p1} \ {pathfinish p3}" + assumes "pathfinish p1 = pathstart p2" "pathfinish p2 = pathstart p3" + shows "simple_path (p1 +++ p2 +++ p3)" + using assms by (intro simple_path_joinI arc_join) (auto simp: path_image_join) + subsection\<^marker>\tag unimportant\\The joining of paths is associative\ lemma path_assoc: "\pathfinish p = pathstart q; pathfinish q = pathstart r\ \ path(p +++ (q +++ r)) \ path((p +++ q) +++ r)" by simp lemma simple_path_assoc: assumes "pathfinish p = pathstart q" "pathfinish q = pathstart r" shows "simple_path (p +++ (q +++ r)) \ simple_path ((p +++ q) +++ r)" proof (cases "pathstart p = pathfinish r") case True show ?thesis proof assume "simple_path (p +++ q +++ r)" with assms True show "simple_path ((p +++ q) +++ r)" by (fastforce simp add: simple_path_join_loop_eq arc_join_eq path_image_join dest: arc_distinct_ends [of r]) next assume 0: "simple_path ((p +++ q) +++ r)" with assms True have q: "pathfinish r \ path_image q" using arc_distinct_ends by (fastforce simp add: simple_path_join_loop_eq arc_join_eq path_image_join) have "pathstart r \ path_image p" using assms by (metis 0 IntI arc_distinct_ends arc_join_eq_alt empty_iff insert_iff pathfinish_in_path_image pathfinish_join simple_path_joinE) with assms 0 q True show "simple_path (p +++ q +++ r)" by (auto simp: simple_path_join_loop_eq arc_join_eq path_image_join dest!: subsetD [OF _ IntI]) qed next case False { fix x :: 'a assume a: "path_image p \ path_image q \ {pathstart q}" "(path_image p \ path_image q) \ path_image r \ {pathstart r}" "x \ path_image p" "x \ path_image r" have "pathstart r \ path_image q" by (metis assms(2) pathfinish_in_path_image) with a have "x = pathstart q" by blast } with False assms show ?thesis by (auto simp: simple_path_eq_arc simple_path_join_loop_eq arc_join_eq path_image_join) qed lemma arc_assoc: "\pathfinish p = pathstart q; pathfinish q = pathstart r\ \ arc(p +++ (q +++ r)) \ arc((p +++ q) +++ r)" by (simp add: arc_simple_path simple_path_assoc) -subsubsection\<^marker>\tag unimportant\\Symmetry and loops\ - -lemma path_sym: - "\pathfinish p = pathstart q; pathfinish q = pathstart p\ \ path(p +++ q) \ path(q +++ p)" - by auto - -lemma simple_path_sym: - "\pathfinish p = pathstart q; pathfinish q = pathstart p\ - \ simple_path(p +++ q) \ simple_path(q +++ p)" - by (metis (full_types) inf_commute insert_commute simple_path_joinE simple_path_join_loop) - -lemma path_image_sym: - "\pathfinish p = pathstart q; pathfinish q = pathstart p\ - \ path_image(p +++ q) = path_image(q +++ p)" - by (simp add: path_image_join sup_commute) - subsection\Subpath\ definition\<^marker>\tag important\ subpath :: "real \ real \ (real \ 'a) \ real \ 'a::real_normed_vector" where "subpath a b g \ \x. g((b - a) * x + a)" lemma path_image_subpath_gen: fixes g :: "_ \ 'a::real_normed_vector" shows "path_image(subpath u v g) = g ` (closed_segment u v)" by (auto simp add: closed_segment_real_eq path_image_def subpath_def) lemma path_image_subpath: fixes g :: "real \ 'a::real_normed_vector" shows "path_image(subpath u v g) = (if u \ v then g ` {u..v} else g ` {v..u})" by (simp add: path_image_subpath_gen closed_segment_eq_real_ivl) lemma path_image_subpath_commute: fixes g :: "real \ 'a::real_normed_vector" shows "path_image(subpath u v g) = path_image(subpath v u g)" by (simp add: path_image_subpath_gen closed_segment_eq_real_ivl) lemma path_subpath [simp]: fixes g :: "real \ 'a::real_normed_vector" assumes "path g" "u \ {0..1}" "v \ {0..1}" shows "path(subpath u v g)" proof - have "continuous_on {u..v} g" "continuous_on {v..u} g" using assms continuous_on_path by fastforce+ then have "continuous_on {0..1} (g \ (\x. ((v-u) * x+ u)))" by (intro continuous_intros; simp add: image_affinity_atLeastAtMost [where c=u]) then show ?thesis by (simp add: path_def subpath_def) qed lemma pathstart_subpath [simp]: "pathstart(subpath u v g) = g(u)" by (simp add: pathstart_def subpath_def) lemma pathfinish_subpath [simp]: "pathfinish(subpath u v g) = g(v)" by (simp add: pathfinish_def subpath_def) lemma subpath_trivial [simp]: "subpath 0 1 g = g" by (simp add: subpath_def) lemma subpath_reversepath: "subpath 1 0 g = reversepath g" by (simp add: reversepath_def subpath_def) lemma reversepath_subpath: "reversepath(subpath u v g) = subpath v u g" by (simp add: reversepath_def subpath_def algebra_simps) lemma subpath_translation: "subpath u v ((\x. a + x) \ g) = (\x. a + x) \ subpath u v g" by (rule ext) (simp add: subpath_def) lemma subpath_image: "subpath u v (f \ g) = f \ subpath u v g" by (rule ext) (simp add: subpath_def) lemma affine_ineq: fixes x :: "'a::linordered_idom" assumes "x \ 1" "v \ u" shows "v + x * u \ u + x * v" proof - have "(1-x)*(u-v) \ 0" using assms by auto then show ?thesis by (simp add: algebra_simps) qed lemma sum_le_prod1: fixes a::real shows "\a \ 1; b \ 1\ \ a + b \ 1 + a * b" by (metis add.commute affine_ineq mult.right_neutral) lemma simple_path_subpath_eq: "simple_path(subpath u v g) \ path(subpath u v g) \ u\v \ (\x y. x \ closed_segment u v \ y \ closed_segment u v \ g x = g y \ x = y \ x = u \ y = v \ x = v \ y = u)" (is "?lhs = ?rhs") proof assume ?lhs then have p: "path (\x. g ((v - u) * x + u))" and sim: "(\x y. \x\{0..1}; y\{0..1}; g ((v - u) * x + u) = g ((v - u) * y + u)\ \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0)" by (auto simp: simple_path_def loop_free_def subpath_def) { fix x y assume "x \ closed_segment u v" "y \ closed_segment u v" "g x = g y" then have "x = y \ x = u \ y = v \ x = v \ y = u" using sim [of "(x-u)/(v-u)" "(y-u)/(v-u)"] p by (auto split: if_split_asm simp add: closed_segment_real_eq image_affinity_atLeastAtMost) (simp_all add: field_split_simps) } moreover have "path(subpath u v g) \ u\v" using sim [of "1/3" "2/3"] p by (auto simp: subpath_def) ultimately show ?rhs by metis next assume ?rhs then have d1: "\x y. \g x = g y; u \ x; x \ v; u \ y; y \ v\ \ x = y \ x = u \ y = v \ x = v \ y = u" and d2: "\x y. \g x = g y; v \ x; x \ u; v \ y; y \ u\ \ x = y \ x = u \ y = v \ x = v \ y = u" and ne: "u < v \ v < u" and psp: "path (subpath u v g)" by (auto simp: closed_segment_real_eq image_affinity_atLeastAtMost) have [simp]: "\x. u + x * v = v + x * u \ u=v \ x=1" by algebra show ?lhs using psp ne unfolding simple_path_def loop_free_def subpath_def by (fastforce simp add: algebra_simps affine_ineq mult_left_mono crossproduct_eq dest: d1 d2) qed lemma arc_subpath_eq: "arc(subpath u v g) \ path(subpath u v g) \ u\v \ inj_on g (closed_segment u v)" by (smt (verit, best) arc_simple_path closed_segment_commute ends_in_segment(2) inj_on_def pathfinish_subpath pathstart_subpath simple_path_subpath_eq) lemma simple_path_subpath: assumes "simple_path g" "u \ {0..1}" "v \ {0..1}" "u \ v" shows "simple_path(subpath u v g)" using assms unfolding simple_path_subpath_eq by (force simp: simple_path_def loop_free_def closed_segment_real_eq image_affinity_atLeastAtMost) lemma arc_simple_path_subpath: "\simple_path g; u \ {0..1}; v \ {0..1}; g u \ g v\ \ arc(subpath u v g)" by (force intro: simple_path_subpath simple_path_imp_arc) lemma arc_subpath_arc: "\arc g; u \ {0..1}; v \ {0..1}; u \ v\ \ arc(subpath u v g)" by (meson arc_def arc_imp_simple_path arc_simple_path_subpath inj_onD) lemma arc_simple_path_subpath_interior: "\simple_path g; u \ {0..1}; v \ {0..1}; u \ v; \u-v\ < 1\ \ arc(subpath u v g)" by (force simp: simple_path_def loop_free_def intro: arc_simple_path_subpath) lemma path_image_subpath_subset: "\u \ {0..1}; v \ {0..1}\ \ path_image(subpath u v g) \ path_image g" by (metis atLeastAtMost_iff atLeastatMost_subset_iff path_image_def path_image_subpath subset_image_iff) lemma join_subpaths_middle: "subpath (0) ((1 / 2)) p +++ subpath ((1 / 2)) 1 p = p" by (rule ext) (simp add: joinpaths_def subpath_def field_split_simps) subsection\<^marker>\tag unimportant\\There is a subpath to the frontier\ lemma subpath_to_frontier_explicit: fixes S :: "'a::metric_space set" assumes g: "path g" and "pathfinish g \ S" obtains u where "0 \ u" "u \ 1" "\x. 0 \ x \ x < u \ g x \ interior S" "(g u \ interior S)" "(u = 0 \ g u \ closure S)" proof - have gcon: "continuous_on {0..1} g" using g by (simp add: path_def) moreover have "bounded ({u. g u \ closure (- S)} \ {0..1})" using compact_eq_bounded_closed by fastforce ultimately have com: "compact ({0..1} \ {u. g u \ closure (- S)})" using closed_vimage_Int by (metis (full_types) Int_commute closed_atLeastAtMost closed_closure compact_eq_bounded_closed vimage_def) have "1 \ {u. g u \ closure (- S)}" using assms by (simp add: pathfinish_def closure_def) then have dis: "{0..1} \ {u. g u \ closure (- S)} \ {}" using atLeastAtMost_iff zero_le_one by blast then obtain u where "0 \ u" "u \ 1" and gu: "g u \ closure (- S)" and umin: "\t. \0 \ t; t \ 1; g t \ closure (- S)\ \ u \ t" using compact_attains_inf [OF com dis] by fastforce then have umin': "\t. \0 \ t; t \ 1; t < u\ \ g t \ S" using closure_def by fastforce have \
: "g u \ closure S" if "u \ 0" proof - have "u > 0" using that \0 \ u\ by auto { fix e::real assume "e > 0" obtain d where "d>0" and d: "\x'. \x' \ {0..1}; dist x' u \ d\ \ dist (g x') (g u) < e" using continuous_onE [OF gcon _ \e > 0\] \0 \ _\ \_ \ 1\ atLeastAtMost_iff by auto have *: "dist (max 0 (u - d / 2)) u \ d" using \0 \ u\ \u \ 1\ \d > 0\ by (simp add: dist_real_def) have "\y\S. dist y (g u) < e" using \0 < u\ \u \ 1\ \d > 0\ by (force intro: d [OF _ *] umin') } then show ?thesis by (simp add: frontier_def closure_approachable) qed show ?thesis proof show "\x. 0 \ x \ x < u \ g x \ interior S" using \u \ 1\ interior_closure umin by fastforce show "g u \ interior S" by (simp add: gu interior_closure) qed (use \0 \ u\ \u \ 1\ \
in auto) qed lemma subpath_to_frontier_strong: assumes g: "path g" and "pathfinish g \ S" obtains u where "0 \ u" "u \ 1" "g u \ interior S" "u = 0 \ (\x. 0 \ x \ x < 1 \ subpath 0 u g x \ interior S) \ g u \ closure S" proof - obtain u where "0 \ u" "u \ 1" and gxin: "\x. 0 \ x \ x < u \ g x \ interior S" and gunot: "(g u \ interior S)" and u0: "(u = 0 \ g u \ closure S)" using subpath_to_frontier_explicit [OF assms] by blast show ?thesis proof show "g u \ interior S" using gunot by blast qed (use \0 \ u\ \u \ 1\ u0 in \(force simp: subpath_def gxin)+\) qed lemma subpath_to_frontier: assumes g: "path g" and g0: "pathstart g \ closure S" and g1: "pathfinish g \ S" obtains u where "0 \ u" "u \ 1" "g u \ frontier S" "path_image(subpath 0 u g) - {g u} \ interior S" proof - obtain u where "0 \ u" "u \ 1" and notin: "g u \ interior S" and disj: "u = 0 \ (\x. 0 \ x \ x < 1 \ subpath 0 u g x \ interior S) \ g u \ closure S" (is "_ \ ?P") using subpath_to_frontier_strong [OF g g1] by blast show ?thesis proof show "g u \ frontier S" by (metis DiffI disj frontier_def g0 notin pathstart_def) show "path_image (subpath 0 u g) - {g u} \ interior S" using disj proof assume "u = 0" then show ?thesis by (simp add: path_image_subpath) next assume P: ?P show ?thesis proof (clarsimp simp add: path_image_subpath_gen) fix y assume y: "y \ closed_segment 0 u" "g y \ interior S" with \0 \ u\ have "0 \ y" "y \ u" by (auto simp: closed_segment_eq_real_ivl split: if_split_asm) then have "y=u \ subpath 0 u g (y/u) \ interior S" using P less_eq_real_def by force then show "g y = g u" using y by (auto simp: subpath_def split: if_split_asm) qed qed qed (use \0 \ u\ \u \ 1\ in auto) qed lemma exists_path_subpath_to_frontier: fixes S :: "'a::real_normed_vector set" assumes "path g" "pathstart g \ closure S" "pathfinish g \ S" obtains h where "path h" "pathstart h = pathstart g" "path_image h \ path_image g" "path_image h - {pathfinish h} \ interior S" "pathfinish h \ frontier S" proof - obtain u where u: "0 \ u" "u \ 1" "g u \ frontier S" "(path_image(subpath 0 u g) - {g u}) \ interior S" using subpath_to_frontier [OF assms] by blast show ?thesis proof show "path_image (subpath 0 u g) \ path_image g" by (simp add: path_image_subpath_subset u) show "pathstart (subpath 0 u g) = pathstart g" by (metis pathstart_def pathstart_subpath) qed (use assms u in \auto simp: path_image_subpath\) qed lemma exists_path_subpath_to_frontier_closed: fixes S :: "'a::real_normed_vector set" assumes S: "closed S" and g: "path g" and g0: "pathstart g \ S" and g1: "pathfinish g \ S" obtains h where "path h" "pathstart h = pathstart g" "path_image h \ path_image g \ S" "pathfinish h \ frontier S" by (smt (verit, del_insts) Diff_iff Int_iff S closure_closed exists_path_subpath_to_frontier frontier_def g g0 g1 interior_subset singletonD subset_eq) subsection \Shift Path to Start at Some Given Point\ definition\<^marker>\tag important\ shiftpath :: "real \ (real \ 'a::topological_space) \ real \ 'a" where "shiftpath a f = (\x. if (a + x) \ 1 then f (a + x) else f (a + x - 1))" lemma shiftpath_alt_def: "shiftpath a f = (\x. if x \ 1-a then f (a + x) else f (a + x - 1))" by (auto simp: shiftpath_def) lemma pathstart_shiftpath: "a \ 1 \ pathstart (shiftpath a g) = g a" unfolding pathstart_def shiftpath_def by auto lemma pathfinish_shiftpath: assumes "0 \ a" and "pathfinish g = pathstart g" shows "pathfinish (shiftpath a g) = g a" using assms unfolding pathstart_def pathfinish_def shiftpath_def by auto lemma endpoints_shiftpath: assumes "pathfinish g = pathstart g" and "a \ {0 .. 1}" shows "pathfinish (shiftpath a g) = g a" and "pathstart (shiftpath a g) = g a" using assms by (simp_all add: pathstart_shiftpath pathfinish_shiftpath) lemma closed_shiftpath: assumes "pathfinish g = pathstart g" and "a \ {0..1}" shows "pathfinish (shiftpath a g) = pathstart (shiftpath a g)" using endpoints_shiftpath[OF assms] by auto lemma path_shiftpath: assumes "path g" and "pathfinish g = pathstart g" and "a \ {0..1}" shows "path (shiftpath a g)" proof - have *: "{0 .. 1} = {0 .. 1-a} \ {1-a .. 1}" using assms(3) by auto have **: "\x. x + a = 1 \ g (x + a - 1) = g (x + a)" by (smt (verit, best) assms(2) pathfinish_def pathstart_def) show ?thesis unfolding path_def shiftpath_def * proof (rule continuous_on_closed_Un) have contg: "continuous_on {0..1} g" using \path g\ path_def by blast show "continuous_on {0..1-a} (\x. if a + x \ 1 then g (a + x) else g (a + x - 1))" proof (rule continuous_on_eq) show "continuous_on {0..1-a} (g \ (+) a)" by (intro continuous_intros continuous_on_subset [OF contg]) (use \a \ {0..1}\ in auto) qed auto show "continuous_on {1-a..1} (\x. if a + x \ 1 then g (a + x) else g (a + x - 1))" proof (rule continuous_on_eq) show "continuous_on {1-a..1} (g \ (+) (a - 1))" by (intro continuous_intros continuous_on_subset [OF contg]) (use \a \ {0..1}\ in auto) qed (auto simp: "**" add.commute add_diff_eq) qed auto qed lemma shiftpath_shiftpath: assumes "pathfinish g = pathstart g" and "a \ {0..1}" and "x \ {0..1}" shows "shiftpath (1 - a) (shiftpath a g) x = g x" using assms unfolding pathfinish_def pathstart_def shiftpath_def by auto lemma path_image_shiftpath: assumes a: "a \ {0..1}" and "pathfinish g = pathstart g" shows "path_image (shiftpath a g) = path_image g" proof - { fix x assume g: "g 1 = g 0" "x \ {0..1::real}" and gne: "\y. y\{0..1} \ {x. \ a + x \ 1} \ g x \ g (a + y - 1)" then have "\y\{0..1} \ {x. a + x \ 1}. g x = g (a + y)" proof (cases "a \ x") case False then show ?thesis apply (rule_tac x="1 + x - a" in bexI) using g gne[of "1 + x - a"] a by (force simp: field_simps)+ next case True then show ?thesis using g a by (rule_tac x="x - a" in bexI) (auto simp: field_simps) qed } then show ?thesis using assms unfolding shiftpath_def path_image_def pathfinish_def pathstart_def by (auto simp: image_iff) qed lemma loop_free_shiftpath: assumes "loop_free g" "pathfinish g = pathstart g" and a: "0 \ a" "a \ 1" shows "loop_free (shiftpath a g)" unfolding loop_free_def proof (intro conjI impI ballI) show "x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" if "x \ {0..1}" "y \ {0..1}" "shiftpath a g x = shiftpath a g y" for x y using that a assms unfolding shiftpath_def loop_free_def by (smt (verit, ccfv_threshold) atLeastAtMost_iff) qed lemma simple_path_shiftpath: assumes "simple_path g" "pathfinish g = pathstart g" and a: "0 \ a" "a \ 1" shows "simple_path (shiftpath a g)" using assms loop_free_shiftpath path_shiftpath simple_path_def by fastforce subsection \Straight-Line Paths\ definition\<^marker>\tag important\ linepath :: "'a::real_normed_vector \ 'a \ real \ 'a" where "linepath a b = (\x. (1 - x) *\<^sub>R a + x *\<^sub>R b)" lemma pathstart_linepath[simp]: "pathstart (linepath a b) = a" unfolding pathstart_def linepath_def by auto lemma pathfinish_linepath[simp]: "pathfinish (linepath a b) = b" unfolding pathfinish_def linepath_def by auto lemma linepath_inner: "linepath a b x \ v = linepath (a \ v) (b \ v) x" by (simp add: linepath_def algebra_simps) lemma Re_linepath': "Re (linepath a b x) = linepath (Re a) (Re b) x" by (simp add: linepath_def) lemma Im_linepath': "Im (linepath a b x) = linepath (Im a) (Im b) x" by (simp add: linepath_def) lemma linepath_0': "linepath a b 0 = a" by (simp add: linepath_def) lemma linepath_1': "linepath a b 1 = b" by (simp add: linepath_def) lemma continuous_linepath_at[intro]: "continuous (at x) (linepath a b)" unfolding linepath_def by (intro continuous_intros) lemma continuous_on_linepath [intro,continuous_intros]: "continuous_on s (linepath a b)" using continuous_linepath_at by (auto intro!: continuous_at_imp_continuous_on) lemma path_linepath[iff]: "path (linepath a b)" unfolding path_def by (rule continuous_on_linepath) lemma path_image_linepath[simp]: "path_image (linepath a b) = closed_segment a b" unfolding path_image_def segment linepath_def by auto lemma reversepath_linepath[simp]: "reversepath (linepath a b) = linepath b a" unfolding reversepath_def linepath_def by auto lemma linepath_0 [simp]: "linepath 0 b x = x *\<^sub>R b" by (simp add: linepath_def) lemma linepath_cnj: "cnj (linepath a b x) = linepath (cnj a) (cnj b) x" by (simp add: linepath_def) lemma arc_linepath: assumes "a \ b" shows [simp]: "arc (linepath a b)" proof - { fix x y :: "real" assume "x *\<^sub>R b + y *\<^sub>R a = x *\<^sub>R a + y *\<^sub>R b" then have "(x - y) *\<^sub>R a = (x - y) *\<^sub>R b" by (simp add: algebra_simps) with assms have "x = y" by simp } then show ?thesis unfolding arc_def inj_on_def by (fastforce simp: algebra_simps linepath_def) qed lemma simple_path_linepath[intro]: "a \ b \ simple_path (linepath a b)" by (simp add: arc_imp_simple_path) lemma linepath_trivial [simp]: "linepath a a x = a" by (simp add: linepath_def real_vector.scale_left_diff_distrib) lemma linepath_refl: "linepath a a = (\x. a)" by auto lemma subpath_refl: "subpath a a g = linepath (g a) (g a)" by (simp add: subpath_def linepath_def algebra_simps) lemma linepath_of_real: "(linepath (of_real a) (of_real b) x) = of_real ((1 - x)*a + x*b)" by (simp add: scaleR_conv_of_real linepath_def) lemma of_real_linepath: "of_real (linepath a b x) = linepath (of_real a) (of_real b) x" by (metis linepath_of_real mult.right_neutral of_real_def real_scaleR_def) lemma inj_on_linepath: assumes "a \ b" shows "inj_on (linepath a b) {0..1}" using arc_imp_inj_on arc_linepath assms by blast lemma linepath_le_1: fixes a::"'a::linordered_idom" shows "\a \ 1; b \ 1; 0 \ u; u \ 1\ \ (1 - u) * a + u * b \ 1" using mult_left_le [of a "1-u"] mult_left_le [of b u] by auto lemma linepath_in_path: shows "x \ {0..1} \ linepath a b x \ closed_segment a b" by (auto simp: segment linepath_def) lemma linepath_image_01: "linepath a b ` {0..1} = closed_segment a b" by (auto simp: segment linepath_def) lemma linepath_in_convex_hull: fixes x::real assumes "a \ convex hull S" and "b \ convex hull S" and "0\x" "x\1" shows "linepath a b x \ convex hull S" by (meson assms atLeastAtMost_iff convex_contains_segment convex_convex_hull linepath_in_path subset_eq) lemma Re_linepath: "Re(linepath (of_real a) (of_real b) x) = (1 - x)*a + x*b" by (simp add: linepath_def) lemma Im_linepath: "Im(linepath (of_real a) (of_real b) x) = 0" by (simp add: linepath_def) lemma bounded_linear_linepath: assumes "bounded_linear f" shows "f (linepath a b x) = linepath (f a) (f b) x" proof - interpret f: bounded_linear f by fact show ?thesis by (simp add: linepath_def f.add f.scale) qed lemma bounded_linear_linepath': assumes "bounded_linear f" shows "f \ linepath a b = linepath (f a) (f b)" using bounded_linear_linepath[OF assms] by (simp add: fun_eq_iff) lemma linepath_cnj': "cnj \ linepath a b = linepath (cnj a) (cnj b)" by (simp add: linepath_def fun_eq_iff) lemma differentiable_linepath [intro]: "linepath a b differentiable at x within A" by (auto simp: linepath_def) lemma has_vector_derivative_linepath_within: "(linepath a b has_vector_derivative (b - a)) (at x within S)" by (force intro: derivative_eq_intros simp add: linepath_def has_vector_derivative_def algebra_simps) subsection\<^marker>\tag unimportant\\Segments via convex hulls\ lemma segments_subset_convex_hull: "closed_segment a b \ (convex hull {a,b,c})" "closed_segment a c \ (convex hull {a,b,c})" "closed_segment b c \ (convex hull {a,b,c})" "closed_segment b a \ (convex hull {a,b,c})" "closed_segment c a \ (convex hull {a,b,c})" "closed_segment c b \ (convex hull {a,b,c})" by (auto simp: segment_convex_hull linepath_of_real elim!: rev_subsetD [OF _ hull_mono]) lemma midpoints_in_convex_hull: assumes "x \ convex hull s" "y \ convex hull s" shows "midpoint x y \ convex hull s" using assms closed_segment_subset_convex_hull csegment_midpoint_subset by blast lemma not_in_interior_convex_hull_3: fixes a :: "complex" shows "a \ interior(convex hull {a,b,c})" "b \ interior(convex hull {a,b,c})" "c \ interior(convex hull {a,b,c})" by (auto simp: card_insert_le_m1 not_in_interior_convex_hull) lemma midpoint_in_closed_segment [simp]: "midpoint a b \ closed_segment a b" using midpoints_in_convex_hull segment_convex_hull by blast lemma midpoint_in_open_segment [simp]: "midpoint a b \ open_segment a b \ a \ b" by (simp add: open_segment_def) lemma continuous_IVT_local_extremum: fixes f :: "'a::euclidean_space \ real" assumes contf: "continuous_on (closed_segment a b) f" and ab: "a \ b" "f a = f b" obtains z where "z \ open_segment a b" "(\w \ closed_segment a b. (f w) \ (f z)) \ (\w \ closed_segment a b. (f z) \ (f w))" proof - obtain c where "c \ closed_segment a b" and c: "\y. y \ closed_segment a b \ f y \ f c" using continuous_attains_sup [of "closed_segment a b" f] contf by auto moreover obtain d where "d \ closed_segment a b" and d: "\y. y \ closed_segment a b \ f d \ f y" using continuous_attains_inf [of "closed_segment a b" f] contf by auto ultimately show ?thesis by (smt (verit) UnE ab closed_segment_eq_open empty_iff insert_iff midpoint_in_open_segment that) qed text\An injective map into R is also an open map w.r.T. the universe, and conversely. \ proposition injective_eq_1d_open_map_UNIV: fixes f :: "real \ real" assumes contf: "continuous_on S f" and S: "is_interval S" shows "inj_on f S \ (\T. open T \ T \ S \ open(f ` T))" (is "?lhs = ?rhs") proof safe fix T assume injf: ?lhs and "open T" and "T \ S" have "\U. open U \ f x \ U \ U \ f ` T" if "x \ T" for x proof - obtain \ where "\ > 0" and \: "cball x \ \ T" using \open T\ \x \ T\ open_contains_cball_eq by blast show ?thesis proof (intro exI conjI) have "closed_segment (x-\) (x+\) = {x-\..x+\}" using \0 < \\ by (auto simp: closed_segment_eq_real_ivl) also have "\ \ S" using \ \T \ S\ by (auto simp: dist_norm subset_eq) finally have "f ` (open_segment (x-\) (x+\)) = open_segment (f (x-\)) (f (x+\))" using continuous_injective_image_open_segment_1 by (metis continuous_on_subset [OF contf] inj_on_subset [OF injf]) then show "open (f ` {x-\<..})" using \0 < \\ by (simp add: open_segment_eq_real_ivl) show "f x \ f ` {x - \<..}" by (auto simp: \\ > 0\) show "f ` {x - \<..} \ f ` T" using \ by (auto simp: dist_norm subset_iff) qed qed with open_subopen show "open (f ` T)" by blast next assume R: ?rhs have False if xy: "x \ S" "y \ S" and "f x = f y" "x \ y" for x y proof - have "open (f ` open_segment x y)" using R by (metis S convex_contains_open_segment is_interval_convex open_greaterThanLessThan open_segment_eq_real_ivl xy) moreover have "continuous_on (closed_segment x y) f" by (meson S closed_segment_subset contf continuous_on_subset is_interval_convex that) then obtain \ where "\ \ open_segment x y" and \: "(\w \ closed_segment x y. (f w) \ (f \)) \ (\w \ closed_segment x y. (f \) \ (f w))" using continuous_IVT_local_extremum [of x y f] \f x = f y\ \x \ y\ by blast ultimately obtain e where "e>0" and e: "\u. dist u (f \) < e \ u \ f ` open_segment x y" using open_dist by (metis image_eqI) have fin: "f \ + (e/2) \ f ` open_segment x y" "f \ - (e/2) \ f ` open_segment x y" using e [of "f \ + (e/2)"] e [of "f \ - (e/2)"] \e > 0\ by (auto simp: dist_norm) show ?thesis using \ \0 < e\ fin open_closed_segment by fastforce qed then show ?lhs by (force simp: inj_on_def) qed subsection\<^marker>\tag unimportant\ \Bounding a point away from a path\ lemma not_on_path_ball: fixes g :: "real \ 'a::heine_borel" assumes "path g" and z: "z \ path_image g" shows "\e > 0. ball z e \ path_image g = {}" proof - have "closed (path_image g)" by (simp add: \path g\ closed_path_image) then obtain a where "a \ path_image g" "\y \ path_image g. dist z a \ dist z y" by (auto intro: distance_attains_inf[OF _ path_image_nonempty, of g z]) then show ?thesis by (rule_tac x="dist z a" in exI) (use dist_commute z in auto) qed lemma not_on_path_cball: fixes g :: "real \ 'a::heine_borel" assumes "path g" and "z \ path_image g" shows "\e>0. cball z e \ (path_image g) = {}" by (smt (verit, ccfv_threshold) open_ball assms centre_in_ball inf.orderE inf_assoc inf_bot_right not_on_path_ball open_contains_cball_eq) subsection \Path component\ text \Original formalization by Tom Hales\ definition\<^marker>\tag important\ "path_component S x y \ (\g. path g \ path_image g \ S \ pathstart g = x \ pathfinish g = y)" abbreviation\<^marker>\tag important\ "path_component_set S x \ Collect (path_component S x)" lemmas path_defs = path_def pathstart_def pathfinish_def path_image_def path_component_def lemma path_component_mem: assumes "path_component S x y" shows "x \ S" and "y \ S" using assms unfolding path_defs by auto lemma path_component_refl: assumes "x \ S" shows "path_component S x x" using assms unfolding path_defs by (metis (full_types) assms continuous_on_const image_subset_iff path_image_def) lemma path_component_refl_eq: "path_component S x x \ x \ S" by (auto intro!: path_component_mem path_component_refl) lemma path_component_sym: "path_component S x y \ path_component S y x" unfolding path_component_def by (metis (no_types) path_image_reversepath path_reversepath pathfinish_reversepath pathstart_reversepath) lemma path_component_trans: assumes "path_component S x y" and "path_component S y z" shows "path_component S x z" using assms unfolding path_component_def by (metis path_join pathfinish_join pathstart_join subset_path_image_join) lemma path_component_of_subset: "S \ T \ path_component S x y \ path_component T x y" unfolding path_component_def by auto lemma path_component_linepath: fixes S :: "'a::real_normed_vector set" shows "closed_segment a b \ S \ path_component S a b" unfolding path_component_def by fastforce subsubsection\<^marker>\tag unimportant\ \Path components as sets\ lemma path_component_set: "path_component_set S x = {y. (\g. path g \ path_image g \ S \ pathstart g = x \ pathfinish g = y)}" by (auto simp: path_component_def) lemma path_component_subset: "path_component_set S x \ S" by (auto simp: path_component_mem(2)) lemma path_component_eq_empty: "path_component_set S x = {} \ x \ S" using path_component_mem path_component_refl_eq by fastforce lemma path_component_mono: "S \ T \ (path_component_set S x) \ (path_component_set T x)" by (simp add: Collect_mono path_component_of_subset) lemma path_component_eq: "y \ path_component_set S x \ path_component_set S y = path_component_set S x" by (metis (no_types, lifting) Collect_cong mem_Collect_eq path_component_sym path_component_trans) subsection \Path connectedness of a space\ definition\<^marker>\tag important\ "path_connected S \ (\x\S. \y\S. \g. path g \ path_image g \ S \ pathstart g = x \ pathfinish g = y)" lemma path_connectedin_iff_path_connected_real [simp]: "path_connectedin euclideanreal S \ path_connected S" by (simp add: path_connectedin path_connected_def path_defs image_subset_iff_funcset) lemma path_connected_component: "path_connected S \ (\x\S. \y\S. path_component S x y)" unfolding path_connected_def path_component_def by auto lemma path_connected_component_set: "path_connected S \ (\x\S. path_component_set S x = S)" unfolding path_connected_component path_component_subset using path_component_mem by blast lemma path_component_maximal: "\x \ T; path_connected T; T \ S\ \ T \ (path_component_set S x)" by (metis path_component_mono path_connected_component_set) lemma convex_imp_path_connected: fixes S :: "'a::real_normed_vector set" assumes "convex S" shows "path_connected S" unfolding path_connected_def using assms convex_contains_segment by fastforce lemma path_connected_UNIV [iff]: "path_connected (UNIV :: 'a::real_normed_vector set)" by (simp add: convex_imp_path_connected) lemma path_component_UNIV: "path_component_set UNIV x = (UNIV :: 'a::real_normed_vector set)" using path_connected_component_set by auto lemma path_connected_imp_connected: assumes "path_connected S" shows "connected S" proof (rule connectedI) fix e1 e2 assume as: "open e1" "open e2" "S \ e1 \ e2" "e1 \ e2 \ S = {}" "e1 \ S \ {}" "e2 \ S \ {}" then obtain x1 x2 where obt:"x1 \ e1 \ S" "x2 \ e2 \ S" by auto then obtain g where g: "path g" "path_image g \ S" and pg: "pathstart g = x1" "pathfinish g = x2" using assms[unfolded path_connected_def,rule_format,of x1 x2] by auto have *: "connected {0..1::real}" by (auto intro!: convex_connected) have "{0..1} \ {x \ {0..1}. g x \ e1} \ {x \ {0..1}. g x \ e2}" using as(3) g(2)[unfolded path_defs] by blast moreover have "{x \ {0..1}. g x \ e1} \ {x \ {0..1}. g x \ e2} = {}" using as(4) g(2)[unfolded path_defs] unfolding subset_eq by auto moreover have "{x \ {0..1}. g x \ e1} \ {} \ {x \ {0..1}. g x \ e2} \ {}" by (smt (verit, ccfv_threshold) IntE atLeastAtMost_iff empty_iff pg mem_Collect_eq obt pathfinish_def pathstart_def) ultimately show False using *[unfolded connected_local not_ex, rule_format, of "{0..1} \ g -` e1" "{0..1} \ g -` e2"] using continuous_openin_preimage_gen[OF g(1)[unfolded path_def] as(1)] using continuous_openin_preimage_gen[OF g(1)[unfolded path_def] as(2)] by auto qed lemma open_path_component: fixes S :: "'a::real_normed_vector set" assumes "open S" shows "open (path_component_set S x)" unfolding open_contains_ball by (metis assms centre_in_ball convex_ball convex_imp_path_connected equals0D openE path_component_eq path_component_eq_empty path_component_maximal) lemma open_non_path_component: fixes S :: "'a::real_normed_vector set" assumes "open S" shows "open (S - path_component_set S x)" unfolding open_contains_ball proof fix y assume y: "y \ S - path_component_set S x" then obtain e where e: "e > 0" "ball y e \ S" using assms openE by auto show "\e>0. ball y e \ S - path_component_set S x" proof (intro exI conjI subsetI DiffI notI) show "\x. x \ ball y e \ x \ S" using e by blast show False if "z \ ball y e" "z \ path_component_set S x" for z by (metis (no_types, lifting) Diff_iff centre_in_ball convex_ball convex_imp_path_connected path_component_eq path_component_maximal subsetD that y e) qed (use e in auto) qed lemma connected_open_path_connected: fixes S :: "'a::real_normed_vector set" assumes "open S" and "connected S" shows "path_connected S" unfolding path_connected_component_set proof (rule, rule, rule path_component_subset, rule) fix x y assume "x \ S" and "y \ S" show "y \ path_component_set S x" proof (rule ccontr) assume "\ ?thesis" moreover have "path_component_set S x \ S \ {}" using \x \ S\ path_component_eq_empty path_component_subset[of S x] by auto ultimately show False using \y \ S\ open_non_path_component[OF \open S\] open_path_component[OF \open S\] using \connected S\[unfolded connected_def not_ex, rule_format, of "path_component_set S x" "S - path_component_set S x"] by auto qed qed lemma path_connected_continuous_image: assumes contf: "continuous_on S f" and "path_connected S" shows "path_connected (f ` S)" unfolding path_connected_def proof clarsimp fix x y assume x: "x \ S" and y: "y \ S" with \path_connected S\ show "\g. path g \ path_image g \ f ` S \ pathstart g = f x \ pathfinish g = f y" unfolding path_defs path_connected_def using continuous_on_subset[OF contf] by (smt (verit, ccfv_threshold) continuous_on_compose2 image_eqI image_subset_iff) qed lemma path_connected_translationI: fixes a :: "'a :: topological_group_add" assumes "path_connected S" shows "path_connected ((\x. a + x) ` S)" by (intro path_connected_continuous_image assms continuous_intros) lemma path_connected_translation: fixes a :: "'a :: topological_group_add" shows "path_connected ((\x. a + x) ` S) = path_connected S" proof - have "\x y. (+) (x::'a) ` (+) (0 - x) ` y = y" by (simp add: image_image) then show ?thesis by (metis (no_types) path_connected_translationI) qed lemma path_connected_segment [simp]: fixes a :: "'a::real_normed_vector" shows "path_connected (closed_segment a b)" by (simp add: convex_imp_path_connected) lemma path_connected_open_segment [simp]: fixes a :: "'a::real_normed_vector" shows "path_connected (open_segment a b)" by (simp add: convex_imp_path_connected) lemma homeomorphic_path_connectedness: "S homeomorphic T \ path_connected S \ path_connected T" unfolding homeomorphic_def homeomorphism_def by (metis path_connected_continuous_image) lemma path_connected_empty [simp]: "path_connected {}" unfolding path_connected_def by auto lemma path_connected_singleton [simp]: "path_connected {a}" unfolding path_connected_def pathstart_def pathfinish_def path_image_def using path_def by fastforce lemma path_connected_Un: assumes "path_connected S" and "path_connected T" and "S \ T \ {}" shows "path_connected (S \ T)" unfolding path_connected_component proof (intro ballI) fix x y assume x: "x \ S \ T" and y: "y \ S \ T" from assms obtain z where z: "z \ S" "z \ T" by auto with x y show "path_component (S \ T) x y" by (smt (verit) assms(1,2) in_mono mem_Collect_eq path_component_eq path_component_maximal sup.bounded_iff sup.cobounded2 sup_ge1) qed lemma path_connected_UNION: assumes "\i. i \ A \ path_connected (S i)" and "\i. i \ A \ z \ S i" shows "path_connected (\i\A. S i)" unfolding path_connected_component proof clarify fix x i y j assume *: "i \ A" "x \ S i" "j \ A" "y \ S j" then have "path_component (S i) x z" and "path_component (S j) z y" using assms by (simp_all add: path_connected_component) then have "path_component (\i\A. S i) x z" and "path_component (\i\A. S i) z y" using *(1,3) by (meson SUP_upper path_component_of_subset)+ then show "path_component (\i\A. S i) x y" by (rule path_component_trans) qed lemma path_component_path_image_pathstart: assumes p: "path p" and x: "x \ path_image p" shows "path_component (path_image p) (pathstart p) x" proof - obtain y where x: "x = p y" and y: "0 \ y" "y \ 1" using x by (auto simp: path_image_def) show ?thesis unfolding path_component_def proof (intro exI conjI) have "continuous_on ((*) y ` {0..1}) p" by (simp add: continuous_on_path image_mult_atLeastAtMost_if p y) then have "continuous_on {0..1} (p \ ((*) y))" using continuous_on_compose continuous_on_mult_const by blast then show "path (\u. p (y * u))" by (simp add: path_def) show "path_image (\u. p (y * u)) \ path_image p" using y mult_le_one by (fastforce simp: path_image_def image_iff) qed (auto simp: pathstart_def pathfinish_def x) qed lemma path_connected_path_image: "path p \ path_connected(path_image p)" unfolding path_connected_component by (meson path_component_path_image_pathstart path_component_sym path_component_trans) lemma path_connected_path_component [simp]: "path_connected (path_component_set S x)" by (smt (verit) mem_Collect_eq path_component_def path_component_eq path_component_maximal path_connected_component path_connected_path_image pathstart_in_path_image) lemma path_component: "path_component S x y \ (\t. path_connected t \ t \ S \ x \ t \ y \ t)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (metis path_component_def path_connected_path_image pathfinish_in_path_image pathstart_in_path_image) next assume ?rhs then show ?lhs by (meson path_component_of_subset path_connected_component) qed lemma path_component_path_component [simp]: "path_component_set (path_component_set S x) x = path_component_set S x" by (metis (full_types) mem_Collect_eq path_component_eq_empty path_component_refl path_connected_component_set path_connected_path_component) lemma path_component_subset_connected_component: "(path_component_set S x) \ (connected_component_set S x)" proof (cases "x \ S") case True show ?thesis by (simp add: True connected_component_maximal path_component_refl path_component_subset path_connected_imp_connected) next case False then show ?thesis using path_component_eq_empty by auto qed subsection\<^marker>\tag unimportant\\Lemmas about path-connectedness\ lemma path_connected_linear_image: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "path_connected S" "bounded_linear f" shows "path_connected(f ` S)" by (auto simp: linear_continuous_on assms path_connected_continuous_image) lemma is_interval_path_connected: "is_interval S \ path_connected S" by (simp add: convex_imp_path_connected is_interval_convex) lemma path_connected_Ioi[simp]: "path_connected {a<..}" for a :: real by (simp add: convex_imp_path_connected) lemma path_connected_Ici[simp]: "path_connected {a..}" for a :: real by (simp add: convex_imp_path_connected) lemma path_connected_Iio[simp]: "path_connected {.. (\x \ topspace X. \y \ topspace X. \S. path_connectedin X S \ x \ S \ y \ S)" by (metis path_connectedin path_connectedin_topspace path_connected_space_def) lemma connectedin_path_image: "pathin X g \ connectedin X (g ` ({0..1}))" by (simp add: path_connectedin_imp_connectedin path_connectedin_path_image) lemma compactin_path_image: "pathin X g \ compactin X (g ` ({0..1}))" unfolding pathin_def by (rule image_compactin [of "top_of_set {0..1}"]) auto lemma linear_homeomorphism_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" obtains g where "homeomorphism (f ` S) S g f" proof - obtain g where "linear g" "g \ f = id" using assms linear_injective_left_inverse by blast then have "homeomorphism (f ` S) S g f" using assms unfolding homeomorphism_def by (auto simp: eq_id_iff [symmetric] image_comp linear_conv_bounded_linear linear_continuous_on) then show thesis .. qed lemma linear_homeomorphic_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "S homeomorphic f ` S" by (meson homeomorphic_def homeomorphic_sym linear_homeomorphism_image [OF assms]) lemma path_connected_Times: assumes "path_connected s" "path_connected t" shows "path_connected (s \ t)" proof (simp add: path_connected_def Sigma_def, clarify) fix x1 y1 x2 y2 assume "x1 \ s" "y1 \ t" "x2 \ s" "y2 \ t" obtain g where "path g" and g: "path_image g \ s" and gs: "pathstart g = x1" and gf: "pathfinish g = x2" using \x1 \ s\ \x2 \ s\ assms by (force simp: path_connected_def) obtain h where "path h" and h: "path_image h \ t" and hs: "pathstart h = y1" and hf: "pathfinish h = y2" using \y1 \ t\ \y2 \ t\ assms by (force simp: path_connected_def) have "path (\z. (x1, h z))" using \path h\ unfolding path_def by (intro continuous_intros continuous_on_compose2 [where g = "Pair _"]; force) moreover have "path (\z. (g z, y2))" using \path g\ unfolding path_def by (intro continuous_intros continuous_on_compose2 [where g = "Pair _"]; force) ultimately have 1: "path ((\z. (x1, h z)) +++ (\z. (g z, y2)))" by (metis hf gs path_join_imp pathstart_def pathfinish_def) have "path_image ((\z. (x1, h z)) +++ (\z. (g z, y2))) \ path_image (\z. (x1, h z)) \ path_image (\z. (g z, y2))" by (rule Path_Connected.path_image_join_subset) also have "\ \ (\x\s. \x1\t. {(x, x1)})" using g h \x1 \ s\ \y2 \ t\ by (force simp: path_image_def) finally have 2: "path_image ((\z. (x1, h z)) +++ (\z. (g z, y2))) \ (\x\s. \x1\t. {(x, x1)})" . show "\g. path g \ path_image g \ (\x\s. \x1\t. {(x, x1)}) \ pathstart g = (x1, y1) \ pathfinish g = (x2, y2)" using 1 2 gf hs by (metis (no_types, lifting) pathfinish_def pathfinish_join pathstart_def pathstart_join) qed lemma is_interval_path_connected_1: fixes s :: "real set" shows "is_interval s \ path_connected s" using is_interval_connected_1 is_interval_path_connected path_connected_imp_connected by blast subsection\<^marker>\tag unimportant\\Path components\ lemma Union_path_component [simp]: "Union {path_component_set S x |x. x \ S} = S" using path_component_subset path_component_refl by blast lemma path_component_disjoint: "disjnt (path_component_set S a) (path_component_set S b) \ (a \ path_component_set S b)" unfolding disjnt_iff using path_component_sym path_component_trans by blast lemma path_component_eq_eq: "path_component S x = path_component S y \ (x \ S) \ (y \ S) \ x \ S \ y \ S \ path_component S x y" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (metis (no_types) path_component_mem(1) path_component_refl) next assume ?rhs then show ?lhs proof assume "x \ S \ y \ S" then show ?lhs by (metis Collect_empty_eq_bot path_component_eq_empty) next assume S: "x \ S \ y \ S \ path_component S x y" show ?lhs by (rule ext) (metis S path_component_trans path_component_sym) qed qed lemma path_component_unique: assumes "x \ C" "C \ S" "path_connected C" "\C'. \x \ C'; C' \ S; path_connected C'\ \ C' \ C" shows "path_component_set S x = C" by (smt (verit, best) Collect_cong assms path_component path_component_of_subset path_connected_component_set) lemma path_component_intermediate_subset: "path_component_set U a \ T \ T \ U \ path_component_set T a = path_component_set U a" by (metis (no_types) path_component_mono path_component_path_component subset_antisym) lemma complement_path_component_Union: fixes x :: "'a :: topological_space" shows "S - path_component_set S x = \({path_component_set S y| y. y \ S} - {path_component_set S x})" proof - have *: "(\x. x \ S - {a} \ disjnt a x) \ \S - a = \(S - {a})" for a::"'a set" and S by (auto simp: disjnt_def) have "\y. y \ {path_component_set S x |x. x \ S} - {path_component_set S x} \ disjnt (path_component_set S x) y" using path_component_disjoint path_component_eq by fastforce then have "\{path_component_set S x |x. x \ S} - path_component_set S x = \({path_component_set S y |y. y \ S} - {path_component_set S x})" by (meson *) then show ?thesis by simp qed subsection\Path components\ definition path_component_of where "path_component_of X x y \ \g. pathin X g \ g 0 = x \ g 1 = y" abbreviation path_component_of_set where "path_component_of_set X x \ Collect (path_component_of X x)" definition path_components_of :: "'a topology \ 'a set set" where "path_components_of X \ path_component_of_set X ` topspace X" lemma pathin_canon_iff: "pathin (top_of_set T) g \ path g \ g \ {0..1} \ T" by (simp add: path_def pathin_def image_subset_iff_funcset) lemma path_component_of_canon_iff [simp]: "path_component_of (top_of_set T) a b \ path_component T a b" by (simp add: path_component_of_def pathin_canon_iff path_defs image_subset_iff_funcset) lemma path_component_in_topspace: "path_component_of X x y \ x \ topspace X \ y \ topspace X" by (auto simp: path_component_of_def pathin_def continuous_map_def) lemma path_component_of_refl: "path_component_of X x x \ x \ topspace X" by (metis path_component_in_topspace path_component_of_def pathin_const) lemma path_component_of_sym: assumes "path_component_of X x y" shows "path_component_of X y x" using assms apply (clarsimp simp: path_component_of_def pathin_def) apply (rule_tac x="g \ (\t. 1 - t)" in exI) apply (auto intro!: continuous_map_compose simp: continuous_map_in_subtopology continuous_on_op_minus) done lemma path_component_of_sym_iff: "path_component_of X x y \ path_component_of X y x" by (metis path_component_of_sym) lemma continuous_map_cases_le: assumes contp: "continuous_map X euclideanreal p" and contq: "continuous_map X euclideanreal q" and contf: "continuous_map (subtopology X {x. x \ topspace X \ p x \ q x}) Y f" and contg: "continuous_map (subtopology X {x. x \ topspace X \ q x \ p x}) Y g" and fg: "\x. \x \ topspace X; p x = q x\ \ f x = g x" shows "continuous_map X Y (\x. if p x \ q x then f x else g x)" proof - have "continuous_map X Y (\x. if q x - p x \ {0..} then f x else g x)" proof (rule continuous_map_cases_function) show "continuous_map X euclideanreal (\x. q x - p x)" by (intro contp contq continuous_intros) show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of {0..}}) Y f" by (simp add: contf) show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of (topspace euclideanreal - {0..})}) Y g" by (simp add: contg flip: Compl_eq_Diff_UNIV) qed (auto simp: fg) then show ?thesis by simp qed lemma continuous_map_cases_lt: assumes contp: "continuous_map X euclideanreal p" and contq: "continuous_map X euclideanreal q" and contf: "continuous_map (subtopology X {x. x \ topspace X \ p x \ q x}) Y f" and contg: "continuous_map (subtopology X {x. x \ topspace X \ q x \ p x}) Y g" and fg: "\x. \x \ topspace X; p x = q x\ \ f x = g x" shows "continuous_map X Y (\x. if p x < q x then f x else g x)" proof - have "continuous_map X Y (\x. if q x - p x \ {0<..} then f x else g x)" proof (rule continuous_map_cases_function) show "continuous_map X euclideanreal (\x. q x - p x)" by (intro contp contq continuous_intros) show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of {0<..}}) Y f" by (simp add: contf) show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of (topspace euclideanreal - {0<..})}) Y g" by (simp add: contg flip: Compl_eq_Diff_UNIV) qed (auto simp: fg) then show ?thesis by simp qed lemma path_component_of_trans: assumes "path_component_of X x y" and "path_component_of X y z" shows "path_component_of X x z" unfolding path_component_of_def pathin_def proof - let ?T01 = "top_of_set {0..1::real}" obtain g1 g2 where g1: "continuous_map ?T01 X g1" "x = g1 0" "y = g1 1" and g2: "continuous_map ?T01 X g2" "g2 0 = g1 1" "z = g2 1" using assms unfolding path_component_of_def pathin_def by blast let ?g = "\x. if x \ 1/2 then (g1 \ (\t. 2 * t)) x else (g2 \ (\t. 2 * t -1)) x" show "\g. continuous_map ?T01 X g \ g 0 = x \ g 1 = z" proof (intro exI conjI) show "continuous_map (subtopology euclideanreal {0..1}) X ?g" proof (intro continuous_map_cases_le continuous_map_compose, force, force) show "continuous_map (subtopology ?T01 {x \ topspace ?T01. x \ 1/2}) ?T01 ((*) 2)" by (auto simp: continuous_map_in_subtopology continuous_map_from_subtopology) have "continuous_map (subtopology (top_of_set {0..1}) {x. 0 \ x \ x \ 1 \ 1 \ x * 2}) euclideanreal (\t. 2 * t - 1)" by (intro continuous_intros) (force intro: continuous_map_from_subtopology) then show "continuous_map (subtopology ?T01 {x \ topspace ?T01. 1/2 \ x}) ?T01 (\t. 2 * t - 1)" by (force simp: continuous_map_in_subtopology) show "(g1 \ (*) 2) x = (g2 \ (\t. 2 * t - 1)) x" if "x \ topspace ?T01" "x = 1/2" for x using that by (simp add: g2(2) mult.commute continuous_map_from_subtopology) qed (auto simp: g1 g2) qed (auto simp: g1 g2) qed lemma path_component_of_mono: "\path_component_of (subtopology X S) x y; S \ T\ \ path_component_of (subtopology X T) x y" unfolding path_component_of_def by (metis subsetD pathin_subtopology) lemma path_component_of: "path_component_of X x y \ (\T. path_connectedin X T \ x \ T \ y \ T)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (metis atLeastAtMost_iff image_eqI order_refl path_component_of_def path_connectedin_path_image zero_le_one) next assume ?rhs then show ?lhs by (metis path_component_of_def path_connectedin) qed lemma path_component_of_set: "path_component_of X x y \ (\g. pathin X g \ g 0 = x \ g 1 = y)" by (auto simp: path_component_of_def) lemma path_component_of_subset_topspace: "Collect(path_component_of X x) \ topspace X" using path_component_in_topspace by fastforce lemma path_component_of_eq_empty: "Collect(path_component_of X x) = {} \ (x \ topspace X)" using path_component_in_topspace path_component_of_refl by fastforce lemma path_connected_space_iff_path_component: "path_connected_space X \ (\x \ topspace X. \y \ topspace X. path_component_of X x y)" by (simp add: path_component_of path_connected_space_subconnected) lemma path_connected_space_imp_path_component_of: "\path_connected_space X; a \ topspace X; b \ topspace X\ \ path_component_of X a b" by (simp add: path_connected_space_iff_path_component) lemma path_connected_space_path_component_set: "path_connected_space X \ (\x \ topspace X. Collect(path_component_of X x) = topspace X)" using path_component_of_subset_topspace path_connected_space_iff_path_component by fastforce lemma path_component_of_maximal: "\path_connectedin X s; x \ s\ \ s \ Collect(path_component_of X x)" using path_component_of by fastforce lemma path_component_of_equiv: "path_component_of X x y \ x \ topspace X \ y \ topspace X \ path_component_of X x = path_component_of X y" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding fun_eq_iff path_component_in_topspace by (metis path_component_in_topspace path_component_of_sym path_component_of_trans) qed (simp add: path_component_of_refl) lemma path_component_of_disjoint: "disjnt (Collect (path_component_of X x)) (Collect (path_component_of X y)) \ ~(path_component_of X x y)" by (force simp: disjnt_def path_component_of_eq_empty path_component_of_equiv) lemma path_component_of_eq: "path_component_of X x = path_component_of X y \ (x \ topspace X) \ (y \ topspace X) \ x \ topspace X \ y \ topspace X \ path_component_of X x y" by (metis Collect_empty_eq_bot path_component_of_eq_empty path_component_of_equiv) lemma path_component_of_aux: "path_component_of X x y \ path_component_of (subtopology X (Collect (path_component_of X x))) x y" by (meson path_component_of path_component_of_maximal path_connectedin_subtopology) lemma path_connectedin_path_component_of: "path_connectedin X (Collect (path_component_of X x))" proof - have "topspace (subtopology X (path_component_of_set X x)) = path_component_of_set X x" by (meson path_component_of_subset_topspace topspace_subtopology_subset) then have "path_connected_space (subtopology X (path_component_of_set X x))" by (metis mem_Collect_eq path_component_of_aux path_component_of_equiv path_connected_space_iff_path_component) then show ?thesis by (simp add: path_component_of_subset_topspace path_connectedin_def) qed lemma path_connectedin_euclidean [simp]: "path_connectedin euclidean S \ path_connected S" by (auto simp: path_connectedin_def path_connected_space_iff_path_component path_connected_component) lemma path_connected_space_euclidean_subtopology [simp]: "path_connected_space(subtopology euclidean S) \ path_connected S" using path_connectedin_topspace by force lemma Union_path_components_of: "\(path_components_of X) = topspace X" by (auto simp: path_components_of_def path_component_of_equiv) lemma path_components_of_maximal: "\C \ path_components_of X; path_connectedin X S; ~disjnt C S\ \ S \ C" by (smt (verit, ccfv_SIG) disjnt_iff imageE mem_Collect_eq path_component_of_equiv path_component_of_maximal path_components_of_def) lemma pairwise_disjoint_path_components_of: "pairwise disjnt (path_components_of X)" by (auto simp: path_components_of_def pairwise_def path_component_of_disjoint path_component_of_equiv) lemma complement_path_components_of_Union: "C \ path_components_of X \ topspace X - C = \(path_components_of X - {C})" by (metis Union_path_components_of bot.extremum ccpo_Sup_singleton diff_Union_pairwise_disjoint insert_subsetI pairwise_disjoint_path_components_of) lemma nonempty_path_components_of: assumes "C \ path_components_of X" shows "C \ {}" by (metis assms imageE path_component_of_eq_empty path_components_of_def) lemma path_components_of_subset: "C \ path_components_of X \ C \ topspace X" by (auto simp: path_components_of_def path_component_of_equiv) lemma path_connectedin_path_components_of: "C \ path_components_of X \ path_connectedin X C" by (auto simp: path_components_of_def path_connectedin_path_component_of) lemma path_component_in_path_components_of: "Collect (path_component_of X a) \ path_components_of X \ a \ topspace X" by (metis imageI nonempty_path_components_of path_component_of_eq_empty path_components_of_def) lemma path_connectedin_Union: assumes \: "\S. S \ \ \ path_connectedin X S" and "\\ \ {}" shows "path_connectedin X (\\)" proof - obtain a where "\S. S \ \ \ a \ S" using assms by blast then have "\x. x \ topspace (subtopology X (\\)) \ path_component_of (subtopology X (\\)) a x" unfolding topspace_subtopology path_component_of by (metis (full_types) IntD2 Union_iff Union_upper \ path_connectedin_subtopology) then show ?thesis using \ unfolding path_connectedin_def by (metis Sup_le_iff path_component_of_equiv path_connected_space_iff_path_component) qed lemma path_connectedin_Un: "\path_connectedin X S; path_connectedin X T; S \ T \ {}\ \ path_connectedin X (S \ T)" by (blast intro: path_connectedin_Union [of "{S,T}", simplified]) lemma path_connected_space_iff_components_eq: "path_connected_space X \ (\C \ path_components_of X. \C' \ path_components_of X. C = C')" unfolding path_components_of_def proof (intro iffI ballI) assume "\C \ path_component_of_set X ` topspace X. \C' \ path_component_of_set X ` topspace X. C = C'" then show "path_connected_space X" using path_component_of_refl path_connected_space_iff_path_component by fastforce qed (auto simp: path_connected_space_path_component_set) lemma path_components_of_eq_empty: "path_components_of X = {} \ X = trivial_topology" by (metis image_is_empty path_components_of_def subtopology_eq_discrete_topology_empty) lemma path_components_of_empty_space: "path_components_of trivial_topology = {}" by (simp add: path_components_of_eq_empty) lemma path_components_of_subset_singleton: "path_components_of X \ {S} \ path_connected_space X \ (topspace X = {} \ topspace X = S)" proof (cases "topspace X = {}") case True then show ?thesis by (auto simp: path_components_of_empty_space path_connected_space_topspace_empty) next case False have "(path_components_of X = {S}) \ (path_connected_space X \ topspace X = S)" by (metis False Set.set_insert ex_in_conv insert_iff path_component_in_path_components_of path_connected_space_iff_components_eq path_connected_space_path_component_set) with False show ?thesis by (simp add: path_components_of_eq_empty subset_singleton_iff) qed lemma path_connected_space_iff_components_subset_singleton: "path_connected_space X \ (\a. path_components_of X \ {a})" by (simp add: path_components_of_subset_singleton) lemma path_components_of_eq_singleton: "path_components_of X = {S} \ path_connected_space X \ topspace X \ {} \ S = topspace X" by (metis cSup_singleton insert_not_empty path_components_of_subset_singleton subset_singleton_iff) lemma path_components_of_path_connected_space: "path_connected_space X \ path_components_of X = (if topspace X = {} then {} else {topspace X})" by (simp add: path_components_of_eq_empty path_components_of_eq_singleton) lemma path_component_subset_connected_component_of: "path_component_of_set X x \ connected_component_of_set X x" proof (cases "x \ topspace X") case True then show ?thesis by (simp add: connected_component_of_maximal path_component_of_refl path_connectedin_imp_connectedin path_connectedin_path_component_of) next case False then show ?thesis using path_component_of_eq_empty by fastforce qed lemma exists_path_component_of_superset: assumes S: "path_connectedin X S" and ne: "topspace X \ {}" obtains C where "C \ path_components_of X" "S \ C" by (metis S ne ex_in_conv path_component_in_path_components_of path_component_of_maximal path_component_of_subset_topspace subset_eq that) lemma path_component_of_eq_overlap: "path_component_of X x = path_component_of X y \ (x \ topspace X) \ (y \ topspace X) \ Collect (path_component_of X x) \ Collect (path_component_of X y) \ {}" by (metis disjnt_def empty_iff inf_bot_right mem_Collect_eq path_component_of_disjoint path_component_of_eq path_component_of_eq_empty) lemma path_component_of_nonoverlap: "Collect (path_component_of X x) \ Collect (path_component_of X y) = {} \ (x \ topspace X) \ (y \ topspace X) \ path_component_of X x \ path_component_of X y" by (metis inf.idem path_component_of_eq_empty path_component_of_eq_overlap) lemma path_component_of_overlap: "Collect (path_component_of X x) \ Collect (path_component_of X y) \ {} \ x \ topspace X \ y \ topspace X \ path_component_of X x = path_component_of X y" by (meson path_component_of_nonoverlap) lemma path_components_of_disjoint: "\C \ path_components_of X; C' \ path_components_of X\ \ disjnt C C' \ C \ C'" by (auto simp: path_components_of_def path_component_of_disjoint path_component_of_equiv) lemma path_components_of_overlap: "\C \ path_components_of X; C' \ path_components_of X\ \ C \ C' \ {} \ C = C'" by (auto simp: path_components_of_def path_component_of_equiv) lemma path_component_of_unique: "\x \ C; path_connectedin X C; \C'. \x \ C'; path_connectedin X C'\ \ C' \ C\ \ Collect (path_component_of X x) = C" by (meson subsetD eq_iff path_component_of_maximal path_connectedin_path_component_of) lemma path_component_of_discrete_topology [simp]: "Collect (path_component_of (discrete_topology U) x) = (if x \ U then {x} else {})" proof - have "\C'. \x \ C'; path_connectedin (discrete_topology U) C'\ \ C' \ {x}" by (metis path_connectedin_discrete_topology subsetD singletonD) then have "x \ U \ Collect (path_component_of (discrete_topology U) x) = {x}" by (simp add: path_component_of_unique) then show ?thesis using path_component_in_topspace by fastforce qed lemma path_component_of_discrete_topology_iff [simp]: "path_component_of (discrete_topology U) x y \ x \ U \ y=x" by (metis empty_iff insertI1 mem_Collect_eq path_component_of_discrete_topology singletonD) lemma path_components_of_discrete_topology [simp]: "path_components_of (discrete_topology U) = (\x. {x}) ` U" by (auto simp: path_components_of_def image_def fun_eq_iff) lemma homeomorphic_map_path_component_of: assumes f: "homeomorphic_map X Y f" and x: "x \ topspace X" shows "Collect (path_component_of Y (f x)) = f ` Collect(path_component_of X x)" proof - obtain g where g: "homeomorphic_maps X Y f g" using f homeomorphic_map_maps by blast show ?thesis proof have "Collect (path_component_of Y (f x)) \ topspace Y" by (simp add: path_component_of_subset_topspace) moreover have "g ` Collect(path_component_of Y (f x)) \ Collect (path_component_of X (g (f x)))" using f g x unfolding homeomorphic_maps_def by (metis image_Collect_subsetI image_eqI mem_Collect_eq path_component_of_equiv path_component_of_maximal path_connectedin_continuous_map_image path_connectedin_path_component_of) ultimately show "Collect (path_component_of Y (f x)) \ f ` Collect (path_component_of X x)" using g x unfolding homeomorphic_maps_def continuous_map_def image_iff subset_iff by metis show "f ` Collect (path_component_of X x) \ Collect (path_component_of Y (f x))" proof (rule path_component_of_maximal) show "path_connectedin Y (f ` Collect (path_component_of X x))" by (meson f homeomorphic_map_path_connectedness_eq path_connectedin_path_component_of) qed (simp add: path_component_of_refl x) qed qed lemma homeomorphic_map_path_components_of: assumes "homeomorphic_map X Y f" shows "path_components_of Y = (image f) ` (path_components_of X)" unfolding path_components_of_def homeomorphic_imp_surjective_map [OF assms, symmetric] using assms homeomorphic_map_path_component_of by fastforce subsection\Paths and path-connectedness\ lemma path_connected_space_quotient_map_image: "\quotient_map X Y q; path_connected_space X\ \ path_connected_space Y" by (metis path_connectedin_continuous_map_image path_connectedin_topspace quotient_imp_continuous_map quotient_imp_surjective_map) lemma path_connected_space_retraction_map_image: "\retraction_map X Y r; path_connected_space X\ \ path_connected_space Y" using path_connected_space_quotient_map_image retraction_imp_quotient_map by blast lemma path_connected_space_prod_topology: "path_connected_space(prod_topology X Y) \ topspace(prod_topology X Y) = {} \ path_connected_space X \ path_connected_space Y" proof (cases "topspace(prod_topology X Y) = {}") case True then show ?thesis using path_connected_space_topspace_empty by force next case False have "path_connected_space (prod_topology X Y)" if X: "path_connected_space X" and Y: "path_connected_space Y" proof (clarsimp simp: path_connected_space_def) fix x y x' y' assume "x \ topspace X" and "y \ topspace Y" and "x' \ topspace X" and "y' \ topspace Y" obtain f where "pathin X f" "f 0 = x" "f 1 = x'" by (meson X \x \ topspace X\ \x' \ topspace X\ path_connected_space_def) obtain g where "pathin Y g" "g 0 = y" "g 1 = y'" by (meson Y \y \ topspace Y\ \y' \ topspace Y\ path_connected_space_def) show "\h. pathin (prod_topology X Y) h \ h 0 = (x,y) \ h 1 = (x',y')" proof (intro exI conjI) show "pathin (prod_topology X Y) (\t. (f t, g t))" using \pathin X f\ \pathin Y g\ by (simp add: continuous_map_paired pathin_def) show "(\t. (f t, g t)) 0 = (x, y)" using \f 0 = x\ \g 0 = y\ by blast show "(\t. (f t, g t)) 1 = (x', y')" using \f 1 = x'\ \g 1 = y'\ by blast qed qed then show ?thesis by (metis False path_connected_space_quotient_map_image prod_topology_trivial1 prod_topology_trivial2 quotient_map_fst quotient_map_snd topspace_discrete_topology) qed lemma path_connectedin_Times: "path_connectedin (prod_topology X Y) (S \ T) \ S = {} \ T = {} \ path_connectedin X S \ path_connectedin Y T" by (auto simp add: path_connectedin_def subtopology_Times path_connected_space_prod_topology) subsection\Path components\ lemma path_component_of_subtopology_eq: "path_component_of (subtopology X U) x = path_component_of X x \ path_component_of_set X x \ U" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (metis path_connectedin_path_component_of path_connectedin_subtopology) next show "?rhs \ ?lhs" unfolding fun_eq_iff by (metis path_connectedin_subtopology path_component_of path_component_of_aux path_component_of_mono) qed lemma path_components_of_subtopology: assumes "C \ path_components_of X" "C \ U" shows "C \ path_components_of (subtopology X U)" using assms path_component_of_refl path_component_of_subtopology_eq topspace_subtopology by (smt (verit) imageE path_component_in_path_components_of path_components_of_def) lemma path_imp_connected_component_of: "path_component_of X x y \ connected_component_of X x y" by (metis in_mono mem_Collect_eq path_component_subset_connected_component_of) lemma finite_path_components_of_finite: "finite(topspace X) \ finite(path_components_of X)" by (simp add: Union_path_components_of finite_UnionD) lemma path_component_of_continuous_image: "\continuous_map X X' f; path_component_of X x y\ \ path_component_of X' (f x) (f y)" by (meson image_eqI path_component_of path_connectedin_continuous_map_image) lemma path_component_of_pair [simp]: "path_component_of_set (prod_topology X Y) (x,y) = path_component_of_set X x \ path_component_of_set Y y" (is "?lhs = ?rhs") proof (cases "?lhs = {}") case True then show ?thesis by (metis Sigma_empty1 Sigma_empty2 mem_Sigma_iff path_component_of_eq_empty topspace_prod_topology) next case False then have "path_component_of X x x" "path_component_of Y y y" using path_component_of_eq_empty path_component_of_refl by fastforce+ moreover have "path_connectedin (prod_topology X Y) (path_component_of_set X x \ path_component_of_set Y y)" by (metis path_connectedin_Times path_connectedin_path_component_of) moreover have "path_component_of X x a" "path_component_of Y y b" if "(x, y) \ C'" "(a,b) \ C'" and "path_connectedin (prod_topology X Y) C'" for C' a b by (smt (verit, best) that continuous_map_fst continuous_map_snd fst_conv snd_conv path_component_of path_component_of_continuous_image)+ ultimately show ?thesis by (intro path_component_of_unique) auto qed lemma path_components_of_prod_topology: "path_components_of (prod_topology X Y) = (\(C,D). C \ D) ` (path_components_of X \ path_components_of Y)" by (force simp add: image_iff path_components_of_def) lemma path_components_of_prod_topology': "path_components_of (prod_topology X Y) = {C \ D |C D. C \ path_components_of X \ D \ path_components_of Y}" by (auto simp: path_components_of_prod_topology) lemma path_component_of_product_topology: "path_component_of_set (product_topology X I) f = (if f \ extensional I then PiE I (\i. path_component_of_set (X i) (f i)) else {})" (is "?lhs = ?rhs") proof (cases "path_component_of_set (product_topology X I) f = {}") case True then show ?thesis by (smt (verit) PiE_eq_empty_iff PiE_iff path_component_of_eq_empty topspace_product_topology) next case False then have [simp]: "f \ extensional I" by (auto simp: path_component_of_eq_empty PiE_iff path_component_of_equiv) show ?thesis proof (intro path_component_of_unique) show "f \ ?rhs" using False path_component_of_eq_empty path_component_of_refl by force show "path_connectedin (product_topology X I) (if f \ extensional I then \\<^sub>E i\I. path_component_of_set (X i) (f i) else {})" by (simp add: path_connectedin_PiE path_connectedin_path_component_of) fix C' assume "f \ C'" and C': "path_connectedin (product_topology X I) C'" show "C' \ ?rhs" proof - have "C' \ extensional I" using PiE_def C' path_connectedin_subset_topspace by fastforce with \f \ C'\ C' show ?thesis apply (clarsimp simp: PiE_iff subset_iff) by (smt (verit, ccfv_threshold) continuous_map_product_projection path_component_of path_component_of_continuous_image) qed qed qed lemma path_components_of_product_topology: "path_components_of (product_topology X I) = {PiE I B |B. \i \ I. B i \ path_components_of(X i)}" (is "?lhs=?rhs") proof show "?lhs \ ?rhs" unfolding path_components_of_def image_subset_iff by (smt (verit) image_iff mem_Collect_eq path_component_of_product_topology topspace_product_topology_alt) next show "?rhs \ ?lhs" proof fix F assume "F \ ?rhs" then obtain B where B: "F = Pi\<^sub>E I B" and "\i\I. \x\topspace (X i). B i = path_component_of_set (X i) x" by (force simp add: path_components_of_def image_iff) then obtain f where ftop: "\i. i \ I \ f i \ topspace (X i)" and BF: "\i. i \ I \ B i = path_component_of_set (X i) (f i)" by metis then have "F = path_component_of_set (product_topology X I) (restrict f I)" by (metis (mono_tags, lifting) B PiE_cong path_component_of_product_topology restrict_apply' restrict_extensional) then show "F \ ?lhs" by (simp add: ftop path_component_in_path_components_of) qed qed subsection \Sphere is path-connected\ lemma path_connected_punctured_universe: assumes "2 \ DIM('a::euclidean_space)" shows "path_connected (- {a::'a})" proof - let ?A = "{x::'a. \i\Basis. x \ i < a \ i}" let ?B = "{x::'a. \i\Basis. a \ i < x \ i}" have A: "path_connected ?A" unfolding Collect_bex_eq proof (rule path_connected_UNION) fix i :: 'a assume "i \ Basis" then show "(\i\Basis. (a \ i - 1)*\<^sub>R i) \ {x::'a. x \ i < a \ i}" by simp show "path_connected {x. x \ i < a \ i}" using convex_imp_path_connected [OF convex_halfspace_lt, of i "a \ i"] by (simp add: inner_commute) qed have B: "path_connected ?B" unfolding Collect_bex_eq proof (rule path_connected_UNION) fix i :: 'a assume "i \ Basis" then show "(\i\Basis. (a \ i + 1) *\<^sub>R i) \ {x::'a. a \ i < x \ i}" by simp show "path_connected {x. a \ i < x \ i}" using convex_imp_path_connected [OF convex_halfspace_gt, of "a \ i" i] by (simp add: inner_commute) qed obtain S :: "'a set" where "S \ Basis" and "card S = Suc (Suc 0)" using obtain_subset_with_card_n[OF assms] by (force simp add: eval_nat_numeral) then obtain b0 b1 :: 'a where "b0 \ Basis" and "b1 \ Basis" and "b0 \ b1" unfolding card_Suc_eq by auto then have "a + b0 - b1 \ ?A \ ?B" by (auto simp: inner_simps inner_Basis) then have "?A \ ?B \ {}" by fast with A B have "path_connected (?A \ ?B)" by (rule path_connected_Un) also have "?A \ ?B = {x. \i\Basis. x \ i \ a \ i}" unfolding neq_iff bex_disj_distrib Collect_disj_eq .. also have "\ = {x. x \ a}" unfolding euclidean_eq_iff [where 'a='a] by (simp add: Bex_def) also have "\ = - {a}" by auto finally show ?thesis . qed corollary connected_punctured_universe: "2 \ DIM('N::euclidean_space) \ connected(- {a::'N})" by (simp add: path_connected_punctured_universe path_connected_imp_connected) proposition path_connected_sphere: fixes a :: "'a :: euclidean_space" assumes "2 \ DIM('a)" shows "path_connected(sphere a r)" proof (cases r "0::real" rule: linorder_cases) case greater then have eq: "(sphere (0::'a) r) = (\x. (r / norm x) *\<^sub>R x) ` (- {0::'a})" by (force simp: image_iff split: if_split_asm) have "continuous_on (- {0::'a}) (\x. (r / norm x) *\<^sub>R x)" by (intro continuous_intros) auto then have "path_connected ((\x. (r / norm x) *\<^sub>R x) ` (- {0::'a}))" by (intro path_connected_continuous_image path_connected_punctured_universe assms) with eq have "path_connected((+) a ` (sphere (0::'a) r))" by (simp add: path_connected_translation) then show ?thesis by (metis add.right_neutral sphere_translation) qed auto lemma connected_sphere: fixes a :: "'a :: euclidean_space" assumes "2 \ DIM('a)" shows "connected(sphere a r)" using path_connected_sphere [OF assms] by (simp add: path_connected_imp_connected) corollary path_connected_complement_bounded_convex: fixes S :: "'a :: euclidean_space set" assumes "bounded S" "convex S" and 2: "2 \ DIM('a)" shows "path_connected (- S)" proof (cases "S = {}") case True then show ?thesis using convex_imp_path_connected by auto next case False then obtain a where "a \ S" by auto have \
[rule_format]: "\y\S. \u. 0 \ u \ u \ 1 \ (1 - u) *\<^sub>R a + u *\<^sub>R y \ S" using \convex S\ \a \ S\ by (simp add: convex_alt) { fix x y assume "x \ S" "y \ S" then have "x \ a" "y \ a" using \a \ S\ by auto then have bxy: "bounded(insert x (insert y S))" by (simp add: \bounded S\) then obtain B::real where B: "0 < B" and Bx: "norm (a - x) < B" and By: "norm (a - y) < B" and "S \ ball a B" using bounded_subset_ballD [OF bxy, of a] by (auto simp: dist_norm) define C where "C = B / norm(x - a)" let ?Cxa = "a + C *\<^sub>R (x - a)" { fix u assume u: "(1 - u) *\<^sub>R x + u *\<^sub>R ?Cxa \ S" and "0 \ u" "u \ 1" have CC: "1 \ 1 + (C - 1) * u" using \x \ a\ \0 \ u\ Bx by (auto simp add: C_def norm_minus_commute) have *: "\v. (1 - u) *\<^sub>R x + u *\<^sub>R (a + v *\<^sub>R (x - a)) = a + (1 + (v - 1) * u) *\<^sub>R (x - a)" by (simp add: algebra_simps) have "a + ((1 / (1 + C * u - u)) *\<^sub>R x + ((u / (1 + C * u - u)) *\<^sub>R a + (C * u / (1 + C * u - u)) *\<^sub>R x)) = (1 + (u / (1 + C * u - u))) *\<^sub>R a + ((1 / (1 + C * u - u)) + (C * u / (1 + C * u - u))) *\<^sub>R x" by (simp add: algebra_simps) also have "\ = (1 + (u / (1 + C * u - u))) *\<^sub>R a + (1 + (u / (1 + C * u - u))) *\<^sub>R x" using CC by (simp add: field_simps) also have "\ = x + (1 + (u / (1 + C * u - u))) *\<^sub>R a + (u / (1 + C * u - u)) *\<^sub>R x" by (simp add: algebra_simps) also have "\ = x + ((1 / (1 + C * u - u)) *\<^sub>R a + ((u / (1 + C * u - u)) *\<^sub>R x + (C * u / (1 + C * u - u)) *\<^sub>R a))" using CC by (simp add: field_simps) (simp add: add_divide_distrib scaleR_add_left) finally have xeq: "(1 - 1 / (1 + (C - 1) * u)) *\<^sub>R a + (1 / (1 + (C - 1) * u)) *\<^sub>R (a + (1 + (C - 1) * u) *\<^sub>R (x - a)) = x" by (simp add: algebra_simps) have False using \
[of "a + (1 + (C - 1) * u) *\<^sub>R (x - a)" "1 / (1 + (C - 1) * u)"] using u \x \ a\ \x \ S\ \0 \ u\ CC by (auto simp: xeq *) } then have pcx: "path_component (- S) x ?Cxa" by (force simp: closed_segment_def intro!: path_component_linepath) define D where "D = B / norm(y - a)" \ \massive duplication with the proof above\ let ?Dya = "a + D *\<^sub>R (y - a)" { fix u assume u: "(1 - u) *\<^sub>R y + u *\<^sub>R ?Dya \ S" and "0 \ u" "u \ 1" have DD: "1 \ 1 + (D - 1) * u" using \y \ a\ \0 \ u\ By by (auto simp add: D_def norm_minus_commute) have *: "\v. (1 - u) *\<^sub>R y + u *\<^sub>R (a + v *\<^sub>R (y - a)) = a + (1 + (v - 1) * u) *\<^sub>R (y - a)" by (simp add: algebra_simps) have "a + ((1 / (1 + D * u - u)) *\<^sub>R y + ((u / (1 + D * u - u)) *\<^sub>R a + (D * u / (1 + D * u - u)) *\<^sub>R y)) = (1 + (u / (1 + D * u - u))) *\<^sub>R a + ((1 / (1 + D * u - u)) + (D * u / (1 + D * u - u))) *\<^sub>R y" by (simp add: algebra_simps) also have "\ = (1 + (u / (1 + D * u - u))) *\<^sub>R a + (1 + (u / (1 + D * u - u))) *\<^sub>R y" using DD by (simp add: field_simps) also have "\ = y + (1 + (u / (1 + D * u - u))) *\<^sub>R a + (u / (1 + D * u - u)) *\<^sub>R y" by (simp add: algebra_simps) also have "\ = y + ((1 / (1 + D * u - u)) *\<^sub>R a + ((u / (1 + D * u - u)) *\<^sub>R y + (D * u / (1 + D * u - u)) *\<^sub>R a))" using DD by (simp add: field_simps) (simp add: add_divide_distrib scaleR_add_left) finally have xeq: "(1 - 1 / (1 + (D - 1) * u)) *\<^sub>R a + (1 / (1 + (D - 1) * u)) *\<^sub>R (a + (1 + (D - 1) * u) *\<^sub>R (y - a)) = y" by (simp add: algebra_simps) have False using \
[of "a + (1 + (D - 1) * u) *\<^sub>R (y - a)" "1 / (1 + (D - 1) * u)"] using u \y \ a\ \y \ S\ \0 \ u\ DD by (auto simp: xeq *) } then have pdy: "path_component (- S) y ?Dya" by (force simp: closed_segment_def intro!: path_component_linepath) have pyx: "path_component (- S) ?Dya ?Cxa" proof (rule path_component_of_subset) show "sphere a B \ - S" using \S \ ball a B\ by (force simp: ball_def dist_norm norm_minus_commute) have aB: "?Dya \ sphere a B" "?Cxa \ sphere a B" using \x \ a\ using \y \ a\ B by (auto simp: dist_norm C_def D_def) then show "path_component (sphere a B) ?Dya ?Cxa" using path_connected_sphere [OF 2] path_connected_component by blast qed have "path_component (- S) x y" by (metis path_component_trans path_component_sym pcx pdy pyx) } then show ?thesis by (auto simp: path_connected_component) qed lemma connected_complement_bounded_convex: fixes S :: "'a :: euclidean_space set" assumes "bounded S" "convex S" "2 \ DIM('a)" shows "connected (- S)" using path_connected_complement_bounded_convex [OF assms] path_connected_imp_connected by blast lemma connected_diff_ball: fixes S :: "'a :: euclidean_space set" assumes "connected S" "cball a r \ S" "2 \ DIM('a)" shows "connected (S - ball a r)" proof (rule connected_diff_open_from_closed [OF ball_subset_cball]) show "connected (cball a r - ball a r)" using assms connected_sphere by (auto simp: cball_diff_eq_sphere) qed (auto simp: assms dist_norm) proposition connected_open_delete: assumes "open S" "connected S" and 2: "2 \ DIM('N::euclidean_space)" shows "connected(S - {a::'N})" proof (cases "a \ S") case True with \open S\ obtain \ where "\ > 0" and \: "cball a \ \ S" using open_contains_cball_eq by blast define b where "b \ a + \ *\<^sub>R (SOME i. i \ Basis)" have "dist a b = \" by (simp add: b_def dist_norm SOME_Basis \0 < \\ less_imp_le) with \ have "b \ \{S - ball a r |r. 0 < r \ r < \}" by auto then have nonemp: "(\{S - ball a r |r. 0 < r \ r < \}) = {} \ False" by auto have con: "\r. r < \ \ connected (S - ball a r)" using \ by (force intro: connected_diff_ball [OF \connected S\ _ 2]) have "x \ \{S - ball a r |r. 0 < r \ r < \}" if "x \ S - {a}" for x using that \0 < \\ by (intro UnionI [of "S - ball a (min \ (dist a x) / 2)"]) auto then have "S - {a} = \{S - ball a r | r. 0 < r \ r < \}" by auto then show ?thesis by (auto intro: connected_Union con dest!: nonemp) next case False then show ?thesis by (simp add: \connected S\) qed corollary path_connected_open_delete: assumes "open S" "connected S" and 2: "2 \ DIM('N::euclidean_space)" shows "path_connected(S - {a::'N})" by (simp add: assms connected_open_delete connected_open_path_connected open_delete) corollary path_connected_punctured_ball: "2 \ DIM('N::euclidean_space) \ path_connected(ball a r - {a::'N})" by (simp add: path_connected_open_delete) corollary connected_punctured_ball: "2 \ DIM('N::euclidean_space) \ connected(ball a r - {a::'N})" by (simp add: connected_open_delete) corollary connected_open_delete_finite: fixes S T::"'a::euclidean_space set" assumes S: "open S" "connected S" and 2: "2 \ DIM('a)" and "finite T" shows "connected(S - T)" using \finite T\ S proof (induct T) case empty show ?case using \connected S\ by simp next case (insert x T) then have "connected (S-T)" by auto moreover have "open (S - T)" using finite_imp_closed[OF \finite T\] \open S\ by auto ultimately have "connected (S - T - {x})" using connected_open_delete[OF _ _ 2] by auto thus ?case by (metis Diff_insert) qed lemma sphere_1D_doubleton_zero: assumes 1: "DIM('a) = 1" and "r > 0" obtains x y::"'a::euclidean_space" where "sphere 0 r = {x,y} \ dist x y = 2*r" proof - obtain b::'a where b: "Basis = {b}" using 1 card_1_singletonE by blast show ?thesis proof (intro that conjI) have "x = norm x *\<^sub>R b \ x = - norm x *\<^sub>R b" if "r = norm x" for x proof - have xb: "(x \ b) *\<^sub>R b = x" using euclidean_representation [of x, unfolded b] by force then have "norm ((x \ b) *\<^sub>R b) = norm x" by simp with b have "\x \ b\ = norm x" using norm_Basis by (simp add: b) with xb show ?thesis by (metis (mono_tags, opaque_lifting) abs_eq_iff abs_norm_cancel) qed with \r > 0\ b show "sphere 0 r = {r *\<^sub>R b, - r *\<^sub>R b}" by (force simp: sphere_def dist_norm) have "dist (r *\<^sub>R b) (- r *\<^sub>R b) = norm (r *\<^sub>R b + r *\<^sub>R b)" by (simp add: dist_norm) also have "\ = norm ((2*r) *\<^sub>R b)" by (metis mult_2 scaleR_add_left) also have "\ = 2*r" using \r > 0\ b norm_Basis by fastforce finally show "dist (r *\<^sub>R b) (- r *\<^sub>R b) = 2*r" . qed qed lemma sphere_1D_doubleton: fixes a :: "'a :: euclidean_space" assumes "DIM('a) = 1" and "r > 0" obtains x y where "sphere a r = {x,y} \ dist x y = 2*r" using sphere_1D_doubleton_zero [OF assms] dist_add_cancel image_empty image_insert by (metis (no_types, opaque_lifting) add.right_neutral sphere_translation) lemma psubset_sphere_Compl_connected: fixes S :: "'a::euclidean_space set" assumes S: "S \ sphere a r" and "0 < r" and 2: "2 \ DIM('a)" shows "connected(- S)" proof - have "S \ sphere a r" using S by blast obtain b where "dist a b = r" and "b \ S" using S mem_sphere by blast have CS: "- S = {x. dist a x \ r \ (x \ S)} \ {x. r \ dist a x \ (x \ S)}" by auto have "{x. dist a x \ r \ x \ S} \ {x. r \ dist a x \ x \ S} \ {}" using \b \ S\ \dist a b = r\ by blast moreover have "connected {x. dist a x \ r \ x \ S}" using assms by (force intro: connected_intermediate_closure [of "ball a r"]) moreover have "connected {x. r \ dist a x \ x \ S}" proof (rule connected_intermediate_closure [of "- cball a r"]) show "{x. r \ dist a x \ x \ S} \ closure (- cball a r)" using interior_closure by (force intro: connected_complement_bounded_convex) qed (use assms connected_complement_bounded_convex in auto) ultimately show ?thesis by (simp add: CS connected_Un) qed subsection\Every annulus is a connected set\ lemma path_connected_2DIM_I: fixes a :: "'N::euclidean_space" assumes 2: "2 \ DIM('N)" and pc: "path_connected {r. 0 \ r \ P r}" shows "path_connected {x. P(norm(x - a))}" proof - have "{x. P(norm(x - a))} = (+) a ` {x. P(norm x)}" by force moreover have "path_connected {x::'N. P(norm x)}" proof - let ?D = "{x. 0 \ x \ P x} \ sphere (0::'N) 1" have "x \ (\z. fst z *\<^sub>R snd z) ` ?D" if "P (norm x)" for x::'N proof (cases "x=0") case True with that show ?thesis apply (simp add: image_iff) by (metis (no_types) mem_sphere_0 order_refl vector_choose_size zero_le_one) next case False with that show ?thesis by (rule_tac x="(norm x, x /\<^sub>R norm x)" in image_eqI) auto qed then have *: "{x::'N. P(norm x)} = (\z. fst z *\<^sub>R snd z) ` ?D" by auto have "continuous_on ?D (\z:: real\'N. fst z *\<^sub>R snd z)" by (intro continuous_intros) moreover have "path_connected ?D" by (metis path_connected_Times [OF pc] path_connected_sphere 2) ultimately show ?thesis by (simp add: "*" path_connected_continuous_image) qed ultimately show ?thesis using path_connected_translation by metis qed proposition path_connected_annulus: fixes a :: "'N::euclidean_space" assumes "2 \ DIM('N)" shows "path_connected {x. r1 < norm(x - a) \ norm(x - a) < r2}" "path_connected {x. r1 < norm(x - a) \ norm(x - a) \ r2}" "path_connected {x. r1 \ norm(x - a) \ norm(x - a) < r2}" "path_connected {x. r1 \ norm(x - a) \ norm(x - a) \ r2}" by (auto simp: is_interval_def intro!: is_interval_convex convex_imp_path_connected path_connected_2DIM_I [OF assms]) proposition connected_annulus: fixes a :: "'N::euclidean_space" assumes "2 \ DIM('N::euclidean_space)" shows "connected {x. r1 < norm(x - a) \ norm(x - a) < r2}" "connected {x. r1 < norm(x - a) \ norm(x - a) \ r2}" "connected {x. r1 \ norm(x - a) \ norm(x - a) < r2}" "connected {x. r1 \ norm(x - a) \ norm(x - a) \ r2}" by (auto simp: path_connected_annulus [OF assms] path_connected_imp_connected) subsection\<^marker>\tag unimportant\\Relations between components and path components\ lemma open_connected_component: fixes S :: "'a::real_normed_vector set" assumes "open S" shows "open (connected_component_set S x)" proof (clarsimp simp: open_contains_ball) fix y assume xy: "connected_component S x y" then obtain e where "e>0" "ball y e \ S" using assms connected_component_in openE by blast then show "\e>0. ball y e \ connected_component_set S x" by (metis xy centre_in_ball connected_ball connected_component_eq_eq connected_component_in connected_component_maximal) qed corollary open_components: fixes S :: "'a::real_normed_vector set" shows "\open u; S \ components u\ \ open S" by (simp add: components_iff) (metis open_connected_component) lemma in_closure_connected_component: fixes S :: "'a::real_normed_vector set" assumes x: "x \ S" and S: "open S" shows "x \ closure (connected_component_set S y) \ x \ connected_component_set S y" proof - have "x islimpt connected_component_set S y \ connected_component S y x" by (metis (no_types, lifting) S connected_component_eq connected_component_refl islimptE mem_Collect_eq open_connected_component x) then show ?thesis by (auto simp: closure_def) qed lemma connected_disjoint_Union_open_pick: assumes "pairwise disjnt B" "\S. S \ A \ connected S \ S \ {}" "\S. S \ B \ open S" "\A \ \B" "S \ A" obtains T where "T \ B" "S \ T" "S \ \(B - {T}) = {}" proof - have "S \ \B" "connected S" "S \ {}" using assms \S \ A\ by blast+ then obtain T where "T \ B" "S \ T \ {}" by (metis Sup_inf_eq_bot_iff inf.absorb_iff2 inf_commute) have 1: "open T" by (simp add: \T \ B\ assms) have 2: "open (\(B-{T}))" using assms by blast have 3: "S \ T \ \(B - {T})" using \S \ \B\ by blast have "T \ \(B - {T}) = {}" using \T \ B\ \pairwise disjnt B\ by (auto simp: pairwise_def disjnt_def) then have 4: "T \ \(B - {T}) \ S = {}" by auto from connectedD [OF \connected S\ 1 2 4 3] have "S \ \(B-{T}) = {}" by (auto simp: Int_commute \S \ T \ {}\) with \T \ B\ 3 that show ?thesis by (metis IntI UnE empty_iff subsetD subsetI) qed lemma connected_disjoint_Union_open_subset: assumes A: "pairwise disjnt A" and B: "pairwise disjnt B" and SA: "\S. S \ A \ open S \ connected S \ S \ {}" and SB: "\S. S \ B \ open S \ connected S \ S \ {}" and eq [simp]: "\A = \B" shows "A \ B" proof fix S assume "S \ A" obtain T where "T \ B" "S \ T" "S \ \(B - {T}) = {}" using SA SB \S \ A\ connected_disjoint_Union_open_pick [OF B, of A] eq order_refl by blast moreover obtain S' where "S' \ A" "T \ S'" "T \ \(A - {S'}) = {}" using SA SB \T \ B\ connected_disjoint_Union_open_pick [OF A, of B] eq order_refl by blast ultimately have "S' = S" by (metis A Int_subset_iff SA \S \ A\ disjnt_def inf.orderE pairwise_def) with \T \ S'\ have "T \ S" by simp with \S \ T\ have "S = T" by blast with \T \ B\ show "S \ B" by simp qed lemma connected_disjoint_Union_open_unique: assumes A: "pairwise disjnt A" and B: "pairwise disjnt B" and SA: "\S. S \ A \ open S \ connected S \ S \ {}" and SB: "\S. S \ B \ open S \ connected S \ S \ {}" and eq [simp]: "\A = \B" shows "A = B" by (metis subset_antisym connected_disjoint_Union_open_subset assms) proposition components_open_unique: fixes S :: "'a::real_normed_vector set" assumes "pairwise disjnt A" "\A = S" "\X. X \ A \ open X \ connected X \ X \ {}" shows "components S = A" proof - have "open S" using assms by blast show ?thesis proof (rule connected_disjoint_Union_open_unique) show "disjoint (components S)" by (simp add: components_eq disjnt_def pairwise_def) qed (use \open S\ in \simp_all add: assms open_components in_components_connected in_components_nonempty\) qed subsection\<^marker>\tag unimportant\\Existence of unbounded components\ lemma cobounded_unbounded_component: fixes S :: "'a :: euclidean_space set" assumes "bounded (-S)" shows "\x. x \ S \ \ bounded (connected_component_set S x)" proof - obtain i::'a where i: "i \ Basis" using nonempty_Basis by blast obtain B where B: "B>0" "-S \ ball 0 B" using bounded_subset_ballD [OF assms, of 0] by auto then have *: "\x. B \ norm x \ x \ S" by (force simp: ball_def dist_norm) have unbounded_inner: "\ bounded {x. inner i x \ B}" proof (clarsimp simp: bounded_def dist_norm) fix e x show "\y. B \ i \ y \ \ norm (x - y) \ e" using i by (rule_tac x="x + (max B e + 1 + \i \ x\) *\<^sub>R i" in exI) (auto simp: inner_right_distrib) qed have \
: "\x. B \ i \ x \ x \ S" using * Basis_le_norm [OF i] by (metis abs_ge_self inner_commute order_trans) have "{x. B \ i \ x} \ connected_component_set S (B *\<^sub>R i)" by (intro connected_component_maximal) (auto simp: i intro: convex_connected convex_halfspace_ge [of B] \
) then have "\ bounded (connected_component_set S (B *\<^sub>R i))" using bounded_subset unbounded_inner by blast moreover have "B *\<^sub>R i \ S" by (rule *) (simp add: norm_Basis [OF i]) ultimately show ?thesis by blast qed lemma cobounded_unique_unbounded_component: fixes S :: "'a :: euclidean_space set" assumes bs: "bounded (-S)" and "2 \ DIM('a)" and bo: "\ bounded(connected_component_set S x)" "\ bounded(connected_component_set S y)" shows "connected_component_set S x = connected_component_set S y" proof - obtain i::'a where i: "i \ Basis" using nonempty_Basis by blast obtain B where "B>0" and B: "-S \ ball 0 B" using bounded_subset_ballD [OF bs, of 0] by auto then have *: "\x. B \ norm x \ x \ S" by (force simp: ball_def dist_norm) obtain x' y' where x': "connected_component S x x'" "norm x' > B" and y': "connected_component S y y'" "norm y' > B" using \B>0\ bo bounded_pos by (metis linorder_not_le mem_Collect_eq) have x'y': "connected_component S x' y'" unfolding connected_component_def proof (intro exI conjI) show "connected (- ball 0 B :: 'a set)" using assms by (auto intro: connected_complement_bounded_convex) qed (use x' y' dist_norm * in auto) show ?thesis using x' y' x'y' by (metis connected_component_eq mem_Collect_eq) qed lemma cobounded_unbounded_components: fixes S :: "'a :: euclidean_space set" shows "bounded (-S) \ \c. c \ components S \ \bounded c" by (metis cobounded_unbounded_component components_def imageI) lemma cobounded_unique_unbounded_components: fixes S :: "'a :: euclidean_space set" shows "\bounded (- S); c \ components S; \ bounded c; c' \ components S; \ bounded c'; 2 \ DIM('a)\ \ c' = c" unfolding components_iff by (metis cobounded_unique_unbounded_component) lemma cobounded_has_bounded_component: fixes S :: "'a :: euclidean_space set" assumes "bounded (- S)" "\ connected S" "2 \ DIM('a)" obtains C where "C \ components S" "bounded C" by (meson cobounded_unique_unbounded_components connected_eq_connected_components_eq assms) subsection\The \inside\ and \outside\ of a Set\ text\<^marker>\tag important\\The inside comprises the points in a bounded connected component of the set's complement. The outside comprises the points in unbounded connected component of the complement.\ definition\<^marker>\tag important\ inside where "inside S \ {x. (x \ S) \ bounded(connected_component_set ( - S) x)}" definition\<^marker>\tag important\ outside where "outside S \ -S \ {x. \ bounded(connected_component_set (- S) x)}" lemma outside: "outside S = {x. \ bounded(connected_component_set (- S) x)}" by (auto simp: outside_def) (metis Compl_iff bounded_empty connected_component_eq_empty) lemma inside_no_overlap [simp]: "inside S \ S = {}" by (auto simp: inside_def) lemma outside_no_overlap [simp]: "outside S \ S = {}" by (auto simp: outside_def) lemma inside_Int_outside [simp]: "inside S \ outside S = {}" by (auto simp: inside_def outside_def) lemma inside_Un_outside [simp]: "inside S \ outside S = (- S)" by (auto simp: inside_def outside_def) lemma inside_eq_outside: "inside S = outside S \ S = UNIV" by (auto simp: inside_def outside_def) lemma inside_outside: "inside S = (- (S \ outside S))" by (force simp: inside_def outside) lemma outside_inside: "outside S = (- (S \ inside S))" by (auto simp: inside_outside) (metis IntI equals0D outside_no_overlap) lemma union_with_inside: "S \ inside S = - outside S" by (auto simp: inside_outside) (simp add: outside_inside) lemma union_with_outside: "S \ outside S = - inside S" by (simp add: inside_outside) lemma outside_mono: "S \ T \ outside T \ outside S" by (auto simp: outside bounded_subset connected_component_mono) lemma inside_mono: "S \ T \ inside S - T \ inside T" by (auto simp: inside_def bounded_subset connected_component_mono) lemma segment_bound_lemma: fixes u::real assumes "x \ B" "y \ B" "0 \ u" "u \ 1" shows "(1 - u) * x + u * y \ B" by (smt (verit) assms convex_bound_le ge_iff_diff_ge_0 minus_add_distrib mult_minus_right neg_le_iff_le) lemma cobounded_outside: fixes S :: "'a :: real_normed_vector set" assumes "bounded S" shows "bounded (- outside S)" proof - obtain B where B: "B>0" "S \ ball 0 B" using bounded_subset_ballD [OF assms, of 0] by auto { fix x::'a and C::real assume Bno: "B \ norm x" and C: "0 < C" have "\y. connected_component (- S) x y \ norm y > C" proof (cases "x = 0") case True with B Bno show ?thesis by force next case False have "closed_segment x (((B + C) / norm x) *\<^sub>R x) \ - ball 0 B" proof fix w assume "w \ closed_segment x (((B + C) / norm x) *\<^sub>R x)" then obtain u where w: "w = (1 - u + u * (B + C) / norm x) *\<^sub>R x" "0 \ u" "u \ 1" by (auto simp add: closed_segment_def real_vector_class.scaleR_add_left [symmetric]) with False B C have "B \ (1 - u) * norm x + u * (B + C)" using segment_bound_lemma [of B "norm x" "B + C" u] Bno by simp with False B C show "w \ - ball 0 B" using distrib_right [of _ _ "norm x"] by (simp add: ball_def w not_less) qed also have "... \ -S" by (simp add: B) finally have "\T. connected T \ T \ - S \ x \ T \ ((B + C) / norm x) *\<^sub>R x \ T" by (rule_tac x="closed_segment x (((B+C)/norm x) *\<^sub>R x)" in exI) simp with False B show ?thesis by (rule_tac x="((B+C)/norm x) *\<^sub>R x" in exI) (simp add: connected_component_def) qed } then show ?thesis apply (simp add: outside_def assms) apply (rule bounded_subset [OF bounded_ball [of 0 B]]) apply (force simp: dist_norm not_less bounded_pos) done qed lemma unbounded_outside: fixes S :: "'a::{real_normed_vector, perfect_space} set" shows "bounded S \ \ bounded(outside S)" using cobounded_imp_unbounded cobounded_outside by blast lemma bounded_inside: fixes S :: "'a::{real_normed_vector, perfect_space} set" shows "bounded S \ bounded(inside S)" by (simp add: bounded_Int cobounded_outside inside_outside) lemma connected_outside: fixes S :: "'a::euclidean_space set" assumes "bounded S" "2 \ DIM('a)" shows "connected(outside S)" apply (clarsimp simp add: connected_iff_connected_component outside) apply (rule_tac S="connected_component_set (- S) x" in connected_component_of_subset) apply (metis (no_types) assms cobounded_unbounded_component cobounded_unique_unbounded_component connected_component_eq_eq connected_component_idemp double_complement mem_Collect_eq) by (simp add: Collect_mono connected_component_eq) lemma outside_connected_component_lt: "outside S = {x. \B. \y. B < norm(y) \ connected_component (- S) x y}" proof - have "\x B. x \ outside S \ \y. B < norm y \ connected_component (- S) x y" by (metis boundedI linorder_not_less mem_Collect_eq outside) moreover have "\x. \B. \y. B < norm y \ connected_component (- S) x y \ x \ outside S" by (metis bounded_pos linorder_not_less mem_Collect_eq outside) ultimately show ?thesis by auto qed lemma outside_connected_component_le: "outside S = {x. \B. \y. B \ norm(y) \ connected_component (- S) x y}" apply (simp add: outside_connected_component_lt Set.set_eq_iff) by (meson gt_ex leD le_less_linear less_imp_le order.trans) lemma not_outside_connected_component_lt: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" and "2 \ DIM('a)" shows "- (outside S) = {x. \B. \y. B < norm(y) \ \ connected_component (- S) x y}" proof - obtain B::real where B: "0 < B" and Bno: "\x. x \ S \ norm x \ B" using S [simplified bounded_pos] by auto have cyz: "connected_component (- S) y z" if yz: "B < norm z" "B < norm y" for y::'a and z::'a proof - have "connected_component (- cball 0 B) y z" using assms yz by (force simp: dist_norm intro: connected_componentI [OF _ subset_refl] connected_complement_bounded_convex) then show ?thesis by (metis connected_component_of_subset Bno Compl_anti_mono mem_cball_0 subset_iff) qed show ?thesis apply (auto simp: outside bounded_pos) apply (metis Compl_iff bounded_iff cobounded_imp_unbounded mem_Collect_eq not_le) by (metis B connected_component_trans cyz not_le) qed lemma not_outside_connected_component_le: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" "2 \ DIM('a)" shows "- (outside S) = {x. \B. \y. B \ norm(y) \ \ connected_component (- S) x y}" unfolding not_outside_connected_component_lt [OF assms] by (metis (no_types, opaque_lifting) dual_order.strict_trans1 gt_ex pinf(8)) lemma inside_connected_component_lt: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" "2 \ DIM('a)" shows "inside S = {x. (x \ S) \ (\B. \y. B < norm(y) \ \ connected_component (- S) x y)}" by (auto simp: inside_outside not_outside_connected_component_lt [OF assms]) lemma inside_connected_component_le: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" "2 \ DIM('a)" shows "inside S = {x. (x \ S) \ (\B. \y. B \ norm(y) \ \ connected_component (- S) x y)}" by (auto simp: inside_outside not_outside_connected_component_le [OF assms]) lemma inside_subset: assumes "connected U" and "\ bounded U" and "T \ U = - S" shows "inside S \ T" using bounded_subset [of "connected_component_set (- S) _" U] assms by (metis (no_types, lifting) ComplI Un_iff connected_component_maximal inside_def mem_Collect_eq subsetI) lemma frontier_not_empty: fixes S :: "'a :: real_normed_vector set" shows "\S \ {}; S \ UNIV\ \ frontier S \ {}" using connected_Int_frontier [of UNIV S] by auto lemma frontier_eq_empty: fixes S :: "'a :: real_normed_vector set" shows "frontier S = {} \ S = {} \ S = UNIV" using frontier_UNIV frontier_empty frontier_not_empty by blast lemma frontier_of_connected_component_subset: fixes S :: "'a::real_normed_vector set" shows "frontier(connected_component_set S x) \ frontier S" proof - { fix y assume y1: "y \ closure (connected_component_set S x)" and y2: "y \ interior (connected_component_set S x)" have "y \ closure S" using y1 closure_mono connected_component_subset by blast moreover have "z \ interior (connected_component_set S x)" if "0 < e" "ball y e \ interior S" "dist y z < e" for e z proof - have "ball y e \ connected_component_set S y" using connected_component_maximal that interior_subset by (metis centre_in_ball connected_ball subset_trans) then show ?thesis using y1 apply (simp add: closure_approachable open_contains_ball_eq [OF open_interior]) by (metis connected_component_eq dist_commute mem_Collect_eq mem_ball mem_interior subsetD \0 < e\ y2) qed then have "y \ interior S" using y2 by (force simp: open_contains_ball_eq [OF open_interior]) ultimately have "y \ frontier S" by (auto simp: frontier_def) } then show ?thesis by (auto simp: frontier_def) qed lemma frontier_Union_subset_closure: fixes F :: "'a::real_normed_vector set set" shows "frontier(\F) \ closure(\t \ F. frontier t)" proof - have "\y\F. \y\frontier y. dist y x < e" if "T \ F" "y \ T" "dist y x < e" "x \ interior (\F)" "0 < e" for x y e T proof (cases "x \ T") case True with that show ?thesis by (metis Diff_iff Sup_upper closure_subset contra_subsetD dist_self frontier_def interior_mono) next case False have \
: "closed_segment x y \ T \ {}" "closed_segment x y - T \ {}" using \y \ T\ False by blast+ obtain c where "c \ closed_segment x y" "c \ frontier T" using False connected_Int_frontier [OF connected_segment \
] by auto with that show ?thesis by (smt (verit) dist_norm segment_bound1) qed then show ?thesis by (fastforce simp add: frontier_def closure_approachable) qed lemma frontier_Union_subset: fixes F :: "'a::real_normed_vector set set" shows "finite F \ frontier(\F) \ (\t \ F. frontier t)" by (metis closed_UN closure_closed frontier_Union_subset_closure frontier_closed) lemma frontier_of_components_subset: fixes S :: "'a::real_normed_vector set" shows "C \ components S \ frontier C \ frontier S" by (metis Path_Connected.frontier_of_connected_component_subset components_iff) lemma frontier_of_components_closed_complement: fixes S :: "'a::real_normed_vector set" shows "\closed S; C \ components (- S)\ \ frontier C \ S" using frontier_complement frontier_of_components_subset frontier_subset_eq by blast lemma frontier_minimal_separating_closed: fixes S :: "'a::real_normed_vector set" assumes "closed S" and nconn: "\ connected(- S)" and C: "C \ components (- S)" and conn: "\T. \closed T; T \ S\ \ connected(- T)" shows "frontier C = S" proof (rule ccontr) assume "frontier C \ S" then have "frontier C \ S" using frontier_of_components_closed_complement [OF \closed S\ C] by blast then have "connected(- (frontier C))" by (simp add: conn) have "\ connected(- (frontier C))" unfolding connected_def not_not proof (intro exI conjI) show "open C" using C \closed S\ open_components by blast show "open (- closure C)" by blast show "C \ - closure C \ - frontier C = {}" using closure_subset by blast show "C \ - frontier C \ {}" using C \open C\ components_eq frontier_disjoint_eq by fastforce show "- frontier C \ C \ - closure C" by (simp add: \open C\ closed_Compl frontier_closures) then show "- closure C \ - frontier C \ {}" by (metis C Compl_Diff_eq Un_Int_eq(4) Un_commute \frontier C \ S\ \open C\ compl_le_compl_iff frontier_def in_components_subset interior_eq leD sup_bot.right_neutral) qed then show False using \connected (- frontier C)\ by blast qed lemma connected_component_UNIV [simp]: fixes x :: "'a::real_normed_vector" shows "connected_component_set UNIV x = UNIV" using connected_iff_eq_connected_component_set [of "UNIV::'a set"] connected_UNIV by auto lemma connected_component_eq_UNIV: fixes x :: "'a::real_normed_vector" shows "connected_component_set s x = UNIV \ s = UNIV" using connected_component_in connected_component_UNIV by blast lemma components_UNIV [simp]: "components UNIV = {UNIV :: 'a::real_normed_vector set}" by (auto simp: components_eq_sing_iff) lemma interior_inside_frontier: fixes S :: "'a::real_normed_vector set" assumes "bounded S" shows "interior S \ inside (frontier S)" proof - { fix x y assume x: "x \ interior S" and y: "y \ S" and cc: "connected_component (- frontier S) x y" have "connected_component_set (- frontier S) x \ frontier S \ {}" proof (rule connected_Int_frontier; simp add: set_eq_iff) show "\u. connected_component (- frontier S) x u \ u \ S" by (meson cc connected_component_in connected_component_refl_eq interior_subset subsetD x) show "\u. connected_component (- frontier S) x u \ u \ S" using y cc by blast qed then have "bounded (connected_component_set (- frontier S) x)" using connected_component_in by auto } then show ?thesis using bounded_subset [OF assms] by (metis (no_types, lifting) Diff_iff frontier_def inside_def mem_Collect_eq subsetI) qed lemma inside_empty [simp]: "inside {} = ({} :: 'a :: {real_normed_vector, perfect_space} set)" by (simp add: inside_def) lemma outside_empty [simp]: "outside {} = (UNIV :: 'a :: {real_normed_vector, perfect_space} set)" using inside_empty inside_Un_outside by blast lemma inside_same_component: "\connected_component (- S) x y; x \ inside S\ \ y \ inside S" using connected_component_eq connected_component_in by (fastforce simp add: inside_def) lemma outside_same_component: "\connected_component (- S) x y; x \ outside S\ \ y \ outside S" using connected_component_eq connected_component_in by (fastforce simp add: outside_def) lemma convex_in_outside: fixes S :: "'a :: {real_normed_vector, perfect_space} set" assumes S: "convex S" and z: "z \ S" shows "z \ outside S" proof (cases "S={}") case True then show ?thesis by simp next case False then obtain a where "a \ S" by blast with z have zna: "z \ a" by auto { assume "bounded (connected_component_set (- S) z)" with bounded_pos_less obtain B where "B>0" and B: "\x. connected_component (- S) z x \ norm x < B" by (metis mem_Collect_eq) define C where "C = (B + 1 + norm z) / norm (z-a)" have "C > 0" using \0 < B\ zna by (simp add: C_def field_split_simps add_strict_increasing) have "\norm (z + C *\<^sub>R (z-a)) - norm (C *\<^sub>R (z-a))\ \ norm z" by (metis add_diff_cancel norm_triangle_ineq3) moreover have "norm (C *\<^sub>R (z-a)) > norm z + B" using zna \B>0\ by (simp add: C_def le_max_iff_disj) ultimately have C: "norm (z + C *\<^sub>R (z-a)) > B" by linarith { fix u::real assume u: "0\u" "u\1" and ins: "(1 - u) *\<^sub>R z + u *\<^sub>R (z + C *\<^sub>R (z - a)) \ S" then have Cpos: "1 + u * C > 0" by (meson \0 < C\ add_pos_nonneg less_eq_real_def zero_le_mult_iff zero_less_one) then have *: "(1 / (1 + u * C)) *\<^sub>R z + (u * C / (1 + u * C)) *\<^sub>R z = z" by (simp add: scaleR_add_left [symmetric] field_split_simps) then have False using convexD_alt [OF S \a \ S\ ins, of "1/(u*C + 1)"] \C>0\ \z \ S\ Cpos u by (simp add: * field_split_simps) } note contra = this have "connected_component (- S) z (z + C *\<^sub>R (z-a))" proof (rule connected_componentI [OF connected_segment]) show "closed_segment z (z + C *\<^sub>R (z - a)) \ - S" using contra by (force simp add: closed_segment_def) qed auto then have False using zna B [of "z + C *\<^sub>R (z-a)"] C by (auto simp: field_split_simps max_mult_distrib_right) } then show ?thesis by (auto simp: outside_def z) qed lemma outside_convex: fixes S :: "'a :: {real_normed_vector, perfect_space} set" assumes "convex S" shows "outside S = - S" by (metis ComplD assms convex_in_outside equalityI inside_Un_outside subsetI sup.cobounded2) lemma outside_singleton [simp]: fixes x :: "'a :: {real_normed_vector, perfect_space}" shows "outside {x} = -{x}" by (auto simp: outside_convex) lemma inside_convex: fixes S :: "'a :: {real_normed_vector, perfect_space} set" shows "convex S \ inside S = {}" by (simp add: inside_outside outside_convex) lemma inside_singleton [simp]: fixes x :: "'a :: {real_normed_vector, perfect_space}" shows "inside {x} = {}" by (auto simp: inside_convex) lemma outside_subset_convex: fixes S :: "'a :: {real_normed_vector, perfect_space} set" shows "\convex T; S \ T\ \ - T \ outside S" using outside_convex outside_mono by blast lemma outside_Un_outside_Un: fixes S :: "'a::real_normed_vector set" assumes "S \ outside(T \ U) = {}" shows "outside(T \ U) \ outside(T \ S)" proof fix x assume x: "x \ outside (T \ U)" have "Y \ - S" if "connected Y" "Y \ - T" "Y \ - U" "x \ Y" "u \ Y" for u Y proof - have "Y \ connected_component_set (- (T \ U)) x" by (simp add: connected_component_maximal that) also have "\ \ outside(T \ U)" by (metis (mono_tags, lifting) Collect_mono mem_Collect_eq outside outside_same_component x) finally have "Y \ outside(T \ U)" . with assms show ?thesis by auto qed with x show "x \ outside (T \ S)" by (simp add: outside_connected_component_lt connected_component_def) meson qed lemma outside_frontier_misses_closure: fixes S :: "'a::real_normed_vector set" assumes "bounded S" shows "outside(frontier S) \ - closure S" using assms frontier_def interior_inside_frontier outside_inside by fastforce lemma outside_frontier_eq_complement_closure: fixes S :: "'a :: {real_normed_vector, perfect_space} set" assumes "bounded S" "convex S" shows "outside(frontier S) = - closure S" by (metis Diff_subset assms convex_closure frontier_def outside_frontier_misses_closure outside_subset_convex subset_antisym) lemma inside_frontier_eq_interior: fixes S :: "'a :: {real_normed_vector, perfect_space} set" shows "\bounded S; convex S\ \ inside(frontier S) = interior S" unfolding inside_outside outside_frontier_eq_complement_closure using closure_subset interior_subset by (auto simp: frontier_def) lemma open_inside: fixes S :: "'a::real_normed_vector set" assumes "closed S" shows "open (inside S)" proof - { fix x assume x: "x \ inside S" have "open (connected_component_set (- S) x)" using assms open_connected_component by blast then obtain e where e: "e>0" and e: "\y. dist y x < e \ connected_component (- S) x y" using dist_not_less_zero apply (simp add: open_dist) by (metis (no_types, lifting) Compl_iff connected_component_refl_eq inside_def mem_Collect_eq x) then have "\e>0. ball x e \ inside S" by (metis e dist_commute inside_same_component mem_ball subsetI x) } then show ?thesis by (simp add: open_contains_ball) qed lemma open_outside: fixes S :: "'a::real_normed_vector set" assumes "closed S" shows "open (outside S)" proof - { fix x assume x: "x \ outside S" have "open (connected_component_set (- S) x)" using assms open_connected_component by blast then obtain e where e: "e>0" and e: "\y. dist y x < e \ connected_component (- S) x y" using dist_not_less_zero x by (auto simp add: open_dist outside_def intro: connected_component_refl) then have "\e>0. ball x e \ outside S" by (metis e dist_commute outside_same_component mem_ball subsetI x) } then show ?thesis by (simp add: open_contains_ball) qed lemma closure_inside_subset: fixes S :: "'a::real_normed_vector set" assumes "closed S" shows "closure(inside S) \ S \ inside S" by (metis assms closure_minimal open_closed open_outside sup.cobounded2 union_with_inside) lemma frontier_inside_subset: fixes S :: "'a::real_normed_vector set" assumes "closed S" shows "frontier(inside S) \ S" using assms closure_inside_subset frontier_closures frontier_disjoint_eq open_inside by fastforce lemma closure_outside_subset: fixes S :: "'a::real_normed_vector set" assumes "closed S" shows "closure(outside S) \ S \ outside S" by (metis assms closed_open closure_minimal inside_outside open_inside sup_ge2) lemma closed_path_image_Un_inside: fixes g :: "real \ 'a :: real_normed_vector" assumes "path g" shows "closed (path_image g \ inside (path_image g))" by (simp add: assms closed_Compl closed_path_image open_outside union_with_inside) lemma frontier_outside_subset: fixes S :: "'a::real_normed_vector set" assumes "closed S" shows "frontier(outside S) \ S" unfolding frontier_def by (metis Diff_subset_conv assms closure_outside_subset interior_eq open_outside sup_aci(1)) lemma inside_complement_unbounded_connected_empty: "\connected (- S); \ bounded (- S)\ \ inside S = {}" using inside_subset by blast lemma inside_bounded_complement_connected_empty: fixes S :: "'a::{real_normed_vector, perfect_space} set" shows "\connected (- S); bounded S\ \ inside S = {}" by (metis inside_complement_unbounded_connected_empty cobounded_imp_unbounded) lemma inside_inside: assumes "S \ inside T" shows "inside S - T \ inside T" unfolding inside_def proof clarify fix x assume x: "x \ T" "x \ S" and bo: "bounded (connected_component_set (- S) x)" show "bounded (connected_component_set (- T) x)" proof (cases "S \ connected_component_set (- T) x = {}") case True then show ?thesis by (metis bounded_subset [OF bo] compl_le_compl_iff connected_component_idemp connected_component_mono disjoint_eq_subset_Compl double_compl) next case False then obtain y where y: "y \ S" "y \ connected_component_set (- T) x" by (meson disjoint_iff) then have "bounded (connected_component_set (- T) y)" using assms [unfolded inside_def] by blast with y show ?thesis by (metis connected_component_eq) qed qed lemma inside_inside_subset: "inside(inside S) \ S" using inside_inside union_with_outside by fastforce lemma inside_outside_intersect_connected: "\connected T; inside S \ T \ {}; outside S \ T \ {}\ \ S \ T \ {}" apply (simp add: inside_def outside_def ex_in_conv [symmetric] disjoint_eq_subset_Compl, clarify) by (metis compl_le_swap1 connected_componentI connected_component_eq mem_Collect_eq) lemma outside_bounded_nonempty: fixes S :: "'a :: {real_normed_vector, perfect_space} set" assumes "bounded S" shows "outside S \ {}" using assms unbounded_outside by force lemma outside_compact_in_open: fixes S :: "'a :: {real_normed_vector,perfect_space} set" assumes S: "compact S" and T: "open T" and "S \ T" "T \ {}" shows "outside S \ T \ {}" proof - have "outside S \ {}" by (simp add: compact_imp_bounded outside_bounded_nonempty S) with assms obtain a b where a: "a \ outside S" and b: "b \ T" by auto show ?thesis proof (cases "a \ T") case True with a show ?thesis by blast next case False have front: "frontier T \ - S" using \S \ T\ frontier_disjoint_eq T by auto { fix \ assume "path \" and pimg_sbs: "path_image \ - {pathfinish \} \ interior (- T)" and pf: "pathfinish \ \ frontier T" and ps: "pathstart \ = a" define c where "c = pathfinish \" have "c \ -S" unfolding c_def using front pf by blast moreover have "open (-S)" using S compact_imp_closed by blast ultimately obtain \::real where "\ > 0" and \: "cball c \ \ -S" using open_contains_cball[of "-S"] S by blast then obtain d where "d \ T" and d: "dist d c < \" using closure_approachable [of c T] pf unfolding c_def by (metis Diff_iff frontier_def) then have "d \ -S" using \ using dist_commute by (metis contra_subsetD mem_cball not_le not_less_iff_gr_or_eq) have pimg_sbs_cos: "path_image \ \ -S" using \c \ - S\ \S \ T\ c_def interior_subset pimg_sbs by fastforce have "closed_segment c d \ cball c \" by (metis \0 < \\ centre_in_cball closed_segment_subset convex_cball d dist_commute less_eq_real_def mem_cball) with \ have "closed_segment c d \ -S" by blast moreover have con_gcd: "connected (path_image \ \ closed_segment c d)" by (rule connected_Un) (auto simp: c_def \path \\ connected_path_image) ultimately have "connected_component (- S) a d" unfolding connected_component_def using pimg_sbs_cos ps by blast then have "outside S \ T \ {}" using outside_same_component [OF _ a] by (metis IntI \d \ T\ empty_iff) } note * = this have pal: "pathstart (linepath a b) \ closure (- T)" by (auto simp: False closure_def) show ?thesis by (rule exists_path_subpath_to_frontier [OF path_linepath pal _ *]) (auto simp: b) qed qed lemma inside_inside_compact_connected: fixes S :: "'a :: euclidean_space set" assumes S: "closed S" and T: "compact T" and "connected T" "S \ inside T" shows "inside S \ inside T" proof (cases "inside T = {}") case True with assms show ?thesis by auto next case False consider "DIM('a) = 1" | "DIM('a) \ 2" using antisym not_less_eq_eq by fastforce then show ?thesis proof cases case 1 then show ?thesis using connected_convex_1_gen assms False inside_convex by blast next case 2 have "bounded S" using assms by (meson bounded_inside bounded_subset compact_imp_bounded) then have coms: "compact S" by (simp add: S compact_eq_bounded_closed) then have bst: "bounded (S \ T)" by (simp add: compact_imp_bounded T) then obtain r where "0 < r" and r: "S \ T \ ball 0 r" using bounded_subset_ballD by blast have outst: "outside S \ outside T \ {}" by (metis bounded_Un bounded_subset bst cobounded_outside disjoint_eq_subset_Compl unbounded_outside) have "S \ T = {}" using assms by (metis disjoint_iff_not_equal inside_no_overlap subsetCE) moreover have "outside S \ inside T \ {}" by (meson False assms(4) compact_eq_bounded_closed coms open_inside outside_compact_in_open T) ultimately have "inside S \ T = {}" using inside_outside_intersect_connected [OF \connected T\, of S] by (metis "2" compact_eq_bounded_closed coms connected_outside inf.commute inside_outside_intersect_connected outst) then show ?thesis using inside_inside [OF \S \ inside T\] by blast qed qed lemma connected_with_inside: fixes S :: "'a :: real_normed_vector set" assumes S: "closed S" and cons: "connected S" shows "connected(S \ inside S)" proof (cases "S \ inside S = UNIV") case True with assms show ?thesis by auto next case False then obtain b where b: "b \ S" "b \ inside S" by blast have *: "\y T. y \ S \ connected T \ a \ T \ y \ T \ T \ (S \ inside S)" if "a \ S \ inside S" for a using that proof assume "a \ S" then show ?thesis using cons by blast next assume a: "a \ inside S" then have ain: "a \ closure (inside S)" by (simp add: closure_def) obtain h where h: "path h" "pathstart h = a" "path_image h - {pathfinish h} \ interior (inside S)" "pathfinish h \ frontier (inside S)" using ain b by (metis exists_path_subpath_to_frontier path_linepath pathfinish_linepath pathstart_linepath) moreover have h1S: "pathfinish h \ S" using S h frontier_inside_subset by blast moreover have "path_image h \ S \ inside S" using IntD1 S h1S h interior_eq open_inside by fastforce ultimately show ?thesis by blast qed show ?thesis apply (simp add: connected_iff_connected_component) apply (clarsimp simp add: connected_component_def dest!: *) subgoal for x y u u' T t' by (rule_tac x = "S \ T \ t'" in exI) (auto intro!: connected_Un cons) done qed text\The proof is virtually the same as that above.\ lemma connected_with_outside: fixes S :: "'a :: real_normed_vector set" assumes S: "closed S" and cons: "connected S" shows "connected(S \ outside S)" proof (cases "S \ outside S = UNIV") case True with assms show ?thesis by auto next case False then obtain b where b: "b \ S" "b \ outside S" by blast have *: "\y T. y \ S \ connected T \ a \ T \ y \ T \ T \ (S \ outside S)" if "a \ (S \ outside S)" for a using that proof assume "a \ S" then show ?thesis by (rule_tac x=a in exI, rule_tac x="{a}" in exI, simp) next assume a: "a \ outside S" then have ain: "a \ closure (outside S)" by (simp add: closure_def) obtain h where h: "path h" "pathstart h = a" "path_image h - {pathfinish h} \ interior (outside S)" "pathfinish h \ frontier (outside S)" using ain b by (metis exists_path_subpath_to_frontier path_linepath pathfinish_linepath pathstart_linepath) moreover have h1S: "pathfinish h \ S" using S frontier_outside_subset h(4) by blast moreover have "path_image h \ S \ outside S" using IntD1 S h1S h interior_eq open_outside by fastforce ultimately show ?thesis by blast qed show ?thesis apply (simp add: connected_iff_connected_component) apply (clarsimp simp add: connected_component_def dest!: *) subgoal for x y u u' T t' by (rule_tac x="(S \ T \ t')" in exI) (auto intro!: connected_Un cons) done qed lemma inside_inside_eq_empty [simp]: fixes S :: "'a :: {real_normed_vector, perfect_space} set" assumes S: "closed S" and cons: "connected S" shows "inside (inside S) = {}" proof - have "connected (- inside S)" by (metis S connected_with_outside cons union_with_outside) then show ?thesis by (metis bounded_Un inside_complement_unbounded_connected_empty unbounded_outside union_with_outside) qed lemma inside_in_components: "inside S \ components (- S) \ connected(inside S) \ inside S \ {}" (is "?lhs = ?rhs") proof assume R: ?rhs then have "\x. \x \ S; x \ inside S\ \ \ connected (inside S)" by (simp add: inside_outside) with R show ?lhs unfolding in_components_maximal by (auto intro: inside_same_component connected_componentI) qed (simp add: in_components_maximal) text\The proof is like that above.\ lemma outside_in_components: "outside S \ components (- S) \ connected(outside S) \ outside S \ {}" (is "?lhs = ?rhs") proof assume R: ?rhs then have "\x. \x \ S; x \ outside S\ \ \ connected (outside S)" by (meson disjoint_iff outside_no_overlap) with R show ?lhs unfolding in_components_maximal by (auto intro: outside_same_component connected_componentI) qed (simp add: in_components_maximal) lemma bounded_unique_outside: fixes S :: "'a :: euclidean_space set" assumes "bounded S" "DIM('a) \ 2" shows "(c \ components (- S) \ \ bounded c) \ c = outside S" using assms by (metis cobounded_unique_unbounded_components connected_outside double_compl outside_bounded_nonempty outside_in_components unbounded_outside) subsection\Condition for an open map's image to contain a ball\ proposition ball_subset_open_map_image: fixes f :: "'a::heine_borel \ 'b :: {real_normed_vector,heine_borel}" assumes contf: "continuous_on (closure S) f" and oint: "open (f ` interior S)" and le_no: "\z. z \ frontier S \ r \ norm(f z - f a)" and "bounded S" "a \ S" "0 < r" shows "ball (f a) r \ f ` S" proof (cases "f ` S = UNIV") case True then show ?thesis by simp next case False then have "closed (frontier (f ` S))" "frontier (f ` S) \ {}" using \a \ S\ by (auto simp: frontier_eq_empty) then obtain w where w: "w \ frontier (f ` S)" and dw_le: "\y. y \ frontier (f ` S) \ norm (f a - w) \ norm (f a - y)" by (auto simp add: dist_norm intro: distance_attains_inf [of "frontier(f ` S)" "f a"]) then obtain \ where \: "\n. \ n \ f ` S" and tendsw: "\ \ w" by (metis Diff_iff frontier_def closure_sequential) then have "\n. \x \ S. \ n = f x" by force then obtain z where zs: "\n. z n \ S" and fz: "\n. \ n = f (z n)" by metis then obtain y K where y: "y \ closure S" and "strict_mono (K :: nat \ nat)" and Klim: "(z \ K) \ y" using \bounded S\ unfolding compact_closure [symmetric] compact_def by (meson closure_subset subset_iff) then have ftendsw: "((\n. f (z n)) \ K) \ w" by (metis LIMSEQ_subseq_LIMSEQ fun.map_cong0 fz tendsw) have zKs: "\n. (z \ K) n \ S" by (simp add: zs) have fz: "f \ z = \" "(\n. f (z n)) = \" using fz by auto then have "(\ \ K) \ f y" by (metis (no_types) Klim zKs y contf comp_assoc continuous_on_closure_sequentially) with fz have wy: "w = f y" using fz LIMSEQ_unique ftendsw by auto have "r \ norm (f y - f a)" proof (rule le_no) show "y \ frontier S" using w wy oint by (force simp: imageI image_mono interiorI interior_subset frontier_def y) qed then have "\y. \norm (f a - y) < r; y \ frontier (f ` S)\ \ False" by (metis dw_le norm_minus_commute not_less order_trans wy) then have "ball (f a) r \ frontier (f ` S) = {}" by (metis disjoint_iff_not_equal dist_norm mem_ball) moreover have "ball (f a) r \ f ` S \ {}" using \a \ S\ \0 < r\ centre_in_ball by blast ultimately show ?thesis by (meson connected_Int_frontier connected_ball diff_shunt_var) qed subsubsection\Special characterizations of classes of functions into and out of R.\ lemma Hausdorff_space_euclidean [simp]: "Hausdorff_space (euclidean :: 'a::metric_space topology)" proof - have "\U V. open U \ open V \ x \ U \ y \ V \ disjnt U V" if "x \ y" for x y :: 'a proof (intro exI conjI) let ?r = "dist x y / 2" have [simp]: "?r > 0" by (simp add: that) show "open (ball x ?r)" "open (ball y ?r)" "x \ (ball x ?r)" "y \ (ball y ?r)" by (auto simp add: that) show "disjnt (ball x ?r) (ball y ?r)" unfolding disjnt_def by (simp add: disjoint_ballI) qed then show ?thesis by (simp add: Hausdorff_space_def) qed proposition embedding_map_into_euclideanreal: assumes "path_connected_space X" shows "embedding_map X euclideanreal f \ continuous_map X euclideanreal f \ inj_on f (topspace X)" proof safe show "continuous_map X euclideanreal f" if "embedding_map X euclideanreal f" using continuous_map_in_subtopology homeomorphic_imp_continuous_map that unfolding embedding_map_def by blast show "inj_on f (topspace X)" if "embedding_map X euclideanreal f" using that homeomorphic_imp_injective_map unfolding embedding_map_def by blast show "embedding_map X euclideanreal f" if cont: "continuous_map X euclideanreal f" and inj: "inj_on f (topspace X)" proof - obtain g where gf: "\x. x \ topspace X \ g (f x) = x" using inv_into_f_f [OF inj] by auto show ?thesis unfolding embedding_map_def homeomorphic_map_maps homeomorphic_maps_def proof (intro exI conjI) show "continuous_map X (top_of_set (f ` topspace X)) f" by (simp add: cont continuous_map_in_subtopology) let ?S = "f ` topspace X" have eq: "{x \ ?S. g x \ U} = f ` U" if "openin X U" for U using openin_subset [OF that] by (auto simp: gf) have 1: "g ` ?S \ topspace X" using eq by blast have "openin (top_of_set ?S) {x \ ?S. g x \ T}" if "openin X T" for T proof - have "T \ topspace X" by (simp add: openin_subset that) have RR: "\x \ ?S \ g -` T. \d>0. \x' \ ?S \ ball x d. g x' \ T" proof (clarsimp simp add: gf) have pcS: "path_connectedin euclidean ?S" using assms cont path_connectedin_continuous_map_image path_connectedin_topspace by blast show "\d>0. \x'\f ` topspace X \ ball (f x) d. g x' \ T" if "x \ T" for x proof - have x: "x \ topspace X" using \T \ topspace X\ \x \ T\ by blast obtain u v d where "0 < d" "u \ topspace X" "v \ topspace X" and sub_fuv: "?S \ {f x - d .. f x + d} \ {f u..f v}" proof (cases "\u \ topspace X. f u < f x") case True then obtain u where u: "u \ topspace X" "f u < f x" .. show ?thesis proof (cases "\v \ topspace X. f x < f v") case True then obtain v where v: "v \ topspace X" "f x < f v" .. show ?thesis proof let ?d = "min (f x - f u) (f v - f x)" show "0 < ?d" by (simp add: \f u < f x\ \f x < f v\) show "f ` topspace X \ {f x - ?d..f x + ?d} \ {f u..f v}" by fastforce qed (auto simp: u v) next case False show ?thesis proof let ?d = "f x - f u" show "0 < ?d" by (simp add: u) show "f ` topspace X \ {f x - ?d..f x + ?d} \ {f u..f x}" using x u False by auto qed (auto simp: x u) qed next case False note no_u = False show ?thesis proof (cases "\v \ topspace X. f x < f v") case True then obtain v where v: "v \ topspace X" "f x < f v" .. show ?thesis proof let ?d = "f v - f x" show "0 < ?d" by (simp add: v) show "f ` topspace X \ {f x - ?d..f x + ?d} \ {f x..f v}" using False by auto qed (auto simp: x v) next case False show ?thesis proof show "f ` topspace X \ {f x - 1..f x + 1} \ {f x..f x}" using False no_u by fastforce qed (auto simp: x) qed qed then obtain h where "pathin X h" "h 0 = u" "h 1 = v" using assms unfolding path_connected_space_def by blast obtain C where "compactin X C" "connectedin X C" "u \ C" "v \ C" proof show "compactin X (h ` {0..1})" using that by (simp add: \pathin X h\ compactin_path_image) show "connectedin X (h ` {0..1})" using \pathin X h\ connectedin_path_image by blast qed (use \h 0 = u\ \h 1 = v\ in auto) have "continuous_map (subtopology euclideanreal (?S \ {f x - d .. f x + d})) (subtopology X C) g" proof (rule continuous_inverse_map) show "compact_space (subtopology X C)" using \compactin X C\ compactin_subspace by blast show "continuous_map (subtopology X C) euclideanreal f" by (simp add: cont continuous_map_from_subtopology) have "{f u .. f v} \ f ` topspace (subtopology X C)" proof (rule connected_contains_Icc) show "connected (f ` topspace (subtopology X C))" using connectedin_continuous_map_image [OF cont] by (simp add: \compactin X C\ \connectedin X C\ compactin_subset_topspace inf_absorb2) show "f u \ f ` topspace (subtopology X C)" by (simp add: \u \ C\ \u \ topspace X\) show "f v \ f ` topspace (subtopology X C)" by (simp add: \v \ C\ \v \ topspace X\) qed then show "f ` topspace X \ {f x - d..f x + d} \ f ` topspace (subtopology X C)" using sub_fuv by blast qed (auto simp: gf) then have contg: "continuous_map (subtopology euclideanreal (?S \ {f x - d .. f x + d})) X g" using continuous_map_in_subtopology by blast have "\e>0. \x \ ?S \ {f x - d .. f x + d} \ ball (f x) e. g x \ T" using openin_continuous_map_preimage [OF contg \openin X T\] x \x \ T\ \0 < d\ unfolding openin_euclidean_subtopology_iff by (force simp: gf dist_commute) then obtain e where "e > 0 \ (\x\f ` topspace X \ {f x - d..f x + d} \ ball (f x) e. g x \ T)" by metis with \0 < d\ have "min d e > 0" "\u. u \ topspace X \ \f x - f u\ < min d e \ u \ T" using dist_real_def gf by force+ then show ?thesis by (metis (full_types) Int_iff dist_real_def image_iff mem_ball gf) qed qed then obtain d where d: "\r. r \ ?S \ g -` T \ d r > 0 \ (\x \ ?S \ ball r (d r). g x \ T)" by metis show ?thesis unfolding openin_subtopology proof (intro exI conjI) show "{x \ ?S. g x \ T} = (\r \ ?S \ g -` T. ball r (d r)) \ f ` topspace X" using d by (auto simp: gf) qed auto qed then show "continuous_map (top_of_set ?S) X g" by (simp add: "1" continuous_map) qed (auto simp: gf) qed qed subsubsection \An injective function into R is a homeomorphism and so an open map.\ lemma injective_into_1d_eq_homeomorphism: fixes f :: "'a::topological_space \ real" assumes f: "continuous_on S f" and S: "path_connected S" shows "inj_on f S \ (\g. homeomorphism S (f ` S) f g)" proof show "\g. homeomorphism S (f ` S) f g" if "inj_on f S" proof - have "embedding_map (top_of_set S) euclideanreal f" using that embedding_map_into_euclideanreal [of "top_of_set S" f] assms by auto then show ?thesis unfolding embedding_map_def topspace_euclidean_subtopology by (metis f homeomorphic_map_closedness_eq homeomorphism_injective_closed_map that) qed qed (metis homeomorphism_def inj_onI) lemma injective_into_1d_imp_open_map: fixes f :: "'a::topological_space \ real" assumes "continuous_on S f" "path_connected S" "inj_on f S" "openin (subtopology euclidean S) T" shows "openin (subtopology euclidean (f ` S)) (f ` T)" using assms homeomorphism_imp_open_map injective_into_1d_eq_homeomorphism by blast lemma homeomorphism_into_1d: fixes f :: "'a::topological_space \ real" assumes "path_connected S" "continuous_on S f" "f ` S = T" "inj_on f S" shows "\g. homeomorphism S T f g" using assms injective_into_1d_eq_homeomorphism by blast subsection\<^marker>\tag unimportant\ \Rectangular paths\ definition\<^marker>\tag unimportant\ rectpath where "rectpath a1 a3 = (let a2 = Complex (Re a3) (Im a1); a4 = Complex (Re a1) (Im a3) in linepath a1 a2 +++ linepath a2 a3 +++ linepath a3 a4 +++ linepath a4 a1)" lemma path_rectpath [simp, intro]: "path (rectpath a b)" by (simp add: Let_def rectpath_def) lemma pathstart_rectpath [simp]: "pathstart (rectpath a1 a3) = a1" by (simp add: rectpath_def Let_def) lemma pathfinish_rectpath [simp]: "pathfinish (rectpath a1 a3) = a1" by (simp add: rectpath_def Let_def) lemma simple_path_rectpath [simp, intro]: assumes "Re a1 \ Re a3" "Im a1 \ Im a3" shows "simple_path (rectpath a1 a3)" unfolding rectpath_def Let_def using assms by (intro simple_path_join_loop arc_join arc_linepath) (auto simp: complex_eq_iff path_image_join closed_segment_same_Re closed_segment_same_Im) lemma path_image_rectpath: assumes "Re a1 \ Re a3" "Im a1 \ Im a3" shows "path_image (rectpath a1 a3) = {z. Re z \ {Re a1, Re a3} \ Im z \ {Im a1..Im a3}} \ {z. Im z \ {Im a1, Im a3} \ Re z \ {Re a1..Re a3}}" (is "?lhs = ?rhs") proof - define a2 a4 where "a2 = Complex (Re a3) (Im a1)" and "a4 = Complex (Re a1) (Im a3)" have "?lhs = closed_segment a1 a2 \ closed_segment a2 a3 \ closed_segment a4 a3 \ closed_segment a1 a4" by (simp_all add: rectpath_def Let_def path_image_join closed_segment_commute a2_def a4_def Un_assoc) also have "\ = ?rhs" using assms by (auto simp: rectpath_def Let_def path_image_join a2_def a4_def closed_segment_same_Re closed_segment_same_Im closed_segment_eq_real_ivl) finally show ?thesis . qed lemma path_image_rectpath_subset_cbox: assumes "Re a \ Re b" "Im a \ Im b" shows "path_image (rectpath a b) \ cbox a b" using assms by (auto simp: path_image_rectpath in_cbox_complex_iff) lemma path_image_rectpath_inter_box: assumes "Re a \ Re b" "Im a \ Im b" shows "path_image (rectpath a b) \ box a b = {}" using assms by (auto simp: path_image_rectpath in_box_complex_iff) lemma path_image_rectpath_cbox_minus_box: assumes "Re a \ Re b" "Im a \ Im b" shows "path_image (rectpath a b) = cbox a b - box a b" using assms by (auto simp: path_image_rectpath in_cbox_complex_iff in_box_complex_iff) end diff --git a/src/HOL/Analysis/Smooth_Paths.thy b/src/HOL/Analysis/Smooth_Paths.thy --- a/src/HOL/Analysis/Smooth_Paths.thy +++ b/src/HOL/Analysis/Smooth_Paths.thy @@ -1,442 +1,534 @@ (* Material originally from HOL Light, ported by Larry Paulson, moved here by Manuel Eberl *) section\<^marker>\tag unimportant\ \Smooth paths\ theory Smooth_Paths - imports - Retracts + imports Retracts begin subsection\<^marker>\tag unimportant\ \Homeomorphisms of arc images\ lemma path_connected_arc_complement: fixes \ :: "real \ 'a::euclidean_space" assumes "arc \" "2 \ DIM('a)" shows "path_connected(- path_image \)" proof - have "path_image \ homeomorphic {0..1::real}" by (simp add: assms homeomorphic_arc_image_interval) then show ?thesis by (intro path_connected_complement_homeomorphic_convex_compact) (auto simp: assms) qed lemma connected_arc_complement: fixes \ :: "real \ 'a::euclidean_space" assumes "arc \" "2 \ DIM('a)" shows "connected(- path_image \)" by (simp add: assms path_connected_arc_complement path_connected_imp_connected) lemma inside_arc_empty: fixes \ :: "real \ 'a::euclidean_space" assumes "arc \" shows "inside(path_image \) = {}" proof (cases "DIM('a) = 1") case True then show ?thesis using assms connected_arc_image connected_convex_1_gen inside_convex by blast next case False then have "connected (- path_image \)" by (metis DIM_ge_Suc0 One_nat_def Suc_1 antisym assms connected_arc_complement not_less_eq_eq) then show ?thesis by (simp add: assms bounded_arc_image inside_bounded_complement_connected_empty) qed lemma inside_simple_curve_imp_closed: fixes \ :: "real \ 'a::euclidean_space" shows "\simple_path \; x \ inside(path_image \)\ \ pathfinish \ = pathstart \" using arc_simple_path inside_arc_empty by blast subsection \Piecewise differentiability of paths\ lemma continuous_on_joinpaths_D1: assumes "continuous_on {0..1} (g1 +++ g2)" shows "continuous_on {0..1} g1" proof (rule continuous_on_eq) have "continuous_on {0..1/2} (g1 +++ g2)" using assms continuous_on_subset split_01 by auto then show "continuous_on {0..1} (g1 +++ g2 \ (*) (inverse 2))" by (intro continuous_intros) force qed (auto simp: joinpaths_def) lemma continuous_on_joinpaths_D2: "\continuous_on {0..1} (g1 +++ g2); pathfinish g1 = pathstart g2\ \ continuous_on {0..1} g2" using path_def path_join by blast lemma piecewise_differentiable_D1: assumes "(g1 +++ g2) piecewise_differentiable_on {0..1}" shows "g1 piecewise_differentiable_on {0..1}" proof - obtain S where cont: "continuous_on {0..1} g1" and "finite S" and S: "\x. x \ {0..1} - S \ g1 +++ g2 differentiable at x within {0..1}" using assms unfolding piecewise_differentiable_on_def by (blast dest!: continuous_on_joinpaths_D1) show ?thesis unfolding piecewise_differentiable_on_def proof (intro exI conjI ballI cont) show "finite (insert 1 (((*)2) ` S))" by (simp add: \finite S\) show "g1 differentiable at x within {0..1}" if "x \ {0..1} - insert 1 ((*) 2 ` S)" for x proof (rule_tac d="dist (x/2) (1/2)" in differentiable_transform_within) have "g1 +++ g2 differentiable at (x / 2) within {0..1/2}" by (rule differentiable_subset [OF S [of "x/2"]] | use that in force)+ then show "g1 +++ g2 \ (*) (inverse 2) differentiable at x within {0..1}" using image_affinity_atLeastAtMost_div [of 2 0 "0::real" 1] by (auto intro: differentiable_chain_within) qed (use that in \auto simp: joinpaths_def\) qed qed lemma piecewise_differentiable_D2: assumes "(g1 +++ g2) piecewise_differentiable_on {0..1}" and eq: "pathfinish g1 = pathstart g2" shows "g2 piecewise_differentiable_on {0..1}" proof - have [simp]: "g1 1 = g2 0" using eq by (simp add: pathfinish_def pathstart_def) obtain S where cont: "continuous_on {0..1} g2" and "finite S" and S: "\x. x \ {0..1} - S \ g1 +++ g2 differentiable at x within {0..1}" using assms unfolding piecewise_differentiable_on_def by (blast dest!: continuous_on_joinpaths_D2) show ?thesis unfolding piecewise_differentiable_on_def proof (intro exI conjI ballI cont) show "finite (insert 0 ((\x. 2*x-1)`S))" by (simp add: \finite S\) show "g2 differentiable at x within {0..1}" if "x \ {0..1} - insert 0 ((\x. 2*x-1)`S)" for x proof (rule_tac d="dist ((x+1)/2) (1/2)" in differentiable_transform_within) have x2: "(x + 1) / 2 \ S" using that apply (clarsimp simp: image_iff) by (metis add.commute add_diff_cancel_left' mult_2 field_sum_of_halves) have "g1 +++ g2 \ (\x. (x+1) / 2) differentiable at x within {0..1}" by (rule differentiable_chain_within differentiable_subset [OF S [of "(x+1)/2"]] | use x2 that in force)+ then show "g1 +++ g2 \ (\x. (x+1) / 2) differentiable at x within {0..1}" by (auto intro: differentiable_chain_within) show "(g1 +++ g2 \ (\x. (x + 1) / 2)) x' = g2 x'" if "x' \ {0..1}" "dist x' x < dist ((x + 1) / 2) (1/2)" for x' proof - have [simp]: "(2*x'+2)/2 = x'+1" by (simp add: field_split_simps) show ?thesis using that by (auto simp: joinpaths_def) qed qed (use that in \auto simp: joinpaths_def\) qed qed lemma piecewise_C1_differentiable_D1: fixes g1 :: "real \ 'a::real_normed_field" assumes "(g1 +++ g2) piecewise_C1_differentiable_on {0..1}" shows "g1 piecewise_C1_differentiable_on {0..1}" proof - obtain S where "finite S" and co12: "continuous_on ({0..1} - S) (\x. vector_derivative (g1 +++ g2) (at x))" and g12D: "\x\{0..1} - S. g1 +++ g2 differentiable at x" using assms by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) have g1D: "g1 differentiable at x" if "x \ {0..1} - insert 1 ((*) 2 ` S)" for x proof (rule differentiable_transform_within) show "g1 +++ g2 \ (*) (inverse 2) differentiable at x" using that g12D unfolding joinpaths_def by (intro differentiable_chain_at derivative_intros | force)+ show "\x'. \dist x' x < dist (x/2) (1/2)\ \ (g1 +++ g2 \ (*) (inverse 2)) x' = g1 x'" using that by (auto simp: dist_real_def joinpaths_def) qed (use that in \auto simp: dist_real_def\) have [simp]: "vector_derivative (g1 \ (*) 2) (at (x/2)) = 2 *\<^sub>R vector_derivative g1 (at x)" if "x \ {0..1} - insert 1 ((*) 2 ` S)" for x apply (subst vector_derivative_chain_at) using that apply (rule derivative_eq_intros g1D | simp)+ done have "continuous_on ({0..1/2} - insert (1/2) S) (\x. vector_derivative (g1 +++ g2) (at x))" using co12 by (rule continuous_on_subset) force then have coDhalf: "continuous_on ({0..1/2} - insert (1/2) S) (\x. vector_derivative (g1 \ (*)2) (at x))" proof (rule continuous_on_eq [OF _ vector_derivative_at]) show "(g1 +++ g2 has_vector_derivative vector_derivative (g1 \ (*) 2) (at x)) (at x)" if "x \ {0..1/2} - insert (1/2) S" for x proof (rule has_vector_derivative_transform_within) show "(g1 \ (*) 2 has_vector_derivative vector_derivative (g1 \ (*) 2) (at x)) (at x)" using that by (force intro: g1D differentiable_chain_at simp: vector_derivative_works [symmetric]) show "\x'. \dist x' x < dist x (1/2)\ \ (g1 \ (*) 2) x' = (g1 +++ g2) x'" using that by (auto simp: dist_norm joinpaths_def) qed (use that in \auto simp: dist_norm\) qed have "continuous_on ({0..1} - insert 1 ((*) 2 ` S)) ((\x. 1/2 * vector_derivative (g1 \ (*)2) (at x)) \ (*)(1/2))" using coDhalf apply (intro continuous_intros) by (simp add: scaleR_conv_of_real image_set_diff image_image) then have con_g1: "continuous_on ({0..1} - insert 1 ((*) 2 ` S)) (\x. vector_derivative g1 (at x))" by (rule continuous_on_eq) (simp add: scaleR_conv_of_real) have "continuous_on {0..1} g1" using continuous_on_joinpaths_D1 assms piecewise_C1_differentiable_on_def by blast with \finite S\ show ?thesis apply (clarsimp simp add: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) apply (rule_tac x="insert 1 (((*)2)`S)" in exI) apply (simp add: g1D con_g1) done qed lemma piecewise_C1_differentiable_D2: fixes g2 :: "real \ 'a::real_normed_field" assumes "(g1 +++ g2) piecewise_C1_differentiable_on {0..1}" "pathfinish g1 = pathstart g2" shows "g2 piecewise_C1_differentiable_on {0..1}" proof - obtain S where "finite S" and co12: "continuous_on ({0..1} - S) (\x. vector_derivative (g1 +++ g2) (at x))" and g12D: "\x\{0..1} - S. g1 +++ g2 differentiable at x" using assms by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) have g2D: "g2 differentiable at x" if "x \ {0..1} - insert 0 ((\x. 2*x-1) ` S)" for x proof (rule differentiable_transform_within) show "g1 +++ g2 \ (\x. (x + 1) / 2) differentiable at x" using g12D that unfolding joinpaths_def apply (drule_tac x= "(x+1) / 2" in bspec, force simp: field_split_simps) apply (rule differentiable_chain_at derivative_intros | force)+ done show "\x'. dist x' x < dist ((x + 1) / 2) (1/2) \ (g1 +++ g2 \ (\x. (x + 1) / 2)) x' = g2 x'" using that by (auto simp: dist_real_def joinpaths_def field_simps) qed (use that in \auto simp: dist_norm\) have [simp]: "vector_derivative (g2 \ (\x. 2*x-1)) (at ((x+1)/2)) = 2 *\<^sub>R vector_derivative g2 (at x)" if "x \ {0..1} - insert 0 ((\x. 2*x-1) ` S)" for x using that by (auto simp: vector_derivative_chain_at field_split_simps g2D) have "continuous_on ({1/2..1} - insert (1/2) S) (\x. vector_derivative (g1 +++ g2) (at x))" using co12 by (rule continuous_on_subset) force then have coDhalf: "continuous_on ({1/2..1} - insert (1/2) S) (\x. vector_derivative (g2 \ (\x. 2*x-1)) (at x))" proof (rule continuous_on_eq [OF _ vector_derivative_at]) show "(g1 +++ g2 has_vector_derivative vector_derivative (g2 \ (\x. 2 * x - 1)) (at x)) (at x)" if "x \ {1 / 2..1} - insert (1 / 2) S" for x proof (rule_tac f="g2 \ (\x. 2*x-1)" and d="dist (3/4) ((x+1)/2)" in has_vector_derivative_transform_within) show "(g2 \ (\x. 2 * x - 1) has_vector_derivative vector_derivative (g2 \ (\x. 2 * x - 1)) (at x)) (at x)" using that by (force intro: g2D differentiable_chain_at simp: vector_derivative_works [symmetric]) show "\x'. \dist x' x < dist (3 / 4) ((x + 1) / 2)\ \ (g2 \ (\x. 2 * x - 1)) x' = (g1 +++ g2) x'" using that by (auto simp: dist_norm joinpaths_def add_divide_distrib) qed (use that in \auto simp: dist_norm\) qed have [simp]: "((\x. (x+1) / 2) ` ({0..1} - insert 0 ((\x. 2 * x - 1) ` S))) = ({1/2..1} - insert (1/2) S)" apply (simp add: image_set_diff inj_on_def image_image) apply (auto simp: image_affinity_atLeastAtMost_div add_divide_distrib) done have "continuous_on ({0..1} - insert 0 ((\x. 2*x-1) ` S)) ((\x. 1/2 * vector_derivative (g2 \ (\x. 2*x-1)) (at x)) \ (\x. (x+1)/2))" by (rule continuous_intros | simp add: coDhalf)+ then have con_g2: "continuous_on ({0..1} - insert 0 ((\x. 2*x-1) ` S)) (\x. vector_derivative g2 (at x))" by (rule continuous_on_eq) (simp add: scaleR_conv_of_real) have "continuous_on {0..1} g2" using continuous_on_joinpaths_D2 assms piecewise_C1_differentiable_on_def by blast with \finite S\ show ?thesis by (meson C1_differentiable_on_eq con_g2 finite_imageI finite_insert g2D piecewise_C1_differentiable_on_def) qed subsection \Valid paths, and their start and finish\ definition\<^marker>\tag important\ valid_path :: "(real \ 'a :: real_normed_vector) \ bool" where "valid_path f \ f piecewise_C1_differentiable_on {0..1::real}" definition closed_path :: "(real \ 'a :: real_normed_vector) \ bool" where "closed_path g \ g 0 = g 1" text\In particular, all results for paths apply\ lemma valid_path_imp_path: "valid_path g \ path g" by (simp add: path_def piecewise_C1_differentiable_on_def valid_path_def) lemma connected_valid_path_image: "valid_path g \ connected(path_image g)" by (metis connected_path_image valid_path_imp_path) lemma compact_valid_path_image: "valid_path g \ compact(path_image g)" by (metis compact_path_image valid_path_imp_path) lemma bounded_valid_path_image: "valid_path g \ bounded(path_image g)" by (metis bounded_path_image valid_path_imp_path) lemma closed_valid_path_image: "valid_path g \ closed(path_image g)" by (metis closed_path_image valid_path_imp_path) lemma valid_path_translation_eq: "valid_path ((+)d \ p) \ valid_path p" by (simp add: valid_path_def piecewise_C1_differentiable_on_translation_eq) lemma valid_path_compose: assumes "valid_path g" and der: "\x. x \ path_image g \ f field_differentiable (at x)" and con: "continuous_on (path_image g) (deriv f)" shows "valid_path (f \ g)" proof - obtain S where "finite S" and g_diff: "g C1_differentiable_on {0..1} - S" using \valid_path g\ unfolding valid_path_def piecewise_C1_differentiable_on_def by auto have "f \ g differentiable at t" when "t\{0..1} - S" for t proof (rule differentiable_chain_at) show "g differentiable at t" using \valid_path g\ by (meson C1_differentiable_on_eq \g C1_differentiable_on {0..1} - S\ that) next have "g t\path_image g" using that DiffD1 image_eqI path_image_def by metis then show "f differentiable at (g t)" using der[THEN field_differentiable_imp_differentiable] by auto qed moreover have "continuous_on ({0..1} - S) (\x. vector_derivative (f \ g) (at x))" proof (rule continuous_on_eq [where f = "\x. vector_derivative g (at x) * deriv f (g x)"], rule continuous_intros) show "continuous_on ({0..1} - S) (\x. vector_derivative g (at x))" using g_diff C1_differentiable_on_eq by auto next have "continuous_on {0..1} (\x. deriv f (g x))" using continuous_on_compose[OF _ con[unfolded path_image_def],unfolded comp_def] \valid_path g\ piecewise_C1_differentiable_on_def valid_path_def by blast then show "continuous_on ({0..1} - S) (\x. deriv f (g x))" using continuous_on_subset by blast next show "vector_derivative g (at t) * deriv f (g t) = vector_derivative (f \ g) (at t)" when "t \ {0..1} - S" for t by (metis C1_differentiable_on_eq DiffD1 der g_diff imageI path_image_def that vector_derivative_chain_at_general) qed ultimately have "f \ g C1_differentiable_on {0..1} - S" using C1_differentiable_on_eq by blast moreover have "path (f \ g)" using der by (simp add: path_continuous_image[OF valid_path_imp_path[OF \valid_path g\]] continuous_at_imp_continuous_on field_differentiable_imp_continuous_at) ultimately show ?thesis unfolding valid_path_def piecewise_C1_differentiable_on_def path_def using \finite S\ by auto qed lemma valid_path_uminus_comp[simp]: fixes g::"real \ 'a ::real_normed_field" shows "valid_path (uminus \ g) \ valid_path g" proof show "valid_path g \ valid_path (uminus \ g)" for g::"real \ 'a" by (auto intro!: valid_path_compose derivative_intros) then show "valid_path g" when "valid_path (uminus \ g)" by (metis fun.map_comp group_add_class.minus_comp_minus id_comp that) qed lemma valid_path_offset[simp]: shows "valid_path (\t. g t - z) \ valid_path g" proof show *: "valid_path (g::real\'a) \ valid_path (\t. g t - z)" for g z unfolding valid_path_def by (fastforce intro:derivative_intros C1_differentiable_imp_piecewise piecewise_C1_differentiable_diff) show "valid_path (\t. g t - z) \ valid_path g" using *[of "\t. g t - z" "-z",simplified] . qed lemma valid_path_imp_reverse: assumes "valid_path g" shows "valid_path(reversepath g)" proof - obtain S where "finite S" and S: "g C1_differentiable_on ({0..1} - S)" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) then have "finite ((-) 1 ` S)" by auto moreover have "(reversepath g C1_differentiable_on ({0..1} - (-) 1 ` S))" unfolding reversepath_def apply (rule C1_differentiable_compose [of "\x::real. 1-x" _ g, unfolded o_def]) using S by (force simp: finite_vimageI inj_on_def C1_differentiable_on_eq elim!: continuous_on_subset)+ ultimately show ?thesis using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def path_def [symmetric]) qed lemma valid_path_reversepath [simp]: "valid_path(reversepath g) \ valid_path g" using valid_path_imp_reverse by force lemma valid_path_join: assumes "valid_path g1" "valid_path g2" "pathfinish g1 = pathstart g2" shows "valid_path(g1 +++ g2)" proof - have "g1 1 = g2 0" using assms by (auto simp: pathfinish_def pathstart_def) moreover have "(g1 \ (\x. 2*x)) piecewise_C1_differentiable_on {0..1/2}" apply (rule piecewise_C1_differentiable_compose) using assms apply (auto simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_joinpaths) apply (force intro: finite_vimageI [where h = "(*)2"] inj_onI) done moreover have "(g2 \ (\x. 2*x-1)) piecewise_C1_differentiable_on {1/2..1}" apply (rule piecewise_C1_differentiable_compose) using assms unfolding valid_path_def piecewise_C1_differentiable_on_def by (auto intro!: continuous_intros finite_vimageI [where h = "(\x. 2*x - 1)"] inj_onI simp: image_affinity_atLeastAtMost_diff continuous_on_joinpaths) ultimately show ?thesis unfolding valid_path_def continuous_on_joinpaths joinpaths_def by (intro piecewise_C1_differentiable_cases) (auto simp: o_def) qed lemma valid_path_join_D1: fixes g1 :: "real \ 'a::real_normed_field" shows "valid_path (g1 +++ g2) \ valid_path g1" unfolding valid_path_def by (rule piecewise_C1_differentiable_D1) lemma valid_path_join_D2: fixes g2 :: "real \ 'a::real_normed_field" shows "\valid_path (g1 +++ g2); pathfinish g1 = pathstart g2\ \ valid_path g2" unfolding valid_path_def by (rule piecewise_C1_differentiable_D2) lemma valid_path_join_eq [simp]: fixes g2 :: "real \ 'a::real_normed_field" shows "pathfinish g1 = pathstart g2 \ (valid_path(g1 +++ g2) \ valid_path g1 \ valid_path g2)" using valid_path_join_D1 valid_path_join_D2 valid_path_join by blast lemma valid_path_shiftpath [intro]: assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" shows "valid_path(shiftpath a g)" using assms unfolding valid_path_def shiftpath_alt_def apply (intro piecewise_C1_differentiable_cases) apply (simp_all add: add.commute) apply (rule piecewise_C1_differentiable_affine [of g 1 a, simplified o_def scaleR_one]) apply (force simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset) apply (rule piecewise_C1_differentiable_affine [of g 1 "a-1", simplified o_def scaleR_one algebra_simps]) apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset) done lemma vector_derivative_linepath_within: "x \ {0..1} \ vector_derivative (linepath a b) (at x within {0..1}) = b - a" by (simp add: has_vector_derivative_linepath_within vector_derivative_at_within_ivl) lemma vector_derivative_linepath_at [simp]: "vector_derivative (linepath a b) (at x) = b - a" by (simp add: has_vector_derivative_linepath_within vector_derivative_at) lemma valid_path_linepath [iff]: "valid_path (linepath a b)" using C1_differentiable_on_eq piecewise_C1_differentiable_on_def valid_path_def by fastforce lemma valid_path_subpath: fixes g :: "real \ 'a :: real_normed_vector" assumes "valid_path g" "u \ {0..1}" "v \ {0..1}" shows "valid_path(subpath u v g)" proof (cases "v=u") case True then show ?thesis unfolding valid_path_def subpath_def by (force intro: C1_differentiable_on_const C1_differentiable_imp_piecewise) next case False let ?f = "\x. ((v-u) * x + u)" have "(g \ ?f) piecewise_C1_differentiable_on {0..1}" proof (rule piecewise_C1_differentiable_compose) show "?f piecewise_C1_differentiable_on {0..1}" by (simp add: C1_differentiable_imp_piecewise) have "g piecewise_C1_differentiable_on (if u \ v then {u..v} else {v..u})" using assms piecewise_C1_differentiable_on_subset valid_path_def by force then show "g piecewise_C1_differentiable_on ?f ` {0..1}" by (simp add: image_affinity_atLeastAtMost split: if_split_asm) show "\x. finite ({0..1} \ ?f -` {x})" using False by (simp add: Int_commute [of "{0..1}"] inj_on_def crossproduct_eq finite_vimage_IntI) qed then show ?thesis by (auto simp: o_def valid_path_def subpath_def) qed lemma valid_path_rectpath [simp, intro]: "valid_path (rectpath a b)" by (simp add: Let_def rectpath_def) +lemma linear_image_valid_path: + fixes p :: "real \ 'a :: euclidean_space" + assumes "valid_path p" "linear f" + shows "valid_path (f \ p)" + unfolding valid_path_def piecewise_C1_differentiable_on_def +proof (intro conjI) + from assms have "path p" + by (simp add: valid_path_imp_path) + thus "continuous_on {0..1} (f \ p)" + unfolding o_def path_def by (intro linear_continuous_on_compose[OF _ assms(2)]) + from assms(1) obtain S where S: "finite S" "p C1_differentiable_on {0..1} - S" + by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) + from S(2) obtain p' :: "real \ 'a" + where p': "\x. x \ {0..1} - S \ (p has_vector_derivative p' x) (at x)" + "continuous_on ({0..1} - S) p'" + by (fastforce simp: C1_differentiable_on_def) + + have "(f \ p has_vector_derivative f (p' x)) (at x)" if "x \ {0..1} - S" for x + by (rule vector_derivative_diff_chain_within [OF p'(1)[OF that]] + linear_imp_has_derivative assms)+ + moreover have "continuous_on ({0..1} - S) (\x. f (p' x))" + by (rule linear_continuous_on_compose [OF p'(2) assms(2)]) + ultimately have "f \ p C1_differentiable_on {0..1} - S" + unfolding C1_differentiable_on_def by (intro exI[of _ "\x. f (p' x)"]) fast + thus "\S. finite S \ f \ p C1_differentiable_on {0..1} - S" + using \finite S\ by blast +qed + +lemma valid_path_times: + fixes \::"real \ 'a ::real_normed_field" + assumes "c\0" + shows "valid_path ((*) c \ \) = valid_path \" +proof + assume "valid_path ((*) c \ \)" + then have "valid_path ((*) (1/c) \ ((*) c \ \))" + by (simp add: valid_path_compose) + then show "valid_path \" + unfolding comp_def using \c\0\ by auto +next + assume "valid_path \" + then show "valid_path ((*) c \ \)" + by (simp add: valid_path_compose) +qed + +lemma path_compose_cnj_iff [simp]: "path (cnj \ p) \ path p" +proof - + have "path (cnj \ p)" if "path p" for p + by (intro path_continuous_image continuous_intros that) + from this[of p] and this[of "cnj \ p"] show ?thesis + by (auto simp: o_def) +qed + +lemma valid_path_cnj: + fixes g::"real \ complex" + shows "valid_path (cnj \ g) = valid_path g" +proof + show valid:"valid_path (cnj \ g)" if "valid_path g" for g + proof - + obtain S where "finite S" and g_diff: "g C1_differentiable_on {0..1} - S" + using \valid_path g\ unfolding valid_path_def piecewise_C1_differentiable_on_def by auto + + have g_diff':"g differentiable at t" when "t\{0..1} - S" for t + by (meson C1_differentiable_on_eq \g C1_differentiable_on {0..1} - S\ that) + then have "(cnj \ g) differentiable at t" when "t\{0..1} - S" for t + using bounded_linear_cnj bounded_linear_imp_differentiable differentiable_chain_at that by blast + moreover have "continuous_on ({0..1} - S) + (\x. vector_derivative (cnj \ g) (at x))" + proof - + have "continuous_on ({0..1} - S) + (\x. vector_derivative (cnj \ g) (at x)) + = continuous_on ({0..1} - S) + (\x. cnj (vector_derivative g (at x)))" + apply (rule continuous_on_cong[OF refl]) + unfolding comp_def using g_diff' + using has_vector_derivative_cnj vector_derivative_at vector_derivative_works by blast + also have "\" + apply (intro continuous_intros) + using C1_differentiable_on_eq g_diff by blast + finally show ?thesis . + qed + ultimately have "cnj \ g C1_differentiable_on {0..1} - S" + using C1_differentiable_on_eq by blast + moreover have "path (cnj \ g)" + apply (rule path_continuous_image[OF valid_path_imp_path[OF \valid_path g\]]) + by (intro continuous_intros) + ultimately show ?thesis unfolding valid_path_def piecewise_C1_differentiable_on_def path_def + using \finite S\ by auto + qed + from this[of "cnj o g"] + show "valid_path (cnj \ g) \ valid_path g" + unfolding comp_def by simp +qed + end diff --git a/src/HOL/Complex_Analysis/Contour_Integration.thy b/src/HOL/Complex_Analysis/Contour_Integration.thy --- a/src/HOL/Complex_Analysis/Contour_Integration.thy +++ b/src/HOL/Complex_Analysis/Contour_Integration.thy @@ -1,1757 +1,1830 @@ section \Contour integration\ theory Contour_Integration imports "HOL-Analysis.Analysis" begin lemma lhopital_complex_simple: assumes "(f has_field_derivative f') (at z)" assumes "(g has_field_derivative g') (at z)" assumes "f z = 0" "g z = 0" "g' \ 0" "f' / g' = c" shows "((\w. f w / g w) \ c) (at z)" proof - have "eventually (\w. w \ z) (at z)" by (auto simp: eventually_at_filter) hence "eventually (\w. ((f w - f z) / (w - z)) / ((g w - g z) / (w - z)) = f w / g w) (at z)" by eventually_elim (simp add: assms field_split_simps) moreover have "((\w. ((f w - f z) / (w - z)) / ((g w - g z) / (w - z))) \ f' / g') (at z)" by (intro tendsto_divide has_field_derivativeD assms) ultimately have "((\w. f w / g w) \ f' / g') (at z)" by (blast intro: Lim_transform_eventually) with assms show ?thesis by simp qed subsection\Definition\ text\ This definition is for complex numbers only, and does not generalise to line integrals in a vector field \ definition\<^marker>\tag important\ has_contour_integral :: "(complex \ complex) \ complex \ (real \ complex) \ bool" (infixr "has'_contour'_integral" 50) where "(f has_contour_integral i) g \ ((\x. f(g x) * vector_derivative g (at x within {0..1})) has_integral i) {0..1}" definition\<^marker>\tag important\ contour_integrable_on (infixr "contour'_integrable'_on" 50) where "f contour_integrable_on g \ \i. (f has_contour_integral i) g" definition\<^marker>\tag important\ contour_integral where "contour_integral g f \ SOME i. (f has_contour_integral i) g \ \ f contour_integrable_on g \ i=0" lemma not_integrable_contour_integral: "\ f contour_integrable_on g \ contour_integral g f = 0" unfolding contour_integrable_on_def contour_integral_def by blast lemma contour_integral_unique: "(f has_contour_integral i) g \ contour_integral g f = i" unfolding contour_integral_def has_contour_integral_def contour_integrable_on_def using has_integral_unique by blast lemma has_contour_integral_eqpath: "\(f has_contour_integral y) p; f contour_integrable_on \; contour_integral p f = contour_integral \ f\ \ (f has_contour_integral y) \" using contour_integrable_on_def contour_integral_unique by auto lemma has_contour_integral_integral: "f contour_integrable_on i \ (f has_contour_integral (contour_integral i f)) i" by (metis contour_integral_unique contour_integrable_on_def) lemma has_contour_integral_unique: "(f has_contour_integral i) g \ (f has_contour_integral j) g \ i = j" using has_integral_unique by (auto simp: has_contour_integral_def) lemma has_contour_integral_integrable: "(f has_contour_integral i) g \ f contour_integrable_on g" using contour_integrable_on_def by blast text\Show that we can forget about the localized derivative.\ lemma has_integral_localized_vector_derivative: "((\x. f (g x) * vector_derivative p (at x within {a..b})) has_integral i) {a..b} \ ((\x. f (g x) * vector_derivative p (at x)) has_integral i) {a..b}" proof - have *: "{a..b} - {a,b} = interior {a..b}" by (simp add: atLeastAtMost_diff_ends) show ?thesis by (rule has_integral_spike_eq [of "{a,b}"]) (auto simp: at_within_interior [of _ "{a..b}"]) qed lemma integrable_on_localized_vector_derivative: "(\x. f (g x) * vector_derivative p (at x within {a..b})) integrable_on {a..b} \ (\x. f (g x) * vector_derivative p (at x)) integrable_on {a..b}" by (simp add: integrable_on_def has_integral_localized_vector_derivative) lemma has_contour_integral: "(f has_contour_integral i) g \ ((\x. f (g x) * vector_derivative g (at x)) has_integral i) {0..1}" by (simp add: has_integral_localized_vector_derivative has_contour_integral_def) lemma contour_integrable_on: "f contour_integrable_on g \ (\t. f(g t) * vector_derivative g (at t)) integrable_on {0..1}" by (simp add: has_contour_integral integrable_on_def contour_integrable_on_def) lemma has_contour_integral_mirror_iff: assumes "valid_path g" shows "(f has_contour_integral I) (-g) \ ((\x. -f (- x)) has_contour_integral I) g" proof - from assms have "g piecewise_differentiable_on {0..1}" by (auto simp: valid_path_def piecewise_C1_imp_differentiable) then obtain S where "finite S" and S: "\x. x \ {0..1} - S \ g differentiable at x within {0..1}" unfolding piecewise_differentiable_on_def by blast have S': "g differentiable at x" if "x \ {0..1} - ({0, 1} \ S)" for x proof - from that have "x \ interior {0..1}" by auto with S[of x] that show ?thesis by (auto simp: at_within_interior[of _ "{0..1}"]) qed have "(f has_contour_integral I) (-g) \ ((\x. f (- g x) * vector_derivative (-g) (at x)) has_integral I) {0..1}" by (simp add: has_contour_integral) also have "\ \ ((\x. -f (- g x) * vector_derivative g (at x)) has_integral I) {0..1}" by (intro has_integral_spike_finite_eq[of "S \ {0, 1}"]) (insert \finite S\ S', auto simp: o_def fun_Compl_def) also have "\ \ ((\x. -f (-x)) has_contour_integral I) g" by (simp add: has_contour_integral) finally show ?thesis . qed lemma contour_integral_on_mirror_iff: assumes "valid_path g" shows "f contour_integrable_on (-g) \ (\x. -f (- x)) contour_integrable_on g" by (auto simp: contour_integrable_on_def has_contour_integral_mirror_iff assms) lemma contour_integral_mirror: assumes "valid_path g" shows "contour_integral (-g) f = contour_integral g (\x. -f (- x))" proof (cases "f contour_integrable_on (-g)") case True with contour_integral_unique assms show ?thesis by (auto simp: contour_integrable_on_def has_contour_integral_mirror_iff) next case False then show ?thesis by (simp add: assms contour_integral_on_mirror_iff not_integrable_contour_integral) qed subsection\<^marker>\tag unimportant\ \Reversing a path\ lemma has_contour_integral_reversepath: assumes "valid_path g" and f: "(f has_contour_integral i) g" shows "(f has_contour_integral (-i)) (reversepath g)" proof - { fix S x assume xs: "g C1_differentiable_on ({0..1} - S)" "x \ (-) 1 ` S" "0 \ x" "x \ 1" have "vector_derivative (\x. g (1 - x)) (at x within {0..1}) = - vector_derivative g (at (1 - x) within {0..1})" proof - obtain f' where f': "(g has_vector_derivative f') (at (1 - x))" using xs by (force simp: has_vector_derivative_def C1_differentiable_on_def) have "(g \ (\x. 1 - x) has_vector_derivative -1 *\<^sub>R f') (at x)" by (intro vector_diff_chain_within has_vector_derivative_at_within [OF f'] derivative_eq_intros | simp)+ then have mf': "((\x. g (1 - x)) has_vector_derivative -f') (at x)" by (simp add: o_def) show ?thesis using xs by (auto simp: vector_derivative_at_within_ivl [OF mf'] vector_derivative_at_within_ivl [OF f']) qed } note * = this obtain S where S: "continuous_on {0..1} g" "finite S" "g C1_differentiable_on {0..1} - S" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) have "((\x. - (f (g (1 - x)) * vector_derivative g (at (1 - x) within {0..1}))) has_integral -i) {0..1}" using has_integral_affinity01 [where m= "-1" and c=1, OF f [unfolded has_contour_integral_def]] by (simp add: has_integral_neg) then show ?thesis using S unfolding reversepath_def has_contour_integral_def by (rule_tac S = "(\x. 1 - x) ` S" in has_integral_spike_finite) (auto simp: *) qed lemma contour_integrable_reversepath: "valid_path g \ f contour_integrable_on g \ f contour_integrable_on (reversepath g)" using has_contour_integral_reversepath contour_integrable_on_def by blast lemma contour_integrable_reversepath_eq: "valid_path g \ (f contour_integrable_on (reversepath g) \ f contour_integrable_on g)" using contour_integrable_reversepath valid_path_reversepath by fastforce lemma contour_integral_reversepath: assumes "valid_path g" shows "contour_integral (reversepath g) f = - (contour_integral g f)" proof (cases "f contour_integrable_on g") case True then show ?thesis by (simp add: assms contour_integral_unique has_contour_integral_integral has_contour_integral_reversepath) next case False then have "\ f contour_integrable_on (reversepath g)" by (simp add: assms contour_integrable_reversepath_eq) with False show ?thesis by (simp add: not_integrable_contour_integral) qed subsection\<^marker>\tag unimportant\ \Joining two paths together\ lemma has_contour_integral_join: assumes "(f has_contour_integral i1) g1" "(f has_contour_integral i2) g2" "valid_path g1" "valid_path g2" shows "(f has_contour_integral (i1 + i2)) (g1 +++ g2)" proof - obtain s1 s2 where s1: "finite s1" "\x\{0..1} - s1. g1 differentiable at x" and s2: "finite s2" "\x\{0..1} - s2. g2 differentiable at x" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) have 1: "((\x. f (g1 x) * vector_derivative g1 (at x)) has_integral i1) {0..1}" and 2: "((\x. f (g2 x) * vector_derivative g2 (at x)) has_integral i2) {0..1}" using assms by (auto simp: has_contour_integral) have i1: "((\x. (2*f (g1 (2*x))) * vector_derivative g1 (at (2*x))) has_integral i1) {0..1/2}" and i2: "((\x. (2*f (g2 (2*x - 1))) * vector_derivative g2 (at (2*x - 1))) has_integral i2) {1/2..1}" using has_integral_affinity01 [OF 1, where m= 2 and c=0, THEN has_integral_cmul [where c=2]] has_integral_affinity01 [OF 2, where m= 2 and c="-1", THEN has_integral_cmul [where c=2]] by (simp_all only: image_affinity_atLeastAtMost_div_diff, simp_all add: scaleR_conv_of_real mult_ac) have g1: "vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at z) = 2 *\<^sub>R vector_derivative g1 (at (z*2))" if "0 \ z" "z*2 < 1" "z*2 \ s1" for z proof (rule vector_derivative_at [OF has_vector_derivative_transform_within]) show "0 < \z - 1/2\" using that by auto have "((*) 2 has_vector_derivative 2) (at z)" by (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) moreover have "(g1 has_vector_derivative vector_derivative g1 (at (z * 2))) (at (2 * z))" using s1 that by (auto simp: algebra_simps vector_derivative_works) ultimately show "((\x. g1 (2 * x)) has_vector_derivative 2 *\<^sub>R vector_derivative g1 (at (z * 2))) (at z)" by (intro vector_diff_chain_at [simplified o_def]) qed (use that in \simp_all add: dist_real_def abs_if split: if_split_asm\) have g2: "vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at z) = 2 *\<^sub>R vector_derivative g2 (at (z*2 - 1))" if "1 < z*2" "z \ 1" "z*2 - 1 \ s2" for z proof (rule vector_derivative_at [OF has_vector_derivative_transform_within]) show "0 < \z - 1/2\" using that by auto have "((\x. 2 * x - 1) has_vector_derivative 2) (at z)" by (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) moreover have "(g2 has_vector_derivative vector_derivative g2 (at (z * 2 - 1))) (at (2 * z - 1))" using s2 that by (auto simp: algebra_simps vector_derivative_works) ultimately show "((\x. g2 (2 * x - 1)) has_vector_derivative 2 *\<^sub>R vector_derivative g2 (at (z * 2 - 1))) (at z)" by (intro vector_diff_chain_at [simplified o_def]) qed (use that in \simp_all add: dist_real_def abs_if split: if_split_asm\) have "((\x. f ((g1 +++ g2) x) * vector_derivative (g1 +++ g2) (at x)) has_integral i1) {0..1/2}" proof (rule has_integral_spike_finite [OF _ _ i1]) show "finite (insert (1/2) ((*) 2 -` s1))" using s1 by (force intro: finite_vimageI [where h = "(*)2"] inj_onI) qed (auto simp add: joinpaths_def scaleR_conv_of_real mult_ac g1) moreover have "((\x. f ((g1 +++ g2) x) * vector_derivative (g1 +++ g2) (at x)) has_integral i2) {1/2..1}" proof (rule has_integral_spike_finite [OF _ _ i2]) show "finite (insert (1/2) ((\x. 2 * x - 1) -` s2))" using s2 by (force intro: finite_vimageI [where h = "\x. 2*x-1"] inj_onI) qed (auto simp add: joinpaths_def scaleR_conv_of_real mult_ac g2) ultimately show ?thesis by (simp add: has_contour_integral has_integral_combine [where c = "1/2"]) qed lemma contour_integrable_joinI: assumes "f contour_integrable_on g1" "f contour_integrable_on g2" "valid_path g1" "valid_path g2" shows "f contour_integrable_on (g1 +++ g2)" using assms by (meson has_contour_integral_join contour_integrable_on_def) lemma contour_integrable_joinD1: assumes "f contour_integrable_on (g1 +++ g2)" "valid_path g1" shows "f contour_integrable_on g1" proof - obtain s1 where s1: "finite s1" "\x\{0..1} - s1. g1 differentiable at x" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) have "(\x. f ((g1 +++ g2) (x/2)) * vector_derivative (g1 +++ g2) (at (x/2))) integrable_on {0..1}" using assms integrable_affinity [of _ 0 "1/2" "1/2" 0] integrable_on_subcbox [where a=0 and b="1/2"] by (fastforce simp: contour_integrable_on) then have *: "(\x. (f ((g1 +++ g2) (x/2))/2) * vector_derivative (g1 +++ g2) (at (x/2))) integrable_on {0..1}" by (auto dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real) have g1: "vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at (z/2)) = 2 *\<^sub>R vector_derivative g1 (at z)" if "0 < z" "z < 1" "z \ s1" for z proof (rule vector_derivative_at [OF has_vector_derivative_transform_within]) show "0 < \(z - 1)/2\" using that by auto have \
: "((\x. x * 2) has_vector_derivative 2) (at (z/2))" using s1 by (auto simp: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) have "(g1 has_vector_derivative vector_derivative g1 (at z)) (at z)" using s1 that by (auto simp: vector_derivative_works) then show "((\x. g1 (2 * x)) has_vector_derivative 2 *\<^sub>R vector_derivative g1 (at z)) (at (z/2))" using vector_diff_chain_at [OF \
] by (auto simp: field_simps o_def) qed (use that in \simp_all add: field_simps dist_real_def abs_if split: if_split_asm\) have fin01: "finite ({0, 1} \ s1)" by (simp add: s1) show ?thesis unfolding contour_integrable_on by (intro integrable_spike_finite [OF fin01 _ *]) (auto simp: joinpaths_def scaleR_conv_of_real g1) qed lemma contour_integrable_joinD2: assumes "f contour_integrable_on (g1 +++ g2)" "valid_path g2" shows "f contour_integrable_on g2" proof - obtain s2 where s2: "finite s2" "\x\{0..1} - s2. g2 differentiable at x" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) have "(\x. f ((g1 +++ g2) (x/2 + 1/2)) * vector_derivative (g1 +++ g2) (at (x/2 + 1/2))) integrable_on {0..1}" using assms integrable_affinity [of _ "1/2::real" 1 "1/2" "1/2"] integrable_on_subcbox [where a="1/2" and b=1] by (fastforce simp: contour_integrable_on image_affinity_atLeastAtMost_diff) then have *: "(\x. (f ((g1 +++ g2) (x/2 + 1/2))/2) * vector_derivative (g1 +++ g2) (at (x/2 + 1/2))) integrable_on {0..1}" by (auto dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real) have g2: "vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at (z/2+1/2)) = 2 *\<^sub>R vector_derivative g2 (at z)" if "0 < z" "z < 1" "z \ s2" for z proof (rule vector_derivative_at [OF has_vector_derivative_transform_within]) show "0 < \z/2\" using that by auto have \
: "((\x. x * 2 - 1) has_vector_derivative 2) (at ((1 + z)/2))" using s2 by (auto simp: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) have "(g2 has_vector_derivative vector_derivative g2 (at z)) (at z)" using s2 that by (auto simp: vector_derivative_works) then show "((\x. g2 (2*x - 1)) has_vector_derivative 2 *\<^sub>R vector_derivative g2 (at z)) (at (z/2 + 1/2))" using vector_diff_chain_at [OF \
] by (auto simp: field_simps o_def) qed (use that in \simp_all add: field_simps dist_real_def abs_if split: if_split_asm\) have fin01: "finite ({0, 1} \ s2)" by (simp add: s2) show ?thesis unfolding contour_integrable_on by (intro integrable_spike_finite [OF fin01 _ *]) (auto simp: joinpaths_def scaleR_conv_of_real g2) qed lemma contour_integrable_join [simp]: "\valid_path g1; valid_path g2\ \ f contour_integrable_on (g1 +++ g2) \ f contour_integrable_on g1 \ f contour_integrable_on g2" using contour_integrable_joinD1 contour_integrable_joinD2 contour_integrable_joinI by blast lemma contour_integral_join [simp]: "\f contour_integrable_on g1; f contour_integrable_on g2; valid_path g1; valid_path g2\ \ contour_integral (g1 +++ g2) f = contour_integral g1 f + contour_integral g2 f" by (simp add: has_contour_integral_integral has_contour_integral_join contour_integral_unique) subsection\<^marker>\tag unimportant\ \Shifting the starting point of a (closed) path\ lemma has_contour_integral_shiftpath: assumes f: "(f has_contour_integral i) g" "valid_path g" and a: "a \ {0..1}" shows "(f has_contour_integral i) (shiftpath a g)" proof - obtain S where S: "finite S" and g: "\x\{0..1} - S. g differentiable at x" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) have *: "((\x. f (g x) * vector_derivative g (at x)) has_integral i) {0..1}" using assms by (auto simp: has_contour_integral) then have i: "i = integral {a..1} (\x. f (g x) * vector_derivative g (at x)) + integral {0..a} (\x. f (g x) * vector_derivative g (at x))" by (smt (verit, ccfv_threshold) Henstock_Kurzweil_Integration.integral_combine a add.commute atLeastAtMost_iff has_integral_iff) have vd1: "vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a))" if "0 \ x" "x + a < 1" "x \ (\x. x - a) ` S" for x unfolding shiftpath_def proof (rule vector_derivative_at [OF has_vector_derivative_transform_within]) have "((\x. g (x + a)) has_vector_derivative vector_derivative g (at (a + x))) (at x)" proof (rule vector_diff_chain_at [of _ 1, simplified o_def scaleR_one]) show "((\x. x + a) has_vector_derivative 1) (at x)" by (rule derivative_eq_intros | simp)+ have "g differentiable at (x + a)" using g a that by force then show "(g has_vector_derivative vector_derivative g (at (a + x))) (at (x + a))" by (metis add.commute vector_derivative_works) qed then show "((\x. g (a + x)) has_vector_derivative vector_derivative g (at (x + a))) (at x)" by (auto simp: field_simps) show "0 < dist (1 - a) x" using that by auto qed (use that in \auto simp: dist_real_def\) have vd2: "vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a - 1))" if "x \ 1" "1 < x + a" "x \ (\x. x - a + 1) ` S" for x unfolding shiftpath_def proof (rule vector_derivative_at [OF has_vector_derivative_transform_within]) have "((\x. g (x + a - 1)) has_vector_derivative vector_derivative g (at (a+x-1))) (at x)" proof (rule vector_diff_chain_at [of _ 1, simplified o_def scaleR_one]) show "((\x. x + a - 1) has_vector_derivative 1) (at x)" by (rule derivative_eq_intros | simp)+ have "g differentiable at (x+a-1)" using g a that by force then show "(g has_vector_derivative vector_derivative g (at (a+x-1))) (at (x + a - 1))" by (metis add.commute vector_derivative_works) qed then show "((\x. g (a + x - 1)) has_vector_derivative vector_derivative g (at (x + a - 1))) (at x)" by (auto simp: field_simps) show "0 < dist (1 - a) x" using that by auto qed (use that in \auto simp: dist_real_def\) have va1: "(\x. f (g x) * vector_derivative g (at x)) integrable_on ({a..1})" using * a by (fastforce intro: integrable_subinterval_real) have v0a: "(\x. f (g x) * vector_derivative g (at x)) integrable_on ({0..a})" using * a by (force intro: integrable_subinterval_real) have "finite ({1 - a} \ (\x. x - a) ` S)" using S by blast then have "((\x. f (shiftpath a g x) * vector_derivative (shiftpath a g) (at x)) has_integral integral {a..1} (\x. f (g x) * vector_derivative g (at x))) {0..1 - a}" apply (rule has_integral_spike_finite [where f = "\x. f(g(a+x)) * vector_derivative g (at(a+x))"]) subgoal using a by (simp add: vd1) (force simp: shiftpath_def add.commute) subgoal using has_integral_affinity [where m=1 and c=a] integrable_integral [OF va1] by (force simp add: add.commute) done moreover have "finite ({1 - a} \ (\x. x - a + 1) ` S)" using S by blast then have "((\x. f (shiftpath a g x) * vector_derivative (shiftpath a g) (at x)) has_integral integral {0..a} (\x. f (g x) * vector_derivative g (at x))) {1 - a..1}" apply (rule has_integral_spike_finite [where f = "\x. f(g(a+x-1)) * vector_derivative g (at(a+x-1))"]) subgoal using a by (simp add: vd2) (force simp: shiftpath_def add.commute) subgoal using has_integral_affinity [where m=1 and c="a-1", simplified, OF integrable_integral [OF v0a]] by (force simp add: algebra_simps) done ultimately show ?thesis using a by (auto simp: i has_contour_integral intro: has_integral_combine [where c = "1-a"]) qed lemma has_contour_integral_shiftpath_D: assumes "(f has_contour_integral i) (shiftpath a g)" "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" shows "(f has_contour_integral i) g" proof - obtain S where S: "finite S" and g: "\x\{0..1} - S. g differentiable at x" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) { fix x assume x: "0 < x" "x < 1" "x \ S" then have gx: "g differentiable at x" using g by auto have \
: "shiftpath (1 - a) (shiftpath a g) differentiable at x" using assms x by (intro differentiable_transform_within [OF gx, of "min x (1-x)"]) (auto simp: dist_real_def shiftpath_shiftpath abs_if split: if_split_asm) have "vector_derivative g (at x within {0..1}) = vector_derivative (shiftpath (1 - a) (shiftpath a g)) (at x within {0..1})" apply (rule vector_derivative_at_within_ivl [OF has_vector_derivative_transform_within_open [where f = "(shiftpath (1 - a) (shiftpath a g))" and S = "{0<..<1}-S"]]) using S assms x \
apply (auto simp: finite_imp_closed open_Diff shiftpath_shiftpath at_within_interior [of _ "{0..1}"] vector_derivative_works [symmetric]) done } note vd = this have fi: "(f has_contour_integral i) (shiftpath (1 - a) (shiftpath a g))" using assms by (auto intro!: has_contour_integral_shiftpath) show ?thesis unfolding has_contour_integral_def proof (rule has_integral_spike_finite [of "{0,1} \ S", OF _ _ fi [unfolded has_contour_integral_def]]) show "finite ({0, 1} \ S)" by (simp add: S) qed (use S assms vd in \auto simp: shiftpath_shiftpath\) qed lemma has_contour_integral_shiftpath_eq: assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" shows "(f has_contour_integral i) (shiftpath a g) \ (f has_contour_integral i) g" using assms has_contour_integral_shiftpath has_contour_integral_shiftpath_D by blast lemma contour_integrable_on_shiftpath_eq: assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" shows "f contour_integrable_on (shiftpath a g) \ f contour_integrable_on g" using assms contour_integrable_on_def has_contour_integral_shiftpath_eq by auto lemma contour_integral_shiftpath: assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" shows "contour_integral (shiftpath a g) f = contour_integral g f" using assms by (simp add: contour_integral_def contour_integrable_on_def has_contour_integral_shiftpath_eq) subsection\<^marker>\tag unimportant\ \More about straight-line paths\ lemma has_contour_integral_linepath: shows "(f has_contour_integral i) (linepath a b) \ ((\x. f(linepath a b x) * (b - a)) has_integral i) {0..1}" by (simp add: has_contour_integral) lemma has_contour_integral_trivial [iff]: "(f has_contour_integral 0) (linepath a a)" by (simp add: has_contour_integral_linepath) lemma has_contour_integral_trivial_iff [simp]: "(f has_contour_integral i) (linepath a a) \ i=0" using has_contour_integral_unique by blast lemma contour_integral_trivial [simp]: "contour_integral (linepath a a) f = 0" using has_contour_integral_trivial contour_integral_unique by blast subsection\Relation to subpath construction\ lemma has_contour_integral_subpath_refl [iff]: "(f has_contour_integral 0) (subpath u u g)" by (simp add: has_contour_integral subpath_def) lemma contour_integrable_subpath_refl [iff]: "f contour_integrable_on (subpath u u g)" using has_contour_integral_subpath_refl contour_integrable_on_def by blast lemma contour_integral_subpath_refl [simp]: "contour_integral (subpath u u g) f = 0" by (simp add: contour_integral_unique) lemma has_contour_integral_subpath: assumes f: "f contour_integrable_on g" and g: "valid_path g" and uv: "u \ {0..1}" "v \ {0..1}" "u \ v" shows "(f has_contour_integral integral {u..v} (\x. f(g x) * vector_derivative g (at x))) (subpath u v g)" proof (cases "v=u") case True then show ?thesis using f by (simp add: contour_integrable_on_def subpath_def has_contour_integral) next case False obtain S where S: "\x. x \ {0..1} - S \ g differentiable at x" and fs: "finite S" using g unfolding piecewise_C1_differentiable_on_def C1_differentiable_on_eq valid_path_def by blast have \
: "(\t. f (g t) * vector_derivative g (at t)) integrable_on {u..v}" using contour_integrable_on f integrable_on_subinterval uv by fastforce then have *: "((\x. f (g ((v - u) * x + u)) * vector_derivative g (at ((v - u) * x + u))) has_integral (1 / (v - u)) * integral {u..v} (\t. f (g t) * vector_derivative g (at t))) {0..1}" using uv False unfolding has_integral_integral apply simp apply (drule has_integral_affinity [where m="v-u" and c=u, simplified]) apply (simp_all add: image_affinity_atLeastAtMost_div_diff scaleR_conv_of_real) apply (simp add: divide_simps) done have vd: "vector_derivative (\x. g ((v-u) * x + u)) (at x) = (v-u) *\<^sub>R vector_derivative g (at ((v-u) * x + u))" if "x \ {0..1}" "x \ (\t. (v-u) *\<^sub>R t + u) -` S" for x proof (rule vector_derivative_at [OF vector_diff_chain_at [simplified o_def]]) show "((\x. (v - u) * x + u) has_vector_derivative v - u) (at x)" by (intro derivative_eq_intros | simp)+ qed (use S uv mult_left_le [of x "v-u"] that in \auto simp: vector_derivative_works\) have fin: "finite ((\t. (v - u) *\<^sub>R t + u) -` S)" using fs by (auto simp: inj_on_def False finite_vimageI) show ?thesis unfolding subpath_def has_contour_integral apply (rule has_integral_spike_finite [OF fin]) using has_integral_cmul [OF *, where c = "v-u"] fs assms by (auto simp: False vd scaleR_conv_of_real) qed lemma contour_integrable_subpath: assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" shows "f contour_integrable_on (subpath u v g)" by (smt (verit, ccfv_threshold) assms contour_integrable_on_def contour_integrable_reversepath_eq has_contour_integral_subpath reversepath_subpath valid_path_subpath) lemma has_integral_contour_integral_subpath: assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "u \ v" shows "((\x. f(g x) * vector_derivative g (at x)) has_integral contour_integral (subpath u v g) f) {u..v}" (is "(?fg has_integral _)_") proof - have "(?fg has_integral integral {u..v} ?fg) {u..v}" using assms contour_integrable_on integrable_on_subinterval by fastforce then show ?thesis by (metis (full_types) assms contour_integral_unique has_contour_integral_subpath) qed lemma contour_integral_subcontour_integral: assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "u \ v" shows "contour_integral (subpath u v g) f = integral {u..v} (\x. f(g x) * vector_derivative g (at x))" using assms has_contour_integral_subpath contour_integral_unique by blast lemma contour_integral_subpath_combine_less: assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "w \ {0..1}" "ux. f (g x) * vector_derivative g (at x)) integrable_on {u..w}" using integrable_on_subcbox [where a=u and b=w and S = "{0..1}"] assms by (auto simp: contour_integrable_on) with assms show ?thesis by (auto simp: contour_integral_subcontour_integral Henstock_Kurzweil_Integration.integral_combine) qed lemma contour_integral_subpath_combine: assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "w \ {0..1}" shows "contour_integral (subpath u v g) f + contour_integral (subpath v w g) f = contour_integral (subpath u w g) f" proof (cases "u\v \ v\w \ u\w") case True have *: "subpath v u g = reversepath(subpath u v g) \ subpath w u g = reversepath(subpath u w g) \ subpath w v g = reversepath(subpath v w g)" by (auto simp: reversepath_subpath) have "u < v \ v < w \ u < w \ w < v \ v < u \ u < w \ v < w \ w < u \ w < u \ u < v \ w < v \ v < u" using True assms by linarith with assms show ?thesis using contour_integral_subpath_combine_less [of f g u v w] contour_integral_subpath_combine_less [of f g u w v] contour_integral_subpath_combine_less [of f g v u w] contour_integral_subpath_combine_less [of f g v w u] contour_integral_subpath_combine_less [of f g w u v] contour_integral_subpath_combine_less [of f g w v u] by (elim disjE) (auto simp: * contour_integral_reversepath contour_integrable_subpath valid_path_subpath algebra_simps) next case False with assms show ?thesis by (metis add.right_neutral contour_integral_reversepath contour_integral_subpath_refl diff_0 eq_diff_eq add_0 reversepath_subpath valid_path_subpath) qed lemma contour_integral_integral: "contour_integral g f = integral {0..1} (\x. f (g x) * vector_derivative g (at x))" by (simp add: contour_integral_def integral_def has_contour_integral contour_integrable_on) lemma contour_integral_cong: assumes "g = g'" "\x. x \ path_image g \ f x = f' x" shows "contour_integral g f = contour_integral g' f'" unfolding contour_integral_integral using assms by (intro integral_cong) (auto simp: path_image_def) lemma contour_integral_spike_finite_simple_path: assumes "finite A" "simple_path g" "g = g'" "\x. x \ path_image g - A \ f x = f' x" shows "contour_integral g f = contour_integral g' f'" unfolding contour_integral_integral proof (rule integral_spike) have "finite (g -` A \ {0<..<1})" using \simple_path g\ \finite A\ by (intro finite_vimage_IntI simple_path_inj_on) auto hence "finite ({0, 1} \ g -` A \ {0<..<1})" by auto thus "negligible ({0, 1} \ g -` A \ {0<..<1})" by (rule negligible_finite) next fix x assume "x \ {0..1} - ({0, 1} \ g -` A \ {0<..<1})" hence "g x \ path_image g - A" by (auto simp: path_image_def) with assms show "f' (g' x) * vector_derivative g' (at x) = f (g x) * vector_derivative g (at x)" by simp qed text \Contour integral along a segment on the real axis\ lemma has_contour_integral_linepath_Reals_iff: fixes a b :: complex and f :: "complex \ complex" assumes "a \ Reals" "b \ Reals" "Re a < Re b" shows "(f has_contour_integral I) (linepath a b) \ ((\x. f (of_real x)) has_integral I) {Re a..Re b}" proof - from assms have [simp]: "of_real (Re a) = a" "of_real (Re b) = b" by (simp_all add: complex_eq_iff) from assms have "a \ b" by auto have "((\x. f (of_real x)) has_integral I) (cbox (Re a) (Re b)) \ ((\x. f (a + b * of_real x - a * of_real x)) has_integral I /\<^sub>R (Re b - Re a)) {0..1}" by (subst has_integral_affinity_iff [of "Re b - Re a" _ "Re a", symmetric]) (insert assms, simp_all add: field_simps scaleR_conv_of_real) also have "(\x. f (a + b * of_real x - a * of_real x)) = (\x. (f (a + b * of_real x - a * of_real x) * (b - a)) /\<^sub>R (Re b - Re a))" using \a \ b\ by (auto simp: field_simps fun_eq_iff scaleR_conv_of_real) also have "(\ has_integral I /\<^sub>R (Re b - Re a)) {0..1} \ ((\x. f (linepath a b x) * (b - a)) has_integral I) {0..1}" using assms by (subst has_integral_cmul_iff) (auto simp: linepath_def scaleR_conv_of_real algebra_simps) also have "\ \ (f has_contour_integral I) (linepath a b)" unfolding has_contour_integral_def by (intro has_integral_cong) (simp add: vector_derivative_linepath_within) finally show ?thesis by simp qed lemma contour_integrable_linepath_Reals_iff: fixes a b :: complex and f :: "complex \ complex" assumes "a \ Reals" "b \ Reals" "Re a < Re b" shows "(f contour_integrable_on linepath a b) \ (\x. f (of_real x)) integrable_on {Re a..Re b}" using has_contour_integral_linepath_Reals_iff[OF assms, of f] by (auto simp: contour_integrable_on_def integrable_on_def) lemma contour_integral_linepath_Reals_eq: fixes a b :: complex and f :: "complex \ complex" assumes "a \ Reals" "b \ Reals" "Re a < Re b" shows "contour_integral (linepath a b) f = integral {Re a..Re b} (\x. f (of_real x))" proof (cases "f contour_integrable_on linepath a b") case True thus ?thesis using has_contour_integral_linepath_Reals_iff[OF assms, of f] using has_contour_integral_integral has_contour_integral_unique by blast next case False thus ?thesis using contour_integrable_linepath_Reals_iff[OF assms, of f] by (simp add: not_integrable_contour_integral not_integrable_integral) qed subsection \Cauchy's theorem where there's a primitive\ lemma contour_integral_primitive_lemma: fixes f :: "complex \ complex" and g :: "real \ complex" assumes "a \ b" and "\x. x \ S \ (f has_field_derivative f' x) (at x within S)" and "g piecewise_differentiable_on {a..b}" "\x. x \ {a..b} \ g x \ S" shows "((\x. f'(g x) * vector_derivative g (at x within {a..b})) has_integral (f(g b) - f(g a))) {a..b}" proof - obtain K where "finite K" and K: "\x\{a..b} - K. g differentiable (at x within {a..b})" and cg: "continuous_on {a..b} g" using assms by (auto simp: piecewise_differentiable_on_def) have "continuous_on (g ` {a..b}) f" using assms by (metis field_differentiable_def field_differentiable_imp_continuous_at continuous_on_eq_continuous_within continuous_on_subset image_subset_iff) then have cfg: "continuous_on {a..b} (\x. f (g x))" by (rule continuous_on_compose [OF cg, unfolded o_def]) { fix x::real assume a: "a < x" and b: "x < b" and xk: "x \ K" then have "g differentiable at x within {a..b}" using K by (simp add: differentiable_at_withinI) then have "(g has_vector_derivative vector_derivative g (at x within {a..b})) (at x within {a..b})" by (simp add: vector_derivative_works has_field_derivative_def scaleR_conv_of_real) then have gdiff: "(g has_derivative (\u. u * vector_derivative g (at x within {a..b}))) (at x within {a..b})" by (simp add: has_vector_derivative_def scaleR_conv_of_real) have "(f has_field_derivative (f' (g x))) (at (g x) within g ` {a..b})" using assms by (metis a atLeastAtMost_iff b DERIV_subset image_subset_iff less_eq_real_def) then have fdiff: "(f has_derivative (*) (f' (g x))) (at (g x) within g ` {a..b})" by (simp add: has_field_derivative_def) have "((\x. f (g x)) has_vector_derivative f' (g x) * vector_derivative g (at x within {a..b})) (at x within {a..b})" using diff_chain_within [OF gdiff fdiff] by (simp add: has_vector_derivative_def scaleR_conv_of_real o_def mult_ac) } then show ?thesis using assms cfg by (force simp: at_within_Icc_at intro: fundamental_theorem_of_calculus_interior_strong [OF \finite K\]) qed lemma contour_integral_primitive: assumes "\x. x \ S \ (f has_field_derivative f' x) (at x within S)" and "valid_path g" "path_image g \ S" shows "(f' has_contour_integral (f(pathfinish g) - f(pathstart g))) g" using assms apply (simp add: valid_path_def path_image_def pathfinish_def pathstart_def has_contour_integral_def) apply (auto intro!: piecewise_C1_imp_differentiable contour_integral_primitive_lemma [of 0 1 S]) done corollary Cauchy_theorem_primitive: assumes "\x. x \ S \ (f has_field_derivative f' x) (at x within S)" and "valid_path g" "path_image g \ S" "pathfinish g = pathstart g" shows "(f' has_contour_integral 0) g" using assms by (metis diff_self contour_integral_primitive) lemma contour_integrable_continuous_linepath: assumes "continuous_on (closed_segment a b) f" shows "f contour_integrable_on (linepath a b)" proof - have "continuous_on (closed_segment a b) (\x. f x * (b - a))" by (rule continuous_intros | simp add: assms)+ then have "continuous_on {0..1} (\x. f (linepath a b x) * (b - a))" by (metis (no_types, lifting) continuous_on_compose continuous_on_cong continuous_on_linepath linepath_image_01 o_apply) then have "(\x. f (linepath a b x) * vector_derivative (linepath a b) (at x within {0..1})) integrable_on {0..1}" by (metis (no_types, lifting) continuous_on_cong integrable_continuous_real vector_derivative_linepath_within) then show ?thesis by (simp add: contour_integrable_on_def has_contour_integral_def integrable_on_def [symmetric]) qed lemma has_field_der_id: "((\x. x\<^sup>2/2) has_field_derivative x) (at x)" by (rule has_derivative_imp_has_field_derivative) (rule derivative_intros | simp)+ lemma contour_integral_id [simp]: "contour_integral (linepath a b) (\y. y) = (b^2 - a^2)/2" using contour_integral_primitive [of UNIV "\x. x^2/2" "\x. x" "linepath a b"] contour_integral_unique by (simp add: has_field_der_id) lemma contour_integrable_on_const [iff]: "(\x. c) contour_integrable_on (linepath a b)" by (simp add: contour_integrable_continuous_linepath) lemma contour_integrable_on_id [iff]: "(\x. x) contour_integrable_on (linepath a b)" by (simp add: contour_integrable_continuous_linepath) subsection\<^marker>\tag unimportant\ \Arithmetical combining theorems\ lemma has_contour_integral_neg: "(f has_contour_integral i) g \ ((\x. -(f x)) has_contour_integral (-i)) g" by (simp add: has_integral_neg has_contour_integral_def) lemma has_contour_integral_add: "\(f1 has_contour_integral i1) g; (f2 has_contour_integral i2) g\ \ ((\x. f1 x + f2 x) has_contour_integral (i1 + i2)) g" by (simp add: has_integral_add has_contour_integral_def algebra_simps) lemma has_contour_integral_diff: "\(f1 has_contour_integral i1) g; (f2 has_contour_integral i2) g\ \ ((\x. f1 x - f2 x) has_contour_integral (i1 - i2)) g" by (simp add: has_integral_diff has_contour_integral_def algebra_simps) lemma has_contour_integral_lmul: "(f has_contour_integral i) g \ ((\x. c * (f x)) has_contour_integral (c*i)) g" by (simp add: has_contour_integral_def algebra_simps has_integral_mult_right) lemma has_contour_integral_rmul: "(f has_contour_integral i) g \ ((\x. (f x) * c) has_contour_integral (i*c)) g" by (simp add: mult.commute has_contour_integral_lmul) lemma has_contour_integral_div: "(f has_contour_integral i) g \ ((\x. f x/c) has_contour_integral (i/c)) g" by (simp add: field_class.field_divide_inverse) (metis has_contour_integral_rmul) lemma has_contour_integral_eq: "\(f has_contour_integral y) p; \x. x \ path_image p \ f x = g x\ \ (g has_contour_integral y) p" by (metis (mono_tags, lifting) has_contour_integral_def has_integral_eq image_eqI path_image_def) lemma has_contour_integral_bound_linepath: assumes "(f has_contour_integral i) (linepath a b)" "0 \ B" and B: "\x. x \ closed_segment a b \ norm(f x) \ B" shows "norm i \ B * norm(b - a)" proof - have "norm i \ (B * norm (b - a)) * content (cbox 0 (1::real))" proof (rule has_integral_bound [of _ "\x. f (linepath a b x) * vector_derivative (linepath a b) (at x within {0..1})"]) show "cmod (f (linepath a b x) * vector_derivative (linepath a b) (at x within {0..1})) \ B * cmod (b - a)" if "x \ cbox 0 1" for x::real using that box_real(2) norm_mult by (metis B linepath_in_path mult_right_mono norm_ge_zero vector_derivative_linepath_within) qed (use assms has_contour_integral_def in auto) then show ?thesis by (auto simp: content_real) qed lemma has_contour_integral_const_linepath: "((\x. c) has_contour_integral c*(b - a))(linepath a b)" unfolding has_contour_integral_linepath by (metis content_real diff_0_right has_integral_const_real lambda_one of_real_1 scaleR_conv_of_real zero_le_one) lemma has_contour_integral_0: "((\x. 0) has_contour_integral 0) g" by (simp add: has_contour_integral_def) lemma has_contour_integral_is_0: "(\z. z \ path_image g \ f z = 0) \ (f has_contour_integral 0) g" by (rule has_contour_integral_eq [OF has_contour_integral_0]) auto lemma has_contour_integral_sum: "\finite s; \a. a \ s \ (f a has_contour_integral i a) p\ \ ((\x. sum (\a. f a x) s) has_contour_integral sum i s) p" by (induction s rule: finite_induct) (auto simp: has_contour_integral_0 has_contour_integral_add) subsection\<^marker>\tag unimportant\ \Operations on path integrals\ lemma contour_integral_const_linepath [simp]: "contour_integral (linepath a b) (\x. c) = c*(b - a)" by (rule contour_integral_unique [OF has_contour_integral_const_linepath]) lemma contour_integral_neg: "contour_integral g (\z. -f z) = -contour_integral g f" by (simp add: contour_integral_integral) lemma contour_integral_add: "f1 contour_integrable_on g \ f2 contour_integrable_on g \ contour_integral g (\x. f1 x + f2 x) = contour_integral g f1 + contour_integral g f2" by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_add) lemma contour_integral_diff: "f1 contour_integrable_on g \ f2 contour_integrable_on g \ contour_integral g (\x. f1 x - f2 x) = contour_integral g f1 - contour_integral g f2" by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_diff) lemma contour_integral_lmul: shows "f contour_integrable_on g \ contour_integral g (\x. c * f x) = c*contour_integral g f" by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_lmul) lemma contour_integral_rmul: shows "f contour_integrable_on g \ contour_integral g (\x. f x * c) = contour_integral g f * c" by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_rmul) lemma contour_integral_div: shows "f contour_integrable_on g \ contour_integral g (\x. f x / c) = contour_integral g f / c" by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_div) lemma contour_integral_eq: "(\x. x \ path_image p \ f x = g x) \ contour_integral p f = contour_integral p g" using contour_integral_cong contour_integral_def by fastforce lemma contour_integral_eq_0: "(\z. z \ path_image g \ f z = 0) \ contour_integral g f = 0" by (simp add: has_contour_integral_is_0 contour_integral_unique) lemma contour_integral_bound_linepath: shows "\f contour_integrable_on (linepath a b); 0 \ B; \x. x \ closed_segment a b \ norm(f x) \ B\ \ norm(contour_integral (linepath a b) f) \ B*norm(b - a)" by (meson has_contour_integral_bound_linepath has_contour_integral_integral) lemma contour_integral_0 [simp]: "contour_integral g (\x. 0) = 0" by (simp add: contour_integral_unique has_contour_integral_0) lemma contour_integral_sum: "\finite s; \a. a \ s \ (f a) contour_integrable_on p\ \ contour_integral p (\x. sum (\a. f a x) s) = sum (\a. contour_integral p (f a)) s" by (auto simp: contour_integral_unique has_contour_integral_sum has_contour_integral_integral) lemma contour_integrable_eq: "\f contour_integrable_on p; \x. x \ path_image p \ f x = g x\ \ g contour_integrable_on p" unfolding contour_integrable_on_def by (metis has_contour_integral_eq) subsection\<^marker>\tag unimportant\ \Arithmetic theorems for path integrability\ lemma contour_integrable_neg: "f contour_integrable_on g \ (\x. -(f x)) contour_integrable_on g" using has_contour_integral_neg contour_integrable_on_def by blast lemma contour_integrable_add: "\f1 contour_integrable_on g; f2 contour_integrable_on g\ \ (\x. f1 x + f2 x) contour_integrable_on g" using has_contour_integral_add contour_integrable_on_def by fastforce lemma contour_integrable_diff: "\f1 contour_integrable_on g; f2 contour_integrable_on g\ \ (\x. f1 x - f2 x) contour_integrable_on g" using has_contour_integral_diff contour_integrable_on_def by fastforce lemma contour_integrable_lmul: "f contour_integrable_on g \ (\x. c * f x) contour_integrable_on g" using has_contour_integral_lmul contour_integrable_on_def by fastforce lemma contour_integrable_rmul: "f contour_integrable_on g \ (\x. f x * c) contour_integrable_on g" using has_contour_integral_rmul contour_integrable_on_def by fastforce lemma contour_integrable_div: "f contour_integrable_on g \ (\x. f x / c) contour_integrable_on g" using has_contour_integral_div contour_integrable_on_def by fastforce lemma contour_integrable_sum: "\finite s; \a. a \ s \ (f a) contour_integrable_on p\ \ (\x. sum (\a. f a x) s) contour_integrable_on p" unfolding contour_integrable_on_def by (metis has_contour_integral_sum) lemma contour_integrable_neg_iff: "(\x. -f x) contour_integrable_on g \ f contour_integrable_on g" using contour_integrable_neg[of f g] contour_integrable_neg[of "\x. -f x" g] by auto lemma contour_integrable_lmul_iff: "c \ 0 \ (\x. c * f x) contour_integrable_on g \ f contour_integrable_on g" using contour_integrable_lmul[of f g c] contour_integrable_lmul[of "\x. c * f x" g "inverse c"] by (auto simp: field_simps) lemma contour_integrable_rmul_iff: "c \ 0 \ (\x. f x * c) contour_integrable_on g \ f contour_integrable_on g" using contour_integrable_rmul[of f g c] contour_integrable_rmul[of "\x. c * f x" g "inverse c"] by (auto simp: field_simps) lemma contour_integrable_div_iff: "c \ 0 \ (\x. f x / c) contour_integrable_on g \ f contour_integrable_on g" using contour_integrable_rmul_iff[of "inverse c"] by (simp add: field_simps) subsection\<^marker>\tag unimportant\ \Reversing a path integral\ lemma has_contour_integral_reverse_linepath: "(f has_contour_integral i) (linepath a b) \ (f has_contour_integral (-i)) (linepath b a)" using has_contour_integral_reversepath valid_path_linepath by fastforce lemma contour_integral_reverse_linepath: "continuous_on (closed_segment a b) f \ contour_integral (linepath a b) f = - (contour_integral(linepath b a) f)" using contour_integral_reversepath by fastforce text \Splitting a path integral in a flat way.*)\ lemma has_contour_integral_split: assumes f: "(f has_contour_integral i) (linepath a c)" "(f has_contour_integral j) (linepath c b)" and k: "0 \ k" "k \ 1" and c: "c - a = k *\<^sub>R (b - a)" shows "(f has_contour_integral (i + j)) (linepath a b)" proof (cases "k = 0 \ k = 1") case True then show ?thesis using assms by auto next case False then have k: "0 < k" "k < 1" "complex_of_real k \ 1" using assms by auto have c': "c = k *\<^sub>R (b - a) + a" by (metis diff_add_cancel c) have bc: "(b - c) = (1 - k) *\<^sub>R (b - a)" by (simp add: algebra_simps c') { assume *: "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R c) * (c - a)) has_integral i) {0..1}" have "\x. (x / k) *\<^sub>R a + ((k - x) / k) *\<^sub>R a = a" using False by (simp add: field_split_simps flip: real_vector.scale_left_distrib) then have "\x. ((k - x) / k) *\<^sub>R a + (x / k) *\<^sub>R c = (1 - x) *\<^sub>R a + x *\<^sub>R b" using False by (simp add: c' algebra_simps) then have "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral i) {0..k}" using k has_integral_affinity01 [OF *, of "inverse k" "0"] by (force dest: has_integral_cmul [where c = "inverse k"] simp add: divide_simps mult.commute [of _ "k"] image_affinity_atLeastAtMost c) } note fi = this { assume *: "((\x. f ((1 - x) *\<^sub>R c + x *\<^sub>R b) * (b - c)) has_integral j) {0..1}" have **: "\x. (((1 - x) / (1 - k)) *\<^sub>R c + ((x - k) / (1 - k)) *\<^sub>R b) = ((1 - x) *\<^sub>R a + x *\<^sub>R b)" using k unfolding c' scaleR_conv_of_real apply (simp add: divide_simps) apply (simp add: distrib_right distrib_left right_diff_distrib left_diff_distrib) done have "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral j) {k..1}" using k has_integral_affinity01 [OF *, of "inverse(1 - k)" "-(k/(1 - k))"] apply (simp add: divide_simps mult.commute [of _ "1-k"] image_affinity_atLeastAtMost ** bc) apply (auto dest: has_integral_cmul [where k = "(1 - k) *\<^sub>R j" and c = "inverse (1 - k)"]) done } then show ?thesis using f k unfolding has_contour_integral_linepath by (simp add: linepath_def has_integral_combine [OF _ _ fi]) qed lemma continuous_on_closed_segment_transform: assumes f: "continuous_on (closed_segment a b) f" and k: "0 \ k" "k \ 1" and c: "c - a = k *\<^sub>R (b - a)" shows "continuous_on (closed_segment a c) f" proof - have c': "c = (1 - k) *\<^sub>R a + k *\<^sub>R b" using c by (simp add: algebra_simps) have "closed_segment a c \ closed_segment a b" by (metis c' ends_in_segment(1) in_segment(1) k subset_closed_segment) then show "continuous_on (closed_segment a c) f" by (rule continuous_on_subset [OF f]) qed lemma contour_integral_split: assumes f: "continuous_on (closed_segment a b) f" and k: "0 \ k" "k \ 1" and c: "c - a = k *\<^sub>R (b - a)" shows "contour_integral(linepath a b) f = contour_integral(linepath a c) f + contour_integral(linepath c b) f" proof - have c': "c = (1 - k) *\<^sub>R a + k *\<^sub>R b" using c by (simp add: algebra_simps) have "closed_segment a c \ closed_segment a b" by (metis c' ends_in_segment(1) in_segment(1) k subset_closed_segment) moreover have "closed_segment c b \ closed_segment a b" by (metis c' ends_in_segment(2) in_segment(1) k subset_closed_segment) ultimately have *: "continuous_on (closed_segment a c) f" "continuous_on (closed_segment c b) f" by (auto intro: continuous_on_subset [OF f]) show ?thesis by (rule contour_integral_unique) (meson "*" c contour_integrable_continuous_linepath has_contour_integral_integral has_contour_integral_split k) qed lemma contour_integral_split_linepath: assumes f: "continuous_on (closed_segment a b) f" and c: "c \ closed_segment a b" shows "contour_integral(linepath a b) f = contour_integral(linepath a c) f + contour_integral(linepath c b) f" using c by (auto simp: closed_segment_def algebra_simps intro!: contour_integral_split [OF f]) subsection\Reversing the order in a double path integral\ text\The condition is stronger than needed but it's often true in typical situations\ lemma fst_im_cbox [simp]: "cbox c d \ {} \ (fst ` cbox (a,c) (b,d)) = cbox a b" by (auto simp: cbox_Pair_eq) lemma snd_im_cbox [simp]: "cbox a b \ {} \ (snd ` cbox (a,c) (b,d)) = cbox c d" by (auto simp: cbox_Pair_eq) proposition contour_integral_swap: assumes fcon: "continuous_on (path_image g \ path_image h) (\(y1,y2). f y1 y2)" and vp: "valid_path g" "valid_path h" and gvcon: "continuous_on {0..1} (\t. vector_derivative g (at t))" and hvcon: "continuous_on {0..1} (\t. vector_derivative h (at t))" shows "contour_integral g (\w. contour_integral h (f w)) = contour_integral h (\z. contour_integral g (\w. f w z))" proof - have gcon: "continuous_on {0..1} g" and hcon: "continuous_on {0..1} h" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) have fgh1: "\x. (\t. f (g x) (h t)) = (\(y1,y2). f y1 y2) \ (\t. (g x, h t))" by (rule ext) simp have fgh2: "\x. (\t. f (g t) (h x)) = (\(y1,y2). f y1 y2) \ (\t. (g t, h x))" by (rule ext) simp have fcon_im1: "\x. 0 \ x \ x \ 1 \ continuous_on ((\t. (g x, h t)) ` {0..1}) (\(x, y). f x y)" by (rule continuous_on_subset [OF fcon]) (auto simp: path_image_def) have fcon_im2: "\x. 0 \ x \ x \ 1 \ continuous_on ((\t. (g t, h x)) ` {0..1}) (\(x, y). f x y)" by (rule continuous_on_subset [OF fcon]) (auto simp: path_image_def) have "continuous_on (cbox (0, 0) (1, 1::real)) ((\x. vector_derivative g (at x)) \ fst)" "continuous_on (cbox (0, 0) (1::real, 1)) ((\x. vector_derivative h (at x)) \ snd)" by (rule continuous_intros | simp add: gvcon hvcon)+ then have gvcon': "continuous_on (cbox (0, 0) (1, 1::real)) (\z. vector_derivative g (at (fst z)))" and hvcon': "continuous_on (cbox (0, 0) (1::real, 1)) (\x. vector_derivative h (at (snd x)))" by auto have "continuous_on ((\x. (g (fst x), h (snd x))) ` cbox (0,0) (1,1)) (\(y1, y2). f y1 y2)" by (auto simp: path_image_def intro: continuous_on_subset [OF fcon]) then have "continuous_on (cbox (0, 0) (1, 1)) ((\(y1, y2). f y1 y2) \ (\w. ((g \ fst) w, (h \ snd) w)))" by (intro gcon hcon continuous_intros | simp)+ then have fgh: "continuous_on (cbox (0, 0) (1, 1)) (\x. f (g (fst x)) (h (snd x)))" by auto have "integral {0..1} (\x. contour_integral h (f (g x)) * vector_derivative g (at x)) = integral {0..1} (\x. contour_integral h (\y. f (g x) y * vector_derivative g (at x)))" proof (rule integral_cong [OF contour_integral_rmul [symmetric]]) have "\x. x \ {0..1} \ continuous_on {0..1} (\xa. f (g x) (h xa))" by (subst fgh1) (rule fcon_im1 hcon continuous_intros | simp)+ then show "\x. x \ {0..1} \ f (g x) contour_integrable_on h" unfolding contour_integrable_on using continuous_on_mult hvcon integrable_continuous_real by blast qed also have "\ = integral {0..1} (\y. contour_integral g (\x. f x (h y) * vector_derivative h (at y)))" unfolding contour_integral_integral apply (subst integral_swap_continuous [where 'a = real and 'b = real, of 0 0 1 1, simplified]) subgoal by (rule fgh gvcon' hvcon' continuous_intros | simp add: split_def)+ subgoal unfolding integral_mult_left [symmetric] by (simp only: mult_ac) done also have "\ = contour_integral h (\z. contour_integral g (\w. f w z))" unfolding contour_integral_integral integral_mult_left [symmetric] by (simp add: algebra_simps) finally show ?thesis by (simp add: contour_integral_integral) qed lemma valid_path_negatepath: "valid_path \ \ valid_path (uminus \ \)" unfolding o_def using piecewise_C1_differentiable_neg valid_path_def by blast lemma has_contour_integral_negatepath: assumes \: "valid_path \" and cint: "((\z. f (- z)) has_contour_integral - i) \" shows "(f has_contour_integral i) (uminus \ \)" proof - obtain S where cont: "continuous_on {0..1} \" and "finite S" and diff: "\ C1_differentiable_on {0..1} - S" using \ by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) have "((\x. - (f (- \ x) * vector_derivative \ (at x within {0..1}))) has_integral i) {0..1}" using cint by (auto simp: has_contour_integral_def dest: has_integral_neg) then have "((\x. f (- \ x) * vector_derivative (uminus \ \) (at x within {0..1})) has_integral i) {0..1}" proof (rule rev_iffD1 [OF _ has_integral_spike_eq]) show "negligible S" by (simp add: \finite S\ negligible_finite) show "f (- \ x) * vector_derivative (uminus \ \) (at x within {0..1}) = - (f (- \ x) * vector_derivative \ (at x within {0..1}))" if "x \ {0..1} - S" for x proof - have "vector_derivative (uminus \ \) (at x within cbox 0 1) = - vector_derivative \ (at x within cbox 0 1)" proof (rule vector_derivative_within_cbox) show "(uminus \ \ has_vector_derivative - vector_derivative \ (at x within cbox 0 1)) (at x within cbox 0 1)" using that unfolding o_def by (metis C1_differentiable_on_eq UNIV_I diff differentiable_subset has_vector_derivative_minus subsetI that vector_derivative_works) qed (use that in auto) then show ?thesis by simp qed qed then show ?thesis by (simp add: has_contour_integral_def) qed lemma contour_integrable_negatepath: assumes \: "valid_path \" and pi: "(\z. f (- z)) contour_integrable_on \" shows "f contour_integrable_on (uminus \ \)" by (metis \ add.inverse_inverse contour_integrable_on_def has_contour_integral_negatepath pi) lemma C1_differentiable_polynomial_function: fixes p :: "real \ 'a::euclidean_space" shows "polynomial_function p \ p C1_differentiable_on S" by (metis continuous_on_polymonial_function C1_differentiable_on_def has_vector_derivative_polynomial_function) lemma valid_path_polynomial_function: fixes p :: "real \ 'a::euclidean_space" shows "polynomial_function p \ valid_path p" by (force simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_polymonial_function C1_differentiable_polynomial_function) lemma valid_path_subpath_trivial [simp]: fixes g :: "real \ 'a::euclidean_space" shows "z \ g x \ valid_path (subpath x x g)" by (simp add: subpath_def valid_path_polynomial_function) subsection\Partial circle path\ definition\<^marker>\tag important\ part_circlepath :: "[complex, real, real, real, real] \ complex" where "part_circlepath z r s t \ \x. z + of_real r * exp (\ * of_real (linepath s t x))" lemma pathstart_part_circlepath [simp]: "pathstart(part_circlepath z r s t) = z + r*exp(\ * s)" by (metis part_circlepath_def pathstart_def pathstart_linepath) lemma pathfinish_part_circlepath [simp]: "pathfinish(part_circlepath z r s t) = z + r*exp(\*t)" by (metis part_circlepath_def pathfinish_def pathfinish_linepath) lemma reversepath_part_circlepath[simp]: "reversepath (part_circlepath z r s t) = part_circlepath z r t s" unfolding part_circlepath_def reversepath_def linepath_def by (auto simp:algebra_simps) lemma has_vector_derivative_part_circlepath [derivative_intros]: "((part_circlepath z r s t) has_vector_derivative (\ * r * (of_real t - of_real s) * exp(\ * linepath s t x))) (at x within X)" unfolding part_circlepath_def linepath_def scaleR_conv_of_real by (rule has_vector_derivative_real_field derivative_eq_intros | simp)+ lemma differentiable_part_circlepath: "part_circlepath c r a b differentiable at x within A" using has_vector_derivative_part_circlepath[of c r a b x A] differentiableI_vector by blast lemma vector_derivative_part_circlepath: "vector_derivative (part_circlepath z r s t) (at x) = \ * r * (of_real t - of_real s) * exp(\ * linepath s t x)" using has_vector_derivative_part_circlepath vector_derivative_at by blast lemma vector_derivative_part_circlepath01: "\0 \ x; x \ 1\ \ vector_derivative (part_circlepath z r s t) (at x within {0..1}) = \ * r * (of_real t - of_real s) * exp(\ * linepath s t x)" using has_vector_derivative_part_circlepath by (auto simp: vector_derivative_at_within_ivl) lemma valid_path_part_circlepath [simp]: "valid_path (part_circlepath z r s t)" unfolding valid_path_def by (auto simp: C1_differentiable_on_eq vector_derivative_works vector_derivative_part_circlepath has_vector_derivative_part_circlepath intro!: C1_differentiable_imp_piecewise continuous_intros) lemma path_part_circlepath [simp]: "path (part_circlepath z r s t)" by (simp add: valid_path_imp_path) proposition path_image_part_circlepath: assumes "s \ t" shows "path_image (part_circlepath z r s t) = {z + r * exp(\ * of_real x) | x. s \ x \ x \ t}" proof - { fix z::real assume "0 \ z" "z \ 1" with \s \ t\ have "\x. (exp (\ * linepath s t z) = exp (\ * of_real x)) \ s \ x \ x \ t" apply (rule_tac x="(1 - z) * s + z * t" in exI) apply (simp add: linepath_def scaleR_conv_of_real algebra_simps) by (metis (no_types) affine_ineq mult.commute mult_left_mono) } moreover { fix z assume "s \ z" "z \ t" then have "z + of_real r * exp (\ * of_real z) \ (\x. z + of_real r * exp (\ * linepath s t x)) ` {0..1}" apply (rule_tac x="(z - s)/(t - s)" in image_eqI) apply (simp add: linepath_def scaleR_conv_of_real divide_simps exp_eq) apply (auto simp: field_split_simps) done } ultimately show ?thesis by (fastforce simp add: path_image_def part_circlepath_def) qed lemma path_image_part_circlepath': "path_image (part_circlepath z r s t) = (\x. z + r * cis x) ` closed_segment s t" proof - have "path_image (part_circlepath z r s t) = (\x. z + r * exp(\ * of_real x)) ` linepath s t ` {0..1}" by (simp add: image_image path_image_def part_circlepath_def) also have "linepath s t ` {0..1} = closed_segment s t" by (rule linepath_image_01) finally show ?thesis by (simp add: cis_conv_exp) qed lemma path_image_part_circlepath_subset: "\s \ t; 0 \ r\ \ path_image(part_circlepath z r s t) \ sphere z r" by (auto simp: path_image_part_circlepath sphere_def dist_norm algebra_simps norm_mult) lemma in_path_image_part_circlepath: assumes "w \ path_image(part_circlepath z r s t)" "s \ t" "0 \ r" shows "norm(w - z) = r" by (smt (verit) assms dist_norm mem_Collect_eq norm_minus_commute path_image_part_circlepath_subset sphere_def subsetD) lemma path_image_part_circlepath_subset': assumes "r \ 0" shows "path_image (part_circlepath z r s t) \ sphere z r" by (smt (verit) assms path_image_part_circlepath_subset reversepath_part_circlepath reversepath_simps(2)) lemma part_circlepath_cnj: "cnj (part_circlepath c r a b x) = part_circlepath (cnj c) r (-a) (-b) x" by (simp add: part_circlepath_def exp_cnj linepath_def algebra_simps) +lemma contour_integrable_on_compose_cnj_iff: + assumes "valid_path \" + shows "f contour_integrable_on (cnj \ \) \ (cnj \ f \ cnj) contour_integrable_on \" +proof - + from assms obtain S where S: "finite S" "\ C1_differentiable_on {0..1} - S" + unfolding valid_path_def piecewise_C1_differentiable_on_def by blast + have "f contour_integrable_on (cnj \ \) \ + ((\t. cnj (cnj (f (cnj (\ t))) * vector_derivative \ (at t))) integrable_on {0..1})" + unfolding contour_integrable_on o_def + proof (intro integrable_spike_finite_eq [OF S(1)]) + fix t :: real assume "t \ {0..1} - S" + hence "\ differentiable at t" + using S(2) by (meson C1_differentiable_on_eq) + hence "vector_derivative (\x. cnj (\ x)) (at t) = cnj (vector_derivative \ (at t))" + by (rule vector_derivative_cnj) + thus "f (cnj (\ t)) * vector_derivative (\x. cnj (\ x)) (at t) = + cnj (cnj (f (cnj (\ t))) * vector_derivative \ (at t))" + by simp + qed + also have "\ \ ((\t. cnj (f (cnj (\ t))) * vector_derivative \ (at t)) integrable_on {0..1})" + by (rule integrable_on_cnj_iff) + also have "\ \ (cnj \ f \ cnj) contour_integrable_on \" + by (simp add: contour_integrable_on o_def) + finally show ?thesis . +qed + +lemma contour_integral_cnj: + assumes "valid_path \" + shows "contour_integral (cnj \ \) f = cnj (contour_integral \ (cnj \ f \ cnj))" +proof - + from assms obtain S where S: "finite S" "\ C1_differentiable_on {0..1} - S" + unfolding valid_path_def piecewise_C1_differentiable_on_def by blast + have "contour_integral (cnj \ \) f = + integral {0..1} (\t. cnj (cnj (f (cnj (\ t))) * vector_derivative \ (at t)))" + unfolding contour_integral_integral + proof (intro integral_spike) + fix t assume "t \ {0..1} - S" + hence "\ differentiable at t" + using S(2) by (meson C1_differentiable_on_eq) + hence "vector_derivative (\x. cnj (\ x)) (at t) = cnj (vector_derivative \ (at t))" + by (rule vector_derivative_cnj) + thus "cnj (cnj (f (cnj (\ t))) * vector_derivative \ (at t)) = + f ((cnj \ \) t) * vector_derivative (cnj \ \) (at t)" + by (simp add: o_def) + qed (use S(1) in auto) + also have "\ = cnj (integral {0..1} (\t. cnj (f (cnj (\ t))) * vector_derivative \ (at t)))" + by (subst integral_cnj [symmetric]) auto + also have "\ = cnj (contour_integral \ (cnj \ f \ cnj))" + by (simp add: contour_integral_integral) + finally show ?thesis . +qed + +lemma contour_integral_negatepath: + assumes "valid_path \" + shows "contour_integral (uminus \ \) f = -(contour_integral \ (\x. f (-x)))" (is "?lhs = ?rhs") +proof (cases "f contour_integrable_on (uminus \ \)") + case True + hence *: "(f has_contour_integral ?lhs) (uminus \ \)" + using has_contour_integral_integral by blast + have "((\z. f (-z)) has_contour_integral - contour_integral (uminus \ \) f) + (uminus \ (uminus \ \))" + by (rule has_contour_integral_negatepath) (use * assms in auto) + hence "((\x. f (-x)) has_contour_integral -?lhs) \" + by (simp add: o_def) + thus ?thesis + by (simp add: contour_integral_unique) +next + case False + hence "\(\z. f (- z)) contour_integrable_on \" + using contour_integrable_negatepath[of \ f] assms by auto + with False show ?thesis + by (simp add: not_integrable_contour_integral) +qed + lemma contour_integral_bound_part_circlepath: assumes "f contour_integrable_on part_circlepath c r a b" assumes "B \ 0" "r \ 0" "\x. x \ path_image (part_circlepath c r a b) \ norm (f x) \ B" shows "norm (contour_integral (part_circlepath c r a b) f) \ B * r * \b - a\" proof - let ?I = "integral {0..1} (\x. f (part_circlepath c r a b x) * \ * of_real (r * (b - a)) * exp (\ * linepath a b x))" have "norm ?I \ integral {0..1} (\x::real. B * 1 * (r * \b - a\) * 1)" proof (rule integral_norm_bound_integral, goal_cases) case 1 with assms(1) show ?case by (simp add: contour_integrable_on vector_derivative_part_circlepath mult_ac) next case (3 x) with assms(2-) show ?case unfolding norm_mult norm_of_real abs_mult by (intro mult_mono) (auto simp: path_image_def) qed auto also have "?I = contour_integral (part_circlepath c r a b) f" by (simp add: contour_integral_integral vector_derivative_part_circlepath mult_ac) finally show ?thesis by simp qed lemma has_contour_integral_part_circlepath_iff: assumes "a < b" shows "(f has_contour_integral I) (part_circlepath c r a b) \ ((\t. f (c + r * cis t) * r * \ * cis t) has_integral I) {a..b}" proof - have "(f has_contour_integral I) (part_circlepath c r a b) \ ((\x. f (part_circlepath c r a b x) * vector_derivative (part_circlepath c r a b) (at x within {0..1})) has_integral I) {0..1}" unfolding has_contour_integral_def .. also have "\ \ ((\x. f (part_circlepath c r a b x) * r * (b - a) * \ * cis (linepath a b x)) has_integral I) {0..1}" by (intro has_integral_cong, subst vector_derivative_part_circlepath01) (simp_all add: cis_conv_exp) also have "\ \ ((\x. f (c + r * exp (\ * linepath (of_real a) (of_real b) x)) * r * \ * exp (\ * linepath (of_real a) (of_real b) x) * vector_derivative (linepath (of_real a) (of_real b)) (at x within {0..1})) has_integral I) {0..1}" by (intro has_integral_cong, subst vector_derivative_linepath_within) (auto simp: part_circlepath_def cis_conv_exp of_real_linepath [symmetric]) also have "\ \ ((\z. f (c + r * exp (\ * z)) * r * \ * exp (\ * z)) has_contour_integral I) (linepath (of_real a) (of_real b))" by (simp add: has_contour_integral_def) also have "\ \ ((\t. f (c + r * cis t) * r * \ * cis t) has_integral I) {a..b}" using assms by (subst has_contour_integral_linepath_Reals_iff) (simp_all add: cis_conv_exp) finally show ?thesis . qed lemma contour_integrable_part_circlepath_iff: assumes "a < b" shows "f contour_integrable_on (part_circlepath c r a b) \ (\t. f (c + r * cis t) * r * \ * cis t) integrable_on {a..b}" using assms by (auto simp: contour_integrable_on_def integrable_on_def has_contour_integral_part_circlepath_iff) lemma contour_integral_part_circlepath_eq: assumes "a < b" shows "contour_integral (part_circlepath c r a b) f = integral {a..b} (\t. f (c + r * cis t) * r * \ * cis t)" proof (cases "f contour_integrable_on part_circlepath c r a b") case True hence "(\t. f (c + r * cis t) * r * \ * cis t) integrable_on {a..b}" using assms by (simp add: contour_integrable_part_circlepath_iff) with True show ?thesis using has_contour_integral_part_circlepath_iff[OF assms] contour_integral_unique has_integral_integrable_integral by blast next case False hence "\(\t. f (c + r * cis t) * r * \ * cis t) integrable_on {a..b}" using assms by (simp add: contour_integrable_part_circlepath_iff) with False show ?thesis by (simp add: not_integrable_contour_integral not_integrable_integral) qed lemma contour_integral_part_circlepath_reverse: "contour_integral (part_circlepath c r a b) f = -contour_integral (part_circlepath c r b a) f" by (metis contour_integral_reversepath reversepath_part_circlepath valid_path_part_circlepath) lemma contour_integral_part_circlepath_reverse': "b < a \ contour_integral (part_circlepath c r a b) f = -contour_integral (part_circlepath c r b a) f" by (rule contour_integral_part_circlepath_reverse) lemma finite_bounded_log: "finite {z::complex. norm z \ b \ exp z = w}" proof (cases "w = 0") case True then show ?thesis by auto next case False have *: "finite {x. cmod ((2 * real_of_int x * pi) * \) \ b + cmod (Ln w)}" proof (simp add: norm_mult finite_int_iff_bounded_le) show "\k. abs ` {x. 2 * \of_int x\ * pi \ b + cmod (Ln w)} \ {..k}" apply (rule_tac x="\(b + cmod (Ln w)) / (2*pi)\" in exI) apply (auto simp: field_split_simps le_floor_iff) done qed have [simp]: "\P f. {z. P z \ (\n. z = f n)} = f ` {n. P (f n)}" by blast have "finite {z. cmod z \ b \ exp z = exp (Ln w)}" using norm_add_leD by (fastforce intro: finite_subset [OF _ *] simp: exp_eq) then show ?thesis using False by auto qed lemma finite_bounded_log2: fixes a::complex assumes "a \ 0" shows "finite {z. norm z \ b \ exp(a*z) = w}" proof - have *: "finite ((\z. z / a) ` {z. cmod z \ b * cmod a \ exp z = w})" by (rule finite_imageI [OF finite_bounded_log]) show ?thesis by (rule finite_subset [OF _ *]) (force simp: assms norm_mult) qed lemma has_contour_integral_bound_part_circlepath_strong: assumes fi: "(f has_contour_integral i) (part_circlepath z r s t)" and "finite k" and le: "0 \ B" "0 < r" "s \ t" and B: "\x. x \ path_image(part_circlepath z r s t) - k \ norm(f x) \ B" shows "cmod i \ B * r * (t - s)" proof - consider "s = t" | "s < t" using \s \ t\ by linarith then show ?thesis proof cases case 1 with fi [unfolded has_contour_integral] have "i = 0" by (simp add: vector_derivative_part_circlepath) with assms show ?thesis by simp next case 2 have [simp]: "\r\ = r" using \r > 0\ by linarith have [simp]: "cmod (complex_of_real t - complex_of_real s) = t-s" by (metis "2" abs_of_pos diff_gt_0_iff_gt norm_of_real of_real_diff) have "finite (part_circlepath z r s t -` {y} \ {0..1})" if "y \ k" for y proof - let ?w = "(y - z)/of_real r / exp(\ * of_real s)" have fin: "finite (of_real -` {z. cmod z \ 1 \ exp (\ * complex_of_real (t - s) * z) = ?w})" using \s < t\ by (intro finite_vimageI [OF finite_bounded_log2]) (auto simp: inj_of_real) show ?thesis unfolding part_circlepath_def linepath_def vimage_def using le by (intro finite_subset [OF _ fin]) (auto simp: algebra_simps scaleR_conv_of_real exp_add exp_diff) qed then have fin01: "finite ((part_circlepath z r s t) -` k \ {0..1})" by (rule finite_finite_vimage_IntI [OF \finite k\]) have **: "((\x. if (part_circlepath z r s t x) \ k then 0 else f(part_circlepath z r s t x) * vector_derivative (part_circlepath z r s t) (at x)) has_integral i) {0..1}" by (rule has_integral_spike [OF negligible_finite [OF fin01]]) (use fi has_contour_integral in auto) have *: "\x. \0 \ x; x \ 1; part_circlepath z r s t x \ k\ \ cmod (f (part_circlepath z r s t x)) \ B" by (auto intro!: B [unfolded path_image_def image_def]) show ?thesis using has_integral_bound [where 'a=real, simplified, OF _ **] using assms le * "2" \r > 0\ by (auto simp add: norm_mult vector_derivative_part_circlepath) qed qed corollary contour_integral_bound_part_circlepath_strong: assumes "f contour_integrable_on part_circlepath z r s t" and "finite k" and "0 \ B" "0 < r" "s \ t" and "\x. x \ path_image(part_circlepath z r s t) - k \ norm(f x) \ B" shows "cmod (contour_integral (part_circlepath z r s t) f) \ B * r * (t - s)" using assms has_contour_integral_bound_part_circlepath_strong has_contour_integral_integral by blast lemma has_contour_integral_bound_part_circlepath: "\(f has_contour_integral i) (part_circlepath z r s t); 0 \ B; 0 < r; s \ t; \x. x \ path_image(part_circlepath z r s t) \ norm(f x) \ B\ \ norm i \ B*r*(t - s)" by (auto intro: has_contour_integral_bound_part_circlepath_strong) lemma contour_integrable_continuous_part_circlepath: "continuous_on (path_image (part_circlepath z r s t)) f \ f contour_integrable_on (part_circlepath z r s t)" unfolding contour_integrable_on has_contour_integral_def vector_derivative_part_circlepath path_image_def by (best intro: integrable_continuous_real path_part_circlepath [unfolded path_def] continuous_intros continuous_on_compose2 [where g=f, OF _ _ order_refl]) lemma simple_path_part_circlepath: "simple_path(part_circlepath z r s t) \ (r \ 0 \ s \ t \ \s - t\ \ 2*pi)" proof (cases "r = 0 \ s = t") case True then show ?thesis unfolding part_circlepath_def simple_path_def loop_free_def by (rule disjE) (force intro: bexI [where x = "1/4"] bexI [where x = "1/3"])+ next case False then have "r \ 0" "s \ t" by auto have *: "\x y z s t. \*((1 - x) * s + x * t) = \*(((1 - y) * s + y * t)) + z \ \*(x - y) * (t - s) = z" by (simp add: algebra_simps) have abs01: "\x y::real. 0 \ x \ x \ 1 \ 0 \ y \ y \ 1 \ (x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0 \ \x - y\ \ {0,1})" by auto have **: "\x y. (\n. (complex_of_real x - of_real y) * (of_real t - of_real s) = 2 * (of_int n * of_real pi)) \ (\n. \x - y\ * (t - s) = 2 * (of_int n * pi))" by (force simp: algebra_simps abs_if dest: arg_cong [where f=Re] arg_cong [where f=complex_of_real] intro: exI [where x = "-n" for n]) have 1: "\s - t\ \ 2 * pi" if "\x. 0 \ x \ x \ 1 \ (\n. x * (t - s) = 2 * (real_of_int n * pi)) \ x = 0 \ x = 1" proof (rule ccontr) assume "\ \s - t\ \ 2 * pi" then have *: "\n. t - s \ of_int n * \s - t\" using False that [of "2*pi / \t - s\"] by (simp add: abs_minus_commute divide_simps) show False using * [of 1] * [of "-1"] by auto qed have 2: "\s - t\ = \2 * (real_of_int n * pi) / x\" if "x \ 0" "x * (t - s) = 2 * (real_of_int n * pi)" for x n proof - have "t-s = 2 * (real_of_int n * pi)/x" using that by (simp add: field_simps) then show ?thesis by (metis abs_minus_commute) qed have abs_away: "\P. (\x\{0..1}. \y\{0..1}. P \x - y\) \ (\x::real. 0 \ x \ x \ 1 \ P x)" by force have "\x n. \s \ t; \s - t\ \ 2 * pi; 0 \ x; x < 1; x * (t - s) = 2 * (real_of_int n * pi)\ \ x = 0" by (rule ccontr) (auto simp: 2 field_split_simps abs_mult dest: of_int_leD) then show ?thesis using False apply (simp add: simple_path_def loop_free_def) apply (simp add: part_circlepath_def linepath_def exp_eq * ** abs01 del: Set.insert_iff) apply (subst abs_away) apply (auto simp: 1) done qed lemma arc_part_circlepath: assumes "r \ 0" "s \ t" "\s - t\ < 2*pi" shows "arc (part_circlepath z r s t)" proof - have *: "x = y" if eq: "\ * (linepath s t x) = \ * (linepath s t y) + 2 * of_int n * complex_of_real pi * \" and x: "x \ {0..1}" and y: "y \ {0..1}" for x y n proof (rule ccontr) assume "x \ y" have "(linepath s t x) = (linepath s t y) + 2 * of_int n * complex_of_real pi" by (metis add_divide_eq_iff complex_i_not_zero mult.commute nonzero_mult_div_cancel_left eq) then have "s*y + t*x = s*x + (t*y + of_int n * (pi * 2))" by (force simp: algebra_simps linepath_def dest: arg_cong [where f=Re]) with \x \ y\ have st: "s-t = (of_int n * (pi * 2) / (y-x))" by (force simp: field_simps) have "\real_of_int n\ < \y - x\" using assms \x \ y\ by (simp add: st abs_mult field_simps) then show False using assms x y st by (auto dest: of_int_lessD) qed then have "inj_on (part_circlepath z r s t) {0..1}" using assms by (force simp add: part_circlepath_def inj_on_def exp_eq) then show ?thesis by (simp add: arc_def) qed subsection\Special case of one complete circle\ definition\<^marker>\tag important\ circlepath :: "[complex, real, real] \ complex" where "circlepath z r \ part_circlepath z r 0 (2*pi)" lemma circlepath: "circlepath z r = (\x. z + r * exp(2 * of_real pi * \ * of_real x))" by (simp add: circlepath_def part_circlepath_def linepath_def algebra_simps) lemma pathstart_circlepath [simp]: "pathstart (circlepath z r) = z + r" by (simp add: circlepath_def) lemma pathfinish_circlepath [simp]: "pathfinish (circlepath z r) = z + r" by (simp add: circlepath_def) (metis exp_two_pi_i mult.commute) lemma circlepath_minus: "circlepath z (-r) x = circlepath z r (x + 1/2)" proof - have "z + of_real r * exp (2 * pi * \ * (x + 1/2)) = z + of_real r * exp (2 * pi * \ * x + pi * \)" by (simp add: divide_simps) (simp add: algebra_simps) also have "\ = z - r * exp (2 * pi * \ * x)" by (simp add: exp_add) finally show ?thesis by (simp add: circlepath path_image_def sphere_def dist_norm) qed lemma circlepath_add1: "circlepath z r (x+1) = circlepath z r x" using circlepath_minus [of z r "x+1/2"] circlepath_minus [of z "-r" x] by (simp add: add.commute) lemma circlepath_add_half: "circlepath z r (x + 1/2) = circlepath z r (x - 1/2)" using circlepath_add1 [of z r "x-1/2"] by (simp add: add.commute) lemma path_image_circlepath_minus_subset: "path_image (circlepath z (-r)) \ path_image (circlepath z r)" proof - have "\x\{0..1}. circlepath z r (y + 1/2) = circlepath z r x" if "0 \ y" "y \ 1" for y proof (cases "y \ 1/2") case False with that show ?thesis by (force simp: circlepath_add_half) qed (use that in force) then show ?thesis by (auto simp add: path_image_def image_def circlepath_minus) qed lemma path_image_circlepath_minus: "path_image (circlepath z (-r)) = path_image (circlepath z r)" using path_image_circlepath_minus_subset by fastforce lemma has_vector_derivative_circlepath [derivative_intros]: "((circlepath z r) has_vector_derivative (2 * pi * \ * r * exp (2 * of_real pi * \ * x))) (at x within X)" unfolding circlepath_def scaleR_conv_of_real by (rule derivative_eq_intros) (simp add: algebra_simps) lemma vector_derivative_circlepath: "vector_derivative (circlepath z r) (at x) = 2 * pi * \ * r * exp(2 * of_real pi * \ * x)" using has_vector_derivative_circlepath vector_derivative_at by blast lemma vector_derivative_circlepath01: "\0 \ x; x \ 1\ \ vector_derivative (circlepath z r) (at x within {0..1}) = 2 * pi * \ * r * exp(2 * of_real pi * \ * x)" using has_vector_derivative_circlepath by (auto simp: vector_derivative_at_within_ivl) lemma valid_path_circlepath [simp]: "valid_path (circlepath z r)" by (simp add: circlepath_def) lemma path_circlepath [simp]: "path (circlepath z r)" by (simp add: valid_path_imp_path) lemma path_image_circlepath_nonneg: assumes "0 \ r" shows "path_image (circlepath z r) = sphere z r" proof - have *: "x \ (\u. z + (cmod (x - z)) * exp (\ * (of_real u * (of_real pi * 2)))) ` {0..1}" for x proof (cases "x = z") case True then show ?thesis by force next case False define w where "w = x - z" then have "w \ 0" by (simp add: False) have **: "\t. \Re w = cos t * cmod w; Im w = sin t * cmod w\ \ w = of_real (cmod w) * exp (\ * t)" using cis_conv_exp complex_eq_iff by auto obtain t where "0 \ t" "t < 2*pi" "Re(w/norm w) = cos t" "Im(w/norm w) = sin t" apply (rule sincos_total_2pi [of "Re(w/(norm w))" "Im(w/(norm w))"]) by (auto simp add: divide_simps \w \ 0\ cmod_power2 [symmetric]) then show ?thesis using False ** w_def \w \ 0\ by (rule_tac x="t / (2*pi)" in image_eqI) (auto simp add: field_simps) qed show ?thesis unfolding circlepath path_image_def sphere_def dist_norm by (force simp: assms algebra_simps norm_mult norm_minus_commute intro: *) qed lemma path_image_circlepath [simp]: "path_image (circlepath z r) = sphere z \r\" using path_image_circlepath_minus by (force simp: path_image_circlepath_nonneg abs_if) lemma has_contour_integral_bound_circlepath_strong: "\(f has_contour_integral i) (circlepath z r); finite k; 0 \ B; 0 < r; \x. \norm(x - z) = r; x \ k\ \ norm(f x) \ B\ \ norm i \ B*(2*pi*r)" unfolding circlepath_def by (auto simp: algebra_simps in_path_image_part_circlepath dest!: has_contour_integral_bound_part_circlepath_strong) lemma has_contour_integral_bound_circlepath: "\(f has_contour_integral i) (circlepath z r); 0 \ B; 0 < r; \x. norm(x - z) = r \ norm(f x) \ B\ \ norm i \ B*(2*pi*r)" by (auto intro: has_contour_integral_bound_circlepath_strong) lemma contour_integrable_continuous_circlepath: "continuous_on (path_image (circlepath z r)) f \ f contour_integrable_on (circlepath z r)" by (simp add: circlepath_def contour_integrable_continuous_part_circlepath) lemma simple_path_circlepath: "simple_path(circlepath z r) \ (r \ 0)" by (simp add: circlepath_def simple_path_part_circlepath) lemma notin_path_image_circlepath [simp]: "cmod (w - z) < r \ w \ path_image (circlepath z r)" by (simp add: sphere_def dist_norm norm_minus_commute) lemma contour_integral_circlepath: assumes "r > 0" shows "contour_integral (circlepath z r) (\w. 1 / (w - z)) = 2 * complex_of_real pi * \" proof (rule contour_integral_unique) show "((\w. 1 / (w - z)) has_contour_integral 2 * complex_of_real pi * \) (circlepath z r)" unfolding has_contour_integral_def using assms has_integral_const_real [of _ 0 1] apply (subst has_integral_cong) apply (simp add: vector_derivative_circlepath01) apply (force simp: circlepath) done qed subsection\ Uniform convergence of path integral\ text\Uniform convergence when the derivative of the path is bounded, and in particular for the special case of a circle.\ proposition contour_integral_uniform_limit: assumes ev_fint: "eventually (\n::'a. (f n) contour_integrable_on \) F" and ul_f: "uniform_limit (path_image \) f l F" and noleB: "\t. t \ {0..1} \ norm (vector_derivative \ (at t)) \ B" and \: "valid_path \" and [simp]: "\ trivial_limit F" shows "l contour_integrable_on \" "((\n. contour_integral \ (f n)) \ contour_integral \ l) F" proof - have "0 \ B" by (meson noleB [of 0] atLeastAtMost_iff norm_ge_zero order_refl order_trans zero_le_one) { fix e::real assume "0 < e" then have "0 < e / (\B\ + 1)" by simp - then have "\\<^sub>F n in F. \x\path_image \. cmod (f n x - l x) < e / (\B\ + 1)" + then have \
: "\\<^sub>F n in F. \x\path_image \. cmod (f n x - l x) < e / (\B\ + 1)" using ul_f [unfolded uniform_limit_iff dist_norm] by auto - with ev_fint obtain a where fga: "\x. x \ {0..1} \ cmod (f a (\ x) - l (\ x)) < e / (\B\ + 1)" and inta: "(\t. f a (\ t) * vector_derivative \ (at t)) integrable_on {0..1}" - using eventually_happens [OF eventually_conj] + using eventually_happens [OF eventually_conj [OF ev_fint \
]] by (fastforce simp: contour_integrable_on path_image_def) have "\h. (\x\{0..1}. cmod (l (\ x) * vector_derivative \ (at x) - h x) \ e) \ h integrable_on {0..1}" proof (intro exI conjI ballI) show "cmod (l (\ x) * vector_derivative \ (at x) - f a (\ x) * vector_derivative \ (at x)) \ e" if "x \ {0..1}" for x proof - have "cmod (l (\ x) * vector_derivative \ (at x) - f a (\ x) * vector_derivative \ (at x)) \ B * e / (\B\ + 1)" using noleB [OF that] fga [OF that] \0 \ B\ \0 < e\ by (fastforce simp: mult_ac dest: mult_mono [OF less_imp_le] simp add: norm_mult left_diff_distrib [symmetric] norm_minus_commute divide_simps) also have "\ \ e" using \0 \ B\ \0 < e\ by (simp add: field_split_simps) finally show ?thesis . qed qed (rule inta) } then show lintg: "l contour_integrable_on \" unfolding contour_integrable_on by (metis (mono_tags, lifting)integrable_uniform_limit_real) { fix e::real define B' where "B' = B + 1" have B': "B' > 0" "B' > B" using \0 \ B\ by (auto simp: B'_def) assume "0 < e" then have ev_no': "\\<^sub>F n in F. \x\path_image \. 2 * cmod (f n x - l x) < e / B'" using ul_f [unfolded uniform_limit_iff dist_norm, rule_format, of "e / B'/2"] B' by (simp add: field_simps) have ie: "integral {0..1::real} (\x. e/2) < e" using \0 < e\ by simp have *: "cmod (f x (\ t) * vector_derivative \ (at t) - l (\ t) * vector_derivative \ (at t)) \ e/2" if t: "t\{0..1}" and leB': "2 * cmod (f x (\ t) - l (\ t)) < e / B'" for x t proof - have "2 * cmod (f x (\ t) - l (\ t)) * cmod (vector_derivative \ (at t)) \ e * (B/ B')" using mult_mono [OF less_imp_le [OF leB'] noleB] B' \0 < e\ t by auto also have "\ < e" by (simp add: B' \0 < e\ mult_imp_div_pos_less) finally have "2 * cmod (f x (\ t) - l (\ t)) * cmod (vector_derivative \ (at t)) < e" . then show ?thesis by (simp add: left_diff_distrib [symmetric] norm_mult) qed have le_e: "\x. \\u\{0..1}. 2 * cmod (f x (\ u) - l (\ u)) < e / B'; f x contour_integrable_on \\ \ cmod (integral {0..1} (\u. f x (\ u) * vector_derivative \ (at u) - l (\ u) * vector_derivative \ (at u))) < e" apply (rule le_less_trans [OF integral_norm_bound_integral ie]) apply (simp add: lintg integrable_diff contour_integrable_on [symmetric]) apply (blast intro: *)+ done have "\\<^sub>F x in F. dist (contour_integral \ (f x)) (contour_integral \ l) < e" apply (rule eventually_mono [OF eventually_conj [OF ev_no' ev_fint]]) apply (simp add: dist_norm contour_integrable_on path_image_def contour_integral_integral) apply (simp add: lintg integral_diff [symmetric] contour_integrable_on [symmetric] le_e) done } then show "((\n. contour_integral \ (f n)) \ contour_integral \ l) F" by (rule tendstoI) qed corollary\<^marker>\tag unimportant\ contour_integral_uniform_limit_circlepath: assumes "\\<^sub>F n::'a in F. (f n) contour_integrable_on (circlepath z r)" and "uniform_limit (sphere z r) f l F" and "\ trivial_limit F" "0 < r" shows "l contour_integrable_on (circlepath z r)" "((\n. contour_integral (circlepath z r) (f n)) \ contour_integral (circlepath z r) l) F" using assms by (auto simp: vector_derivative_circlepath norm_mult intro!: contour_integral_uniform_limit) end diff --git a/src/HOL/Complex_Analysis/Winding_Numbers.thy b/src/HOL/Complex_Analysis/Winding_Numbers.thy --- a/src/HOL/Complex_Analysis/Winding_Numbers.thy +++ b/src/HOL/Complex_Analysis/Winding_Numbers.thy @@ -1,2230 +1,2311 @@ section \Winding numbers\ theory Winding_Numbers imports Cauchy_Integral_Theorem begin subsection \Definition\ definition\<^marker>\tag important\ winding_number_prop :: "[real \ complex, complex, real, real \ complex, complex] \ bool" where "winding_number_prop \ z e p n \ valid_path p \ z \ path_image p \ pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ (\t \ {0..1}. norm(\ t - p t) < e) \ contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" definition\<^marker>\tag important\ winding_number:: "[real \ complex, complex] \ complex" where "winding_number \ z \ SOME n. \e > 0. \p. winding_number_prop \ z e p n" lemma winding_number: assumes "path \" "z \ path_image \" "0 < e" shows "\p. winding_number_prop \ z e p (winding_number \ z)" proof - have "path_image \ \ UNIV - {z}" using assms by blast then obtain d where d: "d>0" and pi_eq: "\h1 h2. valid_path h1 \ valid_path h2 \ (\t\{0..1}. cmod (h1 t - \ t) < d \ cmod (h2 t - \ t) < d) \ pathstart h2 = pathstart h1 \ pathfinish h2 = pathfinish h1 \ path_image h1 \ UNIV - {z} \ path_image h2 \ UNIV - {z} \ (\f. f holomorphic_on UNIV - {z} \ contour_integral h2 f = contour_integral h1 f)" using contour_integral_nearby_ends [of "UNIV - {z}" \] assms by (auto simp: open_delete) then obtain h where h: "polynomial_function h \ pathstart h = pathstart \ \ pathfinish h = pathfinish \ \ (\t \ {0..1}. norm(h t - \ t) < d/2)" using path_approx_polynomial_function [OF \path \\, of "d/2"] d by (metis half_gt_zero_iff) define nn where "nn = 1/(2* pi*\) * contour_integral h (\w. 1/(w - z))" have "\n. \e > 0. \p. winding_number_prop \ z e p n" proof (rule_tac x=nn in exI, clarify) fix e::real assume e: "e>0" obtain p where p: "polynomial_function p \ pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ (\t\{0..1}. cmod (p t - \ t) < min e (d/2))" using path_approx_polynomial_function [OF \path \\, of "min e (d/2)"] d \0 by (metis min_less_iff_conj zero_less_divide_iff zero_less_numeral) have "(\w. 1 / (w - z)) holomorphic_on UNIV - {z}" by (auto simp: intro!: holomorphic_intros) then have "winding_number_prop \ z e p nn" using pi_eq [of h p] h p d by (auto simp: valid_path_polynomial_function norm_minus_commute nn_def winding_number_prop_def) then show "\p. winding_number_prop \ z e p nn" by metis qed then show ?thesis unfolding winding_number_def by (rule someI2_ex) (blast intro: \0) qed lemma winding_number_unique: assumes \: "path \" "z \ path_image \" and pi: "\e. e>0 \ \p. winding_number_prop \ z e p n" shows "winding_number \ z = n" proof - have "path_image \ \ UNIV - {z}" using assms by blast then obtain e where e: "e>0" and pi_eq: "\h1 h2 f. \valid_path h1; valid_path h2; (\t\{0..1}. cmod (h1 t - \ t) < e \ cmod (h2 t - \ t) < e); pathstart h2 = pathstart h1; pathfinish h2 = pathfinish h1; f holomorphic_on UNIV - {z}\ \ contour_integral h2 f = contour_integral h1 f" using contour_integral_nearby_ends [of "UNIV - {z}" \] assms by (auto simp: open_delete) obtain p where p: "winding_number_prop \ z e p n" using pi [OF e] by blast obtain q where q: "winding_number_prop \ z e q (winding_number \ z)" using winding_number [OF \ e] by blast have "2 * complex_of_real pi * \ * n = contour_integral p (\w. 1 / (w - z))" using p by (auto simp: winding_number_prop_def) also have "\ = contour_integral q (\w. 1 / (w - z))" proof (rule pi_eq) show "(\w. 1 / (w - z)) holomorphic_on UNIV - {z}" by (auto intro!: holomorphic_intros) qed (use p q in \auto simp: winding_number_prop_def norm_minus_commute\) also have "\ = 2 * complex_of_real pi * \ * winding_number \ z" using q by (auto simp: winding_number_prop_def) finally have "2 * complex_of_real pi * \ * n = 2 * complex_of_real pi * \ * winding_number \ z" . then show ?thesis by simp qed +lemma winding_number_prop_reversepath: + assumes "winding_number_prop \ z e p n" + shows "winding_number_prop (reversepath \) z e (reversepath p) (-n)" +proof - + have p: "valid_path p" "z \ path_image p" "pathstart p = pathstart \" + "pathfinish p = pathfinish \" "\t. t \ {0..1} \ norm (\ t - p t) < e" + "contour_integral p (\w. 1 / (w - z)) = 2 * complex_of_real pi * \ * n" + using assms by (auto simp: winding_number_prop_def) + show ?thesis + unfolding winding_number_prop_def + proof (intro conjI strip) + show "norm (reversepath \ t - reversepath p t) < e" if "t \ {0..1}" for t + unfolding reversepath_def using p(5)[of "1 - t"] that by auto + show "contour_integral (reversepath p) (\w. 1 / (w - z)) = + complex_of_real (2 * pi) * \ * - n" + using p by (subst contour_integral_reversepath) auto + qed (use p in auto) +qed + +lemma winding_number_prop_reversepath_iff: + "winding_number_prop (reversepath \) z e p n \ winding_number_prop \ z e (reversepath p) (-n)" + using winding_number_prop_reversepath[of "reversepath \" z e p n] + winding_number_prop_reversepath[of \ z e "reversepath p" "-n"] by auto + (*NB not winding_number_prop here due to the loop in p*) lemma winding_number_unique_loop: assumes \: "path \" "z \ path_image \" and loop: "pathfinish \ = pathstart \" and pi: "\e. e>0 \ \p. valid_path p \ z \ path_image p \ pathfinish p = pathstart p \ (\t \ {0..1}. norm (\ t - p t) < e) \ contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" shows "winding_number \ z = n" proof - have "path_image \ \ UNIV - {z}" using assms by blast then obtain e where e: "e>0" and pi_eq: "\h1 h2 f. \valid_path h1; valid_path h2; (\t\{0..1}. cmod (h1 t - \ t) < e \ cmod (h2 t - \ t) < e); pathfinish h1 = pathstart h1; pathfinish h2 = pathstart h2; f holomorphic_on UNIV - {z}\ \ contour_integral h2 f = contour_integral h1 f" using contour_integral_nearby_loops [of "UNIV - {z}" \] assms by (auto simp: open_delete) obtain p where p: "valid_path p \ z \ path_image p \ pathfinish p = pathstart p \ (\t \ {0..1}. norm (\ t - p t) < e) \ contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" using pi [OF e] by blast obtain q where q: "winding_number_prop \ z e q (winding_number \ z)" using winding_number [OF \ e] by blast have "2 * complex_of_real pi * \ * n = contour_integral p (\w. 1 / (w - z))" using p by auto also have "\ = contour_integral q (\w. 1 / (w - z))" proof (rule pi_eq) show "(\w. 1 / (w - z)) holomorphic_on UNIV - {z}" by (auto intro!: holomorphic_intros) qed (use p q loop in \auto simp: winding_number_prop_def norm_minus_commute\) also have "\ = 2 * complex_of_real pi * \ * winding_number \ z" using q by (auto simp: winding_number_prop_def) finally have "2 * complex_of_real pi * \ * n = 2 * complex_of_real pi * \ * winding_number \ z" . then show ?thesis by simp qed proposition winding_number_valid_path: assumes "valid_path \" "z \ path_image \" shows "winding_number \ z = 1/(2*pi*\) * contour_integral \ (\w. 1/(w - z))" by (rule winding_number_unique) (use assms in \auto simp: valid_path_imp_path winding_number_prop_def\) proposition has_contour_integral_winding_number: assumes \: "valid_path \" "z \ path_image \" shows "((\w. 1/(w - z)) has_contour_integral (2*pi*\*winding_number \ z)) \" by (simp add: winding_number_valid_path has_contour_integral_integral contour_integrable_inversediff assms) lemma winding_number_trivial [simp]: "z \ a \ winding_number(linepath a a) z = 0" by (simp add: winding_number_valid_path) lemma winding_number_subpath_trivial [simp]: "z \ g x \ winding_number (subpath x x g) z = 0" by (simp add: path_image_subpath winding_number_valid_path) lemma winding_number_join: assumes \1: "path \1" "z \ path_image \1" and \2: "path \2" "z \ path_image \2" and "pathfinish \1 = pathstart \2" shows "winding_number(\1 +++ \2) z = winding_number \1 z + winding_number \2 z" proof (rule winding_number_unique) show "\p. winding_number_prop (\1 +++ \2) z e p (winding_number \1 z + winding_number \2 z)" if "e > 0" for e proof - obtain p1 where "winding_number_prop \1 z e p1 (winding_number \1 z)" using \0 < e\ \1 winding_number by blast moreover obtain p2 where "winding_number_prop \2 z e p2 (winding_number \2 z)" using \0 < e\ \2 winding_number by blast ultimately have "winding_number_prop (\1+++\2) z e (p1+++p2) (winding_number \1 z + winding_number \2 z)" using assms apply (simp add: winding_number_prop_def not_in_path_image_join contour_integrable_inversediff algebra_simps) apply (auto simp: joinpaths_def) done then show ?thesis by blast qed qed (use assms in \auto simp: not_in_path_image_join\) lemma winding_number_reversepath: assumes "path \" "z \ path_image \" shows "winding_number(reversepath \) z = - (winding_number \ z)" proof (rule winding_number_unique) show "\p. winding_number_prop (reversepath \) z e p (- winding_number \ z)" if "e > 0" for e proof - obtain p where "winding_number_prop \ z e p (winding_number \ z)" using \0 < e\ assms winding_number by blast then have "winding_number_prop (reversepath \) z e (reversepath p) (- winding_number \ z)" using assms unfolding winding_number_prop_def apply (simp add: contour_integral_reversepath contour_integrable_inversediff valid_path_imp_reverse) apply (auto simp: reversepath_def) done then show ?thesis by blast qed qed (use assms in auto) lemma winding_number_shiftpath: assumes \: "path \" "z \ path_image \" and "pathfinish \ = pathstart \" "a \ {0..1}" shows "winding_number(shiftpath a \) z = winding_number \ z" proof (rule winding_number_unique_loop) show "\p. valid_path p \ z \ path_image p \ pathfinish p = pathstart p \ (\t\{0..1}. cmod (shiftpath a \ t - p t) < e) \ contour_integral p (\w. 1 / (w - z)) = 2 * pi * \ * winding_number \ z" if "e > 0" for e proof - obtain p where "winding_number_prop \ z e p (winding_number \ z)" using \0 < e\ assms winding_number by blast then show ?thesis apply (rule_tac x="shiftpath a p" in exI) using assms that apply (auto simp: winding_number_prop_def path_image_shiftpath pathfinish_shiftpath pathstart_shiftpath contour_integral_shiftpath) apply (simp add: shiftpath_def) done qed qed (use assms in \auto simp: path_shiftpath path_image_shiftpath pathfinish_shiftpath pathstart_shiftpath\) lemma winding_number_split_linepath: assumes "c \ closed_segment a b" "z \ closed_segment a b" shows "winding_number(linepath a b) z = winding_number(linepath a c) z + winding_number(linepath c b) z" proof - have "z \ closed_segment a c" "z \ closed_segment c b" using assms by (meson convex_contains_segment convex_segment ends_in_segment subsetCE)+ then show ?thesis using assms by (simp add: winding_number_valid_path contour_integral_split_linepath [symmetric] continuous_on_inversediff field_simps) qed lemma winding_number_cong: "(\t. \0 \ t; t \ 1\ \ p t = q t) \ winding_number p z = winding_number q z" by (simp add: winding_number_def winding_number_prop_def pathstart_def pathfinish_def) lemma winding_number_constI: assumes "c\z" and g: "\t. \0\t; t\1\ \ g t = c" shows "winding_number g z = 0" proof - have "winding_number g z = winding_number (linepath c c) z" using g winding_number_cong by fastforce moreover have "winding_number (linepath c c) z = 0" using \c\z\ by auto ultimately show ?thesis by auto qed lemma winding_number_offset: "winding_number p z = winding_number (\w. p w - z) 0" unfolding winding_number_def proof (intro ext arg_cong [where f = Eps] arg_cong [where f = All] imp_cong refl, safe) fix n e g assume "0 < e" and g: "winding_number_prop p z e g n" then show "\r. winding_number_prop (\w. p w - z) 0 e r n" by (rule_tac x="\t. g t - z" in exI) (force simp: winding_number_prop_def contour_integral_integral valid_path_def path_defs vector_derivative_def has_vector_derivative_diff_const piecewise_C1_differentiable_diff C1_differentiable_imp_piecewise) next fix n e g assume "0 < e" and g: "winding_number_prop (\w. p w - z) 0 e g n" then have "winding_number_prop p z e (\t. g t + z) n" apply (simp add: winding_number_prop_def contour_integral_integral valid_path_def path_defs piecewise_C1_differentiable_add vector_derivative_def has_vector_derivative_add_const C1_differentiable_imp_piecewise) apply (force simp: algebra_simps) done then show "\r. winding_number_prop p z e r n" by metis qed lemma winding_number_negatepath: assumes \: "valid_path \" and 0: "0 \ path_image \" shows "winding_number(uminus \ \) 0 = winding_number \ 0" proof - have "(/) 1 contour_integrable_on \" using "0" \ contour_integrable_inversediff by fastforce then have "((\z. 1/z) has_contour_integral contour_integral \ ((/) 1)) \" by (rule has_contour_integral_integral) then have "((\z. 1 / - z) has_contour_integral - contour_integral \ ((/) 1)) \" using has_contour_integral_neg by auto then have "contour_integral (uminus \ \) ((/) 1) = contour_integral \ ((/) 1)" using \ by (simp add: contour_integral_unique has_contour_integral_negatepath) then show ?thesis using assms by (simp add: winding_number_valid_path valid_path_negatepath image_def path_defs) qed +lemma winding_number_cnj: + assumes "path \" "z \ path_image \" + shows "winding_number (cnj \ \) (cnj z) = -cnj (winding_number \ z)" +proof (rule winding_number_unique) + show "\p. winding_number_prop (cnj \ \) (cnj z) e p (-cnj (winding_number \ z))" + if "e > 0" for e + proof - + from winding_number[OF assms(1,2) \e > 0\] + obtain p where "winding_number_prop \ z e p (winding_number \ z)" + by blast + then have p: "valid_path p" "z \ path_image p" + "pathstart p = pathstart \" + "pathfinish p = pathfinish \" + "\t. t \ {0..1} \ cmod (\ t - p t) < e" and + p_cont:"contour_integral p (\w. 1 / (w - z)) = + complex_of_real (2 * pi) * \ * winding_number \ z" + unfolding winding_number_prop_def by auto + + have "valid_path (cnj \ p)" + using p(1) by (subst valid_path_cnj) auto + moreover have "cnj z \ path_image (cnj \ p)" + using p(2) by (auto simp: path_image_def) + moreover have "pathstart (cnj \ p) = pathstart (cnj \ \)" + using p(3) by (simp add: pathstart_compose) + moreover have "pathfinish (cnj \ p) = pathfinish (cnj \ \)" + using p(4) by (simp add: pathfinish_compose) + moreover have "cmod ((cnj \ \) t - (cnj \ p) t) < e" + if t: "t \ {0..1}" for t + proof - + have "(cnj \ \) t - (cnj \ p) t = cnj (\ t - p t)" + by simp + also have "norm \ = norm (\ t - p t)" + by (subst complex_mod_cnj) auto + also have "\ < e" + using p(5)[OF t] by simp + finally show ?thesis . + qed + moreover have "contour_integral (cnj \ p) (\w. 1 / (w - cnj z)) = + cnj (complex_of_real (2 * pi) * \ * winding_number \ z)" (is "?L=?R") + proof - + have "?L = contour_integral (cnj \ p) (cnj \ (\w. 1 / (cnj w - z)))" + by (simp add: o_def) + also have "\ = cnj (contour_integral p (\x. 1 / (x - z)))" + using p(1) by (subst contour_integral_cnj) (auto simp: o_def) + also have "\ = ?R" + using p_cont by simp + finally show ?thesis . + qed + ultimately show ?thesis + by (intro exI[of _ "cnj \ p"]) (auto simp: winding_number_prop_def) + qed + show "path (cnj \ \)" + by (intro path_continuous_image continuous_intros) (use assms in auto) + show "cnj z \ path_image (cnj \ \)" + using \z \ path_image \\ unfolding path_image_def by auto +qed + text \A combined theorem deducing several things piecewise.\ lemma winding_number_join_pos_combined: "\valid_path \1; z \ path_image \1; 0 < Re(winding_number \1 z); valid_path \2; z \ path_image \2; 0 < Re(winding_number \2 z); pathfinish \1 = pathstart \2\ \ valid_path(\1 +++ \2) \ z \ path_image(\1 +++ \2) \ 0 < Re(winding_number(\1 +++ \2) z)" by (simp add: valid_path_join path_image_join winding_number_join valid_path_imp_path) subsubsection\<^marker>\tag unimportant\ \Useful sufficient conditions for the winding number to be positive\ lemma Re_winding_number: "\valid_path \; z \ path_image \\ \ Re(winding_number \ z) = Im(contour_integral \ (\w. 1/(w - z))) / (2*pi)" by (simp add: winding_number_valid_path field_simps Re_divide power2_eq_square) lemma winding_number_pos_le: assumes \: "valid_path \" "z \ path_image \" and ge: "\x. \0 < x; x < 1\ \ 0 \ Im (vector_derivative \ (at x) * cnj(\ x - z))" shows "0 \ Re(winding_number \ z)" proof - have ge0: "0 \ Im (vector_derivative \ (at x) / (\ x - z))" if x: "0 < x" "x < 1" for x using ge by (simp add: Complex.Im_divide algebra_simps x) let ?vd = "\x. 1 / (\ x - z) * vector_derivative \ (at x)" let ?int = "\z. contour_integral \ (\w. 1 / (w - z))" have "0 \ Im (?int z)" proof (rule has_integral_component_nonneg [of \, simplified]) show "\x. x \ cbox 0 1 \ 0 \ Im (if 0 < x \ x < 1 then ?vd x else 0)" by (force simp: ge0) have "((\a. 1 / (a - z)) has_contour_integral contour_integral \ (\w. 1 / (w - z))) \" using \ by (simp flip: add: contour_integrable_inversediff has_contour_integral_integral) then have hi: "(?vd has_integral ?int z) (cbox 0 1)" using has_contour_integral by auto show "((\x. if 0 < x \ x < 1 then ?vd x else 0) has_integral ?int z) (cbox 0 1)" by (rule has_integral_spike_interior [OF hi]) simp qed then show ?thesis by (simp add: Re_winding_number [OF \] field_simps) qed lemma winding_number_pos_lt_lemma: assumes \: "valid_path \" "z \ path_image \" and e: "0 < e" and ge: "\x. \0 < x; x < 1\ \ e \ Im (vector_derivative \ (at x) / (\ x - z))" shows "0 < Re(winding_number \ z)" proof - let ?vd = "\x. 1 / (\ x - z) * vector_derivative \ (at x)" let ?int = "\z. contour_integral \ (\w. 1 / (w - z))" have "e \ Im (contour_integral \ (\w. 1 / (w - z)))" proof (rule has_integral_component_le [of \ "\x. \*e" "\*e" "{0..1}", simplified]) have "((\a. 1 / (a - z)) has_contour_integral contour_integral \ (\w. 1 / (w - z))) \" thm has_integral_component_le [of \ "\x. \*e" "\*e" "{0..1}", simplified] using \ by (simp add: contour_integrable_inversediff has_contour_integral_integral) then have hi: "(?vd has_integral ?int z) (cbox 0 1)" using has_contour_integral by auto show "((\x. if 0 < x \ x < 1 then ?vd x else \ * e) has_integral ?int z) {0..1}" by (rule has_integral_spike_interior [OF hi, simplified box_real]) (use e in simp) show "\x. 0 \ x \ x \ 1 \ e \ Im (if 0 < x \ x < 1 then ?vd x else \ * e)" by (simp add: ge) qed (use has_integral_const_real [of _ 0 1] in auto) with e show ?thesis by (simp add: Re_winding_number [OF \] field_simps) qed lemma winding_number_pos_lt: assumes \: "valid_path \" "z \ path_image \" and e: "0 < e" and ge: "\x. \0 < x; x < 1\ \ e \ Im (vector_derivative \ (at x) * cnj(\ x - z))" shows "0 < Re (winding_number \ z)" proof - have bm: "bounded ((\w. w - z) ` (path_image \))" using bounded_translation [of _ "-z"] \ by (simp add: bounded_valid_path_image) then obtain B where B: "B > 0" and Bno: "\x. x \ (\w. w - z) ` (path_image \) \ norm x \ B" using bounded_pos [THEN iffD1, OF bm] by blast { fix x::real assume x: "0 < x" "x < 1" then have B2: "cmod (\ x - z)^2 \ B^2" using Bno [of "\ x - z"] by (simp add: path_image_def power2_eq_square mult_mono') with x have "\ x \ z" using \ using path_image_def by fastforce then have "e / B\<^sup>2 \ e / (cmod (\ x - z))\<^sup>2" using B B2 e by (auto simp: divide_left_mono) also have "... \ Im (vector_derivative \ (at x) * cnj (\ x - z)) / (cmod (\ x - z))\<^sup>2" using ge [OF x] by (auto simp: divide_right_mono) finally have "e / B\<^sup>2 \ Im (vector_derivative \ (at x) * cnj (\ x - z)) / (cmod (\ x - z))\<^sup>2" . then have "e / B\<^sup>2 \ Im (vector_derivative \ (at x) / (\ x - z))" by (simp add: complex_div_cnj [of _ "\ x - z" for x] del: complex_cnj_diff times_complex.sel) } note * = this show ?thesis using e B by (simp add: * winding_number_pos_lt_lemma [OF \, of "e/B^2"]) qed subsection\The winding number is an integer\ text\Proof from the book Complex Analysis by Lars V. Ahlfors, Chapter 4, section 2.1, Also on page 134 of Serge Lang's book with the name title, etc.\ lemma exp_fg: fixes z::complex assumes g: "(g has_vector_derivative g') (at x within s)" and f: "(f has_vector_derivative (g' / (g x - z))) (at x within s)" and z: "g x \ z" shows "((\x. exp(-f x) * (g x - z)) has_vector_derivative 0) (at x within s)" proof - have *: "(exp \ (\x. (- f x)) has_vector_derivative - (g' / (g x - z)) * exp (- f x)) (at x within s)" using assms unfolding has_vector_derivative_def scaleR_conv_of_real by (auto intro!: derivative_eq_intros) show ?thesis using z by (auto intro!: derivative_eq_intros * [unfolded o_def] g) qed lemma winding_number_exp_integral: fixes z::complex assumes \: "\ piecewise_C1_differentiable_on {a..b}" and ab: "a \ b" and z: "z \ \ ` {a..b}" shows "(\x. vector_derivative \ (at x) / (\ x - z)) integrable_on {a..b}" (is "?thesis1") "exp (- (integral {a..b} (\x. vector_derivative \ (at x) / (\ x - z)))) * (\ b - z) = \ a - z" (is "?thesis2") proof - let ?D\ = "\x. vector_derivative \ (at x)" have [simp]: "\x. a \ x \ x \ b \ \ x \ z" using z by force have con_g: "continuous_on {a..b} \" using \ by (simp add: piecewise_C1_differentiable_on_def) obtain k where fink: "finite k" and g_C1_diff: "\ C1_differentiable_on ({a..b} - k)" using \ by (force simp: piecewise_C1_differentiable_on_def) have \: "open ({a<..finite k\ by (simp add: finite_imp_closed open_Diff) moreover have "{a<.. {a..b} - k" by force ultimately have g_diff_at: "\x. \x \ k; x \ {a<.. \ \ differentiable at x" by (metis Diff_iff differentiable_on_subset C1_diff_imp_diff [OF g_C1_diff] differentiable_on_def at_within_open) { fix w assume "w \ z" have "continuous_on (ball w (cmod (w - z))) (\w. 1 / (w - z))" by (auto simp: dist_norm intro!: continuous_intros) moreover have "\x. cmod (w - x) < cmod (w - z) \ \f'. ((\w. 1 / (w - z)) has_field_derivative f') (at x)" by (auto simp: intro!: derivative_eq_intros) ultimately have "\h. \y. norm(y - w) < norm(w - z) \ (h has_field_derivative 1/(y - z)) (at y)" using holomorphic_convex_primitive [of "ball w (norm(w - z))" "{}" "\w. 1/(w - z)"] by (force simp: field_differentiable_def Ball_def dist_norm at_within_open_NO_MATCH norm_minus_commute) } then obtain h where h: "\w y. w \ z \ norm(y - w) < norm(w - z) \ (h w has_field_derivative 1/(y - z)) (at y)" by meson have exy: "\y. ((\x. inverse (\ x - z) * ?D\ x) has_integral y) {a..b}" unfolding integrable_on_def [symmetric] proof (rule contour_integral_local_primitive_any [OF piecewise_C1_imp_differentiable [OF \]]) show "\d h. 0 < d \ (\y. cmod (y - w) < d \ (h has_field_derivative inverse (y - z))(at y within - {z}))" if "w \ - {z}" for w using that inverse_eq_divide has_field_derivative_at_within h by (metis Compl_insert DiffD2 insertCI right_minus_eq zero_less_norm_iff) qed simp have vg_int: "(\x. ?D\ x / (\ x - z)) integrable_on {a..b}" unfolding box_real [symmetric] divide_inverse_commute by (auto intro!: exy integrable_subinterval simp add: integrable_on_def ab) with ab show ?thesis1 by (simp add: divide_inverse_commute integral_def integrable_on_def) { fix t assume t: "t \ {a..b}" have cball: "continuous_on (ball (\ t) (dist (\ t) z)) (\x. inverse (x - z))" using z by (auto intro!: continuous_intros simp: dist_norm) have icd: "\x. cmod (\ t - x) < cmod (\ t - z) \ (\w. inverse (w - z)) field_differentiable at x" unfolding field_differentiable_def by (force simp: intro!: derivative_eq_intros) obtain h where h: "\x. cmod (\ t - x) < cmod (\ t - z) \ (h has_field_derivative inverse (x - z)) (at x within {y. cmod (\ t - y) < cmod (\ t - z)})" using holomorphic_convex_primitive [where f = "\w. inverse(w - z)", OF convex_ball finite.emptyI cball icd] by simp (auto simp: ball_def dist_norm that) have "exp (- (integral {a..t} (\x. ?D\ x / (\ x - z)))) * (\ t - z) =\ a - z" proof (rule has_derivative_zero_unique_strong_interval [of "{a,b} \ k" a b]) show "continuous_on {a..b} (\b. exp (- integral {a..b} (\x. ?D\ x / (\ x - z))) * (\ b - z))" by (auto intro!: continuous_intros con_g indefinite_integral_continuous_1 [OF vg_int]) show "((\b. exp (- integral {a..b} (\x. ?D\ x / (\ x - z))) * (\ b - z)) has_derivative (\h. 0)) (at x within {a..b})" if "x \ {a..b} - ({a, b} \ k)" for x proof - have x: "x \ k" "a < x" "x < b" using that by auto then have "x \ interior ({a..b} - k)" using open_subset_interior [OF \] by fastforce then have con: "isCont ?D\ x" using g_C1_diff x by (auto simp: C1_differentiable_on_eq intro: continuous_on_interior) then have con_vd: "continuous (at x within {a..b}) (\x. ?D\ x)" by (rule continuous_at_imp_continuous_within) have gdx: "\ differentiable at x" using x by (simp add: g_diff_at) then obtain d where d: "(\ has_derivative (\x. x *\<^sub>R d)) (at x)" by (auto simp add: differentiable_iff_scaleR) show "((\c. exp (- integral {a..c} (\x. ?D\ x / (\ x - z))) * (\ c - z)) has_derivative (\h. 0)) (at x within {a..b})" proof (rule exp_fg [unfolded has_vector_derivative_def, simplified]) show "(\ has_derivative (\c. c *\<^sub>R d)) (at x within {a..b})" using d by (blast intro: has_derivative_at_withinI) have "((\x. integral {a..x} (\x. ?D\ x / (\ x - z))) has_vector_derivative d / (\ x - z)) (at x within {a..b})" proof (rule has_vector_derivative_eq_rhs [OF integral_has_vector_derivative_continuous_at [where S = "{}", simplified]]) show "continuous (at x within {a..b}) (\x. vector_derivative \ (at x) / (\ x - z))" using continuous_at_imp_continuous_at_within differentiable_imp_continuous_within gdx x by (intro con_vd continuous_intros) (force+) show "vector_derivative \ (at x) / (\ x - z) = d / (\ x - z)" using d vector_derivative_at by (simp add: vector_derivative_at has_vector_derivative_def) qed (use x vg_int in auto) then show "((\x. integral {a..x} (\x. ?D\ x / (\ x - z))) has_derivative (\c. c *\<^sub>R (d / (\ x - z)))) (at x within {a..b})" by (auto simp: has_vector_derivative_def) qed (use x in auto) qed qed (use fink t in auto) } with ab show ?thesis2 by (simp add: divide_inverse_commute integral_def) qed lemma winding_number_exp_2pi: "\path p; z \ path_image p\ \ pathfinish p - z = exp (2 * pi * \ * winding_number p z) * (pathstart p - z)" using winding_number [of p z 1] unfolding valid_path_def path_image_def pathstart_def pathfinish_def winding_number_prop_def by (force dest: winding_number_exp_integral(2) [of _ 0 1 z] simp: field_simps contour_integral_integral exp_minus) lemma integer_winding_number_eq: assumes \: "path \" and z: "z \ path_image \" shows "winding_number \ z \ \ \ pathfinish \ = pathstart \" proof - obtain p where p: "valid_path p" "z \ path_image p" "pathstart p = pathstart \" "pathfinish p = pathfinish \" and eq: "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number \ z" using winding_number [OF assms, of 1] unfolding winding_number_prop_def by auto then have wneq: "winding_number \ z = winding_number p z" using eq winding_number_valid_path by force have iff: "(winding_number \ z \ \) \ (exp (contour_integral p (\w. 1 / (w - z))) = 1)" using eq by (simp add: exp_eq_1 complex_is_Int_iff) have "\ 0 \ z" by (metis pathstart_def pathstart_in_path_image z) then have "exp (contour_integral p (\w. 1 / (w - z))) = (\ 1 - z) / (\ 0 - z)" using p winding_number_exp_integral(2) [of p 0 1 z] by (simp add: valid_path_def path_defs contour_integral_integral exp_minus field_split_simps) then have "winding_number p z \ \ \ pathfinish p = pathstart p" using p wneq iff by (auto simp: path_defs) then show ?thesis using p eq by (auto simp: winding_number_valid_path) qed theorem integer_winding_number: "\path \; pathfinish \ = pathstart \; z \ path_image \\ \ winding_number \ z \ \" by (metis integer_winding_number_eq) text\If the winding number's magnitude is at least one, then the path must contain points in every direction.*) We can thus bound the winding number of a path that doesn't intersect a given ray. \ lemma winding_number_pos_meets: fixes z::complex assumes \: "valid_path \" and z: "z \ path_image \" and 1: "Re (winding_number \ z) \ 1" and w: "w \ z" shows "\a::real. 0 < a \ z + of_real a * (w - z) \ path_image \" proof - have [simp]: "\x. 0 \ x \ x \ 1 \ \ x \ z" using z by (auto simp: path_image_def) have [simp]: "z \ \ ` {0..1}" using path_image_def z by auto have gpd: "\ piecewise_C1_differentiable_on {0..1}" using \ valid_path_def by blast define r where "r = (w - z) / (\ 0 - z)" have [simp]: "r \ 0" using w z by (auto simp: r_def) have cont: "continuous_on {0..1} (\x. Im (integral {0..x} (\x. vector_derivative \ (at x) / (\ x - z))))" by (intro continuous_intros indefinite_integral_continuous_1 winding_number_exp_integral [OF gpd]; simp) have "Arg2pi r \ 2*pi" by (simp add: Arg2pi less_eq_real_def) also have "\ \ Im (integral {0..1} (\x. vector_derivative \ (at x) / (\ x - z)))" using 1 by (simp add: winding_number_valid_path [OF \ z] contour_integral_integral Complex.Re_divide field_simps power2_eq_square) finally have "Arg2pi r \ Im (integral {0..1} (\x. vector_derivative \ (at x) / (\ x - z)))" . then have "\t. t \ {0..1} \ Im(integral {0..t} (\x. vector_derivative \ (at x)/(\ x - z))) = Arg2pi r" by (simp add: Arg2pi_ge_0 cont IVT') then obtain t where t: "t \ {0..1}" and eqArg: "Im (integral {0..t} (\x. vector_derivative \ (at x)/(\ x - z))) = Arg2pi r" by blast define i where "i = integral {0..t} (\x. vector_derivative \ (at x) / (\ x - z))" have gpdt: "\ piecewise_C1_differentiable_on {0..t}" by (metis atLeastAtMost_iff atLeastatMost_subset_iff order_refl piecewise_C1_differentiable_on_subset gpd t) have "exp (- i) * (\ t - z) = \ 0 - z" unfolding i_def proof (rule winding_number_exp_integral [OF gpdt]) show "z \ \ ` {0..t}" using t z unfolding path_image_def by force qed (use t in auto) then have *: "\ t - z = exp i * (\ 0 - z)" by (simp add: exp_minus field_simps) then have "(w - z) = r * (\ 0 - z)" by (simp add: r_def) moreover have "z + exp (Re i) * (exp (\ * Im i) * (\ 0 - z)) = \ t" using * by (simp add: exp_eq_polar field_simps) moreover have "Arg2pi r = Im i" using eqArg by (simp add: i_def) ultimately have "z + complex_of_real (exp (Re i)) * (w - z) / complex_of_real (cmod r) = \ t" using Complex_Transcendental.Arg2pi_eq [of r] \r \ 0\ by (metis mult.left_commute nonzero_mult_div_cancel_left norm_eq_zero of_real_0 of_real_eq_iff times_divide_eq_left) with t show ?thesis by (rule_tac x="exp(Re i) / norm r" in exI) (auto simp: path_image_def) qed lemma winding_number_big_meets: fixes z::complex assumes \: "valid_path \" and z: "z \ path_image \" and "\Re (winding_number \ z)\ \ 1" and w: "w \ z" shows "\a::real. 0 < a \ z + of_real a * (w - z) \ path_image \" proof - { assume "Re (winding_number \ z) \ - 1" then have "Re (winding_number (reversepath \) z) \ 1" by (simp add: \ valid_path_imp_path winding_number_reversepath z) moreover have "valid_path (reversepath \)" using \ valid_path_imp_reverse by auto moreover have "z \ path_image (reversepath \)" by (simp add: z) ultimately have "\a::real. 0 < a \ z + of_real a * (w - z) \ path_image (reversepath \)" using winding_number_pos_meets w by blast then have ?thesis by simp } then show ?thesis using assms by (simp add: abs_if winding_number_pos_meets split: if_split_asm) qed lemma winding_number_less_1: fixes z::complex shows "\valid_path \; z \ path_image \; w \ z; \a::real. 0 < a \ z + of_real a * (w - z) \ path_image \\ \ Re(winding_number \ z) < 1" by (auto simp: not_less dest: winding_number_big_meets) text\One way of proving that WN=1 for a loop.\ lemma winding_number_eq_1: fixes z::complex assumes \: "valid_path \" and z: "z \ path_image \" and loop: "pathfinish \ = pathstart \" and 0: "0 < Re(winding_number \ z)" and 2: "Re(winding_number \ z) < 2" shows "winding_number \ z = 1" proof - have "winding_number \ z \ Ints" by (simp add: \ integer_winding_number loop valid_path_imp_path z) then show ?thesis using 0 2 by (auto simp: Ints_def) qed subsection\Continuity of winding number and invariance on connected sets\ theorem continuous_at_winding_number: fixes z::complex assumes \: "path \" and z: "z \ path_image \" shows "continuous (at z) (winding_number \)" proof - obtain e where "e>0" and cbg: "cball z e \ - path_image \" using open_contains_cball [of "- path_image \"] z by (force simp: closed_def [symmetric] closed_path_image [OF \]) then have ppag: "path_image \ \ - cball z (e/2)" by (force simp: cball_def dist_norm) have oc: "open (- cball z (e/2))" by (simp add: closed_def [symmetric]) obtain d where "d>0" and pi_eq: "\h1 h2. \valid_path h1; valid_path h2; (\t\{0..1}. cmod (h1 t - \ t) < d \ cmod (h2 t - \ t) < d); pathstart h2 = pathstart h1; pathfinish h2 = pathfinish h1\ \ path_image h1 \ - cball z (e/2) \ path_image h2 \ - cball z (e/2) \ (\f. f holomorphic_on - cball z (e/2) \ contour_integral h2 f = contour_integral h1 f)" using contour_integral_nearby_ends [OF oc \ ppag] by metis obtain p where "valid_path p" "z \ path_image p" and p: "pathstart p = pathstart \" "pathfinish p = pathfinish \" and pg: "\t. t\{0..1} \ cmod (\ t - p t) < min d e/2" and pi: "contour_integral p (\x. 1 / (x - z)) = complex_of_real (2 * pi) * \ * winding_number \ z" using winding_number [OF \ z, of "min d e/2"] \d>0\ \e>0\ by (auto simp: winding_number_prop_def) { fix w assume d2: "cmod (w - z) < d/2" and e2: "cmod (w - z) < e/2" have wnotp: "w \ path_image p" proof (clarsimp simp add: path_image_def) show False if w: "w = p x" and "0 \ x" "x \ 1" for x proof - have "cmod (\ x - p x) < min d e/2" using pg that by auto then have "cmod (z - \ x) < e" by (metis e2 less_divide_eq_numeral1(1) min_less_iff_conj norm_minus_commute norm_triangle_half_l w) then show ?thesis using cbg that by (auto simp add: path_image_def cball_def dist_norm less_eq_real_def) qed qed have wnotg: "w \ path_image \" using cbg e2 \e>0\ by (force simp: dist_norm norm_minus_commute) { fix k::real assume k: "k>0" then obtain q where q: "valid_path q" "w \ path_image q" "pathstart q = pathstart \ \ pathfinish q = pathfinish \" and qg: "\t. t \ {0..1} \ cmod (\ t - q t) < min k (min d e) / 2" and qi: "contour_integral q (\u. 1 / (u - w)) = complex_of_real (2 * pi) * \ * winding_number \ w" using winding_number [OF \ wnotg, of "min k (min d e) / 2"] \d>0\ \e>0\ k by (force simp: min_divide_distrib_right winding_number_prop_def) moreover have "\t. t \ {0..1} \ cmod (q t - \ t) < d \ cmod (p t - \ t) < d" using pg qg \0 < d\ by (fastforce simp add: norm_minus_commute) moreover have "(\u. 1 / (u-w)) holomorphic_on - cball z (e/2)" using e2 by (auto simp: dist_norm norm_minus_commute intro!: holomorphic_intros) ultimately have "contour_integral p (\u. 1 / (u - w)) = contour_integral q (\u. 1 / (u - w))" by (metis p \valid_path p\ pi_eq) then have "contour_integral p (\x. 1 / (x - w)) = complex_of_real (2 * pi) * \ * winding_number \ w" by (simp add: pi qi) } note pip = this have "path p" by (simp add: \valid_path p\ valid_path_imp_path) moreover have "\e. e > 0 \ winding_number_prop p w e p (winding_number \ w)" by (simp add: \valid_path p\ pip winding_number_prop_def wnotp) ultimately have "winding_number p w = winding_number \ w" using winding_number_unique wnotp by blast } note wnwn = this obtain pe where "pe>0" and cbp: "cball z (3 / 4 * pe) \ - path_image p" using \valid_path p\ \z \ path_image p\ open_contains_cball [of "- path_image p"] by (force simp: closed_def [symmetric] closed_path_image [OF valid_path_imp_path]) obtain L where "L>0" and L: "\f B. \f holomorphic_on - cball z (3 / 4 * pe); \z \ - cball z (3 / 4 * pe). cmod (f z) \ B\ \ cmod (contour_integral p f) \ L * B" using contour_integral_bound_exists [of "- cball z (3/4*pe)" p] cbp \valid_path p\ by blast { fix e::real and w::complex assume e: "0 < e" and w: "cmod (w - z) < pe/4" "cmod (w - z) < e * pe\<^sup>2 / (8 * L)" then have [simp]: "w \ path_image p" using cbp p(2) \0 < pe\ by (force simp: dist_norm norm_minus_commute path_image_def cball_def) have [simp]: "contour_integral p (\x. 1/(x - w)) - contour_integral p (\x. 1/(x - z)) = contour_integral p (\x. 1/(x - w) - 1/(x - z))" by (simp add: \valid_path p\ \z \ path_image p\ contour_integrable_inversediff contour_integral_diff) { fix x assume pe: "3/4 * pe < cmod (z - x)" have "cmod (w - x) < pe/4 + cmod (z - x)" by (meson add_less_cancel_right norm_diff_triangle_le order_refl order_trans_rules(21) w(1)) then have wx: "cmod (w - x) < 4/3 * cmod (z - x)" using pe by simp have "cmod (z - x) \ cmod (z - w) + cmod (w - x)" using norm_diff_triangle_le by blast also have "\ < pe/4 + cmod (w - x)" using w by (simp add: norm_minus_commute) finally have "pe/2 < cmod (w - x)" using pe by auto then have pe_less: "(pe/2)^2 < cmod (w - x) ^ 2" by (simp add: \0 < pe\ less_eq_real_def power_strict_mono) then have pe2: "pe^2 < 4 * cmod (w - x) ^ 2" by (simp add: power_divide) have "8 * L * cmod (w - z) < e * pe\<^sup>2" using w \L>0\ by (simp add: field_simps) also have "\ < e * 4 * cmod (w - x) * cmod (w - x)" using pe2 \e>0\ by (simp add: power2_eq_square) also have "\ < e * 4 * cmod (w - x) * (4/3 * cmod (z - x))" using \0 < pe\ pe_less e less_eq_real_def wx by fastforce finally have "L * cmod (w - z) < 2/3 * e * cmod (w - x) * cmod (z - x)" by simp also have "\ \ e * cmod (w - x) * cmod (z - x)" using e by simp finally have Lwz: "L * cmod (w - z) < e * cmod (w - x) * cmod (z - x)" . have "L * cmod (1 / (x - w) - 1 / (x - z)) \ e" proof (cases "x=z \ x=w") case True with pe \pe>0\ w \L>0\ show ?thesis by (force simp: norm_minus_commute) next case False with wx w(2) \L>0\ pe pe2 Lwz show ?thesis by (auto simp: divide_simps mult_less_0_iff norm_minus_commute norm_divide norm_mult power2_eq_square) qed } note L_cmod_le = this let ?f = "(\x. 1 / (x - w) - 1 / (x - z))" have "cmod (contour_integral p ?f) \ L * (e * pe\<^sup>2 / L / 4 * (inverse (pe / 2))\<^sup>2)" proof (rule L) show "?f holomorphic_on - cball z (3 / 4 * pe)" using \pe>0\ w by (force simp: dist_norm norm_minus_commute intro!: holomorphic_intros) show " \u \- cball z (3 / 4 * pe). cmod (?f u) \ e * pe\<^sup>2 / L / 4 * (inverse (pe / 2))\<^sup>2" using \pe>0\ w \L>0\ by (auto simp: cball_def dist_norm field_simps L_cmod_le simp del: less_divide_eq_numeral1 le_divide_eq_numeral1) qed also have "\ < 2*e" using \L>0\ e by (force simp: field_simps) finally have "cmod (winding_number p w - winding_number p z) < e" using pi_ge_two e by (force simp: winding_number_valid_path \valid_path p\ \z \ path_image p\ field_simps norm_divide norm_mult intro: less_le_trans) } note cmod_wn_diff = this have "isCont (winding_number p) z" proof (clarsimp simp add: continuous_at_eps_delta) fix e::real assume "e>0" show "\d>0. \x'. dist x' z < d \ dist (winding_number p x') (winding_number p z) < e" using \pe>0\ \L>0\ \e>0\ by (rule_tac x="min (pe/4) (e/2*pe^2/L/4)" in exI) (simp add: dist_norm cmod_wn_diff) qed then show ?thesis apply (rule continuous_transform_within [where d = "min d e/2"]) apply (auto simp: \d>0\ \e>0\ dist_norm wnwn) done qed corollary continuous_on_winding_number: "path \ \ continuous_on (- path_image \) (\w. winding_number \ w)" by (simp add: continuous_at_imp_continuous_on continuous_at_winding_number) subsection\<^marker>\tag unimportant\ \The winding number is constant on a connected region\ lemma winding_number_constant: assumes \: "path \" and loop: "pathfinish \ = pathstart \" and cs: "connected S" and sg: "S \ path_image \ = {}" shows "winding_number \ constant_on S" proof - have *: "1 \ cmod (winding_number \ y - winding_number \ z)" if ne: "winding_number \ y \ winding_number \ z" and "y \ S" "z \ S" for y z proof - have "winding_number \ y \ \" "winding_number \ z \ \" using that integer_winding_number [OF \ loop] sg \y \ S\ by auto with ne show ?thesis by (auto simp: Ints_def simp flip: of_int_diff) qed have cont: "continuous_on S (\w. winding_number \ w)" using continuous_on_winding_number [OF \] sg by (meson continuous_on_subset disjoint_eq_subset_Compl) show ?thesis using "*" zero_less_one by (blast intro: continuous_discrete_range_constant [OF cs cont]) qed lemma winding_number_eq: "\path \; pathfinish \ = pathstart \; w \ S; z \ S; connected S; S \ path_image \ = {}\ \ winding_number \ w = winding_number \ z" using winding_number_constant by (metis constant_on_def) lemma open_winding_number_levelsets: assumes \: "path \" and loop: "pathfinish \ = pathstart \" shows "open {z. z \ path_image \ \ winding_number \ z = k}" proof (clarsimp simp: open_dist) fix z assume z: "z \ path_image \" and k: "k = winding_number \ z" have "open (- path_image \)" by (simp add: closed_path_image \ open_Compl) then obtain e where "e>0" "ball z e \ - path_image \" using open_contains_ball [of "- path_image \"] z by blast then show "\e>0. \y. dist y z < e \ y \ path_image \ \ winding_number \ y = winding_number \ z" using \e>0\ by (force simp: norm_minus_commute dist_norm intro: winding_number_eq [OF assms, where S = "ball z e"]) qed subsection\Winding number is zero "outside" a curve\ proposition winding_number_zero_in_outside: assumes \: "path \" and loop: "pathfinish \ = pathstart \" and z: "z \ outside (path_image \)" shows "winding_number \ z = 0" proof - obtain B::real where "0 < B" and B: "path_image \ \ ball 0 B" using bounded_subset_ballD [OF bounded_path_image [OF \]] by auto obtain w::complex where w: "w \ ball 0 (B + 1)" by (metis abs_of_nonneg le_less less_irrefl mem_ball_0 norm_of_real) have "- ball 0 (B + 1) \ outside (path_image \)" using B subset_ball by (intro outside_subset_convex) auto then have wout: "w \ outside (path_image \)" using w by blast moreover have "winding_number \ constant_on outside (path_image \)" using winding_number_constant [OF \ loop, of "outside(path_image \)"] connected_outside by (metis DIM_complex bounded_path_image dual_order.refl \ outside_no_overlap) ultimately have "winding_number \ z = winding_number \ w" by (metis (no_types, opaque_lifting) constant_on_def z) also have "\ = 0" proof - have wnot: "w \ path_image \" using wout by (simp add: outside_def) { fix e::real assume "0" "pathfinish p = pathfinish \" and pg1: "(\t. \0 \ t; t \ 1\ \ cmod (p t - \ t) < 1)" and pge: "(\t. \0 \ t; t \ 1\ \ cmod (p t - \ t) < e)" using path_approx_polynomial_function [OF \, of "min 1 e"] \e>0\ by (metis atLeastAtMost_iff min_less_iff_conj zero_less_one) have "\p. valid_path p \ w \ path_image p \ pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ (\t\{0..1}. cmod (\ t - p t) < e) \ contour_integral p (\wa. 1 / (wa - w)) = 0" proof (intro exI conjI) have "\x. \0 \ x; x \ 1\ \ cmod (p x) < B + 1" using B unfolding image_subset_iff path_image_def by (meson add_strict_mono atLeastAtMost_iff le_less_trans mem_ball_0 norm_triangle_sub pg1) then have pip: "path_image p \ ball 0 (B + 1)" by (auto simp add: path_image_def dist_norm ball_def) then show "w \ path_image p" using w by blast show vap: "valid_path p" by (simp add: p(1) valid_path_polynomial_function) show "\t\{0..1}. cmod (\ t - p t) < e" by (metis atLeastAtMost_iff norm_minus_commute pge) show "contour_integral p (\wa. 1 / (wa - w)) = 0" proof (rule contour_integral_unique [OF Cauchy_theorem_convex_simple [OF _ convex_ball [of 0 "B+1"]]]) have "\z. cmod z < B + 1 \ z \ w" using mem_ball_0 w by blast then show "(\z. 1 / (z - w)) holomorphic_on ball 0 (B + 1)" by (intro holomorphic_intros; simp add: dist_norm) qed (use p vap pip loop in auto) qed (use p in auto) } then show ?thesis by (auto intro: winding_number_unique [OF \] simp add: winding_number_prop_def wnot) qed finally show ?thesis . qed corollary\<^marker>\tag unimportant\ winding_number_zero_const: "a \ z \ winding_number (\t. a) z = 0" by (rule winding_number_zero_in_outside) (auto simp: pathfinish_def pathstart_def path_polynomial_function) corollary\<^marker>\tag unimportant\ winding_number_zero_outside: "\path \; convex s; pathfinish \ = pathstart \; z \ s; path_image \ \ s\ \ winding_number \ z = 0" by (meson convex_in_outside outside_mono subsetCE winding_number_zero_in_outside) lemma winding_number_zero_at_infinity: assumes \: "path \" and loop: "pathfinish \ = pathstart \" shows "\B. \z. B \ norm z \ winding_number \ z = 0" proof - obtain B::real where "0 < B" and B: "path_image \ \ ball 0 B" using bounded_subset_ballD [OF bounded_path_image [OF \]] by auto have "winding_number \ z = 0" if "B + 1 \ cmod z" for z proof (rule winding_number_zero_outside [OF \ convex_cball loop]) show "z \ cball 0 B" using that by auto show "path_image \ \ cball 0 B" using B order.trans by blast qed then show ?thesis by metis qed lemma bounded_winding_number_nz: assumes "path g" "pathfinish g = pathstart g" shows "bounded {z. winding_number g z \ 0}" proof - obtain B where "\x. norm x \ B \ winding_number g x = 0" using winding_number_zero_at_infinity[OF assms] by auto thus ?thesis unfolding bounded_iff by (intro exI[of _ "B + 1"]) force qed lemma winding_number_zero_point: "\path \; convex S; pathfinish \ = pathstart \; open S; path_image \ \ S\ \ \z. z \ S \ winding_number \ z = 0" using outside_compact_in_open [of "path_image \" S] path_image_nonempty winding_number_zero_in_outside by (fastforce simp add: compact_path_image) text\If a path winds round a set, it winds rounds its inside.\ lemma winding_number_around_inside: assumes \: "path \" and loop: "pathfinish \ = pathstart \" and cls: "closed S" and cos: "connected S" and S_disj: "S \ path_image \ = {}" and z: "z \ S" and wn_nz: "winding_number \ z \ 0" and w: "w \ S \ inside S" shows "winding_number \ w = winding_number \ z" proof - have ssb: "S \ inside(path_image \)" proof fix x :: complex assume "x \ S" hence "x \ path_image \" by (meson disjoint_iff_not_equal S_disj) thus "x \ inside (path_image \)" by (metis Compl_iff S_disj UnE \ \x \ S\ cos inside_outside loop winding_number_eq winding_number_zero_in_outside wn_nz z) qed show ?thesis proof (rule winding_number_eq [OF \ loop w]) show "z \ S \ inside S" using z by blast show "connected (S \ inside S)" by (simp add: cls connected_with_inside cos) show "(S \ inside S) \ path_image \ = {}" unfolding disjoint_iff Un_iff by (metis ComplD UnI1 \ cls compact_path_image connected_path_image inside_Un_outside inside_inside_compact_connected ssb subsetD) qed qed text\Bounding a WN by 1/2 for a path and point in opposite halfspaces.\ lemma winding_number_subpath_continuous: assumes \: "valid_path \" and z: "z \ path_image \" shows "continuous_on {0..1} (\x. winding_number(subpath 0 x \) z)" proof (rule continuous_on_eq) let ?f = "\x. integral {0..x} (\t. vector_derivative \ (at t) / (\ t - z))" show "continuous_on {0..1} (\x. 1 / (2 * pi * \) * ?f x)" proof (intro indefinite_integral_continuous_1 winding_number_exp_integral continuous_intros) show "\ piecewise_C1_differentiable_on {0..1}" using \ valid_path_def by blast qed (use path_image_def z in auto) show "1 / (2 * pi * \) * ?f x = winding_number (subpath 0 x \) z" if x: "x \ {0..1}" for x proof - have "1 / (2*pi*\) * ?f x = 1 / (2*pi*\) * contour_integral (subpath 0 x \) (\w. 1/(w - z))" using assms x by (simp add: contour_integral_subcontour_integral [OF contour_integrable_inversediff]) also have "\ = winding_number (subpath 0 x \) z" proof (subst winding_number_valid_path) show "z \ path_image (subpath 0 x \)" using assms x atLeastAtMost_iff path_image_subpath_subset by force qed (use assms x valid_path_subpath in \force+\) finally show ?thesis . qed qed lemma winding_number_ivt_pos: assumes \: "valid_path \" and z: "z \ path_image \" and "0 \ w" "w \ Re(winding_number \ z)" shows "\t \ {0..1}. Re(winding_number(subpath 0 t \) z) = w" proof - have "continuous_on {0..1} (\x. winding_number (subpath 0 x \) z)" using \ winding_number_subpath_continuous z by blast moreover have "Re (winding_number (subpath 0 0 \) z) \ w" "w \ Re (winding_number (subpath 0 1 \) z)" using assms by (auto simp: path_image_def image_def) ultimately show ?thesis using ivt_increasing_component_on_1[of 0 1, where ?k = "1"] by force qed lemma winding_number_ivt_neg: assumes \: "valid_path \" and z: "z \ path_image \" and "Re(winding_number \ z) \ w" "w \ 0" shows "\t \ {0..1}. Re(winding_number(subpath 0 t \) z) = w" proof - have "continuous_on {0..1} (\x. winding_number (subpath 0 x \) z)" using \ winding_number_subpath_continuous z by blast moreover have "Re (winding_number (subpath 0 0 \) z) \ w" "w \ Re (winding_number (subpath 0 1 \) z)" using assms by (auto simp: path_image_def image_def) ultimately show ?thesis using ivt_decreasing_component_on_1[of 0 1, where ?k = "1"] by force qed lemma winding_number_ivt_abs: assumes \: "valid_path \" and z: "z \ path_image \" and "0 \ w" "w \ \Re(winding_number \ z)\" shows "\t \ {0..1}. \Re (winding_number (subpath 0 t \) z)\ = w" using assms winding_number_ivt_pos [of \ z w] winding_number_ivt_neg [of \ z "-w"] by force lemma winding_number_lt_half_lemma: assumes \: "valid_path \" and z: "z \ path_image \" and az: "a \ z \ b" and pag: "path_image \ \ {w. a \ w > b}" shows "Re(winding_number \ z) < 1/2" proof - { assume "Re(winding_number \ z) \ 1/2" then obtain t::real where t: "0 \ t" "t \ 1" and sub12: "Re (winding_number (subpath 0 t \) z) = 1/2" using winding_number_ivt_pos [OF \ z, of "1/2"] by auto have gt: "\ t - z = - (of_real (exp (- (2 * pi * Im (winding_number (subpath 0 t \) z)))) * (\ 0 - z))" using winding_number_exp_2pi [of "subpath 0 t \" z] apply (simp add: t \ valid_path_imp_path) using closed_segment_eq_real_ivl path_image_def t z by (fastforce simp: path_image_subpath Euler sub12) have "b < a \ \ 0" proof - have "\ 0 \ {c. b < a \ c}" by (metis (no_types) pag atLeastAtMost_iff image_subset_iff order_refl path_image_def zero_le_one) thus ?thesis by blast qed moreover have "b < a \ \ t" by (metis atLeastAtMost_iff image_eqI mem_Collect_eq pag path_image_def subset_iff t) ultimately have "0 < a \ (\ 0 - z)" "0 < a \ (\ t - z)" using az by (simp add: inner_diff_right)+ then have False by (simp add: gt inner_mult_right mult_less_0_iff) } then show ?thesis by force qed lemma winding_number_lt_half: assumes "valid_path \" "a \ z \ b" "path_image \ \ {w. a \ w > b}" shows "\Re (winding_number \ z)\ < 1/2" proof - have "z \ path_image \" using assms by auto with assms have "Re (winding_number \ z) < 0 \ - Re (winding_number \ z) < 1/2" by (metis complex_inner_1_right winding_number_lt_half_lemma [OF valid_path_imp_reverse, of \ z a b] winding_number_reversepath valid_path_imp_path inner_minus_left path_image_reversepath) with assms show ?thesis using \z \ path_image \\ winding_number_lt_half_lemma by fastforce qed lemma winding_number_le_half: assumes \: "valid_path \" and z: "z \ path_image \" and anz: "a \ 0" and azb: "a \ z \ b" and pag: "path_image \ \ {w. a \ w \ b}" shows "\Re (winding_number \ z)\ \ 1/2" proof - { assume wnz_12: "\Re (winding_number \ z)\ > 1/2" have "isCont (winding_number \) z" by (metis continuous_at_winding_number valid_path_imp_path \ z) then obtain d where "d>0" and d: "\x'. dist x' z < d \ dist (winding_number \ x') (winding_number \ z) < \Re(winding_number \ z)\ - 1/2" using continuous_at_eps_delta wnz_12 diff_gt_0_iff_gt by blast define z' where "z' = z - (d / (2 * cmod a)) *\<^sub>R a" have "a \ z * 6 \ d * cmod a + b * 6" by (metis \0 < d\ add_increasing azb less_eq_real_def mult_nonneg_nonneg mult_right_mono norm_ge_zero norm_numeral) with anz have *: "a \ z' \ b - d / 3 * cmod a" unfolding z'_def inner_mult_right' divide_inverse by (simp add: field_split_simps algebra_simps dot_square_norm power2_eq_square) have "cmod (winding_number \ z' - winding_number \ z) < \Re (winding_number \ z)\ - 1/2" using d [of z'] anz \d>0\ by (simp add: dist_norm z'_def) then have "1/2 < \Re (winding_number \ z)\ - cmod (winding_number \ z' - winding_number \ z)" by simp then have "1/2 < \Re (winding_number \ z)\ - \Re (winding_number \ z') - Re (winding_number \ z)\" using abs_Re_le_cmod [of "winding_number \ z' - winding_number \ z"] by simp then have wnz_12': "\Re (winding_number \ z')\ > 1/2" by linarith moreover have "\Re (winding_number \ z')\ < 1/2" proof (rule winding_number_lt_half [OF \ *]) show "path_image \ \ {w. b - d / 3 * cmod a < a \ w}" using azb \d>0\ pag by (auto simp: add_strict_increasing anz field_split_simps dest!: subsetD) qed ultimately have False by simp } then show ?thesis by force qed lemma winding_number_lt_half_linepath: assumes "z \ closed_segment a b" shows "\Re (winding_number (linepath a b) z)\ < 1/2" proof - obtain u v where "u \ z \ v" and uv: "\x. x \ closed_segment a b \ inner u x > v" using separating_hyperplane_closed_point assms closed_segment convex_closed_segment less_eq_real_def by metis moreover have "path_image (linepath a b) \ {w. v < u \ w}" using in_segment(1) uv by auto ultimately show ?thesis using winding_number_lt_half by auto qed text\ Positivity of WN for a linepath.\ lemma winding_number_linepath_pos_lt: assumes "0 < Im ((b - a) * cnj (b - z))" shows "0 < Re(winding_number(linepath a b) z)" proof - have z: "z \ path_image (linepath a b)" using assms by (simp add: closed_segment_def) (force simp: algebra_simps) show ?thesis by (intro winding_number_pos_lt [OF valid_path_linepath z assms]) (simp add: linepath_def algebra_simps) qed subsection\<^marker>\tag unimportant\ \More winding number properties\ text\including the fact that it's +-1 inside a simple closed curve.\ lemma winding_number_homotopic_paths: assumes "homotopic_paths (-{z}) g h" shows "winding_number g z = winding_number h z" proof - have "path g" "path h" using homotopic_paths_imp_path [OF assms] by auto moreover have pag: "z \ path_image g" and pah: "z \ path_image h" using homotopic_paths_imp_subset [OF assms] by auto ultimately obtain d e where "d > 0" "e > 0" and d: "\p. \path p; pathstart p = pathstart g; pathfinish p = pathfinish g; \t\{0..1}. norm (p t - g t) < d\ \ homotopic_paths (-{z}) g p" and e: "\q. \path q; pathstart q = pathstart h; pathfinish q = pathfinish h; \t\{0..1}. norm (q t - h t) < e\ \ homotopic_paths (-{z}) h q" using homotopic_nearby_paths [of g "-{z}"] homotopic_nearby_paths [of h "-{z}"] by force obtain p where p: "valid_path p" "z \ path_image p" "pathstart p = pathstart g" "pathfinish p = pathfinish g" and gp_less:"\t\{0..1}. cmod (g t - p t) < d" and pap: "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number g z" using winding_number [OF \path g\ pag \0 < d\] unfolding winding_number_prop_def by blast obtain q where q: "valid_path q" "z \ path_image q" "pathstart q = pathstart h" "pathfinish q = pathfinish h" and hq_less: "\t\{0..1}. cmod (h t - q t) < e" and paq: "contour_integral q (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number h z" using winding_number [OF \path h\ pah \0 < e\] unfolding winding_number_prop_def by blast have "homotopic_paths (- {z}) g p" by (simp add: d p valid_path_imp_path norm_minus_commute gp_less) moreover have "homotopic_paths (- {z}) h q" by (simp add: e q valid_path_imp_path norm_minus_commute hq_less) ultimately have "homotopic_paths (- {z}) p q" by (blast intro: homotopic_paths_trans homotopic_paths_sym assms) then have "contour_integral p (\w. 1/(w - z)) = contour_integral q (\w. 1/(w - z))" by (rule Cauchy_theorem_homotopic_paths) (auto intro!: holomorphic_intros simp: p q) then show ?thesis by (simp add: pap paq) qed lemma winding_number_homotopic_loops: assumes "homotopic_loops (-{z}) g h" shows "winding_number g z = winding_number h z" proof - have "path g" "path h" using homotopic_loops_imp_path [OF assms] by auto moreover have pag: "z \ path_image g" and pah: "z \ path_image h" using homotopic_loops_imp_subset [OF assms] by auto moreover have gloop: "pathfinish g = pathstart g" and hloop: "pathfinish h = pathstart h" using homotopic_loops_imp_loop [OF assms] by auto ultimately obtain d e where "d > 0" "e > 0" and d: "\p. \path p; pathfinish p = pathstart p; \t\{0..1}. norm (p t - g t) < d\ \ homotopic_loops (-{z}) g p" and e: "\q. \path q; pathfinish q = pathstart q; \t\{0..1}. norm (q t - h t) < e\ \ homotopic_loops (-{z}) h q" using homotopic_nearby_loops [of g "-{z}"] homotopic_nearby_loops [of h "-{z}"] by force obtain p where p: "valid_path p" "z \ path_image p" "pathstart p = pathstart g" "pathfinish p = pathfinish g" and gp_less:"\t\{0..1}. cmod (g t - p t) < d" and pap: "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number g z" using winding_number [OF \path g\ pag \0 < d\] unfolding winding_number_prop_def by blast obtain q where q: "valid_path q" "z \ path_image q" "pathstart q = pathstart h" "pathfinish q = pathfinish h" and hq_less: "\t\{0..1}. cmod (h t - q t) < e" and paq: "contour_integral q (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number h z" using winding_number [OF \path h\ pah \0 < e\] unfolding winding_number_prop_def by blast have gp: "homotopic_loops (- {z}) g p" by (simp add: gloop d gp_less norm_minus_commute p valid_path_imp_path) have hq: "homotopic_loops (- {z}) h q" by (simp add: e hloop hq_less norm_minus_commute q valid_path_imp_path) have "contour_integral p (\w. 1/(w - z)) = contour_integral q (\w. 1/(w - z))" proof (rule Cauchy_theorem_homotopic_loops) show "homotopic_loops (- {z}) p q" by (blast intro: homotopic_loops_trans homotopic_loops_sym gp hq assms) qed (auto intro!: holomorphic_intros simp: p q) then show ?thesis by (simp add: pap paq) qed lemma winding_number_paths_linear_eq: "\path g; path h; pathstart h = pathstart g; pathfinish h = pathfinish g; \t. t \ {0..1} \ z \ closed_segment (g t) (h t)\ \ winding_number h z = winding_number g z" by (blast intro: sym homotopic_paths_linear winding_number_homotopic_paths) lemma winding_number_loops_linear_eq: "\path g; path h; pathfinish g = pathstart g; pathfinish h = pathstart h; \t. t \ {0..1} \ z \ closed_segment (g t) (h t)\ \ winding_number h z = winding_number g z" by (blast intro: sym homotopic_loops_linear winding_number_homotopic_loops) lemma winding_number_nearby_paths_eq: "\path g; path h; pathstart h = pathstart g; pathfinish h = pathfinish g; \t. t \ {0..1} \ norm(h t - g t) < norm(g t - z)\ \ winding_number h z = winding_number g z" by (metis segment_bound(2) norm_minus_commute not_le winding_number_paths_linear_eq) lemma winding_number_nearby_loops_eq: "\path g; path h; pathfinish g = pathstart g; pathfinish h = pathstart h; \t. t \ {0..1} \ norm(h t - g t) < norm(g t - z)\ \ winding_number h z = winding_number g z" by (metis segment_bound(2) norm_minus_commute not_le winding_number_loops_linear_eq) lemma winding_number_subpath_combine: assumes "path g" and notin: "z \ path_image g" and "u \ {0..1}" "v \ {0..1}" "w \ {0..1}" shows "winding_number (subpath u v g) z + winding_number (subpath v w g) z = winding_number (subpath u w g) z" (is "?lhs = ?rhs") proof - have "?lhs = winding_number (subpath u v g +++ subpath v w g) z" using assms by (metis (no_types, lifting) path_image_subpath_subset path_subpath pathfinish_subpath pathstart_subpath subsetD winding_number_join) also have "... = ?rhs" by (meson assms homotopic_join_subpaths subset_Compl_singleton winding_number_homotopic_paths) finally show ?thesis . qed text \Winding numbers of circular contours\ proposition winding_number_part_circlepath_pos_less: assumes "s < t" and no: "norm(w - z) < r" shows "0 < Re (winding_number(part_circlepath z r s t) w)" proof - have "0 < r" by (meson no norm_not_less_zero not_le order.strict_trans2) note valid_path_part_circlepath moreover have " w \ path_image (part_circlepath z r s t)" using assms by (auto simp: path_image_def image_def part_circlepath_def norm_mult linepath_def) moreover have "0 < r * (t - s) * (r - cmod (w - z))" using assms by (metis \0 < r\ diff_gt_0_iff_gt mult_pos_pos) ultimately show ?thesis apply (rule winding_number_pos_lt [where e = "r*(t - s)*(r - norm(w - z))"]) apply (simp add: vector_derivative_part_circlepath right_diff_distrib [symmetric] mult_ac mult_le_cancel_left_pos assms \0) using Re_Im_le_cmod [of "w-z" "linepath s t x" for x] by (simp add: exp_Euler cos_of_real sin_of_real part_circlepath_def algebra_simps cos_squared_eq [unfolded power2_eq_square]) qed lemma winding_number_circlepath_centre: "0 < r \ winding_number (circlepath z r) z = 1" apply (rule winding_number_unique_loop) apply (simp_all add: sphere_def valid_path_imp_path) apply (rule_tac x="circlepath z r" in exI) apply (simp add: sphere_def contour_integral_circlepath) done proposition winding_number_circlepath: assumes "norm(w - z) < r" shows "winding_number(circlepath z r) w = 1" proof (cases "w = z") case True then show ?thesis using assms winding_number_circlepath_centre by auto next case False have [simp]: "r > 0" using assms le_less_trans norm_ge_zero by blast define r' where "r' = norm(w - z)" have "r' < r" by (simp add: assms r'_def) have disjo: "cball z r' \ sphere z r = {}" using \r' < r\ by (force simp: cball_def sphere_def) have "winding_number(circlepath z r) w = winding_number(circlepath z r) z" proof (rule winding_number_around_inside [where S = "cball z r'"]) show "winding_number (circlepath z r) z \ 0" by (simp add: winding_number_circlepath_centre) show "cball z r' \ path_image (circlepath z r) = {}" by (simp add: disjo less_eq_real_def) qed (auto simp: r'_def dist_norm norm_minus_commute) also have "\ = 1" by (simp add: winding_number_circlepath_centre) finally show ?thesis . qed lemma no_bounded_connected_component_imp_winding_number_zero: assumes g: "path g" "path_image g \ S" "pathfinish g = pathstart g" "z \ S" and nb: "\z. bounded (connected_component_set (- S) z) \ z \ S" shows "winding_number g z = 0" proof - have "z \ outside (path_image g)" by (metis nb [of z] \path_image g \ S\ \z \ S\ subsetD mem_Collect_eq outside outside_mono) then show ?thesis by (simp add: g winding_number_zero_in_outside) qed lemma no_bounded_path_component_imp_winding_number_zero: assumes g: "path g" "path_image g \ S" "pathfinish g = pathstart g" "z \ S" and nb: "\z. bounded (path_component_set (- S) z) \ z \ S" shows "winding_number g z = 0" by (simp add: bounded_subset nb path_component_subset_connected_component no_bounded_connected_component_imp_winding_number_zero [OF g]) subsection\Winding number for a triangle\ lemma wn_triangle1: assumes "0 \ interior(convex hull {a,b,c})" shows "\ (Im(a/b) \ 0 \ 0 \ Im(b/c))" proof - { assume 0: "Im(a/b) \ 0" "0 \ Im(b/c)" have "0 \ interior (convex hull {a,b,c})" proof (cases "a=0 \ b=0 \ c=0") case True then show ?thesis by (auto simp: not_in_interior_convex_hull_3) next case False then have "b \ 0" by blast { fix x y::complex and u::real assume eq_f': "Im x * Re b \ Im b * Re x" "Im y * Re b \ Im b * Re y" "0 \ u" "u \ 1" then have "((1 - u) * Im x) * Re b \ Im b * ((1 - u) * Re x)" by (simp add: mult_left_mono mult.assoc mult.left_commute [of "Im b"]) then have "((1 - u) * Im x + u * Im y) * Re b \ Im b * ((1 - u) * Re x + u * Re y)" using eq_f' ordered_comm_semiring_class.comm_mult_left_mono by (fastforce simp add: algebra_simps) } then have "convex {z. Im z * Re b \ Im b * Re z}" by (auto simp: algebra_simps convex_alt) with False 0 have "convex hull {a,b,c} \ {z. Im z * Re b \ Im b * Re z}" by (simp add: subset_hull mult.commute Complex.Im_divide divide_simps complex_neq_0 [symmetric]) moreover have "0 \ interior({z. Im z * Re b \ Im b * Re z})" proof assume "0 \ interior {z. Im z * Re b \ Im b * Re z}" then obtain e where "e>0" and e: "ball 0 e \ {z. Im z * Re b \ Im b * Re z}" by (meson mem_interior) define z where "z \ - sgn (Im b) * (e/3) + sgn (Re b) * (e/3) * \" have "cmod z = cmod (e/3) * cmod (- sgn (Im b) + sgn (Re b) * \)" unfolding z_def norm_mult [symmetric] by (simp add: algebra_simps) also have "... < e" using \e>0\ by (auto simp: norm_mult intro: le_less_trans [OF norm_triangle_ineq4]) finally have "z \ ball 0 e" using \e>0\ by simp then have "Im z * Re b \ Im b * Re z" using e by blast then have b: "0 < Re b" "0 < Im b" and disj: "e * Re b < - (Im b * e) \ e * Re b = - (Im b * e)" using \e>0\ \b \ 0\ by (auto simp add: z_def dist_norm sgn_if less_eq_real_def mult_less_0_iff complex.expand split: if_split_asm) show False \\or just one smt line\ using disj proof assume "e * Re b < - (Im b * e)" with b \0 < e\ less_trans [of _ 0] show False by (metis (no_types) mult_pos_pos neg_less_0_iff_less not_less_iff_gr_or_eq) next assume "e * Re b = - (Im b * e)" with b \0 < e\ show False by (metis mult_pos_pos neg_less_0_iff_less not_less_iff_gr_or_eq) qed qed ultimately show ?thesis using interior_mono by blast qed } with assms show ?thesis by blast qed lemma wn_triangle2_0: assumes "0 \ interior(convex hull {a,b,c})" shows "0 < Im((b - a) * cnj (b)) \ 0 < Im((c - b) * cnj (c)) \ 0 < Im((a - c) * cnj (a)) \ Im((b - a) * cnj (b)) < 0 \ 0 < Im((b - c) * cnj (b)) \ 0 < Im((a - b) * cnj (a)) \ 0 < Im((c - a) * cnj (c))" proof - have [simp]: "{b,c,a} = {a,b,c}" "{c,a,b} = {a,b,c}" by auto show ?thesis using wn_triangle1 [OF assms] wn_triangle1 [of b c a] wn_triangle1 [of c a b] assms by (auto simp: algebra_simps Im_complex_div_gt_0 Im_complex_div_lt_0 not_le not_less) qed lemma wn_triangle2: assumes "z \ interior(convex hull {a,b,c})" shows "0 < Im((b - a) * cnj (b - z)) \ 0 < Im((c - b) * cnj (c - z)) \ 0 < Im((a - c) * cnj (a - z)) \ Im((b - a) * cnj (b - z)) < 0 \ 0 < Im((b - c) * cnj (b - z)) \ 0 < Im((a - b) * cnj (a - z)) \ 0 < Im((c - a) * cnj (c - z))" proof - have 0: "0 \ interior(convex hull {a-z, b-z, c-z})" using assms convex_hull_translation [of "-z" "{a,b,c}"] interior_translation [of "-z"] by (simp cong: image_cong_simp) show ?thesis using wn_triangle2_0 [OF 0] by simp qed lemma wn_triangle3: assumes z: "z \ interior(convex hull {a,b,c})" and "0 < Im((b-a) * cnj (b-z))" "0 < Im((c-b) * cnj (c-z))" "0 < Im((a-c) * cnj (a-z))" shows "winding_number (linepath a b +++ linepath b c +++ linepath c a) z = 1" proof - have znot[simp]: "z \ closed_segment a b" "z \ closed_segment b c" "z \ closed_segment c a" using z interior_of_triangle [of a b c] by (auto simp: closed_segment_def) have gt0: "0 < Re (winding_number (linepath a b +++ linepath b c +++ linepath c a) z)" using assms by (simp add: winding_number_linepath_pos_lt path_image_join winding_number_join_pos_combined) have lt2: "Re (winding_number (linepath a b +++ linepath b c +++ linepath c a) z) < 2" using winding_number_lt_half_linepath [of _ a b] using winding_number_lt_half_linepath [of _ b c] using winding_number_lt_half_linepath [of _ c a] znot by (fastforce simp add: winding_number_join path_image_join) show ?thesis by (rule winding_number_eq_1) (simp_all add: path_image_join gt0 lt2) qed proposition winding_number_triangle: assumes z: "z \ interior(convex hull {a,b,c})" shows "winding_number(linepath a b +++ linepath b c +++ linepath c a) z = (if 0 < Im((b - a) * cnj (b - z)) then 1 else -1)" proof - have [simp]: "{a,c,b} = {a,b,c}" by auto have znot[simp]: "z \ closed_segment a b" "z \ closed_segment b c" "z \ closed_segment c a" using z interior_of_triangle [of a b c] by (auto simp: closed_segment_def) then have [simp]: "z \ closed_segment b a" "z \ closed_segment c b" "z \ closed_segment a c" using closed_segment_commute by blast+ have *: "winding_number (linepath a b +++ linepath b c +++ linepath c a) z = winding_number (reversepath (linepath a c +++ linepath c b +++ linepath b a)) z" by (simp add: reversepath_joinpaths winding_number_join not_in_path_image_join) show ?thesis apply (rule disjE [OF wn_triangle2 [OF z]]) subgoal by (simp add: wn_triangle3 z) subgoal by (simp add: path_image_join winding_number_reversepath * wn_triangle3 z) done qed subsection\Winding numbers for simple closed paths\ lemma winding_number_from_innerpath: assumes "simple_path c1" and c1: "pathstart c1 = a" "pathfinish c1 = b" and "simple_path c2" and c2: "pathstart c2 = a" "pathfinish c2 = b" and "simple_path c" and c: "pathstart c = a" "pathfinish c = b" and c1c2: "path_image c1 \ path_image c2 = {a,b}" and c1c: "path_image c1 \ path_image c = {a,b}" and c2c: "path_image c2 \ path_image c = {a,b}" and ne_12: "path_image c \ inside(path_image c1 \ path_image c2) \ {}" and z: "z \ inside(path_image c1 \ path_image c)" and wn_d: "winding_number (c1 +++ reversepath c) z = d" and "a \ b" "d \ 0" obtains "z \ inside(path_image c1 \ path_image c2)" "winding_number (c1 +++ reversepath c2) z = d" proof - obtain 0: "inside(path_image c1 \ path_image c) \ inside(path_image c2 \ path_image c) = {}" and 1: "inside(path_image c1 \ path_image c) \ inside(path_image c2 \ path_image c) \ (path_image c - {a,b}) = inside(path_image c1 \ path_image c2)" by (rule split_inside_simple_closed_curve [OF \simple_path c1\ c1 \simple_path c2\ c2 \simple_path c\ c \a \ b\ c1c2 c1c c2c ne_12]) have znot: "z \ path_image c" "z \ path_image c1" "z \ path_image c2" using union_with_outside z 1 by auto then have zout: "z \ outside (path_image c \ path_image c2)" by (metis "0" ComplI UnE disjoint_iff_not_equal sup.commute union_with_inside z) have wn_cc2: "winding_number (c +++ reversepath c2) z = 0" by (rule winding_number_zero_in_outside; simp add: zout \simple_path c2\ c c2 \simple_path c\ simple_path_imp_path path_image_join) show ?thesis proof show "z \ inside (path_image c1 \ path_image c2)" using "1" z by blast have "winding_number c1 z - winding_number c z = d " using assms znot by (metis wn_d winding_number_join simple_path_imp_path winding_number_reversepath add.commute path_image_reversepath path_reversepath pathstart_reversepath uminus_add_conv_diff) then show "winding_number (c1 +++ reversepath c2) z = d" using wn_cc2 by (simp add: winding_number_join simple_path_imp_path assms znot winding_number_reversepath) qed qed lemma simple_closed_path_wn1: fixes a::complex and e::real assumes "0 < e" and sp_pl: "simple_path(p +++ linepath (a - e) (a + e))" (is "simple_path ?pae") and psp: "pathstart p = a + e" and pfp: "pathfinish p = a - e" and disj: "ball a e \ path_image p = {}" obtains z where "z \ inside (path_image ?pae)" "cmod (winding_number ?pae z) = 1" proof - have "arc p" and arc_lp: "arc (linepath (a - e) (a + e))" and pap: "path_image p \ path_image (linepath (a - e) (a + e)) \ {pathstart p, a-e}" using simple_path_join_loop_eq [of "linepath (a - e) (a + e)" p] assms by auto have mid_eq_a: "midpoint (a - e) (a + e) = a" by (simp add: midpoint_def) with assms have "a \ path_image ?pae" by (simp add: assms path_image_join) (metis midpoint_in_closed_segment) then have "a \ frontier(inside (path_image ?pae))" using assms by (simp add: Jordan_inside_outside ) with \0 < e\ obtain c where c: "c \ inside (path_image ?pae)" and dac: "dist a c < e" by (auto simp: frontier_straddle) then have "c \ path_image ?pae" using inside_no_overlap by blast then have "c \ path_image p" "c \ closed_segment (a - e) (a + e)" by (simp_all add: assms path_image_join) with \0 < e\ dac have "c \ affine hull {a - e, a + e}" by (simp add: segment_as_ball not_le) with \0 < e\ have *: "\ collinear {a - e, c,a + e}" using collinear_3_affine_hull [of "a-e" "a+e"] by (auto simp: insert_commute) have 13: "1/3 + 1/3 + 1/3 = (1::real)" by simp have "(1/3) *\<^sub>R (a - of_real e) + (1/3) *\<^sub>R c + (1/3) *\<^sub>R (a + of_real e) \ interior(convex hull {a - e, c, a + e})" using interior_convex_hull_3_minimal [OF * DIM_complex] by clarsimp (metis 13 zero_less_divide_1_iff zero_less_numeral) then obtain z where z: "z \ interior(convex hull {a - e, c, a + e})" by force have [simp]: "z \ closed_segment (a - e) c" by (metis DIM_complex Diff_iff IntD2 inf_sup_absorb interior_of_triangle z) have [simp]: "z \ closed_segment (a + e) (a - e)" by (metis DIM_complex DiffD2 Un_iff interior_of_triangle z) have [simp]: "z \ closed_segment c (a + e)" by (metis (no_types, lifting) DIM_complex DiffD2 Un_insert_right inf_sup_aci(5) insertCI interior_of_triangle mk_disjoint_insert z) show thesis proof have "norm (winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z) = 1" using winding_number_triangle [OF z] by simp have zin: "z \ inside (path_image (linepath (a + e) (a - e)) \ path_image p)" and zeq: "winding_number (linepath (a + e) (a - e) +++ reversepath p) z = winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z" proof (rule winding_number_from_innerpath [of "linepath (a + e) (a - e)" "a+e" "a-e" p "linepath (a + e) c +++ linepath c (a - e)" z "winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z"]) show sp_aec: "simple_path (linepath (a + e) c +++ linepath c (a - e))" proof (rule arc_imp_simple_path [OF arc_join]) show "arc (linepath (a + e) c)" by (metis \c \ path_image p\ arc_linepath pathstart_in_path_image psp) show "arc (linepath c (a - e))" by (metis \c \ path_image p\ arc_linepath pathfinish_in_path_image pfp) show "path_image (linepath (a + e) c) \ path_image (linepath c (a - e)) \ {pathstart (linepath c (a - e))}" by clarsimp (metis "*" IntI Int_closed_segment closed_segment_commute singleton_iff) qed auto show "simple_path p" using \arc p\ arc_simple_path by blast show sp_ae2: "simple_path (linepath (a + e) (a - e))" using \arc p\ arc_distinct_ends pfp psp by fastforce show pa: "pathfinish (linepath (a + e) (a - e)) = a - e" "pathstart (linepath (a + e) c +++ linepath c (a - e)) = a + e" "pathfinish (linepath (a + e) c +++ linepath c (a - e)) = a - e" "pathstart p = a + e" "pathfinish p = a - e" "pathstart (linepath (a + e) (a - e)) = a + e" by (simp_all add: assms) show 1: "path_image (linepath (a + e) (a - e)) \ path_image p = {a + e, a - e}" proof show "path_image (linepath (a + e) (a - e)) \ path_image p \ {a + e, a - e}" using pap closed_segment_commute psp segment_convex_hull by fastforce show "{a + e, a - e} \ path_image (linepath (a + e) (a - e)) \ path_image p" using pap pathfinish_in_path_image pathstart_in_path_image pfp psp by fastforce qed show 2: "path_image (linepath (a + e) (a - e)) \ path_image (linepath (a + e) c +++ linepath c (a - e)) = {a + e, a - e}" (is "?lhs = ?rhs") proof have "\ collinear {c, a + e, a - e}" using * by (simp add: insert_commute) then have "convex hull {a + e, a - e} \ convex hull {a + e, c} = {a + e}" "convex hull {a + e, a - e} \ convex hull {c, a - e} = {a - e}" by (metis (full_types) Int_closed_segment insert_commute segment_convex_hull)+ then show "?lhs \ ?rhs" by (metis Int_Un_distrib equalityD1 insert_is_Un path_image_join path_image_linepath path_join_eq path_linepath segment_convex_hull simple_path_def sp_aec) show "?rhs \ ?lhs" using segment_convex_hull by (simp add: path_image_join) qed have "path_image p \ path_image (linepath (a + e) c) \ {a + e}" proof (clarsimp simp: path_image_join) fix x assume "x \ path_image p" and x_ac: "x \ closed_segment (a + e) c" then have "dist x a \ e" by (metis IntI all_not_in_conv disj dist_commute mem_ball not_less) with x_ac dac \e > 0\ show "x = a + e" by (auto simp: norm_minus_commute dist_norm closed_segment_eq_open dest: open_segment_furthest_le [where y=a]) qed moreover have "path_image p \ path_image (linepath c (a - e)) \ {a - e}" proof (clarsimp simp: path_image_join) fix x assume "x \ path_image p" and x_ac: "x \ closed_segment c (a - e)" then have "dist x a \ e" by (metis IntI all_not_in_conv disj dist_commute mem_ball not_less) with x_ac dac \e > 0\ show "x = a - e" by (auto simp: norm_minus_commute dist_norm closed_segment_eq_open dest: open_segment_furthest_le [where y=a]) qed ultimately have "path_image p \ path_image (linepath (a + e) c +++ linepath c (a - e)) \ {a + e, a - e}" by (force simp: path_image_join) then show 3: "path_image p \ path_image (linepath (a + e) c +++ linepath c (a - e)) = {a + e, a - e}" using "1" "2" by blast show 4: "path_image (linepath (a + e) c +++ linepath c (a - e)) \ inside (path_image (linepath (a + e) (a - e)) \ path_image p) \ {}" apply (clarsimp simp: path_image_join segment_convex_hull disjoint_iff_not_equal) by (metis (no_types, opaque_lifting) UnI1 Un_commute c closed_segment_commute ends_in_segment(1) path_image_join path_image_linepath pathstart_linepath pfp segment_convex_hull) show zin_inside: "z \ inside (path_image (linepath (a + e) (a - e)) \ path_image (linepath (a + e) c +++ linepath c (a - e)))" proof (simp add: path_image_join) show "z \ inside (closed_segment (a + e) (a - e) \ (closed_segment (a + e) c \ closed_segment c (a - e)))" by (metis z inside_of_triangle DIM_complex Un_commute closed_segment_commute) qed show 5: "winding_number (linepath (a + e) (a - e) +++ reversepath (linepath (a + e) c +++ linepath c (a - e))) z = winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z" by (simp add: reversepath_joinpaths path_image_join winding_number_join) show 6: "winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z \ 0" by (simp add: winding_number_triangle z) show "winding_number (linepath (a + e) (a - e) +++ reversepath p) z = winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z" by (metis 1 2 3 4 5 6 pa sp_aec sp_ae2 \arc p\ \simple_path p\ arc_distinct_ends winding_number_from_innerpath zin_inside) qed (use assms \e > 0\ in auto) show z_inside: "z \ inside (path_image ?pae)" using zin by (simp add: assms path_image_join Un_commute closed_segment_commute) have "cmod (winding_number ?pae z) = cmod ((winding_number(reversepath ?pae) z))" proof (subst winding_number_reversepath) show "path ?pae" using simple_path_imp_path sp_pl by blast show "z \ path_image ?pae" by (metis IntI emptyE inside_no_overlap z_inside) qed (simp add: inside_def) also have "... = cmod (winding_number(linepath (a + e) (a - e) +++ reversepath p) z)" by (simp add: pfp reversepath_joinpaths) also have "... = cmod (winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z)" by (simp add: zeq) also have "... = 1" using z by (simp add: interior_of_triangle winding_number_triangle) finally show "cmod (winding_number ?pae z) = 1" . qed qed lemma simple_closed_path_wn2: fixes a::complex and d e::real assumes "0 < d" "0 < e" and sp_pl: "simple_path(p +++ linepath (a - d) (a + e))" and psp: "pathstart p = a + e" and pfp: "pathfinish p = a - d" obtains z where "z \ inside (path_image (p +++ linepath (a - d) (a + e)))" "cmod (winding_number (p +++ linepath (a - d) (a + e)) z) = 1" proof - have [simp]: "a + of_real x \ closed_segment (a - \) (a - \) \ x \ closed_segment (-\) (-\)" for x \ \::real using closed_segment_translation_eq [of a] by (metis (no_types, opaque_lifting) add_uminus_conv_diff of_real_minus of_real_closed_segment) have [simp]: "a - of_real x \ closed_segment (a + \) (a + \) \ -x \ closed_segment \ \" for x \ \::real by (metis closed_segment_translation_eq diff_conv_add_uminus of_real_closed_segment of_real_minus) have "arc p" and arc_lp: "arc (linepath (a - d) (a + e))" and "path p" and pap: "path_image p \ closed_segment (a - d) (a + e) \ {a+e, a-d}" using simple_path_join_loop_eq [of "linepath (a - d) (a + e)" p] assms arc_imp_path by auto have "0 \ closed_segment (-d) e" using \0 < d\ \0 < e\ closed_segment_eq_real_ivl by auto then have "a \ path_image (linepath (a - d) (a + e))" using of_real_closed_segment [THEN iffD2] by (force dest: closed_segment_translation_eq [of a, THEN iffD2] simp del: of_real_closed_segment) then have "a \ path_image p" using \0 < d\ \0 < e\ pap by auto then obtain k where "0 < k" and k: "ball a k \ (path_image p) = {}" using \0 < e\ \path p\ not_on_path_ball by blast define kde where "kde \ (min k (min d e)) / 2" have "0 < kde" "kde < k" "kde < d" "kde < e" using \0 < k\ \0 < d\ \0 < e\ by (auto simp: kde_def) let ?q = "linepath (a + kde) (a + e) +++ p +++ linepath (a - d) (a - kde)" have "- kde \ closed_segment (-d) e" using \0 < kde\ \kde < d\ \kde < e\ closed_segment_eq_real_ivl by auto then have a_diff_kde: "a - kde \ closed_segment (a - d) (a + e)" using of_real_closed_segment [THEN iffD2] by (force dest: closed_segment_translation_eq [of a, THEN iffD2] simp del: of_real_closed_segment) then have clsub2: "closed_segment (a - d) (a - kde) \ closed_segment (a - d) (a + e)" by (simp add: subset_closed_segment) then have "path_image p \ closed_segment (a - d) (a - kde) \ {a + e, a - d}" using pap by force moreover have "a + e \ path_image p \ closed_segment (a - d) (a - kde)" using \0 < kde\ \kde < d\ \0 < e\ by (auto simp: closed_segment_eq_real_ivl) ultimately have sub_a_diff_d: "path_image p \ closed_segment (a - d) (a - kde) \ {a - d}" by blast have "kde \ closed_segment (-d) e" using \0 < kde\ \kde < d\ \kde < e\ closed_segment_eq_real_ivl by auto then have a_diff_kde: "a + kde \ closed_segment (a - d) (a + e)" using of_real_closed_segment [THEN iffD2] by (force dest: closed_segment_translation_eq [of "a", THEN iffD2] simp del: of_real_closed_segment) then have clsub1: "closed_segment (a + kde) (a + e) \ closed_segment (a - d) (a + e)" by (simp add: subset_closed_segment) then have "closed_segment (a + kde) (a + e) \ path_image p \ {a + e, a - d}" using pap by force moreover have "closed_segment (a + kde) (a + e) \ closed_segment (a - d) (a - kde) = {}" proof (clarsimp intro!: equals0I) fix y assume y1: "y \ closed_segment (a + kde) (a + e)" and y2: "y \ closed_segment (a - d) (a - kde)" obtain u::real where u: "y = a + u" and "0 < u" proof - obtain \ where \: "y = (1 - \) *\<^sub>R (a + kde) + \ *\<^sub>R (a + e)" and "0 \ \" "\ \ 1" using y1 by (auto simp: in_segment) show thesis proof show "y = a + ((1 - \)*kde + \*e)" using \ by (auto simp: scaleR_conv_of_real algebra_simps) have "(1 - \)*kde + \*e \ 0" using \0 < kde\ \0 \ \\ \\ \ 1\ \0 < e\ by auto moreover have "(1 - \)*kde + \*e \ 0" using \0 < kde\ \0 \ \\ \\ \ 1\ \0 < e\ by (auto simp: add_nonneg_eq_0_iff) ultimately show "(1 - \)*kde + \*e > 0" by simp qed qed moreover obtain v::real where v: "y = a + v" and "v \ 0" proof - obtain \ where \: "y = (1 - \) *\<^sub>R (a - d) + \ *\<^sub>R (a - kde)" and "0 \ \" "\ \ 1" using y2 by (auto simp: in_segment) show thesis proof show "y = a + (- ((1 - \)*d + \*kde))" using \ by (force simp: scaleR_conv_of_real algebra_simps) show "- ((1 - \)*d + \*kde) \ 0" using \0 < kde\ \0 \ \\ \\ \ 1\ \0 < d\ by (metis add.left_neutral add_nonneg_nonneg le_diff_eq less_eq_real_def neg_le_0_iff_le split_mult_pos_le) qed qed ultimately show False by auto qed moreover have "a - d \ closed_segment (a + kde) (a + e)" using \0 < kde\ \kde < d\ \0 < e\ by (auto simp: closed_segment_eq_real_ivl) ultimately have sub_a_plus_e: "closed_segment (a + kde) (a + e) \ (path_image p \ closed_segment (a - d) (a - kde)) \ {a + e}" by auto have "kde \ closed_segment (-kde) e" using \0 < kde\ \kde < d\ \kde < e\ closed_segment_eq_real_ivl by auto then have a_add_kde: "a + kde \ closed_segment (a - kde) (a + e)" using of_real_closed_segment [THEN iffD2] by (force dest: closed_segment_translation_eq [of "a", THEN iffD2] simp del: of_real_closed_segment) have "closed_segment (a - kde) (a + kde) \ closed_segment (a + kde) (a + e) = {a + kde}" by (metis a_add_kde Int_closed_segment) moreover have "path_image p \ closed_segment (a - kde) (a + kde) = {}" proof (rule equals0I, clarify) fix y assume "y \ path_image p" "y \ closed_segment (a - kde) (a + kde)" with equals0D [OF k, of y] \0 < kde\ \kde < k\ show False by (auto simp: dist_norm dest: dist_decreases_closed_segment [where c=a]) qed moreover have "- kde \ closed_segment (-d) kde" using \0 < kde\ \kde < d\ \kde < e\ closed_segment_eq_real_ivl by auto then have a_diff_kde': "a - kde \ closed_segment (a - d) (a + kde)" using of_real_closed_segment [THEN iffD2] by (force dest: closed_segment_translation_eq [of a, THEN iffD2] simp del: of_real_closed_segment) then have "closed_segment (a - d) (a - kde) \ closed_segment (a - kde) (a + kde) = {a - kde}" by (metis Int_closed_segment) ultimately have pa_subset_pm_kde: "path_image ?q \ closed_segment (a - kde) (a + kde) \ {a - kde, a + kde}" by (auto simp: path_image_join assms) have ge_kde1: "\y. x = a + of_real y \ y \ kde" if x: "x \ closed_segment (a + kde) (a + e)" for x proof - obtain u where "0 \ u" "u \ 1" and u: "x = (1 - u) *\<^sub>R (a + kde) + u *\<^sub>R (a + e)" using x by (auto simp: in_segment) then have "kde \ (1 - u) * kde + u * e" using \kde < e\ segment_bound_lemma by auto moreover have "x = a + ((1 - u) * kde + u * e)" by (fastforce simp: u algebra_simps scaleR_conv_of_real) ultimately show ?thesis by blast qed have ge_kde2: "\y. x = a + of_real y \ y \ -kde" if x: "x \ closed_segment (a - d) (a - kde)" for x proof - obtain u where "0 \ u" "u \ 1" and u: "x = (1 - u) *\<^sub>R (a - d) + u *\<^sub>R (a - kde)" using x by (auto simp: in_segment) then have "kde \ ((1-u)*d + u*kde)" using \kde < d\ segment_bound_lemma by auto moreover have "x = a - ((1-u)*d + u*kde)" by (fastforce simp: u algebra_simps scaleR_conv_of_real) ultimately show ?thesis by (metis add_uminus_conv_diff neg_le_iff_le of_real_minus) qed have notin_paq: "x \ path_image ?q" if "dist a x < kde" for x proof - have "x \ path_image p" using k \kde < k\ that by force then show ?thesis using that assms by (auto simp: path_image_join dist_norm dest!: ge_kde1 ge_kde2) qed obtain z where zin: "z \ inside (path_image (?q +++ linepath (a - kde) (a + kde)))" and z1: "cmod (winding_number (?q +++ linepath (a - kde) (a + kde)) z) = 1" proof (rule simple_closed_path_wn1 [of kde ?q a]) show "simple_path (?q +++ linepath (a - kde) (a + kde))" proof (intro simple_path_join_loop conjI) show "arc ?q" proof (rule arc_join) show "arc (linepath (a + kde) (a + e))" using \kde < e\ \arc p\ by (force simp: pfp) show "arc (p +++ linepath (a - d) (a - kde))" using \kde < d\ \kde < e\ \arc p\ sub_a_diff_d by (force simp: pfp intro: arc_join) qed (auto simp: psp pfp path_image_join sub_a_plus_e) show "arc (linepath (a - kde) (a + kde))" using \0 < kde\ by auto qed (use pa_subset_pm_kde in auto) qed (use \0 < kde\ notin_paq in auto) have eq: "path_image (?q +++ linepath (a - kde) (a + kde)) = path_image (p +++ linepath (a - d) (a + e))" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" using clsub1 clsub2 apply (auto simp: path_image_join assms) by (meson subsetCE subset_closed_segment) show "?rhs \ ?lhs" apply (simp add: path_image_join assms Un_ac) by (metis Un_closed_segment Un_assoc a_diff_kde a_diff_kde' le_supI2 subset_refl) qed show thesis proof show zzin: "z \ inside (path_image (p +++ linepath (a - d) (a + e)))" by (metis eq zin) then have znotin: "z \ path_image p" by (metis ComplD Un_iff inside_Un_outside path_image_join pathfinish_linepath pathstart_reversepath pfp reversepath_linepath) have znotin_d_kde: "z \ closed_segment (a - d) (a + kde)" by (metis ComplD Un_iff Un_closed_segment a_diff_kde inside_Un_outside path_image_join path_image_linepath pathstart_linepath pfp zzin) have znotin_d_e: "z \ closed_segment (a - d) (a + e)" by (metis IntI UnCI emptyE inside_no_overlap path_image_join path_image_linepath pathstart_linepath pfp zzin) have znotin_kde_e: "z \ closed_segment (a + kde) (a + e)" and znotin_d_kde': "z \ closed_segment (a - d) (a - kde)" using clsub1 clsub2 zzin by (metis (no_types, opaque_lifting) IntI Un_iff emptyE inside_no_overlap path_image_join path_image_linepath pathstart_linepath pfp subsetD)+ have "winding_number (linepath (a - d) (a + e)) z = winding_number (linepath (a - d) (a + kde)) z + winding_number (linepath (a + kde) (a + e)) z" proof (rule winding_number_split_linepath) show "a + complex_of_real kde \ closed_segment (a - d) (a + e)" by (simp add: a_diff_kde) show "z \ closed_segment (a - d) (a + e)" by (metis ComplD Un_iff inside_Un_outside path_image_join path_image_linepath pathstart_linepath pfp zzin) qed also have "... = winding_number (linepath (a + kde) (a + e)) z + (winding_number (linepath (a - d) (a - kde)) z + winding_number (linepath (a - kde) (a + kde)) z)" by (simp add: winding_number_split_linepath [of "a-kde", symmetric] znotin_d_kde a_diff_kde') finally have "winding_number (p +++ linepath (a - d) (a + e)) z = winding_number p z + winding_number (linepath (a + kde) (a + e)) z + (winding_number (linepath (a - d) (a - kde)) z + winding_number (linepath (a - kde) (a + kde)) z)" by (metis (no_types, lifting) ComplD Un_iff \arc p\ add.assoc arc_imp_path eq path_image_join path_join_path_ends path_linepath simple_path_imp_path sp_pl union_with_outside winding_number_join zin) also have "... = winding_number (linepath (a + kde) (a + e)) z + winding_number (p +++ linepath (a - d) (a - kde)) z + winding_number (linepath (a - kde) (a + kde)) z" using \path p\ znotin assms by simp (metis Un_iff Un_closed_segment a_diff_kde' path_image_linepath path_linepath pathstart_linepath winding_number_join znotin_d_kde) also have "... = winding_number ?q z + winding_number (linepath (a - kde) (a + kde)) z" using \path p\ znotin assms by (simp add: path_image_join winding_number_join znotin_kde_e znotin_d_kde') also have "... = winding_number (?q +++ linepath (a - kde) (a + kde)) z" by (metis (mono_tags, lifting) ComplD UnCI \path p\ eq inside_outside path_image_join path_join_eq path_linepath pathfinish_join pathfinish_linepath pathstart_join pathstart_linepath pfp psp winding_number_join zzin) finally have "winding_number (p +++ linepath (a - d) (a + e)) z = winding_number (?q +++ linepath (a - kde) (a + kde)) z" . then show "cmod (winding_number (p +++ linepath (a - d) (a + e)) z) = 1" by (simp add: z1) qed qed lemma simple_closed_path_wn3: fixes p :: "real \ complex" assumes "simple_path p" and loop: "pathfinish p = pathstart p" obtains z where "z \ inside (path_image p)" "cmod (winding_number p z) = 1" proof - have ins: "inside(path_image p) \ {}" "open(inside(path_image p))" "connected(inside(path_image p))" and out: "outside(path_image p) \ {}" "open(outside(path_image p))" "connected(outside(path_image p))" and bo: "bounded(inside(path_image p))" "\ bounded(outside(path_image p))" and ins_out: "inside(path_image p) \ outside(path_image p) = {}" "inside(path_image p) \ outside(path_image p) = - path_image p" and fro: "frontier(inside(path_image p)) = path_image p" "frontier(outside(path_image p)) = path_image p" using Jordan_inside_outside [OF assms] by auto obtain a where a: "a \ inside(path_image p)" using \inside (path_image p) \ {}\ by blast obtain d::real where "0 < d" and d_fro: "a - d \ frontier (inside (path_image p))" and d_int: "\\. \0 \ \; \ < d\ \ (a - \) \ inside (path_image p)" using ray_to_frontier [of "inside (path_image p)" a "-1"] bo ins a interior_eq by (metis ab_group_add_class.ab_diff_conv_add_uminus of_real_def scale_minus_right zero_neq_neg_one) obtain e::real where "0 < e" and e_fro: "a + e \ frontier (inside (path_image p))" and e_int: "\\. \0 \ \; \ < e\ \ (a + \) \ inside (path_image p)" using ray_to_frontier [of "inside (path_image p)" a 1] bo ins a interior_eq by (metis of_real_def zero_neq_one) obtain t0 where "0 \ t0" "t0 \ 1" and pt: "p t0 = a - d" using a d_fro fro by (auto simp: path_image_def) obtain q where "simple_path q" and q_ends: "pathstart q = a - d" "pathfinish q = a - d" and q_eq_p: "path_image q = path_image p" and wn_q_eq_wn_p: "\z. z \ inside(path_image p) \ winding_number q z = winding_number p z" proof show "simple_path (shiftpath t0 p)" by (simp add: pathstart_shiftpath pathfinish_shiftpath simple_path_shiftpath path_image_shiftpath \0 \ t0\ \t0 \ 1\ assms) show "pathstart (shiftpath t0 p) = a - d" using pt by (simp add: \t0 \ 1\ pathstart_shiftpath) show "pathfinish (shiftpath t0 p) = a - d" by (simp add: \0 \ t0\ loop pathfinish_shiftpath pt) show "path_image (shiftpath t0 p) = path_image p" by (simp add: \0 \ t0\ \t0 \ 1\ loop path_image_shiftpath) show "winding_number (shiftpath t0 p) z = winding_number p z" if "z \ inside (path_image p)" for z by (metis ComplD Un_iff \0 \ t0\ \t0 \ 1\ \simple_path p\ atLeastAtMost_iff inside_Un_outside loop simple_path_imp_path that winding_number_shiftpath) qed have ad_not_ae: "a - d \ a + e" by (metis \0 < d\ \0 < e\ add.left_inverse add_left_cancel add_uminus_conv_diff le_add_same_cancel2 less_eq_real_def not_less of_real_add of_real_def of_real_eq_0_iff pt) have ad_ae_q: "{a - d, a + e} \ path_image q" using \path_image q = path_image p\ d_fro e_fro fro(1) by auto have ada: "open_segment (a - d) a \ inside (path_image p)" proof (clarsimp simp: in_segment) fix u::real assume "0 < u" "u < 1" with d_int have "a - (1 - u) * d \ inside (path_image p)" by (metis \0 < d\ add.commute diff_add_cancel left_diff_distrib' less_add_same_cancel2 less_eq_real_def mult.left_neutral zero_less_mult_iff) then show "(1 - u) *\<^sub>R (a - d) + u *\<^sub>R a \ inside (path_image p)" by (simp add: diff_add_eq of_real_def real_vector.scale_right_diff_distrib) qed have aae: "open_segment a (a + e) \ inside (path_image p)" proof (clarsimp simp: in_segment) fix u::real assume "0 < u" "u < 1" with e_int have "a + u * e \ inside (path_image p)" by (meson \0 < e\ less_eq_real_def mult_less_cancel_right2 not_less zero_less_mult_iff) then show "(1 - u) *\<^sub>R a + u *\<^sub>R (a + e) \ inside (path_image p)" by (metis (mono_tags, lifting) add.assoc of_real_mult pth_6 scaleR_collapse scaleR_conv_of_real) qed have "complex_of_real (d * d + (e * e + d * (e + e))) \ 0" using ad_not_ae by (metis \0 < d\ \0 < e\ add_strict_left_mono less_add_same_cancel1 not_sum_squares_lt_zero of_real_eq_0_iff zero_less_double_add_iff_zero_less_single_add zero_less_mult_iff) moreover have "\u>0. u < 1 \ d = u * d + u * e" using \0 < d\ \0 < e\ by (rule_tac x="d / (d+e)" in exI) (simp add: divide_simps scaleR_conv_of_real flip: distrib_left) ultimately have a_in_de: "a \ open_segment (a - d) (a + e)" using ad_not_ae by (simp add: in_segment algebra_simps scaleR_conv_of_real flip: of_real_add of_real_mult) then have "open_segment (a - d) (a + e) \ inside (path_image p)" using ada a aae Un_open_segment [of a "a-d" "a+e"] by blast then have "path_image q \ open_segment (a - d) (a + e) = {}" using inside_no_overlap by (fastforce simp: q_eq_p) with ad_ae_q have paq_Int_cs: "path_image q \ closed_segment (a - d) (a + e) = {a - d, a + e}" by (simp add: closed_segment_eq_open) obtain t where "0 \ t" "t \ 1" and qt: "q t = a + e" using a e_fro fro ad_ae_q by (auto simp: path_defs) then have "t \ 0" by (metis ad_not_ae pathstart_def q_ends(1)) then have "t \ 1" by (metis ad_not_ae pathfinish_def q_ends(2) qt) have q01: "q 0 = a - d" "q 1 = a - d" using q_ends by (auto simp: pathstart_def pathfinish_def) obtain z where zin: "z \ inside (path_image (subpath t 0 q +++ linepath (a - d) (a + e)))" and z1: "cmod (winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z) = 1" proof (rule simple_closed_path_wn2 [of d e "subpath t 0 q" a], simp_all add: q01) show "simple_path (subpath t 0 q +++ linepath (a - d) (a + e))" proof (rule simple_path_join_loop, simp_all add: qt q01) have "inj_on q (closed_segment t 0)" using \0 \ t\ \simple_path q\ \t \ 1\ \t \ 0\ \t \ 1\ by (fastforce simp: simple_path_def loop_free_def inj_on_def closed_segment_eq_real_ivl) then show "arc (subpath t 0 q)" using \0 \ t\ \simple_path q\ \t \ 1\ \t \ 0\ by (simp add: arc_subpath_eq simple_path_imp_path) show "arc (linepath (a - d) (a + e))" by (simp add: ad_not_ae) show "path_image (subpath t 0 q) \ closed_segment (a - d) (a + e) \ {a + e, a - d}" using qt paq_Int_cs \simple_path q\ \0 \ t\ \t \ 1\ by (force simp: dest: rev_subsetD [OF _ path_image_subpath_subset] intro: simple_path_imp_path) qed qed (auto simp: \0 < d\ \0 < e\ qt) have pa01_Un: "path_image (subpath 0 t q) \ path_image (subpath 1 t q) = path_image q" unfolding path_image_subpath using \0 \ t\ \t \ 1\ by (force simp: path_image_def image_iff) with paq_Int_cs have pa_01q: "(path_image (subpath 0 t q) \ path_image (subpath 1 t q)) \ closed_segment (a - d) (a + e) = {a - d, a + e}" by metis have z_notin_ed: "z \ closed_segment (a + e) (a - d)" using zin q01 by (simp add: path_image_join closed_segment_commute inside_def) have z_notin_0t: "z \ path_image (subpath 0 t q)" by (metis (no_types, opaque_lifting) IntI Un_upper1 subsetD empty_iff inside_no_overlap path_image_join path_image_subpath_commute pathfinish_subpath pathstart_def pathstart_linepath q_ends(1) qt subpath_trivial zin) have [simp]: "- winding_number (subpath t 0 q) z = winding_number (subpath 0 t q) z" by (metis \0 \ t\ \simple_path q\ \t \ 1\ atLeastAtMost_iff zero_le_one path_image_subpath_commute path_subpath real_eq_0_iff_le_ge_0 reversepath_subpath simple_path_imp_path winding_number_reversepath z_notin_0t) obtain z_in_q: "z \ inside(path_image q)" and wn_q: "winding_number (subpath 0 t q +++ subpath t 1 q) z = - winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z" proof (rule winding_number_from_innerpath [of "subpath 0 t q" "a-d" "a+e" "subpath 1 t q" "linepath (a - d) (a + e)" z "- winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z"], simp_all add: q01 qt pa01_Un reversepath_subpath) show "simple_path (subpath 0 t q)" "simple_path (subpath 1 t q)" by (simp_all add: \0 \ t\ \simple_path q\ \t \ 1\ \t \ 0\ \t \ 1\ simple_path_subpath) show "simple_path (linepath (a - d) (a + e))" using ad_not_ae by blast show "path_image (subpath 0 t q) \ path_image (subpath 1 t q) = {a - d, a + e}" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" using \0 \ t\ \simple_path q\ \t \ 1\ \t \ 1\ q_ends qt q01 by (force simp: pathfinish_def qt simple_path_def loop_free_def path_image_subpath) show "?rhs \ ?lhs" using \0 \ t\ \t \ 1\ q01 qt by (force simp: path_image_subpath) qed show "path_image (subpath 0 t q) \ closed_segment (a - d) (a + e) = {a - d, a + e}" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" using paq_Int_cs pa01_Un by fastforce show "?rhs \ ?lhs" using \0 \ t\ \t \ 1\ q01 qt by (force simp: path_image_subpath) qed show "path_image (subpath 1 t q) \ closed_segment (a - d) (a + e) = {a - d, a + e}" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (auto simp: pa_01q [symmetric]) show "?rhs \ ?lhs" using \0 \ t\ \t \ 1\ q01 qt by (force simp: path_image_subpath) qed show "closed_segment (a - d) (a + e) \ inside (path_image q) \ {}" using a a_in_de open_closed_segment pa01_Un q_eq_p by fastforce show "z \ inside (path_image (subpath 0 t q) \ closed_segment (a - d) (a + e))" by (metis path_image_join path_image_linepath path_image_subpath_commute pathfinish_subpath pathstart_linepath q01(1) zin) show "winding_number (subpath 0 t q +++ linepath (a + e) (a - d)) z = - winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z" using z_notin_ed z_notin_0t \0 \ t\ \simple_path q\ \t \ 1\ by (simp add: simple_path_imp_path qt q01 path_image_subpath_commute closed_segment_commute winding_number_join winding_number_reversepath [symmetric]) show "- d \ e" using ad_not_ae by auto show "winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z \ 0" using z1 by auto qed show ?thesis proof show "z \ inside (path_image p)" using q_eq_p z_in_q by auto then have [simp]: "z \ path_image q" by (metis disjoint_iff_not_equal inside_no_overlap q_eq_p) have [simp]: "z \ path_image (subpath 1 t q)" using inside_def pa01_Un z_in_q by fastforce have "winding_number(subpath 0 t q +++ subpath t 1 q) z = winding_number(subpath 0 1 q) z" using z_notin_0t \0 \ t\ \simple_path q\ \t \ 1\ by (simp add: simple_path_imp_path qt path_image_subpath_commute winding_number_join winding_number_subpath_combine) with wn_q have "winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z = - winding_number q z" by auto with z1 have "cmod (winding_number q z) = 1" by simp with z1 wn_q_eq_wn_p show "cmod (winding_number p z) = 1" using z1 wn_q_eq_wn_p by (simp add: \z \ inside (path_image p)\) qed qed proposition simple_closed_path_winding_number_inside: assumes "simple_path \" obtains "\z. z \ inside(path_image \) \ winding_number \ z = 1" | "\z. z \ inside(path_image \) \ winding_number \ z = -1" proof (cases "pathfinish \ = pathstart \") case True have "path \" by (simp add: assms simple_path_imp_path) then have const: "winding_number \ constant_on inside(path_image \)" proof (rule winding_number_constant) show "connected (inside(path_image \))" by (simp add: Jordan_inside_outside True assms) qed (use inside_no_overlap True in auto) obtain z where zin: "z \ inside (path_image \)" and z1: "cmod (winding_number \ z) = 1" using simple_closed_path_wn3 [of \] True assms by blast have "winding_number \ z \ \" using zin integer_winding_number [OF \path \\ True] inside_def by blast moreover have "real_of_int x = - 1 \ x = -1" for x by linarith ultimately consider "winding_number \ z = 1" | "winding_number \ z = -1" using z1 by (auto simp: Ints_def abs_if split: if_split_asm) with that const zin show ?thesis unfolding constant_on_def by metis next case False then show ?thesis using inside_simple_curve_imp_closed assms that(2) by blast qed lemma simple_closed_path_abs_winding_number_inside: assumes "simple_path \" "z \ inside(path_image \)" shows "\Re (winding_number \ z)\ = 1" by (metis assms norm_minus_cancel norm_one one_complex.simps(1) real_norm_def simple_closed_path_winding_number_inside uminus_complex.simps(1)) lemma simple_closed_path_norm_winding_number_inside: assumes "simple_path \" "z \ inside(path_image \)" shows "norm (winding_number \ z) = 1" proof - have "pathfinish \ = pathstart \" using assms inside_simple_curve_imp_closed by blast with assms integer_winding_number have "winding_number \ z \ \" by (simp add: inside_def simple_path_def) then show ?thesis by (metis assms norm_minus_cancel norm_one simple_closed_path_winding_number_inside) qed lemma simple_closed_path_winding_number_cases: assumes "simple_path \" "pathfinish \ = pathstart \" "z \ path_image \" shows "winding_number \ z \ {-1,0,1}" proof - consider "z \ inside (path_image \)" | "z \ outside (path_image \)" by (metis ComplI UnE assms(3) inside_Un_outside) then show ?thesis proof cases case 1 then show ?thesis using assms simple_closed_path_winding_number_inside by auto next case 2 then show ?thesis using assms simple_path_def winding_number_zero_in_outside by blast qed qed lemma simple_closed_path_winding_number_pos: "\simple_path \; pathfinish \ = pathstart \; z \ path_image \; 0 < Re(winding_number \ z)\ \ winding_number \ z = 1" using simple_closed_path_winding_number_cases by fastforce subsection \Winding number for rectangular paths\ proposition winding_number_rectpath: assumes "z \ box a1 a3" shows "winding_number (rectpath a1 a3) z = 1" proof - from assms have less: "Re a1 < Re a3" "Im a1 < Im a3" by (auto simp: in_box_complex_iff) define a2 a4 where "a2 = Complex (Re a3) (Im a1)" and "a4 = Complex (Re a1) (Im a3)" let ?l1 = "linepath a1 a2" and ?l2 = "linepath a2 a3" and ?l3 = "linepath a3 a4" and ?l4 = "linepath a4 a1" from assms and less have "z \ path_image (rectpath a1 a3)" by (auto simp: path_image_rectpath_cbox_minus_box) also have "path_image (rectpath a1 a3) = path_image ?l1 \ path_image ?l2 \ path_image ?l3 \ path_image ?l4" by (simp add: rectpath_def Let_def path_image_join Un_assoc a2_def a4_def) finally have "z \ \" . moreover have "\l\{?l1,?l2,?l3,?l4}. Re (winding_number l z) > 0" unfolding ball_simps HOL.simp_thms a2_def a4_def by (intro conjI; (rule winding_number_linepath_pos_lt; (insert assms, auto simp: a2_def a4_def in_box_complex_iff mult_neg_neg))+) ultimately have "Re (winding_number (rectpath a1 a3) z) > 0" by (simp add: winding_number_join path_image_join rectpath_def Let_def a2_def a4_def) thus "winding_number (rectpath a1 a3) z = 1" using assms less by (intro simple_closed_path_winding_number_pos simple_path_rectpath) (auto simp: path_image_rectpath_cbox_minus_box) qed proposition winding_number_rectpath_outside: assumes "Re a1 \ Re a3" "Im a1 \ Im a3" assumes "z \ cbox a1 a3" shows "winding_number (rectpath a1 a3) z = 0" using assms by (intro winding_number_zero_outside[OF _ _ _ assms(3)] path_image_rectpath_subset_cbox) simp_all text\A per-function version for continuous logs, a kind of monodromy\ proposition\<^marker>\tag unimportant\ winding_number_compose_exp: assumes "path p" shows "winding_number (exp \ p) 0 = (pathfinish p - pathstart p) / (2 * of_real pi * \)" proof - obtain e where "0 < e" and e: "\t. t \ {0..1} \ e \ norm(exp(p t))" proof have "closed (path_image (exp \ p))" by (simp add: assms closed_path_image holomorphic_on_exp holomorphic_on_imp_continuous_on path_continuous_image) then show "0 < setdist {0} (path_image (exp \ p))" by (metis exp_not_eq_zero imageE image_comp infdist_eq_setdist infdist_pos_not_in_closed path_defs(4) path_image_nonempty) next fix t::real assume "t \ {0..1}" have "setdist {0} (path_image (exp \ p)) \ dist 0 (exp (p t))" proof (rule setdist_le_dist) show "exp (p t) \ path_image (exp \ p)" using \t \ {0..1}\ path_image_def by fastforce+ qed auto then show "setdist {0} (path_image (exp \ p)) \ cmod (exp (p t))" by simp qed have "bounded (path_image p)" by (simp add: assms bounded_path_image) then obtain B where "0 < B" and B: "path_image p \ cball 0 B" by (meson bounded_pos mem_cball_0 subsetI) let ?B = "cball (0::complex) (B+1)" have "uniformly_continuous_on ?B exp" using holomorphic_on_exp holomorphic_on_imp_continuous_on by (force intro: compact_uniformly_continuous) then obtain d where "d > 0" and d: "\x x'. \x\?B; x'\?B; dist x' x < d\ \ norm (exp x' - exp x) < e" using \e > 0\ by (auto simp: uniformly_continuous_on_def dist_norm) then have "min 1 d > 0" by force then obtain g where pfg: "polynomial_function g" and "g 0 = p 0" "g 1 = p 1" and gless: "\t. t \ {0..1} \ norm(g t - p t) < min 1 d" using path_approx_polynomial_function [OF \path p\] \d > 0\ \0 < e\ unfolding pathfinish_def pathstart_def by blast have "winding_number (exp \ p) 0 = winding_number (exp \ g) 0" proof (rule winding_number_nearby_paths_eq [symmetric]) show "path (exp \ p)" "path (exp \ g)" by (simp_all add: pfg assms holomorphic_on_exp holomorphic_on_imp_continuous_on path_continuous_image path_polynomial_function) next fix t :: "real" assume t: "t \ {0..1}" with gless have "norm(g t - p t) < 1" using min_less_iff_conj by blast moreover have ptB: "norm (p t) \ B" using B t by (force simp: path_image_def) ultimately have "cmod (g t) \ B + 1" by (meson add_mono_thms_linordered_field(4) le_less_trans less_imp_le norm_triangle_sub) with ptB gless t have "cmod ((exp \ g) t - (exp \ p) t) < e" by (auto simp: dist_norm d) with e t show "cmod ((exp \ g) t - (exp \ p) t) < cmod ((exp \ p) t - 0)" by fastforce qed (use \g 0 = p 0\ \g 1 = p 1\ in \auto simp: pathfinish_def pathstart_def\) also have "... = 1 / (of_real (2 * pi) * \) * contour_integral (exp \ g) (\w. 1 / (w - 0))" proof (rule winding_number_valid_path) have "continuous_on (path_image g) (deriv exp)" by (metis DERIV_exp DERIV_imp_deriv continuous_on_cong holomorphic_on_exp holomorphic_on_imp_continuous_on) then show "valid_path (exp \ g)" by (simp add: field_differentiable_within_exp pfg valid_path_compose valid_path_polynomial_function) show "0 \ path_image (exp \ g)" by (auto simp: path_image_def) qed also have "... = 1 / (of_real (2 * pi) * \) * integral {0..1} (\x. vector_derivative g (at x))" proof (simp add: contour_integral_integral, rule integral_cong) fix t :: "real" assume t: "t \ {0..1}" show "vector_derivative (exp \ g) (at t) / exp (g t) = vector_derivative g (at t)" proof - have "(exp \ g has_vector_derivative vector_derivative (exp \ g) (at t)) (at t)" by (meson DERIV_exp differentiable_def field_vector_diff_chain_at has_vector_derivative_def has_vector_derivative_polynomial_function pfg vector_derivative_works) moreover have "(exp \ g has_vector_derivative vector_derivative g (at t) * exp (g t)) (at t)" by (metis DERIV_exp field_vector_diff_chain_at has_vector_derivative_polynomial_function pfg vector_derivative_at) ultimately show ?thesis by (simp add: divide_simps, rule vector_derivative_unique_at) qed qed also have "... = (pathfinish p - pathstart p) / (2 * of_real pi * \)" proof - have "((\x. vector_derivative g (at x)) has_integral g 1 - g 0) {0..1}" by (meson differentiable_at_polynomial_function fundamental_theorem_of_calculus has_vector_derivative_at_within pfg vector_derivative_works zero_le_one) then show ?thesis unfolding pathfinish_def pathstart_def using \g 0 = p 0\ \g 1 = p 1\ by auto qed finally show ?thesis . qed end \ No newline at end of file diff --git a/src/HOL/Transcendental.thy b/src/HOL/Transcendental.thy --- a/src/HOL/Transcendental.thy +++ b/src/HOL/Transcendental.thy @@ -1,7774 +1,7780 @@ (* Title: HOL/Transcendental.thy Author: Jacques D. Fleuriot, University of Cambridge, University of Edinburgh Author: Lawrence C Paulson Author: Jeremy Avigad *) section \Power Series, Transcendental Functions etc.\ theory Transcendental imports Series Deriv NthRoot begin text \A theorem about the factcorial function on the reals.\ lemma square_fact_le_2_fact: "fact n * fact n \ (fact (2 * n) :: real)" proof (induct n) case 0 then show ?case by simp next case (Suc n) have "(fact (Suc n)) * (fact (Suc n)) = of_nat (Suc n) * of_nat (Suc n) * (fact n * fact n :: real)" by (simp add: field_simps) also have "\ \ of_nat (Suc n) * of_nat (Suc n) * fact (2 * n)" by (rule mult_left_mono [OF Suc]) simp also have "\ \ of_nat (Suc (Suc (2 * n))) * of_nat (Suc (2 * n)) * fact (2 * n)" by (rule mult_right_mono)+ (auto simp: field_simps) also have "\ = fact (2 * Suc n)" by (simp add: field_simps) finally show ?case . qed lemma fact_in_Reals: "fact n \ \" by (induction n) auto lemma of_real_fact [simp]: "of_real (fact n) = fact n" by (metis of_nat_fact of_real_of_nat_eq) lemma pochhammer_of_real: "pochhammer (of_real x) n = of_real (pochhammer x n)" by (simp add: pochhammer_prod) lemma norm_fact [simp]: "norm (fact n :: 'a::real_normed_algebra_1) = fact n" proof - have "(fact n :: 'a) = of_real (fact n)" by simp also have "norm \ = fact n" by (subst norm_of_real) simp finally show ?thesis . qed lemma root_test_convergence: fixes f :: "nat \ 'a::banach" assumes f: "(\n. root n (norm (f n))) \ x" \ \could be weakened to lim sup\ and "x < 1" shows "summable f" proof - have "0 \ x" by (rule LIMSEQ_le[OF tendsto_const f]) (auto intro!: exI[of _ 1]) from \x < 1\ obtain z where z: "x < z" "z < 1" by (metis dense) from f \x < z\ have "eventually (\n. root n (norm (f n)) < z) sequentially" by (rule order_tendstoD) then have "eventually (\n. norm (f n) \ z^n) sequentially" using eventually_ge_at_top proof eventually_elim fix n assume less: "root n (norm (f n)) < z" and n: "1 \ n" from power_strict_mono[OF less, of n] n show "norm (f n) \ z ^ n" by simp qed then show "summable f" unfolding eventually_sequentially using z \0 \ x\ by (auto intro!: summable_comparison_test[OF _ summable_geometric]) qed subsection \More facts about binomial coefficients\ text \ These facts could have been proven before, but having real numbers makes the proofs a lot easier. \ lemma central_binomial_odd: "odd n \ n choose (Suc (n div 2)) = n choose (n div 2)" proof - assume "odd n" hence "Suc (n div 2) \ n" by presburger hence "n choose (Suc (n div 2)) = n choose (n - Suc (n div 2))" by (rule binomial_symmetric) also from \odd n\ have "n - Suc (n div 2) = n div 2" by presburger finally show ?thesis . qed lemma binomial_less_binomial_Suc: assumes k: "k < n div 2" shows "n choose k < n choose (Suc k)" proof - from k have k': "k \ n" "Suc k \ n" by simp_all from k' have "real (n choose k) = fact n / (fact k * fact (n - k))" by (simp add: binomial_fact) also from k' have "n - k = Suc (n - Suc k)" by simp also from k' have "fact \ = (real n - real k) * fact (n - Suc k)" by (subst fact_Suc) (simp_all add: of_nat_diff) also from k have "fact k = fact (Suc k) / (real k + 1)" by (simp add: field_simps) also have "fact n / (fact (Suc k) / (real k + 1) * ((real n - real k) * fact (n - Suc k))) = (n choose (Suc k)) * ((real k + 1) / (real n - real k))" using k by (simp add: field_split_simps binomial_fact) also from assms have "(real k + 1) / (real n - real k) < 1" by simp finally show ?thesis using k by (simp add: mult_less_cancel_left) qed lemma binomial_strict_mono: assumes "k < k'" "2*k' \ n" shows "n choose k < n choose k'" proof - from assms have "k \ k' - 1" by simp thus ?thesis proof (induction rule: inc_induct) case base with assms binomial_less_binomial_Suc[of "k' - 1" n] show ?case by simp next case (step k) from step.prems step.hyps assms have "n choose k < n choose (Suc k)" by (intro binomial_less_binomial_Suc) simp_all also have "\ < n choose k'" by (rule step.IH) finally show ?case . qed qed lemma binomial_mono: assumes "k \ k'" "2*k' \ n" shows "n choose k \ n choose k'" using assms binomial_strict_mono[of k k' n] by (cases "k = k'") simp_all lemma binomial_strict_antimono: assumes "k < k'" "2 * k \ n" "k' \ n" shows "n choose k > n choose k'" proof - from assms have "n choose (n - k) > n choose (n - k')" by (intro binomial_strict_mono) (simp_all add: algebra_simps) with assms show ?thesis by (simp add: binomial_symmetric [symmetric]) qed lemma binomial_antimono: assumes "k \ k'" "k \ n div 2" "k' \ n" shows "n choose k \ n choose k'" proof (cases "k = k'") case False note not_eq = False show ?thesis proof (cases "k = n div 2 \ odd n") case False with assms(2) have "2*k \ n" by presburger with not_eq assms binomial_strict_antimono[of k k' n] show ?thesis by simp next case True have "n choose k' \ n choose (Suc (n div 2))" proof (cases "k' = Suc (n div 2)") case False with assms True not_eq have "Suc (n div 2) < k'" by simp with assms binomial_strict_antimono[of "Suc (n div 2)" k' n] True show ?thesis by auto qed simp_all also from True have "\ = n choose k" by (simp add: central_binomial_odd) finally show ?thesis . qed qed simp_all lemma binomial_maximum: "n choose k \ n choose (n div 2)" proof - have "k \ n div 2 \ 2*k \ n" by linarith consider "2*k \ n" | "2*k \ n" "k \ n" | "k > n" by linarith thus ?thesis proof cases case 1 thus ?thesis by (intro binomial_mono) linarith+ next case 2 thus ?thesis by (intro binomial_antimono) simp_all qed (simp_all add: binomial_eq_0) qed lemma binomial_maximum': "(2*n) choose k \ (2*n) choose n" using binomial_maximum[of "2*n"] by simp lemma central_binomial_lower_bound: assumes "n > 0" shows "4^n / (2*real n) \ real ((2*n) choose n)" proof - from binomial[of 1 1 "2*n"] have "4 ^ n = (\k\2*n. (2*n) choose k)" by (simp add: power_mult power2_eq_square One_nat_def [symmetric] del: One_nat_def) also have "{..2*n} = {0<..<2*n} \ {0,2*n}" by auto also have "(\k\\. (2*n) choose k) = (\k\{0<..<2*n}. (2*n) choose k) + (\k\{0,2*n}. (2*n) choose k)" by (subst sum.union_disjoint) auto also have "(\k\{0,2*n}. (2*n) choose k) \ (\k\1. (n choose k)\<^sup>2)" by (cases n) simp_all also from assms have "\ \ (\k\n. (n choose k)\<^sup>2)" by (intro sum_mono2) auto also have "\ = (2*n) choose n" by (rule choose_square_sum) also have "(\k\{0<..<2*n}. (2*n) choose k) \ (\k\{0<..<2*n}. (2*n) choose n)" by (intro sum_mono binomial_maximum') also have "\ = card {0<..<2*n} * ((2*n) choose n)" by simp also have "card {0<..<2*n} \ 2*n - 1" by (cases n) simp_all also have "(2 * n - 1) * (2 * n choose n) + (2 * n choose n) = ((2*n) choose n) * (2*n)" using assms by (simp add: algebra_simps) finally have "4 ^ n \ (2 * n choose n) * (2 * n)" by simp_all hence "real (4 ^ n) \ real ((2 * n choose n) * (2 * n))" by (subst of_nat_le_iff) with assms show ?thesis by (simp add: field_simps) qed subsection \Properties of Power Series\ lemma powser_zero [simp]: "(\n. f n * 0 ^ n) = f 0" for f :: "nat \ 'a::real_normed_algebra_1" proof - have "(\n<1. f n * 0 ^ n) = (\n. f n * 0 ^ n)" by (subst suminf_finite[where N="{0}"]) (auto simp: power_0_left) then show ?thesis by simp qed lemma powser_sums_zero: "(\n. a n * 0^n) sums a 0" for a :: "nat \ 'a::real_normed_div_algebra" using sums_finite [of "{0}" "\n. a n * 0 ^ n"] by simp lemma powser_sums_zero_iff [simp]: "(\n. a n * 0^n) sums x \ a 0 = x" for a :: "nat \ 'a::real_normed_div_algebra" using powser_sums_zero sums_unique2 by blast text \ Power series has a circle or radius of convergence: if it sums for \x\, then it sums absolutely for \z\ with \<^term>\\z\ < \x\\.\ lemma powser_insidea: fixes x z :: "'a::real_normed_div_algebra" assumes 1: "summable (\n. f n * x^n)" and 2: "norm z < norm x" shows "summable (\n. norm (f n * z ^ n))" proof - from 2 have x_neq_0: "x \ 0" by clarsimp from 1 have "(\n. f n * x^n) \ 0" by (rule summable_LIMSEQ_zero) then have "convergent (\n. f n * x^n)" by (rule convergentI) then have "Cauchy (\n. f n * x^n)" by (rule convergent_Cauchy) then have "Bseq (\n. f n * x^n)" by (rule Cauchy_Bseq) then obtain K where 3: "0 < K" and 4: "\n. norm (f n * x^n) \ K" by (auto simp: Bseq_def) have "\N. \n\N. norm (norm (f n * z ^ n)) \ K * norm (z ^ n) * inverse (norm (x^n))" proof (intro exI allI impI) fix n :: nat assume "0 \ n" have "norm (norm (f n * z ^ n)) * norm (x^n) = norm (f n * x^n) * norm (z ^ n)" by (simp add: norm_mult abs_mult) also have "\ \ K * norm (z ^ n)" by (simp only: mult_right_mono 4 norm_ge_zero) also have "\ = K * norm (z ^ n) * (inverse (norm (x^n)) * norm (x^n))" by (simp add: x_neq_0) also have "\ = K * norm (z ^ n) * inverse (norm (x^n)) * norm (x^n)" by (simp only: mult.assoc) finally show "norm (norm (f n * z ^ n)) \ K * norm (z ^ n) * inverse (norm (x^n))" by (simp add: mult_le_cancel_right x_neq_0) qed moreover have "summable (\n. K * norm (z ^ n) * inverse (norm (x^n)))" proof - from 2 have "norm (norm (z * inverse x)) < 1" using x_neq_0 by (simp add: norm_mult nonzero_norm_inverse divide_inverse [where 'a=real, symmetric]) then have "summable (\n. norm (z * inverse x) ^ n)" by (rule summable_geometric) then have "summable (\n. K * norm (z * inverse x) ^ n)" by (rule summable_mult) then show "summable (\n. K * norm (z ^ n) * inverse (norm (x^n)))" using x_neq_0 by (simp add: norm_mult nonzero_norm_inverse power_mult_distrib power_inverse norm_power mult.assoc) qed ultimately show "summable (\n. norm (f n * z ^ n))" by (rule summable_comparison_test) qed lemma powser_inside: fixes f :: "nat \ 'a::{real_normed_div_algebra,banach}" shows "summable (\n. f n * (x^n)) \ norm z < norm x \ summable (\n. f n * (z ^ n))" by (rule powser_insidea [THEN summable_norm_cancel]) lemma powser_times_n_limit_0: fixes x :: "'a::{real_normed_div_algebra,banach}" assumes "norm x < 1" shows "(\n. of_nat n * x ^ n) \ 0" proof - have "norm x / (1 - norm x) \ 0" using assms by (auto simp: field_split_simps) moreover obtain N where N: "norm x / (1 - norm x) < of_int N" using ex_le_of_int by (meson ex_less_of_int) ultimately have N0: "N>0" by auto then have *: "real_of_int (N + 1) * norm x / real_of_int N < 1" using N assms by (auto simp: field_simps) have **: "real_of_int N * (norm x * (real_of_nat (Suc n) * norm (x ^ n))) \ real_of_nat n * (norm x * ((1 + N) * norm (x ^ n)))" if "N \ int n" for n :: nat proof - from that have "real_of_int N * real_of_nat (Suc n) \ real_of_nat n * real_of_int (1 + N)" by (simp add: algebra_simps) then have "(real_of_int N * real_of_nat (Suc n)) * (norm x * norm (x ^ n)) \ (real_of_nat n * (1 + N)) * (norm x * norm (x ^ n))" using N0 mult_mono by fastforce then show ?thesis by (simp add: algebra_simps) qed show ?thesis using * by (rule summable_LIMSEQ_zero [OF summable_ratio_test, where N1="nat N"]) (simp add: N0 norm_mult field_simps ** del: of_nat_Suc of_int_add) qed corollary lim_n_over_pown: fixes x :: "'a::{real_normed_field,banach}" shows "1 < norm x \ ((\n. of_nat n / x^n) \ 0) sequentially" using powser_times_n_limit_0 [of "inverse x"] by (simp add: norm_divide field_split_simps) lemma sum_split_even_odd: fixes f :: "nat \ real" shows "(\i<2 * n. if even i then f i else g i) = (\iii<2 * Suc n. if even i then f i else g i) = (\ii = (\ii real" assumes "g sums x" shows "(\ n. if even n then 0 else g ((n - 1) div 2)) sums x" unfolding sums_def proof (rule LIMSEQ_I) fix r :: real assume "0 < r" from \g sums x\[unfolded sums_def, THEN LIMSEQ_D, OF this] obtain no where no_eq: "\n. n \ no \ (norm (sum g {.. 2 * no" for m proof - from that have "m div 2 \ no" by auto have sum_eq: "?SUM (2 * (m div 2)) = sum g {..< m div 2}" using sum_split_even_odd by auto then have "(norm (?SUM (2 * (m div 2)) - x) < r)" using no_eq unfolding sum_eq using \m div 2 \ no\ by auto moreover have "?SUM (2 * (m div 2)) = ?SUM m" proof (cases "even m") case True then show ?thesis by (auto simp: even_two_times_div_two) next case False then have eq: "Suc (2 * (m div 2)) = m" by simp then have "even (2 * (m div 2))" using \odd m\ by auto have "?SUM m = ?SUM (Suc (2 * (m div 2)))" unfolding eq .. also have "\ = ?SUM (2 * (m div 2))" using \even (2 * (m div 2))\ by auto finally show ?thesis by auto qed ultimately show ?thesis by auto qed then show "\no. \ m \ no. norm (?SUM m - x) < r" by blast qed lemma sums_if: fixes g :: "nat \ real" assumes "g sums x" and "f sums y" shows "(\ n. if even n then f (n div 2) else g ((n - 1) div 2)) sums (x + y)" proof - let ?s = "\ n. if even n then 0 else f ((n - 1) div 2)" have if_sum: "(if B then (0 :: real) else E) + (if B then T else 0) = (if B then T else E)" for B T E by (cases B) auto have g_sums: "(\ n. if even n then 0 else g ((n - 1) div 2)) sums x" using sums_if'[OF \g sums x\] . have if_eq: "\B T E. (if \ B then T else E) = (if B then E else T)" by auto have "?s sums y" using sums_if'[OF \f sums y\] . from this[unfolded sums_def, THEN LIMSEQ_Suc] have "(\n. if even n then f (n div 2) else 0) sums y" by (simp add: lessThan_Suc_eq_insert_0 sum.atLeast1_atMost_eq image_Suc_lessThan if_eq sums_def cong del: if_weak_cong) from sums_add[OF g_sums this] show ?thesis by (simp only: if_sum) qed subsection \Alternating series test / Leibniz formula\ (* FIXME: generalise these results from the reals via type classes? *) lemma sums_alternating_upper_lower: fixes a :: "nat \ real" assumes mono: "\n. a (Suc n) \ a n" and a_pos: "\n. 0 \ a n" and "a \ 0" shows "\l. ((\n. (\i<2*n. (- 1)^i*a i) \ l) \ (\ n. \i<2*n. (- 1)^i*a i) \ l) \ ((\n. l \ (\i<2*n + 1. (- 1)^i*a i)) \ (\ n. \i<2*n + 1. (- 1)^i*a i) \ l)" (is "\l. ((\n. ?f n \ l) \ _) \ ((\n. l \ ?g n) \ _)") proof (rule nested_sequence_unique) have fg_diff: "\n. ?f n - ?g n = - a (2 * n)" by auto show "\n. ?f n \ ?f (Suc n)" proof show "?f n \ ?f (Suc n)" for n using mono[of "2*n"] by auto qed show "\n. ?g (Suc n) \ ?g n" proof show "?g (Suc n) \ ?g n" for n using mono[of "Suc (2*n)"] by auto qed show "\n. ?f n \ ?g n" proof show "?f n \ ?g n" for n using fg_diff a_pos by auto qed show "(\n. ?f n - ?g n) \ 0" unfolding fg_diff proof (rule LIMSEQ_I) fix r :: real assume "0 < r" with \a \ 0\[THEN LIMSEQ_D] obtain N where "\ n. n \ N \ norm (a n - 0) < r" by auto then have "\n \ N. norm (- a (2 * n) - 0) < r" by auto then show "\N. \n \ N. norm (- a (2 * n) - 0) < r" by auto qed qed lemma summable_Leibniz': fixes a :: "nat \ real" assumes a_zero: "a \ 0" and a_pos: "\n. 0 \ a n" and a_monotone: "\n. a (Suc n) \ a n" shows summable: "summable (\ n. (-1)^n * a n)" and "\n. (\i<2*n. (-1)^i*a i) \ (\i. (-1)^i*a i)" and "(\n. \i<2*n. (-1)^i*a i) \ (\i. (-1)^i*a i)" and "\n. (\i. (-1)^i*a i) \ (\i<2*n+1. (-1)^i*a i)" and "(\n. \i<2*n+1. (-1)^i*a i) \ (\i. (-1)^i*a i)" proof - let ?S = "\n. (-1)^n * a n" let ?P = "\n. \i n. ?f n \ l" and "?f \ l" and above_l: "\ n. l \ ?g n" and "?g \ l" using sums_alternating_upper_lower[OF a_monotone a_pos a_zero] by blast let ?Sa = "\m. \n l" proof (rule LIMSEQ_I) fix r :: real assume "0 < r" with \?f \ l\[THEN LIMSEQ_D] obtain f_no where f: "\n. n \ f_no \ norm (?f n - l) < r" by auto from \0 < r\ \?g \ l\[THEN LIMSEQ_D] obtain g_no where g: "\n. n \ g_no \ norm (?g n - l) < r" by auto have "norm (?Sa n - l) < r" if "n \ (max (2 * f_no) (2 * g_no))" for n proof - from that have "n \ 2 * f_no" and "n \ 2 * g_no" by auto show ?thesis proof (cases "even n") case True then have n_eq: "2 * (n div 2) = n" by (simp add: even_two_times_div_two) with \n \ 2 * f_no\ have "n div 2 \ f_no" by auto from f[OF this] show ?thesis unfolding n_eq atLeastLessThanSuc_atLeastAtMost . next case False then have "even (n - 1)" by simp then have n_eq: "2 * ((n - 1) div 2) = n - 1" by (simp add: even_two_times_div_two) then have range_eq: "n - 1 + 1 = n" using odd_pos[OF False] by auto from n_eq \n \ 2 * g_no\ have "(n - 1) div 2 \ g_no" by auto from g[OF this] show ?thesis by (simp only: n_eq range_eq) qed qed then show "\no. \n \ no. norm (?Sa n - l) < r" by blast qed then have sums_l: "(\i. (-1)^i * a i) sums l" by (simp only: sums_def) then show "summable ?S" by (auto simp: summable_def) have "l = suminf ?S" by (rule sums_unique[OF sums_l]) fix n show "suminf ?S \ ?g n" unfolding sums_unique[OF sums_l, symmetric] using above_l by auto show "?f n \ suminf ?S" unfolding sums_unique[OF sums_l, symmetric] using below_l by auto show "?g \ suminf ?S" using \?g \ l\ \l = suminf ?S\ by auto show "?f \ suminf ?S" using \?f \ l\ \l = suminf ?S\ by auto qed theorem summable_Leibniz: fixes a :: "nat \ real" assumes a_zero: "a \ 0" and "monoseq a" shows "summable (\ n. (-1)^n * a n)" (is "?summable") and "0 < a 0 \ (\n. (\i. (- 1)^i*a i) \ { \i<2*n. (- 1)^i * a i .. \i<2*n+1. (- 1)^i * a i})" (is "?pos") and "a 0 < 0 \ (\n. (\i. (- 1)^i*a i) \ { \i<2*n+1. (- 1)^i * a i .. \i<2*n. (- 1)^i * a i})" (is "?neg") and "(\n. \i<2*n. (- 1)^i*a i) \ (\i. (- 1)^i*a i)" (is "?f") and "(\n. \i<2*n+1. (- 1)^i*a i) \ (\i. (- 1)^i*a i)" (is "?g") proof - have "?summable \ ?pos \ ?neg \ ?f \ ?g" proof (cases "(\n. 0 \ a n) \ (\m. \n\m. a n \ a m)") case True then have ord: "\n m. m \ n \ a n \ a m" and ge0: "\n. 0 \ a n" by auto have mono: "a (Suc n) \ a n" for n using ord[where n="Suc n" and m=n] by auto note leibniz = summable_Leibniz'[OF \a \ 0\ ge0] from leibniz[OF mono] show ?thesis using \0 \ a 0\ by auto next let ?a = "\n. - a n" case False with monoseq_le[OF \monoseq a\ \a \ 0\] have "(\ n. a n \ 0) \ (\m. \n\m. a m \ a n)" by auto then have ord: "\n m. m \ n \ ?a n \ ?a m" and ge0: "\ n. 0 \ ?a n" by auto have monotone: "?a (Suc n) \ ?a n" for n using ord[where n="Suc n" and m=n] by auto note leibniz = summable_Leibniz'[OF _ ge0, of "\x. x", OF tendsto_minus[OF \a \ 0\, unfolded minus_zero] monotone] have "summable (\ n. (-1)^n * ?a n)" using leibniz(1) by auto then obtain l where "(\ n. (-1)^n * ?a n) sums l" unfolding summable_def by auto from this[THEN sums_minus] have "(\ n. (-1)^n * a n) sums -l" by auto then have ?summable by (auto simp: summable_def) moreover have "\- a - - b\ = \a - b\" for a b :: real unfolding minus_diff_minus by auto from suminf_minus[OF leibniz(1), unfolded mult_minus_right minus_minus] have move_minus: "(\n. - ((- 1) ^ n * a n)) = - (\n. (- 1) ^ n * a n)" by auto have ?pos using \0 \ ?a 0\ by auto moreover have ?neg using leibniz(2,4) unfolding mult_minus_right sum_negf move_minus neg_le_iff_le by auto moreover have ?f and ?g using leibniz(3,5)[unfolded mult_minus_right sum_negf move_minus, THEN tendsto_minus_cancel] by auto ultimately show ?thesis by auto qed then show ?summable and ?pos and ?neg and ?f and ?g by safe qed subsection \Term-by-Term Differentiability of Power Series\ definition diffs :: "(nat \ 'a::ring_1) \ nat \ 'a" where "diffs c = (\n. of_nat (Suc n) * c (Suc n))" text \Lemma about distributing negation over it.\ lemma diffs_minus: "diffs (\n. - c n) = (\n. - diffs c n)" by (simp add: diffs_def) lemma diffs_equiv: fixes x :: "'a::{real_normed_vector,ring_1}" shows "summable (\n. diffs c n * x^n) \ (\n. of_nat n * c n * x^(n - Suc 0)) sums (\n. diffs c n * x^n)" unfolding diffs_def by (simp add: summable_sums sums_Suc_imp) lemma lemma_termdiff1: fixes z :: "'a :: {monoid_mult,comm_ring}" shows "(\ppi 0" shows "((z + h) ^ n - z ^ n) / h - of_nat n * z ^ (n - Suc 0) = h * (\p< n - Suc 0. \q< n - Suc 0 - p. (z + h) ^ q * z ^ (n - 2 - q))" (is "?lhs = ?rhs") proof (cases n) case (Suc m) have 0: "\x k. (\njiijppip::nat. p < n \ f p \ K" and K: "0 \ K" shows "sum f {.. of_nat n * K" proof - have "sum f {.. (\i of_nat n * K" by (auto simp: mult_right_mono K) finally show ?thesis . qed lemma lemma_termdiff3: fixes h z :: "'a::real_normed_field" assumes 1: "h \ 0" and 2: "norm z \ K" and 3: "norm (z + h) \ K" shows "norm (((z + h) ^ n - z ^ n) / h - of_nat n * z ^ (n - Suc 0)) \ of_nat n * of_nat (n - Suc 0) * K ^ (n - 2) * norm h" proof - have "norm (((z + h) ^ n - z ^ n) / h - of_nat n * z ^ (n - Suc 0)) = norm (\pq \ of_nat n * (of_nat (n - Suc 0) * K ^ (n - 2)) * norm h" proof (rule mult_right_mono [OF _ norm_ge_zero]) from norm_ge_zero 2 have K: "0 \ K" by (rule order_trans) have le_Kn: "norm ((z + h) ^ i * z ^ j) \ K ^ n" if "i + j = n" for i j n proof - have "norm (z + h) ^ i * norm z ^ j \ K ^ i * K ^ j" by (intro mult_mono power_mono 2 3 norm_ge_zero zero_le_power K) also have "... = K^n" by (metis power_add that) finally show ?thesis by (simp add: norm_mult norm_power) qed then have "\p q. \p < n; q < n - Suc 0\ \ norm ((z + h) ^ q * z ^ (n - 2 - q)) \ K ^ (n - 2)" by (simp del: subst_all) then show "norm (\pq of_nat n * (of_nat (n - Suc 0) * K ^ (n - 2))" by (intro order_trans [OF norm_sum] real_sum_nat_ivl_bounded2 mult_nonneg_nonneg of_nat_0_le_iff zero_le_power K) qed also have "\ = of_nat n * of_nat (n - Suc 0) * K ^ (n - 2) * norm h" by (simp only: mult.assoc) finally show ?thesis . qed lemma lemma_termdiff4: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" and k :: real assumes k: "0 < k" and le: "\h. h \ 0 \ norm h < k \ norm (f h) \ K * norm h" shows "f \0\ 0" proof (rule tendsto_norm_zero_cancel) show "(\h. norm (f h)) \0\ 0" proof (rule real_tendsto_sandwich) show "eventually (\h. 0 \ norm (f h)) (at 0)" by simp show "eventually (\h. norm (f h) \ K * norm h) (at 0)" using k by (auto simp: eventually_at dist_norm le) show "(\h. 0) \(0::'a)\ (0::real)" by (rule tendsto_const) have "(\h. K * norm h) \(0::'a)\ K * norm (0::'a)" by (intro tendsto_intros) then show "(\h. K * norm h) \(0::'a)\ 0" by simp qed qed lemma lemma_termdiff5: fixes g :: "'a::real_normed_vector \ nat \ 'b::banach" and k :: real assumes k: "0 < k" and f: "summable f" and le: "\h n. h \ 0 \ norm h < k \ norm (g h n) \ f n * norm h" shows "(\h. suminf (g h)) \0\ 0" proof (rule lemma_termdiff4 [OF k]) fix h :: 'a assume "h \ 0" and "norm h < k" then have 1: "\n. norm (g h n) \ f n * norm h" by (simp add: le) then have "\N. \n\N. norm (norm (g h n)) \ f n * norm h" by simp moreover from f have 2: "summable (\n. f n * norm h)" by (rule summable_mult2) ultimately have 3: "summable (\n. norm (g h n))" by (rule summable_comparison_test) then have "norm (suminf (g h)) \ (\n. norm (g h n))" by (rule summable_norm) also from 1 3 2 have "(\n. norm (g h n)) \ (\n. f n * norm h)" by (simp add: suminf_le) also from f have "(\n. f n * norm h) = suminf f * norm h" by (rule suminf_mult2 [symmetric]) finally show "norm (suminf (g h)) \ suminf f * norm h" . qed (* FIXME: Long proofs *) lemma termdiffs_aux: fixes x :: "'a::{real_normed_field,banach}" assumes 1: "summable (\n. diffs (diffs c) n * K ^ n)" and 2: "norm x < norm K" shows "(\h. \n. c n * (((x + h) ^ n - x^n) / h - of_nat n * x ^ (n - Suc 0))) \0\ 0" proof - from dense [OF 2] obtain r where r1: "norm x < r" and r2: "r < norm K" by fast from norm_ge_zero r1 have r: "0 < r" by (rule order_le_less_trans) then have r_neq_0: "r \ 0" by simp show ?thesis proof (rule lemma_termdiff5) show "0 < r - norm x" using r1 by simp from r r2 have "norm (of_real r::'a) < norm K" by simp with 1 have "summable (\n. norm (diffs (diffs c) n * (of_real r ^ n)))" by (rule powser_insidea) then have "summable (\n. diffs (diffs (\n. norm (c n))) n * r ^ n)" using r by (simp add: diffs_def norm_mult norm_power del: of_nat_Suc) then have "summable (\n. of_nat n * diffs (\n. norm (c n)) n * r ^ (n - Suc 0))" by (rule diffs_equiv [THEN sums_summable]) also have "(\n. of_nat n * diffs (\n. norm (c n)) n * r ^ (n - Suc 0)) = (\n. diffs (\m. of_nat (m - Suc 0) * norm (c m) * inverse r) n * (r ^ n))" by (simp add: diffs_def r_neq_0 fun_eq_iff split: nat_diff_split) finally have "summable (\n. of_nat n * (of_nat (n - Suc 0) * norm (c n) * inverse r) * r ^ (n - Suc 0))" by (rule diffs_equiv [THEN sums_summable]) also have "(\n. of_nat n * (of_nat (n - Suc 0) * norm (c n) * inverse r) * r ^ (n - Suc 0)) = (\n. norm (c n) * of_nat n * of_nat (n - Suc 0) * r ^ (n - 2))" by (rule ext) (simp add: r_neq_0 split: nat_diff_split) finally show "summable (\n. norm (c n) * of_nat n * of_nat (n - Suc 0) * r ^ (n - 2))" . next fix h :: 'a and n assume h: "h \ 0" assume "norm h < r - norm x" then have "norm x + norm h < r" by simp with norm_triangle_ineq have xh: "norm (x + h) < r" by (rule order_le_less_trans) have "norm (((x + h) ^ n - x ^ n) / h - of_nat n * x ^ (n - Suc 0)) \ real n * (real (n - Suc 0) * (r ^ (n - 2) * norm h))" by (metis (mono_tags, lifting) h mult.assoc lemma_termdiff3 less_eq_real_def r1 xh) then show "norm (c n * (((x + h) ^ n - x^n) / h - of_nat n * x ^ (n - Suc 0))) \ norm (c n) * of_nat n * of_nat (n - Suc 0) * r ^ (n - 2) * norm h" by (simp only: norm_mult mult.assoc mult_left_mono [OF _ norm_ge_zero]) qed qed lemma termdiffs: fixes K x :: "'a::{real_normed_field,banach}" assumes 1: "summable (\n. c n * K ^ n)" and 2: "summable (\n. (diffs c) n * K ^ n)" and 3: "summable (\n. (diffs (diffs c)) n * K ^ n)" and 4: "norm x < norm K" shows "DERIV (\x. \n. c n * x^n) x :> (\n. (diffs c) n * x^n)" unfolding DERIV_def proof (rule LIM_zero_cancel) show "(\h. (suminf (\n. c n * (x + h) ^ n) - suminf (\n. c n * x^n)) / h - suminf (\n. diffs c n * x^n)) \0\ 0" proof (rule LIM_equal2) show "0 < norm K - norm x" using 4 by (simp add: less_diff_eq) next fix h :: 'a assume "norm (h - 0) < norm K - norm x" then have "norm x + norm h < norm K" by simp then have 5: "norm (x + h) < norm K" by (rule norm_triangle_ineq [THEN order_le_less_trans]) have "summable (\n. c n * x^n)" and "summable (\n. c n * (x + h) ^ n)" and "summable (\n. diffs c n * x^n)" using 1 2 4 5 by (auto elim: powser_inside) then have "((\n. c n * (x + h) ^ n) - (\n. c n * x^n)) / h - (\n. diffs c n * x^n) = (\n. (c n * (x + h) ^ n - c n * x^n) / h - of_nat n * c n * x ^ (n - Suc 0))" by (intro sums_unique sums_diff sums_divide diffs_equiv summable_sums) then show "((\n. c n * (x + h) ^ n) - (\n. c n * x^n)) / h - (\n. diffs c n * x^n) = (\n. c n * (((x + h) ^ n - x^n) / h - of_nat n * x ^ (n - Suc 0)))" by (simp add: algebra_simps) next show "(\h. \n. c n * (((x + h) ^ n - x^n) / h - of_nat n * x ^ (n - Suc 0))) \0\ 0" by (rule termdiffs_aux [OF 3 4]) qed qed subsection \The Derivative of a Power Series Has the Same Radius of Convergence\ lemma termdiff_converges: fixes x :: "'a::{real_normed_field,banach}" assumes K: "norm x < K" and sm: "\x. norm x < K \ summable(\n. c n * x ^ n)" shows "summable (\n. diffs c n * x ^ n)" proof (cases "x = 0") case True then show ?thesis using powser_sums_zero sums_summable by auto next case False then have "K > 0" using K less_trans zero_less_norm_iff by blast then obtain r :: real where r: "norm x < norm r" "norm r < K" "r > 0" using K False by (auto simp: field_simps abs_less_iff add_pos_pos intro: that [of "(norm x + K) / 2"]) have to0: "(\n. of_nat n * (x / of_real r) ^ n) \ 0" using r by (simp add: norm_divide powser_times_n_limit_0 [of "x / of_real r"]) obtain N where N: "\n. n\N \ real_of_nat n * norm x ^ n < r ^ n" using r LIMSEQ_D [OF to0, of 1] by (auto simp: norm_divide norm_mult norm_power field_simps) have "summable (\n. (of_nat n * c n) * x ^ n)" proof (rule summable_comparison_test') show "summable (\n. norm (c n * of_real r ^ n))" apply (rule powser_insidea [OF sm [of "of_real ((r+K)/2)"]]) using N r norm_of_real [of "r + K", where 'a = 'a] by auto show "\n. N \ n \ norm (of_nat n * c n * x ^ n) \ norm (c n * of_real r ^ n)" using N r by (fastforce simp add: norm_mult norm_power less_eq_real_def) qed then have "summable (\n. (of_nat (Suc n) * c(Suc n)) * x ^ Suc n)" using summable_iff_shift [of "\n. of_nat n * c n * x ^ n" 1] by simp then have "summable (\n. (of_nat (Suc n) * c(Suc n)) * x ^ n)" using False summable_mult2 [of "\n. (of_nat (Suc n) * c(Suc n) * x ^ n) * x" "inverse x"] by (simp add: mult.assoc) (auto simp: ac_simps) then show ?thesis by (simp add: diffs_def) qed lemma termdiff_converges_all: fixes x :: "'a::{real_normed_field,banach}" assumes "\x. summable (\n. c n * x^n)" shows "summable (\n. diffs c n * x^n)" by (rule termdiff_converges [where K = "1 + norm x"]) (use assms in auto) lemma termdiffs_strong: fixes K x :: "'a::{real_normed_field,banach}" assumes sm: "summable (\n. c n * K ^ n)" and K: "norm x < norm K" shows "DERIV (\x. \n. c n * x^n) x :> (\n. diffs c n * x^n)" proof - have "norm K + norm x < norm K + norm K" using K by force then have K2: "norm ((of_real (norm K) + of_real (norm x)) / 2 :: 'a) < norm K" by (auto simp: norm_triangle_lt norm_divide field_simps) then have [simp]: "norm ((of_real (norm K) + of_real (norm x)) :: 'a) < norm K * 2" by simp have "summable (\n. c n * (of_real (norm x + norm K) / 2) ^ n)" by (metis K2 summable_norm_cancel [OF powser_insidea [OF sm]] add.commute of_real_add) moreover have "\x. norm x < norm K \ summable (\n. diffs c n * x ^ n)" by (blast intro: sm termdiff_converges powser_inside) moreover have "\x. norm x < norm K \ summable (\n. diffs(diffs c) n * x ^ n)" by (blast intro: sm termdiff_converges powser_inside) ultimately show ?thesis by (rule termdiffs [where K = "of_real (norm x + norm K) / 2"]) (use K in \auto simp: field_simps simp flip: of_real_add\) qed lemma termdiffs_strong_converges_everywhere: fixes K x :: "'a::{real_normed_field,banach}" assumes "\y. summable (\n. c n * y ^ n)" shows "((\x. \n. c n * x^n) has_field_derivative (\n. diffs c n * x^n)) (at x)" using termdiffs_strong[OF assms[of "of_real (norm x + 1)"], of x] by (force simp del: of_real_add) lemma termdiffs_strong': fixes z :: "'a :: {real_normed_field,banach}" assumes "\z. norm z < K \ summable (\n. c n * z ^ n)" assumes "norm z < K" shows "((\z. \n. c n * z^n) has_field_derivative (\n. diffs c n * z^n)) (at z)" proof (rule termdiffs_strong) define L :: real where "L = (norm z + K) / 2" have "0 \ norm z" by simp also note \norm z < K\ finally have K: "K \ 0" by simp from assms K have L: "L \ 0" "norm z < L" "L < K" by (simp_all add: L_def) from L show "norm z < norm (of_real L :: 'a)" by simp from L show "summable (\n. c n * of_real L ^ n)" by (intro assms(1)) simp_all qed lemma termdiffs_sums_strong: fixes z :: "'a :: {banach,real_normed_field}" assumes sums: "\z. norm z < K \ (\n. c n * z ^ n) sums f z" assumes deriv: "(f has_field_derivative f') (at z)" assumes norm: "norm z < K" shows "(\n. diffs c n * z ^ n) sums f'" proof - have summable: "summable (\n. diffs c n * z^n)" by (intro termdiff_converges[OF norm] sums_summable[OF sums]) from norm have "eventually (\z. z \ norm -` {..z. (\n. c n * z^n) = f z) (nhds z)" by eventually_elim (insert sums, simp add: sums_iff) have "((\z. \n. c n * z^n) has_field_derivative (\n. diffs c n * z^n)) (at z)" by (intro termdiffs_strong'[OF _ norm] sums_summable[OF sums]) hence "(f has_field_derivative (\n. diffs c n * z^n)) (at z)" by (subst (asm) DERIV_cong_ev[OF refl eq refl]) from this and deriv have "(\n. diffs c n * z^n) = f'" by (rule DERIV_unique) with summable show ?thesis by (simp add: sums_iff) qed lemma isCont_powser: fixes K x :: "'a::{real_normed_field,banach}" assumes "summable (\n. c n * K ^ n)" assumes "norm x < norm K" shows "isCont (\x. \n. c n * x^n) x" using termdiffs_strong[OF assms] by (blast intro!: DERIV_isCont) lemmas isCont_powser' = isCont_o2[OF _ isCont_powser] lemma isCont_powser_converges_everywhere: fixes K x :: "'a::{real_normed_field,banach}" assumes "\y. summable (\n. c n * y ^ n)" shows "isCont (\x. \n. c n * x^n) x" using termdiffs_strong[OF assms[of "of_real (norm x + 1)"], of x] by (force intro!: DERIV_isCont simp del: of_real_add) lemma powser_limit_0: fixes a :: "nat \ 'a::{real_normed_field,banach}" assumes s: "0 < s" and sm: "\x. norm x < s \ (\n. a n * x ^ n) sums (f x)" shows "(f \ a 0) (at 0)" proof - have "norm (of_real s / 2 :: 'a) < s" using s by (auto simp: norm_divide) then have "summable (\n. a n * (of_real s / 2) ^ n)" by (rule sums_summable [OF sm]) then have "((\x. \n. a n * x ^ n) has_field_derivative (\n. diffs a n * 0 ^ n)) (at 0)" by (rule termdiffs_strong) (use s in \auto simp: norm_divide\) then have "isCont (\x. \n. a n * x ^ n) 0" by (blast intro: DERIV_continuous) then have "((\x. \n. a n * x ^ n) \ a 0) (at 0)" by (simp add: continuous_within) moreover have "(\x. f x - (\n. a n * x ^ n)) \0\ 0" apply (clarsimp simp: LIM_eq) apply (rule_tac x=s in exI) using s sm sums_unique by fastforce ultimately show ?thesis by (rule Lim_transform) qed lemma powser_limit_0_strong: fixes a :: "nat \ 'a::{real_normed_field,banach}" assumes s: "0 < s" and sm: "\x. x \ 0 \ norm x < s \ (\n. a n * x ^ n) sums (f x)" shows "(f \ a 0) (at 0)" proof - have *: "((\x. if x = 0 then a 0 else f x) \ a 0) (at 0)" by (rule powser_limit_0 [OF s]) (auto simp: powser_sums_zero sm) show ?thesis using "*" by (auto cong: Lim_cong_within) qed subsection \Derivability of power series\ lemma DERIV_series': fixes f :: "real \ nat \ real" assumes DERIV_f: "\ n. DERIV (\ x. f x n) x0 :> (f' x0 n)" and allf_summable: "\ x. x \ {a <..< b} \ summable (f x)" and x0_in_I: "x0 \ {a <..< b}" and "summable (f' x0)" and "summable L" and L_def: "\n x y. x \ {a <..< b} \ y \ {a <..< b} \ \f x n - f y n\ \ L n * \x - y\" shows "DERIV (\ x. suminf (f x)) x0 :> (suminf (f' x0))" unfolding DERIV_def proof (rule LIM_I) fix r :: real assume "0 < r" then have "0 < r/3" by auto obtain N_L where N_L: "\ n. N_L \ n \ \ \ i. L (i + n) \ < r/3" using suminf_exist_split[OF \0 < r/3\ \summable L\] by auto obtain N_f' where N_f': "\ n. N_f' \ n \ \ \ i. f' x0 (i + n) \ < r/3" using suminf_exist_split[OF \0 < r/3\ \summable (f' x0)\] by auto let ?N = "Suc (max N_L N_f')" have "\ \ i. f' x0 (i + ?N) \ < r/3" (is "?f'_part < r/3") and L_estimate: "\ \ i. L (i + ?N) \ < r/3" using N_L[of "?N"] and N_f' [of "?N"] by auto let ?diff = "\i x. (f (x0 + x) i - f x0 i) / x" let ?r = "r / (3 * real ?N)" from \0 < r\ have "0 < ?r" by simp let ?s = "\n. SOME s. 0 < s \ (\ x. x \ 0 \ \ x \ < s \ \ ?diff n x - f' x0 n \ < ?r)" define S' where "S' = Min (?s ` {..< ?N })" have "0 < S'" unfolding S'_def proof (rule iffD2[OF Min_gr_iff]) show "\x \ (?s ` {..< ?N }). 0 < x" proof fix x assume "x \ ?s ` {.. {..0 < ?r\, unfolded real_norm_def] obtain s where s_bound: "0 < s \ (\x. x \ 0 \ \x\ < s \ \?diff n x - f' x0 n\ < ?r)" by auto have "0 < ?s n" by (rule someI2[where a=s]) (auto simp: s_bound simp del: of_nat_Suc) then show "0 < x" by (simp only: \x = ?s n\) qed qed auto define S where "S = min (min (x0 - a) (b - x0)) S'" then have "0 < S" and S_a: "S \ x0 - a" and S_b: "S \ b - x0" and "S \ S'" using x0_in_I and \0 < S'\ by auto have "\(suminf (f (x0 + x)) - suminf (f x0)) / x - suminf (f' x0)\ < r" if "x \ 0" and "\x\ < S" for x proof - from that have x_in_I: "x0 + x \ {a <..< b}" using S_a S_b by auto note diff_smbl = summable_diff[OF allf_summable[OF x_in_I] allf_summable[OF x0_in_I]] note div_smbl = summable_divide[OF diff_smbl] note all_smbl = summable_diff[OF div_smbl \summable (f' x0)\] note ign = summable_ignore_initial_segment[where k="?N"] note diff_shft_smbl = summable_diff[OF ign[OF allf_summable[OF x_in_I]] ign[OF allf_summable[OF x0_in_I]]] note div_shft_smbl = summable_divide[OF diff_shft_smbl] note all_shft_smbl = summable_diff[OF div_smbl ign[OF \summable (f' x0)\]] have 1: "\(\?diff (n + ?N) x\)\ \ L (n + ?N)" for n proof - have "\?diff (n + ?N) x\ \ L (n + ?N) * \(x0 + x) - x0\ / \x\" using divide_right_mono[OF L_def[OF x_in_I x0_in_I] abs_ge_zero] by (simp only: abs_divide) with \x \ 0\ show ?thesis by auto qed note 2 = summable_rabs_comparison_test[OF _ ign[OF \summable L\]] from 1 have "\ \ i. ?diff (i + ?N) x \ \ (\ i. L (i + ?N))" by (metis (lifting) abs_idempotent order_trans[OF summable_rabs[OF 2] suminf_le[OF _ 2 ign[OF \summable L\]]]) then have "\\i. ?diff (i + ?N) x\ \ r / 3" (is "?L_part \ r/3") using L_estimate by auto have "\\n \ (\n?diff n x - f' x0 n\)" .. also have "\ < (\n {..< ?N}" have "\x\ < S" using \\x\ < S\ . also have "S \ S'" using \S \ S'\ . also have "S' \ ?s n" unfolding S'_def proof (rule Min_le_iff[THEN iffD2]) have "?s n \ (?s ` {.. ?s n \ ?s n" using \n \ {..< ?N}\ by auto then show "\ a \ (?s ` {.. ?s n" by blast qed auto finally have "\x\ < ?s n" . from DERIV_D[OF DERIV_f[where n=n], THEN LIM_D, OF \0 < ?r\, unfolded real_norm_def diff_0_right, unfolded some_eq_ex[symmetric], THEN conjunct2] have "\x. x \ 0 \ \x\ < ?s n \ \?diff n x - f' x0 n\ < ?r" . with \x \ 0\ and \\x\ < ?s n\ show "\?diff n x - f' x0 n\ < ?r" by blast qed auto also have "\ = of_nat (card {.. = real ?N * ?r" by simp also have "\ = r/3" by (auto simp del: of_nat_Suc) finally have "\\n < r / 3" (is "?diff_part < r / 3") . from suminf_diff[OF allf_summable[OF x_in_I] allf_summable[OF x0_in_I]] have "\(suminf (f (x0 + x)) - (suminf (f x0))) / x - suminf (f' x0)\ = \\n. ?diff n x - f' x0 n\" unfolding suminf_diff[OF div_smbl \summable (f' x0)\, symmetric] using suminf_divide[OF diff_smbl, symmetric] by auto also have "\ \ ?diff_part + \(\n. ?diff (n + ?N) x) - (\ n. f' x0 (n + ?N))\" unfolding suminf_split_initial_segment[OF all_smbl, where k="?N"] unfolding suminf_diff[OF div_shft_smbl ign[OF \summable (f' x0)\]] apply (simp only: add.commute) using abs_triangle_ineq by blast also have "\ \ ?diff_part + ?L_part + ?f'_part" using abs_triangle_ineq4 by auto also have "\ < r /3 + r/3 + r/3" using \?diff_part < r/3\ \?L_part \ r/3\ and \?f'_part < r/3\ by (rule add_strict_mono [OF add_less_le_mono]) finally show ?thesis by auto qed then show "\s > 0. \ x. x \ 0 \ norm (x - 0) < s \ norm (((\n. f (x0 + x) n) - (\n. f x0 n)) / x - (\n. f' x0 n)) < r" using \0 < S\ by auto qed lemma DERIV_power_series': fixes f :: "nat \ real" assumes converges: "\x. x \ {-R <..< R} \ summable (\n. f n * real (Suc n) * x^n)" and x0_in_I: "x0 \ {-R <..< R}" and "0 < R" shows "DERIV (\x. (\n. f n * x^(Suc n))) x0 :> (\n. f n * real (Suc n) * x0^n)" (is "DERIV (\x. suminf (?f x)) x0 :> suminf (?f' x0)") proof - have for_subinterval: "DERIV (\x. suminf (?f x)) x0 :> suminf (?f' x0)" if "0 < R'" and "R' < R" and "-R' < x0" and "x0 < R'" for R' proof - from that have "x0 \ {-R' <..< R'}" and "R' \ {-R <..< R}" and "x0 \ {-R <..< R}" by auto show ?thesis proof (rule DERIV_series') show "summable (\ n. \f n * real (Suc n) * R'^n\)" proof - have "(R' + R) / 2 < R" and "0 < (R' + R) / 2" using \0 < R'\ \0 < R\ \R' < R\ by (auto simp: field_simps) then have in_Rball: "(R' + R) / 2 \ {-R <..< R}" using \R' < R\ by auto have "norm R' < norm ((R' + R) / 2)" using \0 < R'\ \0 < R\ \R' < R\ by (auto simp: field_simps) from powser_insidea[OF converges[OF in_Rball] this] show ?thesis by auto qed next fix n x y assume "x \ {-R' <..< R'}" and "y \ {-R' <..< R'}" show "\?f x n - ?f y n\ \ \f n * real (Suc n) * R'^n\ * \x-y\" proof - have "\f n * x ^ (Suc n) - f n * y ^ (Suc n)\ = (\f n\ * \x-y\) * \\p" unfolding right_diff_distrib[symmetric] diff_power_eq_sum abs_mult by auto also have "\ \ (\f n\ * \x-y\) * (\real (Suc n)\ * \R' ^ n\)" proof (rule mult_left_mono) have "\\p \ (\px ^ p * y ^ (n - p)\)" by (rule sum_abs) also have "\ \ (\p {.. n" by auto have "\x^n\ \ R'^n" if "x \ {-R'<..x\ \ R'" by auto then show ?thesis unfolding power_abs by (rule power_mono) auto qed from mult_mono[OF this[OF \x \ {-R'<.., of p] this[OF \y \ {-R'<.., of "n-p"]] and \0 < R'\ have "\x^p * y^(n - p)\ \ R'^p * R'^(n - p)" unfolding abs_mult by auto then show "\x^p * y^(n - p)\ \ R'^n" unfolding power_add[symmetric] using \p \ n\ by auto qed also have "\ = real (Suc n) * R' ^ n" unfolding sum_constant card_atLeastLessThan by auto finally show "\\p \ \real (Suc n)\ * \R' ^ n\" unfolding abs_of_nonneg[OF zero_le_power[OF less_imp_le[OF \0 < R'\]]] by linarith show "0 \ \f n\ * \x - y\" unfolding abs_mult[symmetric] by auto qed also have "\ = \f n * real (Suc n) * R' ^ n\ * \x - y\" unfolding abs_mult mult.assoc[symmetric] by algebra finally show ?thesis . qed next show "DERIV (\x. ?f x n) x0 :> ?f' x0 n" for n by (auto intro!: derivative_eq_intros simp del: power_Suc) next fix x assume "x \ {-R' <..< R'}" then have "R' \ {-R <..< R}" and "norm x < norm R'" using assms \R' < R\ by auto have "summable (\n. f n * x^n)" proof (rule summable_comparison_test, intro exI allI impI) fix n have le: "\f n\ * 1 \ \f n\ * real (Suc n)" by (rule mult_left_mono) auto show "norm (f n * x^n) \ norm (f n * real (Suc n) * x^n)" unfolding real_norm_def abs_mult using le mult_right_mono by fastforce qed (rule powser_insidea[OF converges[OF \R' \ {-R <..< R}\] \norm x < norm R'\]) from this[THEN summable_mult2[where c=x], simplified mult.assoc, simplified mult.commute] show "summable (?f x)" by auto next show "summable (?f' x0)" using converges[OF \x0 \ {-R <..< R}\] . show "x0 \ {-R' <..< R'}" using \x0 \ {-R' <..< R'}\ . qed qed let ?R = "(R + \x0\) / 2" have "\x0\ < ?R" using assms by (auto simp: field_simps) then have "- ?R < x0" proof (cases "x0 < 0") case True then have "- x0 < ?R" using \\x0\ < ?R\ by auto then show ?thesis unfolding neg_less_iff_less[symmetric, of "- x0"] by auto next case False have "- ?R < 0" using assms by auto also have "\ \ x0" using False by auto finally show ?thesis . qed then have "0 < ?R" "?R < R" "- ?R < x0" and "x0 < ?R" using assms by (auto simp: field_simps) from for_subinterval[OF this] show ?thesis . qed lemma geometric_deriv_sums: fixes z :: "'a :: {real_normed_field,banach}" assumes "norm z < 1" shows "(\n. of_nat (Suc n) * z ^ n) sums (1 / (1 - z)^2)" proof - have "(\n. diffs (\n. 1) n * z^n) sums (1 / (1 - z)^2)" proof (rule termdiffs_sums_strong) fix z :: 'a assume "norm z < 1" thus "(\n. 1 * z^n) sums (1 / (1 - z))" by (simp add: geometric_sums) qed (insert assms, auto intro!: derivative_eq_intros simp: power2_eq_square) thus ?thesis unfolding diffs_def by simp qed lemma isCont_pochhammer [continuous_intros]: "isCont (\z. pochhammer z n) z" for z :: "'a::real_normed_field" by (induct n) (auto simp: pochhammer_rec') lemma continuous_on_pochhammer [continuous_intros]: "continuous_on A (\z. pochhammer z n)" for A :: "'a::real_normed_field set" by (intro continuous_at_imp_continuous_on ballI isCont_pochhammer) lemmas continuous_on_pochhammer' [continuous_intros] = continuous_on_compose2[OF continuous_on_pochhammer _ subset_UNIV] subsection \Exponential Function\ definition exp :: "'a \ 'a::{real_normed_algebra_1,banach}" where "exp = (\x. \n. x^n /\<^sub>R fact n)" lemma summable_exp_generic: fixes x :: "'a::{real_normed_algebra_1,banach}" defines S_def: "S \ \n. x^n /\<^sub>R fact n" shows "summable S" proof - have S_Suc: "\n. S (Suc n) = (x * S n) /\<^sub>R (Suc n)" unfolding S_def by (simp del: mult_Suc) obtain r :: real where r0: "0 < r" and r1: "r < 1" using dense [OF zero_less_one] by fast obtain N :: nat where N: "norm x < real N * r" using ex_less_of_nat_mult r0 by auto from r1 show ?thesis proof (rule summable_ratio_test [rule_format]) fix n :: nat assume n: "N \ n" have "norm x \ real N * r" using N by (rule order_less_imp_le) also have "real N * r \ real (Suc n) * r" using r0 n by (simp add: mult_right_mono) finally have "norm x * norm (S n) \ real (Suc n) * r * norm (S n)" using norm_ge_zero by (rule mult_right_mono) then have "norm (x * S n) \ real (Suc n) * r * norm (S n)" by (rule order_trans [OF norm_mult_ineq]) then have "norm (x * S n) / real (Suc n) \ r * norm (S n)" by (simp add: pos_divide_le_eq ac_simps) then show "norm (S (Suc n)) \ r * norm (S n)" by (simp add: S_Suc inverse_eq_divide) qed qed lemma summable_norm_exp: "summable (\n. norm (x^n /\<^sub>R fact n))" for x :: "'a::{real_normed_algebra_1,banach}" proof (rule summable_norm_comparison_test [OF exI, rule_format]) show "summable (\n. norm x^n /\<^sub>R fact n)" by (rule summable_exp_generic) show "norm (x^n /\<^sub>R fact n) \ norm x^n /\<^sub>R fact n" for n by (simp add: norm_power_ineq) qed lemma summable_exp: "summable (\n. inverse (fact n) * x^n)" for x :: "'a::{real_normed_field,banach}" using summable_exp_generic [where x=x] by (simp add: scaleR_conv_of_real nonzero_of_real_inverse) lemma exp_converges: "(\n. x^n /\<^sub>R fact n) sums exp x" unfolding exp_def by (rule summable_exp_generic [THEN summable_sums]) lemma exp_fdiffs: "diffs (\n. inverse (fact n)) = (\n. inverse (fact n :: 'a::{real_normed_field,banach}))" by (simp add: diffs_def mult_ac nonzero_inverse_mult_distrib nonzero_of_real_inverse del: mult_Suc of_nat_Suc) lemma diffs_of_real: "diffs (\n. of_real (f n)) = (\n. of_real (diffs f n))" by (simp add: diffs_def) lemma DERIV_exp [simp]: "DERIV exp x :> exp x" unfolding exp_def scaleR_conv_of_real proof (rule DERIV_cong) have sinv: "summable (\n. of_real (inverse (fact n)) * x ^ n)" for x::'a by (rule exp_converges [THEN sums_summable, unfolded scaleR_conv_of_real]) note xx = exp_converges [THEN sums_summable, unfolded scaleR_conv_of_real] show "((\x. \n. of_real (inverse (fact n)) * x ^ n) has_field_derivative (\n. diffs (\n. of_real (inverse (fact n))) n * x ^ n)) (at x)" by (rule termdiffs [where K="of_real (1 + norm x)"]) (simp_all only: diffs_of_real exp_fdiffs sinv norm_of_real) show "(\n. diffs (\n. of_real (inverse (fact n))) n * x ^ n) = (\n. of_real (inverse (fact n)) * x ^ n)" by (simp add: diffs_of_real exp_fdiffs) qed declare DERIV_exp[THEN DERIV_chain2, derivative_intros] and DERIV_exp[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemmas has_derivative_exp[derivative_intros] = DERIV_exp[THEN DERIV_compose_FDERIV] lemma norm_exp: "norm (exp x) \ exp (norm x)" proof - from summable_norm[OF summable_norm_exp, of x] have "norm (exp x) \ (\n. inverse (fact n) * norm (x^n))" by (simp add: exp_def) also have "\ \ exp (norm x)" using summable_exp_generic[of "norm x"] summable_norm_exp[of x] by (auto simp: exp_def intro!: suminf_le norm_power_ineq) finally show ?thesis . qed lemma isCont_exp: "isCont exp x" for x :: "'a::{real_normed_field,banach}" by (rule DERIV_exp [THEN DERIV_isCont]) lemma isCont_exp' [simp]: "isCont f a \ isCont (\x. exp (f x)) a" for f :: "_ \'a::{real_normed_field,banach}" by (rule isCont_o2 [OF _ isCont_exp]) lemma tendsto_exp [tendsto_intros]: "(f \ a) F \ ((\x. exp (f x)) \ exp a) F" for f:: "_ \'a::{real_normed_field,banach}" by (rule isCont_tendsto_compose [OF isCont_exp]) lemma continuous_exp [continuous_intros]: "continuous F f \ continuous F (\x. exp (f x))" for f :: "_ \'a::{real_normed_field,banach}" unfolding continuous_def by (rule tendsto_exp) lemma continuous_on_exp [continuous_intros]: "continuous_on s f \ continuous_on s (\x. exp (f x))" for f :: "_ \'a::{real_normed_field,banach}" unfolding continuous_on_def by (auto intro: tendsto_exp) subsubsection \Properties of the Exponential Function\ lemma exp_zero [simp]: "exp 0 = 1" unfolding exp_def by (simp add: scaleR_conv_of_real) lemma exp_series_add_commuting: fixes x y :: "'a::{real_normed_algebra_1,banach}" defines S_def: "S \ \x n. x^n /\<^sub>R fact n" assumes comm: "x * y = y * x" shows "S (x + y) n = (\i\n. S x i * S y (n - i))" proof (induct n) case 0 show ?case unfolding S_def by simp next case (Suc n) have S_Suc: "\x n. S x (Suc n) = (x * S x n) /\<^sub>R real (Suc n)" unfolding S_def by (simp del: mult_Suc) then have times_S: "\x n. x * S x n = real (Suc n) *\<^sub>R S x (Suc n)" by simp have S_comm: "\n. S x n * y = y * S x n" by (simp add: power_commuting_commutes comm S_def) have "real (Suc n) *\<^sub>R S (x + y) (Suc n) = (x + y) * (\i\n. S x i * S y (n - i))" by (metis Suc.hyps times_S) also have "\ = x * (\i\n. S x i * S y (n - i)) + y * (\i\n. S x i * S y (n - i))" by (rule distrib_right) also have "\ = (\i\n. x * S x i * S y (n - i)) + (\i\n. S x i * y * S y (n - i))" by (simp add: sum_distrib_left ac_simps S_comm) also have "\ = (\i\n. x * S x i * S y (n - i)) + (\i\n. S x i * (y * S y (n - i)))" by (simp add: ac_simps) also have "\ = (\i\n. real (Suc i) *\<^sub>R (S x (Suc i) * S y (n - i))) + (\i\n. real (Suc n - i) *\<^sub>R (S x i * S y (Suc n - i)))" by (simp add: times_S Suc_diff_le) also have "(\i\n. real (Suc i) *\<^sub>R (S x (Suc i) * S y (n - i))) = (\i\Suc n. real i *\<^sub>R (S x i * S y (Suc n - i)))" by (subst sum.atMost_Suc_shift) simp also have "(\i\n. real (Suc n - i) *\<^sub>R (S x i * S y (Suc n - i))) = (\i\Suc n. real (Suc n - i) *\<^sub>R (S x i * S y (Suc n - i)))" by simp also have "(\i\Suc n. real i *\<^sub>R (S x i * S y (Suc n - i))) + (\i\Suc n. real (Suc n - i) *\<^sub>R (S x i * S y (Suc n - i))) = (\i\Suc n. real (Suc n) *\<^sub>R (S x i * S y (Suc n - i)))" by (simp flip: sum.distrib scaleR_add_left of_nat_add) also have "\ = real (Suc n) *\<^sub>R (\i\Suc n. S x i * S y (Suc n - i))" by (simp only: scaleR_right.sum) finally show "S (x + y) (Suc n) = (\i\Suc n. S x i * S y (Suc n - i))" by (simp del: sum.cl_ivl_Suc) qed lemma exp_add_commuting: "x * y = y * x \ exp (x + y) = exp x * exp y" by (simp only: exp_def Cauchy_product summable_norm_exp exp_series_add_commuting) lemma exp_times_arg_commute: "exp A * A = A * exp A" by (simp add: exp_def suminf_mult[symmetric] summable_exp_generic power_commutes suminf_mult2) lemma exp_add: "exp (x + y) = exp x * exp y" for x y :: "'a::{real_normed_field,banach}" by (rule exp_add_commuting) (simp add: ac_simps) lemma exp_double: "exp(2 * z) = exp z ^ 2" by (simp add: exp_add_commuting mult_2 power2_eq_square) lemmas mult_exp_exp = exp_add [symmetric] lemma exp_of_real: "exp (of_real x) = of_real (exp x)" unfolding exp_def apply (subst suminf_of_real [OF summable_exp_generic]) apply (simp add: scaleR_conv_of_real) done lemmas of_real_exp = exp_of_real[symmetric] corollary exp_in_Reals [simp]: "z \ \ \ exp z \ \" by (metis Reals_cases Reals_of_real exp_of_real) lemma exp_not_eq_zero [simp]: "exp x \ 0" proof have "exp x * exp (- x) = 1" by (simp add: exp_add_commuting[symmetric]) also assume "exp x = 0" finally show False by simp qed lemma exp_minus_inverse: "exp x * exp (- x) = 1" by (simp add: exp_add_commuting[symmetric]) lemma exp_minus: "exp (- x) = inverse (exp x)" for x :: "'a::{real_normed_field,banach}" by (intro inverse_unique [symmetric] exp_minus_inverse) lemma exp_diff: "exp (x - y) = exp x / exp y" for x :: "'a::{real_normed_field,banach}" using exp_add [of x "- y"] by (simp add: exp_minus divide_inverse) lemma exp_of_nat_mult: "exp (of_nat n * x) = exp x ^ n" for x :: "'a::{real_normed_field,banach}" by (induct n) (auto simp: distrib_left exp_add mult.commute) corollary exp_of_nat2_mult: "exp (x * of_nat n) = exp x ^ n" for x :: "'a::{real_normed_field,banach}" by (metis exp_of_nat_mult mult_of_nat_commute) lemma exp_sum: "finite I \ exp (sum f I) = prod (\x. exp (f x)) I" by (induct I rule: finite_induct) (auto simp: exp_add_commuting mult.commute) lemma exp_divide_power_eq: fixes x :: "'a::{real_normed_field,banach}" assumes "n > 0" shows "exp (x / of_nat n) ^ n = exp x" using assms proof (induction n arbitrary: x) case (Suc n) show ?case proof (cases "n = 0") case True then show ?thesis by simp next case False have [simp]: "1 + (of_nat n * of_nat n + of_nat n * 2) \ (0::'a)" using of_nat_eq_iff [of "1 + n * n + n * 2" "0"] by simp from False have [simp]: "x * of_nat n / (1 + of_nat n) / of_nat n = x / (1 + of_nat n)" by simp have [simp]: "x / (1 + of_nat n) + x * of_nat n / (1 + of_nat n) = x" using of_nat_neq_0 by (auto simp add: field_split_simps) show ?thesis using Suc.IH [of "x * of_nat n / (1 + of_nat n)"] False by (simp add: exp_add [symmetric]) qed qed simp lemma exp_power_int: fixes x :: "'a::{real_normed_field,banach}" shows "exp x powi n = exp (of_int n * x)" proof (cases "n \ 0") case True have "exp x powi n = exp x ^ nat n" using True by (simp add: power_int_def) thus ?thesis using True by (subst (asm) exp_of_nat_mult [symmetric]) auto next case False have "exp x powi n = inverse (exp x ^ nat (-n))" using False by (simp add: power_int_def field_simps) also have "exp x ^ nat (-n) = exp (of_nat (nat (-n)) * x)" using False by (subst exp_of_nat_mult) auto also have "inverse \ = exp (-(of_nat (nat (-n)) * x))" by (subst exp_minus) (auto simp: field_simps) also have "-(of_nat (nat (-n)) * x) = of_int n * x" using False by simp finally show ?thesis . qed subsubsection \Properties of the Exponential Function on Reals\ text \Comparisons of \<^term>\exp x\ with zero.\ text \Proof: because every exponential can be seen as a square.\ lemma exp_ge_zero [simp]: "0 \ exp x" for x :: real proof - have "0 \ exp (x/2) * exp (x/2)" by simp then show ?thesis by (simp add: exp_add [symmetric]) qed lemma exp_gt_zero [simp]: "0 < exp x" for x :: real by (simp add: order_less_le) lemma not_exp_less_zero [simp]: "\ exp x < 0" for x :: real by (simp add: not_less) lemma not_exp_le_zero [simp]: "\ exp x \ 0" for x :: real by (simp add: not_le) lemma abs_exp_cancel [simp]: "\exp x\ = exp x" for x :: real by simp text \Strict monotonicity of exponential.\ lemma exp_ge_add_one_self_aux: fixes x :: real assumes "0 \ x" shows "1 + x \ exp x" using order_le_imp_less_or_eq [OF assms] proof assume "0 < x" have "1 + x \ (\n<2. inverse (fact n) * x^n)" by (auto simp: numeral_2_eq_2) also have "\ \ (\n. inverse (fact n) * x^n)" using \0 < x\ by (auto simp add: zero_le_mult_iff intro: sum_le_suminf [OF summable_exp]) finally show "1 + x \ exp x" by (simp add: exp_def) qed auto lemma exp_gt_one: "0 < x \ 1 < exp x" for x :: real proof - assume x: "0 < x" then have "1 < 1 + x" by simp also from x have "1 + x \ exp x" by (simp add: exp_ge_add_one_self_aux) finally show ?thesis . qed lemma exp_less_mono: fixes x y :: real assumes "x < y" shows "exp x < exp y" proof - from \x < y\ have "0 < y - x" by simp then have "1 < exp (y - x)" by (rule exp_gt_one) then have "1 < exp y / exp x" by (simp only: exp_diff) then show "exp x < exp y" by simp qed lemma exp_less_cancel: "exp x < exp y \ x < y" for x y :: real unfolding linorder_not_le [symmetric] by (auto simp: order_le_less exp_less_mono) lemma exp_less_cancel_iff [iff]: "exp x < exp y \ x < y" for x y :: real by (auto intro: exp_less_mono exp_less_cancel) lemma exp_le_cancel_iff [iff]: "exp x \ exp y \ x \ y" for x y :: real by (auto simp: linorder_not_less [symmetric]) lemma exp_inj_iff [iff]: "exp x = exp y \ x = y" for x y :: real by (simp add: order_eq_iff) text \Comparisons of \<^term>\exp x\ with one.\ lemma one_less_exp_iff [simp]: "1 < exp x \ 0 < x" for x :: real using exp_less_cancel_iff [where x = 0 and y = x] by simp lemma exp_less_one_iff [simp]: "exp x < 1 \ x < 0" for x :: real using exp_less_cancel_iff [where x = x and y = 0] by simp lemma one_le_exp_iff [simp]: "1 \ exp x \ 0 \ x" for x :: real using exp_le_cancel_iff [where x = 0 and y = x] by simp lemma exp_le_one_iff [simp]: "exp x \ 1 \ x \ 0" for x :: real using exp_le_cancel_iff [where x = x and y = 0] by simp lemma exp_eq_one_iff [simp]: "exp x = 1 \ x = 0" for x :: real using exp_inj_iff [where x = x and y = 0] by simp lemma lemma_exp_total: "1 \ y \ \x. 0 \ x \ x \ y - 1 \ exp x = y" for y :: real proof (rule IVT) assume "1 \ y" then have "0 \ y - 1" by simp then have "1 + (y - 1) \ exp (y - 1)" by (rule exp_ge_add_one_self_aux) then show "y \ exp (y - 1)" by simp qed (simp_all add: le_diff_eq) lemma exp_total: "0 < y \ \x. exp x = y" for y :: real proof (rule linorder_le_cases [of 1 y]) assume "1 \ y" then show "\x. exp x = y" by (fast dest: lemma_exp_total) next assume "0 < y" and "y \ 1" then have "1 \ inverse y" by (simp add: one_le_inverse_iff) then obtain x where "exp x = inverse y" by (fast dest: lemma_exp_total) then have "exp (- x) = y" by (simp add: exp_minus) then show "\x. exp x = y" .. qed subsection \Natural Logarithm\ class ln = real_normed_algebra_1 + banach + fixes ln :: "'a \ 'a" assumes ln_one [simp]: "ln 1 = 0" definition powr :: "'a \ 'a \ 'a::ln" (infixr "powr" 80) \ \exponentation via ln and exp\ where "x powr a \ if x = 0 then 0 else exp (a * ln x)" lemma powr_0 [simp]: "0 powr z = 0" by (simp add: powr_def) instantiation real :: ln begin definition ln_real :: "real \ real" where "ln_real x = (THE u. exp u = x)" instance by intro_classes (simp add: ln_real_def) end lemma powr_eq_0_iff [simp]: "w powr z = 0 \ w = 0" by (simp add: powr_def) lemma ln_exp [simp]: "ln (exp x) = x" for x :: real by (simp add: ln_real_def) lemma exp_ln [simp]: "0 < x \ exp (ln x) = x" for x :: real by (auto dest: exp_total) lemma exp_ln_iff [simp]: "exp (ln x) = x \ 0 < x" for x :: real by (metis exp_gt_zero exp_ln) lemma ln_unique: "exp y = x \ ln x = y" for x :: real by (erule subst) (rule ln_exp) lemma ln_mult: "0 < x \ 0 < y \ ln (x * y) = ln x + ln y" for x :: real by (rule ln_unique) (simp add: exp_add) lemma ln_prod: "finite I \ (\i. i \ I \ f i > 0) \ ln (prod f I) = sum (\x. ln(f x)) I" for f :: "'a \ real" by (induct I rule: finite_induct) (auto simp: ln_mult prod_pos) lemma ln_inverse: "0 < x \ ln (inverse x) = - ln x" for x :: real by (rule ln_unique) (simp add: exp_minus) lemma ln_div: "0 < x \ 0 < y \ ln (x / y) = ln x - ln y" for x :: real by (rule ln_unique) (simp add: exp_diff) lemma ln_realpow: "0 < x \ ln (x^n) = real n * ln x" by (rule ln_unique) (simp add: exp_of_nat_mult) lemma ln_less_cancel_iff [simp]: "0 < x \ 0 < y \ ln x < ln y \ x < y" for x :: real by (subst exp_less_cancel_iff [symmetric]) simp lemma ln_le_cancel_iff [simp]: "0 < x \ 0 < y \ ln x \ ln y \ x \ y" for x :: real by (simp add: linorder_not_less [symmetric]) lemma ln_mono: "\x::real. \x \ y; 0 < x; 0 < y\ \ ln x \ ln y" using ln_le_cancel_iff by presburger lemma ln_inj_iff [simp]: "0 < x \ 0 < y \ ln x = ln y \ x = y" for x :: real by (simp add: order_eq_iff) lemma ln_add_one_self_le_self: "0 \ x \ ln (1 + x) \ x" for x :: real by (rule exp_le_cancel_iff [THEN iffD1]) (simp add: exp_ge_add_one_self_aux) lemma ln_less_self [simp]: "0 < x \ ln x < x" for x :: real by (rule order_less_le_trans [where y = "ln (1 + x)"]) (simp_all add: ln_add_one_self_le_self) lemma ln_ge_iff: "\x::real. 0 < x \ y \ ln x \ exp y \ x" using exp_le_cancel_iff exp_total by force lemma ln_ge_zero [simp]: "1 \ x \ 0 \ ln x" for x :: real using ln_le_cancel_iff [of 1 x] by simp lemma ln_ge_zero_imp_ge_one: "0 \ ln x \ 0 < x \ 1 \ x" for x :: real using ln_le_cancel_iff [of 1 x] by simp lemma ln_ge_zero_iff [simp]: "0 < x \ 0 \ ln x \ 1 \ x" for x :: real using ln_le_cancel_iff [of 1 x] by simp lemma ln_less_zero_iff [simp]: "0 < x \ ln x < 0 \ x < 1" for x :: real using ln_less_cancel_iff [of x 1] by simp lemma ln_le_zero_iff [simp]: "0 < x \ ln x \ 0 \ x \ 1" for x :: real by (metis less_numeral_extra(1) ln_le_cancel_iff ln_one) lemma ln_gt_zero: "1 < x \ 0 < ln x" for x :: real using ln_less_cancel_iff [of 1 x] by simp lemma ln_gt_zero_imp_gt_one: "0 < ln x \ 0 < x \ 1 < x" for x :: real using ln_less_cancel_iff [of 1 x] by simp lemma ln_gt_zero_iff [simp]: "0 < x \ 0 < ln x \ 1 < x" for x :: real using ln_less_cancel_iff [of 1 x] by simp lemma ln_eq_zero_iff [simp]: "0 < x \ ln x = 0 \ x = 1" for x :: real using ln_inj_iff [of x 1] by simp lemma ln_less_zero: "0 < x \ x < 1 \ ln x < 0" for x :: real by simp lemma ln_neg_is_const: "x \ 0 \ ln x = (THE x. False)" for x :: real by (auto simp: ln_real_def intro!: arg_cong[where f = The]) lemma powr_eq_one_iff [simp]: "a powr x = 1 \ x = 0" if "a > 1" for a x :: real using that by (auto simp: powr_def split: if_splits) lemma isCont_ln: fixes x :: real assumes "x \ 0" shows "isCont ln x" proof (cases "0 < x") case True then have "isCont ln (exp (ln x))" by (intro isCont_inverse_function[where d = "\x\" and f = exp]) auto with True show ?thesis by simp next case False with \x \ 0\ show "isCont ln x" unfolding isCont_def by (subst filterlim_cong[OF _ refl, of _ "nhds (ln 0)" _ "\_. ln 0"]) (auto simp: ln_neg_is_const not_less eventually_at dist_real_def intro!: exI[of _ "\x\"]) qed lemma tendsto_ln [tendsto_intros]: "(f \ a) F \ a \ 0 \ ((\x. ln (f x)) \ ln a) F" for a :: real by (rule isCont_tendsto_compose [OF isCont_ln]) lemma continuous_ln: "continuous F f \ f (Lim F (\x. x)) \ 0 \ continuous F (\x. ln (f x :: real))" unfolding continuous_def by (rule tendsto_ln) lemma isCont_ln' [continuous_intros]: "continuous (at x) f \ f x \ 0 \ continuous (at x) (\x. ln (f x :: real))" unfolding continuous_at by (rule tendsto_ln) lemma continuous_within_ln [continuous_intros]: "continuous (at x within s) f \ f x \ 0 \ continuous (at x within s) (\x. ln (f x :: real))" unfolding continuous_within by (rule tendsto_ln) lemma continuous_on_ln [continuous_intros]: "continuous_on s f \ (\x\s. f x \ 0) \ continuous_on s (\x. ln (f x :: real))" unfolding continuous_on_def by (auto intro: tendsto_ln) lemma DERIV_ln: "0 < x \ DERIV ln x :> inverse x" for x :: real by (rule DERIV_inverse_function [where f=exp and a=0 and b="x+1"]) (auto intro: DERIV_cong [OF DERIV_exp exp_ln] isCont_ln) lemma DERIV_ln_divide: "0 < x \ DERIV ln x :> 1/x" for x :: real by (rule DERIV_ln[THEN DERIV_cong]) (simp_all add: divide_inverse) declare DERIV_ln_divide[THEN DERIV_chain2, derivative_intros] and DERIV_ln_divide[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemmas has_derivative_ln[derivative_intros] = DERIV_ln[THEN DERIV_compose_FDERIV] lemma ln_series: assumes "0 < x" and "x < 2" shows "ln x = (\ n. (-1)^n * (1 / real (n + 1)) * (x - 1)^(Suc n))" (is "ln x = suminf (?f (x - 1))") proof - let ?f' = "\x n. (-1)^n * (x - 1)^n" have "ln x - suminf (?f (x - 1)) = ln 1 - suminf (?f (1 - 1))" proof (rule DERIV_isconst3 [where x = x]) fix x :: real assume "x \ {0 <..< 2}" then have "0 < x" and "x < 2" by auto have "norm (1 - x) < 1" using \0 < x\ and \x < 2\ by auto have "1/x = 1 / (1 - (1 - x))" by auto also have "\ = (\ n. (1 - x)^n)" using geometric_sums[OF \norm (1 - x) < 1\] by (rule sums_unique) also have "\ = suminf (?f' x)" unfolding power_mult_distrib[symmetric] by (rule arg_cong[where f=suminf], rule arg_cong[where f="(^)"], auto) finally have "DERIV ln x :> suminf (?f' x)" using DERIV_ln[OF \0 < x\] unfolding divide_inverse by auto moreover have repos: "\ h x :: real. h - 1 + x = h + x - 1" by auto have "DERIV (\x. suminf (?f x)) (x - 1) :> (\n. (-1)^n * (1 / real (n + 1)) * real (Suc n) * (x - 1) ^ n)" proof (rule DERIV_power_series') show "x - 1 \ {- 1<..<1}" and "(0 :: real) < 1" using \0 < x\ \x < 2\ by auto next fix x :: real assume "x \ {- 1<..<1}" then show "summable (\n. (- 1) ^ n * (1 / real (n + 1)) * real (Suc n) * x^n)" by (simp add: abs_if flip: power_mult_distrib) qed then have "DERIV (\x. suminf (?f x)) (x - 1) :> suminf (?f' x)" unfolding One_nat_def by auto then have "DERIV (\x. suminf (?f (x - 1))) x :> suminf (?f' x)" unfolding DERIV_def repos . ultimately have "DERIV (\x. ln x - suminf (?f (x - 1))) x :> suminf (?f' x) - suminf (?f' x)" by (rule DERIV_diff) then show "DERIV (\x. ln x - suminf (?f (x - 1))) x :> 0" by auto qed (auto simp: assms) then show ?thesis by auto qed lemma exp_first_terms: fixes x :: "'a::{real_normed_algebra_1,banach}" shows "exp x = (\nR (x ^ n)) + (\n. inverse(fact (n + k)) *\<^sub>R (x ^ (n + k)))" proof - have "exp x = suminf (\n. inverse(fact n) *\<^sub>R (x^n))" by (simp add: exp_def) also from summable_exp_generic have "\ = (\ n. inverse(fact(n+k)) *\<^sub>R (x ^ (n + k))) + (\ n::natR (x^n))" (is "_ = _ + ?a") by (rule suminf_split_initial_segment) finally show ?thesis by simp qed lemma exp_first_term: "exp x = 1 + (\n. inverse (fact (Suc n)) *\<^sub>R (x ^ Suc n))" for x :: "'a::{real_normed_algebra_1,banach}" using exp_first_terms[of x 1] by simp lemma exp_first_two_terms: "exp x = 1 + x + (\n. inverse (fact (n + 2)) *\<^sub>R (x ^ (n + 2)))" for x :: "'a::{real_normed_algebra_1,banach}" using exp_first_terms[of x 2] by (simp add: eval_nat_numeral) lemma exp_bound: fixes x :: real assumes a: "0 \ x" and b: "x \ 1" shows "exp x \ 1 + x + x\<^sup>2" proof - have "suminf (\n. inverse(fact (n+2)) * (x ^ (n + 2))) \ x\<^sup>2" proof - have "(\n. x\<^sup>2 / 2 * (1/2) ^ n) sums (x\<^sup>2 / 2 * (1 / (1 - 1/2)))" by (intro sums_mult geometric_sums) simp then have sumsx: "(\n. x\<^sup>2 / 2 * (1/2) ^ n) sums x\<^sup>2" by simp have "suminf (\n. inverse(fact (n+2)) * (x ^ (n + 2))) \ suminf (\n. (x\<^sup>2/2) * ((1/2)^n))" proof (intro suminf_le allI) show "inverse (fact (n + 2)) * x ^ (n + 2) \ (x\<^sup>2/2) * ((1/2)^n)" for n :: nat proof - have "(2::nat) * 2 ^ n \ fact (n + 2)" by (induct n) simp_all then have "real ((2::nat) * 2 ^ n) \ real_of_nat (fact (n + 2))" by (simp only: of_nat_le_iff) then have "((2::real) * 2 ^ n) \ fact (n + 2)" unfolding of_nat_fact by simp then have "inverse (fact (n + 2)) \ inverse ((2::real) * 2 ^ n)" by (rule le_imp_inverse_le) simp then have "inverse (fact (n + 2)) \ 1/(2::real) * (1/2)^n" by (simp add: power_inverse [symmetric]) then have "inverse (fact (n + 2)) * (x^n * x\<^sup>2) \ 1/2 * (1/2)^n * (1 * x\<^sup>2)" by (rule mult_mono) (rule mult_mono, simp_all add: power_le_one a b) then show ?thesis unfolding power_add by (simp add: ac_simps del: fact_Suc) qed show "summable (\n. inverse (fact (n + 2)) * x ^ (n + 2))" by (rule summable_exp [THEN summable_ignore_initial_segment]) show "summable (\n. x\<^sup>2 / 2 * (1/2) ^ n)" by (rule sums_summable [OF sumsx]) qed also have "\ = x\<^sup>2" by (rule sums_unique [THEN sym]) (rule sumsx) finally show ?thesis . qed then show ?thesis unfolding exp_first_two_terms by auto qed corollary exp_half_le2: "exp(1/2) \ (2::real)" using exp_bound [of "1/2"] by (simp add: field_simps) corollary exp_le: "exp 1 \ (3::real)" using exp_bound [of 1] by (simp add: field_simps) lemma exp_bound_half: "norm z \ 1/2 \ norm (exp z) \ 2" by (blast intro: order_trans intro!: exp_half_le2 norm_exp) lemma exp_bound_lemma: assumes "norm z \ 1/2" shows "norm (exp z) \ 1 + 2 * norm z" proof - have *: "(norm z)\<^sup>2 \ norm z * 1" unfolding power2_eq_square by (rule mult_left_mono) (use assms in auto) have "norm (exp z) \ exp (norm z)" by (rule norm_exp) also have "\ \ 1 + (norm z) + (norm z)\<^sup>2" using assms exp_bound by auto also have "\ \ 1 + 2 * norm z" using * by auto finally show ?thesis . qed lemma real_exp_bound_lemma: "0 \ x \ x \ 1/2 \ exp x \ 1 + 2 * x" for x :: real using exp_bound_lemma [of x] by simp lemma ln_one_minus_pos_upper_bound: fixes x :: real assumes a: "0 \ x" and b: "x < 1" shows "ln (1 - x) \ - x" proof - have "(1 - x) * (1 + x + x\<^sup>2) = 1 - x^3" by (simp add: algebra_simps power2_eq_square power3_eq_cube) also have "\ \ 1" by (auto simp: a) finally have "(1 - x) * (1 + x + x\<^sup>2) \ 1" . moreover have c: "0 < 1 + x + x\<^sup>2" by (simp add: add_pos_nonneg a) ultimately have "1 - x \ 1 / (1 + x + x\<^sup>2)" by (elim mult_imp_le_div_pos) also have "\ \ 1 / exp x" by (metis a abs_one b exp_bound exp_gt_zero frac_le less_eq_real_def real_sqrt_abs real_sqrt_pow2_iff real_sqrt_power) also have "\ = exp (- x)" by (auto simp: exp_minus divide_inverse) finally have "1 - x \ exp (- x)" . also have "1 - x = exp (ln (1 - x))" by (metis b diff_0 exp_ln_iff less_iff_diff_less_0 minus_diff_eq) finally have "exp (ln (1 - x)) \ exp (- x)" . then show ?thesis by (auto simp only: exp_le_cancel_iff) qed lemma exp_ge_add_one_self [simp]: "1 + x \ exp x" for x :: real proof (cases "0 \ x \ x \ -1") case True then show ?thesis by (meson exp_ge_add_one_self_aux exp_ge_zero order.trans real_add_le_0_iff) next case False then have ln1: "ln (1 + x) \ x" using ln_one_minus_pos_upper_bound [of "-x"] by simp have "1 + x = exp (ln (1 + x))" using False by auto also have "\ \ exp x" by (simp add: ln1) finally show ?thesis . qed lemma ln_one_plus_pos_lower_bound: fixes x :: real assumes a: "0 \ x" and b: "x \ 1" shows "x - x\<^sup>2 \ ln (1 + x)" proof - have "exp (x - x\<^sup>2) = exp x / exp (x\<^sup>2)" by (rule exp_diff) also have "\ \ (1 + x + x\<^sup>2) / exp (x \<^sup>2)" by (metis a b divide_right_mono exp_bound exp_ge_zero) also have "\ \ (1 + x + x\<^sup>2) / (1 + x\<^sup>2)" by (simp add: a divide_left_mono add_pos_nonneg) also from a have "\ \ 1 + x" by (simp add: field_simps add_strict_increasing zero_le_mult_iff) finally have "exp (x - x\<^sup>2) \ 1 + x" . also have "\ = exp (ln (1 + x))" proof - from a have "0 < 1 + x" by auto then show ?thesis by (auto simp only: exp_ln_iff [THEN sym]) qed finally have "exp (x - x\<^sup>2) \ exp (ln (1 + x))" . then show ?thesis by (metis exp_le_cancel_iff) qed lemma ln_one_minus_pos_lower_bound: fixes x :: real assumes a: "0 \ x" and b: "x \ 1/2" shows "- x - 2 * x\<^sup>2 \ ln (1 - x)" proof - from b have c: "x < 1" by auto then have "ln (1 - x) = - ln (1 + x / (1 - x))" by (auto simp: ln_inverse [symmetric] field_simps intro: arg_cong [where f=ln]) also have "- (x / (1 - x)) \ \" proof - have "ln (1 + x / (1 - x)) \ x / (1 - x)" using a c by (intro ln_add_one_self_le_self) auto then show ?thesis by auto qed also have "- (x / (1 - x)) = - x / (1 - x)" by auto finally have d: "- x / (1 - x) \ ln (1 - x)" . have "0 < 1 - x" using a b by simp then have e: "- x - 2 * x\<^sup>2 \ - x / (1 - x)" using mult_right_le_one_le[of "x * x" "2 * x"] a b by (simp add: field_simps power2_eq_square) from e d show "- x - 2 * x\<^sup>2 \ ln (1 - x)" by (rule order_trans) qed lemma ln_add_one_self_le_self2: fixes x :: real shows "-1 < x \ ln (1 + x) \ x" by (metis diff_gt_0_iff_gt diff_minus_eq_add exp_ge_add_one_self exp_le_cancel_iff exp_ln minus_less_iff) lemma abs_ln_one_plus_x_minus_x_bound_nonneg: fixes x :: real assumes x: "0 \ x" and x1: "x \ 1" shows "\ln (1 + x) - x\ \ x\<^sup>2" proof - from x have "ln (1 + x) \ x" by (rule ln_add_one_self_le_self) then have "ln (1 + x) - x \ 0" by simp then have "\ln(1 + x) - x\ = - (ln(1 + x) - x)" by (rule abs_of_nonpos) also have "\ = x - ln (1 + x)" by simp also have "\ \ x\<^sup>2" proof - from x x1 have "x - x\<^sup>2 \ ln (1 + x)" by (intro ln_one_plus_pos_lower_bound) then show ?thesis by simp qed finally show ?thesis . qed lemma abs_ln_one_plus_x_minus_x_bound_nonpos: fixes x :: real assumes a: "-(1/2) \ x" and b: "x \ 0" shows "\ln (1 + x) - x\ \ 2 * x\<^sup>2" proof - have *: "- (-x) - 2 * (-x)\<^sup>2 \ ln (1 - (- x))" by (metis a b diff_zero ln_one_minus_pos_lower_bound minus_diff_eq neg_le_iff_le) have "\ln (1 + x) - x\ = x - ln (1 - (- x))" using a ln_add_one_self_le_self2 [of x] by (simp add: abs_if) also have "\ \ 2 * x\<^sup>2" using * by (simp add: algebra_simps) finally show ?thesis . qed lemma abs_ln_one_plus_x_minus_x_bound: fixes x :: real assumes "\x\ \ 1/2" shows "\ln (1 + x) - x\ \ 2 * x\<^sup>2" proof (cases "0 \ x") case True then show ?thesis using abs_ln_one_plus_x_minus_x_bound_nonneg assms by fastforce next case False then show ?thesis using abs_ln_one_plus_x_minus_x_bound_nonpos assms by auto qed lemma ln_x_over_x_mono: fixes x :: real assumes x: "exp 1 \ x" "x \ y" shows "ln y / y \ ln x / x" proof - note x moreover have "0 < exp (1::real)" by simp ultimately have a: "0 < x" and b: "0 < y" by (fast intro: less_le_trans order_trans)+ have "x * ln y - x * ln x = x * (ln y - ln x)" by (simp add: algebra_simps) also have "\ = x * ln (y / x)" by (simp only: ln_div a b) also have "y / x = (x + (y - x)) / x" by simp also have "\ = 1 + (y - x) / x" using x a by (simp add: field_simps) also have "x * ln (1 + (y - x) / x) \ x * ((y - x) / x)" using x a by (intro mult_left_mono ln_add_one_self_le_self) simp_all also have "\ = y - x" using a by simp also have "\ = (y - x) * ln (exp 1)" by simp also have "\ \ (y - x) * ln x" using a x exp_total of_nat_1 x(1) by (fastforce intro: mult_left_mono) also have "\ = y * ln x - x * ln x" by (rule left_diff_distrib) finally have "x * ln y \ y * ln x" by arith then have "ln y \ (y * ln x) / x" using a by (simp add: field_simps) also have "\ = y * (ln x / x)" by simp finally show ?thesis using b by (simp add: field_simps) qed lemma ln_le_minus_one: "0 < x \ ln x \ x - 1" for x :: real using exp_ge_add_one_self[of "ln x"] by simp corollary ln_diff_le: "0 < x \ 0 < y \ ln x - ln y \ (x - y) / y" for x :: real by (simp add: ln_div [symmetric] diff_divide_distrib ln_le_minus_one) lemma ln_eq_minus_one: fixes x :: real assumes "0 < x" "ln x = x - 1" shows "x = 1" proof - let ?l = "\y. ln y - y + 1" have D: "\x::real. 0 < x \ DERIV ?l x :> (1/x - 1)" by (auto intro!: derivative_eq_intros) show ?thesis proof (cases rule: linorder_cases) assume "x < 1" from dense[OF \x < 1\] obtain a where "x < a" "a < 1" by blast from \x < a\ have "?l x < ?l a" proof (rule DERIV_pos_imp_increasing) fix y assume "x \ y" "y \ a" with \0 < x\ \a < 1\ have "0 < 1 / y - 1" "0 < y" by (auto simp: field_simps) with D show "\z. DERIV ?l y :> z \ 0 < z" by blast qed also have "\ \ 0" using ln_le_minus_one \0 < x\ \x < a\ by (auto simp: field_simps) finally show "x = 1" using assms by auto next assume "1 < x" from dense[OF this] obtain a where "1 < a" "a < x" by blast from \a < x\ have "?l x < ?l a" proof (rule DERIV_neg_imp_decreasing) fix y assume "a \ y" "y \ x" with \1 < a\ have "1 / y - 1 < 0" "0 < y" by (auto simp: field_simps) with D show "\z. DERIV ?l y :> z \ z < 0" by blast qed also have "\ \ 0" using ln_le_minus_one \1 < a\ by (auto simp: field_simps) finally show "x = 1" using assms by auto next assume "x = 1" then show ?thesis by simp qed qed lemma ln_add_one_self_less_self: fixes x :: real assumes "x > 0" shows "ln (1 + x) < x" by (smt (verit, best) assms ln_eq_minus_one ln_le_minus_one) lemma ln_x_over_x_tendsto_0: "((\x::real. ln x / x) \ 0) at_top" proof (rule lhospital_at_top_at_top[where f' = inverse and g' = "\_. 1"]) from eventually_gt_at_top[of "0::real"] show "\\<^sub>F x in at_top. (ln has_real_derivative inverse x) (at x)" by eventually_elim (auto intro!: derivative_eq_intros simp: field_simps) qed (use tendsto_inverse_0 in \auto simp: filterlim_ident dest!: tendsto_mono[OF at_top_le_at_infinity]\) corollary exp_1_gt_powr: assumes "x > (0::real)" shows "exp 1 > (1 + 1/x) powr x" proof - have "ln (1 + 1/x) < 1/x" using ln_add_one_self_less_self assms by simp thus "exp 1 > (1 + 1/x) powr x" using assms by (simp add: field_simps powr_def) qed lemma exp_ge_one_plus_x_over_n_power_n: assumes "x \ - real n" "n > 0" shows "(1 + x / of_nat n) ^ n \ exp x" proof (cases "x = - of_nat n") case False from assms False have "(1 + x / of_nat n) ^ n = exp (of_nat n * ln (1 + x / of_nat n))" by (subst exp_of_nat_mult, subst exp_ln) (simp_all add: field_simps) also from assms False have "ln (1 + x / real n) \ x / real n" by (intro ln_add_one_self_le_self2) (simp_all add: field_simps) with assms have "exp (of_nat n * ln (1 + x / of_nat n)) \ exp x" by (simp add: field_simps) finally show ?thesis . next case True then show ?thesis by (simp add: zero_power) qed lemma exp_ge_one_minus_x_over_n_power_n: assumes "x \ real n" "n > 0" shows "(1 - x / of_nat n) ^ n \ exp (-x)" using exp_ge_one_plus_x_over_n_power_n[of n "-x"] assms by simp lemma exp_at_bot: "(exp \ (0::real)) at_bot" unfolding tendsto_Zfun_iff proof (rule ZfunI, simp add: eventually_at_bot_dense) fix r :: real assume "0 < r" have "exp x < r" if "x < ln r" for x by (metis \0 < r\ exp_less_mono exp_ln that) then show "\k. \n at_top" by (rule filterlim_at_top_at_top[where Q="\x. True" and P="\x. 0 < x" and g=ln]) (auto intro: eventually_gt_at_top) lemma lim_exp_minus_1: "((\z::'a. (exp(z) - 1) / z) \ 1) (at 0)" for x :: "'a::{real_normed_field,banach}" proof - have "((\z::'a. exp(z) - 1) has_field_derivative 1) (at 0)" by (intro derivative_eq_intros | simp)+ then show ?thesis by (simp add: Deriv.has_field_derivative_iff) qed lemma ln_at_0: "LIM x at_right 0. ln (x::real) :> at_bot" by (rule filterlim_at_bot_at_right[where Q="\x. 0 < x" and P="\x. True" and g=exp]) (auto simp: eventually_at_filter) lemma ln_at_top: "LIM x at_top. ln (x::real) :> at_top" by (rule filterlim_at_top_at_top[where Q="\x. 0 < x" and P="\x. True" and g=exp]) (auto intro: eventually_gt_at_top) lemma filtermap_ln_at_top: "filtermap (ln::real \ real) at_top = at_top" by (intro filtermap_fun_inverse[of exp] exp_at_top ln_at_top) auto lemma filtermap_exp_at_top: "filtermap (exp::real \ real) at_top = at_top" by (intro filtermap_fun_inverse[of ln] exp_at_top ln_at_top) (auto simp: eventually_at_top_dense) lemma filtermap_ln_at_right: "filtermap ln (at_right (0::real)) = at_bot" by (auto intro!: filtermap_fun_inverse[where g="\x. exp x"] ln_at_0 simp: filterlim_at exp_at_bot) lemma tendsto_power_div_exp_0: "((\x. x ^ k / exp x) \ (0::real)) at_top" proof (induct k) case 0 show "((\x. x ^ 0 / exp x) \ (0::real)) at_top" by (simp add: inverse_eq_divide[symmetric]) (metis filterlim_compose[OF tendsto_inverse_0] exp_at_top filterlim_mono at_top_le_at_infinity order_refl) next case (Suc k) show ?case proof (rule lhospital_at_top_at_top) show "eventually (\x. DERIV (\x. x ^ Suc k) x :> (real (Suc k) * x^k)) at_top" by eventually_elim (intro derivative_eq_intros, auto) show "eventually (\x. DERIV exp x :> exp x) at_top" by eventually_elim auto show "eventually (\x. exp x \ 0) at_top" by auto from tendsto_mult[OF tendsto_const Suc, of "real (Suc k)"] show "((\x. real (Suc k) * x ^ k / exp x) \ 0) at_top" by simp qed (rule exp_at_top) qed subsubsection\ A couple of simple bounds\ lemma exp_plus_inverse_exp: fixes x::real shows "2 \ exp x + inverse (exp x)" proof - have "2 \ exp x + exp (-x)" using exp_ge_add_one_self [of x] exp_ge_add_one_self [of "-x"] by linarith then show ?thesis by (simp add: exp_minus) qed lemma real_le_x_sinh: fixes x::real assumes "0 \ x" shows "x \ (exp x - inverse(exp x)) / 2" proof - have *: "exp a - inverse(exp a) - 2*a \ exp b - inverse(exp b) - 2*b" if "a \ b" for a b::real using exp_plus_inverse_exp by (fastforce intro: derivative_eq_intros DERIV_nonneg_imp_nondecreasing [OF that]) show ?thesis using*[OF assms] by simp qed lemma real_le_abs_sinh: fixes x::real shows "abs x \ abs((exp x - inverse(exp x)) / 2)" proof (cases "0 \ x") case True show ?thesis using real_le_x_sinh [OF True] True by (simp add: abs_if) next case False have "-x \ (exp(-x) - inverse(exp(-x))) / 2" by (meson False linear neg_le_0_iff_le real_le_x_sinh) also have "\ \ \(exp x - inverse (exp x)) / 2\" by (metis (no_types, opaque_lifting) abs_divide abs_le_iff abs_minus_cancel add.inverse_inverse exp_minus minus_diff_eq order_refl) finally show ?thesis using False by linarith qed subsection\The general logarithm\ definition log :: "real \ real \ real" \ \logarithm of \<^term>\x\ to base \<^term>\a\\ where "log a x = ln x / ln a" lemma tendsto_log [tendsto_intros]: "(f \ a) F \ (g \ b) F \ 0 < a \ a \ 1 \ 0 < b \ ((\x. log (f x) (g x)) \ log a b) F" unfolding log_def by (intro tendsto_intros) auto lemma continuous_log: assumes "continuous F f" and "continuous F g" and "0 < f (Lim F (\x. x))" and "f (Lim F (\x. x)) \ 1" and "0 < g (Lim F (\x. x))" shows "continuous F (\x. log (f x) (g x))" using assms unfolding continuous_def by (rule tendsto_log) lemma continuous_at_within_log[continuous_intros]: assumes "continuous (at a within s) f" and "continuous (at a within s) g" and "0 < f a" and "f a \ 1" and "0 < g a" shows "continuous (at a within s) (\x. log (f x) (g x))" using assms unfolding continuous_within by (rule tendsto_log) lemma isCont_log[continuous_intros, simp]: assumes "isCont f a" "isCont g a" "0 < f a" "f a \ 1" "0 < g a" shows "isCont (\x. log (f x) (g x)) a" using assms unfolding continuous_at by (rule tendsto_log) lemma continuous_on_log[continuous_intros]: assumes "continuous_on s f" "continuous_on s g" and "\x\s. 0 < f x" "\x\s. f x \ 1" "\x\s. 0 < g x" shows "continuous_on s (\x. log (f x) (g x))" using assms unfolding continuous_on_def by (fast intro: tendsto_log) lemma exp_powr_real: fixes x::real shows "exp x powr y = exp (x*y)" by (simp add: powr_def) lemma powr_one_eq_one [simp]: "1 powr a = 1" by (simp add: powr_def) lemma powr_zero_eq_one [simp]: "x powr 0 = (if x = 0 then 0 else 1)" by (simp add: powr_def) lemma powr_one_gt_zero_iff [simp]: "x powr 1 = x \ 0 \ x" for x :: real by (auto simp: powr_def) declare powr_one_gt_zero_iff [THEN iffD2, simp] lemma powr_diff: fixes w:: "'a::{ln,real_normed_field}" shows "w powr (z1 - z2) = w powr z1 / w powr z2" by (simp add: powr_def algebra_simps exp_diff) lemma powr_mult: "0 \ x \ 0 \ y \ (x * y) powr a = (x powr a) * (y powr a)" for a x y :: real by (simp add: powr_def exp_add [symmetric] ln_mult distrib_left) lemma prod_powr_distrib: fixes x :: "'a \ real" assumes "\i. i\I \ x i \ 0" shows "(prod x I) powr r = (\i\I. x i powr r)" using assms by (induction I rule: infinite_finite_induct) (auto simp add: powr_mult prod_nonneg) lemma powr_ge_pzero [simp]: "0 \ x powr y" for x y :: real by (simp add: powr_def) lemma powr_non_neg[simp]: "\a powr x < 0" for a x::real using powr_ge_pzero[of a x] by arith lemma inverse_powr: "\y::real. 0 \ y \ inverse y powr a = inverse (y powr a)" by (simp add: exp_minus ln_inverse powr_def) lemma powr_divide: "\0 \ x; 0 \ y\ \ (x / y) powr a = (x powr a) / (y powr a)" for a b x :: real by (simp add: divide_inverse powr_mult inverse_powr) lemma powr_add: "x powr (a + b) = (x powr a) * (x powr b)" for a b x :: "'a::{ln,real_normed_field}" by (simp add: powr_def exp_add [symmetric] distrib_right) lemma powr_mult_base: "0 \ x \x * x powr y = x powr (1 + y)" for x :: real by (auto simp: powr_add) lemma powr_powr: "(x powr a) powr b = x powr (a * b)" for a b x :: real by (simp add: powr_def) lemma powr_power: fixes z:: "'a::{real_normed_field,ln}" shows "z \ 0 \ n \ 0 \ (z powr u) ^ n = z powr (of_nat n * u)" by (induction n) (auto simp: algebra_simps powr_add) lemma powr_powr_swap: "(x powr a) powr b = (x powr b) powr a" for a b x :: real by (simp add: powr_powr mult.commute) lemma powr_minus: "x powr (- a) = inverse (x powr a)" for a x :: "'a::{ln,real_normed_field}" by (simp add: powr_def exp_minus [symmetric]) lemma powr_minus_divide: "x powr (- a) = 1/(x powr a)" for a x :: "'a::{ln,real_normed_field}" by (simp add: divide_inverse powr_minus) lemma powr_sum: "x \ 0 \ finite A \ x powr sum f A = (\y\A. x powr f y)" by (simp add: powr_def exp_sum sum_distrib_right) lemma divide_powr_uminus: "a / b powr c = a * b powr (- c)" for a b c :: real by (simp add: powr_minus_divide) lemma powr_less_mono: "a < b \ 1 < x \ x powr a < x powr b" for a b x :: real by (simp add: powr_def) lemma powr_less_cancel: "x powr a < x powr b \ 1 < x \ a < b" for a b x :: real by (simp add: powr_def) lemma powr_less_cancel_iff [simp]: "1 < x \ x powr a < x powr b \ a < b" for a b x :: real by (blast intro: powr_less_cancel powr_less_mono) lemma powr_le_cancel_iff [simp]: "1 < x \ x powr a \ x powr b \ a \ b" for a b x :: real by (simp add: linorder_not_less [symmetric]) lemma powr_realpow: "0 < x \ x powr (real n) = x^n" by (induction n) (simp_all add: ac_simps powr_add) lemma powr_realpow': "(z :: real) \ 0 \ n \ 0 \ z powr of_nat n = z ^ n" by (cases "z = 0") (auto simp: powr_realpow) lemma powr_real_of_int': assumes "x \ 0" "x \ 0 \ n > 0" shows "x powr real_of_int n = power_int x n" by (metis assms exp_ln_iff exp_power_int nless_le power_int_eq_0_iff powr_def) lemma exp_minus_ge: fixes x::real shows "1 - x \ exp (-x)" by (smt (verit) exp_ge_add_one_self) lemma exp_minus_greater: fixes x::real shows "1 - x < exp (-x) \ x \ 0" by (smt (verit) exp_minus_ge exp_eq_one_iff exp_gt_zero ln_eq_minus_one ln_exp) lemma log_ln: "ln x = log (exp(1)) x" by (simp add: log_def) lemma DERIV_log: assumes "x > 0" shows "DERIV (\y. log b y) x :> 1 / (ln b * x)" proof - define lb where "lb = 1 / ln b" moreover have "DERIV (\y. lb * ln y) x :> lb / x" using \x > 0\ by (auto intro!: derivative_eq_intros) ultimately show ?thesis by (simp add: log_def) qed lemmas DERIV_log[THEN DERIV_chain2, derivative_intros] and DERIV_log[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemma powr_log_cancel [simp]: "0 < a \ a \ 1 \ 0 < x \ a powr (log a x) = x" by (simp add: powr_def log_def) lemma log_powr_cancel [simp]: "0 < a \ a \ 1 \ log a (a powr x) = x" by (simp add: log_def powr_def) lemma powr_eq_iff: "\y>0; a>1\ \ a powr x = y \ log a y = x" by auto lemma log_mult: "0 < x \ 0 < y \ log a (x * y) = log a x + log a y" by (simp add: log_def ln_mult divide_inverse distrib_right) lemma log_eq_div_ln_mult_log: "0 < b \ b \ 1 \ 0 < x \ log a x = (ln b/ln a) * log b x" by (simp add: log_def divide_inverse) text\Base 10 logarithms\ lemma log_base_10_eq1: "0 < x \ log 10 x = (ln (exp 1) / ln 10) * ln x" by (simp add: log_def) lemma log_base_10_eq2: "0 < x \ log 10 x = (log 10 (exp 1)) * ln x" by (simp add: log_def) lemma log_one [simp]: "log a 1 = 0" by (simp add: log_def) lemma log_eq_one [simp]: "0 < a \ a \ 1 \ log a a = 1" by (simp add: log_def) lemma log_inverse: "0 < x \ log a (inverse x) = - log a x" using ln_inverse log_def by auto lemma log_divide: "0 < x \ 0 < y \ log a (x/y) = log a x - log a y" by (simp add: log_mult divide_inverse log_inverse) lemma powr_gt_zero [simp]: "0 < x powr a \ x \ 0" for a x :: real by (simp add: powr_def) lemma powr_nonneg_iff[simp]: "a powr x \ 0 \ a = 0" for a x::real by (meson not_less powr_gt_zero) lemma log_add_eq_powr: "0 < b \ b \ 1 \ 0 < x \ log b x + y = log b (x * b powr y)" and add_log_eq_powr: "0 < b \ b \ 1 \ 0 < x \ y + log b x = log b (b powr y * x)" and log_minus_eq_powr: "0 < b \ b \ 1 \ 0 < x \ log b x - y = log b (x * b powr -y)" and minus_log_eq_powr: "0 < b \ b \ 1 \ 0 < x \ y - log b x = log b (b powr y / x)" by (simp_all add: log_mult log_divide) lemma log_less_cancel_iff [simp]: "1 < a \ 0 < x \ 0 < y \ log a x < log a y \ x < y" using powr_less_cancel_iff [of a] powr_log_cancel [of a x] powr_log_cancel [of a y] by (metis less_eq_real_def less_trans not_le zero_less_one) lemma log_inj: assumes "1 < b" shows "inj_on (log b) {0 <..}" proof (rule inj_onI, simp) fix x y assume pos: "0 < x" "0 < y" and *: "log b x = log b y" show "x = y" proof (cases rule: linorder_cases) assume "x = y" then show ?thesis by simp next assume "x < y" then have "log b x < log b y" using log_less_cancel_iff[OF \1 < b\] pos by simp then show ?thesis using * by simp next assume "y < x" then have "log b y < log b x" using log_less_cancel_iff[OF \1 < b\] pos by simp then show ?thesis using * by simp qed qed lemma log_le_cancel_iff [simp]: "1 < a \ 0 < x \ 0 < y \ log a x \ log a y \ x \ y" by (simp flip: linorder_not_less) lemma log_mono: "1 < a \ 0 < x \ x \ y \ log a x \ log a y" by simp lemma log_less: "1 < a \ 0 < x \ x < y \ log a x < log a y" by simp lemma zero_less_log_cancel_iff[simp]: "1 < a \ 0 < x \ 0 < log a x \ 1 < x" using log_less_cancel_iff[of a 1 x] by simp lemma zero_le_log_cancel_iff[simp]: "1 < a \ 0 < x \ 0 \ log a x \ 1 \ x" using log_le_cancel_iff[of a 1 x] by simp lemma log_less_zero_cancel_iff[simp]: "1 < a \ 0 < x \ log a x < 0 \ x < 1" using log_less_cancel_iff[of a x 1] by simp lemma log_le_zero_cancel_iff[simp]: "1 < a \ 0 < x \ log a x \ 0 \ x \ 1" using log_le_cancel_iff[of a x 1] by simp lemma one_less_log_cancel_iff[simp]: "1 < a \ 0 < x \ 1 < log a x \ a < x" using log_less_cancel_iff[of a a x] by simp lemma one_le_log_cancel_iff[simp]: "1 < a \ 0 < x \ 1 \ log a x \ a \ x" using log_le_cancel_iff[of a a x] by simp lemma log_less_one_cancel_iff[simp]: "1 < a \ 0 < x \ log a x < 1 \ x < a" using log_less_cancel_iff[of a x a] by simp lemma log_le_one_cancel_iff[simp]: "1 < a \ 0 < x \ log a x \ 1 \ x \ a" using log_le_cancel_iff[of a x a] by simp lemma le_log_iff: fixes b x y :: real assumes "1 < b" "x > 0" shows "y \ log b x \ b powr y \ x" using assms by (metis less_irrefl less_trans powr_le_cancel_iff powr_log_cancel zero_less_one) lemma less_log_iff: assumes "1 < b" "x > 0" shows "y < log b x \ b powr y < x" by (metis assms dual_order.strict_trans less_irrefl powr_less_cancel_iff powr_log_cancel zero_less_one) lemma assumes "1 < b" "x > 0" shows log_less_iff: "log b x < y \ x < b powr y" and log_le_iff: "log b x \ y \ x \ b powr y" using le_log_iff[OF assms, of y] less_log_iff[OF assms, of y] by auto lemmas powr_le_iff = le_log_iff[symmetric] and powr_less_iff = less_log_iff[symmetric] and less_powr_iff = log_less_iff[symmetric] and le_powr_iff = log_le_iff[symmetric] lemma le_log_of_power: assumes "b ^ n \ m" "1 < b" shows "n \ log b m" proof - from assms have "0 < m" by (metis less_trans zero_less_power less_le_trans zero_less_one) thus ?thesis using assms by (simp add: le_log_iff powr_realpow) qed lemma le_log2_of_power: "2 ^ n \ m \ n \ log 2 m" for m n :: nat using le_log_of_power[of 2] by simp lemma log_of_power_le: "\ m \ b ^ n; b > 1; m > 0 \ \ log b (real m) \ n" by (simp add: log_le_iff powr_realpow) lemma log2_of_power_le: "\ m \ 2 ^ n; m > 0 \ \ log 2 m \ n" for m n :: nat using log_of_power_le[of _ 2] by simp lemma log_of_power_less: "\ m < b ^ n; b > 1; m > 0 \ \ log b (real m) < n" by (simp add: log_less_iff powr_realpow) lemma log2_of_power_less: "\ m < 2 ^ n; m > 0 \ \ log 2 m < n" for m n :: nat using log_of_power_less[of _ 2] by simp lemma less_log_of_power: assumes "b ^ n < m" "1 < b" shows "n < log b m" proof - have "0 < m" by (metis assms less_trans zero_less_power zero_less_one) thus ?thesis using assms by (simp add: less_log_iff powr_realpow) qed lemma less_log2_of_power: "2 ^ n < m \ n < log 2 m" for m n :: nat using less_log_of_power[of 2] by simp lemma gr_one_powr[simp]: fixes x y :: real shows "\ x > 1; y > 0 \ \ 1 < x powr y" by(simp add: less_powr_iff) lemma log_pow_cancel [simp]: "a > 0 \ a \ 1 \ log a (a ^ b) = b" by (simp add: ln_realpow log_def) lemma floor_log_eq_powr_iff: "x > 0 \ b > 1 \ \log b x\ = k \ b powr k \ x \ x < b powr (k + 1)" by (auto simp: floor_eq_iff powr_le_iff less_powr_iff) lemma floor_log_nat_eq_powr_iff: fixes b n k :: nat shows "\ b \ 2; k > 0 \ \ floor (log b (real k)) = n \ b^n \ k \ k < b^(n+1)" by (auto simp: floor_log_eq_powr_iff powr_add powr_realpow of_nat_power[symmetric] of_nat_mult[symmetric] ac_simps simp del: of_nat_power of_nat_mult) lemma floor_log_nat_eq_if: fixes b n k :: nat assumes "b^n \ k" "k < b^(n+1)" "b \ 2" shows "floor (log b (real k)) = n" proof - have "k \ 1" using assms linorder_le_less_linear by force with assms show ?thesis by(simp add: floor_log_nat_eq_powr_iff) qed lemma ceiling_log_eq_powr_iff: "\ x > 0; b > 1 \ \ \log b x\ = int k + 1 \ b powr k < x \ x \ b powr (k + 1)" by (auto simp: ceiling_eq_iff powr_less_iff le_powr_iff) lemma ceiling_log_nat_eq_powr_iff: fixes b n k :: nat shows "\ b \ 2; k > 0 \ \ ceiling (log b (real k)) = int n + 1 \ (b^n < k \ k \ b^(n+1))" using ceiling_log_eq_powr_iff by (auto simp: powr_add powr_realpow of_nat_power[symmetric] of_nat_mult[symmetric] ac_simps simp del: of_nat_power of_nat_mult) lemma ceiling_log_nat_eq_if: fixes b n k :: nat assumes "b^n < k" "k \ b^(n+1)" "b \ 2" shows "\log (real b) (real k)\ = int n + 1" using assms ceiling_log_nat_eq_powr_iff by force lemma floor_log2_div2: fixes n :: nat assumes "n \ 2" shows "\log 2 (real n)\ = \log 2 (n div 2)\ + 1" proof cases assume "n=2" thus ?thesis by simp next let ?m = "n div 2" assume "n\2" hence "1 \ ?m" using assms by arith then obtain i where i: "2 ^ i \ ?m" "?m < 2 ^ (i + 1)" using ex_power_ivl1[of 2 ?m] by auto have "2^(i+1) \ 2*?m" using i(1) by simp also have "2*?m \ n" by arith finally have *: "2^(i+1) \ \" . have "n < 2^(i+1+1)" using i(2) by simp from floor_log_nat_eq_if[OF * this] floor_log_nat_eq_if[OF i] show ?thesis by simp qed lemma ceiling_log2_div2: assumes "n \ 2" shows "ceiling(log 2 (real n)) = ceiling(log 2 ((n-1) div 2 + 1)) + 1" proof cases assume "n=2" thus ?thesis by simp next let ?m = "(n-1) div 2 + 1" assume "n\2" hence "2 \ ?m" using assms by arith then obtain i where i: "2 ^ i < ?m" "?m \ 2 ^ (i + 1)" using ex_power_ivl2[of 2 ?m] by auto have "n \ 2*?m" by arith also have "2*?m \ 2 ^ ((i+1)+1)" using i(2) by simp finally have *: "n \ \" . have "2^(i+1) < n" using i(1) by (auto simp: less_Suc_eq_0_disj) from ceiling_log_nat_eq_if[OF this *] ceiling_log_nat_eq_if[OF i] show ?thesis by simp qed lemma powr_real_of_int: "x > 0 \ x powr real_of_int n = (if n \ 0 then x ^ nat n else inverse (x ^ nat (- n)))" using powr_realpow[of x "nat n"] powr_realpow[of x "nat (-n)"] by (auto simp: field_simps powr_minus) lemma powr_numeral [simp]: "0 \ x \ x powr (numeral n :: real) = x ^ (numeral n)" by (metis less_le power_zero_numeral powr_0 of_nat_numeral powr_realpow) lemma powr_int: assumes "x > 0" shows "x powr i = (if i \ 0 then x ^ nat i else 1/x ^ nat (-i))" by (simp add: assms inverse_eq_divide powr_real_of_int) lemma power_of_nat_log_ge: "b > 1 \ b ^ nat \log b x\ \ x" by (smt (verit) less_log_of_power of_nat_ceiling) lemma power_of_nat_log_le: assumes "b > 1" "x\1" shows "b ^ nat \log b x\ \ x" proof - have "\log b x\ \ 0" using assms by auto then show ?thesis by (smt (verit) assms le_log_iff of_int_floor_le powr_int) qed definition powr_real :: "real \ real \ real" where [code_abbrev, simp]: "powr_real = Transcendental.powr" lemma compute_powr_real [code]: "powr_real b i = (if b \ 0 then Code.abort (STR ''powr_real with nonpositive base'') (\_. powr_real b i) else if \i\ = i then (if 0 \ i then b ^ nat \i\ else 1 / b ^ nat \- i\) else Code.abort (STR ''powr_real with non-integer exponent'') (\_. powr_real b i))" for b i :: real by (auto simp: powr_int) lemma powr_one: "0 \ x \ x powr 1 = x" for x :: real using powr_realpow [of x 1] by simp lemma powr_neg_one: "0 < x \ x powr - 1 = 1/x" for x :: real using powr_int [of x "- 1"] by simp lemma powr_neg_numeral: "0 < x \ x powr - numeral n = 1/x ^ numeral n" for x :: real using powr_int [of x "- numeral n"] by simp lemma root_powr_inverse: "0 < n \ 0 < x \ root n x = x powr (1/n)" by (rule real_root_pos_unique) (auto simp: powr_realpow[symmetric] powr_powr) lemma ln_powr: "x \ 0 \ ln (x powr y) = y * ln x" for x :: real by (simp add: powr_def) lemma ln_root: "n > 0 \ b > 0 \ ln (root n b) = ln b / n" by (simp add: root_powr_inverse ln_powr) lemma ln_sqrt: "0 < x \ ln (sqrt x) = ln x / 2" by (simp add: ln_powr ln_powr[symmetric] mult.commute) lemma log_root: "n > 0 \ a > 0 \ log b (root n a) = log b a / n" by (simp add: log_def ln_root) lemma log_powr: "x \ 0 \ log b (x powr y) = y * log b x" by (simp add: log_def ln_powr) (* [simp] is not worth it, interferes with some proofs *) lemma log_nat_power: "0 < x \ log b (x^n) = real n * log b x" by (simp add: log_powr powr_realpow [symmetric]) lemma log_of_power_eq: assumes "m = b ^ n" "b > 1" shows "n = log b (real m)" proof - have "n = log b (b ^ n)" using assms(2) by (simp add: log_nat_power) also have "\ = log b m" using assms by simp finally show ?thesis . qed lemma log2_of_power_eq: "m = 2 ^ n \ n = log 2 m" for m n :: nat using log_of_power_eq[of _ 2] by simp lemma log_base_change: "0 < a \ a \ 1 \ log b x = log a x / log a b" by (simp add: log_def) lemma log_base_pow: "0 < a \ log (a ^ n) x = log a x / n" by (simp add: log_def ln_realpow) lemma log_base_powr: "a \ 0 \ log (a powr b) x = log a x / b" by (simp add: log_def ln_powr) lemma log_base_root: "n > 0 \ b > 0 \ log (root n b) x = n * (log b x)" by (simp add: log_def ln_root) lemma ln_bound: "0 < x \ ln x \ x" for x :: real using ln_le_minus_one by force lemma powr_less_one: fixes x::real assumes "1 < x" "y < 0" shows "x powr y < 1" using assms less_log_iff by force lemma powr_le_one_le: "\x y::real. 0 < x \ x \ 1 \ 1 \ y \ x powr y \ x" by (smt (verit) ln_gt_zero_imp_gt_one ln_le_cancel_iff ln_powr mult_le_cancel_right2) lemma powr_mono: fixes x :: real assumes "a \ b" and "1 \ x" shows "x powr a \ x powr b" using assms less_eq_real_def by auto lemma ge_one_powr_ge_zero: "1 \ x \ 0 \ a \ 1 \ x powr a" for x :: real using powr_mono by fastforce lemma powr_less_mono2: "0 < a \ 0 \ x \ x < y \ x powr a < y powr a" for x :: real by (simp add: powr_def) lemma powr_less_mono2_neg: "a < 0 \ 0 < x \ x < y \ y powr a < x powr a" for x :: real by (simp add: powr_def) lemma powr_mono2: "x powr a \ y powr a" if "0 \ a" "0 \ x" "x \ y" for x :: real using less_eq_real_def powr_less_mono2 that by auto lemma powr01_less_one: fixes a::real assumes "0 < a" "a < 1" shows "a powr e < 1 \ e>0" proof show "a powr e < 1 \ e>0" using assms not_less_iff_gr_or_eq powr_less_mono2_neg by fastforce show "e>0 \ a powr e < 1" by (metis assms less_eq_real_def powr_less_mono2 powr_one_eq_one) qed lemma powr_le1: "0 \ a \ 0 \ x \ x \ 1 \ x powr a \ 1" for x :: real using powr_mono2 by fastforce lemma powr_mono2': fixes a x y :: real assumes "a \ 0" "x > 0" "x \ y" shows "x powr a \ y powr a" proof - from assms have "x powr - a \ y powr - a" by (intro powr_mono2) simp_all with assms show ?thesis by (auto simp: powr_minus field_simps) qed +lemma powr_mono': "a \ (b::real) \ x \ 0 \ x \ 1 \ x powr b \ x powr a" + using powr_mono[of "-b" "-a" "inverse x"] by (auto simp: powr_def ln_inverse ln_div field_split_simps) + lemma powr_mono_both: fixes x :: real assumes "0 \ a" "a \ b" "1 \ x" "x \ y" shows "x powr a \ y powr b" by (meson assms order.trans powr_mono powr_mono2 zero_le_one) -lemma powr_mono': "a \ (b::real) \ x \ 0 \ x \ 1 \ x powr b \ x powr a" - using powr_mono[of "-b" "-a" "inverse x"] by (auto simp: powr_def ln_inverse ln_div field_split_simps) +lemma powr_mono_both': + fixes x :: real + assumes "a \ b" "b\0" "0 < x" "x \ y" "y \ 1" + shows "x powr a \ y powr b" + by (meson assms nless_le order.trans powr_mono' powr_mono2) lemma powr_less_mono': assumes "(x::real) > 0" "x < 1" "a < b" shows "x powr b < x powr a" by (metis assms log_powr_cancel order.strict_iff_order powr_mono') lemma powr_inj: "0 < a \ a \ 1 \ a powr x = a powr y \ x = y" for x :: real unfolding powr_def exp_inj_iff by simp lemma powr_half_sqrt: "0 \ x \ x powr (1/2) = sqrt x" by (simp add: powr_def root_powr_inverse sqrt_def) lemma powr_half_sqrt_powr: "0 \ x \ x powr (a/2) = sqrt(x powr a)" by (metis divide_inverse mult.left_neutral powr_ge_pzero powr_half_sqrt powr_powr) lemma square_powr_half [simp]: fixes x::real shows "x\<^sup>2 powr (1/2) = \x\" by (simp add: powr_half_sqrt) lemma ln_powr_bound: "1 \ x \ 0 < a \ ln x \ (x powr a) / a" for x :: real by (metis exp_gt_zero linear ln_eq_zero_iff ln_exp ln_less_self ln_powr mult.commute mult_imp_le_div_pos not_less powr_gt_zero) lemma ln_powr_bound2: fixes x :: real assumes "1 < x" and "0 < a" shows "(ln x) powr a \ (a powr a) * x" proof - from assms have "ln x \ (x powr (1 / a)) / (1 / a)" by (metis less_eq_real_def ln_powr_bound zero_less_divide_1_iff) also have "\ = a * (x powr (1 / a))" by simp finally have "(ln x) powr a \ (a * (x powr (1 / a))) powr a" by (metis assms less_imp_le ln_gt_zero powr_mono2) also have "\ = (a powr a) * ((x powr (1 / a)) powr a)" using assms powr_mult by auto also have "(x powr (1 / a)) powr a = x powr ((1 / a) * a)" by (rule powr_powr) also have "\ = x" using assms by auto finally show ?thesis . qed lemma tendsto_powr: fixes a b :: real assumes f: "(f \ a) F" and g: "(g \ b) F" and a: "a \ 0" shows "((\x. f x powr g x) \ a powr b) F" unfolding powr_def proof (rule filterlim_If) from f show "((\x. 0) \ (if a = 0 then 0 else exp (b * ln a))) (inf F (principal {x. f x = 0}))" by simp (auto simp: filterlim_iff eventually_inf_principal elim: eventually_mono dest: t1_space_nhds) from f g a show "((\x. exp (g x * ln (f x))) \ (if a = 0 then 0 else exp (b * ln a))) (inf F (principal {x. f x \ 0}))" by (auto intro!: tendsto_intros intro: tendsto_mono inf_le1) qed lemma tendsto_powr'[tendsto_intros]: fixes a :: real assumes f: "(f \ a) F" and g: "(g \ b) F" and a: "a \ 0 \ (b > 0 \ eventually (\x. f x \ 0) F)" shows "((\x. f x powr g x) \ a powr b) F" proof - from a consider "a \ 0" | "a = 0" "b > 0" "eventually (\x. f x \ 0) F" by auto then show ?thesis proof cases case 1 with f g show ?thesis by (rule tendsto_powr) next case 2 have "((\x. if f x = 0 then 0 else exp (g x * ln (f x))) \ 0) F" proof (intro filterlim_If) have "filterlim f (principal {0<..}) (inf F (principal {z. f z \ 0}))" using \eventually (\x. f x \ 0) F\ by (auto simp: filterlim_iff eventually_inf_principal eventually_principal elim: eventually_mono) moreover have "filterlim f (nhds a) (inf F (principal {z. f z \ 0}))" by (rule tendsto_mono[OF _ f]) simp_all ultimately have f: "filterlim f (at_right 0) (inf F (principal {x. f x \ 0}))" by (simp add: at_within_def filterlim_inf \a = 0\) have g: "(g \ b) (inf F (principal {z. f z \ 0}))" by (rule tendsto_mono[OF _ g]) simp_all show "((\x. exp (g x * ln (f x))) \ 0) (inf F (principal {x. f x \ 0}))" by (rule filterlim_compose[OF exp_at_bot] filterlim_tendsto_pos_mult_at_bot filterlim_compose[OF ln_at_0] f g \b > 0\)+ qed simp_all with \a = 0\ show ?thesis by (simp add: powr_def) qed qed lemma continuous_powr: assumes "continuous F f" and "continuous F g" and "f (Lim F (\x. x)) \ 0" shows "continuous F (\x. (f x) powr (g x :: real))" using assms unfolding continuous_def by (rule tendsto_powr) lemma continuous_at_within_powr[continuous_intros]: fixes f g :: "_ \ real" assumes "continuous (at a within s) f" and "continuous (at a within s) g" and "f a \ 0" shows "continuous (at a within s) (\x. (f x) powr (g x))" using assms unfolding continuous_within by (rule tendsto_powr) lemma isCont_powr[continuous_intros, simp]: fixes f g :: "_ \ real" assumes "isCont f a" "isCont g a" "f a \ 0" shows "isCont (\x. (f x) powr g x) a" using assms unfolding continuous_at by (rule tendsto_powr) lemma continuous_on_powr[continuous_intros]: fixes f g :: "_ \ real" assumes "continuous_on s f" "continuous_on s g" and "\x\s. f x \ 0" shows "continuous_on s (\x. (f x) powr (g x))" using assms unfolding continuous_on_def by (fast intro: tendsto_powr) lemma tendsto_powr2: fixes a :: real assumes f: "(f \ a) F" and g: "(g \ b) F" and "\\<^sub>F x in F. 0 \ f x" and b: "0 < b" shows "((\x. f x powr g x) \ a powr b) F" using tendsto_powr'[of f a F g b] assms by auto lemma has_derivative_powr[derivative_intros]: assumes g[derivative_intros]: "(g has_derivative g') (at x within X)" and f[derivative_intros]:"(f has_derivative f') (at x within X)" assumes pos: "0 < g x" and "x \ X" shows "((\x. g x powr f x::real) has_derivative (\h. (g x powr f x) * (f' h * ln (g x) + g' h * f x / g x))) (at x within X)" proof - have "\\<^sub>F x in at x within X. g x > 0" by (rule order_tendstoD[OF _ pos]) (rule has_derivative_continuous[OF g, unfolded continuous_within]) then obtain d where "d > 0" and pos': "\x'. x' \ X \ dist x' x < d \ 0 < g x'" using pos unfolding eventually_at by force have "((\x. exp (f x * ln (g x))) has_derivative (\h. (g x powr f x) * (f' h * ln (g x) + g' h * f x / g x))) (at x within X)" using pos by (auto intro!: derivative_eq_intros simp: field_split_simps powr_def) then show ?thesis by (rule has_derivative_transform_within[OF _ \d > 0\ \x \ X\]) (auto simp: powr_def dest: pos') qed lemma has_derivative_const_powr [derivative_intros]: assumes "\x. (f has_derivative f') (at x)" "a \ (0::real)" shows "((\x. a powr (f x)) has_derivative (\y. f' y * ln a * a powr (f x))) (at x)" using assms apply (simp add: powr_def) apply (rule assms derivative_eq_intros refl)+ done lemma has_real_derivative_const_powr [derivative_intros]: assumes "\x. (f has_real_derivative f' x) (at x)" "a \ (0::real)" shows "((\x. a powr (f x)) has_real_derivative (f' x * ln a * a powr (f x))) (at x)" using assms apply (simp add: powr_def) apply (rule assms derivative_eq_intros refl | simp)+ done lemma DERIV_powr: fixes r :: real assumes g: "DERIV g x :> m" and pos: "g x > 0" and f: "DERIV f x :> r" shows "DERIV (\x. g x powr f x) x :> (g x powr f x) * (r * ln (g x) + m * f x / g x)" using assms by (auto intro!: derivative_eq_intros ext simp: has_field_derivative_def algebra_simps) lemma DERIV_fun_powr: fixes r :: real assumes g: "DERIV g x :> m" and pos: "g x > 0" shows "DERIV (\x. (g x) powr r) x :> r * (g x) powr (r - of_nat 1) * m" using DERIV_powr[OF g pos DERIV_const, of r] pos by (simp add: powr_diff field_simps) lemma has_real_derivative_powr: assumes "z > 0" shows "((\z. z powr r) has_real_derivative r * z powr (r - 1)) (at z)" proof (subst DERIV_cong_ev[OF refl _ refl]) from assms have "eventually (\z. z \ 0) (nhds z)" by (intro t1_space_nhds) auto then show "eventually (\z. z powr r = exp (r * ln z)) (nhds z)" unfolding powr_def by eventually_elim simp from assms show "((\z. exp (r * ln z)) has_real_derivative r * z powr (r - 1)) (at z)" by (auto intro!: derivative_eq_intros simp: powr_def field_simps exp_diff) qed declare has_real_derivative_powr[THEN DERIV_chain2, derivative_intros] text \A more general version, by Johannes Hölzl\ lemma has_real_derivative_powr': fixes f g :: "real \ real" assumes "(f has_real_derivative f') (at x)" assumes "(g has_real_derivative g') (at x)" assumes "f x > 0" defines "h \ \x. f x powr g x * (g' * ln (f x) + f' * g x / f x)" shows "((\x. f x powr g x) has_real_derivative h x) (at x)" proof (subst DERIV_cong_ev[OF refl _ refl]) from assms have "isCont f x" by (simp add: DERIV_continuous) hence "f \x\ f x" by (simp add: continuous_at) with \f x > 0\ have "eventually (\x. f x > 0) (nhds x)" by (auto simp: tendsto_at_iff_tendsto_nhds dest: order_tendstoD) thus "eventually (\x. f x powr g x = exp (g x * ln (f x))) (nhds x)" by eventually_elim (simp add: powr_def) next from assms show "((\x. exp (g x * ln (f x))) has_real_derivative h x) (at x)" by (auto intro!: derivative_eq_intros simp: h_def powr_def) qed lemma tendsto_zero_powrI: assumes "(f \ (0::real)) F" "(g \ b) F" "\\<^sub>F x in F. 0 \ f x" "0 < b" shows "((\x. f x powr g x) \ 0) F" using tendsto_powr2[OF assms] by simp lemma continuous_on_powr': fixes f g :: "_ \ real" assumes "continuous_on s f" "continuous_on s g" and "\x\s. f x \ 0 \ (f x = 0 \ g x > 0)" shows "continuous_on s (\x. (f x) powr (g x))" unfolding continuous_on_def proof fix x assume x: "x \ s" from assms x show "((\x. f x powr g x) \ f x powr g x) (at x within s)" proof (cases "f x = 0") case True from assms(3) have "eventually (\x. f x \ 0) (at x within s)" by (auto simp: at_within_def eventually_inf_principal) with True x assms show ?thesis by (auto intro!: tendsto_zero_powrI[of f _ g "g x"] simp: continuous_on_def) next case False with assms x show ?thesis by (auto intro!: tendsto_powr' simp: continuous_on_def) qed qed lemma tendsto_neg_powr: assumes "s < 0" and f: "LIM x F. f x :> at_top" shows "((\x. f x powr s) \ (0::real)) F" proof - have "((\x. exp (s * ln (f x))) \ (0::real)) F" (is "?X") by (auto intro!: filterlim_compose[OF exp_at_bot] filterlim_compose[OF ln_at_top] filterlim_tendsto_neg_mult_at_bot assms) also have "?X \ ((\x. f x powr s) \ (0::real)) F" using f filterlim_at_top_dense[of f F] by (intro filterlim_cong[OF refl refl]) (auto simp: neq_iff powr_def elim: eventually_mono) finally show ?thesis . qed lemma tendsto_exp_limit_at_right: "((\y. (1 + x * y) powr (1 / y)) \ exp x) (at_right 0)" for x :: real proof (cases "x = 0") case True then show ?thesis by simp next case False have "((\y. ln (1 + x * y)::real) has_real_derivative 1 * x) (at 0)" by (auto intro!: derivative_eq_intros) then have "((\y. ln (1 + x * y) / y) \ x) (at 0)" by (auto simp: has_field_derivative_def field_has_derivative_at) then have *: "((\y. exp (ln (1 + x * y) / y)) \ exp x) (at 0)" by (rule tendsto_intros) then show ?thesis proof (rule filterlim_mono_eventually) show "eventually (\xa. exp (ln (1 + x * xa) / xa) = (1 + x * xa) powr (1 / xa)) (at_right 0)" unfolding eventually_at_right[OF zero_less_one] using False by (intro exI[of _ "1 / \x\"]) (auto simp: field_simps powr_def abs_if add_nonneg_eq_0_iff) qed (simp_all add: at_eq_sup_left_right) qed lemma tendsto_exp_limit_at_top: "((\y. (1 + x / y) powr y) \ exp x) at_top" for x :: real by (simp add: filterlim_at_top_to_right inverse_eq_divide tendsto_exp_limit_at_right) lemma tendsto_exp_limit_sequentially: "(\n. (1 + x / n) ^ n) \ exp x" for x :: real proof (rule filterlim_mono_eventually) from reals_Archimedean2 [of "\x\"] obtain n :: nat where *: "real n > \x\" .. then have "eventually (\n :: nat. 0 < 1 + x / real n) at_top" by (intro eventually_sequentiallyI [of n]) (auto simp: field_split_simps) then show "eventually (\n. (1 + x / n) powr n = (1 + x / n) ^ n) at_top" by (rule eventually_mono) (erule powr_realpow) show "(\n. (1 + x / real n) powr real n) \ exp x" by (rule filterlim_compose [OF tendsto_exp_limit_at_top filterlim_real_sequentially]) qed auto subsection \Sine and Cosine\ definition sin_coeff :: "nat \ real" where "sin_coeff = (\n. if even n then 0 else (- 1) ^ ((n - Suc 0) div 2) / (fact n))" definition cos_coeff :: "nat \ real" where "cos_coeff = (\n. if even n then ((- 1) ^ (n div 2)) / (fact n) else 0)" definition sin :: "'a \ 'a::{real_normed_algebra_1,banach}" where "sin = (\x. \n. sin_coeff n *\<^sub>R x^n)" definition cos :: "'a \ 'a::{real_normed_algebra_1,banach}" where "cos = (\x. \n. cos_coeff n *\<^sub>R x^n)" lemma sin_coeff_0 [simp]: "sin_coeff 0 = 0" unfolding sin_coeff_def by simp lemma cos_coeff_0 [simp]: "cos_coeff 0 = 1" unfolding cos_coeff_def by simp lemma sin_coeff_Suc: "sin_coeff (Suc n) = cos_coeff n / real (Suc n)" unfolding cos_coeff_def sin_coeff_def by (simp del: mult_Suc) lemma cos_coeff_Suc: "cos_coeff (Suc n) = - sin_coeff n / real (Suc n)" unfolding cos_coeff_def sin_coeff_def by (simp del: mult_Suc) (auto elim: oddE) lemma summable_norm_sin: "summable (\n. norm (sin_coeff n *\<^sub>R x^n))" for x :: "'a::{real_normed_algebra_1,banach}" proof (rule summable_comparison_test [OF _ summable_norm_exp]) show "\N. \n\N. norm (norm (sin_coeff n *\<^sub>R x ^ n)) \ norm (x ^ n /\<^sub>R fact n)" unfolding sin_coeff_def by (auto simp: divide_inverse abs_mult power_abs [symmetric] zero_le_mult_iff) qed lemma summable_norm_cos: "summable (\n. norm (cos_coeff n *\<^sub>R x^n))" for x :: "'a::{real_normed_algebra_1,banach}" proof (rule summable_comparison_test [OF _ summable_norm_exp]) show "\N. \n\N. norm (norm (cos_coeff n *\<^sub>R x ^ n)) \ norm (x ^ n /\<^sub>R fact n)" unfolding cos_coeff_def by (auto simp: divide_inverse abs_mult power_abs [symmetric] zero_le_mult_iff) qed lemma sin_converges: "(\n. sin_coeff n *\<^sub>R x^n) sums sin x" unfolding sin_def by (metis (full_types) summable_norm_cancel summable_norm_sin summable_sums) lemma cos_converges: "(\n. cos_coeff n *\<^sub>R x^n) sums cos x" unfolding cos_def by (metis (full_types) summable_norm_cancel summable_norm_cos summable_sums) lemma sin_of_real: "sin (of_real x) = of_real (sin x)" for x :: real proof - have "(\n. of_real (sin_coeff n *\<^sub>R x^n)) = (\n. sin_coeff n *\<^sub>R (of_real x)^n)" proof show "of_real (sin_coeff n *\<^sub>R x^n) = sin_coeff n *\<^sub>R of_real x^n" for n by (simp add: scaleR_conv_of_real) qed also have "\ sums (sin (of_real x))" by (rule sin_converges) finally have "(\n. of_real (sin_coeff n *\<^sub>R x^n)) sums (sin (of_real x))" . then show ?thesis using sums_unique2 sums_of_real [OF sin_converges] by blast qed corollary sin_in_Reals [simp]: "z \ \ \ sin z \ \" by (metis Reals_cases Reals_of_real sin_of_real) lemma cos_of_real: "cos (of_real x) = of_real (cos x)" for x :: real proof - have "(\n. of_real (cos_coeff n *\<^sub>R x^n)) = (\n. cos_coeff n *\<^sub>R (of_real x)^n)" proof show "of_real (cos_coeff n *\<^sub>R x^n) = cos_coeff n *\<^sub>R of_real x^n" for n by (simp add: scaleR_conv_of_real) qed also have "\ sums (cos (of_real x))" by (rule cos_converges) finally have "(\n. of_real (cos_coeff n *\<^sub>R x^n)) sums (cos (of_real x))" . then show ?thesis using sums_unique2 sums_of_real [OF cos_converges] by blast qed corollary cos_in_Reals [simp]: "z \ \ \ cos z \ \" by (metis Reals_cases Reals_of_real cos_of_real) lemma diffs_sin_coeff: "diffs sin_coeff = cos_coeff" by (simp add: diffs_def sin_coeff_Suc del: of_nat_Suc) lemma diffs_cos_coeff: "diffs cos_coeff = (\n. - sin_coeff n)" by (simp add: diffs_def cos_coeff_Suc del: of_nat_Suc) lemma sin_int_times_real: "sin (of_int m * of_real x) = of_real (sin (of_int m * x))" by (metis sin_of_real of_real_mult of_real_of_int_eq) lemma cos_int_times_real: "cos (of_int m * of_real x) = of_real (cos (of_int m * x))" by (metis cos_of_real of_real_mult of_real_of_int_eq) text \Now at last we can get the derivatives of exp, sin and cos.\ lemma DERIV_sin [simp]: "DERIV sin x :> cos x" for x :: "'a::{real_normed_field,banach}" unfolding sin_def cos_def scaleR_conv_of_real apply (rule DERIV_cong) apply (rule termdiffs [where K="of_real (norm x) + 1 :: 'a"]) apply (simp_all add: norm_less_p1 diffs_of_real diffs_sin_coeff diffs_cos_coeff summable_minus_iff scaleR_conv_of_real [symmetric] summable_norm_sin [THEN summable_norm_cancel] summable_norm_cos [THEN summable_norm_cancel]) done declare DERIV_sin[THEN DERIV_chain2, derivative_intros] and DERIV_sin[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemmas has_derivative_sin[derivative_intros] = DERIV_sin[THEN DERIV_compose_FDERIV] lemma DERIV_cos [simp]: "DERIV cos x :> - sin x" for x :: "'a::{real_normed_field,banach}" unfolding sin_def cos_def scaleR_conv_of_real apply (rule DERIV_cong) apply (rule termdiffs [where K="of_real (norm x) + 1 :: 'a"]) apply (simp_all add: norm_less_p1 diffs_of_real diffs_minus suminf_minus diffs_sin_coeff diffs_cos_coeff summable_minus_iff scaleR_conv_of_real [symmetric] summable_norm_sin [THEN summable_norm_cancel] summable_norm_cos [THEN summable_norm_cancel]) done declare DERIV_cos[THEN DERIV_chain2, derivative_intros] and DERIV_cos[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemmas has_derivative_cos[derivative_intros] = DERIV_cos[THEN DERIV_compose_FDERIV] lemma isCont_sin: "isCont sin x" for x :: "'a::{real_normed_field,banach}" by (rule DERIV_sin [THEN DERIV_isCont]) lemma continuous_on_sin_real: "continuous_on {a..b} sin" for a::real using continuous_at_imp_continuous_on isCont_sin by blast lemma isCont_cos: "isCont cos x" for x :: "'a::{real_normed_field,banach}" by (rule DERIV_cos [THEN DERIV_isCont]) lemma continuous_on_cos_real: "continuous_on {a..b} cos" for a::real using continuous_at_imp_continuous_on isCont_cos by blast context fixes f :: "'a::t2_space \ 'b::{real_normed_field,banach}" begin lemma isCont_sin' [simp]: "isCont f a \ isCont (\x. sin (f x)) a" by (rule isCont_o2 [OF _ isCont_sin]) lemma isCont_cos' [simp]: "isCont f a \ isCont (\x. cos (f x)) a" by (rule isCont_o2 [OF _ isCont_cos]) lemma tendsto_sin [tendsto_intros]: "(f \ a) F \ ((\x. sin (f x)) \ sin a) F" by (rule isCont_tendsto_compose [OF isCont_sin]) lemma tendsto_cos [tendsto_intros]: "(f \ a) F \ ((\x. cos (f x)) \ cos a) F" by (rule isCont_tendsto_compose [OF isCont_cos]) lemma continuous_sin [continuous_intros]: "continuous F f \ continuous F (\x. sin (f x))" unfolding continuous_def by (rule tendsto_sin) lemma continuous_on_sin [continuous_intros]: "continuous_on s f \ continuous_on s (\x. sin (f x))" unfolding continuous_on_def by (auto intro: tendsto_sin) lemma continuous_cos [continuous_intros]: "continuous F f \ continuous F (\x. cos (f x))" unfolding continuous_def by (rule tendsto_cos) lemma continuous_on_cos [continuous_intros]: "continuous_on s f \ continuous_on s (\x. cos (f x))" unfolding continuous_on_def by (auto intro: tendsto_cos) end lemma continuous_within_sin: "continuous (at z within s) sin" for z :: "'a::{real_normed_field,banach}" by (simp add: continuous_within tendsto_sin) lemma continuous_within_cos: "continuous (at z within s) cos" for z :: "'a::{real_normed_field,banach}" by (simp add: continuous_within tendsto_cos) subsection \Properties of Sine and Cosine\ lemma sin_zero [simp]: "sin 0 = 0" by (simp add: sin_def sin_coeff_def scaleR_conv_of_real) lemma cos_zero [simp]: "cos 0 = 1" by (simp add: cos_def cos_coeff_def scaleR_conv_of_real) lemma DERIV_fun_sin: "DERIV g x :> m \ DERIV (\x. sin (g x)) x :> cos (g x) * m" by (fact derivative_intros) lemma DERIV_fun_cos: "DERIV g x :> m \ DERIV (\x. cos(g x)) x :> - sin (g x) * m" by (fact derivative_intros) subsection \Deriving the Addition Formulas\ text \The product of two cosine series.\ lemma cos_x_cos_y: fixes x :: "'a::{real_normed_field,banach}" shows "(\p. \n\p. if even p \ even n then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) sums (cos x * cos y)" proof - have "(cos_coeff n * cos_coeff (p - n)) *\<^sub>R (x^n * y^(p - n)) = (if even p \ even n then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p - n) else 0)" if "n \ p" for n p :: nat proof - from that have *: "even n \ even p \ (-1) ^ (n div 2) * (-1) ^ ((p - n) div 2) = (-1 :: real) ^ (p div 2)" by (metis div_add power_add le_add_diff_inverse odd_add) with that show ?thesis by (auto simp: algebra_simps cos_coeff_def binomial_fact) qed then have "(\p. \n\p. if even p \ even n then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) = (\p. \n\p. (cos_coeff n * cos_coeff (p - n)) *\<^sub>R (x^n * y^(p-n)))" by simp also have "\ = (\p. \n\p. (cos_coeff n *\<^sub>R x^n) * (cos_coeff (p - n) *\<^sub>R y^(p-n)))" by (simp add: algebra_simps) also have "\ sums (cos x * cos y)" using summable_norm_cos by (auto simp: cos_def scaleR_conv_of_real intro!: Cauchy_product_sums) finally show ?thesis . qed text \The product of two sine series.\ lemma sin_x_sin_y: fixes x :: "'a::{real_normed_field,banach}" shows "(\p. \n\p. if even p \ odd n then - ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) sums (sin x * sin y)" proof - have "(sin_coeff n * sin_coeff (p - n)) *\<^sub>R (x^n * y^(p-n)) = (if even p \ odd n then -((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0)" if "n \ p" for n p :: nat proof - have "(-1) ^ ((n - Suc 0) div 2) * (-1) ^ ((p - Suc n) div 2) = - ((-1 :: real) ^ (p div 2))" if np: "odd n" "even p" proof - have "p > 0" using \n \ p\ neq0_conv that(1) by blast then have \
: "(- 1::real) ^ (p div 2 - Suc 0) = - ((- 1) ^ (p div 2))" using \even p\ by (auto simp add: dvd_def power_eq_if) from \n \ p\ np have *: "n - Suc 0 + (p - Suc n) = p - Suc (Suc 0)" "Suc (Suc 0) \ p" by arith+ have "(p - Suc (Suc 0)) div 2 = p div 2 - Suc 0" by simp with \n \ p\ np \
* show ?thesis by (simp add: flip: div_add power_add) qed then show ?thesis using \n\p\ by (auto simp: algebra_simps sin_coeff_def binomial_fact) qed then have "(\p. \n\p. if even p \ odd n then - ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) = (\p. \n\p. (sin_coeff n * sin_coeff (p - n)) *\<^sub>R (x^n * y^(p-n)))" by simp also have "\ = (\p. \n\p. (sin_coeff n *\<^sub>R x^n) * (sin_coeff (p - n) *\<^sub>R y^(p-n)))" by (simp add: algebra_simps) also have "\ sums (sin x * sin y)" using summable_norm_sin by (auto simp: sin_def scaleR_conv_of_real intro!: Cauchy_product_sums) finally show ?thesis . qed lemma sums_cos_x_plus_y: fixes x :: "'a::{real_normed_field,banach}" shows "(\p. \n\p. if even p then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) sums cos (x + y)" proof - have "(\n\p. if even p then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) = cos_coeff p *\<^sub>R ((x + y) ^ p)" for p :: nat proof - have "(\n\p. if even p then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) = (if even p then \n\p. ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0)" by simp also have "\ = (if even p then of_real ((-1) ^ (p div 2) / (fact p)) * (\n\p. (p choose n) *\<^sub>R (x^n) * y^(p-n)) else 0)" by (auto simp: sum_distrib_left field_simps scaleR_conv_of_real nonzero_of_real_divide) also have "\ = cos_coeff p *\<^sub>R ((x + y) ^ p)" by (simp add: cos_coeff_def binomial_ring [of x y] scaleR_conv_of_real atLeast0AtMost) finally show ?thesis . qed then have "(\p. \n\p. if even p then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) = (\p. cos_coeff p *\<^sub>R ((x+y)^p))" by simp also have "\ sums cos (x + y)" by (rule cos_converges) finally show ?thesis . qed theorem cos_add: fixes x :: "'a::{real_normed_field,banach}" shows "cos (x + y) = cos x * cos y - sin x * sin y" proof - have "(if even p \ even n then ((- 1) ^ (p div 2) * int (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) - (if even p \ odd n then - ((- 1) ^ (p div 2) * int (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) = (if even p then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0)" if "n \ p" for n p :: nat by simp then have "(\p. \n\p. (if even p then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0)) sums (cos x * cos y - sin x * sin y)" using sums_diff [OF cos_x_cos_y [of x y] sin_x_sin_y [of x y]] by (simp add: sum_subtractf [symmetric]) then show ?thesis by (blast intro: sums_cos_x_plus_y sums_unique2) qed lemma sin_minus_converges: "(\n. - (sin_coeff n *\<^sub>R (-x)^n)) sums sin x" proof - have [simp]: "\n. - (sin_coeff n *\<^sub>R (-x)^n) = (sin_coeff n *\<^sub>R x^n)" by (auto simp: sin_coeff_def elim!: oddE) show ?thesis by (simp add: sin_def summable_norm_sin [THEN summable_norm_cancel, THEN summable_sums]) qed lemma sin_minus [simp]: "sin (- x) = - sin x" for x :: "'a::{real_normed_algebra_1,banach}" using sin_minus_converges [of x] by (auto simp: sin_def summable_norm_sin [THEN summable_norm_cancel] suminf_minus sums_iff equation_minus_iff) lemma cos_minus_converges: "(\n. (cos_coeff n *\<^sub>R (-x)^n)) sums cos x" proof - have [simp]: "\n. (cos_coeff n *\<^sub>R (-x)^n) = (cos_coeff n *\<^sub>R x^n)" by (auto simp: Transcendental.cos_coeff_def elim!: evenE) show ?thesis by (simp add: cos_def summable_norm_cos [THEN summable_norm_cancel, THEN summable_sums]) qed lemma cos_minus [simp]: "cos (-x) = cos x" for x :: "'a::{real_normed_algebra_1,banach}" using cos_minus_converges [of x] by (metis cos_def sums_unique) lemma cos_abs_real [simp]: "cos \x :: real\ = cos x" by (simp add: abs_if) lemma sin_cos_squared_add [simp]: "(sin x)\<^sup>2 + (cos x)\<^sup>2 = 1" for x :: "'a::{real_normed_field,banach}" using cos_add [of x "-x"] by (simp add: power2_eq_square algebra_simps) lemma sin_cos_squared_add2 [simp]: "(cos x)\<^sup>2 + (sin x)\<^sup>2 = 1" for x :: "'a::{real_normed_field,banach}" by (subst add.commute, rule sin_cos_squared_add) lemma sin_cos_squared_add3 [simp]: "cos x * cos x + sin x * sin x = 1" for x :: "'a::{real_normed_field,banach}" using sin_cos_squared_add2 [unfolded power2_eq_square] . lemma sin_squared_eq: "(sin x)\<^sup>2 = 1 - (cos x)\<^sup>2" for x :: "'a::{real_normed_field,banach}" unfolding eq_diff_eq by (rule sin_cos_squared_add) lemma cos_squared_eq: "(cos x)\<^sup>2 = 1 - (sin x)\<^sup>2" for x :: "'a::{real_normed_field,banach}" unfolding eq_diff_eq by (rule sin_cos_squared_add2) lemma abs_sin_le_one [simp]: "\sin x\ \ 1" for x :: real by (rule power2_le_imp_le) (simp_all add: sin_squared_eq) lemma sin_ge_minus_one [simp]: "- 1 \ sin x" for x :: real using abs_sin_le_one [of x] by (simp add: abs_le_iff) lemma sin_le_one [simp]: "sin x \ 1" for x :: real using abs_sin_le_one [of x] by (simp add: abs_le_iff) lemma abs_cos_le_one [simp]: "\cos x\ \ 1" for x :: real by (rule power2_le_imp_le) (simp_all add: cos_squared_eq) lemma cos_ge_minus_one [simp]: "- 1 \ cos x" for x :: real using abs_cos_le_one [of x] by (simp add: abs_le_iff) lemma cos_le_one [simp]: "cos x \ 1" for x :: real using abs_cos_le_one [of x] by (simp add: abs_le_iff) lemma cos_diff: "cos (x - y) = cos x * cos y + sin x * sin y" for x :: "'a::{real_normed_field,banach}" using cos_add [of x "- y"] by simp lemma cos_double: "cos(2*x) = (cos x)\<^sup>2 - (sin x)\<^sup>2" for x :: "'a::{real_normed_field,banach}" using cos_add [where x=x and y=x] by (simp add: power2_eq_square) lemma sin_cos_le1: "\sin x * sin y + cos x * cos y\ \ 1" for x :: real using cos_diff [of x y] by (metis abs_cos_le_one add.commute) lemma DERIV_fun_pow: "DERIV g x :> m \ DERIV (\x. (g x) ^ n) x :> real n * (g x) ^ (n - 1) * m" by (auto intro!: derivative_eq_intros simp:) lemma DERIV_fun_exp: "DERIV g x :> m \ DERIV (\x. exp (g x)) x :> exp (g x) * m" by (auto intro!: derivative_intros) subsection \The Constant Pi\ definition pi :: real where "pi = 2 * (THE x. 0 \ x \ x \ 2 \ cos x = 0)" text \Show that there's a least positive \<^term>\x\ with \<^term>\cos x = 0\; hence define pi.\ lemma sin_paired: "(\n. (- 1) ^ n / (fact (2 * n + 1)) * x ^ (2 * n + 1)) sums sin x" for x :: real proof - have "(\n. \k = n*2..n. 0 < ?f n" proof fix n :: nat let ?k2 = "real (Suc (Suc (4 * n)))" let ?k3 = "real (Suc (Suc (Suc (4 * n))))" have "x * x < ?k2 * ?k3" using assms by (intro mult_strict_mono', simp_all) then have "x * x * x * x ^ (n * 4) < ?k2 * ?k3 * x * x ^ (n * 4)" by (intro mult_strict_right_mono zero_less_power \0 < x\) then show "0 < ?f n" by (simp add: ac_simps divide_less_eq) qed have sums: "?f sums sin x" by (rule sin_paired [THEN sums_group]) simp show "0 < sin x" unfolding sums_unique [OF sums] using sums_summable [OF sums] pos by (simp add: suminf_pos) qed lemma cos_double_less_one: "0 < x \ x < 2 \ cos (2 * x) < 1" for x :: real using sin_gt_zero_02 [where x = x] by (auto simp: cos_squared_eq cos_double) lemma cos_paired: "(\n. (- 1) ^ n / (fact (2 * n)) * x ^ (2 * n)) sums cos x" for x :: real proof - have "(\n. \k = n * 2.. real" assumes f: "summable f" and fplus: "\d. 0 < f (k + (Suc(Suc 0) * d)) + f (k + ((Suc (Suc 0) * d) + 1))" shows "sum f {..n. \n = n * Suc (Suc 0)..n. f (n + k))" proof (rule sums_group) show "(\n. f (n + k)) sums (\n. f (n + k))" by (simp add: f summable_iff_shift summable_sums) qed auto with fplus have "0 < (\n. f (n + k))" apply (simp add: add.commute) apply (metis (no_types, lifting) suminf_pos summable_def sums_unique) done then show ?thesis by (simp add: f suminf_minus_initial_segment) qed lemma cos_two_less_zero [simp]: "cos 2 < (0::real)" proof - note fact_Suc [simp del] from sums_minus [OF cos_paired] have *: "(\n. - ((- 1) ^ n * 2 ^ (2 * n) / fact (2 * n))) sums - cos (2::real)" by simp then have sm: "summable (\n. - ((- 1::real) ^ n * 2 ^ (2 * n) / (fact (2 * n))))" by (rule sums_summable) have "0 < (\nnn. - ((- 1) ^ n * 2 ^ (2 * n) / (fact (2 * n))))" proof - { fix d let ?six4d = "Suc (Suc (Suc (Suc (Suc (Suc (4 * d))))))" have "(4::real) * (fact (?six4d)) < (Suc (Suc (?six4d)) * fact (Suc (?six4d)))" unfolding of_nat_mult by (rule mult_strict_mono) (simp_all add: fact_less_mono) then have "(4::real) * (fact (?six4d)) < (fact (Suc (Suc (?six4d))))" by (simp only: fact_Suc [of "Suc (?six4d)"] of_nat_mult of_nat_fact) then have "(4::real) * inverse (fact (Suc (Suc (?six4d)))) < inverse (fact (?six4d))" by (simp add: inverse_eq_divide less_divide_eq) } then show ?thesis by (force intro!: sum_pos_lt_pair [OF sm] simp add: divide_inverse algebra_simps) qed ultimately have "0 < (\n. - ((- 1::real) ^ n * 2 ^ (2 * n) / (fact (2 * n))))" by (rule order_less_trans) moreover from * have "- cos 2 = (\n. - ((- 1::real) ^ n * 2 ^ (2 * n) / (fact (2 * n))))" by (rule sums_unique) ultimately have "(0::real) < - cos 2" by simp then show ?thesis by simp qed lemmas cos_two_neq_zero [simp] = cos_two_less_zero [THEN less_imp_neq] lemmas cos_two_le_zero [simp] = cos_two_less_zero [THEN order_less_imp_le] lemma cos_is_zero: "\!x::real. 0 \ x \ x \ 2 \ cos x = 0" proof (rule ex_ex1I) show "\x::real. 0 \ x \ x \ 2 \ cos x = 0" by (rule IVT2) simp_all next fix a b :: real assume ab: "0 \ a \ a \ 2 \ cos a = 0" "0 \ b \ b \ 2 \ cos b = 0" have cosd: "\x::real. cos differentiable (at x)" unfolding real_differentiable_def by (auto intro: DERIV_cos) show "a = b" proof (cases a b rule: linorder_cases) case less then obtain z where "a < z" "z < b" "(cos has_real_derivative 0) (at z)" using Rolle by (metis cosd continuous_on_cos_real ab) then have "sin z = 0" using DERIV_cos DERIV_unique neg_equal_0_iff_equal by blast then show ?thesis by (metis \a < z\ \z < b\ ab order_less_le_trans less_le sin_gt_zero_02) next case greater then obtain z where "b < z" "z < a" "(cos has_real_derivative 0) (at z)" using Rolle by (metis cosd continuous_on_cos_real ab) then have "sin z = 0" using DERIV_cos DERIV_unique neg_equal_0_iff_equal by blast then show ?thesis by (metis \b < z\ \z < a\ ab order_less_le_trans less_le sin_gt_zero_02) qed auto qed lemma pi_half: "pi/2 = (THE x. 0 \ x \ x \ 2 \ cos x = 0)" by (simp add: pi_def) lemma cos_pi_half [simp]: "cos (pi/2) = 0" by (simp add: pi_half cos_is_zero [THEN theI']) lemma cos_of_real_pi_half [simp]: "cos ((of_real pi/2) :: 'a) = 0" if "SORT_CONSTRAINT('a::{real_field,banach,real_normed_algebra_1})" by (metis cos_pi_half cos_of_real eq_numeral_simps(4) nonzero_of_real_divide of_real_0 of_real_numeral) lemma pi_half_gt_zero [simp]: "0 < pi/2" proof - have "0 \ pi/2" by (simp add: pi_half cos_is_zero [THEN theI']) then show ?thesis by (metis cos_pi_half cos_zero less_eq_real_def one_neq_zero) qed lemmas pi_half_neq_zero [simp] = pi_half_gt_zero [THEN less_imp_neq, symmetric] lemmas pi_half_ge_zero [simp] = pi_half_gt_zero [THEN order_less_imp_le] lemma pi_half_less_two [simp]: "pi/2 < 2" proof - have "pi/2 \ 2" by (simp add: pi_half cos_is_zero [THEN theI']) then show ?thesis by (metis cos_pi_half cos_two_neq_zero le_less) qed lemmas pi_half_neq_two [simp] = pi_half_less_two [THEN less_imp_neq] lemmas pi_half_le_two [simp] = pi_half_less_two [THEN order_less_imp_le] lemma pi_gt_zero [simp]: "0 < pi" using pi_half_gt_zero by simp lemma pi_ge_zero [simp]: "0 \ pi" by (rule pi_gt_zero [THEN order_less_imp_le]) lemma pi_neq_zero [simp]: "pi \ 0" by (rule pi_gt_zero [THEN less_imp_neq, symmetric]) lemma pi_not_less_zero [simp]: "\ pi < 0" by (simp add: linorder_not_less) lemma minus_pi_half_less_zero: "-(pi/2) < 0" by simp lemma m2pi_less_pi: "- (2*pi) < pi" by simp lemma sin_pi_half [simp]: "sin(pi/2) = 1" using sin_cos_squared_add2 [where x = "pi/2"] using sin_gt_zero_02 [OF pi_half_gt_zero pi_half_less_two] by (simp add: power2_eq_1_iff) lemma sin_of_real_pi_half [simp]: "sin ((of_real pi/2) :: 'a) = 1" if "SORT_CONSTRAINT('a::{real_field,banach,real_normed_algebra_1})" using sin_pi_half by (metis sin_pi_half eq_numeral_simps(4) nonzero_of_real_divide of_real_1 of_real_numeral sin_of_real) lemma sin_cos_eq: "sin x = cos (of_real pi/2 - x)" for x :: "'a::{real_normed_field,banach}" by (simp add: cos_diff) lemma minus_sin_cos_eq: "- sin x = cos (x + of_real pi/2)" for x :: "'a::{real_normed_field,banach}" by (simp add: cos_add nonzero_of_real_divide) lemma cos_sin_eq: "cos x = sin (of_real pi/2 - x)" for x :: "'a::{real_normed_field,banach}" using sin_cos_eq [of "of_real pi/2 - x"] by simp lemma sin_add: "sin (x + y) = sin x * cos y + cos x * sin y" for x :: "'a::{real_normed_field,banach}" using cos_add [of "of_real pi/2 - x" "-y"] by (simp add: cos_sin_eq) (simp add: sin_cos_eq) lemma sin_diff: "sin (x - y) = sin x * cos y - cos x * sin y" for x :: "'a::{real_normed_field,banach}" using sin_add [of x "- y"] by simp lemma sin_double: "sin(2 * x) = 2 * sin x * cos x" for x :: "'a::{real_normed_field,banach}" using sin_add [where x=x and y=x] by simp lemma cos_of_real_pi [simp]: "cos (of_real pi) = -1" using cos_add [where x = "pi/2" and y = "pi/2"] by (simp add: cos_of_real) lemma sin_of_real_pi [simp]: "sin (of_real pi) = 0" using sin_add [where x = "pi/2" and y = "pi/2"] by (simp add: sin_of_real) lemma cos_pi [simp]: "cos pi = -1" using cos_add [where x = "pi/2" and y = "pi/2"] by simp lemma sin_pi [simp]: "sin pi = 0" using sin_add [where x = "pi/2" and y = "pi/2"] by simp lemma sin_periodic_pi [simp]: "sin (x + pi) = - sin x" by (simp add: sin_add) lemma sin_periodic_pi2 [simp]: "sin (pi + x) = - sin x" by (simp add: sin_add) lemma cos_periodic_pi [simp]: "cos (x + pi) = - cos x" by (simp add: cos_add) lemma cos_periodic_pi2 [simp]: "cos (pi + x) = - cos x" by (simp add: cos_add) lemma sin_periodic [simp]: "sin (x + 2 * pi) = sin x" by (simp add: sin_add sin_double cos_double) lemma cos_periodic [simp]: "cos (x + 2 * pi) = cos x" by (simp add: cos_add sin_double cos_double) lemma cos_npi [simp]: "cos (real n * pi) = (- 1) ^ n" by (induct n) (auto simp: distrib_right) lemma cos_npi2 [simp]: "cos (pi * real n) = (- 1) ^ n" by (metis cos_npi mult.commute) lemma sin_npi [simp]: "sin (real n * pi) = 0" for n :: nat by (induct n) (auto simp: distrib_right) lemma sin_npi2 [simp]: "sin (pi * real n) = 0" for n :: nat by (simp add: mult.commute [of pi]) lemma cos_two_pi [simp]: "cos (2 * pi) = 1" by (simp add: cos_double) lemma sin_two_pi [simp]: "sin (2 * pi) = 0" by (simp add: sin_double) context fixes w :: "'a::{real_normed_field,banach}" begin lemma sin_times_sin: "sin w * sin z = (cos (w - z) - cos (w + z)) / 2" by (simp add: cos_diff cos_add) lemma sin_times_cos: "sin w * cos z = (sin (w + z) + sin (w - z)) / 2" by (simp add: sin_diff sin_add) lemma cos_times_sin: "cos w * sin z = (sin (w + z) - sin (w - z)) / 2" by (simp add: sin_diff sin_add) lemma cos_times_cos: "cos w * cos z = (cos (w - z) + cos (w + z)) / 2" by (simp add: cos_diff cos_add) lemma cos_double_cos: "cos (2 * w) = 2 * cos w ^ 2 - 1" by (simp add: cos_double sin_squared_eq) lemma cos_double_sin: "cos (2 * w) = 1 - 2 * sin w ^ 2" by (simp add: cos_double sin_squared_eq) end lemma sin_plus_sin: "sin w + sin z = 2 * sin ((w + z) / 2) * cos ((w - z) / 2)" for w :: "'a::{real_normed_field,banach}" apply (simp add: mult.assoc sin_times_cos) apply (simp add: field_simps) done lemma sin_diff_sin: "sin w - sin z = 2 * sin ((w - z) / 2) * cos ((w + z) / 2)" for w :: "'a::{real_normed_field,banach}" apply (simp add: mult.assoc sin_times_cos) apply (simp add: field_simps) done lemma cos_plus_cos: "cos w + cos z = 2 * cos ((w + z) / 2) * cos ((w - z) / 2)" for w :: "'a::{real_normed_field,banach,field}" apply (simp add: mult.assoc cos_times_cos) apply (simp add: field_simps) done lemma cos_diff_cos: "cos w - cos z = 2 * sin ((w + z) / 2) * sin ((z - w) / 2)" for w :: "'a::{real_normed_field,banach,field}" apply (simp add: mult.assoc sin_times_sin) apply (simp add: field_simps) done lemma sin_pi_minus [simp]: "sin (pi - x) = sin x" by (metis sin_minus sin_periodic_pi minus_minus uminus_add_conv_diff) lemma cos_pi_minus [simp]: "cos (pi - x) = - (cos x)" by (metis cos_minus cos_periodic_pi uminus_add_conv_diff) lemma sin_minus_pi [simp]: "sin (x - pi) = - (sin x)" by (simp add: sin_diff) lemma cos_minus_pi [simp]: "cos (x - pi) = - (cos x)" by (simp add: cos_diff) lemma sin_2pi_minus [simp]: "sin (2 * pi - x) = - (sin x)" by (metis sin_periodic_pi2 add_diff_eq mult_2 sin_pi_minus) lemma cos_2pi_minus [simp]: "cos (2 * pi - x) = cos x" by (metis (no_types, opaque_lifting) cos_add cos_minus cos_two_pi sin_minus sin_two_pi diff_0_right minus_diff_eq mult_1 mult_zero_left uminus_add_conv_diff) lemma sin_gt_zero2: "0 < x \ x < pi/2 \ 0 < sin x" by (metis sin_gt_zero_02 order_less_trans pi_half_less_two) lemma sin_less_zero: assumes "- pi/2 < x" and "x < 0" shows "sin x < 0" proof - have "0 < sin (- x)" using assms by (simp only: sin_gt_zero2) then show ?thesis by simp qed lemma pi_less_4: "pi < 4" using pi_half_less_two by auto lemma cos_gt_zero: "0 < x \ x < pi/2 \ 0 < cos x" by (simp add: cos_sin_eq sin_gt_zero2) lemma cos_gt_zero_pi: "-(pi/2) < x \ x < pi/2 \ 0 < cos x" using cos_gt_zero [of x] cos_gt_zero [of "-x"] by (cases rule: linorder_cases [of x 0]) auto lemma cos_ge_zero: "-(pi/2) \ x \ x \ pi/2 \ 0 \ cos x" by (auto simp: order_le_less cos_gt_zero_pi) (metis cos_pi_half eq_divide_eq eq_numeral_simps(4)) lemma sin_gt_zero: "0 < x \ x < pi \ 0 < sin x" by (simp add: sin_cos_eq cos_gt_zero_pi) lemma sin_lt_zero: "pi < x \ x < 2 * pi \ sin x < 0" using sin_gt_zero [of "x - pi"] by (simp add: sin_diff) lemma pi_ge_two: "2 \ pi" proof (rule ccontr) assume "\ ?thesis" then have "pi < 2" by auto have "\y > pi. y < 2 \ y < 2 * pi" proof (cases "2 < 2 * pi") case True with dense[OF \pi < 2\] show ?thesis by auto next case False have "pi < 2 * pi" by auto from dense[OF this] and False show ?thesis by auto qed then obtain y where "pi < y" and "y < 2" and "y < 2 * pi" by blast then have "0 < sin y" using sin_gt_zero_02 by auto moreover have "sin y < 0" using sin_gt_zero[of "y - pi"] \pi < y\ and \y < 2 * pi\ sin_periodic_pi[of "y - pi"] by auto ultimately show False by auto qed lemma sin_ge_zero: "0 \ x \ x \ pi \ 0 \ sin x" by (auto simp: order_le_less sin_gt_zero) lemma sin_le_zero: "pi \ x \ x < 2 * pi \ sin x \ 0" using sin_ge_zero [of "x - pi"] by (simp add: sin_diff) lemma sin_pi_divide_n_ge_0 [simp]: assumes "n \ 0" shows "0 \ sin (pi/real n)" by (rule sin_ge_zero) (use assms in \simp_all add: field_split_simps\) lemma sin_pi_divide_n_gt_0: assumes "2 \ n" shows "0 < sin (pi/real n)" by (rule sin_gt_zero) (use assms in \simp_all add: field_split_simps\) text\Proof resembles that of \cos_is_zero\ but with \<^term>\pi\ for the upper bound\ lemma cos_total: assumes y: "-1 \ y" "y \ 1" shows "\!x. 0 \ x \ x \ pi \ cos x = y" proof (rule ex_ex1I) show "\x::real. 0 \ x \ x \ pi \ cos x = y" by (rule IVT2) (simp_all add: y) next fix a b :: real assume ab: "0 \ a \ a \ pi \ cos a = y" "0 \ b \ b \ pi \ cos b = y" have cosd: "\x::real. cos differentiable (at x)" unfolding real_differentiable_def by (auto intro: DERIV_cos) show "a = b" proof (cases a b rule: linorder_cases) case less then obtain z where "a < z" "z < b" "(cos has_real_derivative 0) (at z)" using Rolle by (metis cosd continuous_on_cos_real ab) then have "sin z = 0" using DERIV_cos DERIV_unique neg_equal_0_iff_equal by blast then show ?thesis by (metis \a < z\ \z < b\ ab order_less_le_trans less_le sin_gt_zero) next case greater then obtain z where "b < z" "z < a" "(cos has_real_derivative 0) (at z)" using Rolle by (metis cosd continuous_on_cos_real ab) then have "sin z = 0" using DERIV_cos DERIV_unique neg_equal_0_iff_equal by blast then show ?thesis by (metis \b < z\ \z < a\ ab order_less_le_trans less_le sin_gt_zero) qed auto qed lemma sin_total: assumes y: "-1 \ y" "y \ 1" shows "\!x. - (pi/2) \ x \ x \ pi/2 \ sin x = y" proof - from cos_total [OF y] obtain x where x: "0 \ x" "x \ pi" "cos x = y" and uniq: "\x'. 0 \ x' \ x' \ pi \ cos x' = y \ x' = x " by blast show ?thesis unfolding sin_cos_eq proof (rule ex1I [where a="pi/2 - x"]) show "- (pi/2) \ z \ z \ pi/2 \ cos (of_real pi/2 - z) = y \ z = pi/2 - x" for z using uniq [of "pi/2 -z"] by auto qed (use x in auto) qed lemma cos_zero_lemma: assumes "0 \ x" "cos x = 0" shows "\n. odd n \ x = of_nat n * (pi/2)" proof - have xle: "x < (1 + real_of_int \x/pi\) * pi" using floor_correct [of "x/pi"] by (simp add: add.commute divide_less_eq) obtain n where "real n * pi \ x" "x < real (Suc n) * pi" proof show "real (nat \x / pi\) * pi \ x" using assms floor_divide_lower [of pi x] by auto show "x < real (Suc (nat \x / pi\)) * pi" using assms floor_divide_upper [of pi x] by (simp add: xle) qed then have x: "0 \ x - n * pi" "(x - n * pi) \ pi" "cos (x - n * pi) = 0" by (auto simp: algebra_simps cos_diff assms) then have "\!x. 0 \ x \ x \ pi \ cos x = 0" by (auto simp: intro!: cos_total) then obtain \ where \: "0 \ \" "\ \ pi" "cos \ = 0" and uniq: "\\. 0 \ \ \ \ \ pi \ cos \ = 0 \ \ = \" by blast then have "x - real n * pi = \" using x by blast moreover have "pi/2 = \" using pi_half_ge_zero uniq by fastforce ultimately show ?thesis by (rule_tac x = "Suc (2 * n)" in exI) (simp add: algebra_simps) qed lemma sin_zero_lemma: assumes "0 \ x" "sin x = 0" shows "\n::nat. even n \ x = real n * (pi/2)" proof - obtain n where "odd n" and n: "x + pi/2 = of_nat n * (pi/2)" "n > 0" using cos_zero_lemma [of "x + pi/2"] assms by (auto simp add: cos_add) then have "x = real (n - 1) * (pi/2)" by (simp add: algebra_simps of_nat_diff) then show ?thesis by (simp add: \odd n\) qed lemma cos_zero_iff: "cos x = 0 \ ((\n. odd n \ x = real n * (pi/2)) \ (\n. odd n \ x = - (real n * (pi/2))))" (is "?lhs = ?rhs") proof - have *: "cos (real n * pi/2) = 0" if "odd n" for n :: nat proof - from that obtain m where "n = 2 * m + 1" .. then show ?thesis by (simp add: field_simps) (simp add: cos_add add_divide_distrib) qed show ?thesis proof show ?rhs if ?lhs using that cos_zero_lemma [of x] cos_zero_lemma [of "-x"] by force show ?lhs if ?rhs using that by (auto dest: * simp del: eq_divide_eq_numeral1) qed qed lemma sin_zero_iff: "sin x = 0 \ ((\n. even n \ x = real n * (pi/2)) \ (\n. even n \ x = - (real n * (pi/2))))" (is "?lhs = ?rhs") proof show ?rhs if ?lhs using that sin_zero_lemma [of x] sin_zero_lemma [of "-x"] by force show ?lhs if ?rhs using that by (auto elim: evenE) qed lemma sin_zero_pi_iff: fixes x::real assumes "\x\ < pi" shows "sin x = 0 \ x = 0" proof show "x = 0" if "sin x = 0" using that assms by (auto simp: sin_zero_iff) qed auto lemma cos_zero_iff_int: "cos x = 0 \ (\i. odd i \ x = of_int i * (pi/2))" proof - have 1: "\n. odd n \ \i. odd i \ real n = real_of_int i" by (metis even_of_nat_iff of_int_of_nat_eq) have 2: "\n. odd n \ \i. odd i \ - (real n * pi) = real_of_int i * pi" by (metis even_minus even_of_nat_iff mult.commute mult_minus_right of_int_minus of_int_of_nat_eq) have 3: "\odd i; \n. even n \ real_of_int i \ - (real n)\ \ \n. odd n \ real_of_int i = real n" for i by (cases i rule: int_cases2) auto show ?thesis by (force simp: cos_zero_iff intro!: 1 2 3) qed lemma sin_zero_iff_int: "sin x = 0 \ (\i. even i \ x = of_int i * (pi/2))" (is "?lhs = ?rhs") proof safe assume ?lhs then consider (plus) n where "even n" "x = real n * (pi/2)" | (minus) n where "even n" "x = - (real n * (pi/2))" using sin_zero_iff by auto then show "\n. even n \ x = of_int n * (pi/2)" proof cases case plus then show ?rhs by (metis even_of_nat_iff of_int_of_nat_eq) next case minus then show ?thesis by (rule_tac x="- (int n)" in exI) simp qed next fix i :: int assume "even i" then show "sin (of_int i * (pi/2)) = 0" by (cases i rule: int_cases2, simp_all add: sin_zero_iff) qed lemma sin_zero_iff_int2: "sin x = 0 \ (\i::int. x = of_int i * pi)" proof - have "sin x = 0 \ (\i. even i \ x = real_of_int i * (pi/2))" by (auto simp: sin_zero_iff_int) also have "... = (\j. x = real_of_int (2*j) * (pi/2))" using dvd_triv_left by blast also have "... = (\i::int. x = of_int i * pi)" by auto finally show ?thesis . qed lemma cos_zero_iff_int2: fixes x::real shows "cos x = 0 \ (\n::int. x = n * pi + pi/2)" using sin_zero_iff_int2[of "x-pi/2"] unfolding sin_cos_eq by (auto simp add: algebra_simps) lemma sin_npi_int [simp]: "sin (pi * of_int n) = 0" by (simp add: sin_zero_iff_int2) lemma cos_monotone_0_pi: assumes "0 \ y" and "y < x" and "x \ pi" shows "cos x < cos y" proof - have "- (x - y) < 0" using assms by auto from MVT2[OF \y < x\ DERIV_cos] obtain z where "y < z" and "z < x" and cos_diff: "cos x - cos y = (x - y) * - sin z" by auto then have "0 < z" and "z < pi" using assms by auto then have "0 < sin z" using sin_gt_zero by auto then have "cos x - cos y < 0" unfolding cos_diff minus_mult_commute[symmetric] using \- (x - y) < 0\ by (rule mult_pos_neg2) then show ?thesis by auto qed lemma cos_monotone_0_pi_le: assumes "0 \ y" and "y \ x" and "x \ pi" shows "cos x \ cos y" proof (cases "y < x") case True show ?thesis using cos_monotone_0_pi[OF \0 \ y\ True \x \ pi\] by auto next case False then have "y = x" using \y \ x\ by auto then show ?thesis by auto qed lemma cos_monotone_minus_pi_0: assumes "- pi \ y" and "y < x" and "x \ 0" shows "cos y < cos x" proof - have "0 \ - x" and "- x < - y" and "- y \ pi" using assms by auto from cos_monotone_0_pi[OF this] show ?thesis unfolding cos_minus . qed lemma cos_monotone_minus_pi_0': assumes "- pi \ y" and "y \ x" and "x \ 0" shows "cos y \ cos x" proof (cases "y < x") case True show ?thesis using cos_monotone_minus_pi_0[OF \-pi \ y\ True \x \ 0\] by auto next case False then have "y = x" using \y \ x\ by auto then show ?thesis by auto qed lemma sin_monotone_2pi: assumes "- (pi/2) \ y" and "y < x" and "x \ pi/2" shows "sin y < sin x" unfolding sin_cos_eq using assms by (auto intro: cos_monotone_0_pi) lemma sin_monotone_2pi_le: assumes "- (pi/2) \ y" and "y \ x" and "x \ pi/2" shows "sin y \ sin x" by (metis assms le_less sin_monotone_2pi) lemma sin_x_le_x: fixes x :: real assumes "x \ 0" shows "sin x \ x" proof - let ?f = "\x. x - sin x" have "\u. \0 \ u; u \ x\ \ \y. (?f has_real_derivative 1 - cos u) (at u)" by (auto intro!: derivative_eq_intros simp: field_simps) then have "?f x \ ?f 0" by (metis cos_le_one diff_ge_0_iff_ge DERIV_nonneg_imp_nondecreasing [OF assms]) then show "sin x \ x" by simp qed lemma sin_x_ge_neg_x: fixes x :: real assumes x: "x \ 0" shows "sin x \ - x" proof - let ?f = "\x. x + sin x" have \
: "\u. \0 \ u; u \ x\ \ \y. (?f has_real_derivative 1 + cos u) (at u)" by (auto intro!: derivative_eq_intros simp: field_simps) have "?f x \ ?f 0" by (rule DERIV_nonneg_imp_nondecreasing [OF assms]) (use \
real_0_le_add_iff in force) then show "sin x \ -x" by simp qed lemma abs_sin_x_le_abs_x: "\sin x\ \ \x\" for x :: real using sin_x_ge_neg_x [of x] sin_x_le_x [of x] sin_x_ge_neg_x [of "-x"] sin_x_le_x [of "-x"] by (auto simp: abs_real_def) subsection \More Corollaries about Sine and Cosine\ lemma sin_cos_npi [simp]: "sin (real (Suc (2 * n)) * pi/2) = (-1) ^ n" proof - have "sin ((real n + 1/2) * pi) = cos (real n * pi)" by (auto simp: algebra_simps sin_add) then show ?thesis by (simp add: distrib_right add_divide_distrib add.commute mult.commute [of pi]) qed lemma cos_2npi [simp]: "cos (2 * real n * pi) = 1" for n :: nat by (cases "even n") (simp_all add: cos_double mult.assoc) lemma cos_3over2_pi [simp]: "cos (3/2*pi) = 0" proof - have "cos (3/2*pi) = cos (pi + pi/2)" by simp also have "... = 0" by (subst cos_add, simp) finally show ?thesis . qed lemma sin_2npi [simp]: "sin (2 * real n * pi) = 0" for n :: nat by (auto simp: mult.assoc sin_double) lemma sin_3over2_pi [simp]: "sin (3/2*pi) = - 1" proof - have "sin (3/2*pi) = sin (pi + pi/2)" by simp also have "... = -1" by (subst sin_add, simp) finally show ?thesis . qed lemma cos_pi_eq_zero [simp]: "cos (pi * real (Suc (2 * m)) / 2) = 0" by (simp only: cos_add sin_add of_nat_Suc distrib_right distrib_left add_divide_distrib, auto) lemma DERIV_cos_add [simp]: "DERIV (\x. cos (x + k)) xa :> - sin (xa + k)" by (auto intro!: derivative_eq_intros) lemma sin_zero_norm_cos_one: fixes x :: "'a::{real_normed_field,banach}" assumes "sin x = 0" shows "norm (cos x) = 1" using sin_cos_squared_add [of x, unfolded assms] by (simp add: square_norm_one) lemma sin_zero_abs_cos_one: "sin x = 0 \ \cos x\ = (1::real)" using sin_zero_norm_cos_one by fastforce lemma cos_one_sin_zero: fixes x :: "'a::{real_normed_field,banach}" assumes "cos x = 1" shows "sin x = 0" using sin_cos_squared_add [of x, unfolded assms] by simp lemma sin_times_pi_eq_0: "sin (x * pi) = 0 \ x \ \" by (simp add: sin_zero_iff_int2) (metis Ints_cases Ints_of_int) lemma cos_one_2pi: "cos x = 1 \ (\n::nat. x = n * 2 * pi) \ (\n::nat. x = - (n * 2 * pi))" (is "?lhs = ?rhs") proof assume ?lhs then have "sin x = 0" by (simp add: cos_one_sin_zero) then show ?rhs proof (simp only: sin_zero_iff, elim exE disjE conjE) fix n :: nat assume n: "even n" "x = real n * (pi/2)" then obtain m where m: "n = 2 * m" using dvdE by blast then have me: "even m" using \?lhs\ n by (auto simp: field_simps) (metis one_neq_neg_one power_minus_odd power_one) show ?rhs using m me n by (auto simp: field_simps elim!: evenE) next fix n :: nat assume n: "even n" "x = - (real n * (pi/2))" then obtain m where m: "n = 2 * m" using dvdE by blast then have me: "even m" using \?lhs\ n by (auto simp: field_simps) (metis one_neq_neg_one power_minus_odd power_one) show ?rhs using m me n by (auto simp: field_simps elim!: evenE) qed next assume ?rhs then show "cos x = 1" by (metis cos_2npi cos_minus mult.assoc mult.left_commute) qed lemma cos_one_2pi_int: "cos x = 1 \ (\n::int. x = n * 2 * pi)" (is "?lhs = ?rhs") proof assume "cos x = 1" then show ?rhs by (metis cos_one_2pi mult.commute mult_minus_right of_int_minus of_int_of_nat_eq) next assume ?rhs then show "cos x = 1" by (clarsimp simp add: cos_one_2pi) (metis mult_minus_right of_int_of_nat) qed lemma cos_npi_int [simp]: fixes n::int shows "cos (pi * of_int n) = (if even n then 1 else -1)" by (auto simp: algebra_simps cos_one_2pi_int elim!: oddE evenE) lemma sin_cos_sqrt: "0 \ sin x \ sin x = sqrt (1 - (cos(x) ^ 2))" using sin_squared_eq real_sqrt_unique by fastforce lemma sin_eq_0_pi: "- pi < x \ x < pi \ sin x = 0 \ x = 0" by (metis sin_gt_zero sin_minus minus_less_iff neg_0_less_iff_less not_less_iff_gr_or_eq) lemma cos_treble_cos: "cos (3 * x) = 4 * cos x ^ 3 - 3 * cos x" for x :: "'a::{real_normed_field,banach}" proof - have *: "(sin x * (sin x * 3)) = 3 - (cos x * (cos x * 3))" by (simp add: mult.assoc [symmetric] sin_squared_eq [unfolded power2_eq_square]) have "cos(3 * x) = cos(2*x + x)" by simp also have "\ = 4 * cos x ^ 3 - 3 * cos x" unfolding cos_add cos_double sin_double by (simp add: * field_simps power2_eq_square power3_eq_cube) finally show ?thesis . qed lemma cos_45: "cos (pi/4) = sqrt 2 / 2" proof - let ?c = "cos (pi/4)" let ?s = "sin (pi/4)" have nonneg: "0 \ ?c" by (simp add: cos_ge_zero) have "0 = cos (pi/4 + pi/4)" by simp also have "cos (pi/4 + pi/4) = ?c\<^sup>2 - ?s\<^sup>2" by (simp only: cos_add power2_eq_square) also have "\ = 2 * ?c\<^sup>2 - 1" by (simp add: sin_squared_eq) finally have "?c\<^sup>2 = (sqrt 2 / 2)\<^sup>2" by (simp add: power_divide) then show ?thesis using nonneg by (rule power2_eq_imp_eq) simp qed lemma cos_30: "cos (pi/6) = sqrt 3/2" proof - let ?c = "cos (pi/6)" let ?s = "sin (pi/6)" have pos_c: "0 < ?c" by (rule cos_gt_zero) simp_all have "0 = cos (pi/6 + pi/6 + pi/6)" by simp also have "\ = (?c * ?c - ?s * ?s) * ?c - (?s * ?c + ?c * ?s) * ?s" by (simp only: cos_add sin_add) also have "\ = ?c * (?c\<^sup>2 - 3 * ?s\<^sup>2)" by (simp add: algebra_simps power2_eq_square) finally have "?c\<^sup>2 = (sqrt 3/2)\<^sup>2" using pos_c by (simp add: sin_squared_eq power_divide) then show ?thesis using pos_c [THEN order_less_imp_le] by (rule power2_eq_imp_eq) simp qed lemma sin_45: "sin (pi/4) = sqrt 2 / 2" by (simp add: sin_cos_eq cos_45) lemma sin_60: "sin (pi/3) = sqrt 3/2" by (simp add: sin_cos_eq cos_30) lemma cos_60: "cos (pi/3) = 1/2" proof - have "0 \ cos (pi/3)" by (rule cos_ge_zero) (use pi_half_ge_zero in \linarith+\) then show ?thesis by (simp add: cos_squared_eq sin_60 power_divide power2_eq_imp_eq) qed lemma sin_30: "sin (pi/6) = 1/2" by (simp add: sin_cos_eq cos_60) lemma cos_120: "cos (2 * pi/3) = -1/2" and sin_120: "sin (2 * pi/3) = sqrt 3 / 2" using sin_double[of "pi/3"] cos_double[of "pi/3"] by (simp_all add: power2_eq_square sin_60 cos_60) lemma cos_120': "cos (pi * 2 / 3) = -1/2" using cos_120 by (subst mult.commute) lemma sin_120': "sin (pi * 2 / 3) = sqrt 3 / 2" using sin_120 by (subst mult.commute) lemma cos_integer_2pi: "n \ \ \ cos(2 * pi * n) = 1" by (metis Ints_cases cos_one_2pi_int mult.assoc mult.commute) lemma sin_integer_2pi: "n \ \ \ sin(2 * pi * n) = 0" by (metis sin_two_pi Ints_mult mult.assoc mult.commute sin_times_pi_eq_0) lemma cos_int_2pin [simp]: "cos ((2 * pi) * of_int n) = 1" by (simp add: cos_one_2pi_int) lemma sin_int_2pin [simp]: "sin ((2 * pi) * of_int n) = 0" by (metis Ints_of_int sin_integer_2pi) lemma sin_cos_eq_iff: "sin y = sin x \ cos y = cos x \ (\n::int. y = x + 2 * pi * n)" (is "?L=?R") proof assume ?L then have "cos (y-x) = 1" using cos_add [of y "-x"] by simp then show ?R by (metis cos_one_2pi_int add.commute diff_add_cancel mult.assoc mult.commute) next assume ?R then show ?L by (auto simp: sin_add cos_add) qed lemma sincos_principal_value: "\y. (- pi < y \ y \ pi) \ (sin y = sin x \ cos y = cos x)" proof - define y where "y \ pi - (2 * pi) * frac ((pi - x) / (2 * pi))" have "-pi < y"" y \ pi" by (auto simp: field_simps frac_lt_1 y_def) moreover have "sin y = sin x" "cos y = cos x" by (simp_all add: y_def frac_def divide_simps sin_add cos_add mult_of_int_commute) ultimately show ?thesis by metis qed subsection \Tangent\ definition tan :: "'a \ 'a::{real_normed_field,banach}" where "tan = (\x. sin x / cos x)" lemma tan_of_real: "of_real (tan x) = (tan (of_real x) :: 'a::{real_normed_field,banach})" by (simp add: tan_def sin_of_real cos_of_real) lemma tan_in_Reals [simp]: "z \ \ \ tan z \ \" for z :: "'a::{real_normed_field,banach}" by (simp add: tan_def) lemma tan_zero [simp]: "tan 0 = 0" by (simp add: tan_def) lemma tan_pi [simp]: "tan pi = 0" by (simp add: tan_def) lemma tan_npi [simp]: "tan (real n * pi) = 0" for n :: nat by (simp add: tan_def) lemma tan_pi_half [simp]: "tan (pi / 2) = 0" by (simp add: tan_def) lemma tan_minus [simp]: "tan (- x) = - tan x" by (simp add: tan_def) lemma tan_periodic [simp]: "tan (x + 2 * pi) = tan x" by (simp add: tan_def) lemma lemma_tan_add1: "cos x \ 0 \ cos y \ 0 \ 1 - tan x * tan y = cos (x + y)/(cos x * cos y)" by (simp add: tan_def cos_add field_simps) lemma add_tan_eq: "cos x \ 0 \ cos y \ 0 \ tan x + tan y = sin(x + y)/(cos x * cos y)" for x :: "'a::{real_normed_field,banach}" by (simp add: tan_def sin_add field_simps) lemma tan_eq_0_cos_sin: "tan x = 0 \ cos x = 0 \ sin x = 0" by (auto simp: tan_def) text \Note: half of these zeros would normally be regarded as undefined cases.\ lemma tan_eq_0_Ex: assumes "tan x = 0" obtains k::int where "x = (k/2) * pi" using assms by (metis cos_zero_iff_int mult.commute sin_zero_iff_int tan_eq_0_cos_sin times_divide_eq_left) lemma tan_add: "cos x \ 0 \ cos y \ 0 \ cos (x + y) \ 0 \ tan (x + y) = (tan x + tan y)/(1 - tan x * tan y)" for x :: "'a::{real_normed_field,banach}" by (simp add: add_tan_eq lemma_tan_add1 field_simps) (simp add: tan_def) lemma tan_double: "cos x \ 0 \ cos (2 * x) \ 0 \ tan (2 * x) = (2 * tan x) / (1 - (tan x)\<^sup>2)" for x :: "'a::{real_normed_field,banach}" using tan_add [of x x] by (simp add: power2_eq_square) lemma tan_gt_zero: "0 < x \ x < pi/2 \ 0 < tan x" by (simp add: tan_def zero_less_divide_iff sin_gt_zero2 cos_gt_zero_pi) lemma tan_less_zero: assumes "- pi/2 < x" and "x < 0" shows "tan x < 0" proof - have "0 < tan (- x)" using assms by (simp only: tan_gt_zero) then show ?thesis by simp qed lemma tan_half: "tan x = sin (2 * x) / (cos (2 * x) + 1)" for x :: "'a::{real_normed_field,banach,field}" unfolding tan_def sin_double cos_double sin_squared_eq by (simp add: power2_eq_square) lemma tan_30: "tan (pi/6) = 1 / sqrt 3" unfolding tan_def by (simp add: sin_30 cos_30) lemma tan_45: "tan (pi/4) = 1" unfolding tan_def by (simp add: sin_45 cos_45) lemma tan_60: "tan (pi/3) = sqrt 3" unfolding tan_def by (simp add: sin_60 cos_60) lemma DERIV_tan [simp]: "cos x \ 0 \ DERIV tan x :> inverse ((cos x)\<^sup>2)" for x :: "'a::{real_normed_field,banach}" unfolding tan_def by (auto intro!: derivative_eq_intros, simp add: divide_inverse power2_eq_square) declare DERIV_tan[THEN DERIV_chain2, derivative_intros] and DERIV_tan[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemmas has_derivative_tan[derivative_intros] = DERIV_tan[THEN DERIV_compose_FDERIV] lemma isCont_tan: "cos x \ 0 \ isCont tan x" for x :: "'a::{real_normed_field,banach}" by (rule DERIV_tan [THEN DERIV_isCont]) lemma isCont_tan' [simp,continuous_intros]: fixes a :: "'a::{real_normed_field,banach}" and f :: "'a \ 'a" shows "isCont f a \ cos (f a) \ 0 \ isCont (\x. tan (f x)) a" by (rule isCont_o2 [OF _ isCont_tan]) lemma tendsto_tan [tendsto_intros]: fixes f :: "'a \ 'a::{real_normed_field,banach}" shows "(f \ a) F \ cos a \ 0 \ ((\x. tan (f x)) \ tan a) F" by (rule isCont_tendsto_compose [OF isCont_tan]) lemma continuous_tan: fixes f :: "'a \ 'a::{real_normed_field,banach}" shows "continuous F f \ cos (f (Lim F (\x. x))) \ 0 \ continuous F (\x. tan (f x))" unfolding continuous_def by (rule tendsto_tan) lemma continuous_on_tan [continuous_intros]: fixes f :: "'a \ 'a::{real_normed_field,banach}" shows "continuous_on s f \ (\x\s. cos (f x) \ 0) \ continuous_on s (\x. tan (f x))" unfolding continuous_on_def by (auto intro: tendsto_tan) lemma continuous_within_tan [continuous_intros]: fixes f :: "'a \ 'a::{real_normed_field,banach}" shows "continuous (at x within s) f \ cos (f x) \ 0 \ continuous (at x within s) (\x. tan (f x))" unfolding continuous_within by (rule tendsto_tan) lemma LIM_cos_div_sin: "(\x. cos(x)/sin(x)) \pi/2\ 0" by (rule tendsto_cong_limit, (rule tendsto_intros)+, simp_all) lemma lemma_tan_total: assumes "0 < y" shows "\x. 0 < x \ x < pi/2 \ y < tan x" proof - obtain s where "0 < s" and s: "\x. \x \ pi/2; norm (x - pi/2) < s\ \ norm (cos x / sin x - 0) < inverse y" using LIM_D [OF LIM_cos_div_sin, of "inverse y"] that assms by force obtain e where e: "0 < e" "e < s" "e < pi/2" using \0 < s\ field_lbound_gt_zero pi_half_gt_zero by blast show ?thesis proof (intro exI conjI) have "0 < sin e" "0 < cos e" using e by (auto intro: cos_gt_zero sin_gt_zero2 simp: mult.commute) then show "y < tan (pi/2 - e)" using s [of "pi/2 - e"] e assms by (simp add: tan_def sin_diff cos_diff) (simp add: field_simps split: if_split_asm) qed (use e in auto) qed lemma tan_total_pos: assumes "0 \ y" shows "\x. 0 \ x \ x < pi/2 \ tan x = y" proof (cases "y = 0") case True then show ?thesis using pi_half_gt_zero tan_zero by blast next case False with assms have "y > 0" by linarith obtain x where x: "0 < x" "x < pi/2" "y < tan x" using lemma_tan_total \0 < y\ by blast have "\u\0. u \ x \ tan u = y" proof (intro IVT allI impI) show "isCont tan u" if "0 \ u \ u \ x" for u proof - have "cos u \ 0" using antisym_conv2 cos_gt_zero that x(2) by fastforce with assms show ?thesis by (auto intro!: DERIV_tan [THEN DERIV_isCont]) qed qed (use assms x in auto) then show ?thesis using x(2) by auto qed lemma lemma_tan_total1: "\x. -(pi/2) < x \ x < (pi/2) \ tan x = y" proof (cases "0::real" y rule: le_cases) case le then show ?thesis by (meson less_le_trans minus_pi_half_less_zero tan_total_pos) next case ge with tan_total_pos [of "-y"] obtain x where "0 \ x" "x < pi/2" "tan x = - y" by force then show ?thesis by (rule_tac x="-x" in exI) auto qed proposition tan_total: "\! x. -(pi/2) < x \ x < (pi/2) \ tan x = y" proof - have "u = v" if u: "- (pi/2) < u" "u < pi/2" and v: "- (pi/2) < v" "v < pi/2" and eq: "tan u = tan v" for u v proof (cases u v rule: linorder_cases) case less have "\x. u \ x \ x \ v \ isCont tan x" by (metis cos_gt_zero_pi isCont_tan le_less_trans less_irrefl less_le_trans u(1) v(2)) then have "continuous_on {u..v} tan" by (simp add: continuous_at_imp_continuous_on) moreover have "\x. u < x \ x < v \ tan differentiable (at x)" by (metis DERIV_tan cos_gt_zero_pi real_differentiable_def less_numeral_extra(3) order.strict_trans u(1) v(2)) ultimately obtain z where "u < z" "z < v" "DERIV tan z :> 0" by (metis less Rolle eq) moreover have "cos z \ 0" by (metis (no_types) \u < z\ \z < v\ cos_gt_zero_pi less_le_trans linorder_not_less not_less_iff_gr_or_eq u(1) v(2)) ultimately show ?thesis using DERIV_unique [OF _ DERIV_tan] by fastforce next case greater have "\x. v \ x \ x \ u \ isCont tan x" by (metis cos_gt_zero_pi isCont_tan le_less_trans less_irrefl less_le_trans u(2) v(1)) then have "continuous_on {v..u} tan" by (simp add: continuous_at_imp_continuous_on) moreover have "\x. v < x \ x < u \ tan differentiable (at x)" by (metis DERIV_tan cos_gt_zero_pi real_differentiable_def less_numeral_extra(3) order.strict_trans u(2) v(1)) ultimately obtain z where "v < z" "z < u" "DERIV tan z :> 0" by (metis greater Rolle eq) moreover have "cos z \ 0" by (metis \v < z\ \z < u\ cos_gt_zero_pi less_eq_real_def less_le_trans order_less_irrefl u(2) v(1)) ultimately show ?thesis using DERIV_unique [OF _ DERIV_tan] by fastforce qed auto then have "\!x. - (pi/2) < x \ x < pi/2 \ tan x = y" if x: "- (pi/2) < x" "x < pi/2" "tan x = y" for x using that by auto then show ?thesis using lemma_tan_total1 [where y = y] by auto qed lemma tan_monotone: assumes "- (pi/2) < y" and "y < x" and "x < pi/2" shows "tan y < tan x" proof - have "DERIV tan x' :> inverse ((cos x')\<^sup>2)" if "y \ x'" "x' \ x" for x' proof - have "-(pi/2) < x'" and "x' < pi/2" using that assms by auto with cos_gt_zero_pi have "cos x' \ 0" by force then show "DERIV tan x' :> inverse ((cos x')\<^sup>2)" by (rule DERIV_tan) qed from MVT2[OF \y < x\ this] obtain z where "y < z" and "z < x" and tan_diff: "tan x - tan y = (x - y) * inverse ((cos z)\<^sup>2)" by auto then have "- (pi/2) < z" and "z < pi/2" using assms by auto then have "0 < cos z" using cos_gt_zero_pi by auto then have inv_pos: "0 < inverse ((cos z)\<^sup>2)" by auto have "0 < x - y" using \y < x\ by auto with inv_pos have "0 < tan x - tan y" unfolding tan_diff by auto then show ?thesis by auto qed lemma tan_monotone': assumes "- (pi/2) < y" and "y < pi/2" and "- (pi/2) < x" and "x < pi/2" shows "y < x \ tan y < tan x" proof assume "y < x" then show "tan y < tan x" using tan_monotone and \- (pi/2) < y\ and \x < pi/2\ by auto next assume "tan y < tan x" show "y < x" proof (rule ccontr) assume "\ ?thesis" then have "x \ y" by auto then have "tan x \ tan y" proof (cases "x = y") case True then show ?thesis by auto next case False then have "x < y" using \x \ y\ by auto from tan_monotone[OF \- (pi/2) < x\ this \y < pi/2\] show ?thesis by auto qed then show False using \tan y < tan x\ by auto qed qed lemma tan_inverse: "1 / (tan y) = tan (pi/2 - y)" unfolding tan_def sin_cos_eq[of y] cos_sin_eq[of y] by auto lemma tan_periodic_pi[simp]: "tan (x + pi) = tan x" by (simp add: tan_def) lemma tan_periodic_nat[simp]: "tan (x + real n * pi) = tan x" proof (induct n arbitrary: x) case 0 then show ?case by simp next case (Suc n) have split_pi_off: "x + real (Suc n) * pi = (x + real n * pi) + pi" unfolding Suc_eq_plus1 of_nat_add distrib_right by auto show ?case unfolding split_pi_off using Suc by auto qed lemma tan_periodic_int[simp]: "tan (x + of_int i * pi) = tan x" proof (cases "0 \ i") case False then have i_nat: "of_int i = - of_int (nat (- i))" by auto then show ?thesis by (smt (verit, best) mult_minus_left of_int_of_nat_eq tan_periodic_nat) qed (use zero_le_imp_eq_int in fastforce) lemma tan_periodic_n[simp]: "tan (x + numeral n * pi) = tan x" using tan_periodic_int[of _ "numeral n" ] by simp lemma tan_minus_45 [simp]: "tan (-(pi/4)) = -1" unfolding tan_def by (simp add: sin_45 cos_45) lemma tan_diff: "cos x \ 0 \ cos y \ 0 \ cos (x - y) \ 0 \ tan (x - y) = (tan x - tan y)/(1 + tan x * tan y)" for x :: "'a::{real_normed_field,banach}" using tan_add [of x "-y"] by simp lemma tan_pos_pi2_le: "0 \ x \ x < pi/2 \ 0 \ tan x" using less_eq_real_def tan_gt_zero by auto lemma cos_tan: "\x\ < pi/2 \ cos x = 1 / sqrt (1 + tan x ^ 2)" using cos_gt_zero_pi [of x] by (simp add: field_split_simps tan_def real_sqrt_divide abs_if split: if_split_asm) lemma cos_tan_half: "cos x \0 \ cos (2*x) = (1 - (tan x)^2) / (1 + (tan x)^2)" unfolding cos_double tan_def by (auto simp add:field_simps ) lemma sin_tan: "\x\ < pi/2 \ sin x = tan x / sqrt (1 + tan x ^ 2)" using cos_gt_zero [of "x"] cos_gt_zero [of "-x"] by (force simp: field_split_simps tan_def real_sqrt_divide abs_if split: if_split_asm) lemma sin_tan_half: "sin (2*x) = 2 * tan x / (1 + (tan x)^2)" unfolding sin_double tan_def by (cases "cos x=0") (auto simp add:field_simps power2_eq_square) lemma tan_mono_le: "-(pi/2) < x \ x \ y \ y < pi/2 \ tan x \ tan y" using less_eq_real_def tan_monotone by auto lemma tan_mono_lt_eq: "-(pi/2) < x \ x < pi/2 \ -(pi/2) < y \ y < pi/2 \ tan x < tan y \ x < y" using tan_monotone' by blast lemma tan_mono_le_eq: "-(pi/2) < x \ x < pi/2 \ -(pi/2) < y \ y < pi/2 \ tan x \ tan y \ x \ y" by (meson tan_mono_le not_le tan_monotone) lemma tan_bound_pi2: "\x\ < pi/4 \ \tan x\ < 1" using tan_45 tan_monotone [of x "pi/4"] tan_monotone [of "-x" "pi/4"] by (auto simp: abs_if split: if_split_asm) lemma tan_cot: "tan(pi/2 - x) = inverse(tan x)" by (simp add: tan_def sin_diff cos_diff) subsection \Cotangent\ definition cot :: "'a \ 'a::{real_normed_field,banach}" where "cot = (\x. cos x / sin x)" lemma cot_of_real: "of_real (cot x) = (cot (of_real x) :: 'a::{real_normed_field,banach})" by (simp add: cot_def sin_of_real cos_of_real) lemma cot_in_Reals [simp]: "z \ \ \ cot z \ \" for z :: "'a::{real_normed_field,banach}" by (simp add: cot_def) lemma cot_zero [simp]: "cot 0 = 0" by (simp add: cot_def) lemma cot_pi [simp]: "cot pi = 0" by (simp add: cot_def) lemma cot_npi [simp]: "cot (real n * pi) = 0" for n :: nat by (simp add: cot_def) lemma cot_minus [simp]: "cot (- x) = - cot x" by (simp add: cot_def) lemma cot_periodic [simp]: "cot (x + 2 * pi) = cot x" by (simp add: cot_def) lemma cot_altdef: "cot x = inverse (tan x)" by (simp add: cot_def tan_def) lemma tan_altdef: "tan x = inverse (cot x)" by (simp add: cot_def tan_def) lemma tan_cot': "tan (pi/2 - x) = cot x" by (simp add: tan_cot cot_altdef) lemma cot_gt_zero: "0 < x \ x < pi/2 \ 0 < cot x" by (simp add: cot_def zero_less_divide_iff sin_gt_zero2 cos_gt_zero_pi) lemma cot_less_zero: assumes lb: "- pi/2 < x" and "x < 0" shows "cot x < 0" by (smt (verit) assms cot_gt_zero cot_minus divide_minus_left) lemma DERIV_cot [simp]: "sin x \ 0 \ DERIV cot x :> -inverse ((sin x)\<^sup>2)" for x :: "'a::{real_normed_field,banach}" unfolding cot_def using cos_squared_eq[of x] by (auto intro!: derivative_eq_intros) (simp add: divide_inverse power2_eq_square) lemma isCont_cot: "sin x \ 0 \ isCont cot x" for x :: "'a::{real_normed_field,banach}" by (rule DERIV_cot [THEN DERIV_isCont]) lemma isCont_cot' [simp,continuous_intros]: "isCont f a \ sin (f a) \ 0 \ isCont (\x. cot (f x)) a" for a :: "'a::{real_normed_field,banach}" and f :: "'a \ 'a" by (rule isCont_o2 [OF _ isCont_cot]) lemma tendsto_cot [tendsto_intros]: "(f \ a) F \ sin a \ 0 \ ((\x. cot (f x)) \ cot a) F" for f :: "'a \ 'a::{real_normed_field,banach}" by (rule isCont_tendsto_compose [OF isCont_cot]) lemma continuous_cot: "continuous F f \ sin (f (Lim F (\x. x))) \ 0 \ continuous F (\x. cot (f x))" for f :: "'a \ 'a::{real_normed_field,banach}" unfolding continuous_def by (rule tendsto_cot) lemma continuous_on_cot [continuous_intros]: fixes f :: "'a \ 'a::{real_normed_field,banach}" shows "continuous_on s f \ (\x\s. sin (f x) \ 0) \ continuous_on s (\x. cot (f x))" unfolding continuous_on_def by (auto intro: tendsto_cot) lemma continuous_within_cot [continuous_intros]: fixes f :: "'a \ 'a::{real_normed_field,banach}" shows "continuous (at x within s) f \ sin (f x) \ 0 \ continuous (at x within s) (\x. cot (f x))" unfolding continuous_within by (rule tendsto_cot) subsection \Inverse Trigonometric Functions\ definition arcsin :: "real \ real" where "arcsin y = (THE x. -(pi/2) \ x \ x \ pi/2 \ sin x = y)" definition arccos :: "real \ real" where "arccos y = (THE x. 0 \ x \ x \ pi \ cos x = y)" definition arctan :: "real \ real" where "arctan y = (THE x. -(pi/2) < x \ x < pi/2 \ tan x = y)" lemma arcsin: "- 1 \ y \ y \ 1 \ - (pi/2) \ arcsin y \ arcsin y \ pi/2 \ sin (arcsin y) = y" unfolding arcsin_def by (rule theI' [OF sin_total]) lemma arcsin_pi: "- 1 \ y \ y \ 1 \ - (pi/2) \ arcsin y \ arcsin y \ pi \ sin (arcsin y) = y" by (drule (1) arcsin) (force intro: order_trans) lemma sin_arcsin [simp]: "- 1 \ y \ y \ 1 \ sin (arcsin y) = y" by (blast dest: arcsin) lemma arcsin_bounded: "- 1 \ y \ y \ 1 \ - (pi/2) \ arcsin y \ arcsin y \ pi/2" by (blast dest: arcsin) lemma arcsin_lbound: "- 1 \ y \ y \ 1 \ - (pi/2) \ arcsin y" by (blast dest: arcsin) lemma arcsin_ubound: "- 1 \ y \ y \ 1 \ arcsin y \ pi/2" by (blast dest: arcsin) lemma arcsin_lt_bounded: assumes "- 1 < y" "y < 1" shows "- (pi/2) < arcsin y \ arcsin y < pi/2" proof - have "arcsin y \ pi/2" by (metis arcsin assms not_less not_less_iff_gr_or_eq sin_pi_half) moreover have "arcsin y \ - pi/2" by (metis arcsin assms minus_divide_left not_less not_less_iff_gr_or_eq sin_minus sin_pi_half) ultimately show ?thesis using arcsin_bounded [of y] assms by auto qed lemma arcsin_sin: "- (pi/2) \ x \ x \ pi/2 \ arcsin (sin x) = x" unfolding arcsin_def using the1_equality [OF sin_total] by simp lemma arcsin_unique: assumes "-pi/2 \ x" and "x \ pi/2" and "sin x = y" shows "arcsin y = x" using arcsin_sin[of x] assms by force lemma arcsin_0 [simp]: "arcsin 0 = 0" using arcsin_sin [of 0] by simp lemma arcsin_1 [simp]: "arcsin 1 = pi/2" using arcsin_sin [of "pi/2"] by simp lemma arcsin_minus_1 [simp]: "arcsin (- 1) = - (pi/2)" using arcsin_sin [of "- pi/2"] by simp lemma arcsin_minus: "- 1 \ x \ x \ 1 \ arcsin (- x) = - arcsin x" by (metis (no_types, opaque_lifting) arcsin arcsin_sin minus_minus neg_le_iff_le sin_minus) lemma arcsin_one_half [simp]: "arcsin (1/2) = pi / 6" and arcsin_minus_one_half [simp]: "arcsin (-(1/2)) = -pi / 6" by (intro arcsin_unique; simp add: sin_30 field_simps)+ lemma arcsin_one_over_sqrt_2: "arcsin (1 / sqrt 2) = pi / 4" by (rule arcsin_unique) (auto simp: sin_45 field_simps) lemma arcsin_eq_iff: "\x\ \ 1 \ \y\ \ 1 \ arcsin x = arcsin y \ x = y" by (metis abs_le_iff arcsin minus_le_iff) lemma cos_arcsin_nonzero: "- 1 < x \ x < 1 \ cos (arcsin x) \ 0" using arcsin_lt_bounded cos_gt_zero_pi by force lemma arccos: "- 1 \ y \ y \ 1 \ 0 \ arccos y \ arccos y \ pi \ cos (arccos y) = y" unfolding arccos_def by (rule theI' [OF cos_total]) lemma cos_arccos [simp]: "- 1 \ y \ y \ 1 \ cos (arccos y) = y" by (blast dest: arccos) lemma arccos_bounded: "- 1 \ y \ y \ 1 \ 0 \ arccos y \ arccos y \ pi" by (blast dest: arccos) lemma arccos_lbound: "- 1 \ y \ y \ 1 \ 0 \ arccos y" by (blast dest: arccos) lemma arccos_ubound: "- 1 \ y \ y \ 1 \ arccos y \ pi" by (blast dest: arccos) lemma arccos_lt_bounded: assumes "- 1 < y" "y < 1" shows "0 < arccos y \ arccos y < pi" proof - have "arccos y \ 0" by (metis (no_types) arccos assms(1) assms(2) cos_zero less_eq_real_def less_irrefl) moreover have "arccos y \ -pi" by (metis arccos assms(1) assms(2) cos_minus cos_pi not_less not_less_iff_gr_or_eq) ultimately show ?thesis using arccos_bounded [of y] assms by (metis arccos cos_pi not_less not_less_iff_gr_or_eq) qed lemma arccos_cos: "0 \ x \ x \ pi \ arccos (cos x) = x" by (auto simp: arccos_def intro!: the1_equality cos_total) lemma arccos_cos2: "x \ 0 \ - pi \ x \ arccos (cos x) = -x" by (auto simp: arccos_def intro!: the1_equality cos_total) lemma arccos_unique: assumes "0 \ x" and "x \ pi" and "cos x = y" shows "arccos y = x" using arccos_cos assms by blast lemma cos_arcsin: assumes "- 1 \ x" "x \ 1" shows "cos (arcsin x) = sqrt (1 - x\<^sup>2)" proof (rule power2_eq_imp_eq) show "(cos (arcsin x))\<^sup>2 = (sqrt (1 - x\<^sup>2))\<^sup>2" by (simp add: square_le_1 assms cos_squared_eq) show "0 \ cos (arcsin x)" using arcsin assms cos_ge_zero by blast show "0 \ sqrt (1 - x\<^sup>2)" by (simp add: square_le_1 assms) qed lemma sin_arccos: assumes "- 1 \ x" "x \ 1" shows "sin (arccos x) = sqrt (1 - x\<^sup>2)" proof (rule power2_eq_imp_eq) show "(sin (arccos x))\<^sup>2 = (sqrt (1 - x\<^sup>2))\<^sup>2" by (simp add: square_le_1 assms sin_squared_eq) show "0 \ sin (arccos x)" by (simp add: arccos_bounded assms sin_ge_zero) show "0 \ sqrt (1 - x\<^sup>2)" by (simp add: square_le_1 assms) qed lemma arccos_0 [simp]: "arccos 0 = pi/2" using arccos_cos pi_half_ge_zero by fastforce lemma arccos_1 [simp]: "arccos 1 = 0" using arccos_cos by force lemma arccos_minus_1 [simp]: "arccos (- 1) = pi" by (metis arccos_cos cos_pi order_refl pi_ge_zero) lemma arccos_minus: "-1 \ x \ x \ 1 \ arccos (- x) = pi - arccos x" by (smt (verit, ccfv_threshold) arccos arccos_cos cos_minus cos_minus_pi) lemma arccos_one_half [simp]: "arccos (1/2) = pi / 3" and arccos_minus_one_half [simp]: "arccos (-(1/2)) = 2 * pi / 3" by (intro arccos_unique; simp add: cos_60 cos_120)+ lemma arccos_one_over_sqrt_2: "arccos (1 / sqrt 2) = pi / 4" by (rule arccos_unique) (auto simp: cos_45 field_simps) corollary arccos_minus_abs: assumes "\x\ \ 1" shows "arccos (- x) = pi - arccos x" using assms by (simp add: arccos_minus) lemma sin_arccos_nonzero: "- 1 < x \ x < 1 \ sin (arccos x) \ 0" using arccos_lt_bounded sin_gt_zero by force lemma arctan: "- (pi/2) < arctan y \ arctan y < pi/2 \ tan (arctan y) = y" unfolding arctan_def by (rule theI' [OF tan_total]) lemma tan_arctan: "tan (arctan y) = y" by (simp add: arctan) lemma arctan_bounded: "- (pi/2) < arctan y \ arctan y < pi/2" by (auto simp only: arctan) lemma arctan_lbound: "- (pi/2) < arctan y" by (simp add: arctan) lemma arctan_ubound: "arctan y < pi/2" by (auto simp only: arctan) lemma arctan_unique: assumes "-(pi/2) < x" and "x < pi/2" and "tan x = y" shows "arctan y = x" using assms arctan [of y] tan_total [of y] by (fast elim: ex1E) lemma arctan_tan: "-(pi/2) < x \ x < pi/2 \ arctan (tan x) = x" by (rule arctan_unique) simp_all lemma arctan_zero_zero [simp]: "arctan 0 = 0" by (rule arctan_unique) simp_all lemma arctan_minus: "arctan (- x) = - arctan x" using arctan [of "x"] by (auto simp: arctan_unique) lemma cos_arctan_not_zero [simp]: "cos (arctan x) \ 0" by (intro less_imp_neq [symmetric] cos_gt_zero_pi arctan_lbound arctan_ubound) lemma tan_eq_arctan_Ex: shows "tan x = y \ (\k::int. x = arctan y + k*pi \ (x = pi/2 + k*pi \ y=0))" proof assume lhs: "tan x = y" obtain k::int where k:"-pi/2 < x-k*pi" "x-k*pi \ pi/2" proof define k where "k \ ceiling (x/pi - 1/2)" show "- pi / 2 < x - real_of_int k * pi" using ceiling_divide_lower [of "pi*2" "(x * 2 - pi)"] by (auto simp: k_def field_simps) show "x-k*pi \ pi/2" using ceiling_divide_upper [of "pi*2" "(x * 2 - pi)"] by (auto simp: k_def field_simps) qed have "x = arctan y + of_int k * pi" when "x \ pi/2 + k*pi" proof - have "tan (x - k * pi) = y" using lhs tan_periodic_int[of _ "-k"] by auto then have "arctan y = x - real_of_int k * pi" by (smt (verit) arctan_tan lhs divide_minus_left k mult_minus_left of_int_minus tan_periodic_int that) then show ?thesis by auto qed then show "\k. x = arctan y + of_int k * pi \ (x = pi/2 + k*pi \ y=0)" using lhs k by force qed (auto simp: arctan) lemma arctan_tan_eq_abs_pi: assumes "cos \ \ 0" obtains k where "arctan (tan \) = \ - of_int k * pi" by (metis add.commute assms cos_zero_iff_int2 eq_diff_eq tan_eq_arctan_Ex) lemma tan_eq: assumes "tan x = tan y" "tan x \ 0" obtains k::int where "x = y + k * pi" proof - obtain k0 where k0: "x = arctan (tan y) + real_of_int k0 * pi" using assms tan_eq_arctan_Ex[of x "tan y"] by auto obtain k1 where k1: "arctan (tan y) = y - of_int k1 * pi" using arctan_tan_eq_abs_pi assms tan_eq_0_cos_sin by auto have "x = y + (k0-k1)*pi" using k0 k1 by (auto simp: algebra_simps) with that show ?thesis by blast qed lemma cos_arctan: "cos (arctan x) = 1 / sqrt (1 + x\<^sup>2)" proof (rule power2_eq_imp_eq) have "0 < 1 + x\<^sup>2" by (simp add: add_pos_nonneg) show "0 \ 1 / sqrt (1 + x\<^sup>2)" by simp show "0 \ cos (arctan x)" by (intro less_imp_le cos_gt_zero_pi arctan_lbound arctan_ubound) have "(cos (arctan x))\<^sup>2 * (1 + (tan (arctan x))\<^sup>2) = 1" unfolding tan_def by (simp add: distrib_left power_divide) then show "(cos (arctan x))\<^sup>2 = (1 / sqrt (1 + x\<^sup>2))\<^sup>2" using \0 < 1 + x\<^sup>2\ by (simp add: arctan power_divide eq_divide_eq) qed lemma sin_arctan: "sin (arctan x) = x / sqrt (1 + x\<^sup>2)" using add_pos_nonneg [OF zero_less_one zero_le_power2 [of x]] using tan_arctan [of x] unfolding tan_def cos_arctan by (simp add: eq_divide_eq) lemma tan_sec: "cos x \ 0 \ 1 + (tan x)\<^sup>2 = (inverse (cos x))\<^sup>2" for x :: "'a::{real_normed_field,banach,field}" by (simp add: add_divide_eq_iff inverse_eq_divide power2_eq_square tan_def) lemma arctan_less_iff: "arctan x < arctan y \ x < y" by (metis tan_monotone' arctan_lbound arctan_ubound tan_arctan) lemma arctan_le_iff: "arctan x \ arctan y \ x \ y" by (simp only: not_less [symmetric] arctan_less_iff) lemma arctan_eq_iff: "arctan x = arctan y \ x = y" by (simp only: eq_iff [where 'a=real] arctan_le_iff) lemma zero_less_arctan_iff [simp]: "0 < arctan x \ 0 < x" using arctan_less_iff [of 0 x] by simp lemma arctan_less_zero_iff [simp]: "arctan x < 0 \ x < 0" using arctan_less_iff [of x 0] by simp lemma zero_le_arctan_iff [simp]: "0 \ arctan x \ 0 \ x" using arctan_le_iff [of 0 x] by simp lemma arctan_le_zero_iff [simp]: "arctan x \ 0 \ x \ 0" using arctan_le_iff [of x 0] by simp lemma arctan_eq_zero_iff [simp]: "arctan x = 0 \ x = 0" using arctan_eq_iff [of x 0] by simp lemma continuous_on_arcsin': "continuous_on {-1 .. 1} arcsin" proof - have "continuous_on (sin ` {- pi/2 .. pi/2}) arcsin" by (rule continuous_on_inv) (auto intro: continuous_intros simp: arcsin_sin) also have "sin ` {- pi/2 .. pi/2} = {-1 .. 1}" proof safe fix x :: real assume "x \ {-1..1}" then show "x \ sin ` {- pi/2..pi/2}" using arcsin_lbound arcsin_ubound by (intro image_eqI[where x="arcsin x"]) auto qed simp finally show ?thesis . qed lemma continuous_on_arcsin [continuous_intros]: "continuous_on s f \ (\x\s. -1 \ f x \ f x \ 1) \ continuous_on s (\x. arcsin (f x))" using continuous_on_compose[of s f, OF _ continuous_on_subset[OF continuous_on_arcsin']] by (auto simp: comp_def subset_eq) lemma isCont_arcsin: "-1 < x \ x < 1 \ isCont arcsin x" using continuous_on_arcsin'[THEN continuous_on_subset, of "{ -1 <..< 1 }"] by (auto simp: continuous_on_eq_continuous_at subset_eq) lemma continuous_on_arccos': "continuous_on {-1 .. 1} arccos" proof - have "continuous_on (cos ` {0 .. pi}) arccos" by (rule continuous_on_inv) (auto intro: continuous_intros simp: arccos_cos) also have "cos ` {0 .. pi} = {-1 .. 1}" proof safe fix x :: real assume "x \ {-1..1}" then show "x \ cos ` {0..pi}" using arccos_lbound arccos_ubound by (intro image_eqI[where x="arccos x"]) auto qed simp finally show ?thesis . qed lemma continuous_on_arccos [continuous_intros]: "continuous_on s f \ (\x\s. -1 \ f x \ f x \ 1) \ continuous_on s (\x. arccos (f x))" using continuous_on_compose[of s f, OF _ continuous_on_subset[OF continuous_on_arccos']] by (auto simp: comp_def subset_eq) lemma isCont_arccos: "-1 < x \ x < 1 \ isCont arccos x" using continuous_on_arccos'[THEN continuous_on_subset, of "{ -1 <..< 1 }"] by (auto simp: continuous_on_eq_continuous_at subset_eq) lemma isCont_arctan: "isCont arctan x" proof - obtain u where u: "- (pi/2) < u" "u < arctan x" by (meson arctan arctan_less_iff linordered_field_no_lb) obtain v where v: "arctan x < v" "v < pi/2" by (meson arctan_less_iff arctan_ubound linordered_field_no_ub) have "isCont arctan (tan (arctan x))" proof (rule isCont_inverse_function2 [of u "arctan x" v]) show "\z. \u \ z; z \ v\ \ arctan (tan z) = z" using arctan_unique u(1) v(2) by auto then show "\z. \u \ z; z \ v\ \ isCont tan z" by (metis arctan cos_gt_zero_pi isCont_tan less_irrefl) qed (use u v in auto) then show ?thesis by (simp add: arctan) qed lemma tendsto_arctan [tendsto_intros]: "(f \ x) F \ ((\x. arctan (f x)) \ arctan x) F" by (rule isCont_tendsto_compose [OF isCont_arctan]) lemma continuous_arctan [continuous_intros]: "continuous F f \ continuous F (\x. arctan (f x))" unfolding continuous_def by (rule tendsto_arctan) lemma continuous_on_arctan [continuous_intros]: "continuous_on s f \ continuous_on s (\x. arctan (f x))" unfolding continuous_on_def by (auto intro: tendsto_arctan) lemma DERIV_arcsin: assumes "- 1 < x" "x < 1" shows "DERIV arcsin x :> inverse (sqrt (1 - x\<^sup>2))" proof (rule DERIV_inverse_function) show "(sin has_real_derivative sqrt (1 - x\<^sup>2)) (at (arcsin x))" by (rule derivative_eq_intros | use assms cos_arcsin in force)+ show "sqrt (1 - x\<^sup>2) \ 0" using abs_square_eq_1 assms by force qed (use assms isCont_arcsin in auto) lemma DERIV_arccos: assumes "- 1 < x" "x < 1" shows "DERIV arccos x :> inverse (- sqrt (1 - x\<^sup>2))" proof (rule DERIV_inverse_function) show "(cos has_real_derivative - sqrt (1 - x\<^sup>2)) (at (arccos x))" by (rule derivative_eq_intros | use assms sin_arccos in force)+ show "- sqrt (1 - x\<^sup>2) \ 0" using abs_square_eq_1 assms by force qed (use assms isCont_arccos in auto) lemma DERIV_arctan: "DERIV arctan x :> inverse (1 + x\<^sup>2)" proof (rule DERIV_inverse_function) have "inverse ((cos (arctan x))\<^sup>2) = 1 + x\<^sup>2" by (metis arctan cos_arctan_not_zero power_inverse tan_sec) then show "(tan has_real_derivative 1 + x\<^sup>2) (at (arctan x))" by (auto intro!: derivative_eq_intros) show "\y. \x - 1 < y; y < x + 1\ \ tan (arctan y) = y" using tan_arctan by blast show "1 + x\<^sup>2 \ 0" by (metis power_one sum_power2_eq_zero_iff zero_neq_one) qed (use isCont_arctan in auto) declare DERIV_arcsin[THEN DERIV_chain2, derivative_intros] DERIV_arcsin[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] DERIV_arccos[THEN DERIV_chain2, derivative_intros] DERIV_arccos[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] DERIV_arctan[THEN DERIV_chain2, derivative_intros] DERIV_arctan[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemmas has_derivative_arctan[derivative_intros] = DERIV_arctan[THEN DERIV_compose_FDERIV] and has_derivative_arccos[derivative_intros] = DERIV_arccos[THEN DERIV_compose_FDERIV] and has_derivative_arcsin[derivative_intros] = DERIV_arcsin[THEN DERIV_compose_FDERIV] lemma filterlim_tan_at_right: "filterlim tan at_bot (at_right (- (pi/2)))" by (rule filterlim_at_bot_at_right[where Q="\x. - pi/2 < x \ x < pi/2" and P="\x. True" and g=arctan]) (auto simp: arctan le_less eventually_at dist_real_def simp del: less_divide_eq_numeral1 intro!: tan_monotone exI[of _ "pi/2"]) lemma filterlim_tan_at_left: "filterlim tan at_top (at_left (pi/2))" by (rule filterlim_at_top_at_left[where Q="\x. - pi/2 < x \ x < pi/2" and P="\x. True" and g=arctan]) (auto simp: arctan le_less eventually_at dist_real_def simp del: less_divide_eq_numeral1 intro!: tan_monotone exI[of _ "pi/2"]) lemma tendsto_arctan_at_top: "(arctan \ (pi/2)) at_top" proof (rule tendstoI) fix e :: real assume "0 < e" define y where "y = pi/2 - min (pi/2) e" then have y: "0 \ y" "y < pi/2" "pi/2 \ e + y" using \0 < e\ by auto show "eventually (\x. dist (arctan x) (pi/2) < e) at_top" proof (intro eventually_at_top_dense[THEN iffD2] exI allI impI) fix x assume "tan y < x" then have "arctan (tan y) < arctan x" by (simp add: arctan_less_iff) with y have "y < arctan x" by (subst (asm) arctan_tan) simp_all with arctan_ubound[of x, arith] y \0 < e\ show "dist (arctan x) (pi/2) < e" by (simp add: dist_real_def) qed qed lemma tendsto_arctan_at_bot: "(arctan \ - (pi/2)) at_bot" unfolding filterlim_at_bot_mirror arctan_minus by (intro tendsto_minus tendsto_arctan_at_top) lemma sin_multiple_reduce: "sin (x * numeral n :: 'a :: {real_normed_field, banach}) = sin x * cos (x * of_nat (pred_numeral n)) + cos x * sin (x * of_nat (pred_numeral n))" proof - have "numeral n = of_nat (pred_numeral n) + (1 :: 'a)" by (metis add.commute numeral_eq_Suc of_nat_Suc of_nat_numeral) also have "sin (x * \) = sin (x * of_nat (pred_numeral n) + x)" unfolding of_nat_Suc by (simp add: ring_distribs) finally show ?thesis by (simp add: sin_add) qed lemma cos_multiple_reduce: "cos (x * numeral n :: 'a :: {real_normed_field, banach}) = cos (x * of_nat (pred_numeral n)) * cos x - sin (x * of_nat (pred_numeral n)) * sin x" proof - have "numeral n = of_nat (pred_numeral n) + (1 :: 'a)" by (metis add.commute numeral_eq_Suc of_nat_Suc of_nat_numeral) also have "cos (x * \) = cos (x * of_nat (pred_numeral n) + x)" unfolding of_nat_Suc by (simp add: ring_distribs) finally show ?thesis by (simp add: cos_add) qed lemma arccos_eq_pi_iff: "x \ {-1..1} \ arccos x = pi \ x = -1" by (metis arccos arccos_minus_1 atLeastAtMost_iff cos_pi) lemma arccos_eq_0_iff: "x \ {-1..1} \ arccos x = 0 \ x = 1" by (metis arccos arccos_1 atLeastAtMost_iff cos_zero) subsection \Prove Totality of the Trigonometric Functions\ lemma cos_arccos_abs: "\y\ \ 1 \ cos (arccos y) = y" by (simp add: abs_le_iff) lemma sin_arccos_abs: "\y\ \ 1 \ sin (arccos y) = sqrt (1 - y\<^sup>2)" by (simp add: sin_arccos abs_le_iff) lemma sin_mono_less_eq: "- (pi/2) \ x \ x \ pi/2 \ - (pi/2) \ y \ y \ pi/2 \ sin x < sin y \ x < y" by (metis not_less_iff_gr_or_eq sin_monotone_2pi) lemma sin_mono_le_eq: "- (pi/2) \ x \ x \ pi/2 \ - (pi/2) \ y \ y \ pi/2 \ sin x \ sin y \ x \ y" by (meson leD le_less_linear sin_monotone_2pi sin_monotone_2pi_le) lemma sin_inj_pi: "- (pi/2) \ x \ x \ pi/2 \ - (pi/2) \ y \ y \ pi/2 \ sin x = sin y \ x = y" by (metis arcsin_sin) lemma arcsin_le_iff: assumes "x \ -1" "x \ 1" "y \ -pi/2" "y \ pi/2" shows "arcsin x \ y \ x \ sin y" proof - have "arcsin x \ y \ sin (arcsin x) \ sin y" using arcsin_bounded[of x] assms by (subst sin_mono_le_eq) auto also from assms have "sin (arcsin x) = x" by simp finally show ?thesis . qed lemma le_arcsin_iff: assumes "x \ -1" "x \ 1" "y \ -pi/2" "y \ pi/2" shows "arcsin x \ y \ x \ sin y" proof - have "arcsin x \ y \ sin (arcsin x) \ sin y" using arcsin_bounded[of x] assms by (subst sin_mono_le_eq) auto also from assms have "sin (arcsin x) = x" by simp finally show ?thesis . qed lemma cos_mono_less_eq: "0 \ x \ x \ pi \ 0 \ y \ y \ pi \ cos x < cos y \ y < x" by (meson cos_monotone_0_pi cos_monotone_0_pi_le leD le_less_linear) lemma cos_mono_le_eq: "0 \ x \ x \ pi \ 0 \ y \ y \ pi \ cos x \ cos y \ y \ x" by (metis arccos_cos cos_monotone_0_pi_le eq_iff linear) lemma cos_inj_pi: "0 \ x \ x \ pi \ 0 \ y \ y \ pi \ cos x = cos y \ x = y" by (metis arccos_cos) lemma arccos_le_pi2: "\0 \ y; y \ 1\ \ arccos y \ pi/2" by (metis (mono_tags) arccos_0 arccos cos_le_one cos_monotone_0_pi_le cos_pi cos_pi_half pi_half_ge_zero antisym_conv less_eq_neg_nonpos linear minus_minus order.trans order_refl) lemma sincos_total_pi_half: assumes "0 \ x" "0 \ y" "x\<^sup>2 + y\<^sup>2 = 1" shows "\t. 0 \ t \ t \ pi/2 \ x = cos t \ y = sin t" proof - have x1: "x \ 1" using assms by (metis le_add_same_cancel1 power2_le_imp_le power_one zero_le_power2) with assms have *: "0 \ arccos x" "cos (arccos x) = x" by (auto simp: arccos) from assms have "y = sqrt (1 - x\<^sup>2)" by (metis abs_of_nonneg add.commute add_diff_cancel real_sqrt_abs) with x1 * assms arccos_le_pi2 [of x] show ?thesis by (rule_tac x="arccos x" in exI) (auto simp: sin_arccos) qed lemma sincos_total_pi: assumes "0 \ y" "x\<^sup>2 + y\<^sup>2 = 1" shows "\t. 0 \ t \ t \ pi \ x = cos t \ y = sin t" proof (cases rule: le_cases [of 0 x]) case le from sincos_total_pi_half [OF le] show ?thesis by (metis pi_ge_two pi_half_le_two add.commute add_le_cancel_left add_mono assms) next case ge then have "0 \ -x" by simp then obtain t where t: "t\0" "t \ pi/2" "-x = cos t" "y = sin t" using sincos_total_pi_half assms by auto (metis \0 \ - x\ power2_minus) show ?thesis by (rule exI [where x = "pi -t"]) (use t in auto) qed lemma sincos_total_2pi_le: assumes "x\<^sup>2 + y\<^sup>2 = 1" shows "\t. 0 \ t \ t \ 2 * pi \ x = cos t \ y = sin t" proof (cases rule: le_cases [of 0 y]) case le from sincos_total_pi [OF le] show ?thesis by (metis assms le_add_same_cancel1 mult.commute mult_2_right order.trans) next case ge then have "0 \ -y" by simp then obtain t where t: "t\0" "t \ pi" "x = cos t" "-y = sin t" using sincos_total_pi assms by auto (metis \0 \ - y\ power2_minus) show ?thesis by (rule exI [where x = "2 * pi - t"]) (use t in auto) qed lemma sincos_total_2pi: assumes "x\<^sup>2 + y\<^sup>2 = 1" obtains t where "0 \ t" "t < 2*pi" "x = cos t" "y = sin t" proof - from sincos_total_2pi_le [OF assms] obtain t where t: "0 \ t" "t \ 2*pi" "x = cos t" "y = sin t" by blast show ?thesis by (cases "t = 2 * pi") (use t that in \force+\) qed lemma arcsin_less_mono: "\x\ \ 1 \ \y\ \ 1 \ arcsin x < arcsin y \ x < y" by (rule trans [OF sin_mono_less_eq [symmetric]]) (use arcsin_ubound arcsin_lbound in auto) lemma arcsin_le_mono: "\x\ \ 1 \ \y\ \ 1 \ arcsin x \ arcsin y \ x \ y" using arcsin_less_mono not_le by blast lemma arcsin_less_arcsin: "- 1 \ x \ x < y \ y \ 1 \ arcsin x < arcsin y" using arcsin_less_mono by auto lemma arcsin_le_arcsin: "- 1 \ x \ x \ y \ y \ 1 \ arcsin x \ arcsin y" using arcsin_le_mono by auto lemma arcsin_nonneg: "x \ {0..1} \ arcsin x \ 0" using arcsin_le_arcsin[of 0 x] by simp lemma arccos_less_mono: "\x\ \ 1 \ \y\ \ 1 \ arccos x < arccos y \ y < x" by (rule trans [OF cos_mono_less_eq [symmetric]]) (use arccos_ubound arccos_lbound in auto) lemma arccos_le_mono: "\x\ \ 1 \ \y\ \ 1 \ arccos x \ arccos y \ y \ x" using arccos_less_mono [of y x] by (simp add: not_le [symmetric]) lemma arccos_less_arccos: "- 1 \ x \ x < y \ y \ 1 \ arccos y < arccos x" using arccos_less_mono by auto lemma arccos_le_arccos: "- 1 \ x \ x \ y \ y \ 1 \ arccos y \ arccos x" using arccos_le_mono by auto lemma arccos_eq_iff: "\x\ \ 1 \ \y\ \ 1 \ arccos x = arccos y \ x = y" using cos_arccos_abs by fastforce lemma arccos_cos_eq_abs: assumes "\\\ \ pi" shows "arccos (cos \) = \\\" unfolding arccos_def proof (intro the_equality conjI; clarify?) show "cos \\\ = cos \" by (simp add: abs_real_def) show "x = \\\" if "cos x = cos \" "0 \ x" "x \ pi" for x by (simp add: \cos \\\ = cos \\ assms cos_inj_pi that) qed (use assms in auto) lemma arccos_cos_eq_abs_2pi: obtains k where "arccos (cos \) = \\ - of_int k * (2 * pi)\" proof - define k where "k \ \(\ + pi) / (2 * pi)\" have lepi: "\\ - of_int k * (2 * pi)\ \ pi" using floor_divide_lower [of "2*pi" "\ + pi"] floor_divide_upper [of "2*pi" "\ + pi"] by (auto simp: k_def abs_if algebra_simps) have "arccos (cos \) = arccos (cos (\ - of_int k * (2 * pi)))" using cos_int_2pin sin_int_2pin by (simp add: cos_diff mult.commute) also have "\ = \\ - of_int k * (2 * pi)\" using arccos_cos_eq_abs lepi by blast finally show ?thesis using that by metis qed lemma arccos_arctan: assumes "-1 < x" "x < 1" shows "arccos x = pi/2 - arctan(x / sqrt(1 - x\<^sup>2))" proof - have "arctan(x / sqrt(1 - x\<^sup>2)) - (pi/2 - arccos x) = 0" proof (rule sin_eq_0_pi) show "- pi < arctan (x / sqrt (1 - x\<^sup>2)) - (pi/2 - arccos x)" using arctan_lbound [of "x / sqrt(1 - x\<^sup>2)"] arccos_bounded [of x] assms by (simp add: algebra_simps) next show "arctan (x / sqrt (1 - x\<^sup>2)) - (pi/2 - arccos x) < pi" using arctan_ubound [of "x / sqrt(1 - x\<^sup>2)"] arccos_bounded [of x] assms by (simp add: algebra_simps) next show "sin (arctan (x / sqrt (1 - x\<^sup>2)) - (pi/2 - arccos x)) = 0" using assms by (simp add: algebra_simps sin_diff cos_add sin_arccos sin_arctan cos_arctan power2_eq_square square_eq_1_iff) qed then show ?thesis by simp qed lemma arcsin_plus_arccos: assumes "-1 \ x" "x \ 1" shows "arcsin x + arccos x = pi/2" proof - have "arcsin x = pi/2 - arccos x" apply (rule sin_inj_pi) using assms arcsin [OF assms] arccos [OF assms] by (auto simp: algebra_simps sin_diff) then show ?thesis by (simp add: algebra_simps) qed lemma arcsin_arccos_eq: "-1 \ x \ x \ 1 \ arcsin x = pi/2 - arccos x" using arcsin_plus_arccos by force lemma arccos_arcsin_eq: "-1 \ x \ x \ 1 \ arccos x = pi/2 - arcsin x" using arcsin_plus_arccos by force lemma arcsin_arctan: "-1 < x \ x < 1 \ arcsin x = arctan(x / sqrt(1 - x\<^sup>2))" by (simp add: arccos_arctan arcsin_arccos_eq) lemma arcsin_arccos_sqrt_pos: "0 \ x \ x \ 1 \ arcsin x = arccos(sqrt(1 - x\<^sup>2))" by (smt (verit, del_insts) arccos_cos arcsin_0 arcsin_le_arcsin arcsin_pi cos_arcsin) lemma arcsin_arccos_sqrt_neg: "-1 \ x \ x \ 0 \ arcsin x = -arccos(sqrt(1 - x\<^sup>2))" using arcsin_arccos_sqrt_pos [of "-x"] by (simp add: arcsin_minus) lemma arccos_arcsin_sqrt_pos: "0 \ x \ x \ 1 \ arccos x = arcsin(sqrt(1 - x\<^sup>2))" by (smt (verit, del_insts) arccos_lbound arccos_le_pi2 arcsin_sin sin_arccos) lemma arccos_arcsin_sqrt_neg: "-1 \ x \ x \ 0 \ arccos x = pi - arcsin(sqrt(1 - x\<^sup>2))" using arccos_arcsin_sqrt_pos [of "-x"] by (simp add: arccos_minus) lemma cos_limit_1: assumes "(\j. cos (\ j)) \ 1" shows "\k. (\j. \ j - of_int (k j) * (2 * pi)) \ 0" proof - have "\\<^sub>F j in sequentially. cos (\ j) \ {- 1..1}" by auto then have "(\j. arccos (cos (\ j))) \ arccos 1" using continuous_on_tendsto_compose [OF continuous_on_arccos' assms] by auto moreover have "\j. \k. arccos (cos (\ j)) = \\ j - of_int k * (2 * pi)\" using arccos_cos_eq_abs_2pi by metis then have "\k. \j. arccos (cos (\ j)) = \\ j - of_int (k j) * (2 * pi)\" by metis ultimately have "\k. (\j. \\ j - of_int (k j) * (2 * pi)\) \ 0" by auto then show ?thesis by (simp add: tendsto_rabs_zero_iff) qed lemma cos_diff_limit_1: assumes "(\j. cos (\ j - \)) \ 1" obtains k where "(\j. \ j - of_int (k j) * (2 * pi)) \ \" proof - obtain k where "(\j. (\ j - \) - of_int (k j) * (2 * pi)) \ 0" using cos_limit_1 [OF assms] by auto then have "(\j. \ + ((\ j - \) - of_int (k j) * (2 * pi))) \ \ + 0" by (rule tendsto_add [OF tendsto_const]) with that show ?thesis by auto qed subsection \Machin's formula\ lemma arctan_one: "arctan 1 = pi/4" by (rule arctan_unique) (simp_all add: tan_45 m2pi_less_pi) lemma tan_total_pi4: assumes "\x\ < 1" shows "\z. - (pi/4) < z \ z < pi/4 \ tan z = x" proof show "- (pi/4) < arctan x \ arctan x < pi/4 \ tan (arctan x) = x" unfolding arctan_one [symmetric] arctan_minus [symmetric] unfolding arctan_less_iff using assms by (auto simp: arctan) qed lemma arctan_add: assumes "\x\ \ 1" "\y\ < 1" shows "arctan x + arctan y = arctan ((x + y) / (1 - x * y))" proof (rule arctan_unique [symmetric]) have "- (pi/4) \ arctan x" "- (pi/4) < arctan y" unfolding arctan_one [symmetric] arctan_minus [symmetric] unfolding arctan_le_iff arctan_less_iff using assms by auto from add_le_less_mono [OF this] show 1: "- (pi/2) < arctan x + arctan y" by simp have "arctan x \ pi/4" "arctan y < pi/4" unfolding arctan_one [symmetric] unfolding arctan_le_iff arctan_less_iff using assms by auto from add_le_less_mono [OF this] show 2: "arctan x + arctan y < pi/2" by simp show "tan (arctan x + arctan y) = (x + y) / (1 - x * y)" using cos_gt_zero_pi [OF 1 2] by (simp add: arctan tan_add) qed lemma arctan_double: "\x\ < 1 \ 2 * arctan x = arctan ((2 * x) / (1 - x\<^sup>2))" by (metis arctan_add linear mult_2 not_less power2_eq_square) theorem machin: "pi/4 = 4 * arctan (1 / 5) - arctan (1/239)" proof - have "\1 / 5\ < (1 :: real)" by auto from arctan_add[OF less_imp_le[OF this] this] have "2 * arctan (1 / 5) = arctan (5 / 12)" by auto moreover have "\5 / 12\ < (1 :: real)" by auto from arctan_add[OF less_imp_le[OF this] this] have "2 * arctan (5 / 12) = arctan (120 / 119)" by auto moreover have "\1\ \ (1::real)" and "\1/239\ < (1::real)" by auto from arctan_add[OF this] have "arctan 1 + arctan (1/239) = arctan (120 / 119)" by auto ultimately have "arctan 1 + arctan (1/239) = 4 * arctan (1 / 5)" by auto then show ?thesis unfolding arctan_one by algebra qed lemma machin_Euler: "5 * arctan (1 / 7) + 2 * arctan (3 / 79) = pi/4" proof - have 17: "\1 / 7\ < (1 :: real)" by auto with arctan_double have "2 * arctan (1 / 7) = arctan (7 / 24)" by simp (simp add: field_simps) moreover have "\7 / 24\ < (1 :: real)" by auto with arctan_double have "2 * arctan (7 / 24) = arctan (336 / 527)" by simp (simp add: field_simps) moreover have "\336 / 527\ < (1 :: real)" by auto from arctan_add[OF less_imp_le[OF 17] this] have "arctan(1/7) + arctan (336 / 527) = arctan (2879 / 3353)" by auto ultimately have I: "5 * arctan (1 / 7) = arctan (2879 / 3353)" by auto have 379: "\3 / 79\ < (1 :: real)" by auto with arctan_double have II: "2 * arctan (3 / 79) = arctan (237 / 3116)" by simp (simp add: field_simps) have *: "\2879 / 3353\ < (1 :: real)" by auto have "\237 / 3116\ < (1 :: real)" by auto from arctan_add[OF less_imp_le[OF *] this] have "arctan (2879/3353) + arctan (237/3116) = pi/4" by (simp add: arctan_one) with I II show ?thesis by auto qed (*But could also prove MACHIN_GAUSS: 12 * arctan(1/18) + 8 * arctan(1/57) - 5 * arctan(1/239) = pi/4*) subsection \Introducing the inverse tangent power series\ lemma monoseq_arctan_series: fixes x :: real assumes "\x\ \ 1" shows "monoseq (\n. 1 / real (n * 2 + 1) * x^(n * 2 + 1))" (is "monoseq ?a") proof (cases "x = 0") case True then show ?thesis by (auto simp: monoseq_def) next case False have "norm x \ 1" and "x \ 1" and "-1 \ x" using assms by auto show "monoseq ?a" proof - have mono: "1 / real (Suc (Suc n * 2)) * x ^ Suc (Suc n * 2) \ 1 / real (Suc (n * 2)) * x ^ Suc (n * 2)" if "0 \ x" and "x \ 1" for n and x :: real proof (rule mult_mono) show "1 / real (Suc (Suc n * 2)) \ 1 / real (Suc (n * 2))" by (rule frac_le) simp_all show "0 \ 1 / real (Suc (n * 2))" by auto show "x ^ Suc (Suc n * 2) \ x ^ Suc (n * 2)" by (rule power_decreasing) (simp_all add: \0 \ x\ \x \ 1\) show "0 \ x ^ Suc (Suc n * 2)" by (rule zero_le_power) (simp add: \0 \ x\) qed show ?thesis proof (cases "0 \ x") case True from mono[OF this \x \ 1\, THEN allI] show ?thesis unfolding Suc_eq_plus1[symmetric] by (rule mono_SucI2) next case False then have "0 \ - x" and "- x \ 1" using \-1 \ x\ by auto from mono[OF this] have "1 / real (Suc (Suc n * 2)) * x ^ Suc (Suc n * 2) \ 1 / real (Suc (n * 2)) * x ^ Suc (n * 2)" for n using \0 \ -x\ by auto then show ?thesis unfolding Suc_eq_plus1[symmetric] by (rule mono_SucI1[OF allI]) qed qed qed lemma zeroseq_arctan_series: fixes x :: real assumes "\x\ \ 1" shows "(\n. 1 / real (n * 2 + 1) * x^(n * 2 + 1)) \ 0" (is "?a \ 0") proof (cases "x = 0") case True then show ?thesis by simp next case False have "norm x \ 1" and "x \ 1" and "-1 \ x" using assms by auto show "?a \ 0" proof (cases "\x\ < 1") case True then have "norm x < 1" by auto from tendsto_mult[OF LIMSEQ_inverse_real_of_nat LIMSEQ_power_zero[OF \norm x < 1\, THEN LIMSEQ_Suc]] have "(\n. 1 / real (n + 1) * x ^ (n + 1)) \ 0" unfolding inverse_eq_divide Suc_eq_plus1 by simp then show ?thesis using pos2 by (rule LIMSEQ_linear) next case False then have "x = -1 \ x = 1" using \\x\ \ 1\ by auto then have n_eq: "\ n. x ^ (n * 2 + 1) = x" unfolding One_nat_def by auto from tendsto_mult[OF LIMSEQ_inverse_real_of_nat[THEN LIMSEQ_linear, OF pos2, unfolded inverse_eq_divide] tendsto_const[of x]] show ?thesis unfolding n_eq Suc_eq_plus1 by auto qed qed lemma summable_arctan_series: fixes n :: nat assumes "\x\ \ 1" shows "summable (\ k. (-1)^k * (1 / real (k*2+1) * x ^ (k*2+1)))" (is "summable (?c x)") by (rule summable_Leibniz(1), rule zeroseq_arctan_series[OF assms], rule monoseq_arctan_series[OF assms]) lemma DERIV_arctan_series: assumes "\x\ < 1" shows "DERIV (\x'. \k. (-1)^k * (1 / real (k * 2 + 1) * x' ^ (k * 2 + 1))) x :> (\k. (-1)^k * x^(k * 2))" (is "DERIV ?arctan _ :> ?Int") proof - let ?f = "\n. if even n then (-1)^(n div 2) * 1 / real (Suc n) else 0" have n_even: "even n \ 2 * (n div 2) = n" for n :: nat by presburger then have if_eq: "?f n * real (Suc n) * x'^n = (if even n then (-1)^(n div 2) * x'^(2 * (n div 2)) else 0)" for n x' by auto have summable_Integral: "summable (\ n. (- 1) ^ n * x^(2 * n))" if "\x\ < 1" for x :: real proof - from that have "x\<^sup>2 < 1" by (simp add: abs_square_less_1) have "summable (\ n. (- 1) ^ n * (x\<^sup>2) ^n)" by (rule summable_Leibniz(1)) (auto intro!: LIMSEQ_realpow_zero monoseq_realpow \x\<^sup>2 < 1\ order_less_imp_le[OF \x\<^sup>2 < 1\]) then show ?thesis by (simp only: power_mult) qed have sums_even: "(sums) f = (sums) (\ n. if even n then f (n div 2) else 0)" for f :: "nat \ real" proof - have "f sums x = (\ n. if even n then f (n div 2) else 0) sums x" for x :: real proof assume "f sums x" from sums_if[OF sums_zero this] show "(\n. if even n then f (n div 2) else 0) sums x" by auto next assume "(\ n. if even n then f (n div 2) else 0) sums x" from LIMSEQ_linear[OF this[simplified sums_def] pos2, simplified sum_split_even_odd[simplified mult.commute]] show "f sums x" unfolding sums_def by auto qed then show ?thesis .. qed have Int_eq: "(\n. ?f n * real (Suc n) * x^n) = ?Int" unfolding if_eq mult.commute[of _ 2] suminf_def sums_even[of "\ n. (- 1) ^ n * x ^ (2 * n)", symmetric] by auto have arctan_eq: "(\n. ?f n * x^(Suc n)) = ?arctan x" for x proof - have if_eq': "\n. (if even n then (- 1) ^ (n div 2) * 1 / real (Suc n) else 0) * x ^ Suc n = (if even n then (- 1) ^ (n div 2) * (1 / real (Suc (2 * (n div 2))) * x ^ Suc (2 * (n div 2))) else 0)" using n_even by auto have idx_eq: "\n. n * 2 + 1 = Suc (2 * n)" by auto then show ?thesis unfolding if_eq' idx_eq suminf_def sums_even[of "\ n. (- 1) ^ n * (1 / real (Suc (2 * n)) * x ^ Suc (2 * n))", symmetric] by auto qed have "DERIV (\ x. \ n. ?f n * x^(Suc n)) x :> (\n. ?f n * real (Suc n) * x^n)" proof (rule DERIV_power_series') show "x \ {- 1 <..< 1}" using \\ x \ < 1\ by auto show "summable (\ n. ?f n * real (Suc n) * x'^n)" if x'_bounds: "x' \ {- 1 <..< 1}" for x' :: real proof - from that have "\x'\ < 1" by auto then show ?thesis using that sums_summable sums_if [OF sums_0 [of "\x. 0"] summable_sums [OF summable_Integral]] by (auto simp add: if_distrib [of "\x. x * y" for y] cong: if_cong) qed qed auto then show ?thesis by (simp only: Int_eq arctan_eq) qed lemma arctan_series: assumes "\x\ \ 1" shows "arctan x = (\k. (-1)^k * (1 / real (k * 2 + 1) * x ^ (k * 2 + 1)))" (is "_ = suminf (\ n. ?c x n)") proof - let ?c' = "\x n. (-1)^n * x^(n*2)" have DERIV_arctan_suminf: "DERIV (\ x. suminf (?c x)) x :> (suminf (?c' x))" if "0 < r" and "r < 1" and "\x\ < r" for r x :: real proof (rule DERIV_arctan_series) from that show "\x\ < 1" using \r < 1\ and \\x\ < r\ by auto qed { fix x :: real assume "\x\ \ 1" note summable_Leibniz[OF zeroseq_arctan_series[OF this] monoseq_arctan_series[OF this]] } note arctan_series_borders = this have when_less_one: "arctan x = (\k. ?c x k)" if "\x\ < 1" for x :: real proof - obtain r where "\x\ < r" and "r < 1" using dense[OF \\x\ < 1\] by blast then have "0 < r" and "- r < x" and "x < r" by auto have suminf_eq_arctan_bounded: "suminf (?c x) - arctan x = suminf (?c a) - arctan a" if "-r < a" and "b < r" and "a < b" and "a \ x" and "x \ b" for x a b proof - from that have "\x\ < r" by auto show "suminf (?c x) - arctan x = suminf (?c a) - arctan a" proof (rule DERIV_isconst2[of "a" "b"]) show "a < b" and "a \ x" and "x \ b" using \a < b\ \a \ x\ \x \ b\ by auto have "\x. - r < x \ x < r \ DERIV (\ x. suminf (?c x) - arctan x) x :> 0" proof (rule allI, rule impI) fix x assume "-r < x \ x < r" then have "\x\ < r" by auto with \r < 1\ have "\x\ < 1" by auto have "\- (x\<^sup>2)\ < 1" using abs_square_less_1 \\x\ < 1\ by auto then have "(\n. (- (x\<^sup>2)) ^ n) sums (1 / (1 - (- (x\<^sup>2))))" unfolding real_norm_def[symmetric] by (rule geometric_sums) then have "(?c' x) sums (1 / (1 - (- (x\<^sup>2))))" unfolding power_mult_distrib[symmetric] power_mult mult.commute[of _ 2] by auto then have suminf_c'_eq_geom: "inverse (1 + x\<^sup>2) = suminf (?c' x)" using sums_unique unfolding inverse_eq_divide by auto have "DERIV (\ x. suminf (?c x)) x :> (inverse (1 + x\<^sup>2))" unfolding suminf_c'_eq_geom by (rule DERIV_arctan_suminf[OF \0 < r\ \r < 1\ \\x\ < r\]) from DERIV_diff [OF this DERIV_arctan] show "DERIV (\x. suminf (?c x) - arctan x) x :> 0" by auto qed then have DERIV_in_rball: "\y. a \ y \ y \ b \ DERIV (\x. suminf (?c x) - arctan x) y :> 0" using \-r < a\ \b < r\ by auto then show "\y. \a < y; y < b\ \ DERIV (\x. suminf (?c x) - arctan x) y :> 0" using \\x\ < r\ by auto show "continuous_on {a..b} (\x. suminf (?c x) - arctan x)" using DERIV_in_rball DERIV_atLeastAtMost_imp_continuous_on by blast qed qed have suminf_arctan_zero: "suminf (?c 0) - arctan 0 = 0" unfolding Suc_eq_plus1[symmetric] power_Suc2 mult_zero_right arctan_zero_zero suminf_zero by auto have "suminf (?c x) - arctan x = 0" proof (cases "x = 0") case True then show ?thesis using suminf_arctan_zero by auto next case False then have "0 < \x\" and "- \x\ < \x\" by auto have "suminf (?c (- \x\)) - arctan (- \x\) = suminf (?c 0) - arctan 0" by (rule suminf_eq_arctan_bounded[where x1=0 and a1="-\x\" and b1="\x\", symmetric]) (simp_all only: \\x\ < r\ \-\x\ < \x\\ neg_less_iff_less) moreover have "suminf (?c x) - arctan x = suminf (?c (- \x\)) - arctan (- \x\)" by (rule suminf_eq_arctan_bounded[where x1=x and a1="- \x\" and b1="\x\"]) (simp_all only: \\x\ < r\ \- \x\ < \x\\ neg_less_iff_less) ultimately show ?thesis using suminf_arctan_zero by auto qed then show ?thesis by auto qed show "arctan x = suminf (\n. ?c x n)" proof (cases "\x\ < 1") case True then show ?thesis by (rule when_less_one) next case False then have "\x\ = 1" using \\x\ \ 1\ by auto let ?a = "\x n. \1 / real (n * 2 + 1) * x^(n * 2 + 1)\" let ?diff = "\x n. \arctan x - (\i" have "?diff 1 n \ ?a 1 n" for n :: nat proof - have "0 < (1 :: real)" by auto moreover have "?diff x n \ ?a x n" if "0 < x" and "x < 1" for x :: real proof - from that have "\x\ \ 1" and "\x\ < 1" by auto from \0 < x\ have "0 < 1 / real (0 * 2 + (1::nat)) * x ^ (0 * 2 + 1)" by auto note bounds = mp[OF arctan_series_borders(2)[OF \\x\ \ 1\] this, unfolded when_less_one[OF \\x\ < 1\, symmetric], THEN spec] have "0 < 1 / real (n*2+1) * x^(n*2+1)" by (rule mult_pos_pos) (simp_all only: zero_less_power[OF \0 < x\], auto) then have a_pos: "?a x n = 1 / real (n*2+1) * x^(n*2+1)" by (rule abs_of_pos) show ?thesis proof (cases "even n") case True then have sgn_pos: "(-1)^n = (1::real)" by auto from \even n\ obtain m where "n = 2 * m" .. then have "2 * m = n" .. from bounds[of m, unfolded this atLeastAtMost_iff] have "\arctan x - (\i \ (\ii = ?c x n" by auto also have "\ = ?a x n" unfolding sgn_pos a_pos by auto finally show ?thesis . next case False then have sgn_neg: "(-1)^n = (-1::real)" by auto from \odd n\ obtain m where "n = 2 * m + 1" .. then have m_def: "2 * m + 1 = n" .. then have m_plus: "2 * (m + 1) = n + 1" by auto from bounds[of "m + 1", unfolded this atLeastAtMost_iff, THEN conjunct1] bounds[of m, unfolded m_def atLeastAtMost_iff, THEN conjunct2] have "\arctan x - (\i \ (\ii = - ?c x n" by auto also have "\ = ?a x n" unfolding sgn_neg a_pos by auto finally show ?thesis . qed qed hence "\x \ { 0 <..< 1 }. 0 \ ?a x n - ?diff x n" by auto moreover have "isCont (\ x. ?a x n - ?diff x n) x" for x unfolding diff_conv_add_uminus divide_inverse by (auto intro!: isCont_add isCont_rabs continuous_ident isCont_minus isCont_arctan continuous_at_within_inverse isCont_mult isCont_power continuous_const isCont_sum simp del: add_uminus_conv_diff) ultimately have "0 \ ?a 1 n - ?diff 1 n" by (rule LIM_less_bound) then show ?thesis by auto qed have "?a 1 \ 0" unfolding tendsto_rabs_zero_iff power_one divide_inverse One_nat_def by (auto intro!: tendsto_mult LIMSEQ_linear LIMSEQ_inverse_real_of_nat simp del: of_nat_Suc) have "?diff 1 \ 0" proof (rule LIMSEQ_I) fix r :: real assume "0 < r" obtain N :: nat where N_I: "N \ n \ ?a 1 n < r" for n using LIMSEQ_D[OF \?a 1 \ 0\ \0 < r\] by auto have "norm (?diff 1 n - 0) < r" if "N \ n" for n using \?diff 1 n \ ?a 1 n\ N_I[OF that] by auto then show "\N. \ n \ N. norm (?diff 1 n - 0) < r" by blast qed from this [unfolded tendsto_rabs_zero_iff, THEN tendsto_add [OF _ tendsto_const], of "- arctan 1", THEN tendsto_minus] have "(?c 1) sums (arctan 1)" unfolding sums_def by auto then have "arctan 1 = (\i. ?c 1 i)" by (rule sums_unique) show ?thesis proof (cases "x = 1") case True then show ?thesis by (simp add: \arctan 1 = (\ i. ?c 1 i)\) next case False then have "x = -1" using \\x\ = 1\ by auto have "- (pi/2) < 0" using pi_gt_zero by auto have "- (2 * pi) < 0" using pi_gt_zero by auto have c_minus_minus: "?c (- 1) i = - ?c 1 i" for i by auto have "arctan (- 1) = arctan (tan (-(pi/4)))" unfolding tan_45 tan_minus .. also have "\ = - (pi/4)" by (rule arctan_tan) (auto simp: order_less_trans[OF \- (pi/2) < 0\ pi_gt_zero]) also have "\ = - (arctan (tan (pi/4)))" unfolding neg_equal_iff_equal by (rule arctan_tan[symmetric]) (auto simp: order_less_trans[OF \- (2 * pi) < 0\ pi_gt_zero]) also have "\ = - (arctan 1)" unfolding tan_45 .. also have "\ = - (\ i. ?c 1 i)" using \arctan 1 = (\ i. ?c 1 i)\ by auto also have "\ = (\ i. ?c (- 1) i)" using suminf_minus[OF sums_summable[OF \(?c 1) sums (arctan 1)\]] unfolding c_minus_minus by auto finally show ?thesis using \x = -1\ by auto qed qed qed lemma arctan_half: "arctan x = 2 * arctan (x / (1 + sqrt(1 + x\<^sup>2)))" for x :: real proof - obtain y where low: "- (pi/2) < y" and high: "y < pi/2" and y_eq: "tan y = x" using tan_total by blast then have low2: "- (pi/2) < y / 2" and high2: "y / 2 < pi/2" by auto have "0 < cos y" by (rule cos_gt_zero_pi[OF low high]) then have "cos y \ 0" and cos_sqrt: "sqrt ((cos y)\<^sup>2) = cos y" by auto have "1 + (tan y)\<^sup>2 = 1 + (sin y)\<^sup>2 / (cos y)\<^sup>2" unfolding tan_def power_divide .. also have "\ = (cos y)\<^sup>2 / (cos y)\<^sup>2 + (sin y)\<^sup>2 / (cos y)\<^sup>2" using \cos y \ 0\ by auto also have "\ = 1 / (cos y)\<^sup>2" unfolding add_divide_distrib[symmetric] sin_cos_squared_add2 .. finally have "1 + (tan y)\<^sup>2 = 1 / (cos y)\<^sup>2" . have "sin y / (cos y + 1) = tan y / ((cos y + 1) / cos y)" unfolding tan_def using \cos y \ 0\ by (simp add: field_simps) also have "\ = tan y / (1 + 1 / cos y)" using \cos y \ 0\ unfolding add_divide_distrib by auto also have "\ = tan y / (1 + 1 / sqrt ((cos y)\<^sup>2))" unfolding cos_sqrt .. also have "\ = tan y / (1 + sqrt (1 / (cos y)\<^sup>2))" unfolding real_sqrt_divide by auto finally have eq: "sin y / (cos y + 1) = tan y / (1 + sqrt(1 + (tan y)\<^sup>2))" unfolding \1 + (tan y)\<^sup>2 = 1 / (cos y)\<^sup>2\ . have "arctan x = y" using arctan_tan low high y_eq by auto also have "\ = 2 * (arctan (tan (y/2)))" using arctan_tan[OF low2 high2] by auto also have "\ = 2 * (arctan (sin y / (cos y + 1)))" unfolding tan_half by auto finally show ?thesis unfolding eq \tan y = x\ . qed lemma arctan_monotone: "x < y \ arctan x < arctan y" by (simp only: arctan_less_iff) lemma arctan_monotone': "x \ y \ arctan x \ arctan y" by (simp only: arctan_le_iff) lemma arctan_inverse: assumes "x \ 0" shows "arctan (1/x) = sgn x * pi/2 - arctan x" proof (rule arctan_unique) have \
: "x > 0 \ arctan x < pi" using arctan_bounded [of x] by linarith show "- (pi/2) < sgn x * pi/2 - arctan x" using assms by (auto simp: sgn_real_def arctan algebra_simps \
) show "sgn x * pi/2 - arctan x < pi/2" using arctan_bounded [of "- x"] assms by (auto simp: algebra_simps sgn_real_def arctan_minus) show "tan (sgn x * pi/2 - arctan x) = 1/x" unfolding tan_inverse [of "arctan x", unfolded tan_arctan] sgn_real_def by (simp add: tan_def cos_arctan sin_arctan sin_diff cos_diff) qed theorem pi_series: "pi/4 = (\k. (-1)^k * 1 / real (k * 2 + 1))" (is "_ = ?SUM") proof - have "pi/4 = arctan 1" using arctan_one by auto also have "\ = ?SUM" using arctan_series[of 1] by auto finally show ?thesis by auto qed subsection \Existence of Polar Coordinates\ lemma cos_x_y_le_one: "\x / sqrt (x\<^sup>2 + y\<^sup>2)\ \ 1" by (rule power2_le_imp_le [OF _ zero_le_one]) (simp add: power_divide divide_le_eq not_sum_power2_lt_zero) lemma polar_Ex: "\r::real. \a. x = r * cos a \ y = r * sin a" proof - have polar_ex1: "\r a. x = r * cos a \ y = r * sin a" if "0 < y" for y proof - have "x = sqrt (x\<^sup>2 + y\<^sup>2) * cos (arccos (x / sqrt (x\<^sup>2 + y\<^sup>2)))" by (simp add: cos_arccos_abs [OF cos_x_y_le_one]) moreover have "y = sqrt (x\<^sup>2 + y\<^sup>2) * sin (arccos (x / sqrt (x\<^sup>2 + y\<^sup>2)))" using that by (simp add: sin_arccos_abs [OF cos_x_y_le_one] power_divide right_diff_distrib flip: real_sqrt_mult) ultimately show ?thesis by blast qed show ?thesis proof (cases "0::real" y rule: linorder_cases) case less then show ?thesis by (rule polar_ex1) next case equal then show ?thesis by (force simp: intro!: cos_zero sin_zero) next case greater with polar_ex1 [where y="-y"] show ?thesis by auto (metis cos_minus minus_minus minus_mult_right sin_minus) qed qed subsection \Basics about polynomial functions: products, extremal behaviour and root counts\ lemma polynomial_product_nat: fixes x :: nat assumes m: "\i. i > m \ int (a i) = 0" and n: "\j. j > n \ int (b j) = 0" shows "(\i\m. (a i) * x ^ i) * (\j\n. (b j) * x ^ j) = (\r\m + n. (\k\r. (a k) * (b (r - k))) * x ^ r)" using polynomial_product [of m a n b x] assms by (simp only: of_nat_mult [symmetric] of_nat_power [symmetric] of_nat_eq_iff Int.int_sum [symmetric]) lemma polyfun_diff: (*COMPLEX_SUB_POLYFUN in HOL Light*) fixes x :: "'a::idom" assumes "1 \ n" shows "(\i\n. a i * x^i) - (\i\n. a i * y^i) = (x - y) * (\ji=Suc j..n. a i * y^(i - j - 1)) * x^j)" proof - have h: "bij_betw (\(i,j). (j,i)) ((SIGMA i : atMost n. lessThan i)) (SIGMA j : lessThan n. {Suc j..n})" by (auto simp: bij_betw_def inj_on_def) have "(\i\n. a i * x^i) - (\i\n. a i * y^i) = (\i\n. a i * (x^i - y^i))" by (simp add: right_diff_distrib sum_subtractf) also have "\ = (\i\n. a i * (x - y) * (\j = (\i\n. \j = (\(i,j) \ (SIGMA i : atMost n. lessThan i). a i * (x - y) * (y^(i - Suc j) * x^j))" by (simp add: sum.Sigma) also have "\ = (\(j,i) \ (SIGMA j : lessThan n. {Suc j..n}). a i * (x - y) * (y^(i - Suc j) * x^j))" by (auto simp: sum.reindex_bij_betw [OF h, symmetric] intro: sum.cong_simp) also have "\ = (\ji=Suc j..n. a i * (x - y) * (y^(i - Suc j) * x^j))" by (simp add: sum.Sigma) also have "\ = (x - y) * (\ji=Suc j..n. a i * y^(i - j - 1)) * x^j)" by (simp add: sum_distrib_left mult_ac) finally show ?thesis . qed lemma polyfun_diff_alt: (*COMPLEX_SUB_POLYFUN_ALT in HOL Light*) fixes x :: "'a::idom" assumes "1 \ n" shows "(\i\n. a i * x^i) - (\i\n. a i * y^i) = (x - y) * ((\jki=Suc j..n. a i * y^(i - j - 1)) = (\kk. k < n - j \ k \ (\i. i - Suc j) ` {Suc j..n}" by (rule_tac x="k + Suc j" in image_eqI, auto) then have h: "bij_betw (\i. i - (j + 1)) {Suc j..n} (lessThan (n-j))" by (auto simp: bij_betw_def inj_on_def) then show ?thesis by (auto simp: sum.reindex_bij_betw [OF h, symmetric] intro: sum.cong_simp) qed then show ?thesis by (simp add: polyfun_diff [OF assms] sum_distrib_right) qed lemma polyfun_linear_factor: (*COMPLEX_POLYFUN_LINEAR_FACTOR in HOL Light*) fixes a :: "'a::idom" shows "\b. \z. (\i\n. c(i) * z^i) = (z - a) * (\ii\n. c(i) * a^i)" proof (cases "n = 0") case True then show ?thesis by simp next case False have "(\b. \z. (\i\n. c i * z^i) = (z - a) * (\ii\n. c i * a^i)) \ (\b. \z. (\i\n. c i * z^i) - (\i\n. c i * a^i) = (z - a) * (\i \ (\b. \z. (z - a) * (\ji = Suc j..n. c i * a^(i - Suc j)) * z^j) = (z - a) * (\i = True" by auto finally show ?thesis by simp qed lemma polyfun_linear_factor_root: (*COMPLEX_POLYFUN_LINEAR_FACTOR_ROOT in HOL Light*) fixes a :: "'a::idom" assumes "(\i\n. c(i) * a^i) = 0" obtains b where "\z. (\i\n. c i * z^i) = (z - a) * (\iw. \i\n. c i * w^i) a" for c :: "nat \ 'a::real_normed_div_algebra" by simp lemma zero_polynom_imp_zero_coeffs: fixes c :: "nat \ 'a::{ab_semigroup_mult,real_normed_div_algebra}" assumes "\w. (\i\n. c i * w^i) = 0" "k \ n" shows "c k = 0" using assms proof (induction n arbitrary: c k) case 0 then show ?case by simp next case (Suc n c k) have [simp]: "c 0 = 0" using Suc.prems(1) [of 0] by simp have "(\i\Suc n. c i * w^i) = w * (\i\n. c (Suc i) * w^i)" for w proof - have "(\i\Suc n. c i * w^i) = (\i\n. c (Suc i) * w ^ Suc i)" unfolding Set_Interval.sum.atMost_Suc_shift by simp also have "\ = w * (\i\n. c (Suc i) * w^i)" by (simp add: sum_distrib_left ac_simps) finally show ?thesis . qed then have w: "\w. w \ 0 \ (\i\n. c (Suc i) * w^i) = 0" using Suc by auto then have "(\h. \i\n. c (Suc i) * h^i) \0\ 0" by (simp cong: LIM_cong) \ \the case \w = 0\ by continuity\ then have "(\i\n. c (Suc i) * 0^i) = 0" using isCont_polynom [of 0 "\i. c (Suc i)" n] LIM_unique by (force simp: Limits.isCont_iff) then have "\w. (\i\n. c (Suc i) * w^i) = 0" using w by metis then have "\i. i \ n \ c (Suc i) = 0" using Suc.IH [of "\i. c (Suc i)"] by blast then show ?case using \k \ Suc n\ by (cases k) auto qed lemma polyfun_rootbound: (*COMPLEX_POLYFUN_ROOTBOUND in HOL Light*) fixes c :: "nat \ 'a::{idom,real_normed_div_algebra}" assumes "c k \ 0" "k\n" shows "finite {z. (\i\n. c(i) * z^i) = 0} \ card {z. (\i\n. c(i) * z^i) = 0} \ n" using assms proof (induction n arbitrary: c k) case 0 then show ?case by simp next case (Suc m c k) let ?succase = ?case show ?case proof (cases "{z. (\i\Suc m. c(i) * z^i) = 0} = {}") case True then show ?succase by simp next case False then obtain z0 where z0: "(\i\Suc m. c(i) * z0^i) = 0" by blast then obtain b where b: "\w. (\i\Suc m. c i * w^i) = (w - z0) * (\i\m. b i * w^i)" using polyfun_linear_factor_root [OF z0, unfolded lessThan_Suc_atMost] by blast then have eq: "{z. (\i\Suc m. c i * z^i) = 0} = insert z0 {z. (\i\m. b i * z^i) = 0}" by auto have "\ (\k\m. b k = 0)" proof assume [simp]: "\k\m. b k = 0" then have "\w. (\i\m. b i * w^i) = 0" by simp then have "\w. (\i\Suc m. c i * w^i) = 0" using b by simp then have "\k. k \ Suc m \ c k = 0" using zero_polynom_imp_zero_coeffs by blast then show False using Suc.prems by blast qed then obtain k' where bk': "b k' \ 0" "k' \ m" by blast show ?succase using Suc.IH [of b k'] bk' by (simp add: eq card_insert_if del: sum.atMost_Suc) qed qed lemma fixes c :: "nat \ 'a::{idom,real_normed_div_algebra}" assumes "c k \ 0" "k\n" shows polyfun_roots_finite: "finite {z. (\i\n. c(i) * z^i) = 0}" and polyfun_roots_card: "card {z. (\i\n. c(i) * z^i) = 0} \ n" using polyfun_rootbound assms by auto lemma polyfun_finite_roots: (*COMPLEX_POLYFUN_FINITE_ROOTS in HOL Light*) fixes c :: "nat \ 'a::{idom,real_normed_div_algebra}" shows "finite {x. (\i\n. c i * x^i) = 0} \ (\i\n. c i \ 0)" (is "?lhs = ?rhs") proof assume ?lhs moreover have "\ finite {x. (\i\n. c i * x^i) = 0}" if "\i\n. c i = 0" proof - from that have "\x. (\i\n. c i * x^i) = 0" by simp then show ?thesis using ex_new_if_finite [OF infinite_UNIV_char_0 [where 'a='a]] by auto qed ultimately show ?rhs by metis next assume ?rhs with polyfun_rootbound show ?lhs by blast qed lemma polyfun_eq_0: "(\x. (\i\n. c i * x^i) = 0) \ (\i\n. c i = 0)" for c :: "nat \ 'a::{idom,real_normed_div_algebra}" (*COMPLEX_POLYFUN_EQ_0 in HOL Light*) using zero_polynom_imp_zero_coeffs by auto lemma polyfun_eq_coeffs: "(\x. (\i\n. c i * x^i) = (\i\n. d i * x^i)) \ (\i\n. c i = d i)" for c :: "nat \ 'a::{idom,real_normed_div_algebra}" proof - have "(\x. (\i\n. c i * x^i) = (\i\n. d i * x^i)) \ (\x. (\i\n. (c i - d i) * x^i) = 0)" by (simp add: left_diff_distrib Groups_Big.sum_subtractf) also have "\ \ (\i\n. c i - d i = 0)" by (rule polyfun_eq_0) finally show ?thesis by simp qed lemma polyfun_eq_const: (*COMPLEX_POLYFUN_EQ_CONST in HOL Light*) fixes c :: "nat \ 'a::{idom,real_normed_div_algebra}" shows "(\x. (\i\n. c i * x^i) = k) \ c 0 = k \ (\i \ {1..n}. c i = 0)" (is "?lhs = ?rhs") proof - have *: "\x. (\i\n. (if i=0 then k else 0) * x^i) = k" by (induct n) auto show ?thesis proof assume ?lhs with * have "(\i\n. c i = (if i=0 then k else 0))" by (simp add: polyfun_eq_coeffs [symmetric]) then show ?rhs by simp next assume ?rhs then show ?lhs by (induct n) auto qed qed lemma root_polyfun: fixes z :: "'a::idom" assumes "1 \ n" shows "z^n = a \ (\i\n. (if i = 0 then -a else if i=n then 1 else 0) * z^i) = 0" using assms by (cases n) (simp_all add: sum.atLeast_Suc_atMost atLeast0AtMost [symmetric]) lemma assumes "SORT_CONSTRAINT('a::{idom,real_normed_div_algebra})" and "1 \ n" shows finite_roots_unity: "finite {z::'a. z^n = 1}" and card_roots_unity: "card {z::'a. z^n = 1} \ n" using polyfun_rootbound [of "\i. if i = 0 then -1 else if i=n then 1 else 0" n n] assms(2) by (auto simp: root_polyfun [OF assms(2)]) subsection \Hyperbolic functions\ definition sinh :: "'a :: {banach, real_normed_algebra_1} \ 'a" where "sinh x = (exp x - exp (-x)) /\<^sub>R 2" definition cosh :: "'a :: {banach, real_normed_algebra_1} \ 'a" where "cosh x = (exp x + exp (-x)) /\<^sub>R 2" definition tanh :: "'a :: {banach, real_normed_field} \ 'a" where "tanh x = sinh x / cosh x" definition arsinh :: "'a :: {banach, real_normed_algebra_1, ln} \ 'a" where "arsinh x = ln (x + (x^2 + 1) powr of_real (1/2))" definition arcosh :: "'a :: {banach, real_normed_algebra_1, ln} \ 'a" where "arcosh x = ln (x + (x^2 - 1) powr of_real (1/2))" definition artanh :: "'a :: {banach, real_normed_field, ln} \ 'a" where "artanh x = ln ((1 + x) / (1 - x)) / 2" lemma arsinh_0 [simp]: "arsinh 0 = 0" by (simp add: arsinh_def) lemma arcosh_1 [simp]: "arcosh 1 = 0" by (simp add: arcosh_def) lemma artanh_0 [simp]: "artanh 0 = 0" by (simp add: artanh_def) lemma tanh_altdef: "tanh x = (exp x - exp (-x)) / (exp x + exp (-x))" proof - have "tanh x = (2 *\<^sub>R sinh x) / (2 *\<^sub>R cosh x)" by (simp add: tanh_def scaleR_conv_of_real) also have "2 *\<^sub>R sinh x = exp x - exp (-x)" by (simp add: sinh_def) also have "2 *\<^sub>R cosh x = exp x + exp (-x)" by (simp add: cosh_def) finally show ?thesis . qed lemma tanh_real_altdef: "tanh (x::real) = (1 - exp (- 2 * x)) / (1 + exp (- 2 * x))" proof - have [simp]: "exp (2 * x) = exp x * exp x" "exp (x * 2) = exp x * exp x" by (subst exp_add [symmetric]; simp)+ have "tanh x = (2 * exp (-x) * sinh x) / (2 * exp (-x) * cosh x)" by (simp add: tanh_def) also have "2 * exp (-x) * sinh x = 1 - exp (-2*x)" by (simp add: exp_minus field_simps sinh_def) also have "2 * exp (-x) * cosh x = 1 + exp (-2*x)" by (simp add: exp_minus field_simps cosh_def) finally show ?thesis . qed lemma sinh_converges: "(\n. if even n then 0 else x ^ n /\<^sub>R fact n) sums sinh x" proof - have "(\n. (x ^ n /\<^sub>R fact n - (-x) ^ n /\<^sub>R fact n) /\<^sub>R 2) sums sinh x" unfolding sinh_def by (intro sums_scaleR_right sums_diff exp_converges) also have "(\n. (x ^ n /\<^sub>R fact n - (-x) ^ n /\<^sub>R fact n) /\<^sub>R 2) = (\n. if even n then 0 else x ^ n /\<^sub>R fact n)" by auto finally show ?thesis . qed lemma cosh_converges: "(\n. if even n then x ^ n /\<^sub>R fact n else 0) sums cosh x" proof - have "(\n. (x ^ n /\<^sub>R fact n + (-x) ^ n /\<^sub>R fact n) /\<^sub>R 2) sums cosh x" unfolding cosh_def by (intro sums_scaleR_right sums_add exp_converges) also have "(\n. (x ^ n /\<^sub>R fact n + (-x) ^ n /\<^sub>R fact n) /\<^sub>R 2) = (\n. if even n then x ^ n /\<^sub>R fact n else 0)" by auto finally show ?thesis . qed lemma sinh_0 [simp]: "sinh 0 = 0" by (simp add: sinh_def) lemma cosh_0 [simp]: "cosh 0 = 1" proof - have "cosh 0 = (1/2) *\<^sub>R (1 + 1)" by (simp add: cosh_def) also have "\ = 1" by (rule scaleR_half_double) finally show ?thesis . qed lemma tanh_0 [simp]: "tanh 0 = 0" by (simp add: tanh_def) lemma sinh_minus [simp]: "sinh (- x) = -sinh x" by (simp add: sinh_def algebra_simps) lemma cosh_minus [simp]: "cosh (- x) = cosh x" by (simp add: cosh_def algebra_simps) lemma tanh_minus [simp]: "tanh (-x) = -tanh x" by (simp add: tanh_def) lemma sinh_ln_real: "x > 0 \ sinh (ln x :: real) = (x - inverse x) / 2" by (simp add: sinh_def exp_minus) lemma cosh_ln_real: "x > 0 \ cosh (ln x :: real) = (x + inverse x) / 2" by (simp add: cosh_def exp_minus) lemma tanh_ln_real: "tanh (ln x :: real) = (x ^ 2 - 1) / (x ^ 2 + 1)" if "x > 0" proof - from that have "(x * 2 - inverse x * 2) * (x\<^sup>2 + 1) = (x\<^sup>2 - 1) * (2 * x + 2 * inverse x)" by (simp add: field_simps power2_eq_square) moreover have "x\<^sup>2 + 1 > 0" using that by (simp add: ac_simps add_pos_nonneg) moreover have "2 * x + 2 * inverse x > 0" using that by (simp add: add_pos_pos) ultimately have "(x * 2 - inverse x * 2) / (2 * x + 2 * inverse x) = (x\<^sup>2 - 1) / (x\<^sup>2 + 1)" by (simp add: frac_eq_eq) with that show ?thesis by (simp add: tanh_def sinh_ln_real cosh_ln_real) qed lemma has_field_derivative_scaleR_right [derivative_intros]: "(f has_field_derivative D) F \ ((\x. c *\<^sub>R f x) has_field_derivative (c *\<^sub>R D)) F" unfolding has_field_derivative_def using has_derivative_scaleR_right[of f "\x. D * x" F c] by (simp add: mult_scaleR_left [symmetric] del: mult_scaleR_left) lemma has_field_derivative_sinh [THEN DERIV_chain2, derivative_intros]: "(sinh has_field_derivative cosh x) (at (x :: 'a :: {banach, real_normed_field}))" unfolding sinh_def cosh_def by (auto intro!: derivative_eq_intros) lemma has_field_derivative_cosh [THEN DERIV_chain2, derivative_intros]: "(cosh has_field_derivative sinh x) (at (x :: 'a :: {banach, real_normed_field}))" unfolding sinh_def cosh_def by (auto intro!: derivative_eq_intros) lemma has_field_derivative_tanh [THEN DERIV_chain2, derivative_intros]: "cosh x \ 0 \ (tanh has_field_derivative 1 - tanh x ^ 2) (at (x :: 'a :: {banach, real_normed_field}))" unfolding tanh_def by (auto intro!: derivative_eq_intros simp: power2_eq_square field_split_simps) lemma has_derivative_sinh [derivative_intros]: fixes g :: "'a \ ('a :: {banach, real_normed_field})" assumes "(g has_derivative (\x. Db * x)) (at x within s)" shows "((\x. sinh (g x)) has_derivative (\y. (cosh (g x) * Db) * y)) (at x within s)" proof - have "((\x. - g x) has_derivative (\y. -(Db * y))) (at x within s)" using assms by (intro derivative_intros) also have "(\y. -(Db * y)) = (\x. (-Db) * x)" by (simp add: fun_eq_iff) finally have "((\x. sinh (g x)) has_derivative (\y. (exp (g x) * Db * y - exp (-g x) * (-Db) * y) /\<^sub>R 2)) (at x within s)" unfolding sinh_def by (intro derivative_intros assms) also have "(\y. (exp (g x) * Db * y - exp (-g x) * (-Db) * y) /\<^sub>R 2) = (\y. (cosh (g x) * Db) * y)" by (simp add: fun_eq_iff cosh_def algebra_simps) finally show ?thesis . qed lemma has_derivative_cosh [derivative_intros]: fixes g :: "'a \ ('a :: {banach, real_normed_field})" assumes "(g has_derivative (\y. Db * y)) (at x within s)" shows "((\x. cosh (g x)) has_derivative (\y. (sinh (g x) * Db) * y)) (at x within s)" proof - have "((\x. - g x) has_derivative (\y. -(Db * y))) (at x within s)" using assms by (intro derivative_intros) also have "(\y. -(Db * y)) = (\y. (-Db) * y)" by (simp add: fun_eq_iff) finally have "((\x. cosh (g x)) has_derivative (\y. (exp (g x) * Db * y + exp (-g x) * (-Db) * y) /\<^sub>R 2)) (at x within s)" unfolding cosh_def by (intro derivative_intros assms) also have "(\y. (exp (g x) * Db * y + exp (-g x) * (-Db) * y) /\<^sub>R 2) = (\y. (sinh (g x) * Db) * y)" by (simp add: fun_eq_iff sinh_def algebra_simps) finally show ?thesis . qed lemma sinh_plus_cosh: "sinh x + cosh x = exp x" proof - have "sinh x + cosh x = (1/2) *\<^sub>R (exp x + exp x)" by (simp add: sinh_def cosh_def algebra_simps) also have "\ = exp x" by (rule scaleR_half_double) finally show ?thesis . qed lemma cosh_plus_sinh: "cosh x + sinh x = exp x" by (subst add.commute) (rule sinh_plus_cosh) lemma cosh_minus_sinh: "cosh x - sinh x = exp (-x)" proof - have "cosh x - sinh x = (1/2) *\<^sub>R (exp (-x) + exp (-x))" by (simp add: sinh_def cosh_def algebra_simps) also have "\ = exp (-x)" by (rule scaleR_half_double) finally show ?thesis . qed lemma sinh_minus_cosh: "sinh x - cosh x = -exp (-x)" using cosh_minus_sinh[of x] by (simp add: algebra_simps) context fixes x :: "'a :: {real_normed_field, banach}" begin lemma sinh_zero_iff: "sinh x = 0 \ exp x \ {1, -1}" by (auto simp: sinh_def field_simps exp_minus power2_eq_square square_eq_1_iff) lemma cosh_zero_iff: "cosh x = 0 \ exp x ^ 2 = -1" by (auto simp: cosh_def exp_minus field_simps power2_eq_square eq_neg_iff_add_eq_0) lemma cosh_square_eq: "cosh x ^ 2 = sinh x ^ 2 + 1" by (simp add: cosh_def sinh_def algebra_simps power2_eq_square exp_add [symmetric] scaleR_conv_of_real) lemma sinh_square_eq: "sinh x ^ 2 = cosh x ^ 2 - 1" by (simp add: cosh_square_eq) lemma hyperbolic_pythagoras: "cosh x ^ 2 - sinh x ^ 2 = 1" by (simp add: cosh_square_eq) lemma sinh_add: "sinh (x + y) = sinh x * cosh y + cosh x * sinh y" by (simp add: sinh_def cosh_def algebra_simps scaleR_conv_of_real exp_add [symmetric]) lemma sinh_diff: "sinh (x - y) = sinh x * cosh y - cosh x * sinh y" by (simp add: sinh_def cosh_def algebra_simps scaleR_conv_of_real exp_add [symmetric]) lemma cosh_add: "cosh (x + y) = cosh x * cosh y + sinh x * sinh y" by (simp add: sinh_def cosh_def algebra_simps scaleR_conv_of_real exp_add [symmetric]) lemma cosh_diff: "cosh (x - y) = cosh x * cosh y - sinh x * sinh y" by (simp add: sinh_def cosh_def algebra_simps scaleR_conv_of_real exp_add [symmetric]) lemma tanh_add: "tanh (x + y) = (tanh x + tanh y) / (1 + tanh x * tanh y)" if "cosh x \ 0" "cosh y \ 0" proof - have "(sinh x * cosh y + cosh x * sinh y) * (1 + sinh x * sinh y / (cosh x * cosh y)) = (cosh x * cosh y + sinh x * sinh y) * ((sinh x * cosh y + sinh y * cosh x) / (cosh y * cosh x))" using that by (simp add: field_split_simps) also have "(sinh x * cosh y + sinh y * cosh x) / (cosh y * cosh x) = sinh x / cosh x + sinh y / cosh y" using that by (simp add: field_split_simps) finally have "(sinh x * cosh y + cosh x * sinh y) * (1 + sinh x * sinh y / (cosh x * cosh y)) = (sinh x / cosh x + sinh y / cosh y) * (cosh x * cosh y + sinh x * sinh y)" by simp then show ?thesis using that by (auto simp add: tanh_def sinh_add cosh_add eq_divide_eq) (simp_all add: field_split_simps) qed lemma sinh_double: "sinh (2 * x) = 2 * sinh x * cosh x" using sinh_add[of x] by simp lemma cosh_double: "cosh (2 * x) = cosh x ^ 2 + sinh x ^ 2" using cosh_add[of x] by (simp add: power2_eq_square) end lemma sinh_field_def: "sinh z = (exp z - exp (-z)) / (2 :: 'a :: {banach, real_normed_field})" by (simp add: sinh_def scaleR_conv_of_real) lemma cosh_field_def: "cosh z = (exp z + exp (-z)) / (2 :: 'a :: {banach, real_normed_field})" by (simp add: cosh_def scaleR_conv_of_real) subsubsection \More specific properties of the real functions\ lemma plus_inverse_ge_2: fixes x :: real assumes "x > 0" shows "x + inverse x \ 2" proof - have "0 \ (x - 1) ^ 2" by simp also have "\ = x^2 - 2*x + 1" by (simp add: power2_eq_square algebra_simps) finally show ?thesis using assms by (simp add: field_simps power2_eq_square) qed lemma sinh_real_nonneg_iff [simp]: "sinh (x :: real) \ 0 \ x \ 0" by (simp add: sinh_def) lemma sinh_real_pos_iff [simp]: "sinh (x :: real) > 0 \ x > 0" by (simp add: sinh_def) lemma sinh_real_nonpos_iff [simp]: "sinh (x :: real) \ 0 \ x \ 0" by (simp add: sinh_def) lemma sinh_real_neg_iff [simp]: "sinh (x :: real) < 0 \ x < 0" by (simp add: sinh_def) lemma cosh_real_ge_1: "cosh (x :: real) \ 1" using plus_inverse_ge_2[of "exp x"] by (simp add: cosh_def exp_minus) lemma cosh_real_pos [simp]: "cosh (x :: real) > 0" using cosh_real_ge_1[of x] by simp lemma cosh_real_nonneg[simp]: "cosh (x :: real) \ 0" using cosh_real_ge_1[of x] by simp lemma cosh_real_nonzero [simp]: "cosh (x :: real) \ 0" using cosh_real_ge_1[of x] by simp lemma arsinh_real_def: "arsinh (x::real) = ln (x + sqrt (x^2 + 1))" by (simp add: arsinh_def powr_half_sqrt) lemma arcosh_real_def: "x \ 1 \ arcosh (x::real) = ln (x + sqrt (x^2 - 1))" by (simp add: arcosh_def powr_half_sqrt) lemma arsinh_real_aux: "0 < x + sqrt (x ^ 2 + 1 :: real)" proof (cases "x < 0") case True have "(-x) ^ 2 = x ^ 2" by simp also have "x ^ 2 < x ^ 2 + 1" by simp finally have "sqrt ((-x) ^ 2) < sqrt (x ^ 2 + 1)" by (rule real_sqrt_less_mono) thus ?thesis using True by simp qed (auto simp: add_nonneg_pos) lemma arsinh_minus_real [simp]: "arsinh (-x::real) = -arsinh x" proof - have "arsinh (-x) = ln (sqrt (x\<^sup>2 + 1) - x)" by (simp add: arsinh_real_def) also have "sqrt (x^2 + 1) - x = inverse (sqrt (x^2 + 1) + x)" using arsinh_real_aux[of x] by (simp add: field_split_simps algebra_simps power2_eq_square) also have "ln \ = -arsinh x" using arsinh_real_aux[of x] by (simp add: arsinh_real_def ln_inverse) finally show ?thesis . qed lemma artanh_minus_real [simp]: assumes "abs x < 1" shows "artanh (-x::real) = -artanh x" using assms by (simp add: artanh_def ln_div field_simps) lemma sinh_less_cosh_real: "sinh (x :: real) < cosh x" by (simp add: sinh_def cosh_def) lemma sinh_le_cosh_real: "sinh (x :: real) \ cosh x" by (simp add: sinh_def cosh_def) lemma tanh_real_lt_1: "tanh (x :: real) < 1" by (simp add: tanh_def sinh_less_cosh_real) lemma tanh_real_gt_neg1: "tanh (x :: real) > -1" proof - have "- cosh x < sinh x" by (simp add: sinh_def cosh_def field_split_simps) thus ?thesis by (simp add: tanh_def field_simps) qed lemma tanh_real_bounds: "tanh (x :: real) \ {-1<..<1}" using tanh_real_lt_1 tanh_real_gt_neg1 by simp context fixes x :: real begin lemma arsinh_sinh_real: "arsinh (sinh x) = x" by (simp add: arsinh_real_def powr_def sinh_square_eq sinh_plus_cosh) lemma arcosh_cosh_real: "x \ 0 \ arcosh (cosh x) = x" by (simp add: arcosh_real_def powr_def cosh_square_eq cosh_real_ge_1 cosh_plus_sinh) lemma artanh_tanh_real: "artanh (tanh x) = x" proof - have "artanh (tanh x) = ln (cosh x * (cosh x + sinh x) / (cosh x * (cosh x - sinh x))) / 2" by (simp add: artanh_def tanh_def field_split_simps) also have "cosh x * (cosh x + sinh x) / (cosh x * (cosh x - sinh x)) = (cosh x + sinh x) / (cosh x - sinh x)" by simp also have "\ = (exp x)^2" by (simp add: cosh_plus_sinh cosh_minus_sinh exp_minus field_simps power2_eq_square) also have "ln ((exp x)^2) / 2 = x" by (simp add: ln_realpow) finally show ?thesis . qed lemma sinh_real_zero_iff [simp]: "sinh x = 0 \ x = 0" by (metis arsinh_0 arsinh_sinh_real sinh_0) lemma cosh_real_one_iff [simp]: "cosh x = 1 \ x = 0" by (smt (verit, best) Transcendental.arcosh_cosh_real cosh_0 cosh_minus) lemma tanh_real_nonneg_iff [simp]: "tanh x \ 0 \ x \ 0" by (simp add: tanh_def field_simps) lemma tanh_real_pos_iff [simp]: "tanh x > 0 \ x > 0" by (simp add: tanh_def field_simps) lemma tanh_real_nonpos_iff [simp]: "tanh x \ 0 \ x \ 0" by (simp add: tanh_def field_simps) lemma tanh_real_neg_iff [simp]: "tanh x < 0 \ x < 0" by (simp add: tanh_def field_simps) lemma tanh_real_zero_iff [simp]: "tanh x = 0 \ x = 0" by (simp add: tanh_def field_simps) end lemma sinh_real_strict_mono: "strict_mono (sinh :: real \ real)" by (rule pos_deriv_imp_strict_mono derivative_intros)+ auto lemma cosh_real_strict_mono: assumes "0 \ x" and "x < (y::real)" shows "cosh x < cosh y" proof - from assms have "\z>x. z < y \ cosh y - cosh x = (y - x) * sinh z" by (intro MVT2) (auto dest: connectedD_interval intro!: derivative_eq_intros) then obtain z where z: "z > x" "z < y" "cosh y - cosh x = (y - x) * sinh z" by blast note \cosh y - cosh x = (y - x) * sinh z\ also from \z > x\ and assms have "(y - x) * sinh z > 0" by (intro mult_pos_pos) auto finally show "cosh x < cosh y" by simp qed lemma tanh_real_strict_mono: "strict_mono (tanh :: real \ real)" proof - have *: "tanh x ^ 2 < 1" for x :: real using tanh_real_bounds[of x] by (simp add: abs_square_less_1 abs_if) show ?thesis by (rule pos_deriv_imp_strict_mono) (insert *, auto intro!: derivative_intros) qed lemma sinh_real_abs [simp]: "sinh (abs x :: real) = abs (sinh x)" by (simp add: abs_if) lemma cosh_real_abs [simp]: "cosh (abs x :: real) = cosh x" by (simp add: abs_if) lemma tanh_real_abs [simp]: "tanh (abs x :: real) = abs (tanh x)" by (auto simp: abs_if) lemma sinh_real_eq_iff [simp]: "sinh x = sinh y \ x = (y :: real)" using sinh_real_strict_mono by (simp add: strict_mono_eq) lemma tanh_real_eq_iff [simp]: "tanh x = tanh y \ x = (y :: real)" using tanh_real_strict_mono by (simp add: strict_mono_eq) lemma cosh_real_eq_iff [simp]: "cosh x = cosh y \ abs x = abs (y :: real)" proof - have "cosh x = cosh y \ x = y" if "x \ 0" "y \ 0" for x y :: real using cosh_real_strict_mono[of x y] cosh_real_strict_mono[of y x] that by (cases x y rule: linorder_cases) auto from this[of "abs x" "abs y"] show ?thesis by simp qed lemma sinh_real_le_iff [simp]: "sinh x \ sinh y \ x \ (y::real)" using sinh_real_strict_mono by (simp add: strict_mono_less_eq) lemma cosh_real_nonneg_le_iff: "x \ 0 \ y \ 0 \ cosh x \ cosh y \ x \ (y::real)" using cosh_real_strict_mono[of x y] cosh_real_strict_mono[of y x] by (cases x y rule: linorder_cases) auto lemma cosh_real_nonpos_le_iff: "x \ 0 \ y \ 0 \ cosh x \ cosh y \ x \ (y::real)" using cosh_real_nonneg_le_iff[of "-x" "-y"] by simp lemma tanh_real_le_iff [simp]: "tanh x \ tanh y \ x \ (y::real)" using tanh_real_strict_mono by (simp add: strict_mono_less_eq) lemma sinh_real_less_iff [simp]: "sinh x < sinh y \ x < (y::real)" using sinh_real_strict_mono by (simp add: strict_mono_less) lemma cosh_real_nonneg_less_iff: "x \ 0 \ y \ 0 \ cosh x < cosh y \ x < (y::real)" using cosh_real_strict_mono[of x y] cosh_real_strict_mono[of y x] by (cases x y rule: linorder_cases) auto lemma cosh_real_nonpos_less_iff: "x \ 0 \ y \ 0 \ cosh x < cosh y \ x > (y::real)" using cosh_real_nonneg_less_iff[of "-x" "-y"] by simp lemma tanh_real_less_iff [simp]: "tanh x < tanh y \ x < (y::real)" using tanh_real_strict_mono by (simp add: strict_mono_less) subsubsection \Limits\ lemma sinh_real_at_top: "filterlim (sinh :: real \ real) at_top at_top" proof - have *: "((\x. - exp (- x)) \ (-0::real)) at_top" by (intro tendsto_minus filterlim_compose[OF exp_at_bot] filterlim_uminus_at_bot_at_top) have "filterlim (\x. (1/2) * (-exp (-x) + exp x) :: real) at_top at_top" by (rule filterlim_tendsto_pos_mult_at_top[OF _ _ filterlim_tendsto_add_at_top[OF *]] tendsto_const)+ (auto simp: exp_at_top) also have "(\x. (1/2) * (-exp (-x) + exp x) :: real) = sinh" by (simp add: fun_eq_iff sinh_def) finally show ?thesis . qed lemma sinh_real_at_bot: "filterlim (sinh :: real \ real) at_bot at_bot" proof - have "filterlim (\x. -sinh x :: real) at_bot at_top" by (simp add: filterlim_uminus_at_top [symmetric] sinh_real_at_top) also have "(\x. -sinh x :: real) = (\x. sinh (-x))" by simp finally show ?thesis by (subst filterlim_at_bot_mirror) qed lemma cosh_real_at_top: "filterlim (cosh :: real \ real) at_top at_top" proof - have *: "((\x. exp (- x)) \ (0::real)) at_top" by (intro filterlim_compose[OF exp_at_bot] filterlim_uminus_at_bot_at_top) have "filterlim (\x. (1/2) * (exp (-x) + exp x) :: real) at_top at_top" by (rule filterlim_tendsto_pos_mult_at_top[OF _ _ filterlim_tendsto_add_at_top[OF *]] tendsto_const)+ (auto simp: exp_at_top) also have "(\x. (1/2) * (exp (-x) + exp x) :: real) = cosh" by (simp add: fun_eq_iff cosh_def) finally show ?thesis . qed lemma cosh_real_at_bot: "filterlim (cosh :: real \ real) at_top at_bot" proof - have "filterlim (\x. cosh (-x) :: real) at_top at_top" by (simp add: cosh_real_at_top) thus ?thesis by (subst filterlim_at_bot_mirror) qed lemma tanh_real_at_top: "(tanh \ (1::real)) at_top" proof - have "((\x::real. (1 - exp (- 2 * x)) / (1 + exp (- 2 * x))) \ (1 - 0) / (1 + 0)) at_top" by (intro tendsto_intros filterlim_compose[OF exp_at_bot] filterlim_tendsto_neg_mult_at_bot[OF tendsto_const] filterlim_ident) auto also have "(\x::real. (1 - exp (- 2 * x)) / (1 + exp (- 2 * x))) = tanh" by (rule ext) (simp add: tanh_real_altdef) finally show ?thesis by simp qed lemma tanh_real_at_bot: "(tanh \ (-1::real)) at_bot" proof - have "((\x::real. -tanh x) \ -1) at_top" by (intro tendsto_minus tanh_real_at_top) also have "(\x. -tanh x :: real) = (\x. tanh (-x))" by simp finally show ?thesis by (subst filterlim_at_bot_mirror) qed subsubsection \Properties of the inverse hyperbolic functions\ lemma isCont_sinh: "isCont sinh (x :: 'a :: {real_normed_field, banach})" unfolding sinh_def [abs_def] by (auto intro!: continuous_intros) lemma isCont_cosh: "isCont cosh (x :: 'a :: {real_normed_field, banach})" unfolding cosh_def [abs_def] by (auto intro!: continuous_intros) lemma isCont_tanh: "cosh x \ 0 \ isCont tanh (x :: 'a :: {real_normed_field, banach})" unfolding tanh_def [abs_def] by (auto intro!: continuous_intros isCont_divide isCont_sinh isCont_cosh) lemma continuous_on_sinh [continuous_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" assumes "continuous_on A f" shows "continuous_on A (\x. sinh (f x))" unfolding sinh_def using assms by (intro continuous_intros) lemma continuous_on_cosh [continuous_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" assumes "continuous_on A f" shows "continuous_on A (\x. cosh (f x))" unfolding cosh_def using assms by (intro continuous_intros) lemma continuous_sinh [continuous_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" assumes "continuous F f" shows "continuous F (\x. sinh (f x))" unfolding sinh_def using assms by (intro continuous_intros) lemma continuous_cosh [continuous_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" assumes "continuous F f" shows "continuous F (\x. cosh (f x))" unfolding cosh_def using assms by (intro continuous_intros) lemma continuous_on_tanh [continuous_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" assumes "continuous_on A f" "\x. x \ A \ cosh (f x) \ 0" shows "continuous_on A (\x. tanh (f x))" unfolding tanh_def using assms by (intro continuous_intros) auto lemma continuous_at_within_tanh [continuous_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" assumes "continuous (at x within A) f" "cosh (f x) \ 0" shows "continuous (at x within A) (\x. tanh (f x))" unfolding tanh_def using assms by (intro continuous_intros continuous_divide) auto lemma continuous_tanh [continuous_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" assumes "continuous F f" "cosh (f (Lim F (\x. x))) \ 0" shows "continuous F (\x. tanh (f x))" unfolding tanh_def using assms by (intro continuous_intros continuous_divide) auto lemma tendsto_sinh [tendsto_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" shows "(f \ a) F \ ((\x. sinh (f x)) \ sinh a) F" by (rule isCont_tendsto_compose [OF isCont_sinh]) lemma tendsto_cosh [tendsto_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" shows "(f \ a) F \ ((\x. cosh (f x)) \ cosh a) F" by (rule isCont_tendsto_compose [OF isCont_cosh]) lemma tendsto_tanh [tendsto_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" shows "(f \ a) F \ cosh a \ 0 \ ((\x. tanh (f x)) \ tanh a) F" by (rule isCont_tendsto_compose [OF isCont_tanh]) lemma arsinh_real_has_field_derivative [derivative_intros]: fixes x :: real shows "(arsinh has_field_derivative (1 / (sqrt (x ^ 2 + 1)))) (at x within A)" proof - have pos: "1 + x ^ 2 > 0" by (intro add_pos_nonneg) auto from pos arsinh_real_aux[of x] show ?thesis unfolding arsinh_def [abs_def] by (auto intro!: derivative_eq_intros simp: powr_minus powr_half_sqrt field_split_simps) qed lemma arcosh_real_has_field_derivative [derivative_intros]: fixes x :: real assumes "x > 1" shows "(arcosh has_field_derivative (1 / (sqrt (x ^ 2 - 1)))) (at x within A)" proof - from assms have "x + sqrt (x\<^sup>2 - 1) > 0" by (simp add: add_pos_pos) thus ?thesis using assms unfolding arcosh_def [abs_def] by (auto intro!: derivative_eq_intros simp: powr_minus powr_half_sqrt field_split_simps power2_eq_1_iff) qed lemma artanh_real_has_field_derivative [derivative_intros]: "(artanh has_field_derivative (1 / (1 - x ^ 2))) (at x within A)" if "\x\ < 1" for x :: real proof - from that have "- 1 < x" "x < 1" by linarith+ hence "(artanh has_field_derivative (4 - 4 * x) / ((1 + x) * (1 - x) * (1 - x) * 4)) (at x within A)" unfolding artanh_def [abs_def] by (auto intro!: derivative_eq_intros simp: powr_minus powr_half_sqrt) also have "(4 - 4 * x) / ((1 + x) * (1 - x) * (1 - x) * 4) = 1 / ((1 + x) * (1 - x))" using \-1 < x\ \x < 1\ by (simp add: frac_eq_eq) also have "(1 + x) * (1 - x) = 1 - x ^ 2" by (simp add: algebra_simps power2_eq_square) finally show ?thesis . qed lemma cosh_double_cosh: "cosh (2 * x :: 'a :: {banach, real_normed_field}) = 2 * (cosh x)\<^sup>2 - 1" using cosh_double[of x] by (simp add: sinh_square_eq) lemma sinh_multiple_reduce: "sinh (x * numeral n :: 'a :: {real_normed_field, banach}) = sinh x * cosh (x * of_nat (pred_numeral n)) + cosh x * sinh (x * of_nat (pred_numeral n))" proof - have "numeral n = of_nat (pred_numeral n) + (1 :: 'a)" by (metis add.commute numeral_eq_Suc of_nat_Suc of_nat_numeral) also have "sinh (x * \) = sinh (x * of_nat (pred_numeral n) + x)" unfolding of_nat_Suc by (simp add: ring_distribs) finally show ?thesis by (simp add: sinh_add) qed lemma cosh_multiple_reduce: "cosh (x * numeral n :: 'a :: {real_normed_field, banach}) = cosh (x * of_nat (pred_numeral n)) * cosh x + sinh (x * of_nat (pred_numeral n)) * sinh x" proof - have "numeral n = of_nat (pred_numeral n) + (1 :: 'a)" by (metis add.commute numeral_eq_Suc of_nat_Suc of_nat_numeral) also have "cosh (x * \) = cosh (x * of_nat (pred_numeral n) + x)" unfolding of_nat_Suc by (simp add: ring_distribs) finally show ?thesis by (simp add: cosh_add) qed lemma cosh_arcosh_real [simp]: assumes "x \ (1 :: real)" shows "cosh (arcosh x) = x" proof - have "eventually (\t::real. cosh t \ x) at_top" using cosh_real_at_top by (simp add: filterlim_at_top) then obtain t where "t \ 1" "cosh t \ x" by (metis eventually_at_top_linorder linorder_not_le order_le_less) moreover have "isCont cosh (y :: real)" for y by (intro continuous_intros) ultimately obtain y where "y \ 0" "x = cosh y" using IVT[of cosh 0 x t] assms by auto thus ?thesis by (simp add: arcosh_cosh_real) qed lemma arcosh_eq_0_iff_real [simp]: "x \ 1 \ arcosh x = 0 \ x = (1 :: real)" using cosh_arcosh_real by fastforce lemma arcosh_nonneg_real [simp]: assumes "x \ 1" shows "arcosh (x :: real) \ 0" proof - have "1 + 0 \ x + (x\<^sup>2 - 1) powr (1 / 2)" using assms by (intro add_mono) auto thus ?thesis unfolding arcosh_def by simp qed lemma arcosh_real_strict_mono: fixes x y :: real assumes "1 \ x" "x < y" shows "arcosh x < arcosh y" proof - have "cosh (arcosh x) < cosh (arcosh y)" by (subst (1 2) cosh_arcosh_real) (use assms in auto) thus ?thesis using assms by (subst (asm) cosh_real_nonneg_less_iff) auto qed lemma arcosh_less_iff_real [simp]: fixes x y :: real assumes "1 \ x" "1 \ y" shows "arcosh x < arcosh y \ x < y" using arcosh_real_strict_mono[of x y] arcosh_real_strict_mono[of y x] assms by (cases x y rule: linorder_cases) auto lemma arcosh_real_gt_1_iff [simp]: "x \ 1 \ arcosh x > 0 \ x \ (1 :: real)" using arcosh_less_iff_real[of 1 x] by (auto simp del: arcosh_less_iff_real) lemma sinh_arcosh_real: "x \ 1 \ sinh (arcosh x) = sqrt (x\<^sup>2 - 1)" by (rule sym, rule real_sqrt_unique) (auto simp: sinh_square_eq) lemma sinh_arsinh_real [simp]: "sinh (arsinh x :: real) = x" proof - have "eventually (\t::real. sinh t \ x) at_top" using sinh_real_at_top by (simp add: filterlim_at_top) then obtain t where "sinh t \ x" by (metis eventually_at_top_linorder linorder_not_le order_le_less) moreover have "eventually (\t::real. sinh t \ x) at_bot" using sinh_real_at_bot by (simp add: filterlim_at_bot) then obtain t' where "t' \ t" "sinh t' \ x" by (metis eventually_at_bot_linorder nle_le) moreover have "isCont sinh (y :: real)" for y by (intro continuous_intros) ultimately obtain y where "x = sinh y" using IVT[of sinh t' x t] by auto thus ?thesis by (simp add: arsinh_sinh_real) qed lemma arsinh_real_strict_mono: fixes x y :: real assumes "x < y" shows "arsinh x < arsinh y" proof - have "sinh (arsinh x) < sinh (arsinh y)" by (subst (1 2) sinh_arsinh_real) (use assms in auto) thus ?thesis using assms by (subst (asm) sinh_real_less_iff) auto qed lemma arsinh_less_iff_real [simp]: fixes x y :: real shows "arsinh x < arsinh y \ x < y" using arsinh_real_strict_mono[of x y] arsinh_real_strict_mono[of y x] by (cases x y rule: linorder_cases) auto lemma arsinh_real_eq_0_iff [simp]: "arsinh x = 0 \ x = (0 :: real)" by (metis arsinh_0 sinh_arsinh_real) lemma arsinh_real_pos_iff [simp]: "arsinh x > 0 \ x > (0 :: real)" using arsinh_less_iff_real[of 0 x] by (simp del: arsinh_less_iff_real) lemma arsinh_real_neg_iff [simp]: "arsinh x < 0 \ x < (0 :: real)" using arsinh_less_iff_real[of x 0] by (simp del: arsinh_less_iff_real) lemma cosh_arsinh_real: "cosh (arsinh x) = sqrt (x\<^sup>2 + 1)" by (rule sym, rule real_sqrt_unique) (auto simp: cosh_square_eq) lemma continuous_on_arsinh [continuous_intros]: "continuous_on A (arsinh :: real \ real)" by (rule DERIV_continuous_on derivative_intros)+ lemma continuous_on_arcosh [continuous_intros]: assumes "A \ {1..}" shows "continuous_on A (arcosh :: real \ real)" proof - have pos: "x + sqrt (x ^ 2 - 1) > 0" if "x \ 1" for x using that by (intro add_pos_nonneg) auto show ?thesis unfolding arcosh_def [abs_def] by (intro continuous_on_subset [OF _ assms] continuous_on_ln continuous_on_add continuous_on_id continuous_on_powr') (auto dest: pos simp: powr_half_sqrt intro!: continuous_intros) qed lemma continuous_on_artanh [continuous_intros]: assumes "A \ {-1<..<1}" shows "continuous_on A (artanh :: real \ real)" unfolding artanh_def [abs_def] by (intro continuous_on_subset [OF _ assms]) (auto intro!: continuous_intros) lemma continuous_on_arsinh' [continuous_intros]: fixes f :: "real \ real" assumes "continuous_on A f" shows "continuous_on A (\x. arsinh (f x))" by (rule continuous_on_compose2[OF continuous_on_arsinh assms]) auto lemma continuous_on_arcosh' [continuous_intros]: fixes f :: "real \ real" assumes "continuous_on A f" "\x. x \ A \ f x \ 1" shows "continuous_on A (\x. arcosh (f x))" by (rule continuous_on_compose2[OF continuous_on_arcosh assms(1) order.refl]) (use assms(2) in auto) lemma continuous_on_artanh' [continuous_intros]: fixes f :: "real \ real" assumes "continuous_on A f" "\x. x \ A \ f x \ {-1<..<1}" shows "continuous_on A (\x. artanh (f x))" by (rule continuous_on_compose2[OF continuous_on_artanh assms(1) order.refl]) (use assms(2) in auto) lemma isCont_arsinh [continuous_intros]: "isCont arsinh (x :: real)" using continuous_on_arsinh[of UNIV] by (auto simp: continuous_on_eq_continuous_at) lemma isCont_arcosh [continuous_intros]: assumes "x > 1" shows "isCont arcosh (x :: real)" proof - have "continuous_on {1::real<..} arcosh" by (rule continuous_on_arcosh) auto with assms show ?thesis by (auto simp: continuous_on_eq_continuous_at) qed lemma isCont_artanh [continuous_intros]: assumes "x > -1" "x < 1" shows "isCont artanh (x :: real)" proof - have "continuous_on {-1<..<(1::real)} artanh" by (rule continuous_on_artanh) auto with assms show ?thesis by (auto simp: continuous_on_eq_continuous_at) qed lemma tendsto_arsinh [tendsto_intros]: "(f \ a) F \ ((\x. arsinh (f x)) \ arsinh a) F" for f :: "_ \ real" by (rule isCont_tendsto_compose [OF isCont_arsinh]) lemma tendsto_arcosh_strong [tendsto_intros]: fixes f :: "_ \ real" assumes "(f \ a) F" "a \ 1" "eventually (\x. f x \ 1) F" shows "((\x. arcosh (f x)) \ arcosh a) F" by (rule continuous_on_tendsto_compose[OF continuous_on_arcosh[OF order.refl]]) (use assms in auto) lemma tendsto_arcosh: fixes f :: "_ \ real" assumes "(f \ a) F" "a > 1" shows "((\x. arcosh (f x)) \ arcosh a) F" by (rule isCont_tendsto_compose [OF isCont_arcosh]) (use assms in auto) lemma tendsto_arcosh_at_left_1: "(arcosh \ 0) (at_right (1::real))" proof - have "(arcosh \ arcosh 1) (at_right (1::real))" by (rule tendsto_arcosh_strong) (auto simp: eventually_at intro!: exI[of _ 1]) thus ?thesis by simp qed lemma tendsto_artanh [tendsto_intros]: fixes f :: "'a \ real" assumes "(f \ a) F" "a > -1" "a < 1" shows "((\x. artanh (f x)) \ artanh a) F" by (rule isCont_tendsto_compose [OF isCont_artanh]) (use assms in auto) lemma continuous_arsinh [continuous_intros]: "continuous F f \ continuous F (\x. arsinh (f x :: real))" unfolding continuous_def by (rule tendsto_arsinh) (* TODO: This rule does not work for one-sided continuity at 1 *) lemma continuous_arcosh_strong [continuous_intros]: assumes "continuous F f" "eventually (\x. f x \ 1) F" shows "continuous F (\x. arcosh (f x :: real))" proof (cases "F = bot") case False show ?thesis unfolding continuous_def proof (intro tendsto_arcosh_strong) show "1 \ f (Lim F (\x. x))" using assms False unfolding continuous_def by (rule tendsto_lowerbound) qed (insert assms, auto simp: continuous_def) qed auto lemma continuous_arcosh: "continuous F f \ f (Lim F (\x. x)) > 1 \ continuous F (\x. arcosh (f x :: real))" unfolding continuous_def by (rule tendsto_arcosh) auto lemma continuous_artanh [continuous_intros]: "continuous F f \ f (Lim F (\x. x)) \ {-1<..<1} \ continuous F (\x. artanh (f x :: real))" unfolding continuous_def by (rule tendsto_artanh) auto lemma arsinh_real_at_top: "filterlim (arsinh :: real \ real) at_top at_top" proof (subst filterlim_cong[OF refl refl]) show "filterlim (\x. ln (x + sqrt (1 + x\<^sup>2))) at_top at_top" by (intro filterlim_compose[OF ln_at_top filterlim_at_top_add_at_top] filterlim_ident filterlim_compose[OF sqrt_at_top] filterlim_tendsto_add_at_top[OF tendsto_const] filterlim_pow_at_top) auto qed (auto intro!: eventually_mono[OF eventually_ge_at_top[of 1]] simp: arsinh_real_def add_ac) lemma arsinh_real_at_bot: "filterlim (arsinh :: real \ real) at_bot at_bot" proof - have "filterlim (\x::real. -arsinh x) at_bot at_top" by (subst filterlim_uminus_at_top [symmetric]) (rule arsinh_real_at_top) also have "(\x::real. -arsinh x) = (\x. arsinh (-x))" by simp finally show ?thesis by (subst filterlim_at_bot_mirror) qed lemma arcosh_real_at_top: "filterlim (arcosh :: real \ real) at_top at_top" proof (subst filterlim_cong[OF refl refl]) show "filterlim (\x. ln (x + sqrt (-1 + x\<^sup>2))) at_top at_top" by (intro filterlim_compose[OF ln_at_top filterlim_at_top_add_at_top] filterlim_ident filterlim_compose[OF sqrt_at_top] filterlim_tendsto_add_at_top[OF tendsto_const] filterlim_pow_at_top) auto qed (auto intro!: eventually_mono[OF eventually_ge_at_top[of 1]] simp: arcosh_real_def) lemma artanh_real_at_left_1: "filterlim (artanh :: real \ real) at_top (at_left 1)" proof - have *: "filterlim (\x::real. (1 + x) / (1 - x)) at_top (at_left 1)" by (rule LIM_at_top_divide) (auto intro!: tendsto_eq_intros eventually_mono[OF eventually_at_left_real[of 0]]) have "filterlim (\x::real. (1/2) * ln ((1 + x) / (1 - x))) at_top (at_left 1)" by (intro filterlim_tendsto_pos_mult_at_top[OF tendsto_const] * filterlim_compose[OF ln_at_top]) auto also have "(\x::real. (1/2) * ln ((1 + x) / (1 - x))) = artanh" by (simp add: artanh_def [abs_def]) finally show ?thesis . qed lemma artanh_real_at_right_1: "filterlim (artanh :: real \ real) at_bot (at_right (-1))" proof - have "?thesis \ filterlim (\x::real. -artanh x) at_top (at_right (-1))" by (simp add: filterlim_uminus_at_bot) also have "\ \ filterlim (\x::real. artanh (-x)) at_top (at_right (-1))" by (intro filterlim_cong refl eventually_mono[OF eventually_at_right_real[of "-1" "1"]]) auto also have "\ \ filterlim (artanh :: real \ real) at_top (at_left 1)" by (simp add: filterlim_at_left_to_right) also have \ by (rule artanh_real_at_left_1) finally show ?thesis . qed subsection \Simprocs for root and power literals\ lemma numeral_powr_numeral_real [simp]: "numeral m powr numeral n = (numeral m ^ numeral n :: real)" by (simp add: powr_numeral) context begin private lemma sqrt_numeral_simproc_aux: assumes "m * m \ n" shows "sqrt (numeral n :: real) \ numeral m" proof - have "numeral n \ numeral m * (numeral m :: real)" by (simp add: assms [symmetric]) moreover have "sqrt \ \ numeral m" by (subst real_sqrt_abs2) simp ultimately show "sqrt (numeral n :: real) \ numeral m" by simp qed private lemma root_numeral_simproc_aux: assumes "Num.pow m n \ x" shows "root (numeral n) (numeral x :: real) \ numeral m" by (subst assms [symmetric], subst numeral_pow, subst real_root_pos2) simp_all private lemma powr_numeral_simproc_aux: assumes "Num.pow y n = x" shows "numeral x powr (m / numeral n :: real) \ numeral y powr m" by (subst assms [symmetric], subst numeral_pow, subst powr_numeral [symmetric]) (simp, subst powr_powr, simp_all) private lemma numeral_powr_inverse_eq: "numeral x powr (inverse (numeral n)) = numeral x powr (1 / numeral n :: real)" by simp ML \ signature ROOT_NUMERAL_SIMPROC = sig val sqrt : int option -> int -> int option val sqrt' : int option -> int -> int option val nth_root : int option -> int -> int -> int option val nth_root' : int option -> int -> int -> int option val sqrt_proc : Simplifier.proc val root_proc : int * int -> Simplifier.proc val powr_proc : int * int -> Simplifier.proc end structure Root_Numeral_Simproc : ROOT_NUMERAL_SIMPROC = struct fun iterate NONE p f x = let fun go x = if p x then x else go (f x) in SOME (go x) end | iterate (SOME threshold) p f x = let fun go (threshold, x) = if p x then SOME x else if threshold = 0 then NONE else go (threshold - 1, f x) in go (threshold, x) end fun nth_root _ 1 x = SOME x | nth_root _ _ 0 = SOME 0 | nth_root _ _ 1 = SOME 1 | nth_root threshold n x = let fun newton_step y = ((n - 1) * y + x div Integer.pow (n - 1) y) div n fun is_root y = Integer.pow n y <= x andalso x < Integer.pow n (y + 1) in if x < n then SOME 1 else if x < Integer.pow n 2 then SOME 1 else let val y = Real.floor (Math.pow (Real.fromInt x, Real.fromInt 1 / Real.fromInt n)) in if is_root y then SOME y else iterate threshold is_root newton_step ((x + n - 1) div n) end end fun nth_root' _ 1 x = SOME x | nth_root' _ _ 0 = SOME 0 | nth_root' _ _ 1 = SOME 1 | nth_root' threshold n x = if x < n then NONE else if x < Integer.pow n 2 then NONE else case nth_root threshold n x of NONE => NONE | SOME y => if Integer.pow n y = x then SOME y else NONE fun sqrt _ 0 = SOME 0 | sqrt _ 1 = SOME 1 | sqrt threshold n = let fun aux (a, b) = if n >= b * b then aux (b, b * b) else (a, b) val (lower_root, lower_n) = aux (1, 2) fun newton_step x = (x + n div x) div 2 fun is_sqrt r = r*r <= n andalso n < (r+1)*(r+1) val y = Real.floor (Math.sqrt (Real.fromInt n)) in if is_sqrt y then SOME y else Option.mapPartial (iterate threshold is_sqrt newton_step o (fn x => x * lower_root)) (sqrt threshold (n div lower_n)) end fun sqrt' threshold x = case sqrt threshold x of NONE => NONE | SOME y => if y * y = x then SOME y else NONE fun sqrt_proc ctxt ct = let val n = ct |> Thm.term_of |> dest_comb |> snd |> dest_comb |> snd |> HOLogic.dest_numeral in case sqrt' (SOME 10000) n of NONE => NONE | SOME m => SOME (Thm.instantiate' [] (map (SOME o Thm.cterm_of ctxt o HOLogic.mk_numeral) [m, n]) @{thm sqrt_numeral_simproc_aux}) end handle TERM _ => NONE fun root_proc (threshold1, threshold2) ctxt ct = let val [n, x] = ct |> Thm.term_of |> strip_comb |> snd |> map (dest_comb #> snd #> HOLogic.dest_numeral) in if n > threshold1 orelse x > threshold2 then NONE else case nth_root' (SOME 100) n x of NONE => NONE | SOME m => SOME (Thm.instantiate' [] (map (SOME o Thm.cterm_of ctxt o HOLogic.mk_numeral) [m, n, x]) @{thm root_numeral_simproc_aux}) end handle TERM _ => NONE | Match => NONE fun powr_proc (threshold1, threshold2) ctxt ct = let val eq_thm = Conv.try_conv (Conv.rewr_conv @{thm numeral_powr_inverse_eq}) ct val ct = Thm.dest_equals_rhs (Thm.cprop_of eq_thm) val (_, [x, t]) = strip_comb (Thm.term_of ct) val (_, [m, n]) = strip_comb t val [x, n] = map (dest_comb #> snd #> HOLogic.dest_numeral) [x, n] in if n > threshold1 orelse x > threshold2 then NONE else case nth_root' (SOME 100) n x of NONE => NONE | SOME y => let val [y, n, x] = map HOLogic.mk_numeral [y, n, x] val thm = Thm.instantiate' [] (map (SOME o Thm.cterm_of ctxt) [y, n, x, m]) @{thm powr_numeral_simproc_aux} in SOME (@{thm transitive} OF [eq_thm, thm]) end end handle TERM _ => NONE | Match => NONE end \ end simproc_setup sqrt_numeral ("sqrt (numeral n)") = \K Root_Numeral_Simproc.sqrt_proc\ simproc_setup root_numeral ("root (numeral n) (numeral x)") = \K (Root_Numeral_Simproc.root_proc (200, Integer.pow 200 2))\ simproc_setup powr_divide_numeral ("numeral x powr (m / numeral n :: real)" | "numeral x powr (inverse (numeral n) :: real)") = \K (Root_Numeral_Simproc.powr_proc (200, Integer.pow 200 2))\ lemma "root 100 1267650600228229401496703205376 = 2" by simp lemma "sqrt 196 = 14" by simp lemma "256 powr (7 / 4 :: real) = 16384" by simp lemma "27 powr (inverse 3) = (3::real)" by simp end