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,4128 +1,4146 @@ (* Title: HOL/Analysis/Path_Connected.thy Authors: LC Paulson and Robert Himmelmann (TU Muenchen), based on material from HOL Light *) section \Path-Connectedness\ theory Path_Connected imports Starlike T1_Spaces begin subsection \Paths and Arcs\ definition\<^marker>\tag important\ path :: "(real \ 'a::topological_space) \ bool" where "path g \ continuous_on {0..1} g" definition\<^marker>\tag important\ pathstart :: "(real \ 'a::topological_space) \ 'a" where "pathstart g = g 0" definition\<^marker>\tag important\ pathfinish :: "(real \ 'a::topological_space) \ 'a" where "pathfinish g = g 1" definition\<^marker>\tag important\ path_image :: "(real \ 'a::topological_space) \ 'a set" where "path_image g = g ` {0 .. 1}" definition\<^marker>\tag important\ reversepath :: "(real \ 'a::topological_space) \ real \ 'a" where "reversepath g = (\x. g(1 - x))" definition\<^marker>\tag important\ joinpaths :: "(real \ 'a::topological_space) \ (real \ 'a) \ real \ 'a" (infixr "+++" 75) where "g1 +++ g2 = (\x. if x \ 1/2 then g1 (2 * x) else g2 (2 * x - 1))" definition\<^marker>\tag important\ simple_path :: "(real \ 'a::topological_space) \ bool" where "simple_path g \ path g \ (\x\{0..1}. \y\{0..1}. g x = g y \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0)" definition\<^marker>\tag important\ arc :: "(real \ 'a :: topological_space) \ bool" where "arc g \ path g \ inj_on g {0..1}" subsection\<^marker>\tag unimportant\\Invariance theorems\ lemma path_eq: "path p \ (\t. t \ {0..1} \ p t = q t) \ path q" using continuous_on_eq path_def by blast lemma path_continuous_image: "path g \ continuous_on (path_image g) f \ path(f \ g)" unfolding path_def path_image_def using continuous_on_compose by blast lemma continuous_on_translation_eq: fixes g :: "'a :: real_normed_vector \ 'b :: real_normed_vector" shows "continuous_on A ((+) a \ g) = continuous_on A g" proof - have g: "g = (\x. -a + x) \ ((\x. a + x) \ g)" by (rule ext) simp show ?thesis by (metis (no_types, opaque_lifting) g continuous_on_compose homeomorphism_def homeomorphism_translation) qed lemma path_translation_eq: fixes g :: "real \ 'a :: real_normed_vector" shows "path((\x. a + x) \ g) = path g" using continuous_on_translation_eq path_def by blast lemma path_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "path(f \ g) = path g" proof - from linear_injective_left_inverse [OF assms] obtain h where h: "linear h" "h \ f = id" by blast then have g: "g = h \ (f \ g)" by (metis comp_assoc id_comp) show ?thesis unfolding path_def using h assms by (metis g continuous_on_compose linear_continuous_on linear_conv_bounded_linear) qed lemma pathstart_translation: "pathstart((\x. a + x) \ g) = a + pathstart g" by (simp add: pathstart_def) lemma pathstart_linear_image_eq: "linear f \ pathstart(f \ g) = f(pathstart g)" by (simp add: pathstart_def) lemma pathfinish_translation: "pathfinish((\x. a + x) \ g) = a + pathfinish g" by (simp add: pathfinish_def) lemma pathfinish_linear_image: "linear f \ pathfinish(f \ g) = f(pathfinish g)" by (simp add: pathfinish_def) lemma path_image_translation: "path_image((\x. a + x) \ g) = (\x. a + x) ` (path_image g)" by (simp add: image_comp path_image_def) lemma path_image_linear_image: "linear f \ path_image(f \ g) = f ` (path_image g)" by (simp add: image_comp path_image_def) lemma reversepath_translation: "reversepath((\x. a + x) \ g) = (\x. a + x) \ reversepath g" by (rule ext) (simp add: reversepath_def) lemma reversepath_linear_image: "linear f \ reversepath(f \ g) = f \ reversepath g" by (rule ext) (simp add: reversepath_def) lemma joinpaths_translation: "((\x. a + x) \ g1) +++ ((\x. a + x) \ g2) = (\x. a + x) \ (g1 +++ g2)" by (rule ext) (simp add: joinpaths_def) lemma joinpaths_linear_image: "linear f \ (f \ g1) +++ (f \ g2) = f \ (g1 +++ g2)" by (rule ext) (simp add: joinpaths_def) lemma simple_path_translation_eq: fixes g :: "real \ 'a::euclidean_space" shows "simple_path((\x. a + x) \ g) = simple_path g" by (simp add: simple_path_def path_translation_eq) lemma simple_path_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "simple_path(f \ g) = simple_path g" using assms inj_on_eq_iff [of f] by (auto simp: path_linear_image_eq simple_path_def path_translation_eq) lemma arc_translation_eq: fixes g :: "real \ 'a::euclidean_space" shows "arc((\x. a + x) \ g) = arc g" by (auto simp: arc_def inj_on_def path_translation_eq) lemma arc_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "arc(f \ g) = arc g" using assms inj_on_eq_iff [of f] by (auto simp: arc_def inj_on_def path_linear_image_eq) subsection\<^marker>\tag unimportant\\Basic lemmas about paths\ +lemma 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 arc_imp_simple_path: "arc g \ simple_path g" by (simp add: arc_def inj_on_def simple_path_def) lemma arc_imp_path: "arc g \ path g" using arc_def by blast lemma arc_imp_inj_on: "arc g \ inj_on g {0..1}" by (auto simp: arc_def) lemma simple_path_imp_path: "simple_path g \ path g" using simple_path_def by blast lemma simple_path_cases: "simple_path g \ arc g \ pathfinish g = pathstart g" unfolding simple_path_def arc_def inj_on_def pathfinish_def pathstart_def by force lemma simple_path_imp_arc: "simple_path g \ pathfinish g \ pathstart g \ arc g" using simple_path_cases by auto lemma arc_distinct_ends: "arc g \ pathfinish g \ pathstart g" unfolding arc_def inj_on_def pathfinish_def pathstart_def by fastforce lemma arc_simple_path: "arc g \ simple_path g \ pathfinish g \ pathstart g" using arc_distinct_ends arc_imp_simple_path simple_path_cases by blast lemma simple_path_eq_arc: "pathfinish g \ pathstart g \ (simple_path g = arc g)" by (simp add: arc_simple_path) lemma path_image_const [simp]: "path_image (\t. a) = {a}" by (force simp: path_image_def) lemma path_image_nonempty [simp]: "path_image g \ {}" unfolding path_image_def image_is_empty box_eq_empty by auto lemma pathstart_in_path_image[intro]: "pathstart g \ path_image g" unfolding pathstart_def path_image_def by auto lemma pathfinish_in_path_image[intro]: "pathfinish g \ path_image g" unfolding pathfinish_def path_image_def by auto lemma connected_path_image[intro]: "path g \ connected (path_image g)" unfolding path_def path_image_def using connected_continuous_image connected_Icc by blast lemma compact_path_image[intro]: "path g \ compact (path_image g)" unfolding path_def path_image_def using compact_continuous_image connected_Icc by blast lemma reversepath_reversepath[simp]: "reversepath (reversepath g) = g" unfolding reversepath_def by auto lemma pathstart_reversepath[simp]: "pathstart (reversepath g) = pathfinish g" unfolding pathstart_def reversepath_def pathfinish_def by auto lemma pathfinish_reversepath[simp]: "pathfinish (reversepath g) = pathstart g" unfolding pathstart_def reversepath_def pathfinish_def by auto lemma reversepath_o: "reversepath g = g \ (-)1" by (auto simp: reversepath_def) lemma pathstart_join[simp]: "pathstart (g1 +++ g2) = pathstart g1" unfolding pathstart_def joinpaths_def pathfinish_def by auto lemma pathfinish_join[simp]: "pathfinish (g1 +++ g2) = pathfinish g2" unfolding pathstart_def joinpaths_def pathfinish_def by auto lemma path_image_reversepath[simp]: "path_image (reversepath g) = path_image g" proof - have *: "\g. path_image (reversepath g) \ path_image g" unfolding path_image_def subset_eq reversepath_def Ball_def image_iff by force show ?thesis using *[of g] *[of "reversepath g"] unfolding reversepath_reversepath by auto qed lemma path_reversepath [simp]: "path (reversepath g) \ path g" proof - have *: "\g. path g \ path (reversepath g)" unfolding path_def reversepath_def apply (rule continuous_on_compose[unfolded o_def, of _ "\x. 1 - x"]) apply (auto intro: continuous_intros continuous_on_subset[of "{0..1}"]) done show ?thesis using "*" by force qed lemma arc_reversepath: assumes "arc g" shows "arc(reversepath g)" proof - have injg: "inj_on g {0..1}" using assms by (simp add: arc_def) have **: "\x y::real. 1-x = 1-y \ x = y" by simp show ?thesis using assms by (clarsimp simp: arc_def intro!: inj_onI) (simp add: inj_onD reversepath_def **) qed lemma simple_path_reversepath: "simple_path g \ simple_path (reversepath g)" apply (simp add: simple_path_def) apply (force simp: reversepath_def) done lemmas reversepath_simps = path_reversepath path_image_reversepath pathstart_reversepath pathfinish_reversepath lemma path_join[simp]: assumes "pathfinish g1 = pathstart g2" shows "path (g1 +++ g2) \ path g1 \ path g2" unfolding path_def pathfinish_def pathstart_def proof safe assume cont: "continuous_on {0..1} (g1 +++ g2)" have g1: "continuous_on {0..1} g1 \ continuous_on {0..1} ((g1 +++ g2) \ (\x. x / 2))" by (intro continuous_on_cong refl) (auto simp: joinpaths_def) have g2: "continuous_on {0..1} g2 \ continuous_on {0..1} ((g1 +++ g2) \ (\x. x / 2 + 1/2))" using assms by (intro continuous_on_cong refl) (auto simp: joinpaths_def pathfinish_def pathstart_def) show "continuous_on {0..1} g1" and "continuous_on {0..1} g2" unfolding g1 g2 by (auto intro!: continuous_intros continuous_on_subset[OF cont] simp del: o_apply) next assume g1g2: "continuous_on {0..1} g1" "continuous_on {0..1} g2" have 01: "{0 .. 1} = {0..1/2} \ {1/2 .. 1::real}" by auto { fix x :: real assume "0 \ x" and "x \ 1" then have "x \ (\x. x * 2) ` {0..1 / 2}" by (intro image_eqI[where x="x/2"]) auto } note 1 = this { fix x :: real assume "0 \ x" and "x \ 1" then have "x \ (\x. x * 2 - 1) ` {1 / 2..1}" by (intro image_eqI[where x="x/2 + 1/2"]) auto } note 2 = this show "continuous_on {0..1} (g1 +++ g2)" using assms unfolding joinpaths_def 01 apply (intro continuous_on_cases closed_atLeastAtMost g1g2[THEN continuous_on_compose2] continuous_intros) apply (auto simp: field_simps pathfinish_def pathstart_def intro!: 1 2) done qed subsection\<^marker>\tag unimportant\ \Path Images\ lemma bounded_path_image: "path g \ bounded(path_image g)" by (simp add: compact_imp_bounded compact_path_image) lemma closed_path_image: fixes g :: "real \ 'a::t2_space" shows "path g \ closed(path_image g)" by (metis compact_path_image compact_imp_closed) lemma connected_simple_path_image: "simple_path g \ connected(path_image g)" by (metis connected_path_image simple_path_imp_path) lemma compact_simple_path_image: "simple_path g \ compact(path_image g)" by (metis compact_path_image simple_path_imp_path) lemma bounded_simple_path_image: "simple_path g \ bounded(path_image g)" by (metis bounded_path_image simple_path_imp_path) lemma closed_simple_path_image: fixes g :: "real \ 'a::t2_space" shows "simple_path g \ closed(path_image g)" by (metis closed_path_image simple_path_imp_path) lemma connected_arc_image: "arc g \ connected(path_image g)" by (metis connected_path_image arc_imp_path) lemma compact_arc_image: "arc g \ compact(path_image g)" by (metis compact_path_image arc_imp_path) lemma bounded_arc_image: "arc g \ bounded(path_image g)" by (metis bounded_path_image arc_imp_path) lemma closed_arc_image: fixes g :: "real \ 'a::t2_space" shows "arc g \ closed(path_image g)" by (metis closed_path_image arc_imp_path) lemma path_image_join_subset: "path_image (g1 +++ g2) \ path_image g1 \ path_image g2" unfolding path_image_def joinpaths_def by auto lemma subset_path_image_join: assumes "path_image g1 \ s" and "path_image g2 \ s" shows "path_image (g1 +++ g2) \ s" using path_image_join_subset[of g1 g2] and assms by auto lemma path_image_join: assumes "pathfinish g1 = pathstart g2" shows "path_image(g1 +++ g2) = path_image g1 \ path_image g2" proof - have "path_image g1 \ path_image (g1 +++ g2)" proof (clarsimp simp: path_image_def joinpaths_def) fix u::real assume "0 \ u" "u \ 1" then show "g1 u \ (\x. g1 (2 * x)) ` ({0..1} \ {x. x * 2 \ 1})" by (rule_tac x="u/2" in image_eqI) auto qed moreover have \
: "g2 u \ (\x. g2 (2 * x - 1)) ` ({0..1} \ {x. \ x * 2 \ 1})" if "0 < u" "u \ 1" for u using that assms by (rule_tac x="(u+1)/2" in image_eqI) (auto simp: field_simps pathfinish_def pathstart_def) have "g2 0 \ (\x. g1 (2 * x)) ` ({0..1} \ {x. x * 2 \ 1})" using assms by (rule_tac x="1/2" in image_eqI) (auto simp: pathfinish_def pathstart_def) then have "path_image g2 \ path_image (g1 +++ g2)" by (auto simp: path_image_def joinpaths_def intro!: \
) ultimately show ?thesis using path_image_join_subset by blast qed lemma not_in_path_image_join: assumes "x \ path_image g1" and "x \ path_image g2" shows "x \ path_image (g1 +++ g2)" using assms and path_image_join_subset[of g1 g2] by auto lemma pathstart_compose: "pathstart(f \ p) = f(pathstart p)" by (simp add: pathstart_def) lemma pathfinish_compose: "pathfinish(f \ p) = f(pathfinish p)" by (simp add: pathfinish_def) lemma path_image_compose: "path_image (f \ p) = f ` (path_image p)" by (simp add: image_comp path_image_def) lemma path_compose_join: "f \ (p +++ q) = (f \ p) +++ (f \ q)" by (rule ext) (simp add: joinpaths_def) lemma path_compose_reversepath: "f \ reversepath p = reversepath(f \ p)" by (rule ext) (simp add: reversepath_def) lemma joinpaths_eq: "(\t. t \ {0..1} \ p t = p' t) \ (\t. t \ {0..1} \ q t = q' t) \ t \ {0..1} \ (p +++ q) t = (p' +++ q') t" by (auto simp: joinpaths_def) lemma simple_path_inj_on: "simple_path g \ inj_on g {0<..<1}" by (auto simp: simple_path_def path_image_def inj_on_def less_eq_real_def Ball_def) subsection\<^marker>\tag unimportant\\Simple paths with the endpoints removed\ lemma simple_path_endless: assumes "simple_path c" shows "path_image c - {pathstart c,pathfinish c} = c ` {0<..<1}" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" using less_eq_real_def by (auto simp: path_image_def pathstart_def pathfinish_def) show "?rhs \ ?lhs" using assms apply (auto simp: simple_path_def path_image_def pathstart_def pathfinish_def Ball_def) using less_eq_real_def zero_le_one by blast+ qed lemma connected_simple_path_endless: assumes "simple_path c" shows "connected(path_image c - {pathstart c,pathfinish c})" proof - have "continuous_on {0<..<1} c" using assms by (simp add: simple_path_def continuous_on_path path_def subset_iff) then have "connected (c ` {0<..<1})" using connected_Ioo connected_continuous_image by blast then show ?thesis using assms by (simp add: simple_path_endless) qed lemma nonempty_simple_path_endless: "simple_path c \ path_image c - {pathstart c,pathfinish c} \ {}" by (simp add: simple_path_endless) subsection\<^marker>\tag unimportant\\The operations on paths\ lemma path_image_subset_reversepath: "path_image(reversepath g) \ path_image g" by simp lemma path_imp_reversepath: "path g \ path(reversepath g)" by simp lemma half_bounded_equal: "1 \ x * 2 \ x * 2 \ 1 \ x = (1/2::real)" by simp lemma continuous_on_joinpaths: assumes "continuous_on {0..1} g1" "continuous_on {0..1} g2" "pathfinish g1 = pathstart g2" shows "continuous_on {0..1} (g1 +++ g2)" proof - have "{0..1::real} = {0..1/2} \ {1/2..1}" by auto then show ?thesis using assms by (metis path_def path_join) qed lemma path_join_imp: "\path g1; path g2; pathfinish g1 = pathstart g2\ \ path(g1 +++ g2)" by simp lemma simple_path_join_loop: assumes "arc g1" "arc g2" "pathfinish g1 = pathstart g2" "pathfinish g2 = pathstart g1" "path_image g1 \ path_image g2 \ {pathstart g1, pathstart g2}" shows "simple_path(g1 +++ g2)" proof - have injg1: "inj_on g1 {0..1}" using assms by (simp add: arc_def) have injg2: "inj_on g2 {0..1}" using assms by (simp add: arc_def) have g12: "g1 1 = g2 0" and g21: "g2 1 = g1 0" and sb: "g1 ` {0..1} \ g2 ` {0..1} \ {g1 0, g2 0}" using assms by (simp_all add: arc_def pathfinish_def pathstart_def path_image_def) { fix x and y::real assume g2_eq: "g2 (2 * x - 1) = g1 (2 * y)" and xyI: "x \ 1 \ y \ 0" and xy: "x \ 1" "0 \ y" " y * 2 \ 1" "\ x * 2 \ 1" then consider "g1 (2 * y) = g1 0" | "g1 (2 * y) = g2 0" using sb by force then have False proof cases case 1 then have "y = 0" using xy g2_eq by (auto dest!: inj_onD [OF injg1]) then show ?thesis using xy g2_eq xyI by (auto dest: inj_onD [OF injg2] simp flip: g21) next case 2 then have "2*x = 1" using g2_eq g12 inj_onD [OF injg2] atLeastAtMost_iff xy(1) xy(4) by fastforce with xy show False by auto qed } note * = this { fix x and y::real assume xy: "g1 (2 * x) = g2 (2 * y - 1)" "y \ 1" "0 \ x" "\ y * 2 \ 1" "x * 2 \ 1" then have "x = 0 \ y = 1" using * xy by force } note ** = this show ?thesis using assms apply (simp add: arc_def simple_path_def) apply (auto simp: joinpaths_def split: if_split_asm dest!: * ** dest: inj_onD [OF injg1] inj_onD [OF injg2]) done qed lemma arc_join: assumes "arc g1" "arc g2" "pathfinish g1 = pathstart g2" "path_image g1 \ path_image g2 \ {pathstart g2}" shows "arc(g1 +++ g2)" proof - have injg1: "inj_on g1 {0..1}" using assms by (simp add: arc_def) have injg2: "inj_on g2 {0..1}" using assms by (simp add: arc_def) have g11: "g1 1 = g2 0" and sb: "g1 ` {0..1} \ g2 ` {0..1} \ {g2 0}" using assms by (simp_all add: arc_def pathfinish_def pathstart_def path_image_def) { fix x and y::real assume xy: "g2 (2 * x - 1) = g1 (2 * y)" "x \ 1" "0 \ y" " y * 2 \ 1" "\ x * 2 \ 1" then have "g1 (2 * y) = g2 0" using sb by force then have False using xy inj_onD injg2 by fastforce } note * = this show ?thesis using assms apply (simp add: arc_def inj_on_def) apply (auto simp: joinpaths_def arc_imp_path split: if_split_asm dest: * *[OF sym] inj_onD [OF injg1] inj_onD [OF injg2]) done qed lemma reversepath_joinpaths: "pathfinish g1 = pathstart g2 \ reversepath(g1 +++ g2) = reversepath g2 +++ reversepath g1" unfolding reversepath_def pathfinish_def pathstart_def joinpaths_def by (rule ext) (auto simp: mult.commute) subsection\<^marker>\tag unimportant\\Some reversed and "if and only if" versions of joining theorems\ lemma path_join_path_ends: fixes g1 :: "real \ 'a::metric_space" assumes "path(g1 +++ g2)" "path g2" shows "pathfinish g1 = pathstart g2" proof (rule ccontr) define e where "e = dist (g1 1) (g2 0)" assume Neg: "pathfinish g1 \ pathstart g2" then have "0 < dist (pathfinish g1) (pathstart g2)" by auto then have "e > 0" by (metis e_def pathfinish_def pathstart_def) then have "\e>0. \d>0. \x'\{0..1}. dist x' 0 < d \ dist (g2 x') (g2 0) < e" using \path g2\ atLeastAtMost_iff zero_le_one unfolding path_def continuous_on_iff by blast then obtain d1 where "d1 > 0" and d1: "\x'. \x'\{0..1}; norm x' < d1\ \ dist (g2 x') (g2 0) < e/2" by (metis \0 < e\ half_gt_zero_iff norm_conv_dist) obtain d2 where "d2 > 0" and d2: "\x'. \x'\{0..1}; dist x' (1/2) < d2\ \ dist ((g1 +++ g2) x') (g1 1) < e/2" using assms(1) \e > 0\ unfolding path_def continuous_on_iff apply (drule_tac x="1/2" in bspec, simp) apply (drule_tac x="e/2" in spec, force simp: joinpaths_def) done have int01_1: "min (1/2) (min d1 d2) / 2 \ {0..1}" using \d1 > 0\ \d2 > 0\ by (simp add: min_def) have dist1: "norm (min (1 / 2) (min d1 d2) / 2) < d1" using \d1 > 0\ \d2 > 0\ by (simp add: min_def dist_norm) have int01_2: "1/2 + min (1/2) (min d1 d2) / 4 \ {0..1}" using \d1 > 0\ \d2 > 0\ by (simp add: min_def) have dist2: "dist (1 / 2 + min (1 / 2) (min d1 d2) / 4) (1 / 2) < d2" using \d1 > 0\ \d2 > 0\ by (simp add: min_def dist_norm) have [simp]: "\ min (1 / 2) (min d1 d2) \ 0" using \d1 > 0\ \d2 > 0\ by (simp add: min_def) have "dist (g2 (min (1 / 2) (min d1 d2) / 2)) (g1 1) < e/2" "dist (g2 (min (1 / 2) (min d1 d2) / 2)) (g2 0) < e/2" using d1 [OF int01_1 dist1] d2 [OF int01_2 dist2] by (simp_all add: joinpaths_def) then have "dist (g1 1) (g2 0) < e/2 + e/2" using dist_triangle_half_r e_def by blast then show False by (simp add: e_def [symmetric]) qed lemma path_join_eq [simp]: fixes g1 :: "real \ 'a::metric_space" assumes "path g1" "path g2" shows "path(g1 +++ g2) \ pathfinish g1 = pathstart g2" using assms by (metis path_join_path_ends path_join_imp) lemma simple_path_joinE: assumes "simple_path(g1 +++ g2)" and "pathfinish g1 = pathstart g2" obtains "arc g1" "arc g2" "path_image g1 \ path_image g2 \ {pathstart g1, pathstart g2}" proof - have *: "\x y. \0 \ x; x \ 1; 0 \ y; y \ 1; (g1 +++ g2) x = (g1 +++ g2) y\ \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" using assms by (simp add: simple_path_def) have "path g1" using assms path_join simple_path_imp_path by blast moreover have "inj_on g1 {0..1}" proof (clarsimp simp: inj_on_def) fix x y assume "g1 x = g1 y" "0 \ x" "x \ 1" "0 \ y" "y \ 1" then show "x = y" using * [of "x/2" "y/2"] by (simp add: joinpaths_def split_ifs) qed ultimately have "arc g1" using assms by (simp add: arc_def) have [simp]: "g2 0 = g1 1" using assms by (metis pathfinish_def pathstart_def) have "path g2" using assms path_join simple_path_imp_path by blast moreover have "inj_on g2 {0..1}" proof (clarsimp simp: inj_on_def) fix x y assume "g2 x = g2 y" "0 \ x" "x \ 1" "0 \ y" "y \ 1" then show "x = y" using * [of "(x + 1) / 2" "(y + 1) / 2"] by (force simp: joinpaths_def split_ifs field_split_simps) qed ultimately have "arc g2" using assms by (simp add: arc_def) have "g2 y = g1 0 \ g2 y = g1 1" if "g1 x = g2 y" "0 \ x" "x \ 1" "0 \ y" "y \ 1" for x y using * [of "x / 2" "(y + 1) / 2"] that by (auto simp: joinpaths_def split_ifs field_split_simps) then have "path_image g1 \ path_image g2 \ {pathstart g1, pathstart g2}" by (fastforce simp: pathstart_def pathfinish_def path_image_def) with \arc g1\ \arc g2\ show ?thesis using that by blast qed lemma simple_path_join_loop_eq: assumes "pathfinish g2 = pathstart g1" "pathfinish g1 = pathstart g2" shows "simple_path(g1 +++ g2) \ arc g1 \ arc g2 \ path_image g1 \ path_image g2 \ {pathstart g1, pathstart g2}" by (metis assms simple_path_joinE simple_path_join_loop) lemma arc_join_eq: assumes "pathfinish g1 = pathstart g2" shows "arc(g1 +++ g2) \ arc g1 \ arc g2 \ path_image g1 \ path_image g2 \ {pathstart g2}" (is "?lhs = ?rhs") proof assume ?lhs then have "simple_path(g1 +++ g2)" by (rule arc_imp_simple_path) then have *: "\x y. \0 \ x; x \ 1; 0 \ y; y \ 1; (g1 +++ g2) x = (g1 +++ g2) y\ \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" using assms by (simp add: simple_path_def) have False if "g1 0 = g2 u" "0 \ u" "u \ 1" for u using * [of 0 "(u + 1) / 2"] that assms arc_distinct_ends [OF \?lhs\] by (auto simp: joinpaths_def pathstart_def pathfinish_def split_ifs field_split_simps) then have n1: "pathstart g1 \ path_image g2" unfolding pathstart_def path_image_def using atLeastAtMost_iff by blast show ?rhs using \?lhs\ using \simple_path (g1 +++ g2)\ assms n1 simple_path_joinE by auto next assume ?rhs then show ?lhs using assms by (fastforce simp: pathfinish_def pathstart_def intro!: arc_join) qed lemma arc_join_eq_alt: "pathfinish g1 = pathstart g2 \ (arc(g1 +++ g2) \ arc g1 \ arc g2 \ path_image g1 \ path_image g2 = {pathstart g2})" using pathfinish_in_path_image by (fastforce simp: arc_join_eq) subsection\<^marker>\tag unimportant\\The joining of paths is associative\ lemma path_assoc: "\pathfinish p = pathstart q; pathfinish q = pathstart r\ \ path(p +++ (q +++ r)) \ path((p +++ q) +++ r)" by simp lemma simple_path_assoc: assumes "pathfinish p = pathstart q" "pathfinish q = pathstart r" shows "simple_path (p +++ (q +++ r)) \ simple_path ((p +++ q) +++ r)" proof (cases "pathstart p = pathfinish r") case True show ?thesis proof assume "simple_path (p +++ q +++ r)" with assms True show "simple_path ((p +++ q) +++ r)" by (fastforce simp add: simple_path_join_loop_eq arc_join_eq path_image_join dest: arc_distinct_ends [of r]) next assume 0: "simple_path ((p +++ q) +++ r)" with assms True have q: "pathfinish r \ path_image q" using arc_distinct_ends by (fastforce simp add: simple_path_join_loop_eq arc_join_eq path_image_join) have "pathstart r \ path_image p" using assms by (metis 0 IntI arc_distinct_ends arc_join_eq_alt empty_iff insert_iff pathfinish_in_path_image pathfinish_join simple_path_joinE) with assms 0 q True show "simple_path (p +++ q +++ r)" by (auto simp: simple_path_join_loop_eq arc_join_eq path_image_join dest!: subsetD [OF _ IntI]) qed next case False { fix x :: 'a assume a: "path_image p \ path_image q \ {pathstart q}" "(path_image p \ path_image q) \ path_image r \ {pathstart r}" "x \ path_image p" "x \ path_image r" have "pathstart r \ path_image q" by (metis assms(2) pathfinish_in_path_image) with a have "x = pathstart q" by blast } with False assms show ?thesis by (auto simp: simple_path_eq_arc simple_path_join_loop_eq arc_join_eq path_image_join) qed lemma arc_assoc: "\pathfinish p = pathstart q; pathfinish q = pathstart r\ \ arc(p +++ (q +++ r)) \ arc((p +++ q) +++ r)" by (simp add: arc_simple_path simple_path_assoc) subsubsection\<^marker>\tag unimportant\\Symmetry and loops\ lemma path_sym: "\pathfinish p = pathstart q; pathfinish q = pathstart p\ \ path(p +++ q) \ path(q +++ p)" by auto lemma simple_path_sym: "\pathfinish p = pathstart q; pathfinish q = pathstart p\ \ simple_path(p +++ q) \ simple_path(q +++ p)" by (metis (full_types) inf_commute insert_commute simple_path_joinE simple_path_join_loop) lemma path_image_sym: "\pathfinish p = pathstart q; pathfinish q = pathstart p\ \ path_image(p +++ q) = path_image(q +++ p)" by (simp add: path_image_join sup_commute) subsection\Subpath\ definition\<^marker>\tag important\ subpath :: "real \ real \ (real \ 'a) \ real \ 'a::real_normed_vector" where "subpath a b g \ \x. g((b - a) * x + a)" lemma path_image_subpath_gen: fixes g :: "_ \ 'a::real_normed_vector" shows "path_image(subpath u v g) = g ` (closed_segment u v)" by (auto simp add: closed_segment_real_eq path_image_def subpath_def) lemma path_image_subpath: fixes g :: "real \ 'a::real_normed_vector" shows "path_image(subpath u v g) = (if u \ v then g ` {u..v} else g ` {v..u})" by (simp add: path_image_subpath_gen closed_segment_eq_real_ivl) lemma path_image_subpath_commute: fixes g :: "real \ 'a::real_normed_vector" shows "path_image(subpath u v g) = path_image(subpath v u g)" by (simp add: path_image_subpath_gen closed_segment_eq_real_ivl) lemma path_subpath [simp]: fixes g :: "real \ 'a::real_normed_vector" assumes "path g" "u \ {0..1}" "v \ {0..1}" shows "path(subpath u v g)" proof - have "continuous_on {0..1} (g \ (\x. ((v-u) * x+ u)))" using assms apply (intro continuous_intros; simp add: image_affinity_atLeastAtMost [where c=u]) apply (auto simp: path_def continuous_on_subset) done then show ?thesis by (simp add: path_def subpath_def) qed lemma pathstart_subpath [simp]: "pathstart(subpath u v g) = g(u)" by (simp add: pathstart_def subpath_def) lemma pathfinish_subpath [simp]: "pathfinish(subpath u v g) = g(v)" by (simp add: pathfinish_def subpath_def) lemma subpath_trivial [simp]: "subpath 0 1 g = g" by (simp add: subpath_def) lemma subpath_reversepath: "subpath 1 0 g = reversepath g" by (simp add: reversepath_def subpath_def) lemma reversepath_subpath: "reversepath(subpath u v g) = subpath v u g" by (simp add: reversepath_def subpath_def algebra_simps) lemma subpath_translation: "subpath u v ((\x. a + x) \ g) = (\x. a + x) \ subpath u v g" by (rule ext) (simp add: subpath_def) lemma subpath_image: "subpath u v (f \ g) = f \ subpath u v g" by (rule ext) (simp add: subpath_def) lemma affine_ineq: fixes x :: "'a::linordered_idom" assumes "x \ 1" "v \ u" shows "v + x * u \ u + x * v" proof - have "(1-x)*(u-v) \ 0" using assms by auto then show ?thesis by (simp add: algebra_simps) qed lemma sum_le_prod1: fixes a::real shows "\a \ 1; b \ 1\ \ a + b \ 1 + a * b" by (metis add.commute affine_ineq mult.right_neutral) lemma simple_path_subpath_eq: "simple_path(subpath u v g) \ path(subpath u v g) \ u\v \ (\x y. x \ closed_segment u v \ y \ closed_segment u v \ g x = g y \ x = y \ x = u \ y = v \ x = v \ y = u)" (is "?lhs = ?rhs") proof assume ?lhs then have p: "path (\x. g ((v - u) * x + u))" and sim: "(\x y. \x\{0..1}; y\{0..1}; g ((v - u) * x + u) = g ((v - u) * y + u)\ \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0)" by (auto simp: simple_path_def subpath_def) { fix x y assume "x \ closed_segment u v" "y \ closed_segment u v" "g x = g y" then have "x = y \ x = u \ y = v \ x = v \ y = u" using sim [of "(x-u)/(v-u)" "(y-u)/(v-u)"] p by (auto split: if_split_asm simp add: closed_segment_real_eq image_affinity_atLeastAtMost) (simp_all add: field_split_simps) } moreover have "path(subpath u v g) \ u\v" using sim [of "1/3" "2/3"] p by (auto simp: subpath_def) ultimately show ?rhs by metis next assume ?rhs then have d1: "\x y. \g x = g y; u \ x; x \ v; u \ y; y \ v\ \ x = y \ x = u \ y = v \ x = v \ y = u" and d2: "\x y. \g x = g y; v \ x; x \ u; v \ y; y \ u\ \ x = y \ x = u \ y = v \ x = v \ y = u" and ne: "u < v \ v < u" and psp: "path (subpath u v g)" by (auto simp: closed_segment_real_eq image_affinity_atLeastAtMost) have [simp]: "\x. u + x * v = v + x * u \ u=v \ x=1" by algebra show ?lhs using psp ne unfolding simple_path_def subpath_def by (fastforce simp add: algebra_simps affine_ineq mult_left_mono crossproduct_eq dest: d1 d2) qed lemma arc_subpath_eq: "arc(subpath u v g) \ path(subpath u v g) \ u\v \ inj_on g (closed_segment u v)" (is "?lhs = ?rhs") proof assume ?lhs then have p: "path (\x. g ((v - u) * x + u))" and sim: "(\x y. \x\{0..1}; y\{0..1}; g ((v - u) * x + u) = g ((v - u) * y + u)\ \ x = y)" by (auto simp: arc_def inj_on_def subpath_def) { fix x y assume "x \ closed_segment u v" "y \ closed_segment u v" "g x = g y" then have "x = y" using sim [of "(x-u)/(v-u)" "(y-u)/(v-u)"] p by (cases "v = u") (simp_all split: if_split_asm add: inj_on_def closed_segment_real_eq image_affinity_atLeastAtMost, simp add: field_simps) } moreover have "path(subpath u v g) \ u\v" using sim [of "1/3" "2/3"] p by (auto simp: subpath_def) ultimately show ?rhs unfolding inj_on_def by metis next assume ?rhs then have d1: "\x y. \g x = g y; u \ x; x \ v; u \ y; y \ v\ \ x = y" and d2: "\x y. \g x = g y; v \ x; x \ u; v \ y; y \ u\ \ x = y" and ne: "u < v \ v < u" and psp: "path (subpath u v g)" by (auto simp: inj_on_def closed_segment_real_eq image_affinity_atLeastAtMost) show ?lhs using psp ne unfolding arc_def subpath_def inj_on_def by (auto simp: algebra_simps affine_ineq mult_left_mono crossproduct_eq dest: d1 d2) qed lemma simple_path_subpath: assumes "simple_path g" "u \ {0..1}" "v \ {0..1}" "u \ v" shows "simple_path(subpath u v g)" using assms apply (simp add: simple_path_subpath_eq simple_path_imp_path) apply (simp add: simple_path_def closed_segment_real_eq image_affinity_atLeastAtMost, fastforce) done lemma arc_simple_path_subpath: "\simple_path g; u \ {0..1}; v \ {0..1}; g u \ g v\ \ arc(subpath u v g)" by (force intro: simple_path_subpath simple_path_imp_arc) lemma arc_subpath_arc: "\arc g; u \ {0..1}; v \ {0..1}; u \ v\ \ arc(subpath u v g)" by (meson arc_def arc_imp_simple_path arc_simple_path_subpath inj_onD) lemma arc_simple_path_subpath_interior: "\simple_path g; u \ {0..1}; v \ {0..1}; u \ v; \u-v\ < 1\ \ arc(subpath u v g)" by (force simp: simple_path_def intro: arc_simple_path_subpath) lemma path_image_subpath_subset: "\u \ {0..1}; v \ {0..1}\ \ path_image(subpath u v g) \ path_image g" by (metis atLeastAtMost_iff atLeastatMost_subset_iff path_image_def path_image_subpath subset_image_iff) lemma join_subpaths_middle: "subpath (0) ((1 / 2)) p +++ subpath ((1 / 2)) 1 p = p" by (rule ext) (simp add: joinpaths_def subpath_def field_split_simps) subsection\<^marker>\tag unimportant\\There is a subpath to the frontier\ lemma subpath_to_frontier_explicit: fixes S :: "'a::metric_space set" assumes g: "path g" and "pathfinish g \ S" obtains u where "0 \ u" "u \ 1" "\x. 0 \ x \ x < u \ g x \ interior S" "(g u \ interior S)" "(u = 0 \ g u \ closure S)" proof - have gcon: "continuous_on {0..1} g" using g by (simp add: path_def) moreover have "bounded ({u. g u \ closure (- S)} \ {0..1})" using compact_eq_bounded_closed by fastforce ultimately have com: "compact ({0..1} \ {u. g u \ closure (- S)})" using closed_vimage_Int by (metis (full_types) Int_commute closed_atLeastAtMost closed_closure compact_eq_bounded_closed vimage_def) have "1 \ {u. g u \ closure (- S)}" using assms by (simp add: pathfinish_def closure_def) then have dis: "{0..1} \ {u. g u \ closure (- S)} \ {}" using atLeastAtMost_iff zero_le_one by blast then obtain u where "0 \ u" "u \ 1" and gu: "g u \ closure (- S)" and umin: "\t. \0 \ t; t \ 1; g t \ closure (- S)\ \ u \ t" using compact_attains_inf [OF com dis] by fastforce then have umin': "\t. \0 \ t; t \ 1; t < u\ \ g t \ S" using closure_def by fastforce have \
: "g u \ closure S" if "u \ 0" proof - have "u > 0" using that \0 \ u\ by auto { fix e::real assume "e > 0" obtain d where "d>0" and d: "\x'. \x' \ {0..1}; dist x' u \ d\ \ dist (g x') (g u) < e" using continuous_onE [OF gcon _ \e > 0\] \0 \ _\ \_ \ 1\ atLeastAtMost_iff by auto have *: "dist (max 0 (u - d / 2)) u \ d" using \0 \ u\ \u \ 1\ \d > 0\ by (simp add: dist_real_def) have "\y\S. dist y (g u) < e" using \0 < u\ \u \ 1\ \d > 0\ by (force intro: d [OF _ *] umin') } then show ?thesis by (simp add: frontier_def closure_approachable) qed show ?thesis proof show "\x. 0 \ x \ x < u \ g x \ interior S" using \u \ 1\ interior_closure umin by fastforce show "g u \ interior S" by (simp add: gu interior_closure) qed (use \0 \ u\ \u \ 1\ \
in auto) qed lemma subpath_to_frontier_strong: assumes g: "path g" and "pathfinish g \ S" obtains u where "0 \ u" "u \ 1" "g u \ interior S" "u = 0 \ (\x. 0 \ x \ x < 1 \ subpath 0 u g x \ interior S) \ g u \ closure S" proof - obtain u where "0 \ u" "u \ 1" and gxin: "\x. 0 \ x \ x < u \ g x \ interior S" and gunot: "(g u \ interior S)" and u0: "(u = 0 \ g u \ closure S)" using subpath_to_frontier_explicit [OF assms] by blast show ?thesis proof show "g u \ interior S" using gunot by blast qed (use \0 \ u\ \u \ 1\ u0 in \(force simp: subpath_def gxin)+\) qed lemma subpath_to_frontier: assumes g: "path g" and g0: "pathstart g \ closure S" and g1: "pathfinish g \ S" obtains u where "0 \ u" "u \ 1" "g u \ frontier S" "path_image(subpath 0 u g) - {g u} \ interior S" proof - obtain u where "0 \ u" "u \ 1" and notin: "g u \ interior S" and disj: "u = 0 \ (\x. 0 \ x \ x < 1 \ subpath 0 u g x \ interior S) \ g u \ closure S" (is "_ \ ?P") using subpath_to_frontier_strong [OF g g1] by blast show ?thesis proof show "g u \ frontier S" by (metis DiffI disj frontier_def g0 notin pathstart_def) show "path_image (subpath 0 u g) - {g u} \ interior S" using disj proof assume "u = 0" then show ?thesis by (simp add: path_image_subpath) next assume P: ?P show ?thesis proof (clarsimp simp add: path_image_subpath_gen) fix y assume y: "y \ closed_segment 0 u" "g y \ interior S" with \0 \ u\ have "0 \ y" "y \ u" by (auto simp: closed_segment_eq_real_ivl split: if_split_asm) then have "y=u \ subpath 0 u g (y/u) \ interior S" using P less_eq_real_def by force then show "g y = g u" using y by (auto simp: subpath_def split: if_split_asm) qed qed qed (use \0 \ u\ \u \ 1\ in auto) qed lemma exists_path_subpath_to_frontier: fixes S :: "'a::real_normed_vector set" assumes "path g" "pathstart g \ closure S" "pathfinish g \ S" obtains h where "path h" "pathstart h = pathstart g" "path_image h \ path_image g" "path_image h - {pathfinish h} \ interior S" "pathfinish h \ frontier S" proof - obtain u where u: "0 \ u" "u \ 1" "g u \ frontier S" "(path_image(subpath 0 u g) - {g u}) \ interior S" using subpath_to_frontier [OF assms] by blast show ?thesis proof show "path_image (subpath 0 u g) \ path_image g" by (simp add: path_image_subpath_subset u) show "pathstart (subpath 0 u g) = pathstart g" by (metis pathstart_def pathstart_subpath) qed (use assms u in \auto simp: path_image_subpath\) qed lemma exists_path_subpath_to_frontier_closed: fixes S :: "'a::real_normed_vector set" assumes S: "closed S" and g: "path g" and g0: "pathstart g \ S" and g1: "pathfinish g \ S" obtains h where "path h" "pathstart h = pathstart g" "path_image h \ path_image g \ S" "pathfinish h \ frontier S" proof - obtain h where h: "path h" "pathstart h = pathstart g" "path_image h \ path_image g" "path_image h - {pathfinish h} \ interior S" "pathfinish h \ frontier S" using exists_path_subpath_to_frontier [OF g _ g1] closure_closed [OF S] g0 by auto show ?thesis proof show "path_image h \ path_image g \ S" using assms h interior_subset [of S] by (auto simp: frontier_def) qed (use h in auto) qed subsection \Shift Path to Start at Some Given Point\ definition\<^marker>\tag important\ shiftpath :: "real \ (real \ 'a::topological_space) \ real \ 'a" where "shiftpath a f = (\x. if (a + x) \ 1 then f (a + x) else f (a + x - 1))" lemma shiftpath_alt_def: "shiftpath a f = (\x. if x \ 1-a then f (a + x) else f (a + x - 1))" by (auto simp: shiftpath_def) lemma pathstart_shiftpath: "a \ 1 \ pathstart (shiftpath a g) = g a" unfolding pathstart_def shiftpath_def by auto lemma pathfinish_shiftpath: assumes "0 \ a" and "pathfinish g = pathstart g" shows "pathfinish (shiftpath a g) = g a" using assms unfolding pathstart_def pathfinish_def shiftpath_def by auto lemma endpoints_shiftpath: assumes "pathfinish g = pathstart g" and "a \ {0 .. 1}" shows "pathfinish (shiftpath a g) = g a" and "pathstart (shiftpath a g) = g a" using assms by (auto intro!: pathfinish_shiftpath pathstart_shiftpath) lemma closed_shiftpath: assumes "pathfinish g = pathstart g" and "a \ {0..1}" shows "pathfinish (shiftpath a g) = pathstart (shiftpath a g)" using endpoints_shiftpath[OF assms] by auto lemma path_shiftpath: assumes "path g" and "pathfinish g = pathstart g" and "a \ {0..1}" shows "path (shiftpath a g)" proof - have *: "{0 .. 1} = {0 .. 1-a} \ {1-a .. 1}" using assms(3) by auto have **: "\x. x + a = 1 \ g (x + a - 1) = g (x + a)" using assms(2)[unfolded pathfinish_def pathstart_def] by auto show ?thesis unfolding path_def shiftpath_def * proof (rule continuous_on_closed_Un) have contg: "continuous_on {0..1} g" using \path g\ path_def by blast show "continuous_on {0..1-a} (\x. if a + x \ 1 then g (a + x) else g (a + x - 1))" proof (rule continuous_on_eq) show "continuous_on {0..1-a} (g \ (+) a)" by (intro continuous_intros continuous_on_subset [OF contg]) (use \a \ {0..1}\ in auto) qed auto show "continuous_on {1-a..1} (\x. if a + x \ 1 then g (a + x) else g (a + x - 1))" proof (rule continuous_on_eq) show "continuous_on {1-a..1} (g \ (+) (a - 1))" by (intro continuous_intros continuous_on_subset [OF contg]) (use \a \ {0..1}\ in auto) qed (auto simp: "**" add.commute add_diff_eq) qed auto qed lemma shiftpath_shiftpath: assumes "pathfinish g = pathstart g" and "a \ {0..1}" and "x \ {0..1}" shows "shiftpath (1 - a) (shiftpath a g) x = g x" using assms unfolding pathfinish_def pathstart_def shiftpath_def by auto lemma path_image_shiftpath: assumes a: "a \ {0..1}" and "pathfinish g = pathstart g" shows "path_image (shiftpath a g) = path_image g" proof - { fix x assume g: "g 1 = g 0" "x \ {0..1::real}" and gne: "\y. y\{0..1} \ {x. \ a + x \ 1} \ g x \ g (a + y - 1)" then have "\y\{0..1} \ {x. a + x \ 1}. g x = g (a + y)" proof (cases "a \ x") case False then show ?thesis apply (rule_tac x="1 + x - a" in bexI) using g gne[of "1 + x - a"] a by (force simp: field_simps)+ next case True then show ?thesis using g a by (rule_tac x="x - a" in bexI) (auto simp: field_simps) qed } then show ?thesis using assms unfolding shiftpath_def path_image_def pathfinish_def pathstart_def by (auto simp: image_iff) qed lemma simple_path_shiftpath: assumes "simple_path g" "pathfinish g = pathstart g" and a: "0 \ a" "a \ 1" shows "simple_path (shiftpath a g)" unfolding simple_path_def proof (intro conjI impI ballI) show "path (shiftpath a g)" by (simp add: assms path_shiftpath simple_path_imp_path) have *: "\x y. \g x = g y; x \ {0..1}; y \ {0..1}\ \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" using assms by (simp add: simple_path_def) show "x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" if "x \ {0..1}" "y \ {0..1}" "shiftpath a g x = shiftpath a g y" for x y using that a unfolding shiftpath_def by (force split: if_split_asm dest!: *) qed subsection \Straight-Line Paths\ definition\<^marker>\tag important\ linepath :: "'a::real_normed_vector \ 'a \ real \ 'a" where "linepath a b = (\x. (1 - x) *\<^sub>R a + x *\<^sub>R b)" lemma pathstart_linepath[simp]: "pathstart (linepath a b) = a" unfolding pathstart_def linepath_def by auto lemma pathfinish_linepath[simp]: "pathfinish (linepath a b) = b" unfolding pathfinish_def linepath_def by auto lemma linepath_inner: "linepath a b x \ v = linepath (a \ v) (b \ v) x" by (simp add: linepath_def algebra_simps) lemma Re_linepath': "Re (linepath a b x) = linepath (Re a) (Re b) x" by (simp add: linepath_def) lemma Im_linepath': "Im (linepath a b x) = linepath (Im a) (Im b) x" by (simp add: linepath_def) lemma linepath_0': "linepath a b 0 = a" by (simp add: linepath_def) lemma linepath_1': "linepath a b 1 = b" by (simp add: linepath_def) lemma continuous_linepath_at[intro]: "continuous (at x) (linepath a b)" unfolding linepath_def by (intro continuous_intros) lemma continuous_on_linepath [intro,continuous_intros]: "continuous_on s (linepath a b)" using continuous_linepath_at by (auto intro!: continuous_at_imp_continuous_on) lemma path_linepath[iff]: "path (linepath a b)" unfolding path_def by (rule continuous_on_linepath) lemma path_image_linepath[simp]: "path_image (linepath a b) = closed_segment a b" unfolding path_image_def segment linepath_def by auto lemma reversepath_linepath[simp]: "reversepath (linepath a b) = linepath b a" unfolding reversepath_def linepath_def by auto lemma linepath_0 [simp]: "linepath 0 b x = x *\<^sub>R b" by (simp add: linepath_def) lemma linepath_cnj: "cnj (linepath a b x) = linepath (cnj a) (cnj b) x" by (simp add: linepath_def) lemma arc_linepath: assumes "a \ b" shows [simp]: "arc (linepath a b)" proof - { fix x y :: "real" assume "x *\<^sub>R b + y *\<^sub>R a = x *\<^sub>R a + y *\<^sub>R b" then have "(x - y) *\<^sub>R a = (x - y) *\<^sub>R b" by (simp add: algebra_simps) with assms have "x = y" by simp } then show ?thesis unfolding arc_def inj_on_def by (fastforce simp: algebra_simps linepath_def) qed lemma simple_path_linepath[intro]: "a \ b \ simple_path (linepath a b)" by (simp add: arc_imp_simple_path) lemma linepath_trivial [simp]: "linepath a a x = a" by (simp add: linepath_def real_vector.scale_left_diff_distrib) lemma linepath_refl: "linepath a a = (\x. a)" by auto lemma subpath_refl: "subpath a a g = linepath (g a) (g a)" by (simp add: subpath_def linepath_def algebra_simps) lemma linepath_of_real: "(linepath (of_real a) (of_real b) x) = of_real ((1 - x)*a + x*b)" by (simp add: scaleR_conv_of_real linepath_def) lemma of_real_linepath: "of_real (linepath a b x) = linepath (of_real a) (of_real b) x" by (metis linepath_of_real mult.right_neutral of_real_def real_scaleR_def) lemma inj_on_linepath: assumes "a \ b" shows "inj_on (linepath a b) {0..1}" proof (clarsimp simp: inj_on_def linepath_def) fix x y assume "(1 - x) *\<^sub>R a + x *\<^sub>R b = (1 - y) *\<^sub>R a + y *\<^sub>R b" "0 \ x" "x \ 1" "0 \ y" "y \ 1" then have "x *\<^sub>R (a - b) = y *\<^sub>R (a - b)" by (auto simp: algebra_simps) then show "x=y" using assms by auto qed lemma linepath_le_1: fixes a::"'a::linordered_idom" shows "\a \ 1; b \ 1; 0 \ u; u \ 1\ \ (1 - u) * a + u * b \ 1" using mult_left_le [of a "1-u"] mult_left_le [of b u] by auto lemma linepath_in_path: shows "x \ {0..1} \ linepath a b x \ closed_segment a b" by (auto simp: segment linepath_def) lemma linepath_image_01: "linepath a b ` {0..1} = closed_segment a b" by (auto simp: segment linepath_def) lemma linepath_in_convex_hull: fixes x::real assumes a: "a \ convex hull S" and b: "b \ convex hull S" and x: "0\x" "x\1" shows "linepath a b x \ convex hull S" proof - have "linepath a b x \ closed_segment a b" using x by (auto simp flip: linepath_image_01) then show ?thesis using a b convex_contains_segment by blast qed lemma Re_linepath: "Re(linepath (of_real a) (of_real b) x) = (1 - x)*a + x*b" by (simp add: linepath_def) lemma Im_linepath: "Im(linepath (of_real a) (of_real b) x) = 0" by (simp add: linepath_def) lemma bounded_linear_linepath: assumes "bounded_linear f" shows "f (linepath a b x) = linepath (f a) (f b) x" proof - interpret f: bounded_linear f by fact show ?thesis by (simp add: linepath_def f.add f.scale) qed lemma bounded_linear_linepath': assumes "bounded_linear f" shows "f \ linepath a b = linepath (f a) (f b)" using bounded_linear_linepath[OF assms] by (simp add: fun_eq_iff) lemma linepath_cnj': "cnj \ linepath a b = linepath (cnj a) (cnj b)" by (simp add: linepath_def fun_eq_iff) lemma differentiable_linepath [intro]: "linepath a b differentiable at x within A" by (auto simp: linepath_def) lemma has_vector_derivative_linepath_within: "(linepath a b has_vector_derivative (b - a)) (at x within S)" by (force intro: derivative_eq_intros simp add: linepath_def has_vector_derivative_def algebra_simps) subsection\<^marker>\tag unimportant\\Segments via convex hulls\ lemma segments_subset_convex_hull: "closed_segment a b \ (convex hull {a,b,c})" "closed_segment a c \ (convex hull {a,b,c})" "closed_segment b c \ (convex hull {a,b,c})" "closed_segment b a \ (convex hull {a,b,c})" "closed_segment c a \ (convex hull {a,b,c})" "closed_segment c b \ (convex hull {a,b,c})" by (auto simp: segment_convex_hull linepath_of_real elim!: rev_subsetD [OF _ hull_mono]) lemma midpoints_in_convex_hull: assumes "x \ convex hull s" "y \ convex hull s" shows "midpoint x y \ convex hull s" proof - have "(1 - inverse(2)) *\<^sub>R x + inverse(2) *\<^sub>R y \ convex hull s" by (rule convexD_alt) (use assms in auto) then show ?thesis by (simp add: midpoint_def algebra_simps) qed lemma not_in_interior_convex_hull_3: fixes a :: "complex" shows "a \ interior(convex hull {a,b,c})" "b \ interior(convex hull {a,b,c})" "c \ interior(convex hull {a,b,c})" by (auto simp: card_insert_le_m1 not_in_interior_convex_hull) lemma midpoint_in_closed_segment [simp]: "midpoint a b \ closed_segment a b" using midpoints_in_convex_hull segment_convex_hull by blast lemma midpoint_in_open_segment [simp]: "midpoint a b \ open_segment a b \ a \ b" by (simp add: open_segment_def) lemma continuous_IVT_local_extremum: fixes f :: "'a::euclidean_space \ real" assumes contf: "continuous_on (closed_segment a b) f" and "a \ b" "f a = f b" obtains z where "z \ open_segment a b" "(\w \ closed_segment a b. (f w) \ (f z)) \ (\w \ closed_segment a b. (f z) \ (f w))" proof - obtain c where "c \ closed_segment a b" and c: "\y. y \ closed_segment a b \ f y \ f c" using continuous_attains_sup [of "closed_segment a b" f] contf by auto obtain d where "d \ closed_segment a b" and d: "\y. y \ closed_segment a b \ f d \ f y" using continuous_attains_inf [of "closed_segment a b" f] contf by auto show ?thesis proof (cases "c \ open_segment a b \ d \ open_segment a b") case True then show ?thesis using c d that by blast next case False then have "(c = a \ c = b) \ (d = a \ d = b)" by (simp add: \c \ closed_segment a b\ \d \ closed_segment a b\ open_segment_def) with \a \ b\ \f a = f b\ c d show ?thesis by (rule_tac z = "midpoint a b" in that) (fastforce+) qed qed text\An injective map into R is also an open map w.r.T. the universe, and conversely. \ proposition injective_eq_1d_open_map_UNIV: fixes f :: "real \ real" assumes contf: "continuous_on S f" and S: "is_interval S" shows "inj_on f S \ (\T. open T \ T \ S \ open(f ` T))" (is "?lhs = ?rhs") proof safe fix T assume injf: ?lhs and "open T" and "T \ S" have "\U. open U \ f x \ U \ U \ f ` T" if "x \ T" for x proof - obtain \ where "\ > 0" and \: "cball x \ \ T" using \open T\ \x \ T\ open_contains_cball_eq by blast show ?thesis proof (intro exI conjI) have "closed_segment (x-\) (x+\) = {x-\..x+\}" using \0 < \\ by (auto simp: closed_segment_eq_real_ivl) also have "\ \ S" using \ \T \ S\ by (auto simp: dist_norm subset_eq) finally have "f ` (open_segment (x-\) (x+\)) = open_segment (f (x-\)) (f (x+\))" using continuous_injective_image_open_segment_1 by (metis continuous_on_subset [OF contf] inj_on_subset [OF injf]) then show "open (f ` {x-\<..})" using \0 < \\ by (simp add: open_segment_eq_real_ivl) show "f x \ f ` {x - \<..}" by (auto simp: \\ > 0\) show "f ` {x - \<..} \ f ` T" using \ by (auto simp: dist_norm subset_iff) qed qed with open_subopen show "open (f ` T)" by blast next assume R: ?rhs have False if xy: "x \ S" "y \ S" and "f x = f y" "x \ y" for x y proof - have "open (f ` open_segment x y)" using R by (metis S convex_contains_open_segment is_interval_convex open_greaterThanLessThan open_segment_eq_real_ivl xy) moreover have "continuous_on (closed_segment x y) f" by (meson S closed_segment_subset contf continuous_on_subset is_interval_convex that) then obtain \ where "\ \ open_segment x y" and \: "(\w \ closed_segment x y. (f w) \ (f \)) \ (\w \ closed_segment x y. (f \) \ (f w))" using continuous_IVT_local_extremum [of x y f] \f x = f y\ \x \ y\ by blast ultimately obtain e where "e>0" and e: "\u. dist u (f \) < e \ u \ f ` open_segment x y" using open_dist by (metis image_eqI) have fin: "f \ + (e/2) \ f ` open_segment x y" "f \ - (e/2) \ f ` open_segment x y" using e [of "f \ + (e/2)"] e [of "f \ - (e/2)"] \e > 0\ by (auto simp: dist_norm) show ?thesis using \ \0 < e\ fin open_closed_segment by fastforce qed then show ?lhs by (force simp: inj_on_def) qed subsection\<^marker>\tag unimportant\ \Bounding a point away from a path\ lemma not_on_path_ball: fixes g :: "real \ 'a::heine_borel" assumes "path g" and z: "z \ path_image g" shows "\e > 0. ball z e \ path_image g = {}" proof - have "closed (path_image g)" by (simp add: \path g\ closed_path_image) then obtain a where "a \ path_image g" "\y \ path_image g. dist z a \ dist z y" by (auto intro: distance_attains_inf[OF _ path_image_nonempty, of g z]) then show ?thesis by (rule_tac x="dist z a" in exI) (use dist_commute z in auto) qed lemma not_on_path_cball: fixes g :: "real \ 'a::heine_borel" assumes "path g" and "z \ path_image g" shows "\e>0. cball z e \ (path_image g) = {}" proof - obtain e where "ball z e \ path_image g = {}" "e > 0" using not_on_path_ball[OF assms] by auto moreover have "cball z (e/2) \ ball z e" using \e > 0\ by auto ultimately show ?thesis by (rule_tac x="e/2" in exI) auto qed subsection \Path component\ text \Original formalization by Tom Hales\ definition\<^marker>\tag important\ "path_component S x y \ (\g. path g \ path_image g \ S \ pathstart g = x \ pathfinish g = y)" abbreviation\<^marker>\tag important\ "path_component_set S x \ Collect (path_component S x)" lemmas path_defs = path_def pathstart_def pathfinish_def path_image_def path_component_def lemma path_component_mem: assumes "path_component S x y" shows "x \ S" and "y \ S" using assms unfolding path_defs by auto lemma path_component_refl: assumes "x \ S" shows "path_component S x x" using assms unfolding path_defs by (metis (full_types) assms continuous_on_const image_subset_iff path_image_def) lemma path_component_refl_eq: "path_component S x x \ x \ S" by (auto intro!: path_component_mem path_component_refl) lemma path_component_sym: "path_component S x y \ path_component S y x" unfolding path_component_def by (metis (no_types) path_image_reversepath path_reversepath pathfinish_reversepath pathstart_reversepath) lemma path_component_trans: assumes "path_component S x y" and "path_component S y z" shows "path_component S x z" using assms unfolding path_component_def by (metis path_join pathfinish_join pathstart_join subset_path_image_join) lemma path_component_of_subset: "S \ T \ path_component S x y \ path_component T x y" unfolding path_component_def by auto lemma path_component_linepath: fixes S :: "'a::real_normed_vector set" shows "closed_segment a b \ S \ path_component S a b" unfolding path_component_def by (rule_tac x="linepath a b" in exI, auto) subsubsection\<^marker>\tag unimportant\ \Path components as sets\ lemma path_component_set: "path_component_set S x = {y. (\g. path g \ path_image g \ S \ pathstart g = x \ pathfinish g = y)}" by (auto simp: path_component_def) lemma path_component_subset: "path_component_set S x \ S" by (auto simp: path_component_mem(2)) lemma path_component_eq_empty: "path_component_set S x = {} \ x \ S" using path_component_mem path_component_refl_eq by fastforce lemma path_component_mono: "S \ T \ (path_component_set S x) \ (path_component_set T x)" by (simp add: Collect_mono path_component_of_subset) lemma path_component_eq: "y \ path_component_set S x \ path_component_set S y = path_component_set S x" by (metis (no_types, lifting) Collect_cong mem_Collect_eq path_component_sym path_component_trans) subsection \Path connectedness of a space\ definition\<^marker>\tag important\ "path_connected S \ (\x\S. \y\S. \g. path g \ path_image g \ S \ pathstart g = x \ pathfinish g = y)" lemma path_connectedin_iff_path_connected_real [simp]: "path_connectedin euclideanreal S \ path_connected S" by (simp add: path_connectedin path_connected_def path_defs) lemma path_connected_component: "path_connected S \ (\x\S. \y\S. path_component S x y)" unfolding path_connected_def path_component_def by auto lemma path_connected_component_set: "path_connected S \ (\x\S. path_component_set S x = S)" unfolding path_connected_component path_component_subset using path_component_mem by blast lemma path_component_maximal: "\x \ T; path_connected T; T \ S\ \ T \ (path_component_set S x)" by (metis path_component_mono path_connected_component_set) lemma convex_imp_path_connected: fixes S :: "'a::real_normed_vector set" assumes "convex S" shows "path_connected S" unfolding path_connected_def using assms convex_contains_segment by fastforce lemma path_connected_UNIV [iff]: "path_connected (UNIV :: 'a::real_normed_vector set)" by (simp add: convex_imp_path_connected) lemma path_component_UNIV: "path_component_set UNIV x = (UNIV :: 'a::real_normed_vector set)" using path_connected_component_set by auto lemma path_connected_imp_connected: assumes "path_connected S" shows "connected S" proof (rule connectedI) fix e1 e2 assume as: "open e1" "open e2" "S \ e1 \ e2" "e1 \ e2 \ S = {}" "e1 \ S \ {}" "e2 \ S \ {}" then obtain x1 x2 where obt:"x1 \ e1 \ S" "x2 \ e2 \ S" by auto then obtain g where g: "path g" "path_image g \ S" "pathstart g = x1" "pathfinish g = x2" using assms[unfolded path_connected_def,rule_format,of x1 x2] by auto have *: "connected {0..1::real}" by (auto intro!: convex_connected) have "{0..1} \ {x \ {0..1}. g x \ e1} \ {x \ {0..1}. g x \ e2}" using as(3) g(2)[unfolded path_defs] by blast moreover have "{x \ {0..1}. g x \ e1} \ {x \ {0..1}. g x \ e2} = {}" using as(4) g(2)[unfolded path_defs] unfolding subset_eq by auto moreover have "{x \ {0..1}. g x \ e1} \ {} \ {x \ {0..1}. g x \ e2} \ {}" using g(3,4)[unfolded path_defs] using obt by (simp add: ex_in_conv [symmetric], metis zero_le_one order_refl) ultimately show False using *[unfolded connected_local not_ex, rule_format, of "{0..1} \ g -` e1" "{0..1} \ g -` e2"] using continuous_openin_preimage_gen[OF g(1)[unfolded path_def] as(1)] using continuous_openin_preimage_gen[OF g(1)[unfolded path_def] as(2)] by auto qed lemma open_path_component: fixes S :: "'a::real_normed_vector set" assumes "open S" shows "open (path_component_set S x)" unfolding open_contains_ball proof fix y assume as: "y \ path_component_set S x" then have "y \ S" by (simp add: path_component_mem(2)) then obtain e where e: "e > 0" "ball y e \ S" using assms openE by blast have "\u. dist y u < e \ path_component S x u" by (metis (full_types) as centre_in_ball convex_ball convex_imp_path_connected e mem_Collect_eq mem_ball path_component_eq path_component_of_subset path_connected_component) then show "\e > 0. ball y e \ path_component_set S x" using \e>0\ by auto qed lemma open_non_path_component: fixes S :: "'a::real_normed_vector set" assumes "open S" shows "open (S - path_component_set S x)" unfolding open_contains_ball proof fix y assume y: "y \ S - path_component_set S x" then obtain e where e: "e > 0" "ball y e \ S" using assms openE by auto show "\e>0. ball y e \ S - path_component_set S x" proof (intro exI conjI subsetI DiffI notI) show "\x. x \ ball y e \ x \ S" using e by blast show False if "z \ ball y e" "z \ path_component_set S x" for z proof - have "y \ path_component_set S z" by (meson assms convex_ball convex_imp_path_connected e open_contains_ball_eq open_path_component path_component_maximal that(1)) then have "y \ path_component_set S x" using path_component_eq that(2) by blast then show False using y by blast qed qed (use e in auto) qed lemma connected_open_path_connected: fixes S :: "'a::real_normed_vector set" assumes "open S" and "connected S" shows "path_connected S" unfolding path_connected_component_set proof (rule, rule, rule path_component_subset, rule) fix x y assume "x \ S" and "y \ S" show "y \ path_component_set S x" proof (rule ccontr) assume "\ ?thesis" moreover have "path_component_set S x \ S \ {}" using \x \ S\ path_component_eq_empty path_component_subset[of S x] by auto ultimately show False using \y \ S\ open_non_path_component[OF assms(1)] open_path_component[OF assms(1)] using assms(2)[unfolded connected_def not_ex, rule_format, of "path_component_set S x" "S - path_component_set S x"] by auto qed qed lemma path_connected_continuous_image: assumes contf: "continuous_on S f" and "path_connected S" shows "path_connected (f ` S)" unfolding path_connected_def proof (rule, rule) fix x' y' assume "x' \ f ` S" "y' \ f ` S" then obtain x y where x: "x \ S" and y: "y \ S" and x': "x' = f x" and y': "y' = f y" by auto from x y obtain g where "path g \ path_image g \ S \ pathstart g = x \ pathfinish g = y" using assms(2)[unfolded path_connected_def] by fast then show "\g. path g \ path_image g \ f ` S \ pathstart g = x' \ pathfinish g = y'" unfolding x' y' path_defs by (fastforce intro: continuous_on_compose continuous_on_subset[OF contf]) qed lemma path_connected_translationI: fixes a :: "'a :: topological_group_add" assumes "path_connected S" shows "path_connected ((\x. a + x) ` S)" by (intro path_connected_continuous_image assms continuous_intros) lemma path_connected_translation: fixes a :: "'a :: topological_group_add" shows "path_connected ((\x. a + x) ` S) = path_connected S" proof - have "\x y. (+) (x::'a) ` (+) (0 - x) ` y = y" by (simp add: image_image) then show ?thesis by (metis (no_types) path_connected_translationI) qed lemma path_connected_segment [simp]: fixes a :: "'a::real_normed_vector" shows "path_connected (closed_segment a b)" by (simp add: convex_imp_path_connected) lemma path_connected_open_segment [simp]: fixes a :: "'a::real_normed_vector" shows "path_connected (open_segment a b)" by (simp add: convex_imp_path_connected) lemma homeomorphic_path_connectedness: "S homeomorphic T \ path_connected S \ path_connected T" unfolding homeomorphic_def homeomorphism_def by (metis path_connected_continuous_image) lemma path_connected_empty [simp]: "path_connected {}" unfolding path_connected_def by auto lemma path_connected_singleton [simp]: "path_connected {a}" unfolding path_connected_def pathstart_def pathfinish_def path_image_def using path_def by fastforce lemma path_connected_Un: assumes "path_connected S" and "path_connected T" and "S \ T \ {}" shows "path_connected (S \ T)" unfolding path_connected_component proof (intro ballI) fix x y assume x: "x \ S \ T" and y: "y \ S \ T" from assms obtain z where z: "z \ S" "z \ T" by auto show "path_component (S \ T) x y" using x y proof safe assume "x \ S" "y \ S" then show "path_component (S \ T) x y" by (meson Un_upper1 \path_connected S\ path_component_of_subset path_connected_component) next assume "x \ S" "y \ T" then show "path_component (S \ T) x y" by (metis z assms(1-2) le_sup_iff order_refl path_component_of_subset path_component_trans path_connected_component) next assume "x \ T" "y \ S" then show "path_component (S \ T) x y" by (metis z assms(1-2) le_sup_iff order_refl path_component_of_subset path_component_trans path_connected_component) next assume "x \ T" "y \ T" then show "path_component (S \ T) x y" by (metis Un_upper1 assms(2) path_component_of_subset path_connected_component sup_commute) qed qed lemma path_connected_UNION: assumes "\i. i \ A \ path_connected (S i)" and "\i. i \ A \ z \ S i" shows "path_connected (\i\A. S i)" unfolding path_connected_component proof clarify fix x i y j assume *: "i \ A" "x \ S i" "j \ A" "y \ S j" then have "path_component (S i) x z" and "path_component (S j) z y" using assms by (simp_all add: path_connected_component) then have "path_component (\i\A. S i) x z" and "path_component (\i\A. S i) z y" using *(1,3) by (auto elim!: path_component_of_subset [rotated]) then show "path_component (\i\A. S i) x y" by (rule path_component_trans) qed lemma path_component_path_image_pathstart: assumes p: "path p" and x: "x \ path_image p" shows "path_component (path_image p) (pathstart p) x" proof - obtain y where x: "x = p y" and y: "0 \ y" "y \ 1" using x by (auto simp: path_image_def) show ?thesis unfolding path_component_def proof (intro exI conjI) have "continuous_on ((*) y ` {0..1}) p" by (simp add: continuous_on_path image_mult_atLeastAtMost_if p y) then have "continuous_on {0..1} (p \ ((*) y))" using continuous_on_compose continuous_on_mult_const by blast then show "path (\u. p (y * u))" by (simp add: path_def) show "path_image (\u. p (y * u)) \ path_image p" using y mult_le_one by (fastforce simp: path_image_def image_iff) qed (auto simp: pathstart_def pathfinish_def x) qed lemma path_connected_path_image: "path p \ path_connected(path_image p)" unfolding path_connected_component by (meson path_component_path_image_pathstart path_component_sym path_component_trans) lemma path_connected_path_component [simp]: "path_connected (path_component_set s x)" proof - { fix y z assume pa: "path_component s x y" "path_component s x z" then have pae: "path_component_set s x = path_component_set s y" using path_component_eq by auto have yz: "path_component s y z" using pa path_component_sym path_component_trans by blast then have "\g. path g \ path_image g \ path_component_set s x \ pathstart g = y \ pathfinish g = z" apply (simp add: path_component_def) by (metis pae path_component_maximal path_connected_path_image pathstart_in_path_image) } then show ?thesis by (simp add: path_connected_def) qed lemma path_component: "path_component S x y \ (\t. path_connected t \ t \ S \ x \ t \ y \ t)" apply (intro iffI) apply (metis path_connected_path_image path_defs(5) pathfinish_in_path_image pathstart_in_path_image) using path_component_of_subset path_connected_component by blast lemma path_component_path_component [simp]: "path_component_set (path_component_set S x) x = path_component_set S x" proof (cases "x \ S") case True show ?thesis by (metis True mem_Collect_eq path_component_refl path_connected_component_set path_connected_path_component) next case False then show ?thesis by (metis False empty_iff path_component_eq_empty) qed lemma path_component_subset_connected_component: "(path_component_set S x) \ (connected_component_set S x)" proof (cases "x \ S") case True show ?thesis by (simp add: True connected_component_maximal path_component_refl path_component_subset path_connected_imp_connected) next case False then show ?thesis using path_component_eq_empty by auto qed subsection\<^marker>\tag unimportant\\Lemmas about path-connectedness\ lemma path_connected_linear_image: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "path_connected S" "bounded_linear f" shows "path_connected(f ` S)" by (auto simp: linear_continuous_on assms path_connected_continuous_image) lemma is_interval_path_connected: "is_interval S \ path_connected S" by (simp add: convex_imp_path_connected is_interval_convex) lemma path_connected_Ioi[simp]: "path_connected {a<..}" for a :: real by (simp add: convex_imp_path_connected) lemma path_connected_Ici[simp]: "path_connected {a..}" for a :: real by (simp add: convex_imp_path_connected) lemma path_connected_Iio[simp]: "path_connected {.. (\x \ topspace X. \y \ topspace X. \S. path_connectedin X S \ x \ S \ y \ S)" by (metis path_connectedin path_connectedin_topspace path_connected_space_def) lemma connectedin_path_image: "pathin X g \ connectedin X (g ` ({0..1}))" by (simp add: path_connectedin_imp_connectedin path_connectedin_path_image) lemma compactin_path_image: "pathin X g \ compactin X (g ` ({0..1}))" unfolding pathin_def by (rule image_compactin [of "top_of_set {0..1}"]) auto lemma linear_homeomorphism_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" obtains g where "homeomorphism (f ` S) S g f" proof - obtain g where "linear g" "g \ f = id" using assms linear_injective_left_inverse by blast then have "homeomorphism (f ` S) S g f" using assms unfolding homeomorphism_def by (auto simp: eq_id_iff [symmetric] image_comp linear_conv_bounded_linear linear_continuous_on) then show thesis .. qed lemma linear_homeomorphic_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "S homeomorphic f ` S" by (meson homeomorphic_def homeomorphic_sym linear_homeomorphism_image [OF assms]) lemma path_connected_Times: assumes "path_connected s" "path_connected t" shows "path_connected (s \ t)" proof (simp add: path_connected_def Sigma_def, clarify) fix x1 y1 x2 y2 assume "x1 \ s" "y1 \ t" "x2 \ s" "y2 \ t" obtain g where "path g" and g: "path_image g \ s" and gs: "pathstart g = x1" and gf: "pathfinish g = x2" using \x1 \ s\ \x2 \ s\ assms by (force simp: path_connected_def) obtain h where "path h" and h: "path_image h \ t" and hs: "pathstart h = y1" and hf: "pathfinish h = y2" using \y1 \ t\ \y2 \ t\ assms by (force simp: path_connected_def) have "path (\z. (x1, h z))" using \path h\ unfolding path_def by (intro continuous_intros continuous_on_compose2 [where g = "Pair _"]; force) moreover have "path (\z. (g z, y2))" using \path g\ unfolding path_def by (intro continuous_intros continuous_on_compose2 [where g = "Pair _"]; force) ultimately have 1: "path ((\z. (x1, h z)) +++ (\z. (g z, y2)))" by (metis hf gs path_join_imp pathstart_def pathfinish_def) have "path_image ((\z. (x1, h z)) +++ (\z. (g z, y2))) \ path_image (\z. (x1, h z)) \ path_image (\z. (g z, y2))" by (rule Path_Connected.path_image_join_subset) also have "\ \ (\x\s. \x1\t. {(x, x1)})" using g h \x1 \ s\ \y2 \ t\ by (force simp: path_image_def) finally have 2: "path_image ((\z. (x1, h z)) +++ (\z. (g z, y2))) \ (\x\s. \x1\t. {(x, x1)})" . show "\g. path g \ path_image g \ (\x\s. \x1\t. {(x, x1)}) \ pathstart g = (x1, y1) \ pathfinish g = (x2, y2)" using 1 2 gf hs by (metis (no_types, lifting) pathfinish_def pathfinish_join pathstart_def pathstart_join) qed lemma is_interval_path_connected_1: fixes s :: "real set" shows "is_interval s \ path_connected s" using is_interval_connected_1 is_interval_path_connected path_connected_imp_connected by blast subsection\<^marker>\tag unimportant\\Path components\ lemma Union_path_component [simp]: "Union {path_component_set S x |x. x \ S} = S" apply (rule subset_antisym) using path_component_subset apply force using path_component_refl by auto lemma path_component_disjoint: "disjnt (path_component_set S a) (path_component_set S b) \ (a \ path_component_set S b)" unfolding disjnt_iff using path_component_sym path_component_trans by blast lemma path_component_eq_eq: "path_component S x = path_component S y \ (x \ S) \ (y \ S) \ x \ S \ y \ S \ path_component S x y" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (metis (no_types) path_component_mem(1) path_component_refl) next assume ?rhs then show ?lhs proof assume "x \ S \ y \ S" then show ?lhs by (metis Collect_empty_eq_bot path_component_eq_empty) next assume S: "x \ S \ y \ S \ path_component S x y" show ?lhs by (rule ext) (metis S path_component_trans path_component_sym) qed qed lemma path_component_unique: assumes "x \ c" "c \ S" "path_connected c" "\c'. \x \ c'; c' \ S; path_connected c'\ \ c' \ c" shows "path_component_set S x = c" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" using assms by (metis mem_Collect_eq path_component_refl path_component_subset path_connected_path_component subsetD) qed (simp add: assms path_component_maximal) lemma path_component_intermediate_subset: "path_component_set u a \ t \ t \ u \ path_component_set t a = path_component_set u a" by (metis (no_types) path_component_mono path_component_path_component subset_antisym) lemma complement_path_component_Union: fixes x :: "'a :: topological_space" shows "S - path_component_set S x = \({path_component_set S y| y. y \ S} - {path_component_set S x})" proof - have *: "(\x. x \ S - {a} \ disjnt a x) \ \S - a = \(S - {a})" for a::"'a set" and S by (auto simp: disjnt_def) have "\y. y \ {path_component_set S x |x. x \ S} - {path_component_set S x} \ disjnt (path_component_set S x) y" using path_component_disjoint path_component_eq by fastforce then have "\{path_component_set S x |x. x \ S} - path_component_set S x = \({path_component_set S y |y. y \ S} - {path_component_set S x})" by (meson *) then show ?thesis by simp qed subsection\Path components\ definition path_component_of where "path_component_of X x y \ \g. pathin X g \ g 0 = x \ g 1 = y" abbreviation path_component_of_set where "path_component_of_set X x \ Collect (path_component_of X x)" definition path_components_of :: "'a topology \ 'a set set" where "path_components_of X \ path_component_of_set X ` topspace X" lemma pathin_canon_iff: "pathin (top_of_set T) g \ path g \ g ` {0..1} \ T" by (simp add: path_def pathin_def) lemma path_component_of_canon_iff [simp]: "path_component_of (top_of_set T) a b \ path_component T a b" by (simp add: path_component_of_def pathin_canon_iff path_defs) lemma path_component_in_topspace: "path_component_of X x y \ x \ topspace X \ y \ topspace X" by (auto simp: path_component_of_def pathin_def continuous_map_def) lemma path_component_of_refl: "path_component_of X x x \ x \ topspace X" by (metis path_component_in_topspace path_component_of_def pathin_const) lemma path_component_of_sym: assumes "path_component_of X x y" shows "path_component_of X y x" using assms apply (clarsimp simp: path_component_of_def pathin_def) apply (rule_tac x="g \ (\t. 1 - t)" in exI) apply (auto intro!: continuous_map_compose simp: continuous_map_in_subtopology continuous_on_op_minus) done lemma path_component_of_sym_iff: "path_component_of X x y \ path_component_of X y x" by (metis path_component_of_sym) lemma continuous_map_cases_le: assumes contp: "continuous_map X euclideanreal p" and contq: "continuous_map X euclideanreal q" and contf: "continuous_map (subtopology X {x. x \ topspace X \ p x \ q x}) Y f" and contg: "continuous_map (subtopology X {x. x \ topspace X \ q x \ p x}) Y g" and fg: "\x. \x \ topspace X; p x = q x\ \ f x = g x" shows "continuous_map X Y (\x. if p x \ q x then f x else g x)" proof - have "continuous_map X Y (\x. if q x - p x \ {0..} then f x else g x)" proof (rule continuous_map_cases_function) show "continuous_map X euclideanreal (\x. q x - p x)" by (intro contp contq continuous_intros) show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of {0..}}) Y f" by (simp add: contf) show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of (topspace euclideanreal - {0..})}) Y g" by (simp add: contg flip: Compl_eq_Diff_UNIV) qed (auto simp: fg) then show ?thesis by simp qed lemma continuous_map_cases_lt: assumes contp: "continuous_map X euclideanreal p" and contq: "continuous_map X euclideanreal q" and contf: "continuous_map (subtopology X {x. x \ topspace X \ p x \ q x}) Y f" and contg: "continuous_map (subtopology X {x. x \ topspace X \ q x \ p x}) Y g" and fg: "\x. \x \ topspace X; p x = q x\ \ f x = g x" shows "continuous_map X Y (\x. if p x < q x then f x else g x)" proof - have "continuous_map X Y (\x. if q x - p x \ {0<..} then f x else g x)" proof (rule continuous_map_cases_function) show "continuous_map X euclideanreal (\x. q x - p x)" by (intro contp contq continuous_intros) show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of {0<..}}) Y f" by (simp add: contf) show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of (topspace euclideanreal - {0<..})}) Y g" by (simp add: contg flip: Compl_eq_Diff_UNIV) qed (auto simp: fg) then show ?thesis by simp qed lemma path_component_of_trans: assumes "path_component_of X x y" and "path_component_of X y z" shows "path_component_of X x z" unfolding path_component_of_def pathin_def proof - let ?T01 = "top_of_set {0..1::real}" obtain g1 g2 where g1: "continuous_map ?T01 X g1" "x = g1 0" "y = g1 1" and g2: "continuous_map ?T01 X g2" "g2 0 = g1 1" "z = g2 1" using assms unfolding path_component_of_def pathin_def by blast let ?g = "\x. if x \ 1/2 then (g1 \ (\t. 2 * t)) x else (g2 \ (\t. 2 * t -1)) x" show "\g. continuous_map ?T01 X g \ g 0 = x \ g 1 = z" proof (intro exI conjI) show "continuous_map (subtopology euclideanreal {0..1}) X ?g" proof (intro continuous_map_cases_le continuous_map_compose, force, force) show "continuous_map (subtopology ?T01 {x \ topspace ?T01. x \ 1/2}) ?T01 ((*) 2)" by (auto simp: continuous_map_in_subtopology continuous_map_from_subtopology) have "continuous_map (subtopology (top_of_set {0..1}) {x. 0 \ x \ x \ 1 \ 1 \ x * 2}) euclideanreal (\t. 2 * t - 1)" by (intro continuous_intros) (force intro: continuous_map_from_subtopology) then show "continuous_map (subtopology ?T01 {x \ topspace ?T01. 1/2 \ x}) ?T01 (\t. 2 * t - 1)" by (force simp: continuous_map_in_subtopology) show "(g1 \ (*) 2) x = (g2 \ (\t. 2 * t - 1)) x" if "x \ topspace ?T01" "x = 1/2" for x using that by (simp add: g2(2) mult.commute continuous_map_from_subtopology) qed (auto simp: g1 g2) qed (auto simp: g1 g2) qed lemma path_component_of_mono: "\path_component_of (subtopology X S) x y; S \ T\ \ path_component_of (subtopology X T) x y" unfolding path_component_of_def by (metis subsetD pathin_subtopology) lemma path_component_of: "path_component_of X x y \ (\T. path_connectedin X T \ x \ T \ y \ T)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (metis atLeastAtMost_iff image_eqI order_refl path_component_of_def path_connectedin_path_image zero_le_one) next assume ?rhs then show ?lhs by (metis path_component_of_def path_connectedin) qed lemma path_component_of_set: "path_component_of X x y \ (\g. pathin X g \ g 0 = x \ g 1 = y)" by (auto simp: path_component_of_def) lemma path_component_of_subset_topspace: "Collect(path_component_of X x) \ topspace X" using path_component_in_topspace by fastforce lemma path_component_of_eq_empty: "Collect(path_component_of X x) = {} \ (x \ topspace X)" using path_component_in_topspace path_component_of_refl by fastforce lemma path_connected_space_iff_path_component: "path_connected_space X \ (\x \ topspace X. \y \ topspace X. path_component_of X x y)" by (simp add: path_component_of path_connected_space_subconnected) lemma path_connected_space_imp_path_component_of: "\path_connected_space X; a \ topspace X; b \ topspace X\ \ path_component_of X a b" by (simp add: path_connected_space_iff_path_component) lemma path_connected_space_path_component_set: "path_connected_space X \ (\x \ topspace X. Collect(path_component_of X x) = topspace X)" using path_component_of_subset_topspace path_connected_space_iff_path_component by fastforce lemma path_component_of_maximal: "\path_connectedin X s; x \ s\ \ s \ Collect(path_component_of X x)" using path_component_of by fastforce lemma path_component_of_equiv: "path_component_of X x y \ x \ topspace X \ y \ topspace X \ path_component_of X x = path_component_of X y" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply (simp add: fun_eq_iff path_component_in_topspace) apply (meson path_component_of_sym path_component_of_trans) done qed (simp add: path_component_of_refl) lemma path_component_of_disjoint: "disjnt (Collect (path_component_of X x)) (Collect (path_component_of X y)) \ ~(path_component_of X x y)" by (force simp: disjnt_def path_component_of_eq_empty path_component_of_equiv) lemma path_component_of_eq: "path_component_of X x = path_component_of X y \ (x \ topspace X) \ (y \ topspace X) \ x \ topspace X \ y \ topspace X \ path_component_of X x y" by (metis Collect_empty_eq_bot path_component_of_eq_empty path_component_of_equiv) lemma path_component_of_aux: "path_component_of X x y \ path_component_of (subtopology X (Collect (path_component_of X x))) x y" by (meson path_component_of path_component_of_maximal path_connectedin_subtopology) lemma path_connectedin_path_component_of: "path_connectedin X (Collect (path_component_of X x))" proof - have "topspace (subtopology X (path_component_of_set X x)) = path_component_of_set X x" by (meson path_component_of_subset_topspace topspace_subtopology_subset) then have "path_connected_space (subtopology X (path_component_of_set X x))" by (metis (full_types) path_component_of_aux mem_Collect_eq path_component_of_equiv path_connected_space_iff_path_component) then show ?thesis by (simp add: path_component_of_subset_topspace path_connectedin_def) qed lemma path_connectedin_euclidean [simp]: "path_connectedin euclidean S \ path_connected S" by (auto simp: path_connectedin_def path_connected_space_iff_path_component path_connected_component) lemma path_connected_space_euclidean_subtopology [simp]: "path_connected_space(subtopology euclidean S) \ path_connected S" using path_connectedin_topspace by force lemma Union_path_components_of: "\(path_components_of X) = topspace X" by (auto simp: path_components_of_def path_component_of_equiv) lemma path_components_of_maximal: "\C \ path_components_of X; path_connectedin X S; ~disjnt C S\ \ S \ C" apply (auto simp: path_components_of_def path_component_of_equiv) using path_component_of_maximal path_connectedin_def apply fastforce by (meson disjnt_subset2 path_component_of_disjoint path_component_of_equiv path_component_of_maximal) lemma pairwise_disjoint_path_components_of: "pairwise disjnt (path_components_of X)" by (auto simp: path_components_of_def pairwise_def path_component_of_disjoint path_component_of_equiv) lemma complement_path_components_of_Union: "C \ path_components_of X \ topspace X - C = \(path_components_of X - {C})" by (metis Diff_cancel Diff_subset Union_path_components_of cSup_singleton diff_Union_pairwise_disjoint insert_subset pairwise_disjoint_path_components_of) lemma nonempty_path_components_of: assumes "C \ path_components_of X" shows "C \ {}" proof - have "C \ path_component_of_set X ` topspace X" using assms path_components_of_def by blast then show ?thesis using path_component_of_refl by fastforce qed lemma path_components_of_subset: "C \ path_components_of X \ C \ topspace X" by (auto simp: path_components_of_def path_component_of_equiv) lemma path_connectedin_path_components_of: "C \ path_components_of X \ path_connectedin X C" by (auto simp: path_components_of_def path_connectedin_path_component_of) lemma path_component_in_path_components_of: "Collect (path_component_of X a) \ path_components_of X \ a \ topspace X" by (metis imageI nonempty_path_components_of path_component_of_eq_empty path_components_of_def) lemma path_connectedin_Union: assumes \: "\S. S \ \ \ path_connectedin X S" "\\ \ {}" shows "path_connectedin X (\\)" proof - obtain a where "\S. S \ \ \ a \ S" using assms by blast then have "\x. x \ topspace (subtopology X (\\)) \ path_component_of (subtopology X (\\)) a x" by simp (meson Union_upper \ path_component_of path_connectedin_subtopology) then show ?thesis using \ unfolding path_connectedin_def by (metis Sup_le_iff path_component_of_equiv path_connected_space_iff_path_component) qed lemma path_connectedin_Un: "\path_connectedin X S; path_connectedin X T; S \ T \ {}\ \ path_connectedin X (S \ T)" by (blast intro: path_connectedin_Union [of "{S,T}", simplified]) lemma path_connected_space_iff_components_eq: "path_connected_space X \ (\C \ path_components_of X. \C' \ path_components_of X. C = C')" unfolding path_components_of_def proof (intro iffI ballI) assume "\C \ path_component_of_set X ` topspace X. \C' \ path_component_of_set X ` topspace X. C = C'" then show "path_connected_space X" using path_component_of_refl path_connected_space_iff_path_component by fastforce qed (auto simp: path_connected_space_path_component_set) lemma path_components_of_eq_empty: "path_components_of X = {} \ topspace X = {}" using Union_path_components_of nonempty_path_components_of by fastforce lemma path_components_of_empty_space: "topspace X = {} \ path_components_of X = {}" by (simp add: path_components_of_eq_empty) lemma path_components_of_subset_singleton: "path_components_of X \ {S} \ path_connected_space X \ (topspace X = {} \ topspace X = S)" proof (cases "topspace X = {}") case True then show ?thesis by (auto simp: path_components_of_empty_space path_connected_space_topspace_empty) next case False have "(path_components_of X = {S}) \ (path_connected_space X \ topspace X = S)" proof (intro iffI conjI) assume L: "path_components_of X = {S}" then show "path_connected_space X" by (simp add: path_connected_space_iff_components_eq) show "topspace X = S" by (metis L ccpo_Sup_singleton [of S] Union_path_components_of) next assume R: "path_connected_space X \ topspace X = S" then show "path_components_of X = {S}" using ccpo_Sup_singleton [of S] by (metis False all_not_in_conv insert_iff mk_disjoint_insert path_component_in_path_components_of path_connected_space_iff_components_eq path_connected_space_path_component_set) qed with False show ?thesis by (simp add: path_components_of_eq_empty subset_singleton_iff) qed lemma path_connected_space_iff_components_subset_singleton: "path_connected_space X \ (\a. path_components_of X \ {a})" by (simp add: path_components_of_subset_singleton) lemma path_components_of_eq_singleton: "path_components_of X = {S} \ path_connected_space X \ topspace X \ {} \ S = topspace X" by (metis cSup_singleton insert_not_empty path_components_of_subset_singleton subset_singleton_iff) lemma path_components_of_path_connected_space: "path_connected_space X \ path_components_of X = (if topspace X = {} then {} else {topspace X})" by (simp add: path_components_of_eq_empty path_components_of_eq_singleton) lemma path_component_subset_connected_component_of: "path_component_of_set X x \ connected_component_of_set X x" proof (cases "x \ topspace X") case True then show ?thesis by (simp add: connected_component_of_maximal path_component_of_refl path_connectedin_imp_connectedin path_connectedin_path_component_of) next case False then show ?thesis using path_component_of_eq_empty by fastforce qed lemma exists_path_component_of_superset: assumes S: "path_connectedin X S" and ne: "topspace X \ {}" obtains C where "C \ path_components_of X" "S \ C" proof (cases "S = {}") case True then show ?thesis using ne path_components_of_eq_empty that by fastforce next case False then obtain a where "a \ S" by blast show ?thesis proof show "Collect (path_component_of X a) \ path_components_of X" by (meson \a \ S\ S subsetD path_component_in_path_components_of path_connectedin_subset_topspace) show "S \ Collect (path_component_of X a)" by (simp add: S \a \ S\ path_component_of_maximal) qed qed lemma path_component_of_eq_overlap: "path_component_of X x = path_component_of X y \ (x \ topspace X) \ (y \ topspace X) \ Collect (path_component_of X x) \ Collect (path_component_of X y) \ {}" by (metis disjnt_def empty_iff inf_bot_right mem_Collect_eq path_component_of_disjoint path_component_of_eq path_component_of_eq_empty) lemma path_component_of_nonoverlap: "Collect (path_component_of X x) \ Collect (path_component_of X y) = {} \ (x \ topspace X) \ (y \ topspace X) \ path_component_of X x \ path_component_of X y" by (metis inf.idem path_component_of_eq_empty path_component_of_eq_overlap) lemma path_component_of_overlap: "Collect (path_component_of X x) \ Collect (path_component_of X y) \ {} \ x \ topspace X \ y \ topspace X \ path_component_of X x = path_component_of X y" by (meson path_component_of_nonoverlap) lemma path_components_of_disjoint: "\C \ path_components_of X; C' \ path_components_of X\ \ disjnt C C' \ C \ C'" by (auto simp: path_components_of_def path_component_of_disjoint path_component_of_equiv) lemma path_components_of_overlap: "\C \ path_components_of X; C' \ path_components_of X\ \ C \ C' \ {} \ C = C'" by (auto simp: path_components_of_def path_component_of_equiv) lemma path_component_of_unique: "\x \ C; path_connectedin X C; \C'. \x \ C'; path_connectedin X C'\ \ C' \ C\ \ Collect (path_component_of X x) = C" by (meson subsetD eq_iff path_component_of_maximal path_connectedin_path_component_of) lemma path_component_of_discrete_topology [simp]: "Collect (path_component_of (discrete_topology U) x) = (if x \ U then {x} else {})" proof - have "\C'. \x \ C'; path_connectedin (discrete_topology U) C'\ \ C' \ {x}" by (metis path_connectedin_discrete_topology subsetD singletonD) then have "x \ U \ Collect (path_component_of (discrete_topology U) x) = {x}" by (simp add: path_component_of_unique) then show ?thesis using path_component_in_topspace by fastforce qed lemma path_component_of_discrete_topology_iff [simp]: "path_component_of (discrete_topology U) x y \ x \ U \ y=x" by (metis empty_iff insertI1 mem_Collect_eq path_component_of_discrete_topology singletonD) lemma path_components_of_discrete_topology [simp]: "path_components_of (discrete_topology U) = (\x. {x}) ` U" by (auto simp: path_components_of_def image_def fun_eq_iff) lemma homeomorphic_map_path_component_of: assumes f: "homeomorphic_map X Y f" and x: "x \ topspace X" shows "Collect (path_component_of Y (f x)) = f ` Collect(path_component_of X x)" proof - obtain g where g: "homeomorphic_maps X Y f g" using f homeomorphic_map_maps by blast show ?thesis proof have "Collect (path_component_of Y (f x)) \ topspace Y" by (simp add: path_component_of_subset_topspace) moreover have "g ` Collect(path_component_of Y (f x)) \ Collect (path_component_of X (g (f x)))" using g x unfolding homeomorphic_maps_def by (metis f homeomorphic_imp_surjective_map imageI mem_Collect_eq path_component_of_maximal path_component_of_refl path_connectedin_continuous_map_image path_connectedin_path_component_of) ultimately show "Collect (path_component_of Y (f x)) \ f ` Collect (path_component_of X x)" using g x unfolding homeomorphic_maps_def continuous_map_def image_iff subset_iff by metis show "f ` Collect (path_component_of X x) \ Collect (path_component_of Y (f x))" proof (rule path_component_of_maximal) show "path_connectedin Y (f ` Collect (path_component_of X x))" by (meson f homeomorphic_map_path_connectedness_eq path_connectedin_path_component_of) qed (simp add: path_component_of_refl x) qed qed lemma homeomorphic_map_path_components_of: assumes "homeomorphic_map X Y f" shows "path_components_of Y = (image f) ` (path_components_of X)" (is "?lhs = ?rhs") unfolding path_components_of_def homeomorphic_imp_surjective_map [OF assms, symmetric] using assms homeomorphic_map_path_component_of by fastforce subsection \Sphere is path-connected\ lemma path_connected_punctured_universe: assumes "2 \ DIM('a::euclidean_space)" shows "path_connected (- {a::'a})" proof - let ?A = "{x::'a. \i\Basis. x \ i < a \ i}" let ?B = "{x::'a. \i\Basis. a \ i < x \ i}" have A: "path_connected ?A" unfolding Collect_bex_eq proof (rule path_connected_UNION) fix i :: 'a assume "i \ Basis" then show "(\i\Basis. (a \ i - 1)*\<^sub>R i) \ {x::'a. x \ i < a \ i}" by simp show "path_connected {x. x \ i < a \ i}" using convex_imp_path_connected [OF convex_halfspace_lt, of i "a \ i"] by (simp add: inner_commute) qed have B: "path_connected ?B" unfolding Collect_bex_eq proof (rule path_connected_UNION) fix i :: 'a assume "i \ Basis" then show "(\i\Basis. (a \ i + 1) *\<^sub>R i) \ {x::'a. a \ i < x \ i}" by simp show "path_connected {x. a \ i < x \ i}" using convex_imp_path_connected [OF convex_halfspace_gt, of "a \ i" i] by (simp add: inner_commute) qed obtain S :: "'a set" where "S \ Basis" and "card S = Suc (Suc 0)" using ex_card[OF assms] by auto then obtain b0 b1 :: 'a where "b0 \ Basis" and "b1 \ Basis" and "b0 \ b1" unfolding card_Suc_eq by auto then have "a + b0 - b1 \ ?A \ ?B" by (auto simp: inner_simps inner_Basis) then have "?A \ ?B \ {}" by fast with A B have "path_connected (?A \ ?B)" by (rule path_connected_Un) also have "?A \ ?B = {x. \i\Basis. x \ i \ a \ i}" unfolding neq_iff bex_disj_distrib Collect_disj_eq .. also have "\ = {x. x \ a}" unfolding euclidean_eq_iff [where 'a='a] by (simp add: Bex_def) also have "\ = - {a}" by auto finally show ?thesis . qed corollary connected_punctured_universe: "2 \ DIM('N::euclidean_space) \ connected(- {a::'N})" by (simp add: path_connected_punctured_universe path_connected_imp_connected) proposition path_connected_sphere: fixes a :: "'a :: euclidean_space" assumes "2 \ DIM('a)" shows "path_connected(sphere a r)" proof (cases r "0::real" rule: linorder_cases) case less then show ?thesis by (simp) next case equal then show ?thesis by (simp) next case greater then have eq: "(sphere (0::'a) r) = (\x. (r / norm x) *\<^sub>R x) ` (- {0::'a})" by (force simp: image_iff split: if_split_asm) have "continuous_on (- {0::'a}) (\x. (r / norm x) *\<^sub>R x)" by (intro continuous_intros) auto then have "path_connected ((\x. (r / norm x) *\<^sub>R x) ` (- {0::'a}))" by (intro path_connected_continuous_image path_connected_punctured_universe assms) with eq have "path_connected (sphere (0::'a) r)" by auto then have "path_connected((+) a ` (sphere (0::'a) r))" by (simp add: path_connected_translation) then show ?thesis by (metis add.right_neutral sphere_translation) qed lemma connected_sphere: fixes a :: "'a :: euclidean_space" assumes "2 \ DIM('a)" shows "connected(sphere a r)" using path_connected_sphere [OF assms] by (simp add: path_connected_imp_connected) corollary path_connected_complement_bounded_convex: fixes S :: "'a :: euclidean_space set" assumes "bounded S" "convex S" and 2: "2 \ DIM('a)" shows "path_connected (- S)" proof (cases "S = {}") case True then show ?thesis using convex_imp_path_connected by auto next case False then obtain a where "a \ S" by auto have \
[rule_format]: "\y\S. \u. 0 \ u \ u \ 1 \ (1 - u) *\<^sub>R a + u *\<^sub>R y \ S" using \convex S\ \a \ S\ by (simp add: convex_alt) { fix x y assume "x \ S" "y \ S" then have "x \ a" "y \ a" using \a \ S\ by auto then have bxy: "bounded(insert x (insert y S))" by (simp add: \bounded S\) then obtain B::real where B: "0 < B" and Bx: "norm (a - x) < B" and By: "norm (a - y) < B" and "S \ ball a B" using bounded_subset_ballD [OF bxy, of a] by (auto simp: dist_norm) define C where "C = B / norm(x - a)" let ?Cxa = "a + C *\<^sub>R (x - a)" { fix u assume u: "(1 - u) *\<^sub>R x + u *\<^sub>R ?Cxa \ S" and "0 \ u" "u \ 1" have CC: "1 \ 1 + (C - 1) * u" using \x \ a\ \0 \ u\ Bx by (auto simp add: C_def norm_minus_commute) have *: "\v. (1 - u) *\<^sub>R x + u *\<^sub>R (a + v *\<^sub>R (x - a)) = a + (1 + (v - 1) * u) *\<^sub>R (x - a)" by (simp add: algebra_simps) have "a + ((1 / (1 + C * u - u)) *\<^sub>R x + ((u / (1 + C * u - u)) *\<^sub>R a + (C * u / (1 + C * u - u)) *\<^sub>R x)) = (1 + (u / (1 + C * u - u))) *\<^sub>R a + ((1 / (1 + C * u - u)) + (C * u / (1 + C * u - u))) *\<^sub>R x" by (simp add: algebra_simps) also have "\ = (1 + (u / (1 + C * u - u))) *\<^sub>R a + (1 + (u / (1 + C * u - u))) *\<^sub>R x" using CC by (simp add: field_simps) also have "\ = x + (1 + (u / (1 + C * u - u))) *\<^sub>R a + (u / (1 + C * u - u)) *\<^sub>R x" by (simp add: algebra_simps) also have "\ = x + ((1 / (1 + C * u - u)) *\<^sub>R a + ((u / (1 + C * u - u)) *\<^sub>R x + (C * u / (1 + C * u - u)) *\<^sub>R a))" using CC by (simp add: field_simps) (simp add: add_divide_distrib scaleR_add_left) finally have xeq: "(1 - 1 / (1 + (C - 1) * u)) *\<^sub>R a + (1 / (1 + (C - 1) * u)) *\<^sub>R (a + (1 + (C - 1) * u) *\<^sub>R (x - a)) = x" by (simp add: algebra_simps) have False using \
[of "a + (1 + (C - 1) * u) *\<^sub>R (x - a)" "1 / (1 + (C - 1) * u)"] using u \x \ a\ \x \ S\ \0 \ u\ CC by (auto simp: xeq *) } then have pcx: "path_component (- S) x ?Cxa" by (force simp: closed_segment_def intro!: path_component_linepath) define D where "D = B / norm(y - a)" \ \massive duplication with the proof above\ let ?Dya = "a + D *\<^sub>R (y - a)" { fix u assume u: "(1 - u) *\<^sub>R y + u *\<^sub>R ?Dya \ S" and "0 \ u" "u \ 1" have DD: "1 \ 1 + (D - 1) * u" using \y \ a\ \0 \ u\ By by (auto simp add: D_def norm_minus_commute) have *: "\v. (1 - u) *\<^sub>R y + u *\<^sub>R (a + v *\<^sub>R (y - a)) = a + (1 + (v - 1) * u) *\<^sub>R (y - a)" by (simp add: algebra_simps) have "a + ((1 / (1 + D * u - u)) *\<^sub>R y + ((u / (1 + D * u - u)) *\<^sub>R a + (D * u / (1 + D * u - u)) *\<^sub>R y)) = (1 + (u / (1 + D * u - u))) *\<^sub>R a + ((1 / (1 + D * u - u)) + (D * u / (1 + D * u - u))) *\<^sub>R y" by (simp add: algebra_simps) also have "\ = (1 + (u / (1 + D * u - u))) *\<^sub>R a + (1 + (u / (1 + D * u - u))) *\<^sub>R y" using DD by (simp add: field_simps) also have "\ = y + (1 + (u / (1 + D * u - u))) *\<^sub>R a + (u / (1 + D * u - u)) *\<^sub>R y" by (simp add: algebra_simps) also have "\ = y + ((1 / (1 + D * u - u)) *\<^sub>R a + ((u / (1 + D * u - u)) *\<^sub>R y + (D * u / (1 + D * u - u)) *\<^sub>R a))" using DD by (simp add: field_simps) (simp add: add_divide_distrib scaleR_add_left) finally have xeq: "(1 - 1 / (1 + (D - 1) * u)) *\<^sub>R a + (1 / (1 + (D - 1) * u)) *\<^sub>R (a + (1 + (D - 1) * u) *\<^sub>R (y - a)) = y" by (simp add: algebra_simps) have False using \
[of "a + (1 + (D - 1) * u) *\<^sub>R (y - a)" "1 / (1 + (D - 1) * u)"] using u \y \ a\ \y \ S\ \0 \ u\ DD by (auto simp: xeq *) } then have pdy: "path_component (- S) y ?Dya" by (force simp: closed_segment_def intro!: path_component_linepath) have pyx: "path_component (- S) ?Dya ?Cxa" proof (rule path_component_of_subset) show "sphere a B \ - S" using \S \ ball a B\ by (force simp: ball_def dist_norm norm_minus_commute) have aB: "?Dya \ sphere a B" "?Cxa \ sphere a B" using \x \ a\ using \y \ a\ B by (auto simp: dist_norm C_def D_def) then show "path_component (sphere a B) ?Dya ?Cxa" using path_connected_sphere [OF 2] path_connected_component by blast qed have "path_component (- S) x y" by (metis path_component_trans path_component_sym pcx pdy pyx) } then show ?thesis by (auto simp: path_connected_component) qed lemma connected_complement_bounded_convex: fixes S :: "'a :: euclidean_space set" assumes "bounded S" "convex S" "2 \ DIM('a)" shows "connected (- S)" using path_connected_complement_bounded_convex [OF assms] path_connected_imp_connected by blast lemma connected_diff_ball: fixes S :: "'a :: euclidean_space set" assumes "connected S" "cball a r \ S" "2 \ DIM('a)" shows "connected (S - ball a r)" proof (rule connected_diff_open_from_closed [OF ball_subset_cball]) show "connected (cball a r - ball a r)" using assms connected_sphere by (auto simp: cball_diff_eq_sphere) qed (auto simp: assms dist_norm) proposition connected_open_delete: assumes "open S" "connected S" and 2: "2 \ DIM('N::euclidean_space)" shows "connected(S - {a::'N})" proof (cases "a \ S") case True with \open S\ obtain \ where "\ > 0" and \: "cball a \ \ S" using open_contains_cball_eq by blast define b where "b \ a + \ *\<^sub>R (SOME i. i \ Basis)" have "dist a b = \" by (simp add: b_def dist_norm SOME_Basis \0 < \\ less_imp_le) with \ have "b \ \{S - ball a r |r. 0 < r \ r < \}" by auto then have nonemp: "(\{S - ball a r |r. 0 < r \ r < \}) = {} \ False" by auto have con: "\r. r < \ \ connected (S - ball a r)" using \ by (force intro: connected_diff_ball [OF \connected S\ _ 2]) have "x \ \{S - ball a r |r. 0 < r \ r < \}" if "x \ S - {a}" for x using that \0 < \\ by (intro UnionI [of "S - ball a (min \ (dist a x) / 2)"]) auto then have "S - {a} = \{S - ball a r | r. 0 < r \ r < \}" by auto then show ?thesis by (auto intro: connected_Union con dest!: nonemp) next case False then show ?thesis by (simp add: \connected S\) qed corollary path_connected_open_delete: assumes "open S" "connected S" and 2: "2 \ DIM('N::euclidean_space)" shows "path_connected(S - {a::'N})" by (simp add: assms connected_open_delete connected_open_path_connected open_delete) corollary path_connected_punctured_ball: "2 \ DIM('N::euclidean_space) \ path_connected(ball a r - {a::'N})" by (simp add: path_connected_open_delete) corollary connected_punctured_ball: "2 \ DIM('N::euclidean_space) \ connected(ball a r - {a::'N})" by (simp add: connected_open_delete) corollary connected_open_delete_finite: fixes S T::"'a::euclidean_space set" assumes S: "open S" "connected S" and 2: "2 \ DIM('a)" and "finite T" shows "connected(S - T)" using \finite T\ S proof (induct T) case empty show ?case using \connected S\ by simp next case (insert x F) then have "connected (S-F)" by auto moreover have "open (S - F)" using finite_imp_closed[OF \finite F\] \open S\ by auto ultimately have "connected (S - F - {x})" using connected_open_delete[OF _ _ 2] by auto thus ?case by (metis Diff_insert) qed lemma sphere_1D_doubleton_zero: assumes 1: "DIM('a) = 1" and "r > 0" obtains x y::"'a::euclidean_space" where "sphere 0 r = {x,y} \ dist x y = 2*r" proof - obtain b::'a where b: "Basis = {b}" using 1 card_1_singletonE by blast show ?thesis proof (intro that conjI) have "x = norm x *\<^sub>R b \ x = - norm x *\<^sub>R b" if "r = norm x" for x proof - have xb: "(x \ b) *\<^sub>R b = x" using euclidean_representation [of x, unfolded b] by force then have "norm ((x \ b) *\<^sub>R b) = norm x" by simp with b have "\x \ b\ = norm x" using norm_Basis by (simp add: b) with xb show ?thesis by (metis (mono_tags, 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" proof - have "sphere a r = (+) a ` sphere 0 r" by (metis add.right_neutral sphere_translation) then show ?thesis using sphere_1D_doubleton_zero [OF assms] by (metis (mono_tags, lifting) dist_add_cancel image_empty image_insert that) qed lemma psubset_sphere_Compl_connected: fixes S :: "'a::euclidean_space set" assumes S: "S \ sphere a r" and "0 < r" and 2: "2 \ DIM('a)" shows "connected(- S)" proof - have "S \ sphere a r" using S by blast obtain b where "dist a b = r" and "b \ S" using S mem_sphere by blast have CS: "- S = {x. dist a x \ r \ (x \ S)} \ {x. r \ dist a x \ (x \ S)}" by auto have "{x. dist a x \ r \ x \ S} \ {x. r \ dist a x \ x \ S} \ {}" using \b \ S\ \dist a b = r\ by blast moreover have "connected {x. dist a x \ r \ x \ S}" using assms by (force intro: connected_intermediate_closure [of "ball a r"]) moreover have "connected {x. r \ dist a x \ x \ S}" proof (rule connected_intermediate_closure [of "- cball a r"]) show "{x. r \ dist a x \ x \ S} \ closure (- cball a r)" using interior_closure by (force intro: connected_complement_bounded_convex) qed (use assms connected_complement_bounded_convex in auto) ultimately show ?thesis by (simp add: CS connected_Un) qed subsection\Every annulus is a connected set\ lemma path_connected_2DIM_I: fixes a :: "'N::euclidean_space" assumes 2: "2 \ DIM('N)" and pc: "path_connected {r. 0 \ r \ P r}" shows "path_connected {x. P(norm(x - a))}" proof - have "{x. P(norm(x - a))} = (+) a ` {x. P(norm x)}" by force moreover have "path_connected {x::'N. P(norm x)}" proof - let ?D = "{x. 0 \ x \ P x} \ sphere (0::'N) 1" have "x \ (\z. fst z *\<^sub>R snd z) ` ?D" if "P (norm x)" for x::'N proof (cases "x=0") case True with that show ?thesis apply (simp add: image_iff) by (metis (no_types) mem_sphere_0 order_refl vector_choose_size zero_le_one) next case False with that show ?thesis by (rule_tac x="(norm x, x /\<^sub>R norm x)" in image_eqI) auto qed then have *: "{x::'N. P(norm x)} = (\z. fst z *\<^sub>R snd z) ` ?D" by auto have "continuous_on ?D (\z:: real\'N. fst z *\<^sub>R snd z)" by (intro continuous_intros) moreover have "path_connected ?D" by (metis path_connected_Times [OF pc] path_connected_sphere 2) ultimately show ?thesis by (simp add: "*" path_connected_continuous_image) qed ultimately show ?thesis using path_connected_translation by metis qed proposition path_connected_annulus: fixes a :: "'N::euclidean_space" assumes "2 \ DIM('N)" shows "path_connected {x. r1 < norm(x - a) \ norm(x - a) < r2}" "path_connected {x. r1 < norm(x - a) \ norm(x - a) \ r2}" "path_connected {x. r1 \ norm(x - a) \ norm(x - a) < r2}" "path_connected {x. r1 \ norm(x - a) \ norm(x - a) \ r2}" by (auto simp: is_interval_def intro!: is_interval_convex convex_imp_path_connected path_connected_2DIM_I [OF assms]) proposition connected_annulus: fixes a :: "'N::euclidean_space" assumes "2 \ DIM('N::euclidean_space)" shows "connected {x. r1 < norm(x - a) \ norm(x - a) < r2}" "connected {x. r1 < norm(x - a) \ norm(x - a) \ r2}" "connected {x. r1 \ norm(x - a) \ norm(x - a) < r2}" "connected {x. r1 \ norm(x - a) \ norm(x - a) \ r2}" by (auto simp: path_connected_annulus [OF assms] path_connected_imp_connected) subsection\<^marker>\tag unimportant\\Relations between components and path components\ lemma open_connected_component: fixes S :: "'a::real_normed_vector set" assumes "open S" shows "open (connected_component_set S x)" proof (clarsimp simp: open_contains_ball) fix y assume xy: "connected_component S x y" then obtain e where "e>0" "ball y e \ S" using assms connected_component_in openE by blast then show "\e>0. ball y e \ connected_component_set S x" by (metis xy centre_in_ball connected_ball connected_component_eq_eq connected_component_in connected_component_maximal) qed corollary open_components: fixes S :: "'a::real_normed_vector set" shows "\open u; S \ components u\ \ open S" by (simp add: components_iff) (metis open_connected_component) lemma in_closure_connected_component: fixes S :: "'a::real_normed_vector set" assumes x: "x \ S" and S: "open S" shows "x \ closure (connected_component_set S y) \ x \ connected_component_set S y" proof - { assume "x \ closure (connected_component_set S y)" moreover have "x \ connected_component_set S x" using x by simp ultimately have "x \ connected_component_set S y" using S by (meson Compl_disjoint closure_iff_nhds_not_empty connected_component_disjoint disjoint_eq_subset_Compl open_connected_component) } then show ?thesis by (auto simp: closure_def) qed lemma connected_disjoint_Union_open_pick: assumes "pairwise disjnt B" "\S. S \ A \ connected S \ S \ {}" "\S. S \ B \ open S" "\A \ \B" "S \ A" obtains T where "T \ B" "S \ T" "S \ \(B - {T}) = {}" proof - have "S \ \B" "connected S" "S \ {}" using assms \S \ A\ by blast+ then obtain T where "T \ B" "S \ T \ {}" by (metis Sup_inf_eq_bot_iff inf.absorb_iff2 inf_commute) have 1: "open T" by (simp add: \T \ B\ assms) have 2: "open (\(B-{T}))" using assms by blast have 3: "S \ T \ \(B - {T})" using \S \ \B\ by blast have "T \ \(B - {T}) = {}" using \T \ B\ \pairwise disjnt B\ by (auto simp: pairwise_def disjnt_def) then have 4: "T \ \(B - {T}) \ S = {}" by auto from connectedD [OF \connected S\ 1 2 4 3] have "S \ \(B-{T}) = {}" by (auto simp: Int_commute \S \ T \ {}\) with \T \ B\ have "S \ T" using "3" by auto show ?thesis using \S \ \(B - {T}) = {}\ \S \ T\ \T \ B\ that by auto qed lemma connected_disjoint_Union_open_subset: assumes A: "pairwise disjnt A" and B: "pairwise disjnt B" and SA: "\S. S \ A \ open S \ connected S \ S \ {}" and SB: "\S. S \ B \ open S \ connected S \ S \ {}" and eq [simp]: "\A = \B" shows "A \ B" proof fix S assume "S \ A" obtain T where "T \ B" "S \ T" "S \ \(B - {T}) = {}" using SA SB \S \ A\ connected_disjoint_Union_open_pick [OF B, of A] eq order_refl by blast moreover obtain S' where "S' \ A" "T \ S'" "T \ \(A - {S'}) = {}" using SA SB \T \ B\ connected_disjoint_Union_open_pick [OF A, of B] eq order_refl by blast ultimately have "S' = S" by (metis A Int_subset_iff SA \S \ A\ disjnt_def inf.orderE pairwise_def) with \T \ S'\ have "T \ S" by simp with \S \ T\ have "S = T" by blast with \T \ B\ show "S \ B" by simp qed lemma connected_disjoint_Union_open_unique: assumes A: "pairwise disjnt A" and B: "pairwise disjnt B" and SA: "\S. S \ A \ open S \ connected S \ S \ {}" and SB: "\S. S \ B \ open S \ connected S \ S \ {}" and eq [simp]: "\A = \B" shows "A = B" by (rule subset_antisym; metis connected_disjoint_Union_open_subset assms) proposition components_open_unique: fixes S :: "'a::real_normed_vector set" assumes "pairwise disjnt A" "\A = S" "\X. X \ A \ open X \ connected X \ X \ {}" shows "components S = A" proof - have "open S" using assms by blast show ?thesis proof (rule connected_disjoint_Union_open_unique) show "disjoint (components S)" by (simp add: components_eq disjnt_def pairwise_def) qed (use \open S\ in \simp_all add: assms open_components in_components_connected in_components_nonempty\) qed subsection\<^marker>\tag unimportant\\Existence of unbounded components\ lemma cobounded_unbounded_component: fixes S :: "'a :: euclidean_space set" assumes "bounded (-S)" shows "\x. x \ S \ \ bounded (connected_component_set S x)" proof - obtain i::'a where i: "i \ Basis" using nonempty_Basis by blast obtain B where B: "B>0" "-S \ ball 0 B" using bounded_subset_ballD [OF assms, of 0] by auto then have *: "\x. B \ norm x \ x \ S" by (force simp: ball_def dist_norm) have unbounded_inner: "\ bounded {x. inner i x \ B}" proof (clarsimp simp: bounded_def dist_norm) fix e x show "\y. B \ i \ y \ \ norm (x - y) \ e" using i by (rule_tac x="x + (max B e + 1 + \i \ x\) *\<^sub>R i" in exI) (auto simp: inner_right_distrib) qed have \
: "\x. B \ i \ x \ x \ S" using * Basis_le_norm [OF i] by (metis abs_ge_self inner_commute order_trans) have "{x. B \ i \ x} \ connected_component_set S (B *\<^sub>R i)" by (intro connected_component_maximal) (auto simp: i intro: convex_connected convex_halfspace_ge [of B] \
) then have "\ bounded (connected_component_set S (B *\<^sub>R i))" using bounded_subset unbounded_inner by blast moreover have "B *\<^sub>R i \ S" by (rule *) (simp add: norm_Basis [OF i]) ultimately show ?thesis by blast qed lemma cobounded_unique_unbounded_component: fixes S :: "'a :: euclidean_space set" assumes bs: "bounded (-S)" and "2 \ DIM('a)" and bo: "\ bounded(connected_component_set S x)" "\ bounded(connected_component_set S y)" shows "connected_component_set S x = connected_component_set S y" proof - obtain i::'a where i: "i \ Basis" using nonempty_Basis by blast obtain B where B: "B>0" "-S \ ball 0 B" using bounded_subset_ballD [OF bs, of 0] by auto then have *: "\x. B \ norm x \ x \ S" by (force simp: ball_def dist_norm) obtain x' where x': "connected_component S x x'" "norm x' > B" using bo [unfolded bounded_def dist_norm, simplified, rule_format] by (metis diff_zero norm_minus_commute not_less) obtain y' where y': "connected_component S y y'" "norm y' > B" using bo [unfolded bounded_def dist_norm, simplified, rule_format] by (metis diff_zero norm_minus_commute not_less) have x'y': "connected_component S x' y'" unfolding connected_component_def proof (intro exI conjI) show "connected (- ball 0 B :: 'a set)" using assms by (auto intro: connected_complement_bounded_convex) qed (use x' y' dist_norm * in auto) show ?thesis proof (rule connected_component_eq) show "x \ connected_component_set S y" using x' y' x'y' by (metis (no_types) connected_component_eq_eq connected_component_in mem_Collect_eq) qed qed lemma cobounded_unbounded_components: fixes S :: "'a :: euclidean_space set" shows "bounded (-S) \ \c. c \ components S \ \bounded c" by (metis cobounded_unbounded_component components_def imageI) lemma cobounded_unique_unbounded_components: fixes S :: "'a :: euclidean_space set" shows "\bounded (- S); c \ components S; \ bounded c; c' \ components S; \ bounded c'; 2 \ DIM('a)\ \ c' = c" unfolding components_iff by (metis cobounded_unique_unbounded_component) lemma cobounded_has_bounded_component: fixes S :: "'a :: euclidean_space set" assumes "bounded (- S)" "\ connected S" "2 \ DIM('a)" obtains C where "C \ components S" "bounded C" by (meson cobounded_unique_unbounded_components connected_eq_connected_components_eq assms) subsection\The \inside\ and \outside\ of a Set\ text\<^marker>\tag important\\The inside comprises the points in a bounded connected component of the set's complement. The outside comprises the points in unbounded connected component of the complement.\ definition\<^marker>\tag important\ inside where "inside S \ {x. (x \ S) \ bounded(connected_component_set ( - S) x)}" definition\<^marker>\tag important\ outside where "outside S \ -S \ {x. \ bounded(connected_component_set (- S) x)}" lemma outside: "outside S = {x. \ bounded(connected_component_set (- S) x)}" by (auto simp: outside_def) (metis Compl_iff bounded_empty connected_component_eq_empty) lemma inside_no_overlap [simp]: "inside S \ S = {}" by (auto simp: inside_def) lemma outside_no_overlap [simp]: "outside S \ S = {}" by (auto simp: outside_def) lemma inside_Int_outside [simp]: "inside S \ outside S = {}" by (auto simp: inside_def outside_def) lemma inside_Un_outside [simp]: "inside S \ outside S = (- S)" by (auto simp: inside_def outside_def) lemma inside_eq_outside: "inside S = outside S \ S = UNIV" by (auto simp: inside_def outside_def) lemma inside_outside: "inside S = (- (S \ outside S))" by (force simp: inside_def outside) lemma outside_inside: "outside S = (- (S \ inside S))" by (auto simp: inside_outside) (metis IntI equals0D outside_no_overlap) lemma union_with_inside: "S \ inside S = - outside S" by (auto simp: inside_outside) (simp add: outside_inside) lemma union_with_outside: "S \ outside S = - inside S" by (simp add: inside_outside) lemma outside_mono: "S \ T \ outside T \ outside S" by (auto simp: outside bounded_subset connected_component_mono) lemma inside_mono: "S \ T \ inside S - T \ inside T" by (auto simp: inside_def bounded_subset connected_component_mono) lemma segment_bound_lemma: fixes u::real assumes "x \ B" "y \ B" "0 \ u" "u \ 1" shows "(1 - u) * x + u * y \ B" proof - obtain dx dy where "dx \ 0" "dy \ 0" "x = B + dx" "y = B + dy" using assms by auto (metis add.commute diff_add_cancel) with \0 \ u\ \u \ 1\ show ?thesis by (simp add: add_increasing2 mult_left_le field_simps) qed lemma cobounded_outside: fixes S :: "'a :: real_normed_vector set" assumes "bounded S" shows "bounded (- outside S)" proof - obtain B where B: "B>0" "S \ ball 0 B" using bounded_subset_ballD [OF assms, of 0] by auto { fix x::'a and C::real assume Bno: "B \ norm x" and C: "0 < C" have "\y. connected_component (- S) x y \ norm y > C" proof (cases "x = 0") case True with B Bno show ?thesis by force next case False have "closed_segment x (((B + C) / norm x) *\<^sub>R x) \ - ball 0 B" proof fix w assume "w \ closed_segment x (((B + C) / norm x) *\<^sub>R x)" then obtain u where w: "w = (1 - u + u * (B + C) / norm x) *\<^sub>R x" "0 \ u" "u \ 1" by (auto simp add: closed_segment_def real_vector_class.scaleR_add_left [symmetric]) with False B C have "B \ (1 - u) * norm x + u * (B + C)" using segment_bound_lemma [of B "norm x" "B + C" u] Bno by simp with False B C show "w \ - ball 0 B" using distrib_right [of _ _ "norm x"] by (simp add: ball_def w not_less) qed also have "... \ -S" by (simp add: B) finally have "\T. connected T \ T \ - S \ x \ T \ ((B + C) / norm x) *\<^sub>R x \ T" by (rule_tac x="closed_segment x (((B+C)/norm x) *\<^sub>R x)" in exI) simp with False B show ?thesis by (rule_tac x="((B+C)/norm x) *\<^sub>R x" in exI) (simp add: connected_component_def) qed } then show ?thesis apply (simp add: outside_def assms) apply (rule bounded_subset [OF bounded_ball [of 0 B]]) apply (force simp: dist_norm not_less bounded_pos) done qed lemma unbounded_outside: fixes S :: "'a::{real_normed_vector, perfect_space} set" shows "bounded S \ \ bounded(outside S)" using cobounded_imp_unbounded cobounded_outside by blast lemma bounded_inside: fixes S :: "'a::{real_normed_vector, perfect_space} set" shows "bounded S \ bounded(inside S)" by (simp add: bounded_Int cobounded_outside inside_outside) lemma connected_outside: fixes S :: "'a::euclidean_space set" assumes "bounded S" "2 \ DIM('a)" shows "connected(outside S)" apply (clarsimp simp add: connected_iff_connected_component outside) apply (rule_tac S="connected_component_set (- S) x" in connected_component_of_subset) apply (metis (no_types) assms cobounded_unbounded_component cobounded_unique_unbounded_component connected_component_eq_eq connected_component_idemp double_complement mem_Collect_eq) by (simp add: Collect_mono connected_component_eq) lemma outside_connected_component_lt: "outside S = {x. \B. \y. B < norm(y) \ connected_component (- S) x y}" apply (auto simp: outside bounded_def dist_norm) apply (metis diff_0 norm_minus_cancel not_less) by (metis less_diff_eq norm_minus_commute norm_triangle_ineq2 order.trans pinf(6)) lemma outside_connected_component_le: "outside S = {x. \B. \y. B \ norm(y) \ connected_component (- S) x y}" apply (simp add: outside_connected_component_lt Set.set_eq_iff) by (meson gt_ex leD le_less_linear less_imp_le order.trans) lemma not_outside_connected_component_lt: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" and "2 \ DIM('a)" shows "- (outside S) = {x. \B. \y. B < norm(y) \ \ connected_component (- S) x y}" proof - obtain B::real where B: "0 < B" and Bno: "\x. x \ S \ norm x \ B" using S [simplified bounded_pos] by auto { fix y::'a and z::'a assume yz: "B < norm z" "B < norm y" have "connected_component (- cball 0 B) y z" using assms yz by (force simp: dist_norm intro: connected_componentI [OF _ subset_refl] connected_complement_bounded_convex) then have "connected_component (- S) y z" by (metis connected_component_of_subset Bno Compl_anti_mono mem_cball_0 subset_iff) } note cyz = this show ?thesis apply (auto simp: outside bounded_pos) apply (metis Compl_iff bounded_iff cobounded_imp_unbounded mem_Collect_eq not_le) by (metis B connected_component_trans cyz not_le) qed lemma not_outside_connected_component_le: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" "2 \ DIM('a)" shows "- (outside S) = {x. \B. \y. B \ norm(y) \ \ connected_component (- S) x y}" apply (auto intro: less_imp_le simp: not_outside_connected_component_lt [OF assms]) by (meson gt_ex less_le_trans) lemma inside_connected_component_lt: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" "2 \ DIM('a)" shows "inside S = {x. (x \ S) \ (\B. \y. B < norm(y) \ \ connected_component (- S) x y)}" by (auto simp: inside_outside not_outside_connected_component_lt [OF assms]) lemma inside_connected_component_le: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" "2 \ DIM('a)" shows "inside S = {x. (x \ S) \ (\B. \y. B \ norm(y) \ \ connected_component (- S) x y)}" by (auto simp: inside_outside not_outside_connected_component_le [OF assms]) lemma inside_subset: assumes "connected U" and "\ bounded U" and "T \ U = - S" shows "inside S \ T" apply (auto simp: inside_def) by (metis bounded_subset [of "connected_component_set (- S) _"] connected_component_maximal Compl_iff Un_iff assms subsetI) lemma frontier_not_empty: fixes S :: "'a :: real_normed_vector set" shows "\S \ {}; S \ UNIV\ \ frontier S \ {}" using connected_Int_frontier [of UNIV S] by auto lemma frontier_eq_empty: fixes S :: "'a :: real_normed_vector set" shows "frontier S = {} \ S = {} \ S = UNIV" using frontier_UNIV frontier_empty frontier_not_empty by blast lemma frontier_of_connected_component_subset: fixes S :: "'a::real_normed_vector set" shows "frontier(connected_component_set S x) \ frontier S" proof - { fix y assume y1: "y \ closure (connected_component_set S x)" and y2: "y \ interior (connected_component_set S x)" have "y \ closure S" using y1 closure_mono connected_component_subset by blast moreover have "z \ interior (connected_component_set S x)" if "0 < e" "ball y e \ interior S" "dist y z < e" for e z proof - have "ball y e \ connected_component_set S y" using connected_component_maximal that interior_subset by (metis centre_in_ball connected_ball subset_trans) then show ?thesis using y1 apply (simp add: closure_approachable open_contains_ball_eq [OF open_interior]) by (metis connected_component_eq dist_commute mem_Collect_eq mem_ball mem_interior subsetD \0 < e\ y2) qed then have "y \ interior S" using y2 by (force simp: open_contains_ball_eq [OF open_interior]) ultimately have "y \ frontier S" by (auto simp: frontier_def) } then show ?thesis by (auto simp: frontier_def) qed lemma frontier_Union_subset_closure: fixes F :: "'a::real_normed_vector set set" shows "frontier(\F) \ closure(\t \ F. frontier t)" proof - have "\y\F. \y\frontier y. dist y x < e" if "T \ F" "y \ T" "dist y x < e" "x \ interior (\F)" "0 < e" for x y e T proof (cases "x \ T") case True with that show ?thesis by (metis Diff_iff Sup_upper closure_subset contra_subsetD dist_self frontier_def interior_mono) next case False have 1: "closed_segment x y \ T \ {}" using \y \ T\ by blast have 2: "closed_segment x y - T \ {}" using False by blast obtain c where "c \ closed_segment x y" "c \ frontier T" using False connected_Int_frontier [OF connected_segment 1 2] by auto then show ?thesis proof - have "norm (y - x) < e" by (metis dist_norm \dist y x < e\) moreover have "norm (c - x) \ norm (y - x)" by (simp add: \c \ closed_segment x y\ segment_bound(1)) ultimately have "norm (c - x) < e" by linarith then show ?thesis by (metis (no_types) \c \ frontier T\ dist_norm that(1)) qed qed then show ?thesis by (fastforce simp add: frontier_def closure_approachable) qed lemma frontier_Union_subset: fixes F :: "'a::real_normed_vector set set" shows "finite F \ frontier(\F) \ (\t \ F. frontier t)" by (rule order_trans [OF frontier_Union_subset_closure]) (auto simp: closure_subset_eq) lemma frontier_of_components_subset: fixes S :: "'a::real_normed_vector set" shows "C \ components S \ frontier C \ frontier S" by (metis Path_Connected.frontier_of_connected_component_subset components_iff) lemma frontier_of_components_closed_complement: fixes S :: "'a::real_normed_vector set" shows "\closed S; C \ components (- S)\ \ frontier C \ S" using frontier_complement frontier_of_components_subset frontier_subset_eq by blast lemma frontier_minimal_separating_closed: fixes S :: "'a::real_normed_vector set" assumes "closed S" and nconn: "\ connected(- S)" and C: "C \ components (- S)" and conn: "\T. \closed T; T \ S\ \ connected(- T)" shows "frontier C = S" proof (rule ccontr) assume "frontier C \ S" then have "frontier C \ S" using frontier_of_components_closed_complement [OF \closed S\ C] by blast then have "connected(- (frontier C))" by (simp add: conn) have "\ connected(- (frontier C))" unfolding connected_def not_not proof (intro exI conjI) show "open C" using C \closed S\ open_components by blast show "open (- closure C)" by blast show "C \ - closure C \ - frontier C = {}" using closure_subset by blast show "C \ - frontier C \ {}" using C \open C\ components_eq frontier_disjoint_eq by fastforce show "- frontier C \ C \ - closure C" by (simp add: \open C\ closed_Compl frontier_closures) then show "- closure C \ - frontier C \ {}" by (metis (no_types, lifting) C Compl_subset_Compl_iff \frontier C \ S\ compl_sup frontier_closures in_components_subset psubsetE sup.absorb_iff2 sup.boundedE sup_bot.right_neutral sup_inf_absorb) qed then show False using \connected (- frontier C)\ by blast qed lemma connected_component_UNIV [simp]: fixes x :: "'a::real_normed_vector" shows "connected_component_set UNIV x = UNIV" using connected_iff_eq_connected_component_set [of "UNIV::'a set"] connected_UNIV by auto lemma connected_component_eq_UNIV: fixes x :: "'a::real_normed_vector" shows "connected_component_set s x = UNIV \ s = UNIV" using connected_component_in connected_component_UNIV by blast lemma components_UNIV [simp]: "components UNIV = {UNIV :: 'a::real_normed_vector set}" by (auto simp: components_eq_sing_iff) lemma interior_inside_frontier: fixes S :: "'a::real_normed_vector set" assumes "bounded S" shows "interior S \ inside (frontier S)" proof - { fix x y assume x: "x \ interior S" and y: "y \ S" and cc: "connected_component (- frontier S) x y" have "connected_component_set (- frontier S) x \ frontier S \ {}" proof (rule connected_Int_frontier; simp add: set_eq_iff) show "\u. connected_component (- frontier S) x u \ u \ S" by (meson cc connected_component_in connected_component_refl_eq interior_subset subsetD x) show "\u. connected_component (- frontier S) x u \ u \ S" using y cc by blast qed then have "bounded (connected_component_set (- frontier S) x)" using connected_component_in by auto } then show ?thesis apply (auto simp: inside_def frontier_def) apply (rule classical) apply (rule bounded_subset [OF assms], blast) done qed lemma inside_empty [simp]: "inside {} = ({} :: 'a :: {real_normed_vector, perfect_space} set)" by (simp add: inside_def) lemma outside_empty [simp]: "outside {} = (UNIV :: 'a :: {real_normed_vector, perfect_space} set)" using inside_empty inside_Un_outside by blast lemma inside_same_component: "\connected_component (- S) x y; x \ inside S\ \ y \ inside S" using connected_component_eq connected_component_in by (fastforce simp add: inside_def) lemma outside_same_component: "\connected_component (- S) x y; x \ outside S\ \ y \ outside S" using connected_component_eq connected_component_in by (fastforce simp add: outside_def) lemma convex_in_outside: fixes S :: "'a :: {real_normed_vector, perfect_space} set" assumes S: "convex S" and z: "z \ S" shows "z \ outside S" proof (cases "S={}") case True then show ?thesis by simp next case False then obtain a where "a \ S" by blast with z have zna: "z \ a" by auto { assume "bounded (connected_component_set (- S) z)" with bounded_pos_less obtain B where "B>0" and B: "\x. connected_component (- S) z x \ norm x < B" by (metis mem_Collect_eq) define C where "C = (B + 1 + norm z) / norm (z-a)" have "C > 0" using \0 < B\ zna by (simp add: C_def field_split_simps add_strict_increasing) have "\norm (z + C *\<^sub>R (z-a)) - norm (C *\<^sub>R (z-a))\ \ norm z" by (metis add_diff_cancel norm_triangle_ineq3) moreover have "norm (C *\<^sub>R (z-a)) > norm z + B" using zna \B>0\ by (simp add: C_def le_max_iff_disj) ultimately have C: "norm (z + C *\<^sub>R (z-a)) > B" by linarith { fix u::real assume u: "0\u" "u\1" and ins: "(1 - u) *\<^sub>R z + u *\<^sub>R (z + C *\<^sub>R (z - a)) \ S" then have Cpos: "1 + u * C > 0" by (meson \0 < C\ add_pos_nonneg less_eq_real_def zero_le_mult_iff zero_less_one) then have *: "(1 / (1 + u * C)) *\<^sub>R z + (u * C / (1 + u * C)) *\<^sub>R z = z" by (simp add: scaleR_add_left [symmetric] field_split_simps) then have False using convexD_alt [OF S \a \ S\ ins, of "1/(u*C + 1)"] \C>0\ \z \ S\ Cpos u by (simp add: * field_split_simps) } note contra = this have "connected_component (- S) z (z + C *\<^sub>R (z-a))" proof (rule connected_componentI [OF connected_segment]) show "closed_segment z (z + C *\<^sub>R (z - a)) \ - S" using contra by (force simp add: closed_segment_def) qed auto then have False using zna B [of "z + C *\<^sub>R (z-a)"] C by (auto simp: field_split_simps max_mult_distrib_right) } then show ?thesis by (auto simp: outside_def z) qed lemma outside_convex: fixes S :: "'a :: {real_normed_vector, perfect_space} set" assumes "convex S" shows "outside S = - S" by (metis ComplD assms convex_in_outside equalityI inside_Un_outside subsetI sup.cobounded2) lemma outside_singleton [simp]: fixes x :: "'a :: {real_normed_vector, perfect_space}" shows "outside {x} = -{x}" by (auto simp: outside_convex) lemma inside_convex: fixes S :: "'a :: {real_normed_vector, perfect_space} set" shows "convex S \ inside S = {}" by (simp add: inside_outside outside_convex) lemma inside_singleton [simp]: fixes x :: "'a :: {real_normed_vector, perfect_space}" shows "inside {x} = {}" by (auto simp: inside_convex) lemma outside_subset_convex: fixes S :: "'a :: {real_normed_vector, perfect_space} set" shows "\convex T; S \ T\ \ - T \ outside S" using outside_convex outside_mono by blast lemma outside_Un_outside_Un: fixes S :: "'a::real_normed_vector set" assumes "S \ outside(T \ U) = {}" shows "outside(T \ U) \ outside(T \ S)" proof fix x assume x: "x \ outside (T \ U)" have "Y \ - S" if "connected Y" "Y \ - T" "Y \ - U" "x \ Y" "u \ Y" for u Y proof - have "Y \ connected_component_set (- (T \ U)) x" by (simp add: connected_component_maximal that) also have "\ \ outside(T \ U)" by (metis (mono_tags, lifting) Collect_mono mem_Collect_eq outside outside_same_component x) finally have "Y \ outside(T \ U)" . with assms show ?thesis by auto qed with x show "x \ outside (T \ S)" by (simp add: outside_connected_component_lt connected_component_def) meson qed lemma outside_frontier_misses_closure: fixes S :: "'a::real_normed_vector set" assumes "bounded S" shows "outside(frontier S) \ - closure S" unfolding outside_inside Lattices.boolean_algebra_class.compl_le_compl_iff proof - { assume "interior S \ inside (frontier S)" hence "interior S \ inside (frontier S) = inside (frontier S)" by (simp add: subset_Un_eq) then have "closure S \ frontier S \ inside (frontier S)" using frontier_def by auto } then show "closure S \ frontier S \ inside (frontier S)" using interior_inside_frontier [OF assms] by blast qed lemma outside_frontier_eq_complement_closure: fixes S :: "'a :: {real_normed_vector, perfect_space} set" assumes "bounded S" "convex S" shows "outside(frontier S) = - closure S" by (metis Diff_subset assms convex_closure frontier_def outside_frontier_misses_closure outside_subset_convex subset_antisym) lemma inside_frontier_eq_interior: fixes S :: "'a :: {real_normed_vector, perfect_space} set" shows "\bounded S; convex S\ \ inside(frontier S) = interior S" apply (simp add: inside_outside outside_frontier_eq_complement_closure) using closure_subset interior_subset apply (auto simp: frontier_def) done lemma open_inside: fixes S :: "'a::real_normed_vector set" assumes "closed S" shows "open (inside S)" proof - { fix x assume x: "x \ inside S" have "open (connected_component_set (- S) x)" using assms open_connected_component by blast then obtain e where e: "e>0" and e: "\y. dist y x < e \ connected_component (- S) x y" using dist_not_less_zero apply (simp add: open_dist) by (metis (no_types, lifting) Compl_iff connected_component_refl_eq inside_def mem_Collect_eq x) then have "\e>0. ball x e \ inside S" by (metis e dist_commute inside_same_component mem_ball subsetI x) } then show ?thesis by (simp add: open_contains_ball) qed lemma open_outside: fixes S :: "'a::real_normed_vector set" assumes "closed S" shows "open (outside S)" proof - { fix x assume x: "x \ outside S" have "open (connected_component_set (- S) x)" using assms open_connected_component by blast then obtain e where e: "e>0" and e: "\y. dist y x < e \ connected_component (- S) x y" using dist_not_less_zero x by (auto simp add: open_dist outside_def intro: connected_component_refl) then have "\e>0. ball x e \ outside S" by (metis e dist_commute outside_same_component mem_ball subsetI x) } then show ?thesis by (simp add: open_contains_ball) qed lemma closure_inside_subset: fixes S :: "'a::real_normed_vector set" assumes "closed S" shows "closure(inside S) \ S \ inside S" by (metis assms closure_minimal open_closed open_outside sup.cobounded2 union_with_inside) lemma frontier_inside_subset: fixes S :: "'a::real_normed_vector set" assumes "closed S" shows "frontier(inside S) \ S" proof - have "closure (inside S) \ - inside S = closure (inside S) - interior (inside S)" by (metis (no_types) Diff_Compl assms closure_closed interior_closure open_closed open_inside) moreover have "- inside S \ - outside S = S" by (metis (no_types) compl_sup double_compl inside_Un_outside) moreover have "closure (inside S) \ - outside S" by (metis (no_types) assms closure_inside_subset union_with_inside) ultimately have "closure (inside S) - interior (inside S) \ S" by blast then show ?thesis by (simp add: frontier_def open_inside interior_open) qed lemma closure_outside_subset: fixes S :: "'a::real_normed_vector set" assumes "closed S" shows "closure(outside S) \ S \ outside S" by (metis assms closed_open closure_minimal inside_outside open_inside sup_ge2) lemma frontier_outside_subset: fixes S :: "'a::real_normed_vector set" assumes "closed S" shows "frontier(outside S) \ S" unfolding frontier_def by (metis Diff_subset_conv assms closure_outside_subset interior_eq open_outside sup_aci(1)) lemma inside_complement_unbounded_connected_empty: "\connected (- S); \ bounded (- S)\ \ inside S = {}" using inside_subset by blast lemma inside_bounded_complement_connected_empty: fixes S :: "'a::{real_normed_vector, perfect_space} set" shows "\connected (- S); bounded S\ \ inside S = {}" by (metis inside_complement_unbounded_connected_empty cobounded_imp_unbounded) lemma inside_inside: assumes "S \ inside T" shows "inside S - T \ inside T" unfolding inside_def proof clarify fix x assume x: "x \ T" "x \ S" and bo: "bounded (connected_component_set (- S) x)" show "bounded (connected_component_set (- T) x)" proof (cases "S \ connected_component_set (- T) x = {}") case True then show ?thesis by (metis bounded_subset [OF bo] compl_le_compl_iff connected_component_idemp connected_component_mono disjoint_eq_subset_Compl double_compl) next case False then obtain y where y: "y \ S" "y \ connected_component_set (- T) x" by (meson disjoint_iff) then have "bounded (connected_component_set (- T) y)" using assms [unfolded inside_def] by blast with y show ?thesis by (metis connected_component_eq) qed qed lemma inside_inside_subset: "inside(inside S) \ S" using inside_inside union_with_outside by fastforce lemma inside_outside_intersect_connected: "\connected T; inside S \ T \ {}; outside S \ T \ {}\ \ S \ T \ {}" apply (simp add: inside_def outside_def ex_in_conv [symmetric] disjoint_eq_subset_Compl, clarify) by (metis (no_types, opaque_lifting) Compl_anti_mono connected_component_eq connected_component_maximal contra_subsetD double_compl) lemma outside_bounded_nonempty: fixes S :: "'a :: {real_normed_vector, perfect_space} set" assumes "bounded S" shows "outside S \ {}" by (metis (no_types, lifting) Collect_empty_eq Collect_mem_eq Compl_eq_Diff_UNIV Diff_cancel Diff_disjoint UNIV_I assms ball_eq_empty bounded_diff cobounded_outside convex_ball double_complement order_refl outside_convex outside_def) lemma outside_compact_in_open: fixes S :: "'a :: {real_normed_vector,perfect_space} set" assumes S: "compact S" and T: "open T" and "S \ T" "T \ {}" shows "outside S \ T \ {}" proof - have "outside S \ {}" by (simp add: compact_imp_bounded outside_bounded_nonempty S) with assms obtain a b where a: "a \ outside S" and b: "b \ T" by auto show ?thesis proof (cases "a \ T") case True with a show ?thesis by blast next case False have front: "frontier T \ - S" using \S \ T\ frontier_disjoint_eq T by auto { fix \ assume "path \" and pimg_sbs: "path_image \ - {pathfinish \} \ interior (- T)" and pf: "pathfinish \ \ frontier T" and ps: "pathstart \ = a" define c where "c = pathfinish \" have "c \ -S" unfolding c_def using front pf by blast moreover have "open (-S)" using S compact_imp_closed by blast ultimately obtain \::real where "\ > 0" and \: "cball c \ \ -S" using open_contains_cball[of "-S"] S by blast then obtain d where "d \ T" and d: "dist d c < \" using closure_approachable [of c T] pf unfolding c_def by (metis Diff_iff frontier_def) then have "d \ -S" using \ using dist_commute by (metis contra_subsetD mem_cball not_le not_less_iff_gr_or_eq) have pimg_sbs_cos: "path_image \ \ -S" using \c \ - S\ \S \ T\ c_def interior_subset pimg_sbs by fastforce have "closed_segment c d \ cball c \" by (metis \0 < \\ centre_in_cball closed_segment_subset convex_cball d dist_commute less_eq_real_def mem_cball) with \ have "closed_segment c d \ -S" by blast moreover have con_gcd: "connected (path_image \ \ closed_segment c d)" by (rule connected_Un) (auto simp: c_def \path \\ connected_path_image) ultimately have "connected_component (- S) a d" unfolding connected_component_def using pimg_sbs_cos ps by blast then have "outside S \ T \ {}" using outside_same_component [OF _ a] by (metis IntI \d \ T\ empty_iff) } note * = this have pal: "pathstart (linepath a b) \ closure (- T)" by (auto simp: False closure_def) show ?thesis by (rule exists_path_subpath_to_frontier [OF path_linepath pal _ *]) (auto simp: b) qed qed lemma inside_inside_compact_connected: fixes S :: "'a :: euclidean_space set" assumes S: "closed S" and T: "compact T" and "connected T" "S \ inside T" shows "inside S \ inside T" proof (cases "inside T = {}") case True with assms show ?thesis by auto next case False consider "DIM('a) = 1" | "DIM('a) \ 2" using antisym not_less_eq_eq by fastforce then show ?thesis proof cases case 1 then show ?thesis using connected_convex_1_gen assms False inside_convex by blast next case 2 have "bounded S" using assms by (meson bounded_inside bounded_subset compact_imp_bounded) then have coms: "compact S" by (simp add: S compact_eq_bounded_closed) then have bst: "bounded (S \ T)" by (simp add: compact_imp_bounded T) then obtain r where "0 < r" and r: "S \ T \ ball 0 r" using bounded_subset_ballD by blast have outst: "outside S \ outside T \ {}" proof - have "- ball 0 r \ outside S" by (meson convex_ball le_supE outside_subset_convex r) moreover have "- ball 0 r \ outside T" by (meson convex_ball le_supE outside_subset_convex r) ultimately show ?thesis by (metis Compl_subset_Compl_iff Int_subset_iff bounded_ball inf.orderE outside_bounded_nonempty outside_no_overlap) qed have "S \ T = {}" using assms by (metis disjoint_iff_not_equal inside_no_overlap subsetCE) moreover have "outside S \ inside T \ {}" by (meson False assms(4) compact_eq_bounded_closed coms open_inside outside_compact_in_open T) ultimately have "inside S \ T = {}" using inside_outside_intersect_connected [OF \connected T\, of S] by (metis "2" compact_eq_bounded_closed coms connected_outside inf.commute inside_outside_intersect_connected outst) then show ?thesis using inside_inside [OF \S \ inside T\] by blast qed qed lemma connected_with_inside: fixes S :: "'a :: real_normed_vector set" assumes S: "closed S" and cons: "connected S" shows "connected(S \ inside S)" proof (cases "S \ inside S = UNIV") case True with assms show ?thesis by auto next case False then obtain b where b: "b \ S" "b \ inside S" by blast have *: "\y T. y \ S \ connected T \ a \ T \ y \ T \ T \ (S \ inside S)" if "a \ S \ inside S" for a using that proof assume "a \ S" then show ?thesis by (rule_tac x=a in exI, rule_tac x="{a}" in exI, simp) next assume a: "a \ inside S" then have ain: "a \ closure (inside S)" by (simp add: closure_def) show ?thesis apply (rule exists_path_subpath_to_frontier [OF path_linepath [of a b], of "inside S"]) apply (simp_all add: ain b) subgoal for h apply (rule_tac x="pathfinish h" in exI) apply (simp add: subsetD [OF frontier_inside_subset[OF S]]) apply (rule_tac x="path_image h" in exI) apply (simp add: pathfinish_in_path_image connected_path_image, auto) by (metis Diff_single_insert S frontier_inside_subset insert_iff interior_subset subsetD) done qed show ?thesis apply (simp add: connected_iff_connected_component) apply (clarsimp simp add: connected_component_def dest!: *) subgoal for x y u u' T t' by (rule_tac x="(S \ T \ t')" in exI) (auto intro!: connected_Un cons) done qed text\The proof is virtually the same as that above.\ lemma connected_with_outside: fixes S :: "'a :: real_normed_vector set" assumes S: "closed S" and cons: "connected S" shows "connected(S \ outside S)" proof (cases "S \ outside S = UNIV") case True with assms show ?thesis by auto next case False then obtain b where b: "b \ S" "b \ outside S" by blast have *: "\y T. y \ S \ connected T \ a \ T \ y \ T \ T \ (S \ outside S)" if "a \ (S \ outside S)" for a using that proof assume "a \ S" then show ?thesis by (rule_tac x=a in exI, rule_tac x="{a}" in exI, simp) next assume a: "a \ outside S" then have ain: "a \ closure (outside S)" by (simp add: closure_def) show ?thesis apply (rule exists_path_subpath_to_frontier [OF path_linepath [of a b], of "outside S"]) apply (simp_all add: ain b) subgoal for h apply (rule_tac x="pathfinish h" in exI) apply (simp add: subsetD [OF frontier_outside_subset[OF S]]) apply (rule_tac x="path_image h" in exI) apply (simp add: pathfinish_in_path_image connected_path_image, auto) by (metis (no_types, lifting) frontier_outside_subset insertE insert_Diff interior_eq open_outside pathfinish_in_path_image S subsetCE) done qed show ?thesis apply (simp add: connected_iff_connected_component) apply (clarsimp simp add: connected_component_def dest!: *) subgoal for x y u u' T t' by (rule_tac x="(S \ T \ t')" in exI) (auto intro!: connected_Un cons) done qed lemma inside_inside_eq_empty [simp]: fixes S :: "'a :: {real_normed_vector, perfect_space} set" assumes S: "closed S" and cons: "connected S" shows "inside (inside S) = {}" by (metis (no_types) unbounded_outside connected_with_outside [OF assms] bounded_Un inside_complement_unbounded_connected_empty unbounded_outside union_with_outside) lemma inside_in_components: "inside S \ components (- S) \ connected(inside S) \ inside S \ {}" (is "?lhs = ?rhs") proof assume R: ?rhs then have "\x. \x \ S; x \ inside S\ \ \ connected (inside S)" by (simp add: inside_outside) with R show ?lhs unfolding in_components_maximal by (auto intro: inside_same_component connected_componentI) qed (simp add: in_components_maximal) text\The proof is like that above.\ lemma outside_in_components: "outside S \ components (- S) \ connected(outside S) \ outside S \ {}" (is "?lhs = ?rhs") proof assume R: ?rhs then have "\x. \x \ S; x \ outside S\ \ \ connected (outside S)" by (meson disjoint_iff outside_no_overlap) with R show ?lhs unfolding in_components_maximal by (auto intro: outside_same_component connected_componentI) qed (simp add: in_components_maximal) lemma bounded_unique_outside: fixes S :: "'a :: euclidean_space set" assumes "bounded S" "DIM('a) \ 2" shows "(c \ components (- S) \ \ bounded c \ c = outside S)" using assms by (metis cobounded_unique_unbounded_components connected_outside double_compl outside_bounded_nonempty outside_in_components unbounded_outside) subsection\Condition for an open map's image to contain a ball\ proposition ball_subset_open_map_image: fixes f :: "'a::heine_borel \ 'b :: {real_normed_vector,heine_borel}" assumes contf: "continuous_on (closure S) f" and oint: "open (f ` interior S)" and le_no: "\z. z \ frontier S \ r \ norm(f z - f a)" and "bounded S" "a \ S" "0 < r" shows "ball (f a) r \ f ` S" proof (cases "f ` S = UNIV") case True then show ?thesis by simp next case False then have "closed (frontier (f ` S))" "frontier (f ` S) \ {}" using \a \ S\ by (auto simp: frontier_eq_empty) then obtain w where w: "w \ frontier (f ` S)" and dw_le: "\y. y \ frontier (f ` S) \ norm (f a - w) \ norm (f a - y)" by (auto simp add: dist_norm intro: distance_attains_inf [of "frontier(f ` S)" "f a"]) then obtain \ where \: "\n. \ n \ f ` S" and tendsw: "\ \ w" by (metis Diff_iff frontier_def closure_sequential) then have "\n. \x \ S. \ n = f x" by force then obtain z where zs: "\n. z n \ S" and fz: "\n. \ n = f (z n)" by metis then obtain y K where y: "y \ closure S" and "strict_mono (K :: nat \ nat)" and Klim: "(z \ K) \ y" using \bounded S\ unfolding compact_closure [symmetric] compact_def by (meson closure_subset subset_iff) then have ftendsw: "((\n. f (z n)) \ K) \ w" by (metis LIMSEQ_subseq_LIMSEQ fun.map_cong0 fz tendsw) have zKs: "\n. (z \ K) n \ S" by (simp add: zs) have fz: "f \ z = \" "(\n. f (z n)) = \" using fz by auto then have "(\ \ K) \ f y" by (metis (no_types) Klim zKs y contf comp_assoc continuous_on_closure_sequentially) with fz have wy: "w = f y" using fz LIMSEQ_unique ftendsw by auto have rle: "r \ norm (f y - f a)" proof (rule le_no) show "y \ frontier S" using w wy oint by (force simp: imageI image_mono interiorI interior_subset frontier_def y) qed have **: "(b \ (- S) \ {} \ b - (- S) \ {} \ b \ f \ {}) \ (b \ S \ {}) \ b \ f = {} \ b \ S" for b f and S :: "'b set" by blast have \
: "\y. \norm (f a - y) < r; y \ frontier (f ` S)\ \ False" by (metis dw_le norm_minus_commute not_less order_trans rle wy) show ?thesis apply (rule ** [OF connected_Int_frontier [where t = "f`S", OF connected_ball]]) (*such a horrible mess*) using \a \ S\ \0 < r\ by (auto simp: disjoint_iff_not_equal dist_norm dest: \
) qed subsubsection\Special characterizations of classes of functions into and out of R.\ lemma Hausdorff_space_euclidean [simp]: "Hausdorff_space (euclidean :: 'a::metric_space topology)" proof - have "\U V. open U \ open V \ x \ U \ y \ V \ disjnt U V" if "x \ y" for x y :: 'a proof (intro exI conjI) let ?r = "dist x y / 2" have [simp]: "?r > 0" by (simp add: that) show "open (ball x ?r)" "open (ball y ?r)" "x \ (ball x ?r)" "y \ (ball y ?r)" by (auto simp add: that) show "disjnt (ball x ?r) (ball y ?r)" unfolding disjnt_def by (simp add: disjoint_ballI) qed then show ?thesis by (simp add: Hausdorff_space_def) qed proposition embedding_map_into_euclideanreal: assumes "path_connected_space X" shows "embedding_map X euclideanreal f \ continuous_map X euclideanreal f \ inj_on f (topspace X)" proof safe show "continuous_map X euclideanreal f" if "embedding_map X euclideanreal f" using continuous_map_in_subtopology homeomorphic_imp_continuous_map that unfolding embedding_map_def by blast show "inj_on f (topspace X)" if "embedding_map X euclideanreal f" using that homeomorphic_imp_injective_map unfolding embedding_map_def by blast show "embedding_map X euclideanreal f" if cont: "continuous_map X euclideanreal f" and inj: "inj_on f (topspace X)" proof - obtain g where gf: "\x. x \ topspace X \ g (f x) = x" using inv_into_f_f [OF inj] by auto show ?thesis unfolding embedding_map_def homeomorphic_map_maps homeomorphic_maps_def proof (intro exI conjI) show "continuous_map X (top_of_set (f ` topspace X)) f" by (simp add: cont continuous_map_in_subtopology) let ?S = "f ` topspace X" have eq: "{x \ ?S. g x \ U} = f ` U" if "openin X U" for U using openin_subset [OF that] by (auto simp: gf) have 1: "g ` ?S \ topspace X" using eq by blast have "openin (top_of_set ?S) {x \ ?S. g x \ T}" if "openin X T" for T proof - have "T \ topspace X" by (simp add: openin_subset that) have RR: "\x \ ?S \ g -` T. \d>0. \x' \ ?S \ ball x d. g x' \ T" proof (clarsimp simp add: gf) have pcS: "path_connectedin euclidean ?S" using assms cont path_connectedin_continuous_map_image path_connectedin_topspace by blast show "\d>0. \x'\f ` topspace X \ ball (f x) d. g x' \ T" if "x \ T" for x proof - have x: "x \ topspace X" using \T \ topspace X\ \x \ T\ by blast obtain u v d where "0 < d" "u \ topspace X" "v \ topspace X" and sub_fuv: "?S \ {f x - d .. f x + d} \ {f u..f v}" proof (cases "\u \ topspace X. f u < f x") case True then obtain u where u: "u \ topspace X" "f u < f x" .. show ?thesis proof (cases "\v \ topspace X. f x < f v") case True then obtain v where v: "v \ topspace X" "f x < f v" .. show ?thesis proof let ?d = "min (f x - f u) (f v - f x)" show "0 < ?d" by (simp add: \f u < f x\ \f x < f v\) show "f ` topspace X \ {f x - ?d..f x + ?d} \ {f u..f v}" by fastforce qed (auto simp: u v) next case False show ?thesis proof let ?d = "f x - f u" show "0 < ?d" by (simp add: u) show "f ` topspace X \ {f x - ?d..f x + ?d} \ {f u..f x}" using x u False by auto qed (auto simp: x u) qed next case False note no_u = False show ?thesis proof (cases "\v \ topspace X. f x < f v") case True then obtain v where v: "v \ topspace X" "f x < f v" .. show ?thesis proof let ?d = "f v - f x" show "0 < ?d" by (simp add: v) show "f ` topspace X \ {f x - ?d..f x + ?d} \ {f x..f v}" using False by auto qed (auto simp: x v) next case False show ?thesis proof show "f ` topspace X \ {f x - 1..f x + 1} \ {f x..f x}" using False no_u by fastforce qed (auto simp: x) qed qed then obtain h where "pathin X h" "h 0 = u" "h 1 = v" using assms unfolding path_connected_space_def by blast obtain C where "compactin X C" "connectedin X C" "u \ C" "v \ C" proof show "compactin X (h ` {0..1})" using that by (simp add: \pathin X h\ compactin_path_image) show "connectedin X (h ` {0..1})" using \pathin X h\ connectedin_path_image by blast qed (use \h 0 = u\ \h 1 = v\ in auto) have "continuous_map (subtopology euclideanreal (?S \ {f x - d .. f x + d})) (subtopology X C) g" proof (rule continuous_inverse_map) show "compact_space (subtopology X C)" using \compactin X C\ compactin_subspace by blast show "continuous_map (subtopology X C) euclideanreal f" by (simp add: cont continuous_map_from_subtopology) have "{f u .. f v} \ f ` topspace (subtopology X C)" proof (rule connected_contains_Icc) show "connected (f ` topspace (subtopology X C))" using connectedin_continuous_map_image [OF cont] by (simp add: \compactin X C\ \connectedin X C\ compactin_subset_topspace inf_absorb2) show "f u \ f ` topspace (subtopology X C)" by (simp add: \u \ C\ \u \ topspace X\) show "f v \ f ` topspace (subtopology X C)" by (simp add: \v \ C\ \v \ topspace X\) qed then show "f ` topspace X \ {f x - d..f x + d} \ f ` topspace (subtopology X C)" using sub_fuv by blast qed (auto simp: gf) then have contg: "continuous_map (subtopology euclideanreal (?S \ {f x - d .. f x + d})) X g" using continuous_map_in_subtopology by blast have "\e>0. \x \ ?S \ {f x - d .. f x + d} \ ball (f x) e. g x \ T" using openin_continuous_map_preimage [OF contg \openin X T\] x \x \ T\ \0 < d\ unfolding openin_euclidean_subtopology_iff by (force simp: gf dist_commute) then obtain e where "e > 0 \ (\x\f ` topspace X \ {f x - d..f x + d} \ ball (f x) e. g x \ T)" by metis with \0 < d\ have "min d e > 0" "\u. u \ topspace X \ \f x - f u\ < min d e \ u \ T" using dist_real_def gf by force+ then show ?thesis by (metis (full_types) Int_iff dist_real_def image_iff mem_ball gf) qed qed then obtain d where d: "\r. r \ ?S \ g -` T \ d r > 0 \ (\x \ ?S \ ball r (d r). g x \ T)" by metis show ?thesis unfolding openin_subtopology proof (intro exI conjI) show "{x \ ?S. g x \ T} = (\r \ ?S \ g -` T. ball r (d r)) \ f ` topspace X" using d by (auto simp: gf) qed auto qed then show "continuous_map (top_of_set ?S) X g" by (simp add: continuous_map_def gf) qed (auto simp: gf) qed qed subsubsection \An injective function into R is a homeomorphism and so an open map.\ lemma injective_into_1d_eq_homeomorphism: fixes f :: "'a::topological_space \ real" assumes f: "continuous_on S f" and S: "path_connected S" shows "inj_on f S \ (\g. homeomorphism S (f ` S) f g)" proof show "\g. homeomorphism S (f ` S) f g" if "inj_on f S" proof - have "embedding_map (top_of_set S) euclideanreal f" using that embedding_map_into_euclideanreal [of "top_of_set S" f] assms by auto then show ?thesis by (simp add: embedding_map_def) (metis all_closedin_homeomorphic_image f homeomorphism_injective_closed_map that) qed qed (metis homeomorphism_def inj_onI) lemma injective_into_1d_imp_open_map: fixes f :: "'a::topological_space \ real" assumes "continuous_on S f" "path_connected S" "inj_on f S" "openin (subtopology euclidean S) T" shows "openin (subtopology euclidean (f ` S)) (f ` T)" using assms homeomorphism_imp_open_map injective_into_1d_eq_homeomorphism by blast lemma homeomorphism_into_1d: fixes f :: "'a::topological_space \ real" assumes "path_connected S" "continuous_on S f" "f ` S = T" "inj_on f S" shows "\g. homeomorphism S T f g" using assms injective_into_1d_eq_homeomorphism by blast subsection\<^marker>\tag unimportant\ \Rectangular paths\ definition\<^marker>\tag unimportant\ rectpath where "rectpath a1 a3 = (let a2 = Complex (Re a3) (Im a1); a4 = Complex (Re a1) (Im a3) in linepath a1 a2 +++ linepath a2 a3 +++ linepath a3 a4 +++ linepath a4 a1)" lemma path_rectpath [simp, intro]: "path (rectpath a b)" by (simp add: Let_def rectpath_def) lemma pathstart_rectpath [simp]: "pathstart (rectpath a1 a3) = a1" by (simp add: rectpath_def Let_def) lemma pathfinish_rectpath [simp]: "pathfinish (rectpath a1 a3) = a1" by (simp add: rectpath_def Let_def) lemma simple_path_rectpath [simp, intro]: assumes "Re a1 \ Re a3" "Im a1 \ Im a3" shows "simple_path (rectpath a1 a3)" unfolding rectpath_def Let_def using assms by (intro simple_path_join_loop arc_join arc_linepath) (auto simp: complex_eq_iff path_image_join closed_segment_same_Re closed_segment_same_Im) lemma path_image_rectpath: assumes "Re a1 \ Re a3" "Im a1 \ Im a3" shows "path_image (rectpath a1 a3) = {z. Re z \ {Re a1, Re a3} \ Im z \ {Im a1..Im a3}} \ {z. Im z \ {Im a1, Im a3} \ Re z \ {Re a1..Re a3}}" (is "?lhs = ?rhs") proof - define a2 a4 where "a2 = Complex (Re a3) (Im a1)" and "a4 = Complex (Re a1) (Im a3)" have "?lhs = closed_segment a1 a2 \ closed_segment a2 a3 \ closed_segment a4 a3 \ closed_segment a1 a4" by (simp_all add: rectpath_def Let_def path_image_join closed_segment_commute a2_def a4_def Un_assoc) also have "\ = ?rhs" using assms by (auto simp: rectpath_def Let_def path_image_join a2_def a4_def closed_segment_same_Re closed_segment_same_Im closed_segment_eq_real_ivl) finally show ?thesis . qed lemma path_image_rectpath_subset_cbox: assumes "Re a \ Re b" "Im a \ Im b" shows "path_image (rectpath a b) \ cbox a b" using assms by (auto simp: path_image_rectpath in_cbox_complex_iff) lemma path_image_rectpath_inter_box: assumes "Re a \ Re b" "Im a \ Im b" shows "path_image (rectpath a b) \ box a b = {}" using assms by (auto simp: path_image_rectpath in_box_complex_iff) lemma path_image_rectpath_cbox_minus_box: assumes "Re a \ Re b" "Im a \ Im b" shows "path_image (rectpath a b) = cbox a b - box a b" using assms by (auto simp: path_image_rectpath in_cbox_complex_iff in_box_complex_iff) end diff --git a/src/HOL/Analysis/Starlike.thy b/src/HOL/Analysis/Starlike.thy --- a/src/HOL/Analysis/Starlike.thy +++ b/src/HOL/Analysis/Starlike.thy @@ -1,6404 +1,6345 @@ (* Title: HOL/Analysis/Starlike.thy Author: L C Paulson, University of Cambridge Author: Robert Himmelmann, TU Muenchen Author: Bogdan Grechuk, University of Edinburgh Author: Armin Heller, TU Muenchen Author: Johannes Hoelzl, TU Muenchen *) chapter \Unsorted\ theory Starlike imports Convex_Euclidean_Space Line_Segment begin lemma affine_hull_closed_segment [simp]: "affine hull (closed_segment a b) = affine hull {a,b}" by (simp add: segment_convex_hull) lemma affine_hull_open_segment [simp]: fixes a :: "'a::euclidean_space" shows "affine hull (open_segment a b) = (if a = b then {} else affine hull {a,b})" by (metis affine_hull_convex_hull affine_hull_empty closure_open_segment closure_same_affine_hull segment_convex_hull) lemma rel_interior_closure_convex_segment: fixes S :: "_::euclidean_space set" assumes "convex S" "a \ rel_interior S" "b \ closure S" shows "open_segment a b \ rel_interior S" proof fix x have [simp]: "(1 - u) *\<^sub>R a + u *\<^sub>R b = b - (1 - u) *\<^sub>R (b - a)" for u by (simp add: algebra_simps) assume "x \ open_segment a b" then show "x \ rel_interior S" unfolding closed_segment_def open_segment_def using assms by (auto intro: rel_interior_closure_convex_shrink) qed lemma convex_hull_insert_segments: "convex hull (insert a S) = (if S = {} then {a} else \x \ convex hull S. closed_segment a x)" by (force simp add: convex_hull_insert_alt in_segment) lemma Int_convex_hull_insert_rel_exterior: fixes z :: "'a::euclidean_space" assumes "convex C" "T \ C" and z: "z \ rel_interior C" and dis: "disjnt S (rel_interior C)" shows "S \ (convex hull (insert z T)) = S \ (convex hull T)" (is "?lhs = ?rhs") proof have "T = {} \ z \ S" using dis z by (auto simp add: disjnt_def) then show "?lhs \ ?rhs" proof (clarsimp simp add: convex_hull_insert_segments) fix x y assume "x \ S" and y: "y \ convex hull T" and "x \ closed_segment z y" have "y \ closure C" by (metis y \convex C\ \T \ C\ closure_subset contra_subsetD convex_hull_eq hull_mono) moreover have "x \ rel_interior C" by (meson \x \ S\ dis disjnt_iff) moreover have "x \ open_segment z y \ {z, y}" using \x \ closed_segment z y\ closed_segment_eq_open by blast ultimately show "x \ convex hull T" using rel_interior_closure_convex_segment [OF \convex C\ z] using y z by blast qed show "?rhs \ ?lhs" by (meson hull_mono inf_mono subset_insertI subset_refl) qed subsection\<^marker>\tag unimportant\ \Shrinking towards the interior of a convex set\ lemma mem_interior_convex_shrink: fixes S :: "'a::euclidean_space set" assumes "convex S" and "c \ interior S" and "x \ S" and "0 < e" and "e \ 1" shows "x - e *\<^sub>R (x - c) \ interior S" proof - obtain d where "d > 0" and d: "ball c d \ S" using assms(2) unfolding mem_interior by auto show ?thesis unfolding mem_interior proof (intro exI subsetI conjI) fix y assume "y \ ball (x - e *\<^sub>R (x - c)) (e*d)" then have as: "dist (x - e *\<^sub>R (x - c)) y < e * d" by simp have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" using \e > 0\ by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib) have "c - ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = (1 / e) *\<^sub>R (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)" using \e > 0\ by (auto simp add: euclidean_eq_iff[where 'a='a] field_simps inner_simps) then have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = \1/e\ * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)" by (simp add: dist_norm) also have "\ = \1/e\ * norm (x - e *\<^sub>R (x - c) - y)" by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps) also have "\ < d" using as[unfolded dist_norm] and \e > 0\ by (auto simp add:pos_divide_less_eq[OF \e > 0\] mult.commute) finally have "(1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x \ S" using assms(3-5) d by (intro convexD_alt [OF \convex S\]) (auto intro: convexD_alt [OF \convex S\]) with \e > 0\ show "y \ S" by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib) qed (use \e>0\ \d>0\ in auto) qed lemma mem_interior_closure_convex_shrink: fixes S :: "'a::euclidean_space set" assumes "convex S" and "c \ interior S" and "x \ closure S" and "0 < e" and "e \ 1" shows "x - e *\<^sub>R (x - c) \ interior S" proof - obtain d where "d > 0" and d: "ball c d \ S" using assms(2) unfolding mem_interior by auto have "\y\S. norm (y - x) * (1 - e) < e * d" proof (cases "x \ S") case True then show ?thesis using \e > 0\ \d > 0\ by force next case False then have x: "x islimpt S" using assms(3)[unfolded closure_def] by auto show ?thesis proof (cases "e = 1") case True obtain y where "y \ S" "y \ x" "dist y x < 1" using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto then show ?thesis using True \0 < d\ by auto next case False then have "0 < e * d / (1 - e)" and *: "1 - e > 0" using \e \ 1\ \e > 0\ \d > 0\ by auto then obtain y where "y \ S" "y \ x" "dist y x < e * d / (1 - e)" using islimpt_approachable x by blast then have "norm (y - x) * (1 - e) < e * d" by (metis "*" dist_norm mult_imp_div_pos_le not_less) then show ?thesis using \y \ S\ by blast qed qed then obtain y where "y \ S" and y: "norm (y - x) * (1 - e) < e * d" by auto define z where "z = c + ((1 - e) / e) *\<^sub>R (x - y)" have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)" unfolding z_def using \e > 0\ by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib) have "(1 - e) * norm (x - y) / e < d" using y \0 < e\ by (simp add: field_simps norm_minus_commute) then have "z \ interior (ball c d)" using \0 < e\ \e \ 1\ by (simp add: interior_open[OF open_ball] z_def dist_norm) then have "z \ interior S" using d interiorI interior_ball by blast then show ?thesis unfolding * using mem_interior_convex_shrink \y \ S\ assms by blast qed lemma in_interior_closure_convex_segment: fixes S :: "'a::euclidean_space set" assumes "convex S" and a: "a \ interior S" and b: "b \ closure S" shows "open_segment a b \ interior S" proof (clarsimp simp: in_segment) fix u::real assume u: "0 < u" "u < 1" have "(1 - u) *\<^sub>R a + u *\<^sub>R b = b - (1 - u) *\<^sub>R (b - a)" by (simp add: algebra_simps) also have "... \ interior S" using mem_interior_closure_convex_shrink [OF assms] u by simp finally show "(1 - u) *\<^sub>R a + u *\<^sub>R b \ interior S" . qed lemma convex_closure_interior: fixes S :: "'a::euclidean_space set" assumes "convex S" and int: "interior S \ {}" shows "closure(interior S) = closure S" proof - obtain a where a: "a \ interior S" using int by auto have "closure S \ closure(interior S)" proof fix x assume x: "x \ closure S" show "x \ closure (interior S)" proof (cases "x=a") case True then show ?thesis using \a \ interior S\ closure_subset by blast next case False show ?thesis proof (clarsimp simp add: closure_def islimpt_approachable) fix e::real assume xnotS: "x \ interior S" and "0 < e" show "\x'\interior S. x' \ x \ dist x' x < e" proof (intro bexI conjI) show "x - min (e/2 / norm (x - a)) 1 *\<^sub>R (x - a) \ x" using False \0 < e\ by (auto simp: algebra_simps min_def) show "dist (x - min (e/2 / norm (x - a)) 1 *\<^sub>R (x - a)) x < e" using \0 < e\ by (auto simp: dist_norm min_def) show "x - min (e/2 / norm (x - a)) 1 *\<^sub>R (x - a) \ interior S" using \0 < e\ False by (auto simp add: min_def a intro: mem_interior_closure_convex_shrink [OF \convex S\ a x]) qed qed qed qed then show ?thesis by (simp add: closure_mono interior_subset subset_antisym) qed lemma closure_convex_Int_superset: fixes S :: "'a::euclidean_space set" assumes "convex S" "interior S \ {}" "interior S \ closure T" shows "closure(S \ T) = closure S" proof - have "closure S \ closure(interior S)" by (simp add: convex_closure_interior assms) also have "... \ closure (S \ T)" using interior_subset [of S] assms by (metis (no_types, lifting) Int_assoc Int_lower2 closure_mono closure_open_Int_superset inf.orderE open_interior) finally show ?thesis by (simp add: closure_mono dual_order.antisym) qed subsection\<^marker>\tag unimportant\ \Some obvious but surprisingly hard simplex lemmas\ lemma simplex: assumes "finite S" and "0 \ S" shows "convex hull (insert 0 S) = {y. \u. (\x\S. 0 \ u x) \ sum u S \ 1 \ sum (\x. u x *\<^sub>R x) S = y}" proof (simp add: convex_hull_finite set_eq_iff assms, safe) fix x and u :: "'a \ real" assume "0 \ u 0" "\x\S. 0 \ u x" "u 0 + sum u S = 1" then show "\v. (\x\S. 0 \ v x) \ sum v S \ 1 \ (\x\S. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x)" by force next fix x and u :: "'a \ real" assume "\x\S. 0 \ u x" "sum u S \ 1" then show "\v. 0 \ v 0 \ (\x\S. 0 \ v x) \ v 0 + sum v S = 1 \ (\x\S. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x)" by (rule_tac x="\x. if x = 0 then 1 - sum u S else u x" in exI) (auto simp: sum_delta_notmem assms if_smult) qed lemma substd_simplex: assumes d: "d \ Basis" shows "convex hull (insert 0 d) = {x. (\i\Basis. 0 \ x\i) \ (\i\d. x\i) \ 1 \ (\i\Basis. i \ d \ x\i = 0)}" (is "convex hull (insert 0 ?p) = ?s") proof - let ?D = d have "0 \ ?p" using assms by (auto simp: image_def) from d have "finite d" by (blast intro: finite_subset finite_Basis) show ?thesis unfolding simplex[OF \finite d\ \0 \ ?p\] proof (intro set_eqI; safe) fix u :: "'a \ real" assume as: "\x\?D. 0 \ u x" "sum u ?D \ 1" let ?x = "(\x\?D. u x *\<^sub>R x)" have ind: "\i\Basis. i \ d \ u i = ?x \ i" and notind: "(\i\Basis. i \ d \ ?x \ i = 0)" using substdbasis_expansion_unique[OF assms] by blast+ then have **: "sum u ?D = sum ((\) ?x) ?D" using assms by (auto intro!: sum.cong) show "0 \ ?x \ i" if "i \ Basis" for i using as(1) ind notind that by fastforce show "sum ((\) ?x) ?D \ 1" using "**" as(2) by linarith show "?x \ i = 0" if "i \ Basis" "i \ d" for i using notind that by blast next fix x assume "\i\Basis. 0 \ x \ i" "sum ((\) x) ?D \ 1" "(\i\Basis. i \ d \ x \ i = 0)" with d show "\u. (\x\?D. 0 \ u x) \ sum u ?D \ 1 \ (\x\?D. u x *\<^sub>R x) = x" unfolding substdbasis_expansion_unique[OF assms] by (rule_tac x="inner x" in exI) auto qed qed lemma std_simplex: "convex hull (insert 0 Basis) = {x::'a::euclidean_space. (\i\Basis. 0 \ x\i) \ sum (\i. x\i) Basis \ 1}" using substd_simplex[of Basis] by auto lemma interior_std_simplex: "interior (convex hull (insert 0 Basis)) = {x::'a::euclidean_space. (\i\Basis. 0 < x\i) \ sum (\i. x\i) Basis < 1}" unfolding set_eq_iff mem_interior std_simplex proof (intro allI iffI CollectI; clarify) fix x :: 'a fix e assume "e > 0" and as: "ball x e \ {x. (\i\Basis. 0 \ x \ i) \ sum ((\) x) Basis \ 1}" show "(\i\Basis. 0 < x \ i) \ sum ((\) x) Basis < 1" proof safe fix i :: 'a assume i: "i \ Basis" then show "0 < x \ i" using as[THEN subsetD[where c="x - (e/2) *\<^sub>R i"]] and \e > 0\ by (force simp add: inner_simps) next have **: "dist x (x + (e/2) *\<^sub>R (SOME i. i\Basis)) < e" using \e > 0\ unfolding dist_norm by (auto intro!: mult_strict_left_mono simp: SOME_Basis) have "\i. i \ Basis \ (x + (e/2) *\<^sub>R (SOME i. i\Basis)) \ i = x\i + (if i = (SOME i. i\Basis) then e/2 else 0)" by (auto simp: SOME_Basis inner_Basis inner_simps) then have *: "sum ((\) (x + (e/2) *\<^sub>R (SOME i. i\Basis))) Basis = sum (\i. x\i + (if (SOME i. i\Basis) = i then e/2 else 0)) Basis" by (auto simp: intro!: sum.cong) have "sum ((\) x) Basis < sum ((\) (x + (e/2) *\<^sub>R (SOME i. i\Basis))) Basis" using \e > 0\ DIM_positive by (auto simp: SOME_Basis sum.distrib *) also have "\ \ 1" using ** as by force finally show "sum ((\) x) Basis < 1" by auto qed next fix x :: 'a assume as: "\i\Basis. 0 < x \ i" "sum ((\) x) Basis < 1" obtain a :: 'b where "a \ UNIV" using UNIV_witness .. let ?d = "(1 - sum ((\) x) Basis) / real (DIM('a))" show "\e>0. ball x e \ {x. (\i\Basis. 0 \ x \ i) \ sum ((\) x) Basis \ 1}" proof (rule_tac x="min (Min (((\) x) ` Basis)) D" for D in exI, intro conjI subsetI CollectI) fix y assume y: "y \ ball x (min (Min ((\) x ` Basis)) ?d)" have "sum ((\) y) Basis \ sum (\i. x\i + ?d) Basis" proof (rule sum_mono) fix i :: 'a assume i: "i \ Basis" have "\y\i - x\i\ \ norm (y - x)" by (metis Basis_le_norm i inner_commute inner_diff_right) also have "... < ?d" using y by (simp add: dist_norm norm_minus_commute) finally have "\y\i - x\i\ < ?d" . then show "y \ i \ x \ i + ?d" by auto qed also have "\ \ 1" unfolding sum.distrib sum_constant by (auto simp add: Suc_le_eq) finally show "sum ((\) y) Basis \ 1" . show "(\i\Basis. 0 \ y \ i)" proof safe fix i :: 'a assume i: "i \ Basis" have "norm (x - y) < Min (((\) x) ` Basis)" using y by (auto simp: dist_norm less_eq_real_def) also have "... \ x\i" using i by auto finally have "norm (x - y) < x\i" . then show "0 \ y\i" using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format, OF i] by (auto simp: inner_simps) qed next have "Min (((\) x) ` Basis) > 0" using as by simp moreover have "?d > 0" using as by (auto simp: Suc_le_eq) ultimately show "0 < min (Min ((\) x ` Basis)) ((1 - sum ((\) x) Basis) / real DIM('a))" by linarith qed qed lemma interior_std_simplex_nonempty: obtains a :: "'a::euclidean_space" where "a \ interior(convex hull (insert 0 Basis))" proof - let ?D = "Basis :: 'a set" let ?a = "sum (\b::'a. inverse (2 * real DIM('a)) *\<^sub>R b) Basis" { fix i :: 'a assume i: "i \ Basis" have "?a \ i = inverse (2 * real DIM('a))" by (rule trans[of _ "sum (\j. if i = j then inverse (2 * real DIM('a)) else 0) ?D"]) (simp_all add: sum.If_cases i) } note ** = this show ?thesis proof show "?a \ interior(convex hull (insert 0 Basis))" unfolding interior_std_simplex mem_Collect_eq proof safe fix i :: 'a assume i: "i \ Basis" show "0 < ?a \ i" unfolding **[OF i] by (auto simp add: Suc_le_eq) next have "sum ((\) ?a) ?D = sum (\i. inverse (2 * real DIM('a))) ?D" by (auto intro: sum.cong) also have "\ < 1" unfolding sum_constant divide_inverse[symmetric] by (auto simp add: field_simps) finally show "sum ((\) ?a) ?D < 1" by auto qed qed qed lemma rel_interior_substd_simplex: assumes D: "D \ Basis" shows "rel_interior (convex hull (insert 0 D)) = {x::'a::euclidean_space. (\i\D. 0 < x\i) \ (\i\D. x\i) < 1 \ (\i\Basis. i \ D \ x\i = 0)}" (is "_ = ?s") proof - have "finite D" using D finite_Basis finite_subset by blast show ?thesis proof (cases "D = {}") case True then show ?thesis using rel_interior_sing using euclidean_eq_iff[of _ 0] by auto next case False have h0: "affine hull (convex hull (insert 0 D)) = {x::'a::euclidean_space. (\i\Basis. i \ D \ x\i = 0)}" using affine_hull_convex_hull affine_hull_substd_basis assms by auto have aux: "\x::'a. \i\Basis. (\i\D. 0 \ x\i) \ (\i\Basis. i \ D \ x\i = 0) \ 0 \ x\i" by auto { fix x :: "'a::euclidean_space" assume x: "x \ rel_interior (convex hull (insert 0 D))" then obtain e where "e > 0" and "ball x e \ {xa. (\i\Basis. i \ D \ xa\i = 0)} \ convex hull (insert 0 D)" using mem_rel_interior_ball[of x "convex hull (insert 0 D)"] h0 by auto then have as: "\y. \dist x y < e \ (\i\Basis. i \ D \ y\i = 0)\ \ (\i\D. 0 \ y \ i) \ sum ((\) y) D \ 1" using assms by (force simp: substd_simplex) have x0: "(\i\Basis. i \ D \ x\i = 0)" using x rel_interior_subset substd_simplex[OF assms] by auto have "(\i\D. 0 < x \ i) \ sum ((\) x) D < 1 \ (\i\Basis. i \ D \ x\i = 0)" proof (intro conjI ballI) fix i :: 'a assume "i \ D" then have "\j\D. 0 \ (x - (e/2) *\<^sub>R i) \ j" using D \e > 0\ x0 by (intro as[THEN conjunct1]) (force simp: dist_norm inner_simps inner_Basis) then show "0 < x \ i" using \e > 0\ \i \ D\ D by (force simp: inner_simps inner_Basis) next obtain a where a: "a \ D" using \D \ {}\ by auto then have **: "dist x (x + (e/2) *\<^sub>R a) < e" using \e > 0\ norm_Basis[of a] D by (auto simp: dist_norm) have "\i. i \ Basis \ (x + (e/2) *\<^sub>R a) \ i = x\i + (if i = a then e/2 else 0)" using a D by (auto simp: inner_simps inner_Basis) then have *: "sum ((\) (x + (e/2) *\<^sub>R a)) D = sum (\i. x\i + (if a = i then e/2 else 0)) D" using D by (intro sum.cong) auto have "a \ Basis" using \a \ D\ D by auto then have h1: "(\i\Basis. i \ D \ (x + (e/2) *\<^sub>R a) \ i = 0)" using x0 D \a\D\ by (auto simp add: inner_add_left inner_Basis) have "sum ((\) x) D < sum ((\) (x + (e/2) *\<^sub>R a)) D" using \e > 0\ \a \ D\ \finite D\ by (auto simp add: * sum.distrib) also have "\ \ 1" using ** h1 as[rule_format, of "x + (e/2) *\<^sub>R a"] by auto finally show "sum ((\) x) D < 1" "\i. i\Basis \ i \ D \ x\i = 0" using x0 by auto qed } moreover { fix x :: "'a::euclidean_space" assume as: "x \ ?s" have "\i. 0 < x\i \ 0 = x\i \ 0 \ x\i" by auto moreover have "\i. i \ D \ i \ D" by auto ultimately have "\i. (\i\D. 0 < x\i) \ (\i. i \ D \ x\i = 0) \ 0 \ x\i" by metis then have h2: "x \ convex hull (insert 0 D)" using as assms by (force simp add: substd_simplex) obtain a where a: "a \ D" using \D \ {}\ by auto define d where "d \ (1 - sum ((\) x) D) / real (card D)" have "\e>0. ball x e \ {x. \i\Basis. i \ D \ x \ i = 0} \ convex hull insert 0 D" unfolding substd_simplex[OF assms] proof (intro exI; safe) have "0 < card D" using \D \ {}\ \finite D\ by (simp add: card_gt_0_iff) have "Min (((\) x) ` D) > 0" using as \D \ {}\ \finite D\ by (simp) moreover have "d > 0" using as \0 < card D\ by (auto simp: d_def) ultimately show "min (Min (((\) x) ` D)) d > 0" by auto fix y :: 'a assume y2: "\i\Basis. i \ D \ y\i = 0" assume "y \ ball x (min (Min ((\) x ` D)) d)" then have y: "dist x y < min (Min ((\) x ` D)) d" by auto have "sum ((\) y) D \ sum (\i. x\i + d) D" proof (rule sum_mono) fix i assume "i \ D" with D have i: "i \ Basis" by auto have "\y\i - x\i\ \ norm (y - x)" by (metis i inner_commute inner_diff_right norm_bound_Basis_le order_refl) also have "... < d" by (metis dist_norm min_less_iff_conj norm_minus_commute y) finally have "\y\i - x\i\ < d" . then show "y \ i \ x \ i + d" by auto qed also have "\ \ 1" unfolding sum.distrib sum_constant d_def using \0 < card D\ by auto finally show "sum ((\) y) D \ 1" . fix i :: 'a assume i: "i \ Basis" then show "0 \ y\i" proof (cases "i\D") case True have "norm (x - y) < x\i" using y Min_gr_iff[of "(\) x ` D" "norm (x - y)"] \0 < card D\ \i \ D\ by (simp add: dist_norm card_gt_0_iff) then show "0 \ y\i" using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format] by (auto simp: inner_simps) qed (use y2 in auto) qed then have "x \ rel_interior (convex hull (insert 0 D))" using h0 h2 rel_interior_ball by force } ultimately have "\x. x \ rel_interior (convex hull insert 0 D) \ x \ {x. (\i\D. 0 < x \ i) \ sum ((\) x) D < 1 \ (\i\Basis. i \ D \ x \ i = 0)}" by blast then show ?thesis by (rule set_eqI) qed qed lemma rel_interior_substd_simplex_nonempty: assumes "D \ {}" and "D \ Basis" obtains a :: "'a::euclidean_space" where "a \ rel_interior (convex hull (insert 0 D))" proof - let ?a = "sum (\b::'a::euclidean_space. inverse (2 * real (card D)) *\<^sub>R b) D" have "finite D" using assms finite_Basis infinite_super by blast then have d1: "0 < real (card D)" using \D \ {}\ by auto { fix i assume "i \ D" have "?a \ i = sum (\j. if i = j then inverse (2 * real (card D)) else 0) D" unfolding inner_sum_left using \i \ D\ by (auto simp: inner_Basis subsetD[OF assms(2)] intro: sum.cong) also have "... = inverse (2 * real (card D))" using \i \ D\ \finite D\ by auto finally have "?a \ i = inverse (2 * real (card D))" . } note ** = this show ?thesis proof show "?a \ rel_interior (convex hull (insert 0 D))" unfolding rel_interior_substd_simplex[OF assms(2)] proof safe fix i assume "i \ D" have "0 < inverse (2 * real (card D))" using d1 by auto also have "\ = ?a \ i" using **[of i] \i \ D\ by auto finally show "0 < ?a \ i" by auto next have "sum ((\) ?a) D = sum (\i. inverse (2 * real (card D))) D" by (rule sum.cong) (rule refl, rule **) also have "\ < 1" unfolding sum_constant divide_real_def[symmetric] by (auto simp add: field_simps) finally show "sum ((\) ?a) D < 1" by auto next fix i assume "i \ Basis" and "i \ D" have "?a \ span D" proof (rule span_sum[of D "(\b. b /\<^sub>R (2 * real (card D)))" D]) { fix x :: "'a::euclidean_space" assume "x \ D" then have "x \ span D" using span_base[of _ "D"] by auto then have "x /\<^sub>R (2 * real (card D)) \ span D" using span_mul[of x "D" "(inverse (real (card D)) / 2)"] by auto } then show "\x. x\D \ x /\<^sub>R (2 * real (card D)) \ span D" by auto qed then show "?a \ i = 0 " using \i \ D\ unfolding span_substd_basis[OF assms(2)] using \i \ Basis\ by auto qed qed qed subsection\<^marker>\tag unimportant\ \Relative interior of convex set\ lemma rel_interior_convex_nonempty_aux: fixes S :: "'n::euclidean_space set" assumes "convex S" and "0 \ S" shows "rel_interior S \ {}" proof (cases "S = {0}") case True then show ?thesis using rel_interior_sing by auto next case False obtain B where B: "independent B \ B \ S \ S \ span B \ card B = dim S" using basis_exists[of S] by metis then have "B \ {}" using B assms \S \ {0}\ span_empty by auto have "insert 0 B \ span B" using subspace_span[of B] subspace_0[of "span B"] span_superset by auto then have "span (insert 0 B) \ span B" using span_span[of B] span_mono[of "insert 0 B" "span B"] by blast then have "convex hull insert 0 B \ span B" using convex_hull_subset_span[of "insert 0 B"] by auto then have "span (convex hull insert 0 B) \ span B" using span_span[of B] span_mono[of "convex hull insert 0 B" "span B"] by blast then have *: "span (convex hull insert 0 B) = span B" using span_mono[of B "convex hull insert 0 B"] hull_subset[of "insert 0 B"] by auto then have "span (convex hull insert 0 B) = span S" using B span_mono[of B S] span_mono[of S "span B"] span_span[of B] by auto moreover have "0 \ affine hull (convex hull insert 0 B)" using hull_subset[of "convex hull insert 0 B"] hull_subset[of "insert 0 B"] by auto ultimately have **: "affine hull (convex hull insert 0 B) = affine hull S" using affine_hull_span_0[of "convex hull insert 0 B"] affine_hull_span_0[of "S"] assms hull_subset[of S] by auto obtain d and f :: "'n \ 'n" where fd: "card d = card B" "linear f" "f ` B = d" "f ` span B = {x. \i\Basis. i \ d \ x \ i = (0::real)} \ inj_on f (span B)" and d: "d \ Basis" using basis_to_substdbasis_subspace_isomorphism[of B,OF _ ] B by auto then have "bounded_linear f" using linear_conv_bounded_linear by auto have "d \ {}" using fd B \B \ {}\ by auto have "insert 0 d = f ` (insert 0 B)" using fd linear_0 by auto then have "(convex hull (insert 0 d)) = f ` (convex hull (insert 0 B))" using convex_hull_linear_image[of f "(insert 0 d)"] convex_hull_linear_image[of f "(insert 0 B)"] \linear f\ by auto moreover have "rel_interior (f ` (convex hull insert 0 B)) = f ` rel_interior (convex hull insert 0 B)" proof (rule rel_interior_injective_on_span_linear_image[OF \bounded_linear f\]) show "inj_on f (span (convex hull insert 0 B))" using fd * by auto qed ultimately have "rel_interior (convex hull insert 0 B) \ {}" using rel_interior_substd_simplex_nonempty[OF \d \ {}\ d] by fastforce moreover have "convex hull (insert 0 B) \ S" using B assms hull_mono[of "insert 0 B" "S" "convex"] convex_hull_eq by auto ultimately show ?thesis using subset_rel_interior[of "convex hull insert 0 B" S] ** by auto qed lemma rel_interior_eq_empty: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "rel_interior S = {} \ S = {}" proof - { assume "S \ {}" then obtain a where "a \ S" by auto then have "0 \ (+) (-a) ` S" using assms exI[of "(\x. x \ S \ - a + x = 0)" a] by auto then have "rel_interior ((+) (-a) ` S) \ {}" using rel_interior_convex_nonempty_aux[of "(+) (-a) ` S"] convex_translation[of S "-a"] assms by auto then have "rel_interior S \ {}" using rel_interior_translation [of "- a"] by simp } then show ?thesis by auto qed lemma interior_simplex_nonempty: fixes S :: "'N :: euclidean_space set" assumes "independent S" "finite S" "card S = DIM('N)" obtains a where "a \ interior (convex hull (insert 0 S))" proof - have "affine hull (insert 0 S) = UNIV" by (simp add: hull_inc affine_hull_span_0 dim_eq_full[symmetric] assms(1) assms(3) dim_eq_card_independent) moreover have "rel_interior (convex hull insert 0 S) \ {}" using rel_interior_eq_empty [of "convex hull (insert 0 S)"] by auto ultimately have "interior (convex hull insert 0 S) \ {}" by (simp add: rel_interior_interior) with that show ?thesis by auto qed lemma convex_rel_interior: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "convex (rel_interior S)" proof - { fix x y and u :: real assume assm: "x \ rel_interior S" "y \ rel_interior S" "0 \ u" "u \ 1" then have "x \ S" using rel_interior_subset by auto have "x - u *\<^sub>R (x-y) \ rel_interior S" proof (cases "0 = u") case False then have "0 < u" using assm by auto then show ?thesis using assm rel_interior_convex_shrink[of S y x u] assms \x \ S\ by auto next case True then show ?thesis using assm by auto qed then have "(1 - u) *\<^sub>R x + u *\<^sub>R y \ rel_interior S" by (simp add: algebra_simps) } then show ?thesis unfolding convex_alt by auto qed lemma convex_closure_rel_interior: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "closure (rel_interior S) = closure S" proof - have h1: "closure (rel_interior S) \ closure S" using closure_mono[of "rel_interior S" S] rel_interior_subset[of S] by auto show ?thesis proof (cases "S = {}") case False then obtain a where a: "a \ rel_interior S" using rel_interior_eq_empty assms by auto { fix x assume x: "x \ closure S" { assume "x = a" then have "x \ closure (rel_interior S)" using a unfolding closure_def by auto } moreover { assume "x \ a" { fix e :: real assume "e > 0" define e1 where "e1 = min 1 (e/norm (x - a))" then have e1: "e1 > 0" "e1 \ 1" "e1 * norm (x - a) \ e" using \x \ a\ \e > 0\ le_divide_eq[of e1 e "norm (x - a)"] by simp_all then have *: "x - e1 *\<^sub>R (x - a) \ rel_interior S" using rel_interior_closure_convex_shrink[of S a x e1] assms x a e1_def by auto have "\y. y \ rel_interior S \ y \ x \ dist y x \ e" using "*" \x \ a\ e1 by force } then have "x islimpt rel_interior S" unfolding islimpt_approachable_le by auto then have "x \ closure(rel_interior S)" unfolding closure_def by auto } ultimately have "x \ closure(rel_interior S)" by auto } then show ?thesis using h1 by auto qed auto qed lemma rel_interior_same_affine_hull: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "affine hull (rel_interior S) = affine hull S" by (metis assms closure_same_affine_hull convex_closure_rel_interior) lemma rel_interior_aff_dim: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "aff_dim (rel_interior S) = aff_dim S" by (metis aff_dim_affine_hull2 assms rel_interior_same_affine_hull) lemma rel_interior_rel_interior: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "rel_interior (rel_interior S) = rel_interior S" proof - have "openin (top_of_set (affine hull (rel_interior S))) (rel_interior S)" using openin_rel_interior[of S] rel_interior_same_affine_hull[of S] assms by auto then show ?thesis using rel_interior_def by auto qed lemma rel_interior_rel_open: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "rel_open (rel_interior S)" unfolding rel_open_def using rel_interior_rel_interior assms by auto lemma convex_rel_interior_closure_aux: fixes x y z :: "'n::euclidean_space" assumes "0 < a" "0 < b" "(a + b) *\<^sub>R z = a *\<^sub>R x + b *\<^sub>R y" obtains e where "0 < e" "e < 1" "z = y - e *\<^sub>R (y - x)" proof - define e where "e = a / (a + b)" have "z = (1 / (a + b)) *\<^sub>R ((a + b) *\<^sub>R z)" using assms by (simp add: eq_vector_fraction_iff) also have "\ = (1 / (a + b)) *\<^sub>R (a *\<^sub>R x + b *\<^sub>R y)" using assms scaleR_cancel_left[of "1/(a+b)" "(a + b) *\<^sub>R z" "a *\<^sub>R x + b *\<^sub>R y"] by auto also have "\ = y - e *\<^sub>R (y-x)" using e_def assms by (simp add: divide_simps vector_fraction_eq_iff) (simp add: algebra_simps) finally have "z = y - e *\<^sub>R (y-x)" by auto moreover have "e > 0" "e < 1" using e_def assms by auto ultimately show ?thesis using that[of e] by auto qed lemma convex_rel_interior_closure: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "rel_interior (closure S) = rel_interior S" proof (cases "S = {}") case True then show ?thesis using assms rel_interior_eq_empty by auto next case False have "rel_interior (closure S) \ rel_interior S" using subset_rel_interior[of S "closure S"] closure_same_affine_hull closure_subset by auto moreover { fix z assume z: "z \ rel_interior (closure S)" obtain x where x: "x \ rel_interior S" using \S \ {}\ assms rel_interior_eq_empty by auto have "z \ rel_interior S" proof (cases "x = z") case True then show ?thesis using x by auto next case False obtain e where e: "e > 0" "cball z e \ affine hull closure S \ closure S" using z rel_interior_cball[of "closure S"] by auto hence *: "0 < e/norm(z-x)" using e False by auto define y where "y = z + (e/norm(z-x)) *\<^sub>R (z-x)" have yball: "y \ cball z e" using y_def dist_norm[of z y] e by auto have "x \ affine hull closure S" using x rel_interior_subset_closure hull_inc[of x "closure S"] by blast moreover have "z \ affine hull closure S" using z rel_interior_subset hull_subset[of "closure S"] by blast ultimately have "y \ affine hull closure S" using y_def affine_affine_hull[of "closure S"] mem_affine_3_minus [of "affine hull closure S" z z x "e/norm(z-x)"] by auto then have "y \ closure S" using e yball by auto have "(1 + (e/norm(z-x))) *\<^sub>R z = (e/norm(z-x)) *\<^sub>R x + y" using y_def by (simp add: algebra_simps) then obtain e1 where "0 < e1" "e1 < 1" "z = y - e1 *\<^sub>R (y - x)" using * convex_rel_interior_closure_aux[of "e / norm (z - x)" 1 z x y] by (auto simp add: algebra_simps) then show ?thesis using rel_interior_closure_convex_shrink assms x \y \ closure S\ by fastforce qed } ultimately show ?thesis by auto qed lemma convex_interior_closure: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "interior (closure S) = interior S" using closure_aff_dim[of S] interior_rel_interior_gen[of S] interior_rel_interior_gen[of "closure S"] convex_rel_interior_closure[of S] assms by auto lemma closure_eq_rel_interior_eq: fixes S1 S2 :: "'n::euclidean_space set" assumes "convex S1" and "convex S2" shows "closure S1 = closure S2 \ rel_interior S1 = rel_interior S2" by (metis convex_rel_interior_closure convex_closure_rel_interior assms) lemma closure_eq_between: fixes S1 S2 :: "'n::euclidean_space set" assumes "convex S1" and "convex S2" shows "closure S1 = closure S2 \ rel_interior S1 \ S2 \ S2 \ closure S1" (is "?A \ ?B") proof assume ?A then show ?B by (metis assms closure_subset convex_rel_interior_closure rel_interior_subset) next assume ?B then have "closure S1 \ closure S2" by (metis assms(1) convex_closure_rel_interior closure_mono) moreover from \?B\ have "closure S1 \ closure S2" by (metis closed_closure closure_minimal) ultimately show ?A .. qed lemma open_inter_closure_rel_interior: fixes S A :: "'n::euclidean_space set" assumes "convex S" and "open A" shows "A \ closure S = {} \ A \ rel_interior S = {}" by (metis assms convex_closure_rel_interior open_Int_closure_eq_empty) lemma rel_interior_open_segment: fixes a :: "'a :: euclidean_space" shows "rel_interior(open_segment a b) = open_segment a b" proof (cases "a = b") case True then show ?thesis by auto next case False then have "open_segment a b = affine hull {a, b} \ ball ((a + b) /\<^sub>R 2) (norm (b - a) / 2)" by (simp add: open_segment_as_ball) then show ?thesis unfolding rel_interior_eq openin_open by (metis Elementary_Metric_Spaces.open_ball False affine_hull_open_segment) qed lemma rel_interior_closed_segment: fixes a :: "'a :: euclidean_space" shows "rel_interior(closed_segment a b) = (if a = b then {a} else open_segment a b)" proof (cases "a = b") case True then show ?thesis by auto next case False then show ?thesis by simp (metis closure_open_segment convex_open_segment convex_rel_interior_closure rel_interior_open_segment) qed lemmas rel_interior_segment = rel_interior_closed_segment rel_interior_open_segment subsection\The relative frontier of a set\ definition\<^marker>\tag important\ "rel_frontier S = closure S - rel_interior S" lemma rel_frontier_empty [simp]: "rel_frontier {} = {}" by (simp add: rel_frontier_def) lemma rel_frontier_eq_empty: fixes S :: "'n::euclidean_space set" shows "rel_frontier S = {} \ affine S" unfolding rel_frontier_def using rel_interior_subset_closure by (auto simp add: rel_interior_eq_closure [symmetric]) lemma rel_frontier_sing [simp]: fixes a :: "'n::euclidean_space" shows "rel_frontier {a} = {}" by (simp add: rel_frontier_def) lemma rel_frontier_affine_hull: fixes S :: "'a::euclidean_space set" shows "rel_frontier S \ affine hull S" using closure_affine_hull rel_frontier_def by fastforce lemma rel_frontier_cball [simp]: fixes a :: "'n::euclidean_space" shows "rel_frontier(cball a r) = (if r = 0 then {} else sphere a r)" proof (cases rule: linorder_cases [of r 0]) case less then show ?thesis by (force simp: sphere_def) next case equal then show ?thesis by simp next case greater then show ?thesis by simp (metis centre_in_ball empty_iff frontier_cball frontier_def interior_cball interior_rel_interior_gen rel_frontier_def) qed lemma rel_frontier_translation: fixes a :: "'a::euclidean_space" shows "rel_frontier((\x. a + x) ` S) = (\x. a + x) ` (rel_frontier S)" by (simp add: rel_frontier_def translation_diff rel_interior_translation closure_translation) lemma rel_frontier_nonempty_interior: fixes S :: "'n::euclidean_space set" shows "interior S \ {} \ rel_frontier S = frontier S" by (metis frontier_def interior_rel_interior_gen rel_frontier_def) lemma rel_frontier_frontier: fixes S :: "'n::euclidean_space set" shows "affine hull S = UNIV \ rel_frontier S = frontier S" by (simp add: frontier_def rel_frontier_def rel_interior_interior) lemma closest_point_in_rel_frontier: "\closed S; S \ {}; x \ affine hull S - rel_interior S\ \ closest_point S x \ rel_frontier S" by (simp add: closest_point_in_rel_interior closest_point_in_set rel_frontier_def) lemma closed_rel_frontier [iff]: fixes S :: "'n::euclidean_space set" shows "closed (rel_frontier S)" proof - have *: "closedin (top_of_set (affine hull S)) (closure S - rel_interior S)" by (simp add: closed_subset closedin_diff closure_affine_hull openin_rel_interior) show ?thesis proof (rule closedin_closed_trans[of "affine hull S" "rel_frontier S"]) show "closedin (top_of_set (affine hull S)) (rel_frontier S)" by (simp add: "*" rel_frontier_def) qed simp qed lemma closed_rel_boundary: fixes S :: "'n::euclidean_space set" shows "closed S \ closed(S - rel_interior S)" by (metis closed_rel_frontier closure_closed rel_frontier_def) lemma compact_rel_boundary: fixes S :: "'n::euclidean_space set" shows "compact S \ compact(S - rel_interior S)" by (metis bounded_diff closed_rel_boundary closure_eq compact_closure compact_imp_closed) lemma bounded_rel_frontier: fixes S :: "'n::euclidean_space set" shows "bounded S \ bounded(rel_frontier S)" by (simp add: bounded_closure bounded_diff rel_frontier_def) lemma compact_rel_frontier_bounded: fixes S :: "'n::euclidean_space set" shows "bounded S \ compact(rel_frontier S)" using bounded_rel_frontier closed_rel_frontier compact_eq_bounded_closed by blast lemma compact_rel_frontier: fixes S :: "'n::euclidean_space set" shows "compact S \ compact(rel_frontier S)" by (meson compact_eq_bounded_closed compact_rel_frontier_bounded) lemma convex_same_rel_interior_closure: fixes S :: "'n::euclidean_space set" shows "\convex S; convex T\ \ rel_interior S = rel_interior T \ closure S = closure T" by (simp add: closure_eq_rel_interior_eq) lemma convex_same_rel_interior_closure_straddle: fixes S :: "'n::euclidean_space set" shows "\convex S; convex T\ \ rel_interior S = rel_interior T \ rel_interior S \ T \ T \ closure S" by (simp add: closure_eq_between convex_same_rel_interior_closure) lemma convex_rel_frontier_aff_dim: fixes S1 S2 :: "'n::euclidean_space set" assumes "convex S1" and "convex S2" and "S2 \ {}" and "S1 \ rel_frontier S2" shows "aff_dim S1 < aff_dim S2" proof - have "S1 \ closure S2" using assms unfolding rel_frontier_def by auto then have *: "affine hull S1 \ affine hull S2" using hull_mono[of "S1" "closure S2"] closure_same_affine_hull[of S2] by blast then have "aff_dim S1 \ aff_dim S2" using * aff_dim_affine_hull[of S1] aff_dim_affine_hull[of S2] aff_dim_subset[of "affine hull S1" "affine hull S2"] by auto moreover { assume eq: "aff_dim S1 = aff_dim S2" then have "S1 \ {}" using aff_dim_empty[of S1] aff_dim_empty[of S2] \S2 \ {}\ by auto have **: "affine hull S1 = affine hull S2" by (simp_all add: * eq \S1 \ {}\ affine_dim_equal) obtain a where a: "a \ rel_interior S1" using \S1 \ {}\ rel_interior_eq_empty assms by auto obtain T where T: "open T" "a \ T \ S1" "T \ affine hull S1 \ S1" using mem_rel_interior[of a S1] a by auto then have "a \ T \ closure S2" using a assms unfolding rel_frontier_def by auto then obtain b where b: "b \ T \ rel_interior S2" using open_inter_closure_rel_interior[of S2 T] assms T by auto then have "b \ affine hull S1" using rel_interior_subset hull_subset[of S2] ** by auto then have "b \ S1" using T b by auto then have False using b assms unfolding rel_frontier_def by auto } ultimately show ?thesis using less_le by auto qed lemma convex_rel_interior_if: fixes S :: "'n::euclidean_space set" assumes "convex S" and "z \ rel_interior S" shows "\x\affine hull S. \m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" proof - obtain e1 where e1: "e1 > 0 \ cball z e1 \ affine hull S \ S" using mem_rel_interior_cball[of z S] assms by auto { fix x assume x: "x \ affine hull S" { assume "x \ z" define m where "m = 1 + e1/norm(x-z)" hence "m > 1" using e1 \x \ z\ by auto { fix e assume e: "e > 1 \ e \ m" have "z \ affine hull S" using assms rel_interior_subset hull_subset[of S] by auto then have *: "(1 - e)*\<^sub>R x + e *\<^sub>R z \ affine hull S" using mem_affine[of "affine hull S" x z "(1-e)" e] affine_affine_hull[of S] x by auto have "norm (z + e *\<^sub>R x - (x + e *\<^sub>R z)) = norm ((e - 1) *\<^sub>R (x - z))" by (simp add: algebra_simps) also have "\ = (e - 1) * norm (x-z)" using norm_scaleR e by auto also have "\ \ (m - 1) * norm (x - z)" using e mult_right_mono[of _ _ "norm(x-z)"] by auto also have "\ = (e1 / norm (x - z)) * norm (x - z)" using m_def by auto also have "\ = e1" using \x \ z\ e1 by simp finally have **: "norm (z + e *\<^sub>R x - (x + e *\<^sub>R z)) \ e1" by auto have "(1 - e)*\<^sub>R x+ e *\<^sub>R z \ cball z e1" using m_def ** unfolding cball_def dist_norm by (auto simp add: algebra_simps) then have "(1 - e) *\<^sub>R x+ e *\<^sub>R z \ S" using e * e1 by auto } then have "\m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S )" using \m> 1 \ by auto } moreover { assume "x = z" define m where "m = 1 + e1" then have "m > 1" using e1 by auto { fix e assume e: "e > 1 \ e \ m" then have "(1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using e1 x \x = z\ by (auto simp add: algebra_simps) then have "(1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using e by auto } then have "\m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" using \m > 1\ by auto } ultimately have "\m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S )" by blast } then show ?thesis by auto qed lemma convex_rel_interior_if2: fixes S :: "'n::euclidean_space set" assumes "convex S" assumes "z \ rel_interior S" shows "\x\affine hull S. \e. e > 1 \ (1 - e)*\<^sub>R x + e *\<^sub>R z \ S" using convex_rel_interior_if[of S z] assms by auto lemma convex_rel_interior_only_if: fixes S :: "'n::euclidean_space set" assumes "convex S" and "S \ {}" assumes "\x\S. \e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" shows "z \ rel_interior S" proof - obtain x where x: "x \ rel_interior S" using rel_interior_eq_empty assms by auto then have "x \ S" using rel_interior_subset by auto then obtain e where e: "e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using assms by auto define y where [abs_def]: "y = (1 - e) *\<^sub>R x + e *\<^sub>R z" then have "y \ S" using e by auto define e1 where "e1 = 1/e" then have "0 < e1 \ e1 < 1" using e by auto then have "z =y - (1 - e1) *\<^sub>R (y - x)" using e1_def y_def by (auto simp add: algebra_simps) then show ?thesis using rel_interior_convex_shrink[of S x y "1-e1"] \0 < e1 \ e1 < 1\ \y \ S\ x assms by auto qed lemma convex_rel_interior_iff: fixes S :: "'n::euclidean_space set" assumes "convex S" and "S \ {}" shows "z \ rel_interior S \ (\x\S. \e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" using assms hull_subset[of S "affine"] convex_rel_interior_if[of S z] convex_rel_interior_only_if[of S z] by auto lemma convex_rel_interior_iff2: fixes S :: "'n::euclidean_space set" assumes "convex S" and "S \ {}" shows "z \ rel_interior S \ (\x\affine hull S. \e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" using assms hull_subset[of S] convex_rel_interior_if2[of S z] convex_rel_interior_only_if[of S z] by auto lemma convex_interior_iff: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "z \ interior S \ (\x. \e. e > 0 \ z + e *\<^sub>R x \ S)" proof (cases "aff_dim S = int DIM('n)") case False { assume "z \ interior S" then have False using False interior_rel_interior_gen[of S] by auto } moreover { assume r: "\x. \e. e > 0 \ z + e *\<^sub>R x \ S" { fix x obtain e1 where e1: "e1 > 0 \ z + e1 *\<^sub>R (x - z) \ S" using r by auto obtain e2 where e2: "e2 > 0 \ z + e2 *\<^sub>R (z - x) \ S" using r by auto define x1 where [abs_def]: "x1 = z + e1 *\<^sub>R (x - z)" then have x1: "x1 \ affine hull S" using e1 hull_subset[of S] by auto define x2 where [abs_def]: "x2 = z + e2 *\<^sub>R (z - x)" then have x2: "x2 \ affine hull S" using e2 hull_subset[of S] by auto have *: "e1/(e1+e2) + e2/(e1+e2) = 1" using add_divide_distrib[of e1 e2 "e1+e2"] e1 e2 by simp then have "z = (e2/(e1+e2)) *\<^sub>R x1 + (e1/(e1+e2)) *\<^sub>R x2" by (simp add: x1_def x2_def algebra_simps) (simp add: "*" pth_8) then have z: "z \ affine hull S" using mem_affine[of "affine hull S" x1 x2 "e2/(e1+e2)" "e1/(e1+e2)"] x1 x2 affine_affine_hull[of S] * by auto have "x1 - x2 = (e1 + e2) *\<^sub>R (x - z)" using x1_def x2_def by (auto simp add: algebra_simps) then have "x = z+(1/(e1+e2)) *\<^sub>R (x1-x2)" using e1 e2 by simp then have "x \ affine hull S" using mem_affine_3_minus[of "affine hull S" z x1 x2 "1/(e1+e2)"] x1 x2 z affine_affine_hull[of S] by auto } then have "affine hull S = UNIV" by auto then have "aff_dim S = int DIM('n)" using aff_dim_affine_hull[of S] by (simp) then have False using False by auto } ultimately show ?thesis by auto next case True then have "S \ {}" using aff_dim_empty[of S] by auto have *: "affine hull S = UNIV" using True affine_hull_UNIV by auto { assume "z \ interior S" then have "z \ rel_interior S" using True interior_rel_interior_gen[of S] by auto then have **: "\x. \e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using convex_rel_interior_iff2[of S z] assms \S \ {}\ * by auto fix x obtain e1 where e1: "e1 > 1" "(1 - e1) *\<^sub>R (z - x) + e1 *\<^sub>R z \ S" using **[rule_format, of "z-x"] by auto define e where [abs_def]: "e = e1 - 1" then have "(1 - e1) *\<^sub>R (z - x) + e1 *\<^sub>R z = z + e *\<^sub>R x" by (simp add: algebra_simps) then have "e > 0" "z + e *\<^sub>R x \ S" using e1 e_def by auto then have "\e. e > 0 \ z + e *\<^sub>R x \ S" by auto } moreover { assume r: "\x. \e. e > 0 \ z + e *\<^sub>R x \ S" { fix x obtain e1 where e1: "e1 > 0" "z + e1 *\<^sub>R (z - x) \ S" using r[rule_format, of "z-x"] by auto define e where "e = e1 + 1" then have "z + e1 *\<^sub>R (z - x) = (1 - e) *\<^sub>R x + e *\<^sub>R z" by (simp add: algebra_simps) then have "e > 1" "(1 - e)*\<^sub>R x + e *\<^sub>R z \ S" using e1 e_def by auto then have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" by auto } then have "z \ rel_interior S" using convex_rel_interior_iff2[of S z] assms \S \ {}\ by auto then have "z \ interior S" using True interior_rel_interior_gen[of S] by auto } ultimately show ?thesis by auto qed subsubsection\<^marker>\tag unimportant\ \Relative interior and closure under common operations\ lemma rel_interior_inter_aux: "\{rel_interior S |S. S \ I} \ \I" proof - { fix y assume "y \ \{rel_interior S |S. S \ I}" then have y: "\S \ I. y \ rel_interior S" by auto { fix S assume "S \ I" then have "y \ S" using rel_interior_subset y by auto } then have "y \ \I" by auto } then show ?thesis by auto qed lemma convex_closure_rel_interior_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" shows "\{closure S |S. S \ I} \ closure (\{rel_interior S |S. S \ I})" proof - obtain x where x: "\S\I. x \ rel_interior S" using assms by auto { fix y assume "y \ \{closure S |S. S \ I}" then have y: "\S \ I. y \ closure S" by auto { assume "y = x" then have "y \ closure (\{rel_interior S |S. S \ I})" using x closure_subset[of "\{rel_interior S |S. S \ I}"] by auto } moreover { assume "y \ x" { fix e :: real assume e: "e > 0" define e1 where "e1 = min 1 (e/norm (y - x))" then have e1: "e1 > 0" "e1 \ 1" "e1 * norm (y - x) \ e" using \y \ x\ \e > 0\ le_divide_eq[of e1 e "norm (y - x)"] by simp_all define z where "z = y - e1 *\<^sub>R (y - x)" { fix S assume "S \ I" then have "z \ rel_interior S" using rel_interior_closure_convex_shrink[of S x y e1] assms x y e1 z_def by auto } then have *: "z \ \{rel_interior S |S. S \ I}" by auto have "\z. z \ \{rel_interior S |S. S \ I} \ z \ y \ dist z y \ e" using \y \ x\ z_def * e1 e dist_norm[of z y] by (rule_tac x="z" in exI) auto } then have "y islimpt \{rel_interior S |S. S \ I}" unfolding islimpt_approachable_le by blast then have "y \ closure (\{rel_interior S |S. S \ I})" unfolding closure_def by auto } ultimately have "y \ closure (\{rel_interior S |S. S \ I})" by auto } then show ?thesis by auto qed lemma convex_closure_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" shows "closure (\I) = \{closure S |S. S \ I}" proof - have "\{closure S |S. S \ I} \ closure (\{rel_interior S |S. S \ I})" using convex_closure_rel_interior_inter assms by auto moreover have "closure (\{rel_interior S |S. S \ I}) \ closure (\I)" using rel_interior_inter_aux closure_mono[of "\{rel_interior S |S. S \ I}" "\I"] by auto ultimately show ?thesis using closure_Int[of I] by auto qed lemma convex_inter_rel_interior_same_closure: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" shows "closure (\{rel_interior S |S. S \ I}) = closure (\I)" proof - have "\{closure S |S. S \ I} \ closure (\{rel_interior S |S. S \ I})" using convex_closure_rel_interior_inter assms by auto moreover have "closure (\{rel_interior S |S. S \ I}) \ closure (\I)" using rel_interior_inter_aux closure_mono[of "\{rel_interior S |S. S \ I}" "\I"] by auto ultimately show ?thesis using closure_Int[of I] by auto qed lemma convex_rel_interior_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" shows "rel_interior (\I) \ \{rel_interior S |S. S \ I}" proof - have "convex (\I)" using assms convex_Inter by auto moreover have "convex (\{rel_interior S |S. S \ I})" using assms convex_rel_interior by (force intro: convex_Inter) ultimately have "rel_interior (\{rel_interior S |S. S \ I}) = rel_interior (\I)" using convex_inter_rel_interior_same_closure assms closure_eq_rel_interior_eq[of "\{rel_interior S |S. S \ I}" "\I"] by blast then show ?thesis using rel_interior_subset[of "\{rel_interior S |S. S \ I}"] by auto qed lemma convex_rel_interior_finite_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" and "finite I" shows "rel_interior (\I) = \{rel_interior S |S. S \ I}" proof - have "\I \ {}" using assms rel_interior_inter_aux[of I] by auto have "convex (\I)" using convex_Inter assms by auto show ?thesis proof (cases "I = {}") case True then show ?thesis using Inter_empty rel_interior_UNIV by auto next case False { fix z assume z: "z \ \{rel_interior S |S. S \ I}" { fix x assume x: "x \ \I" { fix S assume S: "S \ I" then have "z \ rel_interior S" "x \ S" using z x by auto then have "\m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e)*\<^sub>R x + e *\<^sub>R z \ S)" using convex_rel_interior_if[of S z] S assms hull_subset[of S] by auto } then obtain mS where mS: "\S\I. mS S > 1 \ (\e. e > 1 \ e \ mS S \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" by metis define e where "e = Min (mS ` I)" then have "e \ mS ` I" using assms \I \ {}\ by simp then have "e > 1" using mS by auto moreover have "\S\I. e \ mS S" using e_def assms by auto ultimately have "\e > 1. (1 - e) *\<^sub>R x + e *\<^sub>R z \ \I" using mS by auto } then have "z \ rel_interior (\I)" using convex_rel_interior_iff[of "\I" z] \\I \ {}\ \convex (\I)\ by auto } then show ?thesis using convex_rel_interior_inter[of I] assms by auto qed qed lemma convex_closure_inter_two: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "convex T" assumes "rel_interior S \ rel_interior T \ {}" shows "closure (S \ T) = closure S \ closure T" using convex_closure_inter[of "{S,T}"] assms by auto lemma convex_rel_interior_inter_two: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "convex T" and "rel_interior S \ rel_interior T \ {}" shows "rel_interior (S \ T) = rel_interior S \ rel_interior T" using convex_rel_interior_finite_inter[of "{S,T}"] assms by auto lemma convex_affine_closure_Int: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "affine T" and "rel_interior S \ T \ {}" shows "closure (S \ T) = closure S \ T" -proof - - have "affine hull T = T" - using assms by auto - then have "rel_interior T = T" - using rel_interior_affine_hull[of T] by metis - moreover have "closure T = T" - using assms affine_closed[of T] by auto - ultimately show ?thesis - using convex_closure_inter_two[of S T] assms affine_imp_convex by auto -qed + by (metis affine_imp_convex assms convex_closure_inter_two rel_interior_affine rel_interior_eq_closure) lemma connected_component_1_gen: fixes S :: "'a :: euclidean_space set" assumes "DIM('a) = 1" shows "connected_component S a b \ closed_segment a b \ S" unfolding connected_component_def by (metis (no_types, lifting) assms subsetD subsetI convex_contains_segment convex_segment(1) ends_in_segment connected_convex_1_gen) lemma connected_component_1: fixes S :: "real set" shows "connected_component S a b \ closed_segment a b \ S" by (simp add: connected_component_1_gen) lemma convex_affine_rel_interior_Int: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "affine T" and "rel_interior S \ T \ {}" shows "rel_interior (S \ T) = rel_interior S \ T" -proof - - have "affine hull T = T" - using assms by auto - then have "rel_interior T = T" - using rel_interior_affine_hull[of T] by metis - moreover have "closure T = T" - using assms affine_closed[of T] by auto - ultimately show ?thesis - using convex_rel_interior_inter_two[of S T] assms affine_imp_convex by auto -qed + by (simp add: affine_imp_convex assms convex_rel_interior_inter_two rel_interior_affine) lemma convex_affine_rel_frontier_Int: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "affine T" and "interior S \ T \ {}" shows "rel_frontier(S \ T) = frontier S \ T" using assms unfolding rel_frontier_def frontier_def using convex_affine_closure_Int convex_affine_rel_interior_Int rel_interior_nonempty_interior by fastforce lemma rel_interior_convex_Int_affine: fixes S :: "'a::euclidean_space set" assumes "convex S" "affine T" "interior S \ T \ {}" - shows "rel_interior(S \ T) = interior S \ T" -proof - - obtain a where aS: "a \ interior S" and aT:"a \ T" - using assms by force - have "rel_interior S = interior S" - by (metis (no_types) aS affine_hull_nonempty_interior equals0D rel_interior_interior) - then show ?thesis - by (metis (no_types) affine_imp_convex assms convex_rel_interior_inter_two hull_same rel_interior_affine_hull) -qed - -lemma closure_convex_Int_affine: - fixes S :: "'a::euclidean_space set" - assumes "convex S" "affine T" "rel_interior S \ T \ {}" - shows "closure(S \ T) = closure S \ T" -proof - have "closure (S \ T) \ closure T" - by (simp add: closure_mono) - also have "... \ T" - by (simp add: affine_closed assms) - finally show "closure(S \ T) \ closure S \ T" - by (simp add: closure_mono) -next - obtain a where "a \ rel_interior S" "a \ T" - using assms by auto - then have ssT: "subspace ((\x. (-a)+x) ` T)" and "a \ S" - using affine_diffs_subspace rel_interior_subset assms by blast+ - show "closure S \ T \ closure (S \ T)" - proof - fix x assume "x \ closure S \ T" - show "x \ closure (S \ T)" - proof (cases "x = a") - case True - then show ?thesis - using \a \ S\ \a \ T\ closure_subset by fastforce - next - case False - then have "x \ closure(open_segment a x)" - by auto - then show ?thesis - using \x \ closure S \ T\ assms convex_affine_closure_Int by blast - qed - qed -qed + shows "rel_interior(S \ T) = interior S \ T" + by (metis Int_emptyI assms convex_affine_rel_interior_Int empty_iff interior_rel_interior_gen) lemma subset_rel_interior_convex: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "convex T" and "S \ closure T" and "\ S \ rel_frontier T" shows "rel_interior S \ rel_interior T" proof - have *: "S \ closure T = S" using assms by auto have "\ rel_interior S \ rel_frontier T" using closure_mono[of "rel_interior S" "rel_frontier T"] closed_rel_frontier[of T] closure_closed[of S] convex_closure_rel_interior[of S] closure_subset[of S] assms by auto then have "rel_interior S \ rel_interior (closure T) \ {}" using assms rel_frontier_def[of T] rel_interior_subset convex_rel_interior_closure[of T] by auto then have "rel_interior S \ rel_interior T = rel_interior (S \ closure T)" using assms convex_closure convex_rel_interior_inter_two[of S "closure T"] convex_rel_interior_closure[of T] by auto also have "\ = rel_interior S" using * by auto finally show ?thesis by auto qed lemma rel_interior_convex_linear_image: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" and "convex S" shows "f ` (rel_interior S) = rel_interior (f ` S)" proof (cases "S = {}") case True then show ?thesis using assms by auto next case False interpret linear f by fact have *: "f ` (rel_interior S) \ f ` S" unfolding image_mono using rel_interior_subset by auto have "f ` S \ f ` (closure S)" unfolding image_mono using closure_subset by auto also have "\ = f ` (closure (rel_interior S))" using convex_closure_rel_interior assms by auto also have "\ \ closure (f ` (rel_interior S))" using closure_linear_image_subset assms by auto finally have "closure (f ` S) = closure (f ` rel_interior S)" using closure_mono[of "f ` S" "closure (f ` rel_interior S)"] closure_closure closure_mono[of "f ` rel_interior S" "f ` S"] * by auto then have "rel_interior (f ` S) = rel_interior (f ` rel_interior S)" using assms convex_rel_interior linear_conv_bounded_linear[of f] convex_linear_image[of _ S] convex_linear_image[of _ "rel_interior S"] closure_eq_rel_interior_eq[of "f ` S" "f ` rel_interior S"] by auto then have "rel_interior (f ` S) \ f ` rel_interior S" using rel_interior_subset by auto moreover { fix z assume "z \ f ` rel_interior S" then obtain z1 where z1: "z1 \ rel_interior S" "f z1 = z" by auto { fix x assume "x \ f ` S" then obtain x1 where x1: "x1 \ S" "f x1 = x" by auto then obtain e where e: "e > 1" "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1 \ S" using convex_rel_interior_iff[of S z1] \convex S\ x1 z1 by auto moreover have "f ((1 - e) *\<^sub>R x1 + e *\<^sub>R z1) = (1 - e) *\<^sub>R x + e *\<^sub>R z" using x1 z1 by (simp add: linear_add linear_scale \linear f\) ultimately have "(1 - e) *\<^sub>R x + e *\<^sub>R z \ f ` S" using imageI[of "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1" S f] by auto then have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ f ` S" using e by auto } then have "z \ rel_interior (f ` S)" using convex_rel_interior_iff[of "f ` S" z] \convex S\ \linear f\ \S \ {}\ convex_linear_image[of f S] linear_conv_bounded_linear[of f] by auto } ultimately show ?thesis by auto qed lemma rel_interior_convex_linear_preimage: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" and "convex S" and "f -` (rel_interior S) \ {}" shows "rel_interior (f -` S) = f -` (rel_interior S)" proof - interpret linear f by fact have "S \ {}" using assms by auto have nonemp: "f -` S \ {}" by (metis assms(3) rel_interior_subset subset_empty vimage_mono) then have "S \ (range f) \ {}" by auto have conv: "convex (f -` S)" using convex_linear_vimage assms by auto then have "convex (S \ range f)" by (simp add: assms(2) convex_Int convex_linear_image linear_axioms) { fix z assume "z \ f -` (rel_interior S)" then have z: "f z \ rel_interior S" by auto { fix x assume "x \ f -` S" then have "f x \ S" by auto then obtain e where e: "e > 1" "(1 - e) *\<^sub>R f x + e *\<^sub>R f z \ S" using convex_rel_interior_iff[of S "f z"] z assms \S \ {}\ by auto moreover have "(1 - e) *\<^sub>R f x + e *\<^sub>R f z = f ((1 - e) *\<^sub>R x + e *\<^sub>R z)" using \linear f\ by (simp add: linear_iff) ultimately have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ f -` S" using e by auto } then have "z \ rel_interior (f -` S)" using convex_rel_interior_iff[of "f -` S" z] conv nonemp by auto } moreover { fix z assume z: "z \ rel_interior (f -` S)" { fix x assume "x \ S \ range f" then obtain y where y: "f y = x" "y \ f -` S" by auto then obtain e where e: "e > 1" "(1 - e) *\<^sub>R y + e *\<^sub>R z \ f -` S" using convex_rel_interior_iff[of "f -` S" z] z conv by auto moreover have "(1 - e) *\<^sub>R x + e *\<^sub>R f z = f ((1 - e) *\<^sub>R y + e *\<^sub>R z)" using \linear f\ y by (simp add: linear_iff) ultimately have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R f z \ S \ range f" using e by auto } then have "f z \ rel_interior (S \ range f)" using \convex (S \ (range f))\ \S \ range f \ {}\ convex_rel_interior_iff[of "S \ (range f)" "f z"] by auto moreover have "affine (range f)" by (simp add: linear_axioms linear_subspace_image subspace_imp_affine) ultimately have "f z \ rel_interior S" using convex_affine_rel_interior_Int[of S "range f"] assms by auto then have "z \ f -` (rel_interior S)" by auto } ultimately show ?thesis by auto qed lemma rel_interior_Times: fixes S :: "'n::euclidean_space set" and T :: "'m::euclidean_space set" assumes "convex S" and "convex T" shows "rel_interior (S \ T) = rel_interior S \ rel_interior T" proof (cases "S = {} \ T = {}") case True then show ?thesis by auto next case False then have "S \ {}" "T \ {}" by auto then have ri: "rel_interior S \ {}" "rel_interior T \ {}" using rel_interior_eq_empty assms by auto then have "fst -` rel_interior S \ {}" using fst_vimage_eq_Times[of "rel_interior S"] by auto then have "rel_interior ((fst :: 'n * 'm \ 'n) -` S) = fst -` rel_interior S" using linear_fst \convex S\ rel_interior_convex_linear_preimage[of fst S] by auto then have s: "rel_interior (S \ (UNIV :: 'm set)) = rel_interior S \ UNIV" by (simp add: fst_vimage_eq_Times) from ri have "snd -` rel_interior T \ {}" using snd_vimage_eq_Times[of "rel_interior T"] by auto then have "rel_interior ((snd :: 'n * 'm \ 'm) -` T) = snd -` rel_interior T" using linear_snd \convex T\ rel_interior_convex_linear_preimage[of snd T] by auto then have t: "rel_interior ((UNIV :: 'n set) \ T) = UNIV \ rel_interior T" by (simp add: snd_vimage_eq_Times) from s t have *: "rel_interior (S \ (UNIV :: 'm set)) \ rel_interior ((UNIV :: 'n set) \ T) = rel_interior S \ rel_interior T" by auto have "S \ T = S \ (UNIV :: 'm set) \ (UNIV :: 'n set) \ T" by auto then have "rel_interior (S \ T) = rel_interior ((S \ (UNIV :: 'm set)) \ ((UNIV :: 'n set) \ T))" by auto also have "\ = rel_interior (S \ (UNIV :: 'm set)) \ rel_interior ((UNIV :: 'n set) \ T)" using * ri assms convex_Times by (subst convex_rel_interior_inter_two) auto finally show ?thesis using * by auto qed lemma rel_interior_scaleR: fixes S :: "'n::euclidean_space set" assumes "c \ 0" shows "((*\<^sub>R) c) ` (rel_interior S) = rel_interior (((*\<^sub>R) c) ` S)" using rel_interior_injective_linear_image[of "((*\<^sub>R) c)" S] linear_conv_bounded_linear[of "(*\<^sub>R) c"] linear_scaleR injective_scaleR[of c] assms by auto lemma rel_interior_convex_scaleR: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "((*\<^sub>R) c) ` (rel_interior S) = rel_interior (((*\<^sub>R) c) ` S)" by (metis assms linear_scaleR rel_interior_convex_linear_image) lemma convex_rel_open_scaleR: fixes S :: "'n::euclidean_space set" assumes "convex S" and "rel_open S" shows "convex (((*\<^sub>R) c) ` S) \ rel_open (((*\<^sub>R) c) ` S)" by (metis assms convex_scaling rel_interior_convex_scaleR rel_open_def) lemma convex_rel_open_finite_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set) \ rel_open S" and "finite I" shows "convex (\I) \ rel_open (\I)" proof (cases "\{rel_interior S |S. S \ I} = {}") case True then have "\I = {}" using assms unfolding rel_open_def by auto then show ?thesis unfolding rel_open_def by auto next case False then have "rel_open (\I)" using assms unfolding rel_open_def using convex_rel_interior_finite_inter[of I] by auto then show ?thesis using convex_Inter assms by auto qed lemma convex_rel_open_linear_image: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" and "convex S" and "rel_open S" shows "convex (f ` S) \ rel_open (f ` S)" by (metis assms convex_linear_image rel_interior_convex_linear_image rel_open_def) lemma convex_rel_open_linear_preimage: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" and "convex S" and "rel_open S" shows "convex (f -` S) \ rel_open (f -` S)" proof (cases "f -` (rel_interior S) = {}") case True then have "f -` S = {}" using assms unfolding rel_open_def by auto then show ?thesis unfolding rel_open_def by auto next case False then have "rel_open (f -` S)" using assms unfolding rel_open_def using rel_interior_convex_linear_preimage[of f S] by auto then show ?thesis using convex_linear_vimage assms by auto qed lemma rel_interior_projection: fixes S :: "('m::euclidean_space \ 'n::euclidean_space) set" and f :: "'m::euclidean_space \ 'n::euclidean_space set" assumes "convex S" and "f = (\y. {z. (y, z) \ S})" shows "(y, z) \ rel_interior S \ (y \ rel_interior {y. (f y \ {})} \ z \ rel_interior (f y))" proof - { fix y assume "y \ {y. f y \ {}}" then obtain z where "(y, z) \ S" using assms by auto then have "\x. x \ S \ y = fst x" by auto then obtain x where "x \ S" "y = fst x" by blast then have "y \ fst ` S" unfolding image_def by auto } then have "fst ` S = {y. f y \ {}}" unfolding fst_def using assms by auto then have h1: "fst ` rel_interior S = rel_interior {y. f y \ {}}" using rel_interior_convex_linear_image[of fst S] assms linear_fst by auto { fix y assume "y \ rel_interior {y. f y \ {}}" then have "y \ fst ` rel_interior S" using h1 by auto then have *: "rel_interior S \ fst -` {y} \ {}" by auto moreover have aff: "affine (fst -` {y})" unfolding affine_alt by (simp add: algebra_simps) ultimately have **: "rel_interior (S \ fst -` {y}) = rel_interior S \ fst -` {y}" using convex_affine_rel_interior_Int[of S "fst -` {y}"] assms by auto have conv: "convex (S \ fst -` {y})" using convex_Int assms aff affine_imp_convex by auto { fix x assume "x \ f y" then have "(y, x) \ S \ (fst -` {y})" using assms by auto moreover have "x = snd (y, x)" by auto ultimately have "x \ snd ` (S \ fst -` {y})" by blast } then have "snd ` (S \ fst -` {y}) = f y" using assms by auto then have ***: "rel_interior (f y) = snd ` rel_interior (S \ fst -` {y})" using rel_interior_convex_linear_image[of snd "S \ fst -` {y}"] linear_snd conv by auto { fix z assume "z \ rel_interior (f y)" then have "z \ snd ` rel_interior (S \ fst -` {y})" using *** by auto moreover have "{y} = fst ` rel_interior (S \ fst -` {y})" using * ** rel_interior_subset by auto ultimately have "(y, z) \ rel_interior (S \ fst -` {y})" by force then have "(y,z) \ rel_interior S" using ** by auto } moreover { fix z assume "(y, z) \ rel_interior S" then have "(y, z) \ rel_interior (S \ fst -` {y})" using ** by auto then have "z \ snd ` rel_interior (S \ fst -` {y})" by (metis Range_iff snd_eq_Range) then have "z \ rel_interior (f y)" using *** by auto } ultimately have "\z. (y, z) \ rel_interior S \ z \ rel_interior (f y)" by auto } then have h2: "\y z. y \ rel_interior {t. f t \ {}} \ (y, z) \ rel_interior S \ z \ rel_interior (f y)" by auto { fix y z assume asm: "(y, z) \ rel_interior S" then have "y \ fst ` rel_interior S" by (metis Domain_iff fst_eq_Domain) then have "y \ rel_interior {t. f t \ {}}" using h1 by auto then have "y \ rel_interior {t. f t \ {}}" and "(z \ rel_interior (f y))" using h2 asm by auto } then show ?thesis using h2 by blast qed lemma rel_frontier_Times: fixes S :: "'n::euclidean_space set" and T :: "'m::euclidean_space set" assumes "convex S" and "convex T" shows "rel_frontier S \ rel_frontier T \ rel_frontier (S \ T)" by (force simp: rel_frontier_def rel_interior_Times assms closure_Times) subsubsection\<^marker>\tag unimportant\ \Relative interior of convex cone\ lemma cone_rel_interior: fixes S :: "'m::euclidean_space set" assumes "cone S" shows "cone ({0} \ rel_interior S)" proof (cases "S = {}") case True then show ?thesis by (simp add: cone_0) next case False then have *: "0 \ S \ (\c. c > 0 \ (*\<^sub>R) c ` S = S)" using cone_iff[of S] assms by auto then have *: "0 \ ({0} \ rel_interior S)" and "\c. c > 0 \ (*\<^sub>R) c ` ({0} \ rel_interior S) = ({0} \ rel_interior S)" by (auto simp add: rel_interior_scaleR) then show ?thesis using cone_iff[of "{0} \ rel_interior S"] by auto qed lemma rel_interior_convex_cone_aux: fixes S :: "'m::euclidean_space set" assumes "convex S" shows "(c, x) \ rel_interior (cone hull ({(1 :: real)} \ S)) \ c > 0 \ x \ (((*\<^sub>R) c) ` (rel_interior S))" proof (cases "S = {}") case True then show ?thesis by (simp add: cone_hull_empty) next case False then obtain s where "s \ S" by auto have conv: "convex ({(1 :: real)} \ S)" using convex_Times[of "{(1 :: real)}" S] assms convex_singleton[of "1 :: real"] by auto define f where "f y = {z. (y, z) \ cone hull ({1 :: real} \ S)}" for y then have *: "(c, x) \ rel_interior (cone hull ({(1 :: real)} \ S)) = (c \ rel_interior {y. f y \ {}} \ x \ rel_interior (f c))" using convex_cone_hull[of "{(1 :: real)} \ S"] conv by (subst rel_interior_projection) auto { fix y :: real assume "y \ 0" then have "y *\<^sub>R (1,s) \ cone hull ({1 :: real} \ S)" using cone_hull_expl[of "{(1 :: real)} \ S"] \s \ S\ by auto then have "f y \ {}" using f_def by auto } then have "{y. f y \ {}} = {0..}" using f_def cone_hull_expl[of "{1 :: real} \ S"] by auto then have **: "rel_interior {y. f y \ {}} = {0<..}" using rel_interior_real_semiline by auto { fix c :: real assume "c > 0" then have "f c = ((*\<^sub>R) c ` S)" using f_def cone_hull_expl[of "{1 :: real} \ S"] by auto then have "rel_interior (f c) = (*\<^sub>R) c ` rel_interior S" using rel_interior_convex_scaleR[of S c] assms by auto } then show ?thesis using * ** by auto qed lemma rel_interior_convex_cone: fixes S :: "'m::euclidean_space set" assumes "convex S" shows "rel_interior (cone hull ({1 :: real} \ S)) = {(c, c *\<^sub>R x) | c x. c > 0 \ x \ rel_interior S}" (is "?lhs = ?rhs") proof - { fix z assume "z \ ?lhs" have *: "z = (fst z, snd z)" by auto then have "z \ ?rhs" using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms \z \ ?lhs\ by fastforce } moreover { fix z assume "z \ ?rhs" then have "z \ ?lhs" using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms by auto } ultimately show ?thesis by blast qed lemma convex_hull_finite_union: assumes "finite I" assumes "\i\I. convex (S i) \ (S i) \ {}" shows "convex hull (\(S ` I)) = {sum (\i. c i *\<^sub>R s i) I | c s. (\i\I. c i \ 0) \ sum c I = 1 \ (\i\I. s i \ S i)}" (is "?lhs = ?rhs") proof - have "?lhs \ ?rhs" proof fix x assume "x \ ?rhs" then obtain c s where *: "sum (\i. c i *\<^sub>R s i) I = x" "sum c I = 1" "(\i\I. c i \ 0) \ (\i\I. s i \ S i)" by auto then have "\i\I. s i \ convex hull (\(S ` I))" using hull_subset[of "\(S ` I)" convex] by auto then show "x \ ?lhs" unfolding *(1)[symmetric] using * assms convex_convex_hull by (subst convex_sum) auto qed { fix i assume "i \ I" with assms have "\p. p \ S i" by auto } then obtain p where p: "\i\I. p i \ S i" by metis { fix i assume "i \ I" { fix x assume "x \ S i" define c where "c j = (if j = i then 1::real else 0)" for j then have *: "sum c I = 1" using \finite I\ \i \ I\ sum.delta[of I i "\j::'a. 1::real"] by auto define s where "s j = (if j = i then x else p j)" for j then have "\j. c j *\<^sub>R s j = (if j = i then x else 0)" using c_def by (auto simp add: algebra_simps) then have "x = sum (\i. c i *\<^sub>R s i) I" using s_def c_def \finite I\ \i \ I\ sum.delta[of I i "\j::'a. x"] by auto moreover have "(\i\I. 0 \ c i) \ sum c I = 1 \ (\i\I. s i \ S i)" using * c_def s_def p \x \ S i\ by auto ultimately have "x \ ?rhs" by force } then have "?rhs \ S i" by auto } then have *: "?rhs \ \(S ` I)" by auto { fix u v :: real assume uv: "u \ 0 \ v \ 0 \ u + v = 1" fix x y assume xy: "x \ ?rhs \ y \ ?rhs" from xy obtain c s where xc: "x = sum (\i. c i *\<^sub>R s i) I \ (\i\I. c i \ 0) \ sum c I = 1 \ (\i\I. s i \ S i)" by auto from xy obtain d t where yc: "y = sum (\i. d i *\<^sub>R t i) I \ (\i\I. d i \ 0) \ sum d I = 1 \ (\i\I. t i \ S i)" by auto define e where "e i = u * c i + v * d i" for i have ge0: "\i\I. e i \ 0" using e_def xc yc uv by simp have "sum (\i. u * c i) I = u * sum c I" by (simp add: sum_distrib_left) moreover have "sum (\i. v * d i) I = v * sum d I" by (simp add: sum_distrib_left) ultimately have sum1: "sum e I = 1" using e_def xc yc uv by (simp add: sum.distrib) define q where "q i = (if e i = 0 then p i else (u * c i / e i) *\<^sub>R s i + (v * d i / e i) *\<^sub>R t i)" for i { fix i assume i: "i \ I" have "q i \ S i" proof (cases "e i = 0") case True then show ?thesis using i p q_def by auto next case False then show ?thesis using mem_convex_alt[of "S i" "s i" "t i" "u * (c i)" "v * (d i)"] mult_nonneg_nonneg[of u "c i"] mult_nonneg_nonneg[of v "d i"] assms q_def e_def i False xc yc uv by (auto simp del: mult_nonneg_nonneg) qed } then have qs: "\i\I. q i \ S i" by auto { fix i assume i: "i \ I" have "(u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i = e i *\<^sub>R q i" proof (cases "e i = 0") case True have ge: "u * (c i) \ 0 \ v * d i \ 0" using xc yc uv i by simp moreover from ge have "u * c i \ 0 \ v * d i \ 0" using True e_def i by simp ultimately have "u * c i = 0 \ v * d i = 0" by auto with True show ?thesis by auto next case False then have "(u * (c i)/(e i))*\<^sub>R (s i)+(v * (d i)/(e i))*\<^sub>R (t i) = q i" using q_def by auto then have "e i *\<^sub>R ((u * (c i)/(e i))*\<^sub>R (s i)+(v * (d i)/(e i))*\<^sub>R (t i)) = (e i) *\<^sub>R (q i)" by auto with False show ?thesis by (simp add: algebra_simps) qed } then have *: "\i\I. (u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i = e i *\<^sub>R q i" by auto have "u *\<^sub>R x + v *\<^sub>R y = sum (\i. (u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i) I" using xc yc by (simp add: algebra_simps scaleR_right.sum sum.distrib) also have "\ = sum (\i. e i *\<^sub>R q i) I" using * by auto finally have "u *\<^sub>R x + v *\<^sub>R y = sum (\i. (e i) *\<^sub>R (q i)) I" by auto then have "u *\<^sub>R x + v *\<^sub>R y \ ?rhs" using ge0 sum1 qs by auto } then have "convex ?rhs" unfolding convex_def by auto then show ?thesis using \?lhs \ ?rhs\ * hull_minimal[of "\(S ` I)" ?rhs convex] by blast qed lemma convex_hull_union_two: fixes S T :: "'m::euclidean_space set" assumes "convex S" and "S \ {}" and "convex T" and "T \ {}" shows "convex hull (S \ T) = {u *\<^sub>R s + v *\<^sub>R t | u v s t. u \ 0 \ v \ 0 \ u + v = 1 \ s \ S \ t \ T}" (is "?lhs = ?rhs") proof define I :: "nat set" where "I = {1, 2}" define s where "s i = (if i = (1::nat) then S else T)" for i have "\(s ` I) = S \ T" using s_def I_def by auto then have "convex hull (\(s ` I)) = convex hull (S \ T)" by auto moreover have "convex hull \(s ` I) = {\ i\I. c i *\<^sub>R sa i | c sa. (\i\I. 0 \ c i) \ sum c I = 1 \ (\i\I. sa i \ s i)}" using assms s_def I_def by (subst convex_hull_finite_union) auto moreover have "{\i\I. c i *\<^sub>R sa i | c sa. (\i\I. 0 \ c i) \ sum c I = 1 \ (\i\I. sa i \ s i)} \ ?rhs" using s_def I_def by auto ultimately show "?lhs \ ?rhs" by auto { fix x assume "x \ ?rhs" then obtain u v s t where *: "x = u *\<^sub>R s + v *\<^sub>R t \ u \ 0 \ v \ 0 \ u + v = 1 \ s \ S \ t \ T" by auto then have "x \ convex hull {s, t}" using convex_hull_2[of s t] by auto then have "x \ convex hull (S \ T)" using * hull_mono[of "{s, t}" "S \ T"] by auto } then show "?lhs \ ?rhs" by blast qed proposition ray_to_rel_frontier: fixes a :: "'a::real_inner" assumes "bounded S" and a: "a \ rel_interior S" and aff: "(a + l) \ affine hull S" and "l \ 0" obtains d where "0 < d" "(a + d *\<^sub>R l) \ rel_frontier S" "\e. \0 \ e; e < d\ \ (a + e *\<^sub>R l) \ rel_interior S" proof - have aaff: "a \ affine hull S" by (meson a hull_subset rel_interior_subset rev_subsetD) let ?D = "{d. 0 < d \ a + d *\<^sub>R l \ rel_interior S}" obtain B where "B > 0" and B: "S \ ball a B" using bounded_subset_ballD [OF \bounded S\] by blast have "a + (B / norm l) *\<^sub>R l \ ball a B" by (simp add: dist_norm \l \ 0\) with B have "a + (B / norm l) *\<^sub>R l \ rel_interior S" using rel_interior_subset subsetCE by blast with \B > 0\ \l \ 0\ have nonMT: "?D \ {}" using divide_pos_pos zero_less_norm_iff by fastforce have bdd: "bdd_below ?D" by (metis (no_types, lifting) bdd_belowI le_less mem_Collect_eq) have relin_Ex: "\x. x \ rel_interior S \ \e>0. \x'\affine hull S. dist x' x < e \ x' \ rel_interior S" using openin_rel_interior [of S] by (simp add: openin_euclidean_subtopology_iff) define d where "d = Inf ?D" obtain \ where "0 < \" and \: "\\. \0 \ \; \ < \\ \ (a + \ *\<^sub>R l) \ rel_interior S" proof - obtain e where "e>0" and e: "\x'. x' \ affine hull S \ dist x' a < e \ x' \ rel_interior S" using relin_Ex a by blast show thesis proof (rule_tac \ = "e / norm l" in that) show "0 < e / norm l" by (simp add: \0 < e\ \l \ 0\) next show "a + \ *\<^sub>R l \ rel_interior S" if "0 \ \" "\ < e / norm l" for \ proof (rule e) show "a + \ *\<^sub>R l \ affine hull S" by (metis (no_types) add_diff_cancel_left' aff affine_affine_hull mem_affine_3_minus aaff) show "dist (a + \ *\<^sub>R l) a < e" using that by (simp add: \l \ 0\ dist_norm pos_less_divide_eq) qed qed qed have inint: "\e. \0 \ e; e < d\ \ a + e *\<^sub>R l \ rel_interior S" unfolding d_def using cInf_lower [OF _ bdd] by (metis (no_types, lifting) a add.right_neutral le_less mem_Collect_eq not_less real_vector.scale_zero_left) have "\ \ d" unfolding d_def using \ dual_order.strict_implies_order le_less_linear by (blast intro: cInf_greatest [OF nonMT]) with \0 < \\ have "0 < d" by simp have "a + d *\<^sub>R l \ rel_interior S" proof assume adl: "a + d *\<^sub>R l \ rel_interior S" obtain e where "e > 0" and e: "\x'. x' \ affine hull S \ dist x' (a + d *\<^sub>R l) < e \ x' \ rel_interior S" using relin_Ex adl by blast have "d + e / norm l \ Inf {d. 0 < d \ a + d *\<^sub>R l \ rel_interior S}" proof (rule cInf_greatest [OF nonMT], clarsimp) fix x::real assume "0 < x" and nonrel: "a + x *\<^sub>R l \ rel_interior S" show "d + e / norm l \ x" proof (cases "x < d") case True with inint nonrel \0 < x\ show ?thesis by auto next case False then have dle: "x < d + e / norm l \ dist (a + x *\<^sub>R l) (a + d *\<^sub>R l) < e" by (simp add: field_simps \l \ 0\) have ain: "a + x *\<^sub>R l \ affine hull S" by (metis add_diff_cancel_left' aff affine_affine_hull mem_affine_3_minus aaff) show ?thesis using e [OF ain] nonrel dle by force qed qed then show False using \0 < e\ \l \ 0\ by (simp add: d_def [symmetric] field_simps) qed moreover have "a + d *\<^sub>R l \ closure S" proof (clarsimp simp: closure_approachable) fix \::real assume "0 < \" have 1: "a + (d - min d (\ / 2 / norm l)) *\<^sub>R l \ S" proof (rule subsetD [OF rel_interior_subset inint]) show "d - min d (\ / 2 / norm l) < d" using \l \ 0\ \0 < d\ \0 < \\ by auto qed auto have "norm l * min d (\ / (norm l * 2)) \ norm l * (\ / (norm l * 2))" by (metis min_def mult_left_mono norm_ge_zero order_refl) also have "... < \" using \l \ 0\ \0 < \\ by (simp add: field_simps) finally have 2: "norm l * min d (\ / (norm l * 2)) < \" . show "\y\S. dist y (a + d *\<^sub>R l) < \" using 1 2 \0 < d\ \0 < \\ by (rule_tac x="a + (d - min d (\ / 2 / norm l)) *\<^sub>R l" in bexI) (auto simp: algebra_simps) qed ultimately have infront: "a + d *\<^sub>R l \ rel_frontier S" by (simp add: rel_frontier_def) show ?thesis by (rule that [OF \0 < d\ infront inint]) qed corollary ray_to_frontier: fixes a :: "'a::euclidean_space" assumes "bounded S" and a: "a \ interior S" and "l \ 0" obtains d where "0 < d" "(a + d *\<^sub>R l) \ frontier S" "\e. \0 \ e; e < d\ \ (a + e *\<^sub>R l) \ interior S" proof - have \
: "interior S = rel_interior S" using a rel_interior_nonempty_interior by auto then have "a \ rel_interior S" using a by simp moreover have "a + l \ affine hull S" using a affine_hull_nonempty_interior by blast ultimately show thesis by (metis \
\bounded S\ \l \ 0\ frontier_def ray_to_rel_frontier rel_frontier_def that) qed lemma segment_to_rel_frontier_aux: fixes x :: "'a::euclidean_space" assumes "convex S" "bounded S" and x: "x \ rel_interior S" and y: "y \ S" and xy: "x \ y" obtains z where "z \ rel_frontier S" "y \ closed_segment x z" "open_segment x z \ rel_interior S" proof - have "x + (y - x) \ affine hull S" using hull_inc [OF y] by auto then obtain d where "0 < d" and df: "(x + d *\<^sub>R (y-x)) \ rel_frontier S" and di: "\e. \0 \ e; e < d\ \ (x + e *\<^sub>R (y-x)) \ rel_interior S" by (rule ray_to_rel_frontier [OF \bounded S\ x]) (use xy in auto) show ?thesis proof show "x + d *\<^sub>R (y - x) \ rel_frontier S" by (simp add: df) next have "open_segment x y \ rel_interior S" using rel_interior_closure_convex_segment [OF \convex S\ x] closure_subset y by blast moreover have "x + d *\<^sub>R (y - x) \ open_segment x y" if "d < 1" using xy \0 < d\ that by (force simp: in_segment algebra_simps) ultimately have "1 \ d" using df rel_frontier_def by fastforce moreover have "x = (1 / d) *\<^sub>R x + ((d - 1) / d) *\<^sub>R x" by (metis \0 < d\ add.commute add_divide_distrib diff_add_cancel divide_self_if less_irrefl scaleR_add_left scaleR_one) ultimately show "y \ closed_segment x (x + d *\<^sub>R (y - x))" unfolding in_segment by (rule_tac x="1/d" in exI) (auto simp: algebra_simps) next show "open_segment x (x + d *\<^sub>R (y - x)) \ rel_interior S" proof (rule rel_interior_closure_convex_segment [OF \convex S\ x]) show "x + d *\<^sub>R (y - x) \ closure S" using df rel_frontier_def by auto qed qed qed lemma segment_to_rel_frontier: fixes x :: "'a::euclidean_space" assumes S: "convex S" "bounded S" and x: "x \ rel_interior S" and y: "y \ S" and xy: "\(x = y \ S = {x})" obtains z where "z \ rel_frontier S" "y \ closed_segment x z" "open_segment x z \ rel_interior S" proof (cases "x=y") case True with xy have "S \ {x}" by blast with True show ?thesis by (metis Set.set_insert all_not_in_conv ends_in_segment(1) insert_iff segment_to_rel_frontier_aux[OF S x] that y) next case False then show ?thesis using segment_to_rel_frontier_aux [OF S x y] that by blast qed proposition rel_frontier_not_sing: fixes a :: "'a::euclidean_space" assumes "bounded S" shows "rel_frontier S \ {a}" proof (cases "S = {}") case True then show ?thesis by simp next case False then obtain z where "z \ S" by blast then show ?thesis proof (cases "S = {z}") case True then show ?thesis by simp next case False then obtain w where "w \ S" "w \ z" using \z \ S\ by blast show ?thesis proof assume "rel_frontier S = {a}" then consider "w \ rel_frontier S" | "z \ rel_frontier S" using \w \ z\ by auto then show False proof cases case 1 then have w: "w \ rel_interior S" using \w \ S\ closure_subset rel_frontier_def by fastforce have "w + (w - z) \ affine hull S" by (metis \w \ S\ \z \ S\ affine_affine_hull hull_inc mem_affine_3_minus scaleR_one) then obtain e where "0 < e" "(w + e *\<^sub>R (w - z)) \ rel_frontier S" using \w \ z\ \z \ S\ by (metis assms ray_to_rel_frontier right_minus_eq w) moreover obtain d where "0 < d" "(w + d *\<^sub>R (z - w)) \ rel_frontier S" using ray_to_rel_frontier [OF \bounded S\ w, of "1 *\<^sub>R (z - w)"] \w \ z\ \z \ S\ by (metis add.commute add.right_neutral diff_add_cancel hull_inc scaleR_one) ultimately have "d *\<^sub>R (z - w) = e *\<^sub>R (w - z)" using \rel_frontier S = {a}\ by force moreover have "e \ -d " using \0 < e\ \0 < d\ by force ultimately show False by (metis (no_types, lifting) \w \ z\ eq_iff_diff_eq_0 minus_diff_eq real_vector.scale_cancel_right real_vector.scale_minus_right scaleR_left.minus) next case 2 then have z: "z \ rel_interior S" using \z \ S\ closure_subset rel_frontier_def by fastforce have "z + (z - w) \ affine hull S" by (metis \z \ S\ \w \ S\ affine_affine_hull hull_inc mem_affine_3_minus scaleR_one) then obtain e where "0 < e" "(z + e *\<^sub>R (z - w)) \ rel_frontier S" using \w \ z\ \w \ S\ by (metis assms ray_to_rel_frontier right_minus_eq z) moreover obtain d where "0 < d" "(z + d *\<^sub>R (w - z)) \ rel_frontier S" using ray_to_rel_frontier [OF \bounded S\ z, of "1 *\<^sub>R (w - z)"] \w \ z\ \w \ S\ by (metis add.commute add.right_neutral diff_add_cancel hull_inc scaleR_one) ultimately have "d *\<^sub>R (w - z) = e *\<^sub>R (z - w)" using \rel_frontier S = {a}\ by force moreover have "e \ -d " using \0 < e\ \0 < d\ by force ultimately show False by (metis (no_types, lifting) \w \ z\ eq_iff_diff_eq_0 minus_diff_eq real_vector.scale_cancel_right real_vector.scale_minus_right scaleR_left.minus) qed qed qed qed subsection\<^marker>\tag unimportant\ \Convexity on direct sums\ lemma closure_sum: fixes S T :: "'a::real_normed_vector set" shows "closure S + closure T \ closure (S + T)" unfolding set_plus_image closure_Times [symmetric] split_def by (intro closure_bounded_linear_image_subset bounded_linear_add bounded_linear_fst bounded_linear_snd) lemma rel_interior_sum: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "convex T" shows "rel_interior (S + T) = rel_interior S + rel_interior T" proof - have "rel_interior S + rel_interior T = (\(x,y). x + y) ` (rel_interior S \ rel_interior T)" by (simp add: set_plus_image) also have "\ = (\(x,y). x + y) ` rel_interior (S \ T)" using rel_interior_Times assms by auto also have "\ = rel_interior (S + T)" using fst_snd_linear convex_Times assms rel_interior_convex_linear_image[of "(\(x,y). x + y)" "S \ T"] by (auto simp add: set_plus_image) finally show ?thesis .. qed lemma rel_interior_sum_gen: fixes S :: "'a \ 'n::euclidean_space set" assumes "\i. i\I \ convex (S i)" shows "rel_interior (sum S I) = sum (\i. rel_interior (S i)) I" using rel_interior_sum rel_interior_sing[of "0"] assms by (subst sum_set_cond_linear[of convex], auto simp add: convex_set_plus) lemma convex_rel_open_direct_sum: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "rel_open S" and "convex T" and "rel_open T" shows "convex (S \ T) \ rel_open (S \ T)" by (metis assms convex_Times rel_interior_Times rel_open_def) lemma convex_rel_open_sum: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "rel_open S" and "convex T" and "rel_open T" shows "convex (S + T) \ rel_open (S + T)" by (metis assms convex_set_plus rel_interior_sum rel_open_def) lemma convex_hull_finite_union_cones: assumes "finite I" and "I \ {}" assumes "\i. i\I \ convex (S i) \ cone (S i) \ S i \ {}" shows "convex hull (\(S ` I)) = sum S I" (is "?lhs = ?rhs") proof - { fix x assume "x \ ?lhs" then obtain c xs where x: "x = sum (\i. c i *\<^sub>R xs i) I \ (\i\I. c i \ 0) \ sum c I = 1 \ (\i\I. xs i \ S i)" using convex_hull_finite_union[of I S] assms by auto define s where "s i = c i *\<^sub>R xs i" for i have "\i\I. s i \ S i" using s_def x assms by (simp add: mem_cone) moreover have "x = sum s I" using x s_def by auto ultimately have "x \ ?rhs" using set_sum_alt[of I S] assms by auto } moreover { fix x assume "x \ ?rhs" then obtain s where x: "x = sum s I \ (\i\I. s i \ S i)" using set_sum_alt[of I S] assms by auto define xs where "xs i = of_nat(card I) *\<^sub>R s i" for i then have "x = sum (\i. ((1 :: real) / of_nat(card I)) *\<^sub>R xs i) I" using x assms by auto moreover have "\i\I. xs i \ S i" using x xs_def assms by (simp add: cone_def) moreover have "\i\I. (1 :: real) / of_nat (card I) \ 0" by auto moreover have "sum (\i. (1 :: real) / of_nat (card I)) I = 1" using assms by auto ultimately have "x \ ?lhs" using assms apply (simp add: convex_hull_finite_union[of I S]) by (rule_tac x = "(\i. 1 / (card I))" in exI) auto } ultimately show ?thesis by auto qed lemma convex_hull_union_cones_two: fixes S T :: "'m::euclidean_space set" assumes "convex S" and "cone S" and "S \ {}" assumes "convex T" and "cone T" and "T \ {}" shows "convex hull (S \ T) = S + T" proof - define I :: "nat set" where "I = {1, 2}" define A where "A i = (if i = (1::nat) then S else T)" for i have "\(A ` I) = S \ T" using A_def I_def by auto then have "convex hull (\(A ` I)) = convex hull (S \ T)" by auto moreover have "convex hull \(A ` I) = sum A I" using A_def I_def by (metis assms convex_hull_finite_union_cones empty_iff finite.emptyI finite.insertI insertI1) moreover have "sum A I = S + T" using A_def I_def by (force simp add: set_plus_def) ultimately show ?thesis by auto qed lemma rel_interior_convex_hull_union: fixes S :: "'a \ 'n::euclidean_space set" assumes "finite I" and "\i\I. convex (S i) \ S i \ {}" shows "rel_interior (convex hull (\(S ` I))) = {sum (\i. c i *\<^sub>R s i) I | c s. (\i\I. c i > 0) \ sum c I = 1 \ (\i\I. s i \ rel_interior(S i))}" (is "?lhs = ?rhs") proof (cases "I = {}") case True then show ?thesis using convex_hull_empty by auto next case False define C0 where "C0 = convex hull (\(S ` I))" have "\i\I. C0 \ S i" unfolding C0_def using hull_subset[of "\(S ` I)"] by auto define K0 where "K0 = cone hull ({1 :: real} \ C0)" define K where "K i = cone hull ({1 :: real} \ S i)" for i have "\i\I. K i \ {}" unfolding K_def using assms by (simp add: cone_hull_empty_iff[symmetric]) have convK: "\i\I. convex (K i)" unfolding K_def by (simp add: assms(2) convex_Times convex_cone_hull) have "K0 \ K i" if "i \ I" for i unfolding K0_def K_def by (simp add: Sigma_mono \\i\I. S i \ C0\ hull_mono that) then have "K0 \ \(K ` I)" by auto moreover have "convex K0" unfolding K0_def by (simp add: C0_def convex_Times convex_cone_hull) ultimately have geq: "K0 \ convex hull (\(K ` I))" using hull_minimal[of _ "K0" "convex"] by blast have "\i\I. K i \ {1 :: real} \ S i" using K_def by (simp add: hull_subset) then have "\(K ` I) \ {1 :: real} \ \(S ` I)" by auto then have "convex hull \(K ` I) \ convex hull ({1 :: real} \ \(S ` I))" by (simp add: hull_mono) then have "convex hull \(K ` I) \ {1 :: real} \ C0" unfolding C0_def using convex_hull_Times[of "{(1 :: real)}" "\(S ` I)"] convex_hull_singleton by auto moreover have "cone (convex hull (\(K ` I)))" by (simp add: K_def cone_Union cone_cone_hull cone_convex_hull) ultimately have "convex hull (\(K ` I)) \ K0" unfolding K0_def using hull_minimal[of _ "convex hull (\(K ` I))" "cone"] by blast then have "K0 = convex hull (\(K ` I))" using geq by auto also have "\ = sum K I" using assms False \\i\I. K i \ {}\ cone_hull_eq convK by (intro convex_hull_finite_union_cones; fastforce simp: K_def) finally have "K0 = sum K I" by auto then have *: "rel_interior K0 = sum (\i. (rel_interior (K i))) I" using rel_interior_sum_gen[of I K] convK by auto { fix x assume "x \ ?lhs" then have "(1::real, x) \ rel_interior K0" using K0_def C0_def rel_interior_convex_cone_aux[of C0 "1::real" x] convex_convex_hull by auto then obtain k where k: "(1::real, x) = sum k I \ (\i\I. k i \ rel_interior (K i))" using \finite I\ * set_sum_alt[of I "\i. rel_interior (K i)"] by auto { fix i assume "i \ I" then have "convex (S i) \ k i \ rel_interior (cone hull {1} \ S i)" using k K_def assms by auto then have "\ci si. k i = (ci, ci *\<^sub>R si) \ 0 < ci \ si \ rel_interior (S i)" using rel_interior_convex_cone[of "S i"] by auto } then obtain c s where cs: "\i\I. k i = (c i, c i *\<^sub>R s i) \ 0 < c i \ s i \ rel_interior (S i)" by metis then have "x = (\i\I. c i *\<^sub>R s i) \ sum c I = 1" using k by (simp add: sum_prod) then have "x \ ?rhs" using k cs by auto } moreover { fix x assume "x \ ?rhs" then obtain c s where cs: "x = sum (\i. c i *\<^sub>R s i) I \ (\i\I. c i > 0) \ sum c I = 1 \ (\i\I. s i \ rel_interior (S i))" by auto define k where "k i = (c i, c i *\<^sub>R s i)" for i { fix i assume "i \ I" then have "k i \ rel_interior (K i)" using k_def K_def assms cs rel_interior_convex_cone[of "S i"] by auto } then have "(1, x) \ rel_interior K0" using * set_sum_alt[of I "(\i. rel_interior (K i))"] assms cs by (simp add: k_def) (metis (mono_tags, lifting) sum_prod) then have "x \ ?lhs" using K0_def C0_def rel_interior_convex_cone_aux[of C0 1 x] by auto } ultimately show ?thesis by blast qed lemma convex_le_Inf_differential: fixes f :: "real \ real" assumes "convex_on I f" and "x \ interior I" and "y \ I" shows "f y \ f x + Inf ((\t. (f x - f t) / (x - t)) ` ({x<..} \ I)) * (y - x)" (is "_ \ _ + Inf (?F x) * (y - x)") proof (cases rule: linorder_cases) assume "x < y" moreover have "open (interior I)" by auto from openE[OF this \x \ interior I\] obtain e where e: "0 < e" "ball x e \ interior I" . moreover define t where "t = min (x + e / 2) ((x + y) / 2)" ultimately have "x < t" "t < y" "t \ ball x e" by (auto simp: dist_real_def field_simps split: split_min) with \x \ interior I\ e interior_subset[of I] have "t \ I" "x \ I" by auto define K where "K = x - e / 2" with \0 < e\ have "K \ ball x e" "K < x" by (auto simp: dist_real_def) then have "K \ I" using \interior I \ I\ e(2) by blast have "Inf (?F x) \ (f x - f y) / (x - y)" proof (intro bdd_belowI cInf_lower2) show "(f x - f t) / (x - t) \ ?F x" using \t \ I\ \x < t\ by auto show "(f x - f t) / (x - t) \ (f x - f y) / (x - y)" using \convex_on I f\ \x \ I\ \y \ I\ \x < t\ \t < y\ by (rule convex_on_diff) next fix y assume "y \ ?F x" with order_trans[OF convex_on_diff[OF \convex_on I f\ \K \ I\ _ \K < x\ _]] show "(f K - f x) / (K - x) \ y" by auto qed then show ?thesis using \x < y\ by (simp add: field_simps) next assume "y < x" moreover have "open (interior I)" by auto from openE[OF this \x \ interior I\] obtain e where e: "0 < e" "ball x e \ interior I" . moreover define t where "t = x + e / 2" ultimately have "x < t" "t \ ball x e" by (auto simp: dist_real_def field_simps) with \x \ interior I\ e interior_subset[of I] have "t \ I" "x \ I" by auto have "(f x - f y) / (x - y) \ Inf (?F x)" proof (rule cInf_greatest) have "(f x - f y) / (x - y) = (f y - f x) / (y - x)" using \y < x\ by (auto simp: field_simps) also fix z assume "z \ ?F x" with order_trans[OF convex_on_diff[OF \convex_on I f\ \y \ I\ _ \y < x\]] have "(f y - f x) / (y - x) \ z" by auto finally show "(f x - f y) / (x - y) \ z" . next have "x + e / 2 \ ball x e" using e by (auto simp: dist_real_def) with e interior_subset[of I] have "x + e / 2 \ {x<..} \ I" by auto then show "?F x \ {}" by blast qed then show ?thesis using \y < x\ by (simp add: field_simps) qed simp subsection\<^marker>\tag unimportant\\Explicit formulas for interior and relative interior of convex hull\ lemma at_within_cbox_finite: assumes "x \ box a b" "x \ S" "finite S" shows "(at x within cbox a b - S) = at x" proof - have "interior (cbox a b - S) = box a b - S" using \finite S\ by (simp add: interior_diff finite_imp_closed) then show ?thesis using at_within_interior assms by fastforce qed lemma affine_independent_convex_affine_hull: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" "T \ S" shows "convex hull T = affine hull T \ convex hull S" proof - have fin: "finite S" "finite T" using assms aff_independent_finite finite_subset by auto have "convex hull T \ affine hull T" using convex_hull_subset_affine_hull by blast moreover have "convex hull T \ convex hull S" using assms hull_mono by blast moreover have "affine hull T \ convex hull S \ convex hull T" proof - have 0: "\u. sum u S = 0 \ (\v\S. u v = 0) \ (\v\S. u v *\<^sub>R v) \ 0" using affine_dependent_explicit_finite assms(1) fin(1) by auto show ?thesis proof (clarsimp simp add: affine_hull_finite fin) fix u assume S: "(\v\T. u v *\<^sub>R v) \ convex hull S" and T1: "sum u T = 1" then obtain v where v: "\x\S. 0 \ v x" "sum v S = 1" "(\x\S. v x *\<^sub>R x) = (\v\T. u v *\<^sub>R v)" by (auto simp add: convex_hull_finite fin) { fix x assume"x \ T" then have S: "S = (S - T) \ T" \ \split into separate cases\ using assms by auto have [simp]: "(\x\T. v x *\<^sub>R x) + (\x\S - T. v x *\<^sub>R x) = (\x\T. u x *\<^sub>R x)" "sum v T + sum v (S - T) = 1" using v fin S by (auto simp: sum.union_disjoint [symmetric] Un_commute) have "(\x\S. if x \ T then v x - u x else v x) = 0" "(\x\S. (if x \ T then v x - u x else v x) *\<^sub>R x) = 0" using v fin T1 by (subst S, subst sum.union_disjoint, auto simp: algebra_simps sum_subtractf)+ } note [simp] = this have "(\x\T. 0 \ u x)" using 0 [of "\x. if x \ T then v x - u x else v x"] \T \ S\ v(1) by fastforce then show "(\v\T. u v *\<^sub>R v) \ convex hull T" using 0 [of "\x. if x \ T then v x - u x else v x"] \T \ S\ T1 by (fastforce simp add: convex_hull_finite fin) qed qed ultimately show ?thesis by blast qed lemma affine_independent_span_eq: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" "card S = Suc (DIM ('a))" shows "affine hull S = UNIV" proof (cases "S = {}") case True then show ?thesis using assms by simp next case False then obtain a T where T: "a \ T" "S = insert a T" by blast then have fin: "finite T" using assms by (metis finite_insert aff_independent_finite) have "UNIV \ (+) a ` span ((\x. x - a) ` T)" proof (intro card_ge_dim_independent Fun.vimage_subsetD) show "independent ((\x. x - a) ` T)" using T affine_dependent_iff_dependent assms(1) by auto show "dim ((+) a -` UNIV) \ card ((\x. x - a) ` T)" using assms T fin by (auto simp: card_image inj_on_def) qed (use surj_plus in auto) then show ?thesis using T(2) affine_hull_insert_span_gen equalityI by fastforce qed lemma affine_independent_span_gt: fixes S :: "'a::euclidean_space set" assumes ind: "\ affine_dependent S" and dim: "DIM ('a) < card S" shows "affine hull S = UNIV" proof (intro affine_independent_span_eq [OF ind] antisym) show "card S \ Suc DIM('a)" using aff_independent_finite affine_dependent_biggerset ind by fastforce show "Suc DIM('a) \ card S" using Suc_leI dim by blast qed lemma empty_interior_affine_hull: fixes S :: "'a::euclidean_space set" assumes "finite S" and dim: "card S \ DIM ('a)" shows "interior(affine hull S) = {}" using assms proof (induct S rule: finite_induct) case (insert x S) then have "dim (span ((\y. y - x) ` S)) < DIM('a)" by (auto simp: Suc_le_lessD card_image_le dual_order.trans intro!: dim_le_card'[THEN le_less_trans]) then show ?case by (simp add: empty_interior_lowdim affine_hull_insert_span_gen interior_translation) qed auto lemma empty_interior_convex_hull: fixes S :: "'a::euclidean_space set" assumes "finite S" and dim: "card S \ DIM ('a)" shows "interior(convex hull S) = {}" by (metis Diff_empty Diff_eq_empty_iff convex_hull_subset_affine_hull interior_mono empty_interior_affine_hull [OF assms]) lemma explicit_subset_rel_interior_convex_hull: fixes S :: "'a::euclidean_space set" shows "finite S \ {y. \u. (\x \ S. 0 < u x \ u x < 1) \ sum u S = 1 \ sum (\x. u x *\<^sub>R x) S = y} \ rel_interior (convex hull S)" by (force simp add: rel_interior_convex_hull_union [where S="\x. {x}" and I=S, simplified]) lemma explicit_subset_rel_interior_convex_hull_minimal: fixes S :: "'a::euclidean_space set" shows "finite S \ {y. \u. (\x \ S. 0 < u x) \ sum u S = 1 \ sum (\x. u x *\<^sub>R x) S = y} \ rel_interior (convex hull S)" by (force simp add: rel_interior_convex_hull_union [where S="\x. {x}" and I=S, simplified]) lemma rel_interior_convex_hull_explicit: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "rel_interior(convex hull S) = {y. \u. (\x \ S. 0 < u x) \ sum u S = 1 \ sum (\x. u x *\<^sub>R x) S = y}" (is "?lhs = ?rhs") proof show "?rhs \ ?lhs" by (simp add: aff_independent_finite explicit_subset_rel_interior_convex_hull_minimal assms) next show "?lhs \ ?rhs" proof (cases "\a. S = {a}") case True then show "?lhs \ ?rhs" by force next case False have fs: "finite S" using assms by (simp add: aff_independent_finite) { fix a b and d::real assume ab: "a \ S" "b \ S" "a \ b" then have S: "S = (S - {a,b}) \ {a,b}" \ \split into separate cases\ by auto have "(\x\S. if x = a then - d else if x = b then d else 0) = 0" "(\x\S. (if x = a then - d else if x = b then d else 0) *\<^sub>R x) = d *\<^sub>R b - d *\<^sub>R a" using ab fs by (subst S, subst sum.union_disjoint, auto)+ } note [simp] = this { fix y assume y: "y \ convex hull S" "y \ ?rhs" have *: False if ua: "\x\S. 0 \ u x" "sum u S = 1" "\ 0 < u a" "a \ S" and yT: "y = (\x\S. u x *\<^sub>R x)" "y \ T" "open T" and sb: "T \ affine hull S \ {w. \u. (\x\S. 0 \ u x) \ sum u S = 1 \ (\x\S. u x *\<^sub>R x) = w}" for u T a proof - have ua0: "u a = 0" using ua by auto obtain b where b: "b\S" "a \ b" using ua False by auto obtain e where e: "0 < e" "ball (\x\S. u x *\<^sub>R x) e \ T" using yT by (auto elim: openE) with b obtain d where d: "0 < d" "norm(d *\<^sub>R (a-b)) < e" by (auto intro: that [of "e / 2 / norm(a-b)"]) have "(\x\S. u x *\<^sub>R x) \ affine hull S" using yT y by (metis affine_hull_convex_hull hull_redundant_eq) then have "(\x\S. u x *\<^sub>R x) - d *\<^sub>R (a - b) \ affine hull S" using ua b by (auto simp: hull_inc intro: mem_affine_3_minus2) then have "y - d *\<^sub>R (a - b) \ T \ affine hull S" using d e yT by auto then obtain v where v: "\x\S. 0 \ v x" "sum v S = 1" "(\x\S. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x) - d *\<^sub>R (a - b)" using subsetD [OF sb] yT by auto have aff: "\u. sum u S = 0 \ (\v\S. u v = 0) \ (\v\S. u v *\<^sub>R v) \ 0" using assms by (simp add: affine_dependent_explicit_finite fs) show False using ua b d v aff [of "\x. (v x - u x) - (if x = a then -d else if x = b then d else 0)"] by (auto simp: algebra_simps sum_subtractf sum.distrib) qed have "y \ rel_interior (convex hull S)" using y apply (simp add: mem_rel_interior) apply (auto simp: convex_hull_finite [OF fs]) apply (drule_tac x=u in spec) apply (auto intro: *) done } with rel_interior_subset show "?lhs \ ?rhs" by blast qed qed lemma interior_convex_hull_explicit_minimal: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "interior(convex hull S) = (if card(S) \ DIM('a) then {} else {y. \u. (\x \ S. 0 < u x) \ sum u S = 1 \ (\x\S. u x *\<^sub>R x) = y})" (is "_ = (if _ then _ else ?rhs)") proof (clarsimp simp: aff_independent_finite empty_interior_convex_hull assms) assume S: "\ card S \ DIM('a)" have "interior (convex hull S) = rel_interior(convex hull S)" using assms S by (simp add: affine_independent_span_gt rel_interior_interior) then show "interior(convex hull S) = ?rhs" by (simp add: assms S rel_interior_convex_hull_explicit) qed lemma interior_convex_hull_explicit: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "interior(convex hull S) = (if card(S) \ DIM('a) then {} else {y. \u. (\x \ S. 0 < u x \ u x < 1) \ sum u S = 1 \ (\x\S. u x *\<^sub>R x) = y})" proof - { fix u :: "'a \ real" and a assume "card Basis < card S" and u: "\x. x\S \ 0 < u x" "sum u S = 1" and a: "a \ S" then have cs: "Suc 0 < card S" by (metis DIM_positive less_trans_Suc) obtain b where b: "b \ S" "a \ b" proof (cases "S \ {a}") case True then show thesis using cs subset_singletonD by fastforce qed blast have "u a + u b \ sum u {a,b}" using a b by simp also have "... \ sum u S" using a b u by (intro Groups_Big.sum_mono2) (auto simp: less_imp_le aff_independent_finite assms) finally have "u a < 1" using \b \ S\ u by fastforce } note [simp] = this show ?thesis using assms by (force simp add: not_le interior_convex_hull_explicit_minimal) qed lemma interior_closed_segment_ge2: fixes a :: "'a::euclidean_space" assumes "2 \ DIM('a)" shows "interior(closed_segment a b) = {}" using assms unfolding segment_convex_hull proof - have "card {a, b} \ DIM('a)" using assms by (simp add: card_insert_if linear not_less_eq_eq numeral_2_eq_2) then show "interior (convex hull {a, b}) = {}" by (metis empty_interior_convex_hull finite.insertI finite.emptyI) qed lemma interior_open_segment: fixes a :: "'a::euclidean_space" shows "interior(open_segment a b) = (if 2 \ DIM('a) then {} else open_segment a b)" proof (simp add: not_le, intro conjI impI) assume "2 \ DIM('a)" then show "interior (open_segment a b) = {}" using interior_closed_segment_ge2 interior_mono segment_open_subset_closed by blast next assume le2: "DIM('a) < 2" show "interior (open_segment a b) = open_segment a b" proof (cases "a = b") case True then show ?thesis by auto next case False with le2 have "affine hull (open_segment a b) = UNIV" by (simp add: False affine_independent_span_gt) then show "interior (open_segment a b) = open_segment a b" using rel_interior_interior rel_interior_open_segment by blast qed qed lemma interior_closed_segment: fixes a :: "'a::euclidean_space" shows "interior(closed_segment a b) = (if 2 \ DIM('a) then {} else open_segment a b)" proof (cases "a = b") case True then show ?thesis by simp next case False then have "closure (open_segment a b) = closed_segment a b" by simp then show ?thesis by (metis (no_types) convex_interior_closure convex_open_segment interior_open_segment) qed lemmas interior_segment = interior_closed_segment interior_open_segment lemma closed_segment_eq [simp]: fixes a :: "'a::euclidean_space" shows "closed_segment a b = closed_segment c d \ {a,b} = {c,d}" proof assume abcd: "closed_segment a b = closed_segment c d" show "{a,b} = {c,d}" proof (cases "a=b \ c=d") case True with abcd show ?thesis by force next case False then have neq: "a \ b \ c \ d" by force have *: "closed_segment c d - {a, b} = rel_interior (closed_segment c d)" using neq abcd by (metis (no_types) open_segment_def rel_interior_closed_segment) have "b \ {c, d}" proof - have "insert b (closed_segment c d) = closed_segment c d" using abcd by blast then show ?thesis by (metis DiffD2 Diff_insert2 False * insertI1 insert_Diff_if open_segment_def rel_interior_closed_segment) qed moreover have "a \ {c, d}" by (metis Diff_iff False * abcd ends_in_segment(1) insertI1 open_segment_def rel_interior_closed_segment) ultimately show "{a, b} = {c, d}" using neq by fastforce qed next assume "{a,b} = {c,d}" then show "closed_segment a b = closed_segment c d" by (simp add: segment_convex_hull) qed lemma closed_open_segment_eq [simp]: fixes a :: "'a::euclidean_space" shows "closed_segment a b \ open_segment c d" by (metis DiffE closed_segment_neq_empty closure_closed_segment closure_open_segment ends_in_segment(1) insertI1 open_segment_def) lemma open_closed_segment_eq [simp]: fixes a :: "'a::euclidean_space" shows "open_segment a b \ closed_segment c d" using closed_open_segment_eq by blast lemma open_segment_eq [simp]: fixes a :: "'a::euclidean_space" shows "open_segment a b = open_segment c d \ a = b \ c = d \ {a,b} = {c,d}" (is "?lhs = ?rhs") proof assume abcd: ?lhs show ?rhs proof (cases "a=b \ c=d") case True with abcd show ?thesis using finite_open_segment by fastforce next case False then have a2: "a \ b \ c \ d" by force with abcd show ?rhs unfolding open_segment_def by (metis (no_types) abcd closed_segment_eq closure_open_segment) qed next assume ?rhs then show ?lhs by (metis Diff_cancel convex_hull_singleton insert_absorb2 open_segment_def segment_convex_hull) qed subsection\<^marker>\tag unimportant\\Similar results for closure and (relative or absolute) frontier\ lemma closure_convex_hull [simp]: fixes S :: "'a::euclidean_space set" shows "compact S ==> closure(convex hull S) = convex hull S" by (simp add: compact_imp_closed compact_convex_hull) lemma rel_frontier_convex_hull_explicit: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "rel_frontier(convex hull S) = {y. \u. (\x \ S. 0 \ u x) \ (\x \ S. u x = 0) \ sum u S = 1 \ sum (\x. u x *\<^sub>R x) S = y}" proof - have fs: "finite S" using assms by (simp add: aff_independent_finite) have "\u y v. \y \ S; u y = 0; sum u S = 1; \x\S. 0 < v x; sum v S = 1; (\x\S. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x)\ \ \u. sum u S = 0 \ (\v\S. u v \ 0) \ (\v\S. u v *\<^sub>R v) = 0" apply (rule_tac x = "\x. u x - v x" in exI) apply (force simp: sum_subtractf scaleR_diff_left) done then show ?thesis using fs assms apply (simp add: rel_frontier_def finite_imp_compact rel_interior_convex_hull_explicit) apply (auto simp: convex_hull_finite) apply (metis less_eq_real_def) by (simp add: affine_dependent_explicit_finite) qed lemma frontier_convex_hull_explicit: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "frontier(convex hull S) = {y. \u. (\x \ S. 0 \ u x) \ (DIM ('a) < card S \ (\x \ S. u x = 0)) \ sum u S = 1 \ sum (\x. u x *\<^sub>R x) S = y}" proof - have fs: "finite S" using assms by (simp add: aff_independent_finite) show ?thesis proof (cases "DIM ('a) < card S") case True with assms fs show ?thesis by (simp add: rel_frontier_def frontier_def rel_frontier_convex_hull_explicit [symmetric] interior_convex_hull_explicit_minimal rel_interior_convex_hull_explicit) next case False then have "card S \ DIM ('a)" by linarith then show ?thesis using assms fs apply (simp add: frontier_def interior_convex_hull_explicit finite_imp_compact) apply (simp add: convex_hull_finite) done qed qed lemma rel_frontier_convex_hull_cases: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "rel_frontier(convex hull S) = \{convex hull (S - {x}) |x. x \ S}" proof - have fs: "finite S" using assms by (simp add: aff_independent_finite) { fix u a have "\x\S. 0 \ u x \ a \ S \ u a = 0 \ sum u S = 1 \ \x v. x \ S \ (\x\S - {x}. 0 \ v x) \ sum v (S - {x}) = 1 \ (\x\S - {x}. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x)" apply (rule_tac x=a in exI) apply (rule_tac x=u in exI) apply (simp add: Groups_Big.sum_diff1 fs) done } moreover { fix a u have "a \ S \ \x\S - {a}. 0 \ u x \ sum u (S - {a}) = 1 \ \v. (\x\S. 0 \ v x) \ (\x\S. v x = 0) \ sum v S = 1 \ (\x\S. v x *\<^sub>R x) = (\x\S - {a}. u x *\<^sub>R x)" apply (rule_tac x="\x. if x = a then 0 else u x" in exI) apply (auto simp: sum.If_cases Diff_eq if_smult fs) done } ultimately show ?thesis using assms apply (simp add: rel_frontier_convex_hull_explicit) apply (auto simp add: convex_hull_finite fs Union_SetCompr_eq) done qed lemma frontier_convex_hull_eq_rel_frontier: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "frontier(convex hull S) = (if card S \ DIM ('a) then convex hull S else rel_frontier(convex hull S))" using assms unfolding rel_frontier_def frontier_def by (simp add: affine_independent_span_gt rel_interior_interior finite_imp_compact empty_interior_convex_hull aff_independent_finite) lemma frontier_convex_hull_cases: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "frontier(convex hull S) = (if card S \ DIM ('a) then convex hull S else \{convex hull (S - {x}) |x. x \ S})" by (simp add: assms frontier_convex_hull_eq_rel_frontier rel_frontier_convex_hull_cases) lemma in_frontier_convex_hull: fixes S :: "'a::euclidean_space set" assumes "finite S" "card S \ Suc (DIM ('a))" "x \ S" shows "x \ frontier(convex hull S)" proof (cases "affine_dependent S") case True with assms obtain y where "y \ S" and y: "y \ affine hull (S - {y})" by (auto simp: affine_dependent_def) moreover have "x \ closure (convex hull S)" by (meson closure_subset hull_inc subset_eq \x \ S\) moreover have "x \ interior (convex hull S)" using assms by (metis Suc_mono affine_hull_convex_hull affine_hull_nonempty_interior \y \ S\ y card.remove empty_iff empty_interior_affine_hull finite_Diff hull_redundant insert_Diff interior_UNIV not_less) ultimately show ?thesis unfolding frontier_def by blast next case False { assume "card S = Suc (card Basis)" then have cs: "Suc 0 < card S" by (simp) with subset_singletonD have "\y \ S. y \ x" by (cases "S \ {x}") fastforce+ } note [dest!] = this show ?thesis using assms unfolding frontier_convex_hull_cases [OF False] Union_SetCompr_eq by (auto simp: le_Suc_eq hull_inc) qed lemma not_in_interior_convex_hull: fixes S :: "'a::euclidean_space set" assumes "finite S" "card S \ Suc (DIM ('a))" "x \ S" shows "x \ interior(convex hull S)" using in_frontier_convex_hull [OF assms] by (metis Diff_iff frontier_def) lemma interior_convex_hull_eq_empty: fixes S :: "'a::euclidean_space set" assumes "card S = Suc (DIM ('a))" shows "interior(convex hull S) = {} \ affine_dependent S" proof show "affine_dependent S \ interior (convex hull S) = {}" proof (clarsimp simp: affine_dependent_def) fix a b assume "b \ S" "b \ affine hull (S - {b})" then have "interior(affine hull S) = {}" using assms by (metis DIM_positive One_nat_def Suc_mono card.remove card.infinite empty_interior_affine_hull eq_iff hull_redundant insert_Diff not_less zero_le_one) then show "interior (convex hull S) = {}" using affine_hull_nonempty_interior by fastforce qed next show "interior (convex hull S) = {} \ affine_dependent S" by (metis affine_hull_convex_hull affine_hull_empty affine_independent_span_eq assms convex_convex_hull empty_not_UNIV rel_interior_eq_empty rel_interior_interior) qed subsection \Coplanarity, and collinearity in terms of affine hull\ definition\<^marker>\tag important\ coplanar where "coplanar S \ \u v w. S \ affine hull {u,v,w}" lemma collinear_affine_hull: "collinear S \ (\u v. S \ affine hull {u,v})" proof (cases "S={}") case True then show ?thesis by simp next case False then obtain x where x: "x \ S" by auto { fix u assume *: "\x y. \x\S; y\S\ \ \c. x - y = c *\<^sub>R u" have "\y c. x - y = c *\<^sub>R u \ \a b. y = a *\<^sub>R x + b *\<^sub>R (x + u) \ a + b = 1" by (rule_tac x="1+c" in exI, rule_tac x="-c" in exI, simp add: algebra_simps) then have "\u v. S \ {a *\<^sub>R u + b *\<^sub>R v |a b. a + b = 1}" using * [OF x] by (rule_tac x=x in exI, rule_tac x="x+u" in exI, force) } moreover { fix u v x y assume *: "S \ {a *\<^sub>R u + b *\<^sub>R v |a b. a + b = 1}" have "\c. x - y = c *\<^sub>R (v-u)" if "x\S" "y\S" proof - obtain a r where "a + r = 1" "x = a *\<^sub>R u + r *\<^sub>R v" using "*" \x \ S\ by blast moreover obtain b s where "b + s = 1" "y = b *\<^sub>R u + s *\<^sub>R v" using "*" \y \ S\ by blast ultimately have "x - y = (r-s) *\<^sub>R (v-u)" by (simp add: algebra_simps) (metis scaleR_left.add) then show ?thesis by blast qed } ultimately show ?thesis unfolding collinear_def affine_hull_2 by blast qed lemma collinear_closed_segment [simp]: "collinear (closed_segment a b)" by (metis affine_hull_convex_hull collinear_affine_hull hull_subset segment_convex_hull) lemma collinear_open_segment [simp]: "collinear (open_segment a b)" unfolding open_segment_def by (metis convex_hull_subset_affine_hull segment_convex_hull dual_order.trans convex_hull_subset_affine_hull Diff_subset collinear_affine_hull) lemma collinear_between_cases: fixes c :: "'a::euclidean_space" shows "collinear {a,b,c} \ between (b,c) a \ between (c,a) b \ between (a,b) c" (is "?lhs = ?rhs") proof assume ?lhs then obtain u v where uv: "\x. x \ {a, b, c} \ \c. x = u + c *\<^sub>R v" by (auto simp: collinear_alt) show ?rhs using uv [of a] uv [of b] uv [of c] by (auto simp: between_1) next assume ?rhs then show ?lhs unfolding between_mem_convex_hull by (metis (no_types, opaque_lifting) collinear_closed_segment collinear_subset hull_redundant hull_subset insert_commute segment_convex_hull) qed lemma subset_continuous_image_segment_1: fixes f :: "'a::euclidean_space \ real" assumes "continuous_on (closed_segment a b) f" shows "closed_segment (f a) (f b) \ image f (closed_segment a b)" by (metis connected_segment convex_contains_segment ends_in_segment imageI is_interval_connected_1 is_interval_convex connected_continuous_image [OF assms]) lemma continuous_injective_image_segment_1: fixes f :: "'a::euclidean_space \ real" assumes contf: "continuous_on (closed_segment a b) f" and injf: "inj_on f (closed_segment a b)" shows "f ` (closed_segment a b) = closed_segment (f a) (f b)" proof show "closed_segment (f a) (f b) \ f ` closed_segment a b" by (metis subset_continuous_image_segment_1 contf) show "f ` closed_segment a b \ closed_segment (f a) (f b)" proof (cases "a = b") case True then show ?thesis by auto next case False then have fnot: "f a \ f b" using inj_onD injf by fastforce moreover have "f a \ open_segment (f c) (f b)" if c: "c \ closed_segment a b" for c proof (clarsimp simp add: open_segment_def) assume fa: "f a \ closed_segment (f c) (f b)" moreover have "closed_segment (f c) (f b) \ f ` closed_segment c b" by (meson closed_segment_subset contf continuous_on_subset convex_closed_segment ends_in_segment(2) subset_continuous_image_segment_1 that) ultimately have "f a \ f ` closed_segment c b" by blast then have a: "a \ closed_segment c b" by (meson ends_in_segment inj_on_image_mem_iff injf subset_closed_segment that) have cb: "closed_segment c b \ closed_segment a b" by (simp add: closed_segment_subset that) show "f a = f c" proof (rule between_antisym) show "between (f c, f b) (f a)" by (simp add: between_mem_segment fa) show "between (f a, f b) (f c)" by (metis a cb between_antisym between_mem_segment between_triv1 subset_iff) qed qed moreover have "f b \ open_segment (f a) (f c)" if c: "c \ closed_segment a b" for c proof (clarsimp simp add: open_segment_def fnot eq_commute) assume fb: "f b \ closed_segment (f a) (f c)" moreover have "closed_segment (f a) (f c) \ f ` closed_segment a c" by (meson contf continuous_on_subset ends_in_segment(1) subset_closed_segment subset_continuous_image_segment_1 that) ultimately have "f b \ f ` closed_segment a c" by blast then have b: "b \ closed_segment a c" by (meson ends_in_segment inj_on_image_mem_iff injf subset_closed_segment that) have ca: "closed_segment a c \ closed_segment a b" by (simp add: closed_segment_subset that) show "f b = f c" proof (rule between_antisym) show "between (f c, f a) (f b)" by (simp add: between_commute between_mem_segment fb) show "between (f b, f a) (f c)" by (metis b between_antisym between_commute between_mem_segment between_triv2 that) qed qed ultimately show ?thesis by (force simp: closed_segment_eq_real_ivl open_segment_eq_real_ivl split: if_split_asm) qed qed lemma continuous_injective_image_open_segment_1: fixes f :: "'a::euclidean_space \ real" assumes contf: "continuous_on (closed_segment a b) f" and injf: "inj_on f (closed_segment a b)" shows "f ` (open_segment a b) = open_segment (f a) (f b)" proof - have "f ` (open_segment a b) = f ` (closed_segment a b) - {f a, f b}" by (metis (no_types, opaque_lifting) empty_subsetI ends_in_segment image_insert image_is_empty inj_on_image_set_diff injf insert_subset open_segment_def segment_open_subset_closed) also have "... = open_segment (f a) (f b)" using continuous_injective_image_segment_1 [OF assms] by (simp add: open_segment_def inj_on_image_set_diff [OF injf]) finally show ?thesis . qed lemma collinear_imp_coplanar: "collinear s ==> coplanar s" by (metis collinear_affine_hull coplanar_def insert_absorb2) lemma collinear_small: assumes "finite s" "card s \ 2" shows "collinear s" proof - have "card s = 0 \ card s = 1 \ card s = 2" using assms by linarith then show ?thesis using assms using card_eq_SucD numeral_2_eq_2 by (force simp: card_1_singleton_iff) qed lemma coplanar_small: assumes "finite s" "card s \ 3" shows "coplanar s" proof - consider "card s \ 2" | "card s = Suc (Suc (Suc 0))" using assms by linarith then show ?thesis proof cases case 1 then show ?thesis by (simp add: \finite s\ collinear_imp_coplanar collinear_small) next case 2 then show ?thesis using hull_subset [of "{_,_,_}"] by (fastforce simp: coplanar_def dest!: card_eq_SucD) qed qed lemma coplanar_empty: "coplanar {}" by (simp add: coplanar_small) lemma coplanar_sing: "coplanar {a}" by (simp add: coplanar_small) lemma coplanar_2: "coplanar {a,b}" by (auto simp: card_insert_if coplanar_small) lemma coplanar_3: "coplanar {a,b,c}" by (auto simp: card_insert_if coplanar_small) lemma collinear_affine_hull_collinear: "collinear(affine hull s) \ collinear s" unfolding collinear_affine_hull by (metis affine_affine_hull subset_hull hull_hull hull_mono) lemma coplanar_affine_hull_coplanar: "coplanar(affine hull s) \ coplanar s" unfolding coplanar_def by (metis affine_affine_hull subset_hull hull_hull hull_mono) lemma coplanar_linear_image: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "coplanar S" "linear f" shows "coplanar(f ` S)" proof - { fix u v w assume "S \ affine hull {u, v, w}" then have "f ` S \ f ` (affine hull {u, v, w})" by (simp add: image_mono) then have "f ` S \ affine hull (f ` {u, v, w})" by (metis assms(2) linear_conv_bounded_linear affine_hull_linear_image) } then show ?thesis by auto (meson assms(1) coplanar_def) qed lemma coplanar_translation_imp: assumes "coplanar S" shows "coplanar ((\x. a + x) ` S)" proof - obtain u v w where "S \ affine hull {u,v,w}" by (meson assms coplanar_def) then have "(+) a ` S \ affine hull {u + a, v + a, w + a}" using affine_hull_translation [of a "{u,v,w}" for u v w] by (force simp: add.commute) then show ?thesis unfolding coplanar_def by blast qed lemma coplanar_translation_eq: "coplanar((\x. a + x) ` S) \ coplanar S" by (metis (no_types) coplanar_translation_imp translation_galois) lemma coplanar_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "coplanar(f ` S) = coplanar S" proof assume "coplanar S" then show "coplanar (f ` S)" using assms(1) coplanar_linear_image by blast next obtain g where g: "linear g" "g \ f = id" using linear_injective_left_inverse [OF assms] by blast assume "coplanar (f ` S)" then show "coplanar S" by (metis coplanar_linear_image g(1) g(2) id_apply image_comp image_id) qed lemma coplanar_subset: "\coplanar t; S \ t\ \ coplanar S" by (meson coplanar_def order_trans) lemma affine_hull_3_imp_collinear: "c \ affine hull {a,b} \ collinear {a,b,c}" by (metis collinear_2 collinear_affine_hull_collinear hull_redundant insert_commute) lemma collinear_3_imp_in_affine_hull: assumes "collinear {a,b,c}" "a \ b" shows "c \ affine hull {a,b}" proof - obtain u x y where "b - a = y *\<^sub>R u" "c - a = x *\<^sub>R u" using assms unfolding collinear_def by auto with \a \ b\ have "\v. c = (1 - x / y) *\<^sub>R a + v *\<^sub>R b \ 1 - x / y + v = 1" by (simp add: algebra_simps) then show ?thesis by (simp add: hull_inc mem_affine) qed lemma collinear_3_affine_hull: assumes "a \ b" shows "collinear {a,b,c} \ c \ affine hull {a,b}" using affine_hull_3_imp_collinear assms collinear_3_imp_in_affine_hull by blast lemma collinear_3_eq_affine_dependent: "collinear{a,b,c} \ a = b \ a = c \ b = c \ affine_dependent {a,b,c}" proof (cases "a = b \ a = c \ b = c") case True then show ?thesis by (auto simp: insert_commute) next case False then have "collinear{a,b,c}" if "affine_dependent {a,b,c}" using that unfolding affine_dependent_def by (auto simp: insert_Diff_if; metis affine_hull_3_imp_collinear insert_commute) moreover have "affine_dependent {a,b,c}" if "collinear{a,b,c}" using False that by (auto simp: affine_dependent_def collinear_3_affine_hull insert_Diff_if) ultimately show ?thesis using False by blast qed lemma affine_dependent_imp_collinear_3: "affine_dependent {a,b,c} \ collinear{a,b,c}" by (simp add: collinear_3_eq_affine_dependent) lemma collinear_3: "NO_MATCH 0 x \ collinear {x,y,z} \ collinear {0, x-y, z-y}" by (auto simp add: collinear_def) lemma collinear_3_expand: "collinear{a,b,c} \ a = c \ (\u. b = u *\<^sub>R a + (1 - u) *\<^sub>R c)" proof - have "collinear{a,b,c} = collinear{a,c,b}" by (simp add: insert_commute) also have "... = collinear {0, a - c, b - c}" by (simp add: collinear_3) also have "... \ (a = c \ b = c \ (\ca. b - c = ca *\<^sub>R (a - c)))" by (simp add: collinear_lemma) also have "... \ a = c \ (\u. b = u *\<^sub>R a + (1 - u) *\<^sub>R c)" by (cases "a = c \ b = c") (auto simp: algebra_simps) finally show ?thesis . qed lemma collinear_aff_dim: "collinear S \ aff_dim S \ 1" proof assume "collinear S" then obtain u and v :: "'a" where "aff_dim S \ aff_dim {u,v}" by (metis \collinear S\ aff_dim_affine_hull aff_dim_subset collinear_affine_hull) then show "aff_dim S \ 1" using order_trans by fastforce next assume "aff_dim S \ 1" then have le1: "aff_dim (affine hull S) \ 1" by simp obtain B where "B \ S" and B: "\ affine_dependent B" "affine hull S = affine hull B" using affine_basis_exists [of S] by auto then have "finite B" "card B \ 2" using B le1 by (auto simp: affine_independent_iff_card) then have "collinear B" by (rule collinear_small) then show "collinear S" by (metis \affine hull S = affine hull B\ collinear_affine_hull_collinear) qed lemma collinear_midpoint: "collinear{a, midpoint a b, b}" proof - have \
: "\a \ midpoint a b; b - midpoint a b \ - 1 *\<^sub>R (a - midpoint a b)\ \ b = midpoint a b" by (simp add: algebra_simps) show ?thesis by (auto simp: collinear_3 collinear_lemma intro: \
) qed lemma midpoint_collinear: fixes a b c :: "'a::real_normed_vector" assumes "a \ c" shows "b = midpoint a c \ collinear{a,b,c} \ dist a b = dist b c" proof - have *: "a - (u *\<^sub>R a + (1 - u) *\<^sub>R c) = (1 - u) *\<^sub>R (a - c)" "u *\<^sub>R a + (1 - u) *\<^sub>R c - c = u *\<^sub>R (a - c)" "\1 - u\ = \u\ \ u = 1/2" for u::real by (auto simp: algebra_simps) have "b = midpoint a c \ collinear{a,b,c}" using collinear_midpoint by blast moreover have "b = midpoint a c \ dist a b = dist b c" if "collinear{a,b,c}" proof - consider "a = c" | u where "b = u *\<^sub>R a + (1 - u) *\<^sub>R c" using \collinear {a,b,c}\ unfolding collinear_3_expand by blast then show ?thesis proof cases case 2 with assms have "dist a b = dist b c \ b = midpoint a c" by (simp add: dist_norm * midpoint_def scaleR_add_right del: divide_const_simps) then show ?thesis by (auto simp: dist_midpoint) qed (use assms in auto) qed ultimately show ?thesis by blast qed lemma between_imp_collinear: fixes x :: "'a :: euclidean_space" assumes "between (a,b) x" shows "collinear {a,x,b}" proof (cases "x = a \ x = b \ a = b") case True with assms show ?thesis by (auto simp: dist_commute) next case False then have False if "\c. b - x \ c *\<^sub>R (a - x)" using that [of "-(norm(b - x) / norm(x - a))"] assms by (simp add: between_norm vector_add_divide_simps flip: real_vector.scale_minus_right) then show ?thesis by (auto simp: collinear_3 collinear_lemma) qed lemma midpoint_between: fixes a b :: "'a::euclidean_space" shows "b = midpoint a c \ between (a,c) b \ dist a b = dist b c" proof (cases "a = c") case False show ?thesis using False between_imp_collinear between_midpoint(1) midpoint_collinear by blast qed (auto simp: dist_commute) lemma collinear_triples: assumes "a \ b" shows "collinear(insert a (insert b S)) \ (\x \ S. collinear{a,b,x})" (is "?lhs = ?rhs") proof safe fix x assume ?lhs and "x \ S" then show "collinear {a, b, x}" using collinear_subset by force next assume ?rhs then have "\x \ S. collinear{a,x,b}" by (simp add: insert_commute) then have *: "\u. x = u *\<^sub>R a + (1 - u) *\<^sub>R b" if "x \ insert a (insert b S)" for x using that assms collinear_3_expand by fastforce+ have "\c. x - y = c *\<^sub>R (b - a)" if x: "x \ insert a (insert b S)" and y: "y \ insert a (insert b S)" for x y proof - obtain u v where "x = u *\<^sub>R a + (1 - u) *\<^sub>R b" "y = v *\<^sub>R a + (1 - v) *\<^sub>R b" using "*" x y by presburger then have "x - y = (v - u) *\<^sub>R (b - a)" by (simp add: scale_left_diff_distrib scale_right_diff_distrib) then show ?thesis .. qed then show ?lhs unfolding collinear_def by metis qed lemma collinear_4_3: assumes "a \ b" shows "collinear {a,b,c,d} \ collinear{a,b,c} \ collinear{a,b,d}" using collinear_triples [OF assms, of "{c,d}"] by (force simp:) lemma collinear_3_trans: assumes "collinear{a,b,c}" "collinear{b,c,d}" "b \ c" shows "collinear{a,b,d}" proof - have "collinear{b,c,a,d}" by (metis (full_types) assms collinear_4_3 insert_commute) then show ?thesis by (simp add: collinear_subset) qed lemma affine_hull_2_alt: fixes a b :: "'a::real_vector" shows "affine hull {a,b} = range (\u. a + u *\<^sub>R (b - a))" proof - have 1: "u *\<^sub>R a + v *\<^sub>R b = a + v *\<^sub>R (b - a)" if "u + v = 1" for u v using that by (simp add: algebra_simps flip: scaleR_add_left) have 2: "a + u *\<^sub>R (b - a) = (1 - u) *\<^sub>R a + u *\<^sub>R b" for u by (auto simp: algebra_simps) show ?thesis by (force simp add: affine_hull_2 dest: 1 intro!: 2) qed lemma interior_convex_hull_3_minimal: fixes a :: "'a::euclidean_space" assumes "\ collinear{a,b,c}" and 2: "DIM('a) = 2" shows "interior(convex hull {a,b,c}) = {v. \x y z. 0 < x \ 0 < y \ 0 < z \ x + y + z = 1 \ x *\<^sub>R a + y *\<^sub>R b + z *\<^sub>R c = v}" (is "?lhs = ?rhs") proof have abc: "a \ b" "a \ c" "b \ c" "\ affine_dependent {a, b, c}" using assms by (auto simp: collinear_3_eq_affine_dependent) with 2 show "?lhs \ ?rhs" by (fastforce simp add: interior_convex_hull_explicit_minimal) show "?rhs \ ?lhs" using abc 2 apply (clarsimp simp add: interior_convex_hull_explicit_minimal) subgoal for x y z by (rule_tac x="\r. (if r=a then x else if r=b then y else if r=c then z else 0)" in exI) auto done qed subsection\<^marker>\tag unimportant\\Basic lemmas about hyperplanes and halfspaces\ lemma halfspace_Int_eq: "{x. a \ x \ b} \ {x. b \ a \ x} = {x. a \ x = b}" "{x. b \ a \ x} \ {x. a \ x \ b} = {x. a \ x = b}" by auto lemma hyperplane_eq_Ex: assumes "a \ 0" obtains x where "a \ x = b" by (rule_tac x = "(b / (a \ a)) *\<^sub>R a" in that) (simp add: assms) lemma hyperplane_eq_empty: "{x. a \ x = b} = {} \ a = 0 \ b \ 0" using hyperplane_eq_Ex by (metis (mono_tags, lifting) empty_Collect_eq inner_zero_left) lemma hyperplane_eq_UNIV: "{x. a \ x = b} = UNIV \ a = 0 \ b = 0" proof - have "a = 0 \ b = 0" if "UNIV \ {x. a \ x = b}" using subsetD [OF that, where c = "((b+1) / (a \ a)) *\<^sub>R a"] by (simp add: field_split_simps split: if_split_asm) then show ?thesis by force qed lemma halfspace_eq_empty_lt: "{x. a \ x < b} = {} \ a = 0 \ b \ 0" proof - have "a = 0 \ b \ 0" if "{x. a \ x < b} \ {}" using subsetD [OF that, where c = "((b-1) / (a \ a)) *\<^sub>R a"] by (force simp add: field_split_simps split: if_split_asm) then show ?thesis by force qed lemma halfspace_eq_empty_gt: "{x. a \ x > b} = {} \ a = 0 \ b \ 0" using halfspace_eq_empty_lt [of "-a" "-b"] by simp lemma halfspace_eq_empty_le: "{x. a \ x \ b} = {} \ a = 0 \ b < 0" proof - have "a = 0 \ b < 0" if "{x. a \ x \ b} \ {}" using subsetD [OF that, where c = "((b-1) / (a \ a)) *\<^sub>R a"] by (force simp add: field_split_simps split: if_split_asm) then show ?thesis by force qed lemma halfspace_eq_empty_ge: "{x. a \ x \ b} = {} \ a = 0 \ b > 0" using halfspace_eq_empty_le [of "-a" "-b"] by simp subsection\<^marker>\tag unimportant\\Use set distance for an easy proof of separation properties\ proposition\<^marker>\tag unimportant\ separation_closures: fixes S :: "'a::euclidean_space set" assumes "S \ closure T = {}" "T \ closure S = {}" obtains U V where "U \ V = {}" "open U" "open V" "S \ U" "T \ V" proof (cases "S = {} \ T = {}") case True with that show ?thesis by auto next case False define f where "f \ \x. setdist {x} T - setdist {x} S" have contf: "continuous_on UNIV f" unfolding f_def by (intro continuous_intros continuous_on_setdist) show ?thesis proof (rule_tac U = "{x. f x > 0}" and V = "{x. f x < 0}" in that) show "{x. 0 < f x} \ {x. f x < 0} = {}" by auto show "open {x. 0 < f x}" by (simp add: open_Collect_less contf) show "open {x. f x < 0}" by (simp add: open_Collect_less contf) have "\x. x \ S \ setdist {x} T \ 0" "\x. x \ T \ setdist {x} S \ 0" by (meson False assms disjoint_iff setdist_eq_0_sing_1)+ then show "S \ {x. 0 < f x}" "T \ {x. f x < 0}" using less_eq_real_def by (fastforce simp add: f_def setdist_sing_in_set)+ qed qed lemma separation_normal: fixes S :: "'a::euclidean_space set" assumes "closed S" "closed T" "S \ T = {}" obtains U V where "open U" "open V" "S \ U" "T \ V" "U \ V = {}" using separation_closures [of S T] by (metis assms closure_closed disjnt_def inf_commute) lemma separation_normal_local: fixes S :: "'a::euclidean_space set" assumes US: "closedin (top_of_set U) S" and UT: "closedin (top_of_set U) T" and "S \ T = {}" obtains S' T' where "openin (top_of_set U) S'" "openin (top_of_set U) T'" "S \ S'" "T \ T'" "S' \ T' = {}" proof (cases "S = {} \ T = {}") case True with that show ?thesis using UT US by (blast dest: closedin_subset) next case False define f where "f \ \x. setdist {x} T - setdist {x} S" have contf: "continuous_on U f" unfolding f_def by (intro continuous_intros) show ?thesis proof (rule_tac S' = "(U \ f -` {0<..})" and T' = "(U \ f -` {..<0})" in that) show "(U \ f -` {0<..}) \ (U \ f -` {..<0}) = {}" by auto show "openin (top_of_set U) (U \ f -` {0<..})" by (rule continuous_openin_preimage [where T=UNIV]) (simp_all add: contf) next show "openin (top_of_set U) (U \ f -` {..<0})" by (rule continuous_openin_preimage [where T=UNIV]) (simp_all add: contf) next have "S \ U" "T \ U" using closedin_imp_subset assms by blast+ then show "S \ U \ f -` {0<..}" "T \ U \ f -` {..<0}" using assms False by (force simp add: f_def setdist_sing_in_set intro!: setdist_gt_0_closedin)+ qed qed lemma separation_normal_compact: fixes S :: "'a::euclidean_space set" assumes "compact S" "closed T" "S \ T = {}" obtains U V where "open U" "compact(closure U)" "open V" "S \ U" "T \ V" "U \ V = {}" proof - have "closed S" "bounded S" using assms by (auto simp: compact_eq_bounded_closed) then obtain r where "r>0" and r: "S \ ball 0 r" by (auto dest!: bounded_subset_ballD) have **: "closed (T \ - ball 0 r)" "S \ (T \ - ball 0 r) = {}" using assms r by blast+ then obtain U V where UV: "open U" "open V" "S \ U" "T \ - ball 0 r \ V" "U \ V = {}" by (meson \closed S\ separation_normal) then have "compact(closure U)" by (meson bounded_ball bounded_subset compact_closure compl_le_swap2 disjoint_eq_subset_Compl le_sup_iff) with UV show thesis using that by auto qed subsection\Connectedness of the intersection of a chain\ proposition connected_chain: fixes \ :: "'a :: euclidean_space set set" assumes cc: "\S. S \ \ \ compact S \ connected S" and linear: "\S T. S \ \ \ T \ \ \ S \ T \ T \ S" shows "connected(\\)" proof (cases "\ = {}") case True then show ?thesis by auto next case False then have cf: "compact(\\)" by (simp add: cc compact_Inter) have False if AB: "closed A" "closed B" "A \ B = {}" and ABeq: "A \ B = \\" and "A \ {}" "B \ {}" for A B proof - obtain U V where "open U" "open V" "A \ U" "B \ V" "U \ V = {}" using separation_normal [OF AB] by metis obtain K where "K \ \" "compact K" using cc False by blast then obtain N where "open N" and "K \ N" by blast let ?\ = "insert (U \ V) ((\S. N - S) ` \)" obtain \ where "\ \ ?\" "finite \" "K \ \\" proof (rule compactE [OF \compact K\]) show "K \ \(insert (U \ V) ((-) N ` \))" using \K \ N\ ABeq \A \ U\ \B \ V\ by auto show "\B. B \ insert (U \ V) ((-) N ` \) \ open B" by (auto simp: \open U\ \open V\ open_Un \open N\ cc compact_imp_closed open_Diff) qed then have "finite(\ - {U \ V})" by blast moreover have "\ - {U \ V} \ (\S. N - S) ` \" using \\ \ ?\\ by blast ultimately obtain \ where "\ \ \" "finite \" and Deq: "\ - {U \ V} = (\S. N-S) ` \" using finite_subset_image by metis obtain J where "J \ \" and J: "(\S\\. N - S) \ N - J" proof (cases "\ = {}") case True with \\ \ {}\ that show ?thesis by auto next case False have "\S T. \S \ \; T \ \\ \ S \ T \ T \ S" by (meson \\ \ \\ in_mono local.linear) with \finite \\ \\ \ {}\ have "\J \ \. (\S\\. N - S) \ N - J" proof induction case (insert X \) show ?case proof (cases "\ = {}") case True then show ?thesis by auto next case False then have "\S T. \S \ \; T \ \\ \ S \ T \ T \ S" by (simp add: insert.prems) with insert.IH False obtain J where "J \ \" and J: "(\Y\\. N - Y) \ N - J" by metis have "N - J \ N - X \ N - X \ N - J" by (meson Diff_mono \J \ \\ insert.prems(2) insert_iff order_refl) then show ?thesis proof assume "N - J \ N - X" with J show ?thesis by auto next assume "N - X \ N - J" with J have "N - X \ \ ((-) N ` \) \ N - J" by auto with \J \ \\ show ?thesis by blast qed qed qed simp with \\ \ \\ show ?thesis by (blast intro: that) qed have "K \ \(insert (U \ V) (\ - {U \ V}))" using \K \ \\\ by auto also have "... \ (U \ V) \ (N - J)" by (metis (no_types, opaque_lifting) Deq Un_subset_iff Un_upper2 J Union_insert order_trans sup_ge1) finally have "J \ K \ U \ V" by blast moreover have "connected(J \ K)" by (metis Int_absorb1 \J \ \\ \K \ \\ cc inf.orderE local.linear) moreover have "U \ (J \ K) \ {}" using ABeq \J \ \\ \K \ \\ \A \ {}\ \A \ U\ by blast moreover have "V \ (J \ K) \ {}" using ABeq \J \ \\ \K \ \\ \B \ {}\ \B \ V\ by blast ultimately show False using connectedD [of "J \ K" U V] \open U\ \open V\ \U \ V = {}\ by auto qed with cf show ?thesis by (auto simp: connected_closed_set compact_imp_closed) qed lemma connected_chain_gen: fixes \ :: "'a :: euclidean_space set set" assumes X: "X \ \" "compact X" and cc: "\T. T \ \ \ closed T \ connected T" and linear: "\S T. S \ \ \ T \ \ \ S \ T \ T \ S" shows "connected(\\)" proof - have "\\ = (\T\\. X \ T)" using X by blast moreover have "connected (\T\\. X \ T)" proof (rule connected_chain) show "\T. T \ (\) X ` \ \ compact T \ connected T" using cc X by auto (metis inf.absorb2 inf.orderE local.linear) show "\S T. S \ (\) X ` \ \ T \ (\) X ` \ \ S \ T \ T \ S" using local.linear by blast qed ultimately show ?thesis by metis qed lemma connected_nest: fixes S :: "'a::linorder \ 'b::euclidean_space set" assumes S: "\n. compact(S n)" "\n. connected(S n)" and nest: "\m n. m \ n \ S n \ S m" shows "connected(\ (range S))" proof (rule connected_chain) show "\A T. A \ range S \ T \ range S \ A \ T \ T \ A" by (metis image_iff le_cases nest) qed (use S in blast) lemma connected_nest_gen: fixes S :: "'a::linorder \ 'b::euclidean_space set" assumes S: "\n. closed(S n)" "\n. connected(S n)" "compact(S k)" and nest: "\m n. m \ n \ S n \ S m" shows "connected(\ (range S))" proof (rule connected_chain_gen [of "S k"]) show "\A T. A \ range S \ T \ range S \ A \ T \ T \ A" by (metis imageE le_cases nest) qed (use S in auto) subsection\Proper maps, including projections out of compact sets\ lemma finite_indexed_bound: assumes A: "finite A" "\x. x \ A \ \n::'a::linorder. P x n" shows "\m. \x \ A. \k\m. P x k" using A proof (induction A) case empty then show ?case by force next case (insert a A) then obtain m n where "\x \ A. \k\m. P x k" "P a n" by force then show ?case by (metis dual_order.trans insert_iff le_cases) qed proposition proper_map: fixes f :: "'a::heine_borel \ 'b::heine_borel" assumes "closedin (top_of_set S) K" and com: "\U. \U \ T; compact U\ \ compact (S \ f -` U)" and "f ` S \ T" shows "closedin (top_of_set T) (f ` K)" proof - have "K \ S" using assms closedin_imp_subset by metis obtain C where "closed C" and Keq: "K = S \ C" using assms by (auto simp: closedin_closed) have *: "y \ f ` K" if "y \ T" and y: "y islimpt f ` K" for y proof - obtain h where "\n. (\x\K. h n = f x) \ h n \ y" "inj h" and hlim: "(h \ y) sequentially" using \y \ T\ y by (force simp: limpt_sequential_inj) then obtain X where X: "\n. X n \ K \ h n = f (X n) \ h n \ y" by metis then have fX: "\n. f (X n) = h n" by metis define \ where "\ \ \n. {a \ K. f a \ insert y (range (\i. f (X (n + i))))}" have "compact (C \ (S \ f -` insert y (range (\i. f(X(n + i))))))" for n proof (intro closed_Int_compact [OF \closed C\ com] compact_sequence_with_limit) show "insert y (range (\i. f (X (n + i)))) \ T" using X \K \ S\ \f ` S \ T\ \y \ T\ by blast show "(\i. f (X (n + i))) \ y" by (simp add: fX add.commute [of n] LIMSEQ_ignore_initial_segment [OF hlim]) qed then have comf: "compact (\ n)" for n by (simp add: Keq Int_def \_def conj_commute) have ne: "\\ \ {}" if "finite \" and \: "\t. t \ \ \ (\n. t = \ n)" for \ proof - obtain m where m: "\t. t \ \ \ \k\m. t = \ k" by (rule exE [OF finite_indexed_bound [OF \finite \\ \]], force+) have "X m \ \\" using X le_Suc_ex by (fastforce simp: \_def dest: m) then show ?thesis by blast qed have "(\n. \ n) \ {}" proof (rule compact_fip_Heine_Borel) show "\\'. \finite \'; \' \ range \\ \ \ \' \ {}" by (meson ne rangeE subset_eq) qed (use comf in blast) then obtain x where "x \ K" "\n. (f x = y \ (\u. f x = h (n + u)))" by (force simp add: \_def fX) then show ?thesis unfolding image_iff by (metis \inj h\ le_add1 not_less_eq_eq rangeI range_ex1_eq) qed with assms closedin_subset show ?thesis by (force simp: closedin_limpt) qed lemma compact_continuous_image_eq: fixes f :: "'a::heine_borel \ 'b::heine_borel" assumes f: "inj_on f S" shows "continuous_on S f \ (\T. compact T \ T \ S \ compact(f ` T))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (metis continuous_on_subset compact_continuous_image) next assume RHS: ?rhs obtain g where gf: "\x. x \ S \ g (f x) = x" by (metis inv_into_f_f f) then have *: "(S \ f -` U) = g ` U" if "U \ f ` S" for U using that by fastforce have gfim: "g ` f ` S \ S" using gf by auto have **: "compact (f ` S \ g -` C)" if C: "C \ S" "compact C" for C proof - obtain h where "h C \ C \ h C \ S \ compact (f ` C)" by (force simp: C RHS) moreover have "f ` C = (f ` S \ g -` C)" using C gf by auto ultimately show ?thesis using C by auto qed show ?lhs using proper_map [OF _ _ gfim] ** by (simp add: continuous_on_closed * closedin_imp_subset) qed subsection\<^marker>\tag unimportant\\Trivial fact: convexity equals connectedness for collinear sets\ lemma convex_connected_collinear: fixes S :: "'a::euclidean_space set" assumes "collinear S" shows "convex S \ connected S" proof assume "convex S" then show "connected S" using convex_connected by blast next assume S: "connected S" show "convex S" proof (cases "S = {}") case True then show ?thesis by simp next case False then obtain a where "a \ S" by auto have "collinear (affine hull S)" by (simp add: assms collinear_affine_hull_collinear) then obtain z where "z \ 0" "\x. x \ affine hull S \ \c. x - a = c *\<^sub>R z" by (meson \a \ S\ collinear hull_inc) then obtain f where f: "\x. x \ affine hull S \ x - a = f x *\<^sub>R z" by metis then have inj_f: "inj_on f (affine hull S)" by (metis diff_add_cancel inj_onI) have diff: "x - y = (f x - f y) *\<^sub>R z" if x: "x \ affine hull S" and y: "y \ affine hull S" for x y proof - have "f x *\<^sub>R z = x - a" by (simp add: f hull_inc x) moreover have "f y *\<^sub>R z = y - a" by (simp add: f hull_inc y) ultimately show ?thesis by (simp add: scaleR_left.diff) qed have cont_f: "continuous_on (affine hull S) f" proof (clarsimp simp: dist_norm continuous_on_iff diff) show "\x e. 0 < e \ \d>0. \y \ affine hull S. \f y - f x\ * norm z < d \ \f y - f x\ < e" by (metis \z \ 0\ mult_pos_pos mult_less_iff1 zero_less_norm_iff) qed then have conn_fS: "connected (f ` S)" by (meson S connected_continuous_image continuous_on_subset hull_subset) show ?thesis proof (clarsimp simp: convex_contains_segment) fix x y z assume "x \ S" "y \ S" "z \ closed_segment x y" have False if "z \ S" proof - have "f ` (closed_segment x y) = closed_segment (f x) (f y)" proof (rule continuous_injective_image_segment_1) show "continuous_on (closed_segment x y) f" by (meson \x \ S\ \y \ S\ convex_affine_hull convex_contains_segment hull_inc continuous_on_subset [OF cont_f]) show "inj_on f (closed_segment x y)" by (meson \x \ S\ \y \ S\ convex_affine_hull convex_contains_segment hull_inc inj_on_subset [OF inj_f]) qed then have fz: "f z \ closed_segment (f x) (f y)" using \z \ closed_segment x y\ by blast have "z \ affine hull S" by (meson \x \ S\ \y \ S\ \z \ closed_segment x y\ convex_affine_hull convex_contains_segment hull_inc subset_eq) then have fz_notin: "f z \ f ` S" using hull_subset inj_f inj_onD that by fastforce moreover have "{.. f ` S \ {}" "{f z<..} \ f ` S \ {}" proof - consider "f x \ f z \ f z \ f y" | "f y \ f z \ f z \ f x" using fz by (auto simp add: closed_segment_eq_real_ivl split: if_split_asm) then have "{.. f ` {x,y} \ {} \ {f z<..} \ f ` {x,y} \ {}" by cases (use fz_notin \x \ S\ \y \ S\ in \auto simp: image_iff\) then show "{.. f ` S \ {}" "{f z<..} \ f ` S \ {}" using \x \ S\ \y \ S\ by blast+ qed ultimately show False using connectedD [OF conn_fS, of "{.. S" by meson qed qed qed lemma compact_convex_collinear_segment_alt: fixes S :: "'a::euclidean_space set" assumes "S \ {}" "compact S" "connected S" "collinear S" obtains a b where "S = closed_segment a b" proof - obtain \ where "\ \ S" using \S \ {}\ by auto have "collinear (affine hull S)" by (simp add: assms collinear_affine_hull_collinear) then obtain z where "z \ 0" "\x. x \ affine hull S \ \c. x - \ = c *\<^sub>R z" by (meson \\ \ S\ collinear hull_inc) then obtain f where f: "\x. x \ affine hull S \ x - \ = f x *\<^sub>R z" by metis let ?g = "\r. r *\<^sub>R z + \" have gf: "?g (f x) = x" if "x \ affine hull S" for x by (metis diff_add_cancel f that) then have inj_f: "inj_on f (affine hull S)" by (metis inj_onI) have diff: "x - y = (f x - f y) *\<^sub>R z" if x: "x \ affine hull S" and y: "y \ affine hull S" for x y proof - have "f x *\<^sub>R z = x - \" by (simp add: f hull_inc x) moreover have "f y *\<^sub>R z = y - \" by (simp add: f hull_inc y) ultimately show ?thesis by (simp add: scaleR_left.diff) qed have cont_f: "continuous_on (affine hull S) f" proof (clarsimp simp: dist_norm continuous_on_iff diff) show "\x e. 0 < e \ \d>0. \y \ affine hull S. \f y - f x\ * norm z < d \ \f y - f x\ < e" by (metis \z \ 0\ mult_pos_pos mult_less_iff1 zero_less_norm_iff) qed then have "connected (f ` S)" by (meson \connected S\ connected_continuous_image continuous_on_subset hull_subset) moreover have "compact (f ` S)" by (meson \compact S\ compact_continuous_image_eq cont_f hull_subset inj_f) ultimately obtain x y where "f ` S = {x..y}" by (meson connected_compact_interval_1) then have fS_eq: "f ` S = closed_segment x y" using \S \ {}\ closed_segment_eq_real_ivl by auto obtain a b where "a \ S" "f a = x" "b \ S" "f b = y" by (metis (full_types) ends_in_segment fS_eq imageE) have "f ` (closed_segment a b) = closed_segment (f a) (f b)" proof (rule continuous_injective_image_segment_1) show "continuous_on (closed_segment a b) f" by (meson \a \ S\ \b \ S\ convex_affine_hull convex_contains_segment hull_inc continuous_on_subset [OF cont_f]) show "inj_on f (closed_segment a b)" by (meson \a \ S\ \b \ S\ convex_affine_hull convex_contains_segment hull_inc inj_on_subset [OF inj_f]) qed then have "f ` (closed_segment a b) = f ` S" by (simp add: \f a = x\ \f b = y\ fS_eq) then have "?g ` f ` (closed_segment a b) = ?g ` f ` S" by simp moreover have "(\x. f x *\<^sub>R z + \) ` closed_segment a b = closed_segment a b" unfolding image_def using \a \ S\ \b \ S\ by (safe; metis (mono_tags, lifting) convex_affine_hull convex_contains_segment gf hull_subset subsetCE) ultimately have "closed_segment a b = S" using gf by (simp add: image_comp o_def hull_inc cong: image_cong) then show ?thesis using that by blast qed lemma compact_convex_collinear_segment: fixes S :: "'a::euclidean_space set" assumes "S \ {}" "compact S" "convex S" "collinear S" obtains a b where "S = closed_segment a b" using assms convex_connected_collinear compact_convex_collinear_segment_alt by blast lemma proper_map_from_compact: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and imf: "f ` S \ T" and "compact S" "closedin (top_of_set T) K" shows "compact (S \ f -` K)" by (rule closedin_compact [OF \compact S\] continuous_closedin_preimage_gen assms)+ lemma proper_map_fst: assumes "compact T" "K \ S" "compact K" shows "compact (S \ T \ fst -` K)" proof - have "(S \ T \ fst -` K) = K \ T" using assms by auto then show ?thesis by (simp add: assms compact_Times) qed lemma closed_map_fst: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "compact T" "closedin (top_of_set (S \ T)) c" shows "closedin (top_of_set S) (fst ` c)" proof - have *: "fst ` (S \ T) \ S" by auto show ?thesis using proper_map [OF _ _ *] by (simp add: proper_map_fst assms) qed lemma proper_map_snd: assumes "compact S" "K \ T" "compact K" shows "compact (S \ T \ snd -` K)" proof - have "(S \ T \ snd -` K) = S \ K" using assms by auto then show ?thesis by (simp add: assms compact_Times) qed lemma closed_map_snd: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "compact S" "closedin (top_of_set (S \ T)) c" shows "closedin (top_of_set T) (snd ` c)" proof - have *: "snd ` (S \ T) \ T" by auto show ?thesis using proper_map [OF _ _ *] by (simp add: proper_map_snd assms) qed lemma closedin_compact_projection: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "compact S" and clo: "closedin (top_of_set (S \ T)) U" shows "closedin (top_of_set T) {y. \x. x \ S \ (x, y) \ U}" proof - have "U \ S \ T" by (metis clo closedin_imp_subset) then have "{y. \x. x \ S \ (x, y) \ U} = snd ` U" by force moreover have "closedin (top_of_set T) (snd ` U)" by (rule closed_map_snd [OF assms]) ultimately show ?thesis by simp qed lemma closed_compact_projection: fixes S :: "'a::euclidean_space set" and T :: "('a * 'b::euclidean_space) set" assumes "compact S" and clo: "closed T" shows "closed {y. \x. x \ S \ (x, y) \ T}" proof - have *: "{y. \x. x \ S \ Pair x y \ T} = {y. \x. x \ S \ Pair x y \ ((S \ UNIV) \ T)}" by auto show ?thesis unfolding * by (intro clo closedin_closed_Int closedin_closed_trans [OF _ closed_UNIV] closedin_compact_projection [OF \compact S\]) qed subsubsection\<^marker>\tag unimportant\\Representing affine hull as a finite intersection of hyperplanes\ proposition\<^marker>\tag unimportant\ affine_hull_convex_Int_nonempty_interior: fixes S :: "'a::real_normed_vector set" assumes "convex S" "S \ interior T \ {}" shows "affine hull (S \ T) = affine hull S" proof show "affine hull (S \ T) \ affine hull S" by (simp add: hull_mono) next obtain a where "a \ S" "a \ T" and at: "a \ interior T" using assms interior_subset by blast then obtain e where "e > 0" and e: "cball a e \ T" using mem_interior_cball by blast have *: "x \ (+) a ` span ((\x. x - a) ` (S \ T))" if "x \ S" for x proof (cases "x = a") case True with that span_0 eq_add_iff image_def mem_Collect_eq show ?thesis by blast next case False define k where "k = min (1/2) (e / norm (x-a))" have k: "0 < k" "k < 1" using \e > 0\ False by (auto simp: k_def) then have xa: "(x-a) = inverse k *\<^sub>R k *\<^sub>R (x-a)" by simp have "e / norm (x - a) \ k" using k_def by linarith then have "a + k *\<^sub>R (x - a) \ cball a e" using \0 < k\ False by (simp add: dist_norm) (simp add: field_simps) then have T: "a + k *\<^sub>R (x - a) \ T" using e by blast have S: "a + k *\<^sub>R (x - a) \ S" using k \a \ S\ convexD [OF \convex S\ \a \ S\ \x \ S\, of "1-k" k] by (simp add: algebra_simps) have "inverse k *\<^sub>R k *\<^sub>R (x-a) \ span ((\x. x - a) ` (S \ T))" by (intro span_mul [OF span_base] image_eqI [where x = "a + k *\<^sub>R (x - a)"]) (auto simp: S T) with xa image_iff show ?thesis by fastforce qed have "S \ affine hull (S \ T)" by (force simp: * \a \ S\ \a \ T\ hull_inc affine_hull_span_gen [of a]) then show "affine hull S \ affine hull (S \ T)" by (simp add: subset_hull) qed corollary affine_hull_convex_Int_open: fixes S :: "'a::real_normed_vector set" assumes "convex S" "open T" "S \ T \ {}" shows "affine hull (S \ T) = affine hull S" using affine_hull_convex_Int_nonempty_interior assms interior_eq by blast corollary affine_hull_affine_Int_nonempty_interior: fixes S :: "'a::real_normed_vector set" assumes "affine S" "S \ interior T \ {}" shows "affine hull (S \ T) = affine hull S" by (simp add: affine_hull_convex_Int_nonempty_interior affine_imp_convex assms) corollary affine_hull_affine_Int_open: fixes S :: "'a::real_normed_vector set" assumes "affine S" "open T" "S \ T \ {}" shows "affine hull (S \ T) = affine hull S" by (simp add: affine_hull_convex_Int_open affine_imp_convex assms) corollary affine_hull_convex_Int_openin: fixes S :: "'a::real_normed_vector set" assumes "convex S" "openin (top_of_set (affine hull S)) T" "S \ T \ {}" shows "affine hull (S \ T) = affine hull S" using assms unfolding openin_open by (metis affine_hull_convex_Int_open hull_subset inf.orderE inf_assoc) corollary affine_hull_openin: fixes S :: "'a::real_normed_vector set" assumes "openin (top_of_set (affine hull T)) S" "S \ {}" shows "affine hull S = affine hull T" using assms unfolding openin_open by (metis affine_affine_hull affine_hull_affine_Int_open hull_hull) corollary affine_hull_open: fixes S :: "'a::real_normed_vector set" assumes "open S" "S \ {}" shows "affine hull S = UNIV" by (metis affine_hull_convex_Int_nonempty_interior assms convex_UNIV hull_UNIV inf_top.left_neutral interior_open) lemma aff_dim_convex_Int_nonempty_interior: fixes S :: "'a::euclidean_space set" shows "\convex S; S \ interior T \ {}\ \ aff_dim(S \ T) = aff_dim S" using aff_dim_affine_hull2 affine_hull_convex_Int_nonempty_interior by blast lemma aff_dim_convex_Int_open: fixes S :: "'a::euclidean_space set" shows "\convex S; open T; S \ T \ {}\ \ aff_dim(S \ T) = aff_dim S" using aff_dim_convex_Int_nonempty_interior interior_eq by blast lemma affine_hull_Diff: fixes S:: "'a::real_normed_vector set" assumes ope: "openin (top_of_set (affine hull S)) S" and "finite F" "F \ S" shows "affine hull (S - F) = affine hull S" proof - have clo: "closedin (top_of_set S) F" using assms finite_imp_closedin by auto moreover have "S - F \ {}" using assms by auto ultimately show ?thesis by (metis ope closedin_def topspace_euclidean_subtopology affine_hull_openin openin_trans) qed lemma affine_hull_halfspace_lt: fixes a :: "'a::euclidean_space" shows "affine hull {x. a \ x < r} = (if a = 0 \ r \ 0 then {} else UNIV)" using halfspace_eq_empty_lt [of a r] by (simp add: open_halfspace_lt affine_hull_open) lemma affine_hull_halfspace_le: fixes a :: "'a::euclidean_space" shows "affine hull {x. a \ x \ r} = (if a = 0 \ r < 0 then {} else UNIV)" proof (cases "a = 0") case True then show ?thesis by simp next case False then have "affine hull closure {x. a \ x < r} = UNIV" using affine_hull_halfspace_lt closure_same_affine_hull by fastforce moreover have "{x. a \ x < r} \ {x. a \ x \ r}" by (simp add: Collect_mono) ultimately show ?thesis using False antisym_conv hull_mono top_greatest by (metis affine_hull_halfspace_lt) qed lemma affine_hull_halfspace_gt: fixes a :: "'a::euclidean_space" shows "affine hull {x. a \ x > r} = (if a = 0 \ r \ 0 then {} else UNIV)" using halfspace_eq_empty_gt [of r a] by (simp add: open_halfspace_gt affine_hull_open) lemma affine_hull_halfspace_ge: fixes a :: "'a::euclidean_space" shows "affine hull {x. a \ x \ r} = (if a = 0 \ r > 0 then {} else UNIV)" using affine_hull_halfspace_le [of "-a" "-r"] by simp lemma aff_dim_halfspace_lt: fixes a :: "'a::euclidean_space" shows "aff_dim {x. a \ x < r} = (if a = 0 \ r \ 0 then -1 else DIM('a))" by simp (metis aff_dim_open halfspace_eq_empty_lt open_halfspace_lt) lemma aff_dim_halfspace_le: fixes a :: "'a::euclidean_space" shows "aff_dim {x. a \ x \ r} = (if a = 0 \ r < 0 then -1 else DIM('a))" proof - have "int (DIM('a)) = aff_dim (UNIV::'a set)" by (simp) then have "aff_dim (affine hull {x. a \ x \ r}) = DIM('a)" if "(a = 0 \ r \ 0)" using that by (simp add: affine_hull_halfspace_le not_less) then show ?thesis by (force) qed lemma aff_dim_halfspace_gt: fixes a :: "'a::euclidean_space" shows "aff_dim {x. a \ x > r} = (if a = 0 \ r \ 0 then -1 else DIM('a))" by simp (metis aff_dim_open halfspace_eq_empty_gt open_halfspace_gt) lemma aff_dim_halfspace_ge: fixes a :: "'a::euclidean_space" shows "aff_dim {x. a \ x \ r} = (if a = 0 \ r > 0 then -1 else DIM('a))" using aff_dim_halfspace_le [of "-a" "-r"] by simp proposition aff_dim_eq_hyperplane: fixes S :: "'a::euclidean_space set" shows "aff_dim S = DIM('a) - 1 \ (\a b. a \ 0 \ affine hull S = {x. a \ x = b})" (is "?lhs = ?rhs") proof (cases "S = {}") case True then show ?thesis by (auto simp: dest: hyperplane_eq_Ex) next case False then obtain c where "c \ S" by blast show ?thesis proof (cases "c = 0") case True have "?lhs \ (\a. a \ 0 \ span ((\x. x - c) ` S) = {x. a \ x = 0})" by (simp add: aff_dim_eq_dim [of c] \c \ S\ hull_inc dim_eq_hyperplane del: One_nat_def) also have "... \ ?rhs" using span_zero [of S] True \c \ S\ affine_hull_span_0 hull_inc by (fastforce simp add: affine_hull_span_gen [of c] \c = 0\) finally show ?thesis . next case False have xc_im: "x \ (+) c ` {y. a \ y = 0}" if "a \ x = a \ c" for a x proof - have "\y. a \ y = 0 \ c + y = x" by (metis that add.commute diff_add_cancel inner_commute inner_diff_left right_minus_eq) then show "x \ (+) c ` {y. a \ y = 0}" by blast qed have 2: "span ((\x. x - c) ` S) = {x. a \ x = 0}" if "(+) c ` span ((\x. x - c) ` S) = {x. a \ x = b}" for a b proof - have "b = a \ c" using span_0 that by fastforce with that have "(+) c ` span ((\x. x - c) ` S) = {x. a \ x = a \ c}" by simp then have "span ((\x. x - c) ` S) = (\x. x - c) ` {x. a \ x = a \ c}" by (metis (no_types) image_cong translation_galois uminus_add_conv_diff) also have "... = {x. a \ x = 0}" by (force simp: inner_distrib inner_diff_right intro: image_eqI [where x="x+c" for x]) finally show ?thesis . qed have "?lhs = (\a. a \ 0 \ span ((\x. x - c) ` S) = {x. a \ x = 0})" by (simp add: aff_dim_eq_dim [of c] \c \ S\ hull_inc dim_eq_hyperplane del: One_nat_def) also have "... = ?rhs" by (fastforce simp add: affine_hull_span_gen [of c] \c \ S\ hull_inc inner_distrib intro: xc_im intro!: 2) finally show ?thesis . qed qed corollary aff_dim_hyperplane [simp]: fixes a :: "'a::euclidean_space" shows "a \ 0 \ aff_dim {x. a \ x = r} = DIM('a) - 1" by (metis aff_dim_eq_hyperplane affine_hull_eq affine_hyperplane) subsection\<^marker>\tag unimportant\\Some stepping theorems\ lemma aff_dim_insert: fixes a :: "'a::euclidean_space" shows "aff_dim (insert a S) = (if a \ affine hull S then aff_dim S else aff_dim S + 1)" proof (cases "S = {}") case True then show ?thesis by simp next case False then obtain x s' where S: "S = insert x s'" "x \ s'" by (meson Set.set_insert all_not_in_conv) show ?thesis using S by (force simp add: affine_hull_insert_span_gen span_zero insert_commute [of a] aff_dim_eq_dim [of x] dim_insert) qed lemma affine_dependent_choose: fixes a :: "'a :: euclidean_space" assumes "\(affine_dependent S)" shows "affine_dependent(insert a S) \ a \ S \ a \ affine hull S" (is "?lhs = ?rhs") proof safe assume "affine_dependent (insert a S)" and "a \ S" then show "False" using \a \ S\ assms insert_absorb by fastforce next assume lhs: "affine_dependent (insert a S)" then have "a \ S" by (metis (no_types) assms insert_absorb) moreover have "finite S" using affine_independent_iff_card assms by blast moreover have "aff_dim (insert a S) \ int (card S)" using \finite S\ affine_independent_iff_card \a \ S\ lhs by fastforce ultimately show "a \ affine hull S" by (metis aff_dim_affine_independent aff_dim_insert assms) next assume "a \ S" and "a \ affine hull S" show "affine_dependent (insert a S)" by (simp add: \a \ affine hull S\ \a \ S\ affine_dependent_def) qed lemma affine_independent_insert: fixes a :: "'a :: euclidean_space" shows "\\ affine_dependent S; a \ affine hull S\ \ \ affine_dependent(insert a S)" by (simp add: affine_dependent_choose) lemma subspace_bounded_eq_trivial: fixes S :: "'a::real_normed_vector set" assumes "subspace S" shows "bounded S \ S = {0}" proof - have "False" if "bounded S" "x \ S" "x \ 0" for x proof - obtain B where B: "\y. y \ S \ norm y < B" "B > 0" using \bounded S\ by (force simp: bounded_pos_less) have "(B / norm x) *\<^sub>R x \ S" using assms subspace_mul \x \ S\ by auto moreover have "norm ((B / norm x) *\<^sub>R x) = B" using that B by (simp add: algebra_simps) ultimately show False using B by force qed then have "bounded S \ S = {0}" using assms subspace_0 by fastforce then show ?thesis by blast qed lemma affine_bounded_eq_trivial: fixes S :: "'a::real_normed_vector set" assumes "affine S" shows "bounded S \ S = {} \ (\a. S = {a})" proof (cases "S = {}") case True then show ?thesis by simp next case False then obtain b where "b \ S" by blast with False assms have "bounded S \ S = {b}" using affine_diffs_subspace [OF assms \b \ S\] by (metis (no_types, lifting) ab_group_add_class.ab_left_minus bounded_translation image_empty image_insert subspace_bounded_eq_trivial translation_invert) then show ?thesis by auto qed lemma affine_bounded_eq_lowdim: fixes S :: "'a::euclidean_space set" assumes "affine S" shows "bounded S \ aff_dim S \ 0" proof show "aff_dim S \ 0 \ bounded S" by (metis aff_dim_sing aff_dim_subset affine_dim_equal affine_sing all_not_in_conv assms bounded_empty bounded_insert dual_order.antisym empty_subsetI insert_subset) qed (use affine_bounded_eq_trivial assms in fastforce) lemma bounded_hyperplane_eq_trivial_0: fixes a :: "'a::euclidean_space" assumes "a \ 0" shows "bounded {x. a \ x = 0} \ DIM('a) = 1" proof assume "bounded {x. a \ x = 0}" then have "aff_dim {x. a \ x = 0} \ 0" by (simp add: affine_bounded_eq_lowdim affine_hyperplane) with assms show "DIM('a) = 1" by (simp add: le_Suc_eq) next assume "DIM('a) = 1" then show "bounded {x. a \ x = 0}" by (simp add: affine_bounded_eq_lowdim affine_hyperplane assms) qed lemma bounded_hyperplane_eq_trivial: fixes a :: "'a::euclidean_space" shows "bounded {x. a \ x = r} \ (if a = 0 then r \ 0 else DIM('a) = 1)" proof (simp add: bounded_hyperplane_eq_trivial_0, clarify) assume "r \ 0" "a \ 0" have "aff_dim {x. y \ x = 0} = aff_dim {x. a \ x = r}" if "y \ 0" for y::'a by (metis that \a \ 0\ aff_dim_hyperplane) then show "bounded {x. a \ x = r} = (DIM('a) = Suc 0)" by (metis One_nat_def \a \ 0\ affine_bounded_eq_lowdim affine_hyperplane bounded_hyperplane_eq_trivial_0) qed subsection\<^marker>\tag unimportant\\General case without assuming closure and getting non-strict separation\ proposition\<^marker>\tag unimportant\ separating_hyperplane_closed_point_inset: fixes S :: "'a::euclidean_space set" assumes "convex S" "closed S" "S \ {}" "z \ S" obtains a b where "a \ S" "(a - z) \ z < b" "\x. x \ S \ b < (a - z) \ x" proof - obtain y where "y \ S" and y: "\u. u \ S \ dist z y \ dist z u" using distance_attains_inf [of S z] assms by auto then have *: "(y - z) \ z < (y - z) \ z + (norm (y - z))\<^sup>2 / 2" using \y \ S\ \z \ S\ by auto show ?thesis proof (rule that [OF \y \ S\ *]) fix x assume "x \ S" have yz: "0 < (y - z) \ (y - z)" using \y \ S\ \z \ S\ by auto { assume 0: "0 < ((z - y) \ (x - y))" with any_closest_point_dot [OF \convex S\ \closed S\] have False using y \x \ S\ \y \ S\ not_less by blast } then have "0 \ ((y - z) \ (x - y))" by (force simp: not_less inner_diff_left) with yz have "0 < 2 * ((y - z) \ (x - y)) + (y - z) \ (y - z)" by (simp add: algebra_simps) then show "(y - z) \ z + (norm (y - z))\<^sup>2 / 2 < (y - z) \ x" by (simp add: field_simps inner_diff_left inner_diff_right dot_square_norm [symmetric]) qed qed lemma separating_hyperplane_closed_0_inset: fixes S :: "'a::euclidean_space set" assumes "convex S" "closed S" "S \ {}" "0 \ S" obtains a b where "a \ S" "a \ 0" "0 < b" "\x. x \ S \ a \ x > b" using separating_hyperplane_closed_point_inset [OF assms] by simp (metis \0 \ S\) proposition\<^marker>\tag unimportant\ separating_hyperplane_set_0_inspan: fixes S :: "'a::euclidean_space set" assumes "convex S" "S \ {}" "0 \ S" obtains a where "a \ span S" "a \ 0" "\x. x \ S \ 0 \ a \ x" proof - define k where [abs_def]: "k c = {x. 0 \ c \ x}" for c :: 'a have "span S \ frontier (cball 0 1) \ \f' \ {}" if f': "finite f'" "f' \ k ` S" for f' proof - obtain C where "C \ S" "finite C" and C: "f' = k ` C" using finite_subset_image [OF f'] by blast obtain a where "a \ S" "a \ 0" using \S \ {}\ \0 \ S\ ex_in_conv by blast then have "norm (a /\<^sub>R (norm a)) = 1" by simp moreover have "a /\<^sub>R (norm a) \ span S" by (simp add: \a \ S\ span_scale span_base) ultimately have ass: "a /\<^sub>R (norm a) \ span S \ sphere 0 1" by simp show ?thesis proof (cases "C = {}") case True with C ass show ?thesis by auto next case False have "closed (convex hull C)" using \finite C\ compact_eq_bounded_closed finite_imp_compact_convex_hull by auto moreover have "convex hull C \ {}" by (simp add: False) moreover have "0 \ convex hull C" by (metis \C \ S\ \convex S\ \0 \ S\ convex_hull_subset hull_same insert_absorb insert_subset) ultimately obtain a b where "a \ convex hull C" "a \ 0" "0 < b" and ab: "\x. x \ convex hull C \ a \ x > b" using separating_hyperplane_closed_0_inset by blast then have "a \ S" by (metis \C \ S\ assms(1) subsetCE subset_hull) moreover have "norm (a /\<^sub>R (norm a)) = 1" using \a \ 0\ by simp moreover have "a /\<^sub>R (norm a) \ span S" by (simp add: \a \ S\ span_scale span_base) ultimately have ass: "a /\<^sub>R (norm a) \ span S \ sphere 0 1" by simp have "\x. \a \ 0; x \ C\ \ 0 \ x \ a" using ab \0 < b\ by (metis hull_inc inner_commute less_eq_real_def less_trans) then have aa: "a /\<^sub>R (norm a) \ (\c\C. {x. 0 \ c \ x})" by (auto simp add: field_split_simps) show ?thesis unfolding C k_def using ass aa Int_iff empty_iff by force qed qed moreover have "\T. T \ k ` S \ closed T" using closed_halfspace_ge k_def by blast ultimately have "(span S \ frontier(cball 0 1)) \ (\ (k ` S)) \ {}" by (metis compact_imp_fip closed_Int_compact closed_span compact_cball compact_frontier) then show ?thesis unfolding set_eq_iff k_def by simp (metis inner_commute norm_eq_zero that zero_neq_one) qed lemma separating_hyperplane_set_point_inaff: fixes S :: "'a::euclidean_space set" assumes "convex S" "S \ {}" and zno: "z \ S" obtains a b where "(z + a) \ affine hull (insert z S)" and "a \ 0" and "a \ z \ b" and "\x. x \ S \ a \ x \ b" proof - from separating_hyperplane_set_0_inspan [of "image (\x. -z + x) S"] have "convex ((+) (- z) ` S)" using \convex S\ by simp moreover have "(+) (- z) ` S \ {}" by (simp add: \S \ {}\) moreover have "0 \ (+) (- z) ` S" using zno by auto ultimately obtain a where "a \ span ((+) (- z) ` S)" "a \ 0" and a: "\x. x \ ((+) (- z) ` S) \ 0 \ a \ x" using separating_hyperplane_set_0_inspan [of "image (\x. -z + x) S"] by blast then have szx: "\x. x \ S \ a \ z \ a \ x" by (metis (no_types, lifting) imageI inner_minus_right inner_right_distrib minus_add neg_le_0_iff_le neg_le_iff_le real_add_le_0_iff) moreover have "z + a \ affine hull insert z S" using \a \ span ((+) (- z) ` S)\ affine_hull_insert_span_gen by blast ultimately show ?thesis using \a \ 0\ szx that by auto qed proposition\<^marker>\tag unimportant\ supporting_hyperplane_rel_boundary: fixes S :: "'a::euclidean_space set" assumes "convex S" "x \ S" and xno: "x \ rel_interior S" obtains a where "a \ 0" and "\y. y \ S \ a \ x \ a \ y" and "\y. y \ rel_interior S \ a \ x < a \ y" proof - obtain a b where aff: "(x + a) \ affine hull (insert x (rel_interior S))" and "a \ 0" and "a \ x \ b" and ageb: "\u. u \ (rel_interior S) \ a \ u \ b" using separating_hyperplane_set_point_inaff [of "rel_interior S" x] assms by (auto simp: rel_interior_eq_empty convex_rel_interior) have le_ay: "a \ x \ a \ y" if "y \ S" for y proof - have con: "continuous_on (closure (rel_interior S)) ((\) a)" by (rule continuous_intros continuous_on_subset | blast)+ have y: "y \ closure (rel_interior S)" using \convex S\ closure_def convex_closure_rel_interior \y \ S\ by fastforce show ?thesis using continuous_ge_on_closure [OF con y] ageb \a \ x \ b\ by fastforce qed have 3: "a \ x < a \ y" if "y \ rel_interior S" for y proof - obtain e where "0 < e" "y \ S" and e: "cball y e \ affine hull S \ S" using \y \ rel_interior S\ by (force simp: rel_interior_cball) define y' where "y' = y - (e / norm a) *\<^sub>R ((x + a) - x)" have "y' \ cball y e" unfolding y'_def using \0 < e\ by force moreover have "y' \ affine hull S" unfolding y'_def by (metis \x \ S\ \y \ S\ \convex S\ aff affine_affine_hull hull_redundant rel_interior_same_affine_hull hull_inc mem_affine_3_minus2) ultimately have "y' \ S" using e by auto have "a \ x \ a \ y" using le_ay \a \ 0\ \y \ S\ by blast moreover have "a \ x \ a \ y" using le_ay [OF \y' \ S\] \a \ 0\ \0 < e\ not_le by (fastforce simp add: y'_def inner_diff dot_square_norm power2_eq_square) ultimately show ?thesis by force qed show ?thesis by (rule that [OF \a \ 0\ le_ay 3]) qed lemma supporting_hyperplane_relative_frontier: fixes S :: "'a::euclidean_space set" assumes "convex S" "x \ closure S" "x \ rel_interior S" obtains a where "a \ 0" and "\y. y \ closure S \ a \ x \ a \ y" and "\y. y \ rel_interior S \ a \ x < a \ y" using supporting_hyperplane_rel_boundary [of "closure S" x] by (metis assms convex_closure convex_rel_interior_closure) subsection\<^marker>\tag unimportant\\ Some results on decomposing convex hulls: intersections, simplicial subdivision\ lemma fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent(S \ T)" shows convex_hull_Int_subset: "convex hull S \ convex hull T \ convex hull (S \ T)" (is ?C) and affine_hull_Int_subset: "affine hull S \ affine hull T \ affine hull (S \ T)" (is ?A) proof - have [simp]: "finite S" "finite T" using aff_independent_finite assms by blast+ have "sum u (S \ T) = 1 \ (\v\S \ T. u v *\<^sub>R v) = (\v\S. u v *\<^sub>R v)" if [simp]: "sum u S = 1" "sum v T = 1" and eq: "(\x\T. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x)" for u v proof - define f where "f x = (if x \ S then u x else 0) - (if x \ T then v x else 0)" for x have "sum f (S \ T) = 0" by (simp add: f_def sum_Un sum_subtractf flip: sum.inter_restrict) moreover have "(\x\(S \ T). f x *\<^sub>R x) = 0" by (simp add: eq f_def sum_Un scaleR_left_diff_distrib sum_subtractf if_smult flip: sum.inter_restrict cong: if_cong) ultimately have "\v. v \ S \ T \ f v = 0" using aff_independent_finite assms unfolding affine_dependent_explicit by blast then have u [simp]: "\x. x \ S \ u x = (if x \ T then v x else 0)" by (simp add: f_def) presburger have "sum u (S \ T) = sum u S" by (simp add: sum.inter_restrict) then have "sum u (S \ T) = 1" using that by linarith moreover have "(\v\S \ T. u v *\<^sub>R v) = (\v\S. u v *\<^sub>R v)" by (auto simp: if_smult sum.inter_restrict intro: sum.cong) ultimately show ?thesis by force qed then show ?A ?C by (auto simp: convex_hull_finite affine_hull_finite) qed proposition\<^marker>\tag unimportant\ affine_hull_Int: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent(S \ T)" shows "affine hull (S \ T) = affine hull S \ affine hull T" by (simp add: affine_hull_Int_subset assms hull_mono subset_antisym) proposition\<^marker>\tag unimportant\ convex_hull_Int: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent(S \ T)" shows "convex hull (S \ T) = convex hull S \ convex hull T" by (simp add: convex_hull_Int_subset assms hull_mono subset_antisym) proposition\<^marker>\tag unimportant\ fixes S :: "'a::euclidean_space set set" assumes "\ affine_dependent (\S)" shows affine_hull_Inter: "affine hull (\S) = (\T\S. affine hull T)" (is "?A") and convex_hull_Inter: "convex hull (\S) = (\T\S. convex hull T)" (is "?C") proof - have "finite S" using aff_independent_finite assms finite_UnionD by blast then have "?A \ ?C" using assms proof (induction S rule: finite_induct) case empty then show ?case by auto next case (insert T F) then show ?case proof (cases "F={}") case True then show ?thesis by simp next case False with "insert.prems" have [simp]: "\ affine_dependent (T \ \F)" by (auto intro: affine_dependent_subset) have [simp]: "\ affine_dependent (\F)" using affine_independent_subset insert.prems by fastforce show ?thesis by (simp add: affine_hull_Int convex_hull_Int insert.IH) qed qed then show "?A" "?C" by auto qed proposition\<^marker>\tag unimportant\ in_convex_hull_exchange_unique: fixes S :: "'a::euclidean_space set" assumes naff: "\ affine_dependent S" and a: "a \ convex hull S" and S: "T \ S" "T' \ S" and x: "x \ convex hull (insert a T)" and x': "x \ convex hull (insert a T')" shows "x \ convex hull (insert a (T \ T'))" proof (cases "a \ S") case True then have "\ affine_dependent (insert a T \ insert a T')" using affine_dependent_subset assms by auto then have "x \ convex hull (insert a T \ insert a T')" by (metis IntI convex_hull_Int x x') then show ?thesis by simp next case False then have anot: "a \ T" "a \ T'" using assms by auto have [simp]: "finite S" by (simp add: aff_independent_finite assms) then obtain b where b0: "\s. s \ S \ 0 \ b s" and b1: "sum b S = 1" and aeq: "a = (\s\S. b s *\<^sub>R s)" using a by (auto simp: convex_hull_finite) have fin [simp]: "finite T" "finite T'" using assms infinite_super \finite S\ by blast+ then obtain c c' where c0: "\t. t \ insert a T \ 0 \ c t" and c1: "sum c (insert a T) = 1" and xeq: "x = (\t \ insert a T. c t *\<^sub>R t)" and c'0: "\t. t \ insert a T' \ 0 \ c' t" and c'1: "sum c' (insert a T') = 1" and x'eq: "x = (\t \ insert a T'. c' t *\<^sub>R t)" using x x' by (auto simp: convex_hull_finite) with fin anot have sumTT': "sum c T = 1 - c a" "sum c' T' = 1 - c' a" and wsumT: "(\t \ T. c t *\<^sub>R t) = x - c a *\<^sub>R a" by simp_all have wsumT': "(\t \ T'. c' t *\<^sub>R t) = x - c' a *\<^sub>R a" using x'eq fin anot by simp define cc where "cc \ \x. if x \ T then c x else 0" define cc' where "cc' \ \x. if x \ T' then c' x else 0" define dd where "dd \ \x. cc x - cc' x + (c a - c' a) * b x" have sumSS': "sum cc S = 1 - c a" "sum cc' S = 1 - c' a" unfolding cc_def cc'_def using S by (simp_all add: Int_absorb1 Int_absorb2 sum_subtractf sum.inter_restrict [symmetric] sumTT') have wsumSS: "(\t \ S. cc t *\<^sub>R t) = x - c a *\<^sub>R a" "(\t \ S. cc' t *\<^sub>R t) = x - c' a *\<^sub>R a" unfolding cc_def cc'_def using S by (simp_all add: Int_absorb1 Int_absorb2 if_smult sum.inter_restrict [symmetric] wsumT wsumT' cong: if_cong) have sum_dd0: "sum dd S = 0" unfolding dd_def using S by (simp add: sumSS' comm_monoid_add_class.sum.distrib sum_subtractf algebra_simps sum_distrib_right [symmetric] b1) have "(\v\S. (b v * x) *\<^sub>R v) = x *\<^sub>R (\v\S. b v *\<^sub>R v)" for x by (simp add: pth_5 real_vector.scale_sum_right mult.commute) then have *: "(\v\S. (b v * x) *\<^sub>R v) = x *\<^sub>R a" for x using aeq by blast have "(\v \ S. dd v *\<^sub>R v) = 0" unfolding dd_def using S by (simp add: * wsumSS sum.distrib sum_subtractf algebra_simps) then have dd0: "dd v = 0" if "v \ S" for v using naff [unfolded affine_dependent_explicit not_ex, rule_format, of S dd] using that sum_dd0 by force consider "c' a \ c a" | "c a \ c' a" by linarith then show ?thesis proof cases case 1 then have "sum cc S \ sum cc' S" by (simp add: sumSS') then have le: "cc x \ cc' x" if "x \ S" for x using dd0 [OF that] 1 b0 mult_left_mono that by (fastforce simp add: dd_def algebra_simps) have cc0: "cc x = 0" if "x \ S" "x \ T \ T'" for x using le [OF \x \ S\] that c0 by (force simp: cc_def cc'_def split: if_split_asm) show ?thesis proof (simp add: convex_hull_finite, intro exI conjI) show "\x\T \ T'. 0 \ (cc(a := c a)) x" by (simp add: c0 cc_def) show "0 \ (cc(a := c a)) a" by (simp add: c0) have "sum (cc(a := c a)) (insert a (T \ T')) = c a + sum (cc(a := c a)) (T \ T')" by (simp add: anot) also have "... = c a + sum (cc(a := c a)) S" using \T \ S\ False cc0 cc_def \a \ S\ by (fastforce intro!: sum.mono_neutral_left split: if_split_asm) also have "... = c a + (1 - c a)" by (metis \a \ S\ fun_upd_other sum.cong sumSS'(1)) finally show "sum (cc(a := c a)) (insert a (T \ T')) = 1" by simp have "(\x\insert a (T \ T'). (cc(a := c a)) x *\<^sub>R x) = c a *\<^sub>R a + (\x \ T \ T'. (cc(a := c a)) x *\<^sub>R x)" by (simp add: anot) also have "... = c a *\<^sub>R a + (\x \ S. (cc(a := c a)) x *\<^sub>R x)" using \T \ S\ False cc0 by (fastforce intro!: sum.mono_neutral_left split: if_split_asm) also have "... = c a *\<^sub>R a + x - c a *\<^sub>R a" by (simp add: wsumSS \a \ S\ if_smult sum_delta_notmem) finally show "(\x\insert a (T \ T'). (cc(a := c a)) x *\<^sub>R x) = x" by simp qed next case 2 then have "sum cc' S \ sum cc S" by (simp add: sumSS') then have le: "cc' x \ cc x" if "x \ S" for x using dd0 [OF that] 2 b0 mult_left_mono that by (fastforce simp add: dd_def algebra_simps) have cc0: "cc' x = 0" if "x \ S" "x \ T \ T'" for x using le [OF \x \ S\] that c'0 by (force simp: cc_def cc'_def split: if_split_asm) show ?thesis proof (simp add: convex_hull_finite, intro exI conjI) show "\x\T \ T'. 0 \ (cc'(a := c' a)) x" by (simp add: c'0 cc'_def) show "0 \ (cc'(a := c' a)) a" by (simp add: c'0) have "sum (cc'(a := c' a)) (insert a (T \ T')) = c' a + sum (cc'(a := c' a)) (T \ T')" by (simp add: anot) also have "... = c' a + sum (cc'(a := c' a)) S" using \T \ S\ False cc0 by (fastforce intro!: sum.mono_neutral_left split: if_split_asm) also have "... = c' a + (1 - c' a)" by (metis \a \ S\ fun_upd_other sum.cong sumSS') finally show "sum (cc'(a := c' a)) (insert a (T \ T')) = 1" by simp have "(\x\insert a (T \ T'). (cc'(a := c' a)) x *\<^sub>R x) = c' a *\<^sub>R a + (\x \ T \ T'. (cc'(a := c' a)) x *\<^sub>R x)" by (simp add: anot) also have "... = c' a *\<^sub>R a + (\x \ S. (cc'(a := c' a)) x *\<^sub>R x)" using \T \ S\ False cc0 by (fastforce intro!: sum.mono_neutral_left split: if_split_asm) also have "... = c a *\<^sub>R a + x - c a *\<^sub>R a" by (simp add: wsumSS \a \ S\ if_smult sum_delta_notmem) finally show "(\x\insert a (T \ T'). (cc'(a := c' a)) x *\<^sub>R x) = x" by simp qed qed qed corollary\<^marker>\tag unimportant\ convex_hull_exchange_Int: fixes a :: "'a::euclidean_space" assumes "\ affine_dependent S" "a \ convex hull S" "T \ S" "T' \ S" shows "(convex hull (insert a T)) \ (convex hull (insert a T')) = convex hull (insert a (T \ T'))" (is "?lhs = ?rhs") proof (rule subset_antisym) show "?lhs \ ?rhs" using in_convex_hull_exchange_unique assms by blast show "?rhs \ ?lhs" by (metis hull_mono inf_le1 inf_le2 insert_inter_insert le_inf_iff) qed lemma Int_closed_segment: fixes b :: "'a::euclidean_space" assumes "b \ closed_segment a c \ \ collinear{a,b,c}" shows "closed_segment a b \ closed_segment b c = {b}" proof (cases "c = a") case True then show ?thesis using assms collinear_3_eq_affine_dependent by fastforce next case False from assms show ?thesis proof assume "b \ closed_segment a c" moreover have "\ affine_dependent {a, c}" by (simp) ultimately show ?thesis using False convex_hull_exchange_Int [of "{a,c}" b "{a}" "{c}"] by (simp add: segment_convex_hull insert_commute) next assume ncoll: "\ collinear {a, b, c}" have False if "closed_segment a b \ closed_segment b c \ {b}" proof - have "b \ closed_segment a b" and "b \ closed_segment b c" by auto with that obtain d where "b \ d" "d \ closed_segment a b" "d \ closed_segment b c" by force then have d: "collinear {a, d, b}" "collinear {b, d, c}" by (auto simp: between_mem_segment between_imp_collinear) have "collinear {a, b, c}" by (metis (full_types) \b \ d\ collinear_3_trans d insert_commute) with ncoll show False .. qed then show ?thesis by blast qed qed lemma affine_hull_finite_intersection_hyperplanes: fixes S :: "'a::euclidean_space set" obtains \ where "finite \" "of_nat (card \) + aff_dim S = DIM('a)" "affine hull S = \\" "\h. h \ \ \ \a b. a \ 0 \ h = {x. a \ x = b}" proof - obtain b where "b \ S" and indb: "\ affine_dependent b" and eq: "affine hull S = affine hull b" using affine_basis_exists by blast obtain c where indc: "\ affine_dependent c" and "b \ c" and affc: "affine hull c = UNIV" by (metis extend_to_affine_basis affine_UNIV hull_same indb subset_UNIV) then have "finite c" by (simp add: aff_independent_finite) then have fbc: "finite b" "card b \ card c" using \b \ c\ infinite_super by (auto simp: card_mono) have imeq: "(\x. affine hull x) ` ((\a. c - {a}) ` (c - b)) = ((\a. affine hull (c - {a})) ` (c - b))" by blast have card_cb: "(card (c - b)) + aff_dim S = DIM('a)" proof - have aff: "aff_dim (UNIV::'a set) = aff_dim c" by (metis aff_dim_affine_hull affc) have "aff_dim b = aff_dim S" by (metis (no_types) aff_dim_affine_hull eq) then have "int (card b) = 1 + aff_dim S" by (simp add: aff_dim_affine_independent indb) then show ?thesis using fbc aff by (simp add: \\ affine_dependent c\ \b \ c\ aff_dim_affine_independent card_Diff_subset of_nat_diff) qed show ?thesis proof (cases "c = b") case True show ?thesis proof show "int (card {}) + aff_dim S = int DIM('a)" using True card_cb by auto show "affine hull S = \ {}" using True affc eq by blast qed auto next case False have ind: "\ affine_dependent (\a\c - b. c - {a})" by (rule affine_independent_subset [OF indc]) auto have *: "1 + aff_dim (c - {t}) = int (DIM('a))" if t: "t \ c" for t proof - have "insert t c = c" using t by blast then show ?thesis by (metis (full_types) add.commute aff_dim_affine_hull aff_dim_insert aff_dim_UNIV affc affine_dependent_def indc insert_Diff_single t) qed let ?\ = "(\x. affine hull x) ` ((\a. c - {a}) ` (c - b))" show ?thesis proof have "card ((\a. affine hull (c - {a})) ` (c - b)) = card (c - b)" proof (rule card_image) show "inj_on (\a. affine hull (c - {a})) (c - b)" unfolding inj_on_def by (metis Diff_eq_empty_iff Diff_iff indc affine_dependent_def hull_subset insert_iff) qed then show "int (card ?\) + aff_dim S = int DIM('a)" by (simp add: imeq card_cb) show "affine hull S = \ ?\" by (metis Diff_eq_empty_iff INT_simps(4) UN_singleton order_refl \b \ c\ False eq double_diff affine_hull_Inter [OF ind]) have "\a. \a \ c; a \ b\ \ aff_dim (c - {a}) = int (DIM('a) - Suc 0)" by (metis "*" DIM_ge_Suc0 One_nat_def add_diff_cancel_left' int_ops(2) of_nat_diff) then show "\h. h \ ?\ \ \a b. a \ 0 \ h = {x. a \ x = b}" by (auto simp only: One_nat_def aff_dim_eq_hyperplane [symmetric]) qed (use \finite c\ in auto) qed qed lemma affine_hyperplane_sums_eq_UNIV_0: fixes S :: "'a :: euclidean_space set" assumes "affine S" and "0 \ S" and "w \ S" and "a \ w \ 0" shows "{x + y| x y. x \ S \ a \ y = 0} = UNIV" proof - have "subspace S" by (simp add: assms subspace_affine) have span1: "span {y. a \ y = 0} \ span {x + y |x y. x \ S \ a \ y = 0}" using \0 \ S\ add.left_neutral by (intro span_mono) force have "w \ span {y. a \ y = 0}" using \a \ w \ 0\ span_induct subspace_hyperplane by auto moreover have "w \ span {x + y |x y. x \ S \ a \ y = 0}" using \w \ S\ by (metis (mono_tags, lifting) inner_zero_right mem_Collect_eq pth_d span_base) ultimately have span2: "span {y. a \ y = 0} \ span {x + y |x y. x \ S \ a \ y = 0}" by blast have "a \ 0" using assms inner_zero_left by blast then have "DIM('a) - 1 = dim {y. a \ y = 0}" by (simp add: dim_hyperplane) also have "... < dim {x + y |x y. x \ S \ a \ y = 0}" using span1 span2 by (blast intro: dim_psubset) finally have "DIM('a) - 1 < dim {x + y |x y. x \ S \ a \ y = 0}" . then have DD: "dim {x + y |x y. x \ S \ a \ y = 0} = DIM('a)" using antisym dim_subset_UNIV lowdim_subset_hyperplane not_le by fastforce have subs: "subspace {x + y| x y. x \ S \ a \ y = 0}" using subspace_sums [OF \subspace S\ subspace_hyperplane] by simp moreover have "span {x + y| x y. x \ S \ a \ y = 0} = UNIV" using DD dim_eq_full by blast ultimately show ?thesis by (simp add: subs) (metis (lifting) span_eq_iff subs) qed proposition\<^marker>\tag unimportant\ affine_hyperplane_sums_eq_UNIV: fixes S :: "'a :: euclidean_space set" assumes "affine S" and "S \ {v. a \ v = b} \ {}" and "S - {v. a \ v = b} \ {}" shows "{x + y| x y. x \ S \ a \ y = b} = UNIV" proof (cases "a = 0") case True with assms show ?thesis by (auto simp: if_splits) next case False obtain c where "c \ S" and c: "a \ c = b" using assms by force with affine_diffs_subspace [OF \affine S\] have "subspace ((+) (- c) ` S)" by blast then have aff: "affine ((+) (- c) ` S)" by (simp add: subspace_imp_affine) have 0: "0 \ (+) (- c) ` S" by (simp add: \c \ S\) obtain d where "d \ S" and "a \ d \ b" and dc: "d-c \ (+) (- c) ` S" using assms by auto then have adc: "a \ (d - c) \ 0" by (simp add: c inner_diff_right) define U where "U \ {x + y |x y. x \ (+) (- c) ` S \ a \ y = 0}" have "u + v \ (+) (c+c) ` U" if "u \ S" "b = a \ v" for u v proof show "u + v = c + c + (u + v - c - c)" by (simp add: algebra_simps) have "\x y. u + v - c - c = x + y \ (\xa\S. x = xa - c) \ a \ y = 0" proof (intro exI conjI) show "u + v - c - c = (u-c) + (v-c)" "a \ (v - c) = 0" by (simp_all add: algebra_simps c that) qed (use that in auto) then show "u + v - c - c \ U" by (auto simp: U_def image_def) qed moreover have "\a \ v = 0; u \ S\ \ \x ya. v + (u + c) = x + ya \ x \ S \ a \ ya = b" for v u by (metis add.left_commute c inner_right_distrib pth_d) ultimately have "{x + y |x y. x \ S \ a \ y = b} = (+) (c+c) ` U" by (fastforce simp: algebra_simps U_def) also have "... = range ((+) (c + c))" by (simp only: U_def affine_hyperplane_sums_eq_UNIV_0 [OF aff 0 dc adc]) also have "... = UNIV" by simp finally show ?thesis . qed lemma aff_dim_sums_Int_0: assumes "affine S" and "affine T" and "0 \ S" "0 \ T" shows "aff_dim {x + y| x y. x \ S \ y \ T} = (aff_dim S + aff_dim T) - aff_dim(S \ T)" proof - have "0 \ {x + y |x y. x \ S \ y \ T}" using assms by force then have 0: "0 \ affine hull {x + y |x y. x \ S \ y \ T}" by (metis (lifting) hull_inc) have sub: "subspace S" "subspace T" using assms by (auto simp: subspace_affine) show ?thesis using dim_sums_Int [OF sub] by (simp add: aff_dim_zero assms 0 hull_inc) qed proposition aff_dim_sums_Int: assumes "affine S" and "affine T" and "S \ T \ {}" shows "aff_dim {x + y| x y. x \ S \ y \ T} = (aff_dim S + aff_dim T) - aff_dim(S \ T)" proof - obtain a where a: "a \ S" "a \ T" using assms by force have aff: "affine ((+) (-a) ` S)" "affine ((+) (-a) ` T)" using affine_translation [symmetric, of "- a"] assms by (simp_all cong: image_cong_simp) have zero: "0 \ ((+) (-a) ` S)" "0 \ ((+) (-a) ` T)" using a assms by auto have "{x + y |x y. x \ (+) (- a) ` S \ y \ (+) (- a) ` T} = (+) (- 2 *\<^sub>R a) ` {x + y| x y. x \ S \ y \ T}" by (force simp: algebra_simps scaleR_2) moreover have "(+) (- a) ` S \ (+) (- a) ` T = (+) (- a) ` (S \ T)" by auto ultimately show ?thesis using aff_dim_sums_Int_0 [OF aff zero] aff_dim_translation_eq by (metis (lifting)) qed lemma aff_dim_affine_Int_hyperplane: fixes a :: "'a::euclidean_space" assumes "affine S" shows "aff_dim(S \ {x. a \ x = b}) = (if S \ {v. a \ v = b} = {} then - 1 else if S \ {v. a \ v = b} then aff_dim S else aff_dim S - 1)" proof (cases "a = 0") case True with assms show ?thesis by auto next case False then have "aff_dim (S \ {x. a \ x = b}) = aff_dim S - 1" if "x \ S" "a \ x \ b" and non: "S \ {v. a \ v = b} \ {}" for x proof - have [simp]: "{x + y| x y. x \ S \ a \ y = b} = UNIV" using affine_hyperplane_sums_eq_UNIV [OF assms non] that by blast show ?thesis using aff_dim_sums_Int [OF assms affine_hyperplane non] by (simp add: of_nat_diff False) qed then show ?thesis by (metis (mono_tags, lifting) inf.orderE aff_dim_empty_eq mem_Collect_eq subsetI) qed lemma aff_dim_lt_full: fixes S :: "'a::euclidean_space set" shows "aff_dim S < DIM('a) \ (affine hull S \ UNIV)" by (metis (no_types) aff_dim_affine_hull aff_dim_le_DIM aff_dim_UNIV affine_hull_UNIV less_le) lemma aff_dim_openin: fixes S :: "'a::euclidean_space set" assumes ope: "openin (top_of_set T) S" and "affine T" "S \ {}" shows "aff_dim S = aff_dim T" proof - show ?thesis proof (rule order_antisym) show "aff_dim S \ aff_dim T" by (blast intro: aff_dim_subset [OF openin_imp_subset] ope) next obtain a where "a \ S" using \S \ {}\ by blast have "S \ T" using ope openin_imp_subset by auto then have "a \ T" using \a \ S\ by auto then have subT': "subspace ((\x. - a + x) ` T)" using affine_diffs_subspace \affine T\ by auto then obtain B where Bsub: "B \ ((\x. - a + x) ` T)" and po: "pairwise orthogonal B" and eq1: "\x. x \ B \ norm x = 1" and "independent B" and cardB: "card B = dim ((\x. - a + x) ` T)" and spanB: "span B = ((\x. - a + x) ` T)" by (rule orthonormal_basis_subspace) auto obtain e where "0 < e" and e: "cball a e \ T \ S" by (meson \a \ S\ openin_contains_cball ope) have "aff_dim T = aff_dim ((\x. - a + x) ` T)" by (metis aff_dim_translation_eq) also have "... = dim ((\x. - a + x) ` T)" using aff_dim_subspace subT' by blast also have "... = card B" by (simp add: cardB) also have "... = card ((\x. e *\<^sub>R x) ` B)" using \0 < e\ by (force simp: inj_on_def card_image) also have "... \ dim ((\x. - a + x) ` S)" proof (simp, rule independent_card_le_dim) have e': "cball 0 e \ (\x. x - a) ` T \ (\x. x - a) ` S" using e by (auto simp: dist_norm norm_minus_commute subset_eq) have "(\x. e *\<^sub>R x) ` B \ cball 0 e \ (\x. x - a) ` T" using Bsub \0 < e\ eq1 subT' \a \ T\ by (auto simp: subspace_def) then show "(\x. e *\<^sub>R x) ` B \ (\x. x - a) ` S" using e' by blast have "inj_on ((*\<^sub>R) e) (span B)" using \0 < e\ inj_on_def by fastforce then show "independent ((\x. e *\<^sub>R x) ` B)" using linear_scale_self \independent B\ linear_dependent_inj_imageD by blast qed also have "... = aff_dim S" using \a \ S\ aff_dim_eq_dim hull_inc by (force cong: image_cong_simp) finally show "aff_dim T \ aff_dim S" . qed qed lemma dim_openin: fixes S :: "'a::euclidean_space set" assumes ope: "openin (top_of_set T) S" and "subspace T" "S \ {}" shows "dim S = dim T" proof (rule order_antisym) show "dim S \ dim T" by (metis ope dim_subset openin_subset topspace_euclidean_subtopology) next have "dim T = aff_dim S" using aff_dim_openin by (metis aff_dim_subspace \subspace T\ \S \ {}\ ope subspace_affine) also have "... \ dim S" by (metis aff_dim_subset aff_dim_subspace dim_span span_superset subspace_span) finally show "dim T \ dim S" by simp qed subsection\Lower-dimensional affine subsets are nowhere dense\ proposition dense_complement_subspace: fixes S :: "'a :: euclidean_space set" assumes dim_less: "dim T < dim S" and "subspace S" shows "closure(S - T) = S" proof - have "closure(S - U) = S" if "dim U < dim S" "U \ S" for U proof - have "span U \ span S" by (metis neq_iff psubsetI span_eq_dim span_mono that) then obtain a where "a \ 0" "a \ span S" and a: "\y. y \ span U \ orthogonal a y" using orthogonal_to_subspace_exists_gen by metis show ?thesis proof have "closed S" by (simp add: \subspace S\ closed_subspace) then show "closure (S - U) \ S" by (simp add: closure_minimal) show "S \ closure (S - U)" proof (clarsimp simp: closure_approachable) fix x and e::real assume "x \ S" "0 < e" show "\y\S - U. dist y x < e" proof (cases "x \ U") case True let ?y = "x + (e/2 / norm a) *\<^sub>R a" show ?thesis proof show "dist ?y x < e" using \0 < e\ by (simp add: dist_norm) next have "?y \ S" by (metis \a \ span S\ \x \ S\ assms(2) span_eq_iff subspace_add subspace_scale) moreover have "?y \ U" proof - have "e/2 / norm a \ 0" using \0 < e\ \a \ 0\ by auto then show ?thesis by (metis True \a \ 0\ a orthogonal_scaleR orthogonal_self real_vector.scale_eq_0_iff span_add_eq span_base) qed ultimately show "?y \ S - U" by blast qed next case False with \0 < e\ \x \ S\ show ?thesis by force qed qed qed qed moreover have "S - S \ T = S-T" by blast moreover have "dim (S \ T) < dim S" by (metis dim_less dim_subset inf.cobounded2 inf.orderE inf.strict_boundedE not_le) ultimately show ?thesis by force qed corollary\<^marker>\tag unimportant\ dense_complement_affine: fixes S :: "'a :: euclidean_space set" assumes less: "aff_dim T < aff_dim S" and "affine S" shows "closure(S - T) = S" proof (cases "S \ T = {}") case True then show ?thesis by (metis Diff_triv affine_hull_eq \affine S\ closure_same_affine_hull closure_subset hull_subset subset_antisym) next case False then obtain z where z: "z \ S \ T" by blast then have "subspace ((+) (- z) ` S)" by (meson IntD1 affine_diffs_subspace \affine S\) moreover have "int (dim ((+) (- z) ` T)) < int (dim ((+) (- z) ` S))" thm aff_dim_eq_dim using z less by (simp add: aff_dim_eq_dim_subtract [of z] hull_inc cong: image_cong_simp) ultimately have "closure(((+) (- z) ` S) - ((+) (- z) ` T)) = ((+) (- z) ` S)" by (simp add: dense_complement_subspace) then show ?thesis by (metis closure_translation translation_diff translation_invert) qed corollary\<^marker>\tag unimportant\ dense_complement_openin_affine_hull: fixes S :: "'a :: euclidean_space set" assumes less: "aff_dim T < aff_dim S" and ope: "openin (top_of_set (affine hull S)) S" shows "closure(S - T) = closure S" proof - have "affine hull S - T \ affine hull S" by blast then have "closure (S \ closure (affine hull S - T)) = closure (S \ (affine hull S - T))" by (rule closure_openin_Int_closure [OF ope]) then show ?thesis by (metis Int_Diff aff_dim_affine_hull affine_affine_hull dense_complement_affine hull_subset inf.orderE less) qed corollary\<^marker>\tag unimportant\ dense_complement_convex: fixes S :: "'a :: euclidean_space set" assumes "aff_dim T < aff_dim S" "convex S" shows "closure(S - T) = closure S" proof show "closure (S - T) \ closure S" by (simp add: closure_mono) have "closure (rel_interior S - T) = closure (rel_interior S)" by (simp add: assms dense_complement_openin_affine_hull openin_rel_interior rel_interior_aff_dim rel_interior_same_affine_hull) then show "closure S \ closure (S - T)" by (metis Diff_mono \convex S\ closure_mono convex_closure_rel_interior order_refl rel_interior_subset) qed corollary\<^marker>\tag unimportant\ dense_complement_convex_closed: fixes S :: "'a :: euclidean_space set" assumes "aff_dim T < aff_dim S" "convex S" "closed S" shows "closure(S - T) = S" by (simp add: assms dense_complement_convex) subsection\<^marker>\tag unimportant\\Parallel slices, etc\ text\ If we take a slice out of a set, we can do it perpendicularly, with the normal vector to the slice parallel to the affine hull.\ proposition\<^marker>\tag unimportant\ affine_parallel_slice: fixes S :: "'a :: euclidean_space set" assumes "affine S" and "S \ {x. a \ x \ b} \ {}" and "\ (S \ {x. a \ x \ b})" obtains a' b' where "a' \ 0" "S \ {x. a' \ x \ b'} = S \ {x. a \ x \ b}" "S \ {x. a' \ x = b'} = S \ {x. a \ x = b}" "\w. w \ S \ (w + a') \ S" proof (cases "S \ {x. a \ x = b} = {}") case True then obtain u v where "u \ S" "v \ S" "a \ u \ b" "a \ v > b" using assms by (auto simp: not_le) define \ where "\ = u + ((b - a \ u) / (a \ v - a \ u)) *\<^sub>R (v - u)" have "\ \ S" by (simp add: \_def \u \ S\ \v \ S\ \affine S\ mem_affine_3_minus) moreover have "a \ \ = b" using \a \ u \ b\ \b < a \ v\ by (simp add: \_def algebra_simps) (simp add: field_simps) ultimately have False using True by force then show ?thesis .. next case False then obtain z where "z \ S" and z: "a \ z = b" using assms by auto with affine_diffs_subspace [OF \affine S\] have sub: "subspace ((+) (- z) ` S)" by blast then have aff: "affine ((+) (- z) ` S)" and span: "span ((+) (- z) ` S) = ((+) (- z) ` S)" by (auto simp: subspace_imp_affine) obtain a' a'' where a': "a' \ span ((+) (- z) ` S)" and a: "a = a' + a''" and "\w. w \ span ((+) (- z) ` S) \ orthogonal a'' w" using orthogonal_subspace_decomp_exists [of "(+) (- z) ` S" "a"] by metis then have "\w. w \ S \ a'' \ (w-z) = 0" by (simp add: span_base orthogonal_def) then have a'': "\w. w \ S \ a'' \ w = (a - a') \ z" by (simp add: a inner_diff_right) then have ba'': "\w. w \ S \ a'' \ w = b - a' \ z" by (simp add: inner_diff_left z) show ?thesis proof (cases "a' = 0") case True with a assms True a'' diff_zero less_irrefl show ?thesis by auto next case False show ?thesis proof show "S \ {x. a' \ x \ a' \ z} = S \ {x. a \ x \ b}" "S \ {x. a' \ x = a' \ z} = S \ {x. a \ x = b}" by (auto simp: a ba'' inner_left_distrib) have "\w. w \ (+) (- z) ` S \ (w + a') \ (+) (- z) ` S" by (metis subspace_add a' span_eq_iff sub) then show "\w. w \ S \ (w + a') \ S" by fastforce qed (use False in auto) qed qed lemma diffs_affine_hull_span: assumes "a \ S" shows "(\x. x - a) ` (affine hull S) = span ((\x. x - a) ` S)" proof - have *: "((\x. x - a) ` (S - {a})) = ((\x. x - a) ` S) - {0}" by (auto simp: algebra_simps) show ?thesis by (auto simp add: algebra_simps affine_hull_span2 [OF assms] *) qed lemma aff_dim_dim_affine_diffs: fixes S :: "'a :: euclidean_space set" assumes "affine S" "a \ S" shows "aff_dim S = dim ((\x. x - a) ` S)" proof - obtain B where aff: "affine hull B = affine hull S" and ind: "\ affine_dependent B" and card: "of_nat (card B) = aff_dim S + 1" using aff_dim_basis_exists by blast then have "B \ {}" using assms by (metis affine_hull_eq_empty ex_in_conv) then obtain c where "c \ B" by auto then have "c \ S" by (metis aff affine_hull_eq \affine S\ hull_inc) have xy: "x - c = y - a \ y = x + 1 *\<^sub>R (a - c)" for x y c and a::'a by (auto simp: algebra_simps) have *: "(\x. x - c) ` S = (\x. x - a) ` S" using assms \c \ S\ by (auto simp: image_iff xy; metis mem_affine_3_minus pth_1) have affS: "affine hull S = S" by (simp add: \affine S\) have "aff_dim S = of_nat (card B) - 1" using card by simp also have "... = dim ((\x. x - c) ` B)" using affine_independent_card_dim_diffs [OF ind \c \ B\] by (simp add: affine_independent_card_dim_diffs [OF ind \c \ B\]) also have "... = dim ((\x. x - c) ` (affine hull B))" by (simp add: diffs_affine_hull_span \c \ B\) also have "... = dim ((\x. x - a) ` S)" by (simp add: affS aff *) finally show ?thesis . qed lemma aff_dim_linear_image_le: assumes "linear f" shows "aff_dim(f ` S) \ aff_dim S" proof - have "aff_dim (f ` T) \ aff_dim T" if "affine T" for T proof (cases "T = {}") case True then show ?thesis by (simp add: aff_dim_geq) next case False then obtain a where "a \ T" by auto have 1: "((\x. x - f a) ` f ` T) = {x - f a |x. x \ f ` T}" by auto have 2: "{x - f a| x. x \ f ` T} = f ` ((\x. x - a) ` T)" by (force simp: linear_diff [OF assms]) have "aff_dim (f ` T) = int (dim {x - f a |x. x \ f ` T})" by (simp add: \a \ T\ hull_inc aff_dim_eq_dim [of "f a"] 1 cong: image_cong_simp) also have "... = int (dim (f ` ((\x. x - a) ` T)))" by (force simp: linear_diff [OF assms] 2) also have "... \ int (dim ((\x. x - a) ` T))" by (simp add: dim_image_le [OF assms]) also have "... \ aff_dim T" by (simp add: aff_dim_dim_affine_diffs [symmetric] \a \ T\ \affine T\) finally show ?thesis . qed then have "aff_dim (f ` (affine hull S)) \ aff_dim (affine hull S)" using affine_affine_hull [of S] by blast then show ?thesis using affine_hull_linear_image assms linear_conv_bounded_linear by fastforce qed lemma aff_dim_injective_linear_image [simp]: assumes "linear f" "inj f" shows "aff_dim (f ` S) = aff_dim S" proof (rule antisym) show "aff_dim (f ` S) \ aff_dim S" by (simp add: aff_dim_linear_image_le assms(1)) next obtain g where "linear g" "g \ f = id" using assms(1) assms(2) linear_injective_left_inverse by blast then have "aff_dim S \ aff_dim(g ` f ` S)" by (simp add: image_comp) also have "... \ aff_dim (f ` S)" by (simp add: \linear g\ aff_dim_linear_image_le) finally show "aff_dim S \ aff_dim (f ` S)" . qed lemma choose_affine_subset: assumes "affine S" "-1 \ d" and dle: "d \ aff_dim S" obtains T where "affine T" "T \ S" "aff_dim T = d" proof (cases "d = -1 \ S={}") case True with assms show ?thesis by (metis aff_dim_empty affine_empty bot.extremum that eq_iff) next case False with assms obtain a where "a \ S" "0 \ d" by auto with assms have ss: "subspace ((+) (- a) ` S)" by (simp add: affine_diffs_subspace_subtract cong: image_cong_simp) have "nat d \ dim ((+) (- a) ` S)" by (metis aff_dim_subspace aff_dim_translation_eq dle nat_int nat_mono ss) then obtain T where "subspace T" and Tsb: "T \ span ((+) (- a) ` S)" and Tdim: "dim T = nat d" using choose_subspace_of_subspace [of "nat d" "(+) (- a) ` S"] by blast then have "affine T" using subspace_affine by blast then have "affine ((+) a ` T)" by (metis affine_hull_eq affine_hull_translation) moreover have "(+) a ` T \ S" proof - have "T \ (+) (- a) ` S" by (metis (no_types) span_eq_iff Tsb ss) then show "(+) a ` T \ S" using add_ac by auto qed moreover have "aff_dim ((+) a ` T) = d" by (simp add: aff_dim_subspace Tdim \0 \ d\ \subspace T\ aff_dim_translation_eq) ultimately show ?thesis by (rule that) qed subsection\Paracompactness\ proposition paracompact: fixes S :: "'a :: {metric_space,second_countable_topology} set" assumes "S \ \\" and opC: "\T. T \ \ \ open T" obtains \' where "S \ \ \'" and "\U. U \ \' \ open U \ (\T. T \ \ \ U \ T)" and "\x. x \ S \ \V. open V \ x \ V \ finite {U. U \ \' \ (U \ V \ {})}" proof (cases "S = {}") case True with that show ?thesis by blast next case False have "\T U. x \ U \ open U \ closure U \ T \ T \ \" if "x \ S" for x proof - obtain T where "x \ T" "T \ \" "open T" using assms \x \ S\ by blast then obtain e where "e > 0" "cball x e \ T" by (force simp: open_contains_cball) then show ?thesis by (meson open_ball \T \ \\ ball_subset_cball centre_in_ball closed_cball closure_minimal dual_order.trans) qed then obtain F G where Gin: "x \ G x" and oG: "open (G x)" and clos: "closure (G x) \ F x" and Fin: "F x \ \" if "x \ S" for x by metis then obtain \ where "\ \ G ` S" "countable \" "\\ = \(G ` S)" using Lindelof [of "G ` S"] by (metis image_iff) then obtain K where K: "K \ S" "countable K" and eq: "\(G ` K) = \(G ` S)" by (metis countable_subset_image) with False Gin have "K \ {}" by force then obtain a :: "nat \ 'a" where "range a = K" by (metis range_from_nat_into \countable K\) then have odif: "\n. open (F (a n) - \{closure (G (a m)) |m. m < n})" using \K \ S\ Fin opC by (fastforce simp add:) let ?C = "range (\n. F(a n) - \{closure(G(a m)) |m. m < n})" have enum_S: "\n. x \ F(a n) \ x \ G(a n)" if "x \ S" for x proof - have "\y \ K. x \ G y" using eq that Gin by fastforce then show ?thesis using clos K \range a = K\ closure_subset by blast qed show ?thesis proof show "S \ Union ?C" proof fix x assume "x \ S" define n where "n \ LEAST n. x \ F(a n)" have n: "x \ F(a n)" using enum_S [OF \x \ S\] by (force simp: n_def intro: LeastI) have notn: "x \ F(a m)" if "m < n" for m using that not_less_Least by (force simp: n_def) then have "x \ \{closure (G (a m)) |m. m < n}" using n \K \ S\ \range a = K\ clos notn by fastforce with n show "x \ Union ?C" by blast qed show "\U. U \ ?C \ open U \ (\T. T \ \ \ U \ T)" using Fin \K \ S\ \range a = K\ by (auto simp: odif) show "\V. open V \ x \ V \ finite {U. U \ ?C \ (U \ V \ {})}" if "x \ S" for x proof - obtain n where n: "x \ F(a n)" "x \ G(a n)" using \x \ S\ enum_S by auto have "{U \ ?C. U \ G (a n) \ {}} \ (\n. F(a n) - \{closure(G(a m)) |m. m < n}) ` atMost n" proof clarsimp fix k assume "(F (a k) - \{closure (G (a m)) |m. m < k}) \ G (a n) \ {}" then have "k \ n" by auto (metis closure_subset not_le subsetCE) then show "F (a k) - \{closure (G (a m)) |m. m < k} \ (\n. F (a n) - \{closure (G (a m)) |m. m < n}) ` {..n}" by force qed moreover have "finite ((\n. F(a n) - \{closure(G(a m)) |m. m < n}) ` atMost n)" by force ultimately have *: "finite {U \ ?C. U \ G (a n) \ {}}" using finite_subset by blast have "a n \ S" using \K \ S\ \range a = K\ by blast then show ?thesis by (blast intro: oG n *) qed qed qed corollary paracompact_closedin: fixes S :: "'a :: {metric_space,second_countable_topology} set" assumes cin: "closedin (top_of_set U) S" and oin: "\T. T \ \ \ openin (top_of_set U) T" and "S \ \\" obtains \' where "S \ \ \'" and "\V. V \ \' \ openin (top_of_set U) V \ (\T. T \ \ \ V \ T)" and "\x. x \ U \ \V. openin (top_of_set U) V \ x \ V \ finite {X. X \ \' \ (X \ V \ {})}" proof - have "\Z. open Z \ (T = U \ Z)" if "T \ \" for T using oin [OF that] by (auto simp: openin_open) then obtain F where opF: "open (F T)" and intF: "U \ F T = T" if "T \ \" for T by metis obtain K where K: "closed K" "U \ K = S" using cin by (auto simp: closedin_closed) have 1: "U \ \(insert (- K) (F ` \))" by clarsimp (metis Int_iff Union_iff \U \ K = S\ \S \ \\\ subsetD intF) have 2: "\T. T \ insert (- K) (F ` \) \ open T" using \closed K\ by (auto simp: opF) obtain \ where "U \ \\" and D1: "\U. U \ \ \ open U \ (\T. T \ insert (- K) (F ` \) \ U \ T)" and D2: "\x. x \ U \ \V. open V \ x \ V \ finite {U \ \. U \ V \ {}}" by (blast intro: paracompact [OF 1 2]) let ?C = "{U \ V |V. V \ \ \ (V \ K \ {})}" show ?thesis proof (rule_tac \' = "{U \ V |V. V \ \ \ (V \ K \ {})}" in that) show "S \ \?C" using \U \ K = S\ \U \ \\\ K by (blast dest!: subsetD) show "\V. V \ ?C \ openin (top_of_set U) V \ (\T. T \ \ \ V \ T)" using D1 intF by fastforce have *: "{X. (\V. X = U \ V \ V \ \ \ V \ K \ {}) \ X \ (U \ V) \ {}} \ (\x. U \ x) ` {U \ \. U \ V \ {}}" for V by blast show "\V. openin (top_of_set U) V \ x \ V \ finite {X \ ?C. X \ V \ {}}" if "x \ U" for x proof - from D2 [OF that] obtain V where "open V" "x \ V" "finite {U \ \. U \ V \ {}}" by auto with * show ?thesis by (rule_tac x="U \ V" in exI) (auto intro: that finite_subset [OF *]) qed qed qed corollary\<^marker>\tag unimportant\ paracompact_closed: fixes S :: "'a :: {metric_space,second_countable_topology} set" assumes "closed S" and opC: "\T. T \ \ \ open T" and "S \ \\" obtains \' where "S \ \\'" and "\U. U \ \' \ open U \ (\T. T \ \ \ U \ T)" and "\x. \V. open V \ x \ V \ finite {U. U \ \' \ (U \ V \ {})}" by (rule paracompact_closedin [of UNIV S \]) (auto simp: assms) subsection\<^marker>\tag unimportant\\Closed-graph characterization of continuity\ lemma continuous_closed_graph_gen: fixes T :: "'b::real_normed_vector set" assumes contf: "continuous_on S f" and fim: "f ` S \ T" shows "closedin (top_of_set (S \ T)) ((\x. Pair x (f x)) ` S)" proof - have eq: "((\x. Pair x (f x)) ` S) = (S \ T \ (\z. (f \ fst)z - snd z) -` {0})" using fim by auto show ?thesis unfolding eq by (intro continuous_intros continuous_closedin_preimage continuous_on_subset [OF contf]) auto qed lemma continuous_closed_graph_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "compact T" and fim: "f ` S \ T" shows "continuous_on S f \ closedin (top_of_set (S \ T)) ((\x. Pair x (f x)) ` S)" (is "?lhs = ?rhs") proof - have "?lhs" if ?rhs proof (clarsimp simp add: continuous_on_closed_gen [OF fim]) fix U assume U: "closedin (top_of_set T) U" have eq: "(S \ f -` U) = fst ` (((\x. Pair x (f x)) ` S) \ (S \ U))" by (force simp: image_iff) show "closedin (top_of_set S) (S \ f -` U)" by (simp add: U closedin_Int closedin_Times closed_map_fst [OF \compact T\] that eq) qed with continuous_closed_graph_gen assms show ?thesis by blast qed lemma continuous_closed_graph: fixes f :: "'a::topological_space \ 'b::real_normed_vector" assumes "closed S" and contf: "continuous_on S f" shows "closed ((\x. Pair x (f x)) ` S)" proof (rule closedin_closed_trans) show "closedin (top_of_set (S \ UNIV)) ((\x. (x, f x)) ` S)" by (rule continuous_closed_graph_gen [OF contf subset_UNIV]) qed (simp add: \closed S\ closed_Times) lemma continuous_from_closed_graph: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "compact T" and fim: "f ` S \ T" and clo: "closed ((\x. Pair x (f x)) ` S)" shows "continuous_on S f" using fim clo by (auto intro: closed_subset simp: continuous_closed_graph_eq [OF \compact T\ fim]) lemma continuous_on_Un_local_open: assumes opS: "openin (top_of_set (S \ T)) S" and opT: "openin (top_of_set (S \ T)) T" and contf: "continuous_on S f" and contg: "continuous_on T f" shows "continuous_on (S \ T) f" using pasting_lemma [of "{S,T}" "top_of_set (S \ T)" id euclidean "\i. f" f] contf contg opS opT by (simp add: subtopology_subtopology) (metis inf.absorb2 openin_imp_subset) lemma continuous_on_cases_local_open: assumes opS: "openin (top_of_set (S \ T)) S" and opT: "openin (top_of_set (S \ T)) T" and contf: "continuous_on S f" and contg: "continuous_on T g" and fg: "\x. x \ S \ \P x \ x \ T \ P x \ f x = g x" shows "continuous_on (S \ T) (\x. if P x then f x else g x)" proof - have "\x. x \ S \ (if P x then f x else g x) = f x" "\x. x \ T \ (if P x then f x else g x) = g x" by (simp_all add: fg) then have "continuous_on S (\x. if P x then f x else g x)" "continuous_on T (\x. if P x then f x else g x)" by (simp_all add: contf contg cong: continuous_on_cong) then show ?thesis by (rule continuous_on_Un_local_open [OF opS opT]) qed subsection\<^marker>\tag unimportant\\The union of two collinear segments is another segment\ proposition\<^marker>\tag unimportant\ in_convex_hull_exchange: fixes a :: "'a::euclidean_space" assumes a: "a \ convex hull S" and xS: "x \ convex hull S" obtains b where "b \ S" "x \ convex hull (insert a (S - {b}))" proof (cases "a \ S") case True with xS insert_Diff that show ?thesis by fastforce next case False show ?thesis proof (cases "finite S \ card S \ Suc (DIM('a))") case True then obtain u where u0: "\i. i \ S \ 0 \ u i" and u1: "sum u S = 1" and ua: "(\i\S. u i *\<^sub>R i) = a" using a by (auto simp: convex_hull_finite) obtain v where v0: "\i. i \ S \ 0 \ v i" and v1: "sum v S = 1" and vx: "(\i\S. v i *\<^sub>R i) = x" using True xS by (auto simp: convex_hull_finite) show ?thesis proof (cases "\b. b \ S \ v b = 0") case True then obtain b where b: "b \ S" "v b = 0" by blast show ?thesis proof have fin: "finite (insert a (S - {b}))" using sum.infinite v1 by fastforce show "x \ convex hull insert a (S - {b})" unfolding convex_hull_finite [OF fin] mem_Collect_eq proof (intro conjI exI ballI) have "(\x \ insert a (S - {b}). if x = a then 0 else v x) = (\x \ S - {b}. if x = a then 0 else v x)" using fin by (force intro: sum.mono_neutral_right) also have "... = (\x \ S - {b}. v x)" using b False by (auto intro!: sum.cong split: if_split_asm) also have "... = (\x\S. v x)" by (metis \v b = 0\ diff_zero sum.infinite sum_diff1 u1 zero_neq_one) finally show "(\x\insert a (S - {b}). if x = a then 0 else v x) = 1" by (simp add: v1) show "\x. x \ insert a (S - {b}) \ 0 \ (if x = a then 0 else v x)" by (auto simp: v0) have "(\x \ insert a (S - {b}). (if x = a then 0 else v x) *\<^sub>R x) = (\x \ S - {b}. (if x = a then 0 else v x) *\<^sub>R x)" using fin by (force intro: sum.mono_neutral_right) also have "... = (\x \ S - {b}. v x *\<^sub>R x)" using b False by (auto intro!: sum.cong split: if_split_asm) also have "... = (\x\S. v x *\<^sub>R x)" by (metis (no_types, lifting) b(2) diff_zero fin finite.emptyI finite_Diff2 finite_insert scale_eq_0_iff sum_diff1) finally show "(\x\insert a (S - {b}). (if x = a then 0 else v x) *\<^sub>R x) = x" by (simp add: vx) qed qed (rule \b \ S\) next case False have le_Max: "u i / v i \ Max ((\i. u i / v i) ` S)" if "i \ S" for i by (simp add: True that) have "Max ((\i. u i / v i) ` S) \ (\i. u i / v i) ` S" using True v1 by (auto intro: Max_in) then obtain b where "b \ S" and beq: "Max ((\b. u b / v b) ` S) = u b / v b" by blast then have "0 \ u b / v b" using le_Max beq divide_le_0_iff le_numeral_extra(2) sum_nonpos u1 by (metis False eq_iff v0) then have "0 < u b" "0 < v b" using False \b \ S\ u0 v0 by force+ have fin: "finite (insert a (S - {b}))" using sum.infinite v1 by fastforce show ?thesis proof show "x \ convex hull insert a (S - {b})" unfolding convex_hull_finite [OF fin] mem_Collect_eq proof (intro conjI exI ballI) have "(\x \ insert a (S - {b}). if x=a then v b / u b else v x - (v b / u b) * u x) = v b / u b + (\x \ S - {b}. v x - (v b / u b) * u x)" using \a \ S\ \b \ S\ True by (auto intro!: sum.cong split: if_split_asm) also have "... = v b / u b + (\x \ S - {b}. v x) - (v b / u b) * (\x \ S - {b}. u x)" by (simp add: Groups_Big.sum_subtractf sum_distrib_left) also have "... = (\x\S. v x)" using \0 < u b\ True by (simp add: Groups_Big.sum_diff1 u1 field_simps) finally show "sum (\x. if x=a then v b / u b else v x - (v b / u b) * u x) (insert a (S - {b})) = 1" by (simp add: v1) show "0 \ (if i = a then v b / u b else v i - v b / u b * u i)" if "i \ insert a (S - {b})" for i using \0 < u b\ \0 < v b\ v0 [of i] le_Max [of i] beq that False by (auto simp: field_simps split: if_split_asm) have "(\x\insert a (S - {b}). (if x=a then v b / u b else v x - v b / u b * u x) *\<^sub>R x) = (v b / u b) *\<^sub>R a + (\x\S - {b}. (v x - v b / u b * u x) *\<^sub>R x)" using \a \ S\ \b \ S\ True by (auto intro!: sum.cong split: if_split_asm) also have "... = (v b / u b) *\<^sub>R a + (\x \ S - {b}. v x *\<^sub>R x) - (v b / u b) *\<^sub>R (\x \ S - {b}. u x *\<^sub>R x)" by (simp add: Groups_Big.sum_subtractf scaleR_left_diff_distrib sum_distrib_left scale_sum_right) also have "... = (\x\S. v x *\<^sub>R x)" using \0 < u b\ True by (simp add: ua vx Groups_Big.sum_diff1 algebra_simps) finally show "(\x\insert a (S - {b}). (if x=a then v b / u b else v x - v b / u b * u x) *\<^sub>R x) = x" by (simp add: vx) qed qed (rule \b \ S\) qed next case False obtain T where "finite T" "T \ S" and caT: "card T \ Suc (DIM('a))" and xT: "x \ convex hull T" using xS by (auto simp: caratheodory [of S]) with False obtain b where b: "b \ S" "b \ T" by (metis antisym subsetI) show ?thesis proof show "x \ convex hull insert a (S - {b})" using \T \ S\ b by (blast intro: subsetD [OF hull_mono xT]) qed (rule \b \ S\) qed qed lemma convex_hull_exchange_Union: fixes a :: "'a::euclidean_space" assumes "a \ convex hull S" shows "convex hull S = (\b \ S. convex hull (insert a (S - {b})))" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (blast intro: in_convex_hull_exchange [OF assms]) show "?rhs \ ?lhs" proof clarify fix x b assume"b \ S" "x \ convex hull insert a (S - {b})" then show "x \ convex hull S" if "b \ S" by (metis (no_types) that assms order_refl hull_mono hull_redundant insert_Diff_single insert_subset subsetCE) qed qed lemma Un_closed_segment: fixes a :: "'a::euclidean_space" assumes "b \ closed_segment a c" shows "closed_segment a b \ closed_segment b c = closed_segment a c" proof (cases "c = a") case True with assms show ?thesis by simp next case False with assms have "convex hull {a, b} \ convex hull {b, c} = (\ba\{a, c}. convex hull insert b ({a, c} - {ba}))" by (auto simp: insert_Diff_if insert_commute) then show ?thesis using convex_hull_exchange_Union by (metis assms segment_convex_hull) qed lemma Un_open_segment: fixes a :: "'a::euclidean_space" assumes "b \ open_segment a c" shows "open_segment a b \ {b} \ open_segment b c = open_segment a c" (is "?lhs = ?rhs") proof - have b: "b \ closed_segment a c" by (simp add: assms open_closed_segment) have *: "?rhs \ insert b (open_segment a b \ open_segment b c)" if "{b,c,a} \ open_segment a b \ open_segment b c = {c,a} \ ?rhs" proof - have "insert a (insert c (insert b (open_segment a b \ open_segment b c))) = insert a (insert c (?rhs))" using that by (simp add: insert_commute) then show ?thesis by (metis (no_types) Diff_cancel Diff_eq_empty_iff Diff_insert2 open_segment_def) qed show ?thesis proof show "?lhs \ ?rhs" by (simp add: assms b subset_open_segment) show "?rhs \ ?lhs" using Un_closed_segment [OF b] * by (simp add: closed_segment_eq_open insert_commute) qed qed subsection\Covering an open set by a countable chain of compact sets\ proposition open_Union_compact_subsets: fixes S :: "'a::euclidean_space set" assumes "open S" obtains C where "\n. compact(C n)" "\n. C n \ S" "\n. C n \ interior(C(Suc n))" "\(range C) = S" "\K. \compact K; K \ S\ \ \N. \n\N. K \ (C n)" proof (cases "S = {}") case True then show ?thesis by (rule_tac C = "\n. {}" in that) auto next case False then obtain a where "a \ S" by auto let ?C = "\n. cball a (real n) - (\x \ -S. \e \ ball 0 (1 / real(Suc n)). {x + e})" have "\N. \n\N. K \ (f n)" if "\n. compact(f n)" and sub_int: "\n. f n \ interior (f(Suc n))" and eq: "\(range f) = S" and "compact K" "K \ S" for f K proof - have *: "\n. f n \ (\n. interior (f n))" by (meson Sup_upper2 UNIV_I \\n. f n \ interior (f (Suc n))\ image_iff) have mono: "\m n. m \ n \f m \ f n" by (meson dual_order.trans interior_subset lift_Suc_mono_le sub_int) obtain I where "finite I" and I: "K \ (\i\I. interior (f i))" proof (rule compactE_image [OF \compact K\]) show "K \ (\n. interior (f n))" using \K \ S\ \\(f ` UNIV) = S\ * by blast qed auto { fix n assume n: "Max I \ n" have "(\i\I. interior (f i)) \ f n" by (rule UN_least) (meson dual_order.trans interior_subset mono I Max_ge [OF \finite I\] n) then have "K \ f n" using I by auto } then show ?thesis by blast qed moreover have "\f. (\n. compact(f n)) \ (\n. (f n) \ S) \ (\n. (f n) \ interior(f(Suc n))) \ ((\(range f) = S))" proof (intro exI conjI allI) show "\n. compact (?C n)" by (auto simp: compact_diff open_sums) show "\n. ?C n \ S" by auto show "?C n \ interior (?C (Suc n))" for n proof (simp add: interior_diff, rule Diff_mono) show "cball a (real n) \ ball a (1 + real n)" by (simp add: cball_subset_ball_iff) have cl: "closed (\x\- S. \e\cball 0 (1 / (2 + real n)). {x + e})" using assms by (auto intro: closed_compact_sums) have "closure (\x\- S. \y\ball 0 (1 / (2 + real n)). {x + y}) \ (\x \ -S. \e \ cball 0 (1 / (2 + real n)). {x + e})" by (intro closure_minimal UN_mono ball_subset_cball order_refl cl) also have "... \ (\x \ -S. \y\ball 0 (1 / (1 + real n)). {x + y})" by (simp add: cball_subset_ball_iff field_split_simps UN_mono) finally show "closure (\x\- S. \y\ball 0 (1 / (2 + real n)). {x + y}) \ (\x \ -S. \y\ball 0 (1 / (1 + real n)). {x + y})" . qed have "S \ \ (range ?C)" proof fix x assume x: "x \ S" then obtain e where "e > 0" and e: "ball x e \ S" using assms open_contains_ball by blast then obtain N1 where "N1 > 0" and N1: "real N1 > 1/e" using reals_Archimedean2 by (metis divide_less_0_iff less_eq_real_def neq0_conv not_le of_nat_0 of_nat_1 of_nat_less_0_iff) obtain N2 where N2: "norm(x - a) \ real N2" by (meson real_arch_simple) have N12: "inverse((N1 + N2) + 1) \ inverse(N1)" using \N1 > 0\ by (auto simp: field_split_simps) have "x \ y + z" if "y \ S" "norm z < 1 / (1 + (real N1 + real N2))" for y z proof - have "e * real N1 < e * (1 + (real N1 + real N2))" by (simp add: \0 < e\) then have "1 / (1 + (real N1 + real N2)) < e" using N1 \e > 0\ by (metis divide_less_eq less_trans mult.commute of_nat_add of_nat_less_0_iff of_nat_Suc) then have "x - z \ ball x e" using that by simp then have "x - z \ S" using e by blast with that show ?thesis by auto qed with N2 show "x \ \ (range ?C)" by (rule_tac a = "N1+N2" in UN_I) (auto simp: dist_norm norm_minus_commute) qed then show "\ (range ?C) = S" by auto qed ultimately show ?thesis using that by metis qed subsection\Orthogonal complement\ definition\<^marker>\tag important\ orthogonal_comp ("_\<^sup>\" [80] 80) where "orthogonal_comp W \ {x. \y \ W. orthogonal y x}" proposition subspace_orthogonal_comp: "subspace (W\<^sup>\)" unfolding subspace_def orthogonal_comp_def orthogonal_def by (auto simp: inner_right_distrib) lemma orthogonal_comp_anti_mono: assumes "A \ B" shows "B\<^sup>\ \ A\<^sup>\" proof fix x assume x: "x \ B\<^sup>\" show "x \ orthogonal_comp A" using x unfolding orthogonal_comp_def by (simp add: orthogonal_def, metis assms in_mono) qed lemma orthogonal_comp_null [simp]: "{0}\<^sup>\ = UNIV" by (auto simp: orthogonal_comp_def orthogonal_def) lemma orthogonal_comp_UNIV [simp]: "UNIV\<^sup>\ = {0}" unfolding orthogonal_comp_def orthogonal_def by auto (use inner_eq_zero_iff in blast) lemma orthogonal_comp_subset: "U \ U\<^sup>\\<^sup>\" by (auto simp: orthogonal_comp_def orthogonal_def inner_commute) lemma subspace_sum_minimal: assumes "S \ U" "T \ U" "subspace U" shows "S + T \ U" proof fix x assume "x \ S + T" then obtain xs xt where "xs \ S" "xt \ T" "x = xs+xt" by (meson set_plus_elim) then show "x \ U" by (meson assms subsetCE subspace_add) qed proposition subspace_sum_orthogonal_comp: fixes U :: "'a :: euclidean_space set" assumes "subspace U" shows "U + U\<^sup>\ = UNIV" proof - obtain B where "B \ U" and ortho: "pairwise orthogonal B" "\x. x \ B \ norm x = 1" and "independent B" "card B = dim U" "span B = U" using orthonormal_basis_subspace [OF assms] by metis then have "finite B" by (simp add: indep_card_eq_dim_span) have *: "\x\B. \y\B. x \ y = (if x=y then 1 else 0)" using ortho norm_eq_1 by (auto simp: orthogonal_def pairwise_def) { fix v let ?u = "\b\B. (v \ b) *\<^sub>R b" have "v = ?u + (v - ?u)" by simp moreover have "?u \ U" by (metis (no_types, lifting) \span B = U\ assms subspace_sum span_base span_mul) moreover have "(v - ?u) \ U\<^sup>\" proof (clarsimp simp: orthogonal_comp_def orthogonal_def) fix y assume "y \ U" with \span B = U\ span_finite [OF \finite B\] obtain u where u: "y = (\b\B. u b *\<^sub>R b)" by auto have "b \ (v - ?u) = 0" if "b \ B" for b using that \finite B\ by (simp add: * algebra_simps inner_sum_right if_distrib [of "(*)v" for v] inner_commute cong: if_cong) then show "y \ (v - ?u) = 0" by (simp add: u inner_sum_left) qed ultimately have "v \ U + U\<^sup>\" using set_plus_intro by fastforce } then show ?thesis by auto qed lemma orthogonal_Int_0: assumes "subspace U" shows "U \ U\<^sup>\ = {0}" using orthogonal_comp_def orthogonal_self by (force simp: assms subspace_0 subspace_orthogonal_comp) lemma orthogonal_comp_self: fixes U :: "'a :: euclidean_space set" assumes "subspace U" shows "U\<^sup>\\<^sup>\ = U" proof have ssU': "subspace (U\<^sup>\)" by (simp add: subspace_orthogonal_comp) have "u \ U" if "u \ U\<^sup>\\<^sup>\" for u proof - obtain v w where "u = v+w" "v \ U" "w \ U\<^sup>\" using subspace_sum_orthogonal_comp [OF assms] set_plus_elim by blast then have "u-v \ U\<^sup>\" by simp moreover have "v \ U\<^sup>\\<^sup>\" using \v \ U\ orthogonal_comp_subset by blast then have "u-v \ U\<^sup>\\<^sup>\" by (simp add: subspace_diff subspace_orthogonal_comp that) ultimately have "u-v = 0" using orthogonal_Int_0 ssU' by blast with \v \ U\ show ?thesis by auto qed then show "U\<^sup>\\<^sup>\ \ U" by auto qed (use orthogonal_comp_subset in auto) lemma ker_orthogonal_comp_adjoint: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" shows "f -` {0} = (range (adjoint f))\<^sup>\" proof - have "\x. \\y. y \ f x = 0\ \ f x = 0" using assms inner_commute all_zero_iff by metis then show ?thesis using assms by (auto simp: orthogonal_comp_def orthogonal_def adjoint_works inner_commute) qed subsection\<^marker>\tag unimportant\ \A non-injective linear function maps into a hyperplane.\ lemma linear_surj_adj_imp_inj: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" "surj (adjoint f)" shows "inj f" proof - have "\x. y = adjoint f x" for y using assms by (simp add: surjD) then show "inj f" using assms unfolding inj_on_def image_def by (metis (no_types) adjoint_works euclidean_eqI) qed \ \\<^url>\https://mathonline.wikidot.com/injectivity-and-surjectivity-of-the-adjoint-of-a-linear-map\\ lemma surj_adjoint_iff_inj [simp]: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" shows "surj (adjoint f) \ inj f" proof assume "surj (adjoint f)" then show "inj f" by (simp add: assms linear_surj_adj_imp_inj) next assume "inj f" have "f -` {0} = {0}" using assms \inj f\ linear_0 linear_injective_0 by fastforce moreover have "f -` {0} = range (adjoint f)\<^sup>\" by (intro ker_orthogonal_comp_adjoint assms) ultimately have "range (adjoint f)\<^sup>\\<^sup>\ = UNIV" by (metis orthogonal_comp_null) then show "surj (adjoint f)" using adjoint_linear \linear f\ by (subst (asm) orthogonal_comp_self) (simp add: adjoint_linear linear_subspace_image) qed lemma inj_adjoint_iff_surj [simp]: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" shows "inj (adjoint f) \ surj f" proof assume "inj (adjoint f)" have "(adjoint f) -` {0} = {0}" by (metis \inj (adjoint f)\ adjoint_linear assms surj_adjoint_iff_inj ker_orthogonal_comp_adjoint orthogonal_comp_UNIV) then have "(range(f))\<^sup>\ = {0}" by (metis (no_types, opaque_lifting) adjoint_adjoint adjoint_linear assms ker_orthogonal_comp_adjoint set_zero) then show "surj f" by (metis \inj (adjoint f)\ adjoint_adjoint adjoint_linear assms surj_adjoint_iff_inj) next assume "surj f" then have "range f = (adjoint f -` {0})\<^sup>\" by (simp add: adjoint_adjoint adjoint_linear assms ker_orthogonal_comp_adjoint) then have "{0} = adjoint f -` {0}" using \surj f\ adjoint_adjoint adjoint_linear assms ker_orthogonal_comp_adjoint by force then show "inj (adjoint f)" by (simp add: \surj f\ adjoint_adjoint adjoint_linear assms linear_surj_adj_imp_inj) qed lemma linear_singular_into_hyperplane: fixes f :: "'n::euclidean_space \ 'n" assumes "linear f" shows "\ inj f \ (\a. a \ 0 \ (\x. a \ f x = 0))" (is "_ = ?rhs") proof assume "\inj f" then show ?rhs using all_zero_iff by (metis (no_types, opaque_lifting) adjoint_clauses(2) adjoint_linear assms linear_injective_0 linear_injective_imp_surjective linear_surj_adj_imp_inj) next assume ?rhs then show "\inj f" by (metis assms linear_injective_isomorphism all_zero_iff) qed lemma linear_singular_image_hyperplane: fixes f :: "'n::euclidean_space \ 'n" assumes "linear f" "\inj f" obtains a where "a \ 0" "\S. f ` S \ {x. a \ x = 0}" using assms by (fastforce simp add: linear_singular_into_hyperplane) end diff --git a/src/HOL/Complex_Analysis/Conformal_Mappings.thy b/src/HOL/Complex_Analysis/Conformal_Mappings.thy --- a/src/HOL/Complex_Analysis/Conformal_Mappings.thy +++ b/src/HOL/Complex_Analysis/Conformal_Mappings.thy @@ -1,1894 +1,1926 @@ section \Conformal Mappings and Consequences of Cauchy's Integral Theorem\ text\By John Harrison et al. Ported from HOL Light by L C Paulson (2016)\ text\Also Cauchy's residue theorem by Wenda Li (2016)\ theory Conformal_Mappings imports Cauchy_Integral_Formula begin subsection \Analytic continuation\ proposition isolated_zeros: assumes holf: "f holomorphic_on S" and "open S" "connected S" "\ \ S" "f \ = 0" "\ \ S" "f \ \ 0" obtains r where "0 < r" and "ball \ r \ S" and "\z. z \ ball \ r - {\} \ f z \ 0" proof - obtain r where "0 < r" and r: "ball \ r \ S" using \open S\ \\ \ S\ open_contains_ball_eq by blast have powf: "((\n. (deriv ^^ n) f \ / (fact n) * (z - \)^n) sums f z)" if "z \ ball \ r" for z by (intro holomorphic_power_series [OF _ that] holomorphic_on_subset [OF holf r]) obtain m where m: "(deriv ^^ m) f \ / (fact m) \ 0" using holomorphic_fun_eq_0_on_connected [OF holf \open S\ \connected S\ _ \\ \ S\ \\ \ S\] \f \ \ 0\ by auto then have "m \ 0" using assms(5) funpow_0 by fastforce obtain s where "0 < s" and s: "\z. z \ cball \ s - {\} \ f z \ 0" using powser_0_nonzero [OF \0 < r\ powf \f \ = 0\ m] by (metis \m \ 0\ dist_norm mem_ball norm_minus_commute not_gr_zero) have "0 < min r s" by (simp add: \0 < r\ \0 < s\) then show thesis apply (rule that) using r s by auto qed proposition analytic_continuation: assumes holf: "f holomorphic_on S" and "open S" and "connected S" and "U \ S" and "\ \ S" and "\ islimpt U" and fU0 [simp]: "\z. z \ U \ f z = 0" and "w \ S" shows "f w = 0" proof - obtain e where "0 < e" and e: "cball \ e \ S" using \open S\ \\ \ S\ open_contains_cball_eq by blast define T where "T = cball \ e \ U" have contf: "continuous_on (closure T) f" by (metis T_def closed_cball closure_minimal e holf holomorphic_on_imp_continuous_on holomorphic_on_subset inf.cobounded1) have fT0 [simp]: "\x. x \ T \ f x = 0" by (simp add: T_def) have "\r. \\e>0. \x'\U. x' \ \ \ dist x' \ < e; 0 < r\ \ \x'\cball \ e \ U. x' \ \ \ dist x' \ < r" by (metis \0 < e\ IntI dist_commute less_eq_real_def mem_cball min_less_iff_conj) then have "\ islimpt T" using \\ islimpt U\ by (auto simp: T_def islimpt_approachable) then have "\ \ closure T" by (simp add: closure_def) then have "f \ = 0" by (auto simp: continuous_constant_on_closure [OF contf]) moreover have "\r. \0 < r; \z. z \ ball \ r - {\} \ f z \ 0\ \ False" by (metis open_ball \\ islimpt T\ centre_in_ball fT0 insertE insert_Diff islimptE) ultimately show ?thesis by (metis \open S\ \connected S\ \\ \ S\ \w \ S\ holf isolated_zeros) qed corollary analytic_continuation_open: assumes "open s" and "open s'" and "s \ {}" and "connected s'" and "s \ s'" assumes "f holomorphic_on s'" and "g holomorphic_on s'" and "\z. z \ s \ f z = g z" assumes "z \ s'" shows "f z = g z" proof - from \s \ {}\ obtain \ where "\ \ s" by auto with \open s\ have \: "\ islimpt s" by (intro interior_limit_point) (auto simp: interior_open) have "f z - g z = 0" by (rule analytic_continuation[of "\z. f z - g z" s' s \]) (insert assms \\ \ s\ \, auto intro: holomorphic_intros) thus ?thesis by simp qed +corollary analytic_continuation': + assumes "f holomorphic_on S" "open S" "connected S" + and "U \ S" "\ \ S" "\ islimpt U" + and "f constant_on U" + shows "f constant_on S" +proof - + obtain c where c: "\x. x \ U \ f x - c = 0" + by (metis \f constant_on U\ constant_on_def diff_self) + have "(\z. f z - c) holomorphic_on S" + using assms by (intro holomorphic_intros) + with c analytic_continuation assms have "\x. x \ S \ f x - c = 0" + by blast + then show ?thesis + unfolding constant_on_def by force +qed + +lemma holomorphic_compact_finite_zeros: + assumes S: "f holomorphic_on S" "open S" "connected S" + and "compact K" "K \ S" + and "\ f constant_on S" + shows "finite {z\K. f z = 0}" +proof (rule ccontr) + assume "infinite {z\K. f z = 0}" + then obtain z where "z \ K" and z: "z islimpt {z\K. f z = 0}" + using \compact K\ by (auto simp: compact_eq_Bolzano_Weierstrass) + moreover have "{z\K. f z = 0} \ S" + using \K \ S\ by blast + ultimately show False + using assms analytic_continuation [OF S] unfolding constant_on_def + by blast +qed + subsection\Open mapping theorem\ lemma holomorphic_contract_to_zero: assumes contf: "continuous_on (cball \ r) f" and holf: "f holomorphic_on ball \ r" and "0 < r" and norm_less: "\z. norm(\ - z) = r \ norm(f \) < norm(f z)" obtains z where "z \ ball \ r" "f z = 0" proof - { assume fnz: "\w. w \ ball \ r \ f w \ 0" then have "0 < norm (f \)" by (simp add: \0 < r\) have fnz': "\w. w \ cball \ r \ f w \ 0" by (metis norm_less dist_norm fnz less_eq_real_def mem_ball mem_cball norm_not_less_zero norm_zero) have "frontier(cball \ r) \ {}" using \0 < r\ by simp define g where [abs_def]: "g z = inverse (f z)" for z have contg: "continuous_on (cball \ r) g" unfolding g_def using contf continuous_on_inverse fnz' by blast have holg: "g holomorphic_on ball \ r" unfolding g_def using fnz holf holomorphic_on_inverse by blast have "frontier (cball \ r) \ cball \ r" by (simp add: subset_iff) then have contf': "continuous_on (frontier (cball \ r)) f" and contg': "continuous_on (frontier (cball \ r)) g" by (blast intro: contf contg continuous_on_subset)+ have froc: "frontier(cball \ r) \ {}" using \0 < r\ by simp moreover have "continuous_on (frontier (cball \ r)) (norm o f)" using contf' continuous_on_compose continuous_on_norm_id by blast ultimately obtain w where w: "w \ frontier(cball \ r)" and now: "\x. x \ frontier(cball \ r) \ norm (f w) \ norm (f x)" using continuous_attains_inf [OF compact_frontier [OF compact_cball]] by (metis comp_apply) then have fw: "0 < norm (f w)" by (simp add: fnz') have "continuous_on (frontier (cball \ r)) (norm o g)" using contg' continuous_on_compose continuous_on_norm_id by blast then obtain v where v: "v \ frontier(cball \ r)" and nov: "\x. x \ frontier(cball \ r) \ norm (g v) \ norm (g x)" using continuous_attains_sup [OF compact_frontier [OF compact_cball] froc] by force then have fv: "0 < norm (f v)" by (simp add: fnz') have "norm ((deriv ^^ 0) g \) \ fact 0 * norm (g v) / r ^ 0" by (rule Cauchy_inequality [OF holg contg \0 < r\]) (simp add: dist_norm nov) then have "cmod (g \) \ cmod (g v)" by simp moreover have "cmod (\ - w) = r" by (metis (no_types) dist_norm frontier_cball mem_sphere w) ultimately obtain wr: "norm (\ - w) = r" and nfw: "norm (f w) \ norm (f \)" unfolding g_def by (metis (no_types) \0 < cmod (f \)\ less_imp_inverse_less norm_inverse not_le now order_trans v) with fw have False using norm_less by force } with that show ?thesis by blast qed theorem open_mapping_thm: assumes holf: "f holomorphic_on S" and S: "open S" and "connected S" and "open U" and "U \ S" and fne: "\ f constant_on S" shows "open (f ` U)" proof - have *: "open (f ` U)" if "U \ {}" and U: "open U" "connected U" and "f holomorphic_on U" and fneU: "\x. \y \ U. f y \ x" for U proof (clarsimp simp: open_contains_ball) fix \ assume \: "\ \ U" show "\e>0. ball (f \) e \ f ` U" proof - have hol: "(\z. f z - f \) holomorphic_on U" by (rule holomorphic_intros that)+ obtain s where "0 < s" and sbU: "ball \ s \ U" and sne: "\z. z \ ball \ s - {\} \ (\z. f z - f \) z \ 0" using isolated_zeros [OF hol U \] by (metis fneU right_minus_eq) obtain r where "0 < r" and r: "cball \ r \ ball \ s" using \0 < s\ by (rule_tac r="s/2" in that) auto have "cball \ r \ U" using sbU r by blast then have frsbU: "frontier (cball \ r) \ U" using Diff_subset frontier_def order_trans by fastforce then have cof: "compact (frontier(cball \ r))" by blast have frne: "frontier (cball \ r) \ {}" using \0 < r\ by auto have contfr: "continuous_on (frontier (cball \ r)) (\z. norm (f z - f \))" by (metis continuous_on_norm continuous_on_subset frsbU hol holomorphic_on_imp_continuous_on) obtain w where "norm (\ - w) = r" and w: "(\z. norm (\ - z) = r \ norm (f w - f \) \ norm(f z - f \))" using continuous_attains_inf [OF cof frne contfr] by (auto simp: dist_norm) moreover define \ where "\ \ norm (f w - f \) / 3" ultimately have "0 < \" using \0 < r\ dist_complex_def r sne by auto have "ball (f \) \ \ f ` U" proof fix \ assume \: "\ \ ball (f \) \" have *: "cmod (\ - f \) < cmod (\ - f z)" if "cmod (\ - z) = r" for z proof - have lt: "cmod (f w - f \) / 3 < cmod (\ - f z)" using w [OF that] \ using dist_triangle2 [of "f \" "\" "f z"] dist_triangle2 [of "f \" "f z" \] by (simp add: \_def dist_norm norm_minus_commute) show ?thesis by (metis \_def dist_commute dist_norm less_trans lt mem_ball \) qed have "continuous_on (cball \ r) (\z. \ - f z)" using \cball \ r \ U\ \f holomorphic_on U\ by (force intro: continuous_intros continuous_on_subset holomorphic_on_imp_continuous_on) moreover have "(\z. \ - f z) holomorphic_on ball \ r" using \cball \ r \ U\ ball_subset_cball holomorphic_on_subset that(4) by (intro holomorphic_intros) blast ultimately obtain z where "z \ ball \ r" "\ - f z = 0" using "*" \0 < r\ holomorphic_contract_to_zero by blast then show "\ \ f ` U" using \cball \ r \ U\ by fastforce qed then show ?thesis using \0 < \\ by blast qed qed have "open (f ` X)" if "X \ components U" for X proof - have holfU: "f holomorphic_on U" using \U \ S\ holf holomorphic_on_subset by blast have "X \ {}" using that by (simp add: in_components_nonempty) moreover have "open X" using that \open U\ open_components by auto moreover have "connected X" using that in_components_maximal by blast moreover have "f holomorphic_on X" by (meson that holfU holomorphic_on_subset in_components_maximal) moreover have "\y\X. f y \ x" for x proof (rule ccontr) assume not: "\ (\y\X. f y \ x)" have "X \ S" using \U \ S\ in_components_subset that by blast obtain w where w: "w \ X" using \X \ {}\ by blast have wis: "w islimpt X" using w \open X\ interior_eq by auto have hol: "(\z. f z - x) holomorphic_on S" by (simp add: holf holomorphic_on_diff) with fne [unfolded constant_on_def] analytic_continuation[OF hol S \connected S\ \X \ S\ _ wis] not \X \ S\ w show False by auto qed ultimately show ?thesis by (rule *) qed then have "open (f ` \(components U))" by (metis (no_types, lifting) imageE image_Union open_Union) then show ?thesis by force qed text\No need for \<^term>\S\ to be connected. But the nonconstant condition is stronger.\ corollary\<^marker>\tag unimportant\ open_mapping_thm2: assumes holf: "f holomorphic_on S" and S: "open S" and "open U" "U \ S" and fnc: "\X. \open X; X \ S; X \ {}\ \ \ f constant_on X" shows "open (f ` U)" proof - have "S = \(components S)" by simp with \U \ S\ have "U = (\C \ components S. C \ U)" by auto then have "f ` U = (\C \ components S. f ` (C \ U))" using image_UN by fastforce moreover { fix C assume "C \ components S" with S \C \ components S\ open_components in_components_connected have C: "open C" "connected C" by auto have "C \ S" by (metis \C \ components S\ in_components_maximal) have nf: "\ f constant_on C" using \open C\ \C \ components S\ \C \ S\ fnc in_components_nonempty by blast have "f holomorphic_on C" by (metis holf holomorphic_on_subset \C \ S\) then have "open (f ` (C \ U))" by (meson C \open U\ inf_le1 nf open_Int open_mapping_thm) } ultimately show ?thesis by force qed corollary\<^marker>\tag unimportant\ open_mapping_thm3: assumes holf: "f holomorphic_on S" and "open S" and injf: "inj_on f S" shows "open (f ` S)" proof (rule open_mapping_thm2 [OF holf]) show "\X. \open X; X \ S; X \ {}\ \ \ f constant_on X" using inj_on_subset injective_not_constant injf by blast qed (use assms in auto) subsection\Maximum modulus principle\ text\If \<^term>\f\ is holomorphic, then its norm (modulus) cannot exhibit a true local maximum that is properly within the domain of \<^term>\f\.\ proposition maximum_modulus_principle: assumes holf: "f holomorphic_on S" and S: "open S" and "connected S" and "open U" and "U \ S" and "\ \ U" and no: "\z. z \ U \ norm(f z) \ norm(f \)" shows "f constant_on S" proof (rule ccontr) assume "\ f constant_on S" then have "open (f ` U)" using open_mapping_thm assms by blast moreover have "\ open (f ` U)" proof - have "\t. cmod (f \ - t) < e \ t \ f ` U" if "0 < e" for e using that apply (rule_tac x="if 0 < Re(f \) then f \ + (e/2) else f \ - (e/2)" in exI) apply (simp add: dist_norm) apply (fastforce simp: cmod_Re_le_iff dest!: no dest: sym) done then show ?thesis unfolding open_contains_ball by (metis \\ \ U\ contra_subsetD dist_norm imageI mem_ball) qed ultimately show False by blast qed proposition maximum_modulus_frontier: assumes holf: "f holomorphic_on (interior S)" and contf: "continuous_on (closure S) f" and bos: "bounded S" and leB: "\z. z \ frontier S \ norm(f z) \ B" and "\ \ S" shows "norm(f \) \ B" proof - have "compact (closure S)" using bos by (simp add: bounded_closure compact_eq_bounded_closed) moreover have "continuous_on (closure S) (cmod \ f)" using contf continuous_on_compose continuous_on_norm_id by blast ultimately obtain z where "z \ closure S" and z: "\y. y \ closure S \ (cmod \ f) y \ (cmod \ f) z" using continuous_attains_sup [of "closure S" "norm o f"] \\ \ S\ by auto then consider "z \ frontier S" | "z \ interior S" using frontier_def by auto then have "norm(f z) \ B" proof cases case 1 then show ?thesis using leB by blast next case 2 have "f constant_on (connected_component_set (interior S) z)" proof (rule maximum_modulus_principle) show "f holomorphic_on connected_component_set (interior S) z" by (metis connected_component_subset holf holomorphic_on_subset) show zin: "z \ connected_component_set (interior S) z" by (simp add: 2) show "\W. W \ connected_component_set (interior S) z \ cmod (f W) \ cmod (f z)" using closure_def connected_component_subset z by fastforce qed (auto simp: open_connected_component) then obtain c where c: "\w. w \ connected_component_set (interior S) z \ f w = c" by (auto simp: constant_on_def) have "f ` closure(connected_component_set (interior S) z) \ {c}" proof (rule image_closure_subset) show "continuous_on (closure (connected_component_set (interior S) z)) f" by (meson closure_mono connected_component_subset contf continuous_on_subset interior_subset) qed (use c in auto) then have cc: "\w. w \ closure(connected_component_set (interior S) z) \ f w = c" by blast have "connected_component (interior S) z z" by (simp add: "2") moreover have "connected_component_set (interior S) z \ UNIV" by (metis bos bounded_interior connected_component_eq_UNIV not_bounded_UNIV) ultimately have "frontier(connected_component_set (interior S) z) \ {}" by (meson "2" connected_component_eq_empty frontier_not_empty) then obtain w where w: "w \ frontier(connected_component_set (interior S) z)" by auto then have "norm (f z) = norm (f w)" by (simp add: "2" c cc frontier_def) also have "... \ B" using w frontier_interior_subset frontier_of_connected_component_subset by (blast intro: leB) finally show ?thesis . qed then show ?thesis using z \\ \ S\ closure_subset by fastforce qed corollary\<^marker>\tag unimportant\ maximum_real_frontier: assumes holf: "f holomorphic_on (interior S)" and contf: "continuous_on (closure S) f" and bos: "bounded S" and leB: "\z. z \ frontier S \ Re(f z) \ B" and "\ \ S" shows "Re(f \) \ B" using maximum_modulus_frontier [of "exp o f" S "exp B"] Transcendental.continuous_on_exp holomorphic_on_compose holomorphic_on_exp assms by auto subsection\<^marker>\tag unimportant\ \Factoring out a zero according to its order\ lemma holomorphic_factor_order_of_zero: assumes holf: "f holomorphic_on S" and os: "open S" and "\ \ S" "0 < n" and dnz: "(deriv ^^ n) f \ \ 0" and dfz: "\i. \0 < i; i < n\ \ (deriv ^^ i) f \ = 0" obtains g r where "0 < r" "g holomorphic_on ball \ r" "\w. w \ ball \ r \ f w - f \ = (w - \)^n * g w" "\w. w \ ball \ r \ g w \ 0" proof - obtain r where "r>0" and r: "ball \ r \ S" using assms by (blast elim!: openE) then have holfb: "f holomorphic_on ball \ r" using holf holomorphic_on_subset by blast define g where "g w = suminf (\i. (deriv ^^ (i + n)) f \ / (fact(i + n)) * (w - \)^i)" for w have sumsg: "(\i. (deriv ^^ (i + n)) f \ / (fact(i + n)) * (w - \)^i) sums g w" and feq: "f w - f \ = (w - \)^n * g w" if w: "w \ ball \ r" for w proof - define powf where "powf = (\i. (deriv ^^ i) f \/(fact i) * (w - \)^i)" have [simp]: "powf 0 = f \" by (simp add: powf_def) have sing: "{.. = 0 then {} else {0})" unfolding powf_def using \0 < n\ dfz by (auto simp: dfz; metis funpow_0 not_gr0) have "powf sums f w" unfolding powf_def by (rule holomorphic_power_series [OF holfb w]) moreover have "(\i" by (subst sum.setdiff_irrelevant [symmetric]; simp add: dfz sing) ultimately have fsums: "(\i. powf (i+n)) sums (f w - f \)" using w sums_iff_shift' by metis then have *: "summable (\i. (w - \) ^ n * ((deriv ^^ (i + n)) f \ * (w - \) ^ i / fact (i + n)))" unfolding powf_def using sums_summable by (auto simp: power_add mult_ac) have "summable (\i. (deriv ^^ (i + n)) f \ * (w - \) ^ i / fact (i + n))" proof (cases "w=\") case False then show ?thesis using summable_mult [OF *, of "1 / (w - \) ^ n"] by simp next case True then show ?thesis by (auto simp: Power.semiring_1_class.power_0_left intro!: summable_finite [of "{0}"] split: if_split_asm) qed then show sumsg: "(\i. (deriv ^^ (i + n)) f \ / (fact(i + n)) * (w - \)^i) sums g w" by (simp add: summable_sums_iff g_def) show "f w - f \ = (w - \)^n * g w" using sums_mult [OF sumsg, of "(w - \) ^ n"] by (intro sums_unique2 [OF fsums]) (auto simp: power_add mult_ac powf_def) qed then have holg: "g holomorphic_on ball \ r" by (meson sumsg power_series_holomorphic) then have contg: "continuous_on (ball \ r) g" by (blast intro: holomorphic_on_imp_continuous_on) have "g \ \ 0" using dnz unfolding g_def by (subst suminf_finite [of "{0}"]) auto obtain d where "0 < d" and d: "\w. w \ ball \ d \ g w \ 0" using \0 < r\ continuous_on_avoid [OF contg _ \g \ \ 0\] by (metis centre_in_ball le_cases mem_ball mem_ball_leI) show ?thesis proof show "g holomorphic_on ball \ (min r d)" using holg by (auto simp: feq holomorphic_on_subset subset_ball d) qed (use \0 < r\ \0 < d\ in \auto simp: feq d\) qed lemma holomorphic_factor_order_of_zero_strong: assumes holf: "f holomorphic_on S" "open S" "\ \ S" "0 < n" and "(deriv ^^ n) f \ \ 0" and "\i. \0 < i; i < n\ \ (deriv ^^ i) f \ = 0" obtains g r where "0 < r" "g holomorphic_on ball \ r" "\w. w \ ball \ r \ f w - f \ = ((w - \) * g w) ^ n" "\w. w \ ball \ r \ g w \ 0" proof - obtain g r where "0 < r" and holg: "g holomorphic_on ball \ r" and feq: "\w. w \ ball \ r \ f w - f \ = (w - \)^n * g w" and gne: "\w. w \ ball \ r \ g w \ 0" by (auto intro: holomorphic_factor_order_of_zero [OF assms]) have con: "continuous_on (ball \ r) (\z. deriv g z / g z)" by (rule continuous_intros) (auto simp: gne holg holomorphic_deriv holomorphic_on_imp_continuous_on) have cd: "(\z. deriv g z / g z) field_differentiable at x" if "dist \ x < r" for x proof (intro derivative_intros) show "deriv g field_differentiable at x" using that holg mem_ball by (blast intro: holomorphic_deriv holomorphic_on_imp_differentiable_at) show "g field_differentiable at x" by (metis that open_ball at_within_open holg holomorphic_on_def mem_ball) qed (simp add: gne that) obtain h where h: "\x. x \ ball \ r \ (h has_field_derivative deriv g x / g x) (at x)" using holomorphic_convex_primitive [of "ball \ r" "{}" "\z. deriv g z / g z"] by (metis (no_types, lifting) Diff_empty at_within_interior cd con convex_ball infinite_imp_nonempty interior_ball mem_ball) then have "continuous_on (ball \ r) h" by (metis open_ball holomorphic_on_imp_continuous_on holomorphic_on_open) then have con: "continuous_on (ball \ r) (\x. exp (h x) / g x)" by (auto intro!: continuous_intros simp add: holg holomorphic_on_imp_continuous_on gne) have 0: "dist \ x < r \ ((\x. exp (h x) / g x) has_field_derivative 0) (at x)" for x apply (rule h derivative_eq_intros DERIV_deriv_iff_field_differentiable [THEN iffD2] | simp)+ using holg by (auto simp: holomorphic_on_imp_differentiable_at gne h) obtain c where c: "\x. x \ ball \ r \ exp (h x) / g x = c" by (rule DERIV_zero_connected_constant [of "ball \ r" "{}" "\x. exp(h x) / g x"]) (auto simp: con 0) have hol: "(\z. exp ((Ln (inverse c) + h z) / of_nat n)) holomorphic_on ball \ r" proof (intro holomorphic_intros holomorphic_on_compose [unfolded o_def, where g = exp]) show "h holomorphic_on ball \ r" using h holomorphic_on_open by blast qed (use \0 < n\ in auto) show ?thesis proof show "\w. w \ ball \ r \ f w - f \ = ((w - \) * exp ((Ln (inverse c) + h w) / of_nat n)) ^ n" using \0 < n\ by (auto simp: feq power_mult_distrib exp_divide_power_eq exp_add gne simp flip: c) qed (use hol \0 < r\ in auto) qed lemma fixes k :: "'a::wellorder" assumes a_def: "a == LEAST x. P x" and P: "P k" shows def_LeastI: "P a" and def_Least_le: "a \ k" unfolding a_def by (rule LeastI Least_le; rule P)+ lemma holomorphic_factor_zero_nonconstant: assumes holf: "f holomorphic_on S" and S: "open S" "connected S" and "\ \ S" "f \ = 0" and nonconst: "\ f constant_on S" obtains g r n where "0 < n" "0 < r" "ball \ r \ S" "g holomorphic_on ball \ r" "\w. w \ ball \ r \ f w = (w - \)^n * g w" "\w. w \ ball \ r \ g w \ 0" proof (cases "\n>0. (deriv ^^ n) f \ = 0") case True then show ?thesis using holomorphic_fun_eq_const_on_connected [OF holf S _ \\ \ S\] nonconst by (simp add: constant_on_def) next case False then obtain n0 where "n0 > 0" and n0: "(deriv ^^ n0) f \ \ 0" by blast obtain r0 where "r0 > 0" "ball \ r0 \ S" using S openE \\ \ S\ by auto define n where "n \ LEAST n. (deriv ^^ n) f \ \ 0" have n_ne: "(deriv ^^ n) f \ \ 0" by (rule def_LeastI [OF n_def]) (rule n0) then have "0 < n" using \f \ = 0\ using funpow_0 by fastforce have n_min: "\k. k < n \ (deriv ^^ k) f \ = 0" using def_Least_le [OF n_def] not_le by blast then obtain g r1 where g: "0 < r1" "g holomorphic_on ball \ r1" and geq: "\w. w \ ball \ r1 \ f w = (w - \) ^ n * g w" and g0: "\w. w \ ball \ r1 \ g w \ 0" by (auto intro: holomorphic_factor_order_of_zero [OF holf \open S\ \\ \ S\ \n > 0\ n_ne] simp: \f \ = 0\) show ?thesis proof show "g holomorphic_on ball \ (min r0 r1)" using g by auto show "\w. w \ ball \ (min r0 r1) \ f w = (w - \) ^ n * g w" by (simp add: geq) qed (use \0 < n\ \0 < r0\ \0 < r1\ \ball \ r0 \ S\ g0 in auto) qed lemma holomorphic_lower_bound_difference: assumes holf: "f holomorphic_on S" and S: "open S" "connected S" and "\ \ S" and "\ \ S" and fne: "f \ \ f \" obtains k n r where "0 < k" "0 < r" "ball \ r \ S" "\w. w \ ball \ r \ k * norm(w - \)^n \ norm(f w - f \)" proof - define n where "n = (LEAST n. 0 < n \ (deriv ^^ n) f \ \ 0)" obtain n0 where "0 < n0" and n0: "(deriv ^^ n0) f \ \ 0" using fne holomorphic_fun_eq_const_on_connected [OF holf S] \\ \ S\ \\ \ S\ by blast then have "0 < n" and n_ne: "(deriv ^^ n) f \ \ 0" unfolding n_def by (metis (mono_tags, lifting) LeastI)+ have n_min: "\k. \0 < k; k < n\ \ (deriv ^^ k) f \ = 0" unfolding n_def by (blast dest: not_less_Least) then obtain g r where "0 < r" and holg: "g holomorphic_on ball \ r" and fne: "\w. w \ ball \ r \ f w - f \ = (w - \) ^ n * g w" and gnz: "\w. w \ ball \ r \ g w \ 0" by (auto intro: holomorphic_factor_order_of_zero [OF holf \open S\ \\ \ S\ \n > 0\ n_ne]) obtain e where "e>0" and e: "ball \ e \ S" using assms by (blast elim!: openE) then have holfb: "f holomorphic_on ball \ e" using holf holomorphic_on_subset by blast define d where "d = (min e r) / 2" have "0 < d" using \0 < r\ \0 < e\ by (simp add: d_def) have "d < r" using \0 < r\ by (auto simp: d_def) then have cbb: "cball \ d \ ball \ r" by (auto simp: cball_subset_ball_iff) then have "g holomorphic_on cball \ d" by (rule holomorphic_on_subset [OF holg]) then have "closed (g ` cball \ d)" by (simp add: compact_imp_closed compact_continuous_image holomorphic_on_imp_continuous_on) moreover have "g ` cball \ d \ {}" using \0 < d\ by auto ultimately obtain x where x: "x \ g ` cball \ d" and "\y. y \ g ` cball \ d \ dist 0 x \ dist 0 y" by (rule distance_attains_inf) blast then have leg: "\w. w \ cball \ d \ norm x \ norm (g w)" by auto have "ball \ d \ cball \ d" by auto also have "... \ ball \ e" using \0 < d\ d_def by auto also have "... \ S" by (rule e) finally have dS: "ball \ d \ S" . have "x \ 0" using gnz x \d < r\ by auto show thesis proof show "\w. w \ ball \ d \ cmod x * cmod (w - \) ^ n \ cmod (f w - f \)" using \d < r\ leg by (auto simp: fne norm_mult norm_power algebra_simps mult_right_mono) qed (use dS \x \ 0\ \d > 0\ in auto) qed lemma assumes holf: "f holomorphic_on (S - {\})" and \: "\ \ interior S" shows holomorphic_on_extend_lim: "(\g. g holomorphic_on S \ (\z \ S - {\}. g z = f z)) \ ((\z. (z - \) * f z) \ 0) (at \)" (is "?P = ?Q") and holomorphic_on_extend_bounded: "(\g. g holomorphic_on S \ (\z \ S - {\}. g z = f z)) \ (\B. eventually (\z. norm(f z) \ B) (at \))" (is "?P = ?R") proof - obtain \ where "0 < \" and \: "ball \ \ \ S" using \ mem_interior by blast have "?R" if holg: "g holomorphic_on S" and gf: "\z. z \ S - {\} \ g z = f z" for g proof - have \
: "cmod (f x) \ cmod (g \) + 1" if "x \ \" "dist x \ < \" "dist (g x) (g \) < 1" for x proof - have "x \ S" by (metis \ dist_commute mem_ball subsetD that(2)) with that gf [of x] show ?thesis using norm_triangle_ineq2 [of "f x" "g \"] dist_complex_def by auto qed then have *: "\\<^sub>F z in at \. dist (g z) (g \) < 1 \ cmod (f z) \ cmod (g \) + 1" using \0 < \\ eventually_at by blast have "continuous_on (interior S) g" by (meson continuous_on_subset holg holomorphic_on_imp_continuous_on interior_subset) then have "\x. x \ interior S \ (g \ g x) (at x)" using continuous_on_interior continuous_within holg holomorphic_on_imp_continuous_on by blast then have "(g \ g \) (at \)" by (simp add: \) then show ?thesis apply (rule_tac x="norm(g \) + 1" in exI) apply (rule eventually_mp [OF * tendstoD [where e=1]], auto) done qed moreover have "?Q" if "\\<^sub>F z in at \. cmod (f z) \ B" for B by (rule lim_null_mult_right_bounded [OF _ that]) (simp add: LIM_zero) moreover have "?P" if "(\z. (z - \) * f z) \\\ 0" proof - define h where [abs_def]: "h z = (z - \)^2 * f z" for z have h0: "(h has_field_derivative 0) (at \)" apply (simp add: h_def has_field_derivative_iff) apply (auto simp: field_split_simps power2_eq_square Lim_transform_within [OF that, of 1]) done have holh: "h holomorphic_on S" proof (simp add: holomorphic_on_def, clarify) fix z assume "z \ S" show "h field_differentiable at z within S" proof (cases "z = \") case True then show ?thesis using field_differentiable_at_within field_differentiable_def h0 by blast next case False then have "f field_differentiable at z within S" using holomorphic_onD [OF holf, of z] \z \ S\ unfolding field_differentiable_def has_field_derivative_iff by (force intro: exI [where x="dist \ z"] elim: Lim_transform_within_set [unfolded eventually_at]) then show ?thesis by (simp add: h_def power2_eq_square derivative_intros) qed qed define g where [abs_def]: "g z = (if z = \ then deriv h \ else (h z - h \) / (z - \))" for z have holg: "g holomorphic_on S" unfolding g_def by (rule pole_lemma [OF holh \]) have \
: "\z\S - {\}. (g z - g \) / (z - \) = f z" using h0 by (auto simp: g_def power2_eq_square divide_simps DERIV_imp_deriv h_def) show ?thesis apply (intro exI conjI) apply (rule pole_lemma [OF holg \]) apply (simp add: \
) done qed ultimately show "?P = ?Q" and "?P = ?R" by meson+ qed lemma pole_at_infinity: assumes holf: "f holomorphic_on UNIV" and lim: "((inverse o f) \ l) at_infinity" obtains a n where "\z. f z = (\i\n. a i * z^i)" proof (cases "l = 0") case False show thesis proof show "f z = (\i\0. inverse l * z ^ i)" for z using False tendsto_inverse [OF lim] by (simp add: Liouville_weak [OF holf]) qed next case True then have [simp]: "l = 0" . show ?thesis proof (cases "\r. 0 < r \ (\z \ ball 0 r - {0}. f(inverse z) \ 0)") case True then obtain r where "0 < r" and r: "\z. z \ ball 0 r - {0} \ f(inverse z) \ 0" by auto have 1: "inverse \ f \ inverse holomorphic_on ball 0 r - {0}" by (rule holomorphic_on_compose holomorphic_intros holomorphic_on_subset [OF holf] | force simp: r)+ have 2: "0 \ interior (ball 0 r)" using \0 < r\ by simp have "\B. 0 eventually (\z. cmod ((inverse \ f \ inverse) z) \ B) (at 0)" apply (rule exI [where x=1]) using tendstoD [OF lim [unfolded lim_at_infinity_0] zero_less_one] by (simp add: eventually_mono) with holomorphic_on_extend_bounded [OF 1 2] obtain g where holg: "g holomorphic_on ball 0 r" and geq: "\z. z \ ball 0 r - {0} \ g z = (inverse \ f \ inverse) z" by meson have ifi0: "(inverse \ f \ inverse) \0\ 0" using \l = 0\ lim lim_at_infinity_0 by blast have g2g0: "g \0\ g 0" using \0 < r\ centre_in_ball continuous_at continuous_on_eq_continuous_at holg by (blast intro: holomorphic_on_imp_continuous_on) have g2g1: "g \0\ 0" apply (rule Lim_transform_within_open [OF ifi0 open_ball [of 0 r]]) using \0 < r\ by (auto simp: geq) have [simp]: "g 0 = 0" by (rule tendsto_unique [OF _ g2g0 g2g1]) simp have "ball 0 r - {0::complex} \ {}" using \0 < r\ by (metis "2" Diff_iff insert_Diff interior_ball interior_singleton) then obtain w::complex where "w \ 0" and w: "norm w < r" by force then have "g w \ 0" by (simp add: geq r) obtain B n e where "0 < B" "0 < e" "e \ r" and leg: "\w. norm w < e \ B * cmod w ^ n \ cmod (g w)" proof (rule holomorphic_lower_bound_difference [OF holg open_ball connected_ball]) show "g w \ g 0" by (simp add: \g w \ 0\) show "w \ ball 0 r" using mem_ball_0 w by blast qed (use \0 < r\ in \auto simp: ball_subset_ball_iff\) have \
: "cmod (f z) \ cmod z ^ n / B" if "2/e \ cmod z" for z proof - have ize: "inverse z \ ball 0 e - {0}" using that \0 < e\ by (auto simp: norm_divide field_split_simps algebra_simps) then have [simp]: "z \ 0" and izr: "inverse z \ ball 0 r - {0}" using \e \ r\ by auto then have [simp]: "f z \ 0" using r [of "inverse z"] by simp have [simp]: "f z = inverse (g (inverse z))" using izr geq [of "inverse z"] by simp show ?thesis using ize leg [of "inverse z"] \0 < B\ \0 < e\ by (simp add: field_split_simps norm_divide algebra_simps) qed show thesis proof show "f z = (\i\n. (deriv ^^ i) f 0 / fact i * z ^ i)" for z using \
by (rule_tac A = "2/e" and B = "1/B" in Liouville_polynomial [OF holf], simp) qed next case False then have fi0: "\r. r > 0 \ \z\ball 0 r - {0}. f (inverse z) = 0" by simp have fz0: "f z = 0" if "0 < r" and lt1: "\x. x \ 0 \ cmod x < r \ inverse (cmod (f (inverse x))) < 1" for z r proof - have f0: "(f \ 0) at_infinity" proof - have DIM_complex[intro]: "2 \ DIM(complex)" \ \should not be necessary!\ by simp have "f (inverse x) \ 0 \ x \ 0 \ cmod x < r \ 1 < cmod (f (inverse x))" for x using lt1[of x] by (auto simp: field_simps) then have **: "cmod (f (inverse x)) \ 1 \ x \ 0 \ cmod x < r \ f (inverse x) = 0" for x by force then have *: "(f \ inverse) ` (ball 0 r - {0}) \ {0} \ - ball 0 1" by force have "continuous_on (inverse ` (ball 0 r - {0})) f" using continuous_on_subset holf holomorphic_on_imp_continuous_on by blast then have "connected ((f \ inverse) ` (ball 0 r - {0}))" using connected_punctured_ball by (intro connected_continuous_image continuous_intros; force) then have "{0} \ (f \ inverse) ` (ball 0 r - {0}) = {} \ - ball 0 1 \ (f \ inverse) ` (ball 0 r - {0}) = {}" by (rule connected_closedD) (use * in auto) then have "f (inverse w) = 0" if "w \ 0" "cmod w < r" for w using **[of w] fi0 \0 < r\ that by force then show ?thesis unfolding lim_at_infinity_0 using eventually_at \r > 0\ by (force simp add: intro: tendsto_eventually) qed obtain w where "w \ ball 0 r - {0}" and "f (inverse w) = 0" using False \0 < r\ by blast then show ?thesis by (auto simp: f0 Liouville_weak [OF holf, of 0]) qed show thesis proof show "\z. f z = (\i\0. 0 * z ^ i)" using lim apply (simp add: lim_at_infinity_0 Lim_at dist_norm norm_inverse) using fz0 zero_less_one by blast qed qed qed subsection\<^marker>\tag unimportant\ \Entire proper functions are precisely the non-trivial polynomials\ lemma proper_map_polyfun: fixes c :: "nat \ 'a::{real_normed_div_algebra,heine_borel}" assumes "closed S" and "compact K" and c: "c i \ 0" "1 \ i" "i \ n" shows "compact (S \ {z. (\i\n. c i * z^i) \ K})" proof - obtain B where "B > 0" and B: "\x. x \ K \ norm x \ B" by (metis compact_imp_bounded \compact K\ bounded_pos) have *: "norm x \ b" if "\x. b \ norm x \ B + 1 \ norm (\i\n. c i * x ^ i)" "(\i\n. c i * x ^ i) \ K" for b x proof - have "norm (\i\n. c i * x ^ i) \ B" using B that by blast moreover have "\ B + 1 \ B" by simp ultimately show "norm x \ b" using that by (metis (no_types) less_eq_real_def not_less order_trans) qed have "bounded {z. (\i\n. c i * z ^ i) \ K}" using Limits.polyfun_extremal [where c=c and B="B+1", OF c] by (auto simp: bounded_pos eventually_at_infinity_pos *) moreover have "closed ((\z. (\i\n. c i * z ^ i)) -` K)" using \compact K\ compact_eq_bounded_closed by (intro allI continuous_closed_vimage continuous_intros; force) ultimately show ?thesis using closed_Int_compact [OF \closed S\] compact_eq_bounded_closed by (auto simp add: vimage_def) qed lemma proper_map_polyfun_univ: fixes c :: "nat \ 'a::{real_normed_div_algebra,heine_borel}" assumes "compact K" "c i \ 0" "1 \ i" "i \ n" shows "compact ({z. (\i\n. c i * z^i) \ K})" using proper_map_polyfun [of UNIV K c i n] assms by simp lemma proper_map_polyfun_eq: assumes "f holomorphic_on UNIV" shows "(\k. compact k \ compact {z. f z \ k}) \ (\c n. 0 < n \ (c n \ 0) \ f = (\z. \i\n. c i * z^i))" (is "?lhs = ?rhs") proof assume compf [rule_format]: ?lhs have 2: "\k. 0 < k \ a k \ 0 \ f = (\z. \i \ k. a i * z ^ i)" if "\z. f z = (\i\n. a i * z ^ i)" for a n proof (cases "\i\n. 0 a i = 0") case True then have [simp]: "\z. f z = a 0" by (simp add: that sum.atMost_shift) have False using compf [of "{a 0}"] by simp then show ?thesis .. next case False then obtain k where k: "0 < k" "k\n" "a k \ 0" by force define m where "m = (GREATEST k. k\n \ a k \ 0)" have m: "m\n \ a m \ 0" unfolding m_def using GreatestI_nat [where b = n] k by (metis (mono_tags, lifting)) have [simp]: "a i = 0" if "m < i" "i \ n" for i using Greatest_le_nat [where b = "n" and P = "\k. k\n \ a k \ 0"] using m_def not_le that by auto have "k \ m" unfolding m_def using Greatest_le_nat [where b = n] k by (metis (mono_tags, lifting)) with k m show ?thesis by (rule_tac x=m in exI) (auto simp: that comm_monoid_add_class.sum.mono_neutral_right) qed have \
: "((inverse \ f) \ 0) at_infinity" proof (rule Lim_at_infinityI) fix e::real assume "0 < e" with compf [of "cball 0 (inverse e)"] show "\B. \x. B \ cmod x \ dist ((inverse \ f) x) 0 \ e" apply simp apply (clarsimp simp add: compact_eq_bounded_closed bounded_pos norm_inverse) by (metis (no_types, opaque_lifting) inverse_inverse_eq le_less_trans less_eq_real_def less_imp_inverse_less linordered_field_no_ub not_less) qed then obtain a n where "\z. f z = (\i\n. a i * z^i)" using assms pole_at_infinity by blast with \
2 show ?rhs by blast next assume ?rhs then obtain c n where "0 < n" "c n \ 0" "f = (\z. \i\n. c i * z ^ i)" by blast then have "compact {z. f z \ k}" if "compact k" for k by (auto intro: proper_map_polyfun_univ [OF that]) then show ?lhs by blast qed subsection \Relating invertibility and nonvanishing of derivative\ lemma has_complex_derivative_locally_injective: assumes holf: "f holomorphic_on S" and S: "\ \ S" "open S" and dnz: "deriv f \ \ 0" obtains r where "r > 0" "ball \ r \ S" "inj_on f (ball \ r)" proof - have *: "\d>0. \x. dist \ x < d \ onorm (\v. deriv f x * v - deriv f \ * v) < e" if "e > 0" for e proof - have contdf: "continuous_on S (deriv f)" by (simp add: holf holomorphic_deriv holomorphic_on_imp_continuous_on \open S\) obtain \ where "\>0" and \: "\x. \x \ S; dist x \ \ \\ \ cmod (deriv f x - deriv f \) \ e/2" using continuous_onE [OF contdf \\ \ S\, of "e/2"] \0 < e\ by (metis dist_complex_def half_gt_zero less_imp_le) have \
: "\\ z. \\ \ S; dist \ \ < \\ \ cmod (deriv f \ - deriv f \) * cmod z \ e/2 * cmod z" by (intro mult_right_mono [OF \]) (auto simp: dist_commute) obtain \ where "\>0" "ball \ \ \ S" by (metis openE [OF \open S\ \\ \ S\]) with \\>0\ have "\\>0. \x. dist \ x < \ \ onorm (\v. deriv f x * v - deriv f \ * v) \ e/2" using \
apply (rule_tac x="min \ \" in exI) apply (intro conjI allI impI Operator_Norm.onorm_le) apply (force simp: mult_right_mono norm_mult [symmetric] left_diff_distrib \)+ done with \e>0\ show ?thesis by force qed have "inj ((*) (deriv f \))" using dnz by simp then obtain g' where g': "linear g'" "g' \ (*) (deriv f \) = id" using linear_injective_left_inverse [of "(*) (deriv f \)"] by auto show ?thesis apply (rule has_derivative_locally_injective [OF S, where f=f and f' = "\z h. deriv f z * h" and g' = g']) using g' * apply (simp_all add: linear_conv_bounded_linear that) using \open S\ has_field_derivative_imp_has_derivative holf holomorphic_derivI by blast qed lemma has_complex_derivative_locally_invertible: assumes holf: "f holomorphic_on S" and S: "\ \ S" "open S" and dnz: "deriv f \ \ 0" obtains r where "r > 0" "ball \ r \ S" "open (f ` (ball \ r))" "inj_on f (ball \ r)" proof - obtain r where "r > 0" "ball \ r \ S" "inj_on f (ball \ r)" by (blast intro: that has_complex_derivative_locally_injective [OF assms]) then have \: "\ \ ball \ r" by simp then have nc: "\ f constant_on ball \ r" using \inj_on f (ball \ r)\ injective_not_constant by fastforce have holf': "f holomorphic_on ball \ r" using \ball \ r \ S\ holf holomorphic_on_subset by blast have "open (f ` ball \ r)" by (simp add: \inj_on f (ball \ r)\ holf' open_mapping_thm3) then show ?thesis using \0 < r\ \ball \ r \ S\ \inj_on f (ball \ r)\ that by blast qed lemma holomorphic_injective_imp_regular: assumes holf: "f holomorphic_on S" and "open S" and injf: "inj_on f S" and "\ \ S" shows "deriv f \ \ 0" proof - obtain r where "r>0" and r: "ball \ r \ S" using assms by (blast elim!: openE) have holf': "f holomorphic_on ball \ r" using \ball \ r \ S\ holf holomorphic_on_subset by blast show ?thesis proof (cases "\n>0. (deriv ^^ n) f \ = 0") case True have fcon: "f w = f \" if "w \ ball \ r" for w by (meson open_ball True \0 < r\ centre_in_ball connected_ball holf' holomorphic_fun_eq_const_on_connected that) have False using fcon [of "\ + r/2"] \0 < r\ r injf unfolding inj_on_def by (metis \\ \ S\ contra_subsetD dist_commute fcon mem_ball perfect_choose_dist) then show ?thesis .. next case False then obtain n0 where n0: "n0 > 0 \ (deriv ^^ n0) f \ \ 0" by blast define n where [abs_def]: "n = (LEAST n. n > 0 \ (deriv ^^ n) f \ \ 0)" have n_ne: "n > 0" "(deriv ^^ n) f \ \ 0" using def_LeastI [OF n_def n0] by auto have n_min: "\k. 0 < k \ k < n \ (deriv ^^ k) f \ = 0" using def_Least_le [OF n_def] not_le by auto obtain g \ where "0 < \" and holg: "g holomorphic_on ball \ \" and fd: "\w. w \ ball \ \ \ f w - f \ = ((w - \) * g w) ^ n" and gnz: "\w. w \ ball \ \ \ g w \ 0" by (blast intro: n_min holomorphic_factor_order_of_zero_strong [OF holf \open S\ \\ \ S\ n_ne]) show ?thesis proof (cases "n=1") case True with n_ne show ?thesis by auto next case False have "g holomorphic_on ball \ (min r \)" using holg by (simp add: holomorphic_on_subset subset_ball) then have holgw: "(\w. (w - \) * g w) holomorphic_on ball \ (min r \)" by (intro holomorphic_intros) have gd: "\w. dist \ w < \ \ (g has_field_derivative deriv g w) (at w)" using holg by (simp add: DERIV_deriv_iff_field_differentiable holomorphic_on_def at_within_open_NO_MATCH) have *: "\w. w \ ball \ (min r \) \ ((\w. (w - \) * g w) has_field_derivative ((w - \) * deriv g w + g w)) (at w)" by (rule gd derivative_eq_intros | simp)+ have [simp]: "deriv (\w. (w - \) * g w) \ \ 0" using * [of \] \0 < \\ \0 < r\ by (simp add: DERIV_imp_deriv gnz) obtain T where "\ \ T" "open T" and Tsb: "T \ ball \ (min r \)" and oimT: "open ((\w. (w - \) * g w) ` T)" using \0 < r\ \0 < \\ has_complex_derivative_locally_invertible [OF holgw, of \] by force define U where "U = (\w. (w - \) * g w) ` T" have "open U" by (metis oimT U_def) moreover have "0 \ U" using \\ \ T\ by (auto simp: U_def intro: image_eqI [where x = \]) ultimately obtain \ where "\>0" and \: "cball 0 \ \ U" using \open U\ open_contains_cball by blast then have "\ * exp(2 * of_real pi * \ * (0/n)) \ cball 0 \" "\ * exp(2 * of_real pi * \ * (1/n)) \ cball 0 \" by (auto simp: norm_mult) with \ have "\ * exp(2 * of_real pi * \ * (0/n)) \ U" "\ * exp(2 * of_real pi * \ * (1/n)) \ U" by blast+ then obtain y0 y1 where "y0 \ T" and y0: "(y0 - \) * g y0 = \ * exp(2 * of_real pi * \ * (0/n))" and "y1 \ T" and y1: "(y1 - \) * g y1 = \ * exp(2 * of_real pi * \ * (1/n))" by (auto simp: U_def) then have "y0 \ ball \ \" "y1 \ ball \ \" using Tsb by auto then have "f y0 - f \ = ((y0 - \) * g y0) ^ n" "f y1 - f \ = ((y1 - \) * g y1) ^ n" using fd by blast+ moreover have "y0 \ y1" using y0 y1 \\ > 0\ complex_root_unity_eq_1 [of n 1] \n > 0\ False by auto moreover have "T \ S" by (meson Tsb min.cobounded1 order_trans r subset_ball) ultimately have False using inj_onD [OF injf, of y0 y1] \y0 \ T\ \y1 \ T\ using complex_root_unity [of n 1] apply (simp add: y0 y1 power_mult_distrib) apply (force simp: algebra_simps) done then show ?thesis .. qed qed qed text\Hence a nice clean inverse function theorem\ lemma has_field_derivative_inverse_strong: fixes f :: "'a::{euclidean_space,real_normed_field} \ 'a" shows "\DERIV f x :> f'; f' \ 0; open S; x \ S; continuous_on S f; \z. z \ S \ g (f z) = z\ \ DERIV g (f x) :> inverse (f')" unfolding has_field_derivative_def by (rule has_derivative_inverse_strong [of S x f g]) auto lemma has_field_derivative_inverse_strong_x: fixes f :: "'a::{euclidean_space,real_normed_field} \ 'a" shows "\DERIV f (g y) :> f'; f' \ 0; open S; continuous_on S f; g y \ S; f(g y) = y; \z. z \ S \ g (f z) = z\ \ DERIV g y :> inverse (f')" unfolding has_field_derivative_def by (rule has_derivative_inverse_strong_x [of S g y f]) auto proposition holomorphic_has_inverse: assumes holf: "f holomorphic_on S" and "open S" and injf: "inj_on f S" obtains g where "g holomorphic_on (f ` S)" "\z. z \ S \ deriv f z * deriv g (f z) = 1" "\z. z \ S \ g(f z) = z" proof - have ofs: "open (f ` S)" by (rule open_mapping_thm3 [OF assms]) have contf: "continuous_on S f" by (simp add: holf holomorphic_on_imp_continuous_on) have *: "(the_inv_into S f has_field_derivative inverse (deriv f z)) (at (f z))" if "z \ S" for z proof - have 1: "(f has_field_derivative deriv f z) (at z)" using DERIV_deriv_iff_field_differentiable \z \ S\ \open S\ holf holomorphic_on_imp_differentiable_at by blast have 2: "deriv f z \ 0" using \z \ S\ \open S\ holf holomorphic_injective_imp_regular injf by blast show ?thesis proof (rule has_field_derivative_inverse_strong [OF 1 2 \open S\ \z \ S\]) show "continuous_on S f" by (simp add: holf holomorphic_on_imp_continuous_on) show "\z. z \ S \ the_inv_into S f (f z) = z" by (simp add: injf the_inv_into_f_f) qed qed show ?thesis proof show "the_inv_into S f holomorphic_on f ` S" by (simp add: holomorphic_on_open ofs) (blast intro: *) next fix z assume "z \ S" have "deriv f z \ 0" using \z \ S\ \open S\ holf holomorphic_injective_imp_regular injf by blast then show "deriv f z * deriv (the_inv_into S f) (f z) = 1" using * [OF \z \ S\] by (simp add: DERIV_imp_deriv) next fix z assume "z \ S" show "the_inv_into S f (f z) = z" by (simp add: \z \ S\ injf the_inv_into_f_f) qed qed subsection\The Schwarz Lemma\ lemma Schwarz1: assumes holf: "f holomorphic_on S" and contf: "continuous_on (closure S) f" and S: "open S" "connected S" and boS: "bounded S" and "S \ {}" obtains w where "w \ frontier S" "\z. z \ closure S \ norm (f z) \ norm (f w)" proof - have connf: "continuous_on (closure S) (norm o f)" using contf continuous_on_compose continuous_on_norm_id by blast have coc: "compact (closure S)" by (simp add: \bounded S\ bounded_closure compact_eq_bounded_closed) then obtain x where x: "x \ closure S" and xmax: "\z. z \ closure S \ norm(f z) \ norm(f x)" using continuous_attains_sup [OF _ _ connf] \S \ {}\ by auto then show ?thesis proof (cases "x \ frontier S") case True then show ?thesis using that xmax by blast next case False then have "x \ S" using \open S\ frontier_def interior_eq x by auto then have "f constant_on S" proof (rule maximum_modulus_principle [OF holf S \open S\ order_refl]) show "\z. z \ S \ cmod (f z) \ cmod (f x)" using closure_subset by (blast intro: xmax) qed then have "f constant_on (closure S)" by (rule constant_on_closureI [OF _ contf]) then obtain c where c: "\x. x \ closure S \ f x = c" by (meson constant_on_def) obtain w where "w \ frontier S" by (metis coc all_not_in_conv assms(6) closure_UNIV frontier_eq_empty not_compact_UNIV) then show ?thesis by (simp add: c frontier_def that) qed qed lemma Schwarz2: "\f holomorphic_on ball 0 r; 0 < s; ball w s \ ball 0 r; \z. norm (w-z) < s \ norm(f z) \ norm(f w)\ \ f constant_on ball 0 r" by (rule maximum_modulus_principle [where U = "ball w s" and \ = w]) (simp_all add: dist_norm) lemma Schwarz3: assumes holf: "f holomorphic_on (ball 0 r)" and [simp]: "f 0 = 0" obtains h where "h holomorphic_on (ball 0 r)" and "\z. norm z < r \ f z = z * (h z)" and "deriv f 0 = h 0" proof - define h where "h z = (if z = 0 then deriv f 0 else f z / z)" for z have d0: "deriv f 0 = h 0" by (simp add: h_def) moreover have "h holomorphic_on (ball 0 r)" by (rule pole_theorem_open_0 [OF holf, of 0]) (auto simp: h_def) moreover have "norm z < r \ f z = z * h z" for z by (simp add: h_def) ultimately show ?thesis using that by blast qed proposition Schwarz_Lemma: assumes holf: "f holomorphic_on (ball 0 1)" and [simp]: "f 0 = 0" and no: "\z. norm z < 1 \ norm (f z) < 1" and \: "norm \ < 1" shows "norm (f \) \ norm \" and "norm(deriv f 0) \ 1" and "((\z. norm z < 1 \ z \ 0 \ norm(f z) = norm z) \ norm(deriv f 0) = 1) \ \\. (\z. norm z < 1 \ f z = \ * z) \ norm \ = 1" (is "?P \ ?Q") proof - obtain h where holh: "h holomorphic_on (ball 0 1)" and fz_eq: "\z. norm z < 1 \ f z = z * (h z)" and df0: "deriv f 0 = h 0" by (rule Schwarz3 [OF holf]) auto have noh_le: "norm (h z) \ 1" if z: "norm z < 1" for z proof - have "norm (h z) < a" if a: "1 < a" for a proof - have "max (inverse a) (norm z) < 1" using z a by (simp_all add: inverse_less_1_iff) then obtain r where r: "max (inverse a) (norm z) < r" and "r < 1" using Rats_dense_in_real by blast then have nzr: "norm z < r" and ira: "inverse r < a" using z a less_imp_inverse_less by force+ then have "0 < r" by (meson norm_not_less_zero not_le order.strict_trans2) have holh': "h holomorphic_on ball 0 r" by (meson holh \r < 1\ holomorphic_on_subset less_eq_real_def subset_ball) have conth': "continuous_on (cball 0 r) h" by (meson \r < 1\ dual_order.trans holh holomorphic_on_imp_continuous_on holomorphic_on_subset mem_ball_0 mem_cball_0 not_less subsetI) obtain w where w: "norm w = r" and lenw: "\z. norm z < r \ norm(h z) \ norm(h w)" apply (rule Schwarz1 [OF holh']) using conth' \0 < r\ by auto have "h w = f w / w" using fz_eq \r < 1\ nzr w by auto then have "cmod (h z) < inverse r" by (metis \0 < r\ \r < 1\ divide_strict_right_mono inverse_eq_divide le_less_trans lenw no norm_divide nzr w) then show ?thesis using ira by linarith qed then show "norm (h z) \ 1" using not_le by blast qed show "cmod (f \) \ cmod \" proof (cases "\ = 0") case True then show ?thesis by auto next case False then show ?thesis by (simp add: noh_le fz_eq \ mult_left_le norm_mult) qed show no_df0: "norm(deriv f 0) \ 1" by (simp add: \\z. cmod z < 1 \ cmod (h z) \ 1\ df0) show "?Q" if "?P" using that proof assume "\z. cmod z < 1 \ z \ 0 \ cmod (f z) = cmod z" then obtain \ where \: "cmod \ < 1" "\ \ 0" "cmod (f \) = cmod \" by blast then have [simp]: "norm (h \) = 1" by (simp add: fz_eq norm_mult) have \
: "ball \ (1 - cmod \) \ ball 0 1" by (simp add: ball_subset_ball_iff) moreover have "\z. cmod (\ - z) < 1 - cmod \ \ cmod (h z) \ cmod (h \)" by (metis \cmod (h \) = 1\ \
dist_0_norm dist_complex_def in_mono mem_ball noh_le) ultimately obtain c where c: "\z. norm z < 1 \ h z = c" using Schwarz2 [OF holh, of "1 - norm \" \, unfolded constant_on_def] \ by auto then have "norm c = 1" using \ by force with c show ?thesis using fz_eq by auto next assume [simp]: "cmod (deriv f 0) = 1" then obtain c where c: "\z. norm z < 1 \ h z = c" using Schwarz2 [OF holh zero_less_one, of 0, unfolded constant_on_def] df0 noh_le by auto moreover have "norm c = 1" using df0 c by auto ultimately show ?thesis using fz_eq by auto qed qed corollary Schwarz_Lemma': assumes holf: "f holomorphic_on (ball 0 1)" and [simp]: "f 0 = 0" and no: "\z. norm z < 1 \ norm (f z) < 1" shows "((\\. norm \ < 1 \ norm (f \) \ norm \) \ norm(deriv f 0) \ 1) \ (((\z. norm z < 1 \ z \ 0 \ norm(f z) = norm z) \ norm(deriv f 0) = 1) \ (\\. (\z. norm z < 1 \ f z = \ * z) \ norm \ = 1))" using Schwarz_Lemma [OF assms] by (metis (no_types) norm_eq_zero zero_less_one) subsection\The Schwarz reflection principle\ lemma hol_pal_lem0: assumes "d \ a \ k" "k \ d \ b" obtains c where "c \ closed_segment a b" "d \ c = k" "\z. z \ closed_segment a c \ d \ z \ k" "\z. z \ closed_segment c b \ k \ d \ z" proof - obtain c where cin: "c \ closed_segment a b" and keq: "k = d \ c" using connected_ivt_hyperplane [of "closed_segment a b" a b d k] by (auto simp: assms) have "closed_segment a c \ {z. d \ z \ k}" "closed_segment c b \ {z. k \ d \ z}" unfolding segment_convex_hull using assms keq by (auto simp: convex_halfspace_le convex_halfspace_ge hull_minimal) then show ?thesis using cin that by fastforce qed lemma hol_pal_lem1: assumes "convex S" "open S" and abc: "a \ S" "b \ S" "c \ S" "d \ 0" and lek: "d \ a \ k" "d \ b \ k" "d \ c \ k" and holf1: "f holomorphic_on {z. z \ S \ d \ z < k}" and contf: "continuous_on S f" shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" proof - have "interior (convex hull {a, b, c}) \ interior(S \ {x. d \ x \ k})" proof (intro interior_mono hull_minimal) show "{a, b, c} \ S \ {x. d \ x \ k}" by (simp add: abc lek) show "convex (S \ {x. d \ x \ k})" by (rule convex_Int [OF \convex S\ convex_halfspace_le]) qed also have "... \ {z \ S. d \ z < k}" by (force simp: interior_open [OF \open S\] \d \ 0\) finally have *: "interior (convex hull {a, b, c}) \ {z \ S. d \ z < k}" . have "continuous_on (convex hull {a,b,c}) f" using \convex S\ contf abc continuous_on_subset subset_hull by fastforce moreover have "f holomorphic_on interior (convex hull {a,b,c})" by (rule holomorphic_on_subset [OF holf1 *]) ultimately show ?thesis using Cauchy_theorem_triangle_interior has_chain_integral_chain_integral3 by blast qed lemma hol_pal_lem2: assumes S: "convex S" "open S" and abc: "a \ S" "b \ S" "c \ S" and "d \ 0" and lek: "d \ a \ k" "d \ b \ k" and holf1: "f holomorphic_on {z. z \ S \ d \ z < k}" and holf2: "f holomorphic_on {z. z \ S \ k < d \ z}" and contf: "continuous_on S f" shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" proof (cases "d \ c \ k") case True show ?thesis by (rule hol_pal_lem1 [OF S abc \d \ 0\ lek True holf1 contf]) next case False then have "d \ c > k" by force obtain a' where a': "a' \ closed_segment b c" and "d \ a' = k" and ba': "\z. z \ closed_segment b a' \ d \ z \ k" and a'c: "\z. z \ closed_segment a' c \ k \ d \ z" using False hol_pal_lem0 [of d b k c, OF \d \ b \ k\] by auto obtain b' where b': "b' \ closed_segment a c" and "d \ b' = k" and ab': "\z. z \ closed_segment a b' \ d \ z \ k" and b'c: "\z. z \ closed_segment b' c \ k \ d \ z" using False hol_pal_lem0 [of d a k c, OF \d \ a \ k\] by auto have a'b': "a' \ S \ b' \ S" using a' abc b' convex_contains_segment \convex S\ by auto have "continuous_on (closed_segment c a) f" by (meson abc contf continuous_on_subset convex_contains_segment \convex S\) then have 1: "contour_integral (linepath c a) f = contour_integral (linepath c b') f + contour_integral (linepath b' a) f" using b' closed_segment_commute contour_integral_split_linepath by blast have "continuous_on (closed_segment b c) f" by (meson abc contf continuous_on_subset convex_contains_segment \convex S\) then have 2: "contour_integral (linepath b c) f = contour_integral (linepath b a') f + contour_integral (linepath a' c) f" by (rule contour_integral_split_linepath [OF _ a']) have 3: "contour_integral (reversepath (linepath b' a')) f = - contour_integral (linepath b' a') f" by (rule contour_integral_reversepath [OF valid_path_linepath]) have fcd_le: "f field_differentiable at x" if "x \ interior S \ x \ interior {x. d \ x \ k}" for x proof - have "f holomorphic_on S \ {c. d \ c < k}" by (metis (no_types) Collect_conj_eq Collect_mem_eq holf1) then have "\C D. x \ interior C \ interior D \ f holomorphic_on interior C \ interior D" using that by (metis Collect_mem_eq Int_Collect \d \ 0\ interior_halfspace_le interior_open \open S\) then show "f field_differentiable at x" by (metis at_within_interior holomorphic_on_def interior_Int interior_interior) qed have ab_le: "\x. x \ closed_segment a b \ d \ x \ k" proof - fix x :: complex assume "x \ closed_segment a b" then have "\C. x \ C \ b \ C \ a \ C \ \ convex C" by (meson contra_subsetD convex_contains_segment) then show "d \ x \ k" by (metis lek convex_halfspace_le mem_Collect_eq) qed have cs: "closed_segment a' b' \ {x. d \ x \ k} \ closed_segment b' a \ {x. d \ x \ k}" by (simp add: \d \ a' = k\ \d \ b' = k\ closed_segment_subset convex_halfspace_le lek(1)) have "continuous_on (S \ {x. d \ x \ k}) f" using contf by (simp add: continuous_on_subset) then have "(f has_contour_integral 0) (linepath a b +++ linepath b a' +++ linepath a' b' +++ linepath b' a)" apply (rule Cauchy_theorem_convex [where K = "{}"]) by (simp_all add: path_image_join convex_Int convex_halfspace_le \convex S\ fcd_le ab_le closed_segment_subset abc a'b' ba' cs) then have 4: "contour_integral (linepath a b) f + contour_integral (linepath b a') f + contour_integral (linepath a' b') f + contour_integral (linepath b' a) f = 0" by (rule has_chain_integral_chain_integral4) have fcd_ge: "f field_differentiable at x" if "x \ interior S \ x \ interior {x. k \ d \ x}" for x proof - have f2: "f holomorphic_on S \ {c. k < d \ c}" by (metis (full_types) Collect_conj_eq Collect_mem_eq holf2) have f3: "interior S = S" by (simp add: interior_open \open S\) then have "x \ S \ interior {c. k \ d \ c}" using that by simp then show "f field_differentiable at x" using f3 f2 unfolding holomorphic_on_def by (metis (no_types) \d \ 0\ at_within_interior interior_Int interior_halfspace_ge interior_interior) qed have cs: "closed_segment c b' \ {x. k \ d \ x} \ closed_segment b' a' \ {x. k \ d \ x}" by (simp add: \d \ a' = k\ b'c closed_segment_subset convex_halfspace_ge) have "continuous_on (S \ {x. k \ d \ x}) f" using contf by (simp add: continuous_on_subset) then have "(f has_contour_integral 0) (linepath a' c +++ linepath c b' +++ linepath b' a')" apply (rule Cauchy_theorem_convex [where K = "{}"]) by (simp_all add: path_image_join convex_Int convex_halfspace_ge \convex S\ fcd_ge closed_segment_subset abc a'b' a'c cs) then have 5: "contour_integral (linepath a' c) f + contour_integral (linepath c b') f + contour_integral (linepath b' a') f = 0" by (rule has_chain_integral_chain_integral3) show ?thesis using 1 2 3 4 5 by (metis add.assoc eq_neg_iff_add_eq_0 reversepath_linepath) qed lemma hol_pal_lem3: assumes S: "convex S" "open S" and abc: "a \ S" "b \ S" "c \ S" and "d \ 0" and lek: "d \ a \ k" and holf1: "f holomorphic_on {z. z \ S \ d \ z < k}" and holf2: "f holomorphic_on {z. z \ S \ k < d \ z}" and contf: "continuous_on S f" shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" proof (cases "d \ b \ k") case True show ?thesis by (rule hol_pal_lem2 [OF S abc \d \ 0\ lek True holf1 holf2 contf]) next case False show ?thesis proof (cases "d \ c \ k") case True have "contour_integral (linepath c a) f + contour_integral (linepath a b) f + contour_integral (linepath b c) f = 0" by (rule hol_pal_lem2 [OF S \c \ S\ \a \ S\ \b \ S\ \d \ 0\ \d \ c \ k\ lek holf1 holf2 contf]) then show ?thesis by (simp add: algebra_simps) next case False have "contour_integral (linepath b c) f + contour_integral (linepath c a) f + contour_integral (linepath a b) f = 0" using hol_pal_lem2 [OF S \b \ S\ \c \ S\ \a \ S\, of "-d" "-k"] using \d \ 0\ \\ d \ b \ k\ False by (simp_all add: holf1 holf2 contf) then show ?thesis by (simp add: algebra_simps) qed qed lemma hol_pal_lem4: assumes S: "convex S" "open S" and abc: "a \ S" "b \ S" "c \ S" and "d \ 0" and holf1: "f holomorphic_on {z. z \ S \ d \ z < k}" and holf2: "f holomorphic_on {z. z \ S \ k < d \ z}" and contf: "continuous_on S f" shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" proof (cases "d \ a \ k") case True show ?thesis by (rule hol_pal_lem3 [OF S abc \d \ 0\ True holf1 holf2 contf]) next case False show ?thesis using \d \ 0\ hol_pal_lem3 [OF S abc, of "-d" "-k"] False by (simp_all add: holf1 holf2 contf) qed lemma holomorphic_on_paste_across_line: assumes S: "open S" and "d \ 0" and holf1: "f holomorphic_on (S \ {z. d \ z < k})" and holf2: "f holomorphic_on (S \ {z. k < d \ z})" and contf: "continuous_on S f" shows "f holomorphic_on S" proof - have *: "\t. open t \ p \ t \ continuous_on t f \ (\a b c. convex hull {a, b, c} \ t \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0)" if "p \ S" for p proof - obtain e where "e>0" and e: "ball p e \ S" using \p \ S\ openE S by blast then have "continuous_on (ball p e) f" using contf continuous_on_subset by blast moreover have "f holomorphic_on {z. dist p z < e \ d \ z < k}" apply (rule holomorphic_on_subset [OF holf1]) using e by auto moreover have "f holomorphic_on {z. dist p z < e \ k < d \ z}" apply (rule holomorphic_on_subset [OF holf2]) using e by auto ultimately show ?thesis apply (rule_tac x="ball p e" in exI) using \e > 0\ e \d \ 0\ hol_pal_lem4 [of "ball p e" _ _ _ d _ k] by (force simp add: subset_hull) qed show ?thesis by (blast intro: * Morera_local_triangle analytic_imp_holomorphic) qed proposition Schwarz_reflection: assumes "open S" and cnjs: "cnj ` S \ S" and holf: "f holomorphic_on (S \ {z. 0 < Im z})" and contf: "continuous_on (S \ {z. 0 \ Im z}) f" and f: "\z. \z \ S; z \ \\ \ (f z) \ \" shows "(\z. if 0 \ Im z then f z else cnj(f(cnj z))) holomorphic_on S" proof - have 1: "(\z. if 0 \ Im z then f z else cnj (f (cnj z))) holomorphic_on (S \ {z. 0 < Im z})" by (force intro: iffD1 [OF holomorphic_cong [OF refl] holf]) have cont_cfc: "continuous_on (S \ {z. Im z \ 0}) (cnj o f o cnj)" using cnjs by (intro continuous_intros continuous_on_compose continuous_on_subset [OF contf]) auto have "cnj \ f \ cnj field_differentiable at x within S \ {z. Im z < 0}" if "x \ S" "Im x < 0" "f field_differentiable at (cnj x) within S \ {z. 0 < Im z}" for x using that apply (clarsimp simp add: field_differentiable_def has_field_derivative_iff Lim_within dist_norm) apply (rule_tac x="cnj f'" in exI) apply (elim all_forward ex_forward conj_forward imp_forward asm_rl, clarify) apply (drule_tac x="cnj xa" in bspec) using cnjs apply force apply (metis complex_cnj_cnj complex_cnj_diff complex_cnj_divide complex_mod_cnj) done then have hol_cfc: "(cnj o f o cnj) holomorphic_on (S \ {z. Im z < 0})" using holf cnjs by (force simp: holomorphic_on_def) have 2: "(\z. if 0 \ Im z then f z else cnj (f (cnj z))) holomorphic_on (S \ {z. Im z < 0})" apply (rule iffD1 [OF holomorphic_cong [OF refl]]) using hol_cfc by auto have [simp]: "(S \ {z. 0 \ Im z}) \ (S \ {z. Im z \ 0}) = S" by force have eq: "\z. \z \ S; Im z \ 0; 0 \ Im z\ \ f z = cnj (f (cnj z))" using f Reals_cnj_iff complex_is_Real_iff by auto have "continuous_on ((S \ {z. 0 \ Im z}) \ (S \ {z. Im z \ 0})) (\z. if 0 \ Im z then f z else cnj (f (cnj z)))" apply (rule continuous_on_cases_local) using cont_cfc contf by (simp_all add: closedin_closed_Int closed_halfspace_Im_le closed_halfspace_Im_ge eq) then have 3: "continuous_on S (\z. if 0 \ Im z then f z else cnj (f (cnj z)))" by force show ?thesis apply (rule holomorphic_on_paste_across_line [OF \open S\, of "- \" _ 0]) using 1 2 3 by auto qed subsection\Bloch's theorem\ lemma Bloch_lemma_0: assumes holf: "f holomorphic_on cball 0 r" and "0 < r" and [simp]: "f 0 = 0" and le: "\z. norm z < r \ norm(deriv f z) \ 2 * norm(deriv f 0)" shows "ball 0 ((3 - 2 * sqrt 2) * r * norm(deriv f 0)) \ f ` ball 0 r" proof - have "sqrt 2 < 3/2" by (rule real_less_lsqrt) (auto simp: power2_eq_square) then have sq3: "0 < 3 - 2 * sqrt 2" by simp show ?thesis proof (cases "deriv f 0 = 0") case True then show ?thesis by simp next case False define C where "C = 2 * norm(deriv f 0)" have "0 < C" using False by (simp add: C_def) have holf': "f holomorphic_on ball 0 r" using holf using ball_subset_cball holomorphic_on_subset by blast then have holdf': "deriv f holomorphic_on ball 0 r" by (rule holomorphic_deriv [OF _ open_ball]) have "Le1": "norm(deriv f z - deriv f 0) \ norm z / (r - norm z) * C" if "norm z < r" for z proof - have T1: "norm(deriv f z - deriv f 0) \ norm z / (R - norm z) * C" if R: "norm z < R" "R < r" for R proof - have "0 < R" using R by (metis less_trans norm_zero zero_less_norm_iff) have df_le: "\x. norm x < r \ norm (deriv f x) \ C" using le by (simp add: C_def) have hol_df: "deriv f holomorphic_on cball 0 R" using R holdf' holomorphic_on_subset by auto have *: "((\w. deriv f w / (w - z)) has_contour_integral 2 * pi * \ * deriv f z) (circlepath 0 R)" if "norm z < R" for z using \0 < R\ that Cauchy_integral_formula_convex_simple [OF convex_cball hol_df, of _ "circlepath 0 R"] by (force simp: winding_number_circlepath) have **: "((\x. deriv f x / (x - z) - deriv f x / x) has_contour_integral of_real (2 * pi) * \ * (deriv f z - deriv f 0)) (circlepath 0 R)" using has_contour_integral_diff [OF * [of z] * [of 0]] \0 < R\ that by (simp add: algebra_simps) have [simp]: "\x. norm x = R \ x \ z" using that(1) by blast have "norm (deriv f x / (x - z) - deriv f x / x) \ C * norm z / (R * (R - norm z))" if "norm x = R" for x proof - have [simp]: "norm (deriv f x * x - deriv f x * (x - z)) = norm (deriv f x) * norm z" by (simp add: norm_mult right_diff_distrib') show ?thesis using \0 < R\ \0 < C\ R that by (auto simp add: norm_mult norm_divide divide_simps df_le mult_mono norm_triangle_ineq2) qed then show ?thesis using has_contour_integral_bound_circlepath [OF **, of "C * norm z/(R*(R - norm z))"] \0 < R\ \0 < C\ R apply (simp add: norm_mult norm_divide) apply (simp add: divide_simps mult.commute) done qed obtain r' where r': "norm z < r'" "r' < r" using Rats_dense_in_real [of "norm z" r] \norm z < r\ by blast then have [simp]: "closure {r'<.. norm(f z)" if r: "norm z < r" for z proof - have 1: "\x. x \ ball 0 r \ ((\z. f z - deriv f 0 * z) has_field_derivative deriv f x - deriv f 0) (at x within ball 0 r)" by (rule derivative_eq_intros holomorphic_derivI holf' | simp)+ have 2: "closed_segment 0 z \ ball 0 r" by (metis \0 < r\ convex_ball convex_contains_segment dist_self mem_ball mem_ball_0 that) have 4: "norm (deriv f (x *\<^sub>R z) - deriv f 0) * norm z \ norm z * norm z * x * C / (r - norm z)" if x: "0 \ x" "x \ 1" for x proof - have [simp]: "x * norm z < r" using r x by (meson le_less_trans mult_le_cancel_right2 norm_not_less_zero) have "norm (deriv f (x *\<^sub>R z) - deriv f 0) \ norm (x *\<^sub>R z) / (r - norm (x *\<^sub>R z)) * C" apply (rule Le1) using r x \0 < r\ by simp also have "... \ norm (x *\<^sub>R z) / (r - norm z) * C" using r x \0 < r\ apply (simp add: field_split_simps) by (simp add: \0 < C\ mult.assoc mult_left_le_one_le ordered_comm_semiring_class.comm_mult_left_mono) finally have "norm (deriv f (x *\<^sub>R z) - deriv f 0) * norm z \ norm (x *\<^sub>R z) / (r - norm z) * C * norm z" by (rule mult_right_mono) simp with x show ?thesis by (simp add: algebra_simps) qed have le_norm: "abc \ norm d - e \ norm(f - d) \ e \ abc \ norm f" for abc d e and f::complex by (metis add_diff_cancel_left' add_diff_eq diff_left_mono norm_diff_ineq order_trans) have "norm (integral {0..1} (\x. (deriv f (x *\<^sub>R z) - deriv f 0) * z)) \ integral {0..1} (\t. (norm z)\<^sup>2 * t / (r - norm z) * C)" proof (rule integral_norm_bound_integral) show "(\x. (deriv f (x *\<^sub>R z) - deriv f 0) * z) integrable_on {0..1}" using contour_integral_primitive [OF 1, of "linepath 0 z"] 2 by (simp add: has_contour_integral_linepath has_integral_integrable_integral) have "(*) ((cmod z)\<^sup>2) integrable_on {0..1}" by (metis ident_integrable_on integrable_0 integrable_eq integrable_on_cmult_iff lambda_zero) then show "(\t. (norm z)\<^sup>2 * t / (r - norm z) * C) integrable_on {0..1}" using integrable_on_cmult_right[where 'b=real, simplified] integrable_on_cdivide [where 'b=real, simplified] by blast qed (simp add: norm_mult power2_eq_square 4) then have int_le: "norm (f z - deriv f 0 * z) \ (norm z)\<^sup>2 * norm(deriv f 0) / ((r - norm z))" using contour_integral_primitive [OF 1, of "linepath 0 z"] 2 by (simp add: has_contour_integral_linepath has_integral_integrable_integral C_def) have "norm z * (norm (deriv f 0) * (r - norm z - norm z)) \ norm z * (norm (deriv f 0) * (r - norm z) - norm (deriv f 0) * norm z)" by (simp add: algebra_simps) then have \
: "(norm z * (r - norm z) - norm z * norm z) * norm (deriv f 0) \ norm (deriv f 0) * norm z * (r - norm z) - norm z * norm z * norm (deriv f 0)" by (simp add: algebra_simps) show ?thesis apply (rule le_norm [OF _ int_le]) using \norm z < r\ by (simp add: power2_eq_square divide_simps C_def norm_mult \
) qed have sq201 [simp]: "0 < (1 - sqrt 2 / 2)" "(1 - sqrt 2 / 2) < 1" by (auto simp: sqrt2_less_2) have 1: "continuous_on (closure (ball 0 ((1 - sqrt 2 / 2) * r))) f" proof (rule continuous_on_subset [OF holomorphic_on_imp_continuous_on [OF holf]]) show "closure (ball 0 ((1 - sqrt 2 / 2) * r)) \ cball 0 r" proof - have "(1 - sqrt 2 / 2) * r \ r" by (simp add: \0 < r\) then show ?thesis by (meson ball_subset_cball closed_cball closure_minimal dual_order.trans subset_ball) qed qed have 2: "open (f ` interior (ball 0 ((1 - sqrt 2 / 2) * r)))" proof (rule open_mapping_thm [OF holf' open_ball connected_ball]) show "interior (ball 0 ((1 - sqrt 2 / 2) * r)) \ ball (0::complex) r" using \0 < r\ mult_pos_pos sq201 by (simp add: ball_subset_ball_iff) show "\ f constant_on ball 0 r" using False \0 < r\ centre_in_ball holf' holomorphic_nonconstant by blast qed auto have "ball 0 ((3 - 2 * sqrt 2) * r * norm (deriv f 0)) = ball (f 0) ((3 - 2 * sqrt 2) * r * norm (deriv f 0))" by simp also have "... \ f ` ball 0 ((1 - sqrt 2 / 2) * r)" proof - have 3: "(3 - 2 * sqrt 2) * r * norm (deriv f 0) \ norm (f z)" if "norm z = (1 - sqrt 2 / 2) * r" for z apply (rule order_trans [OF _ *]) using \0 < r\ apply (simp_all add: field_simps power2_eq_square that) apply (simp add: mult.assoc [symmetric]) done show ?thesis apply (rule ball_subset_open_map_image [OF 1 2 _ bounded_ball]) using \0 < r\ sq201 3 C_def \0 < C\ sq3 by auto qed also have "... \ f ` ball 0 r" proof - have "\x. (1 - sqrt 2 / 2) * r \ r" using \0 < r\ by (auto simp: field_simps) then show ?thesis by auto qed finally show ?thesis . qed qed lemma Bloch_lemma: assumes holf: "f holomorphic_on cball a r" and "0 < r" and le: "\z. z \ ball a r \ norm(deriv f z) \ 2 * norm(deriv f a)" shows "ball (f a) ((3 - 2 * sqrt 2) * r * norm(deriv f a)) \ f ` ball a r" (is "?lhs \ ?rhs") proof - have fz: "(\z. f (a + z)) = f o (\z. (a + z))" by (simp add: o_def) have hol0: "(\z. f (a + z)) holomorphic_on cball 0 r" unfolding fz by (intro holomorphic_intros holf holomorphic_on_compose | simp)+ then have [simp]: "\x. norm x < r \ (\z. f (a + z)) field_differentiable at x" by (metis open_ball at_within_open ball_subset_cball diff_0 dist_norm holomorphic_on_def holomorphic_on_subset mem_ball norm_minus_cancel) have [simp]: "\z. norm z < r \ f field_differentiable at (a + z)" by (metis holf open_ball add_diff_cancel_left' dist_complex_def holomorphic_on_imp_differentiable_at holomorphic_on_subset interior_cball interior_subset mem_ball norm_minus_commute) then have [simp]: "f field_differentiable at a" by (metis add.comm_neutral \0 < r\ norm_eq_zero) have hol1: "(\z. f (a + z) - f a) holomorphic_on cball 0 r" by (intro holomorphic_intros hol0) then have \
: "ball 0 ((3 - 2 * sqrt 2) * r * norm (deriv (\z. f (a + z) - f a) 0)) \ (\z. f (a + z) - f a) ` ball 0 r" apply (rule Bloch_lemma_0) using \0 < r\ apply (simp_all add: \0 < r\) apply (simp add: fz deriv_chain dist_norm le) done show ?thesis proof clarify fix x assume "x \ ?lhs" with subsetD [OF \
, of "x - f a"] show "x \ ?rhs" by (force simp: fz \0 < r\ dist_norm deriv_chain field_differentiable_compose) qed qed proposition Bloch_unit: assumes holf: "f holomorphic_on ball a 1" and [simp]: "deriv f a = 1" obtains b r where "1/12 < r" and "ball b r \ f ` (ball a 1)" proof - define r :: real where "r = 249/256" have "0 < r" "r < 1" by (auto simp: r_def) define g where "g z = deriv f z * of_real(r - norm(z - a))" for z have "deriv f holomorphic_on ball a 1" by (rule holomorphic_deriv [OF holf open_ball]) then have "continuous_on (ball a 1) (deriv f)" using holomorphic_on_imp_continuous_on by blast then have "continuous_on (cball a r) (deriv f)" by (rule continuous_on_subset) (simp add: cball_subset_ball_iff \r < 1\) then have "continuous_on (cball a r) g" by (simp add: g_def continuous_intros) then have 1: "compact (g ` cball a r)" by (rule compact_continuous_image [OF _ compact_cball]) have 2: "g ` cball a r \ {}" using \r > 0\ by auto obtain p where pr: "p \ cball a r" and pge: "\y. y \ cball a r \ norm (g y) \ norm (g p)" using distance_attains_sup [OF 1 2, of 0] by force define t where "t = (r - norm(p - a)) / 2" have "norm (p - a) \ r" using pge [of a] \r > 0\ by (auto simp: g_def norm_mult) then have "norm (p - a) < r" using pr by (simp add: norm_minus_commute dist_norm) then have "0 < t" by (simp add: t_def) have cpt: "cball p t \ ball a r" using \0 < t\ by (simp add: cball_subset_ball_iff dist_norm t_def field_simps) have gen_le_dfp: "norm (deriv f y) * (r - norm (y - a)) / (r - norm (p - a)) \ norm (deriv f p)" if "y \ cball a r" for y proof - have [simp]: "norm (y - a) \ r" using that by (simp add: dist_norm norm_minus_commute) have "norm (g y) \ norm (g p)" using pge [OF that] by simp then have "norm (deriv f y) * abs (r - norm (y - a)) \ norm (deriv f p) * abs (r - norm (p - a))" by (simp only: dist_norm g_def norm_mult norm_of_real) with that \norm (p - a) < r\ show ?thesis by (simp add: dist_norm field_split_simps) qed have le_norm_dfp: "r / (r - norm (p - a)) \ norm (deriv f p)" using gen_le_dfp [of a] \r > 0\ by auto have 1: "f holomorphic_on cball p t" using cpt \r < 1\ order_subst1 subset_ball by (force simp add: intro!: holomorphic_on_subset [OF holf]) have 2: "norm (deriv f z) \ 2 * norm (deriv f p)" if "z \ ball p t" for z proof - have z: "z \ cball a r" by (meson ball_subset_cball subsetD cpt that) then have "norm(z - a) < r" by (metis ball_subset_cball contra_subsetD cpt dist_norm mem_ball norm_minus_commute that) have "norm (deriv f z) * (r - norm (z - a)) / (r - norm (p - a)) \ norm (deriv f p)" using gen_le_dfp [OF z] by simp with \norm (z - a) < r\ \norm (p - a) < r\ have "norm (deriv f z) \ (r - norm (p - a)) / (r - norm (z - a)) * norm (deriv f p)" by (simp add: field_simps) also have "... \ 2 * norm (deriv f p)" proof (rule mult_right_mono) show "(r - cmod (p - a)) / (r - cmod (z - a)) \ 2" using that \norm (p - a) < r\ \norm(z - a) < r\ dist_triangle3 [of z a p] by (simp add: field_simps t_def dist_norm [symmetric]) qed auto finally show ?thesis . qed have sqrt2: "sqrt 2 < 2113/1494" by (rule real_less_lsqrt) (auto simp: power2_eq_square) then have sq3: "0 < 3 - 2 * sqrt 2" by simp have "1 / 12 / ((3 - 2 * sqrt 2) / 2) < r" using sq3 sqrt2 by (auto simp: field_simps r_def) also have "... \ cmod (deriv f p) * (r - cmod (p - a))" using \norm (p - a) < r\ le_norm_dfp by (simp add: pos_divide_le_eq) finally have "1 / 12 < cmod (deriv f p) * (r - cmod (p - a)) * ((3 - 2 * sqrt 2) / 2)" using pos_divide_less_eq half_gt_zero_iff sq3 by blast then have **: "1 / 12 < (3 - 2 * sqrt 2) * t * norm (deriv f p)" using sq3 by (simp add: mult.commute t_def) have "ball (f p) ((3 - 2 * sqrt 2) * t * norm (deriv f p)) \ f ` ball p t" by (rule Bloch_lemma [OF 1 \0 < t\ 2]) also have "... \ f ` ball a 1" proof - have "ball a r \ ball a 1" using \0 < t\ \r < 1\ by (simp add: ball_subset_ball_iff dist_norm) then show ?thesis using ball_subset_cball cpt by blast qed finally have "ball (f p) ((3 - 2 * sqrt 2) * t * norm (deriv f p)) \ f ` ball a 1" . with ** show ?thesis by (rule that) qed theorem Bloch: assumes holf: "f holomorphic_on ball a r" and "0 < r" and r': "r' \ r * norm (deriv f a) / 12" obtains b where "ball b r' \ f ` (ball a r)" proof (cases "deriv f a = 0") case True with r' show ?thesis using ball_eq_empty that by fastforce next case False define C where "C = deriv f a" have "0 < norm C" using False by (simp add: C_def) have dfa: "f field_differentiable at a" using \0 < r\ holomorphic_on_imp_differentiable_at [OF holf] by auto have fo: "(\z. f (a + of_real r * z)) = f o (\z. (a + of_real r * z))" by (simp add: o_def) have holf': "f holomorphic_on (\z. a + complex_of_real r * z) ` ball 0 1" using \0 < r\ holomorphic_on_subset [OF holf] by (force simp: dist_norm norm_mult) have 1: "(\z. f (a + r * z) / (C * r)) holomorphic_on ball 0 1" using \0 < r\ \0 < norm C\ by (intro holomorphic_intros holomorphic_on_compose holf'; simp add: fo)+ have "((\z. f (a + of_real r * z) / (C * of_real r)) has_field_derivative (deriv f (a + of_real r * z) / C)) (at z)" if "norm z < 1" for z proof - have fd: "f field_differentiable at (a + complex_of_real r * z)" using \0 < r\ by (simp_all add: dist_norm norm_mult holomorphic_on_imp_differentiable_at [OF holf] that) have *: "((\x. f (a + of_real r * x)) has_field_derivative (deriv f (a + of_real r * z) * of_real r)) (at z)" by (rule fd DERIV_chain [OF field_differentiable_derivI]derivative_eq_intros | simp add: fo)+ show ?thesis apply (rule derivative_eq_intros * | simp)+ using \0 < r\ by (auto simp: C_def False) qed have "deriv (\z. f (a + of_real r * z) / (C * of_real r)) 0 = deriv (\z. f (a + complex_of_real r * z)) 0 / (C * complex_of_real r)" apply (rule deriv_cdivide_right) by (metis (no_types) DERIV_chain2 add.right_neutral dfa field_differentiable_add_const field_differentiable_def field_differentiable_linear fo mult_zero_right) also have "... = 1" using \0 < r\ by (simp add: C_def False fo derivative_intros dfa deriv_chain) finally have 2: "deriv (\z. f (a + of_real r * z) / (C * of_real r)) 0 = 1" . have sb1: "(*) (C * r) ` (\z. f (a + of_real r * z) / (C * r)) ` ball 0 1 \ f ` ball a r" using \0 < r\ by (auto simp: dist_norm norm_mult C_def False) have sb2: "ball (C * r * b) r' \ (*) (C * r) ` ball b t" if "1 / 12 < t" for b t proof - have *: "r * cmod (deriv f a) / 12 \ r * (t * cmod (deriv f a))" using that \0 < r\ less_eq_real_def mult.commute mult.right_neutral mult_left_mono norm_ge_zero times_divide_eq_right by auto show ?thesis apply clarify apply (rule_tac x="x / (C * r)" in image_eqI) using \0 < r\ apply (simp_all add: dist_norm norm_mult norm_divide C_def False field_simps) using "*" r' by linarith qed show ?thesis apply (rule Bloch_unit [OF 1 2]) apply (rule_tac b="(C * of_real r) * b" in that) using image_mono sb1 sb2 by fastforce qed corollary Bloch_general: assumes holf: "f holomorphic_on S" and "a \ S" and tle: "\z. z \ frontier S \ t \ dist a z" and rle: "r \ t * norm(deriv f a) / 12" obtains b where "ball b r \ f ` S" proof - consider "r \ 0" | "0 < t * norm(deriv f a) / 12" using rle by force then show ?thesis proof cases case 1 then show ?thesis by (simp add: ball_empty that) next case 2 show ?thesis proof (cases "deriv f a = 0") case True then show ?thesis using rle by (simp add: ball_empty that) next case False then have "t > 0" using 2 by (force simp: zero_less_mult_iff) have "\ ball a t \ S \ ball a t \ frontier S \ {}" by (metis Diff_eq_empty_iff \0 < t\ \a \ S\ closure_Int_ball_not_empty closure_subset connected_Int_frontier connected_ball inf.commute) with tle have *: "ball a t \ S" by fastforce then have 1: "f holomorphic_on ball a t" using holf using holomorphic_on_subset by blast show ?thesis apply (rule Bloch [OF 1 \t > 0\ rle]) apply (rule_tac b=b in that) using * apply force done qed qed qed end diff --git a/src/HOL/Conditionally_Complete_Lattices.thy b/src/HOL/Conditionally_Complete_Lattices.thy --- a/src/HOL/Conditionally_Complete_Lattices.thy +++ b/src/HOL/Conditionally_Complete_Lattices.thy @@ -1,799 +1,818 @@ (* Title: HOL/Conditionally_Complete_Lattices.thy Author: Amine Chaieb and L C Paulson, University of Cambridge Author: Johannes Hölzl, TU München Author: Luke S. Serafin, Carnegie Mellon University *) section \Conditionally-complete Lattices\ theory Conditionally_Complete_Lattices imports Finite_Set Lattices_Big Set_Interval begin locale preordering_bdd = preordering begin definition bdd :: \'a set \ bool\ where unfold: \bdd A \ (\M. \x \ A. x \<^bold>\ M)\ lemma empty [simp, intro]: \bdd {}\ by (simp add: unfold) lemma I [intro]: \bdd A\ if \\x. x \ A \ x \<^bold>\ M\ using that by (auto simp add: unfold) lemma E: assumes \bdd A\ obtains M where \\x. x \ A \ x \<^bold>\ M\ using assms that by (auto simp add: unfold) lemma I2: \bdd (f ` A)\ if \\x. x \ A \ f x \<^bold>\ M\ using that by (auto simp add: unfold) lemma mono: \bdd A\ if \bdd B\ \A \ B\ using that by (auto simp add: unfold) lemma Int1 [simp]: \bdd (A \ B)\ if \bdd A\ using mono that by auto lemma Int2 [simp]: \bdd (A \ B)\ if \bdd B\ using mono that by auto end context preorder begin sublocale bdd_above: preordering_bdd \(\)\ \(<)\ defines bdd_above_primitive_def: bdd_above = bdd_above.bdd .. sublocale bdd_below: preordering_bdd \(\)\ \(>)\ defines bdd_below_primitive_def: bdd_below = bdd_below.bdd .. lemma bdd_above_def: \bdd_above A \ (\M. \x \ A. x \ M)\ by (fact bdd_above.unfold) lemma bdd_below_def: \bdd_below A \ (\M. \x \ A. M \ x)\ by (fact bdd_below.unfold) lemma bdd_aboveI: "(\x. x \ A \ x \ M) \ bdd_above A" by (fact bdd_above.I) lemma bdd_belowI: "(\x. x \ A \ m \ x) \ bdd_below A" by (fact bdd_below.I) lemma bdd_aboveI2: "(\x. x \ A \ f x \ M) \ bdd_above (f`A)" by (fact bdd_above.I2) lemma bdd_belowI2: "(\x. x \ A \ m \ f x) \ bdd_below (f`A)" by (fact bdd_below.I2) lemma bdd_above_empty: "bdd_above {}" by (fact bdd_above.empty) lemma bdd_below_empty: "bdd_below {}" by (fact bdd_below.empty) lemma bdd_above_mono: "bdd_above B \ A \ B \ bdd_above A" by (fact bdd_above.mono) lemma bdd_below_mono: "bdd_below B \ A \ B \ bdd_below A" by (fact bdd_below.mono) lemma bdd_above_Int1: "bdd_above A \ bdd_above (A \ B)" by (fact bdd_above.Int1) lemma bdd_above_Int2: "bdd_above B \ bdd_above (A \ B)" by (fact bdd_above.Int2) lemma bdd_below_Int1: "bdd_below A \ bdd_below (A \ B)" by (fact bdd_below.Int1) lemma bdd_below_Int2: "bdd_below B \ bdd_below (A \ B)" by (fact bdd_below.Int2) lemma bdd_above_Ioo [simp, intro]: "bdd_above {a <..< b}" by (auto simp add: bdd_above_def intro!: exI[of _ b] less_imp_le) lemma bdd_above_Ico [simp, intro]: "bdd_above {a ..< b}" by (auto simp add: bdd_above_def intro!: exI[of _ b] less_imp_le) lemma bdd_above_Iio [simp, intro]: "bdd_above {..< b}" by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le) lemma bdd_above_Ioc [simp, intro]: "bdd_above {a <.. b}" by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le) lemma bdd_above_Icc [simp, intro]: "bdd_above {a .. b}" by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le) lemma bdd_above_Iic [simp, intro]: "bdd_above {.. b}" by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le) lemma bdd_below_Ioo [simp, intro]: "bdd_below {a <..< b}" by (auto simp add: bdd_below_def intro!: exI[of _ a] less_imp_le) lemma bdd_below_Ioc [simp, intro]: "bdd_below {a <.. b}" by (auto simp add: bdd_below_def intro!: exI[of _ a] less_imp_le) lemma bdd_below_Ioi [simp, intro]: "bdd_below {a <..}" by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le) lemma bdd_below_Ico [simp, intro]: "bdd_below {a ..< b}" by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le) lemma bdd_below_Icc [simp, intro]: "bdd_below {a .. b}" by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le) lemma bdd_below_Ici [simp, intro]: "bdd_below {a ..}" by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le) end context order_top begin lemma bdd_above_top [simp, intro!]: "bdd_above A" by (rule bdd_aboveI [of _ top]) simp end context order_bot begin lemma bdd_below_bot [simp, intro!]: "bdd_below A" by (rule bdd_belowI [of _ bot]) simp end lemma bdd_above_image_mono: "mono f \ bdd_above A \ bdd_above (f`A)" by (auto simp: bdd_above_def mono_def) lemma bdd_below_image_mono: "mono f \ bdd_below A \ bdd_below (f`A)" by (auto simp: bdd_below_def mono_def) lemma bdd_above_image_antimono: "antimono f \ bdd_below A \ bdd_above (f`A)" by (auto simp: bdd_above_def bdd_below_def antimono_def) lemma bdd_below_image_antimono: "antimono f \ bdd_above A \ bdd_below (f`A)" by (auto simp: bdd_above_def bdd_below_def antimono_def) lemma fixes X :: "'a::ordered_ab_group_add set" shows bdd_above_uminus[simp]: "bdd_above (uminus ` X) \ bdd_below X" and bdd_below_uminus[simp]: "bdd_below (uminus ` X) \ bdd_above X" using bdd_above_image_antimono[of uminus X] bdd_below_image_antimono[of uminus "uminus`X"] using bdd_below_image_antimono[of uminus X] bdd_above_image_antimono[of uminus "uminus`X"] by (auto simp: antimono_def image_image) context lattice begin lemma bdd_above_insert [simp]: "bdd_above (insert a A) = bdd_above A" by (auto simp: bdd_above_def intro: le_supI2 sup_ge1) lemma bdd_below_insert [simp]: "bdd_below (insert a A) = bdd_below A" by (auto simp: bdd_below_def intro: le_infI2 inf_le1) lemma bdd_finite [simp]: assumes "finite A" shows bdd_above_finite: "bdd_above A" and bdd_below_finite: "bdd_below A" using assms by (induct rule: finite_induct, auto) lemma bdd_above_Un [simp]: "bdd_above (A \ B) = (bdd_above A \ bdd_above B)" proof assume "bdd_above (A \ B)" thus "bdd_above A \ bdd_above B" unfolding bdd_above_def by auto next assume "bdd_above A \ bdd_above B" then obtain a b where "\x\A. x \ a" "\x\B. x \ b" unfolding bdd_above_def by auto hence "\x \ A \ B. x \ sup a b" by (auto intro: Un_iff le_supI1 le_supI2) thus "bdd_above (A \ B)" unfolding bdd_above_def .. qed lemma bdd_below_Un [simp]: "bdd_below (A \ B) = (bdd_below A \ bdd_below B)" proof assume "bdd_below (A \ B)" thus "bdd_below A \ bdd_below B" unfolding bdd_below_def by auto next assume "bdd_below A \ bdd_below B" then obtain a b where "\x\A. a \ x" "\x\B. b \ x" unfolding bdd_below_def by auto hence "\x \ A \ B. inf a b \ x" by (auto intro: Un_iff le_infI1 le_infI2) thus "bdd_below (A \ B)" unfolding bdd_below_def .. qed lemma bdd_above_image_sup[simp]: "bdd_above ((\x. sup (f x) (g x)) ` A) \ bdd_above (f`A) \ bdd_above (g`A)" by (auto simp: bdd_above_def intro: le_supI1 le_supI2) lemma bdd_below_image_inf[simp]: "bdd_below ((\x. inf (f x) (g x)) ` A) \ bdd_below (f`A) \ bdd_below (g`A)" by (auto simp: bdd_below_def intro: le_infI1 le_infI2) lemma bdd_below_UN[simp]: "finite I \ bdd_below (\i\I. A i) = (\i \ I. bdd_below (A i))" by (induction I rule: finite.induct) auto lemma bdd_above_UN[simp]: "finite I \ bdd_above (\i\I. A i) = (\i \ I. bdd_above (A i))" by (induction I rule: finite.induct) auto end text \ To avoid name classes with the \<^class>\complete_lattice\-class we prefix \<^const>\Sup\ and \<^const>\Inf\ in theorem names with c. \ class conditionally_complete_lattice = lattice + Sup + Inf + assumes cInf_lower: "x \ X \ bdd_below X \ Inf X \ x" and cInf_greatest: "X \ {} \ (\x. x \ X \ z \ x) \ z \ Inf X" assumes cSup_upper: "x \ X \ bdd_above X \ x \ Sup X" and cSup_least: "X \ {} \ (\x. x \ X \ x \ z) \ Sup X \ z" begin lemma cSup_upper2: "x \ X \ y \ x \ bdd_above X \ y \ Sup X" by (metis cSup_upper order_trans) lemma cInf_lower2: "x \ X \ x \ y \ bdd_below X \ Inf X \ y" by (metis cInf_lower order_trans) lemma cSup_mono: "B \ {} \ bdd_above A \ (\b. b \ B \ \a\A. b \ a) \ Sup B \ Sup A" by (metis cSup_least cSup_upper2) lemma cInf_mono: "B \ {} \ bdd_below A \ (\b. b \ B \ \a\A. a \ b) \ Inf A \ Inf B" by (metis cInf_greatest cInf_lower2) lemma cSup_subset_mono: "A \ {} \ bdd_above B \ A \ B \ Sup A \ Sup B" by (metis cSup_least cSup_upper subsetD) lemma cInf_superset_mono: "A \ {} \ bdd_below B \ A \ B \ Inf B \ Inf A" by (metis cInf_greatest cInf_lower subsetD) lemma cSup_eq_maximum: "z \ X \ (\x. x \ X \ x \ z) \ Sup X = z" by (intro order.antisym cSup_upper[of z X] cSup_least[of X z]) auto lemma cInf_eq_minimum: "z \ X \ (\x. x \ X \ z \ x) \ Inf X = z" by (intro order.antisym cInf_lower[of z X] cInf_greatest[of X z]) auto lemma cSup_le_iff: "S \ {} \ bdd_above S \ Sup S \ a \ (\x\S. x \ a)" by (metis order_trans cSup_upper cSup_least) lemma le_cInf_iff: "S \ {} \ bdd_below S \ a \ Inf S \ (\x\S. a \ x)" by (metis order_trans cInf_lower cInf_greatest) lemma cSup_eq_non_empty: assumes 1: "X \ {}" assumes 2: "\x. x \ X \ x \ a" assumes 3: "\y. (\x. x \ X \ x \ y) \ a \ y" shows "Sup X = a" by (intro 3 1 order.antisym cSup_least) (auto intro: 2 1 cSup_upper) lemma cInf_eq_non_empty: assumes 1: "X \ {}" assumes 2: "\x. x \ X \ a \ x" assumes 3: "\y. (\x. x \ X \ y \ x) \ y \ a" shows "Inf X = a" by (intro 3 1 order.antisym cInf_greatest) (auto intro: 2 1 cInf_lower) lemma cInf_cSup: "S \ {} \ bdd_below S \ Inf S = Sup {x. \s\S. x \ s}" by (rule cInf_eq_non_empty) (auto intro!: cSup_upper cSup_least simp: bdd_below_def) lemma cSup_cInf: "S \ {} \ bdd_above S \ Sup S = Inf {x. \s\S. s \ x}" by (rule cSup_eq_non_empty) (auto intro!: cInf_lower cInf_greatest simp: bdd_above_def) lemma cSup_insert: "X \ {} \ bdd_above X \ Sup (insert a X) = sup a (Sup X)" by (intro cSup_eq_non_empty) (auto intro: le_supI2 cSup_upper cSup_least) lemma cInf_insert: "X \ {} \ bdd_below X \ Inf (insert a X) = inf a (Inf X)" by (intro cInf_eq_non_empty) (auto intro: le_infI2 cInf_lower cInf_greatest) lemma cSup_singleton [simp]: "Sup {x} = x" by (intro cSup_eq_maximum) auto lemma cInf_singleton [simp]: "Inf {x} = x" by (intro cInf_eq_minimum) auto lemma cSup_insert_If: "bdd_above X \ Sup (insert a X) = (if X = {} then a else sup a (Sup X))" using cSup_insert[of X] by simp lemma cInf_insert_If: "bdd_below X \ Inf (insert a X) = (if X = {} then a else inf a (Inf X))" using cInf_insert[of X] by simp lemma le_cSup_finite: "finite X \ x \ X \ x \ Sup X" proof (induct X arbitrary: x rule: finite_induct) case (insert x X y) then show ?case by (cases "X = {}") (auto simp: cSup_insert intro: le_supI2) qed simp lemma cInf_le_finite: "finite X \ x \ X \ Inf X \ x" proof (induct X arbitrary: x rule: finite_induct) case (insert x X y) then show ?case by (cases "X = {}") (auto simp: cInf_insert intro: le_infI2) qed simp lemma cSup_eq_Sup_fin: "finite X \ X \ {} \ Sup X = Sup_fin X" by (induct X rule: finite_ne_induct) (simp_all add: cSup_insert) lemma cInf_eq_Inf_fin: "finite X \ X \ {} \ Inf X = Inf_fin X" by (induct X rule: finite_ne_induct) (simp_all add: cInf_insert) lemma cSup_atMost[simp]: "Sup {..x} = x" by (auto intro!: cSup_eq_maximum) lemma cSup_greaterThanAtMost[simp]: "y < x \ Sup {y<..x} = x" by (auto intro!: cSup_eq_maximum) lemma cSup_atLeastAtMost[simp]: "y \ x \ Sup {y..x} = x" by (auto intro!: cSup_eq_maximum) lemma cInf_atLeast[simp]: "Inf {x..} = x" by (auto intro!: cInf_eq_minimum) lemma cInf_atLeastLessThan[simp]: "y < x \ Inf {y.. x \ Inf {y..x} = y" by (auto intro!: cInf_eq_minimum) lemma cINF_lower: "bdd_below (f ` A) \ x \ A \ \(f ` A) \ f x" using cInf_lower [of _ "f ` A"] by simp lemma cINF_greatest: "A \ {} \ (\x. x \ A \ m \ f x) \ m \ \(f ` A)" using cInf_greatest [of "f ` A"] by auto lemma cSUP_upper: "x \ A \ bdd_above (f ` A) \ f x \ \(f ` A)" using cSup_upper [of _ "f ` A"] by simp lemma cSUP_least: "A \ {} \ (\x. x \ A \ f x \ M) \ \(f ` A) \ M" using cSup_least [of "f ` A"] by auto lemma cINF_lower2: "bdd_below (f ` A) \ x \ A \ f x \ u \ \(f ` A) \ u" by (auto intro: cINF_lower order_trans) lemma cSUP_upper2: "bdd_above (f ` A) \ x \ A \ u \ f x \ u \ \(f ` A)" by (auto intro: cSUP_upper order_trans) lemma cSUP_const [simp]: "A \ {} \ (\x\A. c) = c" by (intro order.antisym cSUP_least) (auto intro: cSUP_upper) lemma cINF_const [simp]: "A \ {} \ (\x\A. c) = c" by (intro order.antisym cINF_greatest) (auto intro: cINF_lower) lemma le_cINF_iff: "A \ {} \ bdd_below (f ` A) \ u \ \(f ` A) \ (\x\A. u \ f x)" by (metis cINF_greatest cINF_lower order_trans) lemma cSUP_le_iff: "A \ {} \ bdd_above (f ` A) \ \(f ` A) \ u \ (\x\A. f x \ u)" by (metis cSUP_least cSUP_upper order_trans) lemma less_cINF_D: "bdd_below (f`A) \ y < (\i\A. f i) \ i \ A \ y < f i" by (metis cINF_lower less_le_trans) lemma cSUP_lessD: "bdd_above (f`A) \ (\i\A. f i) < y \ i \ A \ f i < y" by (metis cSUP_upper le_less_trans) lemma cINF_insert: "A \ {} \ bdd_below (f ` A) \ \(f ` insert a A) = inf (f a) (\(f ` A))" by (simp add: cInf_insert) lemma cSUP_insert: "A \ {} \ bdd_above (f ` A) \ \(f ` insert a A) = sup (f a) (\(f ` A))" by (simp add: cSup_insert) lemma cINF_mono: "B \ {} \ bdd_below (f ` A) \ (\m. m \ B \ \n\A. f n \ g m) \ \(f ` A) \ \(g ` B)" using cInf_mono [of "g ` B" "f ` A"] by auto lemma cSUP_mono: "A \ {} \ bdd_above (g ` B) \ (\n. n \ A \ \m\B. f n \ g m) \ \(f ` A) \ \(g ` B)" using cSup_mono [of "f ` A" "g ` B"] by auto lemma cINF_superset_mono: "A \ {} \ bdd_below (g ` B) \ A \ B \ (\x. x \ B \ g x \ f x) \ \(g ` B) \ \(f ` A)" by (rule cINF_mono) auto lemma cSUP_subset_mono: "\A \ {}; bdd_above (g ` B); A \ B; \x. x \ A \ f x \ g x\ \ \ (f ` A) \ \ (g ` B)" by (rule cSUP_mono) auto lemma less_eq_cInf_inter: "bdd_below A \ bdd_below B \ A \ B \ {} \ inf (Inf A) (Inf B) \ Inf (A \ B)" by (metis cInf_superset_mono lattice_class.inf_sup_ord(1) le_infI1) lemma cSup_inter_less_eq: "bdd_above A \ bdd_above B \ A \ B \ {} \ Sup (A \ B) \ sup (Sup A) (Sup B) " by (metis cSup_subset_mono lattice_class.inf_sup_ord(1) le_supI1) lemma cInf_union_distrib: "A \ {} \ bdd_below A \ B \ {} \ bdd_below B \ Inf (A \ B) = inf (Inf A) (Inf B)" by (intro order.antisym le_infI cInf_greatest cInf_lower) (auto intro: le_infI1 le_infI2 cInf_lower) lemma cINF_union: "A \ {} \ bdd_below (f ` A) \ B \ {} \ bdd_below (f ` B) \ \ (f ` (A \ B)) = \ (f ` A) \ \ (f ` B)" using cInf_union_distrib [of "f ` A" "f ` B"] by (simp add: image_Un) lemma cSup_union_distrib: "A \ {} \ bdd_above A \ B \ {} \ bdd_above B \ Sup (A \ B) = sup (Sup A) (Sup B)" by (intro order.antisym le_supI cSup_least cSup_upper) (auto intro: le_supI1 le_supI2 cSup_upper) lemma cSUP_union: "A \ {} \ bdd_above (f ` A) \ B \ {} \ bdd_above (f ` B) \ \ (f ` (A \ B)) = \ (f ` A) \ \ (f ` B)" using cSup_union_distrib [of "f ` A" "f ` B"] by (simp add: image_Un) lemma cINF_inf_distrib: "A \ {} \ bdd_below (f`A) \ bdd_below (g`A) \ \ (f ` A) \ \ (g ` A) = (\a\A. inf (f a) (g a))" by (intro order.antisym le_infI cINF_greatest cINF_lower2) (auto intro: le_infI1 le_infI2 cINF_greatest cINF_lower le_infI) lemma SUP_sup_distrib: "A \ {} \ bdd_above (f`A) \ bdd_above (g`A) \ \ (f ` A) \ \ (g ` A) = (\a\A. sup (f a) (g a))" by (intro order.antisym le_supI cSUP_least cSUP_upper2) (auto intro: le_supI1 le_supI2 cSUP_least cSUP_upper le_supI) lemma cInf_le_cSup: "A \ {} \ bdd_above A \ bdd_below A \ Inf A \ Sup A" by (auto intro!: cSup_upper2[of "SOME a. a \ A"] intro: someI cInf_lower) + context + fixes f :: "'a \ 'b::conditionally_complete_lattice" + assumes "mono f" + begin + + lemma mono_cInf: "\bdd_below A; A\{}\ \ f (Inf A) \ (INF x\A. f x)" + by (simp add: \mono f\ conditionally_complete_lattice_class.cINF_greatest cInf_lower monoD) + + lemma mono_cSup: "\bdd_above A; A\{}\ \ (SUP x\A. f x) \ f (Sup A)" + by (simp add: \mono f\ conditionally_complete_lattice_class.cSUP_least cSup_upper monoD) + + lemma mono_cINF: "\bdd_below (A`I); I\{}\ \ f (INF i\I. A i) \ (INF x\I. f (A x))" + by (simp add: \mono f\ conditionally_complete_lattice_class.cINF_greatest cINF_lower monoD) + + lemma mono_cSUP: "\bdd_above (A`I); I\{}\ \ (SUP x\I. f (A x)) \ f (SUP i\I. A i)" + by (simp add: \mono f\ conditionally_complete_lattice_class.cSUP_least cSUP_upper monoD) + + end + end instance complete_lattice \ conditionally_complete_lattice by standard (auto intro: Sup_upper Sup_least Inf_lower Inf_greatest) lemma cSup_eq: fixes a :: "'a :: {conditionally_complete_lattice, no_bot}" assumes upper: "\x. x \ X \ x \ a" assumes least: "\y. (\x. x \ X \ x \ y) \ a \ y" shows "Sup X = a" proof cases assume "X = {}" with lt_ex[of a] least show ?thesis by (auto simp: less_le_not_le) qed (intro cSup_eq_non_empty assms) lemma cInf_eq: fixes a :: "'a :: {conditionally_complete_lattice, no_top}" assumes upper: "\x. x \ X \ a \ x" assumes least: "\y. (\x. x \ X \ y \ x) \ y \ a" shows "Inf X = a" proof cases assume "X = {}" with gt_ex[of a] least show ?thesis by (auto simp: less_le_not_le) qed (intro cInf_eq_non_empty assms) class conditionally_complete_linorder = conditionally_complete_lattice + linorder begin lemma less_cSup_iff: "X \ {} \ bdd_above X \ y < Sup X \ (\x\X. y < x)" by (rule iffI) (metis cSup_least not_less, metis cSup_upper less_le_trans) lemma cInf_less_iff: "X \ {} \ bdd_below X \ Inf X < y \ (\x\X. x < y)" by (rule iffI) (metis cInf_greatest not_less, metis cInf_lower le_less_trans) lemma cINF_less_iff: "A \ {} \ bdd_below (f`A) \ (\i\A. f i) < a \ (\x\A. f x < a)" using cInf_less_iff[of "f`A"] by auto lemma less_cSUP_iff: "A \ {} \ bdd_above (f`A) \ a < (\i\A. f i) \ (\x\A. a < f x)" using less_cSup_iff[of "f`A"] by auto lemma less_cSupE: assumes "y < Sup X" "X \ {}" obtains x where "x \ X" "y < x" by (metis cSup_least assms not_le that) lemma less_cSupD: "X \ {} \ z < Sup X \ \x\X. z < x" by (metis less_cSup_iff not_le_imp_less bdd_above_def) lemma cInf_lessD: "X \ {} \ Inf X < z \ \x\X. x < z" by (metis cInf_less_iff not_le_imp_less bdd_below_def) lemma complete_interval: assumes "a < b" and "P a" and "\ P b" shows "\c. a \ c \ c \ b \ (\x. a \ x \ x < c \ P x) \ (\d. (\x. a \ x \ x < d \ P x) \ d \ c)" proof (rule exI [where x = "Sup {d. \x. a \ x \ x < d \ P x}"], auto) show "a \ Sup {d. \c. a \ c \ c < d \ P c}" by (rule cSup_upper, auto simp: bdd_above_def) (metis \a < b\ \\ P b\ linear less_le) next show "Sup {d. \c. a \ c \ c < d \ P c} \ b" apply (rule cSup_least) apply auto apply (metis less_le_not_le) apply (metis \a \\ P b\ linear less_le) done next fix x assume x: "a \ x" and lt: "x < Sup {d. \c. a \ c \ c < d \ P c}" show "P x" apply (rule less_cSupE [OF lt], auto) apply (metis less_le_not_le) apply (metis x) done next fix d assume 0: "\x. a \ x \ x < d \ P x" thus "d \ Sup {d. \c. a \ c \ c < d \ P c}" by (rule_tac cSup_upper, auto simp: bdd_above_def) (metis \a \\ P b\ linear less_le) qed end instance complete_linorder < conditionally_complete_linorder .. lemma cSup_eq_Max: "finite (X::'a::conditionally_complete_linorder set) \ X \ {} \ Sup X = Max X" using cSup_eq_Sup_fin[of X] by (simp add: Sup_fin_Max) lemma cInf_eq_Min: "finite (X::'a::conditionally_complete_linorder set) \ X \ {} \ Inf X = Min X" using cInf_eq_Inf_fin[of X] by (simp add: Inf_fin_Min) lemma cSup_lessThan[simp]: "Sup {.. Sup {y<.. Sup {y.. Inf {y<..x::'a::{conditionally_complete_linorder, dense_linorder}} = y" by (auto intro!: cInf_eq_non_empty intro: dense_ge_bounded) lemma cInf_greaterThanLessThan[simp]: "y < x \ Inf {y<.. Inf (insert x S) = (if S = {} then x else min x (Inf S))" by (simp add: cInf_eq_Min) lemma Sup_insert_finite: fixes S :: "'a::conditionally_complete_linorder set" shows "finite S \ Sup (insert x S) = (if S = {} then x else max x (Sup S))" by (simp add: cSup_insert sup_max) lemma finite_imp_less_Inf: fixes a :: "'a::conditionally_complete_linorder" shows "\finite X; x \ X; \x. x\X \ a < x\ \ a < Inf X" by (induction X rule: finite_induct) (simp_all add: cInf_eq_Min Inf_insert_finite) lemma finite_less_Inf_iff: fixes a :: "'a :: conditionally_complete_linorder" shows "\finite X; X \ {}\ \ a < Inf X \ (\x \ X. a < x)" by (auto simp: cInf_eq_Min) lemma finite_imp_Sup_less: fixes a :: "'a::conditionally_complete_linorder" shows "\finite X; x \ X; \x. x\X \ a > x\ \ a > Sup X" by (induction X rule: finite_induct) (simp_all add: cSup_eq_Max Sup_insert_finite) lemma finite_Sup_less_iff: fixes a :: "'a :: conditionally_complete_linorder" shows "\finite X; X \ {}\ \ a > Sup X \ (\x \ X. a > x)" by (auto simp: cSup_eq_Max) class linear_continuum = conditionally_complete_linorder + dense_linorder + assumes UNIV_not_singleton: "\a b::'a. a \ b" begin lemma ex_gt_or_lt: "\b. a < b \ b < a" by (metis UNIV_not_singleton neq_iff) end instantiation nat :: conditionally_complete_linorder begin definition "Sup (X::nat set) = (if X={} then 0 else Max X)" definition "Inf (X::nat set) = (LEAST n. n \ X)" lemma bdd_above_nat: "bdd_above X \ finite (X::nat set)" proof assume "bdd_above X" then obtain z where "X \ {.. z}" by (auto simp: bdd_above_def) then show "finite X" by (rule finite_subset) simp qed simp instance proof fix x :: nat fix X :: "nat set" show "Inf X \ x" if "x \ X" "bdd_below X" using that by (simp add: Inf_nat_def Least_le) show "x \ Inf X" if "X \ {}" "\y. y \ X \ x \ y" using that unfolding Inf_nat_def ex_in_conv[symmetric] by (rule LeastI2_ex) show "x \ Sup X" if "x \ X" "bdd_above X" using that by (auto simp add: Sup_nat_def bdd_above_nat) show "Sup X \ x" if "X \ {}" "\y. y \ X \ y \ x" proof - from that have "bdd_above X" by (auto simp: bdd_above_def) with that show ?thesis by (simp add: Sup_nat_def bdd_above_nat) qed qed end lemma Inf_nat_def1: fixes K::"nat set" assumes "K \ {}" shows "Inf K \ K" by (auto simp add: Min_def Inf_nat_def) (meson LeastI assms bot.extremum_unique subsetI) lemma Sup_nat_empty [simp]: "Sup {} = (0::nat)" by (auto simp add: Sup_nat_def) instantiation int :: conditionally_complete_linorder begin definition "Sup (X::int set) = (THE x. x \ X \ (\y\X. y \ x))" definition "Inf (X::int set) = - (Sup (uminus ` X))" instance proof { fix x :: int and X :: "int set" assume "X \ {}" "bdd_above X" then obtain x y where "X \ {..y}" "x \ X" by (auto simp: bdd_above_def) then have *: "finite (X \ {x..y})" "X \ {x..y} \ {}" and "x \ y" by (auto simp: subset_eq) have "\!x\X. (\y\X. y \ x)" proof { fix z assume "z \ X" have "z \ Max (X \ {x..y})" proof cases assume "x \ z" with \z \ X\ \X \ {..y}\ *(1) show ?thesis by (auto intro!: Max_ge) next assume "\ x \ z" then have "z < x" by simp also have "x \ Max (X \ {x..y})" using \x \ X\ *(1) \x \ y\ by (intro Max_ge) auto finally show ?thesis by simp qed } note le = this with Max_in[OF *] show ex: "Max (X \ {x..y}) \ X \ (\z\X. z \ Max (X \ {x..y}))" by auto fix z assume *: "z \ X \ (\y\X. y \ z)" with le have "z \ Max (X \ {x..y})" by auto moreover have "Max (X \ {x..y}) \ z" using * ex by auto ultimately show "z = Max (X \ {x..y})" by auto qed then have "Sup X \ X \ (\y\X. y \ Sup X)" unfolding Sup_int_def by (rule theI') } note Sup_int = this { fix x :: int and X :: "int set" assume "x \ X" "bdd_above X" then show "x \ Sup X" using Sup_int[of X] by auto } note le_Sup = this { fix x :: int and X :: "int set" assume "X \ {}" "\y. y \ X \ y \ x" then show "Sup X \ x" using Sup_int[of X] by (auto simp: bdd_above_def) } note Sup_le = this { fix x :: int and X :: "int set" assume "x \ X" "bdd_below X" then show "Inf X \ x" using le_Sup[of "-x" "uminus ` X"] by (auto simp: Inf_int_def) } { fix x :: int and X :: "int set" assume "X \ {}" "\y. y \ X \ x \ y" then show "x \ Inf X" using Sup_le[of "uminus ` X" "-x"] by (force simp: Inf_int_def) } qed end lemma interval_cases: fixes S :: "'a :: conditionally_complete_linorder set" assumes ivl: "\a b x. a \ S \ b \ S \ a \ x \ x \ b \ x \ S" shows "\a b. S = {} \ S = UNIV \ S = {.. S = {..b} \ S = {a<..} \ S = {a..} \ S = {a<.. S = {a<..b} \ S = {a.. S = {a..b}" proof - define lower upper where "lower = {x. \s\S. s \ x}" and "upper = {x. \s\S. x \ s}" with ivl have "S = lower \ upper" by auto moreover have "\a. upper = UNIV \ upper = {} \ upper = {.. a} \ upper = {..< a}" proof cases assume *: "bdd_above S \ S \ {}" from * have "upper \ {.. Sup S}" by (auto simp: upper_def intro: cSup_upper2) moreover from * have "{..< Sup S} \ upper" by (force simp add: less_cSup_iff upper_def subset_eq Ball_def) ultimately have "upper = {.. Sup S} \ upper = {..< Sup S}" unfolding ivl_disj_un(2)[symmetric] by auto then show ?thesis by auto next assume "\ (bdd_above S \ S \ {})" then have "upper = UNIV \ upper = {}" by (auto simp: upper_def bdd_above_def not_le dest: less_imp_le) then show ?thesis by auto qed moreover have "\b. lower = UNIV \ lower = {} \ lower = {b ..} \ lower = {b <..}" proof cases assume *: "bdd_below S \ S \ {}" from * have "lower \ {Inf S ..}" by (auto simp: lower_def intro: cInf_lower2) moreover from * have "{Inf S <..} \ lower" by (force simp add: cInf_less_iff lower_def subset_eq Ball_def) ultimately have "lower = {Inf S ..} \ lower = {Inf S <..}" unfolding ivl_disj_un(1)[symmetric] by auto then show ?thesis by auto next assume "\ (bdd_below S \ S \ {})" then have "lower = UNIV \ lower = {}" by (auto simp: lower_def bdd_below_def not_le dest: less_imp_le) then show ?thesis by auto qed ultimately show ?thesis unfolding greaterThanAtMost_def greaterThanLessThan_def atLeastAtMost_def atLeastLessThan_def by (metis inf_bot_left inf_bot_right inf_top.left_neutral inf_top.right_neutral) qed lemma cSUP_eq_cINF_D: fixes f :: "_ \ 'b::conditionally_complete_lattice" assumes eq: "(\x\A. f x) = (\x\A. f x)" and bdd: "bdd_above (f ` A)" "bdd_below (f ` A)" and a: "a \ A" shows "f a = (\x\A. f x)" apply (rule antisym) using a bdd apply (auto simp: cINF_lower) apply (metis eq cSUP_upper) done lemma cSUP_UNION: fixes f :: "_ \ 'b::conditionally_complete_lattice" assumes ne: "A \ {}" "\x. x \ A \ B(x) \ {}" and bdd_UN: "bdd_above (\x\A. f ` B x)" shows "(\z \ \x\A. B x. f z) = (\x\A. \z\B x. f z)" proof - have bdd: "\x. x \ A \ bdd_above (f ` B x)" using bdd_UN by (meson UN_upper bdd_above_mono) obtain M where "\x y. x \ A \ y \ B(x) \ f y \ M" using bdd_UN by (auto simp: bdd_above_def) then have bdd2: "bdd_above ((\x. \z\B x. f z) ` A)" unfolding bdd_above_def by (force simp: bdd cSUP_le_iff ne(2)) have "(\z \ \x\A. B x. f z) \ (\x\A. \z\B x. f z)" using assms by (fastforce simp add: intro!: cSUP_least intro: cSUP_upper2 simp: bdd2 bdd) moreover have "(\x\A. \z\B x. f z) \ (\ z \ \x\A. B x. f z)" using assms by (fastforce simp add: intro!: cSUP_least intro: cSUP_upper simp: image_UN bdd_UN) ultimately show ?thesis by (rule order_antisym) qed lemma cINF_UNION: fixes f :: "_ \ 'b::conditionally_complete_lattice" assumes ne: "A \ {}" "\x. x \ A \ B(x) \ {}" and bdd_UN: "bdd_below (\x\A. f ` B x)" shows "(\z \ \x\A. B x. f z) = (\x\A. \z\B x. f z)" proof - have bdd: "\x. x \ A \ bdd_below (f ` B x)" using bdd_UN by (meson UN_upper bdd_below_mono) obtain M where "\x y. x \ A \ y \ B(x) \ f y \ M" using bdd_UN by (auto simp: bdd_below_def) then have bdd2: "bdd_below ((\x. \z\B x. f z) ` A)" unfolding bdd_below_def by (force simp: bdd le_cINF_iff ne(2)) have "(\z \ \x\A. B x. f z) \ (\x\A. \z\B x. f z)" using assms by (fastforce simp add: intro!: cINF_greatest intro: cINF_lower simp: bdd2 bdd) moreover have "(\x\A. \z\B x. f z) \ (\z \ \x\A. B x. f z)" using assms by (fastforce simp add: intro!: cINF_greatest intro: cINF_lower2 simp: bdd bdd_UN bdd2) ultimately show ?thesis by (rule order_antisym) qed lemma cSup_abs_le: fixes S :: "('a::{linordered_idom,conditionally_complete_linorder}) set" shows "S \ {} \ (\x. x\S \ \x\ \ a) \ \Sup S\ \ a" apply (auto simp add: abs_le_iff intro: cSup_least) by (metis bdd_aboveI cSup_upper neg_le_iff_le order_trans) end diff --git a/src/HOL/Deriv.thy b/src/HOL/Deriv.thy --- a/src/HOL/Deriv.thy +++ b/src/HOL/Deriv.thy @@ -1,2443 +1,2449 @@ (* Title: HOL/Deriv.thy Author: Jacques D. Fleuriot, University of Cambridge, 1998 Author: Brian Huffman Author: Lawrence C Paulson, 2004 Author: Benjamin Porter, 2005 *) section \Differentiation\ theory Deriv imports Limits begin subsection \Frechet derivative\ definition has_derivative :: "('a::real_normed_vector \ 'b::real_normed_vector) \ ('a \ 'b) \ 'a filter \ bool" (infix "(has'_derivative)" 50) where "(f has_derivative f') F \ bounded_linear f' \ ((\y. ((f y - f (Lim F (\x. x))) - f' (y - Lim F (\x. x))) /\<^sub>R norm (y - Lim F (\x. x))) \ 0) F" text \ Usually the filter \<^term>\F\ is \<^term>\at x within s\. \<^term>\(f has_derivative D) (at x within s)\ means: \<^term>\D\ is the derivative of function \<^term>\f\ at point \<^term>\x\ within the set \<^term>\s\. Where \<^term>\s\ is used to express left or right sided derivatives. In most cases \<^term>\s\ is either a variable or \<^term>\UNIV\. \ text \These are the only cases we'll care about, probably.\ lemma has_derivative_within: "(f has_derivative f') (at x within s) \ bounded_linear f' \ ((\y. (1 / norm(y - x)) *\<^sub>R (f y - (f x + f' (y - x)))) \ 0) (at x within s)" unfolding has_derivative_def tendsto_iff by (subst eventually_Lim_ident_at) (auto simp add: field_simps) lemma has_derivative_eq_rhs: "(f has_derivative f') F \ f' = g' \ (f has_derivative g') F" by simp definition has_field_derivative :: "('a::real_normed_field \ 'a) \ 'a \ 'a filter \ bool" (infix "(has'_field'_derivative)" 50) where "(f has_field_derivative D) F \ (f has_derivative (*) D) F" lemma DERIV_cong: "(f has_field_derivative X) F \ X = Y \ (f has_field_derivative Y) F" by simp definition has_vector_derivative :: "(real \ 'b::real_normed_vector) \ 'b \ real filter \ bool" (infix "has'_vector'_derivative" 50) where "(f has_vector_derivative f') net \ (f has_derivative (\x. x *\<^sub>R f')) net" lemma has_vector_derivative_eq_rhs: "(f has_vector_derivative X) F \ X = Y \ (f has_vector_derivative Y) F" by simp named_theorems derivative_intros "structural introduction rules for derivatives" setup \ let val eq_thms = @{thms has_derivative_eq_rhs DERIV_cong has_vector_derivative_eq_rhs} fun eq_rule thm = get_first (try (fn eq_thm => eq_thm OF [thm])) eq_thms in Global_Theory.add_thms_dynamic (\<^binding>\derivative_eq_intros\, fn context => Named_Theorems.get (Context.proof_of context) \<^named_theorems>\derivative_intros\ |> map_filter eq_rule) end \ text \ The following syntax is only used as a legacy syntax. \ abbreviation (input) FDERIV :: "('a::real_normed_vector \ 'b::real_normed_vector) \ 'a \ ('a \ 'b) \ bool" ("(FDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where "FDERIV f x :> f' \ (f has_derivative f') (at x)" lemma has_derivative_bounded_linear: "(f has_derivative f') F \ bounded_linear f'" by (simp add: has_derivative_def) lemma has_derivative_linear: "(f has_derivative f') F \ linear f'" using bounded_linear.linear[OF has_derivative_bounded_linear] . lemma has_derivative_ident[derivative_intros, simp]: "((\x. x) has_derivative (\x. x)) F" by (simp add: has_derivative_def) lemma has_derivative_id [derivative_intros, simp]: "(id has_derivative id) (at a)" by (metis eq_id_iff has_derivative_ident) lemma has_derivative_const[derivative_intros, simp]: "((\x. c) has_derivative (\x. 0)) F" by (simp add: has_derivative_def) lemma (in bounded_linear) bounded_linear: "bounded_linear f" .. lemma (in bounded_linear) has_derivative: "(g has_derivative g') F \ ((\x. f (g x)) has_derivative (\x. f (g' x))) F" unfolding has_derivative_def by (auto simp add: bounded_linear_compose [OF bounded_linear] scaleR diff dest: tendsto) lemmas has_derivative_scaleR_right [derivative_intros] = bounded_linear.has_derivative [OF bounded_linear_scaleR_right] lemmas has_derivative_scaleR_left [derivative_intros] = bounded_linear.has_derivative [OF bounded_linear_scaleR_left] lemmas has_derivative_mult_right [derivative_intros] = bounded_linear.has_derivative [OF bounded_linear_mult_right] lemmas has_derivative_mult_left [derivative_intros] = bounded_linear.has_derivative [OF bounded_linear_mult_left] lemmas has_derivative_of_real[derivative_intros, simp] = bounded_linear.has_derivative[OF bounded_linear_of_real] lemma has_derivative_add[simp, derivative_intros]: assumes f: "(f has_derivative f') F" and g: "(g has_derivative g') F" shows "((\x. f x + g x) has_derivative (\x. f' x + g' x)) F" unfolding has_derivative_def proof safe let ?x = "Lim F (\x. x)" let ?D = "\f f' y. ((f y - f ?x) - f' (y - ?x)) /\<^sub>R norm (y - ?x)" have "((\x. ?D f f' x + ?D g g' x) \ (0 + 0)) F" using f g by (intro tendsto_add) (auto simp: has_derivative_def) then show "(?D (\x. f x + g x) (\x. f' x + g' x) \ 0) F" by (simp add: field_simps scaleR_add_right scaleR_diff_right) qed (blast intro: bounded_linear_add f g has_derivative_bounded_linear) lemma has_derivative_sum[simp, derivative_intros]: "(\i. i \ I \ (f i has_derivative f' i) F) \ ((\x. \i\I. f i x) has_derivative (\x. \i\I. f' i x)) F" by (induct I rule: infinite_finite_induct) simp_all lemma has_derivative_minus[simp, derivative_intros]: "(f has_derivative f') F \ ((\x. - f x) has_derivative (\x. - f' x)) F" using has_derivative_scaleR_right[of f f' F "-1"] by simp lemma has_derivative_diff[simp, derivative_intros]: "(f has_derivative f') F \ (g has_derivative g') F \ ((\x. f x - g x) has_derivative (\x. f' x - g' x)) F" by (simp only: diff_conv_add_uminus has_derivative_add has_derivative_minus) lemma has_derivative_at_within: "(f has_derivative f') (at x within s) \ (bounded_linear f' \ ((\y. ((f y - f x) - f' (y - x)) /\<^sub>R norm (y - x)) \ 0) (at x within s))" proof (cases "at x within s = bot") case True then show ?thesis by (metis (no_types, lifting) has_derivative_within tendsto_bot) next case False then show ?thesis by (simp add: Lim_ident_at has_derivative_def) qed lemma has_derivative_iff_norm: "(f has_derivative f') (at x within s) \ bounded_linear f' \ ((\y. norm ((f y - f x) - f' (y - x)) / norm (y - x)) \ 0) (at x within s)" using tendsto_norm_zero_iff[of _ "at x within s", where 'b="'b", symmetric] by (simp add: has_derivative_at_within divide_inverse ac_simps) lemma has_derivative_at: "(f has_derivative D) (at x) \ (bounded_linear D \ (\h. norm (f (x + h) - f x - D h) / norm h) \0\ 0)" by (simp add: has_derivative_iff_norm LIM_offset_zero_iff) lemma field_has_derivative_at: fixes x :: "'a::real_normed_field" shows "(f has_derivative (*) D) (at x) \ (\h. (f (x + h) - f x) / h) \0\ D" (is "?lhs = ?rhs") proof - have "?lhs = (\h. norm (f (x + h) - f x - D * h) / norm h) \0 \ 0" by (simp add: bounded_linear_mult_right has_derivative_at) also have "... = (\y. norm ((f (x + y) - f x - D * y) / y)) \0\ 0" by (simp cong: LIM_cong flip: nonzero_norm_divide) also have "... = (\y. norm ((f (x + y) - f x) / y - D / y * y)) \0\ 0" by (simp only: diff_divide_distrib times_divide_eq_left [symmetric]) also have "... = ?rhs" by (simp add: tendsto_norm_zero_iff LIM_zero_iff cong: LIM_cong) finally show ?thesis . qed lemma has_derivative_iff_Ex: "(f has_derivative f') (at x) \ bounded_linear f' \ (\e. (\h. f (x+h) = f x + f' h + e h) \ ((\h. norm (e h) / norm h) \ 0) (at 0))" unfolding has_derivative_at by force lemma has_derivative_at_within_iff_Ex: assumes "x \ S" "open S" shows "(f has_derivative f') (at x within S) \ bounded_linear f' \ (\e. (\h. x+h \ S \ f (x+h) = f x + f' h + e h) \ ((\h. norm (e h) / norm h) \ 0) (at 0))" (is "?lhs = ?rhs") proof safe show "bounded_linear f'" if "(f has_derivative f') (at x within S)" using has_derivative_bounded_linear that by blast show "\e. (\h. x + h \ S \ f (x + h) = f x + f' h + e h) \ (\h. norm (e h) / norm h) \0\ 0" if "(f has_derivative f') (at x within S)" by (metis (full_types) assms that has_derivative_iff_Ex at_within_open) show "(f has_derivative f') (at x within S)" if "bounded_linear f'" and eq [rule_format]: "\h. x + h \ S \ f (x + h) = f x + f' h + e h" and 0: "(\h. norm (e (h::'a)::'b) / norm h) \0\ 0" for e proof - have 1: "f y - f x = f' (y-x) + e (y-x)" if "y \ S" for y using eq [of "y-x"] that by simp have 2: "((\y. norm (e (y-x)) / norm (y - x)) \ 0) (at x within S)" by (simp add: "0" assms tendsto_offset_zero_iff) have "((\y. norm (f y - f x - f' (y - x)) / norm (y - x)) \ 0) (at x within S)" by (simp add: Lim_cong_within 1 2) then show ?thesis by (simp add: has_derivative_iff_norm \bounded_linear f'\) qed qed lemma has_derivativeI: "bounded_linear f' \ ((\y. ((f y - f x) - f' (y - x)) /\<^sub>R norm (y - x)) \ 0) (at x within s) \ (f has_derivative f') (at x within s)" by (simp add: has_derivative_at_within) lemma has_derivativeI_sandwich: assumes e: "0 < e" and bounded: "bounded_linear f'" and sandwich: "(\y. y \ s \ y \ x \ dist y x < e \ norm ((f y - f x) - f' (y - x)) / norm (y - x) \ H y)" and "(H \ 0) (at x within s)" shows "(f has_derivative f') (at x within s)" unfolding has_derivative_iff_norm proof safe show "((\y. norm (f y - f x - f' (y - x)) / norm (y - x)) \ 0) (at x within s)" proof (rule tendsto_sandwich[where f="\x. 0"]) show "(H \ 0) (at x within s)" by fact show "eventually (\n. norm (f n - f x - f' (n - x)) / norm (n - x) \ H n) (at x within s)" unfolding eventually_at using e sandwich by auto qed (auto simp: le_divide_eq) qed fact lemma has_derivative_subset: "(f has_derivative f') (at x within s) \ t \ s \ (f has_derivative f') (at x within t)" by (auto simp add: has_derivative_iff_norm intro: tendsto_within_subset) lemma has_derivative_within_singleton_iff: "(f has_derivative g) (at x within {x}) \ bounded_linear g" by (auto intro!: has_derivativeI_sandwich[where e=1] has_derivative_bounded_linear) subsubsection \Limit transformation for derivatives\ lemma has_derivative_transform_within: assumes "(f has_derivative f') (at x within s)" and "0 < d" and "x \ s" and "\x'. \x' \ s; dist x' x < d\ \ f x' = g x'" shows "(g has_derivative f') (at x within s)" using assms unfolding has_derivative_within by (force simp add: intro: Lim_transform_within) lemma has_derivative_transform_within_open: assumes "(f has_derivative f') (at x within t)" and "open s" and "x \ s" and "\x. x\s \ f x = g x" shows "(g has_derivative f') (at x within t)" using assms unfolding has_derivative_within by (force simp add: intro: Lim_transform_within_open) lemma has_derivative_transform: assumes "x \ s" "\x. x \ s \ g x = f x" assumes "(f has_derivative f') (at x within s)" shows "(g has_derivative f') (at x within s)" using assms by (intro has_derivative_transform_within[OF _ zero_less_one, where g=g]) auto lemma has_derivative_transform_eventually: assumes "(f has_derivative f') (at x within s)" "(\\<^sub>F x' in at x within s. f x' = g x')" assumes "f x = g x" "x \ s" shows "(g has_derivative f') (at x within s)" using assms proof - from assms(2,3) obtain d where "d > 0" "\x'. x' \ s \ dist x' x < d \ f x' = g x'" by (force simp: eventually_at) from has_derivative_transform_within[OF assms(1) this(1) assms(4) this(2)] show ?thesis . qed lemma has_field_derivative_transform_within: assumes "(f has_field_derivative f') (at a within S)" and "0 < d" and "a \ S" and "\x. \x \ S; dist x a < d\ \ f x = g x" shows "(g has_field_derivative f') (at a within S)" using assms unfolding has_field_derivative_def by (metis has_derivative_transform_within) lemma has_field_derivative_transform_within_open: assumes "(f has_field_derivative f') (at a)" and "open S" "a \ S" and "\x. x \ S \ f x = g x" shows "(g has_field_derivative f') (at a)" using assms unfolding has_field_derivative_def by (metis has_derivative_transform_within_open) subsection \Continuity\ lemma has_derivative_continuous: assumes f: "(f has_derivative f') (at x within s)" shows "continuous (at x within s) f" proof - from f interpret F: bounded_linear f' by (rule has_derivative_bounded_linear) note F.tendsto[tendsto_intros] let ?L = "\f. (f \ 0) (at x within s)" have "?L (\y. norm ((f y - f x) - f' (y - x)) / norm (y - x))" using f unfolding has_derivative_iff_norm by blast then have "?L (\y. norm ((f y - f x) - f' (y - x)) / norm (y - x) * norm (y - x))" (is ?m) by (rule tendsto_mult_zero) (auto intro!: tendsto_eq_intros) also have "?m \ ?L (\y. norm ((f y - f x) - f' (y - x)))" by (intro filterlim_cong) (simp_all add: eventually_at_filter) finally have "?L (\y. (f y - f x) - f' (y - x))" by (rule tendsto_norm_zero_cancel) then have "?L (\y. ((f y - f x) - f' (y - x)) + f' (y - x))" by (rule tendsto_eq_intros) (auto intro!: tendsto_eq_intros simp: F.zero) then have "?L (\y. f y - f x)" by simp from tendsto_add[OF this tendsto_const, of "f x"] show ?thesis by (simp add: continuous_within) qed subsection \Composition\ lemma tendsto_at_iff_tendsto_nhds_within: "f x = y \ (f \ y) (at x within s) \ (f \ y) (inf (nhds x) (principal s))" unfolding tendsto_def eventually_inf_principal eventually_at_filter by (intro ext all_cong imp_cong) (auto elim!: eventually_mono) lemma has_derivative_in_compose: assumes f: "(f has_derivative f') (at x within s)" and g: "(g has_derivative g') (at (f x) within (f`s))" shows "((\x. g (f x)) has_derivative (\x. g' (f' x))) (at x within s)" proof - from f interpret F: bounded_linear f' by (rule has_derivative_bounded_linear) from g interpret G: bounded_linear g' by (rule has_derivative_bounded_linear) from F.bounded obtain kF where kF: "\x. norm (f' x) \ norm x * kF" by fast from G.bounded obtain kG where kG: "\x. norm (g' x) \ norm x * kG" by fast note G.tendsto[tendsto_intros] let ?L = "\f. (f \ 0) (at x within s)" let ?D = "\f f' x y. (f y - f x) - f' (y - x)" let ?N = "\f f' x y. norm (?D f f' x y) / norm (y - x)" let ?gf = "\x. g (f x)" and ?gf' = "\x. g' (f' x)" define Nf where "Nf = ?N f f' x" define Ng where [abs_def]: "Ng y = ?N g g' (f x) (f y)" for y show ?thesis proof (rule has_derivativeI_sandwich[of 1]) show "bounded_linear (\x. g' (f' x))" using f g by (blast intro: bounded_linear_compose has_derivative_bounded_linear) next fix y :: 'a assume neq: "y \ x" have "?N ?gf ?gf' x y = norm (g' (?D f f' x y) + ?D g g' (f x) (f y)) / norm (y - x)" by (simp add: G.diff G.add field_simps) also have "\ \ norm (g' (?D f f' x y)) / norm (y - x) + Ng y * (norm (f y - f x) / norm (y - x))" by (simp add: add_divide_distrib[symmetric] divide_right_mono norm_triangle_ineq G.zero Ng_def) also have "\ \ Nf y * kG + Ng y * (Nf y + kF)" proof (intro add_mono mult_left_mono) have "norm (f y - f x) = norm (?D f f' x y + f' (y - x))" by simp also have "\ \ norm (?D f f' x y) + norm (f' (y - x))" by (rule norm_triangle_ineq) also have "\ \ norm (?D f f' x y) + norm (y - x) * kF" using kF by (intro add_mono) simp finally show "norm (f y - f x) / norm (y - x) \ Nf y + kF" by (simp add: neq Nf_def field_simps) qed (use kG in \simp_all add: Ng_def Nf_def neq zero_le_divide_iff field_simps\) finally show "?N ?gf ?gf' x y \ Nf y * kG + Ng y * (Nf y + kF)" . next have [tendsto_intros]: "?L Nf" using f unfolding has_derivative_iff_norm Nf_def .. from f have "(f \ f x) (at x within s)" by (blast intro: has_derivative_continuous continuous_within[THEN iffD1]) then have f': "LIM x at x within s. f x :> inf (nhds (f x)) (principal (f`s))" unfolding filterlim_def by (simp add: eventually_filtermap eventually_at_filter le_principal) have "((?N g g' (f x)) \ 0) (at (f x) within f`s)" using g unfolding has_derivative_iff_norm .. then have g': "((?N g g' (f x)) \ 0) (inf (nhds (f x)) (principal (f`s)))" by (rule tendsto_at_iff_tendsto_nhds_within[THEN iffD1, rotated]) simp have [tendsto_intros]: "?L Ng" unfolding Ng_def by (rule filterlim_compose[OF g' f']) show "((\y. Nf y * kG + Ng y * (Nf y + kF)) \ 0) (at x within s)" by (intro tendsto_eq_intros) auto qed simp qed lemma has_derivative_compose: "(f has_derivative f') (at x within s) \ (g has_derivative g') (at (f x)) \ ((\x. g (f x)) has_derivative (\x. g' (f' x))) (at x within s)" by (blast intro: has_derivative_in_compose has_derivative_subset) lemma has_derivative_in_compose2: assumes "\x. x \ t \ (g has_derivative g' x) (at x within t)" assumes "f ` s \ t" "x \ s" assumes "(f has_derivative f') (at x within s)" shows "((\x. g (f x)) has_derivative (\y. g' (f x) (f' y))) (at x within s)" using assms by (auto intro: has_derivative_subset intro!: has_derivative_in_compose[of f f' x s g]) lemma (in bounded_bilinear) FDERIV: assumes f: "(f has_derivative f') (at x within s)" and g: "(g has_derivative g') (at x within s)" shows "((\x. f x ** g x) has_derivative (\h. f x ** g' h + f' h ** g x)) (at x within s)" proof - from bounded_linear.bounded [OF has_derivative_bounded_linear [OF f]] obtain KF where norm_F: "\x. norm (f' x) \ norm x * KF" by fast from pos_bounded obtain K where K: "0 < K" and norm_prod: "\a b. norm (a ** b) \ norm a * norm b * K" by fast let ?D = "\f f' y. f y - f x - f' (y - x)" let ?N = "\f f' y. norm (?D f f' y) / norm (y - x)" define Ng where "Ng = ?N g g'" define Nf where "Nf = ?N f f'" let ?fun1 = "\y. norm (f y ** g y - f x ** g x - (f x ** g' (y - x) + f' (y - x) ** g x)) / norm (y - x)" let ?fun2 = "\y. norm (f x) * Ng y * K + Nf y * norm (g y) * K + KF * norm (g y - g x) * K" let ?F = "at x within s" show ?thesis proof (rule has_derivativeI_sandwich[of 1]) show "bounded_linear (\h. f x ** g' h + f' h ** g x)" by (intro bounded_linear_add bounded_linear_compose [OF bounded_linear_right] bounded_linear_compose [OF bounded_linear_left] has_derivative_bounded_linear [OF g] has_derivative_bounded_linear [OF f]) next from g have "(g \ g x) ?F" by (intro continuous_within[THEN iffD1] has_derivative_continuous) moreover from f g have "(Nf \ 0) ?F" "(Ng \ 0) ?F" by (simp_all add: has_derivative_iff_norm Ng_def Nf_def) ultimately have "(?fun2 \ norm (f x) * 0 * K + 0 * norm (g x) * K + KF * norm (0::'b) * K) ?F" by (intro tendsto_intros) (simp_all add: LIM_zero_iff) then show "(?fun2 \ 0) ?F" by simp next fix y :: 'd assume "y \ x" have "?fun1 y = norm (f x ** ?D g g' y + ?D f f' y ** g y + f' (y - x) ** (g y - g x)) / norm (y - x)" by (simp add: diff_left diff_right add_left add_right field_simps) also have "\ \ (norm (f x) * norm (?D g g' y) * K + norm (?D f f' y) * norm (g y) * K + norm (y - x) * KF * norm (g y - g x) * K) / norm (y - x)" by (intro divide_right_mono mult_mono' order_trans [OF norm_triangle_ineq add_mono] order_trans [OF norm_prod mult_right_mono] mult_nonneg_nonneg order_refl norm_ge_zero norm_F K [THEN order_less_imp_le]) also have "\ = ?fun2 y" by (simp add: add_divide_distrib Ng_def Nf_def) finally show "?fun1 y \ ?fun2 y" . qed simp qed lemmas has_derivative_mult[simp, derivative_intros] = bounded_bilinear.FDERIV[OF bounded_bilinear_mult] lemmas has_derivative_scaleR[simp, derivative_intros] = bounded_bilinear.FDERIV[OF bounded_bilinear_scaleR] lemma has_derivative_prod[simp, derivative_intros]: fixes f :: "'i \ 'a::real_normed_vector \ 'b::real_normed_field" shows "(\i. i \ I \ (f i has_derivative f' i) (at x within S)) \ ((\x. \i\I. f i x) has_derivative (\y. \i\I. f' i y * (\j\I - {i}. f j x))) (at x within S)" proof (induct I rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case (insert i I) let ?P = "\y. f i x * (\i\I. f' i y * (\j\I - {i}. f j x)) + (f' i y) * (\i\I. f i x)" have "((\x. f i x * (\i\I. f i x)) has_derivative ?P) (at x within S)" using insert by (intro has_derivative_mult) auto also have "?P = (\y. \i'\insert i I. f' i' y * (\j\insert i I - {i'}. f j x))" using insert(1,2) by (auto simp add: sum_distrib_left insert_Diff_if intro!: ext sum.cong) finally show ?case using insert by simp qed lemma has_derivative_power[simp, derivative_intros]: fixes f :: "'a :: real_normed_vector \ 'b :: real_normed_field" assumes f: "(f has_derivative f') (at x within S)" shows "((\x. f x^n) has_derivative (\y. of_nat n * f' y * f x^(n - 1))) (at x within S)" using has_derivative_prod[OF f, of "{..< n}"] by (simp add: prod_constant ac_simps) lemma has_derivative_inverse': fixes x :: "'a::real_normed_div_algebra" assumes x: "x \ 0" shows "(inverse has_derivative (\h. - (inverse x * h * inverse x))) (at x within S)" (is "(_ has_derivative ?f) _") proof (rule has_derivativeI_sandwich) show "bounded_linear (\h. - (inverse x * h * inverse x))" by (simp add: bounded_linear_minus bounded_linear_mult_const bounded_linear_mult_right) show "0 < norm x" using x by simp have "(inverse \ inverse x) (at x within S)" using tendsto_inverse tendsto_ident_at x by auto then show "((\y. norm (inverse y - inverse x) * norm (inverse x)) \ 0) (at x within S)" by (simp add: LIM_zero_iff tendsto_mult_left_zero tendsto_norm_zero) next fix y :: 'a assume h: "y \ x" "dist y x < norm x" then have "y \ 0" by auto have "norm (inverse y - inverse x - ?f (y -x)) / norm (y - x) = norm (- (inverse y * (y - x) * inverse x - inverse x * (y - x) * inverse x)) / norm (y - x)" by (simp add: \y \ 0\ inverse_diff_inverse x) also have "... = norm ((inverse y - inverse x) * (y - x) * inverse x) / norm (y - x)" by (simp add: left_diff_distrib norm_minus_commute) also have "\ \ norm (inverse y - inverse x) * norm (y - x) * norm (inverse x) / norm (y - x)" by (simp add: norm_mult) also have "\ = norm (inverse y - inverse x) * norm (inverse x)" by simp finally show "norm (inverse y - inverse x - ?f (y -x)) / norm (y - x) \ norm (inverse y - inverse x) * norm (inverse x)" . qed lemma has_derivative_inverse[simp, derivative_intros]: fixes f :: "_ \ 'a::real_normed_div_algebra" assumes x: "f x \ 0" and f: "(f has_derivative f') (at x within S)" shows "((\x. inverse (f x)) has_derivative (\h. - (inverse (f x) * f' h * inverse (f x)))) (at x within S)" using has_derivative_compose[OF f has_derivative_inverse', OF x] . lemma has_derivative_divide[simp, derivative_intros]: fixes f :: "_ \ 'a::real_normed_div_algebra" assumes f: "(f has_derivative f') (at x within S)" and g: "(g has_derivative g') (at x within S)" assumes x: "g x \ 0" shows "((\x. f x / g x) has_derivative (\h. - f x * (inverse (g x) * g' h * inverse (g x)) + f' h / g x)) (at x within S)" using has_derivative_mult[OF f has_derivative_inverse[OF x g]] by (simp add: field_simps) lemma has_derivative_power_int': fixes x :: "'a::real_normed_field" assumes x: "x \ 0" shows "((\x. power_int x n) has_derivative (\y. y * (of_int n * power_int x (n - 1)))) (at x within S)" proof (cases n rule: int_cases4) case (nonneg n) thus ?thesis using x by (cases "n = 0") (auto intro!: derivative_eq_intros simp: field_simps power_int_diff fun_eq_iff simp flip: power_Suc) next case (neg n) thus ?thesis using x by (auto intro!: derivative_eq_intros simp: field_simps power_int_diff power_int_minus simp flip: power_Suc power_Suc2 power_add) qed lemma has_derivative_power_int[simp, derivative_intros]: fixes f :: "_ \ 'a::real_normed_field" assumes x: "f x \ 0" and f: "(f has_derivative f') (at x within S)" shows "((\x. power_int (f x) n) has_derivative (\h. f' h * (of_int n * power_int (f x) (n - 1)))) (at x within S)" using has_derivative_compose[OF f has_derivative_power_int', OF x] . text \Conventional form requires mult-AC laws. Types real and complex only.\ lemma has_derivative_divide'[derivative_intros]: fixes f :: "_ \ 'a::real_normed_field" assumes f: "(f has_derivative f') (at x within S)" and g: "(g has_derivative g') (at x within S)" and x: "g x \ 0" shows "((\x. f x / g x) has_derivative (\h. (f' h * g x - f x * g' h) / (g x * g x))) (at x within S)" proof - have "f' h / g x - f x * (inverse (g x) * g' h * inverse (g x)) = (f' h * g x - f x * g' h) / (g x * g x)" for h by (simp add: field_simps x) then show ?thesis using has_derivative_divide [OF f g] x by simp qed subsection \Uniqueness\ text \ This can not generally shown for \<^const>\has_derivative\, as we need to approach the point from all directions. There is a proof in \Analysis\ for \euclidean_space\. \ lemma has_derivative_at2: "(f has_derivative f') (at x) \ bounded_linear f' \ ((\y. (1 / (norm(y - x))) *\<^sub>R (f y - (f x + f' (y - x)))) \ 0) (at x)" using has_derivative_within [of f f' x UNIV] by simp lemma has_derivative_zero_unique: assumes "((\x. 0) has_derivative F) (at x)" shows "F = (\h. 0)" proof - interpret F: bounded_linear F using assms by (rule has_derivative_bounded_linear) let ?r = "\h. norm (F h) / norm h" have *: "?r \0\ 0" using assms unfolding has_derivative_at by simp show "F = (\h. 0)" proof show "F h = 0" for h proof (rule ccontr) assume **: "\ ?thesis" then have h: "h \ 0" by (auto simp add: F.zero) with ** have "0 < ?r h" by simp from LIM_D [OF * this] obtain S where S: "0 < S" and r: "\x. x \ 0 \ norm x < S \ ?r x < ?r h" by auto from dense [OF S] obtain t where t: "0 < t \ t < S" .. let ?x = "scaleR (t / norm h) h" have "?x \ 0" and "norm ?x < S" using t h by simp_all then have "?r ?x < ?r h" by (rule r) then show False using t h by (simp add: F.scaleR) qed qed qed lemma has_derivative_unique: assumes "(f has_derivative F) (at x)" and "(f has_derivative F') (at x)" shows "F = F'" proof - have "((\x. 0) has_derivative (\h. F h - F' h)) (at x)" using has_derivative_diff [OF assms] by simp then have "(\h. F h - F' h) = (\h. 0)" by (rule has_derivative_zero_unique) then show "F = F'" unfolding fun_eq_iff right_minus_eq . qed lemma has_derivative_Uniq: "\\<^sub>\\<^sub>1F. (f has_derivative F) (at x)" by (simp add: Uniq_def has_derivative_unique) subsection \Differentiability predicate\ definition differentiable :: "('a::real_normed_vector \ 'b::real_normed_vector) \ 'a filter \ bool" (infix "differentiable" 50) where "f differentiable F \ (\D. (f has_derivative D) F)" lemma differentiable_subset: "f differentiable (at x within s) \ t \ s \ f differentiable (at x within t)" unfolding differentiable_def by (blast intro: has_derivative_subset) lemmas differentiable_within_subset = differentiable_subset lemma differentiable_ident [simp, derivative_intros]: "(\x. x) differentiable F" unfolding differentiable_def by (blast intro: has_derivative_ident) lemma differentiable_const [simp, derivative_intros]: "(\z. a) differentiable F" unfolding differentiable_def by (blast intro: has_derivative_const) lemma differentiable_in_compose: "f differentiable (at (g x) within (g`s)) \ g differentiable (at x within s) \ (\x. f (g x)) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_in_compose) lemma differentiable_compose: "f differentiable (at (g x)) \ g differentiable (at x within s) \ (\x. f (g x)) differentiable (at x within s)" by (blast intro: differentiable_in_compose differentiable_subset) lemma differentiable_add [simp, derivative_intros]: "f differentiable F \ g differentiable F \ (\x. f x + g x) differentiable F" unfolding differentiable_def by (blast intro: has_derivative_add) lemma differentiable_sum[simp, derivative_intros]: assumes "finite s" "\a\s. (f a) differentiable net" shows "(\x. sum (\a. f a x) s) differentiable net" proof - from bchoice[OF assms(2)[unfolded differentiable_def]] show ?thesis by (auto intro!: has_derivative_sum simp: differentiable_def) qed lemma differentiable_minus [simp, derivative_intros]: "f differentiable F \ (\x. - f x) differentiable F" unfolding differentiable_def by (blast intro: has_derivative_minus) lemma differentiable_diff [simp, derivative_intros]: "f differentiable F \ g differentiable F \ (\x. f x - g x) differentiable F" unfolding differentiable_def by (blast intro: has_derivative_diff) lemma differentiable_mult [simp, derivative_intros]: fixes f g :: "'a::real_normed_vector \ 'b::real_normed_algebra" shows "f differentiable (at x within s) \ g differentiable (at x within s) \ (\x. f x * g x) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_mult) lemma differentiable_cmult_left_iff [simp]: fixes c::"'a::real_normed_field" shows "(\t. c * q t) differentiable at t \ c = 0 \ (\t. q t) differentiable at t" (is "?lhs = ?rhs") proof assume L: ?lhs {assume "c \ 0" then have "q differentiable at t" using differentiable_mult [OF differentiable_const L, of concl: "1/c"] by auto } then show ?rhs by auto qed auto lemma differentiable_cmult_right_iff [simp]: fixes c::"'a::real_normed_field" shows "(\t. q t * c) differentiable at t \ c = 0 \ (\t. q t) differentiable at t" (is "?lhs = ?rhs") by (simp add: mult.commute flip: differentiable_cmult_left_iff) lemma differentiable_inverse [simp, derivative_intros]: fixes f :: "'a::real_normed_vector \ 'b::real_normed_field" shows "f differentiable (at x within s) \ f x \ 0 \ (\x. inverse (f x)) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_inverse) lemma differentiable_divide [simp, derivative_intros]: fixes f g :: "'a::real_normed_vector \ 'b::real_normed_field" shows "f differentiable (at x within s) \ g differentiable (at x within s) \ g x \ 0 \ (\x. f x / g x) differentiable (at x within s)" unfolding divide_inverse by simp lemma differentiable_power [simp, derivative_intros]: fixes f g :: "'a::real_normed_vector \ 'b::real_normed_field" shows "f differentiable (at x within s) \ (\x. f x ^ n) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_power) lemma differentiable_power_int [simp, derivative_intros]: fixes f :: "'a::real_normed_vector \ 'b::real_normed_field" shows "f differentiable (at x within s) \ f x \ 0 \ (\x. power_int (f x) n) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_power_int) lemma differentiable_scaleR [simp, derivative_intros]: "f differentiable (at x within s) \ g differentiable (at x within s) \ (\x. f x *\<^sub>R g x) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_scaleR) lemma has_derivative_imp_has_field_derivative: "(f has_derivative D) F \ (\x. x * D' = D x) \ (f has_field_derivative D') F" unfolding has_field_derivative_def by (rule has_derivative_eq_rhs[of f D]) (simp_all add: fun_eq_iff mult.commute) lemma has_field_derivative_imp_has_derivative: "(f has_field_derivative D) F \ (f has_derivative (*) D) F" by (simp add: has_field_derivative_def) lemma DERIV_subset: "(f has_field_derivative f') (at x within s) \ t \ s \ (f has_field_derivative f') (at x within t)" by (simp add: has_field_derivative_def has_derivative_subset) lemma has_field_derivative_at_within: "(f has_field_derivative f') (at x) \ (f has_field_derivative f') (at x within s)" using DERIV_subset by blast abbreviation (input) DERIV :: "('a::real_normed_field \ 'a) \ 'a \ 'a \ bool" ("(DERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where "DERIV f x :> D \ (f has_field_derivative D) (at x)" abbreviation has_real_derivative :: "(real \ real) \ real \ real filter \ bool" (infix "(has'_real'_derivative)" 50) where "(f has_real_derivative D) F \ (f has_field_derivative D) F" lemma real_differentiable_def: "f differentiable at x within s \ (\D. (f has_real_derivative D) (at x within s))" proof safe assume "f differentiable at x within s" then obtain f' where *: "(f has_derivative f') (at x within s)" unfolding differentiable_def by auto then obtain c where "f' = ((*) c)" by (metis real_bounded_linear has_derivative_bounded_linear mult.commute fun_eq_iff) with * show "\D. (f has_real_derivative D) (at x within s)" unfolding has_field_derivative_def by auto qed (auto simp: differentiable_def has_field_derivative_def) lemma real_differentiableE [elim?]: assumes f: "f differentiable (at x within s)" obtains df where "(f has_real_derivative df) (at x within s)" using assms by (auto simp: real_differentiable_def) lemma has_field_derivative_iff: "(f has_field_derivative D) (at x within S) \ ((\y. (f y - f x) / (y - x)) \ D) (at x within S)" proof - have "((\y. norm (f y - f x - D * (y - x)) / norm (y - x)) \ 0) (at x within S) = ((\y. (f y - f x) / (y - x) - D) \ 0) (at x within S)" apply (subst tendsto_norm_zero_iff[symmetric], rule filterlim_cong) apply (simp_all add: eventually_at_filter field_simps nonzero_norm_divide) done then show ?thesis by (simp add: has_field_derivative_def has_derivative_iff_norm bounded_linear_mult_right LIM_zero_iff) qed lemma DERIV_def: "DERIV f x :> D \ (\h. (f (x + h) - f x) / h) \0\ D" unfolding field_has_derivative_at has_field_derivative_def has_field_derivative_iff .. lemma field_derivative_lim_unique: assumes f: "(f has_field_derivative df) (at z)" and s: "s \ 0" "\n. s n \ 0" and a: "(\n. (f (z + s n) - f z) / s n) \ a" shows "df = a" proof (rule ccontr) assume "df \ a" obtain q where qpos: "\\. \ > 0 \ q \ > 0" and q: "\\ y. \\ > 0; y \ z; dist y z < q \\ \ dist ((f y - f z) / (y - z)) df < \" using f unfolding LIM_def has_field_derivative_iff by metis obtain NA where NA: "\\ n. \\ > 0; n \ NA \\ \ dist ((f (z + s n) - f z) / s n) a < \" using a unfolding LIMSEQ_def by metis obtain NB where NB: "\\ n. \\ > 0; n \ NB \\ \ norm (s n) < \" using s unfolding LIMSEQ_def by (metis norm_conv_dist) have df: "\\ n. \ > 0 \ \0 < \; norm (s n) < q \\ \ dist ((f (z + s n) - f z) / s n) df < \" using add_cancel_left_right add_diff_cancel_left' q s by (metis add_diff_cancel_right' dist_diff(1)) define \ where "\ \ dist df a / 2" with \df \ a\ have "\ > 0" and \: "\+\ \ dist df a" by auto define N where "N \ max (NA \) (NB (q \))" then have "norm (s N) < q \" by (simp add: NB \\ > 0\ qpos) then have "dist ((f (z + s N) - f z) / s N) df < \" by (simp add: \0 < \\ df) moreover have "dist ((f (z + s N) - f z) / s N) a < \" using NA N_def \0 < \\ by force ultimately have "dist df a < dist df a" by (smt (verit, ccfv_SIG) \ dist_commute dist_triangle) then show False .. qed lemma mult_commute_abs: "(\x. x * c) = (*) c" for c :: "'a::ab_semigroup_mult" by (simp add: fun_eq_iff mult.commute) lemma DERIV_compose_FDERIV: fixes f::"real\real" assumes "DERIV f (g x) :> f'" assumes "(g has_derivative g') (at x within s)" shows "((\x. f (g x)) has_derivative (\x. g' x * f')) (at x within s)" using assms has_derivative_compose[of g g' x s f "(*) f'"] by (auto simp: has_field_derivative_def ac_simps) subsection \Vector derivative\ lemma has_field_derivative_iff_has_vector_derivative: "(f has_field_derivative y) F \ (f has_vector_derivative y) F" unfolding has_vector_derivative_def has_field_derivative_def real_scaleR_def mult_commute_abs .. lemma has_field_derivative_subset: "(f has_field_derivative y) (at x within s) \ t \ s \ (f has_field_derivative y) (at x within t)" unfolding has_field_derivative_def by (rule has_derivative_subset) lemma has_vector_derivative_const[simp, derivative_intros]: "((\x. c) has_vector_derivative 0) net" by (auto simp: has_vector_derivative_def) lemma has_vector_derivative_id[simp, derivative_intros]: "((\x. x) has_vector_derivative 1) net" by (auto simp: has_vector_derivative_def) lemma has_vector_derivative_minus[derivative_intros]: "(f has_vector_derivative f') net \ ((\x. - f x) has_vector_derivative (- f')) net" by (auto simp: has_vector_derivative_def) lemma has_vector_derivative_add[derivative_intros]: "(f has_vector_derivative f') net \ (g has_vector_derivative g') net \ ((\x. f x + g x) has_vector_derivative (f' + g')) net" by (auto simp: has_vector_derivative_def scaleR_right_distrib) lemma has_vector_derivative_sum[derivative_intros]: "(\i. i \ I \ (f i has_vector_derivative f' i) net) \ ((\x. \i\I. f i x) has_vector_derivative (\i\I. f' i)) net" by (auto simp: has_vector_derivative_def fun_eq_iff scaleR_sum_right intro!: derivative_eq_intros) lemma has_vector_derivative_diff[derivative_intros]: "(f has_vector_derivative f') net \ (g has_vector_derivative g') net \ ((\x. f x - g x) has_vector_derivative (f' - g')) net" by (auto simp: has_vector_derivative_def scaleR_diff_right) lemma has_vector_derivative_add_const: "((\t. g t + z) has_vector_derivative f') net = ((\t. g t) has_vector_derivative f') net" apply (intro iffI) apply (force dest: has_vector_derivative_diff [where g = "\t. z", OF _ has_vector_derivative_const]) apply (force dest: has_vector_derivative_add [OF _ has_vector_derivative_const]) done lemma has_vector_derivative_diff_const: "((\t. g t - z) has_vector_derivative f') net = ((\t. g t) has_vector_derivative f') net" using has_vector_derivative_add_const [where z = "-z"] by simp lemma (in bounded_linear) has_vector_derivative: assumes "(g has_vector_derivative g') F" shows "((\x. f (g x)) has_vector_derivative f g') F" using has_derivative[OF assms[unfolded has_vector_derivative_def]] by (simp add: has_vector_derivative_def scaleR) lemma (in bounded_bilinear) has_vector_derivative: assumes "(f has_vector_derivative f') (at x within s)" and "(g has_vector_derivative g') (at x within s)" shows "((\x. f x ** g x) has_vector_derivative (f x ** g' + f' ** g x)) (at x within s)" using FDERIV[OF assms(1-2)[unfolded has_vector_derivative_def]] by (simp add: has_vector_derivative_def scaleR_right scaleR_left scaleR_right_distrib) lemma has_vector_derivative_scaleR[derivative_intros]: "(f has_field_derivative f') (at x within s) \ (g has_vector_derivative g') (at x within s) \ ((\x. f x *\<^sub>R g x) has_vector_derivative (f x *\<^sub>R g' + f' *\<^sub>R g x)) (at x within s)" unfolding has_field_derivative_iff_has_vector_derivative by (rule bounded_bilinear.has_vector_derivative[OF bounded_bilinear_scaleR]) lemma has_vector_derivative_mult[derivative_intros]: "(f has_vector_derivative f') (at x within s) \ (g has_vector_derivative g') (at x within s) \ ((\x. f x * g x) has_vector_derivative (f x * g' + f' * g x)) (at x within s)" for f g :: "real \ 'a::real_normed_algebra" by (rule bounded_bilinear.has_vector_derivative[OF bounded_bilinear_mult]) lemma has_vector_derivative_of_real[derivative_intros]: "(f has_field_derivative D) F \ ((\x. of_real (f x)) has_vector_derivative (of_real D)) F" by (rule bounded_linear.has_vector_derivative[OF bounded_linear_of_real]) (simp add: has_field_derivative_iff_has_vector_derivative) lemma has_vector_derivative_real_field: "(f has_field_derivative f') (at (of_real a)) \ ((\x. f (of_real x)) has_vector_derivative f') (at a within s)" using has_derivative_compose[of of_real of_real a _ f "(*) f'"] by (simp add: scaleR_conv_of_real ac_simps has_vector_derivative_def has_field_derivative_def) lemma has_vector_derivative_continuous: "(f has_vector_derivative D) (at x within s) \ continuous (at x within s) f" by (auto intro: has_derivative_continuous simp: has_vector_derivative_def) lemma continuous_on_vector_derivative: "(\x. x \ S \ (f has_vector_derivative f' x) (at x within S)) \ continuous_on S f" by (auto simp: continuous_on_eq_continuous_within intro!: has_vector_derivative_continuous) lemma has_vector_derivative_mult_right[derivative_intros]: fixes a :: "'a::real_normed_algebra" shows "(f has_vector_derivative x) F \ ((\x. a * f x) has_vector_derivative (a * x)) F" by (rule bounded_linear.has_vector_derivative[OF bounded_linear_mult_right]) lemma has_vector_derivative_mult_left[derivative_intros]: fixes a :: "'a::real_normed_algebra" shows "(f has_vector_derivative x) F \ ((\x. f x * a) has_vector_derivative (x * a)) F" by (rule bounded_linear.has_vector_derivative[OF bounded_linear_mult_left]) +lemma has_vector_derivative_divide[derivative_intros]: + fixes a :: "'a::real_normed_field" + shows "(f has_vector_derivative x) F \ ((\x. f x / a) has_vector_derivative (x / a)) F" + using has_vector_derivative_mult_left [of f x F "inverse a"] + by (simp add: field_class.field_divide_inverse) + subsection \Derivatives\ lemma DERIV_D: "DERIV f x :> D \ (\h. (f (x + h) - f x) / h) \0\ D" by (simp add: DERIV_def) lemma has_field_derivativeD: "(f has_field_derivative D) (at x within S) \ ((\y. (f y - f x) / (y - x)) \ D) (at x within S)" by (simp add: has_field_derivative_iff) lemma DERIV_const [simp, derivative_intros]: "((\x. k) has_field_derivative 0) F" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_const]) auto lemma DERIV_ident [simp, derivative_intros]: "((\x. x) has_field_derivative 1) F" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_ident]) auto lemma field_differentiable_add[derivative_intros]: "(f has_field_derivative f') F \ (g has_field_derivative g') F \ ((\z. f z + g z) has_field_derivative f' + g') F" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_add]) (auto simp: has_field_derivative_def field_simps mult_commute_abs) corollary DERIV_add: "(f has_field_derivative D) (at x within s) \ (g has_field_derivative E) (at x within s) \ ((\x. f x + g x) has_field_derivative D + E) (at x within s)" by (rule field_differentiable_add) lemma field_differentiable_minus[derivative_intros]: "(f has_field_derivative f') F \ ((\z. - (f z)) has_field_derivative -f') F" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_minus]) (auto simp: has_field_derivative_def field_simps mult_commute_abs) corollary DERIV_minus: "(f has_field_derivative D) (at x within s) \ ((\x. - f x) has_field_derivative -D) (at x within s)" by (rule field_differentiable_minus) lemma field_differentiable_diff[derivative_intros]: "(f has_field_derivative f') F \ (g has_field_derivative g') F \ ((\z. f z - g z) has_field_derivative f' - g') F" by (simp only: diff_conv_add_uminus field_differentiable_add field_differentiable_minus) corollary DERIV_diff: "(f has_field_derivative D) (at x within s) \ (g has_field_derivative E) (at x within s) \ ((\x. f x - g x) has_field_derivative D - E) (at x within s)" by (rule field_differentiable_diff) lemma DERIV_continuous: "(f has_field_derivative D) (at x within s) \ continuous (at x within s) f" by (drule has_derivative_continuous[OF has_field_derivative_imp_has_derivative]) simp corollary DERIV_isCont: "DERIV f x :> D \ isCont f x" by (rule DERIV_continuous) lemma DERIV_atLeastAtMost_imp_continuous_on: assumes "\x. \a \ x; x \ b\ \ \y. DERIV f x :> y" shows "continuous_on {a..b} f" by (meson DERIV_isCont assms atLeastAtMost_iff continuous_at_imp_continuous_at_within continuous_on_eq_continuous_within) lemma DERIV_continuous_on: "(\x. x \ s \ (f has_field_derivative (D x)) (at x within s)) \ continuous_on s f" unfolding continuous_on_eq_continuous_within by (intro continuous_at_imp_continuous_on ballI DERIV_continuous) lemma DERIV_mult': "(f has_field_derivative D) (at x within s) \ (g has_field_derivative E) (at x within s) \ ((\x. f x * g x) has_field_derivative f x * E + D * g x) (at x within s)" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_mult]) (auto simp: field_simps mult_commute_abs dest: has_field_derivative_imp_has_derivative) lemma DERIV_mult[derivative_intros]: "(f has_field_derivative Da) (at x within s) \ (g has_field_derivative Db) (at x within s) \ ((\x. f x * g x) has_field_derivative Da * g x + Db * f x) (at x within s)" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_mult]) (auto simp: field_simps dest: has_field_derivative_imp_has_derivative) text \Derivative of linear multiplication\ lemma DERIV_cmult: "(f has_field_derivative D) (at x within s) \ ((\x. c * f x) has_field_derivative c * D) (at x within s)" by (drule DERIV_mult' [OF DERIV_const]) simp lemma DERIV_cmult_right: "(f has_field_derivative D) (at x within s) \ ((\x. f x * c) has_field_derivative D * c) (at x within s)" using DERIV_cmult by (auto simp add: ac_simps) lemma DERIV_cmult_Id [simp]: "((*) c has_field_derivative c) (at x within s)" using DERIV_ident [THEN DERIV_cmult, where c = c and x = x] by simp lemma DERIV_cdivide: "(f has_field_derivative D) (at x within s) \ ((\x. f x / c) has_field_derivative D / c) (at x within s)" using DERIV_cmult_right[of f D x s "1 / c"] by simp lemma DERIV_unique: "DERIV f x :> D \ DERIV f x :> E \ D = E" unfolding DERIV_def by (rule LIM_unique) lemma DERIV_Uniq: "\\<^sub>\\<^sub>1D. DERIV f x :> D" by (simp add: DERIV_unique Uniq_def) lemma DERIV_sum[derivative_intros]: "(\ n. n \ S \ ((\x. f x n) has_field_derivative (f' x n)) F) \ ((\x. sum (f x) S) has_field_derivative sum (f' x) S) F" by (rule has_derivative_imp_has_field_derivative [OF has_derivative_sum]) (auto simp: sum_distrib_left mult_commute_abs dest: has_field_derivative_imp_has_derivative) lemma DERIV_inverse'[derivative_intros]: assumes "(f has_field_derivative D) (at x within s)" and "f x \ 0" shows "((\x. inverse (f x)) has_field_derivative - (inverse (f x) * D * inverse (f x))) (at x within s)" proof - have "(f has_derivative (\x. x * D)) = (f has_derivative (*) D)" by (rule arg_cong [of "\x. x * D"]) (simp add: fun_eq_iff) with assms have "(f has_derivative (\x. x * D)) (at x within s)" by (auto dest!: has_field_derivative_imp_has_derivative) then show ?thesis using \f x \ 0\ by (auto intro: has_derivative_imp_has_field_derivative has_derivative_inverse) qed text \Power of \-1\\ lemma DERIV_inverse: "x \ 0 \ ((\x. inverse(x)) has_field_derivative - (inverse x ^ Suc (Suc 0))) (at x within s)" by (drule DERIV_inverse' [OF DERIV_ident]) simp text \Derivative of inverse\ lemma DERIV_inverse_fun: "(f has_field_derivative d) (at x within s) \ f x \ 0 \ ((\x. inverse (f x)) has_field_derivative (- (d * inverse(f x ^ Suc (Suc 0))))) (at x within s)" by (drule (1) DERIV_inverse') (simp add: ac_simps nonzero_inverse_mult_distrib) text \Derivative of quotient\ lemma DERIV_divide[derivative_intros]: "(f has_field_derivative D) (at x within s) \ (g has_field_derivative E) (at x within s) \ g x \ 0 \ ((\x. f x / g x) has_field_derivative (D * g x - f x * E) / (g x * g x)) (at x within s)" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_divide]) (auto dest: has_field_derivative_imp_has_derivative simp: field_simps) lemma DERIV_quotient: "(f has_field_derivative d) (at x within s) \ (g has_field_derivative e) (at x within s)\ g x \ 0 \ ((\y. f y / g y) has_field_derivative (d * g x - (e * f x)) / (g x ^ Suc (Suc 0))) (at x within s)" by (drule (2) DERIV_divide) (simp add: mult.commute) lemma DERIV_power_Suc: "(f has_field_derivative D) (at x within s) \ ((\x. f x ^ Suc n) has_field_derivative (1 + of_nat n) * (D * f x ^ n)) (at x within s)" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_power]) (auto simp: has_field_derivative_def) lemma DERIV_power[derivative_intros]: "(f has_field_derivative D) (at x within s) \ ((\x. f x ^ n) has_field_derivative of_nat n * (D * f x ^ (n - Suc 0))) (at x within s)" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_power]) (auto simp: has_field_derivative_def) lemma DERIV_pow: "((\x. x ^ n) has_field_derivative real n * (x ^ (n - Suc 0))) (at x within s)" using DERIV_power [OF DERIV_ident] by simp lemma DERIV_power_int [derivative_intros]: assumes [derivative_intros]: "(f has_field_derivative d) (at x within s)" and [simp]: "f x \ 0" shows "((\x. power_int (f x) n) has_field_derivative (of_int n * power_int (f x) (n - 1) * d)) (at x within s)" proof (cases n rule: int_cases4) case (nonneg n) thus ?thesis by (cases "n = 0") (auto intro!: derivative_eq_intros simp: field_simps power_int_diff simp flip: power_Suc power_Suc2 power_add) next case (neg n) thus ?thesis by (auto intro!: derivative_eq_intros simp: field_simps power_int_diff power_int_minus simp flip: power_Suc power_Suc2 power_add) qed lemma DERIV_chain': "(f has_field_derivative D) (at x within s) \ DERIV g (f x) :> E \ ((\x. g (f x)) has_field_derivative E * D) (at x within s)" using has_derivative_compose[of f "(*) D" x s g "(*) E"] by (simp only: has_field_derivative_def mult_commute_abs ac_simps) corollary DERIV_chain2: "DERIV f (g x) :> Da \ (g has_field_derivative Db) (at x within s) \ ((\x. f (g x)) has_field_derivative Da * Db) (at x within s)" by (rule DERIV_chain') text \Standard version\ lemma DERIV_chain: "DERIV f (g x) :> Da \ (g has_field_derivative Db) (at x within s) \ (f \ g has_field_derivative Da * Db) (at x within s)" by (drule (1) DERIV_chain', simp add: o_def mult.commute) lemma DERIV_image_chain: "(f has_field_derivative Da) (at (g x) within (g ` s)) \ (g has_field_derivative Db) (at x within s) \ (f \ g has_field_derivative Da * Db) (at x within s)" using has_derivative_in_compose [of g "(*) Db" x s f "(*) Da "] by (simp add: has_field_derivative_def o_def mult_commute_abs ac_simps) (*These two are from HOL Light: HAS_COMPLEX_DERIVATIVE_CHAIN*) lemma DERIV_chain_s: assumes "(\x. x \ s \ DERIV g x :> g'(x))" and "DERIV f x :> f'" and "f x \ s" shows "DERIV (\x. g(f x)) x :> f' * g'(f x)" by (metis (full_types) DERIV_chain' mult.commute assms) lemma DERIV_chain3: (*HAS_COMPLEX_DERIVATIVE_CHAIN_UNIV*) assumes "(\x. DERIV g x :> g'(x))" and "DERIV f x :> f'" shows "DERIV (\x. g(f x)) x :> f' * g'(f x)" by (metis UNIV_I DERIV_chain_s [of UNIV] assms) text \Alternative definition for differentiability\ lemma DERIV_LIM_iff: fixes f :: "'a::{real_normed_vector,inverse} \ 'a" shows "((\h. (f (a + h) - f a) / h) \0\ D) = ((\x. (f x - f a) / (x - a)) \a\ D)" (is "?lhs = ?rhs") proof assume ?lhs then have "(\x. (f (a + (x + - a)) - f a) / (x + - a)) \0 - - a\ D" by (rule LIM_offset) then show ?rhs by simp next assume ?rhs then have "(\x. (f (x+a) - f a) / ((x+a) - a)) \a-a\ D" by (rule LIM_offset) then show ?lhs by (simp add: add.commute) qed lemma has_field_derivative_cong_ev: assumes "x = y" and *: "eventually (\x. x \ S \ f x = g x) (nhds x)" and "u = v" "S = t" "x \ S" shows "(f has_field_derivative u) (at x within S) = (g has_field_derivative v) (at y within t)" unfolding has_field_derivative_iff proof (rule filterlim_cong) from assms have "f y = g y" by (auto simp: eventually_nhds) with * show "\\<^sub>F z in at x within S. (f z - f x) / (z - x) = (g z - g y) / (z - y)" unfolding eventually_at_filter by eventually_elim (auto simp: assms \f y = g y\) qed (simp_all add: assms) lemma has_field_derivative_cong_eventually: assumes "eventually (\x. f x = g x) (at x within S)" "f x = g x" shows "(f has_field_derivative u) (at x within S) = (g has_field_derivative u) (at x within S)" unfolding has_field_derivative_iff proof (rule tendsto_cong) show "\\<^sub>F y in at x within S. (f y - f x) / (y - x) = (g y - g x) / (y - x)" using assms by (auto elim: eventually_mono) qed lemma DERIV_cong_ev: "x = y \ eventually (\x. f x = g x) (nhds x) \ u = v \ DERIV f x :> u \ DERIV g y :> v" by (rule has_field_derivative_cong_ev) simp_all lemma DERIV_mirror: "(DERIV f (- x) :> y) \ (DERIV (\x. f (- x)) x :> - y)" for f :: "real \ real" and x y :: real by (simp add: DERIV_def filterlim_at_split filterlim_at_left_to_right tendsto_minus_cancel_left field_simps conj_commute) lemma DERIV_shift: "(f has_field_derivative y) (at (x + z)) = ((\x. f (x + z)) has_field_derivative y) (at x)" by (simp add: DERIV_def field_simps) lemma DERIV_at_within_shift_lemma: assumes "(f has_field_derivative y) (at (z+x) within (+) z ` S)" shows "(f \ (+)z has_field_derivative y) (at x within S)" proof - have "((+)z has_field_derivative 1) (at x within S)" by (rule derivative_eq_intros | simp)+ with assms DERIV_image_chain show ?thesis by (metis mult.right_neutral) qed lemma DERIV_at_within_shift: "(f has_field_derivative y) (at (z+x) within (+) z ` S) \ ((\x. f (z+x)) has_field_derivative y) (at x within S)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using DERIV_at_within_shift_lemma unfolding o_def by blast next have [simp]: "(\x. x - z) ` (+) z ` S = S" by force assume R: ?rhs have "(f \ (+) z \ (+) (- z) has_field_derivative y) (at (z + x) within (+) z ` S)" by (rule DERIV_at_within_shift_lemma) (use R in \simp add: o_def\) then show ?lhs by (simp add: o_def) qed lemma floor_has_real_derivative: fixes f :: "real \ 'a::{floor_ceiling,order_topology}" assumes "isCont f x" and "f x \ \" shows "((\x. floor (f x)) has_real_derivative 0) (at x)" proof (subst DERIV_cong_ev[OF refl _ refl]) show "((\_. floor (f x)) has_real_derivative 0) (at x)" by simp have "\\<^sub>F y in at x. \f y\ = \f x\" by (rule eventually_floor_eq[OF assms[unfolded continuous_at]]) then show "\\<^sub>F y in nhds x. real_of_int \f y\ = real_of_int \f x\" unfolding eventually_at_filter by eventually_elim auto qed lemmas has_derivative_floor[derivative_intros] = floor_has_real_derivative[THEN DERIV_compose_FDERIV] lemma continuous_floor: fixes x::real shows "x \ \ \ continuous (at x) (real_of_int \ floor)" using floor_has_real_derivative [where f=id] by (auto simp: o_def has_field_derivative_def intro: has_derivative_continuous) lemma continuous_frac: fixes x::real assumes "x \ \" shows "continuous (at x) frac" proof - have "isCont (\x. real_of_int \x\) x" using continuous_floor [OF assms] by (simp add: o_def) then have *: "continuous (at x) (\x. x - real_of_int \x\)" by (intro continuous_intros) moreover have "\\<^sub>F x in nhds x. frac x = x - real_of_int \x\" by (simp add: frac_def) ultimately show ?thesis by (simp add: LIM_imp_LIM frac_def isCont_def) qed text \Caratheodory formulation of derivative at a point\ lemma CARAT_DERIV: "(DERIV f x :> l) \ (\g. (\z. f z - f x = g z * (z - x)) \ isCont g x \ g x = l)" (is "?lhs = ?rhs") proof assume ?lhs show "\g. (\z. f z - f x = g z * (z - x)) \ isCont g x \ g x = l" proof (intro exI conjI) let ?g = "(\z. if z = x then l else (f z - f x) / (z-x))" show "\z. f z - f x = ?g z * (z - x)" by simp show "isCont ?g x" using \?lhs\ by (simp add: isCont_iff DERIV_def cong: LIM_equal [rule_format]) show "?g x = l" by simp qed next assume ?rhs then show ?lhs by (auto simp add: isCont_iff DERIV_def cong: LIM_cong) qed subsection \Local extrema\ text \If \<^term>\0 < f' x\ then \<^term>\x\ is Locally Strictly Increasing At The Right.\ lemma has_real_derivative_pos_inc_right: fixes f :: "real \ real" assumes der: "(f has_real_derivative l) (at x within S)" and l: "0 < l" shows "\d > 0. \h > 0. x + h \ S \ h < d \ f x < f (x + h)" using assms proof - from der [THEN has_field_derivativeD, THEN tendstoD, OF l, unfolded eventually_at] obtain s where s: "0 < s" and all: "\xa. xa\S \ xa \ x \ dist xa x < s \ \(f xa - f x) / (xa - x) - l\ < l" by (auto simp: dist_real_def) then show ?thesis proof (intro exI conjI strip) show "0 < s" by (rule s) next fix h :: real assume "0 < h" "h < s" "x + h \ S" with all [of "x + h"] show "f x < f (x+h)" proof (simp add: abs_if dist_real_def pos_less_divide_eq split: if_split_asm) assume "\ (f (x + h) - f x) / h < l" and h: "0 < h" with l have "0 < (f (x + h) - f x) / h" by arith then show "f x < f (x + h)" by (simp add: pos_less_divide_eq h) qed qed qed lemma DERIV_pos_inc_right: fixes f :: "real \ real" assumes der: "DERIV f x :> l" and l: "0 < l" shows "\d > 0. \h > 0. h < d \ f x < f (x + h)" using has_real_derivative_pos_inc_right[OF assms] by auto lemma has_real_derivative_neg_dec_left: fixes f :: "real \ real" assumes der: "(f has_real_derivative l) (at x within S)" and "l < 0" shows "\d > 0. \h > 0. x - h \ S \ h < d \ f x < f (x - h)" proof - from \l < 0\ have l: "- l > 0" by simp from der [THEN has_field_derivativeD, THEN tendstoD, OF l, unfolded eventually_at] obtain s where s: "0 < s" and all: "\xa. xa\S \ xa \ x \ dist xa x < s \ \(f xa - f x) / (xa - x) - l\ < - l" by (auto simp: dist_real_def) then show ?thesis proof (intro exI conjI strip) show "0 < s" by (rule s) next fix h :: real assume "0 < h" "h < s" "x - h \ S" with all [of "x - h"] show "f x < f (x-h)" proof (simp add: abs_if pos_less_divide_eq dist_real_def split: if_split_asm) assume "- ((f (x-h) - f x) / h) < l" and h: "0 < h" with l have "0 < (f (x-h) - f x) / h" by arith then show "f x < f (x - h)" by (simp add: pos_less_divide_eq h) qed qed qed lemma DERIV_neg_dec_left: fixes f :: "real \ real" assumes der: "DERIV f x :> l" and l: "l < 0" shows "\d > 0. \h > 0. h < d \ f x < f (x - h)" using has_real_derivative_neg_dec_left[OF assms] by auto lemma has_real_derivative_pos_inc_left: fixes f :: "real \ real" shows "(f has_real_derivative l) (at x within S) \ 0 < l \ \d>0. \h>0. x - h \ S \ h < d \ f (x - h) < f x" by (rule has_real_derivative_neg_dec_left [of "\x. - f x" "-l" x S, simplified]) (auto simp add: DERIV_minus) lemma DERIV_pos_inc_left: fixes f :: "real \ real" shows "DERIV f x :> l \ 0 < l \ \d > 0. \h > 0. h < d \ f (x - h) < f x" using has_real_derivative_pos_inc_left by blast lemma has_real_derivative_neg_dec_right: fixes f :: "real \ real" shows "(f has_real_derivative l) (at x within S) \ l < 0 \ \d > 0. \h > 0. x + h \ S \ h < d \ f x > f (x + h)" by (rule has_real_derivative_pos_inc_right [of "\x. - f x" "-l" x S, simplified]) (auto simp add: DERIV_minus) lemma DERIV_neg_dec_right: fixes f :: "real \ real" shows "DERIV f x :> l \ l < 0 \ \d > 0. \h > 0. h < d \ f x > f (x + h)" using has_real_derivative_neg_dec_right by blast lemma DERIV_local_max: fixes f :: "real \ real" assumes der: "DERIV f x :> l" and d: "0 < d" and le: "\y. \x - y\ < d \ f y \ f x" shows "l = 0" proof (cases rule: linorder_cases [of l 0]) case equal then show ?thesis . next case less from DERIV_neg_dec_left [OF der less] obtain d' where d': "0 < d'" and lt: "\h > 0. h < d' \ f x < f (x - h)" by blast obtain e where "0 < e \ e < d \ e < d'" using field_lbound_gt_zero [OF d d'] .. with lt le [THEN spec [where x="x - e"]] show ?thesis by (auto simp add: abs_if) next case greater from DERIV_pos_inc_right [OF der greater] obtain d' where d': "0 < d'" and lt: "\h > 0. h < d' \ f x < f (x + h)" by blast obtain e where "0 < e \ e < d \ e < d'" using field_lbound_gt_zero [OF d d'] .. with lt le [THEN spec [where x="x + e"]] show ?thesis by (auto simp add: abs_if) qed text \Similar theorem for a local minimum\ lemma DERIV_local_min: fixes f :: "real \ real" shows "DERIV f x :> l \ 0 < d \ \y. \x - y\ < d \ f x \ f y \ l = 0" by (drule DERIV_minus [THEN DERIV_local_max]) auto text\In particular, if a function is locally flat\ lemma DERIV_local_const: fixes f :: "real \ real" shows "DERIV f x :> l \ 0 < d \ \y. \x - y\ < d \ f x = f y \ l = 0" by (auto dest!: DERIV_local_max) subsection \Rolle's Theorem\ text \Lemma about introducing open ball in open interval\ lemma lemma_interval_lt: fixes a b x :: real assumes "a < x" "x < b" shows "\d. 0 < d \ (\y. \x - y\ < d \ a < y \ y < b)" using linorder_linear [of "x - a" "b - x"] proof assume "x - a \ b - x" with assms show ?thesis by (rule_tac x = "x - a" in exI) auto next assume "b - x \ x - a" with assms show ?thesis by (rule_tac x = "b - x" in exI) auto qed lemma lemma_interval: "a < x \ x < b \ \d. 0 < d \ (\y. \x - y\ < d \ a \ y \ y \ b)" for a b x :: real by (force dest: lemma_interval_lt) text \Rolle's Theorem. If \<^term>\f\ is defined and continuous on the closed interval \[a,b]\ and differentiable on the open interval \(a,b)\, and \<^term>\f a = f b\, then there exists \x0 \ (a,b)\ such that \<^term>\f' x0 = 0\\ theorem Rolle_deriv: fixes f :: "real \ real" assumes "a < b" and fab: "f a = f b" and contf: "continuous_on {a..b} f" and derf: "\x. \a < x; x < b\ \ (f has_derivative f' x) (at x)" shows "\z. a < z \ z < b \ f' z = (\v. 0)" proof - have le: "a \ b" using \a < b\ by simp have "(a + b) / 2 \ {a..b}" using assms(1) by auto then have *: "{a..b} \ {}" by auto obtain x where x_max: "\z. a \ z \ z \ b \ f z \ f x" and "a \ x" "x \ b" using continuous_attains_sup[OF compact_Icc * contf] by (meson atLeastAtMost_iff) obtain x' where x'_min: "\z. a \ z \ z \ b \ f x' \ f z" and "a \ x'" "x' \ b" using continuous_attains_inf[OF compact_Icc * contf] by (meson atLeastAtMost_iff) consider "a < x" "x < b" | "x = a \ x = b" using \a \ x\ \x \ b\ by arith then show ?thesis proof cases case 1 \ \\<^term>\f\ attains its maximum within the interval\ then obtain l where der: "DERIV f x :> l" using derf differentiable_def real_differentiable_def by blast obtain d where d: "0 < d" and bound: "\y. \x - y\ < d \ a \ y \ y \ b" using lemma_interval [OF 1] by blast then have bound': "\y. \x - y\ < d \ f y \ f x" using x_max by blast \ \the derivative at a local maximum is zero\ have "l = 0" by (rule DERIV_local_max [OF der d bound']) with 1 der derf [of x] show ?thesis by (metis has_derivative_unique has_field_derivative_def mult_zero_left) next case 2 then have fx: "f b = f x" by (auto simp add: fab) consider "a < x'" "x' < b" | "x' = a \ x' = b" using \a \ x'\ \x' \ b\ by arith then show ?thesis proof cases case 1 \ \\<^term>\f\ attains its minimum within the interval\ then obtain l where der: "DERIV f x' :> l" using derf differentiable_def real_differentiable_def by blast from lemma_interval [OF 1] obtain d where d: "0y. \x'-y\ < d \ a \ y \ y \ b" by blast then have bound': "\y. \x' - y\ < d \ f x' \ f y" using x'_min by blast have "l = 0" by (rule DERIV_local_min [OF der d bound']) \ \the derivative at a local minimum is zero\ then show ?thesis using 1 der derf [of x'] by (metis has_derivative_unique has_field_derivative_def mult_zero_left) next case 2 \ \\<^term>\f\ is constant throughout the interval\ then have fx': "f b = f x'" by (auto simp: fab) from dense [OF \a < b\] obtain r where r: "a < r" "r < b" by blast obtain d where d: "0 < d" and bound: "\y. \r - y\ < d \ a \ y \ y \ b" using lemma_interval [OF r] by blast have eq_fb: "f z = f b" if "a \ z" and "z \ b" for z proof (rule order_antisym) show "f z \ f b" by (simp add: fx x_max that) show "f b \ f z" by (simp add: fx' x'_min that) qed have bound': "\y. \r - y\ < d \ f r = f y" proof (intro strip) fix y :: real assume lt: "\r - y\ < d" then have "f y = f b" by (simp add: eq_fb bound) then show "f r = f y" by (simp add: eq_fb r order_less_imp_le) qed obtain l where der: "DERIV f r :> l" using derf differentiable_def r(1) r(2) real_differentiable_def by blast have "l = 0" by (rule DERIV_local_const [OF der d bound']) \ \the derivative of a constant function is zero\ with r der derf [of r] show ?thesis by (metis has_derivative_unique has_field_derivative_def mult_zero_left) qed qed qed corollary Rolle: fixes a b :: real assumes ab: "a < b" "f a = f b" "continuous_on {a..b} f" and dif [rule_format]: "\x. \a < x; x < b\ \ f differentiable (at x)" shows "\z. a < z \ z < b \ DERIV f z :> 0" proof - obtain f' where f': "\x. \a < x; x < b\ \ (f has_derivative f' x) (at x)" using dif unfolding differentiable_def by metis then have "\z. a < z \ z < b \ f' z = (\v. 0)" by (metis Rolle_deriv [OF ab]) then show ?thesis using f' has_derivative_imp_has_field_derivative by fastforce qed subsection \Mean Value Theorem\ theorem mvt: fixes f :: "real \ real" assumes "a < b" and contf: "continuous_on {a..b} f" and derf: "\x. \a < x; x < b\ \ (f has_derivative f' x) (at x)" obtains \ where "a < \" "\ < b" "f b - f a = (f' \) (b - a)" proof - have "\x. a < x \ x < b \ (\y. f' x y - (f b - f a) / (b - a) * y) = (\v. 0)" proof (intro Rolle_deriv[OF \a < b\]) fix x assume x: "a < x" "x < b" show "((\x. f x - (f b - f a) / (b - a) * x) has_derivative (\y. f' x y - (f b - f a) / (b - a) * y)) (at x)" by (intro derivative_intros derf[OF x]) qed (use assms in \auto intro!: continuous_intros simp: field_simps\) then obtain \ where "a < \" "\ < b" "(\y. f' \ y - (f b - f a) / (b - a) * y) = (\v. 0)" by metis then show ?thesis by (metis (no_types, opaque_lifting) that add.right_neutral add_diff_cancel_left' add_diff_eq \a < b\ less_irrefl nonzero_eq_divide_eq) qed theorem MVT: fixes a b :: real assumes lt: "a < b" and contf: "continuous_on {a..b} f" and dif: "\x. \a < x; x < b\ \ f differentiable (at x)" shows "\l z. a < z \ z < b \ DERIV f z :> l \ f b - f a = (b - a) * l" proof - obtain f' :: "real \ real \ real" where derf: "\x. a < x \ x < b \ (f has_derivative f' x) (at x)" using dif unfolding differentiable_def by metis then obtain z where "a < z" "z < b" "f b - f a = (f' z) (b - a)" using mvt [OF lt contf] by blast then show ?thesis by (simp add: ac_simps) (metis derf dif has_derivative_unique has_field_derivative_imp_has_derivative real_differentiable_def) qed corollary MVT2: assumes "a < b" and der: "\x. \a \ x; x \ b\ \ DERIV f x :> f' x" shows "\z::real. a < z \ z < b \ (f b - f a = (b - a) * f' z)" proof - have "\l z. a < z \ z < b \ (f has_real_derivative l) (at z) \ f b - f a = (b - a) * l" proof (rule MVT [OF \a < b\]) show "continuous_on {a..b} f" by (meson DERIV_continuous atLeastAtMost_iff continuous_at_imp_continuous_on der) show "\x. \a < x; x < b\ \ f differentiable (at x)" using assms by (force dest: order_less_imp_le simp add: real_differentiable_def) qed with assms show ?thesis by (blast dest: DERIV_unique order_less_imp_le) qed lemma pos_deriv_imp_strict_mono: assumes "\x. (f has_real_derivative f' x) (at x)" assumes "\x. f' x > 0" shows "strict_mono f" proof (rule strict_monoI) fix x y :: real assume xy: "x < y" from assms and xy have "\z>x. z < y \ f y - f x = (y - x) * f' z" by (intro MVT2) (auto dest: connectedD_interval) then obtain z where z: "z > x" "z < y" "f y - f x = (y - x) * f' z" by blast note \f y - f x = (y - x) * f' z\ also have "(y - x) * f' z > 0" using xy assms by (intro mult_pos_pos) auto finally show "f x < f y" by simp qed proposition deriv_nonneg_imp_mono: assumes deriv: "\x. x \ {a..b} \ (g has_real_derivative g' x) (at x)" assumes nonneg: "\x. x \ {a..b} \ g' x \ 0" assumes ab: "a \ b" shows "g a \ g b" proof (cases "a < b") assume "a < b" from deriv have "\x. \x \ a; x \ b\ \ (g has_real_derivative g' x) (at x)" by simp with MVT2[OF \a < b\] and deriv obtain \ where \_ab: "\ > a" "\ < b" and g_ab: "g b - g a = (b - a) * g' \" by blast from \_ab ab nonneg have "(b - a) * g' \ \ 0" by simp with g_ab show ?thesis by simp qed (insert ab, simp) subsubsection \A function is constant if its derivative is 0 over an interval.\ lemma DERIV_isconst_end: fixes f :: "real \ real" assumes "a < b" and contf: "continuous_on {a..b} f" and 0: "\x. \a < x; x < b\ \ DERIV f x :> 0" shows "f b = f a" using MVT [OF \a < b\] "0" DERIV_unique contf real_differentiable_def by (fastforce simp: algebra_simps) lemma DERIV_isconst2: fixes f :: "real \ real" assumes "a < b" and contf: "continuous_on {a..b} f" and derf: "\x. \a < x; x < b\ \ DERIV f x :> 0" and "a \ x" "x \ b" shows "f x = f a" proof (cases "a < x") case True have *: "continuous_on {a..x} f" using \x \ b\ contf continuous_on_subset by fastforce show ?thesis by (rule DERIV_isconst_end [OF True *]) (use \x \ b\ derf in auto) qed (use \a \ x\ in auto) lemma DERIV_isconst3: fixes a b x y :: real assumes "a < b" and "x \ {a <..< b}" and "y \ {a <..< b}" and derivable: "\x. x \ {a <..< b} \ DERIV f x :> 0" shows "f x = f y" proof (cases "x = y") case False let ?a = "min x y" let ?b = "max x y" have *: "DERIV f z :> 0" if "?a \ z" "z \ ?b" for z proof - have "a < z" and "z < b" using that \x \ {a <..< b}\ and \y \ {a <..< b}\ by auto then have "z \ {a<.. 0" by (rule derivable) qed have isCont: "continuous_on {?a..?b} f" by (meson * DERIV_continuous_on atLeastAtMost_iff has_field_derivative_at_within) have DERIV: "\z. \?a < z; z < ?b\ \ DERIV f z :> 0" using * by auto have "?a < ?b" using \x \ y\ by auto from DERIV_isconst2[OF this isCont DERIV, of x] and DERIV_isconst2[OF this isCont DERIV, of y] show ?thesis by auto qed auto lemma DERIV_isconst_all: fixes f :: "real \ real" shows "\x. DERIV f x :> 0 \ f x = f y" apply (rule linorder_cases [of x y]) apply (metis DERIV_continuous DERIV_isconst_end continuous_at_imp_continuous_on)+ done lemma DERIV_const_ratio_const: fixes f :: "real \ real" assumes "a \ b" and df: "\x. DERIV f x :> k" shows "f b - f a = (b - a) * k" proof (cases a b rule: linorder_cases) case less show ?thesis using MVT [OF less] df by (metis DERIV_continuous DERIV_unique continuous_at_imp_continuous_on real_differentiable_def) next case greater have "f a - f b = (a - b) * k" using MVT [OF greater] df by (metis DERIV_continuous DERIV_unique continuous_at_imp_continuous_on real_differentiable_def) then show ?thesis by (simp add: algebra_simps) qed auto lemma DERIV_const_ratio_const2: fixes f :: "real \ real" assumes "a \ b" and df: "\x. DERIV f x :> k" shows "(f b - f a) / (b - a) = k" using DERIV_const_ratio_const [OF assms] \a \ b\ by auto lemma real_average_minus_first [simp]: "(a + b) / 2 - a = (b - a) / 2" for a b :: real by simp lemma real_average_minus_second [simp]: "(b + a) / 2 - a = (b - a) / 2" for a b :: real by simp text \Gallileo's "trick": average velocity = av. of end velocities.\ lemma DERIV_const_average: fixes v :: "real \ real" and a b :: real assumes neq: "a \ b" and der: "\x. DERIV v x :> k" shows "v ((a + b) / 2) = (v a + v b) / 2" proof (cases rule: linorder_cases [of a b]) case equal with neq show ?thesis by simp next case less have "(v b - v a) / (b - a) = k" by (rule DERIV_const_ratio_const2 [OF neq der]) then have "(b - a) * ((v b - v a) / (b - a)) = (b - a) * k" by simp moreover have "(v ((a + b) / 2) - v a) / ((a + b) / 2 - a) = k" by (rule DERIV_const_ratio_const2 [OF _ der]) (simp add: neq) ultimately show ?thesis using neq by force next case greater have "(v b - v a) / (b - a) = k" by (rule DERIV_const_ratio_const2 [OF neq der]) then have "(b - a) * ((v b - v a) / (b - a)) = (b - a) * k" by simp moreover have " (v ((b + a) / 2) - v a) / ((b + a) / 2 - a) = k" by (rule DERIV_const_ratio_const2 [OF _ der]) (simp add: neq) ultimately show ?thesis using neq by (force simp add: add.commute) qed subsubsection\A function with positive derivative is increasing\ text \A simple proof using the MVT, by Jeremy Avigad. And variants.\ lemma DERIV_pos_imp_increasing_open: fixes a b :: real and f :: "real \ real" assumes "a < b" and "\x. a < x \ x < b \ (\y. DERIV f x :> y \ y > 0)" and con: "continuous_on {a..b} f" shows "f a < f b" proof (rule ccontr) assume f: "\ ?thesis" have "\l z. a < z \ z < b \ DERIV f z :> l \ f b - f a = (b - a) * l" by (rule MVT) (use assms real_differentiable_def in \force+\) then obtain l z where z: "a < z" "z < b" "DERIV f z :> l" and "f b - f a = (b - a) * l" by auto with assms f have "\ l > 0" by (metis linorder_not_le mult_le_0_iff diff_le_0_iff_le) with assms z show False by (metis DERIV_unique) qed lemma DERIV_pos_imp_increasing: fixes a b :: real and f :: "real \ real" assumes "a < b" and der: "\x. \a \ x; x \ b\ \ \y. DERIV f x :> y \ y > 0" shows "f a < f b" by (metis less_le_not_le DERIV_atLeastAtMost_imp_continuous_on DERIV_pos_imp_increasing_open [OF \a < b\] der) lemma DERIV_nonneg_imp_nondecreasing: fixes a b :: real and f :: "real \ real" assumes "a \ b" and "\x. \a \ x; x \ b\ \ \y. DERIV f x :> y \ y \ 0" shows "f a \ f b" proof (rule ccontr, cases "a = b") assume "\ ?thesis" and "a = b" then show False by auto next assume *: "\ ?thesis" assume "a \ b" with \a \ b\ have "a < b" by linarith moreover have "continuous_on {a..b} f" by (meson DERIV_isCont assms(2) atLeastAtMost_iff continuous_at_imp_continuous_on) ultimately have "\l z. a < z \ z < b \ DERIV f z :> l \ f b - f a = (b - a) * l" using assms MVT [OF \a < b\, of f] real_differentiable_def less_eq_real_def by blast then obtain l z where lz: "a < z" "z < b" "DERIV f z :> l" and **: "f b - f a = (b - a) * l" by auto with * have "a < b" "f b < f a" by auto with ** have "\ l \ 0" by (auto simp add: not_le algebra_simps) (metis * add_le_cancel_right assms(1) less_eq_real_def mult_right_mono add_left_mono linear order_refl) with assms lz show False by (metis DERIV_unique order_less_imp_le) qed lemma DERIV_neg_imp_decreasing_open: fixes a b :: real and f :: "real \ real" assumes "a < b" and "\x. a < x \ x < b \ \y. DERIV f x :> y \ y < 0" and con: "continuous_on {a..b} f" shows "f a > f b" proof - have "(\x. -f x) a < (\x. -f x) b" proof (rule DERIV_pos_imp_increasing_open [of a b]) show "\x. \a < x; x < b\ \ \y. ((\x. - f x) has_real_derivative y) (at x) \ 0 < y" using assms by simp (metis field_differentiable_minus neg_0_less_iff_less) show "continuous_on {a..b} (\x. - f x)" using con continuous_on_minus by blast qed (use assms in auto) then show ?thesis by simp qed lemma DERIV_neg_imp_decreasing: fixes a b :: real and f :: "real \ real" assumes "a < b" and der: "\x. \a \ x; x \ b\ \ \y. DERIV f x :> y \ y < 0" shows "f a > f b" by (metis less_le_not_le DERIV_atLeastAtMost_imp_continuous_on DERIV_neg_imp_decreasing_open [OF \a < b\] der) lemma DERIV_nonpos_imp_nonincreasing: fixes a b :: real and f :: "real \ real" assumes "a \ b" and "\x. \a \ x; x \ b\ \ \y. DERIV f x :> y \ y \ 0" shows "f a \ f b" proof - have "(\x. -f x) a \ (\x. -f x) b" using DERIV_nonneg_imp_nondecreasing [of a b "\x. -f x"] assms DERIV_minus by fastforce then show ?thesis by simp qed lemma DERIV_pos_imp_increasing_at_bot: fixes f :: "real \ real" assumes "\x. x \ b \ (\y. DERIV f x :> y \ y > 0)" and lim: "(f \ flim) at_bot" shows "flim < f b" proof - have "\N. \n\N. f n \ f (b - 1)" by (rule_tac x="b - 2" in exI) (force intro: order.strict_implies_order DERIV_pos_imp_increasing assms) then have "flim \ f (b - 1)" by (auto simp: eventually_at_bot_linorder tendsto_upperbound [OF lim]) also have "\ < f b" by (force intro: DERIV_pos_imp_increasing [where f=f] assms) finally show ?thesis . qed lemma DERIV_neg_imp_decreasing_at_top: fixes f :: "real \ real" assumes der: "\x. x \ b \ \y. DERIV f x :> y \ y < 0" and lim: "(f \ flim) at_top" shows "flim < f b" apply (rule DERIV_pos_imp_increasing_at_bot [where f = "\i. f (-i)" and b = "-b", simplified]) apply (metis DERIV_mirror der le_minus_iff neg_0_less_iff_less) apply (metis filterlim_at_top_mirror lim) done text \Derivative of inverse function\ lemma DERIV_inverse_function: fixes f g :: "real \ real" assumes der: "DERIV f (g x) :> D" and neq: "D \ 0" and x: "a < x" "x < b" and inj: "\y. \a < y; y < b\ \ f (g y) = y" and cont: "isCont g x" shows "DERIV g x :> inverse D" unfolding has_field_derivative_iff proof (rule LIM_equal2) show "0 < min (x - a) (b - x)" using x by arith next fix y assume "norm (y - x) < min (x - a) (b - x)" then have "a < y" and "y < b" by (simp_all add: abs_less_iff) then show "(g y - g x) / (y - x) = inverse ((f (g y) - x) / (g y - g x))" by (simp add: inj) next have "(\z. (f z - f (g x)) / (z - g x)) \g x\ D" by (rule der [unfolded has_field_derivative_iff]) then have 1: "(\z. (f z - x) / (z - g x)) \g x\ D" using inj x by simp have 2: "\d>0. \y. y \ x \ norm (y - x) < d \ g y \ g x" proof (rule exI, safe) show "0 < min (x - a) (b - x)" using x by simp next fix y assume "norm (y - x) < min (x - a) (b - x)" then have y: "a < y" "y < b" by (simp_all add: abs_less_iff) assume "g y = g x" then have "f (g y) = f (g x)" by simp then have "y = x" using inj y x by simp also assume "y \ x" finally show False by simp qed have "(\y. (f (g y) - x) / (g y - g x)) \x\ D" using cont 1 2 by (rule isCont_LIM_compose2) then show "(\y. inverse ((f (g y) - x) / (g y - g x))) \x\ inverse D" using neq by (rule tendsto_inverse) qed subsection \Generalized Mean Value Theorem\ theorem GMVT: fixes a b :: real assumes alb: "a < b" and fc: "\x. a \ x \ x \ b \ isCont f x" and fd: "\x. a < x \ x < b \ f differentiable (at x)" and gc: "\x. a \ x \ x \ b \ isCont g x" and gd: "\x. a < x \ x < b \ g differentiable (at x)" shows "\g'c f'c c. DERIV g c :> g'c \ DERIV f c :> f'c \ a < c \ c < b \ (f b - f a) * g'c = (g b - g a) * f'c" proof - let ?h = "\x. (f b - f a) * g x - (g b - g a) * f x" have "\l z. a < z \ z < b \ DERIV ?h z :> l \ ?h b - ?h a = (b - a) * l" proof (rule MVT) from assms show "a < b" by simp show "continuous_on {a..b} ?h" by (simp add: continuous_at_imp_continuous_on fc gc) show "\x. \a < x; x < b\ \ ?h differentiable (at x)" using fd gd by simp qed then obtain l where l: "\z. a < z \ z < b \ DERIV ?h z :> l \ ?h b - ?h a = (b - a) * l" .. then obtain c where c: "a < c \ c < b \ DERIV ?h c :> l \ ?h b - ?h a = (b - a) * l" .. from c have cint: "a < c \ c < b" by auto then obtain g'c where g'c: "DERIV g c :> g'c" using gd real_differentiable_def by blast from c have "a < c \ c < b" by auto then obtain f'c where f'c: "DERIV f c :> f'c" using fd real_differentiable_def by blast from c have "DERIV ?h c :> l" by auto moreover have "DERIV ?h c :> g'c * (f b - f a) - f'c * (g b - g a)" using g'c f'c by (auto intro!: derivative_eq_intros) ultimately have leq: "l = g'c * (f b - f a) - f'c * (g b - g a)" by (rule DERIV_unique) have "?h b - ?h a = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" proof - from c have "?h b - ?h a = (b - a) * l" by auto also from leq have "\ = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" by simp finally show ?thesis by simp qed moreover have "?h b - ?h a = 0" proof - have "?h b - ?h a = ((f b)*(g b) - (f a)*(g b) - (g b)*(f b) + (g a)*(f b)) - ((f b)*(g a) - (f a)*(g a) - (g b)*(f a) + (g a)*(f a))" by (simp add: algebra_simps) then show ?thesis by auto qed ultimately have "(b - a) * (g'c * (f b - f a) - f'c * (g b - g a)) = 0" by auto with alb have "g'c * (f b - f a) - f'c * (g b - g a) = 0" by simp then have "g'c * (f b - f a) = f'c * (g b - g a)" by simp then have "(f b - f a) * g'c = (g b - g a) * f'c" by (simp add: ac_simps) with g'c f'c cint show ?thesis by auto qed lemma GMVT': fixes f g :: "real \ real" assumes "a < b" and isCont_f: "\z. a \ z \ z \ b \ isCont f z" and isCont_g: "\z. a \ z \ z \ b \ isCont g z" and DERIV_g: "\z. a < z \ z < b \ DERIV g z :> (g' z)" and DERIV_f: "\z. a < z \ z < b \ DERIV f z :> (f' z)" shows "\c. a < c \ c < b \ (f b - f a) * g' c = (g b - g a) * f' c" proof - have "\g'c f'c c. DERIV g c :> g'c \ DERIV f c :> f'c \ a < c \ c < b \ (f b - f a) * g'c = (g b - g a) * f'c" using assms by (intro GMVT) (force simp: real_differentiable_def)+ then obtain c where "a < c" "c < b" "(f b - f a) * g' c = (g b - g a) * f' c" using DERIV_f DERIV_g by (force dest: DERIV_unique) then show ?thesis by auto qed subsection \L'Hopitals rule\ lemma isCont_If_ge: fixes a :: "'a :: linorder_topology" assumes "continuous (at_left a) g" and f: "(f \ g a) (at_right a)" shows "isCont (\x. if x \ a then g x else f x) a" (is "isCont ?gf a") proof - have g: "(g \ g a) (at_left a)" using assms continuous_within by blast show ?thesis unfolding isCont_def continuous_within proof (intro filterlim_split_at; simp) show "(?gf \ g a) (at_left a)" by (subst filterlim_cong[OF refl refl, where g=g]) (simp_all add: eventually_at_filter less_le g) show "(?gf \ g a) (at_right a)" by (subst filterlim_cong[OF refl refl, where g=f]) (simp_all add: eventually_at_filter less_le f) qed qed lemma lhopital_right_0: fixes f0 g0 :: "real \ real" assumes f_0: "(f0 \ 0) (at_right 0)" and g_0: "(g0 \ 0) (at_right 0)" and ev: "eventually (\x. g0 x \ 0) (at_right 0)" "eventually (\x. g' x \ 0) (at_right 0)" "eventually (\x. DERIV f0 x :> f' x) (at_right 0)" "eventually (\x. DERIV g0 x :> g' x) (at_right 0)" and lim: "filterlim (\ x. (f' x / g' x)) F (at_right 0)" shows "filterlim (\ x. f0 x / g0 x) F (at_right 0)" proof - define f where [abs_def]: "f x = (if x \ 0 then 0 else f0 x)" for x then have "f 0 = 0" by simp define g where [abs_def]: "g x = (if x \ 0 then 0 else g0 x)" for x then have "g 0 = 0" by simp have "eventually (\x. g0 x \ 0 \ g' x \ 0 \ DERIV f0 x :> (f' x) \ DERIV g0 x :> (g' x)) (at_right 0)" using ev by eventually_elim auto then obtain a where [arith]: "0 < a" and g0_neq_0: "\x. 0 < x \ x < a \ g0 x \ 0" and g'_neq_0: "\x. 0 < x \ x < a \ g' x \ 0" and f0: "\x. 0 < x \ x < a \ DERIV f0 x :> (f' x)" and g0: "\x. 0 < x \ x < a \ DERIV g0 x :> (g' x)" unfolding eventually_at by (auto simp: dist_real_def) have g_neq_0: "\x. 0 < x \ x < a \ g x \ 0" using g0_neq_0 by (simp add: g_def) have f: "DERIV f x :> (f' x)" if x: "0 < x" "x < a" for x using that by (intro DERIV_cong_ev[THEN iffD1, OF _ _ _ f0[OF x]]) (auto simp: f_def eventually_nhds_metric dist_real_def intro!: exI[of _ x]) have g: "DERIV g x :> (g' x)" if x: "0 < x" "x < a" for x using that by (intro DERIV_cong_ev[THEN iffD1, OF _ _ _ g0[OF x]]) (auto simp: g_def eventually_nhds_metric dist_real_def intro!: exI[of _ x]) have "isCont f 0" unfolding f_def by (intro isCont_If_ge f_0 continuous_const) have "isCont g 0" unfolding g_def by (intro isCont_If_ge g_0 continuous_const) have "\\. \x\{0 <..< a}. 0 < \ x \ \ x < x \ f x / g x = f' (\ x) / g' (\ x)" proof (rule bchoice, rule ballI) fix x assume "x \ {0 <..< a}" then have x[arith]: "0 < x" "x < a" by auto with g'_neq_0 g_neq_0 \g 0 = 0\ have g': "\x. 0 < x \ x < a \ 0 \ g' x" "g 0 \ g x" by auto have "\x. 0 \ x \ x < a \ isCont f x" using \isCont f 0\ f by (auto intro: DERIV_isCont simp: le_less) moreover have "\x. 0 \ x \ x < a \ isCont g x" using \isCont g 0\ g by (auto intro: DERIV_isCont simp: le_less) ultimately have "\c. 0 < c \ c < x \ (f x - f 0) * g' c = (g x - g 0) * f' c" using f g \x < a\ by (intro GMVT') auto then obtain c where *: "0 < c" "c < x" "(f x - f 0) * g' c = (g x - g 0) * f' c" by blast moreover from * g'(1)[of c] g'(2) have "(f x - f 0) / (g x - g 0) = f' c / g' c" by (simp add: field_simps) ultimately show "\y. 0 < y \ y < x \ f x / g x = f' y / g' y" using \f 0 = 0\ \g 0 = 0\ by (auto intro!: exI[of _ c]) qed then obtain \ where "\x\{0 <..< a}. 0 < \ x \ \ x < x \ f x / g x = f' (\ x) / g' (\ x)" .. then have \: "eventually (\x. 0 < \ x \ \ x < x \ f x / g x = f' (\ x) / g' (\ x)) (at_right 0)" unfolding eventually_at by (intro exI[of _ a]) (auto simp: dist_real_def) moreover from \ have "eventually (\x. norm (\ x) \ x) (at_right 0)" by eventually_elim auto then have "((\x. norm (\ x)) \ 0) (at_right 0)" by (rule_tac real_tendsto_sandwich[where f="\x. 0" and h="\x. x"]) auto then have "(\ \ 0) (at_right 0)" by (rule tendsto_norm_zero_cancel) with \ have "filterlim \ (at_right 0) (at_right 0)" by (auto elim!: eventually_mono simp: filterlim_at) from this lim have "filterlim (\t. f' (\ t) / g' (\ t)) F (at_right 0)" by (rule_tac filterlim_compose[of _ _ _ \]) ultimately have "filterlim (\t. f t / g t) F (at_right 0)" (is ?P) by (rule_tac filterlim_cong[THEN iffD1, OF refl refl]) (auto elim: eventually_mono) also have "?P \ ?thesis" by (rule filterlim_cong) (auto simp: f_def g_def eventually_at_filter) finally show ?thesis . qed lemma lhopital_right: "(f \ 0) (at_right x) \ (g \ 0) (at_right x) \ eventually (\x. g x \ 0) (at_right x) \ eventually (\x. g' x \ 0) (at_right x) \ eventually (\x. DERIV f x :> f' x) (at_right x) \ eventually (\x. DERIV g x :> g' x) (at_right x) \ filterlim (\ x. (f' x / g' x)) F (at_right x) \ filterlim (\ x. f x / g x) F (at_right x)" for x :: real unfolding eventually_at_right_to_0[of _ x] filterlim_at_right_to_0[of _ _ x] DERIV_shift by (rule lhopital_right_0) lemma lhopital_left: "(f \ 0) (at_left x) \ (g \ 0) (at_left x) \ eventually (\x. g x \ 0) (at_left x) \ eventually (\x. g' x \ 0) (at_left x) \ eventually (\x. DERIV f x :> f' x) (at_left x) \ eventually (\x. DERIV g x :> g' x) (at_left x) \ filterlim (\ x. (f' x / g' x)) F (at_left x) \ filterlim (\ x. f x / g x) F (at_left x)" for x :: real unfolding eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror by (rule lhopital_right[where f'="\x. - f' (- x)"]) (auto simp: DERIV_mirror) lemma lhopital: "(f \ 0) (at x) \ (g \ 0) (at x) \ eventually (\x. g x \ 0) (at x) \ eventually (\x. g' x \ 0) (at x) \ eventually (\x. DERIV f x :> f' x) (at x) \ eventually (\x. DERIV g x :> g' x) (at x) \ filterlim (\ x. (f' x / g' x)) F (at x) \ filterlim (\ x. f x / g x) F (at x)" for x :: real unfolding eventually_at_split filterlim_at_split by (auto intro!: lhopital_right[of f x g g' f'] lhopital_left[of f x g g' f']) lemma lhopital_right_0_at_top: fixes f g :: "real \ real" assumes g_0: "LIM x at_right 0. g x :> at_top" and ev: "eventually (\x. g' x \ 0) (at_right 0)" "eventually (\x. DERIV f x :> f' x) (at_right 0)" "eventually (\x. DERIV g x :> g' x) (at_right 0)" and lim: "((\ x. (f' x / g' x)) \ x) (at_right 0)" shows "((\ x. f x / g x) \ x) (at_right 0)" unfolding tendsto_iff proof safe fix e :: real assume "0 < e" with lim[unfolded tendsto_iff, rule_format, of "e / 4"] have "eventually (\t. dist (f' t / g' t) x < e / 4) (at_right 0)" by simp from eventually_conj[OF eventually_conj[OF ev(1) ev(2)] eventually_conj[OF ev(3) this]] obtain a where [arith]: "0 < a" and g'_neq_0: "\x. 0 < x \ x < a \ g' x \ 0" and f0: "\x. 0 < x \ x \ a \ DERIV f x :> (f' x)" and g0: "\x. 0 < x \ x \ a \ DERIV g x :> (g' x)" and Df: "\t. 0 < t \ t < a \ dist (f' t / g' t) x < e / 4" unfolding eventually_at_le by (auto simp: dist_real_def) from Df have "eventually (\t. t < a) (at_right 0)" "eventually (\t::real. 0 < t) (at_right 0)" unfolding eventually_at by (auto intro!: exI[of _ a] simp: dist_real_def) moreover have "eventually (\t. 0 < g t) (at_right 0)" "eventually (\t. g a < g t) (at_right 0)" using g_0 by (auto elim: eventually_mono simp: filterlim_at_top_dense) moreover have inv_g: "((\x. inverse (g x)) \ 0) (at_right 0)" using tendsto_inverse_0 filterlim_mono[OF g_0 at_top_le_at_infinity order_refl] by (rule filterlim_compose) then have "((\x. norm (1 - g a * inverse (g x))) \ norm (1 - g a * 0)) (at_right 0)" by (intro tendsto_intros) then have "((\x. norm (1 - g a / g x)) \ 1) (at_right 0)" by (simp add: inverse_eq_divide) from this[unfolded tendsto_iff, rule_format, of 1] have "eventually (\x. norm (1 - g a / g x) < 2) (at_right 0)" by (auto elim!: eventually_mono simp: dist_real_def) moreover from inv_g have "((\t. norm ((f a - x * g a) * inverse (g t))) \ norm ((f a - x * g a) * 0)) (at_right 0)" by (intro tendsto_intros) then have "((\t. norm (f a - x * g a) / norm (g t)) \ 0) (at_right 0)" by (simp add: inverse_eq_divide) from this[unfolded tendsto_iff, rule_format, of "e / 2"] \0 < e\ have "eventually (\t. norm (f a - x * g a) / norm (g t) < e / 2) (at_right 0)" by (auto simp: dist_real_def) ultimately show "eventually (\t. dist (f t / g t) x < e) (at_right 0)" proof eventually_elim fix t assume t[arith]: "0 < t" "t < a" "g a < g t" "0 < g t" assume ineq: "norm (1 - g a / g t) < 2" "norm (f a - x * g a) / norm (g t) < e / 2" have "\y. t < y \ y < a \ (g a - g t) * f' y = (f a - f t) * g' y" using f0 g0 t(1,2) by (intro GMVT') (force intro!: DERIV_isCont)+ then obtain y where [arith]: "t < y" "y < a" and D_eq0: "(g a - g t) * f' y = (f a - f t) * g' y" by blast from D_eq0 have D_eq: "(f t - f a) / (g t - g a) = f' y / g' y" using \g a < g t\ g'_neq_0[of y] by (auto simp add: field_simps) have *: "f t / g t - x = ((f t - f a) / (g t - g a) - x) * (1 - g a / g t) + (f a - x * g a) / g t" by (simp add: field_simps) have "norm (f t / g t - x) \ norm (((f t - f a) / (g t - g a) - x) * (1 - g a / g t)) + norm ((f a - x * g a) / g t)" unfolding * by (rule norm_triangle_ineq) also have "\ = dist (f' y / g' y) x * norm (1 - g a / g t) + norm (f a - x * g a) / norm (g t)" by (simp add: abs_mult D_eq dist_real_def) also have "\ < (e / 4) * 2 + e / 2" using ineq Df[of y] \0 < e\ by (intro add_le_less_mono mult_mono) auto finally show "dist (f t / g t) x < e" by (simp add: dist_real_def) qed qed lemma lhopital_right_at_top: "LIM x at_right x. (g::real \ real) x :> at_top \ eventually (\x. g' x \ 0) (at_right x) \ eventually (\x. DERIV f x :> f' x) (at_right x) \ eventually (\x. DERIV g x :> g' x) (at_right x) \ ((\ x. (f' x / g' x)) \ y) (at_right x) \ ((\ x. f x / g x) \ y) (at_right x)" unfolding eventually_at_right_to_0[of _ x] filterlim_at_right_to_0[of _ _ x] DERIV_shift by (rule lhopital_right_0_at_top) lemma lhopital_left_at_top: "LIM x at_left x. g x :> at_top \ eventually (\x. g' x \ 0) (at_left x) \ eventually (\x. DERIV f x :> f' x) (at_left x) \ eventually (\x. DERIV g x :> g' x) (at_left x) \ ((\ x. (f' x / g' x)) \ y) (at_left x) \ ((\ x. f x / g x) \ y) (at_left x)" for x :: real unfolding eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror by (rule lhopital_right_at_top[where f'="\x. - f' (- x)"]) (auto simp: DERIV_mirror) lemma lhopital_at_top: "LIM x at x. (g::real \ real) x :> at_top \ eventually (\x. g' x \ 0) (at x) \ eventually (\x. DERIV f x :> f' x) (at x) \ eventually (\x. DERIV g x :> g' x) (at x) \ ((\ x. (f' x / g' x)) \ y) (at x) \ ((\ x. f x / g x) \ y) (at x)" unfolding eventually_at_split filterlim_at_split by (auto intro!: lhopital_right_at_top[of g x g' f f'] lhopital_left_at_top[of g x g' f f']) lemma lhospital_at_top_at_top: fixes f g :: "real \ real" assumes g_0: "LIM x at_top. g x :> at_top" and g': "eventually (\x. g' x \ 0) at_top" and Df: "eventually (\x. DERIV f x :> f' x) at_top" and Dg: "eventually (\x. DERIV g x :> g' x) at_top" and lim: "((\ x. (f' x / g' x)) \ x) at_top" shows "((\ x. f x / g x) \ x) at_top" unfolding filterlim_at_top_to_right proof (rule lhopital_right_0_at_top) let ?F = "\x. f (inverse x)" let ?G = "\x. g (inverse x)" let ?R = "at_right (0::real)" let ?D = "\f' x. f' (inverse x) * - (inverse x ^ Suc (Suc 0))" show "LIM x ?R. ?G x :> at_top" using g_0 unfolding filterlim_at_top_to_right . show "eventually (\x. DERIV ?G x :> ?D g' x) ?R" unfolding eventually_at_right_to_top using Dg eventually_ge_at_top[where c=1] by eventually_elim (rule derivative_eq_intros DERIV_chain'[where f=inverse] | simp)+ show "eventually (\x. DERIV ?F x :> ?D f' x) ?R" unfolding eventually_at_right_to_top using Df eventually_ge_at_top[where c=1] by eventually_elim (rule derivative_eq_intros DERIV_chain'[where f=inverse] | simp)+ show "eventually (\x. ?D g' x \ 0) ?R" unfolding eventually_at_right_to_top using g' eventually_ge_at_top[where c=1] by eventually_elim auto show "((\x. ?D f' x / ?D g' x) \ x) ?R" unfolding filterlim_at_right_to_top apply (intro filterlim_cong[THEN iffD2, OF refl refl _ lim]) using eventually_ge_at_top[where c=1] by eventually_elim simp qed lemma lhopital_right_at_top_at_top: fixes f g :: "real \ real" assumes f_0: "LIM x at_right a. f x :> at_top" assumes g_0: "LIM x at_right a. g x :> at_top" and ev: "eventually (\x. DERIV f x :> f' x) (at_right a)" "eventually (\x. DERIV g x :> g' x) (at_right a)" and lim: "filterlim (\ x. (f' x / g' x)) at_top (at_right a)" shows "filterlim (\ x. f x / g x) at_top (at_right a)" proof - from lim have pos: "eventually (\x. f' x / g' x > 0) (at_right a)" unfolding filterlim_at_top_dense by blast have "((\x. g x / f x) \ 0) (at_right a)" proof (rule lhopital_right_at_top) from pos show "eventually (\x. f' x \ 0) (at_right a)" by eventually_elim auto from tendsto_inverse_0_at_top[OF lim] show "((\x. g' x / f' x) \ 0) (at_right a)" by simp qed fact+ moreover from f_0 g_0 have "eventually (\x. f x > 0) (at_right a)" "eventually (\x. g x > 0) (at_right a)" unfolding filterlim_at_top_dense by blast+ hence "eventually (\x. g x / f x > 0) (at_right a)" by eventually_elim simp ultimately have "filterlim (\x. inverse (g x / f x)) at_top (at_right a)" by (rule filterlim_inverse_at_top) thus ?thesis by simp qed lemma lhopital_right_at_top_at_bot: fixes f g :: "real \ real" assumes f_0: "LIM x at_right a. f x :> at_top" assumes g_0: "LIM x at_right a. g x :> at_bot" and ev: "eventually (\x. DERIV f x :> f' x) (at_right a)" "eventually (\x. DERIV g x :> g' x) (at_right a)" and lim: "filterlim (\ x. (f' x / g' x)) at_bot (at_right a)" shows "filterlim (\ x. f x / g x) at_bot (at_right a)" proof - from ev(2) have ev': "eventually (\x. DERIV (\x. -g x) x :> -g' x) (at_right a)" by eventually_elim (auto intro: derivative_intros) have "filterlim (\x. f x / (-g x)) at_top (at_right a)" by (rule lhopital_right_at_top_at_top[where f' = f' and g' = "\x. -g' x"]) (insert assms ev', auto simp: filterlim_uminus_at_bot) hence "filterlim (\x. -(f x / g x)) at_top (at_right a)" by simp thus ?thesis by (simp add: filterlim_uminus_at_bot) qed lemma lhopital_left_at_top_at_top: fixes f g :: "real \ real" assumes f_0: "LIM x at_left a. f x :> at_top" assumes g_0: "LIM x at_left a. g x :> at_top" and ev: "eventually (\x. DERIV f x :> f' x) (at_left a)" "eventually (\x. DERIV g x :> g' x) (at_left a)" and lim: "filterlim (\ x. (f' x / g' x)) at_top (at_left a)" shows "filterlim (\ x. f x / g x) at_top (at_left a)" by (insert assms, unfold eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror, rule lhopital_right_at_top_at_top[where f'="\x. - f' (- x)"]) (insert assms, auto simp: DERIV_mirror) lemma lhopital_left_at_top_at_bot: fixes f g :: "real \ real" assumes f_0: "LIM x at_left a. f x :> at_top" assumes g_0: "LIM x at_left a. g x :> at_bot" and ev: "eventually (\x. DERIV f x :> f' x) (at_left a)" "eventually (\x. DERIV g x :> g' x) (at_left a)" and lim: "filterlim (\ x. (f' x / g' x)) at_bot (at_left a)" shows "filterlim (\ x. f x / g x) at_bot (at_left a)" by (insert assms, unfold eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror, rule lhopital_right_at_top_at_bot[where f'="\x. - f' (- x)"]) (insert assms, auto simp: DERIV_mirror) lemma lhopital_at_top_at_top: fixes f g :: "real \ real" assumes f_0: "LIM x at a. f x :> at_top" assumes g_0: "LIM x at a. g x :> at_top" and ev: "eventually (\x. DERIV f x :> f' x) (at a)" "eventually (\x. DERIV g x :> g' x) (at a)" and lim: "filterlim (\ x. (f' x / g' x)) at_top (at a)" shows "filterlim (\ x. f x / g x) at_top (at a)" using assms unfolding eventually_at_split filterlim_at_split by (auto intro!: lhopital_right_at_top_at_top[of f a g f' g'] lhopital_left_at_top_at_top[of f a g f' g']) lemma lhopital_at_top_at_bot: fixes f g :: "real \ real" assumes f_0: "LIM x at a. f x :> at_top" assumes g_0: "LIM x at a. g x :> at_bot" and ev: "eventually (\x. DERIV f x :> f' x) (at a)" "eventually (\x. DERIV g x :> g' x) (at a)" and lim: "filterlim (\ x. (f' x / g' x)) at_bot (at a)" shows "filterlim (\ x. f x / g x) at_bot (at a)" using assms unfolding eventually_at_split filterlim_at_split by (auto intro!: lhopital_right_at_top_at_bot[of f a g f' g'] lhopital_left_at_top_at_bot[of f a g f' g']) end diff --git a/src/HOL/Groups.thy b/src/HOL/Groups.thy --- a/src/HOL/Groups.thy +++ b/src/HOL/Groups.thy @@ -1,1495 +1,1500 @@ (* Title: HOL/Groups.thy Author: Gertrud Bauer Author: Steven Obua Author: Lawrence C Paulson Author: Markus Wenzel Author: Jeremy Avigad *) section \Groups, also combined with orderings\ theory Groups imports Orderings begin subsection \Dynamic facts\ named_theorems ac_simps "associativity and commutativity simplification rules" and algebra_simps "algebra simplification rules for rings" and algebra_split_simps "algebra simplification rules for rings, with potential goal splitting" and field_simps "algebra simplification rules for fields" and field_split_simps "algebra simplification rules for fields, with potential goal splitting" text \ The rewrites accumulated in \algebra_simps\ deal with the classical algebraic structures of groups, rings and family. They simplify terms by multiplying everything out (in case of a ring) and bringing sums and products into a canonical form (by ordered rewriting). As a result it decides group and ring equalities but also helps with inequalities. Of course it also works for fields, but it knows nothing about multiplicative inverses or division. This is catered for by \field_simps\. Facts in \field_simps\ multiply with denominators in (in)equations if they can be proved to be non-zero (for equations) or positive/negative (for inequalities). Can be too aggressive and is therefore separate from the more benign \algebra_simps\. Collections \algebra_split_simps\ and \field_split_simps\ correspond to \algebra_simps\ and \field_simps\ but contain more aggresive rules that may lead to goal splitting. \ subsection \Abstract structures\ text \ These locales provide basic structures for interpretation into bigger structures; extensions require careful thinking, otherwise undesired effects may occur due to interpretation. \ locale semigroup = fixes f :: "'a \ 'a \ 'a" (infixl "\<^bold>*" 70) assumes assoc [ac_simps]: "a \<^bold>* b \<^bold>* c = a \<^bold>* (b \<^bold>* c)" locale abel_semigroup = semigroup + assumes commute [ac_simps]: "a \<^bold>* b = b \<^bold>* a" begin lemma left_commute [ac_simps]: "b \<^bold>* (a \<^bold>* c) = a \<^bold>* (b \<^bold>* c)" proof - have "(b \<^bold>* a) \<^bold>* c = (a \<^bold>* b) \<^bold>* c" by (simp only: commute) then show ?thesis by (simp only: assoc) qed end locale monoid = semigroup + fixes z :: 'a ("\<^bold>1") assumes left_neutral [simp]: "\<^bold>1 \<^bold>* a = a" assumes right_neutral [simp]: "a \<^bold>* \<^bold>1 = a" locale comm_monoid = abel_semigroup + fixes z :: 'a ("\<^bold>1") assumes comm_neutral: "a \<^bold>* \<^bold>1 = a" begin sublocale monoid by standard (simp_all add: commute comm_neutral) end locale group = semigroup + fixes z :: 'a ("\<^bold>1") fixes inverse :: "'a \ 'a" assumes group_left_neutral: "\<^bold>1 \<^bold>* a = a" assumes left_inverse [simp]: "inverse a \<^bold>* a = \<^bold>1" begin lemma left_cancel: "a \<^bold>* b = a \<^bold>* c \ b = c" proof assume "a \<^bold>* b = a \<^bold>* c" then have "inverse a \<^bold>* (a \<^bold>* b) = inverse a \<^bold>* (a \<^bold>* c)" by simp then have "(inverse a \<^bold>* a) \<^bold>* b = (inverse a \<^bold>* a) \<^bold>* c" by (simp only: assoc) then show "b = c" by (simp add: group_left_neutral) qed simp sublocale monoid proof fix a have "inverse a \<^bold>* a = \<^bold>1" by simp then have "inverse a \<^bold>* (a \<^bold>* \<^bold>1) = inverse a \<^bold>* a" by (simp add: group_left_neutral assoc [symmetric]) with left_cancel show "a \<^bold>* \<^bold>1 = a" by (simp only: left_cancel) qed (fact group_left_neutral) lemma inverse_unique: assumes "a \<^bold>* b = \<^bold>1" shows "inverse a = b" proof - from assms have "inverse a \<^bold>* (a \<^bold>* b) = inverse a" by simp then show ?thesis by (simp add: assoc [symmetric]) qed lemma inverse_neutral [simp]: "inverse \<^bold>1 = \<^bold>1" by (rule inverse_unique) simp lemma inverse_inverse [simp]: "inverse (inverse a) = a" by (rule inverse_unique) simp lemma right_inverse [simp]: "a \<^bold>* inverse a = \<^bold>1" proof - have "a \<^bold>* inverse a = inverse (inverse a) \<^bold>* inverse a" by simp also have "\ = \<^bold>1" by (rule left_inverse) then show ?thesis by simp qed lemma inverse_distrib_swap: "inverse (a \<^bold>* b) = inverse b \<^bold>* inverse a" proof (rule inverse_unique) have "a \<^bold>* b \<^bold>* (inverse b \<^bold>* inverse a) = a \<^bold>* (b \<^bold>* inverse b) \<^bold>* inverse a" by (simp only: assoc) also have "\ = \<^bold>1" by simp finally show "a \<^bold>* b \<^bold>* (inverse b \<^bold>* inverse a) = \<^bold>1" . qed lemma right_cancel: "b \<^bold>* a = c \<^bold>* a \ b = c" proof assume "b \<^bold>* a = c \<^bold>* a" then have "b \<^bold>* a \<^bold>* inverse a= c \<^bold>* a \<^bold>* inverse a" by simp then show "b = c" by (simp add: assoc) qed simp end subsection \Generic operations\ class zero = fixes zero :: 'a ("0") class one = fixes one :: 'a ("1") hide_const (open) zero one lemma Let_0 [simp]: "Let 0 f = f 0" unfolding Let_def .. lemma Let_1 [simp]: "Let 1 f = f 1" unfolding Let_def .. setup \ Reorient_Proc.add (fn Const(\<^const_name>\Groups.zero\, _) => true | Const(\<^const_name>\Groups.one\, _) => true | _ => false) \ simproc_setup reorient_zero ("0 = x") = Reorient_Proc.proc simproc_setup reorient_one ("1 = x") = Reorient_Proc.proc typed_print_translation \ let fun tr' c = (c, fn ctxt => fn T => fn ts => if null ts andalso Printer.type_emphasis ctxt T then Syntax.const \<^syntax_const>\_constrain\ $ Syntax.const c $ Syntax_Phases.term_of_typ ctxt T else raise Match); in map tr' [\<^const_syntax>\Groups.one\, \<^const_syntax>\Groups.zero\] end \ \ \show types that are presumably too general\ class plus = fixes plus :: "'a \ 'a \ 'a" (infixl "+" 65) class minus = fixes minus :: "'a \ 'a \ 'a" (infixl "-" 65) class uminus = fixes uminus :: "'a \ 'a" ("- _" [81] 80) class times = fixes times :: "'a \ 'a \ 'a" (infixl "*" 70) subsection \Semigroups and Monoids\ class semigroup_add = plus + assumes add_assoc [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: "(a + b) + c = a + (b + c)" begin sublocale add: semigroup plus by standard (fact add_assoc) end hide_fact add_assoc class ab_semigroup_add = semigroup_add + assumes add_commute [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: "a + b = b + a" begin sublocale add: abel_semigroup plus by standard (fact add_commute) declare add.left_commute [algebra_simps, algebra_split_simps, field_simps, field_split_simps] lemmas add_ac = add.assoc add.commute add.left_commute end hide_fact add_commute lemmas add_ac = add.assoc add.commute add.left_commute class semigroup_mult = times + assumes mult_assoc [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: "(a * b) * c = a * (b * c)" begin sublocale mult: semigroup times by standard (fact mult_assoc) end hide_fact mult_assoc class ab_semigroup_mult = semigroup_mult + assumes mult_commute [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: "a * b = b * a" begin sublocale mult: abel_semigroup times by standard (fact mult_commute) declare mult.left_commute [algebra_simps, algebra_split_simps, field_simps, field_split_simps] lemmas mult_ac = mult.assoc mult.commute mult.left_commute end hide_fact mult_commute lemmas mult_ac = mult.assoc mult.commute mult.left_commute class monoid_add = zero + semigroup_add + assumes add_0_left: "0 + a = a" and add_0_right: "a + 0 = a" begin sublocale add: monoid plus 0 by standard (fact add_0_left add_0_right)+ end lemma zero_reorient: "0 = x \ x = 0" by (fact eq_commute) class comm_monoid_add = zero + ab_semigroup_add + assumes add_0: "0 + a = a" begin subclass monoid_add by standard (simp_all add: add_0 add.commute [of _ 0]) sublocale add: comm_monoid plus 0 by standard (simp add: ac_simps) end class monoid_mult = one + semigroup_mult + assumes mult_1_left: "1 * a = a" and mult_1_right: "a * 1 = a" begin sublocale mult: monoid times 1 by standard (fact mult_1_left mult_1_right)+ end lemma one_reorient: "1 = x \ x = 1" by (fact eq_commute) class comm_monoid_mult = one + ab_semigroup_mult + assumes mult_1: "1 * a = a" begin subclass monoid_mult by standard (simp_all add: mult_1 mult.commute [of _ 1]) sublocale mult: comm_monoid times 1 by standard (simp add: ac_simps) end class cancel_semigroup_add = semigroup_add + assumes add_left_imp_eq: "a + b = a + c \ b = c" assumes add_right_imp_eq: "b + a = c + a \ b = c" begin lemma add_left_cancel [simp]: "a + b = a + c \ b = c" by (blast dest: add_left_imp_eq) lemma add_right_cancel [simp]: "b + a = c + a \ b = c" by (blast dest: add_right_imp_eq) end class cancel_ab_semigroup_add = ab_semigroup_add + minus + assumes add_diff_cancel_left' [simp]: "(a + b) - a = b" assumes diff_diff_add [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: "a - b - c = a - (b + c)" begin lemma add_diff_cancel_right' [simp]: "(a + b) - b = a" using add_diff_cancel_left' [of b a] by (simp add: ac_simps) subclass cancel_semigroup_add proof fix a b c :: 'a assume "a + b = a + c" then have "a + b - a = a + c - a" by simp then show "b = c" by simp next fix a b c :: 'a assume "b + a = c + a" then have "b + a - a = c + a - a" by simp then show "b = c" by simp qed lemma add_diff_cancel_left [simp]: "(c + a) - (c + b) = a - b" unfolding diff_diff_add [symmetric] by simp lemma add_diff_cancel_right [simp]: "(a + c) - (b + c) = a - b" using add_diff_cancel_left [symmetric] by (simp add: ac_simps) lemma diff_right_commute: "a - c - b = a - b - c" by (simp add: diff_diff_add add.commute) end class cancel_comm_monoid_add = cancel_ab_semigroup_add + comm_monoid_add begin lemma diff_zero [simp]: "a - 0 = a" using add_diff_cancel_right' [of a 0] by simp lemma diff_cancel [simp]: "a - a = 0" proof - have "(a + 0) - (a + 0) = 0" by (simp only: add_diff_cancel_left diff_zero) then show ?thesis by simp qed lemma add_implies_diff: assumes "c + b = a" shows "c = a - b" proof - from assms have "(b + c) - (b + 0) = a - b" by (simp add: add.commute) then show "c = a - b" by simp qed lemma add_cancel_right_right [simp]: "a = a + b \ b = 0" (is "?P \ ?Q") proof assume ?Q then show ?P by simp next assume ?P then have "a - a = a + b - a" by simp then show ?Q by simp qed lemma add_cancel_right_left [simp]: "a = b + a \ b = 0" using add_cancel_right_right [of a b] by (simp add: ac_simps) lemma add_cancel_left_right [simp]: "a + b = a \ b = 0" by (auto dest: sym) lemma add_cancel_left_left [simp]: "b + a = a \ b = 0" by (auto dest: sym) end class comm_monoid_diff = cancel_comm_monoid_add + assumes zero_diff [simp]: "0 - a = 0" begin lemma diff_add_zero [simp]: "a - (a + b) = 0" proof - have "a - (a + b) = (a + 0) - (a + b)" by simp also have "\ = 0" by (simp only: add_diff_cancel_left zero_diff) finally show ?thesis . qed end subsection \Groups\ class group_add = minus + uminus + monoid_add + assumes left_minus: "- a + a = 0" assumes add_uminus_conv_diff [simp]: "a + (- b) = a - b" begin lemma diff_conv_add_uminus: "a - b = a + (- b)" by simp sublocale add: group plus 0 uminus by standard (simp_all add: left_minus) lemma minus_unique: "a + b = 0 \ - a = b" by (fact add.inverse_unique) lemma minus_zero: "- 0 = 0" by (fact add.inverse_neutral) lemma minus_minus: "- (- a) = a" by (fact add.inverse_inverse) lemma right_minus: "a + - a = 0" by (fact add.right_inverse) lemma diff_self [simp]: "a - a = 0" using right_minus [of a] by simp subclass cancel_semigroup_add by standard (simp_all add: add.left_cancel add.right_cancel) lemma minus_add_cancel [simp]: "- a + (a + b) = b" by (simp add: add.assoc [symmetric]) lemma add_minus_cancel [simp]: "a + (- a + b) = b" by (simp add: add.assoc [symmetric]) lemma diff_add_cancel [simp]: "a - b + b = a" by (simp only: diff_conv_add_uminus add.assoc) simp lemma add_diff_cancel [simp]: "a + b - b = a" by (simp only: diff_conv_add_uminus add.assoc) simp lemma minus_add: "- (a + b) = - b + - a" by (fact add.inverse_distrib_swap) lemma right_minus_eq [simp]: "a - b = 0 \ a = b" proof assume "a - b = 0" have "a = (a - b) + b" by (simp add: add.assoc) also have "\ = b" using \a - b = 0\ by simp finally show "a = b" . next assume "a = b" then show "a - b = 0" by simp qed lemma eq_iff_diff_eq_0: "a = b \ a - b = 0" by (fact right_minus_eq [symmetric]) lemma diff_0 [simp]: "0 - a = - a" by (simp only: diff_conv_add_uminus add_0_left) lemma diff_0_right [simp]: "a - 0 = a" by (simp only: diff_conv_add_uminus minus_zero add_0_right) lemma diff_minus_eq_add [simp]: "a - - b = a + b" by (simp only: diff_conv_add_uminus minus_minus) lemma neg_equal_iff_equal [simp]: "- a = - b \ a = b" proof assume "- a = - b" then have "- (- a) = - (- b)" by simp then show "a = b" by simp next assume "a = b" then show "- a = - b" by simp qed lemma neg_equal_0_iff_equal [simp]: "- a = 0 \ a = 0" by (subst neg_equal_iff_equal [symmetric]) simp lemma neg_0_equal_iff_equal [simp]: "0 = - a \ 0 = a" by (subst neg_equal_iff_equal [symmetric]) simp text \The next two equations can make the simplifier loop!\ lemma equation_minus_iff: "a = - b \ b = - a" proof - have "- (- a) = - b \ - a = b" by (rule neg_equal_iff_equal) then show ?thesis by (simp add: eq_commute) qed lemma minus_equation_iff: "- a = b \ - b = a" proof - have "- a = - (- b) \ a = -b" by (rule neg_equal_iff_equal) then show ?thesis by (simp add: eq_commute) qed lemma eq_neg_iff_add_eq_0: "a = - b \ a + b = 0" proof assume "a = - b" then show "a + b = 0" by simp next assume "a + b = 0" moreover have "a + (b + - b) = (a + b) + - b" by (simp only: add.assoc) ultimately show "a = - b" by simp qed lemma add_eq_0_iff2: "a + b = 0 \ a = - b" by (fact eq_neg_iff_add_eq_0 [symmetric]) lemma neg_eq_iff_add_eq_0: "- a = b \ a + b = 0" by (auto simp add: add_eq_0_iff2) lemma add_eq_0_iff: "a + b = 0 \ b = - a" by (auto simp add: neg_eq_iff_add_eq_0 [symmetric]) lemma minus_diff_eq [simp]: "- (a - b) = b - a" by (simp only: neg_eq_iff_add_eq_0 diff_conv_add_uminus add.assoc minus_add_cancel) simp lemma add_diff_eq [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: "a + (b - c) = (a + b) - c" by (simp only: diff_conv_add_uminus add.assoc) lemma diff_add_eq_diff_diff_swap: "a - (b + c) = a - c - b" by (simp only: diff_conv_add_uminus add.assoc minus_add) lemma diff_eq_eq [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: "a - b = c \ a = c + b" by auto lemma eq_diff_eq [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: "a = c - b \ a + b = c" by auto lemma diff_diff_eq2 [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: "a - (b - c) = (a + c) - b" by (simp only: diff_conv_add_uminus add.assoc) simp lemma diff_eq_diff_eq: "a - b = c - d \ a = b \ c = d" by (simp only: eq_iff_diff_eq_0 [of a b] eq_iff_diff_eq_0 [of c d]) end class ab_group_add = minus + uminus + comm_monoid_add + assumes ab_left_minus: "- a + a = 0" assumes ab_diff_conv_add_uminus: "a - b = a + (- b)" begin subclass group_add by standard (simp_all add: ab_left_minus ab_diff_conv_add_uminus) subclass cancel_comm_monoid_add proof fix a b c :: 'a have "b + a - a = b" by simp then show "a + b - a = b" by (simp add: ac_simps) show "a - b - c = a - (b + c)" by (simp add: algebra_simps) qed lemma uminus_add_conv_diff [simp]: "- a + b = b - a" by (simp add: add.commute) lemma minus_add_distrib [simp]: "- (a + b) = - a + - b" by (simp add: algebra_simps) lemma diff_add_eq [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: "(a - b) + c = (a + c) - b" by (simp add: algebra_simps) lemma minus_diff_commute: "- b - a = - a - b" by (simp only: diff_conv_add_uminus add.commute) end subsection \(Partially) Ordered Groups\ text \ The theory of partially ordered groups is taken from the books: \<^item> \<^emph>\Lattice Theory\ by Garret Birkhoff, American Mathematical Society, 1979 \<^item> \<^emph>\Partially Ordered Algebraic Systems\, Pergamon Press, 1963 Most of the used notions can also be looked up in \<^item> \<^url>\http://www.mathworld.com\ by Eric Weisstein et. al. \<^item> \<^emph>\Algebra I\ by van der Waerden, Springer \ class ordered_ab_semigroup_add = order + ab_semigroup_add + assumes add_left_mono: "a \ b \ c + a \ c + b" begin lemma add_right_mono: "a \ b \ a + c \ b + c" by (simp add: add.commute [of _ c] add_left_mono) text \non-strict, in both arguments\ lemma add_mono: "a \ b \ c \ d \ a + c \ b + d" by (simp add: add.commute add_left_mono add_right_mono [THEN order_trans]) end +lemma mono_add: + fixes a :: "'a::ordered_ab_semigroup_add" + shows "mono ((+) a)" + by (simp add: add_left_mono monoI) + text \Strict monotonicity in both arguments\ class strict_ordered_ab_semigroup_add = ordered_ab_semigroup_add + assumes add_strict_mono: "a < b \ c < d \ a + c < b + d" class ordered_cancel_ab_semigroup_add = ordered_ab_semigroup_add + cancel_ab_semigroup_add begin lemma add_strict_left_mono: "a < b \ c + a < c + b" by (auto simp add: less_le add_left_mono) lemma add_strict_right_mono: "a < b \ a + c < b + c" by (simp add: add.commute [of _ c] add_strict_left_mono) subclass strict_ordered_ab_semigroup_add proof show "\a b c d. \a < b; c < d\ \ a + c < b + d" by (iprover intro: add_strict_left_mono add_strict_right_mono less_trans) qed lemma add_less_le_mono: "a < b \ c \ d \ a + c < b + d" by (iprover intro: add_left_mono add_strict_right_mono less_le_trans) lemma add_le_less_mono: "a \ b \ c < d \ a + c < b + d" by (iprover intro: add_strict_left_mono add_right_mono less_le_trans) end class ordered_ab_semigroup_add_imp_le = ordered_cancel_ab_semigroup_add + assumes add_le_imp_le_left: "c + a \ c + b \ a \ b" begin lemma add_less_imp_less_left: assumes less: "c + a < c + b" shows "a < b" proof - from less have le: "c + a \ c + b" by (simp add: order_le_less) have "a \ b" using add_le_imp_le_left [OF le] . moreover have "a \ b" proof (rule ccontr) assume "\ ?thesis" then have "a = b" by simp then have "c + a = c + b" by simp with less show "False" by simp qed ultimately show "a < b" by (simp add: order_le_less) qed lemma add_less_imp_less_right: "a + c < b + c \ a < b" by (rule add_less_imp_less_left [of c]) (simp add: add.commute) lemma add_less_cancel_left [simp]: "c + a < c + b \ a < b" by (blast intro: add_less_imp_less_left add_strict_left_mono) lemma add_less_cancel_right [simp]: "a + c < b + c \ a < b" by (blast intro: add_less_imp_less_right add_strict_right_mono) lemma add_le_cancel_left [simp]: "c + a \ c + b \ a \ b" by (auto simp: dest: add_le_imp_le_left add_left_mono) lemma add_le_cancel_right [simp]: "a + c \ b + c \ a \ b" by (simp add: add.commute [of a c] add.commute [of b c]) lemma add_le_imp_le_right: "a + c \ b + c \ a \ b" by simp lemma max_add_distrib_left: "max x y + z = max (x + z) (y + z)" unfolding max_def by auto lemma min_add_distrib_left: "min x y + z = min (x + z) (y + z)" unfolding min_def by auto lemma max_add_distrib_right: "x + max y z = max (x + y) (x + z)" unfolding max_def by auto lemma min_add_distrib_right: "x + min y z = min (x + y) (x + z)" unfolding min_def by auto end subsection \Support for reasoning about signs\ class ordered_comm_monoid_add = comm_monoid_add + ordered_ab_semigroup_add begin lemma add_nonneg_nonneg [simp]: "0 \ a \ 0 \ b \ 0 \ a + b" using add_mono[of 0 a 0 b] by simp lemma add_nonpos_nonpos: "a \ 0 \ b \ 0 \ a + b \ 0" using add_mono[of a 0 b 0] by simp lemma add_nonneg_eq_0_iff: "0 \ x \ 0 \ y \ x + y = 0 \ x = 0 \ y = 0" using add_left_mono[of 0 y x] add_right_mono[of 0 x y] by auto lemma add_nonpos_eq_0_iff: "x \ 0 \ y \ 0 \ x + y = 0 \ x = 0 \ y = 0" using add_left_mono[of y 0 x] add_right_mono[of x 0 y] by auto lemma add_increasing: "0 \ a \ b \ c \ b \ a + c" using add_mono [of 0 a b c] by simp lemma add_increasing2: "0 \ c \ b \ a \ b \ a + c" by (simp add: add_increasing add.commute [of a]) lemma add_decreasing: "a \ 0 \ c \ b \ a + c \ b" using add_mono [of a 0 c b] by simp lemma add_decreasing2: "c \ 0 \ a \ b \ a + c \ b" using add_mono[of a b c 0] by simp lemma add_pos_nonneg: "0 < a \ 0 \ b \ 0 < a + b" using less_le_trans[of 0 a "a + b"] by (simp add: add_increasing2) lemma add_pos_pos: "0 < a \ 0 < b \ 0 < a + b" by (intro add_pos_nonneg less_imp_le) lemma add_nonneg_pos: "0 \ a \ 0 < b \ 0 < a + b" using add_pos_nonneg[of b a] by (simp add: add_commute) lemma add_neg_nonpos: "a < 0 \ b \ 0 \ a + b < 0" using le_less_trans[of "a + b" a 0] by (simp add: add_decreasing2) lemma add_neg_neg: "a < 0 \ b < 0 \ a + b < 0" by (intro add_neg_nonpos less_imp_le) lemma add_nonpos_neg: "a \ 0 \ b < 0 \ a + b < 0" using add_neg_nonpos[of b a] by (simp add: add_commute) lemmas add_sign_intros = add_pos_nonneg add_pos_pos add_nonneg_pos add_nonneg_nonneg add_neg_nonpos add_neg_neg add_nonpos_neg add_nonpos_nonpos end class strict_ordered_comm_monoid_add = comm_monoid_add + strict_ordered_ab_semigroup_add begin lemma pos_add_strict: "0 < a \ b < c \ b < a + c" using add_strict_mono [of 0 a b c] by simp end class ordered_cancel_comm_monoid_add = ordered_comm_monoid_add + cancel_ab_semigroup_add begin subclass ordered_cancel_ab_semigroup_add .. subclass strict_ordered_comm_monoid_add .. lemma add_strict_increasing: "0 < a \ b \ c \ b < a + c" using add_less_le_mono [of 0 a b c] by simp lemma add_strict_increasing2: "0 \ a \ b < c \ b < a + c" using add_le_less_mono [of 0 a b c] by simp end class ordered_ab_semigroup_monoid_add_imp_le = monoid_add + ordered_ab_semigroup_add_imp_le begin lemma add_less_same_cancel1 [simp]: "b + a < b \ a < 0" using add_less_cancel_left [of _ _ 0] by simp lemma add_less_same_cancel2 [simp]: "a + b < b \ a < 0" using add_less_cancel_right [of _ _ 0] by simp lemma less_add_same_cancel1 [simp]: "a < a + b \ 0 < b" using add_less_cancel_left [of _ 0] by simp lemma less_add_same_cancel2 [simp]: "a < b + a \ 0 < b" using add_less_cancel_right [of 0] by simp lemma add_le_same_cancel1 [simp]: "b + a \ b \ a \ 0" using add_le_cancel_left [of _ _ 0] by simp lemma add_le_same_cancel2 [simp]: "a + b \ b \ a \ 0" using add_le_cancel_right [of _ _ 0] by simp lemma le_add_same_cancel1 [simp]: "a \ a + b \ 0 \ b" using add_le_cancel_left [of _ 0] by simp lemma le_add_same_cancel2 [simp]: "a \ b + a \ 0 \ b" using add_le_cancel_right [of 0] by simp subclass cancel_comm_monoid_add by standard auto subclass ordered_cancel_comm_monoid_add by standard end class ordered_ab_group_add = ab_group_add + ordered_ab_semigroup_add begin subclass ordered_cancel_ab_semigroup_add .. subclass ordered_ab_semigroup_monoid_add_imp_le proof fix a b c :: 'a assume "c + a \ c + b" then have "(-c) + (c + a) \ (-c) + (c + b)" by (rule add_left_mono) then have "((-c) + c) + a \ ((-c) + c) + b" by (simp only: add.assoc) then show "a \ b" by simp qed lemma max_diff_distrib_left: "max x y - z = max (x - z) (y - z)" using max_add_distrib_left [of x y "- z"] by simp lemma min_diff_distrib_left: "min x y - z = min (x - z) (y - z)" using min_add_distrib_left [of x y "- z"] by simp lemma le_imp_neg_le: assumes "a \ b" shows "- b \ - a" proof - from assms have "- a + a \ - a + b" by (rule add_left_mono) then have "0 \ - a + b" by simp then have "0 + (- b) \ (- a + b) + (- b)" by (rule add_right_mono) then show ?thesis by (simp add: algebra_simps) qed lemma neg_le_iff_le [simp]: "- b \ - a \ a \ b" proof assume "- b \ - a" then have "- (- a) \ - (- b)" by (rule le_imp_neg_le) then show "a \ b" by simp next assume "a \ b" then show "- b \ - a" by (rule le_imp_neg_le) qed lemma neg_le_0_iff_le [simp]: "- a \ 0 \ 0 \ a" by (subst neg_le_iff_le [symmetric]) simp lemma neg_0_le_iff_le [simp]: "0 \ - a \ a \ 0" by (subst neg_le_iff_le [symmetric]) simp lemma neg_less_iff_less [simp]: "- b < - a \ a < b" by (auto simp add: less_le) lemma neg_less_0_iff_less [simp]: "- a < 0 \ 0 < a" by (subst neg_less_iff_less [symmetric]) simp lemma neg_0_less_iff_less [simp]: "0 < - a \ a < 0" by (subst neg_less_iff_less [symmetric]) simp text \The next several equations can make the simplifier loop!\ lemma less_minus_iff: "a < - b \ b < - a" proof - have "- (- a) < - b \ b < - a" by (rule neg_less_iff_less) then show ?thesis by simp qed lemma minus_less_iff: "- a < b \ - b < a" proof - have "- a < - (- b) \ - b < a" by (rule neg_less_iff_less) then show ?thesis by simp qed lemma le_minus_iff: "a \ - b \ b \ - a" by (auto simp: order.order_iff_strict less_minus_iff) lemma minus_le_iff: "- a \ b \ - b \ a" by (auto simp add: le_less minus_less_iff) lemma diff_less_0_iff_less [simp]: "a - b < 0 \ a < b" proof - have "a - b < 0 \ a + (- b) < b + (- b)" by simp also have "\ \ a < b" by (simp only: add_less_cancel_right) finally show ?thesis . qed lemmas less_iff_diff_less_0 = diff_less_0_iff_less [symmetric] lemma diff_less_eq [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: "a - b < c \ a < c + b" proof (subst less_iff_diff_less_0 [of a]) show "(a - b < c) = (a - (c + b) < 0)" by (simp add: algebra_simps less_iff_diff_less_0 [of _ c]) qed lemma less_diff_eq [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: "a < c - b \ a + b < c" proof (subst less_iff_diff_less_0 [of "a + b"]) show "(a < c - b) = (a + b - c < 0)" by (simp add: algebra_simps less_iff_diff_less_0 [of a]) qed lemma diff_gt_0_iff_gt [simp]: "a - b > 0 \ a > b" by (simp add: less_diff_eq) lemma diff_le_eq [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: "a - b \ c \ a \ c + b" by (auto simp add: le_less diff_less_eq ) lemma le_diff_eq [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: "a \ c - b \ a + b \ c" by (auto simp add: le_less less_diff_eq) lemma diff_le_0_iff_le [simp]: "a - b \ 0 \ a \ b" by (simp add: algebra_simps) lemmas le_iff_diff_le_0 = diff_le_0_iff_le [symmetric] lemma diff_ge_0_iff_ge [simp]: "a - b \ 0 \ a \ b" by (simp add: le_diff_eq) lemma diff_eq_diff_less: "a - b = c - d \ a < b \ c < d" by (auto simp only: less_iff_diff_less_0 [of a b] less_iff_diff_less_0 [of c d]) lemma diff_eq_diff_less_eq: "a - b = c - d \ a \ b \ c \ d" by (auto simp only: le_iff_diff_le_0 [of a b] le_iff_diff_le_0 [of c d]) lemma diff_mono: "a \ b \ d \ c \ a - c \ b - d" by (simp add: field_simps add_mono) lemma diff_left_mono: "b \ a \ c - a \ c - b" by (simp add: field_simps) lemma diff_right_mono: "a \ b \ a - c \ b - c" by (simp add: field_simps) lemma diff_strict_mono: "a < b \ d < c \ a - c < b - d" by (simp add: field_simps add_strict_mono) lemma diff_strict_left_mono: "b < a \ c - a < c - b" by (simp add: field_simps) lemma diff_strict_right_mono: "a < b \ a - c < b - c" by (simp add: field_simps) end locale group_cancel begin lemma add1: "(A::'a::comm_monoid_add) \ k + a \ A + b \ k + (a + b)" by (simp only: ac_simps) lemma add2: "(B::'a::comm_monoid_add) \ k + b \ a + B \ k + (a + b)" by (simp only: ac_simps) lemma sub1: "(A::'a::ab_group_add) \ k + a \ A - b \ k + (a - b)" by (simp only: add_diff_eq) lemma sub2: "(B::'a::ab_group_add) \ k + b \ a - B \ - k + (a - b)" by (simp only: minus_add diff_conv_add_uminus ac_simps) lemma neg1: "(A::'a::ab_group_add) \ k + a \ - A \ - k + - a" by (simp only: minus_add_distrib) lemma rule0: "(a::'a::comm_monoid_add) \ a + 0" by (simp only: add_0_right) end ML_file \Tools/group_cancel.ML\ simproc_setup group_cancel_add ("a + b::'a::ab_group_add") = \fn phi => fn ss => try Group_Cancel.cancel_add_conv\ simproc_setup group_cancel_diff ("a - b::'a::ab_group_add") = \fn phi => fn ss => try Group_Cancel.cancel_diff_conv\ simproc_setup group_cancel_eq ("a = (b::'a::ab_group_add)") = \fn phi => fn ss => try Group_Cancel.cancel_eq_conv\ simproc_setup group_cancel_le ("a \ (b::'a::ordered_ab_group_add)") = \fn phi => fn ss => try Group_Cancel.cancel_le_conv\ simproc_setup group_cancel_less ("a < (b::'a::ordered_ab_group_add)") = \fn phi => fn ss => try Group_Cancel.cancel_less_conv\ class linordered_ab_semigroup_add = linorder + ordered_ab_semigroup_add class linordered_cancel_ab_semigroup_add = linorder + ordered_cancel_ab_semigroup_add begin subclass linordered_ab_semigroup_add .. subclass ordered_ab_semigroup_add_imp_le proof fix a b c :: 'a assume le1: "c + a \ c + b" show "a \ b" proof (rule ccontr) assume *: "\ ?thesis" then have "b \ a" by (simp add: linorder_not_le) then have "c + b \ c + a" by (rule add_left_mono) then have "c + a = c + b" using le1 by (iprover intro: order.antisym) then have "a = b" by simp with * show False by (simp add: linorder_not_le [symmetric]) qed qed end class linordered_ab_group_add = linorder + ordered_ab_group_add begin subclass linordered_cancel_ab_semigroup_add .. lemma equal_neg_zero [simp]: "a = - a \ a = 0" proof assume "a = 0" then show "a = - a" by simp next assume A: "a = - a" show "a = 0" proof (cases "0 \ a") case True with A have "0 \ - a" by auto with le_minus_iff have "a \ 0" by simp with True show ?thesis by (auto intro: order_trans) next case False then have B: "a \ 0" by auto with A have "- a \ 0" by auto with B show ?thesis by (auto intro: order_trans) qed qed lemma neg_equal_zero [simp]: "- a = a \ a = 0" by (auto dest: sym) lemma neg_less_eq_nonneg [simp]: "- a \ a \ 0 \ a" proof assume *: "- a \ a" show "0 \ a" proof (rule classical) assume "\ ?thesis" then have "a < 0" by auto with * have "- a < 0" by (rule le_less_trans) then show ?thesis by auto qed next assume *: "0 \ a" then have "- a \ 0" by (simp add: minus_le_iff) from this * show "- a \ a" by (rule order_trans) qed lemma neg_less_pos [simp]: "- a < a \ 0 < a" by (auto simp add: less_le) lemma less_eq_neg_nonpos [simp]: "a \ - a \ a \ 0" using neg_less_eq_nonneg [of "- a"] by simp lemma less_neg_neg [simp]: "a < - a \ a < 0" using neg_less_pos [of "- a"] by simp lemma double_zero [simp]: "a + a = 0 \ a = 0" proof assume "a + a = 0" then have a: "- a = a" by (rule minus_unique) then show "a = 0" by (simp only: neg_equal_zero) next assume "a = 0" then show "a + a = 0" by simp qed lemma double_zero_sym [simp]: "0 = a + a \ a = 0" using double_zero [of a] by (simp only: eq_commute) lemma zero_less_double_add_iff_zero_less_single_add [simp]: "0 < a + a \ 0 < a" proof assume "0 < a + a" then have "0 - a < a" by (simp only: diff_less_eq) then have "- a < a" by simp then show "0 < a" by simp next assume "0 < a" with this have "0 + 0 < a + a" by (rule add_strict_mono) then show "0 < a + a" by simp qed lemma zero_le_double_add_iff_zero_le_single_add [simp]: "0 \ a + a \ 0 \ a" by (auto simp add: le_less) lemma double_add_less_zero_iff_single_add_less_zero [simp]: "a + a < 0 \ a < 0" proof - have "\ a + a < 0 \ \ a < 0" by (simp add: not_less) then show ?thesis by simp qed lemma double_add_le_zero_iff_single_add_le_zero [simp]: "a + a \ 0 \ a \ 0" proof - have "\ a + a \ 0 \ \ a \ 0" by (simp add: not_le) then show ?thesis by simp qed lemma minus_max_eq_min: "- max x y = min (- x) (- y)" by (auto simp add: max_def min_def) lemma minus_min_eq_max: "- min x y = max (- x) (- y)" by (auto simp add: max_def min_def) end class abs = fixes abs :: "'a \ 'a" ("\_\") class sgn = fixes sgn :: "'a \ 'a" class ordered_ab_group_add_abs = ordered_ab_group_add + abs + assumes abs_ge_zero [simp]: "\a\ \ 0" and abs_ge_self: "a \ \a\" and abs_leI: "a \ b \ - a \ b \ \a\ \ b" and abs_minus_cancel [simp]: "\-a\ = \a\" and abs_triangle_ineq: "\a + b\ \ \a\ + \b\" begin lemma abs_minus_le_zero: "- \a\ \ 0" unfolding neg_le_0_iff_le by simp lemma abs_of_nonneg [simp]: assumes nonneg: "0 \ a" shows "\a\ = a" proof (rule order.antisym) show "a \ \a\" by (rule abs_ge_self) from nonneg le_imp_neg_le have "- a \ 0" by simp from this nonneg have "- a \ a" by (rule order_trans) then show "\a\ \ a" by (auto intro: abs_leI) qed lemma abs_idempotent [simp]: "\\a\\ = \a\" by (rule order.antisym) (auto intro!: abs_ge_self abs_leI order_trans [of "- \a\" 0 "\a\"]) lemma abs_eq_0 [simp]: "\a\ = 0 \ a = 0" proof - have "\a\ = 0 \ a = 0" proof (rule order.antisym) assume zero: "\a\ = 0" with abs_ge_self show "a \ 0" by auto from zero have "\-a\ = 0" by simp with abs_ge_self [of "- a"] have "- a \ 0" by auto with neg_le_0_iff_le show "0 \ a" by auto qed then show ?thesis by auto qed lemma abs_zero [simp]: "\0\ = 0" by simp lemma abs_0_eq [simp]: "0 = \a\ \ a = 0" proof - have "0 = \a\ \ \a\ = 0" by (simp only: eq_ac) then show ?thesis by simp qed lemma abs_le_zero_iff [simp]: "\a\ \ 0 \ a = 0" proof assume "\a\ \ 0" then have "\a\ = 0" by (rule order.antisym) simp then show "a = 0" by simp next assume "a = 0" then show "\a\ \ 0" by simp qed lemma abs_le_self_iff [simp]: "\a\ \ a \ 0 \ a" proof - have "0 \ \a\" using abs_ge_zero by blast then have "\a\ \ a \ 0 \ a" using order.trans by blast then show ?thesis using abs_of_nonneg eq_refl by blast qed lemma zero_less_abs_iff [simp]: "0 < \a\ \ a \ 0" by (simp add: less_le) lemma abs_not_less_zero [simp]: "\ \a\ < 0" proof - have "x \ y \ \ y < x" for x y by auto then show ?thesis by simp qed lemma abs_ge_minus_self: "- a \ \a\" proof - have "- a \ \-a\" by (rule abs_ge_self) then show ?thesis by simp qed lemma abs_minus_commute: "\a - b\ = \b - a\" proof - have "\a - b\ = \- (a - b)\" by (simp only: abs_minus_cancel) also have "\ = \b - a\" by simp finally show ?thesis . qed lemma abs_of_pos: "0 < a \ \a\ = a" by (rule abs_of_nonneg) (rule less_imp_le) lemma abs_of_nonpos [simp]: assumes "a \ 0" shows "\a\ = - a" proof - let ?b = "- a" have "- ?b \ 0 \ \- ?b\ = - (- ?b)" unfolding abs_minus_cancel [of ?b] unfolding neg_le_0_iff_le [of ?b] unfolding minus_minus by (erule abs_of_nonneg) then show ?thesis using assms by auto qed lemma abs_of_neg: "a < 0 \ \a\ = - a" by (rule abs_of_nonpos) (rule less_imp_le) lemma abs_le_D1: "\a\ \ b \ a \ b" using abs_ge_self by (blast intro: order_trans) lemma abs_le_D2: "\a\ \ b \ - a \ b" using abs_le_D1 [of "- a"] by simp lemma abs_le_iff: "\a\ \ b \ a \ b \ - a \ b" by (blast intro: abs_leI dest: abs_le_D1 abs_le_D2) lemma abs_triangle_ineq2: "\a\ - \b\ \ \a - b\" proof - have "\a\ = \b + (a - b)\" by (simp add: algebra_simps) then have "\a\ \ \b\ + \a - b\" by (simp add: abs_triangle_ineq) then show ?thesis by (simp add: algebra_simps) qed lemma abs_triangle_ineq2_sym: "\a\ - \b\ \ \b - a\" by (simp only: abs_minus_commute [of b] abs_triangle_ineq2) lemma abs_triangle_ineq3: "\\a\ - \b\\ \ \a - b\" by (simp add: abs_le_iff abs_triangle_ineq2 abs_triangle_ineq2_sym) lemma abs_triangle_ineq4: "\a - b\ \ \a\ + \b\" proof - have "\a - b\ = \a + - b\" by (simp add: algebra_simps) also have "\ \ \a\ + \- b\" by (rule abs_triangle_ineq) finally show ?thesis by simp qed lemma abs_diff_triangle_ineq: "\a + b - (c + d)\ \ \a - c\ + \b - d\" proof - have "\a + b - (c + d)\ = \(a - c) + (b - d)\" by (simp add: algebra_simps) also have "\ \ \a - c\ + \b - d\" by (rule abs_triangle_ineq) finally show ?thesis . qed lemma abs_add_abs [simp]: "\\a\ + \b\\ = \a\ + \b\" (is "?L = ?R") proof (rule order.antisym) show "?L \ ?R" by (rule abs_ge_self) have "?L \ \\a\\ + \\b\\" by (rule abs_triangle_ineq) also have "\ = ?R" by simp finally show "?L \ ?R" . qed end lemma dense_eq0_I: fixes x::"'a::{dense_linorder,ordered_ab_group_add_abs}" assumes "\e. 0 < e \ \x\ \ e" shows "x = 0" proof (cases "\x\ = 0") case False then have "\x\ > 0" by simp then obtain z where "0 < z" "z < \x\" using dense by force then show ?thesis using assms by (simp flip: not_less) qed auto hide_fact (open) ab_diff_conv_add_uminus add_0 mult_1 ab_left_minus lemmas add_0 = add_0_left (* FIXME duplicate *) lemmas mult_1 = mult_1_left (* FIXME duplicate *) lemmas ab_left_minus = left_minus (* FIXME duplicate *) lemmas diff_diff_eq = diff_diff_add (* FIXME duplicate *) subsection \Canonically ordered monoids\ text \Canonically ordered monoids are never groups.\ class canonically_ordered_monoid_add = comm_monoid_add + order + assumes le_iff_add: "a \ b \ (\c. b = a + c)" begin lemma zero_le[simp]: "0 \ x" by (auto simp: le_iff_add) lemma le_zero_eq[simp]: "n \ 0 \ n = 0" by (auto intro: order.antisym) lemma not_less_zero[simp]: "\ n < 0" by (auto simp: less_le) lemma zero_less_iff_neq_zero: "0 < n \ n \ 0" by (auto simp: less_le) text \This theorem is useful with \blast\\ lemma gr_zeroI: "(n = 0 \ False) \ 0 < n" by (rule zero_less_iff_neq_zero[THEN iffD2]) iprover lemma not_gr_zero[simp]: "\ 0 < n \ n = 0" by (simp add: zero_less_iff_neq_zero) subclass ordered_comm_monoid_add proof qed (auto simp: le_iff_add add_ac) lemma gr_implies_not_zero: "m < n \ n \ 0" by auto lemma add_eq_0_iff_both_eq_0[simp]: "x + y = 0 \ x = 0 \ y = 0" by (intro add_nonneg_eq_0_iff zero_le) lemma zero_eq_add_iff_both_eq_0[simp]: "0 = x + y \ x = 0 \ y = 0" using add_eq_0_iff_both_eq_0[of x y] unfolding eq_commute[of 0] . lemma less_eqE: assumes \a \ b\ obtains c where \b = a + c\ using assms by (auto simp add: le_iff_add) lemma lessE: assumes \a < b\ obtains c where \b = a + c\ and \c \ 0\ proof - from assms have \a \ b\ \a \ b\ by simp_all from \a \ b\ obtain c where \b = a + c\ by (rule less_eqE) moreover have \c \ 0\ using \a \ b\ \b = a + c\ by auto ultimately show ?thesis by (rule that) qed lemmas zero_order = zero_le le_zero_eq not_less_zero zero_less_iff_neq_zero not_gr_zero \ \This should be attributed with \[iff]\, but then \blast\ fails in \Set\.\ end class ordered_cancel_comm_monoid_diff = canonically_ordered_monoid_add + comm_monoid_diff + ordered_ab_semigroup_add_imp_le begin context fixes a b :: 'a assumes le: "a \ b" begin lemma add_diff_inverse: "a + (b - a) = b" using le by (auto simp add: le_iff_add) lemma add_diff_assoc: "c + (b - a) = c + b - a" using le by (auto simp add: le_iff_add add.left_commute [of c]) lemma add_diff_assoc2: "b - a + c = b + c - a" using le by (auto simp add: le_iff_add add.assoc) lemma diff_add_assoc: "c + b - a = c + (b - a)" using le by (simp add: add.commute add_diff_assoc) lemma diff_add_assoc2: "b + c - a = b - a + c" using le by (simp add: add.commute add_diff_assoc) lemma diff_diff_right: "c - (b - a) = c + a - b" by (simp add: add_diff_inverse add_diff_cancel_left [of a c "b - a", symmetric] add.commute) lemma diff_add: "b - a + a = b" by (simp add: add.commute add_diff_inverse) lemma le_add_diff: "c \ b + c - a" by (auto simp add: add.commute diff_add_assoc2 le_iff_add) lemma le_imp_diff_is_add: "a \ b \ b - a = c \ b = c + a" by (auto simp add: add.commute add_diff_inverse) lemma le_diff_conv2: "c \ b - a \ c + a \ b" (is "?P \ ?Q") proof assume ?P then have "c + a \ b - a + a" by (rule add_right_mono) then show ?Q by (simp add: add_diff_inverse add.commute) next assume ?Q then have "a + c \ a + (b - a)" by (simp add: add_diff_inverse add.commute) then show ?P by simp qed end end subsection \Tools setup\ lemma add_mono_thms_linordered_semiring: fixes i j k :: "'a::ordered_ab_semigroup_add" shows "i \ j \ k \ l \ i + k \ j + l" and "i = j \ k \ l \ i + k \ j + l" and "i \ j \ k = l \ i + k \ j + l" and "i = j \ k = l \ i + k = j + l" by (rule add_mono, clarify+)+ lemma add_mono_thms_linordered_field: fixes i j k :: "'a::ordered_cancel_ab_semigroup_add" shows "i < j \ k = l \ i + k < j + l" and "i = j \ k < l \ i + k < j + l" and "i < j \ k \ l \ i + k < j + l" and "i \ j \ k < l \ i + k < j + l" and "i < j \ k < l \ i + k < j + l" by (auto intro: add_strict_right_mono add_strict_left_mono add_less_le_mono add_le_less_mono add_strict_mono) code_identifier code_module Groups \ (SML) Arith and (OCaml) Arith and (Haskell) Arith end diff --git a/src/HOL/Real_Vector_Spaces.thy b/src/HOL/Real_Vector_Spaces.thy --- a/src/HOL/Real_Vector_Spaces.thy +++ b/src/HOL/Real_Vector_Spaces.thy @@ -1,2401 +1,2404 @@ (* Title: HOL/Real_Vector_Spaces.thy Author: Brian Huffman Author: Johannes Hölzl *) section \Vector Spaces and Algebras over the Reals\ theory Real_Vector_Spaces imports Real Topological_Spaces Vector_Spaces begin subsection \Real vector spaces\ class scaleR = fixes scaleR :: "real \ 'a \ 'a" (infixr "*\<^sub>R" 75) begin abbreviation divideR :: "'a \ real \ 'a" (infixl "'/\<^sub>R" 70) where "x /\<^sub>R r \ inverse r *\<^sub>R x" end class real_vector = scaleR + ab_group_add + assumes scaleR_add_right: "a *\<^sub>R (x + y) = a *\<^sub>R x + a *\<^sub>R y" and scaleR_add_left: "(a + b) *\<^sub>R x = a *\<^sub>R x + b *\<^sub>R x" and scaleR_scaleR: "a *\<^sub>R b *\<^sub>R x = (a * b) *\<^sub>R x" and scaleR_one: "1 *\<^sub>R x = x" class real_algebra = real_vector + ring + assumes mult_scaleR_left [simp]: "a *\<^sub>R x * y = a *\<^sub>R (x * y)" and mult_scaleR_right [simp]: "x * a *\<^sub>R y = a *\<^sub>R (x * y)" class real_algebra_1 = real_algebra + ring_1 class real_div_algebra = real_algebra_1 + division_ring class real_field = real_div_algebra + field instantiation real :: real_field begin definition real_scaleR_def [simp]: "scaleR a x = a * x" instance by standard (simp_all add: algebra_simps) end locale linear = Vector_Spaces.linear "scaleR::_\_\'a::real_vector" "scaleR::_\_\'b::real_vector" begin lemmas scaleR = scale end global_interpretation real_vector?: vector_space "scaleR :: real \ 'a \ 'a :: real_vector" rewrites "Vector_Spaces.linear (*\<^sub>R) (*\<^sub>R) = linear" and "Vector_Spaces.linear (*) (*\<^sub>R) = linear" defines dependent_raw_def: dependent = real_vector.dependent and representation_raw_def: representation = real_vector.representation and subspace_raw_def: subspace = real_vector.subspace and span_raw_def: span = real_vector.span and extend_basis_raw_def: extend_basis = real_vector.extend_basis and dim_raw_def: dim = real_vector.dim proof unfold_locales show "Vector_Spaces.linear (*\<^sub>R) (*\<^sub>R) = linear" "Vector_Spaces.linear (*) (*\<^sub>R) = linear" by (force simp: linear_def real_scaleR_def[abs_def])+ qed (use scaleR_add_right scaleR_add_left scaleR_scaleR scaleR_one in auto) hide_const (open)\ \locale constants\ real_vector.dependent real_vector.independent real_vector.representation real_vector.subspace real_vector.span real_vector.extend_basis real_vector.dim abbreviation "independent x \ \ dependent x" global_interpretation real_vector?: vector_space_pair "scaleR::_\_\'a::real_vector" "scaleR::_\_\'b::real_vector" rewrites "Vector_Spaces.linear (*\<^sub>R) (*\<^sub>R) = linear" and "Vector_Spaces.linear (*) (*\<^sub>R) = linear" defines construct_raw_def: construct = real_vector.construct proof unfold_locales show "Vector_Spaces.linear (*) (*\<^sub>R) = linear" unfolding linear_def real_scaleR_def by auto qed (auto simp: linear_def) hide_const (open)\ \locale constants\ real_vector.construct lemma linear_compose: "linear f \ linear g \ linear (g \ f)" unfolding linear_def by (rule Vector_Spaces.linear_compose) text \Recover original theorem names\ lemmas scaleR_left_commute = real_vector.scale_left_commute lemmas scaleR_zero_left = real_vector.scale_zero_left lemmas scaleR_minus_left = real_vector.scale_minus_left lemmas scaleR_diff_left = real_vector.scale_left_diff_distrib lemmas scaleR_sum_left = real_vector.scale_sum_left lemmas scaleR_zero_right = real_vector.scale_zero_right lemmas scaleR_minus_right = real_vector.scale_minus_right lemmas scaleR_diff_right = real_vector.scale_right_diff_distrib lemmas scaleR_sum_right = real_vector.scale_sum_right lemmas scaleR_eq_0_iff = real_vector.scale_eq_0_iff lemmas scaleR_left_imp_eq = real_vector.scale_left_imp_eq lemmas scaleR_right_imp_eq = real_vector.scale_right_imp_eq lemmas scaleR_cancel_left = real_vector.scale_cancel_left lemmas scaleR_cancel_right = real_vector.scale_cancel_right lemma [field_simps]: "c \ 0 \ a = b /\<^sub>R c \ c *\<^sub>R a = b" "c \ 0 \ b /\<^sub>R c = a \ b = c *\<^sub>R a" "c \ 0 \ a + b /\<^sub>R c = (c *\<^sub>R a + b) /\<^sub>R c" "c \ 0 \ a /\<^sub>R c + b = (a + c *\<^sub>R b) /\<^sub>R c" "c \ 0 \ a - b /\<^sub>R c = (c *\<^sub>R a - b) /\<^sub>R c" "c \ 0 \ a /\<^sub>R c - b = (a - c *\<^sub>R b) /\<^sub>R c" "c \ 0 \ - (a /\<^sub>R c) + b = (- a + c *\<^sub>R b) /\<^sub>R c" "c \ 0 \ - (a /\<^sub>R c) - b = (- a - c *\<^sub>R b) /\<^sub>R c" for a b :: "'a :: real_vector" by (auto simp add: scaleR_add_right scaleR_add_left scaleR_diff_right scaleR_diff_left) text \Legacy names\ lemmas scaleR_left_distrib = scaleR_add_left lemmas scaleR_right_distrib = scaleR_add_right lemmas scaleR_left_diff_distrib = scaleR_diff_left lemmas scaleR_right_diff_distrib = scaleR_diff_right lemmas linear_injective_0 = linear_inj_iff_eq_0 and linear_injective_on_subspace_0 = linear_inj_on_iff_eq_0 and linear_cmul = linear_scale and linear_scaleR = linear_scale_self and subspace_mul = subspace_scale and span_linear_image = linear_span_image and span_0 = span_zero and span_mul = span_scale and injective_scaleR = injective_scale lemma scaleR_minus1_left [simp]: "scaleR (-1) x = - x" for x :: "'a::real_vector" using scaleR_minus_left [of 1 x] by simp lemma scaleR_2: fixes x :: "'a::real_vector" shows "scaleR 2 x = x + x" unfolding one_add_one [symmetric] scaleR_left_distrib by simp lemma scaleR_half_double [simp]: fixes a :: "'a::real_vector" shows "(1 / 2) *\<^sub>R (a + a) = a" proof - have "\r. r *\<^sub>R (a + a) = (r * 2) *\<^sub>R a" by (metis scaleR_2 scaleR_scaleR) then show ?thesis by simp qed lemma linear_scale_real: fixes r::real shows "linear f \ f (r * b) = r * f b" using linear_scale by fastforce interpretation scaleR_left: additive "(\a. scaleR a x :: 'a::real_vector)" by standard (rule scaleR_left_distrib) interpretation scaleR_right: additive "(\x. scaleR a x :: 'a::real_vector)" by standard (rule scaleR_right_distrib) lemma nonzero_inverse_scaleR_distrib: "a \ 0 \ x \ 0 \ inverse (scaleR a x) = scaleR (inverse a) (inverse x)" for x :: "'a::real_div_algebra" by (rule inverse_unique) simp lemma inverse_scaleR_distrib: "inverse (scaleR a x) = scaleR (inverse a) (inverse x)" for x :: "'a::{real_div_algebra,division_ring}" by (metis inverse_zero nonzero_inverse_scaleR_distrib scale_eq_0_iff) lemmas sum_constant_scaleR = real_vector.sum_constant_scale\ \legacy name\ named_theorems vector_add_divide_simps "to simplify sums of scaled vectors" lemma [vector_add_divide_simps]: "v + (b / z) *\<^sub>R w = (if z = 0 then v else (z *\<^sub>R v + b *\<^sub>R w) /\<^sub>R z)" "a *\<^sub>R v + (b / z) *\<^sub>R w = (if z = 0 then a *\<^sub>R v else ((a * z) *\<^sub>R v + b *\<^sub>R w) /\<^sub>R z)" "(a / z) *\<^sub>R v + w = (if z = 0 then w else (a *\<^sub>R v + z *\<^sub>R w) /\<^sub>R z)" "(a / z) *\<^sub>R v + b *\<^sub>R w = (if z = 0 then b *\<^sub>R w else (a *\<^sub>R v + (b * z) *\<^sub>R w) /\<^sub>R z)" "v - (b / z) *\<^sub>R w = (if z = 0 then v else (z *\<^sub>R v - b *\<^sub>R w) /\<^sub>R z)" "a *\<^sub>R v - (b / z) *\<^sub>R w = (if z = 0 then a *\<^sub>R v else ((a * z) *\<^sub>R v - b *\<^sub>R w) /\<^sub>R z)" "(a / z) *\<^sub>R v - w = (if z = 0 then -w else (a *\<^sub>R v - z *\<^sub>R w) /\<^sub>R z)" "(a / z) *\<^sub>R v - b *\<^sub>R w = (if z = 0 then -b *\<^sub>R w else (a *\<^sub>R v - (b * z) *\<^sub>R w) /\<^sub>R z)" for v :: "'a :: real_vector" by (simp_all add: divide_inverse_commute scaleR_add_right scaleR_diff_right) lemma eq_vector_fraction_iff [vector_add_divide_simps]: fixes x :: "'a :: real_vector" shows "(x = (u / v) *\<^sub>R a) \ (if v=0 then x = 0 else v *\<^sub>R x = u *\<^sub>R a)" by auto (metis (no_types) divide_eq_1_iff divide_inverse_commute scaleR_one scaleR_scaleR) lemma vector_fraction_eq_iff [vector_add_divide_simps]: fixes x :: "'a :: real_vector" shows "((u / v) *\<^sub>R a = x) \ (if v=0 then x = 0 else u *\<^sub>R a = v *\<^sub>R x)" by (metis eq_vector_fraction_iff) lemma real_vector_affinity_eq: fixes x :: "'a :: real_vector" assumes m0: "m \ 0" shows "m *\<^sub>R x + c = y \ x = inverse m *\<^sub>R y - (inverse m *\<^sub>R c)" (is "?lhs \ ?rhs") proof assume ?lhs then have "m *\<^sub>R x = y - c" by (simp add: field_simps) then have "inverse m *\<^sub>R (m *\<^sub>R x) = inverse m *\<^sub>R (y - c)" by simp then show "x = inverse m *\<^sub>R y - (inverse m *\<^sub>R c)" using m0 by (simp add: scaleR_diff_right) next assume ?rhs with m0 show "m *\<^sub>R x + c = y" by (simp add: scaleR_diff_right) qed lemma real_vector_eq_affinity: "m \ 0 \ y = m *\<^sub>R x + c \ inverse m *\<^sub>R y - (inverse m *\<^sub>R c) = x" for x :: "'a::real_vector" using real_vector_affinity_eq[where m=m and x=x and y=y and c=c] by metis lemma scaleR_eq_iff [simp]: "b + u *\<^sub>R a = a + u *\<^sub>R b \ a = b \ u = 1" for a :: "'a::real_vector" proof (cases "u = 1") case True then show ?thesis by auto next case False have "a = b" if "b + u *\<^sub>R a = a + u *\<^sub>R b" proof - from that have "(u - 1) *\<^sub>R a = (u - 1) *\<^sub>R b" by (simp add: algebra_simps) with False show ?thesis by auto qed then show ?thesis by auto qed lemma scaleR_collapse [simp]: "(1 - u) *\<^sub>R a + u *\<^sub>R a = a" for a :: "'a::real_vector" by (simp add: algebra_simps) subsection \Embedding of the Reals into any \real_algebra_1\: \of_real\\ definition of_real :: "real \ 'a::real_algebra_1" where "of_real r = scaleR r 1" lemma scaleR_conv_of_real: "scaleR r x = of_real r * x" by (simp add: of_real_def) lemma of_real_0 [simp]: "of_real 0 = 0" by (simp add: of_real_def) lemma of_real_1 [simp]: "of_real 1 = 1" by (simp add: of_real_def) lemma of_real_add [simp]: "of_real (x + y) = of_real x + of_real y" by (simp add: of_real_def scaleR_left_distrib) lemma of_real_minus [simp]: "of_real (- x) = - of_real x" by (simp add: of_real_def) lemma of_real_diff [simp]: "of_real (x - y) = of_real x - of_real y" by (simp add: of_real_def scaleR_left_diff_distrib) lemma of_real_mult [simp]: "of_real (x * y) = of_real x * of_real y" by (simp add: of_real_def) lemma of_real_sum[simp]: "of_real (sum f s) = (\x\s. of_real (f x))" by (induct s rule: infinite_finite_induct) auto lemma of_real_prod[simp]: "of_real (prod f s) = (\x\s. of_real (f x))" by (induct s rule: infinite_finite_induct) auto lemma nonzero_of_real_inverse: "x \ 0 \ of_real (inverse x) = inverse (of_real x :: 'a::real_div_algebra)" by (simp add: of_real_def nonzero_inverse_scaleR_distrib) lemma of_real_inverse [simp]: "of_real (inverse x) = inverse (of_real x :: 'a::{real_div_algebra,division_ring})" by (simp add: of_real_def inverse_scaleR_distrib) lemma nonzero_of_real_divide: "y \ 0 \ of_real (x / y) = (of_real x / of_real y :: 'a::real_field)" by (simp add: divide_inverse nonzero_of_real_inverse) lemma of_real_divide [simp]: "of_real (x / y) = (of_real x / of_real y :: 'a::real_div_algebra)" by (simp add: divide_inverse) lemma of_real_power [simp]: "of_real (x ^ n) = (of_real x :: 'a::{real_algebra_1}) ^ n" by (induct n) simp_all lemma of_real_power_int [simp]: "of_real (power_int x n) = power_int (of_real x :: 'a :: {real_div_algebra,division_ring}) n" by (auto simp: power_int_def) lemma of_real_eq_iff [simp]: "of_real x = of_real y \ x = y" by (simp add: of_real_def) lemma inj_of_real: "inj of_real" by (auto intro: injI) lemmas of_real_eq_0_iff [simp] = of_real_eq_iff [of _ 0, simplified] lemmas of_real_eq_1_iff [simp] = of_real_eq_iff [of _ 1, simplified] lemma minus_of_real_eq_of_real_iff [simp]: "-of_real x = of_real y \ -x = y" using of_real_eq_iff[of "-x" y] by (simp only: of_real_minus) lemma of_real_eq_minus_of_real_iff [simp]: "of_real x = -of_real y \ x = -y" using of_real_eq_iff[of x "-y"] by (simp only: of_real_minus) lemma of_real_eq_id [simp]: "of_real = (id :: real \ real)" by (rule ext) (simp add: of_real_def) text \Collapse nested embeddings.\ lemma of_real_of_nat_eq [simp]: "of_real (of_nat n) = of_nat n" by (induct n) auto lemma of_real_of_int_eq [simp]: "of_real (of_int z) = of_int z" by (cases z rule: int_diff_cases) simp lemma of_real_numeral [simp]: "of_real (numeral w) = numeral w" using of_real_of_int_eq [of "numeral w"] by simp lemma of_real_neg_numeral [simp]: "of_real (- numeral w) = - numeral w" using of_real_of_int_eq [of "- numeral w"] by simp lemma numeral_power_int_eq_of_real_cancel_iff [simp]: "power_int (numeral x) n = (of_real y :: 'a :: {real_div_algebra, division_ring}) \ power_int (numeral x) n = y" proof - have "power_int (numeral x) n = (of_real (power_int (numeral x) n) :: 'a)" by simp also have "\ = of_real y \ power_int (numeral x) n = y" by (subst of_real_eq_iff) auto finally show ?thesis . qed lemma of_real_eq_numeral_power_int_cancel_iff [simp]: "(of_real y :: 'a :: {real_div_algebra, division_ring}) = power_int (numeral x) n \ y = power_int (numeral x) n" by (subst (1 2) eq_commute) simp lemma of_real_eq_of_real_power_int_cancel_iff [simp]: "power_int (of_real b :: 'a :: {real_div_algebra, division_ring}) w = of_real x \ power_int b w = x" by (metis of_real_power_int of_real_eq_iff) lemma of_real_in_Ints_iff [simp]: "of_real x \ \ \ x \ \" proof safe fix x assume "(of_real x :: 'a) \ \" then obtain n where "(of_real x :: 'a) = of_int n" by (auto simp: Ints_def) also have "of_int n = of_real (real_of_int n)" by simp finally have "x = real_of_int n" by (subst (asm) of_real_eq_iff) thus "x \ \" by auto qed (auto simp: Ints_def) lemma Ints_of_real [intro]: "x \ \ \ of_real x \ \" by simp text \Every real algebra has characteristic zero.\ instance real_algebra_1 < ring_char_0 proof from inj_of_real inj_of_nat have "inj (of_real \ of_nat)" by (rule inj_compose) then show "inj (of_nat :: nat \ 'a)" by (simp add: comp_def) qed lemma fraction_scaleR_times [simp]: fixes a :: "'a::real_algebra_1" shows "(numeral u / numeral v) *\<^sub>R (numeral w * a) = (numeral u * numeral w / numeral v) *\<^sub>R a" by (metis (no_types, lifting) of_real_numeral scaleR_conv_of_real scaleR_scaleR times_divide_eq_left) lemma inverse_scaleR_times [simp]: fixes a :: "'a::real_algebra_1" shows "(1 / numeral v) *\<^sub>R (numeral w * a) = (numeral w / numeral v) *\<^sub>R a" by (metis divide_inverse_commute inverse_eq_divide of_real_numeral scaleR_conv_of_real scaleR_scaleR) lemma scaleR_times [simp]: fixes a :: "'a::real_algebra_1" shows "(numeral u) *\<^sub>R (numeral w * a) = (numeral u * numeral w) *\<^sub>R a" by (simp add: scaleR_conv_of_real) instance real_field < field_char_0 .. subsection \The Set of Real Numbers\ definition Reals :: "'a::real_algebra_1 set" ("\") where "\ = range of_real" lemma Reals_of_real [simp]: "of_real r \ \" by (simp add: Reals_def) lemma Reals_of_int [simp]: "of_int z \ \" by (subst of_real_of_int_eq [symmetric], rule Reals_of_real) lemma Reals_of_nat [simp]: "of_nat n \ \" by (subst of_real_of_nat_eq [symmetric], rule Reals_of_real) lemma Reals_numeral [simp]: "numeral w \ \" by (subst of_real_numeral [symmetric], rule Reals_of_real) lemma Reals_0 [simp]: "0 \ \" and Reals_1 [simp]: "1 \ \" by (simp_all add: Reals_def) lemma Reals_add [simp]: "a \ \ \ b \ \ \ a + b \ \" by (metis (no_types, opaque_lifting) Reals_def Reals_of_real imageE of_real_add) lemma Reals_minus [simp]: "a \ \ \ - a \ \" by (auto simp: Reals_def) lemma Reals_minus_iff [simp]: "- a \ \ \ a \ \" using Reals_minus by fastforce lemma Reals_diff [simp]: "a \ \ \ b \ \ \ a - b \ \" by (metis Reals_add Reals_minus_iff add_uminus_conv_diff) lemma Reals_mult [simp]: "a \ \ \ b \ \ \ a * b \ \" by (metis (no_types, lifting) Reals_def Reals_of_real imageE of_real_mult) lemma nonzero_Reals_inverse: "a \ \ \ a \ 0 \ inverse a \ \" for a :: "'a::real_div_algebra" by (metis Reals_def Reals_of_real imageE of_real_inverse) lemma Reals_inverse: "a \ \ \ inverse a \ \" for a :: "'a::{real_div_algebra,division_ring}" using nonzero_Reals_inverse by fastforce lemma Reals_inverse_iff [simp]: "inverse x \ \ \ x \ \" for x :: "'a::{real_div_algebra,division_ring}" by (metis Reals_inverse inverse_inverse_eq) lemma nonzero_Reals_divide: "a \ \ \ b \ \ \ b \ 0 \ a / b \ \" for a b :: "'a::real_field" by (simp add: divide_inverse) lemma Reals_divide [simp]: "a \ \ \ b \ \ \ a / b \ \" for a b :: "'a::{real_field,field}" using nonzero_Reals_divide by fastforce lemma Reals_power [simp]: "a \ \ \ a ^ n \ \" for a :: "'a::real_algebra_1" by (metis Reals_def Reals_of_real imageE of_real_power) lemma Reals_cases [cases set: Reals]: assumes "q \ \" obtains (of_real) r where "q = of_real r" unfolding Reals_def proof - from \q \ \\ have "q \ range of_real" unfolding Reals_def . then obtain r where "q = of_real r" .. then show thesis .. qed lemma sum_in_Reals [intro,simp]: "(\i. i \ s \ f i \ \) \ sum f s \ \" proof (induct s rule: infinite_finite_induct) case infinite then show ?case by (metis Reals_0 sum.infinite) qed simp_all lemma prod_in_Reals [intro,simp]: "(\i. i \ s \ f i \ \) \ prod f s \ \" proof (induct s rule: infinite_finite_induct) case infinite then show ?case by (metis Reals_1 prod.infinite) qed simp_all lemma Reals_induct [case_names of_real, induct set: Reals]: "q \ \ \ (\r. P (of_real r)) \ P q" by (rule Reals_cases) auto subsection \Ordered real vector spaces\ class ordered_real_vector = real_vector + ordered_ab_group_add + assumes scaleR_left_mono: "x \ y \ 0 \ a \ a *\<^sub>R x \ a *\<^sub>R y" and scaleR_right_mono: "a \ b \ 0 \ x \ a *\<^sub>R x \ b *\<^sub>R x" begin lemma scaleR_mono: "a \ b \ x \ y \ 0 \ b \ 0 \ x \ a *\<^sub>R x \ b *\<^sub>R y" by (meson order_trans scaleR_left_mono scaleR_right_mono) lemma scaleR_mono': "a \ b \ c \ d \ 0 \ a \ 0 \ c \ a *\<^sub>R c \ b *\<^sub>R d" by (rule scaleR_mono) (auto intro: order.trans) lemma pos_le_divideR_eq [field_simps]: "a \ b /\<^sub>R c \ c *\<^sub>R a \ b" (is "?P \ ?Q") if "0 < c" proof assume ?P with scaleR_left_mono that have "c *\<^sub>R a \ c *\<^sub>R (b /\<^sub>R c)" by simp with that show ?Q by (simp add: scaleR_one scaleR_scaleR inverse_eq_divide) next assume ?Q with scaleR_left_mono that have "c *\<^sub>R a /\<^sub>R c \ b /\<^sub>R c" by simp with that show ?P by (simp add: scaleR_one scaleR_scaleR inverse_eq_divide) qed lemma pos_less_divideR_eq [field_simps]: "a < b /\<^sub>R c \ c *\<^sub>R a < b" if "c > 0" using that pos_le_divideR_eq [of c a b] by (auto simp add: le_less scaleR_scaleR scaleR_one) lemma pos_divideR_le_eq [field_simps]: "b /\<^sub>R c \ a \ b \ c *\<^sub>R a" if "c > 0" using that pos_le_divideR_eq [of "inverse c" b a] by simp lemma pos_divideR_less_eq [field_simps]: "b /\<^sub>R c < a \ b < c *\<^sub>R a" if "c > 0" using that pos_less_divideR_eq [of "inverse c" b a] by simp lemma pos_le_minus_divideR_eq [field_simps]: "a \ - (b /\<^sub>R c) \ c *\<^sub>R a \ - b" if "c > 0" using that by (metis add_minus_cancel diff_0 left_minus minus_minus neg_le_iff_le scaleR_add_right uminus_add_conv_diff pos_le_divideR_eq) lemma pos_less_minus_divideR_eq [field_simps]: "a < - (b /\<^sub>R c) \ c *\<^sub>R a < - b" if "c > 0" using that by (metis le_less less_le_not_le pos_divideR_le_eq pos_divideR_less_eq pos_le_minus_divideR_eq) lemma pos_minus_divideR_le_eq [field_simps]: "- (b /\<^sub>R c) \ a \ - b \ c *\<^sub>R a" if "c > 0" using that by (metis pos_divideR_le_eq pos_le_minus_divideR_eq that inverse_positive_iff_positive le_imp_neg_le minus_minus) lemma pos_minus_divideR_less_eq [field_simps]: "- (b /\<^sub>R c) < a \ - b < c *\<^sub>R a" if "c > 0" using that by (simp add: less_le_not_le pos_le_minus_divideR_eq pos_minus_divideR_le_eq) lemma scaleR_image_atLeastAtMost: "c > 0 \ scaleR c ` {x..y} = {c *\<^sub>R x..c *\<^sub>R y}" apply (auto intro!: scaleR_left_mono simp: image_iff Bex_def) using pos_divideR_le_eq [of c] pos_le_divideR_eq [of c] apply (meson local.order_eq_iff) done end lemma neg_le_divideR_eq [field_simps]: "a \ b /\<^sub>R c \ b \ c *\<^sub>R a" (is "?P \ ?Q") if "c < 0" for a b :: "'a :: ordered_real_vector" using that pos_le_divideR_eq [of "- c" a "- b"] by simp lemma neg_less_divideR_eq [field_simps]: "a < b /\<^sub>R c \ b < c *\<^sub>R a" if "c < 0" for a b :: "'a :: ordered_real_vector" using that neg_le_divideR_eq [of c a b] by (auto simp add: le_less) lemma neg_divideR_le_eq [field_simps]: "b /\<^sub>R c \ a \ c *\<^sub>R a \ b" if "c < 0" for a b :: "'a :: ordered_real_vector" using that pos_divideR_le_eq [of "- c" "- b" a] by simp lemma neg_divideR_less_eq [field_simps]: "b /\<^sub>R c < a \ c *\<^sub>R a < b" if "c < 0" for a b :: "'a :: ordered_real_vector" using that neg_divideR_le_eq [of c b a] by (auto simp add: le_less) lemma neg_le_minus_divideR_eq [field_simps]: "a \ - (b /\<^sub>R c) \ - b \ c *\<^sub>R a" if "c < 0" for a b :: "'a :: ordered_real_vector" using that pos_le_minus_divideR_eq [of "- c" a "- b"] by (simp add: minus_le_iff) lemma neg_less_minus_divideR_eq [field_simps]: "a < - (b /\<^sub>R c) \ - b < c *\<^sub>R a" if "c < 0" for a b :: "'a :: ordered_real_vector" proof - have *: "- b = c *\<^sub>R a \ b = - (c *\<^sub>R a)" by (metis add.inverse_inverse) from that neg_le_minus_divideR_eq [of c a b] show ?thesis by (auto simp add: le_less *) qed lemma neg_minus_divideR_le_eq [field_simps]: "- (b /\<^sub>R c) \ a \ c *\<^sub>R a \ - b" if "c < 0" for a b :: "'a :: ordered_real_vector" using that pos_minus_divideR_le_eq [of "- c" "- b" a] by (simp add: le_minus_iff) lemma neg_minus_divideR_less_eq [field_simps]: "- (b /\<^sub>R c) < a \ c *\<^sub>R a < - b" if "c < 0" for a b :: "'a :: ordered_real_vector" using that by (simp add: less_le_not_le neg_le_minus_divideR_eq neg_minus_divideR_le_eq) lemma [field_split_simps]: "a = b /\<^sub>R c \ (if c = 0 then a = 0 else c *\<^sub>R a = b)" "b /\<^sub>R c = a \ (if c = 0 then a = 0 else b = c *\<^sub>R a)" "a + b /\<^sub>R c = (if c = 0 then a else (c *\<^sub>R a + b) /\<^sub>R c)" "a /\<^sub>R c + b = (if c = 0 then b else (a + c *\<^sub>R b) /\<^sub>R c)" "a - b /\<^sub>R c = (if c = 0 then a else (c *\<^sub>R a - b) /\<^sub>R c)" "a /\<^sub>R c - b = (if c = 0 then - b else (a - c *\<^sub>R b) /\<^sub>R c)" "- (a /\<^sub>R c) + b = (if c = 0 then b else (- a + c *\<^sub>R b) /\<^sub>R c)" "- (a /\<^sub>R c) - b = (if c = 0 then - b else (- a - c *\<^sub>R b) /\<^sub>R c)" for a b :: "'a :: real_vector" by (auto simp add: field_simps) lemma [field_split_simps]: "0 < c \ a \ b /\<^sub>R c \ (if c > 0 then c *\<^sub>R a \ b else if c < 0 then b \ c *\<^sub>R a else a \ 0)" "0 < c \ a < b /\<^sub>R c \ (if c > 0 then c *\<^sub>R a < b else if c < 0 then b < c *\<^sub>R a else a < 0)" "0 < c \ b /\<^sub>R c \ a \ (if c > 0 then b \ c *\<^sub>R a else if c < 0 then c *\<^sub>R a \ b else a \ 0)" "0 < c \ b /\<^sub>R c < a \ (if c > 0 then b < c *\<^sub>R a else if c < 0 then c *\<^sub>R a < b else a > 0)" "0 < c \ a \ - (b /\<^sub>R c) \ (if c > 0 then c *\<^sub>R a \ - b else if c < 0 then - b \ c *\<^sub>R a else a \ 0)" "0 < c \ a < - (b /\<^sub>R c) \ (if c > 0 then c *\<^sub>R a < - b else if c < 0 then - b < c *\<^sub>R a else a < 0)" "0 < c \ - (b /\<^sub>R c) \ a \ (if c > 0 then - b \ c *\<^sub>R a else if c < 0 then c *\<^sub>R a \ - b else a \ 0)" "0 < c \ - (b /\<^sub>R c) < a \ (if c > 0 then - b < c *\<^sub>R a else if c < 0 then c *\<^sub>R a < - b else a > 0)" for a b :: "'a :: ordered_real_vector" by (clarsimp intro!: field_simps)+ lemma scaleR_nonneg_nonneg: "0 \ a \ 0 \ x \ 0 \ a *\<^sub>R x" for x :: "'a::ordered_real_vector" using scaleR_left_mono [of 0 x a] by simp lemma scaleR_nonneg_nonpos: "0 \ a \ x \ 0 \ a *\<^sub>R x \ 0" for x :: "'a::ordered_real_vector" using scaleR_left_mono [of x 0 a] by simp lemma scaleR_nonpos_nonneg: "a \ 0 \ 0 \ x \ a *\<^sub>R x \ 0" for x :: "'a::ordered_real_vector" using scaleR_right_mono [of a 0 x] by simp lemma split_scaleR_neg_le: "(0 \ a \ x \ 0) \ (a \ 0 \ 0 \ x) \ a *\<^sub>R x \ 0" for x :: "'a::ordered_real_vector" by (auto simp: scaleR_nonneg_nonpos scaleR_nonpos_nonneg) lemma le_add_iff1: "a *\<^sub>R e + c \ b *\<^sub>R e + d \ (a - b) *\<^sub>R e + c \ d" for c d e :: "'a::ordered_real_vector" by (simp add: algebra_simps) lemma le_add_iff2: "a *\<^sub>R e + c \ b *\<^sub>R e + d \ c \ (b - a) *\<^sub>R e + d" for c d e :: "'a::ordered_real_vector" by (simp add: algebra_simps) lemma scaleR_left_mono_neg: "b \ a \ c \ 0 \ c *\<^sub>R a \ c *\<^sub>R b" for a b :: "'a::ordered_real_vector" by (drule scaleR_left_mono [of _ _ "- c"], simp_all) lemma scaleR_right_mono_neg: "b \ a \ c \ 0 \ a *\<^sub>R c \ b *\<^sub>R c" for c :: "'a::ordered_real_vector" by (drule scaleR_right_mono [of _ _ "- c"], simp_all) lemma scaleR_nonpos_nonpos: "a \ 0 \ b \ 0 \ 0 \ a *\<^sub>R b" for b :: "'a::ordered_real_vector" using scaleR_right_mono_neg [of a 0 b] by simp lemma split_scaleR_pos_le: "(0 \ a \ 0 \ b) \ (a \ 0 \ b \ 0) \ 0 \ a *\<^sub>R b" for b :: "'a::ordered_real_vector" by (auto simp: scaleR_nonneg_nonneg scaleR_nonpos_nonpos) lemma zero_le_scaleR_iff: fixes b :: "'a::ordered_real_vector" shows "0 \ a *\<^sub>R b \ 0 < a \ 0 \ b \ a < 0 \ b \ 0 \ a = 0" (is "?lhs = ?rhs") proof (cases "a = 0") case True then show ?thesis by simp next case False show ?thesis proof assume ?lhs from \a \ 0\ consider "a > 0" | "a < 0" by arith then show ?rhs proof cases case 1 with \?lhs\ have "inverse a *\<^sub>R 0 \ inverse a *\<^sub>R (a *\<^sub>R b)" by (intro scaleR_mono) auto with 1 show ?thesis by simp next case 2 with \?lhs\ have "- inverse a *\<^sub>R 0 \ - inverse a *\<^sub>R (a *\<^sub>R b)" by (intro scaleR_mono) auto with 2 show ?thesis by simp qed next assume ?rhs then show ?lhs by (auto simp: not_le \a \ 0\ intro!: split_scaleR_pos_le) qed qed lemma scaleR_le_0_iff: "a *\<^sub>R b \ 0 \ 0 < a \ b \ 0 \ a < 0 \ 0 \ b \ a = 0" for b::"'a::ordered_real_vector" by (insert zero_le_scaleR_iff [of "-a" b]) force lemma scaleR_le_cancel_left: "c *\<^sub>R a \ c *\<^sub>R b \ (0 < c \ a \ b) \ (c < 0 \ b \ a)" for b :: "'a::ordered_real_vector" by (auto simp: neq_iff scaleR_left_mono scaleR_left_mono_neg dest: scaleR_left_mono[where a="inverse c"] scaleR_left_mono_neg[where c="inverse c"]) lemma scaleR_le_cancel_left_pos: "0 < c \ c *\<^sub>R a \ c *\<^sub>R b \ a \ b" for b :: "'a::ordered_real_vector" by (auto simp: scaleR_le_cancel_left) lemma scaleR_le_cancel_left_neg: "c < 0 \ c *\<^sub>R a \ c *\<^sub>R b \ b \ a" for b :: "'a::ordered_real_vector" by (auto simp: scaleR_le_cancel_left) lemma scaleR_left_le_one_le: "0 \ x \ a \ 1 \ a *\<^sub>R x \ x" for x :: "'a::ordered_real_vector" and a :: real using scaleR_right_mono[of a 1 x] by simp subsection \Real normed vector spaces\ class dist = fixes dist :: "'a \ 'a \ real" class norm = fixes norm :: "'a \ real" class sgn_div_norm = scaleR + norm + sgn + assumes sgn_div_norm: "sgn x = x /\<^sub>R norm x" class dist_norm = dist + norm + minus + assumes dist_norm: "dist x y = norm (x - y)" class uniformity_dist = dist + uniformity + assumes uniformity_dist: "uniformity = (INF e\{0 <..}. principal {(x, y). dist x y < e})" begin lemma eventually_uniformity_metric: "eventually P uniformity \ (\e>0. \x y. dist x y < e \ P (x, y))" unfolding uniformity_dist by (subst eventually_INF_base) (auto simp: eventually_principal subset_eq intro: bexI[of _ "min _ _"]) end class real_normed_vector = real_vector + sgn_div_norm + dist_norm + uniformity_dist + open_uniformity + assumes norm_eq_zero [simp]: "norm x = 0 \ x = 0" and norm_triangle_ineq: "norm (x + y) \ norm x + norm y" and norm_scaleR [simp]: "norm (scaleR a x) = \a\ * norm x" begin lemma norm_ge_zero [simp]: "0 \ norm x" proof - have "0 = norm (x + -1 *\<^sub>R x)" using scaleR_add_left[of 1 "-1" x] norm_scaleR[of 0 x] by (simp add: scaleR_one) also have "\ \ norm x + norm (-1 *\<^sub>R x)" by (rule norm_triangle_ineq) finally show ?thesis by simp qed +lemma bdd_below_norm_image: "bdd_below (norm ` A)" + by (meson bdd_belowI2 norm_ge_zero) + end class real_normed_algebra = real_algebra + real_normed_vector + assumes norm_mult_ineq: "norm (x * y) \ norm x * norm y" class real_normed_algebra_1 = real_algebra_1 + real_normed_algebra + assumes norm_one [simp]: "norm 1 = 1" lemma (in real_normed_algebra_1) scaleR_power [simp]: "(scaleR x y) ^ n = scaleR (x^n) (y^n)" by (induct n) (simp_all add: scaleR_one scaleR_scaleR mult_ac) class real_normed_div_algebra = real_div_algebra + real_normed_vector + assumes norm_mult: "norm (x * y) = norm x * norm y" class real_normed_field = real_field + real_normed_div_algebra instance real_normed_div_algebra < real_normed_algebra_1 proof show "norm (x * y) \ norm x * norm y" for x y :: 'a by (simp add: norm_mult) next have "norm (1 * 1::'a) = norm (1::'a) * norm (1::'a)" by (rule norm_mult) then show "norm (1::'a) = 1" by simp qed context real_normed_vector begin lemma norm_zero [simp]: "norm (0::'a) = 0" by simp lemma zero_less_norm_iff [simp]: "norm x > 0 \ x \ 0" by (simp add: order_less_le) lemma norm_not_less_zero [simp]: "\ norm x < 0" by (simp add: linorder_not_less) lemma norm_le_zero_iff [simp]: "norm x \ 0 \ x = 0" by (simp add: order_le_less) lemma norm_minus_cancel [simp]: "norm (- x) = norm x" proof - have "- 1 *\<^sub>R x = - (1 *\<^sub>R x)" unfolding add_eq_0_iff2[symmetric] scaleR_add_left[symmetric] using norm_eq_zero by fastforce then have "norm (- x) = norm (scaleR (- 1) x)" by (simp only: scaleR_one) also have "\ = \- 1\ * norm x" by (rule norm_scaleR) finally show ?thesis by simp qed lemma norm_minus_commute: "norm (a - b) = norm (b - a)" proof - have "norm (- (b - a)) = norm (b - a)" by (rule norm_minus_cancel) then show ?thesis by simp qed lemma dist_add_cancel [simp]: "dist (a + b) (a + c) = dist b c" by (simp add: dist_norm) lemma dist_add_cancel2 [simp]: "dist (b + a) (c + a) = dist b c" by (simp add: dist_norm) lemma norm_uminus_minus: "norm (- x - y) = norm (x + y)" by (subst (2) norm_minus_cancel[symmetric], subst minus_add_distrib) simp lemma norm_triangle_ineq2: "norm a - norm b \ norm (a - b)" proof - have "norm (a - b + b) \ norm (a - b) + norm b" by (rule norm_triangle_ineq) then show ?thesis by simp qed lemma norm_triangle_ineq3: "\norm a - norm b\ \ norm (a - b)" proof - have "norm a - norm b \ norm (a - b)" by (simp add: norm_triangle_ineq2) moreover have "norm b - norm a \ norm (a - b)" by (metis norm_minus_commute norm_triangle_ineq2) ultimately show ?thesis by (simp add: abs_le_iff) qed lemma norm_triangle_ineq4: "norm (a - b) \ norm a + norm b" proof - have "norm (a + - b) \ norm a + norm (- b)" by (rule norm_triangle_ineq) then show ?thesis by simp qed lemma norm_triangle_le_diff: "norm x + norm y \ e \ norm (x - y) \ e" by (meson norm_triangle_ineq4 order_trans) lemma norm_diff_ineq: "norm a - norm b \ norm (a + b)" proof - have "norm a - norm (- b) \ norm (a - - b)" by (rule norm_triangle_ineq2) then show ?thesis by simp qed lemma norm_triangle_sub: "norm x \ norm y + norm (x - y)" using norm_triangle_ineq[of "y" "x - y"] by (simp add: field_simps) lemma norm_triangle_le: "norm x + norm y \ e \ norm (x + y) \ e" by (rule norm_triangle_ineq [THEN order_trans]) lemma norm_triangle_lt: "norm x + norm y < e \ norm (x + y) < e" by (rule norm_triangle_ineq [THEN le_less_trans]) lemma norm_add_leD: "norm (a + b) \ c \ norm b \ norm a + c" by (metis ab_semigroup_add_class.add.commute add_commute diff_le_eq norm_diff_ineq order_trans) lemma norm_diff_triangle_ineq: "norm ((a + b) - (c + d)) \ norm (a - c) + norm (b - d)" proof - have "norm ((a + b) - (c + d)) = norm ((a - c) + (b - d))" by (simp add: algebra_simps) also have "\ \ norm (a - c) + norm (b - d)" by (rule norm_triangle_ineq) finally show ?thesis . qed lemma norm_diff_triangle_le: "norm (x - z) \ e1 + e2" if "norm (x - y) \ e1" "norm (y - z) \ e2" proof - have "norm (x - (y + z - y)) \ norm (x - y) + norm (y - z)" using norm_diff_triangle_ineq that diff_diff_eq2 by presburger with that show ?thesis by simp qed lemma norm_diff_triangle_less: "norm (x - z) < e1 + e2" if "norm (x - y) < e1" "norm (y - z) < e2" proof - have "norm (x - z) \ norm (x - y) + norm (y - z)" by (metis norm_diff_triangle_ineq add_diff_cancel_left' diff_diff_eq2) with that show ?thesis by auto qed lemma norm_triangle_mono: "norm a \ r \ norm b \ s \ norm (a + b) \ r + s" by (metis (mono_tags) add_mono_thms_linordered_semiring(1) norm_triangle_ineq order.trans) lemma norm_sum: "norm (sum f A) \ (\i\A. norm (f i))" for f::"'b \ 'a" by (induct A rule: infinite_finite_induct) (auto intro: norm_triangle_mono) lemma sum_norm_le: "norm (sum f S) \ sum g S" if "\x. x \ S \ norm (f x) \ g x" for f::"'b \ 'a" by (rule order_trans [OF norm_sum sum_mono]) (simp add: that) lemma abs_norm_cancel [simp]: "\norm a\ = norm a" by (rule abs_of_nonneg [OF norm_ge_zero]) lemma sum_norm_bound: "norm (sum f S) \ of_nat (card S)*K" if "\x. x \ S \ norm (f x) \ K" for f :: "'b \ 'a" using sum_norm_le[OF that] sum_constant[symmetric] by simp lemma norm_add_less: "norm x < r \ norm y < s \ norm (x + y) < r + s" by (rule order_le_less_trans [OF norm_triangle_ineq add_strict_mono]) end lemma dist_scaleR [simp]: "dist (x *\<^sub>R a) (y *\<^sub>R a) = \x - y\ * norm a" for a :: "'a::real_normed_vector" by (metis dist_norm norm_scaleR scaleR_left.diff) lemma norm_mult_less: "norm x < r \ norm y < s \ norm (x * y) < r * s" for x y :: "'a::real_normed_algebra" by (rule order_le_less_trans [OF norm_mult_ineq]) (simp add: mult_strict_mono') lemma norm_of_real [simp]: "norm (of_real r :: 'a::real_normed_algebra_1) = \r\" by (simp add: of_real_def) lemma norm_numeral [simp]: "norm (numeral w::'a::real_normed_algebra_1) = numeral w" by (subst of_real_numeral [symmetric], subst norm_of_real, simp) lemma norm_neg_numeral [simp]: "norm (- numeral w::'a::real_normed_algebra_1) = numeral w" by (subst of_real_neg_numeral [symmetric], subst norm_of_real, simp) lemma norm_of_real_add1 [simp]: "norm (of_real x + 1 :: 'a :: real_normed_div_algebra) = \x + 1\" by (metis norm_of_real of_real_1 of_real_add) lemma norm_of_real_addn [simp]: "norm (of_real x + numeral b :: 'a :: real_normed_div_algebra) = \x + numeral b\" by (metis norm_of_real of_real_add of_real_numeral) lemma norm_of_int [simp]: "norm (of_int z::'a::real_normed_algebra_1) = \of_int z\" by (subst of_real_of_int_eq [symmetric], rule norm_of_real) lemma norm_of_nat [simp]: "norm (of_nat n::'a::real_normed_algebra_1) = of_nat n" by (metis abs_of_nat norm_of_real of_real_of_nat_eq) lemma nonzero_norm_inverse: "a \ 0 \ norm (inverse a) = inverse (norm a)" for a :: "'a::real_normed_div_algebra" by (metis inverse_unique norm_mult norm_one right_inverse) lemma norm_inverse: "norm (inverse a) = inverse (norm a)" for a :: "'a::{real_normed_div_algebra,division_ring}" by (metis inverse_zero nonzero_norm_inverse norm_zero) lemma nonzero_norm_divide: "b \ 0 \ norm (a / b) = norm a / norm b" for a b :: "'a::real_normed_field" by (simp add: divide_inverse norm_mult nonzero_norm_inverse) lemma norm_divide: "norm (a / b) = norm a / norm b" for a b :: "'a::{real_normed_field,field}" by (simp add: divide_inverse norm_mult norm_inverse) lemma norm_inverse_le_norm: fixes x :: "'a::real_normed_div_algebra" shows "r \ norm x \ 0 < r \ norm (inverse x) \ inverse r" by (simp add: le_imp_inverse_le norm_inverse) lemma norm_power_ineq: "norm (x ^ n) \ norm x ^ n" for x :: "'a::real_normed_algebra_1" proof (induct n) case 0 show "norm (x ^ 0) \ norm x ^ 0" by simp next case (Suc n) have "norm (x * x ^ n) \ norm x * norm (x ^ n)" by (rule norm_mult_ineq) also from Suc have "\ \ norm x * norm x ^ n" using norm_ge_zero by (rule mult_left_mono) finally show "norm (x ^ Suc n) \ norm x ^ Suc n" by simp qed lemma norm_power: "norm (x ^ n) = norm x ^ n" for x :: "'a::real_normed_div_algebra" by (induct n) (simp_all add: norm_mult) lemma norm_power_int: "norm (power_int x n) = power_int (norm x) n" for x :: "'a::real_normed_div_algebra" by (cases n rule: int_cases4) (auto simp: norm_power power_int_minus norm_inverse) lemma power_eq_imp_eq_norm: fixes w :: "'a::real_normed_div_algebra" assumes eq: "w ^ n = z ^ n" and "n > 0" shows "norm w = norm z" proof - have "norm w ^ n = norm z ^ n" by (metis (no_types) eq norm_power) then show ?thesis using assms by (force intro: power_eq_imp_eq_base) qed lemma power_eq_1_iff: fixes w :: "'a::real_normed_div_algebra" shows "w ^ n = 1 \ norm w = 1 \ n = 0" by (metis norm_one power_0_left power_eq_0_iff power_eq_imp_eq_norm power_one) lemma norm_mult_numeral1 [simp]: "norm (numeral w * a) = numeral w * norm a" for a b :: "'a::{real_normed_field,field}" by (simp add: norm_mult) lemma norm_mult_numeral2 [simp]: "norm (a * numeral w) = norm a * numeral w" for a b :: "'a::{real_normed_field,field}" by (simp add: norm_mult) lemma norm_divide_numeral [simp]: "norm (a / numeral w) = norm a / numeral w" for a b :: "'a::{real_normed_field,field}" by (simp add: norm_divide) lemma norm_of_real_diff [simp]: "norm (of_real b - of_real a :: 'a::real_normed_algebra_1) \ \b - a\" by (metis norm_of_real of_real_diff order_refl) text \Despite a superficial resemblance, \norm_eq_1\ is not relevant.\ lemma square_norm_one: fixes x :: "'a::real_normed_div_algebra" assumes "x\<^sup>2 = 1" shows "norm x = 1" by (metis assms norm_minus_cancel norm_one power2_eq_1_iff) lemma norm_less_p1: "norm x < norm (of_real (norm x) + 1 :: 'a)" for x :: "'a::real_normed_algebra_1" proof - have "norm x < norm (of_real (norm x + 1) :: 'a)" by (simp add: of_real_def) then show ?thesis by simp qed lemma prod_norm: "prod (\x. norm (f x)) A = norm (prod f A)" for f :: "'a \ 'b::{comm_semiring_1,real_normed_div_algebra}" by (induct A rule: infinite_finite_induct) (auto simp: norm_mult) lemma norm_prod_le: "norm (prod f A) \ (\a\A. norm (f a :: 'a :: {real_normed_algebra_1,comm_monoid_mult}))" proof (induct A rule: infinite_finite_induct) case empty then show ?case by simp next case (insert a A) then have "norm (prod f (insert a A)) \ norm (f a) * norm (prod f A)" by (simp add: norm_mult_ineq) also have "norm (prod f A) \ (\a\A. norm (f a))" by (rule insert) finally show ?case by (simp add: insert mult_left_mono) next case infinite then show ?case by simp qed lemma norm_prod_diff: fixes z w :: "'i \ 'a::{real_normed_algebra_1, comm_monoid_mult}" shows "(\i. i \ I \ norm (z i) \ 1) \ (\i. i \ I \ norm (w i) \ 1) \ norm ((\i\I. z i) - (\i\I. w i)) \ (\i\I. norm (z i - w i))" proof (induction I rule: infinite_finite_induct) case empty then show ?case by simp next case (insert i I) note insert.hyps[simp] have "norm ((\i\insert i I. z i) - (\i\insert i I. w i)) = norm ((\i\I. z i) * (z i - w i) + ((\i\I. z i) - (\i\I. w i)) * w i)" (is "_ = norm (?t1 + ?t2)") by (auto simp: field_simps) also have "\ \ norm ?t1 + norm ?t2" by (rule norm_triangle_ineq) also have "norm ?t1 \ norm (\i\I. z i) * norm (z i - w i)" by (rule norm_mult_ineq) also have "\ \ (\i\I. norm (z i)) * norm(z i - w i)" by (rule mult_right_mono) (auto intro: norm_prod_le) also have "(\i\I. norm (z i)) \ (\i\I. 1)" by (intro prod_mono) (auto intro!: insert) also have "norm ?t2 \ norm ((\i\I. z i) - (\i\I. w i)) * norm (w i)" by (rule norm_mult_ineq) also have "norm (w i) \ 1" by (auto intro: insert) also have "norm ((\i\I. z i) - (\i\I. w i)) \ (\i\I. norm (z i - w i))" using insert by auto finally show ?case by (auto simp: ac_simps mult_right_mono mult_left_mono) next case infinite then show ?case by simp qed lemma norm_power_diff: fixes z w :: "'a::{real_normed_algebra_1, comm_monoid_mult}" assumes "norm z \ 1" "norm w \ 1" shows "norm (z^m - w^m) \ m * norm (z - w)" proof - have "norm (z^m - w^m) = norm ((\ i < m. z) - (\ i < m. w))" by simp also have "\ \ (\i = m * norm (z - w)" by simp finally show ?thesis . qed subsection \Metric spaces\ class metric_space = uniformity_dist + open_uniformity + assumes dist_eq_0_iff [simp]: "dist x y = 0 \ x = y" and dist_triangle2: "dist x y \ dist x z + dist y z" begin lemma dist_self [simp]: "dist x x = 0" by simp lemma zero_le_dist [simp]: "0 \ dist x y" using dist_triangle2 [of x x y] by simp lemma zero_less_dist_iff: "0 < dist x y \ x \ y" by (simp add: less_le) lemma dist_not_less_zero [simp]: "\ dist x y < 0" by (simp add: not_less) lemma dist_le_zero_iff [simp]: "dist x y \ 0 \ x = y" by (simp add: le_less) lemma dist_commute: "dist x y = dist y x" proof (rule order_antisym) show "dist x y \ dist y x" using dist_triangle2 [of x y x] by simp show "dist y x \ dist x y" using dist_triangle2 [of y x y] by simp qed lemma dist_commute_lessI: "dist y x < e \ dist x y < e" by (simp add: dist_commute) lemma dist_triangle: "dist x z \ dist x y + dist y z" using dist_triangle2 [of x z y] by (simp add: dist_commute) lemma dist_triangle3: "dist x y \ dist a x + dist a y" using dist_triangle2 [of x y a] by (simp add: dist_commute) lemma abs_dist_diff_le: "\dist a b - dist b c\ \ dist a c" using dist_triangle3[of b c a] dist_triangle2[of a b c] by simp lemma dist_pos_lt: "x \ y \ 0 < dist x y" by (simp add: zero_less_dist_iff) lemma dist_nz: "x \ y \ 0 < dist x y" by (simp add: zero_less_dist_iff) declare dist_nz [symmetric, simp] lemma dist_triangle_le: "dist x z + dist y z \ e \ dist x y \ e" by (rule order_trans [OF dist_triangle2]) lemma dist_triangle_lt: "dist x z + dist y z < e \ dist x y < e" by (rule le_less_trans [OF dist_triangle2]) lemma dist_triangle_less_add: "dist x1 y < e1 \ dist x2 y < e2 \ dist x1 x2 < e1 + e2" by (rule dist_triangle_lt [where z=y]) simp lemma dist_triangle_half_l: "dist x1 y < e / 2 \ dist x2 y < e / 2 \ dist x1 x2 < e" by (rule dist_triangle_lt [where z=y]) simp lemma dist_triangle_half_r: "dist y x1 < e / 2 \ dist y x2 < e / 2 \ dist x1 x2 < e" by (rule dist_triangle_half_l) (simp_all add: dist_commute) lemma dist_triangle_third: assumes "dist x1 x2 < e/3" "dist x2 x3 < e/3" "dist x3 x4 < e/3" shows "dist x1 x4 < e" proof - have "dist x1 x3 < e/3 + e/3" by (metis assms(1) assms(2) dist_commute dist_triangle_less_add) then have "dist x1 x4 < (e/3 + e/3) + e/3" by (metis assms(3) dist_commute dist_triangle_less_add) then show ?thesis by simp qed subclass uniform_space proof fix E x assume "eventually E uniformity" then obtain e where E: "0 < e" "\x y. dist x y < e \ E (x, y)" by (auto simp: eventually_uniformity_metric) then show "E (x, x)" "\\<^sub>F (x, y) in uniformity. E (y, x)" by (auto simp: eventually_uniformity_metric dist_commute) show "\D. eventually D uniformity \ (\x y z. D (x, y) \ D (y, z) \ E (x, z))" using E dist_triangle_half_l[where e=e] unfolding eventually_uniformity_metric by (intro exI[of _ "\(x, y). dist x y < e / 2"] exI[of _ "e/2"] conjI) (auto simp: dist_commute) qed lemma open_dist: "open S \ (\x\S. \e>0. \y. dist y x < e \ y \ S)" by (simp add: dist_commute open_uniformity eventually_uniformity_metric) lemma open_ball: "open {y. dist x y < d}" unfolding open_dist proof (intro ballI) fix y assume *: "y \ {y. dist x y < d}" then show "\e>0. \z. dist z y < e \ z \ {y. dist x y < d}" by (auto intro!: exI[of _ "d - dist x y"] simp: field_simps dist_triangle_lt) qed subclass first_countable_topology proof fix x show "\A::nat \ 'a set. (\i. x \ A i \ open (A i)) \ (\S. open S \ x \ S \ (\i. A i \ S))" proof (safe intro!: exI[of _ "\n. {y. dist x y < inverse (Suc n)}"]) fix S assume "open S" "x \ S" then obtain e where e: "0 < e" and "{y. dist x y < e} \ S" by (auto simp: open_dist subset_eq dist_commute) moreover from e obtain i where "inverse (Suc i) < e" by (auto dest!: reals_Archimedean) then have "{y. dist x y < inverse (Suc i)} \ {y. dist x y < e}" by auto ultimately show "\i. {y. dist x y < inverse (Suc i)} \ S" by blast qed (auto intro: open_ball) qed end instance metric_space \ t2_space proof fix x y :: "'a::metric_space" assume xy: "x \ y" let ?U = "{y'. dist x y' < dist x y / 2}" let ?V = "{x'. dist y x' < dist x y / 2}" have *: "d x z \ d x y + d y z \ d y z = d z y \ \ (d x y * 2 < d x z \ d z y * 2 < d x z)" for d :: "'a \ 'a \ real" and x y z :: 'a by arith have "open ?U \ open ?V \ x \ ?U \ y \ ?V \ ?U \ ?V = {}" using dist_pos_lt[OF xy] *[of dist, OF dist_triangle dist_commute] using open_ball[of _ "dist x y / 2"] by auto then show "\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" by blast qed text \Every normed vector space is a metric space.\ instance real_normed_vector < metric_space proof fix x y z :: 'a show "dist x y = 0 \ x = y" by (simp add: dist_norm) show "dist x y \ dist x z + dist y z" using norm_triangle_ineq4 [of "x - z" "y - z"] by (simp add: dist_norm) qed subsection \Class instances for real numbers\ instantiation real :: real_normed_field begin definition dist_real_def: "dist x y = \x - y\" definition uniformity_real_def [code del]: "(uniformity :: (real \ real) filter) = (INF e\{0 <..}. principal {(x, y). dist x y < e})" definition open_real_def [code del]: "open (U :: real set) \ (\x\U. eventually (\(x', y). x' = x \ y \ U) uniformity)" definition real_norm_def [simp]: "norm r = \r\" instance by intro_classes (auto simp: abs_mult open_real_def dist_real_def sgn_real_def uniformity_real_def) end declare uniformity_Abort[where 'a=real, code] lemma dist_of_real [simp]: "dist (of_real x :: 'a) (of_real y) = dist x y" for a :: "'a::real_normed_div_algebra" by (metis dist_norm norm_of_real of_real_diff real_norm_def) declare [[code abort: "open :: real set \ bool"]] instance real :: linorder_topology proof show "(open :: real set \ bool) = generate_topology (range lessThan \ range greaterThan)" proof (rule ext, safe) fix S :: "real set" assume "open S" then obtain f where "\x\S. 0 < f x \ (\y. dist y x < f x \ y \ S)" unfolding open_dist bchoice_iff .. then have *: "(\x\S. {x - f x <..} \ {..< x + f x}) = S" (is "?S = S") by (fastforce simp: dist_real_def) moreover have "generate_topology (range lessThan \ range greaterThan) ?S" by (force intro: generate_topology.Basis generate_topology_Union generate_topology.Int) ultimately show "generate_topology (range lessThan \ range greaterThan) S" by simp next fix S :: "real set" assume "generate_topology (range lessThan \ range greaterThan) S" moreover have "\a::real. open {.. (\y. \y - x\ < a - x \ y \ {..e>0. \y. \y - x\ < e \ y \ {..a::real. open {a <..}" unfolding open_dist dist_real_def proof clarify fix x a :: real assume "a < x" then have "0 < x - a \ (\y. \y - x\ < x - a \ y \ {a<..})" by auto then show "\e>0. \y. \y - x\ < e \ y \ {a<..}" .. qed ultimately show "open S" by induct auto qed qed instance real :: linear_continuum_topology .. lemmas open_real_greaterThan = open_greaterThan[where 'a=real] lemmas open_real_lessThan = open_lessThan[where 'a=real] lemmas open_real_greaterThanLessThan = open_greaterThanLessThan[where 'a=real] lemmas closed_real_atMost = closed_atMost[where 'a=real] lemmas closed_real_atLeast = closed_atLeast[where 'a=real] lemmas closed_real_atLeastAtMost = closed_atLeastAtMost[where 'a=real] instance real :: ordered_real_vector by standard (auto intro: mult_left_mono mult_right_mono) subsection \Extra type constraints\ text \Only allow \<^term>\open\ in class \topological_space\.\ setup \Sign.add_const_constraint (\<^const_name>\open\, SOME \<^typ>\'a::topological_space set \ bool\)\ text \Only allow \<^term>\uniformity\ in class \uniform_space\.\ setup \Sign.add_const_constraint (\<^const_name>\uniformity\, SOME \<^typ>\('a::uniformity \ 'a) filter\)\ text \Only allow \<^term>\dist\ in class \metric_space\.\ setup \Sign.add_const_constraint (\<^const_name>\dist\, SOME \<^typ>\'a::metric_space \ 'a \ real\)\ text \Only allow \<^term>\norm\ in class \real_normed_vector\.\ setup \Sign.add_const_constraint (\<^const_name>\norm\, SOME \<^typ>\'a::real_normed_vector \ real\)\ subsection \Sign function\ lemma norm_sgn: "norm (sgn x) = (if x = 0 then 0 else 1)" for x :: "'a::real_normed_vector" by (simp add: sgn_div_norm) lemma sgn_zero [simp]: "sgn (0::'a::real_normed_vector) = 0" by (simp add: sgn_div_norm) lemma sgn_zero_iff: "sgn x = 0 \ x = 0" for x :: "'a::real_normed_vector" by (simp add: sgn_div_norm) lemma sgn_minus: "sgn (- x) = - sgn x" for x :: "'a::real_normed_vector" by (simp add: sgn_div_norm) lemma sgn_scaleR: "sgn (scaleR r x) = scaleR (sgn r) (sgn x)" for x :: "'a::real_normed_vector" by (simp add: sgn_div_norm ac_simps) lemma sgn_one [simp]: "sgn (1::'a::real_normed_algebra_1) = 1" by (simp add: sgn_div_norm) lemma sgn_of_real: "sgn (of_real r :: 'a::real_normed_algebra_1) = of_real (sgn r)" unfolding of_real_def by (simp only: sgn_scaleR sgn_one) lemma sgn_mult: "sgn (x * y) = sgn x * sgn y" for x y :: "'a::real_normed_div_algebra" by (simp add: sgn_div_norm norm_mult) hide_fact (open) sgn_mult lemma real_sgn_eq: "sgn x = x / \x\" for x :: real by (simp add: sgn_div_norm divide_inverse) lemma zero_le_sgn_iff [simp]: "0 \ sgn x \ 0 \ x" for x :: real by (cases "0::real" x rule: linorder_cases) simp_all lemma sgn_le_0_iff [simp]: "sgn x \ 0 \ x \ 0" for x :: real by (cases "0::real" x rule: linorder_cases) simp_all lemma norm_conv_dist: "norm x = dist x 0" unfolding dist_norm by simp declare norm_conv_dist [symmetric, simp] lemma dist_0_norm [simp]: "dist 0 x = norm x" for x :: "'a::real_normed_vector" by (simp add: dist_norm) lemma dist_diff [simp]: "dist a (a - b) = norm b" "dist (a - b) a = norm b" by (simp_all add: dist_norm) lemma dist_of_int: "dist (of_int m) (of_int n :: 'a :: real_normed_algebra_1) = of_int \m - n\" proof - have "dist (of_int m) (of_int n :: 'a) = dist (of_int m :: 'a) (of_int m - (of_int (m - n)))" by simp also have "\ = of_int \m - n\" by (subst dist_diff, subst norm_of_int) simp finally show ?thesis . qed lemma dist_of_nat: "dist (of_nat m) (of_nat n :: 'a :: real_normed_algebra_1) = of_int \int m - int n\" by (subst (1 2) of_int_of_nat_eq [symmetric]) (rule dist_of_int) subsection \Bounded Linear and Bilinear Operators\ lemma linearI: "linear f" if "\b1 b2. f (b1 + b2) = f b1 + f b2" "\r b. f (r *\<^sub>R b) = r *\<^sub>R f b" using that by unfold_locales (auto simp: algebra_simps) lemma linear_iff: "linear f \ (\x y. f (x + y) = f x + f y) \ (\c x. f (c *\<^sub>R x) = c *\<^sub>R f x)" (is "linear f \ ?rhs") proof assume "linear f" then interpret f: linear f . show "?rhs" by (simp add: f.add f.scale) next assume "?rhs" then show "linear f" by (intro linearI) auto qed lemmas linear_scaleR_left = linear_scale_left lemmas linear_imp_scaleR = linear_imp_scale corollary real_linearD: fixes f :: "real \ real" assumes "linear f" obtains c where "f = (*) c" by (rule linear_imp_scaleR [OF assms]) (force simp: scaleR_conv_of_real) lemma linear_times_of_real: "linear (\x. a * of_real x)" by (auto intro!: linearI simp: distrib_left) (metis mult_scaleR_right scaleR_conv_of_real) locale bounded_linear = linear f for f :: "'a::real_normed_vector \ 'b::real_normed_vector" + assumes bounded: "\K. \x. norm (f x) \ norm x * K" begin lemma pos_bounded: "\K>0. \x. norm (f x) \ norm x * K" proof - obtain K where K: "\x. norm (f x) \ norm x * K" using bounded by blast show ?thesis proof (intro exI impI conjI allI) show "0 < max 1 K" by (rule order_less_le_trans [OF zero_less_one max.cobounded1]) next fix x have "norm (f x) \ norm x * K" using K . also have "\ \ norm x * max 1 K" by (rule mult_left_mono [OF max.cobounded2 norm_ge_zero]) finally show "norm (f x) \ norm x * max 1 K" . qed qed lemma nonneg_bounded: "\K\0. \x. norm (f x) \ norm x * K" using pos_bounded by (auto intro: order_less_imp_le) lemma linear: "linear f" by (fact local.linear_axioms) end lemma bounded_linear_intro: assumes "\x y. f (x + y) = f x + f y" and "\r x. f (scaleR r x) = scaleR r (f x)" and "\x. norm (f x) \ norm x * K" shows "bounded_linear f" by standard (blast intro: assms)+ locale bounded_bilinear = fixes prod :: "'a::real_normed_vector \ 'b::real_normed_vector \ 'c::real_normed_vector" (infixl "**" 70) assumes add_left: "prod (a + a') b = prod a b + prod a' b" and add_right: "prod a (b + b') = prod a b + prod a b'" and scaleR_left: "prod (scaleR r a) b = scaleR r (prod a b)" and scaleR_right: "prod a (scaleR r b) = scaleR r (prod a b)" and bounded: "\K. \a b. norm (prod a b) \ norm a * norm b * K" begin lemma pos_bounded: "\K>0. \a b. norm (a ** b) \ norm a * norm b * K" proof - obtain K where "\a b. norm (a ** b) \ norm a * norm b * K" using bounded by blast then have "norm (a ** b) \ norm a * norm b * (max 1 K)" for a b by (rule order.trans) (simp add: mult_left_mono) then show ?thesis by force qed lemma nonneg_bounded: "\K\0. \a b. norm (a ** b) \ norm a * norm b * K" using pos_bounded by (auto intro: order_less_imp_le) lemma additive_right: "additive (\b. prod a b)" by (rule additive.intro, rule add_right) lemma additive_left: "additive (\a. prod a b)" by (rule additive.intro, rule add_left) lemma zero_left: "prod 0 b = 0" by (rule additive.zero [OF additive_left]) lemma zero_right: "prod a 0 = 0" by (rule additive.zero [OF additive_right]) lemma minus_left: "prod (- a) b = - prod a b" by (rule additive.minus [OF additive_left]) lemma minus_right: "prod a (- b) = - prod a b" by (rule additive.minus [OF additive_right]) lemma diff_left: "prod (a - a') b = prod a b - prod a' b" by (rule additive.diff [OF additive_left]) lemma diff_right: "prod a (b - b') = prod a b - prod a b'" by (rule additive.diff [OF additive_right]) lemma sum_left: "prod (sum g S) x = sum ((\i. prod (g i) x)) S" by (rule additive.sum [OF additive_left]) lemma sum_right: "prod x (sum g S) = sum ((\i. (prod x (g i)))) S" by (rule additive.sum [OF additive_right]) lemma bounded_linear_left: "bounded_linear (\a. a ** b)" proof - obtain K where "\a b. norm (a ** b) \ norm a * norm b * K" using pos_bounded by blast then show ?thesis by (rule_tac K="norm b * K" in bounded_linear_intro) (auto simp: algebra_simps scaleR_left add_left) qed lemma bounded_linear_right: "bounded_linear (\b. a ** b)" proof - obtain K where "\a b. norm (a ** b) \ norm a * norm b * K" using pos_bounded by blast then show ?thesis by (rule_tac K="norm a * K" in bounded_linear_intro) (auto simp: algebra_simps scaleR_right add_right) qed lemma prod_diff_prod: "(x ** y - a ** b) = (x - a) ** (y - b) + (x - a) ** b + a ** (y - b)" by (simp add: diff_left diff_right) lemma flip: "bounded_bilinear (\x y. y ** x)" proof show "\K. \a b. norm (b ** a) \ norm a * norm b * K" by (metis bounded mult.commute) qed (simp_all add: add_right add_left scaleR_right scaleR_left) lemma comp1: assumes "bounded_linear g" shows "bounded_bilinear (\x. (**) (g x))" proof unfold_locales interpret g: bounded_linear g by fact show "\a a' b. g (a + a') ** b = g a ** b + g a' ** b" "\a b b'. g a ** (b + b') = g a ** b + g a ** b'" "\r a b. g (r *\<^sub>R a) ** b = r *\<^sub>R (g a ** b)" "\a r b. g a ** (r *\<^sub>R b) = r *\<^sub>R (g a ** b)" by (auto simp: g.add add_left add_right g.scaleR scaleR_left scaleR_right) from g.nonneg_bounded nonneg_bounded obtain K L where nn: "0 \ K" "0 \ L" and K: "\x. norm (g x) \ norm x * K" and L: "\a b. norm (a ** b) \ norm a * norm b * L" by auto have "norm (g a ** b) \ norm a * K * norm b * L" for a b by (auto intro!: order_trans[OF K] order_trans[OF L] mult_mono simp: nn) then show "\K. \a b. norm (g a ** b) \ norm a * norm b * K" by (auto intro!: exI[where x="K * L"] simp: ac_simps) qed lemma comp: "bounded_linear f \ bounded_linear g \ bounded_bilinear (\x y. f x ** g y)" by (rule bounded_bilinear.flip[OF bounded_bilinear.comp1[OF bounded_bilinear.flip[OF comp1]]]) end lemma bounded_linear_ident[simp]: "bounded_linear (\x. x)" by standard (auto intro!: exI[of _ 1]) lemma bounded_linear_zero[simp]: "bounded_linear (\x. 0)" by standard (auto intro!: exI[of _ 1]) lemma bounded_linear_add: assumes "bounded_linear f" and "bounded_linear g" shows "bounded_linear (\x. f x + g x)" proof - interpret f: bounded_linear f by fact interpret g: bounded_linear g by fact show ?thesis proof from f.bounded obtain Kf where Kf: "norm (f x) \ norm x * Kf" for x by blast from g.bounded obtain Kg where Kg: "norm (g x) \ norm x * Kg" for x by blast show "\K. \x. norm (f x + g x) \ norm x * K" using add_mono[OF Kf Kg] by (intro exI[of _ "Kf + Kg"]) (auto simp: field_simps intro: norm_triangle_ineq order_trans) qed (simp_all add: f.add g.add f.scaleR g.scaleR scaleR_right_distrib) qed lemma bounded_linear_minus: assumes "bounded_linear f" shows "bounded_linear (\x. - f x)" proof - interpret f: bounded_linear f by fact show ?thesis by unfold_locales (simp_all add: f.add f.scaleR f.bounded) qed lemma bounded_linear_sub: "bounded_linear f \ bounded_linear g \ bounded_linear (\x. f x - g x)" using bounded_linear_add[of f "\x. - g x"] bounded_linear_minus[of g] by (auto simp: algebra_simps) lemma bounded_linear_sum: fixes f :: "'i \ 'a::real_normed_vector \ 'b::real_normed_vector" shows "(\i. i \ I \ bounded_linear (f i)) \ bounded_linear (\x. \i\I. f i x)" by (induct I rule: infinite_finite_induct) (auto intro!: bounded_linear_add) lemma bounded_linear_compose: assumes "bounded_linear f" and "bounded_linear g" shows "bounded_linear (\x. f (g x))" proof - interpret f: bounded_linear f by fact interpret g: bounded_linear g by fact show ?thesis proof unfold_locales show "f (g (x + y)) = f (g x) + f (g y)" for x y by (simp only: f.add g.add) show "f (g (scaleR r x)) = scaleR r (f (g x))" for r x by (simp only: f.scaleR g.scaleR) from f.pos_bounded obtain Kf where f: "\x. norm (f x) \ norm x * Kf" and Kf: "0 < Kf" by blast from g.pos_bounded obtain Kg where g: "\x. norm (g x) \ norm x * Kg" by blast show "\K. \x. norm (f (g x)) \ norm x * K" proof (intro exI allI) fix x have "norm (f (g x)) \ norm (g x) * Kf" using f . also have "\ \ (norm x * Kg) * Kf" using g Kf [THEN order_less_imp_le] by (rule mult_right_mono) also have "(norm x * Kg) * Kf = norm x * (Kg * Kf)" by (rule mult.assoc) finally show "norm (f (g x)) \ norm x * (Kg * Kf)" . qed qed qed lemma bounded_bilinear_mult: "bounded_bilinear ((*) :: 'a \ 'a \ 'a::real_normed_algebra)" proof (rule bounded_bilinear.intro) show "\K. \a b::'a. norm (a * b) \ norm a * norm b * K" by (rule_tac x=1 in exI) (simp add: norm_mult_ineq) qed (auto simp: algebra_simps) lemma bounded_linear_mult_left: "bounded_linear (\x::'a::real_normed_algebra. x * y)" using bounded_bilinear_mult by (rule bounded_bilinear.bounded_linear_left) lemma bounded_linear_mult_right: "bounded_linear (\y::'a::real_normed_algebra. x * y)" using bounded_bilinear_mult by (rule bounded_bilinear.bounded_linear_right) lemmas bounded_linear_mult_const = bounded_linear_mult_left [THEN bounded_linear_compose] lemmas bounded_linear_const_mult = bounded_linear_mult_right [THEN bounded_linear_compose] lemma bounded_linear_divide: "bounded_linear (\x. x / y)" for y :: "'a::real_normed_field" unfolding divide_inverse by (rule bounded_linear_mult_left) lemma bounded_bilinear_scaleR: "bounded_bilinear scaleR" proof (rule bounded_bilinear.intro) show "\K. \a b. norm (a *\<^sub>R b) \ norm a * norm b * K" using less_eq_real_def by auto qed (auto simp: algebra_simps) lemma bounded_linear_scaleR_left: "bounded_linear (\r. scaleR r x)" using bounded_bilinear_scaleR by (rule bounded_bilinear.bounded_linear_left) lemma bounded_linear_scaleR_right: "bounded_linear (\x. scaleR r x)" using bounded_bilinear_scaleR by (rule bounded_bilinear.bounded_linear_right) lemmas bounded_linear_scaleR_const = bounded_linear_scaleR_left[THEN bounded_linear_compose] lemmas bounded_linear_const_scaleR = bounded_linear_scaleR_right[THEN bounded_linear_compose] lemma bounded_linear_of_real: "bounded_linear (\r. of_real r)" unfolding of_real_def by (rule bounded_linear_scaleR_left) lemma real_bounded_linear: "bounded_linear f \ (\c::real. f = (\x. x * c))" for f :: "real \ real" proof - { fix x assume "bounded_linear f" then interpret bounded_linear f . from scaleR[of x 1] have "f x = x * f 1" by simp } then show ?thesis by (auto intro: exI[of _ "f 1"] bounded_linear_mult_left) qed instance real_normed_algebra_1 \ perfect_space proof fix x::'a have "\e. 0 < e \ \y. norm (y - x) < e \ y \ x" by (rule_tac x = "x + of_real (e/2)" in exI) auto then show "\ open {x}" by (clarsimp simp: open_dist dist_norm) qed subsection \Filters and Limits on Metric Space\ lemma (in metric_space) nhds_metric: "nhds x = (INF e\{0 <..}. principal {y. dist y x < e})" unfolding nhds_def proof (safe intro!: INF_eq) fix S assume "open S" "x \ S" then obtain e where "{y. dist y x < e} \ S" "0 < e" by (auto simp: open_dist subset_eq) then show "\e\{0<..}. principal {y. dist y x < e} \ principal S" by auto qed (auto intro!: exI[of _ "{y. dist x y < e}" for e] open_ball simp: dist_commute) lemma (in metric_space) tendsto_iff: "(f \ l) F \ (\e>0. eventually (\x. dist (f x) l < e) F)" unfolding nhds_metric filterlim_INF filterlim_principal by auto lemma tendsto_dist_iff: "((f \ l) F) \ (((\x. dist (f x) l) \ 0) F)" unfolding tendsto_iff by simp lemma (in metric_space) tendstoI [intro?]: "(\e. 0 < e \ eventually (\x. dist (f x) l < e) F) \ (f \ l) F" by (auto simp: tendsto_iff) lemma (in metric_space) tendstoD: "(f \ l) F \ 0 < e \ eventually (\x. dist (f x) l < e) F" by (auto simp: tendsto_iff) lemma (in metric_space) eventually_nhds_metric: "eventually P (nhds a) \ (\d>0. \x. dist x a < d \ P x)" unfolding nhds_metric by (subst eventually_INF_base) (auto simp: eventually_principal Bex_def subset_eq intro: exI[of _ "min a b" for a b]) lemma eventually_at: "eventually P (at a within S) \ (\d>0. \x\S. x \ a \ dist x a < d \ P x)" for a :: "'a :: metric_space" by (auto simp: eventually_at_filter eventually_nhds_metric) lemma frequently_at: "frequently P (at a within S) \ (\d>0. \x\S. x \ a \ dist x a < d \ P x)" for a :: "'a :: metric_space" unfolding frequently_def eventually_at by auto lemma eventually_at_le: "eventually P (at a within S) \ (\d>0. \x\S. x \ a \ dist x a \ d \ P x)" for a :: "'a::metric_space" unfolding eventually_at_filter eventually_nhds_metric apply safe apply (rule_tac x="d / 2" in exI, auto) done lemma eventually_at_left_real: "a > (b :: real) \ eventually (\x. x \ {b<.. eventually (\x. x \ {a<.. a) F" and le: "eventually (\x. dist (g x) b \ dist (f x) a) F" shows "(g \ b) F" proof (rule tendstoI) fix e :: real assume "0 < e" with f have "eventually (\x. dist (f x) a < e) F" by (rule tendstoD) with le show "eventually (\x. dist (g x) b < e) F" using le_less_trans by (rule eventually_elim2) qed lemma filterlim_real_sequentially: "LIM x sequentially. real x :> at_top" proof (clarsimp simp: filterlim_at_top) fix Z show "\\<^sub>F x in sequentially. Z \ real x" by (meson eventually_sequentiallyI nat_ceiling_le_eq) qed lemma filterlim_nat_sequentially: "filterlim nat sequentially at_top" proof - have "\\<^sub>F x in at_top. Z \ nat x" for Z by (auto intro!: eventually_at_top_linorderI[where c="int Z"]) then show ?thesis unfolding filterlim_at_top .. qed lemma filterlim_floor_sequentially: "filterlim floor at_top at_top" proof - have "\\<^sub>F x in at_top. Z \ \x\" for Z by (auto simp: le_floor_iff intro!: eventually_at_top_linorderI[where c="of_int Z"]) then show ?thesis unfolding filterlim_at_top .. qed lemma filterlim_sequentially_iff_filterlim_real: "filterlim f sequentially F \ filterlim (\x. real (f x)) at_top F" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using filterlim_compose filterlim_real_sequentially by blast next assume R: ?rhs show ?lhs proof - have "filterlim (\x. nat (floor (real (f x)))) sequentially F" by (intro filterlim_compose[OF filterlim_nat_sequentially] filterlim_compose[OF filterlim_floor_sequentially] R) then show ?thesis by simp qed qed subsubsection \Limits of Sequences\ lemma lim_sequentially: "X \ L \ (\r>0. \no. \n\no. dist (X n) L < r)" for L :: "'a::metric_space" unfolding tendsto_iff eventually_sequentially .. lemmas LIMSEQ_def = lim_sequentially (*legacy binding*) lemma LIMSEQ_iff_nz: "X \ L \ (\r>0. \no>0. \n\no. dist (X n) L < r)" for L :: "'a::metric_space" unfolding lim_sequentially by (metis Suc_leD zero_less_Suc) lemma metric_LIMSEQ_I: "(\r. 0 < r \ \no. \n\no. dist (X n) L < r) \ X \ L" for L :: "'a::metric_space" by (simp add: lim_sequentially) lemma metric_LIMSEQ_D: "X \ L \ 0 < r \ \no. \n\no. dist (X n) L < r" for L :: "'a::metric_space" by (simp add: lim_sequentially) lemma LIMSEQ_norm_0: assumes "\n::nat. norm (f n) < 1 / real (Suc n)" shows "f \ 0" proof (rule metric_LIMSEQ_I) fix \ :: "real" assume "\ > 0" then obtain N::nat where "\ > inverse N" "N > 0" by (metis neq0_conv real_arch_inverse) then have "norm (f n) < \" if "n \ N" for n proof - have "1 / (Suc n) \ 1 / N" using \0 < N\ inverse_of_nat_le le_SucI that by blast also have "\ < \" by (metis (no_types) \inverse (real N) < \\ inverse_eq_divide) finally show ?thesis by (meson assms less_eq_real_def not_le order_trans) qed then show "\no. \n\no. dist (f n) 0 < \" by auto qed subsubsection \Limits of Functions\ lemma LIM_def: "f \a\ L \ (\r > 0. \s > 0. \x. x \ a \ dist x a < s \ dist (f x) L < r)" for a :: "'a::metric_space" and L :: "'b::metric_space" unfolding tendsto_iff eventually_at by simp lemma metric_LIM_I: "(\r. 0 < r \ \s>0. \x. x \ a \ dist x a < s \ dist (f x) L < r) \ f \a\ L" for a :: "'a::metric_space" and L :: "'b::metric_space" by (simp add: LIM_def) lemma metric_LIM_D: "f \a\ L \ 0 < r \ \s>0. \x. x \ a \ dist x a < s \ dist (f x) L < r" for a :: "'a::metric_space" and L :: "'b::metric_space" by (simp add: LIM_def) lemma metric_LIM_imp_LIM: fixes l :: "'a::metric_space" and m :: "'b::metric_space" assumes f: "f \a\ l" and le: "\x. x \ a \ dist (g x) m \ dist (f x) l" shows "g \a\ m" by (rule metric_tendsto_imp_tendsto [OF f]) (auto simp: eventually_at_topological le) lemma metric_LIM_equal2: fixes a :: "'a::metric_space" assumes "g \a\ l" "0 < R" and "\x. x \ a \ dist x a < R \ f x = g x" shows "f \a\ l" proof - have "\S. \open S; l \ S; \\<^sub>F x in at a. g x \ S\ \ \\<^sub>F x in at a. f x \ S" apply (simp add: eventually_at) by (metis assms(2) assms(3) dual_order.strict_trans linorder_neqE_linordered_idom) then show ?thesis using assms by (simp add: tendsto_def) qed lemma metric_LIM_compose2: fixes a :: "'a::metric_space" assumes f: "f \a\ b" and g: "g \b\ c" and inj: "\d>0. \x. x \ a \ dist x a < d \ f x \ b" shows "(\x. g (f x)) \a\ c" using inj by (intro tendsto_compose_eventually[OF g f]) (auto simp: eventually_at) lemma metric_isCont_LIM_compose2: fixes f :: "'a :: metric_space \ _" assumes f [unfolded isCont_def]: "isCont f a" and g: "g \f a\ l" and inj: "\d>0. \x. x \ a \ dist x a < d \ f x \ f a" shows "(\x. g (f x)) \a\ l" by (rule metric_LIM_compose2 [OF f g inj]) subsection \Complete metric spaces\ subsection \Cauchy sequences\ lemma (in metric_space) Cauchy_def: "Cauchy X = (\e>0. \M. \m\M. \n\M. dist (X m) (X n) < e)" proof - have *: "eventually P (INF M. principal {(X m, X n) | n m. m \ M \ n \ M}) \ (\M. \m\M. \n\M. P (X m, X n))" for P apply (subst eventually_INF_base) subgoal by simp subgoal for a b by (intro bexI[of _ "max a b"]) (auto simp: eventually_principal subset_eq) subgoal by (auto simp: eventually_principal, blast) done have "Cauchy X \ (INF M. principal {(X m, X n) | n m. m \ M \ n \ M}) \ uniformity" unfolding Cauchy_uniform_iff le_filter_def * .. also have "\ = (\e>0. \M. \m\M. \n\M. dist (X m) (X n) < e)" unfolding uniformity_dist le_INF_iff by (auto simp: * le_principal) finally show ?thesis . qed lemma (in metric_space) Cauchy_altdef: "Cauchy f \ (\e>0. \M. \m\M. \n>m. dist (f m) (f n) < e)" (is "?lhs \ ?rhs") proof assume ?rhs show ?lhs unfolding Cauchy_def proof (intro allI impI) fix e :: real assume e: "e > 0" with \?rhs\ obtain M where M: "m \ M \ n > m \ dist (f m) (f n) < e" for m n by blast have "dist (f m) (f n) < e" if "m \ M" "n \ M" for m n using M[of m n] M[of n m] e that by (cases m n rule: linorder_cases) (auto simp: dist_commute) then show "\M. \m\M. \n\M. dist (f m) (f n) < e" by blast qed next assume ?lhs show ?rhs proof (intro allI impI) fix e :: real assume e: "e > 0" with \Cauchy f\ obtain M where "\m n. m \ M \ n \ M \ dist (f m) (f n) < e" unfolding Cauchy_def by blast then show "\M. \m\M. \n>m. dist (f m) (f n) < e" by (intro exI[of _ M]) force qed qed lemma (in metric_space) Cauchy_altdef2: "Cauchy s \ (\e>0. \N::nat. \n\N. dist(s n)(s N) < e)" (is "?lhs = ?rhs") proof assume "Cauchy s" then show ?rhs by (force simp: Cauchy_def) next assume ?rhs { fix e::real assume "e>0" with \?rhs\ obtain N where N: "\n\N. dist (s n) (s N) < e/2" by (erule_tac x="e/2" in allE) auto { fix n m assume nm: "N \ m \ N \ n" then have "dist (s m) (s n) < e" using N using dist_triangle_half_l[of "s m" "s N" "e" "s n"] by blast } then have "\N. \m n. N \ m \ N \ n \ dist (s m) (s n) < e" by blast } then have ?lhs unfolding Cauchy_def by blast then show ?lhs by blast qed lemma (in metric_space) metric_CauchyI: "(\e. 0 < e \ \M. \m\M. \n\M. dist (X m) (X n) < e) \ Cauchy X" by (simp add: Cauchy_def) lemma (in metric_space) CauchyI': "(\e. 0 < e \ \M. \m\M. \n>m. dist (X m) (X n) < e) \ Cauchy X" unfolding Cauchy_altdef by blast lemma (in metric_space) metric_CauchyD: "Cauchy X \ 0 < e \ \M. \m\M. \n\M. dist (X m) (X n) < e" by (simp add: Cauchy_def) lemma (in metric_space) metric_Cauchy_iff2: "Cauchy X = (\j. (\M. \m \ M. \n \ M. dist (X m) (X n) < inverse(real (Suc j))))" apply (auto simp add: Cauchy_def) by (metis less_trans of_nat_Suc reals_Archimedean) lemma Cauchy_iff2: "Cauchy X \ (\j. (\M. \m \ M. \n \ M. \X m - X n\ < inverse (real (Suc j))))" by (simp only: metric_Cauchy_iff2 dist_real_def) lemma lim_1_over_n [tendsto_intros]: "((\n. 1 / of_nat n) \ (0::'a::real_normed_field)) sequentially" proof (subst lim_sequentially, intro allI impI exI) fix e::real and n assume e: "e > 0" have "inverse e < of_nat (nat \inverse e + 1\)" by linarith also assume "n \ nat \inverse e + 1\" finally show "dist (1 / of_nat n :: 'a) 0 < e" using e by (simp add: field_split_simps norm_divide) qed lemma (in metric_space) complete_def: shows "complete S = (\f. (\n. f n \ S) \ Cauchy f \ (\l\S. f \ l))" unfolding complete_uniform proof safe fix f :: "nat \ 'a" assume f: "\n. f n \ S" "Cauchy f" and *: "\F\principal S. F \ bot \ cauchy_filter F \ (\x\S. F \ nhds x)" then show "\l\S. f \ l" unfolding filterlim_def using f by (intro *[rule_format]) (auto simp: filtermap_sequentually_ne_bot le_principal eventually_filtermap Cauchy_uniform) next fix F :: "'a filter" assume "F \ principal S" "F \ bot" "cauchy_filter F" assume seq: "\f. (\n. f n \ S) \ Cauchy f \ (\l\S. f \ l)" from \F \ principal S\ \cauchy_filter F\ have FF_le: "F \\<^sub>F F \ uniformity_on S" by (simp add: cauchy_filter_def principal_prod_principal[symmetric] prod_filter_mono) let ?P = "\P e. eventually P F \ (\x. P x \ x \ S) \ (\x y. P x \ P y \ dist x y < e)" have P: "\P. ?P P \" if "0 < \" for \ :: real proof - from that have "eventually (\(x, y). x \ S \ y \ S \ dist x y < \) (uniformity_on S)" by (auto simp: eventually_inf_principal eventually_uniformity_metric) from filter_leD[OF FF_le this] show ?thesis by (auto simp: eventually_prod_same) qed have "\P. \n. ?P (P n) (1 / Suc n) \ P (Suc n) \ P n" proof (rule dependent_nat_choice) show "\P. ?P P (1 / Suc 0)" using P[of 1] by auto next fix P n assume "?P P (1/Suc n)" moreover obtain Q where "?P Q (1 / Suc (Suc n))" using P[of "1/Suc (Suc n)"] by auto ultimately show "\Q. ?P Q (1 / Suc (Suc n)) \ Q \ P" by (intro exI[of _ "\x. P x \ Q x"]) (auto simp: eventually_conj_iff) qed then obtain P where P: "eventually (P n) F" "P n x \ x \ S" "P n x \ P n y \ dist x y < 1 / Suc n" "P (Suc n) \ P n" for n x y by metis have "antimono P" using P(4) unfolding decseq_Suc_iff le_fun_def by blast obtain X where X: "P n (X n)" for n using P(1)[THEN eventually_happens'[OF \F \ bot\]] by metis have "Cauchy X" unfolding metric_Cauchy_iff2 inverse_eq_divide proof (intro exI allI impI) fix j m n :: nat assume "j \ m" "j \ n" with \antimono P\ X have "P j (X m)" "P j (X n)" by (auto simp: antimono_def) then show "dist (X m) (X n) < 1 / Suc j" by (rule P) qed moreover have "\n. X n \ S" using P(2) X by auto ultimately obtain x where "X \ x" "x \ S" using seq by blast show "\x\S. F \ nhds x" proof (rule bexI) have "eventually (\y. dist y x < e) F" if "0 < e" for e :: real proof - from that have "(\n. 1 / Suc n :: real) \ 0 \ 0 < e / 2" by (subst filterlim_sequentially_Suc) (auto intro!: lim_1_over_n) then have "\\<^sub>F n in sequentially. dist (X n) x < e / 2 \ 1 / Suc n < e / 2" using \X \ x\ unfolding tendsto_iff order_tendsto_iff[where 'a=real] eventually_conj_iff by blast then obtain n where "dist x (X n) < e / 2" "1 / Suc n < e / 2" by (auto simp: eventually_sequentially dist_commute) show ?thesis using \eventually (P n) F\ proof eventually_elim case (elim y) then have "dist y (X n) < 1 / Suc n" by (intro X P) also have "\ < e / 2" by fact finally show "dist y x < e" by (rule dist_triangle_half_l) fact qed qed then show "F \ nhds x" unfolding nhds_metric le_INF_iff le_principal by auto qed fact qed text\apparently unused\ lemma (in metric_space) totally_bounded_metric: "totally_bounded S \ (\e>0. \k. finite k \ S \ (\x\k. {y. dist x y < e}))" unfolding totally_bounded_def eventually_uniformity_metric imp_ex apply (subst all_comm) apply (intro arg_cong[where f=All] ext, safe) subgoal for e apply (erule allE[of _ "\(x, y). dist x y < e"]) apply auto done subgoal for e P k apply (intro exI[of _ k]) apply (force simp: subset_eq) done done subsubsection \Cauchy Sequences are Convergent\ (* TODO: update to uniform_space *) class complete_space = metric_space + assumes Cauchy_convergent: "Cauchy X \ convergent X" lemma Cauchy_convergent_iff: "Cauchy X \ convergent X" for X :: "nat \ 'a::complete_space" by (blast intro: Cauchy_convergent convergent_Cauchy) text \To prove that a Cauchy sequence converges, it suffices to show that a subsequence converges.\ lemma Cauchy_converges_subseq: fixes u::"nat \ 'a::metric_space" assumes "Cauchy u" "strict_mono r" "(u \ r) \ l" shows "u \ l" proof - have *: "eventually (\n. dist (u n) l < e) sequentially" if "e > 0" for e proof - have "e/2 > 0" using that by auto then obtain N1 where N1: "\m n. m \ N1 \ n \ N1 \ dist (u m) (u n) < e/2" using \Cauchy u\ unfolding Cauchy_def by blast obtain N2 where N2: "\n. n \ N2 \ dist ((u \ r) n) l < e / 2" using order_tendstoD(2)[OF iffD1[OF tendsto_dist_iff \(u \ r) \ l\] \e/2 > 0\] unfolding eventually_sequentially by auto have "dist (u n) l < e" if "n \ max N1 N2" for n proof - have "dist (u n) l \ dist (u n) ((u \ r) n) + dist ((u \ r) n) l" by (rule dist_triangle) also have "\ < e/2 + e/2" proof (intro add_strict_mono) show "dist (u n) ((u \ r) n) < e / 2" using N1[of n "r n"] N2[of n] that unfolding comp_def by (meson assms(2) le_trans max.bounded_iff strict_mono_imp_increasing) show "dist ((u \ r) n) l < e / 2" using N2 that by auto qed finally show ?thesis by simp qed then show ?thesis unfolding eventually_sequentially by blast qed have "(\n. dist (u n) l) \ 0" by (simp add: less_le_trans * order_tendstoI) then show ?thesis using tendsto_dist_iff by auto qed subsection \The set of real numbers is a complete metric space\ text \ Proof that Cauchy sequences converge based on the one from \<^url>\http://pirate.shu.edu/~wachsmut/ira/numseq/proofs/cauconv.html\ \ text \ If sequence \<^term>\X\ is Cauchy, then its limit is the lub of \<^term>\{r::real. \N. \n\N. r < X n}\ \ lemma increasing_LIMSEQ: fixes f :: "nat \ real" assumes inc: "\n. f n \ f (Suc n)" and bdd: "\n. f n \ l" and en: "\e. 0 < e \ \n. l \ f n + e" shows "f \ l" proof (rule increasing_tendsto) fix x assume "x < l" with dense[of 0 "l - x"] obtain e where "0 < e" "e < l - x" by auto from en[OF \0 < e\] obtain n where "l - e \ f n" by (auto simp: field_simps) with \e < l - x\ \0 < e\ have "x < f n" by simp with incseq_SucI[of f, OF inc] show "eventually (\n. x < f n) sequentially" by (auto simp: eventually_sequentially incseq_def intro: less_le_trans) qed (use bdd in auto) lemma real_Cauchy_convergent: fixes X :: "nat \ real" assumes X: "Cauchy X" shows "convergent X" proof - define S :: "real set" where "S = {x. \N. \n\N. x < X n}" then have mem_S: "\N x. \n\N. x < X n \ x \ S" by auto have bound_isUb: "y \ x" if N: "\n\N. X n < x" and "y \ S" for N and x y :: real proof - from that have "\M. \n\M. y < X n" by (simp add: S_def) then obtain M where "\n\M. y < X n" .. then have "y < X (max M N)" by simp also have "\ < x" using N by simp finally show ?thesis by (rule order_less_imp_le) qed obtain N where "\m\N. \n\N. dist (X m) (X n) < 1" using X[THEN metric_CauchyD, OF zero_less_one] by auto then have N: "\n\N. dist (X n) (X N) < 1" by simp have [simp]: "S \ {}" proof (intro exI ex_in_conv[THEN iffD1]) from N have "\n\N. X N - 1 < X n" by (simp add: abs_diff_less_iff dist_real_def) then show "X N - 1 \ S" by (rule mem_S) qed have [simp]: "bdd_above S" proof from N have "\n\N. X n < X N + 1" by (simp add: abs_diff_less_iff dist_real_def) then show "\s. s \ S \ s \ X N + 1" by (rule bound_isUb) qed have "X \ Sup S" proof (rule metric_LIMSEQ_I) fix r :: real assume "0 < r" then have r: "0 < r/2" by simp obtain N where "\n\N. \m\N. dist (X n) (X m) < r/2" using metric_CauchyD [OF X r] by auto then have "\n\N. dist (X n) (X N) < r/2" by simp then have N: "\n\N. X N - r/2 < X n \ X n < X N + r/2" by (simp only: dist_real_def abs_diff_less_iff) from N have "\n\N. X N - r/2 < X n" by blast then have "X N - r/2 \ S" by (rule mem_S) then have 1: "X N - r/2 \ Sup S" by (simp add: cSup_upper) from N have "\n\N. X n < X N + r/2" by blast from bound_isUb[OF this] have 2: "Sup S \ X N + r/2" by (intro cSup_least) simp_all show "\N. \n\N. dist (X n) (Sup S) < r" proof (intro exI allI impI) fix n assume n: "N \ n" from N n have "X n < X N + r/2" and "X N - r/2 < X n" by simp_all then show "dist (X n) (Sup S) < r" using 1 2 by (simp add: abs_diff_less_iff dist_real_def) qed qed then show ?thesis by (auto simp: convergent_def) qed instance real :: complete_space by intro_classes (rule real_Cauchy_convergent) class banach = real_normed_vector + complete_space instance real :: banach .. lemma tendsto_at_topI_sequentially: fixes f :: "real \ 'b::first_countable_topology" assumes *: "\X. filterlim X at_top sequentially \ (\n. f (X n)) \ y" shows "(f \ y) at_top" proof - obtain A where A: "decseq A" "open (A n)" "y \ A n" "nhds y = (INF n. principal (A n))" for n by (rule nhds_countable[of y]) (rule that) have "\m. \k. \x\k. f x \ A m" proof (rule ccontr) assume "\ (\m. \k. \x\k. f x \ A m)" then obtain m where "\k. \x\k. f x \ A m" by auto then have "\X. \n. (f (X n) \ A m) \ max n (X n) + 1 \ X (Suc n)" by (intro dependent_nat_choice) (auto simp del: max.bounded_iff) then obtain X where X: "\n. f (X n) \ A m" "\n. max n (X n) + 1 \ X (Suc n)" by auto have "1 \ n \ real n \ X n" for n using X[of "n - 1"] by auto then have "filterlim X at_top sequentially" by (force intro!: filterlim_at_top_mono[OF filterlim_real_sequentially] simp: eventually_sequentially) from topological_tendstoD[OF *[OF this] A(2, 3), of m] X(1) show False by auto qed then obtain k where "k m \ x \ f x \ A m" for m x by metis then show ?thesis unfolding at_top_def A by (intro filterlim_base[where i=k]) auto qed lemma tendsto_at_topI_sequentially_real: fixes f :: "real \ real" assumes mono: "mono f" and limseq: "(\n. f (real n)) \ y" shows "(f \ y) at_top" proof (rule tendstoI) fix e :: real assume "0 < e" with limseq obtain N :: nat where N: "N \ n \ \f (real n) - y\ < e" for n by (auto simp: lim_sequentially dist_real_def) have le: "f x \ y" for x :: real proof - obtain n where "x \ real_of_nat n" using real_arch_simple[of x] .. note monoD[OF mono this] also have "f (real_of_nat n) \ y" by (rule LIMSEQ_le_const[OF limseq]) (auto intro!: exI[of _ n] monoD[OF mono]) finally show ?thesis . qed have "eventually (\x. real N \ x) at_top" by (rule eventually_ge_at_top) then show "eventually (\x. dist (f x) y < e) at_top" proof eventually_elim case (elim x) with N[of N] le have "y - f (real N) < e" by auto moreover note monoD[OF mono elim] ultimately show "dist (f x) y < e" using le[of x] by (auto simp: dist_real_def field_simps) qed qed end diff --git a/src/HOL/Rings.thy b/src/HOL/Rings.thy --- a/src/HOL/Rings.thy +++ b/src/HOL/Rings.thy @@ -1,2745 +1,2750 @@ (* Title: HOL/Rings.thy Author: Gertrud Bauer Author: Steven Obua Author: Tobias Nipkow Author: Lawrence C Paulson Author: Markus Wenzel Author: Jeremy Avigad *) section \Rings\ theory Rings imports Groups Set Fun begin subsection \Semirings and rings\ class semiring = ab_semigroup_add + semigroup_mult + assumes distrib_right [algebra_simps, algebra_split_simps]: "(a + b) * c = a * c + b * c" assumes distrib_left [algebra_simps, algebra_split_simps]: "a * (b + c) = a * b + a * c" begin text \For the \combine_numerals\ simproc\ lemma combine_common_factor: "a * e + (b * e + c) = (a + b) * e + c" by (simp add: distrib_right ac_simps) end class mult_zero = times + zero + assumes mult_zero_left [simp]: "0 * a = 0" assumes mult_zero_right [simp]: "a * 0 = 0" begin lemma mult_not_zero: "a * b \ 0 \ a \ 0 \ b \ 0" by auto end class semiring_0 = semiring + comm_monoid_add + mult_zero class semiring_0_cancel = semiring + cancel_comm_monoid_add begin subclass semiring_0 proof fix a :: 'a have "0 * a + 0 * a = 0 * a + 0" by (simp add: distrib_right [symmetric]) then show "0 * a = 0" by (simp only: add_left_cancel) have "a * 0 + a * 0 = a * 0 + 0" by (simp add: distrib_left [symmetric]) then show "a * 0 = 0" by (simp only: add_left_cancel) qed end class comm_semiring = ab_semigroup_add + ab_semigroup_mult + assumes distrib: "(a + b) * c = a * c + b * c" begin subclass semiring proof fix a b c :: 'a show "(a + b) * c = a * c + b * c" by (simp add: distrib) have "a * (b + c) = (b + c) * a" by (simp add: ac_simps) also have "\ = b * a + c * a" by (simp only: distrib) also have "\ = a * b + a * c" by (simp add: ac_simps) finally show "a * (b + c) = a * b + a * c" by blast qed end class comm_semiring_0 = comm_semiring + comm_monoid_add + mult_zero begin subclass semiring_0 .. end class comm_semiring_0_cancel = comm_semiring + cancel_comm_monoid_add begin subclass semiring_0_cancel .. subclass comm_semiring_0 .. end class zero_neq_one = zero + one + assumes zero_neq_one [simp]: "0 \ 1" begin lemma one_neq_zero [simp]: "1 \ 0" by (rule not_sym) (rule zero_neq_one) definition of_bool :: "bool \ 'a" where "of_bool p = (if p then 1 else 0)" lemma of_bool_eq [simp, code]: "of_bool False = 0" "of_bool True = 1" by (simp_all add: of_bool_def) lemma of_bool_eq_iff: "of_bool p = of_bool q \ p = q" by (simp add: of_bool_def) lemma split_of_bool [split]: "P (of_bool p) \ (p \ P 1) \ (\ p \ P 0)" by (cases p) simp_all lemma split_of_bool_asm [split]: "P (of_bool p) \ \ (p \ \ P 1 \ \ p \ \ P 0)" by (cases p) simp_all lemma of_bool_eq_0_iff [simp]: \of_bool P = 0 \ \ P\ by simp lemma of_bool_eq_1_iff [simp]: \of_bool P = 1 \ P\ by simp end class semiring_1 = zero_neq_one + semiring_0 + monoid_mult begin lemma of_bool_conj: "of_bool (P \ Q) = of_bool P * of_bool Q" by auto end lemma lambda_zero: "(\h::'a::mult_zero. 0) = (*) 0" by auto lemma lambda_one: "(\x::'a::monoid_mult. x) = (*) 1" by auto subsection \Abstract divisibility\ class dvd = times begin definition dvd :: "'a \ 'a \ bool" (infix "dvd" 50) where "b dvd a \ (\k. a = b * k)" lemma dvdI [intro?]: "a = b * k \ b dvd a" unfolding dvd_def .. lemma dvdE [elim]: "b dvd a \ (\k. a = b * k \ P) \ P" unfolding dvd_def by blast end context comm_monoid_mult begin subclass dvd . lemma dvd_refl [simp]: "a dvd a" proof show "a = a * 1" by simp qed lemma dvd_trans [trans]: assumes "a dvd b" and "b dvd c" shows "a dvd c" proof - from assms obtain v where "b = a * v" by auto moreover from assms obtain w where "c = b * w" by auto ultimately have "c = a * (v * w)" by (simp add: mult.assoc) then show ?thesis .. qed lemma subset_divisors_dvd: "{c. c dvd a} \ {c. c dvd b} \ a dvd b" by (auto simp add: subset_iff intro: dvd_trans) lemma strict_subset_divisors_dvd: "{c. c dvd a} \ {c. c dvd b} \ a dvd b \ \ b dvd a" by (auto simp add: subset_iff intro: dvd_trans) lemma one_dvd [simp]: "1 dvd a" by (auto intro: dvdI) lemma dvd_mult [simp]: "a dvd (b * c)" if "a dvd c" using that by rule (auto intro: mult.left_commute dvdI) lemma dvd_mult2 [simp]: "a dvd (b * c)" if "a dvd b" using that dvd_mult [of a b c] by (simp add: ac_simps) lemma dvd_triv_right [simp]: "a dvd b * a" by (rule dvd_mult) (rule dvd_refl) lemma dvd_triv_left [simp]: "a dvd a * b" by (rule dvd_mult2) (rule dvd_refl) lemma mult_dvd_mono: assumes "a dvd b" and "c dvd d" shows "a * c dvd b * d" proof - from \a dvd b\ obtain b' where "b = a * b'" .. moreover from \c dvd d\ obtain d' where "d = c * d'" .. ultimately have "b * d = (a * c) * (b' * d')" by (simp add: ac_simps) then show ?thesis .. qed lemma dvd_mult_left: "a * b dvd c \ a dvd c" by (simp add: dvd_def mult.assoc) blast lemma dvd_mult_right: "a * b dvd c \ b dvd c" using dvd_mult_left [of b a c] by (simp add: ac_simps) end class comm_semiring_1 = zero_neq_one + comm_semiring_0 + comm_monoid_mult begin subclass semiring_1 .. lemma dvd_0_left_iff [simp]: "0 dvd a \ a = 0" by auto lemma dvd_0_right [iff]: "a dvd 0" proof show "0 = a * 0" by simp qed lemma dvd_0_left: "0 dvd a \ a = 0" by simp lemma dvd_add [simp]: assumes "a dvd b" and "a dvd c" shows "a dvd (b + c)" proof - from \a dvd b\ obtain b' where "b = a * b'" .. moreover from \a dvd c\ obtain c' where "c = a * c'" .. ultimately have "b + c = a * (b' + c')" by (simp add: distrib_left) then show ?thesis .. qed end class semiring_1_cancel = semiring + cancel_comm_monoid_add + zero_neq_one + monoid_mult begin subclass semiring_0_cancel .. subclass semiring_1 .. end class comm_semiring_1_cancel = comm_semiring + cancel_comm_monoid_add + zero_neq_one + comm_monoid_mult + assumes right_diff_distrib' [algebra_simps, algebra_split_simps]: "a * (b - c) = a * b - a * c" begin subclass semiring_1_cancel .. subclass comm_semiring_0_cancel .. subclass comm_semiring_1 .. lemma left_diff_distrib' [algebra_simps, algebra_split_simps]: "(b - c) * a = b * a - c * a" by (simp add: algebra_simps) lemma dvd_add_times_triv_left_iff [simp]: "a dvd c * a + b \ a dvd b" proof - have "a dvd a * c + b \ a dvd b" (is "?P \ ?Q") proof assume ?Q then show ?P by simp next assume ?P then obtain d where "a * c + b = a * d" .. then have "a * c + b - a * c = a * d - a * c" by simp then have "b = a * d - a * c" by simp then have "b = a * (d - c)" by (simp add: algebra_simps) then show ?Q .. qed then show "a dvd c * a + b \ a dvd b" by (simp add: ac_simps) qed lemma dvd_add_times_triv_right_iff [simp]: "a dvd b + c * a \ a dvd b" using dvd_add_times_triv_left_iff [of a c b] by (simp add: ac_simps) lemma dvd_add_triv_left_iff [simp]: "a dvd a + b \ a dvd b" using dvd_add_times_triv_left_iff [of a 1 b] by simp lemma dvd_add_triv_right_iff [simp]: "a dvd b + a \ a dvd b" using dvd_add_times_triv_right_iff [of a b 1] by simp lemma dvd_add_right_iff: assumes "a dvd b" shows "a dvd b + c \ a dvd c" (is "?P \ ?Q") proof assume ?P then obtain d where "b + c = a * d" .. moreover from \a dvd b\ obtain e where "b = a * e" .. ultimately have "a * e + c = a * d" by simp then have "a * e + c - a * e = a * d - a * e" by simp then have "c = a * d - a * e" by simp then have "c = a * (d - e)" by (simp add: algebra_simps) then show ?Q .. next assume ?Q with assms show ?P by simp qed lemma dvd_add_left_iff: "a dvd c \ a dvd b + c \ a dvd b" using dvd_add_right_iff [of a c b] by (simp add: ac_simps) end class ring = semiring + ab_group_add begin subclass semiring_0_cancel .. text \Distribution rules\ lemma minus_mult_left: "- (a * b) = - a * b" by (rule minus_unique) (simp add: distrib_right [symmetric]) lemma minus_mult_right: "- (a * b) = a * - b" by (rule minus_unique) (simp add: distrib_left [symmetric]) text \Extract signs from products\ lemmas mult_minus_left [simp] = minus_mult_left [symmetric] lemmas mult_minus_right [simp] = minus_mult_right [symmetric] lemma minus_mult_minus [simp]: "- a * - b = a * b" by simp lemma minus_mult_commute: "- a * b = a * - b" by simp lemma right_diff_distrib [algebra_simps, algebra_split_simps]: "a * (b - c) = a * b - a * c" using distrib_left [of a b "-c "] by simp lemma left_diff_distrib [algebra_simps, algebra_split_simps]: "(a - b) * c = a * c - b * c" using distrib_right [of a "- b" c] by simp lemmas ring_distribs = distrib_left distrib_right left_diff_distrib right_diff_distrib lemma eq_add_iff1: "a * e + c = b * e + d \ (a - b) * e + c = d" by (simp add: algebra_simps) lemma eq_add_iff2: "a * e + c = b * e + d \ c = (b - a) * e + d" by (simp add: algebra_simps) end lemmas ring_distribs = distrib_left distrib_right left_diff_distrib right_diff_distrib class comm_ring = comm_semiring + ab_group_add begin subclass ring .. subclass comm_semiring_0_cancel .. lemma square_diff_square_factored: "x * x - y * y = (x + y) * (x - y)" by (simp add: algebra_simps) end class ring_1 = ring + zero_neq_one + monoid_mult begin subclass semiring_1_cancel .. lemma of_bool_not_iff [simp]: \of_bool (\ P) = 1 - of_bool P\ by simp lemma square_diff_one_factored: "x * x - 1 = (x + 1) * (x - 1)" by (simp add: algebra_simps) end class comm_ring_1 = comm_ring + zero_neq_one + comm_monoid_mult begin subclass ring_1 .. subclass comm_semiring_1_cancel by standard (simp add: algebra_simps) lemma dvd_minus_iff [simp]: "x dvd - y \ x dvd y" proof assume "x dvd - y" then have "x dvd - 1 * - y" by (rule dvd_mult) then show "x dvd y" by simp next assume "x dvd y" then have "x dvd - 1 * y" by (rule dvd_mult) then show "x dvd - y" by simp qed lemma minus_dvd_iff [simp]: "- x dvd y \ x dvd y" proof assume "- x dvd y" then obtain k where "y = - x * k" .. then have "y = x * - k" by simp then show "x dvd y" .. next assume "x dvd y" then obtain k where "y = x * k" .. then have "y = - x * - k" by simp then show "- x dvd y" .. qed lemma dvd_diff [simp]: "x dvd y \ x dvd z \ x dvd (y - z)" using dvd_add [of x y "- z"] by simp end subsection \Towards integral domains\ class semiring_no_zero_divisors = semiring_0 + assumes no_zero_divisors: "a \ 0 \ b \ 0 \ a * b \ 0" begin lemma divisors_zero: assumes "a * b = 0" shows "a = 0 \ b = 0" proof (rule classical) assume "\ ?thesis" then have "a \ 0" and "b \ 0" by auto with no_zero_divisors have "a * b \ 0" by blast with assms show ?thesis by simp qed lemma mult_eq_0_iff [simp]: "a * b = 0 \ a = 0 \ b = 0" proof (cases "a = 0 \ b = 0") case False then have "a \ 0" and "b \ 0" by auto then show ?thesis using no_zero_divisors by simp next case True then show ?thesis by auto qed end class semiring_1_no_zero_divisors = semiring_1 + semiring_no_zero_divisors class semiring_no_zero_divisors_cancel = semiring_no_zero_divisors + assumes mult_cancel_right [simp]: "a * c = b * c \ c = 0 \ a = b" and mult_cancel_left [simp]: "c * a = c * b \ c = 0 \ a = b" begin lemma mult_left_cancel: "c \ 0 \ c * a = c * b \ a = b" by simp lemma mult_right_cancel: "c \ 0 \ a * c = b * c \ a = b" by simp end class ring_no_zero_divisors = ring + semiring_no_zero_divisors begin subclass semiring_no_zero_divisors_cancel proof fix a b c have "a * c = b * c \ (a - b) * c = 0" by (simp add: algebra_simps) also have "\ \ c = 0 \ a = b" by auto finally show "a * c = b * c \ c = 0 \ a = b" . have "c * a = c * b \ c * (a - b) = 0" by (simp add: algebra_simps) also have "\ \ c = 0 \ a = b" by auto finally show "c * a = c * b \ c = 0 \ a = b" . qed end class ring_1_no_zero_divisors = ring_1 + ring_no_zero_divisors begin subclass semiring_1_no_zero_divisors .. lemma square_eq_1_iff: "x * x = 1 \ x = 1 \ x = - 1" proof - have "(x - 1) * (x + 1) = x * x - 1" by (simp add: algebra_simps) then have "x * x = 1 \ (x - 1) * (x + 1) = 0" by simp then show ?thesis by (simp add: eq_neg_iff_add_eq_0) qed lemma mult_cancel_right1 [simp]: "c = b * c \ c = 0 \ b = 1" using mult_cancel_right [of 1 c b] by auto lemma mult_cancel_right2 [simp]: "a * c = c \ c = 0 \ a = 1" using mult_cancel_right [of a c 1] by simp lemma mult_cancel_left1 [simp]: "c = c * b \ c = 0 \ b = 1" using mult_cancel_left [of c 1 b] by force lemma mult_cancel_left2 [simp]: "c * a = c \ c = 0 \ a = 1" using mult_cancel_left [of c a 1] by simp end class semidom = comm_semiring_1_cancel + semiring_no_zero_divisors begin subclass semiring_1_no_zero_divisors .. end class idom = comm_ring_1 + semiring_no_zero_divisors begin subclass semidom .. subclass ring_1_no_zero_divisors .. lemma dvd_mult_cancel_right [simp]: "a * c dvd b * c \ c = 0 \ a dvd b" proof - have "a * c dvd b * c \ (\k. b * c = (a * k) * c)" by (auto simp add: ac_simps) also have "(\k. b * c = (a * k) * c) \ c = 0 \ a dvd b" by auto finally show ?thesis . qed lemma dvd_mult_cancel_left [simp]: "c * a dvd c * b \ c = 0 \ a dvd b" using dvd_mult_cancel_right [of a c b] by (simp add: ac_simps) lemma square_eq_iff: "a * a = b * b \ a = b \ a = - b" proof assume "a * a = b * b" then have "(a - b) * (a + b) = 0" by (simp add: algebra_simps) then show "a = b \ a = - b" by (simp add: eq_neg_iff_add_eq_0) next assume "a = b \ a = - b" then show "a * a = b * b" by auto qed end class idom_abs_sgn = idom + abs + sgn + assumes sgn_mult_abs: "sgn a * \a\ = a" and sgn_sgn [simp]: "sgn (sgn a) = sgn a" and abs_abs [simp]: "\\a\\ = \a\" and abs_0 [simp]: "\0\ = 0" and sgn_0 [simp]: "sgn 0 = 0" and sgn_1 [simp]: "sgn 1 = 1" and sgn_minus_1: "sgn (- 1) = - 1" and sgn_mult: "sgn (a * b) = sgn a * sgn b" begin lemma sgn_eq_0_iff: "sgn a = 0 \ a = 0" proof - { assume "sgn a = 0" then have "sgn a * \a\ = 0" by simp then have "a = 0" by (simp add: sgn_mult_abs) } then show ?thesis by auto qed lemma abs_eq_0_iff: "\a\ = 0 \ a = 0" proof - { assume "\a\ = 0" then have "sgn a * \a\ = 0" by simp then have "a = 0" by (simp add: sgn_mult_abs) } then show ?thesis by auto qed lemma abs_mult_sgn: "\a\ * sgn a = a" using sgn_mult_abs [of a] by (simp add: ac_simps) lemma abs_1 [simp]: "\1\ = 1" using sgn_mult_abs [of 1] by simp lemma sgn_abs [simp]: "\sgn a\ = of_bool (a \ 0)" using sgn_mult_abs [of "sgn a"] mult_cancel_left [of "sgn a" "\sgn a\" 1] by (auto simp add: sgn_eq_0_iff) lemma abs_sgn [simp]: "sgn \a\ = of_bool (a \ 0)" using sgn_mult_abs [of "\a\"] mult_cancel_right [of "sgn \a\" "\a\" 1] by (auto simp add: abs_eq_0_iff) lemma abs_mult: "\a * b\ = \a\ * \b\" proof (cases "a = 0 \ b = 0") case True then show ?thesis by auto next case False then have *: "sgn (a * b) \ 0" by (simp add: sgn_eq_0_iff) from abs_mult_sgn [of "a * b"] abs_mult_sgn [of a] abs_mult_sgn [of b] have "\a * b\ * sgn (a * b) = \a\ * sgn a * \b\ * sgn b" by (simp add: ac_simps) then have "\a * b\ * sgn (a * b) = \a\ * \b\ * sgn (a * b)" by (simp add: sgn_mult ac_simps) with * show ?thesis by simp qed lemma sgn_minus [simp]: "sgn (- a) = - sgn a" proof - from sgn_minus_1 have "sgn (- 1 * a) = - 1 * sgn a" by (simp only: sgn_mult) then show ?thesis by simp qed lemma abs_minus [simp]: "\- a\ = \a\" proof - have [simp]: "\- 1\ = 1" using sgn_mult_abs [of "- 1"] by simp then have "\- 1 * a\ = 1 * \a\" by (simp only: abs_mult) then show ?thesis by simp qed end subsection \(Partial) Division\ class divide = fixes divide :: "'a \ 'a \ 'a" (infixl "div" 70) setup \Sign.add_const_constraint (\<^const_name>\divide\, SOME \<^typ>\'a \ 'a \ 'a\)\ context semiring begin lemma [field_simps, field_split_simps]: shows distrib_left_NO_MATCH: "NO_MATCH (x div y) a \ a * (b + c) = a * b + a * c" and distrib_right_NO_MATCH: "NO_MATCH (x div y) c \ (a + b) * c = a * c + b * c" by (rule distrib_left distrib_right)+ end context ring begin lemma [field_simps, field_split_simps]: shows left_diff_distrib_NO_MATCH: "NO_MATCH (x div y) c \ (a - b) * c = a * c - b * c" and right_diff_distrib_NO_MATCH: "NO_MATCH (x div y) a \ a * (b - c) = a * b - a * c" by (rule left_diff_distrib right_diff_distrib)+ end setup \Sign.add_const_constraint (\<^const_name>\divide\, SOME \<^typ>\'a::divide \ 'a \ 'a\)\ text \Algebraic classes with division\ class semidom_divide = semidom + divide + assumes nonzero_mult_div_cancel_right [simp]: "b \ 0 \ (a * b) div b = a" assumes div_by_0 [simp]: "a div 0 = 0" begin lemma nonzero_mult_div_cancel_left [simp]: "a \ 0 \ (a * b) div a = b" using nonzero_mult_div_cancel_right [of a b] by (simp add: ac_simps) subclass semiring_no_zero_divisors_cancel proof show *: "a * c = b * c \ c = 0 \ a = b" for a b c proof (cases "c = 0") case True then show ?thesis by simp next case False have "a = b" if "a * c = b * c" proof - from that have "a * c div c = b * c div c" by simp with False show ?thesis by simp qed then show ?thesis by auto qed show "c * a = c * b \ c = 0 \ a = b" for a b c using * [of a c b] by (simp add: ac_simps) qed lemma div_self [simp]: "a \ 0 \ a div a = 1" using nonzero_mult_div_cancel_left [of a 1] by simp lemma div_0 [simp]: "0 div a = 0" proof (cases "a = 0") case True then show ?thesis by simp next case False then have "a * 0 div a = 0" by (rule nonzero_mult_div_cancel_left) then show ?thesis by simp qed lemma div_by_1 [simp]: "a div 1 = a" using nonzero_mult_div_cancel_left [of 1 a] by simp lemma dvd_div_eq_0_iff: assumes "b dvd a" shows "a div b = 0 \ a = 0" using assms by (elim dvdE, cases "b = 0") simp_all lemma dvd_div_eq_cancel: "a div c = b div c \ c dvd a \ c dvd b \ a = b" by (elim dvdE, cases "c = 0") simp_all lemma dvd_div_eq_iff: "c dvd a \ c dvd b \ a div c = b div c \ a = b" by (elim dvdE, cases "c = 0") simp_all lemma inj_on_mult: "inj_on ((*) a) A" if "a \ 0" proof (rule inj_onI) fix b c assume "a * b = a * c" then have "a * b div a = a * c div a" by (simp only:) with that show "b = c" by simp qed end class idom_divide = idom + semidom_divide begin lemma dvd_neg_div: assumes "b dvd a" shows "- a div b = - (a div b)" proof (cases "b = 0") case True then show ?thesis by simp next case False from assms obtain c where "a = b * c" .. then have "- a div b = (b * - c) div b" by simp from False also have "\ = - c" by (rule nonzero_mult_div_cancel_left) with False \a = b * c\ show ?thesis by simp qed lemma dvd_div_neg: assumes "b dvd a" shows "a div - b = - (a div b)" proof (cases "b = 0") case True then show ?thesis by simp next case False then have "- b \ 0" by simp from assms obtain c where "a = b * c" .. then have "a div - b = (- b * - c) div - b" by simp from \- b \ 0\ also have "\ = - c" by (rule nonzero_mult_div_cancel_left) with False \a = b * c\ show ?thesis by simp qed end class algebraic_semidom = semidom_divide begin text \ Class \<^class>\algebraic_semidom\ enriches a integral domain by notions from algebra, like units in a ring. It is a separate class to avoid spoiling fields with notions which are degenerated there. \ lemma dvd_times_left_cancel_iff [simp]: assumes "a \ 0" shows "a * b dvd a * c \ b dvd c" (is "?lhs \ ?rhs") proof assume ?lhs then obtain d where "a * c = a * b * d" .. with assms have "c = b * d" by (simp add: ac_simps) then show ?rhs .. next assume ?rhs then obtain d where "c = b * d" .. then have "a * c = a * b * d" by (simp add: ac_simps) then show ?lhs .. qed lemma dvd_times_right_cancel_iff [simp]: assumes "a \ 0" shows "b * a dvd c * a \ b dvd c" using dvd_times_left_cancel_iff [of a b c] assms by (simp add: ac_simps) lemma div_dvd_iff_mult: assumes "b \ 0" and "b dvd a" shows "a div b dvd c \ a dvd c * b" proof - from \b dvd a\ obtain d where "a = b * d" .. with \b \ 0\ show ?thesis by (simp add: ac_simps) qed lemma dvd_div_iff_mult: assumes "c \ 0" and "c dvd b" shows "a dvd b div c \ a * c dvd b" proof - from \c dvd b\ obtain d where "b = c * d" .. with \c \ 0\ show ?thesis by (simp add: mult.commute [of a]) qed lemma div_dvd_div [simp]: assumes "a dvd b" and "a dvd c" shows "b div a dvd c div a \ b dvd c" proof (cases "a = 0") case True with assms show ?thesis by simp next case False moreover from assms obtain k l where "b = a * k" and "c = a * l" by blast ultimately show ?thesis by simp qed lemma div_add [simp]: assumes "c dvd a" and "c dvd b" shows "(a + b) div c = a div c + b div c" proof (cases "c = 0") case True then show ?thesis by simp next case False moreover from assms obtain k l where "a = c * k" and "b = c * l" by blast moreover have "c * k + c * l = c * (k + l)" by (simp add: algebra_simps) ultimately show ?thesis by simp qed lemma div_mult_div_if_dvd: assumes "b dvd a" and "d dvd c" shows "(a div b) * (c div d) = (a * c) div (b * d)" proof (cases "b = 0 \ c = 0") case True with assms show ?thesis by auto next case False moreover from assms obtain k l where "a = b * k" and "c = d * l" by blast moreover have "b * k * (d * l) div (b * d) = (b * d) * (k * l) div (b * d)" by (simp add: ac_simps) ultimately show ?thesis by simp qed lemma dvd_div_eq_mult: assumes "a \ 0" and "a dvd b" shows "b div a = c \ b = c * a" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (simp add: assms) next assume ?lhs then have "b div a * a = c * a" by simp moreover from assms have "b div a * a = b" by (auto simp add: ac_simps) ultimately show ?rhs by simp qed lemma dvd_div_mult_self [simp]: "a dvd b \ b div a * a = b" by (cases "a = 0") (auto simp add: ac_simps) lemma dvd_mult_div_cancel [simp]: "a dvd b \ a * (b div a) = b" using dvd_div_mult_self [of a b] by (simp add: ac_simps) lemma div_mult_swap: assumes "c dvd b" shows "a * (b div c) = (a * b) div c" proof (cases "c = 0") case True then show ?thesis by simp next case False from assms obtain d where "b = c * d" .. moreover from False have "a * divide (d * c) c = ((a * d) * c) div c" by simp ultimately show ?thesis by (simp add: ac_simps) qed lemma dvd_div_mult: "c dvd b \ b div c * a = (b * a) div c" using div_mult_swap [of c b a] by (simp add: ac_simps) lemma dvd_div_mult2_eq: assumes "b * c dvd a" shows "a div (b * c) = a div b div c" proof - from assms obtain k where "a = b * c * k" .. then show ?thesis by (cases "b = 0 \ c = 0") (auto, simp add: ac_simps) qed lemma dvd_div_div_eq_mult: assumes "a \ 0" "c \ 0" and "a dvd b" "c dvd d" shows "b div a = d div c \ b * c = a * d" (is "?lhs \ ?rhs") proof - from assms have "a * c \ 0" by simp then have "?lhs \ b div a * (a * c) = d div c * (a * c)" by simp also have "\ \ (a * (b div a)) * c = (c * (d div c)) * a" by (simp add: ac_simps) also have "\ \ (a * b div a) * c = (c * d div c) * a" using assms by (simp add: div_mult_swap) also have "\ \ ?rhs" using assms by (simp add: ac_simps) finally show ?thesis . qed lemma dvd_mult_imp_div: assumes "a * c dvd b" shows "a dvd b div c" proof (cases "c = 0") case True then show ?thesis by simp next case False from \a * c dvd b\ obtain d where "b = a * c * d" .. with False show ?thesis by (simp add: mult.commute [of a] mult.assoc) qed lemma div_div_eq_right: assumes "c dvd b" "b dvd a" shows "a div (b div c) = a div b * c" proof (cases "c = 0 \ b = 0") case True then show ?thesis by auto next case False from assms obtain r s where "b = c * r" and "a = c * r * s" by blast moreover with False have "r \ 0" by auto ultimately show ?thesis using False by simp (simp add: mult.commute [of _ r] mult.assoc mult.commute [of c]) qed lemma div_div_div_same: assumes "d dvd b" "b dvd a" shows "(a div d) div (b div d) = a div b" proof (cases "b = 0 \ d = 0") case True with assms show ?thesis by auto next case False from assms obtain r s where "a = d * r * s" and "b = d * r" by blast with False show ?thesis by simp (simp add: ac_simps) qed text \Units: invertible elements in a ring\ abbreviation is_unit :: "'a \ bool" where "is_unit a \ a dvd 1" lemma not_is_unit_0 [simp]: "\ is_unit 0" by simp lemma unit_imp_dvd [dest]: "is_unit b \ b dvd a" by (rule dvd_trans [of _ 1]) simp_all lemma unit_dvdE: assumes "is_unit a" obtains c where "a \ 0" and "b = a * c" proof - from assms have "a dvd b" by auto then obtain c where "b = a * c" .. moreover from assms have "a \ 0" by auto ultimately show thesis using that by blast qed lemma dvd_unit_imp_unit: "a dvd b \ is_unit b \ is_unit a" by (rule dvd_trans) lemma unit_div_1_unit [simp, intro]: assumes "is_unit a" shows "is_unit (1 div a)" proof - from assms have "1 = 1 div a * a" by simp then show "is_unit (1 div a)" by (rule dvdI) qed lemma is_unitE [elim?]: assumes "is_unit a" obtains b where "a \ 0" and "b \ 0" and "is_unit b" and "1 div a = b" and "1 div b = a" and "a * b = 1" and "c div a = c * b" proof (rule that) define b where "b = 1 div a" then show "1 div a = b" by simp from assms b_def show "is_unit b" by simp with assms show "a \ 0" and "b \ 0" by auto from assms b_def show "a * b = 1" by simp then have "1 = a * b" .. with b_def \b \ 0\ show "1 div b = a" by simp from assms have "a dvd c" .. then obtain d where "c = a * d" .. with \a \ 0\ \a * b = 1\ show "c div a = c * b" by (simp add: mult.assoc mult.left_commute [of a]) qed lemma unit_prod [intro]: "is_unit a \ is_unit b \ is_unit (a * b)" by (subst mult_1_left [of 1, symmetric]) (rule mult_dvd_mono) lemma is_unit_mult_iff: "is_unit (a * b) \ is_unit a \ is_unit b" by (auto dest: dvd_mult_left dvd_mult_right) lemma unit_div [intro]: "is_unit a \ is_unit b \ is_unit (a div b)" by (erule is_unitE [of b a]) (simp add: ac_simps unit_prod) lemma mult_unit_dvd_iff: assumes "is_unit b" shows "a * b dvd c \ a dvd c" proof assume "a * b dvd c" with assms show "a dvd c" by (simp add: dvd_mult_left) next assume "a dvd c" then obtain k where "c = a * k" .. with assms have "c = (a * b) * (1 div b * k)" by (simp add: mult_ac) then show "a * b dvd c" by (rule dvdI) qed lemma mult_unit_dvd_iff': "is_unit a \ (a * b) dvd c \ b dvd c" using mult_unit_dvd_iff [of a b c] by (simp add: ac_simps) lemma dvd_mult_unit_iff: assumes "is_unit b" shows "a dvd c * b \ a dvd c" proof assume "a dvd c * b" with assms have "c * b dvd c * (b * (1 div b))" by (subst mult_assoc [symmetric]) simp also from assms have "b * (1 div b) = 1" by (rule is_unitE) simp finally have "c * b dvd c" by simp with \a dvd c * b\ show "a dvd c" by (rule dvd_trans) next assume "a dvd c" then show "a dvd c * b" by simp qed lemma dvd_mult_unit_iff': "is_unit b \ a dvd b * c \ a dvd c" using dvd_mult_unit_iff [of b a c] by (simp add: ac_simps) lemma div_unit_dvd_iff: "is_unit b \ a div b dvd c \ a dvd c" by (erule is_unitE [of _ a]) (auto simp add: mult_unit_dvd_iff) lemma dvd_div_unit_iff: "is_unit b \ a dvd c div b \ a dvd c" by (erule is_unitE [of _ c]) (simp add: dvd_mult_unit_iff) lemmas unit_dvd_iff = mult_unit_dvd_iff mult_unit_dvd_iff' dvd_mult_unit_iff dvd_mult_unit_iff' div_unit_dvd_iff dvd_div_unit_iff (* FIXME consider named_theorems *) lemma unit_mult_div_div [simp]: "is_unit a \ b * (1 div a) = b div a" by (erule is_unitE [of _ b]) simp lemma unit_div_mult_self [simp]: "is_unit a \ b div a * a = b" by (rule dvd_div_mult_self) auto lemma unit_div_1_div_1 [simp]: "is_unit a \ 1 div (1 div a) = a" by (erule is_unitE) simp lemma unit_div_mult_swap: "is_unit c \ a * (b div c) = (a * b) div c" by (erule unit_dvdE [of _ b]) (simp add: mult.left_commute [of _ c]) lemma unit_div_commute: "is_unit b \ (a div b) * c = (a * c) div b" using unit_div_mult_swap [of b c a] by (simp add: ac_simps) lemma unit_eq_div1: "is_unit b \ a div b = c \ a = c * b" by (auto elim: is_unitE) lemma unit_eq_div2: "is_unit b \ a = c div b \ a * b = c" using unit_eq_div1 [of b c a] by auto lemma unit_mult_left_cancel: "is_unit a \ a * b = a * c \ b = c" using mult_cancel_left [of a b c] by auto lemma unit_mult_right_cancel: "is_unit a \ b * a = c * a \ b = c" using unit_mult_left_cancel [of a b c] by (auto simp add: ac_simps) lemma unit_div_cancel: assumes "is_unit a" shows "b div a = c div a \ b = c" proof - from assms have "is_unit (1 div a)" by simp then have "b * (1 div a) = c * (1 div a) \ b = c" by (rule unit_mult_right_cancel) with assms show ?thesis by simp qed lemma is_unit_div_mult2_eq: assumes "is_unit b" and "is_unit c" shows "a div (b * c) = a div b div c" proof - from assms have "is_unit (b * c)" by (simp add: unit_prod) then have "b * c dvd a" by (rule unit_imp_dvd) then show ?thesis by (rule dvd_div_mult2_eq) qed lemma is_unit_div_mult_cancel_left: assumes "a \ 0" and "is_unit b" shows "a div (a * b) = 1 div b" proof - from assms have "a div (a * b) = a div a div b" by (simp add: mult_unit_dvd_iff dvd_div_mult2_eq) with assms show ?thesis by simp qed lemma is_unit_div_mult_cancel_right: assumes "a \ 0" and "is_unit b" shows "a div (b * a) = 1 div b" using assms is_unit_div_mult_cancel_left [of a b] by (simp add: ac_simps) lemma unit_div_eq_0_iff: assumes "is_unit b" shows "a div b = 0 \ a = 0" by (rule dvd_div_eq_0_iff) (insert assms, auto) lemma div_mult_unit2: "is_unit c \ b dvd a \ a div (b * c) = a div b div c" by (rule dvd_div_mult2_eq) (simp_all add: mult_unit_dvd_iff) text \Coprimality\ definition coprime :: "'a \ 'a \ bool" where "coprime a b \ (\c. c dvd a \ c dvd b \ is_unit c)" lemma coprimeI: assumes "\c. c dvd a \ c dvd b \ is_unit c" shows "coprime a b" using assms by (auto simp: coprime_def) lemma not_coprimeI: assumes "c dvd a" and "c dvd b" and "\ is_unit c" shows "\ coprime a b" using assms by (auto simp: coprime_def) lemma coprime_common_divisor: "is_unit c" if "coprime a b" and "c dvd a" and "c dvd b" using that by (auto simp: coprime_def) lemma not_coprimeE: assumes "\ coprime a b" obtains c where "c dvd a" and "c dvd b" and "\ is_unit c" using assms by (auto simp: coprime_def) lemma coprime_imp_coprime: "coprime a b" if "coprime c d" and "\e. \ is_unit e \ e dvd a \ e dvd b \ e dvd c" and "\e. \ is_unit e \ e dvd a \ e dvd b \ e dvd d" proof (rule coprimeI) fix e assume "e dvd a" and "e dvd b" with that have "e dvd c" and "e dvd d" by (auto intro: dvd_trans) with \coprime c d\ show "is_unit e" by (rule coprime_common_divisor) qed lemma coprime_divisors: "coprime a b" if "a dvd c" "b dvd d" and "coprime c d" using \coprime c d\ proof (rule coprime_imp_coprime) fix e assume "e dvd a" then show "e dvd c" using \a dvd c\ by (rule dvd_trans) assume "e dvd b" then show "e dvd d" using \b dvd d\ by (rule dvd_trans) qed lemma coprime_self [simp]: "coprime a a \ is_unit a" (is "?P \ ?Q") proof assume ?P then show ?Q by (rule coprime_common_divisor) simp_all next assume ?Q show ?P by (rule coprimeI) (erule dvd_unit_imp_unit, rule \?Q\) qed lemma coprime_commute [ac_simps]: "coprime b a \ coprime a b" unfolding coprime_def by auto lemma is_unit_left_imp_coprime: "coprime a b" if "is_unit a" proof (rule coprimeI) fix c assume "c dvd a" with that show "is_unit c" by (auto intro: dvd_unit_imp_unit) qed lemma is_unit_right_imp_coprime: "coprime a b" if "is_unit b" using that is_unit_left_imp_coprime [of b a] by (simp add: ac_simps) lemma coprime_1_left [simp]: "coprime 1 a" by (rule coprimeI) lemma coprime_1_right [simp]: "coprime a 1" by (rule coprimeI) lemma coprime_0_left_iff [simp]: "coprime 0 a \ is_unit a" by (auto intro: coprimeI dvd_unit_imp_unit coprime_common_divisor [of 0 a a]) lemma coprime_0_right_iff [simp]: "coprime a 0 \ is_unit a" using coprime_0_left_iff [of a] by (simp add: ac_simps) lemma coprime_mult_self_left_iff [simp]: "coprime (c * a) (c * b) \ is_unit c \ coprime a b" by (auto intro: coprime_common_divisor) (rule coprimeI, auto intro: coprime_common_divisor simp add: dvd_mult_unit_iff')+ lemma coprime_mult_self_right_iff [simp]: "coprime (a * c) (b * c) \ is_unit c \ coprime a b" using coprime_mult_self_left_iff [of c a b] by (simp add: ac_simps) lemma coprime_absorb_left: assumes "x dvd y" shows "coprime x y \ is_unit x" using assms coprime_common_divisor is_unit_left_imp_coprime by auto lemma coprime_absorb_right: assumes "y dvd x" shows "coprime x y \ is_unit y" using assms coprime_common_divisor is_unit_right_imp_coprime by auto end class unit_factor = fixes unit_factor :: "'a \ 'a" class semidom_divide_unit_factor = semidom_divide + unit_factor + assumes unit_factor_0 [simp]: "unit_factor 0 = 0" and is_unit_unit_factor: "a dvd 1 \ unit_factor a = a" and unit_factor_is_unit: "a \ 0 \ unit_factor a dvd 1" and unit_factor_mult_unit_left: "a dvd 1 \ unit_factor (a * b) = a * unit_factor b" \ \This fine-grained hierarchy will later on allow lean normalization of polynomials\ begin lemma unit_factor_mult_unit_right: "a dvd 1 \ unit_factor (b * a) = unit_factor b * a" using unit_factor_mult_unit_left[of a b] by (simp add: mult_ac) lemmas [simp] = unit_factor_mult_unit_left unit_factor_mult_unit_right end class normalization_semidom = algebraic_semidom + semidom_divide_unit_factor + fixes normalize :: "'a \ 'a" assumes unit_factor_mult_normalize [simp]: "unit_factor a * normalize a = a" and normalize_0 [simp]: "normalize 0 = 0" begin text \ Class \<^class>\normalization_semidom\ cultivates the idea that each integral domain can be split into equivalence classes whose representants are associated, i.e. divide each other. \<^const>\normalize\ specifies a canonical representant for each equivalence class. The rationale behind this is that it is easier to reason about equality than equivalences, hence we prefer to think about equality of normalized values rather than associated elements. \ declare unit_factor_is_unit [iff] lemma unit_factor_dvd [simp]: "a \ 0 \ unit_factor a dvd b" by (rule unit_imp_dvd) simp lemma unit_factor_self [simp]: "unit_factor a dvd a" by (cases "a = 0") simp_all lemma normalize_mult_unit_factor [simp]: "normalize a * unit_factor a = a" using unit_factor_mult_normalize [of a] by (simp add: ac_simps) lemma normalize_eq_0_iff [simp]: "normalize a = 0 \ a = 0" (is "?lhs \ ?rhs") proof assume ?lhs moreover have "unit_factor a * normalize a = a" by simp ultimately show ?rhs by simp next assume ?rhs then show ?lhs by simp qed lemma unit_factor_eq_0_iff [simp]: "unit_factor a = 0 \ a = 0" (is "?lhs \ ?rhs") proof assume ?lhs moreover have "unit_factor a * normalize a = a" by simp ultimately show ?rhs by simp next assume ?rhs then show ?lhs by simp qed lemma div_unit_factor [simp]: "a div unit_factor a = normalize a" proof (cases "a = 0") case True then show ?thesis by simp next case False then have "unit_factor a \ 0" by simp with nonzero_mult_div_cancel_left have "unit_factor a * normalize a div unit_factor a = normalize a" by blast then show ?thesis by simp qed lemma normalize_div [simp]: "normalize a div a = 1 div unit_factor a" proof (cases "a = 0") case True then show ?thesis by simp next case False have "normalize a div a = normalize a div (unit_factor a * normalize a)" by simp also have "\ = 1 div unit_factor a" using False by (subst is_unit_div_mult_cancel_right) simp_all finally show ?thesis . qed lemma is_unit_normalize: assumes "is_unit a" shows "normalize a = 1" proof - from assms have "unit_factor a = a" by (rule is_unit_unit_factor) moreover from assms have "a \ 0" by auto moreover have "normalize a = a div unit_factor a" by simp ultimately show ?thesis by simp qed lemma unit_factor_1 [simp]: "unit_factor 1 = 1" by (rule is_unit_unit_factor) simp lemma normalize_1 [simp]: "normalize 1 = 1" by (rule is_unit_normalize) simp lemma normalize_1_iff: "normalize a = 1 \ is_unit a" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (rule is_unit_normalize) next assume ?lhs then have "unit_factor a * normalize a = unit_factor a * 1" by simp then have "unit_factor a = a" by simp moreover from \?lhs\ have "a \ 0" by auto then have "is_unit (unit_factor a)" by simp ultimately show ?rhs by simp qed lemma div_normalize [simp]: "a div normalize a = unit_factor a" proof (cases "a = 0") case True then show ?thesis by simp next case False then have "normalize a \ 0" by simp with nonzero_mult_div_cancel_right have "unit_factor a * normalize a div normalize a = unit_factor a" by blast then show ?thesis by simp qed lemma mult_one_div_unit_factor [simp]: "a * (1 div unit_factor b) = a div unit_factor b" by (cases "b = 0") simp_all lemma inv_unit_factor_eq_0_iff [simp]: "1 div unit_factor a = 0 \ a = 0" (is "?lhs \ ?rhs") proof assume ?lhs then have "a * (1 div unit_factor a) = a * 0" by simp then show ?rhs by simp next assume ?rhs then show ?lhs by simp qed lemma unit_factor_idem [simp]: "unit_factor (unit_factor a) = unit_factor a" by (cases "a = 0") (auto intro: is_unit_unit_factor) lemma normalize_unit_factor [simp]: "a \ 0 \ normalize (unit_factor a) = 1" by (rule is_unit_normalize) simp lemma normalize_mult_unit_left [simp]: assumes "a dvd 1" shows "normalize (a * b) = normalize b" proof (cases "b = 0") case False have "a * unit_factor b * normalize (a * b) = unit_factor (a * b) * normalize (a * b)" using assms by (subst unit_factor_mult_unit_left) auto also have "\ = a * b" by simp also have "b = unit_factor b * normalize b" by simp hence "a * b = a * unit_factor b * normalize b" by (simp only: mult_ac) finally show ?thesis using assms False by auto qed auto lemma normalize_mult_unit_right [simp]: assumes "b dvd 1" shows "normalize (a * b) = normalize a" using assms by (subst mult.commute) auto lemma normalize_idem [simp]: "normalize (normalize a) = normalize a" proof (cases "a = 0") case False have "normalize a = normalize (unit_factor a * normalize a)" by simp also from False have "\ = normalize (normalize a)" by (subst normalize_mult_unit_left) auto finally show ?thesis .. qed auto lemma unit_factor_normalize [simp]: assumes "a \ 0" shows "unit_factor (normalize a) = 1" proof - from assms have *: "normalize a \ 0" by simp have "unit_factor (normalize a) * normalize (normalize a) = normalize a" by (simp only: unit_factor_mult_normalize) then have "unit_factor (normalize a) * normalize a = normalize a" by simp with * have "unit_factor (normalize a) * normalize a div normalize a = normalize a div normalize a" by simp with * show ?thesis by simp qed lemma normalize_dvd_iff [simp]: "normalize a dvd b \ a dvd b" proof - have "normalize a dvd b \ unit_factor a * normalize a dvd b" using mult_unit_dvd_iff [of "unit_factor a" "normalize a" b] by (cases "a = 0") simp_all then show ?thesis by simp qed lemma dvd_normalize_iff [simp]: "a dvd normalize b \ a dvd b" proof - have "a dvd normalize b \ a dvd normalize b * unit_factor b" using dvd_mult_unit_iff [of "unit_factor b" a "normalize b"] by (cases "b = 0") simp_all then show ?thesis by simp qed lemma normalize_idem_imp_unit_factor_eq: assumes "normalize a = a" shows "unit_factor a = of_bool (a \ 0)" proof (cases "a = 0") case True then show ?thesis by simp next case False then show ?thesis using assms unit_factor_normalize [of a] by simp qed lemma normalize_idem_imp_is_unit_iff: assumes "normalize a = a" shows "is_unit a \ a = 1" using assms by (cases "a = 0") (auto dest: is_unit_normalize) lemma coprime_normalize_left_iff [simp]: "coprime (normalize a) b \ coprime a b" by (rule; rule coprimeI) (auto intro: coprime_common_divisor) lemma coprime_normalize_right_iff [simp]: "coprime a (normalize b) \ coprime a b" using coprime_normalize_left_iff [of b a] by (simp add: ac_simps) text \ We avoid an explicit definition of associated elements but prefer explicit normalisation instead. In theory we could define an abbreviation like \<^prop>\associated a b \ normalize a = normalize b\ but this is counterproductive without suggestive infix syntax, which we do not want to sacrifice for this purpose here. \ lemma associatedI: assumes "a dvd b" and "b dvd a" shows "normalize a = normalize b" proof (cases "a = 0 \ b = 0") case True with assms show ?thesis by auto next case False from \a dvd b\ obtain c where b: "b = a * c" .. moreover from \b dvd a\ obtain d where a: "a = b * d" .. ultimately have "b * 1 = b * (c * d)" by (simp add: ac_simps) with False have "1 = c * d" unfolding mult_cancel_left by simp then have "is_unit c" and "is_unit d" by auto with a b show ?thesis by (simp add: is_unit_normalize) qed lemma associatedD1: "normalize a = normalize b \ a dvd b" using dvd_normalize_iff [of _ b, symmetric] normalize_dvd_iff [of a _, symmetric] by simp lemma associatedD2: "normalize a = normalize b \ b dvd a" using dvd_normalize_iff [of _ a, symmetric] normalize_dvd_iff [of b _, symmetric] by simp lemma associated_unit: "normalize a = normalize b \ is_unit a \ is_unit b" using dvd_unit_imp_unit by (auto dest!: associatedD1 associatedD2) lemma associated_iff_dvd: "normalize a = normalize b \ a dvd b \ b dvd a" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (auto intro!: associatedI) next assume ?lhs then have "unit_factor a * normalize a = unit_factor a * normalize b" by simp then have *: "normalize b * unit_factor a = a" by (simp add: ac_simps) show ?rhs proof (cases "a = 0 \ b = 0") case True with \?lhs\ show ?thesis by auto next case False then have "b dvd normalize b * unit_factor a" and "normalize b * unit_factor a dvd b" by (simp_all add: mult_unit_dvd_iff dvd_mult_unit_iff) with * show ?thesis by simp qed qed lemma associated_eqI: assumes "a dvd b" and "b dvd a" assumes "normalize a = a" and "normalize b = b" shows "a = b" proof - from assms have "normalize a = normalize b" unfolding associated_iff_dvd by simp with \normalize a = a\ have "a = normalize b" by simp with \normalize b = b\ show "a = b" by simp qed lemma normalize_unit_factor_eqI: assumes "normalize a = normalize b" and "unit_factor a = unit_factor b" shows "a = b" proof - from assms have "unit_factor a * normalize a = unit_factor b * normalize b" by simp then show ?thesis by simp qed lemma normalize_mult_normalize_left [simp]: "normalize (normalize a * b) = normalize (a * b)" by (rule associated_eqI) (auto intro!: mult_dvd_mono) lemma normalize_mult_normalize_right [simp]: "normalize (a * normalize b) = normalize (a * b)" by (rule associated_eqI) (auto intro!: mult_dvd_mono) end class normalization_semidom_multiplicative = normalization_semidom + assumes unit_factor_mult: "unit_factor (a * b) = unit_factor a * unit_factor b" begin lemma normalize_mult: "normalize (a * b) = normalize a * normalize b" proof (cases "a = 0 \ b = 0") case True then show ?thesis by auto next case False have "unit_factor (a * b) * normalize (a * b) = a * b" by (rule unit_factor_mult_normalize) then have "normalize (a * b) = a * b div unit_factor (a * b)" by simp also have "\ = a * b div unit_factor (b * a)" by (simp add: ac_simps) also have "\ = a * b div unit_factor b div unit_factor a" using False by (simp add: unit_factor_mult is_unit_div_mult2_eq [symmetric]) also have "\ = a * (b div unit_factor b) div unit_factor a" using False by (subst unit_div_mult_swap) simp_all also have "\ = normalize a * normalize b" using False by (simp add: mult.commute [of a] mult.commute [of "normalize a"] unit_div_mult_swap [symmetric]) finally show ?thesis . qed lemma dvd_unit_factor_div: assumes "b dvd a" shows "unit_factor (a div b) = unit_factor a div unit_factor b" proof - from assms have "a = a div b * b" by simp then have "unit_factor a = unit_factor (a div b * b)" by simp then show ?thesis by (cases "b = 0") (simp_all add: unit_factor_mult) qed lemma dvd_normalize_div: assumes "b dvd a" shows "normalize (a div b) = normalize a div normalize b" proof - from assms have "a = a div b * b" by simp then have "normalize a = normalize (a div b * b)" by simp then show ?thesis by (cases "b = 0") (simp_all add: normalize_mult) qed end text \Syntactic division remainder operator\ class modulo = dvd + divide + fixes modulo :: "'a \ 'a \ 'a" (infixl "mod" 70) text \Arbitrary quotient and remainder partitions\ class semiring_modulo = comm_semiring_1_cancel + divide + modulo + assumes div_mult_mod_eq: "a div b * b + a mod b = a" begin lemma mod_div_decomp: fixes a b obtains q r where "q = a div b" and "r = a mod b" and "a = q * b + r" proof - from div_mult_mod_eq have "a = a div b * b + a mod b" by simp moreover have "a div b = a div b" .. moreover have "a mod b = a mod b" .. note that ultimately show thesis by blast qed lemma mult_div_mod_eq: "b * (a div b) + a mod b = a" using div_mult_mod_eq [of a b] by (simp add: ac_simps) lemma mod_div_mult_eq: "a mod b + a div b * b = a" using div_mult_mod_eq [of a b] by (simp add: ac_simps) lemma mod_mult_div_eq: "a mod b + b * (a div b) = a" using div_mult_mod_eq [of a b] by (simp add: ac_simps) lemma minus_div_mult_eq_mod: "a - a div b * b = a mod b" by (rule add_implies_diff [symmetric]) (fact mod_div_mult_eq) lemma minus_mult_div_eq_mod: "a - b * (a div b) = a mod b" by (rule add_implies_diff [symmetric]) (fact mod_mult_div_eq) lemma minus_mod_eq_div_mult: "a - a mod b = a div b * b" by (rule add_implies_diff [symmetric]) (fact div_mult_mod_eq) lemma minus_mod_eq_mult_div: "a - a mod b = b * (a div b)" by (rule add_implies_diff [symmetric]) (fact mult_div_mod_eq) lemma mod_0_imp_dvd [dest!]: "b dvd a" if "a mod b = 0" proof - have "b dvd (a div b) * b" by simp also have "(a div b) * b = a" using div_mult_mod_eq [of a b] by (simp add: that) finally show ?thesis . qed lemma [nitpick_unfold]: "a mod b = a - a div b * b" by (fact minus_div_mult_eq_mod [symmetric]) end subsection \Quotient and remainder in integral domains\ class semidom_modulo = algebraic_semidom + semiring_modulo begin lemma mod_0 [simp]: "0 mod a = 0" using div_mult_mod_eq [of 0 a] by simp lemma mod_by_0 [simp]: "a mod 0 = a" using div_mult_mod_eq [of a 0] by simp lemma mod_by_1 [simp]: "a mod 1 = 0" proof - from div_mult_mod_eq [of a one] div_by_1 have "a + a mod 1 = a" by simp then have "a + a mod 1 = a + 0" by simp then show ?thesis by (rule add_left_imp_eq) qed lemma mod_self [simp]: "a mod a = 0" using div_mult_mod_eq [of a a] by simp lemma dvd_imp_mod_0 [simp]: "b mod a = 0" if "a dvd b" using that minus_div_mult_eq_mod [of b a] by simp lemma mod_eq_0_iff_dvd: "a mod b = 0 \ b dvd a" by (auto intro: mod_0_imp_dvd) lemma dvd_eq_mod_eq_0 [nitpick_unfold, code]: "a dvd b \ b mod a = 0" by (simp add: mod_eq_0_iff_dvd) lemma dvd_mod_iff: assumes "c dvd b" shows "c dvd a mod b \ c dvd a" proof - from assms have "(c dvd a mod b) \ (c dvd ((a div b) * b + a mod b))" by (simp add: dvd_add_right_iff) also have "(a div b) * b + a mod b = a" using div_mult_mod_eq [of a b] by simp finally show ?thesis . qed lemma dvd_mod_imp_dvd: assumes "c dvd a mod b" and "c dvd b" shows "c dvd a" using assms dvd_mod_iff [of c b a] by simp lemma dvd_minus_mod [simp]: "b dvd a - a mod b" by (simp add: minus_mod_eq_div_mult) lemma cancel_div_mod_rules: "((a div b) * b + a mod b) + c = a + c" "(b * (a div b) + a mod b) + c = a + c" by (simp_all add: div_mult_mod_eq mult_div_mod_eq) end class idom_modulo = idom + semidom_modulo begin subclass idom_divide .. lemma div_diff [simp]: "c dvd a \ c dvd b \ (a - b) div c = a div c - b div c" using div_add [of _ _ "- b"] by (simp add: dvd_neg_div) end subsection \Interlude: basic tool support for algebraic and arithmetic calculations\ named_theorems arith "arith facts -- only ground formulas" ML_file \Tools/arith_data.ML\ ML_file \~~/src/Provers/Arith/cancel_div_mod.ML\ ML \ structure Cancel_Div_Mod_Ring = Cancel_Div_Mod ( val div_name = \<^const_name>\divide\; val mod_name = \<^const_name>\modulo\; val mk_binop = HOLogic.mk_binop; val mk_sum = Arith_Data.mk_sum; val dest_sum = Arith_Data.dest_sum; val div_mod_eqs = map mk_meta_eq @{thms cancel_div_mod_rules}; val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac @{thms diff_conv_add_uminus add_0_left add_0_right ac_simps}) ) \ simproc_setup cancel_div_mod_int ("(a::'a::semidom_modulo) + b") = \K Cancel_Div_Mod_Ring.proc\ subsection \Ordered semirings and rings\ text \ The theory of partially ordered rings is taken from the books: \<^item> \<^emph>\Lattice Theory\ by Garret Birkhoff, American Mathematical Society, 1979 \<^item> \<^emph>\Partially Ordered Algebraic Systems\, Pergamon Press, 1963 Most of the used notions can also be looked up in \<^item> \<^url>\http://www.mathworld.com\ by Eric Weisstein et. al. \<^item> \<^emph>\Algebra I\ by van der Waerden, Springer \ class ordered_semiring = semiring + ordered_comm_monoid_add + assumes mult_left_mono: "a \ b \ 0 \ c \ c * a \ c * b" assumes mult_right_mono: "a \ b \ 0 \ c \ a * c \ b * c" begin lemma mult_mono: "a \ b \ c \ d \ 0 \ b \ 0 \ c \ a * c \ b * d" apply (erule (1) mult_right_mono [THEN order_trans]) apply (erule (1) mult_left_mono) done lemma mult_mono': "a \ b \ c \ d \ 0 \ a \ 0 \ c \ a * c \ b * d" by (rule mult_mono) (fast intro: order_trans)+ end +lemma mono_mult: + fixes a :: "'a::ordered_semiring" + shows "a \ 0 \ mono ((*) a)" + by (simp add: mono_def mult_left_mono) + class ordered_semiring_0 = semiring_0 + ordered_semiring begin lemma mult_nonneg_nonneg [simp]: "0 \ a \ 0 \ b \ 0 \ a * b" using mult_left_mono [of 0 b a] by simp lemma mult_nonneg_nonpos: "0 \ a \ b \ 0 \ a * b \ 0" using mult_left_mono [of b 0 a] by simp lemma mult_nonpos_nonneg: "a \ 0 \ 0 \ b \ a * b \ 0" using mult_right_mono [of a 0 b] by simp text \Legacy -- use @{thm [source] mult_nonpos_nonneg}.\ lemma mult_nonneg_nonpos2: "0 \ a \ b \ 0 \ b * a \ 0" by (drule mult_right_mono [of b 0]) auto lemma split_mult_neg_le: "(0 \ a \ b \ 0) \ (a \ 0 \ 0 \ b) \ a * b \ 0" by (auto simp add: mult_nonneg_nonpos mult_nonneg_nonpos2) end class ordered_cancel_semiring = ordered_semiring + cancel_comm_monoid_add begin subclass semiring_0_cancel .. subclass ordered_semiring_0 .. end class linordered_semiring = ordered_semiring + linordered_cancel_ab_semigroup_add begin subclass ordered_cancel_semiring .. subclass ordered_cancel_comm_monoid_add .. subclass ordered_ab_semigroup_monoid_add_imp_le .. lemma mult_left_less_imp_less: "c * a < c * b \ 0 \ c \ a < b" by (force simp add: mult_left_mono not_le [symmetric]) lemma mult_right_less_imp_less: "a * c < b * c \ 0 \ c \ a < b" by (force simp add: mult_right_mono not_le [symmetric]) end class zero_less_one = order + zero + one + assumes zero_less_one [simp]: "0 < 1" begin subclass zero_neq_one by standard (simp add: less_imp_neq) lemma zero_le_one [simp]: \0 \ 1\ by (rule less_imp_le) simp end class linordered_semiring_1 = linordered_semiring + semiring_1 + zero_less_one begin lemma convex_bound_le: assumes "x \ a" "y \ a" "0 \ u" "0 \ v" "u + v = 1" shows "u * x + v * y \ a" proof- from assms have "u * x + v * y \ u * a + v * a" by (simp add: add_mono mult_left_mono) with assms show ?thesis unfolding distrib_right[symmetric] by simp qed end class linordered_semiring_strict = semiring + comm_monoid_add + linordered_cancel_ab_semigroup_add + assumes mult_strict_left_mono: "a < b \ 0 < c \ c * a < c * b" assumes mult_strict_right_mono: "a < b \ 0 < c \ a * c < b * c" begin subclass semiring_0_cancel .. subclass linordered_semiring proof fix a b c :: 'a assume *: "a \ b" "0 \ c" then show "c * a \ c * b" unfolding le_less using mult_strict_left_mono by (cases "c = 0") auto from * show "a * c \ b * c" unfolding le_less using mult_strict_right_mono by (cases "c = 0") auto qed lemma mult_left_le_imp_le: "c * a \ c * b \ 0 < c \ a \ b" by (auto simp add: mult_strict_left_mono _not_less [symmetric]) lemma mult_right_le_imp_le: "a * c \ b * c \ 0 < c \ a \ b" by (auto simp add: mult_strict_right_mono not_less [symmetric]) lemma mult_pos_pos[simp]: "0 < a \ 0 < b \ 0 < a * b" using mult_strict_left_mono [of 0 b a] by simp lemma mult_pos_neg: "0 < a \ b < 0 \ a * b < 0" using mult_strict_left_mono [of b 0 a] by simp lemma mult_neg_pos: "a < 0 \ 0 < b \ a * b < 0" using mult_strict_right_mono [of a 0 b] by simp text \Legacy -- use @{thm [source] mult_neg_pos}.\ lemma mult_pos_neg2: "0 < a \ b < 0 \ b * a < 0" by (drule mult_strict_right_mono [of b 0]) auto lemma zero_less_mult_pos: assumes "0 < a * b" "0 < a" shows "0 < b" proof (cases "b \ 0") case True then show ?thesis using assms by (auto simp: le_less dest: less_not_sym mult_pos_neg [of a b]) qed (auto simp add: le_less not_less) lemma zero_less_mult_pos2: assumes "0 < b * a" "0 < a" shows "0 < b" proof (cases "b \ 0") case True then show ?thesis using assms by (auto simp: le_less dest: less_not_sym mult_pos_neg2 [of a b]) qed (auto simp add: le_less not_less) text \Strict monotonicity in both arguments\ lemma mult_strict_mono: assumes "a < b" "c < d" "0 < b" "0 \ c" shows "a * c < b * d" proof (cases "c = 0") case True with assms show ?thesis by simp next case False with assms have "a*c < b*c" by (simp add: mult_strict_right_mono [OF \a < b\]) also have "\ < b*d" by (simp add: assms mult_strict_left_mono) finally show ?thesis . qed text \This weaker variant has more natural premises\ lemma mult_strict_mono': assumes "a < b" and "c < d" and "0 \ a" and "0 \ c" shows "a * c < b * d" by (rule mult_strict_mono) (insert assms, auto) lemma mult_less_le_imp_less: assumes "a < b" and "c \ d" and "0 \ a" and "0 < c" shows "a * c < b * d" proof - have "a * c < b * c" by (simp add: assms mult_strict_right_mono) also have "... \ b * d" by (intro mult_left_mono) (use assms in auto) finally show ?thesis . qed lemma mult_le_less_imp_less: assumes "a \ b" and "c < d" and "0 < a" and "0 \ c" shows "a * c < b * d" proof - have "a * c \ b * c" by (simp add: assms mult_right_mono) also have "... < b * d" by (intro mult_strict_left_mono) (use assms in auto) finally show ?thesis . qed end class linordered_semiring_1_strict = linordered_semiring_strict + semiring_1 + zero_less_one begin subclass linordered_semiring_1 .. lemma convex_bound_lt: assumes "x < a" "y < a" "0 \ u" "0 \ v" "u + v = 1" shows "u * x + v * y < a" proof - from assms have "u * x + v * y < u * a + v * a" by (cases "u = 0") (auto intro!: add_less_le_mono mult_strict_left_mono mult_left_mono) with assms show ?thesis unfolding distrib_right[symmetric] by simp qed end class ordered_comm_semiring = comm_semiring_0 + ordered_ab_semigroup_add + assumes comm_mult_left_mono: "a \ b \ 0 \ c \ c * a \ c * b" begin subclass ordered_semiring proof fix a b c :: 'a assume "a \ b" "0 \ c" then show "c * a \ c * b" by (rule comm_mult_left_mono) then show "a * c \ b * c" by (simp only: mult.commute) qed end class ordered_cancel_comm_semiring = ordered_comm_semiring + cancel_comm_monoid_add begin subclass comm_semiring_0_cancel .. subclass ordered_comm_semiring .. subclass ordered_cancel_semiring .. end class linordered_comm_semiring_strict = comm_semiring_0 + linordered_cancel_ab_semigroup_add + assumes comm_mult_strict_left_mono: "a < b \ 0 < c \ c * a < c * b" begin subclass linordered_semiring_strict proof fix a b c :: 'a assume "a < b" "0 < c" then show "c * a < c * b" by (rule comm_mult_strict_left_mono) then show "a * c < b * c" by (simp only: mult.commute) qed subclass ordered_cancel_comm_semiring proof fix a b c :: 'a assume "a \ b" "0 \ c" then show "c * a \ c * b" unfolding le_less using mult_strict_left_mono by (cases "c = 0") auto qed end class ordered_ring = ring + ordered_cancel_semiring begin subclass ordered_ab_group_add .. lemma less_add_iff1: "a * e + c < b * e + d \ (a - b) * e + c < d" by (simp add: algebra_simps) lemma less_add_iff2: "a * e + c < b * e + d \ c < (b - a) * e + d" by (simp add: algebra_simps) lemma le_add_iff1: "a * e + c \ b * e + d \ (a - b) * e + c \ d" by (simp add: algebra_simps) lemma le_add_iff2: "a * e + c \ b * e + d \ c \ (b - a) * e + d" by (simp add: algebra_simps) lemma mult_left_mono_neg: "b \ a \ c \ 0 \ c * a \ c * b" by (auto dest: mult_left_mono [of _ _ "- c"]) lemma mult_right_mono_neg: "b \ a \ c \ 0 \ a * c \ b * c" by (auto dest: mult_right_mono [of _ _ "- c"]) lemma mult_nonpos_nonpos: "a \ 0 \ b \ 0 \ 0 \ a * b" using mult_right_mono_neg [of a 0 b] by simp lemma split_mult_pos_le: "(0 \ a \ 0 \ b) \ (a \ 0 \ b \ 0) \ 0 \ a * b" by (auto simp add: mult_nonpos_nonpos) end class abs_if = minus + uminus + ord + zero + abs + assumes abs_if: "\a\ = (if a < 0 then - a else a)" class linordered_ring = ring + linordered_semiring + linordered_ab_group_add + abs_if begin subclass ordered_ring .. subclass ordered_ab_group_add_abs proof fix a b show "\a + b\ \ \a\ + \b\" by (auto simp add: abs_if not_le not_less algebra_simps simp del: add.commute dest: add_neg_neg add_nonneg_nonneg) qed (auto simp: abs_if) lemma zero_le_square [simp]: "0 \ a * a" using linear [of 0 a] by (auto simp add: mult_nonpos_nonpos) lemma not_square_less_zero [simp]: "\ (a * a < 0)" by (simp add: not_less) proposition abs_eq_iff: "\x\ = \y\ \ x = y \ x = -y" by (auto simp add: abs_if split: if_split_asm) lemma abs_eq_iff': "\a\ = b \ b \ 0 \ (a = b \ a = - b)" by (cases "a \ 0") auto lemma eq_abs_iff': "a = \b\ \ a \ 0 \ (b = a \ b = - a)" using abs_eq_iff' [of b a] by auto lemma sum_squares_ge_zero: "0 \ x * x + y * y" by (intro add_nonneg_nonneg zero_le_square) lemma not_sum_squares_lt_zero: "\ x * x + y * y < 0" by (simp add: not_less sum_squares_ge_zero) end class linordered_ring_strict = ring + linordered_semiring_strict + ordered_ab_group_add + abs_if begin subclass linordered_ring .. lemma mult_strict_left_mono_neg: "b < a \ c < 0 \ c * a < c * b" using mult_strict_left_mono [of b a "- c"] by simp lemma mult_strict_right_mono_neg: "b < a \ c < 0 \ a * c < b * c" using mult_strict_right_mono [of b a "- c"] by simp lemma mult_neg_neg: "a < 0 \ b < 0 \ 0 < a * b" using mult_strict_right_mono_neg [of a 0 b] by simp subclass ring_no_zero_divisors proof fix a b assume "a \ 0" then have a: "a < 0 \ 0 < a" by (simp add: neq_iff) assume "b \ 0" then have b: "b < 0 \ 0 < b" by (simp add: neq_iff) have "a * b < 0 \ 0 < a * b" proof (cases "a < 0") case True show ?thesis proof (cases "b < 0") case True with \a < 0\ show ?thesis by (auto dest: mult_neg_neg) next case False with b have "0 < b" by auto with \a < 0\ show ?thesis by (auto dest: mult_strict_right_mono) qed next case False with a have "0 < a" by auto show ?thesis proof (cases "b < 0") case True with \0 < a\ show ?thesis by (auto dest: mult_strict_right_mono_neg) next case False with b have "0 < b" by auto with \0 < a\ show ?thesis by auto qed qed then show "a * b \ 0" by (simp add: neq_iff) qed lemma zero_less_mult_iff [algebra_split_simps, field_split_simps]: "0 < a * b \ 0 < a \ 0 < b \ a < 0 \ b < 0" by (cases a 0 b 0 rule: linorder_cases[case_product linorder_cases]) (auto simp add: mult_neg_neg not_less le_less dest: zero_less_mult_pos zero_less_mult_pos2) lemma zero_le_mult_iff [algebra_split_simps, field_split_simps]: "0 \ a * b \ 0 \ a \ 0 \ b \ a \ 0 \ b \ 0" by (auto simp add: eq_commute [of 0] le_less not_less zero_less_mult_iff) lemma mult_less_0_iff [algebra_split_simps, field_split_simps]: "a * b < 0 \ 0 < a \ b < 0 \ a < 0 \ 0 < b" using zero_less_mult_iff [of "- a" b] by auto lemma mult_le_0_iff [algebra_split_simps, field_split_simps]: "a * b \ 0 \ 0 \ a \ b \ 0 \ a \ 0 \ 0 \ b" using zero_le_mult_iff [of "- a" b] by auto text \ Cancellation laws for \<^term>\c * a < c * b\ and \<^term>\a * c < b * c\, also with the relations \\\ and equality. \ text \ These ``disjunction'' versions produce two cases when the comparison is an assumption, but effectively four when the comparison is a goal. \ lemma mult_less_cancel_right_disj: "a * c < b * c \ 0 < c \ a < b \ c < 0 \ b < a" proof (cases "c = 0") case False show ?thesis (is "?lhs \ ?rhs") proof assume ?lhs then have "c < 0 \ b < a" "c > 0 \ b > a" by (auto simp flip: not_le intro: mult_right_mono mult_right_mono_neg) with False show ?rhs by (auto simp add: neq_iff) next assume ?rhs with False show ?lhs by (auto simp add: mult_strict_right_mono mult_strict_right_mono_neg) qed qed auto lemma mult_less_cancel_left_disj: "c * a < c * b \ 0 < c \ a < b \ c < 0 \ b < a" proof (cases "c = 0") case False show ?thesis (is "?lhs \ ?rhs") proof assume ?lhs then have "c < 0 \ b < a" "c > 0 \ b > a" by (auto simp flip: not_le intro: mult_left_mono mult_left_mono_neg) with False show ?rhs by (auto simp add: neq_iff) next assume ?rhs with False show ?lhs by (auto simp add: mult_strict_left_mono mult_strict_left_mono_neg) qed qed auto text \ The ``conjunction of implication'' lemmas produce two cases when the comparison is a goal, but give four when the comparison is an assumption. \ lemma mult_less_cancel_right: "a * c < b * c \ (0 \ c \ a < b) \ (c \ 0 \ b < a)" using mult_less_cancel_right_disj [of a c b] by auto lemma mult_less_cancel_left: "c * a < c * b \ (0 \ c \ a < b) \ (c \ 0 \ b < a)" using mult_less_cancel_left_disj [of c a b] by auto lemma mult_le_cancel_right: "a * c \ b * c \ (0 < c \ a \ b) \ (c < 0 \ b \ a)" by (simp add: not_less [symmetric] mult_less_cancel_right_disj) lemma mult_le_cancel_left: "c * a \ c * b \ (0 < c \ a \ b) \ (c < 0 \ b \ a)" by (simp add: not_less [symmetric] mult_less_cancel_left_disj) lemma mult_le_cancel_left_pos: "0 < c \ c * a \ c * b \ a \ b" by (auto simp: mult_le_cancel_left) lemma mult_le_cancel_left_neg: "c < 0 \ c * a \ c * b \ b \ a" by (auto simp: mult_le_cancel_left) lemma mult_less_cancel_left_pos: "0 < c \ c * a < c * b \ a < b" by (auto simp: mult_less_cancel_left) lemma mult_less_cancel_left_neg: "c < 0 \ c * a < c * b \ b < a" by (auto simp: mult_less_cancel_left) end lemmas mult_sign_intros = mult_nonneg_nonneg mult_nonneg_nonpos mult_nonpos_nonneg mult_nonpos_nonpos mult_pos_pos mult_pos_neg mult_neg_pos mult_neg_neg class ordered_comm_ring = comm_ring + ordered_comm_semiring begin subclass ordered_ring .. subclass ordered_cancel_comm_semiring .. end class linordered_nonzero_semiring = ordered_comm_semiring + monoid_mult + linorder + zero_less_one + assumes add_mono1: "a < b \ a + 1 < b + 1" begin subclass zero_neq_one by standard (insert zero_less_one, blast) subclass comm_semiring_1 by standard (rule mult_1_left) lemma zero_le_one [simp]: "0 \ 1" by (rule zero_less_one [THEN less_imp_le]) lemma not_one_le_zero [simp]: "\ 1 \ 0" by (simp add: not_le) lemma not_one_less_zero [simp]: "\ 1 < 0" by (simp add: not_less) lemma of_bool_less_eq_iff [simp]: \of_bool P \ of_bool Q \ (P \ Q)\ by auto lemma of_bool_less_iff [simp]: \of_bool P < of_bool Q \ \ P \ Q\ by auto lemma mult_left_le: "c \ 1 \ 0 \ a \ a * c \ a" using mult_left_mono[of c 1 a] by simp lemma mult_le_one: "a \ 1 \ 0 \ b \ b \ 1 \ a * b \ 1" using mult_mono[of a 1 b 1] by simp lemma zero_less_two: "0 < 1 + 1" using add_pos_pos[OF zero_less_one zero_less_one] . end class linordered_semidom = semidom + linordered_comm_semiring_strict + zero_less_one + assumes le_add_diff_inverse2 [simp]: "b \ a \ a - b + b = a" begin subclass linordered_nonzero_semiring proof show "a + 1 < b + 1" if "a < b" for a b proof (rule ccontr, simp add: not_less) assume "b \ a" with that show False by (simp add: ) qed qed lemma zero_less_eq_of_bool [simp]: \0 \ of_bool P\ by simp lemma zero_less_of_bool_iff [simp]: \0 < of_bool P \ P\ by simp lemma of_bool_less_eq_one [simp]: \of_bool P \ 1\ by simp lemma of_bool_less_one_iff [simp]: \of_bool P < 1 \ \ P\ by simp lemma of_bool_or_iff [simp]: \of_bool (P \ Q) = max (of_bool P) (of_bool Q)\ by (simp add: max_def) text \Addition is the inverse of subtraction.\ lemma le_add_diff_inverse [simp]: "b \ a \ b + (a - b) = a" by (frule le_add_diff_inverse2) (simp add: add.commute) lemma add_diff_inverse: "\ a < b \ b + (a - b) = a" by simp lemma add_le_imp_le_diff: assumes "i + k \ n" shows "i \ n - k" proof - have "n - (i + k) + i + k = n" by (simp add: assms add.assoc) with assms add_implies_diff have "i + k \ n - k + k" by fastforce then show ?thesis by simp qed lemma add_le_add_imp_diff_le: assumes 1: "i + k \ n" and 2: "n \ j + k" shows "i + k \ n \ n \ j + k \ n - k \ j" proof - have "n - (i + k) + i + k = n" using 1 by (simp add: add.assoc) moreover have "n - k = n - k - i + i" using 1 by (simp add: add_le_imp_le_diff) ultimately show ?thesis using 2 add_le_imp_le_diff [of "n-k" k "j + k"] by (simp add: add.commute diff_diff_add) qed lemma less_1_mult: "1 < m \ 1 < n \ 1 < m * n" using mult_strict_mono [of 1 m 1 n] by (simp add: less_trans [OF zero_less_one]) end class linordered_idom = comm_ring_1 + linordered_comm_semiring_strict + ordered_ab_group_add + abs_if + sgn + assumes sgn_if: "sgn x = (if x = 0 then 0 else if 0 < x then 1 else - 1)" begin subclass linordered_ring_strict .. subclass linordered_semiring_1_strict proof have "0 \ 1 * 1" by (fact zero_le_square) then show "0 < 1" by (simp add: le_less) qed subclass ordered_comm_ring .. subclass idom .. subclass linordered_semidom by standard simp subclass idom_abs_sgn by standard (auto simp add: sgn_if abs_if zero_less_mult_iff) lemma abs_bool_eq [simp]: \\of_bool P\ = of_bool P\ by simp lemma linorder_neqE_linordered_idom: assumes "x \ y" obtains "x < y" | "y < x" using assms by (rule neqE) text \These cancellation simp rules also produce two cases when the comparison is a goal.\ lemma mult_le_cancel_right1: "c \ b * c \ (0 < c \ 1 \ b) \ (c < 0 \ b \ 1)" using mult_le_cancel_right [of 1 c b] by simp lemma mult_le_cancel_right2: "a * c \ c \ (0 < c \ a \ 1) \ (c < 0 \ 1 \ a)" using mult_le_cancel_right [of a c 1] by simp lemma mult_le_cancel_left1: "c \ c * b \ (0 < c \ 1 \ b) \ (c < 0 \ b \ 1)" using mult_le_cancel_left [of c 1 b] by simp lemma mult_le_cancel_left2: "c * a \ c \ (0 < c \ a \ 1) \ (c < 0 \ 1 \ a)" using mult_le_cancel_left [of c a 1] by simp lemma mult_less_cancel_right1: "c < b * c \ (0 \ c \ 1 < b) \ (c \ 0 \ b < 1)" using mult_less_cancel_right [of 1 c b] by simp lemma mult_less_cancel_right2: "a * c < c \ (0 \ c \ a < 1) \ (c \ 0 \ 1 < a)" using mult_less_cancel_right [of a c 1] by simp lemma mult_less_cancel_left1: "c < c * b \ (0 \ c \ 1 < b) \ (c \ 0 \ b < 1)" using mult_less_cancel_left [of c 1 b] by simp lemma mult_less_cancel_left2: "c * a < c \ (0 \ c \ a < 1) \ (c \ 0 \ 1 < a)" using mult_less_cancel_left [of c a 1] by simp lemma sgn_0_0: "sgn a = 0 \ a = 0" by (fact sgn_eq_0_iff) lemma sgn_1_pos: "sgn a = 1 \ a > 0" unfolding sgn_if by simp lemma sgn_1_neg: "sgn a = - 1 \ a < 0" unfolding sgn_if by auto lemma sgn_pos [simp]: "0 < a \ sgn a = 1" by (simp only: sgn_1_pos) lemma sgn_neg [simp]: "a < 0 \ sgn a = - 1" by (simp only: sgn_1_neg) lemma abs_sgn: "\k\ = k * sgn k" unfolding sgn_if abs_if by auto lemma sgn_greater [simp]: "0 < sgn a \ 0 < a" unfolding sgn_if by auto lemma sgn_less [simp]: "sgn a < 0 \ a < 0" unfolding sgn_if by auto lemma abs_sgn_eq_1 [simp]: "a \ 0 \ \sgn a\ = 1" by simp lemma abs_sgn_eq: "\sgn a\ = (if a = 0 then 0 else 1)" by (simp add: sgn_if) lemma sgn_mult_self_eq [simp]: "sgn a * sgn a = of_bool (a \ 0)" by (cases "a > 0") simp_all lemma abs_mult_self_eq [simp]: "\a\ * \a\ = a * a" by (cases "a > 0") simp_all lemma same_sgn_sgn_add: "sgn (a + b) = sgn a" if "sgn b = sgn a" proof (cases a 0 rule: linorder_cases) case equal with that show ?thesis by simp next case less with that have "b < 0" by (simp add: sgn_1_neg) with \a < 0\ have "a + b < 0" by (rule add_neg_neg) with \a < 0\ show ?thesis by simp next case greater with that have "b > 0" by (simp add: sgn_1_pos) with \a > 0\ have "a + b > 0" by (rule add_pos_pos) with \a > 0\ show ?thesis by simp qed lemma same_sgn_abs_add: "\a + b\ = \a\ + \b\" if "sgn b = sgn a" proof - have "a + b = sgn a * \a\ + sgn b * \b\" by (simp add: sgn_mult_abs) also have "\ = sgn a * (\a\ + \b\)" using that by (simp add: algebra_simps) finally show ?thesis by (auto simp add: abs_mult) qed lemma sgn_not_eq_imp: "sgn a = - sgn b" if "sgn b \ sgn a" and "sgn a \ 0" and "sgn b \ 0" using that by (cases "a < 0") (auto simp add: sgn_0_0 sgn_1_pos sgn_1_neg) lemma abs_dvd_iff [simp]: "\m\ dvd k \ m dvd k" by (simp add: abs_if) lemma dvd_abs_iff [simp]: "m dvd \k\ \ m dvd k" by (simp add: abs_if) lemma dvd_if_abs_eq: "\l\ = \k\ \ l dvd k" by (subst abs_dvd_iff [symmetric]) simp text \ The following lemmas can be proven in more general structures, but are dangerous as simp rules in absence of @{thm neg_equal_zero}, @{thm neg_less_pos}, @{thm neg_less_eq_nonneg}. \ lemma equation_minus_iff_1 [simp, no_atp]: "1 = - a \ a = - 1" by (fact equation_minus_iff) lemma minus_equation_iff_1 [simp, no_atp]: "- a = 1 \ a = - 1" by (subst minus_equation_iff, auto) lemma le_minus_iff_1 [simp, no_atp]: "1 \ - b \ b \ - 1" by (fact le_minus_iff) lemma minus_le_iff_1 [simp, no_atp]: "- a \ 1 \ - 1 \ a" by (fact minus_le_iff) lemma less_minus_iff_1 [simp, no_atp]: "1 < - b \ b < - 1" by (fact less_minus_iff) lemma minus_less_iff_1 [simp, no_atp]: "- a < 1 \ - 1 < a" by (fact minus_less_iff) lemma add_less_zeroD: shows "x+y < 0 \ x<0 \ y<0" by (auto simp: not_less intro: le_less_trans [of _ "x+y"]) end text \Reasoning about inequalities with division\ context linordered_semidom begin lemma less_add_one: "a < a + 1" proof - have "a + 0 < a + 1" by (blast intro: zero_less_one add_strict_left_mono) then show ?thesis by simp qed end context linordered_idom begin lemma mult_right_le_one_le: "0 \ x \ 0 \ y \ y \ 1 \ x * y \ x" by (rule mult_left_le) lemma mult_left_le_one_le: "0 \ x \ 0 \ y \ y \ 1 \ y * x \ x" by (auto simp add: mult_le_cancel_right2) end text \Absolute Value\ context linordered_idom begin lemma mult_sgn_abs: "sgn x * \x\ = x" by (fact sgn_mult_abs) lemma abs_one: "\1\ = 1" by (fact abs_1) end class ordered_ring_abs = ordered_ring + ordered_ab_group_add_abs + assumes abs_eq_mult: "(0 \ a \ a \ 0) \ (0 \ b \ b \ 0) \ \a * b\ = \a\ * \b\" context linordered_idom begin subclass ordered_ring_abs by standard (auto simp: abs_if not_less mult_less_0_iff) lemma abs_mult_self: "\a\ * \a\ = a * a" by (fact abs_mult_self_eq) lemma abs_mult_less: assumes ac: "\a\ < c" and bd: "\b\ < d" shows "\a\ * \b\ < c * d" proof - from ac have "0 < c" by (blast intro: le_less_trans abs_ge_zero) with bd show ?thesis by (simp add: ac mult_strict_mono) qed lemma abs_less_iff: "\a\ < b \ a < b \ - a < b" by (simp add: less_le abs_le_iff) (auto simp add: abs_if) lemma abs_mult_pos: "0 \ x \ \y\ * x = \y * x\" by (simp add: abs_mult) lemma abs_diff_less_iff: "\x - a\ < r \ a - r < x \ x < a + r" by (auto simp add: diff_less_eq ac_simps abs_less_iff) lemma abs_diff_le_iff: "\x - a\ \ r \ a - r \ x \ x \ a + r" by (auto simp add: diff_le_eq ac_simps abs_le_iff) lemma abs_add_one_gt_zero: "0 < 1 + \x\" by (auto simp: abs_if not_less intro: zero_less_one add_strict_increasing less_trans) end subsection \Dioids\ text \ Dioids are the alternative extensions of semirings, a semiring can either be a ring or a dioid but never both. \ class dioid = semiring_1 + canonically_ordered_monoid_add begin subclass ordered_semiring by standard (auto simp: le_iff_add distrib_left distrib_right) end hide_fact (open) comm_mult_left_mono comm_mult_strict_left_mono distrib code_identifier code_module Rings \ (SML) Arith and (OCaml) Arith and (Haskell) Arith end